diff --git a/Mathlib.lean b/Mathlib.lean index 03b67d0a6f8e58..60a6904ec63fee 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1962,6 +1962,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Grothendieck import Mathlib.CategoryTheory.Limits.Preserves.Limits import Mathlib.CategoryTheory.Limits.Preserves.Opposites import Mathlib.CategoryTheory.Limits.Preserves.Presheaf +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.AbelianImages import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Biproducts import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Equalizers @@ -2000,6 +2001,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Equalizers import Mathlib.CategoryTheory.Limits.Shapes.PiProd import Mathlib.CategoryTheory.Limits.Shapes.Preorder.Basic import Mathlib.CategoryTheory.Limits.Shapes.Preorder.HasIterationOfShape +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.PrincipalSeg import Mathlib.CategoryTheory.Limits.Shapes.Preorder.WellOrderContinuous import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Assoc @@ -3183,7 +3185,9 @@ import Mathlib.FieldTheory.Perfect import Mathlib.FieldTheory.PerfectClosure import Mathlib.FieldTheory.PolynomialGaloisGroup import Mathlib.FieldTheory.PrimitiveElement -import Mathlib.FieldTheory.PurelyInseparable +import Mathlib.FieldTheory.PurelyInseparable.Basic +import Mathlib.FieldTheory.PurelyInseparable.PerfectClosure +import Mathlib.FieldTheory.PurelyInseparable.Tower import Mathlib.FieldTheory.RatFunc.AsPolynomial import Mathlib.FieldTheory.RatFunc.Basic import Mathlib.FieldTheory.RatFunc.Defs @@ -4245,6 +4249,7 @@ import Mathlib.Order.Filter.Ultrafilter.Defs import Mathlib.Order.Filter.ZeroAndBoundedAtFilter import Mathlib.Order.Fin.Basic import Mathlib.Order.Fin.Finset +import Mathlib.Order.Fin.SuccAboveOrderIso import Mathlib.Order.Fin.Tuple import Mathlib.Order.FixedPoints import Mathlib.Order.GaloisConnection.Basic @@ -4276,6 +4281,7 @@ import Mathlib.Order.Interval.Set.Defs import Mathlib.Order.Interval.Set.Disjoint import Mathlib.Order.Interval.Set.Image import Mathlib.Order.Interval.Set.Infinite +import Mathlib.Order.Interval.Set.InitialSeg import Mathlib.Order.Interval.Set.IsoIoo import Mathlib.Order.Interval.Set.Monotone import Mathlib.Order.Interval.Set.OrdConnected @@ -4592,6 +4598,7 @@ import Mathlib.RingTheory.Ideal.Prime import Mathlib.RingTheory.Ideal.Prod import Mathlib.RingTheory.Ideal.Quotient.Basic import Mathlib.RingTheory.Ideal.Quotient.Defs +import Mathlib.RingTheory.Ideal.Quotient.Index import Mathlib.RingTheory.Ideal.Quotient.Nilpotent import Mathlib.RingTheory.Ideal.Quotient.Noetherian import Mathlib.RingTheory.Ideal.Quotient.Operations @@ -4961,6 +4968,7 @@ import Mathlib.Tactic.CategoryTheory.Bicategory.Datatypes import Mathlib.Tactic.CategoryTheory.Bicategory.Normalize import Mathlib.Tactic.CategoryTheory.Bicategory.PureCoherence import Mathlib.Tactic.CategoryTheory.BicategoryCoherence +import Mathlib.Tactic.CategoryTheory.CheckCompositions import Mathlib.Tactic.CategoryTheory.Coherence import Mathlib.Tactic.CategoryTheory.Coherence.Basic import Mathlib.Tactic.CategoryTheory.Coherence.Datatypes @@ -5280,6 +5288,7 @@ 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.Support import Mathlib.Topology.Algebra.Order.UpperLower import Mathlib.Topology.Algebra.Polynomial import Mathlib.Topology.Algebra.PontryaginDual diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 303fe02c038bc6..d4c16489c62cf8 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -9,7 +9,6 @@ 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 @@ -22,6 +21,8 @@ This file could usefully be split further. universe u v w u₁ v₁ +open Function + namespace Algebra variable {R : Type u} {A : Type w} @@ -243,6 +244,11 @@ instance (priority := 99) Semiring.toNatAlgebra : Algebra ℕ R where instance nat_algebra_subsingleton : Subsingleton (Algebra ℕ R) := ⟨fun P Q => by ext; simp⟩ +@[simp] +lemma algebraMap_comp_natCast (R A : Type*) [CommSemiring R] [Semiring A] [Algebra R A] : + algebraMap R A ∘ Nat.cast = Nat.cast := by + ext; simp + end Nat section Int @@ -268,59 +274,105 @@ variable {R} instance int_algebra_subsingleton : Subsingleton (Algebra ℤ R) := ⟨fun P Q => Algebra.algebra_ext P Q <| RingHom.congr_fun <| Subsingleton.elim _ _⟩ +@[simp] +lemma algebraMap_comp_intCast (R A : Type*) [CommRing R] [Ring A] [Algebra R A] : + algebraMap R A ∘ Int.cast = Int.cast := by + ext; simp + end Int -namespace NoZeroSMulDivisors +section FaithfulSMul -variable {R A : Type*} +instance (R : Type*) [NonAssocSemiring R] : FaithfulSMul R R := ⟨fun {r₁ r₂} h ↦ by simpa using h 1⟩ -open Algebra +variable (R A : Type*) [CommSemiring R] [Semiring A] -/-- If `algebraMap R A` is injective and `A` has no zero divisors, -`R`-multiples in `A` are zero only if one of the factors is zero. +lemma faithfulSMul_iff_injective_smul_one [Module R A] [IsScalarTower R A A] : + FaithfulSMul R A ↔ Injective (fun r : R ↦ r • (1 : A)) := by + refine ⟨fun ⟨h⟩ {r₁ r₂} hr ↦ h fun a ↦ ?_, fun h ↦ ⟨fun {r₁ r₂} hr ↦ h ?_⟩⟩ + · simp only at hr + rw [← one_mul a, ← smul_mul_assoc, ← smul_mul_assoc, hr] + · simpa using hr 1 -Cannot be an instance because there is no `Injective (algebraMap R A)` typeclass. --/ -theorem of_algebraMap_injective [CommSemiring R] [Semiring A] [Algebra R A] [NoZeroDivisors A] - (h : Function.Injective (algebraMap R A)) : NoZeroSMulDivisors R A := - ⟨fun hcx => (mul_eq_zero.mp ((smul_def _ _).symm.trans hcx)).imp_left - (map_eq_zero_iff (algebraMap R A) h).mp⟩ +variable [Algebra R A] + +lemma faithfulSMul_iff_algebraMap_injective : FaithfulSMul R A ↔ Injective (algebraMap R A) := by + rw [faithfulSMul_iff_injective_smul_one, Algebra.algebraMap_eq_smul_one'] + +variable [FaithfulSMul R A] -variable (R A) +namespace FaithfulSMul -theorem algebraMap_injective [CommRing R] [Ring A] [Nontrivial A] [Algebra R A] - [NoZeroSMulDivisors R A] : Function.Injective (algebraMap R A) := by - simpa only [algebraMap_eq_smul_one'] using smul_left_injective R one_ne_zero +lemma algebraMap_injective : Injective (algebraMap R A) := + (faithfulSMul_iff_algebraMap_injective R A).mp inferInstance + +@[deprecated (since := "2025-01-31")] +alias _root_.NoZeroSMulDivisors.algebraMap_injective := algebraMap_injective @[simp] -lemma algebraMap_eq_zero_iff [CommRing R] [Ring A] [Nontrivial A] [Algebra R A] - [NoZeroSMulDivisors R A] {r : R} : algebraMap R A r = 0 ↔ r = 0 := - map_eq_zero_iff _ <| algebraMap_injective R A +lemma algebraMap_eq_zero_iff {r : R} : algebraMap R A r = 0 ↔ r = 0 := + map_eq_zero_iff (algebraMap R A) <| algebraMap_injective R A + +@[deprecated (since := "2025-01-31")] +alias _root_.NoZeroSMulDivisors.algebraMap_eq_zero_iff := algebraMap_eq_zero_iff @[simp] -lemma algebraMap_eq_one_iff [CommRing R] [Ring A] [Nontrivial A] [Algebra R A] - [NoZeroSMulDivisors R A] {r : R} : algebraMap R A r = 1 ↔ r = 1 := - map_eq_one_iff _ <| algebraMap_injective R A +lemma algebraMap_eq_one_iff {r : R} : algebraMap R A r = 1 ↔ r = 1 := + map_eq_one_iff _ <| FaithfulSMul.algebraMap_injective R A + +@[deprecated (since := "2025-01-31")] +alias _root_.NoZeroSMulDivisors.algebraMap_eq_one_iff := algebraMap_eq_one_iff -theorem _root_.NeZero.of_noZeroSMulDivisors (n : ℕ) [CommRing R] [NeZero (n : R)] [Ring A] - [Nontrivial A] [Algebra R A] [NoZeroSMulDivisors R A] : NeZero (n : A) := - NeZero.nat_of_injective <| NoZeroSMulDivisors.algebraMap_injective R A +theorem _root_.NeZero.of_faithfulSMul (n : ℕ) [NeZero (n : R)] : + NeZero (n : A) := + NeZero.nat_of_injective <| FaithfulSMul.algebraMap_injective R A -variable {R A} +@[deprecated (since := "2025-01-31")] +alias _root_.NeZero.of_noZeroSMulDivisors := NeZero.of_faithfulSMul -theorem iff_algebraMap_injective [CommRing R] [Ring A] [IsDomain A] [Algebra R A] : - NoZeroSMulDivisors R A ↔ Function.Injective (algebraMap R A) := - ⟨@NoZeroSMulDivisors.algebraMap_injective R A _ _ _ _, NoZeroSMulDivisors.of_algebraMap_injective⟩ +end FaithfulSMul + +lemma Algebra.charZero_of_charZero [CharZero R] : CharZero A := + have := algebraMap_comp_natCast R A + ⟨this ▸ (FaithfulSMul.algebraMap_injective R A).comp CharZero.cast_injective⟩ -- see note [lower instance priority] -instance (priority := 100) CharZero.noZeroSMulDivisors_nat [Semiring R] [NoZeroDivisors R] - [CharZero R] : NoZeroSMulDivisors ℕ R := - NoZeroSMulDivisors.of_algebraMap_injective <| (algebraMap ℕ R).injective_nat +instance (priority := 100) [CharZero R] : FaithfulSMul ℕ R := by + simpa only [faithfulSMul_iff_algebraMap_injective] using (algebraMap ℕ R).injective_nat -- see note [lower instance priority] -instance (priority := 100) CharZero.noZeroSMulDivisors_int [Ring R] [NoZeroDivisors R] - [CharZero R] : NoZeroSMulDivisors ℤ R := - NoZeroSMulDivisors.of_algebraMap_injective <| (algebraMap ℤ R).injective_int +instance (priority := 100) (R : Type*) [Ring R] [CharZero R] : FaithfulSMul ℤ R := by + simpa only [faithfulSMul_iff_algebraMap_injective] using (algebraMap ℤ R).injective_int + +end FaithfulSMul + +namespace NoZeroSMulDivisors + +-- see Note [lower instance priority] +instance (priority := 100) instOfFaithfulSMul {R A : Type*} + [CommSemiring R] [Semiring A] [Algebra R A] [NoZeroDivisors A] [FaithfulSMul R A] : + NoZeroSMulDivisors R A := + ⟨fun hcx => (mul_eq_zero.mp ((Algebra.smul_def _ _).symm.trans hcx)).imp_left + (map_eq_zero_iff (algebraMap R A) <| FaithfulSMul.algebraMap_injective R A).mp⟩ + +@[deprecated (since := "2025-01-31")] +alias of_algebraMap_injective := instOfFaithfulSMul + +variable {R A : Type*} [CommRing R] [Ring A] [Algebra R A] + +instance [Nontrivial A] [NoZeroSMulDivisors R A] : FaithfulSMul R A where + eq_of_smul_eq_smul {r₁ r₂} h := by + specialize h 1 + rw [← sub_eq_zero, ← sub_smul, smul_eq_zero, sub_eq_zero] at h + exact h.resolve_right one_ne_zero + +theorem iff_faithfulSMul [IsDomain A] : NoZeroSMulDivisors R A ↔ FaithfulSMul R A := + ⟨fun _ ↦ inferInstance, fun _ ↦ inferInstance⟩ + +theorem iff_algebraMap_injective [IsDomain A] : + NoZeroSMulDivisors R A ↔ Injective (algebraMap R A) := by + rw [iff_faithfulSMul] + exact faithfulSMul_iff_algebraMap_injective R A end NoZeroSMulDivisors @@ -338,20 +390,16 @@ theorem algebraMap_smul (r : R) (m : M) : (algebraMap R A) r • m = r • m := (algebra_compatible_smul A r m).symm /-- If `M` is `A`-torsion free and `algebraMap R A` is injective, `M` is also `R`-torsion free. -/ -lemma NoZeroSMulDivisors.of_algebraMap_injective' {R A M : Type*} [CommSemiring R] [Semiring A] - [Algebra R A] [AddCommMonoid M] [Module R M] [Module A M] [IsScalarTower R A M] - [NoZeroSMulDivisors A M] (h : Function.Injective (algebraMap R A)) : - NoZeroSMulDivisors R M where +theorem NoZeroSMulDivisors.trans_faithfulSMul (R A M : Type*) [CommRing R] [Ring A] [Algebra R A] + [FaithfulSMul R A] [AddCommGroup M] [Module R M] [Module A M] [IsScalarTower R A M] + [NoZeroSMulDivisors A M] : NoZeroSMulDivisors R M where eq_zero_or_eq_zero_of_smul_eq_zero hx := by rw [← algebraMap_smul (A := A)] at hx - obtain (hc|hx) := eq_zero_or_eq_zero_of_smul_eq_zero hx - · exact Or.inl <| (map_eq_zero_iff _ h).mp hc - · exact Or.inr hx + simpa only [map_eq_zero_iff _ <| FaithfulSMul.algebraMap_injective R A] using + eq_zero_or_eq_zero_of_smul_eq_zero hx -theorem NoZeroSMulDivisors.trans (R A M : Type*) [CommRing R] [Ring A] [IsDomain A] [Algebra R A] - [AddCommGroup M] [Module R M] [Module A M] [IsScalarTower R A M] [NoZeroSMulDivisors R A] - [NoZeroSMulDivisors A M] : NoZeroSMulDivisors R M := - of_algebraMap_injective' (A := A) (NoZeroSMulDivisors.iff_algebraMap_injective.1 inferInstance) +@[deprecated (since := "2025-01-31")] +alias NoZeroSMulDivisors.of_algebraMap_injective' := NoZeroSMulDivisors.trans_faithfulSMul variable {A} @@ -443,15 +491,15 @@ variable {F E : Type*} [CommSemiring F] [Semiring E] [Algebra F E] (b : F →ₗ /-- If `E` is an `F`-algebra, and there exists an injective `F`-linear map from `F` to `E`, then the algebra map from `F` to `E` is also injective. -/ -theorem injective_algebraMap_of_linearMap (hb : Function.Injective b) : - Function.Injective (algebraMap F E) := fun x y e ↦ hb <| by +theorem injective_algebraMap_of_linearMap (hb : Injective b) : + Injective (algebraMap F E) := fun x y e ↦ hb <| by rw [← mul_one x, ← mul_one y, ← smul_eq_mul, ← smul_eq_mul, map_smul, map_smul, Algebra.smul_def, Algebra.smul_def, e] /-- If `E` is an `F`-algebra, and there exists a surjective `F`-linear map from `F` to `E`, then the algebra map from `F` to `E` is also surjective. -/ -theorem surjective_algebraMap_of_linearMap (hb : Function.Surjective b) : - Function.Surjective (algebraMap F E) := fun x ↦ by +theorem surjective_algebraMap_of_linearMap (hb : Surjective b) : + Surjective (algebraMap F E) := fun x ↦ by obtain ⟨x, rfl⟩ := hb x obtain ⟨y, hy⟩ := hb (b 1 * b 1) refine ⟨x * y, ?_⟩ @@ -467,14 +515,14 @@ then the algebra map from `F` to `E` is also bijective. NOTE: The same result can also be obtained if there are two `F`-linear maps from `F` to `E`, one is injective, the other one is surjective. In this case, use `injective_algebraMap_of_linearMap` and `surjective_algebraMap_of_linearMap` separately. -/ -theorem bijective_algebraMap_of_linearMap (hb : Function.Bijective b) : - Function.Bijective (algebraMap F E) := +theorem bijective_algebraMap_of_linearMap (hb : Bijective b) : + Bijective (algebraMap F E) := ⟨injective_algebraMap_of_linearMap b hb.1, surjective_algebraMap_of_linearMap b hb.2⟩ /-- If `E` is an `F`-algebra, there exists an `F`-linear isomorphism from `F` to `E` (namely, `E` is a free `F`-module of rank one), then the algebra map from `F` to `E` is bijective. -/ theorem bijective_algebraMap_of_linearEquiv (b : F ≃ₗ[F] E) : - Function.Bijective (algebraMap F E) := + Bijective (algebraMap F E) := bijective_algebraMap_of_linearMap _ b.bijective end algebraMap @@ -486,7 +534,7 @@ variable {M N} [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module S M] [Is variable [Module R N] [Module S N] [IsScalarTower R S N] /-- If `R →+* S` is surjective, then `S`-linear maps between modules are exactly `R`-linear maps. -/ -def LinearMap.extendScalarsOfSurjectiveEquiv (h : Function.Surjective (algebraMap R S)) : +def LinearMap.extendScalarsOfSurjectiveEquiv (h : Surjective (algebraMap R S)) : (M →ₗ[R] N) ≃ₗ[R] (M →ₗ[S] N) where toFun f := { __ := f, map_smul' := fun r x ↦ by obtain ⟨r, rfl⟩ := h r; simp } map_add' _ _ := rfl @@ -496,17 +544,17 @@ def LinearMap.extendScalarsOfSurjectiveEquiv (h : Function.Surjective (algebraMa right_inv _ := rfl /-- If `R →+* S` is surjective, then `R`-linear maps are also `S`-linear. -/ -abbrev LinearMap.extendScalarsOfSurjective (h : Function.Surjective (algebraMap R S)) +abbrev LinearMap.extendScalarsOfSurjective (h : Surjective (algebraMap R S)) (l : M →ₗ[R] N) : M →ₗ[S] N := extendScalarsOfSurjectiveEquiv h l /-- If `R →+* S` is surjective, then `R`-linear isomorphisms are also `S`-linear. -/ -def LinearEquiv.extendScalarsOfSurjective (h : Function.Surjective (algebraMap R S)) +def LinearEquiv.extendScalarsOfSurjective (h : Surjective (algebraMap R S)) (f : M ≃ₗ[R] N) : M ≃ₗ[S] N where __ := f map_smul' r x := by obtain ⟨r, rfl⟩ := h r; simp -variable (h : Function.Surjective (algebraMap R S)) +variable (h : Surjective (algebraMap R S)) @[simp] lemma LinearMap.extendScalarsOfSurjective_apply (l : M →ₗ[R] N) (x) : diff --git a/Mathlib/Algebra/AlgebraicCard.lean b/Mathlib/Algebra/AlgebraicCard.lean index c9fce882650c05..1d01bc5efa7294 100644 --- a/Mathlib/Algebra/AlgebraicCard.lean +++ b/Mathlib/Algebra/AlgebraicCard.lean @@ -66,7 +66,7 @@ theorem cardinalMk_lift_of_infinite [Infinite R] : Cardinal.lift.{u} #{ x : A // IsAlgebraic R x } = Cardinal.lift.{v} #R := ((cardinalMk_lift_le_max R A).trans_eq (max_eq_left <| aleph0_le_mk _)).antisymm <| lift_mk_le'.2 ⟨⟨fun x => ⟨algebraMap R A x, isAlgebraic_algebraMap _⟩, fun _ _ h => - NoZeroSMulDivisors.algebraMap_injective R A (Subtype.ext_iff.1 h)⟩⟩ + FaithfulSMul.algebraMap_injective R A (Subtype.ext_iff.1 h)⟩⟩ @[deprecated (since := "2024-11-10")] alias cardinal_mk_lift_of_infinite := cardinalMk_lift_of_infinite diff --git a/Mathlib/Algebra/BigOperators/Associated.lean b/Mathlib/Algebra/BigOperators/Associated.lean index f6754712f5eeab..2178d8c747072b 100644 --- a/Mathlib/Algebra/BigOperators/Associated.lean +++ b/Mathlib/Algebra/BigOperators/Associated.lean @@ -188,7 +188,7 @@ variable [CancelCommMonoidWithZero α] theorem exists_mem_multiset_le_of_prime {s : Multiset (Associates α)} {p : Associates α} (hp : Prime p) : p ≤ s.prod → ∃ a ∈ s, p ≤ a := - Multiset.induction_on s (fun ⟨_, Eq⟩ => (hp.ne_one (mul_eq_one.1 Eq.symm).1).elim) + Multiset.induction_on s (fun ⟨_, eq⟩ => (hp.ne_one (mul_eq_one.1 eq.symm).1).elim) fun a s ih h => have : p ≤ a * s.prod := by simpa using h match Prime.le_or_le hp this with diff --git a/Mathlib/Algebra/Category/Grp/Ulift.lean b/Mathlib/Algebra/Category/Grp/Ulift.lean index b2b0a285efc943..fb8a66b2919d1b 100644 --- a/Mathlib/Algebra/Category/Grp/Ulift.lean +++ b/Mathlib/Algebra/Category/Grp/Ulift.lean @@ -200,6 +200,7 @@ noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} whe The functor `uliftFunctor : AddCommGrp.{u} ⥤ AddCommGrp.{max u v}` creates `u`-small colimits. -/ noncomputable instance : CreatesColimitsOfSize.{w, u} uliftFunctor.{v, u} where - CreatesColimitsOfShape := { CreatesColimit := fun {_} ↦ createsColimitOfFullyFaithfulOfPreserves } + CreatesColimitsOfShape := + { CreatesColimit := fun {_} ↦ createsColimitOfReflectsIsomorphismsOfPreserves } end AddCommGrp diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean index 61228a85b17e0b..ddcce4f281924e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean @@ -34,7 +34,7 @@ def monoidalClosedHomEquiv (M N P : ModuleCat.{u} R) : apply TensorProduct.ext' intro m n simp only [Hom.hom₂_ofHom₂, LinearMap.comp_apply, hom_comp, MonoidalCategory.tensorLeft_obj] - erw [MonoidalCategory.braiding_hom_apply, TensorProduct.lift.tmul] + erw [MonoidalCategory.braiding_hom_apply m n, TensorProduct.lift.tmul] right_inv _ := rfl instance : MonoidalClosed (ModuleCat.{u} R) where diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index 1cd71c7594c0a9..1e74ff70e10139 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -32,8 +32,6 @@ namespace CategoryTheory namespace Presieve.FamilyOfElements -attribute [local instance] HasForget.hasCoeToSort HasForget.instFunLike - section smul variable {R : Cᵒᵖ ⥤ RingCat.{u}} {M : PresheafOfModules.{v} R} {X : C} {P : Presieve X} diff --git a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean index d4219c84e035f1..e1d6ad2dd46077 100644 --- a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean @@ -32,7 +32,7 @@ namespace MonCat @[to_additive (attr := simps) "The functor of adjoining a neutral element `zero` to a semigroup"] def adjoinOne : Semigrp.{u} ⥤ MonCat.{u} where obj S := MonCat.of (WithOne S) - map f := ofHom (WithOne.map f) + map f := ofHom (WithOne.map f.hom) map_id _ := MonCat.hom_ext WithOne.map_id map_comp _ _ := MonCat.hom_ext (WithOne.map_comp _ _) @@ -40,13 +40,15 @@ def adjoinOne : Semigrp.{u} ⥤ MonCat.{u} where instance hasForgetToSemigroup : HasForget₂ MonCat Semigrp where forget₂ := { obj := fun M => Semigrp.of M - map f := f.hom.toMulHom } + map f := Semigrp.ofHom f.hom.toMulHom } /-- The `adjoinOne`-forgetful adjunction from `Semigrp` to `MonCat`. -/ @[to_additive "The `adjoinZero`-forgetful adjunction from `AddSemigrp` to `AddMonCat`"] def adjoinOneAdj : adjoinOne ⊣ forget₂ MonCat.{u} Semigrp.{u} := Adjunction.mkOfHomEquiv - { homEquiv := fun _ _ => ConcreteCategory.homEquiv.trans WithOne.lift.symm + { homEquiv X Y := + ConcreteCategory.homEquiv.trans (WithOne.lift.symm.trans + (ConcreteCategory.homEquiv (X := X) (Y := (forget₂ _ _).obj Y)).symm) homEquiv_naturality_left_symm := by intros ext ⟨_|_⟩ <;> simp <;> rfl } diff --git a/Mathlib/Algebra/Category/Semigrp/Basic.lean b/Mathlib/Algebra/Category/Semigrp/Basic.lean index f587fda2666232..4032f9def4d0ca 100644 --- a/Mathlib/Algebra/Category/Semigrp/Basic.lean +++ b/Mathlib/Algebra/Category/Semigrp/Basic.lean @@ -5,7 +5,7 @@ Authors: Julian Kuelshammer -/ import Mathlib.Algebra.PEmptyInstances import Mathlib.Algebra.Group.Equiv.Defs -import Mathlib.CategoryTheory.ConcreteCategory.BundledHom +import Mathlib.CategoryTheory.ConcreteCategory.Basic import Mathlib.CategoryTheory.Functor.ReflectsIso /-! @@ -31,77 +31,163 @@ universe u v open CategoryTheory +/-- The category of additive magmas and additive magma morphisms. -/ +structure AddMagmaCat : Type (u + 1) where + /-- The underlying additive magma. -/ + (carrier : Type u) + [str : Add carrier] + /-- The category of magmas and magma morphisms. -/ @[to_additive] -def MagmaCat : Type (u + 1) := - Bundled Mul +structure MagmaCat : Type (u + 1) where + /-- The underlying magma. -/ + (carrier : Type u) + [str : Mul carrier] -/-- The category of additive magmas and additive magma morphisms. -/ -add_decl_doc AddMagmaCat +attribute [instance] AddMagmaCat.str MagmaCat.str +attribute [to_additive existing] MagmaCat.carrier MagmaCat.str + +initialize_simps_projections AddMagmaCat (carrier → coe, -str) +initialize_simps_projections MagmaCat (carrier → coe, -str) namespace MagmaCat @[to_additive] -instance bundledHom : BundledHom @MulHom := - ⟨@MulHom.toFun, @MulHom.id, @MulHom.comp, - by intros; apply @DFunLike.coe_injective, by aesop_cat, by simp⟩ +instance : CoeSort MagmaCat (Type u) := + ⟨MagmaCat.carrier⟩ --- Porting note: deriving failed for `HasForget`, --- "default handlers have not been implemented yet" --- https://github.com/leanprover-community/mathlib4/issues/5020 -deriving instance LargeCategory for MagmaCat -instance instHasForget : HasForget MagmaCat := BundledHom.hasForget MulHom +attribute [coe] AddMagmaCat.carrier MagmaCat.carrier -attribute [to_additive] instMagmaCatLargeCategory instHasForget +/-- Construct a bundled `MagmaCat` from the underlying type and typeclass. -/ +@[to_additive "Construct a bundled `AddMagmaCat` from the underlying type and typeclass."] +abbrev of (M : Type u) [Mul M] : MagmaCat := ⟨M⟩ -@[to_additive] -instance : CoeSort MagmaCat Type* where - coe X := X.α +end MagmaCat + +/-- The type of morphisms in `AddMagmaCat R`. -/ +@[ext] +structure AddMagmaCat.Hom (A B : AddMagmaCat.{u}) where + private mk :: + /-- The underlying `AddHom`. -/ + hom' : A →ₙ+ B + +/-- The type of morphisms in `MagmaCat R`. -/ +@[to_additive, ext] +structure MagmaCat.Hom (A B : MagmaCat.{u}) where + private mk :: + /-- The underlying `MulHom`. -/ + hom' : A →ₙ* B --- Porting note: Hinting to Lean that `forget R` and `R` are the same -unif_hint forget_obj_eq_coe (R : MagmaCat) where ⊢ - (forget MagmaCat).obj R ≟ R -unif_hint _root_.AddMagmaCat.forget_obj_eq_coe (R : AddMagmaCat) where ⊢ - (forget AddMagmaCat).obj R ≟ R +attribute [to_additive existing AddMagmaCat.Hom.mk] MagmaCat.Hom.mk + +namespace MagmaCat @[to_additive] -instance (X : MagmaCat) : Mul X := X.str +instance : Category MagmaCat.{u} where + Hom X Y := Hom X Y + id X := ⟨MulHom.id X⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ + +@[to_additive] +instance : ConcreteCategory MagmaCat (· →ₙ* ·) where + hom := Hom.hom' + ofHom := Hom.mk + +/-- Turn a morphism in `MagmaCat` back into a `MulHom`. -/ +@[to_additive "Turn a morphism in `AddMagmaCat` back into an `AddHom`."] +abbrev Hom.hom {X Y : MagmaCat.{u}} (f : Hom X Y) := + ConcreteCategory.hom (C := MagmaCat) f + +/-- Typecheck a `MulHom` as a morphism in `MagmaCat`. -/ +@[to_additive "Typecheck an `AddHom` as a morphism in `AddMagmaCat`. "] +abbrev ofHom {X Y : Type u} [Mul X] [Mul Y] (f : X →ₙ* Y) : of X ⟶ of Y := + ConcreteCategory.ofHom (C := MagmaCat) f + +variable {R} in +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (X Y : MagmaCat.{u}) (f : Hom X Y) := + f.hom + +initialize_simps_projections Hom (hom' → hom) +initialize_simps_projections AddMagmaCat.Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ + +@[to_additive (attr := simp)] +lemma coe_id {X : MagmaCat} : (𝟙 X : X → X) = id := rfl + +@[to_additive (attr := simp)] +lemma coe_comp {X Y Z : MagmaCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl + +@[to_additive (attr := simp)] +lemma forget_map {X Y : MagmaCat} (f : X ⟶ Y) : + (forget MagmaCat).map f = f := rfl + +@[to_additive (attr := ext)] +lemma ext {X Y : MagmaCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := + ConcreteCategory.hom_ext _ _ w @[to_additive] -instance instFunLike (X Y : MagmaCat) : FunLike (X ⟶ Y) X Y := - inferInstanceAs <| FunLike (X →ₙ* Y) X Y +-- This is not `simp` to avoid rewriting in types of terms. +theorem coe_of (M : Type u) [Mul M] : (MagmaCat.of M : Type u) = M := rfl + +@[to_additive (attr := simp)] +lemma hom_id {M : MagmaCat} : (𝟙 M : M ⟶ M).hom = MulHom.id M := rfl +/- Provided for rewriting. -/ @[to_additive] -instance instMulHomClass (X Y : MagmaCat) : MulHomClass (X ⟶ Y) X Y := - inferInstanceAs <| MulHomClass (X →ₙ* Y) X Y +lemma id_apply (M : MagmaCat) (x : M) : + (𝟙 M : M ⟶ M) x = x := by simp -/-- Construct a bundled `MagmaCat` from the underlying type and typeclass. -/ +@[to_additive (attr := simp)] +lemma hom_comp {M N T : MagmaCat} (f : M ⟶ N) (g : N ⟶ T) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + +/- Provided for rewriting. -/ @[to_additive] -def of (M : Type u) [Mul M] : MagmaCat := - Bundled.of M +lemma comp_apply {M N T : MagmaCat} (f : M ⟶ N) (g : N ⟶ T) (x : M) : + (f ≫ g) x = g (f x) := by simp -/-- Construct a bundled `AddMagmaCat` from the underlying type and typeclass. -/ -add_decl_doc AddMagmaCat.of +@[to_additive (attr := ext)] +lemma hom_ext {M N : MagmaCat} {f g : M ⟶ N} (hf : f.hom = g.hom) : f = g := + Hom.ext hf @[to_additive (attr := simp)] -theorem coe_of (R : Type u) [Mul R] : (MagmaCat.of R : Type u) = R := - rfl +lemma hom_ofHom {M N : Type u} [Mul M] [Mul N] (f : M →ₙ* N) : + (ofHom f).hom = f := rfl @[to_additive (attr := simp)] -lemma mulEquiv_coe_eq {X Y : Type _} [Mul X] [Mul Y] (e : X ≃* Y) : - (@DFunLike.coe (MagmaCat.of X ⟶ MagmaCat.of Y) _ (fun _ => (forget MagmaCat).obj _) - HasForget.instFunLike (e : X →ₙ* Y) : X → Y) = ↑e := +lemma ofHom_hom {M N : MagmaCat} (f : M ⟶ N) : + ofHom (Hom.hom f) = f := rfl + +@[to_additive (attr := simp)] +lemma ofHom_id {M : Type u} [Mul M] : ofHom (MulHom.id M) = 𝟙 (of M) := rfl + +@[to_additive (attr := simp)] +lemma ofHom_comp {M N P : Type u} [Mul M] [Mul N] [Mul P] + (f : M →ₙ* N) (g : N →ₙ* P) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := rfl -/-- Typecheck a `MulHom` as a morphism in `MagmaCat`. -/ @[to_additive] -def ofHom {X Y : Type u} [Mul X] [Mul Y] (f : X →ₙ* Y) : of X ⟶ of Y := f +lemma ofHom_apply {X Y : Type u} [Mul X] [Mul Y] (f : X →ₙ* Y) (x : X) : + (ofHom f) x = f x := rfl -/-- Typecheck an `AddHom` as a morphism in `AddMagmaCat`. -/ -add_decl_doc AddMagmaCat.ofHom +@[to_additive (attr := simp)] +lemma inv_hom_apply {M N : MagmaCat} (e : M ≅ N) (x : M) : e.inv (e.hom x) = x := by + rw [← comp_apply] + simp -@[to_additive] -theorem ofHom_apply {X Y : Type u} [Mul X] [Mul Y] (f : X →ₙ* Y) (x : X) : ofHom f x = f x := +@[to_additive (attr := simp)] +lemma hom_inv_apply {M N : MagmaCat} (e : M ≅ N) (s : N) : e.hom (e.inv s) = s := by + rw [← comp_apply] + simp + +@[to_additive (attr := simp)] +lemma mulEquiv_coe_eq {X Y : Type _} [Mul X] [Mul Y] (e : X ≃* Y) : + (ofHom (e : X →ₙ* Y)).hom = ↑e := rfl @[to_additive] @@ -110,79 +196,164 @@ instance : Inhabited MagmaCat := end MagmaCat +/-- The category of additive semigroups and semigroup morphisms. -/ +structure AddSemigrp : Type (u + 1) where + /-- The underlying type. -/ + (carrier : Type u) + [str : AddSemigroup carrier] + /-- The category of semigroups and semigroup morphisms. -/ @[to_additive] -def Semigrp : Type (u + 1) := - Bundled Semigroup +structure Semigrp : Type (u + 1) where + /-- The underlying type. -/ + (carrier : Type u) + [str : Semigroup carrier] -/-- The category of additive semigroups and semigroup morphisms. -/ -add_decl_doc AddSemigrp +attribute [instance] AddSemigrp.str Semigrp.str +attribute [to_additive existing] Semigrp.carrier Semigrp.str + +initialize_simps_projections AddSemigrp (carrier → coe, -str) +initialize_simps_projections Semigrp (carrier → coe, -str) namespace Semigrp @[to_additive] -instance : BundledHom.ParentProjection @Semigroup.toMul := ⟨⟩ +instance : CoeSort Semigrp (Type u) := + ⟨Semigrp.carrier⟩ + +attribute [coe] AddSemigrp.carrier Semigrp.carrier + +/-- Construct a bundled `Semigrp` from the underlying type and typeclass. -/ +@[to_additive "Construct a bundled `AddSemigrp` from the underlying type and typeclass."] +abbrev of (M : Type u) [Semigroup M] : Semigrp := ⟨M⟩ -deriving instance LargeCategory for Semigrp +end Semigrp --- Porting note: deriving failed for `HasForget`, --- "default handlers have not been implemented yet" --- https://github.com/leanprover-community/mathlib4/issues/5020 -instance instHasForget : HasForget Semigrp := - BundledHom.hasForget (fun _ _ => _) +/-- The type of morphisms in `AddSemigrp R`. -/ +@[ext] +structure AddSemigrp.Hom (A B : AddSemigrp.{u}) where + private mk :: + /-- The underlying `AddHom`. -/ + hom' : A →ₙ+ B -attribute [to_additive] instSemigrpLargeCategory Semigrp.instHasForget +/-- The type of morphisms in `Semigrp R`. -/ +@[to_additive, ext] +structure Semigrp.Hom (A B : Semigrp.{u}) where + private mk :: + /-- The underlying `MulHom`. -/ + hom' : A →ₙ* B -@[to_additive] -instance : CoeSort Semigrp Type* where - coe X := X.α +attribute [to_additive existing AddSemigrp.Hom.mk] Semigrp.Hom.mk --- Porting note: Hinting to Lean that `forget R` and `R` are the same -unif_hint forget_obj_eq_coe (R : Semigrp) where ⊢ - (forget Semigrp).obj R ≟ R -unif_hint _root_.AddSemigrp.forget_obj_eq_coe (R : AddSemigrp) where ⊢ - (forget AddSemigrp).obj R ≟ R +namespace Semigrp + +@[to_additive] +instance : Category Semigrp.{u} where + Hom X Y := Hom X Y + id X := ⟨MulHom.id X⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ @[to_additive] -instance (X : Semigrp) : Semigroup X := X.str +instance : ConcreteCategory Semigrp (· →ₙ* ·) where + hom := Hom.hom' + ofHom := Hom.mk + +/-- Turn a morphism in `Semigrp` back into a `MulHom`. -/ +@[to_additive "Turn a morphism in `AddSemigrp` back into an `AddHom`."] +abbrev Hom.hom {X Y : Semigrp.{u}} (f : Hom X Y) := + ConcreteCategory.hom (C := Semigrp) f + +/-- Typecheck a `MulHom` as a morphism in `Semigrp`. -/ +@[to_additive "Typecheck an `AddHom` as a morphism in `AddSemigrp`. "] +abbrev ofHom {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) : of X ⟶ of Y := + ConcreteCategory.ofHom (C := Semigrp) f + +variable {R} in +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (X Y : Semigrp.{u}) (f : Hom X Y) := + f.hom + +initialize_simps_projections Hom (hom' → hom) +initialize_simps_projections AddSemigrp.Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ + +@[to_additive (attr := simp)] +lemma coe_id {X : Semigrp} : (𝟙 X : X → X) = id := rfl + +@[to_additive (attr := simp)] +lemma coe_comp {X Y Z : Semigrp} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl + +@[to_additive (attr := deprecated "Use hom_comp instead" (since := "2025-01-28"))] +lemma comp_def {X Y Z : Semigrp} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g).hom = g.hom.comp f.hom := rfl + +@[simp] lemma forget_map {X Y : Semigrp} (f : X ⟶ Y) : (forget Semigrp).map f = (f : X → Y) := rfl + +@[to_additive (attr := ext)] +lemma ext {X Y : Semigrp} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := + ConcreteCategory.hom_ext _ _ w @[to_additive] -instance instFunLike (X Y : Semigrp) : FunLike (X ⟶ Y) X Y := - inferInstanceAs <| FunLike (X →ₙ* Y) X Y +-- This is not `simp` to avoid rewriting in types of terms. +theorem coe_of (R : Type u) [Semigroup R] : ↑(Semigrp.of R) = R := + rfl + +@[to_additive (attr := simp)] +lemma hom_id {X : Semigrp} : (𝟙 X : X ⟶ X).hom = MulHom.id X := rfl +/- Provided for rewriting. -/ @[to_additive] -instance instMulHomClass (X Y : Semigrp) : MulHomClass (X ⟶ Y) X Y := - inferInstanceAs <| MulHomClass (X →ₙ* Y) X Y +lemma id_apply (X : Semigrp) (x : X) : + (𝟙 X : X ⟶ X) x = x := by simp -/-- Construct a bundled `Semigrp` from the underlying type and typeclass. -/ +@[to_additive (attr := simp)] +lemma hom_comp {X Y T : Semigrp} (f : X ⟶ Y) (g : Y ⟶ T) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + +/- Provided for rewriting. -/ @[to_additive] -def of (M : Type u) [Semigroup M] : Semigrp := - Bundled.of M +lemma comp_apply {X Y T : Semigrp} (f : X ⟶ Y) (g : Y ⟶ T) (x : X) : + (f ≫ g) x = g (f x) := by simp -/-- Construct a bundled `AddSemigrp` from the underlying type and typeclass. -/ -add_decl_doc AddSemigrp.of +@[to_additive (attr := ext)] +lemma hom_ext {X Y : Semigrp} {f g : X ⟶ Y} (hf : f.hom = g.hom) : f = g := + Hom.ext hf @[to_additive (attr := simp)] -theorem coe_of (R : Type u) [Semigroup R] : (Semigrp.of R : Type u) = R := - rfl +lemma hom_ofHom {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) : (ofHom f).hom = f := rfl @[to_additive (attr := simp)] -lemma mulEquiv_coe_eq {X Y : Type _} [Semigroup X] [Semigroup Y] (e : X ≃* Y) : - (@DFunLike.coe (Semigrp.of X ⟶ Semigrp.of Y) _ (fun _ => (forget Semigrp).obj _) - HasForget.instFunLike (e : X →ₙ* Y) : X → Y) = ↑e := +lemma ofHom_hom {X Y : Semigrp} (f : X ⟶ Y) : + ofHom (Hom.hom f) = f := rfl + +@[to_additive (attr := simp)] +lemma ofHom_id {X : Type u} [Semigroup X] : ofHom (MulHom.id X) = 𝟙 (of X) := rfl + +@[to_additive (attr := simp)] +lemma ofHom_comp {X Y Z : Type u} [Semigroup X] [Semigroup Y] [Semigroup Z] + (f : X →ₙ* Y) (g : Y →ₙ* Z) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := rfl -/-- Typecheck a `MulHom` as a morphism in `Semigrp`. -/ @[to_additive] -def ofHom {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) : of X ⟶ of Y := - f +lemma ofHom_apply {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) (x : X) : + (ofHom f) x = f x := rfl -/-- Typecheck an `AddHom` as a morphism in `AddSemigrp`. -/ -add_decl_doc AddSemigrp.ofHom +@[to_additive (attr := simp)] +lemma inv_hom_apply {X Y : Semigrp} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by + rw [← comp_apply] + simp -@[to_additive] -theorem ofHom_apply {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) (x : X) : - ofHom f x = f x := +@[to_additive (attr := simp)] +lemma hom_inv_apply {X Y : Semigrp} (e : X ≅ Y) (s : Y) : e.hom (e.inv s) = s := by + rw [← comp_apply] + simp + +@[to_additive (attr := simp)] +lemma mulEquiv_coe_eq {X Y : Type _} [Semigroup X] [Semigroup Y] (e : X ≃* Y) : + (ofHom (e : X →ₙ* Y)).hom = ↑e := rfl @[to_additive] @@ -190,8 +361,10 @@ instance : Inhabited Semigrp := ⟨Semigrp.of PEmpty⟩ @[to_additive] -instance hasForgetToMagmaCat : HasForget₂ Semigrp MagmaCat := - BundledHom.forget₂ _ _ +instance hasForgetToMagmaCat : HasForget₂ Semigrp MagmaCat where + forget₂ := + { obj R := MagmaCat.of R + map f := MagmaCat.ofHom f.hom } end Semigrp @@ -205,12 +378,8 @@ variable [Mul X] [Mul Y] @[to_additive (attr := simps) "Build an isomorphism in the category `AddMagmaCat` from an `AddEquiv` between `Add`s."] def MulEquiv.toMagmaCatIso (e : X ≃* Y) : MagmaCat.of X ≅ MagmaCat.of Y where - hom := e.toMulHom - inv := e.symm.toMulHom - hom_inv_id := by - ext - simp_rw [CategoryTheory.comp_apply, toMulHom_eq_coe, MagmaCat.mulEquiv_coe_eq, symm_apply_apply, - CategoryTheory.id_apply] + hom := MagmaCat.ofHom e.toMulHom + inv := MagmaCat.ofHom e.symm.toMulHom end @@ -223,8 +392,8 @@ variable [Semigroup X] [Semigroup Y] "Build an isomorphism in the category `AddSemigroup` from an `AddEquiv` between `AddSemigroup`s."] def MulEquiv.toSemigrpIso (e : X ≃* Y) : Semigrp.of X ≅ Semigrp.of Y where - hom := e.toMulHom - inv := e.symm.toMulHom + hom := Semigrp.ofHom e.toMulHom + inv := Semigrp.ofHom e.symm.toMulHom end @@ -234,13 +403,13 @@ namespace CategoryTheory.Iso @[to_additive "Build an `AddEquiv` from an isomorphism in the category `AddMagmaCat`."] def magmaCatIsoToMulEquiv {X Y : MagmaCat} (i : X ≅ Y) : X ≃* Y := - MulHom.toMulEquiv i.hom i.inv i.hom_inv_id i.inv_hom_id + MulHom.toMulEquiv i.hom.hom i.inv.hom (by ext; simp) (by ext; simp) /-- Build a `MulEquiv` from an isomorphism in the category `Semigroup`. -/ @[to_additive "Build an `AddEquiv` from an isomorphism in the category `AddSemigroup`."] def semigrpIsoToMulEquiv {X Y : Semigrp} (i : X ≅ Y) : X ≃* Y := - MulHom.toMulEquiv i.hom i.inv i.hom_inv_id i.inv_hom_id + MulHom.toMulEquiv i.hom.hom i.inv.hom (by ext; simp) (by ext; simp) end CategoryTheory.Iso @@ -268,14 +437,14 @@ def mulEquivIsoSemigrpIso {X Y : Type u} [Semigroup X] [Semigroup Y] : instance MagmaCat.forgetReflectsIsos : (forget MagmaCat.{u}).ReflectsIsomorphisms where reflects {X Y} f _ := by let i := asIso ((forget MagmaCat).map f) - let e : X ≃* Y := { f, i.toEquiv with } + let e : X ≃* Y := { f.hom, i.toEquiv with } exact e.toMagmaCatIso.isIso_hom @[to_additive] instance Semigrp.forgetReflectsIsos : (forget Semigrp.{u}).ReflectsIsomorphisms where reflects {X Y} f _ := by let i := asIso ((forget Semigrp).map f) - let e : X ≃* Y := { f, i.toEquiv with } + let e : X ≃* Y := { f.hom, i.toEquiv with } exact e.toSemigrpIso.isIso_hom -- Porting note: this was added in order to ensure that `forget₂ CommMonCat MonCat` @@ -283,7 +452,7 @@ instance Semigrp.forgetReflectsIsos : (forget Semigrp.{u}).ReflectsIsomorphisms -- we could have used `CategoryTheory.HasForget.ReflectsIso` alternatively @[to_additive] instance Semigrp.forget₂_full : (forget₂ Semigrp MagmaCat).Full where - map_surjective f := ⟨f, rfl⟩ + map_surjective f := ⟨ofHom f.hom, rfl⟩ /-! Once we've shown that the forgetful functors to type reflect isomorphisms, diff --git a/Mathlib/Algebra/Central/Basic.lean b/Mathlib/Algebra/Central/Basic.lean index e0386ab247eabb..ef13a87438e02e 100644 --- a/Mathlib/Algebra/Central/Basic.lean +++ b/Mathlib/Algebra/Central/Basic.lean @@ -47,7 +47,7 @@ lemma baseField_essentially_unique simp only [center_eq_bot, mem_bot, Set.mem_range, forall_exists_index] rintro x rfl exact ⟨algebraMap k K x, by simp [algebraMap_eq_smul_one, smul_assoc]⟩ } - refine ⟨NoZeroSMulDivisors.algebraMap_injective k K, fun x => ?_⟩ + refine ⟨FaithfulSMul.algebraMap_injective k K, fun x => ?_⟩ have H : algebraMap K D x ∈ (Subalgebra.center K D : Set D) := Subalgebra.algebraMap_mem _ _ rw [show (Subalgebra.center K D : Set D) = Subalgebra.center k D by rfl] at H simp only [center_eq_bot, coe_bot, Set.mem_range] at H diff --git a/Mathlib/Algebra/Central/TensorProduct.lean b/Mathlib/Algebra/Central/TensorProduct.lean index fc53bce1773bad..0b9cfec58fd423 100644 --- a/Mathlib/Algebra/Central/TensorProduct.lean +++ b/Mathlib/Algebra/Central/TensorProduct.lean @@ -72,13 +72,13 @@ lemma right_of_tensor (inj : Function.Injective (algebraMap K B)) [Module.Flat K non-trivial, then `B` is central. -/ lemma left_of_tensor_of_field (K B C : Type*) [Field K] [Ring B] [Ring C] [Nontrivial C] [Algebra K B] [Algebra K C] [IsCentral K (B ⊗[K] C)] : IsCentral K B := - left_of_tensor K B C <| NoZeroSMulDivisors.algebraMap_injective K C + left_of_tensor K B C <| FaithfulSMul.algebraMap_injective K C /-- Let `B` and `C` be two algebras over a field `K`, if `B ⊗[K] C` is central and `A` is non-trivial, then `B` is central. -/ lemma right_of_tensor_of_field (K B C : Type*) [Field K] [Ring B] [Ring C] [Nontrivial B] [Algebra K B] [Algebra K C] [IsCentral K (B ⊗[K] C)] : IsCentral K C := - right_of_tensor K B C <| NoZeroSMulDivisors.algebraMap_injective K B + right_of_tensor K B C <| FaithfulSMul.algebraMap_injective K B end Algebra.IsCentral diff --git a/Mathlib/Algebra/CharP/LinearMaps.lean b/Mathlib/Algebra/CharP/LinearMaps.lean index 93009eaa7324fe..744919935cef12 100644 --- a/Mathlib/Algebra/CharP/LinearMaps.lean +++ b/Mathlib/Algebra/CharP/LinearMaps.lean @@ -40,13 +40,10 @@ variable {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] theorem charP_end {p : ℕ} [hchar : CharP R p] (htorsion : ∃ x : M, Ideal.torsionOf R M x = ⊥) : CharP (M →ₗ[R] M) p where cast_eq_zero_iff' n := by - have hreduction : ∃ x : M, ∀ r : R, r • x = 0 → r = 0 := - Exists.casesOn htorsion fun x hx ↦ - Exists.intro x fun r a ↦ (hx ▸ Ideal.mem_torsionOf_iff x r).mpr a have exact : (n : M →ₗ[R] M) = (n : R) • 1 := by simp only [Nat.cast_smul_eq_nsmul, nsmul_eq_mul, mul_one] rw [exact, LinearMap.ext_iff, ← hchar.1] - exact ⟨fun h ↦ Exists.casesOn hreduction fun x hx ↦ hx n (h x), + exact ⟨fun h ↦ htorsion.casesOn fun x hx ↦ by simpa [← Ideal.mem_torsionOf_iff, hx] using h x, fun h ↦ (congrArg (fun t ↦ ∀ x, t • x = 0) h).mpr fun x ↦ zero_smul R x⟩ end Module diff --git a/Mathlib/Algebra/Divisibility/Units.lean b/Mathlib/Algebra/Divisibility/Units.lean index 4afd288068460f..7621849857a01f 100644 --- a/Mathlib/Algebra/Divisibility/Units.lean +++ b/Mathlib/Algebra/Divisibility/Units.lean @@ -33,12 +33,12 @@ theorem coe_dvd : ↑u ∣ a := /-- In a monoid, an element `a` divides an element `b` iff `a` divides all associates of `b`. -/ theorem dvd_mul_right : a ∣ b * u ↔ a ∣ b := - Iff.intro (fun ⟨c, Eq⟩ ↦ ⟨c * ↑u⁻¹, by rw [← mul_assoc, ← Eq, Units.mul_inv_cancel_right]⟩) - fun ⟨_, Eq⟩ ↦ Eq.symm ▸ (_root_.dvd_mul_right _ _).mul_right _ + Iff.intro (fun ⟨c, eq⟩ ↦ ⟨c * ↑u⁻¹, by rw [← mul_assoc, ← eq, Units.mul_inv_cancel_right]⟩) + fun ⟨_, eq⟩ ↦ eq.symm ▸ (_root_.dvd_mul_right _ _).mul_right _ /-- In a monoid, an element `a` divides an element `b` iff all associates of `a` divide `b`. -/ theorem mul_right_dvd : a * u ∣ b ↔ a ∣ b := - Iff.intro (fun ⟨c, Eq⟩ => ⟨↑u * c, Eq.trans (mul_assoc _ _ _)⟩) fun h => + Iff.intro (fun ⟨c, eq⟩ => ⟨↑u * c, eq.trans (mul_assoc _ _ _)⟩) fun h => dvd_trans (Dvd.intro (↑u⁻¹) (by rw [mul_assoc, u.mul_inv, mul_one])) h end Monoid diff --git a/Mathlib/Algebra/Field/IsField.lean b/Mathlib/Algebra/Field/IsField.lean index f598ad71864269..d0a69f55c8b70a 100644 --- a/Mathlib/Algebra/Field/IsField.lean +++ b/Mathlib/Algebra/Field/IsField.lean @@ -65,10 +65,11 @@ noncomputable def IsField.toSemifield {R : Type u} [Semiring R] (h : IsField R) nnqsmul_def _ _ := rfl /-- Transferring from `IsField` to `Field`. -/ -noncomputable def IsField.toField {R : Type u} [Ring R] (h : IsField R) : Field R := - { ‹Ring R›, IsField.toSemifield h with - qsmul := _ - qsmul_def := fun _ _ => rfl } +noncomputable def IsField.toField {R : Type u} [Ring R] (h : IsField R) : Field R where + __ := (‹Ring R›:) -- this also works without the `( :)`, but it's slow + __ := h.toSemifield + qsmul := _ + qsmul_def := fun _ _ => rfl /-- For each field, and for each nonzero element of said field, there is a unique inverse. Since `IsField` doesn't remember the data of an `inv` function and as such, diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 24c0f1946cbfad..86fccdf74e237b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -851,7 +851,7 @@ instance normal_inf_normal (H K : Subgroup G) [hH : H.Normal] [hK : K.Normal] : @[to_additive] theorem normal_iInf_normal {ι : Type*} {a : ι → Subgroup G} - (norm : ∀ i : ι , (a i).Normal) : (iInf a).Normal := by + (norm : ∀ i : ι, (a i).Normal) : (iInf a).Normal := by constructor intro g g_in_iInf h rw [Subgroup.mem_iInf] at g_in_iInf ⊢ diff --git a/Mathlib/Algebra/Group/WithOne/Basic.lean b/Mathlib/Algebra/Group/WithOne/Basic.lean index 5382fd1d0651ac..2f1a04997c4ff4 100644 --- a/Mathlib/Algebra/Group/WithOne/Basic.lean +++ b/Mathlib/Algebra/Group/WithOne/Basic.lean @@ -65,7 +65,6 @@ def lift : (α →ₙ* β) ≃ (WithOne α →* β) where invFun F := F.toMulHom.comp coeMulHom left_inv _ := MulHom.ext fun _ => rfl right_inv F := MonoidHom.ext fun x => WithOne.cases_on x F.map_one.symm (fun _ => rfl) --- Porting note: the above proofs were broken because they were parenthesized wrong by mathport? variable (f : α →ₙ* β) @@ -119,8 +118,6 @@ def _root_.MulEquiv.withOneCongr (e : α ≃* β) : WithOne α ≃* WithOne β : left_inv := (by induction · <;> simp) right_inv := (by induction · <;> simp) } --- Porting note: for this declaration and the two below I added the `to_additive` attribute because --- it seemed to be missing from mathlib3 @[to_additive (attr := simp)] theorem _root_.MulEquiv.withOneCongr_refl : (MulEquiv.refl α).withOneCongr = MulEquiv.refl _ := MulEquiv.toMonoidHom_injective map_id diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 1222fa35f4af01..936f7cd3c7d9e8 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -149,7 +149,7 @@ lemma zero_pow_eq (n : ℕ) : (0 : M₀) ^ n = if n = 0 then 1 else 0 := by lemma zero_pow_eq_one₀ [Nontrivial M₀] : (0 : M₀) ^ n = 1 ↔ n = 0 := by rw [zero_pow_eq, one_ne_zero.ite_eq_left_iff] -lemma pow_eq_zero_of_le : ∀ {m n} (_ : m ≤ n) (_ : a ^ m = 0), a ^ n = 0 +lemma pow_eq_zero_of_le : ∀ {m n}, m ≤ n → a ^ m = 0 → a ^ n = 0 | _, _, Nat.le.refl, ha => ha | _, _, Nat.le.step hmn, ha => by rw [pow_succ, pow_eq_zero_of_le hmn ha, zero_mul] diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean index 4aae573cb738de..4baf8de024530d 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean @@ -364,7 +364,7 @@ noncomputable def extFunctor (n : ℕ) : Cᵒᵖ ⥤ C ⥤ AddCommGrp.{w} where symm apply Ext.comp_assoc all_goals omega } - map_comp {X₁ X₂ X₃} f f' := by + map_comp {X₁ X₂ X₃} f f' := by ext Y α dsimp rw [← Ext.mk₀_comp_mk₀] diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 5f4f3fc236c8ab..6db04a3d047af9 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -262,9 +262,9 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu rw [traceForm_apply_apply, LinearMap.zero_apply] let A := AlgebraicClosure (FractionRing R) suffices algebraMap R A (trace R _ ((φ z).comp (φ x))) = 0 by - have _i : NoZeroSMulDivisors R A := NoZeroSMulDivisors.trans R (FractionRing R) A + have _i : NoZeroSMulDivisors R A := NoZeroSMulDivisors.trans_faithfulSMul R (FractionRing R) A rw [← map_zero (algebraMap R A)] at this - exact NoZeroSMulDivisors.algebraMap_injective R A this + exact FaithfulSMul.algebraMap_injective R A this rw [← LinearMap.trace_baseChange, LinearMap.baseChange_comp, ← toEnd_baseChange, ← toEnd_baseChange] replace hz : 1 ⊗ₜ z ∈ lowerCentralSeries A (A ⊗[R] L) (A ⊗[R] L) 1 := by diff --git a/Mathlib/Algebra/Module/CharacterModule.lean b/Mathlib/Algebra/Module/CharacterModule.lean index b4146d2497f6a6..ea00edf9b07353 100644 --- a/Mathlib/Algebra/Module/CharacterModule.lean +++ b/Mathlib/Algebra/Module/CharacterModule.lean @@ -236,8 +236,8 @@ lemma surjective_of_dual_injective (f : A →ₗ[R] A') (hf : Function.Injective obtain ⟨b, rfl⟩ := QuotientAddGroup.mk'_surjective _ a suffices eq : dual (Submodule.mkQ _) c = 0 from congr($eq b) refine hf ?_ - rw [← LinearMap.comp_apply, ← dual_comp, LinearMap.range_mkQ_comp, dual_zero] - rfl + rw [← LinearMap.comp_apply, ← dual_comp, LinearMap.range_mkQ_comp, dual_zero, + LinearMap.zero_apply, dual_apply, AddMonoidHom.zero_comp] lemma dual_injective_iff_surjective {f : A →ₗ[R] A'} : Function.Injective (dual f) ↔ Function.Surjective f := diff --git a/Mathlib/Algebra/Module/Lattice.lean b/Mathlib/Algebra/Module/Lattice.lean index b05bf47a40a736..a4aae942ec8272 100644 --- a/Mathlib/Algebra/Module/Lattice.lean +++ b/Mathlib/Algebra/Module/Lattice.lean @@ -164,9 +164,7 @@ variable [IsPrincipalIdealRing R] Note that under our conditions, `NoZeroSMulDivisors R K` simply says that `algebraMap R K` is injective. -/ instance free [NoZeroSMulDivisors R K] (M : Submodule R V) [IsLattice K M] : Module.Free R M := by - haveI : NoZeroSMulDivisors R V := by - apply NoZeroSMulDivisors.of_algebraMap_injective' (A := K) - exact NoZeroSMulDivisors.algebraMap_injective R K + have := NoZeroSMulDivisors.trans_faithfulSMul R K V -- any torsion free finite module over a PID is free infer_instance diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index 4a3b1f34b2439c..89b681b6bf0d21 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -1270,7 +1270,7 @@ lemma LocalizedModule.subsingleton_iff_ker_eq_top {S : Submonoid R} : rw [← top_le_iff] refine ⟨fun H m _ ↦ Subsingleton.elim _ _, fun H ↦ (subsingleton_iff_forall_eq 0).mpr fun x ↦ ?_⟩ obtain ⟨⟨x, s⟩, rfl⟩ := IsLocalizedModule.mk'_surjective S (LocalizedModule.mkLinearMap S M) x - simpa using @H x trivial + simpa using @H x Submodule.mem_top lemma LocalizedModule.subsingleton_iff {S : Submonoid R} : Subsingleton (LocalizedModule S M) ↔ ∀ m : M, ∃ r ∈ S, r • m = 0 := by diff --git a/Mathlib/Algebra/Order/AbsoluteValue/Basic.lean b/Mathlib/Algebra/Order/AbsoluteValue/Basic.lean index 39413e7dcbea8d..218fb19a9985b7 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue/Basic.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue/Basic.lean @@ -328,9 +328,8 @@ lemma isNontrivial_iff_ne_trivial [DecidablePred fun x : R ↦ x = 0] [NoZeroDiv [Nontrivial S] (v : AbsoluteValue R S) : v.IsNontrivial ↔ v ≠ .trivial := by refine ⟨fun ⟨x, hx₀, hx₁⟩ h ↦ hx₁ <| h.symm ▸ trivial_apply hx₀, fun H ↦ ?_⟩ + simp only [IsNontrivial] contrapose! H - simp only [IsNontrivial] at H - push_neg at H ext1 x rcases eq_or_ne x 0 with rfl | hx · simp diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 411fb19700d736..f137d7941bb7df 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -207,6 +207,17 @@ theorem prod_le_prod_fiberwise_of_prod_fiber_le_one' {t : Finset ι'} {g : ι ∏ x ∈ s, f x ≤ ∏ y ∈ t, ∏ x ∈ s with g x = y, f x := @prod_fiberwise_le_prod_of_one_le_prod_fiber' _ Nᵒᵈ _ _ _ _ _ _ _ h +@[to_additive] +lemma prod_image_le_of_one_le + {g : ι → ι'} {f : ι' → N} (hf : ∀ u ∈ s.image g, 1 ≤ f u) : + ∏ u ∈ s.image g, f u ≤ ∏ u ∈ s, f (g u) := by + rw [prod_comp f g] + refine prod_le_prod' fun a hag ↦ ?_ + obtain ⟨i, hi, hig⟩ := Finset.mem_image.mp hag + apply le_self_pow (hf a hag) + rw [← Nat.pos_iff_ne_zero, card_pos] + exact ⟨i, mem_filter.mpr ⟨hi, hig⟩⟩ + end OrderedCommMonoid @[to_additive] diff --git a/Mathlib/Algebra/Order/GroupWithZero/Bounds.lean b/Mathlib/Algebra/Order/GroupWithZero/Bounds.lean index a697535c762927..ebe1f4440187ce 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Bounds.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Bounds.lean @@ -18,14 +18,13 @@ open Set lemma BddAbove.range_comp_of_nonneg {α β γ : Type*} [Nonempty α] [Preorder β] [Zero β] [Preorder γ] {f : α → β} {g : β → γ} (hf : BddAbove (range f)) (hf0 : 0 ≤ f) (hg : MonotoneOn g {x : β | 0 ≤ x}) : BddAbove (range (fun x => g (f x))) := by - have hg' : BddAbove (g '' (range f)) := by - apply hg.map_bddAbove (by rintro x ⟨a, rfl⟩; exact hf0 a) - · obtain ⟨b, hb⟩ := hf - use b, hb - simp only [mem_upperBounds, mem_range, forall_exists_index, forall_apply_eq_imp_iff] at hb - exact le_trans (hf0 Classical.ofNonempty) (hb Classical.ofNonempty) - change BddAbove (range (g ∘ f)) - simpa only [Set.range_comp] using hg' + suffices hg' : BddAbove (g '' range f) by + rwa [← Function.comp_def, Set.range_comp] + apply hg.map_bddAbove (by rintro x ⟨a, rfl⟩; exact hf0 a) + obtain ⟨b, hb⟩ := hf + use b, hb + simp only [mem_upperBounds, mem_range, forall_exists_index, forall_apply_eq_imp_iff] at hb + exact le_trans (hf0 Classical.ofNonempty) (hb Classical.ofNonempty) /-- If `u v : α → β` are nonnegative and bounded above, then `u * v` is bounded above. -/ theorem bddAbove_range_mul {α β : Type*} [Nonempty α] {u v : α → β} [Preorder β] [Zero β] [Mul β] diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index f538946d0fc944..8330f67d0078d1 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -950,7 +950,7 @@ lemma Monotone.mul [PosMulMono M₀] [MulPosMono M₀] (hf : Monotone f) (hg : M (hf₀ : ∀ x, 0 ≤ f x) (hg₀ : ∀ x, 0 ≤ g x) : Monotone (f * g) := fun _ _ h ↦ mul_le_mul (hf h) (hg h) (hg₀ _) (hf₀ _) -lemma MonotoneOn.mul [PosMulMono M₀] [MulPosMono M₀] {s : Set α } (hf : MonotoneOn f s) +lemma MonotoneOn.mul [PosMulMono M₀] [MulPosMono M₀] {s : Set α} (hf : MonotoneOn f s) (hg : MonotoneOn g s) (hf₀ : ∀ x ∈ s, 0 ≤ f x) (hg₀ : ∀ x ∈ s, 0 ≤ g x) : MonotoneOn (f * g) s := fun _ ha _ hb h ↦ mul_le_mul (hf ha hb h) (hg ha hb h) (hg₀ _ ha) (hf₀ _ hb) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index 6d9fc46645bc15..6b90d6bb4808f3 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -241,6 +241,9 @@ lemma addLECancellable_of_ne_top [Preorder α] [ContravariantClass α α (· + lemma addLECancellable_of_lt_top [Preorder α] [ContravariantClass α α (· + ·) (· ≤ ·)] (ha : a < ⊤) : AddLECancellable a := addLECancellable_of_ne_top ha.ne +lemma addLECancellable_coe [Preorder α] [ContravariantClass α α (· + ·) (· ≤ ·)] (a : α) : + AddLECancellable (a : WithTop α) := addLECancellable_of_ne_top coe_ne_top + lemma addLECancellable_iff_ne_top [Nonempty α] [Preorder α] [ContravariantClass α α (· + ·) (· ≤ ·)] : AddLECancellable a ↔ a ≠ ⊤ where mp := by rintro h rfl; exact (coe_lt_top <| Classical.arbitrary _).not_le <| h <| by simp diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 1047af9fab54a3..430e2e91cbe675 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -397,7 +397,7 @@ theorem modByMonic_eq_zero_iff_dvd (hq : Monic q) : p %ₘ q = 0 ↔ q ∣ p := exact not_lt_of_ge (Nat.le_add_right _ _) (WithBot.coe_lt_coe.1 this)⟩ -/-- See `Polynomial.mul_left_modByMonic` for the other multiplication order. That version, unlike +/-- See `Polynomial.mul_self_modByMonic` for the other multiplication order. That version, unlike this one, requires commutativity. -/ @[simp] lemma self_mul_modByMonic (hq : q.Monic) : (q * p) %ₘ q = 0 := by @@ -476,7 +476,7 @@ section multiplicity /-- An algorithm for deciding polynomial divisibility. The algorithm is "compute `p %ₘ q` and compare to `0`". -See `polynomial.modByMonic` for the algorithm that computes `%ₘ`. +See `Polynomial.modByMonic` for the algorithm that computes `%ₘ`. -/ def decidableDvdMonic [DecidableEq R] (p : R[X]) (hq : Monic q) : Decidable (q ∣ p) := decidable_of_iff (p %ₘ q = 0) (modByMonic_eq_zero_iff_dvd hq) @@ -652,7 +652,7 @@ theorem eval_divByMonic_pow_rootMultiplicity_ne_zero {p : R[X]} (a : R) (hp : p (monic_X_sub_C _) hp).not_pow_dvd_of_multiplicity_lt (Nat.lt_succ_self _) (dvd_of_mul_right_eq _ this) -/-- See `Polynomial.mul_right_modByMonic` for the other multiplication order. This version, unlike +/-- See `Polynomial.self_mul_modByMonic` for the other multiplication order. This version, unlike that one, requires commutativity. -/ @[simp] lemma mul_self_modByMonic (hq : q.Monic) : (p * q) %ₘ q = 0 := by diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 3a5f5ab6b5d84d..620416ce8fc9e7 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -117,7 +117,7 @@ theorem mem_roots_map_of_injective [Semiring S] {p : S[X]} {f : S →+* R} rw [mem_roots ((Polynomial.map_ne_zero_iff hf).mpr hp), IsRoot, eval_map] lemma mem_roots_iff_aeval_eq_zero {x : R} (w : p ≠ 0) : x ∈ roots p ↔ aeval x p = 0 := by - rw [aeval_def, ← mem_roots_map_of_injective (NoZeroSMulDivisors.algebraMap_injective _ _) w, + rw [aeval_def, ← mem_roots_map_of_injective (FaithfulSMul.algebraMap_injective _ _) w, Algebra.id.map_eq_id, map_id] theorem card_le_degree_of_subset_roots {p : R[X]} {Z : Finset R} (h : Z.val ⊆ p.roots) : @@ -392,7 +392,7 @@ theorem mem_aroots' [CommRing S] [IsDomain S] [Algebra T S] {p : T[X]} {a : S} : theorem mem_aroots [CommRing S] [IsDomain S] [Algebra T S] [NoZeroSMulDivisors T S] {p : T[X]} {a : S} : a ∈ p.aroots S ↔ p ≠ 0 ∧ aeval a p = 0 := by rw [mem_aroots', Polynomial.map_ne_zero_iff] - exact NoZeroSMulDivisors.algebraMap_injective T S + exact FaithfulSMul.algebraMap_injective T S theorem aroots_mul [CommRing S] [IsDomain S] [Algebra T S] [NoZeroSMulDivisors T S] {p q : T[X]} (hpq : p * q ≠ 0) : @@ -400,7 +400,7 @@ theorem aroots_mul [CommRing S] [IsDomain S] [Algebra T S] suffices map (algebraMap T S) p * map (algebraMap T S) q ≠ 0 by rw [aroots_def, Polynomial.map_mul, roots_mul this] rwa [← Polynomial.map_mul, Polynomial.map_ne_zero_iff - (NoZeroSMulDivisors.algebraMap_injective T S)] + (FaithfulSMul.algebraMap_injective T S)] @[simp] theorem aroots_X_sub_C [CommRing S] [IsDomain S] [Algebra T S] @@ -436,7 +436,7 @@ theorem aroots_C_mul [CommRing S] [IsDomain S] [Algebra T S] (C a * p).aroots S = p.aroots S := by rw [aroots_def, Polynomial.map_mul, map_C, roots_C_mul] rwa [map_ne_zero_iff] - exact NoZeroSMulDivisors.algebraMap_injective T S + exact FaithfulSMul.algebraMap_injective T S @[simp] theorem aroots_smul_nonzero [CommRing S] [IsDomain S] [Algebra T S] @@ -534,7 +534,7 @@ theorem mem_rootSet' {p : T[X]} {S : Type*} [CommRing S] [IsDomain S] [Algebra T theorem mem_rootSet {p : T[X]} {S : Type*} [CommRing S] [IsDomain S] [Algebra T S] [NoZeroSMulDivisors T S] {a : S} : a ∈ p.rootSet S ↔ p ≠ 0 ∧ aeval a p = 0 := by - rw [mem_rootSet', Polynomial.map_ne_zero_iff (NoZeroSMulDivisors.algebraMap_injective T S)] + rw [mem_rootSet', Polynomial.map_ne_zero_iff (FaithfulSMul.algebraMap_injective T S)] theorem mem_rootSet_of_ne {p : T[X]} {S : Type*} [CommRing S] [IsDomain S] [Algebra T S] [NoZeroSMulDivisors T S] (hp : p ≠ 0) {a : S} : a ∈ p.rootSet S ↔ aeval a p = 0 := @@ -559,7 +559,7 @@ theorem rootSet_mapsTo {p : T[X]} {S S'} [CommRing S] [IsDomain S] [Algebra T S] (p.rootSet S).MapsTo f (p.rootSet S') := by refine rootSet_maps_to' (fun h₀ => ?_) f obtain rfl : p = 0 := - map_injective _ (NoZeroSMulDivisors.algebraMap_injective T S') (by rwa [Polynomial.map_zero]) + map_injective _ (FaithfulSMul.algebraMap_injective T S') (by rwa [Polynomial.map_zero]) exact Polynomial.map_zero _ theorem mem_rootSet_of_injective [CommRing S] {p : S[X]} [Algebra S R] diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 6a4475fd7cc956..c4f440a1475071 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1191,8 +1191,7 @@ instance instInv : Inv ℍ[R] := ⟨fun a => (normSq a)⁻¹ • star a⟩ instance instGroupWithZero : GroupWithZero ℍ[R] := - { Quaternion.instNontrivial, - (by infer_instance : MonoidWithZero ℍ[R]) with + { Quaternion.instNontrivial with inv := Inv.inv inv_zero := by rw [instInv_inv, star_zero, smul_zero] mul_inv_cancel := fun a ha => by @@ -1230,14 +1229,14 @@ instance instRatCast : RatCast ℍ[R] where ratCast q := (q : R) @[norm_cast] lemma coe_ratCast (q : ℚ) : ↑(q : R) = (q : ℍ[R]) := rfl instance instDivisionRing : DivisionRing ℍ[R] where - __ := Quaternion.instGroupWithZero __ := Quaternion.instRing + __ := Quaternion.instGroupWithZero nnqsmul := (· • ·) qsmul := (· • ·) - nnratCast_def q := by rw [← coe_nnratCast, NNRat.cast_def, coe_div, coe_natCast, coe_natCast] - ratCast_def q := by rw [← coe_ratCast, Rat.cast_def, coe_div, coe_intCast, coe_natCast] - nnqsmul_def q x := by rw [← coe_nnratCast, coe_mul_eq_smul]; ext <;> exact NNRat.smul_def _ _ - qsmul_def q x := by rw [← coe_ratCast, coe_mul_eq_smul]; ext <;> exact Rat.smul_def _ _ + nnratCast_def _ := by rw [← coe_nnratCast, NNRat.cast_def, coe_div, coe_natCast, coe_natCast] + ratCast_def _ := by rw [← coe_ratCast, Rat.cast_def, coe_div, coe_intCast, coe_natCast] + nnqsmul_def _ _ := by rw [← coe_nnratCast, coe_mul_eq_smul]; ext <;> exact NNRat.smul_def .. + qsmul_def _ _ := by rw [← coe_ratCast, coe_mul_eq_smul]; ext <;> exact Rat.smul_def .. theorem normSq_inv : normSq a⁻¹ = (normSq a)⁻¹ := map_inv₀ normSq _ diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index f5b3128adee627..b6889e52834068 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -20,11 +20,11 @@ As opposed to `Even`, `Odd` does not have a multiplicative counterpart. Try to generalize `Even` lemmas further. For example, there are still a few lemmas whose `Semiring` assumptions I (DT) am not convinced are necessary. If that turns out to be true, they could be moved -to `Algebra.Group.Even`. +to `Mathlib.Algebra.Group.Even`. ## See also -`Algebra.Group.Even` for the definition of even elements. +`Mathlib.Algebra.Group.Even` for the definition of even elements. -/ assert_not_exists DenselyOrdered OrderedRing diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index c1789b31e7cc86..64b6695b7b5367 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -812,7 +812,7 @@ theorem ofLeftInverseS_symm_apply {g : S → R} {f : R →+* S} (h : Function.Le rfl /-- Given an equivalence `e : R ≃+* S` of semirings and a subsemiring `s` of `R`, -`subsemiring_map e s` is the induced equivalence between `s` and `s.map e` -/ +`subsemiringMap e s` is the induced equivalence between `s` and `s.map e` -/ def subsemiringMap (e : R ≃+* S) (s : Subsemiring R) : s ≃+* s.map (e : R →+* S) := { e.toAddEquiv.addSubmonoidMap s.toAddSubmonoid, e.toMulEquiv.submonoidMap s.toSubmonoid with } @@ -899,7 +899,6 @@ instance mulActionWithZero [Zero α] [MulActionWithZero R' α] (S : Subsemiring MulActionWithZero S α := MulActionWithZero.compHom _ S.subtype.toMonoidWithZeroHom --- Porting note: instance named explicitly for use in `RingTheory/Subring/Basic` /-- The action by a subsemiring is the action by the underlying semiring. -/ instance module [AddCommMonoid α] [Module R' α] (S : Subsemiring R') : Module S α := -- Porting note: copying over the `smul` field causes a timeout diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index cc8afa1e322041..b4abde67aa3743 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -374,7 +374,7 @@ theorem range_fromSpec : delta IsAffineOpen.fromSpec; dsimp [IsAffineOpen.isoSpec_inv] rw [Set.range_comp, Set.range_eq_univ.mpr, Set.image_univ] · exact Subtype.range_coe - rw [← coe_comp, ← TopCat.epi_iff_surjective] + rw [← TopCat.coe_comp, ← TopCat.epi_iff_surjective] infer_instance @[simp] @@ -448,9 +448,9 @@ theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion refine ⟨fun hU => @isAffine_of_isIso _ _ (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] + · rw [Scheme.comp_base, TopCat.coe_comp, Set.range_comp] dsimp [Opens.coe_inclusion', Scheme.restrict] - erw [Subtype.range_coe, Subtype.range_coe] -- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 + rw [Subtype.range_coe, Subtype.range_coe] rfl · infer_instance diff --git a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean index 042468ac880d9b..57cf17e2d2e935 100644 --- a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean @@ -113,7 +113,7 @@ def Cover.bind [P.IsStableUnderComposition] (f : ∀ x : 𝒰.J, (𝒰.obj x).Co rcases (f (𝒰.f x)).covers y with ⟨z, hz⟩ change x ∈ Set.range ((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).base use z - erw [CategoryTheory.comp_apply] + simp only [comp_coeBase, TopCat.hom_comp, ContinuousMap.comp_apply] rw [hz, hy] map_prop _ := P.comp_mem _ _ ((f _).map_prop _) (𝒰.map_prop _) diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index a6ec6c2082cb27..c21c1e5f1caab7 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -274,7 +274,7 @@ theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X) (r : (X.local_affine x).choose_spec.choose) : Set.range (X.affineBasisCover.map ⟨x, r⟩).base = (X.affineCover.map x).base '' (PrimeSpectrum.basicOpen r).1 := by - erw [coe_comp, Set.range_comp] + erw [TopCat.coe_comp, Set.range_comp] -- Porting note: `congr` fails to see the goal is comparing image of the same function refine congr_arg (_ '' ·) ?_ exact (PrimeSpectrum.localization_away_comap_range (Localization.Away r) r :) @@ -291,7 +291,7 @@ theorem affineBasisCover_is_basis (X : Scheme.{u}) : let U' := (X.affineCover.map (X.affineCover.f a)).base ⁻¹' U have hxU' : x ∈ U' := by rw [← e] at haU; exact haU rcases PrimeSpectrum.isBasis_basic_opens.exists_subset_of_mem_open hxU' - ((X.affineCover.map (X.affineCover.f a)).base.continuous_toFun.isOpen_preimage _ + ((X.affineCover.map (X.affineCover.f a)).base.hom.continuous_toFun.isOpen_preimage _ hU) with ⟨_, ⟨_, ⟨s, rfl⟩, rfl⟩, hxV, hVU⟩ refine ⟨_, ⟨⟨_, s⟩, rfl⟩, ?_, ?_⟩ <;> rw [affineBasisCover_map_range] diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index b897c89abc30fd..dd66da6f92dd13 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -88,9 +88,10 @@ theorem toΓSpec_continuous : Continuous X.toΓSpecFun := by /-- The canonical (bundled) continuous map from the underlying topological space of `X` to the prime spectrum of its global sections. -/ -def toΓSpecBase : X.toTopCat ⟶ Spec.topObj (Γ.obj (op X)) where - toFun := X.toΓSpecFun - continuous_toFun := X.toΓSpec_continuous +def toΓSpecBase : X.toTopCat ⟶ Spec.topObj (Γ.obj (op X)) := + TopCat.ofHom + { toFun := X.toΓSpecFun + continuous_toFun := X.toΓSpec_continuous } variable (r : Γ.obj (op X)) @@ -300,7 +301,7 @@ theorem right_triangle (R : CommRingCat) : apply LocallyRingedSpace.comp_ring_hom_ext · ext (p : PrimeSpectrum R) dsimp - ext x + refine PrimeSpectrum.ext (Ideal.ext fun x => ?_) erw [← IsLocalization.AtPrime.to_map_mem_maximal_iff ((structureSheaf R).presheaf.stalk p) p.asIdeal x] rfl diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index 2fd42825b08b67..32baf4533398b9 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -226,7 +226,7 @@ theorem ι_eq_iff (i j : D.J) (x : (D.U i).carrier) (y : (D.U j).carrier) : (TopCat.GlueData.ι_eq_iff_rel D.toLocallyRingedSpaceGlueData.toSheafedSpaceGlueData.toPresheafedSpaceGlueData.toTopGlueData i j x y) - rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff, ← CategoryTheory.comp_apply] + rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff, ← ConcreteCategory.comp_apply] · simp_rw [← D.ι_isoCarrier_inv] rfl -- `rfl` was not needed before https://github.com/leanprover-community/mathlib4/pull/13170 · infer_instance @@ -338,7 +338,7 @@ theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.base := by intro x y h obtain ⟨i, x, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective x obtain ⟨j, y, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective y - rw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at h + rw [← ConcreteCategory.comp_apply, ← ConcreteCategory.comp_apply] at h simp_rw [← Scheme.comp_base] at h rw [ι_fromGlued, ι_fromGlued] at h let e := @@ -347,8 +347,10 @@ theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.base := by rw [𝒰.gluedCover.ι_eq_iff] use e.hom ⟨⟨x, y⟩, h⟩ constructor - · erw [← comp_apply e.hom, IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.left]; rfl - · erw [← comp_apply e.hom, pullbackSymmetry_hom_comp_fst, + · erw [← ConcreteCategory.comp_apply e.hom, + IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.left] + rfl + · erw [← ConcreteCategory.comp_apply e.hom, pullbackSymmetry_hom_comp_fst, IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.right] rfl @@ -371,7 +373,7 @@ theorem fromGlued_open_map : IsOpenMap 𝒰.fromGlued.base := by · rw [← Set.image_preimage_eq_inter_range] apply (show IsOpenImmersion (𝒰.map (𝒰.f x)) from inferInstance).base_open.isOpenMap convert hU (𝒰.f x) using 1 - rw [← ι_fromGlued]; erw [coe_comp]; rw [Set.preimage_comp] + rw [← ι_fromGlued]; erw [TopCat.coe_comp]; rw [Set.preimage_comp] congr! 1 exact Set.preimage_image_eq _ 𝒰.fromGlued_injective · exact ⟨hx, 𝒰.covers x⟩ @@ -387,7 +389,7 @@ instance : Epi 𝒰.fromGlued.base := by intro x obtain ⟨y, h⟩ := 𝒰.covers x use (𝒰.gluedCover.ι (𝒰.f x)).base y - rw [← CategoryTheory.comp_apply] + rw [← ConcreteCategory.comp_apply] rw [← 𝒰.ι_fromGlued (𝒰.f x)] at h exact h diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index e28bfdee6c2901..a6d3e6d69c5588 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -64,7 +64,7 @@ section Initial /-- The map from the empty scheme. -/ @[simps] def Scheme.emptyTo (X : Scheme.{u}) : ∅ ⟶ X := - ⟨{ base := ⟨fun x => PEmpty.elim x, by fun_prop⟩ + ⟨{ base := TopCat.ofHom ⟨fun x => PEmpty.elim x, by fun_prop⟩ c := { app := fun _ => CommRingCat.punitIsTerminal.from _ } }, fun x => PEmpty.elim x⟩ @[ext] @@ -99,7 +99,7 @@ instance (priority := 100) isOpenImmersion_of_isEmpty {X Y : Scheme} (f : X ⟶ instance (priority := 100) isIso_of_isEmpty {X Y : Scheme} (f : X ⟶ Y) [IsEmpty Y] : IsIso f := by - haveI : IsEmpty X := f.base.1.isEmpty + haveI : IsEmpty X := f.base.hom.1.isEmpty have : Epi f.base := by rw [TopCat.epi_iff_surjective]; rintro (x : Y) exact isEmptyElim x @@ -310,7 +310,7 @@ lemma sigmaMk_mk (i) (x : f i) : show ((TopCat.sigmaCofan (fun x ↦ (f x).toTopCat)).inj i ≫ (colimit.isoColimitCocone ⟨_, TopCat.sigmaCofanIsColimit _⟩).inv ≫ _) x = Scheme.forgetToTop.map (Sigma.ι f i) x - congr 1 + congr 2 refine (colimit.isoColimitCocone_ι_inv_assoc ⟨_, TopCat.sigmaCofanIsColimit _⟩ _ _).trans ?_ exact ι_comp_sigmaComparison Scheme.forgetToTop _ _ @@ -388,7 +388,7 @@ lemma coprodMk_inl (x : X) : show ((TopCat.binaryCofan X Y).inl ≫ (colimit.isoColimitCocone ⟨_, TopCat.binaryCofanIsColimit _ _⟩).inv ≫ _) x = Scheme.forgetToTop.map coprod.inl x - congr 1 + congr 2 refine (colimit.isoColimitCocone_ι_inv_assoc ⟨_, TopCat.binaryCofanIsColimit _ _⟩ _ _).trans ?_ exact coprodComparison_inl Scheme.forgetToTop @@ -398,7 +398,7 @@ lemma coprodMk_inr (x : Y) : show ((TopCat.binaryCofan X Y).inr ≫ (colimit.isoColimitCocone ⟨_, TopCat.binaryCofanIsColimit _ _⟩).inv ≫ _) x = Scheme.forgetToTop.map coprod.inr x - congr 1 + congr 2 refine (colimit.isoColimitCocone_ι_inv_assoc ⟨_, TopCat.binaryCofanIsColimit _ _⟩ _ _).trans ?_ exact coprodComparison_inr Scheme.forgetToTop diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index 5b5e0f2bcc6d8d..24705327fb3157 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -195,10 +195,10 @@ namespace IsLocalAtSource attribute [instance] respectsIso /-- -`P` is local at the target if +`P` is local at the source if 1. `P` respects isomorphisms. -2. If `P` holds for `f : X ⟶ Y`, then `P` holds for `f ∣_ U` for any `U`. -3. If `P` holds for `f ∣_ U` for an open cover `U` of `Y`, then `P` holds for `f`. +2. If `P` holds for `f : X ⟶ Y`, then `P` holds for `U.ι ≫ f` for any `U`. +3. If `P` holds for `U.ι ≫ f` for an open cover `U` of `X`, then `P` holds for `f`. -/ protected lemma mk' {P : MorphismProperty Scheme} [P.RespectsIso] (restrict : ∀ {X Y : Scheme} (f : X ⟶ Y) (U : X.Opens), P f → P (U.ι ≫ f)) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index 3cd9d1cbcdddb0..dff41458fbd25d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -229,9 +229,9 @@ lemma topologically_isLocalAtTarget apply IsLocalAtTarget.mk' · intro X Y f U hf simp_rw [topologically, morphismRestrict_base] - exact hP₂ f.base U.carrier f.base.2 U.2 hf + exact hP₂ f.base U.carrier f.base.hom.2 U.2 hf · intro X Y f ι U hU hf - apply hP₃ f.base U hU f.base.continuous fun i ↦ ?_ + apply hP₃ f.base U hU f.base.hom.continuous fun i ↦ ?_ rw [← morphismRestrict_base] exact hf i diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean index a03d46fd9e8300..f3fc346cc6c076 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Immersion.lean @@ -137,7 +137,7 @@ theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsImmersion g] __ := IsPreimmersion.of_comp f g isLocallyClosed_range := by rw [← Set.preimage_image_eq (Set.range _) g.isEmbedding.injective] - have := (f ≫ g).isLocallyClosed_range.preimage g.base.2 + have := (f ≫ g).isLocallyClosed_range.preimage g.base.hom.2 simpa only [Scheme.comp_coeBase, TopCat.coe_comp, Set.range_comp] using this theorem comp_iff {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsImmersion g] : diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index a941013a085e5d..63fb1ec953a187 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -47,7 +47,7 @@ instance (priority := 900) quasiCompact_of_isIso {X Y : Scheme} (f : X ⟶ Y) [I QuasiCompact f := by constructor intro U _ hU' - convert hU'.image (inv f.base).continuous_toFun using 1 + convert hU'.image (inv f.base).hom.continuous_toFun using 1 rw [Set.image_eq_preimage_of_inverse] · delta Function.LeftInverse exact IsIso.inv_hom_id_apply f.base diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 6c6b1821adf90f..96f3f418570754 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -158,7 +158,7 @@ lemma Scheme.Pullback.range_diagonal_subset_diagonalCoverDiagonalRange : rw [← hz₁, ← hy, ← Scheme.comp_base_apply, ← Scheme.comp_base_apply] dsimp only [diagonalCover, Cover.pullbackHom, Cover.bind_obj, openCoverOfLeftRight_obj] rw [← Scheme.comp_base_apply] - congr 4 + congr 5 apply pullback.hom_ext <;> simp lemma isClosedImmersion_diagonal_restrict_diagonalCoverDiagonalRange diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index 3d2e4bc11f7d87..56e5c8780547ac 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -166,7 +166,7 @@ lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [Sur · dsimp only rw [← hx₁', ← hz, ← Scheme.comp_base_apply] erw [← Scheme.comp_base_apply] - congr 4 + congr 5 apply pullback.hom_ext <;> simp [𝓤, ← pullback.condition, ← pullback.condition_assoc] · intro i have := H (S.affineOpenCover.obj i.1) (((𝒰.pullbackCover f).obj i.1).affineOpenCover.obj i.2.1) @@ -177,7 +177,7 @@ lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [Sur ((𝒲 i.1).map i.2.2 ≫ (𝒰.pullbackCover g).map i.1) (𝒰.map i.1) (by simp [pullback.condition]) (by simp [pullback.condition]) inferInstance inferInstance inferInstance - convert this using 6 + convert this using 7 apply pullback.hom_ext <;> simp [𝓤, ← pullback.condition, ← pullback.condition_assoc, Scheme.Cover.pullbackHom] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index 261a9341b50e41..7e46b285cec38c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -174,7 +174,7 @@ lemma Scheme.Hom.denseRange (f : X.Hom Y) [IsDominant f] : DenseRange f.base := instance (priority := 100) [Surjective f] : IsDominant f := ⟨f.surjective.denseRange⟩ instance [IsDominant f] [IsDominant g] : IsDominant (f ≫ g) := - ⟨g.denseRange.comp f.denseRange g.base.2⟩ + ⟨g.denseRange.comp f.denseRange g.base.hom.2⟩ instance : MorphismProperty.IsMultiplicative @IsDominant where id_mem := fun _ ↦ inferInstance diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index f78bd0243ccebc..591a430a5e75ba 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -397,11 +397,11 @@ lemma of_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) [IsOpenImmersion theorem iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : 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₂⟩ + fun ⟨h, _⟩ => IsOpenImmersion.of_stalk_iso f h⟩ theorem _root_.AlgebraicGeometry.isIso_iff_isOpenImmersion {X Y : Scheme.{u}} (f : X ⟶ Y) : IsIso f ↔ IsOpenImmersion f ∧ Epi f.base := - ⟨fun _ => ⟨inferInstance, inferInstance⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.to_iso _ _ f h₁ h₂⟩ + ⟨fun _ => ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ => IsOpenImmersion.to_iso f⟩ theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : IsIso f ↔ IsIso f.base ∧ ∀ x, IsIso (f.stalkMap x) := by diff --git a/Mathlib/AlgebraicGeometry/PointsPi.lean b/Mathlib/AlgebraicGeometry/PointsPi.lean index f430d0a60fda22..a9ebeba3e62e10 100644 --- a/Mathlib/AlgebraicGeometry/PointsPi.lean +++ b/Mathlib/AlgebraicGeometry/PointsPi.lean @@ -76,7 +76,7 @@ lemma isIso_of_comp_eq_sigmaSpec {V : Scheme} have : g.coborderRange = ⊤ := by apply eq_top_of_sigmaSpec_subset_of_isCompact (hVU := subset_coborder) · simpa only [← hU'] using Set.range_comp_subset_range f.base g.base - · exact isCompact_range g.base.2 + · exact isCompact_range g.base.hom.2 have : IsClosedImmersion g := by have : IsIso g.coborderRange.ι := by rw [this, ← Scheme.topIso_hom]; infer_instance rw [← g.liftCoborder_ι] @@ -118,7 +118,7 @@ lemma pointsPi_surjective [CompactSpace X] [∀ i, IsLocalRing (R i)] : have (i) : ∃ j, Set.range (f i).base ⊆ (𝒰.map j).opensRange := by refine ⟨𝒰.f ((f i).base (IsLocalRing.closedPoint (R i))), ?_⟩ rintro _ ⟨x, rfl⟩ - exact ((IsLocalRing.specializes_closedPoint x).map (f i).base.2).mem_open + exact ((IsLocalRing.specializes_closedPoint x).map (f i).base.hom.2).mem_open (𝒰.map _).opensRange.2 (𝒰.covers _) choose j hj using this have (j₀) := pointsPi_surjective_of_isAffine (ι := { i // j i = j₀ }) (R ·) (𝒰.obj j₀) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 91892e67145022..c496fa978f0bb2 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -212,15 +212,16 @@ section /-- The continuous function from the basic open set `D(f)` in `Proj` to the corresponding basic open set in `Spec A⁰_f`. -/ -@[simps! (config := .lemmasOnly) apply_asIdeal] -def toSpec (f : A) : (Proj.T| pbo f) ⟶ Spec.T A⁰_ f where - toFun := ToSpec.toFun f - continuous_toFun := by - rw [PrimeSpectrum.isTopologicalBasis_basic_opens.continuous_iff] - rintro _ ⟨x, rfl⟩ - obtain ⟨x, rfl⟩ := Quotient.mk''_surjective x - rw [ToSpec.preimage_basicOpen] - exact (pbo x.num).2.preimage continuous_subtype_val +@[simps! (config := .lemmasOnly) hom_apply_asIdeal] +def toSpec (f : A) : (Proj.T| pbo f) ⟶ Spec.T A⁰_ f := + TopCat.ofHom + { toFun := ToSpec.toFun f + continuous_toFun := by + rw [PrimeSpectrum.isTopologicalBasis_basic_opens.continuous_iff] + rintro _ ⟨x, rfl⟩ + obtain ⟨x, rfl⟩ := Quotient.mk''_surjective x + rw [ToSpec.preimage_basicOpen] + exact (pbo x.num).2.preimage continuous_subtype_val } variable {𝒜} in lemma toSpec_preimage_basicOpen {f} (z : HomogeneousLocalization.NumDenSameDeg 𝒜 (.powers f)) : @@ -543,21 +544,22 @@ variable {𝒜} in /-- The continuous function `Spec A⁰_f → Proj|D(f)` sending `q` to `{a | aᵢᵐ/fⁱ ∈ q}` where `m` is the degree of `f` -/ def fromSpec {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : - (Spec.T (A⁰_ f)) ⟶ (Proj.T| (pbo f)) where - toFun := FromSpec.toFun f_deg hm - continuous_toFun := by - rw [isTopologicalBasis_subtype (ProjectiveSpectrum.isTopologicalBasis_basic_opens 𝒜) (pbo f).1 - |>.continuous_iff] - rintro s ⟨_, ⟨a, rfl⟩, rfl⟩ - have h₁ : Subtype.val (p := (pbo f).1) ⁻¹' (pbo a) = - ⋃ i : ℕ, Subtype.val (p := (pbo f).1) ⁻¹' (pbo (decompose 𝒜 a i)) := by - simp [ProjectiveSpectrum.basicOpen_eq_union_of_projection 𝒜 a] - let e : _ ≃ _ := - ⟨FromSpec.toFun f_deg hm, ToSpec.toFun f, toSpec_fromSpec _ _ _, fromSpec_toSpec _ _ _⟩ - change IsOpen <| e ⁻¹' _ - rw [Set.preimage_equiv_eq_image_symm, h₁, Set.image_iUnion] - exact isOpen_iUnion fun i ↦ toSpec.image_basicOpen_eq_basicOpen f_deg hm a i ▸ - PrimeSpectrum.isOpen_basicOpen + (Spec.T (A⁰_ f)) ⟶ (Proj.T| (pbo f)) := + TopCat.ofHom + { toFun := FromSpec.toFun f_deg hm + continuous_toFun := by + rw [isTopologicalBasis_subtype (ProjectiveSpectrum.isTopologicalBasis_basic_opens 𝒜) (pbo f).1 + |>.continuous_iff] + rintro s ⟨_, ⟨a, rfl⟩, rfl⟩ + have h₁ : Subtype.val (p := (pbo f).1) ⁻¹' (pbo a) = + ⋃ i : ℕ, Subtype.val (p := (pbo f).1) ⁻¹' (pbo (decompose 𝒜 a i)) := by + simp [ProjectiveSpectrum.basicOpen_eq_union_of_projection 𝒜 a] + let e : _ ≃ _ := + ⟨FromSpec.toFun f_deg hm, ToSpec.toFun f, toSpec_fromSpec _ _ _, fromSpec_toSpec _ _ _⟩ + change IsOpen <| e ⁻¹' _ + rw [Set.preimage_equiv_eq_image_symm, h₁, Set.image_iUnion] + exact isOpen_iUnion fun i ↦ toSpec.image_basicOpen_eq_basicOpen f_deg hm a i ▸ + PrimeSpectrum.isOpen_basicOpen } end ProjIsoSpecTopComponent @@ -672,7 +674,7 @@ lemma toSpec_base_apply_eq {f} (x : Proj| pbo f) : lemma toSpec_base_isIso {f} {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : IsIso (toSpec 𝒜 f).base := by convert (projIsoSpecTopComponent f_deg hm).isIso_hom - exact DFunLike.ext _ _ <| toSpec_base_apply_eq 𝒜 + exact ConcreteCategory.hom_ext _ _ <| toSpec_base_apply_eq 𝒜 lemma mk_mem_toSpec_base_apply {f} (x : Proj| pbo f) (z : NumDenSameDeg 𝒜 (.powers f)) : diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index d82bc0979eb001..ad0701e07ad581 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -243,7 +243,7 @@ theorem Scheme.homOfLE_apply {U V : X.Opens} (e : U ≤ V) (x : U) : theorem Scheme.ι_image_homOfLE_le_ι_image {U V : X.Opens} (e : U ≤ V) (W : Opens V) : U.ι ''ᵁ (X.homOfLE e ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by simp only [← SetLike.coe_subset_coe, IsOpenMap.coe_functor_obj, Set.image_subset_iff, - Scheme.homOfLE_base, Opens.map_coe, Opens.inclusion'_apply] + Scheme.homOfLE_base, Opens.map_coe, Opens.inclusion'_hom_apply] rintro _ h exact ⟨_, h, rfl⟩ diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index ef49d929d8d24d..4d83cc76b9ee34 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -93,7 +93,7 @@ instance {X : Scheme.{u}} : Subsingleton Γ(X, ⊥) := CommRingCat.subsingleton_of_isTerminal X.sheaf.isTerminalOfEmpty @[continuity, fun_prop] -lemma Hom.continuous {X Y : Scheme} (f : X.Hom Y) : Continuous f.base := f.base.2 +lemma Hom.continuous {X Y : Scheme} (f : X.Hom Y) : Continuous f.base := f.base.hom.2 /-- The structure sheaf of a scheme. -/ protected abbrev sheaf (X : Scheme) := @@ -407,7 +407,7 @@ variable {R S : CommRingCat.{u}} (f : R ⟶ S) lemma Spec_carrier (R : CommRingCat.{u}) : (Spec R).carrier = PrimeSpectrum R := rfl lemma Spec_sheaf (R : CommRingCat.{u}) : (Spec R).sheaf = Spec.structureSheaf R := rfl lemma Spec_presheaf (R : CommRingCat.{u}) : (Spec R).presheaf = (Spec.structureSheaf R).1 := rfl -lemma Spec.map_base : (Spec.map f).base = PrimeSpectrum.comap f.hom := rfl +lemma Spec.map_base : (Spec.map f).base = ofHom (PrimeSpectrum.comap f.hom) := rfl lemma Spec.map_base_apply (x : Spec S) : (Spec.map f).base x = PrimeSpectrum.comap f.hom x := rfl lemma Spec.map_app (U) : @@ -765,12 +765,12 @@ lemma stalkMap_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : @[reassoc] lemma stalkSpecializes_stalkMap (x x' : X) - (h : x ⤳ x') : Y.presheaf.stalkSpecializes (f.base.map_specializes h) ≫ f.stalkMap x = + (h : x ⤳ x') : Y.presheaf.stalkSpecializes (f.base.hom.map_specializes h) ≫ f.stalkMap x = f.stalkMap x' ≫ X.presheaf.stalkSpecializes h := PresheafedSpace.stalkMap.stalkSpecializes_stalkMap f.toPshHom h lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : - f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.map_specializes h) y) = + f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.hom.map_specializes h) y) = (X.presheaf.stalkSpecializes h (f.stalkMap x' y)) := DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkSpecializes_stalkMap f x x' h)) y diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 26b9f817d1bf58..702a4628cbccd6 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -60,7 +60,7 @@ def Spec.topObj (R : CommRingCat.{u}) : TopCat := /-- The induced map of a ring homomorphism on the ring spectra, as a morphism of topological spaces. -/ def Spec.topMap {R S : CommRingCat.{u}} (f : R ⟶ S) : Spec.topObj S ⟶ Spec.topObj R := - PrimeSpectrum.comap f.hom + TopCat.ofHom (PrimeSpectrum.comap f.hom) @[simp] theorem Spec.topMap_id (R : CommRingCat.{u}) : Spec.topMap (𝟙 R) = 𝟙 (Spec.topObj R) := diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 54411c05943484..bb906440d71672 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -165,7 +165,7 @@ lemma range_fromSpecStalk {x : X} : ext y constructor · rintro ⟨y, rfl⟩ - exact ((IsLocalRing.specializes_closedPoint y).map (X.fromSpecStalk x).base.2).trans + exact ((IsLocalRing.specializes_closedPoint y).map (X.fromSpecStalk x).base.hom.2).trans (specializes_of_eq fromSpecStalk_closedPoint) · rintro (hy : y ⤳ x) have := fromSpecStalk_closedPoint (x := y) diff --git a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean index 46b7c013d0ec1d..56011e0fdefa0b 100644 --- a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean +++ b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean @@ -126,7 +126,7 @@ lemma specializingMap (H : ValuativeCriterion.Existence f) : dsimp only at hl₁ hl₂ refine ⟨l.base (closedPoint A), ?_, ?_⟩ · simp_rw [← Scheme.fromSpecResidueField_apply x' (closedPoint (X.residueField x')), ← hl₁] - exact (specializes_closedPoint _).map l.base.2 + exact (specializes_closedPoint _).map l.base.hom.2 · rw [← Scheme.comp_base_apply, hl₂] simp only [Scheme.comp_coeBase, TopCat.coe_comp, Function.comp_apply] have : (Spec.map stalk_y_to_A).base (closedPoint A) = closedPoint (Y.presheaf.stalk y) := @@ -304,7 +304,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique RingHom.toMorphismProperty_respectsIso_iff.mp RingHom.injective_respectsIso show P _ rw [← MorphismProperty.arrow_mk_iso_iff (P := P) e] - exact NoZeroSMulDivisors.algebraMap_injective S.R S.K + exact FaithfulSMul.algebraMap_injective S.R S.K rw [Scheme.comp_appTop] at this exact Function.Injective.of_comp this · rw [@HasAffineProperty.iff_of_isAffine @IsClosedImmersion] at this diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean index 441b8fdab6dd2d..8921a149388198 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Basic.lean @@ -336,7 +336,7 @@ def fundamentalGroupoidFunctor : TopCat ⥤ CategoryTheory.Grpd where obj X := { α := FundamentalGroupoid X } map f := { obj := fun x => ⟨f x.as⟩ - map := fun {X Y} p => by exact Path.Homotopic.Quotient.mapFn p f + map := fun {X Y} p => by exact Path.Homotopic.Quotient.mapFn p f.hom map_id := fun _ => rfl map_comp := fun {x y z} p q => by refine Quotient.inductionOn₂ p q fun a b => ?_ @@ -367,7 +367,7 @@ scoped notation "πₓ" => FundamentalGroupoid.fundamentalGroupoidFunctor.obj scoped notation "πₘ" => FundamentalGroupoid.fundamentalGroupoidFunctor.map theorem map_eq {X Y : TopCat} {x₀ x₁ : X} (f : C(X, Y)) (p : Path.Homotopic.Quotient x₀ x₁) : - (πₘ f).map p = p.mapFn f := rfl + (πₘ (TopCat.ofHom f)).map p = p.mapFn f := rfl /-- Help the typechecker by converting a point in a groupoid back to a point in the underlying topological space. -/ diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean index 4ba7cc3bd09126..a4536f3e7d25e9 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean @@ -87,7 +87,8 @@ include hfg /-- If `f(p(t) = g(q(t))` for two paths `p` and `q`, then the induced path homotopy classes `f(p)` and `g(p)` are the same as well, despite having a priori different types -/ -theorem heq_path_of_eq_image : HEq ((πₘ f).map ⟦p⟧) ((πₘ g).map ⟦q⟧) := by +theorem heq_path_of_eq_image : + HEq ((πₘ (TopCat.ofHom f)).map ⟦p⟧) ((πₘ (TopCat.ofHom g)).map ⟦q⟧) := by simp only [map_eq, ← Path.Homotopic.map_lift]; apply Path.Homotopic.hpath_hext; exact hfg private theorem start_path : f x₀ = g x₂ := by convert hfg 0 <;> simp only [Path.source] @@ -95,9 +96,10 @@ private theorem start_path : f x₀ = g x₂ := by convert hfg 0 <;> simp only [ private theorem end_path : f x₁ = g x₃ := by convert hfg 1 <;> simp only [Path.target] theorem eq_path_of_eq_image : - (πₘ f).map ⟦p⟧ = hcast (start_path hfg) ≫ (πₘ g).map ⟦q⟧ ≫ hcast (end_path hfg).symm := by + (πₘ (TopCat.ofHom f)).map ⟦p⟧ = + hcast (start_path hfg) ≫ (πₘ (TopCat.ofHom g)).map ⟦q⟧ ≫ hcast (end_path hfg).symm := by rw [conj_eqToHom_iff_heq - ((πₘ f).map ⟦p⟧) ((πₘ g).map ⟦q⟧) + ((πₘ (TopCat.ofHom f)).map ⟦p⟧) ((πₘ (TopCat.ofHom g)).map ⟦q⟧) (FundamentalGroupoid.ext <| start_path hfg) (FundamentalGroupoid.ext <| end_path hfg)] exact heq_path_of_eq_image hfg @@ -147,33 +149,37 @@ abbrev prodToProdTopI {a₁ a₂ : TopCat.of (ULift I)} {b₁ b₂ : X} (p₁ : /-- The diagonal path `d` of a homotopy `H` on a path `p` -/ def diagonalPath : fromTop (H (0, x₀)) ⟶ fromTop (H (1, x₁)) := - (πₘ H.uliftMap).map (prodToProdTopI uhpath01 p) + (πₘ (TopCat.ofHom H.uliftMap)).map (prodToProdTopI uhpath01 p) /-- The diagonal path, but starting from `f x₀` and going to `g x₁` -/ def diagonalPath' : fromTop (f x₀) ⟶ fromTop (g x₁) := hcast (H.apply_zero x₀).symm ≫ H.diagonalPath p ≫ hcast (H.apply_one x₁) /-- Proof that `f(p) = H(0 ⟶ 0, p)`, with the appropriate casts -/ -theorem apply_zero_path : (πₘ f).map p = hcast (H.apply_zero x₀).symm ≫ - (πₘ H.uliftMap).map (prodToProdTopI (𝟙 (@fromTop (TopCat.of _) (ULift.up 0))) p) ≫ +theorem apply_zero_path : (πₘ (TopCat.ofHom f)).map p = hcast (H.apply_zero x₀).symm ≫ + (πₘ (TopCat.ofHom H.uliftMap)).map + (prodToProdTopI (𝟙 (@fromTop (TopCat.of _) (ULift.up 0))) p) ≫ hcast (H.apply_zero x₁) := Quotient.inductionOn p fun p' => by apply @eq_path_of_eq_image _ _ _ _ H.uliftMap _ _ _ _ _ ((Path.refl (ULift.up _)).prod p') - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [Path.prod_coe]; simp_rw [ulift_apply]; simp + intros + rw [Path.prod_coe, ulift_apply H] + simp /-- Proof that `g(p) = H(1 ⟶ 1, p)`, with the appropriate casts -/ -theorem apply_one_path : (πₘ g).map p = hcast (H.apply_one x₀).symm ≫ - (πₘ H.uliftMap).map (prodToProdTopI (𝟙 (@fromTop (TopCat.of _) (ULift.up 1))) p) ≫ +theorem apply_one_path : (πₘ (TopCat.ofHom g)).map p = hcast (H.apply_one x₀).symm ≫ + (πₘ (TopCat.ofHom H.uliftMap)).map + (prodToProdTopI (𝟙 (@fromTop (TopCat.of _) (ULift.up 1))) p) ≫ hcast (H.apply_one x₁) := Quotient.inductionOn p fun p' => by apply @eq_path_of_eq_image _ _ _ _ H.uliftMap _ _ _ _ _ ((Path.refl (ULift.up _)).prod p') - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [Path.prod_coe]; simp_rw [ulift_apply]; simp + intros + rw [Path.prod_coe, ulift_apply H] + simp /-- Proof that `H.evalAt x = H(0 ⟶ 1, x ⟶ x)`, with the appropriate casts -/ theorem evalAt_eq (x : X) : ⟦H.evalAt x⟧ = hcast (H.apply_zero x).symm ≫ - (πₘ H.uliftMap).map (prodToProdTopI uhpath01 (𝟙 (fromTop x))) ≫ + (πₘ (TopCat.ofHom H.uliftMap)).map (prodToProdTopI uhpath01 (𝟙 (fromTop x))) ≫ hcast (H.apply_one x).symm.symm := by dsimp only [prodToProdTopI, uhpath01, hcast] refine (@conj_eqToHom_iff_heq (πₓ Y) _ _ _ _ _ _ _ _ @@ -183,8 +189,9 @@ theorem evalAt_eq (x : X) : ⟦H.evalAt x⟧ = hcast (H.apply_zero x).symm ≫ apply Path.Homotopic.hpath_hext; intro; rfl -- Finally, we show `d = f(p) ≫ H₁ = H₀ ≫ g(p)` -theorem eq_diag_path : (πₘ f).map p ≫ ⟦H.evalAt x₁⟧ = H.diagonalPath' p ∧ - (⟦H.evalAt x₀⟧ ≫ (πₘ g).map p : fromTop (f x₀) ⟶ fromTop (g x₁)) = H.diagonalPath' p := by +theorem eq_diag_path : (πₘ (TopCat.ofHom f)).map p ≫ ⟦H.evalAt x₁⟧ = H.diagonalPath' p ∧ + (⟦H.evalAt x₀⟧ ≫ (πₘ (TopCat.ofHom g)).map p : + fromTop (f x₀) ⟶ fromTop (g x₁)) = H.diagonalPath' p := by rw [H.apply_zero_path, H.apply_one_path, H.evalAt_eq] erw [H.evalAt_eq] -- Porting note: `rw` didn't work, so using `erw` dsimp only [prodToProdTopI] @@ -211,7 +218,9 @@ variable {X Y : TopCat.{u}} {f g : C(X, Y)} (H : ContinuousMap.Homotopy f g) /-- Given a homotopy H : f ∼ g, we have an associated natural isomorphism between the induced functors `f` and `g` -/ -- Porting note: couldn't use category arrow `\hom` in statement, needed to expand -def homotopicMapsNatIso : @Quiver.Hom _ Functor.category.toQuiver (πₘ f) (πₘ g) where +def homotopicMapsNatIso : @Quiver.Hom _ Functor.category.toQuiver + (πₘ (TopCat.ofHom f)) + (πₘ (TopCat.ofHom g)) where app x := ⟦H.evalAt x.as⟧ -- Porting note: Turned `rw` into `erw` in the line below naturality x y p := by erw [(H.eq_diag_path p).1, (H.eq_diag_path p).2] @@ -222,8 +231,8 @@ open scoped ContinuousMap /-- Homotopy equivalent topological spaces have equivalent fundamental groupoids. -/ def equivOfHomotopyEquiv (hequiv : X ≃ₕ Y) : πₓ X ≌ πₓ Y := by - apply CategoryTheory.Equivalence.mk (πₘ hequiv.toFun : πₓ X ⥤ πₓ Y) - (πₘ hequiv.invFun : πₓ Y ⥤ πₓ X) <;> + apply CategoryTheory.Equivalence.mk (πₘ (TopCat.ofHom hequiv.toFun) : πₓ X ⥤ πₓ Y) + (πₘ (TopCat.ofHom hequiv.invFun) : πₓ Y ⥤ πₓ X) <;> simp only [Grpd.hom_to_functor, Grpd.id_to_functor] · convert (asIso (homotopicMapsNatIso hequiv.left_inv.some)).symm exacts [((π).map_id X).symm, ((π).map_comp _ _).symm] diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean index 2877b1ad5f01d7..6578d88c6565e0 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean @@ -40,7 +40,7 @@ variable {I : Type u} (X : I → TopCat.{u}) /-- The projection map Π i, X i → X i induces a map π(Π i, X i) ⟶ π(X i). -/ def proj (i : I) : πₓ (TopCat.of (∀ i, X i)) ⥤ πₓ (X i) := - πₘ ⟨_, continuous_apply i⟩ + πₘ (TopCat.ofHom ⟨_, continuous_apply i⟩) /-- The projection map is precisely `Path.Homotopic.proj` interpreted as a functor -/ @[simp] @@ -125,11 +125,11 @@ variable (A B : TopCat.{u}) /-- The induced map of the left projection map X × Y → X -/ def projLeft : πₓ (TopCat.of (A × B)) ⥤ πₓ A := - πₘ ⟨_, continuous_fst⟩ + πₘ (TopCat.ofHom ⟨_, continuous_fst⟩) /-- The induced map of the right projection map X × Y → Y -/ def projRight : πₓ (TopCat.of (A × B)) ⥤ πₓ B := - πₘ ⟨_, continuous_snd⟩ + πₘ (TopCat.ofHom ⟨_, continuous_snd⟩) @[simp] theorem projLeft_map (x₀ x₁ : πₓ (TopCat.of (A × B))) (p : x₀ ⟶ x₁) : diff --git a/Mathlib/AlgebraicTopology/TopologicalSimplex.lean b/Mathlib/AlgebraicTopology/TopologicalSimplex.lean index a3767b7e54152f..9e03d618ef9648 100644 --- a/Mathlib/AlgebraicTopology/TopologicalSimplex.lean +++ b/Mathlib/AlgebraicTopology/TopologicalSimplex.lean @@ -66,18 +66,18 @@ theorem continuous_toTopMap {x y : SimplexCategory} (f : x ⟶ y) : Continuous ( @[simps obj map] def toTop : SimplexCategory ⥤ TopCat where obj x := TopCat.of x.toTopObj - map f := ⟨toTopMap f, by continuity⟩ + map f := TopCat.ofHom ⟨toTopMap f, by continuity⟩ map_id := by classical intro Δ - ext f + ext f : 1 apply toTopObj.ext funext i change (Finset.univ.filter (· = i)).sum _ = _ simp [Finset.sum_filter, CategoryTheory.id_apply] map_comp := fun f g => by classical - ext h + ext h : 1 apply toTopObj.ext funext i dsimp diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 6844e170f22e95..dea028ad9f7471 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -1062,7 +1062,8 @@ theorem length_sigmaCompositionAux (a : Composition n) (b : Composition a.length Composition.length (Composition.sigmaCompositionAux a b ⟨i, (length_gather a b).symm ▸ i.2⟩) = Composition.blocksFun b i := show List.length ((splitWrtComposition a.blocks b)[i.1]) = blocksFun b i by - rw [getElem_map_rev List.length, getElem_of_eq (map_length_splitWrtComposition _ _)]; rfl + rw [getElem_map_rev List.length, getElem_of_eq (map_length_splitWrtComposition _ _), blocksFun, + get_eq_getElem] theorem blocksFun_sigmaCompositionAux (a : Composition n) (b : Composition a.length) (i : Fin b.length) (j : Fin (blocksFun b i)) : diff --git a/Mathlib/Analysis/Analytic/IsolatedZeros.lean b/Mathlib/Analysis/Analytic/IsolatedZeros.lean index 0c12a9fea2373d..477f764d36f0e5 100644 --- a/Mathlib/Analysis/Analytic/IsolatedZeros.lean +++ b/Mathlib/Analysis/Analytic/IsolatedZeros.lean @@ -345,6 +345,49 @@ theorem eq_of_frequently_eq [ConnectedSpace 𝕜] (hf : AnalyticOnNhd 𝕜 f uni @[deprecated (since := "2024-09-26")] alias _root_.AnalyticOn.eq_of_frequently_eq := eq_of_frequently_eq +/-- The set where an analytic function has infinite order is clopen in its domain of analyticity. -/ +theorem isClopen_setOf_order_eq_top (h₁f : AnalyticOnNhd 𝕜 f U) : + IsClopen { u : U | (h₁f u.1 u.2).order = ⊤ } := by + constructor + · rw [← isOpen_compl_iff, isOpen_iff_forall_mem_open] + intro z hz + rcases (h₁f z.1 z.2).eventually_eq_zero_or_eventually_ne_zero with h | h + · -- Case: f is locally zero in a punctured neighborhood of z + rw [← (h₁f z.1 z.2).order_eq_top_iff] at h + tauto + · -- Case: f is locally nonzero in a punctured neighborhood of z + obtain ⟨t', h₁t', h₂t', h₃t'⟩ := eventually_nhds_iff.1 (eventually_nhdsWithin_iff.1 h) + use Subtype.val ⁻¹' t' + constructor + · intro w hw + simp only [mem_compl_iff, mem_setOf_eq] + by_cases h₁w : w = z + · rwa [h₁w] + · rw [(h₁f _ w.2).order_eq_zero_iff.2 ((h₁t' w hw) (Subtype.coe_ne_coe.mpr h₁w))] + exact ENat.zero_ne_top + · exact ⟨isOpen_induced h₂t', h₃t'⟩ + · apply isOpen_iff_forall_mem_open.mpr + intro z hz + conv => + arg 1; intro; left; right; arg 1; intro + rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] + simp only [Set.mem_setOf_eq] at hz + rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] at hz + obtain ⟨t', h₁t', h₂t', h₃t'⟩ := hz + use Subtype.val ⁻¹' t' + simp only [Set.mem_compl_iff, Set.mem_singleton_iff, isOpen_induced h₂t', Set.mem_preimage, + h₃t', and_self, and_true] + intro w hw + simp only [mem_setOf_eq] + -- Trivial case: w = z + by_cases h₁w : w = z + · rw [h₁w] + tauto + -- Nontrivial case: w ≠ z + use t' \ {z.1}, fun y h₁y ↦ h₁t' y h₁y.1, h₂t'.sdiff isClosed_singleton + apply (Set.mem_diff w).1 + exact ⟨hw, Set.mem_singleton_iff.not.1 (Subtype.coe_ne_coe.2 h₁w)⟩ + section Mul /-! ### Vanishing of products of analytic functions diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index ff69888be1dc23..caf3414b30d5f6 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Loeffler +Authors: David Loeffler, Stefan Kebekus -/ import Mathlib.Analysis.Analytic.IsolatedZeros import Mathlib.Algebra.Order.AddGroupWithTop @@ -35,6 +35,23 @@ lemma AnalyticAt.meromorphicAt {f : 𝕜 → E} {x : 𝕜} (hf : AnalyticAt 𝕜 MeromorphicAt f x := ⟨0, by simpa only [pow_zero, one_smul]⟩ +/- Analogue of the principle of isolated zeros for an analytic function: if a function is +meromorphic at `z₀`, then either it is identically zero in a punctured neighborhood of `z₀`, or it +does not vanish there at all. -/ +theorem MeromorphicAt.eventually_eq_zero_or_eventually_ne_zero {f : 𝕜 → E} {z₀ : 𝕜} + (hf : MeromorphicAt f z₀) : + (∀ᶠ z in 𝓝[≠] z₀, f z = 0) ∨ (∀ᶠ z in 𝓝[≠] z₀, f z ≠ 0) := by + obtain ⟨n, h⟩ := hf + rcases h.eventually_eq_zero_or_eventually_ne_zero with h₁ | h₂ + · left + filter_upwards [nhdsWithin_le_nhds h₁, self_mem_nhdsWithin] with y h₁y h₂y + rcases (smul_eq_zero.1 h₁y) with h₃ | h₄ + · exact False.elim (h₂y (sub_eq_zero.1 (pow_eq_zero_iff'.1 h₃).1)) + · assumption + · right + filter_upwards [h₂, self_mem_nhdsWithin] with y h₁y h₂y + exact (smul_ne_zero_iff.1 h₁y).2 + namespace MeromorphicAt lemma id (x : 𝕜) : MeromorphicAt id x := analyticAt_id.meromorphicAt @@ -334,6 +351,82 @@ lemma id {U : Set 𝕜} : MeromorphicOn id U := fun x _ ↦ .id x lemma const (e : E) {U : Set 𝕜} : MeromorphicOn (fun _ ↦ e) U := fun x _ ↦ .const e x +/-- The set where a meromorphic function has infinite order is clopen in its domain of meromorphy. +-/ +theorem isClopen_setOf_order_eq_top {U : Set 𝕜} (hf : MeromorphicOn f U) : + IsClopen { u : U | (hf u.1 u.2).order = ⊤ } := by + constructor + · rw [← isOpen_compl_iff, isOpen_iff_forall_mem_open] + intro z hz + rcases (hf z.1 z.2).eventually_eq_zero_or_eventually_ne_zero with h | h + · -- Case: f is locally zero in a punctured neighborhood of z + rw [← (hf z.1 z.2).order_eq_top_iff] at h + tauto + · -- Case: f is locally nonzero in a punctured neighborhood of z + obtain ⟨t', h₁t', h₂t', h₃t'⟩ := eventually_nhds_iff.1 (eventually_nhdsWithin_iff.1 h) + use Subtype.val ⁻¹' t' + constructor + · intro w hw + simp only [Set.mem_compl_iff, Set.mem_setOf_eq] + by_cases h₁w : w = z + · rwa [h₁w] + · rw [MeromorphicAt.order_eq_top_iff, not_eventually] + apply Filter.Eventually.frequently + rw [eventually_nhdsWithin_iff, eventually_nhds_iff] + use t' \ {z.1}, fun y h₁y h₂y ↦ h₁t' y h₁y.1 h₁y.2, h₂t'.sdiff isClosed_singleton, hw, + Set.mem_singleton_iff.not.2 (Subtype.coe_ne_coe.mpr h₁w) + · exact ⟨isOpen_induced h₂t', h₃t'⟩ + · apply isOpen_iff_forall_mem_open.mpr + intro z hz + conv => + arg 1; intro; left; right; arg 1; intro + rw [MeromorphicAt.order_eq_top_iff, eventually_nhdsWithin_iff, eventually_nhds_iff] + simp only [Set.mem_setOf_eq] at hz + rw [MeromorphicAt.order_eq_top_iff, eventually_nhdsWithin_iff, eventually_nhds_iff] at hz + obtain ⟨t', h₁t', h₂t', h₃t'⟩ := hz + use Subtype.val ⁻¹' t' + simp only [Set.mem_compl_iff, Set.mem_singleton_iff, isOpen_induced h₂t', Set.mem_preimage, + h₃t', and_self, and_true] + intro w hw + simp only [Set.mem_setOf_eq] + -- Trivial case: w = z + by_cases h₁w : w = z + · rw [h₁w] + tauto + -- Nontrivial case: w ≠ z + use t' \ {z.1}, fun y h₁y _ ↦ h₁t' y (Set.mem_of_mem_diff h₁y) (Set.mem_of_mem_inter_right h₁y) + constructor + · exact h₂t'.sdiff isClosed_singleton + · apply (Set.mem_diff w).1 + exact ⟨hw, Set.mem_singleton_iff.not.1 (Subtype.coe_ne_coe.2 h₁w)⟩ + +/-- On a connected set, there exists a point where a meromorphic function `f` has finite order iff +`f` has finite order at every point. -/ +theorem exists_order_ne_top_iff_forall {U : Set 𝕜} (hf : MeromorphicOn f U) (hU : IsConnected U) : + (∃ u : U, (hf u u.2).order ≠ ⊤) ↔ (∀ u : U, (hf u u.2).order ≠ ⊤) := by + constructor + · intro h₂f + have := isPreconnected_iff_preconnectedSpace.1 hU.isPreconnected + rcases isClopen_iff.1 hf.isClopen_setOf_order_eq_top with h | h + · intro u + have : u ∉ (∅ : Set U) := by exact fun a => a + rw [← h] at this + tauto + · obtain ⟨u, hU⟩ := h₂f + have : u ∈ Set.univ := by trivial + rw [← h] at this + tauto + · intro h₂f + obtain ⟨v, hv⟩ := hU.nonempty + use ⟨v, hv⟩, h₂f ⟨v, hv⟩ + +/-- On a preconnected set, a meromorphic function has finite order at one point if it has finite +order at another point. -/ +theorem order_ne_top_of_isPreconnected {U : Set 𝕜} {x y : 𝕜} (hf : MeromorphicOn f U) + (hU : IsPreconnected U) (h₁x : x ∈ U) (hy : y ∈ U) (h₂x : (hf x h₁x).order ≠ ⊤) : + (hf y hy).order ≠ ⊤ := + (hf.exists_order_ne_top_iff_forall ⟨Set.nonempty_of_mem h₁x, hU⟩).1 (by use ⟨x, h₁x⟩) ⟨y, hy⟩ + section arithmetic include hf in diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index 62ae49c6840de3..bab4c2a4b7b23e 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -282,7 +282,7 @@ lemma CStarAlgebra.isUnit_of_le {a b : A} (h₀ : IsUnit a) (ha : 0 ≤ a := by nontriviality A have hb := (show 0 ≤ a from ha).trans hab 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 (.of_nonneg ‹_›)] at h₀ ⊢ peel h₀ with r hr _ exact this.trans hab @@ -399,14 +399,15 @@ lemma norm_le_norm_of_nonneg_of_le {a b : A} (ha : 0 ≤ a := by cfc_tac) (hab : simpa only [ge_iff_le, Unitization.norm_inr] using this a b (by simpa) (by rwa [Unitization.inr_le_iff a b]) intro a b ha hab - have hb_nonneg : 0 ≤ b := ha.trans hab - have : 0 ≤ a := by cfc_tac + have hb : 0 ≤ b := ha.trans hab + -- these two `have`s are just for performance + have := IsSelfAdjoint.of_nonneg ha; have := IsSelfAdjoint.of_nonneg hb have h₂ : cfc (id : ℝ → ℝ) a ≤ cfc (fun _ => ‖b‖) a := by calc _ = a := by rw [cfc_id ℝ a] _ ≤ cfc id b := (cfc_id ℝ b) ▸ hab _ ≤ cfc (fun _ => ‖b‖) b := by refine cfc_mono fun x hx => ?_ - calc x = ‖x‖ := (Real.norm_of_nonneg (spectrum_nonneg_of_nonneg hb_nonneg hx)).symm + calc x = ‖x‖ := (Real.norm_of_nonneg (spectrum_nonneg_of_nonneg hb hx)).symm _ ≤ ‖b‖ := spectrum.norm_le_norm_of_mem hx _ = _ := by rw [cfc_const _ _, cfc_const _ _] rw [cfc_le_iff id (fun _ => ‖b‖) a] at h₂ diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index e5879e948ab45d..7ca5af36fcc5e0 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -140,11 +140,10 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L have isoF : Fu ≃ₗᵢ[𝕜] F := LinearIsometryEquiv.ulift 𝕜 F have isoG : Gu ≃ₗᵢ[𝕜] G := LinearIsometryEquiv.ulift 𝕜 G -- lift `f` and `g` to versions `fu` and `gu` on the lifted spaces. - set fu : Du → Eu := isoE.symm ∘ f ∘ isoD with hfu - set gu : Du → Fu := isoF.symm ∘ g ∘ isoD with hgu + let fu : Du → Eu := isoE.symm ∘ f ∘ isoD + let gu : Du → Fu := isoF.symm ∘ g ∘ isoD -- lift the bilinear map `B` to a bilinear map `Bu` on the lifted spaces. - set Bu₀ : Eu →L[𝕜] Fu →L[𝕜] G := ((B.comp (isoE : Eu →L[𝕜] E)).flip.comp (isoF : Fu →L[𝕜] F)).flip - with hBu₀ + let Bu₀ : Eu →L[𝕜] Fu →L[𝕜] G := ((B.comp (isoE : Eu →L[𝕜] E)).flip.comp (isoF : Fu →L[𝕜] F)).flip let Bu : Eu →L[𝕜] Fu →L[𝕜] Gu := ContinuousLinearMap.compL 𝕜 Eu (Fu →L[𝕜] G) (Fu →L[𝕜] Gu) (ContinuousLinearMap.compL 𝕜 Fu G Gu (isoG.symm : G →L[𝕜] Gu)) Bu₀ @@ -152,12 +151,12 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L (ContinuousLinearMap.compL 𝕜 Fu G Gu (isoG.symm : G →L[𝕜] Gu)) Bu₀ := rfl have Bu_eq : (fun y => Bu (fu y) (gu y)) = isoG.symm ∘ (fun y => B (f y) (g y)) ∘ isoD := by ext1 y - simp [Du, Eu, Fu, Gu, hBu, hBu₀, hfu, hgu] + simp [Du, Eu, Fu, Gu, hBu, Bu₀, fu, gu] -- All norms are preserved by the lifting process. have Bu_le : ‖Bu‖ ≤ ‖B‖ := by refine ContinuousLinearMap.opNorm_le_bound _ (norm_nonneg B) fun y => ?_ refine ContinuousLinearMap.opNorm_le_bound _ (by positivity) fun x => ?_ - simp only [Du, Eu, Fu, Gu, hBu, hBu₀, compL_apply, coe_comp', Function.comp_apply, + simp only [Du, Eu, Fu, Gu, hBu, Bu₀, compL_apply, coe_comp', Function.comp_apply, ContinuousLinearEquiv.coe_coe, LinearIsometryEquiv.coe_coe, flip_apply, LinearIsometryEquiv.norm_map] calc diff --git a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean index 892ed8c6dfa95e..1951867327661b 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Prod.lean @@ -352,7 +352,7 @@ theorem hasStrictFDerivAt_apply (i : ι) (f : ∀ i, F' i) : let id' := ContinuousLinearMap.id 𝕜 (∀ i, F' i) have h := ((hasStrictFDerivAt_pi' (Φ := fun (f : ∀ i, F' i) (i' : ι) => f i') (Φ' := id') (x := f))).1 - have h' : comp (proj i) id' = proj i := by rfl + have h' : comp (proj i) id' = proj i := by ext; simp [id'] rw [← h']; apply h; apply hasStrictFDerivAt_id @[simp 1100] -- Porting note: increased priority to make lint happy diff --git a/Mathlib/Analysis/Convex/StrictConvexBetween.lean b/Mathlib/Analysis/Convex/StrictConvexBetween.lean index 08a6eed3ee484e..55871c146e1561 100644 --- a/Mathlib/Analysis/Convex/StrictConvexBetween.lean +++ b/Mathlib/Analysis/Convex/StrictConvexBetween.lean @@ -104,7 +104,7 @@ variable {E F PE PF : Type*} [NormedAddCommGroup E] [NormedAddCommGroup F] [Norm lemma eq_lineMap_of_dist_eq_mul_of_dist_eq_mul (hxy : dist x y = r * dist x z) (hyz : dist y z = (1 - r) * dist x z) : y = AffineMap.lineMap x z r := by have : y -ᵥ x ∈ [(0 : E) -[ℝ] z -ᵥ x] := by - rw [mem_segment_iff_wbtw, ← dist_add_dist_eq_iff, dist_zero_left, dist_vsub_cancel_right, + rw [mem_segment_iff_wbtw, ← dist_add_dist_eq_iff, dist_zero, dist_vsub_cancel_right, ← dist_eq_norm_vsub', ← dist_eq_norm_vsub', hxy, hyz, ← add_mul, add_sub_cancel, one_mul] obtain rfl | hne := eq_or_ne x z diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 09b5b24238fb40..a4062257ae2d47 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -14,7 +14,8 @@ import Mathlib.Data.Real.ConjExponents In this file we prove several inequalities for finite sums, including AM-GM inequality, HM-GM inequality, Young's inequality, Hölder inequality, and Minkowski inequality. Versions for -integrals of some of these inequalities are available in `MeasureTheory.Integral.MeanInequalities`. +integrals of some of these inequalities are available in +`Mathlib.MeasureTheory.Integral.MeanInequalities`. ## Main theorems diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 7bb024ac481353..4373faf39f635e 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -387,9 +387,11 @@ theorem dist_one_right (a : E) : dist a 1 = ‖a‖ := by rw [dist_eq_norm_div, theorem inseparable_one_iff_norm {a : E} : Inseparable a 1 ↔ ‖a‖ = 0 := by rw [Metric.inseparable_iff, dist_one_right] +@[to_additive] +lemma dist_one_left (a : E) : dist 1 a = ‖a‖ := by rw [dist_comm, dist_one_right] + @[to_additive (attr := simp)] -theorem dist_one_left : dist (1 : E) = norm := - funext fun a => by rw [dist_comm, dist_one_right] +lemma dist_one : dist (1 : E) = norm := funext dist_one_left @[to_additive] theorem norm_div_rev (a b : E) : ‖a / b‖ = ‖b / a‖ := by @@ -667,14 +669,14 @@ theorem coe_nnnorm' (a : E) : (‖a‖₊ : ℝ) = ‖a‖ := rfl theorem coe_comp_nnnorm' : (toReal : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥0) = norm := rfl -@[to_additive norm_toNNReal] +@[to_additive (attr := simp) norm_toNNReal] theorem norm_toNNReal' : ‖a‖.toNNReal = ‖a‖₊ := @Real.toNNReal_coe ‖a‖₊ -@[to_additive toReal_enorm] +@[to_additive (attr := simp) toReal_enorm] lemma toReal_enorm' (x : E) : ‖x‖ₑ.toReal = ‖x‖ := by simp [enorm] -@[to_additive ofReal_norm] +@[to_additive (attr := simp) ofReal_norm] lemma ofReal_norm' (x : E) : .ofReal ‖x‖ = ‖x‖ₑ := by simp [enorm, ENNReal.ofReal, Real.toNNReal, nnnorm] diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index ca6dd25e1c9bba..d2feee682ef6b3 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -159,7 +159,7 @@ end NNNorm @[to_additive lipschitzWith_one_norm] theorem lipschitzWith_one_norm' : LipschitzWith 1 (norm : E → ℝ) := by - simpa only [dist_one_left] using LipschitzWith.dist_right (1 : E) + simpa using LipschitzWith.dist_right (1 : E) @[to_additive lipschitzWith_one_nnnorm] theorem lipschitzWith_one_nnnorm' : LipschitzWith 1 (NNNorm.nnnorm : E → ℝ≥0) := diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index efaf322fb9c49b..a96b068f2f45a2 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -879,6 +879,19 @@ theorem ofReal_mul_neg_iff (x : ℝ) (z : K) : x * z < 0 ↔ (x < 0 ∧ 0 < z) ∨ (0 < x ∧ z < 0) := by simpa only [mul_neg, neg_pos, neg_neg_iff_pos] using ofReal_mul_pos_iff x (-z) +lemma instPosMulReflectLE : PosMulReflectLE K := by + constructor + intro a b c (h : _ * _ ≤ _ * _) + obtain ⟨a', ha1, ha2⟩ := pos_iff_exists_ofReal.mp a.2 + rw [← sub_nonneg] + rw [← ha2, ← sub_nonneg, ← mul_sub, le_iff_lt_or_eq] at h + rcases h with h | h + · rw [ofReal_mul_pos_iff] at h + exact le_of_lt <| h.rec (False.elim <| not_lt_of_gt ·.1 ha1) (·.2) + · exact ((mul_eq_zero_iff_left <| ofReal_ne_zero.mpr ha1.ne').mp h.symm).ge + +scoped[ComplexOrder] attribute [instance] RCLike.instPosMulReflectLE + end Order section CleanupLemmas diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index 403dd4f36d6f95..95dd9afb25e88b 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -127,8 +127,7 @@ theorem deriv.log (hf : DifferentiableAt ℝ f x) (hx : f x ≠ 0) : `f x ≠ 0`. -/ lemma Real.deriv_log_comp_eq_logDeriv {f : ℝ → ℝ} {x : ℝ} (h₁ : DifferentiableAt ℝ f x) (h₂ : f x ≠ 0) : deriv (log ∘ f) x = logDeriv f x := by - simp only [ne_eq, logDeriv, Pi.div_apply, ← deriv.log h₁ h₂] - rfl + simp only [ne_eq, logDeriv, Pi.div_apply, ← deriv.log h₁ h₂, Function.comp_def] end deriv diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean index 1e9791803def93..f66f37fbb5a9c2 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean @@ -199,8 +199,7 @@ theorem isBigO_cpow_rpow (hl : IsBoundedUnder (· ≤ ·) l fun x => |(g x).im|) _ =Θ[l] (show α → ℝ from fun x => abs (f x) ^ (g x).re / (1 : ℝ)) := ((isTheta_refl _ _).div (isTheta_exp_arg_mul_im hl)) _ =ᶠ[l] (show α → ℝ from fun x => abs (f x) ^ (g x).re) := by - simp only [ofReal_one, div_one] - rfl + simp only [ofReal_one, div_one, EventuallyEq.rfl] theorem isTheta_cpow_rpow (hl_im : IsBoundedUnder (· ≤ ·) l fun x => |(g x).im|) (hl : ∀ᶠ x in l, f x = 0 → re (g x) = 0 → g x = 0) : @@ -212,8 +211,7 @@ theorem isTheta_cpow_rpow (hl_im : IsBoundedUnder (· ≤ ·) l fun x => |(g x). _ =Θ[l] fun x => abs (f x) ^ (g x).re / (1 : ℝ) := (isTheta_refl _ _).div (isTheta_exp_arg_mul_im hl_im) _ =ᶠ[l] (fun x => abs (f x) ^ (g x).re) := by - simp only [ofReal_one, div_one] - rfl + simp only [ofReal_one, div_one, EventuallyEq.rfl] theorem isTheta_cpow_const_rpow {b : ℂ} (hl : b.re = 0 → b ≠ 0 → ∀ᶠ x in l, f x ≠ 0) : (fun x => f x ^ b) =Θ[l] fun x => abs (f x) ^ b.re := diff --git a/Mathlib/CategoryTheory/Abelian/Images.lean b/Mathlib/CategoryTheory/Abelian/Images.lean index 8c3bf7213ffe12..0b4b6a16873afe 100644 --- a/Mathlib/CategoryTheory/Abelian/Images.lean +++ b/Mathlib/CategoryTheory/Abelian/Images.lean @@ -105,7 +105,7 @@ theorem coimage_image_factorisation : coimage.π f ≫ coimageImageComparison f simp [coimageImageComparison] /-- The coimage-image comparison morphism is functorial. -/ -@[simps!] +@[simps! obj map] def coimageImageComparisonFunctor : Arrow C ⥤ Arrow C where obj f := Arrow.mk (coimageImageComparison f.hom) map {f g} η := Arrow.homMk diff --git a/Mathlib/CategoryTheory/Abelian/Refinements.lean b/Mathlib/CategoryTheory/Abelian/Refinements.lean index 75582a7bcc0249..9b37b8436a3d69 100644 --- a/Mathlib/CategoryTheory/Abelian/Refinements.lean +++ b/Mathlib/CategoryTheory/Abelian/Refinements.lean @@ -26,7 +26,7 @@ In this file, the basic result is `epi_iff_surjective_up_to_refinements` which states that `f : X ⟶ Y` is a morphism in an abelian category, then it is an epimorphism if and only if for all `y : A ⟶ Y`, there exists an epimorphism `π : A' ⟶ A` and `x : A' ⟶ X` such -that `π ≫ y = x ≫ f`. In order words, if we allow a precomposition +that `π ≫ y = x ≫ f`. In other words, if we allow a precomposition with an epimorphism, we may lift a morphism to `Y` to a morphism to `X`. Following unpublished notes by George Bergman, we shall say that the precomposition by an epimorphism `π ≫ y` is a refinement of `y`. Then, diff --git a/Mathlib/CategoryTheory/Abelian/Transfer.lean b/Mathlib/CategoryTheory/Abelian/Transfer.lean index 114d4d2a53902d..bc71c7a1b37c8b 100644 --- a/Mathlib/CategoryTheory/Abelian/Transfer.lean +++ b/Mathlib/CategoryTheory/Abelian/Transfer.lean @@ -3,11 +3,9 @@ Copyright (c) 2022 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Algebra.Equiv.TransferInstance import Mathlib.CategoryTheory.Abelian.Basic import Mathlib.CategoryTheory.Adjunction.Limits -import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Kernels -import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.AbelianImages import Mathlib.CategoryTheory.Preadditive.Transfer /-! @@ -73,99 +71,32 @@ theorem hasCokernels (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) : HasCokernels C : haveI : HasCokernel (G.map (F.map f) ≫ i.hom.app _) := Limits.hasCokernel_comp_iso _ _ apply Limits.hasCokernel_epi_comp } -variable [Limits.HasCokernels C] - -/-- Auxiliary construction for `coimageIsoImage` -/ -def cokernelIso (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : - G.obj (cokernel (F.map f)) ≅ cokernel f := by - -- We have to write an explicit `PreservesColimits` type here, - -- as `leftAdjointPreservesColimits` has universe variables. - have : PreservesColimits G := adj.leftAdjoint_preservesColimits - calc - G.obj (cokernel (F.map f)) ≅ cokernel (G.map (F.map f)) := - (asIso (cokernelComparison _ G)).symm - _ ≅ cokernel (i.hom.app X ≫ f ≫ i.inv.app Y) := cokernelIsoOfEq (NatIso.naturality_2 i f).symm - _ ≅ cokernel (f ≫ i.inv.app Y) := cokernelEpiComp (i.hom.app X) (f ≫ i.inv.app Y) - _ ≅ cokernel f := cokernelCompIsIso f (i.inv.app Y) - -variable [Limits.HasKernels C] [PreservesFiniteLimits G] - -/-- Auxiliary construction for `coimageIsoImage` -/ -def coimageIsoImageAux (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : - kernel (G.map (cokernel.π (F.map f))) ≅ kernel (cokernel.π f) := by - have : PreservesColimits G := adj.leftAdjoint_preservesColimits - calc - kernel (G.map (cokernel.π (F.map f))) ≅ - kernel (cokernel.π (G.map (F.map f)) ≫ cokernelComparison (F.map f) G) := - kernelIsoOfEq (π_comp_cokernelComparison _ _).symm - _ ≅ kernel (cokernel.π (G.map (F.map f))) := kernelCompMono _ _ - _ ≅ kernel (cokernel.π (_ ≫ f ≫ _) ≫ (cokernelIsoOfEq _).hom) := - (kernelIsoOfEq (π_comp_cokernelIsoOfEq_hom (NatIso.naturality_2 i f)).symm) - _ ≅ kernel (cokernel.π (_ ≫ f ≫ _)) := kernelCompMono _ _ - _ ≅ kernel (cokernel.π (f ≫ i.inv.app Y) ≫ (cokernelEpiComp (i.hom.app X) _).inv) := - (kernelIsoOfEq (by simp only [cokernel.π_desc, cokernelEpiComp_inv])) - _ ≅ kernel (cokernel.π (f ≫ _)) := kernelCompMono _ _ - _ ≅ kernel (inv (i.inv.app Y) ≫ cokernel.π f ≫ (cokernelCompIsIso f (i.inv.app Y)).inv) := - (kernelIsoOfEq - (by simp only [cokernel.π_desc, cokernelCompIsIso_inv, Iso.hom_inv_id_app_assoc, - NatIso.inv_inv_app])) - _ ≅ kernel (cokernel.π f ≫ _) := kernelIsIsoComp _ _ - _ ≅ kernel (cokernel.π f) := kernelCompMono _ _ - -variable [Functor.PreservesZeroMorphisms F] - -/-- Auxiliary definition: the abelian coimage and abelian image agree. -We still need to check that this agrees with the canonical morphism. --/ -def coimageIsoImage (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : - Abelian.coimage f ≅ Abelian.image f := by - have : PreservesLimits F := adj.rightAdjoint_preservesLimits - calc - Abelian.coimage f ≅ cokernel (kernel.ι f) := Iso.refl _ - _ ≅ G.obj (cokernel (F.map (kernel.ι f))) := (cokernelIso _ _ i adj _).symm - _ ≅ G.obj (cokernel (kernelComparison f F ≫ kernel.ι (F.map f))) := - (G.mapIso (cokernelIsoOfEq (by simp))) - _ ≅ G.obj (cokernel (kernel.ι (F.map f))) := G.mapIso (cokernelEpiComp _ _) - _ ≅ G.obj (Abelian.coimage (F.map f)) := Iso.refl _ - _ ≅ G.obj (Abelian.image (F.map f)) := G.mapIso (Abelian.coimageIsoImage _) - _ ≅ G.obj (kernel (cokernel.π (F.map f))) := Iso.refl _ - _ ≅ kernel (G.map (cokernel.π (F.map f))) := PreservesKernel.iso _ _ - _ ≅ kernel (cokernel.π f) := coimageIsoImageAux F G i adj f - _ ≅ Abelian.image f := Iso.refl _ - --- The account of this proof in the Stacks project omits this calculation. -theorem coimageIsoImage_hom (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : - (coimageIsoImage F G i adj f).hom = Abelian.coimageImageComparison f := by - dsimp [coimageIsoImage, cokernelIso, cokernelEpiComp, cokernelCompIsIso_inv, - coimageIsoImageAux, kernelCompMono] - simpa only [← cancel_mono (Abelian.image.ι f), ← cancel_epi (Abelian.coimage.π f), - Category.assoc, Category.id_comp, cokernel.π_desc_assoc, - π_comp_cokernelIsoOfEq_inv_assoc, PreservesKernel.iso_hom, - π_comp_cokernelComparison_assoc, ← G.map_comp_assoc, kernel.lift_ι, - Abelian.coimage_image_factorisation, lift_comp_kernelIsoOfEq_hom_assoc, - kernelIsIsoComp_hom, kernel.lift_ι_assoc, kernelIsoOfEq_hom_comp_ι_assoc, - kernelComparison_comp_ι_assoc, π_comp_cokernelIsoOfEq_hom_assoc, - asIso_hom, NatIso.inv_inv_app] using NatIso.naturality_1 i f - end AbelianOfAdjunction open AbelianOfAdjunction /-- If `C` is an additive category, `D` is an abelian category, -we have `F : C ⥤ D` `G : D ⥤ C` (both preserving zero morphisms), +we have `F : C ⥤ D` `G : D ⥤ C` (with `G` preserving zero morphisms), `G` is left exact (that is, preserves finite limits), and further we have `adj : G ⊣ F` and `i : F ⋙ G ≅ 𝟭 C`, then `C` is also abelian. -/ @[stacks 03A3] def abelianOfAdjunction {C : Type u₁} [Category.{v₁} C] [Preadditive C] [HasFiniteProducts C] - {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] + {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) (G : D ⥤ C) [Functor.PreservesZeroMorphisms G] [PreservesFiniteLimits G] (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) : Abelian C := by haveI := hasKernels F G i haveI := hasCokernels F G i adj have : ∀ {X Y : C} (f : X ⟶ Y), IsIso (Abelian.coimageImageComparison f) := by intro X Y f - rw [← coimageIsoImage_hom F G i adj f] + let arrowIso : Arrow.mk (G.map (F.map f)) ≅ Arrow.mk f := + ((Functor.mapArrowFunctor _ _).mapIso i).app (Arrow.mk f) + have : PreservesColimits G := adj.leftAdjoint_preservesColimits + let iso : Arrow.mk (G.map (Abelian.coimageImageComparison (F.map f))) ≅ + Arrow.mk (Abelian.coimageImageComparison f) := + Abelian.PreservesCoimageImageComparison.iso G (F.map f) ≪≫ + Abelian.coimageImageComparisonFunctor.mapIso arrowIso + rw [Arrow.isIso_iff_isIso_of_isIso iso.inv] infer_instance apply Abelian.ofCoimageImageComparisonIsIso @@ -174,7 +105,7 @@ via a functor that preserves zero morphisms, then `C` is also abelian. -/ def abelianOfEquivalence {C : Type u₁} [Category.{v₁} C] [Preadditive C] [HasFiniteProducts C] - {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] + {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) [F.IsEquivalence] : Abelian C := abelianOfAdjunction F F.inv F.asEquivalence.unitIso.symm F.asEquivalence.symm.toAdjunction diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index 1779eaa423ca12..dedb872367a90f 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -91,6 +91,35 @@ theorem mk_inj (A B : T) {f g : A ⟶ B} : Arrow.mk f = Arrow.mk g ↔ f = g := instance {X Y : T} : CoeOut (X ⟶ Y) (Arrow T) where coe := mk +lemma mk_eq_mk_iff {X Y X' Y' : T} (f : X ⟶ Y) (f' : X' ⟶ Y') : + Arrow.mk f = Arrow.mk f' ↔ + ∃ (hX : X = X') (hY : Y = Y'), f = eqToHom hX ≫ f' ≫ eqToHom hY.symm := by + constructor + · intro h + refine ⟨congr_arg Comma.left h, congr_arg Comma.right h, ?_⟩ + have := (eqToIso h).hom.w + dsimp at this + rw [Comma.eqToHom_left, Comma.eqToHom_right] at this + rw [reassoc_of% this, eqToHom_trans, eqToHom_refl, Category.comp_id] + · rintro ⟨rfl, rfl, h⟩ + simp only [eqToHom_refl, Category.comp_id, Category.id_comp] at h + rw [h] + +lemma ext {f g : Arrow T} + (h₁ : f.left = g.left) (h₂ : f.right = g.right) + (h₃ : f.hom = eqToHom h₁ ≫ g.hom ≫ eqToHom h₂.symm) : f = g := + (mk_eq_mk_iff _ _).2 (by aesop) + +@[simp] +lemma arrow_mk_comp_eqToHom {X Y Y' : T} (f : X ⟶ Y) (h : Y = Y') : + Arrow.mk (f ≫ eqToHom h) = Arrow.mk f := + ext rfl h.symm (by simp) + +@[simp] +lemma arrow_mk_eqToHom_comp {X' X Y : T} (f : X ⟶ Y) (h : X' = X) : + Arrow.mk (eqToHom h ≫ f) = Arrow.mk f := + ext h rfl (by simp) + /-- A morphism in the arrow category is a commutative square connecting two objects of the arrow category. -/ @[simps] @@ -294,16 +323,6 @@ def rightFunc : Arrow C ⥤ C := @[simps] def leftToRight : (leftFunc : Arrow C ⥤ C) ⟶ rightFunc where app f := f.hom -lemma ext {f g : Arrow C} - (h₁ : f.left = g.left) (h₂ : f.right = g.right) - (h₃ : f.hom = eqToHom h₁ ≫ g.hom ≫ eqToHom h₂.symm) : f = g := by - obtain ⟨X, Y, f⟩ := f - obtain ⟨X', Y', g⟩ := g - obtain rfl : X = X' := h₁ - obtain rfl : Y = Y' := h₂ - obtain rfl : f = g := by simpa using h₃ - rfl - end Arrow namespace Functor @@ -315,10 +334,7 @@ variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] /-- A functor `C ⥤ D` induces a functor between the corresponding arrow categories. -/ @[simps] def mapArrow (F : C ⥤ D) : Arrow C ⥤ Arrow D where - obj a := - { left := F.obj a.left - right := F.obj a.right - hom := F.map a.hom } + obj a := Arrow.mk (F.map a.hom) map f := { left := F.map f.left right := F.map f.right @@ -355,9 +371,11 @@ instance isEquivalence_mapArrow (F : C ⥤ D) [IsEquivalence F] : end Functor +variable {C D : Type*} [Category C] [Category D] + /-- The images of `f : Arrow C` by two isomorphic functors `F : C ⥤ D` are isomorphic arrows in `D`. -/ -def Arrow.isoOfNatIso {C D : Type*} [Category C] [Category D] {F G : C ⥤ D} (e : F ≅ G) +def Arrow.isoOfNatIso {F G : C ⥤ D} (e : F ≅ G) (f : Arrow C) : F.mapArrow.obj f ≅ G.mapArrow.obj f := Arrow.isoMk (e.app f.left) (e.app f.right) @@ -382,4 +400,14 @@ def Arrow.discreteEquiv (S : Type u) : Arrow (Discrete S) ≃ S where rfl right_inv _ := rfl +/-- Extensionality lemma for functors `C ⥤ D` which uses as an assumption +that the induced maps `Arrow C → Arrow D` coincide. -/ +lemma Arrow.functor_ext {F G : C ⥤ D} (h : ∀ ⦃X Y : C⦄ (f : X ⟶ Y), + F.mapArrow.obj (Arrow.mk f) = G.mapArrow.obj (Arrow.mk f)) : + F = G := + Functor.ext (fun X ↦ congr_arg Comma.left (h (𝟙 X))) (fun X Y f ↦ by + have := h f + simp only [Functor.mapArrow_obj, mk_eq_mk_iff] at this + tauto) + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Comma/CardinalArrow.lean b/Mathlib/CategoryTheory/Comma/CardinalArrow.lean index 14443de419c695..0f9f4d3bf09c4e 100644 --- a/Mathlib/CategoryTheory/Comma/CardinalArrow.lean +++ b/Mathlib/CategoryTheory/Comma/CardinalArrow.lean @@ -77,17 +77,16 @@ noncomputable def Arrow.shrinkHomsEquiv (C : Type u) [Category.{v} C] [LocallySm Arrow.{w} (ShrinkHoms C) ≃ Arrow C where toFun := (ShrinkHoms.equivalence C).inverse.mapArrow.obj invFun := (ShrinkHoms.equivalence C).functor.mapArrow.obj - left_inv _ := by simp [Functor.mapArrow]; rfl - right_inv _ := by simp [Functor.mapArrow]; rfl + left_inv _ := by simp + right_inv _ := by simp /-- The bijection `Arrow (Shrink C) ≃ Arrow C`. -/ noncomputable def Arrow.shrinkEquiv (C : Type u) [Category.{v} C] [Small.{w} C] : Arrow (Shrink.{w} C) ≃ Arrow C where toFun := (Shrink.equivalence C).inverse.mapArrow.obj invFun := (Shrink.equivalence C).functor.mapArrow.obj - left_inv f := by - refine Arrow.ext (Equiv.apply_symm_apply _ _) - ((Equiv.apply_symm_apply _ _)) (by simp; rfl) + left_inv _ := Arrow.ext (Equiv.apply_symm_apply _ _) + ((Equiv.apply_symm_apply _ _)) (by simp ; rfl) right_inv _ := Arrow.ext (by simp [Shrink.equivalence]) (by simp [Shrink.equivalence]) (by simp [Shrink.equivalence]) diff --git a/Mathlib/CategoryTheory/EqToHom.lean b/Mathlib/CategoryTheory/EqToHom.lean index e9ede8e6dc5e4c..d1f6ccdc36a022 100644 --- a/Mathlib/CategoryTheory/EqToHom.lean +++ b/Mathlib/CategoryTheory/EqToHom.lean @@ -252,6 +252,7 @@ theorem hext {F G : C ⥤ D} (h_obj : ∀ X, F.obj X = G.obj X) -- Using equalities between functors. theorem congr_obj {F G : C ⥤ D} (h : F = G) (X) : F.obj X = G.obj X := by rw [h] +@[reassoc] theorem congr_hom {F G : C ⥤ D} (h : F = G) {X Y} (f : X ⟶ Y) : F.map f = eqToHom (congr_obj h X) ≫ G.map f ≫ eqToHom (congr_obj h Y).symm := by subst h; simp diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index f7766dcde513a9..73bd7f0e18a9d8 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -294,9 +294,9 @@ noncomputable def finitaryExtensiveTopCatAux (Z : TopCat.{u}) · rintro x ⟨⟨⟩, hx⟩; refine ⟨⟨⟨x, PUnit.unit⟩, hx.symm⟩, rfl⟩ refine ((TopCat.binaryCofan_isColimit_iff _).mpr ⟨?_, ?_, ?_⟩).some · refine ⟨(Homeomorph.prodPUnit Z).isEmbedding.comp .subtypeVal, ?_⟩ - convert f.2.1 _ isOpen_range_inl + convert f.hom.2.1 _ isOpen_range_inl · refine ⟨(Homeomorph.prodPUnit Z).isEmbedding.comp .subtypeVal, ?_⟩ - convert f.2.1 _ isOpen_range_inr + convert f.hom.2.1 _ isOpen_range_inr · convert Set.isCompl_range_inl_range_inr.preimage f instance finitaryExtensive_TopCat : FinitaryExtensive TopCat.{u} := by @@ -317,11 +317,12 @@ instance finitaryExtensive_TopCat : FinitaryExtensive TopCat.{u} := by (ConcreteCategory.congr_hom hαY val :).symm delta ExistsUnique at this choose l hl hl' using this - refine ⟨⟨l, ?_⟩, ContinuousMap.ext fun a => (hl a).symm, TopCat.isTerminalPUnit.hom_ext _ _, - fun {l'} h₁ _ => ContinuousMap.ext fun x => + refine ⟨TopCat.ofHom ⟨l, ?_⟩, TopCat.ext fun a => (hl a).symm, + TopCat.isTerminalPUnit.hom_ext _ _, + fun {l'} h₁ _ => TopCat.ext fun x => hl' x (l' x) (ConcreteCategory.congr_hom h₁ x).symm⟩ apply (IsEmbedding.inl (X := X') (Y := Y')).isInducing.continuous_iff.mpr - convert s.fst.2 using 1 + convert s.fst.hom.2 using 1 exact (funext hl).symm · refine ⟨⟨hαY.symm⟩, ⟨PullbackCone.isLimitAux' _ ?_⟩⟩ intro s @@ -334,11 +335,12 @@ instance finitaryExtensive_TopCat : FinitaryExtensive TopCat.{u} := by · exact ⟨val, rfl, fun y h => Sum.inr_injective h.symm⟩ delta ExistsUnique at this choose l hl hl' using this - refine ⟨⟨l, ?_⟩, ContinuousMap.ext fun a => (hl a).symm, TopCat.isTerminalPUnit.hom_ext _ _, + refine ⟨TopCat.ofHom ⟨l, ?_⟩, TopCat.ext fun a => (hl a).symm, + TopCat.isTerminalPUnit.hom_ext _ _, fun {l'} h₁ _ => - ContinuousMap.ext fun x => hl' x (l' x) (ConcreteCategory.congr_hom h₁ x).symm⟩ + TopCat.ext fun x => hl' x (l' x) (ConcreteCategory.congr_hom h₁ x).symm⟩ apply (IsEmbedding.inr (X := X') (Y := Y')).isInducing.continuous_iff.mpr - convert s.fst.2 using 1 + convert s.fst.hom.2 using 1 exact (funext hl).symm · intro Z f exact finitaryExtensiveTopCatAux Z f diff --git a/Mathlib/CategoryTheory/Functor/Basic.lean b/Mathlib/CategoryTheory/Functor/Basic.lean index 13a3440fd5fb98..59fdcbc05e00ca 100644 --- a/Mathlib/CategoryTheory/Functor/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/Basic.lean @@ -5,6 +5,7 @@ Authors: Tim Baumann, Stephen Morgan, Kim Morrison -/ import Mathlib.CategoryTheory.Category.Basic import Mathlib.Combinatorics.Quiver.Prefunctor +import Mathlib.Tactic.CategoryTheory.CheckCompositions /-! # Functors diff --git a/Mathlib/CategoryTheory/Grothendieck.lean b/Mathlib/CategoryTheory/Grothendieck.lean index 799069e98d5d95..f3614f0ca47a91 100644 --- a/Mathlib/CategoryTheory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Grothendieck.lean @@ -28,6 +28,16 @@ Really we should treat `Cat` as a 2-category, and allow `F` to be a 2-functor. There is also a closely related construction starting with `G : Cᵒᵖ ⥤ Cat`, where morphisms consists again of `β : b ⟶ b'` and `φ : f ⟶ (F.map (op β)).obj f'`. +## Notable constructions + +- `Grothendieck F` is the Grothendieck construction. +- Elements of `Grothendieck F` whose base is `c : C` can be transported along `f : c ⟶ d` using +`transport`. +- A natural transformation `α : F ⟶ G` induces `map α : Grothendieck F ⥤ Grothendieck G`. +- The Grothendieck construction and `map` together form a functor (`functor`) from the functor +category `E ⥤ Cat` to the over category `Over E`. +- A functor `G : D ⥤ C` induces `pre F G : Grothendieck (G ⋙ F) ⥤ Grothendieck F`. + ## References See also `CategoryTheory.Functor.Elements` for the category of elements of functor `F : C ⥤ Type`. @@ -144,10 +154,64 @@ theorem congr {X Y : Grothendieck F} {f g : X ⟶ Y} (h : f = g) : dsimp simp +@[simp] +theorem base_eqToHom {X Y : Grothendieck F} (h : X = Y) : + (eqToHom h).base = eqToHom (congrArg Grothendieck.base h) := by subst h; rfl + +@[simp] +theorem fiber_eqToHom {X Y : Grothendieck F} (h : X = Y) : + (eqToHom h).fiber = eqToHom (by subst h; simp) := by subst h; rfl + lemma eqToHom_eq {X Y : Grothendieck F} (hF : X = Y) : eqToHom hF = { base := eqToHom (by subst hF; rfl), fiber := eqToHom (by subst hF; simp) } := by subst hF rfl + +section Transport + +/-- +If `F : C ⥤ Cat` is a functor and `t : c ⟶ d` is a morphism in `C`, then `transport` maps each +`c`-based element of `Grothendieck F` to a `d`-based element. +-/ +@[simps] +def transport (x : Grothendieck F) {c : C} (t : x.base ⟶ c) : Grothendieck F := + ⟨c, (F.map t).obj x.fiber⟩ + +/-- +If `F : C ⥤ Cat` is a functor and `t : c ⟶ d` is a morphism in `C`, then `transport` maps each +`c`-based element `x` of `Grothendieck F` to a `d`-based element `x.transport t`. + +`transport_hom` is the morphism `x ⟶ x.transport t` induced by `t` and the identity on fibers. +-/ +@[simps] +def toTransport (x : Grothendieck F) {c : C} (t : x.base ⟶ c) : x ⟶ x.transport t := + ⟨t, 𝟙 _⟩ + +/-- +Construct an isomorphism in a Grothendieck construction from isomorphisms in its base and fiber. +-/ +@[simps] +def isoMk {X Y : Grothendieck F} (e₁ : X.base ≅ Y.base) + (e₂ : (F.map e₁.hom).obj X.fiber ≅ Y.fiber) : + X ≅ Y where + hom := ⟨e₁.hom, e₂.hom⟩ + inv := ⟨e₁.inv, (F.map e₁.inv).map e₂.inv ≫ + eqToHom (Functor.congr_obj (F.mapIso e₁).hom_inv_id X.fiber)⟩ + hom_inv_id := Grothendieck.ext _ _ (by simp) (by simp) + inv_hom_id := Grothendieck.ext _ _ (by simp) (by + have := Functor.congr_hom (F.mapIso e₁).inv_hom_id e₂.inv + dsimp at this + simp [this]) + +/-- +If `F : C ⥤ Cat` and `x : Grothendieck F`, then every `C`-isomorphism `α : x.base ≅ c` induces +an isomorphism between `x` and its transport along `α` +-/ +@[simps!] +def transportIso (x : Grothendieck F) {c : C} (α : x.base ≅ c) : + x.transport α.hom ≅ x := (isoMk α (Iso.refl _)).symm + +end Transport section variable (F) @@ -355,7 +419,10 @@ def grothendieckTypeToCat : Grothendieck (G ⋙ typeToCat) ≌ G.Elements where simp rfl -variable (F) in +section Pre + +variable (F) + /-- Applying a functor `G : D ⥤ C` to the base of the Grothendieck construction induces a functor `Grothendieck (G ⋙ F) ⥤ Grothendieck F`. -/ @[simps] @@ -365,6 +432,93 @@ def pre (G : D ⥤ C) : Grothendieck (G ⋙ F) ⥤ Grothendieck F where map_id X := Grothendieck.ext _ _ (G.map_id _) (by simp) map_comp f g := Grothendieck.ext _ _ (G.map_comp _ _) (by simp) +@[simp] +theorem pre_id : pre F (𝟭 C) = 𝟭 _ := rfl + +/-- +An natural isomorphism between functors `G ≅ H` induces a natural isomorphism between the canonical +morphism `pre F G` and `pre F H`, up to composition with +`Grothendieck (G ⋙ F) ⥤ Grothendieck (H ⋙ F)`. +-/ +def preNatIso {G H : D ⥤ C} (α : G ≅ H) : + pre F G ≅ map (whiskerRight α.hom F) ⋙ (pre F H) := + NatIso.ofComponents + (fun X => (transportIso ⟨G.obj X.base, X.fiber⟩ (α.app X.base)).symm) + (fun f => by fapply Grothendieck.ext <;> simp) + +/-- +Given an equivalence of categories `G`, `preInv _ G` is the (weak) inverse of the `pre _ G.functor`. +-/ +def preInv (G : D ≌ C) : Grothendieck F ⥤ Grothendieck (G.functor ⋙ F) := + map (whiskerRight G.counitInv F) ⋙ Grothendieck.pre (G.functor ⋙ F) G.inverse + +variable {F} in +lemma pre_comp_map (G: D ⥤ C) {H : C ⥤ Cat} (α : F ⟶ H) : + pre F G ⋙ map α = map (whiskerLeft G α) ⋙ pre H G := rfl + +variable {F} in +lemma pre_comp_map_assoc (G: D ⥤ C) {H : C ⥤ Cat} (α : F ⟶ H) {E : Type*} [Category E] + (K : Grothendieck H ⥤ E) : pre F G ⋙ map α ⋙ K= map (whiskerLeft G α) ⋙ pre H G ⋙ K := rfl + +variable {E : Type*} [Category E] in +@[simp] +lemma pre_comp (G : D ⥤ C) (H : E ⥤ D) : pre F (H ⋙ G) = pre (G ⋙ F) H ⋙ pre F G := rfl + +/-- +Let `G` be an equivalence of categories. The functor induced via `pre` by `G.functor ⋙ G.inverse` +is naturally isomorphic to the functor induced via `map` by a whiskered version of `G`'s inverse +unit. +-/ +protected def preUnitIso (G : D ≌ C) : + map (whiskerRight G.unitInv _) ≅ pre (G.functor ⋙ F) (G.functor ⋙ G.inverse) := + preNatIso _ G.unitIso.symm |>.symm + +/-- +Given a functor `F : C ⥤ Cat` and an equivalence of categories `G : D ≌ C`, the functor +`pre F G.functor` is an equivalence between `Grothendieck (G.functor ⋙ F)` and `Grothendieck F`. +-/ +def preEquivalence (G : D ≌ C) : Grothendieck (G.functor ⋙ F) ≌ Grothendieck F where + functor := pre F G.functor + inverse := preInv F G + unitIso := by + refine (eqToIso ?_) + ≪≫ (Grothendieck.preUnitIso F G |> isoWhiskerLeft (map _)) + ≪≫ (pre_comp_map_assoc G.functor _ _ |> Eq.symm |> eqToIso) + calc + _ = map (𝟙 _) := map_id_eq.symm + _ = map _ := ?_ + _ = map _ ⋙ map _ := map_comp_eq _ _ + congr; ext X + simp only [Functor.comp_obj, Functor.comp_map, ← Functor.map_comp, Functor.id_obj, + Functor.map_id, NatTrans.comp_app, NatTrans.id_app, whiskerLeft_app, whiskerRight_app, + Equivalence.counitInv_functor_comp] + counitIso := preNatIso F G.counitIso.symm |>.symm + functor_unitIso_comp := by + intro X + simp only [preInv, Grothendieck.preUnitIso, eq_mpr_eq_cast, cast_eq, pre_id, id_eq, + Iso.trans_hom, eqToIso.hom, eqToHom_app, eqToHom_refl, isoWhiskerLeft_hom, NatTrans.comp_app] + fapply Grothendieck.ext <;> simp [preNatIso, transportIso] + +variable {F} in +/-- +Let `F, F' : C ⥤ Cat` be functor, `G : D ≌ C` an equivalence and `α : F ⟶ F'` a natural +transformation. + +Left-whiskering `α` by `G` and then taking the Grothendieck construction is, up to isomorphism, +the same as taking the Grothendieck construction of `α` and using the equivalences `pre F G` +and `pre F' G` to match the expected type: + +``` +Grothendieck (G.functor ⋙ F) ≌ Grothendieck F ⥤ Grothendieck F' ≌ Grothendieck (G.functor ⋙ F') +``` +-/ +def mapWhiskerLeftIsoConjPreMap {F' : C ⥤ Cat} (G : D ≌ C) (α : F ⟶ F') : + map (whiskerLeft G.functor α) ≅ + (preEquivalence F G).functor ⋙ map α ⋙ (preEquivalence F' G).inverse := + (Functor.rightUnitor _).symm ≪≫ isoWhiskerLeft _ (preEquivalence F' G).unitIso + +end Pre + section FunctorFrom variable {E : Type*} [Category E] diff --git a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean index 43c2773df6cc1f..a8bd0651b3c8b6 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean @@ -7,7 +7,9 @@ import Mathlib.CategoryTheory.Limits.Constructions.BinaryProducts import Mathlib.CategoryTheory.Limits.Constructions.Equalizers import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts import Mathlib.CategoryTheory.Limits.Preserves.Finite +import Mathlib.CategoryTheory.Limits.Preserves.Creates.Finite import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Equalizers +import Mathlib.CategoryTheory.Limits.Creates import Mathlib.Data.Fintype.Prod import Mathlib.Data.Fintype.Sigma @@ -217,6 +219,60 @@ lemma preservesLimits_of_preservesEqualizers_and_products [HasEqualizers C] [∀ J, PreservesLimitsOfShape (Discrete.{w} J) G] : PreservesLimitsOfSize.{w, w} G where preservesLimitsOfShape := preservesLimit_of_preservesEqualizers_and_product G +section + +variable [HasLimitsOfShape (Discrete J) D] [HasLimitsOfShape (Discrete (Σp : J × J, p.1 ⟶ p.2)) D] + [HasEqualizers D] + +variable (G : C ⥤ D) [G.ReflectsIsomorphisms] [CreatesLimitsOfShape WalkingParallelPair G] + [CreatesLimitsOfShape (Discrete.{w} J) G] + [CreatesLimitsOfShape (Discrete.{w} (Σp : J × J, p.1 ⟶ p.2)) G] + +attribute [local instance] preservesLimit_of_preservesEqualizers_and_product in +/-- If a functor creates equalizers and the appropriate products, it creates limits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating limits" in the literature, and whether or not the condition can be dropped seems to depend +on the specific definition that is used. -/ +noncomputable def createsLimitsOfShapeOfCreatesEqualizersAndProducts : + CreatesLimitsOfShape J G where + CreatesLimit {K} := + have : HasLimitsOfShape (Discrete J) C := + hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape G + have : HasLimitsOfShape (Discrete (Σp : J × J, p.1 ⟶ p.2)) C := + hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape G + have : HasEqualizers C := + hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape G + have : HasLimit K := hasLimit_of_equalizer_and_product K + createsLimitOfReflectsIsomorphismsOfPreserves + +end + +/-- If a functor creates equalizers and finite products, it creates finite limits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating limits" in the literature, and whether or not the condition can be dropped seems to depend +on the specific definition that is used. -/ +noncomputable def createsFiniteLimitsOfCreatesEqualizersAndFiniteProducts [HasEqualizers D] + [HasFiniteProducts D] (G : C ⥤ D) [G.ReflectsIsomorphisms] + [CreatesLimitsOfShape WalkingParallelPair G] + [CreatesFiniteProducts G] : CreatesFiniteLimits G where + createsFiniteLimits _ _ _ := createsLimitsOfShapeOfCreatesEqualizersAndProducts G + +/-- If a functor creates equalizers and products, it creates limits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating limits" in the literature, and whether or not the condition can be dropped seems to depend +on the specific definition that is used. -/ +noncomputable def createsLimitsOfSizeOfCreatesEqualizersAndProducts [HasEqualizers D] + [HasProducts.{w} D] (G : C ⥤ D) [G.ReflectsIsomorphisms] + [CreatesLimitsOfShape WalkingParallelPair G] [∀ J, CreatesLimitsOfShape (Discrete.{w} J) G] : + CreatesLimitsOfSize.{w, w} G where + CreatesLimitsOfShape := createsLimitsOfShapeOfCreatesEqualizersAndProducts G + theorem hasFiniteLimits_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] : HasFiniteLimits C := @hasFiniteLimits_of_hasEqualizers_and_finite_products C _ @@ -239,6 +295,24 @@ lemma preservesFiniteLimits_of_preservesTerminal_and_pullbacks [HasTerminal C] apply PreservesFiniteProducts.mk apply preservesFiniteProducts_of_preserves_binary_and_terminal G +attribute [local instance] preservesFiniteLimits_of_preservesTerminal_and_pullbacks in +/-- If a functor creates terminal objects and pullbacks, it creates finite limits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating limits" in the literature, and whether or not the condition can be dropped seems to depend +on the specific definition that is used. -/ +noncomputable def createsFiniteLimitsOfCreatesTerminalAndPullbacks [HasTerminal D] + [HasPullbacks D] (G : C ⥤ D) [G.ReflectsIsomorphisms] + [CreatesLimitsOfShape (Discrete.{0} PEmpty) G] [CreatesLimitsOfShape WalkingCospan G] : + CreatesFiniteLimits G where + createsFiniteLimits _ _ _ := + { CreatesLimit := + have : HasTerminal C := hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape G + have : HasPullbacks C := hasLimitsOfShape_of_hasLimitsOfShape_createsLimitsOfShape G + have : HasFiniteLimits C := hasFiniteLimits_of_hasTerminal_and_pullbacks + createsLimitOfReflectsIsomorphismsOfPreserves } + /-! We now dualize the above constructions, resorting to copy-paste. -/ @@ -433,6 +507,60 @@ lemma preservesColimits_of_preservesCoequalizers_and_coproducts [HasCoequalizers [∀ J, PreservesColimitsOfShape (Discrete.{w} J) G] : PreservesColimitsOfSize.{w, w} G where preservesColimitsOfShape := preservesColimit_of_preservesCoequalizers_and_coproduct G +section + +variable [HasColimitsOfShape (Discrete J) D] + [HasColimitsOfShape (Discrete (Σp : J × J, p.1 ⟶ p.2)) D] [HasCoequalizers D] + +variable (G : C ⥤ D) [G.ReflectsIsomorphisms] [CreatesColimitsOfShape WalkingParallelPair G] + [CreatesColimitsOfShape (Discrete.{w} J) G] + [CreatesColimitsOfShape (Discrete.{w} (Σp : J × J, p.1 ⟶ p.2)) G] + +attribute [local instance] preservesColimit_of_preservesCoequalizers_and_coproduct in +/-- If a functor creates coequalizers and the appropriate coproducts, it creates colimits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating colimits" in the literature, and whether or not the condition can be dropped seems to +depend on the specific definition that is used. -/ +noncomputable def createsColimitsOfShapeOfCreatesCoequalizersAndCoproducts : + CreatesColimitsOfShape J G where + CreatesColimit {K} := + have : HasColimitsOfShape (Discrete J) C := + hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape G + have : HasColimitsOfShape (Discrete (Σp : J × J, p.1 ⟶ p.2)) C := + hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape G + have : HasCoequalizers C := + hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape G + have : HasColimit K := hasColimit_of_coequalizer_and_coproduct K + createsColimitOfReflectsIsomorphismsOfPreserves + +end + +/-- If a functor creates coequalizers and finite coproducts, it creates finite colimits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating colimits" in the literature, and whether or not the condition can be dropped seems to +depend on the specific definition that is used. -/ +noncomputable def createsFiniteColimitsOfCreatesCoequalizersAndFiniteCoproducts [HasCoequalizers D] + [HasFiniteCoproducts D] (G : C ⥤ D) [G.ReflectsIsomorphisms] + [CreatesColimitsOfShape WalkingParallelPair G] + [CreatesFiniteCoproducts G] : CreatesFiniteColimits G where + createsFiniteColimits _ _ _ := createsColimitsOfShapeOfCreatesCoequalizersAndCoproducts G + +/-- If a functor creates coequalizers and coproducts, it creates colimits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating colimits" in the literature, and whether or not the condition can be dropped seems to +depend on the specific definition that is used. -/ +noncomputable def createsColimitsOfSizeOfCreatesCoequalizersAndCoproducts [HasCoequalizers D] + [HasCoproducts.{w} D] (G : C ⥤ D) [G.ReflectsIsomorphisms] + [CreatesColimitsOfShape WalkingParallelPair G] + [∀ J, CreatesColimitsOfShape (Discrete.{w} J) G] : CreatesColimitsOfSize.{w, w} G where + CreatesColimitsOfShape := createsColimitsOfShapeOfCreatesCoequalizersAndCoproducts G + theorem hasFiniteColimits_of_hasInitial_and_pushouts [HasInitial C] [HasPushouts C] : HasFiniteColimits C := @hasFiniteColimits_of_hasCoequalizers_and_finite_coproducts C _ @@ -455,4 +583,22 @@ lemma preservesFiniteColimits_of_preservesInitial_and_pushouts [HasInitial C] apply PreservesFiniteCoproducts.mk apply preservesFiniteCoproductsOfPreservesBinaryAndInitial G +attribute [local instance] preservesFiniteColimits_of_preservesInitial_and_pushouts in +/-- If a functor creates initial objects and pushouts, it creates finite colimits. + +We additionally require the rather strong condition that the functor reflects isomorphisms. It is +unclear whether the statement remains true without this condition. There are various definitions of +"creating colimits" in the literature, and whether or not the condition can be dropped seems to +depend on the specific definition that is used. -/ +noncomputable def createsFiniteColimitsOfCreatesInitialAndPushouts [HasInitial D] + [HasPushouts D] (G : C ⥤ D) [G.ReflectsIsomorphisms] + [CreatesColimitsOfShape (Discrete.{0} PEmpty) G] [CreatesColimitsOfShape WalkingSpan G] : + CreatesFiniteColimits G where + createsFiniteColimits _ _ _ := + { CreatesColimit := + have : HasInitial C := hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape G + have : HasPushouts C := hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape G + have : HasFiniteColimits C := hasFiniteColimits_of_hasInitial_and_pushouts + createsColimitOfReflectsIsomorphismsOfPreserves } + end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 2ed9c2de711ca5..d8bb0a5989fbe7 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -263,6 +263,13 @@ def createsLimitOfReflectsIso' {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphis validLift := h.validLift ≪≫ IsLimit.uniqueUpToIso hc t makesLimit := h.makesLimit } +/-- If `F` reflects isomorphisms, and we already know that the limit exists in the source and `F` +preserves it, then `F` creates that limit. -/ +def createsLimitOfReflectsIsomorphismsOfPreserves {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphisms] + [HasLimit K] [PreservesLimit K F] : CreatesLimit K F := + createsLimitOfReflectsIso' (isLimitOfPreserves F (limit.isLimit _)) + ⟨⟨_, Iso.refl _⟩, limit.isLimit _⟩ + -- Notice however that even if the isomorphism is `Iso.refl _`, -- this construction will insert additional identity morphisms in the cone maps, -- so the constructed limits may not be ideal, definitionally. @@ -273,11 +280,7 @@ of a limit cone for `K ⋙ F`. def createsLimitOfFullyFaithfulOfLift' {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] {l : Cone (K ⋙ F)} (hl : IsLimit l) (c : Cone K) (i : F.mapCone c ≅ l) : CreatesLimit K F := - createsLimitOfReflectsIso fun _ t => - { liftedCone := c - validLift := i ≪≫ IsLimit.uniqueUpToIso hl t - makesLimit := - IsLimit.ofFaithful F (IsLimit.ofIsoLimit hl i.symm) _ fun _ => F.map_preimage _ } + createsLimitOfReflectsIso' hl ⟨⟨c, i⟩, isLimitOfReflects F (IsLimit.ofIsoLimit hl i.symm)⟩ -- Notice however that even if the isomorphism is `Iso.refl _`, -- this construction will insert additional identity morphisms in the cone maps, @@ -392,6 +395,17 @@ def createsColimitOfReflectsIso' {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorph validLift := h.validLift ≪≫ IsColimit.uniqueUpToIso hc t makesColimit := h.makesColimit } +/-- If `F` reflects isomorphisms, and we already know that the colimit exists in the source and `F` +preserves it, then `F` creates that colimit. -/ +def createsColimitOfReflectsIsomorphismsOfPreserves {K : J ⥤ C} {F : C ⥤ D} + [F.ReflectsIsomorphisms] [HasColimit K] [PreservesColimit K F] : CreatesColimit K F := + createsColimitOfReflectsIso' (isColimitOfPreserves F (colimit.isColimit _)) + ⟨⟨_, Iso.refl _⟩, colimit.isColimit _⟩ + +@[deprecated (since := "2025-02-01")] +noncomputable alias createsColimitOfFullyFaithfulOfPreserves := + createsColimitOfReflectsIsomorphismsOfPreserves + -- Notice however that even if the isomorphism is `Iso.refl _`, -- this construction will insert additional identity morphisms in the cocone maps, -- so the constructed colimits may not be ideal, definitionally. @@ -402,11 +416,7 @@ lift of a colimit cocone for `K ⋙ F`. def createsColimitOfFullyFaithfulOfLift' {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] {l : Cocone (K ⋙ F)} (hl : IsColimit l) (c : Cocone K) (i : F.mapCocone c ≅ l) : CreatesColimit K F := - createsColimitOfReflectsIso fun _ t => - { liftedCocone := c - validLift := i ≪≫ IsColimit.uniqueUpToIso hl t - makesColimit := - IsColimit.ofFaithful F (IsColimit.ofIsoColimit hl i.symm) _ fun _ => F.map_preimage _ } + createsColimitOfReflectsIso' hl ⟨⟨c, i⟩, isColimitOfReflects F (IsColimit.ofIsoColimit hl i.symm)⟩ -- Notice however that even if the isomorphism is `Iso.refl _`, -- this construction will insert additional identity morphisms in the cocone maps, @@ -420,12 +430,6 @@ def createsColimitOfFullyFaithfulOfLift {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F. CreatesColimit K F := createsColimitOfFullyFaithfulOfLift' (colimit.isColimit _) c i -/-- A fully faithful functor that preserves a colimit that exists also creates the colimit. -/ -def createsColimitOfFullyFaithfulOfPreserves {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] - [HasColimit K] [PreservesColimit K F] : CreatesColimit K F := - createsColimitOfFullyFaithfulOfLift' (isColimitOfPreserves _ (colimit.isColimit K)) _ - (Iso.refl _) - -- Notice however that even if the isomorphism is `Iso.refl _`, -- this construction will insert additional identity morphisms in the cocone maps, -- so the constructed colimits may not be ideal, definitionally. diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean index 89f2b68d3e8565..2013b3ba3f84f3 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean @@ -78,7 +78,6 @@ lemma preservesFiniteLimits_of_preservesFiniteLimitsOfSize (F : C ⥤ D) ∀ (J : Type w) {𝒥 : SmallCategory J} (_ : @FinCategory J 𝒥), PreservesLimitsOfShape J F) : PreservesFiniteLimits F where preservesFiniteLimits J (_ : SmallCategory J) _ := by - letI : Category (ULiftHom (ULift J)) := ULiftHom.category haveI := h (ULiftHom (ULift J)) CategoryTheory.finCategoryUlift exact preservesLimitsOfShape_of_equiv (ULiftHomULiftCategory.equiv J).symm F diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/AbelianImages.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/AbelianImages.lean new file mode 100644 index 00000000000000..6c1057877a816d --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/AbelianImages.lean @@ -0,0 +1,95 @@ +/- +Copyright (c) 2025 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Abelian.Images +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Kernels + +/-! +# Preservation of coimage-image comparisons + +If a functor preserves kernels and cokernels, then it preserves abelian images, abelian coimages +and coimage-image comparisons. +-/ + +noncomputable section + +universe v₁ v₂ u₁ u₂ + +open CategoryTheory Limits + +namespace CategoryTheory.Abelian + +variable {C : Type u₁} [Category.{v₁} C] [HasZeroMorphisms C] [HasKernels C] [HasCokernels C] +variable {D : Type u₂} [Category.{v₂} D] [HasZeroMorphisms D] [HasKernels D] [HasCokernels D] +variable (F : C ⥤ D) [F.PreservesZeroMorphisms] [PreservesLimitsOfShape WalkingParallelPair F] + [PreservesColimitsOfShape WalkingParallelPair F] +variable {X Y : C} (f : X ⟶ Y) + +/-- If a functor preserves kernels and cokernels, it preserves abelian images. -/ +def PreservesImage.iso : F.obj (Abelian.image f) ≅ Abelian.image (F.map f) := + PreservesKernel.iso F _ ≪≫ kernel.mapIso _ _ (Iso.refl _) (PreservesCokernel.iso F _) (by simp) + +@[reassoc (attr := simp)] +theorem PreservesImage.iso_hom_ι : + (PreservesImage.iso F f).hom ≫ Abelian.image.ι (F.map f) = F.map (Abelian.image.ι f) := by + simp [iso] + +@[reassoc (attr := simp)] +theorem PreservesImage.factorThruImage_iso_hom : + F.map (Abelian.factorThruImage f) ≫ (PreservesImage.iso F f).hom = + Abelian.factorThruImage (F.map f) := by + ext; simp [iso] + +@[reassoc (attr := simp)] +theorem PreservesImage.iso_inv_ι : + (PreservesImage.iso F f).inv ≫ F.map (Abelian.image.ι f) = Abelian.image.ι (F.map f) := by + simp [iso] + +@[reassoc (attr := simp)] +theorem PreservesImage.factorThruImage_iso_inv : + Abelian.factorThruImage (F.map f) ≫ (PreservesImage.iso F f).inv = + F.map (Abelian.factorThruImage f) := by + simp [Iso.comp_inv_eq] + +/-- If a functor preserves kernels and cokernels, it preserves abelian coimages. -/ +def PreservesCoimage.iso : F.obj (Abelian.coimage f) ≅ Abelian.coimage (F.map f) := + PreservesCokernel.iso F _ ≪≫ cokernel.mapIso _ _ (PreservesKernel.iso F _) (Iso.refl _) (by simp) + +@[reassoc (attr := simp)] +theorem PreservesCoimage.iso_hom_π : + F.map (Abelian.coimage.π f) ≫ (PreservesCoimage.iso F f).hom = Abelian.coimage.π (F.map f) := by + simp [iso] + +@[reassoc (attr := simp)] +theorem PreservesCoimage.factorThruCoimage_iso_inv : + (PreservesCoimage.iso F f).inv ≫ F.map (Abelian.factorThruCoimage f) = + Abelian.factorThruCoimage (F.map f) := by + ext; simp [iso] + +@[reassoc (attr := simp)] +theorem PreservesCoimage.factorThruCoimage_iso_hom : + (PreservesCoimage.iso F f).hom ≫ Abelian.factorThruCoimage (F.map f) = + F.map (Abelian.factorThruCoimage f) := by + simp [← Iso.eq_inv_comp] + +@[reassoc (attr := simp)] +theorem PreservesCoimage.iso_inv_π : + Abelian.coimage.π (F.map f) ≫ (PreservesCoimage.iso F f).inv = F.map (Abelian.coimage.π f) := by + simp [Iso.comp_inv_eq] + +theorem PreservesCoimage.hom_coimageImageComparison : + (PreservesCoimage.iso F f).hom ≫ coimageImageComparison (F.map f) = + F.map (coimageImageComparison f) ≫ (PreservesImage.iso F f).hom := by + simp [← Functor.map_comp, ← Iso.eq_inv_comp, ← cancel_epi (Abelian.coimage.π (F.map f)), + ← cancel_mono (Abelian.image.ι (F.map f))] + +/-- If a functor preserves kernels and cokernels, it perserves coimage-image comparisons. -/ +@[simps!] +def PreservesCoimageImageComparison.iso : + Arrow.mk (F.map (coimageImageComparison f)) ≅ Arrow.mk (coimageImageComparison (F.map f)) := + Arrow.isoMk' _ _ (PreservesCoimage.iso F f) (PreservesImage.iso F f) + (PreservesCoimage.hom_coimageImageComparison F f) + +end CategoryTheory.Abelian diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean index 034c5e96b668a2..0bf8b6d24a4e04 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean @@ -128,6 +128,12 @@ an isomorphism. def PreservesKernel.iso : G.obj (kernel f) ≅ kernel (G.map f) := IsLimit.conePointUniqueUpToIso (isLimitOfHasKernelOfPreservesLimit G f) (limit.isLimit _) +@[reassoc (attr := simp)] +theorem PreservesKernel.iso_inv_ι : + (PreservesKernel.iso G f).inv ≫ G.map (kernel.ι f) = kernel.ι (G.map f) := + IsLimit.conePointUniqueUpToIso_inv_comp (isLimitOfHasKernelOfPreservesLimit G f) + (limit.isLimit _) (WalkingParallelPair.zero) + @[simp] theorem PreservesKernel.iso_hom : (PreservesKernel.iso G f).hom = kernelComparison f G := by rw [← cancel_mono (kernel.ι _)] @@ -251,6 +257,11 @@ def PreservesCokernel.iso : G.obj (cokernel f) ≅ cokernel (G.map f) := IsColimit.coconePointUniqueUpToIso (isColimitOfHasCokernelOfPreservesColimit G f) (colimit.isColimit _) +@[reassoc (attr := simp)] +theorem PreservesCokernel.π_iso_hom : G.map (cokernel.π f) ≫ (iso G f).hom = cokernel.π (G.map f) := + IsColimit.comp_coconePointUniqueUpToIso_hom (isColimitOfHasCokernelOfPreservesColimit G f) + (colimit.isColimit _) (WalkingParallelPair.one) + @[simp] theorem PreservesCokernel.iso_inv : (PreservesCokernel.iso G f).inv = cokernelComparison f G := by rw [← cancel_epi (cokernel.π _)] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index 2b845755423e4f..f8819e7377ca7b 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -72,6 +72,7 @@ noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} whe The functor `uliftFunctor : Type u ⥤ Type (max u v)` creates `u`-small colimits. -/ noncomputable instance : CreatesColimitsOfSize.{w, u} uliftFunctor.{v, u} where - CreatesColimitsOfShape := { CreatesColimit := fun {_} ↦ createsColimitOfFullyFaithfulOfPreserves } + CreatesColimitsOfShape := + { CreatesColimit := fun {_} ↦ createsColimitOfReflectsIsomorphismsOfPreserves } end CategoryTheory.Limits.Types diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean index b255543b53a02b..3992e66f2c72ac 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean @@ -3,9 +3,11 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic import Mathlib.CategoryTheory.Limits.Comma +import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.Basic import Mathlib.Order.SuccPred.Limit +import Mathlib.Order.Interval.Set.InitialSeg /-! # An assumption for constructions by transfinite induction @@ -20,12 +22,12 @@ universe w v v' u u' namespace CategoryTheory.Limits -variable (J : Type w) [Preorder J] (C : Type u) [Category.{v} C] +variable (J : Type w) [LinearOrder J] (C : Type u) [Category.{v} C] (K : Type u') [Category.{v'} K] -/-- A category `C` has iterations of shape a preordered type `J` +/-- A category `C` has iterations of shape a linearly ordered type `J` when certain specific shapes of colimits exists: colimits indexed by `J`, -and by `Set.Iio j` for `j : J`. -/ +and by `Set.Iio j` for `j : J`. -/ class HasIterationOfShape : Prop where hasColimitsOfShape_of_isSuccLimit (j : J) (hj : Order.IsSuccLimit j) : HasColimitsOfShape (Set.Iio j) C := by infer_instance @@ -41,6 +43,13 @@ lemma hasColimitsOfShape_of_isSuccLimit (j : J) HasColimitsOfShape (Set.Iio j) C := HasIterationOfShape.hasColimitsOfShape_of_isSuccLimit j hj +variable {J} in +lemma hasColimitsOfShape_of_isSuccLimit' + {α : Type*} [PartialOrder α] (h : α + subst hi₀ + exact (hi.not_lt (s.lt_top (Classical.arbitrary _))).elim + | hs i hi _ => + obtain ⟨a, rfl⟩ := (s.mem_range_iff_rel (b := i)).2 (by + simpa only [← hi₀] using Order.lt_succ_of_not_isMax hi) + have : OrderTop α := + { top := a + le_top b := by + rw [← s.le_iff_le] + exact Order.le_of_lt_succ (by simpa only [hi₀] using s.lt_top b) } + infer_instance + | hl i hi => + subst hi₀ + exact hasColimitsOfShape_of_isSuccLimit' C s hi + +lemma hasIterationOfShape_of_initialSeg {α : Type*} [LinearOrder α] + (h : α ≤i J) [Nonempty α] : + HasIterationOfShape α C where + hasColimitsOfShape := hasColimitsOfShape_of_initialSeg C h + hasColimitsOfShape_of_isSuccLimit j hj := by + have : Nonempty (Set.Iio j) := by + obtain ⟨a, ha⟩ := not_isMin_iff.1 hj.1 + exact ⟨⟨a, ha⟩⟩ + exact hasColimitsOfShape_of_initialSeg _ + (InitialSeg.trans (Set.principalSegIio j) h) + end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/PrincipalSeg.lean b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/PrincipalSeg.lean new file mode 100644 index 00000000000000..bc73cdfb62e8d4 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/PrincipalSeg.lean @@ -0,0 +1,33 @@ +/- +Copyright (c) 2025 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Order.InitialSeg +import Mathlib.CategoryTheory.Category.Preorder +import Mathlib.CategoryTheory.Limits.Cones + +/-! +# Cocones associated to principal segments + +If `f : α cases i · simp only [Walk.toSubgraph, Walk.getVert_zero, zero_add, getVert_cons_succ, Subgraph.sup_adj, - subgraphOfAdj_adj, true_or] + subgraphOfAdj_adj, true_or] · simp only [Walk.toSubgraph, getVert_cons_succ, Subgraph.sup_adj, subgraphOfAdj_adj, Sym2.eq, - Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] + Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] right exact ih (Nat.succ_lt_succ_iff.mp hi) diff --git a/Mathlib/Combinatorics/SimpleGraph/Operations.lean b/Mathlib/Combinatorics/SimpleGraph/Operations.lean index 4af95c2d77aa83..05d3c1c45655e9 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Operations.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Operations.lean @@ -148,6 +148,9 @@ def edge : SimpleGraph V := fromEdgeSet {s(s, t)} lemma edge_adj (v w : V) : (edge s t).Adj v w ↔ (v = s ∧ w = t ∨ v = t ∧ w = s) ∧ v ≠ w := by rw [edge, fromEdgeSet_adj, Set.mem_singleton_iff, Sym2.eq_iff] +lemma edge_comm : edge s t = edge t s := by + rw [edge, edge, Sym2.eq_swap] + variable [DecidableEq V] in instance : DecidableRel (edge s t).Adj := fun _ _ ↦ by rw [edge_adj]; infer_instance @@ -159,6 +162,9 @@ lemma edge_self_eq_bot : edge s s = ⊥ := by lemma sup_edge_self : G ⊔ edge s s = G := by rw [edge_self_eq_bot, sup_of_le_left bot_le] +lemma lt_sup_edge (hne : s ≠ t) (hn : ¬ G.Adj s t) : G < G ⊔ edge s t := + left_lt_sup.2 fun h ↦ hn <| h <| (edge_adj ..).mpr ⟨Or.inl ⟨rfl, rfl⟩, hne⟩ + variable {s t} lemma edge_edgeSet_of_ne (h : s ≠ t) : (edge s t).edgeSet = {s(s, t)} := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index d016e6d94ad996..a8747f5def9c04 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -312,8 +312,8 @@ lemma IsCycle.getVert_injOn {p : G.Walk u u} (hpc : p.IsCycle) : (p.not_nil_of_tail_not_nil (not_nil_of_isCycle_cons hpc)), Set.mem_setOf] at hn hm have := ((Walk.cons_isCycle_iff _ _).mp hpc).1.getVert_injOn (by omega : n - 1 ≤ p.tail.length) (by omega : m - 1 ≤ p.tail.length) - (by simp_all [SimpleGraph.Walk.getVert_tail, show n - 1 + 1 = n from by omega, - show m - 1 + 1 = m from by omega]) + (by simp_all [SimpleGraph.Walk.getVert_tail, show n - 1 + 1 = n by omega, + show m - 1 + 1 = m by omega]) omega lemma IsCycle.getVert_injOn' {p : G.Walk u u} (hpc : p.IsCycle) : @@ -324,8 +324,8 @@ lemma IsCycle.getVert_injOn' {p : G.Walk u u} (hpc : p.IsCycle) : have : p.length - n = p.length - m := Walk.length_reverse _ ▸ hpc.reverse.getVert_injOn (by simp only [Walk.length_reverse, Set.mem_setOf_eq]; omega) (by simp only [Walk.length_reverse, Set.mem_setOf_eq]; omega) - (by simp [Walk.getVert_reverse, show p.length - (p.length - n) = n from by omega, hnm, - show p.length - (p.length - m) = m from by omega]) + (by simp [Walk.getVert_reverse, show p.length - (p.length - n) = n by omega, hnm, + show p.length - (p.length - m) = m by omega]) omega lemma IsCycle.snd_ne_penultimate {p : G.Walk u u} (hp : p.IsCycle) : p.snd ≠ p.penultimate := by @@ -337,12 +337,10 @@ lemma IsCycle.snd_ne_penultimate {p : G.Walk u u} (hp : p.IsCycle) : p.snd ≠ p lemma IsCycle.getVert_endpoint_iff {i : ℕ} {p : G.Walk u u} (hpc : p.IsCycle) (hl : i ≤ p.length) : p.getVert i = u ↔ i = 0 ∨ i = p.length := by refine ⟨?_, by aesop⟩ - intro h - by_cases hi : i = 0 - · left; exact hi - · right - exact hpc.getVert_injOn (by simp only [Set.mem_setOf_eq]; omega) - (by simp only [Set.mem_setOf_eq]; omega) (h.symm ▸ (Walk.getVert_length p).symm) + rw [or_iff_not_imp_left] + intro h hi + exact hpc.getVert_injOn (by simp only [Set.mem_setOf_eq]; omega) + (by simp only [Set.mem_setOf_eq]; omega) (h.symm ▸ (Walk.getVert_length p).symm) lemma IsCycle.getVert_sub_one_neq_getVert_add_one {i : ℕ} {p : G.Walk u u} (hpc : p.IsCycle) (h : i ≤ p.length) : p.getVert (i - 1) ≠ p.getVert (i + 1) := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index 86e1e71cf4a52d..947d13e1ebf70e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -1149,25 +1149,23 @@ theorem mem_support_iff_exists_getVert {u v w : V} {p : G.Walk v w} : Nat.le_add_right, and_self] · rintro ⟨n, hn⟩ rw [SimpleGraph.Walk.mem_support_iff] - by_cases h0 : n = 0 - · rw [h0, getVert_zero] at hn + cases n with + | zero => + rw [getVert_zero] at hn left exact hn.1.symm - · right + | succ n => + right have hnp : ¬ p.Nil := by rw [nil_iff_length_eq] omega - rw [← support_tail_of_not_nil _ hnp] - rw [mem_support_iff_exists_getVert] - use n - 1 - simp only [Nat.sub_le_iff_le_add] - rw [getVert_tail, length_tail_add_one hnp] - have : (n - 1 + 1) = n := by omega - rwa [this] + rw [← support_tail_of_not_nil _ hnp, mem_support_iff_exists_getVert] + use n + rwa [getVert_tail, ← Nat.add_one_le_add_one_iff, length_tail_add_one hnp] termination_by p.length decreasing_by · simp_wf - rw [@Nat.lt_iff_add_one_le] + rw [Nat.lt_iff_add_one_le] rw [length_tail_add_one hnp] end Walk diff --git a/Mathlib/Condensed/Discrete/Characterization.lean b/Mathlib/Condensed/Discrete/Characterization.lean index 166add8efd026e..67da90011bbd03 100644 --- a/Mathlib/Condensed/Discrete/Characterization.lean +++ b/Mathlib/Condensed/Discrete/Characterization.lean @@ -33,8 +33,6 @@ universe u open CategoryTheory Limits Functor FintypeCat -attribute [local instance] HasForget.instFunLike - namespace Condensed variable {C : Type*} [Category C] [HasWeakSheafify (coherentTopology CompHaus.{u}) C] diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index fd62250dd9340c..9d8180a466bf6e 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -21,8 +21,6 @@ noncomputable section open CategoryTheory Functor Limits FintypeCat CompHausLike.LocallyConstant -attribute [local instance] HasForget.instFunLike - namespace Condensed section LocallyConstantAsColimit @@ -45,7 +43,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi obtain ⟨j, h⟩ := Profinite.exists_locallyConstant.{_, u} c hc f exact ⟨⟨j⟩, h⟩ · intro ⟨i⟩ ⟨j⟩ (fi : LocallyConstant _ _) (fj : LocallyConstant _ _) - (h : fi.comap (c.π.app i) = fj.comap (c.π.app j)) + (h : fi.comap (c.π.app i).hom = fj.comap (c.π.app j).hom) obtain ⟨k, ki, kj, _⟩ := IsCofilteredOrEmpty.cone_objs i j refine ⟨⟨k⟩, ki.op, kj.op, ?_⟩ dsimp @@ -61,7 +59,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi lemma isColimitLocallyConstantPresheaf_desc_apply (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (s : Cocone ((F ⋙ toProfinite).op ⋙ locallyConstantPresheaf X)) (i : I) (f : LocallyConstant (toProfinite.obj (F.obj i)) X) : - (isColimitLocallyConstantPresheaf c X hc).desc s (f.comap (c.π.app i)) = s.ι.app ⟨i⟩ f := by + (isColimitLocallyConstantPresheaf c X hc).desc s (f.comap (c.π.app i).hom) = s.ι.app ⟨i⟩ f := by change ((((locallyConstantPresheaf X).mapCocone c.op).ι.app ⟨i⟩) ≫ (isColimitLocallyConstantPresheaf c X hc).desc s) _ = _ rw [(isColimitLocallyConstantPresheaf c X hc).fac] @@ -75,7 +73,7 @@ noncomputable def isColimitLocallyConstantPresheafDiagram (S : Profinite) : lemma isColimitLocallyConstantPresheafDiagram_desc_apply (S : Profinite) (s : Cocone (S.diagram.op ⋙ locallyConstantPresheaf X)) (i : DiscreteQuotient S) (f : LocallyConstant (S.diagram.obj i) X) : - (isColimitLocallyConstantPresheafDiagram X S).desc s (f.comap (S.asLimitCone.π.app i)) = + (isColimitLocallyConstantPresheafDiagram X S).desc s (f.comap (S.asLimitCone.π.app i).hom) = s.ι.app ⟨i⟩ f := isColimitLocallyConstantPresheaf_desc_apply S.asLimitCone X S.asLimit s i f @@ -185,13 +183,13 @@ def locallyConstantIsoFinYoneda : /-- A finite set as a coproduct cocone in `Profinite` over itself. -/ def fintypeCatAsCofan (X : Profinite) : Cofan (fun (_ : X) ↦ (Profinite.of (PUnit.{u+1}))) := - Cofan.mk X (fun x ↦ (ContinuousMap.const _ x)) + Cofan.mk X (fun x ↦ TopCat.ofHom (ContinuousMap.const _ x)) /-- A finite set is the coproduct of its points in `Profinite`. -/ def fintypeCatAsCofanIsColimit (X : Profinite) [Finite X] : IsColimit (fintypeCatAsCofan X) := by - refine mkCofanColimit _ (fun t ↦ ⟨fun x ↦ t.inj x PUnit.unit, ?_⟩) ?_ - (fun _ _ h ↦ by ext x; exact ContinuousMap.congr_fun (h x) _) + refine mkCofanColimit _ (fun t ↦ TopCat.ofHom ⟨fun x ↦ t.inj x PUnit.unit, ?_⟩) ?_ + (fun _ _ h ↦ by ext x; exact CategoryTheory.congr_fun (h x) _) · apply continuous_of_discreteTopology (α := X) · aesop @@ -311,7 +309,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi (isLimitOfPreserves lightToProfinite hc) f exact ⟨⟨j⟩, h⟩ · intro ⟨i⟩ ⟨j⟩ (fi : LocallyConstant _ _) (fj : LocallyConstant _ _) - (h : fi.comap (c.π.app i) = fj.comap (c.π.app j)) + (h : fi.comap (c.π.app i).hom = fj.comap (c.π.app j).hom) obtain ⟨k, ki, kj, _⟩ := IsCofilteredOrEmpty.cone_objs i j refine ⟨⟨k⟩, ki.op, kj.op, ?_⟩ dsimp @@ -327,7 +325,7 @@ noncomputable def isColimitLocallyConstantPresheaf (hc : IsLimit c) [∀ i, Epi lemma isColimitLocallyConstantPresheaf_desc_apply (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (s : Cocone ((F ⋙ toLightProfinite).op ⋙ locallyConstantPresheaf X)) (n : ℕᵒᵖ) (f : LocallyConstant (toLightProfinite.obj (F.obj n)) X) : - (isColimitLocallyConstantPresheaf c X hc).desc s (f.comap (c.π.app n)) = s.ι.app ⟨n⟩ f := by + (isColimitLocallyConstantPresheaf c X hc).desc s (f.comap (c.π.app n).hom) = s.ι.app ⟨n⟩ f := by change ((((locallyConstantPresheaf X).mapCocone c.op).ι.app ⟨n⟩) ≫ (isColimitLocallyConstantPresheaf c X hc).desc s) _ = _ rw [(isColimitLocallyConstantPresheaf c X hc).fac] @@ -342,7 +340,7 @@ noncomputable def isColimitLocallyConstantPresheafDiagram (S : LightProfinite) : lemma isColimitLocallyConstantPresheafDiagram_desc_apply (S : LightProfinite) (s : Cocone (S.diagram.rightOp ⋙ locallyConstantPresheaf X)) (n : ℕ) (f : LocallyConstant (S.diagram.obj ⟨n⟩) X) : - (isColimitLocallyConstantPresheafDiagram X S).desc s (f.comap (S.asLimitCone.π.app ⟨n⟩)) = + (isColimitLocallyConstantPresheafDiagram X S).desc s (f.comap (S.asLimitCone.π.app ⟨n⟩).hom) = s.ι.app n f := by change ((((locallyConstantPresheaf X).mapCocone (coconeRightOpOfCone S.asLimitCone)).ι.app n) ≫ (isColimitLocallyConstantPresheafDiagram X S).desc s) _ = _ @@ -457,13 +455,13 @@ def locallyConstantIsoFinYoneda : toLightProfinite.op ⋙ /-- A finite set as a coproduct cocone in `LightProfinite` over itself. -/ def fintypeCatAsCofan (X : LightProfinite) : Cofan (fun (_ : X) ↦ (LightProfinite.of (PUnit.{u+1}))) := - Cofan.mk X (fun x ↦ (ContinuousMap.const _ x)) + Cofan.mk X (fun x ↦ TopCat.ofHom (ContinuousMap.const _ x)) /-- A finite set is the coproduct of its points in `LightProfinite`. -/ def fintypeCatAsCofanIsColimit (X : LightProfinite) [Finite X] : IsColimit (fintypeCatAsCofan X) := by - refine mkCofanColimit _ (fun t ↦ ⟨fun x ↦ t.inj x PUnit.unit, ?_⟩) ?_ - (fun _ _ h ↦ by ext x; exact ContinuousMap.congr_fun (h x) _) + refine mkCofanColimit _ (fun t ↦ TopCat.ofHom ⟨fun x ↦ t.inj x PUnit.unit, ?_⟩) ?_ + (fun _ _ h ↦ by ext x; exact CategoryTheory.congr_fun (h x) _) · apply continuous_of_discreteTopology (α := X) · aesop diff --git a/Mathlib/Condensed/Discrete/LocallyConstant.lean b/Mathlib/Condensed/Discrete/LocallyConstant.lean index bec3c4eba50686..e31971393c97f3 100644 --- a/Mathlib/Condensed/Discrete/LocallyConstant.lean +++ b/Mathlib/Condensed/Discrete/LocallyConstant.lean @@ -69,8 +69,6 @@ universe u w open CategoryTheory Limits LocallyConstant TopologicalSpace.Fiber Opposite Function Fiber -attribute [local instance] HasForget.instFunLike - variable {P : TopCat.{u} → Prop} namespace CompHausLike.LocallyConstant @@ -83,7 +81,7 @@ maps. def functorToPresheaves : Type (max u w) ⥤ ((CompHausLike.{u} P)ᵒᵖ ⥤ Type max u w) where obj X := { obj := fun ⟨S⟩ ↦ LocallyConstant S X - map := fun f g ↦ g.comap f.unop } + map := fun f g ↦ g.comap f.unop.hom } map f := { app := fun _ t ↦ t.map f } /-- @@ -112,12 +110,12 @@ def fiber : CompHausLike.{u} P := CompHausLike.of P a.val instance : HasProp P (fiber r a) := inferInstanceAs (HasProp P (Subtype _)) /-- The inclusion map from a component of the coproduct induced by `f` into `S`. -/ -def sigmaIncl : fiber r a ⟶ Q := TopologicalSpace.Fiber.sigmaIncl _ a +def sigmaIncl : fiber r a ⟶ Q := ofHom _ (TopologicalSpace.Fiber.sigmaIncl _ a) /-- The canonical map from the coproduct induced by `f` to `S` as an isomorphism in `CompHausLike P`. -/ noncomputable def sigmaIso [HasExplicitFiniteCoproducts.{u} P] : (finiteCoproduct (fiber r)) ≅ Q := - isoOfBijective (sigmaIsoHom r) ⟨sigmaIsoHom_inj r, sigmaIsoHom_surj r⟩ + isoOfBijective (ofHom _ (sigmaIsoHom r)) ⟨sigmaIsoHom_inj r, sigmaIsoHom_surj r⟩ lemma sigmaComparison_comp_sigmaIso [HasExplicitFiniteCoproducts.{u} P] (X : (CompHausLike.{u} P)ᵒᵖ ⥤ Type max u w) : @@ -180,19 +178,21 @@ variable {T : CompHausLike.{u} P} (g : T ⟶ S) This is an auxiliary definition, the details do not matter. What's important is that this map exists so that the lemma `incl_comap` works. -/ -def componentHom (a : Fiber (f.comap g)) : - fiber _ a ⟶ fiber _ (Fiber.mk f (g a.preimage)) where - toFun x := ⟨g x.val, by - simp only [Fiber.mk, Set.mem_preimage, Set.mem_singleton_iff] - convert map_eq_image _ _ x - exact map_preimage_eq_image_map _ _ a⟩ - continuous_toFun := by exact Continuous.subtype_mk (g.continuous.comp continuous_subtype_val) _ - -- term mode gives "unknown free variable" error. +def componentHom (a : Fiber (f.comap g.hom)) : + fiber _ a ⟶ fiber _ (Fiber.mk f (g a.preimage)) := + TopCat.ofHom + { toFun x := ⟨g x.val, by + simp only [Fiber.mk, Set.mem_preimage, Set.mem_singleton_iff] + convert map_eq_image _ _ x + exact map_preimage_eq_image_map _ _ a⟩ + continuous_toFun := by + exact Continuous.subtype_mk (g.hom.continuous.comp continuous_subtype_val) _ } + -- term mode gives "unknown free variable" error. lemma incl_comap {S T : (CompHausLike P)ᵒᵖ} (f : LocallyConstant S.unop (Y.obj (op (CompHausLike.of P PUnit.{u+1})))) - (g : S ⟶ T) (a : Fiber (f.comap g.unop)) : - g ≫ (sigmaIncl (f.comap g.unop) a).op = + (g : S ⟶ T) (a : Fiber (f.comap g.unop.hom)) : + g ≫ (sigmaIncl (f.comap g.unop.hom) a).op = (sigmaIncl f _).op ≫ (componentHom f g.unop a).op := rfl @@ -205,7 +205,7 @@ noncomputable def counitApp [HasExplicitFiniteCoproducts.{u} P] naturality := by intro S T g ext f - apply presheaf_ext (f.comap g.unop) + apply presheaf_ext (f.comap g.unop.hom) intro a simp only [op_unop, functorToPresheaves_obj_obj, types_comp_apply, functorToPresheaves_obj_map, incl_of_counitAppApp, ← FunctorToTypes.map_comp_apply, incl_comap] @@ -264,14 +264,15 @@ noncomputable def counit [HasExplicitFiniteCoproducts.{u} P] : haveI := CompHaus apply presheaf_ext (f.map (g.val.app (op (CompHausLike.of P PUnit.{u+1})))) intro a simp only [op_unop, functorToPresheaves_map_app, incl_of_counitAppApp] - apply presheaf_ext (f.comap (sigmaIncl _ _)) + apply presheaf_ext (f.comap (sigmaIncl _ _).hom) intro b simp only [counitAppAppImage, ← FunctorToTypes.map_comp_apply, ← op_comp, CompHausLike.coe_of, map_apply, IsTerminal.comp_from, ← map_preimage_eq_image_map] change (_ ≫ Y.val.map _) _ = (_ ≫ Y.val.map _) _ - simp only [← g.val.naturality, - show sigmaIncl (f.comap (sigmaIncl (f.map _) a)) b ≫ sigmaIncl (f.map _) a = - (sigmaInclIncl f _ a b) ≫ sigmaIncl f (Fiber.mk f _) from rfl] + simp only [← g.val.naturality] + rw [show sigmaIncl (f.comap (sigmaIncl (f.map _) a).hom) b ≫ sigmaIncl (f.map _) a = + CompHausLike.ofHom P (X := fiber _ b) (sigmaInclIncl f _ a b) ≫ sigmaIncl f (Fiber.mk f _) + by ext; rfl] simp only [op_comp, Functor.map_comp, types_comp_apply, incl_of_counitAppApp] simp only [counitAppAppImage, ← FunctorToTypes.map_comp_apply, ← op_comp, terminal.comp_from] rw [mk_image] diff --git a/Mathlib/Condensed/Discrete/Module.lean b/Mathlib/Condensed/Discrete/Module.lean index 79c0f44b9350da..eb153f31358e0a 100644 --- a/Mathlib/Condensed/Discrete/Module.lean +++ b/Mathlib/Condensed/Discrete/Module.lean @@ -24,8 +24,6 @@ universe w u open CategoryTheory LocallyConstant CompHausLike Functor Category Functor Opposite -attribute [local instance] HasForget.instFunLike - variable {P : TopCat.{u} → Prop} namespace CompHausLike.LocallyConstantModule @@ -40,7 +38,7 @@ constant maps. def functorToPresheaves : ModuleCat.{max u w} R ⥤ ((CompHausLike.{u} P)ᵒᵖ ⥤ ModuleCat R) where obj X := { obj := fun ⟨S⟩ ↦ ModuleCat.of R (LocallyConstant S X) - map := fun f ↦ ModuleCat.ofHom (comapₗ R f.unop) } + map := fun f ↦ ModuleCat.ofHom (comapₗ R f.unop.hom) } map f := { app := fun S ↦ ModuleCat.ofHom (mapₗ R f.hom) } variable [HasExplicitFiniteCoproducts.{0} P] [HasExplicitPullbacks.{u} P] diff --git a/Mathlib/Condensed/Light/Epi.lean b/Mathlib/Condensed/Light/Epi.lean index 5f757b75530cab..e5c1c087d4b4fa 100644 --- a/Mathlib/Condensed/Light/Epi.lean +++ b/Mathlib/Condensed/Light/Epi.lean @@ -88,7 +88,7 @@ lemma epi_π_app_zero_of_epi : Epi (c.π.app ⟨0⟩) := by change Epi (((forget R).mapCone c).π.app ⟨0⟩) apply coherentTopology.epi_π_app_zero_of_epi · simp only [LightProfinite.effectiveEpi_iff_surjective] - exact fun _ h ↦ Concrete.surjective_π_app_zero_of_surjective_map (limit.isLimit _) h + exact fun x h ↦ Concrete.surjective_π_app_zero_of_surjective_map (limit.isLimit x) h · have := (freeForgetAdjunction R).isRightAdjoint exact isLimitOfPreserves _ hc · exact fun _ ↦ (forget R).map_epi _ diff --git a/Mathlib/Condensed/Light/TopCatAdjunction.lean b/Mathlib/Condensed/Light/TopCatAdjunction.lean index 2201055ab0ffc4..9a7834ba22ed77 100644 --- a/Mathlib/Condensed/Light/TopCatAdjunction.lean +++ b/Mathlib/Condensed/Light/TopCatAdjunction.lean @@ -22,8 +22,6 @@ universe u open LightCondensed LightCondSet CategoryTheory LightProfinite -attribute [local instance] HasForget.instFunLike - namespace LightCondSet variable (X : LightCondSet.{u}) @@ -42,7 +40,7 @@ local instance underlyingTopologicalSpace : TopologicalSpace.coinduced (coinducingCoprod X) inferInstance /-- The object part of the functor `LightCondSet ⥤ TopCat` -/ -def toTopCat : TopCat.{u} := TopCat.of (X.val.obj ⟨LightProfinite.of PUnit⟩) +abbrev toTopCat : TopCat.{u} := TopCat.of (X.val.obj ⟨LightProfinite.of PUnit⟩) lemma continuous_coinducingCoprod {S : LightProfinite.{u}} (x : X.val.obj ⟨S⟩) : Continuous fun a ↦ (X.coinducingCoprod ⟨⟨S, x⟩, a⟩) := by @@ -54,17 +52,18 @@ lemma continuous_coinducingCoprod {S : LightProfinite.{u}} (x : X.val.obj ⟨S variable {X} {Y : LightCondSet} (f : X ⟶ Y) /-- The map part of the functor `LightCondSet ⥤ TopCat` -/ -@[simps] -def toTopCatMap : X.toTopCat ⟶ Y.toTopCat where - toFun := f.val.app ⟨LightProfinite.of PUnit⟩ - continuous_toFun := by - rw [continuous_coinduced_dom] - apply continuous_sigma - intro ⟨S, x⟩ - simp only [Function.comp_apply, coinducingCoprod] - rw [show (fun (a : S) ↦ f.val.app ⟨of PUnit⟩ (X.val.map ((of PUnit.{u+1}).const a).op x)) = _ - from funext fun a ↦ NatTrans.naturality_apply f.val ((of PUnit.{u+1}).const a).op x] - exact continuous_coinducingCoprod _ _ +@[simps!] +def toTopCatMap : X.toTopCat ⟶ Y.toTopCat := + TopCat.ofHom + { toFun := f.val.app ⟨LightProfinite.of PUnit⟩ + continuous_toFun := by + rw [continuous_coinduced_dom] + apply continuous_sigma + intro ⟨S, x⟩ + simp only [Function.comp_apply, coinducingCoprod] + rw [show (fun (a : S) ↦ f.val.app ⟨of PUnit⟩ (X.val.map ((of PUnit.{u+1}).const a).op x)) = _ + from funext fun a ↦ NatTrans.naturality_apply f.val ((of PUnit.{u+1}).const a).op x] + exact continuous_coinducingCoprod _ _ } /-- The functor `LightCondSet ⥤ TopCat` -/ @[simps] @@ -73,11 +72,12 @@ def _root_.lightCondSetToTopCat : LightCondSet.{u} ⥤ TopCat.{u} where map f := toTopCatMap f /-- The counit of the adjunction `lightCondSetToTopCat ⊣ topCatToLightCondSet` -/ -def topCatAdjunctionCounit (X : TopCat.{u}) : X.toLightCondSet.toTopCat ⟶ X where - toFun x := x.1 PUnit.unit - continuous_toFun := by - rw [continuous_coinduced_dom] - continuity +def topCatAdjunctionCounit (X : TopCat.{u}) : X.toLightCondSet.toTopCat ⟶ X := + TopCat.ofHom + { toFun x := x.1 PUnit.unit + continuous_toFun := by + rw [continuous_coinduced_dom] + continuity } /-- The counit of the adjunction `lightCondSetToTopCat ⊣ topCatToLightCondSet` is always bijective, but not an isomorphism in general (the inverse isn't continuous unless `X` is sequential). @@ -105,9 +105,8 @@ def topCatAdjunctionUnit (X : LightCondSet.{u}) : X ⟶ X.toTopCat.toLightCondSe apply continuous_coinduced_rng } naturality := fun _ _ _ ↦ by ext - simp only [TopCat.toSheafCompHausLike_val_obj, CompHausLike.compHausLikeToTop_obj, - Opposite.op_unop, types_comp_apply, TopCat.toSheafCompHausLike_val_map, - ← FunctorToTypes.map_comp_apply] + simp only [TopCat.toSheafCompHausLike_val_obj, Opposite.op_unop, types_comp_apply, + TopCat.toSheafCompHausLike_val_map, ← FunctorToTypes.map_comp_apply] rfl } /-- The adjunction `lightCondSetToTopCat ⊣ topCatToLightCondSet` -/ @@ -165,7 +164,7 @@ can only prove this for `X : TopCat.{0}`. def sequentialAdjunctionHomeo (X : TopCat.{0}) [SequentialSpace X] : X.toLightCondSet.toTopCat ≃ₜ X where toEquiv := topCatAdjunctionCounitEquiv X - continuous_toFun := (topCatAdjunctionCounit X).continuous + continuous_toFun := (topCatAdjunctionCounit X).hom.continuous continuous_invFun := by apply SeqContinuous.continuous unfold SeqContinuous diff --git a/Mathlib/Condensed/TopCatAdjunction.lean b/Mathlib/Condensed/TopCatAdjunction.lean index af00f6cd791793..8dfbaaa415b0b5 100644 --- a/Mathlib/Condensed/TopCatAdjunction.lean +++ b/Mathlib/Condensed/TopCatAdjunction.lean @@ -21,8 +21,6 @@ universe u open Condensed CondensedSet CategoryTheory CompHaus -attribute [local instance] HasForget.instFunLike - variable (X : CondensedSet.{u}) /-- Auxiliary definition to define the topology on `X(*)` for a condensed set `X`. -/ @@ -37,7 +35,7 @@ local instance : TopologicalSpace (X.val.obj ⟨CompHaus.of PUnit⟩) := TopologicalSpace.coinduced (coinducingCoprod X) inferInstance /-- The object part of the functor `CondensedSet ⥤ TopCat` -/ -def CondensedSet.toTopCat : TopCat.{u+1} := TopCat.of (X.val.obj ⟨of PUnit⟩) +abbrev CondensedSet.toTopCat : TopCat.{u+1} := TopCat.of (X.val.obj ⟨of PUnit⟩) namespace CondensedSet @@ -51,17 +49,18 @@ lemma continuous_coinducingCoprod {S : CompHaus.{u}} (x : X.val.obj ⟨S⟩) : variable {X} {Y : CondensedSet} (f : X ⟶ Y) /-- The map part of the functor `CondensedSet ⥤ TopCat` -/ -@[simps] -def toTopCatMap : X.toTopCat ⟶ Y.toTopCat where - toFun := f.val.app ⟨of PUnit⟩ - continuous_toFun := by - rw [continuous_coinduced_dom] - apply continuous_sigma - intro ⟨S, x⟩ - simp only [Function.comp_apply, coinducingCoprod] - rw [show (fun (a : S) ↦ f.val.app ⟨of PUnit⟩ (X.val.map ((of PUnit.{u+1}).const a).op x)) = _ - from funext fun a ↦ NatTrans.naturality_apply f.val ((of PUnit.{u+1}).const a).op x] - exact continuous_coinducingCoprod Y _ +@[simps!] +def toTopCatMap : X.toTopCat ⟶ Y.toTopCat := + TopCat.ofHom + { toFun := f.val.app ⟨of PUnit⟩ + continuous_toFun := by + rw [continuous_coinduced_dom] + apply continuous_sigma + intro ⟨S, x⟩ + simp only [Function.comp_apply, coinducingCoprod] + rw [show (fun (a : S) ↦ f.val.app ⟨of PUnit⟩ (X.val.map ((of PUnit.{u+1}).const a).op x)) = _ + from funext fun a ↦ NatTrans.naturality_apply f.val ((of PUnit.{u+1}).const a).op x] + exact continuous_coinducingCoprod Y _ } end CondensedSet @@ -74,12 +73,20 @@ def condensedSetToTopCat : CondensedSet.{u} ⥤ TopCat.{u+1} where namespace CondensedSet /-- The counit of the adjunction `condensedSetToTopCat ⊣ topCatToCondensedSet` -/ -@[simps] -def topCatAdjunctionCounit (X : TopCat.{u+1}) : X.toCondensedSet.toTopCat ⟶ X where - toFun x := x.1 PUnit.unit - continuous_toFun := by - rw [continuous_coinduced_dom] - continuity +def topCatAdjunctionCounit (X : TopCat.{u+1}) : X.toCondensedSet.toTopCat ⟶ X := + TopCat.ofHom + { toFun x := x.1 PUnit.unit + continuous_toFun := by + rw [continuous_coinduced_dom] + continuity } + +/-- `simp`-normal form of the lemma that `@[simps]` would generate. -/ +@[simp] lemma topCatAdjunctionCounit_hom_apply (X : TopCat) (x) : + -- We have to specify here to not infer the `TopologicalSpace` instance on `C(PUnit, X)`, + -- which suggests type synonyms are being unfolded too far somewhere. + DFunLike.coe (F := @ContinuousMap C(PUnit, X) X (_) _) + (TopCat.Hom.hom (topCatAdjunctionCounit X)) x = + x PUnit.unit := rfl /-- The counit of the adjunction `condensedSetToTopCat ⊣ topCatToCondensedSet` is always bijective, but not an isomorphism in general (the inverse isn't continuous unless `X` is compactly generated). @@ -107,15 +114,15 @@ def topCatAdjunctionUnit (X : CondensedSet.{u}) : X ⟶ X.toTopCat.toCondensedSe apply continuous_coinduced_rng } naturality := fun _ _ _ ↦ by ext - simp only [TopCat.toSheafCompHausLike_val_obj, CompHausLike.compHausLikeToTop_obj, + simp only [TopCat.toSheafCompHausLike_val_obj, Opposite.op_unop, types_comp_apply, TopCat.toSheafCompHausLike_val_map, ← FunctorToTypes.map_comp_apply] rfl } /-- The adjunction `condensedSetToTopCat ⊣ topCatToCondensedSet` -/ noncomputable def topCatAdjunction : condensedSetToTopCat.{u} ⊣ topCatToCondensedSet where - unit := { app := topCatAdjunctionUnit } - counit := { app := topCatAdjunctionCounit } + unit.app := topCatAdjunctionUnit + counit.app := topCatAdjunctionCounit left_triangle_components Y := by ext change Y.val.map (𝟙 _) _ = _ @@ -168,7 +175,7 @@ is a homeomorphism. def compactlyGeneratedAdjunctionCounitHomeo (X : TopCat.{u+1}) [UCompactlyGeneratedSpace.{u} X] : X.toCondensedSet.toTopCat ≃ₜ X where toEquiv := topCatAdjunctionCounitEquiv X - continuous_toFun := (topCatAdjunctionCounit X).continuous + continuous_toFun := (topCatAdjunctionCounit X).hom.continuous continuous_invFun := by apply continuous_from_uCompactlyGeneratedSpace exact fun _ _ ↦ continuous_coinducingCoprod X.toCondensedSet _ diff --git a/Mathlib/Condensed/TopComparison.lean b/Mathlib/Condensed/TopComparison.lean index b849bc48d6588d..d8ea3491adde1c 100644 --- a/Mathlib/Condensed/TopComparison.lean +++ b/Mathlib/Condensed/TopComparison.lean @@ -25,8 +25,6 @@ universe w w' v u open CategoryTheory Opposite Limits regularTopology ContinuousMap Topology -attribute [local instance] HasForget.instFunLike - variable {C : Type u} [Category.{v} C] (G : C ⥤ TopCat.{w}) (X : Type w') [TopologicalSpace X] @@ -113,7 +111,7 @@ def TopCat.toSheafCompHausLike : apply (config := { allowSynthFailures := true }) equalizerCondition_yonedaPresheaf (CompHausLike.compHausLikeToTop.{u} P) X intro Z B π he - apply IsQuotientMap.of_surjective_continuous (hs _ he) π.continuous + apply IsQuotientMap.of_surjective_continuous (hs _ he) π.hom.continuous /-- `TopCat.toSheafCompHausLike` yields a functor from `TopCat.{max u w}` to @@ -124,7 +122,7 @@ noncomputable def topCatToSheafCompHausLike : have := CompHausLike.preregular hs TopCat.{max u w} ⥤ Sheaf (coherentTopology (CompHausLike.{u} P)) (Type (max u w)) where obj X := X.toSheafCompHausLike P hs - map f := ⟨⟨fun _ g ↦ f.comp g, by aesop⟩⟩ + map f := ⟨⟨fun _ g ↦ f.hom.comp g, by aesop⟩⟩ end diff --git a/Mathlib/Data/DFinsupp/Small.lean b/Mathlib/Data/DFinsupp/Small.lean index 04278d86c6531d..a0779ecf449e75 100644 --- a/Mathlib/Data/DFinsupp/Small.lean +++ b/Mathlib/Data/DFinsupp/Small.lean @@ -21,8 +21,7 @@ variable {ι : Type u} {π : ι → Type v} [∀ i, Zero (π i)] section Small instance DFinsupp.small [Small.{w} ι] [∀ (i : ι), Small.{w} (π i)] : - Small.{w} (DFinsupp π) := small_of_injective (f := fun x j ↦ x j) (fun f f' eq ↦ by - ext j - exact congr_fun eq j) + Small.{w} (DFinsupp π) := + small_of_injective (f := fun x j ↦ x j) (fun f f' eq ↦ by ext j; exact congr_fun eq j) end Small diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 835dd0a62d9f31..621606be58e0e7 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -14,9 +14,10 @@ import Mathlib.Data.NNReal.Basic In this file we prove elementary properties of algebraic operations on `ℝ≥0∞`, including addition, multiplication, natural powers and truncated subtraction, as well as how these interact with the order structure on `ℝ≥0∞`. Notably excluded from this list are inversion and division, the -definitions and properties of which can be found in `Data.ENNReal.Inv`. +definitions and properties of which can be found in `Mathlib.Data.ENNReal.Inv`. -Note: the definitions of the operations included in this file can be found in `Data.ENNReal.Basic`. +Note: the definitions of the operations included in this file can be found in +`Mathlib.Data.ENNReal.Basic`. -/ open Set NNReal ENNReal diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 24f664b0250d95..6de87c5125c52e 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -311,6 +311,7 @@ lemma sub_ne_top_iff : a - b ≠ ⊤ ↔ a ≠ ⊤ ∨ b = ⊤ := WithTop.sub_ne lemma addLECancellable_of_ne_top : a ≠ ⊤ → AddLECancellable a := WithTop.addLECancellable_of_ne_top lemma addLECancellable_of_lt_top : a < ⊤ → AddLECancellable a := WithTop.addLECancellable_of_lt_top +lemma addLECancellable_coe (a : ℕ) : AddLECancellable (a : ℕ∞) := WithTop.addLECancellable_coe _ protected lemma le_sub_of_add_le_left (ha : a ≠ ⊤) : a + b ≤ c → b ≤ c - a := (addLECancellable_of_ne_top ha).le_tsub_of_add_le_left diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index b06760d8dfe35b..41d00bd2897614 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1265,6 +1265,17 @@ lemma succAbove_predAbove {p : Fin n} {i : Fin (n + 1)} (h : i ≠ castSucc p) : · rw [predAbove_of_le_castSucc _ _ (Fin.le_of_lt h), succAbove_castPred_of_lt _ _ h] · rw [predAbove_of_castSucc_lt _ _ h, succAbove_pred_of_lt _ _ h] +/-- Sending `Fin (n+1)` to `Fin n` by subtracting one from anything above `p` +then back to `Fin (n+1)` with a gap around `p.succ` is the identity away from `p.succ`. -/ +@[simp] +lemma succ_succAbove_predAbove {n : ℕ} {p : Fin n} {i : Fin (n + 1)} (h : i ≠ p.succ) : + p.succ.succAbove (p.predAbove i) = i := by + obtain h | h := Fin.lt_or_lt_of_ne h + · rw [predAbove_of_le_castSucc _ _ (le_castSucc_iff.2 h), + succAbove_castPred_of_lt _ _ h] + · rw [predAbove_of_castSucc_lt _ _ (Fin.lt_of_le_of_lt (p.castSucc_le_succ) h), + succAbove_pred_of_lt _ _ h] + /-- Sending `Fin n` into `Fin (n + 1)` with a gap at `p` then back to `Fin n` by subtracting one from anything above `p` is the identity. -/ @[simp] diff --git a/Mathlib/Data/List/ReduceOption.lean b/Mathlib/Data/List/ReduceOption.lean index 4e1c64731f5ace..63f75871fb36c5 100644 --- a/Mathlib/Data/List/ReduceOption.lean +++ b/Mathlib/Data/List/ReduceOption.lean @@ -43,8 +43,7 @@ theorem reduceOption_append (l l' : List (Option α)) : @[simp] theorem reduceOption_replicate_none {n : ℕ} : (replicate n (@none α)).reduceOption = [] := by dsimp [reduceOption] - rw [filterMap_replicate_of_none] - rfl + rw [filterMap_replicate_of_none (id_def _)] theorem reduceOption_eq_nil_iff (l : List (Option α)) : l.reduceOption = [] ↔ ∃ n, l = replicate n none := by @@ -77,12 +76,7 @@ theorem reduceOption_eq_append_iff (l : List (Option α)) (l'₁ l'₂ : List α l.reduceOption = l'₁ ++ l'₂ ↔ ∃ l₁ l₂, l = l₁ ++ l₂ ∧ l₁.reduceOption = l'₁ ∧ l₂.reduceOption = l'₂ := by dsimp [reduceOption] - constructor - · intro h - rw [filterMap_eq_append_iff] at h - trivial - · intro ⟨_, _, h, hl₁, hl₂⟩ - rw [h, filterMap_append, hl₁, hl₂] + exact filterMap_eq_append_iff theorem reduceOption_eq_concat_iff (l : List (Option α)) (l' : List α) (a : α) : l.reduceOption = l'.concat a ↔ @@ -96,8 +90,7 @@ theorem reduceOption_eq_concat_iff (l : List (Option α)) (l' : List α) (a : α obtain ⟨m, n, hl₂⟩ := hl₂ use l₁ ++ replicate m none, replicate n none simp_rw [h, reduceOption_append, reduceOption_replicate_none, append_assoc, append_nil, hl₁, - hl₂] - trivial + hl₂, and_self] · intro ⟨_, _, h, hl₁, hl₂⟩ rw [h, reduceOption_append, reduceOption_cons_of_some, hl₁, hl₂] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index d3c3bf82924e02..b9c45b65d4de3a 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1963,9 +1963,9 @@ lemma count_sub (a : α) (s t : Multiset α) : count a (s - t) = count a s - cou /-- This is a special case of `tsub_le_iff_right`, which should be used instead of this. This is needed to prove `OrderedSub (Multiset α)`. -/ protected lemma sub_le_iff_le_add : s - t ≤ u ↔ s ≤ u + t := by - revert s - exact @(Multiset.induction_on t (by simp [Multiset.sub_zero]) fun a t IH s => by - simp [IH, erase_le_iff_le_cons]) + induction t using Multiset.induction_on generalizing s with + | empty => simp [Multiset.sub_zero] + | cons a s IH => simp [IH, erase_le_iff_le_cons] /-- This is a special case of `tsub_le_iff_left`, which should be used instead of this. -/ protected lemma sub_le_iff_le_add' : s - t ≤ u ↔ s ≤ t + u := by diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index 8b1f848ab5ca8c..937cba89b97638 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -77,8 +77,7 @@ theorem choose_modEq_choose_mul_prod_range_choose (a : ℕ) : rw [prod_range_succ, cast_mul, ← mul_assoc, mul_right_comm] gcongr apply choose_modEq_choose_mod_mul_choose_div.trans - simp_rw [pow_succ, Nat.div_div_eq_div_mul, mul_comm] - rfl + simp_rw [pow_succ, Nat.div_div_eq_div_mul, mul_comm, Int.ModEq.refl] /-- **Lucas's Theorem**: For primes `p`, `choose n k` is congruent to the product of `choose (⌊n / p ^ i⌋ % p) (⌊k / p ^ i⌋ % p)` over `i` modulo `p`. -/ diff --git a/Mathlib/Data/Set/Accumulate.lean b/Mathlib/Data/Set/Accumulate.lean index 72726156c50141..962b0e6cd928b8 100644 --- a/Mathlib/Data/Set/Accumulate.lean +++ b/Mathlib/Data/Set/Accumulate.lean @@ -52,6 +52,10 @@ theorem iUnion_accumulate [Preorder α] : ⋃ x, Accumulate s x = ⋃ x, s x := exact ⟨x', hz⟩ · exact iUnion_mono fun i => subset_accumulate +@[simp] +lemma accumulate_bot [PartialOrder α] [OrderBot α] (s : α → Set β) : Accumulate s ⊥ = s ⊥ := by + simp [Set.accumulate_def] + @[simp] lemma accumulate_zero_nat (s : ℕ → Set β) : Accumulate s 0 = s 0 := by simp [accumulate_def] diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 3a93bcdb1c839b..300c6423b58e81 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -689,7 +689,7 @@ theorem exists_subset_range_and_iff {f : α → β} {p : Set β → Prop} : @[simp] theorem forall_subset_range_iff {f : α → β} {p : Set β → Prop} : (∀ s, s ⊆ range f → p s) ↔ ∀ s, p (f '' s) := by - rw [← forall_mem_range, range_image]; rfl + rw [← forall_mem_range, range_image]; simp only [mem_powerset_iff] @[simp] theorem preimage_subset_preimage_iff {s t : Set α} {f : β → α} (hs : s ⊆ range f) : diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index dac555b84fad52..0759b4e495bcf1 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -136,13 +136,13 @@ equivalence relations. -/ @[simps] def prodQuotientEquiv (r : Setoid α) (s : Setoid β) : Quotient r × Quotient s ≃ Quotient (r.prod s) where - toFun := fun (x, y) ↦ Quotient.map₂ Prod.mk (fun _ _ hx _ _ hy ↦ ⟨hx, hy⟩) x y - invFun := fun q ↦ Quotient.liftOn' q (fun xy ↦ (Quotient.mk'' xy.1, Quotient.mk'' xy.2)) + toFun | (x, y) => Quotient.map₂ Prod.mk (fun _ _ hx _ _ hy ↦ ⟨hx, hy⟩) x y + invFun q := Quotient.liftOn' q (fun xy ↦ (Quotient.mk'' xy.1, Quotient.mk'' xy.2)) fun x y hxy ↦ Prod.ext (by simpa using hxy.1) (by simpa using hxy.2) - left_inv := fun q ↦ by + left_inv q := by rcases q with ⟨qa, qb⟩ exact Quotient.inductionOn₂' qa qb fun _ _ ↦ rfl - right_inv := fun q ↦ by + right_inv q := by simp only refine Quotient.inductionOn' q fun _ ↦ rfl @@ -151,14 +151,14 @@ equivalence relations. -/ @[simps] noncomputable def piQuotientEquiv {ι : Sort*} {α : ι → Sort*} (r : ∀ i, Setoid (α i)) : (∀ i, Quotient (r i)) ≃ Quotient (@piSetoid _ _ r) where - toFun := fun x ↦ Quotient.mk'' fun i ↦ (x i).out - invFun := fun q ↦ Quotient.liftOn' q (fun x i ↦ Quotient.mk'' (x i)) fun x y hxy ↦ by + toFun x := Quotient.mk'' fun i ↦ (x i).out + invFun q := Quotient.liftOn' q (fun x i ↦ Quotient.mk'' (x i)) fun x y hxy ↦ by ext i simpa using hxy i - left_inv := fun q ↦ by + left_inv q := by ext i simp - right_inv := fun q ↦ by + right_inv q := by refine Quotient.inductionOn' q fun _ ↦ ?_ simp only [Quotient.liftOn'_mk'', Quotient.eq''] intro i @@ -346,8 +346,7 @@ def liftEquiv (r : Setoid α) : { f : α → β // r ≤ ker f } ≃ (Quotient r theorem lift_unique {r : Setoid α} {f : α → β} (H : r ≤ ker f) (g : Quotient r → β) (Hg : f = g ∘ Quotient.mk'') : Quotient.lift f H = g := by ext ⟨x⟩ - erw [Quotient.lift_mk f H, Hg] - rfl + rw [← Quotient.mk, Quotient.lift_mk f H, Hg, Function.comp_apply, Quotient.mk''_eq_mk] /-- Given a map f from α to β, the natural map from the quotient of α by the kernel of f is injective. -/ @@ -464,8 +463,8 @@ def correspondence (r : Setoid α) : { s // r ≤ s } ≃o Setoid (Quotient r) w toFun s := ⟨Quotient.lift₂ s.1.1 fun _ _ _ _ h₁ h₂ ↦ Eq.propIntro (fun h ↦ s.1.trans' (s.1.trans' (s.1.symm' (s.2 h₁)) h) (s.2 h₂)) (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⟩⟩ + ⟨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, Quotient.eq'.2 h]⟩ left_inv _ := rfl right_inv _ := ext fun x y ↦ Quotient.inductionOn₂ x y fun _ _ ↦ Iff.rfl diff --git a/Mathlib/FieldTheory/CardinalEmb.lean b/Mathlib/FieldTheory/CardinalEmb.lean index 3f51c66e061a4d..307b1bf0c022c9 100644 --- a/Mathlib/FieldTheory/CardinalEmb.lean +++ b/Mathlib/FieldTheory/CardinalEmb.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ import Mathlib.FieldTheory.SeparableClosure -import Mathlib.FieldTheory.PurelyInseparable +import Mathlib.FieldTheory.PurelyInseparable.Basic import Mathlib.LinearAlgebra.FreeAlgebra import Mathlib.Order.Interval.Set.WithBotTop import Mathlib.Order.DirectedInverseSystem diff --git a/Mathlib/FieldTheory/Differential/Liouville.lean b/Mathlib/FieldTheory/Differential/Liouville.lean index cf9795aae7da6b..4d84efa15e4013 100644 --- a/Mathlib/FieldTheory/Differential/Liouville.lean +++ b/Mathlib/FieldTheory/Differential/Liouville.lean @@ -66,7 +66,7 @@ lemma IsLiouville.trans {A : Type*} [Field A] [Algebra K A] [Algebra F A] apply_fun ((↑) : F → K) · simp only [Function.comp_apply, coe_deriv, hc, algebraMap.coe_zero] apply hc₀ - · apply NoZeroSMulDivisors.algebraMap_injective + · apply FaithfulSMul.algebraMap_injective section Algebraic /- @@ -120,7 +120,7 @@ private local instance isLiouville_of_finiteDimensional_galois [FiniteDimensiona isLiouville (a : F) (ι : Type) [Fintype ι] (c : ι → F) (hc : ∀ x, (c x)′ = 0) (u : ι → K) (v : K) (h : a = ∑ x, c x * logDeriv (u x) + v′) := by haveI : CharZero K := charZero_of_injective_algebraMap - (NoZeroSMulDivisors.algebraMap_injective F K) + (FaithfulSMul.algebraMap_injective F K) -- We sum `e x` over all isomorphisms `e : K ≃ₐ[F] K`. -- Because this is a Galois extension each of the relevant values will be in `F`. -- We need to divide by `Fintype.card (K ≃ₐ[F] K)` to get the original answer. @@ -159,7 +159,7 @@ private local instance isLiouville_of_finiteDimensional_galois [FiniteDimensiona · -- Proving that this works is mostly straightforward algebraic manipulation, apply_fun (algebraMap F K) case inj => - exact NoZeroSMulDivisors.algebraMap_injective F K + exact FaithfulSMul.algebraMap_injective F K simp only [map_add, map_sum, map_mul, ← logDeriv_algebraMap, hu₀, ← deriv_algebraMap, hv₀] unfold u₁ v₁ c₀ clear c₀ u₁ u₀ hu₀ v₁ v₀ hv₀ diff --git a/Mathlib/FieldTheory/IntermediateField/Adjoin/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Adjoin/Basic.lean index b7716cd8544784..0852aad6f9be59 100644 --- a/Mathlib/FieldTheory/IntermediateField/Adjoin/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Adjoin/Basic.lean @@ -110,13 +110,11 @@ theorem toSubalgebra_iSup_of_directed (dir : Directed (· ≤ ·) t) : instance finiteDimensional_iSup_of_finite [h : Finite ι] [∀ i, FiniteDimensional K (t i)] : FiniteDimensional K (⨆ i, t i : IntermediateField K L) := by rw [← iSup_univ] - refine Set.Finite.induction_on - (motive := fun s _ => FiniteDimensional K (⨆ i ∈ s, t i : IntermediateField K L)) - _ Set.finite_univ ?_ ?_ - all_goals dsimp - · rw [iSup_emptyset] + induction Set.univ, Set.finite_univ (α := ι) using Set.Finite.induction_on with + | empty => + rw [iSup_emptyset] exact (botEquiv K L).symm.toLinearEquiv.finiteDimensional - · intro _ s _ _ hs + | insert s hs => rw [iSup_insert] exact IntermediateField.finiteDimensional_sup _ _ diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index f64c0cb3f0efdc..a691ac27315c17 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -150,7 +150,7 @@ theorem of_ringEquiv (k' : Type u) [Field k'] (e : k ≃+* k') have hpe : degree (p.map e.symm.toRingHom) ≠ 0 := by rw [degree_map] exact ne_of_gt (degree_pos_of_irreducible hp) - rcases IsAlgClosed.exists_root (k := k) (p.map e.symm) hpe with ⟨x, hx⟩ + rcases IsAlgClosed.exists_root (k := k) (p.map e.symm.toRingHom) hpe with ⟨x, hx⟩ use e x rw [IsRoot] at hx apply e.symm.injective @@ -279,10 +279,10 @@ variable {S : Type v} [CommRing S] [IsDomain S] [Algebra R S] [Algebra R M] [NoZ variable {M} private instance FractionRing.isAlgebraic : - letI : IsDomain R := (NoZeroSMulDivisors.algebraMap_injective R S).isDomain _ + letI : IsDomain R := (FaithfulSMul.algebraMap_injective R S).isDomain _ letI : Algebra (FractionRing R) (FractionRing S) := FractionRing.liftAlgebra R _ Algebra.IsAlgebraic (FractionRing R) (FractionRing S) := by - letI : IsDomain R := (NoZeroSMulDivisors.algebraMap_injective R S).isDomain _ + letI : IsDomain R := (FaithfulSMul.algebraMap_injective R S).isDomain _ letI : Algebra (FractionRing R) (FractionRing S) := FractionRing.liftAlgebra R _ have := FractionRing.isScalarTower_liftAlgebra R (FractionRing S) have := (IsFractionRing.isAlgebraic_iff' R S (FractionRing S)).1 inferInstance @@ -295,7 +295,7 @@ private instance FractionRing.isAlgebraic : closed extension of R. -/ @[stacks 09GU] noncomputable irreducible_def lift : S →ₐ[R] M := by - letI : IsDomain R := (NoZeroSMulDivisors.algebraMap_injective R S).isDomain _ + letI : IsDomain R := (FaithfulSMul.algebraMap_injective R S).isDomain _ letI := FractionRing.liftAlgebra R M letI := FractionRing.liftAlgebra R (FractionRing S) have := FractionRing.isScalarTower_liftAlgebra R M @@ -370,11 +370,8 @@ theorem ofAlgebraic [hKJ : Algebra.IsAlgebraic K J] : IsAlgClosure K L := an algebraic extension of `R` -/ noncomputable def equivOfAlgebraic' [Nontrivial S] [NoZeroSMulDivisors R S] [Algebra.IsAlgebraic R L] : L ≃ₐ[R] M := by - letI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.of_algebraMap_injective <| by - rw [IsScalarTower.algebraMap_eq R S L] - exact (Function.Injective.comp (NoZeroSMulDivisors.algebraMap_injective S L) - (NoZeroSMulDivisors.algebraMap_injective R S) :) - letI : IsAlgClosure R L := + have : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans_faithfulSMul R S L + have : IsAlgClosure R L := { isAlgClosed := IsAlgClosure.isAlgClosed S isAlgebraic := ‹_› } exact IsAlgClosure.equiv _ _ _ @@ -396,13 +393,13 @@ noncomputable def equivOfEquivAux (hSR : S ≃+* R) : { e : L ≃+* M // e.toRingHom.comp (algebraMap S L) = (algebraMap R M).comp hSR.toRingHom } := by letI : Algebra R S := RingHom.toAlgebra hSR.symm.toRingHom letI : Algebra S R := RingHom.toAlgebra hSR.toRingHom - letI : IsDomain R := (NoZeroSMulDivisors.algebraMap_injective R M).isDomain _ - letI : IsDomain S := (NoZeroSMulDivisors.algebraMap_injective S L).isDomain _ + letI : IsDomain R := (FaithfulSMul.algebraMap_injective R M).isDomain _ + letI : IsDomain S := (FaithfulSMul.algebraMap_injective S L).isDomain _ letI : Algebra R L := RingHom.toAlgebra ((algebraMap S L).comp (algebraMap R S)) haveI : IsScalarTower R S L := IsScalarTower.of_algebraMap_eq fun _ => rfl haveI : IsScalarTower S R L := IsScalarTower.of_algebraMap_eq (by simp [RingHom.algebraMap_toAlgebra]) - haveI : NoZeroSMulDivisors R S := NoZeroSMulDivisors.of_algebraMap_injective hSR.symm.injective + have : FaithfulSMul R S := (faithfulSMul_iff_algebraMap_injective R S).mpr hSR.symm.injective have : Algebra.IsAlgebraic R L := (IsAlgClosure.isAlgebraic.extendScalars (show Function.Injective (algebraMap S R) from hSR.injective)) refine diff --git a/Mathlib/FieldTheory/IsPerfectClosure.lean b/Mathlib/FieldTheory/IsPerfectClosure.lean index db75df108ee8cd..247993e778f5fa 100644 --- a/Mathlib/FieldTheory/IsPerfectClosure.lean +++ b/Mathlib/FieldTheory/IsPerfectClosure.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Jz Pan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jz Pan -/ -import Mathlib.FieldTheory.PurelyInseparable +import Mathlib.FieldTheory.PurelyInseparable.Basic import Mathlib.FieldTheory.PerfectClosure /-! diff --git a/Mathlib/FieldTheory/JacobsonNoether.lean b/Mathlib/FieldTheory/JacobsonNoether.lean index e019eaf9f42310..44f523cd54a239 100644 --- a/Mathlib/FieldTheory/JacobsonNoether.lean +++ b/Mathlib/FieldTheory/JacobsonNoether.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.CharP.LinearMaps import Mathlib.Algebra.CharP.Subring import Mathlib.Algebra.GroupWithZero.Conj import Mathlib.Algebra.Lie.OfAssociative -import Mathlib.FieldTheory.PurelyInseparable +import Mathlib.FieldTheory.PurelyInseparable.Basic /-! # The Jacobson-Noether theorem diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 352ba6623e838b..21d45c381d28d5 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -73,7 +73,7 @@ lemma X_pow_sub_C_eq_prod {R : Type*} [CommRing R] [IsDomain R] (X ^ n - C a) = ∏ i ∈ Finset.range n, (X - C (ζ ^ i * α)) := by let K := FractionRing R let i := algebraMap R K - have h := NoZeroSMulDivisors.algebraMap_injective R K + have h := FaithfulSMul.algebraMap_injective R K apply_fun Polynomial.map i using map_injective i h simpa only [Polynomial.map_sub, Polynomial.map_pow, map_X, map_C, map_mul, map_pow, Polynomial.map_prod, Polynomial.map_mul] diff --git a/Mathlib/FieldTheory/Normal/Defs.lean b/Mathlib/FieldTheory/Normal/Defs.lean index a8977634afdf3d..7d3a150c439528 100644 --- a/Mathlib/FieldTheory/Normal/Defs.lean +++ b/Mathlib/FieldTheory/Normal/Defs.lean @@ -188,7 +188,7 @@ def Normal.algHomEquivAut [Normal F E] : (E →ₐ[F] K₁) ≃ E ≃ₐ[F] E wh right_inv σ := by ext simp only [AlgHom.restrictNormal', AlgEquiv.toAlgHom_eq_coe, AlgEquiv.coe_ofBijective] - apply NoZeroSMulDivisors.algebraMap_injective E K₁ + apply FaithfulSMul.algebraMap_injective E K₁ rw [AlgHom.restrictNormal_commutes] simp diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean deleted file mode 100644 index e38f330fe860dc..00000000000000 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ /dev/null @@ -1,1180 +0,0 @@ -/- -Copyright (c) 2024 Jz Pan. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jz Pan --/ -import Mathlib.Algebra.CharP.ExpChar -import Mathlib.Algebra.CharP.IntermediateField -import Mathlib.FieldTheory.SeparableClosure - -/-! - -# Purely inseparable extension and relative perfect closure - -This file contains basics about purely inseparable extensions and the relative perfect closure -of fields. - -## Main definitions - -- `IsPurelyInseparable`: typeclass for purely inseparable field extensions: an algebraic extension - `E / F` is purely inseparable if and only if the minimal polynomial of every element of `E ∖ F` - is not separable. - -- `perfectClosure`: the relative perfect closure of `F` in `E`, it consists of the elements - `x` of `E` such that there exists a natural number `n` such that `x ^ (ringExpChar F) ^ n` - is contained in `F`, where `ringExpChar F` is the exponential characteristic of `F`. - It is also the maximal purely inseparable subextension of `E / F` (`le_perfectClosure_iff`). - -## Main results - -- `IsPurelyInseparable.surjective_algebraMap_of_isSeparable`, - `IsPurelyInseparable.bijective_algebraMap_of_isSeparable`, - `IntermediateField.eq_bot_of_isPurelyInseparable_of_isSeparable`: - if `E / F` is both purely inseparable and separable, then `algebraMap F E` is surjective - (hence bijective). In particular, if an intermediate field of `E / F` is both purely inseparable - and separable, then it is equal to `F`. - -- `isPurelyInseparable_iff_pow_mem`: a field extension `E / F` of exponential characteristic `q` is - purely inseparable if and only if for every element `x` of `E`, there exists a natural number `n` - such that `x ^ (q ^ n)` is contained in `F`. - -- `IsPurelyInseparable.trans`: if `E / F` and `K / E` are both purely inseparable extensions, then - `K / F` is also purely inseparable. - -- `isPurelyInseparable_iff_natSepDegree_eq_one`: `E / F` is purely inseparable if and only if for - every element `x` of `E`, its minimal polynomial has separable degree one. - -- `isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C`: a field extension `E / F` of exponential - characteristic `q` is purely inseparable if and only if for every element `x` of `E`, the minimal - polynomial of `x` over `F` is of form `X ^ (q ^ n) - y` for some natural number `n` and some - element `y` of `F`. - -- `isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow`: a field extension `E / F` of exponential - characteristic `q` is purely inseparable if and only if for every element `x` of `E`, the minimal - polynomial of `x` over `F` is of form `(X - x) ^ (q ^ n)` for some natural number `n`. - -- `isPurelyInseparable_iff_finSepDegree_eq_one`: an extension is purely inseparable - if and only if it has finite separable degree (`Field.finSepDegree`) one. - -- `IsPurelyInseparable.normal`: a purely inseparable extension is normal. - -- `separableClosure.isPurelyInseparable`: if `E / F` is algebraic, then `E` is purely inseparable - over the separable closure of `F` in `E`. - -- `separableClosure_le_iff`: if `E / F` is algebraic, then an intermediate field of `E / F` contains - the separable closure of `F` in `E` if and only if `E` is purely inseparable over it. - -- `eq_separableClosure_iff`: if `E / F` is algebraic, then an intermediate field of `E / F` is equal - to the separable closure of `F` in `E` if and only if it is separable over `F`, and `E` - is purely inseparable over it. - -- `le_perfectClosure_iff`: an intermediate field of `E / F` is contained in the relative perfect - closure of `F` in `E` if and only if it is purely inseparable over `F`. - -- `perfectClosure.perfectRing`, `perfectClosure.perfectField`: if `E` is a perfect field, then the - (relative) perfect closure `perfectClosure F E` is perfect. - -- `IsPurelyInseparable.injective_comp_algebraMap`: if `E / F` is purely inseparable, then for any - reduced ring `L`, the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is injective. - In particular, a purely inseparable field extension is an epimorphism in the category of fields. - -- `IsPurelyInseparable.of_injective_comp_algebraMap`: if `L` is an algebraically closed field - containing `E`, such that the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is - injective, then `E / F` is purely inseparable. As a corollary, epimorphisms in the category of - fields must be purely inseparable extensions. - -- `IntermediateField.isPurelyInseparable_adjoin_iff_pow_mem`: if `F` is of exponential - characteristic `q`, then `F(S) / F` is a purely inseparable extension if and only if for any - `x ∈ S`, `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. - -- `Field.finSepDegree_eq`: if `E / F` is algebraic, then the `Field.finSepDegree F E` is equal to - `Field.sepDegree F E` as a natural number. This means that the cardinality of `Field.Emb F E` - and the degree of `(separableClosure F E) / F` are both finite or infinite, and when they are - finite, they coincide. - -- `Field.finSepDegree_mul_finInsepDegree`: the finite separable degree multiply by the finite - inseparable degree is equal to the (finite) field extension degree. - -- `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`: the separable degrees satisfy the - tower law: $[E:F]_s [K:E]_s = [K:F]_s$. - -- `IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable`, - `IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable'`: - if `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then - for any subset `S` of `K` such that `F(S) / F` is algebraic, the `E(S) / E` and `F(S) / F` have - the same separable degree. In particular, if `S` is an intermediate field of `K / F` such that - `S / F` is algebraic, the `E(S) / E` and `S / F` have the same separable degree. - -- `minpoly.map_eq_of_isSeparable_of_isPurelyInseparable`: if `K / E / F` is a field extension tower, - such that `E / F` is purely inseparable, then for any element `x` of `K` separable over `F`, - it has the same minimal polynomials over `F` and over `E`. - -- `Polynomial.Separable.map_irreducible_of_isPurelyInseparable`: if `E / F` is purely inseparable, - `f` is a separable irreducible polynomial over `F`, then it is also irreducible over `E`. - -## Tags - -separable degree, degree, separable closure, purely inseparable - -## TODO - -- Restate some intermediate result in terms of linearly disjointness. - -- Prove that the inseparable degrees satisfy the tower law: $[E:F]_i [K:E]_i = [K:F]_i$. - Probably an argument using linearly disjointness is needed. - --/ - -open Module Polynomial IntermediateField Field Finsupp - -noncomputable section - -universe u v w - -section IsPurelyInseparable - -variable (F : Type u) (E : Type v) [CommRing F] [Ring E] [Algebra F E] -variable (K : Type w) [Ring K] [Algebra F K] - -/-- Typeclass for purely inseparable field extensions: an algebraic extension `E / F` is purely -inseparable if and only if the minimal polynomial of every element of `E ∖ F` is not separable. - -We define this for general (commutative) rings and only assume `F` and `E` are fields -if this is needed for a proof. -/ -class IsPurelyInseparable : Prop where - isIntegral : Algebra.IsIntegral F E - inseparable' (x : E) : IsSeparable F x → x ∈ (algebraMap F E).range - -attribute [instance] IsPurelyInseparable.isIntegral - -variable {E} in -theorem IsPurelyInseparable.isIntegral' [IsPurelyInseparable F E] (x : E) : IsIntegral F x := - Algebra.IsIntegral.isIntegral _ - -theorem IsPurelyInseparable.isAlgebraic [Nontrivial F] [IsPurelyInseparable F E] : - Algebra.IsAlgebraic F E := inferInstance - -variable {E} - -theorem IsPurelyInseparable.inseparable [IsPurelyInseparable F E] : - ∀ x : E, IsSeparable F x → x ∈ (algebraMap F E).range := - IsPurelyInseparable.inseparable' - -variable {F K} - -theorem isPurelyInseparable_iff : IsPurelyInseparable F E ↔ ∀ x : E, - IsIntegral F x ∧ (IsSeparable F x → x ∈ (algebraMap F E).range) := - ⟨fun h x ↦ ⟨h.isIntegral' _ x, h.inseparable' x⟩, fun h ↦ ⟨⟨fun x ↦ (h x).1⟩, fun x ↦ (h x).2⟩⟩ - -/-- Transfer `IsPurelyInseparable` across an `AlgEquiv`. -/ -theorem AlgEquiv.isPurelyInseparable (e : K ≃ₐ[F] E) [IsPurelyInseparable F K] : - IsPurelyInseparable F E := by - refine ⟨⟨fun _ ↦ by rw [← isIntegral_algEquiv e.symm]; exact IsPurelyInseparable.isIntegral' F _⟩, - fun x h ↦ ?_⟩ - rw [IsSeparable, ← minpoly.algEquiv_eq e.symm] at h - simpa only [RingHom.mem_range, algebraMap_eq_apply] using IsPurelyInseparable.inseparable F _ h - -theorem AlgEquiv.isPurelyInseparable_iff (e : K ≃ₐ[F] E) : - IsPurelyInseparable F K ↔ IsPurelyInseparable F E := - ⟨fun _ ↦ e.isPurelyInseparable, fun _ ↦ e.symm.isPurelyInseparable⟩ - -/-- If `E / F` is an algebraic extension, `F` is separably closed, -then `E / F` is purely inseparable. -/ -instance Algebra.IsAlgebraic.isPurelyInseparable_of_isSepClosed - {F : Type u} {E : Type v} [Field F] [Ring E] [IsDomain E] [Algebra F E] - [Algebra.IsAlgebraic F E] - [IsSepClosed F] : IsPurelyInseparable F E := - ⟨inferInstance, fun x h ↦ minpoly.mem_range_of_degree_eq_one F x <| - IsSepClosed.degree_eq_one_of_irreducible F (minpoly.irreducible - (Algebra.IsIntegral.isIntegral _)) h⟩ - -variable (F E K) - -/-- If `E / F` is both purely inseparable and separable, then `algebraMap F E` is surjective. -/ -theorem IsPurelyInseparable.surjective_algebraMap_of_isSeparable - [IsPurelyInseparable F E] [Algebra.IsSeparable F E] : Function.Surjective (algebraMap F E) := - fun x ↦ IsPurelyInseparable.inseparable F x (Algebra.IsSeparable.isSeparable F x) - -/-- If `E / F` is both purely inseparable and separable, then `algebraMap F E` is bijective. -/ -theorem IsPurelyInseparable.bijective_algebraMap_of_isSeparable - [Nontrivial E] [NoZeroSMulDivisors F E] - [IsPurelyInseparable F E] [Algebra.IsSeparable F E] : Function.Bijective (algebraMap F E) := - ⟨NoZeroSMulDivisors.algebraMap_injective F E, surjective_algebraMap_of_isSeparable F E⟩ - -variable {F E} in -/-- If a subalgebra of `E / F` is both purely inseparable and separable, then it is equal -to `F`. -/ -theorem Subalgebra.eq_bot_of_isPurelyInseparable_of_isSeparable (L : Subalgebra F E) - [IsPurelyInseparable F L] [Algebra.IsSeparable F L] : L = ⊥ := bot_unique fun x hx ↦ by - obtain ⟨y, hy⟩ := IsPurelyInseparable.surjective_algebraMap_of_isSeparable F L ⟨x, hx⟩ - exact ⟨y, congr_arg (Subalgebra.val _) hy⟩ - -/-- If an intermediate field of `E / F` is both purely inseparable and separable, then it is equal -to `F`. -/ -theorem IntermediateField.eq_bot_of_isPurelyInseparable_of_isSeparable - {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] (L : IntermediateField F E) - [IsPurelyInseparable F L] [Algebra.IsSeparable F L] : L = ⊥ := bot_unique fun x hx ↦ by - obtain ⟨y, hy⟩ := IsPurelyInseparable.surjective_algebraMap_of_isSeparable F L ⟨x, hx⟩ - exact ⟨y, congr_arg (algebraMap L E) hy⟩ - -/-- If `E / F` is purely inseparable, then the separable closure of `F` in `E` is -equal to `F`. -/ -theorem separableClosure.eq_bot_of_isPurelyInseparable - (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [IsPurelyInseparable F E] : - separableClosure F E = ⊥ := - bot_unique fun x h ↦ IsPurelyInseparable.inseparable F x (mem_separableClosure_iff.1 h) - -/-- If `E / F` is an algebraic extension, then the separable closure of `F` in `E` is -equal to `F` if and only if `E / F` is purely inseparable. -/ -theorem separableClosure.eq_bot_iff - {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] [Algebra.IsAlgebraic F E] : - separableClosure F E = ⊥ ↔ IsPurelyInseparable F E := - ⟨fun h ↦ isPurelyInseparable_iff.2 fun x ↦ ⟨Algebra.IsIntegral.isIntegral x, fun hs ↦ by - simpa only [h] using mem_separableClosure_iff.2 hs⟩, fun _ ↦ eq_bot_of_isPurelyInseparable F E⟩ - -instance isPurelyInseparable_self : IsPurelyInseparable F F := - ⟨inferInstance, fun x _ ↦ ⟨x, rfl⟩⟩ - -section - -variable (F : Type u) {E : Type v} [Field F] [Ring E] [IsDomain E] [Algebra F E] -variable (q : ℕ) [ExpChar F q] (x : E) - -/-- A field extension `E / F` of exponential characteristic `q` is purely inseparable -if and only if for every element `x` of `E`, there exists a natural number `n` such that -`x ^ (q ^ n)` is contained in `F`. -/ -@[stacks 09HE] -theorem isPurelyInseparable_iff_pow_mem : - IsPurelyInseparable F E ↔ ∀ x : E, ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by - rw [isPurelyInseparable_iff] - refine ⟨fun h x ↦ ?_, fun h x ↦ ?_⟩ - · obtain ⟨g, h1, n, h2⟩ := (minpoly.irreducible (h x).1).hasSeparableContraction q - exact ⟨n, (h _).2 <| h1.of_dvd <| minpoly.dvd F _ <| by - simpa only [expand_aeval, minpoly.aeval] using congr_arg (aeval x) h2⟩ - have hdeg := (minpoly.natSepDegree_eq_one_iff_pow_mem q).2 (h x) - have halg : IsIntegral F x := by_contra fun h' ↦ by - simp only [minpoly.eq_zero h', natSepDegree_zero, zero_ne_one] at hdeg - refine ⟨halg, fun hsep ↦ ?_⟩ - rwa [hsep.natSepDegree_eq_natDegree, minpoly.natDegree_eq_one_iff] at hdeg - -theorem IsPurelyInseparable.pow_mem [IsPurelyInseparable F E] : - ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := - (isPurelyInseparable_iff_pow_mem F q).1 ‹_› x - -end - -end IsPurelyInseparable - -variable (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] -variable (K : Type w) [Field K] [Algebra F K] - -section perfectClosure - -/-- The relative perfect closure of `F` in `E`, consists of the elements `x` of `E` such that there -exists a natural number `n` such that `x ^ (ringExpChar F) ^ n` is contained in `F`, where -`ringExpChar F` is the exponential characteristic of `F`. It is also the maximal purely inseparable -subextension of `E / F` (`le_perfectClosure_iff`). -/ -@[stacks 09HH] -def perfectClosure : IntermediateField F E where - carrier := {x : E | ∃ n : ℕ, x ^ (ringExpChar F) ^ n ∈ (algebraMap F E).range} - add_mem' := by - rintro x y ⟨n, hx⟩ ⟨m, hy⟩ - use n + m - have := expChar_of_injective_algebraMap (algebraMap F E).injective (ringExpChar F) - rw [add_pow_expChar_pow, pow_add, pow_mul, mul_comm (_ ^ n), pow_mul] - exact add_mem (pow_mem hx _) (pow_mem hy _) - mul_mem' := by - rintro x y ⟨n, hx⟩ ⟨m, hy⟩ - use n + m - rw [mul_pow, pow_add, pow_mul, mul_comm (_ ^ n), pow_mul] - exact mul_mem (pow_mem hx _) (pow_mem hy _) - inv_mem' := by - rintro x ⟨n, hx⟩ - use n; rw [inv_pow] - apply inv_mem (id hx : _ ∈ (⊥ : IntermediateField F E)) - algebraMap_mem' := fun x ↦ ⟨0, by rw [pow_zero, pow_one]; exact ⟨x, rfl⟩⟩ - -variable {F E} - -theorem mem_perfectClosure_iff {x : E} : - x ∈ perfectClosure F E ↔ ∃ n : ℕ, x ^ (ringExpChar F) ^ n ∈ (algebraMap F E).range := Iff.rfl - -theorem mem_perfectClosure_iff_pow_mem (q : ℕ) [ExpChar F q] {x : E} : - x ∈ perfectClosure F E ↔ ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by - rw [mem_perfectClosure_iff, ringExpChar.eq F q] - -/-- An element is contained in the relative perfect closure if and only if its minimal polynomial -has separable degree one. -/ -theorem mem_perfectClosure_iff_natSepDegree_eq_one {x : E} : - x ∈ perfectClosure F E ↔ (minpoly F x).natSepDegree = 1 := by - rw [mem_perfectClosure_iff, minpoly.natSepDegree_eq_one_iff_pow_mem (ringExpChar F)] - -/-- A field extension `E / F` is purely inseparable if and only if the relative perfect closure of -`F` in `E` is equal to `E`. -/ -theorem isPurelyInseparable_iff_perfectClosure_eq_top : - IsPurelyInseparable F E ↔ perfectClosure F E = ⊤ := by - rw [isPurelyInseparable_iff_pow_mem F (ringExpChar F)] - exact ⟨fun H ↦ top_unique fun x _ ↦ H x, fun H _ ↦ H.ge trivial⟩ - -variable (F E) - -/-- The relative perfect closure of `F` in `E` is purely inseparable over `F`. -/ -instance perfectClosure.isPurelyInseparable : IsPurelyInseparable F (perfectClosure F E) := by - rw [isPurelyInseparable_iff_pow_mem F (ringExpChar F)] - exact fun ⟨_, n, y, h⟩ ↦ ⟨n, y, (algebraMap _ E).injective h⟩ - -/-- The relative perfect closure of `F` in `E` is algebraic over `F`. -/ -instance perfectClosure.isAlgebraic : Algebra.IsAlgebraic F (perfectClosure F E) := - IsPurelyInseparable.isAlgebraic F _ - -/-- If `E / F` is separable, then the perfect closure of `F` in `E` is equal to `F`. Note that - the converse is not necessarily true (see https://math.stackexchange.com/a/3009197) - even when `E / F` is algebraic. -/ -theorem perfectClosure.eq_bot_of_isSeparable [Algebra.IsSeparable F E] : perfectClosure F E = ⊥ := - haveI := Algebra.isSeparable_tower_bot_of_isSeparable F (perfectClosure F E) E - eq_bot_of_isPurelyInseparable_of_isSeparable _ - -/-- An intermediate field of `E / F` is contained in the relative perfect closure of `F` in `E` -if it is purely inseparable over `F`. -/ -theorem le_perfectClosure (L : IntermediateField F E) [h : IsPurelyInseparable F L] : - L ≤ perfectClosure F E := by - rw [isPurelyInseparable_iff_pow_mem F (ringExpChar F)] at h - intro x hx - obtain ⟨n, y, hy⟩ := h ⟨x, hx⟩ - exact ⟨n, y, congr_arg (algebraMap L E) hy⟩ - -/-- An intermediate field of `E / F` is contained in the relative perfect closure of `F` in `E` -if and only if it is purely inseparable over `F`. -/ -theorem le_perfectClosure_iff (L : IntermediateField F E) : - L ≤ perfectClosure F E ↔ IsPurelyInseparable F L := by - refine ⟨fun h ↦ (isPurelyInseparable_iff_pow_mem F (ringExpChar F)).2 fun x ↦ ?_, - fun _ ↦ le_perfectClosure F E L⟩ - obtain ⟨n, y, hy⟩ := h x.2 - exact ⟨n, y, (algebraMap L E).injective hy⟩ - -theorem separableClosure_inf_perfectClosure : separableClosure F E ⊓ perfectClosure F E = ⊥ := - haveI := (le_separableClosure_iff F E _).mp (inf_le_left (b := perfectClosure F E)) - haveI := (le_perfectClosure_iff F E _).mp (inf_le_right (a := separableClosure F E)) - eq_bot_of_isPurelyInseparable_of_isSeparable _ - -section map - -variable {F E K} - -/-- If `i` is an `F`-algebra homomorphism from `E` to `K`, then `i x` is contained in -`perfectClosure F K` if and only if `x` is contained in `perfectClosure F E`. -/ -theorem map_mem_perfectClosure_iff (i : E →ₐ[F] K) {x : E} : - i x ∈ perfectClosure F K ↔ x ∈ perfectClosure F E := by - simp_rw [mem_perfectClosure_iff] - refine ⟨fun ⟨n, y, h⟩ ↦ ⟨n, y, ?_⟩, fun ⟨n, y, h⟩ ↦ ⟨n, y, ?_⟩⟩ - · apply_fun i using i.injective - rwa [AlgHom.commutes, map_pow] - simpa only [AlgHom.commutes, map_pow] using congr_arg i h - -/-- If `i` is an `F`-algebra homomorphism from `E` to `K`, then the preimage of `perfectClosure F K` -under the map `i` is equal to `perfectClosure F E`. -/ -theorem perfectClosure.comap_eq_of_algHom (i : E →ₐ[F] K) : - (perfectClosure F K).comap i = perfectClosure F E := by - ext x - exact map_mem_perfectClosure_iff i - -/-- If `i` is an `F`-algebra homomorphism from `E` to `K`, then the image of `perfectClosure F E` -under the map `i` is contained in `perfectClosure F K`. -/ -theorem perfectClosure.map_le_of_algHom (i : E →ₐ[F] K) : - (perfectClosure F E).map i ≤ perfectClosure F K := - map_le_iff_le_comap.mpr (perfectClosure.comap_eq_of_algHom i).ge - -/-- If `i` is an `F`-algebra isomorphism of `E` and `K`, then the image of `perfectClosure F E` -under the map `i` is equal to in `perfectClosure F K`. -/ -theorem perfectClosure.map_eq_of_algEquiv (i : E ≃ₐ[F] K) : - (perfectClosure F E).map i.toAlgHom = perfectClosure F K := - (map_le_of_algHom i.toAlgHom).antisymm (fun x hx ↦ ⟨i.symm x, - (map_mem_perfectClosure_iff i.symm.toAlgHom).2 hx, i.right_inv x⟩) - -/-- If `E` and `K` are isomorphic as `F`-algebras, then `perfectClosure F E` and -`perfectClosure F K` are also isomorphic as `F`-algebras. -/ -def perfectClosure.algEquivOfAlgEquiv (i : E ≃ₐ[F] K) : - perfectClosure F E ≃ₐ[F] perfectClosure F K := - (intermediateFieldMap i _).trans (equivOfEq (map_eq_of_algEquiv i)) - -alias AlgEquiv.perfectClosure := perfectClosure.algEquivOfAlgEquiv - -end map - -/-- If `E` is a perfect field of exponential characteristic `p`, then the (relative) perfect closure -`perfectClosure F E` is perfect. -/ -instance perfectClosure.perfectRing (p : ℕ) [ExpChar E p] - [PerfectRing E p] : PerfectRing (perfectClosure F E) p := .ofSurjective _ p fun x ↦ by - haveI := RingHom.expChar _ (algebraMap F E).injective p - obtain ⟨x', hx⟩ := surjective_frobenius E p x.1 - obtain ⟨n, y, hy⟩ := (mem_perfectClosure_iff_pow_mem p).1 x.2 - rw [frobenius_def] at hx - rw [← hx, ← pow_mul, ← pow_succ'] at hy - exact ⟨⟨x', (mem_perfectClosure_iff_pow_mem p).2 ⟨n + 1, y, hy⟩⟩, by - simp_rw [frobenius_def, SubmonoidClass.mk_pow, hx]⟩ - -/-- If `E` is a perfect field, then the (relative) perfect closure -`perfectClosure F E` is perfect. -/ -instance perfectClosure.perfectField [PerfectField E] : PerfectField (perfectClosure F E) := - PerfectRing.toPerfectField _ (ringExpChar E) - -end perfectClosure - -section IsPurelyInseparable - -/-- If `K / E / F` is a field extension tower such that `K / F` is purely inseparable, -then `E / F` is also purely inseparable. -/ -theorem IsPurelyInseparable.tower_bot [Algebra E K] [IsScalarTower F E K] - [IsPurelyInseparable F K] : IsPurelyInseparable F E := by - refine ⟨⟨fun x ↦ (isIntegral' F (algebraMap E K x)).tower_bot_of_field⟩, fun x h ↦ ?_⟩ - rw [IsSeparable, ← minpoly.algebraMap_eq (algebraMap E K).injective] at h - obtain ⟨y, h⟩ := inseparable F _ h - exact ⟨y, (algebraMap E K).injective (h.symm ▸ (IsScalarTower.algebraMap_apply F E K y).symm)⟩ - -/-- If `K / E / F` is a field extension tower such that `K / F` is purely inseparable, -then `K / E` is also purely inseparable. -/ -theorem IsPurelyInseparable.tower_top [Algebra E K] [IsScalarTower F E K] - [h : IsPurelyInseparable F K] : IsPurelyInseparable E K := by - obtain ⟨q, _⟩ := ExpChar.exists F - haveI := expChar_of_injective_algebraMap (algebraMap F E).injective q - rw [isPurelyInseparable_iff_pow_mem _ q] at h ⊢ - intro x - obtain ⟨n, y, h⟩ := h x - exact ⟨n, (algebraMap F E) y, h.symm ▸ (IsScalarTower.algebraMap_apply F E K y).symm⟩ - -/-- If `E / F` and `K / E` are both purely inseparable extensions, then `K / F` is also -purely inseparable. -/ -@[stacks 02JJ "See also 00GM"] -theorem IsPurelyInseparable.trans [Algebra E K] [IsScalarTower F E K] - [h1 : IsPurelyInseparable F E] [h2 : IsPurelyInseparable E K] : IsPurelyInseparable F K := by - obtain ⟨q, _⟩ := ExpChar.exists F - haveI := expChar_of_injective_algebraMap (algebraMap F E).injective q - rw [isPurelyInseparable_iff_pow_mem _ q] at h1 h2 ⊢ - intro x - obtain ⟨n, y, h2⟩ := h2 x - obtain ⟨m, z, h1⟩ := h1 y - refine ⟨n + m, z, ?_⟩ - rw [IsScalarTower.algebraMap_apply F E K, h1, map_pow, h2, ← pow_mul, ← pow_add] - -namespace IntermediateField - -variable (M : IntermediateField F K) - -instance isPurelyInseparable_tower_bot [IsPurelyInseparable F K] : IsPurelyInseparable F M := - IsPurelyInseparable.tower_bot F M K - -instance isPurelyInseparable_tower_top [IsPurelyInseparable F K] : IsPurelyInseparable M K := - IsPurelyInseparable.tower_top F M K - -end IntermediateField - -variable {E} - -/-- A field extension `E / F` is purely inseparable if and only if for every element `x` of `E`, -its minimal polynomial has separable degree one. -/ -theorem isPurelyInseparable_iff_natSepDegree_eq_one : - IsPurelyInseparable F E ↔ ∀ x : E, (minpoly F x).natSepDegree = 1 := by - obtain ⟨q, _⟩ := ExpChar.exists F - simp_rw [isPurelyInseparable_iff_pow_mem F q, minpoly.natSepDegree_eq_one_iff_pow_mem q] - -theorem IsPurelyInseparable.natSepDegree_eq_one [IsPurelyInseparable F E] (x : E) : - (minpoly F x).natSepDegree = 1 := - (isPurelyInseparable_iff_natSepDegree_eq_one F).1 ‹_› x - -/-- A field extension `E / F` of exponential characteristic `q` is purely inseparable -if and only if for every element `x` of `E`, the minimal polynomial of `x` over `F` is of form -`X ^ (q ^ n) - y` for some natural number `n` and some element `y` of `F`. -/ -theorem isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C (q : ℕ) [hF : ExpChar F q] : - IsPurelyInseparable F E ↔ ∀ x : E, ∃ (n : ℕ) (y : F), minpoly F x = X ^ q ^ n - C y := by - simp_rw [isPurelyInseparable_iff_natSepDegree_eq_one, - minpoly.natSepDegree_eq_one_iff_eq_X_pow_sub_C q] - -theorem IsPurelyInseparable.minpoly_eq_X_pow_sub_C (q : ℕ) [ExpChar F q] [IsPurelyInseparable F E] - (x : E) : ∃ (n : ℕ) (y : F), minpoly F x = X ^ q ^ n - C y := - (isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C F q).1 ‹_› x - -/-- A field extension `E / F` of exponential characteristic `q` is purely inseparable -if and only if for every element `x` of `E`, the minimal polynomial of `x` over `F` is of form -`(X - x) ^ (q ^ n)` for some natural number `n`. -/ -theorem isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow (q : ℕ) [hF : ExpChar F q] : - IsPurelyInseparable F E ↔ - ∀ x : E, ∃ n : ℕ, (minpoly F x).map (algebraMap F E) = (X - C x) ^ q ^ n := by - simp_rw [isPurelyInseparable_iff_natSepDegree_eq_one, - minpoly.natSepDegree_eq_one_iff_eq_X_sub_C_pow q] - -theorem IsPurelyInseparable.minpoly_eq_X_sub_C_pow (q : ℕ) [ExpChar F q] [IsPurelyInseparable F E] - (x : E) : ∃ n : ℕ, (minpoly F x).map (algebraMap F E) = (X - C x) ^ q ^ n := - (isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow F q).1 ‹_› x - -variable (E) - -variable {F E} in -/-- If an extension has finite separable degree one, then it is purely inseparable. -/ -theorem isPurelyInseparable_of_finSepDegree_eq_one - (hdeg : finSepDegree F E = 1) : IsPurelyInseparable F E := by - by_cases H : Algebra.IsAlgebraic F E - · rw [isPurelyInseparable_iff] - refine fun x ↦ ⟨Algebra.IsIntegral.isIntegral x, fun hsep ↦ ?_⟩ - have : Algebra.IsAlgebraic F⟮x⟯ E := Algebra.IsAlgebraic.tower_top (K := F) F⟮x⟯ - have := finSepDegree_mul_finSepDegree_of_isAlgebraic F F⟮x⟯ E - rw [hdeg, mul_eq_one, (finSepDegree_adjoin_simple_eq_finrank_iff F E x - (Algebra.IsAlgebraic.isAlgebraic x)).2 hsep, - IntermediateField.finrank_eq_one_iff] at this - simpa only [this.1] using mem_adjoin_simple_self F x - · rw [← Algebra.transcendental_iff_not_isAlgebraic] at H - simp [finSepDegree_eq_zero_of_transcendental F E] at hdeg - -namespace IsPurelyInseparable - -variable [IsPurelyInseparable F E] (R L : Type*) [CommSemiring R] [Algebra R F] [Algebra R E] - -/-- If `E / F` is purely inseparable, then for any reduced ring `L`, the map `(E →+* L) → (F →+* L)` -induced by `algebraMap F E` is injective. In particular, a purely inseparable field extension -is an epimorphism in the category of fields. -/ -theorem injective_comp_algebraMap [CommRing L] [IsReduced L] : - Function.Injective fun f : E →+* L ↦ f.comp (algebraMap F E) := fun f g heq ↦ by - ext x - let q := ringExpChar F - obtain ⟨n, y, h⟩ := IsPurelyInseparable.pow_mem F q x - replace heq := congr($heq y) - simp_rw [RingHom.comp_apply, h, map_pow] at heq - nontriviality L - haveI := expChar_of_injective_ringHom (f.comp (algebraMap F E)).injective q - exact iterateFrobenius_inj L q n heq - -theorem injective_restrictDomain [CommRing L] [IsReduced L] [Algebra R L] [IsScalarTower R F E] : - Function.Injective (AlgHom.restrictDomain (A := R) F (C := E) (D := L)) := fun _ _ eq ↦ - AlgHom.coe_ringHom_injective <| injective_comp_algebraMap F E L <| congr_arg AlgHom.toRingHom eq - -instance [Field L] [PerfectField L] [Algebra F L] : Nonempty (E →ₐ[F] L) := - nonempty_algHom_of_splits fun x ↦ ⟨IsPurelyInseparable.isIntegral' _ _, - have ⟨q, _⟩ := ExpChar.exists F - PerfectField.splits_of_natSepDegree_eq_one (algebraMap F L) - ((minpoly.natSepDegree_eq_one_iff_eq_X_pow_sub_C q).mpr <| - IsPurelyInseparable.minpoly_eq_X_pow_sub_C F q x)⟩ - -theorem bijective_comp_algebraMap [Field L] [PerfectField L] : - Function.Bijective fun f : E →+* L ↦ f.comp (algebraMap F E) := - ⟨injective_comp_algebraMap F E L, fun g ↦ let _ := g.toAlgebra - ⟨_, (Classical.arbitrary <| E →ₐ[F] L).comp_algebraMap⟩⟩ - -theorem bijective_restrictDomain [Field L] [PerfectField L] [Algebra R L] [IsScalarTower R F E] : - Function.Bijective (AlgHom.restrictDomain (A := R) F (C := E) (D := L)) := - ⟨injective_restrictDomain F E R L, fun g ↦ let _ := g.toAlgebra - let f := Classical.arbitrary (E →ₐ[F] L) - ⟨f.restrictScalars R, AlgHom.coe_ringHom_injective f.comp_algebraMap⟩⟩ - -end IsPurelyInseparable - -/-- If `E / F` is purely inseparable, then for any reduced `F`-algebra `L`, there exists at most one -`F`-algebra homomorphism from `E` to `L`. -/ -instance instSubsingletonAlgHomOfIsPurelyInseparable [IsPurelyInseparable F E] (L : Type w) - [CommRing L] [IsReduced L] [Algebra F L] : Subsingleton (E →ₐ[F] L) where - allEq f g := AlgHom.coe_ringHom_injective <| - IsPurelyInseparable.injective_comp_algebraMap F E L (by simp_rw [AlgHom.comp_algebraMap]) - -instance instUniqueAlgHomOfIsPurelyInseparable [IsPurelyInseparable F E] (L : Type w) - [CommRing L] [IsReduced L] [Algebra F L] [Algebra E L] [IsScalarTower F E L] : - Unique (E →ₐ[F] L) := uniqueOfSubsingleton (IsScalarTower.toAlgHom F E L) - -/-- If `E / F` is purely inseparable, then `Field.Emb F E` has exactly one element. -/ -instance instUniqueEmbOfIsPurelyInseparable [IsPurelyInseparable F E] : - Unique (Emb F E) := instUniqueAlgHomOfIsPurelyInseparable F E _ - -/-- A purely inseparable extension has finite separable degree one. -/ -theorem IsPurelyInseparable.finSepDegree_eq_one [IsPurelyInseparable F E] : - finSepDegree F E = 1 := Nat.card_unique - -/-- A purely inseparable extension has separable degree one. -/ -theorem IsPurelyInseparable.sepDegree_eq_one [IsPurelyInseparable F E] : - sepDegree F E = 1 := by - rw [sepDegree, separableClosure.eq_bot_of_isPurelyInseparable, IntermediateField.rank_bot] - -/-- A purely inseparable extension has inseparable degree equal to degree. -/ -theorem IsPurelyInseparable.insepDegree_eq [IsPurelyInseparable F E] : - insepDegree F E = Module.rank F E := by - rw [insepDegree, separableClosure.eq_bot_of_isPurelyInseparable, rank_bot'] - -/-- A purely inseparable extension has finite inseparable degree equal to degree. -/ -theorem IsPurelyInseparable.finInsepDegree_eq [IsPurelyInseparable F E] : - finInsepDegree F E = finrank F E := congr(Cardinal.toNat $(insepDegree_eq F E)) - -/-- An extension is purely inseparable if and only if it has finite separable degree one. -/ -theorem isPurelyInseparable_iff_finSepDegree_eq_one : - IsPurelyInseparable F E ↔ finSepDegree F E = 1 := - ⟨fun _ ↦ IsPurelyInseparable.finSepDegree_eq_one F E, - fun h ↦ isPurelyInseparable_of_finSepDegree_eq_one h⟩ - -variable {F E} in -/-- An algebraic extension is purely inseparable if and only if all of its finite dimensional -subextensions are purely inseparable. -/ -theorem isPurelyInseparable_iff_fd_isPurelyInseparable [Algebra.IsAlgebraic F E] : - IsPurelyInseparable F E ↔ - ∀ L : IntermediateField F E, FiniteDimensional F L → IsPurelyInseparable F L := by - refine ⟨fun _ _ _ ↦ IsPurelyInseparable.tower_bot F _ E, - fun h ↦ isPurelyInseparable_iff.2 fun x ↦ ?_⟩ - have hx : IsIntegral F x := Algebra.IsIntegral.isIntegral x - refine ⟨hx, fun _ ↦ ?_⟩ - obtain ⟨y, h⟩ := (h _ (adjoin.finiteDimensional hx)).inseparable' _ <| - show Separable (minpoly F (AdjoinSimple.gen F x)) by rwa [minpoly_eq] - exact ⟨y, congr_arg (algebraMap _ E) h⟩ - -/-- A purely inseparable extension is normal. -/ -instance IsPurelyInseparable.normal [IsPurelyInseparable F E] : Normal F E where - toIsAlgebraic := isAlgebraic F E - splits' x := by - obtain ⟨n, h⟩ := IsPurelyInseparable.minpoly_eq_X_sub_C_pow F (ringExpChar F) x - rw [← splits_id_iff_splits, h] - exact splits_pow _ (splits_X_sub_C _) _ - -/-- If `E / F` is algebraic, then `E` is purely inseparable over the -separable closure of `F` in `E`. -/ -@[stacks 030K "$E/E_{sep}$ is purely inseparable."] -instance separableClosure.isPurelyInseparable [Algebra.IsAlgebraic F E] : - IsPurelyInseparable (separableClosure F E) E := isPurelyInseparable_iff.2 fun x ↦ by - set L := separableClosure F E - refine ⟨(IsAlgebraic.tower_top L (Algebra.IsAlgebraic.isAlgebraic (R := F) x)).isIntegral, - fun h ↦ ?_⟩ - haveI := (isSeparable_adjoin_simple_iff_isSeparable L E).2 h - haveI : Algebra.IsSeparable F (restrictScalars F L⟮x⟯) := Algebra.IsSeparable.trans F L L⟮x⟯ - have hx : x ∈ restrictScalars F L⟮x⟯ := mem_adjoin_simple_self _ x - exact ⟨⟨x, mem_separableClosure_iff.2 <| isSeparable_of_mem_isSeparable F E hx⟩, rfl⟩ - -open Cardinal in -theorem Field.Emb.cardinal_separableClosure [Algebra.IsAlgebraic F E] : - #(Field.Emb F <| separableClosure F E) = #(Field.Emb F E) := by - rw [← (embProdEmbOfIsAlgebraic F (separableClosure F E) E).cardinal_eq, - mk_prod, mk_eq_one (Emb _ E), lift_one, mul_one, lift_id] - -/-- An intermediate field of `E / F` contains the separable closure of `F` in `E` -if `E` is purely inseparable over it. -/ -theorem separableClosure_le (L : IntermediateField F E) - [h : IsPurelyInseparable L E] : separableClosure F E ≤ L := fun x hx ↦ by - obtain ⟨y, rfl⟩ := h.inseparable' _ <| - IsSeparable.tower_top L (mem_separableClosure_iff.1 hx) - exact y.2 - -/-- If `E / F` is algebraic, then an intermediate field of `E / F` contains the -separable closure of `F` in `E` if and only if `E` is purely inseparable over it. -/ -theorem separableClosure_le_iff [Algebra.IsAlgebraic F E] (L : IntermediateField F E) : - separableClosure F E ≤ L ↔ IsPurelyInseparable L E := by - refine ⟨fun h ↦ ?_, fun _ ↦ separableClosure_le F E L⟩ - have := separableClosure.isPurelyInseparable F E - letI := (inclusion h).toAlgebra - letI : SMul (separableClosure F E) L := Algebra.toSMul - haveI : IsScalarTower (separableClosure F E) L E := IsScalarTower.of_algebraMap_eq (congrFun rfl) - exact IsPurelyInseparable.tower_top (separableClosure F E) L E - -/-- If an intermediate field of `E / F` is separable over `F`, and `E` is purely inseparable -over it, then it is equal to the separable closure of `F` in `E`. -/ -theorem eq_separableClosure (L : IntermediateField F E) - [Algebra.IsSeparable F L] [IsPurelyInseparable L E] : L = separableClosure F E := - le_antisymm (le_separableClosure F E L) (separableClosure_le F E L) - -open separableClosure in -/-- If `E / F` is algebraic, then an intermediate field of `E / F` is equal to the separable closure -of `F` in `E` if and only if it is separable over `F`, and `E` is purely inseparable -over it. -/ -theorem eq_separableClosure_iff [Algebra.IsAlgebraic F E] (L : IntermediateField F E) : - L = separableClosure F E ↔ Algebra.IsSeparable F L ∧ IsPurelyInseparable L E := - ⟨by rintro rfl; exact ⟨isSeparable F E, isPurelyInseparable F E⟩, - fun ⟨_, _⟩ ↦ eq_separableClosure F E L⟩ - -/-- If `L` is an algebraically closed field containing `E`, such that the map -`(E →+* L) → (F →+* L)` induced by `algebraMap F E` is injective, then `E / F` is -purely inseparable. As a corollary, epimorphisms in the category of fields must be -purely inseparable extensions. -/ -theorem IsPurelyInseparable.of_injective_comp_algebraMap (L : Type w) [Field L] [IsAlgClosed L] - [Nonempty (E →+* L)] (h : Function.Injective fun f : E →+* L ↦ f.comp (algebraMap F E)) : - IsPurelyInseparable F E := by - rw [isPurelyInseparable_iff_finSepDegree_eq_one, finSepDegree, Nat.card_eq_one_iff_unique] - letI := (Classical.arbitrary (E →+* L)).toAlgebra - let j : AlgebraicClosure E →ₐ[E] L := IsAlgClosed.lift - exact ⟨⟨fun f g ↦ DFunLike.ext' <| j.injective.comp_left (congr_arg (⇑) <| - @h (j.toRingHom.comp f) (j.toRingHom.comp g) (by ext; simp))⟩, inferInstance⟩ - -end IsPurelyInseparable - -namespace IntermediateField - -instance isPurelyInseparable_bot : IsPurelyInseparable F (⊥ : IntermediateField F E) := - (botEquiv F E).symm.isPurelyInseparable - -/-- `F⟮x⟯ / F` is a purely inseparable extension if and only if the minimal polynomial of `x` -has separable degree one. -/ -theorem isPurelyInseparable_adjoin_simple_iff_natSepDegree_eq_one {x : E} : - IsPurelyInseparable F F⟮x⟯ ↔ (minpoly F x).natSepDegree = 1 := by - rw [← le_perfectClosure_iff, adjoin_simple_le_iff, mem_perfectClosure_iff_natSepDegree_eq_one] - -/-- If `F` is of exponential characteristic `q`, then `F⟮x⟯ / F` is a purely inseparable extension -if and only if `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. -/ -theorem isPurelyInseparable_adjoin_simple_iff_pow_mem (q : ℕ) [hF : ExpChar F q] {x : E} : - IsPurelyInseparable F F⟮x⟯ ↔ ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by - rw [← le_perfectClosure_iff, adjoin_simple_le_iff, mem_perfectClosure_iff_pow_mem q] - -/-- If `F` is of exponential characteristic `q`, then `F(S) / F` is a purely inseparable extension -if and only if for any `x ∈ S`, `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. -/ -theorem isPurelyInseparable_adjoin_iff_pow_mem (q : ℕ) [hF : ExpChar F q] {S : Set E} : - IsPurelyInseparable F (adjoin F S) ↔ ∀ x ∈ S, ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by - simp_rw [← le_perfectClosure_iff, adjoin_le_iff, ← mem_perfectClosure_iff_pow_mem q, - Set.subset_def, SetLike.mem_coe] - -/-- A compositum of two purely inseparable extensions is purely inseparable. -/ -instance isPurelyInseparable_sup (L1 L2 : IntermediateField F E) - [h1 : IsPurelyInseparable F L1] [h2 : IsPurelyInseparable F L2] : - IsPurelyInseparable F (L1 ⊔ L2 : IntermediateField F E) := by - rw [← le_perfectClosure_iff] at h1 h2 ⊢ - exact sup_le h1 h2 - -/-- A compositum of purely inseparable extensions is purely inseparable. -/ -instance isPurelyInseparable_iSup {ι : Sort*} {t : ι → IntermediateField F E} - [h : ∀ i, IsPurelyInseparable F (t i)] : - IsPurelyInseparable F (⨆ i, t i : IntermediateField F E) := by - simp_rw [← le_perfectClosure_iff] at h ⊢ - exact iSup_le h - -/-- If `F` is a field of exponential characteristic `q`, `F(S) / F` is separable, then -`F(S) = F(S ^ (q ^ n))` for any natural number `n`. -/ -theorem adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable (S : Set E) - [Algebra.IsSeparable F (adjoin F S)] (q : ℕ) [ExpChar F q] (n : ℕ) : - adjoin F S = adjoin F ((· ^ q ^ n) '' S) := by - set L := adjoin F S - set M := adjoin F ((· ^ q ^ n) '' S) - have hi : M ≤ L := by - rw [adjoin_le_iff] - rintro _ ⟨y, hy, rfl⟩ - exact pow_mem (subset_adjoin F S hy) _ - letI := (inclusion hi).toAlgebra - haveI : Algebra.IsSeparable M (extendScalars hi) := - Algebra.isSeparable_tower_top_of_isSeparable F M L - haveI : IsPurelyInseparable M (extendScalars hi) := by - haveI := expChar_of_injective_algebraMap (algebraMap F M).injective q - rw [extendScalars_adjoin hi, isPurelyInseparable_adjoin_iff_pow_mem M _ q] - exact fun x hx ↦ ⟨n, ⟨x ^ q ^ n, subset_adjoin F _ ⟨x, hx, rfl⟩⟩, rfl⟩ - simpa only [extendScalars_restrictScalars, restrictScalars_bot_eq_self] using congr_arg - (restrictScalars F) (extendScalars hi).eq_bot_of_isPurelyInseparable_of_isSeparable - -/-- If `E / F` is a separable field extension of exponential characteristic `q`, then -`F(S) = F(S ^ (q ^ n))` for any subset `S` of `E` and any natural number `n`. -/ -theorem adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable' [Algebra.IsSeparable F E] (S : Set E) - (q : ℕ) [ExpChar F q] (n : ℕ) : adjoin F S = adjoin F ((· ^ q ^ n) '' S) := - haveI := Algebra.isSeparable_tower_bot_of_isSeparable F (adjoin F S) E - adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E S q n - --- TODO: prove the converse when `F(S) / F` is finite -/-- If `F` is a field of exponential characteristic `q`, `F(S) / F` is separable, then -`F(S) = F(S ^ q)`. -/ -theorem adjoin_eq_adjoin_pow_expChar_of_isSeparable (S : Set E) [Algebra.IsSeparable F (adjoin F S)] - (q : ℕ) [ExpChar F q] : adjoin F S = adjoin F ((· ^ q) '' S) := - pow_one q ▸ adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E S q 1 - -/-- If `E / F` is a separable field extension of exponential characteristic `q`, then -`F(S) = F(S ^ q)` for any subset `S` of `E`. -/ -theorem adjoin_eq_adjoin_pow_expChar_of_isSeparable' [Algebra.IsSeparable F E] (S : Set E) - (q : ℕ) [ExpChar F q] : adjoin F S = adjoin F ((· ^ q) '' S) := - pow_one q ▸ adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable' F E S q 1 - --- Special cases for simple adjoin - -/-- If `F` is a field of exponential characteristic `q`, `a : E` is separable over `F`, then -`F⟮a⟯ = F⟮a ^ q ^ n⟯` for any natural number `n`. -/ -theorem adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable {a : E} (ha : IsSeparable F a) - (q : ℕ) [ExpChar F q] (n : ℕ) : F⟮a⟯ = F⟮a ^ q ^ n⟯ := by - haveI := (isSeparable_adjoin_simple_iff_isSeparable F E).mpr ha - simpa using adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E {a} q n - -/-- If `E / F` is a separable field extension of exponential characteristic `q`, then -`F⟮a⟯ = F⟮a ^ q ^ n⟯` for any subset `a : E` and any natural number `n`. -/ -theorem adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable' [Algebra.IsSeparable F E] (a : E) - (q : ℕ) [ExpChar F q] (n : ℕ) : F⟮a⟯ = F⟮a ^ q ^ n⟯ := by - haveI := Algebra.isSeparable_tower_bot_of_isSeparable F F⟮a⟯ E - simpa using adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E {a} q n - -/-- If `F` is a field of exponential characteristic `q`, `a : E` is separable over `F`, then -`F⟮a⟯ = F⟮a ^ q⟯`. -/ -theorem adjoin_simple_eq_adjoin_pow_expChar_of_isSeparable {a : E} (ha : IsSeparable F a) - (q : ℕ) [ExpChar F q] : F⟮a⟯ = F⟮a ^ q⟯ := - pow_one q ▸ adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable F E ha q 1 - -/-- If `E / F` is a separable field extension of exponential characteristic `q`, then -`F⟮a⟯ = F⟮a ^ q⟯` for any `a : E`. -/ -theorem adjoin_simple_eq_adjoin_pow_expChar_of_isSeparable' [Algebra.IsSeparable F E] (a : E) - (q : ℕ) [ExpChar F q] : F⟮a⟯ = F⟮a ^ q⟯ := - pow_one q ▸ adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable' F E a q 1 - -end IntermediateField - -section - -variable (q n : ℕ) [hF : ExpChar F q] {ι : Type*} {v : ι → E} {F E} - -/-- If `E / F` is a separable extension of exponential characteristic `q`, if `{ u_i }` is a family -of elements of `E` which `F`-linearly spans `E`, then `{ u_i ^ (q ^ n) }` also `F`-linearly spans -`E` for any natural number `n`. -/ -theorem Field.span_map_pow_expChar_pow_eq_top_of_isSeparable [Algebra.IsSeparable F E] - (h : Submodule.span F (Set.range v) = ⊤) : - Submodule.span F (Set.range (v · ^ q ^ n)) = ⊤ := by - erw [← Algebra.top_toSubmodule, ← top_toSubalgebra, ← adjoin_univ, - adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable' F E _ q n, - adjoin_algebraic_toSubalgebra fun x _ ↦ Algebra.IsAlgebraic.isAlgebraic x, - Set.image_univ, Algebra.adjoin_eq_span, (powMonoidHom _).mrange.closure_eq] - refine (Submodule.span_mono <| Set.range_comp_subset_range _ _).antisymm (Submodule.span_le.2 ?_) - rw [Set.range_comp, ← Set.image_univ] - haveI := expChar_of_injective_algebraMap (algebraMap F E).injective q - apply h ▸ Submodule.image_span_subset_span (LinearMap.iterateFrobenius F E q n) _ - -/-- If `E / F` is a finite separable extension of exponential characteristic `q`, if `{ u_i }` is a -family of elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` is also -`F`-linearly independent for any natural number `n`. A special case of -`LinearIndependent.map_pow_expChar_pow_of_isSeparable` -and is an intermediate result used to prove it. -/ -private theorem LinearIndependent.map_pow_expChar_pow_of_fd_isSeparable - [FiniteDimensional F E] [Algebra.IsSeparable F E] - (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by - have h' := h.coe_range - let ι' := h'.extend (Set.range v).subset_univ - let b : Basis ι' F E := Basis.extend h' - letI : Fintype ι' := FiniteDimensional.fintypeBasisIndex b - have H := linearIndependent_of_top_le_span_of_card_eq_finrank - (span_map_pow_expChar_pow_eq_top_of_isSeparable q n b.span_eq).ge - (finrank_eq_card_basis b).symm - let f (i : ι) : ι' := ⟨v i, h'.subset_extend _ ⟨i, rfl⟩⟩ - convert H.comp f fun _ _ heq ↦ h.injective (by simpa only [f, Subtype.mk.injEq] using heq) - simp_rw [Function.comp_apply, b] - rw [Basis.extend_apply_self] - -/-- If `E / F` is a separable extension of exponential characteristic `q`, if `{ u_i }` is a -family of elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` is also -`F`-linearly independent for any natural number `n`. -/ -theorem LinearIndependent.map_pow_expChar_pow_of_isSeparable [Algebra.IsSeparable F E] - (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by - classical - have halg := Algebra.IsSeparable.isAlgebraic F E - rw [linearIndependent_iff_finset_linearIndependent] at h ⊢ - intro s - let E' := adjoin F (s.image v : Set E) - haveI : FiniteDimensional F E' := finiteDimensional_adjoin - fun x _ ↦ Algebra.IsIntegral.isIntegral x - haveI : Algebra.IsSeparable F E' := Algebra.isSeparable_tower_bot_of_isSeparable F E' E - let v' (i : s) : E' := ⟨v i.1, subset_adjoin F _ (Finset.mem_image.2 ⟨i.1, i.2, rfl⟩)⟩ - have h' : LinearIndependent F v' := (h s).of_comp E'.val.toLinearMap - exact (h'.map_pow_expChar_pow_of_fd_isSeparable q n).map' - E'.val.toLinearMap (LinearMap.ker_eq_bot_of_injective E'.val.injective) - -/-- If `E / F` is a field extension of exponential characteristic `q`, if `{ u_i }` is a -family of separable elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` -is also `F`-linearly independent for any natural number `n`. -/ -theorem LinearIndependent.map_pow_expChar_pow_of_isSeparable' - (hsep : ∀ i : ι, IsSeparable F (v i)) - (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by - let E' := adjoin F (Set.range v) - haveI : Algebra.IsSeparable F E' := (isSeparable_adjoin_iff_isSeparable F _).2 <| by - rintro _ ⟨y, rfl⟩; exact hsep y - let v' (i : ι) : E' := ⟨v i, subset_adjoin F _ ⟨i, rfl⟩⟩ - have h' : LinearIndependent F v' := h.of_comp E'.val.toLinearMap - exact (h'.map_pow_expChar_pow_of_isSeparable q n).map' - E'.val.toLinearMap (LinearMap.ker_eq_bot_of_injective E'.val.injective) - -/-- If `E / F` is a separable extension of exponential characteristic `q`, if `{ u_i }` is an -`F`-basis of `E`, then `{ u_i ^ (q ^ n) }` is also an `F`-basis of `E` -for any natural number `n`. -/ -def Basis.mapPowExpCharPowOfIsSeparable [Algebra.IsSeparable F E] - (b : Basis ι F E) : Basis ι F E := - Basis.mk (b.linearIndependent.map_pow_expChar_pow_of_isSeparable q n) - (span_map_pow_expChar_pow_eq_top_of_isSeparable q n b.span_eq).ge - -end - -/-- If `E` is an algebraic closure of `F`, then `F` is separably closed if and only if `E / F` is -purely inseparable. -/ -theorem isSepClosed_iff_isPurelyInseparable_algebraicClosure [IsAlgClosure F E] : - IsSepClosed F ↔ IsPurelyInseparable F E := - ⟨fun _ ↦ inferInstance, fun H ↦ by - haveI := IsAlgClosure.isAlgClosed F (K := E) - rwa [← separableClosure.eq_bot_iff, IsSepClosed.separableClosure_eq_bot_iff] at H⟩ - -variable {F E} in -/-- If `E / F` is an algebraic extension, `F` is separably closed, -then `E` is also separably closed. -/ -theorem Algebra.IsAlgebraic.isSepClosed [Algebra.IsAlgebraic F E] - [IsSepClosed F] : IsSepClosed E := - have : Algebra.IsAlgebraic F (AlgebraicClosure E) := Algebra.IsAlgebraic.trans (L := E) - (isSepClosed_iff_isPurelyInseparable_algebraicClosure E _).mpr - (IsPurelyInseparable.tower_top F E <| AlgebraicClosure E) - -theorem perfectField_of_perfectClosure_eq_bot [h : PerfectField E] (eq : perfectClosure F E = ⊥) : - PerfectField F := by - let p := ringExpChar F - haveI := expChar_of_injective_algebraMap (algebraMap F E).injective p - haveI := PerfectRing.ofSurjective F p fun x ↦ by - obtain ⟨y, h⟩ := surjective_frobenius E p (algebraMap F E x) - have : y ∈ perfectClosure F E := ⟨1, x, by rw [← h, pow_one, frobenius_def, ringExpChar.eq F p]⟩ - obtain ⟨z, rfl⟩ := eq ▸ this - exact ⟨z, (algebraMap F E).injective (by erw [RingHom.map_frobenius, h])⟩ - exact PerfectRing.toPerfectField F p - -/-- If `E / F` is a separable extension, `E` is perfect, then `F` is also prefect. -/ -theorem perfectField_of_isSeparable_of_perfectField_top [Algebra.IsSeparable F E] [PerfectField E] : - PerfectField F := - perfectField_of_perfectClosure_eq_bot F E (perfectClosure.eq_bot_of_isSeparable F E) - -/-- If `E` is an algebraic closure of `F`, then `F` is perfect if and only if `E / F` is -separable. -/ -theorem perfectField_iff_isSeparable_algebraicClosure [IsAlgClosure F E] : - PerfectField F ↔ Algebra.IsSeparable F E := - ⟨fun _ ↦ IsSepClosure.separable, fun _ ↦ haveI : IsAlgClosed E := IsAlgClosure.isAlgClosed F - perfectField_of_isSeparable_of_perfectField_top F E⟩ - -namespace Field - -/-- If `E / F` is algebraic, then the `Field.finSepDegree F E` is equal to `Field.sepDegree F E` -as a natural number. This means that the cardinality of `Field.Emb F E` and the degree of -`(separableClosure F E) / F` are both finite or infinite, and when they are finite, they -coincide. -/ -@[stacks 09HJ "`sepDegree` is defined as the right hand side of 09HJ"] -theorem finSepDegree_eq [Algebra.IsAlgebraic F E] : - finSepDegree F E = Cardinal.toNat (sepDegree F E) := by - have : Algebra.IsAlgebraic (separableClosure F E) E := Algebra.IsAlgebraic.tower_top (K := F) _ - have h := finSepDegree_mul_finSepDegree_of_isAlgebraic F (separableClosure F E) E |>.symm - haveI := separableClosure.isSeparable F E - haveI := separableClosure.isPurelyInseparable F E - rwa [finSepDegree_eq_finrank_of_isSeparable F (separableClosure F E), - IsPurelyInseparable.finSepDegree_eq_one (separableClosure F E) E, mul_one] at h - -/-- The finite separable degree multiply by the finite inseparable degree is equal -to the (finite) field extension degree. -/ -theorem finSepDegree_mul_finInsepDegree : finSepDegree F E * finInsepDegree F E = finrank F E := by - by_cases halg : Algebra.IsAlgebraic F E - · have := congr_arg Cardinal.toNat (sepDegree_mul_insepDegree F E) - rwa [Cardinal.toNat_mul, ← finSepDegree_eq F E] at this - rw [finInsepDegree, finrank_of_infinite_dimensional (K := F) (V := E) fun _ ↦ - halg (Algebra.IsAlgebraic.of_finite F E), - finrank_of_infinite_dimensional (K := separableClosure F E) (V := E) fun _ ↦ - halg ((separableClosure.isAlgebraic F E).trans), - mul_zero] - -end Field - -namespace separableClosure - -variable [Algebra E K] [IsScalarTower F E K] {F E} - -/-- If `K / E / F` is a field extension tower, such that `E / F` is algebraic and `K / E` -is separable, then `E` adjoin `separableClosure F K` is equal to `K`. It is a special case of -`separableClosure.adjoin_eq_of_isAlgebraic`, and is an intermediate result used to prove it. -/ -lemma adjoin_eq_of_isAlgebraic_of_isSeparable [Algebra.IsAlgebraic F E] - [Algebra.IsSeparable E K] : adjoin E (separableClosure F K : Set K) = ⊤ := - top_unique fun x _ ↦ by - set S := separableClosure F K - set L := adjoin E (S : Set K) - have := Algebra.isSeparable_tower_top_of_isSeparable E L K - let i : S →+* L := Subsemiring.inclusion fun x hx ↦ subset_adjoin E (S : Set K) hx - let _ : Algebra S L := i.toAlgebra - let _ : SMul S L := Algebra.toSMul - have : IsScalarTower S L K := IsScalarTower.of_algebraMap_eq (congrFun rfl) - have : Algebra.IsAlgebraic F K := Algebra.IsAlgebraic.trans (L := E) - have : IsPurelyInseparable S K := separableClosure.isPurelyInseparable F K - have := IsPurelyInseparable.tower_top S L K - obtain ⟨y, rfl⟩ := IsPurelyInseparable.surjective_algebraMap_of_isSeparable L K x - exact y.2 - -/-- If `K / E / F` is a field extension tower, such that `E / F` is algebraic, then -`E` adjoin `separableClosure F K` is equal to `separableClosure E K`. -/ -theorem adjoin_eq_of_isAlgebraic [Algebra.IsAlgebraic F E] : - adjoin E (separableClosure F K) = separableClosure E K := by - set S := separableClosure E K - have h := congr_arg lift (adjoin_eq_of_isAlgebraic_of_isSeparable (F := F) S) - rw [lift_top, lift_adjoin] at h - haveI : IsScalarTower F S K := IsScalarTower.of_algebraMap_eq (congrFun rfl) - rw [← h, ← map_eq_of_separableClosure_eq_bot F (separableClosure_eq_bot E K)] - simp only [S, coe_map, IsScalarTower.coe_toAlgHom', IntermediateField.algebraMap_apply] - -end separableClosure - -section TowerLaw - -variable [Algebra E K] [IsScalarTower F E K] - -variable {F K} in -/-- If `K / E / F` is a field extension tower such that `E / F` is purely inseparable, -if `{ u_i }` is a family of separable elements of `K` which is `F`-linearly independent, -then it is also `E`-linearly independent. -/ -theorem LinearIndependent.map_of_isPurelyInseparable_of_isSeparable [IsPurelyInseparable F E] - {ι : Type*} {v : ι → K} (hsep : ∀ i : ι, IsSeparable F (v i)) - (h : LinearIndependent F v) : LinearIndependent E v := by - obtain ⟨q, _⟩ := ExpChar.exists F - haveI := expChar_of_injective_algebraMap (algebraMap F K).injective q - refine linearIndependent_iff.mpr fun l hl ↦ Finsupp.ext fun i ↦ ?_ - choose f hf using fun i ↦ (isPurelyInseparable_iff_pow_mem F q).1 ‹_› (l i) - let n := l.support.sup f - have := (expChar_pow_pos F q n).ne' - replace hf (i : ι) : l i ^ q ^ n ∈ (algebraMap F E).range := by - by_cases hs : i ∈ l.support - · convert pow_mem (hf i) (q ^ (n - f i)) using 1 - rw [← pow_mul, ← pow_add, Nat.add_sub_of_le (Finset.le_sup hs)] - exact ⟨0, by rw [map_zero, Finsupp.not_mem_support_iff.1 hs, zero_pow this]⟩ - choose lF hlF using hf - let lF₀ := Finsupp.onFinset l.support lF fun i ↦ by - contrapose! - refine fun hs ↦ (injective_iff_map_eq_zero _).mp (algebraMap F E).injective _ ?_ - rw [hlF, Finsupp.not_mem_support_iff.1 hs, zero_pow this] - replace h := linearIndependent_iff.1 (h.map_pow_expChar_pow_of_isSeparable' q n hsep) lF₀ <| by - replace hl := congr($hl ^ q ^ n) - rw [linearCombination_apply, Finsupp.sum, sum_pow_char_pow, zero_pow this] at hl - rw [← hl, linearCombination_apply, onFinset_sum _ (fun _ ↦ by exact zero_smul _ _)] - refine Finset.sum_congr rfl fun i _ ↦ ?_ - simp_rw [Algebra.smul_def, mul_pow, IsScalarTower.algebraMap_apply F E K, hlF, map_pow] - refine pow_eq_zero ((hlF _).symm.trans ?_) - convert map_zero (algebraMap F E) - exact congr($h i) - -namespace Field - -/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable and `K / E` -is separable, then the separable degree of `K / F` is equal to the degree of `K / E`. -It is a special case of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`, and is an -intermediate result used to prove it. -/ -lemma sepDegree_eq_of_isPurelyInseparable_of_isSeparable - [IsPurelyInseparable F E] [Algebra.IsSeparable E K] : sepDegree F K = Module.rank E K := by - let S := separableClosure F K - have h := S.adjoin_rank_le_of_isAlgebraic_right E - rw [separableClosure.adjoin_eq_of_isAlgebraic_of_isSeparable K, rank_top'] at h - obtain ⟨ι, ⟨b⟩⟩ := Basis.exists_basis F S - exact h.antisymm' (b.mk_eq_rank'' ▸ (b.linearIndependent.map' S.val.toLinearMap - (LinearMap.ker_eq_bot_of_injective S.val.injective) - |>.map_of_isPurelyInseparable_of_isSeparable E (fun i ↦ - by simpa only [IsSeparable, minpoly_eq] using Algebra.IsSeparable.isSeparable F (b i)) - |>.cardinal_le_rank)) - -/-- If `K / E / F` is a field extension tower, such that `E / F` is separable, -then $[E:F] [K:E]_s = [K:F]_s$. -It is a special case of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`, and is an -intermediate result used to prove it. -/ -lemma lift_rank_mul_lift_sepDegree_of_isSeparable [Algebra.IsSeparable F E] : - Cardinal.lift.{w} (Module.rank F E) * Cardinal.lift.{v} (sepDegree E K) = - Cardinal.lift.{v} (sepDegree F K) := by - rw [sepDegree, sepDegree, separableClosure.eq_restrictScalars_of_isSeparable F E K] - exact lift_rank_mul_lift_rank F E (separableClosure E K) - -/-- The same-universe version of `Field.lift_rank_mul_lift_sepDegree_of_isSeparable`. -/ -lemma rank_mul_sepDegree_of_isSeparable (K : Type v) [Field K] [Algebra F K] - [Algebra E K] [IsScalarTower F E K] [Algebra.IsSeparable F E] : - Module.rank F E * sepDegree E K = sepDegree F K := by - simpa only [Cardinal.lift_id] using lift_rank_mul_lift_sepDegree_of_isSeparable F E K - -/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, -then $[K:F]_s = [K:E]_s$. -It is a special case of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`, and is an -intermediate result used to prove it. -/ -lemma sepDegree_eq_of_isPurelyInseparable [IsPurelyInseparable F E] : - sepDegree F K = sepDegree E K := by - convert sepDegree_eq_of_isPurelyInseparable_of_isSeparable F E (separableClosure E K) - haveI : IsScalarTower F (separableClosure E K) K := IsScalarTower.of_algebraMap_eq (congrFun rfl) - rw [sepDegree, ← separableClosure.map_eq_of_separableClosure_eq_bot F - (separableClosure.separableClosure_eq_bot E K)] - exact (separableClosure F (separableClosure E K)).equivMap - (IsScalarTower.toAlgHom F (separableClosure E K) K) |>.symm.toLinearEquiv.rank_eq - -/-- If `K / E / F` is a field extension tower, such that `E / F` is algebraic, then their -separable degrees satisfy the tower law: $[E:F]_s [K:E]_s = [K:F]_s$. -/ -theorem lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic [Algebra.IsAlgebraic F E] : - Cardinal.lift.{w} (sepDegree F E) * Cardinal.lift.{v} (sepDegree E K) = - Cardinal.lift.{v} (sepDegree F K) := by - have h := lift_rank_mul_lift_sepDegree_of_isSeparable F (separableClosure F E) K - haveI := separableClosure.isPurelyInseparable F E - rwa [sepDegree_eq_of_isPurelyInseparable (separableClosure F E) E K] at h - -/-- The same-universe version of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`. -/ -@[stacks 09HK "Part 1"] -theorem sepDegree_mul_sepDegree_of_isAlgebraic (K : Type v) [Field K] [Algebra F K] - [Algebra E K] [IsScalarTower F E K] [Algebra.IsAlgebraic F E] : - sepDegree F E * sepDegree E K = sepDegree F K := by - simpa only [Cardinal.lift_id] using lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic F E K - -end Field - -variable {F K} in -/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then -for any subset `S` of `K` such that `F(S) / F` is algebraic, the `E(S) / E` and `F(S) / F` have -the same separable degree. -/ -theorem IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable - (S : Set K) [Algebra.IsAlgebraic F (adjoin F S)] [IsPurelyInseparable F E] : - sepDegree E (adjoin E S) = sepDegree F (adjoin F S) := by - set M := adjoin F S - set L := adjoin E S - let E' := (IsScalarTower.toAlgHom F E K).fieldRange - let j : E ≃ₐ[F] E' := AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K) - have hi : M ≤ L.restrictScalars F := by - rw [restrictScalars_adjoin_of_algEquiv (E := K) j rfl, restrictScalars_adjoin] - exact adjoin.mono _ _ _ Set.subset_union_right - let i : M →+* L := Subsemiring.inclusion hi - letI : Algebra M L := i.toAlgebra - letI : SMul M L := Algebra.toSMul - haveI : IsScalarTower F M L := IsScalarTower.of_algebraMap_eq (congrFun rfl) - haveI : IsPurelyInseparable M L := by - change IsPurelyInseparable M (extendScalars hi) - obtain ⟨q, _⟩ := ExpChar.exists F - have : extendScalars hi = adjoin M (E' : Set K) := restrictScalars_injective F <| by - conv_lhs => rw [extendScalars_restrictScalars, restrictScalars_adjoin_of_algEquiv - (E := K) j rfl, ← adjoin_self F E', adjoin_adjoin_comm] - rw [this, isPurelyInseparable_adjoin_iff_pow_mem _ _ q] - rintro x ⟨y, hy⟩ - obtain ⟨n, z, hz⟩ := IsPurelyInseparable.pow_mem F q y - refine ⟨n, algebraMap F M z, ?_⟩ - rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply F E K, hz, ← hy, map_pow, - AlgHom.toRingHom_eq_coe, IsScalarTower.coe_toAlgHom] - have h := lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic F E L - rw [IsPurelyInseparable.sepDegree_eq_one F E, Cardinal.lift_one, one_mul] at h - rw [Cardinal.lift_injective h, ← sepDegree_mul_sepDegree_of_isAlgebraic F M L, - IsPurelyInseparable.sepDegree_eq_one M L, mul_one] - -variable {F K} in -/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then -for any intermediate field `S` of `K / F` such that `S / F` is algebraic, the `E(S) / E` and -`S / F` have the same separable degree. -/ -theorem IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable' - (S : IntermediateField F K) [Algebra.IsAlgebraic F S] [IsPurelyInseparable F E] : - sepDegree E (adjoin E (S : Set K)) = sepDegree F S := by - have : Algebra.IsAlgebraic F (adjoin F (S : Set K)) := by rwa [adjoin_self] - have := sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable (F := F) E (S : Set K) - rwa [adjoin_self] at this - -variable {F K} in -/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then -for any element `x` of `K` separable over `F`, it has the same minimal polynomials over `F` and -over `E`. -/ -theorem minpoly.map_eq_of_isSeparable_of_isPurelyInseparable (x : K) - (hsep : IsSeparable F x) [IsPurelyInseparable F E] : - (minpoly F x).map (algebraMap F E) = minpoly E x := by - have hi := IsSeparable.isIntegral hsep - have hi' : IsIntegral E x := IsIntegral.tower_top hi - refine eq_of_monic_of_dvd_of_natDegree_le (monic hi') ((monic hi).map (algebraMap F E)) - (dvd_map_of_isScalarTower F E x) (le_of_eq ?_) - have hsep' := IsSeparable.tower_top E hsep - haveI := (isSeparable_adjoin_simple_iff_isSeparable _ _).2 hsep - haveI := (isSeparable_adjoin_simple_iff_isSeparable _ _).2 hsep' - have := Algebra.IsSeparable.isAlgebraic F F⟮x⟯ - have := Algebra.IsSeparable.isAlgebraic E E⟮x⟯ - rw [Polynomial.natDegree_map, ← adjoin.finrank hi, ← adjoin.finrank hi', - ← finSepDegree_eq_finrank_of_isSeparable F _, ← finSepDegree_eq_finrank_of_isSeparable E _, - finSepDegree_eq, finSepDegree_eq, - sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable (F := F) E] - -variable {F} in -/-- If `E / F` is a purely inseparable field extension, `f` is a separable irreducible polynomial -over `F`, then it is also irreducible over `E`. -/ -theorem Polynomial.Separable.map_irreducible_of_isPurelyInseparable {f : F[X]} (hsep : f.Separable) - (hirr : Irreducible f) [IsPurelyInseparable F E] : Irreducible (f.map (algebraMap F E)) := by - let K := AlgebraicClosure E - obtain ⟨x, hx⟩ := IsAlgClosed.exists_aeval_eq_zero K f - (natDegree_pos_iff_degree_pos.1 hirr.natDegree_pos).ne' - have ha : Associated f (minpoly F x) := by - have := isUnit_C.2 (leadingCoeff_ne_zero.2 hirr.ne_zero).isUnit.inv - exact ⟨this.unit, by rw [IsUnit.unit_spec, minpoly.eq_of_irreducible hirr hx]⟩ - have ha' : Associated (f.map (algebraMap F E)) ((minpoly F x).map (algebraMap F E)) := - ha.map (mapRingHom (algebraMap F E)).toMonoidHom - have heq := minpoly.map_eq_of_isSeparable_of_isPurelyInseparable E x (ha.separable hsep) - rw [ha'.irreducible_iff, heq] - exact minpoly.irreducible (Algebra.IsIntegral.isIntegral x) - -end TowerLaw diff --git a/Mathlib/FieldTheory/PurelyInseparable/Basic.lean b/Mathlib/FieldTheory/PurelyInseparable/Basic.lean new file mode 100644 index 00000000000000..5db00ddb901d62 --- /dev/null +++ b/Mathlib/FieldTheory/PurelyInseparable/Basic.lean @@ -0,0 +1,593 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.FieldTheory.SeparableClosure + +/-! + +# Basic results about purely inseparable extensions + +This file contains basic definitions and results about purely inseparable extensions. + +## Main definitions + +- `IsPurelyInseparable`: typeclass for purely inseparable field extensions: an algebraic extension + `E / F` is purely inseparable if and only if the minimal polynomial of every element of `E ∖ F` + is not separable. + +## Main results + +- `IsPurelyInseparable.surjective_algebraMap_of_isSeparable`, + `IsPurelyInseparable.bijective_algebraMap_of_isSeparable`, + `IntermediateField.eq_bot_of_isPurelyInseparable_of_isSeparable`: + if `E / F` is both purely inseparable and separable, then `algebraMap F E` is surjective + (hence bijective). In particular, if an intermediate field of `E / F` is both purely inseparable + and separable, then it is equal to `F`. + +- `isPurelyInseparable_iff_pow_mem`: a field extension `E / F` of exponential characteristic `q` is + purely inseparable if and only if for every element `x` of `E`, there exists a natural number `n` + such that `x ^ (q ^ n)` is contained in `F`. + +- `IsPurelyInseparable.trans`: if `E / F` and `K / E` are both purely inseparable extensions, then + `K / F` is also purely inseparable. + +- `isPurelyInseparable_iff_natSepDegree_eq_one`: `E / F` is purely inseparable if and only if for + every element `x` of `E`, its minimal polynomial has separable degree one. + +- `isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C`: a field extension `E / F` of exponential + characteristic `q` is purely inseparable if and only if for every element `x` of `E`, the minimal + polynomial of `x` over `F` is of form `X ^ (q ^ n) - y` for some natural number `n` and some + element `y` of `F`. + +- `isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow`: a field extension `E / F` of exponential + characteristic `q` is purely inseparable if and only if for every element `x` of `E`, the minimal + polynomial of `x` over `F` is of form `(X - x) ^ (q ^ n)` for some natural number `n`. + +- `isPurelyInseparable_iff_finSepDegree_eq_one`: an extension is purely inseparable + if and only if it has finite separable degree (`Field.finSepDegree`) one. + +- `IsPurelyInseparable.normal`: a purely inseparable extension is normal. + +- `separableClosure.isPurelyInseparable`: if `E / F` is algebraic, then `E` is purely inseparable + over the separable closure of `F` in `E`. + +- `separableClosure_le_iff`: if `E / F` is algebraic, then an intermediate field of `E / F` contains + the separable closure of `F` in `E` if and only if `E` is purely inseparable over it. + +- `eq_separableClosure_iff`: if `E / F` is algebraic, then an intermediate field of `E / F` is equal + to the separable closure of `F` in `E` if and only if it is separable over `F`, and `E` + is purely inseparable over it. + +- `IsPurelyInseparable.injective_comp_algebraMap`: if `E / F` is purely inseparable, then for any + reduced ring `L`, the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is injective. + In particular, a purely inseparable field extension is an epimorphism in the category of fields. + +- `IsPurelyInseparable.of_injective_comp_algebraMap`: if `L` is an algebraically closed field + containing `E`, such that the map `(E →+* L) → (F →+* L)` induced by `algebraMap F E` is + injective, then `E / F` is purely inseparable. As a corollary, epimorphisms in the category of + fields must be purely inseparable extensions. + +- `Field.finSepDegree_eq`: if `E / F` is algebraic, then the `Field.finSepDegree F E` is equal to + `Field.sepDegree F E` as a natural number. This means that the cardinality of `Field.Emb F E` + and the degree of `(separableClosure F E) / F` are both finite or infinite, and when they are + finite, they coincide. + +- `Field.finSepDegree_mul_finInsepDegree`: the finite separable degree multiply by the finite + inseparable degree is equal to the (finite) field extension degree. + +## Tags + +separable degree, degree, separable closure, purely inseparable + +-/ + +open Module Polynomial IntermediateField Field + +noncomputable section + +universe u v w + +section General + +variable (F E : Type*) [CommRing F] [Ring E] [Algebra F E] +variable (K : Type*) [Ring K] [Algebra F K] + +/-- Typeclass for purely inseparable field extensions: an algebraic extension `E / F` is purely +inseparable if and only if the minimal polynomial of every element of `E ∖ F` is not separable. + +We define this for general (commutative) rings and only assume `F` and `E` are fields +if this is needed for a proof. -/ +class IsPurelyInseparable : Prop where + isIntegral : Algebra.IsIntegral F E + inseparable' (x : E) : IsSeparable F x → x ∈ (algebraMap F E).range + +attribute [instance] IsPurelyInseparable.isIntegral + +variable {E} in +theorem IsPurelyInseparable.isIntegral' [IsPurelyInseparable F E] (x : E) : IsIntegral F x := + Algebra.IsIntegral.isIntegral _ + +theorem IsPurelyInseparable.isAlgebraic [Nontrivial F] [IsPurelyInseparable F E] : + Algebra.IsAlgebraic F E := inferInstance + +variable {E} + +theorem IsPurelyInseparable.inseparable [IsPurelyInseparable F E] : + ∀ x : E, IsSeparable F x → x ∈ (algebraMap F E).range := + IsPurelyInseparable.inseparable' + +variable {F} + +theorem isPurelyInseparable_iff : IsPurelyInseparable F E ↔ ∀ x : E, + IsIntegral F x ∧ (IsSeparable F x → x ∈ (algebraMap F E).range) := + ⟨fun h x ↦ ⟨h.isIntegral' _ x, h.inseparable' x⟩, fun h ↦ ⟨⟨fun x ↦ (h x).1⟩, fun x ↦ (h x).2⟩⟩ + +variable {K} + +/-- Transfer `IsPurelyInseparable` across an `AlgEquiv`. -/ +theorem AlgEquiv.isPurelyInseparable (e : K ≃ₐ[F] E) [IsPurelyInseparable F K] : + IsPurelyInseparable F E := by + refine ⟨⟨fun _ ↦ by rw [← isIntegral_algEquiv e.symm]; exact IsPurelyInseparable.isIntegral' F _⟩, + fun x h ↦ ?_⟩ + rw [IsSeparable, ← minpoly.algEquiv_eq e.symm] at h + simpa only [RingHom.mem_range, algebraMap_eq_apply] using IsPurelyInseparable.inseparable F _ h + +theorem AlgEquiv.isPurelyInseparable_iff (e : K ≃ₐ[F] E) : + IsPurelyInseparable F K ↔ IsPurelyInseparable F E := + ⟨fun _ ↦ e.isPurelyInseparable, fun _ ↦ e.symm.isPurelyInseparable⟩ + +/-- If `E / F` is an algebraic extension, `F` is separably closed, +then `E / F` is purely inseparable. -/ +instance Algebra.IsAlgebraic.isPurelyInseparable_of_isSepClosed + {F : Type u} {E : Type v} [Field F] [Ring E] [IsDomain E] [Algebra F E] + [Algebra.IsAlgebraic F E] [IsSepClosed F] : IsPurelyInseparable F E := + ⟨inferInstance, fun x h ↦ minpoly.mem_range_of_degree_eq_one F x <| + IsSepClosed.degree_eq_one_of_irreducible F (minpoly.irreducible + (Algebra.IsIntegral.isIntegral _)) h⟩ + +variable (F E K) + +/-- If `E / F` is both purely inseparable and separable, then `algebraMap F E` is surjective. -/ +theorem IsPurelyInseparable.surjective_algebraMap_of_isSeparable + [IsPurelyInseparable F E] [Algebra.IsSeparable F E] : Function.Surjective (algebraMap F E) := + fun x ↦ IsPurelyInseparable.inseparable F x (Algebra.IsSeparable.isSeparable F x) + +/-- If `E / F` is both purely inseparable and separable, then `algebraMap F E` is bijective. -/ +theorem IsPurelyInseparable.bijective_algebraMap_of_isSeparable + [Nontrivial E] [NoZeroSMulDivisors F E] + [IsPurelyInseparable F E] [Algebra.IsSeparable F E] : Function.Bijective (algebraMap F E) := + ⟨FaithfulSMul.algebraMap_injective F E, surjective_algebraMap_of_isSeparable F E⟩ + +variable {F E} in +/-- If a subalgebra of `E / F` is both purely inseparable and separable, then it is equal +to `F`. -/ +theorem Subalgebra.eq_bot_of_isPurelyInseparable_of_isSeparable (L : Subalgebra F E) + [IsPurelyInseparable F L] [Algebra.IsSeparable F L] : L = ⊥ := bot_unique fun x hx ↦ by + obtain ⟨y, hy⟩ := IsPurelyInseparable.surjective_algebraMap_of_isSeparable F L ⟨x, hx⟩ + exact ⟨y, congr_arg (Subalgebra.val _) hy⟩ + +/-- If an intermediate field of `E / F` is both purely inseparable and separable, then it is equal +to `F`. -/ +theorem IntermediateField.eq_bot_of_isPurelyInseparable_of_isSeparable + {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] (L : IntermediateField F E) + [IsPurelyInseparable F L] [Algebra.IsSeparable F L] : L = ⊥ := bot_unique fun x hx ↦ by + obtain ⟨y, hy⟩ := IsPurelyInseparable.surjective_algebraMap_of_isSeparable F L ⟨x, hx⟩ + exact ⟨y, congr_arg (algebraMap L E) hy⟩ + +/-- If `E / F` is purely inseparable, then the separable closure of `F` in `E` is +equal to `F`. -/ +theorem separableClosure.eq_bot_of_isPurelyInseparable + (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] [IsPurelyInseparable F E] : + separableClosure F E = ⊥ := + bot_unique fun x h ↦ IsPurelyInseparable.inseparable F x (mem_separableClosure_iff.1 h) + +/-- If `E / F` is an algebraic extension, then the separable closure of `F` in `E` is +equal to `F` if and only if `E / F` is purely inseparable. -/ +theorem separableClosure.eq_bot_iff + {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] [Algebra.IsAlgebraic F E] : + separableClosure F E = ⊥ ↔ IsPurelyInseparable F E := + ⟨fun h ↦ isPurelyInseparable_iff.2 fun x ↦ ⟨Algebra.IsIntegral.isIntegral x, fun hs ↦ by + simpa only [h] using mem_separableClosure_iff.2 hs⟩, fun _ ↦ eq_bot_of_isPurelyInseparable F E⟩ + +instance isPurelyInseparable_self : IsPurelyInseparable F F := + ⟨inferInstance, fun x _ ↦ ⟨x, rfl⟩⟩ + +section + +variable (F : Type u) {E : Type v} [Field F] [Ring E] [IsDomain E] [Algebra F E] +variable (q : ℕ) [ExpChar F q] (x : E) + +/-- A field extension `E / F` of exponential characteristic `q` is purely inseparable +if and only if for every element `x` of `E`, there exists a natural number `n` such that +`x ^ (q ^ n)` is contained in `F`. -/ +@[stacks 09HE] +theorem isPurelyInseparable_iff_pow_mem : + IsPurelyInseparable F E ↔ ∀ x : E, ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by + rw [isPurelyInseparable_iff] + refine ⟨fun h x ↦ ?_, fun h x ↦ ?_⟩ + · obtain ⟨g, h1, n, h2⟩ := (minpoly.irreducible (h x).1).hasSeparableContraction q + exact ⟨n, (h _).2 <| h1.of_dvd <| minpoly.dvd F _ <| by + simpa only [expand_aeval, minpoly.aeval] using congr_arg (aeval x) h2⟩ + have hdeg := (minpoly.natSepDegree_eq_one_iff_pow_mem q).2 (h x) + have halg : IsIntegral F x := by_contra fun h' ↦ by + simp only [minpoly.eq_zero h', natSepDegree_zero, zero_ne_one] at hdeg + refine ⟨halg, fun hsep ↦ ?_⟩ + rwa [hsep.natSepDegree_eq_natDegree, minpoly.natDegree_eq_one_iff] at hdeg + +theorem IsPurelyInseparable.pow_mem [IsPurelyInseparable F E] : + ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := + (isPurelyInseparable_iff_pow_mem F q).1 ‹_› x + +end + +end General + +variable (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] +variable (K : Type w) [Field K] [Algebra F K] + +section Field + +/-- If `K / E / F` is a field extension tower such that `K / F` is purely inseparable, +then `E / F` is also purely inseparable. -/ +theorem IsPurelyInseparable.tower_bot [Algebra E K] [IsScalarTower F E K] + [IsPurelyInseparable F K] : IsPurelyInseparable F E := by + refine ⟨⟨fun x ↦ (isIntegral' F (algebraMap E K x)).tower_bot_of_field⟩, fun x h ↦ ?_⟩ + rw [IsSeparable, ← minpoly.algebraMap_eq (algebraMap E K).injective] at h + obtain ⟨y, h⟩ := inseparable F _ h + exact ⟨y, (algebraMap E K).injective (h.symm ▸ (IsScalarTower.algebraMap_apply F E K y).symm)⟩ + +/-- If `K / E / F` is a field extension tower such that `K / F` is purely inseparable, +then `K / E` is also purely inseparable. -/ +theorem IsPurelyInseparable.tower_top [Algebra E K] [IsScalarTower F E K] + [h : IsPurelyInseparable F K] : IsPurelyInseparable E K := by + obtain ⟨q, _⟩ := ExpChar.exists F + haveI := expChar_of_injective_algebraMap (algebraMap F E).injective q + rw [isPurelyInseparable_iff_pow_mem _ q] at h ⊢ + intro x + obtain ⟨n, y, h⟩ := h x + exact ⟨n, (algebraMap F E) y, h.symm ▸ (IsScalarTower.algebraMap_apply F E K y).symm⟩ + +/-- If `E / F` and `K / E` are both purely inseparable extensions, then `K / F` is also +purely inseparable. -/ +@[stacks 02JJ "See also 00GM"] +theorem IsPurelyInseparable.trans [Algebra E K] [IsScalarTower F E K] + [h1 : IsPurelyInseparable F E] [h2 : IsPurelyInseparable E K] : IsPurelyInseparable F K := by + obtain ⟨q, _⟩ := ExpChar.exists F + haveI := expChar_of_injective_algebraMap (algebraMap F E).injective q + rw [isPurelyInseparable_iff_pow_mem _ q] at h1 h2 ⊢ + intro x + obtain ⟨n, y, h2⟩ := h2 x + obtain ⟨m, z, h1⟩ := h1 y + refine ⟨n + m, z, ?_⟩ + rw [IsScalarTower.algebraMap_apply F E K, h1, map_pow, h2, ← pow_mul, ← pow_add] + +namespace IntermediateField + +variable (M : IntermediateField F K) + +instance isPurelyInseparable_tower_bot [IsPurelyInseparable F K] : IsPurelyInseparable F M := + IsPurelyInseparable.tower_bot F M K + +instance isPurelyInseparable_tower_top [IsPurelyInseparable F K] : IsPurelyInseparable M K := + IsPurelyInseparable.tower_top F M K + +end IntermediateField + +variable {E} + +/-- A field extension `E / F` is purely inseparable if and only if for every element `x` of `E`, +its minimal polynomial has separable degree one. -/ +theorem isPurelyInseparable_iff_natSepDegree_eq_one : + IsPurelyInseparable F E ↔ ∀ x : E, (minpoly F x).natSepDegree = 1 := by + obtain ⟨q, _⟩ := ExpChar.exists F + simp_rw [isPurelyInseparable_iff_pow_mem F q, minpoly.natSepDegree_eq_one_iff_pow_mem q] + +theorem IsPurelyInseparable.natSepDegree_eq_one [IsPurelyInseparable F E] (x : E) : + (minpoly F x).natSepDegree = 1 := + (isPurelyInseparable_iff_natSepDegree_eq_one F).1 ‹_› x + +/-- A field extension `E / F` of exponential characteristic `q` is purely inseparable +if and only if for every element `x` of `E`, the minimal polynomial of `x` over `F` is of form +`X ^ (q ^ n) - y` for some natural number `n` and some element `y` of `F`. -/ +theorem isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C (q : ℕ) [hF : ExpChar F q] : + IsPurelyInseparable F E ↔ ∀ x : E, ∃ (n : ℕ) (y : F), minpoly F x = X ^ q ^ n - C y := by + simp_rw [isPurelyInseparable_iff_natSepDegree_eq_one, + minpoly.natSepDegree_eq_one_iff_eq_X_pow_sub_C q] + +theorem IsPurelyInseparable.minpoly_eq_X_pow_sub_C (q : ℕ) [ExpChar F q] [IsPurelyInseparable F E] + (x : E) : ∃ (n : ℕ) (y : F), minpoly F x = X ^ q ^ n - C y := + (isPurelyInseparable_iff_minpoly_eq_X_pow_sub_C F q).1 ‹_› x + +/-- A field extension `E / F` of exponential characteristic `q` is purely inseparable +if and only if for every element `x` of `E`, the minimal polynomial of `x` over `F` is of form +`(X - x) ^ (q ^ n)` for some natural number `n`. -/ +theorem isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow (q : ℕ) [hF : ExpChar F q] : + IsPurelyInseparable F E ↔ + ∀ x : E, ∃ n : ℕ, (minpoly F x).map (algebraMap F E) = (X - C x) ^ q ^ n := by + simp_rw [isPurelyInseparable_iff_natSepDegree_eq_one, + minpoly.natSepDegree_eq_one_iff_eq_X_sub_C_pow q] + +theorem IsPurelyInseparable.minpoly_eq_X_sub_C_pow (q : ℕ) [ExpChar F q] [IsPurelyInseparable F E] + (x : E) : ∃ n : ℕ, (minpoly F x).map (algebraMap F E) = (X - C x) ^ q ^ n := + (isPurelyInseparable_iff_minpoly_eq_X_sub_C_pow F q).1 ‹_› x + +variable (E) + +variable {F E} in +/-- If an extension has finite separable degree one, then it is purely inseparable. -/ +theorem isPurelyInseparable_of_finSepDegree_eq_one + (hdeg : finSepDegree F E = 1) : IsPurelyInseparable F E := by + by_cases H : Algebra.IsAlgebraic F E + · rw [isPurelyInseparable_iff] + refine fun x ↦ ⟨Algebra.IsIntegral.isIntegral x, fun hsep ↦ ?_⟩ + have : Algebra.IsAlgebraic F⟮x⟯ E := Algebra.IsAlgebraic.tower_top (K := F) F⟮x⟯ + have := finSepDegree_mul_finSepDegree_of_isAlgebraic F F⟮x⟯ E + rw [hdeg, mul_eq_one, (finSepDegree_adjoin_simple_eq_finrank_iff F E x + (Algebra.IsAlgebraic.isAlgebraic x)).2 hsep, + IntermediateField.finrank_eq_one_iff] at this + simpa only [this.1] using mem_adjoin_simple_self F x + · rw [← Algebra.transcendental_iff_not_isAlgebraic] at H + simp [finSepDegree_eq_zero_of_transcendental F E] at hdeg + +namespace IsPurelyInseparable + +variable [IsPurelyInseparable F E] (R L : Type*) [CommSemiring R] [Algebra R F] [Algebra R E] + +/-- If `E / F` is purely inseparable, then for any reduced ring `L`, the map `(E →+* L) → (F →+* L)` +induced by `algebraMap F E` is injective. In particular, a purely inseparable field extension +is an epimorphism in the category of fields. -/ +theorem injective_comp_algebraMap [CommRing L] [IsReduced L] : + Function.Injective fun f : E →+* L ↦ f.comp (algebraMap F E) := fun f g heq ↦ by + ext x + let q := ringExpChar F + obtain ⟨n, y, h⟩ := IsPurelyInseparable.pow_mem F q x + replace heq := congr($heq y) + simp_rw [RingHom.comp_apply, h, map_pow] at heq + nontriviality L + haveI := expChar_of_injective_ringHom (f.comp (algebraMap F E)).injective q + exact iterateFrobenius_inj L q n heq + +theorem injective_restrictDomain [CommRing L] [IsReduced L] [Algebra R L] [IsScalarTower R F E] : + Function.Injective (AlgHom.restrictDomain (A := R) F (C := E) (D := L)) := fun _ _ eq ↦ + AlgHom.coe_ringHom_injective <| injective_comp_algebraMap F E L <| congr_arg AlgHom.toRingHom eq + +instance [Field L] [PerfectField L] [Algebra F L] : Nonempty (E →ₐ[F] L) := + nonempty_algHom_of_splits fun x ↦ ⟨IsPurelyInseparable.isIntegral' _ _, + have ⟨q, _⟩ := ExpChar.exists F + PerfectField.splits_of_natSepDegree_eq_one (algebraMap F L) + ((minpoly.natSepDegree_eq_one_iff_eq_X_pow_sub_C q).mpr <| + IsPurelyInseparable.minpoly_eq_X_pow_sub_C F q x)⟩ + +theorem bijective_comp_algebraMap [Field L] [PerfectField L] : + Function.Bijective fun f : E →+* L ↦ f.comp (algebraMap F E) := + ⟨injective_comp_algebraMap F E L, fun g ↦ let _ := g.toAlgebra + ⟨_, (Classical.arbitrary <| E →ₐ[F] L).comp_algebraMap⟩⟩ + +theorem bijective_restrictDomain [Field L] [PerfectField L] [Algebra R L] [IsScalarTower R F E] : + Function.Bijective (AlgHom.restrictDomain (A := R) F (C := E) (D := L)) := + ⟨injective_restrictDomain F E R L, fun g ↦ let _ := g.toAlgebra + let f := Classical.arbitrary (E →ₐ[F] L) + ⟨f.restrictScalars R, AlgHom.coe_ringHom_injective f.comp_algebraMap⟩⟩ + +end IsPurelyInseparable + +/-- If `E / F` is purely inseparable, then for any reduced `F`-algebra `L`, there exists at most one +`F`-algebra homomorphism from `E` to `L`. -/ +instance instSubsingletonAlgHomOfIsPurelyInseparable [IsPurelyInseparable F E] (L : Type w) + [CommRing L] [IsReduced L] [Algebra F L] : Subsingleton (E →ₐ[F] L) where + allEq f g := AlgHom.coe_ringHom_injective <| + IsPurelyInseparable.injective_comp_algebraMap F E L (by simp_rw [AlgHom.comp_algebraMap]) + +instance instUniqueAlgHomOfIsPurelyInseparable [IsPurelyInseparable F E] (L : Type w) + [CommRing L] [IsReduced L] [Algebra F L] [Algebra E L] [IsScalarTower F E L] : + Unique (E →ₐ[F] L) := uniqueOfSubsingleton (IsScalarTower.toAlgHom F E L) + +/-- If `E / F` is purely inseparable, then `Field.Emb F E` has exactly one element. -/ +instance instUniqueEmbOfIsPurelyInseparable [IsPurelyInseparable F E] : + Unique (Emb F E) := instUniqueAlgHomOfIsPurelyInseparable F E _ + +/-- A purely inseparable extension has finite separable degree one. -/ +theorem IsPurelyInseparable.finSepDegree_eq_one [IsPurelyInseparable F E] : + finSepDegree F E = 1 := Nat.card_unique + +/-- A purely inseparable extension has separable degree one. -/ +theorem IsPurelyInseparable.sepDegree_eq_one [IsPurelyInseparable F E] : + sepDegree F E = 1 := by + rw [sepDegree, separableClosure.eq_bot_of_isPurelyInseparable, IntermediateField.rank_bot] + +/-- A purely inseparable extension has inseparable degree equal to degree. -/ +theorem IsPurelyInseparable.insepDegree_eq [IsPurelyInseparable F E] : + insepDegree F E = Module.rank F E := by + rw [insepDegree, separableClosure.eq_bot_of_isPurelyInseparable, rank_bot'] + +/-- A purely inseparable extension has finite inseparable degree equal to degree. -/ +theorem IsPurelyInseparable.finInsepDegree_eq [IsPurelyInseparable F E] : + finInsepDegree F E = finrank F E := congr(Cardinal.toNat $(insepDegree_eq F E)) + +/-- An extension is purely inseparable if and only if it has finite separable degree one. -/ +theorem isPurelyInseparable_iff_finSepDegree_eq_one : + IsPurelyInseparable F E ↔ finSepDegree F E = 1 := + ⟨fun _ ↦ IsPurelyInseparable.finSepDegree_eq_one F E, + fun h ↦ isPurelyInseparable_of_finSepDegree_eq_one h⟩ + +variable {F E} in +/-- An algebraic extension is purely inseparable if and only if all of its finite dimensional +subextensions are purely inseparable. -/ +theorem isPurelyInseparable_iff_fd_isPurelyInseparable [Algebra.IsAlgebraic F E] : + IsPurelyInseparable F E ↔ + ∀ L : IntermediateField F E, FiniteDimensional F L → IsPurelyInseparable F L := by + refine ⟨fun _ _ _ ↦ IsPurelyInseparable.tower_bot F _ E, + fun h ↦ isPurelyInseparable_iff.2 fun x ↦ ?_⟩ + have hx : IsIntegral F x := Algebra.IsIntegral.isIntegral x + refine ⟨hx, fun _ ↦ ?_⟩ + obtain ⟨y, h⟩ := (h _ (adjoin.finiteDimensional hx)).inseparable' _ <| + show Separable (minpoly F (AdjoinSimple.gen F x)) by rwa [minpoly_eq] + exact ⟨y, congr_arg (algebraMap _ E) h⟩ + +/-- A purely inseparable extension is normal. -/ +instance IsPurelyInseparable.normal [IsPurelyInseparable F E] : Normal F E where + toIsAlgebraic := isAlgebraic F E + splits' x := by + obtain ⟨n, h⟩ := IsPurelyInseparable.minpoly_eq_X_sub_C_pow F (ringExpChar F) x + rw [← splits_id_iff_splits, h] + exact splits_pow _ (splits_X_sub_C _) _ + +/-- If `E / F` is algebraic, then `E` is purely inseparable over the +separable closure of `F` in `E`. -/ +@[stacks 030K "$E/E_{sep}$ is purely inseparable."] +instance separableClosure.isPurelyInseparable [Algebra.IsAlgebraic F E] : + IsPurelyInseparable (separableClosure F E) E := isPurelyInseparable_iff.2 fun x ↦ by + set L := separableClosure F E + refine ⟨(IsAlgebraic.tower_top L (Algebra.IsAlgebraic.isAlgebraic (R := F) x)).isIntegral, + fun h ↦ ?_⟩ + haveI := (isSeparable_adjoin_simple_iff_isSeparable L E).2 h + haveI : Algebra.IsSeparable F (restrictScalars F L⟮x⟯) := Algebra.IsSeparable.trans F L L⟮x⟯ + have hx : x ∈ restrictScalars F L⟮x⟯ := mem_adjoin_simple_self _ x + exact ⟨⟨x, mem_separableClosure_iff.2 <| isSeparable_of_mem_isSeparable F E hx⟩, rfl⟩ + +open Cardinal in +theorem Field.Emb.cardinal_separableClosure [Algebra.IsAlgebraic F E] : + #(Field.Emb F <| separableClosure F E) = #(Field.Emb F E) := by + rw [← (embProdEmbOfIsAlgebraic F (separableClosure F E) E).cardinal_eq, + mk_prod, mk_eq_one (Emb _ E), lift_one, mul_one, lift_id] + +/-- An intermediate field of `E / F` contains the separable closure of `F` in `E` +if `E` is purely inseparable over it. -/ +theorem separableClosure_le (L : IntermediateField F E) + [h : IsPurelyInseparable L E] : separableClosure F E ≤ L := fun x hx ↦ by + obtain ⟨y, rfl⟩ := h.inseparable' _ <| + IsSeparable.tower_top L (mem_separableClosure_iff.1 hx) + exact y.2 + +/-- If `E / F` is algebraic, then an intermediate field of `E / F` contains the +separable closure of `F` in `E` if and only if `E` is purely inseparable over it. -/ +theorem separableClosure_le_iff [Algebra.IsAlgebraic F E] (L : IntermediateField F E) : + separableClosure F E ≤ L ↔ IsPurelyInseparable L E := by + refine ⟨fun h ↦ ?_, fun _ ↦ separableClosure_le F E L⟩ + have := separableClosure.isPurelyInseparable F E + letI := (inclusion h).toAlgebra + letI : SMul (separableClosure F E) L := Algebra.toSMul + haveI : IsScalarTower (separableClosure F E) L E := IsScalarTower.of_algebraMap_eq (congrFun rfl) + exact IsPurelyInseparable.tower_top (separableClosure F E) L E + +/-- If an intermediate field of `E / F` is separable over `F`, and `E` is purely inseparable +over it, then it is equal to the separable closure of `F` in `E`. -/ +theorem eq_separableClosure (L : IntermediateField F E) + [Algebra.IsSeparable F L] [IsPurelyInseparable L E] : L = separableClosure F E := + le_antisymm (le_separableClosure F E L) (separableClosure_le F E L) + +open separableClosure in +/-- If `E / F` is algebraic, then an intermediate field of `E / F` is equal to the separable closure +of `F` in `E` if and only if it is separable over `F`, and `E` is purely inseparable +over it. -/ +theorem eq_separableClosure_iff [Algebra.IsAlgebraic F E] (L : IntermediateField F E) : + L = separableClosure F E ↔ Algebra.IsSeparable F L ∧ IsPurelyInseparable L E := + ⟨by rintro rfl; exact ⟨isSeparable F E, isPurelyInseparable F E⟩, + fun ⟨_, _⟩ ↦ eq_separableClosure F E L⟩ + +/-- If `L` is an algebraically closed field containing `E`, such that the map +`(E →+* L) → (F →+* L)` induced by `algebraMap F E` is injective, then `E / F` is +purely inseparable. As a corollary, epimorphisms in the category of fields must be +purely inseparable extensions. -/ +theorem IsPurelyInseparable.of_injective_comp_algebraMap (L : Type w) [Field L] [IsAlgClosed L] + [Nonempty (E →+* L)] (h : Function.Injective fun f : E →+* L ↦ f.comp (algebraMap F E)) : + IsPurelyInseparable F E := by + rw [isPurelyInseparable_iff_finSepDegree_eq_one, finSepDegree, Nat.card_eq_one_iff_unique] + letI := (Classical.arbitrary (E →+* L)).toAlgebra + let j : AlgebraicClosure E →ₐ[E] L := IsAlgClosed.lift + exact ⟨⟨fun f g ↦ DFunLike.ext' <| j.injective.comp_left (congr_arg (⇑) <| + @h (j.toRingHom.comp f) (j.toRingHom.comp g) (by ext; simp))⟩, inferInstance⟩ + +end Field + +namespace IntermediateField + +instance isPurelyInseparable_bot : IsPurelyInseparable F (⊥ : IntermediateField F E) := + (botEquiv F E).symm.isPurelyInseparable + +end IntermediateField + +/-- If `E` is an algebraic closure of `F`, then `F` is separably closed if and only if `E / F` is +purely inseparable. -/ +theorem isSepClosed_iff_isPurelyInseparable_algebraicClosure [IsAlgClosure F E] : + IsSepClosed F ↔ IsPurelyInseparable F E := + ⟨fun _ ↦ inferInstance, fun H ↦ by + haveI := IsAlgClosure.isAlgClosed F (K := E) + rwa [← separableClosure.eq_bot_iff, IsSepClosed.separableClosure_eq_bot_iff] at H⟩ + +variable {F E} in +/-- If `E / F` is an algebraic extension, `F` is separably closed, +then `E` is also separably closed. -/ +theorem Algebra.IsAlgebraic.isSepClosed [Algebra.IsAlgebraic F E] + [IsSepClosed F] : IsSepClosed E := + have : Algebra.IsAlgebraic F (AlgebraicClosure E) := Algebra.IsAlgebraic.trans (L := E) + (isSepClosed_iff_isPurelyInseparable_algebraicClosure E _).mpr + (IsPurelyInseparable.tower_top F E <| AlgebraicClosure E) + +namespace Field + +/-- If `E / F` is algebraic, then the `Field.finSepDegree F E` is equal to `Field.sepDegree F E` +as a natural number. This means that the cardinality of `Field.Emb F E` and the degree of +`(separableClosure F E) / F` are both finite or infinite, and when they are finite, they +coincide. -/ +@[stacks 09HJ "`sepDegree` is defined as the right hand side of 09HJ"] +theorem finSepDegree_eq [Algebra.IsAlgebraic F E] : + finSepDegree F E = Cardinal.toNat (sepDegree F E) := by + have : Algebra.IsAlgebraic (separableClosure F E) E := Algebra.IsAlgebraic.tower_top (K := F) _ + have h := finSepDegree_mul_finSepDegree_of_isAlgebraic F (separableClosure F E) E |>.symm + haveI := separableClosure.isSeparable F E + haveI := separableClosure.isPurelyInseparable F E + rwa [finSepDegree_eq_finrank_of_isSeparable F (separableClosure F E), + IsPurelyInseparable.finSepDegree_eq_one (separableClosure F E) E, mul_one] at h + +/-- The finite separable degree multiply by the finite inseparable degree is equal +to the (finite) field extension degree. -/ +theorem finSepDegree_mul_finInsepDegree : finSepDegree F E * finInsepDegree F E = finrank F E := by + by_cases halg : Algebra.IsAlgebraic F E + · have := congr_arg Cardinal.toNat (sepDegree_mul_insepDegree F E) + rwa [Cardinal.toNat_mul, ← finSepDegree_eq F E] at this + rw [finInsepDegree, finrank_of_infinite_dimensional (K := F) (V := E) fun _ ↦ + halg (Algebra.IsAlgebraic.of_finite F E), + finrank_of_infinite_dimensional (K := separableClosure F E) (V := E) fun _ ↦ + halg ((separableClosure.isAlgebraic F E).trans), + mul_zero] + +end Field + +namespace separableClosure + +variable [Algebra E K] [IsScalarTower F E K] {F E} + +/-- If `K / E / F` is a field extension tower, such that `E / F` is algebraic and `K / E` +is separable, then `E` adjoin `separableClosure F K` is equal to `K`. It is a special case of +`separableClosure.adjoin_eq_of_isAlgebraic`, and is an intermediate result used to prove it. -/ +lemma adjoin_eq_of_isAlgebraic_of_isSeparable [Algebra.IsAlgebraic F E] + [Algebra.IsSeparable E K] : adjoin E (separableClosure F K : Set K) = ⊤ := + top_unique fun x _ ↦ by + set S := separableClosure F K + set L := adjoin E (S : Set K) + have := Algebra.isSeparable_tower_top_of_isSeparable E L K + let i : S →+* L := Subsemiring.inclusion fun x hx ↦ subset_adjoin E (S : Set K) hx + let _ : Algebra S L := i.toAlgebra + let _ : SMul S L := Algebra.toSMul + have : IsScalarTower S L K := IsScalarTower.of_algebraMap_eq (congrFun rfl) + have : Algebra.IsAlgebraic F K := Algebra.IsAlgebraic.trans (L := E) + have : IsPurelyInseparable S K := separableClosure.isPurelyInseparable F K + have := IsPurelyInseparable.tower_top S L K + obtain ⟨y, rfl⟩ := IsPurelyInseparable.surjective_algebraMap_of_isSeparable L K x + exact y.2 + +/-- If `K / E / F` is a field extension tower, such that `E / F` is algebraic, then +`E` adjoin `separableClosure F K` is equal to `separableClosure E K`. -/ +theorem adjoin_eq_of_isAlgebraic [Algebra.IsAlgebraic F E] : + adjoin E (separableClosure F K) = separableClosure E K := by + set S := separableClosure E K + have h := congr_arg lift (adjoin_eq_of_isAlgebraic_of_isSeparable (F := F) S) + rw [lift_top, lift_adjoin] at h + haveI : IsScalarTower F S K := IsScalarTower.of_algebraMap_eq (congrFun rfl) + rw [← h, ← map_eq_of_separableClosure_eq_bot F (separableClosure_eq_bot E K)] + simp only [S, coe_map, IsScalarTower.coe_toAlgHom', IntermediateField.algebraMap_apply] + +end separableClosure diff --git a/Mathlib/FieldTheory/PurelyInseparable/PerfectClosure.lean b/Mathlib/FieldTheory/PurelyInseparable/PerfectClosure.lean new file mode 100644 index 00000000000000..fb3adb8fe2e6aa --- /dev/null +++ b/Mathlib/FieldTheory/PurelyInseparable/PerfectClosure.lean @@ -0,0 +1,410 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.Algebra.CharP.ExpChar +import Mathlib.Algebra.CharP.IntermediateField +import Mathlib.FieldTheory.PurelyInseparable.Basic + +/-! + +# Basic results about relative perfect closure + +This file contains basic results about relative perfect closures. + +## Main definitions + +- `perfectClosure`: the relative perfect closure of `F` in `E`, it consists of the elements + `x` of `E` such that there exists a natural number `n` such that `x ^ (ringExpChar F) ^ n` + is contained in `F`, where `ringExpChar F` is the exponential characteristic of `F`. + It is also the maximal purely inseparable subextension of `E / F` (`le_perfectClosure_iff`). + +## Main results + +- `le_perfectClosure_iff`: an intermediate field of `E / F` is contained in the relative perfect + closure of `F` in `E` if and only if it is purely inseparable over `F`. + +- `perfectClosure.perfectRing`, `perfectClosure.perfectField`: if `E` is a perfect field, then the + (relative) perfect closure `perfectClosure F E` is perfect. + +- `IntermediateField.isPurelyInseparable_adjoin_iff_pow_mem`: if `F` is of exponential + characteristic `q`, then `F(S) / F` is a purely inseparable extension if and only if for any + `x ∈ S`, `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. + +## Tags + +separable degree, degree, separable closure, purely inseparable + +-/ + +open IntermediateField + +noncomputable section + +universe u v w + +variable (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] +variable (K : Type w) [Field K] [Algebra F K] + +section perfectClosure + +/-- The relative perfect closure of `F` in `E`, consists of the elements `x` of `E` such that there +exists a natural number `n` such that `x ^ (ringExpChar F) ^ n` is contained in `F`, where +`ringExpChar F` is the exponential characteristic of `F`. It is also the maximal purely inseparable +subextension of `E / F` (`le_perfectClosure_iff`). -/ +@[stacks 09HH] +def perfectClosure : IntermediateField F E where + carrier := {x : E | ∃ n : ℕ, x ^ (ringExpChar F) ^ n ∈ (algebraMap F E).range} + add_mem' := by + rintro x y ⟨n, hx⟩ ⟨m, hy⟩ + use n + m + have := expChar_of_injective_algebraMap (algebraMap F E).injective (ringExpChar F) + rw [add_pow_expChar_pow, pow_add, pow_mul, mul_comm (_ ^ n), pow_mul] + exact add_mem (pow_mem hx _) (pow_mem hy _) + mul_mem' := by + rintro x y ⟨n, hx⟩ ⟨m, hy⟩ + use n + m + rw [mul_pow, pow_add, pow_mul, mul_comm (_ ^ n), pow_mul] + exact mul_mem (pow_mem hx _) (pow_mem hy _) + inv_mem' := by + rintro x ⟨n, hx⟩ + use n; rw [inv_pow] + apply inv_mem (id hx : _ ∈ (⊥ : IntermediateField F E)) + algebraMap_mem' := fun x ↦ ⟨0, by rw [pow_zero, pow_one]; exact ⟨x, rfl⟩⟩ + +variable {F E} + +theorem mem_perfectClosure_iff {x : E} : + x ∈ perfectClosure F E ↔ ∃ n : ℕ, x ^ (ringExpChar F) ^ n ∈ (algebraMap F E).range := Iff.rfl + +theorem mem_perfectClosure_iff_pow_mem (q : ℕ) [ExpChar F q] {x : E} : + x ∈ perfectClosure F E ↔ ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by + rw [mem_perfectClosure_iff, ringExpChar.eq F q] + +/-- An element is contained in the relative perfect closure if and only if its minimal polynomial +has separable degree one. -/ +theorem mem_perfectClosure_iff_natSepDegree_eq_one {x : E} : + x ∈ perfectClosure F E ↔ (minpoly F x).natSepDegree = 1 := by + rw [mem_perfectClosure_iff, minpoly.natSepDegree_eq_one_iff_pow_mem (ringExpChar F)] + +/-- A field extension `E / F` is purely inseparable if and only if the relative perfect closure of +`F` in `E` is equal to `E`. -/ +theorem isPurelyInseparable_iff_perfectClosure_eq_top : + IsPurelyInseparable F E ↔ perfectClosure F E = ⊤ := by + rw [isPurelyInseparable_iff_pow_mem F (ringExpChar F)] + exact ⟨fun H ↦ top_unique fun x _ ↦ H x, fun H _ ↦ H.ge trivial⟩ + +variable (F E) + +/-- The relative perfect closure of `F` in `E` is purely inseparable over `F`. -/ +instance perfectClosure.isPurelyInseparable : IsPurelyInseparable F (perfectClosure F E) := by + rw [isPurelyInseparable_iff_pow_mem F (ringExpChar F)] + exact fun ⟨_, n, y, h⟩ ↦ ⟨n, y, (algebraMap _ E).injective h⟩ + +/-- The relative perfect closure of `F` in `E` is algebraic over `F`. -/ +instance perfectClosure.isAlgebraic : Algebra.IsAlgebraic F (perfectClosure F E) := + IsPurelyInseparable.isAlgebraic F _ + +/-- If `E / F` is separable, then the perfect closure of `F` in `E` is equal to `F`. Note that + the converse is not necessarily true (see https://math.stackexchange.com/a/3009197) + even when `E / F` is algebraic. -/ +theorem perfectClosure.eq_bot_of_isSeparable [Algebra.IsSeparable F E] : perfectClosure F E = ⊥ := + haveI := Algebra.isSeparable_tower_bot_of_isSeparable F (perfectClosure F E) E + eq_bot_of_isPurelyInseparable_of_isSeparable _ + +/-- An intermediate field of `E / F` is contained in the relative perfect closure of `F` in `E` +if it is purely inseparable over `F`. -/ +theorem le_perfectClosure (L : IntermediateField F E) [h : IsPurelyInseparable F L] : + L ≤ perfectClosure F E := by + rw [isPurelyInseparable_iff_pow_mem F (ringExpChar F)] at h + intro x hx + obtain ⟨n, y, hy⟩ := h ⟨x, hx⟩ + exact ⟨n, y, congr_arg (algebraMap L E) hy⟩ + +/-- An intermediate field of `E / F` is contained in the relative perfect closure of `F` in `E` +if and only if it is purely inseparable over `F`. -/ +theorem le_perfectClosure_iff (L : IntermediateField F E) : + L ≤ perfectClosure F E ↔ IsPurelyInseparable F L := by + refine ⟨fun h ↦ (isPurelyInseparable_iff_pow_mem F (ringExpChar F)).2 fun x ↦ ?_, + fun _ ↦ le_perfectClosure F E L⟩ + obtain ⟨n, y, hy⟩ := h x.2 + exact ⟨n, y, (algebraMap L E).injective hy⟩ + +theorem separableClosure_inf_perfectClosure : separableClosure F E ⊓ perfectClosure F E = ⊥ := + haveI := (le_separableClosure_iff F E _).mp (inf_le_left (b := perfectClosure F E)) + haveI := (le_perfectClosure_iff F E _).mp (inf_le_right (a := separableClosure F E)) + eq_bot_of_isPurelyInseparable_of_isSeparable _ + +section map + +variable {F E K} + +/-- If `i` is an `F`-algebra homomorphism from `E` to `K`, then `i x` is contained in +`perfectClosure F K` if and only if `x` is contained in `perfectClosure F E`. -/ +theorem map_mem_perfectClosure_iff (i : E →ₐ[F] K) {x : E} : + i x ∈ perfectClosure F K ↔ x ∈ perfectClosure F E := by + simp_rw [mem_perfectClosure_iff] + refine ⟨fun ⟨n, y, h⟩ ↦ ⟨n, y, ?_⟩, fun ⟨n, y, h⟩ ↦ ⟨n, y, ?_⟩⟩ + · apply_fun i using i.injective + rwa [AlgHom.commutes, map_pow] + simpa only [AlgHom.commutes, map_pow] using congr_arg i h + +/-- If `i` is an `F`-algebra homomorphism from `E` to `K`, then the preimage of `perfectClosure F K` +under the map `i` is equal to `perfectClosure F E`. -/ +theorem perfectClosure.comap_eq_of_algHom (i : E →ₐ[F] K) : + (perfectClosure F K).comap i = perfectClosure F E := by + ext x + exact map_mem_perfectClosure_iff i + +/-- If `i` is an `F`-algebra homomorphism from `E` to `K`, then the image of `perfectClosure F E` +under the map `i` is contained in `perfectClosure F K`. -/ +theorem perfectClosure.map_le_of_algHom (i : E →ₐ[F] K) : + (perfectClosure F E).map i ≤ perfectClosure F K := + map_le_iff_le_comap.mpr (perfectClosure.comap_eq_of_algHom i).ge + +/-- If `i` is an `F`-algebra isomorphism of `E` and `K`, then the image of `perfectClosure F E` +under the map `i` is equal to in `perfectClosure F K`. -/ +theorem perfectClosure.map_eq_of_algEquiv (i : E ≃ₐ[F] K) : + (perfectClosure F E).map i.toAlgHom = perfectClosure F K := + (map_le_of_algHom i.toAlgHom).antisymm (fun x hx ↦ ⟨i.symm x, + (map_mem_perfectClosure_iff i.symm.toAlgHom).2 hx, i.right_inv x⟩) + +/-- If `E` and `K` are isomorphic as `F`-algebras, then `perfectClosure F E` and +`perfectClosure F K` are also isomorphic as `F`-algebras. -/ +def perfectClosure.algEquivOfAlgEquiv (i : E ≃ₐ[F] K) : + perfectClosure F E ≃ₐ[F] perfectClosure F K := + (intermediateFieldMap i _).trans (equivOfEq (map_eq_of_algEquiv i)) + +alias AlgEquiv.perfectClosure := perfectClosure.algEquivOfAlgEquiv + +end map + +/-- If `E` is a perfect field of exponential characteristic `p`, then the (relative) perfect closure +`perfectClosure F E` is perfect. -/ +instance perfectClosure.perfectRing (p : ℕ) [ExpChar E p] + [PerfectRing E p] : PerfectRing (perfectClosure F E) p := .ofSurjective _ p fun x ↦ by + haveI := RingHom.expChar _ (algebraMap F E).injective p + obtain ⟨x', hx⟩ := surjective_frobenius E p x.1 + obtain ⟨n, y, hy⟩ := (mem_perfectClosure_iff_pow_mem p).1 x.2 + rw [frobenius_def] at hx + rw [← hx, ← pow_mul, ← pow_succ'] at hy + exact ⟨⟨x', (mem_perfectClosure_iff_pow_mem p).2 ⟨n + 1, y, hy⟩⟩, by + simp_rw [frobenius_def, SubmonoidClass.mk_pow, hx]⟩ + +/-- If `E` is a perfect field, then the (relative) perfect closure +`perfectClosure F E` is perfect. -/ +instance perfectClosure.perfectField [PerfectField E] : PerfectField (perfectClosure F E) := + PerfectRing.toPerfectField _ (ringExpChar E) + +end perfectClosure + +namespace IntermediateField + +/-- `F⟮x⟯ / F` is a purely inseparable extension if and only if the minimal polynomial of `x` +has separable degree one. -/ +theorem isPurelyInseparable_adjoin_simple_iff_natSepDegree_eq_one {x : E} : + IsPurelyInseparable F F⟮x⟯ ↔ (minpoly F x).natSepDegree = 1 := by + rw [← le_perfectClosure_iff, adjoin_simple_le_iff, mem_perfectClosure_iff_natSepDegree_eq_one] + +/-- If `F` is of exponential characteristic `q`, then `F⟮x⟯ / F` is a purely inseparable extension +if and only if `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. -/ +theorem isPurelyInseparable_adjoin_simple_iff_pow_mem (q : ℕ) [hF : ExpChar F q] {x : E} : + IsPurelyInseparable F F⟮x⟯ ↔ ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by + rw [← le_perfectClosure_iff, adjoin_simple_le_iff, mem_perfectClosure_iff_pow_mem q] + +/-- If `F` is of exponential characteristic `q`, then `F(S) / F` is a purely inseparable extension +if and only if for any `x ∈ S`, `x ^ (q ^ n)` is contained in `F` for some `n : ℕ`. -/ +theorem isPurelyInseparable_adjoin_iff_pow_mem (q : ℕ) [hF : ExpChar F q] {S : Set E} : + IsPurelyInseparable F (adjoin F S) ↔ ∀ x ∈ S, ∃ n : ℕ, x ^ q ^ n ∈ (algebraMap F E).range := by + simp_rw [← le_perfectClosure_iff, adjoin_le_iff, ← mem_perfectClosure_iff_pow_mem q, + Set.subset_def, SetLike.mem_coe] + +/-- A compositum of two purely inseparable extensions is purely inseparable. -/ +instance isPurelyInseparable_sup (L1 L2 : IntermediateField F E) + [h1 : IsPurelyInseparable F L1] [h2 : IsPurelyInseparable F L2] : + IsPurelyInseparable F (L1 ⊔ L2 : IntermediateField F E) := by + rw [← le_perfectClosure_iff] at h1 h2 ⊢ + exact sup_le h1 h2 + +/-- A compositum of purely inseparable extensions is purely inseparable. -/ +instance isPurelyInseparable_iSup {ι : Sort*} {t : ι → IntermediateField F E} + [h : ∀ i, IsPurelyInseparable F (t i)] : + IsPurelyInseparable F (⨆ i, t i : IntermediateField F E) := by + simp_rw [← le_perfectClosure_iff] at h ⊢ + exact iSup_le h + +/-- If `F` is a field of exponential characteristic `q`, `F(S) / F` is separable, then +`F(S) = F(S ^ (q ^ n))` for any natural number `n`. -/ +theorem adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable (S : Set E) + [Algebra.IsSeparable F (adjoin F S)] (q : ℕ) [ExpChar F q] (n : ℕ) : + adjoin F S = adjoin F ((· ^ q ^ n) '' S) := by + set L := adjoin F S + set M := adjoin F ((· ^ q ^ n) '' S) + have hi : M ≤ L := by + rw [adjoin_le_iff] + rintro _ ⟨y, hy, rfl⟩ + exact pow_mem (subset_adjoin F S hy) _ + letI := (inclusion hi).toAlgebra + haveI : Algebra.IsSeparable M (extendScalars hi) := + Algebra.isSeparable_tower_top_of_isSeparable F M L + haveI : IsPurelyInseparable M (extendScalars hi) := by + haveI := expChar_of_injective_algebraMap (algebraMap F M).injective q + rw [extendScalars_adjoin hi, isPurelyInseparable_adjoin_iff_pow_mem M _ q] + exact fun x hx ↦ ⟨n, ⟨x ^ q ^ n, subset_adjoin F _ ⟨x, hx, rfl⟩⟩, rfl⟩ + simpa only [extendScalars_restrictScalars, restrictScalars_bot_eq_self] using congr_arg + (restrictScalars F) (extendScalars hi).eq_bot_of_isPurelyInseparable_of_isSeparable + +/-- If `E / F` is a separable field extension of exponential characteristic `q`, then +`F(S) = F(S ^ (q ^ n))` for any subset `S` of `E` and any natural number `n`. -/ +theorem adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable' [Algebra.IsSeparable F E] (S : Set E) + (q : ℕ) [ExpChar F q] (n : ℕ) : adjoin F S = adjoin F ((· ^ q ^ n) '' S) := + haveI := Algebra.isSeparable_tower_bot_of_isSeparable F (adjoin F S) E + adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E S q n + +-- TODO: prove the converse when `F(S) / F` is finite +/-- If `F` is a field of exponential characteristic `q`, `F(S) / F` is separable, then +`F(S) = F(S ^ q)`. -/ +theorem adjoin_eq_adjoin_pow_expChar_of_isSeparable (S : Set E) [Algebra.IsSeparable F (adjoin F S)] + (q : ℕ) [ExpChar F q] : adjoin F S = adjoin F ((· ^ q) '' S) := + pow_one q ▸ adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E S q 1 + +/-- If `E / F` is a separable field extension of exponential characteristic `q`, then +`F(S) = F(S ^ q)` for any subset `S` of `E`. -/ +theorem adjoin_eq_adjoin_pow_expChar_of_isSeparable' [Algebra.IsSeparable F E] (S : Set E) + (q : ℕ) [ExpChar F q] : adjoin F S = adjoin F ((· ^ q) '' S) := + pow_one q ▸ adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable' F E S q 1 + +-- Special cases for simple adjoin + +/-- If `F` is a field of exponential characteristic `q`, `a : E` is separable over `F`, then +`F⟮a⟯ = F⟮a ^ q ^ n⟯` for any natural number `n`. -/ +theorem adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable {a : E} (ha : IsSeparable F a) + (q : ℕ) [ExpChar F q] (n : ℕ) : F⟮a⟯ = F⟮a ^ q ^ n⟯ := by + haveI := (isSeparable_adjoin_simple_iff_isSeparable F E).mpr ha + simpa using adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E {a} q n + +/-- If `E / F` is a separable field extension of exponential characteristic `q`, then +`F⟮a⟯ = F⟮a ^ q ^ n⟯` for any subset `a : E` and any natural number `n`. -/ +theorem adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable' [Algebra.IsSeparable F E] (a : E) + (q : ℕ) [ExpChar F q] (n : ℕ) : F⟮a⟯ = F⟮a ^ q ^ n⟯ := by + haveI := Algebra.isSeparable_tower_bot_of_isSeparable F F⟮a⟯ E + simpa using adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable F E {a} q n + +/-- If `F` is a field of exponential characteristic `q`, `a : E` is separable over `F`, then +`F⟮a⟯ = F⟮a ^ q⟯`. -/ +theorem adjoin_simple_eq_adjoin_pow_expChar_of_isSeparable {a : E} (ha : IsSeparable F a) + (q : ℕ) [ExpChar F q] : F⟮a⟯ = F⟮a ^ q⟯ := + pow_one q ▸ adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable F E ha q 1 + +/-- If `E / F` is a separable field extension of exponential characteristic `q`, then +`F⟮a⟯ = F⟮a ^ q⟯` for any `a : E`. -/ +theorem adjoin_simple_eq_adjoin_pow_expChar_of_isSeparable' [Algebra.IsSeparable F E] (a : E) + (q : ℕ) [ExpChar F q] : F⟮a⟯ = F⟮a ^ q⟯ := + pow_one q ▸ adjoin_simple_eq_adjoin_pow_expChar_pow_of_isSeparable' F E a q 1 + +end IntermediateField + +section + +variable (q n : ℕ) [hF : ExpChar F q] {ι : Type*} {v : ι → E} {F E} + +/-- If `E / F` is a separable extension of exponential characteristic `q`, if `{ u_i }` is a family +of elements of `E` which `F`-linearly spans `E`, then `{ u_i ^ (q ^ n) }` also `F`-linearly spans +`E` for any natural number `n`. -/ +theorem Field.span_map_pow_expChar_pow_eq_top_of_isSeparable [Algebra.IsSeparable F E] + (h : Submodule.span F (Set.range v) = ⊤) : + Submodule.span F (Set.range (v · ^ q ^ n)) = ⊤ := by + erw [← Algebra.top_toSubmodule, ← top_toSubalgebra, ← adjoin_univ, + adjoin_eq_adjoin_pow_expChar_pow_of_isSeparable' F E _ q n, + adjoin_algebraic_toSubalgebra fun x _ ↦ Algebra.IsAlgebraic.isAlgebraic x, + Set.image_univ, Algebra.adjoin_eq_span, (powMonoidHom _).mrange.closure_eq] + refine (Submodule.span_mono <| Set.range_comp_subset_range _ _).antisymm (Submodule.span_le.2 ?_) + rw [Set.range_comp, ← Set.image_univ] + haveI := expChar_of_injective_algebraMap (algebraMap F E).injective q + apply h ▸ Submodule.image_span_subset_span (LinearMap.iterateFrobenius F E q n) _ + +/-- If `E / F` is a finite separable extension of exponential characteristic `q`, if `{ u_i }` is a +family of elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` is also +`F`-linearly independent for any natural number `n`. A special case of +`LinearIndependent.map_pow_expChar_pow_of_isSeparable` +and is an intermediate result used to prove it. -/ +private theorem LinearIndependent.map_pow_expChar_pow_of_fd_isSeparable + [FiniteDimensional F E] [Algebra.IsSeparable F E] + (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by + have h' := h.coe_range + let ι' := h'.extend (Set.range v).subset_univ + let b : Basis ι' F E := Basis.extend h' + letI : Fintype ι' := FiniteDimensional.fintypeBasisIndex b + have H := linearIndependent_of_top_le_span_of_card_eq_finrank + (Field.span_map_pow_expChar_pow_eq_top_of_isSeparable q n b.span_eq).ge + (Module.finrank_eq_card_basis b).symm + let f (i : ι) : ι' := ⟨v i, h'.subset_extend _ ⟨i, rfl⟩⟩ + convert H.comp f fun _ _ heq ↦ h.injective (by simpa only [f, Subtype.mk.injEq] using heq) + simp_rw [Function.comp_apply, b] + rw [Basis.extend_apply_self] + +/-- If `E / F` is a separable extension of exponential characteristic `q`, if `{ u_i }` is a +family of elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` is also +`F`-linearly independent for any natural number `n`. -/ +theorem LinearIndependent.map_pow_expChar_pow_of_isSeparable [Algebra.IsSeparable F E] + (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by + classical + have halg := Algebra.IsSeparable.isAlgebraic F E + rw [linearIndependent_iff_finset_linearIndependent] at h ⊢ + intro s + let E' := adjoin F (s.image v : Set E) + haveI : FiniteDimensional F E' := finiteDimensional_adjoin + fun x _ ↦ Algebra.IsIntegral.isIntegral x + haveI : Algebra.IsSeparable F E' := Algebra.isSeparable_tower_bot_of_isSeparable F E' E + let v' (i : s) : E' := ⟨v i.1, subset_adjoin F _ (Finset.mem_image.2 ⟨i.1, i.2, rfl⟩)⟩ + have h' : LinearIndependent F v' := (h s).of_comp E'.val.toLinearMap + exact (h'.map_pow_expChar_pow_of_fd_isSeparable q n).map' + E'.val.toLinearMap (LinearMap.ker_eq_bot_of_injective E'.val.injective) + +/-- If `E / F` is a field extension of exponential characteristic `q`, if `{ u_i }` is a +family of separable elements of `E` which is `F`-linearly independent, then `{ u_i ^ (q ^ n) }` +is also `F`-linearly independent for any natural number `n`. -/ +theorem LinearIndependent.map_pow_expChar_pow_of_isSeparable' + (hsep : ∀ i : ι, IsSeparable F (v i)) + (h : LinearIndependent F v) : LinearIndependent F (v · ^ q ^ n) := by + let E' := adjoin F (Set.range v) + haveI : Algebra.IsSeparable F E' := (isSeparable_adjoin_iff_isSeparable F _).2 <| by + rintro _ ⟨y, rfl⟩; exact hsep y + let v' (i : ι) : E' := ⟨v i, subset_adjoin F _ ⟨i, rfl⟩⟩ + have h' : LinearIndependent F v' := h.of_comp E'.val.toLinearMap + exact (h'.map_pow_expChar_pow_of_isSeparable q n).map' + E'.val.toLinearMap (LinearMap.ker_eq_bot_of_injective E'.val.injective) + +/-- If `E / F` is a separable extension of exponential characteristic `q`, if `{ u_i }` is an +`F`-basis of `E`, then `{ u_i ^ (q ^ n) }` is also an `F`-basis of `E` +for any natural number `n`. -/ +def Basis.mapPowExpCharPowOfIsSeparable [Algebra.IsSeparable F E] + (b : Basis ι F E) : Basis ι F E := + Basis.mk (b.linearIndependent.map_pow_expChar_pow_of_isSeparable q n) + (Field.span_map_pow_expChar_pow_eq_top_of_isSeparable q n b.span_eq).ge + +end + +theorem perfectField_of_perfectClosure_eq_bot [h : PerfectField E] (eq : perfectClosure F E = ⊥) : + PerfectField F := by + let p := ringExpChar F + haveI := expChar_of_injective_algebraMap (algebraMap F E).injective p + haveI := PerfectRing.ofSurjective F p fun x ↦ by + obtain ⟨y, h⟩ := surjective_frobenius E p (algebraMap F E x) + have : y ∈ perfectClosure F E := ⟨1, x, by rw [← h, pow_one, frobenius_def, ringExpChar.eq F p]⟩ + obtain ⟨z, rfl⟩ := eq ▸ this + exact ⟨z, (algebraMap F E).injective (by erw [RingHom.map_frobenius, h])⟩ + exact PerfectRing.toPerfectField F p + +/-- If `E / F` is a separable extension, `E` is perfect, then `F` is also prefect. -/ +theorem perfectField_of_isSeparable_of_perfectField_top [Algebra.IsSeparable F E] [PerfectField E] : + PerfectField F := + perfectField_of_perfectClosure_eq_bot F E (perfectClosure.eq_bot_of_isSeparable F E) + +/-- If `E` is an algebraic closure of `F`, then `F` is perfect if and only if `E / F` is +separable. -/ +theorem perfectField_iff_isSeparable_algebraicClosure [IsAlgClosure F E] : + PerfectField F ↔ Algebra.IsSeparable F E := + ⟨fun _ ↦ IsSepClosure.separable, fun _ ↦ haveI : IsAlgClosed E := IsAlgClosure.isAlgClosed F + perfectField_of_isSeparable_of_perfectField_top F E⟩ diff --git a/Mathlib/FieldTheory/PurelyInseparable/Tower.lean b/Mathlib/FieldTheory/PurelyInseparable/Tower.lean new file mode 100644 index 00000000000000..715d67ee39e47f --- /dev/null +++ b/Mathlib/FieldTheory/PurelyInseparable/Tower.lean @@ -0,0 +1,242 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.FieldTheory.PurelyInseparable.PerfectClosure + +/-! + +# Tower law for purely inseparable extensions + +This file contains results related to `Field.finIsepDegree` and the tower law. + +## Main results + +- `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`: the separable degrees satisfy the + tower law: $[E:F]_s [K:E]_s = [K:F]_s$. + +- `IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable`, + `IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable'`: + if `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then + for any subset `S` of `K` such that `F(S) / F` is algebraic, the `E(S) / E` and `F(S) / F` have + the same separable degree. In particular, if `S` is an intermediate field of `K / F` such that + `S / F` is algebraic, the `E(S) / E` and `S / F` have the same separable degree. + +- `minpoly.map_eq_of_isSeparable_of_isPurelyInseparable`: if `K / E / F` is a field extension tower, + such that `E / F` is purely inseparable, then for any element `x` of `K` separable over `F`, + it has the same minimal polynomials over `F` and over `E`. + +- `Polynomial.Separable.map_irreducible_of_isPurelyInseparable`: if `E / F` is purely inseparable, + `f` is a separable irreducible polynomial over `F`, then it is also irreducible over `E`. + +## Tags + +separable degree, degree, separable closure, purely inseparable + +## TODO + +- Restate some intermediate result in terms of linearly disjointness. + +- Prove that the inseparable degrees satisfy the tower law: $[E:F]_i [K:E]_i = [K:F]_i$. + Probably an argument using linearly disjointness is needed. + +-/ + +open Polynomial IntermediateField Field + +noncomputable section + +universe u v w + +section TowerLaw + +variable (F : Type u) (E : Type v) [Field F] [Field E] [Algebra F E] +variable (K : Type w) [Field K] [Algebra F K] + +variable [Algebra E K] [IsScalarTower F E K] + +variable {F K} in +/-- If `K / E / F` is a field extension tower such that `E / F` is purely inseparable, +if `{ u_i }` is a family of separable elements of `K` which is `F`-linearly independent, +then it is also `E`-linearly independent. -/ +theorem LinearIndependent.map_of_isPurelyInseparable_of_isSeparable [IsPurelyInseparable F E] + {ι : Type*} {v : ι → K} (hsep : ∀ i : ι, IsSeparable F (v i)) + (h : LinearIndependent F v) : LinearIndependent E v := by + obtain ⟨q, _⟩ := ExpChar.exists F + haveI := expChar_of_injective_algebraMap (algebraMap F K).injective q + refine linearIndependent_iff.mpr fun l hl ↦ Finsupp.ext fun i ↦ ?_ + choose f hf using fun i ↦ (isPurelyInseparable_iff_pow_mem F q).1 ‹_› (l i) + let n := l.support.sup f + have := (expChar_pow_pos F q n).ne' + replace hf (i : ι) : l i ^ q ^ n ∈ (algebraMap F E).range := by + by_cases hs : i ∈ l.support + · convert pow_mem (hf i) (q ^ (n - f i)) using 1 + rw [← pow_mul, ← pow_add, Nat.add_sub_of_le (Finset.le_sup hs)] + exact ⟨0, by rw [map_zero, Finsupp.not_mem_support_iff.1 hs, zero_pow this]⟩ + choose lF hlF using hf + let lF₀ := Finsupp.onFinset l.support lF fun i ↦ by + contrapose! + refine fun hs ↦ (injective_iff_map_eq_zero _).mp (algebraMap F E).injective _ ?_ + rw [hlF, Finsupp.not_mem_support_iff.1 hs, zero_pow this] + replace h := linearIndependent_iff.1 (h.map_pow_expChar_pow_of_isSeparable' q n hsep) lF₀ <| by + replace hl := congr($hl ^ q ^ n) + rw [Finsupp.linearCombination_apply, Finsupp.sum, sum_pow_char_pow, zero_pow this] at hl + rw [← hl, Finsupp.linearCombination_apply, + Finsupp.onFinset_sum _ (fun _ ↦ by exact zero_smul _ _)] + refine Finset.sum_congr rfl fun i _ ↦ ?_ + simp_rw [Algebra.smul_def, mul_pow, IsScalarTower.algebraMap_apply F E K, hlF, map_pow] + refine pow_eq_zero ((hlF _).symm.trans ?_) + convert map_zero (algebraMap F E) + exact congr($h i) + +namespace Field + +/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable and `K / E` +is separable, then the separable degree of `K / F` is equal to the degree of `K / E`. +It is a special case of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`, and is an +intermediate result used to prove it. -/ +lemma sepDegree_eq_of_isPurelyInseparable_of_isSeparable + [IsPurelyInseparable F E] [Algebra.IsSeparable E K] : sepDegree F K = Module.rank E K := by + let S := separableClosure F K + have h := S.adjoin_rank_le_of_isAlgebraic_right E + rw [separableClosure.adjoin_eq_of_isAlgebraic_of_isSeparable K, rank_top'] at h + obtain ⟨ι, ⟨b⟩⟩ := Basis.exists_basis F S + exact h.antisymm' (b.mk_eq_rank'' ▸ (b.linearIndependent.map' S.val.toLinearMap + (LinearMap.ker_eq_bot_of_injective S.val.injective) + |>.map_of_isPurelyInseparable_of_isSeparable E (fun i ↦ + by simpa only [IsSeparable, minpoly_eq] using Algebra.IsSeparable.isSeparable F (b i)) + |>.cardinal_le_rank)) + +/-- If `K / E / F` is a field extension tower, such that `E / F` is separable, +then $[E:F] [K:E]_s = [K:F]_s$. +It is a special case of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`, and is an +intermediate result used to prove it. -/ +lemma lift_rank_mul_lift_sepDegree_of_isSeparable [Algebra.IsSeparable F E] : + Cardinal.lift.{w} (Module.rank F E) * Cardinal.lift.{v} (sepDegree E K) = + Cardinal.lift.{v} (sepDegree F K) := by + rw [sepDegree, sepDegree, separableClosure.eq_restrictScalars_of_isSeparable F E K] + exact lift_rank_mul_lift_rank F E (separableClosure E K) + +/-- The same-universe version of `Field.lift_rank_mul_lift_sepDegree_of_isSeparable`. -/ +lemma rank_mul_sepDegree_of_isSeparable (K : Type v) [Field K] [Algebra F K] + [Algebra E K] [IsScalarTower F E K] [Algebra.IsSeparable F E] : + Module.rank F E * sepDegree E K = sepDegree F K := by + simpa only [Cardinal.lift_id] using lift_rank_mul_lift_sepDegree_of_isSeparable F E K + +/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, +then $[K:F]_s = [K:E]_s$. +It is a special case of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`, and is an +intermediate result used to prove it. -/ +lemma sepDegree_eq_of_isPurelyInseparable [IsPurelyInseparable F E] : + sepDegree F K = sepDegree E K := by + convert sepDegree_eq_of_isPurelyInseparable_of_isSeparable F E (separableClosure E K) + haveI : IsScalarTower F (separableClosure E K) K := IsScalarTower.of_algebraMap_eq (congrFun rfl) + rw [sepDegree, ← separableClosure.map_eq_of_separableClosure_eq_bot F + (separableClosure.separableClosure_eq_bot E K)] + exact (separableClosure F (separableClosure E K)).equivMap + (IsScalarTower.toAlgHom F (separableClosure E K) K) |>.symm.toLinearEquiv.rank_eq + +/-- If `K / E / F` is a field extension tower, such that `E / F` is algebraic, then their +separable degrees satisfy the tower law: $[E:F]_s [K:E]_s = [K:F]_s$. -/ +theorem lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic [Algebra.IsAlgebraic F E] : + Cardinal.lift.{w} (sepDegree F E) * Cardinal.lift.{v} (sepDegree E K) = + Cardinal.lift.{v} (sepDegree F K) := by + have h := lift_rank_mul_lift_sepDegree_of_isSeparable F (separableClosure F E) K + haveI := separableClosure.isPurelyInseparable F E + rwa [sepDegree_eq_of_isPurelyInseparable (separableClosure F E) E K] at h + +/-- The same-universe version of `Field.lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic`. -/ +@[stacks 09HK "Part 1"] +theorem sepDegree_mul_sepDegree_of_isAlgebraic (K : Type v) [Field K] [Algebra F K] + [Algebra E K] [IsScalarTower F E K] [Algebra.IsAlgebraic F E] : + sepDegree F E * sepDegree E K = sepDegree F K := by + simpa only [Cardinal.lift_id] using lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic F E K + +end Field + +variable {F K} in +/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then +for any subset `S` of `K` such that `F(S) / F` is algebraic, the `E(S) / E` and `F(S) / F` have +the same separable degree. -/ +theorem IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable + (S : Set K) [Algebra.IsAlgebraic F (adjoin F S)] [IsPurelyInseparable F E] : + sepDegree E (adjoin E S) = sepDegree F (adjoin F S) := by + set M := adjoin F S + set L := adjoin E S + let E' := (IsScalarTower.toAlgHom F E K).fieldRange + let j : E ≃ₐ[F] E' := AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K) + have hi : M ≤ L.restrictScalars F := by + rw [restrictScalars_adjoin_of_algEquiv (E := K) j rfl, restrictScalars_adjoin] + exact adjoin.mono _ _ _ Set.subset_union_right + let i : M →+* L := Subsemiring.inclusion hi + letI : Algebra M L := i.toAlgebra + letI : SMul M L := Algebra.toSMul + haveI : IsScalarTower F M L := IsScalarTower.of_algebraMap_eq (congrFun rfl) + haveI : IsPurelyInseparable M L := by + change IsPurelyInseparable M (extendScalars hi) + obtain ⟨q, _⟩ := ExpChar.exists F + have : extendScalars hi = adjoin M (E' : Set K) := restrictScalars_injective F <| by + conv_lhs => rw [extendScalars_restrictScalars, restrictScalars_adjoin_of_algEquiv + (E := K) j rfl, ← adjoin_self F E', adjoin_adjoin_comm] + rw [this, isPurelyInseparable_adjoin_iff_pow_mem _ _ q] + rintro x ⟨y, hy⟩ + obtain ⟨n, z, hz⟩ := IsPurelyInseparable.pow_mem F q y + refine ⟨n, algebraMap F M z, ?_⟩ + rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply F E K, hz, ← hy, map_pow, + AlgHom.toRingHom_eq_coe, IsScalarTower.coe_toAlgHom] + have h := lift_sepDegree_mul_lift_sepDegree_of_isAlgebraic F E L + rw [IsPurelyInseparable.sepDegree_eq_one F E, Cardinal.lift_one, one_mul] at h + rw [Cardinal.lift_injective h, ← sepDegree_mul_sepDegree_of_isAlgebraic F M L, + IsPurelyInseparable.sepDegree_eq_one M L, mul_one] + +variable {F K} in +/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then +for any intermediate field `S` of `K / F` such that `S / F` is algebraic, the `E(S) / E` and +`S / F` have the same separable degree. -/ +theorem IntermediateField.sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable' + (S : IntermediateField F K) [Algebra.IsAlgebraic F S] [IsPurelyInseparable F E] : + sepDegree E (adjoin E (S : Set K)) = sepDegree F S := by + have : Algebra.IsAlgebraic F (adjoin F (S : Set K)) := by rwa [adjoin_self] + have := sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable (F := F) E (S : Set K) + rwa [adjoin_self] at this + +variable {F K} in +/-- If `K / E / F` is a field extension tower, such that `E / F` is purely inseparable, then +for any element `x` of `K` separable over `F`, it has the same minimal polynomials over `F` and +over `E`. -/ +theorem minpoly.map_eq_of_isSeparable_of_isPurelyInseparable (x : K) + (hsep : IsSeparable F x) [IsPurelyInseparable F E] : + (minpoly F x).map (algebraMap F E) = minpoly E x := by + have hi := IsSeparable.isIntegral hsep + have hi' : IsIntegral E x := IsIntegral.tower_top hi + refine eq_of_monic_of_dvd_of_natDegree_le (monic hi') ((monic hi).map (algebraMap F E)) + (dvd_map_of_isScalarTower F E x) (le_of_eq ?_) + have hsep' := IsSeparable.tower_top E hsep + haveI := (isSeparable_adjoin_simple_iff_isSeparable _ _).2 hsep + haveI := (isSeparable_adjoin_simple_iff_isSeparable _ _).2 hsep' + have := Algebra.IsSeparable.isAlgebraic F F⟮x⟯ + have := Algebra.IsSeparable.isAlgebraic E E⟮x⟯ + rw [Polynomial.natDegree_map, ← adjoin.finrank hi, ← adjoin.finrank hi', + ← finSepDegree_eq_finrank_of_isSeparable F _, ← finSepDegree_eq_finrank_of_isSeparable E _, + finSepDegree_eq, finSepDegree_eq, + sepDegree_adjoin_eq_of_isAlgebraic_of_isPurelyInseparable (F := F) E] + +variable {F} in +/-- If `E / F` is a purely inseparable field extension, `f` is a separable irreducible polynomial +over `F`, then it is also irreducible over `E`. -/ +theorem Polynomial.Separable.map_irreducible_of_isPurelyInseparable {f : F[X]} (hsep : f.Separable) + (hirr : Irreducible f) [IsPurelyInseparable F E] : Irreducible (f.map (algebraMap F E)) := by + let K := AlgebraicClosure E + obtain ⟨x, hx⟩ := IsAlgClosed.exists_aeval_eq_zero K f + (natDegree_pos_iff_degree_pos.1 hirr.natDegree_pos).ne' + have ha : Associated f (minpoly F x) := by + have := isUnit_C.2 (leadingCoeff_ne_zero.2 hirr.ne_zero).isUnit.inv + exact ⟨this.unit, by rw [IsUnit.unit_spec, minpoly.eq_of_irreducible hirr hx]⟩ + have ha' : Associated (f.map (algebraMap F E)) ((minpoly F x).map (algebraMap F E)) := + ha.map (mapRingHom (algebraMap F E)).toMonoidHom + have heq := minpoly.map_eq_of_isSeparable_of_isPurelyInseparable E x (ha.separable hsep) + rw [ha'.irreducible_iff, heq] + exact minpoly.irreducible (Algebra.IsIntegral.isIntegral x) + +end TowerLaw diff --git a/Mathlib/FieldTheory/RatFunc/Basic.lean b/Mathlib/FieldTheory/RatFunc/Basic.lean index 9a1cde7fbd08bb..badeed639d44e1 100644 --- a/Mathlib/FieldTheory/RatFunc/Basic.lean +++ b/Mathlib/FieldTheory/RatFunc/Basic.lean @@ -663,7 +663,7 @@ instance : IsFractionRing K[X] (RatFunc K) where simp only [← ofFractionRing_algebraMap, Function.comp_apply, ← ofFractionRing_mul, ofFractionRing.injEq] -@[deprecated "Use NoZeroSMulDivisors.algebraMap_eq_zero_iff instead." (since := "2024-09-08")] +@[deprecated "Use FaithfulSMul.algebraMap_eq_zero_iff instead." (since := "2024-09-08")] theorem algebraMap_eq_zero_iff {x : K[X]} : algebraMap K[X] (RatFunc K) x = 0 ↔ x = 0 := by simp diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index c5f523f226b655..d430b51986bf66 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -861,26 +861,24 @@ instance ChartedSpace.sum : ChartedSpace H (M ⊕ M') where chartAt := Sum.elim (fun x ↦ (cm.chartAt x).lift_openEmbedding IsOpenEmbedding.inl) (fun x ↦ (cm'.chartAt x).lift_openEmbedding IsOpenEmbedding.inr) mem_chart_source p := by - by_cases h : Sum.isLeft p - · let x := Sum.getLeft p h - rw [Sum.eq_left_getLeft_of_isLeft h, Sum.elim_inl, lift_openEmbedding_source, + cases p with + | inl x => + rw [Sum.elim_inl, lift_openEmbedding_source, ← PartialHomeomorph.lift_openEmbedding_source _ IsOpenEmbedding.inl] use x, cm.mem_chart_source x - · have h' : Sum.isRight p := Sum.not_isLeft.mp h - let x := Sum.getRight p h' - rw [Sum.eq_right_getRight_of_isRight h', Sum.elim_inr, lift_openEmbedding_source, + | inr x => + rw [Sum.elim_inr, lift_openEmbedding_source, ← PartialHomeomorph.lift_openEmbedding_source _ IsOpenEmbedding.inr] use x, cm'.mem_chart_source x chart_mem_atlas p := by - by_cases h : Sum.isLeft p - · rw [Sum.eq_left_getLeft_of_isLeft h, Sum.elim_inl] + cases p with + | inl x => + rw [Sum.elim_inl] left - let x := Sum.getLeft p h use ChartedSpace.chartAt x, cm.chart_mem_atlas x - · have h' : Sum.isRight p := Sum.not_isLeft.mp h - rw [Sum.eq_right_getRight_of_isRight h', Sum.elim_inr] + | inr x => + rw [Sum.elim_inr] right - let x := Sum.getRight p h' use ChartedSpace.chartAt x, cm'.chart_mem_atlas x lemma ChartedSpace.sum_chartAt_inl (x : M) : diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index 1c23a54eb9e67d..58dad19553fec1 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -82,7 +82,7 @@ theorem UniqueMDiffOn.uniqueMDiffOn_preimage (hs : UniqueMDiffOn I s) {e : Parti (he : e.MDifferentiable I I') : UniqueMDiffOn I' (e.target ∩ e.symm ⁻¹' s) := fun _x hx ↦ e.right_inv hx.1 ▸ (hs _ hx.2).preimage_partialHomeomorph he (e.map_target hx.1) -variable [IsManifold I 1 M] in +variable [IsManifold I 1 M] in /-- If a set in a manifold has the unique derivative property, then its pullback by any extended chart, in the vector space, also has the unique derivative property. -/ theorem UniqueMDiffOn.uniqueMDiffOn_target_inter (hs : UniqueMDiffOn I s) (x : M) : diff --git a/Mathlib/Geometry/Manifold/Sheaf/Basic.lean b/Mathlib/Geometry/Manifold/Sheaf/Basic.lean index 2368052ad3acf2..4df8d6b6d2a5c1 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Basic.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Basic.lean @@ -60,7 +60,7 @@ def StructureGroupoid.LocalInvariantProp.localPredicate (hG : LocalInvariantProp apply h locality := by intro V f h x - obtain ⟨U, hxU, i, hU : ChartedSpace.LiftProp P (f ∘ i)⟩ := h x + obtain ⟨U, hxU, i, hU : ChartedSpace.LiftProp P (f ∘ _)⟩ := h x let x' : U := ⟨x, hxU⟩ have hUV : U ≤ V := CategoryTheory.leOfHom i have : ChartedSpace.LiftPropAt P f (Opens.inclusion hUV x') := by diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index 36647bd9a55ca2..537ff7ff313b80 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -273,7 +273,7 @@ instance : EmptyCollection LocallyRingedSpace.{u} := ⟨LocallyRingedSpace.empty /-- The canonical map from the empty locally ringed space. -/ def emptyTo (X : LocallyRingedSpace.{u}) : ∅ ⟶ X := - ⟨⟨⟨fun x => PEmpty.elim x, by fun_prop⟩, + ⟨⟨ofHom ⟨fun x => PEmpty.elim x, by fun_prop⟩, { app := fun U => CommRingCat.ofHom <| by refine ⟨⟨⟨0, ?_⟩, ?_⟩, ?_, ?_⟩ <;> intros <;> rfl }⟩, fun x => PEmpty.elim x⟩ @@ -354,12 +354,12 @@ lemma stalkMap_comp (x : X) : @[reassoc] lemma stalkSpecializes_stalkMap (x x' : X) (h : x ⤳ x') : - Y.presheaf.stalkSpecializes (f.base.map_specializes h) ≫ f.stalkMap x = + Y.presheaf.stalkSpecializes (f.base.hom.map_specializes h) ≫ f.stalkMap x = f.stalkMap x' ≫ X.presheaf.stalkSpecializes h := PresheafedSpace.stalkMap.stalkSpecializes_stalkMap f.toShHom h lemma stalkSpecializes_stalkMap_apply (x x' : X) (h : x ⤳ x') (y) : - f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.map_specializes h) y) = + f.stalkMap x (Y.presheaf.stalkSpecializes (f.base.hom.map_specializes h) y) = (X.presheaf.stalkSpecializes h (f.stalkMap x' y)) := DFunLike.congr_fun (CommRingCat.hom_ext_iff.mp (stalkSpecializes_stalkMap f x x' h)) y diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index a4c26869dcdfeb..154dc93a8b6e6a 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -189,7 +189,7 @@ theorem imageBasicOpen_image_preimage : (g.base : X.carrier.1 ⟶ Y.carrier.1) · ext simp_rw [types_comp_apply, ← TopCat.comp_app, ← PresheafedSpace.comp_base] - congr 2 + congr 3 exact coequalizer.condition f.toShHom g.toShHom · apply isColimitCoforkMapOfIsColimit (forget TopCat) apply isColimitCoforkMapOfIsColimit (SheafedSpace.forget _) diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index aa63a2c5e4a93d..8d53322f85c513 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -311,12 +311,7 @@ def pullbackConeOfLeftFst : erw [TopCat.pullbackIsoProdSubtype_inv_fst_apply] · erw [TopCat.pullbackIsoProdSubtype_inv_snd_apply] · rintro _ ⟨x, h₁, rfl⟩ - -- next line used to be - -- `exact ⟨_, h₁, ConcreteCategory.congr_hom pullback.condition x⟩))` - -- before https://github.com/leanprover-community/mathlib4/pull/13170 - refine ⟨_, h₁, ?_⟩ - change (_ ≫ f.base) _ = (_ ≫ g.base) _ - rw [pullback.condition])) + exact ⟨_, h₁, CategoryTheory.congr_fun pullback.condition x⟩)) naturality := by intro U V i induction U using Opposite.rec' @@ -855,7 +850,7 @@ theorem sigma_ι_isOpenEmbedding : IsOpenEmbedding (colimit.ι F i).base := by -- 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.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_comp_isIso, + rw [TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_comp_isIso, TopCat.isOpenEmbedding_iff_isIso_comp] exact .sigmaMk @@ -1126,10 +1121,9 @@ 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, Set.range_comp, + rw [LocallyRingedSpace.comp_base, ← this, ← Category.assoc, TopCat.coe_comp, Set.range_comp, Set.range_eq_univ.mpr, Set.image_univ] - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` on this lemma - · erw [TopCat.pullback_fst_range] + · rw [TopCat.pullback_fst_range] ext constructor · rintro ⟨y, eq⟩; exact ⟨y, eq.symm⟩ diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index 2edeb618bd4222..818fe3e6d019b4 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -3,8 +3,9 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.Topology.Sheaves.Presheaf import Mathlib.CategoryTheory.Adjunction.FullyFaithful +import Mathlib.CategoryTheory.Elementwise +import Mathlib.Topology.Sheaves.Presheaf /-! # Presheafed spaces @@ -52,22 +53,12 @@ variable {C} namespace PresheafedSpace --- Porting note: using `Coe` here triggers an error, `CoeOut` seems an acceptable alternative instance coeCarrier : CoeOut (PresheafedSpace C) TopCat where coe X := X.carrier attribute [coe] PresheafedSpace.carrier --- Porting note: we add this instance, as Lean does not reliably use the `CoeOut` instance above --- in downstream files. instance : CoeSort (PresheafedSpace C) Type* where coe X := X.carrier --- Porting note: the following lemma is removed because it is a syntactic tauto -/-@[simp] -theorem as_coe (X : PresheafedSpace.{w, v, u} C) : X.carrier = (X : TopCat.{w}) := - rfl-/ - --- Porting note: removed @[simp] as the `simpVarHead` linter complains --- @[simp] theorem mk_coe (carrier) (presheaf) : (({ carrier presheaf } : PresheafedSpace C) : TopCat) = carrier := @@ -91,10 +82,6 @@ structure Hom (X Y : PresheafedSpace C) where base : (X : TopCat) ⟶ (Y : TopCat) c : Y.presheaf ⟶ base _* X.presheaf --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): eventually, the `ext` lemma shall be applied to terms in `X ⟶ Y` --- rather than `Hom X Y`, this one was renamed `Hom.ext` instead of `ext`, --- and the more practical lemma `ext` is defined just after the definition --- of the `Category` instance @[ext (iff := false)] theorem Hom.ext {X Y : PresheafedSpace C} (α β : Hom X Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := by @@ -114,7 +101,6 @@ theorem hext {X Y : PresheafedSpace C} (α β : Hom X Y) (w : α.base = β.base) cases β congr --- Porting note: `eqToHom` is no longer necessary in the definition of `c` /-- The identity morphism of a `PresheafedSpace`. -/ def id (X : PresheafedSpace C) : Hom X X where base := 𝟙 (X : TopCat) @@ -138,23 +124,12 @@ section attribute [local simp] id comp --- Porting note: in mathlib3, `tidy` could (almost) prove the category axioms, but proofs --- were included because `tidy` was slow. Here, `aesop_cat` succeeds reasonably quickly --- for `comp_id` and `assoc` /-- The category of PresheafedSpaces. Morphisms are pairs, a continuous map and a presheaf map from the presheaf on the target to the pushforward of the presheaf on the source. -/ instance categoryOfPresheafedSpaces : Category (PresheafedSpace C) where Hom := Hom id := id comp := comp - id_comp _ := by - dsimp - ext - · dsimp - simp - · dsimp - simp only [map_id, whiskerRight_id', assoc] - erw [comp_id, comp_id] variable {C} @@ -176,7 +151,6 @@ attribute [local simp] eqToHom_map theorem id_base (X : PresheafedSpace C) : (𝟙 X : X ⟶ X).base = 𝟙 (X : TopCat) := rfl --- Porting note: `eqToHom` is no longer needed in the statements of `id_c` and `id_c_app` theorem id_c (X : PresheafedSpace C) : (𝟙 X : X ⟶ X).c = 𝟙 X.presheaf := rfl @@ -195,9 +169,10 @@ theorem comp_base {X Y Z : PresheafedSpace C} (f : X ⟶ Y) (g : Y ⟶ Z) : instance (X Y : PresheafedSpace C) : CoeFun (X ⟶ Y) fun _ => (↑X → ↑Y) := ⟨fun f => f.base⟩ --- Porting note: removed as this is a syntactic tauto ---theorem coe_to_fun_eq {X Y : PresheafedSpace.{v, v, u} C} (f : X ⟶ Y) : (f : ↑X → ↑Y) = f.base := --- rfl +/-! +Note that we don't include a `ConcreteCategory` instance, since equality of morphisms `X ⟶ Y` +does not follow from equality of their coercions `X → Y`. +-/ -- The `reassoc` attribute was added despite the LHS not being a composition of two homs, -- for the reasons explained in the docstring. @@ -247,7 +222,7 @@ def isoOfComponents (H : X.1 ≅ Y.1) (α : H.hom _* X.2 ≅ Y.2) : X ≅ Y wher inv_hom_id := by ext · dsimp - rw [H.inv_hom_id] + exact H.inv_hom_id_apply _ dsimp simp only [Presheaf.toPushforwardOfIso_app, assoc, ← α.hom.naturality] simp only [eqToHom_map, eqToHom_app, eqToHom_trans_assoc, eqToHom_refl, id_comp] diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 912cb224dd41a2..0a1ef6d97b543a 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -120,9 +120,7 @@ abbrev toTopGlueData : TopCat.GlueData := toGlueData := 𝖣.mapGlueData (forget C) } 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] + rw [← show _ = (𝖣.ι i).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) _, TopCat.coe_comp] exact (TopCat.homeoOfIso (𝖣.gluedIso (PresheafedSpace.forget _)).symm).isOpenEmbedding.comp (D.toTopGlueData.ι_isOpenEmbedding i) @@ -133,14 +131,9 @@ 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 have eq₁ : _ = (π₁ i, j, k).base := PreservesPullback.iso_hom_fst (forget C) _ _ have eq₂ : _ = (π₂ i, j, k).base := PreservesPullback.iso_hom_snd (forget C) _ _ - rw [← eq₁, ← eq₂] - -- Porting note: `rw` to `erw` on `coe_comp` - erw [coe_comp] - rw [Set.image_comp] - -- Porting note: `rw` to `erw` on `coe_comp` - erw [coe_comp] - rw [Set.preimage_comp, Set.image_preimage_eq, TopCat.pullback_snd_image_fst_preimage] - · rfl + rw [← eq₁, ← eq₂, TopCat.coe_comp, Set.image_comp, TopCat.coe_comp, Set.preimage_comp, + Set.image_preimage_eq] + · simp only [forget_obj, forget_map, TopCat.pullback_snd_image_fst_preimage] rw [← TopCat.epi_iff_surjective] infer_instance @@ -186,22 +179,22 @@ theorem snd_invApp_t_app' (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k) have := (𝖣.t_fac k i j).symm rw [← IsIso.inv_comp_eq] at this replace this := (congr_arg ((PresheafedSpace.Hom.base ·)) this).symm + replace this := congr_arg (TopCat.Hom.hom ·) this replace this := congr_arg (ContinuousMap.toFun ·) this dsimp at this - rw [coe_comp, coe_comp] at this rw [this, Set.image_comp, Set.image_comp, Set.preimage_image_eq] swap · refine Function.HasLeftInverse.injective ⟨(D.t i k).base, fun x => ?_⟩ - rw [← CategoryTheory.comp_apply, ← comp_base, D.t_inv, id_base, CategoryTheory.id_apply] + rw [← ConcreteCategory.comp_apply, ← comp_base, D.t_inv, id_base, ConcreteCategory.id_apply] refine congr_arg (_ '' ·) ?_ refine congr_fun ?_ _ refine Set.image_eq_preimage_of_inverse ?_ ?_ · intro x - rw [← CategoryTheory.comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, - CategoryTheory.id_apply] + rw [← ConcreteCategory.comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, + ConcreteCategory.id_apply] · intro x - rw [← CategoryTheory.comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, - CategoryTheory.id_apply] + rw [← ConcreteCategory.comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, + ConcreteCategory.id_apply] · rw [← IsIso.eq_inv_comp, IsOpenImmersion.inv_invApp, Category.assoc, (D.t' k i j).c.naturality_assoc] simp_rw [← Category.assoc] @@ -219,7 +212,7 @@ theorem snd_invApp_t_app' (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k) subst eq use (inv (D.t' k i j)).base y change (inv (D.t' k i j) ≫ π₁ k, i, j).base y = _ - congr 2 + congr 3 rw [IsIso.inv_comp_eq, 𝖣.t_fac_assoc, 𝖣.t_inv, Category.comp_id] /-- The red and the blue arrows in ![this diagram](https://i.imgur.com/q6X1GJ9.png) commute. -/ @@ -245,7 +238,7 @@ theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : rw [← show _ = (𝖣.ι i).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) i, ← show _ = (𝖣.ι j).base from 𝖣.ι_gluedIso_inv (PresheafedSpace.forget _) j] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` on `coe_comp` - erw [coe_comp, coe_comp, coe_comp] + erw [TopCat.coe_comp, TopCat.coe_comp, TopCat.coe_comp] rw [Set.image_comp, Set.preimage_comp] erw [Set.preimage_image_eq] · refine Eq.trans (D.toTopGlueData.preimage_image_eq_image' _ _ _) ?_ @@ -258,7 +251,7 @@ theorem ι_image_preimage_eq (i j : D.J) (U : Opens (D.U i).carrier) : change (D.t i j ≫ D.t j i).base '' _ = _ rw [𝖣.t_inv] simp - · rw [← coe_comp, ← TopCat.mono_iff_injective] + · rw [← TopCat.coe_comp, ← TopCat.mono_iff_injective] infer_instance /-- (Implementation). The map `Γ(𝒪_{U_i}, U) ⟶ Γ(𝒪_{U_j}, 𝖣.ι j ⁻¹' (𝖣.ι i '' U))` -/ @@ -409,7 +402,7 @@ theorem ιInvApp_π {i : D.J} (U : Opens (D.U i).carrier) : · exact h2.symm · have := D.ι_gluedIso_inv (PresheafedSpace.forget _) i dsimp at this - rw [← this, coe_comp] + rw [← this, TopCat.coe_comp] refine Function.Injective.comp ?_ (TopCat.GlueData.ι_injective D.toTopGlueData i) rw [← TopCat.mono_iff_injective] infer_instance @@ -515,8 +508,8 @@ def vPullbackConeIsLimit (i j : D.J) : IsLimit (𝖣.vPullbackCone i j) := exact this _ rw [← Set.image_subset_iff, ← Set.image_univ, ← Set.image_comp, Set.image_univ] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` - erw [← coe_comp] - rw [this, coe_comp, ← Set.image_univ, Set.image_comp] + erw [← TopCat.coe_comp] + rw [this, TopCat.coe_comp, ← Set.image_univ, Set.image_comp] exact Set.image_subset_range _ _ · apply IsOpenImmersion.lift_fac · rw [← cancel_mono (𝖣.ι j), Category.assoc, ← (𝖣.vPullbackCone i j).condition] diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index 1f4e3891132018..1d1a7084377028 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -248,7 +248,9 @@ def colimitCoconeIsColimit (F : J ⥤ PresheafedSpace.{_, _, v} C) : m.base = colimit.desc (F ⋙ PresheafedSpace.forget C) ((PresheafedSpace.forget C).mapCocone s) := by dsimp - ext j + -- `colimit.hom_ext` used to be automatically applied by `ext` before https://github.com/leanprover-community/mathlib4/pull/21302 + apply colimit.hom_ext fun j => ?_ + ext rw [colimit.ι_desc, mapCocone_ι_app, ← w j] simp ext : 1 diff --git a/Mathlib/Geometry/RingedSpace/Stalks.lean b/Mathlib/Geometry/RingedSpace/Stalks.lean index 9574828c1ed897..828a1f82add078 100644 --- a/Mathlib/Geometry/RingedSpace/Stalks.lean +++ b/Mathlib/Geometry/RingedSpace/Stalks.lean @@ -185,7 +185,7 @@ def stalkIso {X Y : PresheafedSpace.{_, _, v} C} (α : X ≅ Y) (x : X) : @[reassoc, elementwise, simp, nolint simpNF] theorem stalkSpecializes_stalkMap {X Y : PresheafedSpace.{_, _, v} C} (f : X ⟶ Y) {x y : X} (h : x ⤳ y) : - Y.presheaf.stalkSpecializes (f.base.map_specializes h) ≫ f.stalkMap x = + Y.presheaf.stalkSpecializes (f.base.hom.map_specializes h) ≫ f.stalkMap x = f.stalkMap y ≫ X.presheaf.stalkSpecializes h := by -- Porting note: the original one liner `dsimp [stalkMap]; simp [stalkMap]` doesn't work, -- I had to uglify this diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index a609b7d46308dc..67fdde3821c7cc 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -697,6 +697,11 @@ theorem sum_card_orderOf_eq_card_pow_eq_one [Fintype G] [DecidableEq G] (hn : n theorem orderOf_le_card_univ [Fintype G] : orderOf x ≤ Fintype.card G := Finset.le_card_of_inj_on_range (x ^ ·) (fun _ _ ↦ Finset.mem_univ _) pow_injOn_Iio_orderOf +@[to_additive] +theorem orderOf_le_card [Finite G] : orderOf x ≤ Nat.card G := by + obtain ⟨⟩ := nonempty_fintype G + simpa using orderOf_le_card_univ + end FiniteMonoid section FiniteCancelMonoid diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 028fb46fc8ecba..d548b5dc7293de 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -113,6 +113,14 @@ lemma isCyclic_iff_exists_orderOf_eq_natCard [Finite α] : IsCyclic α ↔ ∃ g : α, orderOf g = Nat.card α := by simp_rw [isCyclic_iff_exists_zpowers_eq_top, ← card_eq_iff_eq_top, Nat.card_zpowers] +@[to_additive] +lemma isCyclic_iff_exists_natCard_le_orderOf [Finite α] : + IsCyclic α ↔ ∃ g : α, Nat.card α ≤ orderOf g := by + rw [isCyclic_iff_exists_orderOf_eq_natCard] + apply exists_congr + intro g + exact ⟨Eq.ge, le_antisymm orderOf_le_card⟩ + @[deprecated (since := "2024-12-20")] alias isCyclic_iff_exists_ofOrder_eq_natCard := isCyclic_iff_exists_orderOf_eq_natCard @@ -132,6 +140,11 @@ theorem isCyclic_of_orderOf_eq_card [Finite α] (x : α) (hx : orderOf x = Nat.c IsCyclic α := isCyclic_iff_exists_orderOf_eq_natCard.mpr ⟨x, hx⟩ +@[to_additive] +theorem isCyclic_of_card_le_orderOf [Finite α] (x : α) (hx : Nat.card α ≤ orderOf x) : + IsCyclic α := + isCyclic_iff_exists_natCard_le_orderOf.mpr ⟨x, hx⟩ + @[to_additive] theorem Subgroup.eq_bot_or_eq_top_of_prime_card (H : Subgroup G) [hp : Fact (Nat.card G).Prime] : H = ⊥ ∨ H = ⊤ := by diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index f3eee53f08930b..8556b9f1edbc71 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -90,6 +90,9 @@ theorem vsub_mem_vectorSpan {s : Set P} {p1 p2 : P} (hp1 : p1 ∈ s) (hp2 : p2 p1 -ᵥ p2 ∈ vectorSpan k s := vsub_set_subset_vectorSpan k s (vsub_mem_vsub hp1 hp2) +@[simp] lemma vectorSpan_vadd (s : Set P) (v : V) : vectorSpan k (v +ᵥ s) = vectorSpan k s := by + simp [vectorSpan] + /-- The points in the affine span of a (possibly empty) set of points. Use `affineSpan` instead to get an `AffineSubspace k P`. -/ def spanPoints (s : Set P) : Set P := diff --git a/Mathlib/LinearAlgebra/Basis/Exact.lean b/Mathlib/LinearAlgebra/Basis/Exact.lean index 88e09f3929bbd7..896513a905aca5 100644 --- a/Mathlib/LinearAlgebra/Basis/Exact.lean +++ b/Mathlib/LinearAlgebra/Basis/Exact.lean @@ -57,7 +57,7 @@ private lemma top_le_span_of_aux (v : κ ⊕ σ → M) replace hs := DFunLike.congr_fun hs (s m) simp only [LinearMap.coe_comp, Function.comp_apply, LinearMap.id_coe, id_eq] at hs simp [hs] - have : m ∈ Submodule.span R (Set.range v) := hsp trivial + have : m ∈ Submodule.span R (Set.range v) := hsp Submodule.mem_top obtain ⟨c, rfl⟩ := Finsupp.mem_span_range_iff_exists_finsupp.mp this simp only [LinearMap.mem_ker, Finsupp.sum, map_sum, map_smul, Finset.sum_sum_eq_sum_toLeft_add_sum_toRight, map_add, hslzero, smul_zero, diff --git a/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean b/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean index 8fba55a7ccff65..7bfb7c488ce3bc 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean @@ -66,12 +66,12 @@ def dualSubmoduleToDual [NoZeroSMulDivisors R S] (N : Submodule R M) : B.dualSubmodule N →ₗ[R] Module.Dual R N := { toFun := fun x ↦ { toFun := B.dualSubmoduleParing x - map_add' := fun x y ↦ NoZeroSMulDivisors.algebraMap_injective R S (by simp) - map_smul' := fun r m ↦ NoZeroSMulDivisors.algebraMap_injective R S + map_add' := fun x y ↦ FaithfulSMul.algebraMap_injective R S (by simp) + map_smul' := fun r m ↦ FaithfulSMul.algebraMap_injective R S (by simp [← Algebra.smul_def]) } - map_add' := fun x y ↦ LinearMap.ext fun z ↦ NoZeroSMulDivisors.algebraMap_injective R S + map_add' := fun x y ↦ LinearMap.ext fun z ↦ FaithfulSMul.algebraMap_injective R S (by simp) - map_smul' := fun r x ↦ LinearMap.ext fun y ↦ NoZeroSMulDivisors.algebraMap_injective R S + map_smul' := fun r x ↦ LinearMap.ext fun y ↦ FaithfulSMul.algebraMap_injective R S (by simp [← Algebra.smul_def]) } lemma dualSubmoduleToDual_injective (hB : B.Nondegenerate) [NoZeroSMulDivisors R S] diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 295a03c7af209f..fe66eb87c9846b 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -716,7 +716,8 @@ instance instFiniteDimensionalOfIsReflexive (K V : Type*) rw [lift_id'] exact lt_trans h₁ h₂ -instance [IsDomain R] : NoZeroSMulDivisors R M := by +-- see Note [lower instance priority] +instance (priority := 100) [IsDomain R] : NoZeroSMulDivisors R M := by refine (noZeroSMulDivisors_iff R M).mpr ?_ intro r m hrm rw [or_iff_not_imp_left] diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 5aeea273a0985e..c4adfb4fb4c3b6 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -639,7 +639,7 @@ lemma disjoint_genEigenspace [NoZeroSMulDivisors R M] apply mapsTo_genEigenspace_of_comm (Algebra.mul_sub_algebraMap_commutes f _) apply isNilpotent_restrict_genEigenspace_nat have hf₁₂ : f₂ - f₁ = algebraMap R (End R p) (μ₁ - μ₂) := by ext; simp [f₁, f₂, sub_smul] - rw [hf₁₂, IsNilpotent.map_iff (NoZeroSMulDivisors.algebraMap_injective R (End R p)), + rw [hf₁₂, IsNilpotent.map_iff (FaithfulSMul.algebraMap_injective R (End R p)), isNilpotent_iff_eq_zero, sub_eq_zero] at this contradiction diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 1b216c6e08dc5e..cf92c3c6fe93a4 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1300,7 +1300,7 @@ lemma linearIndependent_algHom_toLinearMap' (K M L) [CommRing K] LinearIndependent K (AlgHom.toLinearMap : (M →ₐ[K] L) → M →ₗ[K] L) := by apply (linearIndependent_algHom_toLinearMap K M L).restrict_scalars simp_rw [Algebra.smul_def, mul_one] - exact NoZeroSMulDivisors.algebraMap_injective K L + exact FaithfulSMul.algebraMap_injective K L end Module diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index 5248086346b66e..947720038a3dbc 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -39,7 +39,7 @@ variable {n 𝕜 : Type*} [RCLike 𝕜] [Fintype n] [DecidableEq n] {A : Matrix lemma finite_real_spectrum : (spectrum ℝ A).Finite := by rw [← spectrum.preimage_algebraMap 𝕜] - exact A.finite_spectrum.preimage (NoZeroSMulDivisors.algebraMap_injective ℝ 𝕜).injOn + exact A.finite_spectrum.preimage (FaithfulSMul.algebraMap_injective ℝ 𝕜).injOn instance : Finite (spectrum ℝ A) := A.finite_real_spectrum diff --git a/Mathlib/LinearAlgebra/PerfectPairing/Restrict.lean b/Mathlib/LinearAlgebra/PerfectPairing/Restrict.lean index f9b217399a9775..91153233ec4dc7 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing/Restrict.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing/Restrict.lean @@ -95,7 +95,7 @@ private def restrictScalarsAux (hp : ∀ m n, p (i m) (j n) ∈ (algebraMap S R).range) : M' →ₗ[S] N' →ₗ[S] S := LinearMap.restrictScalarsRange i j (Algebra.linearMap S R) - (NoZeroSMulDivisors.algebraMap_injective S R) p.toLin hp + (FaithfulSMul.algebraMap_injective S R) p.toLin hp private lemma restrictScalarsAux_injective (hi : Injective i) @@ -103,7 +103,7 @@ private lemma restrictScalarsAux_injective (hp : ∀ m n, p (i m) (j n) ∈ (algebraMap S R).range) : Injective (p.restrictScalarsAux i j hp) := by let f := LinearMap.restrictScalarsRange i j (Algebra.linearMap S R) - (NoZeroSMulDivisors.algebraMap_injective S R) p.toLin hp + (FaithfulSMul.algebraMap_injective S R) p.toLin hp rw [← LinearMap.ker_eq_bot] refine (Submodule.eq_bot_iff _).mpr fun x (hx : f x = 0) ↦ ?_ replace hx (n : N) : p (i x) n = 0 := by @@ -128,7 +128,7 @@ private lemma restrictScalarsAux_surjective obtain ⟨m, hm⟩ := h g refine ⟨m, ?_⟩ ext n - apply NoZeroSMulDivisors.algebraMap_injective S R + apply FaithfulSMul.algebraMap_injective S R change Algebra.linearMap S R _ = _ simpa [restrictScalarsAux] using LinearMap.congr_fun hm n @@ -180,7 +180,7 @@ lemma exists_basis_basis_of_span_eq_top_of_mem_algebraMap have hv' : LinearIndependent K v' := by replace hv₃ := hv₃.restrict_scalars (R := K) <| by simp_rw [← Algebra.algebraMap_eq_smul_one] - exact NoZeroSMulDivisors.algebraMap_injective K L + exact FaithfulSMul.algebraMap_injective K L rw [show ((↑) : v → M) = M'.subtype ∘ v' from rfl] at hv₃ exact hv₃.of_comp suffices span K (Set.range v') = ⊤ by @@ -259,7 +259,7 @@ def restrictScalarsField (x : M') (y : N') : algebraMap K L (p.restrictScalarsField i j hi hj hij hp x y) = p (i x) (j y) := LinearMap.restrictScalarsRange_apply i j (Algebra.linearMap K L) - (NoZeroSMulDivisors.algebraMap_injective K L) p.toLin hp x y + (FaithfulSMul.algebraMap_injective K L) p.toLin hp x y end Field diff --git a/Mathlib/LinearAlgebra/Projectivization/Cardinality.lean b/Mathlib/LinearAlgebra/Projectivization/Cardinality.lean index 6c66f723eb964c..e62e38df83ac4b 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Cardinality.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Cardinality.lean @@ -91,11 +91,9 @@ variable (k V : Type*) [Field k] [AddCommGroup V] [Module k V] /-- Cardinality formula for the points of `ℙ k V` if `k` and `V` are finite expressed as a fraction. -/ lemma card'' [Finite k] : Nat.card (ℙ k V) = (Nat.card V - 1) / (Nat.card k - 1) := by - haveI : Fintype k := Fintype.ofFinite k - rw [card k] have : 1 < Nat.card k := Finite.one_lt_card - have h : 0 ≠ (Nat.card k - 1) := by omega - exact Nat.eq_div_of_mul_eq_left (Ne.symm h) rfl + rw [card k, Nat.mul_div_cancel] + omega lemma card_of_finrank [Finite k] {n : ℕ} (h : Module.finrank k V = n) : Nat.card (ℙ k V) = ∑ i ∈ Finset.range n, Nat.card k ^ i := by @@ -115,8 +113,7 @@ lemma card_of_finrank [Finite k] {n : ℕ} (h : Module.finrank k V = n) : let e : V ≃ₗ[k] (Fin n → k) := LinearEquiv.ofFinrankEq _ _ (by simpa) have hc : Nat.card V = Nat.card k ^ n := by simp [Nat.card_congr e.toEquiv, Nat.card_fun] zify - have hn : 1 ≤ Nat.card k := Nat.one_le_of_lt Finite.one_lt_card - conv_rhs => rw [Int.natCast_sub hn, Int.natCast_one, geom_sum_mul] + conv_rhs => rw [Int.natCast_sub this.le, Int.natCast_one, geom_sum_mul] rw [← Int.natCast_mul, ← card k V, hc] simp diff --git a/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean b/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean index 60e98abdd2d75a..759fcc3bb473a0 100644 --- a/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean +++ b/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean @@ -75,7 +75,7 @@ def restrictScalars' : root_coroot_two i := by have : algebraMap K L 2 = 2 := by rw [← Int.cast_two (R := K), ← Int.cast_two (R := L), map_intCast] - exact NoZeroSMulDivisors.algebraMap_injective K L <| by simp [this] + exact FaithfulSMul.algebraMap_injective K L <| by simp [this] reflection_perm := P.reflection_perm reflection_perm_root i j := by ext; simpa [algebra_compatible_smul L] using P.reflection_perm_root i j diff --git a/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean b/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean index 3c23104e4e2bbd..863271814fe03d 100644 --- a/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean +++ b/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean @@ -48,6 +48,6 @@ variable [Nontrivial R] [NoZeroSMulDivisors S R] @[simp] lemma cartanMatrixIn_apply_same (i : b.support) : b.cartanMatrixIn S i i = 2 := - NoZeroSMulDivisors.algebraMap_injective S R <| by simp [cartanMatrixIn_def, map_ofNat] + FaithfulSMul.algebraMap_injective S R <| by simp [cartanMatrixIn_def, map_ofNat] end RootPairing.Base diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index a71bc3defb5517..494bda6e3b380c 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -84,10 +84,7 @@ way. -/ def ofEquiv (α) {β} [Denumerable α] (e : β ≃ α) : Denumerable β := { Encodable.ofEquiv _ e with decode_inv := fun n => by - -- Porting note: replaced `simp` - simp_rw [Option.mem_def, decode_ofEquiv e, encode_ofEquiv e, decode_eq_ofNat, - Option.map_some', Option.some_inj, exists_eq_left', Equiv.apply_symm_apply, - Denumerable.encode_ofNat] } + simp [decode_ofEquiv, encode_ofEquiv] } @[simp] theorem ofEquiv_ofNat (α) {β} [Denumerable α] (e : β ≃ α) (n) : diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index 4e516587e5f5aa..bb90812cec01d2 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -34,8 +34,7 @@ instance {α : Sort u} {β : Sort v} : EmbeddingLike (α ↪ β) α β where initialize_simps_projections Embedding (toFun → apply) --- Porting note: this needs `tactic.lift`. ---instance {α β : Sort*} : CanLift (α → β) (α ↪ β) coeFn Injective where prf f hf := ⟨⟨f, hf⟩, rfl⟩ +instance {α β : Sort*} : CanLift (α → β) (α ↪ β) (↑) Injective where prf f hf := ⟨⟨f, hf⟩, rfl⟩ theorem exists_surjective_iff {α β : Sort*} : (∃ f : α → β, Surjective f) ↔ Nonempty (α → β) ∧ Nonempty (β ↪ α) := @@ -197,9 +196,6 @@ lemma setValue_right_apply_eq {α β} (f : α ↪ β) (a c : α) [∀ a', Decida protected def some {α} : α ↪ Option α := ⟨some, Option.some_injective α⟩ --- Porting note: Lean 4 unfolds coercion `α → Option α` to `some`, so there is no separate --- `Function.Embedding.coeOption`. - /-- A version of `Option.map` for `Function.Embedding`s. -/ @[simps (config := .asFn)] def optionMap {α β} (f : α ↪ β) : Option α ↪ Option β := diff --git a/Mathlib/Logic/Encodable/Basic.lean b/Mathlib/Logic/Encodable/Basic.lean index e1558f646d24a4..2504ac90134c1f 100644 --- a/Mathlib/Logic/Encodable/Basic.lean +++ b/Mathlib/Logic/Encodable/Basic.lean @@ -43,8 +43,6 @@ wish to enforce infiniteness. -/ class Encodable (α : Type*) where /-- Encoding from Type α to ℕ -/ encode : α → ℕ - -- Porting note: was `decode [] : ℕ → Option α`. This means that `decode` does not take the type - --explicitly in Lean4 /-- Decoding from ℕ to Option α-/ decode : ℕ → Option α /-- Invariant relationship between encoding and decoding -/ @@ -94,12 +92,10 @@ def ofLeftInverse [Encodable α] (f : β → α) (finv : α → β) (linv : ∀ def ofEquiv (α) [Encodable α] (e : β ≃ α) : Encodable β := ofLeftInverse e e.symm e.left_inv --- Porting note: removing @[simp], too powerful theorem encode_ofEquiv {α β} [Encodable α] (e : β ≃ α) (b : β) : @encode _ (ofEquiv _ e) b = encode (e b) := rfl --- Porting note: removing @[simp], too powerful theorem decode_ofEquiv {α β} [Encodable α] (e : β ≃ α) (n : ℕ) : @decode _ (ofEquiv _ e) n = (decode n).map e.symm := show Option.bind _ _ = Option.map _ _ @@ -224,7 +220,6 @@ section Sum variable [Encodable α] [Encodable β] --- Porting note: removing bit0 and bit1 /-- Explicit encoding function for the sum of two encodable types. -/ def encodeSum : α ⊕ β → ℕ | Sum.inl a => 2 * encode a @@ -240,12 +235,10 @@ def decodeSum (n : ℕ) : Option (α ⊕ β) := instance _root_.Sum.encodable : Encodable (α ⊕ β) := ⟨encodeSum, decodeSum, fun s => by cases s <;> simp [encodeSum, div2_val, decodeSum, encodek]⟩ --- Porting note: removing bit0 and bit1 from statement @[simp] theorem encode_inl (a : α) : @encode (α ⊕ β) _ (Sum.inl a) = 2 * (encode a) := rfl --- Porting note: removing bit0 and bit1 from statement @[simp] theorem encode_inr (b : β) : @encode (α ⊕ β) _ (Sum.inr b) = 2 * (encode b) + 1 := rfl diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 42d71d763a3249..dd6b85f1b20401 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -27,7 +27,7 @@ In this file we continue the work on equivalences begun in `Mathlib/Logic/Equiv/ * canonical isomorphisms between various types: e.g., - `Equiv.sumEquivSigmaBool` is the canonical equivalence between the sum of two types `α ⊕ β` - and the sigma-type `Σ b : Bool, b.casesOn α β`; + and the sigma-type `Σ b, bif b then β else α`; - `Equiv.prodSumDistrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum satisfy the distributive law up to a canonical equivalence; @@ -476,7 +476,7 @@ def piOptionEquivProd {α} {β : Option α → Type*} : `β` to be types from the same universe, so it cannot be used directly to transfer theorems about sigma types to theorems about sum types. In many cases one can use `ULift` to work around this difficulty. -/ -def sumEquivSigmaBool (α β) : α ⊕ β ≃ Σ b : Bool, b.casesOn α β := +def sumEquivSigmaBool (α β) : α ⊕ β ≃ Σ b, bif b then β else α := ⟨fun s => s.elim (fun x => ⟨false, x⟩) fun x => ⟨true, x⟩, fun s => match s with | ⟨false, a⟩ => inl a @@ -1429,10 +1429,7 @@ theorem swapCore_swapCore (r a b : α) : swapCore a b (swapCore a b r) = r := by unfold swapCore; split_ifs <;> cc theorem swapCore_comm (r a b : α) : swapCore a b r = swapCore b a r := by - unfold swapCore - -- Porting note: whatever solution works for `swapCore_swapCore` will work here too. - split_ifs with h₁ h₂ h₃ <;> try simp - · cases h₁; cases h₂; rfl + unfold swapCore; split_ifs <;> cc /-- `swap a b` is the permutation that swaps `a` and `b` and leaves other values as is. -/ diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index b3ca33b7c9ca30..59336397cf0f69 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -164,8 +164,6 @@ protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := instance : Trans Equiv Equiv Equiv where trans := Equiv.trans --- Porting note: this is not a syntactic tautology any more because --- the coercion from `e` to a function is now `DFunLike.coe` not `e.toFun` @[simp, mfld_simps] theorem toFun_as_coe (e : α ≃ β) : e.toFun = e := rfl @[simp, mfld_simps] theorem invFun_as_coe (e : α ≃ β) : e.invFun = e.symm := rfl @@ -224,14 +222,10 @@ protected def cast {α β : Sort _} (h : α = β) : α ≃ β := theorem Perm.coe_subsingleton {α : Type*} [Subsingleton α] (e : Perm α) : (e : α → α) = id := by rw [Perm.subsingleton_eq_refl e, coe_refl] --- Porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_refl` --- in an expression such as `Equiv.refl a x` @[simp] theorem refl_apply (x : α) : Equiv.refl α x = x := rfl @[simp] theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : (f.trans g : α → γ) = g ∘ f := rfl --- Porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_trans` --- in an expression such as `Equiv.trans f g x` @[simp] theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl @[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := e.right_inv x diff --git a/Mathlib/Logic/Equiv/Fin.lean b/Mathlib/Logic/Equiv/Fin.lean index c3622dc3f66bf0..02a7f1a2a6d226 100644 --- a/Mathlib/Logic/Equiv/Fin.lean +++ b/Mathlib/Logic/Equiv/Fin.lean @@ -303,9 +303,8 @@ theorem finRotate_one : finRotate 1 = Equiv.refl _ := simp only [Fin.lt_iff_val_lt_val, Fin.val_last, Fin.val_mk] at h simp [finRotate_of_lt h, Fin.ext_iff, Fin.add_def, Nat.mod_eq_of_lt (Nat.succ_lt_succ h)] --- Porting note: was a @[simp] theorem finRotate_apply_zero : finRotate n.succ 0 = 1 := by - rw [finRotate_succ_apply, Fin.zero_add] + simp theorem coe_finRotate_of_ne_last {i : Fin n.succ} (h : i ≠ Fin.last n) : (finRotate (n + 1) i : ℕ) = i + 1 := by diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 9c0c3d02003eed..231ce6a316449e 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -177,10 +177,7 @@ def optionSubtype [DecidableEq β] (x : β) : ext a cases a · simpa using e.property.symm - -- Porting note: this cases had been by `simpa`, - -- but `simp` here is mysteriously slow, even after squeezing. - -- `rfl` closes the goal quickly, so we use that. - · rfl + · simp right_inv e := by ext a rfl diff --git a/Mathlib/Logic/Equiv/PartialEquiv.lean b/Mathlib/Logic/Equiv/PartialEquiv.lean index 753807e85c86e8..8355147d52dbe1 100644 --- a/Mathlib/Logic/Equiv/PartialEquiv.lean +++ b/Mathlib/Logic/Equiv/PartialEquiv.lean @@ -164,20 +164,14 @@ def Simps.symm_apply (e : PartialEquiv α β) : β → α := initialize_simps_projections PartialEquiv (toFun → apply, invFun → symm_apply) --- Porting note: this can be proven with `dsimp only` --- @[simp, mfld_simps] --- theorem coe_mk (f : α → β) (g s t ml mr il ir) : --- (PartialEquiv.mk f g s t ml mr il ir : α → β) = f := by dsimp only +theorem coe_mk (f : α → β) (g s t ml mr il ir) : + (PartialEquiv.mk f g s t ml mr il ir : α → β) = f := rfl @[simp, mfld_simps] theorem coe_symm_mk (f : α → β) (g s t ml mr il ir) : ((PartialEquiv.mk f g s t ml mr il ir).symm : β → α) = g := rfl --- Porting note: this is now a syntactic tautology --- @[simp, mfld_simps] --- theorem toFun_as_coe : e.toFun = e := rfl - @[simp, mfld_simps] theorem invFun_as_coe : e.invFun = e.symm := rfl diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 1eee4fc7e87598..5bbe6198e2a15f 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -43,8 +43,6 @@ theorem const_injective [Nonempty α] : Injective (const α : β → α → β) theorem const_inj [Nonempty α] {y₁ y₂ : β} : const α y₁ = const α y₂ ↔ y₁ = y₂ := ⟨fun h ↦ const_injective h, fun h ↦ h ▸ rfl⟩ --- Porting note: `Function.onFun` is now reducible --- @[simp] theorem onFun_apply (f : β → β → γ) (g : α → β) (a b : α) : onFun f g a b = f (g a) (g b) := rfl diff --git a/Mathlib/Logic/Function/Defs.lean b/Mathlib/Logic/Function/Defs.lean index 8196c2377e17e6..79adace923e878 100644 --- a/Mathlib/Logic/Function/Defs.lean +++ b/Mathlib/Logic/Function/Defs.lean @@ -50,9 +50,6 @@ abbrev swap {φ : α → β → Sort u₃} (f : ∀ x y, φ x y) : ∀ y x, φ x theorem swap_def {φ : α → β → Sort u₃} (f : ∀ x y, φ x y) : swap f = fun y x => f x y := rfl --- Porting note: removed, it was never used --- notation f " -[" op "]- " g => combine f op g - @[simp, mfld_simps] theorem id_comp (f : α → β) : id ∘ f = f := rfl diff --git a/Mathlib/MeasureTheory/Category/MeasCat.lean b/Mathlib/MeasureTheory/Category/MeasCat.lean index 1d461ae5229578..ed9cde0e7bb8fd 100644 --- a/Mathlib/MeasureTheory/Category/MeasCat.lean +++ b/Mathlib/MeasureTheory/Category/MeasCat.lean @@ -114,8 +114,9 @@ def Integral : Giry.Algebra where end MeasCat -instance TopCat.hasForgetToMeasCat : HasForget₂ TopCat.{u} MeasCat.{u} := - BundledHom.mkHasForget₂ borel (fun f => ⟨f.1, f.2.borel_measurable⟩) (fun _ => rfl) +instance TopCat.hasForgetToMeasCat : HasForget₂ TopCat.{u} MeasCat.{u} where + forget₂.obj X := @MeasCat.of _ (borel X) + forget₂.map f := ⟨f.1, f.hom.2.borel_measurable⟩ /-- The Borel functor, the canonical embedding of topological spaces into measurable spaces. -/ abbrev Borel : TopCat.{u} ⥤ MeasCat.{u} := diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index b793c87ee48e0a..8591fb29495a53 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -511,7 +511,7 @@ theorem lintegral_rnDeriv_mul [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ (hf : AEMeasurable f ν) : ∫⁻ x, μ.rnDeriv ν x * f x ∂ν = ∫⁻ x, f x ∂μ := by nth_rw 2 [← withDensity_rnDeriv_eq μ ν hμν] rw [lintegral_withDensity_eq_lintegral_mul₀ (measurable_rnDeriv μ ν).aemeasurable hf] - rfl + simp only [Pi.mul_apply] lemma setLIntegral_rnDeriv_mul [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) {f : α → ℝ≥0∞} (hf : AEMeasurable f ν) {s : Set α} (hs : MeasurableSet s) : diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index bfb9326f21a71d..b118e35f16fbac 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -798,8 +798,9 @@ theorem addHaar_image_le_lintegral_abs_det_fderiv_aux1 (hs : MeasurableSet s) have I : ENNReal.ofReal |A.det| < m := by simp only [m, ENNReal.ofReal, lt_add_iff_pos_right, εpos, ENNReal.coe_lt_coe] rcases ((addHaar_image_le_mul_of_det_lt μ A I).and self_mem_nhdsWithin).exists with ⟨δ, h, δpos⟩ - obtain ⟨δ', δ'pos, hδ'⟩ : ∃ (δ' : ℝ), 0 < δ' ∧ ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := - continuousAt_iff.1 (ContinuousLinearMap.continuous_det (E := E)).continuousAt ε εpos + obtain ⟨δ', δ'pos, hδ'⟩ : ∃ (δ' : ℝ), 0 < δ' ∧ ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := by + refine continuousAt_iff.1 ?_ ε εpos + exact ContinuousLinearMap.continuous_det.continuousAt let δ'' : ℝ≥0 := ⟨δ' / 2, (half_pos δ'pos).le⟩ refine ⟨min δ δ'', lt_min δpos (half_pos δ'pos), ?_, ?_⟩ · intro B hB @@ -920,8 +921,9 @@ theorem lintegral_abs_det_fderiv_le_addHaar_image_aux1 (hs : MeasurableSet s) ∀ (t : Set E) (g : E → E), ApproximatesLinearOn g A t δ → ENNReal.ofReal |A.det| * μ t ≤ μ (g '' t) + ε * μ t := by intro A - obtain ⟨δ', δ'pos, hδ'⟩ : ∃ (δ' : ℝ), 0 < δ' ∧ ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := - continuousAt_iff.1 (ContinuousLinearMap.continuous_det (E := E)).continuousAt ε εpos + obtain ⟨δ', δ'pos, hδ'⟩ : ∃ (δ' : ℝ), 0 < δ' ∧ ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := by + refine continuousAt_iff.1 ?_ ε εpos + exact ContinuousLinearMap.continuous_det.continuousAt let δ'' : ℝ≥0 := ⟨δ' / 2, (half_pos δ'pos).le⟩ have I'' : ∀ B : E →L[ℝ] E, ‖B - A‖ ≤ ↑δ'' → |B.det - A.det| ≤ ↑ε := by intro B hB diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index dcb4374b447f53..99a81102f651b2 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -592,11 +592,9 @@ lemma eLpNorm_indicator_eq_eLpNorm_restrict (hs : MeasurableSet s) : rw [← lintegral_indicator hs] congr simp_rw [enorm_indicator_eq_indicator_enorm] - have h_zero : (fun x => x ^ p.toReal) (0 : ℝ≥0∞) = 0 := by - simp [ENNReal.toReal_pos hp_zero hp_top] - -- Porting note: The implicit argument should be specified because the elaborator can't deal with - -- `∘` well. - exact (Set.indicator_comp_of_zero (g := fun x : ℝ≥0∞ => x ^ p.toReal) h_zero).symm + rw [eq_comm, ← Function.comp_def (fun x : ℝ≥0∞ => x ^ p.toReal), Set.indicator_comp_of_zero, + Function.comp_def] + simp [ENNReal.toReal_pos hp_zero hp_top] @[deprecated (since := "2025-01-07")] alias eLpNorm_indicator_eq_restrict := eLpNorm_indicator_eq_eLpNorm_restrict @@ -611,7 +609,7 @@ lemma eLpNorm_restrict_le (f : α → F) (p : ℝ≥0∞) (μ : Measure α) (s : lemma eLpNorm_indicator_le (f : α → E) : eLpNorm (s.indicator f) p μ ≤ eLpNorm f p μ := by refine eLpNorm_mono_ae <| .of_forall fun x ↦ ?_ - suffices ‖s.indicator f x‖₊ ≤ ‖f x‖₊ by exact NNReal.coe_mono this + suffices ‖s.indicator f x‖₊ ≤ ‖f x‖₊ from NNReal.coe_mono this rw [nnnorm_indicator_eq_indicator_nnnorm] exact s.indicator_le_self _ x diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 80239aa069a9a6..6d6742753612f3 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -1185,11 +1185,9 @@ theorem _root_.Measurable.add_simpleFunc induction' f using SimpleFunc.induction with c s hs f f' hff' hf hf' · simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero] - change Measurable (g + s.piecewise (Function.const α c) (0 : α → E)) rw [← s.piecewise_same g, ← piecewise_add] exact Measurable.piecewise hs (hg.add_const _) (hg.add_const _) - · have : (g + ↑(f + f')) - = (Function.support f).piecewise (g + (f : α → E)) (g + f') := by + · have : (g + ↑(f + f')) = (Function.support f).piecewise (g + (f : α → E)) (g + f') := by ext x by_cases hx : x ∈ Function.support f · simpa only [SimpleFunc.coe_add, Pi.add_apply, Function.mem_support, ne_eq, not_not, @@ -1210,11 +1208,9 @@ theorem _root_.Measurable.simpleFunc_add induction' f using SimpleFunc.induction with c s hs f f' hff' hf hf' · simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero] - change Measurable (s.piecewise (Function.const α c) (0 : α → E) + g) rw [← s.piecewise_same g, ← piecewise_add] exact Measurable.piecewise hs (hg.const_add _) (hg.const_add _) - · have : (↑(f + f') + g) - = (Function.support f).piecewise ((f : α → E) + g) (f' + g) := by + · have : (↑(f + f') + g) = (Function.support f).piecewise ((f : α → E) + g) (f' + g) := by ext x by_cases hx : x ∈ Function.support f · simpa only [coe_add, Pi.add_apply, Function.mem_support, ne_eq, not_not, diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index 986010b3e928a6..e64dc4d0e948da 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Peter Pfaffelhuber -/ import Mathlib.MeasureTheory.SetSemiring -import Mathlib.Topology.Algebra.InfiniteSum.Defs -import Mathlib.Topology.Instances.ENNReal.Lemmas +import Mathlib.MeasureTheory.OuterMeasure.Induced /-! # Additive Contents @@ -28,25 +27,31 @@ Let `m` be an `AddContent C`. If `C` is a set semi-ring (`IsSetSemiring C`) we h * `MeasureTheory.sum_addContent_le_of_subset`: if `I` is a finset of pairwise disjoint sets in `C` and `⋃₀ I ⊆ t` for `t ∈ C`, then `∑ s ∈ I, m s ≤ m t`. * `MeasureTheory.addContent_mono`: if `s ⊆ t` for two sets in `C`, then `m s ≤ m t`. +* `MeasureTheory.addContent_sUnion_le_sum`: an `addContent C` on a `SetSemiring C` is + sub-additive. +* `MeasureTheory.addContent_iUnion_eq_tsum_of_disjoint_of_addContent_iUnion_le`: if an + `AddContent` is σ-subadditive on a semi-ring of sets, then it is σ-additive. * `MeasureTheory.addContent_union'`: if `s, t ∈ C` are disjoint and `s ∪ t ∈ C`, then `m (s ∪ t) = m s + m t`. If `C` is a set ring (`IsSetRing`), then `addContent_union` gives the same conclusion without the hypothesis `s ∪ t ∈ C` (since it is a consequence of `IsSetRing C`). -If `C` is a set ring (`MeasureTheory.IsSetRing C`), we have, for `s, t ∈ C`, +If `C` is a set ring (`MeasureTheory.IsSetRing C`), we have -* `MeasureTheory.addContent_union_le`: `m (s ∪ t) ≤ m s + m t` -* `MeasureTheory.addContent_le_diff`: `m s - m t ≤ m (s \ t)` +* `MeasureTheory.addContent_union_le`: for `s, t ∈ C`, `m (s ∪ t) ≤ m s + m t` +* `MeasureTheory.addContent_le_diff`: for `s, t ∈ C`, `m s - m t ≤ m (s \ t)` * `IsSetRing.addContent_of_union`: a function on a ring of sets which is additive on pairs of disjoint sets defines an additive content * `addContent_iUnion_eq_sum_of_tendsto_zero`: if an additive content is continuous at `∅`, then its value on a countable disjoint union is the sum of the values +* `MeasureTheory.addContent_iUnion_le_of_addContent_iUnion_eq_tsum`: if an `AddContent` is + σ-additive on a set ring, then it is σ-subadditive. -/ open Set Finset Function Filter -open scoped ENNReal Topology +open scoped ENNReal Topology Function namespace MeasureTheory @@ -124,6 +129,8 @@ lemma addContent_eq_add_disjointOfDiffUnion_of_subset (hC : IsSetSemiring C) exact hC.pairwiseDisjoint_union_disjointOfDiffUnion hs hI h_dis · rwa [hC.sUnion_union_disjointOfDiffUnion_of_subset hs hI hI_ss] +/-- For an `m : addContent C` on a `SetSemiring C`, if `I` is a `Finset` of pairwise disjoint + sets in `C` and `⋃₀ I ⊆ t` for `t ∈ C`, then `∑ s ∈ I, m s ≤ m t`.-/ lemma sum_addContent_le_of_subset (hC : IsSetSemiring C) (h_ss : ↑I ⊆ C) (h_dis : PairwiseDisjoint (I : Set (Set α)) id) (ht : t ∈ C) (hJt : ∀ s ∈ I, s ⊆ t) : @@ -132,6 +139,7 @@ lemma sum_addContent_le_of_subset (hC : IsSetSemiring C) rw [addContent_eq_add_disjointOfDiffUnion_of_subset hC ht h_ss hJt h_dis] exact le_add_right le_rfl +/-- An `addContent C` on a `SetSemiring C` is monotone. -/ lemma addContent_mono (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) (hst : s ⊆ t) : m s ≤ m t := by @@ -141,8 +149,127 @@ lemma addContent_mono (hC : IsSetSemiring C) (hs : s ∈ C) (ht : t ∈ C) · simp only [coe_singleton, pairwiseDisjoint_singleton] · simp [hst] +/-- For an `m : addContent C` on a `SetSemiring C` and `s t : Set α` with `s ⊆ t`, we can write +`m t = m s + ∑ i in hC.disjointOfDiff ht hs, m i`.-/ +theorem eq_add_disjointOfDiff_of_subset (hC : IsSetSemiring C) + (hs : s ∈ C) (ht : t ∈ C) (hst : s ⊆ t) : + m t = m s + ∑ i ∈ hC.disjointOfDiff ht hs, m i := by + classical + conv_lhs => rw [← hC.sUnion_insert_disjointOfDiff ht hs hst] + rw [← coe_insert, addContent_sUnion] + · rw [sum_insert] + exact hC.nmem_disjointOfDiff ht hs + · rw [coe_insert] + exact Set.insert_subset hs (hC.subset_disjointOfDiff ht hs) + · rw [coe_insert] + exact hC.pairwiseDisjoint_insert_disjointOfDiff ht hs + · rw [coe_insert] + rwa [hC.sUnion_insert_disjointOfDiff ht hs hst] + +/-- An `addContent C` on a `SetSemiring C` is sub-additive.-/ +lemma addContent_sUnion_le_sum {m : AddContent C} (hC : IsSetSemiring C) + (J : Finset (Set α)) (h_ss : ↑J ⊆ C) (h_mem : ⋃₀ ↑J ∈ C) : + m (⋃₀ ↑J) ≤ ∑ u ∈ J, m u := by + classical + have h1 : (disjiUnion J (hC.disjointOfUnion h_ss) + (hC.pairwiseDisjoint_disjointOfUnion h_ss) : Set (Set α)) ⊆ C := by + simp only [disjiUnion_eq_biUnion, coe_biUnion, mem_coe, iUnion_subset_iff] + exact fun _ x ↦ hC.disjointOfUnion_subset h_ss x + have h2 : PairwiseDisjoint (disjiUnion J (hC.disjointOfUnion h_ss) + ((hC.pairwiseDisjoint_disjointOfUnion h_ss)) : Set (Set α)) id := by + simp only [disjiUnion_eq_biUnion, coe_biUnion, mem_coe] + exact hC.pairwiseDisjoint_biUnion_disjointOfUnion h_ss + have h3 : ⋃₀ J = ⋃₀ ((disjiUnion J (hC.disjointOfUnion h_ss) + (hC.pairwiseDisjoint_disjointOfUnion h_ss)) : Set (Set α)) := by + simp only [disjiUnion_eq_biUnion, coe_biUnion, mem_coe] + exact (Exists.choose_spec (hC.disjointOfUnion_props h_ss)).2.2.2.2.2 + rw [h3, addContent_sUnion h1 h2, sum_disjiUnion] + · apply sum_le_sum + intro x hx + refine sum_addContent_le_of_subset hC (hC.disjointOfUnion_subset h_ss hx) + (hC.pairwiseDisjoint_disjointOfUnion_of_mem h_ss hx) (h_ss hx) + (fun _ s ↦ hC.subset_of_mem_disjointOfUnion h_ss hx s) + · simp only [disjiUnion_eq_biUnion, coe_biUnion, mem_coe] at * + exact h3.symm ▸ h_mem + +lemma addContent_le_sum_of_subset_sUnion {m : AddContent C} (hC : IsSetSemiring C) + {J : Finset (Set α)} (h_ss : ↑J ⊆ C) (ht : t ∈ C) (htJ : t ⊆ ⋃₀ ↑J) : + m t ≤ ∑ u ∈ J, m u := by + -- we can't apply `addContent_mono` and `addContent_sUnion_le_sum` because `⋃₀ ↑J` might not + -- be in `C` + classical + let Jt := J.image (fun u ↦ t ∩ u) + have ht_eq : t = ⋃₀ Jt := by + rw [coe_image, sUnion_image, ← inter_iUnion₂, inter_eq_self_of_subset_left] + rwa [← sUnion_eq_biUnion] + rw [ht_eq] + refine (addContent_sUnion_le_sum hC Jt ?_ ?_).trans ?_ + · intro s + simp only [Jt, coe_image, Set.mem_image, mem_coe, forall_exists_index, and_imp] + rintro u hu rfl + exact hC.inter_mem _ ht _ (h_ss hu) + · rwa [← ht_eq] + · refine (Finset.sum_image_le_of_nonneg fun _ _ ↦ zero_le _).trans (sum_le_sum fun u hu ↦ ?_) + exact addContent_mono hC (hC.inter_mem _ ht _ (h_ss hu)) (h_ss hu) inter_subset_right + +/-- If an `AddContent` is σ-subadditive on a semi-ring of sets, then it is σ-additive. -/ +theorem addContent_iUnion_eq_tsum_of_disjoint_of_addContent_iUnion_le {m : AddContent C} + (hC : IsSetSemiring C) + (m_subadd : ∀ (f : ℕ → Set α) (_ : ∀ i, f i ∈ C) (_ : ⋃ i, f i ∈ C) + (_hf_disj : Pairwise (Disjoint on f)), m (⋃ i, f i) ≤ ∑' i, m (f i)) + (f : ℕ → Set α) (hf : ∀ i, f i ∈ C) (hf_Union : (⋃ i, f i) ∈ C) + (hf_disj : Pairwise (Disjoint on f)) : + m (⋃ i, f i) = ∑' i, m (f i) := by + refine le_antisymm (m_subadd f hf hf_Union hf_disj) ?_ + refine tsum_le_of_sum_le ENNReal.summable fun I ↦ ?_ + classical + rw [← Finset.sum_image_of_disjoint addContent_empty (hf_disj.pairwiseDisjoint _)] + refine sum_addContent_le_of_subset hC (I := I.image f) ?_ ?_ hf_Union ?_ + · simp only [coe_image, Set.image_subset_iff] + refine (subset_preimage_image f I).trans (preimage_mono ?_) + rintro i ⟨j, _, rfl⟩ + exact hf j + · simp only [coe_image] + intro s hs t ht hst + rw [Set.mem_image] at hs ht + obtain ⟨i, _, rfl⟩ := hs + obtain ⟨j, _, rfl⟩ := ht + have hij : i ≠ j := by intro h_eq; rw [h_eq] at hst; exact hst rfl + exact hf_disj hij + · simp only [Finset.mem_image, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] + exact fun i _ ↦ subset_iUnion _ i + end IsSetSemiring +section AddContentExtend + +/-- An additive content obtained from another one on the same semiring of sets by setting the value +of each set not in the semiring at `∞`. -/ +protected noncomputable +def AddContent.extend (hC : IsSetSemiring C) (m : AddContent C) : AddContent C where + toFun := extend (fun x (_ : x ∈ C) ↦ m x) + empty' := by rw [extend_eq, addContent_empty]; exact hC.empty_mem + sUnion' I h_ss h_dis h_mem := by + rw [extend_eq] + swap; · exact h_mem + rw [addContent_sUnion h_ss h_dis h_mem] + refine Finset.sum_congr rfl (fun s hs ↦ ?_) + rw [extend_eq] + exact h_ss hs + +protected theorem AddContent.extend_eq_extend (hC : IsSetSemiring C) (m : AddContent C) : + m.extend hC = extend (fun x (_ : x ∈ C) ↦ m x) := rfl + +protected theorem AddContent.extend_eq (hC : IsSetSemiring C) (m : AddContent C) (hs : s ∈ C) : + m.extend hC s = m s := by + rwa [m.extend_eq_extend, extend_eq] + +protected theorem AddContent.extend_eq_top (hC : IsSetSemiring C) (m : AddContent C) (hs : s ∉ C) : + m.extend hC s = ∞ := by + rwa [m.extend_eq_extend, extend_eq_top] + +end AddContentExtend + section IsSetRing lemma addContent_union (hC : IsSetRing C) (hs : s ∈ C) (ht : t ∈ C) diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index 3a147318917637..ea268bfc743a08 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -272,8 +272,7 @@ lemma exists_integral_isMulLeftInvariant_eq_smul_of_hasCompactSupport (μ' μ : integral_isMulLeftInvariant_isMulRightInvariant_combo f_cont f_comp g_cont g_comp g_nonneg g_one /- Since the `ν`-factor is the same for `μ` and `μ'`, this gives the result. -/ rw [← A, mul_assoc, mul_comm] at B - simp only [B, integral_smul_nnreal_measure] - rfl + simp [B, integral_smul_nnreal_measure, c, NNReal.smul_def] open scoped Classical in /-- Given two left-invariant measures which are finite on compacts, `haarScalarFactor μ' μ` is a diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 5bd12b8bc56832..0dceda1950733f 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -470,7 +470,7 @@ theorem mkMetric_apply (m : ℝ≥0∞ → ℝ≥0∞) (s : Set X) : simp only [← OuterMeasure.coe_mkMetric, OuterMeasure.mkMetric, OuterMeasure.mkMetric', OuterMeasure.iSup_apply, OuterMeasure.mkMetric'.pre, OuterMeasure.boundedBy_apply, extend] refine - surjective_id.iSup_congr (id) fun r => + surjective_id.iSup_congr id fun r => iSup_congr_Prop Iff.rfl fun _ => surjective_id.iInf_congr _ fun t => iInf_congr_Prop Iff.rfl fun ht => ?_ dsimp diff --git a/Mathlib/MeasureTheory/SetSemiring.lean b/Mathlib/MeasureTheory/SetSemiring.lean index 976dcb2c87dc3e..21a16de95ad05c 100644 --- a/Mathlib/MeasureTheory/SetSemiring.lean +++ b/Mathlib/MeasureTheory/SetSemiring.lean @@ -431,8 +431,8 @@ lemma empty_nmem_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) (hj : (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.2.2.1 j hj lemma sUnion_disjointOfUnion (hC : IsSetSemiring C) (hJ : ↑J ⊆ C) : - ⋃₀ ⋃ x ∈ J, (hC.disjointOfUnion hJ x : Set (Set α)) = ⋃₀ J - := (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.2.2.2.symm + ⋃₀ ⋃ x ∈ J, (hC.disjointOfUnion hJ x : Set (Set α)) = ⋃₀ J := + (Exists.choose_spec (hC.disjointOfUnion_props hJ)).2.2.2.2.2.symm end disjointOfUnion diff --git a/Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean b/Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean index e0642d458ddf38..5cb49ca9935c60 100644 --- a/Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean +++ b/Mathlib/MeasureTheory/VectorMeasure/WithDensity.lean @@ -67,7 +67,7 @@ theorem withDensityᵥ_neg : μ.withDensityᵥ (-f) = -μ.withDensityᵥ f := by · ext1 i hi rw [VectorMeasure.neg_apply, withDensityᵥ_apply hf hi, ← integral_neg, withDensityᵥ_apply hf.neg hi] - rfl + simp only [Pi.neg_apply] · rw [withDensityᵥ, withDensityᵥ, dif_neg hf, dif_neg, neg_zero] rwa [integrable_neg_iff] @@ -105,7 +105,7 @@ theorem withDensityᵥ_smul {𝕜 : Type*} [NontriviallyNormedField 𝕜] [Norme · ext1 i hi rw [withDensityᵥ_apply (hf.smul r) hi, VectorMeasure.smul_apply, withDensityᵥ_apply hf hi, ← integral_smul r f] - rfl + simp only [Pi.smul_apply] · by_cases hr : r = 0 · rw [hr, zero_smul, zero_smul, withDensityᵥ_zero] · rw [withDensityᵥ, withDensityᵥ, dif_neg hf, dif_neg, smul_zero] @@ -123,7 +123,7 @@ theorem withDensityᵥ_smul_eq_withDensityᵥ_withDensity {f : α → ℝ≥0} { rw [withDensityᵥ_apply hfg hs, withDensityᵥ_apply ((integrable_withDensity_iff_integrable_smul₀ hf).mpr hfg) hs, setIntegral_withDensity_eq_setIntegral_smul₀ hf.restrict _ hs] - rfl + simp only [Pi.smul_apply'] theorem withDensityᵥ_smul_eq_withDensityᵥ_withDensity' {f : α → ℝ≥0∞} {g : α → E} (hf : AEMeasurable f μ) (hflt : ∀ᵐ x ∂μ, f x < ∞) @@ -131,7 +131,9 @@ theorem withDensityᵥ_smul_eq_withDensityᵥ_withDensity' {f : α → ℝ≥0 μ.withDensityᵥ (fun x ↦ (f x).toReal • g x) = (μ.withDensity f).withDensityᵥ g := by rw [← withDensity_congr_ae (coe_toNNReal_ae_eq hflt), ← withDensityᵥ_smul_eq_withDensityᵥ_withDensity hf.ennreal_toNNReal hfg] - rfl + apply congr_arg + ext + simp [NNReal.smul_def, ENNReal.coe_toNNReal_eq_toReal] theorem Measure.withDensityᵥ_absolutelyContinuous (μ : Measure α) (f : α → ℝ) : μ.withDensityᵥ f ≪ᵥ μ.toENNRealVectorMeasure := by diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index 5d5e4e175bd03e..4f5d86bc154e46 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -299,19 +299,18 @@ theorem finite_of_singleton [IsDomain B] [h : IsCyclotomicExtension {n} A B] : /-- If `S` is finite and `IsCyclotomicExtension S A B`, then `B` is a finite `A`-algebra. -/ protected theorem finite [IsDomain B] [h₁ : Finite S] [h₂ : IsCyclotomicExtension S A B] : Module.Finite A B := by - cases' nonempty_fintype S with h - revert h₂ A B - refine Set.Finite.induction_on _ h₁ (fun A B => ?_) @fun n S _ _ H A B => ?_ - · intro _ _ _ _ _ + rw [finite_coe_iff] at h₁ + induction S, h₁ using Set.Finite.induction_on generalizing h₂ A B with + | empty => refine Module.finite_def.2 ⟨({1} : Finset B), ?_⟩ simp [← top_toSubmodule, ← empty, toSubmodule_bot, Submodule.one_eq_span] - · intro _ _ _ _ h + | @insert n S _ _ H => haveI : IsCyclotomicExtension S A (adjoin A {b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1}) := union_left _ (insert n S) _ _ (subset_insert n S) haveI := H A (adjoin A {b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1}) have : Module.Finite (adjoin A {b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1}) B := by - rw [← union_singleton] at h - letI := @union_right S {n} A B _ _ _ h + rw [← union_singleton] at h₂ + let _ := union_right S {n} A B exact finite_of_singleton n _ _ exact Module.Finite.trans (adjoin A {b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1}) _ @@ -542,12 +541,11 @@ instance CyclotomicField.algebra' {R : Type*} [CommRing R] [Algebra R K] : instance {R : Type*} [CommRing R] [Algebra R K] : IsScalarTower R K (CyclotomicField n K) := SplittingField.isScalarTower _ -instance CyclotomicField.noZeroSMulDivisors [IsFractionRing A K] : - NoZeroSMulDivisors A (CyclotomicField n K) := by - refine NoZeroSMulDivisors.of_algebraMap_injective ?_ - rw [IsScalarTower.algebraMap_eq A K (CyclotomicField n K)] +instance [IsFractionRing A K] : NoZeroSMulDivisors A (CyclotomicField n K) := by + rw [NoZeroSMulDivisors.iff_faithfulSMul, faithfulSMul_iff_algebraMap_injective, + IsScalarTower.algebraMap_eq A K (CyclotomicField n K)] exact - (Function.Injective.comp (NoZeroSMulDivisors.algebraMap_injective K (CyclotomicField n K)) + (Function.Injective.comp (FaithfulSMul.algebraMap_injective K (CyclotomicField n K)) (IsFractionRing.injective A K) :) /-- If `A` is a domain with fraction field `K` and `n : ℕ+`, we define `CyclotomicRing n A K` as @@ -586,7 +584,7 @@ instance [IsFractionRing A K] : theorem algebraBase_injective [IsFractionRing A K] : Function.Injective <| algebraMap A (CyclotomicRing n A K) := - NoZeroSMulDivisors.algebraMap_injective _ _ + FaithfulSMul.algebraMap_injective _ _ instance : Algebra (CyclotomicRing n A K) (CyclotomicField n K) := (adjoin A _).toAlgebra @@ -596,7 +594,7 @@ theorem adjoin_algebra_injective : Subtype.val_injective instance : NoZeroSMulDivisors (CyclotomicRing n A K) (CyclotomicField n K) := - NoZeroSMulDivisors.of_algebraMap_injective (adjoin_algebra_injective n A K) + NoZeroSMulDivisors.iff_algebraMap_injective.mpr (adjoin_algebra_injective n A K) instance : IsScalarTower A (CyclotomicRing n A K) (CyclotomicField n K) := IsScalarTower.subalgebra' _ _ _ _ @@ -606,8 +604,8 @@ instance isCyclotomicExtension [IsFractionRing A K] [NeZero ((n : ℕ) : A)] : exists_prim_root := @fun a han => by rw [mem_singleton_iff] at han subst a - haveI := NeZero.of_noZeroSMulDivisors A K n - haveI := NeZero.of_noZeroSMulDivisors A (CyclotomicField n K) n + have := NeZero.of_faithfulSMul A K n + have := NeZero.of_faithfulSMul A (CyclotomicField n K) n obtain ⟨μ, hμ⟩ := (CyclotomicField.isCyclotomicExtension n K).exists_prim_root (mem_singleton n) refine ⟨⟨μ, subset_adjoin ?_⟩, ?_⟩ · apply (isRoot_of_unity_iff n.pos (CyclotomicField n K)).mpr diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index 0e61f7398af63f..e15aabe4533a76 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -161,7 +161,7 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( rw [pow_one] at hζ hcycl have : natDegree (minpoly K ζ) = 1 := by rw [hζ.eq_neg_one_of_two_right, show (-1 : L) = algebraMap K L (-1) by simp, - minpoly.eq_X_sub_C_of_algebraMap_inj _ (NoZeroSMulDivisors.algebraMap_injective K L)] + minpoly.eq_X_sub_C_of_algebraMap_inj _ (FaithfulSMul.algebraMap_injective K L)] exact natDegree_X_sub_C (-1) rcases Fin.equiv_iff_eq.2 this with ⟨e⟩ rw [← Algebra.discr_reindex K (hζ.powerBasis K).basis e, coe_basis, powerBasis_gen]; norm_num diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 2767a4973edb3c..2b096f99910472 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -139,7 +139,7 @@ noncomputable def embeddingsEquivPrimitiveRoots (C : Type*) [CommRing C] [IsDoma (hζ.powerBasis K).liftEquiv.trans { toFun := fun x => by haveI := IsCyclotomicExtension.neZero' n K L - haveI hn := NeZero.of_noZeroSMulDivisors K C n + haveI hn := NeZero.of_faithfulSMul K C n refine ⟨x.1, ?_⟩ cases x rwa [mem_primitiveRoots n.pos, ← isRoot_cyclotomic_iff, IsRoot.def, @@ -147,7 +147,7 @@ noncomputable def embeddingsEquivPrimitiveRoots (C : Type*) [CommRing C] [IsDoma ← eval₂_eq_eval_map, ← aeval_def] invFun := fun x => by haveI := IsCyclotomicExtension.neZero' n K L - haveI hn := NeZero.of_noZeroSMulDivisors K C n + haveI hn := NeZero.of_faithfulSMul K C n refine ⟨x.1, ?_⟩ cases x rwa [aeval_def, eval₂_eq_eval_map, hζ.powerBasis_gen K, ← @@ -347,7 +347,7 @@ theorem sub_one_norm_eq_eval_cyclotomic [IsCyclotomicExtension {n} K L] (h : 2 < rw [cyclotomic', eval_prod, ← @Finset.prod_attach E E, ← univ_eq_attach] refine Fintype.prod_equiv (hζ.embeddingsEquivPrimitiveRoots E hirr) _ _ fun σ => ?_ simp - haveI : NeZero ((n : ℕ) : E) := NeZero.of_noZeroSMulDivisors K _ (n : ℕ) + haveI : NeZero ((n : ℕ) : E) := NeZero.of_faithfulSMul K _ (n : ℕ) rw [Hprod, cyclotomic', ← cyclotomic_eq_prod_X_sub_primitiveRoots (isRoot_cyclotomic_iff.1 hz), ← map_cyclotomic_int, _root_.map_intCast, ← Int.cast_one, eval_intCast_map, eq_intCast, Int.cast_id] diff --git a/Mathlib/NumberTheory/FunctionField.lean b/Mathlib/NumberTheory/FunctionField.lean index cd2a6cc65454d3..28a58c87bd73f2 100644 --- a/Mathlib/NumberTheory/FunctionField.lean +++ b/Mathlib/NumberTheory/FunctionField.lean @@ -207,12 +207,12 @@ theorem inftyValuation.C {k : Fq} (hk : k ≠ 0) : theorem inftyValuation.X : inftyValuationDef Fq RatFunc.X = Multiplicative.ofAdd (1 : ℤ) := by rw [inftyValuationDef, if_neg RatFunc.X_ne_zero, RatFunc.intDegree_X] -@[simp] +-- Dropped attribute `@[simp]` due to issue described here: +-- https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/.60synthInstance.2EmaxHeartbeats.60.20error.20but.20only.20in.20.60simpNF.60 theorem inftyValuation.polynomial {p : Fq[X]} (hp : p ≠ 0) : inftyValuationDef Fq (algebraMap Fq[X] (RatFunc Fq) p) = Multiplicative.ofAdd (p.natDegree : ℤ) := by - have hp' : algebraMap Fq[X] (RatFunc Fq) p ≠ 0 := by - rw [Ne, NoZeroSMulDivisors.algebraMap_eq_zero_iff]; exact hp + have hp' : algebraMap Fq[X] (RatFunc Fq) p ≠ 0 := by simpa rw [inftyValuationDef, if_neg hp', RatFunc.intDegree_polynomial] /-- The valued field `Fq(t)` with the valuation at infinity. -/ diff --git a/Mathlib/NumberTheory/KummerDedekind.lean b/Mathlib/NumberTheory/KummerDedekind.lean index 664ad5684a29a3..37e5831f0e6d39 100644 --- a/Mathlib/NumberTheory/KummerDedekind.lean +++ b/Mathlib/NumberTheory/KummerDedekind.lean @@ -100,7 +100,7 @@ lemma mem_coeSubmodule_conductor {L} [CommRing L] [Algebra S L] [Algebra R L] obtain ⟨y, _, e⟩ := H 1 rw [map_one, mul_one] at e subst e - simp only [← _root_.map_mul, (NoZeroSMulDivisors.algebraMap_injective S L).eq_iff, + simp only [← _root_.map_mul, (FaithfulSMul.algebraMap_injective S L).eq_iff, exists_eq_right] at H exact ⟨_, H, rfl⟩ · rw [AlgHom.map_adjoin, Set.image_singleton]; rfl @@ -233,7 +233,7 @@ attribute [local instance] Ideal.Quotient.field private noncomputable def f (hx : (conductor R x).comap (algebraMap R S) ⊔ I = ⊤) (hx' : IsIntegral R x) : S ⧸ I.map (algebraMap R S) ≃+* (R ⧸ I)[X] ⧸ span {(minpoly R x).map (Ideal.Quotient.mk I)} := - (quotAdjoinEquivQuotMap hx (NoZeroSMulDivisors.algebraMap_injective + (quotAdjoinEquivQuotMap hx (FaithfulSMul.algebraMap_injective (Algebra.adjoin R {x}) S)).symm.trans <| ((Algebra.adjoin.powerBasis' hx').quotientEquivQuotientMinpolyMap I).toRingEquiv.trans <| quotEquivOfEq (by rw [Algebra.adjoin.powerBasis'_minpoly_gen hx']) @@ -248,7 +248,7 @@ private lemma f_symm_aux (hx : (conductor R x).comap (algebraMap R S) ⊔ I = congr convert (adjoin.powerBasis' hx').quotientEquivQuotientMinpolyMap_symm_apply_mk I Q apply (quotAdjoinEquivQuotMap hx - (NoZeroSMulDivisors.algebraMap_injective ((adjoin R {x})) S)).injective + (FaithfulSMul.algebraMap_injective ((adjoin R {x})) S)).injective simp only [RingEquiv.apply_symm_apply, adjoin.powerBasis'_gen, quotAdjoinEquivQuotMap_apply_mk, coe_aeval_mk_apply] @@ -262,7 +262,7 @@ noncomputable def normalizedFactorsMapEquivNormalizedFactorsMinPolyMk (hI : IsMa {d : (R ⧸ I)[X] | d ∈ normalizedFactors (Polynomial.map (Ideal.Quotient.mk I) (minpoly R x))} := by refine (normalizedFactorsEquivOfQuotEquiv (f hx hx') ?_ ?_).trans ?_ - · rwa [Ne, map_eq_bot_iff_of_injective (NoZeroSMulDivisors.algebraMap_injective R S), ← Ne] + · rwa [Ne, map_eq_bot_iff_of_injective (FaithfulSMul.algebraMap_injective R S), ← Ne] · by_contra h exact (show Polynomial.map (Ideal.Quotient.mk I) (minpoly R x) ≠ 0 from Polynomial.map_monic_ne_zero (minpoly.monic hx')) (span_singleton_eq_bot.mp h) @@ -326,7 +326,7 @@ theorem normalizedFactors_ideal_map_eq_normalizedFactors_min_poly_mk_map (hI : I · exact Polynomial.map_monic_ne_zero (minpoly.monic hx') · exact irreducible_of_normalized_factor _ hJ · rwa [← bot_eq_zero, Ne, - map_eq_bot_iff_of_injective (NoZeroSMulDivisors.algebraMap_injective R S)] + map_eq_bot_iff_of_injective (FaithfulSMul.algebraMap_injective R S)] theorem Ideal.irreducible_map_of_irreducible_minpoly (hI : IsMaximal I) (hI' : I ≠ ⊥) (hx : (conductor R x).comap (algebraMap R S) ⊔ I = ⊤) (hx' : IsIntegral R x) @@ -340,7 +340,7 @@ theorem Ideal.irreducible_map_of_irreducible_minpoly (hI : IsMaximal I) (hI' : I obtain ⟨y, hy⟩ := this have h := prod_normalizedFactors (show I.map (algebraMap R S) ≠ 0 by rwa [← bot_eq_zero, Ne, - map_eq_bot_iff_of_injective (NoZeroSMulDivisors.algebraMap_injective R S)]) + map_eq_bot_iff_of_injective (FaithfulSMul.algebraMap_injective R S)]) rw [associated_iff_eq, hy, Multiset.prod_singleton] at h rw [← h] exact diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index 65ce05f28dc8ed..e0b835f9c10c81 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -206,15 +206,15 @@ variable {K} /-- The canonical map from `𝓞 K` to `K` is injective. -This is a convenient abbreviation for `NoZeroSMulDivisors.algebraMap_injective`. +This is a convenient abbreviation for `FaithfulSMul.algebraMap_injective`. -/ lemma coe_injective : Function.Injective (algebraMap (𝓞 K) K) := - NoZeroSMulDivisors.algebraMap_injective _ _ + FaithfulSMul.algebraMap_injective _ _ /-- The canonical map from `𝓞 K` to `K` is injective. This is a convenient abbreviation for `map_eq_zero_iff` applied to -`NoZeroSMulDivisors.algebraMap_injective`. +`FaithfulSMul.algebraMap_injective`. -/ lemma coe_eq_zero_iff {x : 𝓞 K} : algebraMap _ K x = 0 ↔ x = 0 := map_eq_zero_iff _ coe_injective @@ -222,7 +222,7 @@ lemma coe_eq_zero_iff {x : 𝓞 K} : algebraMap _ K x = 0 ↔ x = 0 := /-- The canonical map from `𝓞 K` to `K` is injective. This is a convenient abbreviation for `map_ne_zero_iff` applied to -`NoZeroSMulDivisors.algebraMap_injective`. +`FaithfulSMul.algebraMap_injective`. -/ lemma coe_ne_zero_iff {x : 𝓞 K} : algebraMap _ K x ≠ 0 ↔ x ≠ 0 := map_ne_zero_iff _ coe_injective @@ -340,10 +340,10 @@ 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) + NoZeroSMulDivisors.iff_algebraMap_injective.mpr <| algebraMap.injective K L instance : NoZeroSMulDivisors (𝓞 K) L := - NoZeroSMulDivisors.trans (𝓞 K) (𝓞 L) L + NoZeroSMulDivisors.trans_faithfulSMul (𝓞 K) (𝓞 L) L end extension diff --git a/Mathlib/NumberTheory/NumberField/EquivReindex.lean b/Mathlib/NumberTheory/NumberField/EquivReindex.lean index 76ece2db3d216a..d3862754f4aec4 100644 --- a/Mathlib/NumberTheory/NumberField/EquivReindex.lean +++ b/Mathlib/NumberTheory/NumberField/EquivReindex.lean @@ -26,9 +26,9 @@ open Module.Free Module canonicalEmbedding Matrix Finset /-- An equivalence between the set of embeddings of `K` into `ℂ` and the index set of the chosen basis of the ring of integers of `K`. -/ -abbrev equivReindex : (K →+* ℂ) ≃ (ChooseBasisIndex ℤ (𝓞 K)) := - Fintype.equivOfCardEq <| - by rw [Embeddings.card, ← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank] +abbrev equivReindex : (K →+* ℂ) ≃ ChooseBasisIndex ℤ (𝓞 K) := + Fintype.equivOfCardEq <| by + rw [Embeddings.card, ← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank] /-- The basis matrix for the embeddings of `K` into `ℂ`. This matrix is formed by taking the lattice basis vectors of `K` and reindexing them according to the diff --git a/Mathlib/NumberTheory/NumberField/FinitePlaces.lean b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean index 5d5cf658791b0c..cd928c8084a1c8 100644 --- a/Mathlib/NumberTheory/NumberField/FinitePlaces.lean +++ b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean @@ -91,7 +91,7 @@ noncomputable instance instRankOneValuedAdicCompletion : dsimp [adicCompletion] rw [valuedAdicCompletion_eq_valuation' v (x : K)] constructor - · simpa only [ne_eq, map_eq_zero, NoZeroSMulDivisors.algebraMap_eq_zero_iff] + · simpa only [ne_eq, map_eq_zero, FaithfulSMul.algebraMap_eq_zero_iff] · apply ne_of_lt rw [valuation_eq_intValuationDef, intValuation_lt_one_iff_dvd] exact dvd_span_singleton.mpr hx1 @@ -235,8 +235,7 @@ theorem mulSupport_finite_int {x : 𝓞 K} (h_x_nezero : x ≠ 0) : theorem mulSupport_finite {x : K} (h_x_nezero : x ≠ 0) : (Function.mulSupport fun w : FinitePlace K ↦ w x).Finite := by rcases IsFractionRing.div_surjective (A := 𝓞 K) x with ⟨a, b, hb, rfl⟩ - simp_all only [ne_eq, div_eq_zero_iff, NoZeroSMulDivisors.algebraMap_eq_zero_iff, not_or, - map_div₀] + simp_all only [ne_eq, div_eq_zero_iff, FaithfulSMul.algebraMap_eq_zero_iff, not_or, map_div₀] obtain ⟨ha, hb⟩ := h_x_nezero simp_rw [← RingOfIntegers.coe_eq_algebraMap] apply ((mulSupport_finite_int ha).union (mulSupport_finite_int hb)).subset diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean index f5cd7772f5a6a3..50211fbf7068c4 100644 --- a/Mathlib/NumberTheory/NumberField/House.lean +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -133,7 +133,7 @@ private theorem asiegel_ne_0 : asiegel K a ≠ 0 := by simp (config := { unfoldPartialApp := true }) only [asiegel, a'] simp only [ne_eq] rw [funext_iff]; intros hs - simp only [Prod.forall] at hs; + simp only [Prod.forall] at hs apply ha rw [← Matrix.ext_iff]; intros k' l specialize hs k' diff --git a/Mathlib/NumberTheory/RamificationInertia/Basic.lean b/Mathlib/NumberTheory/RamificationInertia/Basic.lean index 137fee8752dbb4..4c8addfccce8ab 100644 --- a/Mathlib/NumberTheory/RamificationInertia/Basic.lean +++ b/Mathlib/NumberTheory/RamificationInertia/Basic.lean @@ -271,7 +271,7 @@ theorem FinrankQuotientMap.span_eq_top [IsDomain R] [IsDomain S] [Algebra K L] [ Submodule.span K (algebraMap S L '' b) = ⊤ := by have hRL : Function.Injective (algebraMap R L) := by rw [IsScalarTower.algebraMap_eq R K L] - exact (algebraMap K L).injective.comp (NoZeroSMulDivisors.algebraMap_injective R K) + exact (algebraMap K L).injective.comp (FaithfulSMul.algebraMap_injective R K) -- Let `M` be the `R`-module spanned by the proposed basis elements. let M : Submodule R S := Submodule.span R b -- Then `S / M` is generated by some finite set of `n` vectors `a`. diff --git a/Mathlib/Order/Fin/SuccAboveOrderIso.lean b/Mathlib/Order/Fin/SuccAboveOrderIso.lean new file mode 100644 index 00000000000000..9f02a2b2448f47 --- /dev/null +++ b/Mathlib/Order/Fin/SuccAboveOrderIso.lean @@ -0,0 +1,37 @@ +/- +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.Order.Fin.Basic +import Mathlib.Data.Fintype.Basic +import Mathlib.Tactic.FinCases + +/-! +# The order isomorphism `Fin (n + 1) ≃o {i}ᶜ` + +Given `i : Fin (n + 2)`, we show that `Fin.succAboveOrderEmb` induces +an order isomorphism `Fin (n + 1) ≃o ({i}ᶜ : Finset (Fin (n + 2)))`. + +-/ + +open Finset + +/-- Given `i : Fin (n + 2)`, this is the order isomorphism +between `Fin (n + 1)` and the finite set `{i}ᶜ`. -/ +noncomputable def Fin.succAboveOrderIso {n : ℕ} (i : Fin (n + 2)) : + Fin (n + 1) ≃o ({i}ᶜ : Finset (Fin (n + 2))) where + toEquiv := + Equiv.ofBijective (f := fun a ↦ ⟨Fin.succAboveOrderEmb i a, + by simpa using Fin.succAbove_ne i a⟩) (by + constructor + · intro a b h + exact (Fin.succAboveOrderEmb i).injective (by simpa using h) + · rintro ⟨j, hj⟩ + simp only [mem_compl, mem_singleton] at hj + obtain rfl | ⟨i, rfl⟩ := Fin.eq_zero_or_eq_succ i + · exact ⟨j.pred hj, by simp⟩ + · exact ⟨i.predAbove j, by aesop⟩) + map_rel_iff' {a b} := by + simp only [Equiv.ofBijective_apply, Subtype.mk_le_mk, OrderEmbedding.le_iff_le] diff --git a/Mathlib/Order/Interval/Set/InitialSeg.lean b/Mathlib/Order/Interval/Set/InitialSeg.lean new file mode 100644 index 00000000000000..399bb0faee906a --- /dev/null +++ b/Mathlib/Order/Interval/Set/InitialSeg.lean @@ -0,0 +1,83 @@ +/- +Copyright (c) 2025 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Order.InitialSeg + +/-! +# Intervals as initial segments + +We show that `Set.Iic` and `Set.Iio` are respectively initial and +principal segments, and that any principal segment `f` is order +isomorphic to `Set.Iio f.top`. + +-/ + +namespace Set + +variable {α : Type*} [Preorder α] + +/-- `Set.Iic j` is an initial segment. -/ +@[simps] +def initialSegIic (j : α) : + Set.Iic j ≤i α where + toFun := fun ⟨j, hj⟩ ↦ j + inj' _ _ _ := by aesop + map_rel_iff' := by aesop + mem_range_of_rel' x k h := by simpa using h.le.trans x.2 + +/-- `Set.Iio j` is a principal segment. -/ +@[simps] +def principalSegIio (j : α) : + Set.Iio j h (a, b))) ∧ (∀ b, Monotone (fun a => h (a, b))) where + mp h := ⟨fun _ _ _ hab => h (Prod.mk_le_mk_iff_right.mpr hab), + fun _ _ _ hab => h (Prod.mk_le_mk_iff_left.mpr hab)⟩ + mpr h _ _ hab := le_trans (h.1 _ (Prod.mk_le_mk.mp hab).2) (h.2 _ (Prod.mk_le_mk.mp hab).1) + +lemma antitone_prod_iff {h : α × β → γ} : + Antitone h ↔ (∀ a, Antitone (fun b => h (a, b))) ∧ (∀ b, Antitone (fun a => h (a, b))) where + mp h := ⟨fun _ _ _ hab => h (Prod.mk_le_mk_iff_right.mpr hab), + fun _ _ _ hab => h (Prod.mk_le_mk_iff_left.mpr hab)⟩ + mpr h _ _ hab:= le_trans (h.1 _ (Prod.mk_le_mk.mp hab).2) (h.2 _ (Prod.mk_le_mk.mp hab).1) + end Preorder section PartialOrder diff --git a/Mathlib/Order/Nucleus.lean b/Mathlib/Order/Nucleus.lean index e867d9f21930b9..8f3b54ef36a4ef 100644 --- a/Mathlib/Order/Nucleus.lean +++ b/Mathlib/Order/Nucleus.lean @@ -55,6 +55,7 @@ instance : FunLike (Nucleus X) X X where lemma toFun_eq_coe (n : Nucleus X) : n.toFun = n := rfl @[simp] lemma coe_toInfHom (n : Nucleus X) : ⇑n.toInfHom = n := rfl +@[simp] lemma coe_mk (f : InfHom X X) (h1 h2) : ⇑(mk f h1 h2) = f := rfl instance : NucleusClass (Nucleus X) X where idempotent _ _ := idempotent' .. diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index ec410d68a63d8d..4703ef101dec41 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -219,10 +219,8 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] theorem integrable_pdf_smul_iff [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X ℙ μ] {f : E → F} (hf : AEStronglyMeasurable f μ) : Integrable (fun x => (pdf X ℙ μ x).toReal • f x) μ ↔ Integrable (fun x => f (X x)) ℙ := by - -- Porting note: using `erw` because `rw` doesn't recognize `(f <| X ·)` as `f ∘ X` - -- https://github.com/leanprover-community/mathlib4/issues/5164 - erw [← integrable_map_measure (hf.mono_ac HasPDF.absolutelyContinuous) - (HasPDF.aemeasurable X ℙ μ), + rw [← Function.comp_def, + ← integrable_map_measure (hf.mono_ac HasPDF.absolutelyContinuous) (HasPDF.aemeasurable X ℙ μ), map_eq_withDensity_pdf X ℙ μ, pdf_def, integrable_rnDeriv_smul_iff HasPDF.absolutelyContinuous] rw [withDensity_rnDeriv_eq _ _ HasPDF.absolutelyContinuous] diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 8b3c99c7d763c6..e807d75af5d28e 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -283,8 +283,8 @@ lemma gaussianReal_map_const_mul (c : ℝ) : rw [Measure.map_const] simp only [ne_eq, measure_univ, one_smul, mul_eq_zero] convert (gaussianReal_zero_var 0).symm - simp only [ne_eq, zero_pow, mul_eq_zero, hv, or_false, not_false_eq_true, reduceCtorEq] - rfl + simp only [ne_eq, zero_pow, mul_eq_zero, hv, or_false, not_false_eq_true, reduceCtorEq, + NNReal.mk_zero] let e : ℝ ≃ᵐ ℝ := (Homeomorph.mulLeft₀ c hc).symm.toMeasurableEquiv have he' : ∀ x, HasDerivAt e ((fun _ ↦ c⁻¹) x) x := by suffices ∀ x, HasDerivAt (fun x => c⁻¹ * x) (c⁻¹ * 1) x by rwa [mul_one] at this diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index a1aa1648a693ef..8197f13c31cc9f 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -131,15 +131,16 @@ lemma measurable_densityProcess (κ : Kernel α (γ × β)) (ν : Kernel α γ) Measurable (fun (p : α × γ) ↦ densityProcess κ ν n p.1 p.2 s) := (measurable_densityProcess_aux κ ν n hs).ennreal_toReal +-- The following two lemmas also work without the `( :)`, but they are slow. lemma measurable_densityProcess_left (κ : Kernel α (γ × β)) (ν : Kernel α γ) (n : ℕ) (x : γ) {s : Set β} (hs : MeasurableSet s) : Measurable (fun a ↦ densityProcess κ ν n a x s) := - (measurable_densityProcess κ ν n hs).comp (measurable_id.prod_mk measurable_const) + ((measurable_densityProcess κ ν n hs).comp (measurable_id.prod_mk measurable_const):) lemma measurable_densityProcess_right (κ : Kernel α (γ × β)) (ν : Kernel α γ) (n : ℕ) {s : Set β} (a : α) (hs : MeasurableSet s) : Measurable (fun x ↦ densityProcess κ ν n a x s) := - (measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id) + ((measurable_densityProcess κ ν n hs).comp (measurable_const.prod_mk measurable_id):) lemma measurable_countableFiltration_densityProcess (κ : Kernel α (γ × β)) (ν : Kernel α γ) (n : ℕ) (a : α) {s : Set β} (hs : MeasurableSet s) : diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index b2e5e6476ce211..25eddf3912d4a0 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -278,7 +278,7 @@ open groupCohomology.resolution /-- Given a `k`-linear `G`-representation `A`, the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` is `k`-linearly isomorphic to the set of functions `Gⁿ → A`. -/ noncomputable def diagonalHomEquiv : - (Rep.ofMulAction k G (Fin (n + 1) → G) ⟶ A) ≃ₗ[k] (Fin n → G) → A := + (Rep.diagonal k G (n + 1) ⟶ A) ≃ₗ[k] (Fin n → G) → A := Linear.homCongr k ((diagonalSucc k G n).trans ((Representation.ofMulAction k G G).repOfTprodIso 1)) (Iso.refl _) ≪≫ₗ @@ -291,7 +291,7 @@ variable {n A} the set of representation morphisms `Hom(k[Gⁿ⁺¹], A)` with `Fun(Gⁿ, A)`. This lemma says that this sends a morphism of representations `f : k[Gⁿ⁺¹] ⟶ A` to the function `(g₁, ..., gₙ) ↦ f(1, g₁, g₁g₂, ..., g₁g₂...gₙ).` -/ -theorem diagonalHomEquiv_apply (f : Rep.ofMulAction k G (Fin (n + 1) → G) ⟶ A) (x : Fin n → G) : +theorem diagonalHomEquiv_apply (f : Rep.diagonal k G (n + 1) ⟶ A) (x : Fin n → G) : diagonalHomEquiv n A f x = f.hom (Finsupp.single (Fin.partialProd x) 1) := by /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was unfold diagonalHomEquiv @@ -324,8 +324,7 @@ theorem diagonalHomEquiv_symm_apply (f : (Fin n → G) → A) (x : Fin (n + 1) leftRegularHomEquiv_symm_apply, Linear.homCongr_symm_apply, Iso.trans_hom, Iso.refl_inv, Category.comp_id, Action.comp_hom, MonoidalClosed.linearHomEquivComm_symm_hom, ModuleCat.hom_comp, LinearMap.comp_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [diagonalSucc_hom_single] + rw [diagonalSucc_hom_single] -- The prototype linter that checks if `erw` could be replaced with `rw` would time out -- if it replaces the next `erw`s with `rw`s. So we focus down on the relevant part. conv_lhs => diff --git a/Mathlib/RingTheory/AlgebraTower.lean b/Mathlib/RingTheory/AlgebraTower.lean index ef70b261a16453..0b9fce27c1b707 100644 --- a/Mathlib/RingTheory/AlgebraTower.lean +++ b/Mathlib/RingTheory/AlgebraTower.lean @@ -167,7 +167,7 @@ variable [CommRing R] [Ring S] [Algebra R S] theorem Basis.algebraMap_injective {ι : Type*} [NoZeroDivisors R] [Nontrivial S] (b : @Basis ι R S _ _ Algebra.toModule) : Function.Injective (algebraMap R S) := have : NoZeroSMulDivisors R S := b.noZeroSMulDivisors - NoZeroSMulDivisors.algebraMap_injective R S + FaithfulSMul.algebraMap_injective R S end Ring diff --git a/Mathlib/RingTheory/Algebraic/Cardinality.lean b/Mathlib/RingTheory/Algebraic/Cardinality.lean index 84f3f265104fcd..eef75eac83d8f8 100644 --- a/Mathlib/RingTheory/Algebraic/Cardinality.lean +++ b/Mathlib/RingTheory/Algebraic/Cardinality.lean @@ -32,7 +32,7 @@ theorem lift_cardinalMk_le_sigma_polynomial : let p := Classical.indefiniteDescription _ (Algebra.IsAlgebraic.isAlgebraic x) ⟨p.1, x, by dsimp - have := (Polynomial.map_ne_zero_iff (NoZeroSMulDivisors.algebraMap_injective R L)).2 p.2.1 + have := (Polynomial.map_ne_zero_iff (FaithfulSMul.algebraMap_injective R L)).2 p.2.1 rw [Polynomial.mem_roots this, Polynomial.IsRoot, Polynomial.eval_map, ← Polynomial.aeval_def, p.2.2]⟩) fun x y => by diff --git a/Mathlib/RingTheory/Algebraic/Integral.lean b/Mathlib/RingTheory/Algebraic/Integral.lean index 625d736514f0f0..6dc35de98b4bd7 100644 --- a/Mathlib/RingTheory/Algebraic/Integral.lean +++ b/Mathlib/RingTheory/Algebraic/Integral.lean @@ -251,7 +251,7 @@ theorem restrictScalars [Algebra.IsAlgebraic R S] exact int _ (Finset.mem_image_of_mem _ <| support_smul _ _ hn) have : IsAlgebraic (integralClosure R S) a := by refine ⟨p, ?_, ?_⟩ - · have := NoZeroSMulDivisors.of_algebraMap_injective hRS + · have : FaithfulSMul R S := (faithfulSMul_iff_algebraMap_injective R S).mpr hRS simpa only [← Polynomial.map_ne_zero_iff (f := Subring.subtype _) Subtype.val_injective, p, map_toSubring, smul_ne_zero_iff] using And.intro hr hp rw [← eval_map_algebraMap, Subalgebra.algebraMap_eq, ← map_map, ← Subalgebra.toSubring_subtype, diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index 5ab55f608ea3d7..5295cf6eec2d48 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -120,7 +120,7 @@ variable [IsFractionRing A K] variable (A K) in lemma map_equiv_traceDual [IsDomain A] [IsFractionRing B L] [IsDomain B] - [NoZeroSMulDivisors A B] (I : Submodule B (FractionRing B)) : + [FaithfulSMul A B] (I : Submodule B (FractionRing B)) : (traceDual A (FractionRing A) I).map (FractionRing.algEquiv B L) = traceDual A K (I.map (FractionRing.algEquiv B L)) := by show Submodule.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap _ = @@ -580,7 +580,7 @@ lemma pow_sub_one_dvd_differentIdeal_aux [IsFractionRing B L] [IsDedekindDomain (hP : P ^ e ∣ p.map (algebraMap A B)) : P ^ (e - 1) ∣ differentIdeal A B := by obtain ⟨a, ha⟩ := (pow_dvd_pow _ (Nat.sub_le e 1)).trans hP have hp' := (Ideal.map_eq_bot_iff_of_injective - (NoZeroSMulDivisors.algebraMap_injective A B)).not.mpr hp + (FaithfulSMul.algebraMap_injective A B)).not.mpr hp have habot : a ≠ ⊥ := fun ha' ↦ hp' (by simpa [ha'] using ha) have hPbot : P ≠ ⊥ := by rintro rfl; apply hp' diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index 45e9ac50e0e503..12909e94394959 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -343,19 +343,23 @@ instance : Coe (FiniteAdeleRing R K) (K_hat R K) where theorem coe_one : (1 : FiniteAdeleRing R K) = (1 : K_hat R K) := rfl @[simp, norm_cast] -theorem coe_zero: (0 : FiniteAdeleRing R K) = (0 : K_hat R K) := rfl +theorem coe_zero : (0 : FiniteAdeleRing R K) = (0 : K_hat R K) := rfl @[simp, norm_cast] -theorem coe_add (x y : FiniteAdeleRing R K) : (x + y : FiniteAdeleRing R K) = - (x : K_hat R K) + (y : K_hat R K) := rfl +theorem coe_add (x y : FiniteAdeleRing R K) : + (x + y : FiniteAdeleRing R K) = (x : K_hat R K) + (y : K_hat R K) := + rfl @[simp, norm_cast] -theorem coe_mul (x y : FiniteAdeleRing R K) : (x * y : FiniteAdeleRing R K) = - (x : K_hat R K) * (y : K_hat R K) := rfl +theorem coe_mul (x y : FiniteAdeleRing R K) : + (x * y : FiniteAdeleRing R K) = (x : K_hat R K) * (y : K_hat R K) := + rfl @[simp, norm_cast] -theorem coe_algebraMap (x : K) : (((algebraMap K (FiniteAdeleRing R K)) x) : K_hat R K) = - (algebraMap K (ProdAdicCompletions R K)) x := rfl +theorem coe_algebraMap (x : K) : + (((algebraMap K (FiniteAdeleRing R K)) x) : K_hat R K) = + (algebraMap K (ProdAdicCompletions R K)) x := + rfl @[ext] lemma ext {a₁ a₂ : FiniteAdeleRing R K} (h : (a₁ : K_hat R K) = a₂) : a₁ = a₂ := diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 6092f3c84aea48..fb14158a971a9f 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -142,7 +142,7 @@ theorem coe_ideal_span_singleton_div_self {x : R₁} (hx : x ≠ 0) : (Ideal.span ({x} : Set R₁) : FractionalIdeal R₁⁰ K) / Ideal.span ({x} : Set R₁) = 1 := by rw [coeIdeal_span_singleton, spanSingleton_div_self K <| - (map_ne_zero_iff _ <| NoZeroSMulDivisors.algebraMap_injective R₁ K).mpr hx] + (map_ne_zero_iff _ <| FaithfulSMul.algebraMap_injective R₁ K).mpr hx] theorem spanSingleton_mul_inv {x : K} (hx : x ≠ 0) : spanSingleton R₁⁰ x * (spanSingleton R₁⁰ x)⁻¹ = 1 := by @@ -153,7 +153,7 @@ theorem coe_ideal_span_singleton_mul_inv {x : R₁} (hx : x ≠ 0) : (Ideal.span ({x} : Set R₁) : FractionalIdeal R₁⁰ K)⁻¹ = 1 := by rw [coeIdeal_span_singleton, spanSingleton_mul_inv K <| - (map_ne_zero_iff _ <| NoZeroSMulDivisors.algebraMap_injective R₁ K).mpr hx] + (map_ne_zero_iff _ <| FaithfulSMul.algebraMap_injective R₁ K).mpr hx] theorem spanSingleton_inv_mul {x : K} (hx : x ≠ 0) : (spanSingleton R₁⁰ x)⁻¹ * spanSingleton R₁⁰ x = 1 := by @@ -215,7 +215,7 @@ lemma den_mem_inv {I : FractionalIdeal R₁⁰ K} (hI : I ≠ ⊥) : lemma num_le_mul_inv (I : FractionalIdeal R₁⁰ K) : I.num ≤ I * I⁻¹ := by by_cases hI : I = 0 - · rw [hI, num_zero_eq <| NoZeroSMulDivisors.algebraMap_injective R₁ K, zero_mul, zero_eq_bot, + · rw [hI, num_zero_eq <| FaithfulSMul.algebraMap_injective R₁ K, zero_mul, zero_eq_bot, coeIdeal_bot] · rw [mul_comm, ← den_mul_self_eq_num'] exact mul_right_mono I <| spanSingleton_le_iff_mem.2 (den_mem_inv hI) @@ -592,9 +592,9 @@ theorem Ideal.dvd_iff_le {I J : Ideal A} : I ∣ J ↔ J ≤ I := · have hJ : J = ⊥ := by rwa [hI, ← eq_bot_iff] at h rw [hI, hJ] have hI' : (I : FractionalIdeal A⁰ (FractionRing A)) ≠ 0 := coeIdeal_ne_zero.mpr hI - have : (I : FractionalIdeal A⁰ (FractionRing A))⁻¹ * J ≤ 1 := - le_trans (mul_left_mono (↑I)⁻¹ ((coeIdeal_le_coeIdeal _).mpr h)) - (le_of_eq (inv_mul_cancel₀ hI')) + have : (I : FractionalIdeal A⁰ (FractionRing A))⁻¹ * J ≤ 1 := by + rw [← inv_mul_cancel₀ hI'] + exact mul_left_mono _ ((coeIdeal_le_coeIdeal _).mpr h) obtain ⟨H, hH⟩ := le_one_iff_exists_coeIdeal.mp this use H refine coeIdeal_injective (show (J : FractionalIdeal A⁰ (FractionRing A)) = ↑(I * H) from ?_) diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 22c54ebc367361..d6361c0e303a4e 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -63,12 +63,12 @@ theorem IsIntegralClosure.isLocalization [IsDomain A] [Algebra.IsAlgebraic K L] IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := by haveI : IsDomain C := (IsIntegralClosure.equiv A C L (integralClosure A L)).toMulEquiv.isDomain (integralClosure A L) - haveI : NoZeroSMulDivisors A L := NoZeroSMulDivisors.trans A K L + haveI : NoZeroSMulDivisors A L := NoZeroSMulDivisors.trans_faithfulSMul A K L haveI : NoZeroSMulDivisors A C := IsIntegralClosure.noZeroSMulDivisors A L refine ⟨?_, fun z => ?_, fun {x y} h => ⟨1, ?_⟩⟩ · rintro ⟨_, x, hx, rfl⟩ rw [isUnit_iff_ne_zero, map_ne_zero_iff _ (IsIntegralClosure.algebraMap_injective C A L), - Subtype.coe_mk, map_ne_zero_iff _ (NoZeroSMulDivisors.algebraMap_injective A C)] + Subtype.coe_mk, map_ne_zero_iff _ (FaithfulSMul.algebraMap_injective A C)] exact mem_nonZeroDivisors_iff_ne_zero.mp hx · obtain ⟨m, hm⟩ := IsIntegral.exists_multiple_integral_of_isLocalization A⁰ z @@ -247,4 +247,9 @@ instance integralClosure.isDedekindDomain_fractionRing [IsDedekindDomain A] : IsDedekindDomain (integralClosure A L) := integralClosure.isDedekindDomain A (FractionRing A) L +attribute [local instance] FractionRing.liftAlgebra in +instance [NoZeroSMulDivisors A C] [Module.Finite A C] [IsIntegrallyClosed C] : + IsLocalization (Algebra.algebraMapSubmonoid C A⁰) (FractionRing C) := + IsIntegralClosure.isLocalization _ (FractionRing A) _ _ + end IsIntegralClosure diff --git a/Mathlib/RingTheory/DedekindDomain/PID.lean b/Mathlib/RingTheory/DedekindDomain/PID.lean index b9bad1bafa86b2..393d049b6791e2 100644 --- a/Mathlib/RingTheory/DedekindDomain/PID.lean +++ b/Mathlib/RingTheory/DedekindDomain/PID.lean @@ -183,7 +183,7 @@ theorem IsPrincipalIdealRing.of_finite_primes [IsDedekindDomain R] variable [IsDedekindDomain R] variable (S : Type*) [CommRing S] -variable [Algebra R S] [Module.Free R S] [Module.Finite R S] +variable [Algebra R S] [NoZeroSMulDivisors R S] [Module.Finite R S] variable (p : Ideal R) (hp0 : p ≠ ⊥) [IsPrime p] variable {Sₚ : Type*} [CommRing Sₚ] [Algebra S Sₚ] variable [IsLocalization (Algebra.algebraMapSubmonoid S p.primeCompl) Sₚ] @@ -200,7 +200,7 @@ theorem IsLocalization.OverPrime.mem_normalizedFactors_of_isPrime [IsDomain S] {P : Ideal Sₚ} (hP : IsPrime P) (hP0 : P ≠ ⊥) : P ∈ normalizedFactors (Ideal.map (algebraMap R Sₚ) p) := by have non_zero_div : Algebra.algebraMapSubmonoid S p.primeCompl ≤ S⁰ := - map_le_nonZeroDivisors_of_injective _ (NoZeroSMulDivisors.algebraMap_injective _ _) + map_le_nonZeroDivisors_of_injective _ (FaithfulSMul.algebraMap_injective _ _) p.primeCompl_le_nonZeroDivisors letI : Algebra (Localization.AtPrime p) Sₚ := localizationAlgebra p.primeCompl S haveI : IsScalarTower R (Localization.AtPrime p) Sₚ := @@ -233,7 +233,7 @@ theorem IsLocalization.OverPrime.mem_normalizedFactors_of_isPrime [IsDomain S] · assumption rw [IsScalarTower.algebraMap_eq R S Sₚ] exact - (IsLocalization.injective Sₚ non_zero_div).comp (NoZeroSMulDivisors.algebraMap_injective _ _) + (IsLocalization.injective Sₚ non_zero_div).comp (FaithfulSMul.algebraMap_injective _ _) /-- Let `p` be a prime in the Dedekind domain `R` and `S` be an integral extension of `R`, then the localization `Sₚ` of `S` at `p` is a PID. -/ diff --git a/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean b/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean index 95dc12c2a93416..63e7b60137c546 100644 --- a/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean +++ b/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean @@ -211,12 +211,12 @@ theorem fromUnit_ker [hn : Fact <| 0 < n] : rcases IsIntegrallyClosed.exists_algebraMap_eq_of_isIntegral_pow (R := R) (x := i) hn.out (hi.symm ▸ isIntegral_algebraMap) with ⟨i', rfl⟩ - rw [← map_mul, map_eq_one_iff _ <| NoZeroSMulDivisors.algebraMap_injective R K] at vi - rw [← map_mul, map_eq_one_iff _ <| NoZeroSMulDivisors.algebraMap_injective R K] at iv + rw [← map_mul, map_eq_one_iff _ <| FaithfulSMul.algebraMap_injective R K] at vi + rw [← map_mul, map_eq_one_iff _ <| FaithfulSMul.algebraMap_injective R K] at iv rw [Units.val_mk, ← map_pow] at hv exact ⟨⟨v', i', vi, iv⟩, by simpa only [Units.ext_iff, powMonoidHom_apply, Units.val_pow_eq_pow_val] using - NoZeroSMulDivisors.algebraMap_injective R K hv⟩ + FaithfulSMul.algebraMap_injective R K hv⟩ · rintro ⟨x, hx⟩ rw [← hx] exact Subtype.mk_eq_mk.mpr <| (QuotientGroup.eq_one_iff _).mpr ⟨Units.map (algebraMap R K) x, diff --git a/Mathlib/RingTheory/FiniteLength.lean b/Mathlib/RingTheory/FiniteLength.lean index 90ac5c06d9d02a..8b866ed99ebdea 100644 --- a/Mathlib/RingTheory/FiniteLength.lean +++ b/Mathlib/RingTheory/FiniteLength.lean @@ -34,9 +34,11 @@ variable {R} {M N : Type*} [AddCommGroup M] [Module R M] [AddCommGroup N] [Modul theorem LinearEquiv.isFiniteLength (e : M ≃ₗ[R] N) (h : IsFiniteLength R M) : IsFiniteLength R N := by - induction' h with M _ _ _ M _ _ S _ _ ih generalizing N - · have := e.symm.toEquiv.subsingleton; exact .of_subsingleton - · have : IsSimpleModule R (N ⧸ Submodule.map (e : M →ₗ[R] N) S) := + induction h generalizing N with + | of_subsingleton => + have := e.symm.toEquiv.subsingleton; exact .of_subsingleton + | @of_simple_quotient M _ _ S _ _ ih => + have : IsSimpleModule R (N ⧸ Submodule.map (e : M →ₗ[R] N) S) := IsSimpleModule.congr (Submodule.Quotient.equiv S _ e rfl).symm exact .of_simple_quotient (ih <| e.submoduleMap S) diff --git a/Mathlib/RingTheory/FractionalIdeal/Basic.lean b/Mathlib/RingTheory/FractionalIdeal/Basic.lean index 17d1057a60b9b8..c508c05e3b4061 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Basic.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Basic.lean @@ -132,7 +132,7 @@ noncomputable def equivNum [Nontrivial P] [NoZeroSMulDivisors R P] (LinearEquiv.ofBijective ((DistribMulAction.toLinearMap R P I.den).restrict fun _ hx ↦ ?_) ⟨fun _ _ hxy ↦ ?_, fun ⟨y, hy⟩ ↦ ?_⟩) (Submodule.equivMapOfInjective (Algebra.linearMap R P) - (NoZeroSMulDivisors.algebraMap_injective R P) (num I)).symm + (FaithfulSMul.algebraMap_injective R P) (num I)).symm · rw [← den_mul_self_eq_num] exact Submodule.smul_mem_pointwise_smul _ _ _ hx · simp_rw [LinearMap.restrict_apply, DistribMulAction.toLinearMap_apply, Subtype.mk.injEq] at hxy diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 67dad3b20e060e..4ef12677ebde3e 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -903,7 +903,7 @@ variable [CommRing R] [CommRing S] theorem map_ne_bot_of_ne_bot {S : Type*} [Ring S] [Nontrivial S] [Algebra R S] [NoZeroSMulDivisors R S] {I : Ideal R} (h : I ≠ ⊥) : map (algebraMap R S) I ≠ ⊥ := - (map_eq_bot_iff_of_injective (NoZeroSMulDivisors.algebraMap_injective R S)).mp.mt h + (map_eq_bot_iff_of_injective (FaithfulSMul.algebraMap_injective R S)).mp.mt h theorem map_eq_iff_sup_ker_eq_of_surjective {I J : Ideal R} (f : R →+* S) (hf : Function.Surjective f) : map f I = map f J ↔ I ⊔ RingHom.ker f = J ⊔ RingHom.ker f := by @@ -1039,18 +1039,16 @@ def idealMap (I : Ideal R) : I →ₗ[R] I.map (algebraMap R S) := 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) +@[simp] +theorem FaithfulSMul.ker_algebraMap_eq_bot (R A : Type*) [CommSemiring R] [Semiring A] + [Algebra R A] [FaithfulSMul R A] : RingHom.ker (algebraMap R A) = ⊥ := by + ext; simp -theorem ker_algebraMap_eq_bot (R A : Type*) [CommSemiring R] [Semiring A] [Nontrivial A] - [Algebra R A] [NoZeroSMulDivisors R A] : RingHom.ker (algebraMap R A) = ⊥ := by - ext; simp [Algebra.algebraMap_eq_smul_one'] +@[deprecated (since := "2025-01-31")] +alias NoZeroSMulDivisors.iff_ker_algebraMap_eq_bot := FaithfulSMul.ker_algebraMap_eq_bot -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)) +@[deprecated (since := "2025-01-31")] +alias NoZeroSMulDivisors.of_ker_algebraMap_eq_bot := FaithfulSMul.ker_algebraMap_eq_bot -end NoZeroSMulDivisors +@[deprecated (since := "2025-01-31")] +alias NoZeroSMulDivisors.ker_algebraMap_eq_bot := FaithfulSMul.ker_algebraMap_eq_bot diff --git a/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean b/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean index ccb7c0362da1eb..a650aa9e98e157 100644 --- a/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean +++ b/Mathlib/RingTheory/Ideal/Norm/RelNorm.lean @@ -6,6 +6,7 @@ Authors: Anne Baanen, Alex J. Best import Mathlib.LinearAlgebra.FreeModule.PID import Mathlib.RingTheory.DedekindDomain.PID import Mathlib.RingTheory.Localization.NormTrace +import Mathlib.RingTheory.IntegralClosure.IntegralRestrict /-! @@ -34,92 +35,132 @@ namespace Ideal open Submodule -variable (R : Type*) [CommRing R] {S : Type*} [CommRing S] [Algebra R S] +attribute [local instance] FractionRing.liftAlgebra -/-- `Ideal.spanNorm R (I : Ideal S)` is the ideal generated by mapping `Algebra.norm R` over `I`. +variable (R S : Type*) [CommRing R] [IsDomain R] {S : Type*} [CommRing S] [IsDomain S] +variable [IsIntegrallyClosed R] [IsIntegrallyClosed S] [Algebra R S] [Module.Finite R S] +variable [NoZeroSMulDivisors R S] [Algebra.IsSeparable (FractionRing R) (FractionRing S)] + +/-- `Ideal.spanNorm R (I : Ideal S)` is the ideal generated by mapping `Algebra.intNorm R S` +over `I`. See also `Ideal.relNorm`. -/ def spanNorm (I : Ideal S) : Ideal R := - Ideal.span (Algebra.norm R '' (I : Set S)) + Ideal.map (Algebra.intNorm R S) I @[simp] -theorem spanNorm_bot [Nontrivial S] [Module.Free R S] [Module.Finite R S] : +theorem spanNorm_bot : spanNorm R (⊥ : Ideal S) = ⊥ := span_eq_bot.mpr fun x hx => by simpa using hx variable {R} @[simp] -theorem spanNorm_eq_bot_iff [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S] - {I : Ideal S} : spanNorm R I = ⊥ ↔ I = ⊥ := by - simp only [spanNorm, Ideal.span_eq_bot, Set.mem_image, SetLike.mem_coe, forall_exists_index, - and_imp, forall_apply_eq_imp_iff₂, - Algebra.norm_eq_zero_iff_of_basis (Module.Free.chooseBasis R S), @eq_bot_iff _ _ _ I, - SetLike.le_def] +theorem spanNorm_eq_bot_iff {I : Ideal S} : spanNorm R I = ⊥ ↔ I = ⊥ := by + simp only [spanNorm, span_eq_bot, Set.mem_image, SetLike.mem_coe, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂, Algebra.intNorm_eq_zero, @eq_bot_iff _ _ _ I, SetLike.le_def, map] rfl variable (R) -theorem norm_mem_spanNorm {I : Ideal S} (x : S) (hx : x ∈ I) : Algebra.norm R x ∈ I.spanNorm R := +theorem intNorm_mem_spanNorm {I : Ideal S} {x : S} (hx : x ∈ I) : + Algebra.intNorm R S x ∈ I.spanNorm R := subset_span (Set.mem_image_of_mem _ hx) +theorem norm_mem_spanNorm [Module.Free R S] {I : Ideal S} (x : S) (hx : x ∈ I) : + Algebra.norm R x ∈ I.spanNorm R := by + refine subset_span ⟨x, hx, ?_⟩ + rw [Algebra.intNorm_eq_norm] + @[simp] -theorem spanNorm_singleton {r : S} : spanNorm R (span ({r} : Set S)) = span {Algebra.norm R r} := +theorem spanNorm_singleton {r : S} : + spanNorm R (span ({r} : Set S)) = span {Algebra.intNorm R S r} := le_antisymm (span_le.mpr fun x hx => mem_span_singleton.mpr (by obtain ⟨x, hx', rfl⟩ := (Set.mem_image _ _ _).mp hx exact map_dvd _ (mem_span_singleton.mp hx'))) - ((span_singleton_le_iff_mem _).mpr (norm_mem_spanNorm _ _ (mem_span_singleton_self _))) + ((span_singleton_le_iff_mem _).mpr (intNorm_mem_spanNorm _ (mem_span_singleton_self _))) @[simp] theorem spanNorm_top : spanNorm R (⊤ : Ideal S) = ⊤ := by - -- Porting note: was - -- simp [← Ideal.span_singleton_one] - rw [← Ideal.span_singleton_one, spanNorm_singleton] - simp - -theorem map_spanNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : - map f (spanNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := by - rw [spanNorm, map_span, Set.image_image] - -- Porting note: `Function.comp` reducibility - rfl + simp [← Ideal.span_singleton_one] + +theorem map_spanIntNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : + map f (spanNorm R I) = span (f ∘ Algebra.intNorm R S '' (I : Set S)) := by + rw [spanNorm] + nth_rw 2 [map] + simp [map_span, Set.image_image] @[mono] theorem spanNorm_mono {I J : Ideal S} (h : I ≤ J) : spanNorm R I ≤ spanNorm R J := Ideal.span_mono (Set.monotone_image h) -theorem spanNorm_localization (I : Ideal S) [Module.Finite R S] [Module.Free R S] (M : Submonoid R) +theorem spanIntNorm_localization (I : Ideal S) (M : Submonoid R) (hM : M ≤ R⁰) {Rₘ : Type*} (Sₘ : Type*) [CommRing Rₘ] [Algebra R Rₘ] [CommRing Sₘ] [Algebra S Sₘ] [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] - [IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] : + [IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] + [IsIntegrallyClosed Rₘ] [IsDomain Rₘ] [IsDomain Sₘ] [NoZeroSMulDivisors Rₘ Sₘ] + [Module.Finite Rₘ Sₘ] [IsIntegrallyClosed Sₘ] + [Algebra.IsSeparable (FractionRing Rₘ) (FractionRing Sₘ)] : spanNorm Rₘ (I.map (algebraMap S Sₘ)) = (spanNorm R I).map (algebraMap R Rₘ) := by - cases subsingleton_or_nontrivial R - · haveI := IsLocalization.unique R Rₘ M - simp [eq_iff_true_of_subsingleton] - let b := Module.Free.chooseBasis R S - rw [map_spanNorm] + let K := FractionRing R + let f : Rₘ →+* K := IsLocalization.map _ (T := R⁰) (RingHom.id R) hM + let L := FractionRing S + let g : Sₘ →+* L := IsLocalization.map _ (M := Algebra.algebraMapSubmonoid S M) (T := S⁰) + (RingHom.id S) (Submonoid.map_le_of_le_comap _ <| hM.trans + (nonZeroDivisors_le_comap_nonZeroDivisors_of_injective _ + (FaithfulSMul.algebraMap_injective _ _))) + algebraize [f, g, (algebraMap K L).comp f] + have : IsScalarTower R Rₘ K := IsScalarTower.of_algebraMap_eq' + (by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHomCompTriple.comp_eq]) + let _ := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization M Rₘ K + have : IsScalarTower S Sₘ L := IsScalarTower.of_algebraMap_eq' + (by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHomCompTriple.comp_eq]) + have : IsScalarTower Rₘ Sₘ L := by + apply IsScalarTower.of_algebraMap_eq' + apply IsLocalization.ringHom_ext M + rw [RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra (R := Sₘ), RingHom.comp_assoc, + RingHom.comp_assoc, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R S Sₘ, + IsLocalization.map_comp, RingHom.comp_id, ← RingHom.comp_assoc, IsLocalization.map_comp, + RingHom.comp_id, ← IsScalarTower.algebraMap_eq, ← IsScalarTower.algebraMap_eq] + let _ := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization + (Algebra.algebraMapSubmonoid S M) Sₘ L + have : IsIntegralClosure Sₘ Rₘ L := + IsIntegralClosure.of_isIntegrallyClosed _ _ _ + rw [map_spanIntNorm] refine span_eq_span (Set.image_subset_iff.mpr ?_) (Set.image_subset_iff.mpr ?_) - · rintro a' ha' - simp only [Set.mem_preimage, submodule_span_eq, ← map_spanNorm, SetLike.mem_coe, + · intro a' ha' + simp only [Set.mem_preimage, submodule_span_eq, ← map_spanIntNorm, SetLike.mem_coe, IsLocalization.mem_map_algebraMap_iff (Algebra.algebraMapSubmonoid S M) Sₘ, IsLocalization.mem_map_algebraMap_iff M Rₘ, Prod.exists] at ha' ⊢ obtain ⟨⟨a, ha⟩, ⟨_, ⟨s, hs, rfl⟩⟩, has⟩ := ha' - refine ⟨⟨Algebra.norm R a, norm_mem_spanNorm _ _ ha⟩, - ⟨s ^ Fintype.card (Module.Free.ChooseBasisIndex R S), pow_mem hs _⟩, ?_⟩ + refine ⟨⟨Algebra.intNorm R S a, intNorm_mem_spanNorm _ ha⟩, + ⟨s ^ Module.finrank K L, pow_mem hs _⟩, ?_⟩ simp only [Submodule.coe_mk, Subtype.coe_mk, map_pow] at has ⊢ - apply_fun Algebra.norm Rₘ at has - rwa [_root_.map_mul, ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R Rₘ, - Algebra.norm_algebraMap_of_basis (b.localizationLocalization Rₘ M Sₘ), - Algebra.norm_localization R M a] at has + apply_fun algebraMap _ L at has + apply_fun Algebra.norm K at has + simp only [_root_.map_mul, IsScalarTower.algebraMap_apply R Rₘ Sₘ] at has + rw [← IsScalarTower.algebraMap_apply, ← IsScalarTower.algebraMap_apply, + ← IsScalarTower.algebraMap_apply, + IsScalarTower.algebraMap_apply R K L, + Algebra.norm_algebraMap] at has + apply IsFractionRing.injective Rₘ K + simp only [_root_.map_mul, map_pow] + have : FiniteDimensional K L := Module.Finite_of_isLocalization R S _ _ R⁰ + rwa [Algebra.algebraMap_intNorm (L := L), ← IsScalarTower.algebraMap_apply, + ← IsScalarTower.algebraMap_apply, Algebra.algebraMap_intNorm (L := L)] · intro a ha - rw [Set.mem_preimage, Function.comp_apply, ← Algebra.norm_localization (Sₘ := Sₘ) R M a] + rw [Set.mem_preimage, Function.comp_apply, Algebra.intNorm_eq_of_isLocalization + (A := R) (B := S) M (Aₘ := Rₘ) (Bₘ := Sₘ)] exact subset_span (Set.mem_image_of_mem _ (mem_map_of_mem _ ha)) theorem spanNorm_mul_spanNorm_le (I J : Ideal S) : spanNorm R I * spanNorm R J ≤ spanNorm R (I * J) := by - rw [spanNorm, spanNorm, spanNorm, Ideal.span_mul_span', ← Set.image_mul] + rw [spanNorm, spanNorm, spanNorm] + nth_rw 1 [map]; nth_rw 1 [map] + rw [Ideal.span_mul_span', ← Set.image_mul] refine Ideal.span_mono (Set.monotone_image ?_) rintro _ ⟨x, hxI, y, hyJ, rfl⟩ exact Ideal.mul_mem_mul hxI hyJ @@ -127,60 +168,73 @@ theorem spanNorm_mul_spanNorm_le (I J : Ideal S) : /-- This condition `eq_bot_or_top` is equivalent to being a field. However, `Ideal.spanNorm_mul_of_field` is harder to apply since we'd need to upgrade a `CommRing R` instance to a `Field R` instance. -/ -theorem spanNorm_mul_of_bot_or_top [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S] - (eq_bot_or_top : ∀ I : Ideal R, I = ⊥ ∨ I = ⊤) (I J : Ideal S) : +theorem spanNorm_mul_of_bot_or_top (eq_bot_or_top : ∀ I : Ideal R, I = ⊥ ∨ I = ⊤) (I J : Ideal S) : spanNorm R (I * J) = spanNorm R I * spanNorm R J := by refine le_antisymm ?_ (spanNorm_mul_spanNorm_le R _ _) - cases' eq_bot_or_top (spanNorm R I) with hI hI + rcases eq_bot_or_top (spanNorm R I) with hI | hI · rw [hI, spanNorm_eq_bot_iff.mp hI, bot_mul, spanNorm_bot] exact bot_le rw [hI, Ideal.top_mul] - cases' eq_bot_or_top (spanNorm R J) with hJ hJ + rcases eq_bot_or_top (spanNorm R J) with hJ | hJ · rw [hJ, spanNorm_eq_bot_iff.mp hJ, mul_bot, spanNorm_bot] rw [hJ] exact le_top -@[simp] -theorem spanNorm_mul_of_field {K : Type*} [Field K] [Algebra K S] [IsDomain S] [Module.Finite K S] - (I J : Ideal S) : spanNorm K (I * J) = spanNorm K I * spanNorm K J := - spanNorm_mul_of_bot_or_top K eq_bot_or_top I J - -variable [IsDomain R] [IsDomain S] [IsDedekindDomain R] [IsDedekindDomain S] -variable [Module.Finite R S] [Module.Free R S] +variable [IsDedekindDomain R] [IsDedekindDomain S] /-- Multiplicativity of `Ideal.spanNorm`. simp-normal form is `map_mul (Ideal.relNorm R)`. -/ theorem spanNorm_mul (I J : Ideal S) : spanNorm R (I * J) = spanNorm R I * spanNorm R J := by nontriviality R cases subsingleton_or_nontrivial S - · have : ∀ I : Ideal S, I = ⊤ := fun I => Subsingleton.elim I ⊤ + · have : ∀ I : Ideal S, I = ⊤ := fun I ↦ Subsingleton.elim I ⊤ simp [this I, this J, this (I * J)] - refine eq_of_localization_maximal ?_ - intro P hP + refine eq_of_localization_maximal (fun P hP ↦ ?_) by_cases hP0 : P = ⊥ · subst hP0 rw [spanNorm_mul_of_bot_or_top] intro I - refine or_iff_not_imp_right.mpr fun hI => ?_ - exact (hP.eq_of_le hI bot_le).symm + exact or_iff_not_imp_right.mpr fun hI ↦ (hP.eq_of_le hI bot_le).symm let P' := Algebra.algebraMapSubmonoid S P.primeCompl - letI : Algebra (Localization.AtPrime P) (Localization P') := localizationAlgebra P.primeCompl S - haveI : IsScalarTower R (Localization.AtPrime P) (Localization P') := + let Rₚ := Localization.AtPrime P + let Sₚ := Localization P' + let _ : Algebra Rₚ Sₚ := localizationAlgebra P.primeCompl S + have : IsScalarTower R Rₚ Sₚ := IsScalarTower.of_algebraMap_eq (fun x => (IsLocalization.map_eq (T := P') (Q := Localization P') P.primeCompl.le_comap_map x).symm) have h : P' ≤ S⁰ := - map_le_nonZeroDivisors_of_injective _ (NoZeroSMulDivisors.algebraMap_injective _ _) + map_le_nonZeroDivisors_of_injective _ (FaithfulSMul.algebraMap_injective _ _) P.primeCompl_le_nonZeroDivisors - haveI : IsDomain (Localization P') := IsLocalization.isDomain_localization h - haveI : IsDedekindDomain (Localization P') := IsLocalization.isDedekindDomain S h _ - letI := Classical.decEq (Ideal (Localization P')) - haveI : IsPrincipalIdealRing (Localization P') := + have : IsDomain Sₚ := IsLocalization.isDomain_localization h + have : IsDedekindDomain Sₚ := IsLocalization.isDedekindDomain S h _ + have : IsPrincipalIdealRing Sₚ := IsDedekindDomain.isPrincipalIdealRing_localization_over_prime S P hP0 - rw [Ideal.map_mul, ← spanNorm_localization R I P.primeCompl (Localization P'), - ← spanNorm_localization R J P.primeCompl (Localization P'), - ← spanNorm_localization R (I * J) P.primeCompl (Localization P'), Ideal.map_mul, - ← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator, + have := NoZeroSMulDivisors_of_isLocalization R S Rₚ Sₚ P.primeCompl_le_nonZeroDivisors + have := Module.Finite_of_isLocalization R S Rₚ Sₚ P.primeCompl + let L := FractionRing S + let g : Sₚ →+* L := IsLocalization.map _ (M := P') (T := S⁰) (RingHom.id S) h + algebraize [g] + have : IsScalarTower S Sₚ (FractionRing S) := IsScalarTower.of_algebraMap_eq' + (by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHom.comp_id]) + have := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization P' Sₚ (FractionRing S) + have : Algebra.IsSeparable (FractionRing Rₚ) (FractionRing Sₚ) := by + apply Algebra.IsSeparable.of_equiv_equiv + (FractionRing.algEquiv Rₚ (FractionRing R)).symm.toRingEquiv + (FractionRing.algEquiv Sₚ (FractionRing S)).symm.toRingEquiv + apply IsLocalization.ringHom_ext R⁰ + ext + simp only [AlgEquiv.toRingEquiv_eq_coe, RingHom.coe_comp, + RingHom.coe_coe, Function.comp_apply, ← IsScalarTower.algebraMap_apply] + rw [IsScalarTower.algebraMap_apply R Rₚ (FractionRing R), AlgEquiv.coe_ringEquiv, + AlgEquiv.commutes, IsScalarTower.algebraMap_apply R S L, + IsScalarTower.algebraMap_apply S Sₚ L, AlgEquiv.coe_ringEquiv, AlgEquiv.commutes] + simp only [← IsScalarTower.algebraMap_apply] + rw [IsScalarTower.algebraMap_apply R Rₚ (FractionRing Rₚ), + ← IsScalarTower.algebraMap_apply Rₚ, ← IsScalarTower.algebraMap_apply] + simp only [Ideal.map_mul, ← spanIntNorm_localization (R := R) (S := S) + (Rₘ := Localization.AtPrime P) (Sₘ := Localization P') _ _ P.primeCompl_le_nonZeroDivisors] + rw [← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator, span_singleton_mul_span_singleton, spanNorm_singleton, spanNorm_singleton, - spanNorm_singleton, span_singleton_mul_span_singleton, _root_.map_mul] + spanNorm_singleton, span_singleton_mul_span_singleton, _root_.map_mul] /-- The relative norm `Ideal.relNorm R (I : Ideal S)`, where `R` and `S` are Dedekind domains, and `S` is an extension of `R` that is finite and free as a module. -/ @@ -190,7 +244,8 @@ def relNorm : Ideal S →*₀ Ideal R where map_one' := by dsimp only; rw [one_eq_top, spanNorm_top R, one_eq_top] map_mul' := spanNorm_mul R -theorem relNorm_apply (I : Ideal S) : relNorm R I = span (Algebra.norm R '' (I : Set S) : Set R) := +theorem relNorm_apply (I : Ideal S) : + relNorm R I = span (Algebra.intNorm R S '' (I : Set S) : Set R) := rfl @[simp] @@ -212,16 +267,17 @@ theorem relNorm_eq_bot_iff {I : Ideal S} : relNorm R I = ⊥ ↔ I = ⊥ := variable (R) -theorem norm_mem_relNorm (I : Ideal S) {x : S} (hx : x ∈ I) : Algebra.norm R x ∈ relNorm R I := +theorem norm_mem_relNorm [Module.Free R S] (I : Ideal S) {x : S} (hx : x ∈ I) : + Algebra.norm R x ∈ relNorm R I := norm_mem_spanNorm R x hx @[simp] -theorem relNorm_singleton (r : S) : relNorm R (span ({r} : Set S)) = span {Algebra.norm R r} := +theorem relNorm_singleton (r : S) : relNorm R (span ({r} : Set S)) = span {Algebra.intNorm R S r} := spanNorm_singleton R theorem map_relNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) : - map f (relNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := - map_spanNorm R I f + map f (relNorm R I) = span (f ∘ Algebra.intNorm R S '' (I : Set S)) := + map_spanIntNorm R I f @[mono] theorem relNorm_mono {I J : Ideal S} (h : I ≤ J) : relNorm R I ≤ relNorm R J := diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index 48ef1022ba6d00..a48fdd2dbd260f 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -541,7 +541,7 @@ variable (A : Type*) [CommRing A] (B : Type*) [Ring B] [Nontrivial B] @[simp] theorem under_bot : under A (⊥ : Ideal B) = ⊥ := - comap_bot_of_injective (algebraMap A B) (NoZeroSMulDivisors.algebraMap_injective A B) + comap_bot_of_injective (algebraMap A B) (FaithfulSMul.algebraMap_injective A B) instance bot_liesOver_bot : (⊥ : Ideal B).LiesOver (⊥ : Ideal A) where over := (under_bot A B).symm @@ -583,14 +583,15 @@ instance algebra_finiteType_of_liesOver [Algebra.FiniteType A B] : 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 +instance instFaithfulSMul : FaithfulSMul (A ⧸ p) (B ⧸ P) := by + rw [faithfulSMul_iff_algebraMap_injective] 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 (algebraMap_injective_of_liesOver P p) +@[deprecated (since := "2025-01-31")] +alias algebraMap_injective_of_liesOver := instFaithfulSMul variable {p} in theorem nontrivial_of_liesOver_of_ne_top (hp : p ≠ ⊤) : Nontrivial (B ⧸ P) := @@ -668,8 +669,7 @@ instance Quotient.algebra_isIntegral_of_liesOver : Algebra.IsIntegral (A ⧸ p) theorem exists_ideal_liesOver_maximal_of_isIntegral [p.IsMaximal] (B : Type*) [CommRing B] [Nontrivial B] [Algebra A B] [NoZeroSMulDivisors A B] [Algebra.IsIntegral A B] : ∃ P : Ideal B, P.IsMaximal ∧ P.LiesOver p := by - rcases exists_ideal_over_maximal_of_isIntegral p <| - (NoZeroSMulDivisors.ker_algebraMap_eq_bot A B).trans_le bot_le with ⟨P, hm, hP⟩ + obtain ⟨P, hm, hP⟩ := exists_ideal_over_maximal_of_isIntegral (S := B) p <| by simp exact ⟨P, hm, ⟨hP.symm⟩⟩ end IsIntegral diff --git a/Mathlib/RingTheory/Ideal/Quotient/Index.lean b/Mathlib/RingTheory/Ideal/Quotient/Index.lean new file mode 100644 index 00000000000000..d4f8ace34aac5d --- /dev/null +++ b/Mathlib/RingTheory/Ideal/Quotient/Index.lean @@ -0,0 +1,104 @@ +/- +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.Algebra.GeomSum +import Mathlib.Data.Finsupp.Fintype +import Mathlib.GroupTheory.Index +import Mathlib.LinearAlgebra.DirectSum.Finsupp +import Mathlib.LinearAlgebra.TensorProduct.Quotient +import Mathlib.LinearAlgebra.TensorProduct.RightExactness +import Mathlib.RingTheory.Finiteness.Cardinality +import Mathlib.RingTheory.Finiteness.TensorProduct +import Mathlib.RingTheory.Ideal.Quotient.Operations + +/-! +# Indices of ideals + +## Main result +- `Submodule.finite_quotient_smul`: + Let `N` be a finite index f.g. `R`-submodule, and `I` be a finite index ideal. + Then `I • N` also has finite index. +- `Ideal.index_quotient_pow_le`: + If `I` is generated by `k` elements, + the index of `I ^ n` is bounded by `#(R ⧸ I) ^ (k⁰ + k¹ + ⋯ + kⁿ⁻¹)`. + +-/ + +variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] +variable (I : Ideal R) {N : Submodule R M} + +open TensorProduct in +/-- Let `N` be a finite index f.g. `R`-submodule, and `I` be a finite index ideal. +Then `I • N` also has finite index. -/ +lemma Submodule.finite_quotient_smul [Finite (R ⧸ I)] [Finite (M ⧸ N)] (hN : N.FG) : + Finite (M ⧸ I • N) := by + suffices (I • N).toAddSubgroup.FiniteIndex by + exact (I • N).toAddSubgroup.finite_quotient_of_finiteIndex + suffices Nat.card (N ⧸ (I • N).comap N.subtype) ≠ 0 by + constructor + rw [← AddSubgroup.relindex_mul_index + (H := (I • N).toAddSubgroup) (K := N.toAddSubgroup) Submodule.smul_le_right] + have inst : Finite (M ⧸ N.toAddSubgroup) := ‹_› + exact mul_ne_zero this AddSubgroup.index_ne_zero_of_finite + let e : (N ⧸ (I • N).comap N.subtype) ≃ₗ[R] (R ⧸ I) ⊗[R] N := + Submodule.quotEquivOfEq _ (I • (⊤ : Submodule R N)) (Submodule.map_injective_of_injective + N.injective_subtype (by simp [Submodule.smul_le_right])) ≪≫ₗ + (quotTensorEquivQuotSMul N I).symm + rw [Nat.card_congr e.toEquiv] + have : Module.Finite R N := Module.Finite.iff_fg.mpr hN + have : Finite ((R ⧸ I) ⊗[R] N) := Module.finite_of_finite (R ⧸ I) + exact Nat.card_pos.ne' + +-- We have `hs` and `N` instead of using `span R s` in the goal to make it easier to use. +-- Usually we would like to bound the index of some abstract `I • N`, and we may construct `s` while +-- applying this lemma instead of having to provide it beforehand. +open TensorProduct in +lemma Submodule.index_smul_le [Finite (R ⧸ I)] + (s : Finset M) (hs : Submodule.span R s = N) : + (I • N).toAddSubgroup.index ≤ I.toAddSubgroup.index ^ s.card * N.toAddSubgroup.index := by + classical + cases nonempty_fintype (R ⧸ I) + rw [← AddSubgroup.relindex_mul_index + (H := (I • N).toAddSubgroup) (K := N.toAddSubgroup) Submodule.smul_le_right] + gcongr + show (Nat.card (N ⧸ (I • N).comap N.subtype)) ≤ Nat.card (R ⧸ I) ^ s.card + let e : (N ⧸ (I • N).comap N.subtype) ≃ₗ[R] (R ⧸ I) ⊗[R] N := + Submodule.quotEquivOfEq _ (I • (⊤ : Submodule R N)) (Submodule.map_injective_of_injective + N.injective_subtype (by simp [Submodule.smul_le_right])) ≪≫ₗ + (quotTensorEquivQuotSMul N I).symm + rw [Nat.card_congr e.toEquiv] + have H : LinearMap.range (Finsupp.linearCombination R (α := s) (↑)) = N := by + rw [Finsupp.range_linearCombination, ← hs, Subtype.range_val]; rfl + let f : (s →₀ R) →ₗ[R] N := (Finsupp.linearCombination R (↑)).codRestrict _ + (Set.range_subset_iff (s := N.carrier).mp <| by exact H.le) + have hf : Function.Surjective f := fun x ↦ by + obtain ⟨y, hy⟩ := H.ge x.2; exact ⟨y, Subtype.ext hy⟩ + have : Function.Surjective + (f.lTensor (R ⧸ I) ∘ₗ (finsuppScalarRight R (R ⧸ I) s).symm.toLinearMap) := + (LinearMap.lTensor_surjective (R ⧸ I) hf).comp (LinearEquiv.surjective _) + refine (Nat.card_le_card_of_surjective _ this).trans ?_ + simp only [Nat.card_eq_fintype_card, Fintype.card_finsupp, Fintype.card_coe, le_rfl] + +variable {I} + +lemma Ideal.finite_quotient_pow (hI : I.FG) [Finite (R ⧸ I)] (n) : Finite (R ⧸ I ^ n) := by + induction n with + | zero => + simp only [pow_zero, Ideal.one_eq_top] + infer_instance + | succ n _ => + exact Submodule.finite_quotient_smul (I ^ n) hI + +lemma Ideal.index_pow_le + (s : Finset R) (hs : Ideal.span s = I) [Finite (R ⧸ I)] (n) : + (I ^ n).toAddSubgroup.index ≤ I.toAddSubgroup.index ^ ∑ i ∈ Finset.range n, s.card ^ i := by + have := Ideal.finite_quotient_pow ⟨s, hs⟩ + induction n with + | zero => + simp + | succ n IH => + refine (Submodule.index_smul_le (I ^ n) s hs).trans ?_ + refine (Nat.mul_le_mul (Nat.pow_le_pow_left IH _) le_rfl).trans ?_ + rw [← pow_mul, ← pow_succ, geom_sum_succ, mul_comm] diff --git a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean index a71d12054e32b8..8b758cb2564b10 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Operations.lean @@ -188,11 +188,11 @@ def quotientInfToPiQuotient (I : ι → Ideal R) : (R ⧸ ⨅ i, I i) →+* ∀ lemma quotientInfToPiQuotient_mk (I : ι → Ideal R) (x : R) : quotientInfToPiQuotient I (Quotient.mk _ x) = fun i : ι ↦ Quotient.mk (I i) x := -rfl + rfl lemma quotientInfToPiQuotient_mk' (I : ι → Ideal R) (x : R) (i : ι) : quotientInfToPiQuotient I (Quotient.mk _ x) i = Quotient.mk (I i) x := -rfl + rfl lemma quotientInfToPiQuotient_inj (I : ι → Ideal R) : Injective (quotientInfToPiQuotient I) := by rw [quotientInfToPiQuotient, injective_lift_iff, ker_Pi_Quotient_mk] diff --git a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean index b820b75229efdd..63f417f2adbe59 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean @@ -39,13 +39,13 @@ This is inverse to the restriction. See `galRestrictHom`. -/ noncomputable def galLift (σ : B →ₐ[A] B) : L →ₐ[K] L := haveI := (IsFractionRing.injective A K).isDomain - haveI := NoZeroSMulDivisors.trans A K L + haveI := NoZeroSMulDivisors.trans_faithfulSMul A K L haveI := IsIntegralClosure.isLocalization A K L B haveI H : ∀ (y : Algebra.algebraMapSubmonoid B A⁰), IsUnit (((algebraMap B L).comp σ) (y : B)) := by rintro ⟨_, x, hx, rfl⟩ simpa only [RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, AlgHom.commutes, - isUnit_iff_ne_zero, ne_eq, map_eq_zero_iff _ (NoZeroSMulDivisors.algebraMap_injective _ _), + isUnit_iff_ne_zero, ne_eq, map_eq_zero_iff _ (FaithfulSMul.algebraMap_injective _ _), ← IsScalarTower.algebraMap_apply] using nonZeroDivisors.ne_zero hx haveI H_eq : (IsLocalization.lift (S := L) H).comp (algebraMap K L) = (algebraMap K L) := by apply IsLocalization.ringHom_ext A⁰ diff --git a/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean b/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean index 42fff4d0e0234e..f567814549c9d7 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean @@ -165,7 +165,7 @@ variable (R) @[simp] theorem integralClosure_eq_bot [IsIntegrallyClosedIn R A] [NoZeroSMulDivisors R A] [Nontrivial A] : integralClosure R A = ⊥ := - (integralClosure_eq_bot_iff A (NoZeroSMulDivisors.algebraMap_injective _ _)).mpr ‹_› + (integralClosure_eq_bot_iff A (FaithfulSMul.algebraMap_injective _ _)).mpr ‹_› variable {A} {B : Type*} [CommRing B] diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index ff4bf59739f942..abfec7850aa3db 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -551,7 +551,7 @@ theorem Algebra.IsIntegral.tower_bot [Algebra R S] [Algebra R T] [Algebra 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) + (FaithfulSMul.algebraMap_injective S T) rw [← IsScalarTower.algebraMap_eq R S T] exact h.isIntegral diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index d88c930f3e8eec..57d9cf6e6bfb41 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -89,8 +89,8 @@ variable [Ring R] [IsDomain R] [Fintype R] `Mathlib.RingTheory.LittleWedderburn`. -/ def Fintype.divisionRingOfIsDomain (R : Type*) [Ring R] [IsDomain R] [DecidableEq R] [Fintype R] : DivisionRing R where + __ := (‹Ring R›:) -- this also works without the `( :)`, but it's slightly slow __ := Fintype.groupWithZeroOfCancel R - __ := ‹Ring R› nnqsmul := _ nnqsmul_def := fun _ _ => rfl qsmul := _ diff --git a/Mathlib/RingTheory/Invariant.lean b/Mathlib/RingTheory/Invariant.lean index fe8623572f7508..1da0a8bb04211b 100644 --- a/Mathlib/RingTheory/Invariant.lean +++ b/Mathlib/RingTheory/Invariant.lean @@ -72,11 +72,11 @@ theorem Algebra.isInvariant_of_isGalois [FiniteDimensional K L] [h : IsGalois K rw [h, IntermediateField.mem_bot] at hb obtain ⟨k, hk⟩ := hb have hb : IsIntegral A b := IsIntegralClosure.isIntegral A L b - rw [← isIntegral_algebraMap_iff (NoZeroSMulDivisors.algebraMap_injective B L), ← hk, - isIntegral_algebraMap_iff (NoZeroSMulDivisors.algebraMap_injective K L)] at hb + rw [← isIntegral_algebraMap_iff (FaithfulSMul.algebraMap_injective B L), ← hk, + isIntegral_algebraMap_iff (FaithfulSMul.algebraMap_injective K L)] at hb obtain ⟨a, rfl⟩ := IsIntegrallyClosed.algebraMap_eq_of_integral hb rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply A B L, - (NoZeroSMulDivisors.algebraMap_injective B L).eq_iff] at hk + (FaithfulSMul.algebraMap_injective B L).eq_iff] at hk exact ⟨a, hk⟩ end Galois @@ -163,7 +163,7 @@ end transitivity section surjectivity -open IsScalarTower NoZeroSMulDivisors Polynomial +open FaithfulSMul IsScalarTower Polynomial variable {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] (G : Type*) [Group G] [Finite G] [MulSemiringAction G B] [SMulCommClass G A B] diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index 4ef9b0f8f9357d..b1595411dbd27f 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -66,9 +66,9 @@ and the unit ball inside the `X`-adic completion of `RatFunc K`. ## Implementation details -* Since `LaurentSeries` is just an abbreviation of `HahnSeries ℤ _`, the definition of the +* Since `LaurentSeries` is just an abbreviation of `HahnSeries ℤ R`, the definition of the coefficients is given in terms of `HahnSeries.coeff` and this forces sometimes to go back-and-forth -from `X : _⸨X⸩` to `single 1 1 : HahnSeries ℤ _`. +from `X : R⸨X⸩` to `single 1 1 : HahnSeries ℤ R`. * To prove the isomorphism between the `X`-adic completion of `RatFunc K` and `K⸨X⸩` we construct two completions of `RatFunc K`: the first (`LaurentSeries.ratfuncAdicComplPkg`) is its abstract uniform completion; the second (`LaurentSeries.LaurentSeriesPkg`) is simply `K⸨X⸩`, once we prove @@ -198,11 +198,6 @@ variable [Semiring R] instance : Coe R⟦X⟧ R⸨X⸩ := ⟨HahnSeries.ofPowerSeries ℤ R⟩ -/- Porting note: now a syntactic tautology and not needed elsewhere -theorem coe_powerSeries (x : R⟦X⟧) : - (x : R⸨X⸩) = HahnSeries.ofPowerSeries ℤ R x := - rfl -/ - @[simp] theorem coeff_coe_powerSeries (x : R⟦X⟧) (n : ℕ) : HahnSeries.coeff (x : R⸨X⸩) n = PowerSeries.coeff R n x := by @@ -360,8 +355,6 @@ theorem coe_smul {S : Type*} [Semiring S] [Module R S] (r : R) (x : S⟦X⟧) : ext simp [coeff_coe, coeff_smul, smul_ite] --- Porting note: RingHom.map_bit0 and RingHom.map_bit1 no longer exist - @[norm_cast] theorem coe_pow (n : ℕ) : ((f ^ n : R⟦X⟧) : R⸨X⸩) = (ofPowerSeries ℤ R f) ^ n := (ofPowerSeries ℤ R).map_pow _ _ diff --git a/Mathlib/RingTheory/LocalProperties/IntegrallyClosed.lean b/Mathlib/RingTheory/LocalProperties/IntegrallyClosed.lean index 3753f3c70b5be5..bb0b3fd60f8d2a 100644 --- a/Mathlib/RingTheory/LocalProperties/IntegrallyClosed.lean +++ b/Mathlib/RingTheory/LocalProperties/IntegrallyClosed.lean @@ -54,7 +54,7 @@ theorem IsIntegrallyClosed.of_localization_maximal {R : Type*} [CommRing R] [IsD apply mem_span_singleton'.mpr ⟨b * y.1, _⟩ rw [← hb, ← mul_assoc, mul_comm y.2.1 b, mul_assoc, mul_assoc] exact congrArg (HMul.hMul b) <| (mul_comm y.1 x.2.1).trans <| - NoZeroSMulDivisors.algebraMap_injective R (Localization R⁰) <| mk'_eq_iff_eq.mp <| + FaithfulSMul.algebraMap_injective R (Localization R⁰) <| mk'_eq_iff_eq.mp <| (mk'_eq_algebraMap_mk'_of_submonoid_le _ _ p.primeCompl_le_nonZeroDivisors y.1 y.2).trans <| show algebraMap (Localization.AtPrime p) _ (mk' _ y.1 y.2) = mk' _ x.1 x.2 by simpa only [← mk_eq_mk', ← hy] using by rfl diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 309f147722debf..3fc7f9b8d2bd8a 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -81,6 +81,9 @@ variable (R K) protected theorem injective : Function.Injective (algebraMap R K) := IsLocalization.injective _ (le_of_eq rfl) +instance (priority := 100) : FaithfulSMul R K := + (faithfulSMul_iff_algebraMap_injective R K).mpr <| IsFractionRing.injective R K + variable {R K} @[norm_cast, simp] @@ -88,9 +91,6 @@ variable {R K} theorem coe_inj {a b : R} : (Algebra.cast a : K) = Algebra.cast b ↔ a = b := (IsFractionRing.injective R K).eq_iff -instance (priority := 100) [NoZeroDivisors K] : NoZeroSMulDivisors R K := - NoZeroSMulDivisors.of_algebraMap_injective <| IsFractionRing.injective R K - protected theorem to_map_ne_zero_of_mem_nonZeroDivisors [Nontrivial R] {x : R} (hx : x ∈ nonZeroDivisors R) : algebraMap R K x ≠ 0 := IsLocalization.to_map_ne_zero_of_mem_nonZeroDivisors _ le_rfl hx @@ -468,10 +468,14 @@ theorem algebraMap_injective_of_field_isFractionRing (K L : Type*) [Field K] [Se rw [← RingHom.coe_comp, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R K L] exact (algebraMap K L).injective.comp (IsFractionRing.injective R K) -theorem NoZeroSMulDivisors.of_field_isFractionRing [NoZeroDivisors S] (K L : Type*) [Field K] - [Semiring L] [Nontrivial L] [Algebra R K] [IsFractionRing R K] [Algebra S L] [Algebra K L] - [Algebra R L] [IsScalarTower R S L] [IsScalarTower R K L] : NoZeroSMulDivisors R S := - of_algebraMap_injective (algebraMap_injective_of_field_isFractionRing R S K L) +theorem FaithfulSMul.of_field_isFractionRing (K L : Type*) [Field K] [Semiring L] + [Nontrivial L] [Algebra R K] [IsFractionRing R K] [Algebra S L] [Algebra K L] [Algebra R L] + [IsScalarTower R S L] [IsScalarTower R K L] : FaithfulSMul R S := + (faithfulSMul_iff_algebraMap_injective R S).mpr <| + algebraMap_injective_of_field_isFractionRing R S K L + +@[deprecated (since := "2025-01-31")] +alias NoZeroSMulDivisors.of_field_isFractionRing := FaithfulSMul.of_field_isFractionRing end algebraMap_injective @@ -509,15 +513,15 @@ theorem mk_eq_div {r s} : Should usually be introduced locally along with `isScalarTower_liftAlgebra` See note [reducible non-instances]. -/ noncomputable abbrev liftAlgebra [IsDomain R] [Field K] [Algebra R K] - [NoZeroSMulDivisors R K] : Algebra (FractionRing R) K := - RingHom.toAlgebra (IsFractionRing.lift (NoZeroSMulDivisors.algebraMap_injective R _)) + [FaithfulSMul R K] : Algebra (FractionRing R) K := + RingHom.toAlgebra (IsFractionRing.lift (FaithfulSMul.algebraMap_injective R _)) -- Porting note: had to fill in the `_` by hand for this instance -instance isScalarTower_liftAlgebra [IsDomain R] [Field K] [Algebra R K] [NoZeroSMulDivisors R K] : +instance isScalarTower_liftAlgebra [IsDomain R] [Field K] [Algebra R K] [FaithfulSMul R K] : by letI := liftAlgebra R K; exact IsScalarTower R (FractionRing R) K := by letI := liftAlgebra R K exact IsScalarTower.of_algebraMap_eq fun x => - (IsFractionRing.lift_algebraMap (NoZeroSMulDivisors.algebraMap_injective R K) x).symm + (IsFractionRing.lift_algebraMap (FaithfulSMul.algebraMap_injective R K) x).symm /-- Given a ring `A` and a localization map to a fraction ring `f : A →+* K`, we get an `A`-isomorphism between the fraction ring of `A` as a quotient @@ -526,10 +530,9 @@ noncomputable def algEquiv (K : Type*) [CommRing K] [Algebra A K] [IsFractionRin FractionRing A ≃ₐ[A] K := Localization.algEquiv (nonZeroDivisors A) K -instance [Algebra R A] [NoZeroSMulDivisors R A] : NoZeroSMulDivisors R (FractionRing A) := by - apply NoZeroSMulDivisors.of_algebraMap_injective - rw [IsScalarTower.algebraMap_eq R A] - apply Function.Injective.comp (NoZeroSMulDivisors.algebraMap_injective A (FractionRing A)) - (NoZeroSMulDivisors.algebraMap_injective R A) +instance [Algebra R A] [FaithfulSMul R A] : FaithfulSMul R (FractionRing A) := by + rw [faithfulSMul_iff_algebraMap_injective, IsScalarTower.algebraMap_eq R A] + exact (FaithfulSMul.algebraMap_injective A (FractionRing A)).comp + (FaithfulSMul.algebraMap_injective R A) end FractionRing diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 5238e642ff74e1..17ab0c536c1ad3 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -6,6 +6,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan import Mathlib.GroupTheory.MonoidLocalization.Away import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Localization.Defs +import Mathlib.Algebra.Algebra.Tower /-! # Ideals in localizations of commutative rings @@ -253,6 +254,32 @@ theorem ideal_eq_iInf_comap_map_away {S : Finset R} (hS : Ideal.span (α := R) S rw [pow_add, mul_assoc, ← mul_comm x, e] exact I.mul_mem_left _ y.2 +variable (R) in +lemma _root_.NoZeroSMulDivisors_of_isLocalization (Rₚ Sₚ : Type*) [CommRing Rₚ] [CommRing Sₚ] + [Algebra R Rₚ] [Algebra R Sₚ] [Algebra S Sₚ] [Algebra Rₚ Sₚ] [IsScalarTower R S Sₚ] + [IsScalarTower R Rₚ Sₚ] {M : Submonoid R} (hM : M ≤ R⁰) [IsLocalization M Rₚ] + [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₚ] [NoZeroSMulDivisors R S] [IsDomain S] : + NoZeroSMulDivisors Rₚ Sₚ := by + have e : Algebra.algebraMapSubmonoid S M ≤ S⁰ := + Submonoid.map_le_of_le_comap _ <| hM.trans + (nonZeroDivisors_le_comap_nonZeroDivisors_of_injective _ + (FaithfulSMul.algebraMap_injective _ _)) + have : IsDomain Sₚ := IsLocalization.isDomain_of_le_nonZeroDivisors S e + have : algebraMap Rₚ Sₚ = IsLocalization.map (T := Algebra.algebraMapSubmonoid S M) Sₚ + (algebraMap R S) (Submonoid.le_comap_map M) := by + apply IsLocalization.ringHom_ext M + simp only [IsLocalization.map_comp, ← IsScalarTower.algebraMap_eq] + rw [NoZeroSMulDivisors.iff_algebraMap_injective, RingHom.injective_iff_ker_eq_bot, + RingHom.ker_eq_bot_iff_eq_zero] + intro x hx + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective M x + simp only [RingHom.algebraMap_toAlgebra, IsLocalization.map_mk', IsLocalization.mk'_eq_zero_iff, + Subtype.exists, exists_prop, this] at hx ⊢ + obtain ⟨_, ⟨a, ha, rfl⟩, H⟩ := hx + simp only [← _root_.map_mul, + (injective_iff_map_eq_zero' _).mp (FaithfulSMul.algebraMap_injective R S)] at H + exact ⟨a, ha, H⟩ + end CommRing end IsLocalization diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index 5e9059d15096e8..abf87d3d9d06d8 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -364,14 +364,14 @@ theorem isAlgebraic_iff' [Field K] [IsDomain R] [IsDomain S] [Algebra R K] [Alge refine IsIntegral.mul ?_ ?_ · rw [← isAlgebraic_iff_isIntegral] refine .extendScalars - (NoZeroSMulDivisors.algebraMap_injective R (FractionRing R)) ?_ + (FaithfulSMul.algebraMap_injective R (FractionRing R)) ?_ exact .algebraMap (h a) · rw [← isAlgebraic_iff_isIntegral] use (f.map (algebraMap R (FractionRing R))).reverse constructor · rwa [Ne, Polynomial.reverse_eq_zero, ← Polynomial.degree_eq_bot, Polynomial.degree_map_eq_of_injective - (NoZeroSMulDivisors.algebraMap_injective R (FractionRing R)), + (FaithfulSMul.algebraMap_injective R (FractionRing R)), Polynomial.degree_eq_bot] · have : Invertible (algebraMap S K b) := IsUnit.invertible @@ -379,7 +379,7 @@ theorem isAlgebraic_iff' [Field K] [IsDomain R] [IsDomain S] [Algebra R K] [Alge (mem_nonZeroDivisors_iff_ne_zero.2 fun h => nonZeroDivisors.ne_zero ha ((injective_iff_map_eq_zero (algebraMap S K)).1 - (NoZeroSMulDivisors.algebraMap_injective _ _) b h))) + (FaithfulSMul.algebraMap_injective _ _) b h))) rw [Polynomial.aeval_def, ← invOf_eq_inv, Polynomial.eval₂_reverse_eq_zero_iff, Polynomial.eval₂_map, ← IsScalarTower.algebraMap_eq, ← Polynomial.aeval_def, Polynomial.aeval_algebraMap_apply, hf₂, RingHom.map_zero] @@ -388,7 +388,7 @@ theorem isAlgebraic_iff' [Field K] [IsDomain R] [IsDomain S] [Algebra R K] [Alge use f, hf₁ rw [Polynomial.aeval_algebraMap_apply] at hf₂ exact - (injective_iff_map_eq_zero (algebraMap S K)).1 (NoZeroSMulDivisors.algebraMap_injective _ _) _ + (injective_iff_map_eq_zero (algebraMap S K)).1 (FaithfulSMul.algebraMap_injective _ _) _ hf₂ open nonZeroDivisors diff --git a/Mathlib/RingTheory/Localization/NumDen.lean b/Mathlib/RingTheory/Localization/NumDen.lean index 4b94d345d32e98..44d30a02b6e6e3 100644 --- a/Mathlib/RingTheory/Localization/NumDen.lean +++ b/Mathlib/RingTheory/Localization/NumDen.lean @@ -85,9 +85,9 @@ lemma num_zero : IsFractionRing.num A (0 : K) = 0 := by have := mk'_num_den' A (0 : K) simp only [div_eq_zero_iff] at this rcases this with h | h - · exact NoZeroSMulDivisors.algebraMap_injective A K (by convert h; simp) + · exact FaithfulSMul.algebraMap_injective A K (by convert h; simp) · replace h : algebraMap A K (den A (0 : K)) = algebraMap A K 0 := by convert h; simp - absurd NoZeroSMulDivisors.algebraMap_injective A K h + absurd FaithfulSMul.algebraMap_injective A K h apply nonZeroDivisors.coe_ne_zero @[simp] @@ -116,8 +116,8 @@ theorem isUnit_den_iff (x : K) : IsUnit (den A x : A) ↔ IsLocalization.IsInteg · simp only [mk'_num_den'] intro h replace h : algebraMap A K (den A x : A) = algebraMap A K 0 := by convert h; simp - exact nonZeroDivisors.coe_ne_zero _ <| NoZeroSMulDivisors.algebraMap_injective A K h - exact NoZeroSMulDivisors.algebraMap_injective A K + exact nonZeroDivisors.coe_ne_zero _ <| FaithfulSMul.algebraMap_injective A K h + exact FaithfulSMul.algebraMap_injective A K theorem isUnit_den_zero : IsUnit (den A (0 : K) : A) := by simp [isUnit_den_iff, IsLocalization.isInteger_zero] @@ -126,11 +126,11 @@ lemma associated_den_num_inv (x : K) (hx : x ≠ 0) : Associated (den A x : A) ( associated_of_dvd_dvd (IsRelPrime.dvd_of_dvd_mul_right (IsFractionRing.num_den_reduced A x).symm <| dvd_of_mul_left_dvd (a := (den A x⁻¹ : A)) <| dvd_of_eq <| - NoZeroSMulDivisors.algebraMap_injective A K <| Eq.symm <| eq_of_div_eq_one + FaithfulSMul.algebraMap_injective A K <| Eq.symm <| eq_of_div_eq_one (by simp [mul_div_mul_comm, hx])) (IsRelPrime.dvd_of_dvd_mul_right (IsFractionRing.num_den_reduced A x⁻¹) <| dvd_of_mul_left_dvd (a := (num A x : A)) <| dvd_of_eq <| - NoZeroSMulDivisors.algebraMap_injective A K <| eq_of_div_eq_one + FaithfulSMul.algebraMap_injective A K <| eq_of_div_eq_one (by simp [mul_div_mul_comm, hx])) lemma associated_num_den_inv (x : K) (hx : x ≠ 0) : Associated (num A x : A) (den A x⁻¹) := by diff --git a/Mathlib/RingTheory/Nilpotent/Defs.lean b/Mathlib/RingTheory/Nilpotent/Defs.lean index cbdde6aab35bc3..adebd3a741323b 100644 --- a/Mathlib/RingTheory/Nilpotent/Defs.lean +++ b/Mathlib/RingTheory/Nilpotent/Defs.lean @@ -52,26 +52,26 @@ theorem not_isNilpotent_one [MonoidWithZero R] [Nontrivial R] : lemma IsNilpotent.pow_succ (n : ℕ) {S : Type*} [MonoidWithZero S] {x : S} (hx : IsNilpotent x) : IsNilpotent (x ^ n.succ) := by - obtain ⟨N,hN⟩ := hx + obtain ⟨N, hN⟩ := hx use N rw [← pow_mul, Nat.succ_mul, pow_add, hN, mul_zero] -theorem IsNilpotent.of_pow [MonoidWithZero R] {x : R} {m : ℕ} +theorem IsNilpotent.of_pow [MonoidWithZero R] {x : R} {m : ℕ} (h : IsNilpotent (x ^ m)) : IsNilpotent x := by obtain ⟨n, h⟩ := h - use m*n + use m * n rw [← h, pow_mul x m n] lemma IsNilpotent.pow_of_pos {n} {S : Type*} [MonoidWithZero S] {x : S} (hx : IsNilpotent x) (hn : n ≠ 0) : IsNilpotent (x ^ n) := by cases n with | zero => contradiction - | succ => exact IsNilpotent.pow_succ _ hx + | succ => exact IsNilpotent.pow_succ _ hx @[simp] -lemma IsNilpotent.pow_iff_pos {n} {S : Type*} [MonoidWithZero S] {x : S} - (hn : n ≠ 0) : IsNilpotent (x ^ n) ↔ IsNilpotent x := - ⟨fun h => of_pow h, fun h => pow_of_pos h hn⟩ +lemma IsNilpotent.pow_iff_pos {n} {S : Type*} [MonoidWithZero S] {x : S} (hn : n ≠ 0) : + IsNilpotent (x ^ n) ↔ IsNilpotent x := + ⟨of_pow, (pow_of_pos · hn)⟩ theorem IsNilpotent.map [MonoidWithZero R] [MonoidWithZero S] {r : R} {F : Type*} [FunLike F R S] [MonoidWithZeroHomClass F R S] (hr : IsNilpotent r) (f : F) : diff --git a/Mathlib/RingTheory/Norm/Basic.lean b/Mathlib/RingTheory/Norm/Basic.lean index 77ed251d0b1113..2b04897d0af568 100644 --- a/Mathlib/RingTheory/Norm/Basic.lean +++ b/Mathlib/RingTheory/Norm/Basic.lean @@ -230,7 +230,7 @@ theorem norm_eq_prod_embeddings [FiniteDimensional K L] [Algebra.IsSeparable K L theorem norm_eq_prod_automorphisms [FiniteDimensional K L] [IsGalois K L] (x : L) : algebraMap K L (norm K x) = ∏ σ : L ≃ₐ[K] L, σ x := by - apply NoZeroSMulDivisors.algebraMap_injective L (AlgebraicClosure L) + apply FaithfulSMul.algebraMap_injective L (AlgebraicClosure L) rw [map_prod (algebraMap L (AlgebraicClosure L))] rw [← Fintype.prod_equiv (Normal.algHomEquivAut K (AlgebraicClosure L) L)] · rw [← norm_eq_prod_embeddings _ _ x, ← IsScalarTower.algebraMap_apply] diff --git a/Mathlib/RingTheory/Polynomial/Content.lean b/Mathlib/RingTheory/Polynomial/Content.lean index 394d392597e7be..003050fdd07b20 100644 --- a/Mathlib/RingTheory/Polynomial/Content.lean +++ b/Mathlib/RingTheory/Polynomial/Content.lean @@ -264,7 +264,7 @@ theorem aeval_primPart_eq_zero {S : Type*} [Ring S] [IsDomain S] [Algebra R S] aeval s p.primPart = 0 := by rw [eq_C_content_mul_primPart p, map_mul, aeval_C] at hp have hcont : p.content ≠ 0 := fun h => hpzero (content_eq_zero_iff.1 h) - replace hcont := Function.Injective.ne (NoZeroSMulDivisors.algebraMap_injective R S) hcont + replace hcont := Function.Injective.ne (FaithfulSMul.algebraMap_injective R S) hcont rw [map_zero] at hcont exact eq_zero_of_ne_zero_of_mul_left_eq_zero hcont hp diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean index 198c98c81527be..fd90e039de176f 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean @@ -616,17 +616,17 @@ where `μ` varies over the `n`-th roots of unity. -/ theorem _root_.IsPrimitiveRoot.pow_sub_pow_eq_prod_sub_mul (hpos : 0 < n) (h : IsPrimitiveRoot ζ n) : x ^ n - y ^ n = ∏ ζ ∈ nthRootsFinset n R, (x - ζ * y) := by let K := FractionRing R - apply NoZeroSMulDivisors.algebraMap_injective R K + apply FaithfulSMul.algebraMap_injective R K rw [map_sub, map_pow, map_pow, map_prod] simp_rw [map_sub, map_mul] have h' : IsPrimitiveRoot (algebraMap R K ζ) n := - h.map_of_injective <| NoZeroSMulDivisors.algebraMap_injective R K + h.map_of_injective <| FaithfulSMul.algebraMap_injective R K rw [h'.pow_sub_pow_eq_prod_sub_mul_field _ _ hpos] refine (prod_nbij (algebraMap R K) (fun a ha ↦ map_mem_nthRootsFinset ha _) (fun a _ b _ H ↦ - NoZeroSMulDivisors.algebraMap_injective R K H) (fun a ha ↦ ?_) (fun _ _ ↦ rfl)).symm + FaithfulSMul.algebraMap_injective R K H) (fun a ha ↦ ?_) (fun _ _ ↦ rfl)).symm have := Set.surj_on_of_inj_on_of_ncard_le (s := nthRootsFinset n R) (t := nthRootsFinset n K) _ (fun _ hr ↦ map_mem_nthRootsFinset hr _) - (fun a _ b _ H ↦ NoZeroSMulDivisors.algebraMap_injective R K H) + (fun a _ b _ H ↦ FaithfulSMul.algebraMap_injective R K H) (by simp [h.card_nthRootsFinset, h'.card_nthRootsFinset]) obtain ⟨x, hx, hx1⟩ := this _ ha exact ⟨x, hx, hx1.symm⟩ diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Roots.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Roots.lean index 74c4705dddc413..e82cfc397703d1 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Roots.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Roots.lean @@ -163,7 +163,7 @@ open IsPrimitiveRoot Complex theorem _root_.IsPrimitiveRoot.minpoly_eq_cyclotomic_of_irreducible {K : Type*} [Field K] {R : Type*} [CommRing R] [IsDomain R] {μ : R} {n : ℕ} [Algebra K R] (hμ : IsPrimitiveRoot μ n) (h : Irreducible <| cyclotomic n K) [NeZero (n : K)] : cyclotomic n K = minpoly K μ := by - haveI := NeZero.of_noZeroSMulDivisors K R n + haveI := NeZero.of_faithfulSMul K R n refine minpoly.eq_of_irreducible_of_monic h ?_ (cyclotomic.monic n K) rwa [aeval_def, eval₂_eq_eval_map, map_cyclotomic, ← IsRoot.def, isRoot_cyclotomic_iff] diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean index a1e63fad9a3288..e5bb59e900c6a2 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean @@ -173,7 +173,7 @@ theorem dvd_pow_natDegree_of_eval₂_eq_zero {f : R →+* A} (hf : Function.Inje theorem dvd_pow_natDegree_of_aeval_eq_zero [Algebra R A] [Nontrivial A] [NoZeroSMulDivisors R A] {p : R[X]} (hp : p.Monic) (x y : R) (z : A) (h : Polynomial.aeval z p = 0) (hz : z * algebraMap R A x = algebraMap R A y) : x ∣ y ^ p.natDegree := - dvd_pow_natDegree_of_eval₂_eq_zero (NoZeroSMulDivisors.algebraMap_injective R A) hp x y z h + dvd_pow_natDegree_of_eval₂_eq_zero (FaithfulSMul.algebraMap_injective R A) hp x y z h ((mul_comm _ _).trans hz) end ScaleRoots diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index ca5b849f7101c3..b72c8d4e3eee82 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -238,7 +238,7 @@ theorem mem_adjoin_of_smul_prime_smul_of_minpoly_isEisensteinAt {B : PowerBasis have := B.finite set P := minpoly R B.gen with hP obtain ⟨n, hn⟩ := Nat.exists_eq_succ_of_ne_zero B.dim_pos.ne' - haveI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans R K L + haveI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans_faithfulSMul R K L let _ := P.map (algebraMap R L) -- There is a polynomial `Q` such that `p • z = aeval B.gen Q`. We can assume that -- `Q.degree < P.degree` and `Q ≠ 0`. diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index 97eaca3724e772..c1ccd93811bc6b 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -315,7 +315,7 @@ variable [CommRing R] [IsDomain R] theorem order_mul (φ ψ : R⟦X⟧) : order (φ * ψ) = order φ + order ψ := by classical simp only [order_eq_emultiplicity_X] - rw [emultiplicity_mul X_prime] + exact emultiplicity_mul X_prime -- Dividing `X` by the maximal power of `X` dividing it leaves `1`. @[simp] diff --git a/Mathlib/RingTheory/RingHom/FinitePresentation.lean b/Mathlib/RingTheory/RingHom/FinitePresentation.lean index f2503b3b0b8f95..4aa353d93690de 100644 --- a/Mathlib/RingTheory/RingHom/FinitePresentation.lean +++ b/Mathlib/RingTheory/RingHom/FinitePresentation.lean @@ -58,7 +58,7 @@ theorem finitePresentation_holdsForLocalizationAway : introv R _ suffices Algebra.FinitePresentation R S by rw [RingHom.FinitePresentation] - convert this; ext; + convert this; ext rw [Algebra.smul_def]; rfl exact IsLocalization.Away.finitePresentation r @@ -96,7 +96,7 @@ theorem finitePresentation_ofLocalizationSpanTarget : classical letI := f.toAlgebra replace H : ∀ r : s, Algebra.FinitePresentation R (Localization.Away (r : S)) := by - intro r; simp_rw [RingHom.FinitePresentation] at H; + intro r; simp_rw [RingHom.FinitePresentation] at H convert H r; ext; simp_rw [Algebra.smul_def]; rfl /- We already know that `S` is of finite type over `R`, so we have a surjection diff --git a/Mathlib/RingTheory/Spectrum/Prime/Module.lean b/Mathlib/RingTheory/Spectrum/Prime/Module.lean index e7f608704c4c89..28fe87edbab06e 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Module.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Module.lean @@ -31,8 +31,7 @@ lemma LocalizedModule.subsingleton_iff_disjoint {f : R} : Subsingleton (LocalizedModule (.powers f) M) ↔ Disjoint ↑(PrimeSpectrum.basicOpen f) (Module.support R M) := by rw [subsingleton_iff_support_subset, PrimeSpectrum.basicOpen_eq_zeroLocus_compl, - disjoint_compl_left_iff] - rfl + disjoint_compl_left_iff, Set.le_iff_subset] lemma Module.stableUnderSpecialization_support : StableUnderSpecialization (Module.support R M) := by diff --git a/Mathlib/RingTheory/Support.lean b/Mathlib/RingTheory/Support.lean index cbb8bb690f8a47..ec284eeb4dd94a 100644 --- a/Mathlib/RingTheory/Support.lean +++ b/Mathlib/RingTheory/Support.lean @@ -50,8 +50,8 @@ lemma Module.not_mem_support_iff : lemma Module.not_mem_support_iff' : p ∉ Module.support R M ↔ ∀ m : M, ∃ r ∉ p.asIdeal, r • m = 0 := by - rw [not_mem_support_iff, LocalizedModule.subsingleton_iff] - rfl + simp only [not_mem_support_iff, Ideal.primeCompl, LocalizedModule.subsingleton_iff, + Submonoid.mem_mk, Subsemigroup.mem_mk, Set.mem_compl_iff, SetLike.mem_coe] lemma Module.mem_support_iff' : p ∈ Module.support R M ↔ ∃ m : M, ∀ r ∉ p.asIdeal, r • m ≠ 0 := by diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index 5b322e66a5a7b8..849dcdfa2234c2 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -263,7 +263,7 @@ theorem trace_eq_sum_embeddings [FiniteDimensional K L] [Algebra.IsSeparable K L theorem trace_eq_sum_automorphisms (x : L) [FiniteDimensional K L] [IsGalois K L] : algebraMap K L (Algebra.trace K L x) = ∑ σ : L ≃ₐ[K] L, σ x := by - apply NoZeroSMulDivisors.algebraMap_injective L (AlgebraicClosure L) + apply FaithfulSMul.algebraMap_injective L (AlgebraicClosure L) rw [_root_.map_sum (algebraMap L (AlgebraicClosure L))] rw [← Fintype.sum_equiv (Normal.algHomEquivAut K (AlgebraicClosure L) L)] · rw [← trace_eq_sum_embeddings (AlgebraicClosure L) (x := x)] diff --git a/Mathlib/RingTheory/Trace/Quotient.lean b/Mathlib/RingTheory/Trace/Quotient.lean index 2ab1cd6fe80de3..54bb340675a08f 100644 --- a/Mathlib/RingTheory/Trace/Quotient.lean +++ b/Mathlib/RingTheory/Trace/Quotient.lean @@ -200,7 +200,7 @@ lemma Algebra.trace_quotient_eq_of_isDedekindDomain (x) [IsDedekindDomain R] [Is have e : Algebra.algebraMapSubmonoid S p.primeCompl ≤ S⁰ := Submonoid.map_le_of_le_comap _ <| p.primeCompl_le_nonZeroDivisors.trans (nonZeroDivisors_le_comap_nonZeroDivisors_of_injective _ - (NoZeroSMulDivisors.algebraMap_injective _ _)) + (FaithfulSMul.algebraMap_injective _ _)) haveI : IsDomain Sₚ := IsLocalization.isDomain_of_le_nonZeroDivisors S e haveI : NoZeroSMulDivisors Rₚ Sₚ := by rw [NoZeroSMulDivisors.iff_algebraMap_injective, RingHom.injective_iff_ker_eq_bot, @@ -210,7 +210,7 @@ lemma Algebra.trace_quotient_eq_of_isDedekindDomain (x) [IsDedekindDomain R] [Is simp only [Sₚ, RingHom.algebraMap_toAlgebra, IsLocalization.map_mk', IsLocalization.mk'_eq_zero_iff, mul_eq_zero, Subtype.exists, exists_prop] at hx ⊢ obtain ⟨_, ⟨a, ha, rfl⟩, H⟩ := hx - simp only [(injective_iff_map_eq_zero' _).mp (NoZeroSMulDivisors.algebraMap_injective R S)] at H + simp only [(injective_iff_map_eq_zero' _).mp (FaithfulSMul.algebraMap_injective R S)] at H refine ⟨a, ha, H⟩ haveI : Module.Finite Rₚ Sₚ := Module.Finite_of_isLocalization R S _ _ p.primeCompl haveI : IsIntegrallyClosed Sₚ := isIntegrallyClosed_of_isLocalization _ _ e diff --git a/Mathlib/RingTheory/Unramified/Field.lean b/Mathlib/RingTheory/Unramified/Field.lean index d2ba91b1bb3d95..35b6536cfcb589 100644 --- a/Mathlib/RingTheory/Unramified/Field.lean +++ b/Mathlib/RingTheory/Unramified/Field.lean @@ -3,7 +3,7 @@ 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.FieldTheory.PurelyInseparable +import Mathlib.FieldTheory.PurelyInseparable.Basic import Mathlib.RingTheory.Artinian.Ring import Mathlib.RingTheory.LocalProperties.Basic import Mathlib.Algebra.Polynomial.Taylor diff --git a/Mathlib/RingTheory/Valuation/AlgebraInstances.lean b/Mathlib/RingTheory/Valuation/AlgebraInstances.lean index 91ff35d526b162..58b8aafadc5793 100644 --- a/Mathlib/RingTheory/Valuation/AlgebraInstances.lean +++ b/Mathlib/RingTheory/Valuation/AlgebraInstances.lean @@ -35,7 +35,7 @@ namespace ValuationSubring instance : Algebra v.valuationSubring L := Algebra.ofSubring v.valuationSubring.toSubring theorem algebraMap_injective : Injective (algebraMap v.valuationSubring L) := - (NoZeroSMulDivisors.algebraMap_injective K L).comp (IsFractionRing.injective _ _) + (FaithfulSMul.algebraMap_injective K L).comp (IsFractionRing.injective _ _) theorem isIntegral_of_mem_ringOfIntegers {x : L} (hx : x ∈ integralClosure v.valuationSubring L) : IsIntegral v.valuationSubring (⟨x, hx⟩ : integralClosure v.valuationSubring L) := by diff --git a/Mathlib/RingTheory/Valuation/LocalSubring.lean b/Mathlib/RingTheory/Valuation/LocalSubring.lean index ac0c4253b0f3e7..ee34567001f699 100644 --- a/Mathlib/RingTheory/Valuation/LocalSubring.lean +++ b/Mathlib/RingTheory/Valuation/LocalSubring.lean @@ -54,8 +54,9 @@ lemma LocalSubring.mem_of_isMax_of_isIntegral {R : LocalSubring K} (hR : IsMax R) {x : K} (hx : IsIntegral R.toSubring x) : x ∈ R.toSubring := by let S := Algebra.adjoin R.toSubring {x} have : Algebra.IsIntegral R.toSubring S := Algebra.IsIntegral.adjoin (by simpa) + have : FaithfulSMul R.toSubring S := NoZeroSMulDivisors.instFaithfulSMulOfNontrivial obtain ⟨Q : Ideal S.toSubring, hQ, e⟩ := Ideal.exists_ideal_over_maximal_of_isIntegral - (R := R.toSubring) (S := S) (maximalIdeal _) (le_maximalIdeal (by simp [Ideal.eq_top_iff_one])) + (R := R.toSubring) (S := S) (maximalIdeal _) (le_maximalIdeal (by simp)) have : R = .ofPrime S.toSubring Q := by have hRS : R.toSubring ≤ S.toSubring := fun r hr ↦ algebraMap_mem S ⟨r, hr⟩ apply hR.eq_of_le ⟨hRS.trans (LocalSubring.le_ofPrime _ _), ⟨?_⟩⟩ diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index fda897921c1fe2..988829d93d43a0 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -29,6 +29,7 @@ import Mathlib.Tactic.CategoryTheory.Bicategory.Datatypes import Mathlib.Tactic.CategoryTheory.Bicategory.Normalize import Mathlib.Tactic.CategoryTheory.Bicategory.PureCoherence import Mathlib.Tactic.CategoryTheory.BicategoryCoherence +import Mathlib.Tactic.CategoryTheory.CheckCompositions import Mathlib.Tactic.CategoryTheory.Coherence import Mathlib.Tactic.CategoryTheory.Coherence.Basic import Mathlib.Tactic.CategoryTheory.Coherence.Datatypes diff --git a/Mathlib/Tactic/CategoryTheory/CheckCompositions.lean b/Mathlib/Tactic/CategoryTheory/CheckCompositions.lean new file mode 100644 index 00000000000000..8f346c55bbdbc7 --- /dev/null +++ b/Mathlib/Tactic/CategoryTheory/CheckCompositions.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2025 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Mathlib.CategoryTheory.Category.Basic + +/-! +The `check_compositions` tactic, +which checks the typing of categorical compositions in the goal, +reporting discrepancies at "instances and reducible" transparency. + +Reports from this tactic do not necessarily indicate a problem, +although typically `simp` should reduce rather than increase the reported discrepancies. + +`check_compositions` may be useful in diagnosing uses of `erw` in the category theory library. +-/ + +namespace Mathlib.Tactic.CheckCompositions + +open CategoryTheory + +open Lean Meta Elab Tactic + +/-- Find appearances of `CategoryStruct.comp C inst X Y Z f g`, and apply `f` to each. -/ +def forEachComposition (e : Expr) (f : Expr → MetaM Unit) : MetaM Unit := do + e.forEach (fun e ↦ if e.isAppOfArity ``CategoryStruct.comp 7 then f e else pure ()) + +/-- Given a composition `CategoryStruct.comp _ _ X Y Z f g`, +infer the types of `f` and `g` and check whether their sources and targets agree, +at "instances and reducible" transparency, with `X`, `Y`, and `Z`, +reporting any discrepancies. -/ +def checkComposition (e : Expr) : MetaM Unit := do + match_expr e with + | CategoryStruct.comp _ _ X Y Z f g => + match_expr ← inferType f with + | Quiver.Hom _ _ X' Y' => + withReducibleAndInstances do + if !(← isDefEq X' X) then + logInfo m!"In composition\n {e}\nthe source of\n {f}\nis\n {X'}\nbut should be\n {X}" + if !(← isDefEq Y' Y) then + logInfo m!"In composition\n {e}\nthe target of\n {f}\nis\n {Y'}\nbut should be\n {Y}" + | _ => throwError "In composition\n {e}\nthe type of\n {f}\nis not a morphism." + match_expr ← inferType g with + | Quiver.Hom _ _ Y' Z' => + withReducibleAndInstances do + if !(← isDefEq Y' Y) then + logInfo m!"In composition\n {e}\nthe source of\n {g}\nis\n {Y'}\nbut should be\n {Y}" + if !(← isDefEq Z' Z) then + logInfo m!"In composition\n {e}\nthe target of\n {g}\nis\n {Z'}\nbut should be\n {Z}" + | _ => throwError "In composition\n {e}\nthe type of\n {g}\nis not a morphism." + | _ => throwError "{e} is not a composition." + +/-- Check the typing of categorical compositions in an expression.-/ +def checkCompositions (e : Expr) : MetaM Unit := do + forEachComposition e checkComposition + +/-- Check the typing of categorical compositions in the goal.-/ +def checkCompositionsTac : TacticM Unit := withMainContext do + let e ← getMainTarget + checkCompositions e + +/-- For each composition `f ≫ g` in the goal, +which internally is represented as `CategoryStruct.comp C inst X Y Z f g`, +infer the types of `f` and `g` and check whether their sources and targets agree +with `X`, `Y`, and `Z` at "instances and reducible" transparency, +reporting any discrepancies. + +An example: + +``` +example (j : J) : + colimit.ι ((F ⋙ G) ⋙ H) j ≫ (preservesColimitIso (G ⋙ H) F).inv = + H.map (G.map (colimit.ι F j)) := by + + -- We know which lemma we want to use, and it's even a simp lemma, but `rw` + -- won't let us apply it + fail_if_success rw [ι_preservesColimitIso_inv] + fail_if_success rw [ι_preservesColimitIso_inv (G ⋙ H)] + fail_if_success simp only [ι_preservesColimitIso_inv] + + -- This would work: + -- erw [ι_preservesColimitIso_inv (G ⋙ H)] + + -- `check_compositions` checks if the two morphisms we're composing are + -- composed by abusing defeq, and indeed it tells us that we are abusing + -- definitional associativity of composition of functors here: it prints + -- the following. + + -- info: In composition + -- colimit.ι ((F ⋙ G) ⋙ H) j ≫ (preservesColimitIso (G ⋙ H) F).inv + -- the source of + -- (preservesColimitIso (G ⋙ H) F).inv + -- is + -- colimit (F ⋙ G ⋙ H) + -- but should be + -- colimit ((F ⋙ G) ⋙ H) + + check_compositions + + -- In this case, we can "fix" this by reassociating in the goal, but + -- usually at this point the right thing to do is to back off and + -- check how we ended up with a bad goal in the first place. + dsimp only [Functor.assoc] + + -- This would work now, but it is not needed, because simp works as well + -- rw [ι_preservesColimitIso_inv] + + simp +``` +-/ +elab "check_compositions" : tactic => checkCompositionsTac + +end Mathlib.Tactic.CheckCompositions diff --git a/Mathlib/Tactic/Linter/OldObtain.lean b/Mathlib/Tactic/Linter/OldObtain.lean index f9d77e63e79e8e..14e631d59f79a2 100644 --- a/Mathlib/Tactic/Linter/OldObtain.lean +++ b/Mathlib/Tactic/Linter/OldObtain.lean @@ -10,9 +10,9 @@ import Lean.Elab.Command import Mathlib.Tactic.Linter.Header /-! -# The `oldObtain` linter, against stream-of-conciousness `obtain` +# The `oldObtain` linter, against stream-of-consciousness `obtain` -The `oldObtain` linter flags any occurrences of "stream-of-conciousness" `obtain`, +The `oldObtain` linter flags any occurrences of "stream-of-consciousness" `obtain`, i.e. uses of the `obtain` tactic which do not immediately provide a proof. ## Example @@ -64,7 +64,7 @@ def isObtainWithoutProof : Syntax → Bool @[deprecated isObtainWithoutProof (since := "2024-12-07")] def is_obtain_without_proof := @isObtainWithoutProof -/-- The `oldObtain` linter emits a warning upon uses of the "stream-of-conciousness" variants +/-- The `oldObtain` linter emits a warning upon uses of the "stream-of-consciousness" variants of the `obtain` tactic, i.e. with the proof postponed. -/ register_option linter.oldObtain : Bool := { defValue := false @@ -78,7 +78,7 @@ def oldObtainLinter : Linter where run := withSetOptionIn fun stx => do if (← MonadState.get).messages.hasErrors then return if let some head := stx.find? isObtainWithoutProof then - Linter.logLint linter.oldObtain head m!"Please remove stream-of-conciousness `obtain` syntax" + Linter.logLint linter.oldObtain head m!"Please remove stream-of-consciousness `obtain` syntax" initialize addLinter oldObtainLinter diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean index d71d9584de4711..534a4dbf5c1c6a 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean @@ -25,12 +25,6 @@ disconnected. * `ofClosedSubgroup` : A closed subgroup of a profinite group is profinite. -# TODO - -As discussion in `https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Refactor.20 -Category.20of.20ProfiniteGrp.20and.20ContinuousMulEquiv/near/493290115` - -* Prove `(forget ProfiniteGrp.{u}).ReflectsIsomorphisms` using `profiniteGrpToProfinite` -/ universe u v @@ -80,13 +74,13 @@ is a closed set, thus implying Hausdorff in a topological group.)-/ compact and totally disconnected topological additive group. (The condition of being Hausdorff can be omitted here because totally disconnected implies that {0} is a closed set, thus implying Hausdorff in a topological additive group.)"] -def ProfiniteGrp.of (G : Type u) [Group G] [TopologicalSpace G] [TopologicalGroup G] +abbrev ProfiniteGrp.of (G : Type u) [Group G] [TopologicalSpace G] [TopologicalGroup G] [CompactSpace G] [TotallyDisconnectedSpace G] : ProfiniteGrp.{u} where toProfinite := .of G group := ‹_› topologicalGroup := ‹_› -@[to_additive (attr := simp)] +@[to_additive] lemma ProfiniteGrp.coe_of (G : Type u) [Group G] [TopologicalSpace G] [TopologicalGroup G] [CompactSpace G] [TotallyDisconnectedSpace G] : (ProfiniteGrp.of G : Type u) = G := rfl @@ -96,14 +90,14 @@ lemma ProfiniteGrp.coe_of (G : Type u) [Group G] [TopologicalSpace G] [Topologic structure ProfiniteAddGrp.Hom (A B : ProfiniteAddGrp.{u}) where private mk :: /-- The underlying `ContinuousAddMonoidHom`. -/ - hom : ContinuousAddMonoidHom A B + hom' : ContinuousAddMonoidHom A B /-- The type of morphisms in `ProfiniteGrp`. -/ @[to_additive (attr := ext) existing] structure ProfiniteGrp.Hom (A B : ProfiniteGrp.{u}) where private mk :: /-- The underlying `ContinuousMonoidHom`. -/ - hom : ContinuousMonoidHom A B + hom' : ContinuousMonoidHom A B attribute [to_additive existing ProfiniteAddGrp.Hom.mk] ProfiniteGrp.Hom.mk @@ -111,14 +105,26 @@ attribute [to_additive existing ProfiniteAddGrp.Hom.mk] ProfiniteGrp.Hom.mk instance : Category ProfiniteGrp where Hom A B := ProfiniteGrp.Hom A B id A := ⟨ContinuousMonoidHom.id A⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ + +@[to_additive] +instance : ConcreteCategory ProfiniteGrp (fun X Y => ContinuousMonoidHom X Y) where + hom f := f.hom' + ofHom f := ⟨f⟩ + + /-- The underlying `ContinuousMonoidHom`. -/ +@[to_additive "The underlying `ContinuousAddMonoidHom`."] +abbrev ProfiniteGrp.Hom.hom {M N : ProfiniteGrp.{u}} (f : ProfiniteGrp.Hom M N) : + ContinuousMonoidHom M N := + ConcreteCategory.hom (C := ProfiniteGrp) f /-- Typecheck a `ContinuousMonoidHom` as a morphism in `ProfiniteGrp`. -/ @[to_additive "Typecheck a `ContinuousAddMonoidHom` as a morphism in `ProfiniteAddGrp`."] -abbrev ProfiniteGrp.ofHom {X Y: Type u} [Group X] [TopologicalSpace X] [TopologicalGroup X] +abbrev ProfiniteGrp.ofHom {X Y : Type u} [Group X] [TopologicalSpace X] [TopologicalGroup X] [CompactSpace X] [TotallyDisconnectedSpace X] [Group Y] [TopologicalSpace Y] [TopologicalGroup Y] [CompactSpace Y] [TotallyDisconnectedSpace Y] - (f : ContinuousMonoidHom X Y) : ProfiniteGrp.of X ⟶ ProfiniteGrp.of Y := ⟨f⟩ + (f : ContinuousMonoidHom X Y) : ProfiniteGrp.of X ⟶ ProfiniteGrp.of Y := + ConcreteCategory.ofHom f namespace ProfiniteGrp @@ -153,7 +159,7 @@ variable {X Y Z : Type u} [Group X] [TopologicalSpace X] [TopologicalGroup X] [TopologicalGroup Y] [CompactSpace Y] [TotallyDisconnectedSpace Y] [Group Z] [TopologicalSpace Z] [TopologicalGroup Z] [CompactSpace Z] [TotallyDisconnectedSpace Z] -@[to_additive] +@[to_additive (attr := simp)] lemma hom_ofHom (f : ContinuousMonoidHom X Y) : (ofHom f).hom = f := rfl @[to_additive (attr := simp)] @@ -181,20 +187,13 @@ lemma hom_inv_apply {A B : ProfiniteGrp.{u}} (e : A ≅ B) (x : B) : e.hom (e.in rw [← comp_apply] simp -@[to_additive] -instance : HasForget ProfiniteGrp where - forget := - { obj := fun G => G - map := fun f => f } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - @[to_additive (attr := simp)] -theorem coe_id (X : ProfiniteGrp) : (𝟙 ((forget ProfiniteGrp).obj X)) = id := +theorem coe_id (X : ProfiniteGrp) : (𝟙 X : X → X) = id := rfl @[to_additive (attr := simp)] theorem coe_comp {X Y Z : ProfiniteGrp} (f : X ⟶ Y) (g : Y ⟶ Z) : - ((forget ProfiniteGrp).map f ≫ (forget ProfiniteGrp).map g) = g ∘ f := + (f ≫ g : X → Z) = g ∘ f := rfl /-- Construct a term of `ProfiniteGrp` from a type endowed with the structure of a @@ -249,17 +248,34 @@ def ofContinuousMulEquiv {G : ProfiniteGrp.{u}} {H : Type v} [TopologicalSpace H let _ : TotallyDisconnectedSpace H := Homeomorph.totallyDisconnectedSpace e.toHomeomorph .of H +/-- Build an isomorphism in the category `ProfiniteGrp` from +a `ContinuousMulEquiv` between `ProfiniteGrp`s. -/ +def ContinuousMulEquiv.toProfiniteGrpIso {X Y : ProfiniteGrp} (e : X ≃ₜ* Y) : X ≅ Y where + hom := ofHom e + inv := ofHom e.symm + /-- The functor mapping a profinite group to its underlying profinite space. -/ @[to_additive] instance : HasForget₂ ProfiniteGrp Profinite where forget₂ := { obj G := G.toProfinite - map f := ⟨f, by continuity⟩} + map f := CompHausLike.ofHom _ ⟨f, by continuity⟩} @[to_additive] instance : (forget₂ ProfiniteGrp Profinite).Faithful := { map_injective := fun {_ _} _ _ h => - ConcreteCategory.hom_ext_iff.mpr (congrFun (congrArg ContinuousMap.toFun h)) } + ConcreteCategory.hom_ext _ _ (CategoryTheory.congr_fun h) } + +instance : (forget₂ ProfiniteGrp Profinite).ReflectsIsomorphisms where + reflects {X Y} f _ := by + let i := asIso ((forget₂ ProfiniteGrp Profinite).map f) + let e : X ≃ₜ* Y := + { CompHausLike.homeoOfIso i with + map_mul' := map_mul f.hom } + exact (ContinuousMulEquiv.toProfiniteGrpIso e).isIso_hom + +instance : (forget ProfiniteGrp.{u}).ReflectsIsomorphisms := + CategoryTheory.reflectsIsomorphisms_comp (forget₂ ProfiniteGrp Profinite) (forget Profinite) end ProfiniteGrp @@ -318,13 +334,13 @@ abbrev limitCone : Limits.Cone F where /-- `ProfiniteGrp.limitCone` is a limit cone. -/ def limitConeIsLimit : Limits.IsLimit (limitCone F) where - lift cone := ⟨{ - ((Profinite.limitConeIsLimit (F ⋙ (forget₂ ProfiniteGrp Profinite))).lift - ((forget₂ ProfiniteGrp Profinite).mapCone cone)) with - map_one' := Subtype.ext (funext fun j ↦ map_one (cone.π.app j).hom) - -- TODO: investigate whether it's possible to set up `ext` lemmas for the `TopCat`-related - -- categories so that `by ext j; exact map_one (cone.π.app j)` works here, similarly below. - map_mul' := fun _ _ ↦ Subtype.ext (funext fun j ↦ map_mul (cone.π.app j).hom _ _) }⟩ + lift cone := ofHom + { ((Profinite.limitConeIsLimit (F ⋙ (forget₂ ProfiniteGrp Profinite))).lift + ((forget₂ ProfiniteGrp Profinite).mapCone cone)).hom with + map_one' := Subtype.ext (funext fun j ↦ map_one (cone.π.app j).hom) + -- TODO: investigate whether it's possible to set up `ext` lemmas for the `TopCat`-related + -- categories so that `by ext j; exact map_one (cone.π.app j)` works here, similarly below. + map_mul' := fun _ _ ↦ Subtype.ext (funext fun j ↦ map_mul (cone.π.app j).hom _ _) } uniq cone m h := by apply (forget₂ ProfiniteGrp Profinite).map_injective simpa using (Profinite.limitConeIsLimit (F ⋙ (forget₂ ProfiniteGrp Profinite))).uniq diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Limits.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Limits.lean index d6277f4439ca63..f9eccd2caf7327 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Limits.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Limits.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Nailin Guan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nailin Guan, Youle Fang, Jujian Zhang, Yuyang Zhao -/ +import Mathlib.CategoryTheory.ConcreteCategory.EpiMono import Mathlib.Topology.Algebra.Category.ProfiniteGrp.Basic import Mathlib.Topology.Algebra.ClopenNhdofOne @@ -41,14 +42,14 @@ instance (P : ProfiniteGrp) : SmallCategory (OpenNormalSubgroup P) := /-- The functor from `OpenNormalSubgroup P` to `FiniteGrp` sending `U` to `P ⧸ U`, where `P : ProfiniteGrp`. -/ -def toFiniteQuotientFunctor (P : ProfiniteGrp) : OpenNormalSubgroup P ⥤ FiniteGrp := { - obj := fun H => FiniteGrp.of (P ⧸ H.toSubgroup) - map := fun fHK => FiniteGrp.ofHom (QuotientGroup.map _ _ (.id _) (leOfHom fHK)) - map_id _ := ConcreteCategory.ext <| QuotientGroup.map_id _ - map_comp f g := ConcreteCategory.ext <| (QuotientGroup.map_comp_map - _ _ _ (.id _) (.id _) (leOfHom f) (leOfHom g)).symm } - -/--The `MonoidHom` from a profinite group `P` to the projective limit of its quotients by +def toFiniteQuotientFunctor (P : ProfiniteGrp) : OpenNormalSubgroup P ⥤ FiniteGrp where + obj := fun H => FiniteGrp.of (P ⧸ H.toSubgroup) + map := fun fHK => FiniteGrp.ofHom (QuotientGroup.map _ _ (.id _) (leOfHom fHK)) + map_id _ := ConcreteCategory.ext <| QuotientGroup.map_id _ + map_comp f g := ConcreteCategory.ext <| (QuotientGroup.map_comp_map + _ _ _ (.id _) (.id _) (leOfHom f) (leOfHom g)).symm + +/-- The `MonoidHom` from a profinite group `P` to the projective limit of its quotients by open normal subgroups ordered by inclusion.-/ def toLimit_fun (P : ProfiniteGrp.{u}) : P →* limit (toFiniteQuotientFunctor P ⋙ forget₂ FiniteGrp ProfiniteGrp) where @@ -59,10 +60,9 @@ def toLimit_fun (P : ProfiniteGrp.{u}) : P →* lemma toLimit_fun_continuous (P : ProfiniteGrp.{u}) : Continuous (toLimit_fun P) := by apply continuous_induced_rng.mpr (continuous_pi _) intro H - dsimp only [Functor.comp_obj, CompHausLike.toCompHausLike_obj, CompHausLike.compHausLikeToTop_obj, - CompHausLike.coe_of, Functor.comp_map, CompHausLike.toCompHausLike_map, - CompHausLike.compHausLikeToTop_map, Set.mem_setOf_eq, toLimit_fun, - MonoidHom.coe_mk, OneHom.coe_mk, Function.comp_apply] + dsimp only [Functor.comp_obj, CompHausLike.coe_of, Functor.comp_map, + CompHausLike.toCompHausLike_map, CompHausLike.compHausLikeToTop_map, Set.mem_setOf_eq, + toLimit_fun, MonoidHom.coe_mk, OneHom.coe_mk, Function.comp_apply] apply Continuous.mk intro s _ rw [← (Set.biUnion_preimage_singleton QuotientGroup.mk s)] @@ -123,27 +123,19 @@ theorem toLimit_injective (P : ProfiniteGrp.{u}) : Function.Injective (toLimit P its quotients by open normal subgroups -/ noncomputable def continuousMulEquivLimittoFiniteQuotientFunctor (P : ProfiniteGrp.{u}) : P ≃ₜ* (limit (toFiniteQuotientFunctor P ⋙ forget₂ FiniteGrp ProfiniteGrp)) := { - (Continuous.homeoOfEquivCompactToT2 (f := Equiv.ofBijective _ - ⟨toLimit_injective P, toLimit_surjective P⟩) - P.toLimit.hom.continuous_toFun) - with + (Continuous.homeoOfEquivCompactToT2 + (f := Equiv.ofBijective _ ⟨toLimit_injective P, toLimit_surjective P⟩) + P.toLimit.hom.continuous_toFun) with map_mul' := (toLimit P).hom.map_mul' } ---TODO : Refactor using `(forget ProfiniteGrp.{u}).ReflectsIsomorphisms` after it is proved. +instance isIso_toLimit (P : ProfiniteGrp.{u}) : IsIso (toLimit P) := by + rw [CategoryTheory.ConcreteCategory.isIso_iff_bijective] + exact ⟨toLimit_injective P, toLimit_surjective P⟩ + /-- The isomorphism in the category of profinite group between a profinite group and the projective limit of its quotients by open normal subgroups -/ noncomputable def isoLimittoFiniteQuotientFunctor (P : ProfiniteGrp.{u}) : - P ≅ (limit (toFiniteQuotientFunctor P ⋙ forget₂ FiniteGrp ProfiniteGrp)) where - hom := P.toLimit - inv := ofHom { (continuousMulEquivLimittoFiniteQuotientFunctor P).symm.toMonoidHom with - continuous_toFun := (continuousMulEquivLimittoFiniteQuotientFunctor P).continuous_invFun} - hom_inv_id := by - ext x - exact ContinuousMulEquiv.symm_apply_apply - (continuousMulEquivLimittoFiniteQuotientFunctor P) x - inv_hom_id := by - ext x - exact ContinuousMulEquiv.apply_symm_apply - (continuousMulEquivLimittoFiniteQuotientFunctor P) x + P ≅ (limit (toFiniteQuotientFunctor P ⋙ forget₂ FiniteGrp ProfiniteGrp)) := + ContinuousMulEquiv.toProfiniteGrpIso (continuousMulEquivLimittoFiniteQuotientFunctor P) end ProfiniteGrp diff --git a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean index 83fcac52c04191..31ac4d376124eb 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean @@ -35,13 +35,11 @@ section Nat section Monoid -namespace HasProd - /-- If `f : ℕ → M` has product `m`, then the partial products `∏ i ∈ range n, f i` converge to `m`. -/ @[to_additive "If `f : ℕ → M` has sum `m`, then the partial sums `∑ i ∈ range n, f i` converge to `m`."] -theorem tendsto_prod_nat {f : ℕ → M} (h : HasProd f m) : +theorem HasProd.tendsto_prod_nat {f : ℕ → M} (h : HasProd f m) : Tendsto (fun n ↦ ∏ i ∈ range n, f i) atTop (𝓝 m) := h.comp tendsto_finset_range @@ -51,7 +49,15 @@ to `∏' i, f i`. -/ to `∑' i, f i`."] theorem Multipliable.tendsto_prod_tprod_nat {f : ℕ → M} (h : Multipliable f) : Tendsto (fun n ↦ ∏ i ∈ range n, f i) atTop (𝓝 (∏' i, f i)) := - tendsto_prod_nat h.hasProd + h.hasProd.tendsto_prod_nat + +@[deprecated (since := "2025-02-02")] +alias HasProd.Multipliable.tendsto_prod_tprod_nat := Multipliable.tendsto_prod_tprod_nat + +@[deprecated (since := "2025-02-02")] +alias HasSum.Multipliable.tendsto_sum_tsum_nat := Summable.tendsto_sum_tsum_nat + +namespace HasProd section ContinuousMul diff --git a/Mathlib/Topology/Algebra/LinearTopology.lean b/Mathlib/Topology/Algebra/LinearTopology.lean index b89ad49aba8593..4c30c6e05f89ec 100644 --- a/Mathlib/Topology/Algebra/LinearTopology.lean +++ b/Mathlib/Topology/Algebra/LinearTopology.lean @@ -28,7 +28,7 @@ hence our definition agrees with [N. Bourbaki, *Algebra II*, chapter 4, §2, n° * `IsLinearTopology R M`: the topology on `M` is `R`-linear, meaning that there exists a basis of neighborhoods of 0 consisting of `R`-submodules. Note that we don't impose that the topology is invariant by translation, so you'll often want to add `ContinuousConstVAdd M M` to get -something meaningless. To express that the topology of a ring `R` is linear, use +something meaningful. To express that the topology of a ring `R` is linear, use `[IsLinearTopology R R] [IsLinearTopology Rᵐᵒᵖ R]`. * `IsLinearTopology.mk_of_hasBasis`: a convenient constructor for `IsLinearTopology`. See also `IsLinearTopology.mk_of_hasBasis'`. @@ -134,7 +134,7 @@ variable (R) in /-- To show that `M` is linearly-topologized as an `R`-module, it suffices to show that it has a basis of neighborhoods of zero made of `R`-submodules. -Note: for technical reasons detailed in the module docstring, Lean sometimes struggle to find the +Note: for technical reasons detailed in the module docstring, Lean sometimes struggles to find the right `SMulMemClass` instance. See `IsLinearTopology.mk_of_hasBasis'` for a more explicit variant. -/ lemma mk_of_hasBasis {ι : Sort*} {S : Type*} [SetLike S M] diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 46fa435868f93d..e3d1356d4a232d 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -8,7 +8,6 @@ import Mathlib.LinearAlgebra.FreeModule.Finite.Matrix import Mathlib.Topology.Algebra.Module.Simple import Mathlib.Topology.Algebra.Module.Determinant import Mathlib.RingTheory.LocalRing.Basic -import Mathlib.RingTheory.Localization.FractionRing /-! # Finite dimensional topological vector spaces over complete fields diff --git a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean index d93e0243f2c58f..f0e637c7c1fdf7 100644 --- a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean +++ b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean @@ -448,7 +448,7 @@ instance instProd : IsModuleTopology R (M × N) := by let i₂ : N →ₗ[R] P := LinearMap.inr R M N rw [show (i : M × N → P) = (fun abcd ↦ abcd.1 + abcd.2 : P × P → P) ∘ - (fun ab ↦ (i₁ ab.1,i₂ ab.2)) by + (fun ab ↦ (i₁ ab.1, i₂ ab.2)) by ext ⟨a, b⟩ <;> aesop] -- and these maps are all continuous, hence `i` is too fun_prop diff --git a/Mathlib/Topology/Algebra/Order/Support.lean b/Mathlib/Topology/Algebra/Order/Support.lean new file mode 100644 index 00000000000000..bae2dcc6e5b15e --- /dev/null +++ b/Mathlib/Topology/Algebra/Order/Support.lean @@ -0,0 +1,46 @@ +/- +Copyright (c) 2025 Yoh Tanimoto. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yoh Tanimoto +-/ +import Mathlib.Algebra.Order.Group.Indicator +import Mathlib.Topology.Algebra.Support + +/-! +# The topological support of sup and inf of functions + +In a topological space `X` and a space `M` with `Sup` structure, for `f g : X → M` with compact +support, we show that `f ⊔ g` has compact support. Similarly, in `β` with `Inf` structure, `f ⊓ g` +has compact support if so do `f` and `g`. + +-/ + +variable {X M : Type*} [TopologicalSpace X] [One M] + +section SemilatticeSup + +variable [SemilatticeSup M] + +@[to_additive] +theorem HasCompactMulSupport.sup {f g : X → M} (hf : HasCompactMulSupport f) + (hg : HasCompactMulSupport g) : HasCompactMulSupport (f ⊔ g) := by + apply IsCompact.of_isClosed_subset (IsCompact.union hf hg) (isClosed_mulTSupport _) + rw [mulTSupport, mulTSupport, mulTSupport, ← closure_union] + apply closure_mono + exact Function.mulSupport_sup f g + +end SemilatticeSup + +section SemilatticeInf + +variable [SemilatticeInf M] + +@[to_additive] +theorem HasCompactMulSupport.inf {f g : X → M} (hf : HasCompactMulSupport f) + (hg : HasCompactMulSupport g) : HasCompactMulSupport (f ⊓ g) := by + apply IsCompact.of_isClosed_subset (IsCompact.union hf hg) (isClosed_mulTSupport _) + rw [mulTSupport, mulTSupport, mulTSupport, ← closure_union] + apply closure_mono + exact Function.mulSupport_inf f g + +end SemilatticeInf diff --git a/Mathlib/Topology/Algebra/StarSubalgebra.lean b/Mathlib/Topology/Algebra/StarSubalgebra.lean index 593aedee1802e7..3cde0521a3cb30 100644 --- a/Mathlib/Topology/Algebra/StarSubalgebra.lean +++ b/Mathlib/Topology/Algebra/StarSubalgebra.lean @@ -33,7 +33,7 @@ variable [TopologicalSpace A] [Semiring A] [Algebra R A] [StarRing A] [StarModul instance [TopologicalSemiring A] (s : StarSubalgebra R A) : TopologicalSemiring s := s.toSubalgebra.topologicalSemiring -/-- The `StarSubalgebra.inclusion` of a star subalgebra is an `Embedding`. -/ +/-- The `StarSubalgebra.inclusion` of a star subalgebra is an embedding. -/ lemma isEmbedding_inclusion {S₁ S₂ : StarSubalgebra R A} (h : S₁ ≤ S₂) : IsEmbedding (inclusion h) where eq_induced := Eq.symm induced_compose diff --git a/Mathlib/Topology/Algebra/Support.lean b/Mathlib/Topology/Algebra/Support.lean index e255b35efc4166..0c96aafbdfcca0 100644 --- a/Mathlib/Topology/Algebra/Support.lean +++ b/Mathlib/Topology/Algebra/Support.lean @@ -21,12 +21,11 @@ Furthermore, we say that `f` has compact support if the topological support of ` * `mulTSupport` & `tsupport` * `HasCompactMulSupport` & `HasCompactSupport` -## Implementation Notes +## TODO -* We write all lemmas for multiplicative functions, and use `@[to_additive]` to get the more common - additive versions. -* We do not put the definitions in the `Function` namespace, following many other topological - definitions that are in the root namespace (compare `Embedding` vs `Function.Embedding`). +The definitions have been put in the root namespace following many other topological definitions, +like `Embedding`. Since then, `Embedding` was renamed to `Topology.IsEmbedding`, so it might be +worth reconsidering namespacing the definitions here. -/ diff --git a/Mathlib/Topology/CWComplex.lean b/Mathlib/Topology/CWComplex.lean index 8dae230580f73c..5223795380a0e9 100644 --- a/Mathlib/Topology/CWComplex.lean +++ b/Mathlib/Topology/CWComplex.lean @@ -38,11 +38,12 @@ namespace RelativeCWComplex /-- The inclusion map from the `n`-sphere to the `(n + 1)`-disk. (For `n = -1`, this involves the empty space `𝕊 (-1)`. This is the reason why `sphere` takes `n : ℤ` as an input rather than `n : ℕ`.) -/ -def sphereInclusion (n : ℤ) : 𝕊 n ⟶ 𝔻 (n + 1) where - toFun := fun ⟨p, hp⟩ ↦ ⟨p, le_of_eq hp⟩ - continuous_toFun := ⟨fun t ⟨s, ⟨r, hro, hrs⟩, hst⟩ ↦ by - rw [isOpen_induced_iff, ← hst, ← hrs] - tauto⟩ +def sphereInclusion (n : ℤ) : 𝕊 n ⟶ 𝔻 (n + 1) := + TopCat.ofHom + { toFun := fun ⟨p, hp⟩ ↦ ⟨p, le_of_eq hp⟩ + continuous_toFun := ⟨fun t ⟨s, ⟨r, hro, hrs⟩, hst⟩ ↦ by + rw [isOpen_induced_iff, ← hst, ← hrs] + tauto⟩ } /-- A type witnessing that `X'` is obtained from `X` by attaching generalized cells `f : S ⟶ D` -/ structure AttachGeneralizedCells {S D : TopCat.{u}} (f : S ⟶ D) (X X' : TopCat.{u}) where diff --git a/Mathlib/Topology/Category/CompHaus/Basic.lean b/Mathlib/Topology/Category/CompHaus/Basic.lean index 705abf604355da..af6a4a57382146 100644 --- a/Mathlib/Topology/Category/CompHaus/Basic.lean +++ b/Mathlib/Topology/Category/CompHaus/Basic.lean @@ -34,9 +34,6 @@ The category `CompHaus` is defined using the structure `CompHausLike`. See the f universe v u --- This was a global instance prior to https://github.com/leanprover-community/mathlib4/pull/13170. We may experiment with removing it. -attribute [local instance] CategoryTheory.HasForget.instFunLike - open CategoryTheory CompHausLike /-- The category of compact Hausdorff spaces. -/ @@ -45,7 +42,7 @@ abbrev CompHaus := CompHausLike (fun _ ↦ True) namespace CompHaus instance : Inhabited CompHaus := - ⟨{ toTop := { α := PEmpty }, prop := trivial}⟩ + ⟨{ toTop := TopCat.of PEmpty, prop := trivial}⟩ instance : CoeSort CompHaus Type* := ⟨fun X => X.toTop⟩ @@ -85,17 +82,15 @@ Hausdorff spaces in topological spaces. -/ noncomputable def stoneCechEquivalence (X : TopCat.{u}) (Y : CompHaus.{u}) : (stoneCechObj X ⟶ Y) ≃ (X ⟶ compHausToTop.obj Y) where - toFun f := + toFun f := TopCat.ofHom { toFun := f ∘ stoneCechUnit - continuous_toFun := f.2.comp (@continuous_stoneCechUnit X _) } - invFun f := - { toFun := stoneCechExtend f.2 - continuous_toFun := continuous_stoneCechExtend f.2 } + continuous_toFun := f.hom.2.comp (@continuous_stoneCechUnit X _) } + invFun f := CompHausLike.ofHom _ + { toFun := stoneCechExtend f.hom.2 + continuous_toFun := continuous_stoneCechExtend f.hom.2 } left_inv := by rintro ⟨f : StoneCech X ⟶ Y, hf : Continuous f⟩ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` fails. - apply ContinuousMap.ext - intro (x : StoneCech X) + ext x refine congr_fun ?_ x apply Continuous.ext_on denseRange_stoneCechUnit (continuous_stoneCechExtend _) hf · rintro _ ⟨y, rfl⟩ @@ -103,9 +98,7 @@ noncomputable def stoneCechEquivalence (X : TopCat.{u}) (Y : CompHaus.{u}) : apply continuous_stoneCechUnit right_inv := by rintro ⟨f : (X : Type _) ⟶ Y, hf : Continuous f⟩ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` fails. - apply ContinuousMap.ext - intro + ext exact congr_fun (stoneCechExtend_extends hf) _ /-- The Stone-Cech compactification functor from topological spaces to compact Hausdorff spaces, @@ -157,7 +150,7 @@ def limitCone {J : Type v} [SmallCategory J] (F : J ⥤ CompHaus.{max v u}) : Li apply isClosed_iInter intro f apply isClosed_eq - · exact (ContinuousMap.continuous (F.map f)).comp (continuous_apply i) + · exact ((F.map f).hom.continuous).comp (continuous_apply i) · exact continuous_apply j is_hausdorff := show T2Space { u : ∀ j, F.obj j | ∀ {i j : J} (f : i ⟶ j), (F.map f) (u i) = u j } from @@ -185,7 +178,7 @@ theorem epi_iff_surjective {X Y : CompHaus.{u}} (f : X ⟶ Y) : Epi f ↔ Functi contrapose! rintro ⟨y, hy⟩ hf let C := Set.range f - have hC : IsClosed C := (isCompact_range f.continuous).isClosed + have hC : IsClosed C := (isCompact_range f.hom.continuous).isClosed let D := ({y} : Set Y) have hD : IsClosed D := isClosed_singleton have hCD : Disjoint C D := by @@ -196,27 +189,17 @@ theorem epi_iff_surjective {X Y : CompHaus.{u}} (f : X ⟶ Y) : Epi f ↔ Functi haveI : CompactSpace (ULift.{u} <| Set.Icc (0 : ℝ) 1) := Homeomorph.ulift.symm.compactSpace haveI : T2Space (ULift.{u} <| Set.Icc (0 : ℝ) 1) := Homeomorph.ulift.symm.t2Space let Z := of (ULift.{u} <| Set.Icc (0 : ℝ) 1) - let g : Y ⟶ Z := + let g : Y ⟶ Z := ofHom _ ⟨fun y' => ⟨⟨φ y', hφ01 y'⟩⟩, continuous_uLift_up.comp (φ.continuous.subtype_mk fun y' => hφ01 y')⟩ - let h : Y ⟶ Z := ⟨fun _ => ⟨⟨0, Set.left_mem_Icc.mpr zero_le_one⟩⟩, continuous_const⟩ + let h : Y ⟶ Z := ofHom _ + ⟨fun _ => ⟨⟨0, Set.left_mem_Icc.mpr zero_le_one⟩⟩, continuous_const⟩ have H : h = g := by rw [← cancel_epi f] - ext x - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` doesn't apply these two lemmas. - apply ULift.ext - apply Subtype.ext - dsimp - -- Porting note: This `change` is not ideal. - -- I think lean is having issues understanding when a `ContinuousMap` should be considered - -- as a morphism. - -- TODO(?): Make morphisms in `CompHaus` (and other topological categories) - -- into a one-field-structure. - change 0 = φ (f x) - simp only [hφ0 (Set.mem_range_self x), Pi.zero_apply] + ext x : 4 + simp [g, h, Z, hφ0 (Set.mem_range_self x)] apply_fun fun e => (e y).down.1 at H - dsimp [Z] at H - change 0 = φ y at H + dsimp [g, h, Z] at H simp only [hφ1 (Set.mem_singleton y), Pi.one_apply] at H exact zero_ne_one H · rw [← CategoryTheory.epi_iff_surjective] diff --git a/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean b/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean index 0c6a81f860057f..ae813659bce41d 100644 --- a/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/CompHaus/EffectiveEpi.lean @@ -29,8 +29,6 @@ universe u open CategoryTheory Limits CompHausLike -attribute [local instance] HasForget.instFunLike - namespace CompHaus open List in diff --git a/Mathlib/Topology/Category/CompHaus/Frm.lean b/Mathlib/Topology/Category/CompHaus/Frm.lean index 3e048912779004..b1cc66b4ec488c 100644 --- a/Mathlib/Topology/Category/CompHaus/Frm.lean +++ b/Mathlib/Topology/Category/CompHaus/Frm.lean @@ -18,9 +18,10 @@ open TopologicalSpace Opposite CategoryTheory @[simps] def topCatOpToFrm : TopCatᵒᵖ ⥤ Frm where obj X := Frm.of (Opens (unop X : TopCat)) - map f := Opens.comap <| Quiver.Hom.unop f + map f := Opens.comap <| (Quiver.Hom.unop f).hom map_id _ := Opens.comap_id -- Note, `CompHaus` is too strong. We only need `T0Space`. instance CompHausOpToFrame.faithful : (compHausToTop.op ⋙ topCatOpToFrm.{u}).Faithful := - ⟨fun h => Quiver.Hom.unop_inj <| Opens.comap_injective h⟩ + ⟨fun {X _ _ _} h => Quiver.Hom.unop_inj <| ConcreteCategory.ext <| + Opens.comap_injective (β := (unop X).toTop) h⟩ diff --git a/Mathlib/Topology/Category/CompHaus/Projective.lean b/Mathlib/Topology/Category/CompHaus/Projective.lean index 94949dbecb705a..e891c0742ab6c3 100644 --- a/Mathlib/Topology/Category/CompHaus/Projective.lean +++ b/Mathlib/Topology/Category/CompHaus/Projective.lean @@ -34,8 +34,6 @@ open CategoryTheory Function namespace CompHaus -attribute [local instance] HasForget.instFunLike - instance projective_ultrafilter (X : Type*) : Projective (of <| Ultrafilter X) where factors {Y Z} f g hg := by rw [epi_iff_surjective] at hg @@ -43,14 +41,10 @@ instance projective_ultrafilter (X : Type*) : Projective (of <| Ultrafilter X) w let t : X → Y := g' ∘ f ∘ (pure : X → Ultrafilter X) let h : Ultrafilter X → Y := Ultrafilter.extend t have hh : Continuous h := continuous_ultrafilter_extend _ - use ⟨h, hh⟩ - apply (forget CompHaus).map_injective - simp only [Functor.map_comp, ContinuousMap.coe_mk, coe_comp] - convert denseRange_pure.equalizer (g.continuous.comp hh) f.continuous _ - -- Porting note: We need to get the coercions to functions under control. - -- The next two lines should not be needed. - let g'' : ContinuousMap Y Z := g - have : g'' ∘ g' = id := hg'.comp_eq_id + use CompHausLike.ofHom _ ⟨h, hh⟩ + apply ConcreteCategory.coe_ext + have : g.hom ∘ g' = id := hg'.comp_eq_id + convert denseRange_pure.equalizer (g.hom.continuous.comp hh) f.hom.continuous _ -- This used to be `rw`, but we need `rw; rfl` after https://github.com/leanprover/lean4/pull/2644 rw [comp_assoc, ultrafilter_extend_extends, ← comp_assoc, this, id_comp] rfl @@ -59,7 +53,7 @@ instance projective_ultrafilter (X : Type*) : Projective (of <| Ultrafilter X) w the natural map `Ultrafilter X → X` is a projective presentation. -/ def projectivePresentation (X : CompHaus) : ProjectivePresentation X where p := of <| Ultrafilter X - f := ⟨_, continuous_ultrafilter_extend id⟩ + f := CompHausLike.ofHom _ ⟨_, continuous_ultrafilter_extend id⟩ projective := CompHaus.projective_ultrafilter X epi := ConcreteCategory.epi_of_surjective _ fun x => diff --git a/Mathlib/Topology/Category/CompHausLike/Basic.lean b/Mathlib/Topology/Category/CompHausLike/Basic.lean index 7dd9915be9c18c..d4bb77dfb91d8d 100644 --- a/Mathlib/Topology/Category/CompHausLike/Basic.lean +++ b/Mathlib/Topology/Category/CompHausLike/Basic.lean @@ -64,8 +64,6 @@ universe u open CategoryTheory -attribute [local instance] HasForget.instFunLike - variable (P : TopCat.{u} → Prop) /-- The type of Compact Hausdorff topological spaces satisfying an additional property `P`. -/ @@ -89,8 +87,8 @@ instance : CoeSort (CompHausLike P) (Type u) := instance category : Category (CompHausLike P) := InducedCategory.category toTop -instance hasForget : HasForget (CompHausLike P) := - InducedCategory.hasForget _ +instance concreteCategory : ConcreteCategory (CompHausLike P) (C(·, ·)) := + InducedCategory.concreteCategory toTop instance hasForget₂ : HasForget₂ (CompHausLike P) TopCat := InducedCategory.hasForget₂ _ @@ -101,30 +99,47 @@ variable (X : Type u) [TopologicalSpace X] [CompactSpace X] [T2Space X] class HasProp : Prop where hasProp : P (TopCat.of X) +instance (X : CompHausLike P) : HasProp P X := ⟨X.4⟩ + variable [HasProp P X] /-- A constructor for objects of the category `CompHausLike P`, taking a type, and bundling the compact Hausdorff topology found by typeclass inference. -/ -def of : CompHausLike P where +abbrev of : CompHausLike P where toTop := TopCat.of X is_compact := ‹_› is_hausdorff := ‹_› prop := HasProp.hasProp -@[simp] -theorem coe_of : (CompHausLike.of P X : Type _) = X := - rfl +theorem coe_of : (CompHausLike.of P X : Type _) = X := rfl @[simp] -theorem coe_id (X : CompHausLike P) : (𝟙 ((forget (CompHausLike P)).obj X)) = id := +theorem coe_id (X : CompHausLike P) : (𝟙 X : X → X) = id := rfl @[simp] theorem coe_comp {X Y Z : CompHausLike P} (f : X ⟶ Y) (g : Y ⟶ Z) : - ((forget (CompHausLike P)).map f ≫ (forget (CompHausLike P)).map g) = g ∘ f := + (f ≫ g : X → Z) = g ∘ f := rfl +section + +variable {X} {Y : Type u} [TopologicalSpace Y] [CompactSpace Y] [T2Space Y] [HasProp P Y] +variable {Z : Type u} [TopologicalSpace Z] [CompactSpace Z] [T2Space Z] [HasProp P Z] + +/-- Typecheck a continous map as a morphism in the category `CompHausLike P`. -/ +abbrev ofHom (f : C(X, Y)) : of P X ⟶ of P Y := ConcreteCategory.ofHom f + +@[simp] lemma hom_ofHom (f : C(X, Y)) : ConcreteCategory.hom (ofHom P f) = f := rfl + +@[simp] lemma ofHom_id : ofHom P (ContinuousMap.id X) = 𝟙 (of _ X) := rfl + +@[simp] lemma ofHom_comp (f : C(X, Y)) (g : C(Y, Z)) : + ofHom P (g.comp f) = ofHom _ f ≫ ofHom _ g := rfl + +end + -- Note (https://github.com/leanprover-community/mathlib4/issues/10754): Lean does not see through the forgetful functor here instance (X : CompHausLike.{u} P) : TopologicalSpace ((forget (CompHausLike P)).obj X) := inferInstanceAs (TopologicalSpace X.toTop) @@ -140,7 +155,7 @@ instance (X : CompHausLike.{u} P) : T2Space ((forget (CompHausLike P)).obj X) := variable {P} /-- If `P` imples `P'`, then there is a functor from `CompHausLike P` to `CompHausLike P'`. -/ -@[simps] +@[simps map] def toCompHausLike {P P' : TopCat → Prop} (h : ∀ (X : CompHausLike P), P X.toTop → P' X.toTop) : CompHausLike P ⥤ CompHausLike P' where obj X := @@ -166,7 +181,7 @@ end variable (P) /-- The fully faithful embedding of `CompHausLike P` in `TopCat`. -/ -@[simps!] +@[simps! map] def compHausLikeToTop : CompHausLike.{u} P ⥤ TopCat.{u} := inducedFunctor _ -- deriving Full, Faithful -- Porting note: deriving fails, adding manually. @@ -200,16 +215,16 @@ theorem mono_iff_injective {X Y : CompHausLike.{u} P} (f : X ⟶ Y) : Mono f ↔ Function.Injective f := by constructor · intro hf x₁ x₂ h - let g₁ : X ⟶ X := ⟨fun _ => x₁, continuous_const⟩ - let g₂ : X ⟶ X := ⟨fun _ => x₂, continuous_const⟩ + let g₁ : X ⟶ X := ofHom _ ⟨fun _ => x₁, continuous_const⟩ + let g₂ : X ⟶ X := ofHom _ ⟨fun _ => x₂, continuous_const⟩ have : g₁ ≫ f = g₂ ≫ f := by ext; exact h - exact ContinuousMap.congr_fun ((cancel_mono _).mp this) x₁ + exact CategoryTheory.congr_fun ((cancel_mono _).mp this) x₁ · rw [← CategoryTheory.mono_iff_injective] apply (forget (CompHausLike P)).mono_of_mono_map /-- Any continuous function on compact Hausdorff spaces is a closed map. -/ theorem isClosedMap {X Y : CompHausLike.{u} P} (f : X ⟶ Y) : IsClosedMap f := fun _ hC => - (hC.isCompact.image f.continuous).isClosed + (hC.isCompact.image f.hom.continuous).isClosed /-- Any continuous bijection of compact Hausdorff spaces is an isomorphism. -/ theorem isIso_of_bijective {X Y : CompHausLike.{u} P} (f : X ⟶ Y) (bij : Function.Bijective f) : @@ -220,7 +235,7 @@ theorem isIso_of_bijective {X Y : CompHausLike.{u} P} (f : X ⟶ Y) (bij : Funct intro S hS rw [← E.image_eq_preimage] exact isClosedMap f S hS - refine ⟨⟨⟨E.symm, hE⟩, ?_, ?_⟩⟩ + refine ⟨⟨ofHom _ ⟨E.symm, hE⟩, ?_, ?_⟩⟩ · ext x apply E.symm_apply_apply · ext x @@ -258,7 +273,7 @@ def isoEquivHomeo {X Y : CompHausLike.{u} P} : (X ≅ Y) ≃ (X ≃ₜ Y) where /-- A constant map as a morphism in `CompHausLike` -/ def const {P : TopCat.{u} → Prop} (T : CompHausLike.{u} P) {S : CompHausLike.{u} P} (s : S) : T ⟶ S := - ContinuousMap.const _ s + ofHom _ (ContinuousMap.const _ s) lemma const_comp {P : TopCat.{u} → Prop} {S T U : CompHausLike.{u} P} (s : S) (g : S ⟶ U) : T.const s ≫ g = T.const (g s) := diff --git a/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean b/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean index 967d8531b8e53d..91a9ffc3bf187f 100644 --- a/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/CompHausLike/EffectiveEpi.lean @@ -21,8 +21,6 @@ universe u open CategoryTheory Limits Topology -attribute [local instance] HasForget.instFunLike - namespace CompHausLike variable {P : TopCat.{u} → Prop} @@ -33,18 +31,30 @@ If `π` is a surjective morphism in `CompHausLike P`, then it is an effective ep noncomputable def effectiveEpiStruct {B X : CompHausLike P} (π : X ⟶ B) (hπ : Function.Surjective π) : EffectiveEpiStruct π where - desc e h := (IsQuotientMap.of_surjective_continuous hπ π.continuous).lift e fun a b hab ↦ - DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ - (by ext; exact hab)) a - fac e h := ((IsQuotientMap.of_surjective_continuous hπ π.continuous).lift_comp e - fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ + desc e h := ofHom _ ((IsQuotientMap.of_surjective_continuous hπ π.hom.continuous).lift e.hom + fun a b hab ↦ + CategoryTheory.congr_fun (h + (ofHom _ ⟨fun _ ↦ a, continuous_const⟩) + (ofHom _ ⟨fun _ ↦ b, continuous_const⟩) + (by ext; exact hab)) a) + fac e h := TopCat.hom_ext ((IsQuotientMap.of_surjective_continuous hπ π.hom.continuous).lift_comp + e.hom + fun a b hab ↦ CategoryTheory.congr_fun (h + (ofHom _ ⟨fun _ ↦ a, continuous_const⟩) + (ofHom _ ⟨fun _ ↦ b, continuous_const⟩) (by ext; exact hab)) a) uniq e h g hm := by - suffices g = (IsQuotientMap.of_surjective_continuous hπ π.continuous).liftEquiv ⟨e, - fun a b hab ↦ DFunLike.congr_fun - (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) - a⟩ by assumption - rw [← Equiv.symm_apply_eq (IsQuotientMap.of_surjective_continuous hπ π.continuous).liftEquiv] + suffices g = ofHom _ ((IsQuotientMap.of_surjective_continuous hπ π.hom.continuous).liftEquiv + ⟨e.hom, + fun a b hab ↦ CategoryTheory.congr_fun + (h + (ofHom _ ⟨fun _ ↦ a, continuous_const⟩) + (ofHom _ ⟨fun _ ↦ b, continuous_const⟩) + (by ext; exact hab)) + a⟩) by assumption + apply ConcreteCategory.ext + rw [hom_ofHom, + ← Equiv.symm_apply_eq (IsQuotientMap.of_surjective_continuous hπ π.hom.continuous).liftEquiv] ext simp only [IsQuotientMap.liftEquiv_symm_apply_coe, ContinuousMap.comp_apply, ← hm] rfl diff --git a/Mathlib/Topology/Category/CompHausLike/Limits.lean b/Mathlib/Topology/Category/CompHausLike/Limits.lean index a818fb4571e8ba..92c8d6db22fc37 100644 --- a/Mathlib/Topology/Category/CompHausLike/Limits.lean +++ b/Mathlib/Topology/Category/CompHausLike/Limits.lean @@ -40,8 +40,6 @@ namespace CompHausLike universe w u -attribute [local instance] HasForget.instFunLike - section FiniteCoproducts variable {P : TopCat.{max u w} → Prop} {α : Type w} [Finite α] (X : α → CompHausLike P) @@ -63,9 +61,10 @@ def finiteCoproduct : CompHausLike P := CompHausLike.of P (Σ (a : α), X a) /-- The inclusion of one of the factors into the explicit finite coproduct. -/ -def finiteCoproduct.ι (a : α) : X a ⟶ finiteCoproduct X where - toFun := fun x ↦ ⟨a, x⟩ - continuous_toFun := continuous_sigmaMk (σ := fun a ↦ X a) +def finiteCoproduct.ι (a : α) : X a ⟶ finiteCoproduct X := + ofHom _ + { toFun := fun x ↦ ⟨a, x⟩ + continuous_toFun := continuous_sigmaMk (σ := fun a ↦ X a) } /-- To construct a morphism from the explicit finite coproduct, it suffices to @@ -73,11 +72,12 @@ specify a morphism from each of its factors. This is essentially the universal property of the coproduct. -/ def finiteCoproduct.desc {B : CompHausLike P} (e : (a : α) → (X a ⟶ B)) : - finiteCoproduct X ⟶ B where - toFun := fun ⟨a, x⟩ ↦ e a x - continuous_toFun := by - apply continuous_sigma - intro a; exact (e a).continuous + finiteCoproduct X ⟶ B := + ofHom _ + { toFun := fun ⟨a, x⟩ ↦ e a x + continuous_toFun := by + apply continuous_sigma + intro a; exact (e a).hom.continuous } @[reassoc (attr := simp)] lemma finiteCoproduct.ι_desc {B : CompHausLike P} (e : (a : α) → (X a ⟶ B)) (a : α) : @@ -100,7 +100,7 @@ def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cofan X) := (fun s ↦ desc _ fun a ↦ s.inj a) (fun _ _ ↦ ι_desc _ _ _) fun _ _ hm ↦ finiteCoproduct.hom_ext _ _ _ fun a ↦ - (DFunLike.ext _ _ fun t ↦ congrFun (congrArg DFunLike.coe (hm a)) t) + (ConcreteCategory.hom_ext _ _ fun t ↦ congrFun (congrArg _ (hm a)) t) lemma finiteCoproduct.ι_injective (a : α) : Function.Injective (finiteCoproduct.ι X a) := by intro x y hxy @@ -207,23 +207,25 @@ pairs `(x,y)` such that `f x = g y`, with the topology induced by the product. def pullback : CompHausLike P := letI set := { xy : X × Y | f xy.fst = g xy.snd } haveI : CompactSpace set := - isCompact_iff_compactSpace.mp (isClosed_eq (f.continuous.comp continuous_fst) - (g.continuous.comp continuous_snd)).isCompact + isCompact_iff_compactSpace.mp (isClosed_eq (f.hom.continuous.comp continuous_fst) + (g.hom.continuous.comp continuous_snd)).isCompact CompHausLike.of P set /-- The projection from the pullback to the first component. -/ -def pullback.fst : pullback f g ⟶ X where - toFun := fun ⟨⟨x, _⟩, _⟩ ↦ x - continuous_toFun := Continuous.comp continuous_fst continuous_subtype_val +def pullback.fst : pullback f g ⟶ X := + TopCat.ofHom + { toFun := fun ⟨⟨x, _⟩, _⟩ ↦ x + continuous_toFun := Continuous.comp continuous_fst continuous_subtype_val } /-- The projection from the pullback to the second component. -/ -def pullback.snd : pullback f g ⟶ Y where - toFun := fun ⟨⟨_,y⟩,_⟩ ↦ y - continuous_toFun := Continuous.comp continuous_snd continuous_subtype_val +def pullback.snd : pullback f g ⟶ Y := + TopCat.ofHom + { toFun := fun ⟨⟨_,y⟩,_⟩ ↦ y + continuous_toFun := Continuous.comp continuous_snd continuous_subtype_val } @[reassoc] lemma pullback.condition : pullback.fst f g ≫ f = pullback.snd f g ≫ g := by @@ -235,12 +237,13 @@ which are compatible with the maps to the base. This is essentially the universal property of the pullback. -/ def pullback.lift {Z : CompHausLike P} (a : Z ⟶ X) (b : Z ⟶ Y) (w : a ≫ f = b ≫ g) : - Z ⟶ pullback f g where - toFun := fun z ↦ ⟨⟨a z, b z⟩, by apply_fun (fun q ↦ q z) at w; exact w⟩ - continuous_toFun := by - apply Continuous.subtype_mk - rw [continuous_prod_mk] - exact ⟨a.continuous, b.continuous⟩ + Z ⟶ pullback f g := + TopCat.ofHom + { toFun := fun z ↦ ⟨⟨a z, b z⟩, by apply_fun (fun q ↦ q z) at w; exact w⟩ + continuous_toFun := by + apply Continuous.subtype_mk + rw [continuous_prod_mk] + exact ⟨a.hom.continuous, b.hom.continuous⟩ } @[reassoc (attr := simp)] lemma pullback.lift_fst {Z : CompHausLike P} (a : Z ⟶ X) (b : Z ⟶ Y) (w : a ≫ f = b ≫ g) : @@ -371,7 +374,7 @@ variable {P : TopCat.{u} → Prop} def isTerminalPUnit [HasProp P PUnit.{u+1}] : IsTerminal (CompHausLike.of P PUnit.{u + 1}) := haveI : ∀ X, Unique (X ⟶ CompHausLike.of P PUnit.{u + 1}) := fun _ ↦ - ⟨⟨⟨fun _ ↦ PUnit.unit, continuous_const⟩⟩, fun _ ↦ rfl⟩ + ⟨⟨ofHom _ ⟨fun _ ↦ PUnit.unit, continuous_const⟩⟩, fun _ ↦ rfl⟩ Limits.IsTerminal.ofUnique _ end Terminal diff --git a/Mathlib/Topology/Category/CompHausLike/SigmaComparison.lean b/Mathlib/Topology/Category/CompHausLike/SigmaComparison.lean index fc574ef66ea359..85db117945a69d 100644 --- a/Mathlib/Topology/Category/CompHausLike/SigmaComparison.lean +++ b/Mathlib/Topology/Category/CompHausLike/SigmaComparison.lean @@ -36,7 +36,7 @@ The comparison map from the value of a condensed set on a finite coproduct to th values on the components. -/ def sigmaComparison : X.obj ⟨(of P ((a : α) × σ a))⟩ ⟶ ((a : α) → X.obj ⟨of P (σ a)⟩) := - fun x a ↦ X.map ⟨Sigma.mk a, continuous_sigmaMk⟩ x + fun x a ↦ X.map (ofHom _ ⟨Sigma.mk a, continuous_sigmaMk⟩).op x theorem sigmaComparison_eq_comp_isos : sigmaComparison X σ = (X.mapIso (opCoproductIsoProduct' diff --git a/Mathlib/Topology/Category/CompactlyGenerated.lean b/Mathlib/Topology/Category/CompactlyGenerated.lean index 3defff92c5ca02..b98907c3e182e4 100644 --- a/Mathlib/Topology/Category/CompactlyGenerated.lean +++ b/Mathlib/Topology/Category/CompactlyGenerated.lean @@ -37,7 +37,7 @@ structure CompactlyGenerated where namespace CompactlyGenerated instance : Inhabited CompactlyGenerated.{u, w} := - ⟨{ toTop := { α := ULift (Fin 37) } }⟩ + ⟨{ toTop := TopCat.of (ULift (Fin 37)) }⟩ instance : CoeSort CompactlyGenerated Type* := ⟨fun X => X.toTop⟩ @@ -47,16 +47,25 @@ attribute [instance] is_compactly_generated instance : Category.{w, w+1} CompactlyGenerated.{u, w} := InducedCategory.category toTop -instance : HasForget.{w} CompactlyGenerated.{u, w} := - InducedCategory.hasForget _ +instance : ConcreteCategory.{w} CompactlyGenerated.{u, w} (C(·, ·)) := + InducedCategory.concreteCategory toTop variable (X : Type w) [TopologicalSpace X] [UCompactlyGeneratedSpace.{u} X] /-- Constructor for objects of the category `CompactlyGenerated`. -/ -def of : CompactlyGenerated.{u, w} where +abbrev of : CompactlyGenerated.{u, w} where toTop := TopCat.of X is_compactly_generated := ‹_› +section + +variable {X} {Y : Type w} [TopologicalSpace Y] [UCompactlyGeneratedSpace.{u} Y] + +/-- Typecheck a `ContinuousMap` as a morphism in `CompactlyGenerated`. -/ +abbrev ofHom (f : C(X, Y)) : of X ⟶ of Y := ConcreteCategory.ofHom f + +end + /-- The fully faithful embedding of `CompactlyGenerated` in `TopCat`. -/ @[simps!] def compactlyGeneratedToTop : CompactlyGenerated.{u, w} ⥤ TopCat.{w} := @@ -73,8 +82,8 @@ instance : compactlyGeneratedToTop.{u, w}.Faithful := fullyFaithfulCompactlyGene /-- Construct an isomorphism from a homeomorphism. -/ @[simps hom inv] def isoOfHomeo {X Y : CompactlyGenerated.{u, w}} (f : X ≃ₜ Y) : X ≅ Y where - hom := ⟨f, f.continuous⟩ - inv := ⟨f.symm, f.symm.continuous⟩ + hom := ofHom ⟨f, f.continuous⟩ + inv := ofHom ⟨f.symm, f.symm.continuous⟩ hom_inv_id := by ext x exact f.symm_apply_apply x @@ -87,10 +96,10 @@ def isoOfHomeo {X Y : CompactlyGenerated.{u, w}} (f : X ≃ₜ Y) : X ≅ Y wher def homeoOfIso {X Y : CompactlyGenerated.{u, w}} (f : X ≅ Y) : X ≃ₜ Y where toFun := f.hom invFun := f.inv - left_inv x := by simp - right_inv x := by simp - continuous_toFun := f.hom.continuous - continuous_invFun := f.inv.continuous + left_inv := f.hom_inv_id_apply + right_inv := f.inv_hom_id_apply + continuous_toFun := f.hom.hom.continuous + continuous_invFun := f.inv.hom.continuous /-- The equivalence between isomorphisms in `CompactlyGenerated` and homeomorphisms of topological spaces. -/ diff --git a/Mathlib/Topology/Category/Compactum.lean b/Mathlib/Topology/Category/Compactum.lean index 32dbfd00182a5a..e336d80b202f0e 100644 --- a/Mathlib/Topology/Category/Compactum.lean +++ b/Mathlib/Topology/Category/Compactum.lean @@ -407,8 +407,8 @@ end Compactum /-- The functor functor from Compactum to CompHaus. -/ def compactumToCompHaus : Compactum ⥤ CompHaus where - obj X := { toTop := { α := X }, prop := trivial } - map := fun f => + obj X := { toTop := TopCat.of X, prop := trivial } + map := fun f => CompHausLike.ofHom _ { toFun := f continuous_toFun := Compactum.continuous_of_hom _ } @@ -416,7 +416,7 @@ namespace compactumToCompHaus /-- The functor `compactumToCompHaus` is full. -/ instance full : compactumToCompHaus.{u}.Full where - map_surjective f := ⟨Compactum.homOfContinuous f.1 f.2, rfl⟩ + map_surjective f := ⟨Compactum.homOfContinuous f.1 f.hom.2, rfl⟩ /-- The functor `compactumToCompHaus` is faithful. -/ instance faithful : compactumToCompHaus.Faithful where @@ -425,18 +425,18 @@ instance faithful : compactumToCompHaus.Faithful where intro _ _ _ _ h -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` gets confused by coercion using forget. apply Monad.Algebra.Hom.ext - apply congrArg (fun f => f.toFun) h + apply congrArg (fun f => f.hom.toFun) h /-- This definition is used to prove essential surjectivity of `compactumToCompHaus`. -/ def isoOfTopologicalSpace {D : CompHaus} : compactumToCompHaus.obj (Compactum.ofTopologicalSpace D) ≅ D where - hom := + hom := CompHausLike.ofHom _ { toFun := id continuous_toFun := continuous_def.2 fun _ h => by rw [isOpen_iff_ultrafilter'] at h exact h } - inv := + inv := CompHausLike.ofHom _ { toFun := id continuous_toFun := continuous_def.2 fun _ h1 => by diff --git a/Mathlib/Topology/Category/DeltaGenerated.lean b/Mathlib/Topology/Category/DeltaGenerated.lean index 7024dfe3ba5b2f..42cdecc4815f8c 100644 --- a/Mathlib/Topology/Category/DeltaGenerated.lean +++ b/Mathlib/Topology/Category/DeltaGenerated.lean @@ -41,11 +41,11 @@ attribute [instance] deltaGenerated instance : LargeCategory.{u} DeltaGenerated.{u} := InducedCategory.category toTop -instance : HasForget.{u} DeltaGenerated.{u} := - InducedCategory.hasForget _ +instance : ConcreteCategory.{u} DeltaGenerated.{u} (C(·, ·)) := + InducedCategory.concreteCategory toTop /-- Constructor for objects of the category `DeltaGenerated` -/ -def of (X : Type u) [TopologicalSpace X] [DeltaGeneratedSpace X] : DeltaGenerated.{u} where +abbrev of (X : Type u) [TopologicalSpace X] [DeltaGeneratedSpace X] : DeltaGenerated.{u} where toTop := TopCat.of X deltaGenerated := ‹_› @@ -67,19 +67,19 @@ instance : deltaGeneratedToTop.{u}.Faithful := fullyFaithfulDeltaGeneratedToTop. @[simps!] def topToDeltaGenerated : TopCat.{u} ⥤ DeltaGenerated.{u} where obj X := of (DeltaGeneratedSpace.of X) - map {_ Y} f := ⟨f, (continuous_to_deltaGenerated (Y := Y)).mpr <| - continuous_le_dom deltaGenerated_le f.continuous⟩ + map {_ Y} f := TopCat.ofHom ⟨f, (continuous_to_deltaGenerated (Y := Y)).mpr <| + continuous_le_dom deltaGenerated_le f.hom.continuous⟩ instance : topToDeltaGenerated.{u}.Faithful := - ⟨fun h ↦ by ext x; exact congrFun (congrArg ContinuousMap.toFun h) x⟩ + ⟨fun h ↦ by ext x; exact CategoryTheory.congr_fun h x⟩ /-- The adjunction between the forgetful functor `DeltaGenerated ⥤ TopCat` and its coreflector. -/ def coreflectorAdjunction : deltaGeneratedToTop ⊣ topToDeltaGenerated := Adjunction.mkOfUnitCounit { unit := { - app X := ⟨id, continuous_iff_coinduced_le.mpr (eq_deltaGenerated (X := X)).le⟩ } + app X := TopCat.ofHom ⟨id, continuous_iff_coinduced_le.mpr (eq_deltaGenerated (X := X)).le⟩ } counit := { - app X := ⟨DeltaGeneratedSpace.counit, DeltaGeneratedSpace.continuous_counit⟩ }} + app X := TopCat.ofHom ⟨DeltaGeneratedSpace.counit, DeltaGeneratedSpace.continuous_counit⟩ }} /-- The category of delta-generated spaces is coreflective in the category of topological spaces. -/ instance deltaGeneratedToTop.coreflective : Coreflective deltaGeneratedToTop where diff --git a/Mathlib/Topology/Category/FinTopCat.lean b/Mathlib/Topology/Category/FinTopCat.lean index ff492e48cdfcfe..34fca348c9e7f8 100644 --- a/Mathlib/Topology/Category/FinTopCat.lean +++ b/Mathlib/Topology/Category/FinTopCat.lean @@ -22,13 +22,13 @@ open CategoryTheory /-- A bundled finite topological space. -/ structure FinTopCat where /-- carrier of a finite topological space. -/ - toTop : TopCat.{u} + toTop : TopCat.{u} -- TODO: turn this into an `extends`? [fintype : Fintype toTop] namespace FinTopCat instance : Inhabited FinTopCat := - ⟨{ toTop := { α := PEmpty } }⟩ + ⟨{ toTop := TopCat.of PEmpty }⟩ instance : CoeSort FinTopCat (Type u) := ⟨fun X => X.toTop⟩ @@ -38,14 +38,8 @@ attribute [instance] fintype instance : Category FinTopCat := InducedCategory.category toTop -instance : HasForget FinTopCat := - InducedCategory.hasForget _ - -instance (X : FinTopCat) : TopologicalSpace ((forget FinTopCat).obj X) := - inferInstanceAs <| TopologicalSpace X - -instance (X : FinTopCat) : Fintype ((forget FinTopCat).obj X) := - X.fintype +instance : ConcreteCategory FinTopCat (C(·, ·)) := + InducedCategory.concreteCategory toTop /-- Construct a bundled `FinTopCat` from the underlying type and the appropriate typeclasses. -/ def of (X : Type u) [Fintype X] [TopologicalSpace X] : FinTopCat where @@ -59,7 +53,7 @@ theorem coe_of (X : Type u) [Fintype X] [TopologicalSpace X] : /-- The forgetful functor to `FintypeCat`. -/ instance : HasForget₂ FinTopCat FintypeCat := - HasForget₂.mk' (fun X ↦ FintypeCat.of X) (fun _ ↦ rfl) (fun f ↦ f.toFun) HEq.rfl + HasForget₂.mk' (fun X ↦ FintypeCat.of X) (fun _ ↦ rfl) (fun f ↦ f.hom.toFun) HEq.rfl instance (X : FinTopCat) : TopologicalSpace ((forget₂ FinTopCat FintypeCat).obj X) := inferInstanceAs <| TopologicalSpace X diff --git a/Mathlib/Topology/Category/LightProfinite/AsLimit.lean b/Mathlib/Topology/Category/LightProfinite/AsLimit.lean index a7680b1ad8bb98..a7201f66bea5d3 100644 --- a/Mathlib/Topology/Category/LightProfinite/AsLimit.lean +++ b/Mathlib/Topology/Category/LightProfinite/AsLimit.lean @@ -20,8 +20,6 @@ noncomputable section open CategoryTheory Limits CompHausLike -attribute [local instance] HasForget.instFunLike - namespace LightProfinite universe u @@ -79,7 +77,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 only [toCompHausLike_obj, Functor.comp_obj, toCompHausLike_map, coe_of] + simp only [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 7644318842d136..02867f3283bab5 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -32,12 +32,6 @@ where possible, try to keep them in sync -/ universe v u -/- -Previously, this had accidentally been made a global instance, -and we now turn it on locally when convenient. --/ -attribute [local instance] CategoryTheory.HasForget.instFunLike - open CategoryTheory Limits Opposite FintypeCat Topology TopologicalSpace CompHausLike /-- `LightProfinite` is the category of second countable profinite spaces. -/ @@ -100,10 +94,10 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `LightProfinite`, endowing a finite type with the discrete topology. -/ -@[simps! -isSimp map_apply] +@[simps! -isSimp map_hom_apply] def FintypeCat.toLightProfinite : FintypeCat ⥤ LightProfinite where obj A := LightProfinite.of A - map f := ⟨f, by continuity⟩ + map f := CompHausLike.ofHom _ ⟨f, by continuity⟩ /-- `FintypeCat.toLightProfinite` is fully faithful. -/ def FintypeCat.toLightProfiniteFullyFaithful : toLightProfinite.FullyFaithful where @@ -205,7 +199,7 @@ theorem epi_iff_surjective {X Y : LightProfinite.{u}} (f : X ⟶ Y) : contrapose! rintro ⟨y, hy⟩ hf let C := Set.range f - have hC : IsClosed C := (isCompact_range f.continuous).isClosed + have hC : IsClosed C := (isCompact_range f.hom.continuous).isClosed let U := Cᶜ have hyU : y ∈ U := by refine Set.mem_compl ?_ @@ -215,22 +209,20 @@ theorem epi_iff_surjective {X Y : LightProfinite.{u}} (f : X ⟶ Y) : obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_isClopen.mem_nhds_iff.mp hUy classical let Z := of (ULift.{u} <| Fin 2) - let g : Y ⟶ Z := ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ - let h : Y ⟶ Z := ⟨fun _ => ⟨1⟩, continuous_const⟩ + let g : Y ⟶ Z := CompHausLike.ofHom _ + ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ + let h : Y ⟶ Z := CompHausLike.ofHom _ ⟨fun _ => ⟨1⟩, continuous_const⟩ have H : h = g := by rw [← cancel_epi f] ext x - apply ULift.ext dsimp [g, LocallyConstant.ofIsClopen] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [CategoryTheory.comp_apply, ContinuousMap.coe_mk, - CategoryTheory.comp_apply, ContinuousMap.coe_mk, Function.comp_apply, if_neg] + rw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, hom_ofHom, ContinuousMap.coe_mk, + Function.comp_apply, if_neg] refine mt (fun α => hVU α) ?_ - simp only [U, C, Set.mem_range_self, not_true, not_false_iff, Set.mem_compl_iff] + simp [U, C] apply_fun fun e => (e y).down at H dsimp [g, LocallyConstant.ofIsClopen] at H - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, Function.comp_apply, if_pos hyV] at H + rw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, Function.comp_apply, if_pos hyV] at H exact top_ne_bot H · rw [← CategoryTheory.epi_iff_surjective] apply (forget LightProfinite).epi_of_epi_map @@ -257,7 +249,8 @@ def toProfinite (S : LightDiagram) : Profinite := S.cone.pt @[simps!] instance : Category LightDiagram := InducedCategory.category toProfinite -instance hasForget : HasForget LightDiagram := InducedCategory.hasForget _ +instance hasForget : ConcreteCategory LightDiagram (fun X Y => C(X.toProfinite, Y.toProfinite)) := + InducedCategory.concreteCategory toProfinite end LightDiagram @@ -321,13 +314,13 @@ instance (S : LightDiagram.{u}) : SecondCountableTopology S.cone.pt := by refine @Pi.finite _ _ ?_ _ simp only [Functor.comp_obj] exact show (Finite (S.diagram.obj _)) from inferInstance - · exact fun a ↦ a.snd.comap (S.cone.π.app ⟨a.fst⟩) + · exact fun a ↦ a.snd.comap (S.cone.π.app ⟨a.fst⟩).hom · intro a obtain ⟨n, g, h⟩ := Profinite.exists_locallyConstant S.cone S.isLimit a exact ⟨⟨unop n, g⟩, h.symm⟩ /-- The inverse part of the equivalence `LightProfinite ≌ LightDiagram` -/ -@[simps] +@[simps obj map] def lightDiagramToLightProfinite : LightDiagram.{u} ⥤ LightProfinite.{u} where obj X := LightProfinite.of X.cone.pt map f := f diff --git a/Mathlib/Topology/Category/LightProfinite/EffectiveEpi.lean b/Mathlib/Topology/Category/LightProfinite/EffectiveEpi.lean index 156e8287d0dc72..fac00ec92aca85 100644 --- a/Mathlib/Topology/Category/LightProfinite/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/LightProfinite/EffectiveEpi.lean @@ -19,8 +19,6 @@ universe u open CategoryTheory Limits CompHausLike -attribute [local instance] HasForget.instFunLike - namespace LightProfinite theorem effectiveEpi_iff_surjective {X Y : LightProfinite.{u}} (f : X ⟶ Y) : diff --git a/Mathlib/Topology/Category/Locale.lean b/Mathlib/Topology/Category/Locale.lean index 208b21c604fbc8..9991b17f693f1e 100644 --- a/Mathlib/Topology/Category/Locale.lean +++ b/Mathlib/Topology/Category/Locale.lean @@ -53,4 +53,4 @@ def topToLocale : TopCat ⥤ Locale := instance CompHausToLocale.faithful : (compHausToTop ⋙ topToLocale.{u}).Faithful := ⟨fun h => by dsimp at h - exact Opens.comap_injective (Quiver.Hom.op_inj h)⟩ + exact ConcreteCategory.ext (Opens.comap_injective (Quiver.Hom.op_inj h))⟩ diff --git a/Mathlib/Topology/Category/Profinite/AsLimit.lean b/Mathlib/Topology/Category/Profinite/AsLimit.lean index a417cb6fdd2840..c8ea63847dd6b7 100644 --- a/Mathlib/Topology/Category/Profinite/AsLimit.lean +++ b/Mathlib/Topology/Category/Profinite/AsLimit.lean @@ -54,7 +54,8 @@ abbrev diagram : DiscreteQuotient X ⥤ Profinite := /-- A cone over `X.diagram` whose cone point is `X`. -/ def asLimitCone : CategoryTheory.Limits.Cone X.diagram := { pt := X - π := { app := fun S => ⟨S.proj, IsLocallyConstant.continuous (S.proj_isLocallyConstant)⟩ } } + π := { app := fun S => CompHausLike.ofHom (Y := X.diagram.obj S) _ + ⟨S.proj, IsLocallyConstant.continuous (S.proj_isLocallyConstant)⟩ } } instance isIso_asLimitCone_lift : IsIso ((limitConeIsLimit.{u, u} X.diagram).lift X.asLimitCone) := CompHausLike.isIso_of_bijective _ diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index 30959b584c678b..3e4dfad34be93a 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -109,12 +109,12 @@ spaces in compact Hausdorff spaces. -/ def Profinite.toCompHausEquivalence (X : CompHaus.{u}) (Y : Profinite.{u}) : (CompHaus.toProfiniteObj X ⟶ Y) ≃ (X ⟶ profiniteToCompHaus.obj Y) where - toFun f := f.comp ⟨Quotient.mk'', continuous_quotient_mk'⟩ - invFun g := - { toFun := Continuous.connectedComponentsLift g.2 - continuous_toFun := Continuous.connectedComponentsLift_continuous g.2 } - left_inv _ := ContinuousMap.ext <| ConnectedComponents.surjective_coe.forall.2 fun _ => rfl - right_inv _ := ContinuousMap.ext fun _ => rfl + toFun f := ofHom _ (f.hom.comp ⟨Quotient.mk'', continuous_quotient_mk'⟩) + invFun g := TopCat.ofHom + { toFun := Continuous.connectedComponentsLift g.hom.2 + continuous_toFun := Continuous.connectedComponentsLift_continuous g.hom.2 } + left_inv _ := TopCat.ext <| ConnectedComponents.surjective_coe.forall.2 fun _ => rfl + right_inv _ := TopCat.ext fun _ => rfl /-- The connected_components functor from compact Hausdorff spaces to profinite spaces, left adjoint to the inclusion functor. @@ -140,10 +140,10 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `Profinite`, endowing a finite type with the discrete topology. -/ -@[simps -isSimp map_apply] +@[simps! -isSimp map_hom_apply] def FintypeCat.toProfinite : FintypeCat ⥤ Profinite where obj A := Profinite.of A - map f := ⟨f, by continuity⟩ + map f := ofHom _ ⟨f, by continuity⟩ /-- `FintypeCat.toLightProfinite` is fully faithful. -/ def FintypeCat.toProfiniteFullyFaithful : toProfinite.FullyFaithful where @@ -222,7 +222,7 @@ theorem epi_iff_surjective {X Y : Profinite.{u}} (f : X ⟶ Y) : Epi f ↔ Funct contrapose! rintro ⟨y, hy⟩ hf let C := Set.range f - have hC : IsClosed C := (isCompact_range f.continuous).isClosed + have hC : IsClosed C := (isCompact_range f.hom.continuous).isClosed let U := Cᶜ have hyU : y ∈ U := by refine Set.mem_compl ?_ @@ -232,22 +232,20 @@ theorem epi_iff_surjective {X Y : Profinite.{u}} (f : X ⟶ Y) : Epi f ↔ Funct obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_isClopen.mem_nhds_iff.mp hUy classical let Z := of (ULift.{u} <| Fin 2) - let g : Y ⟶ Z := ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ - let h : Y ⟶ Z := ⟨fun _ => ⟨1⟩, continuous_const⟩ + let g : Y ⟶ Z := ofHom _ + ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ + let h : Y ⟶ Z := ofHom _ ⟨fun _ => ⟨1⟩, continuous_const⟩ have H : h = g := by rw [← cancel_epi f] ext x - apply ULift.ext dsimp [g, LocallyConstant.ofIsClopen] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [CategoryTheory.comp_apply, ContinuousMap.coe_mk, CategoryTheory.comp_apply, + rw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, ConcreteCategory.hom_ofHom, ContinuousMap.coe_mk, Function.comp_apply, if_neg] refine mt (fun α => hVU α) ?_ - simp only [U, C, Set.mem_range_self, not_true, not_false_iff, Set.mem_compl_iff] + simp [U, C] apply_fun fun e => (e y).down at H dsimp [g, LocallyConstant.ofIsClopen] at H - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, Function.comp_apply, if_pos hyV] at H + rw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, Function.comp_apply, if_pos hyV] at H exact top_ne_bot H · rw [← CategoryTheory.epi_iff_surjective] apply (forget Profinite).epi_of_epi_map diff --git a/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean b/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean index 35974372c30e3f..2aea64f2cf8325 100644 --- a/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean +++ b/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean @@ -26,9 +26,6 @@ namespace Profinite open CategoryTheory Limits --- This was a global instance prior to https://github.com/leanprover-community/mathlib4/pull/13170. We may experiment with removing it. -attribute [local instance] HasForget.instFunLike - universe u v variable {J : Type v} [SmallCategory J] [IsCofiltered J] {F : J ⥤ Profinite.{max u v}} (C : Cone F) @@ -48,8 +45,8 @@ theorem exists_isClopen_of_cofiltered {U : Set C.pt} (hC : IsLimit C) (hU : IsCl change TopologicalSpace.IsTopologicalBasis {W : Set (F.obj i) | IsClopen W} apply isTopologicalBasis_isClopen · rintro i j f V (hV : IsClopen _) - exact ⟨hV.1.preimage ((F ⋙ toTopCat).map f).continuous, - hV.2.preimage ((F ⋙ toTopCat).map f).continuous⟩ + exact ⟨hV.1.preimage ((F ⋙ toTopCat).map f).hom.continuous, + hV.2.preimage ((F ⋙ toTopCat).map f).hom.continuous⟩ -- Porting note: `<;> continuity` fails -- Using this, since `U` is open, we can write `U` as a union of clopen sets all of which -- are preimages of clopens from the factors in the limit. @@ -64,13 +61,13 @@ theorem exists_isClopen_of_cofiltered {U : Set C.pt} (hC : IsLimit C) (hU : IsCl -- clopens constructed in the previous step. have hUo : ∀ (i : ↑S), IsOpen ((fun s ↦ (forget Profinite).map (C.π.app (j s)) ⁻¹' V s) i) := by intro s - exact (hV s).1.2.preimage (C.π.app (j s)).continuous + exact (hV s).1.2.preimage (C.π.app (j s)).hom.continuous have hsU : U ⊆ ⋃ (i : ↑S), (fun s ↦ (forget Profinite).map (C.π.app (j s)) ⁻¹' V s) i := by dsimp only rw [h] rintro x ⟨T, hT, hx⟩ refine ⟨_, ⟨⟨T, hT⟩, rfl⟩, ?_⟩ - dsimp only [forget_map_eq_coe] + dsimp only [ConcreteCategory.forget_map_eq_coe] rwa [← (hV ⟨T, hT⟩).2] have := hU.1.isCompact.elim_finite_subcover (fun s : S => C.π.app (j s) ⁻¹' V s) hUo hsU -- Porting note: same remark as after `hB` @@ -90,24 +87,25 @@ theorem exists_isClopen_of_cofiltered {U : Set C.pt} (hC : IsLimit C) (hU : IsCl intro s hs dsimp [W] rw [dif_pos hs] - exact ⟨(hV s).1.1.preimage (F.map _).continuous, (hV s).1.2.preimage (F.map _).continuous⟩ + exact ⟨(hV s).1.1.preimage (F.map _).hom.continuous, + (hV s).1.2.preimage (F.map _).hom.continuous⟩ · ext x constructor · intro hx simp_rw [W, Set.preimage_iUnion, Set.mem_iUnion] obtain ⟨_, ⟨s, rfl⟩, _, ⟨hs, rfl⟩, hh⟩ := hG hx refine ⟨s, hs, ?_⟩ - rwa [dif_pos hs, ← Set.preimage_comp, ← CompHausLike.coe_comp, ← Functor.map_comp, C.w] + rwa [dif_pos hs, ← Set.preimage_comp, ← CompHausLike.coe_comp, C.w] · intro hx simp_rw [W, Set.preimage_iUnion, Set.mem_iUnion] at hx obtain ⟨s, hs, hx⟩ := hx rw [h] refine ⟨s.1, s.2, ?_⟩ rw [(hV s).2] - rwa [dif_pos hs, ← Set.preimage_comp, ← CompHausLike.coe_comp, ← Functor.map_comp, C.w] at hx + rwa [dif_pos hs, ← Set.preimage_comp, ← CompHausLike.coe_comp, C.w] at hx theorem exists_locallyConstant_fin_two (hC : IsLimit C) (f : LocallyConstant C.pt (Fin 2)) : - ∃ (j : J) (g : LocallyConstant (F.obj j) (Fin 2)), f = g.comap (C.π.app _) := by + ∃ (j : J) (g : LocallyConstant (F.obj j) (Fin 2)), f = g.comap (C.π.app _).hom := by let U := f ⁻¹' {0} have hU : IsClopen U := f.isLocallyConstant.isClopen_fiber _ obtain ⟨j, V, hV, h⟩ := exists_isClopen_of_cofiltered C hC hU @@ -122,7 +120,7 @@ theorem exists_locallyConstant_fin_two (hC : IsLimit C) (f : LocallyConstant C.p open Classical in theorem exists_locallyConstant_finite_aux {α : Type*} [Finite α] (hC : IsLimit C) (f : LocallyConstant C.pt α) : ∃ (j : J) (g : LocallyConstant (F.obj j) (α → Fin 2)), - (f.map fun a b => if a = b then (0 : Fin 2) else 1) = g.comap (C.π.app _) := by + (f.map fun a b => if a = b then (0 : Fin 2) else 1) = g.comap (C.π.app _).hom := by cases nonempty_fintype α let ι : α → α → Fin 2 := fun x y => if x = y then 0 else 1 let ff := (f.map ι).flip @@ -134,14 +132,14 @@ theorem exists_locallyConstant_finite_aux {α : Type*} [Finite α] (hC : IsLimit intro a simp only [Finset.mem_image, Finset.mem_univ, true_and, exists_apply_eq_apply] let fs : ∀ a : α, j0 ⟶ j a := fun a => (hj0 (hj a)).some - let gg : α → LocallyConstant (F.obj j0) (Fin 2) := fun a => (g a).comap (F.map (fs _)) + let gg : α → LocallyConstant (F.obj j0) (Fin 2) := fun a => (g a).comap (F.map (fs _)).hom let ggg := LocallyConstant.unflip gg refine ⟨j0, ggg, ?_⟩ have : f.map ι = LocallyConstant.unflip (f.map ι).flip := by simp rw [this]; clear this have : - LocallyConstant.comap (C.π.app j0) ggg = - LocallyConstant.unflip (LocallyConstant.comap (C.π.app j0) ggg).flip := by + LocallyConstant.comap (C.π.app j0).hom ggg = + LocallyConstant.unflip (LocallyConstant.comap (C.π.app j0).hom ggg).flip := by simp rw [this]; clear this congr 1 @@ -155,7 +153,7 @@ theorem exists_locallyConstant_finite_aux {α : Type*} [Finite α] (hC : IsLimit theorem exists_locallyConstant_finite_nonempty {α : Type*} [Finite α] [Nonempty α] (hC : IsLimit C) (f : LocallyConstant C.pt α) : - ∃ (j : J) (g : LocallyConstant (F.obj j) α), f = g.comap (C.π.app _) := by + ∃ (j : J) (g : LocallyConstant (F.obj j) α), f = g.comap (C.π.app _).hom := by inhabit α obtain ⟨j, gg, h⟩ := exists_locallyConstant_finite_aux _ hC f classical @@ -187,7 +185,7 @@ theorem exists_locallyConstant_finite_nonempty {α : Type*} [Finite α] [Nonempt /-- Any locally constant function from a cofiltered limit of profinite sets factors through one of the components. -/ theorem exists_locallyConstant {α : Type*} (hC : IsLimit C) (f : LocallyConstant C.pt α) : - ∃ (j : J) (g : LocallyConstant (F.obj j) α), f = g.comap (C.π.app _) := by + ∃ (j : J) (g : LocallyConstant (F.obj j) α), f = g.comap (C.π.app _).hom := by let S := f.discreteQuotient let ff : S → α := f.lift cases isEmpty_or_nonempty S diff --git a/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean b/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean index ecdeb71c73112f..c96bf1481284cf 100644 --- a/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/Profinite/EffectiveEpi.lean @@ -25,8 +25,6 @@ universe u open CategoryTheory Limits -attribute [local instance] HasForget.instFunLike - namespace Profinite open List in diff --git a/Mathlib/Topology/Category/Profinite/Extend.lean b/Mathlib/Topology/Category/Profinite/Extend.lean index cfb6aa9416fa0e..1f88e04d2474f2 100644 --- a/Mathlib/Topology/Category/Profinite/Extend.lean +++ b/Mathlib/Topology/Category/Profinite/Extend.lean @@ -27,8 +27,6 @@ universe u w open CategoryTheory Limits FintypeCat Functor -attribute [local instance] HasForget.instFunLike - namespace Profinite variable {I : Type u} [SmallCategory I] [IsCofiltered I] @@ -43,7 +41,7 @@ lemma exists_hom (hc : IsLimit c) {X : FintypeCat} (f : c.pt ⟶ toProfinite.obj let _ : TopologicalSpace X := ⊥ have : DiscreteTopology (toProfinite.obj X) := ⟨rfl⟩ let f' : LocallyConstant c.pt (toProfinite.obj X) := - ⟨f, (IsLocallyConstant.iff_continuous _).mpr f.continuous⟩ + ⟨f, (IsLocallyConstant.iff_continuous _).mpr f.hom.continuous⟩ obtain ⟨i, g, h⟩ := exists_locallyConstant.{_, u} c hc f' refine ⟨i, (g : _ → _), ?_⟩ ext x diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index fc08069c2269f0..361f9b2704c8ff 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -209,9 +209,10 @@ def spanFunctor [∀ (s : Finset I) (i : I), Decidable (i ∈ s)] (hC : IsCompac (Finset I)ᵒᵖ ⥤ Profinite.{u} where obj s := @Profinite.of (π C (· ∈ (unop s))) _ (by rw [← isCompact_iff_compactSpace]; exact hC.image (continuous_proj _)) _ _ - map h := ⟨(ProjRestricts C (leOfHom h.unop)), continuous_projRestricts _ _⟩ + map h := @CompHausLike.ofHom _ _ _ (_) (_) (_) (_) (_) (_) (_) (_) + ⟨(ProjRestricts C (leOfHom h.unop)), continuous_projRestricts _ _⟩ map_id J := by simp only [projRestricts_eq_id C (· ∈ (unop J))]; rfl - map_comp _ _ := by dsimp; congr; dsimp; rw [projRestricts_eq_comp] + map_comp _ _ := by dsimp; rw [← CompHausLike.ofHom_comp]; congr; dsimp; rw [projRestricts_eq_comp] /-- The limit cone on `spanFunctor` with point `C`. -/ noncomputable @@ -219,7 +220,7 @@ def spanCone [∀ (s : Finset I) (i : I), Decidable (i ∈ s)] (hC : IsCompact C Cone (spanFunctor hC) where pt := @Profinite.of C _ (by rwa [← isCompact_iff_compactSpace]) _ _ π := - { app := fun s ↦ ⟨ProjRestrict C (· ∈ unop s), continuous_projRestrict _ _⟩ + { app := fun s ↦ TopCat.ofHom ⟨ProjRestrict C (· ∈ unop s), continuous_projRestrict _ _⟩ naturality := by intro X Y h simp only [Functor.const_obj_obj, Homeomorph.setCongr, Homeomorph.homeomorph_mk_coe, diff --git a/Mathlib/Topology/Category/Profinite/Product.lean b/Mathlib/Topology/Category/Profinite/Product.lean index 0224264e33b984..f46ed3c1884318 100644 --- a/Mathlib/Topology/Category/Profinite/Product.lean +++ b/Mathlib/Topology/Category/Profinite/Product.lean @@ -86,13 +86,13 @@ noncomputable def indexFunctor (hC : IsCompact C) : (Finset ι)ᵒᵖ ⥤ Profinite.{u} where obj J := @Profinite.of (obj C (· ∈ (unop J))) _ (by rw [← isCompact_iff_compactSpace]; exact hC.image (Pi.continuous_precomp' _)) _ _ - map h := map C (leOfHom h.unop) + map h := TopCat.ofHom (map C (leOfHom h.unop)) /-- The limit cone on `indexFunctor` -/ noncomputable def indexCone (hC : IsCompact C) : Cone (indexFunctor hC) where pt := @Profinite.of C _ (by rwa [← isCompact_iff_compactSpace]) _ _ - π := { app := fun J ↦ π_app C (· ∈ unop J) } + π := { app := fun J ↦ TopCat.ofHom (π_app C (· ∈ unop J)) } variable (hC : IsCompact C) diff --git a/Mathlib/Topology/Category/Profinite/Projective.lean b/Mathlib/Topology/Category/Profinite/Projective.lean index 4a53a7e2751871..c15935e8aa6727 100644 --- a/Mathlib/Topology/Category/Profinite/Projective.lean +++ b/Mathlib/Topology/Category/Profinite/Projective.lean @@ -30,12 +30,8 @@ universe u v w open CategoryTheory Function --- This was a global instance prior to https://github.com/leanprover-community/mathlib4/pull/13170. We may experiment with removing it. -attribute [local instance] HasForget.instFunLike - namespace Profinite - instance projective_ultrafilter (X : Type u) : Projective (of <| Ultrafilter X) where factors {Y Z} f g hg := by rw [epi_iff_surjective] at hg @@ -43,21 +39,18 @@ instance projective_ultrafilter (X : Type u) : Projective (of <| Ultrafilter X) let t : X → Y := g' ∘ f ∘ (pure : X → Ultrafilter X) let h : Ultrafilter X → Y := Ultrafilter.extend t have hh : Continuous h := continuous_ultrafilter_extend _ - use ⟨h, hh⟩ - apply (forget Profinite).map_injective + use CompHausLike.ofHom _ ⟨h, hh⟩ + apply ConcreteCategory.coe_ext simp only [h, ContinuousMap.coe_mk, coe_comp] - convert denseRange_pure.equalizer (g.continuous.comp hh) f.continuous _ - -- Porting note: same fix as in `Topology.Category.CompHaus.Projective` - let g'' : ContinuousMap Y Z := g - have : g'' ∘ g' = id := hg'.comp_eq_id - -- This used to be `rw`, but we need `rw; rfl` after https://github.com/leanprover/lean4/pull/2644 + convert denseRange_pure.equalizer (g.hom.continuous.comp hh) f.hom.continuous _ + have : g.hom ∘ g' = id := hg'.comp_eq_id rw [comp_assoc, ultrafilter_extend_extends, ← comp_assoc, this, id_comp] rfl /-- For any profinite `X`, the natural map `Ultrafilter X → X` is a projective presentation. -/ def projectivePresentation (X : Profinite.{u}) : ProjectivePresentation X where p := of <| Ultrafilter X - f := ⟨_, continuous_ultrafilter_extend id⟩ + f := CompHausLike.ofHom _ ⟨_, continuous_ultrafilter_extend id⟩ projective := Profinite.projective_ultrafilter X epi := ConcreteCategory.epi_of_surjective _ fun x => ⟨(pure x : Ultrafilter X), congr_fun (ultrafilter_extend_extends (𝟙 X)) x⟩ diff --git a/Mathlib/Topology/Category/Sequential.lean b/Mathlib/Topology/Category/Sequential.lean index 0d0d338a0a45da..c9cacc180776b1 100644 --- a/Mathlib/Topology/Category/Sequential.lean +++ b/Mathlib/Topology/Category/Sequential.lean @@ -18,21 +18,19 @@ for defining categories of topological spaces, by giving it the induced category open CategoryTheory -attribute [local instance] HasForget.instFunLike - universe u /-- The type sequential topological spaces. -/ structure Sequential where /-- The underlying topological space of an object of `Sequential`. -/ - toTop : TopCat.{u} + toTop : TopCat.{u} -- TODO: turn this into `extends` /-- The underlying topological space is sequential. -/ [is_sequential : SequentialSpace toTop] namespace Sequential instance : Inhabited Sequential.{u} := - ⟨{ toTop := { α := ULift (Fin 37) } }⟩ + ⟨{ toTop := TopCat.of (ULift (Fin 37)) }⟩ instance : CoeSort Sequential Type* := ⟨fun X => X.toTop⟩ @@ -42,13 +40,13 @@ attribute [instance] is_sequential instance : Category.{u, u+1} Sequential.{u} := InducedCategory.category toTop -instance : HasForget.{u} Sequential.{u} := - InducedCategory.hasForget _ +instance : ConcreteCategory.{u} Sequential.{u} (C(·, ·)) := + InducedCategory.concreteCategory toTop variable (X : Type u) [TopologicalSpace X] [SequentialSpace X] /-- Constructor for objects of the category `Sequential`. -/ -def of : Sequential.{u} where +abbrev of : Sequential.{u} where toTop := TopCat.of X is_sequential := ‹_› @@ -70,8 +68,8 @@ instance : sequentialToTop.{u}.Faithful := /-- Construct an isomorphism from a homeomorphism. -/ @[simps hom inv] def isoOfHomeo {X Y : Sequential.{u}} (f : X ≃ₜ Y) : X ≅ Y where - hom := ⟨f, f.continuous⟩ - inv := ⟨f.symm, f.symm.continuous⟩ + hom := TopCat.ofHom ⟨f, f.continuous⟩ + inv := TopCat.ofHom ⟨f.symm, f.symm.continuous⟩ hom_inv_id := by ext x exact f.symm_apply_apply x @@ -84,10 +82,10 @@ def isoOfHomeo {X Y : Sequential.{u}} (f : X ≃ₜ Y) : X ≅ Y where def homeoOfIso {X Y : Sequential.{u}} (f : X ≅ Y) : X ≃ₜ Y where toFun := f.hom invFun := f.inv - left_inv x := by simp - right_inv x := by simp - continuous_toFun := f.hom.continuous - continuous_invFun := f.inv.continuous + left_inv := f.hom_inv_id_apply + right_inv := f.inv_hom_id_apply + continuous_toFun := f.hom.hom.continuous + continuous_invFun := f.inv.hom.continuous /-- The equivalence between isomorphisms in `Sequential` and homeomorphisms of topological spaces. -/ diff --git a/Mathlib/Topology/Category/Stonean/Basic.lean b/Mathlib/Topology/Category/Stonean/Basic.lean index 421a0327cf878b..bf7f0aa66a44b3 100644 --- a/Mathlib/Topology/Category/Stonean/Basic.lean +++ b/Mathlib/Topology/Category/Stonean/Basic.lean @@ -42,9 +42,6 @@ universe u open CategoryTheory open scoped Topology --- This was a global instance prior to https://github.com/leanprover-community/mathlib4/pull/13170. We may experiment with removing it. -attribute [local instance] HasForget.instFunLike - /-- `Stonean` is the category of extremally disconnected compact Hausdorff spaces. -/ abbrev Stonean := CompHausLike (fun X ↦ ExtremallyDisconnected X) @@ -56,13 +53,13 @@ instance (X : CompHaus.{u}) [Projective X] : ExtremallyDisconnected X := by intro A B _ _ _ _ _ _ f g hf hg hsurj let A' : CompHaus := CompHaus.of A let B' : CompHaus := CompHaus.of B - let f' : X ⟶ B' := ⟨f, hf⟩ - let g' : A' ⟶ B' := ⟨g,hg⟩ + let f' : X ⟶ B' := CompHausLike.ofHom _ ⟨f, hf⟩ + let g' : A' ⟶ B' := CompHausLike.ofHom _ ⟨g,hg⟩ have : Epi g' := by rw [CompHaus.epi_iff_surjective] assumption obtain ⟨h, hh⟩ := Projective.factors f' g' - refine ⟨h, h.2, ?_⟩ + refine ⟨h, h.hom.2, ?_⟩ ext t apply_fun (fun e => e t) at hh exact hh @@ -125,20 +122,20 @@ A morphism in `Stonean` is an epi iff it is surjective. -/ lemma epi_iff_surjective {X Y : Stonean} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := by - refine ⟨?_, ConcreteCategory.epi_of_surjective _⟩ + refine ⟨?_, fun h => ConcreteCategory.epi_of_surjective f h⟩ dsimp [Function.Surjective] intro h y by_contra! hy let C := Set.range f - have hC : IsClosed C := (isCompact_range f.continuous).isClosed + have hC : IsClosed C := (isCompact_range f.hom.continuous).isClosed let U := Cᶜ have hUy : U ∈ 𝓝 y := by simp only [U, C, Set.mem_range, hy, exists_false, not_false_eq_true, hC.compl_mem_nhds] obtain ⟨V, hV, hyV, hVU⟩ := isTopologicalBasis_isClopen.mem_nhds_iff.mp hUy classical - let g : Y ⟶ mkFinite (ULift (Fin 2)) := + let g : Y ⟶ mkFinite (ULift (Fin 2)) := TopCat.ofHom ⟨(LocallyConstant.ofIsClopen hV).map ULift.up, LocallyConstant.continuous _⟩ - let h : Y ⟶ mkFinite (ULift (Fin 2)) := ⟨fun _ => ⟨1⟩, continuous_const⟩ + let h : Y ⟶ mkFinite (ULift (Fin 2)) := TopCat.ofHom ⟨fun _ => ⟨1⟩, continuous_const⟩ have H : h = g := by rw [← cancel_epi f] ext x @@ -159,8 +156,9 @@ instance instProjectiveCompHausCompHaus (X : Stonean) : Projective (toCompHaus.o intro B C φ f _ haveI : ExtremallyDisconnected (toCompHaus.obj X).toTop := X.prop have hf : Function.Surjective f := by rwa [← CompHaus.epi_iff_surjective] - obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.continuous f.continuous hf - use ⟨f', h.left⟩ + obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.hom.continuous f.hom.continuous + hf + use ofHom _ ⟨f', h.left⟩ ext exact congr_fun h.right _ @@ -170,8 +168,9 @@ instance (X : Stonean) : Projective (toProfinite.obj X) where intro B C φ f _ haveI : ExtremallyDisconnected (toProfinite.obj X) := X.prop have hf : Function.Surjective f := by rwa [← Profinite.epi_iff_surjective] - obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.continuous f.continuous hf - use ⟨f', h.left⟩ + obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.hom.continuous f.hom.continuous + hf + use ofHom _ ⟨f', h.left⟩ ext exact congr_fun h.right _ @@ -181,8 +180,9 @@ instance (X : Stonean) : Projective X where intro B C φ f _ haveI : ExtremallyDisconnected X.toTop := X.prop have hf : Function.Surjective f := by rwa [← Stonean.epi_iff_surjective] - obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.continuous f.continuous hf - use ⟨f', h.left⟩ + obtain ⟨f', h⟩ := CompactT2.ExtremallyDisconnected.projective φ.hom.continuous f.hom.continuous + hf + use ofHom _ ⟨f', h.left⟩ ext exact congr_fun h.right _ @@ -199,10 +199,10 @@ def presentation (X : CompHaus) : Stonean where prop := by refine CompactT2.Projective.extremallyDisconnected (@fun Y Z _ _ _ _ _ _ f g hfcont hgcont hgsurj => ?_) - let g₁ : (CompHaus.of Y) ⟶ (CompHaus.of Z) := ⟨g, hgcont⟩ - let f₁ : (projectivePresentation X).p ⟶ (CompHaus.of Z) := ⟨f, hfcont⟩ + let g₁ : (CompHaus.of Y) ⟶ (CompHaus.of Z) := CompHausLike.ofHom _ ⟨g, hgcont⟩ + let f₁ : (projectivePresentation X).p ⟶ (CompHaus.of Z) := CompHausLike.ofHom _ ⟨f, hfcont⟩ have hg₁ : Epi g₁ := (epi_iff_surjective _).2 hgsurj - refine ⟨Projective.factorThru f₁ g₁, (Projective.factorThru f₁ g₁).2, funext (fun _ => ?_)⟩ + refine ⟨Projective.factorThru f₁ g₁, (Projective.factorThru f₁ g₁).hom.2, funext (fun _ => ?_)⟩ change (Projective.factorThru f₁ g₁ ≫ g₁) _ = f _ rw [Projective.factorThru_comp] rfl diff --git a/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean b/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean index 9d1ef396385b62..5b16f02450c687 100644 --- a/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/Stonean/EffectiveEpi.lean @@ -24,8 +24,6 @@ universe u open CategoryTheory Limits CompHausLike -attribute [local instance] HasForget.instFunLike - namespace Stonean open List in diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index 6e9904e8c16da7..f3b4f280e534cd 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -17,8 +17,6 @@ universe w u open CategoryTheory Limits CompHausLike Topology -attribute [local instance] HasForget.instFunLike - namespace Stonean instance : HasExplicitFiniteCoproducts.{w, u} (fun Y ↦ ExtremallyDisconnected Y) where @@ -30,8 +28,8 @@ include hi lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.range f)) where open_closure U hU := by have h : IsClopen (i ⁻¹' (Set.range f)) := - ⟨IsClosed.preimage i.continuous (isCompact_range f.continuous).isClosed, - IsOpen.preimage i.continuous hi.isOpen_range⟩ + ⟨IsClosed.preimage i.hom.continuous (isCompact_range f.hom.continuous).isClosed, + IsOpen.preimage i.hom.continuous hi.isOpen_range⟩ rw [← (closure U).preimage_image_eq Subtype.coe_injective, ← h.1.isClosedEmbedding_subtypeVal.closure_image_eq U] exact isOpen_induced (ExtremallyDisconnected.open_closure _ @@ -39,7 +37,7 @@ lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.ra lemma extremallyDisconnected_pullback : ExtremallyDisconnected {xy : X × Y | f xy.1 = i xy.2} := have := extremallyDisconnected_preimage i hi - let e := (TopCat.pullbackHomeoPreimage i i.2 f hi.isEmbedding).symm + let e := (TopCat.pullbackHomeoPreimage i i.hom.2 f hi.isEmbedding).symm let e' : {xy : X × Y | f xy.1 = i xy.2} ≃ₜ {xy : Y × X | i xy.1 = f xy.2} := by exact TopCat.homeoOfIso ((TopCat.pullbackIsoProdSubtype f i).symm ≪≫ pullbackSymmetry _ _ ≪≫ diff --git a/Mathlib/Topology/Category/TopCat/Adjunctions.lean b/Mathlib/Topology/Category/TopCat/Adjunctions.lean index 52f1591638bd4a..1ca62757df9a9a 100644 --- a/Mathlib/Topology/Category/TopCat/Adjunctions.lean +++ b/Mathlib/Topology/Category/TopCat/Adjunctions.lean @@ -28,13 +28,13 @@ namespace TopCat @[simps! unit counit] def adj₁ : discrete ⊣ forget TopCat.{u} where unit := { app := fun _ => id } - counit := { app := fun _ => ⟨id, continuous_bot⟩ } + counit := { app := fun X => TopCat.ofHom (X := discrete.obj X) ⟨id, continuous_bot⟩ } /-- Equipping a type with the trivial topology is right adjoint to the forgetful functor `Top ⥤ Type`. -/ @[simps! unit counit] def adj₂ : forget TopCat.{u} ⊣ trivial where - unit := { app := fun _ => ⟨id, continuous_top⟩ } + unit := { app := fun X => TopCat.ofHom (Y := trivial.obj X) ⟨id, continuous_top⟩ } counit := { app := fun _ => id } instance : (forget TopCat.{u}).IsRightAdjoint := diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 31d2fe85ecda8f..42c84fbb1dd72d 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Kim Morrison, Mario Carneiro -/ -import Mathlib.CategoryTheory.ConcreteCategory.BundledHom +import Mathlib.CategoryTheory.ConcreteCategory.Basic import Mathlib.Topology.ContinuousMap.Basic /-! @@ -20,72 +20,123 @@ open CategoryTheory TopologicalSpace Topology universe u -/-- The category of topological spaces and continuous maps. -/ -@[to_additive existing TopCat] -def TopCat : Type (u + 1) := - Bundled TopologicalSpace +/-- The category of semirings. -/ +structure TopCat where + private mk :: + /-- The underlying type. -/ + carrier : Type u + [str : TopologicalSpace carrier] + +attribute [instance] TopCat.str + +initialize_simps_projections TopCat (-str) namespace TopCat -instance bundledHom : BundledHom @ContinuousMap where - toFun := @ContinuousMap.toFun - id := @ContinuousMap.id - comp := @ContinuousMap.comp +instance : CoeSort (TopCat) (Type u) := + ⟨TopCat.carrier⟩ + +attribute [coe] TopCat.carrier + +/-- The object in `TopCat` associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `TopCat`. -/ +abbrev of (X : Type u) [TopologicalSpace X] : TopCat := + ⟨X⟩ + +lemma coe_of (X : Type u) [TopologicalSpace X] : (of X : Type u) = X := + rfl + +lemma of_carrier (X : TopCat.{u}) : of X = X := rfl + +variable {X} in +/-- The type of morphisms in `TopCat`. -/ +@[ext] +structure Hom (X Y : TopCat.{u}) where + private mk :: + /-- The underlying `ContinuousMap`. -/ + hom' : C(X, Y) -deriving instance LargeCategory for TopCat +instance : Category TopCat where + Hom X Y := Hom X Y + id X := ⟨ContinuousMap.id X⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ --- Porting note: currently no derive handler for HasForget --- see https://github.com/leanprover-community/mathlib4/issues/5020 -instance hasForget : HasForget TopCat := - inferInstanceAs <| HasForget (Bundled TopologicalSpace) +instance : ConcreteCategory.{u} TopCat (fun X Y => C(X, Y)) where + hom := Hom.hom' + ofHom f := ⟨f⟩ -instance : CoeSort TopCat Type* where - coe X := X.α +/-- Turn a morphism in `TopCat` back into a `ContinuousMap`. -/ +abbrev Hom.hom {X Y : TopCat.{u}} (f : Hom X Y) := + ConcreteCategory.hom (C := TopCat) f -instance topologicalSpaceUnbundled (X : TopCat) : TopologicalSpace X := - X.str +/-- Typecheck a `ContinuousMap` as a morphism in `TopCat`. -/ +abbrev ofHom {X Y : Type u} [TopologicalSpace X] [TopologicalSpace Y] (f : C(X, Y)) : of X ⟶ of Y := + ConcreteCategory.ofHom (C := TopCat) f -instance instFunLike (X Y : TopCat) : FunLike (X ⟶ Y) X Y := - inferInstanceAs <| FunLike C(X, Y) X Y +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (X Y : TopCat) (f : Hom X Y) := + f.hom -instance instContinuousMapClass (X Y : TopCat) : ContinuousMapClass (X ⟶ Y) X Y := - inferInstanceAs <| ContinuousMapClass C(X, Y) X Y +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ + +@[simp] +lemma hom_id {X : TopCat.{u}} : (𝟙 X : X ⟶ X).hom = ContinuousMap.id X := rfl @[simp] theorem id_app (X : TopCat.{u}) (x : ↑X) : (𝟙 X : X ⟶ X) x = x := rfl +@[simp] theorem coe_id (X : TopCat.{u}) : (𝟙 X : X → X) = id := rfl + +@[simp] +lemma hom_comp {X Y Z : TopCat.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + @[simp] theorem comp_app {X Y Z : TopCat.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g : X → Z) x = g (f x) := rfl -@[simp] theorem coe_id (X : TopCat.{u}) : (𝟙 X : X → X) = id := rfl - @[simp] theorem coe_comp {X Y Z : TopCat.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g : X → Z) = g ∘ f := rfl +@[ext] +lemma hom_ext {X Y : TopCat} {f g : X ⟶ Y} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +@[ext] +lemma ext {X Y : TopCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := + ConcreteCategory.hom_ext _ _ w + @[simp] -lemma hom_inv_id_apply {X Y : TopCat} (f : X ≅ Y) (x : X) : f.inv (f.hom x) = x := - DFunLike.congr_fun f.hom_inv_id x +lemma hom_ofHom {X Y : Type u} [TopologicalSpace X] [TopologicalSpace Y] (f : C(X, Y)) : + (ofHom f).hom = f := rfl @[simp] -lemma inv_hom_id_apply {X Y : TopCat} (f : X ≅ Y) (y : Y) : f.hom (f.inv y) = y := - DFunLike.congr_fun f.inv_hom_id y +lemma ofHom_hom {X Y : TopCat} (f : X ⟶ Y) : + ofHom (Hom.hom f) = f := rfl -/-- Construct a bundled `Top` from the underlying type and the typeclass. -/ -def of (X : Type u) [TopologicalSpace X] : TopCat := - -- Porting note: needed to call inferInstance - ⟨X, inferInstance⟩ +@[simp] +lemma ofHom_id {X : Type u} [TopologicalSpace X] : ofHom (ContinuousMap.id X) = 𝟙 (of X) := rfl -instance topologicalSpace_coe (X : TopCat) : TopologicalSpace X := - X.str +@[simp] +lemma ofHom_comp {X Y Z : Type u} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] + (f : C(X, Y)) (g : C(Y, Z)) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := + rfl --- Porting note: cannot see through forget; made reducible to get closer to Lean 3 behavior -@[instance] abbrev topologicalSpace_forget - (X : TopCat) : TopologicalSpace <| (forget TopCat).obj X := - X.str +lemma ofHom_apply {X Y : Type u} [TopologicalSpace X] [TopologicalSpace Y] (f : C(X, Y)) (x : X) : + (ofHom f) x = f x := rfl @[simp] -theorem coe_of (X : Type u) [TopologicalSpace X] : (of X : Type u) = X := rfl +lemma hom_inv_id_apply {X Y : TopCat} (f : X ≅ Y) (x : X) : f.inv (f.hom x) = x := + CategoryTheory.congr_fun f.hom_inv_id x + +@[simp] +lemma inv_hom_id_apply {X Y : TopCat} (f : X ≅ Y) (y : Y) : f.hom (f.inv y) = y := + CategoryTheory.congr_fun f.inv_hom_id y /-- Replace a function coercion for a morphism `TopCat.of X ⟶ TopCat.of Y` with the definitionally @@ -95,7 +146,7 @@ equal function coercion for a continuous map `C(X, Y)`. {f : C(X, Y)} {x} : @DFunLike.coe (TopCat.of X ⟶ TopCat.of Y) ((CategoryTheory.forget TopCat).obj (TopCat.of X)) (fun _ ↦ (CategoryTheory.forget TopCat).obj (TopCat.of Y)) HasForget.instFunLike - f x = + (ofHom f) x = @DFunLike.coe C(X, Y) X (fun _ ↦ Y) _ f x := @@ -104,30 +155,29 @@ equal function coercion for a continuous map `C(X, Y)`. instance inhabited : Inhabited TopCat := ⟨TopCat.of Empty⟩ --- Porting note: added to ease the port of `AlgebraicTopology.TopologicalSimplex` -lemma hom_apply {X Y : TopCat} (f : X ⟶ Y) (x : X) : f x = ContinuousMap.toFun f x := rfl +@[deprecated + "Simply remove this from the `simp`/`rw` set: the LHS and RHS are now identical." + (since := "2025-01-30")] +lemma hom_apply {X Y : TopCat} (f : X ⟶ Y) (x : X) : f x = ContinuousMap.toFun f.hom x := rfl /-- The discrete topology on any type. -/ def discrete : Type u ⥤ TopCat.{u} where - obj X := ⟨X , ⊥⟩ - map f := @ContinuousMap.mk _ _ ⊥ ⊥ f continuous_bot + obj X := @of X ⊥ + map f := @ofHom _ _ ⊥ ⊥ <| @ContinuousMap.mk _ _ ⊥ ⊥ f continuous_bot instance {X : Type u} : DiscreteTopology (discrete.obj X) := ⟨rfl⟩ /-- The trivial topology on any type. -/ def trivial : Type u ⥤ TopCat.{u} where - obj X := ⟨X, ⊤⟩ - map f := @ContinuousMap.mk _ _ ⊤ ⊤ f continuous_top + obj X := @of X ⊤ + map f := @ofHom _ _ ⊤ ⊤ <| @ContinuousMap.mk _ _ ⊤ ⊤ f continuous_top /-- Any homeomorphisms induces an isomorphism in `Top`. -/ @[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 : 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 _ + hom := ofHom f + inv := ofHom f.symm /-- Any isomorphism in `Top` induces a homeomorphism. -/ @[simps] @@ -136,33 +186,29 @@ def homeoOfIso {X Y : TopCat.{u}} (f : X ≅ Y) : X ≃ₜ Y where invFun := f.inv left_inv x := by simp right_inv x := by simp - continuous_toFun := f.hom.continuous - continuous_invFun := f.inv.continuous + continuous_toFun := f.hom.hom.continuous + continuous_invFun := f.inv.hom.continuous @[simp] theorem of_isoOfHomeo {X Y : TopCat.{u}} (f : X ≃ₜ Y) : homeoOfIso (isoOfHomeo f) = f := by - -- Porting note: unfold some defs now - dsimp [homeoOfIso, isoOfHomeo] ext rfl @[simp] theorem of_homeoOfIso {X Y : TopCat.{u}} (f : X ≅ Y) : isoOfHomeo (homeoOfIso f) = f := by - -- Porting note: unfold some defs now - dsimp [homeoOfIso, isoOfHomeo] ext rfl lemma isIso_of_bijective_of_isOpenMap {X Y : TopCat.{u}} (f : X ⟶ Y) (hfbij : Function.Bijective f) (hfcl : IsOpenMap f) : IsIso f := let e : X ≃ₜ Y := Homeomorph.homeomorphOfContinuousOpen - (Equiv.ofBijective f hfbij) f.continuous hfcl + (Equiv.ofBijective f hfbij) f.hom.continuous hfcl inferInstanceAs <| IsIso (TopCat.isoOfHomeo e).hom lemma isIso_of_bijective_of_isClosedMap {X Y : TopCat.{u}} (f : X ⟶ Y) (hfbij : Function.Bijective f) (hfcl : IsClosedMap f) : IsIso f := let e : X ≃ₜ Y := Homeomorph.homeomorphOfContinuousClosed - (Equiv.ofBijective f hfbij) f.continuous hfcl + (Equiv.ofBijective f hfbij) f.hom.continuous hfcl inferInstanceAs <| IsIso (TopCat.isoOfHomeo e).hom theorem isOpenEmbedding_iff_comp_isIso {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : @@ -174,7 +220,7 @@ alias openEmbedding_iff_comp_isIso := isOpenEmbedding_iff_comp_isIso @[simp] 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 + IsOpenEmbedding (g ∘ f) ↔ IsOpenEmbedding f := by simp only [← Functor.map_comp] exact isOpenEmbedding_iff_comp_isIso f g @@ -186,7 +232,7 @@ theorem isOpenEmbedding_iff_isIso_comp {X Y Z : TopCat} (f : X ⟶ Y) (g : Y ⟶ constructor · intro h convert h.comp (TopCat.homeoOfIso (asIso f).symm).isOpenEmbedding - exact congrArg _ (IsIso.inv_hom_id_assoc f g).symm + exact congr_arg (DFunLike.coe ∘ ConcreteCategory.hom) (IsIso.inv_hom_id_assoc f g).symm · exact fun h => h.comp (TopCat.homeoOfIso (asIso f)).isOpenEmbedding @[deprecated (since := "2024-10-18")] @@ -194,7 +240,7 @@ alias openEmbedding_iff_isIso_comp := isOpenEmbedding_iff_isIso_comp @[simp] 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 + IsOpenEmbedding (g ∘ f) ↔ IsOpenEmbedding g := by simp only [← Functor.map_comp] exact isOpenEmbedding_iff_isIso_comp f g diff --git a/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean b/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean index a42db3835f71f8..057fcab0da98ce 100644 --- a/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean +++ b/Mathlib/Topology/Category/TopCat/EffectiveEpi.lean @@ -29,21 +29,28 @@ noncomputable def effectiveEpiStructOfQuotientMap {B X : TopCat.{u}} (π : X ⟶ B) (hπ : IsQuotientMap π) : EffectiveEpiStruct π where /- `IsQuotientMap.lift` gives the required morphism -/ - desc e h := hπ.lift e fun a b hab ↦ - DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ + desc e h := ofHom <| hπ.lift e.hom fun a b hab ↦ + CategoryTheory.congr_fun (h + (ofHom ⟨fun _ ↦ a, continuous_const⟩) + (ofHom ⟨fun _ ↦ b, continuous_const⟩) (by ext; exact hab)) a /- `IsQuotientMap.lift_comp` gives the factorisation -/ - fac e h := (hπ.lift_comp e - fun a b hab ↦ DFunLike.congr_fun (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ + fac e h := hom_ext (hπ.lift_comp e.hom + fun a b hab ↦ CategoryTheory.congr_fun (h + (ofHom ⟨fun _ ↦ a, continuous_const⟩) + (ofHom ⟨fun _ ↦ b, continuous_const⟩) (by ext; exact hab)) a) /- Uniqueness follows from the fact that `IsQuotientMap.lift` is an equivalence (given by `IsQuotientMap.liftEquiv`). -/ uniq e h g hm := by - suffices g = hπ.liftEquiv ⟨e, - fun a b hab ↦ DFunLike.congr_fun - (h ⟨fun _ ↦ a, continuous_const⟩ ⟨fun _ ↦ b, continuous_const⟩ (by ext; exact hab)) - a⟩ by assumption - rw [← Equiv.symm_apply_eq hπ.liftEquiv] + suffices g = ofHom (hπ.liftEquiv ⟨e.hom, + fun a b hab ↦ CategoryTheory.congr_fun (h + (ofHom ⟨fun _ ↦ a, continuous_const⟩) + (ofHom ⟨fun _ ↦ b, continuous_const⟩) + (by ext; exact hab)) + a⟩) by assumption + apply hom_ext + rw [hom_ofHom, ← Equiv.symm_apply_eq hπ.liftEquiv] ext simp only [IsQuotientMap.liftEquiv_symm_apply_coe, ContinuousMap.comp_apply, ← hm] rfl diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index 99c6519501cba9..445b8f3488846d 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -34,16 +34,12 @@ Generally you should just use `limit.cone F`, unless you need the actual definit def limitCone (F : J ⥤ TopCat.{max v u}) : Cone F where pt := TopCat.of { u : ∀ j : J, F.obj j | ∀ {i j : J} (f : i ⟶ j), F.map f (u i) = u j } π := - { app := fun j => + { app := fun j => ofHom { toFun := fun u => u.val j -- Porting note: `continuity` from the original mathlib3 proof failed here. continuous_toFun := Continuous.comp (continuous_apply _) (continuous_subtype_val) } naturality := fun X Y f => by - -- Automation fails in various ways in this proof. Why?! - dsimp - rw [Category.id_comp] - apply ContinuousMap.ext - intro a + ext a exact (a.2 f).symm } /-- A choice of limit cone for a functor `F : J ⥤ TopCat` whose topology is defined as an @@ -52,35 +48,35 @@ Generally you should just use `limit.cone F`, unless you need the actual definit (which is in terms of `Types.limitCone`). -/ def limitConeInfi (F : J ⥤ TopCat.{max v u}) : Cone F where - pt := - ⟨(Types.limitCone.{v,u} (F ⋙ forget)).pt, - ⨅ j, (F.obj j).str.induced ((Types.limitCone.{v,u} (F ⋙ forget)).π.app j)⟩ + pt := @of + ((Types.limitCone.{v,u} (F ⋙ forget)).pt) + (⨅ j, (F.obj j).str.induced ((Types.limitCone.{v,u} (F ⋙ forget)).π.app j)) π := - { app := fun j => - ⟨(Types.limitCone.{v,u} (F ⋙ forget)).π.app j, continuous_iff_le_induced.mpr (iInf_le _ _)⟩ + { app := fun j => @ofHom _ _ (_) (_) <| @ContinuousMap.mk _ _ (_) (_) + ((Types.limitCone.{v,u} (F ⋙ forget)).π.app j) + (continuous_iff_le_induced.mpr (iInf_le _ _)) naturality := fun _ _ f => - ContinuousMap.coe_injective ((Types.limitCone.{v,u} (F ⋙ forget)).π.naturality f) } + ConcreteCategory.coe_ext ((Types.limitCone.{v,u} (F ⋙ forget)).π.naturality f) } /-- The chosen cone `TopCat.limitCone F` for a functor `F : J ⥤ TopCat` is a limit cone. Generally you should just use `limit.isLimit F`, unless you need the actual definition (which is in terms of `Types.limitConeIsLimit`). -/ def limitConeIsLimit (F : J ⥤ TopCat.{max v u}) : IsLimit (limitCone.{v,u} F) where - lift S := + lift S := ofHom { toFun := fun x => ⟨fun _ => S.π.app _ x, fun f => by dsimp rw [← S.w f] rfl⟩ continuous_toFun := - Continuous.subtype_mk (continuous_pi fun j => (S.π.app j).2) fun x i j f => by + Continuous.subtype_mk (continuous_pi fun j => (S.π.app j).hom.2) fun x i j f => by dsimp rw [← S.w f] rfl } uniq S m h := by - apply ContinuousMap.ext; intros a; apply Subtype.ext; funext j - dsimp - rw [← h] + ext a + simp [← h] rfl /-- The chosen cone `TopCat.limitConeInfi F` for a functor `F : J ⥤ TopCat` is a limit cone. @@ -90,17 +86,18 @@ Generally you should just use `limit.isLimit F`, unless you need the actual defi def limitConeInfiIsLimit (F : J ⥤ TopCat.{max v u}) : IsLimit (limitConeInfi.{v,u} F) := by refine IsLimit.ofFaithful forget (Types.limitConeIsLimit.{v,u} (F ⋙ forget)) -- Porting note: previously could infer all ?_ except continuity - (fun s => ⟨fun v => ⟨fun j => (Functor.mapCone forget s).π.app j v, ?_⟩, ?_⟩) fun s => ?_ + (fun s => @ofHom _ _ (_) (_) (@ContinuousMap.mk _ _ (_) (_) + (fun v => ⟨fun j => (Functor.mapCone forget s).π.app j v, ?_⟩) + ?_)) fun s => ?_ · dsimp [Functor.sections] intro _ _ _ - rw [← comp_apply', forget_map_eq_coe, ← s.π.naturality, forget_map_eq_coe] + rw [← ConcreteCategory.comp_apply, ← s.π.naturality] dsimp - rw [Category.id_comp] · exact continuous_iff_coinduced_le.mpr (le_iInf fun j => coinduced_le_iff_le_induced.mp <| - (continuous_iff_coinduced_le.mp (s.π.app j).continuous :)) + (continuous_iff_coinduced_le.mp (s.π.app j).hom.continuous :)) · rfl instance topCat_hasLimitsOfSize : HasLimitsOfSize.{w, v} TopCat.{max v u} where @@ -128,17 +125,18 @@ Generally you should just use `colimit.cocone F`, unless you need the actual def (which is in terms of `Types.colimitCocone`). -/ def colimitCocone (F : J ⥤ TopCat.{max v u}) : Cocone F where - pt := - ⟨(Types.TypeMax.colimitCocone.{v,u} (F ⋙ forget)).pt, - ⨆ j, (F.obj j).str.coinduced ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.app j)⟩ + pt := @of + ((Types.TypeMax.colimitCocone.{v,u} (F ⋙ forget)).pt) + (⨆ j, (F.obj j).str.coinduced ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.app j)) ι := - { app := fun j => - ⟨(Types.TypeMax.colimitCocone (F ⋙ forget)).ι.app j, continuous_iff_coinduced_le.mpr <| + { app := fun j => @ofHom _ _ (_) (_) <| @ContinuousMap.mk _ _ (_) (_) + ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.app j) + (continuous_iff_coinduced_le.mpr <| -- Porting note: didn't need function before le_iSup (fun j => - coinduced ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.app j) (F.obj j).str) j⟩ + coinduced ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.app j) (F.obj j).str) j) naturality := fun _ _ f => - ContinuousMap.coe_injective ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.naturality f) } + ConcreteCategory.coe_ext ((Types.TypeMax.colimitCocone (F ⋙ forget)).ι.naturality f) } /-- The chosen cocone `TopCat.colimitCocone F` for a functor `F : J ⥤ TopCat` is a colimit cocone. Generally you should just use `colimit.isColimit F`, unless you need the actual definition @@ -149,17 +147,16 @@ def colimitCoconeIsColimit (F : J ⥤ TopCat.{max v u}) : IsColimit (colimitCoco IsColimit.ofFaithful forget (Types.TypeMax.colimitCoconeIsColimit.{v, u} _) (fun s => -- Porting note: it appears notation for forget breaks dot notation (also above) -- Porting note: previously function was inferred - ⟨Quot.lift (fun p => (Functor.mapCocone forget s).ι.app p.fst p.snd) ?_, ?_⟩) fun s => ?_ + @ofHom _ _ (_) (_) <| @ContinuousMap.mk _ _ (_) (_) + (Quot.lift (fun p => (Functor.mapCocone forget s).ι.app p.fst p.snd) ?_) + ?_) fun s => ?_ · intro _ _ ⟨_, h⟩ - dsimp - rw [h, Functor.comp_map, ← comp_apply', s.ι.naturality] - dsimp - rw [Category.comp_id] + simp [h, ← ConcreteCategory.comp_apply, s.ι.naturality] · exact continuous_iff_le_induced.mpr (iSup_le fun j => coinduced_le_iff_le_induced.mp <| - (continuous_iff_coinduced_le.mp (s.ι.app j).continuous :)) + (continuous_iff_coinduced_le.mp (s.ι.app j).hom.continuous :)) · rfl instance topCat_hasColimitsOfSize : HasColimitsOfSize.{w,v} TopCat.{max v u} where @@ -185,7 +182,7 @@ instance forget_preservesColimits : PreservesColimits (forget : TopCat.{u} ⥤ T /-- The terminal object of `Top` is `PUnit`. -/ def isTerminalPUnit : IsTerminal (TopCat.of PUnit.{u + 1}) := haveI : ∀ X, Unique (X ⟶ TopCat.of PUnit.{u + 1}) := fun X => - ⟨⟨⟨fun _ => PUnit.unit, continuous_const⟩⟩, fun f => by ext; aesop⟩ + ⟨⟨ofHom ⟨fun _ => PUnit.unit, continuous_const⟩⟩, fun f => by ext⟩ Limits.IsTerminal.ofUnique _ /-- The terminal object of `Top` is `PUnit`. -/ @@ -195,7 +192,7 @@ def terminalIsoPUnit : ⊤_ TopCat.{u} ≅ TopCat.of PUnit := /-- The initial object of `Top` is `PEmpty`. -/ def isInitialPEmpty : IsInitial (TopCat.of PEmpty.{u + 1}) := haveI : ∀ X, Unique (TopCat.of PEmpty.{u + 1} ⟶ X) := fun X => - ⟨⟨⟨fun x => x.elim, by continuity⟩⟩, fun f => by ext ⟨⟩⟩ + ⟨⟨ofHom ⟨fun x => x.elim, by continuity⟩⟩, fun f => by ext ⟨⟩⟩ Limits.IsInitial.ofUnique _ /-- The initial object of `Top` is `PEmpty`. -/ diff --git a/Mathlib/Topology/Category/TopCat/Limits/Konig.lean b/Mathlib/Topology/Category/TopCat/Limits/Konig.lean index 0d6f7501547f2b..a1d9385502d9a2 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Konig.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Konig.lean @@ -29,7 +29,6 @@ discrete topology) in lemmas `nonempty_sections_of_finite_cofiltered_system` and (See for the Set version.) -/ - open CategoryTheory open CategoryTheory.Limits @@ -108,13 +107,7 @@ theorem partialSections.closed [∀ j : J, T2Space (F.obj j)] {G : Finset J} rw [this] apply isClosed_biInter intro f _ - -- Porting note: can't see through forget - have : T2Space ((forget TopCat).obj (F.obj f.snd.fst)) := - inferInstanceAs (T2Space (F.obj f.snd.fst)) - apply isClosed_eq - -- Porting note: used to be a single `continuity` that closed both goals - · exact (F.map f.snd.snd.snd.snd).continuous.comp (continuous_apply f.fst) - · continuity + apply isClosed_eq <;> fun_prop /-- Cofiltered limits of nonempty compact Hausdorff spaces are nonempty topological spaces. -/ diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index e9caaf169d458f..a1177ebe29171f 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -26,7 +26,7 @@ variable {J : Type v} [Category.{w} J] /-- The projection from the product as a bundled continuous map. -/ abbrev piπ {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) : TopCat.of (∀ i, α i) ⟶ α i := - ⟨fun f => f i, continuous_apply i⟩ + ofHom ⟨fun f => f i, continuous_apply i⟩ /-- The explicit fan of a family of topological spaces given by the pi type. -/ @[simps! pt π_app] @@ -35,15 +35,14 @@ def piFan {ι : Type v} (α : ι → TopCat.{max v u}) : Fan α := /-- The constructed fan is indeed a limit -/ def piFanIsLimit {ι : Type v} (α : ι → TopCat.{max v u}) : IsLimit (piFan α) where - lift S := + lift S := ofHom { toFun := fun s i => S.π.app ⟨i⟩ s - continuous_toFun := continuous_pi (fun i => (S.π.app ⟨i⟩).2) } + continuous_toFun := continuous_pi (fun i => (S.π.app ⟨i⟩).hom.2) } uniq := by intro S m h - apply ContinuousMap.ext; intro x + ext x funext i simp [ContinuousMap.coe_mk, ← h ⟨i⟩] - rfl fac _ _ := rfl /-- The product is homeomorphic to the product of the underlying spaces, @@ -68,10 +67,9 @@ theorem piIsoPi_hom_apply {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) rw [Iso.inv_comp_eq] at this exact ConcreteCategory.congr_hom this x --- Porting note: Lean doesn't automatically reduce TopCat.of X|>.α to X now /-- The inclusion to the coproduct as a bundled continuous map. -/ abbrev sigmaι {ι : Type v} (α : ι → TopCat.{max v u}) (i : ι) : α i ⟶ TopCat.of (Σi, α i) := by - refine ContinuousMap.mk ?_ ?_ + refine ofHom (ContinuousMap.mk ?_ ?_) · dsimp apply Sigma.mk i · dsimp; continuity @@ -83,13 +81,13 @@ def sigmaCofan {ι : Type v} (α : ι → TopCat.{max v u}) : Cofan α := /-- The constructed cofan is indeed a colimit -/ def sigmaCofanIsColimit {ι : Type v} (β : ι → TopCat.{max v u}) : IsColimit (sigmaCofan β) where - desc S := + desc S := ofHom { toFun := fun (s : of (Σ i, β i)) => S.ι.app ⟨s.1⟩ s.2 - continuous_toFun := continuous_sigma fun i => (S.ι.app ⟨i⟩).continuous_toFun } + continuous_toFun := by continuity } uniq := by intro S m h ext ⟨i, x⟩ - simp only [hom_apply, ← h] + simp only [← h] congr fac s j := by cases j @@ -114,7 +112,6 @@ theorem sigmaIsoSigma_inv_apply {ι : Type v} (α : ι → TopCat.{max v u}) (i rw [← sigmaIsoSigma_hom_ι_apply, ← comp_app, ← comp_app, Iso.hom_inv_id, Category.comp_id] --- Porting note: cannot use .topologicalSpace in place .str theorem induced_of_isLimit {F : J ⥤ TopCat.{max v u}} (C : Cone F) (hC : IsLimit C) : C.pt.str = ⨅ j, (F.obj j).str.induced (C.π.app j) := by let homeo := homeoOfIso (hC.conePointUniqueUpToIso (limitConeInfiIsLimit F)) @@ -132,11 +129,11 @@ section Prod -- Porting note: why is autoParam not firing? /-- The first projection from the product. -/ abbrev prodFst {X Y : TopCat.{u}} : TopCat.of (X × Y) ⟶ X := - ⟨Prod.fst, by continuity⟩ + ofHom ⟨Prod.fst, by continuity⟩ /-- The second projection from the product. -/ abbrev prodSnd {X Y : TopCat.{u}} : TopCat.of (X × Y) ⟶ Y := - ⟨Prod.snd, by continuity⟩ + ofHom ⟨Prod.snd, by continuity⟩ /-- The explicit binary cofan of `X, Y` given by `X × Y`. -/ def prodBinaryFan (X Y : TopCat.{u}) : BinaryFan X Y := @@ -144,18 +141,16 @@ def prodBinaryFan (X Y : TopCat.{u}) : BinaryFan X Y := /-- The constructed binary fan is indeed a limit -/ def prodBinaryFanIsLimit (X Y : TopCat.{u}) : IsLimit (prodBinaryFan X Y) where - lift := fun S : BinaryFan X Y => { + lift := fun S : BinaryFan X Y => ofHom { toFun := fun s => (S.fst s, S.snd s) - -- Porting note: continuity failed again here. Lean cannot infer - -- ContinuousMapClass (X ⟶ Y) X Y for X Y : TopCat which may be one of the problems - continuous_toFun := Continuous.prod_mk - (BinaryFan.fst S).continuous_toFun (BinaryFan.snd S).continuous_toFun } + continuous_toFun := by continuity } fac := by rintro S (_ | _) <;> {dsimp; ext; rfl} uniq := by intro S m h - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): used to be `ext x` - refine ContinuousMap.ext (fun (x : ↥(S.pt)) => Prod.ext ?_ ?_) + ext x + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): used to be part of `ext x` + refine Prod.ext ?_ ?_ · specialize h ⟨WalkingPair.left⟩ apply_fun fun e => e x at h exact h @@ -181,12 +176,11 @@ theorem prodIsoProd_hom_snd (X Y : TopCat.{u}) : simp [← Iso.eq_inv_comp, prodIsoProd] rfl --- Porting note: need to force Lean to coerce X × Y to a type -theorem prodIsoProd_hom_apply {X Y : TopCat.{u}} (x : ↑ (X ⨯ Y)) : +-- Note that `(x : X ⨯ Y)` would mean `(x : ↑X × ↑Y)` below: +theorem prodIsoProd_hom_apply {X Y : TopCat.{u}} (x : ↑(X ⨯ Y)) : (prodIsoProd X Y).hom x = ((Limits.prod.fst : X ⨯ Y ⟶ _) x, (Limits.prod.snd : X ⨯ Y ⟶ _) x) := by - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` didn't pick this up. - apply Prod.ext + ext · exact ConcreteCategory.congr_hom (prodIsoProd_hom_fst X Y) x · exact ConcreteCategory.congr_hom (prodIsoProd_hom_snd X Y) x @@ -215,24 +209,21 @@ theorem range_prod_map {W X Y Z : TopCat.{u}} (f : W ⟶ Y) (g : X ⟶ Z) : ext x constructor · rintro ⟨y, rfl⟩ - simp_rw [Set.mem_inter_iff, Set.mem_preimage, Set.mem_range] - -- sizable changes in this proof after https://github.com/leanprover-community/mathlib4/pull/13170 - rw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] - simp_rw [Limits.prod.map_fst, - Limits.prod.map_snd, CategoryTheory.comp_apply] - exact ⟨exists_apply_eq_apply _ _, exists_apply_eq_apply _ _⟩ + simp_rw [Set.mem_inter_iff, Set.mem_preimage, Set.mem_range, ← ConcreteCategory.comp_apply, + Limits.prod.map_fst, Limits.prod.map_snd, ConcreteCategory.comp_apply, exists_apply_eq_apply, + and_self_iff] · rintro ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩ use (prodIsoProd W X).inv (x₁, x₂) change (forget TopCat).map _ _ = _ apply Concrete.limit_ext rintro ⟨⟨⟩⟩ · change limit.π (pair Y Z) _ ((prod.map f g) _) = _ - erw [← CategoryTheory.comp_apply, Limits.prod.map_fst] + erw [← ConcreteCategory.comp_apply, Limits.prod.map_fst] change (_ ≫ _ ≫ f) _ = _ rw [TopCat.prodIsoProd_inv_fst_assoc,TopCat.comp_app] exact hx₁ · change limit.π (pair Y Z) _ ((prod.map f g) _) = _ - erw [← CategoryTheory.comp_apply, Limits.prod.map_snd] + erw [← ConcreteCategory.comp_apply, Limits.prod.map_snd] change (_ ≫ _ ≫ g) _ = _ rw [TopCat.prodIsoProd_inv_snd_assoc,TopCat.comp_app] exact hx₂ @@ -240,10 +231,9 @@ theorem range_prod_map {W X Y Z : TopCat.{u}} (f : W ⟶ Y) (g : X ⟶ Z) : theorem isInducing_prodMap {W X Y Z : TopCat.{u}} {f : W ⟶ X} {g : Y ⟶ Z} (hf : IsInducing f) (hg : IsInducing g) : IsInducing (Limits.prod.map f g) := by constructor - simp_rw [topologicalSpace_coe, prod_topology, induced_inf, induced_compose, ← coe_comp, + simp_rw [prod_topology, induced_inf, induced_compose, ← coe_comp, prod.map_fst, prod.map_snd, coe_comp, ← induced_compose (g := f), ← induced_compose (g := g)] - erw [← hf.eq_induced, ← hg.eq_induced] -- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 - rfl -- `rfl` was not needed before https://github.com/leanprover-community/mathlib4/pull/13170 + rw [← hf.eq_induced, ← hg.eq_induced] @[deprecated (since := "2024-10-28")] alias inducing_prod_map := isInducing_prodMap @@ -261,14 +251,13 @@ end Prod /-- The binary coproduct cofan in `TopCat`. -/ protected def binaryCofan (X Y : TopCat.{u}) : BinaryCofan X Y := - BinaryCofan.mk (⟨Sum.inl, by continuity⟩ : X ⟶ TopCat.of (X ⊕ Y)) ⟨Sum.inr, by continuity⟩ + BinaryCofan.mk (ofHom ⟨Sum.inl, by continuity⟩) (ofHom ⟨Sum.inr, by continuity⟩) /-- The constructed binary coproduct cofan in `TopCat` is the coproduct. -/ def binaryCofanIsColimit (X Y : TopCat.{u}) : IsColimit (TopCat.binaryCofan X Y) := by - refine Limits.BinaryCofan.isColimitMk (fun s => - {toFun := Sum.elim s.inl s.inr, continuous_toFun := ?_ }) ?_ ?_ ?_ - · apply - Continuous.sum_elim (BinaryCofan.inl s).continuous_toFun (BinaryCofan.inr s).continuous_toFun + refine Limits.BinaryCofan.isColimitMk (fun s => ofHom + { toFun := Sum.elim s.inl s.inr, continuous_toFun := ?_ }) ?_ ?_ ?_ + · continuity · intro s ext rfl @@ -277,7 +266,7 @@ def binaryCofanIsColimit (X Y : TopCat.{u}) : IsColimit (TopCat.binaryCofan X Y) rfl · intro s m h₁ h₂ ext (x | x) - exacts [(ConcreteCategory.congr_hom h₁ x :), (ConcreteCategory.congr_hom h₂ x :)] + exacts [ConcreteCategory.congr_hom h₁ x, ConcreteCategory.congr_hom h₂ x] theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : Nonempty (IsColimit c) ↔ @@ -303,7 +292,7 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : exact fun _ => or_not refine ⟨BinaryCofan.IsColimit.mk _ ?_ ?_ ?_ ?_⟩ · intro T f g - refine ContinuousMap.mk ?_ ?_ + refine ofHom (ContinuousMap.mk ?_ ?_) · exact fun x => if h : x ∈ Set.range c.inl then f ((Equiv.ofInjective _ h₁.injective).symm ⟨x, h⟩) else g ((Equiv.ofInjective _ h₂.injective).symm ⟨x, (this x).resolve_left h⟩) @@ -316,9 +305,7 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : convert_to Continuous (f ∘ (Homeomorph.ofIsEmbedding _ h₁.isEmbedding).symm) · ext ⟨x, hx⟩ exact dif_pos hx - apply Continuous.comp - · exact f.continuous_toFun - · continuity + continuity · exact h₁.isOpen_range · revert h x apply (IsOpen.continuousOn_iff _).mp @@ -331,7 +318,7 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : · ext ⟨x, hx⟩ exact dif_neg hx apply Continuous.comp - · exact g.continuous_toFun + · exact g.hom.continuous_toFun · apply Continuous.comp · continuity · rw [IsEmbedding.subtypeVal.isInducing.continuous_iff] @@ -341,17 +328,17 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : exact h₂.isOpen_range · intro T f g ext x - refine (dif_pos ?_).trans ?_ - · exact ⟨x, rfl⟩ - · dsimp - conv_lhs => rw [Equiv.ofInjective_symm_apply] + dsimp + rw [dif_pos] + conv_lhs => rw [Equiv.ofInjective_symm_apply] · intro T f g ext x - refine (dif_neg ?_).trans ?_ + dsimp + rw [dif_neg] + · exact congr_arg g (Equiv.ofInjective_symm_apply _ _) · rintro ⟨y, e⟩ have : c.inr x ∈ Set.range c.inl ⊓ Set.range c.inr := ⟨⟨_, e⟩, ⟨_, rfl⟩⟩ rwa [disjoint_iff.mp h₃.1] at this - · exact congr_arg g (Equiv.ofInjective_symm_apply _ _) · rintro T _ _ m rfl rfl ext x change m x = dite _ _ _ diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index 910c964c28b26b..9f15b3544fecb8 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -29,13 +29,13 @@ variable {X Y Z : TopCat.{u}} /-- The first projection from the pullback. -/ abbrev pullbackFst (f : X ⟶ Z) (g : Y ⟶ Z) : TopCat.of { p : X × Y // f p.1 = g p.2 } ⟶ X := - ⟨Prod.fst ∘ Subtype.val, by fun_prop⟩ + ofHom ⟨Prod.fst ∘ Subtype.val, by fun_prop⟩ lemma pullbackFst_apply (f : X ⟶ Z) (g : Y ⟶ Z) (x) : pullbackFst f g x = x.1.1 := rfl /-- The second projection from the pullback. -/ abbrev pullbackSnd (f : X ⟶ Z) (g : Y ⟶ Z) : TopCat.of { p : X × Y // f p.1 = g p.2 } ⟶ Y := - ⟨Prod.snd ∘ Subtype.val, by fun_prop⟩ + ofHom ⟨Prod.snd ∘ Subtype.val, by fun_prop⟩ lemma pullbackSnd_apply (f : X ⟶ Z) (g : Y ⟶ Z) (x) : pullbackSnd f g x = x.1.2 := rfl @@ -45,11 +45,7 @@ def pullbackCone (f : X ⟶ Z) (g : Y ⟶ Z) : PullbackCone f g := (by dsimp [pullbackFst, pullbackSnd, Function.comp_def] ext ⟨x, h⟩ - -- Next 2 lines were - -- `rw [comp_apply, ContinuousMap.coe_mk, comp_apply, ContinuousMap.coe_mk]` - -- `exact h` before https://github.com/leanprover/lean4/pull/2644 - rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply] - congr!) + simpa) /-- The constructed cone is a limit. -/ def pullbackConeIsLimit (f : X ⟶ Z) (g : Y ⟶ Z) : IsLimit (pullbackCone f g) := @@ -57,28 +53,24 @@ def pullbackConeIsLimit (f : X ⟶ Z) (g : Y ⟶ Z) : IsLimit (pullbackCone f g) (by intro S constructor; swap - · exact + · exact ofHom { toFun := fun x => ⟨⟨S.fst x, S.snd x⟩, by simpa using ConcreteCategory.congr_hom S.condition x⟩ continuous_toFun := by apply Continuous.subtype_mk <| Continuous.prod_mk ?_ ?_ - · exact (PullbackCone.fst S)|>.continuous_toFun - · exact (PullbackCone.snd S)|>.continuous_toFun + · exact (PullbackCone.fst S).hom.continuous_toFun + · exact (PullbackCone.snd S).hom.continuous_toFun } refine ⟨?_, ?_, ?_⟩ · delta pullbackCone ext a - -- This used to be `rw`, but we need `rw; rfl` after https://github.com/leanprover/lean4/pull/2644 - rw [CategoryTheory.comp_apply, ContinuousMap.coe_mk] - rfl + dsimp · delta pullbackCone ext a - -- This used to be `rw`, but we need `rw; rfl` after https://github.com/leanprover/lean4/pull/2644 - rw [CategoryTheory.comp_apply, ContinuousMap.coe_mk] - rfl + dsimp · intro m h₁ h₂ + ext x -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): used to be `ext x`. - apply ContinuousMap.ext; intro x apply Subtype.ext apply Prod.ext · simpa using ConcreteCategory.congr_hom h₁ x @@ -117,12 +109,11 @@ theorem pullbackIsoProdSubtype_hom_snd (f : X ⟶ Z) (g : Y ⟶ Z) : (pullbackIsoProdSubtype f g).hom ≫ pullbackSnd f g = pullback.snd _ _ := by rw [← Iso.eq_inv_comp, pullbackIsoProdSubtype_inv_snd] --- Porting note: why do I need to tell Lean to coerce pullback to a type theorem pullbackIsoProdSubtype_hom_apply {f : X ⟶ Z} {g : Y ⟶ Z} - (x : HasForget.forget.obj (pullback f g)) : + (x : ↑(pullback f g)) : (pullbackIsoProdSubtype f g).hom x = ⟨⟨pullback.fst f g x, pullback.snd f g x⟩, by - simpa using ConcreteCategory.congr_hom pullback.condition x⟩ := by + simpa using CategoryTheory.congr_fun pullback.condition x⟩ := by apply Subtype.ext; apply Prod.ext exacts [ConcreteCategory.congr_hom (pullbackIsoProdSubtype_hom_fst f g) x, ConcreteCategory.congr_hom (pullbackIsoProdSubtype_hom_snd f g) x] @@ -143,11 +134,13 @@ theorem range_pullback_to_prod {X Y Z : TopCat} (f : X ⟶ Z) (g : Y ⟶ Z) : ext x constructor · rintro ⟨y, rfl⟩ - change (_ ≫ _ ≫ f) _ = (_ ≫ _ ≫ g) _ -- new `change` after https://github.com/leanprover-community/mathlib4/pull/13170 + simp only [← ConcreteCategory.comp_apply, Set.mem_setOf_eq] simp [pullback.condition] · rintro (h : f (_, _).1 = g (_, _).2) use (pullbackIsoProdSubtype f g).inv ⟨⟨_, _⟩, h⟩ - change (forget TopCat).map _ _ = _ -- new `change` after https://github.com/leanprover-community/mathlib4/pull/13170 + -- new `change` after https://github.com/leanprover-community/mathlib4/pull/13170 + -- should be removed when we redo limits for `ConcreteCategory` instead of `HasForget` + change (forget TopCat).map _ _ = _ apply Concrete.limit_ext rintro ⟨⟨⟩⟩ <;> erw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply, limit.lift_π] <;> -- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 @@ -181,7 +174,7 @@ def pullbackHomeoPreimage theorem isInducing_pullback_to_prod {X Y Z : TopCat.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) : IsInducing <| ⇑(prod.lift (pullback.fst f g) (pullback.snd f g)) := - ⟨by simp [topologicalSpace_coe, prod_topology, pullback_topology, induced_compose, ← coe_comp]⟩ + ⟨by simp [prod_topology, pullback_topology, induced_compose, ← coe_comp]⟩ @[deprecated (since := "2024-10-28")] alias inducing_pullback_to_prod := isInducing_pullback_to_prod @@ -202,39 +195,46 @@ theorem range_pullback_map {W X Y Z S T : TopCat} (f₁ : W ⟶ S) (f₂ : X ⟶ constructor · rintro ⟨y, rfl⟩ simp only [Set.mem_inter_iff, Set.mem_preimage, Set.mem_range] - rw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] + rw [← ConcreteCategory.comp_apply, ← ConcreteCategory.comp_apply] simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, CategoryTheory.comp_apply] exact ⟨exists_apply_eq_apply _ _, exists_apply_eq_apply _ _⟩ rintro ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩ have : f₁ x₁ = f₂ x₂ := by apply (TopCat.mono_iff_injective _).mp H₃ - rw [← CategoryTheory.comp_apply, eq₁, ← CategoryTheory.comp_apply, eq₂, - CategoryTheory.comp_apply, CategoryTheory.comp_apply, hx₁, hx₂, ← CategoryTheory.comp_apply, - pullback.condition] - rfl -- `rfl` was not needed before https://github.com/leanprover-community/mathlib4/pull/13170 + rw [← ConcreteCategory.comp_apply, eq₁, ← ConcreteCategory.comp_apply, eq₂, + ConcreteCategory.comp_apply, ConcreteCategory.comp_apply, hx₁, hx₂, + ← ConcreteCategory.comp_apply, pullback.condition, ConcreteCategory.comp_apply] use (pullbackIsoProdSubtype f₁ f₂).inv ⟨⟨x₁, x₂⟩, this⟩ + -- `change` should be removed when we redo limits for `ConcreteCategory` instead of `HasForget` change (forget TopCat).map _ _ = _ apply Concrete.limit_ext rintro (_ | _ | _) <;> erw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] -- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 · simp only [Category.assoc, limit.lift_π, PullbackCone.mk_π_app_one] - simp only [cospan_one, pullbackIsoProdSubtype_inv_fst_assoc, CategoryTheory.comp_apply] - rw [pullbackFst_apply, hx₁, ← limit.w _ WalkingCospan.Hom.inl, cospan_map_inl, - CategoryTheory.comp_apply (g := g₁)] + simp only [cospan_one, pullbackIsoProdSubtype_inv_fst_assoc, ConcreteCategory.comp_apply, + CategoryTheory.comp_apply] + -- Work around the `ConcreteCategory`/`HasForget` mismatch: + change g₁ (i₁ ((pullbackFst f₁ f₂) ⟨(x₁, x₂), this⟩)) = (limit.π (cospan g₁ g₂) none) _ + simp [pullbackFst_apply, hx₁, ← limit.w _ WalkingCospan.Hom.inl, cospan_map_inl] · simp only [cospan_left, limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, pullbackIsoProdSubtype_inv_fst_assoc, CategoryTheory.comp_apply] - erw [hx₁] -- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 + -- Work around the `ConcreteCategory`/`HasForget` mismatch: + change i₁ ((pullbackFst f₁ f₂) ⟨(x₁, x₂), this⟩) = + limit.π (cospan g₁ g₂) (some WalkingPair.left) _ + simp [hx₁] · simp only [cospan_right, limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, pullbackIsoProdSubtype_inv_snd_assoc, CategoryTheory.comp_apply] - erw [hx₂] -- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 + change i₂ (pullbackSnd f₁ f₂ ⟨(x₁, x₂), this⟩) = + limit.π (cospan g₁ g₂) (some WalkingPair.right) _ + simp [hx₂] theorem pullback_fst_range {X Y S : TopCat} (f : X ⟶ S) (g : Y ⟶ S) : Set.range (pullback.fst f g) = { x : X | ∃ y : Y, f x = g y } := by ext x constructor - · rintro ⟨(y : (forget TopCat).obj _), rfl⟩ - use (pullback.snd f g) y - exact ConcreteCategory.congr_hom pullback.condition y + · rintro ⟨y, rfl⟩ + use pullback.snd f g y + exact CategoryTheory.congr_fun pullback.condition y · rintro ⟨y, eq⟩ use (TopCat.pullbackIsoProdSubtype f g).inv ⟨⟨x, y⟩, eq⟩ rw [pullbackIsoProdSubtype_inv_fst_apply] @@ -243,9 +243,9 @@ theorem pullback_snd_range {X Y S : TopCat} (f : X ⟶ S) (g : Y ⟶ S) : Set.range (pullback.snd f g) = { y : Y | ∃ x : X, f x = g y } := by ext y constructor - · rintro ⟨(x : (forget TopCat).obj _), rfl⟩ - use (pullback.fst f g) x - exact ConcreteCategory.congr_hom pullback.condition x + · rintro ⟨x, rfl⟩ + use pullback.fst f g x + exact CategoryTheory.congr_fun pullback.condition x · rintro ⟨x, eq⟩ use (TopCat.pullbackIsoProdSubtype f g).inv ⟨⟨x, y⟩, eq⟩ rw [pullbackIsoProdSubtype_inv_snd_apply] @@ -406,9 +406,9 @@ theorem pullback_snd_image_fst_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : Set X) g ⁻¹' (f '' U) := by ext x constructor - · rintro ⟨(y : (forget TopCat).obj _), hy, rfl⟩ + · rintro ⟨y, hy, rfl⟩ exact - ⟨(pullback.fst f g) y, hy, ConcreteCategory.congr_hom pullback.condition y⟩ + ⟨(pullback.fst f g) y, hy, CategoryTheory.congr_fun pullback.condition y⟩ · rintro ⟨y, hy, eq⟩ -- next 5 lines were -- `exact ⟨(TopCat.pullbackIsoProdSubtype f g).inv ⟨⟨_, _⟩, eq⟩, by simpa, by simp⟩` before https://github.com/leanprover-community/mathlib4/pull/13170 @@ -423,10 +423,10 @@ theorem pullback_fst_image_snd_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : Set Y) f ⁻¹' (g '' U) := by ext x constructor - · rintro ⟨(y : (forget TopCat).obj _), hy, rfl⟩ + · rintro ⟨y, hy, rfl⟩ exact ⟨(pullback.snd f g) y, hy, - (ConcreteCategory.congr_hom pullback.condition y).symm⟩ + (CategoryTheory.congr_fun pullback.condition y).symm⟩ · rintro ⟨y, hy, eq⟩ -- next 5 lines were -- `exact ⟨(TopCat.pullbackIsoProdSubtype f g).inv ⟨⟨_, _⟩, eq.symm⟩, by simpa, by simp⟩` @@ -434,7 +434,7 @@ theorem pullback_fst_image_snd_preimage (f : X ⟶ Z) (g : Y ⟶ Z) (U : Set Y) refine ⟨(TopCat.pullbackIsoProdSubtype f g).inv ⟨⟨_, _⟩, eq.symm⟩, ?_, ?_⟩ · simp only [coe_of, Set.mem_preimage] convert hy - erw [pullbackIsoProdSubtype_inv_snd_apply] + rw [pullbackIsoProdSubtype_inv_snd_apply] · rw [pullbackIsoProdSubtype_inv_fst_apply] end Pullback @@ -453,7 +453,6 @@ theorem colimit_topology (F : J ⥤ TopCat.{max v u}) : theorem colimit_isOpen_iff (F : J ⥤ TopCat.{max v u}) (U : Set ((colimit F :) : Type max v u)) : IsOpen U ↔ ∀ j, IsOpen (colimit.ι F j ⁻¹' U) := by - dsimp [topologicalSpace_coe] conv_lhs => rw [colimit_topology F] exact isOpen_iSup_iff @@ -467,7 +466,7 @@ theorem coequalizer_isOpen_iff (F : WalkingParallelPair ⥤ TopCat.{u}) · intro H j cases j · rw [← colimit.w F WalkingParallelPairHom.left] - exact (F.map WalkingParallelPairHom.left).continuous_toFun.isOpen_preimage _ H + exact (F.map WalkingParallelPairHom.left).hom.continuous_toFun.isOpen_preimage _ H · exact H end TopCat diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index cd24b21e34cb9b..f7ded3baca8503 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -98,8 +98,9 @@ theorem leSupr_apply_mk {ι : Type*} (U : ι → Opens X) (i : ι) (x) (m) : realising each open set as a topological space itself. -/ def toTopCat (X : TopCat.{u}) : Opens X ⥤ TopCat where - obj U := ⟨U, inferInstance⟩ - map i := ⟨fun x ↦ ⟨x.1, i.le x.2⟩, IsEmbedding.subtypeVal.continuous_iff.2 continuous_induced_dom⟩ + obj U := TopCat.of U + map i := TopCat.ofHom ⟨fun x ↦ ⟨x.1, i.le x.2⟩, + IsEmbedding.subtypeVal.continuous_iff.2 continuous_induced_dom⟩ @[simp] theorem toTopCat_map (X : TopCat.{u}) {U V : Opens X} {f : U ⟶ V} {x} {h} : @@ -108,10 +109,11 @@ theorem toTopCat_map (X : TopCat.{u}) {U V : Opens X} {f : U ⟶ V} {x} {h} : /-- The inclusion map from an open subset to the whole space, as a morphism in `TopCat`. -/ -@[simps (config := .asFn)] -def inclusion' {X : TopCat.{u}} (U : Opens X) : (toTopCat X).obj U ⟶ X where - toFun := _ - continuous_toFun := continuous_subtype_val +@[simps! (config := .asFn)] +def inclusion' {X : TopCat.{u}} (U : Opens X) : (toTopCat X).obj U ⟶ X := + TopCat.ofHom + { toFun := _ + continuous_toFun := continuous_subtype_val } @[simp] theorem coe_inclusion' {X : TopCat} {U : Opens X} : @@ -127,12 +129,12 @@ alias openEmbedding := isOpenEmbedding -/ def inclusionTopIso (X : TopCat.{u}) : (toTopCat X).obj ⊤ ≅ X where hom := inclusion' ⊤ - inv := ⟨fun x => ⟨x, trivial⟩, continuous_def.2 fun _ ⟨_, hS, hSU⟩ => hSU ▸ hS⟩ + inv := TopCat.ofHom ⟨fun x => ⟨x, trivial⟩, continuous_def.2 fun _ ⟨_, hS, hSU⟩ => hSU ▸ hS⟩ /-- `Opens.map f` gives the functor from open sets in Y to open set in X, given by taking preimages under f. -/ def map (f : X ⟶ Y) : Opens Y ⥤ Opens X where - obj U := ⟨f ⁻¹' (U : Set Y), U.isOpen.preimage f.continuous⟩ + obj U := ⟨f ⁻¹' (U : Set Y), U.isOpen.preimage f.hom.continuous⟩ map i := ⟨⟨fun _ h => i.le h⟩⟩ @[simp] @@ -140,7 +142,7 @@ theorem map_coe (f : X ⟶ Y) (U : Opens Y) : ((map f).obj U : Set X) = f ⁻¹' rfl @[simp] -theorem map_obj (f : X ⟶ Y) (U) (p) : (map f).obj ⟨U, p⟩ = ⟨f ⁻¹' U, p.preimage f.continuous⟩ := +theorem map_obj (f : X ⟶ Y) (U) (p) : (map f).obj ⟨U, p⟩ = ⟨f ⁻¹' U, p.preimage f.hom.continuous⟩ := rfl @[simp] diff --git a/Mathlib/Topology/Category/TopCat/Yoneda.lean b/Mathlib/Topology/Category/TopCat/Yoneda.lean index f724695c46e335..5542ec352bdf88 100644 --- a/Mathlib/Topology/Category/TopCat/Yoneda.lean +++ b/Mathlib/Topology/Category/TopCat/Yoneda.lean @@ -32,7 +32,7 @@ A universe polymorphic "Yoneda presheaf" on `C` given by continuous maps into a @[simps] def yonedaPresheaf : Cᵒᵖ ⥤ Type (max w w') where obj X := C(F.obj (unop X), Y) - map f g := ContinuousMap.comp g (F.map f.unop) + map f g := ContinuousMap.comp g (F.map f.unop).hom /-- A universe polymorphic Yoneda presheaf on `TopCat` given by continuous maps into a topoological @@ -41,7 +41,7 @@ space `Y`. @[simps] def yonedaPresheaf' : TopCat.{w}ᵒᵖ ⥤ Type (max w w') where obj X := C((unop X).1, Y) - map f g := ContinuousMap.comp g f.unop + map f g := ContinuousMap.comp g f.unop.hom theorem comp_yonedaPresheaf' : yonedaPresheaf F Y = F.op ⋙ yonedaPresheaf' Y := rfl diff --git a/Mathlib/Topology/Category/TopCommRingCat.lean b/Mathlib/Topology/Category/TopCommRingCat.lean index eedae4582314d9..6544499fdfee33 100644 --- a/Mathlib/Topology/Category/TopCommRingCat.lean +++ b/Mathlib/Topology/Category/TopCommRingCat.lean @@ -86,7 +86,7 @@ instance forgetToCommRingCatTopologicalSpace (R : TopCommRingCat) : /-- The forgetful functor to `TopCat`. -/ instance hasForgetToTopCat : HasForget₂ TopCommRingCat TopCat := - HasForget₂.mk' (fun R => TopCat.of R) (fun _ => rfl) (fun f => ⟨⇑f.1, f.2⟩) HEq.rfl + HasForget₂.mk' (fun R => TopCat.of R) (fun _ => rfl) (fun f => TopCat.ofHom ⟨⇑f.1, f.2⟩) HEq.rfl instance forgetToTopCatCommRing (R : TopCommRingCat) : CommRing ((forget₂ TopCommRingCat TopCat).obj R) := @@ -107,7 +107,7 @@ instance : (forget₂ TopCommRingCat.{u} TopCat.{u}).ReflectsIsomorphisms where let e_Ring : X ≃+* Y := { f.1, ((forget TopCat).mapIso i_Top).toEquiv with } -- Putting these together we obtain the isomorphism we're after: exact - ⟨⟨⟨e_Ring.symm, i_Top.inv.2⟩, + ⟨⟨⟨e_Ring.symm, i_Top.inv.hom.2⟩, ⟨by ext x exact e_Ring.left_inv x, by diff --git a/Mathlib/Topology/Category/UniformSpace.lean b/Mathlib/Topology/Category/UniformSpace.lean index 37b17fb3b293d0..e2e2b390df7a8c 100644 --- a/Mathlib/Topology/Category/UniformSpace.lean +++ b/Mathlib/Topology/Category/UniformSpace.lean @@ -78,7 +78,7 @@ theorem hom_ext {X Y : UniformSpaceCat} {f g : X ⟶ Y} : (f : X → Y) = g → instance hasForgetToTop : HasForget₂ UniformSpaceCat.{u} TopCat.{u} where forget₂ := { obj := fun X => TopCat.of X - map := fun f => + map := fun f => TopCat.ofHom { toFun := f continuous_toFun := f.property.continuous } } diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index d0db78c7c2d298..d83763afd6e29d 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -906,8 +906,8 @@ theorem Topology.IsInducing.isCompact_iff {f : X → Y} (hf : IsInducing f) : @[deprecated (since := "2024-10-28")] alias Inducing.isCompact_iff := IsInducing.isCompact_iff -/-- If `f : X → Y` is an `Embedding`, the image `f '' s` of a set `s` is compact - if and only if `s` is compact. -/ +/-- If `f : X → Y` is an embedding, the image `f '' s` of a set `s` is compact +if and only if `s` is compact. -/ theorem Topology.IsEmbedding.isCompact_iff {f : X → Y} (hf : IsEmbedding f) : IsCompact s ↔ IsCompact (f '' s) := hf.isInducing.isCompact_iff diff --git a/Mathlib/Topology/Compactness/Lindelof.lean b/Mathlib/Topology/Compactness/Lindelof.lean index 258f4db27756f0..d74d5d157ddc5d 100644 --- a/Mathlib/Topology/Compactness/Lindelof.lean +++ b/Mathlib/Topology/Compactness/Lindelof.lean @@ -610,8 +610,8 @@ theorem Topology.IsInducing.isLindelof_iff {f : X → Y} (hf : IsInducing f) : @[deprecated (since := "2024-10-28")] alias Inducing.isLindelof_iff := IsInducing.isLindelof_iff -/-- If `f : X → Y` is an `Embedding`, the image `f '' s` of a set `s` is Lindelöf - if and only if `s` is Lindelöf. -/ +/-- If `f : X → Y` is an embedding, the image `f '' s` of a set `s` is Lindelöf +if and only if `s` is Lindelöf. -/ theorem Topology.IsEmbedding.isLindelof_iff {f : X → Y} (hf : IsEmbedding f) : IsLindelof s ↔ IsLindelof (f '' s) := hf.isInducing.isLindelof_iff diff --git a/Mathlib/Topology/Compactness/SigmaCompact.lean b/Mathlib/Topology/Compactness/SigmaCompact.lean index 92f225b9e60be6..adf7cf56cc86c4 100644 --- a/Mathlib/Topology/Compactness/SigmaCompact.lean +++ b/Mathlib/Topology/Compactness/SigmaCompact.lean @@ -121,8 +121,8 @@ lemma Topology.IsInducing.isSigmaCompact_iff {f : X → Y} {s : Set X} @[deprecated (since := "2024-10-28")] alias Inducing.isSigmaCompact_iff := IsInducing.isSigmaCompact_iff -/-- If `f : X → Y` is an `Embedding`, the image `f '' s` of a set `s` is σ-compact - if and only `s` is σ-compact. -/ +/-- If `f : X → Y` is an embedding, the image `f '' s` of a set `s` is σ-compact +if and only `s` is σ-compact. -/ lemma Topology.IsEmbedding.isSigmaCompact_iff {f : X → Y} {s : Set X} (hf : IsEmbedding f) : IsSigmaCompact s ↔ IsSigmaCompact (f '' s) := hf.isInducing.isSigmaCompact_iff diff --git a/Mathlib/Topology/Constructible.lean b/Mathlib/Topology/Constructible.lean index 4ad63357cdd714..5fcb1b6feb84f4 100644 --- a/Mathlib/Topology/Constructible.lean +++ b/Mathlib/Topology/Constructible.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.BooleanSubalgebra +import Mathlib.Topology.QuasiSeparated import Mathlib.Topology.Spectral.Hom /-! @@ -296,10 +297,13 @@ variable [CompactSpace X] {P : ∀ s : Set X, IsConstructible s → Prop} {B : S lemma _root_.IsRetrocompact.isCompact (hs : IsRetrocompact s) : IsCompact s := by simpa using hs CompactSpace.isCompact_univ +variable [QuasiSeparatedSpace X] + /-- Variant of `TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact` for a non-indexed topological basis. -/ +@[stacks 0069 "Iff form of (2). Note that Stacks doesn't define quasi-separated spaces."] lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact' - (basis : IsTopologicalBasis B) (compact_inter : ∀ U ∈ B, ∀ V ∈ B, IsCompact (U ∩ V)) + (basis : IsTopologicalBasis B) (isCompact_basis : ∀ U ∈ B, IsCompact U) (hU : IsOpen U) : IsRetrocompact U ↔ IsCompact U := by refine ⟨IsRetrocompact.isCompact, fun hU' {V} hV' hV ↦ ?_⟩ obtain ⟨s, rfl⟩ := eq_sUnion_finset_of_isTopologicalBasis_of_isCompact_open _ basis _ hU' hU @@ -308,43 +312,44 @@ lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact' refine ((s.finite_toSet.image _).prod (t.finite_toSet.image _)).isCompact_biUnion ?_ simp only [mem_prod, mem_image, Finset.mem_coe, Subtype.exists, exists_and_right, exists_eq_right, and_imp, forall_exists_index, Prod.forall] - exact fun u v hu _ hv _ ↦ compact_inter _ hu _ hv + exact fun u v hu _ hv _ ↦ (isCompact_basis _ hu).inter_of_isOpen (isCompact_basis _ hv) + (basis.isOpen hu) (basis.isOpen hv) +@[stacks 0069 "Iff form of (2). Note that Stacks doesn't define quasi-separated spaces."] lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact - (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) + (basis : IsTopologicalBasis (range b)) (isCompact_basis : ∀ i, IsCompact (b i)) (hU : IsOpen U) : IsRetrocompact U ↔ IsCompact U := - basis.isRetrocompact_iff_isCompact' (by simpa using compact_inter) hU + basis.isRetrocompact_iff_isCompact' (by simpa using isCompact_basis) hU /-- Variant of `TopologicalSpace.IsTopologicalBasis.isRetrocompact` for a non-indexed topological basis. -/ lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact' (basis : IsTopologicalBasis B) - (compact_inter : ∀ U ∈ B, ∀ V ∈ B, IsCompact (U ∩ V)) (hU : U ∈ B) : IsRetrocompact U := - (basis.isRetrocompact_iff_isCompact' compact_inter <| basis.isOpen hU).2 <| by - simpa using compact_inter _ hU _ hU + (isCompact_basis : ∀ U ∈ B, IsCompact U) (hU : U ∈ B) : IsRetrocompact U := + (basis.isRetrocompact_iff_isCompact' isCompact_basis <| basis.isOpen hU).2 <| isCompact_basis _ hU lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact - (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) (i : ι) : + (basis : IsTopologicalBasis (range b)) (isCompact_basis : ∀ i, IsCompact (b i)) (i : ι) : IsRetrocompact (b i) := - (basis.isRetrocompact_iff_isCompact compact_inter <| basis.isOpen <| mem_range_self _).2 <| by - simpa using compact_inter i i + (basis.isRetrocompact_iff_isCompact isCompact_basis <| basis.isOpen <| mem_range_self _).2 <| + isCompact_basis _ /-- Variant of `TopologicalSpace.IsTopologicalBasis.isConstructible` for a non-indexed topological basis. -/ lemma _root_.TopologicalSpace.IsTopologicalBasis.isConstructible' (basis : IsTopologicalBasis B) - (compact_inter : ∀ U ∈ B, ∀ V ∈ B, IsCompact (U ∩ V)) (hU : U ∈ B) : IsConstructible U := - (basis.isRetrocompact' compact_inter hU).isConstructible <| basis.isOpen hU + (isCompact_basis : ∀ U ∈ B, IsCompact U) (hU : U ∈ B) : IsConstructible U := + (basis.isRetrocompact' isCompact_basis hU).isConstructible <| basis.isOpen hU lemma _root_.TopologicalSpace.IsTopologicalBasis.isConstructible - (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) (i : ι) : + (basis : IsTopologicalBasis (range b)) (isCompact_basis : ∀ i, IsCompact (b i)) (i : ι) : IsConstructible (b i) := - (basis.isRetrocompact compact_inter _).isConstructible <| basis.isOpen <| mem_range_self _ + (basis.isRetrocompact isCompact_basis _).isConstructible <| basis.isOpen <| mem_range_self _ @[elab_as_elim] lemma IsConstructible.induction_of_isTopologicalBasis {ι : Type*} [Nonempty ι] (b : ι → Set X) - (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) + (basis : IsTopologicalBasis (range b)) (isCompact_basis : ∀ i, IsCompact (b i)) (sdiff : ∀ i s (hs : Set.Finite s), P (b i \ ⋃ j ∈ s, b j) - ((basis.isConstructible compact_inter _).sdiff <| .biUnion hs fun _ _ ↦ - basis.isConstructible compact_inter _)) + ((basis.isConstructible isCompact_basis _).sdiff <| .biUnion hs fun _ _ ↦ + basis.isConstructible isCompact_basis _)) (union : ∀ s hs t ht, P s hs → P t ht → P (s ∪ t) (hs.union ht)) (s : Set X) (hs : IsConstructible s) : P s hs := by induction s, hs using BooleanSubalgebra.closure_sdiff_sup_induction with @@ -354,8 +359,7 @@ lemma IsConstructible.induction_of_isTopologicalBasis {ι : Type*} [Nonempty ι] | bot_mem => exact ⟨isOpen_empty, .empty⟩ | top_mem => exact ⟨isOpen_univ, .univ⟩ | sdiff U hU V hV => - have := isCompact_open_iff_eq_finite_iUnion_of_isTopologicalBasis _ basis - fun i ↦ by simpa using compact_inter i i + have := isCompact_open_iff_eq_finite_iUnion_of_isTopologicalBasis _ basis isCompact_basis obtain ⟨s, hs, rfl⟩ := (this _).1 ⟨hU.2.isCompact, hU.1⟩ obtain ⟨t, ht, rfl⟩ := (this _).1 ⟨hV.2.isCompact, hV.1⟩ simp_rw [iUnion_diff] @@ -364,11 +368,11 @@ lemma IsConstructible.induction_of_isTopologicalBasis {ι : Type*} [Nonempty ι] | @insert i s hi hs ih => simp_rw [biUnion_insert] exact union _ _ _ - (.biUnion hs fun i _ ↦ (basis.isConstructible compact_inter _).sdiff <| - .biUnion ht fun j _ ↦ basis.isConstructible compact_inter _) + (.biUnion hs fun i _ ↦ (basis.isConstructible isCompact_basis _).sdiff <| + .biUnion ht fun j _ ↦ basis.isConstructible isCompact_basis _) (sdiff _ _ ht) (ih ⟨isOpen_biUnion fun _ _ ↦ basis.isOpen ⟨_, rfl⟩, .biUnion hs - fun i _ ↦ basis.isRetrocompact compact_inter _⟩) + fun i _ ↦ basis.isRetrocompact isCompact_basis _⟩) | sup s _ t _ hs' ht' => exact union _ _ _ _ hs' ht' end CompactSpace diff --git a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean index 331fef8c6d1015..b7975118f44fa3 100644 --- a/Mathlib/Topology/ContinuousMap/CompactlySupported.lean +++ b/Mathlib/Topology/ContinuousMap/CompactlySupported.lean @@ -3,8 +3,7 @@ Copyright (c) 2024 Yoh Tanimoto. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yoh Tanimoto -/ -import Mathlib.Topology.Algebra.Support -import Mathlib.Topology.ContinuousMap.CocompactMap +import Mathlib.Topology.Algebra.Order.Support import Mathlib.Topology.ContinuousMap.ZeroAtInfty /-! @@ -124,6 +123,29 @@ def ContinuousMap.liftCompactlySupported [CompactSpace α] : C(α, β) ≃ C_c( left_inv _ := rfl right_inv _ := rfl +variable {γ : Type*} [TopologicalSpace γ] [Zero γ] + +/-- Composition of a continuous function `f` with compact support with another continuous function +`g` sending `0` to `0` from the left yields another continuous function `g ∘ f` with compact +support. + +If `g` doesn't send `0` to `0`, `f.compLeft g` defaults to `0`. -/ +noncomputable def compLeft (g : C(β, γ)) (f : C_c(α, β)) : C_c(α, γ) where + toContinuousMap := by classical exact if g 0 = 0 then g.comp f else 0 + hasCompactSupport' := by + split_ifs with hg + · exact f.hasCompactSupport'.comp_left hg + · exact .zero + +lemma toContinuousMap_compLeft {g : C(β, γ)} (hg : g 0 = 0) (f : C_c(α, β)) : + (f.compLeft g).toContinuousMap = g.comp f := if_pos hg + +lemma coe_compLeft {g : C(β, γ)} (hg : g 0 = 0) (f : C_c(α, β)) : f.compLeft g = g ∘ f := by + simp [compLeft, if_pos hg] + +lemma compLeft_apply {g : C(β, γ)} (hg : g 0 = 0) (f : C_c(α, β)) (a : α) : + f.compLeft g a = g (f a) := by simp [coe_compLeft hg f] + end Basics /-! ### Algebraic structure @@ -399,6 +421,84 @@ instance : StarRing C_c(α, β) := end StarRing +section PartialOrder + +/-! ### The partial order in `C_c` +When `β` is equipped with a partial order, `C_c(α, β)` is given the pointwise partial order. +-/ + +variable {β : Type*} [TopologicalSpace β] [Zero β] [PartialOrder β] + +instance partialOrder : PartialOrder C_c(α, β) := PartialOrder.lift (⇑) DFunLike.coe_injective + +theorem le_def {f g : C_c(α, β)} : f ≤ g ↔ ∀ a, f a ≤ g a := Pi.le_def + +theorem lt_def {f g : C_c(α, β)} : f < g ↔ (∀ a, f a ≤ g a) ∧ ∃ a, f a < g a := Pi.lt_def + +end PartialOrder + +section SemilatticeSup + +variable [SemilatticeSup β] [Zero β] [TopologicalSpace β] [ContinuousSup β] + +instance instSup : Max C_c(α, β) where max f g := + { toFun := f ⊔ g + continuous_toFun := Continuous.sup f.continuous g.continuous + hasCompactSupport' := f.hasCompactSupport.sup g.hasCompactSupport } + +@[simp, norm_cast] lemma coe_sup (f g : C_c(α, β)) : ⇑(f ⊔ g) = ⇑f ⊔ g := rfl + +@[simp] lemma sup_apply (f g : C_c(α, β)) (a : α) : (f ⊔ g) a = f a ⊔ g a := rfl + +instance semilatticeSup : SemilatticeSup C_c(α, β) := + DFunLike.coe_injective.semilatticeSup _ coe_sup + +lemma finsetSup'_apply {ι : Type*} {s : Finset ι} (H : s.Nonempty) (f : ι → C_c(α, β)) (a : α) : + s.sup' H f a = s.sup' H fun i ↦ f i a := + Finset.comp_sup'_eq_sup'_comp H (fun g : C_c(α, β) ↦ g a) fun _ _ ↦ rfl + +@[simp, norm_cast] +lemma coe_finsetSup' {ι : Type*} {s : Finset ι} (H : s.Nonempty) (f : ι → C_c(α, β)) : + ⇑(s.sup' H f) = s.sup' H fun i ↦ ⇑(f i) := by ext; simp [finsetSup'_apply] + +end SemilatticeSup + +section SemilatticeInf + +variable [SemilatticeInf β] [Zero β] [TopologicalSpace β] [ContinuousInf β] + +instance instInf : Min C_c(α, β) where min f g := + { toFun := f ⊓ g + continuous_toFun := Continuous.inf f.continuous g.continuous + hasCompactSupport' := f.hasCompactSupport.inf g.hasCompactSupport } + +@[simp, norm_cast] lemma coe_inf (f g : C_c(α, β)) : ⇑(f ⊓ g) = ⇑f ⊓ g := rfl + +@[simp] lemma inf_apply (f g : C_c(α, β)) (a : α) : (f ⊓ g) a = f a ⊓ g a := rfl + +instance semilatticeInf : SemilatticeInf C_c(α, β) := + DFunLike.coe_injective.semilatticeInf _ coe_inf + +lemma finsetInf'_apply {ι : Type*} {s : Finset ι} (H : s.Nonempty) (f : ι → C_c(α, β)) (a : α) : + s.inf' H f a = s.inf' H fun i ↦ f i a := + Finset.comp_inf'_eq_inf'_comp H (fun g : C_c(α, β) ↦ g a) fun _ _ ↦ rfl + +@[simp, norm_cast] +lemma coe_finsetInf' {ι : Type*} {s : Finset ι} (H : s.Nonempty) (f : ι → C_c(α, β)) : + ⇑(s.inf' H f) = s.inf' H fun i ↦ ⇑(f i) := by ext; simp [finsetInf'_apply] + +end SemilatticeInf + +section Lattice + +instance [Lattice β] [TopologicalSpace β] [TopologicalLattice β] [Zero β] : + Lattice C_c(α, β) := + DFunLike.coe_injective.lattice _ coe_sup coe_inf + +-- TODO transfer this lattice structure to `BoundedContinuousFunction` + +end Lattice + /-! ### `C_c` as a functor For each `β` with sufficient structure, there is a contravariant functor `C_c(-, β)` from the @@ -541,6 +641,58 @@ noncomputable def nnrealPart (f : C_c(α, ℝ)) : C_c(α, ℝ≥0) where lemma nnrealPart_apply (f : C_c(α, ℝ)) (x : α) : f.nnrealPart x = Real.toNNReal (f x) := rfl +/-- The compactly supported continuous `ℝ≥0`-valued function as a compactly supported `ℝ`-valued +function. -/ +noncomputable def toReal (f : C_c(α, ℝ≥0)) : C_c(α, ℝ) := + f.compLeft ContinuousMap.coeNNRealReal + +@[simp] +lemma toReal_apply (f : C_c(α, ℝ≥0)) (x : α) : f.toReal x = f x := compLeft_apply rfl _ _ + +@[simp] lemma toReal_nonneg {f : C_c(α, ℝ≥0)} : 0 ≤ f.toReal := fun _ ↦ by simp + +@[simp] lemma toReal_add (f g : C_c(α, ℝ≥0)) : (f + g).toReal = f.toReal + g.toReal := by ext; simp +@[simp] lemma toReal_smul (r : ℝ≥0) (f : C_c(α, ℝ≥0)) : (r • f).toReal = r • f.toReal := by + ext; simp [NNReal.smul_def] + +lemma nnrealPart_sub_nnrealPart_neg (f : C_c(α, ℝ)) : + (nnrealPart f).toReal - (nnrealPart (-f)).toReal = f := by + ext x + simp + +/-- The compactly supported continuous `ℝ≥0`-valued function as a compactly supported `ℝ`-valued +function. -/ +noncomputable def toRealLinearMap : C_c(α, ℝ≥0) →ₗ[ℝ≥0] C_c(α, ℝ) where + toFun := toReal + map_add' f g := by ext x; simp + map_smul' a f := by ext x; simp [NNReal.smul_def] + +@[simp, norm_cast] +lemma coe_toRealLinearMap : (toRealLinearMap : C_c(α, ℝ≥0) → C_c(α, ℝ)) = toReal := rfl + +lemma toRealLinearMap_apply (f : C_c(α, ℝ≥0)) : toRealLinearMap f = f.toReal := rfl + +lemma toRealLinearMap_apply_apply (f : C_c(α, ℝ≥0)) (x : α) : + toRealLinearMap f x = (f x).toReal := by simp + +/-- For a positive linear functional `Λ : C_c(α, ℝ) → ℝ`, define a `ℝ≥0`-linear map. -/ +noncomputable def toNNRealLinear (Λ : C_c(α, ℝ) →ₗ[ℝ] ℝ) (hΛ : ∀ f, 0 ≤ f → 0 ≤ Λ f) : + C_c(α, ℝ≥0) →ₗ[ℝ≥0] ℝ≥0 where + toFun f := ⟨Λ (toRealLinearMap f), hΛ _ <| by simp⟩ + map_add' f g := by ext; simp + map_smul' a f := by ext; simp [NNReal.smul_def] + +@[simp] +lemma toNNRealLinear_apply (Λ : C_c(α, ℝ) →ₗ[ℝ] ℝ) (hΛ) (f : C_c(α, ℝ≥0)) : + toNNRealLinear Λ hΛ f = Λ (toReal f) := rfl + +@[simp] lemma toNNRealLinear_inj (Λ₁ Λ₂ : C_c(α, ℝ) →ₗ[ℝ] ℝ) (hΛ₁ hΛ₂) : + toNNRealLinear Λ₁ hΛ₁ = toNNRealLinear Λ₂ hΛ₂ ↔ Λ₁ = Λ₂ := by + simp only [LinearMap.ext_iff, NNReal.eq_iff, toNNRealLinear_apply] + refine ⟨fun h f ↦ ?_, fun h f ↦ by rw [LinearMap.ext h]⟩ + rw [← nnrealPart_sub_nnrealPart_neg f] + simp_rw [map_sub, h] + end CompactlySupportedContinuousMap end NonnegativePart diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index 338841fc6393b3..4e561c7a9e60ea 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -25,7 +25,7 @@ as well as topology inducing maps, topological embeddings, and quotient maps. * `IsInducing`: a map `f : X → Y` is called *inducing*, if the topology on the domain is equal to the induced topology. -* `Embedding`: a map `f : X → Y` is an *embedding*, +* `IsEmbedding`: a map `f : X → Y` is an *embedding*, if it is a topology inducing map and it is injective. * `IsOpenEmbedding`: a map `f : X → Y` is an *open embedding*, diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 1e3280f06bdbef..804c803d834ffd 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -112,16 +112,20 @@ def Rel (a b : Σ i, ((D.U i : TopCat) : Type _)) : Prop := ∃ x : D.V (a.1, b.1), D.f _ _ x = a.2 ∧ D.f _ _ (D.t _ _ x) = b.2 theorem rel_equiv : Equivalence D.Rel := - ⟨fun x => ⟨inv (D.f _ _) x.2, IsIso.inv_hom_id_apply _ _, - by simp [IsIso.inv_hom_id_apply _ _]⟩, by + ⟨fun x => ⟨inv (D.f _ _) x.2, IsIso.inv_hom_id_apply (D.f x.fst x.fst) _, + -- Use `elementwise_of%` elaborator instead of `IsIso.inv_hom_id_apply` to work around + -- `ConcreteCategory`/`HasForget` mismatch: + by simp [elementwise_of% IsIso.inv_hom_id (D.f x.fst x.fst)]⟩, by rintro a b ⟨x, e₁, e₂⟩ - exact ⟨D.t _ _ x, e₂, by rw [← e₁, D.t_inv_apply]⟩, by + -- `erw` works around the `ConcreteCategory`/`HasForget` mismatch: + exact ⟨D.t _ _ x, e₂, by erw [← e₁, D.t_inv_apply]⟩, by rintro ⟨i, a⟩ ⟨j, b⟩ ⟨k, c⟩ ⟨x, e₁, e₂⟩ rintro ⟨y, e₃, e₄⟩ let z := (pullbackIsoProdSubtype (D.f j i) (D.f j k)).inv ⟨⟨_, _⟩, e₂.trans e₃.symm⟩ have eq₁ : (D.t j i) ((pullback.fst _ _ : _ /-(D.f j k)-/ ⟶ D.V (j, i)) z) = x := by dsimp only [coe_of, z] - erw [pullbackIsoProdSubtype_inv_fst_apply, D.t_inv_apply]-- now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 + -- `erw` works around the `ConcreteCategory`/`HasForget` mismatch: + erw [pullbackIsoProdSubtype_inv_fst_apply, D.t_inv_apply] have eq₂ : (pullback.snd _ _ : _ ⟶ D.V _) z = y := pullbackIsoProdSubtype_inv_snd_apply _ _ _ clear_value z use (pullback.fst _ _ : _ ⟶ D.V (i, k)) (D.t' _ _ _ z) @@ -135,7 +139,7 @@ theorem rel_equiv : Equivalence D.Rel := apply @Epi.left_cancellation _ _ _ _ (D.t' k j i) rw [𝖣.cocycle_assoc, 𝖣.t_fac_assoc, 𝖣.t_inv_assoc] exact pullback.condition.symm - exact ⟨ContinuousMap.congr_fun h₁ z, ContinuousMap.congr_fun h₂ z⟩⟩ + exact ⟨CategoryTheory.congr_fun h₁ z, CategoryTheory.congr_fun h₂ z⟩⟩ open CategoryTheory.Limits.WalkingParallelPair @@ -152,12 +156,15 @@ theorem eqvGen_of_π_eq x y := by delta GlueData.π Multicoequalizer.sigmaπ at h -- Porting note: inlined `inferInstance` instead of leaving as a side goal. - replace h := (TopCat.mono_iff_injective (Multicoequalizer.isoCoequalizer 𝖣.diagram).inv).mp + replace h : coequalizer.π D.diagram.fstSigmaMap D.diagram.sndSigmaMap x = + coequalizer.π D.diagram.fstSigmaMap D.diagram.sndSigmaMap y := + (TopCat.mono_iff_injective (Multicoequalizer.isoCoequalizer 𝖣.diagram).inv).mp inferInstance h let diagram := parallelPair 𝖣.diagram.fstSigmaMap 𝖣.diagram.sndSigmaMap ⋙ forget _ have : colimit.ι diagram one x = colimit.ι diagram one y := by - dsimp only [coequalizer.π, ContinuousMap.toFun_eq_coe] at h - rw [← ι_preservesColimitIso_hom, forget_map_eq_coe, types_comp_apply, h] + dsimp only [coequalizer.π] at h + rw [← ι_preservesColimitIso_hom, ConcreteCategory.forget_map_eq_coe, types_comp_apply] + erw [h] simp have : (colimit.ι diagram _ ≫ colim.map _ ≫ (colimit.isoColimitCocone _).hom) _ = @@ -194,9 +201,9 @@ theorem ι_eq_iff_rel (i j : D.J) (x : D.U i) (y : D.U j) : (ConcreteCategory.bijective_of_isIso (sigmaIsoSigma.{u, u} _).inv).2 x unfold InvImage MultispanIndex.fstSigmaMap MultispanIndex.sndSigmaMap simp only [forget_map_eq_coe] - erw [TopCat.comp_app, sigmaIsoSigma_inv_apply, ← CategoryTheory.comp_apply, - ← CategoryTheory.comp_apply, colimit.ι_desc_assoc, ← CategoryTheory.comp_apply, - ← CategoryTheory.comp_apply, colimit.ι_desc_assoc] + erw [TopCat.comp_app, sigmaIsoSigma_inv_apply, ← ConcreteCategory.comp_apply, + ← ConcreteCategory.comp_apply, colimit.ι_desc_assoc, ← ConcreteCategory.comp_apply, + ← ConcreteCategory.comp_apply, colimit.ι_desc_assoc] -- previous line now `erw` after https://github.com/leanprover-community/mathlib4/pull/13170 erw [sigmaIsoSigma_hom_ι_apply, sigmaIsoSigma_hom_ι_apply] exact ⟨y, ⟨rfl, rfl⟩⟩ @@ -204,7 +211,9 @@ theorem ι_eq_iff_rel (i j : D.J) (x : D.U i) (y : D.U j) : dsimp only at * -- Porting note: there were `subst e₁` and `subst e₂`, instead of the `rw` rw [← e₁, ← e₂] at * - rw [D.glue_condition_apply] + -- `erw; rfl` works around the `ConcreteCategory`/`HasForget` mismatch: + erw [D.glue_condition_apply] + rfl theorem ι_injective (i : D.J) : Function.Injective (𝖣.ι i) := by intro x y h @@ -227,7 +236,8 @@ theorem image_inter (i j : D.J) : exact ⟨y, by simp [e₁]⟩ · rintro ⟨x, hx⟩ refine ⟨⟨D.f i j x, hx⟩, ⟨D.f j i (D.t _ _ x), ?_⟩⟩ - rw [D.glue_condition_apply] + -- `erw` works around the `ConcreteCategory`/`HasForget` mismatch: + erw [D.glue_condition_apply] exact hx theorem preimage_range (i j : D.J) : 𝖣.ι j ⁻¹' Set.range (𝖣.ι i) = Set.range (D.f j i) := by @@ -243,7 +253,9 @@ theorem preimage_image_eq_image (i j : D.J) (U : Set (𝖣.U i)) : generalize 𝖣.ι i '' U = U' -- next 4 lines were `simp` before https://github.com/leanprover-community/mathlib4/pull/13170 simp only [GlueData.diagram_l, GlueData.diagram_r, Set.mem_preimage, coe_comp, Function.comp_apply] - rw [D.glue_condition_apply] + -- `erw; rfl` works around the `ConcreteCategory`/`HasForget` mismatch: + erw [D.glue_condition_apply] + rfl rw [← this, Set.image_preimage_eq_inter_range] symm apply Set.inter_eq_self_of_subset_left @@ -269,11 +281,11 @@ theorem open_image_open (i : D.J) (U : Opens (𝖣.U i)) : IsOpen (𝖣.ι i '' intro j rw [preimage_image_eq_image] apply (D.f_open _ _).isOpenMap - apply (D.t j i ≫ D.f i j).continuous_toFun.isOpen_preimage + apply (D.t j i ≫ D.f i j).hom.continuous_toFun.isOpen_preimage exact U.isOpen theorem ι_isOpenEmbedding (i : D.J) : IsOpenEmbedding (𝖣.ι i) := - .of_continuous_injective_isOpenMap (𝖣.ι i).continuous_toFun (D.ι_injective i) fun U h => + .of_continuous_injective_isOpenMap (𝖣.ι i).hom.continuous_toFun (D.ι_injective i) fun U h => D.open_image_open i ⟨U, h⟩ @[deprecated (since := "2024-10-18")] @@ -319,7 +331,7 @@ instance (h : MkCore.{u}) (i j : h.J) : IsIso (h.t i j) := by def MkCore.t' (h : MkCore.{u}) (i j k : h.J) : pullback (h.V i j).inclusion' (h.V i k).inclusion' ⟶ pullback (h.V j k).inclusion' (h.V j i).inclusion' := by - refine (pullbackIsoProdSubtype _ _).hom ≫ ⟨?_, ?_⟩ ≫ (pullbackIsoProdSubtype _ _).inv + refine (pullbackIsoProdSubtype _ _).hom ≫ ofHom ⟨?_, ?_⟩ ≫ (pullbackIsoProdSubtype _ _).inv · intro x refine ⟨⟨⟨(h.t i j x.1.1).1, ?_⟩, h.t i j x.1.1⟩, rfl⟩ rcases x with ⟨⟨⟨x, hx⟩, ⟨x', hx'⟩⟩, rfl : x = x'⟩ @@ -355,8 +367,9 @@ def mk' (h : MkCore.{u}) : TopCat.GlueData where simp only [Iso.inv_hom_id_assoc, Category.assoc, Category.id_comp] rw [← Iso.eq_inv_comp, Iso.inv_hom_id] ext1 ⟨⟨⟨x, hx⟩, ⟨x', hx'⟩⟩, rfl : x = x'⟩ - rw [comp_app, ContinuousMap.coe_mk, comp_app, id_app, ContinuousMap.coe_mk, Subtype.mk_eq_mk, - Prod.mk.inj_iff, Subtype.mk_eq_mk, Subtype.ext_iff, and_self_iff] + dsimp only [Opens.coe_inclusion', hom_comp, hom_ofHom, ContinuousMap.comp_assoc, + ContinuousMap.comp_apply, ContinuousMap.coe_mk, hom_id, ContinuousMap.id_apply] + rw [Subtype.mk_eq_mk, Prod.mk.inj_iff, Subtype.mk_eq_mk, Subtype.ext_iff, and_self_iff] convert congr_arg Subtype.val (h.t_inv k i ⟨x, hx'⟩) using 3 refine Subtype.ext ?_ exact h.cocycle i j k ⟨x, hx⟩ hx' @@ -372,7 +385,7 @@ def ofOpenSubsets : TopCat.GlueData.{u} := { J U := fun i => (Opens.toTopCat <| TopCat.of α).obj (U i) V := fun _ j => (Opens.map <| Opens.inclusion' _).obj (U j) - t := fun i j => ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, by fun_prop⟩ + t := fun i j => ofHom ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, by fun_prop⟩ V_id := fun i => by ext; simp t_id := fun i => by ext; rfl t_inter := fun _ _ _ _ hx => hx @@ -420,7 +433,7 @@ theorem fromOpenSubsetsGlue_isOpenMap : IsOpenMap (fromOpenSubsetsGlue U) := by exact Set.preimage_image_eq _ (fromOpenSubsetsGlue_injective U) · refine ⟨Set.mem_image_of_mem _ hx, ?_⟩ rw [ι_fromOpenSubsetsGlue_apply] - exact Set.mem_range_self _ + exact Set.mem_range_self (f := (Opens.inclusion' _).hom) ⟨x, hx'⟩ theorem fromOpenSubsetsGlue_isOpenEmbedding : IsOpenEmbedding (fromOpenSubsetsGlue U) := .of_continuous_injective_isOpenMap (ContinuousMap.continuous_toFun _) @@ -447,7 +460,7 @@ def openCoverGlueHomeo (h : ⋃ i, (U i : Set α) = Set.univ) : (Equiv.ofBijective (fromOpenSubsetsGlue U) ⟨fromOpenSubsetsGlue_injective U, Set.range_eq_univ.mp ((range_fromOpenSubsetsGlue U).symm ▸ h)⟩) - (fromOpenSubsetsGlue U).2 (fromOpenSubsetsGlue_isOpenMap U) + (fromOpenSubsetsGlue U).hom.2 (fromOpenSubsetsGlue_isOpenMap U) end GlueData diff --git a/Mathlib/Topology/Instances/NNReal/Defs.lean b/Mathlib/Topology/Instances/NNReal/Defs.lean index 4ab30b14fb4b0c..5f3144de0b9ed1 100644 --- a/Mathlib/Topology/Instances/NNReal/Defs.lean +++ b/Mathlib/Topology/Instances/NNReal/Defs.lean @@ -85,6 +85,9 @@ instance [TopologicalSpace α] [MulAction ℝ α] [ContinuousSMul ℝ α] : def _root_.ContinuousMap.coeNNRealReal : C(ℝ≥0, ℝ) := ⟨(↑), continuous_coe⟩ +@[simp] +lemma coeNNRealReal_zero : ContinuousMap.coeNNRealReal 0 = 0 := rfl + instance ContinuousMap.canLift {X : Type*} [TopologicalSpace X] : CanLift C(X, ℝ) C(X, ℝ≥0) ContinuousMap.coeNNRealReal.comp fun f => ∀ x, 0 ≤ f x where prf f hf := ⟨⟨fun x => ⟨f x, hf x⟩, f.2.subtype_mk _⟩, DFunLike.ext' rfl⟩ diff --git a/Mathlib/Topology/MetricSpace/HolderNorm.lean b/Mathlib/Topology/MetricSpace/HolderNorm.lean index 424128d4adb859..b7fdaf6f47dff2 100644 --- a/Mathlib/Topology/MetricSpace/HolderNorm.lean +++ b/Mathlib/Topology/MetricSpace/HolderNorm.lean @@ -61,7 +61,8 @@ lemma HolderWith.memHolder {C : ℝ≥0} (hf : HolderWith C r f) : MemHolder r f refine ⟨fun h => ?_, fun hf => let ⟨C, hC⟩ := hf; iInf_lt_top.2 ⟨C, iInf_lt_top.2 ⟨hC, coe_lt_top⟩⟩⟩ simp_rw [eHolderNorm, iInf_lt_top] at h - exact let ⟨C, hC, _⟩ := h; ⟨C, hC⟩ + let ⟨C, hC, _⟩ := h + exact ⟨C, hC⟩ lemma eHolderNorm_ne_top : eHolderNorm r f ≠ ∞ ↔ MemHolder r f := by rw [← eHolderNorm_lt_top, lt_top_iff_ne_top] @@ -92,8 +93,7 @@ variable (X) in lemma nnHolderNorm_const (r : ℝ≥0) (c : Y) : nnHolderNorm r (Function.const X c) = 0 := by refine le_antisymm (ENNReal.coe_le_coe.1 <| le_trans coe_nnHolderNorm_le_eHolderNorm ?_) (zero_le _) - rw [eHolderNorm_const] - rfl + rw [eHolderNorm_const, ENNReal.coe_zero] variable (X) in @[simp] diff --git a/Mathlib/Topology/Order/Category/AlexDisc.lean b/Mathlib/Topology/Order/Category/AlexDisc.lean index fd44e293d43370..da80cb58877d52 100644 --- a/Mathlib/Topology/Order/Category/AlexDisc.lean +++ b/Mathlib/Topology/Order/Category/AlexDisc.lean @@ -14,45 +14,50 @@ maps, and proves it's equivalent to the category of preorders. open CategoryTheory Topology -/-- Auxiliary typeclass to define the category of Alexandrov-discrete spaces. Do not use this -directly. Use `AlexandrovDiscrete` instead. -/ -class AlexandrovDiscreteSpace (α : Type*) extends TopologicalSpace α, AlexandrovDiscrete α - /-- The category of Alexandrov-discrete spaces. -/ -def AlexDisc := Bundled AlexandrovDiscreteSpace +structure AlexDisc extends TopCat where + [is_alexandrovDiscrete : AlexandrovDiscrete carrier] namespace AlexDisc -instance instCoeSort : CoeSort AlexDisc Type* := Bundled.coeSort -instance instTopologicalSpace (α : AlexDisc) : TopologicalSpace α := α.2.1 -instance instAlexandrovDiscrete (α : AlexDisc) : AlexandrovDiscrete α := α.2.2 +attribute [instance] is_alexandrovDiscrete -instance : BundledHom.ParentProjection @AlexandrovDiscreteSpace.toTopologicalSpace := ⟨⟩ +instance : CoeSort AlexDisc (Type _) := + ⟨fun X => X.toTopCat⟩ -deriving instance LargeCategory for AlexDisc +instance category : Category AlexDisc := + InducedCategory.category toTopCat -instance instHasForget : HasForget AlexDisc := BundledHom.hasForget _ -instance instHasForgetToTop : HasForget₂ AlexDisc TopCat := BundledHom.forget₂ _ _ -instance forgetToTop_full : (forget₂ AlexDisc TopCat).Full := BundledHom.forget₂_full _ _ -instance forgetToTop_faithful : (forget₂ AlexDisc TopCat).Faithful where +instance concreteCategory : ConcreteCategory AlexDisc (C(·, ·)) := + InducedCategory.concreteCategory toTopCat -@[simp] lemma coe_forgetToTop (X : AlexDisc) : ↥((forget₂ _ TopCat).obj X) = X := rfl +instance instHasForgetToTop : HasForget₂ AlexDisc TopCat := InducedCategory.hasForget₂ toTopCat + +-- TODO: generalize to `InducedCategory.forget₂_full`? +instance forgetToTop_full : (forget₂ AlexDisc TopCat).Full where + map_surjective f := ⟨f, rfl⟩ + +instance forgetToTop_faithful : (forget₂ AlexDisc TopCat).Faithful where /-- Construct a bundled `AlexDisc` from the underlying topological space. -/ -def of (α : Type*) [TopologicalSpace α] [AlexandrovDiscrete α] : AlexDisc := ⟨α, ⟨⟩⟩ +abbrev of (X : Type*) [TopologicalSpace X] [AlexandrovDiscrete X] : AlexDisc where + toTopCat := TopCat.of X + +lemma coe_of (α : Type*) [TopologicalSpace α] [AlexandrovDiscrete α] : ↥(of α) = α := rfl -@[simp] lemma coe_of (α : Type*) [TopologicalSpace α] [AlexandrovDiscrete α] : ↥(of α) = α := rfl @[simp] lemma forgetToTop_of (α : Type*) [TopologicalSpace α] [AlexandrovDiscrete α] : (forget₂ AlexDisc TopCat).obj (of α) = TopCat.of α := rfl +@[simp] lemma coe_forgetToTop (X : AlexDisc) : ↥((forget₂ _ TopCat).obj X) = X := rfl + -- This was a global instance prior to https://github.com/leanprover-community/mathlib4/pull/13170. We may experiment with removing it. attribute [local instance] CategoryTheory.HasForget.instFunLike /-- Constructs an equivalence between preorders from an order isomorphism between them. -/ @[simps] def Iso.mk {α β : AlexDisc} (e : α ≃ₜ β) : α ≅ β where - hom := (e : ContinuousMap α β) - inv := (e.symm : ContinuousMap β α) + hom := TopCat.ofHom (e : ContinuousMap α β) + inv := TopCat.ofHom (e.symm : ContinuousMap β α) hom_inv_id := DFunLike.ext _ _ e.symm_apply_apply inv_hom_id := DFunLike.ext _ _ e.apply_symm_apply @@ -62,7 +67,8 @@ end AlexDisc @[simps] def alexDiscEquivPreord : AlexDisc ≌ Preord where functor := forget₂ _ _ ⋙ topToPreord - inverse := { obj := fun X ↦ AlexDisc.of (WithUpperSet X), map := WithUpperSet.map } + inverse.obj X := AlexDisc.of (WithUpperSet X) + inverse.map f := TopCat.ofHom (WithUpperSet.map f) unitIso := NatIso.ofComponents fun X ↦ AlexDisc.Iso.mk <| by dsimp; exact homeoWithUpperSetTopologyorderIso X counitIso := NatIso.ofComponents fun X ↦ Preord.Iso.mk <| by diff --git a/Mathlib/Topology/Order/Category/FrameAdjunction.lean b/Mathlib/Topology/Order/Category/FrameAdjunction.lean index 2501e640cac851..192dfc9fe540fd 100644 --- a/Mathlib/Topology/Order/Category/FrameAdjunction.lean +++ b/Mathlib/Topology/Order/Category/FrameAdjunction.lean @@ -84,8 +84,9 @@ topological spaces, which sends a locale `L` to the topological space `PT L` of from `L` to `Prop` and a locale homomorphism `f` to a continuous function between the spaces of points. -/ def pt : Locale ⥤ TopCat where - obj L := ⟨PT L.unop, inferInstance⟩ - map f := ⟨fun p ↦ p.comp f.unop, continuous_def.2 <| by rintro s ⟨u, rfl⟩; use f.unop u; rfl⟩ + obj L := TopCat.of (PT L.unop) + map f := TopCat.ofHom ⟨fun p ↦ p.comp f.unop, + continuous_def.2 <| by rintro s ⟨u, rfl⟩; use f.unop u; rfl⟩ end pt_definition section locale_top_adjunction @@ -110,7 +111,7 @@ def counitAppCont : FrameHom L (Opens <| PT L) where /-- The forgetful functor `topToLocale` is left adjoint to the functor `pt`. -/ def adjunctionTopToLocalePT : topToLocale ⊣ pt where - unit := { app := fun X ↦ ⟨localePointOfSpacePoint X, continuous_def.2 <| + unit := { app := fun X ↦ TopCat.ofHom ⟨localePointOfSpacePoint X, continuous_def.2 <| by rintro _ ⟨u, rfl⟩; simpa using u.2⟩ } counit := { app := fun L ↦ ⟨counitAppCont L⟩ } diff --git a/Mathlib/Topology/Order/UpperLowerSetTopology.lean b/Mathlib/Topology/Order/UpperLowerSetTopology.lean index 1c2c2390d6ebec..a7364240ec2728 100644 --- a/Mathlib/Topology/Order/UpperLowerSetTopology.lean +++ b/Mathlib/Topology/Order/UpperLowerSetTopology.lean @@ -54,7 +54,7 @@ namespace Topology /-- Topology whose open sets are upper sets. Note: In general the upper set topology does not coincide with the upper topology. -/ -def upperSet (α : Type*) [Preorder α] : TopologicalSpace α where +def upperSet (α : Type*) [Preorder α] : TopologicalSpace α where IsOpen := IsUpperSet isOpen_univ := isUpperSet_univ isOpen_inter _ _ := IsUpperSet.inter @@ -63,7 +63,7 @@ def upperSet (α : Type*) [Preorder α] : TopologicalSpace α where /-- Topology whose open sets are lower sets. Note: In general the lower set topology does not coincide with the lower topology. -/ -def lowerSet (α : Type*) [Preorder α] : TopologicalSpace α where +def lowerSet (α : Type*) [Preorder α] : TopologicalSpace α where IsOpen := IsLowerSet isOpen_univ := isLowerSet_univ isOpen_inter _ _ := IsLowerSet.inter @@ -407,7 +407,7 @@ def map (f : α →o β) : C(WithLowerSet α, WithLowerSet β) where @[simp] lemma map_comp (g : β →o γ) (f : α →o β) : map (g.comp f) = (map g).comp (map f) := rfl @[simp] lemma toLowerSet_specializes_toLowerSet {a b : α} : - toLowerSet a ⤳ toLowerSet b ↔ a ≤ b := by + toLowerSet a ⤳ toLowerSet b ↔ a ≤ b := by simp_rw [specializes_iff_closure_subset, IsLowerSet.closure_singleton, Ici_subset_Ici, toLowerSet_le_iff] diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index c0004b4df3cbab..956f5b021f4413 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -1363,7 +1363,7 @@ noncomputable def lift_openEmbedding (e : PartialHomeomorph X Z) (hf : IsOpenEmb rw [← hx'x, hf.toPartialHomeomorph_left_inv]; exact hx have : ContinuousOn (hf.toPartialHomeomorph).symm (f '' univ) := (hf.toPartialHomeomorph).continuousOn_invFun - apply this.mono; gcongr; exact fun ⦃a⦄ a ↦ trivial + exact this.mono <| image_mono <| subset_univ _ exact ContinuousOn.congr this heq continuousOn_invFun := hf.continuous.comp_continuousOn e.continuousOn_invFun diff --git a/Mathlib/Topology/QuasiSeparated.lean b/Mathlib/Topology/QuasiSeparated.lean index 26da564a4682f2..154b7fe4c19010 100644 --- a/Mathlib/Topology/QuasiSeparated.lean +++ b/Mathlib/Topology/QuasiSeparated.lean @@ -25,8 +25,7 @@ of compact open subsets are still compact. a quasi-separated space, then so is `α`. -/ - -open TopologicalSpace Topology +open Set TopologicalSpace Topology variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {f : α → β} @@ -111,14 +110,32 @@ instance (priority := 100) NoetherianSpace.to_quasiSeparatedSpace [NoetherianSpa QuasiSeparatedSpace α := ⟨fun _ _ _ _ _ _ => NoetherianSpace.isCompact _⟩ -theorem IsQuasiSeparated.of_quasiSeparatedSpace (s : Set α) [QuasiSeparatedSpace α] : - IsQuasiSeparated s := +lemma QuasiSeparatedSpace.of_isTopologicalBasis {ι : Type*} {b : ι → Set α} + (basis : IsTopologicalBasis (range b)) (isCompact_inter : ∀ i j, IsCompact (b i ∩ b j)) : + QuasiSeparatedSpace α where + inter_isCompact U V hUopen hUcomp hVopen hVcomp := by + have aux := isCompact_open_iff_eq_finite_iUnion_of_isTopologicalBasis b basis fun i ↦ by + simpa using isCompact_inter i i + obtain ⟨s, hs, rfl⟩ := (aux _).1 ⟨hUcomp, hUopen⟩ + obtain ⟨t, ht, rfl⟩ := (aux _).1 ⟨hVcomp, hVopen⟩ + rw [iUnion₂_inter_iUnion₂] + exact hs.isCompact_biUnion fun i hi ↦ ht.isCompact_biUnion fun j hj ↦ isCompact_inter .. + +section QuasiSeparatedSpace +variable [QuasiSeparatedSpace α] {U V : Set α} + +lemma IsQuasiSeparated.of_quasiSeparatedSpace (s : Set α) : IsQuasiSeparated s := isQuasiSeparated_univ.of_subset (Set.subset_univ _) -theorem QuasiSeparatedSpace.of_isOpenEmbedding (h : IsOpenEmbedding f) [QuasiSeparatedSpace β] : - QuasiSeparatedSpace α := - isQuasiSeparated_univ_iff.mp - (h.isQuasiSeparated_iff.mpr <| IsQuasiSeparated.of_quasiSeparatedSpace _) +lemma QuasiSeparatedSpace.of_isOpenEmbedding {f : β → α} (h : IsOpenEmbedding f) : + QuasiSeparatedSpace β := + isQuasiSeparated_univ_iff.mp (h.isQuasiSeparated_iff.mpr <| .of_quasiSeparatedSpace _) @[deprecated (since := "2024-10-18")] alias QuasiSeparatedSpace.of_openEmbedding := QuasiSeparatedSpace.of_isOpenEmbedding + +lemma IsCompact.inter_of_isOpen (hUcomp : IsCompact U) (hVcomp : IsCompact V) (hUopen : IsOpen U) + (hVopen : IsOpen V) : IsCompact (U ∩ V) := + QuasiSeparatedSpace.inter_isCompact _ _ hUopen hUcomp hVopen hVcomp + +end QuasiSeparatedSpace diff --git a/Mathlib/Topology/Sheaves/CommRingCat.lean b/Mathlib/Topology/Sheaves/CommRingCat.lean index 2bdba51313e3d1..9ee14ac28953d7 100644 --- a/Mathlib/Topology/Sheaves/CommRingCat.lean +++ b/Mathlib/Topology/Sheaves/CommRingCat.lean @@ -196,23 +196,54 @@ section ContinuousFunctions namespace TopCat -variable (X : TopCat.{v}) +variable (X : TopCat.{v}) (R : TopCommRingCat.{v}) + +instance : NatCast (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + natCast n := ofHom n + +instance : IntCast (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + intCast n := ofHom n + +instance : Zero (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + zero := ofHom 0 + +instance : One (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + one := ofHom 1 + +instance : Neg (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + neg f := ofHom (-f.hom) + +instance : Sub (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + sub f g := ofHom (f.hom - g.hom) + +instance : Add (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + add f g := ofHom (f.hom + g.hom) + +instance : Mul (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + mul f g := ofHom (f.hom * g.hom) + +instance : SMul ℕ (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + smul n f := ofHom (n • f.hom) + +instance : SMul ℤ (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) where + smul n f := ofHom (n • f.hom) + +instance : Pow (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) ℕ where + pow f n := ofHom (f.hom ^ n) + +instance : CommRing (X ⟶ (forget₂ TopCommRingCat TopCat).obj R) := + Function.Injective.commRing _ ConcreteCategory.hom_injective + rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) -- TODO upgrade the result to TopCommRing? /-- The (bundled) commutative ring of continuous functions from a topological space to a topological commutative ring, with pointwise multiplication. -/ def continuousFunctions (X : TopCat.{v}ᵒᵖ) (R : TopCommRingCat.{v}) : CommRingCat.{v} := - -- Porting note: Lean did not see through that `X.unop ⟶ R` is just continuous functions - -- hence forms a ring - @CommRingCat.of (X.unop ⟶ (forget₂ TopCommRingCat TopCat).obj R) <| - inferInstanceAs (CommRing (ContinuousMap _ _)) + CommRingCat.of (X.unop ⟶ (forget₂ TopCommRingCat TopCat).obj R) namespace continuousFunctions -instance (X : TopCat.{v}ᵒᵖ) (R : TopCommRingCat.{v}) : - CommRing (unop X ⟶ (forget₂ TopCommRingCat TopCat).obj R) := - inferInstanceAs (CommRing (ContinuousMap _ _)) - /-- Pulling back functions into a topological ring along a continuous map is a ring homomorphism. -/ def pullback {X Y : TopCatᵒᵖ} (f : X ⟶ Y) (R : TopCommRingCat) : continuousFunctions X R ⟶ continuousFunctions Y R := CommRingCat.ofHom @@ -227,12 +258,10 @@ this is a ring homomorphism (with respect to the pointwise ring operations on fu def map (X : TopCat.{u}ᵒᵖ) {R S : TopCommRingCat.{u}} (φ : R ⟶ S) : continuousFunctions X R ⟶ continuousFunctions X S := CommRingCat.ofHom { toFun g := g ≫ (forget₂ TopCommRingCat TopCat).map φ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` tactic does not work, since Lean can't see through `R ⟶ S` is just - -- continuous ring homomorphism - map_one' := ContinuousMap.ext fun _ => φ.1.map_one - map_zero' := ContinuousMap.ext fun _ => φ.1.map_zero - map_add' := fun _ _ => ContinuousMap.ext fun _ => φ.1.map_add _ _ - map_mul' := fun _ _ => ContinuousMap.ext fun _ => φ.1.map_mul _ _ } + map_one' := by ext; exact φ.1.map_one + map_zero' := by ext; exact φ.1.map_zero + map_add' _ _ := by ext; exact φ.1.map_add _ _ + map_mul' _ _ := by ext; exact φ.1.map_mul _ _ } end continuousFunctions @@ -353,9 +382,8 @@ theorem objSupIsoProdEqLocus_inv_eq_iff {X : TopCat.{u}} (F : X.Sheaf CommRingCa constructor · rintro rfl rw [← TopCat.Sheaf.objSupIsoProdEqLocus_inv_fst, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_snd] - -- `simp` doesn't see through the type equality of objects in `CommRingCat`, so use `rw` https://github.com/leanprover-community/mathlib4/pull/8386 - repeat rw [← CommRingCat.comp_apply] - simp only [← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp, and_self] + simp only [← CommRingCat.comp_apply, ← Functor.map_comp, ← op_comp, Category.assoc, + homOfLE_comp, and_self] · rintro ⟨e₁, e₂⟩ refine F.eq_of_locally_eq₂ (homOfLE (inf_le_right : U ⊓ W ≤ W)) (homOfLE (inf_le_right : V ⊓ W ≤ W)) ?_ _ _ ?_ ?_ @@ -364,15 +392,13 @@ theorem objSupIsoProdEqLocus_inv_eq_iff {X : TopCat.{u}} (F : X.Sheaf CommRingCa · change (F.val.map _) ((F.val.map (homOfLE e).op).hom ((F.objSupIsoProdEqLocus U V).inv.hom x)) = (F.val.map _) y rw [← e₁, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_fst] - -- `simp` doesn't see through the type equality of objects in `CommRingCat`, so use `rw` https://github.com/leanprover-community/mathlib4/pull/8386 - repeat rw [← CommRingCat.comp_apply] - simp only [← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp] + simp only [← CommRingCat.comp_apply, ← Functor.map_comp, ← op_comp, Category.assoc, + homOfLE_comp] · show (F.val.map _) ((F.val.map (homOfLE e).op).hom ((F.objSupIsoProdEqLocus U V).inv.hom x)) = (F.val.map _) y rw [← e₂, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_snd] - -- `simp` doesn't see through the type equality of objects in `CommRingCat`, so use `rw` https://github.com/leanprover-community/mathlib4/pull/8386 - repeat rw [← CommRingCat.comp_apply] - simp only [← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp] + simp only [← CommRingCat.comp_apply, ← Functor.map_comp, ← op_comp, Category.assoc, + homOfLE_comp] end TopCat.Sheaf diff --git a/Mathlib/Topology/Sheaves/LocalPredicate.lean b/Mathlib/Topology/Sheaves/LocalPredicate.lean index f212a015ac7f6e..b23ec369670d98 100644 --- a/Mathlib/Topology/Sheaves/LocalPredicate.lean +++ b/Mathlib/Topology/Sheaves/LocalPredicate.lean @@ -269,7 +269,7 @@ the presheaf of continuous functions. def subpresheafContinuousPrelocalIsoPresheafToTop (T : TopCat.{v}) : subpresheafToTypes (continuousPrelocal X T) ≅ presheafToTop X T := NatIso.ofComponents fun X => - { hom := by rintro ⟨f, c⟩; exact ⟨f, c⟩ + { hom := by rintro ⟨f, c⟩; exact ofHom ⟨f, c⟩ inv := by rintro ⟨f, c⟩; exact ⟨f, c⟩ } /-- The sheaf of continuous functions on `X` with values in a space `T`. diff --git a/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean b/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean index 3ff71a0720c23c..bca89d92fce460 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean @@ -135,8 +135,6 @@ end TypeValued section -attribute [local instance] HasForget.hasCoeToSort HasForget.instFunLike - variable [HasLimits C] [(forget C).ReflectsIsomorphisms] [PreservesLimits (forget C)] variable {X : TopCat.{v}} (F : Presheaf C X) diff --git a/Mathlib/Topology/Sheaves/Skyscraper.lean b/Mathlib/Topology/Sheaves/Skyscraper.lean index 259b736775448b..64f9541a6587ae 100644 --- a/Mathlib/Topology/Sheaves/Skyscraper.lean +++ b/Mathlib/Topology/Sheaves/Skyscraper.lean @@ -66,9 +66,10 @@ def skyscraperPresheaf : Presheaf C X where theorem skyscraperPresheaf_eq_pushforward [hd : ∀ U : Opens (TopCat.of PUnit.{u + 1}), Decidable (PUnit.unit ∈ U)] : skyscraperPresheaf p₀ A = - ContinuousMap.const (TopCat.of PUnit) p₀ _* + (ofHom (ContinuousMap.const (TopCat.of PUnit) p₀)) _* skyscraperPresheaf (X := TopCat.of PUnit) PUnit.unit A := by - convert_to @skyscraperPresheaf X p₀ (fun U => hd <| (Opens.map <| ContinuousMap.const _ p₀).obj U) + convert_to @skyscraperPresheaf X p₀ (fun U => hd <| (Opens.map <| ofHom <| + ContinuousMap.const _ p₀).obj U) C _ _ A = _ <;> congr /-- Taking skyscraper presheaf at a point is functorial: `c ↦ skyscraper p₀ c` defines a functor by diff --git a/Mathlib/Topology/Sheaves/Stalks.lean b/Mathlib/Topology/Sheaves/Stalks.lean index c977d29a84f28a..a6f2e68ef6c47f 100644 --- a/Mathlib/Topology/Sheaves/Stalks.lean +++ b/Mathlib/Topology/Sheaves/Stalks.lean @@ -374,7 +374,7 @@ theorem stalkSpecializes_stalkFunctor_map {F G : X.Presheaf C} (f : F ⟶ G) {x -- See https://github.com/leanprover-community/batteries/issues/365 for the simpNF issue. @[reassoc, elementwise, simp, nolint simpNF] theorem stalkSpecializes_stalkPushforward (f : X ⟶ Y) (F : X.Presheaf C) {x y : X} (h : x ⤳ y) : - (f _* F).stalkSpecializes (f.map_specializes h) ≫ F.stalkPushforward _ f x = + (f _* F).stalkSpecializes (f.hom.map_specializes h) ≫ F.stalkPushforward _ f x = F.stalkPushforward _ f y ≫ F.stalkSpecializes h := by change (_ : colimit _ ⟶ _) = (_ : colimit _ ⟶ _) ext; delta stalkPushforward diff --git a/Mathlib/Topology/Specialization.lean b/Mathlib/Topology/Specialization.lean index bd57a351804bdc..1ffdefb59b68c4 100644 --- a/Mathlib/Topology/Specialization.lean +++ b/Mathlib/Topology/Specialization.lean @@ -87,4 +87,4 @@ def homeoWithUpperSetTopologyorderIso (α : Type*) [TopologicalSpace α] [Alexan @[simps] def topToPreord : TopCat ⥤ Preord where obj X := Preord.of <| Specialization X - map := Specialization.map + map {X Y} f := Specialization.map (α := X) (β := Y) f.hom diff --git a/MathlibTest/CategoryTheory/CheckCompositions.lean b/MathlibTest/CategoryTheory/CheckCompositions.lean new file mode 100644 index 00000000000000..2598018866aed4 --- /dev/null +++ b/MathlibTest/CategoryTheory/CheckCompositions.lean @@ -0,0 +1,49 @@ +import Mathlib.CategoryTheory.Limits.Preserves.Limits +import Mathlib.Tactic.Recall + +universe v₁ v₂ v₃ v u₁ u₂ u₃ u + +open CategoryTheory Limits + +variable {J : Type u} [Category.{v} J] {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + {E : Type u₃} [Category.{v₃} E] (F : J ⥤ C) (G : C ⥤ D) (H : D ⥤ E) + +variable [HasColimitsOfShape J C] [HasColimitsOfShape J E] [PreservesColimit F (G ⋙ H)] + +/-- +info: In composition + colimit.ι ((F ⋙ G) ⋙ H) j ≫ (preservesColimitIso (G ⋙ H) F).inv +the source of + (preservesColimitIso (G ⋙ H) F).inv +is + colimit (F ⋙ G ⋙ H) +but should be + colimit ((F ⋙ G) ⋙ H) +-/ +#guard_msgs in +set_option linter.unusedTactic false in +example (j : J) : + colimit.ι ((F ⋙ G) ⋙ H) j ≫ (preservesColimitIso (G ⋙ H) F).inv = + H.map (G.map (colimit.ι F j)) := by + + -- We know which lemma we want to use, and it's even a simp lemma, but `rw` won't let us apply it + fail_if_success rw [ι_preservesColimitIso_inv] + fail_if_success rw [ι_preservesColimitIso_inv (G ⋙ H)] + fail_if_success simp only [ι_preservesColimitIso_inv] + + -- This would work: + -- erw [ι_preservesColimitIso_inv (G ⋙ H)] + + -- `check_compositions` checks if the two morphisms we're composing are composed by abusing defeq, + -- and indeed it tells us that we are abusing definitional associativity of composition of + -- functors here! + check_compositions + + -- In this case, we can "fix" this by reassociating in the goal, but usually at this point the + -- right thing to do is to back off and check how we ended up with a bad goal in the first place. + dsimp only [Functor.assoc] + + -- This would work now, but it is not needed, because simp works as well + -- rw [ι_preservesColimitIso_inv] + + simp diff --git a/MathlibTest/CategoryTheory/ConcreteCategory/Semigrp.lean b/MathlibTest/CategoryTheory/ConcreteCategory/Semigrp.lean new file mode 100644 index 00000000000000..2707c28a358882 --- /dev/null +++ b/MathlibTest/CategoryTheory/ConcreteCategory/Semigrp.lean @@ -0,0 +1,47 @@ +import Mathlib.Algebra.Category.Semigrp.Basic + +universe v u + +open CategoryTheory Semigrp + +set_option maxHeartbeats 10000 +set_option synthInstance.maxHeartbeats 2000 + +/- We test if all the coercions and `map_mul` lemmas trigger correctly. -/ + +example (X : Type u) [Semigroup X] : ⇑(𝟙 (of X)) = id := by simp + +example {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) : + ⇑(ofHom f) = ⇑f := by simp + +example {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) + (x : X) : (ofHom f) x = f x := by simp + +example {X Y Z : Semigrp} (f : X ⟶ Y) (g : Y ⟶ Z) : ⇑(f ≫ g) = ⇑g ∘ ⇑f := by simp + +example {X Y Z : Type u} [Semigroup X] [Semigroup Y] [Semigroup Z] + (f : X →ₙ* Y) (g : Y →ₙ* Z) : + ⇑(ofHom f ≫ ofHom g) = g ∘ f := by simp + +example {X Y : Type u} [Semigroup X] [Semigroup Y] {Z : Semigrp} + (f : X →ₙ* Y) (g : of Y ⟶ Z) : + ⇑(ofHom f ≫ g) = g ∘ f := by simp + +example {X Y : Semigrp} {Z : Type u} [Semigroup Z] (f : X ⟶ Y) (g : Y ⟶ of Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {Y Z : Semigrp} {X : Type u} [Semigroup X] (f : of X ⟶ Y) (g : Y ⟶ Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {X Y Z : Semigrp} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := by simp + +example {X Y : Semigrp} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by simp + +example {X Y : Semigrp} (e : X ≅ Y) (y : Y) : e.hom (e.inv y) = y := by simp + +example (X : Semigrp) : ⇑(𝟙 X) = id := by simp + +example {X : Type*} [Semigroup X] : ⇑(MulHom.id X) = id := by simp + +example {M N : Semigrp} (f : M ⟶ N) (x y : M) : f (x * y) = f x * f y := by + simp diff --git a/MathlibTest/oldObtain.lean b/MathlibTest/oldObtain.lean index 5a7bb3201701c8..46c0e3f7f44f86 100644 --- a/MathlibTest/oldObtain.lean +++ b/MathlibTest/oldObtain.lean @@ -20,7 +20,7 @@ theorem foo : True := by -- These cases are linted against. /-- -warning: Please remove stream-of-conciousness `obtain` syntax +warning: Please remove stream-of-consciousness `obtain` syntax note: this linter can be disabled with `set_option linter.oldObtain false` -/ #guard_msgs in @@ -30,7 +30,7 @@ theorem foo' : True := by trivial /-- -warning: Please remove stream-of-conciousness `obtain` syntax +warning: Please remove stream-of-consciousness `obtain` syntax note: this linter can be disabled with `set_option linter.oldObtain false` -/ #guard_msgs in diff --git a/MathlibTest/slow_simp.lean b/MathlibTest/slow_simp.lean index 124ec08fb86dc5..587f98f81ac999 100644 --- a/MathlibTest/slow_simp.lean +++ b/MathlibTest/slow_simp.lean @@ -51,16 +51,16 @@ instance : Category PointedSpace where end PointedSpace -set_option maxHeartbeats 40000 in +set_option maxHeartbeats 20000 in def PointedSpaceEquiv_inverse : Under (TopCat.of Unit) ⥤ PointedSpace where obj := fun X => { carrier := X.right base := X.hom () } map := fun f => - { map := f.right + { map := f.right.hom base := by have := f.w - replace this := DFunLike.congr_fun this () + replace this := CategoryTheory.congr_fun this () simp [-Under.w] at this simp exact this.symm } diff --git a/docs/1000.yaml b/docs/1000.yaml index ca2e502ee182b5..cbd2a725c1b8a1 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -1738,12 +1738,9 @@ Q1765521: Q1766814: title: Smn theorem - decls: - - Nat.Partrec.Code.smn - - Nat.Partrec'.part_iff + decl: Nat.Partrec.Code.smn authors: Mario Carneiro date: 2018 - comment: "This theorem is trivial when using Mathlib's computability definition, but this definition is equivalent to the usual one, as shown in `Nat.Partrec'.part_iff`." Q1785610: title: Poncelet's closure theorem diff --git a/lake-manifest.json b/lake-manifest.json index 4eb64458fbdf0b..f263b2cfc76dc0 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "1622a8693b31523c8f82db48e01b14c74bc1f155", + "rev": "e104724db91693a3a088c5fdd76cff4fac331739", "name": "plausible", "manifestFile": "lake-manifest.json", - "inputRev": "v4.16.0-rc1", + "inputRev": "main", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/LeanSearchClient", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f72319c9686788305a8ab059f3c4d8c724785c83", + "rev": "1a6613663c3eb08c401ce0fd1a408412f2c2321e", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -35,37 +35,37 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "07f60e90998dfd6592688a14cd67bd4e384b77b2", + "rev": "dafff53912eac07a1a983080df61df10f622fe25", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.50", + "inputRev": "main", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "79402ad9ab4be9a2286701a9880697e2351e4955", + "rev": "44dab9f36aa73d5f789629b55528c5a1501877a6", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "v4.16.0-rc1", + "inputRev": "master", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f0c584bcb14c5adfb53079781eeea75b26ebbd32", + "rev": "b1311119f5f7c79c818d3e06621cc81f4bb8973f", "name": "Qq", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0", + "inputRev": "master", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/batteries", "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4d5ece38d480adedff34d632a4bc0ebeb9dad9de", + "rev": "f53b5c1f9151dce5971165b4302b5b38a35be58d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "fin-find", diff --git a/lakefile.lean b/lakefile.lean index 306cb339189d54..88910f482d8d8d 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -13,7 +13,7 @@ require "leanprover-community" / "aesop" @ git "v4.16.0-rc1" require "leanprover-community" / "proofwidgets" @ git "v0.0.50" require "leanprover-community" / "importGraph" @ git "main" require "leanprover-community" / "LeanSearchClient" @ git "main" -require "leanprover-community" / "plausible" @ git "v4.16.0-rc1" +require "leanprover-community" / "plausible" @ git "main" /-! ## Options for building mathlib diff --git a/lean-toolchain b/lean-toolchain index 2ffc30ceba59b0..8b4f470c73e011 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.16.0-rc2 \ No newline at end of file +leanprover/lean4:v4.16.0 \ No newline at end of file diff --git a/scripts/nolints.json b/scripts/nolints.json index d52df230c49964..c7f6d49d1395e4 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -66,9 +66,6 @@ ["docBlame", "CancelDenoms.synthesizeUsingNormNum"], ["docBlame", "CategoryTheory.«term_⟶[_]_»"], ["docBlame", "ChangeOfRings.«term_⊗ₜ[_,_]_»"], - ["docBlame", "Combinator.I"], - ["docBlame", "Combinator.K"], - ["docBlame", "Combinator.S"], ["docBlame", "CompleteBooleanAlgebra.toCompleteAtomicBooleanAlgebra"], ["docBlame", "Computation.parallelRec"], ["docBlame", "Congr!.elabConfig"], @@ -166,7 +163,6 @@ ["docBlame", "LinearRecurrence.order"], ["docBlame", "MagmaCat.forget_obj_eq_coe"], ["docBlame", "MaximalSpectrum.asIdeal"], - ["docBlame", "MeasureTheory.«term_[_|_]»"], ["docBlame", "MeasureTheory.«term_→₁[_]_»"], ["docBlame", "MeasureTheory.«term_→₂[_]_»"], ["docBlame", "MeasureTheory.«term_≤[_]_»"], @@ -215,7 +211,6 @@ ["docBlame", "Nat.rfind"], ["docBlame", "Nat.rfindOpt"], ["docBlame", "Nat.rfindX"], - ["docBlame", "Nat.subInduction"], ["docBlame", "One.one"], ["docBlame", "OptionT.callCC"], ["docBlame", "OptionT.mkLabel"], @@ -275,7 +270,6 @@ ["docBlame", "Stream'.corec'"], ["docBlame", "Stream'.corecOn"], ["docBlame", "Stream'.unfolds"], - ["docBlame", "StrictWeakOrder.Equiv"], ["docBlame", "Submodule.«term_∙_»"], ["docBlame", "Traversable.foldMap"], ["docBlame", "Traversable.foldl"], @@ -646,5 +640,4 @@ ["docBlame", "Mathlib.Meta.NormNum.evalLT.core.intArm"], ["docBlame", "Mathlib.Meta.NormNum.evalLT.core.ratArm"], ["docBlame", "Mathlib.Meta.NormNum.evalMul.core.intArm"], - ["docBlame", "Mathlib.Meta.NormNum.evalMul.core.ratArm"], - ["unusedArguments", "Combinator.K"]] + ["docBlame", "Mathlib.Meta.NormNum.evalMul.core.ratArm"]] diff --git a/scripts/noshake.json b/scripts/noshake.json index 2732e7a608dbe0..c41910752e88bc 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -300,13 +300,14 @@ "Mathlib.Tactic.CategoryTheory.Coherence": ["Mathlib.CategoryTheory.Monoidal.Free.Coherence", "Mathlib.Tactic.CategoryTheory.MonoidalComp"], + "Mathlib.Tactic.CategoryTheory.CheckCompositions": + ["Mathlib.CategoryTheory.Category.Basic"], "Mathlib.Tactic.CategoryTheory.BicategoryCoherence": ["Mathlib.CategoryTheory.Bicategory.Coherence", "Mathlib.Tactic.CategoryTheory.BicategoricalComp"], + "Mathlib.Tactic.CC.MkProof": ["Mathlib.Tactic.CC.Lemmas"], "Mathlib.Tactic.CC.Datatypes": ["Mathlib.Lean.Meta.CongrTheorems", "Mathlib.Lean.Meta.Datatypes"], - "Mathlib.Tactic.CC.MkProof": - ["Mathlib.Tactic.CC.Lemmas"], "Mathlib.Tactic.ByContra": ["Batteries.Tactic.Init"], "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], "Mathlib.Tactic.Bound.Attribute": @@ -431,6 +432,8 @@ ["Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback", "Mathlib.CategoryTheory.Limits.Shapes.Pullbacks"], "Mathlib.CategoryTheory.Limits.IsLimit": ["Batteries.Tactic.Congr"], + "Mathlib.CategoryTheory.Functor.Basic": + ["Mathlib.Tactic.CategoryTheory.CheckCompositions"], "Mathlib.CategoryTheory.Bicategory.Functor.Oplax": ["Mathlib.Tactic.CategoryTheory.ToApp"], "Mathlib.CategoryTheory.Bicategory.Functor.Lax":