diff --git a/.github/build.in.yml b/.github/build.in.yml index 141b5316926ab..6a1d17a8cf495 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -127,16 +127,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 3bd02efbb94dd..87e3093f9b131 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -137,16 +137,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index b5f11a1ea1f03..5ae1b33476e94 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -144,16 +144,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 4da591010b8b7..79bfe1d0f3d0a 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -141,16 +141,6 @@ jobs: run: | lake build cache - - name: prune ProofWidgets .lake - run: | - lake build proofwidgets:release - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - - name: get cache id: get run: | diff --git a/.github/workflows/lean4checker.yml b/.github/workflows/lean4checker.yml index 066648826282a..c521c9164c994 100644 --- a/.github/workflows/lean4checker.yml +++ b/.github/workflows/lean4checker.yml @@ -1,4 +1,3 @@ -# https://chat.openai.com/share/26102e95-ac9c-4d03-a000-73e4f6cee8cd name: lean4checker Workflow on: @@ -6,14 +5,19 @@ on: - cron: '0 0 * * *' # Runs at 00:00 UTC every day workflow_dispatch: +env: + DEFAULT_BRANCH: master + TAG_PATTERN: '^nightly-testing-[0-9]{4}-[0-9]{2}-[0-9]{2}$' + jobs: check-lean4checker: - runs-on: pr + runs-on: ubuntu-latest + strategy: + matrix: + branch_type: [master, nightly] steps: - - - name: cleanup + - name: Cleanup run: | - find . -name . -o -prune -exec rm -rf -- {} + # Delete all but the 5 most recent toolchains. # Make sure to delete both the `~/.elan/toolchains/X` directory and the `~/.elan/update-hashes/X` file. # Skip symbolic links (`-type d`), the current directory (`! -name .`), and `nightly` and `stable`. @@ -23,21 +27,24 @@ jobs: : # Do nothing on failure, but suppress errors fi - # The Hoskinson runners may not have jq installed, so do that now. - - name: 'Setup jq' - uses: dcarbone/install-jq-action@v2.1.0 - - - name: install elan + - name: Install elan run: | set -o pipefail curl -sSfL https://github.com/leanprover/elan/releases/download/v3.1.1/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: Checkout master branch + - name: Fetch latest tags (if nightly) + if: matrix.branch_type == 'nightly' + run: | + git fetch --tags + LATEST_TAG=$(git tag | grep -E "${{ env.TAG_PATTERN }}" | sort -r | head -n 1) + echo "LATEST_TAG=${LATEST_TAG}" >> "$GITHUB_ENV" + + - name: Checkout branch or tag uses: actions/checkout@v4 with: - ref: 'master' + ref: ${{ matrix.branch_type == 'master' && env.DEFAULT_BRANCH || env.LATEST_TAG }} - name: If using a lean-pr-release toolchain, uninstall run: | @@ -46,7 +53,7 @@ jobs: elan toolchain uninstall "$(cat lean-toolchain)" fi - - name: print lean and lake versions + - name: Print Lean and Lake versions run: | lean --version lake --version @@ -55,16 +62,6 @@ jobs: run: | lake exe cache get - - name: prune ProofWidgets .lake - run: | - # The ProofWidgets release contains not just the `.js` (which we need in order to build) - # but also `.oleans`, which may have been built with the wrong toolchain. - # This removes them. - # See discussion at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/nightly-testing/near/411225235 - rm -rf .lake/packages/proofwidgets/.lake/build/lib - rm -rf .lake/packages/proofwidgets/.lake/build/ir - lake build ProofWidgets - - name: Check environments using lean4checker id: lean4checker run: | @@ -100,7 +97,7 @@ jobs: type: 'stream' topic: 'lean4checker' content: | - ✅ lean4checker [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} + ✅ lean4checker [succeeded](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) - name: Post failure message on Zulip if: failure() @@ -113,7 +110,7 @@ jobs: type: 'stream' topic: 'lean4checker failure' content: | - ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} + ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) continue-on-error: true - name: Post failure message on Zulip main topic @@ -127,5 +124,5 @@ jobs: type: 'stream' topic: 'lean4checker' content: | - ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} + ❌ lean4checker [failed](https://github.com/${{ github.repository }}/actions/runs/${{ github.run_id }}) on ${{ github.sha }} (branch: ${{ matrix.branch_type == 'master' && 'master' || env.LATEST_TAG }}) continue-on-error: true diff --git a/Archive/Imo/Imo1964Q1.lean b/Archive/Imo/Imo1964Q1.lean index dda0e7f1e11cb..6a3e8cba29dfd 100644 --- a/Archive/Imo/Imo1964Q1.lean +++ b/Archive/Imo/Imo1964Q1.lean @@ -5,25 +5,17 @@ Authors: Kevin Buzzard -/ import Mathlib.Tactic.IntervalCases import Mathlib.Data.Nat.ModEq -import Mathlib.Tactic.Ring /-! # IMO 1964 Q1 (a) Find all positive integers $n$ for which $2^n-1$ is divisible by $7$. - (b) Prove that there is no positive integer $n$ for which $2^n+1$ is divisible by $7$. -We define a predicate for the solutions in (a), and prove that it is the set of positive -integers which are a multiple of 3. --/ - - -/-! -## Intermediate lemmas +For (a), we find that the order of $2$ mod $7$ is $3$. Therefore for (b), it suffices to check +$n = 0, 1, 2$. -/ - open Nat namespace Imo1964Q1 @@ -33,17 +25,13 @@ theorem two_pow_mod_seven (n : ℕ) : 2 ^ n ≡ 2 ^ (n % 3) [MOD 7] := calc 2 ^ n = 2 ^ (3 * (n / 3) + t) := by rw [Nat.div_add_mod] _ = (2 ^ 3) ^ (n / 3) * 2 ^ t := by rw [pow_add, pow_mul] _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; decide - _ = 2 ^ t := by ring - -/-! -## The question --/ + _ = 2 ^ t := by rw [one_pow, one_mul] +end Imo1964Q1 -def ProblemPredicate (n : ℕ) : Prop := - 7 ∣ 2 ^ n - 1 +open Imo1964Q1 -theorem imo1964_q1a (n : ℕ) (_ : 0 < n) : ProblemPredicate n ↔ 3 ∣ n := by +theorem imo1964_q1a (n : ℕ) (_ : 0 < n) : 7 ∣ 2 ^ n - 1 ↔ 3 ∣ n := by let t := n % 3 have : t < 3 := Nat.mod_lt _ (by decide) calc 7 ∣ 2 ^ n - 1 ↔ 2 ^ n ≡ 1 [MOD 7] := by @@ -53,15 +41,11 @@ theorem imo1964_q1a (n : ℕ) (_ : 0 < n) : ProblemPredicate n ↔ 3 ∣ n := by _ ↔ t = 0 := by interval_cases t <;> decide _ ↔ 3 ∣ n := by rw [dvd_iff_mod_eq_zero] -end Imo1964Q1 - -open Imo1964Q1 - theorem imo1964_q1b (n : ℕ) : ¬7 ∣ 2 ^ n + 1 := by intro h let t := n % 3 have : t < 3 := Nat.mod_lt _ (by decide) have H : 2 ^ t + 1 ≡ 0 [MOD 7] := calc - 2 ^ t + 1 ≡ 2 ^ n + 1 [MOD 7 ] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm + 2 ^ t + 1 ≡ 2 ^ n + 1 [MOD 7] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm _ ≡ 0 [MOD 7] := h.modEq_zero_nat interval_cases t <;> contradiction diff --git a/Archive/Imo/Imo2019Q4.lean b/Archive/Imo/Imo2019Q4.lean index fe10780d6586f..1bcfbe0c77878 100644 --- a/Archive/Imo/Imo2019Q4.lean +++ b/Archive/Imo/Imo2019Q4.lean @@ -30,7 +30,7 @@ open scoped Nat open Nat hiding zero_le Prime -open Finset multiplicity +open Finset namespace Imo2019Q4 diff --git a/Archive/Wiedijk100Theorems/PerfectNumbers.lean b/Archive/Wiedijk100Theorems/PerfectNumbers.lean index a97ff90b5c744..d9d5bb09b3335 100644 --- a/Archive/Wiedijk100Theorems/PerfectNumbers.lean +++ b/Archive/Wiedijk100Theorems/PerfectNumbers.lean @@ -55,7 +55,7 @@ theorem even_two_pow_mul_mersenne_of_prime (k : ℕ) (pr : (mersenne (k + 1)).Pr Even (2 ^ k * mersenne (k + 1)) := by simp [ne_zero_of_prime_mersenne k pr, parity_simps] theorem eq_two_pow_mul_odd {n : ℕ} (hpos : 0 < n) : ∃ k m : ℕ, n = 2 ^ k * m ∧ ¬Even m := by - have h := Nat.multiplicity_finite_iff.2 ⟨Nat.prime_two.ne_one, hpos⟩ + have h := Nat.finiteMultiplicity_iff.2 ⟨Nat.prime_two.ne_one, hpos⟩ cases' pow_multiplicity_dvd 2 n with m hm use multiplicity 2 n, m refine ⟨hm, ?_⟩ diff --git a/Cache/Requests.lean b/Cache/Requests.lean index 295d65cd96693..38f3db11e9406 100644 --- a/Cache/Requests.lean +++ b/Cache/Requests.lean @@ -161,14 +161,27 @@ into the `lean-toolchain` file at the root directory of your project" /-- Fetches the ProofWidgets cloud release and prunes non-JS files. -/ def getProofWidgets (buildDir : FilePath) : IO Unit := do - if (← buildDir.pathExists) then return - -- Unpack the ProofWidgets cloud release (for its `.js` files) + if (← buildDir.pathExists) then + -- Check if the ProofWidgets build is out-of-date via `lake`. + -- This is done through Lake as cache has no simple heuristic + -- to determine whether the ProofWidgets JS is out-of-date. + let exitCode ← (← IO.Process.spawn {cmd := "lake", args := #["-q", "build", "--no-build", "proofwidgets:release"]}).wait + if exitCode == 0 then -- up-to-date + return + else if exitCode == 3 then -- needs fetch (`--no-build` triggered) + pure () + else + throw <| IO.userError s!"Failed to validate ProofWidgets cloud release: lake failed with error code {exitCode}" + -- Download and unpack the ProofWidgets cloud release (for its `.js` files) let exitCode ← (← IO.Process.spawn {cmd := "lake", args := #["-q", "build", "proofwidgets:release"]}).wait if exitCode != 0 then throw <| IO.userError s!"Failed to fetch ProofWidgets cloud release: lake failed with error code {exitCode}" - -- prune non-js ProofWidgets files (e.g., `olean`, `.c`) - IO.FS.removeDirAll (buildDir / "lib") - IO.FS.removeDirAll (buildDir / "ir") + -- Prune non-JS ProofWidgets files (e.g., `olean`, `.c`) + try + IO.FS.removeDirAll (buildDir / "lib") + IO.FS.removeDirAll (buildDir / "ir") + catch e => + throw <| IO.userError s!"Failed to prune ProofWidgets cloud release: {e}" /-- Downloads missing files, and unpacks files. -/ def getFiles (hashMap : IO.HashMap) (forceDownload forceUnpack parallel decompress : Bool) : diff --git a/Counterexamples/Pseudoelement.lean b/Counterexamples/Pseudoelement.lean index af58937a6308c..46052cc97e77a 100644 --- a/Counterexamples/Pseudoelement.lean +++ b/Counterexamples/Pseudoelement.lean @@ -74,26 +74,24 @@ theorem x_not_pseudo_eq : ¬PseudoEqual _ x y := by replace h := ModuleCat.eq_range_of_pseudoequal h dsimp [x, y] at h let φ := biprod.lift (𝟙 (of ℤ ℚ)) (2 • 𝟙 (of ℤ ℚ)) - have mem_range := mem_range_self φ (1 : ℚ) + have mem_range := mem_range_self φ.hom (1 : ℚ) rw [h] at mem_range obtain ⟨a, ha⟩ := mem_range - erw [← ModuleCat.id_apply (φ (1 : ℚ)), ← biprod.total, ← LinearMap.comp_apply, ← comp_def, - Preadditive.comp_add] at ha + rw [← ModuleCat.id_apply _ (φ (1 : ℚ)), ← biprod.total, ← LinearMap.comp_apply, + ← ModuleCat.hom_comp, Preadditive.comp_add] at ha let π₁ := (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) have ha₁ := congr_arg π₁ ha - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at ha₁ + rw [← ModuleCat.comp_apply, ← ModuleCat.comp_apply] at ha₁ simp only [π₁, φ, BinaryBiproduct.bicone_fst, biprod.lift_fst, CategoryTheory.id_apply, biprod.lift_fst_assoc, Category.id_comp, biprod.lift_snd_assoc, Linear.smul_comp, Preadditive.add_comp, BinaryBicone.inl_fst, BinaryBicone.inr_fst, smul_zero, add_zero] at ha₁ let π₂ := (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) have ha₂ := congr_arg π₂ ha - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at ha₂ + rw [← ModuleCat.comp_apply, ← ModuleCat.comp_apply] at ha₂ simp only [π₁, π₂, φ, BinaryBiproduct.bicone_snd, biprod.lift_snd, CategoryTheory.id_apply, biprod.lift_fst_assoc, Category.id_comp, biprod.lift_snd_assoc, Linear.smul_comp, Preadditive.add_comp, BinaryBicone.inl_snd, BinaryBicone.inr_snd, zero_add, two_smul] at ha₂ - erw [add_apply, CategoryTheory.id_apply] at ha₂ + erw [add_apply, ModuleCat.id_apply] at ha₂ subst ha₁ simp only [self_eq_add_right] at ha₂ exact one_ne_zero' ℚ ha₂ @@ -104,12 +102,12 @@ open scoped Pseudoelement /-- `biprod.fst ⟦x⟧ = biprod.fst ⟦y⟧`. -/ theorem fst_mk'_x_eq_fst_mk'_y : - (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦x⟧ = (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦y⟧ := + pseudoApply biprod.fst ⟦x⟧ = pseudoApply biprod.fst ⟦y⟧ := Quotient.eq.2 fst_x_pseudo_eq_fst_y /-- `biprod.snd ⟦x⟧ = biprod.snd ⟦y⟧`. -/ theorem snd_mk'_x_eq_snd_mk'_y : - (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦x⟧ = (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) ⟦y⟧ := + pseudoApply biprod.snd ⟦x⟧ = pseudoApply biprod.snd ⟦y⟧ := Quotient.eq.2 snd_x_pseudo_eq_snd_y -- Porting note: needs explicit type ascription `: Quotient <| Pseudoelement.setoid _` @@ -125,8 +123,8 @@ theorem exist_ne_and_fst_eq_fst_and_snd_eq_snd : ∃ x y, -- Porting note: removed type ascription `: of ℤ ℚ ⊞ of ℤ ℚ`, it gave an error about -- `Type` not having zero morphisms. jmc: I don't understand where the error came from x ≠ y ∧ - (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) x = (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) y ∧ - (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) x = (biprod.snd : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) y := + pseudoApply (biprod.fst : of ℤ ℚ ⊞ of ℤ ℚ ⟶ _) x = pseudoApply biprod.fst y ∧ + pseudoApply biprod.snd x = pseudoApply biprod.snd y := ⟨⟦x⟧, ⟦y⟧, mk'_x_ne_mk'_y, fst_mk'_x_eq_fst_mk'_y, snd_mk'_x_eq_snd_mk'_y⟩ end diff --git a/Mathlib.lean b/Mathlib.lean index f1a2458fab371..4ec2f1daeaf33 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -554,6 +554,7 @@ import Mathlib.Algebra.Module.Presentation.Differentials import Mathlib.Algebra.Module.Presentation.DirectSum import Mathlib.Algebra.Module.Presentation.Finite import Mathlib.Algebra.Module.Presentation.Free +import Mathlib.Algebra.Module.Presentation.RestrictScalars import Mathlib.Algebra.Module.Presentation.Tautological import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Projective @@ -937,6 +938,7 @@ import Mathlib.Algebra.Tropical.Basic import Mathlib.Algebra.Tropical.BigOperators import Mathlib.Algebra.Tropical.Lattice import Mathlib.Algebra.Vertex.HVertexOperator +import Mathlib.Algebra.Vertex.VertexOperator import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.AffineSpace import Mathlib.AlgebraicGeometry.Cover.MorphismProperty @@ -1054,6 +1056,7 @@ import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal import Mathlib.AlgebraicTopology.SimplicialSet.Nerve @@ -2005,6 +2008,7 @@ import Mathlib.CategoryTheory.MorphismProperty.IsInvertedBy import Mathlib.CategoryTheory.MorphismProperty.Limits import Mathlib.CategoryTheory.MorphismProperty.OverAdjunction import Mathlib.CategoryTheory.MorphismProperty.Representable +import Mathlib.CategoryTheory.MorphismProperty.Retract import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans import Mathlib.CategoryTheory.Noetherian @@ -2043,6 +2047,7 @@ import Mathlib.CategoryTheory.Products.Unitor import Mathlib.CategoryTheory.Quotient import Mathlib.CategoryTheory.Quotient.Linear import Mathlib.CategoryTheory.Quotient.Preadditive +import Mathlib.CategoryTheory.Retract import Mathlib.CategoryTheory.Shift.Basic import Mathlib.CategoryTheory.Shift.CommShift import Mathlib.CategoryTheory.Shift.Induced @@ -2482,6 +2487,7 @@ import Mathlib.Data.Finsupp.Indicator import Mathlib.Data.Finsupp.Interval import Mathlib.Data.Finsupp.Lex import Mathlib.Data.Finsupp.MonomialOrder +import Mathlib.Data.Finsupp.MonomialOrder.DegLex import Mathlib.Data.Finsupp.Multiset import Mathlib.Data.Finsupp.NeLocus import Mathlib.Data.Finsupp.Notation @@ -3332,6 +3338,7 @@ import Mathlib.LinearAlgebra.Eigenspace.Zero import Mathlib.LinearAlgebra.ExteriorAlgebra.Basic import Mathlib.LinearAlgebra.ExteriorAlgebra.Grading import Mathlib.LinearAlgebra.ExteriorAlgebra.OfAlternating +import Mathlib.LinearAlgebra.ExteriorPower.Basic import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.LinearAlgebra.FiniteDimensional.Defs import Mathlib.LinearAlgebra.FiniteSpan @@ -3355,6 +3362,7 @@ import Mathlib.LinearAlgebra.FreeModule.PID import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.LinearAlgebra.FreeProduct.Basic import Mathlib.LinearAlgebra.GeneralLinearGroup +import Mathlib.LinearAlgebra.Goursat import Mathlib.LinearAlgebra.InvariantBasisNumber import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.LinearAlgebra.JordanChevalley @@ -3876,6 +3884,7 @@ import Mathlib.NumberTheory.NumberField.Discriminant.Basic import Mathlib.NumberTheory.NumberField.Discriminant.Defs import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.NumberTheory.NumberField.EquivReindex +import Mathlib.NumberTheory.NumberField.FinitePlaces import Mathlib.NumberTheory.NumberField.FractionalIdeal import Mathlib.NumberTheory.NumberField.House import Mathlib.NumberTheory.NumberField.Norm @@ -3988,6 +3997,7 @@ import Mathlib.Order.Extension.Well import Mathlib.Order.Filter.AtTopBot import Mathlib.Order.Filter.AtTopBot.Archimedean import Mathlib.Order.Filter.AtTopBot.BigOperators +import Mathlib.Order.Filter.AtTopBot.CountablyGenerated import Mathlib.Order.Filter.AtTopBot.Field import Mathlib.Order.Filter.AtTopBot.Floor import Mathlib.Order.Filter.AtTopBot.Group @@ -4001,6 +4011,7 @@ import Mathlib.Order.Filter.Cocardinal import Mathlib.Order.Filter.Cofinite import Mathlib.Order.Filter.CountableInter import Mathlib.Order.Filter.CountableSeparatingOn +import Mathlib.Order.Filter.CountablyGenerated import Mathlib.Order.Filter.Curry import Mathlib.Order.Filter.Defs import Mathlib.Order.Filter.ENNReal @@ -4154,7 +4165,9 @@ import Mathlib.Probability.Independence.Kernel import Mathlib.Probability.Independence.ZeroOne import Mathlib.Probability.Integration import Mathlib.Probability.Kernel.Basic -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic +import Mathlib.Probability.Kernel.Composition.IntegralCompProd +import Mathlib.Probability.Kernel.Composition.MeasureCompProd import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Defs @@ -4167,10 +4180,8 @@ import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes import Mathlib.Probability.Kernel.Disintegration.StandardBorel import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.Integral -import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance import Mathlib.Probability.Kernel.MeasurableIntegral -import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.RadonNikodym import Mathlib.Probability.Kernel.WithDensity import Mathlib.Probability.Martingale.Basic @@ -4304,7 +4315,6 @@ import Mathlib.RingTheory.Finiteness.Projective import Mathlib.RingTheory.Finiteness.Subalgebra import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.Fintype -import Mathlib.RingTheory.Flat.Algebra import Mathlib.RingTheory.Flat.Basic import Mathlib.RingTheory.Flat.CategoryTheory import Mathlib.RingTheory.Flat.Equalizer @@ -4375,8 +4385,9 @@ import Mathlib.RingTheory.IntegralDomain import Mathlib.RingTheory.IsAdjoinRoot import Mathlib.RingTheory.IsPrimary import Mathlib.RingTheory.IsTensorProduct -import Mathlib.RingTheory.Jacobson -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal +import Mathlib.RingTheory.Jacobson.Polynomial +import Mathlib.RingTheory.Jacobson.Ring import Mathlib.RingTheory.Kaehler.Basic import Mathlib.RingTheory.Kaehler.CotangentComplex import Mathlib.RingTheory.Kaehler.Polynomial @@ -4430,10 +4441,13 @@ import Mathlib.RingTheory.MaximalSpectrum import Mathlib.RingTheory.Multiplicity import Mathlib.RingTheory.MvPolynomial import Mathlib.RingTheory.MvPolynomial.Basic +import Mathlib.RingTheory.MvPolynomial.EulerIdentity import Mathlib.RingTheory.MvPolynomial.FreeCommRing +import Mathlib.RingTheory.MvPolynomial.Groebner import Mathlib.RingTheory.MvPolynomial.Homogeneous import Mathlib.RingTheory.MvPolynomial.Ideal import Mathlib.RingTheory.MvPolynomial.Localization +import Mathlib.RingTheory.MvPolynomial.MonomialOrder import Mathlib.RingTheory.MvPolynomial.Symmetric.Defs import Mathlib.RingTheory.MvPolynomial.Symmetric.FundamentalTheorem import Mathlib.RingTheory.MvPolynomial.Symmetric.NewtonIdentities @@ -4524,6 +4538,7 @@ import Mathlib.RingTheory.Regular.RegularSequence import Mathlib.RingTheory.RingHom.Finite import Mathlib.RingTheory.RingHom.FinitePresentation import Mathlib.RingTheory.RingHom.FiniteType +import Mathlib.RingTheory.RingHom.Flat import Mathlib.RingTheory.RingHom.Injective import Mathlib.RingTheory.RingHom.Integral import Mathlib.RingTheory.RingHom.Locally @@ -4946,6 +4961,7 @@ import Mathlib.Topology.Algebra.Group.SubmonoidClosure import Mathlib.Topology.Algebra.Group.TopologicalAbelianization import Mathlib.Topology.Algebra.GroupCompletion import Mathlib.Topology.Algebra.GroupWithZero +import Mathlib.Topology.Algebra.Indicator import Mathlib.Topology.Algebra.InfiniteSum.Basic import Mathlib.Topology.Algebra.InfiniteSum.Constructions import Mathlib.Topology.Algebra.InfiniteSum.Defs @@ -5353,6 +5369,7 @@ import Mathlib.Topology.UniformSpace.CompactConvergence import Mathlib.Topology.UniformSpace.CompareReals import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Completion +import Mathlib.Topology.UniformSpace.Dini import Mathlib.Topology.UniformSpace.Equicontinuity import Mathlib.Topology.UniformSpace.Equiv import Mathlib.Topology.UniformSpace.HeineCantor diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 7c503761b174e..24dddc1deceec 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -93,24 +93,138 @@ theorem one_le {P : Submodule R A} : (1 : Submodule R A) ≤ P ↔ (1 : A) ∈ P -- Porting note: simpa no longer closes refl goals, so added `SetLike.mem_coe` simp only [one_eq_span, span_le, Set.singleton_subset_iff, SetLike.mem_coe] +variable {M : Type*} [AddCommMonoid M] [Module R M] [Module A M] [IsScalarTower R A M] + +instance : SMul (Submodule R A) (Submodule R M) where + smul A' M' := + { __ := A'.toAddSubmonoid • M'.toAddSubmonoid + smul_mem' := fun r m hm ↦ AddSubmonoid.smul_induction_on hm + (fun a ha m hm ↦ by rw [← smul_assoc]; exact AddSubmonoid.smul_mem_smul (A'.smul_mem r ha) hm) + fun m₁ m₂ h₁ h₂ ↦ by rw [smul_add]; exact (A'.1 • M'.1).add_mem h₁ h₂ } + +section + +variable {I J : Submodule R A} {N P : Submodule R M} + +theorem smul_toAddSubmonoid : (I • N).toAddSubmonoid = I.toAddSubmonoid • N.toAddSubmonoid := rfl + +theorem smul_mem_smul {r} {n} (hr : r ∈ I) (hn : n ∈ N) : r • n ∈ I • N := + AddSubmonoid.smul_mem_smul hr hn + +theorem smul_le : I • N ≤ P ↔ ∀ r ∈ I, ∀ n ∈ N, r • n ∈ P := + AddSubmonoid.smul_le + +@[simp, norm_cast] +lemma coe_set_smul : (I : Set A) • N = I • N := + set_smul_eq_of_le _ _ _ + (fun _ _ hr hx ↦ smul_mem_smul hr hx) + (smul_le.mpr fun _ hr _ hx ↦ mem_set_smul_of_mem_mem hr hx) + +@[elab_as_elim] +theorem smul_induction_on {p : M → Prop} {x} (H : x ∈ I • N) (smul : ∀ r ∈ I, ∀ n ∈ N, p (r • n)) + (add : ∀ x y, p x → p y → p (x + y)) : p x := + AddSubmonoid.smul_induction_on H smul add + +/-- Dependent version of `Submodule.smul_induction_on`. -/ +@[elab_as_elim] +theorem smul_induction_on' {x : M} (hx : x ∈ I • N) {p : ∀ x, x ∈ I • N → Prop} + (smul : ∀ (r : A) (hr : r ∈ I) (n : M) (hn : n ∈ N), p (r • n) (smul_mem_smul hr hn)) + (add : ∀ x hx y hy, p x hx → p y hy → p (x + y) (add_mem ‹_› ‹_›)) : p x hx := by + refine Exists.elim ?_ fun (h : x ∈ I • N) (H : p x h) ↦ H + exact smul_induction_on hx (fun a ha x hx ↦ ⟨_, smul _ ha _ hx⟩) + fun x y ⟨_, hx⟩ ⟨_, hy⟩ ↦ ⟨_, add _ _ _ _ hx hy⟩ + +theorem smul_mono (hij : I ≤ J) (hnp : N ≤ P) : I • N ≤ J • P := + AddSubmonoid.smul_le_smul hij hnp + +theorem smul_mono_left (h : I ≤ J) : I • N ≤ J • N := + smul_mono h le_rfl + +instance : CovariantClass (Submodule R A) (Submodule R M) HSMul.hSMul LE.le := + ⟨fun _ _ => smul_mono le_rfl⟩ + +@[deprecated smul_mono_right (since := "2024-03-31")] +protected theorem smul_mono_right (h : N ≤ P) : I • N ≤ I • P := + _root_.smul_mono_right I h + +variable (I J N P) + +@[simp] +theorem smul_bot : I • (⊥ : Submodule R M) = ⊥ := + toAddSubmonoid_injective <| AddSubmonoid.addSubmonoid_smul_bot _ + +@[simp] +theorem bot_smul : (⊥ : Submodule R A) • N = ⊥ := + le_bot_iff.mp <| smul_le.mpr <| by rintro _ rfl _ _; rw [zero_smul]; exact zero_mem _ + +theorem smul_sup : I • (N ⊔ P) = I • N ⊔ I • P := + toAddSubmonoid_injective <| by + simp only [smul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.addSubmonoid_smul_sup] + +theorem sup_smul : (I ⊔ J) • N = I • N ⊔ J • N := + le_antisymm (smul_le.mpr fun mn hmn p hp ↦ by + obtain ⟨m, hm, n, hn, rfl⟩ := mem_sup.mp hmn + rw [add_smul]; exact add_mem_sup (smul_mem_smul hm hp) <| smul_mem_smul hn hp) + (sup_le (smul_mono_left le_sup_left) <| smul_mono_left le_sup_right) + +protected theorem smul_assoc {B} [Semiring B] [Module R B] [Module A B] [Module B M] + [IsScalarTower R A B] [IsScalarTower R B M] [IsScalarTower A B M] + (I : Submodule R A) (J : Submodule R B) (N : Submodule R M) : + (I • J) • N = I • J • N := + le_antisymm + (smul_le.2 fun _ hrsij t htn ↦ smul_induction_on hrsij + (fun r hr s hs ↦ smul_assoc r s t ▸ smul_mem_smul hr (smul_mem_smul hs htn)) + fun x y ↦ (add_smul x y t).symm ▸ add_mem) + (smul_le.2 fun r hr _ hsn ↦ smul_induction_on hsn + (fun j hj n hn ↦ (smul_assoc r j n).symm ▸ smul_mem_smul (smul_mem_smul hr hj) hn) + fun m₁ m₂ ↦ (smul_add r m₁ m₂) ▸ add_mem) + +@[deprecated smul_inf_le (since := "2024-03-31")] +protected theorem smul_inf_le (M₁ M₂ : Submodule R M) : + I • (M₁ ⊓ M₂) ≤ I • M₁ ⊓ I • M₂ := smul_inf_le _ _ _ + +theorem smul_iSup {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : + I • (⨆ i, t i)= ⨆ i, I • t i := + toAddSubmonoid_injective <| by + simp only [smul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.smul_iSup] + +theorem iSup_smul {ι : Sort*} {t : ι → Submodule R A} {N : Submodule R M} : + (⨆ i, t i) • N = ⨆ i, t i • N := + le_antisymm (smul_le.mpr fun t ht s hs ↦ iSup_induction _ (C := (· • s ∈ _)) ht + (fun i t ht ↦ mem_iSup_of_mem i <| smul_mem_smul ht hs) + (by simp_rw [zero_smul]; apply zero_mem) fun x y ↦ by simp_rw [add_smul]; apply add_mem) + (iSup_le fun i ↦ Submodule.smul_mono_left <| le_iSup _ i) + +@[deprecated smul_iInf_le (since := "2024-03-31")] +protected theorem smul_iInf_le {ι : Sort*} {I : Submodule R A} {t : ι → Submodule R M} : + I • iInf t ≤ ⨅ i, I • t i := + smul_iInf_le + +protected theorem one_smul : (1 : Submodule R A) • N = N := by + refine le_antisymm (smul_le.mpr fun r hr m hm ↦ ?_) fun m hm ↦ ?_ + · obtain ⟨r, rfl⟩ := hr + rw [LinearMap.toSpanSingleton_apply, smul_one_smul]; exact N.smul_mem r hm + · rw [← one_smul A m]; exact smul_mem_smul (one_le.mp le_rfl) hm + +theorem smul_subset_smul : (↑I : Set A) • (↑N : Set M) ⊆ (↑(I • N) : Set M) := + AddSubmonoid.smul_subset_smul + +end + variable [IsScalarTower R A A] /-- Multiplication of sub-R-modules of an R-module A that is also a semiring. The submodule `M * N` consists of finite sums of elements `m * n` for `m ∈ M` and `n ∈ N`. -/ instance mul : Mul (Submodule R A) where - mul M N := - { __ := M.toAddSubmonoid * N.toAddSubmonoid - smul_mem' := fun r a ha ↦ AddSubmonoid.mul_induction_on ha - (fun m h n h' ↦ by rw [← smul_mul_assoc]; exact AddSubmonoid.mul_mem_mul (M.smul_mem r h) h') - fun a₁ a₂ h₁ h₂ ↦ by rw [smul_add]; apply (M.1 * N.1).add_mem h₁ h₂ } + mul := (· • ·) variable (S T : Set A) {M N P Q : Submodule R A} {m n : A} theorem mul_mem_mul (hm : m ∈ M) (hn : n ∈ N) : m * n ∈ M * N := - AddSubmonoid.mul_mem_mul hm hn + smul_mem_smul hm hn theorem mul_le : M * N ≤ P ↔ ∀ m ∈ M, ∀ n ∈ N, m * n ∈ P := - AddSubmonoid.mul_le + smul_le theorem mul_toAddSubmonoid (M N : Submodule R A) : (M * N).toAddSubmonoid = M.toAddSubmonoid * N.toAddSubmonoid := rfl @@ -118,46 +232,40 @@ theorem mul_toAddSubmonoid (M N : Submodule R A) : @[elab_as_elim] protected theorem mul_induction_on {C : A → Prop} {r : A} (hr : r ∈ M * N) (hm : ∀ m ∈ M, ∀ n ∈ N, C (m * n)) (ha : ∀ x y, C x → C y → C (x + y)) : C r := - AddSubmonoid.mul_induction_on hr hm ha + smul_induction_on hr hm ha /-- A dependent version of `mul_induction_on`. -/ @[elab_as_elim] protected theorem mul_induction_on' {C : ∀ r, r ∈ M * N → Prop} (mem_mul_mem : ∀ m (hm : m ∈ M) n (hn : n ∈ N), C (m * n) (mul_mem_mul hm hn)) (add : ∀ x hx y hy, C x hx → C y hy → C (x + y) (add_mem hx hy)) {r : A} (hr : r ∈ M * N) : - C r hr := by - refine Exists.elim ?_ fun (hr : r ∈ M * N) (hc : C r hr) ↦ hc - exact Submodule.mul_induction_on hr - (fun x hx y hy ↦ ⟨_, mem_mul_mem _ hx _ hy⟩) - fun x y ⟨_, hx⟩ ⟨_, hy⟩ ↦ ⟨_, add _ _ _ _ hx hy⟩ + C r hr := + smul_induction_on' hr mem_mul_mem add variable (M) @[simp] theorem mul_bot : M * ⊥ = ⊥ := - toAddSubmonoid_injective (AddSubmonoid.mul_bot _) + smul_bot _ @[simp] theorem bot_mul : ⊥ * M = ⊥ := - toAddSubmonoid_injective (AddSubmonoid.bot_mul _) + bot_smul _ -protected theorem one_mul : (1 : Submodule R A) * M = M := by - refine le_antisymm (mul_le.mpr fun r hr m hm ↦ ?_) fun m hm ↦ ?_ - · obtain ⟨r, rfl⟩ := hr - rw [LinearMap.toSpanSingleton_apply, smul_one_mul]; exact M.smul_mem r hm - · rw [← one_mul m]; exact mul_mem_mul (one_le.mp le_rfl) hm +protected theorem one_mul : (1 : Submodule R A) * M = M := + Submodule.one_smul _ variable {M} @[mono] theorem mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := - AddSubmonoid.mul_le_mul hmp hnq + smul_mono hmp hnq theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := - AddSubmonoid.mul_le_mul_left h + smul_mono_left h theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := - AddSubmonoid.mul_le_mul_right h + smul_mono_right _ h theorem mul_comm_of_commute (h : ∀ m ∈ M, ∀ n ∈ N, Commute m n) : M * N = N * M := toAddSubmonoid_injective <| AddSubmonoid.mul_comm_of_commute h @@ -165,15 +273,13 @@ theorem mul_comm_of_commute (h : ∀ m ∈ M, ∀ n ∈ N, Commute m n) : M * N variable (M N P) theorem mul_sup : M * (N ⊔ P) = M * N ⊔ M * P := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.mul_sup] + smul_sup _ _ _ theorem sup_mul : (M ⊔ N) * P = M * P ⊔ N * P := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.sup_mul] + sup_smul _ _ _ theorem mul_subset_mul : (↑M : Set A) * (↑N : Set A) ⊆ (↑(M * N) : Set A) := - AddSubmonoid.mul_subset_mul + smul_subset_smul _ _ lemma restrictScalars_mul {A B C} [Semiring A] [Semiring B] [Semiring C] [SMul A B] [Module A C] [Module B C] [IsScalarTower A C C] [IsScalarTower B C C] @@ -184,12 +290,10 @@ lemma restrictScalars_mul {A B C} [Semiring A] [Semiring B] [Semiring C] variable {ι : Sort uι} theorem iSup_mul (s : ι → Submodule R A) (t : Submodule R A) : (⨆ i, s i) * t = ⨆ i, s i * t := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.iSup_mul] + iSup_smul theorem mul_iSup (t : Submodule R A) (s : ι → Submodule R A) : (t * ⨆ i, s i) = ⨆ i, t * s i := - toAddSubmonoid_injective <| by - simp only [mul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.mul_iSup] + smul_iSup /-- Sub-`R`-modules of an `R`-module form an idempotent semiring. -/ instance : NonUnitalSemiring (Submodule R A) where diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index f336a553c7f57..045c79a9465a3 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -162,7 +162,7 @@ theorem prod_Ioi_succ {M : Type*} [CommMonoid M] {n : ℕ} (i : Fin n) (v : Fin @[to_additive] theorem prod_congr' {M : Type*} [CommMonoid M] {a b : ℕ} (f : Fin b → M) (h : a = b) : - (∏ i : Fin a, f (cast h i)) = ∏ i : Fin b, f i := by + (∏ i : Fin a, f (i.cast h)) = ∏ i : Fin b, f i := by subst h congr diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index c235e6c2bd74a..34cd0b147e819 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -426,6 +426,13 @@ lemma prod_filter_not_mul_prod_filter (s : Finset α) (p : α → Prop) [Decidab (∏ x ∈ s.filter fun x ↦ ¬p x, f x) * ∏ x ∈ s.filter p, f x = ∏ x ∈ s, f x := by rw [mul_comm, prod_filter_mul_prod_filter_not] +@[to_additive] +theorem prod_filter_xor (p q : α → Prop) [DecidableEq α] [DecidablePred p] [DecidablePred q] : + (∏ x ∈ s with (Xor' (p x) (q x)), f x) = + (∏ x ∈ s with (p x ∧ ¬ q x), f x) * (∏ x ∈ s with (q x ∧ ¬ p x), f x) := by + rw [← prod_union (disjoint_filter_and_not_filter _ _), ← filter_or] + rfl + section ToList @[to_additive (attr := simp)] diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 1053dac09d757..1f4d18462f038 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -153,7 +153,7 @@ instance hasForgetToRing : HasForget₂ (AlgebraCat.{v} R) RingCat.{v} where instance hasForgetToModule : HasForget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => ModuleCat.asHom f.hom.toLinearMap } + map := fun f => ModuleCat.ofHom f.hom.toLinearMap } @[simp] lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @@ -162,7 +162,7 @@ lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @[simp] lemma forget₂_module_map {X Y : AlgebraCat.{v} R} (f : X ⟶ Y) : - (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.hom.toLinearMap := + (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.ofHom f.hom.toLinearMap := rfl variable {R} in diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index fa16a8aac4bcd..48a1d4c9ff2f5 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -60,9 +60,10 @@ noncomputable instance instMonoidalCategory : MonoidalCategory (AlgebraCat.{u} R (forget₂ (AlgebraCat R) (ModuleCat R)) { μIso := fun _ _ => Iso.refl _ εIso := Iso.refl _ - associator_eq := fun _ _ _ => TensorProduct.ext_threefold (fun _ _ _ => rfl) - leftUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) - rightUnitor_eq := fun _ => TensorProduct.ext' (fun _ _ => rfl) } + associator_eq := fun _ _ _ => + ModuleCat.hom_ext <| TensorProduct.ext_threefold (fun _ _ _ => rfl) + leftUnitor_eq := fun _ => ModuleCat.hom_ext <| TensorProduct.ext' (fun _ _ => rfl) + rightUnitor_eq := fun _ => ModuleCat.hom_ext <| TensorProduct.ext' (fun _ _ => rfl) } /-- `forget₂ (AlgebraCat R) (ModuleCat R)` as a monoidal functor. -/ example : (forget₂ (AlgebraCat R) (ModuleCat R)).Monoidal := inferInstance diff --git a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean index 3d24b8db74884..362312b348546 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean @@ -46,8 +46,9 @@ variable (R) /-- The object in the category of `R`-coalgebras associated to an `R`-coalgebra. -/ @[simps] def of (X : Type v) [AddCommGroup X] [Module R X] [Coalgebra R X] : - CoalgebraCat R where - instCoalgebra := (inferInstance : Coalgebra R X) + CoalgebraCat R := + { ModuleCat.of R X with + instCoalgebra := (inferInstance : Coalgebra R X) } variable {R} @@ -105,7 +106,7 @@ instance concreteCategory : ConcreteCategory.{v} (CoalgebraCat.{v} R) where instance hasForgetToModule : HasForget₂ (CoalgebraCat R) (ModuleCat R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => f.toCoalgHom.toLinearMap } + map := fun f => ModuleCat.ofHom f.toCoalgHom.toLinearMap } @[simp] theorem forget₂_obj (X : CoalgebraCat R) : @@ -114,7 +115,7 @@ theorem forget₂_obj (X : CoalgebraCat R) : @[simp] theorem forget₂_map (X Y : CoalgebraCat R) (f : X ⟶ Y) : - (forget₂ (CoalgebraCat R) (ModuleCat R)).map f = (f.toCoalgHom : X →ₗ[R] Y) := + (forget₂ (CoalgebraCat R) (ModuleCat R)).map f = ModuleCat.ofHom (f.toCoalgHom : X →ₗ[R] Y) := rfl end CoalgebraCat diff --git a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean index c150eabdf07e7..51ff35c7035db 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/ComonEquivalence.lean @@ -36,13 +36,13 @@ open CategoryTheory MonoidalCategory variable {R : Type u} [CommRing R] /-- An `R`-coalgebra is a comonoid object in the category of `R`-modules. -/ -@[simps] def toComonObj (X : CoalgebraCat R) : Comon_ (ModuleCat R) where +@[simps X counit comul] def toComonObj (X : CoalgebraCat R) : Comon_ (ModuleCat R) where X := ModuleCat.of R X - counit := ModuleCat.asHom Coalgebra.counit - comul := ModuleCat.asHom Coalgebra.comul - counit_comul := by simpa only [ModuleCat.of_coe] using Coalgebra.rTensor_counit_comp_comul - comul_counit := by simpa only [ModuleCat.of_coe] using Coalgebra.lTensor_counit_comp_comul - comul_assoc := by simp_rw [ModuleCat.of_coe]; exact Coalgebra.coassoc.symm + counit := ModuleCat.ofHom Coalgebra.counit + comul := ModuleCat.ofHom Coalgebra.comul + counit_comul := ModuleCat.hom_ext <| by simpa using Coalgebra.rTensor_counit_comp_comul + comul_counit := ModuleCat.hom_ext <| by simpa using Coalgebra.lTensor_counit_comp_comul + comul_assoc := ModuleCat.hom_ext <| by simp_rw [ModuleCat.of_coe]; exact Coalgebra.coassoc.symm variable (R) in /-- The natural functor from `R`-coalgebras to comonoid objects in the category of `R`-modules. -/ @@ -50,27 +50,27 @@ variable (R) in def toComon : CoalgebraCat R ⥤ Comon_ (ModuleCat R) where obj X := toComonObj X map f := - { hom := ModuleCat.asHom f.1 - hom_counit := f.1.counit_comp - hom_comul := f.1.map_comp_comul.symm } + { hom := ModuleCat.ofHom f.1 + hom_counit := ModuleCat.hom_ext f.1.counit_comp + hom_comul := ModuleCat.hom_ext f.1.map_comp_comul.symm } /-- A comonoid object in the category of `R`-modules has a natural comultiplication and counit. -/ @[simps] instance ofComonObjCoalgebraStruct (X : Comon_ (ModuleCat R)) : CoalgebraStruct R X.X where - comul := X.comul - counit := X.counit + comul := X.comul.hom + counit := X.counit.hom /-- A comonoid object in the category of `R`-modules has a natural `R`-coalgebra structure. -/ -def ofComonObj (X : Comon_ (ModuleCat R)) : CoalgebraCat R where - carrier := X.X - instCoalgebra := - { ofComonObjCoalgebraStruct X with - coassoc := X.comul_assoc.symm - rTensor_counit_comp_comul := X.counit_comul - lTensor_counit_comp_comul := X.comul_counit } +def ofComonObj (X : Comon_ (ModuleCat R)) : CoalgebraCat R := + { ModuleCat.of R X.X with + instCoalgebra := + { ofComonObjCoalgebraStruct X with + coassoc := ModuleCat.hom_ext_iff.mp X.comul_assoc.symm + rTensor_counit_comp_comul := ModuleCat.hom_ext_iff.mp X.counit_comul + lTensor_counit_comp_comul := ModuleCat.hom_ext_iff.mp X.comul_counit } } variable (R) @@ -79,9 +79,9 @@ def ofComon : Comon_ (ModuleCat R) ⥤ CoalgebraCat R where obj X := ofComonObj X map f := { toCoalgHom := - { f.hom with - counit_comp := f.hom_counit - map_comp_comul := f.hom_comul.symm }} + { f.hom.hom with + counit_comp := ModuleCat.hom_ext_iff.mp f.hom_counit + map_comp_comul := ModuleCat.hom_ext_iff.mp f.hom_comul.symm }} /-- The natural category equivalence between `R`-coalgebras and comonoid objects in the category of `R`-modules. -/ diff --git a/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean index dbd37c50e0b59..be3ac68381c8e 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/Monoidal.lean @@ -50,9 +50,9 @@ noncomputable def MonoidalCategory.inducingFunctorData : whiskerRight_eq X f := by ext; rfl tensorHom_eq f g := by ext; rfl εIso := Iso.refl _ - associator_eq X Y Z := TensorProduct.ext <| TensorProduct.ext <| by ext; rfl - leftUnitor_eq X := TensorProduct.ext <| by ext; rfl - rightUnitor_eq X := TensorProduct.ext <| by ext; rfl + associator_eq X Y Z := ModuleCat.hom_ext <| TensorProduct.ext <| TensorProduct.ext <| by ext; rfl + leftUnitor_eq X := ModuleCat.hom_ext <| TensorProduct.ext <| by ext; rfl + rightUnitor_eq X := ModuleCat.hom_ext <| TensorProduct.ext <| by ext; rfl noncomputable instance instMonoidalCategory : MonoidalCategory (CoalgebraCat R) := Monoidal.induced (forget₂ _ (ModuleCat R)) (MonoidalCategory.inducingFunctorData R) diff --git a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean index e6577b9339afb..07f1ca89ad5e6 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Basic.lean @@ -72,12 +72,6 @@ instance : LargeCategory (FGModuleCat R) := by dsimp [FGModuleCat] infer_instance -instance {M N : FGModuleCat R} : FunLike (M ⟶ N) M N := - LinearMap.instFunLike - -instance {M N : FGModuleCat R} : LinearMapClass (M ⟶ N) R M N := - LinearMap.semilinearMapClass - instance : ConcreteCategory (FGModuleCat R) := by dsimp [FGModuleCat] infer_instance @@ -101,9 +95,20 @@ instance : Inhabited (FGModuleCat R) := ⟨⟨ModuleCat.of R R, Module.Finite.self R⟩⟩ /-- Lift an unbundled finitely generated module to `FGModuleCat R`. -/ -def of (V : Type u) [AddCommGroup V] [Module R V] [Module.Finite R V] : FGModuleCat R := +abbrev of (V : Type u) [AddCommGroup V] [Module R V] [Module.Finite R V] : FGModuleCat R := ⟨ModuleCat.of R V, by change Module.Finite R V; infer_instance⟩ +variable {R} in +/-- Lift a linear map between finitely generated modules to `FGModuleCat R`. -/ +abbrev ofHom {V W : Type u} [AddCommGroup V] [Module R V] [Module.Finite R V] + [AddCommGroup W] [Module R W] [Module.Finite R W] + (f : V →ₗ[R] W) : of R V ⟶ of R W := + ModuleCat.ofHom f + +variable {R} in +@[ext] lemma hom_ext {V W : FGModuleCat R} {f g : V ⟶ W} (h : f.hom = g.hom) : f = g := + ModuleCat.hom_ext h + instance (V : FGModuleCat R) : Module.Finite R V := V.property @@ -127,8 +132,8 @@ def _root_.LinearEquiv.toFGModuleCatIso {V W : Type u} [AddCommGroup V] [Module R V] [Module.Finite R V] [AddCommGroup W] [Module R W] [Module.Finite R W] (e : V ≃ₗ[R] W) : FGModuleCat.of R V ≅ FGModuleCat.of R W where - hom := e.toLinearMap - inv := e.symm.toLinearMap + hom := ModuleCat.ofHom e.toLinearMap + inv := ModuleCat.ofHom e.symm.toLinearMap hom_inv_id := by ext x; exact e.left_inv x inv_hom_id := by ext x; exact e.right_inv x @@ -176,7 +181,11 @@ instance : (forget₂ (FGModuleCat.{u} R) (ModuleCat.{u} R)).Additive where instance : (forget₂ (FGModuleCat.{u} R) (ModuleCat.{u} R)).Linear R where theorem Iso.conj_eq_conj {V W : FGModuleCat R} (i : V ≅ W) (f : End V) : - Iso.conj i f = LinearEquiv.conj (isoToLinearEquiv i) f := + Iso.conj i f = FGModuleCat.ofHom (LinearEquiv.conj (isoToLinearEquiv i) f.hom) := + rfl + +theorem Iso.conj_hom_eq_conj {V W : FGModuleCat R} (i : V ≅ W) (f : End V) : + (Iso.conj i f).hom = (LinearEquiv.conj (isoToLinearEquiv i) f.hom) := rfl end CommRing @@ -186,11 +195,12 @@ section Field variable (K : Type u) [Field K] instance (V W : FGModuleCat K) : Module.Finite K (V ⟶ W) := - (by infer_instance : Module.Finite K (V →ₗ[K] W)) + (inferInstanceAs <| Module.Finite K (V →ₗ[K] W)).equiv ModuleCat.homLinearEquiv.symm instance closedPredicateModuleFinite : MonoidalCategory.ClosedPredicate fun V : ModuleCat.{u} K ↦ Module.Finite K V where - prop_ihom {X Y} _ _ := Module.Finite.linearMap K K X Y + prop_ihom {X Y} _ _ := + (inferInstanceAs <| Module.Finite K (X →ₗ[K] Y)).equiv ModuleCat.homLinearEquiv.symm instance : MonoidalClosed (FGModuleCat K) := by dsimp [FGModuleCat] @@ -201,7 +211,7 @@ instance : MonoidalClosed (FGModuleCat K) := by variable (V W : FGModuleCat K) @[simp] -theorem ihom_obj : (ihom V).obj W = FGModuleCat.of K (V →ₗ[K] W) := +theorem ihom_obj : (ihom V).obj W = FGModuleCat.of K (V ⟶ W) := rfl /-- The dual module is the dual in the rigid monoidal category `FGModuleCat K`. -/ @@ -216,33 +226,43 @@ open CategoryTheory.MonoidalCategory /-- The coevaluation map is defined in `LinearAlgebra.coevaluation`. -/ def FGModuleCatCoevaluation : 𝟙_ (FGModuleCat K) ⟶ V ⊗ FGModuleCatDual K V := - coevaluation K V + ModuleCat.ofHom <| coevaluation K V theorem FGModuleCatCoevaluation_apply_one : - FGModuleCatCoevaluation K V (1 : K) = + (FGModuleCatCoevaluation K V).hom (1 : K) = ∑ i : Basis.ofVectorSpaceIndex K V, (Basis.ofVectorSpace K V) i ⊗ₜ[K] (Basis.ofVectorSpace K V).coord i := coevaluation_apply_one K V /-- The evaluation morphism is given by the contraction map. -/ def FGModuleCatEvaluation : FGModuleCatDual K V ⊗ V ⟶ 𝟙_ (FGModuleCat K) := - contractLeft K V + ModuleCat.ofHom <| contractLeft K V -@[simp] theorem FGModuleCatEvaluation_apply (f : FGModuleCatDual K V) (x : V) : - (FGModuleCatEvaluation K V) (f ⊗ₜ x) = f.toFun x := + (FGModuleCatEvaluation K V).hom (f ⊗ₜ x) = f.toFun x := + contractLeft_apply f x + +/-- `@[simp]`-normal form of `FGModuleCatEvaluation_apply`, where the carriers have been unfolded. +-/ +@[simp] +theorem FGModuleCatEvaluation_apply' (f : FGModuleCatDual K V) (x : V) : + DFunLike.coe + (F := ((ModuleCat.of K (Module.Dual K V) ⊗ V.obj).carrier →ₗ[K] (𝟙_ (ModuleCat K)))) + (FGModuleCatEvaluation K V).hom (f ⊗ₜ x) = f.toFun x := contractLeft_apply f x private theorem coevaluation_evaluation : letI V' : FGModuleCat K := FGModuleCatDual K V V' ◁ FGModuleCatCoevaluation K V ≫ (α_ V' V V').inv ≫ FGModuleCatEvaluation K V ▷ V' = (ρ_ V').hom ≫ (λ_ V').inv := by + ext : 1 apply contractLeft_assoc_coevaluation K V private theorem evaluation_coevaluation : FGModuleCatCoevaluation K V ▷ V ≫ (α_ V (FGModuleCatDual K V) V).hom ≫ V ◁ FGModuleCatEvaluation K V = (λ_ V).hom ≫ (ρ_ V).inv := by + ext : 1 apply contractLeft_assoc_coevaluation' K V instance exactPairing : ExactPairing V (FGModuleCatDual K V) where @@ -266,9 +286,10 @@ end FGModuleCat @[simp] theorem LinearMap.comp_id_fgModuleCat {R} [Ring R] {G : FGModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] - (f : G →ₗ[R] H) : f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.asHom f) + (f : G →ₗ[R] H) : f.comp (ModuleCat.Hom.hom (𝟙 G)) = f := + ModuleCat.hom_ext_iff.mp <| Category.id_comp (ModuleCat.ofHom f) + @[simp] theorem LinearMap.id_fgModuleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : FGModuleCat.{u} R} - (f : G →ₗ[R] H) : LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.asHom f) + (f : G →ₗ[R] H) : LinearMap.comp (ModuleCat.Hom.hom (𝟙 H)) f = f := + ModuleCat.hom_ext_iff.mp <| Category.comp_id (ModuleCat.ofHom f) diff --git a/Mathlib/Algebra/Category/FGModuleCat/Limits.lean b/Mathlib/Algebra/Category/FGModuleCat/Limits.lean index 3e5af1918343b..7b88d667c1bdb 100644 --- a/Mathlib/Algebra/Category/FGModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/FGModuleCat/Limits.lean @@ -38,7 +38,7 @@ variable {k : Type v} [Field k] instance {J : Type} [Finite J] (Z : J → ModuleCat.{v} k) [∀ j, FiniteDimensional k (Z j)] : FiniteDimensional k (∏ᶜ fun j => Z j : ModuleCat.{v} k) := haveI : FiniteDimensional k (ModuleCat.of k (∀ j, Z j)) := by unfold ModuleCat.of; infer_instance - FiniteDimensional.of_injective (ModuleCat.piIsoPi _).hom + FiniteDimensional.of_injective (ModuleCat.piIsoPi _).hom.hom ((ModuleCat.mono_iff_injective _).1 (by infer_instance)) /-- Finite limits of finite dimensional vectors spaces are finite dimensional, @@ -48,7 +48,7 @@ instance (F : J ⥤ FGModuleCat k) : haveI : ∀ j, FiniteDimensional k ((F ⋙ forget₂ (FGModuleCat k) (ModuleCat.{v} k)).obj j) := by intro j; change FiniteDimensional k (F.obj j); infer_instance FiniteDimensional.of_injective - (limitSubobjectProduct (F ⋙ forget₂ (FGModuleCat k) (ModuleCat.{v} k))) + (limitSubobjectProduct (F ⋙ forget₂ (FGModuleCat k) (ModuleCat.{v} k))).hom ((ModuleCat.mono_iff_injective _).1 inferInstance) /-- The forgetful functor from `FGModuleCat k` to `ModuleCat k` creates all finite limits. -/ diff --git a/Mathlib/Algebra/Category/Grp/AB5.lean b/Mathlib/Algebra/Category/Grp/AB5.lean index a5230f8bf90ff..81a72475dbf30 100644 --- a/Mathlib/Algebra/Category/Grp/AB5.lean +++ b/Mathlib/Algebra/Category/Grp/AB5.lean @@ -3,14 +3,9 @@ Copyright (c) 2023 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata, Moritz Firsching, Nikolas Kuhn, Amelia Livingston -/ - -import Mathlib.Algebra.Homology.ShortComplex.Ab -import Mathlib.Algebra.Homology.ShortComplex.ExactFunctor -import Mathlib.CategoryTheory.Abelian.Exact import Mathlib.Algebra.Category.Grp.FilteredColimits -import Mathlib.CategoryTheory.Abelian.FunctorCategory +import Mathlib.Algebra.Homology.ShortComplex.Ab import Mathlib.CategoryTheory.Abelian.GrothendieckAxioms - /-! # The category of abelian groups satisfies Grothendieck's axiom AB5 @@ -51,4 +46,4 @@ instance : HasFilteredColimits (AddCommGrp.{u}) where HasColimitsOfShape := inferInstance noncomputable instance : AB5 (AddCommGrp.{u}) where - preservesFiniteLimits := fun _ => inferInstance + ofShape _ := { preservesFiniteLimits := inferInstance } diff --git a/Mathlib/Algebra/Category/Grp/Injective.lean b/Mathlib/Algebra/Category/Grp/Injective.lean index f8ea96a6004ef..a5f8e576a886a 100644 --- a/Mathlib/Algebra/Category/Grp/Injective.lean +++ b/Mathlib/Algebra/Category/Grp/Injective.lean @@ -57,9 +57,9 @@ theorem Module.Baer.of_divisible [DivisibleBy A ℤ] : Module.Baer ℤ A := fun namespace AddCommGrp -theorem injective_as_module_iff : Injective (⟨A⟩ : ModuleCat ℤ) ↔ +theorem injective_as_module_iff : Injective (ModuleCat.of ℤ A) ↔ Injective (⟨A,inferInstance⟩ : AddCommGrp) := - ((forget₂ (ModuleCat ℤ) AddCommGrp).asEquivalence.map_injective_iff ⟨A⟩).symm + ((forget₂ (ModuleCat ℤ) AddCommGrp).asEquivalence.map_injective_iff (ModuleCat.of ℤ A)).symm instance injective_of_divisible [DivisibleBy A ℤ] : Injective (⟨A,inferInstance⟩ : AddCommGrp) := diff --git a/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean index 242349db0fa0c..cd88f753ea5b8 100644 --- a/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean +++ b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean @@ -29,12 +29,13 @@ instance forget₂_addCommGroup_full : (forget₂ (ModuleCat ℤ) AddCommGrp.{u} -- `AddMonoidHom.toIntLinearMap` doesn't work here because `A` and `B` are not -- definitionally equal to the canonical `AddCommGroup.toIntModule` module -- instances it expects. - f := ⟨@LinearMap.mk _ _ _ _ _ _ _ _ _ A.isModule B.isModule - { toFun := f, - map_add' := AddMonoidHom.map_add (show A.carrier →+ B.carrier from f) } - (fun n x => by - convert AddMonoidHom.map_zsmul (show A.carrier →+ B.carrier from f) x n <;> - ext <;> apply int_smul_eq_zsmul), rfl⟩ + f := ⟨@ModuleCat.ofHom _ _ _ _ _ A.isModule _ B.isModule <| + @LinearMap.mk _ _ _ _ _ _ _ _ _ A.isModule B.isModule + { toFun := f, + map_add' := AddMonoidHom.map_add (show A.carrier →+ B.carrier from f) } + (fun n x => by + convert AddMonoidHom.map_zsmul (show A.carrier →+ B.carrier from f) x n <;> + ext <;> apply int_smul_eq_zsmul), rfl⟩ /-- The forgetful functor from `ℤ` modules to `AddCommGrp` is essentially surjective. -/ instance forget₂_addCommGrp_essSurj : (forget₂ (ModuleCat ℤ) AddCommGrp.{u}).EssSurj where diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index fcb99e727a6c4..2bb2a6d564458 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -29,9 +29,9 @@ variable {R : Type u} [Ring R] {M N : ModuleCat.{v} R} (f : M ⟶ N) /-- In the category of modules, every monomorphism is normal. -/ def normalMono (hf : Mono f) : NormalMono f where - Z := of R (N ⧸ LinearMap.range f) - g := f.range.mkQ - w := LinearMap.range_mkQ_comp _ + Z := of R (N ⧸ LinearMap.range f.hom) + g := ofHom f.hom.range.mkQ + w := hom_ext <| LinearMap.range_mkQ_comp _ isLimit := /- The following [invalid Lean code](https://github.com/leanprover-community/lean/issues/341) might help you understand what's going on here: @@ -43,16 +43,16 @@ def normalMono (hf : Mono f) : NormalMono f where ``` -/ IsKernel.isoKernel _ _ (kernelIsLimit _) - (LinearEquiv.toModuleIso' + (LinearEquiv.toModuleIso ((Submodule.quotEquivOfEqBot _ (ker_eq_bot_of_mono _)).symm ≪≫ₗ - (LinearMap.quotKerEquivRange f ≪≫ₗ + (LinearMap.quotKerEquivRange f.hom ≪≫ₗ LinearEquiv.ofEq _ _ (Submodule.ker_mkQ _).symm))) <| by ext; rfl /-- In the category of modules, every epimorphism is normal. -/ def normalEpi (hf : Epi f) : NormalEpi f where - W := of R (LinearMap.ker f) - g := (LinearMap.ker f).subtype - w := LinearMap.comp_ker_subtype _ + W := of R (LinearMap.ker f.hom) + g := ofHom (LinearMap.ker f.hom).subtype + w := hom_ext <| LinearMap.comp_ker_subtype _ isColimit := /- The following invalid Lean code might help you understand what's going on here: ``` @@ -63,9 +63,9 @@ def normalEpi (hf : Epi f) : NormalEpi f where ``` -/ IsCokernel.cokernelIso _ _ (cokernelIsColimit _) - (LinearEquiv.toModuleIso' + (LinearEquiv.toModuleIso (Submodule.quotEquivOfEq _ _ (Submodule.range_subtype _) ≪≫ₗ - LinearMap.quotKerEquivRange f ≪≫ₗ + LinearMap.quotKerEquivRange f.hom ≪≫ₗ LinearEquiv.ofTop _ (range_eq_top_of_epi _))) <| by ext; rfl /-- The category of R-modules is abelian. -/ diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index 12126e451178b..41760fe6fdefd 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -36,9 +36,9 @@ free `R`-module with generators `x : X`, implemented as the type `X →₀ R`. -/ def free : Type u ⥤ ModuleCat R where obj X := ModuleCat.of R (X →₀ R) - map {_ _} f := Finsupp.lmapDomain _ _ f - map_id := by intros; exact Finsupp.lmapDomain_id _ _ - map_comp := by intros; exact Finsupp.lmapDomain_comp _ _ _ _ + map {_ _} f := ofHom <| Finsupp.lmapDomain _ _ f + map_id := by intros; ext : 1; exact Finsupp.lmapDomain_id _ _ + map_comp := by intros; ext : 1; exact Finsupp.lmapDomain_comp _ _ _ _ variable {R} @@ -49,13 +49,13 @@ noncomputable def freeMk {X : Type u} (x : X) : (free R).obj X := Finsupp.single lemma free_hom_ext {X : Type u} {M : ModuleCat.{u} R} {f g : (free R).obj X ⟶ M} (h : ∀ (x : X), f (freeMk x) = g (freeMk x)) : f = g := - (Finsupp.lhom_ext' (fun x ↦ LinearMap.ext_ring (h x))) + ModuleCat.hom_ext (Finsupp.lhom_ext' (fun x ↦ LinearMap.ext_ring (h x))) /-- The morphism of modules `(free R).obj X ⟶ M` corresponding to a map `f : X ⟶ M`. -/ noncomputable def freeDesc {X : Type u} {M : ModuleCat.{u} R} (f : X ⟶ M) : (free R).obj X ⟶ M := - Finsupp.lift M R X f + ofHom <| Finsupp.lift M R X f @[simp] lemma freeDesc_apply {X : Type u} {M : ModuleCat.{u} R} (f : X ⟶ M) (x : X) : @@ -109,12 +109,11 @@ namespace FreeMonoidal (This should not be used directly: it is part of the implementation of the monoidal structure on the functor `free R`.) -/ def εIso : 𝟙_ (ModuleCat R) ≅ (free R).obj (𝟙_ (Type u)) where - hom := Finsupp.lsingle PUnit.unit - inv := Finsupp.lapply PUnit.unit + hom := ofHom <| Finsupp.lsingle PUnit.unit + inv := ofHom <| Finsupp.lapply PUnit.unit hom_inv_id := by - ext x - dsimp - erw [Finsupp.lapply_apply, Finsupp.lsingle_apply, Finsupp.single_eq_same] + ext + simp [free] inv_hom_id := by ext ⟨⟩ dsimp [freeMk] diff --git a/Mathlib/Algebra/Category/ModuleCat/Algebra.lean b/Mathlib/Algebra/Category/ModuleCat/Algebra.lean index d2d388822bfff..275ae2fa12b4f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Algebra.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Algebra.lean @@ -51,18 +51,9 @@ theorem isScalarTower_of_algebra_moduleCat (M : ModuleCat.{v} A) : IsScalarTower attribute [scoped instance] ModuleCat.isScalarTower_of_algebra_moduleCat -- We verify that the morphism spaces become `k`-modules. -example (M N : ModuleCat.{v} A) : Module k (M ⟶ N) := LinearMap.module --- Porting note: used to be `by infer_instance` instead of `LinearMap.module` +example (M N : ModuleCat.{v} A) : Module k (M ⟶ N) := inferInstance instance linearOverField : Linear k (ModuleCat.{v} A) where - -- Porting note: used to be `by infer_instance` instead of `LinearMap.module` - homModule _ _ := LinearMap.module - smul_comp := by - -- Porting note: this was automatic by `aesop_cat` - intros - ext - dsimp only [coe_comp, Function.comp_apply] - rw [LinearMap.smul_apply, LinearMap.map_smul_of_tower] - rfl + homModule _ _ := inferInstance end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 417c5c7e12758..9d0fdd30373c1 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -22,33 +22,15 @@ respectively. To construct an object in the category of `R`-modules from a type `M` with an instance of the `Module` typeclass, write `of R M`. There is a coercion in the other direction. +The roundtrip `↑(of R M)` is definitionally equal to `M` itself (when `M` is a type with `Module` +instance), and so is `of R ↑M` (when `M : ModuleCat R M`). -Similarly, there is a coercion from morphisms in `Module R` to linear maps. - -Porting note: the next two paragraphs should be revised. - -Unfortunately, Lean is not smart enough to see that, given an object `M : ModuleCat R`, -the expression `of R M`, where we coerce `M` to the carrier type, -is definitionally equal to `M` itself. -This means that to go the other direction, i.e., from linear maps/equivalences to (iso)morphisms -in the category of `R`-modules, we have to take care not to inadvertently end up with an -`of R M` where `M` is already an object. Hence, given `f : M →ₗ[R] N`, -* if `M N : ModuleCat R`, simply use `f`; -* if `M : ModuleCat R` and `N` is an unbundled `R`-module, use `↿f` or `asHomLeft f`; -* if `M` is an unbundled `R`-module and `N : ModuleCat R`, use `↾f` or `asHomRight f`; -* if `M` and `N` are unbundled `R`-modules, use `↟f` or `asHom f`. - -Similarly, given `f : M ≃ₗ[R] N`, use `toModuleIso`, `toModuleIso'Left`, `toModuleIso'Right` -or `toModuleIso'`, respectively. - -The arrow notations are localized, so you may have to `open ModuleCat` (or `open scoped ModuleCat`) -to use them. Note that the notation for `asHomLeft` clashes with the notation used to promote -functions between types to morphisms in the category `Type`, so to avoid confusion, it is probably a -good idea to avoid having the locales `ModuleCat` and `CategoryTheory.Type` open at the same time. - -If you get an error when trying to apply a theorem and the `convert` tactic produces goals of the -form `M = of R M`, then you probably used an incorrect variant of `asHom` or `toModuleIso`. +The morphisms are given their own type, not identified with `LinearMap`. +There is a cast from morphisms in `Module R` to linear maps, written `f.hom` (`ModuleCat.Hom.hom`). +To go from linear maps to morphisms in `Module R`, use `ModuleCat.ofHom`. +Similarly, given an isomorphism `f : M ≅ N` use `f.toLinearEquiv` and given a linear equiv +`f : M ≃ₗ[R] N`, use `f.toModuleIso`. -/ @@ -69,6 +51,7 @@ impose here that the `ℤ`-multiplication field from the module structure is def from the `isAddCommGroup` structure (contrary to what we do for all module structures in mathlib), which creates some difficulties down the road. -/ structure ModuleCat where + private mk :: /-- the underlying type of an object in `ModuleCat R` -/ carrier : Type v [isAddCommGroup : AddCommGroup carrier] @@ -88,27 +71,137 @@ instance : CoeSort (ModuleCat.{v} R) (Type v) := attribute [coe] ModuleCat.carrier +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `ModuleCat R`. -/ +abbrev of (X : Type v) [AddCommGroup X] [Module R X] : ModuleCat.{v} R := + ⟨X⟩ + +lemma coe_of (X : Type v) [Ring X] [Module R X] : (of R X : Type v) = X := + rfl + +-- Ensure the roundtrips are reducibly defeq (so tactics like `rw` can see through them). +example (X : Type v) [Ring X] [Module R X] : (of R X : Type v) = X := by with_reducible rfl +example (M : ModuleCat.{v} R) : of R M = M := by with_reducible rfl + +variable {R} in +/-- The type of morphisms in `ModuleCat R`. -/ +@[ext] +structure Hom (M N : ModuleCat.{v} R) where + private mk :: + /-- The underlying linear map. -/ + hom : M →ₗ[R] N + instance moduleCategory : Category.{v, max (v+1) u} (ModuleCat.{v} R) where - Hom M N := M →ₗ[R] N - id _ := LinearMap.id - comp f g := g.comp f - id_comp _ := LinearMap.id_comp _ - comp_id _ := LinearMap.comp_id _ - assoc f g h := LinearMap.comp_assoc (f := f) (g := g) (h := h) + Hom M N := Hom M N + id _ := ⟨LinearMap.id⟩ + comp f g := ⟨g.hom.comp f.hom⟩ + +instance {M N : ModuleCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where + coe f := f.hom + +section + +variable {R} + +@[simp] +lemma hom_id {M : ModuleCat.{v} R} : (𝟙 M : M ⟶ M).hom = LinearMap.id := rfl + +/- Provided for rewriting. -/ +lemma id_apply (M : ModuleCat.{v} R) (x : M) : + (𝟙 M : M ⟶ M) x = x := by simp + +@[simp] +lemma hom_comp {M N O : ModuleCat.{v} R} (f : M ⟶ N) (g : N ⟶ O) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + +/- Provided for rewriting. -/ +lemma comp_apply {M N O : ModuleCat.{v} R} (f : M ⟶ N) (g : N ⟶ O) (x : M) : + (f ≫ g) x = g (f x) := by simp + +@[ext] +lemma hom_ext {M N : ModuleCat.{v} R} {f g : M ⟶ N} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +lemma hom_bijective {M N : ModuleCat.{v} R} : + Function.Bijective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) where + left f g h := by cases f; cases g; simpa using h + right f := ⟨⟨f⟩, rfl⟩ + +/-- Convenience shortcut for `ModuleCat.hom_bijective.injective`. -/ +lemma hom_injective {M N : ModuleCat.{v} R} : + Function.Injective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) := + hom_bijective.injective + +/-- Convenience shortcut for `ModuleCat.hom_bijective.surjective`. -/ +lemma hom_surjective {M N : ModuleCat.{v} R} : + Function.Surjective (Hom.hom : (M ⟶ N) → (M →ₗ[R] N)) := + hom_bijective.surjective + +/-- Typecheck a `LinearMap` as a morphism in `ModuleCat R`. -/ +abbrev ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] + (f : X →ₗ[R] Y) : of R X ⟶ of R Y := + ⟨f⟩ + +@[deprecated (since := "2024-10-06")] alias asHom := ModuleCat.ofHom -instance {M N : ModuleCat.{v} R} : FunLike (M ⟶ N) M N := - LinearMap.instFunLike +/- Doesn't need to be `@[simp]` since the `simp` tactic applies this rewrite automatically: +`ofHom` and `hom` are reducibly equal to the constructor and projection respectively. -/ +lemma hom_ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] + [Module R Y] (f : X →ₗ[R] Y) : (ofHom f).hom = f := rfl -instance {M N : ModuleCat.{v} R} : LinearMapClass (M ⟶ N) R M N := - LinearMap.semilinearMapClass +@[simp] +lemma ofHom_hom {M N : ModuleCat.{v} R} (f : M ⟶ N) : + ofHom (Hom.hom f) = f := rfl + +@[simp] +lemma ofHom_id {M : Type v} [AddCommGroup M] [Module R M] : ofHom LinearMap.id = 𝟙 (of R M) := rfl + +@[simp] +lemma ofHom_comp {M N O : Type v} [AddCommGroup M] [AddCommGroup N] [AddCommGroup O] [Module R M] + [Module R N] [Module R O] (f : M →ₗ[R] N) (g : N →ₗ[R] O) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := + rfl + +/- Doesn't need to be `@[simp]` since `simp only` can solve this. -/ +lemma ofHom_apply {M N : Type v} [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] + (f : M →ₗ[R] N) (x : M) : ofHom f x = f x := rfl + +@[simp] +lemma inv_hom_apply {M N : ModuleCat.{v} R} (e : M ≅ N) (x : M) : e.inv (e.hom x) = x := by + rw [← comp_apply] + simp + +@[simp] +lemma hom_inv_apply {M N : ModuleCat.{v} R} (e : M ≅ N) (x : N) : e.hom (e.inv x) = x := by + rw [← comp_apply] + simp + +/-- `ModuleCat.Hom.hom` bundled as an `Equiv`. -/ +def homEquiv {M N : ModuleCat.{v} R} : (M ⟶ N) ≃ (M →ₗ[R] N) where + toFun := Hom.hom + invFun := ofHom + left_inv _ := rfl + right_inv _ := rfl + +end + +instance : Inhabited (ModuleCat R) := + ⟨of R R⟩ instance moduleConcreteCategory : ConcreteCategory.{v} (ModuleCat.{v} R) where forget := { obj := fun R => R - map := fun f => f.toFun } - forget_faithful := ⟨fun h => LinearMap.ext (fun x => by - dsimp at h - rw [h])⟩ + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ + +/- Not a `@[simp]` lemma since it will rewrite the (co)domain of maps and cause +definitional equality issues. -/ +lemma forget_obj {M : ModuleCat.{v} R} : (forget (ModuleCat.{v} R)).obj M = M := rfl + +/- Not a `@[simp]` lemma since the LHS is a categorical arrow and the RHS is a plain function. -/ +lemma forget_map {M N : ModuleCat.{v} R} (f : M ⟶ N) : + (forget (ModuleCat.{v} R)).map f = f := + rfl -- Porting note: -- One might hope these two instances would not be needed, @@ -119,18 +212,10 @@ instance {M : ModuleCat.{v} R} : AddCommGroup ((forget (ModuleCat R)).obj M) := instance {M : ModuleCat.{v} R} : Module R ((forget (ModuleCat R)).obj M) := (inferInstance : Module R M) -@[ext] -lemma ext {M N : ModuleCat.{v} R} {f₁ f₂ : M ⟶ N} (h : ∀ (x : M), f₁ x = f₂ x) : f₁ = f₂ := - DFunLike.ext _ _ h - instance hasForgetToAddCommGroup : HasForget₂ (ModuleCat R) AddCommGrp where forget₂ := { obj := fun M => AddCommGrp.of M - map := fun f => AddCommGrp.ofHom f.toAddMonoidHom } - -/-- The object in the category of R-modules associated to an R-module -/ -def of (X : Type v) [AddCommGroup X] [Module R X] : ModuleCat R := - ⟨X⟩ + map := fun f => AddCommGrp.ofHom f.hom.toAddMonoidHom } @[simp] theorem forget₂_obj (X : ModuleCat R) : @@ -148,22 +233,14 @@ theorem forget₂_obj_moduleCat_of (X : Type v) [AddCommGroup X] [Module R X] : @[simp] theorem forget₂_map (X Y : ModuleCat R) (f : X ⟶ Y) : - (forget₂ (ModuleCat R) AddCommGrp).map f = LinearMap.toAddMonoidHom f := + (forget₂ (ModuleCat R) AddCommGrp).map f = f.hom.toAddMonoidHom := rfl instance : Inhabited (ModuleCat R) := ⟨of R PUnit⟩ -instance ofUnique {X : Type v} [AddCommGroup X] [Module R X] [i : Unique X] : Unique (of R X) := - i - @[simp] theorem of_coe (X : ModuleCat R) : of R X = X := rfl --- Porting note: the simpNF linter complains, but we really need this?! --- @[simp, nolint simpNF] -theorem coe_of (X : Type v) [AddCommGroup X] [Module R X] : (of R X : Type v) = X := - rfl - variable {R} /-- Forgetting to the underlying type and then building the bundled object returns the original @@ -174,33 +251,18 @@ def ofSelfIso (M : ModuleCat R) : ModuleCat.of R M ≅ M where inv := 𝟙 M theorem isZero_of_subsingleton (M : ModuleCat R) [Subsingleton M] : IsZero M where - unique_to X := ⟨⟨⟨(0 : M →ₗ[R] X)⟩, fun f => by + unique_to X := ⟨⟨⟨ofHom (0 : M →ₗ[R] X)⟩, fun f => by ext x rw [Subsingleton.elim x (0 : M)] dsimp simp⟩⟩ - unique_from X := ⟨⟨⟨(0 : X →ₗ[R] M)⟩, fun f => by + unique_from X := ⟨⟨⟨ofHom (0 : X →ₗ[R] M)⟩, fun f => by ext x subsingleton⟩⟩ instance : HasZeroObject (ModuleCat.{v} R) := ⟨⟨of R PUnit, isZero_of_subsingleton _⟩⟩ -variable {M N U : ModuleCat.{v} R} - -@[simp] -theorem id_apply (m : M) : (𝟙 M : M → M) m = m := - rfl - -@[simp] -theorem coe_comp (f : M ⟶ N) (g : N ⟶ U) : (f ≫ g : M → U) = g ∘ f := - rfl - -theorem comp_def (f : M ⟶ N) (g : N ⟶ U) : f ≫ g = g.comp f := - rfl - -@[simp] lemma forget_map (f : M ⟶ N) : (forget (ModuleCat R)).map f = (f : M → N) := rfl - end ModuleCat variable {R} @@ -208,37 +270,23 @@ variable {X₁ X₂ : Type v} open ModuleCat -/-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.asHom [AddCommGroup X₁] [Module R X₁] [AddCommGroup X₂] [Module R X₂] : - (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ ModuleCat.of R X₂) := - id - -@[deprecated (since := "2024-10-06")] alias ModuleCat.ofHom := ModuleCat.asHom - /-- Reinterpreting a linear map in the category of `R`-modules -/ -scoped[ModuleCat] notation "↟" f:1024 => ModuleCat.asHom f +scoped[ModuleCat] notation "↟" f:1024 => ModuleCat.ofHom f -@[simp 1100] -theorem ModuleCat.asHom_apply {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] - [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) (x : X) : (↟ f) x = f x := - rfl - -@[deprecated (since := "2024-10-06")] alias ModuleCat.ofHom_apply := ModuleCat.asHom_apply +@[deprecated (since := "2024-10-06")] alias ModuleCat.asHom_apply := ModuleCat.ofHom_apply -/-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.asHomRight [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat.{v} R} : - (X₁ →ₗ[R] X₂) → (ModuleCat.of R X₁ ⟶ X₂) := - id +-- Since `of` and the coercion now roundtrip reducibly, we don't need to distinguish in which place +-- we need to add `of` when coercing from linear maps to morphisms. +@[deprecated ModuleCat.ofHom (since := "2024-11-29")] alias ModuleCat.asHomRight := ModuleCat.ofHom +@[deprecated ModuleCat.ofHom (since := "2024-11-29")] alias ModuleCat.asHomLeft := ModuleCat.ofHom -/-- Reinterpreting a linear map in the category of `R`-modules. -/ +/-- Reinterpreting a linear map in the category of `R`-modules. +This notation is deprecated: use `↟` instead. +-/ scoped[ModuleCat] notation "↾" f:1024 => ModuleCat.asHomRight f - -/-- Reinterpreting a linear map in the category of `R`-modules. -/ -def ModuleCat.asHomLeft {X₁ : ModuleCat.{v} R} [AddCommGroup X₂] [Module R X₂] : - (X₁ →ₗ[R] X₂) → (X₁ ⟶ ModuleCat.of R X₂) := - id - -/-- Reinterpreting a linear map in the category of `R`-modules. -/ +/-- Reinterpreting a linear map in the category of `R`-modules. +This notation is deprecated: use `↟` instead. +-/ scoped[ModuleCat] notation "↿" f:1024 => ModuleCat.asHomLeft f section @@ -247,21 +295,24 @@ section @[simps] def LinearEquiv.toModuleIso {g₁ : AddCommGroup X₁} {g₂ : AddCommGroup X₂} {m₁ : Module R X₁} {m₂ : Module R X₂} (e : X₁ ≃ₗ[R] X₂) : ModuleCat.of R X₁ ≅ ModuleCat.of R X₂ where - hom := (e : X₁ →ₗ[R] X₂) - inv := (e.symm : X₂ →ₗ[R] X₁) + hom := ofHom (e : X₁ →ₗ[R] X₂) + inv := ofHom (e.symm : X₂ →ₗ[R] X₁) hom_inv_id := by ext; apply e.left_inv inv_hom_id := by ext; apply e.right_inv /-- Build an isomorphism in the category `Module R` from a `LinearEquiv` between `Module`s. -/ +@[deprecated LinearEquiv.toModuleIso (since := "2024-11-29")] abbrev LinearEquiv.toModuleIso' {M N : ModuleCat.{v} R} (i : M ≃ₗ[R] N) : M ≅ N := i.toModuleIso /-- Build an isomorphism in the category `ModuleCat R` from a `LinearEquiv` between `Module`s. -/ +@[deprecated LinearEquiv.toModuleIso (since := "2024-11-29")] abbrev LinearEquiv.toModuleIso'Left {X₁ : ModuleCat.{v} R} [AddCommGroup X₂] [Module R X₂] (e : X₁ ≃ₗ[R] X₂) : X₁ ≅ ModuleCat.of R X₂ := e.toModuleIso /-- Build an isomorphism in the category `ModuleCat R` from a `LinearEquiv` between `Module`s. -/ +@[deprecated LinearEquiv.toModuleIso (since := "2024-11-29")] abbrev LinearEquiv.toModuleIso'Right [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat.{v} R} (e : X₁ ≃ₗ[R] X₂) : ModuleCat.of R X₁ ≅ X₂ := e.toModuleIso @@ -270,7 +321,7 @@ namespace CategoryTheory.Iso /-- Build a `LinearEquiv` from an isomorphism in the category `ModuleCat R`. -/ def toLinearEquiv {X Y : ModuleCat R} (i : X ≅ Y) : X ≃ₗ[R] Y := - LinearEquiv.ofLinear i.hom i.inv i.inv_hom_id i.hom_inv_id + LinearEquiv.ofLinear i.hom.hom i.inv.hom (by aesop) (by aesop) end CategoryTheory.Iso @@ -286,45 +337,122 @@ end namespace ModuleCat -instance {M N : ModuleCat.{v} R} : AddCommGroup (M ⟶ N) := LinearMap.addCommGroup +section AddCommGroup + +variable {M N : ModuleCat.{v} R} + +instance : Add (M ⟶ N) where + add f g := ⟨f.hom + g.hom⟩ + +@[simp] lemma hom_add (f g : M ⟶ N) : (f + g).hom = f.hom + g.hom := rfl + +instance : Zero (M ⟶ N) where + zero := ⟨0⟩ + +@[simp] lemma hom_zero : (0 : M ⟶ N).hom = 0 := rfl + +instance : SMul ℕ (M ⟶ N) where + smul n f := ⟨n • f.hom⟩ + +@[simp] lemma hom_nsmul (n : ℕ) (f : M ⟶ N) : (n • f).hom = n • f.hom := rfl + +instance : Neg (M ⟶ N) where + neg f := ⟨-f.hom⟩ + +@[simp] lemma hom_neg (f : M ⟶ N) : (-f).hom = -f.hom := rfl + +instance : Sub (M ⟶ N) where + sub f g := ⟨f.hom - g.hom⟩ + +@[simp] lemma hom_sub (f g : M ⟶ N) : (f - g).hom = f.hom - g.hom := rfl + +instance : SMul ℤ (M ⟶ N) where + smul n f := ⟨n • f.hom⟩ + +@[simp] lemma hom_zsmul (n : ℕ) (f : M ⟶ N) : (n • f).hom = n • f.hom := rfl + +instance : AddCommGroup (M ⟶ N) := + Function.Injective.addCommGroup (Hom.hom) hom_injective + rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) + +@[simp] lemma hom_sum {ι : Type*} (f : ι → (M ⟶ N)) (s : Finset ι) : + (∑ i in s, f i).hom = ∑ i in s, (f i).hom := + map_sum ({ toFun := ModuleCat.Hom.hom, map_zero' := ModuleCat.hom_zero, map_add' := hom_add } : + (M ⟶ N) →+ (M →ₗ[R] N)) _ _ instance : Preadditive (ModuleCat.{v} R) where - add_comp P Q R f f' g := by - ext - dsimp - erw [map_add] - rfl instance forget₂_addCommGrp_additive : (forget₂ (ModuleCat.{v} R) AddCommGrp).Additive where +/-- `ModuleCat.Hom.hom` bundled as an additive equivalence. -/ +@[simps!] +def homAddEquiv : (M ⟶ N) ≃+ (M →ₗ[R] N) := + { homEquiv with + map_add' := fun _ _ => rfl } + +end AddCommGroup + +section SMul + +variable {M N : ModuleCat.{v} R} {S : Type*} [Monoid S] [DistribMulAction S N] [SMulCommClass R S N] + +instance : SMul S (M ⟶ N) where + smul c f := ⟨c • f.hom⟩ + +@[simp] lemma hom_smul (s : S) (f : M ⟶ N) : (s • f).hom = s • f.hom := rfl + +end SMul + +section Module + +variable {M N : ModuleCat.{v} R} {S : Type*} [Semiring S] [Module S N] [SMulCommClass R S N] + +instance Hom.instModule : Module S (M ⟶ N) := + Function.Injective.module S + { toFun := Hom.hom, map_zero' := hom_zero, map_add' := hom_add } + hom_injective + (fun _ _ => rfl) + +/-- `ModuleCat.Hom.hom` bundled as a linear equivalence. -/ +@[simps] +def homLinearEquiv : (M ⟶ N) ≃ₗ[S] (M →ₗ[R] N) := + { homAddEquiv with + map_smul' := fun _ _ => rfl } + +end Module + section variable {S : Type u} [CommRing S] +variable {M N : ModuleCat.{v} S} + instance : Linear S (ModuleCat.{v} S) where - homModule _ _ := LinearMap.module - smul_comp := by - intros - ext - dsimp - rw [LinearMap.smul_apply, LinearMap.smul_apply, map_smul] - rfl variable {X Y X' Y' : ModuleCat.{v} S} theorem Iso.homCongr_eq_arrowCongr (i : X ≅ X') (j : Y ≅ Y') (f : X ⟶ Y) : - Iso.homCongr i j f = LinearEquiv.arrowCongr i.toLinearEquiv j.toLinearEquiv f := + Iso.homCongr i j f = ⟨LinearEquiv.arrowCongr i.toLinearEquiv j.toLinearEquiv f.hom⟩ := rfl theorem Iso.conj_eq_conj (i : X ≅ X') (f : End X) : - Iso.conj i f = LinearEquiv.conj i.toLinearEquiv f := + Iso.conj i f = ⟨LinearEquiv.conj i.toLinearEquiv f.hom⟩ := rfl end variable (M N : ModuleCat.{v} R) +/-- `ModuleCat.Hom.hom` as an isomorphism of monoids. -/ +@[simps] +def endMulEquiv : End M ≃* (M →ₗ[R] M) where + toFun := ModuleCat.Hom.hom + invFun := ModuleCat.ofHom + map_mul' _ _ := rfl + left_inv _ := rfl + right_inv _ := rfl + /-- The scalar multiplication on an object of `ModuleCat R` considered as a morphism of rings from `R` to the endomorphisms of the underlying abelian group. -/ def smul : R →+* End ((forget₂ (ModuleCat R) AddCommGrp).obj M) where @@ -341,7 +469,7 @@ lemma smul_naturality {M N : ModuleCat.{v} R} (f : M ⟶ N) (r : R) : (forget₂ (ModuleCat R) AddCommGrp).map f ≫ N.smul r = M.smul r ≫ (forget₂ (ModuleCat R) AddCommGrp).map f := by ext x - exact (f.map_smul r x).symm + exact (f.hom.map_smul r x).symm variable (R) @@ -409,9 +537,9 @@ a morphism between the underlying objects in `AddCommGrp` and the compatibility with the scalar multiplication. -/ @[simps] def homMk : M ⟶ N where - toFun := φ - map_add' _ _ := φ.map_add _ _ - map_smul' r x := (congr_hom (hφ r) x).symm + hom.toFun := φ + hom.map_add' _ _ := φ.map_add _ _ + hom.map_smul' r x := (congr_hom (hφ r) x).symm lemma forget₂_map_homMk : (forget₂ (ModuleCat R) AddCommGrp).map (homMk φ hφ) = φ := rfl @@ -420,7 +548,7 @@ end instance : (forget (ModuleCat.{v} R)).ReflectsIsomorphisms where reflects f _ := - (inferInstance : IsIso ((LinearEquiv.mk f + (inferInstance : IsIso ((LinearEquiv.mk f.hom (asIso ((forget (ModuleCat R)).map f)).toEquiv.invFun (Equiv.left_inv _) (Equiv.right_inv _)).toModuleIso).hom) @@ -433,15 +561,45 @@ instance : (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}).ReflectsIsomorphisms whe end ModuleCat +section Bilinear + +variable {R : Type*} [CommRing R] + +namespace ModuleCat + +/-- Turn a bilinear map into a homomorphism. -/ +@[simps] +def ofHom₂ {M N P : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) : + M ⟶ of R (N ⟶ P) := + ofHom <| homLinearEquiv.symm.toLinearMap ∘ₗ f + +/-- Turn a homomorphism into a bilinear map. -/ +@[simps!] +def Hom.hom₂ {M N P : ModuleCat.{u} R} + -- We write `Hom` instead of `M ⟶ (of R (N ⟶ P))`, otherwise dot notation breaks + -- since it is expecting the type of `f` to be `ModuleCat.Hom`, not `Quiver.Hom`. + (f : Hom M (of R (N ⟶ P))) : + M →ₗ[R] N →ₗ[R] P := + Hom.hom (by convert (f ≫ ofHom homLinearEquiv.toLinearMap)) + +@[simp] lemma Hom.hom₂_ofHom₂ {M N P : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) : + (ofHom₂ f).hom₂ = f := rfl + +@[simp] lemma ofHom₂_hom₂ {M N P : ModuleCat.{u} R} (f : M ⟶ of R (N ⟶ P)) : + ofHom₂ f.hom₂ = f := rfl + +end ModuleCat + +end Bilinear + /-! `@[simp]` lemmas for `LinearMap.comp` and categorical identities. -/ @[simp] theorem LinearMap.comp_id_moduleCat {R} [Ring R] {G : ModuleCat.{u} R} {H : Type u} [AddCommGroup H] [Module R H] (f : G →ₗ[R] H) : - f.comp (𝟙 G) = f := - Category.id_comp (ModuleCat.asHom f) + f.comp (𝟙 G : G ⟶ G).hom = f := by simp + @[simp] theorem LinearMap.id_moduleCat_comp {R} [Ring R] {G : Type u} [AddCommGroup G] [Module R G] {H : ModuleCat.{u} R} (f : G →ₗ[R] H) : - LinearMap.comp (𝟙 H) f = f := - Category.comp_id (ModuleCat.asHom f) + LinearMap.comp (𝟙 H : H ⟶ H).hom f = f := by simp diff --git a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean index b1862b543607c..4387064834c7e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean @@ -41,10 +41,12 @@ def binaryProductLimitCone (M N : ModuleCat.{v} R) : Limits.LimitCone (pair M N) π := { app := fun j => Discrete.casesOn j fun j => - WalkingPair.casesOn j (LinearMap.fst R M N) (LinearMap.snd R M N) + WalkingPair.casesOn j (ofHom <| LinearMap.fst R M N) (ofHom <| LinearMap.snd R M N) naturality := by rintro ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨⟨⟨⟩⟩⟩ <;> rfl } } isLimit := - { lift := fun s => LinearMap.prod (s.π.app ⟨WalkingPair.left⟩) (s.π.app ⟨WalkingPair.right⟩) + { lift := fun s => ofHom <| LinearMap.prod + (s.π.app ⟨WalkingPair.left⟩).hom + (s.π.app ⟨WalkingPair.right⟩).hom fac := by rintro s (⟨⟩ | ⟨⟩) <;> rfl uniq := fun s m w => by simp_rw [← w ⟨WalkingPair.left⟩, ← w ⟨WalkingPair.right⟩] @@ -52,12 +54,12 @@ def binaryProductLimitCone (M N : ModuleCat.{v} R) : Limits.LimitCone (pair M N) @[simp] theorem binaryProductLimitCone_cone_π_app_left (M N : ModuleCat.{v} R) : - (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.left⟩ = LinearMap.fst R M N := + (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.left⟩ = ofHom (LinearMap.fst R M N) := rfl @[simp] theorem binaryProductLimitCone_cone_π_app_right (M N : ModuleCat.{v} R) : - (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.right⟩ = LinearMap.snd R M N := + (binaryProductLimitCone M N).cone.π.app ⟨WalkingPair.right⟩ = ofHom (LinearMap.snd R M N) := rfl /-- We verify that the biproduct in `ModuleCat R` is isomorphic to @@ -69,12 +71,12 @@ noncomputable def biprodIsoProd (M N : ModuleCat.{v} R) : @[simp, elementwise] theorem biprodIsoProd_inv_comp_fst (M N : ModuleCat.{v} R) : - (biprodIsoProd M N).inv ≫ biprod.fst = LinearMap.fst R M N := + (biprodIsoProd M N).inv ≫ biprod.fst = ofHom (LinearMap.fst R M N) := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk WalkingPair.left) @[simp, elementwise] theorem biprodIsoProd_inv_comp_snd (M N : ModuleCat.{v} R) : - (biprodIsoProd M N).inv ≫ biprod.snd = LinearMap.snd R M N := + (biprodIsoProd M N).inv ≫ biprod.snd = ofHom (LinearMap.snd R M N) := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk WalkingPair.right) namespace HasLimit @@ -85,14 +87,15 @@ variable {J : Type w} (f : J → ModuleCat.{max w v} R) to the cartesian product of those groups. -/ @[simps] -def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) where - toFun x j := s.π.app ⟨j⟩ x - map_add' x y := by - simp only [Functor.const_obj_obj, map_add] - rfl - map_smul' r x := by - simp only [Functor.const_obj_obj, map_smul] - rfl +def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) := + ofHom + { toFun := fun x j => s.π.app ⟨j⟩ x + map_add' := fun x y => by + simp only [Functor.const_obj_obj, map_add] + rfl + map_smul' := fun r x => by + simp only [Functor.const_obj_obj, map_smul] + rfl } /-- Construct limit data for a product in `ModuleCat R`, using `ModuleCat.of R (∀ j, F.obj j)`. -/ @@ -100,13 +103,12 @@ def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) where def productLimitCone : Limits.LimitCone (Discrete.functor f) where cone := { pt := ModuleCat.of R (∀ j, f j) - π := Discrete.natTrans fun j => (LinearMap.proj j.as : (∀ j, f j) →ₗ[R] f j.as) } + π := Discrete.natTrans fun j => ofHom (LinearMap.proj j.as : (∀ j, f j) →ₗ[R] f j.as) } isLimit := { lift := lift.{_, v} f fac := fun _ _ => rfl uniq := fun s m w => by - ext x - funext j + ext x j exact congr_arg (fun g : s.pt ⟶ f j => (g : s.pt → f j) x) (w ⟨j⟩) } end HasLimit @@ -124,7 +126,7 @@ noncomputable def biproductIsoPi [Finite J] (f : J → ModuleCat.{v} R) : @[simp, elementwise] theorem biproductIsoPi_inv_comp_π [Finite J] (f : J → ModuleCat.{v} R) (j : J) : - (biproductIsoPi f).inv ≫ biproduct.π f j = (LinearMap.proj j : (∀ j, f j) →ₗ[R] f j) := + (biproductIsoPi f).inv ≫ biproduct.π f j = ofHom (LinearMap.proj j : (∀ j, f j) →ₗ[R] f j) := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk j) end ModuleCat @@ -144,8 +146,8 @@ of modules. -/ noncomputable def lequivProdOfRightSplitExact {f : B →ₗ[R] M} (hj : Function.Injective j) (exac : LinearMap.range j = LinearMap.ker g) (h : g.comp f = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfSection _ - (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.asHom j) - (ModuleCat.asHom g) exac) (asHom f) h + (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) + (ModuleCat.ofHom g) exac) (ofHom f) (hom_ext h) (by simpa only [ModuleCat.mono_iff_injective])).isoBinaryBiproduct ≪≫ biprodIsoProd _ _ ).symm.toLinearEquiv @@ -154,8 +156,8 @@ of modules. -/ noncomputable def lequivProdOfLeftSplitExact {f : M →ₗ[R] A} (hg : Function.Surjective g) (exac : LinearMap.range j = LinearMap.ker g) (h : f.comp j = LinearMap.id) : (A × B) ≃ₗ[R] M := ((ShortComplex.Splitting.ofExactOfRetraction _ - (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.asHom j) - (ModuleCat.asHom g) exac) (ModuleCat.asHom f) h + (ShortComplex.Exact.moduleCat_of_range_eq_ker (ModuleCat.ofHom j) + (ModuleCat.ofHom g) exac) (ModuleCat.ofHom f) (hom_ext h) (by simpa only [ModuleCat.epi_iff_surjective] using hg)).isoBinaryBiproduct ≪≫ biprodIsoProd _ _).symm.toLinearEquiv diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 4bbcee39a959e..1d40a7cc808f9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -56,15 +56,18 @@ variable (M : ModuleCat.{v} S) /-- Any `S`-module M is also an `R`-module via a ring homomorphism `f : R ⟶ S` by defining `r • m := f r • m` (`Module.compHom`). This is called restriction of scalars. -/ -def obj' : ModuleCat R where - carrier := M - isModule := Module.compHom M f +def obj' : ModuleCat R := + let _ := Module.compHom M f + of R M /-- Given an `S`-linear map `g : M → M'` between `S`-modules, `g` is also `R`-linear between `M` and `M'` by means of restriction of scalars. -/ def map' {M M' : ModuleCat.{v} S} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := - { g with map_smul' := fun r => g.map_smul (f r) } + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)`. + -- This suggests `RestrictScalars.obj'` needs to be redesigned. + ofHom (X := obj' f M) (Y := obj' f M') + { g.hom with map_smul' := fun r => g.hom.map_smul (f r) } end RestrictScalars @@ -76,24 +79,21 @@ def restrictScalars {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R → ModuleCat.{v} S ⥤ ModuleCat.{v} R where obj := RestrictScalars.obj' f map := RestrictScalars.map' f - map_id _ := LinearMap.ext fun _ => rfl - map_comp _ _ := LinearMap.ext fun _ => rfl instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars.{v} f).Faithful where - map_injective h := - LinearMap.ext fun x => by simpa only using DFunLike.congr_fun h x + map_injective h := by + ext x + simpa only using DFunLike.congr_fun (ModuleCat.hom_ext_iff.mp h) x instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars.{v} f).PreservesMonomorphisms where preserves _ h := by rwa [mono_iff_injective] at h ⊢ -- Porting note: this should be automatic -instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] {f : R →+* S} - {M : ModuleCat.{v} S} : Module R <| (restrictScalars f).obj M := - inferInstanceAs <| Module R <| RestrictScalars.obj' f M - --- Porting note: this should be automatic +-- TODO: this instance gives diamonds if `f : S →+* S`, see `PresheafOfModules.pushforward₀`. +-- The correct solution is probably to define explicit maps between `M` and +-- `(restrictScalars f).obj M`. instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] {f : R →+* S} {M : ModuleCat.{v} S} : Module S <| (restrictScalars f).obj M := inferInstanceAs <| Module S M @@ -105,20 +105,18 @@ theorem restrictScalars.map_apply {R : Type u₁} {S : Type u₂} [Ring R] [Ring @[simp] theorem restrictScalars.smul_def {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) - {M : ModuleCat.{v} S} (r : R) (m : (restrictScalars f).obj M) : r • m = (f r • m : M) := + {M : ModuleCat.{v} S} (r : R) (m : (restrictScalars f).obj M) : r • m = f r • show M from m := rfl theorem restrictScalars.smul_def' {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) {M : ModuleCat.{v} S} (r : R) (m : M) : - -- Porting note: clumsy way to coerce - let m' : (restrictScalars f).obj M := m - (r • m' : (restrictScalars f).obj M) = (f r • m : M) := + r • (show (restrictScalars f).obj M from m) = f r • m := rfl instance (priority := 100) sMulCommClass_mk {R : Type u₁} {S : Type u₂} [Ring R] [CommRing S] (f : R →+* S) (M : Type v) [I : AddCommGroup M] [Module S M] : - haveI : SMul R M := (RestrictScalars.obj' f (ModuleCat.mk M)).isModule.toSMul + haveI : SMul R M := (RestrictScalars.obj' f (ModuleCat.of S M)).isModule.toSMul SMulCommClass R S M := @SMulCommClass.mk R S M (_) _ fun r s m => (by simp [← mul_smul, mul_comm] : f r • s • m = s • f r • m) @@ -129,14 +127,16 @@ morphisms `M ⟶ (ModuleCat.restrictScalars f).obj N`. -/ def semilinearMapAddEquiv {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) (M : ModuleCat.{v} R) (N : ModuleCat.{v} S) : (M →ₛₗ[f] N) ≃+ (M ⟶ (ModuleCat.restrictScalars f).obj N) where - toFun g := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + toFun g := ofHom (Y := (ModuleCat.restrictScalars f).obj N) <| { toFun := g map_add' := by simp map_smul' := by simp } invFun g := { toFun := g map_add' := by simp - map_smul' := g.map_smul } + map_smul' := g.hom.map_smul } left_inv _ := rfl right_inv _ := rfl map_add' _ _ := rfl @@ -149,7 +149,7 @@ variable {R : Type u₁} [Ring R] (f : R →+* R) to `M`. -/ def restrictScalarsId'App (hf : f = RingHom.id R) (M : ModuleCat R) : (restrictScalars f).obj M ≅ M := - LinearEquiv.toModuleIso' <| + LinearEquiv.toModuleIso <| @AddEquiv.toLinearEquiv _ _ _ _ _ _ (((restrictScalars f).obj M).isModule) _ (by rfl) (fun r x ↦ by subst hf; rfl) @@ -198,7 +198,11 @@ variable {R₁ : Type u₁} {R₂ : Type u₂} {R₃ : Type u₃} [Ring R₁] [R identifies to successively restricting scalars. -/ def restrictScalarsComp'App (hgf : gf = g.comp f) (M : ModuleCat R₃) : (restrictScalars gf).obj M ≅ (restrictScalars f).obj ((restrictScalars g).obj M) := - (AddEquiv.toLinearEquiv (by rfl) (fun r x ↦ by subst hgf; rfl)).toModuleIso' + (AddEquiv.toLinearEquiv + (M := ↑((restrictScalars gf).obj M)) + (M₂ := ↑((restrictScalars f).obj ((restrictScalars g).obj M))) + (by rfl) + (fun r x ↦ by subst hgf; rfl)).toModuleIso variable (hgf : gf = g.comp f) @@ -243,10 +247,14 @@ def restrictScalarsEquivalenceOfRingEquiv {R S} [Ring R] [Ring S] (e : R ≃+* S ModuleCat S ≌ ModuleCat R where functor := ModuleCat.restrictScalars e.toRingHom inverse := ModuleCat.restrictScalars e.symm - unitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso' + unitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso + (X₁ := M) + (X₂ := (restrictScalars e.symm.toRingHom).obj ((restrictScalars e.toRingHom).obj M)) { __ := AddEquiv.refl M map_smul' := fun s m ↦ congr_arg (· • m) (e.right_inv s).symm }) (by intros; rfl) - counitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso' + counitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso + (X₁ := (restrictScalars e.toRingHom).obj ((restrictScalars e.symm.toRingHom).obj M)) + (X₂ := M) { __ := AddEquiv.refl M map_smul' := fun r _ ↦ congr_arg (· • (_ : M)) (e.left_inv r)}) (by intros; rfl) functor_unitIso_comp := by intros; rfl @@ -279,33 +287,26 @@ variable (M : ModuleCat.{v} R) /-- Extension of scalars turn an `R`-module into `S`-module by M ↦ S ⨂ M -/ def obj' : ModuleCat S := - ⟨TensorProduct R ((restrictScalars f).obj ⟨S⟩) M⟩ + of _ (TensorProduct R ((restrictScalars f).obj (of _ S)) M) /-- Extension of scalars is a functor where an `R`-module `M` is sent to `S ⊗ M` and `l : M1 ⟶ M2` is sent to `s ⊗ m ↦ s ⊗ l m` -/ def map' {M1 M2 : ModuleCat.{v} R} (l : M1 ⟶ M2) : obj' f M1 ⟶ obj' f M2 := - by-- The "by apply" part makes this require 75% fewer heartbeats to process (https://github.com/leanprover-community/mathlib4/pull/16371). - apply @LinearMap.baseChange R S M1 M2 _ _ ((algebraMap S _).comp f).toAlgebra _ _ _ _ l - -theorem map'_id {M : ModuleCat.{v} R} : map' f (𝟙 M) = 𝟙 _ := - LinearMap.ext fun x : obj' f M => by - dsimp only [map'] - rw [ModuleCat.id_apply] -- Porting note: this got put in the dsimp by mathport - induction x using TensorProduct.induction_on with - | zero => rw [map_zero] - | tmul => -- Porting note: issues with synthesizing Algebra R S - erw [@LinearMap.baseChange_tmul R S M M _ _ (_), ModuleCat.id_apply] - | add _ _ ihx ihy => rw [map_add, ihx, ihy] + ofHom (@LinearMap.baseChange R S M1 M2 _ _ ((algebraMap S _).comp f).toAlgebra _ _ _ _ l.hom) + +theorem map'_id {M : ModuleCat.{v} R} : map' f (𝟙 M) = 𝟙 _ := by + ext x + simp [map'] theorem map'_comp {M₁ M₂ M₃ : ModuleCat.{v} R} (l₁₂ : M₁ ⟶ M₂) (l₂₃ : M₂ ⟶ M₃) : - map' f (l₁₂ ≫ l₂₃) = map' f l₁₂ ≫ map' f l₂₃ := - LinearMap.ext fun x : obj' f M₁ => by - dsimp only [map'] - induction x using TensorProduct.induction_on with - | zero => rfl - | tmul => rfl - | add _ _ ihx ihy => rw [map_add, map_add, ihx, ihy] + map' f (l₁₂ ≫ l₂₃) = map' f l₁₂ ≫ map' f l₂₃ := by + ext x + dsimp only [map'] + induction x using TensorProduct.induction_on with + | zero => rfl + | tmul => rfl + | add _ _ ihx ihy => rw [map_add, map_add, ihx, ihy] end ExtendScalars @@ -352,7 +353,7 @@ variable (M : Type v) [AddCommMonoid M] [Module R M] /-- Given an `R`-module M, consider Hom(S, M) -- the `R`-linear maps between S (as an `R`-module by means of restriction of scalars) and M. `S` acts on Hom(S, M) by `s • g = x ↦ g (x • s)` -/ -instance hasSMul : SMul S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M where +instance hasSMul : SMul S <| (restrictScalars f).obj (of _ S) →ₗ[R] M where smul s g := { toFun := fun s' : S => g (s' * s : S) map_add' := fun x y : S => by dsimp; rw [add_mul, map_add] @@ -363,16 +364,16 @@ instance hasSMul : SMul S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M where erw [smul_eq_mul, smul_eq_mul, mul_assoc] } @[simp] -theorem smul_apply' (s : S) (g : (restrictScalars f).obj ⟨S⟩ →ₗ[R] M) (s' : S) : +theorem smul_apply' (s : S) (g : (restrictScalars f).obj (of _ S) →ₗ[R] M) (s' : S) : (s • g) s' = g (s' * s : S) := rfl -instance mulAction : MulAction S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M := +instance mulAction : MulAction S <| (restrictScalars f).obj (of _ S) →ₗ[R] M := { CoextendScalars.hasSMul f _ with one_smul := fun g => LinearMap.ext fun s : S => by simp mul_smul := fun (s t : S) g => LinearMap.ext fun x : S => by simp [mul_assoc] } -instance distribMulAction : DistribMulAction S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M := +instance distribMulAction : DistribMulAction S <| (restrictScalars f).obj (of _ S) →ₗ[R] M := { CoextendScalars.mulAction f _ with smul_add := fun s g h => LinearMap.ext fun _ : S => by simp smul_zero := fun _ => LinearMap.ext fun _ : S => by simp } @@ -380,7 +381,7 @@ instance distribMulAction : DistribMulAction S <| (restrictScalars f).obj ⟨S /-- `S` acts on Hom(S, M) by `s • g = x ↦ g (x • s)`, this action defines an `S`-module structure on Hom(S, M). -/ -instance isModule : Module S <| (restrictScalars f).obj ⟨S⟩ →ₗ[R] M := +instance isModule : Module S <| (restrictScalars f).obj (of _ S) →ₗ[R] M := { CoextendScalars.distribMulAction f _ with add_smul := fun s1 s2 g => LinearMap.ext fun x : S => by simp [mul_add, LinearMap.map_add] zero_smul := fun g => LinearMap.ext fun x : S => by simp [LinearMap.map_zero] } @@ -392,18 +393,19 @@ variable (M : ModuleCat.{v} R) /-- If `M` is an `R`-module, then the set of `R`-linear maps `S →ₗ[R] M` is an `S`-module with scalar multiplication defined by `s • l := x ↦ l (x • s)`-/ def obj' : ModuleCat S := - ⟨(restrictScalars f).obj ⟨S⟩ →ₗ[R] M⟩ + of _ ((restrictScalars f).obj (of _ S) →ₗ[R] M) -instance : CoeFun (obj' f M) fun _ => S → M where coe g := g.toFun +instance : CoeFun (obj' f M) fun _ => S → M := + inferInstanceAs <| CoeFun ((restrictScalars f).obj (of _ S) →ₗ[R] M) _ /-- If `M, M'` are `R`-modules, then any `R`-linear map `g : M ⟶ M'` induces an `S`-linear map `(S →ₗ[R] M) ⟶ (S →ₗ[R] M')` defined by `h ↦ g ∘ h`-/ @[simps] -def map' {M M' : ModuleCat R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' where - toFun h := g.comp h - map_add' _ _ := LinearMap.comp_add _ _ _ - map_smul' s h := LinearMap.ext fun t : S => by dsimp; rw [smul_apply',smul_apply']; simp - -- Porting note: smul_apply' not working in simp +def map' {M M' : ModuleCat R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := + ofHom + { toFun := fun h => g.hom.comp h + map_add' := fun _ _ => LinearMap.comp_add _ _ _ + map_smul' := fun s h => by ext; simp } end CoextendScalars @@ -416,14 +418,13 @@ def coextendScalars {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R → ModuleCat R ⥤ ModuleCat S where obj := CoextendScalars.obj' f map := CoextendScalars.map' f - map_id _ := LinearMap.ext fun _ => LinearMap.ext fun _ => rfl - map_comp _ _ := LinearMap.ext fun _ => LinearMap.ext fun _ => rfl + map_id _ := by ext; rfl + map_comp _ _ := by ext; rfl namespace CoextendScalars variable {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) --- Porting note: this coercion doesn't line up well with task below instance (M : ModuleCat R) : CoeFun ((coextendScalars f).obj M) fun _ => S → M := inferInstanceAs <| CoeFun (CoextendScalars.obj' f M) _ @@ -445,51 +446,50 @@ variable {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) /-- Given `R`-module X and `S`-module Y, any `g : (restrictScalars f).obj Y ⟶ X` corresponds to `Y ⟶ (coextendScalars f).obj X` by sending `y ↦ (s ↦ g (s • y))` -/ -@[simps apply_apply] def HomEquiv.fromRestriction {X : ModuleCat R} {Y : ModuleCat S} - (g : (restrictScalars f).obj Y ⟶ X) : Y ⟶ (coextendScalars f).obj X where - toFun := fun y : Y => - { toFun := fun s : S => g <| (s • y : Y) - map_add' := fun s1 s2 : S => by simp only [add_smul]; rw [LinearMap.map_add] - map_smul' := fun r (s : S) => by - -- Porting note: dsimp clears out some rw's but less eager to apply others with Lean 4 - dsimp - rw [← g.map_smul] - erw [smul_eq_mul, mul_smul] - rfl} - map_add' := fun y1 y2 : Y => - LinearMap.ext fun s : S => by - -- Porting note: double dsimp seems odd - dsimp only [id_eq, eq_mpr_eq_cast, AddHom.toFun_eq_coe, AddHom.coe_mk, RingHom.id_apply, - RingHom.toMonoidHom_eq_coe, OneHom.toFun_eq_coe, - MonoidHom.toOneHom_coe, MonoidHom.coe_coe, ZeroHom.coe_mk, smul_eq_mul, cast_eq, - LinearMap.coe_mk] - rw [LinearMap.add_apply, LinearMap.coe_mk, LinearMap.coe_mk] - dsimp only [AddHom.coe_mk] - rw [smul_add, map_add] - map_smul' := fun (s : S) (y : Y) => LinearMap.ext fun t : S => by - -- Porting note: used to be simp [mul_smul] - simp only [LinearMap.coe_mk, AddHom.coe_mk, RingHom.id_apply] - rw [ModuleCat.CoextendScalars.smul_apply', LinearMap.coe_mk] - dsimp - rw [mul_smul] + (g : (restrictScalars f).obj Y ⟶ X) : Y ⟶ (coextendScalars f).obj X := + ofHom + { toFun := fun y : Y => + { toFun := fun s : S => g <| (s • y : Y) + map_add' := fun s1 s2 : S => by simp only [add_smul]; rw [LinearMap.map_add] + map_smul' := fun r (s : S) => by + -- Porting note: dsimp clears out some rw's but less eager to apply others with Lean 4 + dsimp + rw [← g.hom.map_smul] + erw [smul_eq_mul, mul_smul] + rfl } + map_add' := fun y1 y2 : Y => + LinearMap.ext fun s : S => by + simp [smul_add, map_add] + map_smul' := fun (s : S) (y : Y) => LinearMap.ext fun t : S => by + simp [mul_smul] } + +/-- This should be autogenerated by `@[simps]` but we need to give `s` the correct type here. -/ +@[simp] lemma HomEquiv.fromRestriction_hom_apply_apply {X : ModuleCat R} {Y : ModuleCat S} + (g : (restrictScalars f).obj Y ⟶ X) (y) (s : S) : + (HomEquiv.fromRestriction f g).hom y s = g (s • y) := rfl /-- Given `R`-module X and `S`-module Y, any `g : Y ⟶ (coextendScalars f).obj X` corresponds to `(restrictScalars f).obj Y ⟶ X` by `y ↦ g y 1` -/ -@[simps apply] def HomEquiv.toRestriction {X Y} (g : Y ⟶ (coextendScalars f).obj X) : - (restrictScalars f).obj Y ⟶ X where - toFun := fun y : Y => (g y) (1 : S) - map_add' x y := by dsimp; rw [g.map_add, LinearMap.add_apply] - map_smul' r (y : Y) := by - dsimp - rw [← LinearMap.map_smul] - erw [smul_eq_mul, mul_one, LinearMap.map_smul] - -- Porting note: should probably change CoeFun for obj above - rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] - rw [CoextendScalars.smul_apply (s := f r) (g := g y) (s' := 1), one_mul] - simp + (restrictScalars f).obj Y ⟶ X := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ofHom (X := (restrictScalars f).obj Y) + { toFun := fun y : Y => (g y) (1 : S) + map_add' := fun x y => by dsimp; rw [g.hom.map_add, LinearMap.add_apply] + map_smul' := fun r (y : Y) => by + dsimp + rw [← LinearMap.map_smul] + erw [smul_eq_mul, mul_one, LinearMap.map_smul] + rw [CoextendScalars.smul_apply (s := f r) (g := g y) (s' := 1), one_mul] + simp } + +/-- This should be autogenerated by `@[simps]` but we need to give `1` the correct type here. -/ +@[simp] lemma HomEquiv.toRestriction_hom_apply {X : ModuleCat R} {Y : ModuleCat S} + (g : Y ⟶ (coextendScalars f).obj X) (y) : + (HomEquiv.toRestriction f g).hom y = g.hom y (1 : S) := rfl -- Porting note: add to address timeout in unit' /-- Auxiliary definition for `unit'` -/ @@ -521,17 +521,14 @@ def app' (Y : ModuleCat S) : Y →ₗ[S] (restrictScalars f ⋙ coextendScalars The natural transformation from identity functor to the composition of restriction and coextension of scalars. -/ --- @[simps] Porting note: not in normal form and not used +@[simps] protected def unit' : 𝟭 (ModuleCat S) ⟶ restrictScalars f ⋙ coextendScalars f where - app Y := app' f Y + app Y := ofHom (app' f Y) naturality Y Y' g := - LinearMap.ext fun y : Y => LinearMap.ext fun s : S => by + hom_ext <| LinearMap.ext fun y : Y => LinearMap.ext fun s : S => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): previously simp [CoextendScalars.map_apply] - simp only [ModuleCat.coe_comp, Functor.id_map, Functor.id_obj, Functor.comp_obj, - Functor.comp_map] - rw [coe_comp, coe_comp, Function.comp, Function.comp] - conv_rhs => rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] - rw [CoextendScalars.map_apply, AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, + simp only [ModuleCat.hom_comp, Functor.id_map, Functor.id_obj, Functor.comp_obj, + Functor.comp_map, LinearMap.coe_comp, Function.comp, CoextendScalars.map_apply, restrictScalars.map_apply f] change s • (g y) = g (s • y) rw [map_smul] @@ -539,24 +536,21 @@ protected def unit' : 𝟭 (ModuleCat S) ⟶ restrictScalars f ⋙ coextendScala /-- The natural transformation from the composition of coextension and restriction of scalars to identity functor. -/ --- @[simps] Porting note: not in normal form and not used +@[simps] protected def counit' : coextendScalars f ⋙ restrictScalars f ⟶ 𝟭 (ModuleCat R) where - app X := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + app X := ofHom (X := (restrictScalars f).obj ((coextendScalars f).obj X)) { toFun := fun g => g.toFun (1 : S) map_add' := fun x1 x2 => by dsimp rw [LinearMap.add_apply] map_smul' := fun r (g : (restrictScalars f).obj ((coextendScalars f).obj X)) => by dsimp - rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] - rw [CoextendScalars.smul_apply (s := f r) (g := g) (s' := 1), one_mul, ← LinearMap.map_smul] - rw [← LinearMap.coe_toAddHom, ← AddHom.toFun_eq_coe] + rw [CoextendScalars.smul_apply, one_mul, ← LinearMap.map_smul] congr change f r = (f r) • (1 : S) rw [smul_eq_mul (a := f r) (a' := 1), mul_one] } - naturality X X' g := LinearMap.ext fun h => by - rw [ModuleCat.coe_comp] - rfl end RestrictionCoextensionAdj @@ -569,26 +563,18 @@ def restrictCoextendScalarsAdj {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] homEquiv := fun X Y ↦ { toFun := RestrictionCoextensionAdj.HomEquiv.fromRestriction.{u₁,u₂,v} f invFun := RestrictionCoextensionAdj.HomEquiv.toRestriction.{u₁,u₂,v} f - left_inv := fun g => LinearMap.ext fun x : X => by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): once just simp - rw [RestrictionCoextensionAdj.HomEquiv.toRestriction_apply, AddHom.toFun_eq_coe, - LinearMap.coe_toAddHom, RestrictionCoextensionAdj.HomEquiv.fromRestriction_apply_apply, - one_smul] - right_inv := fun g => LinearMap.ext fun x => LinearMap.ext fun s : S => by + left_inv := fun g => by ext; simp + right_inv := fun g => hom_ext <| LinearMap.ext fun x => LinearMap.ext fun s : S => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): once just simp - rw [RestrictionCoextensionAdj.HomEquiv.fromRestriction_apply_apply, - RestrictionCoextensionAdj.HomEquiv.toRestriction_apply, AddHom.toFun_eq_coe, - LinearMap.coe_toAddHom, LinearMap.map_smulₛₗ, RingHom.id_apply, - CoextendScalars.smul_apply', one_mul] } + rw [RestrictionCoextensionAdj.HomEquiv.fromRestriction_hom_apply_apply, + RestrictionCoextensionAdj.HomEquiv.toRestriction_hom_apply, LinearMap.map_smulₛₗ, + RingHom.id_apply, CoextendScalars.smul_apply', one_mul] } unit := RestrictionCoextensionAdj.unit'.{u₁,u₂,v} f counit := RestrictionCoextensionAdj.counit'.{u₁,u₂,v} f - homEquiv_unit := LinearMap.ext fun _ => rfl - homEquiv_counit := fun {X Y g} => LinearMap.ext <| by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10745): previously simp [RestrictionCoextensionAdj.counit'] - intro x; dsimp - rw [coe_comp, Function.comp] - change _ = (((restrictScalars f).map g) x).toFun (1 : S) - rw [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, restrictScalars.map_apply] } + homEquiv_unit := hom_ext <| LinearMap.ext fun _ => rfl + homEquiv_counit := fun {X Y g} => by + ext + simp [RestrictionCoextensionAdj.counit'] } instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars.{max u₂ w} f).IsLeftAdjoint := @@ -611,17 +597,20 @@ Given `R`-module X and `S`-module Y and a map `g : (extendScalars f).obj X ⟶ Y map `S ⨂ X → Y`, there is a `X ⟶ (restrictScalars f).obj Y`, i.e. `R`-linear map `X ⟶ Y` by `x ↦ g (1 ⊗ x)`. -/ -@[simps apply] +@[simps hom_apply] def HomEquiv.toRestrictScalars {X Y} (g : (extendScalars f).obj X ⟶ Y) : - X ⟶ (restrictScalars f).obj Y where - toFun x := g <| (1 : S)⊗ₜ[R,f]x - map_add' _ _ := by dsimp; rw [tmul_add, map_add] - map_smul' r x := by - letI : Module R S := Module.compHom S f - letI : Module R Y := Module.compHom Y f - dsimp - erw [RestrictScalars.smul_def, ← LinearMap.map_smul, tmul_smul] - congr + X ⟶ (restrictScalars f).obj Y := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ofHom (Y := (restrictScalars f).obj Y) + { toFun := fun x => g <| (1 : S)⊗ₜ[R,f]x + map_add' := fun _ _ => by dsimp; rw [tmul_add, map_add] + map_smul' := fun r s => by + letI : Module R S := Module.compHom S f + letI : Module R Y := Module.compHom Y f + dsimp + erw [RestrictScalars.smul_def, ← LinearMap.map_smul, tmul_smul] + congr } -- Porting note: forced to break apart fromExtendScalars due to timeouts /-- @@ -636,7 +625,7 @@ def HomEquiv.evalAt {X : ModuleCat R} {Y : ModuleCat S} (s : S) map_add' := by intros dsimp only - rw [map_add,smul_add] } + rw [map_add, smul_add] } (by intros r x rw [AddHom.toFun_eq_coe, AddHom.coe_mk, RingHom.id_apply, @@ -647,11 +636,11 @@ Given `R`-module X and `S`-module Y and a map `X ⟶ (restrictScalars f).obj Y`, `X ⟶ Y`, there is a map `(extend_scalars f).obj X ⟶ Y`, i.e `S`-linear map `S ⨂ X → Y` by `s ⊗ x ↦ s • g x`. -/ -@[simps apply] +@[simps hom_apply] def HomEquiv.fromExtendScalars {X Y} (g : X ⟶ (restrictScalars f).obj Y) : (extendScalars f).obj X ⟶ Y := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f - refine {toFun := fun z => TensorProduct.lift ?_ z, map_add' := ?_, map_smul' := ?_} + refine ofHom {toFun := fun z => TensorProduct.lift ?_ z, map_add' := ?_, map_smul' := ?_} · refine {toFun := fun s => HomEquiv.evalAt f s g, map_add' := fun (s₁ s₂ : S) => ?_, map_smul' := fun (r : R) (s : S) => ?_} @@ -685,6 +674,7 @@ def homEquiv {X Y} : invFun := HomEquiv.fromExtendScalars.{u₁,u₂,v} f left_inv g := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f + apply hom_ext apply LinearMap.ext; intro z induction z using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] @@ -698,9 +688,11 @@ def homEquiv {X Y} : | add _ _ ih1 ih2 => rw [map_add, map_add, ih1, ih2] right_inv g := by letI m1 : Module R S := Module.compHom S f; letI m2 : Module R Y := Module.compHom Y f - apply LinearMap.ext; intro x - rw [HomEquiv.toRestrictScalars_apply, HomEquiv.fromExtendScalars_apply, lift.tmul, - LinearMap.coe_mk, LinearMap.coe_mk] + ext x + rw [HomEquiv.toRestrictScalars_hom_apply] + -- This needs to be `erw` because of some unfolding in `fromExtendScalars` + erw [HomEquiv.fromExtendScalars_hom_apply] + rw [lift.tmul, LinearMap.coe_mk, LinearMap.coe_mk] dsimp rw [one_smul] @@ -708,13 +700,16 @@ def homEquiv {X Y} : For any `R`-module X, there is a natural `R`-linear map from `X` to `X ⨂ S` by sending `x ↦ x ⊗ 1` -/ -- @[simps] Porting note: not in normal form and not used -def Unit.map {X} : X ⟶ (extendScalars f ⋙ restrictScalars f).obj X where - toFun x := (1 : S)⊗ₜ[R,f]x - map_add' x x' := by dsimp; rw [TensorProduct.tmul_add] - map_smul' r x := by - letI m1 : Module R S := Module.compHom S f - -- Porting note: used to be rfl - dsimp; rw [← TensorProduct.smul_tmul,TensorProduct.smul_tmul'] +def Unit.map {X} : X ⟶ (extendScalars f ⋙ restrictScalars f).obj X := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ofHom (Y := (extendScalars f ⋙ restrictScalars f).obj X) + { toFun := fun x => (1 : S)⊗ₜ[R,f]x + map_add' := fun x x' => by dsimp; rw [TensorProduct.tmul_add] + map_smul' := fun r x => by + letI m1 : Module R S := Module.compHom S f + -- Porting note: used to be rfl + dsimp; rw [← TensorProduct.smul_tmul,TensorProduct.smul_tmul'] } /-- The natural transformation from identity functor on `R`-module to the composition of extension and @@ -726,45 +721,41 @@ def unit : 𝟭 (ModuleCat R) ⟶ extendScalars f ⋙ restrictScalars.{max v u /-- For any `S`-module Y, there is a natural `R`-linear map from `S ⨂ Y` to `Y` by `s ⊗ y ↦ s • y` -/ -@[simps apply] -def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y where - toFun := - letI m1 : Module R S := Module.compHom S f - letI m2 : Module R Y := Module.compHom Y f - TensorProduct.lift - { toFun := fun s : S => - { toFun := fun y : Y => s • y, - map_add' := smul_add _ - map_smul' := fun r y => by - change s • f r • y = f r • s • y - rw [← mul_smul, mul_comm, mul_smul] }, - map_add' := fun s₁ s₂ => by - ext y - change (s₁ + s₂) • y = s₁ • y + s₂ • y - rw [add_smul] - map_smul' := fun r s => by - ext y - change (f r • s) • y = (f r) • s • y - rw [smul_eq_mul, mul_smul] } - map_add' _ _ := by rw [map_add] - map_smul' s z := by - letI m1 : Module R S := Module.compHom S f - letI m2 : Module R Y := Module.compHom Y f - dsimp only - induction z using TensorProduct.induction_on with - | zero => rw [smul_zero, map_zero, smul_zero] - | tmul s' y => - rw [ExtendScalars.smul_tmul, LinearMap.coe_mk] - erw [TensorProduct.lift.tmul, TensorProduct.lift.tmul] - set s' : S := s' - change (s * s') • y = s • s' • y - rw [mul_smul] - | add _ _ ih1 ih2 => rw [smul_add, map_add, map_add, ih1, ih2, smul_add] - --- Porting note: this file has to probably be reworked when --- coercions and instance synthesis are fixed for concrete categories --- so I say nolint now and move on -attribute [nolint simpNF] Counit.map_apply +@[simps hom_apply] +def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y := + ofHom + { toFun := + letI m1 : Module R S := Module.compHom S f + letI m2 : Module R Y := Module.compHom Y f + TensorProduct.lift + { toFun := fun s : S => + { toFun := fun y : Y => s • y, + map_add' := smul_add _ + map_smul' := fun r y => by + change s • f r • y = f r • s • y + rw [← mul_smul, mul_comm, mul_smul] }, + map_add' := fun s₁ s₂ => by + ext y + change (s₁ + s₂) • y = s₁ • y + s₂ • y + rw [add_smul] + map_smul' := fun r s => by + ext y + change (f r • s) • y = (f r) • s • y + rw [smul_eq_mul, mul_smul] } + map_add' := fun _ _ => by rw [map_add] + map_smul' := fun s z => by + letI m1 : Module R S := Module.compHom S f + letI m2 : Module R Y := Module.compHom Y f + dsimp only + induction z using TensorProduct.induction_on with + | zero => rw [smul_zero, map_zero, smul_zero] + | tmul s' y => + rw [ExtendScalars.smul_tmul, LinearMap.coe_mk] + erw [TensorProduct.lift.tmul, TensorProduct.lift.tmul] + set s' : S := s' + change (s * s') • y = s • s' • y + rw [mul_smul] + | add _ _ ih1 ih2 => rw [smul_add, map_add, map_add, ih1, ih2, smul_add] } /-- The natural transformation from the composition of restriction and extension of scalars to the identity functor on `S`-module. @@ -777,15 +768,13 @@ def counit : restrictScalars.{max v u₂,u₁,u₂} f ⋙ extendScalars f ⟶ letI m1 : Module R S := Module.compHom S f letI m2 : Module R Y := Module.compHom Y f letI m2 : Module R Y' := Module.compHom Y' f - apply LinearMap.ext; intro z + ext z induction z using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] | tmul s' y => dsimp - rw [ModuleCat.coe_comp, ModuleCat.coe_comp, Function.comp_apply, Function.comp_apply, - ExtendScalars.map_tmul, restrictScalars.map_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [Counit.map_apply] + erw [Counit.map_hom_apply] rw [lift.tmul, LinearMap.coe_mk, LinearMap.coe_mk] set s' : S := s' change s' • g y = g (s' • y) @@ -803,20 +792,18 @@ def extendRestrictScalarsAdj {R : Type u₁} {S : Type u₂} [CommRing R] [CommR homEquiv := fun _ _ ↦ ExtendRestrictScalarsAdj.homEquiv.{v,u₁,u₂} f unit := ExtendRestrictScalarsAdj.unit.{v,u₁,u₂} f counit := ExtendRestrictScalarsAdj.counit.{v,u₁,u₂} f - homEquiv_unit := fun {X Y g} ↦ LinearMap.ext fun x => by + homEquiv_unit := fun {X Y g} ↦ hom_ext <| LinearMap.ext fun x => by dsimp - rw [ModuleCat.coe_comp, Function.comp_apply, restrictScalars.map_apply] rfl - homEquiv_counit := fun {X Y g} ↦ LinearMap.ext fun x => by + homEquiv_counit := fun {X Y g} ↦ hom_ext <| LinearMap.ext fun x => by induction x using TensorProduct.induction_on with | zero => rw [map_zero, map_zero] | tmul => rw [ExtendRestrictScalarsAdj.homEquiv_symm_apply] dsimp - rw [ModuleCat.coe_comp, Function.comp_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [ExtendRestrictScalarsAdj.Counit.map_apply] - dsimp + erw [ExtendRestrictScalarsAdj.Counit.map_hom_apply, + ExtendRestrictScalarsAdj.HomEquiv.fromExtendScalars_hom_apply] | add => rw [map_add, map_add]; congr 1 } instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) : diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean index 323f048f1af4a..cb4c1a9e98df5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Basic.lean @@ -112,12 +112,13 @@ noncomputable abbrev d (b : B) : KaehlerDifferential f := (D f).d b lemma ext {M : ModuleCat B} {α β : KaehlerDifferential f ⟶ M} (h : ∀ (b : B), α (d b) = β (d b)) : α = β := by rw [← sub_eq_zero] - have : ⊤ ≤ LinearMap.ker (α - β) := by + have : ⊤ ≤ LinearMap.ker (α - β).hom := by rw [← KaehlerDifferential.span_range_derivation, Submodule.span_le] rintro _ ⟨y, rfl⟩ - rw [SetLike.mem_coe, LinearMap.mem_ker, LinearMap.sub_apply, sub_eq_zero] + rw [SetLike.mem_coe, LinearMap.mem_ker, ModuleCat.hom_sub, LinearMap.sub_apply, sub_eq_zero] apply h rw [top_le_iff, LinearMap.ker_eq_top] at this + ext : 1 exact this /-- The map `KaehlerDifferential f ⟶ (ModuleCat.restrictScalars g').obj (KaehlerDifferential f')` @@ -133,6 +134,9 @@ noncomputable def map : letI := (g ≫ f').toAlgebra have : IsScalarTower A A' B' := IsScalarTower.of_algebraMap_eq' rfl have := IsScalarTower.of_algebraMap_eq' fac + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + ModuleCat.ofHom (Y := (ModuleCat.restrictScalars g').obj (KaehlerDifferential f')) { toFun := fun x ↦ _root_.KaehlerDifferential.map A A' B B' x map_add' := by simp map_smul' := by simp } @@ -163,7 +167,7 @@ morphism `CommRingCat.KaehlerDifferential f ⟶ M`. -/ noncomputable def desc : CommRingCat.KaehlerDifferential f ⟶ M := letI := f.toAlgebra letI := Module.compHom M f - D.liftKaehlerDifferential + ofHom D.liftKaehlerDifferential @[simp] lemma desc_d (b : B) : D.desc (CommRingCat.KaehlerDifferential.d b) = D.d b := by diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index 5ea1401c768c2..9df2e8a46993e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -77,7 +77,7 @@ variable (d : M.Derivation φ) /-- The postcomposition of a derivation by a morphism of presheaves of modules. -/ @[simps! d_apply] def postcomp (f : M ⟶ N) : N.Derivation φ where - d := (f.app _).toAddMonoidHom.comp d.d + d := (f.app _).hom.toAddMonoidHom.comp d.d d_map {X Y} g x := by simpa using naturality_apply f g (d.d x) d_app {X} a := by dsimp @@ -191,7 +191,7 @@ attribute [simp] relativeDifferentials'_obj lemma relativeDifferentials'_map_d {X Y : Dᵒᵖ} (f : X ⟶ Y) (x : R.obj X) : DFunLike.coe (α := CommRingCat.KaehlerDifferential (φ'.app X)) (β := fun _ ↦ CommRingCat.KaehlerDifferential (φ'.app Y)) - ((relativeDifferentials' φ').map f) (CommRingCat.KaehlerDifferential.d x) = + ((relativeDifferentials' φ').map f).hom (CommRingCat.KaehlerDifferential.d x) = CommRingCat.KaehlerDifferential.d (R.map f x) := CommRingCat.KaehlerDifferential.map_d (φ'.naturality f) _ diff --git a/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean b/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean index 4ba5ef78289de..c6db67923b764 100644 --- a/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean +++ b/Mathlib/Algebra/Category/ModuleCat/EpiMono.lean @@ -24,22 +24,24 @@ namespace ModuleCat variable {R : Type u} [Ring R] {X Y : ModuleCat.{v} R} (f : X ⟶ Y) variable {M : Type v} [AddCommGroup M] [Module R M] -theorem ker_eq_bot_of_mono [Mono f] : LinearMap.ker f = ⊥ := - LinearMap.ker_eq_bot_of_cancel fun u v => (@cancel_mono _ _ _ _ _ f _ (↟u) (↟v)).1 +theorem ker_eq_bot_of_mono [Mono f] : LinearMap.ker f.hom = ⊥ := + LinearMap.ker_eq_bot_of_cancel fun u v h => ModuleCat.hom_ext_iff.mp <| + (@cancel_mono _ _ _ _ _ f _ (↟u) (↟v)).1 <| ModuleCat.hom_ext_iff.mpr h -theorem range_eq_top_of_epi [Epi f] : LinearMap.range f = ⊤ := - LinearMap.range_eq_top_of_cancel fun u v => (@cancel_epi _ _ _ _ _ f _ (↟u) (↟v)).1 +theorem range_eq_top_of_epi [Epi f] : LinearMap.range f.hom = ⊤ := + LinearMap.range_eq_top_of_cancel fun u v h => ModuleCat.hom_ext_iff.mp <| + (@cancel_epi _ _ _ _ _ f _ (↟u) (↟v)).1 <| ModuleCat.hom_ext_iff.mpr h -theorem mono_iff_ker_eq_bot : Mono f ↔ LinearMap.ker f = ⊥ := +theorem mono_iff_ker_eq_bot : Mono f ↔ LinearMap.ker f.hom = ⊥ := ⟨fun _ => ker_eq_bot_of_mono _, fun hf => ConcreteCategory.mono_of_injective _ <| by convert LinearMap.ker_eq_bot.1 hf⟩ theorem mono_iff_injective : Mono f ↔ Function.Injective f := by rw [mono_iff_ker_eq_bot, LinearMap.ker_eq_bot] -theorem epi_iff_range_eq_top : Epi f ↔ LinearMap.range f = ⊤ := +theorem epi_iff_range_eq_top : Epi f ↔ LinearMap.range f.hom = ⊤ := ⟨fun _ => range_eq_top_of_epi _, fun hf => - ConcreteCategory.epi_of_surjective _ <| LinearMap.range_eq_top.1 hf⟩ + ConcreteCategory.epi_of_surjective _ <| by convert LinearMap.range_eq_top.1 hf⟩ theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := by rw [epi_iff_range_eq_top, LinearMap.range_eq_top] @@ -48,10 +50,10 @@ theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := by def uniqueOfEpiZero (X) [h : Epi (0 : X ⟶ of R M)] : Unique M := uniqueOfSurjectiveZero X ((ModuleCat.epi_iff_surjective _).mp h) -instance mono_as_hom'_subtype (U : Submodule R X) : Mono (ModuleCat.asHomRight U.subtype) := +instance mono_as_hom'_subtype (U : Submodule R X) : Mono (ModuleCat.ofHom U.subtype) := (mono_iff_ker_eq_bot _).mpr (Submodule.ker_subtype U) -instance epi_as_hom''_mkQ (U : Submodule R X) : Epi (↿U.mkQ) := +instance epi_as_hom''_mkQ (U : Submodule R X) : Epi (ModuleCat.ofHom U.mkQ) := (epi_iff_range_eq_top _).mpr <| Submodule.range_mkQ _ instance forget_preservesEpimorphisms : (forget (ModuleCat.{v} R)).PreservesEpimorphisms where diff --git a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean index 553ce84f4002e..08ca1ad384315 100644 --- a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean @@ -132,7 +132,8 @@ def colimit : ModuleCatMax.{v, u, u} R := /-- The linear map from a given `R`-module in the diagram to the colimit module. -/ def coconeMorphism (j : J) : F.obj j ⟶ colimit F := - { (AddCommGrp.FilteredColimits.colimitCocone + ofHom + { (AddCommGrp.FilteredColimits.colimitCocone (F ⋙ forget₂ (ModuleCat R) AddCommGrp.{max v u})).ι.app j with map_smul' := fun r x => by erw [colimit_smul_mk_eq F r ⟨j, x⟩]; rfl } @@ -142,7 +143,7 @@ def colimitCocone : Cocone F where ι := { app := coconeMorphism F naturality := fun _ _' f => - LinearMap.coe_injective + hom_ext <| LinearMap.coe_injective ((Types.TypeMax.colimitCocone (F ⋙ forget (ModuleCat R))).ι.naturality f) } /-- Given a cocone `t` of `F`, the induced monoid linear map from the colimit to the cocone point. @@ -150,25 +151,27 @@ We already know that this is a morphism between additive groups. The only thing it is a linear map, i.e. preserves scalar multiplication. -/ def colimitDesc (t : Cocone F) : colimit F ⟶ t.pt := - { (AddCommGrp.FilteredColimits.colimitCoconeIsColimit + ofHom + { (AddCommGrp.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ (ModuleCatMax.{v, u} R) AddCommGrp.{max v u})).desc ((forget₂ (ModuleCat R) AddCommGrp.{max v u}).mapCocone t) with map_smul' := fun r x => by refine Quot.inductionOn x ?_; clear x; intro x; obtain ⟨j, x⟩ := x erw [colimit_smul_mk_eq] - exact LinearMap.map_smul (t.ι.app j) r x } + exact LinearMap.map_smul (t.ι.app j).hom r x } /-- The proposed colimit cocone is a colimit in `ModuleCat R`. -/ def colimitCoconeIsColimit : IsColimit (colimitCocone F) where desc := colimitDesc F fac t j := - LinearMap.coe_injective <| + hom_ext <| LinearMap.coe_injective <| (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget (ModuleCat R))).fac ((forget (ModuleCat R)).mapCocone t) j uniq t _ h := - LinearMap.coe_injective <| + hom_ext <| LinearMap.coe_injective <| (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget (ModuleCat R))).uniq - ((forget (ModuleCat R)).mapCocone t) _ fun j => funext fun x => LinearMap.congr_fun (h j) x + ((forget (ModuleCat R)).mapCocone t) _ fun j => funext fun x => LinearMap.congr_fun + (ModuleCat.hom_ext_iff.mp (h j)) x instance forget₂AddCommGroup_preservesFilteredColimits : PreservesFilteredColimits (forget₂ (ModuleCat.{u} R) AddCommGrp.{u}) where diff --git a/Mathlib/Algebra/Category/ModuleCat/Free.lean b/Mathlib/Algebra/Category/ModuleCat/Free.lean index 9f822325b9425..27fcdedc44ea4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Free.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Free.lean @@ -48,7 +48,7 @@ theorem disjoint_span_sum : Disjoint (span R (range (u ∘ Sum.inl))) (span R (range (u ∘ Sum.inr))) := by rw [huv, disjoint_comm] refine Disjoint.mono_right (span_mono (range_comp_subset_range _ _)) ?_ - rw [← LinearMap.range_coe, span_eq (LinearMap.range S.f), hS.moduleCat_range_eq_ker] + rw [← LinearMap.range_coe, span_eq (LinearMap.range S.f.hom), hS.moduleCat_range_eq_ker] exact range_ker_disjoint hw include hv hm in @@ -66,8 +66,8 @@ where the top row is an exact sequence of modules and the maps on the bottom are independent. -/ theorem linearIndependent_leftExact : LinearIndependent R u := by rw [linearIndependent_sum] - refine ⟨?_, LinearIndependent.of_comp S.g hw, disjoint_span_sum hS hw huv⟩ - rw [huv, LinearMap.linearIndependent_iff S.f]; swap + refine ⟨?_, LinearIndependent.of_comp S.g.hom hw, disjoint_span_sum hS hw huv⟩ + rw [huv, LinearMap.linearIndependent_iff S.f.hom]; swap · rw [LinearMap.ker_eq_bot, ← mono_iff_injective] infer_instance exact hv @@ -78,7 +78,7 @@ include hS' hv in /-- Given a short exact sequence `0 ⟶ X₁ ⟶ X₂ ⟶ X₃ ⟶ 0` of `R`-modules and linearly independent families `v : ι → N` and `w : ι' → P`, we get a linearly independent family `ι ⊕ ι' → M` -/ theorem linearIndependent_shortExact {w : ι' → S.X₃} (hw : LinearIndependent R w) : - LinearIndependent R (Sum.elim (S.f ∘ v) (S.g.toFun.invFun ∘ w)) := by + LinearIndependent R (Sum.elim (S.f ∘ v) (S.g.hom.toFun.invFun ∘ w)) := by apply linearIndependent_leftExact hS'.exact hv _ hS'.mono_f rfl dsimp convert hw @@ -109,7 +109,7 @@ theorem span_exact {β : Type*} {u : ι ⊕ β → S.X₂} (huv : u ∘ Sum.inl rw [Finsupp.mem_span_range_iff_exists_finsupp] at hgm obtain ⟨cm, hm⟩ := hgm let m' : S.X₂ := Finsupp.sum cm fun j a ↦ a • (u (Sum.inr j)) - have hsub : m - m' ∈ LinearMap.range S.f := by + have hsub : m - m' ∈ LinearMap.range S.f.hom := by rw [hS.moduleCat_range_eq_ker] simp only [LinearMap.mem_ker, map_sub, sub_eq_zero] rw [← hm, map_finsupp_sum] @@ -138,7 +138,7 @@ include hS in families `v : ι → X₁` and `w : ι' → X₃`, we get a spanning family `ι ⊕ ι' → X₂` -/ theorem span_rightExact {w : ι' → S.X₃} (hv : ⊤ ≤ span R (range v)) (hw : ⊤ ≤ span R (range w)) (hE : Epi S.g) : - ⊤ ≤ span R (range (Sum.elim (S.f ∘ v) (S.g.toFun.invFun ∘ w))) := by + ⊤ ≤ span R (range (Sum.elim (S.f ∘ v) (S.g.hom.toFun.invFun ∘ w))) := by refine span_exact hS ?_ hv ?_ · simp only [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, Sum.elim_comp_inl] · convert hw diff --git a/Mathlib/Algebra/Category/ModuleCat/Images.lean b/Mathlib/Algebra/Category/ModuleCat/Images.lean index 9a516ef12b9b4..17f0387b646f9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Images.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Images.lean @@ -32,18 +32,18 @@ section -- implementation details of `HasImage` for ModuleCat; use the API, not these /-- The image of a morphism in `ModuleCat R` is just the bundling of `LinearMap.range f` -/ def image : ModuleCat R := - ModuleCat.of R (LinearMap.range f) + ModuleCat.of R (LinearMap.range f.hom) /-- The inclusion of `image f` into the target -/ def image.ι : image f ⟶ H := - f.range.subtype + ofHom f.hom.range.subtype instance : Mono (image.ι f) := ConcreteCategory.mono_of_injective (image.ι f) Subtype.val_injective /-- The corestriction map to the image -/ def factorThruImage : G ⟶ image f := - f.rangeRestrict + ofHom f.hom.rangeRestrict theorem image.fac : factorThruImage f ≫ image.ι f = f := rfl @@ -53,22 +53,23 @@ attribute [local simp] image.fac variable {f} /-- The universal property for the image factorisation -/ -noncomputable def image.lift (F' : MonoFactorisation f) : image f ⟶ F'.I where - toFun := (fun x => F'.e (Classical.indefiniteDescription _ x.2).1 : image f → F'.I) - map_add' x y := by - apply (mono_iff_injective F'.m).1 - · infer_instance - rw [LinearMap.map_add] - change (F'.e ≫ F'.m) _ = (F'.e ≫ F'.m) _ + (F'.e ≫ F'.m) _ - simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] - rfl - map_smul' c x := by - apply (mono_iff_injective F'.m).1 - · infer_instance - rw [LinearMap.map_smul] - change (F'.e ≫ F'.m) _ = _ • (F'.e ≫ F'.m) _ - simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] - rfl +noncomputable def image.lift (F' : MonoFactorisation f) : image f ⟶ F'.I := + ofHom + { toFun := (fun x => F'.e (Classical.indefiniteDescription _ x.2).1 : image f → F'.I) + map_add' := fun x y => by + apply (mono_iff_injective F'.m).1 + · infer_instance + rw [LinearMap.map_add] + change (F'.e ≫ F'.m) _ = (F'.e ≫ F'.m) _ + (F'.e ≫ F'.m) _ + simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] + rfl + map_smul' := fun c x => by + apply (mono_iff_injective F'.m).1 + · infer_instance + rw [LinearMap.map_smul] + change (F'.e ≫ F'.m) _ = _ • (F'.e ≫ F'.m) _ + simp_rw [F'.fac, (Classical.indefiniteDescription (fun z => f z = _) _).2] + rfl } theorem image.lift_fac (F' : MonoFactorisation f) : image.lift F' ≫ F'.m = image.ι f := by ext x @@ -92,17 +93,17 @@ noncomputable def isImage : IsImage (monoFactorisation f) where /-- The categorical image of a morphism in `ModuleCat R` agrees with the linear algebraic range. -/ noncomputable def imageIsoRange {G H : ModuleCat.{v} R} (f : G ⟶ H) : - Limits.image f ≅ ModuleCat.of R (LinearMap.range f) := + Limits.image f ≅ ModuleCat.of R (LinearMap.range f.hom) := IsImage.isoExt (Image.isImage f) (isImage f) @[simp, reassoc, elementwise] theorem imageIsoRange_inv_image_ι {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.asHom f.range.subtype := + (imageIsoRange f).inv ≫ Limits.image.ι f = ModuleCat.ofHom f.hom.range.subtype := IsImage.isoExt_inv_m _ _ @[simp, reassoc, elementwise] theorem imageIsoRange_hom_subtype {G H : ModuleCat.{v} R} (f : G ⟶ H) : - (imageIsoRange f).hom ≫ ModuleCat.asHom f.range.subtype = Limits.image.ι f := by + (imageIsoRange f).hom ≫ ModuleCat.ofHom f.hom.range.subtype = Limits.image.ι f := by rw [← imageIsoRange_inv_image_ι f, Iso.hom_inv_id_assoc] end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Injective.lean b/Mathlib/Algebra/Category/ModuleCat/Injective.lean index f579799501b6a..e11b9d7185a1f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Injective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Injective.lean @@ -20,16 +20,17 @@ namespace Module theorem injective_object_of_injective_module [inj : Injective R M] : CategoryTheory.Injective (ModuleCat.of R M) where factors g f m := - have ⟨l, h⟩ := inj.out f ((ModuleCat.mono_iff_injective f).mp m) g - ⟨l, LinearMap.ext h⟩ + have ⟨l, h⟩ := inj.out f.hom ((ModuleCat.mono_iff_injective f).mp m) g.hom + ⟨ModuleCat.ofHom l, by ext x; simpa using h x⟩ theorem injective_module_of_injective_object [inj : CategoryTheory.Injective <| ModuleCat.of R M] : Module.Injective R M where out X Y _ _ _ _ f hf g := by - have : CategoryTheory.Mono (ModuleCat.asHom f) := (ModuleCat.mono_iff_injective _).mpr hf - obtain ⟨l, rfl⟩ := inj.factors (ModuleCat.asHom g) (ModuleCat.asHom f) - exact ⟨l, fun _ ↦ rfl⟩ + have : CategoryTheory.Mono (ModuleCat.ofHom f) := (ModuleCat.mono_iff_injective _).mpr hf + obtain ⟨l, h⟩ := inj.factors (ModuleCat.ofHom g) (ModuleCat.ofHom f) + obtain rfl := ModuleCat.hom_ext_iff.mp h + exact ⟨l.hom, fun _ => rfl⟩ theorem injective_iff_injective_object : Module.Injective R M ↔ diff --git a/Mathlib/Algebra/Category/ModuleCat/Kernels.lean b/Mathlib/Algebra/Category/ModuleCat/Kernels.lean index 0a06c74a2fe68..6f2398e63fd4f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Kernels.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Kernels.lean @@ -26,40 +26,38 @@ variable {M N : ModuleCat.{v} R} (f : M ⟶ N) /-- The kernel cone induced by the concrete kernel. -/ def kernelCone : KernelFork f := -- Porting note: previously proven by tidy - KernelFork.ofι (asHom f.ker.subtype) <| by ext x; cases x; assumption + KernelFork.ofι (ofHom f.hom.ker.subtype) <| by ext x; cases x; assumption /-- The kernel of a linear map is a kernel in the categorical sense. -/ def kernelIsLimit : IsLimit (kernelCone f) := Fork.IsLimit.mk _ - (fun s => + (fun s => ofHom <| -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation on LinearMap.ker - LinearMap.codRestrict (LinearMap.ker f) (Fork.ι s) fun c => + LinearMap.codRestrict (LinearMap.ker f.hom) (Fork.ι s).hom fun c => LinearMap.mem_ker.2 <| by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← @Function.comp_apply _ _ _ f (Fork.ι s) c, ← coe_comp] - rw [Fork.condition, HasZeroMorphisms.comp_zero (Fork.ι s) N] + erw [← @Function.comp_apply _ _ _ f (Fork.ι s) c, ← LinearMap.coe_comp] + rw [← ModuleCat.hom_comp, Fork.condition, HasZeroMorphisms.comp_zero (Fork.ι s) N] rfl) - (fun _ => LinearMap.subtype_comp_codRestrict _ _ _) fun s m h => - LinearMap.ext fun x => Subtype.ext_iff_val.2 (by simp [← h]; rfl) + (fun _ => hom_ext <| LinearMap.subtype_comp_codRestrict _ _ _) fun s m h => + hom_ext <| LinearMap.ext fun x => Subtype.ext_iff_val.2 (by simp [← h]; rfl) /-- The cokernel cocone induced by the projection onto the quotient. -/ def cokernelCocone : CokernelCofork f := - CokernelCofork.ofπ (asHom f.range.mkQ) <| LinearMap.range_mkQ_comp _ + CokernelCofork.ofπ (ofHom f.hom.range.mkQ) <| hom_ext <| LinearMap.range_mkQ_comp _ /-- The projection onto the quotient is a cokernel in the categorical sense. -/ def cokernelIsColimit : IsColimit (cokernelCocone f) := Cofork.IsColimit.mk _ - (fun s => - f.range.liftQ (Cofork.π s) <| LinearMap.range_le_ker_iff.2 <| CokernelCofork.condition s) - (fun s => f.range.liftQ_mkQ (Cofork.π s) _) fun s m h => by + (fun s => ofHom <| f.hom.range.liftQ (Cofork.π s).hom <| + LinearMap.range_le_ker_iff.2 <| ModuleCat.hom_ext_iff.mp <| CokernelCofork.condition s) + (fun s => hom_ext <| f.hom.range.liftQ_mkQ (Cofork.π s).hom _) fun s m h => by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - haveI : Epi (asHom (LinearMap.range f).mkQ) := + haveI : Epi (ofHom (LinearMap.range f.hom).mkQ) := (epi_iff_range_eq_top _).mpr (Submodule.range_mkQ _) -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - apply (cancel_epi (asHom (LinearMap.range f).mkQ)).1 - convert h - -- Porting note: no longer necessary - -- exact Submodule.liftQ_mkQ _ _ _ + apply (cancel_epi (ofHom (LinearMap.range f.hom).mkQ)).1 + exact h end @@ -84,20 +82,20 @@ agrees with the usual module-theoretical kernel. -/ noncomputable def kernelIsoKer {G H : ModuleCat.{v} R} (f : G ⟶ H) : -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - kernel f ≅ ModuleCat.of R (LinearMap.ker f) := + kernel f ≅ ModuleCat.of R (LinearMap.ker f.hom) := limit.isoLimitCone ⟨_, kernelIsLimit f⟩ -- We now show this isomorphism commutes with the inclusion of the kernel into the source. @[simp, elementwise] -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation theorem kernelIsoKer_inv_kernel_ι : (kernelIsoKer f).inv ≫ kernel.ι f = - (LinearMap.ker f).subtype := + ofHom (LinearMap.ker f.hom).subtype := limit.isoLimitCone_inv_π _ _ @[simp, elementwise] theorem kernelIsoKer_hom_ker_subtype : -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - (kernelIsoKer f).hom ≫ (LinearMap.ker f).subtype = kernel.ι f := + (kernelIsoKer f).hom ≫ ofHom (LinearMap.ker f.hom).subtype = kernel.ι f := IsLimit.conePointUniqueUpToIso_inv_comp _ (limit.isLimit _) WalkingParallelPair.zero /-- The categorical cokernel of a morphism in `ModuleCat` @@ -105,18 +103,18 @@ agrees with the usual module-theoretical quotient. -/ noncomputable def cokernelIsoRangeQuotient {G H : ModuleCat.{v} R} (f : G ⟶ H) : -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11036): broken dot notation - cokernel f ≅ ModuleCat.of R (H ⧸ LinearMap.range f) := + cokernel f ≅ ModuleCat.of R (H ⧸ LinearMap.range f.hom) := colimit.isoColimitCocone ⟨_, cokernelIsColimit f⟩ -- We now show this isomorphism commutes with the projection of target to the cokernel. @[simp, elementwise] theorem cokernel_π_cokernelIsoRangeQuotient_hom : - cokernel.π f ≫ (cokernelIsoRangeQuotient f).hom = f.range.mkQ := + cokernel.π f ≫ (cokernelIsoRangeQuotient f).hom = ofHom f.hom.range.mkQ := colimit.isoColimitCocone_ι_hom _ _ @[simp, elementwise] theorem range_mkQ_cokernelIsoRangeQuotient_inv : - ↿f.range.mkQ ≫ (cokernelIsoRangeQuotient f).inv = cokernel.π f := + ofHom f.hom.range.mkQ ≫ (cokernelIsoRangeQuotient f).inv = cokernel.π f := colimit.isoColimitCocone_ι_inv ⟨_, cokernelIsColimit f⟩ WalkingParallelPair.one theorem cokernel_π_ext {M N : ModuleCat.{u} R} (f : M ⟶ N) {x y : N} (m : M) (w : x = y + f m) : diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index 862547dd00ba8..607ca41120826 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -45,9 +45,7 @@ def sectionsSubmodule : Submodule R (∀ j, F.obj j) := forget₂ AddCommGrp AddGrp.{w}) with carrier := (F ⋙ forget (ModuleCat R)).sections smul_mem' := fun r s sh j j' f => by - simp only [forget_map, Functor.comp_map, Pi.smul_apply, map_smul] - dsimp [Functor.sections] at sh - rw [sh f] } + simpa [Functor.sections, forget_map] using congr_arg (r • ·) (sh f) } instance : AddCommMonoid (F ⋙ forget (ModuleCat R)).sections := inferInstanceAs <| AddCommMonoid (sectionsSubmodule F) @@ -101,16 +99,16 @@ namespace HasLimits def limitCone : Cone F where pt := ModuleCat.of R (Types.Small.limitCone.{v, w} (F ⋙ forget _)).pt π := - { app := limitπLinearMap F - naturality := fun _ _ f => - LinearMap.coe_injective ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } + { app := fun j => ofHom (limitπLinearMap F j) + naturality := fun _ _ f => hom_ext <| LinearMap.coe_injective <| + ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } /-- Witness that the limit cone in `ModuleCat R` is a limit cone. (Internal use only; use the limits API.) -/ def limitConeIsLimit : IsLimit (limitCone.{t, v, w} F) := by refine IsLimit.ofFaithful (forget (ModuleCat R)) (Types.Small.limitConeIsLimit.{v, w} _) - (fun s => ⟨⟨(Types.Small.limitConeIsLimit.{v, w} _).lift + (fun s => ofHom ⟨⟨(Types.Small.limitConeIsLimit.{v, w} _).lift ((forget (ModuleCat R)).mapCone s), ?_⟩, ?_⟩) (fun s => rfl) · intro x y @@ -216,14 +214,12 @@ variable (f : ∀ i j, i ≤ j → G i →ₗ[R] G j) [DirectedSystem G fun i j @[simps] def directLimitDiagram : ι ⥤ ModuleCat R where obj i := ModuleCat.of R (G i) - map hij := f _ _ hij.le + map hij := ofHom (f _ _ hij.le) map_id i := by - apply LinearMap.ext - intro x + ext apply Module.DirectedSystem.map_self map_comp hij hjk := by - apply LinearMap.ext - intro x + ext symm apply Module.DirectedSystem.map_map @@ -237,35 +233,33 @@ In `directLimitIsColimit` we show that it is a colimit cocone. -/ def directLimitCocone : Cocone (directLimitDiagram G f) where pt := ModuleCat.of R <| DirectLimit G f ι := - { app := Module.DirectLimit.of R ι G f + { app := fun x => ofHom (Module.DirectLimit.of R ι G f x) naturality := fun _ _ hij => by - apply LinearMap.ext - intro x + ext exact DirectLimit.of_f } /-- The unbundled `directLimit` of modules is a colimit in the sense of `CategoryTheory`. -/ @[simps] def directLimitIsColimit [IsDirected ι (· ≤ ·)] : IsColimit (directLimitCocone G f) where - desc s := - DirectLimit.lift R ι G f s.ι.app fun i j h x => by + desc s := ofHom <| + DirectLimit.lift R ι G f (fun i => (s.ι.app i).hom) fun i j h x => by + simp only [Functor.const_obj_obj] rw [← s.w (homOfLE h)] rfl fac s i := by - apply LinearMap.ext - intro x + ext dsimp only [directLimitCocone, CategoryStruct.comp] rw [LinearMap.comp_apply] apply DirectLimit.lift_of uniq s m h := by have : s.ι.app = fun i => - LinearMap.comp m (DirectLimit.of R ι (fun i => G i) (fun i j H => f i j H) i) := by + (ofHom (DirectLimit.of R ι (fun i => G i) (fun i j H => f i j H) i)) ≫ m := by funext i rw [← h] rfl - apply LinearMap.ext - intro x + ext simp only [this] apply Module.DirectLimit.lift_unique diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index d6e565553b5e4..89bdba339ff9f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -54,25 +54,27 @@ def tensorObj (M N : ModuleCat R) : ModuleCat R := /-- (implementation) tensor product of morphisms R-modules -/ def tensorHom {M N M' N' : ModuleCat R} (f : M ⟶ N) (g : M' ⟶ N') : tensorObj M M' ⟶ tensorObj N N' := - TensorProduct.map f g + ofHom <| TensorProduct.map f.hom g.hom /-- (implementation) left whiskering for R-modules -/ def whiskerLeft (M : ModuleCat R) {N₁ N₂ : ModuleCat R} (f : N₁ ⟶ N₂) : tensorObj M N₁ ⟶ tensorObj M N₂ := - f.lTensor M + ofHom <| f.hom.lTensor M /-- (implementation) right whiskering for R-modules -/ def whiskerRight {M₁ M₂ : ModuleCat R} (f : M₁ ⟶ M₂) (N : ModuleCat R) : tensorObj M₁ N ⟶ tensorObj M₂ N := - f.rTensor N + ofHom <| f.hom.rTensor N theorem tensor_id (M N : ModuleCat R) : tensorHom (𝟙 M) (𝟙 N) = 𝟙 (ModuleCat.of R (M ⊗ N)) := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): even with high priority `ext` fails to find this. apply TensorProduct.ext rfl theorem tensor_comp {X₁ Y₁ Z₁ X₂ Y₂ Z₂ : ModuleCat R} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (g₁ : Y₁ ⟶ Z₁) (g₂ : Y₂ ⟶ Z₂) : tensorHom (f₁ ≫ g₁) (f₂ ≫ g₂) = tensorHom f₁ f₂ ≫ tensorHom g₁ g₂ := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): even with high priority `ext` fails to find this. apply TensorProduct.ext rfl @@ -95,63 +97,36 @@ instance instMonoidalCategoryStruct : MonoidalCategoryStruct (ModuleCat.{u} R) w tensorObj := tensorObj whiskerLeft := whiskerLeft whiskerRight := whiskerRight - tensorHom f g := TensorProduct.map f g + tensorHom f g := ofHom <| TensorProduct.map f.hom g.hom tensorUnit := ModuleCat.of R R associator := associator leftUnitor := leftUnitor rightUnitor := rightUnitor -section - -/-! The `associator_naturality` and `pentagon` lemmas below are very slow to elaborate. - -We give them some help by expressing the lemmas first non-categorically, then using -`convert _aux using 1` to have the elaborator work as little as possible. -/ - - -open TensorProduct (assoc map) - -private theorem associator_naturality_aux {X₁ X₂ X₃ : Type*} [AddCommMonoid X₁] [AddCommMonoid X₂] - [AddCommMonoid X₃] [Module R X₁] [Module R X₂] [Module R X₃] {Y₁ Y₂ Y₃ : Type*} - [AddCommMonoid Y₁] [AddCommMonoid Y₂] [AddCommMonoid Y₃] [Module R Y₁] [Module R Y₂] - [Module R Y₃] (f₁ : X₁ →ₗ[R] Y₁) (f₂ : X₂ →ₗ[R] Y₂) (f₃ : X₃ →ₗ[R] Y₃) : - ↑(assoc R Y₁ Y₂ Y₃) ∘ₗ map (map f₁ f₂) f₃ = map f₁ (map f₂ f₃) ∘ₗ ↑(assoc R X₁ X₂ X₃) := by - apply TensorProduct.ext_threefold - intro x y z - rfl - -variable (R) - -private theorem pentagon_aux (W X Y Z : Type*) [AddCommMonoid W] [AddCommMonoid X] - [AddCommMonoid Y] [AddCommMonoid Z] [Module R W] [Module R X] [Module R Y] [Module R Z] : - (((assoc R X Y Z).toLinearMap.lTensor W).comp - (assoc R W (X ⊗[R] Y) Z).toLinearMap).comp - ((assoc R W X Y).toLinearMap.rTensor Z) = - (assoc R W X (Y ⊗[R] Z)).toLinearMap.comp (assoc R (W ⊗[R] X) Y Z).toLinearMap := by - apply TensorProduct.ext_fourfold - intro w x y z - rfl - -end - theorem associator_naturality {X₁ X₂ X₃ Y₁ Y₂ Y₃ : ModuleCat R} (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (f₃ : X₃ ⟶ Y₃) : tensorHom (tensorHom f₁ f₂) f₃ ≫ (associator Y₁ Y₂ Y₃).hom = (associator X₁ X₂ X₃).hom ≫ tensorHom f₁ (tensorHom f₂ f₃) := by - convert associator_naturality_aux f₁ f₂ f₃ using 1 + ext : 1 + apply TensorProduct.ext_threefold + intro x y z + rfl theorem pentagon (W X Y Z : ModuleCat R) : whiskerRight (associator W X Y).hom Z ≫ (associator W (tensorObj X Y) Z).hom ≫ whiskerLeft W (associator X Y Z).hom = (associator (tensorObj W X) Y Z).hom ≫ (associator W X (tensorObj Y Z)).hom := by - convert pentagon_aux R W X Y Z using 1 + ext : 1 + apply TensorProduct.ext_fourfold + intro w x y z + rfl theorem leftUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : tensorHom (𝟙 (ModuleCat.of R R)) f ≫ (leftUnitor N).hom = (leftUnitor M).hom ≫ f := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): broken ext apply TensorProduct.ext - apply LinearMap.ext_ring - apply LinearMap.ext; intro x + ext x -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10934): used to be dsimp change ((leftUnitor N).hom) ((tensorHom (𝟙 (of R R)) f) ((1 : R) ⊗ₜ[R] x)) = f (((leftUnitor M).hom) (1 ⊗ₜ[R] x)) @@ -161,13 +136,11 @@ theorem leftUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : theorem rightUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : tensorHom f (𝟙 (ModuleCat.of R R)) ≫ (rightUnitor N).hom = (rightUnitor M).hom ≫ f := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): broken ext apply TensorProduct.ext - apply LinearMap.ext; intro x - apply LinearMap.ext_ring - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10934): used to be dsimp - change ((rightUnitor N).hom) ((tensorHom f (𝟙 (of R R))) (x ⊗ₜ[R] (1 : R))) = - f (((rightUnitor M).hom) (x ⊗ₜ[R] 1)) + ext x + dsimp erw [TensorProduct.rid_tmul, TensorProduct.rid_tmul] rw [LinearMap.map_smul] rfl @@ -175,9 +148,9 @@ theorem rightUnitor_naturality {M N : ModuleCat R} (f : M ⟶ N) : theorem triangle (M N : ModuleCat.{u} R) : (associator M (ModuleCat.of R R) N).hom ≫ tensorHom (𝟙 M) (leftUnitor N).hom = tensorHom (rightUnitor M).hom (𝟙 N) := by + ext : 1 apply TensorProduct.ext_threefold intro x y z - change R at y -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10934): used to be dsimp [tensorHom, associator] change x ⊗ₜ[R] ((leftUnitor N).hom) (y ⊗ₜ[R] z) = ((rightUnitor M).hom) (x ⊗ₜ[R] y) ⊗ₜ[R] z erw [TensorProduct.lid_tmul, TensorProduct.rid_tmul] @@ -263,7 +236,7 @@ variable (f : M₁ → M₂ → M₃) (h₁ : ∀ m₁ m₂ n, f (m₁ + m₂) n /-- Construct for morphisms from the tensor product of two objects in `ModuleCat`. -/ noncomputable def tensorLift : M₁ ⊗ M₂ ⟶ M₃ := - TensorProduct.lift (LinearMap.mk₂ R f h₁ h₂ h₃ h₄) + ofHom <| TensorProduct.lift (LinearMap.mk₂ R f h₁ h₂ h₃ h₄) @[simp] lemma tensorLift_tmul (m : M₁) (n : M₂) : @@ -273,13 +246,13 @@ end lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f (m ⊗ₜ n) = g (m ⊗ₜ n)) : f = g := - TensorProduct.ext (by ext; apply h) + hom_ext <| TensorProduct.ext (by ext; apply h) /-- Extensionality lemma for morphisms from a module of the form `(M₁ ⊗ M₂) ⊗ M₃`. -/ lemma tensor_ext₃' {f g : (M₁ ⊗ M₂) ⊗ M₃ ⟶ M₄} (h : ∀ m₁ m₂ m₃, f (m₁ ⊗ₜ m₂ ⊗ₜ m₃) = g (m₁ ⊗ₜ m₂ ⊗ₜ m₃)) : f = g := - TensorProduct.ext_threefold h + hom_ext <| TensorProduct.ext_threefold h /-- Extensionality lemma for morphisms from a module of the form `M₁ ⊗ (M₂ ⊗ M₃)`. -/ lemma tensor_ext₃ {f g : M₁ ⊗ (M₂ ⊗ M₃) ⟶ M₄} @@ -292,56 +265,57 @@ end MonoidalCategory open Opposite --- Porting note: simp wasn't firing but rw was, annoying instance : MonoidalPreadditive (ModuleCat.{u} R) := by refine ⟨?_, ?_, ?_, ?_⟩ · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_zero, LinearMap.zero_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerLeft_apply] - rw [LinearMap.zero_apply, TensorProduct.tmul_zero] + simp · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.zero_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_zero, LinearMap.zero_apply, ] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerRight_apply] - rw [LinearMap.zero_apply, TensorProduct.zero_tmul] + simp · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.add_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_add, LinearMap.add_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerLeft_apply, MonoidalCategory.whiskerLeft_apply] erw [MonoidalCategory.whiskerLeft_apply] - rw [LinearMap.add_apply, TensorProduct.tmul_add] + simp [TensorProduct.tmul_add] · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.add_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_add, LinearMap.add_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerRight_apply, MonoidalCategory.whiskerRight_apply] erw [MonoidalCategory.whiskerRight_apply] - rw [LinearMap.add_apply, TensorProduct.add_tmul] + simp [TensorProduct.add_tmul] --- Porting note: simp wasn't firing but rw was, annoying instance : MonoidalLinear R (ModuleCat.{u} R) := by refine ⟨?_, ?_⟩ · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_smul, LinearMap.smul_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerLeft_apply, MonoidalCategory.whiskerLeft_apply] - rw [LinearMap.smul_apply, TensorProduct.tmul_smul] + simp · intros + ext : 1 refine TensorProduct.ext (LinearMap.ext fun x => LinearMap.ext fun y => ?_) - simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply] - rw [LinearMap.smul_apply] + simp only [LinearMap.compr₂_apply, TensorProduct.mk_apply, hom_smul, LinearMap.smul_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [MonoidalCategory.whiskerRight_apply, MonoidalCategory.whiskerRight_apply] - rw [LinearMap.smul_apply, TensorProduct.smul_tmul, TensorProduct.tmul_smul] + simp [TensorProduct.smul_tmul, TensorProduct.tmul_smul] + +@[simp] lemma ofHom₂_compr₂ {M N P Q : ModuleCat.{u} R} (f : M →ₗ[R] N →ₗ[R] P) (g : P →ₗ[R] Q): + ofHom₂ (f.compr₂ g) = ofHom₂ f ≫ ofHom (Linear.rightComp R _ (ofHom g)) := rfl end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean index 3b8592dcc5ce7..61228a85b17e0 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean @@ -27,15 +27,13 @@ variable {R : Type u} [CommRing R] def monoidalClosedHomEquiv (M N P : ModuleCat.{u} R) : ((MonoidalCategory.tensorLeft M).obj N ⟶ P) ≃ (N ⟶ ((linearCoyoneda R (ModuleCat R)).obj (op M)).obj P) where - toFun f := LinearMap.compr₂ (TensorProduct.mk R N M) ((β_ N M).hom ≫ f) - invFun f := (β_ M N).hom ≫ TensorProduct.lift f + toFun f := ofHom₂ <| LinearMap.compr₂ (TensorProduct.mk R N M) ((β_ N M).hom ≫ f).hom + invFun f := (β_ M N).hom ≫ ofHom (TensorProduct.lift f.hom₂) left_inv f := by + ext : 1 apply TensorProduct.ext' intro m n - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [coe_comp] - rw [Function.comp_apply] - -- This used to be `rw` and was longer (?), but we need `erw` after https://github.com/leanprover/lean4/pull/2644 + simp only [Hom.hom₂_ofHom₂, LinearMap.comp_apply, hom_comp, MonoidalCategory.tensorLeft_obj] erw [MonoidalCategory.braiding_hom_apply, TensorProduct.lift.tmul] right_inv _ := rfl @@ -47,6 +45,7 @@ instance : MonoidalClosed (ModuleCat.{u} R) where -- Porting note: this proof was automatic in mathlib3 homEquiv_naturality_left_symm := by intros + ext : 1 apply TensorProduct.ext' intro m n rfl } } @@ -57,26 +56,24 @@ theorem ihom_map_apply {M N P : ModuleCat.{u} R} (f : N ⟶ P) (g : ModuleCat.of open MonoidalCategory --- Porting note: `CoeFun` was replaced by `DFunLike` --- I can't seem to express the function coercion here without writing `@DFunLike.coe`. theorem monoidalClosed_curry {M N P : ModuleCat.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) : - @DFunLike.coe _ _ _ LinearMap.instFunLike - ((MonoidalClosed.curry f : N →ₗ[R] M →ₗ[R] P) y) x = f (x ⊗ₜ[R] y) := + ((MonoidalClosed.curry f).hom y).hom x = f (x ⊗ₜ[R] y) := rfl @[simp] theorem monoidalClosed_uncurry {M N P : ModuleCat.{u} R} (f : N ⟶ M ⟶[ModuleCat.{u} R] P) (x : M) (y : N) : - MonoidalClosed.uncurry f (x ⊗ₜ[R] y) = - @DFunLike.coe _ _ _ LinearMap.instFunLike (f y) x := + MonoidalClosed.uncurry f (x ⊗ₜ[R] y) = (f y).hom x := rfl /-- Describes the counit of the adjunction `M ⊗ - ⊣ Hom(M, -)`. Given an `R`-module `N` this should give a map `M ⊗ Hom(M, N) ⟶ N`, so we flip the order of the arguments in the identity map `Hom(M, N) ⟶ (M ⟶ N)` and uncurry the resulting map `M ⟶ Hom(M, N) ⟶ N.` -/ theorem ihom_ev_app (M N : ModuleCat.{u} R) : - (ihom.ev M).app N = TensorProduct.uncurry _ _ _ _ LinearMap.id.flip := by + (ihom.ev M).app N = ModuleCat.ofHom (TensorProduct.uncurry R M ((ihom M).obj N) N + (LinearMap.lcomp _ _ homLinearEquiv.toLinearMap ∘ₗ LinearMap.id.flip)) := by rw [← MonoidalClosed.uncurry_id_eq_ev] + ext : 1 apply TensorProduct.ext' apply monoidalClosed_uncurry @@ -84,11 +81,12 @@ theorem ihom_ev_app (M N : ModuleCat.{u} R) : define a map `N ⟶ Hom(M, M ⊗ N)`, which is given by flipping the arguments in the natural `R`-bilinear map `M ⟶ N ⟶ M ⊗ N`. -/ theorem ihom_coev_app (M N : ModuleCat.{u} R) : - (ihom.coev M).app N = (TensorProduct.mk _ _ _).flip := + (ihom.coev M).app N = ModuleCat.ofHom₂ (TensorProduct.mk _ _ _).flip := rfl theorem monoidalClosed_pre_app {M N : ModuleCat.{u} R} (P : ModuleCat.{u} R) (f : N ⟶ M) : - (MonoidalClosed.pre f).app P = LinearMap.lcomp R _ f := + (MonoidalClosed.pre f).app P = ofHom (homLinearEquiv.symm.toLinearMap ∘ₗ + LinearMap.lcomp _ _ f.hom ∘ₗ homLinearEquiv.toLinearMap) := rfl end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean index d6b5c46f9dba1..42370456d2060 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Symmetric.lean @@ -29,6 +29,7 @@ namespace MonoidalCategory @[simp] theorem braiding_naturality {X₁ X₂ Y₁ Y₂ : ModuleCat.{u} R} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂) : (f ⊗ g) ≫ (Y₁.braiding Y₂).hom = (X₁.braiding X₂).hom ≫ (g ⊗ f) := by + ext : 1 apply TensorProduct.ext' intro x y rfl @@ -49,6 +50,7 @@ theorem braiding_naturality_right (X : ModuleCat R) {Y Z : ModuleCat R} (f : Y theorem hexagon_forward (X Y Z : ModuleCat.{u} R) : (α_ X Y Z).hom ≫ (braiding X _).hom ≫ (α_ Y Z X).hom = (braiding X Y).hom ▷ Z ≫ (α_ Y X Z).hom ≫ Y ◁ (braiding X Z).hom := by + ext : 1 apply TensorProduct.ext_threefold intro x y z rfl @@ -58,6 +60,7 @@ theorem hexagon_reverse (X Y Z : ModuleCat.{u} R) : (α_ X Y Z).inv ≫ (braiding _ Z).hom ≫ (α_ Z X Y).inv = X ◁ (Y.braiding Z).hom ≫ (α_ X Z Y).inv ≫ (X.braiding Z).hom ▷ Y := by apply (cancel_epi (α_ X Y Z).hom).1 + ext : 1 apply TensorProduct.ext_threefold intro x y z rfl @@ -74,6 +77,7 @@ instance symmetricCategory : SymmetricCategory (ModuleCat.{u} R) where -- Porting note: this proof was automatic in Lean3 -- now `aesop` is applying `ModuleCat.ext` in favour of `TensorProduct.ext`. symmetry _ _ := by + ext : 1 apply TensorProduct.ext' aesop_cat @@ -88,8 +92,8 @@ theorem braiding_inv_apply {M N : ModuleCat.{u} R} (m : M) (n : N) : rfl theorem tensorμ_eq_tensorTensorTensorComm {A B C D : ModuleCat R} : - tensorμ A B C D = (TensorProduct.tensorTensorTensorComm R A B C D).toLinearMap := - TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext₂ fun _ _ => + tensorμ A B C D = ofHom (TensorProduct.tensorTensorTensorComm R A B C D).toLinearMap := + ModuleCat.hom_ext <| TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext₂ fun _ _ => TensorProduct.ext <| LinearMap.ext₂ fun _ _ => rfl @[simp] diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index 653280c6b57bd..8755d8d2ac2f7 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -161,7 +161,10 @@ when the preferred constructor `PresheafOfModules.mk` is not as convenient as th @[simps] def ofPresheaf : PresheafOfModules.{v} R where obj X := ModuleCat.of _ (M.obj X) - map f := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + map {X Y} f := ModuleCat.ofHom + (Y := (ModuleCat.restrictScalars (R.map f)).obj (ModuleCat.of _ (M.obj Y))) { toFun := fun x ↦ M.map f x map_add' := by simp map_smul' := fun r m ↦ map_smul f r m } @@ -178,7 +181,7 @@ which satisfy a suitable linearity condition. -/ def homMk (φ : M₁.presheaf ⟶ M₂.presheaf) (hφ : ∀ (X : Cᵒᵖ) (r : R.obj X) (m : M₁.obj X), φ.app X (r • m) = r • φ.app X m) : M₁ ⟶ M₂ where - app X := + app X := ModuleCat.ofHom { toFun := φ.app X map_add' := by simp map_smul' := hφ X } @@ -197,30 +200,21 @@ instance : Neg (M₁ ⟶ M₂) where { app := fun X ↦ -f.app X naturality := fun {X Y} h ↦ by ext x - dsimp - erw [map_neg] - rw [← naturality_apply] - rfl } + simp [← naturality_apply] } instance : Add (M₁ ⟶ M₂) where add f g := { app := fun X ↦ f.app X + g.app X naturality := fun {X Y} h ↦ by ext x - dsimp - erw [map_add] - rw [← naturality_apply, ← naturality_apply] - rfl } + simp [← naturality_apply] } instance : Sub (M₁ ⟶ M₂) where sub f g := { app := fun X ↦ f.app X - g.app X naturality := fun {X Y} h ↦ by ext x - dsimp - erw [map_sub] - rw [← naturality_apply, ← naturality_apply] - rfl } + simp [← naturality_apply] } @[simp] lemma neg_app (f : M₁ ⟶ M₂) (X : Cᵒᵖ) : (-f).app X = -f.app X := rfl @[simp] lemma add_app (f g : M₁ ⟶ M₂) (X : Cᵒᵖ) : (f + g).app X = f.app X + g.app X := rfl @@ -267,7 +261,10 @@ noncomputable def restriction {X Y : Cᵒᵖ} (f : X ⟶ Y) : /-- The obvious free presheaf of modules of rank `1`. -/ def unit : PresheafOfModules R where obj X := ModuleCat.of _ (R.obj X) - map {X Y } f := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + map {X Y} f := ModuleCat.ofHom + (Y := (ModuleCat.restrictScalars (R.map f)).obj (ModuleCat.of (R.obj Y) (R.obj Y))) { toFun := fun x ↦ R.map f x map_add' := by simp map_smul' := by aesop_cat } @@ -321,14 +318,16 @@ def unitHomEquiv (M : PresheafOfModules R) : toFun f := sectionsMk (fun X ↦ Hom.app f X (1 : R.obj X)) (by intros; rw [← naturality_apply, unit_map_one]) invFun s := - { app := fun X ↦ (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm (s.val X) + { app := fun X ↦ ModuleCat.ofHom + ((LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm (s.val X)) naturality := fun {X Y} f ↦ by - ext (x : R.obj X) - change R.map f x • s.eval Y = M.map f (x • s.eval X) + ext + dsimp + change R.map f 1 • s.eval Y = M.map f (1 • s.eval X) simp } left_inv f := by - ext1 X - exact (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm_apply_apply (f.app X) + ext X : 2 + exact (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).symm_apply_apply (f.app X).hom right_inv s := by ext X exact (LinearMap.ringLmapEquivSelf (R.obj X) ℤ (M.obj X)).apply_symm_apply (s.val X) @@ -352,27 +351,29 @@ variable (M : PresheafOfModules.{v} R) noncomputable abbrev forgetToPresheafModuleCatObjObj (Y : Cᵒᵖ) : ModuleCat (R.obj X) := (ModuleCat.restrictScalars (R.map (hX.to Y))).obj (M.obj Y) -@[simp] +-- This should not be a `simp` lemma because `M.obj Y` is missing the `Module (R.obj X)` instance, +-- so `simp`ing breaks downstream proofs. lemma forgetToPresheafModuleCatObjObj_coe (Y : Cᵒᵖ) : (forgetToPresheafModuleCatObjObj X hX M Y : Type _) = M.obj Y := rfl /-- Auxiliary definition for `forgetToPresheafModuleCatObj`. -/ def forgetToPresheafModuleCatObjMap {Y Z : Cᵒᵖ} (f : Y ⟶ Z) : forgetToPresheafModuleCatObjObj X hX M Y ⟶ - forgetToPresheafModuleCatObjObj X hX M Z where - toFun x := M.map f x - map_add' := by simp - map_smul' r x := by - simp only [ModuleCat.restrictScalars.smul_def, AddHom.toFun_eq_coe, AddHom.coe_mk, - RingHom.id_apply, M.map_smul] - rw [← CategoryTheory.comp_apply, ← R.map_comp] - congr - apply hX.hom_ext + forgetToPresheafModuleCatObjObj X hX M Z := + ModuleCat.ofHom + (X := forgetToPresheafModuleCatObjObj X hX M Y) (Y := forgetToPresheafModuleCatObjObj X hX M Z) + { toFun := fun x => M.map f x + map_add' := by simp + map_smul' := fun r x => by + simp only [ModuleCat.restrictScalars.smul_def, AddHom.toFun_eq_coe, AddHom.coe_mk, + RingHom.id_apply, M.map_smul] + rw [← CategoryTheory.comp_apply, ← R.map_comp] + congr + apply hX.hom_ext } @[simp] lemma forgetToPresheafModuleCatObjMap_apply {Y Z : Cᵒᵖ} (f : Y ⟶ Z) (m : M.obj Y) : - DFunLike.coe (α := M.obj Y) (β := fun _ ↦ M.obj Z) - (forgetToPresheafModuleCatObjMap X hX M f) m = M.map f m := rfl + (forgetToPresheafModuleCatObjMap X hX M f).hom m = M.map f m := rfl /-- Implementation of the functor `PresheafOfModules R ⥤ Cᵒᵖ ⥤ ModuleCat (R.obj X)` @@ -402,10 +403,12 @@ morphism level `(f : M ⟶ N) ↦ (c ↦ f(c))`. noncomputable def forgetToPresheafModuleCatMap (X : Cᵒᵖ) (hX : Limits.IsInitial X) {M N : PresheafOfModules.{v} R} (f : M ⟶ N) : forgetToPresheafModuleCatObj X hX M ⟶ forgetToPresheafModuleCatObj X hX N where - app Y := + app Y := ModuleCat.ofHom + (X := (forgetToPresheafModuleCatObj X hX M).obj Y) + (Y := (forgetToPresheafModuleCatObj X hX N).obj Y) { toFun := f.app Y map_add' := by simp - map_smul' := fun r ↦ (f.app Y).map_smul (R.1.map (hX.to Y) _) } + map_smul' := fun r ↦ (f.app Y).hom.map_smul (R.1.map (hX.to Y) _) } naturality Y Z g := by ext x exact naturality_apply f g x diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean index 987578da8c709..2b6d781c0892b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/ChangeOfRings.lean @@ -28,7 +28,12 @@ variable {C : Type u'} [Category.{v'} C] {R R' : Cᵒᵖ ⥤ RingCat.{u}} noncomputable def restrictScalarsObj (M' : PresheafOfModules.{v} R') (α : R ⟶ R') : PresheafOfModules R where obj := fun X ↦ (ModuleCat.restrictScalars (α.app X)).obj (M'.obj X) - map := fun {X Y} f ↦ + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)`. + -- This suggests `restrictScalars` needs to be redesigned. + map := fun {X Y} f ↦ ModuleCat.ofHom + (X := (ModuleCat.restrictScalars (α.app X)).obj (M'.obj X)) + (Y := (ModuleCat.restrictScalars (R.map f)).obj + ((ModuleCat.restrictScalars (α.app Y)).obj (M'.obj Y))) { toFun := M'.map f map_add' := map_add _ map_smul' := fun r x ↦ (M'.map_smul f (α.app _ r) x).trans (by diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean index 2b9de6dc8f4e9..f4d23fbb2a3e3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -60,7 +60,7 @@ noncomputable def tensorObj : PresheafOfModules (R ⋙ forget₂ _ _) where intro m₁ m₂ dsimp [tensorObjMap] simp only [map_comp, Functor.comp_obj, CommRingCat.forgetToRingCat_obj, Functor.comp_map, - ModuleCat.restrictScalarsComp'_inv_app, ModuleCat.coe_comp, Function.comp_apply, + ModuleCat.restrictScalarsComp'_inv_app, ModuleCat.hom_comp, LinearMap.comp_apply, ModuleCat.restrictScalars.map_apply, ModuleCat.restrictScalarsComp'App_inv_apply] rfl) @@ -71,7 +71,7 @@ lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X : _)) (β := fun _ ↦ (ModuleCat.restrictScalars ((forget₂ CommRingCat RingCat).map (R.map f))).obj (M₁.obj Y ⊗ M₂.obj Y)) - ((tensorObj M₁ M₂).map f) (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl + ((tensorObj M₁ M₂).map f).hom (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl /-- The tensor product of two morphisms of presheaves of modules. -/ @[simps] @@ -79,10 +79,8 @@ noncomputable def tensorHom (f : M₁ ⟶ M₂) (g : M₃ ⟶ M₄) : tensorObj app X := f.app X ⊗ g.app X naturality {X Y} φ := ModuleCat.MonoidalCategory.tensor_ext (fun m₁ m₃ ↦ by dsimp - rw [tensorObj_map_tmul] - erw [ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, - naturality_apply, naturality_apply] - rfl) + rw [tensorObj_map_tmul, ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, + naturality_apply, naturality_apply]) end Monoidal diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean index ae8b9e272a34a..099ab943de7ac 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean @@ -37,10 +37,14 @@ def pushforward₀ (R : Dᵒᵖ ⥤ RingCat.{u}) : { obj := fun X ↦ ModuleCat.of _ (M.obj (F.op.obj X)) map := fun {X Y} f ↦ M.map (F.op.map f) map_id := fun X ↦ by - ext x + refine ModuleCat.hom_ext + -- Work around an instance diamond for `restrictScalarsId'` + (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) exact (M.congr_map_apply (F.op.map_id X) x).trans (by simp) map_comp := fun f g ↦ by - ext x + refine ModuleCat.hom_ext + -- Work around an instance diamond for `restrictScalarsId'` + (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) exact (M.congr_map_apply (F.op.map_comp f g) x).trans (by simp) } map {M₁ M₂} φ := { app := fun X ↦ φ.app _ } @@ -66,22 +70,29 @@ noncomputable def pushforwardCompToPresheaf : pushforward.{v} φ ⋙ toPresheaf _ ≅ toPresheaf _ ⋙ (whiskeringLeft _ _ _).obj F.op := Iso.refl _ -@[simp] lemma pushforward_obj_map_apply (M : PresheafOfModules.{v} R) {X Y : Cᵒᵖ} (f : X ⟶ Y) (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : - DFunLike.coe - (α := (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) - (β := fun _ ↦ (ModuleCat.restrictScalars (φ.app Y)).obj - (M.obj (Opposite.op (F.obj Y.unop)))) (((pushforward φ).obj M).map f) m = - M.map (F.map f.unop).op m := rfl + (((pushforward φ).obj M).map f).hom m = M.map (F.map f.unop).op m := rfl +/-- `@[simp]`-normal form of `pushforward_obj_map_apply`. -/ @[simp] +lemma pushforward_obj_map_apply' (M : PresheafOfModules.{v} R) {X Y : Cᵒᵖ} (f : X ⟶ Y) + (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + DFunLike.coe + (F := ↑((ModuleCat.restrictScalars _).obj _) →ₗ[_] + ↑((ModuleCat.restrictScalars (S.map f)).obj ((ModuleCat.restrictScalars _).obj _))) + (((pushforward φ).obj M).map f).hom m = M.map (F.map f.unop).op m := rfl + lemma pushforward_map_app_apply {M N : PresheafOfModules.{v} R} (α : M ⟶ N) (X : Cᵒᵖ) + (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : + (((pushforward φ).map α).app X).hom m = α.app (Opposite.op (F.obj X.unop)) m := rfl + +/-- `@[simp]`-normal form of `pushforward_map_app_apply`. -/ +@[simp] +lemma pushforward_map_app_apply' {M N : PresheafOfModules.{v} R} (α : M ⟶ N) (X : Cᵒᵖ) (m : (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) : DFunLike.coe - (α := (ModuleCat.restrictScalars (φ.app X)).obj (M.obj (Opposite.op (F.obj X.unop)))) - (β := fun _ ↦ (ModuleCat.restrictScalars (φ.app X)).obj - (N.obj (Opposite.op (F.obj X.unop)))) - (((pushforward φ).map α).app X) m = α.app (Opposite.op (F.obj X.unop)) m := rfl + (F := ↑((ModuleCat.restrictScalars _).obj _) →ₗ[_] ↑((ModuleCat.restrictScalars _).obj _)) + (((pushforward φ).map α).app X).hom m = α.app (Opposite.op (F.obj X.unop)) m := rfl end PresheafOfModules diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index c00b766a8362d..bcf7dd4cfc6a3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -314,10 +314,14 @@ def toSheafify : M₀ ⟶ (restrictScalars α).obj (sheafify α φ).val := simpa using (Sheafify.map_smul_eq α φ (α.app _ r₀) (φ.app _ m₀) (𝟙 _) r₀ (by aesop) m₀ (by simp)).symm) -@[simp] lemma toSheafify_app_apply (X : Cᵒᵖ) (x : M₀.obj X) : - DFunLike.coe (α := M₀.obj X) (β := fun _ ↦ A.val.obj X) - ((toSheafify α φ).app X) x = φ.app X x := rfl + ((toSheafify α φ).app X).hom x = φ.app X x := rfl + +/-- `@[simp]`-normal form of `toSheafify_app_apply`. -/ +@[simp] +lemma toSheafify_app_apply' (X : Cᵒᵖ) (x : M₀.obj X) : + DFunLike.coe (F := (_ →ₗ[_] ↑((ModuleCat.restrictScalars (α.app X)).obj _))) + ((toSheafify α φ).app X).hom x = φ.app X x := rfl @[simp] lemma toPresheaf_map_toSheafify : (toPresheaf R₀).map (toSheafify α φ) = φ := rfl @@ -370,7 +374,7 @@ def sheafifyMap (fac : (toPresheaf R₀).map τ₀ ≫ φ' = φ ≫ τ.val) : sheafify α φ ⟶ sheafify α φ' where val := homMk τ.val (fun X r m ↦ by let f := (sheafifyHomEquiv' α φ (by exact A'.cond)).symm (τ₀ ≫ toSheafify α φ') - suffices τ.val = (toPresheaf _).map f by simpa only [this] using (f.app X).map_smul r m + suffices τ.val = (toPresheaf _).map f by simpa only [this] using (f.app X).hom.map_smul r m apply ((J.W_of_isLocallyBijective φ).homEquiv _ A'.cond).injective dsimp [f] erw [comp_toPresheaf_map_sheafifyHomEquiv'_symm_hom] diff --git a/Mathlib/Algebra/Category/ModuleCat/Products.lean b/Mathlib/Algebra/Category/ModuleCat/Products.lean index 9b504d0c1dbe0..937317e6ea0a5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Products.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Products.lean @@ -28,18 +28,16 @@ section product /-- The product cone induced by the concrete product. -/ def productCone : Fan Z := - Fan.mk (ModuleCat.of R (∀ i : ι, Z i)) fun i => (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) + Fan.mk (ModuleCat.of R (∀ i : ι, Z i)) fun i => + ofHom (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) /-- The concrete product cone is limiting. -/ def productConeIsLimit : IsLimit (productCone Z) where - lift s := (LinearMap.pi fun j => s.π.app ⟨j⟩ : s.pt →ₗ[R] ∀ i : ι, Z i) - fac s j := by - cases j - aesop + lift s := ofHom (LinearMap.pi fun j => (s.π.app ⟨j⟩).hom : s.pt →ₗ[R] ∀ i : ι, Z i) uniq s m w := by ext x funext i - exact LinearMap.congr_fun (w ⟨i⟩) x + exact DFunLike.congr_fun (congr_arg Hom.hom (w ⟨i⟩)) x -- While we could use this to construct a `HasProducts (ModuleCat R)` instance, -- we already have `HasLimits (ModuleCat R)` in `Algebra.Category.ModuleCat.Limits`. @@ -54,12 +52,12 @@ noncomputable def piIsoPi : ∏ᶜ Z ≅ ModuleCat.of R (∀ i, Z i) := -- We now show this isomorphism commutes with the inclusion of the kernel into the source. @[simp, elementwise] theorem piIsoPi_inv_kernel_ι (i : ι) : - (piIsoPi Z).inv ≫ Pi.π Z i = (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) := + (piIsoPi Z).inv ≫ Pi.π Z i = ofHom (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) := limit.isoLimitCone_inv_π _ _ @[simp, elementwise] theorem piIsoPi_hom_ker_subtype (i : ι) : - (piIsoPi Z).hom ≫ (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) = Pi.π Z i := + (piIsoPi Z).hom ≫ ofHom (LinearMap.proj i : (∀ i : ι, Z i) →ₗ[R] Z i) = Pi.π Z i := IsLimit.conePointUniqueUpToIso_inv_comp _ (limit.isLimit _) (Discrete.mk i) end product @@ -72,19 +70,20 @@ variable [DecidableEq ι] /-- The coproduct cone induced by the concrete product. -/ def coproductCocone : Cofan Z := - Cofan.mk (ModuleCat.of R (⨁ i : ι, Z i)) fun i => (DirectSum.lof R ι (fun i ↦ Z i) i) + Cofan.mk (ModuleCat.of R (⨁ i : ι, Z i)) fun i => ofHom (DirectSum.lof R ι (fun i ↦ Z i) i) /-- The concrete coproduct cone is limiting. -/ def coproductCoconeIsColimit : IsColimit (coproductCocone Z) where - desc s := DirectSum.toModule R ι _ fun i ↦ s.ι.app ⟨i⟩ + desc s := ofHom <| DirectSum.toModule R ι _ fun i ↦ (s.ι.app ⟨i⟩).hom fac := by rintro s ⟨i⟩ ext (x : Z i) simpa only [Discrete.functor_obj_eq_as, coproductCocone, Cofan.mk_pt, Functor.const_obj_obj, - Cofan.mk_ι_app, coe_comp, Function.comp_apply] using + Cofan.mk_ι_app, hom_comp, LinearMap.coe_comp, Function.comp_apply] using DirectSum.toModule_lof (ι := ι) R (M := fun i ↦ Z i) i x uniq := by rintro s f h + ext : 1 refine DirectSum.linearMap_ext _ fun i ↦ ?_ ext x simpa only [LinearMap.coe_comp, Function.comp_apply, toModule_lof] using @@ -100,12 +99,12 @@ noncomputable def coprodIsoDirectSum : ∐ Z ≅ ModuleCat.of R (⨁ i, Z i) := @[simp, elementwise] theorem ι_coprodIsoDirectSum_hom (i : ι) : - Sigma.ι Z i ≫ (coprodIsoDirectSum Z).hom = DirectSum.lof R ι (fun i ↦ Z i) i := + Sigma.ι Z i ≫ (coprodIsoDirectSum Z).hom = ofHom (DirectSum.lof R ι (fun i ↦ Z i) i) := colimit.isoColimitCocone_ι_hom _ _ @[simp, elementwise] theorem lof_coprodIsoDirectSum_inv (i : ι) : - DirectSum.lof R ι (fun i ↦ Z i) i ≫ (coprodIsoDirectSum Z).inv = Sigma.ι Z i := + ofHom (DirectSum.lof R ι (fun i ↦ Z i) i) ≫ (coprodIsoDirectSum Z).inv = Sigma.ι Z i := (coproductCoconeIsColimit Z).comp_coconePointUniqueUpToIso_hom (colimit.isColimit _) _ end coproduct diff --git a/Mathlib/Algebra/Category/ModuleCat/Projective.lean b/Mathlib/Algebra/Category/ModuleCat/Projective.lean index e8bfc43488358..1947ab1f393c3 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Projective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Projective.lean @@ -30,13 +30,16 @@ theorem IsProjective.iff_projective {R : Type u} [Ring R] {P : Type max u v} [Ad [Module R P] : Module.Projective R P ↔ Projective (ModuleCat.of R P) := by refine ⟨fun h => ?_, fun h => ?_⟩ · letI : Module.Projective R (ModuleCat.of R P) := h - exact ⟨fun E X epi => Module.projective_lifting_property _ _ - ((ModuleCat.epi_iff_surjective _).mp epi)⟩ + refine ⟨fun E X epi => ?_⟩ + obtain ⟨f, h⟩ := Module.projective_lifting_property X.hom E.hom + ((ModuleCat.epi_iff_surjective _).mp epi) + exact ⟨ofHom f, hom_ext h⟩ · refine Module.Projective.of_lifting_property.{u,v} ?_ intro E X mE mX sE sX f g s haveI : Epi (↟f) := (ModuleCat.epi_iff_surjective (↟f)).mpr s letI : Projective (ModuleCat.of R P) := h - exact ⟨Projective.factorThru (↟g) (↟f), Projective.factorThru_comp (↟g) (↟f)⟩ + exact ⟨(Projective.factorThru (↟g) (↟f)).hom, + ModuleCat.hom_ext_iff.mp <| Projective.factorThru_comp (↟g) (↟f)⟩ namespace ModuleCat @@ -56,15 +59,9 @@ instance moduleCat_enoughProjectives : EnoughProjectives (ModuleCat.{max u v} R) projective := projective_of_free.{v,u} (ι := M) (M := ModuleCat.of R (M →₀ R)) <| Finsupp.basisSingleOne - f := Finsupp.basisSingleOne.constr ℕ _root_.id + f := ofHom <| Finsupp.basisSingleOne.constr ℕ _root_.id epi := (epi_iff_range_eq_top _).mpr (range_eq_top.2 fun m => ⟨Finsupp.single m (1 : R), by - -- Porting note: simp [Finsupp.linearCombination_single] fails but rw succeeds - dsimp [Basis.constr] - simp only [Finsupp.lmapDomain_id, comp_id] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [Finsupp.linearCombination_single] - rw [one_smul] - rfl ⟩) }⟩ + simp [Finsupp.linearCombination_single, Basis.constr] ⟩)}⟩ end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean index 3a23ba4b4c02e..cf467ff8b32d1 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean @@ -219,7 +219,7 @@ noncomputable def homEquivOfIsLocallyBijective : (M₂ ⟶ N) ≃ (M₁ ⟶ N) w have hφ' : ∀ (z : M₂.obj X), φ.app _ (M₂.map p.op z) = N.map p.op (φ.app _ z) := congr_fun ((forget _).congr_map (φ.naturality p.op)) change N.map p.op (φ.app X (r • y)) = N.map p.op (r • φ.app X y) - rw [← hφ', M₂.map_smul, ← hx, ← (f.app _).map_smul, hφ, (ψ.app _).map_smul, + rw [← hφ', M₂.map_smul, ← hx, ← (f.app _).hom.map_smul, hφ, (ψ.app _).hom.map_smul, ← hφ, hx, N.map_smul, hφ']) left_inv φ := (toPresheaf _).map_injective (((J.W_of_isLocallyBijective diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean index 53503eb352661..a73bfc742c383 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/ChangeOfRings.lean @@ -59,7 +59,7 @@ noncomputable def restrictHomEquivOfIsLocallySurjective change M₂.map p.op (g.app X (r' • m)) = M₂.map p.op (r' • show M₂.obj X from g.app X m) dsimp at hg ⊢ rw [← hg, M₂.map_smul, ← hg, ← hr] - erw [← (g.app _).map_smul] + erw [← (g.app _).hom.map_smul] rw [M₁.map_smul, ← hr] rfl) left_inv _ := rfl diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 0facef08450f1..3c70d7b6228ef 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -35,39 +35,37 @@ variable {R : Type u} [Ring R] (M : ModuleCat.{v} R) submodules. -/ noncomputable def subobjectModule : Subobject M ≃o Submodule R M := OrderIso.symm - { invFun := fun S => LinearMap.range S.arrow - toFun := fun N => Subobject.mk (↾N.subtype) + { invFun := fun S => LinearMap.range S.arrow.hom + toFun := fun N => Subobject.mk (ofHom N.subtype) right_inv := fun S => Eq.symm (by fapply eq_mk_of_comm - · apply LinearEquiv.toModuleIso'Left - apply LinearEquiv.ofBijective (LinearMap.codRestrict (LinearMap.range S.arrow) S.arrow _) + · apply LinearEquiv.toModuleIso + apply LinearEquiv.ofBijective (LinearMap.codRestrict + (LinearMap.range S.arrow.hom) S.arrow.hom _) constructor · simp [← LinearMap.ker_eq_bot, LinearMap.ker_codRestrict] rw [ker_eq_bot_of_mono] · rw [← LinearMap.range_eq_top, LinearMap.range_codRestrict, Submodule.comap_subtype_self] exact LinearMap.mem_range_self _ - · apply LinearMap.ext - intro x + · ext x rfl) left_inv := fun N => by - -- Porting note: The type of `↾N.subtype` was ambiguous. Not entirely sure, I made the right - -- choice here - convert congr_arg LinearMap.range - (underlyingIso_arrow (↾N.subtype : of R { x // x ∈ N } ⟶ M)) using 1 + convert congr_arg LinearMap.range (ModuleCat.hom_ext_iff.mp + (underlyingIso_arrow (ofHom N.subtype))) using 1 · have : -- Porting note: added the `.toLinearEquiv.toLinearMap` - (underlyingIso (↾N.subtype : of R _ ⟶ M)).inv = - (underlyingIso (↾N.subtype : of R _ ⟶ M)).symm.toLinearEquiv.toLinearMap := by - apply LinearMap.ext - intro x + (underlyingIso (ofHom N.subtype)).inv = + ofHom (underlyingIso (ofHom N.subtype)).symm.toLinearEquiv.toLinearMap := by + ext x rfl - rw [this, comp_def, LinearEquiv.range_comp] + rw [this, hom_comp, LinearEquiv.range_comp] · exact (Submodule.range_subtype _).symm map_rel_iff' := fun {S T} => by refine ⟨fun h => ?_, fun h => mk_le_mk_of_comm (↟(Submodule.inclusion h)) rfl⟩ - convert LinearMap.range_comp_le_range (ofMkLEMk _ _ h) (↾T.subtype) - · simpa only [← comp_def, ofMkLEMk_comp] using (Submodule.range_subtype _).symm + convert LinearMap.range_comp_le_range (ofMkLEMk _ _ h).hom (ofHom T.subtype).hom + · rw [← hom_comp, ofMkLEMk_comp] + exact (Submodule.range_subtype _).symm · exact (Submodule.range_subtype _).symm } instance wellPowered_moduleCat : WellPowered (ModuleCat.{v} R) := @@ -77,16 +75,16 @@ attribute [local instance] hasKernels_moduleCat /-- Bundle an element `m : M` such that `f m = 0` as a term of `kernelSubobject f`. -/ noncomputable def toKernelSubobject {M N : ModuleCat.{v} R} {f : M ⟶ N} : - LinearMap.ker f →ₗ[R] kernelSubobject f := - (kernelSubobjectIso f ≪≫ ModuleCat.kernelIsoKer f).inv + LinearMap.ker f.hom →ₗ[R] kernelSubobject f := + (kernelSubobjectIso f ≪≫ ModuleCat.kernelIsoKer f).inv.hom @[simp] -theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap.ker f) : +theorem toKernelSubobject_arrow {M N : ModuleCat R} {f : M ⟶ N} (x : LinearMap.ker f.hom) : (kernelSubobject f).arrow (toKernelSubobject x) = x.1 := by -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10959): the whole proof was just `simp [toKernelSubobject]`. suffices ((arrow ((kernelSubobject f))) ∘ (kernelSubobjectIso f ≪≫ kernelIsoKer f).inv) x = x by convert this - rw [Iso.trans_inv, ← coe_comp, Category.assoc] + rw [Iso.trans_inv, ← LinearMap.coe_comp, ← hom_comp, Category.assoc] simp only [Category.assoc, kernelSubobject_arrow', kernelIsoKer_inv_kernel_ι] aesop_cat @@ -108,9 +106,13 @@ theorem cokernel_π_imageSubobject_ext {L M N : ModuleCat.{v} R} (f : L ⟶ M) [ subst w -- Porting note (https://github.com/leanprover-community/mathlib4/pull/10959): The proof from here used to just be `simp`. simp only [map_add, add_right_eq_self] - change ((cokernel.π g) ∘ (g) ∘ (factorThruImageSubobject f)) l = 0 - rw [← coe_comp, ← coe_comp, Category.assoc] - simp only [cokernel.condition, comp_zero] - rfl + -- TODO: add a `@[simp]` lemma along the lines of: + -- ``` + -- lemma ModuleCat.Hom.cokernel_condition : (cokernel.π g).hom (g.hom x) = 0 + -- ``` + -- ideally generated for all concrete categories (using a metaprogram like `@[elementwise]`?). + -- See also: https://github.com/leanprover-community/mathlib4/pull/19511#discussion_r1867083077 + change ((factorThruImageSubobject f) ≫ g ≫ (cokernel.π g)).hom l = 0 + simp end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean b/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean index a22898dd2560d..febe6bfa7f0fc 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean @@ -26,10 +26,7 @@ def ringEquivEndForget₂ (R : Type u) [Ring R] : R ≃+* End (AdditiveFunctor.of (forget₂ (ModuleCat.{u} R) AddCommGrp.{u})) where toFun r := { app := fun M => - @AddCommGrp.ofHom M.carrier M.carrier _ _ (DistribMulAction.toAddMonoidHom M r) - naturality := fun M N f => by - ext - exact (f.map_smul _ _).symm } + @AddCommGrp.ofHom M.carrier M.carrier _ _ (DistribMulAction.toAddMonoidHom M r) } invFun φ := φ.app (ModuleCat.of R R) (1 : R) left_inv := by intro r @@ -39,19 +36,19 @@ def ringEquivEndForget₂ (R : Type u) [Ring R] : apply NatTrans.ext ext M (x : M) have w := congr_fun ((forget _).congr_map - (φ.naturality (ModuleCat.asHomRight (LinearMap.toSpanSingleton R M x)))) (1 : R) + (φ.naturality (ModuleCat.ofHom (LinearMap.toSpanSingleton R M x)))) (1 : R) exact w.symm.trans (congr_arg (φ.app M) (one_smul R x)) map_add' := by intros apply NatTrans.ext ext - dsimp - simp only [AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, add_smul] + simp only [AdditiveFunctor.of_fst, ModuleCat.forget₂_obj, AddCommGrp.coe_of, + AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, add_smul] rfl map_mul' := by intros apply NatTrans.ext ext - dsimp - simp only [AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, mul_smul] + simp only [AdditiveFunctor.of_fst, ModuleCat.forget₂_obj, AddCommGrp.coe_of, + AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, mul_smul] rfl diff --git a/Mathlib/Algebra/Category/Ring/Under/Limits.lean b/Mathlib/Algebra/Category/Ring/Under/Limits.lean index 810501a4eca05..29babdcccee80 100644 --- a/Mathlib/Algebra/Category/Ring/Under/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Under/Limits.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Category.Ring.Under.Basic import Mathlib.CategoryTheory.Limits.Constructions.LimitsOfProductsAndEqualizers import Mathlib.CategoryTheory.Limits.Over import Mathlib.RingTheory.TensorProduct.Pi -import Mathlib.RingTheory.Flat.Algebra +import Mathlib.RingTheory.RingHom.Flat import Mathlib.RingTheory.Flat.Equalizer /-! @@ -208,6 +208,7 @@ lemma preservesFiniteLimits_of_flat (hf : RingHom.Flat f) : PreservesFiniteLimits (Under.pushout f) where preservesFiniteLimits _ := letI : Algebra R S := RingHom.toAlgebra f + haveI : Module.Flat R S := hf preservesLimitsOfShape_of_natIso (tensorProdIsoPushout R S) end CommRingCat.Under diff --git a/Mathlib/Algebra/CharP/LocalRing.lean b/Mathlib/Algebra/CharP/LocalRing.lean index e2fb6ba42ddaa..b85b347fd0fd2 100644 --- a/Mathlib/Algebra/CharP/LocalRing.lean +++ b/Mathlib/Algebra/CharP/LocalRing.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Jon Eugster. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jon Eugster -/ +import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.IsPrimePow import Mathlib.Data.Nat.Factorization.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index 917885981d942..b6ff248ab4b41 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -1315,8 +1315,8 @@ namespace Associates variable [CancelCommMonoidWithZero α] [GCDMonoid α] instance instGCDMonoid : GCDMonoid (Associates α) where - gcd := Quotient.map₂' gcd fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.gcd hb - lcm := Quotient.map₂' lcm fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.lcm hb + gcd := Quotient.map₂ gcd fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.gcd hb + lcm := Quotient.map₂ lcm fun _ _ (ha : Associated _ _) _ _ (hb : Associated _ _) => ha.lcm hb gcd_dvd_left := by rintro ⟨a⟩ ⟨b⟩; exact mk_le_mk_of_dvd (gcd_dvd_left _ _) gcd_dvd_right := by rintro ⟨a⟩ ⟨b⟩; exact mk_le_mk_of_dvd (gcd_dvd_right _ _) dvd_gcd := by diff --git a/Mathlib/Algebra/Group/Indicator.lean b/Mathlib/Algebra/Group/Indicator.lean index 8dfed5c563fb9..604906c6e8770 100644 --- a/Mathlib/Algebra/Group/Indicator.lean +++ b/Mathlib/Algebra/Group/Indicator.lean @@ -120,9 +120,17 @@ set. -/ theorem mem_of_mulIndicator_ne_one (h : mulIndicator s f a ≠ 1) : a ∈ s := not_imp_comm.1 (fun hn => mulIndicator_of_not_mem hn f) h -@[to_additive] +/-- See `Set.eqOn_mulIndicator'` for the version with `sᶜ`. -/ +@[to_additive + "See `Set.eqOn_indicator'` for the version with `sᶜ`"] theorem eqOn_mulIndicator : EqOn (mulIndicator s f) f s := fun _ hx => mulIndicator_of_mem hx f +/-- See `Set.eqOn_mulIndicator` for the version with `s`. -/ +@[to_additive + "See `Set.eqOn_indicator` for the version with `s`."] +theorem eqOn_mulIndicator' : EqOn (mulIndicator s f) 1 sᶜ := + fun _ hx => mulIndicator_of_not_mem hx f + @[to_additive] theorem mulSupport_mulIndicator_subset : mulSupport (s.mulIndicator f) ⊆ s := fun _ hx => hx.imp_symm fun h => mulIndicator_of_not_mem h f diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index c6b11e2b701e3..f828e651230b1 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -479,6 +479,7 @@ theorem normalClosure_le_normal {N : Subgroup G} [N.Normal] (h : s ⊆ N) : norm theorem normalClosure_subset_iff {N : Subgroup G} [N.Normal] : s ⊆ N ↔ normalClosure s ≤ N := ⟨normalClosure_le_normal, Set.Subset.trans subset_normalClosure⟩ +@[gcongr] theorem normalClosure_mono {s t : Set G} (h : s ⊆ t) : normalClosure s ≤ normalClosure t := normalClosure_le_normal (Set.Subset.trans h subset_normalClosure) diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index d5827d8bec2d1..b5e3e4461137b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -202,6 +202,21 @@ theorem closure_mul_le (S T : Set G) : closure (S * T) ≤ closure S ⊔ closure (closure S ⊔ closure T).mul_mem (SetLike.le_def.mp le_sup_left <| subset_closure hs) (SetLike.le_def.mp le_sup_right <| subset_closure ht) +@[to_additive] +lemma closure_pow_le : ∀ {n}, n ≠ 0 → closure (s ^ n) ≤ closure s + | 1, _ => by simp + | n + 2, _ => + calc + closure (s ^ (n + 2)) + _ = closure (s ^ (n + 1) * s) := by rw [pow_succ] + _ ≤ closure (s ^ (n + 1)) ⊔ closure s := closure_mul_le .. + _ ≤ closure s ⊔ closure s := by gcongr ?_ ⊔ _; exact closure_pow_le n.succ_ne_zero + _ = closure s := sup_idem _ + +@[to_additive] +lemma closure_pow {n : ℕ} (hs : 1 ∈ s) (hn : n ≠ 0) : closure (s ^ n) = closure s := + (closure_pow_le hn).antisymm <| by gcongr; exact subset_pow hs hn + @[to_additive] theorem sup_eq_closure_mul (H K : Subgroup G) : H ⊔ K = closure ((H : Set G) * (K : Set G)) := le_antisymm diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index cc8c81e07b75a..567936d4f37ef 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -96,6 +96,21 @@ theorem closure_mul_le (S T : Set M) : closure (S * T) ≤ closure S ⊔ closure (closure S ⊔ closure T).mul_mem (SetLike.le_def.mp le_sup_left <| subset_closure hs) (SetLike.le_def.mp le_sup_right <| subset_closure ht) +@[to_additive] +lemma closure_pow_le : ∀ {n}, n ≠ 0 → closure (s ^ n) ≤ closure s + | 1, _ => by simp + | n + 2, _ => + calc + closure (s ^ (n + 2)) + _ = closure (s ^ (n + 1) * s) := by rw [pow_succ] + _ ≤ closure (s ^ (n + 1)) ⊔ closure s := closure_mul_le .. + _ ≤ closure s ⊔ closure s := by gcongr ?_ ⊔ _; exact closure_pow_le n.succ_ne_zero + _ = closure s := sup_idem _ + +@[to_additive] +lemma closure_pow {n : ℕ} (hs : 1 ∈ s) (hn : n ≠ 0) : closure (s ^ n) = closure s := + (closure_pow_le hn).antisymm <| by gcongr; exact subset_pow hs hn + @[to_additive] theorem sup_eq_closure_mul (H K : Submonoid M) : H ⊔ K = closure ((H : Set M) * (K : Set M)) := le_antisymm @@ -572,15 +587,10 @@ theorem bot_mul (S : AddSubmonoid R) : ⊥ * S = ⊥ := variable {M N P Q : AddSubmonoid R} -@[mono] -theorem mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := - smul_le_smul hmp hnq - -theorem mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := - smul_le_smul_left h +@[mono, gcongr] lemma mul_le_mul (hmp : M ≤ P) (hnq : N ≤ Q) : M * N ≤ P * Q := smul_le_smul hmp hnq -theorem mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := - smul_le_smul_right h +@[gcongr] lemma mul_le_mul_left (h : M ≤ N) : M * P ≤ N * P := smul_le_smul_left h +@[gcongr] lemma mul_le_mul_right (h : N ≤ P) : M * N ≤ M * P := smul_le_smul_right h theorem mul_subset_mul : (↑M : Set R) * (↑N : Set R) ⊆ (↑(M * N) : Set R) := smul_subset_smul diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean index c112e0a96c541..0de5a19d7fbc0 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean @@ -185,12 +185,18 @@ open scoped RightActions @[simp] lemma inv_smul_finset_distrib₀ (a : α) (s : Finset α) : (a • s)⁻¹ = s⁻¹ <• a⁻¹ := by obtain rfl | ha := eq_or_ne a 0 · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] - · ext; simp [← inv_smul_mem_iff₀, *] + -- was `simp` and very slow (https://github.com/leanprover-community/mathlib4/issues/19751) + · ext; simp only [mem_inv', ne_eq, not_false_eq_true, ← inv_smul_mem_iff₀, smul_eq_mul, + MulOpposite.op_inv, inv_eq_zero, MulOpposite.op_eq_zero_iff, inv_inv, + MulOpposite.smul_eq_mul_unop, MulOpposite.unop_op, mul_inv_rev, ha] @[simp] lemma inv_op_smul_finset_distrib₀ (a : α) (s : Finset α) : (s <• a)⁻¹ = a⁻¹ • s⁻¹ := by obtain rfl | ha := eq_or_ne a 0 · obtain rfl | hs := s.eq_empty_or_nonempty <;> simp [*] - · ext; simp [← inv_smul_mem_iff₀, *] + -- was `simp` and very slow (https://github.com/leanprover-community/mathlib4/issues/19751) + · ext; simp only [mem_inv', ne_eq, MulOpposite.op_eq_zero_iff, not_false_eq_true, ← + inv_smul_mem_iff₀, MulOpposite.smul_eq_mul_unop, MulOpposite.unop_inv, MulOpposite.unop_op, + inv_eq_zero, inv_inv, smul_eq_mul, mul_inv_rev, ha] end GroupWithZero diff --git a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean index c93d58a0a6142..2d7511f6ec174 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean @@ -12,18 +12,20 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic assert_not_exists DenselyOrdered -variable {G₀ : Type*} [GroupWithZero G₀] +variable {G₀ : Type*} + +namespace Equiv +section GroupWithZero +variable [GroupWithZero G₀] /-- In a `GroupWithZero` `G₀`, the unit group `G₀ˣ` is equivalent to the subtype of nonzero elements. -/ -@[simps] def unitsEquivNeZero : G₀ˣ ≃ {a : G₀ // a ≠ 0} where +@[simps] def _root_.unitsEquivNeZero : G₀ˣ ≃ {a : G₀ // a ≠ 0} where toFun a := ⟨a, a.ne_zero⟩ invFun a := Units.mk0 _ a.prop left_inv _ := Units.ext rfl right_inv _ := rfl -namespace Equiv - /-- Left multiplication by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @[simps! (config := .asFn)] @@ -45,10 +47,25 @@ theorem _root_.mulRight_bijective₀ (a : G₀) (ha : a ≠ 0) : Function.Biject /-- Right division by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @[simps! (config := { simpRhs := true })] -def divRight₀ (a : G₀) (ha : a ≠ 0) : G₀ ≃ G₀ where +def divRight₀ (a : G₀) (ha : a ≠ 0) : Perm G₀ where toFun := (· / a) invFun := (· * a) left_inv _ := by simp [ha] right_inv _ := by simp [ha] +end GroupWithZero + +section CommGroupWithZero +variable [CommGroupWithZero G₀] + +/-- Left division by a nonzero element in a `CommGroupWithZero` is a permutation of the underlying +type. -/ +@[simps! (config := { simpRhs := true })] +def divLeft₀ (a : G₀) (ha : a ≠ 0) : Perm G₀ where + toFun := (a / ·) + invFun := (a / ·) + left_inv _ := by simp [ha] + right_inv _ := by simp [ha] + +end CommGroupWithZero end Equiv diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index a66ee204117c8..c98cac0916bc0 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -66,9 +66,7 @@ variable {R : Type u} [CommRing R] {D : Type v} [SmallCategory D] determined by the functor `I` -/ def ringModIdeals (I : D ⥤ Ideal R) : D ⥤ ModuleCat.{u} R where obj t := ModuleCat.of R <| R ⧸ I.obj t - map w := Submodule.mapQ _ _ LinearMap.id (I.map w).down.down - -- Porting note: was 'obviously' - map_comp f g := by apply Submodule.linearMap_qext; rfl + map w := ModuleCat.ofHom <| Submodule.mapQ _ _ LinearMap.id (I.map w).down.down -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: Once this file is ported, move this instance to the right location. instance moduleCat_enoughProjectives' : EnoughProjectives (ModuleCat.{u} R) := diff --git a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean index 59034f288455b..87236bc52bb4d 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ModuleCat.lean @@ -33,7 +33,7 @@ linear maps `f` and `g` and the vanishing of their composition. -/ def moduleCatMk {X₁ X₂ X₃ : Type v} [AddCommGroup X₁] [AddCommGroup X₂] [AddCommGroup X₃] [Module R X₁] [Module R X₂] [Module R X₃] (f : X₁ →ₗ[R] X₂) (g : X₂ →ₗ[R] X₃) (hfg : g.comp f = 0) : ShortComplex (ModuleCat.{v} R) := - ShortComplex.mk (ModuleCat.asHom f) (ModuleCat.asHom g) hfg + ShortComplex.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) (ModuleCat.hom_ext hfg) variable (S : ShortComplex (ModuleCat.{v} R)) @@ -46,33 +46,19 @@ lemma moduleCat_exact_iff : S.exact_iff_of_concreteCategory lemma moduleCat_exact_iff_ker_sub_range : - S.Exact ↔ LinearMap.ker S.g ≤ LinearMap.range S.f := by + S.Exact ↔ LinearMap.ker S.g.hom ≤ LinearMap.range S.f.hom := by rw [moduleCat_exact_iff] - constructor - · intro h x₂ hx₂ - exact h x₂ hx₂ - · intro h x₂ hx₂ - exact h hx₂ + aesop lemma moduleCat_exact_iff_range_eq_ker : - S.Exact ↔ LinearMap.range S.f = LinearMap.ker S.g := by + S.Exact ↔ LinearMap.range S.f.hom = LinearMap.ker S.g.hom := by rw [moduleCat_exact_iff_ker_sub_range] - constructor - · intro h - ext x - constructor - · rintro ⟨y, hy⟩ - rw [← hy] - simp only [LinearMap.mem_ker, moduleCat_zero_apply] - · intro hx - exact h hx - · intro h - rw [h] + aesop variable {S} lemma Exact.moduleCat_range_eq_ker (hS : S.Exact) : - LinearMap.range S.f = LinearMap.ker S.g := by + LinearMap.range S.f.hom = LinearMap.ker S.g.hom := by simpa only [moduleCat_exact_iff_range_eq_ker] using hS lemma ShortExact.moduleCat_injective_f (hS : S.ShortExact) : @@ -94,19 +80,17 @@ lemma ShortExact.moduleCat_exact_iff_function_exact : morphisms `f` and `g` and the assumption `LinearMap.range f ≤ LinearMap.ker g`. -/ @[simps] def moduleCatMkOfKerLERange {X₁ X₂ X₃ : ModuleCat.{v} R} (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) - (hfg : LinearMap.range f ≤ LinearMap.ker g) : ShortComplex (ModuleCat.{v} R) := - ShortComplex.mk f g (by - ext - exact hfg ⟨_, rfl⟩) + (hfg : LinearMap.range f.hom ≤ LinearMap.ker g.hom) : ShortComplex (ModuleCat.{v} R) := + ShortComplex.mk f g (by aesop) lemma Exact.moduleCat_of_range_eq_ker {X₁ X₂ X₃ : ModuleCat.{v} R} - (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) (hfg : LinearMap.range f = LinearMap.ker g) : + (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) (hfg : LinearMap.range f.hom = LinearMap.ker g.hom) : (moduleCatMkOfKerLERange f g (by rw [hfg])).Exact := by simpa only [moduleCat_exact_iff_range_eq_ker] using hfg /-- The canonical linear map `S.X₁ →ₗ[R] LinearMap.ker S.g` induced by `S.f`. -/ @[simps] -def moduleCatToCycles : S.X₁ →ₗ[R] LinearMap.ker S.g where +def moduleCatToCycles : S.X₁ →ₗ[R] LinearMap.ker S.g.hom where toFun x := ⟨S.f x, S.moduleCat_zero_apply x⟩ map_add' x y := by aesop map_smul' a x := by aesop @@ -114,35 +98,28 @@ def moduleCatToCycles : S.X₁ →ₗ[R] LinearMap.ker S.g where /-- The homology of `S`, defined as the quotient of the kernel of `S.g` by the image of `S.moduleCatToCycles` -/ abbrev moduleCatHomology := - ModuleCat.of R (LinearMap.ker S.g ⧸ LinearMap.range S.moduleCatToCycles) + ModuleCat.of R (LinearMap.ker S.g.hom ⧸ LinearMap.range S.moduleCatToCycles) /-- The canonical map `ModuleCat.of R (LinearMap.ker S.g) ⟶ S.moduleCatHomology`. -/ -abbrev moduleCatHomologyπ : ModuleCat.of R (LinearMap.ker S.g) ⟶ S.moduleCatHomology := - (LinearMap.range S.moduleCatToCycles).mkQ +abbrev moduleCatHomologyπ : ModuleCat.of R (LinearMap.ker S.g.hom) ⟶ S.moduleCatHomology := + ModuleCat.ofHom (LinearMap.range S.moduleCatToCycles).mkQ /-- The explicit left homology data of a short complex of modules that is given by a kernel and a quotient given by the `LinearMap` API. -/ -@[simps] +@[simps K H i π] def moduleCatLeftHomologyData : S.LeftHomologyData where - K := ModuleCat.of R (LinearMap.ker S.g) + K := ModuleCat.of R (LinearMap.ker S.g.hom) H := S.moduleCatHomology - i := (LinearMap.ker S.g).subtype + i := ModuleCat.ofHom (LinearMap.ker S.g.hom).subtype π := S.moduleCatHomologyπ - wi := by - ext ⟨_, hx⟩ - exact hx + wi := by aesop hi := ModuleCat.kernelIsLimit _ - wπ := by - ext (x : S.X₁) - dsimp - erw [Submodule.Quotient.mk_eq_zero] - rw [LinearMap.mem_range] - apply exists_apply_eq_apply - hπ := ModuleCat.cokernelIsColimit (ModuleCat.asHom S.moduleCatToCycles) + wπ := by aesop + hπ := ModuleCat.cokernelIsColimit (ModuleCat.ofHom S.moduleCatToCycles) @[simp] lemma moduleCatLeftHomologyData_f' : - S.moduleCatLeftHomologyData.f' = S.moduleCatToCycles := rfl + S.moduleCatLeftHomologyData.f' = ModuleCat.ofHom S.moduleCatToCycles := rfl instance : Epi S.moduleCatHomologyπ := (inferInstance : Epi S.moduleCatLeftHomologyData.π) @@ -150,22 +127,22 @@ instance : Epi S.moduleCatHomologyπ := /-- Given a short complex `S` of modules, this is the isomorphism between the abstract `S.cycles` of the homology API and the more concrete description as `LinearMap.ker S.g`. -/ -noncomputable def moduleCatCyclesIso : S.cycles ≅ ModuleCat.of R (LinearMap.ker S.g) := +noncomputable def moduleCatCyclesIso : S.cycles ≅ ModuleCat.of R (LinearMap.ker S.g.hom) := S.moduleCatLeftHomologyData.cyclesIso @[reassoc (attr := simp, elementwise)] lemma moduleCatCyclesIso_hom_subtype : - S.moduleCatCyclesIso.hom ≫ (LinearMap.ker S.g).subtype = S.iCycles := + S.moduleCatCyclesIso.hom ≫ ModuleCat.ofHom (LinearMap.ker S.g.hom).subtype = S.iCycles := S.moduleCatLeftHomologyData.cyclesIso_hom_comp_i @[reassoc (attr := simp, elementwise)] lemma moduleCatCyclesIso_inv_iCycles : - S.moduleCatCyclesIso.inv ≫ S.iCycles = (LinearMap.ker S.g).subtype := + S.moduleCatCyclesIso.inv ≫ S.iCycles = ModuleCat.ofHom (LinearMap.ker S.g.hom).subtype := S.moduleCatLeftHomologyData.cyclesIso_inv_comp_iCycles @[reassoc (attr := simp, elementwise)] lemma toCycles_moduleCatCyclesIso_hom : - S.toCycles ≫ S.moduleCatCyclesIso.hom = S.moduleCatToCycles := by + S.toCycles ≫ S.moduleCatCyclesIso.hom = ModuleCat.ofHom S.moduleCatToCycles := by rw [← cancel_mono S.moduleCatLeftHomologyData.i, moduleCatLeftHomologyData_i, Category.assoc, S.moduleCatCyclesIso_hom_subtype, toCycles_i] rfl diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index 1ce5fd8fdcf58..4a9632102afc4 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -77,7 +77,7 @@ theorem derivedSeriesOfIdeal_add (k l : ℕ) : D (k + l) I = D k (D l I) := by | zero => rw [Nat.zero_add, derivedSeriesOfIdeal_zero] | succ k ih => rw [Nat.succ_add k l, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ, ih] -@[mono] +@[gcongr, mono] theorem derivedSeriesOfIdeal_le {I J : LieIdeal R L} {k l : ℕ} (h₁ : I ≤ J) (h₂ : l ≤ k) : D k I ≤ D l J := by revert l; induction' k with k ih <;> intro l h₂ diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 1201232787862..32a16eb6c461c 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -838,7 +838,7 @@ theorem comap_incl_eq_bot : N₂.comap N.incl = ⊥ ↔ N ⊓ N₂ = ⊥ := by inf_coe_toSubmodule] rw [← Submodule.disjoint_iff_comap_eq_bot, disjoint_iff] -@[mono] +@[gcongr, mono] theorem map_mono (h : N ≤ N₂) : N.map f ≤ N₂.map f := Set.image_subset _ h diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index 1acfd102224c7..01285cf51e542 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -132,8 +132,11 @@ lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right [LieAlgebra.IsN obtain ⟨k, hk⟩ := k suffices genWeightSpace M ((k + 1) • α + χ) ≤ genWeightSpaceChain M α χ p q by apply this - simpa using (rootSpaceWeightSpaceProduct R L H M α (k • α + χ) ((k + 1) • α + χ) - (by rw [add_smul]; abel) (⟨x, hx⟩ ⊗ₜ ⟨z, hz⟩)).property + -- was `simpa using [...]` and very slow + -- (https://github.com/leanprover-community/mathlib4/issues/19751) + simpa only [zsmul_eq_mul, Int.cast_add, Pi.intCast_def, Int.cast_one] using + (rootSpaceWeightSpaceProduct R L H M α (k • α + χ) ((k + 1) • α + χ) + (by rw [add_smul]; abel) (⟨x, hx⟩ ⊗ₜ ⟨z, hz⟩)).property rw [genWeightSpaceChain] rcases eq_or_ne (k + 1) q with rfl | hk'; · simp only [hq, bot_le] replace hk' : k + 1 ∈ Ioo p q := ⟨by linarith [hk.1], lt_of_le_of_ne hk.2 hk'⟩ diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 67ae9340c5ef4..1ccd71d6cbcbb 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -96,6 +96,8 @@ theorem coe_pow (e : M ≃ₗ[R] M) (n : ℕ) : ⇑(e ^ n) = e^[n] := hom_coe_po theorem pow_apply (e : M ≃ₗ[R] M) (n : ℕ) (m : M) : (e ^ n) m = e^[n] m := congr_fun (coe_pow e n) m +@[simp] lemma mul_apply (f : M ≃ₗ[R] M) (g : M ≃ₗ[R] M) (x : M) : (f * g) x = f (g x) := rfl + /-- Restriction from `R`-linear automorphisms of `M` to `R`-linear endomorphisms of `M`, promoted to a monoid hom. -/ @[simps] diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index efe62ee72caee..b7f0ab08698d3 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -914,6 +914,11 @@ theorem is_universal : ∃! l : M' →ₗ[R] M'', l.comp f = g := fun g h => ⟨lift S f g h, lift_comp S f g h, fun l hl => (lift_unique S f g h l hl).symm⟩ +theorem linearMap_ext {N N'} [AddCommMonoid N] [Module R N] [AddCommMonoid N'] [Module R N'] + (f' : N →ₗ[R] N') [IsLocalizedModule S f'] ⦃g g' : M' →ₗ[R] N'⦄ + (h : g ∘ₗ f = g' ∘ₗ f) : g = g' := + (is_universal S f _ <| map_units f').unique h rfl + theorem ringHom_ext (map_unit : ∀ x : S, IsUnit ((algebraMap R (Module.End R M'')) x)) ⦃j k : M' →ₗ[R] M''⦄ (h : j.comp f = k.comp f) : j = k := by rw [← lift_unique S f (k.comp f) map_unit j h, lift_unique] diff --git a/Mathlib/Algebra/Module/Presentation/RestrictScalars.lean b/Mathlib/Algebra/Module/Presentation/RestrictScalars.lean new file mode 100644 index 0000000000000..108320b3d56d8 --- /dev/null +++ b/Mathlib/Algebra/Module/Presentation/RestrictScalars.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Module.Presentation.DirectSum +import Mathlib.Algebra.Module.Presentation.Cokernel + +/-! +# Presentation of the restriction of scalars of a module + +Given a morphism of rings `A → B` and a `B`-module `M`, we obtain a presentation +of `M` as a `A`-module from a presentation of `M` as `B`-module, +a presentation of `B` as a `A`-module (and some additional data). + +## TODO +* deduce that if `B` is a finitely presented as an `A`-module and `M` is +finitely presented as an `B`-module, then `M` is finitely presented as an `A`-module + +-/ + +namespace Module + +variable {B : Type*} [Ring B] {M : Type*} [AddCommGroup M] [Module B M] + [DecidableEq B] + (presM : Presentation B M) [DecidableEq presM.G] + {A : Type*} [CommRing A] [Algebra A B] [Module A M] [IsScalarTower A B M] + (presB : Presentation A B) + +namespace Presentation + +/-- The additional data that is necessary in order to obtain a presentation +of the restriction of scalars of a module. -/ +abbrev RestrictScalarsData : Type _ := + (presB.finsupp presM.G).CokernelData + (LinearMap.restrictScalars A presM.map) + (fun (⟨g, g'⟩ : presB.G × presM.R) ↦ presB.var g • Finsupp.single g' (1 : B)) + +variable (data : presM.RestrictScalarsData presB) + +/-- A presentation of the restriction of scalars from `B` to `A` of a `B`-module `M`, +given a presentation of `M` as a `B`-module, a presentation of `B` as an `A`-module, +and an additional data. -/ +noncomputable def restrictScalars : Presentation A M := + ofExact (g := LinearMap.restrictScalars A presM.π) (presB.finsupp presM.G) data + presM.exact presM.surjective_π (by + ext v + dsimp + simp only [Submodule.mem_top, iff_true] + apply Finsupp.induction + · simp + · intro r b w _ _ hw + refine Submodule.add_mem _ ?_ hw + obtain ⟨β, rfl⟩ := presB.surjective_π b + apply Finsupp.induction (p := fun β ↦ Finsupp.single r (presB.π β) ∈ _) + · simp + · intro g a f _ _ hf + rw [map_add, Finsupp.single_add] + refine Submodule.add_mem _ ?_ hf + rw [← Finsupp.smul_single_one, ← Finsupp.smul_single_one, + map_smul, Relations.Solution.π_single, smul_assoc] + exact Submodule.smul_mem _ _ (Submodule.subset_span ⟨⟨g, r⟩, rfl⟩)) + +end Presentation + +end Module diff --git a/Mathlib/Algebra/Module/Submodule/Bilinear.lean b/Mathlib/Algebra/Module/Submodule/Bilinear.lean index 1564df8114cd7..0f4c48b6a1522 100644 --- a/Mathlib/Algebra/Module/Submodule/Bilinear.lean +++ b/Mathlib/Algebra/Module/Submodule/Bilinear.lean @@ -84,7 +84,7 @@ theorem map₂_bot_left (f : M →ₗ[R] N →ₗ[R] P) (q : Submodule R N) : ma rw [Submodule.mem_bot] at hm ⊢ rw [hm, LinearMap.map_zero₂] -@[mono] +@[gcongr, mono] theorem map₂_le_map₂ {f : M →ₗ[R] N →ₗ[R] P} {p₁ p₂ : Submodule R M} {q₁ q₂ : Submodule R N} (hp : p₁ ≤ p₂) (hq : q₁ ≤ q₂) : map₂ f p₁ q₁ ≤ map₂ f p₂ q₂ := map₂_le.2 fun _m hm _n hn => apply_mem_map₂ _ (hp hm) (hq hn) diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index cadfde60a1179..b8b6b0b57e2f6 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -62,11 +62,17 @@ variable (b : Basis ι K E) theorem span_top : span K (span ℤ (Set.range b) : Set E) = ⊤ := by simp [span_span_of_tower] - theorem map {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] (f : E ≃ₗ[K] F) : Submodule.map (f.restrictScalars ℤ) (span ℤ (Set.range b)) = span ℤ (Set.range (b.map f)) := by simp_rw [Submodule.map_span, LinearEquiv.restrictScalars_apply, Basis.coe_map, Set.range_comp] +open scoped Pointwise in +theorem smul {c : K} (hc : c ≠ 0) : + c • span ℤ (Set.range b) = span ℤ (Set.range (b.isUnitSMul (fun _ ↦ hc.isUnit))) := by + rw [smul_span, Set.smul_set_range] + congr! + rw [Basis.isUnitSMul_apply] + /-- The fundamental domain of the ℤ-lattice spanned by `b`. See `ZSpan.isAddFundamentalDomain` for the proof that it is a fundamental domain. -/ def fundamentalDomain : Set E := {m | ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1} @@ -309,6 +315,13 @@ instance [Finite ι] : DiscreteTopology (span ℤ (Set.range b)) := by instance [Finite ι] : DiscreteTopology (span ℤ (Set.range b)).toAddSubgroup := inferInstanceAs <| DiscreteTopology (span ℤ (Set.range b)) +theorem setFinite_inter [ProperSpace E] [Finite ι] {s : Set E} (hs : Bornology.IsBounded s) : + Set.Finite (s ∩ span ℤ (Set.range b)) := by + have : DiscreteTopology (span ℤ (Set.range b)) := inferInstance + refine Metric.finite_isBounded_inter_isClosed hs ?_ + change IsClosed ((span ℤ (Set.range b)).toAddSubgroup : Set E) + exact AddSubgroup.isClosed_of_discrete + @[measurability] theorem fundamentalDomain_measurableSet [MeasurableSpace E] [OpensMeasurableSpace E] [Finite ι] : MeasurableSet (fundamentalDomain b) := by diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 41e3c6d67e134..da1c6d0c0a369 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -345,6 +345,10 @@ variable {S : Type*} instance smulZeroClass [Semiring k] [SMulZeroClass R k] : SMulZeroClass R (MonoidAlgebra k G) := Finsupp.smulZeroClass +instance noZeroSMulDivisors [Zero R] [Semiring k] [SMulZeroClass R k] [NoZeroSMulDivisors R k] : + NoZeroSMulDivisors R (MonoidAlgebra k G) := + Finsupp.noZeroSMulDivisors + instance distribSMul [Semiring k] [DistribSMul R k] : DistribSMul R (MonoidAlgebra k G) := Finsupp.distribSMul _ _ @@ -1063,6 +1067,10 @@ section Semiring instance smulZeroClass [Semiring k] [SMulZeroClass R k] : SMulZeroClass R k[G] := Finsupp.smulZeroClass +instance noZeroSMulDivisors [Zero R] [Semiring k] [SMulZeroClass R k] [NoZeroSMulDivisors R k] : + NoZeroSMulDivisors R k[G] := + Finsupp.noZeroSMulDivisors + variable [Semiring k] [AddMonoid G] instance semiring : Semiring k[G] := diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 45521d9bd0bb7..5227f6ddbbe60 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -500,6 +500,15 @@ theorem coeff_eq_zero_of_totalDegree_lt {f : MvPolynomial σ R} {d : σ →₀ exact lt_irrefl _ · exact lt_of_le_of_lt (Nat.zero_le _) h +theorem totalDegree_eq_zero_iff_eq_C {p : MvPolynomial σ R} : + p.totalDegree = 0 ↔ p = C (p.coeff 0) := by + constructor <;> intro h + · ext m; classical rw [coeff_C]; split_ifs with hm; · rw [← hm] + apply coeff_eq_zero_of_totalDegree_lt; rw [h] + exact Finset.sum_pos (fun i hi ↦ Nat.pos_of_ne_zero <| Finsupp.mem_support_iff.mp hi) + (Finsupp.support_nonempty_iff.mpr <| Ne.symm hm) + · rw [h, totalDegree_C] + theorem totalDegree_rename_le (f : σ → τ) (p : MvPolynomial σ R) : (rename f p).totalDegree ≤ p.totalDegree := Finset.sup_le fun b => by diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index 9c6e2c3c43e7c..5cafce58bab3e 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -72,6 +72,13 @@ theorem pderiv_monomial {i : σ} : · rw [Finsupp.not_mem_support_iff] at hi; simp [hi] · simp +lemma X_mul_pderiv_monomial {i : σ} {m : σ →₀ ℕ} {r : R} : + X i * pderiv i (monomial m r) = m i • monomial m r := by + rw [pderiv_monomial, X, monomial_mul, smul_monomial] + by_cases h : m i = 0 + · simp_rw [h, Nat.cast_zero, mul_zero, zero_smul, monomial_zero] + rw [one_mul, mul_comm, nsmul_eq_mul, add_comm, sub_add_single_one_cancel h] + theorem pderiv_C {i : σ} : pderiv i (C a) = 0 := derivation_C _ _ diff --git a/Mathlib/Algebra/Order/Antidiag/Nat.lean b/Mathlib/Algebra/Order/Antidiag/Nat.lean index 40c8db2fe9e60..88a8cc741a284 100644 --- a/Mathlib/Algebra/Order/Antidiag/Nat.lean +++ b/Mathlib/Algebra/Order/Antidiag/Nat.lean @@ -11,11 +11,13 @@ import Mathlib.NumberTheory.ArithmeticFunction # Sets of tuples with a fixed product This file defines the finite set of `d`-tuples of natural numbers with a fixed product `n` as -`Nat.finMulAntidiagonal`. +`Nat.finMulAntidiag`. ## Main Results * There are `d^(ω n)` ways to write `n` as a product of `d` natural numbers, when `n` is squarefree -(`card_finMulAntidiagonal_of_squarefree`) +(`card_finMulAntidiag_of_squarefree`) +* There are `3^(ω n)` pairs of natural numbers whose `lcm` is `n`, when `n` is squarefree +(`card_pair_lcm_eq`) -/ open Finset @@ -239,4 +241,78 @@ theorem card_finMulAntidiag_of_squarefree {d n : ℕ} (hn : Squarefree n) : ArithmeticFunction.cardDistinctFactors_apply, ← List.card_toFinset, toFinset_factors, Finset.card_fin] +theorem finMulAntidiag_three {n : ℕ} : + ∀ a ∈ finMulAntidiag 3 n, a 0 * a 1 * a 2 = n := by + intro a ha + rw [← (mem_finMulAntidiag.mp ha).1, Fin.prod_univ_three a] + +namespace card_pair_lcm_eq + +/-! +The following private declarations are ingredients for the proof of `card_pair_lcm_eq`. +-/ + +@[reducible] +private def f {n : ℕ} : ∀ a ∈ finMulAntidiag 3 n, ℕ × ℕ := fun a _ => (a 0 * a 1, a 0 * a 2) + +private theorem f_img {n : ℕ} (hn : Squarefree n) (a : Fin 3 → ℕ) + (ha : a ∈ finMulAntidiag 3 n) : + f a ha ∈ Finset.filter (fun ⟨x, y⟩ => x.lcm y = n) (n.divisors ×ˢ n.divisors) := by + rw [mem_filter, Finset.mem_product, mem_divisors, mem_divisors] + refine ⟨⟨⟨?_, hn.ne_zero⟩, ⟨?_, hn.ne_zero⟩⟩, ?_⟩ <;> rw [f, ← finMulAntidiag_three a ha] + · apply dvd_mul_right + · use a 1; ring + dsimp only + rw [lcm_mul_left, Nat.Coprime.lcm_eq_mul] + · ring + refine coprime_of_squarefree_mul (hn.squarefree_of_dvd ?_) + use a 0; rw [← finMulAntidiag_three a ha]; ring + +private theorem f_inj {n : ℕ} (a : Fin 3 → ℕ) (ha : a ∈ finMulAntidiag 3 n) + (b : Fin 3 → ℕ) (hb : b ∈ finMulAntidiag 3 n) (hfab : f a ha = f b hb) : + a = b := by + obtain ⟨hfab1, hfab2⟩ := Prod.mk.inj hfab + have hprods : a 0 * a 1 * a 2 = a 0 * a 1 * b 2 := by + rw [finMulAntidiag_three a ha, hfab1, finMulAntidiag_three b hb] + have hab2 : a 2 = b 2 := by + rw [← mul_right_inj' <| mul_ne_zero (ne_zero_of_mem_finMulAntidiag ha 0) + (ne_zero_of_mem_finMulAntidiag ha 1)] + exact hprods + have hab0 : a 0 = b 0 := by + rw [hab2] at hfab2 + exact (mul_left_inj' <| ne_zero_of_mem_finMulAntidiag hb 2).mp hfab2; + have hab1 : a 1 = b 1 := by + rw [hab0] at hfab1 + exact (mul_right_inj' <| ne_zero_of_mem_finMulAntidiag hb 0).mp hfab1; + funext i; fin_cases i <;> assumption + +private theorem f_surj {n : ℕ} (hn : n ≠ 0) (b : ℕ × ℕ) + (hb : b ∈ Finset.filter (fun ⟨x, y⟩ => x.lcm y = n) (n.divisors ×ˢ n.divisors)) : + ∃ (a : Fin 3 → ℕ) (ha : a ∈ finMulAntidiag 3 n), f a ha = b := by + dsimp only at hb + let g := b.fst.gcd b.snd + let a := ![g, b.fst/g, b.snd/g] + have ha : a ∈ finMulAntidiag 3 n := by + rw [mem_finMulAntidiag] + rw [mem_filter, Finset.mem_product] at hb + refine ⟨?_, hn⟩ + · rw [Fin.prod_univ_three a] + simp only [a, Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons, + Matrix.cons_val_two, Matrix.tail_cons] + rw [Nat.mul_div_cancel_left' (Nat.gcd_dvd_left _ _), ← hb.2, lcm, + Nat.mul_div_assoc b.fst (Nat.gcd_dvd_right b.fst b.snd)] + use a; use ha + apply Prod.ext <;> simp only [Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons] + <;> apply Nat.mul_div_cancel' + · apply Nat.gcd_dvd_left + · apply Nat.gcd_dvd_right + +end card_pair_lcm_eq + +open card_pair_lcm_eq in +theorem card_pair_lcm_eq {n : ℕ} (hn : Squarefree n) : + #{p ∈ (n.divisors ×ˢ n.divisors) | p.1.lcm p.2 = n} = 3 ^ ω n := by + rw [← card_finMulAntidiag_of_squarefree hn, eq_comm] + apply Finset.card_bij f (f_img hn) (f_inj) (f_surj hn.ne_zero) + end Nat diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index f2c11d7931ca5..aa2f6b9fb2e4a 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -223,7 +223,8 @@ lemma nsmul_piAntidiag [DecidableEq (ι → ℕ)] (s : Finset ι) (m : ℕ) {n : · rw [not_imp_comm.1 (hfsup _) hi] exact dvd_zero _ refine ⟨fun i ↦ f i / n, ?_⟩ - simpa [Nat.sum_div, Nat.div_ne_zero_iff_of_dvd, funext_iff, Nat.mul_div_cancel', ← Nat.sum_div, *] + simp [funext_iff, Nat.mul_div_cancel', ← Nat.sum_div, *] + aesop lemma map_nsmul_piAntidiag (s : Finset ι) (m : ℕ) {n : ℕ} (hn : n ≠ 0) : (piAntidiag s m).map diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 5b90caf6e8964..122d83453934b 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,7 +537,8 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} -set_option linter.docPrime false in +/-- See also `div_le_div_iff` for a version that works for `LinearOrderedSemifield` with +additional assumptions. -/ @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index b40dacaccb13e..d7ee5915f5f32 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -90,8 +90,8 @@ structure. When possible, instead of parametrizing results over `(f : α ≃+o β)`, you should parametrize over `(F : Type*) [FunLike F M N] [AddEquivClass F M N] [OrderIsoClass F M N] (f : F)`. -/ -structure OrderAddMonoidIso (α β : Type*) [Preorder α] [Preorder β] [AddZeroClass α] - [AddZeroClass β] extends α ≃+ β where +structure OrderAddMonoidIso (α β : Type*) [Preorder α] [Preorder β] [Add α] [Add β] + extends α ≃+ β where /-- An `OrderAddMonoidIso` respects `≤`. -/ map_le_map_iff' {a b : α} : toFun a ≤ toFun b ↔ a ≤ b @@ -146,8 +146,8 @@ When possible, instead of parametrizing results over `(f : α ≃*o β)`, you should parametrize over `(F : Type*) [FunLike F M N] [MulEquivClass F M N] [OrderIsoClass F M N] (f : F)`. -/ @[to_additive] -structure OrderMonoidIso (α β : Type*) [Preorder α] [Preorder β] [MulOneClass α] - [MulOneClass β] extends α ≃* β where +structure OrderMonoidIso (α β : Type*) [Preorder α] [Preorder β] [Mul α] [Mul β] + extends α ≃* β where /-- An `OrderMonoidIso` respects `≤`. -/ map_le_map_iff' {a b : α} : toFun a ≤ toFun b ↔ a ≤ b @@ -513,8 +513,8 @@ namespace OrderMonoidIso section Preorder -variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] [MulOneClass α] [MulOneClass β] - [MulOneClass γ] [MulOneClass δ] {f g : α ≃*o β} +variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] [Mul α] [Mul β] + [Mul γ] [Mul δ] {f g : α ≃*o β} @[to_additive] instance : EquivLike (α ≃*o β) α β where diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index 0e9d89e4f69f8..a6577a705c6e3 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -293,9 +293,17 @@ lemma mul_max [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) : a * max b c = max (a * b) (a * c) := mul_left_mono.map_max @[to_additive] -lemma max_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : +lemma max_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : max a b * c = max (a * c) (b * c) := mul_right_mono.map_max +@[to_additive] +lemma mul_min [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) : + a * min b c = min (a * b) (a * c) := mul_left_mono.map_min + +@[to_additive] +lemma min_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : + min a b * c = min (a * c) (b * c) := mul_right_mono.map_min + @[to_additive] lemma min_lt_max_of_mul_lt_mul [MulLeftMono α] [MulRightMono α] (h : a * b < c * d) : min a b < max c d := by diff --git a/Mathlib/Algebra/Order/Quantale.lean b/Mathlib/Algebra/Order/Quantale.lean index 96bfa1624aa35..13108eac3152d 100644 --- a/Mathlib/Algebra/Order/Quantale.lean +++ b/Mathlib/Algebra/Order/Quantale.lean @@ -176,4 +176,21 @@ theorem rightMulResiduation_le_iff_mul_le : x ≤ y ⇨ᵣ z ↔ y * x ≤ z whe iSup_le_iff, implies_true] mpr h1 := le_sSup h1 +section Zero + +variable {α : Type*} [Semigroup α] [CompleteLattice α] [IsQuantale α] +variable {x : α} + +@[to_additive (attr := simp)] +theorem bot_mul : ⊥ * x = ⊥ := by + rw [← sSup_empty, sSup_mul_distrib] + simp only [Set.mem_empty_iff_false, not_false_eq_true, iSup_neg, iSup_bot, sSup_empty] + +@[to_additive (attr := simp)] +theorem mul_bot : x * ⊥ = ⊥ := by + rw [← sSup_empty, mul_sSup_distrib] + simp only [Set.mem_empty_iff_false, not_false_eq_true, iSup_neg, iSup_bot, sSup_empty] + +end Zero + end IsQuantale diff --git a/Mathlib/Algebra/Order/SuccPred.lean b/Mathlib/Algebra/Order/SuccPred.lean index 185ea8250d6f9..faf2e279f4aa8 100644 --- a/Mathlib/Algebra/Order/SuccPred.lean +++ b/Mathlib/Algebra/Order/SuccPred.lean @@ -6,7 +6,7 @@ Authors: Violeta Hernández Palacios, Yaël Dillies import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Order.ZeroLEOne import Mathlib.Data.Int.Cast.Defs -import Mathlib.Order.SuccPred.Archimedean +import Mathlib.Order.SuccPred.Limit /-! # Interaction between successors and arithmetic @@ -138,6 +138,46 @@ theorem covBy_iff_sub_one_eq [Sub α] [One α] [PredSubOrder α] [NoMinOrder α] rw [← pred_eq_sub_one] exact pred_eq_iff_covBy.symm +theorem IsSuccPrelimit.add_one_lt [Add α] [One α] [SuccAddOrder α] + (hx : IsSuccPrelimit x) (hy : y < x) : y + 1 < x := by + rw [← succ_eq_add_one] + exact hx.succ_lt hy + +theorem IsPredPrelimit.lt_sub_one [Sub α] [One α] [PredSubOrder α] + (hx : IsPredPrelimit x) (hy : x < y) : x < y - 1 := by + rw [← pred_eq_sub_one] + exact hx.lt_pred hy + +theorem IsSuccLimit.add_one_lt [Add α] [One α] [SuccAddOrder α] + (hx : IsSuccLimit x) (hy : y < x) : y + 1 < x := + hx.isSuccPrelimit.add_one_lt hy + +theorem IsPredLimit.lt_sub_one [Sub α] [One α] [PredSubOrder α] + (hx : IsPredLimit x) (hy : x < y) : x < y - 1 := + hx.isPredPrelimit.lt_sub_one hy + +theorem IsSuccPrelimit.add_natCast_lt [AddMonoidWithOne α] [SuccAddOrder α] + (hx : IsSuccPrelimit x) (hy : y < x) : ∀ n : ℕ, y + n < x + | 0 => by simpa + | n + 1 => by + rw [Nat.cast_add_one, ← add_assoc] + exact hx.add_one_lt (hx.add_natCast_lt hy n) + +theorem IsPredPrelimit.lt_sub_natCast [AddCommGroupWithOne α] [PredSubOrder α] + (hx : IsPredPrelimit x) (hy : x < y) : ∀ n : ℕ, x < y - n + | 0 => by simpa + | n + 1 => by + rw [Nat.cast_add_one, ← sub_sub] + exact hx.lt_sub_one (hx.lt_sub_natCast hy n) + +theorem IsSuccLimit.add_natCast_lt [AddMonoidWithOne α] [SuccAddOrder α] + (hx : IsSuccLimit x) (hy : y < x) : ∀ n : ℕ, y + n < x := + hx.isSuccPrelimit.add_natCast_lt hy + +theorem IsPredLimit.lt_sub_natCast [AddCommGroupWithOne α] [PredSubOrder α] + (hx : IsPredLimit x) (hy : x < y) : ∀ n : ℕ, x < y - n := + hx.isPredPrelimit.lt_sub_natCast hy + end PartialOrder section LinearOrder diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index db335123a7319..93d9b0870d987 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -135,6 +135,11 @@ instance smulZeroClass {S : Type*} [SMulZeroClass S R] : SMulZeroClass S R[X] wh smul r p := ⟨r • p.toFinsupp⟩ smul_zero a := congr_arg ofFinsupp (smul_zero a) +instance {S : Type*} [Zero S] [SMulZeroClass S R] [NoZeroSMulDivisors S R] : + NoZeroSMulDivisors S R[X] where + eq_zero_or_eq_zero_of_smul_eq_zero eq := + (eq_zero_or_eq_zero_of_smul_eq_zero <| congr_arg toFinsupp eq).imp id (congr_arg ofFinsupp) + -- to avoid a bug in the `ring` tactic instance (priority := 1) pow : Pow R[X] ℕ where pow p n := npowRec n p diff --git a/Mathlib/Algebra/Polynomial/Degree/Domain.lean b/Mathlib/Algebra/Polynomial/Degree/Domain.lean index cad96a6912720..04519630e17d6 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Domain.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Domain.lean @@ -40,6 +40,15 @@ lemma natDegree_mul (hp : p ≠ 0) (hq : q ≠ 0) : (p*q).natDegree = p.natDegre rw [← Nat.cast_inj (R := WithBot ℕ), ← degree_eq_natDegree (mul_ne_zero hp hq), Nat.cast_add, ← degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul] +variable (p) in +lemma natDegree_smul (ha : a ≠ 0) : (a • p).natDegree = p.natDegree := by + by_cases hp : p = 0 + · simp only [hp, smul_zero] + · apply natDegree_eq_of_le_of_coeff_ne_zero + · exact (natDegree_smul_le _ _).trans (le_refl _) + · simpa only [coeff_smul, coeff_natDegree, smul_eq_mul, ne_eq, mul_eq_zero, + leadingCoeff_eq_zero, not_or] using ⟨ha, hp⟩ + @[simp] lemma natDegree_pow (p : R[X]) (n : ℕ) : natDegree (p ^ n) = n * natDegree p := by classical diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 42daf78354251..41009d47c194c 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -53,8 +53,8 @@ theorem X_pow_dvd_iff {f : R[X]} {n : ℕ} : X ^ n ∣ f ↔ ∀ d < n, f.coeff variable {p q : R[X]} -theorem multiplicity_finite_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < degree p) (hmp : Monic p) - (hq : q ≠ 0) : multiplicity.Finite p q := +theorem finiteMultiplicity_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < degree p) (hmp : Monic p) + (hq : q ≠ 0) : FiniteMultiplicity p q := have zn0 : (0 : R) ≠ 1 := haveI := Nontrivial.of_polynomial_ne hq zero_ne_one @@ -76,6 +76,9 @@ theorem multiplicity_finite_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < deg (add_pos_of_pos_of_nonneg (by rwa [one_mul]) (Nat.zero_le _))) this⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity_finite_of_degree_pos_of_monic := finiteMultiplicity_of_degree_pos_of_monic + end Semiring section Ring @@ -470,12 +473,15 @@ 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) -theorem multiplicity_X_sub_C_finite (a : R) (h0 : p ≠ 0) : multiplicity.Finite (X - C a) p := by +theorem finiteMultiplicity_X_sub_C (a : R) (h0 : p ≠ 0) : FiniteMultiplicity (X - C a) p := by haveI := Nontrivial.of_polynomial_ne h0 - refine multiplicity_finite_of_degree_pos_of_monic ?_ (monic_X_sub_C _) h0 + refine finiteMultiplicity_of_degree_pos_of_monic ?_ (monic_X_sub_C _) h0 rw [degree_X_sub_C] decide +@[deprecated (since := "2024-11-30")] +alias multiplicity_X_sub_C_finite := finiteMultiplicity_X_sub_C + /- Porting note: stripping out classical for decidability instance parameter might make for better ergonomics -/ /-- The largest power of `X - C a` which divides `p`. @@ -488,7 +494,7 @@ def rootMultiplicity (a : R) (p : R[X]) : ℕ := let _ : DecidablePred fun n : ℕ => ¬(X - C a) ^ (n + 1) ∣ p := fun n => have := decidableDvdMonic p ((monic_X_sub_C a).pow (n + 1)) inferInstanceAs (Decidable ¬_) - Nat.find (multiplicity_X_sub_C_finite a h0) + Nat.find (finiteMultiplicity_X_sub_C a h0) /- Porting note: added the following due to diamond with decidableProp and decidableDvdMonic see also [Zulip] @@ -497,7 +503,7 @@ theorem rootMultiplicity_eq_nat_find_of_nonzero [DecidableEq R] {p : R[X]} (p0 : letI : DecidablePred fun n : ℕ => ¬(X - C a) ^ (n + 1) ∣ p := fun n => have := decidableDvdMonic p ((monic_X_sub_C a).pow (n + 1)) inferInstanceAs (Decidable ¬_) - rootMultiplicity a p = Nat.find (multiplicity_X_sub_C_finite a p0) := by + rootMultiplicity a p = Nat.find (finiteMultiplicity_X_sub_C a p0) := by dsimp [rootMultiplicity] cases Subsingleton.elim ‹DecidableEq R› (Classical.decEq R) rw [dif_neg p0] @@ -510,7 +516,7 @@ theorem rootMultiplicity_eq_multiplicity [DecidableEq R] split · rfl rename_i h - simp only [multiplicity_X_sub_C_finite a h, ↓reduceDIte] + simp only [finiteMultiplicity_X_sub_C a h, ↓reduceDIte] rw [← ENat.some_eq_coe, WithTop.untop'_coe] congr @@ -548,7 +554,7 @@ theorem exists_eq_pow_rootMultiplicity_mul_and_not_dvd (p : R[X]) (hp : p ≠ 0) ∃ q : R[X], p = (X - C a) ^ p.rootMultiplicity a * q ∧ ¬ (X - C a) ∣ q := by classical rw [rootMultiplicity_eq_multiplicity, if_neg hp] - apply (multiplicity_X_sub_C_finite a hp).exists_eq_pow_mul_and_not_dvd + apply (finiteMultiplicity_X_sub_C a hp).exists_eq_pow_mul_and_not_dvd end multiplicity @@ -633,7 +639,7 @@ theorem eval_divByMonic_pow_rootMultiplicity_ne_zero {p : R[X]} (a : R) (hp : p have := pow_mul_divByMonic_rootMultiplicity_eq p a rw [hq, ← mul_assoc, ← pow_succ, rootMultiplicity_eq_multiplicity, if_neg hp] at this exact - (multiplicity_finite_of_degree_pos_of_monic + (finiteMultiplicity_of_degree_pos_of_monic (show (0 : WithBot ℕ) < degree (X - C a) by rw [degree_X_sub_C]; decide) (monic_X_sub_C _) hp).not_pow_dvd_of_multiplicity_lt (Nat.lt_succ_self _) (dvd_of_mul_right_eq _ this) diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index bc0f49ad785ba..bb20dc5de2f78 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -175,11 +175,10 @@ See `isRoot_of_isRoot_iff_dvd_derivative_mul` -/ theorem isRoot_of_isRoot_of_dvd_derivative_mul [CharZero R] {f g : R[X]} (hf0 : f ≠ 0) (hfd : f ∣ f.derivative * g) {a : R} (haf : f.IsRoot a) : g.IsRoot a := by rcases hfd with ⟨r, hr⟩ - by_cases hdf0 : derivative f = 0 - · rw [eq_C_of_derivative_eq_zero hdf0] at haf hf0 - simp only [eval_C, derivative_C, zero_mul, dvd_zero, iff_true, IsRoot.def] at haf - rw [haf, map_zero] at hf0 - exact (hf0 rfl).elim + have hdf0 : derivative f ≠ 0 := by + contrapose! haf + rw [eq_C_of_derivative_eq_zero haf] at hf0 ⊢ + exact not_isRoot_C _ _ <| C_ne_zero.mp hf0 by_contra hg have hdfg0 : f.derivative * g ≠ 0 := mul_ne_zero hdf0 (by rintro rfl; simp at hg) have hr' := congr_arg (rootMultiplicity a) hr diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index b8bd2125c5b79..ad3c8e6a0e543 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -230,7 +230,7 @@ theorem rootMultiplicity_mul {p q : R[X]} {x : R} (hpq : p * q ≠ 0) : have hq : q ≠ 0 := right_ne_zero_of_mul hpq rw [rootMultiplicity_eq_multiplicity (p * q), if_neg hpq, rootMultiplicity_eq_multiplicity p, if_neg hp, rootMultiplicity_eq_multiplicity q, if_neg hq, - multiplicity_mul (prime_X_sub_C x) (multiplicity_X_sub_C_finite _ hpq)] + multiplicity_mul (prime_X_sub_C x) (finiteMultiplicity_X_sub_C _ hpq)] open Multiset in set_option linter.unusedVariables false in diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 26820e434059b..d81eb33fca30f 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -66,6 +66,7 @@ variable {R : Type u} {S : Type v} {T : Type w} [Ring R] variable [Ring S] [Ring T] namespace Subring +variable {s t : Subring R} -- Porting note: there is no `Subring.toSubmonoid` but we can't define it because there is a -- projection `s.toSubmonoid` @@ -78,6 +79,12 @@ theorem toSubsemiring_strictMono : StrictMono (toSubsemiring : Subring R → Sub theorem toSubsemiring_mono : Monotone (toSubsemiring : Subring R → Subsemiring R) := toSubsemiring_strictMono.monotone +@[gcongr] +lemma toSubsemiring_lt_toSubsemiring (hst : s < t) : s.toSubsemiring < t.toSubsemiring := hst + +@[gcongr] +lemma toSubsemiring_le_toSubsemiring (hst : s ≤ t) : s.toSubsemiring ≤ t.toSubsemiring := hst + @[mono] theorem toAddSubgroup_strictMono : StrictMono (toAddSubgroup : Subring R → AddSubgroup R) := fun _ _ => id @@ -86,6 +93,12 @@ theorem toAddSubgroup_strictMono : StrictMono (toAddSubgroup : Subring R → Add theorem toAddSubgroup_mono : Monotone (toAddSubgroup : Subring R → AddSubgroup R) := toAddSubgroup_strictMono.monotone +@[gcongr] +lemma toAddSubgroup_lt_toAddSubgroup (hst : s < t) : s.toAddSubgroup < t.toAddSubgroup := hst + +@[gcongr] +lemma toAddSubgroup_le_toAddSubgroup (hst : s ≤ t) : s.toAddSubgroup ≤ t.toAddSubgroup := hst + @[mono] theorem toSubmonoid_strictMono : StrictMono (fun s : Subring R => s.toSubmonoid) := fun _ _ => id @@ -656,7 +669,7 @@ theorem coe_prod (s : Subring R) (t : Subring S) : theorem mem_prod {s : Subring R} {t : Subring S} {p : R × S} : p ∈ s.prod t ↔ p.1 ∈ s ∧ p.2 ∈ t := Iff.rfl -@[mono] +@[gcongr, mono] theorem prod_mono ⦃s₁ s₂ : Subring R⦄ (hs : s₁ ≤ s₂) ⦃t₁ t₂ : Subring S⦄ (ht : t₁ ≤ t₂) : s₁.prod t₁ ≤ s₂.prod t₂ := Set.prod_mono hs ht diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 65c4d305a67a1..88035633f260c 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -589,7 +589,7 @@ theorem mem_prod {s : Subsemiring R} {t : Subsemiring S} {p : R × S} : p ∈ s.prod t ↔ p.1 ∈ s ∧ p.2 ∈ t := Iff.rfl -@[mono] +@[gcongr, mono] theorem prod_mono ⦃s₁ s₂ : Subsemiring R⦄ (hs : s₁ ≤ s₂) ⦃t₁ t₂ : Subsemiring S⦄ (ht : t₁ ≤ t₂) : s₁.prod t₁ ≤ s₂.prod t₂ := Set.prod_mono hs ht diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 9d68d11fc173a..bdc6c55ec4a16 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -103,13 +103,7 @@ theorem Squarefree.gcd_left {a : α} (b : α) (ha : Squarefree a) : Squarefree ( end SquarefreeGcdOfSquarefree -namespace multiplicity - -section CommMonoid - -variable [CommMonoid R] - -theorem squarefree_iff_emultiplicity_le_one (r : R) : +theorem squarefree_iff_emultiplicity_le_one [CommMonoid R] (r : R) : Squarefree r ↔ ∀ x : R, emultiplicity x r ≤ 1 ∨ IsUnit x := by refine forall_congr' fun a => ?_ rw [← sq, pow_dvd_iff_le_emultiplicity, or_iff_not_imp_left, not_le, imp_congr _ Iff.rfl] @@ -117,18 +111,8 @@ theorem squarefree_iff_emultiplicity_le_one (r : R) : rw [← one_add_one_eq_two] exact Order.add_one_le_iff_of_not_isMax (by simp) -end CommMonoid - -section CancelCommMonoidWithZero - -variable [CancelCommMonoidWithZero R] [WfDvdMonoid R] - -theorem finite_prime_left {a b : R} (ha : Prime a) (hb : b ≠ 0) : multiplicity.Finite a b := - finite_of_not_isUnit ha.not_unit hb - -end CancelCommMonoidWithZero - -end multiplicity +@[deprecated (since := "2024-11-30")] +alias multiplicity.squarefree_iff_emultiplicity_le_one := squarefree_iff_emultiplicity_le_one section Irreducible @@ -266,7 +250,7 @@ lemma _root_.exists_squarefree_dvd_pow_of_ne_zero {x : R} (hx : x ≠ 0) : theorem squarefree_iff_nodup_normalizedFactors [NormalizationMonoid R] {x : R} (x0 : x ≠ 0) : Squarefree x ↔ Multiset.Nodup (normalizedFactors x) := by classical - rw [multiplicity.squarefree_iff_emultiplicity_le_one, Multiset.nodup_iff_count_le_one] + rw [squarefree_iff_emultiplicity_le_one, Multiset.nodup_iff_count_le_one] haveI := nontrivial_of_ne x 0 x0 constructor <;> intro h a · by_cases hmem : a ∈ normalizedFactors x diff --git a/Mathlib/Algebra/Vertex/HVertexOperator.lean b/Mathlib/Algebra/Vertex/HVertexOperator.lean index 4248cbf1e971d..d68241ff4baf6 100644 --- a/Mathlib/Algebra/Vertex/HVertexOperator.lean +++ b/Mathlib/Algebra/Vertex/HVertexOperator.lean @@ -92,6 +92,16 @@ def of_coeff (f : Γ → V →ₗ[R] W) @[deprecated (since := "2024-06-18")] alias _root_.VertexAlg.HetVertexOperator.of_coeff := of_coeff +@[simp] +theorem add_coeff (A B : HVertexOperator Γ R V W) : (A + B).coeff = A.coeff + B.coeff := by + ext + simp + +@[simp] +theorem smul_coeff (A : HVertexOperator Γ R V W) (r : R) : (r • A).coeff = r • (A.coeff) := by + ext + simp + end Coeff section Products diff --git a/Mathlib/Algebra/Vertex/VertexOperator.lean b/Mathlib/Algebra/Vertex/VertexOperator.lean new file mode 100644 index 0000000000000..e7eaef0b16c08 --- /dev/null +++ b/Mathlib/Algebra/Vertex/VertexOperator.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2024 Scott Carnahan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Carnahan +-/ +import Mathlib.Algebra.Vertex.HVertexOperator +import Mathlib.Data.Int.Interval + +/-! +# Vertex operators +In this file we introduce vertex operators as linear maps to Laurent series. + +## Definitions + * `VertexOperator` is an `R`-linear map from an `R`-module `V` to `LaurentSeries V`. + * `VertexOperator.ncoeff` is the coefficient of a vertex operator under normalized indexing. + +## TODO + * `HasseDerivative` : A divided-power derivative. + * `Locality` : A weak form of commutativity. + * `Residue products` : A family of products on `VertexOperator R V` parametrized by integers. + +## References + * [G. Mason, *Vertex rings and Pierce bundles*][mason2017] + * [A. Matsuo, K. Nagatomo, *On axioms for a vertex algebra and locality of quantum + fields*][matsuo1997] +-/ + +noncomputable section + +variable {R V : Type*} [CommRing R] [AddCommGroup V] [Module R V] + +/-- A vertex operator over a commutative ring `R` is an `R`-linear map from an `R`-module `V` to +Laurent series with coefficients in `V`. We write this as a specialization of the heterogeneous +case. -/ +abbrev VertexOperator (R : Type*) (V : Type*) [CommRing R] [AddCommGroup V] + [Module R V] := HVertexOperator ℤ R V V + +namespace VertexOperator + +open HVertexOperator + +@[ext] +theorem ext (A B : VertexOperator R V) (h : ∀ v : V, A v = B v) : + A = B := LinearMap.ext h + +/-- The coefficient of a vertex operator under normalized indexing. -/ +def ncoeff {R} [CommRing R] [AddCommGroup V] [Module R V] (A : VertexOperator R V) (n : ℤ) : + Module.End R V := HVertexOperator.coeff A (-n - 1) + +/-- In the literature, the `n`th normalized coefficient of a vertex operator `A` is written as +either `Aₙ` or `A(n)`. -/ +scoped[VertexOperator] notation A "[[" n "]]" => ncoeff A n + +@[simp] +theorem coeff_eq_ncoeff (A : VertexOperator R V) + (n : ℤ) : HVertexOperator.coeff A n = A [[-n - 1]] := by + rw [ncoeff, neg_sub, sub_neg_eq_add, add_sub_cancel_left] + +@[simp] +theorem ncoeff_add (A B : VertexOperator R V) (n : ℤ) : (A + B) [[n]] = A [[n]] + B [[n]] := by + rw [ncoeff, ncoeff, ncoeff, add_coeff, Pi.add_apply] + +@[simp] +theorem ncoeff_smul (A : VertexOperator R V) (r : R) (n : ℤ) : (r • A) [[n]] = r • A [[n]] := by + rw [ncoeff, ncoeff, smul_coeff, Pi.smul_apply] + +theorem ncoeff_eq_zero_of_lt_order (A : VertexOperator R V) (n : ℤ) (x : V) + (h : -n - 1 < HahnSeries.order ((HahnModule.of R).symm (A x))) : (A [[n]]) x = 0 := by + simp only [ncoeff, HVertexOperator.coeff, LinearMap.coe_mk, AddHom.coe_mk] + exact HahnSeries.coeff_eq_zero_of_lt_order h + +theorem coeff_eq_zero_of_lt_order (A : VertexOperator R V) (n : ℤ) (x : V) + (h : n < HahnSeries.order ((HahnModule.of R).symm (A x))) : coeff A n x = 0 := by + rw [coeff_eq_ncoeff, ncoeff_eq_zero_of_lt_order A (-n - 1) x] + omega + +/-- Given an endomorphism-valued function on integers satisfying a pointwise bounded-pole condition, +we produce a vertex operator. -/ +noncomputable def of_coeff (f : ℤ → Module.End R V) + (hf : ∀ (x : V), ∃ (n : ℤ), ∀ (m : ℤ), m < n → (f m) x = 0) : VertexOperator R V := + HVertexOperator.of_coeff f + (fun x => HahnSeries.suppBddBelow_supp_PWO (fun n => (f n) x) + (HahnSeries.forallLTEqZero_supp_BddBelow (fun n => (f n) x) + (Exists.choose (hf x)) (Exists.choose_spec (hf x)))) + +@[simp] +theorem of_coeff_apply_coeff (f : ℤ → Module.End R V) + (hf : ∀ (x : V), ∃ n, ∀ m < n, (f m) x = 0) (x : V) (n : ℤ) : + ((HahnModule.of R).symm ((of_coeff f hf) x)).coeff n = (f n) x := by + rfl + +@[simp] +theorem ncoeff_of_coeff (f : ℤ → Module.End R V) + (hf : ∀(x : V), ∃(n : ℤ), ∀(m : ℤ), m < n → (f m) x = 0) (n : ℤ) : + (of_coeff f hf) [[n]] = f (-n - 1) := by + ext v + rw [ncoeff, coeff_apply, of_coeff_apply_coeff] + +end VertexOperator diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 7751d57d64eda..ac5248a0b208b 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -295,9 +295,7 @@ lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : map n (f ≫ g) = map n f ≫ map n g := by ext1 · simp - · simp only [TopologicalSpace.Opens.map_top, Scheme.comp_coeBase, - TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, CommRingCat.comp_apply] - rw [map_appTop_coord, map_appTop_coord, map_appTop_coord] + · simp lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : map n (Spec.map φ) = diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean index ac58cdee1d2ca..c92f4e2898103 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Degree.lean @@ -249,7 +249,7 @@ lemma natDegree_preΨ' {n : ℕ} (h : (n : R) ≠ 0) : natDegree_eq_of_le_of_coeff_ne_zero (W.natDegree_preΨ'_le n) <| W.coeff_preΨ'_ne_zero h lemma natDegree_preΨ'_pos {n : ℕ} (hn : 2 < n) (h : (n : R) ≠ 0) : 0 < (W.preΨ' n).natDegree := by - rw [W.natDegree_preΨ' h, Nat.div_pos_iff two_ne_zero] + simp only [W.natDegree_preΨ' h, Nat.div_pos_iff, zero_lt_two, true_and] split_ifs <;> exact Nat.AtLeastTwo.prop.trans <| Nat.sub_le_sub_right (Nat.pow_le_pow_of_le_left hn 2) _ diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index 12264f5de8c43..763031d694da3 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -158,7 +158,11 @@ noncomputable instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) : noncomputable def tilde : (Spec (CommRingCat.of R)).Modules where val := { obj := fun U ↦ ModuleCat.of _ (M.tildeInType.val.obj U) - map := fun {U V} i ↦ + map := fun {U V} i ↦ ofHom + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)` + -- This suggests `restrictScalars` needs to be redesigned. + (Y := (restrictScalars ((Spec (CommRingCat.of R)).ringCatSheaf.val.map i)).obj + (of ((Spec (CommRingCat.of R)).ringCatSheaf.val.obj V) (M.tildeInType.val.obj V))) { toFun := M.tildeInType.val.map i map_smul' := by intros; rfl map_add' := by intros; rfl } } @@ -200,15 +204,18 @@ If `U` is an open subset of `Spec R`, this is the morphism of `R`-modules from ` `M^~(U)`. -/ def toOpen (U : Opens (PrimeSpectrum.Top R)) : - ModuleCat.of R M ⟶ (tildeInModuleCat M).1.obj (op U) where - toFun f := - ⟨fun x ↦ LocalizedModule.mkLinearMap _ _ f, fun x ↦ - ⟨U, x.2, 𝟙 _, f, 1, fun y ↦ ⟨(Ideal.ne_top_iff_one _).1 y.1.2.1, by simp⟩⟩⟩ - map_add' f g := Subtype.eq <| funext fun x ↦ LinearMap.map_add _ _ _ - map_smul' r m := by - simp only [isLocallyFraction_pred, LocalizedModule.mkLinearMap_apply, LinearMapClass.map_smul, - RingHom.id_apply] - rfl + ModuleCat.of R M ⟶ (tildeInModuleCat M).1.obj (op U) := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)` + -- This suggests `restrictScalars` needs to be redesigned. + ModuleCat.ofHom (Y := (tildeInModuleCat M).1.obj (op U)) + { toFun := fun f => + ⟨fun x ↦ LocalizedModule.mkLinearMap _ _ f, fun x ↦ + ⟨U, x.2, 𝟙 _, f, 1, fun y ↦ ⟨(Ideal.ne_top_iff_one _).1 y.1.2.1, by simp⟩⟩⟩ + map_add' := fun f g => Subtype.eq <| funext fun x ↦ LinearMap.map_add _ _ _ + map_smul' := fun r m => by + simp only [isLocallyFraction_pred, LocalizedModule.mkLinearMap_apply, LinearMapClass.map_smul, + RingHom.id_apply] + rfl } @[simp] theorem toOpen_res (U V : Opens (PrimeSpectrum.Top R)) (i : V ⟶ U) : @@ -235,7 +242,8 @@ lemma isUnit_toStalk (x : PrimeSpectrum.Top R) (r : x.asIdeal.primeCompl) : ⟨fun q ↦ (Localization.mk 1 ⟨r, q.2.2⟩ : Localization.AtPrime q.1.asIdeal) • s.1 ⟨q.1, q.2.1⟩, fun q ↦ ?_⟩, by simpa only [Module.algebraMap_end_apply, ← map_smul] using - germ_ext (W := O) (hxW := ⟨mem, r.2⟩) (iWU := 𝟙 _) (iWV := homOfLE inf_le_left) _ <| + germ_ext (C := ModuleCat R) (W := O) (hxW := ⟨mem, r.2⟩) (iWU := 𝟙 _) + (iWV := homOfLE inf_le_left) _ <| Subtype.eq <| funext fun y ↦ smul_eq_iff_of_mem (S := y.1.1.primeCompl) r _ _ _ |>.2 rfl⟩ obtain ⟨V, mem_V, iV, num, den, hV⟩ := s.2 ⟨q.1, q.2.1⟩ refine ⟨V ⊓ O, ⟨mem_V, q.2⟩, homOfLE inf_le_right, num, r * den, fun y ↦ ?_⟩ @@ -252,7 +260,7 @@ to the stalk of `M^~` at `x`. noncomputable def localizationToStalk (x : PrimeSpectrum.Top R) : ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) ⟶ (TopCat.Presheaf.stalk (tildeInModuleCat M) x) := - LocalizedModule.lift _ (toStalk M x) <| isUnit_toStalk M x + ModuleCat.ofHom <| LocalizedModule.lift _ (toStalk M x).hom <| isUnit_toStalk M x /-- The ring homomorphism that takes a section of the structure sheaf of `R` on the open set `U`, @@ -260,10 +268,15 @@ implemented as a subtype of dependent functions to localizations at prime ideals the section on the point corresponding to a given prime ideal. -/ def openToLocalization (U : Opens (PrimeSpectrum R)) (x : PrimeSpectrum R) (hx : x ∈ U) : (tildeInModuleCat M).obj (op U) ⟶ - ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where - toFun s := (s.1 ⟨x, hx⟩ : _) - map_add' _ _ := rfl - map_smul' _ _ := rfl + ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) := + -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(X := ...)` and `(Y := ...)` + -- This suggests `restrictScalars` needs to be redesigned. + ModuleCat.ofHom + (X := (tildeInModuleCat M).obj (op U)) + (Y := ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M)) + { toFun := fun s => (s.1 ⟨x, hx⟩ : _) + map_add' := fun _ _ => rfl + map_smul' := fun _ _ => rfl } /-- The morphism of `R`-modules from the stalk of `M^~` at `x` to the localization of `M` at the @@ -286,9 +299,9 @@ theorem germ_comp_stalkToFiberLinearMap (U : Opens (PrimeSpectrum.Top R)) (x) (h @[simp] theorem stalkToFiberLinearMap_germ (U : Opens (PrimeSpectrum.Top R)) (x : PrimeSpectrum.Top R) (hx : x ∈ U) (s : (tildeInModuleCat M).1.obj (op U)) : - stalkToFiberLinearMap M x + (stalkToFiberLinearMap M x).hom (TopCat.Presheaf.germ (tildeInModuleCat M) U x hx s) = (s.1 ⟨x, hx⟩ : _) := - DFunLike.ext_iff.1 (germ_comp_stalkToFiberLinearMap M U x hx) s + DFunLike.ext_iff.1 (ModuleCat.hom_ext_iff.mp (germ_comp_stalkToFiberLinearMap M U x hx)) s @[reassoc (attr := simp), elementwise (attr := simp)] theorem toOpen_germ (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : @@ -298,13 +311,13 @@ theorem toOpen_germ (U : Opens (PrimeSpectrum.Top R)) (x) (hx : x ∈ U) : @[reassoc (attr := simp)] theorem toStalk_comp_stalkToFiberLinearMap (x : PrimeSpectrum.Top R) : toStalk M x ≫ stalkToFiberLinearMap M x = - LocalizedModule.mkLinearMap x.asIdeal.primeCompl M := by + ofHom (LocalizedModule.mkLinearMap x.asIdeal.primeCompl M) := by rw [toStalk, Category.assoc, germ_comp_stalkToFiberLinearMap]; rfl theorem stalkToFiberLinearMap_toStalk (x : PrimeSpectrum.Top R) (m : M) : - (stalkToFiberLinearMap M x) (toStalk M x m) = + (stalkToFiberLinearMap M x).hom (toStalk M x m) = LocalizedModule.mk m 1 := - LinearMap.ext_iff.1 (toStalk_comp_stalkToFiberLinearMap M x) _ + LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (toStalk_comp_stalkToFiberLinearMap M x)) _ /-- If `U` is an open subset of `Spec R`, `m` is an element of `M` and `r` is an element of `R` @@ -347,7 +360,7 @@ theorem res_const (f : M) (g : R) (U hu V hv i) : @[simp] theorem localizationToStalk_mk (x : PrimeSpectrum.Top R) (f : M) (s : x.asIdeal.primeCompl) : - localizationToStalk M x (LocalizedModule.mk f s) = + (localizationToStalk M x).hom (LocalizedModule.mk f s) = (tildeInModuleCat M).germ (PrimeSpectrum.basicOpen (s : R)) x s.2 (const M f s (PrimeSpectrum.basicOpen s) fun _ => id) := (Module.End_isUnit_iff _ |>.1 (isUnit_toStalk M x s)).injective <| by @@ -356,6 +369,7 @@ theorem localizationToStalk_mk (x : PrimeSpectrum.Top R) (f : M) (s : x.asIdeal. show (M.tildeInModuleCat.germ ⊤ x ⟨⟩) ((toOpen M ⊤) f) = _ rw [← map_smul] fapply TopCat.Presheaf.germ_ext (W := PrimeSpectrum.basicOpen s.1) (hxW := s.2) + (F := M.tildeInModuleCat) · exact homOfLE le_top · exact 𝟙 _ refine Subtype.eq <| funext fun y => show LocalizedModule.mk f 1 = _ from ?_ @@ -376,7 +390,8 @@ noncomputable def stalkIso (x : PrimeSpectrum.Top R) : ModuleCat.of R (LocalizedModule x.asIdeal.primeCompl M) where hom := stalkToFiberLinearMap M x inv := localizationToStalk M x - hom_inv_id := TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ext _ fun s ↦ by + hom_inv_id := TopCat.Presheaf.stalk_hom_ext _ fun U hxU ↦ ModuleCat.hom_ext <| + LinearMap.ext fun s ↦ by show localizationToStalk M x (stalkToFiberLinearMap M x (M.tildeInModuleCat.germ U x hxU s)) = M.tildeInModuleCat.germ U x hxU s rw [stalkToFiberLinearMap_germ] @@ -384,15 +399,19 @@ noncomputable def stalkIso (x : PrimeSpectrum.Top R) : exists_const _ _ s x hxU rw [← res_apply M U V iVU s ⟨x, hxV⟩, ← hs, const_apply, localizationToStalk_mk] exact (tildeInModuleCat M).germ_ext V hxV (homOfLE hg) iVU <| hs ▸ rfl - inv_hom_id := by ext x; exact x.induction_on (fun _ _ => by simp) + inv_hom_id := by ext x; exact x.induction_on (fun _ _ => by + simp only [hom_comp, LinearMap.coe_comp, Function.comp_apply, hom_id, LinearMap.id_coe, id_eq] + rw [localizationToStalk_mk, stalkToFiberLinearMap_germ] + simp) instance (x : PrimeSpectrum.Top R) : - IsLocalizedModule x.asIdeal.primeCompl (toStalk M x) := by + IsLocalizedModule x.asIdeal.primeCompl (toStalk M x).hom := by convert IsLocalizedModule.of_linearEquiv (hf := localizedModuleIsLocalizedModule (M := M) x.asIdeal.primeCompl) (e := (stalkIso M x).symm.toLinearEquiv) - simp only [of_coe, show (stalkIso M x).symm.toLinearEquiv.toLinearMap = (stalkIso M x).inv by rfl, - stalkIso_inv] + ext + simp only [of_coe, + show (stalkIso M x).symm.toLinearEquiv.toLinearMap = (stalkIso M x).inv.hom by rfl] erw [LocalizedModule.lift_comp] end Tilde diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean index 57274ae5ae060..f7cf9b2d50040 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Jacobson.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian -import Mathlib.RingTheory.Jacobson +import Mathlib.RingTheory.Jacobson.Ring import Mathlib.Topology.JacobsonSpace /-! diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 115b2bc0a1851..156dc3423a3ca 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -231,6 +231,10 @@ def mkOfLe {n} (i j : Fin (n+1)) (h : i ≤ j) : ([1] : SimplexCategory) ⟶ [n] | 0, 1, _ => h } +@[simp] +lemma mkOfLe_refl {n} (j : Fin (n + 1)) : + mkOfLe j j (by omega) = [1].const [n] j := Hom.ext_one_left _ _ + /-- The morphism `[1] ⟶ [n]` that picks out the "diagonal composite" edge-/ def diag (n : ℕ) : ([1] : SimplexCategory) ⟶ [n] := mkOfLe 0 n (Fin.zero_le _) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 6e545ab1fa48e..be2e7617b9773 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -384,10 +384,10 @@ protected abbrev Truncated.cosk (n : ℕ) : SSet.Truncated n ⥤ SSet.{u} := SimplicialObject.Truncated.cosk n /-- The n-skeleton as an endofunctor on `SSet`. -/ -abbrev sk (n : ℕ) : SSet ⥤ SSet := SimplicialObject.sk n +abbrev sk (n : ℕ) : SSet.{u} ⥤ SSet.{u} := SimplicialObject.sk n /-- The n-coskeleton as an endofunctor on `SSet`. -/ -abbrev cosk (n : ℕ) : SSet ⥤ SSet := SimplicialObject.cosk n +abbrev cosk (n : ℕ) : SSet.{u} ⥤ SSet.{u} := SimplicialObject.cosk n end diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean new file mode 100644 index 0000000000000..93eed2ed3a913 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Coskeletal.lean @@ -0,0 +1,256 @@ +/- +Copyright (c) 2024 Emily Riehl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Emily Riehl, Joël Riou +-/ +import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal +import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal +import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction +import Mathlib.CategoryTheory.Functor.KanExtension.Basic + +/-! +# Coskeletal simplicial sets + +In this file, we prove that if `X` is `StrictSegal` then `X` is 2-coskeletal, +i.e. `X ≅ (cosk 2).obj X`. In particular, nerves of categories are 2-coskeletal. + +This isomorphism follows from the fact that `rightExtensionInclusion X 2` is a right Kan +extension. In fact, we show that when `X` is `StrictSegal` then +`(rightExtensionInclusion X n).IsPointwiseRightKanExtension` holds. + +As an example, `SimplicialObject.IsCoskeletal (nerve C) 2` shows that that nerves of categories +are 2-coskeletal. +-/ + + +universe v u + +open CategoryTheory Simplicial SimplexCategory Opposite Category Functor Limits + +namespace SSet + +namespace Truncated + +/-- The identity natural transformation exhibits a simplicial set as a right extension of its +restriction along `(Truncated.inclusion (n := n)).op`.-/ +@[simps!] +def rightExtensionInclusion (X : SSet.{u}) (n : ℕ) : + RightExtension (Truncated.inclusion (n := n)).op + ((Truncated.inclusion n).op ⋙ X) := RightExtension.mk _ (𝟙 _) + +end Truncated + +section + +local notation (priority := high) "[" n "]" => SimplexCategory.mk n + +local macro:max (priority := high) "[" n:term "]₂" : term => + `((⟨SimplexCategory.mk $n, by dsimp; omega⟩ : SimplexCategory.Truncated 2)) + +open StructuredArrow + +namespace StrictSegal +variable (X : SSet.{u}) [StrictSegal X] + +namespace isPointwiseRightKanExtensionAt + +/-- A morphism in `SimplexCategory` with domain `[0]`, `[1]`, or `[2]` defines an object in the +comma category `StructuredArrow (op [n]) (Truncated.inclusion (n := 2)).op`.-/ +abbrev strArrowMk₂ {i : ℕ} {n : ℕ} (φ : [i] ⟶ [n]) (hi : i ≤ 2) : + StructuredArrow (op [n]) (Truncated.inclusion (n := 2)).op := + StructuredArrow.mk (Y := op ⟨[i], hi⟩) (by exact φ.op) + +/-- Given a term in the cone over the diagram +`(proj (op [n]) ((Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)` where `X` is +Strict Segal, one can produce an `n`-simplex in `X`. -/ +@[simp] +noncomputable def lift {X : SSet.{u}} [StrictSegal X] {n} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ + (Truncated.inclusion 2).op ⋙ X)) (x : s.pt) : X _[n] := + StrictSegal.spineToSimplex { + vertex := fun i ↦ s.π.app (.mk (Y := op [0]₂) (.op (SimplexCategory.const _ _ i))) x + arrow := fun i ↦ s.π.app (.mk (Y := op [1]₂) (.op (mkOfLe _ _ (Fin.castSucc_le_succ i)))) x + arrow_src := fun i ↦ by + let φ : strArrowMk₂ (mkOfLe _ _ (Fin.castSucc_le_succ i)) (by simp) ⟶ + strArrowMk₂ ([0].const _ i.castSucc) (by simp) := + StructuredArrow.homMk (δ 1).op + (Quiver.Hom.unop_inj (by ext x; fin_cases x; rfl)) + exact congr_fun (s.w φ) x + arrow_tgt := fun i ↦ by + dsimp + let φ : strArrowMk₂ (mkOfLe _ _ (Fin.castSucc_le_succ i)) (by simp) ⟶ + strArrowMk₂ ([0].const _ i.succ) (by simp) := + StructuredArrow.homMk (δ 0).op + (Quiver.Hom.unop_inj (by ext x; fin_cases x; rfl)) + exact congr_fun (s.w φ) x } + +lemma fac_aux₁ {n : ℕ} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)) + (x : s.pt) (i : ℕ) (hi : i < n) : + X.map (mkOfSucc ⟨i, hi⟩).op (lift s x) = + s.π.app (strArrowMk₂ (mkOfSucc ⟨i, hi⟩) (by omega)) x := by + dsimp [lift] + rw [spineToSimplex_arrow] + rfl + +lemma fac_aux₂ {n : ℕ} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)) + (x : s.pt) (i j : ℕ) (hij : i ≤ j) (hj : j ≤ n) : + X.map (mkOfLe ⟨i, by omega⟩ ⟨j, by omega⟩ hij).op (lift s x) = + s.π.app (strArrowMk₂ (mkOfLe ⟨i, by omega⟩ ⟨j, by omega⟩ hij) (by omega)) x := by + obtain ⟨k, hk⟩ := Nat.le.dest hij + revert i j + induction k with + | zero => + rintro i j hij hj hik + obtain rfl : i = j := by omega + have : mkOfLe ⟨i, Nat.lt_add_one_of_le hj⟩ ⟨i, Nat.lt_add_one_of_le hj⟩ (by omega) = + [1].const [0] 0 ≫ [0].const [n] ⟨i, Nat.lt_add_one_of_le hj⟩ := Hom.ext_one_left _ _ + rw [this] + let α : (strArrowMk₂ ([0].const [n] ⟨i, Nat.lt_add_one_of_le hj⟩) (by omega)) ⟶ + (strArrowMk₂ ([1].const [0] 0 ≫ [0].const [n] ⟨i, Nat.lt_add_one_of_le hj⟩) (by omega)) := + StructuredArrow.homMk (([1].const [0] 0).op) (by simp; rfl) + have nat := congr_fun (s.π.naturality α) x + dsimp only [Fin.val_zero, Nat.add_zero, id_eq, Int.reduceNeg, Int.Nat.cast_ofNat_Int, + Int.reduceAdd, Fin.eta, comp_obj, StructuredArrow.proj_obj, op_obj, const_obj_obj, + const_obj_map, types_comp_apply, types_id_apply, Functor.comp_map, StructuredArrow.proj_map, + op_map] at nat + rw [nat, op_comp, Functor.map_comp] + simp only [types_comp_apply] + refine congrArg (X.map ([1].const [0] 0).op) ?_ + unfold strArrowMk₂ + rw [lift, StrictSegal.spineToSimplex_vertex] + congr + | succ k hk => + intro i j hij hj hik + let α := strArrowMk₂ (mkOfLeComp (n := n) ⟨i, by omega⟩ ⟨i + k, by omega⟩ + ⟨j, by omega⟩ (by simp) + (by simp only [Fin.mk_le_mk]; omega)) (by rfl) + let α₀ := strArrowMk₂ (mkOfLe (n := n) ⟨i + k, by omega⟩ ⟨j, by omega⟩ + (by simp only [Fin.mk_le_mk]; omega)) (by simp) + let α₁ := strArrowMk₂ (mkOfLe (n := n) ⟨i, by omega⟩ ⟨j, by omega⟩ + (by simp only [Fin.mk_le_mk]; omega)) (by simp) + let α₂ := strArrowMk₂ (mkOfLe (n := n) ⟨i, by omega⟩ ⟨i + k, by omega⟩ (by simp)) (by simp) + let β₀ : α ⟶ α₀ := StructuredArrow.homMk ((mkOfSucc 1).op) (Quiver.Hom.unop_inj + (by ext x; fin_cases x <;> rfl)) + let β₁ : α ⟶ α₁ := StructuredArrow.homMk ((δ 1).op) (Quiver.Hom.unop_inj + (by ext x; fin_cases x <;> rfl)) + let β₂ : α ⟶ α₂ := StructuredArrow.homMk ((mkOfSucc 0).op) (Quiver.Hom.unop_inj + (by ext x; fin_cases x <;> rfl)) + have h₀ : X.map α₀.hom (lift s x) = s.π.app α₀ x := by + obtain rfl : j = (i + k) + 1 := by omega + exact fac_aux₁ _ _ _ _ (by omega) + have h₂ : X.map α₂.hom (lift s x) = s.π.app α₂ x := + hk i (i + k) (by simp) (by omega) rfl + change X.map α₁.hom (lift s x) = s.π.app α₁ x + have : X.map α.hom (lift s x) = s.π.app α x := by + apply StrictSegal.spineInjective + apply Path.ext' + intro t + dsimp only [spineEquiv] + rw [Equiv.coe_fn_mk, spine_arrow, spine_arrow, + ← FunctorToTypes.map_comp_apply] + match t with + | 0 => + have : α.hom ≫ (mkOfSucc 0).op = α₂.hom := + Quiver.Hom.unop_inj (by ext x ; fin_cases x <;> rfl) + rw [this, h₂, ← congr_fun (s.w β₂) x] + rfl + | 1 => + have : α.hom ≫ (mkOfSucc 1).op = α₀.hom := + Quiver.Hom.unop_inj (by ext x ; fin_cases x <;> rfl) + rw [this, h₀, ← congr_fun (s.w β₀) x] + rfl + rw [← StructuredArrow.w β₁, FunctorToTypes.map_comp_apply, this, ← s.w β₁] + dsimp + +lemma fac_aux₃ {n : ℕ} + (s : Cone (proj (op [n]) (Truncated.inclusion 2).op ⋙ (Truncated.inclusion 2).op ⋙ X)) + (x : s.pt) (φ : [1] ⟶ [n]) : + X.map φ.op (lift s x) = s.π.app (strArrowMk₂ φ (by omega)) x := by + obtain ⟨i, j, hij, rfl⟩ : ∃ i j hij, φ = mkOfLe i j hij := + ⟨φ.toOrderHom 0, φ.toOrderHom 1, φ.toOrderHom.monotone (by simp), + Hom.ext_one_left _ _ rfl rfl⟩ + exact fac_aux₂ _ _ _ _ _ _ (by omega) + +end isPointwiseRightKanExtensionAt + +open Truncated + +open isPointwiseRightKanExtensionAt in +/-- A strict Segal simplicial set is 2-coskeletal. -/ +noncomputable def isPointwiseRightKanExtensionAt (n : ℕ) : + (rightExtensionInclusion X 2).IsPointwiseRightKanExtensionAt ⟨[n]⟩ where + lift s x := lift (X := X) s x + fac s j := by + ext x + obtain ⟨⟨i, hi⟩, ⟨f : _ ⟶ _⟩, rfl⟩ := j.mk_surjective + obtain ⟨i, rfl⟩ : ∃ j, SimplexCategory.mk j = i := ⟨_, i.mk_len⟩ + dsimp at hi ⊢ + apply StrictSegal.spineInjective + dsimp + ext k + · dsimp only [spineEquiv, Equiv.coe_fn_mk] + erw [spine_map_vertex] + rw [spine_spineToSimplex, spine_vertex] + let α : strArrowMk₂ f hi ⟶ strArrowMk₂ ([0].const [n] (f.toOrderHom k)) (by omega) := + StructuredArrow.homMk (([0].const _ (by exact k)).op) (by simp; rfl) + exact congr_fun (s.w α).symm x + · dsimp only [spineEquiv, Equiv.coe_fn_mk, spine_arrow] + rw [← FunctorToTypes.map_comp_apply] + let α : strArrowMk₂ f hi ⟶ strArrowMk₂ (mkOfSucc k ≫ f) (by omega) := + StructuredArrow.homMk (mkOfSucc k).op (by simp; rfl) + exact (isPointwiseRightKanExtensionAt.fac_aux₃ _ _ _ _).trans (congr_fun (s.w α).symm x) + uniq s m hm := by + ext x + apply StrictSegal.spineInjective (X := X) + dsimp [spineEquiv] + rw [StrictSegal.spine_spineToSimplex] + ext i + · exact congr_fun (hm (StructuredArrow.mk (Y := op [0]₂) ([0].const [n] i).op)) x + · exact congr_fun (hm (.mk (Y := op [1]₂) (.op (mkOfLe _ _ (Fin.castSucc_le_succ i))))) x + +/-- Since `StrictSegal.isPointwiseRightKanExtensionAt` proves that the appropriate +cones are limit cones, `rightExtensionInclusion X 2` is a pointwise right Kan extension.-/ +noncomputable def isPointwiseRightKanExtension : + (rightExtensionInclusion X 2).IsPointwiseRightKanExtension := + fun Δ => isPointwiseRightKanExtensionAt X Δ.unop.len + +theorem isRightKanExtension : + X.IsRightKanExtension (𝟙 ((inclusion 2).op ⋙ X)) := + RightExtension.IsPointwiseRightKanExtension.isRightKanExtension + (isPointwiseRightKanExtension X) + +/-- When `X` is `StrictSegal`, `X` is 2-coskeletal. -/ +instance isCoskeletal : SimplicialObject.IsCoskeletal X 2 where + isRightKanExtension := isRightKanExtension X + +end StrictSegal + +end + +end SSet + +namespace CategoryTheory + +namespace Nerve + +open SSet + +example (C : Type u) [Category.{v} C] : + SimplicialObject.IsCoskeletal (nerve C) 2 := by infer_instance + +/-- The essential data of the nerve functor is contained in the 2-truncation, which is +recorded by the composite functor `nerveFunctor₂`.-/ +def nerveFunctor₂ : Cat.{v, u} ⥤ SSet.Truncated 2 := nerveFunctor ⋙ truncation 2 + +/-- The natural isomorphism between `nerveFunctor` and `nerveFunctor₂ ⋙ Truncated.cosk 2` whose +components `nerve C ≅ (Truncated.cosk 2).obj (nerveFunctor₂.obj C)` shows that nerves of categories +are 2-coskeletal.-/ +noncomputable def cosk₂Iso : nerveFunctor.{v, u} ≅ nerveFunctor₂.{v, u} ⋙ Truncated.cosk 2 := + NatIso.ofComponents (fun C ↦ (nerve C).isoCoskOfIsCoskeletal 2) + (fun _ ↦ (coskAdj 2).unit.naturality _) + +end Nerve + +end CategoryTheory diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean index c82254637aaf3..a24814b5a2a14 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Nerve.lean @@ -37,7 +37,7 @@ instance {C : Type*} [Category C] {Δ : SimplexCategoryᵒᵖ} : Category ((nerv /-- The nerve of a category, as a functor `Cat ⥤ SSet` -/ @[simps] -def nerveFunctor : Cat ⥤ SSet where +def nerveFunctor : Cat.{v, u} ⥤ SSet where obj C := nerve C map F := { app := fun _ => (F.mapComposableArrows _).obj } diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index 33e4d07da1bce..ff920be3bfa29 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -64,6 +64,13 @@ def spine (n : ℕ) (Δ : X _[n]) : X.Path n where simp only [← FunctorToTypes.map_comp_apply, ← op_comp] rw [SimplexCategory.δ_zero_mkOfSucc] +lemma spine_map_vertex {n : ℕ} (x : X _[n]) {m : ℕ} (φ : ([m] : SimplexCategory) ⟶ [n]) + (i : Fin (m + 1)) : + (spine X m (X.map φ.op x)).vertex i = (spine X n x).vertex (φ.toOrderHom i) := by + dsimp [spine] + rw [← FunctorToTypes.map_comp_apply] + rfl + lemma spine_map_subinterval {n : ℕ} (j l : ℕ) (hjl : j + l ≤ n) (Δ : X _[n]) : X.spine l (X.map (subinterval j l (by omega)).op Δ) = (X.spine n Δ).interval j l (by omega) := by diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index 7436832a79255..d45282b986c0d 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -19,6 +19,10 @@ Examples of `StrictSegal` simplicial sets are given by nerves of categories. TODO: Show that these are the only examples: that a `StrictSegal` simplicial set is isomorphic to the nerve of its homotopy category. + +`StrictSegal` simplicial sets have an important property of being 2-coskeletal which is proven +in `Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal`. + -/ universe v u diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 472d06dd79024..29001e428114c 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -1051,16 +1051,15 @@ theorem length_sigmaCompositionAux (a : Composition n) (b : Composition a.length 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 -set_option linter.deprecated false in theorem blocksFun_sigmaCompositionAux (a : Composition n) (b : Composition a.length) (i : Fin b.length) (j : Fin (blocksFun b i)) : blocksFun (sigmaCompositionAux a b ⟨i, (length_gather a b).symm ▸ i.2⟩) ⟨j, (length_sigmaCompositionAux a b i).symm ▸ j.2⟩ = - blocksFun a (embedding b i j) := - show get (get _ ⟨_, _⟩) ⟨_, _⟩ = a.blocks.get ⟨_, _⟩ by - rw [get_of_eq (get_splitWrtComposition _ _ _), get_drop', get_take']; rfl + blocksFun a (embedding b i j) := by + unfold sigmaCompositionAux + rw [blocksFun, get_eq_getElem, getElem_of_eq (getElem_splitWrtComposition _ _ _ _), + getElem_drop, getElem_take]; rfl -set_option linter.deprecated false in /-- Auxiliary lemma to prove that the composition of formal multilinear series is associative. Consider a composition `a` of `n` and a composition `b` of `a.length`. Grouping together some diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index fa0d86469c73b..622d8b9eb3b5d 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -357,7 +357,7 @@ def face {n} (I : Box (Fin (n + 1))) (i : Fin (n + 1)) : Box (Fin n) := theorem face_mk {n} (l u : Fin (n + 1) → ℝ) (h : ∀ i, l i < u i) (i : Fin (n + 1)) : face ⟨l, u, h⟩ i = ⟨l ∘ Fin.succAbove i, u ∘ Fin.succAbove i, fun _ ↦ h _⟩ := rfl -@[mono] +@[gcongr, mono] theorem face_mono {n} {I J : Box (Fin (n + 1))} (h : I ≤ J) (i : Fin (n + 1)) : face I i ≤ face J i := fun _ hx _ ↦ Ioc_subset_Ioc ((le_iff_bounds.1 h).1 _) ((le_iff_bounds.1 h).2 _) (hx _) diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index c1b1d7ea49571..f9ca9c7272245 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -408,19 +408,19 @@ nonrec theorem RCond.min {ι : Type*} {r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ (h₂ : l.RCond r₂) : l.RCond fun x => min (r₁ x) (r₂ x) := fun hR x => congr_arg₂ min (h₁ hR x) (h₂ hR x) -@[mono] +@[gcongr, mono] theorem toFilterDistortion_mono (I : Box ι) (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) : l₁.toFilterDistortion I c₁ ≤ l₂.toFilterDistortion I c₂ := iInf_mono fun _ => iInf_mono' fun hr => ⟨hr.mono h, principal_mono.2 fun _ => MemBaseSet.mono I h hc fun _ _ => le_rfl⟩ -@[mono] +@[gcongr, mono] theorem toFilter_mono (I : Box ι) {l₁ l₂ : IntegrationParams} (h : l₁ ≤ l₂) : l₁.toFilter I ≤ l₂.toFilter I := iSup_mono fun _ => toFilterDistortion_mono I h le_rfl -@[mono] +@[gcongr, mono] theorem toFilteriUnion_mono (I : Box ι) {l₁ l₂ : IntegrationParams} (h : l₁ ≤ l₂) (π₀ : Prepartition I) : l₁.toFilteriUnion I π₀ ≤ l₂.toFilteriUnion I π₀ := iSup_mono fun _ => inf_le_inf_right _ <| toFilterDistortion_mono _ h le_rfl diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean index 092ae83e8c61f..9c0e70ce75075 100644 --- a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -3,6 +3,8 @@ Copyright (c) 2024 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.Algebra.Module.ZLattice.Basic +import Mathlib.Analysis.BoxIntegral.Integrability import Mathlib.Analysis.BoxIntegral.Partition.Measure import Mathlib.Analysis.BoxIntegral.Partition.Tagged @@ -11,13 +13,12 @@ import Mathlib.Analysis.BoxIntegral.Partition.Tagged Fix `n` a positive integer. `BoxIntegral.unitPartition.box` are boxes in `ι → ℝ` obtained by dividing the unit box uniformly into boxes of side length `1 / n` and translating the boxes by -the lattice `ι → ℤ` so that they cover the whole space. For fixed `n`, there are indexed by vectors `ν : ι → ℤ`. Let `B` be a `BoxIntegral`. A `unitPartition.box` is admissible for `B` (more precisely its index is admissible) if it is contained in `B`. There are finitely many admissible `unitPartition.box` for `B` and thus we can form the corresponing tagged prepartition, see -`BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.Box` comes with its +`BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.box` comes with its tag situated at its "upper most" vertex). If `B` satisfies `hasIntegralVertices`, that is its vertices are in `ι → ℤ`, then the corresponding prepartition is actually a partition. @@ -35,6 +36,20 @@ coordinates in `ℤ` * `BoxIntegral.unitPartition.prepartition_isPartition`: For `B : BoxIntegral.Box`, if `B` has integral vertices, then the prepartition of `unitPartition.box` admissible for `B` is a partition of `B`. + +* `tendsto_tsum_div_pow_atTop_integral`: let `s` be a bounded, measurable set of `ι → ℝ` +whose frontier has zero volume and let `F` be a continuous function. Then the limit as `n → ∞` +of `∑ F x / n ^ card ι`, where the sum is over the points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the +integral of `F` over `s`. + +* `tendsto_card_div_pow_atTop_volume`: let `s` be a bounded, measurable set of `ι → ℝ` whose +frontier has zero volume. Then the limit as `n → ∞` of `card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` +tends to the volume of `s`. + +* `tendsto_card_div_pow_atTop_volume'`: a version of `tendsto_card_div_pow_atTop_volume` where we +assume in addition that `x • s ⊆ y • s` whenever `0 < x ≤ y`. Then we get the same limit +`card (s ∩ x⁻¹ • (ι → ℤ)) / x ^ card ι → volume s` but the limit is over a real variable `x`. + -/ noncomputable section @@ -186,7 +201,7 @@ theorem setFinite_index {s : Set (ι → ℝ)} (hs₁ : NullMeasurableSet s) (hs (by simp only [Set.iUnion_subset_iff, Set.inter_subset_right, implies_true]) hs₂ · rw [Set.mem_setOf, Set.inter_eq_self_of_subset_left hν, volume_box] -/-- For `B : BoxIntegral.Box`, the set of indices of `unitPartition.Box` that are subsets of `B`. +/-- For `B : BoxIntegral.Box`, the set of indices of `unitPartition.box` that are subsets of `B`. This is a finite set. These boxes cover `B` if it has integral vertices, see `unitPartition.prepartition_isPartition`. -/ def admissibleIndex (B : Box ι) : Finset (ι → ℤ) := by @@ -293,4 +308,180 @@ theorem prepartition_isPartition {B : Box ι} (hB : hasIntegralVertices B) : rw [TaggedPrepartition.mem_toPrepartition, mem_prepartition_iff] exact ⟨index n x, mem_admissibleIndex_of_mem_box n hB hx, rfl⟩ +open Submodule Pointwise BigOperators + +open scoped Pointwise + +variable (c : ℝ) (s : Set (ι → ℝ)) (F : (ι → ℝ) → ℝ) + +-- The image of `ι → ℤ` inside `ι → ℝ` +local notation "L" => span ℤ (Set.range (Pi.basisFun ℝ ι)) + +variable {n} in +theorem mem_smul_span_iff {v : ι → ℝ} : + v ∈ (n : ℝ)⁻¹ • L ↔ ∀ i, n * v i ∈ Set.range (algebraMap ℤ ℝ) := by + rw [ZSpan.smul _ (inv_ne_zero (NeZero.ne _)), Basis.mem_span_iff_repr_mem] + simp_rw [Basis.repr_isUnitSMul, Pi.basisFun_repr, Units.smul_def, Units.val_inv_eq_inv_val, + IsUnit.unit_spec, inv_inv, smul_eq_mul] + +theorem tag_mem_smul_span (ν : ι → ℤ) : + tag n ν ∈ (n : ℝ)⁻¹ • L := by + refine mem_smul_span_iff.mpr fun i ↦ ⟨ν i + 1, ?_⟩ + rw [tag_apply, div_eq_inv_mul, ← mul_assoc, mul_inv_cancel_of_invertible, one_mul, map_add, + map_one, eq_intCast] + +theorem tag_index_eq_self_of_mem_smul_span {x : ι → ℝ} (hx : x ∈ (n : ℝ)⁻¹ • L) : + tag n (index n x) = x := by + rw [mem_smul_span_iff] at hx + ext i + obtain ⟨a, ha⟩ : ∃ a : ℤ, a = n * x i := hx i + rwa [tag_apply, index_apply, Int.cast_sub, Int.cast_one, sub_add_cancel, ← ha, Int.ceil_intCast, + div_eq_iff (NeZero.ne _), mul_comm] + +theorem eq_of_mem_smul_span_of_index_eq_index {x y : ι → ℝ} (hx : x ∈ (n : ℝ)⁻¹ • L) + (hy : y ∈ (n : ℝ)⁻¹ • L) (h : index n x = index n y) : x = y := by + rw [← tag_index_eq_self_of_mem_smul_span n hx, ← tag_index_eq_self_of_mem_smul_span n hy, h] + +theorem integralSum_eq_tsum_div {B : Box ι} (hB : hasIntegralVertices B) (hs₀ : s ≤ B) : + integralSum (Set.indicator s F) (BoxAdditiveMap.toSMul (Measure.toBoxAdditive volume)) + (prepartition n B) = (∑' x : ↑(s ∩ (n : ℝ)⁻¹ • L), F x) / n ^ card ι := by + classical + unfold integralSum + have : Fintype ↑(s ∩ (n : ℝ)⁻¹ • L) := by + apply Set.Finite.fintype + rw [← coe_pointwise_smul, ZSpan.smul _ (inv_ne_zero (NeZero.ne _))] + exact ZSpan.setFinite_inter _ (B.isBounded.subset hs₀) + rw [tsum_fintype, Finset.sum_set_coe, Finset.sum_div, eq_comm] + simp_rw [Set.indicator_apply, apply_ite, BoxAdditiveMap.toSMul_apply, Measure.toBoxAdditive_apply, + smul_eq_mul, mul_zero, Finset.sum_ite, Finset.sum_const_zero, add_zero] + refine Finset.sum_bij (fun x _ ↦ box n (index n x)) (fun _ hx ↦ Finset.mem_filter.mpr ?_) + (fun _ hx _ hy h ↦ ?_) (fun I hI ↦ ?_) (fun _ hx ↦ ?_) + · rw [Set.mem_toFinset] at hx + refine ⟨mem_prepartition_boxes_iff.mpr + ⟨index n _, mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1), rfl⟩, ?_⟩ + simp_rw [prepartition_tag n (mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1)), + tag_index_eq_self_of_mem_smul_span n hx.2, hx.1] + · rw [Set.mem_toFinset] at hx hy + exact eq_of_mem_smul_span_of_index_eq_index n hx.2 hy.2 (box_injective n h) + · rw [Finset.mem_filter] at hI + refine ⟨(prepartition n B).tag I, Set.mem_toFinset.mpr ⟨hI.2, ?_⟩, box_index_tag_eq_self n hI.1⟩ + rw [← box_index_tag_eq_self n hI.1, prepartition_tag n + (mem_admissibleIndex_of_mem_box n hB (hs₀ hI.2))] + exact tag_mem_smul_span _ _ + · rw [Set.mem_toFinset] at hx + rw [volume_box, prepartition_tag n (mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1)), + tag_index_eq_self_of_mem_smul_span n hx.2, ENNReal.toReal_div, + ENNReal.one_toReal, ENNReal.toReal_pow, ENNReal.toReal_nat, mul_comm_div, one_mul] + +open Filter + +/-- Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume and let `F` +be a continuous function. Then the limit as `n → ∞` of `∑ F x / n ^ card ι`, where the sum is +over the points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the integral of `F` over `s`. -/ +theorem _root_.tendsto_tsum_div_pow_atTop_integral (hF : Continuous F) (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (∑' x : ↑(s ∩ (n : ℝ)⁻¹ • L), F x) / n ^ card ι) + atTop (nhds (∫ x in s, F x)) := by + obtain ⟨B, hB, hs₀⟩ := le_hasIntegralVertices_of_isBounded hs₁ + refine Metric.tendsto_atTop.mpr fun ε hε ↦ ?_ + have h₁ : ∃ C, ∀ x ∈ Box.Icc B, ‖Set.indicator s F x‖ ≤ C := by + obtain ⟨C₀, h₀⟩ := (Box.isCompact_Icc B).exists_bound_of_continuousOn hF.continuousOn + refine ⟨max 0 C₀, fun x hx ↦ ?_⟩ + rw [Set.indicator] + split_ifs with hs + · exact le_max_of_le_right (h₀ x hx) + · exact norm_zero.trans_le <|le_max_left 0 _ + have h₂ : ∀ᵐ x, ContinuousAt (s.indicator F) x := by + filter_upwards [compl_mem_ae_iff.mpr hs₃] with _ h + using (hF.continuousOn).continuousAt_indicator h + obtain ⟨r, hr₁, hr₂⟩ := (hasIntegral_iff.mp <| + AEContinuous.hasBoxIntegral (volume : Measure (ι → ℝ)) h₁ h₂ + IntegrationParams.Riemann) (ε / 2) (half_pos hε) + refine ⟨⌈(r 0 0 : ℝ)⁻¹⌉₊, fun n hn ↦ lt_of_le_of_lt ?_ (half_lt_self_iff.mpr hε)⟩ + have : NeZero n := + ⟨Nat.ne_zero_iff_zero_lt.mpr <| (Nat.ceil_pos.mpr (inv_pos.mpr (r 0 0).prop)).trans_le hn⟩ + rw [← integralSum_eq_tsum_div _ s F hB hs₀, ← Measure.restrict_restrict_of_subset hs₀, + ← integral_indicator hs₂] + refine hr₂ 0 _ ⟨?_, fun _ ↦ ?_, fun h ↦ ?_, fun h ↦ ?_⟩ (prepartition_isPartition _ hB) + · rw [show r 0 = fun _ ↦ r 0 0 from funext_iff.mpr (hr₁ 0 rfl)] + apply prepartition_isSubordinate n B + rw [one_div, inv_le_comm₀ (mod_cast (Nat.pos_of_neZero n)) (r 0 0).prop] + exact le_trans (Nat.le_ceil _) (Nat.cast_le.mpr hn) + · exact prepartition_isHenstock n B + · simp only [IntegrationParams.Riemann, Bool.false_eq_true] at h + · simp only [IntegrationParams.Riemann, Bool.false_eq_true] at h + +/-- Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume. Then the limit +as `n → ∞` of `card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` tends to the volume of `s`. This is a +special case of `tendsto_card_div_pow` with `F = 1`. -/ +theorem _root_.tendsto_card_div_pow_atTop_volume (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card ↑(s ∩ (n : ℝ)⁻¹ • L) : ℝ) / n ^ card ι) + atTop (nhds (volume s).toReal) := by + convert tendsto_tsum_div_pow_atTop_integral s (fun _ ↦ 1) continuous_const hs₁ hs₂ hs₃ + · rw [tsum_const, nsmul_eq_mul, mul_one, Nat.cast_inj] + · rw [setIntegral_const, smul_eq_mul, mul_one] + +private def tendsto_card_div_pow₁ {c : ℝ} (hc : c ≠ 0) : + ↑(s ∩ c⁻¹ • L) ≃ ↑(c • s ∩ L) := + Equiv.subtypeEquiv (Equiv.smulRight hc) (fun x ↦ by + simp_rw [Set.mem_inter_iff, Equiv.smulRight_apply, Set.smul_mem_smul_set_iff₀ hc, + ← Set.mem_inv_smul_set_iff₀ hc]) + +private theorem tendsto_card_div_pow₂ (hs₁ : IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) {x y : ℝ} (hx : 0 < x) (hy : x ≤ y) : + Nat.card ↑(s ∩ x⁻¹ • L) ≤ Nat.card ↑(s ∩ y⁻¹ • L) := by + rw [Nat.card_congr (tendsto_card_div_pow₁ s hx.ne'), + Nat.card_congr (tendsto_card_div_pow₁ s (hx.trans_le hy).ne')] + refine Nat.card_mono ?_ ?_ + · exact ZSpan.setFinite_inter _ (IsBounded.smul₀ hs₁ y) + · exact Set.inter_subset_inter_left _ <| hs₄ hx hy + +private theorem tendsto_card_div_pow₃ (hs₁ : IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + ∀ᶠ x : ℝ in atTop, (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι ≤ + (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι := by + filter_upwards [eventually_ge_atTop 1] with x hx + gcongr + exact tendsto_card_div_pow₂ s hs₁ hs₄ (Nat.cast_pos.mpr (Nat.floor_pos.mpr hx)) + (Nat.floor_le (zero_le_one.trans hx)) + +private theorem tendsto_card_div_pow₄ (hs₁ : IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + ∀ᶠ x : ℝ in atTop, (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι ≤ + (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι := by + filter_upwards [eventually_gt_atTop 0] with x hx + gcongr + exact tendsto_card_div_pow₂ s hs₁ hs₄ hx (Nat.le_ceil _) + +private theorem tendsto_card_div_pow₅ : + (fun x ↦ (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / ⌊x⌋₊ ^ card ι * (⌊x⌋₊ / x) ^ card ι) + =ᶠ[atTop] (fun x ↦ (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι) := by + filter_upwards [eventually_ge_atTop 1] with x hx + have : 0 < ⌊x⌋₊ := Nat.floor_pos.mpr hx + rw [div_pow, mul_div, div_mul_cancel₀ _ (by positivity)] + +private theorem tendsto_card_div_pow₆ : + (fun x ↦ (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / ⌈x⌉₊ ^ card ι * (⌈x⌉₊ / x) ^ card ι) + =ᶠ[atTop] (fun x ↦ (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι) := by + filter_upwards [eventually_ge_atTop 1] with x hx + have : 0 < ⌊x⌋₊ := Nat.floor_pos.mpr hx + rw [div_pow, mul_div, div_mul_cancel₀ _ (by positivity)] + +/-- A version of `tendsto_card_div_pow_atTop_volume` for a real variable. -/ +theorem _root_.tendsto_card_div_pow_atTop_volume' (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + Tendsto (fun x : ℝ ↦ (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι) + atTop (nhds (volume s).toReal) := by + rw [show (volume s).toReal = (volume s).toReal * 1 ^ card ι by ring] + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ + (tendsto_card_div_pow₃ s hs₁ hs₄) (tendsto_card_div_pow₄ s hs₁ hs₄) + · refine Tendsto.congr' (tendsto_card_div_pow₅ s) (Tendsto.mul ?_ (Tendsto.pow ?_ _)) + · exact Tendsto.comp (tendsto_card_div_pow_atTop_volume s hs₁ hs₂ hs₃) tendsto_nat_floor_atTop + · exact tendsto_nat_floor_div_atTop + · refine Tendsto.congr' (tendsto_card_div_pow₆ s) (Tendsto.mul ?_ (Tendsto.pow ?_ _)) + · exact Tendsto.comp (tendsto_card_div_pow_atTop_volume s hs₁ hs₂ hs₃) tendsto_nat_ceil_atTop + · exact tendsto_nat_ceil_div_atTop + end BoxIntegral.unitPartition diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 58258d4bbd896..6f8cc0631548c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -107,12 +107,7 @@ attribute [local instance 1001] open Set Fin Filter Function -/-- Smoothness exponent for analytic functions. -/ -scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) -/-- Smoothness exponent for infinitely differentiable functions. -/ -scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) - -open ContDiff +open scoped ContDiff universe u uE uF uG uX diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index abdac0002a6e2..a66bfc2242b36 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -98,10 +98,7 @@ In this file, we denote `⊤ : ℕ∞` with `∞`. noncomputable section -open scoped Classical -open ENat NNReal Topology Filter - -local notation "∞" => (⊤ : ℕ∞) +open ENat NNReal Topology Filter Set Fin Filter Function /- Porting note: These lines are not required in Mathlib4. @@ -109,7 +106,12 @@ attribute [local instance 1001] NormedAddCommGroup.toAddCommGroup NormedSpace.toModule' AddCommGroup.toAddCommMonoid -/ -open Set Fin Filter Function +/-- Smoothness exponent for analytic functions. -/ +scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) +/-- Smoothness exponent for infinitely differentiable functions. -/ +scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) + +open scoped ContDiff universe u uE uF @@ -129,7 +131,7 @@ Notice that `p` does not sum up to `f` on the diagonal (`FormalMultilinearSeries structure HasFTaylorSeriesUpToOn (n : WithTop ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) (s : Set E) : Prop where zero_eq : ∀ x ∈ s, (p x 0).curry0 = f x - protected fderivWithin : ∀ m : ℕ, (m : ℕ∞) < n → ∀ x ∈ s, + protected fderivWithin : ∀ m : ℕ, m < n → ∀ x ∈ s, HasFDerivWithinAt (p · m) (p x m.succ).curryLeft s x cont : ∀ m : ℕ, m ≤ n → ContinuousOn (p · m) s @@ -146,6 +148,16 @@ theorem HasFTaylorSeriesUpToOn.congr (h : HasFTaylorSeriesUpToOn n f p s) rw [h₁ x hx] exact h.zero_eq x hx +theorem HasFTaylorSeriesUpToOn.congr_series {q} (hp : HasFTaylorSeriesUpToOn n f p s) + (hpq : ∀ m : ℕ, m ≤ n → EqOn (p · m) (q · m) s) : + HasFTaylorSeriesUpToOn n f q s where + zero_eq x hx := by simp only [← (hpq 0 (zero_le n) hx), hp.zero_eq x hx] + fderivWithin m hm x hx := by + refine ((hp.fderivWithin m hm x hx).congr' (hpq m hm.le).symm hx).congr_fderiv ?_ + refine congrArg _ (hpq (m + 1) ?_ hx) + exact ENat.add_one_natCast_le_withTop_of_lt hm + cont m hm := (hp.cont m hm).congr (hpq m hm).symm + theorem HasFTaylorSeriesUpToOn.mono (h : HasFTaylorSeriesUpToOn n f p s) {t : Set E} (hst : t ⊆ s) : HasFTaylorSeriesUpToOn n f p t := ⟨fun x hx => h.zero_eq x (hst hx), fun m hm x hx => (h.fderivWithin m hm x (hst hx)).mono hst, @@ -549,10 +561,31 @@ theorem iteratedFDerivWithin_eventually_congr_set (h : s =ᶠ[𝓝 x] t) (n : iteratedFDerivWithin 𝕜 n f s =ᶠ[𝓝 x] iteratedFDerivWithin 𝕜 n f t := iteratedFDerivWithin_eventually_congr_set' x (h.filter_mono inf_le_left) n +/-- If two sets coincide in a punctured neighborhood of `x`, +then the corresponding iterated derivatives are equal. + +Note that we also allow to puncture the neighborhood of `x` at `y`. +If `y ≠ x`, then this is a no-op. -/ +theorem iteratedFDerivWithin_congr_set' {y} (h : s =ᶠ[𝓝[{y}ᶜ] x] t) (n : ℕ) : + iteratedFDerivWithin 𝕜 n f s x = iteratedFDerivWithin 𝕜 n f t x := + (iteratedFDerivWithin_eventually_congr_set' y h n).self_of_nhds + +@[simp] +theorem iteratedFDerivWithin_insert {n y} : + iteratedFDerivWithin 𝕜 n f (insert x s) y = iteratedFDerivWithin 𝕜 n f s y := + iteratedFDerivWithin_congr_set' (y := x) + (eventually_mem_nhdsWithin.mono <| by intros; simp_all).set_eq _ + theorem iteratedFDerivWithin_congr_set (h : s =ᶠ[𝓝 x] t) (n : ℕ) : iteratedFDerivWithin 𝕜 n f s x = iteratedFDerivWithin 𝕜 n f t x := (iteratedFDerivWithin_eventually_congr_set h n).self_of_nhds +@[simp] +theorem ftaylorSeriesWithin_insert : + ftaylorSeriesWithin 𝕜 f (insert x s) = ftaylorSeriesWithin 𝕜 f s := by + ext y n : 2 + apply iteratedFDerivWithin_insert + /-- The iterated differential within a set `s` at a point `x` is not modified if one intersects `s` with a neighborhood of `x` within `s`. -/ theorem iteratedFDerivWithin_inter' {n : ℕ} (hu : u ∈ 𝓝[s] x) : @@ -578,7 +611,7 @@ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn (hx : x ∈ s) : p x m = iteratedFDerivWithin 𝕜 m f s x := by induction' m with m IH generalizing x · rw [h.zero_eq' hx, iteratedFDerivWithin_zero_eq_comp]; rfl - · have A : (m : ℕ∞) < n := lt_of_lt_of_le (mod_cast lt_add_one m) hmn + · have A : m < n := lt_of_lt_of_le (mod_cast lt_add_one m) hmn have : HasFDerivWithinAt (fun y : E => iteratedFDerivWithin 𝕜 m f s y) (ContinuousMultilinearMap.curryLeft (p x (Nat.succ m))) s x := @@ -776,7 +809,7 @@ theorem iteratedFDerivWithin_univ {n : ℕ} : rw [iteratedFDeriv_succ_apply_left, iteratedFDerivWithin_succ_apply_left, IH, fderivWithin_univ] theorem HasFTaylorSeriesUpTo.eq_iteratedFDeriv - (h : HasFTaylorSeriesUpTo n f p) {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (x : E) : + (h : HasFTaylorSeriesUpTo n f p) {m : ℕ} (hmn : m ≤ n) (x : E) : p x m = iteratedFDeriv 𝕜 m f x := by rw [← iteratedFDerivWithin_univ] rw [← hasFTaylorSeriesUpToOn_univ_iff] at h diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 06469937e9249..b7cfacef1e22b 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -401,10 +401,10 @@ def extendMiddle (c : OrderedFinpartition n) (k : Fin c.length) : OrderedFinpart rcases eq_or_ne (c.index i) k with rfl | hi · have A : update c.partSize (c.index i) (c.partSize (c.index i) + 1) (c.index i) = c.partSize (c.index i) + 1 := by simp - exact ⟨c.index i, cast A.symm (succ (c.invEmbedding i)), by simp⟩ + exact ⟨c.index i, (succ (c.invEmbedding i)).cast A.symm , by simp⟩ · have A : update c.partSize k (c.partSize k + 1) (c.index i) = c.partSize (c.index i) := by simp [hi] - exact ⟨c.index i, cast A.symm (c.invEmbedding i), by simp [hi]⟩ + exact ⟨c.index i, (c.invEmbedding i).cast A.symm, by simp [hi]⟩ lemma index_extendMiddle_zero (c : OrderedFinpartition n) (i : Fin c.length) : (c.extendMiddle i).index 0 = i := by @@ -678,8 +678,8 @@ def extendEquiv (n : ℕ) : · simp only [↓reduceDIte, comp_apply] rcases eq_or_ne j 0 with rfl | hj · simpa using c.emb_zero - · let j' := Fin.pred (cast B.symm j) (by simpa using hj) - have : j = cast B (succ j') := by simp [j'] + · let j' := Fin.pred (j.cast B.symm) (by simpa using hj) + have : j = (succ j').cast B := by simp [j'] simp only [this, coe_cast, val_succ, cast_mk, cases_succ', comp_apply, succ_mk, Nat.succ_eq_add_one, succ_pred] rfl diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index ac81c76f071c7..4fe7a8fcd272e 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -268,6 +268,31 @@ protected theorem Convex.strictConvex {s : Set E} (hs : Convex 𝕜 s) end ContinuousConstSMul +section ContinuousSMul + +variable [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] + [TopologicalAddGroup E] [TopologicalSpace 𝕜] [OrderTopology 𝕜] [ContinuousSMul 𝕜 E] + +theorem Convex.closure_interior_eq_closure_of_nonempty_interior {s : Set E} (hs : Convex 𝕜 s) + (hs' : (interior s).Nonempty) : closure (interior s) = closure s := + subset_antisymm (closure_mono interior_subset) + fun _ h ↦ closure_mono (hs.openSegment_interior_closure_subset_interior hs'.choose_spec h) + (segment_subset_closure_openSegment (right_mem_segment ..)) + +theorem Convex.interior_closure_eq_interior_of_nonempty_interior {s : Set E} (hs : Convex 𝕜 s) + (hs' : (interior s).Nonempty) : interior (closure s) = interior s := by + refine subset_antisymm ?_ (interior_mono subset_closure) + intro y hy + rcases hs' with ⟨x, hx⟩ + have h := AffineMap.lineMap_apply_one (k := 𝕜) x y + obtain ⟨t, ht1, ht⟩ := AffineMap.lineMap_continuous.tendsto' _ _ h |>.eventually_mem + (mem_interior_iff_mem_nhds.1 hy) |>.exists_gt + apply hs.openSegment_interior_closure_subset_interior hx ht + nth_rw 1 [← AffineMap.lineMap_apply_zero (k := 𝕜) x y, ← image_openSegment] + exact ⟨1, Ioo_subset_openSegment ⟨zero_lt_one, ht1⟩, h⟩ + +end ContinuousSMul + section TopologicalSpace variable [OrderedSemiring 𝕜] [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index e18e2e98d3690..660294ee97883 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -336,8 +336,11 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] · exact hu.comp (by convert contDiff_update 1 x i) · exact h2u.comp_isClosedEmbedding (isClosedEmbedding_update x i) _ ≤ ∫⁻ xᵢ, (‖fderiv ℝ u (update x i xᵢ)‖₊ : ℝ≥0∞) := ?_ - gcongr with y; swap + gcongr · exact Measure.restrict_le_self + intro y + dsimp + gcongr -- bound the derivative which appears calc ‖deriv (u ∘ update x i) y‖₊ = ‖fderiv ℝ u (update x i y) (deriv (update x i) y)‖₊ := by rw [fderiv_comp_deriv _ (hu.differentiable le_rfl).differentiableAt diff --git a/Mathlib/Analysis/Normed/Ring/WithAbs.lean b/Mathlib/Analysis/Normed/Ring/WithAbs.lean index d616c2ef77867..4cbcad5fc3e00 100644 --- a/Mathlib/Analysis/Normed/Ring/WithAbs.lean +++ b/Mathlib/Analysis/Normed/Ring/WithAbs.lean @@ -22,7 +22,7 @@ being used to define Archimedean completions of a number field. to assign and infer instances on a semiring that depend on absolute values. - `WithAbs.equiv v` : the canonical (type) equivalence between `WithAbs v` and `R`. - `WithAbs.ringEquiv v` : The canonical ring equivalence between `WithAbs v` and `R`. - - `AbsoluteValue.completion` : the uniform space completion of a field `K` according to the + - `AbsoluteValue.Completion` : the uniform space completion of a field `K` according to the uniform structure defined by the specified real absolute value. -/ @@ -141,18 +141,20 @@ open WithAbs variable {K : Type*} [Field K] (v : AbsoluteValue K ℝ) /-- The completion of a field with respect to a real absolute value. -/ -abbrev completion := UniformSpace.Completion (WithAbs v) +abbrev Completion := UniformSpace.Completion (WithAbs v) + +@[deprecated (since := "2024-12-01")] alias completion := Completion namespace Completion -instance : Coe K v.completion := +instance : Coe K v.Completion := inferInstanceAs <| Coe (WithAbs v) (UniformSpace.Completion (WithAbs v)) variable {L : Type*} [NormedField L] [CompleteSpace L] {f : WithAbs v →+* L} {v} /-- If the absolute value of a normed field factors through an embedding into another normed field -`L`, then we can extend that embedding to an embedding on the completion `v.completion →+* L`. -/ -abbrev extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : v.completion →+* L := +`L`, then we can extend that embedding to an embedding on the completion `v.Completion →+* L`. -/ +abbrev extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : v.Completion →+* L := UniformSpace.Completion.extensionHom _ (WithAbs.isUniformInducing_of_comp h).uniformContinuous.continuous @@ -162,8 +164,8 @@ theorem extensionEmbedding_of_comp_coe (h : ∀ x, ‖f x‖ = v x) (x : K) : (WithAbs.isUniformInducing_of_comp h).uniformContinuous.continuous] /-- If the absolute value of a normed field factors through an embedding into another normed field, -then the extended embedding `v.completion →+* L` preserves distances. -/ -theorem extensionEmbedding_dist_eq_of_comp (h : ∀ x, ‖f x‖ = v x) (x y : v.completion) : +then the extended embedding `v.Completion →+* L` preserves distances. -/ +theorem extensionEmbedding_dist_eq_of_comp (h : ∀ x, ‖f x‖ = v x) (x y : v.Completion) : dist (extensionEmbedding_of_comp h x) (extensionEmbedding_of_comp h y) = dist x y := by refine UniformSpace.Completion.induction_on₂ x y ?_ (fun x y => ?_) @@ -173,13 +175,13 @@ theorem extensionEmbedding_dist_eq_of_comp (h : ∀ x, ‖f x‖ = v x) (x y : v exact UniformSpace.Completion.dist_eq x y ▸ (WithAbs.isometry_of_comp h).dist_eq _ _ /-- If the absolute value of a normed field factors through an embedding into another normed field, -then the extended embedding `v.completion →+* L` is an isometry. -/ +then the extended embedding `v.Completion →+* L` is an isometry. -/ theorem isometry_extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : Isometry (extensionEmbedding_of_comp h) := Isometry.of_dist_eq <| extensionEmbedding_dist_eq_of_comp h /-- If the absolute value of a normed field factors through an embedding into another normed field, -then the extended embedding `v.completion →+* L` is a closed embedding. -/ +then the extended embedding `v.Completion →+* L` is a closed embedding. -/ theorem isClosedEmbedding_extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x) : IsClosedEmbedding (extensionEmbedding_of_comp h) := (isometry_extensionEmbedding_of_comp h).isClosedEmbedding @@ -187,7 +189,7 @@ theorem isClosedEmbedding_extensionEmbedding_of_comp (h : ∀ x, ‖f x‖ = v x /-- If the absolute value of a normed field factors through an embedding into another normed field that is locally compact, then the completion of the first normed field is also locally compact. -/ theorem locallyCompactSpace [LocallyCompactSpace L] (h : ∀ x, ‖f x‖ = v x) : - LocallyCompactSpace (v.completion) := + LocallyCompactSpace (v.Completion) := (isClosedEmbedding_extensionEmbedding_of_comp h).locallyCompactSpace end AbsoluteValue.Completion diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index a0af1f9a0076e..0a7f64d0a15aa 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -698,14 +698,50 @@ theorem one_lt_rpow_iff (hx : 0 ≤ x) : 1 < x ^ y ↔ 1 < x ∧ 0 < y ∨ 0 < x · rcases _root_.em (y = 0) with (rfl | hy) <;> simp [*, lt_irrefl, (zero_lt_one' ℝ).not_lt] · simp [one_lt_rpow_iff_of_pos hx, hx] -theorem rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ z) (hyz : z ≤ y) : +/-- This is a more general but less convenient version of `rpow_le_rpow_of_exponent_ge`. +This version allows `x = 0`, so it explicitly forbids `x = y = 0`, `z ≠ 0`. -/ +theorem rpow_le_rpow_of_exponent_ge_of_imp (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hyz : z ≤ y) + (h : x = 0 → y = 0 → z = 0) : x ^ y ≤ x ^ z := by rcases eq_or_lt_of_le hx0 with (rfl | hx0') - · rcases eq_or_lt_of_le hz with (rfl | hz') - · exact (rpow_zero 0).symm ▸ rpow_le_one hx0 hx1 hyz - rw [zero_rpow, zero_rpow] <;> linarith + · rcases eq_or_ne y 0 with rfl | hy0 + · rw [h rfl rfl] + · rw [zero_rpow hy0] + apply zero_rpow_nonneg · exact rpow_le_rpow_of_exponent_ge hx0' hx1 hyz +/-- This version of `rpow_le_rpow_of_exponent_ge` allows `x = 0` but requires `0 ≤ z`. +See also `rpow_le_rpow_of_exponent_ge_of_imp` for the most general version. -/ +theorem rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ z) (hyz : z ≤ y) : + x ^ y ≤ x ^ z := + rpow_le_rpow_of_exponent_ge_of_imp hx0 hx1 hyz fun _ hy ↦ le_antisymm (hyz.trans_eq hy) hz + +theorem self_le_rpow_of_le_one (h₁ : 0 ≤ x) (h₂ : x ≤ 1) (h₃ : y ≤ 1) : x ≤ x ^ y := by + simpa only [rpow_one] + using rpow_le_rpow_of_exponent_ge_of_imp h₁ h₂ h₃ fun _ ↦ (absurd · one_ne_zero) + +theorem self_le_rpow_of_one_le (h₁ : 1 ≤ x) (h₂ : 1 ≤ y) : x ≤ x ^ y := by + simpa only [rpow_one] using rpow_le_rpow_of_exponent_le h₁ h₂ + +theorem rpow_le_self_of_le_one (h₁ : 0 ≤ x) (h₂ : x ≤ 1) (h₃ : 1 ≤ y) : x ^ y ≤ x := by + simpa only [rpow_one] + using rpow_le_rpow_of_exponent_ge_of_imp h₁ h₂ h₃ fun _ ↦ (absurd · (one_pos.trans_le h₃).ne') + +theorem rpow_le_self_of_one_le (h₁ : 1 ≤ x) (h₂ : y ≤ 1) : x ^ y ≤ x := by + simpa only [rpow_one] using rpow_le_rpow_of_exponent_le h₁ h₂ + +theorem self_lt_rpow_of_lt_one (h₁ : 0 < x) (h₂ : x < 1) (h₃ : y < 1) : x < x ^ y := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_gt h₁ h₂ h₃ + +theorem self_lt_rpow_of_one_lt (h₁ : 1 < x) (h₂ : 1 < y) : x < x ^ y := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_lt h₁ h₂ + +theorem rpow_lt_self_of_lt_one (h₁ : 0 < x) (h₂ : x < 1) (h₃ : 1 < y) : x ^ y < x := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_gt h₁ h₂ h₃ + +theorem rpow_lt_self_of_one_lt (h₁ : 1 < x) (h₂ : y < 1) : x ^ y < x := by + simpa only [rpow_one] using rpow_lt_rpow_of_exponent_lt h₁ h₂ + theorem rpow_left_injOn {x : ℝ} (hx : x ≠ 0) : InjOn (fun y : ℝ => y ^ x) { y : ℝ | 0 ≤ y } := by rintro y hy z hz (hyz : y ^ x = z ^ x) rw [← rpow_one y, ← rpow_one z, ← mul_inv_cancel₀ hx, rpow_mul hy, rpow_mul hz, hyz] diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean index 6919f9dcbb728..7e4c7f3d67170 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean @@ -145,6 +145,13 @@ theorem arctan_zero : arctan 0 = 0 := by simp [arctan_eq_arcsin] @[mono] theorem arctan_strictMono : StrictMono arctan := tanOrderIso.symm.strictMono +@[gcongr] +lemma arctan_lt_arctan {x y : ℝ} (hxy : x < y) : arctan x < arctan y := arctan_strictMono hxy + +@[gcongr] +lemma arctan_le_arctan {x y : ℝ} (hxy : x ≤ y) : arctan x ≤ arctan y := + arctan_strictMono.monotone hxy + theorem arctan_injective : arctan.Injective := arctan_strictMono.injective @[simp] diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean index 9e2609f0ac9b5..23c49de796496 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean @@ -94,6 +94,18 @@ theorem U_complex_cos (n : ℤ) : (U ℂ n).eval (cos θ) * sin θ = sin ((n + 1 push_cast ring_nf +/-- The `n`-th rescaled Chebyshev polynomial of the first kind (Vieta–Lucas polynomial) evaluates on +`2 * cos θ` to the value `2 * cos (n * θ)`. -/ +@[simp] +theorem C_two_mul_complex_cos (n : ℤ) : (C ℂ n).eval (2 * cos θ) = 2 * cos (n * θ) := by + simp [C_eq_two_mul_T_comp_half_mul_X] + +/-- The `n`-th rescaled Chebyshev polynomial of the second kind (Vieta–Fibonacci polynomial) +evaluates on `2 * cos θ` to the value `sin ((n + 1) * θ) / sin θ`. -/ +@[simp] +theorem S_two_mul_complex_cos (n : ℤ) : (S ℂ n).eval (2 * cos θ) * sin θ = sin ((n + 1) * θ) := by + simp [S_eq_U_comp_half_mul_X] + end Complex /-! ### Real versions -/ @@ -115,6 +127,18 @@ value `sin ((n + 1) * θ) / sin θ`. -/ theorem U_real_cos : (U ℝ n).eval (cos θ) * sin θ = sin ((n + 1) * θ) := mod_cast U_complex_cos θ n +/-- The `n`-th rescaled Chebyshev polynomial of the first kind (Vieta–Lucas polynomial) evaluates on +`2 * cos θ` to the value `2 * cos (n * θ)`. -/ +@[simp] +theorem C_two_mul_real_cos : (C ℝ n).eval (2 * cos θ) = 2 * cos (n * θ) := by + simp [C_eq_two_mul_T_comp_half_mul_X] + +/-- The `n`-th rescaled Chebyshev polynomial of the second kind (Vieta–Fibonacci polynomial) +evaluates on `2 * cos θ` to the value `sin ((n + 1) * θ) / sin θ`. -/ +@[simp] +theorem S_two_mul_real_cos : (S ℝ n).eval (2 * cos θ) * sin θ = sin ((n + 1) * θ) := by + simp [S_eq_U_comp_half_mul_X] + end Real end Polynomial.Chebyshev diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean index e11ac111c917c..9e81a28ffd8e2 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean @@ -3,11 +3,12 @@ Copyright (c) 2023 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Isaac Hernando, Coleton Kotch, Adam Topaz -/ - +import Mathlib.Algebra.Homology.ShortComplex.ExactFunctor +import Mathlib.CategoryTheory.Abelian.FunctorCategory import Mathlib.CategoryTheory.Limits.Constructions.Filtered -import Mathlib.CategoryTheory.Limits.Shapes.Biproducts import Mathlib.CategoryTheory.Limits.Preserves.FunctorCategory - +import Mathlib.CategoryTheory.Limits.Shapes.Countable +import Mathlib.Logic.Equiv.List /-! # Grothendieck Axioms @@ -17,9 +18,10 @@ basic facts about them. ## Definitions -- `AB4` -- a category satisfies `AB4` provided that coproducts are exact. -- `AB5` -- a category satisfies `AB5` provided that filtered colimits are exact. -- The duals of the above definitions, called `AB4Star` and `AB5Star`. +- `HasExactColimitsOfShape J` -- colimits of shape `J` are exact. +- The dual of the above definitions, called `HasExactLimitsOfShape`. +- `AB4` -- coproducts are exact (this is formulated in terms of `HasExactColimitsOfShape`). +- `AB5` -- filtered colimits are exact (this is formulated in terms of `HasExactColimitsOfShape`). ## Theorems @@ -50,63 +52,226 @@ namespace CategoryTheory open Limits -universe v v' u u' w +attribute [instance] comp_preservesFiniteLimits comp_preservesFiniteColimits + +universe w w' w₂ w₂' v v' u u' variable (C : Type u) [Category.{v} C] -attribute [local instance] hasCoproducts_of_finite_and_filtered +/-- +A category `C` is said to have exact colimits of shape `J` provided that colimits of shape `J` +exist and are exact (in the sense that they preserve finite limits). +-/ +class HasExactColimitsOfShape (J : Type u') [Category.{v'} J] (C : Type u) [Category.{v} C] + [HasColimitsOfShape J C] where + /-- Exactness of `J`-shaped colimits stated as `colim : (J ⥤ C) ⥤ C` preserving finite limits. -/ + preservesFiniteLimits : PreservesFiniteLimits (colim (J := J) (C := C)) + +/-- +A category `C` is said to have exact limits of shape `J` provided that limits of shape `J` +exist and are exact (in the sense that they preserve finite colimits). +-/ +class HasExactLimitsOfShape (J : Type u') [Category.{v'} J] (C : Type u) [Category.{v} C] + [HasLimitsOfShape J C] where + /-- Exactness of `J`-shaped limits stated as `lim : (J ⥤ C) ⥤ C` preserving finite colimits. -/ + preservesFiniteColimits : PreservesFiniteColimits (lim (J := J) (C := C)) + +attribute [instance] HasExactColimitsOfShape.preservesFiniteLimits + HasExactLimitsOfShape.preservesFiniteColimits + +/-- +Transport a `HasExactColimitsOfShape` along an equivalence of the shape. + +Note: When `C` has finite limits, this lemma holds with the equivalence replaced by a final +functor, see `hasExactColimitsOfShape_of_final` below. +-/ +lemma hasExactColimitsOfShape_of_equiv {J J' : Type*} [Category J] [Category J'] (e : J ≌ J') + [HasColimitsOfShape J C] [HasExactColimitsOfShape J C] : + haveI : HasColimitsOfShape J' C := hasColimitsOfShape_of_equivalence e + HasExactColimitsOfShape J' C := + haveI : HasColimitsOfShape J' C := hasColimitsOfShape_of_equivalence e + ⟨preservesFiniteLimits_of_natIso (Functor.Final.colimIso e.functor)⟩ + +/-- +Transport a `HasExactLimitsOfShape` along an equivalence of the shape. + +Note: When `C` has finite colimits, this lemma holds with the equivalence replaced by a initial +functor, see `hasExactLimitsOfShape_of_initial` below. +-/ +lemma hasExactLimitsOfShape_of_equiv {J J' : Type*} [Category J] [Category J'] (e : J ≌ J') + [HasLimitsOfShape J C] [HasExactLimitsOfShape J C] : + haveI : HasLimitsOfShape J' C := hasLimitsOfShape_of_equivalence e + HasExactLimitsOfShape J' C := + haveI : HasLimitsOfShape J' C := hasLimitsOfShape_of_equivalence e + ⟨preservesFiniteColimits_of_natIso (Functor.Initial.limIso e.functor)⟩ + +/-- +A category `C` which has coproducts is said to have `AB4` of size `w` provided that +coproducts of size `w` are exact. +-/ +@[pp_with_univ] +class AB4OfSize [HasCoproducts.{w} C] where + ofShape (α : Type w) : HasExactColimitsOfShape (Discrete α) C + +attribute [instance] AB4OfSize.ofShape /-- A category `C` which has coproducts is said to have `AB4` provided that coproducts are exact. -/ -class AB4 [HasCoproducts C] where - /-- Exactness of coproducts stated as `colim : (Discrete α ⥤ C) ⥤ C` preserving limits. -/ - preservesFiniteLimits (α : Type v) : - PreservesFiniteLimits (colim (J := Discrete α) (C := C)) +abbrev AB4 [HasCoproducts C] := AB4OfSize.{v} C + +lemma AB4OfSize_shrink [HasCoproducts.{max w w'} C] [AB4OfSize.{max w w'} C] : + haveI : HasCoproducts.{w} C := hasCoproducts_shrink.{w, w'} + AB4OfSize.{w} C := + haveI := hasCoproducts_shrink.{w, w'} (C := C) + ⟨fun J ↦ hasExactColimitsOfShape_of_equiv C + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _)⟩ + +instance (priority := 100) [HasCoproducts.{w} C] [AB4OfSize.{w} C] : + haveI : HasCoproducts.{0} C := hasCoproducts_shrink + AB4OfSize.{0} C := AB4OfSize_shrink C + +/-- A category `C` which has products is said to have `AB4Star` (in literature `AB4*`) +provided that products are exact. -/ +@[pp_with_univ] +class AB4StarOfSize [HasProducts.{w} C] where + ofShape (α : Type w) : HasExactLimitsOfShape (Discrete α) C -attribute [instance] AB4.preservesFiniteLimits +attribute [instance] AB4StarOfSize.ofShape /-- A category `C` which has products is said to have `AB4Star` (in literature `AB4*`) provided that products are exact. -/ -class AB4Star [HasProducts C] where - /-- Exactness of products stated as `lim : (Discrete α ⥤ C) ⥤ C` preserving colimits. -/ - preservesFiniteColimits (α : Type v) : - PreservesFiniteColimits (lim (J := Discrete α) (C := C)) +abbrev AB4Star [HasProducts C] := AB4StarOfSize.{v} C + +lemma AB4StarOfSize_shrink [HasProducts.{max w w'} C] [AB4StarOfSize.{max w w'} C] : + haveI : HasProducts.{w} C := hasProducts_shrink.{w, w'} + AB4StarOfSize.{w} C := + haveI := hasProducts_shrink.{w, w'} (C := C) + ⟨fun J ↦ hasExactLimitsOfShape_of_equiv C + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _)⟩ + +instance (priority := 100) [HasProducts.{w} C] [AB4StarOfSize.{w} C] : + haveI : HasProducts.{0} C := hasProducts_shrink + AB4StarOfSize.{0} C := AB4StarOfSize_shrink C + +/-- +A category `C` which has countable coproducts is said to have countable `AB4` provided that +countable coproducts are exact. +-/ +class CountableAB4 [HasCountableCoproducts C] where + ofShape (α : Type) [Countable α] : HasExactColimitsOfShape (Discrete α) C + +instance (priority := 100) [HasCoproducts.{0} C] [AB4OfSize.{0} C] : CountableAB4 C := + ⟨inferInstance⟩ + +/-- +A category `C` which has countable coproducts is said to have countable `AB4Star` provided that +countable products are exact. +-/ +class CountableAB4Star [HasCountableProducts C] where + ofShape (α : Type) [Countable α] : HasExactLimitsOfShape (Discrete α) C + +instance (priority := 100) [HasProducts.{0} C] [AB4StarOfSize.{0} C] : CountableAB4Star C := + ⟨inferInstance⟩ -attribute [instance] AB4Star.preservesFiniteColimits +attribute [instance] CountableAB4.ofShape CountableAB4Star.ofShape + +/-- +A category `C` which has filtered colimits of a given size is said to have `AB5` of that size +provided that these filtered colimits are exact. + +`AB5OfSize.{w, w'} C` means that `C` has exact colimits of shape `J : Type w'` with +`Category.{w} J` such that `J` is filtered. +-/ +@[pp_with_univ] +class AB5OfSize [HasFilteredColimitsOfSize.{w, w'} C] where + ofShape (J : Type w') [Category.{w} J] [IsFiltered J] : HasExactColimitsOfShape J C + +attribute [instance] AB5OfSize.ofShape /-- A category `C` which has filtered colimits is said to have `AB5` provided that filtered colimits are exact. -/ -class AB5 [HasFilteredColimits C] where - /-- Exactness of filtered colimits stated as `colim : (J ⥤ C) ⥤ C` on filtered `J` - preserving limits. -/ - preservesFiniteLimits (J : Type v) [SmallCategory J] [IsFiltered J] : - PreservesFiniteLimits (colim (J := J) (C := C)) - -attribute [instance] AB5.preservesFiniteLimits +abbrev AB5 [HasFilteredColimits C] := AB5OfSize.{v, v} C + +lemma AB5OfSize_of_univLE [HasFilteredColimitsOfSize.{w₂, w₂'} C] [UnivLE.{w, w₂}] + [UnivLE.{w', w₂'}] [AB5OfSize.{w₂, w₂'} C] : + haveI : HasFilteredColimitsOfSize.{w, w'} C := hasFilteredColimitsOfSize_of_univLE.{w} + AB5OfSize.{w, w'} C := by + haveI : HasFilteredColimitsOfSize.{w, w'} C := hasFilteredColimitsOfSize_of_univLE.{w} + constructor + intro J _ _ + haveI := IsFiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)) + exact hasExactColimitsOfShape_of_equiv _ ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)).symm + +lemma AB5OfSize_shrink [HasFilteredColimitsOfSize.{max w w₂, max w' w₂'} C] + [AB5OfSize.{max w w₂, max w' w₂'} C] : + haveI : HasFilteredColimitsOfSize.{w, w'} C := hasFilteredColimitsOfSize_shrink + AB5OfSize.{w, w'} C := + AB5OfSize_of_univLE C /-- A category `C` which has cofiltered limits is said to have `AB5Star` (in literature `AB5*`) provided that cofiltered limits are exact. -/ -class AB5Star [HasCofilteredLimits C] where - /-- Exactness of cofiltered limits stated as `lim : (J ⥤ C) ⥤ C` on cofiltered `J` - preserving colimits. -/ - preservesFiniteColimits (J : Type v) [SmallCategory J] [IsCofiltered J] : - PreservesFiniteColimits (lim (J := J) (C := C)) +@[pp_with_univ] +class AB5StarOfSize [HasCofilteredLimitsOfSize.{w, w'} C] where + ofShape (J : Type w') [Category.{w} J] [IsCofiltered J] : HasExactLimitsOfShape J C -attribute [instance] AB5Star.preservesFiniteColimits +attribute [instance] AB5StarOfSize.ofShape -noncomputable section +/-- +A category `C` which has cofiltered limits is said to have `AB5Star` (in literature `AB5*`) +provided that cofiltered limits are exact. +-/ +abbrev AB5Star [HasCofilteredLimits C] := AB5StarOfSize.{v, v} C + +lemma AB5StarOfSize_of_univLE [HasCofilteredLimitsOfSize.{w₂, w₂'} C] [UnivLE.{w, w₂}] + [UnivLE.{w', w₂'}] [AB5StarOfSize.{w₂, w₂'} C] : + haveI : HasCofilteredLimitsOfSize.{w, w'} C := hasCofilteredLimitsOfSize_of_univLE.{w} + AB5StarOfSize.{w, w'} C := by + haveI : HasCofilteredLimitsOfSize.{w, w'} C := hasCofilteredLimitsOfSize_of_univLE.{w} + constructor + intro J _ _ + haveI := IsCofiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)) + exact hasExactLimitsOfShape_of_equiv _ ((ShrinkHoms.equivalence.{w₂} J).trans <| + Shrink.equivalence.{w₂'} (ShrinkHoms.{w'} J)).symm + +lemma AB5StarOfSize_shrink [HasCofilteredLimitsOfSize.{max w w₂, max w' w₂'} C] + [AB5StarOfSize.{max w w₂, max w' w₂'} C] : + haveI : HasCofilteredLimitsOfSize.{w, w'} C := hasCofilteredLimitsOfSize_shrink + AB5StarOfSize.{w, w'} C := + AB5StarOfSize_of_univLE C + +/-- `HasExactColimitsOfShape` can be "pushed forward" along final functors -/ +lemma hasExactColimitsOfShape_of_final [HasFiniteLimits C] {J J' : Type*} [Category J] [Category J'] + (F : J ⥤ J') [F.Final] [HasColimitsOfShape J' C] [HasColimitsOfShape J C] + [HasExactColimitsOfShape J C] : HasExactColimitsOfShape J' C where + preservesFiniteLimits := + letI : PreservesFiniteLimits ((whiskeringLeft J J' C).obj F) := ⟨fun _ ↦ inferInstance⟩ + letI := comp_preservesFiniteLimits ((whiskeringLeft J J' C).obj F) colim + preservesFiniteLimits_of_natIso (Functor.Final.colimIso F) + +/-- `HasExactLimitsOfShape` can be "pushed forward" along initial functors -/ +lemma hasExactLimitsOfShape_of_initial [HasFiniteColimits C] {J J' : Type*} [Category J] + [Category J'] (F : J ⥤ J') [F.Initial] [HasLimitsOfShape J' C] [HasLimitsOfShape J C] + [HasExactLimitsOfShape J C] : HasExactLimitsOfShape J' C where + preservesFiniteColimits := + letI : PreservesFiniteColimits ((whiskeringLeft J J' C).obj F) := ⟨fun _ ↦ inferInstance⟩ + letI := comp_preservesFiniteColimits ((whiskeringLeft J J' C).obj F) lim + preservesFiniteColimits_of_natIso (Functor.Initial.limIso F) + +section AB4OfAB5 + +variable {α : Type w} [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] open CoproductsFromFiniteFiltered -variable {α : Type w} -variable [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] - instance preservesFiniteLimits_liftToFinset : PreservesFiniteLimits (liftToFinset C α) := preservesFiniteLimits_of_evaluation _ fun I => letI : PreservesFiniteLimits (colim (J := Discrete I) (C := C)) := @@ -117,15 +282,192 @@ instance preservesFiniteLimits_liftToFinset : PreservesFiniteLimits (liftToFinse letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj (Discrete.functor (·.val)) ⋙ colim) := comp_preservesFiniteLimits _ _ - preservesFiniteLimits_of_natIso (liftToFinsetEvaluationIso I).symm + preservesFiniteLimits_of_natIso (liftToFinsetEvaluationIso I).symm -/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ -lemma AB4.of_AB5 [HasFilteredColimits C] [AB5 C] : AB4 C where - preservesFiniteLimits J := +variable (J : Type*) + +/-- +`HasExactColimitsOfShape (Finset (Discrete J)) C` implies `HasExactColimitsOfShape (Discrete J) C` +-/ +lemma hasExactColimitsOfShape_discrete_of_hasExactColimitsOfShape_finset_discrete + [HasColimitsOfShape (Discrete J) C] [HasColimitsOfShape (Finset (Discrete J)) C] + [HasExactColimitsOfShape (Finset (Discrete J)) C] : HasExactColimitsOfShape (Discrete J) C where + preservesFiniteLimits := letI : PreservesFiniteLimits (liftToFinset C J ⋙ colim) := comp_preservesFiniteLimits _ _ preservesFiniteLimits_of_natIso (liftToFinsetColimIso) +attribute [local instance] hasCoproducts_of_finite_and_filtered in +/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ +lemma AB4.of_AB5 [HasFilteredColimitsOfSize.{w, w} C] + [AB5OfSize.{w, w} C] : AB4OfSize.{w} C where + ofShape _ := hasExactColimitsOfShape_discrete_of_hasExactColimitsOfShape_finset_discrete _ _ + +/-- +A category with finite biproducts and finite limits has countable AB4 if sequential colimits are +exact. +-/ +lemma CountableAB4.of_countableAB5 [HasColimitsOfShape ℕ C] [HasExactColimitsOfShape ℕ C] + [HasCountableCoproducts C] : CountableAB4 C where + ofShape J := + have : HasColimitsOfShape (Finset (Discrete J)) C := + Functor.Final.hasColimitsOfShape_of_final + (IsFiltered.sequentialFunctor (Finset (Discrete J))) + have := hasExactColimitsOfShape_of_final C (IsFiltered.sequentialFunctor (Finset (Discrete J))) + hasExactColimitsOfShape_discrete_of_hasExactColimitsOfShape_finset_discrete _ _ + +end AB4OfAB5 + +section AB4StarOfAB5Star + +variable {α : Type w} [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteColimits C] + +open ProductsFromFiniteCofiltered + +instance preservesFiniteColimits_liftToFinset : PreservesFiniteColimits (liftToFinset C α) := + preservesFiniteColimits_of_evaluation _ fun ⟨I⟩ => + letI : PreservesFiniteColimits (lim (J := Discrete I) (C := C)) := + preservesFiniteColimits_of_natIso HasBiproductsOfShape.colimIsoLim + letI : PreservesFiniteColimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj + (Discrete.functor fun x ↦ ↑x)) := ⟨fun _ _ _ => inferInstance⟩ + letI : PreservesFiniteColimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj + (Discrete.functor (·.val)) ⋙ lim) := + comp_preservesFiniteColimits _ _ + preservesFiniteColimits_of_natIso (liftToFinsetEvaluationIso _ _ I).symm + +variable (J : Type*) + +/-- +`HasExactLimitsOfShape (Finset (Discrete J))ᵒᵖ C` implies `HasExactLimitsOfShape (Discrete J) C` +-/ +lemma hasExactLimitsOfShape_discrete_of_hasExactLimitsOfShape_finset_discrete_op + [HasLimitsOfShape (Discrete J) C] [HasLimitsOfShape (Finset (Discrete J))ᵒᵖ C] + [HasExactLimitsOfShape (Finset (Discrete J))ᵒᵖ C] : + HasExactLimitsOfShape (Discrete J) C where + preservesFiniteColimits := + letI : PreservesFiniteColimits (ProductsFromFiniteCofiltered.liftToFinset C J ⋙ lim) := + comp_preservesFiniteColimits _ _ + preservesFiniteColimits_of_natIso (ProductsFromFiniteCofiltered.liftToFinsetLimIso _ _) + +attribute [local instance] hasProducts_of_finite_and_cofiltered in +/-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ +lemma AB4Star.of_AB5Star [HasCofilteredLimitsOfSize.{w, w} C] [AB5StarOfSize.{w, w} C] : + AB4StarOfSize.{w} C where + ofShape _ := hasExactLimitsOfShape_discrete_of_hasExactLimitsOfShape_finset_discrete_op _ _ + +/-- +A category with finite biproducts and finite limits has countable AB4* if sequential limits are +exact. +-/ +lemma CountableAB4Star.of_countableAB5Star [HasLimitsOfShape ℕᵒᵖ C] [HasExactLimitsOfShape ℕᵒᵖ C] + [HasCountableProducts C] : CountableAB4Star C where + ofShape J := + have : HasLimitsOfShape (Finset (Discrete J))ᵒᵖ C := + Functor.Initial.hasLimitsOfShape_of_initial + (IsFiltered.sequentialFunctor (Finset (Discrete J))).op + have := hasExactLimitsOfShape_of_initial C + (IsFiltered.sequentialFunctor (Finset (Discrete J))).op + hasExactLimitsOfShape_discrete_of_hasExactLimitsOfShape_finset_discrete_op _ _ + +end AB4StarOfAB5Star + +/-- +Checking exactness of colimits of shape `Discrete ℕ` and `Discrete J` for finite `J` is enough for +countable AB4. +-/ +lemma CountableAB4.of_hasExactColimitsOfShape_nat_and_finite [HasCountableCoproducts C] + [HasFiniteLimits C] [∀ (J : Type) [Finite J], HasExactColimitsOfShape (Discrete J) C] + [HasExactColimitsOfShape (Discrete ℕ) C] : + CountableAB4 C where + ofShape J := by + by_cases h : Finite J + · infer_instance + · have : Infinite J := ⟨h⟩ + let _ := Encodable.ofCountable J + let _ := Denumerable.ofEncodableOfInfinite J + exact hasExactColimitsOfShape_of_final C (Discrete.equivalence (Denumerable.eqv J)).inverse + +/-- +Checking exactness of limits of shape `Discrete ℕ` and `Discrete J` for finite `J` is enough for +countable AB4*. +-/ +lemma CountableAB4Star.of_hasExactLimitsOfShape_nat_and_finite [HasCountableProducts C] + [HasFiniteColimits C] [∀ (J : Type) [Finite J], HasExactLimitsOfShape (Discrete J) C] + [HasExactLimitsOfShape (Discrete ℕ) C] : + CountableAB4Star C where + ofShape J := by + by_cases h : Finite J + · infer_instance + · have : Infinite J := ⟨h⟩ + let _ := Encodable.ofCountable J + let _ := Denumerable.ofEncodableOfInfinite J + exact hasExactLimitsOfShape_of_initial C (Discrete.equivalence (Denumerable.eqv J)).inverse + +section EpiMono + +open Functor + +section + +variable [HasZeroMorphisms C] [HasFiniteBiproducts C] + +noncomputable instance hasExactColimitsOfShape_discrete_finite (J : Type*) [Finite J] : + HasExactColimitsOfShape (Discrete J) C where + preservesFiniteLimits := preservesFiniteLimits_of_natIso HasBiproductsOfShape.colimIsoLim.symm + +noncomputable instance hasExactLimitsOfShape_discrete_finite {J : Type*} [Finite J] : + HasExactLimitsOfShape (Discrete J) C where + preservesFiniteColimits := preservesFiniteColimits_of_natIso HasBiproductsOfShape.colimIsoLim + +/-- +Checking AB of shape `Discrete ℕ` is enough for countable AB4, provided that the category has +finite biproducts and finite limits. +-/ +lemma CountableAB4.of_hasExactColimitsOfShape_nat [HasFiniteLimits C] [HasCountableCoproducts C] + [HasExactColimitsOfShape (Discrete ℕ) C] : CountableAB4 C := by + apply (config := { allowSynthFailures := true }) + CountableAB4.of_hasExactColimitsOfShape_nat_and_finite + exact fun _ ↦ inferInstance + +/-- +Checking AB* of shape `Discrete ℕ` is enough for countable AB4*, provided that the category has +finite biproducts and finite colimits. +-/ +lemma CountableAB4Star.of_hasExactLimitsOfShape_nat [HasFiniteColimits C] + [HasCountableProducts C] [HasExactLimitsOfShape (Discrete ℕ) C] : CountableAB4Star C := by + apply (config := { allowSynthFailures := true }) + CountableAB4Star.of_hasExactLimitsOfShape_nat_and_finite + exact fun _ ↦ inferInstance + end +variable [Abelian C] (J : Type u') [Category.{v'} J] + +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryCoproducts + preservesBinaryBiproducts_of_preservesBinaryProducts + +/-- +If `colim` of shape `J` into an abelian category `C` preserves monomorphisms, then `C` has AB of +shape `J`. +-/ +lemma hasExactColimitsOfShape_of_preservesMono [HasColimitsOfShape J C] + [PreservesMonomorphisms (colim (J := J) (C := C))] : HasExactColimitsOfShape J C where + preservesFiniteLimits := by + apply (config := { allowSynthFailures := true }) preservesFiniteLimits_of_preservesHomology + · exact preservesHomology_of_preservesMonos_and_cokernels _ + · exact additive_of_preservesBinaryBiproducts _ + +/-- +If `lim` of shape `J` into an abelian category `C` preserves epimorphisms, then `C` has AB* of +shape `J`. +-/ +lemma hasExactLimitsOfShape_of_preservesEpi [HasLimitsOfShape J C] + [PreservesEpimorphisms (lim (J := J) (C := C))] : HasExactLimitsOfShape J C where + preservesFiniteColimits := by + apply (config := { allowSynthFailures := true }) preservesFiniteColimits_of_preservesHomology + · exact preservesHomology_of_preservesEpis_and_kernels _ + · exact additive_of_preservesBinaryBiproducts _ + +end EpiMono + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean index 447bcadee24a4..8aa0fd2f01e4b 100644 --- a/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean +++ b/Mathlib/CategoryTheory/Abelian/Pseudoelements.lean @@ -427,20 +427,20 @@ section Module /-- In the category `Module R`, if `x` and `y` are pseudoequal, then the range of the associated morphisms is the same. -/ theorem ModuleCat.eq_range_of_pseudoequal {R : Type*} [CommRing R] {G : ModuleCat R} {x y : Over G} - (h : PseudoEqual G x y) : LinearMap.range x.hom = LinearMap.range y.hom := by + (h : PseudoEqual G x y) : LinearMap.range x.hom.hom = LinearMap.range y.hom.hom := by obtain ⟨P, p, q, hp, hq, H⟩ := h refine Submodule.ext fun a => ⟨fun ha => ?_, fun ha => ?_⟩ · obtain ⟨a', ha'⟩ := ha obtain ⟨a'', ha''⟩ := (ModuleCat.epi_iff_surjective p).1 hp a' refine ⟨q a'', ?_⟩ - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← LinearMap.comp_apply, ← ModuleCat.comp_def, ← H, - ModuleCat.comp_def, LinearMap.comp_apply, ha'', ha'] + dsimp at ha' ⊢ + rw [← LinearMap.comp_apply, ← ModuleCat.hom_comp, ← H, + ModuleCat.hom_comp, LinearMap.comp_apply, ha'', ha'] · obtain ⟨a', ha'⟩ := ha obtain ⟨a'', ha''⟩ := (ModuleCat.epi_iff_surjective q).1 hq a' refine ⟨p a'', ?_⟩ - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [← LinearMap.comp_apply, ← ModuleCat.comp_def, H, ModuleCat.comp_def, LinearMap.comp_apply, + dsimp at ha' ⊢ + rw [← LinearMap.comp_apply, ← ModuleCat.hom_comp, H, ModuleCat.hom_comp, LinearMap.comp_apply, ha'', ha'] end Module diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean index dd4c4c861c62c..1496b467c1f8a 100644 --- a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean +++ b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean @@ -4,12 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic -import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.Tactic.CategoryTheory.Elementwise /-! - -# Colimits in concrete categories with algebraic structures +# Colimits in ModuleCat Let `C` be a concrete category and `F : J ⥤ C` a filtered diagram in `C`. We discuss some results about `colimit F` when objects and morphisms in `C` have some algebraic structures. @@ -25,6 +24,11 @@ about `colimit F` when objects and morphisms in `C` have some algebraic structur and `x : Fᵢ` we have `r • x = 0` implies `x = 0`, then if `r • x = 0` for `x : colimit F`, then `x = 0`. +## Implementation details + +For now, we specialize our results to `C = ModuleCat R`, which is the only place they are used. +In the future they might be generalized by assuming a `HasForget₂ C (ModuleCat R)` instance, +plus assertions that the module structures induced by `HasForget₂` coincide. -/ universe t w v u r @@ -33,20 +37,20 @@ open CategoryTheory namespace CategoryTheory.Limits.Concrete -attribute [local instance] ConcreteCategory.instFunLike ConcreteCategory.hasCoeToSort - -variable {C : Type u} [Category.{v} C] [ConcreteCategory.{max t w} C] {J : Type w} [Category.{r} J] +variable (R : Type*) [Ring R] {J : Type w} [Category.{r} J] section zero theorem colimit_rep_eq_zero - (F : J ⥤ C) [PreservesColimit F (forget C)] [IsFiltered J] - [∀ c : C, Zero c] [∀ {c c' : C}, ZeroHomClass (c ⟶ c') c c'] [HasColimit F] - (j : J) (x : F.obj j) (hx : colimit.ι F j x = 0) : - ∃ (j' : J) (i : j ⟶ j'), F.map i x = 0 := by - rw [show 0 = colimit.ι F j 0 by simp, colimit_rep_eq_iff_exists] at hx + (F : J ⥤ ModuleCat.{max t w} R) [PreservesColimit F (forget (ModuleCat R))] [IsFiltered J] + [HasColimit F] (j : J) (x : F.obj j) (hx : colimit.ι F j x = 0) : + ∃ (j' : J) (i : j ⟶ j'), (F.map i).hom x = 0 := by + -- Break the abstraction barrier between homs and functions for `colimit_rep_eq_iff_exists`. + have : ∀ (X Y : ModuleCat R) (f : X ⟶ Y), + DFunLike.coe f.hom = DFunLike.coe (self := ConcreteCategory.instFunLike) f := fun _ _ _ => rfl + rw [show 0 = colimit.ι F j 0 by simp, this, colimit_rep_eq_iff_exists] at hx obtain ⟨j', i, y, g⟩ := hx - exact ⟨j', i, g ▸ by simp⟩ + exact ⟨j', i, g ▸ by simp [← this]⟩ end zero @@ -57,23 +61,25 @@ if `r` has no zero smul divisors for all small-enough sections, then `r` has no in the colimit. -/ lemma colimit_no_zero_smul_divisor - (F : J ⥤ C) [PreservesColimit F (forget C)] [IsFiltered J] [HasColimit F] - (R : Type*) [Semiring R] - [∀ c : C, AddCommMonoid c] [∀ c : C, Module R c] - [∀ {c c' : C}, LinearMapClass (c ⟶ c') R c c'] + (F : J ⥤ ModuleCat.{max t w} R) [PreservesColimit F (forget (ModuleCat R))] + [IsFiltered J] [HasColimit F] (r : R) (H : ∃ (j' : J), ∀ (j : J) (_ : j' ⟶ j), ∀ (c : F.obj j), r • c = 0 → c = 0) - (x : (forget C).obj (colimit F)) (hx : r • x = 0) : x = 0 := by + (x : (forget (ModuleCat R)).obj (colimit F)) (hx : r • x = 0) : x = 0 := by + + -- Break the abstraction barrier between homs and functions for `Concrete.colimit_exists_rep`. + have : ∀ (X Y : ModuleCat R) (f : X ⟶ Y), + DFunLike.coe f.hom = DFunLike.coe (self := ConcreteCategory.instFunLike) f := fun _ _ _ => rfl classical obtain ⟨j, x, rfl⟩ := Concrete.colimit_exists_rep F x - rw [← map_smul] at hx + rw [← this, ← map_smul (colimit.ι F j).hom] at hx obtain ⟨j', i, h⟩ := Concrete.colimit_rep_eq_zero (hx := hx) obtain ⟨j'', H⟩ := H - simpa [elementwise_of% (colimit.w F), map_zero] using congr(colimit.ι F _ + simpa [elementwise_of% (colimit.w F), this, map_zero] using congr(colimit.ι F _ $(H (IsFiltered.sup {j, j', j''} { ⟨j, j', by simp, by simp, i⟩ }) (IsFiltered.toSup _ _ <| by simp) (F.map (IsFiltered.toSup _ _ <| by simp) x) (by rw [← IsFiltered.toSup_commutes (f := i) (mY := by simp) (mf := by simp), F.map_comp, - comp_apply, ← map_smul, ← map_smul, h, map_zero]))) + ModuleCat.comp_apply, ← map_smul, ← map_smul, h, map_zero]))) end module diff --git a/Mathlib/CategoryTheory/Limits/Filtered.lean b/Mathlib/CategoryTheory/Limits/Filtered.lean index 536b215594402..f30d4ab31ca83 100644 --- a/Mathlib/CategoryTheory/Limits/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Filtered.lean @@ -17,7 +17,7 @@ Furthermore, we define the type classes `HasCofilteredLimitsOfSize` and `HasFilt -/ -universe w' w v u +universe w' w w₂' w₂ v u noncomputable section @@ -103,6 +103,32 @@ instance (priority := 100) hasColimitsOfShape_of_has_filtered_colimits HasColimitsOfShape I C := HasFilteredColimitsOfSize.HasColimitsOfShape _ +lemma hasCofilteredLimitsOfSize_of_univLE [UnivLE.{w, w₂}] [UnivLE.{w', w₂'}] + [HasCofilteredLimitsOfSize.{w₂', w₂} C] : + HasCofilteredLimitsOfSize.{w', w} C where + HasLimitsOfShape J := + haveI := IsCofiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)) + hasLimitsOfShape_of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)).symm + +lemma hasCofilteredLimitsOfSize_shrink [HasCofilteredLimitsOfSize.{max w' w₂', max w w₂} C] : + HasCofilteredLimitsOfSize.{w', w} C := + hasCofilteredLimitsOfSize_of_univLE.{w', w, max w' w₂', max w w₂} + +lemma hasFilteredColimitsOfSize_of_univLE [UnivLE.{w, w₂}] [UnivLE.{w', w₂'}] + [HasFilteredColimitsOfSize.{w₂', w₂} C] : + HasFilteredColimitsOfSize.{w', w} C where + HasColimitsOfShape J := + haveI := IsFiltered.of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)) + hasColimitsOfShape_of_equivalence ((ShrinkHoms.equivalence.{w₂'} J).trans <| + Shrink.equivalence.{w₂} (ShrinkHoms.{w} J)).symm + +lemma hasFilteredColimitsOfSize_shrink [HasFilteredColimitsOfSize.{max w' w₂', max w w₂} C] : + HasFilteredColimitsOfSize.{w', w} C := + hasFilteredColimitsOfSize_of_univLE.{w', w, max w' w₂', max w w₂} + end Limits end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index ed9bd8a77cbda..02cc099007693 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -653,12 +653,17 @@ abbrev HasCoproducts := variable {C} -theorem has_smallest_products_of_hasProducts [HasProducts.{w} C] : HasProducts.{0} C := fun J => - hasLimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w} J) ≌ _) +lemma hasProducts_shrink [HasProducts.{max w w'} C] : HasProducts.{w} C := fun J => + hasLimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _) + +lemma hasCoproducts_shrink [HasCoproducts.{max w w'} C] : HasCoproducts.{w} C := fun J => + hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w'} J) ≌ _) + +theorem has_smallest_products_of_hasProducts [HasProducts.{w} C] : HasProducts.{0} C := + hasProducts_shrink theorem has_smallest_coproducts_of_hasCoproducts [HasCoproducts.{w} C] : HasCoproducts.{0} C := - fun J => - hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift : Discrete (ULift.{w} J) ≌ _) + hasCoproducts_shrink theorem hasProducts_of_limit_fans (lf : ∀ {J : Type w} (f : J → C), Fan f) (lf_isLimit : ∀ {J : Type w} (f : J → C), IsLimit (lf f)) : HasProducts.{w} C := diff --git a/Mathlib/CategoryTheory/Linear/Yoneda.lean b/Mathlib/CategoryTheory/Linear/Yoneda.lean index ee468da7d7708..ae48cbcbb39b7 100644 --- a/Mathlib/CategoryTheory/Linear/Yoneda.lean +++ b/Mathlib/CategoryTheory/Linear/Yoneda.lean @@ -28,9 +28,6 @@ namespace CategoryTheory variable (R : Type w) [Ring R] {C : Type u} [Category.{v} C] [Preadditive C] [Linear R C] variable (C) --- Porting note: inserted specific `ModuleCat.asHom` in the definition of `linearYoneda` --- and similarly in `linearCoyoneda`, otherwise many simp lemmas are not triggered automatically. --- Eventually, doing so allows more proofs to be automatic! /-- The Yoneda embedding for `R`-linear categories `C`, sending an object `X : C` to the `ModuleCat R`-valued presheaf on `C`, with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ @@ -38,9 +35,9 @@ with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ def linearYoneda : C ⥤ Cᵒᵖ ⥤ ModuleCat R where obj X := { obj := fun Y => ModuleCat.of R (unop Y ⟶ X) - map := fun f => ModuleCat.asHom (Linear.leftComp R _ f.unop) } + map := fun f => ModuleCat.ofHom (Linear.leftComp R _ f.unop) } map {X₁ X₂} f := - { app := fun Y => @ModuleCat.asHom R _ (Y.unop ⟶ X₁) (Y.unop ⟶ X₂) _ _ _ _ + { app := fun Y => @ModuleCat.ofHom R _ (Y.unop ⟶ X₁) (Y.unop ⟶ X₂) _ _ _ _ (Linear.rightComp R _ f) } /-- The Yoneda embedding for `R`-linear categories `C`, @@ -50,9 +47,9 @@ with value on `X : C` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ def linearCoyoneda : Cᵒᵖ ⥤ C ⥤ ModuleCat R where obj Y := { obj := fun X => ModuleCat.of R (unop Y ⟶ X) - map := fun f => ModuleCat.asHom (Linear.rightComp R _ f) } + map := fun f => ModuleCat.ofHom (Linear.rightComp R _ f) } map {Y₁ Y₂} f := - { app := fun X => @ModuleCat.asHom R _ (unop Y₁ ⟶ X) (unop Y₂ ⟶ X) _ _ _ _ + { app := fun X => @ModuleCat.ofHom R _ (unop Y₁ ⟶ X) (unop Y₂ ⟶ X) _ _ _ _ (Linear.leftComp _ _ f.unop) } instance linearYoneda_obj_additive (X : C) : ((linearYoneda R C).obj X).Additive where diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 4b3678918fac2..69026eb1f7464 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -157,14 +157,20 @@ def trivial : Bimon_ C := Comon_.trivial (Mon_ C) /-- The bimonoid morphism from the trivial bimonoid to any bimonoid. -/ @[simps] -def trivial_to (A : Bimon_ C) : trivial C ⟶ A := +def trivialTo (A : Bimon_ C) : trivial C ⟶ A := { hom := (default : Mon_.trivial C ⟶ A.X), } +@[deprecated (since := "2024-12-07")] alias trivial_to := trivialTo +@[deprecated (since := "2024-12-07")] alias trivial_to_hom := trivialTo_hom + /-- The bimonoid morphism from any bimonoid to the trivial bimonoid. -/ @[simps!] -def to_trivial (A : Bimon_ C) : A ⟶ trivial C := +def toTrivial (A : Bimon_ C) : A ⟶ trivial C := (default : @Quiver.Hom (Comon_ (Mon_ C)) _ A (Comon_.trivial (Mon_ C))) +@[deprecated (since := "2024-12-07")] alias to_trivial := toTrivial +@[deprecated (since := "2024-12-07")] alias to_trivial_hom := toTrivial_hom + /-! # Additional lemmas -/ variable {C} diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index 9341953c58518..94382ab42f817 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -45,19 +45,19 @@ instance Ring_of_Mon_ (A : Mon_ (ModuleCat.{u} R)) : Ring A.X := one := A.one (1 : R) mul := fun x y => A.mul (x ⊗ₜ y) one_mul := fun x => by - convert LinearMap.congr_fun A.one_mul ((1 : R) ⊗ₜ x) + convert LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul) ((1 : R) ⊗ₜ x) rw [MonoidalCategory.leftUnitor_hom_apply, one_smul] mul_one := fun x => by - convert LinearMap.congr_fun A.mul_one (x ⊗ₜ (1 : R)) - erw [MonoidalCategory.leftUnitor_hom_apply, one_smul] + convert LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.mul_one) (x ⊗ₜ (1 : R)) + rw [MonoidalCategory.rightUnitor_hom_apply, one_smul] mul_assoc := fun x y z => by - convert LinearMap.congr_fun A.mul_assoc (x ⊗ₜ y ⊗ₜ z) + convert LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.mul_assoc) (x ⊗ₜ y ⊗ₜ z) left_distrib := fun x y z => by - convert A.mul.map_add (x ⊗ₜ y) (x ⊗ₜ z) + convert A.mul.hom.map_add (x ⊗ₜ y) (x ⊗ₜ z) rw [← TensorProduct.tmul_add] rfl right_distrib := fun x y z => by - convert A.mul.map_add (x ⊗ₜ z) (y ⊗ₜ z) + convert A.mul.hom.map_add (x ⊗ₜ z) (y ⊗ₜ z) rw [← TensorProduct.add_tmul] rfl zero_mul := fun x => show A.mul _ = 0 by @@ -66,18 +66,19 @@ instance Ring_of_Mon_ (A : Mon_ (ModuleCat.{u} R)) : Ring A.X := rw [TensorProduct.tmul_zero, map_zero] } instance Algebra_of_Mon_ (A : Mon_ (ModuleCat.{u} R)) : Algebra R A.X := - { A.one with - map_zero' := A.one.map_zero + { A.one.hom with + map_zero' := A.one.hom.map_zero map_one' := rfl map_mul' := fun x y => by - have h := LinearMap.congr_fun A.one_mul.symm (x ⊗ₜ A.one y) - rwa [MonoidalCategory.leftUnitor_hom_apply, ← A.one.map_smul] at h + have h := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul.symm) (x ⊗ₜ A.one y) + rwa [MonoidalCategory.leftUnitor_hom_apply, ← A.one.hom.map_smul] at h commutes' := fun r a => by dsimp - have h₁ := LinearMap.congr_fun A.one_mul (r ⊗ₜ a) - have h₂ := LinearMap.congr_fun A.mul_one (a ⊗ₜ r) + have h₁ := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul) (r ⊗ₜ a) + have h₂ := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.mul_one) (a ⊗ₜ r) exact h₁.trans h₂.symm - smul_def' := fun r a => (LinearMap.congr_fun A.one_mul (r ⊗ₜ a)).symm } + smul_def' := fun r a => + (LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp A.one_mul) (r ⊗ₜ a)).symm } @[simp] theorem algebraMap (A : Mon_ (ModuleCat.{u} R)) (r : R) : algebraMap R A.X r = A.one r := @@ -89,33 +90,35 @@ theorem algebraMap (A : Mon_ (ModuleCat.{u} R)) (r : R) : algebraMap R A.X r = A def functor : Mon_ (ModuleCat.{u} R) ⥤ AlgebraCat R where obj A := AlgebraCat.of R A.X map {_ _} f := AlgebraCat.ofHom - { f.hom.toAddMonoidHom with + { f.hom.hom.toAddMonoidHom with toFun := f.hom - map_one' := LinearMap.congr_fun f.one_hom (1 : R) - map_mul' := fun x y => LinearMap.congr_fun f.mul_hom (x ⊗ₜ y) - commutes' := fun r => LinearMap.congr_fun f.one_hom r } + map_one' := LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp f.one_hom) (1 : R) + map_mul' := fun x y => LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp f.mul_hom) (x ⊗ₜ y) + commutes' := fun r => LinearMap.congr_fun (ModuleCat.hom_ext_iff.mp f.one_hom) r } /-- Converting a bundled algebra to a monoid object in `ModuleCat R`. -/ @[simps] def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where X := ModuleCat.of R A - one := Algebra.linearMap R A - mul := LinearMap.mul' R A + one := ofHom <| Algebra.linearMap R A + mul := ofHom <| LinearMap.mul' R A one_mul := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext_ring <| LinearMap.ext fun x => ?_ - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [compr₂_apply, compr₂_apply, CategoryTheory.comp_apply] + rw [compr₂_apply, compr₂_apply, hom_comp, LinearMap.comp_apply] -- Porting note: this `dsimp` does nothing -- dsimp [AlgebraCat.id_apply, TensorProduct.mk_apply, Algebra.linearMap_apply, - -- LinearMap.compr₂_apply, Function.comp_apply, RingHom.map_one, - -- ModuleCat.MonoidalCategory.hom_apply, AlgebraCat.coe_comp, - -- ModuleCat.MonoidalCategory.leftUnitor_hom_apply] + -- LinearMap.compr₂_apply, Function.comp_apply, RingHom.map_one, + -- ModuleCat.MonoidalCategory.tensorHom_tmul, AlgebraCat.hom_comp, + -- ModuleCat.MonoidalCategory.leftUnitor_hom_apply] -- Porting note: because `dsimp` is not effective, `rw` needs to be changed to `erw` + dsimp erw [LinearMap.mul'_apply, MonoidalCategory.leftUnitor_hom_apply, ← Algebra.smul_def] - erw [id_apply] + dsimp mul_one := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext_ring ?_ -- Porting note: this `dsimp` does nothing @@ -124,23 +127,22 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where -- AlgebraCat.coe_comp] -- Porting note: because `dsimp` is not effective, `rw` needs to be changed to `erw` erw [compr₂_apply, compr₂_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [CategoryTheory.comp_apply] + rw [ModuleCat.hom_comp, LinearMap.comp_apply] erw [LinearMap.mul'_apply, ModuleCat.MonoidalCategory.rightUnitor_hom_apply, ← Algebra.commutes, ← Algebra.smul_def] - erw [id_apply] + dsimp mul_assoc := by + ext : 1 set_option tactic.skipAssignedInstances false in -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => LinearMap.ext fun z => ?_ dsimp only [compr₂_apply, TensorProduct.mk_apply] rw [compr₂_apply, compr₂_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [CategoryTheory.comp_apply, - CategoryTheory.comp_apply, CategoryTheory.comp_apply] + rw [hom_comp, LinearMap.comp_apply, hom_comp, LinearMap.comp_apply, hom_comp, + LinearMap.comp_apply] erw [LinearMap.mul'_apply, LinearMap.mul'_apply] - erw [id_apply] + dsimp only [id_coe, id_eq] erw [TensorProduct.mk_apply, TensorProduct.mk_apply, mul'_apply, LinearMap.id_apply, mul'_apply] simp only [LinearMap.mul'_apply, mul_assoc] @@ -150,9 +152,9 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where def inverse : AlgebraCat.{u} R ⥤ Mon_ (ModuleCat.{u} R) where obj := inverseObj map f := - { hom := f.hom.toLinearMap - one_hom := LinearMap.ext f.hom.commutes - mul_hom := TensorProduct.ext <| LinearMap.ext₂ <| map_mul f.hom } + { hom := ofHom <| f.hom.toLinearMap + one_hom := hom_ext <| LinearMap.ext f.hom.commutes + mul_hom := hom_ext <| TensorProduct.ext <| LinearMap.ext₂ <| map_mul f.hom } end MonModuleEquivalenceAlgebra @@ -169,21 +171,23 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where NatIso.ofComponents (fun A => { hom := - { hom := + { hom := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } mul_hom := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext ?_ dsimp at * rfl } inv := - { hom := + { hom := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } mul_hom := by + ext : 1 -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext ?_ dsimp at * @@ -214,11 +218,11 @@ def monModuleEquivalenceAlgebraForget : Mon_.forget (ModuleCat.{u} R) := NatIso.ofComponents (fun A => - { hom := + { hom := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } - inv := + inv := ofHom { toFun := _root_.id map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } }) diff --git a/Mathlib/CategoryTheory/MorphismProperty/Retract.lean b/Mathlib/CategoryTheory/MorphismProperty/Retract.lean new file mode 100644 index 0000000000000..af653b25df92e --- /dev/null +++ b/Mathlib/CategoryTheory/MorphismProperty/Retract.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2024 Jack McKoen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jack McKoen +-/ +import Mathlib.CategoryTheory.Retract +import Mathlib.CategoryTheory.MorphismProperty.Basic + +/-! +# Stability under retracts + +Given `P : MorphismProperty C`, we introduce a typeclass `P.IsStableUnderRetracts` which +is the property that `P` is stable under retracts. + +-/ + +universe v u + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] + +namespace MorphismProperty + +/-- A class of morphisms is stable under retracts if a retract of such a morphism still +lies in the class. -/ +class IsStableUnderRetracts (P : MorphismProperty C) : Prop where + of_retract {X Y Z W : C} {f : X ⟶ Y} {g : Z ⟶ W} (h : RetractArrow f g) (hg : P g) : P f + +lemma of_retract {P : MorphismProperty C} [P.IsStableUnderRetracts] + {X Y Z W : C} {f : X ⟶ Y} {g : Z ⟶ W} (h : RetractArrow f g) (hg : P g) : P f := + IsStableUnderRetracts.of_retract h hg + +instance IsStableUnderRetracts.monomorphisms : (monomorphisms C).IsStableUnderRetracts where + of_retract {_ _ _ _ f g} h (hg : Mono g) := ⟨fun α β w ↦ by + rw [← cancel_mono h.i.left, ← cancel_mono g, Category.assoc, Category.assoc, + h.i_w, reassoc_of% w]⟩ + +instance IsStableUnderRetracts.epimorphisms : (epimorphisms C).IsStableUnderRetracts where + of_retract {_ _ _ _ f g} h (hg : Epi g) := ⟨fun α β w ↦ by + rw [← cancel_epi h.r.right, ← cancel_epi g, ← Category.assoc, ← Category.assoc, ← h.r_w, + Category.assoc, Category.assoc, w]⟩ + +instance IsStableUnderRetracts.isomorphisms : (isomorphisms C).IsStableUnderRetracts where + of_retract {X Y Z W f g} h (_ : IsIso _) := by + refine ⟨h.i.right ≫ inv g ≫ h.r.left, ?_, ?_⟩ + · rw [← h.i_w_assoc, IsIso.hom_inv_id_assoc, h.retract_left] + · rw [Category.assoc, Category.assoc, h.r_w, IsIso.inv_hom_id_assoc, h.retract_right] + +end MorphismProperty + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean index 812bfc902cea1..beedbdb8d0a67 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean @@ -39,7 +39,7 @@ object `X` to the `End Y`-module of morphisms `X ⟶ Y`. @[simps] def preadditiveYonedaObj (Y : C) : Cᵒᵖ ⥤ ModuleCat.{v} (End Y) where obj X := ModuleCat.of _ (X.unop ⟶ Y) - map f := ModuleCat.asHom + map f := ModuleCat.ofHom { toFun := fun g => f.unop ≫ g map_add' := fun _ _ => comp_add _ _ _ _ _ _ map_smul' := fun _ _ => Eq.symm <| Category.assoc _ _ _ } @@ -66,7 +66,7 @@ object `Y` to the `End X`-module of morphisms `X ⟶ Y`. @[simps] def preadditiveCoyonedaObj (X : Cᵒᵖ) : C ⥤ ModuleCat.{v} (End X) where obj Y := ModuleCat.of _ (unop X ⟶ Y) - map f := ModuleCat.asHom + map f := ModuleCat.ofHom { toFun := fun g => g ≫ f map_add' := fun _ _ => add_comp _ _ _ _ _ _ map_smul' := fun _ _ => Category.assoc _ _ _ } diff --git a/Mathlib/CategoryTheory/Retract.lean b/Mathlib/CategoryTheory/Retract.lean new file mode 100644 index 0000000000000..952554b768d5a --- /dev/null +++ b/Mathlib/CategoryTheory/Retract.lean @@ -0,0 +1,106 @@ +/- +Copyright (c) 2024 Jack McKoen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jack McKoen +-/ +import Mathlib.CategoryTheory.Comma.Arrow +import Mathlib.CategoryTheory.EpiMono + +/-! +# Retracts + +Defines retracts of objects and morphisms. + +-/ + +universe v v' u u' + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D] + +/-- An object `X` is a retract of `Y` if there are morphisms `i : X ⟶ Y` and `r : Y ⟶ X` such +that `i ≫ r = 𝟙 X`. -/ +structure Retract (X Y : C) where + /-- the split monomorphism -/ + i : X ⟶ Y + /-- the split epimorphism -/ + r : Y ⟶ X + retract : i ≫ r = 𝟙 X := by aesop_cat + +namespace Retract + +attribute [reassoc (attr := simp)] retract + +variable {X Y : C} (h : Retract X Y) + +/-- If `X` is a retract of `Y`, then `F.obj X` is a retract of `F.obj Y`. -/ +@[simps] +def map (F : C ⥤ D) : Retract (F.obj X) (F.obj Y) where + i := F.map h.i + r := F.map h.r + retract := by rw [← F.map_comp h.i h.r, h.retract, F.map_id] + +/-- a retract determines a split epimorphism. -/ +@[simps] def splitEpi : SplitEpi h.r where + section_ := h.i + +/-- a retract determines a split monomorphism. -/ +@[simps] def splitMono : SplitMono h.i where + retraction := h.r + +instance : IsSplitEpi h.r := ⟨⟨h.splitEpi⟩⟩ + +instance : IsSplitMono h.i := ⟨⟨h.splitMono⟩⟩ + +end Retract + +/-- +``` + X -------> Z -------> X + | | | + f g f + | | | + v v v + Y -------> W -------> Y + +``` +A morphism `f : X ⟶ Y` is a retract of `g : Z ⟶ W` if there are morphisms `i : f ⟶ g` +and `r : g ⟶ f` in the arrow category such that `i ≫ r = 𝟙 f`. -/ +abbrev RetractArrow {X Y Z W : C} (f : X ⟶ Y) (g : Z ⟶ W) := Retract (Arrow.mk f) (Arrow.mk g) + +namespace RetractArrow + +variable {X Y Z W : C} {f : X ⟶ Y} {g : Z ⟶ W} (h : RetractArrow f g) + +@[reassoc] +lemma i_w : h.i.left ≫ g = f ≫ h.i.right := h.i.w + +@[reassoc] +lemma r_w : h.r.left ≫ f = g ≫ h.r.right := h.r.w + +/-- The top of a retract diagram of morphisms determines a retract of objects. -/ +@[simps!] +def left : Retract X Z := h.map Arrow.leftFunc + +/-- The bottom of a retract diagram of morphisms determines a retract of objects. -/ +@[simps!] +def right : Retract Y W := h.map Arrow.rightFunc + +@[reassoc (attr := simp)] +lemma retract_left : h.i.left ≫ h.r.left = 𝟙 X := h.left.retract + +@[reassoc (attr := simp)] +lemma retract_right : h.i.right ≫ h.r.right = 𝟙 Y := h.right.retract + +instance : IsSplitEpi h.r.left := ⟨⟨h.left.splitEpi⟩⟩ + +instance : IsSplitEpi h.r.right := ⟨⟨h.right.splitEpi⟩⟩ + +instance : IsSplitMono h.i.left := ⟨⟨h.left.splitMono⟩⟩ + +instance : IsSplitMono h.i.right := ⟨⟨h.right.splitMono⟩⟩ + +end RetractArrow + +end CategoryTheory diff --git a/Mathlib/Combinatorics/Enumerative/Catalan.lean b/Mathlib/Combinatorics/Enumerative/Catalan.lean index 21a578fab89bf..d7dd958d1d693 100644 --- a/Mathlib/Combinatorics/Enumerative/Catalan.lean +++ b/Mathlib/Combinatorics/Enumerative/Catalan.lean @@ -30,7 +30,7 @@ triangulations of convex polygons. * `catalan_eq_centralBinom_div`: The explicit formula for the Catalan number using the central binomial coefficient, `catalan n = Nat.centralBinom n / (n + 1)`. -* `treesOfNodesEq_card_eq_catalan`: The number of binary trees with `n` internal nodes +* `treesOfNumNodesEq_card_eq_catalan`: The number of binary trees with `n` internal nodes is `catalan n` ## Implementation details diff --git a/Mathlib/Combinatorics/Hall/Basic.lean b/Mathlib/Combinatorics/Hall/Basic.lean index da15e7bee4a35..b7de187e22a39 100644 --- a/Mathlib/Combinatorics/Hall/Basic.lean +++ b/Mathlib/Combinatorics/Hall/Basic.lean @@ -160,8 +160,7 @@ instance {α : Type u} {β : Type v} [DecidableEq β] (r : α → β → Prop) [∀ a : α, Fintype (Rel.image r {a})] (A : Finset α) : Fintype (Rel.image r A) := by have h : Rel.image r A = (A.biUnion fun a => (Rel.image r {a}).toFinset : Set β) := by ext - -- Porting note: added `Set.mem_toFinset` - simp [Rel.image, (Set.mem_toFinset)] + simp [Rel.image] rw [h] apply FinsetCoe.fintype @@ -185,9 +184,9 @@ theorem Fintype.all_card_le_rel_image_card_iff_exists_injective {α : Type u} { rw [← Set.toFinset_card] apply congr_arg ext b - -- Porting note: added `Set.mem_toFinset` + -- Porting note: added `Set.mem_toFinset` (fixed by lean#6123) simp [Rel.image, (Set.mem_toFinset)] - -- Porting note: added `Set.mem_toFinset` + -- Porting note: added `Set.mem_toFinset` (fixed by lean#6123) have h' : ∀ (f : α → β) (x), r x (f x) ↔ f x ∈ r' x := by simp [Rel.image, (Set.mem_toFinset)] simp only [h, h'] apply Finset.all_card_le_biUnion_card_iff_exists_injective diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 1da6a51b3e613..fbf6473d0fddf 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -603,7 +603,7 @@ theorem fromEdgeSet_sdiff (s t : Set (Sym2 V)) : ext v w constructor <;> simp +contextual -@[mono] +@[gcongr, mono] theorem fromEdgeSet_mono {s t : Set (Sym2 V)} (h : s ⊆ t) : fromEdgeSet s ≤ fromEdgeSet t := by rintro v w simp +contextual only [fromEdgeSet_adj, Ne, not_false_iff, diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index dd588e24a5611..ea2c7d7ec4dbc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -493,7 +493,7 @@ theorem cliqueSet_eq_empty_iff : G.cliqueSet n = ∅ ↔ G.CliqueFree n := by protected alias ⟨_, CliqueFree.cliqueSet⟩ := cliqueSet_eq_empty_iff -@[mono] +@[gcongr, mono] theorem cliqueSet_mono (h : G ≤ H) : G.cliqueSet n ⊆ H.cliqueSet n := fun _ ↦ IsNClique.mono h @@ -636,7 +636,7 @@ theorem card_cliqueFinset_le : #(G.cliqueFinset n) ≤ (card α).choose n := by variable [DecidableRel H.Adj] -@[mono] +@[gcongr, mono] theorem cliqueFinset_mono (h : G ≤ H) : G.cliqueFinset n ⊆ H.cliqueFinset n := monotone_filter_right _ fun _ ↦ IsNClique.mono h diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 5d6ee34bb110b..1e27af2e43ad5 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -333,11 +333,9 @@ to the same vertex. lemma IsCycles.other_adj_of_adj (h : G.IsCycles) (hadj : G.Adj v w) : ∃ w', w ≠ w' ∧ G.Adj v w' := by simp_rw [← SimpleGraph.mem_neighborSet] at hadj ⊢ - cases' (G.neighborSet v).eq_empty_or_nonempty with hl hr - · exact ((Set.eq_empty_iff_forall_not_mem.mp hl) _ hadj).elim - · have := h hr - obtain ⟨w', hww'⟩ := (G.neighborSet v).exists_ne_of_one_lt_ncard (by omega) w - exact ⟨w', ⟨hww'.2.symm, hww'.1⟩⟩ + have := h ⟨w, hadj⟩ + obtain ⟨w', hww'⟩ := (G.neighborSet v).exists_ne_of_one_lt_ncard (by omega) w + exact ⟨w', ⟨hww'.2.symm, hww'.1⟩⟩ open scoped symmDiff diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 3a54a207a63cd..3ee260f7663ee 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -1077,11 +1077,11 @@ theorem induce_mono (hg : G' ≤ G'') (hs : s ⊆ s') : G'.induce s ≤ G''.indu intro v w hv hw ha exact ⟨hs hv, hs hw, hg.2 ha⟩ -@[mono] +@[gcongr, mono] theorem induce_mono_left (hg : G' ≤ G'') : G'.induce s ≤ G''.induce s := induce_mono hg subset_rfl -@[mono] +@[gcongr, mono] theorem induce_mono_right (hs : s ⊆ s') : G'.induce s ≤ G'.induce s' := induce_mono le_rfl hs @@ -1159,12 +1159,12 @@ theorem deleteVerts_empty : G'.deleteVerts ∅ = G' := by theorem deleteVerts_le : G'.deleteVerts s ≤ G' := by constructor <;> simp [Set.diff_subset] -@[mono] +@[gcongr, mono] theorem deleteVerts_mono {G' G'' : G.Subgraph} (h : G' ≤ G'') : G'.deleteVerts s ≤ G''.deleteVerts s := induce_mono h (Set.diff_subset_diff_left h.1) -@[mono] +@[gcongr, mono] theorem deleteVerts_anti {s s' : Set V} (h : s ⊆ s') : G'.deleteVerts s' ≤ G'.deleteVerts s := induce_mono (le_refl _) (Set.diff_subset_diff_right h) diff --git a/Mathlib/Combinatorics/Young/SemistandardTableau.lean b/Mathlib/Combinatorics/Young/SemistandardTableau.lean index fac50829baf05..52168fe27c5c2 100644 --- a/Mathlib/Combinatorics/Young/SemistandardTableau.lean +++ b/Mathlib/Combinatorics/Young/SemistandardTableau.lean @@ -29,8 +29,8 @@ for all pairs `(i, j) ∉ μ` and to satisfy the row-weak and column-strict cond - `SemistandardYoungTableau (μ : YoungDiagram)`: semistandard Young tableaux of shape `μ`. There is a `coe` instance such that `T i j` is value of the `(i, j)` entry of the semistandard Young tableau `T`. -- `Ssyt.highestWeight (μ : YoungDiagram)`: the semistandard Young tableau whose `i`th row - consists entirely of `i`s, for each `i`. +- `SemistandardYoungTableau.highestWeight (μ : YoungDiagram)`: the semistandard Young tableau whose + `i`th row consists entirely of `i`s, for each `i`. ## Tags diff --git a/Mathlib/Condensed/Discrete/Module.lean b/Mathlib/Condensed/Discrete/Module.lean index 5657c6d51579e..75760ce80c094 100644 --- a/Mathlib/Condensed/Discrete/Module.lean +++ b/Mathlib/Condensed/Discrete/Module.lean @@ -40,8 +40,8 @@ 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 ↦ comapₗ R f.unop } - map f := { app := fun S ↦ mapₗ R f } + map := fun f ↦ ModuleCat.ofHom (comapₗ R f.unop) } + map f := { app := fun S ↦ ModuleCat.ofHom (mapₗ R f.hom) } variable [HasExplicitFiniteCoproducts.{0} P] [HasExplicitPullbacks.{u} P] (hs : ∀ ⦃X Y : CompHausLike P⦄ (f : X ⟶ Y), EffectiveEpi f → Function.Surjective f) @@ -79,8 +79,8 @@ abbrev functor : ModuleCat R ⥤ CondensedMod.{u} R := /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u+1} R) : M ≅ (ModuleCat.of R (LocallyConstant (CompHaus.of PUnit.{u+1}) M)) where - hom := constₗ R - inv := evalₗ R PUnit.unit + hom := ModuleCat.ofHom (constₗ R) + inv := ModuleCat.ofHom (evalₗ R PUnit.unit) /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat R) : @@ -186,8 +186,8 @@ abbrev functor : ModuleCat R ⥤ LightCondMod.{u} R := /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₁ (M : ModuleCat.{u} R) : M ≅ (ModuleCat.of R (LocallyConstant (LightProfinite.of PUnit.{u+1}) M)) where - hom := constₗ R - inv := evalₗ R PUnit.unit + hom := ModuleCat.ofHom (constₗ R) + inv := ModuleCat.ofHom (evalₗ R PUnit.unit) /-- Auxiliary definition for `functorIsoDiscrete`. -/ noncomputable def functorIsoDiscreteAux₂ (M : ModuleCat.{u} R) : diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 4fcd646907769..fac3a3ec93421 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -269,7 +269,7 @@ This one instead uses a `NeZero n` typeclass hypothesis. theorem pos_iff_ne_zero' [NeZero n] (a : Fin n) : 0 < a ↔ a ≠ 0 := by rw [← val_fin_lt, val_zero', Nat.pos_iff_ne_zero, Ne, Ne, Fin.ext_iff, val_zero'] -@[simp] lemma cast_eq_self (a : Fin n) : cast rfl a = a := rfl +@[simp] lemma cast_eq_self (a : Fin n) : a.cast rfl = a := rfl @[simp] theorem cast_eq_zero {k l : ℕ} [NeZero k] [NeZero l] (h : k = l) (x : Fin k) : Fin.cast h x = 0 ↔ x = 0 := by simp [← val_eq_val] @@ -299,7 +299,7 @@ theorem revPerm_symm : (@revPerm n).symm = revPerm := rfl theorem cast_rev (i : Fin n) (h : n = m) : - cast h i.rev = (i.cast h).rev := by + i.rev.cast h = (i.cast h).rev := by subst h; simp theorem rev_eq_iff {i j : Fin n} : rev i = j ↔ i = rev j := by @@ -619,23 +619,23 @@ theorem coe_of_injective_castLE_symm {n k : ℕ} (h : n ≤ k) (i : Fin k) (hi) rw [← coe_castLE h] exact congr_arg Fin.val (Equiv.apply_ofInjective_symm _ _) -theorem leftInverse_cast (eq : n = m) : LeftInverse (cast eq.symm) (cast eq) := +theorem leftInverse_cast (eq : n = m) : LeftInverse (Fin.cast eq.symm) (Fin.cast eq) := fun _ => rfl -theorem rightInverse_cast (eq : n = m) : RightInverse (cast eq.symm) (cast eq) := +theorem rightInverse_cast (eq : n = m) : RightInverse (Fin.cast eq.symm) (Fin.cast eq) := fun _ => rfl -theorem cast_lt_cast (eq : n = m) {a b : Fin n} : cast eq a < cast eq b ↔ a < b := +theorem cast_lt_cast (eq : n = m) {a b : Fin n} : a.cast eq < b.cast eq ↔ a < b := Iff.rfl -theorem cast_le_cast (eq : n = m) {a b : Fin n} : cast eq a ≤ cast eq b ↔ a ≤ b := +theorem cast_le_cast (eq : n = m) {a b : Fin n} : a.cast eq ≤ b.cast eq ↔ a ≤ b := Iff.rfl /-- The 'identity' equivalence between `Fin m` and `Fin n` when `m = n`. -/ @[simps] def _root_.finCongr (eq : n = m) : Fin n ≃ Fin m where - toFun := cast eq - invFun := cast eq.symm + toFun := Fin.cast eq + invFun := Fin.cast eq.symm left_inv := leftInverse_cast eq right_inv := rightInverse_cast eq @@ -656,12 +656,12 @@ a generic theorem about `cast`. -/ lemma _root_.finCongr_eq_equivCast (h : n = m) : finCongr h = .cast (h ▸ rfl) := by subst h; simp @[simp] -theorem cast_zero {n' : ℕ} [NeZero n] {h : n = n'} : cast h (0 : Fin n) = +theorem cast_zero {n' : ℕ} [NeZero n] {h : n = n'} : (0 : Fin n).cast h = by { haveI : NeZero n' := by {rw [← h]; infer_instance}; exact 0} := rfl /-- While in many cases `Fin.cast` is better than `Equiv.cast`/`cast`, sometimes we want to apply a generic theorem about `cast`. -/ -theorem cast_eq_cast (h : n = m) : (cast h : Fin n → Fin m) = _root_.cast (h ▸ rfl) := by +theorem cast_eq_cast (h : n = m) : (Fin.cast h : Fin n → Fin m) = _root_.cast (h ▸ rfl) := by subst h ext rfl diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 430153e9bdb75..e8bfa47eb7aa6 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -395,7 +395,7 @@ theorem cons_eq_append (x : α) (xs : Fin n → α) : subst h; simp lemma append_rev {m n} (xs : Fin m → α) (ys : Fin n → α) (i : Fin (m + n)) : - append xs ys (rev i) = append (ys ∘ rev) (xs ∘ rev) (cast (Nat.add_comm ..) i) := by + append xs ys (rev i) = append (ys ∘ rev) (xs ∘ rev) (i.cast (Nat.add_comm ..)) := by rcases rev_surjective i with ⟨i, rfl⟩ rw [rev_rev] induction i using Fin.addCases @@ -403,7 +403,7 @@ lemma append_rev {m n} (xs : Fin m → α) (ys : Fin n → α) (i : Fin (m + n)) · simp [cast_rev, rev_addNat] lemma append_comp_rev {m n} (xs : Fin m → α) (ys : Fin n → α) : - append xs ys ∘ rev = append (ys ∘ rev) (xs ∘ rev) ∘ cast (Nat.add_comm ..) := + append xs ys ∘ rev = append (ys ∘ rev) (xs ∘ rev) ∘ Fin.cast (Nat.add_comm ..) := funext <| append_rev xs ys theorem append_castAdd_natAdd {f : Fin (m + n) → α} : @@ -430,11 +430,11 @@ theorem repeat_apply (a : Fin n → α) (i : Fin (m * n)) : @[simp] theorem repeat_zero (a : Fin n → α) : - Fin.repeat 0 a = Fin.elim0 ∘ cast (Nat.zero_mul _) := - funext fun x => (cast (Nat.zero_mul _) x).elim0 + Fin.repeat 0 a = Fin.elim0 ∘ Fin.cast (Nat.zero_mul _) := + funext fun x => (x.cast (Nat.zero_mul _)).elim0 @[simp] -theorem repeat_one (a : Fin n → α) : Fin.repeat 1 a = a ∘ cast (Nat.one_mul _) := by +theorem repeat_one (a : Fin n → α) : Fin.repeat 1 a = a ∘ Fin.cast (Nat.one_mul _) := by generalize_proofs h apply funext rw [(Fin.rightInverse_cast h.symm).surjective.forall] @@ -443,7 +443,7 @@ theorem repeat_one (a : Fin n → α) : Fin.repeat 1 a = a ∘ cast (Nat.one_mul theorem repeat_succ (a : Fin n → α) (m : ℕ) : Fin.repeat m.succ a = - append a (Fin.repeat m a) ∘ cast ((Nat.succ_mul _ _).trans (Nat.add_comm ..)) := by + append a (Fin.repeat m a) ∘ Fin.cast ((Nat.succ_mul _ _).trans (Nat.add_comm ..)) := by generalize_proofs h apply funext rw [(Fin.rightInverse_cast h.symm).surjective.forall] @@ -453,7 +453,7 @@ theorem repeat_succ (a : Fin n → α) (m : ℕ) : @[simp] theorem repeat_add (a : Fin n → α) (m₁ m₂ : ℕ) : Fin.repeat (m₁ + m₂) a = - append (Fin.repeat m₁ a) (Fin.repeat m₂ a) ∘ cast (Nat.add_mul ..) := by + append (Fin.repeat m₁ a) (Fin.repeat m₂ a) ∘ Fin.cast (Nat.add_mul ..) := by generalize_proofs h apply funext rw [(Fin.rightInverse_cast h.symm).surjective.forall] diff --git a/Mathlib/Data/Fin/Tuple/Take.lean b/Mathlib/Data/Fin/Tuple/Take.lean index 462351c63a8d1..f54535ef33286 100644 --- a/Mathlib/Data/Fin/Tuple/Take.lean +++ b/Mathlib/Data/Fin/Tuple/Take.lean @@ -132,7 +132,7 @@ theorem take_addCases_right {n' : ℕ} {motive : Fin (n + n') → Sort*} (m : by_cases h' : i < n · simp only [h', ↓reduceDIte] congr - · simp only [h', ↓reduceDIte, subNat, castLE, cast, eqRec_eq_cast] + · simp only [h', ↓reduceDIte, subNat, castLE, Fin.cast, eqRec_eq_cast] /-- Version of `take_addCases_right` that specializes `addCases` to `append`. -/ theorem take_append_right {n' : ℕ} {α : Sort*} (m : ℕ) (h : m ≤ n') (u : (i : Fin n) → α) diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index 1b8ca2379a123..aeb31bb6c4a94 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -167,7 +167,6 @@ theorem range_cons_cons_empty (x y : α) (u : Fin 0 → α) : Set.range (vecCons x <| vecCons y u) = {x, y} := by rw [range_cons, range_cons_empty, Set.singleton_union] -@[simp] theorem vecCons_const (a : α) : (vecCons a fun _ : Fin n => a) = fun _ => a := funext <| Fin.forall_iff_succ.2 ⟨rfl, cons_val_succ _ _⟩ diff --git a/Mathlib/Data/Finset/Filter.lean b/Mathlib/Data/Finset/Filter.lean index d42d8d22414fa..746efb1723405 100644 --- a/Mathlib/Data/Finset/Filter.lean +++ b/Mathlib/Data/Finset/Filter.lean @@ -214,6 +214,15 @@ lemma _root_.Set.pairwiseDisjoint_filter [DecidableEq β] (f : α → β) (s : S obtain ⟨-, rfl⟩ : x ∈ t ∧ f x = j := by simpa using hj hx contradiction +theorem disjoint_filter_and_not_filter : + Disjoint (s.filter (fun x ↦ p x ∧ ¬q x)) (s.filter (fun x ↦ q x ∧ ¬p x)) := by + intro _ htp htq + simp [bot_eq_empty, le_eq_subset, subset_empty] + by_contra hn + rw [← not_nonempty_iff_eq_empty, not_not] at hn + obtain ⟨_, hx⟩ := hn + exact (mem_filter.mp (htq hx)).2.2 (mem_filter.mp (htp hx)).2.1 + variable {p q} lemma filter_inj : s.filter p = t.filter p ↔ ∀ ⦃a⦄, p a → (a ∈ s ↔ a ∈ t) := by diff --git a/Mathlib/Data/Finset/Insert.lean b/Mathlib/Data/Finset/Insert.lean index 1b93ae1f21cc9..8d5f792fbf340 100644 --- a/Mathlib/Data/Finset/Insert.lean +++ b/Mathlib/Data/Finset/Insert.lean @@ -216,6 +216,10 @@ instance (i : α) : Unique ({i} : Finset α) where @[simp] lemma default_singleton (i : α) : ((default : ({i} : Finset α)) : α) = i := rfl +instance Nontrivial.instDecidablePred [DecidableEq α] : + DecidablePred (Finset.Nontrivial (α := α)) := + inferInstanceAs (DecidablePred fun s ↦ ∃ a ∈ s, ∃ b ∈ s, a ≠ b) + end Singleton /-! ### cons -/ @@ -411,9 +415,11 @@ theorem insert_ne_self : insert a s ≠ s ↔ a ∉ s := theorem pair_eq_singleton (a : α) : ({a, a} : Finset α) = {a} := insert_eq_of_mem <| mem_singleton_self _ -theorem Insert.comm (a b : α) (s : Finset α) : insert a (insert b s) = insert b (insert a s) := +theorem insert_comm (a b : α) (s : Finset α) : insert a (insert b s) = insert b (insert a s) := ext fun x => by simp only [mem_insert, or_left_comm] +@[deprecated (since := "2024-11-29")] alias Insert.comm := insert_comm + @[norm_cast] theorem coe_pair {a b : α} : (({a, b} : Finset α) : Set α) = {a, b} := by ext @@ -424,7 +430,7 @@ theorem coe_eq_pair {s : Finset α} {a b : α} : (s : Set α) = {a, b} ↔ s = { rw [← coe_pair, coe_inj] theorem pair_comm (a b : α) : ({a, b} : Finset α) = {b, a} := - Insert.comm a b ∅ + insert_comm a b ∅ theorem insert_idem (a : α) (s : Finset α) : insert a (insert a s) = insert a s := ext fun x => by simp only [mem_insert, ← or_assoc, or_self_iff] diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index a9cc52080c652..97a141e5d4d22 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -634,6 +634,22 @@ protected theorem sup_lt_iff (ha : ⊥ < a) : s.sup f < a ↔ ∀ b ∈ s, f b < Finset.cons_induction_on s (fun _ => ha) fun c t hc => by simpa only [sup_cons, sup_lt_iff, mem_cons, forall_eq_or_imp] using And.imp_right⟩ +theorem sup_mem_of_nonempty (hs : s.Nonempty) : s.sup f ∈ f '' s := by + classical + induction s using Finset.induction with + | empty => exfalso; simp only [Finset.not_nonempty_empty] at hs + | @insert a s _ h => + rw [Finset.sup_insert (b := a) (s := s) (f := f)] + by_cases hs : s = ∅ + · simp [hs] + · rw [← ne_eq, ← Finset.nonempty_iff_ne_empty] at hs + simp only [Finset.coe_insert] + rcases le_total (f a) (s.sup f) with (ha | ha) + · rw [sup_eq_right.mpr ha] + exact Set.image_mono (Set.subset_insert a s) (h hs) + · rw [sup_eq_left.mpr ha] + apply Set.mem_image_of_mem _ (Set.mem_insert a ↑s) + end OrderBot section OrderTop diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index efd119c83d79e..6ba35e3f0ddf1 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -1427,11 +1427,10 @@ theorem sum_smul_index_addMonoidHom [AddMonoid M] [AddCommMonoid N] [DistribSMul {b : R} {h : α → M →+ N} : ((b • g).sum fun a => h a) = g.sum fun i c => h i (b • c) := sum_mapRange_index fun i => (h i).map_zero -instance noZeroSMulDivisors [Semiring R] [AddCommMonoid M] [Module R M] {ι : Type*} +instance noZeroSMulDivisors [Zero R] [Zero M] [SMulZeroClass R M] {ι : Type*} [NoZeroSMulDivisors R M] : NoZeroSMulDivisors R (ι →₀ M) := - ⟨fun h => - or_iff_not_imp_left.mpr fun hc => - Finsupp.ext fun i => (smul_eq_zero.mp (DFunLike.ext_iff.mp h i)).resolve_left hc⟩ + ⟨fun h => or_iff_not_imp_left.mpr fun hc => Finsupp.ext fun i => + (eq_zero_or_eq_zero_of_smul_eq_zero (DFunLike.ext_iff.mp h i)).resolve_left hc⟩ section DistribMulActionSemiHom variable [Monoid R] [AddMonoid M] [AddMonoid N] [DistribMulAction R M] [DistribMulAction R N] diff --git a/Mathlib/Data/Finsupp/Lex.lean b/Mathlib/Data/Finsupp/Lex.lean index 6b2a45dd8483b..87e567ec57773 100644 --- a/Mathlib/Data/Finsupp/Lex.lean +++ b/Mathlib/Data/Finsupp/Lex.lean @@ -75,6 +75,26 @@ instance Lex.linearOrder [LinearOrder N] : LinearOrder (Lex (α →₀ N)) where le := (· ≤ ·) __ := LinearOrder.lift' (toLex ∘ toDFinsupp ∘ ofLex) finsuppEquivDFinsupp.injective +theorem Lex.single_strictAnti : StrictAnti (fun (a : α) ↦ toLex (single a 1)) := by + intro a b h + simp only [LT.lt, Finsupp.lex_def] + simp only [ofLex_toLex, Nat.lt_eq] + use a + constructor + · intro d hd + simp only [Finsupp.single_eq_of_ne (ne_of_lt hd).symm, + Finsupp.single_eq_of_ne (ne_of_gt (lt_trans hd h))] + · simp only [single_eq_same, single_eq_of_ne (ne_of_lt h).symm, zero_lt_one] + +theorem Lex.single_lt_iff {a b : α} : toLex (single b 1) < toLex (single a 1) ↔ a < b := + Lex.single_strictAnti.lt_iff_lt + +theorem Lex.single_le_iff {a b : α} : toLex (single b 1) ≤ toLex (single a 1) ↔ a ≤ b := + Lex.single_strictAnti.le_iff_le + +theorem Lex.single_antitone : Antitone (fun (a : α) ↦ toLex (single a 1)) := + Lex.single_strictAnti.antitone + variable [PartialOrder N] theorem toLex_monotone : Monotone (@toLex (α →₀ N)) := diff --git a/Mathlib/Data/Finsupp/MonomialOrder.lean b/Mathlib/Data/Finsupp/MonomialOrder.lean index 431fcb7d3ce34..788df94797c35 100644 --- a/Mathlib/Data/Finsupp/MonomialOrder.lean +++ b/Mathlib/Data/Finsupp/MonomialOrder.lean @@ -135,7 +135,6 @@ example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 1) := by example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 2) := by use 0; simp - variable {σ : Type*} [LinearOrder σ] /-- The lexicographic order on `σ →₀ ℕ`, as a `MonomialOrder` -/ diff --git a/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean new file mode 100644 index 0000000000000..0deee8bfefe7c --- /dev/null +++ b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean @@ -0,0 +1,258 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ + +import Mathlib.Data.Finsupp.MonomialOrder +import Mathlib.Data.Finsupp.Weight +import Mathlib.Logic.Equiv.TransferInstance + +/-! Homogeneous lexicographic monomial ordering + +* `MonomialOrder.degLex`: a variant of the lexicographic ordering that first compares degrees. +For this, `σ` needs to be embedded with an ordering relation which satisfies `WellFoundedGT σ`. +(This last property is automatic when `σ` is finite). + +The type synonym is `DegLex (σ →₀ ℕ)` and the two lemmas `MonomialOrder.degLex_le_iff` +and `MonomialOrder.degLex_lt_iff` rewrite the ordering as comparisons in the type `Lex (σ →₀ ℕ)`. + +## References + +* [Cox, Little and O'Shea, *Ideals, varieties, and algorithms*][coxlittleoshea1997] +* [Becker and Weispfenning, *Gröbner bases*][Becker-Weispfenning1993] + +-/ + +section degLex + +/-- A type synonym to equip a type with its lexicographic order sorted by degrees. -/ +def DegLex (α : Type*) := α + +variable {α : Type*} + +/-- `toDegLex` is the identity function to the `DegLex` of a type. -/ +@[match_pattern] def toDegLex : α ≃ DegLex α := Equiv.refl _ + +theorem toDegLex_injective : Function.Injective (toDegLex (α := α)) := fun _ _ ↦ _root_.id + +theorem toDegLex_inj {a b : α} : toDegLex a = toDegLex b ↔ a = b := Iff.rfl + +/-- `ofDegLex` is the identity function from the `DegLex` of a type. -/ +@[match_pattern] def ofDegLex : DegLex α ≃ α := Equiv.refl _ + +theorem ofDegLex_injective : Function.Injective (ofDegLex (α := α)) := fun _ _ ↦ _root_.id + +theorem ofDegLex_inj {a b : DegLex α} : ofDegLex a = ofDegLex b ↔ a = b := Iff.rfl + +@[simp] theorem ofDegLex_symm_eq : (@ofDegLex α).symm = toDegLex := rfl + +@[simp] theorem toDegLex_symm_eq : (@toDegLex α).symm = ofDegLex := rfl + +@[simp] theorem ofDegLex_toDegLex (a : α) : ofDegLex (toDegLex a) = a := rfl + +@[simp] theorem toDegLex_ofDegLex (a : DegLex α) : toDegLex (ofDegLex a) = a := rfl + +/-- A recursor for `DegLex`. Use as `induction x`. -/ +@[elab_as_elim, induction_eliminator, cases_eliminator] +protected def DegLex.rec {β : DegLex α → Sort*} (h : ∀ a, β (toDegLex a)) : + ∀ a, β a := fun a => h (ofDegLex a) + +@[simp] lemma DegLex.forall_iff {p : DegLex α → Prop} : (∀ a, p a) ↔ ∀ a, p (toDegLex a) := Iff.rfl +@[simp] lemma DegLex.exists_iff {p : DegLex α → Prop} : (∃ a, p a) ↔ ∃ a, p (toDegLex a) := Iff.rfl + +noncomputable instance [AddCommMonoid α] : + AddCommMonoid (DegLex α) := ofDegLex.addCommMonoid + +theorem toDegLex_add [AddCommMonoid α] (a b : α) : + toDegLex (a + b) = toDegLex a + toDegLex b := rfl + +theorem ofDegLex_add [AddCommMonoid α] (a b : DegLex α) : + ofDegLex (a + b) = ofDegLex a + ofDegLex b := rfl + +namespace Finsupp + +/-- `Finsupp.DegLex r s` is the homogeneous lexicographic order on `α →₀ M`, +where `α` is ordered by `r` and `M` is ordered by `s`. +The type synonym `DegLex (α →₀ M)` has an order given by `Finsupp.DegLex (· < ·) (· < ·)`. -/ +protected def DegLex (r : α → α → Prop) (s : ℕ → ℕ → Prop) : + (α →₀ ℕ) → (α →₀ ℕ) → Prop := + (Prod.Lex s (Finsupp.Lex r s)) on (fun x ↦ (x.degree, x)) + +theorem degLex_def {r : α → α → Prop} {s : ℕ → ℕ → Prop} {a b : α →₀ ℕ} : + Finsupp.DegLex r s a b ↔ Prod.Lex s (Finsupp.Lex r s) (a.degree, a) (b.degree, b) := + Iff.rfl + +namespace DegLex + +theorem wellFounded + {r : α → α → Prop} [IsTrichotomous α r] (hr : WellFounded (Function.swap r)) + {s : ℕ → ℕ → Prop} (hs : WellFounded s) (hs0 : ∀ ⦃n⦄, ¬ s n 0) : + WellFounded (Finsupp.DegLex r s) := by + have wft := WellFounded.prod_lex hs (Finsupp.Lex.wellFounded' hs0 hs hr) + rw [← Set.wellFoundedOn_univ] at wft + unfold Finsupp.DegLex + rw [← Set.wellFoundedOn_range] + exact Set.WellFoundedOn.mono wft (le_refl _) (fun _ _ ↦ trivial) + +instance [LT α] : LT (DegLex (α →₀ ℕ)) := + ⟨fun f g ↦ Finsupp.DegLex (· < ·) (· < ·) (ofDegLex f) (ofDegLex g)⟩ + +theorem lt_def [LT α] {a b : DegLex (α →₀ ℕ)} : + a < b ↔ (toLex ((ofDegLex a).degree, toLex (ofDegLex a))) < + (toLex ((ofDegLex b).degree, toLex (ofDegLex b))) := + Iff.rfl + +theorem lt_iff [LT α] {a b : DegLex (α →₀ ℕ)} : + a < b ↔ (ofDegLex a).degree < (ofDegLex b).degree ∨ + (((ofDegLex a).degree = (ofDegLex b).degree) ∧ toLex (ofDegLex a) < toLex (ofDegLex b)) := by + simp only [lt_def, Prod.Lex.lt_iff] + +variable [LinearOrder α] + +instance isStrictOrder : IsStrictOrder (DegLex (α →₀ ℕ)) (· < ·) where + irrefl := fun a ↦ by simp [lt_def] + trans := by + intro a b c hab hbc + simp only [lt_iff] at hab hbc ⊢ + rcases hab with (hab | hab) + · rcases hbc with (hbc | hbc) + · left; exact lt_trans hab hbc + · left; exact lt_of_lt_of_eq hab hbc.1 + · rcases hbc with (hbc | hbc) + · left; exact lt_of_eq_of_lt hab.1 hbc + · right; exact ⟨Eq.trans hab.1 hbc.1, lt_trans hab.2 hbc.2⟩ + +/-- The linear order on `Finsupp`s obtained by the homogeneous lexicographic ordering. -/ +instance : LinearOrder (DegLex (α →₀ ℕ)) := + LinearOrder.lift' + (fun (f : DegLex (α →₀ ℕ)) ↦ toLex ((ofDegLex f).degree, toLex (ofDegLex f))) + (fun f g ↦ by simp) + +theorem le_iff {x y : DegLex (α →₀ ℕ)} : + x ≤ y ↔ (ofDegLex x).degree < (ofDegLex y).degree ∨ + (ofDegLex x).degree = (ofDegLex y).degree ∧ toLex (ofDegLex x) ≤ toLex (ofDegLex y) := by + simp only [le_iff_eq_or_lt, lt_iff, EmbeddingLike.apply_eq_iff_eq] + by_cases h : x = y + · simp [h] + · by_cases k : (ofDegLex x).degree < (ofDegLex y).degree + · simp [k] + · simp only [h, k, false_or] + +noncomputable instance : OrderedCancelAddCommMonoid (DegLex (α →₀ ℕ)) where + le_of_add_le_add_left a b c h := by + rw [le_iff] at h ⊢ + simpa only [ofDegLex_add, degree_add, add_lt_add_iff_left, add_right_inj, toLex_add, + add_le_add_iff_left] using h + add_le_add_left a b h c := by + rw [le_iff] at h ⊢ + simpa [ofDegLex_add, degree_add] using h + +/-- The linear order on `Finsupp`s obtained by the homogeneous lexicographic ordering. -/ +noncomputable instance : + LinearOrderedCancelAddCommMonoid (DegLex (α →₀ ℕ)) where + le_total := instLinearOrderDegLexNat.le_total + decidableLE := instLinearOrderDegLexNat.decidableLE + compare_eq_compareOfLessAndEq := instLinearOrderDegLexNat.compare_eq_compareOfLessAndEq + +theorem single_strictAnti : StrictAnti (fun (a : α) ↦ toDegLex (single a 1)) := by + intro _ _ h + simp only [lt_iff, ofDegLex_toDegLex, degree_single, lt_self_iff_false, Lex.single_lt_iff, h, + and_self, or_true] + +theorem single_antitone : Antitone (fun (a : α) ↦ toDegLex (single a 1)) := + single_strictAnti.antitone + +theorem single_lt_iff {a b : α} : + toDegLex (Finsupp.single b 1) < toDegLex (Finsupp.single a 1) ↔ a < b := + single_strictAnti.lt_iff_lt + +theorem single_le_iff {a b : α} : + toDegLex (Finsupp.single b 1) ≤ toDegLex (Finsupp.single a 1) ↔ a ≤ b := + single_strictAnti.le_iff_le + +theorem monotone_degree : + Monotone (fun (x : DegLex (α →₀ ℕ)) ↦ (ofDegLex x).degree) := by + intro x y + rw [le_iff] + rintro (h | h) + · apply le_of_lt h + · apply le_of_eq h.1 + +instance orderBot : OrderBot (DegLex (α →₀ ℕ)) where + bot := toDegLex (0 : α →₀ ℕ) + bot_le x := by + simp only [le_iff, ofDegLex_toDegLex, toLex_zero, degree_zero] + rcases eq_zero_or_pos (ofDegLex x).degree with (h | h) + · simp only [h, lt_self_iff_false, true_and, false_or, ge_iff_le] + exact bot_le + · simp [h] + +instance wellFoundedLT [WellFoundedGT α] : + WellFoundedLT (DegLex (α →₀ ℕ)) := + ⟨wellFounded wellFounded_gt wellFounded_lt fun n ↦ (zero_le n).not_lt⟩ + +end DegLex + +end Finsupp + +namespace MonomialOrder + +open Finsupp + +variable {σ : Type*} [LinearOrder σ] [WellFoundedGT σ] + +/-- The deg-lexicographic order on `σ →₀ ℕ`, as a `MonomialOrder` -/ +noncomputable def degLex : + MonomialOrder σ where + syn := DegLex (σ →₀ ℕ) + toSyn := { toEquiv := toDegLex, map_add' := toDegLex_add } + toSyn_monotone a b h := by + simp only [AddEquiv.coe_mk, DegLex.le_iff, ofDegLex_toDegLex] + by_cases ha : a.degree < b.degree + · exact Or.inl ha + · refine Or.inr ⟨le_antisymm ?_ (not_lt.mp ha), toLex_monotone h⟩ + rw [← add_tsub_cancel_of_le h, degree_add] + exact Nat.le_add_right a.degree (b - a).degree + +theorem degLex_le_iff {a b : σ →₀ ℕ} : + a ≼[degLex] b ↔ toDegLex a ≤ toDegLex b := + Iff.rfl + +theorem degLex_lt_iff {a b : σ →₀ ℕ} : + a ≺[degLex] b ↔ toDegLex a < toDegLex b := + Iff.rfl + +theorem degLex_single_le_iff {a b : σ} : + single a 1 ≼[degLex] single b 1 ↔ b ≤ a := by + rw [MonomialOrder.degLex_le_iff, DegLex.single_le_iff] + +theorem degLex_single_lt_iff {a b : σ} : + single a 1 ≺[degLex] single b 1 ↔ b < a := by + rw [MonomialOrder.degLex_lt_iff, DegLex.single_lt_iff] + +end MonomialOrder + +section Examples + +open Finsupp MonomialOrder DegLex + +/-- for the deg-lexicographic ordering, X 1 < X 0 -/ +example : single (1 : Fin 2) 1 ≺[degLex] single 0 1 := by + rw [degLex_lt_iff, single_lt_iff] + exact Nat.one_pos + +/-- for the deg-lexicographic ordering, X 0 * X 1 < X 0 ^ 2 -/ +example : (single 0 1 + single 1 1) ≺[degLex] single (0 : Fin 2) 2 := by + simp only [degLex_lt_iff, lt_iff, ofDegLex_toDegLex, degree_add, + degree_single, Nat.reduceAdd, lt_self_iff_false, true_and, false_or] + use 0 + simp + +/-- for the deg-lexicographic ordering, X 0 < X 1 ^ 2 -/ +example : single (0 : Fin 2) 1 ≺[degLex] single 1 2 := by + simp [degLex_lt_iff, lt_iff] + +end Examples + +end degLex diff --git a/Mathlib/Data/Finsupp/Order.lean b/Mathlib/Data/Finsupp/Order.lean index b06f454769ac2..ad9bfbd4b5485 100644 --- a/Mathlib/Data/Finsupp/Order.lean +++ b/Mathlib/Data/Finsupp/Order.lean @@ -322,6 +322,10 @@ theorem add_sub_single_one {a : ι} {u u' : ι →₀ ℕ} (h : u' a ≠ 0) : u + (u' - single a 1) = u + u' - single a 1 := (add_tsub_assoc_of_le (single_le_iff.mpr <| Nat.one_le_iff_ne_zero.mpr h) _).symm +lemma sub_add_single_one_cancel {u : ι →₀ ℕ} {i : ι} (h : u i ≠ 0) : + u - single i 1 + single i 1 = u := by + rw [sub_single_one_add h, add_tsub_cancel_right] + end Nat end Finsupp diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 2278beedc1aec..36ec2f64e9d94 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -105,6 +105,13 @@ theorem NonTorsionWeight.ne_zero [NonTorsionWeight w] (s : σ) : apply Nat.zero_ne_one.symm exact NonTorsionWeight.eq_zero_of_smul_eq_zero h +variable {w} in +lemma weight_sub_single_add {f : σ →₀ ℕ} {i : σ} (hi : f i ≠ 0) : + (f - single i 1).weight w + w i = f.weight w := by + conv_rhs => rw [← sub_add_single_one_cancel hi, weight_apply] + rw [sum_add_index', sum_single_index, one_smul, weight_apply] + exacts [zero_smul .., fun _ ↦ zero_smul .., fun _ _ _ ↦ add_smul ..] + end AddCommMonoid section OrderedAddCommMonoid @@ -195,6 +202,20 @@ def degree (d : σ →₀ ℕ) := ∑ i ∈ d.support, d i @[deprecated degree (since := "2024-07-20")] alias _root_.MvPolynomial.degree := degree +@[simp] +theorem degree_add (a b : σ →₀ ℕ) : (a + b).degree = a.degree + b.degree := + sum_add_index' (h := fun _ ↦ id) (congrFun rfl) fun _ _ ↦ congrFun rfl + +@[simp] +theorem degree_single (a : σ) (m : ℕ) : (Finsupp.single a m).degree = m := by + rw [degree, Finset.sum_eq_single a] + · simp only [single_eq_same] + · intro b _ hba + exact single_eq_of_ne hba.symm + · intro ha + simp only [mem_support_iff, single_eq_same, ne_eq, Decidable.not_not] at ha + rw [single_eq_same, ha] + lemma degree_eq_zero_iff (d : σ →₀ ℕ) : degree d = 0 ↔ d = 0 := by simp only [degree, Finset.sum_eq_zero_iff, Finsupp.mem_support_iff, ne_eq, Decidable.not_imp_self, DFunLike.ext_iff, Finsupp.coe_zero, Pi.zero_apply] diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index e7ab67fafc45c..3e4f9127bad67 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -399,8 +399,6 @@ theorem shiftRight_add' : ∀ (m : ℤ) (n k : ℕ), m >>> (n + k : ℤ) = (m >> /-! ### bitwise ops -/ -attribute [local simp] Int.zero_div - theorem shiftLeft_add : ∀ (m : ℤ) (n : ℕ) (k : ℤ), m <<< (n + k) = (m <<< (n : ℤ)) <<< k | (m : ℕ), n, (k : ℕ) => congr_arg ofNat (by simp [Nat.shiftLeft_eq, Nat.pow_add, mul_assoc]) diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index 42c4167321de5..4676d73593af0 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -112,9 +112,6 @@ theorem toNat_of_nonpos : ∀ {z : ℤ}, z ≤ 0 → z.toNat = 0 This lemma is orphaned from `Data.Int.Bitwise` as it also requires material from `Data.Int.Order`. -/ - -attribute [local simp] Int.zero_div - @[simp] theorem div2_bit (b n) : div2 (bit b n) = n := by rw [bit_val, div2_val, add_comm, Int.add_mul_ediv_left, (_ : (_ / 2 : ℤ) = 0), zero_add] diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 5b5ca26520e79..6b8c7c19967e6 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -831,17 +831,13 @@ theorem get_reverse (l : List α) (i : Nat) (h1 h2) : dsimp omega -set_option linter.deprecated false - theorem get_reverse' (l : List α) (n) (hn') : l.reverse.get n = l.get ⟨l.length - 1 - n, hn'⟩ := by - rw [eq_comm] - convert get_reverse l.reverse n (by simpa) n.2 using 1 simp theorem eq_cons_of_length_one {l : List α} (h : l.length = 1) : l = [l.get ⟨0, by omega⟩] := by refine ext_get (by convert h) fun n h₁ h₂ => ?_ - simp only [get_singleton] + simp congr omega diff --git a/Mathlib/Data/List/Destutter.lean b/Mathlib/Data/List/Destutter.lean index df3a1ad3b5778..f021d2425e7a0 100644 --- a/Mathlib/Data/List/Destutter.lean +++ b/Mathlib/Data/List/Destutter.lean @@ -25,8 +25,11 @@ Note that we make no guarantees of being the longest sublist with this property; adjacent, chain, duplicates, remove, list, stutter, destutter -/ +open Function -variable {α : Type*} (l : List α) (R : α → α → Prop) [DecidableRel R] {a b : α} +variable {α β : Type*} (l l₁ l₂ : List α) (R : α → α → Prop) [DecidableRel R] {a b : α} + +variable {R₂ : β → β → Prop} [DecidableRel R₂] namespace List @@ -148,4 +151,116 @@ theorem destutter_eq_nil : ∀ {l : List α}, destutter R l = [] ↔ l = [] | [] => Iff.rfl | _ :: l => ⟨fun h => absurd h <| l.destutter'_ne_nil R, fun h => nomatch h⟩ +variable {R} + +/-- For a relation-preserving map, `destutter` commutes with `map`. -/ +theorem map_destutter {f : α → β} : ∀ {l : List α}, (∀ a ∈ l, ∀ b ∈ l, R a b ↔ R₂ (f a) (f b)) → + (l.destutter R).map f = (l.map f).destutter R₂ + | [], hl => by simp + | [a], hl => by simp + | a :: b :: l, hl => by + have := hl a (by simp) b (by simp) + simp_rw [map_cons, destutter_cons_cons, ← this] + by_cases hr : R a b <;> + simp [hr, ← destutter_cons', map_destutter fun c hc d hd ↦ hl _ (cons_subset_cons _ + (subset_cons_self _ _) hc) _ (cons_subset_cons _ (subset_cons_self _ _) hd), + map_destutter fun c hc d hd ↦ hl _ (subset_cons_self _ _ hc) _ (subset_cons_self _ _ hd)] + +/-- For a injective function `f`, `destutter' (·≠·)` commutes with `map f`. -/ +theorem map_destutter_ne {f : α → β} (h : Injective f) [DecidableEq α] [DecidableEq β] : + (l.destutter (·≠·)).map f = (l.map f).destutter (·≠·) := + map_destutter fun _ _ _ _ ↦ h.ne_iff.symm + +/-- `destutter'` on a relation like ≠ or <, whose negation is transitive, has length monotone +under a `¬R` changing of the first element. -/ +theorem length_destutter'_cotrans_ge [i : IsTrans α Rᶜ] : + ∀ {a} {l : List α}, ¬R b a → (l.destutter' R b).length ≤ (l.destutter' R a).length + | a, [], hba => by simp + | a, c :: l, hba => by + by_cases hbc : R b c + case pos => + have hac : ¬Rᶜ a c := (mt (_root_.trans hba)) (not_not.2 hbc) + simp_rw [destutter', if_pos (not_not.1 hac), if_pos hbc, length_cons, le_refl] + case neg => + simp only [destutter', if_neg hbc] + by_cases hac : R a c + case pos => + simp only [if_pos hac, length_cons] + exact Nat.le_succ_of_le (length_destutter'_cotrans_ge hbc) + case neg => + simp only [if_neg hac] + exact length_destutter'_cotrans_ge hba + +/-- `List.destutter'` on a relation like `≠`, whose negation is an equivalence, gives the same +length if the first elements are not related. -/ +theorem length_destutter'_congr [IsEquiv α Rᶜ] (hab : ¬R a b) : + (l.destutter' R a).length = (l.destutter' R b).length := + (length_destutter'_cotrans_ge hab).antisymm <| length_destutter'_cotrans_ge (symm hab : Rᶜ b a) + +/-- `List.destutter'` on a relation like ≠, whose negation is an equivalence, has length + monotonic under List.cons -/ +/- +TODO: Replace this lemma by the more general version: +theorem Sublist.length_destutter'_mono [IsEquiv α Rᶜ] (h : a :: l₁ <+ b :: l₂) : + (List.destutter' R a l₁).length ≤ (List.destutter' R b l₂).length +-/ +theorem le_length_destutter'_cons [IsEquiv α Rᶜ] : + ∀ {l : List α}, (l.destutter' R b).length ≤ ((b :: l).destutter' R a).length + | [] => by by_cases hab : (R a b) <;> simp_all [Nat.le_succ] + | c :: cs => by + by_cases hab : R a b + case pos => simp [destutter', if_pos hab, Nat.le_succ] + obtain hac | hac : R a c ∨ Rᶜ a c := em _ + · have hbc : ¬Rᶜ b c := mt (_root_.trans hab) (not_not.2 hac) + simp [destutter', if_pos hac, if_pos (not_not.1 hbc), if_neg hab] + · have hbc : ¬R b c := trans (symm hab) hac + simp only [destutter', if_neg hbc, if_neg hac, if_neg hab] + exact (length_destutter'_congr cs hab).ge + +/-- `List.destutter` on a relation like ≠, whose negation is an equivalence, has length +monotone under List.cons -/ +theorem length_destutter_le_length_destutter_cons [IsEquiv α Rᶜ] : + ∀ {l : List α}, (l.destutter R).length ≤ ((a :: l).destutter R).length + | [] => by simp [destutter] + | b :: l => le_length_destutter'_cons + +variable {l l₁ l₂} + +/-- `destutter ≠` has length monotone under `List.cons`. -/ +theorem length_destutter_ne_le_length_destutter_cons [DecidableEq α] : + (l.destutter (· ≠ ·)).length ≤ ((a :: l).destutter (· ≠ ·)).length := + length_destutter_le_length_destutter_cons + +/-- `destutter` of relations like `≠`, whose negation is an equivalence relation, +gives a list of maximal length over any chain. + +In other words, `l.destutter R` is an `R`-chain sublist of `l`, and is at least as long as any other +`R`-chain sublist. -/ +lemma Chain'.length_le_length_destutter [IsEquiv α Rᶜ] : + ∀ {l₁ l₂ : List α}, l₁ <+ l₂ → l₁.Chain' R → l₁.length ≤ (l₂.destutter R).length + -- `l₁ := []`, `l₂ := []` + | [], [], _, _ => by simp + -- `l₁ := l₁`, `l₂ := a :: l₂` + | l₁, _, .cons (l₂ := l₂) a hl, hl₁ => + (hl₁.length_le_length_destutter hl).trans length_destutter_le_length_destutter_cons + -- `l₁ := [a]`, `l₂ := a :: l₂` + | _, _, .cons₂ (l₁ := []) (l₂ := l₁) a hl, hl₁ => by simp [Nat.one_le_iff_ne_zero] + -- `l₁ := a :: l₁`, `l₂ := a :: b :: l₂` + | _, _, .cons₂ a <| .cons (l₁ := l₁) (l₂ := l₂) b hl, hl₁ => by + by_cases hab : R a b + · simpa [destutter_cons_cons, hab] using hl₁.tail.length_le_length_destutter (hl.cons _) + · simpa [destutter_cons_cons, hab] using hl₁.length_le_length_destutter (hl.cons₂ _) + -- `l₁ := a :: b :: l₁`, `l₂ := a :: b :: l₂` + | _, _, .cons₂ a <| .cons₂ (l₁ := l₁) (l₂ := l₂) b hl, hl₁ => by + simpa [destutter_cons_cons, rel_of_chain_cons hl₁] + using hl₁.tail.length_le_length_destutter (hl.cons₂ _) + +/-- `destutter` of `≠` gives a list of maximal length over any chain. + +In other words, `l.destutter (· ≠ ·)` is a `≠`-chain sublist of `l`, and is at least as long as any +other `≠`-chain sublist. -/ +lemma Chain'.length_le_length_destutter_ne [DecidableEq α] (hl : l₁ <+ l₂) + (hl₁ : l₁.Chain' (· ≠ ·)) : l₁.length ≤ (l₂.destutter (· ≠ ·)).length := + hl₁.length_le_length_destutter hl + end List diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 5aaf7611ee1bd..9e93d425fe171 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -49,6 +49,11 @@ theorem getD_eq_default {n : ℕ} (hn : l.length ≤ n) : l.getD n d = d := by · simp at hn · exact ih (Nat.le_of_succ_le_succ hn) +theorem getD_reverse {l : List α} (i) (h : i < length l) : + getD l.reverse i = getD l (l.length - 1 - i) := by + funext a + rwa [List.getD_eq_getElem?_getD, List.getElem?_reverse, ← List.getD_eq_getElem?_getD] + /-- An empty list can always be decidably checked for the presence of an element. Not an instance because it would clash with `DecidableEq α`. -/ def decidableGetDNilNe (a : α) : DecidablePred fun i : ℕ => getD ([] : List α) i a ≠ a := @@ -69,6 +74,11 @@ theorem getElem?_getD_replicate_default_eq (r n : ℕ) : (replicate r d)[n]?.get @[deprecated (since := "2024-06-12")] alias getD_replicate_default_eq := getElem?_getD_replicate_default_eq +theorem getD_replicate {y i n} (h : i < n) : + getD (replicate n x) i y = x := by + rw [getD_eq_getElem, getElem_replicate] + rwa [length_replicate] + theorem getD_append (l l' : List α) (d : α) (n : ℕ) (h : n < l.length) : (l ++ l').getD n d = l.getD n d := by rw [getD_eq_getElem _ _ (Nat.lt_of_lt_of_le h (length_append _ _ ▸ Nat.le_add_right _ _)), diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index a0866034e24e9..49f6dd007cd72 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -3,14 +3,13 @@ Copyright (c) 2014 Parikshit Khanna. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro -/ -import Mathlib.Data.List.Flatten +import Mathlib.Data.List.Lemmas import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.List.Count import Mathlib.Data.List.Duplicate import Mathlib.Data.List.InsertIdx import Batteries.Data.List.Perm import Mathlib.Data.List.Perm.Basic -import Mathlib.Data.List.Lemmas /-! # Permutations of a list @@ -182,23 +181,19 @@ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α) simp only [mem_permutationsAux2', ← this, or_comm, and_left_comm, mem_append, mem_flatMap, append_assoc, cons_append, exists_prop] -set_option linter.deprecated false in theorem length_foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : length (foldr (fun y r => (permutationsAux2 t ts r y id).2) r L) = - Nat.sum (map length L) + length r := by - simp [foldr_permutationsAux2, Function.comp_def, length_permutationsAux2, length_flatMap', - Nat.sum_eq_listSum] + (map length L).sum + length r := by + simp [foldr_permutationsAux2, Function.comp_def, length_permutationsAux2, length_flatMap] -set_option linter.deprecated false in theorem length_foldr_permutationsAux2' (t : α) (ts : List α) (r L : List (List α)) (n) (H : ∀ l ∈ L, length l = n) : length (foldr (fun y r => (permutationsAux2 t ts r y id).2) r L) = n * length L + length r := by - rw [length_foldr_permutationsAux2, (_ : Nat.sum (map length L) = n * length L)] + rw [length_foldr_permutationsAux2, (_ : (map length L).sum = n * length L)] induction' L with l L ih · simp - have sum_map : Nat.sum (map length L) = n * length L := ih fun l m => H l (mem_cons_of_mem _ m) + have sum_map : (map length L).sum = n * length L := ih fun l m => H l (mem_cons_of_mem _ m) have length_l : length l = n := H _ (mem_cons_self _ _) - simp only [Nat.sum_eq_listSum] at sum_map simp [sum_map, length_l, Nat.mul_add, Nat.add_comm, mul_succ] @[simp] diff --git a/Mathlib/Data/Nat/Choose/Factorization.lean b/Mathlib/Data/Nat/Choose/Factorization.lean index 6bd3a1dac00f3..e639e46d275bf 100644 --- a/Mathlib/Data/Nat/Choose/Factorization.lean +++ b/Mathlib/Data/Nat/Choose/Factorization.lean @@ -39,10 +39,10 @@ theorem factorization_choose_le_log : (choose n k).factorization p ≤ log p n : refine le_of_not_lt fun hnk => h ?_ simp [choose_eq_zero_of_lt hnk] rw [factorization_def _ hp, @padicValNat_def _ ⟨hp⟩ _ (choose_pos hkn)] - rw [← Nat.cast_le (α := ℕ∞), ← multiplicity.Finite.emultiplicity_eq_multiplicity] + rw [← Nat.cast_le (α := ℕ∞), ← FiniteMultiplicity.emultiplicity_eq_multiplicity] · simp only [hp.emultiplicity_choose hkn (lt_add_one _), Nat.cast_le] exact (Finset.card_filter_le _ _).trans (le_of_eq (Nat.card_Ico _ _)) - apply Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, choose_pos hkn⟩ + apply Nat.finiteMultiplicity_iff.2 ⟨hp.ne_one, choose_pos hkn⟩ /-- A `pow` form of `Nat.factorization_choose_le` -/ theorem pow_factorization_choose_le (hn : 0 < n) : p ^ (choose n k).factorization p ≤ n := diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index fc518c6dbbc1c..d4a9697c1a97e 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -438,12 +438,23 @@ protected lemma div_le_div_right (h : a ≤ b) : a / c ≤ b / c := lemma lt_of_div_lt_div (h : a / c < b / c) : a < b := Nat.lt_of_not_le fun hab ↦ Nat.not_le_of_lt h <| Nat.div_le_div_right hab -protected lemma div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := - Nat.pos_of_ne_zero fun h ↦ Nat.lt_irrefl a <| - calc - a = a % b := by simpa [h] using (mod_add_div a b).symm - _ < b := mod_lt a hb - _ ≤ a := hba +@[simp] protected lemma div_eq_zero_iff : a / b = 0 ↔ b = 0 ∨ a < b where + mp h := by + rw [← mod_add_div a b, h, Nat.mul_zero, Nat.add_zero, or_iff_not_imp_left] + exact mod_lt _ ∘ Nat.pos_iff_ne_zero.2 + mpr := by + obtain rfl | hb := eq_or_ne b 0 + · simp + simp only [hb, false_or] + rw [← Nat.mul_right_inj hb, ← Nat.add_left_cancel_iff, mod_add_div] + simp +contextual [mod_eq_of_lt] + +protected lemma div_ne_zero_iff : a / b ≠ 0 ↔ b ≠ 0 ∧ b ≤ a := by simp + +@[simp] protected lemma div_pos_iff : 0 < a / b ↔ 0 < b ∧ b ≤ a := by + simp [Nat.pos_iff_ne_zero] + +protected lemma div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := Nat.div_pos_iff.2 ⟨hb, hba⟩ lemma lt_mul_of_div_lt (h : a / c < b) (hc : 0 < c) : a < b * c := Nat.lt_of_not_ge <| Nat.not_le_of_gt h ∘ (Nat.le_div_iff_mul_le hc).2 @@ -993,17 +1004,6 @@ lemma div_eq_iff_eq_of_dvd_dvd (hn : n ≠ 0) (ha : a ∣ n) (hb : b ∣ n) : n exact Nat.eq_mul_of_div_eq_right ha h · rw [h] -protected lemma div_eq_zero_iff (hb : 0 < b) : a / b = 0 ↔ a < b where - mp h := by rw [← mod_add_div a b, h, Nat.mul_zero, Nat.add_zero]; exact mod_lt _ hb - mpr h := by rw [← Nat.mul_right_inj (Nat.ne_of_gt hb), ← Nat.add_left_cancel_iff, mod_add_div, - mod_eq_of_lt h, Nat.mul_zero, Nat.add_zero] - -protected lemma div_ne_zero_iff (hb : b ≠ 0) : a / b ≠ 0 ↔ b ≤ a := by - rw [ne_eq, Nat.div_eq_zero_iff (Nat.pos_of_ne_zero hb), not_lt] - -protected lemma div_pos_iff (hb : b ≠ 0) : 0 < a / b ↔ b ≤ a := by - rw [Nat.pos_iff_ne_zero, Nat.div_ne_zero_iff hb] - lemma le_iff_ne_zero_of_dvd (ha : a ≠ 0) (hab : a ∣ b) : a ≤ b ↔ b ≠ 0 where mp := by rw [← Nat.pos_iff_ne_zero] at ha ⊢; exact Nat.lt_of_lt_of_le ha mpr hb := Nat.le_of_dvd (Nat.pos_iff_ne_zero.2 hb) hab @@ -1266,7 +1266,7 @@ lemma dvd_mul_of_div_dvd (h : b ∣ a) (hdiv : a / b ∣ c) : a ∣ b * c := by /-- If a small natural number is divisible by a larger natural number, the small number is zero. -/ lemma eq_zero_of_dvd_of_lt (w : a ∣ b) (h : b < a) : b = 0 := - Nat.eq_zero_of_dvd_of_div_eq_zero w ((Nat.div_eq_zero_iff (lt_of_le_of_lt (zero_le b) h)).mpr h) + Nat.eq_zero_of_dvd_of_div_eq_zero w (by simp [h]) lemma le_of_lt_add_of_dvd (h : a < b + n) : n ∣ a → n ∣ b → a ≤ b := by rintro ⟨a, rfl⟩ ⟨b, rfl⟩ diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index f9385aca645c1..5adc574882628 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -304,8 +304,8 @@ theorem digits_len (b n : ℕ) (hb : 1 < b) (hn : n ≠ 0) : (b.digits n).length induction' n using Nat.strong_induction_on with n IH rw [digits_eq_cons_digits_div hb hn, List.length] by_cases h : n / b = 0 - · have hb0 : b ≠ 0 := (Nat.succ_le_iff.1 hb).ne_bot - simp [h, log_eq_zero_iff, ← Nat.div_eq_zero_iff hb0.bot_lt] + · simp [IH, h] + aesop · have : n / b < n := div_lt_self (Nat.pos_of_ne_zero hn) hb rw [IH _ this h, log_div_base, tsub_add_cancel_of_le] refine Nat.succ_le_of_lt (log_pos hb ?_) @@ -821,8 +821,7 @@ theorem digits_one (b n) (n0 : 0 < n) (nb : n < b) : Nat.digits b n = [n] ∧ 1 have b2 : 1 < b := lt_iff_add_one_le.mpr (le_trans (add_le_add_right (lt_iff_add_one_le.mp n0) 1) nb) refine ⟨?_, b2, n0⟩ - rw [Nat.digits_def' b2 n0, Nat.mod_eq_of_lt nb, - (Nat.div_eq_zero_iff ((zero_le n).trans_lt nb)).2 nb, Nat.digits_zero] + rw [Nat.digits_def' b2 n0, Nat.mod_eq_of_lt nb, Nat.div_eq_zero_iff.2 <| .inr nb, Nat.digits_zero] /- Porting note: this part of the file is tactic related. diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 15659840ed9b9..a56075aac773f 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -270,7 +270,7 @@ and `n'` such that `n'` is not divisible by `p` and `n = p^e * n'`. -/ theorem exists_eq_pow_mul_and_not_dvd {n : ℕ} (hn : n ≠ 0) (p : ℕ) (hp : p ≠ 1) : ∃ e n' : ℕ, ¬p ∣ n' ∧ n = p ^ e * n' := let ⟨a', h₁, h₂⟩ := - (Nat.multiplicity_finite_iff.mpr ⟨hp, Nat.pos_of_ne_zero hn⟩).exists_eq_pow_mul_and_not_dvd + (Nat.finiteMultiplicity_iff.mpr ⟨hp, Nat.pos_of_ne_zero hn⟩).exists_eq_pow_mul_and_not_dvd ⟨_, a', h₂, h₁⟩ /-- Any nonzero natural number is the product of an odd part `m` and a power of @@ -284,7 +284,7 @@ theorem dvd_iff_div_factorization_eq_tsub {d n : ℕ} (hd : d ≠ 0) (hdn : d d ∣ n ↔ (n / d).factorization = n.factorization - d.factorization := by refine ⟨factorization_div, ?_⟩ rcases eq_or_lt_of_le hdn with (rfl | hd_lt_n); · simp - have h1 : n / d ≠ 0 := fun H => Nat.lt_asymm hd_lt_n ((Nat.div_eq_zero_iff hd.bot_lt).mp H) + have h1 : n / d ≠ 0 := by simp [*] intro h rw [dvd_iff_le_div_mul n d] by_contra h2 diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index 5d4c83169c971..67bf2864fa197 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -37,7 +37,7 @@ def primeFactorsList : ℕ → List ℕ | k + 2 => let m := minFac (k + 2) m :: primeFactorsList ((k + 2) / m) -decreasing_by show (k + 2) / m < (k + 2); exact factors_lemma +decreasing_by exact factors_lemma @[deprecated (since := "2024-06-14")] alias factors := primeFactorsList diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index 77e18051deac1..ece38a314c41c 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -46,7 +46,7 @@ Legendre, p-adic -/ -open Finset Nat multiplicity +open Finset Nat open Nat @@ -57,7 +57,7 @@ divides `n`. This set is expressed by filtering `Ico 1 b` where `b` is any bound `log m n`. -/ theorem emultiplicity_eq_card_pow_dvd {m n b : ℕ} (hm : m ≠ 1) (hn : 0 < n) (hb : log m n < b) : emultiplicity m n = #{i ∈ Ico 1 b | m ^ i ∣ n} := - have fin := Nat.multiplicity_finite_iff.2 ⟨hm, hn⟩ + have fin := Nat.finiteMultiplicity_iff.2 ⟨hm, hn⟩ calc emultiplicity m n = #(Ico 1 <| multiplicity m n + 1) := by simp [fin.emultiplicity_eq_multiplicity] @@ -91,7 +91,7 @@ theorem emultiplicity_pow {p m n : ℕ} (hp : p.Prime) : _root_.emultiplicity_pow hp.prime theorem emultiplicity_self {p : ℕ} (hp : p.Prime) : emultiplicity p p = 1 := - (Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, hp.pos⟩).emultiplicity_self + (Nat.finiteMultiplicity_iff.2 ⟨hp.ne_one, hp.pos⟩).emultiplicity_self theorem emultiplicity_pow_self {p n : ℕ} (hp : p.Prime) : emultiplicity p (p ^ n) = n := _root_.emultiplicity_pow_self hp.ne_zero hp.prime.not_unit n @@ -136,7 +136,7 @@ theorem emultiplicity_factorial_mul_succ {n p : ℕ} (hp : p.Prime) : have h2 : p * n + 1 ≤ p * (n + 1) := by linarith have h3 : p * n + 1 ≤ p * (n + 1) + 1 := by omega have hm : emultiplicity p (p * n)! ≠ ⊤ := by - rw [Ne, emultiplicity_eq_top, Classical.not_not, Nat.multiplicity_finite_iff] + rw [Ne, emultiplicity_eq_top, Classical.not_not, Nat.finiteMultiplicity_iff] exact ⟨hp.ne_one, factorial_pos _⟩ revert hm have h4 : ∀ m ∈ Ico (p * n + 1) (p * (n + 1)), emultiplicity p m = 0 := by @@ -201,8 +201,8 @@ theorem emultiplicity_choose' {p n k b : ℕ} (hp : p.Prime) (hnb : log p (n + k (add_comm n k ▸ hnb)), multiplicity_choose_aux hp (le_add_left k n)] simp [add_comm] refine (WithTop.add_right_cancel_iff ?_).1 h₁ - apply finite_iff_emultiplicity_ne_top.1 - exact Nat.multiplicity_finite_iff.2 ⟨hp.ne_one, mul_pos (factorial_pos k) (factorial_pos n)⟩ + apply finiteMultiplicity_iff_emultiplicity_ne_top.1 + exact Nat.finiteMultiplicity_iff.2 ⟨hp.ne_one, mul_pos (factorial_pos k) (factorial_pos n)⟩ /-- The multiplicity of `p` in `choose n k` is the number of carries when `k` and `n - k` are added in base `p`. The set is expressed by filtering `Ico 1 b` where `b` @@ -251,8 +251,8 @@ theorem emultiplicity_choose_prime_pow {p n k : ℕ} (hp : p.Prime) (hkn : k ≤ emultiplicity p (choose (p ^ n) k) = ↑(n - multiplicity p k) := by push_cast rw [← emultiplicity_choose_prime_pow_add_emultiplicity hp hkn hk0, - (multiplicity_finite_iff.2 ⟨hp.ne_one, Nat.pos_of_ne_zero hk0⟩).emultiplicity_eq_multiplicity, - (multiplicity_finite_iff.2 ⟨hp.ne_one, choose_pos hkn⟩).emultiplicity_eq_multiplicity] + (finiteMultiplicity_iff.2 ⟨hp.ne_one, Nat.pos_of_ne_zero hk0⟩).emultiplicity_eq_multiplicity, + (finiteMultiplicity_iff.2 ⟨hp.ne_one, choose_pos hkn⟩).emultiplicity_eq_multiplicity] norm_cast rw [Nat.add_sub_cancel_right] diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 53bb8bb29fec5..b1463cae82ce5 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -50,7 +50,7 @@ theorem _root_.Squarefree.natFactorization_le_one {n : ℕ} (p : ℕ) (hn : Squa n.factorization p ≤ 1 := by rcases eq_or_ne n 0 with (rfl | hn') · simp - rw [multiplicity.squarefree_iff_emultiplicity_le_one] at hn + rw [squarefree_iff_emultiplicity_le_one] at hn by_cases hp : p.Prime · have := hn p rw [← multiplicity_eq_factorization hp hn'] @@ -387,8 +387,7 @@ lemma primeFactors_prod (hs : ∀ p ∈ s, p.Prime) : primeFactors (∏ p ∈ s, lemma primeFactors_div_gcd (hm : Squarefree m) (hn : n ≠ 0) : primeFactors (m / m.gcd n) = primeFactors m \ primeFactors n := by ext p - have : m / m.gcd n ≠ 0 := - (Nat.div_ne_zero_iff <| gcd_ne_zero_right hn).2 <| gcd_le_left _ hm.ne_zero.bot_lt + have : m / m.gcd n ≠ 0 := by simp [gcd_ne_zero_right hn, gcd_le_left _ hm.ne_zero.bot_lt] simp only [mem_primeFactors, ne_eq, this, not_false_eq_true, and_true, not_and, mem_sdiff, hm.ne_zero, hn, dvd_div_iff_mul_dvd (gcd_dvd_left _ _)] refine ⟨fun hp ↦ ⟨⟨hp.1, dvd_of_mul_left_dvd hp.2⟩, fun _ hpn ↦ hp.1.not_unit <| hm _ <| diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 9508faa6a1ad6..ccad963fe4394 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -79,8 +79,8 @@ protected def hrecOn₂ (qa : Quot ra) (qb : Quot rb) (f : ∀ a b, φ ⟦a⟧ /-- Map a function `f : α → β` such that `ra x y` implies `rb (f x) (f y)` to a map `Quot ra → Quot rb`. -/ -protected def map (f : α → β) (h : (ra ⇒ rb) f f) : Quot ra → Quot rb := - (Quot.lift fun x ↦ ⟦f x⟧) fun x y (h₁ : ra x y) ↦ Quot.sound <| h h₁ +protected def map (f : α → β) (h : ∀ ⦃a b : α⦄, ra a b → rb (f a) (f b)) : Quot ra → Quot rb := + Quot.lift (fun x => Quot.mk rb (f x)) fun _ _ hra ↦ Quot.sound <| h hra /-- If `ra` is a subrelation of `ra'`, then we have a natural map `Quot ra → Quot ra'`. -/ protected def mapRight {ra' : α → α → Prop} (h : ∀ a₁ a₂, ra a₁ a₂ → ra' a₁ a₂) : @@ -227,11 +227,11 @@ protected def hrecOn₂ (qa : Quotient sa) (qb : Quotient sb) (f : ∀ a b, φ /-- Map a function `f : α → β` that sends equivalent elements to equivalent elements to a function `Quotient sa → Quotient sb`. Useful to define unary operations on quotients. -/ -protected def map (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) : Quotient sa → Quotient sb := +protected def map (f : α → β) (h : ∀ ⦃a b : α⦄, a ≈ b → f a ≈ f b) : Quotient sa → Quotient sb := Quot.map f h @[simp] -theorem map_mk (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) : +theorem map_mk (f : α → β) (h) (x : α) : Quotient.map f h (⟦x⟧ : Quotient sa) = (⟦f x⟧ : Quotient sb) := rfl @@ -240,12 +240,13 @@ variable {γ : Sort*} {sc : Setoid γ} /-- Map a function `f : α → β → γ` that sends equivalent elements to equivalent elements to a function `f : Quotient sa → Quotient sb → Quotient sc`. Useful to define binary operations on quotients. -/ -protected def map₂ (f : α → β → γ) (h : ((· ≈ ·) ⇒ (· ≈ ·) ⇒ (· ≈ ·)) f f) : +protected def map₂ (f : α → β → γ) + (h : ∀ ⦃a₁ a₂⦄, a₁ ≈ a₂ → ∀ ⦃b₁ b₂⦄, b₁ ≈ b₂ → f a₁ b₁ ≈ f a₂ b₂) : Quotient sa → Quotient sb → Quotient sc := Quotient.lift₂ (fun x y ↦ ⟦f x y⟧) fun _ _ _ _ h₁ h₂ ↦ Quot.sound <| h h₁ h₂ @[simp] -theorem map₂_mk (f : α → β → γ) (h : ((· ≈ ·) ⇒ (· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) (y : β) : +theorem map₂_mk (f : α → β → γ) (h) (x : α) (y : β) : Quotient.map₂ f h (⟦x⟧ : Quotient sa) (⟦y⟧ : Quotient sb) = (⟦f x y⟧ : Quotient sc) := rfl @@ -693,7 +694,8 @@ theorem hrecOn₂'_mk'' {φ : Quotient s₁ → Quotient s₂ → Sort*} /-- Map a function `f : α → β` that sends equivalent elements to equivalent elements to a function `Quotient sa → Quotient sb`. Useful to define unary operations on quotients. -/ -protected def map' (f : α → β) (h : (s₁.r ⇒ s₂.r) f f) : Quotient s₁ → Quotient s₂ := +protected def map' (f : α → β) (h : ∀ a b, s₁.r a b → s₂.r (f a) (f b)) : + Quotient s₁ → Quotient s₂ := Quot.map f h @[simp] @@ -702,15 +704,7 @@ theorem map'_mk'' (f : α → β) (h) (x : α) : rfl /-- A version of `Quotient.map₂` using curly braces and unification. -/ -protected def map₂' (f : α → β → γ) (h : (s₁.r ⇒ s₂.r ⇒ s₃.r) f f) : - Quotient s₁ → Quotient s₂ → Quotient s₃ := - Quotient.map₂ f h - -@[simp] -theorem map₂'_mk'' (f : α → β → γ) (h) (x : α) : - (Quotient.mk'' x : Quotient s₁).map₂' f h = - (Quotient.map' (f x) (h (Setoid.refl x)) : Quotient s₂ → Quotient s₃) := - rfl +@[deprecated (since := "2024-12-01")] protected alias map₂' := Quotient.map₂ theorem exact' {a b : α} : (Quotient.mk'' a : Quotient s₁) = Quotient.mk'' b → s₁ a b := diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 235dc1f4d78d5..12f99f0fcb687 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -25,7 +25,7 @@ but this only works if you `unseal Nat.sqrt.iter in` before the theorem where yo -/ -open Rat Real multiplicity +open Rat Real /-- A real number is irrational if it is not equal to any rational number. -/ def Irrational (x : ℝ) := @@ -79,7 +79,7 @@ theorem irrational_nrt_of_n_not_dvd_multiplicity {x : ℝ} (n : ℕ) {m : ℤ} ( rw [← Int.cast_pow, Int.cast_inj] at hxr subst m have : y ≠ 0 := by rintro rfl; rw [zero_pow hnpos.ne'] at hm; exact hm rfl - rw [(Int.multiplicity_finite_iff.2 ⟨by simp [hp.1.ne_one], this⟩).multiplicity_pow + rw [(Int.finiteMultiplicity_iff.2 ⟨by simp [hp.1.ne_one], this⟩).multiplicity_pow (Nat.prime_iff_prime_int.1 hp.1), Nat.mul_mod_right] at hv exact hv rfl diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 41676f753e553..b77cd5d86cdec 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -924,10 +924,12 @@ theorem sUnion_powerset_gc : gc_sSup_Iic /-- `⋃₀` and `𝒫` form a Galois insertion. -/ -def sUnion_powerset_gi : +def sUnionPowersetGI : GaloisInsertion (⋃₀ · : Set (Set α) → Set α) (𝒫 · : Set α → Set (Set α)) := gi_sSup_Iic +@[deprecated (since := "2024-12-07")] alias sUnion_powerset_gi := sUnionPowersetGI + /-- If all sets in a collection are either `∅` or `Set.univ`, then so is their union. -/ theorem sUnion_mem_empty_univ {S : Set (Set α)} (h : S ⊆ {∅, univ}) : ⋃₀ S ∈ ({∅, univ} : Set (Set α)) := by diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index bbaa5b239fbfe..d860f84f98bd6 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -133,7 +133,7 @@ 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 + toFun := fun (x, y) ↦ Quotient.map₂ Prod.mk (fun _ _ hx _ _ hy ↦ ⟨hx, hy⟩) x y invFun := fun q ↦ Quotient.liftOn' q (fun xy ↦ (Quotient.mk'' xy.1, Quotient.mk'' xy.2)) fun x y hxy ↦ Prod.ext (by simpa using hxy.1) (by simpa using hxy.2) left_inv := fun q ↦ by diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 43e41f0704f39..37dffbc1303eb 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -1575,6 +1575,6 @@ def Nat.residueClassesEquiv (N : ℕ) [NeZero N] : ℕ ≃ ZMod N × ℕ where · simp only [add_comm p.1.val, cast_add, cast_mul, natCast_self, zero_mul, natCast_val, cast_id', id_eq, zero_add] · simp only [add_comm p.1.val, mul_add_div (NeZero.pos _), - (Nat.div_eq_zero_iff <| (NeZero.pos _)).2 p.1.val_lt, add_zero] + (Nat.div_eq_zero_iff).2 <| .inr p.1.val_lt, add_zero] set_option linter.style.longFile 1700 diff --git a/Mathlib/Deprecated/LazyList.lean b/Mathlib/Deprecated/LazyList.lean index 5ba81d38999a8..15de4bb396935 100644 --- a/Mathlib/Deprecated/LazyList.lean +++ b/Mathlib/Deprecated/LazyList.lean @@ -102,7 +102,7 @@ instance : LawfulMonad LazyList := LawfulMonad.mk' (bind_pure_comp := by intro _ _ f xs simp only [bind, Functor.map, pure, singleton] - induction xs using LazyList.traverse.induct (m := @Id) (f := f) with + induction xs using LazyList.traverse.induct with | case1 => simp only [Thunk.pure, LazyList.bind, LazyList.traverse, Id.pure_eq] | case2 _ _ ih => diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index 745637e6ec963..3842d160c48ac 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -600,6 +600,7 @@ theorem normalClosure_subset_iff {s t : Set G} (ht : IsNormalSubgroup t) : s ⊆ t ↔ normalClosure s ⊆ t := ⟨normalClosure_subset ht, Set.Subset.trans subset_normalClosure⟩ +@[gcongr] theorem normalClosure_mono {s t : Set G} : s ⊆ t → normalClosure s ⊆ normalClosure t := fun h => normalClosure_subset normalClosure.is_normal (Set.Subset.trans h subset_normalClosure) diff --git a/Mathlib/FieldTheory/AlgebraicClosure.lean b/Mathlib/FieldTheory/AlgebraicClosure.lean index 9a4a0902bc605..056881d31cde2 100644 --- a/Mathlib/FieldTheory/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/AlgebraicClosure.lean @@ -38,6 +38,9 @@ def algebraicClosure : IntermediateField F E := variable {F E} +theorem algebraicClosure_toSubalgebra : + (algebraicClosure F E).toSubalgebra = integralClosure F E := rfl + /-- An element is contained in the algebraic closure of `F` in `E` if and only if it is an integral element. -/ theorem mem_algebraicClosure_iff' {x : E} : @@ -100,8 +103,7 @@ variable (F E K) /-- The algebraic closure of `F` in `E` is algebraic over `F`. -/ instance isAlgebraic : Algebra.IsAlgebraic F (algebraicClosure F E) := - ⟨fun x ↦ - isAlgebraic_iff.mpr (IsAlgebraic.isIntegral (mem_algebraicClosure_iff.mp x.2)).isAlgebraic⟩ + ⟨fun x ↦ isAlgebraic_iff.mpr x.2.isAlgebraic⟩ /-- The algebraic closure of `F` in `E` is the integral closure of `F` in `E`. -/ instance isIntegralClosure : IsIntegralClosure (algebraicClosure F E) F E := @@ -109,6 +111,10 @@ instance isIntegralClosure : IsIntegralClosure (algebraicClosure F E) F E := end algebraicClosure +protected theorem Transcendental.algebraicClosure {a : E} (ha : Transcendental F a) : + Transcendental (algebraicClosure F E) a := + ha.extendScalars Subtype.val_injective + variable (F E K) /-- An intermediate field of `E / F` is contained in the algebraic closure of `F` in `E` diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index 5cfd3ecd0a861..a7bc963365950 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -160,7 +160,11 @@ theorem realize_genericPolyMapSurjOnOfInjOn Set.MapsTo, Set.mem_def, injOnAlt, funext_iff, Set.SurjOn, Set.image, setOf, Set.subset_def, Equiv.forall_congr_left (mvPolynomialSupportLEEquiv mons)] simp +singlePass only [← Sum.elim_comp_inl_inr] - simp [Set.mem_def, Function.comp_def] + -- was `simp` and very slow (https://github.com/leanprover-community/mathlib4/issues/19751) + simp only [Function.comp_def, Sum.elim_inl, Sum.elim_inr, Fin.castAdd_zero, Fin.cast_eq_self, + Nat.add_zero, Term.realize_var, Term.realize_relabel, realize_termOfFreeCommRing, + lift_genericPolyMap, Nat.reduceAdd, Fin.isValue, Function.uncurry_apply_pair, Fin.cons_zero, + Fin.cons_one, ↓reduceIte, one_ne_zero] theorem ACF_models_genericPolyMapSurjOnOfInjOn_of_prime [Fintype ι] {p : ℕ} (hp : p.Prime) (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index 7488ef0514690..39db5e7df739e 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -184,7 +184,7 @@ See `PerfectField.separable_iff_squarefree` for the converse when the coefficien field. -/ theorem Separable.squarefree {p : R[X]} (hsep : Separable p) : Squarefree p := by classical - rw [multiplicity.squarefree_iff_emultiplicity_le_one p] + rw [squarefree_iff_emultiplicity_le_one p] exact fun f => or_iff_not_imp_right.mpr fun hunit => emultiplicity_le_one_of_separable hunit hsep end CommSemiring @@ -282,7 +282,7 @@ theorem rootMultiplicity_le_one_of_separable [Nontrivial R] {p : R[X]} (hsep : S by_cases hp : p = 0 · simp [hp] rw [rootMultiplicity_eq_multiplicity, if_neg hp, ← Nat.cast_le (α := ℕ∞), - Nat.cast_one, ← (multiplicity_X_sub_C_finite x hp).emultiplicity_eq_multiplicity] + Nat.cast_one, ← (finiteMultiplicity_X_sub_C x hp).emultiplicity_eq_multiplicity] apply emultiplicity_le_one_of_separable (not_isUnit_X_sub_C _) hsep end CommRing diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 1b85e1f399fbf..8cc00db3d0c2b 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -592,7 +592,7 @@ theorem eq_X_pow_char_pow_sub_C_pow_of_natSepDegree_eq_one (q : ℕ) [ExpChar F have hD := (h ▸ natSepDegree_le_of_dvd p f hf hm.ne_zero).antisymm <| Nat.pos_of_ne_zero <| (natSepDegree_ne_zero_iff _).2 hI.natDegree_pos.ne' obtain ⟨n, y, H, hp⟩ := hM.eq_X_pow_char_pow_sub_C_of_natSepDegree_eq_one_of_irreducible q hI hD - have hF := multiplicity_finite_of_degree_pos_of_monic (degree_pos_of_irreducible hI) hM hm.ne_zero + have hF := finiteMultiplicity_of_degree_pos_of_monic (degree_pos_of_irreducible hI) hM hm.ne_zero classical have hne := (multiplicity_pos_of_dvd hf).ne' refine ⟨_, n, y, hne, H, ?_⟩ diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index cae9a80c72889..2997b5c92169b 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -211,7 +211,7 @@ instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] : SmoothInv₀ 𝓘(𝕜) exact contDiffAt_inv 𝕜 hx variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalSpace H] {E : Type*} - [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) {G : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {I : ModelWithCorners 𝕜 E H} {G : Type*} [TopologicalSpace G] [ChartedSpace H G] [Inv G] [Zero G] [SmoothInv₀ I G] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] @@ -227,23 +227,23 @@ include I in This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ theorem hasContinuousInv₀_of_hasSmoothInv₀ : HasContinuousInv₀ G := - { continuousAt_inv₀ := fun _ hx ↦ (contMDiffAt_inv₀ I hx).continuousAt } + { continuousAt_inv₀ := fun _ hx ↦ (contMDiffAt_inv₀ (I := I) hx).continuousAt } theorem contMDiffOn_inv₀ : ContMDiffOn I I ⊤ (Inv.inv : G → G) {0}ᶜ := fun _x hx => - (contMDiffAt_inv₀ I hx).contMDiffWithinAt + (contMDiffAt_inv₀ hx).contMDiffWithinAt @[deprecated (since := "2024-11-21")] alias smoothOn_inv₀ := contMDiffOn_inv₀ @[deprecated (since := "2024-11-21")] alias SmoothOn_inv₀ := contMDiffOn_inv₀ -variable {I} {s : Set M} {a : M} +variable {s : Set M} {a : M} theorem ContMDiffWithinAt.inv₀ (hf : ContMDiffWithinAt I' I n f s a) (ha : f a ≠ 0) : ContMDiffWithinAt I' I n (fun x => (f x)⁻¹) s a := - ((contMDiffAt_inv₀ I ha).of_le le_top).comp_contMDiffWithinAt a hf + ((contMDiffAt_inv₀ ha).of_le le_top).comp_contMDiffWithinAt a hf theorem ContMDiffAt.inv₀ (hf : ContMDiffAt I' I n f a) (ha : f a ≠ 0) : ContMDiffAt I' I n (fun x ↦ (f x)⁻¹) a := - ((contMDiffAt_inv₀ I ha).of_le le_top).comp a hf + ((contMDiffAt_inv₀ ha).of_le le_top).comp a hf theorem ContMDiff.inv₀ (hf : ContMDiff I' I n f) (h0 : ∀ x, f x ≠ 0) : ContMDiff I' I n (fun x ↦ (f x)⁻¹) := diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 93c0039fd5c2f..2c3d9ef063c40 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -1051,6 +1051,10 @@ theorem StructureGroupoid.mem_maximalAtlas_of_mem_groupoid {f : PartialHomeomorp rintro e (rfl : e = PartialHomeomorph.refl H) exact ⟨G.trans (G.symm hf) G.id_mem, G.trans (G.symm G.id_mem) hf⟩ +theorem StructureGroupoid.maximalAtlas_mono {G G' : StructureGroupoid H} (h : G ≤ G') : + G.maximalAtlas M ⊆ G'.maximalAtlas M := + fun _ he e' he' ↦ ⟨h (he e' he').1, h (he e' he').2⟩ + end MaximalAtlas section Singleton diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index 112fab89cc177..3ab6587a89afc 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -479,6 +479,21 @@ theorem contMDiffOn_iff : convert hdiff x (f x) (extChartAt I x x) (by simp only [hx, mfld_simps]) using 1 mfld_set_tac +/-- zero-smoothness on a set is equivalent to continuity on this set. -/ +theorem contMDiffOn_zero_iff : + ContMDiffOn I I' 0 f s ↔ ContinuousOn f s := by + rw [contMDiffOn_iff] + refine ⟨fun h ↦ h.1, fun h ↦ ⟨h, ?_⟩⟩ + intro x y + rw [show ((0 : ℕ∞) : WithTop ℕ∞) = 0 by rfl, contDiffOn_zero] + apply (continuousOn_extChartAt _).comp + · apply h.comp ((continuousOn_extChartAt_symm _).mono inter_subset_left) (fun z hz ↦ ?_) + simp only [preimage_inter, mem_inter_iff, mem_preimage] at hz + exact hz.2.1 + · intro z hz + simp only [preimage_inter, mem_inter_iff, mem_preimage] at hz + exact hz.2.2 + /-- One can reformulate smoothness on a set as continuity on this set, and smoothness in any extended chart in the target. -/ theorem contMDiffOn_iff_target : @@ -525,6 +540,11 @@ theorem contMDiff_iff_target : @[deprecated (since := "2024-11-20")] alias smooth_iff_target := contMDiff_iff_target +/-- zero-smoothness is equivalent to continuity. -/ +theorem contMDiff_zero_iff : + ContMDiff I I' 0 f ↔ Continuous f := by + rw [← contMDiffOn_univ, continuous_iff_continuousOn_univ, contMDiffOn_zero_iff] + end SmoothManifoldWithCorners /-! ### Deducing smoothness from smoothness one step beyond -/ diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 6a6dd316cf875..ced591a0a2f5f 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -79,9 +79,7 @@ scoped[Manifold] notation M " ≃ₘ^" n:1000 "⟮" I ", " J "⟯ " N => Diffeom scoped[Manifold] notation M " ≃ₘ⟮" I ", " J "⟯ " N => Diffeomorph I J M N ⊤ /-- `n`-times continuously differentiable diffeomorphism between `E` and `E'`. -/ -scoped[Manifold] - notation E " ≃ₘ^" n:1000 "[" 𝕜 "] " E' => - Diffeomorph (modelWithCornersSelf 𝕜 E) (modelWithCornersSelf 𝕜 E') E E' n +scoped[Manifold] notation E " ≃ₘ^" n:1000 "[" 𝕜 "] " E' => Diffeomorph 𝓘(𝕜, E) 𝓘(𝕜, E') E E' n /-- Infinitely differentiable diffeomorphism between `E` and `E'`. -/ scoped[Manifold] diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 8e667f22b2b24..7f10c097d74fe 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -14,16 +14,18 @@ or with boundary or with corners. As a concrete example, we construct explicitly boundary structure on the real interval `[x, y]`. More specifically, we introduce -* `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space +* `modelWithCornersEuclideanHalfSpace n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space used to define `n`-dimensional real manifolds with boundary -* `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used +* `modelWithCornersEuclideanQuadrant n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used to define `n`-dimensional real manifolds with corners ## Notations In the locale `Manifold`, we introduce the notations * `𝓡 n` for the identity model with corners on `EuclideanSpace ℝ (Fin n)` -* `𝓡∂ n` for `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)`. +* `𝓡∂ n` for `modelWithCornersEuclideanHalfSpace n`. For instance, if a manifold `M` is boundaryless, smooth and modelled on `EuclideanSpace ℝ (Fin m)`, and `N` is smooth with boundary modelled on `EuclideanHalfSpace n`, and `f : M → N` is a smooth diff --git a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean index b30e07d620aa0..2e727c4ef5022 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/NormedSpace.lean @@ -79,7 +79,7 @@ theorem Differentiable.comp_mdifferentiable {g : F → F'} {f : M → F} (hg : D end Module -/-! ### Linear maps between normed spaces are smooth -/ +/-! ### Linear maps between normed spaces are differentiable -/ #adaptation_note /-- @@ -189,8 +189,8 @@ theorem MDifferentiable.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F (hg : MDifferentiable I 𝓘(𝕜, F₁ →L[𝕜] F₃) g) (hf : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₁) f) : MDifferentiable I 𝓘(𝕜, F₂ →L[𝕜] F₃) fun x => (g x).comp (f x) := fun x => (hg x).clm_comp (hf x) -/-- Applying a linear map to a vector is smooth within a set. Version in vector spaces. For a -version in nontrivial vector bundles, see `MDifferentiableWithinAt.clm_apply_of_inCoordinates`. -/ +/-- Applying a linear map to a vector is differentiable within a set. Version in vector spaces. For +a version in nontrivial vector bundles, see `MDifferentiableWithinAt.clm_apply_of_inCoordinates`. -/ theorem MDifferentiableWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} {x : M} (hg : MDifferentiableWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) g s x) (hf : MDifferentiableWithinAt I 𝓘(𝕜, F₁) f s x) : @@ -201,7 +201,7 @@ theorem MDifferentiableWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : exact differentiable_fst.clm_apply differentiable_snd) (hg.prod_mk_space hf) (by simp_rw [mapsTo_univ]) -/-- Applying a linear map to a vector is smooth. Version in vector spaces. For a +/-- Applying a linear map to a vector is differentiable. Version in vector spaces. For a version in nontrivial vector bundles, see `MDifferentiableAt.clm_apply_of_inCoordinates`. -/ theorem MDifferentiableAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {x : M} (hg : MDifferentiableAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) g x) (hf : MDifferentiableAt I 𝓘(𝕜, F₁) f x) : @@ -282,7 +282,7 @@ theorem MDifferentiable.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → MDifferentiable I 𝓘(𝕜, F₁ × F₂ →L[𝕜] F₃ × F₄) fun x => (g x).prodMap (f x) := fun x => (hg x).clm_prodMap (hf x) -/-! ### Smoothness of scalar multiplication -/ +/-! ### Differentiability of scalar multiplication -/ variable {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean index 7afe6dc78bd73..e3df71c1fd846 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Tangent.lean @@ -62,7 +62,7 @@ lemma mfderiv_chartAt_eq_tangentCoordChange {x y : M} (hsrc : x ∈ (chartAt H y the basis also has unique differential. -/ theorem UniqueMDiffOn.tangentBundle_proj_preimage {s : Set M} (hs : UniqueMDiffOn I s) : UniqueMDiffOn I.tangent (π E (TangentSpace I) ⁻¹' s) := - hs.smooth_bundle_preimage _ + hs.bundle_preimage _ /-- To write a linear map between tangent spaces in coordinates amounts to precomposing and postcomposing it with derivatives of extended charts. diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index e090b101ab9e7..187e4cca99ec5 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -115,36 +115,78 @@ end open Bundle variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type*} - [TopologicalSpace (TotalSpace F Z)] [∀ b, TopologicalSpace (Z b)] [∀ b, AddCommMonoid (Z b)] - [∀ b, Module 𝕜 (Z b)] [FiberBundle F Z] [VectorBundle 𝕜 F Z] [SmoothVectorBundle F Z I] + [TopologicalSpace (TotalSpace F Z)] [∀ b, TopologicalSpace (Z b)] [FiberBundle F Z] -theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : - e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := - ⟨e.contMDiffOn.mdifferentiableOn le_top, e.contMDiffOn_symm.mdifferentiableOn le_top⟩ - -theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} +private lemma UniqueMDiffWithinAt.bundle_preimage_aux {p : TotalSpace F Z} + (hs : UniqueMDiffWithinAt I s p.proj) (h's : s ⊆ (trivializationAt F Z p.proj).baseSet) : + UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) p := by + suffices ((extChartAt I p.proj).symm ⁻¹' s ∩ range I) ×ˢ univ ⊆ + (extChartAt (I.prod 𝓘(𝕜, F)) p).symm ⁻¹' (TotalSpace.proj ⁻¹' s) ∩ range (I.prod 𝓘(𝕜, F)) by + let w := (extChartAt (I.prod 𝓘(𝕜, F)) p p).2 + have A : extChartAt (I.prod 𝓘(𝕜, F)) p p = (extChartAt I p.1 p.1, w) := by + ext + · simp [FiberBundle.chartedSpace_chartAt] + · rfl + simp only [UniqueMDiffWithinAt, A] at hs ⊢ + exact (hs.prod (uniqueDiffWithinAt_univ (x := w))).mono this + rcases p with ⟨x, v⟩ + dsimp + rintro ⟨z, w⟩ ⟨hz, -⟩ + simp only [ModelWithCorners.target_eq, mem_inter_iff, mem_preimage, Function.comp_apply, + mem_range] at hz + simp only [FiberBundle.chartedSpace_chartAt, PartialHomeomorph.coe_trans_symm, mem_inter_iff, + mem_preimage, Function.comp_apply, mem_range] + constructor + · rw [PartialEquiv.prod_symm, PartialEquiv.refl_symm, PartialEquiv.prod_coe, + ModelWithCorners.toPartialEquiv_coe_symm, PartialEquiv.refl_coe, + PartialHomeomorph.prod_symm, PartialHomeomorph.refl_symm, PartialHomeomorph.prod_apply, + PartialHomeomorph.refl_apply] + convert hz.1 + apply Trivialization.proj_symm_apply' + exact h's hz.1 + · rcases hz.2 with ⟨u, rfl⟩ + exact ⟨(u, w), rfl⟩ + +/-- In a fiber bundle, the preimage under the projection of a set with unique differentials +in the base has unique differentials in the bundle. -/ +theorem UniqueMDiffWithinAt.bundle_preimage {p : TotalSpace F Z} (hs : UniqueMDiffWithinAt I s p.proj) : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) p := by - set e := trivializationAt F Z p.proj - have hp : p ∈ e.source := FiberBundle.mem_trivializationAt_proj_source - have : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (s ×ˢ univ) (e p) := by - rw [← Prod.mk.eta (p := e p), FiberBundle.trivializationAt_proj_fst] - exact hs.prod (uniqueMDiffWithinAt_univ _) - rw [← e.left_inv hp] - refine (this.preimage_partialHomeomorph e.mdifferentiable.symm (e.map_source hp)).mono ?_ - rintro y ⟨hy, hys, -⟩ - rwa [PartialHomeomorph.symm_symm, e.coe_coe, e.coe_fst hy] at hys + suffices UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) + (π F Z ⁻¹' (s ∩ (trivializationAt F Z p.proj).baseSet)) p from this.mono (by simp) + apply UniqueMDiffWithinAt.bundle_preimage_aux (hs.inter _) inter_subset_right + exact IsOpen.mem_nhds (trivializationAt F Z p.proj).open_baseSet + (FiberBundle.mem_baseSet_trivializationAt' p.proj) + +@[deprecated (since := "2024-12-02")] +alias UniqueMDiffWithinAt.smooth_bundle_preimage := UniqueMDiffWithinAt.bundle_preimage variable (Z) -theorem UniqueMDiffWithinAt.smooth_bundle_preimage' {b : M} (hs : UniqueMDiffWithinAt I s b) +/-- In a fiber bundle, the preimage under the projection of a set with unique differentials +in the base has unique differentials in the bundle. Version with a point `⟨b, x⟩`. -/ +theorem UniqueMDiffWithinAt.bundle_preimage' {b : M} (hs : UniqueMDiffWithinAt I s b) (x : Z b) : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) ⟨b, x⟩ := - hs.smooth_bundle_preimage (p := ⟨b, x⟩) + hs.bundle_preimage (p := ⟨b, x⟩) + +@[deprecated (since := "2024-12-02")] +alias UniqueMDiffWithinAt.smooth_bundle_preimage' := UniqueMDiffWithinAt.bundle_preimage' -/-- In a smooth fiber bundle, the preimage under the projection of a set with -unique differential in the basis also has unique differential. -/ -theorem UniqueMDiffOn.smooth_bundle_preimage (hs : UniqueMDiffOn I s) : +/-- In a fiber bundle, the preimage under the projection of a set with unique differentials +in the base has unique differentials in the bundle. -/ +theorem UniqueMDiffOn.bundle_preimage (hs : UniqueMDiffOn I s) : UniqueMDiffOn (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) := fun _p hp ↦ - (hs _ hp).smooth_bundle_preimage + (hs _ hp).bundle_preimage + +@[deprecated (since := "2024-12-02")] +alias UniqueMDiffOn.smooth_bundle_preimage := UniqueMDiffOn.bundle_preimage + +/- TODO: move me to `Mathlib.Geometry.Manifold.VectorBundle.MDifferentiable` once #19636 is in. -/ +variable [∀ b, AddCommMonoid (Z b)] [∀ b, Module 𝕜 (Z b)] [VectorBundle 𝕜 F Z] + +theorem Trivialization.mdifferentiable [SmoothVectorBundle F Z I] + (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : + e.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := + ⟨e.contMDiffOn.mdifferentiableOn le_top, e.contMDiffOn_symm.mdifferentiableOn le_top⟩ end UniqueMDiff diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 7511ccb7ad482..9fcbaff7d4fac 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -48,12 +48,14 @@ but add these assumptions later as needed. (Quite a few results still do not req `extChartAt I x` is the canonical such partial equiv around `x`. As specific examples of models with corners, we define (in `Geometry.Manifold.Instances.Real`) -* `modelWithCornersSelf ℝ (EuclideanSpace ℝ (Fin n))` for the model space used to define +* `modelWithCornersEuclideanHalfSpace n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space `n`-dimensional real manifolds without boundary (with notation `𝓡 n` in the locale `Manifold`) * `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space used to define `n`-dimensional real manifolds with boundary (with notation `𝓡∂ n` in the locale `Manifold`) -* `ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used +* `modelWithCornersEuclideanQuadrant n : + ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanQuadrant n)` for the model space used to define `n`-dimensional real manifolds with corners With these definitions at hand, to invoke an `n`-dimensional real manifold without boundary, @@ -369,7 +371,7 @@ protected theorem secondCountableTopology [SecondCountableTopology E] (I : Model I.isClosedEmbedding.isEmbedding.secondCountableTopology include I in -/-- Every smooth manifold is a Fréchet space (T1 space) -- regardless of whether it is +/-- Every manifold is a Fréchet space (T1 space) -- regardless of whether it is Hausdorff. -/ protected theorem t1Space (M : Type*) [TopologicalSpace M] [ChartedSpace H M] : T1Space M := by have : T2Space H := I.isClosedEmbedding.toIsEmbedding.t2Space diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index 0026330ad4b73..9ecaf70c760a9 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -281,7 +281,7 @@ protected theorem eq {a b : M} : (a : c.Quotient) = (b : c.Quotient) ↔ c a b : @[to_additive "The addition induced on the quotient by an additive congruence relation on a type with an addition."] instance hasMul : Mul c.Quotient := - ⟨Quotient.map₂' (· * ·) fun _ _ h1 _ _ h2 => c.mul h1 h2⟩ + ⟨Quotient.map₂ (· * ·) fun _ _ h1 _ _ h2 => c.mul h1 h2⟩ /-- The kernel of the quotient map induced by a congruence relation `c` equals `c`. -/ @[to_additive (attr := simp) "The kernel of the quotient map induced by an additive congruence @@ -739,7 +739,7 @@ instance hasInv : Inv c.Quotient := @[to_additive "The subtraction induced on the quotient by an additive congruence relation on a type with a subtraction."] instance hasDiv : Div c.Quotient := - ⟨(Quotient.map₂' (· / ·)) fun _ _ h₁ _ _ h₂ => c.div h₁ h₂⟩ + ⟨(Quotient.map₂ (· / ·)) fun _ _ h₁ _ _ h₂ => c.div h₁ h₂⟩ /-- The integer scaling induced on the quotient by a congruence relation on a type with a subtraction. -/ diff --git a/Mathlib/GroupTheory/Coxeter/Basic.lean b/Mathlib/GroupTheory/Coxeter/Basic.lean index 1ab0f3e0bef58..c33678eadac13 100644 --- a/Mathlib/GroupTheory/Coxeter/Basic.lean +++ b/Mathlib/GroupTheory/Coxeter/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Newell Jensen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Newell Jensen, Mitchell Lee +Authors: Newell Jensen, Mitchell Lee, Óscar Álvarez -/ import Mathlib.Algebra.Group.Subgroup.Pointwise import Mathlib.Algebra.Ring.Int.Parity @@ -411,6 +411,81 @@ theorem length_alternatingWord (i i' : B) (m : ℕ) : · dsimp [alternatingWord] · simpa [alternatingWord] using ih i' i +lemma getElem_alternatingWord (i j : B) (p k : ℕ) (hk : k < p) : + (alternatingWord i j p)[k]'(by simp; exact hk) = (if Even (p + k) then i else j) := by + revert k + induction p with + | zero => + intro k hk + simp only [not_lt_zero'] at hk + | succ n h => + intro k hk + simp_rw [alternatingWord_succ' i j n] + match k with + | 0 => + by_cases h2 : Even n + · simp only [h2, ↓reduceIte, getElem_cons_zero, add_zero, + (by simp [Even.add_one, h2] : ¬Even (n + 1))] + · simp only [h2, ↓reduceIte, getElem_cons_zero, add_zero, + Odd.add_one (Nat.not_even_iff_odd.mp h2)] + | k + 1 => + simp only [add_lt_add_iff_right] at hk h + simp only [getElem_cons_succ, h k hk] + ring_nf + have even_add_two (m : ℕ) : Even (2 + m) ↔ Even m := by + simp only [add_tsub_cancel_right, even_two, (Nat.even_sub (by omega : m ≤ 2 + m)).mp] + by_cases h_even : Even (n + k) + · rw [if_pos h_even] + rw [← even_add_two (n+k), ← Nat.add_assoc 2 n k] at h_even + rw [if_pos h_even] + · rw [if_neg h_even] + rw [← even_add_two (n+k), ← Nat.add_assoc 2 n k] at h_even + rw [if_neg h_even] + +lemma getElem_alternatingWord_swapIndices (i j : B) (p k : ℕ) (h : k + 1 < p) : + (alternatingWord i j p)[k+1]'(by simp; exact h) = + (alternatingWord j i p)[k]'(by simp [h]; omega) := by + rw [getElem_alternatingWord i j p (k+1) (by omega), getElem_alternatingWord j i p k (by omega)] + by_cases h_even : Even (p + k) + · rw [if_pos h_even, ← add_assoc] + simp only [ite_eq_right_iff, isEmpty_Prop, Nat.not_even_iff_odd, Even.add_one h_even, + IsEmpty.forall_iff] + · rw [if_neg h_even, ← add_assoc] + simp [Odd.add_one (Nat.not_even_iff_odd.mp h_even)] + +lemma listTake_alternatingWord (i j : B) (p k : ℕ) (h : k < 2 * p) : + List.take k (alternatingWord i j (2 * p)) = + if Even k then alternatingWord i j k else alternatingWord j i k := by + induction k with + | zero => + simp only [take_zero, even_zero, ↓reduceIte, alternatingWord] + | succ k h' => + have hk : k < 2 * p := by omega + apply h' at hk + by_cases h_even : Even k + · simp only [h_even, ↓reduceIte] at hk + simp only [Nat.not_even_iff_odd.mpr (Even.add_one h_even), ↓reduceIte] + rw [← List.take_concat_get _ _ (by simp[h]; omega), alternatingWord_succ, ← hk] + apply congr_arg + rw [getElem_alternatingWord i j (2*p) k (by omega)] + simp [(by apply Nat.even_add.mpr; simp[h_even]: Even (2 * p + k))] + · simp only [h_even, ↓reduceIte] at hk + simp only [(by simp at h_even; exact Odd.add_one h_even : Even (k + 1)), ↓reduceIte] + rw [← List.take_concat_get _ _ (by simp[h]; omega), alternatingWord_succ, hk] + apply congr_arg + rw [getElem_alternatingWord i j (2*p) k (by omega)] + simp [(by apply Nat.odd_add.mpr; simp[h_even]: Odd (2 * p + k))] + +lemma listTake_succ_alternatingWord (i j : B) (p : ℕ) (k : ℕ) (h : k + 1 < 2 * p) : + List.take (k + 1) (alternatingWord i j (2 * p)) = + i :: (List.take k (alternatingWord j i (2 * p))) := by + rw [listTake_alternatingWord j i p k (by omega), listTake_alternatingWord i j p (k+1) h] + + by_cases h_even : Even k + · simp [h_even, Nat.not_even_iff_odd.mpr (Even.add_one h_even), alternatingWord_succ', h_even] + · simp [h_even, (by simp at h_even; exact Odd.add_one h_even: Even (k + 1)), + alternatingWord_succ', h_even] + theorem prod_alternatingWord_eq_mul_pow (i i' : B) (m : ℕ) : π (alternatingWord i i' m) = (if Even m then 1 else s i') * (s i * s i') ^ (m / 2) := by induction' m with m ih diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 6afee6ebfd411..b757715af2158 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2024 Mitchell Lee. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mitchell Lee +Authors: Mitchell Lee, Óscar Álvarez -/ import Mathlib.GroupTheory.Coxeter.Length import Mathlib.Data.List.GetD @@ -255,6 +255,13 @@ theorem getD_rightInvSeq (ω : List B) (j : ℕ) : · simp only [getD_eq_getElem?_getD, get?_eq_getElem?] at ih simp [getD_cons_succ, ih j'] +lemma getElem_rightInvSeq (ω : List B) (j : ℕ) (h : j < ω.length) : + (ris ω)[j]'(by simp[h]) = + (π (ω.drop (j + 1)))⁻¹ + * (Option.map (cs.simple) (ω.get? j)).getD 1 + * π (ω.drop (j + 1)) := by + rw [← List.getD_eq_getElem (ris ω) 1, getD_rightInvSeq] + theorem getD_leftInvSeq (ω : List B) (j : ℕ) : (lis ω).getD j 1 = π (ω.take j) @@ -271,6 +278,12 @@ theorem getD_leftInvSeq (ω : List B) (j : ℕ) : rw [ih j'] simp [← mul_assoc, wordProd_cons] +lemma getElem_leftInvSeq (ω : List B) (j : ℕ) (h : j < ω.length) : + (lis ω)[j]'(by simp[h]) = + cs.wordProd (List.take j ω) * s ω[j] * (cs.wordProd (List.take j ω))⁻¹ := by + rw [← List.getD_eq_getElem (lis ω) 1, getD_leftInvSeq] + simp [h] + theorem getD_rightInvSeq_mul_self (ω : List B) (j : ℕ) : ((ris ω).getD j 1) * ((ris ω).getD j 1) = 1 := by simp_rw [getD_rightInvSeq, mul_assoc] @@ -439,4 +452,38 @@ theorem IsReduced.nodup_leftInvSeq {ω : List B} (rω : cs.IsReduced ω) : List. apply nodup_rightInvSeq rwa [isReduced_reverse] +lemma getElem_succ_leftInvSeq_alternatingWord + (i j : B) (p k : ℕ) (h : k + 1 < 2 * p) : + (lis (alternatingWord i j (2 * p)))[k + 1]'(by simp; exact h) = + MulAut.conj (s i) ((lis (alternatingWord j i (2 * p)))[k]'(by simp; linarith)) := by + rw [cs.getElem_leftInvSeq (alternatingWord i j (2 * p)) (k + 1) (by simp[h]), + cs.getElem_leftInvSeq (alternatingWord j i (2 * p)) k (by simp[h]; omega)] + simp only [MulAut.conj, listTake_succ_alternatingWord i j p k h, cs.wordProd_cons, mul_assoc, + mul_inv_rev, inv_simple, MonoidHom.coe_mk, OneHom.coe_mk, MulEquiv.coe_mk, Equiv.coe_fn_mk, + mul_right_inj, mul_left_inj] + rw [getElem_alternatingWord_swapIndices i j (2 * p) k] + omega + +theorem getElem_leftInvSeq_alternatingWord + (i j : B) (p k : ℕ) (h : k < 2 * p) : + (lis (alternatingWord i j (2 * p)))[k]'(by simp; linarith) = + π alternatingWord j i (2 * k + 1) := by + revert i j + induction k with + | zero => + intro i j + simp only [CoxeterSystem.getElem_leftInvSeq cs (alternatingWord i j (2 * p)) 0 (by simp [h]), + take_zero, wordProd_nil, one_mul, inv_one, mul_one, alternatingWord, concat_eq_append, + nil_append, wordProd_singleton] + apply congr_arg + simp only [getElem_alternatingWord i j (2 * p) 0 (by simp [h]), add_zero, even_two, + Even.mul_right, ↓reduceIte] + | succ k hk => + intro i j + simp only [getElem_succ_leftInvSeq_alternatingWord cs i j p k h, hk (by omega), + MulAut.conj_apply, inv_simple, alternatingWord_succ' j i, even_two, Even.mul_right, + ↓reduceIte, wordProd_cons] + rw [(by ring: 2 * (k + 1) = 2 * k + 1 + 1), alternatingWord_succ j i, wordProd_concat] + simp [mul_assoc] + end CoxeterSystem diff --git a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean index 6cb4009aa11c0..f2d850b5a8d40 100644 --- a/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean +++ b/Mathlib/GroupTheory/SpecificGroups/ZGroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.Algebra.Squarefree.Basic -import Mathlib.GroupTheory.Sylow +import Mathlib.GroupTheory.Nilpotent /-! # Z-Groups @@ -30,6 +30,9 @@ variable {G G' f} namespace IsZGroup +instance [IsZGroup G] {p : ℕ} [Fact p.Prime] (P : Sylow p G) : IsCyclic P := + isZGroup p Fact.out P + theorem of_squarefree (hG : Squarefree (Nat.card G)) : IsZGroup G := by have : Finite G := Nat.finite_of_card_ne_zero hG.ne_zero refine ⟨fun p hp P ↦ ?_⟩ @@ -59,4 +62,24 @@ theorem of_surjective [Finite G] [hG : IsZGroup G] (hf : Function.Surjective f) instance [Finite G] [IsZGroup G] (H : Subgroup G) [H.Normal] : IsZGroup (G ⧸ H) := of_surjective (QuotientGroup.mk'_surjective H) +theorem exponent_eq_card [Finite G] [IsZGroup G] : Monoid.exponent G = Nat.card G := by + refine dvd_antisymm Group.exponent_dvd_nat_card ?_ + rw [← Nat.factorization_prime_le_iff_dvd Nat.card_pos.ne' Monoid.exponent_ne_zero_of_finite] + intro p hp + have := Fact.mk hp + let P : Sylow p G := default + rw [← hp.pow_dvd_iff_le_factorization Monoid.exponent_ne_zero_of_finite, + ← P.card_eq_multiplicity, ← (isZGroup p hp P).exponent_eq_card] + exact Monoid.exponent_dvd_of_monoidHom P.1.subtype P.1.subtype_injective + +instance [Finite G] [IsZGroup G] [hG : Group.IsNilpotent G] : IsCyclic G := by + have (p : { x // x ∈ (Nat.card G).primeFactors }) : Fact p.1.Prime := + ⟨Nat.prime_of_mem_primeFactors p.2⟩ + let h (p : { x // x ∈ (Nat.card G).primeFactors }) (P : Sylow p G) : CommGroup P := + IsCyclic.commGroup + obtain ⟨ϕ⟩ := ((isNilpotent_of_finite_tfae (G := G)).out 0 4).mp hG + let _ : CommGroup G := + ⟨fun g h ↦ by rw [← ϕ.symm.injective.eq_iff, map_mul, mul_comm, ← map_mul]⟩ + exact IsCyclic.of_exponent_eq_card exponent_eq_card + end IsZGroup diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index 76c5638c80a15..b62ed22785418 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -1250,7 +1250,7 @@ theorem affineSpan_pair_le_of_right_mem {p₁ p₂ p₃ : P} (h : p₁ ∈ line[ variable (k) /-- `affineSpan` is monotone. -/ -@[mono] +@[gcongr, mono] theorem affineSpan_mono {s₁ s₂ : Set P} (h : s₁ ⊆ s₂) : affineSpan k s₁ ≤ affineSpan k s₂ := spanPoints_subset_coe_of_subset_coe (Set.Subset.trans h (subset_affineSpan k _)) diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index e3322e7d5c035..662046a8624b8 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -359,6 +359,10 @@ theorem isUnitSMul_apply {v : Basis ι R M} {w : ι → R} (hw : ∀ i, IsUnit ( v.isUnitSMul hw i = w i • v i := unitsSMul_apply i +theorem repr_isUnitSMul {v : Basis ι R₂ M} {w : ι → R₂} (hw : ∀ i, IsUnit (w i)) (x : M) (i : ι) : + (v.isUnitSMul hw).repr x i = (hw i).unit⁻¹ • v.repr x i := + repr_unitsSMul _ _ _ _ + section Fin /-- Let `b` be a basis for a submodule `N` of `M`. If `y : M` is linear independent of `N` diff --git a/Mathlib/LinearAlgebra/Basis/Flag.lean b/Mathlib/LinearAlgebra/Basis/Flag.lean index 4ffe11d49a060..2d138c697c2e9 100644 --- a/Mathlib/LinearAlgebra/Basis/Flag.lean +++ b/Mathlib/LinearAlgebra/Basis/Flag.lean @@ -22,7 +22,8 @@ namespace Basis section Semiring -variable {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] {n : ℕ} +variable {R M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] {n : ℕ} {b : Basis (Fin n) R M} + {i j : Fin (n + 1)} /-- The subspace spanned by the first `k` vectors of the basis `b`. -/ def flag (b : Basis (Fin n) R M) (k : Fin (n + 1)) : Submodule R M := @@ -64,6 +65,11 @@ theorem isChain_range_flag (b : Basis (Fin n) R M) : IsChain (· ≤ ·) (range theorem flag_strictMono [Nontrivial R] (b : Basis (Fin n) R M) : StrictMono b.flag := Fin.strictMono_iff_lt_succ.2 fun _ ↦ by simp [flag_succ] +@[gcongr] lemma flag_le_flag (hij : i ≤ j) : b.flag i ≤ b.flag j := flag_mono _ hij + +@[gcongr] +lemma flag_lt_flag [Nontrivial R] (hij : i < j) : b.flag i < b.flag j := flag_strictMono _ hij + end Semiring section CommRing diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index d88ad52680183..65ef7c84f0abe 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -99,8 +99,7 @@ theorem GradedAlgebra.lift_ι_eq (i' : ZMod 2) (x' : evenOdd Q i') : · rw [Nat.succ_eq_add_one, add_comm, Nat.cast_add, Nat.cast_one] rfl | zero => - set_option linter.deprecated false in - rw [AlgHom.map_zero] + rw [map_zero] apply Eq.symm apply DFinsupp.single_eq_zero.mpr; rfl | add x y hx hy ihx ihy => diff --git a/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean b/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean new file mode 100644 index 0000000000000..e02ccb29fa8ea --- /dev/null +++ b/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean @@ -0,0 +1,201 @@ +/- +Copyright (c) 2024 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel, Joël Riou +-/ +import Mathlib.Algebra.Module.Presentation.Basic +import Mathlib.LinearAlgebra.ExteriorAlgebra.OfAlternating + +/-! +# Exterior powers + +We study the exterior powers of a module `M` over a commutative ring `R`. + +## Definitions + +* `exteriorPower.ιMulti` is the canonical alternating map on `M` with values in `⋀[R]^n M`. + +* `exteriorPower.presentation R n M` is the standard presentation of the `R`-module `⋀[R]^n M`. + +## Theorems +* `ιMulti_span`: The image of `exteriorPower.ιMulti` spans `⋀[R]^n M`. + +* We construct `exteriorPower.alternatingMapLinearEquiv` which +expresses the universal property of the exterior power as a +linear equivalence `(M [⋀^Fin n]→ₗ[R] N) ≃ₗ[R] ⋀[R]^n M →ₗ[R] N` between +alternating maps and linear maps from the exterior power. + +-/ + +open scoped TensorProduct + +universe u v uM uN uN' uN'' uE uF + +variable (R : Type u) [CommRing R] (n : ℕ) {M N N' : Type*} + [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + [AddCommGroup N'] [Module R N'] + +namespace exteriorPower + +open Function + +/-! The canonical alternating map from `Fin n → M` to `⋀[R]^n M`. -/ + +/-- `exteriorAlgebra.ιMulti` is the alternating map from `Fin n → M` to `⋀[r]^n M` +induced by `exteriorAlgebra.ιMulti`, i.e. sending a family of vectors `m : Fin n → M` to the +product of its entries. -/ +def ιMulti : M [⋀^Fin n]→ₗ[R] (⋀[R]^n M) := + (ExteriorAlgebra.ιMulti R n).codRestrict (⋀[R]^n M) fun _ => + ExteriorAlgebra.ιMulti_range R n <| Set.mem_range_self _ + +@[simp] lemma ιMulti_apply_coe (a : Fin n → M) : ιMulti R n a = ExteriorAlgebra.ιMulti R n a := rfl + +variable (M) +/-- The image of `ExteriorAlgebra.ιMulti R n` spans the `n`th exterior power. Variant of +`ExteriorAlgebra.ιMulti_span_fixedDegree`, useful in rewrites. -/ +lemma ιMulti_span_fixedDegree : + Submodule.span R (Set.range (ExteriorAlgebra.ιMulti R n)) = ⋀[R]^n M := + ExteriorAlgebra.ιMulti_span_fixedDegree R n + +/-- The image of `exteriorPower.ιMulti` spans `⋀[R]^n M`. -/ +lemma ιMulti_span : + Submodule.span R (Set.range (ιMulti R n)) = (⊤ : Submodule R (⋀[R]^n M)) := by + apply LinearMap.map_injective (Submodule.ker_subtype (⋀[R]^n M)) + rw [LinearMap.map_span, ← Set.image_univ, Set.image_image] + simp only [Submodule.coe_subtype, ιMulti_apply_coe, Set.image_univ, Submodule.map_top, + Submodule.range_subtype] + exact ExteriorAlgebra.ιMulti_span_fixedDegree R n + +namespace presentation + +/-- The index type for the relations in the standard presentation of `⋀[R]^n M`, +in the particular case `ι` is `Fin n`. -/ +inductive Rels (ι : Type*) (M : Type*) + | add (m : ι → M) (i : ι) (x y : M) + | smul (m : ι → M) (i : ι) (r : R) (x : M) + | alt (m : ι → M) (i j : ι) (hm : m i = m j) (hij : i ≠ j) + +/-- The relations in the standard presentation of `⋀[R]^n M` with generators and relations. -/ +@[simps] +noncomputable def relations (ι : Type*) [DecidableEq ι] (M : Type*) + [AddCommGroup M] [Module R M] : + Module.Relations R where + G := ι → M + R := Rels R ι M + relation r := match r with + | .add m i x y => Finsupp.single (update m i x) 1 + + Finsupp.single (update m i y) 1 - + Finsupp.single (update m i (x + y)) 1 + | .smul m i r x => Finsupp.single (update m i (r • x)) 1 - + r • Finsupp.single (update m i x) 1 + | .alt m _ _ _ _ => Finsupp.single m 1 + +variable {R} in +/-- The solutions in a module `N` to the linear equations +given by `exteriorPower.relations R ι M` identify to alternating maps to `N`. -/ +@[simps!] +def relationsSolutionEquiv {ι : Type*} [DecidableEq ι] {M : Type*} + [AddCommGroup M] [Module R M] : + (relations R ι M).Solution N ≃ AlternatingMap R M N ι where + toFun s := + { toFun := fun m ↦ s.var m + map_update_add' := fun m i x y ↦ by + have := s.linearCombination_var_relation (.add m i x y) + dsimp at this ⊢ + rw [map_sub, map_add, Finsupp.linearCombination_single, one_smul, + Finsupp.linearCombination_single, one_smul, + Finsupp.linearCombination_single, one_smul, sub_eq_zero] at this + convert this.symm -- `convert` is necessary due to the implementation of `MultilinearMap` + map_update_smul' := fun m i r x ↦ by + have := s.linearCombination_var_relation (.smul m i r x) + dsimp at this ⊢ + rw [Finsupp.smul_single, smul_eq_mul, mul_one, map_sub, + Finsupp.linearCombination_single, one_smul, + Finsupp.linearCombination_single, sub_eq_zero] at this + convert this + map_eq_zero_of_eq' := fun v i j hm hij ↦ + by simpa using s.linearCombination_var_relation (.alt v i j hm hij) } + invFun f := + { var := fun m ↦ f m + linearCombination_var_relation := by + rintro (⟨m, i, x, y⟩ | ⟨m, i, r, x⟩ | ⟨v, i, j, hm, hij⟩) + · simp + · simp + · simpa using f.map_eq_zero_of_eq v hm hij } + left_inv _ := rfl + right_inv _ := rfl + +/-- The universal property of the exterior power. -/ +def isPresentationCore : + (relationsSolutionEquiv.symm (ιMulti R n (M := M))).IsPresentationCore where + desc s := LinearMap.comp (ExteriorAlgebra.liftAlternating + (Function.update 0 n (relationsSolutionEquiv s))) (Submodule.subtype _) + postcomp_desc s := by aesop + postcomp_injective {N _ _ f f' h} := by + rw [Submodule.linearMap_eq_iff_of_span_eq_top _ _ (ιMulti_span R n M)] + rintro ⟨_, ⟨f, rfl⟩⟩ + exact Module.Relations.Solution.congr_var h f + +end presentation + +/-- The standard presentation of the `R`-module `⋀[R]^n M`. -/ +@[simps! G R relation var] +noncomputable def presentation : Module.Presentation R (⋀[R]^n M) := + .ofIsPresentation (presentation.isPresentationCore R n M).isPresentation + +variable {R M n} + +/-- Two linear maps on `⋀[R]^n M` that agree on the image of `exteriorPower.ιMulti` +are equal. -/ +@[ext] +lemma linearMap_ext {f : ⋀[R]^n M →ₗ[R] N} {g : ⋀[R]^n M →ₗ[R] N} + (heq : f.compAlternatingMap (ιMulti R n) = g.compAlternatingMap (ιMulti R n)) : f = g := + (presentation R n M).postcomp_injective (by ext f; apply DFunLike.congr_fun heq ) + +/-- The linear equivalence between `n`-fold alternating maps from `M` to `N` and linear maps from +`⋀[R]^n M` to `N`: this is the universal property of the `n`th exterior power of `M`. -/ +noncomputable def alternatingMapLinearEquiv : (M [⋀^Fin n]→ₗ[R] N) ≃ₗ[R] ⋀[R]^n M →ₗ[R] N := + LinearEquiv.symm + (Equiv.toLinearEquiv + ((presentation R n M).linearMapEquiv.trans presentation.relationsSolutionEquiv) + { map_add := fun _ _ => rfl + map_smul := fun _ _ => rfl }) + +@[simp] +lemma alternatingMapLinearEquiv_comp_ιMulti (f : M [⋀^Fin n]→ₗ[R] N) : + (alternatingMapLinearEquiv f).compAlternatingMap (ιMulti R n) = f := by + obtain ⟨φ, rfl⟩ := alternatingMapLinearEquiv.symm.surjective f + dsimp [alternatingMapLinearEquiv] + simp only [LinearEquiv.symm_apply_apply] + rfl + +@[simp] +lemma alternatingMapLinearEquiv_apply_ιMulti (f : M [⋀^Fin n]→ₗ[R] N) (a : Fin n → M) : + alternatingMapLinearEquiv f (ιMulti R n a) = f a := + DFunLike.congr_fun (alternatingMapLinearEquiv_comp_ιMulti f) a + +@[simp] +lemma alternatingMapLinearEquiv_symm_apply (F : ⋀[R]^n M →ₗ[R] N) (m : Fin n → M) : + alternatingMapLinearEquiv.symm F m = F.compAlternatingMap (ιMulti R n) m := by + obtain ⟨f, rfl⟩ := alternatingMapLinearEquiv.surjective F + simp only [LinearEquiv.symm_apply_apply, alternatingMapLinearEquiv_comp_ιMulti] + +@[simp] +lemma alternatingMapLinearEquiv_ιMulti : + alternatingMapLinearEquiv (ιMulti R n (M := M)) = LinearMap.id := by + ext + simp only [alternatingMapLinearEquiv_comp_ιMulti, ιMulti_apply_coe, + LinearMap.compAlternatingMap_apply, LinearMap.id_coe, id_eq] + +/-- If `f` is an alternating map from `M` to `N`, +`alternatingMapLinearEquiv f` is the corresponding linear map from `⋀[R]^n M` to `N`, +and if `g` is a linear map from `N` to `N'`, then +the alternating map `g.compAlternatingMap f` from `M` to `N'` corresponds to the linear +map `g.comp (alternatingMapLinearEquiv f)` on `⋀[R]^n M`. -/ +lemma alternatingMapLinearEquiv_comp (g : N →ₗ[R] N') (f : M [⋀^Fin n]→ₗ[R] N) : + alternatingMapLinearEquiv (g.compAlternatingMap f) = g.comp (alternatingMapLinearEquiv f) := by + ext + simp only [alternatingMapLinearEquiv_comp_ιMulti, LinearMap.compAlternatingMap_apply, + LinearMap.coe_comp, comp_apply, alternatingMapLinearEquiv_apply_ιMulti] + +end exteriorPower diff --git a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean index bfa28160ab5c3..a6d1297916a19 100644 --- a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean @@ -68,7 +68,7 @@ universe uS uA uB then their quotients are also `R`-equivalent. (Special case of the third isomorphism theorem.)-/ -def algEquiv_quot_algEquiv +def algEquivQuotAlgEquiv {R : Type u} [CommSemiring R] {A B : Type v} [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (f : A ≃ₐ[R] B) (rel : A → A → Prop) : RingQuot rel ≃ₐ[R] RingQuot (rel on f.symm) := @@ -83,16 +83,20 @@ def algEquiv_quot_algEquiv fun x y h ↦ by apply RingQuot.mkAlgHom_rel; simpa⟩)) (by ext b; simp) (by ext a; simp) +@[deprecated (since := "2024-12-07")] alias algEquiv_quot_algEquiv := algEquivQuotAlgEquiv + /--If two (semi)rings are equivalent and their quotients by a relation `rel` are defined, then their quotients are also equivalent. (Special case of `algEquiv_quot_algEquiv` when `R = ℕ`, which in turn is a special case of the third isomorphism theorem.)-/ -def equiv_quot_equiv {A B : Type v} [Semiring A] [Semiring B] (f : A ≃+* B) (rel : A → A → Prop) : +def equivQuotEquiv {A B : Type v} [Semiring A] [Semiring B] (f : A ≃+* B) (rel : A → A → Prop) : RingQuot rel ≃+* RingQuot (rel on f.symm) := let f_alg : A ≃ₐ[ℕ] B := AlgEquiv.ofRingEquiv (f := f) (fun n ↦ by simp) - algEquiv_quot_algEquiv f_alg rel |>.toRingEquiv + algEquivQuotAlgEquiv f_alg rel |>.toRingEquiv + +@[deprecated (since := "2024-12-07")] alias equiv_quot_equiv := equivQuotEquiv end RingQuot @@ -143,11 +147,14 @@ theorem rel_id (i : I) : rel R A (ι R <| lof R I A i 1) 1 := rel.id @[reducible] def _root_.LinearAlgebra.FreeProduct := RingQuot <| FreeProduct.rel R A /--The free product of the collection of `R`-algebras `A i`, as a quotient of `PowerAlgebra R A`-/ -@[reducible] def _root_.LinearAlgebra.FreeProduct_ofPowers := RingQuot <| FreeProduct.rel' R A +@[reducible] def _root_.LinearAlgebra.FreeProductOfPowers := RingQuot <| FreeProduct.rel' R A + +@[deprecated (since := "2024-12-07")] +alias _root_.LinearAlgebra.FreeProduct_ofPowers := LinearAlgebra.FreeProductOfPowers /--The `R`-algebra equivalence relating `FreeProduct` and `FreeProduct_ofPowers`-/ -noncomputable def equivPowerAlgebra : FreeProduct_ofPowers R A ≃ₐ[R] FreeProduct R A := - RingQuot.algEquiv_quot_algEquiv +noncomputable def equivPowerAlgebra : FreeProductOfPowers R A ≃ₐ[R] FreeProduct R A := + RingQuot.algEquivQuotAlgEquiv (FreeProduct.powerAlgebra_equiv_freeAlgebra R A |>.symm) (FreeProduct.rel R A) |>.symm diff --git a/Mathlib/LinearAlgebra/Goursat.lean b/Mathlib/LinearAlgebra/Goursat.lean new file mode 100644 index 0000000000000..0316c6952d5d5 --- /dev/null +++ b/Mathlib/LinearAlgebra/Goursat.lean @@ -0,0 +1,137 @@ +/- +Copyright (c) 2024 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import Mathlib.GroupTheory.Goursat +import Mathlib.LinearAlgebra.Prod +import Mathlib.LinearAlgebra.Quotient.Basic + +/-! +# Goursat's lemma for submodules + +Let `M, N` be modules over a ring `R`. If `L` is a submodule of `M × N` which projects fully on +to both factors, then there exist submodules `M' ≤ M` and `N' ≤ N` such that `M' × N' ≤ L` and the +image of `L` in `(M ⧸ M') × (N ⧸ N')` is the graph of an isomorphism `M ⧸ M' ≃ₗ[R] N ⧸ N'`. +Equivalently, `L` is equal to the preimage in `M × N` of the graph of this isomorphism +`M ⧸ M' ≃ₗ[R] N ⧸ N'`. + +`M'` and `N'` can be explicitly constructed as `Submodule.goursatFst L` and `Submodule.goursatSnd L` +respectively. +-/ + +open Function Set LinearMap + +namespace Submodule +variable {R M N : Type*} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + {L : Submodule R (M × N)} + (hL₁ : Surjective (Prod.fst ∘ L.subtype)) (hL₂ : Surjective (Prod.snd ∘ L.subtype)) + +variable (L) in +/-- For `L` a submodule of `M × N`, `L.goursatFst` is the kernel of the projection map `L → N`, +considered as a submodule of `M`. + +This is the first submodule appearing in Goursat's lemma. See `Subgroup.goursat`. -/ +def goursatFst : Submodule R M := + ((LinearMap.snd R M N).comp L.subtype).ker.map ((LinearMap.fst R M N).comp L.subtype) + + +variable (L) in +/-- For `L` a subgroup of `M × N`, `L.goursatSnd` is the kernel of the projection map `L → M`, +considered as a subgroup of `N`. + +This is the second subgroup appearing in Goursat's lemma. See `Subgroup.goursat`. -/ +def goursatSnd : Submodule R N := + ((LinearMap.fst R M N).comp L.subtype).ker.map ((LinearMap.snd R M N).comp L.subtype) + +lemma goursatFst_toAddSubgroup : + (goursatFst L).toAddSubgroup = L.toAddSubgroup.goursatFst := by + ext x + simp [mem_toAddSubgroup, goursatFst, AddSubgroup.mem_goursatFst] + +lemma goursatSnd_toAddSubgroup : + (goursatSnd L).toAddSubgroup = L.toAddSubgroup.goursatSnd := by + ext x + simp [mem_toAddSubgroup, goursatSnd, AddSubgroup.mem_goursatSnd] + +variable (L) in +lemma goursatFst_prod_goursatSnd_le : L.goursatFst.prod L.goursatSnd ≤ L := by + simpa only [← toAddSubgroup_le, goursatFst_toAddSubgroup, goursatSnd_toAddSubgroup] + using L.toAddSubgroup.goursatFst_prod_goursatSnd_le + +include hL₁ hL₂ in +/-- **Goursat's lemma** for a submodule of a product with surjective projections. + +If `L` is a submodule of `M × N` which projects fully on both factors, then there exist submodules +`M' ≤ M` and `N' ≤ N` such that `M' × N' ≤ L` and the image of `L` in `(M ⧸ M') × (N ⧸ N')` is the +graph of an isomorphism of `R`-modules `(M ⧸ M') ≃ (N ⧸ N')`. + +`M` and `N` can be explicitly constructed as `L.goursatFst` and `L.goursatSnd` respectively. -/ +lemma goursat_surjective : ∃ e : (M ⧸ L.goursatFst) ≃ₗ[R] N ⧸ L.goursatSnd, + LinearMap.range ((L.goursatFst.mkQ.prodMap L.goursatSnd.mkQ).comp L.subtype) = e.graph := by + -- apply add-group result + obtain ⟨(e : M ⧸ L.goursatFst ≃+ N ⧸ L.goursatSnd), he⟩ := + L.toAddSubgroup.goursat_surjective hL₁ hL₂ + -- check R-linearity of the map + have (r : R) (x : M ⧸ L.goursatFst) : e (r • x) = r • e x := by + show (r • x, r • e x) ∈ e.toAddMonoidHom.graph + rw [← he, ← Prod.smul_mk] + have : (x, e x) ∈ e.toAddMonoidHom.graph := rfl + rw [← he, AddMonoidHom.mem_range] at this + rcases this with ⟨⟨l, hl⟩, hl'⟩ + use ⟨r • l, L.smul_mem r hl⟩ + rw [← hl'] + rfl + -- define the map as an R-linear equiv + use { e with map_smul' := this } + rw [← toAddSubgroup_injective.eq_iff] + convert he using 1 + ext v + rw [mem_toAddSubgroup, mem_graph_iff, Eq.comm] + rfl + +/-- **Goursat's lemma** for an arbitrary submodule of a product. + +If `L` is a submodule of `M × N`, then there exist submodules `M'' ≤ M' ≤ M` and `N'' ≤ N' ≤ N` such +that `L ≤ M' × N'`, and `L` is (the image in `M × N` of) the preimage of the graph of an `R`-linear +isomorphism `M' ⧸ M'' ≃ N' ⧸ N''`. -/ +lemma goursat : ∃ (M' : Submodule R M) (N' : Submodule R N) (M'' : Submodule R M') + (N'' : Submodule R N') (e : (M' ⧸ M'') ≃ₗ[R] N' ⧸ N''), + L = (e.graph.comap <| M''.mkQ.prodMap N''.mkQ).map (M'.subtype.prodMap N'.subtype) := by + let M' := L.map (LinearMap.fst ..) + let N' := L.map (LinearMap.snd ..) + let P : L →ₗ[R] M' := (LinearMap.fst ..).submoduleMap L + let Q : L →ₗ[R] N' := (LinearMap.snd ..).submoduleMap L + let L' : Submodule R (M' × N') := LinearMap.range (P.prod Q) + have hL₁' : Surjective (Prod.fst ∘ L'.subtype) := by + simp only [← coe_fst (R := R), ← coe_comp, ← range_eq_top, LinearMap.range_comp, range_subtype] + simpa only [L', ← LinearMap.range_comp, fst_prod, range_eq_top] using + (LinearMap.fst ..).submoduleMap_surjective L + have hL₂' : Surjective (Prod.snd ∘ L'.subtype) := by + simp only [← coe_snd (R := R), ← coe_comp, ← range_eq_top, LinearMap.range_comp, range_subtype] + simpa only [L', ← LinearMap.range_comp, snd_prod, range_eq_top] using + (LinearMap.snd ..).submoduleMap_surjective L + obtain ⟨e, he⟩ := goursat_surjective hL₁' hL₂' + use M', N', L'.goursatFst, L'.goursatSnd, e + rw [← he] + simp only [LinearMap.range_comp, Submodule.range_subtype, L'] + rw [comap_map_eq_self] + · ext ⟨m, n⟩ + constructor + · intro hmn + simp only [mem_map, LinearMap.mem_range, prod_apply, Subtype.exists, Prod.exists, coe_prodMap, + coe_subtype, Prod.map_apply, Prod.mk.injEq, exists_and_right, exists_eq_right_right, + exists_eq_right, M', N', fst_apply, snd_apply] + exact ⟨⟨n, hmn⟩, ⟨m, hmn⟩, ⟨m, n, hmn, rfl⟩⟩ + · simp only [mem_map, LinearMap.mem_range, prod_apply, Subtype.exists, Prod.exists, + coe_prodMap, coe_subtype, Prod.map_apply, Prod.mk.injEq, exists_and_right, + exists_eq_right_right, exists_eq_right, forall_exists_index, Pi.prod] + rintro hm hn m₁ n₁ hm₁n₁ ⟨hP, hQ⟩ + simp only [Subtype.ext_iff] at hP hQ + rwa [← hP, ← hQ] + · convert goursatFst_prod_goursatSnd_le (range <| P.prod Q) + ext ⟨m, n⟩ + simp_rw [mem_ker, coe_prodMap, Prod.map_apply, Submodule.mem_prod, Prod.zero_eq_mk, + Prod.ext_iff, ← mem_ker, ker_mkQ] + +end Submodule diff --git a/Mathlib/LinearAlgebra/Matrix/Ideal.lean b/Mathlib/LinearAlgebra/Matrix/Ideal.lean index b4de0f855d6dc..ef88aeec52f02 100644 --- a/Mathlib/LinearAlgebra/Matrix/Ideal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Ideal.lean @@ -6,7 +6,7 @@ Authors: Jujian Zhang, Wojciech Nawrocki import Mathlib.Data.Matrix.Basis import Mathlib.RingTheory.Ideal.Lattice import Mathlib.RingTheory.TwoSidedIdeal.Operations -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal /-! # Ideals in a matrix ring diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 33f885c1bc407..8aa773428cf39 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -76,6 +76,10 @@ theorem fst_apply (x : M × M₂) : fst R M M₂ x = x.1 := theorem snd_apply (x : M × M₂) : snd R M M₂ x = x.2 := rfl +@[simp, norm_cast] lemma coe_fst : ⇑(fst R M M₂) = Prod.fst := rfl + +@[simp, norm_cast] lemma coe_snd : ⇑(snd R M M₂) = Prod.snd := rfl + theorem fst_surjective : Function.Surjective (fst R M M₂) := fun x => ⟨(x, 0), rfl⟩ theorem snd_surjective : Function.Surjective (snd R M M₂) := fun x => ⟨(0, x), rfl⟩ @@ -830,8 +834,7 @@ variable {N : Type*} [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] open Function -set_option linter.deprecated false - +set_option linter.deprecated false in /-- An auxiliary construction for `tunnel`. The composition of `f`, followed by the isomorphism back to `K`, followed by the inclusion of this submodule back into `M`. -/ @@ -839,11 +842,13 @@ followed by the inclusion of this submodule back into `M`. -/ def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := (Subtype.val_injective.comp Kφ.2.symm.injective).comp i +set_option linter.deprecated false in /-- Auxiliary definition for `tunnel`. -/ @[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M @@ -853,6 +858,7 @@ def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule ((Submodule.fst R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.fstEquiv R M N)⟩ +set_option linter.deprecated false in /-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ @@ -865,6 +871,7 @@ def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M) rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le⟩ +set_option linter.deprecated false in /-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules all isomorphic to `N`. -/ @@ -872,12 +879,14 @@ all isomorphic to `N`. def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) +set_option linter.deprecated false in /-- Each `tailing f i n` is a copy of `N`. -/ @[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.sndEquiv R M N) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by @@ -885,6 +894,7 @@ theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by @@ -894,6 +904,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, Submodule.fst_inf_snd, Submodule.map_bot] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ @@ -902,19 +913,23 @@ theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injectiv rw [← Submodule.map_sup, sup_comm, Submodule.fst_sup_snd, Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le +set_option linter.deprecated false in /-- The supremum of all the copies of `N` found inside the tunnel. -/ @[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := partialSups (tailing f i) +set_option linter.deprecated false in @[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by simp [tailings] +set_option linter.deprecated false in @[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by @@ -927,6 +942,7 @@ theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : · apply Disjoint.mono_right _ ih apply tailing_sup_tunnel_succ_le_tunnel +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (tailing f i (n + 1)) := diff --git a/Mathlib/LinearAlgebra/Projectivization/Constructions.lean b/Mathlib/LinearAlgebra/Projectivization/Constructions.lean index fd7470292b540..cdbb83b605a70 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Constructions.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Constructions.lean @@ -67,7 +67,7 @@ variable [DecidableEq F] /-- Cross product on the projective plane. -/ def cross : ℙ F (Fin 3 → F) → ℙ F (Fin 3 → F) → ℙ F (Fin 3 → F) := - Quotient.map₂' (fun v w ↦ if h : crossProduct v.1 w.1 = 0 then v else ⟨crossProduct v.1 w.1, h⟩) + Quotient.map₂ (fun v w ↦ if h : crossProduct v.1 w.1 = 0 then v else ⟨crossProduct v.1 w.1, h⟩) (fun _ _ ⟨a, ha⟩ _ _ ⟨b, hb⟩ ↦ by simp_rw [← ha, ← hb, LinearMap.map_smul_of_tower, LinearMap.smul_apply, smul_smul, mul_comm b a, smul_eq_zero_iff_eq] diff --git a/Mathlib/LinearAlgebra/Projectivization/Subspace.lean b/Mathlib/LinearAlgebra/Projectivization/Subspace.lean index 2ae201b95720b..b0a0703432616 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Subspace.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Subspace.lean @@ -154,6 +154,9 @@ span of that set. -/ theorem monotone_span : Monotone (span : Set (ℙ K V) → Subspace K V) := gi.gc.monotone_l +@[gcongr] +lemma span_le_span {s t : Set (ℙ K V)} (hst : s ⊆ t) : span s ≤ span t := monotone_span hst + theorem subset_span_trans {S T U : Set (ℙ K V)} (hST : S ⊆ span T) (hTU : T ⊆ span U) : S ⊆ span U := gi.gc.le_u_l_trans hST hTU diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean index 7e6c8fa28492e..5d6cf88a2bfcf 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean @@ -37,8 +37,9 @@ instance : CoeSort (QuadraticModuleCat.{v} R) (Type v) := /-- The object in the category of quadratic R-modules associated to a quadratic R-module. -/ @[simps form] def of {X : Type v} [AddCommGroup X] [Module R X] (Q : QuadraticForm R X) : - QuadraticModuleCat R where - form := Q + QuadraticModuleCat R := + { ModuleCat.of R X with + form := Q } /-- A type alias for `QuadraticForm.LinearIsometry` to avoid confusion between the categorical and algebraic spellings of composition. -/ @@ -89,7 +90,7 @@ instance concreteCategory : ConcreteCategory.{v} (QuadraticModuleCat.{v} R) wher instance hasForgetToModule : HasForget₂ (QuadraticModuleCat R) (ModuleCat R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => f.toIsometry.toLinearMap } + map := fun f => ModuleCat.ofHom f.toIsometry.toLinearMap } @[simp] theorem forget₂_obj (X : QuadraticModuleCat R) : @@ -98,7 +99,8 @@ theorem forget₂_obj (X : QuadraticModuleCat R) : @[simp] theorem forget₂_map (X Y : QuadraticModuleCat R) (f : X ⟶ Y) : - (forget₂ (QuadraticModuleCat R) (ModuleCat R)).map f = f.toIsometry.toLinearMap := + (forget₂ (QuadraticModuleCat R) (ModuleCat R)).map f = + ModuleCat.ofHom f.toIsometry.toLinearMap := rfl variable {X Y Z : Type v} diff --git a/Mathlib/LinearAlgebra/Reflection.lean b/Mathlib/LinearAlgebra/Reflection.lean index d807bcfeec42e..79d13d900e5b4 100644 --- a/Mathlib/LinearAlgebra/Reflection.lean +++ b/Mathlib/LinearAlgebra/Reflection.lean @@ -1,12 +1,13 @@ /- Copyright (c) 2023 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Oliver Nash, Deepro Choudhury +Authors: Oliver Nash, Deepro Choudhury, Mitchell Lee, Johan Commelin -/ import Mathlib.Algebra.Module.LinearMap.Basic import Mathlib.GroupTheory.OrderOfElement import Mathlib.LinearAlgebra.Dual import Mathlib.LinearAlgebra.FiniteSpan +import Mathlib.RingTheory.Polynomial.Chebyshev /-! # Reflections in linear algebra @@ -28,6 +29,9 @@ is characterised by properties 1 and 2 above, and is a linear isometry. * `Module.reflection`: the definition of the map `y ↦ y - (f y) • x`. This requires the assumption that `f x = 2` but by way of compensation it produces a linear equivalence rather than a mere linear map. + * `Module.reflection_mul_reflection_pow_apply`: a formula for $(r_1 r_2)^m z$, where $r_1$ and + $r_2$ are reflections and $z \in M$. It involves the Chebyshev polynomials and holds over any + commutative ring. This is used to define reflection representations of Coxeter groups. * `Module.Dual.eq_of_preReflection_mapsTo`: a uniqueness result about reflections preserving finite spanning sets that is useful in the theory of root data / systems. @@ -101,6 +105,9 @@ lemma involutive_reflection (h : f x = 2) : Involutive (reflection h) := involutive_preReflection h +@[simp] +lemma reflection_inv (h : f x = 2) : (reflection h)⁻¹ = reflection h := rfl + @[simp] lemma reflection_symm (h : f x = 2) : (reflection h).symm = reflection h := @@ -114,6 +121,194 @@ lemma bijOn_reflection_of_mapsTo {Φ : Set M} (h : f x = 2) (h' : MapsTo (reflec BijOn (reflection h) Φ Φ := (invOn_reflection_of_mapsTo h).bijOn h' h' +/-! ### Powers of the product of two reflections + +Let $M$ be a module over a commutative ring $R$. Let $x, y \in M$ and $f, g \in M^*$ with +$f(x) = g(y) = 2$. The corresponding reflections $r_1, r_2 \colon M \to M$ (`Module.reflection`) are +given by $r_1z = z - f(z) x$ and $r_2 z = z - g(z) y$. These are linear automorphisms of $M$. + +To define reflection representations of a Coxeter group, it is important to be able to compute the +order of the composition $r_1 r_2$. + +Note that if $M$ is a real inner product space and $r_1$ and $r_2$ are both orthogonal +reflections (i.e. $f(z) = 2 \langle x, z \rangle / \langle x, x \rangle$ and +$g(z) = 2 \langle y, z\rangle / \langle y, y\rangle$ for all $z \in M$), +then $r_1 r_2$ is a rotation by the angle +$$\cos^{-1}\left(\frac{f(y) g(x) - 2}{2}\right)$$ +and one may determine the order of $r_1 r_2$ accordingly. + +However, if $M$ does not have an inner product, and even if $R$ is not $\mathbb{R}$, then we may +instead use the formulas in this section. These formulas all involve evaluating Chebyshev +$S$-polynomials (`Polynomial.Chebyshev.S`) at $t = f(y) g(x) - 2$, and they hold over any +commutative ring. -/ +section + +open Int Polynomial.Chebyshev + +variable {x y : M} {f g : Dual R M} (hf : f x = 2) (hg : g y = 2) + +/-- A formula for $(r_1 r_2)^m z$, where $m$ is a natural number and $z \in M$. -/ +lemma reflection_mul_reflection_pow_apply (m : ℕ) (z : M) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) z = + z + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x * f z - g z) • y - f z • x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y * g z - f z) • x - g z • y) := by + induction m with + | zero => simp + | succ m ih => + /- Now, let us collect two facts about the evaluations of `S r k`. These easily follow from the + properties of the `S` polynomials. -/ + have S_eval_t_sub_two (k : ℤ) : + (S R (k - 2)).eval t = t * (S R (k - 1)).eval t - (S R k).eval t := by + simp [S_sub_two] + have S_eval_t_sq_add_S_eval_t_sq (k : ℤ) : + (S R k).eval t ^ 2 + (S R (k + 1)).eval t ^ 2 - t * (S R k).eval t * (S R (k + 1)).eval t + = 1 := by + simpa using congr_arg (Polynomial.eval t) (S_sq_add_S_sq R k) + -- Apply the inductive hypothesis. + rw [pow_succ', LinearEquiv.mul_apply, ih, LinearEquiv.mul_apply] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- `m` can be written in the form `2 * k + e`, where `e` is `0` or `1`. + push_cast + set k : ℤ := m / 2 + set e : ℤ := m % 2 + rw [show m = 2 * k + e from (Int.ediv_add_emod m 2).symm] + simp_rw [add_assoc (2 * k), add_sub_assoc (2 * k), add_comm (2 * k), + add_mul_ediv_left _ k (by norm_num : (2 : ℤ) ≠ 0)] + have he : e = 0 ∨ e = 1 := by omega + clear_value e + /- Now, equate the coefficients on both sides. These linear combinations were + found using `polyrith`. -/ + match_scalars + · rfl + · linear_combination (norm := skip) (-g z * f y * (S R (e - 1 + k)).eval t + + f z * (S R (e - 1 + k)).eval t) * S_eval_t_sub_two (e + k) + + (-g z * f y + f z) * S_eval_t_sq_add_S_eval_t_sq (k - 1) + subst ht + obtain rfl | rfl : e = 0 ∨ e = 1 := he <;> ring_nf + · linear_combination (norm := skip) + g z * (S R (e - 1 + k)).eval t * S_eval_t_sub_two (e + k) + + g z * S_eval_t_sq_add_S_eval_t_sq (k - 1) + subst ht + obtain rfl | rfl : e = 0 ∨ e = 1 := he <;> ring_nf + +/-- A formula for $(r_1 r_2)^m$, where $m$ is a natural number. -/ +lemma reflection_mul_reflection_pow (m : ℕ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m).toLinearMap = + LinearMap.id (R := R) (M := M) + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x • f - g).smulRight y - f.smulRight x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y • g - f).smulRight x - g.smulRight y) := by + ext z + simpa using reflection_mul_reflection_pow_apply hf hg m z t ht + +/-- A formula for $(r_1 r_2)^m z$, where $m$ is an integer and $z \in M$. -/ +lemma reflection_mul_reflection_zpow_apply (m : ℤ) (z : M) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) z = + z + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x * f z - g z) • y - f z • x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y * g z - f z) • x - g z • y) := by + induction m using Int.negInduction with + | nat m => exact_mod_cast reflection_mul_reflection_pow_apply hf hg m z t ht + | neg _ m => + have ht' : t = g x * f y - 2 := by rwa [mul_comm (g x)] + rw [zpow_neg, ← inv_zpow, mul_inv_rev, reflection_inv, reflection_inv, zpow_natCast, + reflection_mul_reflection_pow_apply hg hf m z t ht', add_right_comm z] + have aux {a b : ℤ} (hab : a + b = -3) : a / 2 = -(b / 2) - 2 := by + rw [← mul_right_inj' (by norm_num : (2 : ℤ) ≠ 0), mul_sub, mul_neg, + eq_sub_of_add_eq (Int.ediv_add_emod _ _), eq_sub_of_add_eq (Int.ediv_add_emod _ _)] + omega + rw [aux (by omega : (-m - 3) + m = (-3 : ℤ)), + aux (by omega : (-m - 2) + (m - 1) = (-3 : ℤ)), + aux (by omega : (-m - 1) + (m - 2) = (-3 : ℤ)), + aux (by omega : -m + (m - 3) = (-3 : ℤ))] + simp only [S_neg_sub_two, Polynomial.eval_neg] + ring_nf + +/-- A formula for $(r_1 r_2)^m$, where $m$ is an integer. -/ +lemma reflection_mul_reflection_zpow (m : ℤ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m).toLinearMap = + LinearMap.id (R := R) (M := M) + + ((S R ((m - 2) / 2)).eval t * ((S R ((m - 1) / 2)).eval t + (S R ((m - 3) / 2)).eval t)) • + ((g x • f - g).smulRight y - f.smulRight x) + + ((S R ((m - 1) / 2)).eval t * ((S R (m / 2)).eval t + (S R ((m - 2) / 2)).eval t)) • + ((f y • g - f).smulRight x - g.smulRight y) := by + ext z + simpa using reflection_mul_reflection_zpow_apply hf hg m z t ht + +/-- A formula for $(r_1 r_2)^m x$, where $m$ is an integer. This is the special case of +`Module.reflection_mul_reflection_zpow_apply` with $z = x$. -/ +lemma reflection_mul_reflection_zpow_apply_self (m : ℤ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R (m - 1)).eval t * -g x) • y := by + /- Even though this is a special case of `Module.reflection_mul_reflection_zpow_apply`, it is + easier to prove it from scratch. -/ + have S_eval_t_sub_two (k : ℤ) : + (S R (k - 2)).eval t = (f y * g x - 2) * (S R (k - 1)).eval t - (S R k).eval t := by + simp [S_sub_two, ht] + induction m using Int.induction_on with + | hz => simp + | hp m ih => + -- Apply the inductive hypothesis. + rw [add_comm (m : ℤ) 1, zpow_one_add, LinearEquiv.mul_apply, LinearEquiv.mul_apply, ih] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- Equate coefficients of `x` and `y`. + match_scalars + · linear_combination (norm := ring_nf) -S_eval_t_sub_two (m + 1) + · ring_nf + | hn m ih => + -- Apply the inductive hypothesis. + rw [sub_eq_add_neg (-m : ℤ) 1, add_comm (-m : ℤ) (-1), zpow_add, zpow_neg_one, mul_inv_rev, + reflection_inv, reflection_inv, LinearEquiv.mul_apply, LinearEquiv.mul_apply, ih] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- Equate coefficients of `x` and `y`. + match_scalars + · linear_combination (norm := ring_nf) -S_eval_t_sub_two (-m) + · linear_combination (norm := ring_nf) g x * S_eval_t_sub_two (-m) + +/-- A formula for $(r_1 r_2)^m x$, where $m$ is a natural number. This is the special case of +`Module.reflection_mul_reflection_pow_apply` with $z = x$. -/ +lemma reflection_mul_reflection_pow_apply_self (m : ℕ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + ((reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R (m - 1)).eval t * -g x) • y := + mod_cast reflection_mul_reflection_zpow_apply_self hf hg m t ht + +/-- A formula for $r_2 (r_1 r_2)^m x$, where $m$ is an integer. -/ +lemma reflection_mul_reflection_mul_reflection_zpow_apply_self (m : ℤ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + (reflection hg * (reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R m).eval t * -g x) • y := by + rw [LinearEquiv.mul_apply, reflection_mul_reflection_zpow_apply_self hf hg m t ht] + -- Expand out all the reflections and use `hf`, `hg`. + simp only [reflection_apply, map_add, map_sub, map_smul, hf, hg] + -- Equate coefficients of `x` and `y`. + module + +/-- A formula for $r_2 (r_1 r_2)^m x$, where $m$ is a natural number. -/ +lemma reflection_mul_reflection_mul_reflection_pow_apply_self (m : ℕ) + (t : R := f y * g x - 2) (ht : t = f y * g x - 2 := by rfl) : + (reflection hg * (reflection hf * reflection hg) ^ m) x = + ((S R m).eval t + (S R (m - 1)).eval t) • x + ((S R m).eval t * -g x) • y := + mod_cast reflection_mul_reflection_mul_reflection_zpow_apply_self hf hg m t ht + +end + +/-! ### Lemmas used to prove uniqueness results for root data -/ + /-- See also `Module.Dual.eq_of_preReflection_mapsTo'` for a variant of this lemma which applies when `Φ` does not span. diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index 5a32efb2dc700..ba9290e8307a8 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -73,7 +73,7 @@ protected def equiv_of_mapsTo : end reflection_perm -lemma infinite_of_linearly_independent_coxeterWeight_four [CharZero R] [NoZeroSMulDivisors ℤ M] +lemma infinite_of_linearIndependent_coxeterWeight_four [CharZero R] [NoZeroSMulDivisors ℤ M] (P : RootPairing ι R M N) (i j : ι) (hl : LinearIndependent R ![P.root i, P.root j]) (hc : P.coxeterWeight i j = 4) : Infinite ι := by refine (infinite_range_iff (Embedding.injective P.root)).mp (Infinite.mono ?_ @@ -94,6 +94,13 @@ lemma infinite_of_linearly_independent_coxeterWeight_four [CharZero R] [NoZeroSM variable [Finite ι] (P : RootPairing ι R M N) (i j : ι) +lemma coxeterWeight_ne_four_of_linearIndependent [CharZero R] [NoZeroSMulDivisors ℤ M] + (hl : LinearIndependent R ![P.root i, P.root j]) : + P.coxeterWeight i j ≠ 4 := by + intro contra + have := P.infinite_of_linearIndependent_coxeterWeight_four i j hl contra + exact not_finite ι + /-- Even though the roots may not span, coroots are distinguished by their pairing with the roots. The proof depends crucially on the fact that there are finitely-many roots. diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 3970a7e11cf75..66e76b9408a4a 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -541,6 +541,12 @@ def coxeterWeight : R := pairing P i j * pairing P j i lemma coxeterWeight_swap : coxeterWeight P i j = coxeterWeight P j i := by simp only [coxeterWeight, mul_comm] +lemma exists_int_eq_coxeterWeight (h : P.IsCrystallographic) (i j : ι) : + ∃ z : ℤ, P.coxeterWeight i j = z := by + obtain ⟨a, ha⟩ := P.isCrystallographic_iff.mp h i j + obtain ⟨b, hb⟩ := P.isCrystallographic_iff.mp h j i + exact ⟨a * b, by simp [coxeterWeight, ha, hb]⟩ + /-- Two roots are orthogonal when they are fixed by each others' reflections. -/ def IsOrthogonal : Prop := pairing P i j = 0 ∧ pairing P j i = 0 diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index 9917885d01294..3b1ebfefe84f2 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -3,9 +3,8 @@ Copyright (c) 2024 Scott Carnahan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Carnahan -/ -import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Ring.SumsOfSquares -import Mathlib.LinearAlgebra.RootSystem.Defs +import Mathlib.LinearAlgebra.RootSystem.RootPositive /-! # The canonical bilinear form on a finite root pairing @@ -98,6 +97,15 @@ lemma corootForm_apply_apply (x y : N) : P.CorootForm x y = ∑ i, P.root' i x * P.root' i y := by simp [CorootForm] +lemma toPerfectPairing_apply_apply_Polarization (x y : M) : + P.toPerfectPairing y (P.Polarization x) = P.RootForm x y := by + simp [RootForm] + +lemma toPerfectPairing_apply_CoPolarization (x : N) : + P.toPerfectPairing (P.CoPolarization x) = P.CorootForm x := by + ext y + exact P.flip.toPerfectPairing_apply_apply_Polarization x y + lemma flip_comp_polarization_eq_rootForm : P.flip.toLin ∘ₗ P.Polarization = P.RootForm := by ext; simp [rootForm_apply_apply, RootPairing.flip] @@ -150,6 +158,23 @@ lemma rootForm_self_smul_coroot (i : ι) : rw [Finset.sum_smul, add_neg_eq_zero.mpr rfl] exact sub_eq_zero_of_eq rfl +lemma four_smul_rootForm_sq_eq_coxeterWeight_smul (i j : ι) : + 4 • (P.RootForm (P.root i) (P.root j)) ^ 2 = P.coxeterWeight i j • + (P.RootForm (P.root i) (P.root i) * P.RootForm (P.root j) (P.root j)) := by + have hij : 4 • (P.RootForm (P.root i)) (P.root j) = + 2 • P.toPerfectPairing (P.root j) (2 • P.Polarization (P.root i)) := by + rw [← toPerfectPairing_apply_apply_Polarization, LinearMap.map_smul_of_tower, ← smul_assoc, + Nat.nsmul_eq_mul] + have hji : 2 • (P.RootForm (P.root i)) (P.root j) = + P.toPerfectPairing (P.root i) (2 • P.Polarization (P.root j)) := by + rw [show (P.RootForm (P.root i)) (P.root j) = (P.RootForm (P.root j)) (P.root i) by + apply rootForm_symmetric, ← toPerfectPairing_apply_apply_Polarization, + LinearMap.map_smul_of_tower] + rw [sq, nsmul_eq_mul, ← mul_assoc, ← nsmul_eq_mul, hij, ← rootForm_self_smul_coroot, + smul_mul_assoc 2, ← mul_smul_comm, hji, ← rootForm_self_smul_coroot, map_smul, ← pairing, + map_smul, ← pairing, smul_eq_mul, smul_eq_mul, smul_eq_mul, coxeterWeight] + ring + lemma corootForm_self_smul_root (i : ι) : (P.CorootForm (P.coroot i) (P.coroot i)) • P.root i = 2 • P.CoPolarization (P.coroot i) := rootForm_self_smul_coroot (P.flip) i @@ -229,6 +254,38 @@ lemma rootForm_root_self_pos (j : ι) : use j simp +/-- SGA3 XXI Prop. 2.3.1 -/ +lemma coxeterWeight_le_four (i j : ι) : P.coxeterWeight i j ≤ 4 := by + set li := P.RootForm (P.root i) (P.root i) + set lj := P.RootForm (P.root j) (P.root j) + set lij := P.RootForm (P.root i) (P.root j) + have hi := P.rootForm_root_self_pos i + have hj := P.rootForm_root_self_pos j + have cs : 4 * lij ^ 2 ≤ 4 * (li * lj) := by + rw [mul_le_mul_left four_pos] + exact LinearMap.BilinForm.apply_sq_le_of_symm P.RootForm P.rootForm_self_non_neg + P.rootForm_symmetric (P.root i) (P.root j) + have key : 4 • lij ^ 2 = _ • (li * lj) := P.four_smul_rootForm_sq_eq_coxeterWeight_smul i j + simp only [nsmul_eq_mul, smul_eq_mul, Nat.cast_ofNat] at key + rwa [key, mul_le_mul_right (by positivity)] at cs + +instance instIsRootPositiveRootForm : IsRootPositive P P.RootForm where + zero_lt_apply_root i := P.rootForm_root_self_pos i + symm := P.rootForm_symmetric + apply_reflection_eq := P.rootForm_reflection_reflection_apply + +lemma coxeterWeight_mem_set_of_isCrystallographic (i j : ι) (hP : P.IsCrystallographic) : + P.coxeterWeight i j ∈ ({0, 1, 2, 3, 4} : Set R) := by + obtain ⟨n, hcn⟩ : ∃ n : ℕ, P.coxeterWeight i j = n := by + obtain ⟨z, hz⟩ := P.exists_int_eq_coxeterWeight hP i j + have hz₀ : 0 ≤ z := by simpa [hz] using P.coxeterWeight_non_neg P.RootForm i j + obtain ⟨n, rfl⟩ := Int.eq_ofNat_of_zero_le hz₀ + exact ⟨n, by simp [hz]⟩ + have : P.coxeterWeight i j ≤ 4 := P.coxeterWeight_le_four i j + simp only [hcn, mem_insert_iff, mem_singleton_iff] at this ⊢ + norm_cast at this ⊢ + omega + lemma prod_rootForm_root_self_pos : 0 < ∏ i, P.RootForm (P.root i) (P.root i) := Finset.prod_pos fun i _ => rootForm_root_self_pos P i diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 821094cd58377..65fb8e8be87a7 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -7,7 +7,6 @@ import Mathlib.LinearAlgebra.BilinearForm.Basic import Mathlib.LinearAlgebra.Dimension.Localization import Mathlib.LinearAlgebra.QuadraticForm.Basic import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear -import Mathlib.LinearAlgebra.RootSystem.RootPositive /-! # Nondegeneracy of the polarization on a finite root pairing @@ -55,11 +54,6 @@ variable {ι R M N : Type*} variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] (P : RootPairing ι R M N) -lemma rootForm_rootPositive : IsRootPositive P P.RootForm where - zero_lt_apply_root i := P.rootForm_root_self_pos i - symm := P.rootForm_symmetric - apply_reflection_eq := P.rootForm_reflection_reflection_apply - @[simp] lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : finrank R (P.rootSpan.map P.Polarization) = finrank R P.corootSpan := by diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 99617d933657d..f4f55354935c4 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -798,4 +798,97 @@ end CommRing end Nondegenerate +namespace BilinForm + +lemma apply_smul_sub_smul_sub_eq [CommRing R] [AddCommGroup M] [Module R M] + (B : LinearMap.BilinForm R M) (x y : M) : + B ((B x y) • x - (B x x) • y) ((B x y) • x - (B x x) • y) = + (B x x) * ((B x x) * (B y y) - (B x y) * (B y x)) := by + simp only [map_sub, map_smul, sub_apply, smul_apply, smul_eq_mul, mul_sub, + mul_comm (B x y) (B x x), mul_left_comm (B x y) (B x x)] + abel + +variable [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] (B : LinearMap.BilinForm R M) + +/-- The **Cauchy-Schwarz inequality** for positive semidefinite forms. -/ +lemma apply_mul_apply_le_of_forall_zero_le (hs : ∀ x, 0 ≤ B x x) (x y : M) : + (B x y) * (B y x) ≤ (B x x) * (B y y) := by + have aux (x y : M) : 0 ≤ (B x x) * ((B x x) * (B y y) - (B x y) * (B y x)) := by + rw [← apply_smul_sub_smul_sub_eq B x y] + exact hs (B x y • x - B x x • y) + rcases lt_or_le 0 (B x x) with hx | hx + · exact sub_nonneg.mp <| nonneg_of_mul_nonneg_right (aux x y) hx + · replace hx : B x x = 0 := le_antisymm hx (hs x) + rcases lt_or_le 0 (B y y) with hy | hy + · rw [mul_comm (B x y), mul_comm (B x x)] + exact sub_nonneg.mp <| nonneg_of_mul_nonneg_right (aux y x) hy + · replace hy : B y y = 0 := le_antisymm hy (hs y) + suffices B x y = - B y x by simpa [this, hx, hy] using mul_self_nonneg (B y x) + rw [eq_neg_iff_add_eq_zero] + apply le_antisymm + · simpa [hx, hy, le_neg_iff_add_nonpos_left] using hs (x - y) + · simpa [hx, hy] using hs (x + y) + +/-- The **Cauchy-Schwarz inequality** for positive semidefinite symmetric forms. -/ +lemma apply_sq_le_of_symm (hs : ∀ x, 0 ≤ B x x) (hB : B.IsSymm) (x y : M) : + (B x y) ^ 2 ≤ (B x x) * (B y y) := by + rw [show (B x y) ^ 2 = (B x y) * (B y x) by rw [sq, ← hB, RingHom.id_apply]] + exact apply_mul_apply_le_of_forall_zero_le B hs x y + +/-- The equality case of **Cauchy-Schwarz**. -/ +lemma not_linearIndependent_of_apply_mul_apply_eq (hp : ∀ x, x ≠ 0 → 0 < B x x) + (x y : M) (he : (B x y) * (B y x) = (B x x) * (B y y)) : + ¬ LinearIndependent R ![x, y] := by + have hz : (B x y) • x - (B x x) • y = 0 := by + by_contra hc + exact (ne_of_lt (hp ((B x) y • x - (B x) x • y) hc)).symm <| + (apply_smul_sub_smul_sub_eq B x y).symm ▸ (mul_eq_zero_of_right ((B x) x) + (sub_eq_zero_of_eq he.symm)) + by_contra hL + by_cases hx : x = 0 + · simpa [hx] using LinearIndependent.ne_zero 0 hL + · have h := sub_eq_zero.mpr (sub_eq_zero.mp hz).symm + rw [sub_eq_add_neg, ← neg_smul, add_comm] at h + exact (Ne.symm (ne_of_lt (hp x hx))) (LinearIndependent.eq_zero_of_pair hL h).2 + +/-- Strict **Cauchy-Schwarz** is equivalent to linear independence for positive definite forms. -/ +lemma apply_mul_apply_lt_iff_linearIndependent [NoZeroSMulDivisors R M] + (hp : ∀ x, x ≠ 0 → 0 < B x x) (x y : M) : + (B x y) * (B y x) < (B x x) * (B y y) ↔ LinearIndependent R ![x, y] := by + have hle : ∀ z, 0 ≤ B z z := by + intro z + by_cases hz : z = 0; simp [hz] + exact le_of_lt (hp z hz) + constructor + · contrapose! + intro h + rw [LinearIndependent.pair_iff] at h + push_neg at h + obtain ⟨r, s, hl, h0⟩ := h + by_cases hr : r = 0; · simp_all + by_cases hs : s = 0; · simp_all + suffices + (B (r • x) (r • x)) * (B (s • y) (s • y)) = (B (r • x) (s • y)) * (B (s • y) (r • x)) by + simp only [map_smul, smul_apply, smul_eq_mul] at this + rw [show r * (r * (B x) x) * (s * (s * (B y) y)) = (r * r * s * s) * ((B x) x * (B y) y) by + ring, show s * (r * (B x) y) * (r * (s * (B y) x)) = (r * r * s * s) * ((B x) y * (B y) x) + by ring] at this + have hrs : r * r * s * s ≠ 0 := by simp [hr, hs] + exact le_of_eq <| mul_right_injective₀ hrs this + simp [show s • y = - r • x by rwa [neg_smul, ← add_eq_zero_iff_eq_neg']] + · contrapose! + intro h + refine not_linearIndependent_of_apply_mul_apply_eq B hp x y (le_antisymm + (apply_mul_apply_le_of_forall_zero_le B hle x y) h) + +/-- Strict **Cauchy-Schwarz** is equivalent to linear independence for positive definite symmetric +forms. -/ +lemma apply_sq_lt_iff_linearIndependent_of_symm [NoZeroSMulDivisors R M] + (hp : ∀ x, x ≠ 0 → 0 < B x x) (hB: B.IsSymm) (x y : M) : + (B x y) ^ 2 < (B x x) * (B y y) ↔ LinearIndependent R ![x, y] := by + rw [show (B x y) ^ 2 = (B x y) * (B y x) by rw [sq, ← hB, RingHom.id_apply]] + exact apply_mul_apply_lt_iff_linearIndependent B hp x y + +end BilinForm + end LinearMap diff --git a/Mathlib/LinearAlgebra/Span/Basic.lean b/Mathlib/LinearAlgebra/Span/Basic.lean index 0a1c4349e44cc..997c8ed3189d4 100644 --- a/Mathlib/LinearAlgebra/Span/Basic.lean +++ b/Mathlib/LinearAlgebra/Span/Basic.lean @@ -80,6 +80,57 @@ theorem span_preimage_le (f : F) (s : Set M₂) : alias _root_.LinearMap.span_preimage_le := Submodule.span_preimage_le +section + +variable {N : Type*} [AddCommMonoid N] [Module R N] + +lemma linearMap_eq_iff_of_eq_span {V : Submodule R M} (f g : V →ₗ[R] N) + {S : Set M} (hV : V = span R S) : + f = g ↔ ∀ (s : S), f ⟨s, by simpa only [hV] using subset_span (by simp)⟩ = + g ⟨s, by simpa only [hV] using subset_span (by simp)⟩ := by + constructor + · rintro rfl _ + rfl + · intro h + subst hV + suffices ∀ (x : M) (hx : x ∈ span R S), f ⟨x, hx⟩ = g ⟨x, hx⟩ by + ext ⟨x, hx⟩ + exact this x hx + intro x hx + induction hx using span_induction with + | mem x hx => exact h ⟨x, hx⟩ + | zero => erw [map_zero, map_zero] + | add x y hx hy hx' hy' => + erw [f.map_add ⟨x, hx⟩ ⟨y, hy⟩, g.map_add ⟨x, hx⟩ ⟨y, hy⟩] + rw [hx', hy'] + | smul a x hx hx' => + erw [f.map_smul a ⟨x, hx⟩, g.map_smul a ⟨x, hx⟩] + rw [hx'] + +lemma linearMap_eq_iff_of_span_eq_top (f g : M →ₗ[R] N) + {S : Set M} (hM : span R S = ⊤) : + f = g ↔ ∀ (s : S), f s = g s := by + convert linearMap_eq_iff_of_eq_span (f.comp (Submodule.subtype _)) + (g.comp (Submodule.subtype _)) hM.symm + constructor + · rintro rfl + rfl + · intro h + ext x + exact DFunLike.congr_fun h ⟨x, by simp⟩ + +lemma linearMap_eq_zero_iff_of_span_eq_top (f : M →ₗ[R] N) + {S : Set M} (hM : span R S = ⊤) : + f = 0 ↔ ∀ (s : S), f s = 0 := + linearMap_eq_iff_of_span_eq_top f 0 hM + +lemma linearMap_eq_zero_iff_of_eq_span {V : Submodule R M} (f : V →ₗ[R] N) + {S : Set M} (hV : V = span R S) : + f = 0 ↔ ∀ (s : S), f ⟨s, by simpa only [hV] using subset_span (by simp)⟩ = 0 := + linearMap_eq_iff_of_eq_span f 0 hV + +end + /-- See `Submodule.span_smul_eq` (in `RingTheory.Ideal.Operations`) for `span R (r • s) = r • span R s` that holds for arbitrary `r` in a `CommSemiring`. -/ theorem span_smul_eq_of_isUnit (s : Set M) (r : R) (hr : IsUnit r) : span R (r • s) = span R s := by diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 13c8fd29ba91a..d63de1913084e 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -353,6 +353,10 @@ theorem xor_iff_iff_not : Xor' a b ↔ (a ↔ ¬b) := by simp only [← @xor_not theorem xor_iff_not_iff' : Xor' a b ↔ (¬a ↔ b) := by simp only [← @xor_not_left _ b, not_not] +theorem xor_iff_or_and_not_and (a b : Prop) : Xor' a b ↔ (a ∨ b) ∧ (¬ (a ∧ b)) := by + rw [Xor', or_and_right, not_and_or, and_or_left, and_not_self_iff, false_or, + and_or_left, and_not_self_iff, or_false] + end Propositional /-! ### Declarations about equality -/ diff --git a/Mathlib/Logic/Equiv/TransferInstance.lean b/Mathlib/Logic/Equiv/TransferInstance.lean index e1e5ce7366f08..2b78ff6212187 100644 --- a/Mathlib/Logic/Equiv/TransferInstance.lean +++ b/Mathlib/Logic/Equiv/TransferInstance.lean @@ -621,10 +621,9 @@ protected abbrev algebra (e : α ≃ β) [Semiring β] : simp only [apply_symm_apply, Algebra.mul_smul_comm] lemma algebraMap_def (e : α ≃ β) [Semiring β] [Algebra R β] (r : R) : - let _ := Equiv.semiring e - let _ := Equiv.algebra R e - (algebraMap R α) r = e.symm ((algebraMap R β) r) := by - intros + (@algebraMap R α _ (Equiv.semiring e) (Equiv.algebra R e)) r = e.symm ((algebraMap R β) r) := by + let _ := Equiv.semiring e + let _ := Equiv.algebra R e simp only [Algebra.algebraMap_eq_smul_one] show e.symm (r • e 1) = e.symm (r • 1) simp only [Equiv.one_def, apply_symm_apply] diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 4c9f13b1bf99b..707d7858aebbe 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -735,7 +735,7 @@ theorem memℒp_of_bounded [IsFiniteMeasure μ] have hb : ∀ᵐ x ∂μ, f x ≤ b := h.mono fun ω h => h.2 (memℒp_const (max |a| |b|)).mono' hX (by filter_upwards [ha, hb] with x using abs_le_max_abs_abs) -@[mono] +@[gcongr, mono] theorem eLpNorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) : eLpNorm' f q ν ≤ eLpNorm' f q μ := by simp_rw [eLpNorm'] @@ -745,7 +745,7 @@ theorem eLpNorm'_mono_measure (f : α → F) (hμν : ν ≤ μ) (hq : 0 ≤ q) @[deprecated (since := "2024-07-27")] alias snorm'_mono_measure := eLpNorm'_mono_measure -@[mono] +@[gcongr, mono] theorem eLpNormEssSup_mono_measure (f : α → F) (hμν : ν ≪ μ) : eLpNormEssSup f ν ≤ eLpNormEssSup f μ := by simp_rw [eLpNormEssSup] @@ -754,7 +754,7 @@ theorem eLpNormEssSup_mono_measure (f : α → F) (hμν : ν ≪ μ) : @[deprecated (since := "2024-07-27")] alias snormEssSup_mono_measure := eLpNormEssSup_mono_measure -@[mono] +@[gcongr, mono] theorem eLpNorm_mono_measure (f : α → F) (hμν : ν ≤ μ) : eLpNorm f p ν ≤ eLpNorm f p μ := by by_cases hp0 : p = 0 · simp [hp0] diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 74809c39984b2..f8ef14327a38d 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -665,7 +665,7 @@ theorem mem_image_of_mem_range_restrict {r : β} {s : Set α} {f : α →ₛ β} rw [restrict_of_not_measurable hs] at hr exact (h0 <| eq_zero_of_mem_range_zero hr).elim -@[mono] +@[gcongr, mono] theorem restrict_mono [Preorder β] (s : Set α) {f g : α →ₛ β} (H : f ≤ g) : f.restrict s ≤ g.restrict s := if hs : MeasurableSet s then fun x => by diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index a3a28a6bca1d4..00aeccd87421d 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1334,7 +1334,7 @@ theorem integral_mono_ae {f g : α → ℝ} (hf : Integrable f μ) (hg : Integra exact setToFun_mono (dominatedFinMeasAdditive_weightedSMul μ) (fun s _ _ => weightedSMul_nonneg s) hf hg h -@[mono] +@[gcongr, mono] theorem integral_mono {f g : α → ℝ} (hf : Integrable f μ) (hg : Integrable g μ) (h : f ≤ g) : ∫ a, f a ∂μ ≤ ∫ a, g a ∂μ := integral_mono_ae hf hg <| Eventually.of_forall h diff --git a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean index 9740edb400613..a326a634d81ca 100644 --- a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean +++ b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean @@ -589,7 +589,7 @@ theorem continuous_parametric_primitive_of_continuous (uIcc_subset_Icc ⟨a_lt.1.le, lt_b.1.le⟩ ⟨a_lt.2.le, lt_b.2.le⟩) exact Eventually.of_forall this _ ≤ ∫ t in Icc (b₀ - δ) (b₀ + δ), M + 1 ∂μ + ∫ _t in Icc a b, δ ∂μ := by - gcongr + gcongr ?_ + ?_ · apply setIntegral_mono_on · exact (hf.uncurry_left _).norm.integrableOn_Icc · exact continuous_const.integrableOn_Icc diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index aff81c505d96f..e5d534014c897 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -78,7 +78,7 @@ theorem SimpleFunc.lintegral_eq_lintegral {m : MeasurableSpace α} (f : α → exact le_antisymm (iSup₂_le fun g hg => lintegral_mono hg <| le_rfl) (le_iSup₂_of_le f le_rfl le_rfl) -@[mono] +@[gcongr, mono] theorem lintegral_mono' {m : MeasurableSpace α} ⦃μ ν : Measure α⦄ (hμν : μ ≤ ν) ⦃f g : α → ℝ≥0∞⦄ (hfg : f ≤ g) : ∫⁻ a, f a ∂μ ≤ ∫⁻ a, g a ∂ν := by rw [lintegral, lintegral] diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 9693de1e0a8bf..f0a6159c042b5 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -12,6 +12,7 @@ import Mathlib.GroupTheory.Coset.Defs import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.Order.Filter.SmallSets import Mathlib.Order.LiminfLimsup +import Mathlib.Order.Filter.AtTopBot.CountablyGenerated import Mathlib.Tactic.FinCases /-! diff --git a/Mathlib/MeasureTheory/Measure/Stieltjes.lean b/Mathlib/MeasureTheory/Measure/Stieltjes.lean index 262bd709454e1..a381f2f16d3fb 100644 --- a/Mathlib/MeasureTheory/Measure/Stieltjes.lean +++ b/Mathlib/MeasureTheory/Measure/Stieltjes.lean @@ -129,6 +129,8 @@ instance : Module ℝ≥0 StieltjesFunction where add_smul _ _ _ := ext fun _ ↦ add_mul _ _ _ zero_smul _ := ext fun _ ↦ zero_mul _ +@[simp] lemma zero_apply (x : ℝ) : (0 : StieltjesFunction) x = 0 := rfl + @[simp] lemma add_apply (f g : StieltjesFunction) (x : ℝ) : (f + g) x = f x + g x := rfl /-- If a function `f : ℝ → ℝ` is monotone, then the function mapping `x` to the right limit of `f` @@ -430,6 +432,12 @@ theorem measure_Iic {l : ℝ} (hf : Tendsto f atBot (𝓝 l)) (x : ℝ) : simp_rw [measure_Ioc] exact ENNReal.tendsto_ofReal (Tendsto.const_sub _ hf) +lemma measure_Iio {l : ℝ} (hf : Tendsto f atBot (𝓝 l)) (x : ℝ) : + f.measure (Iio x) = ofReal (leftLim f x - l) := by + rw [← Iic_diff_right, measure_diff _ (nullMeasurableSet_singleton x), measure_singleton, + f.measure_Iic hf, ← ofReal_sub _ (sub_nonneg.mpr <| Monotone.leftLim_le f.mono' le_rfl)] + <;> simp + theorem measure_Ici {l : ℝ} (hf : Tendsto f atTop (𝓝 l)) (x : ℝ) : f.measure (Ici x) = ofReal (l - leftLim f x) := by refine tendsto_nhds_unique (tendsto_measure_Ico_atTop _ _) ?_ @@ -441,12 +449,52 @@ theorem measure_Ici {l : ℝ} (hf : Tendsto f atTop (𝓝 l)) (x : ℝ) : rw [tendsto_atTop_atTop] exact fun y => ⟨y + 1, fun z hyz => by rwa [le_sub_iff_add_le]⟩ +lemma measure_Ioi {l : ℝ} (hf : Tendsto f atTop (𝓝 l)) (x : ℝ) : + f.measure (Ioi x) = ofReal (l - f x) := by + rw [← Ici_diff_left, measure_diff _ (nullMeasurableSet_singleton x), measure_singleton, + f.measure_Ici hf, ← ofReal_sub _ (sub_nonneg.mpr <| Monotone.leftLim_le f.mono' le_rfl)] + <;> simp + +lemma measure_Ioi_of_tendsto_atTop_atTop (hf : Tendsto f atTop atTop) (x : ℝ) : + f.measure (Ioi x) = ∞ := by + refine ENNReal.eq_top_of_forall_nnreal_le fun r ↦ ?_ + obtain ⟨N, hN⟩ := eventually_atTop.mp (tendsto_atTop.mp hf (r + f x)) + exact (f.measure_Ioc x (max x N) ▸ ENNReal.coe_nnreal_eq r ▸ (ENNReal.ofReal_le_ofReal <| + le_tsub_of_add_le_right <| hN _ (le_max_right x N))).trans (measure_mono Ioc_subset_Ioi_self) + +lemma measure_Ici_of_tendsto_atTop_atTop (hf : Tendsto f atTop atTop) (x : ℝ) : + f.measure (Ici x) = ∞ := by + rw [← top_le_iff, ← f.measure_Ioi_of_tendsto_atTop_atTop hf x] + exact measure_mono Ioi_subset_Ici_self + +lemma measure_Iic_of_tendsto_atBot_atBot (hf : Tendsto f atBot atBot) (x : ℝ) : + f.measure (Iic x) = ∞ := by + refine ENNReal.eq_top_of_forall_nnreal_le fun r ↦ ?_ + obtain ⟨N, hN⟩ := eventually_atBot.mp (tendsto_atBot.mp hf (f x - r)) + exact (f.measure_Ioc (min x N) x ▸ ENNReal.coe_nnreal_eq r ▸ (ENNReal.ofReal_le_ofReal <| + le_sub_comm.mp <| hN _ (min_le_right x N))).trans (measure_mono Ioc_subset_Iic_self) + +lemma measure_Iio_of_tendsto_atBot_atBot (hf : Tendsto f atBot atBot) (x : ℝ) : + f.measure (Iio x) = ∞ := by + rw [← top_le_iff, ← f.measure_Iic_of_tendsto_atBot_atBot hf (x - 1)] + exact measure_mono <| Set.Iic_subset_Iio.mpr <| sub_one_lt x + theorem measure_univ {l u : ℝ} (hfl : Tendsto f atBot (𝓝 l)) (hfu : Tendsto f atTop (𝓝 u)) : f.measure univ = ofReal (u - l) := by refine tendsto_nhds_unique (tendsto_measure_Iic_atTop _) ?_ simp_rw [measure_Iic f hfl] exact ENNReal.tendsto_ofReal (Tendsto.sub_const hfu _) +lemma measure_univ_of_tendsto_atTop_atTop (hf : Tendsto f atTop atTop) : + f.measure univ = ∞ := by + rw [← top_le_iff, ← f.measure_Ioi_of_tendsto_atTop_atTop hf 0] + exact measure_mono (subset_univ _) + +lemma measure_univ_of_tendsto_atBot_atBot (hf : Tendsto f atBot atBot) : + f.measure univ = ∞ := by + rw [← top_le_iff, ← f.measure_Iio_of_tendsto_atBot_atBot hf 0] + exact measure_mono (subset_univ _) + lemma isFiniteMeasure {l u : ℝ} (hfl : Tendsto f atBot (𝓝 l)) (hfu : Tendsto f atTop (𝓝 u)) : IsFiniteMeasure f.measure := ⟨by simp [f.measure_univ hfl hfu]⟩ @@ -490,9 +538,13 @@ lemma eq_of_measure_of_eq (g : StieltjesFunction) {y : ℝ} · rw [sub_nonneg] exact f.mono hxy +@[simp] +lemma measure_zero : StieltjesFunction.measure 0 = 0 := + Measure.ext_of_Ioc _ _ (by simp) + @[simp] lemma measure_const (c : ℝ) : (StieltjesFunction.const c).measure = 0 := - Measure.ext_of_Ioc _ _ (fun _ _ _ ↦ by simp) + Measure.ext_of_Ioc _ _ (by simp) @[simp] lemma measure_add (f g : StieltjesFunction) : (f + g).measure = f.measure + g.measure := by diff --git a/Mathlib/ModelTheory/Fraisse.lean b/Mathlib/ModelTheory/Fraisse.lean index ce7356ee4aeb5..11b8204c05fa6 100644 --- a/Mathlib/ModelTheory/Fraisse.lean +++ b/Mathlib/ModelTheory/Fraisse.lean @@ -29,8 +29,8 @@ Fraïssé limit - the countable ultrahomogeneous structure with that age. - A class `K` has `FirstOrder.Language.Amalgamation` when for any pair of embeddings of a structure `M` in `K` into other structures in `K`, those two structures can be embedded into a fourth structure in `K` such that the resulting square of embeddings commutes. -- `FirstOrder.Language.IsFraisse` indicates that a class is nonempty, isomorphism-invariant, - essentially countable, and satisfies the hereditary, joint embedding, and amalgamation properties. +- `FirstOrder.Language.IsFraisse` indicates that a class is nonempty, essentially countable, + and satisfies the hereditary, joint embedding, and amalgamation properties. - `FirstOrder.Language.IsFraisseLimit` indicates that a structure is a Fraïssé limit for a given class. @@ -109,12 +109,11 @@ def Amalgamation : Prop := M ∈ K → N ∈ K → P ∈ K → ∃ (Q : Bundled.{w} L.Structure) (NQ : N ↪[L] Q) (PQ : P ↪[L] Q), Q ∈ K ∧ NQ.comp MN = PQ.comp MP -/-- A Fraïssé class is a nonempty, isomorphism-invariant, essentially countable class of structures -satisfying the hereditary, joint embedding, and amalgamation properties. -/ +/-- A Fraïssé class is a nonempty, essentially countable class of structures satisfying the +hereditary, joint embedding, and amalgamation properties. -/ class IsFraisse : Prop where is_nonempty : K.Nonempty FG : ∀ M : Bundled.{w} L.Structure, M ∈ K → Structure.FG L M - is_equiv_invariant : ∀ M N : Bundled.{w} L.Structure, Nonempty (M ≃[L] N) → (M ∈ K ↔ N ∈ K) is_essentially_countable : (Quotient.mk' '' K).Countable hereditary : Hereditary K jointEmbedding : JointEmbedding K @@ -146,6 +145,10 @@ theorem Hereditary.is_equiv_invariant_of_fg (h : Hereditary K) ⟨fun MK => h M MK ((fg M MK).mem_age_of_equiv hn), fun NK => h N NK ((fg N NK).mem_age_of_equiv ⟨hn.some.symm⟩)⟩ +theorem IsFraisse.is_equiv_invariant [h : IsFraisse K] {M N : Bundled.{w} L.Structure} + (hn : Nonempty (M ≃[L] N)) : M ∈ K ↔ N ∈ K := + h.hereditary.is_equiv_invariant_of_fg h.FG M N hn + variable (M) theorem age.nonempty : (L.age M).Nonempty := @@ -221,7 +224,6 @@ theorem age_directLimit {ι : Type w} [Preorder ι] [IsDirected ι (· ≤ ·)] /-- Sufficient conditions for a class to be the age of a countably-generated structure. -/ theorem exists_cg_is_age_of (hn : K.Nonempty) - (h : ∀ M N : Bundled.{w} L.Structure, Nonempty (M ≃[L] N) → (M ∈ K ↔ N ∈ K)) (hc : (Quotient.mk' '' K).Countable) (fg : ∀ M : Bundled.{w} L.Structure, M ∈ K → Structure.FG L M) (hp : Hereditary K) (jep : JointEmbedding K) : ∃ M : Bundled.{w} L.Structure, Structure.CG L M ∧ L.age M = K := by @@ -234,7 +236,7 @@ theorem exists_cg_is_age_of (hn : K.Nonempty) -- Porting note: fix hP2 because `Quotient.out (Quotient.mk' x) ≈ a` was not simplified -- to `x ≈ a` in hF replace hP2 := Setoid.trans (Setoid.symm (Quotient.mk_out P)) hP2 - exact (h _ _ hP2).1 hP1 + exact (hp.is_equiv_invariant_of_fg fg _ _ hP2).1 hP1 choose P hPK hP hFP using fun (N : K) (n : ℕ) => jep N N.2 (F (n + 1)).out (hF' _) let G : ℕ → K := @Nat.rec (fun _ => K) ⟨(F 0).out, hF' 0⟩ fun n N => ⟨P N n, hPK N n⟩ -- Poting note: was @@ -265,8 +267,8 @@ theorem exists_countable_is_age_of_iff [Countable (Σ l, L.Functions l)] : · rintro ⟨M, h1, h2, rfl⟩ refine ⟨age.nonempty M, age.is_equiv_invariant L M, age.countable_quotient M, fun N hN => hN.1, age.hereditary M, age.jointEmbedding M⟩ - · rintro ⟨Kn, eqinv, cq, hfg, hp, jep⟩ - obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn eqinv cq hfg hp jep + · rintro ⟨Kn, _, cq, hfg, hp, jep⟩ + obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn cq hfg hp jep exact ⟨M, Structure.cg_iff_countable.1 hM, rfl⟩ variable (L) @@ -358,7 +360,7 @@ theorem IsUltrahomogeneous.amalgamation_age (h : L.IsUltrahomogeneous M) : theorem IsUltrahomogeneous.age_isFraisse [Countable M] (h : L.IsUltrahomogeneous M) : IsFraisse (L.age M) := - ⟨age.nonempty M, fun _ hN => hN.1, age.is_equiv_invariant L M, age.countable_quotient M, + ⟨age.nonempty M, fun _ hN => hN.1, age.countable_quotient M, age.hereditary M, age.jointEmbedding M, h.amalgamation_age⟩ namespace IsFraisseLimit diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 75203851ec43e..25239a538fd47 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -310,7 +310,7 @@ theorem realize_iff : (φ.iff ψ).Realize v xs ↔ (φ.Realize v xs ↔ ψ.Reali simp only [BoundedFormula.iff, realize_inf, realize_imp, and_imp, ← iff_def] theorem realize_castLE_of_eq {m n : ℕ} (h : m = n) {h' : m ≤ n} {φ : L.BoundedFormula α m} - {v : α → M} {xs : Fin n → M} : (φ.castLE h').Realize v xs ↔ φ.Realize v (xs ∘ cast h) := by + {v : α → M} {xs : Fin n → M} : (φ.castLE h').Realize v xs ↔ φ.Realize v (xs ∘ Fin.cast h) := by subst h simp only [castLE_rfl, cast_refl, OrderIso.coe_refl, Function.comp_id] diff --git a/Mathlib/NumberTheory/FLT/Three.lean b/Mathlib/NumberTheory/FLT/Three.lean index 7f70138544592..fc457ccd599f2 100644 --- a/Mathlib/NumberTheory/FLT/Three.lean +++ b/Mathlib/NumberTheory/FLT/Three.lean @@ -85,7 +85,7 @@ private lemma three_dvd_b_of_dvd_a_of_gcd_eq_one_of_case2 {a b c : ℤ} (ha : a rw [add_comm (a ^ 3), add_assoc, add_comm (a ^ 3), ← add_assoc] at HF refine isCoprime_of_gcd_eq_one_of_FLT ?_ HF convert Hgcd using 2 - rw [Finset.pair_comm, Finset.Insert.comm] + rw [Finset.pair_comm, Finset.insert_comm] by_contra! h3b by_cases h3c : 3 ∣ c · apply h3b @@ -132,15 +132,15 @@ theorem fermatLastTheoremThree_of_three_dvd_only_c rw [← sub_eq_zero, sub_eq_add_neg, ← (show Odd 3 by decide).neg_pow] at hF rcases h1 with (h3a | h3b) | h3c · refine fermatLastTheoremThree_of_dvd_a_of_gcd_eq_one_of_case2 ha h3a ?_ H hF - simp only [← Hgcd, Insert.comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] + simp only [← Hgcd, insert_comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] · rw [add_comm (a ^ 3)] at hF refine fermatLastTheoremThree_of_dvd_a_of_gcd_eq_one_of_case2 hb h3b ?_ H hF - simp only [← Hgcd, Insert.comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] + simp only [← Hgcd, insert_comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] · rw [add_comm _ ((-c) ^ 3), ← add_assoc] at hF refine fermatLastTheoremThree_of_dvd_a_of_gcd_eq_one_of_case2 (neg_ne_zero.2 hc) (by simp [h3c]) ?_ H hF - rw [Finset.Insert.comm (-c), Finset.pair_comm (-c) b] - simp only [← Hgcd, Insert.comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] + rw [Finset.insert_comm (-c), Finset.pair_comm (-c) b] + simp only [← Hgcd, insert_comm, gcd_insert, gcd_singleton, id_eq, ← abs_eq_normalize, abs_neg] section eisenstein @@ -210,8 +210,8 @@ variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] /-- For any `S' : Solution'`, the multiplicity of `λ` in `S'.c` is finite. -/ lemma Solution'.multiplicity_lambda_c_finite : - multiplicity.Finite (hζ.toInteger - 1) S'.c := - multiplicity.finite_of_not_isUnit hζ.zeta_sub_one_prime'.not_unit S'.hc + FiniteMultiplicity (hζ.toInteger - 1) S'.c := + .of_not_isUnit hζ.zeta_sub_one_prime'.not_unit S'.hc section DecidableRel @@ -328,14 +328,14 @@ lemma lambda_sq_dvd_or_dvd_or_dvd : by_contra! h rcases h with ⟨h1, h2, h3⟩ rw [← emultiplicity_lt_iff_not_dvd] at h1 h2 h3 - have h1' : multiplicity.Finite (hζ.toInteger - 1) (S'.a + S'.b) := - finite_iff_emultiplicity_ne_top.2 (fun ht ↦ by simp [ht] at h1) - have h2' : multiplicity.Finite (hζ.toInteger - 1) (S'.a + η * S'.b) := by - refine finite_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) + have h1' : FiniteMultiplicity (hζ.toInteger - 1) (S'.a + S'.b) := + finiteMultiplicity_iff_emultiplicity_ne_top.2 (fun ht ↦ by simp [ht] at h1) + have h2' : FiniteMultiplicity (hζ.toInteger - 1) (S'.a + η * S'.b) := by + refine finiteMultiplicity_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) rw [coe_eta] at ht simp [ht] at h2 - have h3' : multiplicity.Finite (hζ.toInteger - 1) (S'.a + η ^ 2 * S'.b) := by - refine finite_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) + have h3' : FiniteMultiplicity (hζ.toInteger - 1) (S'.a + η ^ 2 * S'.b) := by + refine finiteMultiplicity_iff_emultiplicity_ne_top.2 (fun ht ↦ ?_) rw [coe_eta] at ht simp [ht] at h3 rw [h1'.emultiplicity_eq_multiplicity, Nat.cast_lt] at h1 diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 85a6e500e4a62..ab751110fdab3 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -141,21 +141,24 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : refine Ideal.Quotient.eq_zero_iff_mem.mpr ?_ simp [mem_span_singleton] -namespace multiplicity - section IntegralDomain variable [IsDomain R] -theorem pow_sub_pow_of_prime {p : R} (hp : Prime p) {x y : R} (hxy : p ∣ x - y) (hx : ¬p ∣ x) - {n : ℕ} (hn : ¬p ∣ n) : emultiplicity p (x ^ n - y ^ n) = emultiplicity p (x - y) := by +theorem emultiplicity_pow_sub_pow_of_prime {p : R} (hp : Prime p) {x y : R} + (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} (hn : ¬p ∣ n) : + emultiplicity p (x ^ n - y ^ n) = emultiplicity p (x - y) := by rw [← geom_sum₂_mul, emultiplicity_mul hp, emultiplicity_eq_zero.2 (not_dvd_geom_sum₂ hp hxy hx hn), zero_add] +@[deprecated (since := "2024-11-30")] +alias multiplicity.pow_sub_pow_of_prime := emultiplicity_pow_sub_pow_of_prime + variable (hp : Prime (p : R)) (hp1 : Odd p) (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) include hp hp1 hxy hx -theorem geom_sum₂_eq_one : emultiplicity (↑p) (∑ i ∈ range p, x ^ i * y ^ (p - 1 - i)) = 1 := by +theorem emultiplicity_geom_sum₂_eq_one : + emultiplicity (↑p) (∑ i ∈ range p, x ^ i * y ^ (p - 1 - i)) = 1 := by rw [← Nat.cast_one] refine emultiplicity_eq_coe.2 ⟨?_, ?_⟩ · rw [pow_one] @@ -167,20 +170,29 @@ theorem geom_sum₂_eq_one : emultiplicity (↑p) (∑ i ∈ range p, x ^ i * y rw [pow_two, mul_dvd_mul_iff_left hp.ne_zero] exact mt hp.dvd_of_dvd_pow hx -theorem pow_prime_sub_pow_prime : +@[deprecated (since := "2024-11-30")] +alias multiplicity.geom_sum₂_eq_one := emultiplicity_geom_sum₂_eq_one + +theorem emultiplicity_pow_prime_sub_pow_prime : emultiplicity (↑p) (x ^ p - y ^ p) = emultiplicity (↑p) (x - y) + 1 := by - rw [← geom_sum₂_mul, emultiplicity_mul hp, geom_sum₂_eq_one hp hp1 hxy hx, add_comm] + rw [← geom_sum₂_mul, emultiplicity_mul hp, emultiplicity_geom_sum₂_eq_one hp hp1 hxy hx, add_comm] -theorem pow_prime_pow_sub_pow_prime_pow (a : ℕ) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.pow_prime_sub_pow_prime := emultiplicity_pow_prime_sub_pow_prime + +theorem emultiplicity_pow_prime_pow_sub_pow_prime_pow (a : ℕ) : emultiplicity (↑p) (x ^ p ^ a - y ^ p ^ a) = emultiplicity (↑p) (x - y) + a := by induction' a with a h_ind · rw [Nat.cast_zero, add_zero, pow_zero, pow_one, pow_one] rw [Nat.cast_add, Nat.cast_one, ← add_assoc, ← h_ind, pow_succ, pow_mul, pow_mul] - apply pow_prime_sub_pow_prime hp hp1 + apply emultiplicity_pow_prime_sub_pow_prime hp hp1 · rw [← geom_sum₂_mul] exact dvd_mul_of_dvd_right hxy _ · exact fun h => hx (hp.dvd_of_dvd_pow h) +@[deprecated (since := "2024-11-30")] +alias multiplicity.pow_prime_pow_sub_pow_prime_pow := emultiplicity_pow_prime_pow_sub_pow_prime_pow + end IntegralDomain section LiftingTheExponent @@ -189,17 +201,17 @@ variable (hp : Nat.Prime p) (hp1 : Odd p) include hp hp1 /-- **Lifting the exponent lemma** for odd primes. -/ -theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : +theorem Int.emultiplicity_pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : emultiplicity (↑p) (x ^ n - y ^ n) = emultiplicity (↑p) (x - y) + emultiplicity p n := by cases' n with n · simp only [emultiplicity_zero, add_top, pow_zero, sub_self] - have h : Finite _ _ := Nat.multiplicity_finite_iff.mpr ⟨hp.ne_one, n.succ_pos⟩ + have h : FiniteMultiplicity _ _ := Nat.finiteMultiplicity_iff.mpr ⟨hp.ne_one, n.succ_pos⟩ simp only [Nat.succ_eq_add_one] at h rcases emultiplicity_eq_coe.mp h.emultiplicity_eq_multiplicity with ⟨⟨k, hk⟩, hpn⟩ conv_lhs => rw [hk, pow_mul, pow_mul] rw [Nat.prime_iff_prime_int] at hp - rw [pow_sub_pow_of_prime hp, pow_prime_pow_sub_pow_prime_pow hp hp1 hxy hx, - h.emultiplicity_eq_multiplicity] + rw [emultiplicity_pow_sub_pow_of_prime hp, + emultiplicity_pow_prime_pow_sub_pow_prime_pow hp hp1 hxy hx, h.emultiplicity_eq_multiplicity] · rw [← geom_sum₂_mul] exact dvd_mul_of_dvd_right hxy _ · exact fun h => hx (hp.dvd_of_dvd_pow h) @@ -208,13 +220,20 @@ theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) ( refine hpn ⟨c, ?_⟩ rwa [pow_succ, mul_assoc] -theorem Int.pow_add_pow {x y : ℤ} (hxy : ↑p ∣ x + y) (hx : ¬↑p ∣ x) {n : ℕ} (hn : Odd n) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.Int.pow_sub_pow := Int.emultiplicity_pow_sub_pow + +theorem Int.emultiplicity_pow_add_pow {x y : ℤ} (hxy : ↑p ∣ x + y) (hx : ¬↑p ∣ x) + {n : ℕ} (hn : Odd n) : emultiplicity (↑p) (x ^ n + y ^ n) = emultiplicity (↑p) (x + y) + emultiplicity p n := by rw [← sub_neg_eq_add] at hxy rw [← sub_neg_eq_add, ← sub_neg_eq_add, ← Odd.neg_pow hn] - exact Int.pow_sub_pow hp hp1 hxy hx n + exact Int.emultiplicity_pow_sub_pow hp hp1 hxy hx n + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Int.pow_add_pow := Int.emultiplicity_pow_add_pow -theorem Nat.pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : ℕ) : +theorem Nat.emultiplicity_pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : ℕ) : emultiplicity p (x ^ n - y ^ n) = emultiplicity p (x - y) + emultiplicity p n := by obtain hyx | hyx := le_total y x · iterate 2 rw [← Int.natCast_emultiplicity] @@ -222,20 +241,25 @@ theorem Nat.pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : rw [← Int.natCast_dvd_natCast] at hxy hx rw [Int.natCast_sub hyx] at * push_cast at * - exact Int.pow_sub_pow hp hp1 hxy hx n + exact Int.emultiplicity_pow_sub_pow hp hp1 hxy hx n · simp only [Nat.sub_eq_zero_iff_le.mpr (Nat.pow_le_pow_left hyx n), emultiplicity_zero, Nat.sub_eq_zero_iff_le.mpr hyx, top_add] -theorem Nat.pow_add_pow {x y : ℕ} (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : Odd n) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.Nat.pow_sub_pow := Nat.emultiplicity_pow_sub_pow + +theorem Nat.emultiplicity_pow_add_pow {x y : ℕ} (hxy : p ∣ x + y) (hx : ¬p ∣ x) + {n : ℕ} (hn : Odd n) : emultiplicity p (x ^ n + y ^ n) = emultiplicity p (x + y) + emultiplicity p n := by iterate 2 rw [← Int.natCast_emultiplicity] rw [← Int.natCast_dvd_natCast] at hxy hx push_cast at * - exact Int.pow_add_pow hp hp1 hxy hx hn + exact Int.emultiplicity_pow_add_pow hp hp1 hxy hx hn -end LiftingTheExponent +@[deprecated (since := "2024-11-30")] +alias multiplicity.Nat.pow_add_pow := Nat.emultiplicity_pow_add_pow -end multiplicity +end LiftingTheExponent end CommRing @@ -294,10 +318,10 @@ theorem Int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬2 have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even cases' n with n · simp only [pow_zero, sub_self, emultiplicity_zero, Int.ofNat_zero, add_top] - have h : multiplicity.Finite 2 n.succ := Nat.multiplicity_finite_iff.mpr ⟨by norm_num, n.succ_pos⟩ + have h : FiniteMultiplicity 2 n.succ := Nat.finiteMultiplicity_iff.mpr ⟨by norm_num, n.succ_pos⟩ simp only [Nat.succ_eq_add_one] at h rcases emultiplicity_eq_coe.mp h.emultiplicity_eq_multiplicity with ⟨⟨k, hk⟩, hpn⟩ - rw [hk, pow_mul, pow_mul, multiplicity.pow_sub_pow_of_prime, + rw [hk, pow_mul, pow_mul, emultiplicity_pow_sub_pow_of_prime, Int.two_pow_two_pow_sub_pow_two_pow _ hxy hx, ← hk] · norm_cast rw [h.emultiplicity_eq_multiplicity] @@ -332,8 +356,8 @@ theorem Int.two_pow_sub_pow {x y : ℤ} {n : ℕ} (hxy : 2 ∣ x - y) (hx : ¬2 emultiplicity_mul Int.prime_two, emultiplicity_mul Int.prime_two] · suffices emultiplicity (2 : ℤ) ↑(2 : ℕ) = 1 by rw [this, add_comm 1, ← add_assoc] norm_cast - rw [multiplicity.Finite.emultiplicity_self] - rw [Nat.multiplicity_finite_iff] + rw [FiniteMultiplicity.emultiplicity_self] + rw [Nat.finiteMultiplicity_iff] decide · rw [← even_iff_two_dvd, Int.not_even_iff_odd] apply Odd.pow @@ -378,7 +402,7 @@ theorem pow_sub_pow (hyx : y < x) (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} padicValNat p (x ^ n - y ^ n) = padicValNat p (x - y) + padicValNat p n := by rw [← Nat.cast_inj (R := ℕ∞), Nat.cast_add] iterate 3 rw [padicValNat_eq_emultiplicity] - · exact multiplicity.Nat.pow_sub_pow hp.out hp1 hxy hx n + · exact Nat.emultiplicity_pow_sub_pow hp.out hp1 hxy hx n · exact hn.bot_lt · exact Nat.sub_pos_of_lt hyx · exact Nat.sub_pos_of_lt (Nat.pow_lt_pow_left hyx hn) @@ -389,7 +413,7 @@ theorem pow_add_pow (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : Odd n) · contradiction rw [← Nat.cast_inj (R := ℕ∞), Nat.cast_add] iterate 3 rw [padicValNat_eq_emultiplicity] - · exact multiplicity.Nat.pow_add_pow hp.out hp1 hxy hx hn + · exact Nat.emultiplicity_pow_add_pow hp.out hp1 hxy hx hn · exact Odd.pos hn · simp only [add_pos_iff, Nat.succ_pos', or_true] · exact Nat.lt_add_left _ (pow_pos y.succ_pos _) diff --git a/Mathlib/NumberTheory/NumberField/AdeleRing.lean b/Mathlib/NumberTheory/NumberField/AdeleRing.lean index 7c5add17d6901..1ee25ce15fd5f 100644 --- a/Mathlib/NumberTheory/NumberField/AdeleRing.lean +++ b/Mathlib/NumberTheory/NumberField/AdeleRing.lean @@ -47,11 +47,11 @@ variable (K : Type*) [Field K] The infinite adele ring is the finite product of completions of a number field over its infinite places. See `NumberField.InfinitePlace` for the definition of an infinite place and -`NumberField.InfinitePlace.completion` for the associated completion. +`NumberField.InfinitePlace.Completion` for the associated completion. -/ /-- The infinite adele ring of a number field. -/ -def InfiniteAdeleRing := (v : InfinitePlace K) → v.completion +def InfiniteAdeleRing := (v : InfinitePlace K) → v.Completion namespace InfiniteAdeleRing @@ -82,11 +82,11 @@ abbrev ringEquiv_mixedSpace : InfiniteAdeleRing K ≃+* mixedEmbedding.mixedSpace K := RingEquiv.trans (RingEquiv.piEquivPiSubtypeProd (fun (v : InfinitePlace K) => IsReal v) - (fun (v : InfinitePlace K) => v.completion)) + (fun (v : InfinitePlace K) => v.Completion)) (RingEquiv.prodCongr - (RingEquiv.piCongrRight (fun ⟨_, hv⟩ => Completion.ringEquiv_real_of_isReal hv)) + (RingEquiv.piCongrRight (fun ⟨_, hv⟩ => Completion.ringEquivRealOfIsReal hv)) (RingEquiv.trans - (RingEquiv.piCongrRight (fun v => Completion.ringEquiv_complex_of_isComplex + (RingEquiv.piCongrRight (fun v => Completion.ringEquivComplexOfIsComplex ((not_isReal_iff_isComplex.1 v.2)))) (RingEquiv.piCongrLeft (fun _ => ℂ) <| Equiv.subtypeEquivRight (fun _ => not_isReal_iff_isComplex)))) @@ -94,7 +94,7 @@ abbrev ringEquiv_mixedSpace : @[simp] theorem ringEquiv_mixedSpace_apply (x : InfiniteAdeleRing K) : ringEquiv_mixedSpace K x = - (fun (v : {w : InfinitePlace K // IsReal w}) => extensionEmbedding_of_isReal v.2 (x v), + (fun (v : {w : InfinitePlace K // IsReal w}) => extensionEmbeddingOfIsReal v.2 (x v), fun (v : {w : InfinitePlace K // IsComplex w}) => extensionEmbedding v.1 (x v)) := rfl /-- Transfers the embedding of `x ↦ (x)ᵥ` of the number field `K` into its infinite adele @@ -103,8 +103,8 @@ ring to the mixed embedding `x ↦ (φᵢ(x))ᵢ` of `K` into the space `ℝ ^ r theorem mixedEmbedding_eq_algebraMap_comp {x : K} : mixedEmbedding K x = ringEquiv_mixedSpace K (algebraMap K _ x) := by ext v <;> simp only [ringEquiv_mixedSpace_apply, algebraMap_apply, - ringEquiv_real_of_isReal, ringEquiv_complex_of_isComplex, extensionEmbedding, - extensionEmbedding_of_isReal, extensionEmbedding_of_comp, RingEquiv.coe_ofBijective, + ringEquivRealOfIsReal, ringEquivComplexOfIsComplex, extensionEmbedding, + extensionEmbeddingOfIsReal, extensionEmbedding_of_comp, RingEquiv.coe_ofBijective, RingHom.coe_mk, MonoidHom.coe_mk, OneHom.coe_mk, UniformSpace.Completion.extensionHom] · rw [UniformSpace.Completion.extension_coe (WithAbs.isUniformInducing_of_comp <| v.1.norm_embedding_of_isReal v.2).uniformContinuous x] diff --git a/Mathlib/NumberTheory/NumberField/Completion.lean b/Mathlib/NumberTheory/NumberField/Completion.lean index 901f318194033..0d23896da11ed 100644 --- a/Mathlib/NumberTheory/NumberField/Completion.lean +++ b/Mathlib/NumberTheory/NumberField/Completion.lean @@ -19,35 +19,35 @@ of instances is through the use of type synonyms. In this case, we use the type of a semiring. In particular this type synonym depends on an absolute value, which provides a systematic way of assigning and inferring instances of the semiring that also depend on an absolute value. The completion of a field at multiple absolute values is defined in -`Mathlib.Algebra.Ring.WithAbs` as `AbsoluteValue.completion`. The completion of a number +`Mathlib.Algebra.Ring.WithAbs` as `AbsoluteValue.Completion`. The completion of a number field at an infinite place is then derived in this file, as `InfinitePlace` is a subtype of `AbsoluteValue`. ## Main definitions - - `NumberField.InfinitePlace.completion` : the completion of a number field `K` at an infinite + - `NumberField.InfinitePlace.Completion` : the completion of a number field `K` at an infinite place, obtained by completing `K` with respect to the absolute value associated to the infinite place. - `NumberField.InfinitePlace.Completion.extensionEmbedding` : the embedding `v.embedding : K →+* ℂ` - extended to `v.completion →+* ℂ`. - - `NumberField.InfinitePlace.Completion.extensionEmbedding_of_isReal` : if the infinite place `v` + extended to `v.Completion →+* ℂ`. + - `NumberField.InfinitePlace.Completion.extensionEmbeddingOfIsReal` : if the infinite place `v` is real, then this extends the embedding `v.embedding_of_isReal : K →+* ℝ` to - `v.completion →+* ℝ`. + `v.Completion →+* ℝ`. - `NumberField.InfinitePlace.Completion.equiv_real_of_isReal` : the ring isomorphism - `v.completion ≃+* ℝ` when `v` is a real infinite place; the forward direction of this is - `extensionEmbedding_of_isReal`. + `v.Completion ≃+* ℝ` when `v` is a real infinite place; the forward direction of this is + `extensionEmbeddingOfIsReal`. - `NumberField.InfinitePlace.Completion.equiv_complex_of_isComplex` : the ring isomorphism - `v.completion ≃+* ℂ` when `v` is a complex infinite place; the forward direction of this is + `v.Completion ≃+* ℂ` when `v` is a complex infinite place; the forward direction of this is `extensionEmbedding`. ## Main results - `NumberField.Completion.locallyCompactSpace` : the completion of a number field at an infinite place is locally compact. - - `NumberField.Completion.isometry_extensionEmbedding` : the embedding `v.completion →+* ℂ` is + - `NumberField.Completion.isometry_extensionEmbedding` : the embedding `v.Completion →+* ℂ` is an isometry. See also `isometry_extensionEmbedding_of_isReal` for the corresponding result on - `v.completion →+* ℝ` when `v` is real. + `v.Completion →+* ℝ` when `v` is real. - `NumberField.Completion.bijective_extensionEmbedding_of_isComplex` : the embedding - `v.completion →+* ℂ` is bijective when `v` is complex. See also - `bijective_extensionEmebdding_of_isReal` for the corresponding result for `v.completion →+* ℝ` + `v.Completion →+* ℂ` is bijective when `v` is complex. See also + `bijective_extensionEmebdding_of_isReal` for the corresponding result for `v.Completion →+* ℝ` when `v` is real. ## Tags @@ -62,70 +62,75 @@ open AbsoluteValue.Completion variable {K : Type*} [Field K] (v : InfinitePlace K) /-- The completion of a number field at an infinite place. -/ -abbrev completion := v.1.completion +abbrev Completion := v.1.Completion + +@[deprecated (since := "2024-12-01")] alias completion := Completion namespace Completion -instance : NormedField v.completion := +instance : NormedField v.Completion := letI := (WithAbs.isUniformInducing_of_comp v.norm_embedding_eq).completableTopField UniformSpace.Completion.instNormedFieldOfCompletableTopField (WithAbs v.1) lemma norm_coe (x : WithAbs v.1) : - ‖(x : v.completion)‖ = v (WithAbs.equiv v.1 x) := + ‖(x : v.Completion)‖ = v (WithAbs.equiv v.1 x) := UniformSpace.Completion.norm_coe x -instance : Algebra K v.completion := - inferInstanceAs <| Algebra (WithAbs v.1) v.1.completion +instance : Algebra K v.Completion := + inferInstanceAs <| Algebra (WithAbs v.1) v.1.Completion /-- The coercion from the rationals to its completion along an infinite place is `Rat.cast`. -/ lemma WithAbs.ratCast_equiv (v : InfinitePlace ℚ) (x : WithAbs v.1) : - Rat.cast (WithAbs.equiv _ x) = (x : v.completion) := + Rat.cast (WithAbs.equiv _ x) = (x : v.Completion) := (eq_ratCast (UniformSpace.Completion.coeRingHom.comp (WithAbs.ringEquiv v.1).symm.toRingHom) x).symm lemma Rat.norm_infinitePlace_completion (v : InfinitePlace ℚ) (x : ℚ) : - ‖(x : v.completion)‖ = |x| := by + ‖(x : v.Completion)‖ = |x| := by rw [← (WithAbs.equiv v.1).apply_symm_apply x, WithAbs.ratCast_equiv, norm_coe, (WithAbs.equiv v.1).apply_symm_apply, Rat.infinitePlace_apply] /-- The completion of a number field at an infinite place is locally compact. -/ -instance locallyCompactSpace : LocallyCompactSpace (v.completion) := +instance locallyCompactSpace : LocallyCompactSpace (v.Completion) := AbsoluteValue.Completion.locallyCompactSpace v.norm_embedding_eq -/-- The embedding associated to an infinite place extended to an embedding `v.completion →+* ℂ`. -/ -def extensionEmbedding : v.completion →+* ℂ := extensionEmbedding_of_comp v.norm_embedding_eq +/-- The embedding associated to an infinite place extended to an embedding `v.Completion →+* ℂ`. -/ +def extensionEmbedding : v.Completion →+* ℂ := extensionEmbedding_of_comp v.norm_embedding_eq -/-- The embedding `K →+* ℝ` associated to a real infinite place extended to `v.completion →+* ℝ`. -/ -def extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.completion →+* ℝ := +/-- The embedding `K →+* ℝ` associated to a real infinite place extended to `v.Completion →+* ℝ`. -/ +def extensionEmbeddingOfIsReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion →+* ℝ := extensionEmbedding_of_comp <| v.norm_embedding_of_isReal hv +@[deprecated (since := "2024-12-07")] +noncomputable alias extensionEmbedding_of_isReal := extensionEmbeddingOfIsReal + @[simp] theorem extensionEmbedding_coe (x : K) : extensionEmbedding v x = v.embedding x := extensionEmbedding_of_comp_coe v.norm_embedding_eq x @[simp] theorem extensionEmbedding_of_isReal_coe {v : InfinitePlace K} (hv : IsReal v) (x : K) : - extensionEmbedding_of_isReal hv x = embedding_of_isReal hv x := + extensionEmbeddingOfIsReal hv x = embedding_of_isReal hv x := extensionEmbedding_of_comp_coe (v.norm_embedding_of_isReal hv) x -/-- The embedding `v.completion →+* ℂ` is an isometry. -/ +/-- The embedding `v.Completion →+* ℂ` is an isometry. -/ theorem isometry_extensionEmbedding : Isometry (extensionEmbedding v) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp v.norm_embedding_eq) -/-- The embedding `v.completion →+* ℝ` at a real infinite palce is an isometry. -/ +/-- The embedding `v.Completion →+* ℝ` at a real infinite palce is an isometry. -/ theorem isometry_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - Isometry (extensionEmbedding_of_isReal hv) := + Isometry (extensionEmbeddingOfIsReal hv) := Isometry.of_dist_eq (extensionEmbedding_dist_eq_of_comp <| v.norm_embedding_of_isReal hv) -/-- The embedding `v.completion →+* ℂ` has closed image inside `ℂ`. -/ +/-- The embedding `v.Completion →+* ℂ` has closed image inside `ℂ`. -/ theorem isClosed_image_extensionEmbedding : IsClosed (Set.range (extensionEmbedding v)) := (isClosedEmbedding_extensionEmbedding_of_comp v.norm_embedding_eq).isClosed_range -/-- The embedding `v.completion →+* ℝ` associated to a real infinite place has closed image +/-- The embedding `v.Completion →+* ℝ` associated to a real infinite place has closed image inside `ℝ`. -/ theorem isClosed_image_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - IsClosed (Set.range (extensionEmbedding_of_isReal hv)) := + IsClosed (Set.range (extensionEmbeddingOfIsReal hv)) := (isClosedEmbedding_extensionEmbedding_of_comp <| v.norm_embedding_of_isReal hv).isClosed_range theorem subfield_ne_real_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : @@ -136,49 +141,61 @@ theorem subfield_ne_real_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : obtain ⟨r, hr⟩ := hv ▸ extensionEmbedding_coe v x ▸ RingHom.mem_fieldRange_self _ _ simp only [ComplexEmbedding.conjugate_coe_eq, ← hr, Complex.ofRealHom_eq_coe, Complex.conj_ofReal] -/-- If `v` is a complex infinite place, then the embedding `v.completion →+* ℂ` is surjective. -/ +/-- If `v` is a complex infinite place, then the embedding `v.Completion →+* ℂ` is surjective. -/ theorem surjective_extensionEmbedding_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : Function.Surjective (extensionEmbedding v) := by rw [← RingHom.fieldRange_eq_top_iff] exact (Complex.subfield_eq_of_closed <| isClosed_image_extensionEmbedding v).resolve_left <| subfield_ne_real_of_isComplex hv -/-- If `v` is a complex infinite place, then the embedding `v.completion →+* ℂ` is bijective. -/ +/-- If `v` is a complex infinite place, then the embedding `v.Completion →+* ℂ` is bijective. -/ theorem bijective_extensionEmbedding_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : Function.Bijective (extensionEmbedding v) := ⟨(extensionEmbedding v).injective, surjective_extensionEmbedding_of_isComplex hv⟩ -/-- The ring isomorphism `v.completion ≃+* ℂ`, when `v` is complex, given by the bijection -`v.completion →+* ℂ`. -/ -def ringEquiv_complex_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : - v.completion ≃+* ℂ := +/-- The ring isomorphism `v.Completion ≃+* ℂ`, when `v` is complex, given by the bijection +`v.Completion →+* ℂ`. -/ +def ringEquivComplexOfIsComplex {v : InfinitePlace K} (hv : IsComplex v) : + v.Completion ≃+* ℂ := RingEquiv.ofBijective _ (bijective_extensionEmbedding_of_isComplex hv) -/-- If the infinite place `v` is complex, then `v.completion` is isometric to `ℂ`. -/ -def isometryEquiv_complex_of_isComplex {v : InfinitePlace K} (hv : IsComplex v) : - v.completion ≃ᵢ ℂ where - toEquiv := ringEquiv_complex_of_isComplex hv +@[deprecated (since := "2024-12-07")] +noncomputable alias ringEquiv_complex_of_isComplex := ringEquivComplexOfIsComplex + +/-- If the infinite place `v` is complex, then `v.Completion` is isometric to `ℂ`. -/ +def isometryEquivComplexOfIsComplex {v : InfinitePlace K} (hv : IsComplex v) : + v.Completion ≃ᵢ ℂ where + toEquiv := ringEquivComplexOfIsComplex hv isometry_toFun := isometry_extensionEmbedding v -/-- If `v` is a real infinite place, then the embedding `v.completion →+* ℝ` is surjective. -/ +@[deprecated (since := "2024-12-07")] +noncomputable alias isometryEquiv_complex_of_isComplex := isometryEquivComplexOfIsComplex + +/-- If `v` is a real infinite place, then the embedding `v.Completion →+* ℝ` is surjective. -/ theorem surjective_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - Function.Surjective (extensionEmbedding_of_isReal hv) := by + Function.Surjective (extensionEmbeddingOfIsReal hv) := by rw [← RingHom.fieldRange_eq_top_iff, ← Real.subfield_eq_of_closed] exact isClosed_image_extensionEmbedding_of_isReal hv -/-- If `v` is a real infinite place, then the embedding `v.completion →+* ℝ` is bijective. -/ +/-- If `v` is a real infinite place, then the embedding `v.Completion →+* ℝ` is bijective. -/ theorem bijective_extensionEmbedding_of_isReal {v : InfinitePlace K} (hv : IsReal v) : - Function.Bijective (extensionEmbedding_of_isReal hv) := - ⟨(extensionEmbedding_of_isReal hv).injective, surjective_extensionEmbedding_of_isReal hv⟩ + Function.Bijective (extensionEmbeddingOfIsReal hv) := + ⟨(extensionEmbeddingOfIsReal hv).injective, surjective_extensionEmbedding_of_isReal hv⟩ -/-- The ring isomorphism `v.completion ≃+* ℝ`, when `v` is real, given by the bijection -`v.completion →+* ℝ`. -/ -def ringEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.completion ≃+* ℝ := +/-- The ring isomorphism `v.Completion ≃+* ℝ`, when `v` is real, given by the bijection +`v.Completion →+* ℝ`. -/ +def ringEquivRealOfIsReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃+* ℝ := RingEquiv.ofBijective _ (bijective_extensionEmbedding_of_isReal hv) -/-- If the infinite place `v` is real, then `v.completion` is isometric to `ℝ`. -/ -def isometryEquiv_real_of_isReal {v : InfinitePlace K} (hv : IsReal v) : v.completion ≃ᵢ ℝ where - toEquiv := ringEquiv_real_of_isReal hv +@[deprecated (since := "2024-12-07")] +noncomputable alias ringEquiv_real_of_isReal := ringEquivRealOfIsReal + +/-- If the infinite place `v` is real, then `v.Completion` is isometric to `ℝ`. -/ +def isometryEquivRealOfIsReal {v : InfinitePlace K} (hv : IsReal v) : v.Completion ≃ᵢ ℝ where + toEquiv := ringEquivRealOfIsReal hv isometry_toFun := isometry_extensionEmbedding_of_isReal hv +@[deprecated (since := "2024-12-07")] +noncomputable alias isometryEquiv_real_of_isReal := isometryEquivRealOfIsReal + end NumberField.InfinitePlace.Completion diff --git a/Mathlib/NumberTheory/NumberField/FinitePlaces.lean b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean new file mode 100644 index 0000000000000..834648c1e128e --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/FinitePlaces.lean @@ -0,0 +1,236 @@ +/- +Copyright (c) 2024 Fabrizio Barroero. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Fabrizio Barroero +-/ +import Mathlib.Data.Int.WithZero +import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.RingTheory.DedekindDomain.AdicValuation +import Mathlib.RingTheory.DedekindDomain.Factorization +import Mathlib.RingTheory.Ideal.Norm.AbsNorm +import Mathlib.Topology.Algebra.Valued.NormedValued + +/-! +# Finite places of number fields +This file defines finite places of a number field `K` as absolute values coming from an embedding +into a completion of `K` associated to a non-zero prime ideal of `𝓞 K`. + +## Main Definitions and Results +* `NumberField.vadicAbv`: a `v`-adic absolute value on `K`. +* `NumberField.FinitePlace`: the type of finite places of a number field `K`. +* `NumberField.FinitePlace.mulSupport_finite`: the `v`-adic absolute value of a non-zero element of +`K` is different from 1 for at most finitely many `v`. + +## Tags +number field, places, finite places +-/ + +open Ideal IsDedekindDomain HeightOneSpectrum NumberField WithZeroMulInt + +namespace NumberField + +section absoluteValue + +variable {K : Type*} [Field K] [NumberField K] (v : HeightOneSpectrum (𝓞 K)) + +/-- The norm of a maximal ideal as an element of `ℝ≥0` is `> 1` -/ +lemma one_lt_norm : 1 < (absNorm v.asIdeal : NNReal) := by + norm_cast + by_contra! h + apply IsPrime.ne_top v.isPrime + rw [← absNorm_eq_one_iff] + have : 0 < absNorm v.asIdeal := by + rw [Nat.pos_iff_ne_zero, absNorm_ne_zero_iff] + exact (v.asIdeal.fintypeQuotientOfFreeOfNeBot v.ne_bot).finite + omega + +private lemma norm_ne_zero : (absNorm v.asIdeal : NNReal) ≠ 0 := ne_zero_of_lt (one_lt_norm v) + +/-- The `v`-adic absolute value on `K` defined as the norm of `v` raised to negative `v`-adic +valuation.-/ +noncomputable def vadicAbv : AbsoluteValue K ℝ where + toFun x := toNNReal (norm_ne_zero v) (v.valuation x) + map_mul' _ _ := by simp only [_root_.map_mul, NNReal.coe_mul] + nonneg' _ := NNReal.zero_le_coe + eq_zero' _ := by simp only [NNReal.coe_eq_zero, map_eq_zero] + add_le' x y := by + -- the triangle inequality is implied by the ultrametric one + apply le_trans _ <| max_le_add_of_nonneg (zero_le ((toNNReal (norm_ne_zero v)) (v.valuation x))) + (zero_le ((toNNReal (norm_ne_zero v)) (v.valuation y))) + have h_mono := (toNNReal_strictMono (one_lt_norm v)).monotone + rw [← h_mono.map_max] --max goes inside withZeroMultIntToNNReal + exact h_mono (v.valuation.map_add x y) + +theorem vadicAbv_def {x : K} : vadicAbv v x = toNNReal (norm_ne_zero v) (v.valuation x) := rfl + +end absoluteValue + +section FinitePlace +variable {K : Type*} [Field K] [NumberField K] (v : HeightOneSpectrum (𝓞 K)) + +/-- The embedding of a number field inside its completion with respect to `v`. -/ +def embedding : K →+* adicCompletion K v := + @UniformSpace.Completion.coeRingHom K _ v.adicValued.toUniformSpace _ _ + +noncomputable instance instRankOneValuedAdicCompletion : + Valuation.RankOne (valuedAdicCompletion K v).v where + hom := { + toFun := toNNReal (norm_ne_zero v) + map_zero' := rfl + map_one' := rfl + map_mul' := MonoidWithZeroHom.map_mul (toNNReal (norm_ne_zero v)) + } + strictMono' := toNNReal_strictMono (one_lt_norm v) + nontrivial' := by + rcases Submodule.exists_mem_ne_zero_of_ne_bot v.ne_bot with ⟨x, hx1, hx2⟩ + use (x : K) + rw [valuedAdicCompletion_eq_valuation' v (x : K)] + constructor + · simpa only [ne_eq, map_eq_zero, NoZeroSMulDivisors.algebraMap_eq_zero_iff] + · apply ne_of_lt + rw [valuation_eq_intValuationDef, intValuation_lt_one_iff_dvd] + exact dvd_span_singleton.mpr hx1 + +/-- The `v`-adic completion of `K` is a normed field. -/ +noncomputable instance instNormedFieldValuedAdicCompletion : NormedField (adicCompletion K v) := + Valued.toNormedField (adicCompletion K v) (WithZero (Multiplicative ℤ)) + +/-- A finite place of a number field `K` is a place associated to an embedding into a completion +with respect to a maximal ideal. -/ +def FinitePlace (K : Type*) [Field K] [NumberField K] := + {w : AbsoluteValue K ℝ // ∃ v : HeightOneSpectrum (𝓞 K), place (embedding v) = w} + +/-- Return the finite place defined by a maximal ideal `v`. -/ +noncomputable def FinitePlace.mk (v : HeightOneSpectrum (𝓞 K)) : FinitePlace K := + ⟨place (embedding v), ⟨v, rfl⟩⟩ + +lemma toNNReal_Valued_eq_vadicAbv (x : K) : + toNNReal (norm_ne_zero v) (Valued.v (self:=v.adicValued) x) = vadicAbv v x := rfl + +/-- The norm of the image after the embedding associated to `v` is equal to the `v`-adic absolute +value. -/ +theorem FinitePlace.norm_def (x : K) : ‖embedding v x‖ = vadicAbv v x := by + simp only [NormedField.toNorm, instNormedFieldValuedAdicCompletion, Valued.toNormedField, + instFieldAdicCompletion, Valued.norm, Valuation.RankOne.hom, MonoidWithZeroHom.coe_mk, + ZeroHom.coe_mk, embedding, UniformSpace.Completion.coeRingHom, RingHom.coe_mk, MonoidHom.coe_mk, + OneHom.coe_mk, Valued.valuedCompletion_apply, toNNReal_Valued_eq_vadicAbv] + +/-- The norm of the image after the embedding associated to `v` is equal to the norm of `v` raised +to the power of the `v`-adic valuation. -/ +theorem FinitePlace.norm_def' (x : K) : ‖embedding v x‖ = toNNReal (norm_ne_zero v) + (v.valuation x) := by + rw [norm_def, vadicAbv_def] + +/-- The norm of the image after the embedding associated to `v` is equal to the norm of `v` raised +to the power of the `v`-adic valuation for integers. -/ +theorem FinitePlace.norm_def_int (x : 𝓞 K) : ‖embedding v x‖ = toNNReal (norm_ne_zero v) + (v.intValuationDef x) := by + rw [norm_def, vadicAbv_def, valuation_eq_intValuationDef] + +open FinitePlace + +/-- The `v`-adic norm of an integer is at most 1. -/ +theorem norm_le_one (x : 𝓞 K) : ‖embedding v x‖ ≤ 1 := by + rw [norm_def', NNReal.coe_le_one, toNNReal_le_one_iff (one_lt_norm v)] + exact valuation_le_one v x + +/-- The `v`-adic norm of an integer is 1 if and only if it is not in the ideal. -/ +theorem norm_eq_one_iff_not_mem (x : 𝓞 K) : ‖(embedding v) x‖ = 1 ↔ x ∉ v.asIdeal := by + rw [norm_def_int, NNReal.coe_eq_one, toNNReal_eq_one_iff (v.intValuationDef x) + (norm_ne_zero v) (one_lt_norm v).ne', ← dvd_span_singleton, + ← intValuation_lt_one_iff_dvd, not_lt] + exact (intValuation_le_one v x).ge_iff_eq.symm + +/-- The `v`-adic norm of an integer is less than 1 if and only if it is in the ideal. -/ +theorem norm_lt_one_iff_mem (x : 𝓞 K) : ‖embedding v x‖ < 1 ↔ x ∈ v.asIdeal := by + rw [norm_def_int, NNReal.coe_lt_one, toNNReal_lt_one_iff (one_lt_norm v), + intValuation_lt_one_iff_dvd] + exact dvd_span_singleton + +end FinitePlace + +namespace FinitePlace +variable {K : Type*} [Field K] [NumberField K] + +instance : FunLike (FinitePlace K) K ℝ where + coe w x := w.1 x + coe_injective' _ _ h := Subtype.eq (AbsoluteValue.ext <| congr_fun h) + +instance : MonoidWithZeroHomClass (FinitePlace K) K ℝ where + map_mul w := w.1.map_mul + map_one w := w.1.map_one + map_zero w := w.1.map_zero + +instance : NonnegHomClass (FinitePlace K) K ℝ where + apply_nonneg w := w.1.nonneg + +@[simp] +theorem apply (v : HeightOneSpectrum (𝓞 K)) (x : K) : mk v x = ‖embedding v x‖ := rfl + +/-- For a finite place `w`, return a maximal ideal `v` such that `w = finite_place v` . -/ +noncomputable def maximalIdeal (w : FinitePlace K) : HeightOneSpectrum (𝓞 K) := w.2.choose + +@[simp] +theorem mk_maximalIdeal (w : FinitePlace K) : mk (maximalIdeal w) = w := Subtype.ext w.2.choose_spec + +@[simp] +theorem norm_embedding_eq (w : FinitePlace K) (x : K) : + ‖embedding (maximalIdeal w) x‖ = w x := by + conv_rhs => rw [← mk_maximalIdeal w, apply] + +theorem pos_iff {w : FinitePlace K} {x : K} : 0 < w x ↔ x ≠ 0 := AbsoluteValue.pos_iff w.1 + +@[simp] +theorem mk_eq_iff {v₁ v₂ : HeightOneSpectrum (𝓞 K)} : mk v₁ = mk v₂ ↔ v₁ = v₂ := by + refine ⟨?_, fun a ↦ by rw [a]⟩ + contrapose! + intro h + rw [DFunLike.ne_iff] + have ⟨x, hx1, hx2⟩ : ∃ x : 𝓞 K, x ∈ v₁.asIdeal ∧ x ∉ v₂.asIdeal := by + by_contra! H + exact h <| HeightOneSpectrum.ext_iff.mpr <| IsMaximal.eq_of_le (isMaximal v₁) IsPrime.ne_top' H + use x + simp only [apply] + rw [← norm_lt_one_iff_mem ] at hx1 + rw [← norm_eq_one_iff_not_mem] at hx2 + linarith + +theorem maximalIdeal_mk (v : HeightOneSpectrum (𝓞 K)) : maximalIdeal (mk v) = v := by + rw [← mk_eq_iff, mk_maximalIdeal] + +lemma maximalIdeal_injective : (fun w : FinitePlace K ↦ maximalIdeal w).Injective := + Function.HasLeftInverse.injective ⟨mk, mk_maximalIdeal⟩ + +lemma maximalIdeal_inj (w₁ w₂ : FinitePlace K) : maximalIdeal w₁ = maximalIdeal w₂ ↔ w₁ = w₂ := + maximalIdeal_injective.eq_iff + +theorem mulSupport_finite_int {x : 𝓞 K} (h_x_nezero : x ≠ 0) : + (Function.mulSupport fun w : FinitePlace K ↦ w x).Finite := by + have (w : FinitePlace K) : w x ≠ 1 ↔ w x < 1 := + ne_iff_lt_iff_le.mpr <| norm_embedding_eq w x ▸ norm_le_one w.maximalIdeal x + simp_rw [Function.mulSupport, this, ← norm_embedding_eq, norm_lt_one_iff_mem, + ← Ideal.dvd_span_singleton] + have h : {v : HeightOneSpectrum (𝓞 K) | v.asIdeal ∣ span {x}}.Finite := by + apply Ideal.finite_factors + simp only [Submodule.zero_eq_bot, ne_eq, span_singleton_eq_bot, h_x_nezero, not_false_eq_true] + have h_inj : Set.InjOn FinitePlace.maximalIdeal {w | w.maximalIdeal.asIdeal ∣ span {x}} := + Function.Injective.injOn maximalIdeal_injective + refine (h.subset ?_).of_finite_image h_inj + simp only [dvd_span_singleton, Set.image_subset_iff, Set.preimage_setOf_eq, subset_refl] + +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₀] + obtain ⟨ha, hb⟩ := h_x_nezero + simp_rw [← RingOfIntegers.coe_eq_algebraMap] + apply ((mulSupport_finite_int ha).union (mulSupport_finite_int hb)).subset + intro w + simp only [Function.mem_mulSupport, ne_eq, Set.mem_union] + contrapose! + simp +contextual only [ne_eq, one_ne_zero, not_false_eq_true, div_self, implies_true] + +end FinitePlace + +end NumberField diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index a50570c41e9f1..5e8647a2a1598 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -227,9 +227,9 @@ theorem dvd_iff_norm_le {n : ℕ} {z : ℤ} : ↑(p ^ n) ∣ z ↔ padicNorm p z · rw [zpow_le_zpow_iff_right₀, neg_le_neg_iff, padicValRat.of_int, padicValInt.of_ne_one_ne_zero hp.1.ne_one _] · norm_cast - rw [← multiplicity.Finite.pow_dvd_iff_le_multiplicity] + rw [← FiniteMultiplicity.pow_dvd_iff_le_multiplicity] · norm_cast - · apply Int.multiplicity_finite_iff.2 ⟨by simp [hp.out.ne_one], mod_cast hz⟩ + · apply Int.finiteMultiplicity_iff.2 ⟨by simp [hp.out.ne_one], mod_cast hz⟩ · exact_mod_cast hz · exact_mod_cast hp.out.one_lt diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index d86fe1b155206..7e461266cf61d 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -60,7 +60,7 @@ p-adic, p adic, padic, norm, valuation, cauchy, completion, p-adic completion noncomputable section -open Nat multiplicity padicNorm CauSeq CauSeq.Completion Metric +open Nat padicNorm CauSeq CauSeq.Completion Metric /-- The type of Cauchy sequences of rationals with respect to the `p`-adic norm. -/ abbrev PadicSeq (p : ℕ) := diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean index d2df11023e6f1..07c80e89edce7 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean @@ -66,12 +66,8 @@ open Nat open Rat -open multiplicity - namespace padicValNat -open multiplicity - variable {p : ℕ} /-- If `p ≠ 0` and `p ≠ 1`, then `padicValNat p p` is `1`. -/ @@ -91,7 +87,7 @@ theorem maxPowDiv_eq_emultiplicity {p n : ℕ} (hp : 1 < p) (hn : 0 < n) : apply Nat.not_lt.mpr <| le_of_dvd hp hn h simp -theorem maxPowDiv_eq_multiplicity {p n : ℕ} (hp : 1 < p) (hn : 0 < n) (h : Finite p n) : +theorem maxPowDiv_eq_multiplicity {p n : ℕ} (hp : 1 < p) (hn : 0 < n) (h : FiniteMultiplicity p n) : p.maxPowDiv n = multiplicity p n := by exact_mod_cast h.emultiplicity_eq_multiplicity ▸ maxPowDiv_eq_emultiplicity hp hn @@ -101,7 +97,7 @@ theorem padicValNat_eq_maxPowDiv : @padicValNat = @maxPowDiv := by ext p n by_cases h : 1 < p ∧ 0 < n · rw [padicValNat_def' h.1.ne' h.2, maxPowDiv_eq_multiplicity h.1 h.2] - exact Nat.multiplicity_finite_iff.2 ⟨h.1.ne', h.2⟩ + exact Nat.finiteMultiplicity_iff.2 ⟨h.1.ne', h.2⟩ · simp only [not_and_or,not_gt_eq,Nat.le_zero] at h apply h.elim · intro h @@ -121,8 +117,6 @@ def padicValInt (p : ℕ) (z : ℤ) : ℕ := namespace padicValInt -open multiplicity - variable {p : ℕ} theorem of_ne_one_ne_zero {z : ℤ} (hp : p ≠ 1) (hz : z ≠ 0) : @@ -163,8 +157,6 @@ lemma padicValRat_def (p : ℕ) (q : ℚ) : namespace padicValRat -open multiplicity - variable {p : ℕ} /-- `padicValRat p q` is symmetric in `q`. -/ @@ -232,13 +224,11 @@ end padicValNat namespace padicValRat -open multiplicity - variable {p : ℕ} [hp : Fact p.Prime] /-- The multiplicity of `p : ℕ` in `a : ℤ` is finite exactly when `a ≠ 0`. -/ -theorem finite_int_prime_iff {a : ℤ} : Finite (p : ℤ) a ↔ a ≠ 0 := by - simp [Int.multiplicity_finite_iff, hp.1.ne_one] +theorem finite_int_prime_iff {a : ℤ} : FiniteMultiplicity (p : ℤ) a ↔ a ≠ 0 := by + simp [Int.finiteMultiplicity_iff, hp.1.ne_one] /-- A rewrite lemma for `padicValRat p q` when `q` is expressed in terms of `Rat.mk`. -/ protected theorem defn (p : ℕ) [hp : Fact p.Prime] {q : ℚ} {n d : ℤ} (hqz : q ≠ 0) @@ -296,8 +286,8 @@ theorem padicValRat_le_padicValRat_iff {n₁ n₂ d₁ d₂ : ℤ} (hn₁ : n₁ (hd₁ : d₁ ≠ 0) (hd₂ : d₂ ≠ 0) : padicValRat p (n₁ /. d₁) ≤ padicValRat p (n₂ /. d₂) ↔ ∀ n : ℕ, (p : ℤ) ^ n ∣ n₁ * d₂ → (p : ℤ) ^ n ∣ n₂ * d₁ := by - have hf1 : Finite (p : ℤ) (n₁ * d₂) := finite_int_prime_iff.2 (mul_ne_zero hn₁ hd₂) - have hf2 : Finite (p : ℤ) (n₂ * d₁) := finite_int_prime_iff.2 (mul_ne_zero hn₂ hd₁) + have hf1 : FiniteMultiplicity (p : ℤ) (n₁ * d₂) := finite_int_prime_iff.2 (mul_ne_zero hn₁ hd₂) + have hf2 : FiniteMultiplicity (p : ℤ) (n₂ * d₁) := finite_int_prime_iff.2 (mul_ne_zero hn₂ hd₁) conv => lhs rw [padicValRat.defn p (Rat.divInt_ne_zero_of_ne_zero hn₁ hd₁) rfl, diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean b/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean index cba87c009c504..39a5bdc501f09 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Defs.lean @@ -22,18 +22,16 @@ universe u open Nat -open multiplicity - variable {p : ℕ} /-- For `p ≠ 1`, the `p`-adic valuation of a natural `n ≠ 0` is the largest natural number `k` such that `p^k` divides `n`. If `n = 0` or `p = 1`, then `padicValNat p q` defaults to `0`. -/ def padicValNat (p : ℕ) (n : ℕ) : ℕ := - if h : p ≠ 1 ∧ 0 < n then Nat.find (multiplicity_finite_iff.2 h) else 0 + if h : p ≠ 1 ∧ 0 < n then Nat.find (finiteMultiplicity_iff.2 h) else 0 theorem padicValNat_def' {n : ℕ} (hp : p ≠ 1) (hn : 0 < n) : padicValNat p n = multiplicity p n := by - simp [padicValNat, hp, hn, multiplicity, emultiplicity, multiplicity_finite_iff.2 ⟨hp, hn⟩] + simp [padicValNat, hp, hn, multiplicity, emultiplicity, finiteMultiplicity_iff.2 ⟨hp, hn⟩] convert (WithTop.untop'_coe ..).symm /-- A simplification of `padicValNat` when one input is prime, by analogy with @@ -46,13 +44,11 @@ theorem padicValNat_def [hp : Fact p.Prime] {n : ℕ} (hn : 0 < n) : `padicValRat_def`. -/ theorem padicValNat_eq_emultiplicity [hp : Fact p.Prime] {n : ℕ} (hn : 0 < n) : padicValNat p n = emultiplicity p n := by - rw [(multiplicity_finite_iff.2 ⟨hp.out.ne_one, hn⟩).emultiplicity_eq_multiplicity] + rw [(finiteMultiplicity_iff.2 ⟨hp.out.ne_one, hn⟩).emultiplicity_eq_multiplicity] exact_mod_cast padicValNat_def hn namespace padicValNat -open multiplicity - open List /-- `padicValNat p 0` is `0` for any `p`. -/ @@ -86,7 +82,7 @@ theorem le_padicValNat_iff_replicate_subperm_primeFactorsList {a b : ℕ} {n : (hb : b ≠ 0) : n ≤ padicValNat a b ↔ replicate n a <+~ b.primeFactorsList := by rw [← le_emultiplicity_iff_replicate_subperm_primeFactorsList ha hb, - Nat.multiplicity_finite_iff.2 ⟨ha.ne_one, Nat.pos_of_ne_zero hb⟩ + Nat.finiteMultiplicity_iff.2 ⟨ha.ne_one, Nat.pos_of_ne_zero hb⟩ |>.emultiplicity_eq_multiplicity, ← padicValNat_def' ha.ne_one (Nat.pos_of_ne_zero hb), Nat.cast_le] diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index caa54da5b5024..8e5a610e2930f 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -529,11 +529,20 @@ theorem PartialOrder.ext {A B : PartialOrder α} ext x y exact H x y +theorem PartialOrder.ext_lt {A B : PartialOrder α} + (H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B := by + ext x y + rw [le_iff_lt_or_eq, @le_iff_lt_or_eq _ A, H] + theorem LinearOrder.ext {A B : LinearOrder α} (H : ∀ x y : α, (haveI := A; x ≤ y) ↔ x ≤ y) : A = B := by ext x y exact H x y +theorem LinearOrder.ext_lt {A B : LinearOrder α} + (H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B := + LinearOrder.toPartialOrder_injective (PartialOrder.ext_lt H) + /-- Given a relation `R` on `β` and a function `f : α → β`, the preimage relation on `α` is defined by `x ≤ y ↔ f x ≤ f y`. It is the unique relation on `α` making `f` a `RelEmbedding` (assuming `f` is injective). -/ @@ -689,6 +698,9 @@ instance (α : Type*) [LE α] : LE αᵒᵈ := instance (α : Type*) [LT α] : LT αᵒᵈ := ⟨fun x y : α ↦ y < x⟩ +instance instOrd (α : Type*) [Ord α] : Ord αᵒᵈ where + compare := fun (a b : α) ↦ compare b a + instance instSup (α : Type*) [Min α] : Max αᵒᵈ := ⟨((· ⊓ ·) : α → α → α)⟩ @@ -706,6 +718,7 @@ instance instPartialOrder (α : Type*) [PartialOrder α] : PartialOrder αᵒᵈ instance instLinearOrder (α : Type*) [LinearOrder α] : LinearOrder αᵒᵈ where __ := inferInstanceAs (PartialOrder αᵒᵈ) + __ := inferInstanceAs (Ord αᵒᵈ) le_total := fun a b : α ↦ le_total b a max := fun a b ↦ (min a b : α) min := fun a b ↦ (max a b : α) @@ -713,6 +726,10 @@ instance instLinearOrder (α : Type*) [LinearOrder α] : LinearOrder αᵒᵈ wh max_def := fun a b ↦ show (min .. : α) = _ by rw [min_comm, min_def]; rfl decidableLE := (inferInstance : DecidableRel (fun a b : α ↦ b ≤ a)) decidableLT := (inferInstance : DecidableRel (fun a b : α ↦ b < a)) + decidableEq := (inferInstance : DecidableEq α) + compare_eq_compareOfLessAndEq a b := by + simp only [compare, LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq, eq_comm] + rfl /-- The opposite linear order to a given linear order -/ def _root_.LinearOrder.swap (α : Type*) (_ : LinearOrder α) : LinearOrder α := @@ -720,16 +737,19 @@ def _root_.LinearOrder.swap (α : Type*) (_ : LinearOrder α) : LinearOrder α : instance : ∀ [Inhabited α], Inhabited αᵒᵈ := fun [x : Inhabited α] => x +theorem Ord.dual_dual (α : Type*) [H : Ord α] : OrderDual.instOrd αᵒᵈ = H := + rfl + theorem Preorder.dual_dual (α : Type*) [H : Preorder α] : OrderDual.instPreorder αᵒᵈ = H := - Preorder.ext fun _ _ ↦ Iff.rfl + rfl theorem instPartialOrder.dual_dual (α : Type*) [H : PartialOrder α] : OrderDual.instPartialOrder αᵒᵈ = H := - PartialOrder.ext fun _ _ ↦ Iff.rfl + rfl theorem instLinearOrder.dual_dual (α : Type*) [H : LinearOrder α] : OrderDual.instLinearOrder αᵒᵈ = H := - LinearOrder.ext fun _ _ ↦ Iff.rfl + rfl end OrderDual @@ -762,6 +782,10 @@ theorem compl_le [LinearOrder α] : (· ≤ · : α → α → _)ᶜ = (· > ·) theorem compl_gt [LinearOrder α] : (· > · : α → α → _)ᶜ = (· ≤ ·) := by ext; simp [compl] theorem compl_ge [LinearOrder α] : (· ≥ · : α → α → _)ᶜ = (· < ·) := by ext; simp [compl] +instance Ne.instIsEquiv_compl : IsEquiv α (· ≠ ·)ᶜ := by + convert eq_isEquiv α + simp [compl] + /-! ### Order instances on the function space -/ diff --git a/Mathlib/Order/BoundedOrder/Lattice.lean b/Mathlib/Order/BoundedOrder/Lattice.lean index 42c1cf7c0ac57..3c06e09a51844 100644 --- a/Mathlib/Order/BoundedOrder/Lattice.lean +++ b/Mathlib/Order/BoundedOrder/Lattice.lean @@ -7,19 +7,19 @@ import Mathlib.Order.BoundedOrder.Basic import Mathlib.Order.Lattice /-! -# bounded lattices +# Bounded lattices This file defines top and bottom elements (greatest and least elements) of a type, the bounded -variants of different kinds of lattices, sets up the typeclass hierarchy between them and provides +variants of different kinds of lattices, sets up the typeclass hierarchy between them, and provides instances for `Prop` and `fun`. ## Common lattices -* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]` +* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]`. It captures the properties of `Disjoint` that are common to `GeneralizedBooleanAlgebra` and `DistribLattice` when `OrderBot`. * Bounded and distributive lattice. Notated by `[DistribLattice α] [BoundedOrder α]`. - Typical examples include `Prop` and `Det α`. + Typical examples include `Prop` and `Set α`. -/ open Function OrderDual diff --git a/Mathlib/Order/Bounds/Image.lean b/Mathlib/Order/Bounds/Image.lean index 910c466332520..cb97202b9d3ce 100644 --- a/Mathlib/Order/Bounds/Image.lean +++ b/Mathlib/Order/Bounds/Image.lean @@ -196,6 +196,45 @@ theorem map_isLeast : IsLeast s a → IsGreatest (f '' s) (f a) := end Antitone +section StrictMono + +variable [LinearOrder α] [Preorder β] {f : α → β} {a : α} {s : Set α} + +lemma StrictMono.mem_upperBounds_image (hf : StrictMono f) : + f a ∈ upperBounds (f '' s) ↔ a ∈ upperBounds s := by simp [upperBounds, hf.le_iff_le] + +lemma StrictMono.mem_lowerBounds_image (hf : StrictMono f) : + f a ∈ lowerBounds (f '' s) ↔ a ∈ lowerBounds s := by simp [lowerBounds, hf.le_iff_le] + +lemma StrictMono.map_isLeast (hf : StrictMono f) : IsLeast (f '' s) (f a) ↔ IsLeast s a := by + simp [IsLeast, hf.injective.eq_iff, hf.mem_lowerBounds_image] + +lemma StrictMono.map_isGreatest (hf : StrictMono f) : + IsGreatest (f '' s) (f a) ↔ IsGreatest s a := by + simp [IsGreatest, hf.injective.eq_iff, hf.mem_upperBounds_image] + +end StrictMono + +section StrictAnti + +variable [LinearOrder α] [Preorder β] {f : α → β} {a : α} {s : Set α} + +lemma StrictAnti.mem_upperBounds_image (hf : StrictAnti f) : + f a ∈ upperBounds (f '' s) ↔ a ∈ lowerBounds s := by + simp [upperBounds, lowerBounds, hf.le_iff_le] + +lemma StrictAnti.mem_lowerBounds_image (hf : StrictAnti f) : + f a ∈ lowerBounds (f '' s) ↔ a ∈ upperBounds s := by + simp [upperBounds, lowerBounds, hf.le_iff_le] + +lemma StrictAnti.map_isLeast (hf : StrictAnti f) : IsLeast (f '' s) (f a) ↔ IsGreatest s a := by + simp [IsLeast, IsGreatest, hf.injective.eq_iff, hf.mem_lowerBounds_image] + +lemma StrictAnti.map_isGreatest (hf : StrictAnti f) : IsGreatest (f '' s) (f a) ↔ IsLeast s a := by + simp [IsLeast, IsGreatest, hf.injective.eq_iff, hf.mem_upperBounds_image] + +end StrictAnti + section Image2 variable [Preorder α] [Preorder β] [Preorder γ] {f : α → β → γ} {s : Set α} {t : Set β} {a : α} diff --git a/Mathlib/Order/Cofinal.lean b/Mathlib/Order/Cofinal.lean index 572a4ffc3cedb..532481a8f4cbe 100644 --- a/Mathlib/Order/Cofinal.lean +++ b/Mathlib/Order/Cofinal.lean @@ -3,9 +3,7 @@ Copyright (c) 2024 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ -import Mathlib.Data.Set.Basic -import Mathlib.Tactic.Contrapose -import Mathlib.Order.Bounds.Defs +import Mathlib.Order.GaloisConnection /-! # Cofinal sets @@ -21,7 +19,7 @@ For the cofinality of a set as a cardinal, see `Mathlib.SetTheory.Cardinal.Cofin - Deprecate `Order.Cofinal` in favor of this predicate. -/ -variable {α : Type*} +variable {α β : Type*} section LE variable [LE α] @@ -60,6 +58,16 @@ theorem IsCofinal.trans {s : Set α} {t : Set s} (hs : IsCofinal s) (ht : IsCofi obtain ⟨c, hc, hc'⟩ := ht ⟨b, hb⟩ exact ⟨c, Set.mem_image_of_mem _ hc, hb'.trans hc'⟩ +theorem GaloisConnection.map_cofinal [Preorder β] {f : β → α} {g : α → β} + (h : GaloisConnection f g) {s : Set α} (hs : IsCofinal s) : IsCofinal (g '' s) := by + intro a + obtain ⟨b, hb, hb'⟩ := hs (f a) + exact ⟨g b, Set.mem_image_of_mem _ hb, h.le_iff_le.1 hb'⟩ + +theorem OrderIso.map_cofinal [Preorder β] (e : α ≃o β) {s : Set α} (hs : IsCofinal s) : + IsCofinal (e '' s) := + e.symm.to_galoisConnection.map_cofinal hs + end Preorder section PartialOrder @@ -105,4 +113,14 @@ theorem not_isCofinal_iff_bddAbove [NoMaxOrder α] {s : Set α} : ¬ IsCofinal s theorem not_bddAbove_iff_isCofinal [NoMaxOrder α] {s : Set α} : ¬ BddAbove s ↔ IsCofinal s := not_iff_comm.1 not_isCofinal_iff_bddAbove +/-- The set of "records" (the smallest inputs yielding the highest values) with respect to a +well-ordering of `α` is a cofinal set. -/ +theorem isCofinal_setOf_imp_lt (r : α → α → Prop) [h : IsWellFounded α r] : + IsCofinal { a | ∀ b, r b a → b < a } := by + intro a + obtain ⟨b, hb, hb'⟩ := h.wf.has_min (Set.Ici a) Set.nonempty_Ici + refine ⟨b, fun c hc ↦ ?_, hb⟩ + by_contra! hc' + exact hb' c (hb.trans hc') hc + end LinearOrder diff --git a/Mathlib/Order/Extension/Well.lean b/Mathlib/Order/Extension/Well.lean index 1b5175e593d0c..3d74ec49658eb 100644 --- a/Mathlib/Order/Extension/Well.lean +++ b/Mathlib/Order/Extension/Well.lean @@ -70,10 +70,9 @@ end IsWellFounded namespace WellFounded -set_option linter.deprecated false - variable (hwf : WellFounded r) +set_option linter.deprecated false in /-- An arbitrary well order on `α` that extends `r`. The construction maps `r` into two well-orders: the first map is `WellFounded.rank`, which is not @@ -89,12 +88,14 @@ noncomputable def wellOrderExtension : LinearOrder α := @LinearOrder.lift' α (Ordinal ×ₗ Cardinal) _ (fun a : α => (hwf.rank a, embeddingToCardinal a)) fun _ _ h => embeddingToCardinal.injective <| congr_arg Prod.snd h +set_option linter.deprecated false in @[deprecated IsWellFounded.wellOrderExtension.isWellFounded_lt (since := "2024-09-07")] instance wellOrderExtension.isWellFounded_lt : IsWellFounded α hwf.wellOrderExtension.lt := ⟨InvImage.wf (fun a : α => (hwf.rank a, embeddingToCardinal a)) <| Ordinal.lt_wf.prod_lex Cardinal.lt_wf⟩ include hwf in +set_option linter.deprecated false in /-- Any well-founded relation can be extended to a well-ordering on that type. -/ @[deprecated IsWellFounded.exists_well_order_ge (since := "2024-09-07")] theorem exists_well_order_ge : ∃ s, r ≤ s ∧ IsWellOrder α s := diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index ed602722331c0..d8c1fe0ab97b7 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -172,20 +172,6 @@ theorem Tendsto.eventually_forall_le_atBot [Preorder β] {l : Filter α} ∀ᶠ x in l, ∀ y, y ≤ f x → p y := by rw [← Filter.eventually_forall_le_atBot] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap -instance (priority := 200) atTop.isCountablyGenerated [Preorder α] [Countable α] : - (atTop : Filter <| α).IsCountablyGenerated := - isCountablyGenerated_seq _ - -instance (priority := 200) atBot.isCountablyGenerated [Preorder α] [Countable α] : - (atBot : Filter <| α).IsCountablyGenerated := - isCountablyGenerated_seq _ - -instance _root_.OrderDual.instIsCountablyGeneratedAtTop [Preorder α] - [IsCountablyGenerated (atBot : Filter α)] : IsCountablyGenerated (atTop : Filter αᵒᵈ) := ‹_› - -instance _root_.OrderDual.instIsCountablyGeneratedAtBot [Preorder α] - [IsCountablyGenerated (atTop : Filter α)] : IsCountablyGenerated (atBot : Filter αᵒᵈ) := ‹_› - theorem _root_.IsTop.atTop_eq [Preorder α] {a : α} (ha : IsTop a) : atTop = 𝓟 (Ici a) := (iInf_le _ _).antisymm <| le_iInf fun b ↦ principal_mono.2 <| Ici_subset_Ici.2 <| ha b @@ -291,10 +277,6 @@ theorem map_atTop_eq {f : α → β} : atTop.map f = ⨅ a, 𝓟 (f '' { a' | a theorem frequently_atTop' [NoMaxOrder α] : (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b > a, p b := atTop_basis_Ioi.frequently_iff.trans <| by simp -lemma atTop_countable_basis [Countable α] : - HasCountableBasis (atTop : Filter α) (fun _ => True) Ici := - { atTop_basis with countable := to_countable _ } - end IsDirected section IsCodirected @@ -333,10 +315,6 @@ theorem map_atBot_eq {f : α → β} : atBot.map f = ⨅ a, 𝓟 (f '' { a' | a' theorem frequently_atBot' [NoMinOrder α] : (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b < a, p b := frequently_atTop' (α := αᵒᵈ) -lemma atBot_countable_basis [Countable α] : - HasCountableBasis (atBot : Filter α) (fun _ => True) Iic := - { atBot_basis with countable := to_countable _ } - end IsCodirected theorem tendsto_atTop [Preorder β] {m : α → β} {f : Filter α} : @@ -918,12 +896,6 @@ theorem prod_atTop_atTop_eq [Preorder α] [Preorder β] : · subsingleton simpa [atTop, prod_iInf_left, prod_iInf_right, iInf_prod] using iInf_comm -instance instIsCountablyGeneratedAtTopProd [Preorder α] [IsCountablyGenerated (atTop : Filter α)] - [Preorder β] [IsCountablyGenerated (atTop : Filter β)] : - IsCountablyGenerated (atTop : Filter (α × β)) := by - rw [← prod_atTop_atTop_eq] - infer_instance - lemma tendsto_finset_prod_atTop : Tendsto (fun (p : Finset ι × Finset ι') ↦ p.1 ×ˢ p.2) atTop atTop := by classical @@ -938,12 +910,6 @@ theorem prod_atBot_atBot_eq [Preorder α] [Preorder β] : (atBot : Filter α) ×ˢ (atBot : Filter β) = (atBot : Filter (α × β)) := @prod_atTop_atTop_eq αᵒᵈ βᵒᵈ _ _ -instance instIsCountablyGeneratedAtBotProd [Preorder α] [IsCountablyGenerated (atBot : Filter α)] - [Preorder β] [IsCountablyGenerated (atBot : Filter β)] : - IsCountablyGenerated (atBot : Filter (α × β)) := by - rw [← prod_atBot_atBot_eq] - infer_instance - theorem prod_map_atTop_eq {α₁ α₂ β₁ β₂ : Type*} [Preorder β₁] [Preorder β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : map u₁ atTop ×ˢ map u₂ atTop = map (Prod.map u₁ u₂) atTop := by rw [prod_map_map_eq, prod_atTop_atTop_eq, Prod.map_def] @@ -1296,106 +1262,12 @@ theorem HasAntitoneBasis.subbasis_with_rel {f : Filter α} {s : ℕ → Set α} simp only [forall_mem_image, forall_and, mem_Iio] at hφ exact ⟨φ, forall_swap.2 hφ.1, forall_swap.2 hφ.2⟩ -/-- If `f` is a nontrivial countably generated filter, then there exists a sequence that converges -to `f`. -/ -theorem exists_seq_tendsto (f : Filter α) [IsCountablyGenerated f] [NeBot f] : - ∃ x : ℕ → α, Tendsto x atTop f := by - obtain ⟨B, h⟩ := f.exists_antitone_basis - choose x hx using fun n => Filter.nonempty_of_mem (h.mem n) - exact ⟨x, h.tendsto hx⟩ - -theorem exists_seq_monotone_tendsto_atTop_atTop (α : Type*) [Preorder α] [Nonempty α] - [IsDirected α (· ≤ ·)] [(atTop : Filter α).IsCountablyGenerated] : - ∃ xs : ℕ → α, Monotone xs ∧ Tendsto xs atTop atTop := by - obtain ⟨ys, h⟩ := exists_seq_tendsto (atTop : Filter α) - choose c hleft hright using exists_ge_ge (α := α) - set xs : ℕ → α := fun n => (List.range n).foldl (fun x n ↦ c x (ys n)) (ys 0) - have hsucc (n : ℕ) : xs (n + 1) = c (xs n) (ys n) := by simp [xs, List.range_succ] - refine ⟨xs, ?_, ?_⟩ - · refine monotone_nat_of_le_succ fun n ↦ ?_ - rw [hsucc] - apply hleft - · refine (tendsto_add_atTop_iff_nat 1).1 <| tendsto_atTop_mono (fun n ↦ ?_) h - rw [hsucc] - apply hright - -theorem exists_seq_antitone_tendsto_atTop_atBot (α : Type*) [Preorder α] [Nonempty α] - [IsDirected α (· ≥ ·)] [(atBot : Filter α).IsCountablyGenerated] : - ∃ xs : ℕ → α, Antitone xs ∧ Tendsto xs atTop atBot := - exists_seq_monotone_tendsto_atTop_atTop αᵒᵈ - -/-- An abstract version of continuity of sequentially continuous functions on metric spaces: -if a filter `k` is countably generated then `Tendsto f k l` iff for every sequence `u` -converging to `k`, `f ∘ u` tends to `l`. -/ -theorem tendsto_iff_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : - Tendsto f k l ↔ ∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l := by - refine ⟨fun h x hx => h.comp hx, fun H s hs => ?_⟩ - contrapose! H - have : NeBot (k ⊓ 𝓟 (f ⁻¹' sᶜ)) := by simpa [neBot_iff, inf_principal_eq_bot] - rcases (k ⊓ 𝓟 (f ⁻¹' sᶜ)).exists_seq_tendsto with ⟨x, hx⟩ - rw [tendsto_inf, tendsto_principal] at hx - refine ⟨x, hx.1, fun h => ?_⟩ - rcases (hx.2.and (h hs)).exists with ⟨N, hnmem, hmem⟩ - exact hnmem hmem - -theorem tendsto_of_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : - (∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l) → Tendsto f k l := - tendsto_iff_seq_tendsto.2 - -theorem eventually_iff_seq_eventually {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] : - (∀ᶠ n in l, p n) ↔ ∀ x : ℕ → ι, Tendsto x atTop l → ∀ᶠ n : ℕ in atTop, p (x n) := by - simpa using tendsto_iff_seq_tendsto (f := id) (l := 𝓟 {x | p x}) - -theorem frequently_iff_seq_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] : - (∃ᶠ n in l, p n) ↔ ∃ x : ℕ → ι, Tendsto x atTop l ∧ ∃ᶠ n : ℕ in atTop, p (x n) := by - simp only [Filter.Frequently, eventually_iff_seq_eventually (l := l)] - push_neg; rfl - theorem subseq_forall_of_frequently {ι : Type*} {x : ℕ → ι} {p : ι → Prop} {l : Filter ι} (h_tendsto : Tendsto x atTop l) (h : ∃ᶠ n in atTop, p (x n)) : ∃ ns : ℕ → ℕ, Tendsto (fun n => x (ns n)) atTop l ∧ ∀ n, p (x (ns n)) := by choose ns hge hns using frequently_atTop.1 h exact ⟨ns, h_tendsto.comp (tendsto_atTop_mono hge tendsto_id), hns⟩ -theorem exists_seq_forall_of_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] (h : ∃ᶠ n in l, p n) : - ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := by - rw [frequently_iff_seq_frequently] at h - obtain ⟨x, hx_tendsto, hx_freq⟩ := h - obtain ⟨n_to_n, h_tendsto, h_freq⟩ := subseq_forall_of_frequently hx_tendsto hx_freq - exact ⟨x ∘ n_to_n, h_tendsto, h_freq⟩ - -lemma frequently_iff_seq_forall {ι : Type*} {l : Filter ι} {p : ι → Prop} - [l.IsCountablyGenerated] : - (∃ᶠ n in l, p n) ↔ ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := - ⟨exists_seq_forall_of_frequently, fun ⟨_ns, hnsl, hpns⟩ ↦ - hnsl.frequently <| Frequently.of_forall hpns⟩ - -/-- A sequence converges if every subsequence has a convergent subsequence. -/ -theorem tendsto_of_subseq_tendsto {ι : Type*} {x : ι → α} {f : Filter α} {l : Filter ι} - [l.IsCountablyGenerated] - (hxy : ∀ ns : ℕ → ι, Tendsto ns atTop l → - ∃ ms : ℕ → ℕ, Tendsto (fun n => x (ns <| ms n)) atTop f) : - Tendsto x l f := by - contrapose! hxy - obtain ⟨s, hs, hfreq⟩ : ∃ s ∈ f, ∃ᶠ n in l, x n ∉ s := by - rwa [not_tendsto_iff_exists_frequently_nmem] at hxy - obtain ⟨y, hy_tendsto, hy_freq⟩ := exists_seq_forall_of_frequently hfreq - refine ⟨y, hy_tendsto, fun ms hms_tendsto ↦ ?_⟩ - rcases (hms_tendsto.eventually_mem hs).exists with ⟨n, hn⟩ - exact absurd hn <| hy_freq _ - -theorem subseq_tendsto_of_neBot {f : Filter α} [IsCountablyGenerated f] {u : ℕ → α} - (hx : NeBot (f ⊓ map u atTop)) : ∃ θ : ℕ → ℕ, StrictMono θ ∧ Tendsto (u ∘ θ) atTop f := by - rw [← Filter.push_pull', map_neBot_iff] at hx - rcases exists_seq_tendsto (comap u f ⊓ atTop) with ⟨φ, hφ⟩ - rw [tendsto_inf, tendsto_comap_iff] at hφ - obtain ⟨ψ, hψ, hψφ⟩ : ∃ ψ : ℕ → ℕ, StrictMono ψ ∧ StrictMono (φ ∘ ψ) := - strictMono_subseq_of_tendsto_atTop hφ.2 - exact ⟨φ ∘ ψ, hψφ, hφ.1.comp hψ.tendsto_atTop⟩ - end Filter open Filter Finset diff --git a/Mathlib/Order/Filter/AtTopBot/Archimedean.lean b/Mathlib/Order/Filter/AtTopBot/Archimedean.lean index ecc8b33a34a1d..167f93687a47d 100644 --- a/Mathlib/Order/Filter/AtTopBot/Archimedean.lean +++ b/Mathlib/Order/Filter/AtTopBot/Archimedean.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Order.Filter.AtTopBot.Group +import Mathlib.Order.Filter.CountablyGenerated import Mathlib.Tactic.GCongr /-! diff --git a/Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean b/Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean new file mode 100644 index 0000000000000..96978e899bf72 --- /dev/null +++ b/Mathlib/Order/Filter/AtTopBot/CountablyGenerated.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov, Patrick Massot +-/ +import Mathlib.Order.Filter.AtTopBot +import Mathlib.Order.Filter.CountablyGenerated + +/-! +# Convergence to infinity and countably generated filters + +In this file we prove that + +- `Filter.atTop` and `Filter.atBot` filters on a countable type are countably generated; +- `Filter.exists_seq_tendsto`: if `f` is a nontrivial countably generated filter, + then there exists a sequence that converges. to `f`; +- `Filter.tendsto_iff_seq_tendsto`: convergence along a countably generated filter + is equivalent to convergence along all sequences that converge to this filter. +-/ + +open Set + +namespace Filter + +variable {α β : Type*} + +instance (priority := 200) atTop.isCountablyGenerated [Preorder α] [Countable α] : + (atTop : Filter <| α).IsCountablyGenerated := + isCountablyGenerated_seq _ + +instance (priority := 200) atBot.isCountablyGenerated [Preorder α] [Countable α] : + (atBot : Filter <| α).IsCountablyGenerated := + isCountablyGenerated_seq _ + +instance instIsCountablyGeneratedAtTopProd [Preorder α] [IsCountablyGenerated (atTop : Filter α)] + [Preorder β] [IsCountablyGenerated (atTop : Filter β)] : + IsCountablyGenerated (atTop : Filter (α × β)) := by + rw [← prod_atTop_atTop_eq] + infer_instance + +instance instIsCountablyGeneratedAtBotProd [Preorder α] [IsCountablyGenerated (atBot : Filter α)] + [Preorder β] [IsCountablyGenerated (atBot : Filter β)] : + IsCountablyGenerated (atBot : Filter (α × β)) := by + rw [← prod_atBot_atBot_eq] + infer_instance + +instance _root_.OrderDual.instIsCountablyGeneratedAtTop [Preorder α] + [IsCountablyGenerated (atBot : Filter α)] : IsCountablyGenerated (atTop : Filter αᵒᵈ) := ‹_› + +instance _root_.OrderDual.instIsCountablyGeneratedAtBot [Preorder α] + [IsCountablyGenerated (atTop : Filter α)] : IsCountablyGenerated (atBot : Filter αᵒᵈ) := ‹_› + +lemma atTop_countable_basis [Preorder α] [IsDirected α (· ≤ ·)] [Nonempty α] [Countable α] : + HasCountableBasis (atTop : Filter α) (fun _ => True) Ici := + { atTop_basis with countable := to_countable _ } + +lemma atBot_countable_basis [Preorder α] [IsDirected α (· ≥ ·)] [Nonempty α] [Countable α] : + HasCountableBasis (atBot : Filter α) (fun _ => True) Iic := + { atBot_basis with countable := to_countable _ } + +/-- If `f` is a nontrivial countably generated filter, then there exists a sequence that converges +to `f`. -/ +theorem exists_seq_tendsto (f : Filter α) [IsCountablyGenerated f] [NeBot f] : + ∃ x : ℕ → α, Tendsto x atTop f := by + obtain ⟨B, h⟩ := f.exists_antitone_basis + choose x hx using fun n => Filter.nonempty_of_mem (h.mem n) + exact ⟨x, h.tendsto hx⟩ + +theorem exists_seq_monotone_tendsto_atTop_atTop (α : Type*) [Preorder α] [Nonempty α] + [IsDirected α (· ≤ ·)] [(atTop : Filter α).IsCountablyGenerated] : + ∃ xs : ℕ → α, Monotone xs ∧ Tendsto xs atTop atTop := by + obtain ⟨ys, h⟩ := exists_seq_tendsto (atTop : Filter α) + choose c hleft hright using exists_ge_ge (α := α) + set xs : ℕ → α := fun n => (List.range n).foldl (fun x n ↦ c x (ys n)) (ys 0) + have hsucc (n : ℕ) : xs (n + 1) = c (xs n) (ys n) := by simp [xs, List.range_succ] + refine ⟨xs, ?_, ?_⟩ + · refine monotone_nat_of_le_succ fun n ↦ ?_ + rw [hsucc] + apply hleft + · refine (tendsto_add_atTop_iff_nat 1).1 <| tendsto_atTop_mono (fun n ↦ ?_) h + rw [hsucc] + apply hright + +theorem exists_seq_antitone_tendsto_atTop_atBot (α : Type*) [Preorder α] [Nonempty α] + [IsDirected α (· ≥ ·)] [(atBot : Filter α).IsCountablyGenerated] : + ∃ xs : ℕ → α, Antitone xs ∧ Tendsto xs atTop atBot := + exists_seq_monotone_tendsto_atTop_atTop αᵒᵈ + +/-- An abstract version of continuity of sequentially continuous functions on metric spaces: +if a filter `k` is countably generated then `Tendsto f k l` iff for every sequence `u` +converging to `k`, `f ∘ u` tends to `l`. -/ +theorem tendsto_iff_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : + Tendsto f k l ↔ ∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l := by + refine ⟨fun h x hx => h.comp hx, fun H s hs => ?_⟩ + contrapose! H + have : NeBot (k ⊓ 𝓟 (f ⁻¹' sᶜ)) := by simpa [neBot_iff, inf_principal_eq_bot] + rcases (k ⊓ 𝓟 (f ⁻¹' sᶜ)).exists_seq_tendsto with ⟨x, hx⟩ + rw [tendsto_inf, tendsto_principal] at hx + refine ⟨x, hx.1, fun h => ?_⟩ + rcases (hx.2.and (h hs)).exists with ⟨N, hnmem, hmem⟩ + exact hnmem hmem + +theorem tendsto_of_seq_tendsto {f : α → β} {k : Filter α} {l : Filter β} [k.IsCountablyGenerated] : + (∀ x : ℕ → α, Tendsto x atTop k → Tendsto (f ∘ x) atTop l) → Tendsto f k l := + tendsto_iff_seq_tendsto.2 + +theorem eventually_iff_seq_eventually {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] : + (∀ᶠ n in l, p n) ↔ ∀ x : ℕ → ι, Tendsto x atTop l → ∀ᶠ n : ℕ in atTop, p (x n) := by + simpa using tendsto_iff_seq_tendsto (f := id) (l := 𝓟 {x | p x}) + +theorem frequently_iff_seq_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] : + (∃ᶠ n in l, p n) ↔ ∃ x : ℕ → ι, Tendsto x atTop l ∧ ∃ᶠ n : ℕ in atTop, p (x n) := by + simp only [Filter.Frequently, eventually_iff_seq_eventually (l := l)] + push_neg; rfl + +theorem exists_seq_forall_of_frequently {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] (h : ∃ᶠ n in l, p n) : + ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := by + rw [frequently_iff_seq_frequently] at h + obtain ⟨x, hx_tendsto, hx_freq⟩ := h + obtain ⟨n_to_n, h_tendsto, h_freq⟩ := subseq_forall_of_frequently hx_tendsto hx_freq + exact ⟨x ∘ n_to_n, h_tendsto, h_freq⟩ + +lemma frequently_iff_seq_forall {ι : Type*} {l : Filter ι} {p : ι → Prop} + [l.IsCountablyGenerated] : + (∃ᶠ n in l, p n) ↔ ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := + ⟨exists_seq_forall_of_frequently, fun ⟨_ns, hnsl, hpns⟩ ↦ + hnsl.frequently <| Frequently.of_forall hpns⟩ + +/-- A sequence converges if every subsequence has a convergent subsequence. -/ +theorem tendsto_of_subseq_tendsto {ι : Type*} {x : ι → α} {f : Filter α} {l : Filter ι} + [l.IsCountablyGenerated] + (hxy : ∀ ns : ℕ → ι, Tendsto ns atTop l → + ∃ ms : ℕ → ℕ, Tendsto (fun n => x (ns <| ms n)) atTop f) : + Tendsto x l f := by + contrapose! hxy + obtain ⟨s, hs, hfreq⟩ : ∃ s ∈ f, ∃ᶠ n in l, x n ∉ s := by + rwa [not_tendsto_iff_exists_frequently_nmem] at hxy + obtain ⟨y, hy_tendsto, hy_freq⟩ := exists_seq_forall_of_frequently hfreq + refine ⟨y, hy_tendsto, fun ms hms_tendsto ↦ ?_⟩ + rcases (hms_tendsto.eventually_mem hs).exists with ⟨n, hn⟩ + exact absurd hn <| hy_freq _ + +theorem subseq_tendsto_of_neBot {f : Filter α} [IsCountablyGenerated f] {u : ℕ → α} + (hx : NeBot (f ⊓ map u atTop)) : ∃ θ : ℕ → ℕ, StrictMono θ ∧ Tendsto (u ∘ θ) atTop f := by + rw [← Filter.push_pull', map_neBot_iff] at hx + rcases exists_seq_tendsto (comap u f ⊓ atTop) with ⟨φ, hφ⟩ + rw [tendsto_inf, tendsto_comap_iff] at hφ + obtain ⟨ψ, hψ, hψφ⟩ : ∃ ψ : ℕ → ℕ, StrictMono ψ ∧ StrictMono (φ ∘ ψ) := + strictMono_subseq_of_tendsto_atTop hφ.2 + exact ⟨φ ∘ ψ, hψφ, hφ.1.comp hψ.tendsto_atTop⟩ + +end Filter diff --git a/Mathlib/Order/Filter/AtTopBot/Field.lean b/Mathlib/Order/Filter/AtTopBot/Field.lean index 519009d383724..1112b9f9d9e92 100644 --- a/Mathlib/Order/Filter/AtTopBot/Field.lean +++ b/Mathlib/Order/Filter/AtTopBot/Field.lean @@ -94,7 +94,7 @@ theorem tendsto_const_mul_pow_atTop_iff : exact pos_of_mul_pos_left hck (pow_nonneg hk _) lemma tendsto_zpow_atTop_atTop {n : ℤ} (hn : 0 < n) : Tendsto (fun x : α ↦ x ^ n) atTop atTop := by - lift n to ℕ+ using hn; simp + lift n to ℕ using hn.le; simp [(Int.ofNat_pos.mp hn).ne'] end LinearOrderedSemifield diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index ab8dd13872eb6..42dfb3d129084 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Data.Prod.PProd -import Mathlib.Data.Set.Countable import Mathlib.Order.Filter.Finite /-! @@ -77,8 +76,6 @@ machinery, e.g., `simp only [true_and_iff]` or `simp only [forall_const]` can he open Set Filter -section sort - variable {α β γ : Type*} {ι ι' : Sort*} /-- A filter basis `B` on a type `α` is a nonempty collection of sets of `α` @@ -850,215 +847,3 @@ theorem map_sigma_mk_comap {π : α → Type*} {π' : β → Type*} {f : α → apply image_sigmaMk_preimage_sigmaMap hf end Filter - -end sort - -namespace Filter - -variable {α β γ ι : Type*} {ι' : Sort*} - -/-- `IsCountablyGenerated f` means `f = generate s` for some countable `s`. -/ -class IsCountablyGenerated (f : Filter α) : Prop where - /-- There exists a countable set that generates the filter. -/ - out : ∃ s : Set (Set α), s.Countable ∧ f = generate s - -/-- `IsCountableBasis p s` means the image of `s` bounded by `p` is a countable filter basis. -/ -structure IsCountableBasis (p : ι → Prop) (s : ι → Set α) extends IsBasis p s : Prop where - /-- The set of `i` that satisfy the predicate `p` is countable. -/ - countable : (setOf p).Countable - -/-- We say that a filter `l` has a countable basis `s : ι → Set α` bounded by `p : ι → Prop`, -if `t ∈ l` if and only if `t` includes `s i` for some `i` such that `p i`, and the set -defined by `p` is countable. -/ -structure HasCountableBasis (l : Filter α) (p : ι → Prop) (s : ι → Set α) - extends HasBasis l p s : Prop where - /-- The set of `i` that satisfy the predicate `p` is countable. -/ - countable : (setOf p).Countable - -/-- A countable filter basis `B` on a type `α` is a nonempty countable collection of sets of `α` -such that the intersection of two elements of this collection contains some element -of the collection. -/ -structure CountableFilterBasis (α : Type*) extends FilterBasis α where - /-- The set of sets of the filter basis is countable. -/ - countable : sets.Countable - --- For illustration purposes, the countable filter basis defining `(atTop : Filter ℕ)` -instance Nat.inhabitedCountableFilterBasis : Inhabited (CountableFilterBasis ℕ) := - ⟨⟨default, countable_range fun n => Ici n⟩⟩ - -theorem HasCountableBasis.isCountablyGenerated {f : Filter α} {p : ι → Prop} {s : ι → Set α} - (h : f.HasCountableBasis p s) : f.IsCountablyGenerated := - ⟨⟨{ t | ∃ i, p i ∧ s i = t }, h.countable.image s, h.toHasBasis.eq_generate⟩⟩ - -theorem HasBasis.isCountablyGenerated [Countable ι] {f : Filter α} {p : ι → Prop} {s : ι → Set α} - (h : f.HasBasis p s) : f.IsCountablyGenerated := - HasCountableBasis.isCountablyGenerated ⟨h, to_countable _⟩ - -theorem antitone_seq_of_seq (s : ℕ → Set α) : - ∃ t : ℕ → Set α, Antitone t ∧ ⨅ i, 𝓟 (s i) = ⨅ i, 𝓟 (t i) := by - use fun n => ⋂ m ≤ n, s m; constructor - · exact fun i j hij => biInter_mono (Iic_subset_Iic.2 hij) fun n _ => Subset.rfl - apply le_antisymm <;> rw [le_iInf_iff] <;> intro i - · rw [le_principal_iff] - refine (biInter_mem (finite_le_nat _)).2 fun j _ => ?_ - exact mem_iInf_of_mem j (mem_principal_self _) - · refine iInf_le_of_le i (principal_mono.2 <| iInter₂_subset i ?_) - rfl - -theorem countable_biInf_eq_iInf_seq [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) - (Bne : B.Nonempty) (f : ι → α) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := - let ⟨g, hg⟩ := Bcbl.exists_eq_range Bne - ⟨g, hg.symm ▸ iInf_range⟩ - -theorem countable_biInf_eq_iInf_seq' [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) - (f : ι → α) {i₀ : ι} (h : f i₀ = ⊤) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := by - rcases B.eq_empty_or_nonempty with hB | Bnonempty - · rw [hB, iInf_emptyset] - use fun _ => i₀ - simp [h] - · exact countable_biInf_eq_iInf_seq Bcbl Bnonempty f - -theorem countable_biInf_principal_eq_seq_iInf {B : Set (Set α)} (Bcbl : B.Countable) : - ∃ x : ℕ → Set α, ⨅ t ∈ B, 𝓟 t = ⨅ i, 𝓟 (x i) := - countable_biInf_eq_iInf_seq' Bcbl 𝓟 principal_univ - -section IsCountablyGenerated - -protected theorem HasAntitoneBasis.mem_iff [Preorder ι] {l : Filter α} {s : ι → Set α} - (hs : l.HasAntitoneBasis s) {t : Set α} : t ∈ l ↔ ∃ i, s i ⊆ t := - hs.toHasBasis.mem_iff.trans <| by simp only [exists_prop, true_and] - -protected theorem HasAntitoneBasis.mem [Preorder ι] {l : Filter α} {s : ι → Set α} - (hs : l.HasAntitoneBasis s) (i : ι) : s i ∈ l := - hs.toHasBasis.mem_of_mem trivial - -theorem HasAntitoneBasis.hasBasis_ge [Preorder ι] [IsDirected ι (· ≤ ·)] {l : Filter α} - {s : ι → Set α} (hs : l.HasAntitoneBasis s) (i : ι) : l.HasBasis (fun j => i ≤ j) s := - hs.1.to_hasBasis (fun j _ => (exists_ge_ge i j).imp fun _k hk => ⟨hk.1, hs.2 hk.2⟩) fun j _ => - ⟨j, trivial, Subset.rfl⟩ - -/-- If `f` is countably generated and `f.HasBasis p s`, then `f` admits a decreasing basis -enumerated by natural numbers such that all sets have the form `s i`. More precisely, there is a -sequence `i n` such that `p (i n)` for all `n` and `s (i n)` is a decreasing sequence of sets which -forms a basis of `f`-/ -theorem HasBasis.exists_antitone_subbasis {f : Filter α} [h : f.IsCountablyGenerated] - {p : ι' → Prop} {s : ι' → Set α} (hs : f.HasBasis p s) : - ∃ x : ℕ → ι', (∀ i, p (x i)) ∧ f.HasAntitoneBasis fun i => s (x i) := by - obtain ⟨x', hx'⟩ : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i) := by - rcases h with ⟨s, hsc, rfl⟩ - rw [generate_eq_biInf] - exact countable_biInf_principal_eq_seq_iInf hsc - have : ∀ i, x' i ∈ f := fun i => hx'.symm ▸ (iInf_le (fun i => 𝓟 (x' i)) i) (mem_principal_self _) - let x : ℕ → { i : ι' // p i } := fun n => - Nat.recOn n (hs.index _ <| this 0) fun n xn => - hs.index _ <| inter_mem (this <| n + 1) (hs.mem_of_mem xn.2) - have x_anti : Antitone fun i => s (x i).1 := - antitone_nat_of_succ_le fun i => (hs.set_index_subset _).trans inter_subset_right - have x_subset : ∀ i, s (x i).1 ⊆ x' i := by - rintro (_ | i) - exacts [hs.set_index_subset _, (hs.set_index_subset _).trans inter_subset_left] - refine ⟨fun i => (x i).1, fun i => (x i).2, ?_⟩ - have : (⨅ i, 𝓟 (s (x i).1)).HasAntitoneBasis fun i => s (x i).1 := .iInf_principal x_anti - convert this - exact - le_antisymm (le_iInf fun i => le_principal_iff.2 <| by cases i <;> apply hs.set_index_mem) - (hx'.symm ▸ - le_iInf fun i => le_principal_iff.2 <| this.1.mem_iff.2 ⟨i, trivial, x_subset i⟩) - -/-- A countably generated filter admits a basis formed by an antitone sequence of sets. -/ -theorem exists_antitone_basis (f : Filter α) [f.IsCountablyGenerated] : - ∃ x : ℕ → Set α, f.HasAntitoneBasis x := - let ⟨x, _, hx⟩ := f.basis_sets.exists_antitone_subbasis - ⟨x, hx⟩ - -theorem exists_antitone_seq (f : Filter α) [f.IsCountablyGenerated] : - ∃ x : ℕ → Set α, Antitone x ∧ ∀ {s}, s ∈ f ↔ ∃ i, x i ⊆ s := - let ⟨x, hx⟩ := f.exists_antitone_basis - ⟨x, hx.antitone, by simp [hx.1.mem_iff]⟩ - -instance Inf.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] - [IsCountablyGenerated g] : IsCountablyGenerated (f ⊓ g) := by - rcases f.exists_antitone_basis with ⟨s, hs⟩ - rcases g.exists_antitone_basis with ⟨t, ht⟩ - exact HasCountableBasis.isCountablyGenerated ⟨hs.1.inf ht.1, Set.to_countable _⟩ - -instance map.isCountablyGenerated (l : Filter α) [l.IsCountablyGenerated] (f : α → β) : - (map f l).IsCountablyGenerated := - let ⟨_x, hxl⟩ := l.exists_antitone_basis - (hxl.map _).isCountablyGenerated - -instance comap.isCountablyGenerated (l : Filter β) [l.IsCountablyGenerated] (f : α → β) : - (comap f l).IsCountablyGenerated := - let ⟨_x, hxl⟩ := l.exists_antitone_basis - (hxl.comap _).isCountablyGenerated - -instance Sup.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] - [IsCountablyGenerated g] : IsCountablyGenerated (f ⊔ g) := by - rcases f.exists_antitone_basis with ⟨s, hs⟩ - rcases g.exists_antitone_basis with ⟨t, ht⟩ - exact HasCountableBasis.isCountablyGenerated ⟨hs.1.sup ht.1, Set.to_countable _⟩ - -instance prod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] - [IsCountablyGenerated lb] : IsCountablyGenerated (la ×ˢ lb) := - Filter.Inf.isCountablyGenerated _ _ - -instance coprod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] - [IsCountablyGenerated lb] : IsCountablyGenerated (la.coprod lb) := - Filter.Sup.isCountablyGenerated _ _ - -end IsCountablyGenerated - -theorem isCountablyGenerated_seq [Countable ι'] (x : ι' → Set α) : - IsCountablyGenerated (⨅ i, 𝓟 (x i)) := by - use range x, countable_range x - rw [generate_eq_biInf, iInf_range] - -theorem isCountablyGenerated_of_seq {f : Filter α} (h : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i)) : - f.IsCountablyGenerated := by - rcases h with ⟨x, rfl⟩ - apply isCountablyGenerated_seq - -theorem isCountablyGenerated_biInf_principal {B : Set (Set α)} (h : B.Countable) : - IsCountablyGenerated (⨅ s ∈ B, 𝓟 s) := - isCountablyGenerated_of_seq (countable_biInf_principal_eq_seq_iInf h) - -theorem isCountablyGenerated_iff_exists_antitone_basis {f : Filter α} : - IsCountablyGenerated f ↔ ∃ x : ℕ → Set α, f.HasAntitoneBasis x := by - constructor - · intro h - exact f.exists_antitone_basis - · rintro ⟨x, h⟩ - rw [h.1.eq_iInf] - exact isCountablyGenerated_seq x - -@[instance] -theorem isCountablyGenerated_principal (s : Set α) : IsCountablyGenerated (𝓟 s) := - isCountablyGenerated_of_seq ⟨fun _ => s, iInf_const.symm⟩ - -@[instance] -theorem isCountablyGenerated_pure (a : α) : IsCountablyGenerated (pure a) := by - rw [← principal_singleton] - exact isCountablyGenerated_principal _ - -@[instance] -theorem isCountablyGenerated_bot : IsCountablyGenerated (⊥ : Filter α) := - @principal_empty α ▸ isCountablyGenerated_principal _ - -@[instance] -theorem isCountablyGenerated_top : IsCountablyGenerated (⊤ : Filter α) := - @principal_univ α ▸ isCountablyGenerated_principal _ - --- Porting note: without explicit `Sort u` and `Type v`, Lean 4 uses `ι : Prop` -universe u v - -instance iInf.isCountablyGenerated {ι : Sort u} {α : Type v} [Countable ι] (f : ι → Filter α) - [∀ i, IsCountablyGenerated (f i)] : IsCountablyGenerated (⨅ i, f i) := by - choose s hs using fun i => exists_antitone_basis (f i) - rw [← PLift.down_surjective.iInf_comp] - refine HasCountableBasis.isCountablyGenerated ⟨hasBasis_iInf fun n => (hs _).1, ?_⟩ - refine (countable_range <| Sigma.map ((↑) : Finset (PLift ι) → Set (PLift ι)) fun _ => id).mono ?_ - rintro ⟨I, f⟩ ⟨hI, -⟩ - lift I to Finset (PLift ι) using hI - exact ⟨⟨I, f⟩, rfl⟩ - -end Filter diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index 50fbb43ffe460..abcac7aaded93 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov import Mathlib.Data.Finite.Prod import Mathlib.Data.Fintype.Pi import Mathlib.Order.Filter.AtTopBot +import Mathlib.Order.Filter.CountablyGenerated import Mathlib.Order.Filter.Ker import Mathlib.Order.Filter.Pi diff --git a/Mathlib/Order/Filter/CountablyGenerated.lean b/Mathlib/Order/Filter/CountablyGenerated.lean new file mode 100644 index 0000000000000..dd106b65ece2d --- /dev/null +++ b/Mathlib/Order/Filter/CountablyGenerated.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2019 Gabriel Ebner. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gabriel Ebner, Yury Kudryashov, Patrick Massot +-/ +import Mathlib.Data.Set.Countable +import Mathlib.Order.Filter.Bases + +/-! +# Countably generated filters + +In this file we define a typeclass `Filter.IsCountablyGenerated` +saying that a filter is generated by a countable family of sets. + +We also define predicates `Filter.IsCountableBasis` and `Filter.HasCountableBasis` +saying that a specific family of sets is a countable basis. +-/ + +open Set + +namespace Filter + +variable {α β γ ι : Type*} {ι' : Sort*} + +/-- `IsCountablyGenerated f` means `f = generate s` for some countable `s`. -/ +class IsCountablyGenerated (f : Filter α) : Prop where + /-- There exists a countable set that generates the filter. -/ + out : ∃ s : Set (Set α), s.Countable ∧ f = generate s + +/-- `IsCountableBasis p s` means the image of `s` bounded by `p` is a countable filter basis. -/ +structure IsCountableBasis (p : ι → Prop) (s : ι → Set α) extends IsBasis p s : Prop where + /-- The set of `i` that satisfy the predicate `p` is countable. -/ + countable : (setOf p).Countable + +/-- We say that a filter `l` has a countable basis `s : ι → Set α` bounded by `p : ι → Prop`, +if `t ∈ l` if and only if `t` includes `s i` for some `i` such that `p i`, and the set +defined by `p` is countable. -/ +structure HasCountableBasis (l : Filter α) (p : ι → Prop) (s : ι → Set α) + extends HasBasis l p s : Prop where + /-- The set of `i` that satisfy the predicate `p` is countable. -/ + countable : (setOf p).Countable + +/-- A countable filter basis `B` on a type `α` is a nonempty countable collection of sets of `α` +such that the intersection of two elements of this collection contains some element +of the collection. -/ +structure CountableFilterBasis (α : Type*) extends FilterBasis α where + /-- The set of sets of the filter basis is countable. -/ + countable : sets.Countable + +-- For illustration purposes, the countable filter basis defining `(atTop : Filter ℕ)` +instance Nat.inhabitedCountableFilterBasis : Inhabited (CountableFilterBasis ℕ) := + ⟨⟨default, countable_range fun n => Ici n⟩⟩ + +theorem HasCountableBasis.isCountablyGenerated {f : Filter α} {p : ι → Prop} {s : ι → Set α} + (h : f.HasCountableBasis p s) : f.IsCountablyGenerated := + ⟨⟨{ t | ∃ i, p i ∧ s i = t }, h.countable.image s, h.toHasBasis.eq_generate⟩⟩ + +theorem HasBasis.isCountablyGenerated [Countable ι] {f : Filter α} {p : ι → Prop} {s : ι → Set α} + (h : f.HasBasis p s) : f.IsCountablyGenerated := + HasCountableBasis.isCountablyGenerated ⟨h, to_countable _⟩ + +theorem antitone_seq_of_seq (s : ℕ → Set α) : + ∃ t : ℕ → Set α, Antitone t ∧ ⨅ i, 𝓟 (s i) = ⨅ i, 𝓟 (t i) := by + use fun n => ⋂ m ≤ n, s m; constructor + · exact fun i j hij => biInter_mono (Iic_subset_Iic.2 hij) fun n _ => Subset.rfl + apply le_antisymm <;> rw [le_iInf_iff] <;> intro i + · rw [le_principal_iff] + refine (biInter_mem (finite_le_nat _)).2 fun j _ => ?_ + exact mem_iInf_of_mem j (mem_principal_self _) + · refine iInf_le_of_le i (principal_mono.2 <| iInter₂_subset i ?_) + rfl + +theorem countable_biInf_eq_iInf_seq [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) + (Bne : B.Nonempty) (f : ι → α) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := + let ⟨g, hg⟩ := Bcbl.exists_eq_range Bne + ⟨g, hg.symm ▸ iInf_range⟩ + +theorem countable_biInf_eq_iInf_seq' [CompleteLattice α] {B : Set ι} (Bcbl : B.Countable) + (f : ι → α) {i₀ : ι} (h : f i₀ = ⊤) : ∃ x : ℕ → ι, ⨅ t ∈ B, f t = ⨅ i, f (x i) := by + rcases B.eq_empty_or_nonempty with hB | Bnonempty + · rw [hB, iInf_emptyset] + use fun _ => i₀ + simp [h] + · exact countable_biInf_eq_iInf_seq Bcbl Bnonempty f + +theorem countable_biInf_principal_eq_seq_iInf {B : Set (Set α)} (Bcbl : B.Countable) : + ∃ x : ℕ → Set α, ⨅ t ∈ B, 𝓟 t = ⨅ i, 𝓟 (x i) := + countable_biInf_eq_iInf_seq' Bcbl 𝓟 principal_univ + +section IsCountablyGenerated + +protected theorem HasAntitoneBasis.mem_iff [Preorder ι] {l : Filter α} {s : ι → Set α} + (hs : l.HasAntitoneBasis s) {t : Set α} : t ∈ l ↔ ∃ i, s i ⊆ t := + hs.toHasBasis.mem_iff.trans <| by simp only [exists_prop, true_and] + +protected theorem HasAntitoneBasis.mem [Preorder ι] {l : Filter α} {s : ι → Set α} + (hs : l.HasAntitoneBasis s) (i : ι) : s i ∈ l := + hs.toHasBasis.mem_of_mem trivial + +theorem HasAntitoneBasis.hasBasis_ge [Preorder ι] [IsDirected ι (· ≤ ·)] {l : Filter α} + {s : ι → Set α} (hs : l.HasAntitoneBasis s) (i : ι) : l.HasBasis (fun j => i ≤ j) s := + hs.1.to_hasBasis (fun j _ => (exists_ge_ge i j).imp fun _k hk => ⟨hk.1, hs.2 hk.2⟩) fun j _ => + ⟨j, trivial, Subset.rfl⟩ + +/-- If `f` is countably generated and `f.HasBasis p s`, then `f` admits a decreasing basis +enumerated by natural numbers such that all sets have the form `s i`. More precisely, there is a +sequence `i n` such that `p (i n)` for all `n` and `s (i n)` is a decreasing sequence of sets which +forms a basis of `f`-/ +theorem HasBasis.exists_antitone_subbasis {f : Filter α} [h : f.IsCountablyGenerated] + {p : ι' → Prop} {s : ι' → Set α} (hs : f.HasBasis p s) : + ∃ x : ℕ → ι', (∀ i, p (x i)) ∧ f.HasAntitoneBasis fun i => s (x i) := by + obtain ⟨x', hx'⟩ : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i) := by + rcases h with ⟨s, hsc, rfl⟩ + rw [generate_eq_biInf] + exact countable_biInf_principal_eq_seq_iInf hsc + have : ∀ i, x' i ∈ f := fun i => hx'.symm ▸ (iInf_le (fun i => 𝓟 (x' i)) i) (mem_principal_self _) + let x : ℕ → { i : ι' // p i } := fun n => + Nat.recOn n (hs.index _ <| this 0) fun n xn => + hs.index _ <| inter_mem (this <| n + 1) (hs.mem_of_mem xn.2) + have x_anti : Antitone fun i => s (x i).1 := + antitone_nat_of_succ_le fun i => (hs.set_index_subset _).trans inter_subset_right + have x_subset : ∀ i, s (x i).1 ⊆ x' i := by + rintro (_ | i) + exacts [hs.set_index_subset _, (hs.set_index_subset _).trans inter_subset_left] + refine ⟨fun i => (x i).1, fun i => (x i).2, ?_⟩ + have : (⨅ i, 𝓟 (s (x i).1)).HasAntitoneBasis fun i => s (x i).1 := .iInf_principal x_anti + convert this + exact + le_antisymm (le_iInf fun i => le_principal_iff.2 <| by cases i <;> apply hs.set_index_mem) + (hx'.symm ▸ + le_iInf fun i => le_principal_iff.2 <| this.1.mem_iff.2 ⟨i, trivial, x_subset i⟩) + +/-- A countably generated filter admits a basis formed by an antitone sequence of sets. -/ +theorem exists_antitone_basis (f : Filter α) [f.IsCountablyGenerated] : + ∃ x : ℕ → Set α, f.HasAntitoneBasis x := + let ⟨x, _, hx⟩ := f.basis_sets.exists_antitone_subbasis + ⟨x, hx⟩ + +theorem exists_antitone_seq (f : Filter α) [f.IsCountablyGenerated] : + ∃ x : ℕ → Set α, Antitone x ∧ ∀ {s}, s ∈ f ↔ ∃ i, x i ⊆ s := + let ⟨x, hx⟩ := f.exists_antitone_basis + ⟨x, hx.antitone, by simp [hx.1.mem_iff]⟩ + +instance Inf.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] + [IsCountablyGenerated g] : IsCountablyGenerated (f ⊓ g) := by + rcases f.exists_antitone_basis with ⟨s, hs⟩ + rcases g.exists_antitone_basis with ⟨t, ht⟩ + exact HasCountableBasis.isCountablyGenerated ⟨hs.1.inf ht.1, Set.to_countable _⟩ + +instance map.isCountablyGenerated (l : Filter α) [l.IsCountablyGenerated] (f : α → β) : + (map f l).IsCountablyGenerated := + let ⟨_x, hxl⟩ := l.exists_antitone_basis + (hxl.map _).isCountablyGenerated + +instance comap.isCountablyGenerated (l : Filter β) [l.IsCountablyGenerated] (f : α → β) : + (comap f l).IsCountablyGenerated := + let ⟨_x, hxl⟩ := l.exists_antitone_basis + (hxl.comap _).isCountablyGenerated + +instance Sup.isCountablyGenerated (f g : Filter α) [IsCountablyGenerated f] + [IsCountablyGenerated g] : IsCountablyGenerated (f ⊔ g) := by + rcases f.exists_antitone_basis with ⟨s, hs⟩ + rcases g.exists_antitone_basis with ⟨t, ht⟩ + exact HasCountableBasis.isCountablyGenerated ⟨hs.1.sup ht.1, Set.to_countable _⟩ + +instance prod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] + [IsCountablyGenerated lb] : IsCountablyGenerated (la ×ˢ lb) := + Filter.Inf.isCountablyGenerated _ _ + +instance coprod.isCountablyGenerated (la : Filter α) (lb : Filter β) [IsCountablyGenerated la] + [IsCountablyGenerated lb] : IsCountablyGenerated (la.coprod lb) := + Filter.Sup.isCountablyGenerated _ _ + +end IsCountablyGenerated + +theorem isCountablyGenerated_seq [Countable ι'] (x : ι' → Set α) : + IsCountablyGenerated (⨅ i, 𝓟 (x i)) := by + use range x, countable_range x + rw [generate_eq_biInf, iInf_range] + +theorem isCountablyGenerated_of_seq {f : Filter α} (h : ∃ x : ℕ → Set α, f = ⨅ i, 𝓟 (x i)) : + f.IsCountablyGenerated := by + rcases h with ⟨x, rfl⟩ + apply isCountablyGenerated_seq + +theorem isCountablyGenerated_biInf_principal {B : Set (Set α)} (h : B.Countable) : + IsCountablyGenerated (⨅ s ∈ B, 𝓟 s) := + isCountablyGenerated_of_seq (countable_biInf_principal_eq_seq_iInf h) + +theorem isCountablyGenerated_iff_exists_antitone_basis {f : Filter α} : + IsCountablyGenerated f ↔ ∃ x : ℕ → Set α, f.HasAntitoneBasis x := by + constructor + · intro h + exact f.exists_antitone_basis + · rintro ⟨x, h⟩ + rw [h.1.eq_iInf] + exact isCountablyGenerated_seq x + +@[instance] +theorem isCountablyGenerated_principal (s : Set α) : IsCountablyGenerated (𝓟 s) := + isCountablyGenerated_of_seq ⟨fun _ => s, iInf_const.symm⟩ + +@[instance] +theorem isCountablyGenerated_pure (a : α) : IsCountablyGenerated (pure a) := by + rw [← principal_singleton] + exact isCountablyGenerated_principal _ + +@[instance] +theorem isCountablyGenerated_bot : IsCountablyGenerated (⊥ : Filter α) := + @principal_empty α ▸ isCountablyGenerated_principal _ + +@[instance] +theorem isCountablyGenerated_top : IsCountablyGenerated (⊤ : Filter α) := + @principal_univ α ▸ isCountablyGenerated_principal _ + +-- Porting note: without explicit `Sort u` and `Type v`, Lean 4 uses `ι : Prop` +universe u v + +instance iInf.isCountablyGenerated {ι : Sort u} {α : Type v} [Countable ι] (f : ι → Filter α) + [∀ i, IsCountablyGenerated (f i)] : IsCountablyGenerated (⨅ i, f i) := by + choose s hs using fun i => exists_antitone_basis (f i) + rw [← PLift.down_surjective.iInf_comp] + refine HasCountableBasis.isCountablyGenerated ⟨hasBasis_iInf fun n => (hs _).1, ?_⟩ + refine (countable_range <| Sigma.map ((↑) : Finset (PLift ι) → Set (PLift ι)) fun _ => id).mono ?_ + rintro ⟨I, f⟩ ⟨hI, -⟩ + lift I to Finset (PLift ι) using hI + exact ⟨⟨I, f⟩, rfl⟩ + +instance pi.isCountablyGenerated {ι : Type*} {α : ι → Type*} [Countable ι] + (f : ∀ i, Filter (α i)) [∀ i, IsCountablyGenerated (f i)] : IsCountablyGenerated (pi f) := + iInf.isCountablyGenerated _ + +end Filter diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index 432915dcbac1d..769d8eb31e206 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -319,6 +319,10 @@ protected def prod (f : Filter α) (g : Filter β) : Filter (α × β) := f ×ˢ theorem prod_eq_inf (f : Filter α) (g : Filter β) : f ×ˢ g = f.comap Prod.fst ⊓ g.comap Prod.snd := rfl +/-- The product of an indexed family of filters. -/ +def pi {ι : Type*} {α : ι → Type*} (f : ∀ i, Filter (α i)) : Filter (∀ i, α i) := + ⨅ i, comap (Function.eval i) (f i) + /-- The monadic bind operation on filter is defined the usual way in terms of `map` and `join`. Unfortunately, this `bind` does not result in the expected applicative. See `Filter.seq` for the diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index b7df9bc2a0c66..fe099e6b204fa 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -198,7 +198,7 @@ theorem map_map (op₁ : γ → δ) (op₂ : β → γ) (f : Germ l β) : /-- Lift a binary function `β → γ → δ` to a function `Germ l β → Germ l γ → Germ l δ`. -/ def map₂ (op : β → γ → δ) : Germ l β → Germ l γ → Germ l δ := - Quotient.map₂' (fun f g x => op (f x) (g x)) fun f f' Hf g g' Hg => + Quotient.map₂ (fun f g x => op (f x) (g x)) fun f f' Hf g g' Hg => Hg.mp <| Hf.mono fun x Hf Hg => by simp only [Hf, Hg] @[simp] diff --git a/Mathlib/Order/Filter/Pi.lean b/Mathlib/Order/Filter/Pi.lean index 10716f6ff6edc..e583584b9123e 100644 --- a/Mathlib/Order/Filter/Pi.lean +++ b/Mathlib/Order/Filter/Pi.lean @@ -30,14 +30,6 @@ variable {ι : Type*} {α : ι → Type*} {f f₁ f₂ : (i : ι) → Filter (α section Pi -/-- The product of an indexed family of filters. -/ -def pi (f : ∀ i, Filter (α i)) : Filter (∀ i, α i) := - ⨅ i, comap (eval i) (f i) - -instance pi.isCountablyGenerated [Countable ι] [∀ i, IsCountablyGenerated (f i)] : - IsCountablyGenerated (pi f) := - iInf.isCountablyGenerated _ - theorem tendsto_eval_pi (f : ∀ i, Filter (α i)) (i : ι) : Tendsto (eval i) (pi f) (f i) := tendsto_iInf' i tendsto_comap diff --git a/Mathlib/Order/Fin/Basic.lean b/Mathlib/Order/Fin/Basic.lean index 10983f1527c03..4e47785db09e5 100644 --- a/Mathlib/Order/Fin/Basic.lean +++ b/Mathlib/Order/Fin/Basic.lean @@ -122,7 +122,7 @@ end FromFin /-! #### Monotonicity -/ lemma val_strictMono : StrictMono (val : Fin n → ℕ) := fun _ _ ↦ id -lemma cast_strictMono {k l : ℕ} (h : k = l) : StrictMono (cast h) := fun {_ _} h ↦ h +lemma cast_strictMono {k l : ℕ} (h : k = l) : StrictMono (Fin.cast h) := fun {_ _} h ↦ h lemma strictMono_succ : StrictMono (succ : Fin n → Fin (n + 1)) := fun _ _ ↦ succ_lt_succ lemma strictMono_castLE (h : n ≤ m) : StrictMono (castLE h : Fin n → Fin m) := fun _ _ ↦ id @@ -182,7 +182,7 @@ def orderIsoSubtype : Fin n ≃o {i // i < n} := `castOrderIso eq i` embeds `i` into an equal `Fin` type. -/ @[simps] def castOrderIso (eq : n = m) : Fin n ≃o Fin m where - toEquiv := ⟨cast eq, cast eq.symm, leftInverse_cast eq, rightInverse_cast eq⟩ + toEquiv := ⟨Fin.cast eq, Fin.cast eq.symm, leftInverse_cast eq, rightInverse_cast eq⟩ map_rel_iff' := cast_le_cast eq @[deprecated (since := "2024-05-23")] alias castIso := castOrderIso diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index c34eaa969055e..9e76dd904111b 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -296,8 +296,7 @@ lemma ωScottContinuous.comp (hg : ωScottContinuous g) (hf : ωScottContinuous lemma ωScottContinuous.const {x : β} : ωScottContinuous (Function.const α x) := by simp [ωScottContinuous, ScottContinuousOn, Set.range_nonempty] -set_option linter.deprecated false - +set_option linter.deprecated false in /-- A monotone function `f : α →o β` is continuous if it distributes over ωSup. In order to distinguish it from the (more commonly used) continuity from topology @@ -308,6 +307,7 @@ in Scott topological spaces (not defined here). -/ def Continuous (f : α →o β) : Prop := ∀ c : Chain α, f (ωSup c) = ωSup (c.map f) +set_option linter.deprecated false in /-- `Continuous' f` asserts that `f` is both monotone and continuous. -/ @[deprecated ωScottContinuous (since := "2024-05-29")] def Continuous' (f : α → β) : Prop := @@ -318,6 +318,7 @@ lemma isLUB_of_scottContinuous {c : Chain α} {f : α → β} (hf : ScottContinu IsLUB (Set.range (Chain.map c ⟨f, (ScottContinuous.monotone hf)⟩)) (f (ωSup c)) := ωScottContinuous.isLUB hf.scottContinuousOn +set_option linter.deprecated false in @[deprecated ScottContinuous.ωScottContinuous (since := "2024-05-29")] lemma ScottContinuous.continuous' {f : α → β} (hf : ScottContinuous f) : Continuous' f := by constructor @@ -325,45 +326,55 @@ lemma ScottContinuous.continuous' {f : α → β} (hf : ScottContinuous f) : Con rw [← (ωSup_eq_of_isLUB (isLUB_of_scottContinuous hf))] simp only [OrderHom.coe_mk] +set_option linter.deprecated false in @[deprecated ωScottContinuous.monotone (since := "2024-05-29")] theorem Continuous'.to_monotone {f : α → β} (hf : Continuous' f) : Monotone f := hf.fst +set_option linter.deprecated false in @[deprecated ωScottContinuous.of_monotone_map_ωSup (since := "2024-05-29")] theorem Continuous.of_bundled (f : α → β) (hf : Monotone f) (hf' : Continuous ⟨f, hf⟩) : Continuous' f := ⟨hf, hf'⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous.of_monotone_map_ωSup (since := "2024-05-29")] theorem Continuous.of_bundled' (f : α →o β) (hf' : Continuous f) : Continuous' f := ⟨f.mono, hf'⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous_iff_monotone_map_ωSup (since := "2024-05-29")] theorem Continuous'.to_bundled (f : α → β) (hf : Continuous' f) : Continuous ⟨f, hf.to_monotone⟩ := hf.snd +set_option linter.deprecated false in @[simp, norm_cast, deprecated ωScottContinuous_iff_monotone_map_ωSup (since := "2024-05-29")] theorem continuous'_coe : ∀ {f : α →o β}, Continuous' f ↔ Continuous f | ⟨_, hf⟩ => ⟨fun ⟨_, hc⟩ => hc, fun hc => ⟨hf, hc⟩⟩ variable (f : α →o β) (g : β →o γ) +set_option linter.deprecated false in @[deprecated ωScottContinuous.id (since := "2024-05-29")] theorem continuous_id : Continuous (@OrderHom.id α _) := by intro c; rw [c.map_id]; rfl +set_option linter.deprecated false in @[deprecated ωScottContinuous.comp (since := "2024-05-29")] theorem continuous_comp (hfc : Continuous f) (hgc : Continuous g) : Continuous (g.comp f) := by dsimp [Continuous] at *; intro rw [hfc, hgc, Chain.map_comp] +set_option linter.deprecated false in @[deprecated ωScottContinuous.id (since := "2024-05-29")] theorem id_continuous' : Continuous' (@id α) := continuous_id.of_bundled' _ +set_option linter.deprecated false in @[deprecated ωScottContinuous.const (since := "2024-05-29")] theorem continuous_const (x : β) : Continuous (OrderHom.const α x) := fun c => eq_of_forall_ge_iff fun z => by rw [ωSup_le_iff, Chain.map_coe, OrderHom.const_coe_coe]; simp +set_option linter.deprecated false in @[deprecated ωScottContinuous.const (since := "2024-05-29")] theorem const_continuous' (x : β) : Continuous' (Function.const α x) := Continuous.of_bundled' (OrderHom.const α x) (continuous_const x) @@ -484,13 +495,13 @@ lemma ωScottContinuous.of_apply₂ (hf : ∀ a, ωScottContinuous (f · a)) : lemma ωScottContinuous_iff_apply₂ : ωScottContinuous f ↔ ∀ a, ωScottContinuous (f · a) := ⟨ωScottContinuous.apply₂, ωScottContinuous.of_apply₂⟩ -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.apply₂ (since := "2024-05-29")] theorem flip₁_continuous' (f : ∀ x : α, γ → β x) (a : α) (hf : Continuous' fun x y => f y x) : Continuous' (f a) := Continuous.of_bundled _ (fun _ _ h => hf.to_monotone h a) fun c => congr_fun (hf.to_bundled _ c) a +set_option linter.deprecated false in @[deprecated ωScottContinuous.of_apply₂ (since := "2024-05-29")] theorem flip₂_continuous' (f : γ → ∀ x, β x) (hf : ∀ x, Continuous' fun g => f g x) : Continuous' f := @@ -571,8 +582,7 @@ lemma ωScottContinuous.top : ωScottContinuous (⊤ : α → β) := lemma ωScottContinuous.bot : ωScottContinuous (⊥ : α → β) := by rw [← sSup_empty]; exact ωScottContinuous.sSup (by simp) -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.sSup (since := "2024-05-29")] theorem sSup_continuous (s : Set <| α →o β) (hs : ∀ f ∈ s, Continuous f) : Continuous (sSup s) := by intro c @@ -582,11 +592,13 @@ theorem sSup_continuous (s : Set <| α →o β) (hs : ∀ f ∈ s, Continuous f) simpa (config := { contextual := true }) [ωSup_le_iff, hs _ _ _] using this exact ⟨fun H n f hf => H f hf n, fun H f hf n => H n f hf⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous.iSup (since := "2024-05-29")] theorem iSup_continuous {ι : Sort*} {f : ι → α →o β} (h : ∀ i, Continuous (f i)) : Continuous (⨆ i, f i) := sSup_continuous _ <| Set.forall_mem_range.2 h +set_option linter.deprecated false in @[deprecated ωScottContinuous.sSup (since := "2024-05-29")] theorem sSup_continuous' (s : Set (α → β)) (hc : ∀ f ∈ s, Continuous' f) : Continuous' (sSup s) := by @@ -596,18 +608,21 @@ theorem sSup_continuous' (s : Set (α → β)) (hc : ∀ f ∈ s, Continuous' f) norm_cast exact iSup_continuous fun f ↦ iSup_continuous fun hf ↦ hc hf +set_option linter.deprecated false in @[deprecated ωScottContinuous.sup (since := "2024-05-29")] theorem sup_continuous {f g : α →o β} (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊔ g) := by rw [← sSup_pair]; apply sSup_continuous rintro f (rfl | rfl | _) <;> assumption +set_option linter.deprecated false in @[deprecated ωScottContinuous.top (since := "2024-05-29")] theorem top_continuous : Continuous (⊤ : α →o β) := by intro c; apply eq_of_forall_ge_iff; intro z simp only [OrderHom.instTopOrderHom_top, OrderHom.const_coe_coe, Function.const, top_le_iff, ωSup_le_iff, Chain.map_coe, Function.comp, forall_const] +set_option linter.deprecated false in @[deprecated ωScottContinuous.bot (since := "2024-05-29")] theorem bot_continuous : Continuous (⊥ : α →o β) := by rw [← sSup_empty] @@ -635,8 +650,7 @@ lemma ωScottContinuous.inf (hf : ωScottContinuous f) (hg : ωScottContinuous g (h (max j i)).imp (le_trans <| hf.monotone <| c.mono <| le_max_left _ _) (le_trans <| hg.monotone <| c.mono <| le_max_right _ _)⟩ -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.inf (since := "2024-05-29")] theorem inf_continuous (f g : α →o β) (hf : Continuous f) (hg : Continuous g) : Continuous (f ⊓ g) := by @@ -647,6 +661,7 @@ theorem inf_continuous (f g : α →o β) (hf : Continuous f) (hg : Continuous g (h (max j i)).imp (le_trans <| f.mono <| c.mono <| le_max_left _ _) (le_trans <| g.mono <| c.mono <| le_max_right _ _)⟩ +set_option linter.deprecated false in @[deprecated ωScottContinuous.inf (since := "2024-05-29")] theorem inf_continuous' {f g : α → β} (hf : Continuous' f) (hg : Continuous' g) : Continuous' (f ⊓ g) := @@ -779,8 +794,7 @@ lemma ωScottContinuous.seq {β γ} {f : α → Part (β → γ)} {g : α → Pa simp only [seq_eq_bind_map] exact ωScottContinuous.bind hf <| ωScottContinuous.of_apply₂ fun _ ↦ ωScottContinuous.map hg -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated ωScottContinuous.bind (since := "2024-05-29")] theorem bind_continuous' {β γ : Type v} (f : α → Part β) (g : α → β → Part γ) : Continuous' f → Continuous' g → Continuous' fun x => f x >>= g x @@ -788,11 +802,13 @@ theorem bind_continuous' {β γ : Type v} (f : α → Part β) (g : α → β Continuous.of_bundled' (OrderHom.partBind ⟨f, hf⟩ ⟨g, hg⟩) (by intro c; rw [ωSup_bind, ← hf', ← hg']; rfl) +set_option linter.deprecated false in @[deprecated ωScottContinuous.map (since := "2024-05-29")] theorem map_continuous' {β γ : Type v} (f : β → γ) (g : α → Part β) (hg : Continuous' g) : Continuous' fun x => f <$> g x := by simp only [map_eq_bind_pure_comp]; apply bind_continuous' _ _ hg; apply const_continuous' +set_option linter.deprecated false in @[deprecated ωScottContinuous.seq (since := "2024-05-29")] theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α → Part β) (hf : Continuous' f) (hg : Continuous' g) : Continuous' fun x => f x <*> g x := by @@ -802,8 +818,6 @@ theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α intro apply map_continuous' _ _ hg -set_option linter.deprecated true - theorem continuous (F : α →𝒄 β) (C : Chain α) : F (ωSup C) = ωSup (C.map F) := F.ωScottContinuous.map_ωSup _ diff --git a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean index c8ab06d7b669a..eea7681180495 100644 --- a/Mathlib/Order/SuccPred/CompleteLinearOrder.lean +++ b/Mathlib/Order/SuccPred/CompleteLinearOrder.lean @@ -11,7 +11,7 @@ import Mathlib.Order.SuccPred.Limit -/ -open Order +open Order Set variable {ι α : Type*} @@ -144,6 +144,32 @@ lemma IsLUB.exists_of_not_isSuccPrelimit (hf : IsLUB (Set.range f) x) (hx : ¬ I @[deprecated IsLUB.exists_of_not_isSuccPrelimit (since := "2024-09-05")] alias IsLUB.exists_of_not_isSuccLimit := IsLUB.exists_of_not_isSuccPrelimit +theorem Order.IsSuccPrelimit.sSup_Iio (h : IsSuccPrelimit x) : sSup (Iio x) = x := by + obtain rfl | hx := eq_bot_or_bot_lt x + · simp + · refine (csSup_le ⟨⊥, hx⟩ fun a ha ↦ ha.le).antisymm <| le_of_forall_lt fun a ha ↦ ?_ + rw [lt_csSup_iff' bddAbove_Iio] + obtain ⟨b, hb', hb⟩ := (not_covBy_iff ha).1 (h a) + use b, hb + +theorem Order.IsSuccPrelimit.iSup_Iio (h : IsSuccPrelimit x) : ⨆ a : Iio x, a.1 = x := by + rw [← sSup_eq_iSup', h.sSup_Iio] + +theorem Order.IsSuccLimit.sSup_Iio (h : IsSuccLimit x) : sSup (Iio x) = x := + h.isSuccPrelimit.sSup_Iio + +theorem Order.IsSuccLimit.iSup_Iio (h : IsSuccLimit x) : ⨆ a : Iio x, a.1 = x := + h.isSuccPrelimit.iSup_Iio + +theorem sSup_Iio_eq_self_iff_isSuccPrelimit : sSup (Iio x) = x ↔ IsSuccPrelimit x := by + refine ⟨fun h ↦ ?_, IsSuccPrelimit.sSup_Iio⟩ + by_contra hx + rw [← h] at hx + simpa [h] using csSup_mem_of_not_isSuccPrelimit' bddAbove_Iio hx + +theorem iSup_Iio_eq_self_iff_isSuccPrelimit : ⨆ a : Iio x, a.1 = x ↔ IsSuccPrelimit x := by + rw [← sSup_eq_iSup', sSup_Iio_eq_self_iff_isSuccPrelimit] + end ConditionallyCompleteLinearOrderBot section CompleteLinearOrder diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index 8929a9eebe09a..6170a1e80ff57 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -193,6 +193,9 @@ variable [PartialOrder α] theorem isSuccLimit_iff [OrderBot α] : IsSuccLimit a ↔ a ≠ ⊥ ∧ IsSuccPrelimit a := by rw [IsSuccLimit, isMin_iff_eq_bot] +theorem IsSuccLimit.bot_lt [OrderBot α] (h : IsSuccLimit a) : ⊥ < a := + h.ne_bot.bot_lt + variable [SuccOrder α] theorem isSuccPrelimit_of_succ_ne (h : ∀ b, succ b ≠ a) : IsSuccPrelimit a := fun b hba => @@ -223,6 +226,12 @@ theorem mem_range_succ_or_isSuccPrelimit (a) : a ∈ range (succ : α → α) @[deprecated mem_range_succ_or_isSuccPrelimit (since := "2024-09-05")] alias mem_range_succ_or_isSuccLimit := mem_range_succ_or_isSuccPrelimit +theorem isMin_or_mem_range_succ_or_isSuccLimit (a) : + IsMin a ∨ a ∈ range (succ : α → α) ∨ IsSuccLimit a := by + rw [IsSuccLimit] + have := mem_range_succ_or_isSuccPrelimit a + tauto + theorem isSuccPrelimit_of_succ_lt (H : ∀ a < b, succ a < b) : IsSuccPrelimit b := fun a hab => (H a hab.lt).ne (CovBy.succ_eq hab) @@ -490,6 +499,9 @@ variable [PartialOrder α] theorem isPredLimit_iff [OrderTop α] : IsPredLimit a ↔ a ≠ ⊤ ∧ IsPredPrelimit a := by rw [IsPredLimit, isMax_iff_eq_top] +theorem IsPredLimit.lt_top [OrderTop α] (h : IsPredLimit a) : a < ⊤ := + h.ne_top.lt_top + variable [PredOrder α] theorem isPredPrelimit_of_pred_ne (h : ∀ b, pred b ≠ a) : IsPredPrelimit a := fun b hba => diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 0511b5790ae02..17b5c843b5ec7 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -90,9 +90,8 @@ protected theorem lt_sup {r : α → α → Prop} (wf : WellFounded r) {s : Set section deprecated -set_option linter.deprecated false - open Classical in +set_option linter.deprecated false in /-- A successor of an element `x` in a well-founded order is a minimal element `y` such that `x < y` if one exists. Otherwise it is `x` itself. -/ @[deprecated "If you have a linear order, consider defining a `SuccOrder` instance through @@ -100,12 +99,14 @@ open Classical in protected noncomputable def succ {r : α → α → Prop} (wf : WellFounded r) (x : α) : α := if h : ∃ y, r x y then wf.min { y | r x y } h else x +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-25")] protected theorem lt_succ {r : α → α → Prop} (wf : WellFounded r) {x : α} (h : ∃ y, r x y) : r x (wf.succ x) := by rw [WellFounded.succ, dif_pos h] apply min_mem +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-25")] protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] {x : α} (h : ∃ y, r x y) (y : α) : r y (wo.wf.succ x) ↔ r y x ∨ y = x := by diff --git a/Mathlib/Order/Zorn.lean b/Mathlib/Order/Zorn.lean index b0418fa34a94b..833cc2e53c890 100644 --- a/Mathlib/Order/Zorn.lean +++ b/Mathlib/Order/Zorn.lean @@ -178,3 +178,23 @@ theorem IsChain.exists_maxChain (hc : IsChain r c) : ∃ M, @IsMaxChain _ r M cases' hcs₁ hsy hsz hsseq with h h · exact (hcs₀ hsz).right (h hysy) hzsz hyz · exact (hcs₀ hsy).right hysy (h hzsz) hyz + +/-! ### Flags -/ + +namespace Flag + +variable [Preorder α] {c : Set α} {s : Flag α} {a b : α} + +lemma _root_.IsChain.exists_subset_flag (hc : IsChain (· ≤ ·) c) : ∃ s : Flag α, c ⊆ s := + let ⟨s, hs, hcs⟩ := hc.exists_maxChain; ⟨ofIsMaxChain s hs, hcs⟩ + +lemma exists_mem (a : α) : ∃ s : Flag α, a ∈ s := + let ⟨s, hs⟩ := Set.subsingleton_singleton (a := a).isChain.exists_subset_flag + ⟨s, hs rfl⟩ + +lemma exists_mem_mem (hab : a ≤ b) : ∃ s : Flag α, a ∈ s ∧ b ∈ s := by + simpa [Set.insert_subset_iff] using (IsChain.pair hab).exists_subset_flag + +instance : Nonempty (Flag α) := ⟨.ofIsMaxChain _ maxChain_spec⟩ + +end Flag diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition/Basic.lean similarity index 100% rename from Mathlib/Probability/Kernel/Composition.lean rename to Mathlib/Probability/Kernel/Composition/Basic.lean diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean similarity index 99% rename from Mathlib/Probability/Kernel/IntegralCompProd.lean rename to Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean index e88e7a0f5f916..0fab071f69bea 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/IntegralCompProd.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.MeasureTheory.Integral.SetIntegral /-! diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean similarity index 99% rename from Mathlib/Probability/Kernel/MeasureCompProd.lean rename to Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean index 9a7d41a2b61ca..e15b968b5a1ba 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/Composition/MeasureCompProd.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.IntegralCompProd +import Mathlib.Probability.Kernel.Composition.IntegralCompProd /-! # Composition-Product of a measure and a kernel diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 22934beefaf08..8d443a75e791c 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Yaël Dillies, Kin Yau James Wong. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Kin Yau James Wong, Rémy Degenne -/ -import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Composition.MeasureCompProd /-! # Disintegration of measures and kernels diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 4bdb5515cd2ee..a8558585a52da 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Function.AEEqOfIntegral -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes /-! diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 94af1027b716b..316c64c2fce93 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Process.PartitionFiltration diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index d89d6a2aefef8..9f1b75f64926a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Composition.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.Basic import Mathlib.Probability.Kernel.Disintegration.CondCDF import Mathlib.Probability.Kernel.Disintegration.Density diff --git a/Mathlib/Probability/Kernel/Invariance.lean b/Mathlib/Probability/Kernel/Invariance.lean index e9d40e85a0dd9..c58725ae3dd90 100644 --- a/Mathlib/Probability/Kernel/Invariance.lean +++ b/Mathlib/Probability/Kernel/Invariance.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import Mathlib.Probability.Kernel.Composition +import Mathlib.Probability.Kernel.Composition.Basic /-! # Invariance of measures along a kernel diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index 5479f9417977f..cb98ce35c607a 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -116,8 +116,8 @@ theorem char_orthonormal (V W : FDRep k G) [Simple V] [Simple W] : rw [char_iso (FDRep.dualTensorIsoLinHom W.ρ V)] -- The average over the group of the character of a representation equals the dimension of the -- space of invariants. - rw [average_char_eq_finrank_invariants] - rw [show (of (linHom W.ρ V.ρ)).ρ = linHom W.ρ V.ρ from FDRep.of_ρ (linHom W.ρ V.ρ)] + rw [average_char_eq_finrank_invariants, ← FDRep.endMulEquiv_comp_ρ (of _), + FDRep.of_ρ (linHom W.ρ V.ρ)] -- The space of invariants of `Hom(W, V)` is the subspace of `G`-equivariant linear maps, -- `Hom_G(W, V)`. erw [(linHom.invariantsEquivFDRepHom W V).finrank_eq] -- Porting note: Changed `rw` to `erw` diff --git a/Mathlib/RepresentationTheory/FDRep.lean b/Mathlib/RepresentationTheory/FDRep.lean index 1f3f8535f0114..c3217c02d5c69 100644 --- a/Mathlib/RepresentationTheory/FDRep.lean +++ b/Mathlib/RepresentationTheory/FDRep.lean @@ -82,7 +82,18 @@ instance (V W : FDRep k G) : FiniteDimensional k (V ⟶ W) := /-- The monoid homomorphism corresponding to the action of `G` onto `V : FDRep k G`. -/ def ρ (V : FDRep k G) : G →* V →ₗ[k] V := - Action.ρ V + (ModuleCat.endMulEquiv _).toMonoidHom.comp (Action.ρ V) + +@[simp] +lemma endMulEquiv_symm_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endMulEquiv V.V.obj).symm).comp (ρ V) = Action.ρ V := rfl + +@[simp] +lemma endMulEquiv_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endMulEquiv V.V.obj)).comp (Action.ρ V) = ρ V := rfl + +@[simp] +lemma hom_action_ρ (V : FDRep k G) (g : G) : (Action.ρ V g).hom = ρ V g := rfl /-- The underlying `LinearEquiv` of an isomorphism of representations. -/ def isoToLinearEquiv {V W : FDRep k G} (i : V ≅ W) : V ≃ₗ[k] W := @@ -91,15 +102,16 @@ def isoToLinearEquiv {V W : FDRep k G} (i : V ≅ W) : V ≃ₗ[k] W := theorem Iso.conj_ρ {V W : FDRep k G} (i : V ≅ W) (g : G) : W.ρ g = (FDRep.isoToLinearEquiv i).conj (V.ρ g) := by -- Porting note: Changed `rw` to `erw` - erw [FDRep.isoToLinearEquiv, ← FGModuleCat.Iso.conj_eq_conj, Iso.conj_apply] - rw [Iso.eq_inv_comp ((Action.forget (FGModuleCat k) (MonCat.of G)).mapIso i)] + erw [FDRep.isoToLinearEquiv, ← hom_action_ρ V, ← FGModuleCat.Iso.conj_hom_eq_conj, Iso.conj_apply] + rw [← ModuleCat.hom_ofHom (W.ρ g), ← ModuleCat.hom_ext_iff, + Iso.eq_inv_comp ((Action.forget (FGModuleCat k) (MonCat.of G)).mapIso i)] exact (i.hom.comm g).symm /-- Lift an unbundled representation to `FDRep`. -/ @[simps ρ] def of {V : Type u} [AddCommGroup V] [Module k V] [FiniteDimensional k V] (ρ : Representation k G V) : FDRep k G := - ⟨FGModuleCat.of k V, ρ⟩ + ⟨FGModuleCat.of k V, ρ ≫ MonCat.ofHom (ModuleCat.endMulEquiv _).symm.toMonoidHom⟩ instance : HasForget₂ (FDRep k G) (Rep k G) where forget₂ := (forget₂ (FGModuleCat k) (ModuleCat k)).mapAction (MonCat.of G) @@ -179,11 +191,13 @@ noncomputable def dualTensorIsoLinHomAux : /-- When `V` and `W` are finite dimensional representations of a group `G`, the isomorphism `dualTensorHomEquiv k V W` of vector spaces induces an isomorphism of representations. -/ noncomputable def dualTensorIsoLinHom : FDRep.of ρV.dual ⊗ W ≅ FDRep.of (linHom ρV W.ρ) := by - refine Action.mkIso (dualTensorIsoLinHomAux ρV W) ?_ - convert dualTensorHom_comm ρV W.ρ + refine Action.mkIso (dualTensorIsoLinHomAux ρV W) (fun g => ?_) + ext : 1 + exact dualTensorHom_comm ρV W.ρ g @[simp] -theorem dualTensorIsoLinHom_hom_hom : (dualTensorIsoLinHom ρV W).hom.hom = dualTensorHom k V W := +theorem dualTensorIsoLinHom_hom_hom : + (dualTensorIsoLinHom ρV W).hom.hom = ModuleCat.ofHom (dualTensorHom k V W) := rfl end FDRep diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 2b6bd6ab54337..756771dc9da34 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -121,9 +121,9 @@ morphisms in `Rep k G`) commutes with the differentials in the complex of inhomo and the homogeneous `linearYonedaObjResolution`. -/ @[nolint checkType] theorem d_eq : d n A = - (diagonalHomEquiv n A).toModuleIso.inv ≫ + ((diagonalHomEquiv n A).toModuleIso.inv ≫ (linearYonedaObjResolution A).d n (n + 1) ≫ - (diagonalHomEquiv (n + 1) A).toModuleIso.hom := by + (diagonalHomEquiv (n + 1) A).toModuleIso.hom).hom := by ext f g /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was simp only [ModuleCat.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, @@ -146,7 +146,7 @@ and the homogeneous `linearYonedaObjResolution`. -/ -- https://github.com/leanprover-community/mathlib4/issues/5164 change d n A f g = diagonalHomEquiv (n + 1) A ((resolution k G).d (n + 1) n ≫ (diagonalHomEquiv n A).symm f) g - rw [diagonalHomEquiv_apply, Action.comp_hom, ModuleCat.comp_def, LinearMap.comp_apply, + rw [diagonalHomEquiv_apply, Action.comp_hom, ModuleCat.hom_comp, LinearMap.comp_apply, resolution.d_eq] erw [resolution.d_of (Fin.partialProd g)] simp only [map_sum, ← Finsupp.smul_single_one _ ((-1 : k) ^ _)] @@ -175,7 +175,7 @@ $$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \ which calculates the group cohomology of `A`. -/ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := CochainComplex.of (fun n => ModuleCat.of k ((Fin n → G) → A)) - (fun n => inhomogeneousCochains.d n A) fun n => by + (fun n => ModuleCat.ofHom (inhomogeneousCochains.d n A)) fun n => by /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was ext x y have := LinearMap.ext_iff.1 ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2)) @@ -184,20 +184,23 @@ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := LinearEquiv.toModuleIso_inv, LinearEquiv.coe_coe, LinearEquiv.symm_apply_apply, this, LinearMap.zero_apply, map_zero, Pi.zero_apply] -/ ext x - have := LinearMap.ext_iff.1 ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2)) - simp only [ModuleCat.comp_def, LinearMap.comp_apply] at this + have : ∀ x, _ = (0 : _ →ₗ[_] _) x := LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp + ((linearYonedaObjResolution A).d_comp_d n (n + 1) (n + 2))) + simp only [ModuleCat.hom_comp, LinearMap.comp_apply] at this dsimp only - simp only [d_eq, LinearEquiv.toModuleIso_inv, LinearEquiv.toModuleIso_hom, ModuleCat.coe_comp, - Function.comp_apply] + simp only [d_eq, LinearEquiv.toModuleIso_inv_hom, LinearEquiv.toModuleIso_hom_hom, + ModuleCat.hom_comp, LinearMap.comp_apply, LinearEquiv.coe_coe, ModuleCat.hom_zero] /- Porting note: I can see I need to rewrite `LinearEquiv.coe_coe` twice to at least reduce the need for `symm_apply_apply` to be an `erw`. However, even `erw` refuses to rewrite the second `coe_coe`... -/ erw [LinearEquiv.symm_apply_apply, this] - exact map_zero _ + simp only [LinearMap.zero_apply, ChainComplex.linearYonedaObj_X, linearYoneda_obj_obj_carrier, + map_zero, Pi.zero_apply, LinearMap.zero_apply] + rfl @[simp] theorem inhomogeneousCochains.d_def (n : ℕ) : - (inhomogeneousCochains A).d n (n + 1) = inhomogeneousCochains.d n A := + (inhomogeneousCochains A).d n (n + 1) = ModuleCat.ofHom (inhomogeneousCochains.d n A) := CochainComplex.of_d _ _ _ _ /-- Given a `k`-linear `G`-representation `A`, the complex of inhomogeneous cochains is isomorphic @@ -207,8 +210,15 @@ def inhomogeneousCochainsIso : inhomogeneousCochains A ≅ linearYonedaObjResolu (Rep.diagonalHomEquiv i A).toModuleIso.symm) ?_ rintro i j (h : i + 1 = j) subst h - simp only [CochainComplex.of_d, d_eq, Category.assoc, Iso.symm_hom, Iso.hom_inv_id, - Category.comp_id] + ext + simp only [ChainComplex.linearYonedaObj_X, linearYoneda_obj_obj_carrier, CochainComplex.of_x, + linearYoneda_obj_obj_isAddCommGroup, linearYoneda_obj_obj_isModule, Iso.symm_hom, + ChainComplex.linearYonedaObj_d, ModuleCat.hom_comp, linearYoneda_obj_map_hom, + Quiver.Hom.unop_op, LinearEquiv.toModuleIso_inv_hom, LinearMap.coe_comp, Function.comp_apply, + Linear.leftComp_apply, inhomogeneousCochains.d_def, d_eq, LinearEquiv.toModuleIso_hom_hom, + ModuleCat.ofHom_comp, Category.assoc, LinearEquiv.comp_coe, LinearEquiv.self_trans_symm, + LinearEquiv.refl_toLinearMap, LinearMap.id_comp, LinearEquiv.coe_coe] + rfl /-- The `n`-cocycles `Zⁿ(G, A)` of a `k`-linear `G`-representation `A`, i.e. the kernel of the `n`th differential in the complex of inhomogeneous cochains. -/ diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 24d8d8f5bd37c..ff50887857737 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -142,7 +142,7 @@ square commutes: where the vertical arrows are `zeroCochainsLequiv` and `oneCochainsLequiv` respectively. -/ theorem dZero_comp_eq : dZero A ∘ₗ (zeroCochainsLequiv A) = - oneCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 0 1 := by + oneCochainsLequiv A ∘ₗ ((inhomogeneousCochains A).d 0 1).hom := by ext x y show A.ρ y (x default) - x default = _ + ({0} : Finset _).sum _ simp_rw [Fin.val_eq_zero, zero_add, pow_one, neg_smul, one_smul, @@ -163,7 +163,7 @@ square commutes: where the vertical arrows are `oneCochainsLequiv` and `twoCochainsLequiv` respectively. -/ theorem dOne_comp_eq : dOne A ∘ₗ oneCochainsLequiv A = - twoCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 1 2 := by + twoCochainsLequiv A ∘ₗ ((inhomogeneousCochains A).d 1 2).hom := by ext x y show A.ρ y.1 (x _) - x _ + x _ = _ + _ rw [Fin.sum_univ_two] @@ -185,7 +185,8 @@ square commutes: where the vertical arrows are `twoCochainsLequiv` and `threeCochainsLequiv` respectively. -/ theorem dTwo_comp_eq : - dTwo A ∘ₗ twoCochainsLequiv A = threeCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 2 3 := by + dTwo A ∘ₗ twoCochainsLequiv A = + threeCochainsLequiv A ∘ₗ ((inhomogeneousCochains A).d 2 3).hom := by ext x y show A.ρ y.1 (x _) - x _ + x _ - x _ = _ + _ dsimp @@ -201,12 +202,13 @@ theorem dOne_comp_dZero : dOne A ∘ₗ dZero A = 0 := by rfl theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by - show ModuleCat.asHom (dOne A) ≫ ModuleCat.asHom (dTwo A) = _ - have h1 : _ ≫ ModuleCat.asHom (dOne A) = _ ≫ _ := congr_arg ModuleCat.asHom (dOne_comp_eq A) - have h2 : _ ≫ ModuleCat.asHom (dTwo A) = _ ≫ _ := congr_arg ModuleCat.asHom (dTwo_comp_eq A) - simp only [← LinearEquiv.toModuleIso_hom] at h1 h2 - simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, - Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, zero_comp, comp_zero] + show (ModuleCat.ofHom (dOne A) ≫ ModuleCat.ofHom (dTwo A)).hom = _ + have h1 := congr_arg ModuleCat.ofHom (dOne_comp_eq A) + have h2 := congr_arg ModuleCat.ofHom (dTwo_comp_eq A) + simp only [ModuleCat.ofHom_comp, ModuleCat.ofHom_comp, ← LinearEquiv.toModuleIso_hom_hom] at h1 h2 + simp only [(Iso.eq_inv_comp _).2 h2, (Iso.eq_inv_comp _).2 h1, ModuleCat.ofHom_hom, + ModuleCat.hom_ofHom, Category.assoc, Iso.hom_inv_id_assoc, HomologicalComplex.d_comp_d_assoc, + zero_comp, comp_zero, ModuleCat.hom_zero] end Differentials @@ -749,9 +751,9 @@ lemma shortComplexH0_exact : (shortComplexH0 A).Exact := by `(inhomogeneousCochains A).d 0 1` of the complex of inhomogeneous cochains of `A`. -/ @[simps! hom_left hom_right inv_left inv_right] def dZeroArrowIso : Arrow.mk ((inhomogeneousCochains A).d 0 1) ≅ - Arrow.mk (ModuleCat.asHom (dZero A)) := + Arrow.mk (ModuleCat.ofHom (dZero A)) := Arrow.isoMk (zeroCochainsLequiv A).toModuleIso - (oneCochainsLequiv A).toModuleIso (dZero_comp_eq A) + (oneCochainsLequiv A).toModuleIso (ModuleCat.hom_ext (dZero_comp_eq A)) /-- The 0-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `A.ρ.invariants`, which is a simpler type. -/ @@ -761,7 +763,7 @@ def isoZeroCocycles : cocycles A 0 ≅ ModuleCat.of k A.ρ.invariants := (dZeroArrowIso A) lemma isoZeroCocycles_hom_comp_subtype : - (isoZeroCocycles A).hom ≫ A.ρ.invariants.subtype = + (isoZeroCocycles A).hom ≫ ModuleCat.ofHom A.ρ.invariants.subtype = iCocycles A 0 ≫ (zeroCochainsLequiv A).toModuleIso.hom := by dsimp [isoZeroCocycles] apply KernelFork.mapOfIsLimit_ι @@ -788,7 +790,9 @@ short complex associated to the complex of inhomogeneous cochains of `A`. -/ @[simps! hom inv] def shortComplexH1Iso : (inhomogeneousCochains A).sc' 0 1 2 ≅ shortComplexH1 A := isoMk (zeroCochainsLequiv A).toModuleIso (oneCochainsLequiv A).toModuleIso - (twoCochainsLequiv A).toModuleIso (dZero_comp_eq A) (dOne_comp_eq A) + (twoCochainsLequiv A).toModuleIso + (ModuleCat.hom_ext (dZero_comp_eq A)) + (ModuleCat.hom_ext (dOne_comp_eq A)) /-- The 1-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `oneCocycles A`, which is a simpler type. -/ @@ -797,7 +801,7 @@ def isoOneCocycles : cocycles A 1 ≅ ModuleCat.of k (oneCocycles A) := cyclesMapIso (shortComplexH1Iso A) ≪≫ (shortComplexH1 A).moduleCatCyclesIso lemma isoOneCocycles_hom_comp_subtype : - (isoOneCocycles A).hom ≫ ModuleCat.asHom (oneCocycles A).subtype = + (isoOneCocycles A).hom ≫ ModuleCat.ofHom (oneCocycles A).subtype = iCocycles A 1 ≫ (oneCochainsLequiv A).toModuleIso.hom := by dsimp [isoOneCocycles] rw [Category.assoc, Category.assoc] @@ -807,9 +811,8 @@ lemma isoOneCocycles_hom_comp_subtype : lemma toCocycles_comp_isoOneCocycles_hom : toCocycles A 0 1 ≫ (isoOneCocycles A).hom = (zeroCochainsLequiv A).toModuleIso.hom ≫ - ModuleCat.asHom (shortComplexH1 A).moduleCatToCycles := by + ModuleCat.ofHom (shortComplexH1 A).moduleCatToCycles := by simp [isoOneCocycles] - rfl /-- The 1st group cohomology of `A`, defined as the 1st cohomology of the complex of inhomogeneous cochains, is isomorphic to `oneCocycles A ⧸ oneCoboundaries A`, which is a simpler type. -/ @@ -836,7 +839,9 @@ isomorphic to the 2nd short complex associated to the complex of inhomogeneous c def shortComplexH2Iso : (inhomogeneousCochains A).sc' 1 2 3 ≅ shortComplexH2 A := isoMk (oneCochainsLequiv A).toModuleIso (twoCochainsLequiv A).toModuleIso - (threeCochainsLequiv A).toModuleIso (dOne_comp_eq A) (dTwo_comp_eq A) + (threeCochainsLequiv A).toModuleIso + (ModuleCat.hom_ext (dOne_comp_eq A)) + (ModuleCat.hom_ext (dTwo_comp_eq A)) /-- The 2-cocycles of the complex of inhomogeneous cochains of `A` are isomorphic to `twoCocycles A`, which is a simpler type. -/ @@ -845,7 +850,7 @@ def isoTwoCocycles : cocycles A 2 ≅ ModuleCat.of k (twoCocycles A) := cyclesMapIso (shortComplexH2Iso A) ≪≫ (shortComplexH2 A).moduleCatCyclesIso lemma isoTwoCocycles_hom_comp_subtype : - (isoTwoCocycles A).hom ≫ ModuleCat.asHom (twoCocycles A).subtype = + (isoTwoCocycles A).hom ≫ ModuleCat.ofHom (twoCocycles A).subtype = iCocycles A 2 ≫ (twoCochainsLequiv A).toModuleIso.hom := by dsimp [isoTwoCocycles] rw [Category.assoc, Category.assoc] @@ -855,9 +860,8 @@ lemma isoTwoCocycles_hom_comp_subtype : lemma toCocycles_comp_isoTwoCocycles_hom : toCocycles A 1 2 ≫ (isoTwoCocycles A).hom = (oneCochainsLequiv A).toModuleIso.hom ≫ - ModuleCat.asHom (shortComplexH2 A).moduleCatToCycles := by + ModuleCat.ofHom (shortComplexH2 A).moduleCatToCycles := by simp [isoTwoCocycles] - rfl /-- The 2nd group cohomology of `A`, defined as the 2nd cohomology of the complex of inhomogeneous cochains, is isomorphic to `twoCocycles A ⧸ twoCoboundaries A`, which is a simpler type. -/ diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index 564819a98563f..fedc31c3c9a40 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -320,11 +320,8 @@ theorem diagonalHomEquiv_symm_apply (f : (Fin n → G) → A) (x : Fin (n + 1) one_smul, Rep.of_ρ, Rep.Action_ρ_eq_ρ, Rep.trivial_def (x 0)⁻¹, Finsupp.llift_apply A k k] -/ simp only [LinearEquiv.trans_symm, LinearEquiv.symm_symm, LinearEquiv.trans_apply, leftRegularHomEquiv_symm_apply, Linear.homCongr_symm_apply, Iso.trans_hom, Iso.refl_inv, - Category.comp_id, Action.comp_hom, MonoidalClosed.linearHomEquivComm_symm_hom] - -- Porting note: This is a sure sign that coercions for morphisms in `ModuleCat` - -- are still not set up properly. - rw [ModuleCat.coe_comp] - simp only [ModuleCat.coe_comp, Function.comp_apply] + 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] -- The prototype linter that checks if `erw` could be replaced with `rw` would time out @@ -482,7 +479,9 @@ instance x_projective (G : Type u) [Group G] (n : ℕ) : /-- Simpler expression for the differential in the standard resolution of `k` as a `G`-representation. It sends `(g₀, ..., gₙ₊₁) ↦ ∑ (-1)ⁱ • (g₀, ..., ĝᵢ, ..., gₙ₊₁)`. -/ -theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = d k G (n + 1) := by +theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = + ModuleCat.ofHom (d k G (n + 1)) := by + ext : 1 refine Finsupp.lhom_ext' fun x => LinearMap.ext_ring ?_ dsimp [groupCohomology.resolution] /- Porting note (https://github.com/leanprover-community/mathlib4/issues/11039): broken proof was @@ -490,11 +489,12 @@ theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = d simp_rw [alternatingFaceMapComplex_obj_d, AlternatingFaceMapComplex.objD, SimplicialObject.δ, Functor.comp_map, ← Int.cast_smul_eq_zsmul k ((-1) ^ _ : ℤ), Int.cast_pow, Int.cast_neg, Int.cast_one, Action.sum_hom, Action.smul_hom, Rep.linearization_map_hom] - rw [LinearMap.coeFn_sum, Fintype.sum_apply] + rw [ModuleCat.hom_sum, LinearMap.coeFn_sum, Fintype.sum_apply] erw [d_of (k := k) x] /- Porting note: want to rewrite `LinearMap.smul_apply` but simp/simp_rw won't do it; I need erw, so using Finset.sum_congr to get rid of the binder -/ refine Finset.sum_congr rfl fun _ _ => ?_ + simp only [ModuleCat.hom_smul, SimplexCategory.len_mk] erw [LinearMap.smul_apply] rw [Finsupp.lmapDomain_apply, Finsupp.mapDomain_single, Finsupp.smul_single', mul_one] rfl @@ -533,8 +533,8 @@ def forget₂ToModuleCatHomotopyEquiv : /-- The hom of `k`-linear `G`-representations `k[G¹] → k` sending `∑ nᵢgᵢ ↦ ∑ nᵢ`. -/ def ε : Rep.ofMulAction k G (Fin 1 → G) ⟶ Rep.trivial k G k where - hom := Finsupp.linearCombination _ fun _ => (1 : k) - comm g := Finsupp.lhom_ext' fun _ => LinearMap.ext_ring (by + hom := ModuleCat.ofHom <| Finsupp.linearCombination _ fun _ => (1 : k) + comm g := ModuleCat.hom_ext <| Finsupp.lhom_ext' fun _ => LinearMap.ext_ring (by show Finsupp.linearCombination k (fun _ => (1 : k)) (Finsupp.mapDomain _ (Finsupp.single _ _)) = Finsupp.linearCombination k (fun _ => (1 : k)) (Finsupp.single _ _) @@ -551,7 +551,8 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : convert Category.id_comp (X := (forget₂ToModuleCat k G).X 0) _ · dsimp only [HomotopyEquiv.ofIso, compForgetAugmentedIso] simp only [Iso.symm_hom, eqToIso.inv, HomologicalComplex.eqToHom_f, eqToHom_refl] - trans (linearCombination _ fun _ => (1 : k)).comp ((ModuleCat.free k).map (terminal.from _)) + ext : 1 + trans (linearCombination _ fun _ => (1 : k)).comp ((ModuleCat.free k).map (terminal.from _)).hom · erw [Finsupp.lmapDomain_linearCombination (α := Fin 1 → G) (R := k) (α' := ⊤_ Type u) (v := fun _ => (1 : k)) (v' := fun _ => (1 : k)) (terminal.from @@ -563,21 +564,19 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : · ext x dsimp (config := { unfoldPartialApp := true }) [HomotopyEquiv.ofIso, Finsupp.LinearEquiv.finsuppUnique] - rw [linearCombination_single, one_smul, - @Unique.eq_default _ Types.terminalIso.toEquiv.unique x, - ChainComplex.single₀_map_f_zero, LinearMap.coe_mk, AddHom.coe_mk, Function.comp_apply, - Finsupp.equivFunOnFinite_apply, Finsupp.single_eq_same] + rw [@Unique.eq_default _ Types.terminalIso.toEquiv.unique x] + simp · exact @Subsingleton.elim _ (@Unique.instSubsingleton _ (Limits.uniqueToTerminal _)) _ _ theorem d_comp_ε : (groupCohomology.resolution k G).d 1 0 ≫ ε k G = 0 := by ext : 1 - refine LinearMap.ext fun x => ?_ + refine ModuleCat.hom_ext <| LinearMap.ext fun x => ?_ have : (forget₂ToModuleCat k G).d 1 0 ≫ (forget₂ (Rep k G) (ModuleCat.{u} k)).map (ε k G) = 0 := by rw [← forget₂ToModuleCatHomotopyEquiv_f_0_eq, ← (forget₂ToModuleCatHomotopyEquiv k G).1.2 1 0 rfl] exact comp_zero - exact LinearMap.ext_iff.1 this _ + exact LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp this) _ /-- The chain map from the standard resolution of `k` to `k[0]` given by `∑ nᵢgᵢ ↦ ∑ nᵢ` in degree zero. -/ diff --git a/Mathlib/RepresentationTheory/Invariants.lean b/Mathlib/RepresentationTheory/Invariants.lean index 898d8d3427513..acf708f7c06b9 100644 --- a/Mathlib/RepresentationTheory/Invariants.lean +++ b/Mathlib/RepresentationTheory/Invariants.lean @@ -121,19 +121,23 @@ variable {k : Type u} [CommRing k] {G : Grp.{u}} theorem mem_invariants_iff_comm {X Y : Rep k G} (f : X.V →ₗ[k] Y.V) (g : G) : (linHom X.ρ Y.ρ) g f = f ↔ f.comp (X.ρ g) = (Y.ρ g).comp f := by dsimp - erw [← ρAut_apply_inv] - rw [← LinearMap.comp_assoc, ← ModuleCat.comp_def, ← ModuleCat.comp_def, Iso.inv_comp_eq, - ρAut_apply_hom] + rw [← LinearMap.comp_assoc, ← ModuleCat.hom_ofHom (Y.ρ g), ← ModuleCat.hom_ofHom f, + ← ModuleCat.hom_comp, ← ModuleCat.hom_ofHom (X.ρ g⁻¹), ← ModuleCat.hom_comp, + Rep.ofHom_ρ, ← ρAut_apply_inv X g, Rep.ofHom_ρ, ← ρAut_apply_hom Y g, ← ModuleCat.hom_ext_iff, + Iso.inv_comp_eq, ρAut_apply_hom, ← ModuleCat.hom_ofHom (X.ρ g), + ← ModuleCat.hom_comp, ← ModuleCat.hom_ext_iff] exact comm /-- The invariants of the representation `linHom X.ρ Y.ρ` correspond to the representation homomorphisms from `X` to `Y`. -/ @[simps] def invariantsEquivRepHom (X Y : Rep k G) : (linHom X.ρ Y.ρ).invariants ≃ₗ[k] X ⟶ Y where - toFun f := ⟨f.val, fun g => (mem_invariants_iff_comm _ g).1 (f.property g)⟩ + toFun f := ⟨ModuleCat.ofHom f.val, fun g => + ModuleCat.hom_ext ((mem_invariants_iff_comm _ g).1 (f.property g))⟩ map_add' _ _ := rfl map_smul' _ _ := rfl - invFun f := ⟨f.hom, fun g => (mem_invariants_iff_comm _ g).2 (f.comm g)⟩ + invFun f := ⟨f.hom.hom, fun g => + (mem_invariants_iff_comm _ g).2 (ModuleCat.hom_ext_iff.mp (f.comm g))⟩ left_inv _ := by ext; rfl right_inv _ := by ext; rfl diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 0a14328b88a96..342bac0e56cf4 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -61,11 +61,11 @@ instance (V : Rep k G) : Module k V := by -/ def ρ (V : Rep k G) : Representation k G V := -- Porting note: was `V.ρ` - Action.ρ V + (ModuleCat.endMulEquiv V.V).toMonoidHom.comp (Action.ρ V) /-- Lift an unbundled representation to `Rep`. -/ def of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : Rep k G := - ⟨ModuleCat.of k V, ρ⟩ + ⟨ModuleCat.of k V, MonCat.ofHom ((ModuleCat.endMulEquiv _).symm.toMonoidHom.comp ρ) ⟩ @[simp] theorem coe_of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : @@ -76,9 +76,16 @@ theorem coe_of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[ theorem of_ρ {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : (of ρ).ρ = ρ := rfl -theorem Action_ρ_eq_ρ {A : Rep k G} : Action.ρ A = A.ρ := +theorem Action_ρ_eq_ρ {A : Rep k G} : + Action.ρ A = (ModuleCat.endMulEquiv _).symm.toMonoidHom.comp A.ρ := rfl +@[simp] +lemma ρ_hom {X : Rep k G} (g : G) : (Action.ρ X g).hom = X.ρ g := rfl + +@[simp] +lemma ofHom_ρ {X : Rep k G} (g : G) : ModuleCat.ofHom (X.ρ g) = Action.ρ X g := rfl + /-- Allows us to apply lemmas about the underlying `ρ`, which would take an element `g : G` rather than `g : MonCat.of G` as an argument. -/ theorem of_ρ_apply {V : Type u} [AddCommGroup V] [Module k V] (ρ : Representation k G V) @@ -97,7 +104,7 @@ theorem ρ_self_inv_apply {G : Type u} [Group G] {A : Rep k G} (g : G) (x : A) : theorem hom_comm_apply {A B : Rep k G} (f : A ⟶ B) (g : G) (x : A) : f.hom (A.ρ g x) = B.ρ g (f.hom x) := - LinearMap.ext_iff.1 (f.comm g) x + LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (f.comm g)) x variable (k G) @@ -172,7 +179,8 @@ theorem linearization_single (X : Action (Type u) (MonCat.of G)) (g : G) (x : X. variable {X Y : Action (Type u) (MonCat.of G)} (f : X ⟶ Y) @[simp] -theorem linearization_map_hom : ((linearization k G).map f).hom = Finsupp.lmapDomain k k f.hom := +theorem linearization_map_hom : ((linearization k G).map f).hom = + ModuleCat.ofHom (Finsupp.lmapDomain k k f.hom) := rfl theorem linearization_map_hom_single (x : X.V) (r : k) : @@ -183,16 +191,19 @@ open Functor.LaxMonoidal Functor.OplaxMonoidal Functor.Monoidal @[simp] theorem linearization_μ_hom (X Y : Action (Type u) (MonCat.of G)) : - (μ (linearization k G) X Y).hom = (finsuppTensorFinsupp' k X.V Y.V).toLinearMap := + (μ (linearization k G) X Y).hom = + ModuleCat.ofHom (finsuppTensorFinsupp' k X.V Y.V).toLinearMap := rfl @[simp] theorem linearization_δ_hom (X Y : Action (Type u) (MonCat.of G)) : - (δ (linearization k G) X Y).hom = (finsuppTensorFinsupp' k X.V Y.V).symm.toLinearMap := + (δ (linearization k G) X Y).hom = + ModuleCat.ofHom (finsuppTensorFinsupp' k X.V Y.V).symm.toLinearMap := rfl @[simp] -theorem linearization_ε_hom : (ε (linearization k G)).hom = Finsupp.lsingle PUnit.unit := +theorem linearization_ε_hom : (ε (linearization k G)).hom = + ModuleCat.ofHom (Finsupp.lsingle PUnit.unit) := rfl theorem linearization_η_hom_apply (r : k) : @@ -206,7 +217,7 @@ on `k[X]`. -/ @[simps!] noncomputable def linearizationTrivialIso (X : Type u) : (linearization k G).obj (Action.mk X 1) ≅ trivial k G (X →₀ k) := - Action.mkIso (Iso.refl _) fun _ => Finsupp.lhom_ext' fun _ => LinearMap.ext + Action.mkIso (Iso.refl _) fun _ => ModuleCat.hom_ext <| Finsupp.lhom_ext' fun _ => LinearMap.ext fun _ => linearization_single .. /-- Given a `G`-action on `H`, this is `k[H]` bundled with the natural representation @@ -269,14 +280,15 @@ variable {k G} `g ↦ A.ρ(g)(x).` -/ @[simps] noncomputable def leftRegularHom (A : Rep k G) (x : A) : Rep.ofMulAction k G G ⟶ A where - hom := Finsupp.lift _ _ _ fun g => A.ρ g x + hom := ModuleCat.ofHom <| Finsupp.lift _ _ _ fun g => A.ρ g x comm g := by + ext : 1 refine Finsupp.lhom_ext' fun y => LinearMap.ext_ring ?_ /- Porting note: rest of broken proof was simpa only [LinearMap.comp_apply, ModuleCat.comp_def, Finsupp.lsingle_apply, Finsupp.lift_apply, Action_ρ_eq_ρ, of_ρ_apply, Representation.ofMulAction_single, Finsupp.sum_single_index, zero_smul, one_smul, smul_eq_mul, A.ρ.map_mul] -/ - simp only [LinearMap.comp_apply, ModuleCat.comp_def, Finsupp.lsingle_apply] + simp only [LinearMap.comp_apply, ModuleCat.hom_comp, Finsupp.lsingle_apply] erw [Finsupp.lift_apply, Finsupp.lift_apply, Representation.ofMulAction_single (G := G)] simp only [Finsupp.sum_single_index, zero_smul, one_smul, smul_eq_mul, A.ρ.map_mul, of_ρ] rfl @@ -284,7 +296,7 @@ noncomputable def leftRegularHom (A : Rep k G) (x : A) : Rep.ofMulAction k G G theorem leftRegularHom_apply {A : Rep k G} (x : A) : (leftRegularHom A x).hom (Finsupp.single 1 1) = x := by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [leftRegularHom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, + erw [leftRegularHom_hom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, A.ρ.map_one, LinearMap.one_apply] rw [zero_smul] @@ -297,14 +309,14 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ map_smul' _ _ := rfl invFun x := leftRegularHom A x left_inv f := by - refine Action.Hom.ext (Finsupp.lhom_ext' fun x : G => LinearMap.ext_ring ?_) + refine Action.Hom.ext (ModuleCat.hom_ext (Finsupp.lhom_ext' fun x : G => LinearMap.ext_ring ?_)) have : f.hom ((ofMulAction k G G).ρ x (Finsupp.single (1 : G) (1 : k))) = A.ρ x (f.hom (Finsupp.single (1 : G) (1 : k))) := - LinearMap.ext_iff.1 (f.comm x) (Finsupp.single 1 1) - simp only [leftRegularHom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, Finsupp.lift_apply, - ← this, coe_of, of_ρ, Representation.ofMulAction_single x (1 : G) (1 : k), smul_eq_mul, - mul_one, zero_smul, Finsupp.sum_single_index, one_smul] + LinearMap.ext_iff.1 (ModuleCat.hom_ext_iff.mp (f.comm x)) (Finsupp.single 1 1) + simp only [leftRegularHom_hom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, + Finsupp.lift_apply, ← this, coe_of, of_ρ, Representation.ofMulAction_single x (1 : G) (1 : k), + smul_eq_mul, mul_one, zero_smul, Finsupp.sum_single_index, one_smul] -- Mismatched `Zero k` instances rfl right_inv x := leftRegularHom_apply x @@ -312,7 +324,7 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ theorem leftRegularHomEquiv_symm_single {A : Rep k G} (x : A) (g : G) : ((leftRegularHomEquiv A).symm x).hom (Finsupp.single g 1) = A.ρ g x := by -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom, Finsupp.lift_apply, + erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul] rw [zero_smul] @@ -332,8 +344,8 @@ variable [Group G] (A B C : Rep k G) protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G where obj B := Rep.of (Representation.linHom A.ρ B.ρ) map := fun {X} {Y} f => - { hom := ModuleCat.asHom (LinearMap.llcomp k _ _ _ f.hom) - comm := fun g => LinearMap.ext fun x => LinearMap.ext fun y => by + { hom := ModuleCat.ofHom (LinearMap.llcomp k _ _ _ f.hom.hom) + comm := fun g => ModuleCat.hom_ext <| LinearMap.ext fun x => LinearMap.ext fun y => by show f.hom (X.ρ g _) = _ simp only [hom_comm_apply]; rfl } map_id := fun _ => by ext; rfl @@ -348,32 +360,33 @@ protected def ihom (A : Rep k G) : Rep k G ⥤ Rep k G where `k`-linear map underlying `f`, giving a map `A →ₗ[k] B →ₗ[k] C`, then flipping the arguments. -/ def homEquiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) where toFun f := - { hom := (TensorProduct.curry f.hom).flip + { hom := ModuleCat.ofHom <| (TensorProduct.curry f.hom.hom).flip comm := fun g => by - refine LinearMap.ext fun x => LinearMap.ext fun y => ?_ + ext x : 2 + refine LinearMap.ext fun y => ?_ change f.hom (_ ⊗ₜ[k] _) = C.ρ g (f.hom (_ ⊗ₜ[k] _)) rw [← hom_comm_apply] change _ = f.hom ((A.ρ g * A.ρ g⁻¹) y ⊗ₜ[k] _) simp only [← map_mul, mul_inv_cancel, map_one] rfl } invFun f := - { hom := TensorProduct.uncurry k _ _ _ f.hom.flip - comm := fun g => TensorProduct.ext' fun x y => by -/- Porting note: rest of broken proof was + { hom := ModuleCat.ofHom <| TensorProduct.uncurry k _ _ _ f.hom.hom.flip + comm := fun g => ModuleCat.hom_ext <| TensorProduct.ext' fun x y => by + /- Porting note: rest of broken proof was dsimp only [MonoidalCategory.tensorLeft_obj, ModuleCat.comp_def, LinearMap.comp_apply, tensor_ρ, ModuleCat.MonoidalCategory.hom_apply, TensorProduct.map_tmul] simp only [TensorProduct.uncurry_apply f.hom.flip, LinearMap.flip_apply, Action_ρ_eq_ρ, hom_comm_apply f g y, Rep.ihom_obj_ρ_apply, LinearMap.comp_apply, ρ_inv_self_apply] -/ - change TensorProduct.uncurry k _ _ _ f.hom.flip (A.ρ g x ⊗ₜ[k] B.ρ g y) = - C.ρ g (TensorProduct.uncurry k _ _ _ f.hom.flip (x ⊗ₜ[k] y)) + change TensorProduct.uncurry k _ _ _ f.hom.hom.flip (A.ρ g x ⊗ₜ[k] B.ρ g y) = + C.ρ g (TensorProduct.uncurry k _ _ _ f.hom.hom.flip (x ⊗ₜ[k] y)) -- The next 3 tactics used to be `rw` before https://github.com/leanprover/lean4/pull/2644 erw [TensorProduct.uncurry_apply, LinearMap.flip_apply, hom_comm_apply, Rep.ihom_obj_ρ_apply, LinearMap.comp_apply, LinearMap.comp_apply] --, ρ_inv_self_apply (A := C)] dsimp rw [ρ_inv_self_apply] - rfl} - left_inv _ := Action.Hom.ext (TensorProduct.ext' fun _ _ => rfl) + rfl } + left_inv _ := Action.Hom.ext (ModuleCat.hom_ext (TensorProduct.ext' fun _ _ => rfl)) right_inv f := by ext; rfl variable {A B C} @@ -381,12 +394,13 @@ variable {A B C} /-- Porting note: if we generate this with `@[simps]` the linter complains some types in the LHS simplify. -/ theorem homEquiv_apply_hom (f : A ⊗ B ⟶ C) : - (homEquiv A B C f).hom = (TensorProduct.curry f.hom).flip := rfl + (homEquiv A B C f).hom = ModuleCat.ofHom (TensorProduct.curry f.hom.hom).flip := rfl /-- Porting note: if we generate this with `@[simps]` the linter complains some types in the LHS simplify. -/ theorem homEquiv_symm_apply_hom (f : B ⟶ (Rep.ihom A).obj C) : - ((homEquiv A B C).symm f).hom = TensorProduct.uncurry k A B C f.hom.flip := rfl + ((homEquiv A B C).symm f).hom = + ModuleCat.ofHom (TensorProduct.uncurry k A B C f.hom.hom.flip) := rfl instance : MonoidalClosed (Rep k G) where closed A := @@ -394,9 +408,9 @@ instance : MonoidalClosed (Rep k G) where adj := Adjunction.mkOfHomEquiv ( { homEquiv := Rep.homEquiv A homEquiv_naturality_left_symm := fun _ _ => Action.Hom.ext - (TensorProduct.ext' fun _ _ => rfl) - homEquiv_naturality_right := fun _ _ => Action.Hom.ext (LinearMap.ext - fun _ => LinearMap.ext fun _ => rfl) })} + (ModuleCat.hom_ext (TensorProduct.ext' fun _ _ => rfl)) + homEquiv_naturality_right := fun _ _ => Action.Hom.ext (ModuleCat.hom_ext (LinearMap.ext + fun _ => LinearMap.ext fun _ => rfl)) })} @[simp] theorem ihom_obj_ρ_def (A B : Rep k G) : ((ihom A).obj B).ρ = ((Rep.ihom A).obj B).ρ := @@ -408,13 +422,13 @@ theorem homEquiv_def (A B C : Rep k G) : (ihom.adjunction A).homEquiv B C = Rep. @[simp] theorem ihom_ev_app_hom (A B : Rep k G) : - Action.Hom.hom ((ihom.ev A).app B) - = TensorProduct.uncurry k A (A →ₗ[k] B) B LinearMap.id.flip := by + Action.Hom.hom ((ihom.ev A).app B) = ModuleCat.ofHom + (TensorProduct.uncurry k A (A →ₗ[k] B) B LinearMap.id.flip) := by ext; rfl @[simp] theorem ihom_coev_app_hom (A B : Rep k G) : - Action.Hom.hom ((ihom.coev A).app B) = (TensorProduct.mk k _ _).flip := - LinearMap.ext fun _ => LinearMap.ext fun _ => rfl + Action.Hom.hom ((ihom.coev A).app B) = ModuleCat.ofHom (TensorProduct.mk k _ _).flip := + ModuleCat.hom_ext <| LinearMap.ext fun _ => LinearMap.ext fun _ => rfl variable (A B C) @@ -435,25 +449,27 @@ variable {A B C} -- `simpNF` times out @[simp, nolint simpNF] theorem MonoidalClosed.linearHomEquiv_hom (f : A ⊗ B ⟶ C) : - (MonoidalClosed.linearHomEquiv A B C f).hom = (TensorProduct.curry f.hom).flip := + (MonoidalClosed.linearHomEquiv A B C f).hom = + ModuleCat.ofHom (TensorProduct.curry f.hom.hom).flip := rfl -- `simpNF` times out @[simp, nolint simpNF] theorem MonoidalClosed.linearHomEquivComm_hom (f : A ⊗ B ⟶ C) : - (MonoidalClosed.linearHomEquivComm A B C f).hom = TensorProduct.curry f.hom := + (MonoidalClosed.linearHomEquivComm A B C f).hom = + ModuleCat.ofHom (TensorProduct.curry f.hom.hom) := rfl theorem MonoidalClosed.linearHomEquiv_symm_hom (f : B ⟶ A ⟶[Rep k G] C) : ((MonoidalClosed.linearHomEquiv A B C).symm f).hom = - TensorProduct.uncurry k A B C f.hom.flip := by + ModuleCat.ofHom (TensorProduct.uncurry k A B C f.hom.hom.flip) := by simp [linearHomEquiv] rfl theorem MonoidalClosed.linearHomEquivComm_symm_hom (f : A ⟶ B ⟶[Rep k G] C) : - ((MonoidalClosed.linearHomEquivComm A B C).symm f).hom - = TensorProduct.uncurry k A B C f.hom := - TensorProduct.ext' fun _ _ => rfl + ((MonoidalClosed.linearHomEquivComm A B C).symm f).hom = + ModuleCat.ofHom (TensorProduct.uncurry k A B C f.hom.hom) := + ModuleCat.hom_ext <| TensorProduct.ext' fun _ _ => rfl end MonoidalClosed @@ -512,8 +528,10 @@ theorem to_Module_monoidAlgebra_map_aux {k G : Type*} [CommRing k] [Monoid G] (V /-- Auxiliary definition for `toModuleMonoidAlgebra`. -/ def toModuleMonoidAlgebraMap {V W : Rep k G} (f : V ⟶ W) : ModuleCat.of (MonoidAlgebra k G) V.ρ.asModule ⟶ ModuleCat.of (MonoidAlgebra k G) W.ρ.asModule := - { f.hom with - map_smul' := fun r x => to_Module_monoidAlgebra_map_aux V.V W.V V.ρ W.ρ f.hom f.comm r x } + ModuleCat.ofHom + { f.hom.hom with + map_smul' := fun r x => to_Module_monoidAlgebra_map_aux V.V W.V V.ρ W.ρ f.hom.hom + (fun g => ModuleCat.hom_ext_iff.mp (f.comm g)) r x } /-- Functorially convert a representation of `G` into a module over `MonoidAlgebra k G`. -/ def toModuleMonoidAlgebra : Rep k G ⥤ ModuleCat.{u} (MonoidAlgebra k G) where @@ -524,8 +542,10 @@ def toModuleMonoidAlgebra : Rep k G ⥤ ModuleCat.{u} (MonoidAlgebra k G) where def ofModuleMonoidAlgebra : ModuleCat.{u} (MonoidAlgebra k G) ⥤ Rep k G where obj M := Rep.of (Representation.ofModule M) map f := - { hom := { f with map_smul' := fun r x => f.map_smul (algebraMap k _ r) x } - comm := fun g => by ext; apply f.map_smul } + { hom := ModuleCat.ofHom + { f.hom with + map_smul' := fun r x => f.hom.map_smul (algebraMap k _ r) x } + comm := fun g => by ext; apply f.hom.map_smul } theorem ofModuleMonoidAlgebra_obj_coe (M : ModuleCat.{u} (MonoidAlgebra k G)) : (ofModuleMonoidAlgebra.obj M : Type u) = RestrictScalars k (MonoidAlgebra k G) M := @@ -551,7 +571,7 @@ def unitIsoAddEquiv {V : Rep k G} : V ≃+ (toModuleMonoidAlgebra ⋙ ofModuleMo /-- Auxiliary definition for `equivalenceModuleMonoidAlgebra`. -/ def counitIso (M : ModuleCat.{u} (MonoidAlgebra k G)) : (ofModuleMonoidAlgebra ⋙ toModuleMonoidAlgebra).obj M ≅ M := - LinearEquiv.toModuleIso' + LinearEquiv.toModuleIso { counitIsoAddEquiv with map_smul' := fun r x => by set_option tactic.skipAssignedInstances false in @@ -566,16 +586,13 @@ theorem unit_iso_comm (V : Rep k G) (g : G) (x : V) : unitIsoAddEquiv ((V.ρ g).toFun x) = ((ofModuleMonoidAlgebra.obj (toModuleMonoidAlgebra.obj V)).ρ g).toFun (unitIsoAddEquiv x) := by dsimp [unitIsoAddEquiv, ofModuleMonoidAlgebra, toModuleMonoidAlgebra] -/- Porting note: rest of broken proof was simp only [AddEquiv.apply_eq_iff_eq, AddEquiv.apply_symm_apply, - Representation.asModuleEquiv_symm_map_rho, Representation.ofModule_asModule_act] -/ - rw [Representation.asModuleEquiv_symm_map_rho] - rfl + Representation.asModuleEquiv_symm_map_rho, Representation.ofModule_asModule_act] /-- Auxiliary definition for `equivalenceModuleMonoidAlgebra`. -/ def unitIso (V : Rep k G) : V ≅ (toModuleMonoidAlgebra ⋙ ofModuleMonoidAlgebra).obj V := Action.mkIso - (LinearEquiv.toModuleIso' + (LinearEquiv.toModuleIso { unitIsoAddEquiv with map_smul' := fun r x => by dsimp [unitIsoAddEquiv] diff --git a/Mathlib/RingTheory/AdicCompletion/Algebra.lean b/Mathlib/RingTheory/AdicCompletion/Algebra.lean index 6f7c6168e2c6c..9c8babe6957b6 100644 --- a/Mathlib/RingTheory/AdicCompletion/Algebra.lean +++ b/Mathlib/RingTheory/AdicCompletion/Algebra.lean @@ -25,7 +25,7 @@ suppress_compilation open Submodule -variable {R : Type*} [CommRing R] (I : Ideal R) +variable {R S : Type*} [CommRing R] [CommRing S] (I : Ideal R) variable {M : Type*} [AddCommGroup M] [Module R M] namespace AdicCompletion @@ -87,8 +87,10 @@ instance : CommRing (AdicCompletion I R) := (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) -instance : Algebra R (AdicCompletion I R) where - toFun r := ⟨algebraMap R (∀ n, R ⧸ (I ^ n • ⊤ : Ideal R)) r, by simp⟩ +instance [Algebra S R] : Algebra S (AdicCompletion I R) where + toFun r := ⟨algebraMap S (∀ n, R ⧸ (I ^ n • ⊤ : Ideal R)) r, by + simp [-Ideal.Quotient.mk_algebraMap, + IsScalarTower.algebraMap_apply S R (R ⧸ (I ^ _ • ⊤ : Ideal R))]⟩ map_one' := Subtype.ext <| map_one _ map_mul' x y := Subtype.ext <| map_mul _ x y map_zero' := Subtype.ext <| map_zero _ @@ -252,23 +254,15 @@ instance module : Module (AdicCompletion I R) (AdicCompletion I M) where mul_smul r s x := by ext n simp only [smul_eval, val_mul, mul_smul] - smul_zero r := by - ext n - rw [smul_eval, val_zero, smul_zero] - smul_add r x y := by - ext n - simp only [smul_eval, val_add, smul_add] - add_smul r s x := by - ext n - simp only [coe_eval, smul_eval, map_add, add_smul, val_add] - zero_smul x := by - ext n - simp only [smul_eval, _root_.map_zero, zero_smul, val_zero] + smul_zero r := by ext n; simp + smul_add r x y := by ext n; simp + add_smul r s x := by ext n; simp [val_smul, add_smul] + zero_smul x := by ext n; simp instance : IsScalarTower R (AdicCompletion I R) (AdicCompletion I M) where smul_assoc r s x := by ext n - rw [smul_eval, val_smul, val_smul, smul_eval, smul_assoc] + rw [smul_eval, val_smul_apply, val_smul_apply, smul_eval, smul_assoc] /-- A priori `AdicCompletion I R` has two `AdicCompletion I R`-module instances. Both agree definitionally. -/ diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index 1c4dfed5cf15a..7327b43e29424 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -8,7 +8,7 @@ import Mathlib.CategoryTheory.Abelian.DiagramLemmas.Four import Mathlib.LinearAlgebra.TensorProduct.Pi import Mathlib.LinearAlgebra.TensorProduct.RightExactness import Mathlib.RingTheory.AdicCompletion.Exactness -import Mathlib.RingTheory.Flat.Algebra +import Mathlib.RingTheory.Flat.Basic /-! @@ -101,7 +101,7 @@ private lemma piEquivOfFintype_comp_ofTensorProduct_eq : ext i j k suffices h : (if j = i then 1 else 0) = (if j = i then 1 else 0 : AdicCompletion I R).val k by simpa [Pi.single_apply, -smul_eq_mul, -Algebra.id.smul_eq_mul] - split <;> simp only [smul_eq_mul, val_zero, val_one] + split <;> simp private lemma ofTensorProduct_eq : ofTensorProduct I (ι → R) = (piEquivOfFintype I (ι := ι) (fun _ : ι ↦ R)).symm.toLinearMap ∘ₗ @@ -241,25 +241,25 @@ private instance : AddCommGroup (AdicCompletion I R ⊗[R] (LinearMap.ker f)) := private def firstRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := ComposableArrows.mk₄ - (ModuleCat.asHom <| lTensorKerIncl I M f) - (ModuleCat.asHom <| lTensorf I M f) - (ModuleCat.asHom (0 : AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] PUnit)) - (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom <| lTensorKerIncl I M f) + (ModuleCat.ofHom <| lTensorf I M f) + (ModuleCat.ofHom (0 : AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) private def secondRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := ComposableArrows.mk₄ - (ModuleCat.asHom (map I <| (LinearMap.ker f).subtype)) - (ModuleCat.asHom (map I f)) - (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) - (ModuleCat.asHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom (map I <| (LinearMap.ker f).subtype)) + (ModuleCat.ofHom (map I f)) + (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) + (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) include hf private lemma firstRow_exact : (firstRow I M f).Exact where zero k _ := match k with - | 0 => (tens_exact I M f hf).linearMap_comp_eq_zero - | 1 => LinearMap.zero_comp _ - | 2 => LinearMap.zero_comp 0 + | 0 => ModuleCat.hom_ext (tens_exact I M f hf).linearMap_comp_eq_zero + | 1 => ModuleCat.hom_ext (LinearMap.zero_comp _) + | 2 => ModuleCat.hom_ext (LinearMap.zero_comp 0) exact k _ := by rw [ShortComplex.moduleCat_exact_iff] match k with @@ -269,9 +269,9 @@ private lemma firstRow_exact : (firstRow I M f).Exact where private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M f).Exact where zero k _ := match k with - | 0 => (adic_exact I M f hf).linearMap_comp_eq_zero - | 1 => LinearMap.zero_comp (map I f) - | 2 => LinearMap.zero_comp 0 + | 0 => ModuleCat.hom_ext (adic_exact I M f hf).linearMap_comp_eq_zero + | 1 => ModuleCat.hom_ext (LinearMap.zero_comp (map I f)) + | 2 => ModuleCat.hom_ext (LinearMap.zero_comp 0) exact k _ := by rw [ShortComplex.moduleCat_exact_iff] match k with @@ -282,25 +282,25 @@ private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M /- The compatible vertical maps between the first and the second row. -/ private def firstRowToSecondRow : firstRow I M f ⟶ secondRow I M f := ComposableArrows.homMk₄ - (ModuleCat.asHom (ofTensorProduct I (LinearMap.ker f))) - (ModuleCat.asHom (ofTensorProduct I (ι → R))) - (ModuleCat.asHom (ofTensorProduct I M)) - (ModuleCat.asHom 0) - (ModuleCat.asHom 0) - (ofTensorProduct_naturality I <| (LinearMap.ker f).subtype).symm - (ofTensorProduct_naturality I f).symm + (ModuleCat.ofHom (ofTensorProduct I (LinearMap.ker f))) + (ModuleCat.ofHom (ofTensorProduct I (ι → R))) + (ModuleCat.ofHom (ofTensorProduct I M)) + (ModuleCat.ofHom 0) + (ModuleCat.ofHom 0) + (ModuleCat.hom_ext (ofTensorProduct_naturality I <| (LinearMap.ker f).subtype).symm) + (ModuleCat.hom_ext (ofTensorProduct_naturality I f).symm) rfl rfl private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : - IsIso (ModuleCat.asHom (ofTensorProduct I M)) := by + IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := by refine Abelian.isIso_of_epi_of_isIso_of_isIso_of_mono (firstRow_exact I M f hf) (secondRow_exact I M f hf) (firstRowToSecondRow I M f) ?_ ?_ ?_ ?_ · apply ConcreteCategory.epi_of_surjective exact ofTensorProduct_surjective_of_finite I (LinearMap.ker f) · apply (ConcreteCategory.isIso_iff_bijective _).mpr exact ofTensorProduct_bijective_of_pi_of_fintype I ι - · show IsIso (ModuleCat.asHom 0) + · show IsIso (ModuleCat.ofHom 0) apply Limits.isIso_of_isTerminal <;> exact Limits.IsZero.isTerminal (ModuleCat.isZero_of_subsingleton _) · apply ConcreteCategory.mono_of_injective @@ -310,9 +310,9 @@ private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : private lemma ofTensorProduct_bijective_of_map_from_fin [Fintype ι] [IsNoetherianRing R] : Function.Bijective (ofTensorProduct I M) := by - have : IsIso (ModuleCat.asHom (ofTensorProduct I M)) := + have : IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := ofTensorProduct_iso I M f hf - exact ConcreteCategory.bijective_of_isIso (ModuleCat.asHom (ofTensorProduct I M)) + exact ConcreteCategory.bijective_of_isIso (ModuleCat.ofHom (ofTensorProduct I M)) end @@ -375,8 +375,8 @@ lemma tensor_map_id_left_injective_of_injective (hf : Function.Injective f) : end /-- Adic completion of a Noetherian ring `R` is flat over `R`. -/ -instance flat_of_isNoetherian [IsNoetherianRing R] : Algebra.Flat R (AdicCompletion I R) where - out := (Module.Flat.iff_lTensor_injective' R (AdicCompletion I R)).mpr <| fun J ↦ +instance flat_of_isNoetherian [IsNoetherianRing R] : Module.Flat R (AdicCompletion I R) := + (Module.Flat.iff_lTensor_injective' R (AdicCompletion I R)).mpr fun J ↦ tensor_map_id_left_injective_of_injective I (Submodule.injective_subtype J) end Noetherian diff --git a/Mathlib/RingTheory/AdicCompletion/Basic.lean b/Mathlib/RingTheory/AdicCompletion/Basic.lean index a680f5898fdf6..62028936b50db 100644 --- a/Mathlib/RingTheory/AdicCompletion/Basic.lean +++ b/Mathlib/RingTheory/AdicCompletion/Basic.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Judith Ludwig, Christian Merten -/ import Mathlib.Algebra.GeomSum import Mathlib.LinearAlgebra.SModEq -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal /-! # Completion of a module with respect to an ideal. @@ -29,7 +29,7 @@ suppress_compilation open Submodule -variable {R : Type*} [CommRing R] (I : Ideal R) +variable {R S T : Type*} [CommRing R] (I : Ideal R) variable (M : Type*) [AddCommGroup M] [Module R M] variable {N : Type*} [AddCommGroup N] [Module R N] @@ -202,24 +202,62 @@ instance : Neg (AdicCompletion I M) where instance : Sub (AdicCompletion I M) where sub x y := ⟨x.val - y.val, by simp [x.property, y.property]⟩ -instance : SMul ℕ (AdicCompletion I M) where - smul n x := ⟨n • x.val, by simp [x.property]⟩ +instance instSMul [SMul S R] [SMul S M] [IsScalarTower S R M] : SMul S (AdicCompletion I M) where + smul r x := ⟨r • x.val, by simp [x.property]⟩ + +@[simp, norm_cast] lemma val_zero : (0 : AdicCompletion I M).val = 0 := rfl + +lemma val_zero_apply (n : ℕ) : (0 : AdicCompletion I M).val n = 0 := rfl + +variable {I M} + +@[simp, norm_cast] lemma val_add (f g : AdicCompletion I M) : (f + g).val = f.val + g.val := rfl +@[simp, norm_cast] lemma val_sub (f g : AdicCompletion I M) : (f - g).val = f.val - g.val := rfl +@[simp, norm_cast] lemma val_neg (f : AdicCompletion I M) : (-f).val = -f.val := rfl + +lemma val_add_apply (f g : AdicCompletion I M) (n : ℕ) : (f + g).val n = f.val n + g.val n := rfl +lemma val_sub_apply (f g : AdicCompletion I M) (n : ℕ) : (f - g).val n = f.val n - g.val n := rfl +lemma val_neg_apply (f : AdicCompletion I M) (n : ℕ) : (-f).val n = -f.val n := rfl + +/- No `simp` attribute, since it causes `simp` unification timeouts when considering +the `Module (AdicCompletion I R) (AdicCompletion I M)` instance (see `AdicCompletion/Algebra`). -/ +@[norm_cast] +lemma val_smul [SMul S R] [SMul S M] [IsScalarTower S R M] (s : S) (f : AdicCompletion I M) : + (s • f).val = s • f.val := rfl + +lemma val_smul_apply [SMul S R] [SMul S M] [IsScalarTower S R M] (s : S) (f : AdicCompletion I M) + (n : ℕ) : (s • f).val n = s • f.val n := rfl + +@[ext] +lemma ext {x y : AdicCompletion I M} (h : ∀ n, x.val n = y.val n) : x = y := Subtype.eq <| funext h -instance : SMul ℤ (AdicCompletion I M) where - smul n x := ⟨n • x.val, by simp [x.property]⟩ +variable (I M) instance : AddCommGroup (AdicCompletion I M) := let f : AdicCompletion I M → ∀ n, M ⧸ (I ^ n • ⊤ : Submodule R M) := Subtype.val - Subtype.val_injective.addCommGroup f rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) - (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) + Subtype.val_injective.addCommGroup f rfl val_add val_neg val_sub (fun _ _ ↦ val_smul ..) + (fun _ _ ↦ val_smul ..) -instance : SMul R (AdicCompletion I M) where - smul r x := ⟨r • x.val, by simp [x.property]⟩ - -instance : Module R (AdicCompletion I M) := +instance [Semiring S] [SMul S R] [Module S M] [IsScalarTower S R M] : + Module S (AdicCompletion I M) := let f : AdicCompletion I M →+ ∀ n, M ⧸ (I ^ n • ⊤ : Submodule R M) := { toFun := Subtype.val, map_zero' := rfl, map_add' := fun _ _ ↦ rfl } - Subtype.val_injective.module R f (fun _ _ ↦ rfl) + Subtype.val_injective.module S f val_smul + +instance instIsScalarTower [SMul S T] [SMul S R] [SMul T R] [SMul S M] [SMul T M] + [IsScalarTower S R M] [IsScalarTower T R M] [IsScalarTower S T M] : + IsScalarTower S T (AdicCompletion I M) where + smul_assoc s t f := by ext; simp [val_smul] + +instance instSMulCommClass [SMul S R] [SMul T R] [SMul S M] [SMul T M] + [IsScalarTower S R M] [IsScalarTower T R M] [SMulCommClass S T M] : + SMulCommClass S T (AdicCompletion I M) where + smul_comm s t f := by ext; simp [val_smul, smul_comm] + +instance instIsCentralScalar [SMul S R] [SMul Sᵐᵒᵖ R] [SMul S M] [SMul Sᵐᵒᵖ M] + [IsScalarTower S R M] [IsScalarTower Sᵐᵒᵖ R M] [IsCentralScalar S M] : + IsCentralScalar S (AdicCompletion I M) where + op_smul_eq_smul s f := by ext; simp [val_smul, op_smul_eq_smul] /-- The canonical inclusion from the completion to the product. -/ @[simps] @@ -228,6 +266,18 @@ def incl : AdicCompletion I M →ₗ[R] (∀ n, M ⧸ (I ^ n • ⊤ : Submodule map_add' _ _ := rfl map_smul' _ _ := rfl +variable {I M} + +@[simp, norm_cast] +lemma val_sum {ι : Type*} (s : Finset ι) (f : ι → AdicCompletion I M) : + (∑ i ∈ s, f i).val = ∑ i ∈ s, (f i).val := by + simp_rw [← funext (incl_apply _ _ _), map_sum] + +lemma val_sum_apply {ι : Type*} (s : Finset ι) (f : ι → AdicCompletion I M) (n : ℕ) : + (∑ i ∈ s, f i).val n = ∑ i ∈ s, (f i).val n := by simp + +variable (I M) + /-- The canonical linear map to the completion. -/ def of : M →ₗ[R] AdicCompletion I M where toFun x := ⟨fun n => mkQ (I ^ n • ⊤ : Submodule R M) x, fun _ => rfl⟩ @@ -266,43 +316,17 @@ theorem eval_surjective (n : ℕ) : Function.Surjective (eval I M n) := fun x theorem range_eval (n : ℕ) : LinearMap.range (eval I M n) = ⊤ := LinearMap.range_eq_top.2 (eval_surjective I M n) -@[simp] -theorem val_zero (n : ℕ) : (0 : AdicCompletion I M).val n = 0 := - rfl - variable {I M} -@[simp] -theorem val_add (n : ℕ) (f g : AdicCompletion I M) : (f + g).val n = f.val n + g.val n := - rfl - -@[simp] -theorem val_sub (n : ℕ) (f g : AdicCompletion I M) : (f - g).val n = f.val n - g.val n := - rfl - -@[simp] -theorem val_sum {α : Type*} (s : Finset α) (f : α → AdicCompletion I M) (n : ℕ) : - (Finset.sum s f).val n = Finset.sum s (fun a ↦ (f a).val n) := by - simp_rw [← incl_apply, map_sum, Finset.sum_apply] - -/- No `simp` attribute, since it causes `simp` unification timeouts when considering -the `AdicCompletion I R` module instance on `AdicCompletion I M` (see `AdicCompletion/Algebra`). -/ -theorem val_smul (n : ℕ) (r : R) (f : AdicCompletion I M) : (r • f).val n = r • f.val n := - rfl - -@[ext] -theorem ext {x y : AdicCompletion I M} (h : ∀ n, x.val n = y.val n) : x = y := - Subtype.eq <| funext h - variable (I M) instance : IsHausdorff I (AdicCompletion I M) where haus' x h := ext fun n ↦ by refine smul_induction_on (SModEq.zero.1 <| h n) (fun r hr x _ ↦ ?_) (fun x y hx hy ↦ ?_) - · simp only [val_smul, val_zero] + · simp only [val_smul_apply, val_zero] exact Quotient.inductionOn' (x.val n) (fun a ↦ SModEq.zero.2 <| smul_mem_smul hr mem_top) - · simp only [val_add, hx, val_zero, hy, add_zero] + · simp only [val_add_apply, hx, val_zero_apply, hy, add_zero] @[simp] theorem transitionMap_mk {m n : ℕ} (hmn : m ≤ n) (x : M) : diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index 03b236a454123..4ac8d7b974ccc 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -302,7 +302,7 @@ theorem sum_comp_sumInv : sum I M ∘ₗ sumInv I M = LinearMap.id := by simp only [LinearMap.coe_comp, Function.comp_apply, LinearMap.id_coe, id_eq, mk_apply_coe, Submodule.mkQ_apply] rw [← DirectSum.sum_univ_of (((sumInv I M) ((AdicCompletion.mk I (⨁ (j : ι), M j)) f)))] - simp only [sumInv_apply, map_mk, map_sum, sum_of, val_sum, mk_apply_coe, + simp only [sumInv_apply, map_mk, map_sum, sum_of, val_sum_apply, mk_apply_coe, AdicCauchySequence.map_apply_coe, Submodule.mkQ_apply] simp only [← Submodule.mkQ_apply, ← map_sum] erw [DirectSum.sum_univ_of] diff --git a/Mathlib/RingTheory/Algebraic/Basic.lean b/Mathlib/RingTheory/Algebraic/Basic.lean index 2768c2dca2d4d..617f4e02b975c 100644 --- a/Mathlib/RingTheory/Algebraic/Basic.lean +++ b/Mathlib/RingTheory/Algebraic/Basic.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import Mathlib.Algebra.Polynomial.Expand import Mathlib.Algebra.Polynomial.Roots +import Mathlib.RingTheory.Adjoin.Polynomial import Mathlib.RingTheory.Algebraic.Defs import Mathlib.RingTheory.Polynomial.Tower @@ -31,10 +32,25 @@ theorem is_transcendental_of_subsingleton [Subsingleton R] (x : A) : Transcenden variable {R} -variable (R) in +theorem IsAlgebraic.nontrivial {a : A} (h : IsAlgebraic R a) : Nontrivial R := by + contrapose! h + rw [not_nontrivial_iff_subsingleton] at h + apply is_transcendental_of_subsingleton + +variable (R A) + +theorem Algebra.IsAlgebraic.nontrivial [alg : Algebra.IsAlgebraic R A] : Nontrivial R := + (alg.1 0).nontrivial + +instance (priority := low) Algebra.transcendental_of_subsingleton [Subsingleton R] : + Algebra.Transcendental R A := + ⟨⟨0, is_transcendental_of_subsingleton R 0⟩⟩ + theorem Polynomial.transcendental_X : Transcendental R (X (R := R)) := by simp [transcendental_iff] +variable {R A} + theorem IsAlgebraic.of_aeval {r : A} (f : R[X]) (hf : f.natDegree ≠ 0) (hf' : f.leadingCoeff ∈ nonZeroDivisors R) (H : IsAlgebraic R (aeval r f)) : IsAlgebraic R r := by @@ -93,6 +109,17 @@ theorem transcendental_iff_ker_eq_bot {x : A} : Transcendental R x ↔ RingHom.ker (aeval (R := R) x) = ⊥ := by rw [transcendental_iff_injective, RingHom.injective_iff_ker_eq_bot] +theorem Algebra.isAlgebraic_of_not_injective (h : ¬ Function.Injective (algebraMap R A)) : + Algebra.IsAlgebraic R A where + isAlgebraic a := isAlgebraic_iff_not_injective.mpr + fun inj ↦ h <| by convert inj.comp C_injective; ext; simp + +theorem Algebra.injective_of_transcendental [h : Algebra.Transcendental R A] : + Function.Injective (algebraMap R A) := by + rw [transcendental_iff_not_isAlgebraic] at h + contrapose! h + exact isAlgebraic_of_not_injective h + end section zero_ne_one @@ -129,6 +156,16 @@ theorem isAlgebraic_of_mem_rootSet {R : Type u} {A : Type v} [Field R] [Field A] {p : R[X]} {x : A} (hx : x ∈ p.rootSet A) : IsAlgebraic R x := ⟨p, ne_zero_of_mem_rootSet hx, aeval_eq_zero_of_mem_rootSet hx⟩ +variable (S) in +theorem IsLocalization.isAlgebraic [Nontrivial R] (M : Submonoid R) [IsLocalization M S] : + Algebra.IsAlgebraic R S where + isAlgebraic x := by + obtain rfl | hx := eq_or_ne x 0 + · exact isAlgebraic_zero + have ⟨⟨r, m⟩, h⟩ := surj M x + refine ⟨C m.1 * X - C r, fun eq ↦ hx ?_, by simpa [sub_eq_zero, mul_comm x] using h⟩ + rwa [← eq_mk'_iff_mul_eq, show r = 0 by simpa using congr(coeff $eq 0), mk'_zero] at h + open IsScalarTower protected theorem IsAlgebraic.algebraMap {a : S} : @@ -318,11 +355,9 @@ instance algebra_isAlgebraic_bot_right [Nontrivial R] : end Subalgebra theorem IsAlgebraic.of_pow {r : A} {n : ℕ} (hn : 0 < n) (ht : IsAlgebraic R (r ^ n)) : - IsAlgebraic R r := by - obtain ⟨p, p_nonzero, hp⟩ := ht - refine ⟨Polynomial.expand _ n p, ?_, ?_⟩ - · rwa [Polynomial.expand_ne_zero hn] - · rwa [Polynomial.expand_aeval n p r] + IsAlgebraic R r := + have ⟨p, p_nonzero, hp⟩ := ht + ⟨_, by rwa [expand_ne_zero hn], by rwa [expand_aeval n p r]⟩ theorem Transcendental.pow {r : A} (ht : Transcendental R r) {n : ℕ} (hn : 0 < n) : Transcendental R (r ^ n) := fun ht' ↦ ht <| ht'.of_pow hn @@ -449,8 +484,7 @@ section NoZeroSMulDivisors namespace Algebra.IsAlgebraic -variable [CommRing K] [Field L] -variable [Algebra K L] +variable [CommRing K] [Field L] [Algebra K L] theorem algHom_bijective [NoZeroSMulDivisors K L] [Algebra.IsAlgebraic K L] (f : L →ₐ[K] L) : Function.Bijective f := by @@ -502,42 +536,49 @@ variable {R S : Type*} [CommRing R] [Ring S] [Algebra R S] theorem IsAlgebraic.exists_nonzero_coeff_and_aeval_eq_zero {s : S} (hRs : IsAlgebraic R s) (hs : s ∈ nonZeroDivisors S) : - ∃ (q : Polynomial R), q.coeff 0 ≠ 0 ∧ aeval s q = 0 := by + ∃ q : R[X], q.coeff 0 ≠ 0 ∧ aeval s q = 0 := by obtain ⟨p, hp0, hp⟩ := hRs - obtain ⟨q, hpq, hq⟩ := Polynomial.exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp0 0 + obtain ⟨q, hpq, hq⟩ := exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp0 0 simp only [C_0, sub_zero, X_pow_mul, X_dvd_iff] at hpq hq rw [hpq, map_mul, aeval_X_pow] at hp exact ⟨q, hq, (nonZeroDivisors S).pow_mem hs (rootMultiplicity 0 p) (aeval s q) hp⟩ +theorem IsAlgebraic.exists_nonzero_eq_adjoin_mul + {s : S} (hRs : IsAlgebraic R s) (hs : s ∈ nonZeroDivisors S) : + ∃ᵉ (t ∈ Algebra.adjoin R {s}) (r ≠ (0 : R)), s * t = algebraMap R S r := by + have ⟨q, hq0, hq⟩ := hRs.exists_nonzero_coeff_and_aeval_eq_zero hs + have ⟨p, hp⟩ := X_dvd_sub_C (p := q) + refine ⟨aeval s p, aeval_mem_adjoin_singleton _ _, _, neg_ne_zero.mpr hq0, ?_⟩ + apply_fun aeval s at hp + rwa [map_sub, hq, zero_sub, map_mul, aeval_X, aeval_C, ← map_neg, eq_comm] at hp + theorem IsAlgebraic.exists_nonzero_dvd {s : S} (hRs : IsAlgebraic R s) (hs : s ∈ nonZeroDivisors S) : - ∃ r : R, r ≠ 0 ∧ s ∣ (algebraMap R S) r := by + ∃ r : R, r ≠ 0 ∧ s ∣ algebraMap R S r := by obtain ⟨q, hq0, hq⟩ := hRs.exists_nonzero_coeff_and_aeval_eq_zero hs - have key := map_dvd (Polynomial.aeval s) (Polynomial.X_dvd_sub_C (p := q)) - rw [map_sub, hq, zero_sub, dvd_neg, Polynomial.aeval_X, Polynomial.aeval_C] at key + have key := map_dvd (aeval s) (X_dvd_sub_C (p := q)) + rw [map_sub, hq, zero_sub, dvd_neg, aeval_X, aeval_C] at key exact ⟨q.coeff 0, hq0, key⟩ /-- A fraction `(a : S) / (b : S)` can be reduced to `(c : S) / (d : R)`, if `b` is algebraic over `R`. -/ theorem IsAlgebraic.exists_smul_eq_mul (a : S) {b : S} (hRb : IsAlgebraic R b) (hb : b ∈ nonZeroDivisors S) : - ∃ᵉ (c : S) (d ≠ (0 : R)), d • a = b * c := by - obtain ⟨r, hr, s, h⟩ := IsAlgebraic.exists_nonzero_dvd (R := R) (S := S) hRb hb - exact ⟨s * a, r, hr, by rw [Algebra.smul_def, h, mul_assoc]⟩ + ∃ᵉ (c : S) (d ≠ (0 : R)), d • a = b * c := + have ⟨r, hr, s, h⟩ := hRb.exists_nonzero_dvd hb + ⟨s * a, r, hr, by rw [Algebra.smul_def, h, mul_assoc]⟩ variable (R) /-- A fraction `(a : S) / (b : S)` can be reduced to `(c : S) / (d : R)`, if `b` is algebraic over `R`. -/ -theorem Algebra.IsAlgebraic.exists_smul_eq_mul [IsDomain S] [Algebra.IsAlgebraic R S] +theorem Algebra.IsAlgebraic.exists_smul_eq_mul [NoZeroDivisors S] [Algebra.IsAlgebraic R S] (a : S) {b : S} (hb : b ≠ 0) : ∃ᵉ (c : S) (d ≠ (0 : R)), d • a = b * c := - (Algebra.IsAlgebraic.isAlgebraic b).exists_smul_eq_mul a (mem_nonZeroDivisors_iff_ne_zero.mpr hb) + (isAlgebraic b).exists_smul_eq_mul a (mem_nonZeroDivisors_of_ne_zero hb) end -variable {R S : Type*} [CommRing R] [CommRing S] - section Field variable {K L : Type*} [Field K] [Field L] [Algebra K L] (A : Subalgebra K L) @@ -601,12 +642,13 @@ end Field section Infinite -theorem Transcendental.infinite {R A : Type*} [CommRing R] [Ring A] [Algebra R A] - [Nontrivial R] {x : A} (hx : Transcendental R x) : Infinite A := +variable {R A : Type*} [CommRing R] [Ring A] [Algebra R A] [Nontrivial R] + +theorem Transcendental.infinite {x : A} (hx : Transcendental R x) : Infinite A := .of_injective _ (transcendental_iff_injective.mp hx) -theorem Algebra.Transcendental.infinite (R A : Type*) [CommRing R] [Ring A] [Algebra R A] - [Nontrivial R] [Algebra.Transcendental R A] : Infinite A := +variable (R A) in +theorem Algebra.Transcendental.infinite [Algebra.Transcendental R A] : Infinite A := have ⟨x, hx⟩ := ‹Algebra.Transcendental R A› hx.infinite diff --git a/Mathlib/RingTheory/Algebraic/Integral.lean b/Mathlib/RingTheory/Algebraic/Integral.lean index b45ecf7023b2f..5a9c3fa6e04f6 100644 --- a/Mathlib/RingTheory/Algebraic/Integral.lean +++ b/Mathlib/RingTheory/Algebraic/Integral.lean @@ -19,7 +19,17 @@ is algebraic and that every algebraic element over a field is integral. over a field. * `IsAlgebraic.of_finite`, `Algebra.IsAlgebraic.of_finite`: finite-dimensional (as module) implies algebraic. -* `exists_integral_multiple`: an algebraic element has a multiple which is integral +* `IsAlgebraic.exists_integral_multiple`: an algebraic element has a multiple which is integral +* `IsAlgebraic.iff_exists_smul_integral`: If `R` is reduced and `S` is an `R`-algebra with + injective `algebraMap`, then an element of `S` is algebraic over `R` iff some `R`-multiple + is integral over `R`. +* `Algebra.IsAlgebraic.trans'`: If `A/S/R` is a tower of algebras and both `A/S` and `S/R` are + algebraic, then `A/R` is also algebraic, provided that `S` has no zero divisors and + `algebraMap S A` is injective (so `S` can be regarded as an `R`-subalgebra of `A`). +* `Subalgebra.algebraicClosure`: If `R` is a domain and `S` is an arbitrary `R`-algebra, + then the elements of `S` that are algebraic over `R` form a subalgebra. +* `Transcendental.extendScalars`: an element of an `R`-algebra that is transcendental over `R` + remains transcendental over any algebraic `R`-subalgebra that has no zero divisors. -/ assert_not_exists LocalRing @@ -66,11 +76,11 @@ alias ⟨IsAlgebraic.isIntegral, _⟩ := isAlgebraic_iff_isIntegral protected instance Algebra.IsAlgebraic.isIntegral [Algebra.IsAlgebraic K A] : Algebra.IsIntegral K A := Algebra.isAlgebraic_iff_isIntegral.mp ‹_› -variable (K) in -theorem Algebra.IsAlgebraic.of_isIntegralClosure (B C : Type*) - [CommRing B] [CommRing C] [Algebra K B] [Algebra K C] [Algebra B C] - [IsScalarTower K B C] [IsIntegralClosure B K C] : Algebra.IsAlgebraic K B := - Algebra.isAlgebraic_iff_isIntegral.mpr (IsIntegralClosure.isIntegral_algebra K C) +theorem Algebra.IsAlgebraic.of_isIntegralClosure (R B C : Type*) [CommRing R] [Nontrivial R] + [CommRing B] [CommRing C] [Algebra R B] [Algebra R C] [Algebra B C] + [IsScalarTower R B C] [IsIntegralClosure B R C] : Algebra.IsAlgebraic R B := + have := IsIntegralClosure.isIntegral_algebra R (A := B) C + inferInstance end Field @@ -146,15 +156,249 @@ end Field end -variable {R S : Type*} [CommRing R] [CommRing S] +variable {R S A : Type*} [CommRing R] [CommRing S] [Ring A] +variable [Algebra R S] [Algebra R A] [Algebra S A] [IsScalarTower R S A] +variable {z : A} {z' : S} -theorem exists_integral_multiple [Algebra R S] {z : S} (hz : IsAlgebraic R z) - (inj : ∀ x, algebraMap R S x = 0 → x = 0) : - ∃ᵉ (x : integralClosure R S) (y ≠ (0 : R)), algebraMap R S y * z = x := by - rcases hz with ⟨p, p_ne_zero, px⟩ +namespace IsAlgebraic + +theorem exists_integral_multiple (hz : IsAlgebraic R z) + (inj : Function.Injective (algebraMap R A)) : + ∃ y ≠ (0 : R), IsIntegral R (y • z) := by + have ⟨p, p_ne_zero, px⟩ := hz set a := p.leadingCoeff have a_ne_zero : a ≠ 0 := mt Polynomial.leadingCoeff_eq_zero.mp p_ne_zero - have x_integral : IsIntegral R (algebraMap R S a * z) := + have x_integral : IsIntegral R (algebraMap R A a * z) := ⟨p.integralNormalization, monic_integralNormalization p_ne_zero, - integralNormalization_aeval_eq_zero px inj⟩ - exact ⟨⟨_, x_integral⟩, a, a_ne_zero, rfl⟩ + integralNormalization_aeval_eq_zero px fun _ ↦ (map_eq_zero_iff _ inj).mp⟩ + exact ⟨_, a_ne_zero, Algebra.smul_def a z ▸ x_integral⟩ + +@[deprecated (since := "2024-11-30")] +alias _root_.exists_integral_multiple := exists_integral_multiple + +theorem _root_.Algebra.IsAlgebraic.exists_integral_multiples [NoZeroDivisors R] + [alg : Algebra.IsAlgebraic R A] (inj : Function.Injective (algebraMap R A)) (s : Finset A) : + ∃ y ≠ (0 : R), ∀ z ∈ s, IsIntegral R (y • z) := by + have := Algebra.IsAlgebraic.nontrivial R A + choose r hr int using fun x ↦ (alg.1 x).exists_integral_multiple inj + refine ⟨∏ x ∈ s, r x, Finset.prod_ne_zero_iff.mpr fun _ _ ↦ hr _, fun _ h ↦ ?_⟩ + classical rw [← Finset.prod_erase_mul _ _ h, mul_smul] + exact (int _).smul _ + +theorem of_smul_isIntegral {y : R} (hy : ¬ IsNilpotent y) + (h : IsIntegral R (y • z)) : IsAlgebraic R z := by + have ⟨p, monic, eval0⟩ := h + refine ⟨p.comp (C y * X), fun h ↦ ?_, by simpa [aeval_comp, Algebra.smul_def] using eval0⟩ + apply_fun (coeff · p.natDegree) at h + have hy0 : y ≠ 0 := by rintro rfl; exact hy .zero + rw [coeff_zero, ← mul_one p.natDegree, ← natDegree_C_mul_X y hy0, + coeff_comp_degree_mul_degree, monic, one_mul, leadingCoeff_C_mul_X] at h + · exact hy ⟨_, h⟩ + · rw [natDegree_C_mul_X _ hy0]; rintro ⟨⟩ + +theorem of_smul {y : R} (hy : y ∈ nonZeroDivisors R) + (h : IsAlgebraic R (y • z)) : IsAlgebraic R z := + have ⟨p, hp, eval0⟩ := h + ⟨_, mt (comp_C_mul_X_eq_zero_iff hy).mp hp, by simpa [aeval_comp, Algebra.smul_def] using eval0⟩ + +theorem iff_exists_smul_integral [IsReduced R] (inj : Function.Injective (algebraMap R A)) : + IsAlgebraic R z ↔ ∃ y ≠ (0 : R), IsIntegral R (y • z) := + ⟨(exists_integral_multiple · inj), fun ⟨_, hy, int⟩ ↦ + of_smul_isIntegral (by rwa [isNilpotent_iff_eq_zero]) int⟩ + +section trans + +variable (R) [NoZeroDivisors S] (inj : Function.Injective (algebraMap S A)) +include inj + +theorem restrictScalars_of_isIntegral [int : Algebra.IsIntegral R S] + {a : A} (h : IsAlgebraic S a) : IsAlgebraic R a := by + by_cases hRS : Function.Injective (algebraMap R S) + on_goal 2 => exact (Algebra.isAlgebraic_of_not_injective + fun h ↦ hRS <| .of_comp (IsScalarTower.algebraMap_eq R S A ▸ h)).1 _ + have := hRS.noZeroDivisors _ (map_zero _) (map_mul _) + have ⟨s, hs, int_s⟩ := h.exists_integral_multiple inj + cases subsingleton_or_nontrivial R + · have := Module.subsingleton R S + exact (is_transcendental_of_subsingleton _ _ h).elim + have ⟨r, hr, _, e⟩ := (int.1 s).isAlgebraic.exists_nonzero_dvd (mem_nonZeroDivisors_of_ne_zero hs) + refine .of_smul_isIntegral (y := r) (by rwa [isNilpotent_iff_eq_zero]) ?_ + rw [Algebra.smul_def, IsScalarTower.algebraMap_apply R S, + e, ← Algebra.smul_def, mul_comm, mul_smul] + exact isIntegral_trans _ (int_s.smul _) + +theorem restrictScalars [Algebra.IsAlgebraic R S] + {a : A} (h : IsAlgebraic S a) : IsAlgebraic R a := by + have ⟨p, hp, eval0⟩ := h + by_cases hRS : Function.Injective (algebraMap R S) + on_goal 2 => exact (Algebra.isAlgebraic_of_not_injective + fun h ↦ hRS <| .of_comp (IsScalarTower.algebraMap_eq R S A ▸ h)).1 _ + have := hRS.noZeroDivisors _ (map_zero _) (map_mul _) + classical + have ⟨r, hr, int⟩ := Algebra.IsAlgebraic.exists_integral_multiples hRS (p.support.image (coeff p)) + let p := (r • p).toSubring (integralClosure R S).toSubring fun s hs ↦ by + obtain ⟨n, hn, rfl⟩ := mem_coeffs_iff.mp hs + 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 + 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, + map_toSubring, eval_map_algebraMap, ← AlgHom.restrictScalars_apply R, + map_smul, AlgHom.restrictScalars_apply, eval0, smul_zero] + exact restrictScalars_of_isIntegral _ (by exact inj.comp Subtype.val_injective) this + +theorem _root_.IsIntegral.trans_isAlgebraic [alg : Algebra.IsAlgebraic R S] + {a : A} (h : IsIntegral S a) : IsAlgebraic R a := by + cases subsingleton_or_nontrivial A + · have := Algebra.IsAlgebraic.nontrivial R S + exact Subsingleton.elim a 0 ▸ isAlgebraic_zero + · have := Module.nontrivial S A + exact h.isAlgebraic.restrictScalars _ inj + +end trans + +variable [nzd : NoZeroDivisors R] {a b : S} (ha : IsAlgebraic R a) (hb : IsAlgebraic R b) +include ha +omit nzd + +protected lemma neg : IsAlgebraic R (-a) := + have ⟨p, h, eval0⟩ := ha + ⟨algEquivAevalNegX p, EmbeddingLike.map_ne_zero_iff.mpr h, by simpa [← comp_eq_aeval, aeval_comp]⟩ + +protected lemma smul (r : R) : IsAlgebraic R (r • a) := + have ⟨_, hp, eval0⟩ := ha + ⟨_, scaleRoots_ne_zero hp r, Algebra.smul_def r a ▸ scaleRoots_aeval_eq_zero eval0⟩ + +protected lemma nsmul (n : ℕ) : IsAlgebraic R (n • a) := + Nat.cast_smul_eq_nsmul R n a ▸ ha.smul _ + +protected lemma zsmul (n : ℤ) : IsAlgebraic R (n • a) := + Int.cast_smul_eq_zsmul R n a ▸ ha.smul _ + +include hb nzd + +protected lemma mul : IsAlgebraic R (a * b) := by + refine (em _).elim (fun h ↦ ?_) fun h ↦ (Algebra.isAlgebraic_of_not_injective h).1 _ + have ⟨ra, a0, int_a⟩ := ha.exists_integral_multiple h + have ⟨rb, b0, int_b⟩ := hb.exists_integral_multiple h + refine (IsAlgebraic.iff_exists_smul_integral h).mpr ⟨_, mul_ne_zero a0 b0, ?_⟩ + simp_rw [Algebra.smul_def, map_mul, mul_mul_mul_comm, ← Algebra.smul_def] + exact int_a.mul int_b + +protected lemma add : IsAlgebraic R (a + b) := by + refine (em _).elim (fun h ↦ ?_) fun h ↦ (Algebra.isAlgebraic_of_not_injective h).1 _ + have ⟨ra, a0, int_a⟩ := ha.exists_integral_multiple h + have ⟨rb, b0, int_b⟩ := hb.exists_integral_multiple h + refine (IsAlgebraic.iff_exists_smul_integral h).mpr ⟨_, mul_ne_zero b0 a0, ?_⟩ + rw [smul_add, mul_smul, mul_comm, mul_smul] + exact (int_a.smul _).add (int_b.smul _) + +protected lemma sub : IsAlgebraic R (a - b) := + sub_eq_add_neg a b ▸ ha.add hb.neg + +omit hb +protected lemma pow (n : ℕ) : IsAlgebraic R (a ^ n) := + have := ha.nontrivial + n.rec (pow_zero a ▸ isAlgebraic_one) fun _ h ↦ pow_succ a _ ▸ h.mul ha + +end IsAlgebraic + +namespace Algebra + +variable (R) [NoZeroDivisors S] (inj : Function.Injective (algebraMap S A)) +include inj + +/-- Transitivity of algebraicity for algebras over domains. -/ +theorem IsAlgebraic.trans' [Algebra.IsAlgebraic R S] [alg : Algebra.IsAlgebraic S A] : + Algebra.IsAlgebraic R A := + ⟨fun _ ↦ (alg.1 _).restrictScalars _ inj⟩ + +theorem IsIntegral.trans_isAlgebraic [Algebra.IsIntegral R S] [alg : Algebra.IsAlgebraic S A] : + Algebra.IsAlgebraic R A := + ⟨fun _ ↦ (alg.1 _).restrictScalars_of_isIntegral _ inj⟩ + +theorem IsAlgebraic.trans_isIntegral [Algebra.IsAlgebraic R S] [int : Algebra.IsIntegral S A] : + Algebra.IsAlgebraic R A := + ⟨fun _ ↦ (int.1 _).trans_isAlgebraic _ inj⟩ + +end Algebra + +variable (R S) +/-- If `R` is a domain and `S` is an arbitrary `R`-algebra, then the elements of `S` +that are algebraic over `R` form a subalgebra. -/ +def Subalgebra.algebraicClosure [IsDomain R] : Subalgebra R S where + carrier := {s | _root_.IsAlgebraic R s} + mul_mem' ha hb := ha.mul hb + add_mem' ha hb := ha.add hb + algebraMap_mem' := isAlgebraic_algebraMap + +theorem integralClosure_le_algebraicClosure [IsDomain R] : + integralClosure R S ≤ Subalgebra.algebraicClosure R S := + fun _ ↦ IsIntegral.isAlgebraic + +theorem Subalgebra.algebraicClosure_eq_integralClosure {K} [Field K] [Algebra K S] : + algebraicClosure K S = integralClosure K S := + SetLike.ext fun _ ↦ isAlgebraic_iff_isIntegral + +instance [IsDomain R] : Algebra.IsAlgebraic R (Subalgebra.algebraicClosure R S) := + (Subalgebra.isAlgebraic_iff _).mp fun _ ↦ id + +variable {R S} + +theorem Algebra.isAlgebraic_adjoin_iff [IsDomain R] {s : Set S} : + (adjoin R s).IsAlgebraic ↔ ∀ x ∈ s, IsAlgebraic R x := + Algebra.adjoin_le_iff (S := Subalgebra.algebraicClosure R S) + +theorem Algebra.isAlgebraic_adjoin_of_nonempty [NoZeroDivisors R] {s : Set S} (hs : s.Nonempty) : + (adjoin R s).IsAlgebraic ↔ ∀ x ∈ s, IsAlgebraic R x := + ⟨fun h x hx ↦ h _ (subset_adjoin hx), fun h ↦ + have ⟨x, hx⟩ := hs + have := (isDomain_iff_noZeroDivisors_and_nontrivial _).mpr ⟨‹_›, (h x hx).nontrivial⟩ + isAlgebraic_adjoin_iff.mpr h⟩ + +theorem Algebra.isAlgebraic_adjoin_singleton_iff [NoZeroDivisors R] {s : S} : + (adjoin R {s}).IsAlgebraic ↔ IsAlgebraic R s := + (isAlgebraic_adjoin_of_nonempty <| Set.singleton_nonempty s).trans forall_eq + +theorem IsAlgebraic.of_mul [NoZeroDivisors R] {y z : S} (hy : y ∈ nonZeroDivisors S) + (alg_y : IsAlgebraic R y) (alg_yz : IsAlgebraic R (y * z)) : IsAlgebraic R z := by + have ⟨t, ht, r, hr, eq⟩ := alg_y.exists_nonzero_eq_adjoin_mul hy + have := alg_yz.mul (Algebra.isAlgebraic_adjoin_singleton_iff.mpr alg_y _ ht) + rw [mul_right_comm, eq, ← Algebra.smul_def] at this + exact this.of_smul (mem_nonZeroDivisors_of_ne_zero hr) + +namespace Transcendental + +section + +variable {a : A} (ha : Transcendental R a) +include ha + +lemma extendScalars_of_isIntegral [NoZeroDivisors S] [Algebra.IsIntegral R S] + (inj : Function.Injective (algebraMap S A)) : Transcendental S a := by + contrapose ha + rw [Transcendental, not_not] at ha ⊢ + exact ha.restrictScalars_of_isIntegral _ inj + +lemma extendScalars [NoZeroDivisors S] [Algebra.IsAlgebraic R S] + (inj : Function.Injective (algebraMap S A)) : Transcendental S a := by + contrapose ha + rw [Transcendental, not_not] at ha ⊢ + exact ha.restrictScalars _ inj + +end + +variable {a : S} (ha : Transcendental R a) +include ha + +protected lemma integralClosure [NoZeroDivisors S] : + Transcendental (integralClosure R S) a := + ha.extendScalars_of_isIntegral Subtype.val_injective + +lemma subalgebraAlgebraicClosure [IsDomain R] [NoZeroDivisors S] : + Transcendental (Subalgebra.algebraicClosure R S) a := + ha.extendScalars Subtype.val_injective + +end Transcendental diff --git a/Mathlib/RingTheory/ChainOfDivisors.lean b/Mathlib/RingTheory/ChainOfDivisors.lean index a1921d9752ce6..68af812321e63 100644 --- a/Mathlib/RingTheory/ChainOfDivisors.lean +++ b/Mathlib/RingTheory/ChainOfDivisors.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen, Paul Lezeau -/ import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.Algebra.IsPrimePow -import Mathlib.Algebra.Squarefree.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity import Mathlib.Data.ZMod.Defs import Mathlib.Order.Atoms import Mathlib.Order.Hom.Bounded @@ -58,7 +58,7 @@ theorem Associates.isAtom_iff {p : Associates M} (h₁ : p ≠ 0) : IsAtom p ↔ ⟨(ha.unit⁻¹ : Units _), by rw [hab, mul_assoc, IsUnit.mul_val_inv ha, mul_one]⟩) hb⟩⟩ -open UniqueFactorizationMonoid multiplicity Irreducible Associates +open UniqueFactorizationMonoid Irreducible Associates namespace DivisorChain @@ -311,8 +311,8 @@ theorem emultiplicity_prime_le_emultiplicity_image_by_factor_orderIso {m p : Ass · simp [hn] by_cases hm : m = 0 · simp [hm] at hp - rw [(finite_prime_left (prime_of_normalized_factor p hp) hm).emultiplicity_eq_multiplicity, - ← pow_dvd_iff_le_emultiplicity] + rw [FiniteMultiplicity.of_prime_left (prime_of_normalized_factor p hp) hm + |>.emultiplicity_eq_multiplicity, ← pow_dvd_iff_le_emultiplicity] apply pow_image_of_prime_by_factor_orderIso_dvd hn hp d (pow_multiplicity_dvd ..) theorem emultiplicity_prime_eq_emultiplicity_image_by_factor_orderIso {m p : Associates M} diff --git a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean index c7ac28b39f0ee..dafc8270e1007 100644 --- a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean +++ b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean @@ -57,8 +57,8 @@ noncomputable instance TensorProduct.instCoalgebra : Coalgebra R (M ⊗[R] N) := map_comp_comul := by rw [CoalgebraCat.ofComonObjCoalgebraStruct_comul] simp [-Mon_.monMonoidalStruct_tensorObj_X, - ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom, - ModuleCat.comp_def, ModuleCat.of, ModuleCat.asHom, + ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom_hom, + ModuleCat.hom_comp, ModuleCat.of, ModuleCat.ofHom, ModuleCat.MonoidalCategory.tensorμ_eq_tensorTensorTensorComm] } end diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 2533459ec4771..5925c3017eae3 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -911,7 +911,7 @@ theorem irreducible_pow_sup_of_ge (hI : I ≠ ⊥) (hJ : Irreducible J) (n : ℕ classical rw [irreducible_pow_sup hI hJ, min_eq_left] · congr - rw [← Nat.cast_inj (R := ℕ∞), ← multiplicity.Finite.emultiplicity_eq_multiplicity, + rw [← Nat.cast_inj (R := ℕ∞), ← FiniteMultiplicity.emultiplicity_eq_multiplicity, emultiplicity_eq_count_normalizedFactors hJ hI, normalize_eq J] rw [← emultiplicity_lt_top] apply hn.trans_lt @@ -1317,7 +1317,7 @@ end ChineseRemainder section PID -open multiplicity UniqueFactorizationMonoid Ideal +open UniqueFactorizationMonoid Ideal variable {R} variable [IsDomain R] [IsPrincipalIdealRing R] @@ -1358,18 +1358,18 @@ theorem singleton_span_mem_normalizedFactors_of_mem_normalizedFactors [Normaliza theorem emultiplicity_eq_emultiplicity_span {a b : R} : emultiplicity (Ideal.span {a}) (Ideal.span ({b} : Set R)) = emultiplicity a b := by - by_cases h : Finite a b + by_cases h : FiniteMultiplicity a b · rw [h.emultiplicity_eq_multiplicity] apply emultiplicity_eq_of_dvd_of_not_dvd <;> rw [Ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd] · exact pow_multiplicity_dvd a b · apply h.not_pow_dvd_of_multiplicity_lt apply lt_add_one - · suffices ¬Finite (Ideal.span ({a} : Set R)) (Ideal.span ({b} : Set R)) by + · suffices ¬FiniteMultiplicity (Ideal.span ({a} : Set R)) (Ideal.span ({b} : Set R)) by rw [emultiplicity_eq_top.2 h, emultiplicity_eq_top.2 this] - exact Finite.not_iff_forall.mpr fun n => by + exact FiniteMultiplicity.not_iff_forall.mpr fun n => by rw [Ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd] - exact Finite.not_iff_forall.mp h n + exact FiniteMultiplicity.not_iff_forall.mp h n section NormalizationMonoid variable [NormalizationMonoid R] diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index b5c1195243656..22c54ebc36736 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -112,25 +112,11 @@ variable (A K) /-- Send a set of `x`s in a finite extension `L` of the fraction field of `R` to `(y : R) • x ∈ integralClosure R L`. -/ theorem exists_integral_multiples (s : Finset L) : - ∃ y ≠ (0 : A), ∀ x ∈ s, IsIntegral A (y • x) := by - haveI := Classical.decEq L - refine s.induction ?_ ?_ - · use 1, one_ne_zero - rintro x ⟨⟩ - · rintro x s hx ⟨y, hy, hs⟩ - have := exists_integral_multiple - ((IsFractionRing.isAlgebraic_iff A K L).mpr (.of_finite _ x)) - ((injective_iff_map_eq_zero (algebraMap A L)).mp ?_) - · rcases this with ⟨x', y', hy', hx'⟩ - refine ⟨y * y', mul_ne_zero hy hy', fun x'' hx'' => ?_⟩ - rcases Finset.mem_insert.mp hx'' with (rfl | hx'') - · rw [mul_smul, Algebra.smul_def, Algebra.smul_def, hx'] - exact isIntegral_algebraMap.mul x'.2 - · rw [mul_comm, mul_smul, Algebra.smul_def] - exact isIntegral_algebraMap.mul (hs _ hx'') - · rw [IsScalarTower.algebraMap_eq A K L] - apply (algebraMap K L).injective.comp - exact IsFractionRing.injective _ _ + ∃ y ≠ (0 : A), ∀ x ∈ s, IsIntegral A (y • x) := + have := IsLocalization.isAlgebraic K (nonZeroDivisors A) + have := Algebra.IsAlgebraic.trans' A (algebraMap K L).injective + Algebra.IsAlgebraic.exists_integral_multiples (IsScalarTower.algebraMap_eq A K L ▸ + (algebraMap K L).injective.comp (IsFractionRing.injective _ _)) _ variable (L) diff --git a/Mathlib/RingTheory/Filtration.lean b/Mathlib/RingTheory/Filtration.lean index 49f369df496ee..ea599e0713283 100644 --- a/Mathlib/RingTheory/Filtration.lean +++ b/Mathlib/RingTheory/Filtration.lean @@ -40,10 +40,7 @@ This file contains the definitions and basic results around (stable) `I`-filtrat -/ - -universe u v - -variable {R M : Type u} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) +variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) open Polynomial @@ -52,7 +49,7 @@ open scoped Polynomial /-- An `I`-filtration on the module `M` is a sequence of decreasing submodules `N i` such that `I • (N i) ≤ N (i + 1)`. Note that we do not require the filtration to start from `⊤`. -/ @[ext] -structure Ideal.Filtration (M : Type u) [AddCommGroup M] [Module R M] where +structure Ideal.Filtration (M : Type*) [AddCommGroup M] [Module R M] where N : ℕ → Submodule R M mono : ∀ i, N (i + 1) ≤ N i smul_le : ∀ i, I • N i ≤ N (i + 1) diff --git a/Mathlib/RingTheory/Flat/Algebra.lean b/Mathlib/RingTheory/Flat/Algebra.lean deleted file mode 100644 index bf379a8806436..0000000000000 --- a/Mathlib/RingTheory/Flat/Algebra.lean +++ /dev/null @@ -1,85 +0,0 @@ -/- -Copyright (c) 2024 Christian Merten. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Christian Merten --/ -import Mathlib.RingTheory.Flat.Stability -import Mathlib.LinearAlgebra.TensorProduct.Tower - -/-! -# Flat algebras and flat ring homomorphisms - -We define flat algebras and flat ring homomorphisms. - -## Main definitions - -* `Algebra.Flat`: an `R`-algebra `S` is flat if it is flat as `R`-module -* `RingHom.Flat`: a ring homomorphism `f : R → S` is flat, if it makes `S` into a flat `R`-algebra - --/ - -universe u v w t - -open Function (Injective Surjective) - -open LinearMap (lsmul rTensor lTensor) - -open TensorProduct - -/-- An `R`-algebra `S` is flat if it is flat as `R`-module. -/ -class Algebra.Flat (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] : Prop where - out : Module.Flat R S - -namespace Algebra.Flat - -attribute [instance] out - -instance self (R : Type u) [CommRing R] : Algebra.Flat R R where - out := Module.Flat.self R - -variable (R : Type u) (S : Type v) [CommRing R] [CommRing S] - -/-- If `T` is a flat `S`-algebra and `S` is a flat `R`-algebra, -then `T` is a flat `R`-algebra. -/ -theorem trans (T : Type w) [CommRing T] [Algebra R S] [Algebra R T] [Algebra S T] - [IsScalarTower R S T] [Algebra.Flat R S] [Algebra.Flat S T] : Algebra.Flat R T where - out := Module.Flat.trans R S T - -@[deprecated (since := "2024-11-08")] alias comp := trans - -/-- If `S` is a flat `R`-algebra and `T` is any `R`-algebra, -then `T ⊗[R] S` is a flat `T`-algebra. -/ -instance baseChange (T : Type w) [CommRing T] [Algebra R S] [Algebra R T] [Algebra.Flat R S] : - Algebra.Flat T (T ⊗[R] S) where - out := Module.Flat.baseChange R T S - -/-- A base change of a flat algebra is flat. -/ -theorem isBaseChange [Algebra R S] (R' : Type w) (S' : Type t) [CommRing R'] [CommRing S'] - [Algebra R R'] [Algebra S S'] [Algebra R' S'] [Algebra R S'] [IsScalarTower R R' S'] - [IsScalarTower R S S'] [h : IsPushout R S R' S'] [Algebra.Flat R R'] : Algebra.Flat S S' where - out := Module.Flat.isBaseChange R S R' S' h.out - -end Algebra.Flat - -/-- A ring homomorphism `f : R →+* S` is flat if `S` is flat as an `R` algebra. -/ -@[algebraize RingHom.Flat.out] -class RingHom.Flat {R : Type u} {S : Type v} [CommRing R] [CommRing S] (f : R →+* S) : Prop where - out : f.toAlgebra.Flat := by infer_instance - -namespace RingHom.Flat - -attribute [instance] out - -/-- The identity of a ring is flat. -/ -instance identity (R : Type u) [CommRing R] : RingHom.Flat (1 : R →+* R) where - -variable {R : Type u} {S : Type v} {T : Type w} [CommRing R] [CommRing S] [CommRing T] - (f : R →+* S) (g : S →+* T) - -/-- Composition of flat ring homomorphisms is flat. -/ -instance comp [RingHom.Flat f] [RingHom.Flat g] : RingHom.Flat (g.comp f) where - out := by - algebraize_only [f, g, g.comp f] - exact Algebra.Flat.trans R S T - -end RingHom.Flat diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 9509f2f734fbb..a36eb5785365c 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -423,3 +423,28 @@ theorem tensorProduct_mapIncl_injective_of_left end Flat end Module + +section Injective + +open scoped TensorProduct + +variable {R S A B : Type*} [CommRing R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] + [CommSemiring S] [Algebra S A] [SMulCommClass R S A] + +namespace Algebra.TensorProduct + +theorem includeLeft_injective [Module.Flat R A] (hb : Function.Injective (algebraMap R B)) : + Function.Injective (includeLeft : A →ₐ[S] A ⊗[R] B) := by + convert Module.Flat.lTensor_preserves_injective_linearMap (M := A) (Algebra.linearMap R B) hb + |>.comp (_root_.TensorProduct.rid R A).symm.injective + ext; simp + +theorem includeRight_injective [Module.Flat R B] (ha : Function.Injective (algebraMap R A)) : + Function.Injective (includeRight : B →ₐ[R] A ⊗[R] B) := by + convert Module.Flat.rTensor_preserves_injective_linearMap (M := B) (Algebra.linearMap R A) ha + |>.comp (_root_.TensorProduct.lid R B).symm.injective + ext; simp + +end Algebra.TensorProduct + +end Injective diff --git a/Mathlib/RingTheory/Flat/CategoryTheory.lean b/Mathlib/RingTheory/Flat/CategoryTheory.lean index ea6d1442dad38..e8aaed01f1d4d 100644 --- a/Mathlib/RingTheory/Flat/CategoryTheory.lean +++ b/Mathlib/RingTheory/Flat/CategoryTheory.lean @@ -52,8 +52,8 @@ lemma iff_lTensor_preserves_shortComplex_exact : ⟨fun _ _ ↦ lTensor_shortComplex_exact _ _, fun H ↦ iff_lTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| - H (.mk (ModuleCat.asHom f) (ModuleCat.asHom g) - (DFunLike.ext _ _ h.apply_apply_eq_zero)) + H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) + (ModuleCat.hom_ext (DFunLike.ext _ _ h.apply_apply_eq_zero))) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ lemma iff_rTensor_preserves_shortComplex_exact : @@ -62,8 +62,8 @@ lemma iff_rTensor_preserves_shortComplex_exact : ⟨fun _ _ ↦ rTensor_shortComplex_exact _ _, fun H ↦ iff_rTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ moduleCat_exact_iff_function_exact _ |>.1 <| - H (.mk (ModuleCat.asHom f) (ModuleCat.asHom g) - (DFunLike.ext _ _ h.apply_apply_eq_zero)) + H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) + (ModuleCat.hom_ext (DFunLike.ext _ _ h.apply_apply_eq_zero))) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ end Module.Flat diff --git a/Mathlib/RingTheory/Flat/Localization.lean b/Mathlib/RingTheory/Flat/Localization.lean index 9269fb9b8bd69..2de407505a926 100644 --- a/Mathlib/RingTheory/Flat/Localization.lean +++ b/Mathlib/RingTheory/Flat/Localization.lean @@ -3,21 +3,30 @@ Copyright (c) 2024 Junyan Xu. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ -import Mathlib.Algebra.Module.LocalizedModule.Submodule -import Mathlib.RingTheory.Flat.Basic -import Mathlib.RingTheory.Localization.BaseChange +import Mathlib.RingTheory.Flat.Stability +import Mathlib.RingTheory.LocalProperties.Exactness /-! # Flatness and localization -In this file we show that localizations are flat, and flatness is a local property (TODO). +In this file we show that localizations are flat, and flatness is a local property. ## Main result -- `IsLocalization.flat`: a localization of a commutative ring is flat over it. +* `IsLocalization.flat`: a localization of a commutative ring is flat over it. +* `Module.flat_iff_of_isLocalization` : Let `Rₚ` a localization of a commutative ring `R` + and `M` be a module over `Rₚ`. Then `M` is flat over `R` if and only if `M` is flat over `Rₚ`. +* `Module.flat_of_isLocalized_maximal` : Let `M` be a module over a commutative ring `R`. + If the localization of `M` at each maximal ideal `P` is flat over `Rₚ`, then `M` is flat over `R`. +* `Module.flat_of_isLocalized_span` : Let `M` be a module over a commutative ring `R` + and `S` be a set that spans `R`. If the localization of `M` at each `s : S` is flat + over `Localization.Away s`, then `M` is flat over `R`. -/ +open IsLocalizedModule LocalizedModule LinearMap TensorProduct + variable {R : Type*} (S : Type*) [CommRing R] [CommRing S] [Algebra R S] variable (p : Submonoid R) [IsLocalization p S] +variable (M : Type*) [AddCommGroup M] [Module R M] [Module S M] [IsScalarTower R S M] include p in theorem IsLocalization.flat : Module.Flat R S := @@ -30,3 +39,66 @@ theorem IsLocalization.flat : Module.Flat R S := simpa [this, - Subtype.val_injective] using Subtype.val_injective instance Localization.flat : Module.Flat R (Localization p) := IsLocalization.flat _ p + +namespace Module + +include p in +theorem flat_iff_of_isLocalization : Flat S M ↔ Flat R M := + have := isLocalizedModule_id p M S + have := IsLocalization.flat S p + ⟨fun _ ↦ .trans R S M, fun _ ↦ .of_isLocalizedModule S p .id⟩ + +private lemma aux (I : Ideal R) (s : Submonoid R) : + have hM := isBaseChange s (Localization s) (mkLinearMap s M) + have hIM := isBaseChange s (Localization s) (mkLinearMap s (I ⊗[R] M)) + let e := (hIM.equiv.restrictScalars R).symm ≪≫ₗ + leftComm _ _ _ _ ≪≫ₗ (hM.equiv.restrictScalars R).lTensor I + LocalizedModule.map s (TensorProduct.lift <| lsmul R M ∘ₗ I.subtype) = + TensorProduct.lift (lsmul R (LocalizedModule s M) ∘ₗ I.subtype) ∘ₗ e.toLinearMap := by + refine linearMap_ext s (mkLinearMap s _) (mkLinearMap s _) ?_ + ext m i + rw [AlgebraTensorModule.curry_apply, curry_apply, restrictScalars_apply, LinearMap.comp_apply, + restrictScalars_apply, mkLinearMap_apply, AlgebraTensorModule.curry_apply, curry_apply, + restrictScalars_apply] + simpa [-mkLinearMap_apply, IsBaseChange.equiv_symm_apply, IsBaseChange.equiv_tmul] using + (mkLinearMap_apply _ _ _).symm.trans (map_smul (mkLinearMap s M) _ _) + +variable (Mₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommGroup (Mₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Mₚ P)] + (f : ∀ (P : Ideal R) [P.IsMaximal], M →ₗ[R] Mₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + +theorem flat_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Flat R (LocalizedModule J.primeCompl M)) : + Flat R M := + (flat_iff _ _).mpr fun I fg ↦ injective_of_localized_maximal _ fun J hJ ↦ by + rw [← LinearMap.coe_restrictScalars R, aux] + simpa using (flat_iff _ _).mp (h J) fg + +include f in +theorem flat_of_isLocalized_maixmal (H : ∀ (P : Ideal R) [P.IsMaximal], Flat R (Mₚ P)) : + Module.Flat R M := + flat_of_localized_maximal M fun P _ ↦ .of_linearEquiv _ _ _ (iso _ (f P)) + +variable (s : Set R) (spn : Ideal.span s = ⊤) + (Mₛ : ∀ _ : s, Type*) + [∀ r : s, AddCommGroup (Mₛ r)] + [∀ r : s, Module R (Mₛ r)] + (g : ∀ r : s, M →ₗ[R] Mₛ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (g r)] +include spn + +theorem flat_of_localized_span + (h : ∀ r : s, Flat R (LocalizedModule (.powers r.1) M)) : + Flat R M := + (Module.flat_iff _ _).mpr fun I fg ↦ injective_of_localized_span s spn _ fun r ↦ by + rw [← LinearMap.coe_restrictScalars R, aux] + simpa using (Module.flat_iff _ _).mp (h r) fg + +include g in +theorem flat_of_isLocalized_span (H : ∀ r : s, Module.Flat R (Mₛ r)) : + Module.Flat R M := + flat_of_localized_span M s spn fun r ↦ .of_linearEquiv _ _ _ (iso _ (g r)) + +end Module diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 82a5c54c0a72c..1b14cfce13ed0 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -363,7 +363,7 @@ instance hasPow : Pow (HomogeneousLocalization 𝒜 x) ℕ where instance : Add (HomogeneousLocalization 𝒜 x) where add := - Quotient.map₂' (· + ·) + Quotient.map₂ (· + ·) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) c3 c4 (h' : Localization.mk _ _ = Localization.mk _ _) => by change Localization.mk _ _ = Localization.mk _ _ @@ -376,7 +376,7 @@ instance : Sub (HomogeneousLocalization 𝒜 x) where sub z1 z2 := z1 + -z2 instance : Mul (HomogeneousLocalization 𝒜 x) where mul := - Quotient.map₂' (· * ·) + Quotient.map₂ (· * ·) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) c3 c4 (h' : Localization.mk _ _ = Localization.mk _ _) => by change Localization.mk _ _ = Localization.mk _ _ diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index c9ea82db6087b..41f0ce1b8f513 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -267,7 +267,22 @@ end AddCommGroup section SMul -variable [PartialOrder Γ] [PartialOrder Γ'] [AddCommMonoid R] [AddCommMonoid V] [SMulWithZero R V] +variable [PartialOrder Γ] [PartialOrder Γ'] [AddCommMonoid V] + +instance [Zero R] [SMulWithZero R V] : SMul R (SummableFamily Γ' V β) := + ⟨fun r t => + { toFun := r • t + isPWO_iUnion_support' := t.isPWO_iUnion_support.mono (Set.iUnion_mono fun i => + Pi.smul_apply r t i ▸ Function.support_const_smul_subset r _) + finite_co_support' := by + intro g + refine (t.finite_co_support g).subset ?_ + intro i hi + simp only [Pi.smul_apply, smul_coeff, ne_eq, Set.mem_setOf_eq] at hi + simp only [Function.mem_support, ne_eq] + exact right_ne_zero_of_smul hi } ⟩ + +variable [AddCommMonoid R] [SMulWithZero R V] theorem smul_support_subset_prod (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (gh : Γ × Γ') : @@ -322,13 +337,15 @@ theorem finite_co_support_prod_smul (s : SummableFamily Γ R α) /-- An elementwise scalar multiplication of one summable family on another. -/ @[simps] -def FamilySMul (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : +def smul (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : (SummableFamily Γ' V (α × β)) where toFun ab := (of R).symm (s (ab.1) • ((of R) (t (ab.2)))) isPWO_iUnion_support' := isPWO_iUnion_support_prod_smul s.isPWO_iUnion_support t.isPWO_iUnion_support finite_co_support' g := finite_co_support_prod_smul s t g +@[deprecated (since := "2024-11-17")] noncomputable alias FamilySMul := smul + theorem sum_vAddAntidiagonal_eq (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') (a : α × β) : ∑ x ∈ VAddAntidiagonal (s a.1).isPWO_support' (t a.2).isPWO_support' g, (s a.1).coeff x.1 • @@ -341,12 +358,12 @@ theorem sum_vAddAntidiagonal_eq (s : SummableFamily Γ R α) (t : SummableFamily · exact smul_eq_zero_of_left hs ((t a.2).coeff gh.2) · simp_all -theorem family_smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] +theorem smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') : - (FamilySMul s t).hsum.coeff g = ∑ gh ∈ VAddAntidiagonal s.isPWO_iUnion_support + (smul s t).hsum.coeff g = ∑ gh ∈ VAddAntidiagonal s.isPWO_iUnion_support t.isPWO_iUnion_support g, (s.hsum.coeff gh.1) • (t.hsum.coeff gh.2) := by rw [hsum_coeff] - simp only [hsum_coeff_eq_sum, FamilySMul_toFun, HahnModule.smul_coeff, Equiv.symm_apply_apply] + simp only [hsum_coeff_eq_sum, smul_toFun, HahnModule.smul_coeff, Equiv.symm_apply_apply] simp_rw [sum_vAddAntidiagonal_eq, Finset.smul_sum, Finset.sum_smul] rw [← sum_finsum_comm _ _ <| fun gh _ => smul_support_finite s t gh] refine sum_congr rfl fun gh _ => ?_ @@ -358,11 +375,13 @@ theorem family_smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] Set.mem_setOf_eq, Prod.forall, coeff_support, mem_product] exact hsupp ab.1 ab.2 hab -theorem hsum_family_smul {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] +@[deprecated (since := "2024-11-17")] alias family_smul_coeff := smul_coeff + +theorem smul_hsum {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : - (FamilySMul s t).hsum = (of R).symm (s.hsum • (of R) (t.hsum)) := by + (smul s t).hsum = (of R).symm (s.hsum • (of R) (t.hsum)) := by ext g - rw [family_smul_coeff s t g, HahnModule.smul_coeff, Equiv.symm_apply_apply] + rw [smul_coeff s t g, HahnModule.smul_coeff, Equiv.symm_apply_apply] refine Eq.symm (sum_of_injOn (fun a ↦ a) (fun _ _ _ _ h ↦ h) (fun _ hgh => ?_) (fun gh _ hgh => ?_) fun _ _ => by simp) · simp_all only [mem_coe, mem_vaddAntidiagonal, mem_support, ne_eq, Set.mem_iUnion, and_true] @@ -377,11 +396,13 @@ theorem hsum_family_smul {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] · exact smul_eq_zero_of_left h (t.hsum.coeff gh.2) · simp_all +@[deprecated (since := "2024-11-17")] alias hsum_family_smul := smul_hsum + instance [AddCommMonoid R] [SMulWithZero R V] : SMul (HahnSeries Γ R) (SummableFamily Γ' V β) where - smul x t := Equiv (Equiv.punitProd β) <| FamilySMul (single x) t + smul x t := Equiv (Equiv.punitProd β) <| smul (single x) t theorem smul_eq {x : HahnSeries Γ R} {t : SummableFamily Γ' V β} : - x • t = Equiv (Equiv.punitProd β) (FamilySMul (single x) t) := + x • t = Equiv (Equiv.punitProd β) (smul (single x) t) := rfl @[simp] @@ -393,7 +414,7 @@ theorem smul_apply {x : HahnSeries Γ R} {s : SummableFamily Γ' V α} {a : α} theorem hsum_smul_module {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] {x : HahnSeries Γ R} {s : SummableFamily Γ' V α} : (x • s).hsum = (of R).symm (x • of R s.hsum) := by - rw [smul_eq, hsum_equiv, hsum_family_smul, hsum_single] + rw [smul_eq, hsum_equiv, smul_hsum, hsum_single] end SMul @@ -427,6 +448,40 @@ theorem hsum_sub {R : Type*} [Ring R] {s t : SummableFamily Γ R α} : (s - t).hsum = s.hsum - t.hsum := by rw [← lsum_apply, LinearMap.map_sub, lsum_apply, lsum_apply] +theorem isPWO_iUnion_support_prod_mul {s : α → HahnSeries Γ R} {t : β → HahnSeries Γ R} + (hs : (⋃ a, (s a).support).IsPWO) (ht : (⋃ b, (t b).support).IsPWO) : + (⋃ (a : α × β), ((fun a ↦ ((s a.1) * (t a.2))) a).support).IsPWO := + isPWO_iUnion_support_prod_smul hs ht + +theorem finite_co_support_prod_mul (s : SummableFamily Γ R α) + (t : SummableFamily Γ R β) (g : Γ) : + Finite {(a : α × β) | ((fun (a : α × β) ↦ (s a.1 * t a.2)) a).coeff g ≠ 0} := + finite_co_support_prod_smul s t g + +/-- A summable family given by pointwise multiplication of a pair of summable families. -/ +@[simps] +def mul (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) : + (SummableFamily Γ R (α × β)) where + toFun a := s (a.1) * t (a.2) + isPWO_iUnion_support' := + isPWO_iUnion_support_prod_mul s.isPWO_iUnion_support t.isPWO_iUnion_support + finite_co_support' g := finite_co_support_prod_mul s t g + +theorem mul_eq_smul {β : Type*} (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) : + mul s t = smul s t := + rfl + +theorem mul_coeff {β : Type*} (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) (g : Γ) : + (mul s t).hsum.coeff g = ∑ gh ∈ addAntidiagonal s.isPWO_iUnion_support + t.isPWO_iUnion_support g, (s.hsum.coeff gh.1) * (t.hsum.coeff gh.2) := by + simp_rw [← smul_eq_mul, mul_eq_smul] + exact smul_coeff s t g + +theorem hsum_mul {β : Type*} (s : SummableFamily Γ R α) (t : SummableFamily Γ R β) : + (mul s t).hsum = s.hsum * t.hsum := by + rw [← smul_eq_mul, mul_eq_smul] + exact smul_hsum s t + end Semiring section OfFinsupp diff --git a/Mathlib/RingTheory/Ideal/Cotangent.lean b/Mathlib/RingTheory/Ideal/Cotangent.lean index 04fd59250aa5c..9f2e89438bbbc 100644 --- a/Mathlib/RingTheory/Ideal/Cotangent.lean +++ b/Mathlib/RingTheory/Ideal/Cotangent.lean @@ -196,6 +196,7 @@ def mapCotangent (I₁ : Ideal A) (I₂ : Ideal B) (f : A →ₐ[R] B) (h : I₁ ((I₂ • ⊤ : Submodule B I₂).restrictScalars R) ?_ ?_ · exact f.toLinearMap.restrict (p := I₁.restrictScalars R) (q := I₂.restrictScalars R) h · intro x hx + rw [Submodule.restrictScalars_mem] at hx refine Submodule.smul_induction_on hx ?_ (fun _ _ ↦ add_mem) rintro a ha ⟨b, hb⟩ - simp only [SetLike.mk_smul_mk, smul_eq_mul, Submodule.mem_comap, Submodule.restrictScalars_mem] diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index d627344ccb869..a04061874ab5e 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -56,116 +56,28 @@ section Semiring variable [Semiring R] [AddCommMonoid M] [Module R M] -instance : SMul (Ideal R) (Submodule R M) where - smul I N := - { __ := I.toAddSubmonoid • N.toAddSubmonoid - smul_mem' := fun r m hm ↦ AddSubmonoid.smul_induction_on hm - (fun m hm n ↦ by rw [smul_smul]; exact AddSubmonoid.smul_mem_smul <| I.smul_mem _ hm) - fun m₁ m₂ h₁ h₂ ↦ by rw [smul_add]; exact (I.1 • N.1).add_mem h₁ h₂ } - /-- This duplicates the global `smul_eq_mul`, but doesn't have to unfold anywhere near as much to apply. -/ protected theorem _root_.Ideal.smul_eq_mul (I J : Ideal R) : I • J = I * J := rfl -variable {I J : Ideal R} {N P : Submodule R M} - -theorem smul_toAddSubmonoid : (I • N).toAddSubmonoid = I.toAddSubmonoid • N.toAddSubmonoid := rfl - -theorem smul_mem_smul {r} {n} (hr : r ∈ I) (hn : n ∈ N) : r • n ∈ I • N := - AddSubmonoid.smul_mem_smul hr hn - -theorem smul_le : I • N ≤ P ↔ ∀ r ∈ I, ∀ n ∈ N, r • n ∈ P := - AddSubmonoid.smul_le - -@[simp, norm_cast] -lemma coe_set_smul : (I : Set R) • N = I • N := - set_smul_eq_of_le _ _ _ - (fun _ _ hr hx ↦ smul_mem_smul hr hx) - (smul_le.mpr fun _ hr _ hx ↦ mem_set_smul_of_mem_mem hr hx) - -@[elab_as_elim] -theorem smul_induction_on {p : M → Prop} {x} (H : x ∈ I • N) (smul : ∀ r ∈ I, ∀ n ∈ N, p (r • n)) - (add : ∀ x y, p x → p y → p (x + y)) : p x := - AddSubmonoid.smul_induction_on H smul add - -/-- Dependent version of `Submodule.smul_induction_on`. -/ -@[elab_as_elim] -theorem smul_induction_on' {x : M} (hx : x ∈ I • N) {p : ∀ x, x ∈ I • N → Prop} - (smul : ∀ (r : R) (hr : r ∈ I) (n : M) (hn : n ∈ N), p (r • n) (smul_mem_smul hr hn)) - (add : ∀ x hx y hy, p x hx → p y hy → p (x + y) (add_mem ‹_› ‹_›)) : p x hx := by - refine Exists.elim ?_ fun (h : x ∈ I • N) (H : p x h) ↦ H - exact smul_induction_on hx (fun a ha x hx ↦ ⟨_, smul _ ha _ hx⟩) - fun x y ⟨_, hx⟩ ⟨_, hy⟩ ↦ ⟨_, add _ _ _ _ hx hy⟩ +variable {I : Ideal R} {N : Submodule R M} theorem smul_le_right : I • N ≤ N := smul_le.2 fun r _ _ ↦ N.smul_mem r -theorem smul_mono (hij : I ≤ J) (hnp : N ≤ P) : I • N ≤ J • P := - AddSubmonoid.smul_le_smul hij hnp - -theorem smul_mono_left (h : I ≤ J) : I • N ≤ J • N := - smul_mono h le_rfl - -instance : CovariantClass (Ideal R) (Submodule R M) HSMul.hSMul LE.le := - ⟨fun _ _ => smul_mono le_rfl⟩ - -@[deprecated smul_mono_right (since := "2024-03-31")] -protected theorem smul_mono_right (h : N ≤ P) : I • N ≤ I • P := - _root_.smul_mono_right I h - theorem map_le_smul_top (I : Ideal R) (f : R →ₗ[R] M) : Submodule.map f I ≤ I • (⊤ : Submodule R M) := by rintro _ ⟨y, hy, rfl⟩ rw [← mul_one y, ← smul_eq_mul, f.map_smul] exact smul_mem_smul hy mem_top -variable (I J N P) - -@[simp] -theorem smul_bot : I • (⊥ : Submodule R M) = ⊥ := - toAddSubmonoid_injective <| AddSubmonoid.addSubmonoid_smul_bot _ - -@[simp] -theorem bot_smul : (⊥ : Ideal R) • N = ⊥ := - le_bot_iff.mp <| smul_le.mpr <| by rintro _ rfl _ _; rw [zero_smul]; exact zero_mem _ +variable (I N) @[simp] theorem top_smul : (⊤ : Ideal R) • N = N := le_antisymm smul_le_right fun r hri => one_smul R r ▸ smul_mem_smul mem_top hri -theorem smul_sup : I • (N ⊔ P) = I • N ⊔ I • P := - toAddSubmonoid_injective <| by - simp only [smul_toAddSubmonoid, sup_toAddSubmonoid, AddSubmonoid.addSubmonoid_smul_sup] - -theorem sup_smul : (I ⊔ J) • N = I • N ⊔ J • N := - le_antisymm (smul_le.mpr fun mn hmn p hp ↦ by - obtain ⟨m, hm, n, hn, rfl⟩ := mem_sup.mp hmn - rw [add_smul]; exact add_mem_sup (smul_mem_smul hm hp) <| smul_mem_smul hn hp) - (sup_le (smul_mono_left le_sup_left) <| smul_mono_left le_sup_right) - -protected theorem smul_assoc : (I • J) • N = I • J • N := - le_antisymm - (smul_le.2 fun _ hrsij t htn ↦ smul_induction_on hrsij - (fun r hr s hs ↦ smul_assoc r s t ▸ smul_mem_smul hr (smul_mem_smul hs htn)) - fun x y ↦ (add_smul x y t).symm ▸ add_mem) - (smul_le.2 fun r hr _ hsn ↦ smul_induction_on hsn - (fun j hj n hn ↦ (smul_assoc r j n).symm ▸ smul_mem_smul (smul_mem_smul hr hj) hn) - fun m₁ m₂ ↦ (smul_add r m₁ m₂) ▸ add_mem) - -@[deprecated smul_inf_le (since := "2024-03-31")] -protected theorem smul_inf_le (M₁ M₂ : Submodule R M) : - I • (M₁ ⊓ M₂) ≤ I • M₁ ⊓ I • M₂ := smul_inf_le _ _ _ - -theorem smul_iSup {ι : Sort*} {I : Ideal R} {t : ι → Submodule R M} : I • iSup t = ⨆ i, I • t i := - toAddSubmonoid_injective <| by - simp only [smul_toAddSubmonoid, iSup_toAddSubmonoid, AddSubmonoid.smul_iSup] - -@[deprecated smul_iInf_le (since := "2024-03-31")] -protected theorem smul_iInf_le {ι : Sort*} {I : Ideal R} {t : ι → Submodule R M} : - I • iInf t ≤ ⨅ i, I • t i := - smul_iInf_le - theorem mem_of_span_top_of_smul_mem (M' : Submodule R M) (s : Set R) (hs : Ideal.span s = ⊤) (x : M) (H : ∀ r : s, (r : R) • x ∈ M') : x ∈ M' := by suffices LinearMap.range (LinearMap.toSpanSingleton R M x) ≤ M' by @@ -245,16 +157,10 @@ theorem ideal_span_singleton_smul (r : R) (N : Submodule R M) : submodule `M'` of `x`, we only need to show that `r ^ n • x ∈ M'` for some `n` for each `r : s`. -/ theorem mem_of_span_eq_top_of_smul_pow_mem (M' : Submodule R M) (s : Set R) (hs : Ideal.span s = ⊤) (x : M) (H : ∀ r : s, ∃ n : ℕ, ((r : R) ^ n : R) • x ∈ M') : x ∈ M' := by - obtain ⟨s', hs₁, hs₂⟩ := (Ideal.span_eq_top_iff_finite _).mp hs - replace H : ∀ r : s', ∃ n : ℕ, ((r : R) ^ n : R) • x ∈ M' := fun r => H ⟨_, hs₁ r.2⟩ - choose n₁ n₂ using H - let N := s'.attach.sup n₁ - have hs' := Ideal.span_pow_eq_top (s' : Set R) hs₂ N - apply M'.mem_of_span_top_of_smul_mem _ hs' + choose f hf using H + apply M'.mem_of_span_top_of_smul_mem _ (Ideal.span_range_pow_eq_top s hs f) rintro ⟨_, r, hr, rfl⟩ - convert M'.smul_mem (r ^ (N - n₁ ⟨r, hr⟩)) (n₂ ⟨r, hr⟩) using 1 - simp only [Subtype.coe_mk, smul_smul, ← pow_add] - rw [tsub_add_cancel_of_le (Finset.le_sup (s'.mem_attach _) : n₁ ⟨r, hr⟩ ≤ N)] + exact hf r open Pointwise in @[simp] diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index ebb045e9944ca..3442658846e0f 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -88,7 +88,7 @@ theorem injective_quotient_le_comap_map (P : Ideal R[X]) : R[x] / P → (R / (P ∩ R))[x] / (P / (P ∩ R)) ``` commutes. It is used, for instance, in the proof of `quotient_mk_comp_C_is_integral_of_jacobson`, -in the file `RingTheory.Jacobson`. +in the file `Mathlib.RingTheory.Jacobson.Polynomial`. -/ theorem quotient_mk_maps_eq (P : Ideal R[X]) : ((Quotient.mk (map (mapRingHom (Quotient.mk (P.comap (C : R →+* R[X])))) P)).comp C).comp diff --git a/Mathlib/RingTheory/JacobsonIdeal.lean b/Mathlib/RingTheory/Jacobson/Ideal.lean similarity index 86% rename from Mathlib/RingTheory/JacobsonIdeal.lean rename to Mathlib/RingTheory/Jacobson/Ideal.lean index 804cf08cea5f3..181f68bbf0591 100644 --- a/Mathlib/RingTheory/JacobsonIdeal.lean +++ b/Mathlib/RingTheory/Jacobson/Ideal.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Devon Tuma, Wojciech Nawrocki -/ import Mathlib.RingTheory.Ideal.IsPrimary -import Mathlib.RingTheory.Ideal.Quotient.Basic -import Mathlib.RingTheory.Polynomial.Quotient +import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.TwoSidedIdeal.Operations /-! @@ -52,8 +51,6 @@ namespace Ideal variable {R : Type u} {S : Type v} -open Polynomial - section Jacobson section Ring @@ -99,8 +96,7 @@ theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, z * y * x let ⟨p, hpi, q, hq, hpq⟩ := Submodule.mem_sup.1 ((eq_top_iff_one _).1 hxy) let ⟨r, hr⟩ := mem_span_singleton'.1 hq ⟨r, by - -- Porting note: supply `mul_add_one` with explicit variables - rw [mul_assoc, ← mul_add_one r (y * x), hr, ← hpq, ← neg_sub, add_sub_cancel_right] + rw [mul_assoc, ← mul_add_one, hr, ← hpq, ← neg_sub, add_sub_cancel_right] exact I.neg_mem hpi⟩) fun hxy : I ⊔ span {y * x + 1} ≠ ⊤ => let ⟨M, hm1, hm2⟩ := exists_le_maximal _ hxy suffices x ∉ M from (this <| mem_sInf.1 hx ⟨le_trans le_sup_left hm2, hm1⟩).elim @@ -111,8 +107,7 @@ theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, z * y * x let ⟨z, hz⟩ := hx (-y) hm.1.1 <| (eq_top_iff_one _).2 <| sub_sub_cancel (z * -y * x + z) 1 ▸ M.sub_mem (by - -- Porting note: supply `mul_add_one` with explicit variables - rw [mul_assoc, ← mul_add_one z, neg_mul, ← sub_eq_iff_eq_add.mpr df.symm, neg_sub, + rw [mul_assoc, ← mul_add_one, neg_mul, ← sub_eq_iff_eq_add.mpr df.symm, neg_sub, sub_add_cancel] exact M.mul_mem_left _ hi) <| him hz⟩ @@ -160,15 +155,15 @@ theorem eq_jacobson_iff_sInf_maximal' : /-- An ideal `I` equals its Jacobson radical if and only if every element outside `I` also lies outside of a maximal ideal containing `I`. -/ theorem eq_jacobson_iff_not_mem : - I.jacobson = I ↔ ∀ (x) (_ : x ∉ I), ∃ M : Ideal R, (I ≤ M ∧ M.IsMaximal) ∧ x ∉ M := by + I.jacobson = I ↔ ∀ x ∉ I, ∃ M : Ideal R, (I ≤ M ∧ M.IsMaximal) ∧ x ∉ M := by constructor · intro h x hx - erw [← h, mem_sInf] at hx + rw [← h, Ideal.jacobson, mem_sInf] at hx push_neg at hx exact hx · refine fun h => le_antisymm (fun x hx => ?_) le_jacobson contrapose hx - erw [mem_sInf] + rw [Ideal.jacobson, mem_sInf] push_neg exact h x hx @@ -338,45 +333,6 @@ end CommRing end Jacobson -section Polynomial - -open Polynomial - -variable [CommRing R] - -theorem jacobson_bot_polynomial_le_sInf_map_maximal : - jacobson (⊥ : Ideal R[X]) ≤ sInf (map (C : R →+* R[X]) '' { J : Ideal R | J.IsMaximal }) := by - refine le_sInf fun J => exists_imp.2 fun j hj => ?_ - haveI : j.IsMaximal := hj.1 - refine Trans.trans (jacobson_mono bot_le) (le_of_eq ?_ : J.jacobson ≤ J) - suffices t : (⊥ : Ideal (Polynomial (R ⧸ j))).jacobson = ⊥ by - rw [← hj.2, jacobson_eq_iff_jacobson_quotient_eq_bot] - replace t := congr_arg (map (polynomialQuotientEquivQuotientPolynomial j).toRingHom) t - rwa [map_jacobson_of_bijective _, map_bot] at t - exact RingEquiv.bijective (polynomialQuotientEquivQuotientPolynomial j) - refine eq_bot_iff.2 fun f hf => ?_ - have r1 : (X : (R ⧸ j)[X]) ≠ 0 := fun hX => by - replace hX := congr_arg (fun f => coeff f 1) hX - simp only [coeff_X_one, coeff_zero] at hX - exact zero_ne_one hX.symm - have r2 := eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit ((mem_jacobson_bot.1 hf) X)) - simp only [coeff_add, mul_coeff_zero, coeff_X_zero, mul_zero, coeff_one_zero, zero_add] at r2 - erw [add_left_eq_self] at r2 - simpa using (mul_eq_zero.mp r2).resolve_right r1 - -- Porting note: this is golfed to much - -- simpa [(fun hX => by simpa using congr_arg (fun f => coeff f 1) hX : (X : (R ⧸ j)[X]) ≠ 0)] - -- using eq_C_of_degree_eq_zero (degree_eq_zero_of_is_unit ((mem_jacobson_bot.1 hf) X)) - -theorem jacobson_bot_polynomial_of_jacobson_bot (h : jacobson (⊥ : Ideal R) = ⊥) : - jacobson (⊥ : Ideal R[X]) = ⊥ := by - refine eq_bot_iff.2 (le_trans jacobson_bot_polynomial_le_sInf_map_maximal ?_) - refine fun f hf => (Submodule.mem_bot R[X]).2 <| Polynomial.ext fun n => - Trans.trans (?_ : coeff f n = 0) (coeff_zero n).symm - suffices f.coeff n ∈ Ideal.jacobson ⊥ by rwa [h, Submodule.mem_bot] at this - exact mem_sInf.2 fun j hj => (mem_map_C_iff.1 ((mem_sInf.1 hf) ⟨j, ⟨hj.2, rfl⟩⟩)) n - -end Polynomial - section IsLocal variable [CommRing R] diff --git a/Mathlib/RingTheory/Jacobson/Polynomial.lean b/Mathlib/RingTheory/Jacobson/Polynomial.lean new file mode 100644 index 0000000000000..c102d0db7133b --- /dev/null +++ b/Mathlib/RingTheory/Jacobson/Polynomial.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2020 Devon Tuma. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Devon Tuma +-/ +import Mathlib.RingTheory.Jacobson.Ideal +import Mathlib.RingTheory.Polynomial.Quotient +/-! +# Jacobson radical of polynomial ring + +-/ + +namespace Ideal + +section Polynomial + +open Polynomial + +variable {R : Type*} [CommRing R] + +theorem jacobson_bot_polynomial_le_sInf_map_maximal : + jacobson (⊥ : Ideal R[X]) ≤ sInf (map (C : R →+* R[X]) '' { J : Ideal R | J.IsMaximal }) := by + refine le_sInf fun J => exists_imp.2 fun j hj => ?_ + haveI : j.IsMaximal := hj.1 + refine Trans.trans (jacobson_mono bot_le) (le_of_eq ?_ : J.jacobson ≤ J) + suffices t : (⊥ : Ideal (Polynomial (R ⧸ j))).jacobson = ⊥ by + rw [← hj.2, jacobson_eq_iff_jacobson_quotient_eq_bot] + replace t := congr_arg (map (polynomialQuotientEquivQuotientPolynomial j).toRingHom) t + rwa [map_jacobson_of_bijective _, map_bot] at t + exact RingEquiv.bijective (polynomialQuotientEquivQuotientPolynomial j) + refine eq_bot_iff.2 fun f hf => ?_ + have r1 : (X : (R ⧸ j)[X]) ≠ 0 := ne_of_apply_ne (coeff · 1) <| by simp + simpa [r1] using eq_C_of_degree_eq_zero (degree_eq_zero_of_isUnit ((mem_jacobson_bot.1 hf) X)) + +theorem jacobson_bot_polynomial_of_jacobson_bot (h : jacobson (⊥ : Ideal R) = ⊥) : + jacobson (⊥ : Ideal R[X]) = ⊥ := by + refine eq_bot_iff.2 (le_trans jacobson_bot_polynomial_le_sInf_map_maximal ?_) + refine fun f hf => (Submodule.mem_bot R[X]).2 <| Polynomial.ext fun n => + Trans.trans (?_ : coeff f n = 0) (coeff_zero n).symm + suffices f.coeff n ∈ Ideal.jacobson ⊥ by rwa [h, Submodule.mem_bot] at this + exact mem_sInf.2 fun j hj => (mem_map_C_iff.1 ((mem_sInf.1 hf) ⟨j, ⟨hj.2, rfl⟩⟩)) n + +end Polynomial + +end Ideal diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson/Ring.lean similarity index 99% rename from Mathlib/RingTheory/Jacobson.lean rename to Mathlib/RingTheory/Jacobson/Ring.lean index 5fa8e8a0b7c09..a8165786b21e0 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson/Ring.lean @@ -5,7 +5,7 @@ Authors: Devon Tuma -/ import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.RingTheory.Ideal.Over -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Polynomial import Mathlib.RingTheory.Artinian /-! diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index 96de1da9e85a4..d668bcecc861a 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -259,7 +259,7 @@ def Derivation.liftKaehlerDifferential (D : Derivation R S M) : Ω[S⁄R] →ₗ · exact D.tensorProductTo.comp ((KaehlerDifferential.ideal R S).subtype.restrictScalars S) · intro x hx rw [LinearMap.mem_ker] - refine Submodule.smul_induction_on hx ?_ ?_ + refine Submodule.smul_induction_on ((Submodule.restrictScalars_mem _ _ _).mp hx) ?_ ?_ · rintro x hx y - rw [RingHom.mem_ker] at hx dsimp @@ -839,7 +839,8 @@ theorem KaehlerDifferential.range_kerCotangentToTensor constructor · rintro ⟨x, rfl⟩ obtain ⟨x, rfl⟩ := Ideal.toCotangent_surjective _ x - simp [kerCotangentToTensor_toCotangent, RingHom.mem_ker.mp x.2] + simp only [kerCotangentToTensor_toCotangent, Submodule.restrictScalars_mem, LinearMap.mem_ker, + mapBaseChange_tmul, map_D, RingHom.mem_ker.mp x.2, map_zero, smul_zero] · intro hx obtain ⟨x, rfl⟩ := LinearMap.rTensor_surjective (Ω[A⁄R]) (g := Algebra.linearMap A B) h x obtain ⟨x, rfl⟩ := (TensorProduct.lid _ _).symm.surjective x @@ -869,6 +870,8 @@ theorem KaehlerDifferential.range_kerCotangentToTensor simp [RingHom.mem_ker, ha, this.2] · simp only [map_sum, LinearMapClass.map_smul, kerCotangentToTensor_toCotangent, map_sub] simp_rw [← TensorProduct.tmul_smul] + -- was `simp [kerCotangentToTensor_toCotangent, RingHom.mem_ker.mp x.2]` and very slow + -- (https://github.com/leanprover-community/mathlib4/issues/19751) simp only [smul_sub, TensorProduct.tmul_sub, Finset.sum_sub_distrib, ← TensorProduct.tmul_sum, ← Finset.sum_smul, Finset.sum_attach, sub_eq_self, Finset.sum_attach (f := fun i ↦ x i • KaehlerDifferential.D R A i)] diff --git a/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean b/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean index 2a707385da691..1919f202cbf8f 100644 --- a/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/MaximalIdeal/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Chris Hughes, Mario Carneiro -/ -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal import Mathlib.RingTheory.LocalRing.MaximalIdeal.Defs import Mathlib.RingTheory.Localization.Basic import Mathlib.RingTheory.Nilpotent.Lemmas diff --git a/Mathlib/RingTheory/LocalRing/Subring.lean b/Mathlib/RingTheory/LocalRing/Subring.lean index bafa2eb662812..90d074bc764e8 100644 --- a/Mathlib/RingTheory/LocalRing/Subring.lean +++ b/Mathlib/RingTheory/LocalRing/Subring.lean @@ -18,9 +18,6 @@ import Mathlib.RingTheory.Localization.AtPrime open IsLocalRing Set -open scoped Polynomial - - variable {R S : Type*} [CommRing R] [CommRing S] variable {K : Type*} [Field K] diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 137b20d13c55e..309f147722deb 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -223,6 +223,21 @@ such that `z = f x * (f y)⁻¹`. -/ noncomputable def lift (hg : Injective g) : K →+* L := IsLocalization.lift fun y : nonZeroDivisors A => isUnit_map_of_injective hg y +theorem lift_unique (hg : Function.Injective g) {f : K →+* L} + (hf1 : ∀ x, f (algebraMap A K x) = g x) : IsFractionRing.lift hg = f := + IsLocalization.lift_unique _ hf1 + +/-- Another version of unique to give two lift maps should be equal -/ +theorem ringHom_ext {f1 f2 : K →+* L} + (hf : ∀ x : A, f1 (algebraMap A K x) = f2 (algebraMap A K x)) : f1 = f2 := by + ext z + obtain ⟨x, y, hy, rfl⟩ := IsFractionRing.div_surjective (A := A) z + rw [map_div₀, map_div₀, hf, hf] + +theorem injective_comp_algebraMap : + Function.Injective fun (f : K →+* L) => f.comp (algebraMap A K) := + fun _ _ h => ringHom_ext (fun x => RingHom.congr_fun h x) + section liftAlgHom variable [Algebra R A] [Algebra R K] [IsScalarTower R A K] [Algebra R L] diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index 595fd2c074e5a..cb310c7e3e585 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -300,13 +300,11 @@ theorem isFractionRing_of_algebraic [Algebra.IsAlgebraic A L] mem_nonZeroDivisors_iff_ne_zero.mp hy ((injective_iff_map_eq_zero (algebraMap C L)).mp (algebraMap_injective C A L) _ h)) surj' := fun z => - let ⟨x, y, hy, hxy⟩ := exists_integral_multiple (Algebra.IsAlgebraic.isAlgebraic z) inj - ⟨⟨mk' C (x : L) x.2, algebraMap _ _ y, - mem_nonZeroDivisors_iff_ne_zero.mpr fun h => - hy (inj _ (by rw [IsScalarTower.algebraMap_apply A C L, h, RingHom.map_zero]))⟩, - by - simp only - rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply A C L, mul_comm, hxy]⟩ + let ⟨x, hx, int⟩ := (Algebra.IsAlgebraic.isAlgebraic z).exists_integral_multiple + ((injective_iff_map_eq_zero _).mpr inj) + ⟨⟨mk' C _ int, algebraMap _ _ x, mem_nonZeroDivisors_of_ne_zero fun h ↦ + hx (inj _ <| by rw [IsScalarTower.algebraMap_apply A C L, h, RingHom.map_zero])⟩, by + rw [algebraMap_mk', ← IsScalarTower.algebraMap_apply A C L, Algebra.smul_def, mul_comm]⟩ exists_of_eq := fun {x y} h => ⟨1, by simpa using algebraMap_injective C A L h⟩ } variable (K L) @@ -320,8 +318,7 @@ theorem isFractionRing_of_finite_extension [IsDomain A] [Algebra K L] [IsScalarT isFractionRing_of_algebraic A C fun _ hx => IsFractionRing.to_map_eq_zero_iff.mp - ((_root_.map_eq_zero <| - algebraMap K L).mp <| (IsScalarTower.algebraMap_apply _ _ _ _).symm.trans hx) + ((map_eq_zero <| algebraMap K L).mp <| (IsScalarTower.algebraMap_apply _ _ _ _).symm.trans hx) end IsIntegralClosure diff --git a/Mathlib/RingTheory/MaximalSpectrum.lean b/Mathlib/RingTheory/MaximalSpectrum.lean index b9b3ccfdbd9db..959eae0a6633e 100644 --- a/Mathlib/RingTheory/MaximalSpectrum.lean +++ b/Mathlib/RingTheory/MaximalSpectrum.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ -import Mathlib.RingTheory.Ideal.Colon import Mathlib.RingTheory.Localization.AsSubring import Mathlib.RingTheory.PrimeSpectrum @@ -70,23 +69,12 @@ theorem iInf_localization_eq_bot : (⨅ v : MaximalSpectrum R, constructor · contrapose intro hrange hlocal - let denom : Ideal R := (Submodule.span R {1} : Submodule R K).colon (Submodule.span R {x}) - have hdenom : (1 : R) ∉ denom := by - intro hdenom - rcases Submodule.mem_span_singleton.mp - (Submodule.mem_colon.mp hdenom x <| Submodule.mem_span_singleton_self x) with ⟨y, hy⟩ - exact hrange ⟨y, by rw [← mul_one <| algebraMap R K y, ← Algebra.smul_def, hy, one_smul]⟩ - rcases denom.exists_le_maximal fun h => (h ▸ hdenom) Submodule.mem_top with ⟨max, hmax, hle⟩ + let denom : Ideal R := (1 : Submodule R K).comap (LinearMap.toSpanSingleton R K x) + have hdenom : (1 : R) ∉ denom := by simpa [denom] using hrange + rcases denom.exists_le_maximal (denom.ne_top_iff_one.mpr hdenom) with ⟨max, hmax, hle⟩ rcases hlocal ⟨max, hmax⟩ with ⟨n, d, hd, rfl⟩ - apply hd (hle <| Submodule.mem_colon.mpr fun _ hy => _) - intro _ hy - rcases Submodule.mem_span_singleton.mp hy with ⟨y, rfl⟩ - exact Submodule.mem_span_singleton.mpr ⟨y * n, by - rw [Algebra.smul_def, mul_one, map_mul, smul_comm, Algebra.smul_def, Algebra.smul_def, - mul_comm <| algebraMap R K d, - inv_mul_cancel_right₀ <| - (map_ne_zero_iff _ <| NoZeroSMulDivisors.algebraMap_injective R K).mpr fun h => - (h ▸ hd) max.zero_mem]⟩ + exact hd (hle ⟨n, by simp [denom, Algebra.smul_def, mul_left_comm, mul_inv_cancel₀ <| + (map_ne_zero_iff _ <| IsFractionRing.injective R K).mpr fun h ↦ hd (h ▸ max.zero_mem :)]⟩) · rintro ⟨y, rfl⟩ ⟨v, hv⟩ exact ⟨y, 1, v.ne_top_iff_one.mp hv.ne_top, by rw [map_one, inv_one, mul_one]⟩ @@ -101,13 +89,7 @@ variable [IsDomain R] (K : Type v) [Field K] [Algebra R K] [IsFractionRing R K] viewed as subalgebras of its field of fractions. -/ theorem iInf_localization_eq_bot : ⨅ v : PrimeSpectrum R, Localization.subalgebra.ofField K _ (v.asIdeal.primeCompl_le_nonZeroDivisors) = ⊥ := by - ext x - rw [Algebra.mem_iInf] - constructor - · rw [← MaximalSpectrum.iInf_localization_eq_bot, Algebra.mem_iInf] - exact fun hx ⟨v, hv⟩ => hx ⟨v, hv.isPrime⟩ - · rw [Algebra.mem_bot] - rintro ⟨y, rfl⟩ ⟨v, hv⟩ - exact ⟨y, 1, v.ne_top_iff_one.mp hv.ne_top, by rw [map_one, inv_one, mul_one]⟩ + refine bot_unique (.trans (fun _ ↦ ?_) (MaximalSpectrum.iInf_localization_eq_bot R K).le) + simpa only [Algebra.mem_iInf] using fun hx ⟨v, hv⟩ ↦ hx ⟨v, hv.isPrime⟩ end PrimeSpectrum diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index f6528327fbc78..19c092dd8f1ed 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -6,8 +6,8 @@ Authors: Robert Y. Lewis, Chris Hughes, Daniel Weber import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int.Defs import Mathlib.Data.ENat.Basic -import Mathlib.Tactic.Linarith /-! # Multiplicity of a divisor @@ -22,7 +22,7 @@ several basic results on it. `n`. * `multiplicity a b`: a `ℕ`-valued version of `multiplicity`, defaulting for `1` instead of `⊤`. The reason for using `1` as a default value instead of `0` is to have `multiplicity_eq_zero_iff`. -* `multiplicity.Finite a b`: a predicate denoting that the multiplicity of `a` in `b` is finite. +* `FiniteMultiplicity a b`: a predicate denoting that the multiplicity of `a` in `b` is finite. -/ @@ -31,47 +31,56 @@ variable {α β : Type*} open Nat /-- `multiplicity.Finite a b` indicates that the multiplicity of `a` in `b` is finite. -/ -abbrev multiplicity.Finite [Monoid α] (a b : α) : Prop := +abbrev FiniteMultiplicity [Monoid α] (a b : α) : Prop := ∃ n : ℕ, ¬a ^ (n + 1) ∣ b +@[deprecated (since := "2024-11-30")] alias multiplicity.Finite := FiniteMultiplicity + open scoped Classical in /-- `emultiplicity a b` returns the largest natural number `n` such that `a ^ n ∣ b`, as an `ℕ∞`. If `∀ n, a ^ n ∣ b` then it returns `⊤`. -/ noncomputable def emultiplicity [Monoid α] (a b : α) : ℕ∞ := - if h : multiplicity.Finite a b then Nat.find h else ⊤ + if h : FiniteMultiplicity a b then Nat.find h else ⊤ /-- A `ℕ`-valued version of `emultiplicity`, returning `1` instead of `⊤`. -/ noncomputable def multiplicity [Monoid α] (a b : α) : ℕ := (emultiplicity a b).untop' 1 -open multiplicity - section Monoid variable [Monoid α] [Monoid β] {a b : α} @[simp] theorem emultiplicity_eq_top : - emultiplicity a b = ⊤ ↔ ¬Finite a b := by + emultiplicity a b = ⊤ ↔ ¬FiniteMultiplicity a b := by simp [emultiplicity] -theorem emultiplicity_lt_top {a b : α} : emultiplicity a b < ⊤ ↔ Finite a b := by +theorem emultiplicity_lt_top {a b : α} : emultiplicity a b < ⊤ ↔ FiniteMultiplicity a b := by simp [lt_top_iff_ne_top, emultiplicity_eq_top] -theorem finite_iff_emultiplicity_ne_top : - Finite a b ↔ emultiplicity a b ≠ ⊤ := by simp +theorem finiteMultiplicity_iff_emultiplicity_ne_top : + FiniteMultiplicity a b ↔ emultiplicity a b ≠ ⊤ := by simp + +@[deprecated (since := "2024-11-30")] +alias finite_iff_emultiplicity_ne_top := finiteMultiplicity_iff_emultiplicity_ne_top -alias ⟨multiplicity.Finite.emultiplicity_ne_top, _⟩ := finite_iff_emultiplicity_ne_top +alias ⟨FiniteMultiplicity.emultiplicity_ne_top, _⟩ := finite_iff_emultiplicity_ne_top + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_ne_top := FiniteMultiplicity.emultiplicity_ne_top @[deprecated (since := "2024-11-08")] -alias Finite.emultiplicity_ne_top := multiplicity.Finite.emultiplicity_ne_top +alias Finite.emultiplicity_ne_top := FiniteMultiplicity.emultiplicity_ne_top -theorem finite_of_emultiplicity_eq_natCast {n : ℕ} (h : emultiplicity a b = n) : - Finite a b := by +theorem finiteMultiplicity_of_emultiplicity_eq_natCast {n : ℕ} (h : emultiplicity a b = n) : + FiniteMultiplicity a b := by by_contra! nh rw [← emultiplicity_eq_top, h] at nh trivial +@[deprecated (since := "2024-11-30")] +alias finite_of_emultiplicity_eq_natCast := finiteMultiplicity_of_emultiplicity_eq_natCast + theorem multiplicity_eq_of_emultiplicity_eq_some {n : ℕ} (h : emultiplicity a b = n) : multiplicity a b = n := by simp [multiplicity, h] @@ -81,16 +90,24 @@ theorem emultiplicity_ne_of_multiplicity_ne {n : ℕ} : multiplicity a b ≠ n → emultiplicity a b ≠ n := mt multiplicity_eq_of_emultiplicity_eq_some -theorem multiplicity.Finite.emultiplicity_eq_multiplicity (h : Finite a b) : +theorem FiniteMultiplicity.emultiplicity_eq_multiplicity (h : FiniteMultiplicity a b) : emultiplicity a b = multiplicity a b := by cases hm : emultiplicity a b · simp [h] at hm rw [multiplicity_eq_of_emultiplicity_eq_some hm] -theorem multiplicity.Finite.emultiplicity_eq_iff_multiplicity_eq {n : ℕ} (h : Finite a b) : - emultiplicity a b = n ↔ multiplicity a b = n := by +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_eq_multiplicity := + FiniteMultiplicity.emultiplicity_eq_multiplicity + +theorem FiniteMultiplicity.emultiplicity_eq_iff_multiplicity_eq {n : ℕ} + (h : FiniteMultiplicity a b) : emultiplicity a b = n ↔ multiplicity a b = n := by simp [h.emultiplicity_eq_multiplicity] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_eq_iff_multiplicity_eq := + FiniteMultiplicity.emultiplicity_eq_iff_multiplicity_eq + theorem emultiplicity_eq_iff_multiplicity_eq_of_ne_one {n : ℕ} (h : n ≠ 1) : emultiplicity a b = n ↔ multiplicity a b = n := by constructor @@ -103,14 +120,18 @@ theorem emultiplicity_eq_zero_iff_multiplicity_eq_zero : emultiplicity_eq_iff_multiplicity_eq_of_ne_one zero_ne_one @[simp] -theorem multiplicity_eq_one_of_not_finite (h : ¬Finite a b) : +theorem multiplicity_eq_one_of_not_finiteMultiplicity (h : ¬FiniteMultiplicity a b) : multiplicity a b = 1 := by simp [multiplicity, emultiplicity_eq_top.2 h] +@[deprecated (since := "2024-11-30")] +alias multiplicity_eq_one_of_not_finite := + multiplicity_eq_one_of_not_finiteMultiplicity + @[simp] theorem multiplicity_le_emultiplicity : multiplicity a b ≤ emultiplicity a b := by - by_cases hf : Finite a b + by_cases hf : FiniteMultiplicity a b · simp [hf.emultiplicity_eq_multiplicity] · simp [hf, emultiplicity_eq_top.2] @@ -124,52 +145,73 @@ theorem multiplicity_le_of_emultiplicity_le {n : ℕ} (h : emultiplicity a b ≤ multiplicity a b ≤ n := by exact_mod_cast multiplicity_le_emultiplicity.trans h -theorem multiplicity.Finite.emultiplicity_le_of_multiplicity_le (hfin : Finite a b) +theorem FiniteMultiplicity.emultiplicity_le_of_multiplicity_le (hfin : FiniteMultiplicity a b) {n : ℕ} (h : multiplicity a b ≤ n) : emultiplicity a b ≤ n := by rw [emultiplicity_eq_multiplicity hfin] assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_le_of_multiplicity_le := + FiniteMultiplicity.emultiplicity_le_of_multiplicity_le + theorem le_emultiplicity_of_le_multiplicity {n : ℕ} (h : n ≤ multiplicity a b) : n ≤ emultiplicity a b := by exact_mod_cast (WithTop.coe_mono h).trans multiplicity_le_emultiplicity -theorem multiplicity.Finite.le_multiplicity_of_le_emultiplicity (hfin : Finite a b) +theorem FiniteMultiplicity.le_multiplicity_of_le_emultiplicity (hfin : FiniteMultiplicity a b) {n : ℕ} (h : n ≤ emultiplicity a b) : n ≤ multiplicity a b := by rw [emultiplicity_eq_multiplicity hfin] at h assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.le_multiplicity_of_le_emultiplicity := + FiniteMultiplicity.le_multiplicity_of_le_emultiplicity + theorem multiplicity_lt_of_emultiplicity_lt {n : ℕ} (h : emultiplicity a b < n) : multiplicity a b < n := by exact_mod_cast multiplicity_le_emultiplicity.trans_lt h -theorem multiplicity.Finite.emultiplicity_lt_of_multiplicity_lt (hfin : Finite a b) +theorem FiniteMultiplicity.emultiplicity_lt_of_multiplicity_lt (hfin : FiniteMultiplicity a b) {n : ℕ} (h : multiplicity a b < n) : emultiplicity a b < n := by rw [emultiplicity_eq_multiplicity hfin] assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_lt_of_multiplicity_lt := + FiniteMultiplicity.emultiplicity_lt_of_multiplicity_lt + theorem lt_emultiplicity_of_lt_multiplicity {n : ℕ} (h : n < multiplicity a b) : n < emultiplicity a b := by exact_mod_cast (WithTop.coe_strictMono h).trans_le multiplicity_le_emultiplicity -theorem multiplicity.Finite.lt_multiplicity_of_lt_emultiplicity (hfin : Finite a b) +theorem FiniteMultiplicity.lt_multiplicity_of_lt_emultiplicity (hfin : FiniteMultiplicity a b) {n : ℕ} (h : n < emultiplicity a b) : n < multiplicity a b := by rw [emultiplicity_eq_multiplicity hfin] at h assumption_mod_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.lt_multiplicity_of_lt_emultiplicity := + FiniteMultiplicity.lt_multiplicity_of_lt_emultiplicity + theorem emultiplicity_pos_iff : 0 < emultiplicity a b ↔ 0 < multiplicity a b := by simp [pos_iff_ne_zero, pos_iff_ne_zero, emultiplicity_eq_zero_iff_multiplicity_eq_zero] -theorem multiplicity.Finite.def : Finite a b ↔ ∃ n : ℕ, ¬a ^ (n + 1) ∣ b := +theorem FiniteMultiplicity.def : FiniteMultiplicity a b ↔ ∃ n : ℕ, ¬a ^ (n + 1) ∣ b := Iff.rfl -theorem multiplicity.Finite.not_dvd_of_one_right : Finite a 1 → ¬a ∣ 1 := fun ⟨n, hn⟩ ⟨d, hd⟩ => - hn ⟨d ^ (n + 1), (pow_mul_pow_eq_one (n + 1) hd.symm).symm⟩ +@[deprecated (since := "2024-11-30")] alias multiplicity.Finite.def := FiniteMultiplicity.def + +theorem FiniteMultiplicity.not_dvd_of_one_right : FiniteMultiplicity a 1 → ¬a ∣ 1 := + fun ⟨n, hn⟩ ⟨d, hd⟩ => hn ⟨d ^ (n + 1), (pow_mul_pow_eq_one (n + 1) hd.symm).symm⟩ + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_dvd_of_one_right := FiniteMultiplicity.not_dvd_of_one_right @[norm_cast] theorem Int.natCast_emultiplicity (a b : ℕ) : emultiplicity (a : ℤ) (b : ℤ) = emultiplicity a b := by - unfold emultiplicity multiplicity.Finite + unfold emultiplicity FiniteMultiplicity congr! <;> norm_cast @[norm_cast] @@ -178,22 +220,32 @@ theorem Int.natCast_multiplicity (a b : ℕ) : multiplicity (a : ℤ) (b : ℤ) @[deprecated (since := "2024-04-05")] alias Int.coe_nat_multiplicity := Int.natCast_multiplicity -theorem multiplicity.Finite.not_iff_forall : ¬Finite a b ↔ ∀ n : ℕ, a ^ n ∣ b := +theorem FiniteMultiplicity.not_iff_forall : ¬FiniteMultiplicity a b ↔ ∀ n : ℕ, a ^ n ∣ b := ⟨fun h n => Nat.casesOn n (by rw [_root_.pow_zero] exact one_dvd _) - (by simpa [multiplicity.Finite, Classical.not_not] using h), - by simp [multiplicity.Finite, multiplicity, Classical.not_not]; tauto⟩ + (by simpa [FiniteMultiplicity] using h), + by simp [FiniteMultiplicity, multiplicity]; tauto⟩ -theorem multiplicity.Finite.not_unit (h : Finite a b) : ¬IsUnit a := +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_iff_forall := FiniteMultiplicity.not_iff_forall + +theorem FiniteMultiplicity.not_unit (h : FiniteMultiplicity a b) : ¬IsUnit a := let ⟨n, hn⟩ := h hn ∘ IsUnit.dvd ∘ IsUnit.pow (n + 1) -theorem multiplicity.Finite.mul_left {c : α} : Finite a (b * c) → Finite a b := fun ⟨n, hn⟩ => +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_unit := FiniteMultiplicity.not_unit + +theorem FiniteMultiplicity.mul_left {c : α} : + FiniteMultiplicity a (b * c) → FiniteMultiplicity a b := fun ⟨n, hn⟩ => ⟨n, fun h => hn (h.trans (dvd_mul_right _ _))⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.mul_left := FiniteMultiplicity.mul_left + theorem pow_dvd_of_le_emultiplicity {k : ℕ} (hk : k ≤ emultiplicity a b) : a ^ k ∣ b := by classical cases k @@ -202,7 +254,7 @@ theorem pow_dvd_of_le_emultiplicity {k : ℕ} (hk : k ≤ emultiplicity a b) : split at hk · norm_cast at hk simpa using (Nat.find_min _ (lt_of_succ_le hk)) - · apply Finite.not_iff_forall.mp ‹_› + · apply FiniteMultiplicity.not_iff_forall.mp ‹_› theorem pow_dvd_of_le_multiplicity {k : ℕ} (hk : k ≤ multiplicity a b) : a ^ k ∣ b := pow_dvd_of_le_emultiplicity (le_emultiplicity_of_le_multiplicity hk) @@ -220,15 +272,19 @@ theorem not_pow_dvd_of_emultiplicity_lt {m : ℕ} (hm : emultiplicity a b < m) : exact hn2 ((pow_dvd_pow _ hn1).trans nh) · simp at hm -theorem multiplicity.Finite.not_pow_dvd_of_multiplicity_lt (hf : Finite a b) {m : ℕ} +theorem FiniteMultiplicity.not_pow_dvd_of_multiplicity_lt (hf : FiniteMultiplicity a b) {m : ℕ} (hm : multiplicity a b < m) : ¬a ^ m ∣ b := by apply not_pow_dvd_of_emultiplicity_lt rw [hf.emultiplicity_eq_multiplicity] norm_cast +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_pow_dvd_of_multiplicity_lt := + FiniteMultiplicity.not_pow_dvd_of_multiplicity_lt + theorem multiplicity_pos_of_dvd (hdiv : a ∣ b) : 0 < multiplicity a b := by refine zero_lt_iff.2 fun h => ?_ - simpa [hdiv] using Finite.not_pow_dvd_of_multiplicity_lt + simpa [hdiv] using FiniteMultiplicity.not_pow_dvd_of_multiplicity_lt (by by_contra! nh; simp [nh] at h) (lt_one_iff.mpr h) theorem emultiplicity_pos_of_dvd (hdiv : a ∣ b) : 0 < emultiplicity a b := @@ -236,7 +292,7 @@ theorem emultiplicity_pos_of_dvd (hdiv : a ∣ b) : 0 < emultiplicity a b := theorem emultiplicity_eq_of_dvd_of_not_dvd {k : ℕ} (hk : a ^ k ∣ b) (hsucc : ¬a ^ (k + 1) ∣ b) : emultiplicity a b = k := by classical - have : Finite a b := ⟨k, hsucc⟩ + have : FiniteMultiplicity a b := ⟨k, hsucc⟩ simp only [emultiplicity, this, ↓reduceDIte, Nat.cast_inj, find_eq_iff, hsucc, not_false_eq_true, Decidable.not_not, true_and] exact fun n hn ↦ (pow_dvd_pow _ hn).trans hk @@ -249,24 +305,36 @@ theorem le_emultiplicity_of_pow_dvd {k : ℕ} (hk : a ^ k ∣ b) : k ≤ emultiplicity a b := le_of_not_gt fun hk' => not_pow_dvd_of_emultiplicity_lt hk' hk -theorem multiplicity.Finite.le_multiplicity_of_pow_dvd (hf : Finite a b) {k : ℕ} (hk : a ^ k ∣ b) : - k ≤ multiplicity a b := +theorem FiniteMultiplicity.le_multiplicity_of_pow_dvd (hf : FiniteMultiplicity a b) + {k : ℕ} (hk : a ^ k ∣ b) : k ≤ multiplicity a b := hf.le_multiplicity_of_le_emultiplicity (le_emultiplicity_of_pow_dvd hk) +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.le_multiplicity_of_pow_dvd := + FiniteMultiplicity.le_multiplicity_of_pow_dvd + theorem pow_dvd_iff_le_emultiplicity {k : ℕ} : a ^ k ∣ b ↔ k ≤ emultiplicity a b := ⟨le_emultiplicity_of_pow_dvd, pow_dvd_of_le_emultiplicity⟩ -theorem multiplicity.Finite.pow_dvd_iff_le_multiplicity (hf : Finite a b) {k : ℕ} : +theorem FiniteMultiplicity.pow_dvd_iff_le_multiplicity (hf : FiniteMultiplicity a b) {k : ℕ} : a ^ k ∣ b ↔ k ≤ multiplicity a b := by exact_mod_cast hf.emultiplicity_eq_multiplicity ▸ pow_dvd_iff_le_emultiplicity +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.pow_dvd_iff_le_multiplicity := + FiniteMultiplicity.pow_dvd_iff_le_multiplicity + theorem emultiplicity_lt_iff_not_dvd {k : ℕ} : emultiplicity a b < k ↔ ¬a ^ k ∣ b := by rw [pow_dvd_iff_le_emultiplicity, not_le] -theorem multiplicity.Finite.multiplicity_lt_iff_not_dvd {k : ℕ} (hf : Finite a b) : +theorem FiniteMultiplicity.multiplicity_lt_iff_not_dvd {k : ℕ} (hf : FiniteMultiplicity a b) : multiplicity a b < k ↔ ¬a ^ k ∣ b := by rw [hf.pow_dvd_iff_le_multiplicity, not_le] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_lt_iff_not_dvd := + FiniteMultiplicity.multiplicity_lt_iff_not_dvd + theorem emultiplicity_eq_coe {n : ℕ} : emultiplicity a b = n ↔ a ^ n ∣ b ∧ ¬a ^ (n + 1) ∣ b := by constructor @@ -281,33 +349,48 @@ theorem emultiplicity_eq_coe {n : ℕ} : · rw [and_imp] apply emultiplicity_eq_of_dvd_of_not_dvd -theorem multiplicity.Finite.multiplicity_eq_iff (hf : Finite a b) {n : ℕ} : +theorem FiniteMultiplicity.multiplicity_eq_iff (hf : FiniteMultiplicity a b) {n : ℕ} : multiplicity a b = n ↔ a ^ n ∣ b ∧ ¬a ^ (n + 1) ∣ b := by simp [← emultiplicity_eq_coe, hf.emultiplicity_eq_multiplicity] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_eq_iff := FiniteMultiplicity.multiplicity_eq_iff + @[simp] -theorem multiplicity.Finite.not_of_isUnit_left (b : α) (ha : IsUnit a) : ¬Finite a b := +theorem FiniteMultiplicity.not_of_isUnit_left (b : α) (ha : IsUnit a) : ¬FiniteMultiplicity a b := (·.not_unit ha) -theorem multiplicity.Finite.not_of_one_left (b : α) : ¬ Finite 1 b := by simp +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_of_isUnit_left := FiniteMultiplicity.not_of_isUnit_left + +theorem FiniteMultiplicity.not_of_one_left (b : α) : ¬ FiniteMultiplicity 1 b := by simp + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_of_one_left := FiniteMultiplicity.not_of_one_left @[simp] theorem emultiplicity_one_left (b : α) : emultiplicity 1 b = ⊤ := - emultiplicity_eq_top.2 (Finite.not_of_one_left _) + emultiplicity_eq_top.2 (FiniteMultiplicity.not_of_one_left _) @[simp] -theorem multiplicity.Finite.one_right (ha : Finite a 1) : multiplicity a 1 = 0 := by +theorem FiniteMultiplicity.one_right (ha : FiniteMultiplicity a 1) : multiplicity a 1 = 0 := by simp [ha.multiplicity_eq_iff, ha.not_dvd_of_one_right] -theorem multiplicity.Finite.not_of_unit_left (a : α) (u : αˣ) : ¬ Finite (u : α) a := - Finite.not_of_isUnit_left a u.isUnit +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.one_right := FiniteMultiplicity.one_right + +theorem FiniteMultiplicity.not_of_unit_left (a : α) (u : αˣ) : ¬ FiniteMultiplicity (u : α) a := + FiniteMultiplicity.not_of_isUnit_left a u.isUnit + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.not_of_unit_left := FiniteMultiplicity.not_of_unit_left theorem emultiplicity_eq_zero : emultiplicity a b = 0 ↔ ¬a ∣ b := by - by_cases hf : Finite a b + by_cases hf : FiniteMultiplicity a b · rw [← ENat.coe_zero, emultiplicity_eq_coe] simp - · simpa [emultiplicity_eq_top.2 hf] using Finite.not_iff_forall.1 hf 1 + · simpa [emultiplicity_eq_top.2 hf] using FiniteMultiplicity.not_iff_forall.1 hf 1 theorem multiplicity_eq_zero : multiplicity a b = 0 ↔ ¬a ∣ b := @@ -321,7 +404,7 @@ theorem multiplicity_ne_zero : multiplicity a b ≠ 0 ↔ a ∣ b := by simp [multiplicity_eq_zero] -theorem multiplicity.Finite.exists_eq_pow_mul_and_not_dvd (hfin : Finite a b) : +theorem FiniteMultiplicity.exists_eq_pow_mul_and_not_dvd (hfin : FiniteMultiplicity a b) : ∃ c : α, b = a ^ multiplicity a b * c ∧ ¬a ∣ c := by obtain ⟨c, hc⟩ := pow_multiplicity_dvd a b refine ⟨c, hc, ?_⟩ @@ -330,6 +413,10 @@ theorem multiplicity.Finite.exists_eq_pow_mul_and_not_dvd (hfin : Finite a b) : have h₁ : a ^ (multiplicity a b + 1) ∣ b := ⟨k, hc⟩ exact (hfin.multiplicity_eq_iff.1 (by simp)).2 h₁ +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.exists_eq_pow_mul_and_not_dvd := + FiniteMultiplicity.exists_eq_pow_mul_and_not_dvd + theorem emultiplicity_le_emultiplicity_iff {c d : β} : emultiplicity a b ≤ emultiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b → c ^ n ∣ d := by classical constructor @@ -349,12 +436,17 @@ theorem emultiplicity_le_emultiplicity_iff {c d : β} : simp_all only [not_exists, Decidable.not_not, not_true_eq_false, top_le_iff, dite_eq_right_iff, ENat.coe_ne_top, imp_false, not_false_eq_true, implies_true] -theorem multiplicity.Finite.multiplicity_le_multiplicity_iff {c d : β} (hab : Finite a b) - (hcd : Finite c d) : multiplicity a b ≤ multiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b → c ^ n ∣ d := by +theorem FiniteMultiplicity.multiplicity_le_multiplicity_iff {c d : β} (hab : FiniteMultiplicity a b) + (hcd : FiniteMultiplicity c d) : + multiplicity a b ≤ multiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b → c ^ n ∣ d := by rw [← WithTop.coe_le_coe, ENat.some_eq_coe, ← hab.emultiplicity_eq_multiplicity, ← hcd.emultiplicity_eq_multiplicity] apply emultiplicity_le_emultiplicity_iff +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_le_multiplicity_iff := + FiniteMultiplicity.multiplicity_le_multiplicity_iff + theorem emultiplicity_eq_emultiplicity_iff {c d : β} : emultiplicity a b = emultiplicity c d ↔ ∀ n : ℕ, a ^ n ∣ b ↔ c ^ n ∣ d := ⟨fun h n => @@ -400,8 +492,8 @@ theorem dvd_iff_multiplicity_pos {a b : α} : 0 < multiplicity a b ↔ a ∣ b : theorem dvd_iff_emultiplicity_pos {a b : α} : 0 < emultiplicity a b ↔ a ∣ b := emultiplicity_pos_iff.trans dvd_iff_multiplicity_pos -theorem Nat.multiplicity_finite_iff {a b : ℕ} : Finite a b ↔ a ≠ 1 ∧ 0 < b := by - rw [← not_iff_not, Finite.not_iff_forall, not_and_or, Ne, Classical.not_not, not_lt, +theorem Nat.finiteMultiplicity_iff {a b : ℕ} : FiniteMultiplicity a b ↔ a ≠ 1 ∧ 0 < b := by + rw [← not_iff_not, FiniteMultiplicity.not_iff_forall, not_and_or, not_ne_iff, not_lt, Nat.le_zero] exact ⟨fun h => @@ -417,6 +509,9 @@ theorem Nat.multiplicity_finite_iff {a b : ℕ} : Finite a b ↔ a ≠ 1 ∧ 0 < not_lt_of_ge (le_of_dvd (Nat.pos_of_ne_zero hb) (h b)) (b.lt_pow_self ha_gt_one), fun h => by cases h <;> simp [*]⟩ +@[deprecated (since := "2024-11-30")] +alias Nat.multiplicity_finite_iff := Nat.finiteMultiplicity_iff + alias ⟨_, Dvd.multiplicity_pos⟩ := dvd_iff_multiplicity_pos end Monoid @@ -425,8 +520,11 @@ section CommMonoid variable [CommMonoid α] -theorem multiplicity.Finite.mul_right {a b c : α} (hf : Finite a (b * c)) : Finite a c := - (mul_comm b c ▸ hf).mul_left +theorem FiniteMultiplicity.mul_right {a b c : α} (hf : FiniteMultiplicity a (b * c)) : + FiniteMultiplicity a c := (mul_comm b c ▸ hf).mul_left + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.mul_right := FiniteMultiplicity.mul_right theorem emultiplicity_of_isUnit_right {a b : α} (ha : ¬IsUnit a) (hb : IsUnit b) : emultiplicity a b = 0 := @@ -471,10 +569,13 @@ section MonoidWithZero variable [MonoidWithZero α] -theorem multiplicity.Finite.ne_zero {a b : α} (h : Finite a b) : b ≠ 0 := +theorem FiniteMultiplicity.ne_zero {a b : α} (h : FiniteMultiplicity a b) : b ≠ 0 := let ⟨n, hn⟩ := h fun hb => by simp [hb] at hn +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.ne_zero := FiniteMultiplicity.ne_zero + @[simp] theorem emultiplicity_zero (a : α) : emultiplicity a 0 = ⊤ := emultiplicity_eq_top.2 (fun v ↦ v.ne_zero rfl) @@ -493,12 +594,15 @@ section Semiring variable [Semiring α] -theorem multiplicity.Finite.or_of_add {p a b : α} (hf : Finite p (a + b)) : - Finite p a ∨ Finite p b := by +theorem FiniteMultiplicity.or_of_add {p a b : α} (hf : FiniteMultiplicity p (a + b)) : + FiniteMultiplicity p a ∨ FiniteMultiplicity p b := by by_contra! nh obtain ⟨c, hc⟩ := hf simp_all [dvd_add] +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.or_of_add := FiniteMultiplicity.or_of_add + theorem min_le_emultiplicity_add {p a b : α} : min (emultiplicity p a) (emultiplicity p b) ≤ emultiplicity p (a + b) := by cases hm : min (emultiplicity p a) (emultiplicity p b) @@ -516,12 +620,19 @@ section Ring variable [Ring α] @[simp] -theorem multiplicity.Finite.neg_iff {a b : α} : Finite a (-b) ↔ Finite a b := by - unfold Finite +theorem FiniteMultiplicity.neg_iff {a b : α} : + FiniteMultiplicity a (-b) ↔ FiniteMultiplicity a b := by + unfold FiniteMultiplicity congr! 3 simp only [dvd_neg] -alias ⟨_, multiplicity.Finite.neg⟩ := Finite.neg_iff +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.neg_iff := FiniteMultiplicity.neg_iff + +alias ⟨_, FiniteMultiplicity.neg⟩ := FiniteMultiplicity.neg_iff + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.neg := FiniteMultiplicity.neg @[simp] theorem emultiplicity_neg (a b : α) : emultiplicity a (-b) = emultiplicity a b := by @@ -544,7 +655,7 @@ theorem Int.multiplicity_natAbs (a : ℕ) (b : ℤ) : theorem emultiplicity_add_of_gt {p a b : α} (h : emultiplicity p b < emultiplicity p a) : emultiplicity p (a + b) = emultiplicity p b := by - have : Finite p b := finite_iff_emultiplicity_ne_top.2 (by simp [·] at h) + have : FiniteMultiplicity p b := finiteMultiplicity_iff_emultiplicity_ne_top.2 (by simp [·] at h) rw [this.emultiplicity_eq_multiplicity] at * apply emultiplicity_eq_of_dvd_of_not_dvd · apply dvd_add @@ -557,18 +668,21 @@ theorem emultiplicity_add_of_gt {p a b : α} (h : emultiplicity p b < emultiplic apply pow_dvd_of_le_emultiplicity exact Order.add_one_le_of_lt h -theorem multiplicity.Finite.multiplicity_add_of_gt {p a b : α} (hf : Finite p b) +theorem FiniteMultiplicity.multiplicity_add_of_gt {p a b : α} (hf : FiniteMultiplicity p b) (h : multiplicity p b < multiplicity p a) : multiplicity p (a + b) = multiplicity p b := multiplicity_eq_of_emultiplicity_eq <| emultiplicity_add_of_gt (hf.emultiplicity_eq_multiplicity ▸ (WithTop.coe_strictMono h).trans_le multiplicity_le_emultiplicity) +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_add_of_gt := FiniteMultiplicity.multiplicity_add_of_gt + theorem emultiplicity_sub_of_gt {p a b : α} (h : emultiplicity p b < emultiplicity p a) : emultiplicity p (a - b) = emultiplicity p b := by rw [sub_eq_add_neg, emultiplicity_add_of_gt] <;> rw [emultiplicity_neg]; assumption theorem multiplicity_sub_of_gt {p a b : α} (h : multiplicity p b < multiplicity p a) - (hfin : Finite p b) : multiplicity p (a - b) = multiplicity p b := by + (hfin : FiniteMultiplicity p b) : multiplicity p (a - b) = multiplicity p b := by rw [sub_eq_add_neg, hfin.neg.multiplicity_add_of_gt] <;> rw [multiplicity_neg]; assumption theorem emultiplicity_add_eq_min {p a b : α} @@ -581,8 +695,8 @@ theorem emultiplicity_add_eq_min {p a b : α} · rw [emultiplicity_add_of_gt hab, min_eq_right] exact le_of_lt hab -theorem multiplicity_add_eq_min {p a b : α} (ha : Finite p a) (hb : Finite p b) - (h : multiplicity p a ≠ multiplicity p b) : +theorem multiplicity_add_eq_min {p a b : α} (ha : FiniteMultiplicity p a) + (hb : FiniteMultiplicity p b) (h : multiplicity p a ≠ multiplicity p b) : multiplicity p (a + b) = min (multiplicity p a) (multiplicity p b) := by rcases lt_trichotomy (multiplicity p a) (multiplicity p b) with (hab | _ | hab) · rw [add_comm, ha.multiplicity_add_of_gt hab, min_eq_left] @@ -599,7 +713,7 @@ variable [CancelCommMonoidWithZero α] /- Porting note: Pulled a b intro parameters since Lean parses that more easily -/ -theorem multiplicity.finite_mul_aux {p : α} (hp : Prime p) {a b : α} : +theorem finiteMultiplicity_mul_aux {p : α} (hp : Prime p) {a b : α} : ∀ {n m : ℕ}, ¬p ^ (n + 1) ∣ a → ¬p ^ (m + 1) ∣ b → ¬p ^ (n + m + 1) ∣ a * b | n, m => fun ha hb ⟨s, hs⟩ => have : p ∣ a * b := ⟨p ^ (n + m) * s, by simp [hs, pow_add, mul_comm, mul_assoc, mul_left_comm]⟩ @@ -612,7 +726,7 @@ theorem multiplicity.finite_mul_aux {p : α} (hp : Prime p) {a b : α} : rw [tsub_add_cancel_of_le (succ_le_of_lt hn0)] at hy simp [hy, pow_add, mul_comm, mul_assoc, mul_left_comm]⟩) have : 1 ≤ n + m := le_trans hn0 (Nat.le_add_right n m) - finite_mul_aux hp hpx hb + finiteMultiplicity_mul_aux hp hpx hb ⟨s, mul_right_cancel₀ hp.1 (by rw [tsub_add_eq_add_tsub (succ_le_of_lt hn0), tsub_add_cancel_of_le this] simp_all [mul_comm, mul_assoc, mul_left_comm, pow_add])⟩) @@ -626,29 +740,40 @@ theorem multiplicity.finite_mul_aux {p : α} (hp : Prime p) {a b : α} : mul_right_cancel₀ hp.1 <| by rw [tsub_add_cancel_of_le (succ_le_of_lt hm0)] at hy simp [hy, pow_add, mul_comm, mul_assoc, mul_left_comm]⟩) - finite_mul_aux hp ha hpx + finiteMultiplicity_mul_aux hp ha hpx ⟨s, mul_right_cancel₀ hp.1 (by rw [add_assoc, tsub_add_cancel_of_le (succ_le_of_lt hm0)] simp_all [mul_comm, mul_assoc, mul_left_comm, pow_add])⟩ -theorem Prime.multiplicity_finite_mul {p a b : α} (hp : Prime p) : - Finite p a → Finite p b → Finite p (a * b) := - fun ⟨n, hn⟩ ⟨m, hm⟩ => ⟨n + m, finite_mul_aux hp hn hm⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity.finite_mul_aux := finiteMultiplicity_mul_aux -theorem multiplicity.Finite.mul_iff {p a b : α} (hp : Prime p) : - Finite p (a * b) ↔ Finite p a ∧ Finite p b := +theorem Prime.finiteMultiplicity_mul {p a b : α} (hp : Prime p) : + FiniteMultiplicity p a → FiniteMultiplicity p b → FiniteMultiplicity p (a * b) := + fun ⟨n, hn⟩ ⟨m, hm⟩ => ⟨n + m, finiteMultiplicity_mul_aux hp hn hm⟩ + +@[deprecated (since := "2024-11-30")] +alias Prime.multiplicity_finite_mul := Prime.finiteMultiplicity_mul + +theorem FiniteMultiplicity.mul_iff {p a b : α} (hp : Prime p) : + FiniteMultiplicity p (a * b) ↔ FiniteMultiplicity p a ∧ FiniteMultiplicity p b := ⟨fun h => ⟨h.mul_left, h.mul_right⟩, fun h => - hp.multiplicity_finite_mul h.1 h.2⟩ + hp.finiteMultiplicity_mul h.1 h.2⟩ + +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.mul_iff := FiniteMultiplicity.mul_iff -theorem multiplicity.Finite.pow {p a : α} (hp : Prime p) - (hfin : Finite p a) {k : ℕ} : Finite p (a ^ k) := +theorem FiniteMultiplicity.pow {p a : α} (hp : Prime p) + (hfin : FiniteMultiplicity p a) {k : ℕ} : FiniteMultiplicity p (a ^ k) := match k, hfin with | 0, _ => ⟨0, by simp [mt isUnit_iff_dvd_one.2 hp.2.1]⟩ - | k + 1, ha => by rw [_root_.pow_succ']; exact hp.multiplicity_finite_mul ha (ha.pow hp) + | k + 1, ha => by rw [_root_.pow_succ']; exact hp.finiteMultiplicity_mul ha (ha.pow hp) + +@[deprecated (since := "2024-11-30")] alias multiplicity.Finite.pow := FiniteMultiplicity.pow @[simp] theorem multiplicity_self {a : α} : multiplicity a a = 1 := by - by_cases ha : Finite a a + by_cases ha : FiniteMultiplicity a a · rw [ha.multiplicity_eq_iff] simp only [pow_one, dvd_refl, reduceAdd, true_and] rintro ⟨v, hv⟩ @@ -661,11 +786,14 @@ theorem multiplicity_self {a : α} : multiplicity a a = 1 := by · simp [ha] @[simp] -theorem multiplicity.Finite.emultiplicity_self {a : α} (hfin : Finite a a) : +theorem FiniteMultiplicity.emultiplicity_self {a : α} (hfin : FiniteMultiplicity a a) : emultiplicity a a = 1 := by simp [hfin.emultiplicity_eq_multiplicity] -theorem multiplicity_mul {p a b : α} (hp : Prime p) (hfin : Finite p (a * b)) : +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.emultiplicity_self := FiniteMultiplicity.emultiplicity_self + +theorem multiplicity_mul {p a b : α} (hp : Prime p) (hfin : FiniteMultiplicity p (a * b)) : multiplicity p (a * b) = multiplicity p a + multiplicity p b := by have hdiva : p ^ multiplicity p a ∣ a := pow_multiplicity_dvd .. have hdivb : p ^ multiplicity p b ∣ b := pow_multiplicity_dvd .. @@ -681,14 +809,14 @@ theorem multiplicity_mul {p a b : α} (hp : Prime p) (hfin : Finite p (a * b)) : theorem emultiplicity_mul {p a b : α} (hp : Prime p) : emultiplicity p (a * b) = emultiplicity p a + emultiplicity p b := by - by_cases hfin : Finite p (a * b) + by_cases hfin : FiniteMultiplicity p (a * b) · rw [hfin.emultiplicity_eq_multiplicity, hfin.mul_left.emultiplicity_eq_multiplicity, hfin.mul_right.emultiplicity_eq_multiplicity] norm_cast exact multiplicity_mul hp hfin · rw [emultiplicity_eq_top.2 hfin, eq_comm, WithTop.add_eq_top, emultiplicity_eq_top, emultiplicity_eq_top] - simpa only [Finite.mul_iff hp, not_and_or] using hfin + simpa only [FiniteMultiplicity.mul_iff hp, not_and_or] using hfin theorem Finset.emultiplicity_prod {β : Type*} {p : α} (hp : Prime p) (s : Finset β) (f : β → α) : emultiplicity p (∏ x ∈ s, f x) = ∑ x ∈ s, emultiplicity p (f x) := by classical @@ -703,11 +831,14 @@ theorem emultiplicity_pow {p a : α} (hp : Prime p) {k : ℕ} : · simp [emultiplicity_of_one_right hp.not_unit] · simp [pow_succ, emultiplicity_mul hp, hk, add_mul] -protected theorem multiplicity.Finite.multiplicity_pow {p a : α} (hp : Prime p) - (ha : Finite p a) {k : ℕ} : multiplicity p (a ^ k) = k * multiplicity p a := by +protected theorem FiniteMultiplicity.multiplicity_pow {p a : α} (hp : Prime p) + (ha : FiniteMultiplicity p a) {k : ℕ} : multiplicity p (a ^ k) = k * multiplicity p a := by exact_mod_cast (ha.pow hp).emultiplicity_eq_multiplicity ▸ ha.emultiplicity_eq_multiplicity ▸ emultiplicity_pow hp +@[deprecated (since := "2024-11-30")] +alias multiplicity.Finite.multiplicity_pow := FiniteMultiplicity.multiplicity_pow + theorem emultiplicity_pow_self {p : α} (h0 : p ≠ 0) (hu : ¬IsUnit p) (n : ℕ) : emultiplicity p (p ^ n) = n := by apply emultiplicity_eq_of_dvd_of_not_dvd @@ -731,8 +862,6 @@ end CancelCommMonoidWithZero section Nat -open multiplicity - theorem multiplicity_eq_zero_of_coprime {p a b : ℕ} (hp : p ≠ 1) (hle : multiplicity p a ≤ multiplicity p b) (hab : Nat.Coprime a b) : multiplicity p a = 0 := by apply Nat.eq_zero_of_not_pos @@ -745,16 +874,29 @@ theorem multiplicity_eq_zero_of_coprime {p a b : ℕ} (hp : p ≠ 1) end Nat -theorem Int.multiplicity_finite_iff_natAbs_finite {a b : ℤ} : - Finite a b ↔ Finite a.natAbs b.natAbs := by - simp only [Finite.def, ← Int.natAbs_dvd_natAbs, Int.natAbs_pow] +theorem Int.finiteMultiplicity_iff_finiteMultiplicity_natAbs {a b : ℤ} : + FiniteMultiplicity a b ↔ FiniteMultiplicity a.natAbs b.natAbs := by + simp only [FiniteMultiplicity.def, ← Int.natAbs_dvd_natAbs, Int.natAbs_pow] + +@[deprecated (since := "2024-11-30")] +alias Int.multiplicity_finite_iff_natAbs_finite := + Int.finiteMultiplicity_iff_finiteMultiplicity_natAbs -theorem Int.multiplicity_finite_iff {a b : ℤ} : Finite a b ↔ a.natAbs ≠ 1 ∧ b ≠ 0 := by - rw [multiplicity_finite_iff_natAbs_finite, Nat.multiplicity_finite_iff, +theorem Int.finiteMultiplicity_iff {a b : ℤ} : FiniteMultiplicity a b ↔ a.natAbs ≠ 1 ∧ b ≠ 0 := by + rw [finiteMultiplicity_iff_finiteMultiplicity_natAbs, Nat.finiteMultiplicity_iff, pos_iff_ne_zero, Int.natAbs_ne_zero] -instance Nat.decidableMultiplicityFinite : DecidableRel fun a b : ℕ => Finite a b := fun _ _ => - decidable_of_iff' _ Nat.multiplicity_finite_iff +@[deprecated (since := "2024-11-30")] +alias Int.multiplicity_finite_iff := Int.finiteMultiplicity_iff + +instance Nat.decidableFiniteMultiplicity : DecidableRel fun a b : ℕ => FiniteMultiplicity a b := + fun _ _ ↦ decidable_of_iff' _ Nat.finiteMultiplicity_iff + +@[deprecated (since := "2024-11-30")] +alias Nat.decidableMultiplicityFinite := Nat.decidableFiniteMultiplicity + +instance Int.decidableMultiplicityFinite : DecidableRel fun a b : ℤ => FiniteMultiplicity a b := + fun _ _ ↦ decidable_of_iff' _ Int.finiteMultiplicity_iff -instance Int.decidableMultiplicityFinite : DecidableRel fun a b : ℤ => Finite a b := fun _ _ => - decidable_of_iff' _ Int.multiplicity_finite_iff +@[deprecated (since := "2024-11-30")] +alias Int.decidableFiniteMultiplicity := Int.decidableMultiplicityFinite diff --git a/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean b/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean new file mode 100644 index 0000000000000..80d599cc70bfb --- /dev/null +++ b/Mathlib/RingTheory/MvPolynomial/EulerIdentity.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.Algebra.MvPolynomial.PDeriv +import Mathlib.RingTheory.MvPolynomial.Homogeneous + +/-! +# Euler's homogeneous identity + +## Main resutls + +* `IsHomogeneous.sum_X_mul_pderiv`: Euler's identity for homogeneous polynomials: + for a multivariate homogeneous polynomial, + the product of each variable with the derivative with respect to that variable + sums up to the degree times the polynomial. +* `IsWeightedHomogeneous.sum_weight_X_mul_pderiv`: the weighted version of Euler's identity. +-/ + +namespace MvPolynomial + +open Finsupp + +variable {R σ M : Type*} [CommSemiring R] {φ : MvPolynomial σ R} + +protected lemma IsWeightedHomogeneous.pderiv [AddCancelCommMonoid M] {w : σ → M} {n n' : M} {i : σ} + (h : φ.IsWeightedHomogeneous w n) (h' : n' + w i = n) : + (pderiv i φ).IsWeightedHomogeneous w n' := by + rw [← mem_weightedHomogeneousSubmodule, weightedHomogeneousSubmodule_eq_finsupp_supported, + Finsupp.supported_eq_span_single] at h + refine Submodule.span_induction ?_ ?_ (fun p q _ _ hp hq ↦ ?_) (fun r p _ h ↦ ?_) h + · rintro _ ⟨m, hm, rfl⟩ + simp_rw [single_eq_monomial, pderiv_monomial, one_mul] + by_cases hi : m i = 0 + · rw [hi, Nat.cast_zero, monomial_zero]; apply isWeightedHomogeneous_zero + convert isWeightedHomogeneous_monomial .. + rw [← add_right_cancel_iff (a := w i), h', ← hm, weight_sub_single_add hi] + · rw [map_zero]; apply isWeightedHomogeneous_zero + · rw [map_add]; exact hp.add hq + · rw [(pderiv i).map_smul]; exact (weightedHomogeneousSubmodule ..).smul_mem _ h + +protected lemma IsHomogeneous.pderiv {n : ℕ} {i : σ} (h : φ.IsHomogeneous n) : + (pderiv i φ).IsHomogeneous (n - 1) := by + obtain _ | n := n + · rw [← totalDegree_zero_iff_isHomogeneous, totalDegree_eq_zero_iff_eq_C] at h + rw [h, pderiv_C]; apply isHomogeneous_zero + · exact IsWeightedHomogeneous.pderiv h rfl + +variable [Fintype σ] {n : ℕ} + +open Finset in +/-- Euler's identity for weighted homogeneous polynomials. -/ +theorem IsWeightedHomogeneous.sum_weight_X_mul_pderiv {w : σ → ℕ} + (h : φ.IsWeightedHomogeneous w n) : ∑ i : σ, w i • (X i * pderiv i φ) = n • φ := by + rw [← mem_weightedHomogeneousSubmodule, weightedHomogeneousSubmodule_eq_finsupp_supported, + supported_eq_span_single] at h + refine Submodule.span_induction ?_ ?_ (fun p q _ _ hp hq ↦ ?_) (fun r p _ h ↦ ?_) h + · rintro _ ⟨m, hm, rfl⟩ + simp_rw [single_eq_monomial, X_mul_pderiv_monomial, smul_smul, ← sum_smul, mul_comm (w _)] + congr + rwa [Set.mem_setOf, weight_apply, sum_fintype] at hm + intro; apply zero_smul + · simp + · simp_rw [map_add, left_distrib, smul_add, sum_add_distrib, hp, hq] + · simp_rw [(pderiv _).map_smul, nsmul_eq_mul, mul_smul_comm, ← Finset.smul_sum, ← nsmul_eq_mul, h] + +/-- Euler's identity for homogeneous polynomials. -/ +theorem IsHomogeneous.sum_X_mul_pderiv (h : φ.IsHomogeneous n) : + ∑ i : σ, X i * pderiv i φ = n • φ := by + simp_rw [← h.sum_weight_X_mul_pderiv, Pi.one_apply, one_smul] + +end MvPolynomial diff --git a/Mathlib/RingTheory/MvPolynomial/Groebner.lean b/Mathlib/RingTheory/MvPolynomial/Groebner.lean new file mode 100644 index 0000000000000..a2149146b532f --- /dev/null +++ b/Mathlib/RingTheory/MvPolynomial/Groebner.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ +import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.MonomialOrder +import Mathlib.Data.Finsupp.WellFounded +import Mathlib.Data.List.TFAE +import Mathlib.RingTheory.MvPolynomial.Homogeneous +import Mathlib.RingTheory.MvPolynomial.MonomialOrder + +/-! # Division algorithm with respect to monomial orders + +We provide a division algorithm with respect to monomial orders in polynomial rings. +Let `R` be a commutative ring, `σ` a type of indeterminates and `m : MonomialOrder σ` +a monomial ordering on `σ →₀ ℕ`. + +Consider a family of polynomials `b : ι → MvPolynomial σ R` with invertible leading coefficients +(with respect to `m`) : we assume `hb : ∀ i, IsUnit (m.lCoeff (b i))`). + +* `MonomialOrder.div hb f` furnishes + - a finitely supported family `g : ι →₀ MvPolynomial σ R` + - and a “remainder” `r : MvPolynomial σ R` +such that the three properties hold: + (1) One has `f = ∑ (g i) * (b i) + r` + (2) For every `i`, `m.degree ((g i) * (b i)` is less than or equal to that of `f` + (3) For every `i`, every monomial in the support of `r` is strictly smaller + than the leading term of `b i`, + +The proof is done by induction, using two standard constructions + +* `MonomialOrder.subLTerm f` deletes the leading term of a polynomial `f` + +* `MonomialOrder.reduce hb f` subtracts from `f` the appropriate multiple of `b : MvPolynomial σ R`, +provided `IsUnit (m.lCoeff b)`. + +* `MonomialOrder.div_set` is the variant of `MonomialOrder.div` for a set of polynomials. + +## Reference : [Becker-Weispfenning1993] + +## TODO + +* Prove that under `Field F`, `IsUnit (m.lCoeff (b i))` is equivalent to `b i ≠ 0`. + +-/ + +namespace MonomialOrder + +open MvPolynomial + +open scoped MonomialOrder + +variable {σ : Type*} {m : MonomialOrder σ} {R : Type*} [CommRing R] + +variable (m) in +/-- Delete the leading term in a multivariate polynomial (for some monomial order) -/ +noncomputable def subLTerm (f : MvPolynomial σ R) : MvPolynomial σ R := + f - monomial (m.degree f) (m.lCoeff f) + +theorem degree_sub_LTerm_le (f : MvPolynomial σ R) : + m.degree (m.subLTerm f) ≼[m] m.degree f := by + apply le_trans degree_sub_le + simp only [sup_le_iff, le_refl, true_and] + apply degree_monomial_le + +theorem degree_sub_LTerm_lt {f : MvPolynomial σ R} (hf : m.degree f ≠ 0) : + m.degree (m.subLTerm f) ≺[m] m.degree f := by + rw [lt_iff_le_and_ne] + refine ⟨degree_sub_LTerm_le f, ?_⟩ + classical + intro hf' + simp only [EmbeddingLike.apply_eq_iff_eq] at hf' + have : m.subLTerm f ≠ 0 := by + intro h + simp only [h, degree_zero] at hf' + exact hf hf'.symm + rw [← coeff_degree_ne_zero_iff (m := m), hf'] at this + apply this + simp [subLTerm, coeff_monomial, lCoeff] + +variable (m) in +/-- Reduce a polynomial modulo a polynomial with unit leading term (for some monomial order) -/ +noncomputable def reduce {b : MvPolynomial σ R} (hb : IsUnit (m.lCoeff b)) (f : MvPolynomial σ R) : + MvPolynomial σ R := + f - monomial (m.degree f - m.degree b) (hb.unit⁻¹ * m.lCoeff f) * b + +theorem degree_reduce_lt {f b : MvPolynomial σ R} (hb : IsUnit (m.lCoeff b)) + (hbf : m.degree b ≤ m.degree f) (hf : m.degree f ≠ 0) : + m.degree (m.reduce hb f) ≺[m] m.degree f := by + have H : m.degree f = + m.degree ((monomial (m.degree f - m.degree b)) (hb.unit⁻¹ * m.lCoeff f)) + + m.degree b := by + classical + rw [degree_monomial, if_neg] + · ext d + rw [tsub_add_cancel_of_le hbf] + · simp only [Units.mul_right_eq_zero, lCoeff_eq_zero_iff] + intro hf0 + apply hf + simp [hf0] + have H' : coeff (m.degree f) (m.reduce hb f) = 0 := by + simp only [reduce, coeff_sub, sub_eq_zero] + nth_rewrite 2 [H] + rw [coeff_mul_of_degree_add (m := m), lCoeff_monomial] + rw [mul_comm, ← mul_assoc] + simp only [IsUnit.mul_val_inv, one_mul] + rfl + rw [lt_iff_le_and_ne] + constructor + · classical + apply le_trans degree_sub_le + simp only [sup_le_iff, le_refl, true_and] + apply le_of_le_of_eq degree_mul_le + rw [m.toSyn.injective.eq_iff] + exact H.symm + · intro K + simp only [EmbeddingLike.apply_eq_iff_eq] at K + nth_rewrite 1 [← K] at H' + change lCoeff m _ = 0 at H' + rw [lCoeff_eq_zero_iff] at H' + rw [H', degree_zero] at K + exact hf K.symm + +theorem div {ι : Type*} {b : ι → MvPolynomial σ R} + (hb : ∀ i, IsUnit (m.lCoeff (b i))) (f : MvPolynomial σ R) : + ∃ (g : ι →₀ (MvPolynomial σ R)) (r : MvPolynomial σ R), + f = Finsupp.linearCombination _ b g + r ∧ + (∀ i, m.degree (b i * (g i)) ≼[m] m.degree f) ∧ + (∀ c ∈ r.support, ∀ i, ¬ (m.degree (b i) ≤ c)) := by + by_cases hb' : ∃ i, m.degree (b i) = 0 + · obtain ⟨i, hb0⟩ := hb' + use Finsupp.single i ((hb i).unit⁻¹ • f), 0 + constructor + · simp only [Finsupp.linearCombination_single, smul_eq_mul, add_zero] + simp only [smul_mul_assoc, ← smul_eq_iff_eq_inv_smul, Units.smul_isUnit] + nth_rewrite 2 [eq_C_of_degree_eq_zero hb0] + rw [mul_comm, smul_eq_C_mul] + constructor + · intro j + by_cases hj : j = i + · apply le_trans degree_mul_le + simp only [hj, hb0, Finsupp.single_eq_same, zero_add] + apply le_of_eq + simp only [EmbeddingLike.apply_eq_iff_eq] + apply degree_smul (Units.isRegular _) + · simp only [Finsupp.single_eq_of_ne (Ne.symm hj), mul_zero, degree_zero, map_zero] + apply bot_le + · simp + push_neg at hb' + by_cases hf0 : f = 0 + · refine ⟨0, 0, by simp [hf0], ?_, by simp⟩ + intro b + simp only [Finsupp.coe_zero, Pi.zero_apply, mul_zero, degree_zero, map_zero] + exact bot_le + by_cases hf : ∃ i, m.degree (b i) ≤ m.degree f + · obtain ⟨i, hf⟩ := hf + have deg_reduce : m.degree (m.reduce (hb i) f) ≺[m] m.degree f := by + apply degree_reduce_lt (hb i) hf + intro hf0' + apply hb' i + simpa [hf0'] using hf + obtain ⟨g', r', H'⟩ := div hb (m.reduce (hb i) f) + use g' + + Finsupp.single i (monomial (m.degree f - m.degree (b i)) ((hb i).unit⁻¹ * m.lCoeff f)) + use r' + constructor + · rw [map_add, add_assoc, add_comm _ r', ← add_assoc, ← H'.1] + simp [reduce] + constructor + · rintro j + simp only [Finsupp.coe_add, Pi.add_apply] + rw [mul_add] + apply le_trans degree_add_le + simp only [sup_le_iff] + constructor + · exact le_trans (H'.2.1 _) (le_of_lt deg_reduce) + · classical + rw [Finsupp.single_apply] + split_ifs with hc + · apply le_trans degree_mul_le + simp only [map_add] + apply le_of_le_of_eq (add_le_add_left (degree_monomial_le _) _) + simp only [← hc] + rw [← map_add, m.toSyn.injective.eq_iff] + rw [add_tsub_cancel_of_le] + exact hf + · simp only [mul_zero, degree_zero, map_zero] + exact bot_le + · exact H'.2.2 + · push_neg at hf + suffices ∃ (g' : ι →₀ MvPolynomial σ R), ∃ r', + (m.subLTerm f = Finsupp.linearCombination (MvPolynomial σ R) b g' + r') ∧ + (∀ i, m.degree ((b i) * (g' i)) ≼[m] m.degree (m.subLTerm f)) ∧ + (∀ c ∈ r'.support, ∀ i, ¬ m.degree (b i) ≤ c) by + obtain ⟨g', r', H'⟩ := this + use g', r' + monomial (m.degree f) (m.lCoeff f) + constructor + · simp [← add_assoc, ← H'.1, subLTerm] + constructor + · exact fun b ↦ le_trans (H'.2.1 b) (degree_sub_LTerm_le f) + · intro c hc i + by_cases hc' : c ∈ r'.support + · exact H'.2.2 c hc' i + · convert hf i + classical + have := MvPolynomial.support_add hc + rw [Finset.mem_union, Classical.or_iff_not_imp_left] at this + simpa only [Finset.mem_singleton] using support_monomial_subset (this hc') + by_cases hf'0 : m.subLTerm f = 0 + · refine ⟨0, 0, by simp [hf'0], ?_, by simp⟩ + intro b + simp only [Finsupp.coe_zero, Pi.zero_apply, mul_zero, degree_zero, map_zero] + exact bot_le + · exact (div hb) (m.subLTerm f) +termination_by WellFounded.wrap + ((isWellFounded_iff m.syn fun x x_1 ↦ x < x_1).mp m.wf) (m.toSyn (m.degree f)) +decreasing_by +· exact deg_reduce +· apply degree_sub_LTerm_lt + intro hf0 + apply hf'0 + simp only [subLTerm, sub_eq_zero] + nth_rewrite 1 [eq_C_of_degree_eq_zero hf0, hf0] + simp + +theorem div_set {B : Set (MvPolynomial σ R)} + (hB : ∀ b ∈ B, IsUnit (m.lCoeff b)) (f : MvPolynomial σ R) : + ∃ (g : B →₀ (MvPolynomial σ R)) (r : MvPolynomial σ R), + f = Finsupp.linearCombination _ (fun (b : B) ↦ (b : MvPolynomial σ R)) g + r ∧ + (∀ (b : B), m.degree ((b : MvPolynomial σ R) * (g b)) ≼[m] m.degree f) ∧ + (∀ c ∈ r.support, ∀ b ∈ B, ¬ (m.degree b ≤ c)) := by + obtain ⟨g, r, H⟩ := m.div (b := fun (p : B) ↦ p) (fun b ↦ hB b b.prop) f + exact ⟨g, r, H.1, H.2.1, fun c hc b hb ↦ H.2.2 c hc ⟨b, hb⟩⟩ + +end MonomialOrder + + diff --git a/Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean b/Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean new file mode 100644 index 0000000000000..674a8ee44cb9d --- /dev/null +++ b/Mathlib/RingTheory/MvPolynomial/MonomialOrder.lean @@ -0,0 +1,388 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ +import Mathlib.RingTheory.MvPolynomial.Homogeneous +import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.WellFounded +import Mathlib.Data.List.TFAE +import Mathlib.Data.Finsupp.MonomialOrder + +/-! # Degree and leading coefficient of polynomials with respect to a monomial order + +We consider a type `σ` of indeterminates and a commutative semiring `R` +and a monomial order `m : MonomialOrder σ`. + +* `m.degree f` is the degree of `f` for the monomial ordering `m` + +* `m.lCoeff f` is the leading coefficient of `f` for the monomial ordering `m` + +* `m.lCoeff_ne_zero_iff f` asserts that this coefficient is nonzero iff `f ≠ 0`. + +* in a field, `m.lCoeff_is_unit_iff f` asserts that this coefficient is a unit iff `f ≠ 0`. + +* `m.degree_add_le` : the `m.degree` of `f + g` is smaller than or equal to the supremum +of those of `f` and `g` + +* `m.degree_add_of_lt h` : the `m.degree` of `f + g` is equal to that of `f` +if the `m.degree` of `g` is strictly smaller than that `f` + +* `m.lCoeff_add_of_lt h`: then, the leading coefficient of `f + g` is that of `f` . + +* `m.degree_add_of_ne h` : the `m.degree` of `f + g` is equal to that the supremum +of those of `f` and `g` if they are distinct + +* `m.degree_sub_le` : the `m.degree` of `f - g` is smaller than or equal to the supremum +of those of `f` and `g` + +* `m.degree_sub_of_lt h` : the `m.degree` of `f - g` is equal to that of `f` +if the `m.degree` of `g` is strictly smaller than that `f` + +* `m.lCoeff_sub_of_lt h`: then, the leading coefficient of `f - g` is that of `f` . + +* `m.degree_mul_le`: the `m.degree` of `f * g` is smaller than or equal to the sum of those of +`f` and `g`. + +* `m.degree_mul_of_isRegular_left`, `m.degree_mul_of_isRegular_right` and `m.degree_mul` +assert the equality when the leading coefficient of `f` or `g` is regular, +or when `R` is a domain and `f` and `g` are nonzero. + +* `m.lCoeff_mul_of_isRegular_left`, `m.lCoeff_mul_of_isRegular_right` and `m.lCoeff_mul` +say that `m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g` + +## Reference + +[Becker-Weispfenning1993] + +-/ + +namespace MonomialOrder + +open MvPolynomial + +open scoped MonomialOrder + +variable {σ : Type*} {m : MonomialOrder σ} + +section Semiring + +variable {R : Type*} [CommSemiring R] + +variable (m) in +/-- the degree of a multivariate polynomial with respect to a monomial ordering -/ +def degree {R : Type*} [CommSemiring R] (f : MvPolynomial σ R) : σ →₀ ℕ := + m.toSyn.symm (f.support.sup m.toSyn) + +variable (m) in +/-- the leading coefficient of a multivariate polynomial with respect to a monomial ordering -/ +def lCoeff {R : Type*} [CommSemiring R] (f : MvPolynomial σ R) : R := + f.coeff (m.degree f) + +@[simp] +theorem degree_zero : m.degree (0 : MvPolynomial σ R) = 0 := by + simp [degree] + +@[simp] +theorem lCoeff_zero : m.lCoeff (0 : MvPolynomial σ R) = 0 := by + simp [degree, lCoeff] + +theorem degree_monomial_le {d : σ →₀ ℕ} (c : R) : + m.degree (monomial d c) ≼[m] d := by + simp only [degree, AddEquiv.apply_symm_apply] + apply le_trans (Finset.sup_mono support_monomial_subset) + simp only [Finset.sup_singleton, le_refl] + +theorem degree_monomial {d : σ →₀ ℕ} (c : R) [Decidable (c = 0)] : + m.degree (monomial d c) = if c = 0 then 0 else d := by + simp only [degree, support_monomial] + split_ifs with hc <;> simp + +@[simp] +theorem lCoeff_monomial {d : σ →₀ ℕ} (c : R) : + m.lCoeff (monomial d c) = c := by + classical + simp only [lCoeff, degree_monomial] + split_ifs with hc <;> simp [hc] + +theorem degree_le_iff {f : MvPolynomial σ R} {d : σ →₀ ℕ} : + m.degree f ≼[m] d ↔ ∀ c ∈ f.support, c ≼[m] d := by + unfold degree + simp only [AddEquiv.apply_symm_apply, Finset.sup_le_iff, mem_support_iff, ne_eq] + +theorem degree_lt_iff {f : MvPolynomial σ R} {d : σ →₀ ℕ} (hd : 0 ≺[m] d) : + m.degree f ≺[m] d ↔ ∀ c ∈ f.support, c ≺[m] d := by + simp only [map_zero] at hd + unfold degree + simp only [AddEquiv.apply_symm_apply] + exact Finset.sup_lt_iff hd + +theorem le_degree {f : MvPolynomial σ R} {d : σ →₀ ℕ} (hd : d ∈ f.support) : + d ≼[m] m.degree f := by + unfold degree + simp only [AddEquiv.apply_symm_apply, Finset.le_sup hd] + +theorem coeff_eq_zero_of_lt {f : MvPolynomial σ R} {d : σ →₀ ℕ} (hd : m.degree f ≺[m] d) : + f.coeff d = 0 := by + rw [← not_le] at hd + by_contra hf + apply hd (m.le_degree (mem_support_iff.mpr hf)) +theorem lCoeff_ne_zero_iff {f : MvPolynomial σ R} : + m.lCoeff f ≠ 0 ↔ f ≠ 0 := by + constructor + · rw [not_imp_not] + intro hf + rw [hf, lCoeff_zero] + · intro hf + rw [← support_nonempty] at hf + rw [lCoeff, ← mem_support_iff, degree] + suffices f.support.sup m.toSyn ∈ m.toSyn '' f.support by + obtain ⟨d, hd, hd'⟩ := this + rw [← hd', AddEquiv.symm_apply_apply] + exact hd + exact Finset.sup_mem_of_nonempty hf + +@[simp] +theorem lCoeff_eq_zero_iff {f : MvPolynomial σ R} : + lCoeff m f = 0 ↔ f = 0 := by + simp only [← not_iff_not, lCoeff_ne_zero_iff] + +theorem coeff_degree_ne_zero_iff {f : MvPolynomial σ R} : + f.coeff (m.degree f) ≠ 0 ↔ f ≠ 0 := + m.lCoeff_ne_zero_iff + +@[simp] +theorem coeff_degree_eq_zero_iff {f : MvPolynomial σ R} : + f.coeff (m.degree f) = 0 ↔ f = 0 := + m.lCoeff_eq_zero_iff + +theorem degree_eq_zero_iff_totalDegree_eq_zero {f : MvPolynomial σ R} : + m.degree f = 0 ↔ f.totalDegree = 0 := by + rw [← m.toSyn.injective.eq_iff] + rw [map_zero, ← m.bot_eq_zero, eq_bot_iff, m.bot_eq_zero, ← m.toSyn.map_zero] + rw [degree_le_iff] + rw [totalDegree_eq_zero_iff] + apply forall_congr' + intro d + apply imp_congr (rfl.to_iff) + rw [map_zero, ← m.bot_eq_zero, ← eq_bot_iff, m.bot_eq_zero] + simp only [EmbeddingLike.map_eq_zero_iff] + exact Finsupp.ext_iff + +@[simp] +theorem degree_C (r : R) : + m.degree (C r) = 0 := by + rw [degree_eq_zero_iff_totalDegree_eq_zero, totalDegree_C] + +theorem degree_add_le {f g : MvPolynomial σ R} : + m.toSyn (m.degree (f + g)) ≤ m.toSyn (m.degree f) ⊔ m.toSyn (m.degree g) := by + conv_rhs => rw [← m.toSyn.apply_symm_apply (_ ⊔ _)] + rw [degree_le_iff] + simp only [AddEquiv.apply_symm_apply, le_sup_iff] + intro b hb + by_cases hf : b ∈ f.support + · left + exact m.le_degree hf + · right + apply m.le_degree + simp only [not_mem_support_iff] at hf + simpa only [mem_support_iff, coeff_add, hf, zero_add] using hb + +theorem degree_add_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.degree (f + g) = m.degree f := by + apply m.toSyn.injective + apply le_antisymm + · apply le_trans degree_add_le + simp only [sup_le_iff, le_refl, true_and, le_of_lt h] + · apply le_degree + rw [mem_support_iff, coeff_add, m.coeff_eq_zero_of_lt h, add_zero, ← lCoeff, lCoeff_ne_zero_iff] + intro hf + rw [← not_le, hf] at h + apply h + simp only [degree_zero, map_zero] + apply bot_le + +theorem lCoeff_add_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.lCoeff (f + g) = m.lCoeff f := by + simp only [lCoeff, m.degree_add_of_lt h, coeff_add, coeff_eq_zero_of_lt h, add_zero] + +theorem degree_add_of_ne {f g : MvPolynomial σ R} + (h : m.degree f ≠ m.degree g) : + m.toSyn (m.degree (f + g)) = m.toSyn (m.degree f) ⊔ m.toSyn (m.degree g) := by + by_cases h' : m.degree g ≺[m] m.degree f + · simp [degree_add_of_lt h', left_eq_sup, le_of_lt h'] + · rw [not_lt, le_iff_eq_or_lt, Classical.or_iff_not_imp_left, EmbeddingLike.apply_eq_iff_eq] at h' + rw [add_comm, degree_add_of_lt (h' h), right_eq_sup] + simp only [le_of_lt (h' h)] + +theorem degree_mul_le {f g : MvPolynomial σ R} : + m.degree (f * g) ≼[m] m.degree f + m.degree g := by + classical + rw [degree_le_iff] + intro c + rw [← not_lt, mem_support_iff, not_imp_not] + intro hc + rw [coeff_mul] + apply Finset.sum_eq_zero + rintro ⟨d, e⟩ hde + simp only [Finset.mem_antidiagonal] at hde + dsimp only + by_cases hd : m.degree f ≺[m] d + · rw [m.coeff_eq_zero_of_lt hd, zero_mul] + · suffices m.degree g ≺[m] e by + rw [m.coeff_eq_zero_of_lt this, mul_zero] + simp only [not_lt] at hd + apply lt_of_add_lt_add_left (a := m.toSyn d) + simp only [← map_add, hde] + apply lt_of_le_of_lt _ hc + simp only [map_add] + exact add_le_add_right hd _ + +/-- Multiplicativity of leading coefficients -/ +theorem coeff_mul_of_degree_add {f g : MvPolynomial σ R} : + (f * g).coeff (m.degree f + m.degree g) = m.lCoeff f * m.lCoeff g := by + classical + rw [coeff_mul] + rw [Finset.sum_eq_single (m.degree f, m.degree g)] + · rfl + · rintro ⟨c, d⟩ hcd h + simp only [Finset.mem_antidiagonal] at hcd + by_cases hf : m.degree f ≺[m] c + · rw [m.coeff_eq_zero_of_lt hf, zero_mul] + · suffices m.degree g ≺[m] d by + rw [coeff_eq_zero_of_lt this, mul_zero] + apply lt_of_add_lt_add_left (a := m.toSyn c) + simp only [← map_add, hcd] + simp only [map_add] + rw [← not_le] + intro h'; apply hf + simp only [le_iff_eq_or_lt] at h' + cases h' with + | inl h' => + simp only [← map_add, EmbeddingLike.apply_eq_iff_eq, add_left_inj] at h' + exfalso + apply h + simp only [h', Prod.mk.injEq, true_and] + simpa [h'] using hcd + | inr h' => + exact lt_of_add_lt_add_right h' + · simp + +/-- Multiplicativity of leading coefficients -/ +theorem degree_mul_of_isRegular_left {f g : MvPolynomial σ R} + (hf : IsRegular (m.lCoeff f)) (hg : g ≠ 0) : + m.degree (f * g) = m.degree f + m.degree g := by + apply m.toSyn.injective + apply le_antisymm degree_mul_le + apply le_degree + rw [mem_support_iff, coeff_mul_of_degree_add] + simp only [ne_eq, hf, IsRegular.left, IsLeftRegular.mul_left_eq_zero_iff, + lCoeff_eq_zero_iff] + exact hg + +/-- Multiplicativity of leading coefficients -/ +theorem lCoeff_mul_of_isRegular_left {f g : MvPolynomial σ R} + (hf : IsRegular (m.lCoeff f)) (hg : g ≠ 0) : + m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g := by + simp only [lCoeff, degree_mul_of_isRegular_left hf hg, coeff_mul_of_degree_add] + +/-- Multiplicativity of leading coefficients -/ +theorem degree_mul_of_isRegular_right {f g : MvPolynomial σ R} + (hf : f ≠ 0) (hg : IsRegular (m.lCoeff g)) : + m.degree (f * g) = m.degree f + m.degree g := by + rw [mul_comm, m.degree_mul_of_isRegular_left hg hf, add_comm] + +/-- Multiplicativity of leading coefficients -/ +theorem lCoeff_mul_of_isRegular_right {f g : MvPolynomial σ R} + (hf : f ≠ 0) (hg : IsRegular (m.lCoeff g)) : + m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g := by + simp only [lCoeff, degree_mul_of_isRegular_right hf hg, coeff_mul_of_degree_add] + +/-- Degree of product -/ +theorem degree_mul [IsDomain R] {f g : MvPolynomial σ R} (hf : f ≠ 0) (hg : g ≠ 0) : + m.degree (f * g) = m.degree f + m.degree g := + degree_mul_of_isRegular_left (isRegular_of_ne_zero (lCoeff_ne_zero_iff.mpr hf)) hg + +/-- Degree of of product -/ +theorem degree_mul_of_nonzero_mul [IsDomain R] {f g : MvPolynomial σ R} (hfg : f * g ≠ 0) : + m.degree (f * g) = m.degree f + m.degree g := + degree_mul (left_ne_zero_of_mul hfg) (right_ne_zero_of_mul hfg) + +/-- Multiplicativity of leading coefficients -/ +theorem lCoeff_mul [IsDomain R] {f g : MvPolynomial σ R} + (hf : f ≠ 0) (hg : g ≠ 0) : + m.lCoeff (f * g) = m.lCoeff f * m.lCoeff g := by + rw [lCoeff, degree_mul hf hg, ← coeff_mul_of_degree_add] + +theorem degree_smul_le {r : R} {f : MvPolynomial σ R} : + m.degree (r • f) ≼[m] m.degree f := by + rw [smul_eq_C_mul] + apply le_of_le_of_eq degree_mul_le + simp + +theorem degree_smul {r : R} (hr : IsRegular r) {f : MvPolynomial σ R} : + m.degree (r • f) = m.degree f := by + by_cases hf : f = 0 + · simp [hf] + apply m.toSyn.injective + apply le_antisymm degree_smul_le + apply le_degree + simp only [mem_support_iff, smul_eq_C_mul] + rw [← zero_add (degree m f), ← degree_C r, coeff_mul_of_degree_add] + simp [lCoeff, hr.left.mul_left_eq_zero_iff, hf] + +theorem eq_C_of_degree_eq_zero {f : MvPolynomial σ R} (hf : m.degree f = 0) : + f = C (m.lCoeff f) := by + ext d + simp only [lCoeff, hf] + classical + by_cases hd : d = 0 + · simp [hd] + · rw [coeff_C, if_neg (Ne.symm hd)] + apply coeff_eq_zero_of_lt (m := m) + rw [hf, map_zero, lt_iff_le_and_ne, ne_eq, eq_comm, EmbeddingLike.map_eq_zero_iff] + exact ⟨bot_le, hd⟩ + +end Semiring + +section Ring + +variable {R : Type*} [CommRing R] + +@[simp] +theorem degree_neg {f : MvPolynomial σ R} : + m.degree (-f) = m.degree f := by + unfold degree + rw [support_neg] + +theorem degree_sub_le {f g : MvPolynomial σ R} : + m.toSyn (m.degree (f - g)) ≤ m.toSyn (m.degree f) ⊔ m.toSyn (m.degree g) := by + rw [sub_eq_add_neg] + apply le_of_le_of_eq m.degree_add_le + rw [degree_neg] + +theorem degree_sub_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.degree (f - g) = m.degree f := by + rw [sub_eq_add_neg] + apply degree_add_of_lt + simp only [degree_neg, h] + +theorem lCoeff_sub_of_lt {f g : MvPolynomial σ R} (h : m.degree g ≺[m] m.degree f) : + m.lCoeff (f - g) = m.lCoeff f := by + rw [sub_eq_add_neg] + apply lCoeff_add_of_lt + simp only [degree_neg, h] + +end Ring + +section Field + +variable {R : Type*} [Field R] + +theorem lCoeff_is_unit_iff {f : MvPolynomial σ R} : + IsUnit (m.lCoeff f) ↔ f ≠ 0 := by + simp only [isUnit_iff_ne_zero, ne_eq, lCoeff_eq_zero_iff] + +end Field + +end MonomialOrder diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index 18bd357e9e117..c0b99872fcf13 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -251,6 +251,12 @@ theorem mul {w : σ → M} (hφ : IsWeightedHomogeneous w φ m) (hψ : IsWeighte IsWeightedHomogeneous w (φ * ψ) (m + n) := weightedHomogeneousSubmodule_mul w m n <| Submodule.mul_mem_mul hφ hψ +theorem pow {w : σ → M} (hφ : IsWeightedHomogeneous w φ m) (n : ℕ) : + IsWeightedHomogeneous w (φ ^ n) (n • m) := by + induction n with + | zero => rw [pow_zero, zero_smul]; exact isWeightedHomogeneous_one R w + | succ n ih => rw [pow_succ, succ_nsmul]; exact ih.mul hφ + /-- A product of weighted homogeneous polynomials is weighted homogeneous, with weighted degree equal to the sum of the weighted degrees. -/ theorem prod {ι : Type*} (s : Finset ι) (φ : ι → MvPolynomial σ R) (n : ι → M) {w : σ → M} : diff --git a/Mathlib/RingTheory/Nakayama.lean b/Mathlib/RingTheory/Nakayama.lean index 5d859e8e2189a..05507c2e511a2 100644 --- a/Mathlib/RingTheory/Nakayama.lean +++ b/Mathlib/RingTheory/Nakayama.lean @@ -3,8 +3,9 @@ Copyright (c) 2021 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import Mathlib.RingTheory.Finiteness.Basic import Mathlib.RingTheory.Finiteness.Nakayama -import Mathlib.RingTheory.JacobsonIdeal +import Mathlib.RingTheory.Jacobson.Ideal /-! # Nakayama's lemma diff --git a/Mathlib/RingTheory/Nullstellensatz.lean b/Mathlib/RingTheory/Nullstellensatz.lean index 71d381cfeb340..6e41ad765593b 100644 --- a/Mathlib/RingTheory/Nullstellensatz.lean +++ b/Mathlib/RingTheory/Nullstellensatz.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Devon Tuma. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Devon Tuma -/ -import Mathlib.RingTheory.Jacobson +import Mathlib.RingTheory.Jacobson.Ring import Mathlib.FieldTheory.IsAlgClosed.Basic import Mathlib.RingTheory.MvPolynomial import Mathlib.RingTheory.PrimeSpectrum diff --git a/Mathlib/RingTheory/Polynomial/Chebyshev.lean b/Mathlib/RingTheory/Polynomial/Chebyshev.lean index 4e9eb5a7fe9d2..b8d50d06cb70c 100644 --- a/Mathlib/RingTheory/Polynomial/Chebyshev.lean +++ b/Mathlib/RingTheory/Polynomial/Chebyshev.lean @@ -16,16 +16,21 @@ with integral coefficients. * `Polynomial.Chebyshev.T`: the Chebyshev polynomials of the first kind. * `Polynomial.Chebyshev.U`: the Chebyshev polynomials of the second kind. +* `Polynomial.Chebyshev.C`: the rescaled Chebyshev polynomials of the first kind (also known as the + Vieta–Lucas polynomials), given by $C_n(2x) = 2T_n(x)$. +* `Polynomial.Chebyshev.S`: the rescaled Chebyshev polynomials of the second kind (also known as the + Vieta–Fibonacci polynomials), given by $S_n(2x) = U_n(x)$. ## Main statements * The formal derivative of the Chebyshev polynomials of the first kind is a scalar multiple of the Chebyshev polynomials of the second kind. -* `Polynomial.Chebyshev.mul_T`, twice the product of the `m`-th and `k`-th Chebyshev polynomials of - the first kind is the sum of the `m + k`-th and `m - k`-th Chebyshev polynomials of the first - kind. +* `Polynomial.Chebyshev.T_mul_T`, twice the product of the `m`-th and `k`-th Chebyshev polynomials + of the first kind is the sum of the `m + k`-th and `m - k`-th Chebyshev polynomials of the first + kind. There is a similar statement `Polynomial.Chebyshev.C_mul_C` for the `C` polynomials. * `Polynomial.Chebyshev.T_mul`, the `(m * n)`-th Chebyshev polynomial of the first kind is the - composition of the `m`-th and `n`-th Chebyshev polynomials of the first kind. + composition of the `m`-th and `n`-th Chebyshev polynomials of the first kind. There is a similar + statement `Polynomial.Chebyshev.C_mul` for the `C` polynomials. ## Implementation details @@ -53,7 +58,7 @@ namespace Polynomial.Chebyshev open Polynomial -variable (R S : Type*) [CommRing R] [CommRing S] +variable (R R' : Type*) [CommRing R] [CommRing R'] /-- `T n` is the `n`-th Chebyshev polynomial of the first kind. -/ -- Well-founded definitions are now irreducible by default; @@ -74,7 +79,7 @@ protected theorem induct (motive : ℤ → Prop) (add_two : ∀ (n : ℕ), motive (↑n + 1) → motive ↑n → motive (↑n + 2)) (neg_add_one : ∀ (n : ℕ), motive (-↑n) → motive (-↑n + 1) → motive (-↑n - 1)) : ∀ (a : ℤ), motive a := - T.induct Unit motive zero one add_two fun n hn hnm => by + T.induct motive zero one add_two fun n hn hnm => by simpa only [Int.negSucc_eq, neg_add] using neg_add_one n hn hnm @[simp] @@ -100,7 +105,7 @@ theorem T_zero : T R 0 = 1 := rfl @[simp] theorem T_one : T R 1 = X := rfl -theorem T_neg_one : T R (-1) = X := (by ring : 2 * X * 1 - X = X) +theorem T_neg_one : T R (-1) = X := show 2 * X * 1 - X = X by ring theorem T_two : T R 2 = 2 * X ^ 2 - 1 := by simpa [pow_two, mul_assoc] using T_add_two R 0 @@ -217,10 +222,195 @@ theorem one_sub_X_sq_mul_U_eq_pol_in_T (n : ℤ) : (1 - X ^ 2) * U R n = X * T R (n + 1) - T R (n + 2) := by linear_combination T_eq_X_mul_T_sub_pol_U R n -variable {R S} +/-- `C n` is the `n`th rescaled Chebyshev polynomial of the first kind (also known as a Vieta–Lucas +polynomial), given by $C_n(2x) = 2T_n(x)$. See `Polynomial.Chebyshev.C_comp_two_mul_X`. -/ +@[semireducible] noncomputable def C : ℤ → R[X] + | 0 => 2 + | 1 => X + | (n : ℕ) + 2 => X * C (n + 1) - C n + | -((n : ℕ) + 1) => X * C (-n) - C (-n + 1) + termination_by n => Int.natAbs n + Int.natAbs (n - 1) + +@[simp] +theorem C_add_two : ∀ n, C R (n + 2) = X * C R (n + 1) - C R n + | (k : ℕ) => C.eq_3 R k + | -(k + 1 : ℕ) => by linear_combination (norm := (simp [Int.negSucc_eq]; ring_nf)) C.eq_4 R k + +theorem C_add_one (n : ℤ) : C R (n + 1) = X * C R n - C R (n - 1) := by + linear_combination (norm := ring_nf) C_add_two R (n - 1) + +theorem C_sub_two (n : ℤ) : C R (n - 2) = X * C R (n - 1) - C R n := by + linear_combination (norm := ring_nf) C_add_two R (n - 2) + +theorem C_sub_one (n : ℤ) : C R (n - 1) = X * C R n - C R (n + 1) := by + linear_combination (norm := ring_nf) C_add_two R (n - 1) + +theorem C_eq (n : ℤ) : C R n = X * C R (n - 1) - C R (n - 2) := by + linear_combination (norm := ring_nf) C_add_two R (n - 2) + +@[simp] +theorem C_zero : C R 0 = 2 := rfl + +@[simp] +theorem C_one : C R 1 = X := rfl + +theorem C_neg_one : C R (-1) = X := show X * 2 - X = X by ring + +theorem C_two : C R 2 = X ^ 2 - 2 := by + simpa [pow_two, mul_assoc] using C_add_two R 0 @[simp] -theorem map_T (f : R →+* S) (n : ℤ) : map f (T R n) = T S n := by +theorem C_neg (n : ℤ) : C R (-n) = C R n := by + induction n using Polynomial.Chebyshev.induct with + | zero => rfl + | one => show X * 2 - X = X; ring + | add_two n ih1 ih2 => + have h₁ := C_add_two R n + have h₂ := C_sub_two R (-n) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 - h₁ + h₂ + | neg_add_one n ih1 ih2 => + have h₁ := C_add_one R n + have h₂ := C_sub_one R (-n) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 + h₁ - h₂ + +theorem C_natAbs (n : ℤ) : C R n.natAbs = C R n := by + obtain h | h := Int.natAbs_eq n <;> nth_rw 2 [h]; simp + +theorem C_neg_two : C R (-2) = X ^ 2 - 2 := by simp [C_two] + +theorem C_comp_two_mul_X (n : ℤ) : (C R n).comp (2 * X) = 2 * T R n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp_rw [C_add_two, T_add_two, sub_comp, mul_comp, X_comp, ih1, ih2] + ring + | neg_add_one n ih1 ih2 => + simp_rw [C_sub_one, T_sub_one, sub_comp, mul_comp, X_comp, ih1, ih2] + ring + +theorem C_eq_two_mul_T_comp_half_mul_X [Invertible (2 : R)] (n : ℤ) : + C R n = 2 * (T R n).comp (Polynomial.C ⅟2 * X) := by + have := congr_arg (·.comp (Polynomial.C ⅟2 * X)) (C_comp_two_mul_X R n) + simp_rw [comp_assoc, mul_comp, ofNat_comp, X_comp, ← mul_assoc, ← C_eq_natCast, ← C_mul, + Nat.cast_ofNat, mul_invOf_self', map_one, one_mul, comp_X, map_ofNat] at this + assumption + +theorem T_eq_half_mul_C_comp_two_mul_X [Invertible (2 : R)] (n : ℤ) : + T R n = Polynomial.C ⅟2 * (C R n).comp (2 * X) := by + rw [C_comp_two_mul_X, ← mul_assoc, ← map_ofNat Polynomial.C 2, ← map_mul, invOf_mul_self', + map_one, one_mul] + +/-- `S n` is the `n`th rescaled Chebyshev polynomial of the second kind (also known as a +Vieta–Fibonacci polynomial), given by $S_n(2x) = U_n(x)$. See +`Polynomial.Chebyshev.S_comp_two_mul_X`. -/ +@[semireducible] noncomputable def S : ℤ → R[X] + | 0 => 1 + | 1 => X + | (n : ℕ) + 2 => X * S (n + 1) - S n + | -((n : ℕ) + 1) => X * S (-n) - S (-n + 1) + termination_by n => Int.natAbs n + Int.natAbs (n - 1) + +@[simp] +theorem S_add_two : ∀ n, S R (n + 2) = X * S R (n + 1) - S R n + | (k : ℕ) => S.eq_3 R k + | -(k + 1 : ℕ) => by linear_combination (norm := (simp [Int.negSucc_eq]; ring_nf)) S.eq_4 R k + +theorem S_add_one (n : ℤ) : S R (n + 1) = X * S R n - S R (n - 1) := by + linear_combination (norm := ring_nf) S_add_two R (n - 1) + +theorem S_sub_two (n : ℤ) : S R (n - 2) = X * S R (n - 1) - S R n := by + linear_combination (norm := ring_nf) S_add_two R (n - 2) + +theorem S_sub_one (n : ℤ) : S R (n - 1) = X * S R n - S R (n + 1) := by + linear_combination (norm := ring_nf) S_add_two R (n - 1) + +theorem S_eq (n : ℤ) : S R n = X * S R (n - 1) - S R (n - 2) := by + linear_combination (norm := ring_nf) S_add_two R (n - 2) + +@[simp] +theorem S_zero : S R 0 = 1 := rfl + +@[simp] +theorem S_one : S R 1 = X := rfl + +@[simp] +theorem S_neg_one : S R (-1) = 0 := by simpa using S_sub_one R 0 + +theorem S_two : S R 2 = X ^ 2 - 1 := by + have := S_add_two R 0 + simp only [zero_add, S_one, S_zero] at this + linear_combination this + +@[simp] +theorem S_neg_two : S R (-2) = -1 := by + simpa [zero_sub, Int.reduceNeg, S_neg_one, mul_zero, S_zero] using S_sub_two R 0 + +theorem S_neg_sub_one (n : ℤ) : S R (-n - 1) = -S R (n - 1) := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + have h₁ := S_add_one R n + have h₂ := S_sub_two R (-n - 1) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 + h₁ + h₂ + | neg_add_one n ih1 ih2 => + have h₁ := S_eq R n + have h₂ := S_sub_two R (-n) + linear_combination (norm := ring_nf) (X:R[X]) * ih1 - ih2 + h₁ + h₂ + +theorem S_neg (n : ℤ) : S R (-n) = -S R (n - 2) := by simpa [sub_sub] using S_neg_sub_one R (n - 1) + +@[simp] +theorem S_neg_sub_two (n : ℤ) : S R (-n - 2) = -S R n := by + simpa [sub_eq_add_neg, add_comm] using S_neg R (n + 2) + +theorem S_comp_two_mul_X (n : ℤ) : (S R n).comp (2 * X) = U R n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => simp_rw [U_add_two, S_add_two, sub_comp, mul_comp, X_comp, ih1, ih2] + | neg_add_one n ih1 ih2 => simp_rw [U_sub_one, S_sub_one, sub_comp, mul_comp, X_comp, ih1, ih2] + +theorem S_sq_add_S_sq (n : ℤ) : S R n ^ 2 + S R (n + 1) ^ 2 - X * S R n * S R (n + 1) = 1 := by + induction n using Int.induction_on with + | hz => simp; ring + | hp n ih => + have h₁ := S_add_two R n + linear_combination (norm := ring_nf) (S R (2 + n) - S R n) * h₁ + ih + | hn n ih => + have h₁ := S_sub_one R (-n) + linear_combination (norm := ring_nf) (S R (-1 - n) - S R (1 - n)) * h₁ + ih + +theorem S_eq_U_comp_half_mul_X [Invertible (2 : R)] (n : ℤ) : + S R n = (U R n).comp (Polynomial.C ⅟2 * X) := by + have := congr_arg (·.comp (Polynomial.C ⅟2 * X)) (S_comp_two_mul_X R n) + simp_rw [comp_assoc, mul_comp, ofNat_comp, X_comp, ← mul_assoc, ← C_eq_natCast, ← C_mul, + Nat.cast_ofNat, mul_invOf_self', map_one, one_mul, comp_X] at this + assumption + +theorem S_eq_X_mul_S_add_C (n : ℤ) : 2 * S R (n + 1) = X * S R n + C R (n + 1) := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp [two_mul] + | one => simp [S_two, C_two]; ring + | add_two n ih1 ih2 => + have h₁ := S_add_two R (n + 1) + have h₂ := S_add_two R n + have h₃ := C_add_two R (n + 1) + linear_combination (norm := ring_nf) -h₃ - (X:R[X]) * h₂ + 2 * h₁ + (X:R[X]) * ih1 - ih2 + | neg_add_one n ih1 ih2 => + have h₁ := S_add_two R (-n - 1) + have h₂ := S_add_two R (-n) + have h₃ := C_add_two R (-n) + linear_combination (norm := ring_nf) -h₃ + 2 * h₂ - (X:R[X]) * h₁ - ih2 + (X:R[X]) * ih1 + +theorem C_eq_S_sub_X_mul_S (n : ℤ) : C R n = 2 * S R n - X * S R (n - 1) := by + linear_combination (norm := ring_nf) - S_eq_X_mul_S_add_C R (n - 1) + +variable {R R'} + +@[simp] +theorem map_T (f : R →+* R') (n : ℤ) : map f (T R n) = T R' n := by induction n using Polynomial.Chebyshev.induct with | zero => simp | one => simp @@ -232,7 +422,7 @@ theorem map_T (f : R →+* S) (n : ℤ) : map f (T R n) = T S n := by ih2] @[simp] -theorem map_U (f : R →+* S) (n : ℤ) : map f (U R n) = U S n := by +theorem map_U (f : R →+* R') (n : ℤ) : map f (U R n) = U R' n := by induction n using Polynomial.Chebyshev.induct with | zero => simp | one => simp @@ -243,6 +433,26 @@ theorem map_U (f : R →+* S) (n : ℤ) : map f (U R n) = U S n := by simp_rw [U_sub_one, Polynomial.map_sub, Polynomial.map_mul, Polynomial.map_ofNat, map_X, ih1, ih2] +@[simp] +theorem map_C (f : R →+* R') (n : ℤ) : map f (C R n) = C R' n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp_rw [C_add_two, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + | neg_add_one n ih1 ih2 => + simp_rw [C_sub_one, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + +@[simp] +theorem map_S (f : R →+* R') (n : ℤ) : map f (S R n) = S R' n := by + induction n using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two n ih1 ih2 => + simp_rw [S_add_two, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + | neg_add_one n ih1 ih2 => + simp_rw [S_sub_one, Polynomial.map_sub, Polynomial.map_mul, map_X, ih1, ih2] + theorem T_derivative_eq_U (n : ℤ) : derivative (T R n) = n * U R (n - 1) := by induction n using Polynomial.Chebyshev.induct with | zero => simp @@ -284,7 +494,7 @@ variable (R) /-- Twice the product of two Chebyshev `T` polynomials is the sum of two other Chebyshev `T` polynomials. -/ -theorem mul_T (m k : ℤ) : 2 * T R m * T R k = T R (m + k) + T R (m - k) := by +theorem T_mul_T (m k : ℤ) : 2 * T R m * T R k = T R (m + k) + T R (m - k) := by induction k using Polynomial.Chebyshev.induct with | zero => simp [two_mul] | one => rw [T_add_one, T_one]; ring @@ -299,20 +509,55 @@ theorem mul_T (m k : ℤ) : 2 * T R m * T R k = T R (m + k) + T R (m - k) := by have h₃ := T_add_two R (-k - 1) linear_combination (norm := ring_nf) 2 * T R m * h₃ - h₂ - h₁ - ih2 + 2 * (X : R[X]) * ih1 +@[deprecated (since := "2024-12-03")] alias mul_T := T_mul_T + +/-- The product of two Chebyshev `C` polynomials is the sum of two other Chebyshev `C` polynomials. +-/ +theorem C_mul_C (m k : ℤ) : C R m * C R k = C R (m + k) + C R (m - k) := by + induction k using Polynomial.Chebyshev.induct with + | zero => simp [mul_two] + | one => rw [C_add_one, C_one]; ring + | add_two k ih1 ih2 => + have h₁ := C_add_two R (m + k) + have h₂ := C_sub_two R (m - k) + have h₃ := C_add_two R k + linear_combination (norm := ring_nf) C R m * h₃ - h₂ - h₁ - ih2 + (X:R[X]) * ih1 + | neg_add_one k ih1 ih2 => + have h₁ := C_add_two R (m + (-k - 1)) + have h₂ := C_sub_two R (m - (-k - 1)) + have h₃ := C_add_two R (-k - 1) + linear_combination (norm := ring_nf) C R m * h₃ - h₂ - h₁ - ih2 + (X:R[X]) * ih1 + /-- The `(m * n)`-th Chebyshev `T` polynomial is the composition of the `m`-th and `n`-th. -/ theorem T_mul (m n : ℤ) : T R (m * n) = (T R m).comp (T R n) := by induction m using Polynomial.Chebyshev.induct with | zero => simp | one => simp | add_two m ih1 ih2 => - have h₁ := mul_T R ((m + 1) * n) n + have h₁ := T_mul_T R ((m + 1) * n) n have h₂ := congr_arg (comp · (T R n)) <| T_add_two R m simp only [sub_comp, mul_comp, ofNat_comp, X_comp] at h₂ linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + 2 * T R n * ih1 | neg_add_one m ih1 ih2 => - have h₁ := mul_T R ((-m) * n) n + have h₁ := T_mul_T R ((-m) * n) n have h₂ := congr_arg (comp · (T R n)) <| T_add_two R (-m - 1) simp only [sub_comp, mul_comp, ofNat_comp, X_comp] at h₂ linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + 2 * T R n * ih1 +/-- The `(m * n)`-th Chebyshev `C` polynomial is the composition of the `m`-th and `n`-th. -/ +theorem C_mul (m n : ℤ) : C R (m * n) = (C R m).comp (C R n) := by + induction m using Polynomial.Chebyshev.induct with + | zero => simp + | one => simp + | add_two m ih1 ih2 => + have h₁ := C_mul_C R ((m + 1) * n) n + have h₂ := congr_arg (comp · (C R n)) <| C_add_two R m + simp only [sub_comp, mul_comp, X_comp] at h₂ + linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + C R n * ih1 + | neg_add_one m ih1 ih2 => + have h₁ := C_mul_C R ((-m) * n) n + have h₂ := congr_arg (comp · (C R n)) <| C_add_two R (-m - 1) + simp only [sub_comp, mul_comp, X_comp] at h₂ + linear_combination (norm := ring_nf) -ih2 - h₂ - h₁ + C R n * ih1 + end Polynomial.Chebyshev diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index 1872e6ac8a9e2..b563e9a572c2a 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -135,25 +135,45 @@ private theorem two_mul_C_half_eq_one [Invertible (2 : R)] : 2 * C (⅟ 2 : R) = private theorem C_half_mul_two_eq_one [Invertible (2 : R)] : C (⅟ 2 : R) * 2 = 1 := by rw [mul_comm, two_mul_C_half_eq_one] -theorem dickson_one_one_eq_chebyshev_T [Invertible (2 : R)] : - ∀ n, dickson 1 (1 : R) n = 2 * (Chebyshev.T R n).comp (C (⅟ 2) * X) +theorem dickson_one_one_eq_chebyshev_C : ∀ n, dickson 1 (1 : R) n = Chebyshev.C R n | 0 => by - simp only [Chebyshev.T_zero, mul_one, one_comp, dickson_zero] + simp only [Chebyshev.C_zero, mul_one, one_comp, dickson_zero] norm_num | 1 => by - rw [dickson_one, Nat.cast_one, Chebyshev.T_one, X_comp, ← mul_assoc, two_mul_C_half_eq_one, - one_mul] + rw [dickson_one, Nat.cast_one, Chebyshev.C_one] | n + 2 => by - rw [dickson_add_two, C_1, Nat.cast_add, Nat.cast_two, Chebyshev.T_add_two, - dickson_one_one_eq_chebyshev_T (n + 1), dickson_one_one_eq_chebyshev_T n, sub_comp, mul_comp, - mul_comp, X_comp, ofNat_comp] - simp_rw [← mul_assoc, Nat.cast_ofNat, two_mul_C_half_eq_one, Nat.cast_add, Nat.cast_one] + rw [dickson_add_two, C_1, Nat.cast_add, Nat.cast_two, Chebyshev.C_add_two, + dickson_one_one_eq_chebyshev_C (n + 1), dickson_one_one_eq_chebyshev_C n] + push_cast ring +theorem dickson_one_one_eq_chebyshev_T [Invertible (2 : R)] (n : ℕ) : + dickson 1 (1 : R) n = 2 * (Chebyshev.T R n).comp (C (⅟ 2) * X) := + (dickson_one_one_eq_chebyshev_C R n).trans (Chebyshev.C_eq_two_mul_T_comp_half_mul_X R n) + theorem chebyshev_T_eq_dickson_one_one [Invertible (2 : R)] (n : ℕ) : - Chebyshev.T R n = C (⅟ 2) * (dickson 1 1 n).comp (2 * X) := by - rw [dickson_one_one_eq_chebyshev_T, mul_comp, ofNat_comp, comp_assoc, mul_comp, C_comp, X_comp] - simp_rw [← mul_assoc, Nat.cast_ofNat, C_half_mul_two_eq_one, one_mul, comp_X] + Chebyshev.T R n = C (⅟ 2) * (dickson 1 1 n).comp (2 * X) := + dickson_one_one_eq_chebyshev_C R n ▸ Chebyshev.T_eq_half_mul_C_comp_two_mul_X R n + +theorem dickson_two_one_eq_chebyshev_S : ∀ n, dickson 2 (1 : R) n = Chebyshev.S R n + | 0 => by + simp only [Chebyshev.S_zero, mul_one, one_comp, dickson_zero] + norm_num + | 1 => by + rw [dickson_one, Nat.cast_one, Chebyshev.S_one] + | n + 2 => by + rw [dickson_add_two, C_1, Nat.cast_add, Nat.cast_two, Chebyshev.S_add_two, + dickson_two_one_eq_chebyshev_S (n + 1), dickson_two_one_eq_chebyshev_S n] + push_cast + ring + +theorem dickson_two_one_eq_chebyshev_U [Invertible (2 : R)] (n : ℕ) : + dickson 2 (1 : R) n = (Chebyshev.U R n).comp (C (⅟ 2) * X) := + (dickson_two_one_eq_chebyshev_S R n).trans (Chebyshev.S_eq_U_comp_half_mul_X R n) + +theorem chebyshev_U_eq_dickson_two_one (n : ℕ) : + Chebyshev.U R n = (dickson 2 (1 : R) n).comp (2 * X) := + dickson_two_one_eq_chebyshev_S R n ▸ (Chebyshev.S_comp_two_mul_X R n).symm /-- The `(m * n)`-th Dickson polynomial of the first kind is the composition of the `m`-th and `n`-th. -/ diff --git a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean index 302a9660f7465..806c1a591b0d2 100644 --- a/Mathlib/RingTheory/Polynomial/HilbertPoly.lean +++ b/Mathlib/RingTheory/Polynomial/HilbertPoly.lean @@ -5,6 +5,8 @@ Authors: Fangming Li, Jujian Zhang -/ import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Algebra.Polynomial.Eval.SMul +import Mathlib.Algebra.Polynomial.Roots +import Mathlib.Order.Interval.Set.Infinite import Mathlib.RingTheory.Polynomial.Pochhammer import Mathlib.RingTheory.PowerSeries.WellKnown import Mathlib.Tactic.FieldSimp @@ -17,9 +19,9 @@ given any `p : F[X]` and `d : ℕ`, there exists some `h : F[X]` such that for a `n : ℕ`, `h(n)` is equal to the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. This `h` is unique and is denoted as `Polynomial.hilbertPoly p d`. -For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` is -`Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence -`Polynomial.hilbertPoly (1 : F[X]) d` is the polynomial `(n + 1)···(n + d)/d!`. Note that +For example, given `d : ℕ`, the power series expansion of `1/(1-X)ᵈ⁺¹` in `F[X]` +is `Σₙ ((d + n).choose d)Xⁿ`, which equals `Σₙ ((n + 1)···(n + d)/d!)Xⁿ` and hence +`Polynomial.hilbertPoly (1 : F[X]) (d + 1)` is the polynomial `(n + 1)···(n + d)/d!`. Note that if `d! = 0` in `F`, then the polynomial `(n + 1)···(n + d)/d!` no longer works, so we do not want the characteristic of `F` to be divisible by `d!`. As `Polynomial.hilbertPoly` may take any `p : F[X]` and `d : ℕ` as its inputs, it is necessary for us to assume that `CharZero F`. @@ -56,9 +58,33 @@ equals `(n - k + d).choose d`. noncomputable def preHilbertPoly (d k : ℕ) : F[X] := (d.factorial : F)⁻¹ • ((ascPochhammer F d).comp (Polynomial.X - (C (k : F)) + 1)) +lemma natDegree_preHilbertPoly [CharZero F] (d k : ℕ) : + (preHilbertPoly F d k).natDegree = d := by + have hne : (d ! : F) ≠ 0 := by norm_cast; positivity + rw [preHilbertPoly, natDegree_smul _ (inv_ne_zero hne), natDegree_comp, ascPochhammer_natDegree, + add_comm_sub, ← C_1, ← map_sub, natDegree_add_C, natDegree_X, mul_one] + +lemma coeff_preHilbertPoly_self [CharZero F] (d k : ℕ) : + (preHilbertPoly F d k).coeff d = (d ! : F)⁻¹ := by + delta preHilbertPoly + have hne : (d ! : F) ≠ 0 := by norm_cast; positivity + have heq : d = ((ascPochhammer F d).comp (X - C (k : F) + 1)).natDegree := + (natDegree_preHilbertPoly F d k).symm.trans (natDegree_smul _ (inv_ne_zero hne)) + nth_rw 3 [heq] + calc + _ = (d ! : F)⁻¹ • ((ascPochhammer F d).comp (X - C ((k : F) - 1))).leadingCoeff := by + simp only [sub_add, ← C_1, ← map_sub, coeff_smul, coeff_natDegree] + _ = (d ! : F)⁻¹ := by + simp only [leadingCoeff_comp (ne_of_eq_of_ne (natDegree_X_sub_C _) one_ne_zero), Monic.def.1 + (monic_ascPochhammer _ _), one_mul, leadingCoeff_X_sub_C, one_pow, smul_eq_mul, mul_one] + +lemma leadingCoeff_preHilbertPoly [CharZero F] (d k : ℕ) : + (preHilbertPoly F d k).leadingCoeff = (d ! : F)⁻¹ := by + rw [leadingCoeff, natDegree_preHilbertPoly, coeff_preHilbertPoly_self] + lemma preHilbertPoly_eq_choose_sub_add [CharZero F] (d : ℕ) {k n : ℕ} (hkn : k ≤ n): (preHilbertPoly F d k).eval (n : F) = (n - k + d).choose d := by - have : ((d ! : ℕ) : F) ≠ 0 := by norm_cast; positivity + have : (d ! : F) ≠ 0 := by norm_cast; positivity calc _ = (↑d !)⁻¹ * eval (↑(n - k + 1)) (ascPochhammer F d) := by simp [cast_sub hkn, preHilbertPoly] _ = (n - k + d).choose d := by @@ -73,7 +99,7 @@ is defined as `∑ i in p.support, (p.coeff i) • Polynomial.preHilbertPoly F d `M` is a graded module whose Poincaré series can be written as `p(X)/(1 - X)ᵈ` for some `p : ℚ[X]` with integer coefficients, then `Polynomial.hilbertPoly p d` is the Hilbert polynomial of `M`. See also `Polynomial.coeff_mul_invOneSubPow_eq_hilbertPoly_eval`, -which says that `PowerSeries.coeff F n (p * (PowerSeries.invOneSubPow F d))` equals +which says that `PowerSeries.coeff F n (p * PowerSeries.invOneSubPow F d)` equals `(Polynomial.hilbertPoly p d).eval (n : F)` for any large enough `n : ℕ`. -/ noncomputable def hilbertPoly (p : F[X]) : (d : ℕ) → F[X] @@ -96,14 +122,16 @@ lemma hilbertPoly_X_pow_succ (d k : ℕ) : hilbertPoly ((X : F[X]) ^ k) (d + 1) = preHilbertPoly F d k := by delta hilbertPoly; simp +variable [CharZero F] + /-- The key property of Hilbert polynomials. If `F` is a field with characteristic `0`, `p : F[X]` and `d : ℕ`, then for any large enough `n : ℕ`, `(Polynomial.hilbertPoly p d).eval (n : F)` equals the coefficient of `Xⁿ` in the power series expansion of `p/(1 - X)ᵈ`. -/ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval - [CharZero F] {p : F[X]} (d : ℕ) {n : ℕ} (hn : p.natDegree < n) : - PowerSeries.coeff F n (p * (invOneSubPow F d)) = (hilbertPoly p d).eval (n : F) := by + {p : F[X]} (d : ℕ) {n : ℕ} (hn : p.natDegree < n) : + PowerSeries.coeff F n (p * invOneSubPow F d) = (hilbertPoly p d).eval (n : F) := by delta hilbertPoly; induction d with | zero => simp only [invOneSubPow_zero, Units.val_one, mul_one, coeff_coe, eval_zero] exact coeff_eq_zero_of_natDegree_lt hn @@ -127,4 +155,91 @@ theorem coeff_mul_invOneSubPow_eq_hilbertPoly_eval rw [hx.2, zero_mul] · rw [add_comm, Nat.add_sub_assoc (h_le ⟨x, hx⟩), succ_eq_add_one, add_tsub_cancel_right] +/-- +The polynomial satisfying the key property of `Polynomial.hilbertPoly p d` is unique. +-/ +theorem exists_unique_hilbertPoly (p : F[X]) (d : ℕ) : + ∃! h : F[X], ∃ N : ℕ, ∀ n > N, + PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F) := by + use hilbertPoly p d; constructor + · use p.natDegree + exact fun n => coeff_mul_invOneSubPow_eq_hilbertPoly_eval d + · rintro h ⟨N, hhN⟩ + apply eq_of_infinite_eval_eq h (hilbertPoly p d) + apply ((Set.Ioi_infinite (max N p.natDegree)).image cast_injective.injOn).mono + rintro x ⟨n, hn, rfl⟩ + simp only [Set.mem_Ioi, sup_lt_iff, Set.mem_setOf_eq] at hn ⊢ + rw [← coeff_mul_invOneSubPow_eq_hilbertPoly_eval d hn.2, hhN n hn.1] + +/-- +If `h : F[X]` and there exists some `N : ℕ` such that for any number `n : ℕ` bigger than `N` +we have `PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F)`, then `h` is exactly +`Polynomial.hilbertPoly p d`. +-/ +theorem eq_hilbertPoly_of_forall_coeff_eq_eval + {p h : F[X]} {d : ℕ} (N : ℕ) (hhN : ∀ n > N, + PowerSeries.coeff F n (p * invOneSubPow F d) = h.eval (n : F)) : + h = hilbertPoly p d := + ExistsUnique.unique (exists_unique_hilbertPoly p d) ⟨N, hhN⟩ + ⟨p.natDegree, fun _ x => coeff_mul_invOneSubPow_eq_hilbertPoly_eval d x⟩ + +lemma hilbertPoly_mul_one_sub_succ (p : F[X]) (d : ℕ) : + hilbertPoly (p * (1 - X)) (d + 1) = hilbertPoly p d := by + apply eq_hilbertPoly_of_forall_coeff_eq_eval (p * (1 - X)).natDegree + intro n hn + have heq : 1 - PowerSeries.X = ((1 - X : F[X]) : F⟦X⟧) := by simp only [coe_sub, coe_one, coe_X] + rw [← one_sub_pow_mul_invOneSubPow_val_add_eq_invOneSubPow_val F d 1, pow_one, ← mul_assoc, heq, + ← coe_mul, coeff_mul_invOneSubPow_eq_hilbertPoly_eval (d + 1) hn] + +lemma hilbertPoly_mul_one_sub_pow_add (p : F[X]) (d e : ℕ) : + hilbertPoly (p * (1 - X) ^ e) (d + e) = hilbertPoly p d := by + induction e with + | zero => simp + | succ e he => rw [pow_add, pow_one, ← mul_assoc, ← add_assoc, hilbertPoly_mul_one_sub_succ, he] + +lemma hilbertPoly_eq_zero_of_le_rootMultiplicity_one + {p : F[X]} {d : ℕ} (hdp : d ≤ p.rootMultiplicity 1) : + hilbertPoly p d = 0 := by + by_cases hp : p = 0 + · rw [hp, hilbertPoly_zero_nat] + · rcases exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp 1 with ⟨q, hq1, hq2⟩ + have heq : p = q * (- 1) ^ p.rootMultiplicity 1 * (1 - X) ^ p.rootMultiplicity 1 := by + simp only [mul_assoc, ← mul_pow, neg_mul, one_mul, neg_sub] + exact hq1.trans (mul_comm _ _) + rw [heq, ← zero_add d, ← Nat.sub_add_cancel hdp, pow_add (1 - X), ← mul_assoc, + hilbertPoly_mul_one_sub_pow_add, hilbertPoly] + +theorem natDegree_hilbertPoly_of_ne_zero_of_rootMultiplicity_lt + {p : F[X]} {d : ℕ} (hp : p ≠ 0) (hpd : p.rootMultiplicity 1 < d) : + (hilbertPoly p d).natDegree = d - p.rootMultiplicity 1 - 1 := by + rcases exists_eq_pow_rootMultiplicity_mul_and_not_dvd p hp 1 with ⟨q, hq1, hq2⟩ + have heq : p = q * (- 1) ^ p.rootMultiplicity 1 * (1 - X) ^ p.rootMultiplicity 1 := by + simp only [mul_assoc, ← mul_pow, neg_mul, one_mul, neg_sub] + exact hq1.trans (mul_comm _ _) + nth_rw 1 [heq, ← Nat.sub_add_cancel (le_of_lt hpd), hilbertPoly_mul_one_sub_pow_add, + ← Nat.sub_add_cancel (Nat.le_sub_of_add_le' <| add_one_le_of_lt hpd)] + delta hilbertPoly + apply natDegree_eq_of_le_of_coeff_ne_zero + · apply natDegree_sum_le_of_forall_le _ _ <| fun _ _ => ?_ + apply le_trans (natDegree_smul_le _ _) + rw [natDegree_preHilbertPoly] + · have : (fun (x : ℕ) (a : F) => a) = fun x a => a * 1 ^ x := by simp only [one_pow, mul_one] + simp only [finset_sum_coeff, coeff_smul, smul_eq_mul, coeff_preHilbertPoly_self, + ← Finset.sum_mul, ← sum_def _ (fun _ a => a), this, ← eval_eq_sum, eval_mul, eval_pow, + eval_neg, eval_one, _root_.mul_eq_zero, pow_eq_zero_iff', neg_eq_zero, one_ne_zero, ne_eq, + false_and, or_false, inv_eq_zero, cast_eq_zero, not_or] + exact ⟨(not_iff_not.2 dvd_iff_isRoot).1 hq2, factorial_ne_zero _⟩ + +theorem natDegree_hilbertPoly_of_ne_zero + {p : F[X]} {d : ℕ} (hh : hilbertPoly p d ≠ 0) : + (hilbertPoly p d).natDegree = d - p.rootMultiplicity 1 - 1 := by + have hp : p ≠ 0 := by + intro h + rw [h] at hh + exact hh (hilbertPoly_zero_nat F d) + have hpd : p.rootMultiplicity 1 < d := by + by_contra h + exact hh (hilbertPoly_eq_zero_of_le_rootMultiplicity_one <| not_lt.1 h) + exact natDegree_hilbertPoly_of_ne_zero_of_rootMultiplicity_lt hp hpd + end Polynomial diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index 70348c70fd666..381b894105712 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -38,8 +38,6 @@ variable {R : Type*} section OrderBasic -open multiplicity - variable [Semiring R] {φ : R⟦X⟧} theorem exists_coeff_ne_zero_iff_ne_zero : (∃ n : ℕ, coeff R n φ ≠ 0) ↔ φ ≠ 0 := by diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index ca5bd0b442330..783e6348c6aa9 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -159,6 +159,18 @@ theorem mk_add_choose_mul_one_sub_pow_eq_one : (mk fun n ↦ Nat.choose (d + n) d : S⟦X⟧) * ((1 - X) ^ (d + 1)) = 1 := (invOneSubPow S (d + 1)).val_inv +theorem invOneSubPow_add (e : ℕ) : + invOneSubPow S (d + e) = invOneSubPow S d * invOneSubPow S e := by + simp_rw [invOneSubPow_eq_inv_one_sub_pow, pow_add] + +theorem one_sub_pow_mul_invOneSubPow_val_add_eq_invOneSubPow_val (e : ℕ) : + (1 - X) ^ e * (invOneSubPow S (d + e)).val = (invOneSubPow S d).val := by + simp [invOneSubPow_add, Units.val_mul, mul_comm, mul_assoc, ← invOneSubPow_inv_eq_one_sub_pow] + +theorem one_sub_pow_add_mul_invOneSubPow_val_eq_one_sub_pow (e : ℕ) : + (1 - X) ^ (d + e) * (invOneSubPow S e).val = (1 - X) ^ d := by + simp [pow_add, mul_assoc, ← invOneSubPow_inv_eq_one_sub_pow S e] + end invOneSubPow section Field diff --git a/Mathlib/RingTheory/RingHom/Flat.lean b/Mathlib/RingTheory/RingHom/Flat.lean new file mode 100644 index 0000000000000..ee70f2de16d6b --- /dev/null +++ b/Mathlib/RingTheory/RingHom/Flat.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.RingTheory.Flat.Localization +import Mathlib.RingTheory.LocalProperties.Basic + +/-! +# Flat ring homomorphisms + +In this file we define flat ring homomorphisms and show their meta properties. + +-/ + +universe u v + +open TensorProduct + +/-- A ring homomorphism `f : R →+* S` is flat if `S` is flat as an `R` module. -/ +@[algebraize Module.Flat] +def RingHom.Flat {R : Type u} {S : Type v} [CommRing R] [CommRing S] (f : R →+* S) : Prop := + letI : Algebra R S := f.toAlgebra + Module.Flat R S + +namespace RingHom.Flat + +variable {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] + +variable (R) in +/-- The identity of a ring is flat. -/ +lemma id : RingHom.Flat (RingHom.id R) := + Module.Flat.self R + +/-- Composition of flat ring homomorphisms is flat. -/ +lemma comp {f : R →+* S} {g : S →+* T} (hf : f.Flat) (hg : g.Flat) : Flat (g.comp f) := by + algebraize [f, g, (g.comp f)] + exact Module.Flat.trans R S T + +/-- Bijective ring maps are flat. -/ +lemma of_bijective {f : R →+* S} (hf : Function.Bijective f) : Flat f := by + algebraize [f] + exact Module.Flat.of_linearEquiv R R S (LinearEquiv.ofBijective (Algebra.linearMap R S) hf).symm + +lemma containsIdentities : ContainsIdentities Flat := id + +lemma stableUnderComposition : StableUnderComposition Flat := by + introv R hf hg + exact hf.comp hg + +lemma respectsIso : RespectsIso Flat := by + apply stableUnderComposition.respectsIso + introv + exact of_bijective e.bijective + +lemma isStableUnderBaseChange : IsStableUnderBaseChange Flat := by + apply IsStableUnderBaseChange.mk _ respectsIso + introv h + replace h : Module.Flat R T := by + rw [RingHom.Flat] at h; convert h; ext; simp_rw [Algebra.smul_def]; rfl + suffices Module.Flat S (S ⊗[R] T) by + rw [RingHom.Flat]; convert this; congr; ext; simp_rw [Algebra.smul_def]; rfl + exact inferInstance + +lemma holdsForLocalizationAway : HoldsForLocalizationAway Flat := by + introv R h + suffices Module.Flat R S by + rw [RingHom.Flat]; convert this; ext; simp_rw [Algebra.smul_def]; rfl + exact IsLocalization.flat _ (Submonoid.powers r) + +end RingHom.Flat diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index df43e22b8bed7..2da9825d08c25 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -150,7 +150,7 @@ theorem minpoly_eq_pow {p : ℕ} [hprime : Fact p.Prime] (hdiv : ¬p ∣ n) : (separable_X_pow_sub_C 1 (fun h => hdiv <| (ZMod.natCast_zmod_eq_zero_iff_dvd n p).1 h) one_ne_zero).squarefree cases' - (multiplicity.squarefree_iff_emultiplicity_le_one (X ^ n - 1)).1 hfree + (squarefree_iff_emultiplicity_le_one (X ^ n - 1)).1 hfree (map (Int.castRingHom (ZMod p)) P) with hle hunit · rw [Nat.cast_one] at habs; exact hle.not_lt habs diff --git a/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean b/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean index 8399113e64d64..f1f8616f7cb18 100644 --- a/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean +++ b/Mathlib/RingTheory/RootsOfUnity/PrimitiveRoots.lean @@ -363,7 +363,7 @@ theorem primitiveRoots_one : primitiveRoots 1 R = {(1 : R)} := by theorem neZero' {n : ℕ} [NeZero n] (hζ : IsPrimitiveRoot ζ n) : NeZero ((n : ℕ) : R) := by let p := ringChar R - have hfin := Nat.multiplicity_finite_iff.2 ⟨CharP.char_ne_one R p, NeZero.pos n⟩ + have hfin := Nat.finiteMultiplicity_iff.2 ⟨CharP.char_ne_one R p, NeZero.pos n⟩ obtain ⟨m, hm⟩ := hfin.exists_eq_pow_mul_and_not_dvd by_cases hp : p ∣ n · obtain ⟨k, hk⟩ := Nat.exists_eq_succ_of_ne_zero (multiplicity_pos_of_dvd hp).ne' diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean index dbc52cb56fad1..a6504ae1d7bda 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -33,11 +33,21 @@ theorem WfDvdMonoid.max_power_factor [CommMonoidWithZero α] [WfDvdMonoid α] {a (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := max_power_factor' h hx.not_unit -theorem multiplicity.finite_of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] - {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : multiplicity.Finite a b := by +theorem FiniteMultiplicity.of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] + {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : FiniteMultiplicity a b := by obtain ⟨n, c, ndvd, rfl⟩ := WfDvdMonoid.max_power_factor' hb ha exact ⟨n, by rwa [pow_succ, mul_dvd_mul_iff_left (left_ne_zero_of_mul hb)]⟩ +@[deprecated (since := "2024-11-30")] +alias multiplicity.finite_of_not_isUnit := FiniteMultiplicity.of_not_isUnit + +theorem FiniteMultiplicity.of_prime_left [CancelCommMonoidWithZero α] [WfDvdMonoid α] + {a b : α} (ha : Prime a) (hb : b ≠ 0) : FiniteMultiplicity a b := + .of_not_isUnit ha.not_unit hb + +@[deprecated (since := "2024-11-30")] +alias multiplicity.finite_prime_left := FiniteMultiplicity.of_prime_left + namespace UniqueFactorizationMonoid variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] diff --git a/Mathlib/RingTheory/Valuation/LocalSubring.lean b/Mathlib/RingTheory/Valuation/LocalSubring.lean index a3c86e708bac1..ae5fef7208b40 100644 --- a/Mathlib/RingTheory/Valuation/LocalSubring.lean +++ b/Mathlib/RingTheory/Valuation/LocalSubring.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Andrew Yang, Yaël Dillies, Javier López-Contreras. All righ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, Yaël Dillies, Javier López-Contreras -/ -import Mathlib.RingTheory.LocalRing.Subring import Mathlib.RingTheory.Ideal.Over +import Mathlib.RingTheory.LocalRing.Subring +import Mathlib.RingTheory.Polynomial.Ideal import Mathlib.RingTheory.Valuation.ValuationSubring /-! diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index ae9eeee389945..52eabc3d86978 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -504,7 +504,6 @@ theorem lift_eq_aleph1 {c : Cardinal.{u}} : lift.{v} c = ℵ₁ ↔ c = ℵ₁ : section deprecated -set_option linter.deprecated false set_option linter.docPrime false @[deprecated preAleph (since := "2024-10-22")] @@ -522,6 +521,7 @@ noncomputable alias aleph' := preAleph def alephIdx.initialSeg : @InitialSeg Cardinal Ordinal (· < ·) (· < ·) := @RelEmbedding.collapse Cardinal Ordinal (· < ·) (· < ·) _ Cardinal.ord.orderEmbedding.ltEmbedding +set_option linter.deprecated false in /-- The `aleph'` index function, which gives the ordinal index of a cardinal. (The `aleph'` part is because unlike `aleph` this counts also the finite stages. So `alephIdx n = n`, `alephIdx ℵ₀ = ω`, @@ -533,6 +533,7 @@ def alephIdx.initialSeg : @InitialSeg Cardinal Ordinal (· < ·) (· < ·) := def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := aleph'.symm.toRelIsoLT +set_option linter.deprecated false in /-- The `aleph'` index function, which gives the ordinal index of a cardinal. (The `aleph'` part is because unlike `aleph` this counts also the finite stages. So `alephIdx n = n`, `alephIdx ω = ω`, @@ -542,10 +543,12 @@ def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := def alephIdx : Cardinal → Ordinal := aleph'.symm +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := rfl +set_option linter.deprecated false in /-- The `aleph'` function gives the cardinals listed by their ordinal index, and is the inverse of `aleph_idx`. `aleph' n = n`, `aleph' ω = ω`, `aleph' (ω + 1) = succ ℵ₀`, etc. @@ -556,55 +559,68 @@ theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephId def Aleph'.relIso := aleph' +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := rfl +set_option linter.deprecated false in @[deprecated preAleph_lt_preAleph (since := "2024-10-22")] theorem aleph'_lt {o₁ o₂ : Ordinal} : aleph' o₁ < aleph' o₂ ↔ o₁ < o₂ := aleph'.lt_iff_lt +set_option linter.deprecated false in @[deprecated preAleph_le_preAleph (since := "2024-10-22")] theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ ≤ o₂ := aleph'.le_iff_le +set_option linter.deprecated false in @[deprecated preAleph_max (since := "2024-10-22")] theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := aleph'.monotone.map_max +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o +set_option linter.deprecated false in @[deprecated preAleph_zero (since := "2024-10-22")] theorem aleph'_zero : aleph' 0 = 0 := aleph'.map_bot +set_option linter.deprecated false in @[deprecated preAleph_succ (since := "2024-10-22")] theorem aleph'_succ (o : Ordinal) : aleph' (succ o) = succ (aleph' o) := aleph'.map_succ o +set_option linter.deprecated false in @[deprecated preAleph_nat (since := "2024-10-22")] theorem aleph'_nat : ∀ n : ℕ, aleph' n = n := preAleph_nat +set_option linter.deprecated false in @[deprecated lift_preAleph (since := "2024-10-22")] theorem lift_aleph' (o : Ordinal.{u}) : lift.{v} (aleph' o) = aleph' (Ordinal.lift.{v} o) := lift_preAleph o +set_option linter.deprecated false in @[deprecated preAleph_le_of_isLimit (since := "2024-10-22")] theorem aleph'_le_of_limit {o : Ordinal} (l : o.IsLimit) {c} : aleph' o ≤ c ↔ ∀ o' < o, aleph' o' ≤ c := preAleph_le_of_isLimit l +set_option linter.deprecated false in @[deprecated preAleph_limit (since := "2024-10-22")] theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, aleph' a := preAleph_limit ho +set_option linter.deprecated false in @[deprecated preAleph_omega0 (since := "2024-10-22")] theorem aleph'_omega0 : aleph' ω = ℵ₀ := preAleph_omega0 @@ -612,6 +628,7 @@ theorem aleph'_omega0 : aleph' ω = ℵ₀ := @[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias aleph'_omega := aleph'_omega0 +set_option linter.deprecated false in /-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ @[deprecated aleph' (since := "2024-08-28")] def aleph'Equiv : Ordinal ≃ Cardinal := @@ -621,14 +638,17 @@ def aleph'Equiv : Ordinal ≃ Cardinal := theorem aleph_eq_aleph' (o : Ordinal) : ℵ_ o = preAleph (ω + o) := rfl +set_option linter.deprecated false in @[deprecated aleph0_le_preAleph (since := "2024-10-22")] theorem aleph0_le_aleph' {o : Ordinal} : ℵ₀ ≤ aleph' o ↔ ω ≤ o := by rw [← aleph'_omega0, aleph'_le] +set_option linter.deprecated false in @[deprecated preAleph_pos (since := "2024-10-22")] theorem aleph'_pos {o : Ordinal} (ho : 0 < o) : 0 < aleph' o := by rwa [← aleph'_zero, aleph'_lt] +set_option linter.deprecated false in @[deprecated preAleph_isNormal (since := "2024-10-22")] theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := preAleph_isNormal @@ -647,15 +667,18 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ +set_option linter.deprecated false in /-- Infinite ordinals that are cardinals are unbounded. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (ℵ_ a).ord = o := by diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index 0c071383f0154..8e9e3236e07e2 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -82,7 +82,7 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by · exact (mul_lt_aleph0 qo qo).trans_le ol · suffices (succ (typein LT.lt (g p))).card < #α from (IH _ this qo).trans_lt this rw [← lt_ord] - apply (isLimit_ord ol).2 + apply (isLimit_ord ol).succ_lt rw [e] apply typein_lt_type diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index d4660d1f437be..12ba9cbde683e 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -177,54 +177,6 @@ def map₂ (f : Type u → Type v → Type w) (hf : ∀ α β γ δ, α ≃ β Cardinal.{u} → Cardinal.{v} → Cardinal.{w} := Quotient.map₂ f fun α β ⟨e₁⟩ γ δ ⟨e₂⟩ => ⟨hf α β γ δ e₁ e₂⟩ -/-- We define the order on cardinal numbers by `#α ≤ #β` if and only if - there exists an embedding (injective function) from α to β. -/ -instance : LE Cardinal.{u} := - ⟨fun q₁ q₂ => - Quotient.liftOn₂ q₁ q₂ (fun α β => Nonempty <| α ↪ β) fun _ _ _ _ ⟨e₁⟩ ⟨e₂⟩ => - propext ⟨fun ⟨e⟩ => ⟨e.congr e₁ e₂⟩, fun ⟨e⟩ => ⟨e.congr e₁.symm e₂.symm⟩⟩⟩ - -instance partialOrder : PartialOrder Cardinal.{u} where - le := (· ≤ ·) - le_refl := by - rintro ⟨α⟩ - exact ⟨Embedding.refl _⟩ - le_trans := by - rintro ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩ - exact ⟨e₁.trans e₂⟩ - le_antisymm := by - rintro ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩ - exact Quotient.sound (e₁.antisymm e₂) - -instance linearOrder : LinearOrder Cardinal.{u} := - { Cardinal.partialOrder with - le_total := by - rintro ⟨α⟩ ⟨β⟩ - apply Embedding.total - decidableLE := Classical.decRel _ } - -theorem le_def (α β : Type u) : #α ≤ #β ↔ Nonempty (α ↪ β) := - Iff.rfl - -theorem mk_le_of_injective {α β : Type u} {f : α → β} (hf : Injective f) : #α ≤ #β := - ⟨⟨f, hf⟩⟩ - -theorem _root_.Function.Embedding.cardinal_le {α β : Type u} (f : α ↪ β) : #α ≤ #β := - ⟨f⟩ - -theorem mk_le_of_surjective {α β : Type u} {f : α → β} (hf : Surjective f) : #β ≤ #α := - ⟨Embedding.ofSurjective f hf⟩ - -theorem le_mk_iff_exists_set {c : Cardinal} {α : Type u} : c ≤ #α ↔ ∃ p : Set α, #p = c := - ⟨inductionOn c fun _ ⟨⟨f, hf⟩⟩ => ⟨Set.range f, (Equiv.ofInjective f hf).cardinal_eq.symm⟩, - fun ⟨_, e⟩ => e ▸ ⟨⟨Subtype.val, fun _ _ => Subtype.eq⟩⟩⟩ - -theorem mk_subtype_le {α : Type u} (p : α → Prop) : #(Subtype p) ≤ #α := - ⟨Embedding.subtype p⟩ - -theorem mk_set_le (s : Set α) : #s ≤ #α := - mk_subtype_le s - /-! ### Lifting cardinals to a higher universe -/ /-- The universe lift operation on cardinals. You can specify the universes explicitly with @@ -281,21 +233,6 @@ lemma mk_preimage_down {s : Set α} : #(ULift.down.{v} ⁻¹' s) = lift.{v} (#s) ULift.up_bijective.comp (restrictPreimage_bijective _ (ULift.down_bijective)) exact Equiv.ofBijective f this -theorem out_embedding {c c' : Cardinal} : c ≤ c' ↔ Nonempty (c.out ↪ c'.out) := by - conv_lhs => rw [← Cardinal.mk_out c, ← Cardinal.mk_out c', le_def] - -theorem lift_mk_le {α : Type v} {β : Type w} : - lift.{max u w} #α ≤ lift.{max u v} #β ↔ Nonempty (α ↪ β) := - ⟨fun ⟨f⟩ => ⟨Embedding.congr Equiv.ulift Equiv.ulift f⟩, fun ⟨f⟩ => - ⟨Embedding.congr Equiv.ulift.symm Equiv.ulift.symm f⟩⟩ - -/-- A variant of `Cardinal.lift_mk_le` with specialized universes. -Because Lean often can not realize it should use this specialization itself, -we provide this statement separately so you don't have to solve the specialization problem either. --/ -theorem lift_mk_le' {α : Type u} {β : Type v} : lift.{v} #α ≤ lift.{u} #β ↔ Nonempty (α ↪ β) := - lift_mk_le.{0} - theorem lift_mk_eq {α : Type u} {β : Type v} : lift.{max v w} #α = lift.{max u w} #β ↔ Nonempty (α ≃ β) := Quotient.eq'.trans @@ -325,6 +262,73 @@ theorem lift_mk_shrink'' (α : Type max u v) [Small.{v} α] : Cardinal.lift.{u} #(Shrink.{v} α) = #α := by rw [← lift_umax, lift_mk_shrink.{max u v, v, 0} α, ← lift_umax, lift_id] +/-! ### Order on cardinals -/ + +/-- We define the order on cardinal numbers by `#α ≤ #β` if and only if + there exists an embedding (injective function) from α to β. -/ +instance : LE Cardinal.{u} := + ⟨fun q₁ q₂ => + Quotient.liftOn₂ q₁ q₂ (fun α β => Nonempty <| α ↪ β) fun _ _ _ _ ⟨e₁⟩ ⟨e₂⟩ => + propext ⟨fun ⟨e⟩ => ⟨e.congr e₁ e₂⟩, fun ⟨e⟩ => ⟨e.congr e₁.symm e₂.symm⟩⟩⟩ + +instance partialOrder : PartialOrder Cardinal.{u} where + le := (· ≤ ·) + le_refl := by + rintro ⟨α⟩ + exact ⟨Embedding.refl _⟩ + le_trans := by + rintro ⟨α⟩ ⟨β⟩ ⟨γ⟩ ⟨e₁⟩ ⟨e₂⟩ + exact ⟨e₁.trans e₂⟩ + le_antisymm := by + rintro ⟨α⟩ ⟨β⟩ ⟨e₁⟩ ⟨e₂⟩ + exact Quotient.sound (e₁.antisymm e₂) + +instance linearOrder : LinearOrder Cardinal.{u} := + { Cardinal.partialOrder with + le_total := by + rintro ⟨α⟩ ⟨β⟩ + apply Embedding.total + decidableLE := Classical.decRel _ } + +theorem le_def (α β : Type u) : #α ≤ #β ↔ Nonempty (α ↪ β) := + Iff.rfl + +theorem mk_le_of_injective {α β : Type u} {f : α → β} (hf : Injective f) : #α ≤ #β := + ⟨⟨f, hf⟩⟩ + +theorem _root_.Function.Embedding.cardinal_le {α β : Type u} (f : α ↪ β) : #α ≤ #β := + ⟨f⟩ + +theorem mk_le_of_surjective {α β : Type u} {f : α → β} (hf : Surjective f) : #β ≤ #α := + ⟨Embedding.ofSurjective f hf⟩ + +theorem le_mk_iff_exists_set {c : Cardinal} {α : Type u} : c ≤ #α ↔ ∃ p : Set α, #p = c := + ⟨inductionOn c fun _ ⟨⟨f, hf⟩⟩ => ⟨Set.range f, (Equiv.ofInjective f hf).cardinal_eq.symm⟩, + fun ⟨_, e⟩ => e ▸ ⟨⟨Subtype.val, fun _ _ => Subtype.eq⟩⟩⟩ + +theorem mk_subtype_le {α : Type u} (p : α → Prop) : #(Subtype p) ≤ #α := + ⟨Embedding.subtype p⟩ + +theorem mk_set_le (s : Set α) : #s ≤ #α := + mk_subtype_le s + +theorem out_embedding {c c' : Cardinal} : c ≤ c' ↔ Nonempty (c.out ↪ c'.out) := by + conv_lhs => rw [← Cardinal.mk_out c, ← Cardinal.mk_out c', le_def] + +theorem lift_mk_le {α : Type v} {β : Type w} : + lift.{max u w} #α ≤ lift.{max u v} #β ↔ Nonempty (α ↪ β) := + ⟨fun ⟨f⟩ => ⟨Embedding.congr Equiv.ulift Equiv.ulift f⟩, fun ⟨f⟩ => + ⟨Embedding.congr Equiv.ulift.symm Equiv.ulift.symm f⟩⟩ + +/-- A variant of `Cardinal.lift_mk_le` with specialized universes. +Because Lean often can not realize it should use this specialization itself, +we provide this statement separately so you don't have to solve the specialization problem either. +-/ +theorem lift_mk_le' {α : Type u} {β : Type v} : lift.{v} #α ≤ lift.{u} #β ↔ Nonempty (α ↪ β) := + lift_mk_le.{0} + +/-! ### `lift` sends `Cardinal.{u}` to an initial segment of `Cardinal.{max u v}`. -/ + /-- `Cardinal.lift` as an `InitialSeg`. -/ @[simps!] def liftInitialSeg : Cardinal.{u} ≤i Cardinal.{max u v} := by @@ -815,19 +819,21 @@ protected theorem isSuccLimit_iff {c : Cardinal} : IsSuccLimit c ↔ c ≠ 0 ∧ section deprecated -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated IsSuccLimit.isSuccPrelimit (since := "2024-09-17")] protected theorem IsLimit.isSuccPrelimit {c} (h : IsLimit c) : IsSuccPrelimit c := h.2 +set_option linter.deprecated false in @[deprecated ne_zero_of_isSuccLimit (since := "2024-09-17")] protected theorem IsLimit.ne_zero {c} (h : IsLimit c) : c ≠ 0 := h.1 +set_option linter.deprecated false in @[deprecated IsLimit.isSuccPrelimit (since := "2024-09-05")] alias IsLimit.isSuccLimit := IsLimit.isSuccPrelimit +set_option linter.deprecated false in @[deprecated IsSuccLimit.succ_lt (since := "2024-09-17")] theorem IsLimit.succ_lt {x c} (h : IsLimit c) : x < c → succ x < c := h.isSuccPrelimit.succ_lt @@ -1519,17 +1525,18 @@ theorem aleph0_le_of_isSuccLimit {c : Cardinal} (h : IsSuccLimit c) : ℵ₀ ≤ section deprecated -set_option linter.deprecated false - +set_option linter.deprecated false in @[deprecated isSuccLimit_aleph0 (since := "2024-09-17")] theorem isLimit_aleph0 : IsLimit ℵ₀ := ⟨aleph0_ne_zero, isSuccPrelimit_aleph0⟩ +set_option linter.deprecated false in @[deprecated not_isSuccLimit_natCast (since := "2024-09-17")] lemma not_isLimit_natCast : (n : ℕ) → ¬ IsLimit (n : Cardinal.{u}) | 0, e => e.1 rfl | Nat.succ n, e => Order.not_isSuccPrelimit_succ _ (nat_succ n ▸ e.2) +set_option linter.deprecated false in @[deprecated aleph0_le_of_isSuccLimit (since := "2024-09-17")] theorem IsLimit.aleph0_le {c : Cardinal} (h : IsLimit c) : ℵ₀ ≤ c := by by_contra! h' diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index df29d51ac8b46..786a594aa7434 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -696,7 +696,7 @@ theorem cof_eq' (r : α → α → Prop) [IsWellOrder α r] (h : IsLimit (type r ∃ S : Set α, (∀ a, ∃ b ∈ S, r a b) ∧ #S = cof (type r) := let ⟨S, H, e⟩ := cof_eq r ⟨S, fun a => - let a' := enum r ⟨_, h.2 _ (typein_lt_type r a)⟩ + let a' := enum r ⟨_, h.succ_lt (typein_lt_type r a)⟩ let ⟨b, h, ab⟩ := H a' ⟨b, h, (IsOrderConnected.conn a b a' <| @@ -834,11 +834,13 @@ theorem isStrongLimit_beth {o : Ordinal} (H : IsSuccPrelimit o) : IsStrongLimit · rw [beth_zero] exact isStrongLimit_aleph0 · refine ⟨beth_ne_zero o, fun a ha => ?_⟩ - rw [beth_limit ⟨h, isSuccPrelimit_iff_succ_lt.1 H⟩] at ha - rcases exists_lt_of_lt_ciSup' ha with ⟨⟨i, hi⟩, ha⟩ - have := power_le_power_left two_ne_zero ha.le - rw [← beth_succ] at this - exact this.trans_lt (beth_lt.2 (H.succ_lt hi)) + rw [beth_limit] at ha + · rcases exists_lt_of_lt_ciSup' ha with ⟨⟨i, hi⟩, ha⟩ + have := power_le_power_left two_ne_zero ha.le + rw [← beth_succ] at this + exact this.trans_lt (beth_lt.2 (H.succ_lt hi)) + · rw [isLimit_iff] + exact ⟨h, H⟩ theorem mk_bounded_subset {α : Type*} (h : ∀ x < #α, (2^x) < #α) {r : α → α → Prop} [IsWellOrder α r] (hr : (#α).ord = type r) : #{ s : Set α // Bounded r s } = #α := by @@ -1140,7 +1142,7 @@ theorem derivFamily_lt_ord_lift {ι : Type u} {f : ι → Ordinal → Ordinal} { rw [derivFamily_succ] exact nfpFamily_lt_ord_lift hω (by rwa [hc.cof_eq]) hf - ((isLimit_ord hc.1).2 _ (hb ((lt_succ b).trans hb'))) + ((isLimit_ord hc.1).succ_lt (hb ((lt_succ b).trans hb'))) | H₃ b hb H => intro hb' -- TODO: generalize the universes of the lemmas in this file so we don't have to rely on bsup diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 3563908538d09..b01376da2f420 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -3,10 +3,11 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn, Violeta Hernández Palacios -/ -import Mathlib.SetTheory.Ordinal.Basic -import Mathlib.Data.Nat.SuccPred import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Data.Nat.SuccPred +import Mathlib.Order.SuccPred.InitialSeg import Mathlib.SetTheory.Cardinal.UnivLE +import Mathlib.SetTheory.Ordinal.Basic /-! # Ordinal arithmetic @@ -197,33 +198,36 @@ theorem lift_pred (o : Ordinal.{v}) : lift.{u} (pred o) = pred (lift.{u} o) := b TODO: deprecate this in favor of `Order.IsSuccLimit`. -/ def IsLimit (o : Ordinal) : Prop := - o ≠ 0 ∧ ∀ a < o, succ a < o + IsSuccLimit o + +theorem isLimit_iff {o} : IsLimit o ↔ o ≠ 0 ∧ IsSuccPrelimit o := by + simp [IsLimit, IsSuccLimit] theorem IsLimit.isSuccPrelimit {o} (h : IsLimit o) : IsSuccPrelimit o := - isSuccPrelimit_iff_succ_lt.mpr h.2 + IsSuccLimit.isSuccPrelimit h @[deprecated IsLimit.isSuccPrelimit (since := "2024-09-05")] alias IsLimit.isSuccLimit := IsLimit.isSuccPrelimit theorem IsLimit.succ_lt {o a : Ordinal} (h : IsLimit o) : a < o → succ a < o := - h.2 a + IsSuccLimit.succ_lt h theorem isSuccPrelimit_zero : IsSuccPrelimit (0 : Ordinal) := isSuccPrelimit_bot @[deprecated isSuccPrelimit_zero (since := "2024-09-05")] alias isSuccLimit_zero := isSuccPrelimit_zero -theorem not_zero_isLimit : ¬IsLimit 0 - | ⟨h, _⟩ => h rfl +theorem not_zero_isLimit : ¬IsLimit 0 := + not_isSuccLimit_bot -theorem not_succ_isLimit (o) : ¬IsLimit (succ o) - | ⟨_, h⟩ => lt_irrefl _ (h _ (lt_succ o)) +theorem not_succ_isLimit (o) : ¬IsLimit (succ o) := + not_isSuccLimit_succ o theorem not_succ_of_isLimit {o} (h : IsLimit o) : ¬∃ a, o = succ a | ⟨a, e⟩ => not_succ_isLimit a (e ▸ h) theorem succ_lt_of_isLimit {o a : Ordinal} (h : IsLimit o) : succ a < o ↔ a < o := - ⟨(lt_succ a).trans, h.2 _⟩ + IsSuccLimit.succ_lt_iff h theorem le_succ_of_isLimit {o} (h : IsLimit o) {a} : o ≤ succ a ↔ o ≤ a := le_iff_le_iff_lt_iff_lt.2 <| succ_lt_of_isLimit h @@ -238,29 +242,23 @@ theorem lt_limit {o} (h : IsLimit o) {a} : a < o ↔ ∃ x < o, a < x := by @[simp] theorem lift_isLimit (o : Ordinal.{v}) : IsLimit (lift.{u,v} o) ↔ IsLimit o := - and_congr (not_congr <| by simpa only [lift_zero] using @lift_inj o 0) - ⟨fun H a h => (lift_lt.{u,v}).1 <| - by simpa only [lift_succ] using H _ (lift_lt.2 h), fun H a h => by - obtain ⟨a', rfl⟩ := mem_range_lift_of_le h.le - rw [← lift_succ, lift_lt] - exact H a' (lift_lt.1 h)⟩ + liftInitialSeg.isSuccLimit_apply_iff theorem IsLimit.pos {o : Ordinal} (h : IsLimit o) : 0 < o := - lt_of_le_of_ne (Ordinal.zero_le _) h.1.symm + IsSuccLimit.bot_lt h + +theorem IsLimit.ne_zero {o : Ordinal} (h : IsLimit o) : o ≠ 0 := + h.pos.ne' theorem IsLimit.one_lt {o : Ordinal} (h : IsLimit o) : 1 < o := by - simpa only [succ_zero] using h.2 _ h.pos + simpa only [succ_zero] using h.succ_lt h.pos theorem IsLimit.nat_lt {o : Ordinal} (h : IsLimit o) : ∀ n : ℕ, (n : Ordinal) < o | 0 => h.pos - | n + 1 => h.2 _ (IsLimit.nat_lt h n) + | n + 1 => h.succ_lt (IsLimit.nat_lt h n) theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ IsLimit o := by - classical - exact if o0 : o = 0 then Or.inl o0 - else - if h : ∃ a, o = succ a then Or.inr (Or.inl h) - else Or.inr <| Or.inr ⟨o0, fun _a => (succ_lt_of_not_succ h).2⟩ + simpa [eq_comm] using isMin_or_mem_range_succ_or_isSuccLimit o theorem isLimit_of_not_succ_of_ne_zero {o : Ordinal} (h : ¬∃ a, o = succ a) (h' : o ≠ 0) : IsLimit o := ((zero_or_succ_or_limit o).resolve_left h').resolve_left h @@ -279,23 +277,24 @@ theorem IsLimit.iSup_Iio {o : Ordinal} (h : IsLimit o) : ⨆ a : Iio o, a.1 = o induction at successor ordinals and at limit ordinals, then it holds for all ordinals. -/ @[elab_as_elim] def limitRecOn {C : Ordinal → Sort*} (o : Ordinal) (H₁ : C 0) (H₂ : ∀ o, C o → C (succ o)) - (H₃ : ∀ o, IsLimit o → (∀ o' < o, C o') → C o) : C o := - SuccOrder.prelimitRecOn o (fun o _ ↦ H₂ o) fun o hl ↦ - if h : o = 0 then fun _ ↦ h ▸ H₁ else H₃ o ⟨h, fun _ ↦ hl.succ_lt⟩ + (H₃ : ∀ o, IsLimit o → (∀ o' < o, C o') → C o) : C o := by + refine SuccOrder.limitRecOn o (fun a ha ↦ ?_) (fun a _ ↦ H₂ a) H₃ + convert H₁ + simpa using ha @[simp] -theorem limitRecOn_zero {C} (H₁ H₂ H₃) : @limitRecOn C 0 H₁ H₂ H₃ = H₁ := by - rw [limitRecOn, SuccOrder.prelimitRecOn_of_isSuccPrelimit _ _ isSuccPrelimit_zero, dif_pos rfl] +theorem limitRecOn_zero {C} (H₁ H₂ H₃) : @limitRecOn C 0 H₁ H₂ H₃ = H₁ := + SuccOrder.limitRecOn_isMin _ _ _ isMin_bot @[simp] theorem limitRecOn_succ {C} (o H₁ H₂ H₃) : - @limitRecOn C (succ o) H₁ H₂ H₃ = H₂ o (@limitRecOn C o H₁ H₂ H₃) := by - rw [limitRecOn, limitRecOn, SuccOrder.prelimitRecOn_succ] + @limitRecOn C (succ o) H₁ H₂ H₃ = H₂ o (@limitRecOn C o H₁ H₂ H₃) := + SuccOrder.limitRecOn_succ .. @[simp] theorem limitRecOn_limit {C} (o H₁ H₂ H₃ h) : - @limitRecOn C o H₁ H₂ H₃ = H₃ o h fun x _h => @limitRecOn C x H₁ H₂ H₃ := by - simp_rw [limitRecOn, SuccOrder.prelimitRecOn_of_isSuccPrelimit _ _ h.isSuccPrelimit, dif_neg h.1] + @limitRecOn C o H₁ H₂ H₃ = H₃ o h fun x _h => @limitRecOn C x H₁ H₂ H₃ := + SuccOrder.limitRecOn_of_isSuccLimit .. /-- Bounded recursion on ordinals. Similar to `limitRecOn`, with the assumption `o < l` added to all cases. The final term's domain is the ordinals below `l`. -/ @@ -347,7 +346,7 @@ alias out_no_max_of_succ_lt := toType_noMax_of_succ_lt theorem bounded_singleton {r : α → α → Prop} [IsWellOrder α r] (hr : (type r).IsLimit) (x) : Bounded r {x} := by - refine ⟨enum r ⟨succ (typein r x), hr.2 _ (typein_lt_type r x)⟩, ?_⟩ + refine ⟨enum r ⟨succ (typein r x), hr.succ_lt (typein_lt_type r x)⟩, ?_⟩ intro b hb rw [mem_singleton_iff.1 hb] nth_rw 1 [← enum_typein r x] @@ -398,7 +397,7 @@ theorem IsNormal.strictMono {f} (H : IsNormal f) : StrictMono f := fun a b => limitRecOn b (Not.elim (not_lt_of_le <| Ordinal.zero_le _)) (fun _b IH h => (lt_or_eq_of_le (le_of_lt_succ h)).elim (fun h => (IH h).trans (H.1 _)) fun e => e ▸ H.1 _) - fun _b l _IH h => lt_of_lt_of_le (H.1 a) ((H.2 _ l _).1 le_rfl _ (l.2 _ h)) + fun _b l _IH h => lt_of_lt_of_le (H.1 a) ((H.2 _ l _).1 le_rfl _ (l.succ_lt h)) theorem IsNormal.monotone {f} (H : IsNormal f) : Monotone f := H.strictMono.monotone @@ -460,10 +459,14 @@ theorem IsNormal.trans {f g} (H₁ : IsNormal f) (H₂ : IsNormal g) : IsNormal ⟨fun _x => H₁.lt_iff.2 (H₂.1 _), fun o l _a => H₁.le_set' (· < o) ⟨0, l.pos⟩ g _ fun _c => H₂.2 _ l _⟩ -theorem IsNormal.isLimit {f} (H : IsNormal f) {o} (l : IsLimit o) : IsLimit (f o) := - ⟨ne_of_gt <| (Ordinal.zero_le _).trans_lt <| H.lt_iff.2 l.pos, fun _ h => - let ⟨_b, h₁, h₂⟩ := (H.limit_lt l).1 h - (succ_le_of_lt h₂).trans_lt (H.lt_iff.2 h₁)⟩ +theorem IsNormal.isLimit {f} (H : IsNormal f) {o} (ho : IsLimit o) : IsLimit (f o) := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] + use (H.lt_iff.2 ho.pos).ne_bot + intro a ha + obtain ⟨b, hb, hab⟩ := (H.limit_lt ho).1 ha + rw [← succ_le_iff] at hab + apply hab.trans_lt + rwa [H.lt_iff] theorem add_le_of_limit {a b c : Ordinal} (h : IsLimit b) : a + b ≤ c ↔ ∀ b' < b, a + b' ≤ c := ⟨fun h _ l => (add_le_add_left l.le _).trans h, fun H => @@ -482,7 +485,7 @@ theorem add_le_of_limit {a b c : Ordinal} (h : IsLimit b) : a + b ≤ c ↔ ∀ · exact irrefl _ (this _) intro x rw [← typein_lt_typein (Sum.Lex r s), typein_enum] - have := H _ (h.2 _ (typein_lt_type s x)) + have := H _ (h.succ_lt (typein_lt_type s x)) rw [add_succ, succ_le_iff] at this refine (RelEmbedding.ofMonotone (fun a => ?_) fun a b => ?_).ordinal_type_le.trans_lt this @@ -566,6 +569,9 @@ protected theorem sub_eq_zero_iff_le {a b : Ordinal} : a - b = 0 ↔ a ≤ b := ⟨fun h => by simpa only [h, add_zero] using le_add_sub a b, fun h => by rwa [← Ordinal.le_zero, sub_le, add_zero]⟩ +protected theorem sub_ne_zero_iff_lt {a b : Ordinal} : a - b ≠ 0 ↔ b < a := by + simpa using Ordinal.sub_eq_zero_iff_le.not + theorem sub_sub (a b c : Ordinal) : a - b - c = a - (b + c) := eq_of_forall_ge_iff fun d => by rw [sub_le, sub_le, sub_le, add_assoc] @@ -587,9 +593,12 @@ theorem lt_add_iff {a b c : Ordinal} (hc : c ≠ 0) : a < b + c ↔ ∃ d < c, a rintro ⟨d, hd, ha⟩ exact ha.trans_lt (add_lt_add_left hd b) -theorem isLimit_sub {a b} (l : IsLimit a) (h : b < a) : IsLimit (a - b) := - ⟨ne_of_gt <| lt_sub.2 <| by rwa [add_zero], fun c h => by - rw [lt_sub, add_succ]; exact l.2 _ (lt_sub.1 h)⟩ +theorem isLimit_sub {a b} (ha : IsLimit a) (h : b < a) : IsLimit (a - b) := by + rw [isLimit_iff, Ordinal.sub_ne_zero_iff_lt, isSuccPrelimit_iff_succ_lt] + refine ⟨h, fun c hc ↦ ?_⟩ + rw [lt_sub] at hc ⊢ + rw [add_succ] + exact ha.succ_lt hc @[deprecated isLimit_sub (since := "2024-10-11")] alias sub_isLimit := isLimit_sub @@ -732,7 +741,7 @@ private theorem mul_le_of_limit_aux {α β r s} [IsWellOrder α r] [IsWellOrder exact irrefl _ (this _ _) intro a b rw [← typein_lt_typein (Prod.Lex s r), typein_enum] - have := H _ (h.2 _ (typein_lt_type s b)) + have := H _ (h.succ_lt (typein_lt_type s b)) rw [mul_succ] at this have := ((add_lt_add_iff_left _).2 (typein_lt_type _ a)).trans_le this refine (RelEmbedding.ofMonotone (fun a => ?_) fun a b => ?_).ordinal_type_le.trans_lt this @@ -1464,8 +1473,8 @@ theorem sInf_compl_lt_ord_succ {ι : Type u} (f : ι → Ordinal.{u}) : -- TODO: remove `bsup` in favor of `iSup` in a future refactor. section bsup -set_option linter.deprecated false +set_option linter.deprecated false in private theorem sup_le_sup {ι ι' : Type u} (r : ι → ι → Prop) (r' : ι' → ι' → Prop) [IsWellOrder ι r] [IsWellOrder ι' r'] {o} (ho : type r = o) (ho' : type r' = o) (f : ∀ a < o, Ordinal.{max u v}) : @@ -1480,23 +1489,27 @@ private theorem sup_le_sup {ι ι' : Type u} (r : ι → ι → Prop) (r' : ι' simp_rw [familyOfBFamily', ← hj] apply le_sup +set_option linter.deprecated false in theorem sup_eq_sup {ι ι' : Type u} (r : ι → ι → Prop) (r' : ι' → ι' → Prop) [IsWellOrder ι r] [IsWellOrder ι' r'] {o : Ordinal.{u}} (ho : type r = o) (ho' : type r' = o) (f : ∀ a < o, Ordinal.{max u v}) : sup.{_, v} (familyOfBFamily' r ho f) = sup.{_, v} (familyOfBFamily' r' ho' f) := sup_eq_of_range_eq.{u, u, v} (by simp) +set_option linter.deprecated false in /-- The supremum of a family of ordinals indexed by the set of ordinals less than some `o : Ordinal.{u}`. This is a special case of `sup` over the family provided by `familyOfBFamily`. -/ def bsup (o : Ordinal.{u}) (f : ∀ a < o, Ordinal.{max u v}) : Ordinal.{max u v} := sup.{_, v} (familyOfBFamily o f) +set_option linter.deprecated false in @[simp] theorem sup_eq_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : sup.{_, v} (familyOfBFamily o f) = bsup.{_, v} o f := rfl +set_option linter.deprecated false in @[simp] theorem sup_eq_bsup' {o : Ordinal.{u}} {ι} (r : ι → ι → Prop) [IsWellOrder ι r] (ho : type r = o) (f : ∀ a < o, Ordinal.{max u v}) : sup.{_, v} (familyOfBFamily' r ho f) = bsup.{_, v} o f := @@ -1507,6 +1520,7 @@ theorem sSup_eq_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : congr rw [range_familyOfBFamily] +set_option linter.deprecated false in @[simp] theorem bsup_eq_sup' {ι : Type u} (r : ι → ι → Prop) [IsWellOrder ι r] (f : ι → Ordinal.{max u v}) : bsup.{_, v} _ (bfamilyOfFamily' r f) = sup.{_, v} f := by @@ -1518,6 +1532,7 @@ theorem bsup_eq_bsup {ι : Type u} (r r' : ι → ι → Prop) [IsWellOrder ι r bsup.{_, v} _ (bfamilyOfFamily' r f) = bsup.{_, v} _ (bfamilyOfFamily' r' f) := by rw [bsup_eq_sup', bsup_eq_sup'] +set_option linter.deprecated false in @[simp] theorem bsup_eq_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : bsup.{_, v} _ (bfamilyOfFamily f) = sup.{_, v} f := @@ -1530,6 +1545,7 @@ theorem bsup_congr {o₁ o₂ : Ordinal.{u}} (f : ∀ a < o₁, Ordinal.{max u v -- Porting note: `rfl` is required. rfl +set_option linter.deprecated false in theorem bsup_le_iff {o f a} : bsup.{u, v} o f ≤ a ↔ ∀ i h, f i h ≤ a := sup_le_iff.trans ⟨fun h i hi => by @@ -1547,6 +1563,7 @@ theorem lt_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {a} : a < bsup.{_, v} o f ↔ ∃ i hi, a < f i hi := by simpa only [not_forall, not_le] using not_congr (@bsup_le_iff.{_, v} _ f a) +set_option linter.deprecated false in theorem IsNormal.bsup {f : Ordinal.{max u v} → Ordinal.{max u w}} (H : IsNormal f) {o : Ordinal.{u}} : ∀ (g : ∀ a < o, Ordinal), o ≠ 0 → f (bsup.{_, v} o g) = bsup.{_, w} o fun a h => f (g a h) := @@ -1558,6 +1575,7 @@ theorem lt_bsup_of_ne_bsup {o : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} (∀ i h, f i h ≠ bsup.{_, v} o f) ↔ ∀ i h, f i h < bsup.{_, v} o f := ⟨fun hf _ _ => lt_of_le_of_ne (le_bsup _ _ _) (hf _ _), fun hf _ _ => ne_of_lt (hf _ _)⟩ +set_option linter.deprecated false in theorem bsup_not_succ_of_ne_bsup {o : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} (hf : ∀ {i : Ordinal} (h : i < o), f i h ≠ bsup.{_, v} o f) (a) : a < bsup.{_, v} o f → succ a < bsup.{_, v} o f := by @@ -1589,6 +1607,7 @@ theorem bsup_const {o : Ordinal.{u}} (ho : o ≠ 0) (a : Ordinal.{max u v}) : (bsup.{_, v} o fun _ _ => a) = a := le_antisymm (bsup_le fun _ _ => le_rfl) (le_bsup _ 0 (Ordinal.pos_iff_ne_zero.2 ho)) +set_option linter.deprecated false in @[simp] theorem bsup_one (f : ∀ a < (1 : Ordinal), Ordinal) : bsup 1 f = f 0 zero_lt_one := by simp_rw [← sup_eq_bsup, sup_unique, familyOfBFamily, familyOfBFamily', typein_one_toType] @@ -1604,6 +1623,7 @@ theorem bsup_eq_of_brange_eq {o o'} {f : ∀ a < o, Ordinal} {g : ∀ a < o', Or (h : brange o f = brange o' g) : bsup.{u, max v w} o f = bsup.{v, max u w} o' g := (bsup_le_of_brange_subset.{u, v, w} h.le).antisymm (bsup_le_of_brange_subset.{v, u, w} h.ge) +set_option linter.deprecated false in theorem iSup_eq_bsup {o} {f : ∀ a < o, Ordinal} : ⨆ a : Iio o, f a.1 a.2 = bsup o f := by simp_rw [Iio, bsup, sup, iSup, range_familyOfBFamily, brange, range, Subtype.exists, mem_setOf] @@ -1612,17 +1632,19 @@ end bsup -- TODO: bring the lsub API in line with the sSup / iSup API, or deprecate it altogether. section lsub -set_option linter.deprecated false +set_option linter.deprecated false in /-- The least strict upper bound of a family of ordinals. -/ def lsub {ι} (f : ι → Ordinal) : Ordinal := sup (succ ∘ f) +set_option linter.deprecated false in @[simp] theorem sup_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} (succ ∘ f) = lsub.{_, v} f := rfl +set_option linter.deprecated false in theorem lsub_le_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : lsub.{_, v} f ≤ a ↔ ∀ i, f i < a := by convert sup_le_iff.{_, v} (f := succ ∘ f) (a := a) using 2 @@ -1632,6 +1654,7 @@ theorem lsub_le_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : theorem lsub_le {ι} {f : ι → Ordinal} {a} : (∀ i, f i < a) → lsub f ≤ a := lsub_le_iff.2 +set_option linter.deprecated false in theorem lt_lsub {ι} (f : ι → Ordinal) (i) : f i < lsub f := succ_le_iff.1 (le_sup _ i) @@ -1639,19 +1662,23 @@ theorem lt_lsub_iff {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : a < lsub.{_, v} f ↔ ∃ i, a ≤ f i := by simpa only [not_forall, not_lt, not_le] using not_congr (@lsub_le_iff.{_, v} _ f a) +set_option linter.deprecated false in theorem sup_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f ≤ lsub.{_, v} f := sup_le fun i => (lt_lsub f i).le +set_option linter.deprecated false in theorem lsub_le_sup_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : lsub.{_, v} f ≤ succ (sup.{_, v} f) := lsub_le fun i => lt_succ_iff.2 (le_sup f i) +set_option linter.deprecated false in theorem sup_eq_lsub_or_sup_succ_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f = lsub.{_, v} f ∨ succ (sup.{_, v} f) = lsub.{_, v} f := by cases' eq_or_lt_of_le (sup_le_lsub.{_, v} f) with h h · exact Or.inl h · exact Or.inr ((succ_le_of_lt h).antisymm (lsub_le_sup_succ f)) +set_option linter.deprecated false in theorem sup_succ_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : succ (sup.{_, v} f) ≤ lsub.{_, v} f ↔ ∃ i, f i = sup.{_, v} f := by refine ⟨fun h => ?_, ?_⟩ @@ -1661,10 +1688,12 @@ theorem sup_succ_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : rw [succ_le_iff, ← hf] exact lt_lsub _ _ +set_option linter.deprecated false in theorem sup_succ_eq_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : succ (sup.{_, v} f) = lsub.{_, v} f ↔ ∃ i, f i = sup.{_, v} f := (lsub_le_sup_succ f).le_iff_eq.symm.trans (sup_succ_le_lsub f) +set_option linter.deprecated false in theorem sup_eq_lsub_iff_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f = lsub.{_, v} f ↔ ∀ a < lsub.{_, v} f, succ a < lsub.{_, v} f := by refine ⟨fun h => ?_, fun hf => le_antisymm (sup_le_lsub f) (lsub_le fun i => ?_)⟩ @@ -1680,6 +1709,7 @@ theorem sup_eq_lsub_iff_succ {ι : Type u} (f : ι → Ordinal.{max u v}) : rw [heq] at this exact this.false +set_option linter.deprecated false in theorem sup_eq_lsub_iff_lt_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : sup.{_, v} f = lsub.{_, v} f ↔ ∀ i, f i < sup.{_, v} f := ⟨fun h i => by @@ -1702,14 +1732,17 @@ theorem lsub_eq_zero_iff {ι : Type u} (f : ι → Ordinal.{max u v}) : rw [h] at this exact this.false +set_option linter.deprecated false in @[simp] theorem lsub_const {ι} [Nonempty ι] (o : Ordinal) : (lsub fun _ : ι => o) = succ o := sup_const (succ o) +set_option linter.deprecated false in @[simp] theorem lsub_unique {ι} [Unique ι] (f : ι → Ordinal) : lsub f = succ (f default) := sup_unique _ +set_option linter.deprecated false in theorem lsub_le_of_range_subset {ι ι'} {f : ι → Ordinal} {g : ι' → Ordinal} (h : Set.range f ⊆ Set.range g) : lsub.{u, max v w} f ≤ lsub.{v, max u w} g := sup_le_of_range_subset.{u, v, w} (by convert Set.image_subset succ h <;> apply Set.range_comp) @@ -1718,6 +1751,7 @@ theorem lsub_eq_of_range_eq {ι ι'} {f : ι → Ordinal} {g : ι' → Ordinal} (h : Set.range f = Set.range g) : lsub.{u, max v w} f = lsub.{v, max u w} g := (lsub_le_of_range_subset.{u, v, w} h.le).antisymm (lsub_le_of_range_subset.{v, u, w} h.ge) +set_option linter.deprecated false in @[simp] theorem lsub_sum {α : Type u} {β : Type v} (f : α ⊕ β → Ordinal) : lsub.{max u v, w} f = @@ -1731,6 +1765,7 @@ theorem lsub_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : theorem nonempty_compl_range {ι : Type u} (f : ι → Ordinal.{max u v}) : (Set.range f)ᶜ.Nonempty := ⟨_, lsub_not_mem_range.{_, v} f⟩ +set_option linter.deprecated false in @[simp] theorem lsub_typein (o : Ordinal) : lsub.{u, u} (typein (α := o.toType) (· < ·)) = o := (lsub_le.{u, u} typein_lt_self).antisymm @@ -1740,11 +1775,13 @@ theorem lsub_typein (o : Ordinal) : lsub.{u, u} (typein (α := o.toType) (· < conv_rhs at h => rw [← type_lt o] simpa [typein_enum] using lt_lsub.{u, u} (typein (· < ·)) (enum (· < ·) ⟨_, h⟩)) +set_option linter.deprecated false in theorem sup_typein_limit {o : Ordinal} (ho : ∀ a, a < o → succ a < o) : sup.{u, u} (typein ((· < ·) : o.toType → o.toType → Prop)) = o := by -- Porting note: `rwa` → `rw` & `assumption` rw [(sup_eq_lsub_iff_succ.{u, u} (typein (· < ·))).2] <;> rw [lsub_typein o]; assumption +set_option linter.deprecated false in @[simp] theorem sup_typein_succ {o : Ordinal} : sup.{u, u} (typein ((· < ·) : (succ o).toType → (succ o).toType → Prop)) = o := by @@ -1764,7 +1801,6 @@ end lsub -- both of them at once. section blsub -set_option linter.deprecated false /-- The least strict upper bound of a family of ordinals indexed by the set of ordinals less than some `o : Ordinal.{u}`. @@ -1870,7 +1906,7 @@ theorem bsup_eq_blsub_iff_lt_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max apply lt_blsub, fun h => le_antisymm (bsup_le_blsub f) (blsub_le h)⟩ theorem bsup_eq_blsub_of_lt_succ_limit {o : Ordinal.{u}} (ho : IsLimit o) - {f : ∀ a < o, Ordinal.{max u v}} (hf : ∀ a ha, f a ha < f (succ a) (ho.2 a ha)) : + {f : ∀ a < o, Ordinal.{max u v}} (hf : ∀ a ha, f a ha < f (succ a) (ho.succ_lt ha)) : bsup.{_, v} o f = blsub.{_, v} o f := by rw [bsup_eq_blsub_iff_lt_bsup] exact fun i hi => (hf i hi).trans_le (le_bsup f _ _) @@ -1948,7 +1984,7 @@ theorem blsub_comp {o o' : Ordinal.{max u v}} {f : ∀ a < o, Ordinal.{max u v w theorem IsNormal.bsup_eq {f : Ordinal.{u} → Ordinal.{max u v}} (H : IsNormal f) {o : Ordinal.{u}} (h : IsLimit o) : (Ordinal.bsup.{_, v} o fun x _ => f x) = f o := by - rw [← IsNormal.bsup.{u, u, v} H (fun x _ => x) h.1, bsup_id_limit h.2] + rw [← IsNormal.bsup.{u, u, v} H (fun x _ => x) h.ne_bot, bsup_id_limit fun _ ↦ h.succ_lt] theorem IsNormal.blsub_eq {f : Ordinal.{u} → Ordinal.{max u v}} (H : IsNormal f) {o : Ordinal.{u}} (h : IsLimit o) : (blsub.{_, v} o fun x _ => f x) = f o := by @@ -1992,6 +2028,7 @@ def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : O Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} @@ -2003,7 +2040,6 @@ theorem lt_blsub₂ {o₁ o₂ : Ordinal} end blsub section mex -set_option linter.deprecated false /-! ### Minimum excluded ordinals -/ @@ -2013,33 +2049,40 @@ set_option linter.deprecated false def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by @@ -2048,6 +2091,7 @@ theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → rw [← hj] at hi exact ne_mex g j hi +set_option linter.deprecated false in @[deprecated sInf_compl_lt_ord_succ (since := "2024-09-20")] theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : mex.{_, u} f < (succ #ι).ord := by @@ -2064,6 +2108,7 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : convert Cardinal.mk_le_of_injective hg rw [Cardinal.mk_ord_toType (succ #ι)] +set_option linter.deprecated false in /-- The minimum excluded ordinal of a family of ordinals indexed by the set of ordinals less than some `o : Ordinal.{u}`. This is a special case of `mex` over the family provided by `familyOfBFamily`. @@ -2073,17 +2118,20 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by @@ -2092,28 +2140,33 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := mex_monotone (by rwa [range_familyOfBFamily, range_familyOfBFamily]) +set_option linter.deprecated false in @[deprecated sInf_compl_lt_ord_succ (since := "2024-09-20")] theorem bmex_lt_ord_succ_card {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{u}) : bmex.{_, u} o f < (succ o.card).ord := by @@ -2316,10 +2369,11 @@ theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omeg @[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 -theorem isLimit_omega0 : IsLimit ω := - ⟨omega0_ne_zero, fun o h => by - let ⟨n, e⟩ := lt_omega0.1 h - rw [e]; exact nat_lt_omega0 (n + 1)⟩ +theorem isLimit_omega0 : IsLimit ω := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] + refine ⟨omega0_ne_zero, fun o h => ?_⟩ + obtain ⟨n, rfl⟩ := lt_omega0.1 h + exact nat_lt_omega0 (n + 1) @[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias omega0_isLimit := isLimit_omega0 @@ -2349,8 +2403,8 @@ theorem sup_natCast : sup Nat.cast = ω := alias sup_nat_cast := sup_natCast theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o - | 0 => lt_of_le_of_ne (Ordinal.zero_le o) h.1.symm - | n + 1 => h.2 _ (nat_lt_limit h n) + | 0 => h.pos + | n + 1 => h.succ_lt (nat_lt_limit h n) theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := omega0_le.2 fun n => le_of_lt <| nat_lt_limit h n @@ -2359,7 +2413,7 @@ theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := alias omega_le_of_isLimit := omega0_le_of_isLimit theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by - refine ⟨fun l => ⟨l.1, ⟨a / ω, le_antisymm ?_ (mul_div_le _ _)⟩⟩, fun h => ?_⟩ + refine ⟨fun l => ⟨l.ne_zero, ⟨a / ω, le_antisymm ?_ (mul_div_le _ _)⟩⟩, fun h => ?_⟩ · refine (limit_le l).2 fun x hx => le_of_lt ?_ rw [← div_lt omega0_ne_zero, ← succ_le_iff, le_div omega0_ne_zero, mul_succ, add_le_of_limit isLimit_omega0] @@ -2384,7 +2438,7 @@ theorem add_mul_limit_aux {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) rw [IH _ h] apply (add_le_add_left _ _).trans · rw [← mul_succ] - exact mul_le_mul_left' (succ_le_of_lt <| l.2 _ h) _ + exact mul_le_mul_left' (succ_le_of_lt <| l.succ_lt h) _ · rw [← ba] exact le_add_right _ _) (mul_le_mul_right' (le_add_right _ _) _) @@ -2451,6 +2505,7 @@ namespace Cardinal open Ordinal theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] refine ⟨fun h => aleph0_ne_zero ?_, fun a => lt_imp_lt_of_le_imp_le fun h => ?_⟩ · rw [← Ordinal.le_zero, ord_le] at h simpa only [card_zero, nonpos_iff_eq_zero] using co.trans h @@ -2465,7 +2520,7 @@ theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := - toType_noMax_of_succ_lt (isLimit_ord h).2 + toType_noMax_of_succ_lt fun _ ↦ (isLimit_ord h).succ_lt end Cardinal diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 84f2f8db36b0d..00a1b45e0c167 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -226,8 +226,6 @@ end section -set_option linter.deprecated false - variable {o : Ordinal.{u}} {f : ∀ b < o, Ordinal.{max u v} → Ordinal.{max u v}} /-- The next common fixed point, at least `a`, for a family of normal functions indexed by ordinals. @@ -238,47 +236,56 @@ def nfpBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal.{ Ordinal.{max u v} → Ordinal.{max u v} := nfpFamily (familyOfBFamily o f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_nfpFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : nfpBFamily.{u, v} o f = nfpFamily (familyOfBFamily o f) := rfl +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem foldr_le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a l) : List.foldr (familyOfBFamily o f) a l ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a) : a ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup (fun _ ↦ List.foldr _ a _) [] +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem lt_nfpBFamily {a b} : a < nfpBFamily.{u, v} o f b ↔ ∃ l, a < List.foldr (familyOfBFamily o f) b l := Ordinal.lt_iSup_iff +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_iff {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : nfpBFamily.{u, v} o f a ≤ b ↔ ∀ l, List.foldr (familyOfBFamily o f) a l ≤ b := Ordinal.iSup_le_iff +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : (∀ l, List.foldr (familyOfBFamily o f) a l ≤ b) → nfpBFamily.{u, v} o f a ≤ b := Ordinal.iSup_le +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBFamily.{u, v} o f) := nfpFamily_monotone fun _ => hf _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b < nfpBFamily.{u, v} o f a) ↔ b < nfpBFamily.{u, v} o f a := @@ -287,6 +294,7 @@ theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) refine (apply_lt_nfpFamily_iff ?_).1 fun _ => h _ _ exact fun _ => H _ _, apply_lt_nfpBFamily H⟩ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∃ i hi, nfpBFamily.{u, v} o f a ≤ f i hi b) ↔ nfpBFamily.{u, v} o f a ≤ b := by @@ -294,11 +302,13 @@ theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a push_neg exact apply_lt_nfpBFamily_iff.{u, v} ho H +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_fp (H : ∀ i hi, Monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfpBFamily.{u, v} o f a ≤ b := nfpFamily_le_fp (fun _ => H _ _) ab fun _ => h _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : f i hi (nfpBFamily.{u, v} o f a) = nfpBFamily.{u, v} o f a := by @@ -307,6 +317,7 @@ theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : rw [familyOfBFamily_enum] exact H +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfpBFamily.{u, v} o f a) ↔ b ≤ nfpBFamily.{u, v} o f a := by @@ -316,10 +327,12 @@ theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a · rw [← nfpBFamily_fp (H i hi)] exact (H i hi).monotone h +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} o f a = a := nfpFamily_eq_self fun _ => h _ _ +set_option linter.deprecated false in /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ @[deprecated "No deprecation message was provided." (since := "2024-10-14")] @@ -330,6 +343,7 @@ theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : rw [Set.mem_iInter₂] exact fun i hi ↦ nfpBFamily_fp (H i hi) _ +set_option linter.deprecated false in /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ @[deprecated not_bddAbove_fp_bfamily (since := "2024-09-20")] @@ -347,11 +361,13 @@ def derivBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal Ordinal.{max u v} → Ordinal.{max u v} := derivFamily (familyOfBFamily o f) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_derivFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : derivBFamily.{u, v} o f = derivFamily (familyOfBFamily o f) := rfl +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : IsNormal (derivBFamily o f) := @@ -360,6 +376,7 @@ theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) @[deprecated isNormal_derivBFamily (since := "2024-10-11")] alias derivBFamily_isNormal := isNormal_derivBFamily +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : f i hi (derivBFamily.{u, v} o f a) = derivBFamily.{u, v} o f a := by @@ -368,6 +385,7 @@ theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : rw [familyOfBFamily_enum] exact H +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a ≤ a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by @@ -378,6 +396,7 @@ theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : apply h · exact fun _ => H _ _ +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a = a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by @@ -386,6 +405,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : rw [← (H i hi).le_iff_eq] exact h i hi +set_option linter.deprecated false in /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ @[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 02806888024eb..52e8d862840e0 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -802,8 +802,8 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr · apply (mul_le_mul_left' (le_succ b) _).trans rw [← add_one_eq_succ, add_mul_succ _ (one_add_of_omega0_le h), add_one_eq_succ, succ_le_iff, Ordinal.mul_lt_mul_iff_left (Ordinal.pos_iff_ne_zero.2 e0)] - exact isLimit_omega0.2 _ l - · apply (principal_mul_omega0 (isLimit_omega0.2 _ h) l).le.trans + exact isLimit_omega0.succ_lt l + · apply (principal_mul_omega0 (isLimit_omega0.succ_lt h) l).le.trans simpa using mul_le_mul_right' (one_le_iff_ne_zero.2 e0) ω section diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 194abeabc7c50..5ae52a926bb59 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -162,8 +162,9 @@ theorem principal_add_of_le_one (ho : o ≤ 1) : Principal (· + ·) o := by · exact principal_zero · exact principal_add_one -theorem isLimit_of_principal_add (ho₁ : 1 < o) (ho : Principal (· + ·) o) : o.IsLimit := - ⟨ho₁.ne_bot, fun _ ha ↦ ho ha ho₁⟩ +theorem isLimit_of_principal_add (ho₁ : 1 < o) (ho : Principal (· + ·) o) : o.IsLimit := by + rw [isLimit_iff, isSuccPrelimit_iff_succ_lt] + exact ⟨ho₁.ne_bot, fun _ ha ↦ ho ha ho₁⟩ @[deprecated (since := "2024-10-16")] alias principal_add_isLimit := isLimit_of_principal_add @@ -388,7 +389,7 @@ theorem mul_lt_omega0_opow (c0 : 0 < c) (ha : a < ω ^ c) (hb : b < ω) : a * b · rcases ((isNormal_opow one_lt_omega0).limit_lt l).1 ha with ⟨x, hx, ax⟩ refine (mul_le_mul' (le_of_lt ax) (le_of_lt hb)).trans_lt ?_ rw [← opow_succ, opow_lt_opow_iff_right one_lt_omega0] - exact l.2 _ hx + exact l.succ_lt hx @[deprecated (since := "2024-09-30")] alias mul_lt_omega_opow := mul_lt_omega0_opow @@ -457,7 +458,7 @@ theorem mul_eq_opow_log_succ (ha : a ≠ 0) (hb : Principal (· * ·) b) (hb₂ have hbo₀ : b ^ log b a ≠ 0 := Ordinal.pos_iff_ne_zero.1 (opow_pos _ (zero_lt_one.trans hb₁)) apply (mul_le_mul_right' (le_of_lt (lt_mul_succ_div a hbo₀)) c).trans rw [mul_assoc, opow_succ] - refine mul_le_mul_left' (hb (hbl.2 _ ?_) hcb).le _ + refine mul_le_mul_left' (hb (hbl.succ_lt ?_) hcb).le _ rw [div_lt hbo₀, ← opow_succ] exact lt_opow_succ_log_self hb₁ _ · rw [opow_succ] diff --git a/Mathlib/SetTheory/Ordinal/Rank.lean b/Mathlib/SetTheory/Ordinal/Rank.lean index 1eb486c71b504..fb8b097d185a7 100644 --- a/Mathlib/SetTheory/Ordinal/Rank.lean +++ b/Mathlib/SetTheory/Ordinal/Rank.lean @@ -97,8 +97,6 @@ theorem IsWellFounded.rank_eq_typein (r) [IsWellOrder α r] : rank r = Ordinal.t namespace WellFounded -set_option linter.deprecated false - variable {r : α → α → Prop} (hwf : WellFounded r) /-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the @@ -108,18 +106,22 @@ smallest ordinal greater than the ranks of all elements below it (i.e. elements noncomputable def rank (a : α) : Ordinal.{u} := (hwf.apply a).rank +set_option linter.deprecated false in @[deprecated IsWellFounded.rank_eq (since := "2024-09-07")] theorem rank_eq : hwf.rank a = ⨆ b : { b // r b a }, Order.succ (hwf.rank b) := (hwf.apply a).rank_eq +set_option linter.deprecated false in @[deprecated IsWellFounded.rank_lt_of_rel (since := "2024-09-07")] theorem rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := Acc.rank_lt_of_rel _ h +set_option linter.deprecated false in @[deprecated WellFoundedLT.rank_strictMono (since := "2024-09-07")] theorem rank_strictMono [Preorder α] [WellFoundedLT α] : StrictMono (rank <| @wellFounded_lt α _ _) := fun _ _ => rank_lt_of_rel _ +set_option linter.deprecated false in @[deprecated WellFoundedGT.rank_strictAnti (since := "2024-09-07")] theorem rank_strictAnti [Preorder α] [WellFoundedGT α] : StrictAnti (rank <| @wellFounded_gt α _ _) := fun _ _ => rank_lt_of_rel wellFounded_gt diff --git a/Mathlib/SetTheory/Ordinal/Topology.lean b/Mathlib/SetTheory/Ordinal/Topology.lean index a1dd880e86d62..515769035484d 100644 --- a/Mathlib/SetTheory/Ordinal/Topology.lean +++ b/Mathlib/SetTheory/Ordinal/Topology.lean @@ -37,11 +37,11 @@ instance : TopologicalSpace Ordinal.{u} := Preorder.topology Ordinal.{u} instance : OrderTopology Ordinal.{u} := ⟨rfl⟩ theorem isOpen_singleton_iff : IsOpen ({a} : Set Ordinal) ↔ ¬IsLimit a := by - refine ⟨fun h ⟨h₀, hsucc⟩ => ?_, fun ha => ?_⟩ + refine ⟨fun h ha => ?_, fun ha => ?_⟩ · obtain ⟨b, c, hbc, hbc'⟩ := - (mem_nhds_iff_exists_Ioo_subset' ⟨0, Ordinal.pos_iff_ne_zero.2 h₀⟩ ⟨_, lt_succ a⟩).1 + (mem_nhds_iff_exists_Ioo_subset' ⟨0, ha.pos⟩ ⟨_, lt_succ a⟩).1 (h.mem_nhds rfl) - have hba := hsucc b hbc.1 + have hba := ha.succ_lt hbc.1 exact hba.ne (hbc' ⟨lt_succ b, hba.trans hbc.2⟩) · rcases zero_or_succ_or_limit a with (rfl | ⟨b, rfl⟩ | ha') · rw [← bot_eq_zero, ← Set.Iic_bot, ← Iio_succ] @@ -73,7 +73,7 @@ theorem nhds_eq_pure : 𝓝 a = pure a ↔ ¬IsLimit a := theorem isOpen_iff : IsOpen s ↔ ∀ o ∈ s, IsLimit o → ∃ a < o, Set.Ioo a o ⊆ s := by refine isOpen_iff_mem_nhds.trans <| forall₂_congr fun o ho => ?_ by_cases ho' : IsLimit o - · simp only [(nhdsBasis_Ioc ho'.1).mem_iff, ho', true_implies] + · simp only [(nhdsBasis_Ioc ho'.ne_zero).mem_iff, ho', true_implies] refine exists_congr fun a => and_congr_right fun ha => ?_ simp only [← Set.Ioo_insert_right ha, Set.insert_subset_iff, ho, true_and] · simp [nhds_eq_pure.2 ho', ho, ho'] @@ -216,8 +216,8 @@ theorem isNormal_iff_strictMono_and_continuous (f : Ordinal.{u} → Ordinal.{u}) suffices o ∈ f ⁻¹' Set.Iic a from Set.mem_preimage.1 this rw [mem_iff_iSup_of_isClosed (IsClosed.preimage h' (@isClosed_Iic _ _ _ _ a))] exact - ⟨_, toType_nonempty_iff_ne_zero.2 ho.1, typein (· < ·), fun i => h _ (typein_lt_self i), - sup_typein_limit ho.2⟩ + ⟨_, toType_nonempty_iff_ne_zero.2 ho.ne_zero, typein (· < ·), fun i => h _ (typein_lt_self i), + sup_typein_limit fun _ ↦ ho.succ_lt⟩ theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : IsNormal (enumOrd s) ↔ IsClosed s := by @@ -237,14 +237,14 @@ theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : · rw [isClosed_iff_bsup] at h suffices enumOrd s a ≤ bsup.{u, u} a fun b (_ : b < a) => enumOrd s b from this.trans (bsup_le H) - obtain ⟨b, hb⟩ := enumOrd_surjective hs (h ha.1 (fun b _ => enumOrd s b) + obtain ⟨b, hb⟩ := enumOrd_surjective hs (h ha.ne_zero (fun b _ => enumOrd s b) fun b _ => enumOrd_mem hs b) rw [← hb] apply Hs.monotone by_contra! hba apply (Hs (lt_succ b)).not_le rw [hb] - exact le_bsup.{u, u} _ _ (ha.2 _ hba) + exact le_bsup.{u, u} _ _ (ha.succ_lt hba) open Set Filter Set.Notation diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index a2b3354096bf0..00484448b82b2 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -443,8 +443,6 @@ def embed : PSet.{max (u + 1) v} := theorem lift_mem_embed : ∀ x : PSet.{u}, PSet.Lift.{u, max (u + 1) v} x ∈ embed.{u, v} := fun x => ⟨⟨x⟩, Equiv.rfl⟩ -set_option linter.deprecated false - /-- Function equivalence is defined so that `f ~ g` iff `∀ x y, x ~ y → f x ~ g y`. This extends to equivalence of `n`-ary functions. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] @@ -452,37 +450,44 @@ def Arity.Equiv : ∀ {n}, OfArity PSet.{u} PSet.{u} n → OfArity PSet.{u} PSet | 0, a, b => PSet.Equiv a b | _ + 1, a, b => ∀ x y : PSet, PSet.Equiv x y → Arity.Equiv (a x) (b y) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Arity.equiv_const {a : PSet.{u}} : ∀ n, Arity.Equiv (OfArity.const PSet.{u} a n) (OfArity.const PSet.{u} a n) | 0 => Equiv.rfl | _ + 1 => fun _ _ _ => Arity.equiv_const _ +set_option linter.deprecated false in /-- `resp n` is the collection of n-ary functions on `PSet` that respect equivalence, i.e. when the inputs are equivalent the output is as well. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp (n) := { x : OfArity PSet.{u} PSet.{u} n // Arity.Equiv x x } +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.inhabited {n} : Inhabited (Resp n) := ⟨⟨OfArity.const _ default _, Arity.equiv_const _⟩⟩ +set_option linter.deprecated false in /-- The `n`-ary image of a `(n + 1)`-ary function respecting equivalence as a function respecting equivalence. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.f {n} (f : Resp (n + 1)) (x : PSet) : Resp n := ⟨f.1 x, f.2 _ _ <| Equiv.refl x⟩ +set_option linter.deprecated false in /-- Function equivalence for functions respecting equivalence. See `PSet.Arity.Equiv`. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.Equiv {n} (a b : Resp n) : Prop := Arity.Equiv a.1 b.1 +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02"), refl] protected theorem Resp.Equiv.refl {n} (a : Resp n) : Resp.Equiv a a := a.2 +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] protected theorem Resp.Equiv.euc : ∀ {n} {a b c : Resp n}, Resp.Equiv a b → Resp.Equiv c b → Resp.Equiv a c @@ -490,15 +495,18 @@ protected theorem Resp.Equiv.euc : | n + 1, a, b, c, hab, hcb => fun x y h => @Resp.Equiv.euc n (a.f x) (b.f y) (c.f y) (hab _ _ h) (hcb _ _ <| PSet.Equiv.refl y) +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02"), symm] protected theorem Resp.Equiv.symm {n} {a b : Resp n} : Resp.Equiv a b → Resp.Equiv b a := (Resp.Equiv.refl b).euc +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02"), trans] protected theorem Resp.Equiv.trans {n} {x y z : Resp n} (h1 : Resp.Equiv x y) (h2 : Resp.Equiv y z) : Resp.Equiv x z := h1.euc h2.symm +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.setoid {n} : Setoid (Resp n) := ⟨Resp.Equiv, Resp.Equiv.refl, Resp.Equiv.symm, Resp.Equiv.trans⟩ @@ -606,10 +614,9 @@ end ZFSet namespace PSet -set_option linter.deprecated false - namespace Resp +set_option linter.deprecated false in /-- Helper function for `PSet.eval`. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def evalAux : @@ -625,11 +632,13 @@ def evalAux : (@Quotient.ind _ _ fun q => F b q = F c q) fun z => evalAux.2 (Resp.f b z) (Resp.f c z) (h _ _ (PSet.Equiv.refl z))⟩ +set_option linter.deprecated false in /-- An equivalence-respecting function yields an n-ary ZFC set function. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def eval (n) : Resp n → OfArity ZFSet.{u} ZFSet.{u} n := evalAux.1 +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem eval_val {n f x} : (@eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) ⟦x⟧ = eval n (Resp.f f x) := @@ -637,6 +646,7 @@ theorem eval_val {n f x} : end Resp +set_option linter.deprecated false in /-- A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image. -/ @@ -647,17 +657,20 @@ class inductive Definable (n) : OfArity ZFSet.{u} ZFSet.{u} n → Type (u + 1) attribute [deprecated "No deprecation message was provided." (since := "2024-09-02"), instance] Definable.mk +set_option linter.deprecated false in /-- The evaluation of a function respecting equivalence is definable, by that same function. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.EqMk {n} (f) : ∀ {s : OfArity ZFSet.{u} ZFSet.{u} n} (_ : Resp.eval _ f = s), Definable n s | _, rfl => ⟨f⟩ +set_option linter.deprecated false in /-- Turns a definable function into a function that respects equivalence. -/ @[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.Resp {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [Definable n s], Resp n | _, ⟨f⟩ => f +set_option linter.deprecated false in @[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Definable.eq {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [H : Definable n s], (@Definable.Resp n s H).eval _ = s diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 3c61c4a884912..8619b96da076e 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -13,7 +13,6 @@ import Plausible import ImportGraph.Imports -- Import common Batteries tactics and commands -import Batteries.Tactic.Where import Batteries.Tactic.Basic import Batteries.Tactic.HelpCmd diff --git a/Mathlib/Tactic/Linter/FlexibleLinter.lean b/Mathlib/Tactic/Linter/FlexibleLinter.lean index 724936f1bdbee..f2a568cca5b34 100644 --- a/Mathlib/Tactic/Linter/FlexibleLinter.lean +++ b/Mathlib/Tactic/Linter/FlexibleLinter.lean @@ -343,8 +343,7 @@ Otherwise, if an `FVarId` with the same `userName` exists in the new context, us If both of these fail, return `default` (i.e. "fail"). -/ def persistFVars (fv : FVarId) (before after : LocalContext) : FVarId := let ldecl := (before.find? fv).getD default - let name := ldecl.userName - (getFVarIdCandidates fv name after).getD 0 default + (getFVarIdCandidates fv ldecl.userName after).getD 0 default /-- `reallyPersist` converts an array of pairs `(fvar, mvar)` to another array of the same type. -/ def reallyPersist @@ -376,11 +375,11 @@ def flexibleLinter : Linter where run := withSetOptionIn fun _stx => do if (← MonadState.get).messages.hasErrors then return let trees ← getInfoTrees - let x := trees.toList.map (extractCtxAndGoals (fun _ => true)) + let x := trees.map (extractCtxAndGoals (fun _ => true)) -- `stains` records pairs `(location, mvar)`, where -- * `location` is either a hypothesis or the main goal modified by a flexible tactic and -- * `mvar` is the metavariable containing the modified location - let mut stains : Array ((FVarId × MVarId) × (Stained × Syntax)) := .empty + let mut stains : Array ((FVarId × MVarId) × (Stained × Syntax)) := #[] let mut msgs : Array (Syntax × Syntax × Stained) := #[] for d in x do for (s, ctx0, ctx1, mvs0, mvs1) in d do let skind := s.getKind @@ -389,17 +388,14 @@ def flexibleLinter : Linter where run := withSetOptionIn fun _stx => do for d in getStained! s do if shouldStain? then for currMVar1 in mvs1 do - let lctx1 := ((ctx1.decls.find? currMVar1).getD default).lctx + let lctx1 := (ctx1.decls.findD currMVar1 default).lctx let locsAfter := d.toFMVarId currMVar1 lctx1 - - for l in locsAfter do - stains := stains.push (l, (d, s)) - + stains := stains ++ locsAfter.map (fun l ↦ (l, (d, s))) else let stained_in_syntax := if usesGoal? skind then (toStained s).insert d else toStained s if !flexible.contains skind then for currMv0 in mvs0 do - let lctx0 := ((ctx0.decls.find? currMv0).getD default).lctx + let lctx0 := (ctx0.decls.findD currMv0 default).lctx let mut foundFvs : Std.HashSet (FVarId × MVarId):= {} for st in stained_in_syntax do for d in st.toFMVarId currMv0 lctx0 do diff --git a/Mathlib/Tactic/Linter/HaveLetLinter.lean b/Mathlib/Tactic/Linter/HaveLetLinter.lean index fe586741660b6..d1752900da94f 100644 --- a/Mathlib/Tactic/Linter/HaveLetLinter.lean +++ b/Mathlib/Tactic/Linter/HaveLetLinter.lean @@ -49,10 +49,9 @@ register_option linter.haveLet : Nat := { namespace haveLet /-- find the `have` syntax. -/ -partial def isHave? : Syntax → Bool | .node _ ``Lean.Parser.Tactic.tacticHave_ _ => true - |_ => false + | _ => false end haveLet diff --git a/Mathlib/Tactic/Linter/Multigoal.lean b/Mathlib/Tactic/Linter/Multigoal.lean index e1d838948d7f9..cd119e103387b 100644 --- a/Mathlib/Tactic/Linter/Multigoal.lean +++ b/Mathlib/Tactic/Linter/Multigoal.lean @@ -61,7 +61,7 @@ There is some overlap in scope between `ignoreBranch` and `exclusions`. Tactic combinators like `repeat` or `try` are a mix of both. -/ -abbrev exclusions : Std.HashSet SyntaxNodeKind := .ofList [ +abbrev exclusions : Std.HashSet SyntaxNodeKind := .ofArray #[ -- structuring a proof ``Lean.Parser.Term.cdot, ``cdot, @@ -112,7 +112,7 @@ Reasons for ignoring these tactics include There is some overlap in scope between `exclusions` and `ignoreBranch`. -/ -abbrev ignoreBranch : Std.HashSet SyntaxNodeKind := .ofList [ +abbrev ignoreBranch : Std.HashSet SyntaxNodeKind := .ofArray #[ ``Lean.Parser.Tactic.Conv.conv, `Mathlib.Tactic.Conv.convLHS, `Mathlib.Tactic.Conv.convRHS, @@ -161,7 +161,7 @@ def multiGoalLinter : Linter where run := withSetOptionIn fun _stx ↦ do if (← get).messages.hasErrors then return let trees ← getInfoTrees - for t in trees.toArray do + for t in trees do for (s, before, after, n) in getManyGoals t do let goals (k : Nat) := if k == 1 then f!"1 goal" else f!"{k} goals" let fmt ← Command.liftCoreM diff --git a/Mathlib/Tactic/Linter/OldObtain.lean b/Mathlib/Tactic/Linter/OldObtain.lean index 34997e37d0fc6..f9d77e63e79e8 100644 --- a/Mathlib/Tactic/Linter/OldObtain.lean +++ b/Mathlib/Tactic/Linter/OldObtain.lean @@ -54,12 +54,15 @@ open Lean Elab namespace Mathlib.Linter.Style /-- Whether a syntax element is an `obtain` tactic call without a provided proof. -/ -def is_obtain_without_proof : Syntax → Bool +def isObtainWithoutProof : Syntax → Bool -- Using the `obtain` tactic without a proof requires proving a type; -- a pattern is optional. | `(tactic|obtain : $_type) | `(tactic|obtain $_pat : $_type) => true | _ => false +/-- Deprecated alias for `Mathlib.Linter.Style.isObtainWithoutProof`. -/ +@[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 of the `obtain` tactic, i.e. with the proof postponed. -/ @@ -74,7 +77,7 @@ def oldObtainLinter : Linter where run := withSetOptionIn fun stx => do return if (← MonadState.get).messages.hasErrors then return - if let some head := stx.find? is_obtain_without_proof then + if let some head := stx.find? isObtainWithoutProof then Linter.logLint linter.oldObtain head m!"Please remove stream-of-conciousness `obtain` syntax" initialize addLinter oldObtainLinter diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index 65272ba2d6d0a..bbe2f16f58a98 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -52,16 +52,24 @@ namespace Style.setOption /-- Whether a syntax element is a `set_option` command, tactic or term: Return the name of the option being set, if any. -/ -def parse_set_option : Syntax → Option Name +def parseSetOption : Syntax → Option Name -- This handles all four possibilities of `_val`: a string, number, `true` and `false`. | `(command|set_option $name:ident $_val) => some name.getId | `(set_option $name:ident $_val in $_x) => some name.getId | `(tactic|set_option $name:ident $_val in $_x) => some name.getId | _ => none +/-- Deprecated alias for `Mathlib.Linter.Style.setOption.parseSetOption`. -/ +@[deprecated parseSetOption (since := "2024-12-07")] +def parse_set_option := @parseSetOption + /-- Whether a given piece of syntax is a `set_option` command, tactic or term. -/ -def is_set_option : Syntax → Bool := - fun stx ↦ parse_set_option stx matches some _name +def isSetOption : Syntax → Bool := + fun stx ↦ parseSetOption stx matches some _name + +/-- Deprecated alias for `Mathlib.Linter.Style.setOption.isSetOption`. -/ +@[deprecated isSetOption (since := "2024-12-07")] +def is_set_option := @isSetOption /-- The `setOption` linter: this lints any `set_option` command, term or tactic which sets a `pp`, `profiler` or `trace` option. @@ -76,8 +84,8 @@ def setOptionLinter : Linter where run := withSetOptionIn fun stx => do return if (← MonadState.get).messages.hasErrors then return - if let some head := stx.find? is_set_option then - if let some name := parse_set_option head then + if let some head := stx.find? isSetOption then + if let some name := parseSetOption head then let forbidden := [`debug, `pp, `profiler, `trace] if forbidden.contains name.getRoot then Linter.logLint linter.style.setOption head diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 81dea423b38b4..01c6ba0915ce9 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -72,32 +72,33 @@ abbrev M := StateRefT (Std.HashMap String.Range Syntax) IO This can be increased dynamically, using `#allow_unused_tactic`. -/ initialize allowedRef : IO.Ref (Std.HashSet SyntaxNodeKind) ← - IO.mkRef <| Std.HashSet.empty - |>.insert `Mathlib.Tactic.Says.says - |>.insert `Batteries.Tactic.«tacticOn_goal-_=>_» + IO.mkRef <| .ofArray #[ + `Mathlib.Tactic.Says.says, + `Batteries.Tactic.«tacticOn_goal-_=>_», -- attempt to speed up, by ignoring more tactics - |>.insert `by - |>.insert `null - |>.insert `«]» - |>.insert ``Lean.Parser.Term.byTactic - |>.insert ``Lean.Parser.Tactic.tacticSeq - |>.insert ``Lean.Parser.Tactic.tacticSeq1Indented - |>.insert ``Lean.Parser.Tactic.tacticTry_ + `by, + `null, + `«]», + ``Lean.Parser.Term.byTactic, + ``Lean.Parser.Tactic.tacticSeq, + ``Lean.Parser.Tactic.tacticSeq1Indented, + ``Lean.Parser.Tactic.tacticTry_, -- the following `SyntaxNodeKind`s play a role in silencing `test`s - |>.insert ``Lean.Parser.Tactic.guardHyp - |>.insert ``Lean.Parser.Tactic.guardTarget - |>.insert ``Lean.Parser.Tactic.failIfSuccess - |>.insert `Mathlib.Tactic.successIfFailWithMsg - |>.insert `Mathlib.Tactic.failIfNoProgress - |>.insert `Mathlib.Tactic.ExtractGoal.extractGoal - |>.insert `Mathlib.Tactic.Propose.propose' - |>.insert `Lean.Parser.Tactic.traceState - |>.insert `Mathlib.Tactic.tacticMatch_target_ - |>.insert ``Lean.Parser.Tactic.change - |>.insert `change? - |>.insert `«tactic#adaptation_note_» - |>.insert `tacticSleep_heartbeats_ - |>.insert `Mathlib.Tactic.«tacticRename_bvar_→__» + ``Lean.Parser.Tactic.guardHyp, + ``Lean.Parser.Tactic.guardTarget, + ``Lean.Parser.Tactic.failIfSuccess, + `Mathlib.Tactic.successIfFailWithMsg, + `Mathlib.Tactic.failIfNoProgress, + `Mathlib.Tactic.ExtractGoal.extractGoal, + `Mathlib.Tactic.Propose.propose', + `Lean.Parser.Tactic.traceState, + `Mathlib.Tactic.tacticMatch_target_, + ``Lean.Parser.Tactic.change, + `change?, + `«tactic#adaptation_note_», + `tacticSleep_heartbeats_, + `Mathlib.Tactic.«tacticRename_bvar_→__» + ] /-- `#allow_unused_tactic` takes an input a space-separated list of identifiers. These identifiers are then allowed by the unused tactic linter: @@ -119,35 +120,35 @@ A list of blacklisted syntax kinds, which are expected to have subterms that con unevaluated tactics. -/ initialize ignoreTacticKindsRef : IO.Ref NameHashSet ← - IO.mkRef <| Std.HashSet.empty - |>.insert `Mathlib.Tactic.Says.says - |>.insert ``Parser.Term.binderTactic - |>.insert ``Lean.Parser.Term.dynamicQuot - |>.insert ``Lean.Parser.Tactic.quotSeq - |>.insert ``Lean.Parser.Tactic.tacticStop_ - |>.insert ``Lean.Parser.Command.notation - |>.insert ``Lean.Parser.Command.mixfix - |>.insert ``Lean.Parser.Tactic.discharger - |>.insert ``Lean.Parser.Tactic.Conv.conv - |>.insert `Batteries.Tactic.seq_focus - |>.insert `Mathlib.Tactic.Hint.registerHintStx - |>.insert `Mathlib.Tactic.LinearCombination.linearCombination - |>.insert `Mathlib.Tactic.LinearCombination'.linearCombination' - |>.insert `Aesop.Frontend.Parser.addRules - |>.insert `Aesop.Frontend.Parser.aesopTactic - |>.insert `Aesop.Frontend.Parser.aesopTactic? + IO.mkRef <| .ofArray #[ + `Mathlib.Tactic.Says.says, + ``Parser.Term.binderTactic, + ``Lean.Parser.Term.dynamicQuot, + ``Lean.Parser.Tactic.quotSeq, + ``Lean.Parser.Tactic.tacticStop_, + ``Lean.Parser.Command.notation, + ``Lean.Parser.Command.mixfix, + ``Lean.Parser.Tactic.discharger, + ``Lean.Parser.Tactic.Conv.conv, + `Batteries.Tactic.seq_focus, + `Mathlib.Tactic.Hint.registerHintStx, + `Mathlib.Tactic.LinearCombination.linearCombination, + `Mathlib.Tactic.LinearCombination'.linearCombination', + `Aesop.Frontend.Parser.addRules, + `Aesop.Frontend.Parser.aesopTactic, + `Aesop.Frontend.Parser.aesopTactic?, -- the following `SyntaxNodeKind`s play a role in silencing `test`s - |>.insert ``Lean.Parser.Tactic.failIfSuccess - |>.insert `Mathlib.Tactic.successIfFailWithMsg - |>.insert `Mathlib.Tactic.failIfNoProgress + ``Lean.Parser.Tactic.failIfSuccess, + `Mathlib.Tactic.successIfFailWithMsg, + `Mathlib.Tactic.failIfNoProgress + ] /-- Is this a syntax kind that contains intentionally unused tactic subterms? -/ def isIgnoreTacticKind (ignoreTacticKinds : NameHashSet) (k : SyntaxNodeKind) : Bool := k.components.contains `Conv || "slice".isPrefixOf k.toString || - match k with - | .str _ "quot" => true - | _ => ignoreTacticKinds.contains k + k matches .str _ "quot" || + ignoreTacticKinds.contains k /-- Adds a new syntax kind whose children will be ignored by the `unusedTactic` linter. diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 12f7051a82e2c..76c4709fd49a4 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -336,7 +336,7 @@ def pairUp : List (Expr × Bool × Syntax) → List Expr → To support a new binary operation, extend the list in this definition, so that it contains enough lemmas to allow `simp` to close a generic permutation goal for the new binary operation. -/ -def move_oper_simpCtx : MetaM Simp.Context := do +def moveOperSimpCtx : MetaM Simp.Context := do let simpNames := Elab.Tactic.simpOnlyBuiltins ++ [ ``add_comm, ``add_assoc, ``add_left_comm, -- for `HAdd.hAdd` ``mul_comm, ``mul_assoc, ``mul_left_comm, -- for `HMul.hMul` @@ -348,6 +348,9 @@ def move_oper_simpCtx : MetaM Simp.Context := do let simpThms ← simpNames.foldlM (·.addConst ·) ({} : SimpTheorems) Simp.mkContext {} (simpTheorems := #[simpThms]) +@[deprecated (since := "2024-12-07")] +alias move_oper_simpCtx := moveOperSimpCtx + /-- `reorderAndSimp mv op instr` takes as input an `MVarId` `mv`, the name `op` of a binary operation and a list of "instructions" `instr` that it passes to `permuteExpr`. @@ -368,7 +371,7 @@ def reorderAndSimp (mv : MVarId) (instr : List (Expr × Bool)) : throwError m!"There should only be 2 goals, instead of {twoGoals.length}" -- `permGoal` is the single goal `mv_permuted`, possibly more operations will be permuted later on let permGoal ← twoGoals.filterM fun v => return !(← v.isAssigned) - match ← (simpGoal (permGoal[1]!) (← move_oper_simpCtx)) with + match ← (simpGoal (permGoal[1]!) (← moveOperSimpCtx)) with | (some x, _) => throwError m!"'move_oper' could not solve {indentD x.2}" | (none, _) => return permGoal diff --git a/Mathlib/Tactic/Ring/RingNF.lean b/Mathlib/Tactic/Ring/RingNF.lean index 3d57d68676c46..095a8b819b8d4 100644 --- a/Mathlib/Tactic/Ring/RingNF.lean +++ b/Mathlib/Tactic/Ring/RingNF.lean @@ -259,9 +259,17 @@ example (x : ℕ) (h : x * 2 > 5): x + x > 5 := by ring; assumption -- suggests ``` -/ macro (name := ring) "ring" : tactic => - `(tactic| first | ring1 | try_this ring_nf) + `(tactic| first | ring1 | try_this ring_nf + "\n\nThe `ring` tactic failed to close the goal. Use `ring_nf` to obtain a normal form. + \nNote that `ring` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") @[inherit_doc ring] macro "ring!" : tactic => - `(tactic| first | ring1! | try_this ring_nf!) + `(tactic| first | ring1! | try_this ring_nf! + "\n\nThe `ring!` tactic failed to close the goal. Use `ring_nf!` to obtain a normal form. + \nNote that `ring!` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") /-- The tactic `ring` evaluates expressions in *commutative* (semi)rings. @@ -270,9 +278,17 @@ This is the conv tactic version, which rewrites a target which is a ring equalit See also the `ring` tactic. -/ macro (name := ringConv) "ring" : conv => - `(conv| first | discharge => ring1 | try_this ring_nf) + `(conv| first | discharge => ring1 | try_this ring_nf + "\n\nThe `ring` tactic failed to close the goal. Use `ring_nf` to obtain a normal form. + \nNote that `ring` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") @[inherit_doc ringConv] macro "ring!" : conv => - `(conv| first | discharge => ring1! | try_this ring_nf!) + `(conv| first | discharge => ring1! | try_this ring_nf! + "\n\nThe `ring!` tactic failed to close the goal. Use `ring_nf!` to obtain a normal form. + \nNote that `ring!` works primarily in *commutative* rings. \ + If you have a noncommutative ring, abelian group or module, consider using \ + `noncomm_ring`, `abel` or `module` instead.") end RingNF diff --git a/Mathlib/Tactic/ToExpr.lean b/Mathlib/Tactic/ToExpr.lean index 8997800729532..d915f530ccf7d 100644 --- a/Mathlib/Tactic/ToExpr.lean +++ b/Mathlib/Tactic/ToExpr.lean @@ -16,37 +16,6 @@ that come from core Lean 4 that do not handle universe polymorphism. In addition, we provide some additional `ToExpr` instances for core definitions. -/ -section override -- Note: this section uses `autoImplicit` pervasively -namespace Lean - -attribute [-instance] Lean.instToExprOption - -set_option autoImplicit true in -deriving instance ToExpr for Option - -attribute [-instance] Lean.instToExprList - -set_option autoImplicit true in -deriving instance ToExpr for List - -attribute [-instance] Lean.instToExprArray - -universe u in -instance {α : Type u} [ToExpr α] [ToLevel.{u}] : ToExpr (Array α) := - let type := toTypeExpr α - { toExpr := fun as => mkApp2 (mkConst ``List.toArray [toLevel.{u}]) type (toExpr as.toList) - toTypeExpr := mkApp (mkConst ``Array [toLevel.{u}]) type } - -attribute [-instance] Lean.instToExprProd - -set_option autoImplicit true in -deriving instance ToExpr for Prod - -deriving instance ToExpr for System.FilePath - -end Lean -end override - namespace Mathlib open Lean @@ -62,7 +31,6 @@ instance [ToLevel.{u}] : ToExpr PUnit.{u+1} where deriving instance ToExpr for String.Pos deriving instance ToExpr for Substring deriving instance ToExpr for SourceInfo -deriving instance ToExpr for Syntax.Preresolved deriving instance ToExpr for Syntax open DataValue in @@ -86,12 +54,10 @@ instance : ToExpr MData where toExpr := toExprMData toTypeExpr := mkConst ``MData -deriving instance ToExpr for FVarId deriving instance ToExpr for MVarId deriving instance ToExpr for LevelMVarId deriving instance ToExpr for Level deriving instance ToExpr for BinderInfo -deriving instance ToExpr for Literal deriving instance ToExpr for Expr end Mathlib diff --git a/Mathlib/Tactic/ToLevel.lean b/Mathlib/Tactic/ToLevel.lean index 17b243651eb54..fd85a5198e967 100644 --- a/Mathlib/Tactic/ToLevel.lean +++ b/Mathlib/Tactic/ToLevel.lean @@ -17,31 +17,6 @@ override the ones from Lean 4 core. namespace Lean -/-- A class to create `Level` expressions that denote particular universe levels in Lean. -`Lean.ToLevel.toLevel.{u}` evaluates to a `Lean.Level` term representing `u` -/ -@[pp_with_univ] -class ToLevel.{u} where - /-- A `Level` that represents the universe level `u`. -/ - toLevel : Level - /-- The universe itself. This is only here to avoid the "unused universe parameter" error. -/ - univ : Type u := Sort u -export ToLevel (toLevel) attribute [pp_with_univ] toLevel -instance : ToLevel.{0} where - toLevel := .zero - -universe u v - -instance [ToLevel.{u}] : ToLevel.{u+1} where - toLevel := .succ toLevel.{u} - -/-- `ToLevel` for `max u v`. This is not an instance since it causes divergence. -/ -def ToLevel.max [ToLevel.{u}] [ToLevel.{v}] : ToLevel.{max u v} where - toLevel := .max toLevel.{u} toLevel.{v} - -/-- `ToLevel` for `imax u v`. This is not an instance since it causes divergence. -/ -def ToLevel.imax [ToLevel.{u}] [ToLevel.{v}] : ToLevel.{imax u v} where - toLevel := .imax toLevel.{u} toLevel.{v} - end Lean diff --git a/Mathlib/Tactic/TryThis.lean b/Mathlib/Tactic/TryThis.lean index 3d4814c3d4087..f34ea59970cb9 100644 --- a/Mathlib/Tactic/TryThis.lean +++ b/Mathlib/Tactic/TryThis.lean @@ -18,13 +18,17 @@ namespace Mathlib.Tactic open Lean /-- Produces the text `Try this: ` with the given tactic, and then executes it. -/ -elab tk:"try_this" tac:tactic : tactic => do +elab tk:"try_this" tac:tactic info:(str)? : tactic => do Elab.Tactic.evalTactic tac - Meta.Tactic.TryThis.addSuggestion tk tac (origSpan? := ← getRef) + Meta.Tactic.TryThis.addSuggestion tk + { suggestion := tac, postInfo? := TSyntax.getString <$> info } + (origSpan? := ← getRef) /-- Produces the text `Try this: ` with the given conv tactic, and then executes it. -/ -elab tk:"try_this" tac:conv : conv => do +elab tk:"try_this" tac:conv info:(str)? : conv => do Elab.Tactic.evalTactic tac - Meta.Tactic.TryThis.addSuggestion tk tac (origSpan? := ← getRef) + Meta.Tactic.TryThis.addSuggestion tk + { suggestion := tac, postInfo? := TSyntax.getString <$> info } + (origSpan? := ← getRef) end Mathlib.Tactic diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 4b003dcf49f11..6746a7d253f9f 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -233,6 +233,9 @@ theorem ContinuousWithinAt.inv (hf : ContinuousWithinAt f s x) : ContinuousWithinAt (fun x => (f x)⁻¹) s x := Filter.Tendsto.inv hf +@[to_additive] +instance OrderDual.instContinuousInv : ContinuousInv Gᵒᵈ := ‹ContinuousInv G› + @[to_additive] instance Prod.continuousInv [TopologicalSpace H] [Inv H] [ContinuousInv H] : ContinuousInv (G × H) := @@ -536,6 +539,9 @@ end OrderedCommGroup instance [TopologicalSpace H] [Group H] [TopologicalGroup H] : TopologicalGroup (G × H) where continuous_inv := continuous_inv.prodMap continuous_inv +@[to_additive] +instance OrderDual.instTopologicalGroup : TopologicalGroup Gᵒᵈ where + @[to_additive] instance Pi.topologicalGroup {C : β → Type*} [∀ b, TopologicalSpace (C b)] [∀ b, Group (C b)] [∀ b, TopologicalGroup (C b)] : TopologicalGroup (∀ b, C b) where diff --git a/Mathlib/Topology/Algebra/Indicator.lean b/Mathlib/Topology/Algebra/Indicator.lean new file mode 100644 index 0000000000000..189c9d6f29038 --- /dev/null +++ b/Mathlib/Topology/Algebra/Indicator.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2024 PFR contributors. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: PFR contributors +-/ +import Mathlib.Algebra.Group.Indicator +import Mathlib.Topology.ContinuousOn + +/-! +# Continuity of indicator functions +-/ + +open Set +open scoped Topology + +variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {f : α → β} {s : Set α} [One β] + +@[to_additive] +lemma continuous_mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : ContinuousOn f (closure s)) : + Continuous (mulIndicator s f) := by + classical exact continuous_piecewise hs hf continuousOn_const + +@[to_additive] +protected lemma Continuous.mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : Continuous f) : + Continuous (mulIndicator s f) := by + classical exact hf.piecewise hs continuous_const + +@[to_additive] +theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s)) {x : α} + (hx : x ∉ frontier s) : + ContinuousAt (s.mulIndicator f) x := by + rw [← Set.mem_compl_iff, compl_frontier_eq_union_interior] at hx + obtain h | h := hx + · have hs : interior s ∈ 𝓝 x := mem_interior_iff_mem_nhds.mp (by rwa [interior_interior]) + exact ContinuousAt.congr (hf.continuousAt hs) <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨interior s, hs, Set.eqOn_mulIndicator.symm.mono interior_subset⟩ + · exact ContinuousAt.congr continuousAt_const <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨sᶜ, mem_interior_iff_mem_nhds.mp h, Set.eqOn_mulIndicator'.symm⟩ diff --git a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean index 1f2427e29b56f..6790b8c087127 100644 --- a/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean +++ b/Mathlib/Topology/Algebra/SeparationQuotient/Basic.lean @@ -85,7 +85,7 @@ variable {M : Type*} [TopologicalSpace M] @[to_additive] instance instMul [Mul M] [ContinuousMul M] : Mul (SeparationQuotient M) where - mul := Quotient.map₂' (· * ·) fun _ _ h₁ _ _ h₂ ↦ Inseparable.mul h₁ h₂ + mul := Quotient.map₂ (· * ·) fun _ _ h₁ _ _ h₂ ↦ Inseparable.mul h₁ h₂ @[to_additive (attr := simp)] theorem mk_mul [Mul M] [ContinuousMul M] (a b : M) : mk (a * b) = mk a * mk b := rfl @@ -168,7 +168,7 @@ instance instInvOneClass [InvOneClass G] [ContinuousInv G] : @[to_additive] instance instDiv [Div G] [ContinuousDiv G] : Div (SeparationQuotient G) where - div := Quotient.map₂' (· / ·) fun _ _ h₁ _ _ h₂ ↦ (Inseparable.prod h₁ h₂).map continuous_div' + div := Quotient.map₂ (· / ·) fun _ _ h₁ _ _ h₂ ↦ (Inseparable.prod h₁ h₂).map continuous_div' @[to_additive (attr := simp)] theorem mk_div [Div G] [ContinuousDiv G] (x y : G) : mk (x / y) = mk x / mk y := rfl diff --git a/Mathlib/Topology/Bases.lean b/Mathlib/Topology/Bases.lean index 4db6b21371950..969cd3b6dd3fa 100644 --- a/Mathlib/Topology/Bases.lean +++ b/Mathlib/Topology/Bases.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Data.Set.Constructions +import Mathlib.Order.Filter.AtTopBot.CountablyGenerated import Mathlib.Topology.Constructions import Mathlib.Topology.ContinuousOn diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 42749f3df74e4..686efbf1e5992 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1055,10 +1055,9 @@ theorem Products.limitOrdinal (l : Products I) : l.isGood (π C (ord I · < o)) ∃ (o' : Ordinal), o' < o ∧ l.isGood (π C (ord I · < o')) := by refine ⟨fun h ↦ ?_, fun ⟨o', ⟨ho', hl⟩⟩ ↦ isGood_mono C (le_of_lt ho') hl⟩ use Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) - have ha : ⊥ < o := by rw [Ordinal.bot_eq_zero, Ordinal.pos_iff_ne_zero]; exact ho.1 have hslt : Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) < o := by - simp only [Finset.sup_lt_iff ha, List.mem_toFinset] - exact fun b hb ↦ ho.2 _ (prop_of_isGood C (ord I · < o) h b hb) + simp only [Finset.sup_lt_iff ho.pos, List.mem_toFinset] + exact fun b hb ↦ ho.succ_lt (prop_of_isGood C (ord I · < o) h b hb) refine ⟨hslt, fun he ↦ h ?_⟩ have hlt : ∀ i ∈ l.val, ord I i < Finset.sup l.val.toFinset (fun a ↦ Order.succ (ord I a)) := by intro i hi @@ -1354,14 +1353,14 @@ theorem CC_exact {f : LocallyConstant C ℤ} (hf : Linear_CC' C hsC ho f = 0) : exact C1_projOrd C hsC ho hx₁ variable (o) in -theorem succ_mono : CategoryTheory.Mono (ModuleCat.asHom (πs C o)) := by +theorem succ_mono : CategoryTheory.Mono (ModuleCat.ofHom (πs C o)) := by rw [ModuleCat.mono_iff_injective] exact injective_πs _ _ include hC in theorem succ_exact : - (ShortComplex.mk (ModuleCat.asHom (πs C o)) (ModuleCat.asHom (Linear_CC' C hsC ho)) - (by ext; apply CC_comp_zero)).Exact := by + (ShortComplex.mk (ModuleCat.ofHom (πs C o)) (ModuleCat.ofHom (Linear_CC' C hsC ho)) + (by ext : 2; apply CC_comp_zero)).Exact := by rw [ShortComplex.moduleCat_exact_iff] intro f exact CC_exact C hC hsC ho @@ -1478,7 +1477,7 @@ theorem span_sum : Set.range (eval C) = Set.range (Sum.elim theorem square_commutes : SumEval C ho ∘ Sum.inl = - ModuleCat.asHom (πs C o) ∘ eval (π C (ord I · < o)) := by + ModuleCat.ofHom (πs C o) ∘ eval (π C (ord I · < o)) := by ext l dsimp [SumEval] rw [← Products.eval_πs C (Products.prop_of_isGood _ _ l.prop)] @@ -1644,7 +1643,7 @@ theorem maxTail_isGood (l : MaxProducts C ho) rfl have hse := succ_exact C hC hsC ho rw [ShortComplex.moduleCat_exact_iff_range_eq_ker] at hse - dsimp [ModuleCat.asHom] at hse + dsimp [ModuleCat.ofHom] at hse -- Rewrite `this` using exact sequence manipulations to conclude that a term is in the range of -- the linear map `πs`: @@ -1662,7 +1661,7 @@ theorem maxTail_isGood (l : MaxProducts C ho) apply Submodule.add_mem · apply Submodule.finsupp_sum_mem intro q _ - erw [LinearMap.map_smul (fₗ := πs C o) (c := w q) (x := eval (π C (ord I · < o)) q)] + rw [LinearMap.map_smul] apply Submodule.smul_mem apply Submodule.subset_span dsimp only [eval] @@ -1701,9 +1700,9 @@ include hC in theorem linearIndependent_comp_of_eval (h₁ : ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : LinearIndependent ℤ (eval (C' C ho)) → - LinearIndependent ℤ (ModuleCat.asHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by - dsimp [SumEval, ModuleCat.asHom] - erw [max_eq_eval_unapply C hsC ho] + LinearIndependent ℤ (ModuleCat.ofHom (Linear_CC' C hsC ho) ∘ SumEval C ho ∘ Sum.inr) := by + dsimp [SumEval, ModuleCat.ofHom] + rw [max_eq_eval_unapply C hsC ho] intro h let f := MaxToGood C hC hsC ho h₁ have hf : f.Injective := maxToGood_injective C hC hsC ho h₁ diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 1149080e64cd9..15e1b303bf533 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -698,7 +698,7 @@ theorem nhdsSet_prod_le_of_disjoint_cocompact {f : Filter Y} (hs : IsCompact s) _ = 𝓝ˢ (s ×ˢ K) := (hs.nhdsSet_prod_eq hK).symm _ ≤ 𝓝ˢ (s ×ˢ Set.univ) := nhdsSet_mono (prod_mono_right le_top) -theorem prod_nhdsSet_le_of_disjoint_cocompact {f : Filter X} (ht : IsCompact t) +theorem prod_nhdsSet_le_of_disjoint_cocompact {t : Set Y} {f : Filter X} (ht : IsCompact t) (hf : Disjoint f (Filter.cocompact X)) : f ×ˢ 𝓝ˢ t ≤ 𝓝ˢ (Set.univ ×ˢ t) := by obtain ⟨K, hKf, hK⟩ := (disjoint_cocompact_right f).mp hf @@ -709,6 +709,16 @@ theorem prod_nhdsSet_le_of_disjoint_cocompact {f : Filter X} (ht : IsCompact t) _ = 𝓝ˢ (K ×ˢ t) := (hK.nhdsSet_prod_eq ht).symm _ ≤ 𝓝ˢ (Set.univ ×ˢ t) := nhdsSet_mono (prod_mono_left le_top) +theorem nhds_prod_le_of_disjoint_cocompact {f : Filter Y} (x : X) + (hf : Disjoint f (Filter.cocompact Y)) : + 𝓝 x ×ˢ f ≤ 𝓝ˢ ({x} ×ˢ Set.univ) := by + simpa using nhdsSet_prod_le_of_disjoint_cocompact isCompact_singleton hf + +theorem prod_nhds_le_of_disjoint_cocompact {f : Filter X} (y : Y) + (hf : Disjoint f (Filter.cocompact X)) : + f ×ˢ 𝓝 y ≤ 𝓝ˢ (Set.univ ×ˢ {y}) := by + simpa using prod_nhdsSet_le_of_disjoint_cocompact isCompact_singleton hf + /-- If `s` and `t` are compact sets and `n` is an open neighborhood of `s × t`, then there exist open neighborhoods `u ⊇ s` and `v ⊇ t` such that `u × v ⊆ n`. @@ -867,6 +877,16 @@ theorem Filter.comap_cocompact_le {f : X → Y} (hf : Continuous f) : refine ⟨f '' t, ht.image hf, ?_⟩ simpa using t.subset_preimage_image f +/-- If a filter is disjoint from the cocompact filter, so is its image under any continuous +function. -/ +theorem disjoint_map_cocompact {g : X → Y} {f : Filter X} (hg : Continuous g) + (hf : Disjoint f (Filter.cocompact X)) : Disjoint (map g f) (Filter.cocompact Y) := by + rw [← Filter.disjoint_comap_iff_map, disjoint_iff_inf_le] + calc + f ⊓ (comap g (cocompact Y)) + _ ≤ f ⊓ Filter.cocompact X := inf_le_inf_left f (Filter.comap_cocompact_le hg) + _ = ⊥ := disjoint_iff.mp hf + theorem isCompact_range [CompactSpace X] {f : X → Y} (hf : Continuous f) : IsCompact (range f) := by rw [← image_univ]; exact isCompact_univ.image hf diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index 81b40ea72345a..5c2ccedeed4ea 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -4,11 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Mario Carneiro, Yury Kudryashov, Heather Macbeth -/ import Mathlib.Algebra.Module.MinimalAxioms -import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic -import Mathlib.Topology.Bornology.BoundedOperation import Mathlib.Tactic.Monotonicity +import Mathlib.Topology.Algebra.Indicator +import Mathlib.Topology.Bornology.BoundedOperation +import Mathlib.Topology.ContinuousMap.Algebra /-! # Bounded continuous functions diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 3e1d97b1dedd0..56c76196a7d53 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Algebra.Group.Indicator import Mathlib.Topology.Constructions /-! @@ -1403,21 +1402,6 @@ theorem Continuous.piecewise [∀ a, Decidable (a ∈ s)] Continuous (piecewise s f g) := hf.if hs hg -section Indicator -variable [One β] - -@[to_additive] -lemma continuous_mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : ContinuousOn f (closure s)) : - Continuous (mulIndicator s f) := by - classical exact continuous_piecewise hs hf continuousOn_const - -@[to_additive] -protected lemma Continuous.mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf : Continuous f) : - Continuous (mulIndicator s f) := by - classical exact hf.piecewise hs continuous_const - -end Indicator - theorem IsOpen.ite' (hs : IsOpen s) (hs' : IsOpen s') (ht : ∀ x ∈ frontier t, x ∈ s ↔ x ∈ s') : IsOpen (t.ite s s') := by classical diff --git a/Mathlib/Topology/IndicatorConstPointwise.lean b/Mathlib/Topology/IndicatorConstPointwise.lean index deb980b52a7e8..0ec0a97c72bba 100644 --- a/Mathlib/Topology/IndicatorConstPointwise.lean +++ b/Mathlib/Topology/IndicatorConstPointwise.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ +import Mathlib.Algebra.Group.Indicator import Mathlib.Topology.Separation.Basic /-! diff --git a/Mathlib/Topology/NhdsSet.lean b/Mathlib/Topology/NhdsSet.lean index 9b361653607e1..f0aa111474378 100644 --- a/Mathlib/Topology/NhdsSet.lean +++ b/Mathlib/Topology/NhdsSet.lean @@ -122,7 +122,7 @@ theorem mem_nhdsSet_empty : s ∈ 𝓝ˢ (∅ : Set X) := by simp @[simp] theorem nhdsSet_univ : 𝓝ˢ (univ : Set X) = ⊤ := by rw [isOpen_univ.nhdsSet_eq, principal_univ] -@[mono] +@[gcongr, mono] theorem nhdsSet_mono (h : s ⊆ t) : 𝓝ˢ s ≤ 𝓝ˢ t := sSup_le_sSup <| image_subset _ h diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index 59abac51b6b7c..38a08ce558259 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -472,9 +472,10 @@ theorem atBot_le_nhds_bot [OrderBot α] : (atBot : Filter α) ≤ 𝓝 ⊥ := by rw [OrderBot.atBot_eq] apply pure_le_nhds +set_option linter.deprecated false in @[deprecated OrderTop.atTop_eq (since := "2024-02-14")] theorem atTop_le_nhds_top [OrderTop α] : (atTop : Filter α) ≤ 𝓝 ⊤ := - set_option linter.deprecated false in @atBot_le_nhds_bot αᵒᵈ _ _ _ + @atBot_le_nhds_bot αᵒᵈ _ _ _ variable (α) diff --git a/Mathlib/Topology/Order/LeftRight.lean b/Mathlib/Topology/Order/LeftRight.lean index e18f75c1f4fa6..065632d5f1161 100644 --- a/Mathlib/Topology/Order/LeftRight.lean +++ b/Mathlib/Topology/Order/LeftRight.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ +import Mathlib.Order.Antichain import Mathlib.Topology.ContinuousOn /-! diff --git a/Mathlib/Topology/Separation/Basic.lean b/Mathlib/Topology/Separation/Basic.lean index b56a3a27aa9a9..827ce1e31923c 100644 --- a/Mathlib/Topology/Separation/Basic.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.Group.Support import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 9f28773dd89af..ef9363d4616d6 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1217,6 +1217,10 @@ theorem UniformContinuous.continuous [UniformSpace α] [UniformSpace β] {f : α instance ULift.uniformSpace [UniformSpace α] : UniformSpace (ULift α) := UniformSpace.comap ULift.down ‹_› +/-- Uniform space structure on `αᵒᵈ`. -/ +instance OrderDual.instUniformSpace [UniformSpace α] : UniformSpace (αᵒᵈ) := + ‹UniformSpace α› + section UniformContinuousInfi -- Porting note: renamed for dot notation; add an `iff` lemma? diff --git a/Mathlib/Topology/UniformSpace/Dini.lean b/Mathlib/Topology/UniformSpace/Dini.lean new file mode 100644 index 0000000000000..9e90b66608f20 --- /dev/null +++ b/Mathlib/Topology/UniformSpace/Dini.lean @@ -0,0 +1,159 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.Normed.Order.Lattice +import Mathlib.Topology.ContinuousMap.Ordered +import Mathlib.Topology.UniformSpace.CompactConvergence + +/-! # Dini's Theorem + +This file proves Dini's theorem, which states that if `F n` is a monotone increasing sequence of +continuous real-valued functions on a compact set `s` converging pointwise to a continuous function +`f`, then `F n` converges uniformly to `f`. + +We generalize the codomain from `ℝ` to a normed lattice additive commutative group `G`. +This theorem is true in a different generality as well: when `G` is a linearly ordered topological +group with the order topology. This weakens the norm assumption, in exchange for strengthening to +a linear order. This separate generality is not included in this file, but that generality was +included in initial drafts of the original +[PR #19068](https://github.com/leanprover-community/mathlib4/pull/19068) and can be recovered if +necessary. + +The key idea of the proof is to use a particular basis of `𝓝 0` which consists of open sets that +are somehow monotone in the sense that if `s` is in the basis, and `0 ≤ x ≤ y`, then +`y ∈ s → x ∈ s`, and so the proof would work on any topological ordered group possessing +such a basis. In the case of a linearly ordered topological group with the order topology, this +basis is `nhds_basis_Ioo`. In the case of a normed lattice additive commutative group, this basis +is `nhds_basis_ball`, and the fact that this basis satisfies the monotonicity criterion +corresponds to `HasSolidNorm`. +-/ + +open Filter Topology + +variable {ι α G : Type*} [Preorder ι] [TopologicalSpace α] [NormedLatticeAddCommGroup G] + +section Unbundled + +open Metric + +variable {F : ι → α → G} {f : α → G} + +namespace Monotone + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions +converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ +lemma tendstoLocallyUniformly_of_forall_tendsto + (hF_cont : ∀ i, Continuous (F i)) (hF_mono : Monotone F) (hf : Continuous f) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformly F f atTop := by + refine (atTop : Filter ι).eq_or_neBot.elim (fun h ↦ ?eq_bot) (fun _ ↦ ?_) + case eq_bot => simp [h, tendstoLocallyUniformly_iff_forall_tendsto] + have F_le_f (x : α) (n : ι) : F n x ≤ f x := by + refine ge_of_tendsto (h_tendsto x) ?_ + filter_upwards [Ici_mem_atTop n] with m hnm + exact hF_mono hnm x + simp_rw [Metric.tendstoLocallyUniformly_iff, dist_eq_norm'] + intro ε ε_pos x + simp_rw +singlePass [tendsto_iff_norm_sub_tendsto_zero] at h_tendsto + obtain ⟨n, hn⟩ := (h_tendsto x).eventually (eventually_lt_nhds ε_pos) |>.exists + refine ⟨{y | ‖F n y - f y‖ < ε}, ⟨isOpen_lt (by fun_prop) continuous_const |>.mem_nhds hn, ?_⟩⟩ + filter_upwards [eventually_ge_atTop n] with m hnm z hz + refine norm_le_norm_of_abs_le_abs ?_ |>.trans_lt hz + simp only [abs_of_nonpos (sub_nonpos_of_le (F_le_f _ _)), neg_sub, sub_le_sub_iff_left] + exact hF_mono hnm z + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +set `s` converging pointwise to a continuous funciton `f`, then `F n` converges locally uniformly +to `f`. -/ +lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_mono : ∀ x ∈ s, Monotone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformlyOn F f atTop s := by + rw [tendstoLocallyUniformlyOn_iff_tendstoLocallyUniformly_comp_coe] + exact tendstoLocallyUniformly_of_forall_tendsto (hF_cont · |>.restrict) + (fun _ _ h x ↦ hF_mono _ x.2 h) hf.restrict (fun x ↦ h_tendsto x x.2) + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +compact space converging pointwise to a continuous function `f`, then `F n` converges uniformly to +`f`. -/ +lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Continuous (F i)) + (hF_mono : Monotone F) (hf : Continuous f) (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformly F f atTop := + tendstoLocallyUniformly_iff_tendstoUniformly_of_compactSpace.mp <| + tendstoLocallyUniformly_of_forall_tendsto hF_cont hF_mono hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions on a +compact set `s` converging pointwise to a continuous function `f`, then `F n` converges uniformly +to `f`. -/ +lemma tendstoUniformlyOn_of_forall_tendsto {s : Set α} (hs : IsCompact s) + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_mono : ∀ x ∈ s, Monotone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformlyOn F f atTop s := + tendstoLocallyUniformlyOn_iff_tendstoUniformlyOn_of_compact hs |>.mp <| + tendstoLocallyUniformlyOn_of_forall_tendsto hF_cont hF_mono hf h_tendsto + +end Monotone + +namespace Antitone + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +converging pointwise to a continuous function `f`, then `F n` converges locally uniformly to `f`. -/ +lemma tendstoLocallyUniformly_of_forall_tendsto + (hF_cont : ∀ i, Continuous (F i)) (hF_anti : Antitone F) (hf : Continuous f) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformly F f atTop := + Monotone.tendstoLocallyUniformly_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +set `s` converging pointwise to a continuous function `f`, then `F n` converges locally uniformly +to `f`. -/ +lemma tendstoLocallyUniformlyOn_of_forall_tendsto {s : Set α} + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_anti : ∀ x ∈ s, Antitone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoLocallyUniformlyOn F f atTop s := + Monotone.tendstoLocallyUniformlyOn_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +compact space converging pointwise to a continuous function `f`, then `F n` converges uniformly +to `f`. -/ +lemma tendstoUniformly_of_forall_tendsto [CompactSpace α] (hF_cont : ∀ i, Continuous (F i)) + (hF_anti : Antitone F) (hf : Continuous f) (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformly F f atTop := + Monotone.tendstoUniformly_of_forall_tendsto (G := Gᵒᵈ) hF_cont hF_anti hf h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions on a +compact set `s` converging pointwise to a continuous `f`, then `F n` converges uniformly to `f`. -/ +lemma tendstoUniformlyOn_of_forall_tendsto {s : Set α} (hs : IsCompact s) + (hF_cont : ∀ i, ContinuousOn (F i) s) (hF_anti : ∀ x ∈ s, Antitone (F · x)) + (hf : ContinuousOn f s) (h_tendsto : ∀ x ∈ s, Tendsto (F · x) atTop (𝓝 (f x))) : + TendstoUniformlyOn F f atTop s := + Monotone.tendstoUniformlyOn_of_forall_tendsto (G := Gᵒᵈ) hs hF_cont hF_anti hf h_tendsto + +end Antitone + +end Unbundled + +namespace ContinuousMap + +variable {F : ι → C(α, G)} {f : C(α, G)} + +/-- **Dini's theorem** If `F n` is a monotone increasing collection of continuous functions +converging pointwise to a continuous function `f`, then `F n` converges to `f` in the +compact-open topology. -/ +lemma tendsto_of_monotone_of_pointwise (hF_mono : Monotone F) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + Tendsto F atTop (𝓝 f) := + tendsto_of_tendstoLocallyUniformly <| + hF_mono.tendstoLocallyUniformly_of_forall_tendsto (F · |>.continuous) f.continuous h_tendsto + +/-- **Dini's theorem** If `F n` is a monotone decreasing collection of continuous functions +converging pointwise to a continuous function `f`, then `F n` converges to `f` in the +compact-open topology. -/ +lemma tendsto_of_antitone_of_pointwise (hF_anti : Antitone F) + (h_tendsto : ∀ x, Tendsto (F · x) atTop (𝓝 (f x))) : + Tendsto F atTop (𝓝 f) := + tendsto_of_monotone_of_pointwise (G := Gᵒᵈ) hF_anti h_tendsto + +end ContinuousMap diff --git a/MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean b/MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean new file mode 100644 index 0000000000000..650d59aeef1ba --- /dev/null +++ b/MathlibTest/CategoryTheory/ConcreteCategory/ModuleCat.lean @@ -0,0 +1,53 @@ +import Mathlib.Algebra.Category.ModuleCat.Basic + +universe v u + +open CategoryTheory ModuleCat + +set_option maxHeartbeats 10000 +set_option synthInstance.maxHeartbeats 2000 + +variable (R : Type u) [CommRing R] + +/- We test if all the coercions and `map_add` lemmas trigger correctly. -/ + +example (X : Type v) [AddCommGroup X] [Module R X] : ⇑(𝟙 (of R X)) = id := by simp + +example {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) : + ⇑(ModuleCat.ofHom f) = ⇑f := by simp + +example {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) + (x : X) : (ModuleCat.ofHom f) x = f x := by simp + +example {X Y Z : ModuleCat R} (f : X ⟶ Y) (g : Y ⟶ Z) : ⇑(f ≫ g) = ⇑g ∘ ⇑f := by simp + +example {X Y Z : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] [Ring Z] + [Algebra R Z] (f : X →ₗ[R] Y) (g : Y →ₗ[R] Z) : + ⇑(ModuleCat.ofHom f ≫ ModuleCat.ofHom g) = g ∘ f := by simp + +example {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] {Z : ModuleCat R} + (f : X →ₗ[R] Y) (g : of R Y ⟶ Z) : + ⇑(ModuleCat.ofHom f ≫ g) = g ∘ f := by simp + +example {X Y : ModuleCat R} {Z : Type v} [Ring Z] [Algebra R Z] (f : X ⟶ Y) (g : Y ⟶ of R Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {Y Z : ModuleCat R} {X : Type v} [AddCommGroup X] [Module R X] (f : of R X ⟶ Y) (g : Y ⟶ Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {X Y Z : ModuleCat R} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := by simp + +example {X Y : ModuleCat R} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by simp + +example {X Y : ModuleCat R} (e : X ≅ Y) (y : Y) : e.hom (e.inv y) = y := by simp + +example (X : ModuleCat R) : ⇑(𝟙 X) = id := by simp + +example {M N : ModuleCat.{v} R} (f : M ⟶ N) (x y : M) : f (x + y) = f x + f y := by + simp + +example {M N : ModuleCat.{v} R} (f : M ⟶ N) : f 0 = 0 := by + simp + +example {M N : ModuleCat.{v} R} (f : M ⟶ N) (r : R) (m : M) : f (r • m) = r • f m := by + simp diff --git a/MathlibTest/HashCommandLinter.lean b/MathlibTest/HashCommandLinter.lean index 4193342406c14..cf5304e11e734 100644 --- a/MathlibTest/HashCommandLinter.lean +++ b/MathlibTest/HashCommandLinter.lean @@ -2,6 +2,12 @@ import Lean.Elab.GuardMsgs import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Linter.HashCommandLinter + +-- #adaptation_note +-- The hashCommand linter started failing after https://github.com/leanprover/lean4/pull/6299 +-- Disabling aync elaboration fixes it, but of course we're not going to do that globally. +set_option Elab.async false + set_option linter.hashCommand true section ignored_commands diff --git a/MathlibTest/HaveLetLinter.lean b/MathlibTest/HaveLetLinter.lean index ec7409b66ccde..96d2aae3e74e1 100644 --- a/MathlibTest/HaveLetLinter.lean +++ b/MathlibTest/HaveLetLinter.lean @@ -1,6 +1,12 @@ import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Tauto + +-- #adaptation_note +-- The haveLet linter started failing after https://github.com/leanprover/lean4/pull/6299 +-- Disabling aync elaboration fixes it, but of course we're not going to do that globally. +set_option Elab.async false + set_option linter.haveLet 1 /-- diff --git a/docs/references.bib b/docs/references.bib index 976afa9ae6d69..69f11f4a17b95 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2602,6 +2602,16 @@ @Article{ markowsky1976 pages = {53-68} } +@Article{ mason2017, + title = {Vertex rings and Pierce bundles}, + author = {Geoffrey Mason}, + year = {2017}, + eprint = {1707.00328}, + archiveprefix = {arXiv}, + primaryclass = {math.RA}, + url = {https://arxiv.org/abs/1707.00328v1} +} + @InProceedings{ mathlib2020, author = {The mathlib Community}, title = {The Lean Mathematical Library}, @@ -2620,6 +2630,17 @@ @InProceedings{ mathlib2020 series = {CPP 2020} } +@Article{ matsuo1997, + title = {On axioms for a vertex algebra and locality of quantum + fields}, + author = {Atsushi Matsuo, Kiyokazu Nagatomo}, + year = {1997}, + eprint = {9706118}, + archiveprefix = {arXiv}, + primaryclass = {hep-th}, + url = {https://arxiv.org/abs/hep-th/9706118} +} + @Book{ mattila1995, author = {Mattila, Pertti}, title = {Geometry of sets and measures in {E}uclidean spaces}, diff --git a/lake-manifest.json b/lake-manifest.json index 64ff25861fc7b..9e03305365c4a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -45,10 +45,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8b6048aa0a4a4b6bcf83597802d8dee734e64b7e", + "rev": "43bcb1964528411e47bfa4edd0c87d1face1fce4", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "v4.15.0-rc1", + "inputRev": "master", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b3a24e0c53d4343e02df0dab6197a030471be686", + "rev": "309c346abe681fddb669a0197580fb6ef5a77123", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6334", diff --git a/lakefile.lean b/lakefile.lean index b522389365bde..70aa7fadaa5b1 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -9,7 +9,7 @@ open Lake DSL require "leanprover-community" / "batteries" @ git "lean-pr-testing-6334" require "leanprover-community" / "Qq" @ git "v4.14.0" -require "leanprover-community" / "aesop" @ git "v4.15.0-rc1" +require "leanprover-community" / "aesop" @ git "master" require "leanprover-community" / "proofwidgets" @ git "v0.0.48" require "leanprover-community" / "importGraph" @ git "v4.15.0-rc1" require "leanprover-community" / "LeanSearchClient" @ git "main" diff --git a/scripts/nolints.json b/scripts/nolints.json index 4f5d00a571a78..58730127f796d 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -352,7 +352,6 @@ ["docBlame", "VectorPrebundle.pretrivializationAt"], ["docBlame", "VectorPrebundle.pretrivializationAtlas"], ["docBlame", "WithIdeal.i"], - ["docBlame", "WithZeroTopology.instHasContinuousInv₀"], ["docBlame", "WriterT.adapt"], ["docBlame", "WriterT.callCC"], ["docBlame", "WriterT.callCC'"],