diff --git a/.github/workflows/PR_summary.yml b/.github/workflows/PR_summary.yml index 1c095f8a66f89..111d7058adaa3 100644 --- a/.github/workflows/PR_summary.yml +++ b/.github/workflows/PR_summary.yml @@ -14,6 +14,27 @@ jobs: with: fetch-depth: 0 + - name: Update the merge-conflict label + run: | + printf 'PR number: "%s"\n' "${{ github.event.pull_request.number }}" + if git merge origin/master --no-ff --no-commit + then + git merge --abort || true + echo "This PR does not have merge conflicts with master." + # we use curl rather than octokit/request-action so that the job won't fail + # (and send an annoying email) if the labels don't exist + curl --request DELETE \ + --url https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.pull_request.number }}/labels/merge-conflict \ + --header 'authorization: Bearer ${{ secrets.GITHUB_TOKEN }}' + else + echo "This PR has merge conflicts with main." + # we use curl rather than octokit/request-action so that the job won't fail + # (and send an annoying email) if the labels don't exist + curl --request POST \ + --url https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.pull_request.number }}/labels/merge-conflict \ + --header 'authorization: Bearer ${{ secrets.GITHUB_TOKEN }}' + fi + - name: Set up Python uses: actions/setup-python@v5 with: diff --git a/.github/workflows/latest_import.yml b/.github/workflows/latest_import.yml index aa7a52cd7f981..6e9b5b0cf2cfd 100644 --- a/.github/workflows/latest_import.yml +++ b/.github/workflows/latest_import.yml @@ -48,8 +48,10 @@ jobs: - name: add minImports linter option run: | - # set `linter.minImports option` to true in `lakefile` - sed -i -- '/^ -- '\`'latest_import.yml'\`' uses this comment/{s=^= ⟨`linter.minImports, true⟩,\n=}' lakefile.lean + # we disable checking for backticks inside single quotes with the next line + # shellcheck disable=SC2016 + # set `linter.minImports option` to true and `Elab.async` to false in `lakefile` + sed -i -- '/^ -- '\`'latest_import.yml'\`' uses this comment/{s=^= ⟨`linter.minImports, true⟩,\n ⟨`Elab.async, false⟩,\n=}' lakefile.lean # import the `minImports` linter in `Mathlib.Init` sed -i -z 's=^=import Mathlib.Tactic.Linter.MinImports\n=' Mathlib/Init.lean diff --git a/.github/workflows/zulip_emoji_closed_pr.yaml b/.github/workflows/zulip_emoji_closed_pr.yaml new file mode 100644 index 0000000000000..9504e7395ca0a --- /dev/null +++ b/.github/workflows/zulip_emoji_closed_pr.yaml @@ -0,0 +1,61 @@ +name: Add "closed-pr" emoji in Zulip +# adds a reaction to Zulip messages that refer to a PR that was closed, but not merged + +# triggers the action when +on: + pull_request: + # the pull request is closed or reopened, to add or remove the emoji + types: [closed, reopened] + +jobs: + add_closed_pr_emoji: + # we set the `TITLE` of the PR as a variable, this shields from possible code injection + env: + TITLE: ${{ github.event.pull_request.title }} + + name: Add closed-pr emoji in Zulip + runs-on: ubuntu-latest + steps: + - name: Debugging information + run: | + # may be superfluous: GitHub may print the values of the environment variables by default + printf '%s' "${TITLE}" | hexdump -cC + printf 'PR title:"%s"\n' "${TITLE}" + printf 'issue number: "%s"\npull request number: "%s"\n' "${{ github.event.issue.number }}" "${{ github.event.pull_request.number }}" + + - name: Set up Python + if: ${{ ! startsWith(github.event.pull_request.title, '[Merged by Bors]') || + github.event_name == 'reopened' }} + uses: actions/setup-python@v5 + with: + python-version: '3.x' + + - name: Install dependencies + if: ${{ ! startsWith(github.event.pull_request.title, '[Merged by Bors]') || + github.event_name == 'reopened' }} + run: | + python -m pip install --upgrade pip + pip install zulip + + - name: checkout zulip_emoji_merge_delegate script + if: ${{ ! startsWith(github.event.pull_request.title, '[Merged by Bors]') || + github.event_name == 'reopened' }} + uses: actions/checkout@v4 + with: + sparse-checkout: | + scripts/zulip_emoji_merge_delegate.py + + - name: Run Zulip Emoji Merge Delegate Script + if: ${{ ! startsWith(github.event.pull_request.title, '[Merged by Bors]') || + github.event_name == 'reopened' }} + env: + ZULIP_API_KEY: ${{ secrets.ZULIP_API_KEY }} + ZULIP_EMAIL: github-mathlib4-bot@leanprover.zulipchat.com + ZULIP_SITE: https://leanprover.zulipchat.com + run: | + # ${{ github.event.action }} is either "closed" or "reopened" + # however, it only matters that it is either "closed" or + # something different from "closed" + # the effect is to add the "closed-pr" emoji to the message, if it is "closed" + # and to remove it otherwise + python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "${{ github.event.action }}" "${{ github.event.pull_request.number }}" diff --git a/Archive.lean b/Archive.lean index fc8da3829ecc9..9b6a77f6d4afd 100644 --- a/Archive.lean +++ b/Archive.lean @@ -46,6 +46,7 @@ import Archive.Imo.Imo2020Q2 import Archive.Imo.Imo2021Q1 import Archive.Imo.Imo2024Q1 import Archive.Imo.Imo2024Q2 +import Archive.Imo.Imo2024Q3 import Archive.Imo.Imo2024Q5 import Archive.Imo.Imo2024Q6 import Archive.MiuLanguage.Basic diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index 94378d663c520..88e40d89763aa 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -39,9 +39,6 @@ the lower bound: `a(b-1)^2/2 ≤ |A|`. Rearranging gives the result. -/ - -open scoped Classical - variable {C J : Type*} (r : C → J → Prop) namespace Imo1998Q2 @@ -86,6 +83,7 @@ theorem JudgePair.agree_iff_same_rating (p : JudgePair J) (c : C) : p.Agree r c ↔ (r c p.judge₁ ↔ r c p.judge₂) := Iff.rfl +open scoped Classical in /-- The set of contestants on which two judges agree. -/ def agreedContestants [Fintype C] (p : JudgePair J) : Finset C := Finset.univ.filter fun c => p.Agree r c @@ -93,14 +91,17 @@ section variable [Fintype J] [Fintype C] +open scoped Classical in /-- All incidences of agreement. -/ def A : Finset (AgreedTriple C J) := Finset.univ.filter @fun (a : AgreedTriple C J) => (a.judgePair.Agree r a.contestant ∧ a.judgePair.Distinct) +open scoped Classical in theorem A_maps_to_offDiag_judgePair (a : AgreedTriple C J) : a ∈ A r → a.judgePair ∈ Finset.offDiag (@Finset.univ J _) := by simp [A, Finset.mem_offDiag] +open scoped Classical in theorem A_fibre_over_contestant (c : C) : (Finset.univ.filter fun p : JudgePair J => p.Agree r c ∧ p.Distinct) = ((A r).filter fun a : AgreedTriple C J => a.contestant = c).image Prod.snd := by @@ -110,6 +111,7 @@ theorem A_fibre_over_contestant (c : C) : · rintro ⟨_, h₂⟩; refine ⟨(c, p), ?_⟩; tauto · intro h; aesop +open scoped Classical in theorem A_fibre_over_contestant_card (c : C) : (Finset.univ.filter fun p : JudgePair J => p.Agree r c ∧ p.Distinct).card = ((A r).filter fun a : AgreedTriple C J => a.contestant = c).card := by @@ -119,6 +121,7 @@ theorem A_fibre_over_contestant_card (c : C) : rintro ⟨a, p⟩ h ⟨a', p'⟩ h' rfl aesop (add simp AgreedTriple.contestant) +open scoped Classical in theorem A_fibre_over_judgePair {p : JudgePair J} (h : p.Distinct) : agreedContestants r p = ((A r).filter fun a : AgreedTriple C J => a.judgePair = p).image AgreedTriple.contestant := by @@ -126,6 +129,7 @@ theorem A_fibre_over_judgePair {p : JudgePair J} (h : p.Distinct) : · rw [Finset.mem_image]; refine ⟨⟨c, p⟩, ?_⟩; aesop · aesop +open scoped Classical in theorem A_fibre_over_judgePair_card {p : JudgePair J} (h : p.Distinct) : (agreedContestants r p).card = ((A r).filter fun a : AgreedTriple C J => a.judgePair = p).card := by @@ -138,6 +142,7 @@ theorem A_card_upper_bound {k : ℕ} (hk : ∀ p : JudgePair J, p.Distinct → (agreedContestants r p).card ≤ k) : (A r).card ≤ k * (Fintype.card J * Fintype.card J - Fintype.card J) := by change _ ≤ k * (Finset.card _ * Finset.card _ - Finset.card _) + classical rw [← Finset.offDiag_card] apply Finset.card_le_mul_card_image_of_maps_to (A_maps_to_offDiag_judgePair r) intro p hp @@ -162,6 +167,7 @@ section variable [Fintype J] +open scoped Classical in theorem judge_pairs_card_lower_bound {z : ℕ} (hJ : Fintype.card J = 2 * z + 1) (c : C) : 2 * z * z + 2 * z + 1 ≤ (Finset.univ.filter fun p : JudgePair J => p.Agree r c).card := by let x := (Finset.univ.filter fun j => r c j).card @@ -173,6 +179,7 @@ theorem judge_pairs_card_lower_bound {z : ℕ} (hJ : Fintype.card J = 2 * z + 1) suffices x + y = 2 * z + 1 by simp [← Int.ofNat_add, this] rw [Finset.filter_card_add_filter_neg_card_eq_card, ← hJ]; rfl +open scoped Classical in theorem distinct_judge_pairs_card_lower_bound {z : ℕ} (hJ : Fintype.card J = 2 * z + 1) (c : C) : 2 * z * z ≤ (Finset.univ.filter fun p : JudgePair J => p.Agree r c ∧ p.Distinct).card := by let s := Finset.univ.filter fun p : JudgePair J => p.Agree r c @@ -192,6 +199,7 @@ theorem distinct_judge_pairs_card_lower_bound {z : ℕ} (hJ : Fintype.card J = 2 theorem A_card_lower_bound [Fintype C] {z : ℕ} (hJ : Fintype.card J = 2 * z + 1) : 2 * z * z * Fintype.card C ≤ (A r).card := by + classical have h : ∀ a, a ∈ A r → Prod.fst a ∈ @Finset.univ C _ := by intros; apply Finset.mem_univ apply Finset.mul_card_image_le_card_of_maps_to h intro c _ diff --git a/Archive/Imo/Imo2024Q3.lean b/Archive/Imo/Imo2024Q3.lean new file mode 100644 index 0000000000000..abfb5b06bf769 --- /dev/null +++ b/Archive/Imo/Imo2024Q3.lean @@ -0,0 +1,961 @@ +/- +Copyright (c) 2024 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import Mathlib.Data.Nat.Nth + +/-! +# IMO 2024 Q3 + +Let $a_1, a_2, a_3, \ldots$ be an infinite sequence of positive integers, and let $N$ be a +positive integer. Suppose that, for each $n > N$, $a_n$ is equal to the number of times +$a_{n-1}$ appears in the list $a_1, a_2, \ldots, a_{n-1}$. + +Prove that at least one of the sequences $a_1, a_3, a_5, \ldots$ and $a_2, a_4, a_6, \ldots$ +is eventually periodic. + +We follow Solution 1 from the +[official solutions](https://www.imo2024.uk/s/IMO2024-solutions-updated.pdf). We show that +the positive integers up to some positive integer $k$ (referred to as small numbers) all appear +infinitely often in the given sequence while all larger positive integers appear only finitely +often and then that the sequence eventually alternates between small numbers and larger integers. +A further detailed analysis of the eventual behavior of the sequence ends up showing that the +sequence of small numbers is eventually periodic with period at most $k$ (in fact exactly $k$). +-/ + + +open scoped Finset + +namespace Imo2024Q3 + +/-- The condition of the problem. Following usual Lean conventions, this is expressed with +indices starting from 0, rather than from 1 as in the informal statment (but `N` remains as the +index of the last term for which `a n` is not defined in terms of previous terms). -/ +def Condition (a : ℕ → ℕ) (N : ℕ) : Prop := + (∀ i, 0 < a i) ∧ ∀ n, N < n → a n = #{i ∈ Finset.range n | a i = a (n - 1)} + +/-- The property of a sequence being eventually periodic. -/ +def EventuallyPeriodic (b : ℕ → ℕ) : Prop := ∃ p M, 0 < p ∧ ∀ m, M ≤ m → b (m + p) = b m + +/-! ### Definitions and lemmas about the sequence that do not actually need the condition of +the problem -/ + +/-- A number greater than any of the initial terms `a 0` through `a N` (otherwise arbitrary). -/ +def M (a : ℕ → ℕ) (N : ℕ) : ℕ := (Finset.range (N + 1)).sup a + 1 + +lemma M_pos (a : ℕ → ℕ) (N : ℕ) : 0 < M a N := + Nat.add_one_pos _ + +lemma one_le_M (a : ℕ → ℕ) (N : ℕ) : 1 ≤ M a N := + Nat.lt_iff_add_one_le.1 (M_pos a N) + +lemma apply_lt_M_of_le_N (a : ℕ → ℕ) {N i : ℕ} (h : i ≤ N) : a i < M a N := + Nat.lt_add_one_iff.2 (Finset.le_sup (Finset.mem_range_succ_iff.2 h)) + +lemma N_lt_of_M_le_apply {a : ℕ → ℕ} {N i : ℕ} (h : M a N ≤ a i) : N < i := by + by_contra! hi + exact Nat.not_succ_le_self _ (h.trans (Finset.le_sup (Finset.mem_range_succ_iff.2 hi))) + +lemma ne_zero_of_M_le_apply {a : ℕ → ℕ} {N i : ℕ} (h : M a N ≤ a i) : i ≠ 0 := + Nat.not_eq_zero_of_lt (N_lt_of_M_le_apply h) + +lemma apply_lt_of_M_le_apply {a : ℕ → ℕ} {N i j : ℕ} (hi : M a N ≤ a i) (hj : j ≤ N) : + a j < a i := + (apply_lt_M_of_le_N a hj).trans_le hi + +lemma apply_ne_of_M_le_apply {a : ℕ → ℕ} {N i j : ℕ} (hi : M a N ≤ a i) (hj : j ≤ N) : + a j ≠ a i := + (apply_lt_of_M_le_apply hi hj).ne + +lemma toFinset_card_pos {a : ℕ → ℕ} {i : ℕ} (hf : {j | a j = a i}.Finite) : 0 < #hf.toFinset := + Finset.card_pos.mpr ((Set.Finite.toFinset_nonempty _).mpr ⟨i, rfl⟩) + +lemma apply_nth_zero (a : ℕ → ℕ) (i : ℕ) : a (Nat.nth (a · = a i) 0) = a i := + Nat.nth_mem (p := (a · = a i)) 0 toFinset_card_pos + +lemma map_add_one_range (p : ℕ → Prop) [DecidablePred p] (n : ℕ) (h0 : ¬ p 0) : + {x ∈ Finset.range n | p (x + 1)}.map ⟨(· + 1), add_left_injective 1⟩ = + {x ∈ Finset.range (n + 1) | p x } := by + ext x + simp only [Finset.mem_map] + constructor + · aesop + · intro hx + use x - 1 + cases x <;> simp_all + +namespace Condition + +/-! ### The basic structure of the sequence, eventually alternating small and large numbers -/ + +variable {a : ℕ → ℕ} {N : ℕ} (hc : Condition a N) +include hc + +protected lemma pos (n : ℕ) : 0 < a n := hc.1 n + +@[simp] lemma apply_ne_zero (n : ℕ) : a n ≠ 0 := + (hc.pos _).ne' + +lemma one_le_apply (n : ℕ) : 1 ≤ a n := + Nat.one_le_iff_ne_zero.2 (hc.apply_ne_zero n) + +lemma apply_eq_card {n : ℕ} (h : N < n) : a n = #{i ∈ Finset.range n | a i = a (n - 1)} := + hc.2 n h + +lemma apply_add_one_eq_card {n : ℕ} (h : N ≤ n) : + a (n + 1) = #{i ∈ Finset.range (n + 1) | a i = a n} := by + rw [hc.apply_eq_card (Nat.lt_add_one_of_le h)] + simp + +@[simp] lemma nth_apply_eq_zero (n : ℕ) : Nat.nth (a · = 0) n = 0 := by + convert Nat.nth_false _ with i + simp only [(hc.pos i).ne'] + +lemma nth_apply_add_one_eq {n : ℕ} (h : N ≤ n) : Nat.nth (a · = a n) (a (n + 1) - 1) = n := by + rw [hc.apply_add_one_eq_card h] + nth_rw 5 [← Nat.nth_count (p := (a · = a n)) rfl] + simp [Finset.range_add_one, Finset.filter_insert, Nat.count_eq_card_filter_range] + +lemma apply_nth_add_one_eq {m n : ℕ} (hfc : ∀ hf : {i | a i = m}.Finite, n < #hf.toFinset) + (hn : N ≤ Nat.nth (a · = m) n) : a (Nat.nth (a · = m) n + 1) = n + 1 := by + rw [hc.apply_eq_card (Nat.lt_add_one_of_le hn), add_tsub_cancel_right, + ← Nat.count_eq_card_filter_range, Nat.nth_mem n hfc, Nat.count_nth_succ hfc] + +lemma apply_nth_add_one_eq_of_infinite {m n : ℕ} (hi : {i | a i = m}.Infinite) + (hn : N ≤ Nat.nth (a · = m) n) : a (Nat.nth (a · = m) n + 1) = n + 1 := + hc.apply_nth_add_one_eq (fun hf ↦ absurd hf hi) hn + +lemma apply_nth_add_one_eq_of_lt {m n : ℕ} (hn : N < Nat.nth (a · = m) n) : + a (Nat.nth (a · = m) n + 1) = n + 1 := by + refine hc.apply_nth_add_one_eq ?_ hn.le + by_contra! hf + have := Nat.nth_eq_zero.2 (.inr hf) + omega + +lemma lt_toFinset_card {j : ℕ} (h : M a N ≤ a (j + 1)) (hf : {i | a i = a j}.Finite) : + M a N - 1 < #hf.toFinset := by + rw [Nat.sub_lt_iff_lt_add (M_pos _ _), Nat.lt_one_add_iff] + exact (hc.apply_eq_card (N_lt_of_M_le_apply h) ▸ h).trans (Finset.card_le_card (by simp)) + +lemma nth_ne_zero_of_M_le_of_lt {i k : ℕ} (hi : M a N ≤ a i) (hk : k < a (i + 1)) : + Nat.nth (a · = a i) k ≠ 0 := + Nat.nth_ne_zero_anti (apply_ne_of_M_le_apply hi (Nat.zero_le _)) (by omega) + (hc.nth_apply_add_one_eq (N_lt_of_M_le_apply hi).le ▸ ne_zero_of_M_le_apply hi) + +lemma apply_add_one_lt_of_apply_eq {i j : ℕ} (hi : N ≤ i) (hij : i < j) (ha : a i = a j) : + a (i + 1) < a (j + 1) := by + rw [hc.apply_add_one_eq_card hi, hc.apply_add_one_eq_card (by omega), ha] + refine Finset.card_lt_card (Finset.ssubset_def.mp ⟨Finset.filter_subset_filter _ + (by simp [hij.le]), Finset.not_subset.mpr ⟨j, ?_⟩⟩) + simp only [Finset.mem_filter, Finset.mem_range, lt_add_iff_pos_right, zero_lt_one, true_and, + and_true] + omega + +lemma apply_add_one_ne_of_apply_eq {i j : ℕ} (hi : N ≤ i) (hj : N ≤ j) (hij : i ≠ j) + (ha : a i = a j) : a (i + 1) ≠ a (j + 1) := + hij.lt_or_lt.elim (fun h ↦ (hc.apply_add_one_lt_of_apply_eq hi h ha).ne) fun h ↦ + (hc.apply_add_one_lt_of_apply_eq hj h ha.symm).ne' + +lemma exists_infinite_setOf_apply_eq : ∃ m, {i | a i = m}.Infinite := by + by_contra hi + have hr : (Set.range a).Infinite := by + contrapose! hi with hr + rw [Set.not_infinite, ← Set.finite_coe_iff] at hr + obtain ⟨n, hn⟩ := Finite.exists_infinite_fiber (Set.rangeFactorization a) + rw [Set.infinite_coe_iff, Set.preimage] at hn + simp only [Set.mem_singleton_iff, Set.rangeFactorization, Subtype.ext_iff] at hn + exact ⟨↑n, hn⟩ + simp only [not_exists, Set.not_infinite] at hi + have hinj : Set.InjOn (fun i ↦ Nat.nth (a · = i) 0 + 1) (Set.range a \ Set.Ico 0 (M a N)) := by + rintro _ ⟨⟨_, rfl⟩, hi⟩ _ ⟨⟨_, rfl⟩, hj⟩ h + simp only [Set.mem_diff, Set.mem_range, Set.mem_Ico, zero_le, true_and, not_lt] at hi hj + simp only [add_left_inj] at h + convert congr(a $h) using 1 <;> simp [apply_nth_zero] + refine Set.not_infinite.2 (hi 1) (Set.infinite_of_injOn_mapsTo hinj (fun i hi ↦ ?_) + (hr.diff (Set.finite_Ico _ _))) + simp only [Set.mem_diff, Set.mem_range, Set.mem_Ico, zero_le, true_and, not_lt] at hi + rcases hi with ⟨⟨_, rfl⟩, hi⟩ + exact hc.apply_nth_add_one_eq toFinset_card_pos + (N_lt_of_M_le_apply (a := a) (by simp only [apply_nth_zero, hi])).le + +lemma nonempty_setOf_infinite_setOf_apply_eq : {m | {i | a i = m}.Infinite}.Nonempty := + hc.exists_infinite_setOf_apply_eq + +lemma injOn_setOf_apply_add_one_eq_of_M_le {n : ℕ} (h : M a N ≤ n) : + Set.InjOn a {i | a (i + 1) = n} := by + intro i hi j hj hij + have hi' := hi ▸ hc.nth_apply_add_one_eq (Nat.lt_add_one_iff.mp (N_lt_of_M_le_apply (hi ▸ h))) + have hj' := hj ▸ hc.nth_apply_add_one_eq (Nat.lt_add_one_iff.mp (N_lt_of_M_le_apply (hj ▸ h))) + rw [← hi', ← hj', hij] + +lemma empty_consecutive_apply_ge_M : {i | M a N ≤ a i ∧ M a N ≤ a (i + 1)} = ∅ := by + rw [Set.eq_empty_iff_forall_not_mem] + intro i + induction i using Nat.strong_induction_on with | h i ih => + -- Let i be the first index where both `a i` and `a (i + 1)` are at least M. + rintro ⟨hi1, hi2⟩ + have hi : ∀ j < i, M a N ≤ a j → a (j + 1) < M a N := by simp_all + -- t is the set of indices before an appearance of the integer (a i). For each j ∈ t, (a j) + -- is the (a i)th appearance of that value, so each such value before index i appears at least + -- M times before that index; since (a i) is the (at least) Mth appearance of that value, there + -- are at least M positive integers appearing M times before (a i), a contradiction because one of + -- those must be at least M. + let t : Finset ℕ := {j ∈ Finset.range i | a (j + 1) = a i} + let t' : Finset ℕ := {j ∈ Finset.range (i + 1) | a j = a i} + have t_map_eq_t' : t.map ⟨(· + 1), add_left_injective 1⟩ = t' := by + refine map_add_one_range (a · = a i) i ?_ + intro H + rw [←H, M] at hi1 + have a0_le : a 0 ≤ (Finset.range (N + 1)).sup a := Finset.le_sup (by simp) + omega + have card_t_eq_card_t' : #t = #t' := by simp [← t_map_eq_t', t] + have htM : ∀ j ∈ t, a j < M a N := by + intro j hj + simp only [t, Finset.mem_filter, Finset.mem_range] at hj + obtain ⟨hj, hji⟩ := hj + by_contra! hjM + exact (lt_self_iff_false _).mp ((hji ▸ hi j hj hjM).trans_le hi1) + have N_le_i : N ≤ i := by + unfold M at hi1 + by_contra! HH + have i_in_range : i ∈ Finset.range (N + 1) := by rw [Finset.mem_range]; omega + have ai_le_sup : a i ≤ (Finset.range (N + 1)).sup a := Finset.le_sup i_in_range + omega + have ht' : a (i + 1) = #t' := hc.apply_add_one_eq_card N_le_i + rw [← card_t_eq_card_t'] at ht' + have ht'inj : Set.InjOn a t := by + refine (hc.injOn_setOf_apply_add_one_eq_of_M_le hi1).mono ?_ + simp_all [t, t'] + have card_image_eq_card_t : #(Finset.image a t) = #t := Finset.card_image_of_injOn ht'inj + have card_image_lt_M : #(Finset.image a t) < M a N := by + refine (Finset.card_le_card (t := Finset.Ico 1 (M a N)) ?_).trans_lt ?_ + · simp only [Finset.subset_iff, Finset.mem_image, Finset.mem_Ico, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂] + exact fun j hj ↦ ⟨hc.pos _, htM j hj⟩ + · simpa using M_pos a N + omega + +lemma card_lt_M_of_M_le {n : ℕ} (h : M a N ≤ n) : + ∃ hf : {i | a i = n}.Finite, #hf.toFinset < M a N := by + have := empty_consecutive_apply_ge_M hc + contrapose! this with hin + use Nat.nth (a · = n) (M a N - 1) + have hin' := fun hf ↦ Nat.sub_one_lt_of_le (M_pos a N) (hin hf) + have ha : M a N ≤ a (Nat.nth (a · = n) (M a N - 1)) := (Nat.nth_mem _ hin').symm ▸ h + refine ⟨ha, ?_⟩ + suffices H : a (Nat.nth (fun x ↦ a x = n) (M a N - 1) + 1) = M a N from Nat.le_of_eq H.symm + convert hc.apply_nth_add_one_eq hin' (N_lt_of_M_le_apply ha).le using 1 + +lemma bddAbove_setOf_infinite_setOf_apply_eq : BddAbove {m | {i | a i = m}.Infinite} := by + refine ⟨M a N, fun x hi ↦ ?_⟩ + by_contra hx + exact hi (hc.card_lt_M_of_M_le (not_le.mp hx).le).1 + +lemma infinite_setOf_apply_eq_anti {j k : ℕ} (hj : 0 < j) (hk : {i | a i = k}.Infinite) + (hjk : j ≤ k) : {i | a i = j}.Infinite := by + have hk' : {i | a (i + 1) = k}.Infinite := by + have hinj : Set.InjOn (· + 1) {i | a (i + 1) = k} := (add_left_injective _).injOn + rw [← Set.infinite_image_iff hinj] + have hk0 : ({i | a i = k} \ {0}).Infinite := hk.diff (Set.finite_singleton _) + convert hk0 using 1 + ext i + simp only [Set.mem_image, Set.mem_setOf_eq, Set.mem_diff, Set.mem_singleton_iff] + refine ⟨?_, ?_⟩ + · rintro ⟨j, rfl, rfl⟩ + simp + · rintro ⟨rfl, h⟩ + exact ⟨i - 1, by simp [(by omega : i - 1 + 1 = i)]⟩ + have hinj : Set.InjOn (fun x ↦ Nat.nth (a · = a x) (j - 1) + 1) + ({i | a (i + 1) = k} \ Set.Ico 0 N) := by + intro x hx y hy h + simp only [Set.mem_diff, Set.mem_setOf_eq, Set.mem_Ico, zero_le, true_and, not_lt] at hx hy + rcases hx with ⟨hxk, hNx⟩ + rcases hy with ⟨hyk, hNy⟩ + simp only [add_left_inj] at h + have hxk' : Nat.nth (a · = a x) (k - 1) = x := by rw [← hxk, hc.nth_apply_add_one_eq hNx] + have hyk' : Nat.nth (a · = a y) (k - 1) = y := by rw [← hyk, hc.nth_apply_add_one_eq hNy] + have hjk' : j - 1 ≤ k - 1 := by omega + apply_fun a at hxk' hyk' + have hyj : a (Nat.nth (a · = a y) (j - 1)) = a y := + Nat.nth_mem_anti (p := (a · = a y)) hjk' hyk' + rw [← h, Nat.nth_mem_anti (p := (a · = a x)) hjk' hxk'] at hyj + by_contra hxy + exact hc.apply_add_one_ne_of_apply_eq hNx hNy hxy hyj (hyk ▸ hxk) + have hk'' : (_ \ Set.Ico 0 (N + 2)).Infinite := + ((Set.infinite_image_iff hinj).mpr (hk'.diff (Set.finite_Ico _ _))).diff (Set.finite_Ico _ _) + refine hk''.mono fun _ hi ↦ ?_ + simp only [Set.mem_image, Set.mem_diff, Set.mem_setOf_eq, Set.mem_Ico, zero_le, true_and, + not_lt] at hi + rcases hi with ⟨⟨x, -, rfl⟩, _⟩ + rw [Set.mem_setOf_eq, hc.apply_nth_add_one_eq_of_lt (by omega), Nat.sub_add_cancel hj] + +/-! ### The definitions of small, medium and big numbers and the eventual alternation -/ + +variable (a) + +/-- The largest number to appear infinitely often. -/ +noncomputable def k : ℕ := sSup {m | {i | a i = m}.Infinite} + +/-- Small numbers are those that are at most `k` (that is, those that appear infinitely often). -/ +def Small (j : ℕ) : Prop := j ≤ k a + +variable {a} + +lemma infinite_setOf_apply_eq_k : {i | a i = k a}.Infinite := + Nat.sSup_mem hc.nonempty_setOf_infinite_setOf_apply_eq hc.bddAbove_setOf_infinite_setOf_apply_eq + +lemma infinite_setOf_apply_eq_iff_small {j : ℕ} (hj : 0 < j) : + {i | a i = j}.Infinite ↔ Small a j := + ⟨fun h ↦ le_csSup hc.bddAbove_setOf_infinite_setOf_apply_eq h, + fun h ↦ hc.infinite_setOf_apply_eq_anti hj hc.infinite_setOf_apply_eq_k h⟩ + +lemma finite_setOf_apply_eq_iff_not_small {j : ℕ} (hj : 0 < j) : + {i | a i = j}.Finite ↔ ¬Small a j := by + simpa only [Set.not_infinite] using (hc.infinite_setOf_apply_eq_iff_small hj).not + +lemma finite_setOf_apply_eq_k_add_one : {i | a i = k a + 1}.Finite := by + rw [hc.finite_setOf_apply_eq_iff_not_small (by omega), Small] + omega + +/-- There are only finitely many `m` that appear more than `k` times. -/ +lemma finite_setOf_k_lt_card : {m | ∀ hf : {i | a i = m}.Finite, k a < #hf.toFinset}.Finite := by + rw [← Set.finite_image_iff] + · refine Set.Finite.of_diff (hc.finite_setOf_apply_eq_k_add_one.subset fun i hi ↦ ?_) + (Set.finite_Iic N) + simp only [Set.mem_diff, Set.mem_image, Set.mem_setOf_eq, Set.mem_Iic, not_le] at hi + rcases hi with ⟨⟨j, hjf, rfl⟩, hNi⟩ + rw [Set.mem_setOf_eq, hc.apply_nth_add_one_eq hjf (by omega)] + · intro i hi j hj hij + simp only [add_left_inj] at hij + apply_fun a at hij + rwa [Nat.nth_mem _ hi, Nat.nth_mem _ hj] at hij + +lemma bddAbove_setOf_k_lt_card : BddAbove {m | ∀ hf : {i | a i = m}.Finite, k a < #hf.toFinset} := + hc.finite_setOf_k_lt_card.bddAbove + +lemma k_pos : 0 < k a := by + by_contra! hn + apply nonpos_iff_eq_zero.mp hn ▸ hc.infinite_setOf_apply_eq_k + convert Set.finite_empty + ext i + simp [(hc.pos i).ne'] + +lemma small_one : Small a 1 := by + by_contra hns + simp only [Small, not_le, Nat.lt_one_iff, hc.k_pos.ne'] at hns + +lemma infinite_setOf_apply_eq_one : {i | a i = 1}.Infinite := + (hc.infinite_setOf_apply_eq_iff_small (by decide)).mpr hc.small_one + +variable (a) + +/-- The largest number to appear more than `k` times. -/ +noncomputable def l : ℕ := sSup {m | ∀ hf : {i | a i = m}.Finite, k a < #hf.toFinset} + +/-- Medium numbers are those that are more than `k` but at most `l` (and include all numbers +appearing finitely often but more than `k` times). -/ +def Medium (j : ℕ) : Prop := k a < j ∧ j ≤ l a + +/-- Big numbers are those greater than `l` (thus, appear at most `k` times). -/ +def Big (j : ℕ) : Prop := l a < j + +variable {a} + +lemma k_le_l : k a ≤ l a := + le_csSup hc.bddAbove_setOf_k_lt_card (fun hf ↦ absurd hf hc.infinite_setOf_apply_eq_k) + +lemma k_lt_of_big {j : ℕ} (h : Big a j) : k a < j := + hc.k_le_l.trans_lt h + +lemma pos_of_big {j : ℕ} (h : Big a j) : 0 < j := + (Nat.zero_le _).trans_lt (hc.k_lt_of_big h) + +lemma not_small_of_big {j : ℕ} (h : Big a j) : ¬Small a j := by simp [Small, hc.k_lt_of_big h] + +lemma exists_card_le_of_big {j : ℕ} (h : Big a j) : + ∃ hf : {i | a i = j}.Finite, #hf.toFinset ≤ k a := by + have hns := hc.not_small_of_big h + rw [← hc.finite_setOf_apply_eq_iff_not_small (hc.pos_of_big h)] at hns + use hns + by_contra! hlt + exact not_mem_of_csSup_lt h hc.bddAbove_setOf_k_lt_card fun _ ↦ hlt + +variable (a N) + +/-- `N'aux` is such that, by position `N'aux`, every medium number has made all its appearances +and every small number has appeared more than max(k, N+1) times. -/ +noncomputable def N'aux : ℕ := + sSup {i | Medium a (a i)} ⊔ sSup ((fun i ↦ Nat.nth (a · = i) (k a ⊔ (N + 1))) '' Set.Iic (k a)) + +/-- `N'` is such that, by position `N'`, every medium number has made all its appearances +and every small number has appeared more than max(k, N+1) times; furthermore, `a N'` is small +(which means that every subsequent big number is preceded by a small number). -/ +noncomputable def N' : ℕ := by + classical + exact N'aux a N + (if Small a (a (N'aux a N + 1)) then 1 else 2) + +variable {a N} + +lemma not_medium_of_N'aux_lt {j : ℕ} (h : N'aux a N < j) : ¬Medium a (a j) := by + let s : Set ℕ := ⋃ i ∈ Set.Ioc (k a) (l a), {j | a j = i} + have hf : s.Finite := by + refine (Set.finite_Ioc _ _).biUnion ?_ + rintro i ⟨hk, -⟩ + rwa [hc.finite_setOf_apply_eq_iff_not_small (by omega), Small, not_le] + exact fun hm ↦ not_mem_of_csSup_lt (le_sup_left.trans_lt h) + (hf.subset fun i hi ↦ (by simpa [s] using hi)).bddAbove hm + +lemma small_or_big_of_N'aux_lt {j : ℕ} (h : N'aux a N < j) : Small a (a j) ∨ Big a (a j) := by + have _ := hc.not_medium_of_N'aux_lt h + rw [Small, Medium, Big] at * + omega + +lemma small_or_big_of_N'_le {j : ℕ} (h : N' a N ≤ j) : Small a (a j) ∨ Big a (a j) := by + refine hc.small_or_big_of_N'aux_lt ?_ + rw [N'] at h + split_ifs at h <;> omega + +omit hc + +lemma nth_sup_k_N_add_one_le_N'aux_of_small {j : ℕ} (h : Small a j) : + Nat.nth (a · = j) (k a ⊔ (N + 1)) ≤ N'aux a N := by + by_contra! hn + exact not_mem_of_csSup_lt (le_sup_right.trans_lt hn) ((Set.finite_Iic _).image _).bddAbove + ⟨j, h, rfl⟩ + +include hc + +lemma nth_sup_k_le_N'aux_of_small {j : ℕ} (h : Small a j) : + Nat.nth (a · = j) (k a) ≤ N'aux a N := + match j with + | 0 => by simp only [hc.nth_apply_eq_zero, zero_le] + | j + 1 => ((Nat.nth_le_nth ((hc.infinite_setOf_apply_eq_iff_small (Nat.zero_lt_succ j)).mpr h)).2 + le_sup_left).trans (nth_sup_k_N_add_one_le_N'aux_of_small h) + +lemma nth_sup_N_add_one_le_N'aux_of_small {j : ℕ} (h : Small a j) : + Nat.nth (a · = j) (N + 1) ≤ N'aux a N := + match j with + | 0 => by simp only [hc.nth_apply_eq_zero, zero_le] + | j + 1 => ((Nat.nth_le_nth ((hc.infinite_setOf_apply_eq_iff_small (Nat.zero_lt_succ j)).mpr h)).2 + le_sup_right).trans (nth_sup_k_N_add_one_le_N'aux_of_small h) + +lemma N_lt_N'aux : N < N'aux a N := + Nat.add_one_le_iff.mp ((Nat.le_nth fun hf ↦ absurd hf hc.infinite_setOf_apply_eq_one).trans + (hc.nth_sup_N_add_one_le_N'aux_of_small hc.small_one)) + +/-- `N` is less than `N'`. -/ +lemma N_lt_N' : N < N' a N := hc.N_lt_N'aux.trans_le (Nat.le_add_right _ _) + +lemma lt_card_filter_eq_of_small_nth_lt {i j t : ℕ} (hj0 : 0 < j) (h : Small a j) + (ht : Nat.nth (a · = j) t < i) : t < #{m ∈ Finset.range i | a m = j} := by + rw [← hc.infinite_setOf_apply_eq_iff_small hj0] at h + rw [← Nat.count_eq_card_filter_range] + exact (Nat.nth_lt_nth h).mp (ht.trans_le (Nat.le_nth_count h _)) + +lemma k_lt_card_filter_eq_of_small_of_N'aux_le {i j : ℕ} (hj0 : 0 < j) (h : Small a j) + (hN'aux : N'aux a N < i) : k a < #{m ∈ Finset.range i | a m = j} := + hc.lt_card_filter_eq_of_small_nth_lt hj0 h ((hc.nth_sup_k_le_N'aux_of_small h).trans_lt hN'aux) + +lemma N_add_one_lt_card_filter_eq_of_small_of_N'aux_le {i j : ℕ} (hj0 : 0 < j) (h : Small a j) + (hN'aux : N'aux a N < i) : N + 1 < #{m ∈ Finset.range i | a m = j} := + hc.lt_card_filter_eq_of_small_nth_lt hj0 h + ((hc.nth_sup_N_add_one_le_N'aux_of_small h).trans_lt hN'aux) + +lemma N_add_one_lt_card_filter_eq_of_small_of_N'_le {i j : ℕ} (hj0 : 0 < j) (h : Small a j) + (hN' : N' a N < i) : N + 1 < #{m ∈ Finset.range i | a m = j} := by + refine hc.N_add_one_lt_card_filter_eq_of_small_of_N'aux_le hj0 h ?_ + rw [N'] at hN' + split_ifs at hN' <;> omega + +lemma apply_add_one_big_of_apply_small_of_N'aux_le {i : ℕ} (h : Small a (a i)) + (hN'aux : N'aux a N ≤ i) : Big a (a (i + 1)) := by + have hN'' : N'aux a N < i + 1 := by omega + suffices ¬Small a (a (i + 1)) by simpa [this] using hc.small_or_big_of_N'aux_lt hN'' + rw [hc.apply_add_one_eq_card (hc.N_lt_N'aux.le.trans hN'aux), Small, not_le] + exact hc.k_lt_card_filter_eq_of_small_of_N'aux_le (hc.pos _) h hN'' + +lemma apply_add_one_big_of_apply_small_of_N'_le {i : ℕ} (h : Small a (a i)) (hN' : N' a N ≤ i) : + Big a (a (i + 1)) := + hc.apply_add_one_big_of_apply_small_of_N'aux_le h ((Nat.le_add_right _ _).trans hN') + +lemma apply_add_one_small_of_apply_big_of_N'aux_le {i : ℕ} (h : Big a (a i)) + (hN'aux : N'aux a N ≤ i) : Small a (a (i + 1)) := by + obtain ⟨hf, hfc⟩ := hc.exists_card_le_of_big h + rw [hc.apply_add_one_eq_card (hc.N_lt_N'aux.le.trans hN'aux)] + exact (Finset.card_le_card (by simp)).trans hfc + +lemma apply_add_one_small_of_apply_big_of_N'_le {i : ℕ} (h : Big a (a i)) (hN' : N' a N ≤ i) : + Small a (a (i + 1)) := + hc.apply_add_one_small_of_apply_big_of_N'aux_le h ((Nat.le_add_right _ _).trans hN') + +lemma apply_add_two_small_of_apply_small_of_N'_le {i : ℕ} (h : Small a (a i)) (hN' : N' a N ≤ i) : + Small a (a (i + 2)) := + hc.apply_add_one_small_of_apply_big_of_N'_le (hc.apply_add_one_big_of_apply_small_of_N'_le h hN') + (by omega) + +/-- `a (N' a N)` is a small number. -/ +lemma small_apply_N' : Small a (a (N' a N)) := by + rw [N'] + split_ifs with hi + · exact hi + · have hb : Big a (a (N'aux a N + 1)) := by + simpa [hi] using hc.small_or_big_of_N'aux_lt (Nat.lt_add_one (N'aux a N)) + exact hc.apply_add_one_small_of_apply_big_of_N'aux_le hb (by omega) + +lemma small_apply_N'_add_iff_even {n : ℕ} : Small a (a (N' a N + n)) ↔ Even n := by + induction n with + | zero => simpa using hc.small_apply_N' + | succ n ih => + by_cases he : Even n <;> rw [← add_assoc] <;> simp only [he, iff_true, iff_false] at ih + · have hne : ¬ Even (n + 1) := by simp [Nat.not_even_iff_odd, he] + simp only [hne, iff_false] + exact hc.not_small_of_big (hc.apply_add_one_big_of_apply_small_of_N'_le ih (by omega)) + · have hb : Big a (a (N' a N + n)) := by + simpa [ih] using hc.small_or_big_of_N'_le (j := N' a N + n) (by omega) + simp [hc.apply_add_one_small_of_apply_big_of_N'_le hb (by omega), Nat.not_even_iff_odd.mp he] + +lemma small_apply_add_two_mul_iff_small {n : ℕ} (m : ℕ) (hN' : N' a N ≤ n) : + Small a (a (n + 2 * m)) ↔ Small a (a n) := by + rw [show n = N' a N + (n - N' a N) by omega, add_assoc, hc.small_apply_N'_add_iff_even, + hc.small_apply_N'_add_iff_even] + simp [Nat.even_add] + +lemma apply_sub_one_small_of_apply_big_of_N'_le {i : ℕ} (h : Big a (a i)) (hN' : N' a N ≤ i) : + Small a (a (i - 1)) := by + have h0i : 1 ≤ i := by + have := hc.N_lt_N' + omega + have h' : N' a N ≤ i - 1 := by + by_contra hi + have hi' : i = N' a N := by omega + exact hc.not_small_of_big (hi' ▸ h) hc.small_apply_N' + exact (hc.small_or_big_of_N'_le h').elim id fun hb ↦ + False.elim (hc.not_small_of_big h (Nat.sub_add_cancel h0i ▸ + (hc.apply_add_one_small_of_apply_big_of_N'_le hb h'))) + +lemma apply_sub_one_big_of_apply_small_of_N'_lt {i : ℕ} (h : Small a (a i)) (hN' : N' a N < i) : + Big a (a (i - 1)) := by + have h0i : 1 ≤ i := by omega + have h' : N' a N ≤ i - 1 := by omega + exact (hc.small_or_big_of_N'_le h').elim (fun hs ↦ False.elim (hc.not_small_of_big + (Nat.sub_add_cancel h0i ▸ hc.apply_add_one_big_of_apply_small_of_N'_le hs h') h)) id + +lemma apply_sub_two_small_of_apply_small_of_N'_lt {i : ℕ} (h : Small a (a i)) (hN' : N' a N < i) : + Small a (a (i - 2)) := by + convert hc.apply_sub_one_small_of_apply_big_of_N'_le + (hc.apply_sub_one_big_of_apply_small_of_N'_lt h hN') (by omega) using 1 + +lemma N_add_one_lt_apply_of_apply_big_of_N'_le {i : ℕ} (h : Big a (a i)) (hN' : N' a N ≤ i) : + N + 1 < a i := by + refine hc.apply_eq_card (hc.N_lt_N'.trans_le hN') ▸ + hc.N_add_one_lt_card_filter_eq_of_small_of_N'_le (hc.pos _) + (hc.apply_sub_one_small_of_apply_big_of_N'_le h hN') ?_ + by_contra + exact hc.not_small_of_big ((by omega : i = N' a N) ▸ h) hc.small_apply_N' + +lemma setOf_apply_eq_of_apply_big_of_N'_le {i : ℕ} (h : Big a (a i)) (hN' : N' a N ≤ i) : + {j | a j = a i} = {j | N < j ∧ Small a (a (j - 1)) ∧ + a i = #{t ∈ Finset.range j | a t = a (j - 1)}} := by + have hs : {j | N < j ∧ Small a (a (j - 1)) ∧ a i = #{t ∈ Finset.range j | a t = a (j - 1)}} ⊆ + {j | a j = a i} := by + rintro _ ⟨hNj, -, hj⟩ + exact hj ▸ hc.apply_eq_card hNj + rcases hc.exists_card_le_of_big h with ⟨hf, hck⟩ + have hf' : {j | N < j ∧ Small a (a (j - 1)) ∧ + a i = #{t ∈ Finset.range j | a t = a (j - 1)}}.Finite := hf.subset hs + suffices hf.toFinset = hf'.toFinset by simpa using this + rw [← Set.Finite.toFinset_subset_toFinset (hs := hf') (ht := hf)] at hs + refine (Finset.eq_of_subset_of_card_le hs (hck.trans ?_)).symm + have hs : #((Finset.Icc 1 (k a)).image (fun t ↦ Nat.nth (a · = t) (a i - 1) + 1)) = k a := by + convert Finset.card_image_of_injOn fun t ht u hu htu ↦ ?_ + · simp only [Nat.card_Icc, add_tsub_cancel_right] + · simp only [add_left_inj] at htu + simp only [Finset.coe_Icc, Set.mem_Icc] at ht hu + rw [← Small, ← hc.infinite_setOf_apply_eq_iff_small (by omega)] at ht hu + apply_fun a at htu + rwa [Nat.nth_mem_of_infinite ht.2, Nat.nth_mem_of_infinite hu.2] at htu + refine hs ▸ Finset.card_le_card (Finset.subset_iff.2 fun j hj ↦ ?_) + simp only [Set.Finite.mem_toFinset, Set.mem_setOf_eq] + simp only [Finset.mem_image, Finset.mem_Icc] at hj + rcases hj with ⟨t, ⟨ht1, htk⟩, rfl⟩ + have hN1 : N < a i - 1 := by + have := hc.N_add_one_lt_apply_of_apply_big_of_N'_le h hN' + omega + simp only [add_tsub_cancel_right] + rw [← Small] at htk + have htki := htk + rw [← hc.infinite_setOf_apply_eq_iff_small (by omega)] at htki + rw [Nat.nth_mem_of_infinite htki] + simp only [htk, true_and] + refine ⟨Nat.lt_add_one_iff.mpr ((Nat.le_nth (fun hf ↦ absurd hf htki)).trans + ((Nat.nth_le_nth htki).2 hN1.le)), ?_⟩ + rw [← Nat.count_eq_card_filter_range, Nat.count_nth_succ_of_infinite htki] + omega + +lemma N_lt_of_apply_eq_of_apply_big_of_N'_le {i j : ℕ} (hj : a j = a i) (h : Big a (a i)) + (hN' : N' a N ≤ i) : N < j := + have hj' : j ∈ {t | a t = a i} := by simpa using hj + (hc.setOf_apply_eq_of_apply_big_of_N'_le h hN' ▸ hj').1 + +lemma small_apply_sub_one_of_apply_eq_of_apply_big_of_N'_le {i j : ℕ} (hj : a j = a i) + (h : Big a (a i)) (hN' : N' a N ≤ i) : Small a (a (j - 1)) := + have hj' : j ∈ {t | a t = a i} := by simpa using hj + (hc.setOf_apply_eq_of_apply_big_of_N'_le h hN' ▸ hj').2.1 + +/-! ### The main lemmas leading to the required result -/ + +/-- Lemma 1 from the informal solution. -/ +lemma apply_add_one_eq_card_small_le_card_eq {i : ℕ} (hi : N' a N < i) (hib : Big a (a i)) : + a (i + 1) = #{m ∈ Finset.range (k a + 1) | a i ≤ #{j ∈ Finset.range i | a j = m}} := by + rw [hc.apply_add_one_eq_card (hc.N_lt_N'.trans hi).le] + convert Finset.card_image_of_injOn (f := fun j ↦ Nat.nth (a · = j) (a i - 1) + 1) ?_ using 1 + · congr + ext j + simp only [Finset.mem_filter, Finset.mem_range, Finset.mem_image] + refine ⟨fun ⟨hj, hji⟩ ↦ ⟨a (j - 1), hji ▸ ?_⟩, fun ⟨t, ⟨hts, htr⟩, ht⟩ ↦ ?_⟩ + · have hjN : N < j := hc.N_lt_of_apply_eq_of_apply_big_of_N'_le hji hib hi.le + refine ⟨⟨Nat.lt_add_one_iff.mpr (hc.small_apply_sub_one_of_apply_eq_of_apply_big_of_N'_le + hji hib hi.le), ?_⟩, ?_⟩ + · rw [hc.apply_eq_card hjN] + have : j ≤ i := by omega + gcongr + · have hj1 : j = j - 1 + 1 := by omega + nth_rw 2 [hj1] + rw [hc.nth_apply_add_one_eq (by omega), hj1.symm] + · subst ht + rw [Nat.lt_add_one_iff, ← Small] at hts + have ht0 : 0 < t := by + by_contra! h0 + simp [nonpos_iff_eq_zero.mp h0, hc.apply_ne_zero] at htr + rw [← hc.infinite_setOf_apply_eq_iff_small ht0] at hts + rw [← Nat.count_eq_card_filter_range] at htr + constructor + · rwa [add_lt_add_iff_right, ← Nat.lt_nth_iff_count_lt hts, + Nat.sub_lt_iff_lt_add (hc.one_le_apply _), Nat.lt_one_add_iff] + · rw [hc.apply_nth_add_one_eq_of_infinite hts] + · exact Nat.sub_add_cancel (hc.one_le_apply _) + · refine (Nat.le_nth fun hf ↦ absurd hf hts).trans ((Nat.nth_le_nth hts).2 ?_) + have := hc.N_add_one_lt_apply_of_apply_big_of_N'_le hib hi.le + omega + · intro t ht u hu htu + simp only [Finset.coe_filter, Finset.mem_range, Set.mem_setOf_eq, Nat.lt_add_one_iff] at ht hu + rw [← Small] at ht hu + have ht0 : 0 < t := by + by_contra! h0 + simp only [nonpos_iff_eq_zero] at h0 + simp [h0, hc.apply_ne_zero] at ht + have hu0 : 0 < u := by + by_contra! h0 + simp only [nonpos_iff_eq_zero] at h0 + simp [h0, hc.apply_ne_zero] at hu + rw [← hc.infinite_setOf_apply_eq_iff_small ht0] at ht + rw [← hc.infinite_setOf_apply_eq_iff_small hu0] at hu + simp only [add_left_inj] at htu + apply_fun a at htu + rwa [Nat.nth_mem_of_infinite ht.1, Nat.nth_mem_of_infinite hu.1] at htu + +/-- Similar to Lemma 1 from the informal solution, but with a `Small` hypothesis instead of `Big` +and considering a range one larger (the form needed for Lemma 2). -/ +lemma apply_eq_card_small_le_card_eq_of_small {i : ℕ} (hi : N' a N + 1 < i) + (his : Small a (a i)) : + a i = #{m ∈ Finset.range (k a + 1) | a (i - 1) ≤ #{j ∈ Finset.range i | a j = m}} := by + have hib : Big a (a (i - 1)) := hc.apply_sub_one_big_of_apply_small_of_N'_lt his (by omega) + nth_rw 1 [show i = i - 1 + 1 by omega] + rw [hc.apply_add_one_eq_card_small_le_card_eq (by omega) hib] + congr 1 + ext j + simp only [Finset.mem_filter, Finset.mem_range, and_congr_right_iff] + intro hj + convert Iff.rfl using 2 + congr 1 + ext t + simp only [Finset.mem_filter, Finset.mem_range, and_congr_right_iff] + refine ⟨fun ⟨hti, rfl⟩ ↦ ⟨?_, rfl⟩, fun ⟨_, rfl⟩ ↦ ⟨by omega, rfl⟩⟩ + by_contra hti1 + have htieq : t = i - 1 := by omega + subst htieq + exact hc.not_small_of_big hib (Nat.lt_add_one_iff.mp hj) + +/-- Lemma 2 from the informal solution. -/ +lemma exists_apply_sub_two_eq_of_apply_eq {i j : ℕ} (hi : N' a N + 2 < i) (hijlt : i < j) + (his : Small a (a i)) (hijeq : a i = a j) + (hij1 : ∀ t, Small a t → #{x ∈ Finset.Ico i j | a x = t} ≤ 1) : + ∃ t, t ∈ Finset.Ico i j ∧ a (i - 2) = a t := by + let I : Finset ℕ := {t ∈ Finset.range (k a + 1) | a (i - 1) ≤ #{u ∈ Finset.range i | a u = t}} + let J : Finset ℕ := {t ∈ Finset.range (k a + 1) | a (j - 1) ≤ #{u ∈ Finset.range j | a u = t}} + have hIc : a i = #I := hc.apply_eq_card_small_le_card_eq_of_small (by omega) his + have hJc : a j = #J := hc.apply_eq_card_small_le_card_eq_of_small (by omega) (hijeq ▸ his) + have hIJc : #I = #J := hIc ▸ hJc ▸ hijeq + have := hc.N_lt_N' + have hiju : Finset.range i ∪ Finset.Ico i j = Finset.range j := by + rw [Finset.range_eq_Ico, Finset.Ico_union_Ico' (by omega) (by omega)] + simp [hijlt.le] + have hi2s : a (i - 2) < k a + 1 := + Nat.lt_add_one_iff.mpr (hc.apply_sub_two_small_of_apply_small_of_N'_lt his (by omega)) + have hiI : a (i - 2) ∈ I := by + simp only [I, Finset.mem_filter, Finset.mem_range, hi2s, true_and] + rw [hc.apply_eq_card (by omega), show i - 1 - 1 = i - 2 by omega] + exact Finset.card_le_card (Finset.filter_subset_filter _ (by simp)) + have hj2s : a (j - 2) < k a + 1 := + Nat.lt_add_one_iff.mpr (hc.apply_sub_two_small_of_apply_small_of_N'_lt (hijeq ▸ his) (by omega)) + have hjJ : a (j - 2) ∈ J := by + simp only [J, Finset.mem_filter, Finset.mem_range, hj2s, true_and] + rw [hc.apply_eq_card (by omega), show j - 1 - 1 = j - 2 by omega] + exact Finset.card_le_card (Finset.filter_subset_filter _ (by simp)) + have hjI : a (j - 2) ∈ I := by + by_contra hjI + have hjIf := hjI + simp only [Finset.mem_filter, Finset.mem_range, hj2s, true_and, not_le, I, + Nat.lt_iff_add_one_le] at hjIf + have hjI' : a (j - 1) ≤ a (i - 1) := by + calc a (j - 1) ≤ #{u ∈ Finset.range j | a u = a (j - 2)} := + (Finset.mem_filter.mp hjJ).2 + _ = #{u ∈ Finset.range i ∪ Finset.Ico i j | a u = a (j - 2)} := by rw [hiju] + _ ≤ #{u ∈ Finset.range i | a u = a (j - 2)} + #{u ∈ Finset.Ico i j | a u = a (j - 2)} := by + rw [Finset.filter_union] + exact Finset.card_union_le _ _ + _ ≤ #{u ∈ Finset.range i | a u = a (j - 2)} + 1 := by + gcongr + exact hij1 _ (by rwa [Nat.lt_add_one_iff, ← Small] at hj2s) + _ ≤ a (i - 1) := hjIf + refine hjI (Finset.eq_of_subset_of_card_le (fun x hxI ↦ ?_) hIJc.symm.le ▸ hjJ) + simp only [Finset.mem_filter, I, J] at * + refine ⟨hxI.1, ?_⟩ + calc a (j - 1) ≤ a (i - 1) := hjI' + _ ≤ #{u ∈ Finset.range i | a u = x} := hxI.2 + _ ≤ #{u ∈ Finset.range j | a u = x} := + Finset.card_le_card (Finset.filter_subset_filter _ (by simp [hijlt.le])) + have hi1j1 : a (i - 1) + 1 ≤ a (j - 1) := by + calc a (i - 1) + 1 ≤ #{u ∈ Finset.range i | a u = a (j - 2)} + 1 := by + gcongr + simp only [Finset.mem_filter, I] at hjI + exact hjI.2 + _ ≤ #{u ∈ Finset.range i | a u = a (j - 2)} + #{u ∈ Finset.Ico i j | a u = a (j - 2)} := by + gcongr + simp only [Finset.one_le_card] + refine ⟨j - 2, ?_⟩ + simp only [Finset.mem_filter, Finset.mem_Ico, and_true] + refine ⟨?_, by omega⟩ + by_contra hj + have hj' : j = i + 1 := by omega + subst hj' + exact hc.not_small_of_big (hc.apply_add_one_big_of_apply_small_of_N'_le his (by omega)) + (hijeq ▸ his) + _ = #({u ∈ Finset.range i | a u = a (j - 2)} ∪ {u ∈ Finset.Ico i j | a u = a (j - 2)}) := by + refine (Finset.card_union_of_disjoint ?_).symm + simp only [Finset.disjoint_iff_ne, Finset.mem_filter, Finset.mem_range, Finset.mem_Ico, + and_imp] + rintro t hti - u hiu - - + omega + _ = #{u ∈ Finset.range j | a u = a (j - 2)} := by + rw [← Finset.filter_union, hiju] + _ = a (j - 1) := by + rw [hc.apply_eq_card (show N < j - 1 by omega)] + congr 1 + ext t + simp only [Finset.mem_filter, Finset.mem_range] + refine ⟨fun ⟨htj, htj'⟩ ↦ ⟨?_, by convert htj' using 1⟩, + fun ⟨htj, htj'⟩ ↦ ⟨by omega, by convert htj' using 1⟩⟩ + by_contra htj'' + have ht1 : t = j - 1 := by omega + subst ht1 + exact hc.not_small_of_big (htj' ▸ hc.apply_sub_one_big_of_apply_small_of_N'_lt + (hijeq ▸ his) (by omega)) (hc.apply_sub_two_small_of_apply_small_of_N'_lt + (hijeq ▸ his) (by omega)) + have hIJ : I = J := by + refine (Finset.eq_of_subset_of_card_le (Finset.subset_iff.mp fun x hxJ ↦ ?_) hIJc.le).symm + simp only [Finset.mem_filter, Finset.mem_range, I, J, Nat.lt_add_one_iff] at * + refine ⟨hxJ.1, (add_le_add_iff_right 1).mp ?_⟩ + calc a (i - 1) + 1 ≤ a (j - 1) := hi1j1 + _ ≤ #{u ∈ Finset.range j | a u = x} := hxJ.2 + _ = #({u ∈ Finset.range i | a u = x} ∪ {u ∈ Finset.Ico i j | a u = x}) := by + rw [← Finset.filter_union, hiju] + _ ≤ #{u ∈ Finset.range i | a u = x} + #{u ∈ Finset.Ico i j | a u = x} := + Finset.card_union_le _ _ + _ ≤ #{u ∈ Finset.range i | a u = x} + 1 := by + gcongr + exact hij1 _ hxJ.1 + simp only [hIJ, J, Finset.mem_filter] at hiI + have hiI' := hi1j1.trans hiI.2 + rw [hc.apply_eq_card (by omega), show i - 1 - 1 = i - 2 by omega, Nat.add_one_le_iff, + ← not_le] at hiI' + rcases Finset.not_subset.mp (mt Finset.card_le_card hiI') with ⟨t, htj, hti⟩ + simp only [Finset.mem_filter, Finset.mem_range] at htj hti + simp only [htj.2, and_true, not_lt, tsub_le_iff_right] at hti + refine ⟨t, Finset.mem_Ico.mpr ⟨?_, htj.1⟩, htj.2.symm⟩ + by_contra + have hti' : t = i - 1 := by omega + subst hti' + exact hc.not_small_of_big (hc.apply_sub_one_big_of_apply_small_of_N'_lt his (by omega)) (htj.2 ▸ + (hc.apply_sub_two_small_of_apply_small_of_N'_lt his (by omega))) + +variable (a) + +/-- The indices, minus `n`, of small numbers appearing for the second or subsequent time at or +after `a n`. -/ +def pSet (n : ℕ) : Set ℕ := {t | ∃ i ∈ Finset.Ico n (n + t), Small a (a i) ∧ a (n + t) = a i} + +/-- The index, minus `n`, of the second appearance of the first small number to appear twice at +or after `a n`. This is only used for small `a n` with `N' a N + 2 < n`. -/ +noncomputable def p (n : ℕ) : ℕ := sInf (pSet a n) + +variable {a} + +lemma nonempty_pSet (n : ℕ) : (pSet a n).Nonempty := by + rcases hc.infinite_setOf_apply_eq_one.exists_gt n with ⟨i, hi1, hni⟩ + rcases hc.infinite_setOf_apply_eq_one.exists_gt i with ⟨j, hj1, hij⟩ + refine ⟨j - n, ?_⟩ + simp only [pSet, Finset.mem_Ico, Set.mem_setOf_eq] + exact ⟨i, ⟨hni.le, by omega⟩, hi1 ▸ ⟨hc.small_one, hj1 ▸ (by congr; omega)⟩⟩ + +lemma exists_mem_Ico_small_and_apply_add_p_eq (n : ℕ) : + ∃ i ∈ Finset.Ico n (n + p a n), Small a (a i) ∧ a (n + p a n) = a i := + csInf_mem (hc.nonempty_pSet _) + +lemma p_pos (n : ℕ) : 0 < p a n := by + by_contra! h + have hn := hc.exists_mem_Ico_small_and_apply_add_p_eq n + simp [h] at hn + +lemma card_filter_apply_eq_Ico_add_p_le_one (n : ℕ) {j : ℕ} (hjs : Small a j) : + #{i ∈ Finset.Ico n (n + p a n) | a i = j} ≤ 1 := by + have h : IsLeast (pSet a n) (p a n) := isLeast_csInf (hc.nonempty_pSet n) + simp only [IsLeast, pSet, Set.mem_setOf_eq, mem_lowerBounds, forall_exists_index, and_imp, + Finset.mem_Ico] at h + rw [Finset.card_le_one_iff] + intro x y hx hy + simp only [Finset.mem_filter, Finset.mem_Ico] at hx hy + rcases lt_trichotomy x y with hxy | rfl | hxy + · replace h := h.2 (y - n) x hx.1.1 (by omega) (hx.2 ▸ hjs) + rw [show n + (y - n) = y by omega, hx.2, hy.2] at h + omega + · rfl + · replace h := h.2 (x - n) y hy.1.1 (by omega) (hy.2 ▸ hjs) + rw [show n + (x - n) = x by omega, hx.2, hy.2] at h + omega + +lemma apply_add_p_eq {n : ℕ} (hn : N' a N + 2 < n) (hs : Small a (a n)) : a (n + p a n) = a n := by + rcases hc.exists_mem_Ico_small_and_apply_add_p_eq n with ⟨i, hiIco, his, hin⟩ + suffices i = n by rw [hin, this] + simp only [Finset.mem_Ico] at hiIco + by_contra hin' + have hf (t : ℕ) (hts : Small a t) : #{x ∈ Finset.Ico i (n + p a n) | a x = t} ≤ 1 := + calc #{x ∈ Finset.Ico i (n + p a n) | a x = t} ≤ #{x ∈ Finset.Ico n (n + p a n) | a x = t} := + Finset.card_le_card (Finset.filter_subset_filter _ (Finset.Ico_subset_Ico hiIco.1 le_rfl)) + _ ≤ 1 := hc.card_filter_apply_eq_Ico_add_p_le_one _ hts + obtain ⟨t, hti, hi2t⟩ := hc.exists_apply_sub_two_eq_of_apply_eq (j := n + p a n) (by omega) + (by omega) his hin.symm hf + have h1 := hc.card_filter_apply_eq_Ico_add_p_le_one n + (hi2t ▸ hc.apply_sub_two_small_of_apply_small_of_N'_lt his (by omega)) + revert h1 + simp only [imp_false, not_le, Finset.one_lt_card_iff, Finset.mem_filter, Finset.mem_Ico, ne_eq, + exists_and_left] + simp only [Finset.mem_Ico] at hti + refine ⟨i - 2, ⟨⟨?_, by omega⟩, hi2t⟩, t, by omega⟩ + by_contra hi2 + have hi1 : n = i - 1 := by omega + subst hi1 + exact hc.not_small_of_big (hc.apply_sub_one_big_of_apply_small_of_N'_lt his (by omega)) hs + +lemma even_p {n : ℕ} (hn : N' a N + 2 < n) (hs : Small a (a n)) : Even (p a n) := by + have hna : n = N' a N + (n - (N' a N)) := by omega + have hs' := hc.apply_add_p_eq hn hs ▸ hs + rw [hna, hc.small_apply_N'_add_iff_even] at hs + nth_rw 1 [hna] at hs' + rw [add_assoc, hc.small_apply_N'_add_iff_even] at hs' + simpa [Nat.even_add, hs] using hs' + +lemma p_le_two_mul_k {n : ℕ} (hn : N' a N + 2 < n) (hs : Small a (a n)) : p a n ≤ 2 * k a := by + by_contra hlt + obtain ⟨x, hx, y, hy, hxyne, hxy⟩ : ∃ x ∈ Finset.range (k a + 1), ∃ y ∈ Finset.range (k a + 1), + x ≠ y ∧ a (n + 2 * x) = a (n + 2 * y) := by + convert Finset.exists_ne_map_eq_of_card_lt_of_maps_to (t := Finset.Icc 1 (k a)) ?_ ?_ + · simp + · rintro i - + simp [Finset.mem_Icc, Nat.lt_add_one_iff] + rw [← Small, hc.small_apply_add_two_mul_iff_small i (by omega)] + simp [hs, hc.one_le_apply] + have hs' : Small a (a (n + 2 * x)) := by rwa [hc.small_apply_add_two_mul_iff_small x (by omega)] + have hj := hc.card_filter_apply_eq_Ico_add_p_le_one n hs' + revert hj + simp only [imp_false, not_le, Finset.one_lt_card_iff, Finset.mem_filter, Finset.mem_Ico, ne_eq, + exists_and_left] + simp only [Finset.mem_range] at hx hy + exact ⟨n + 2 * x, by omega, n + 2 * y, by omega⟩ + +lemma p_apply_sub_two_le_p_apply {n : ℕ} (hn : N' a N + 4 < n) (hs : Small a (a n)) : + p a (n - 2) ≤ p a n := by + obtain ⟨t, hti, _⟩ := hc.exists_apply_sub_two_eq_of_apply_eq (j := n + p a n) (by omega) + ((lt_add_iff_pos_right n).mpr (hc.p_pos n)) hs + (hc.apply_add_p_eq (by omega) hs).symm (fun _ ↦ hc.card_filter_apply_eq_Ico_add_p_le_one _) + by_contra + have hn2 := (hc.apply_sub_two_small_of_apply_small_of_N'_lt hs (by omega)) + have : p a n ≤ p a (n - 2) - 2 := by + obtain ⟨_, _⟩ := hc.even_p (by omega) hs + obtain ⟨_, _⟩ := hc.even_p (by omega) hn2 + omega + have h := hc.card_filter_apply_eq_Ico_add_p_le_one (n - 2) hn2 + revert h + simp only [imp_false, not_le, Finset.one_lt_card_iff, Finset.mem_filter, Finset.mem_Ico, + ne_eq, exists_and_left] + simp only [Finset.mem_Ico] at hti + exact ⟨n - 2, ⟨⟨le_rfl, by omega⟩, rfl⟩, t, by omega⟩ + +lemma p_apply_le_p_apply_add_two {n : ℕ} (hn : N' a N + 2 < n) (hs : Small a (a n)) : + p a n ≤ p a (n + 2) := + hc.p_apply_sub_two_le_p_apply (n := n + 2) (by omega) + (hc.apply_add_two_small_of_apply_small_of_N'_le hs (by omega)) + +variable (a N) + +lemma exists_p_eq : ∃ b c, ∀ n, b < n → p a (N' a N + 2 * n) = c := by + let c : ℕ := sSup (Set.range (fun i ↦ p a (N' a N + 2 * (2 + i)))) + have hk : 2 * k a ∈ upperBounds (Set.range (fun i ↦ p a (N' a N + 2 * (2 + i)))) := by + simp only [mem_upperBounds, Set.mem_range, forall_exists_index, forall_apply_eq_imp_iff] + exact fun i ↦ hc.p_le_two_mul_k (by omega) (hc.small_apply_N'_add_iff_even.mpr (by simp)) + have hlec : ∀ j ∈ Set.range (fun i ↦ p a (N' a N + 2 * (2 + i))), j ≤ c := + fun _ hj ↦ le_csSup ⟨_, hk⟩ hj + obtain ⟨t, ht⟩ := Set.Nonempty.csSup_mem (Set.range_nonempty _) (BddAbove.finite ⟨2 * k a, hk⟩) + have heqc (u : ℕ) : p a (N' a N + 2 * (2 + t + u)) = c := by + induction u with + | zero => simpa using ht + | succ u ih => + refine le_antisymm ?_ (ih ▸ ?_) + · simp only [Set.mem_range, forall_exists_index, forall_apply_eq_imp_iff] at hlec + exact add_assoc _ _ (u + 1) ▸ hlec (t + (u + 1)) + · have hs : Small a (a (N' a N + 2 * (2 + t + u))) := by + rw [hc.small_apply_N'_add_iff_even] + simp + convert hc.p_apply_le_p_apply_add_two (by omega) hs using 1 + refine ⟨1 + t, c, fun n hn ↦ ?_⟩ + rw [show n = 2 + t + (n - (2 + t)) by omega] + exact heqc _ + +lemma exists_a_apply_add_eq : ∃ b c, 0 < c ∧ ∀ n, b < n → + a (N' a N + 2 * n + 2 * c) = a (N' a N + 2 * n) := by + obtain ⟨b, c', hbc'⟩ := hc.exists_p_eq a N + have hs (n : ℕ) : Small a (a (N' a N + 2 * n)) := hc.small_apply_N'_add_iff_even.mpr (by simp) + refine ⟨b + 2, c' / 2, ?_, fun n hbn ↦ hbc' n (by omega) ▸ ?_⟩ + · have := hbc' (b + 2) (by omega) + have := hc.p_pos (N' a N + 2 * (b + 2)) + rcases hc.even_p (by omega) (hs (b + 2)) with ⟨_, _⟩ + omega + · convert hc.apply_add_p_eq (by omega) (hs n) using 3 + rcases hc.even_p (by omega) (hs n) with ⟨_, ht⟩ + simp [ht, ← two_mul] + +variable {a N} + +end Condition + +theorem result {a : ℕ → ℕ} {N : ℕ} (h : Condition a N) : + EventuallyPeriodic (fun i ↦ a (2 * i)) ∨ EventuallyPeriodic (fun i ↦ a (2 * i + 1)) := by + obtain ⟨b, c, hc, hbc⟩ := h.exists_a_apply_add_eq a N + obtain ⟨t, _⟩ | ⟨t, _⟩ := Nat.even_or_odd (Condition.N' a N) + · refine .inl ⟨c, Condition.N' a N / 2 + b + 1, hc, fun m hm ↦ ?_⟩ + convert hbc (m - t) (by omega) using 1 <;> dsimp only <;> congr <;> omega + · refine .inr ⟨c, Condition.N' a N / 2 + b + 1, hc, fun m hm ↦ ?_⟩ + convert hbc (m - t) (by omega) using 1 <;> dsimp only <;> congr 1 <;> omega + +end Imo2024Q3 diff --git a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean index b4e7f16d0b902..c5d0604033257 100644 --- a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean +++ b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean @@ -38,8 +38,6 @@ be phrased in terms of counting walks. -/ -open scoped Classical - namespace Theorems100 noncomputable section @@ -54,6 +52,7 @@ section FriendshipDef variable (G : SimpleGraph V) +open scoped Classical in /-- This property of a graph is the hypothesis of the friendship theorem: every pair of nonadjacent vertices has exactly one common friend, a vertex to which both are adjacent. @@ -74,6 +73,7 @@ namespace Friendship variable (R) +open scoped Classical in include hG in /-- One characterization of a friendship graph is that there is exactly one walk of length 2 between distinct vertices. These walks are counted in off-diagonal entries of the square of @@ -86,6 +86,7 @@ theorem adjMatrix_sq_of_ne {v w : V} (hvw : v ≠ w) : Fintype.card_ofFinset (s := filter (fun x ↦ x ∈ G.neighborSet v ∩ G.neighborSet w) univ), Set.mem_inter_iff, mem_neighborSet] +open scoped Classical in include hG in /-- This calculation amounts to counting the number of length 3 walks between nonadjacent vertices. We use it to show that nonadjacent vertices have equal degrees. -/ @@ -101,6 +102,7 @@ theorem adjMatrix_pow_three_of_not_adj {v w : V} (non_adj : ¬G.Adj v w) : variable {R} +open scoped Classical in include hG in /-- As `v` and `w` not being adjacent implies `degree G v = ((G.adjMatrix R) ^ 3) v w` and `degree G w = ((G.adjMatrix R) ^ 3) v w`, @@ -115,6 +117,7 @@ theorem degree_eq_of_not_adj {v w : V} (hvw : ¬G.Adj v w) : degree G v = degree simp only [pow_succ _ 2, sq, ← transpose_mul, transpose_apply] simp only [mul_assoc] +open scoped Classical in include hG in /-- Let `A` be the adjacency matrix of a graph `G`. If `G` is a friendship graph, then all of the off-diagonal entries of `A^2` are 1. @@ -126,6 +129,7 @@ theorem adjMatrix_sq_of_regular (hd : G.IsRegularOfDegree d) : · rw [h, sq, adjMatrix_mul_self_apply_self, hd]; simp · rw [adjMatrix_sq_of_ne R hG h, of_apply, if_neg h] +open scoped Classical in include hG in theorem adjMatrix_sq_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) (hd : G.IsRegularOfDegree d) : G.adjMatrix (ZMod p) ^ 2 = of fun _ _ => 1 := by @@ -135,6 +139,7 @@ section Nonempty variable [Nonempty V] +open scoped Classical in include hG in /-- If `G` is a friendship graph without a politician (a vertex adjacent to all others), then it is regular. We have shown that nonadjacent vertices of a friendship graph have the same degree, @@ -168,6 +173,7 @@ theorem isRegularOf_not_existsPolitician (hG' : ¬ExistsPolitician G) : rw [key ((mem_commonNeighbors G).mpr ⟨hvx, G.symm hxw⟩), key ((mem_commonNeighbors G).mpr ⟨hvy, G.symm hcontra⟩)] +open scoped Classical in include hG in /-- Let `A` be the adjacency matrix of a `d`-regular friendship graph, and let `v` be a vector all of whose components are `1`. Then `v` is an eigenvector of `A ^ 2`, and we can compute @@ -187,6 +193,7 @@ theorem card_of_regular (hd : G.IsRegularOfDegree d) : d + (Fintype.card V - 1) card_neighborSet_eq_degree, hd v, Function.const_def, adjMatrix_mulVec_apply _ _ (mulVec _ _), mul_one, sum_const, Set.toFinset_card, Algebra.id.smul_eq_mul, Nat.cast_id] +open scoped Classical in include hG in /-- The size of a `d`-regular friendship graph is `1 mod (d-1)`, and thus `1 mod p` for a factor `p ∣ d-1`. -/ @@ -200,17 +207,20 @@ theorem card_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) (hd : G.IsRegu end Nonempty +open scoped Classical in theorem adjMatrix_sq_mul_const_one_of_regular (hd : G.IsRegularOfDegree d) : G.adjMatrix R * of (fun _ _ => 1) = of (fun _ _ => (d : R)) := by ext x simp only [← hd x, degree, adjMatrix_mul_apply, sum_const, Nat.smul_one_eq_cast, of_apply] +open scoped Classical in theorem adjMatrix_mul_const_one_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) (hd : G.IsRegularOfDegree d) : G.adjMatrix (ZMod p) * of (fun _ _ => 1) = of (fun _ _ => 1) := by rw [adjMatrix_sq_mul_const_one_of_regular hd, dmod] +open scoped Classical in include hG in /-- Modulo a factor of `d-1`, the square and all higher powers of the adjacency matrix of a `d`-regular friendship graph reduce to the matrix whose entries are all 1. -/ @@ -227,6 +237,7 @@ theorem adjMatrix_pow_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) variable [Nonempty V] +open scoped Classical in include hG in /-- This is the main proof. Assuming that `3 ≤ d`, we take `p` to be a prime factor of `d-1`. Then the `p`th power of the adjacency matrix of a `d`-regular friendship graph must have trace 1 @@ -260,6 +271,7 @@ theorem false_of_three_le_degree (hd : G.IsRegularOfDegree d) (h : 3 ≤ d) : Fa Nat.dvd_one, Nat.minFac_eq_one_iff] omega +open scoped Classical in include hG in /-- If `d ≤ 1`, a `d`-regular friendship graph has at most one vertex, which is trivially a politician. -/ @@ -277,6 +289,7 @@ theorem existsPolitician_of_degree_le_one (hd : G.IsRegularOfDegree d) (hd1 : d apply h apply Fintype.card_le_one_iff.mp this +open scoped Classical in include hG in /-- If `d = 2`, a `d`-regular friendship graph has 3 vertices, so it must be complete graph, and all the vertices are politicians. -/ @@ -295,6 +308,7 @@ theorem neighborFinset_eq_of_degree_eq_two (hd : G.IsRegularOfDegree 2) (v : V) · dsimp only [IsRegularOfDegree, degree] at hd rw [hd] +open scoped Classical in include hG in theorem existsPolitician_of_degree_eq_two (hd : G.IsRegularOfDegree 2) : ExistsPolitician G := by have v := Classical.arbitrary V @@ -303,6 +317,7 @@ theorem existsPolitician_of_degree_eq_two (hd : G.IsRegularOfDegree 2) : ExistsP rw [← mem_neighborFinset, neighborFinset_eq_of_degree_eq_two hG hd v, Finset.mem_erase] exact ⟨hvw.symm, Finset.mem_univ _⟩ +open scoped Classical in include hG in theorem existsPolitician_of_degree_le_two (hd : G.IsRegularOfDegree d) (h : d ≤ 2) : ExistsPolitician G := by diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 96f1877c186a5..caf136306ff69 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -64,8 +64,6 @@ variable {α : Type*} open Finset -open scoped Classical - /-- The partial product for the generating function for odd partitions. TODO: As `m` tends to infinity, this converges (in the `X`-adic topology). @@ -92,10 +90,12 @@ open Finset.HasAntidiagonal universe u variable {ι : Type u} +open scoped Classical in /-- A convenience constructor for the power series whose coefficients indicate a subset. -/ def indicatorSeries (α : Type*) [Semiring α] (s : Set ℕ) : PowerSeries α := PowerSeries.mk fun n => if n ∈ s then 1 else 0 +open scoped Classical in theorem coeff_indicator (s : Set ℕ) [Semiring α] (n : ℕ) : coeff α n (indicatorSeries _ s) = if n ∈ s then 1 else 0 := coeff_mk _ _ @@ -106,6 +106,7 @@ theorem coeff_indicator_pos (s : Set ℕ) [Semiring α] (n : ℕ) (h : n ∈ s) theorem coeff_indicator_neg (s : Set ℕ) [Semiring α] (n : ℕ) (h : n ∉ s) : coeff α n (indicatorSeries _ s) = 0 := by rw [coeff_indicator, if_neg h] +open scoped Classical in theorem constantCoeff_indicator (s : Set ℕ) [Semiring α] : constantCoeff α (indicatorSeries _ s) = if 0 ∈ s then 1 else 0 := rfl @@ -162,6 +163,7 @@ theorem num_series' [Field α] (i : ℕ) : def mkOdd : ℕ ↪ ℕ := ⟨fun i => 2 * i + 1, fun x y h => by linarith⟩ +open scoped Classical in -- The main workhorse of the partition theorem proof. theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) (hs : ∀ i ∈ s, 0 < i) (c : ℕ → Set ℕ) (hc : ∀ i, i ∉ s → 0 ∈ c i) : diff --git a/Counterexamples/IrrationalPowerOfIrrational.lean b/Counterexamples/IrrationalPowerOfIrrational.lean index 75e822d68b149..6061a15b832af 100644 --- a/Counterexamples/IrrationalPowerOfIrrational.lean +++ b/Counterexamples/IrrationalPowerOfIrrational.lean @@ -17,7 +17,7 @@ If `c` is irrational, then `c^√2 = 2` is rational, so we are done. -/ -open Classical Real +open Real namespace Counterexample diff --git a/Counterexamples/LinearOrderWithPosMulPosEqZero.lean b/Counterexamples/LinearOrderWithPosMulPosEqZero.lean index 3967fef0982e9..a20b56bb93669 100644 --- a/Counterexamples/LinearOrderWithPosMulPosEqZero.lean +++ b/Counterexamples/LinearOrderWithPosMulPosEqZero.lean @@ -68,13 +68,16 @@ instance commMonoid : CommMonoid Foo where mul_comm := by boom mul_assoc := by boom -instance : LinearOrderedCommMonoidWithZero Foo := - { Foo.linearOrder, Foo.commMonoid with - zero := 0 - zero_mul := by boom - mul_zero := by boom - mul_le_mul_left := by rintro ⟨⟩ ⟨⟩ h ⟨⟩ <;> revert h <;> decide - zero_le_one := by decide } +instance : LinearOrderedCommMonoidWithZero Foo where + __ := linearOrder + __ := commMonoid + zero := 0 + zero_mul := by boom + mul_zero := by boom + mul_le_mul_left := by rintro ⟨⟩ ⟨⟩ h ⟨⟩ <;> revert h <;> decide + zero_le_one := by decide + bot := 0 + bot_le := by boom theorem not_mul_pos : ¬∀ {M : Type} [LinearOrderedCommMonoidWithZero M], ∀ a b : M, 0 < a → 0 < b → 0 < a * b := by diff --git a/Counterexamples/MapFloor.lean b/Counterexamples/MapFloor.lean index 3263f315ca27f..5d7fef98f8954 100644 --- a/Counterexamples/MapFloor.lean +++ b/Counterexamples/MapFloor.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Order.Floor +import Mathlib.Algebra.Order.Round import Mathlib.Algebra.Order.Group.PiLex import Mathlib.Algebra.Order.Hom.Ring import Mathlib.Algebra.Polynomial.Reverse diff --git a/Counterexamples/SorgenfreyLine.lean b/Counterexamples/SorgenfreyLine.lean index da5f2a78f9496..496df14aaa0b7 100644 --- a/Counterexamples/SorgenfreyLine.lean +++ b/Counterexamples/SorgenfreyLine.lean @@ -42,7 +42,8 @@ namespace Counterexample noncomputable section -/-- The Sorgenfrey line. It is the real line with the topology space structure generated by +/-- The Sorgenfrey line (denoted as `ℝₗ` within the `SorgenfreyLine` namespace). +It is the real line with the topology space structure generated by half-open intervals `Set.Ico a b`. -/ def SorgenfreyLine : Type := ℝ -- Porting note: was deriving ConditionallyCompleteLinearOrder, LinearOrderedField, Archimedean diff --git a/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean b/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean index b13cc110a5128..227ffb987b92f 100644 --- a/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean +++ b/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean @@ -235,7 +235,7 @@ example {α} [Ring α] [Nontrivial α] : ∃ f g : AddMonoidAlgebra α F, f ≠ example {α} [Zero α] : 2 • (Finsupp.single 0 1 : α →₀ F) = (Finsupp.single 0 1 : α →₀ F) ∧ (Finsupp.single 0 1 : α →₀ F) ≠ 0 := - ⟨smul_single _ _ _, by simp [Ne, Finsupp.single_eq_zero, z01.ne]⟩ + ⟨Finsupp.smul_single _ _ _, by simp [Ne, Finsupp.single_eq_zero, z01.ne]⟩ end F diff --git a/Mathlib.lean b/Mathlib.lean index a2e33bd5bd606..4dfe6dc5afdd4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1,3 +1,4 @@ +import Std import Batteries import Mathlib.Algebra.AddConstMap.Basic import Mathlib.Algebra.AddConstMap.Equiv @@ -20,6 +21,7 @@ import Mathlib.Algebra.Algebra.Rat import Mathlib.Algebra.Algebra.RestrictScalars import Mathlib.Algebra.Algebra.Spectrum import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Algebra.Algebra.Subalgebra.Centralizer import Mathlib.Algebra.Algebra.Subalgebra.Directed import Mathlib.Algebra.Algebra.Subalgebra.IsSimpleOrder import Mathlib.Algebra.Algebra.Subalgebra.MulOpposite @@ -35,6 +37,7 @@ import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.Algebra.Unitization import Mathlib.Algebra.Algebra.ZMod import Mathlib.Algebra.AlgebraicCard +import Mathlib.Algebra.Azumaya.Defs import Mathlib.Algebra.BigOperators.Associated import Mathlib.Algebra.BigOperators.Balance import Mathlib.Algebra.BigOperators.Expect @@ -62,6 +65,7 @@ import Mathlib.Algebra.BigOperators.Ring.Nat import Mathlib.Algebra.BigOperators.RingEquiv import Mathlib.Algebra.BigOperators.Sym import Mathlib.Algebra.BigOperators.WithTop +import Mathlib.Algebra.BrauerGroup.Defs import Mathlib.Algebra.Category.AlgebraCat.Basic import Mathlib.Algebra.Category.AlgebraCat.Limits import Mathlib.Algebra.Category.AlgebraCat.Monoidal @@ -88,6 +92,7 @@ import Mathlib.Algebra.Category.Grp.ForgetCorepresentable import Mathlib.Algebra.Category.Grp.Images import Mathlib.Algebra.Category.Grp.Injective import Mathlib.Algebra.Category.Grp.Kernels +import Mathlib.Algebra.Category.Grp.LargeColimits import Mathlib.Algebra.Category.Grp.Limits import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.Algebra.Category.Grp.Subobject @@ -108,6 +113,7 @@ import Mathlib.Algebra.Category.ModuleCat.Differentials.Basic import Mathlib.Algebra.Category.ModuleCat.Differentials.Presheaf import Mathlib.Algebra.Category.ModuleCat.EnoughInjectives import Mathlib.Algebra.Category.ModuleCat.EpiMono +import Mathlib.Algebra.Category.ModuleCat.ExteriorPower import Mathlib.Algebra.Category.ModuleCat.FilteredColimits import Mathlib.Algebra.Category.ModuleCat.Free import Mathlib.Algebra.Category.ModuleCat.Images @@ -227,6 +233,7 @@ import Mathlib.Algebra.EuclideanDomain.Field import Mathlib.Algebra.EuclideanDomain.Int import Mathlib.Algebra.Exact import Mathlib.Algebra.Expr +import Mathlib.Algebra.Field.Action.ConjAct import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.Field.Equiv @@ -307,6 +314,7 @@ import Mathlib.Algebra.Group.Invertible.Defs import Mathlib.Algebra.Group.MinimalAxioms import Mathlib.Algebra.Group.Nat.Defs import Mathlib.Algebra.Group.Nat.Even +import Mathlib.Algebra.Group.Nat.Hom import Mathlib.Algebra.Group.Nat.TypeTags import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.NatPowAssoc @@ -348,6 +356,7 @@ import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Algebra.Group.Submonoid.DistribMulAction import Mathlib.Algebra.Group.Submonoid.Finsupp import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.Group.Submonoid.MulAction import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.Group.Submonoid.Pointwise @@ -373,6 +382,7 @@ import Mathlib.Algebra.Group.WithOne.Basic import Mathlib.Algebra.Group.WithOne.Defs import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Algebra.GroupWithZero.Action.ConjAct import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.GroupWithZero.Action.Faithful @@ -430,6 +440,7 @@ import Mathlib.Algebra.Homology.DerivedCategory.SingleTriangle import Mathlib.Algebra.Homology.DifferentialObject import Mathlib.Algebra.Homology.Embedding.Basic import Mathlib.Algebra.Homology.Embedding.Boundary +import Mathlib.Algebra.Homology.Embedding.CochainComplex import Mathlib.Algebra.Homology.Embedding.Extend import Mathlib.Algebra.Homology.Embedding.ExtendHomology import Mathlib.Algebra.Homology.Embedding.HomEquiv @@ -488,6 +499,7 @@ import Mathlib.Algebra.Homology.ShortComplex.ModuleCat import Mathlib.Algebra.Homology.ShortComplex.Preadditive import Mathlib.Algebra.Homology.ShortComplex.PreservesHomology import Mathlib.Algebra.Homology.ShortComplex.QuasiIso +import Mathlib.Algebra.Homology.ShortComplex.Retract import Mathlib.Algebra.Homology.ShortComplex.RightHomology import Mathlib.Algebra.Homology.ShortComplex.ShortExact import Mathlib.Algebra.Homology.ShortComplex.SnakeLemma @@ -660,6 +672,7 @@ import Mathlib.Algebra.Order.Antidiag.Pi import Mathlib.Algebra.Order.Antidiag.Prod import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.Archimedean.Hom +import Mathlib.Algebra.Order.Archimedean.IndicatorCard import Mathlib.Algebra.Order.Archimedean.Submonoid import Mathlib.Algebra.Order.BigOperators.Expect import Mathlib.Algebra.Order.BigOperators.Group.Finset @@ -676,6 +689,7 @@ import Mathlib.Algebra.Order.CauSeq.BigOperators import Mathlib.Algebra.Order.CauSeq.Completion import Mathlib.Algebra.Order.Chebyshev import Mathlib.Algebra.Order.CompleteField +import Mathlib.Algebra.Order.Disjointed import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Canonical import Mathlib.Algebra.Order.Field.Defs @@ -800,6 +814,7 @@ import Mathlib.Algebra.Order.Ring.Synonym import Mathlib.Algebra.Order.Ring.Unbundled.Basic import Mathlib.Algebra.Order.Ring.Unbundled.Rat import Mathlib.Algebra.Order.Ring.WithTop +import Mathlib.Algebra.Order.Round import Mathlib.Algebra.Order.Star.Basic import Mathlib.Algebra.Order.Star.Conjneg import Mathlib.Algebra.Order.Star.Prod @@ -891,6 +906,7 @@ import Mathlib.Algebra.Regular.Basic import Mathlib.Algebra.Regular.Pow import Mathlib.Algebra.Regular.SMul import Mathlib.Algebra.Ring.Action.Basic +import Mathlib.Algebra.Ring.Action.ConjAct import Mathlib.Algebra.Ring.Action.Field import Mathlib.Algebra.Ring.Action.Group import Mathlib.Algebra.Ring.Action.Invariant @@ -1020,6 +1036,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.Finite import Mathlib.AlgebraicGeometry.Morphisms.FinitePresentation import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.Morphisms.Flat +import Mathlib.AlgebraicGeometry.Morphisms.FormallyUnramified import Mathlib.AlgebraicGeometry.Morphisms.Immersion import Mathlib.AlgebraicGeometry.Morphisms.Integral import Mathlib.AlgebraicGeometry.Morphisms.IsIso @@ -1102,12 +1119,15 @@ import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialObject.Split import Mathlib.AlgebraicTopology.SimplicialSet.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.Boundary import Mathlib.AlgebraicTopology.SimplicialSet.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.HomotopyCat +import Mathlib.AlgebraicTopology.SimplicialSet.Horn import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path +import Mathlib.AlgebraicTopology.SimplicialSet.StdSimplex import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal import Mathlib.AlgebraicTopology.SingularSet import Mathlib.AlgebraicTopology.TopologicalSimplex @@ -1557,7 +1577,8 @@ import Mathlib.Analysis.SpecialFunctions.Complex.LogBounds import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.ExpLog import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.PosPart -import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow.Basic +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow.Isometric import Mathlib.Analysis.SpecialFunctions.Exp import Mathlib.Analysis.SpecialFunctions.ExpDeriv import Mathlib.Analysis.SpecialFunctions.Exponential @@ -1580,6 +1601,7 @@ import Mathlib.Analysis.SpecialFunctions.Log.ENNRealLogExp import Mathlib.Analysis.SpecialFunctions.Log.ERealExp import Mathlib.Analysis.SpecialFunctions.Log.Monotone import Mathlib.Analysis.SpecialFunctions.Log.NegMulLog +import Mathlib.Analysis.SpecialFunctions.Log.Summable import Mathlib.Analysis.SpecialFunctions.NonIntegrable import Mathlib.Analysis.SpecialFunctions.OrdinaryHypergeometric import Mathlib.Analysis.SpecialFunctions.PolarCoord @@ -1633,6 +1655,7 @@ import Mathlib.CategoryTheory.Abelian.LeftDerived import Mathlib.CategoryTheory.Abelian.NonPreadditive import Mathlib.CategoryTheory.Abelian.Opposite import Mathlib.CategoryTheory.Abelian.Projective +import Mathlib.CategoryTheory.Abelian.ProjectiveDimension import Mathlib.CategoryTheory.Abelian.ProjectiveResolution import Mathlib.CategoryTheory.Abelian.Pseudoelements import Mathlib.CategoryTheory.Abelian.Refinements @@ -1665,7 +1688,8 @@ import Mathlib.CategoryTheory.Adjunction.Triple import Mathlib.CategoryTheory.Adjunction.Unique import Mathlib.CategoryTheory.Adjunction.Whiskering import Mathlib.CategoryTheory.Balanced -import Mathlib.CategoryTheory.Bicategory.Adjunction +import Mathlib.CategoryTheory.Bicategory.Adjunction.Basic +import Mathlib.CategoryTheory.Bicategory.Adjunction.Mate import Mathlib.CategoryTheory.Bicategory.Basic import Mathlib.CategoryTheory.Bicategory.Coherence import Mathlib.CategoryTheory.Bicategory.End @@ -1727,6 +1751,7 @@ import Mathlib.CategoryTheory.CofilteredSystem import Mathlib.CategoryTheory.CommSq import Mathlib.CategoryTheory.Comma.Arrow import Mathlib.CategoryTheory.Comma.Basic +import Mathlib.CategoryTheory.Comma.CardinalArrow import Mathlib.CategoryTheory.Comma.Final import Mathlib.CategoryTheory.Comma.LocallySmall import Mathlib.CategoryTheory.Comma.Over @@ -1768,7 +1793,7 @@ import Mathlib.CategoryTheory.Enriched.Basic import Mathlib.CategoryTheory.Enriched.FunctorCategory import Mathlib.CategoryTheory.Enriched.HomCongr import Mathlib.CategoryTheory.Enriched.Opposite -import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Enriched.Ordinary.Basic import Mathlib.CategoryTheory.EpiMono import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Equivalence @@ -1783,7 +1808,9 @@ import Mathlib.CategoryTheory.FiberedCategory.Fibered import Mathlib.CategoryTheory.FiberedCategory.HomLift import Mathlib.CategoryTheory.Filtered.Basic import Mathlib.CategoryTheory.Filtered.Connected +import Mathlib.CategoryTheory.Filtered.CostructuredArrow import Mathlib.CategoryTheory.Filtered.Final +import Mathlib.CategoryTheory.Filtered.Flat import Mathlib.CategoryTheory.Filtered.Grothendieck import Mathlib.CategoryTheory.Filtered.OfColimitCommutesFiniteLimit import Mathlib.CategoryTheory.Filtered.Small @@ -1821,7 +1848,7 @@ import Mathlib.CategoryTheory.Galois.Prorepresentability import Mathlib.CategoryTheory.Galois.Topology import Mathlib.CategoryTheory.Generator.Abelian import Mathlib.CategoryTheory.Generator.Basic -import Mathlib.CategoryTheory.Generator.Coproduct +import Mathlib.CategoryTheory.Generator.Indization import Mathlib.CategoryTheory.Generator.Preadditive import Mathlib.CategoryTheory.Generator.Presheaf import Mathlib.CategoryTheory.Generator.Sheaf @@ -1906,6 +1933,7 @@ import Mathlib.CategoryTheory.Limits.Indization.Equalizers import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.IndObject import Mathlib.CategoryTheory.Limits.Indization.LocallySmall +import Mathlib.CategoryTheory.Limits.Indization.ParallelPair import Mathlib.CategoryTheory.Limits.Indization.Products import Mathlib.CategoryTheory.Limits.IsConnected import Mathlib.CategoryTheory.Limits.IsLimit @@ -1919,6 +1947,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Basic import Mathlib.CategoryTheory.Limits.Preserves.Filtered import Mathlib.CategoryTheory.Limits.Preserves.Finite import Mathlib.CategoryTheory.Limits.Preserves.FunctorCategory +import Mathlib.CategoryTheory.Limits.Preserves.Grothendieck import Mathlib.CategoryTheory.Limits.Preserves.Limits import Mathlib.CategoryTheory.Limits.Preserves.Opposites import Mathlib.CategoryTheory.Limits.Preserves.Presheaf @@ -1958,6 +1987,9 @@ import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Basic import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Equalizers import Mathlib.CategoryTheory.Limits.Shapes.PiProd +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.Basic +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.HasIterationOfShape +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.WellOrderContinuous import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Assoc import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq @@ -2010,6 +2042,7 @@ import Mathlib.CategoryTheory.Localization.FiniteProducts import Mathlib.CategoryTheory.Localization.HasLocalization import Mathlib.CategoryTheory.Localization.HomEquiv import Mathlib.CategoryTheory.Localization.LocalizerMorphism +import Mathlib.CategoryTheory.Localization.LocallySmall import Mathlib.CategoryTheory.Localization.Opposite import Mathlib.CategoryTheory.Localization.Pi import Mathlib.CategoryTheory.Localization.Predicate @@ -2019,6 +2052,7 @@ import Mathlib.CategoryTheory.Localization.SmallHom import Mathlib.CategoryTheory.Localization.SmallShiftedHom import Mathlib.CategoryTheory.Localization.StructuredArrow import Mathlib.CategoryTheory.Localization.Triangulated +import Mathlib.CategoryTheory.Localization.Trifunctor import Mathlib.CategoryTheory.Monad.Adjunction import Mathlib.CategoryTheory.Monad.Algebra import Mathlib.CategoryTheory.Monad.Basic @@ -2117,6 +2151,8 @@ import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic import Mathlib.CategoryTheory.Preadditive.Yoneda.Injective import Mathlib.CategoryTheory.Preadditive.Yoneda.Limits import Mathlib.CategoryTheory.Preadditive.Yoneda.Projective +import Mathlib.CategoryTheory.Presentable.Basic +import Mathlib.CategoryTheory.Presentable.IsCardinalFiltered import Mathlib.CategoryTheory.Products.Associator import Mathlib.CategoryTheory.Products.Basic import Mathlib.CategoryTheory.Products.Bifunctor @@ -2230,6 +2266,8 @@ import Mathlib.CategoryTheory.Subobject.MonoOver import Mathlib.CategoryTheory.Subobject.Types import Mathlib.CategoryTheory.Subobject.WellPowered import Mathlib.CategoryTheory.Subpresheaf.Basic +import Mathlib.CategoryTheory.Subpresheaf.Image +import Mathlib.CategoryTheory.Subpresheaf.Sieves import Mathlib.CategoryTheory.Subterminal import Mathlib.CategoryTheory.Sums.Associator import Mathlib.CategoryTheory.Sums.Basic @@ -2379,6 +2417,7 @@ import Mathlib.Computability.Encoding import Mathlib.Computability.EpsilonNFA import Mathlib.Computability.Halting import Mathlib.Computability.Language +import Mathlib.Computability.MyhillNerode import Mathlib.Computability.NFA import Mathlib.Computability.Partrec import Mathlib.Computability.PartrecCode @@ -2465,6 +2504,7 @@ import Mathlib.Data.Complex.FiniteDimensional import Mathlib.Data.Complex.Module import Mathlib.Data.Complex.Order import Mathlib.Data.Complex.Orientation +import Mathlib.Data.Complex.Trigonometric import Mathlib.Data.Countable.Basic import Mathlib.Data.Countable.Defs import Mathlib.Data.Countable.Small @@ -2500,6 +2540,7 @@ import Mathlib.Data.Fin.Basic import Mathlib.Data.Fin.Fin2 import Mathlib.Data.Fin.FlagRange import Mathlib.Data.Fin.Parity +import Mathlib.Data.Fin.Rev import Mathlib.Data.Fin.SuccPred import Mathlib.Data.Fin.Tuple.Basic import Mathlib.Data.Fin.Tuple.BubbleSortInduction @@ -2681,6 +2722,7 @@ import Mathlib.Data.List.Iterate import Mathlib.Data.List.Lattice import Mathlib.Data.List.Lemmas import Mathlib.Data.List.Lex +import Mathlib.Data.List.Map2 import Mathlib.Data.List.MinMax import Mathlib.Data.List.Monad import Mathlib.Data.List.NatAntidiagonal @@ -2738,12 +2780,14 @@ import Mathlib.Data.Matrix.Rank import Mathlib.Data.Matrix.Reflection import Mathlib.Data.Matrix.RowCol import Mathlib.Data.Matroid.Basic +import Mathlib.Data.Matroid.Circuit import Mathlib.Data.Matroid.Closure import Mathlib.Data.Matroid.Constructions import Mathlib.Data.Matroid.Dual import Mathlib.Data.Matroid.IndepAxioms import Mathlib.Data.Matroid.Init import Mathlib.Data.Matroid.Map +import Mathlib.Data.Matroid.Rank.Cardinal import Mathlib.Data.Matroid.Restrict import Mathlib.Data.Matroid.Sum import Mathlib.Data.Multiset.Antidiagonal @@ -3010,6 +3054,7 @@ import Mathlib.Data.Sym.Sym2.Order import Mathlib.Data.Tree.Basic import Mathlib.Data.Tree.Get import Mathlib.Data.Tree.RBMap +import Mathlib.Data.Tree.Traversable import Mathlib.Data.TwoPointing import Mathlib.Data.TypeMax import Mathlib.Data.TypeVec @@ -3031,7 +3076,8 @@ import Mathlib.Data.ZMod.Coprime import Mathlib.Data.ZMod.Defs import Mathlib.Data.ZMod.Factorial import Mathlib.Data.ZMod.IntUnitsPower -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientGroup +import Mathlib.Data.ZMod.QuotientRing import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.AlgebraClasses import Mathlib.Deprecated.Aliases @@ -3122,8 +3168,9 @@ import Mathlib.FieldTheory.Minpoly.IsConjRoot import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed import Mathlib.FieldTheory.Minpoly.MinpolyDiv import Mathlib.FieldTheory.MvRatFunc.Rank -import Mathlib.FieldTheory.Normal -import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.Normal.Basic +import Mathlib.FieldTheory.Normal.Closure +import Mathlib.FieldTheory.Normal.Defs import Mathlib.FieldTheory.Perfect import Mathlib.FieldTheory.PerfectClosure import Mathlib.FieldTheory.PolynomialGaloisGroup @@ -3161,6 +3208,7 @@ import Mathlib.Geometry.Euclidean.Sphere.Power import Mathlib.Geometry.Euclidean.Sphere.Ptolemy import Mathlib.Geometry.Euclidean.Sphere.SecondInter import Mathlib.Geometry.Euclidean.Triangle +import Mathlib.Geometry.Group.Growth.LinearLowerBound import Mathlib.Geometry.Group.Growth.QuotientInter import Mathlib.Geometry.Manifold.Algebra.LeftInvariantDerivation import Mathlib.Geometry.Manifold.Algebra.LieGroup @@ -3213,6 +3261,7 @@ import Mathlib.Geometry.Manifold.VectorBundle.MDifferentiable import Mathlib.Geometry.Manifold.VectorBundle.Pullback import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection import Mathlib.Geometry.Manifold.VectorBundle.Tangent +import Mathlib.Geometry.Manifold.VectorField import Mathlib.Geometry.Manifold.WhitneyEmbedding import Mathlib.Geometry.RingedSpace.Basic import Mathlib.Geometry.RingedSpace.LocallyRingedSpace @@ -3583,7 +3632,9 @@ import Mathlib.LinearAlgebra.Quotient.Pi import Mathlib.LinearAlgebra.Ray import Mathlib.LinearAlgebra.Reflection import Mathlib.LinearAlgebra.RootSystem.Base +import Mathlib.LinearAlgebra.RootSystem.BaseChange import Mathlib.LinearAlgebra.RootSystem.Basic +import Mathlib.LinearAlgebra.RootSystem.CartanMatrix import Mathlib.LinearAlgebra.RootSystem.Defs import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear import Mathlib.LinearAlgebra.RootSystem.Finite.Nondegenerate @@ -3696,6 +3747,7 @@ import Mathlib.MeasureTheory.Covering.OneDim import Mathlib.MeasureTheory.Covering.Vitali import Mathlib.MeasureTheory.Covering.VitaliFamily import Mathlib.MeasureTheory.Decomposition.Exhaustion +import Mathlib.MeasureTheory.Decomposition.IntegralRNDeriv import Mathlib.MeasureTheory.Decomposition.Jordan import Mathlib.MeasureTheory.Decomposition.Lebesgue import Mathlib.MeasureTheory.Decomposition.RadonNikodym @@ -3973,6 +4025,7 @@ import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.NumberTheory.LSeries.Positivity import Mathlib.NumberTheory.LSeries.PrimesInAP import Mathlib.NumberTheory.LSeries.RiemannZeta +import Mathlib.NumberTheory.LSeries.SumCoeff import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter import Mathlib.NumberTheory.LegendreSymbol.Basic @@ -4079,6 +4132,7 @@ import Mathlib.Order.BoundedOrder.Monotone import Mathlib.Order.Bounds.Basic import Mathlib.Order.Bounds.Defs import Mathlib.Order.Bounds.Image +import Mathlib.Order.Bounds.Lattice import Mathlib.Order.Bounds.OrderIso import Mathlib.Order.Category.BddDistLat import Mathlib.Order.Category.BddLat @@ -4108,6 +4162,7 @@ import Mathlib.Order.Compare import Mathlib.Order.CompleteBooleanAlgebra import Mathlib.Order.CompleteLattice import Mathlib.Order.CompleteLattice.Finset +import Mathlib.Order.CompleteLattice.SetLike import Mathlib.Order.CompleteLatticeIntervals import Mathlib.Order.CompletePartialOrder import Mathlib.Order.CompleteSublattice @@ -4171,12 +4226,14 @@ import Mathlib.Order.Filter.Ring import Mathlib.Order.Filter.SmallSets import Mathlib.Order.Filter.Subsingleton import Mathlib.Order.Filter.Tendsto -import Mathlib.Order.Filter.Ultrafilter +import Mathlib.Order.Filter.Ultrafilter.Basic +import Mathlib.Order.Filter.Ultrafilter.Defs import Mathlib.Order.Filter.ZeroAndBoundedAtFilter import Mathlib.Order.Fin.Basic import Mathlib.Order.Fin.Tuple import Mathlib.Order.FixedPoints -import Mathlib.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Basic +import Mathlib.Order.GaloisConnection.Defs import Mathlib.Order.GameAdd import Mathlib.Order.Grade import Mathlib.Order.Height @@ -4235,6 +4292,7 @@ import Mathlib.Order.Monotone.Odd import Mathlib.Order.Monotone.Union import Mathlib.Order.Nat import Mathlib.Order.Notation +import Mathlib.Order.Nucleus import Mathlib.Order.OmegaCompletePartialOrder import Mathlib.Order.OrdContinuous import Mathlib.Order.OrderIsoNat @@ -4333,7 +4391,8 @@ import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Martingale.OptionalSampling import Mathlib.Probability.Martingale.OptionalStopping import Mathlib.Probability.Martingale.Upcrossing -import Mathlib.Probability.Moments +import Mathlib.Probability.Moments.Basic +import Mathlib.Probability.Moments.IntegrableExpMul import Mathlib.Probability.Notation import Mathlib.Probability.ProbabilityMassFunction.Basic import Mathlib.Probability.ProbabilityMassFunction.Binomial @@ -4491,6 +4550,7 @@ import Mathlib.RingTheory.Ideal.BigOperators import Mathlib.RingTheory.Ideal.Colon import Mathlib.RingTheory.Ideal.Cotangent import Mathlib.RingTheory.Ideal.Defs +import Mathlib.RingTheory.Ideal.Height import Mathlib.RingTheory.Ideal.IdempotentFG import Mathlib.RingTheory.Ideal.IsPrimary import Mathlib.RingTheory.Ideal.IsPrincipal @@ -4580,6 +4640,7 @@ import Mathlib.RingTheory.Localization.LocalizationLocalization import Mathlib.RingTheory.Localization.Module import Mathlib.RingTheory.Localization.NormTrace import Mathlib.RingTheory.Localization.NumDen +import Mathlib.RingTheory.Localization.Pi import Mathlib.RingTheory.Localization.Submodule import Mathlib.RingTheory.MatrixAlgebra import Mathlib.RingTheory.Multiplicity @@ -4706,12 +4767,15 @@ import Mathlib.RingTheory.SimpleRing.Matrix import Mathlib.RingTheory.Smooth.Basic import Mathlib.RingTheory.Smooth.Kaehler import Mathlib.RingTheory.Smooth.Local +import Mathlib.RingTheory.Smooth.Locus import Mathlib.RingTheory.Smooth.Pi import Mathlib.RingTheory.Smooth.StandardSmooth import Mathlib.RingTheory.Spectrum.Maximal.Basic import Mathlib.RingTheory.Spectrum.Maximal.Defs +import Mathlib.RingTheory.Spectrum.Maximal.Localization import Mathlib.RingTheory.Spectrum.Maximal.Topology import Mathlib.RingTheory.Spectrum.Prime.Basic +import Mathlib.RingTheory.Spectrum.Prime.ConstructibleSet import Mathlib.RingTheory.Spectrum.Prime.Defs import Mathlib.RingTheory.Spectrum.Prime.FreeLocus import Mathlib.RingTheory.Spectrum.Prime.IsOpenComapC @@ -4729,6 +4793,7 @@ import Mathlib.RingTheory.TensorProduct.Free import Mathlib.RingTheory.TensorProduct.MvPolynomial import Mathlib.RingTheory.TensorProduct.Nontrivial import Mathlib.RingTheory.TensorProduct.Pi +import Mathlib.RingTheory.TensorProduct.Quotient import Mathlib.RingTheory.Trace.Basic import Mathlib.RingTheory.Trace.Defs import Mathlib.RingTheory.Trace.Quotient @@ -4804,6 +4869,7 @@ import Mathlib.SetTheory.Cardinal.SchroederBernstein import Mathlib.SetTheory.Cardinal.Subfield import Mathlib.SetTheory.Cardinal.ToNat import Mathlib.SetTheory.Cardinal.UnivLE +import Mathlib.SetTheory.Descriptive.Tree import Mathlib.SetTheory.Game.Basic import Mathlib.SetTheory.Game.Birthday import Mathlib.SetTheory.Game.Domineering @@ -4918,6 +4984,7 @@ import Mathlib.Tactic.ExtractGoal import Mathlib.Tactic.ExtractLets import Mathlib.Tactic.FBinop import Mathlib.Tactic.FailIfNoProgress +import Mathlib.Tactic.FastInstance import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.FinCases import Mathlib.Tactic.Find @@ -5107,6 +5174,7 @@ import Mathlib.Testing.Plausible.Testable import Mathlib.Topology.AlexandrovDiscrete import Mathlib.Topology.Algebra.Affine import Mathlib.Topology.Algebra.Algebra +import Mathlib.Topology.Algebra.Algebra.Equiv import Mathlib.Topology.Algebra.Algebra.Rat import Mathlib.Topology.Algebra.Category.ProfiniteGrp.Basic import Mathlib.Topology.Algebra.Category.ProfiniteGrp.Limits @@ -5142,6 +5210,7 @@ import Mathlib.Topology.Algebra.InfiniteSum.Nonarchimedean import Mathlib.Topology.Algebra.InfiniteSum.Order import Mathlib.Topology.Algebra.InfiniteSum.Real import Mathlib.Topology.Algebra.InfiniteSum.Ring +import Mathlib.Topology.Algebra.IntermediateField import Mathlib.Topology.Algebra.LinearTopology import Mathlib.Topology.Algebra.Localization import Mathlib.Topology.Algebra.Module.Alternating.Basic @@ -5197,6 +5266,7 @@ import Mathlib.Topology.Algebra.SeparationQuotient.Section import Mathlib.Topology.Algebra.Star import Mathlib.Topology.Algebra.StarSubalgebra import Mathlib.Topology.Algebra.Support +import Mathlib.Topology.Algebra.TopologicallyNilpotent import Mathlib.Topology.Algebra.UniformConvergence import Mathlib.Topology.Algebra.UniformField import Mathlib.Topology.Algebra.UniformFilterBasis @@ -5292,6 +5362,7 @@ import Mathlib.Topology.Connected.LocallyConnected import Mathlib.Topology.Connected.PathConnected import Mathlib.Topology.Connected.Separation import Mathlib.Topology.Connected.TotallyDisconnected +import Mathlib.Topology.Constructible import Mathlib.Topology.Constructions import Mathlib.Topology.ContinuousMap.Algebra import Mathlib.Topology.ContinuousMap.Basic diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index dc5674fb8f1ef..1f7087d7e049d 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -32,7 +32,8 @@ assert_not_exists Finset open Function Set -/-- A bundled map `f : G → H` such that `f (x + a) = f x + b` for all `x`. +/-- A bundled map `f : G → H` such that `f (x + a) = f x + b` for all `x`, +denoted as `f: G →+c[a, b] H`. One can think about `f` as a lift to `G` of a map between two `AddCircle`s. -/ structure AddConstMap (G H : Type*) [Add G] [Add H] (a : G) (b : H) where diff --git a/Mathlib/Algebra/AddConstMap/Equiv.lean b/Mathlib/Algebra/AddConstMap/Equiv.lean index 28779a4f482cc..7c5fbbb2766d2 100644 --- a/Mathlib/Algebra/AddConstMap/Equiv.lean +++ b/Mathlib/Algebra/AddConstMap/Equiv.lean @@ -20,7 +20,8 @@ assert_not_exists Finset open Function open scoped AddConstMap -/-- An equivalence between `G` and `H` conjugating `(· + a)` to `(· + b)`. -/ +/-- An equivalence between `G` and `H` conjugating `(· + a)` to `(· + b)`, +denoted as `G ≃+c[a, b] H`. -/ structure AddConstEquiv (G H : Type*) [Add G] [Add H] (a : G) (b : H) extends G ≃ H, G →+c[a, b] H diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index c853185785a74..7977a04a58c14 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -38,6 +38,8 @@ multiplicative group actions). -/ +open scoped Pointwise + /-- An `AddTorsor G P` gives a structure to the nonempty type `P`, acted on by an `AddGroup G` with a transitive and free action given by the `+ᵥ` operation and a corresponding subtraction given by the @@ -162,8 +164,6 @@ theorem vadd_eq_vadd_iff_neg_add_eq_vsub {v₁ v₂ : G} {p₁ p₂ : P} : namespace Set -open Pointwise - theorem singleton_vsub_self (p : P) : ({p} : Set P) -ᵥ {p} = {(0 : G)} := by rw [Set.singleton_vsub_singleton, vsub_self] @@ -231,6 +231,13 @@ theorem vadd_eq_vadd_iff_sub_eq_vsub {v₁ v₂ : G} {p₁ p₂ : P} : theorem vsub_sub_vsub_comm (p₁ p₂ p₃ p₄ : P) : p₁ -ᵥ p₂ - (p₃ -ᵥ p₄) = p₁ -ᵥ p₃ - (p₂ -ᵥ p₄) := by rw [← vsub_vadd_eq_vsub_sub, vsub_vadd_comm, vsub_vadd_eq_vsub_sub] +namespace Set + +@[simp] lemma vadd_set_vsub_vadd_set (v : G) (s t : Set P) : (v +ᵥ s) -ᵥ (v +ᵥ t) = s -ᵥ t := by + ext; simp [mem_vsub, mem_vadd_set] + +end Set + end comm namespace Prod diff --git a/Mathlib/Algebra/Algebra/Defs.lean b/Mathlib/Algebra/Algebra/Defs.lean index 8f9470e1325d7..7b3f54fdd852d 100644 --- a/Mathlib/Algebra/Algebra/Defs.lean +++ b/Mathlib/Algebra/Algebra/Defs.lean @@ -95,7 +95,6 @@ section Prio See the implementation notes in this file for discussion of the details of this definition. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): unsupported @[nolint has_nonempty_instance] class Algebra (R : Type u) (A : Type v) [CommSemiring R] [Semiring A] extends SMul R A where /-- Embedding `R →+* A` given by `Algebra` structure. Use `algebraMap` from the root namespace instead.-/ diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index cd52e4afdabda..9a0fa07bd87e7 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -22,7 +22,8 @@ This file defines bundled isomorphisms of `R`-algebras. universe u v w u₁ v₁ -/-- An equivalence of algebras is an equivalence of rings commuting with the actions of scalars. -/ +/-- An equivalence of algebras (denoted as `A ≃ₐ[R] B`) +is an equivalence of rings commuting with the actions of scalars. -/ structure AlgEquiv (R : Type u) (A : Type v) (B : Type w) [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] extends A ≃ B, A ≃* B, A ≃+ B, A ≃+* B where /-- An equivalence of algebras commutes with the action of scalars. -/ diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index b813755c63742..355643e02267e 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -22,8 +22,7 @@ This file defines bundled homomorphisms of `R`-algebras. universe u v w u₁ v₁ -/-- Defining the homomorphism in the category R-Alg. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet +/-- Defining the homomorphism in the category R-Alg, denoted `A →ₐ[R] B`. -/ structure AlgHom (R : Type u) (A : Type v) (B : Type w) [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] extends RingHom A B where commutes' : ∀ r : R, toFun (algebraMap R A r) = algebraMap R B r @@ -414,8 +413,8 @@ end RingHom namespace Algebra -variable (R : Type u) (A : Type v) -variable [CommSemiring R] [Semiring A] [Algebra R A] +variable (R : Type u) (A : Type v) (B : Type w) +variable [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] /-- `AlgebraMap` as an `AlgHom`. -/ def ofId : R →ₐ[R] A := @@ -434,6 +433,9 @@ instance subsingleton_id : Subsingleton (R →ₐ[R] A) := @[ext high] theorem ext_id (f g : R →ₐ[R] A) : f = g := Subsingleton.elim _ _ +@[simp] +theorem comp_ofId (φ : A →ₐ[R] B) : φ.comp (Algebra.ofId R A) = Algebra.ofId R B := by ext + section MulDistribMulAction instance : MulDistribMulAction (A →ₐ[R] A) Aˣ where diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index d4e25752e01d4..0563a75d8d15b 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -46,8 +46,10 @@ universe u u₁ v w w₁ w₂ w₃ variable {R : Type u} {S : Type u₁} -/-- A morphism respecting addition, multiplication, and scalar multiplication. When these arise from -algebra structures, this is the same as a not-necessarily-unital morphism of algebras. -/ +/-- A morphism respecting addition, multiplication, and scalar multiplication +(denoted as `A →ₛₙₐ[φ] B`, or `A →ₙₐ[R] B` when `φ` is the identity on `R`). +When these arise from algebra structures, this is the same +as a not-necessarily-unital morphism of algebras. -/ structure NonUnitalAlgHom [Monoid R] [Monoid S] (φ : R →* S) (A : Type v) (B : Type w) [NonUnitalNonAssocSemiring A] [DistribMulAction R A] [NonUnitalNonAssocSemiring B] [DistribMulAction S B] extends A →ₑ+[φ] B, A →ₙ* B diff --git a/Mathlib/Algebra/Algebra/Pi.lean b/Mathlib/Algebra/Algebra/Pi.lean index 188a138794019..d04b77f4ffb48 100644 --- a/Mathlib/Algebra/Algebra/Pi.lean +++ b/Mathlib/Algebra/Algebra/Pi.lean @@ -48,6 +48,15 @@ theorem algebraMap_apply {_ : CommSemiring R} [_s : ∀ i, Semiring (f i)] [∀ (a : R) (i : I) : algebraMap R (∀ i, f i) a i = algebraMap R (f i) a := rfl +variable {I} in +instance (g : I → Type*) [∀ i, CommSemiring (f i)] [∀ i, Semiring (g i)] + [∀ i, Algebra (f i) (g i)] : Algebra (∀ i, f i) (∀ i, g i) where + algebraMap := Pi.ringHom fun _ ↦ (algebraMap _ _).comp (Pi.evalRingHom f _) + commutes' _ _ := funext fun _ ↦ Algebra.commutes _ _ + smul_def' _ _ := funext fun _ ↦ Algebra.smul_def _ _ + +example [∀ i, CommSemiring (f i)] : Pi.instAlgebraForall f f = Algebra.id _ := rfl + -- One could also build a `∀ i, R i`-algebra structure on `∀ i, A i`, -- when each `A i` is an `R i`-algebra, although I'm not sure that it's useful. variable {I} (R) @@ -70,6 +79,16 @@ def evalAlgHom {_ : CommSemiring R} [∀ i, Semiring (f i)] [∀ i, Algebra R (f toFun := fun f => f i commutes' := fun _ => rfl } +@[simp] +theorem algHom_evalAlgHom [CommSemiring R] [s : ∀ i, Semiring (f i)] [∀ i, Algebra R (f i)] : + algHom R f (evalAlgHom R f) = AlgHom.id R (Π i, f i) := rfl + +/-- `Pi.algHom` commutes with composition. -/ +theorem algHom_comp [CommSemiring R] [∀ i, Semiring (f i)] [∀ i, Algebra R (f i)] + {A B : Type*} [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] + (g : ∀ i, B →ₐ[R] f i) (h : A →ₐ[R] B) : + (algHom R f g).comp h = algHom R f (fun i ↦ (g i).comp h) := rfl + variable (A B : Type*) [CommSemiring R] [Semiring B] [Algebra R B] /-- `Function.const` as an `AlgHom`. The name matches `Pi.constRingHom`, `Pi.constMonoidHom`, diff --git a/Mathlib/Algebra/Algebra/Prod.lean b/Mathlib/Algebra/Algebra/Prod.lean index 0c726579527a5..bee86dbf00c8b 100644 --- a/Mathlib/Algebra/Algebra/Prod.lean +++ b/Mathlib/Algebra/Algebra/Prod.lean @@ -61,7 +61,15 @@ def fst : A × B →ₐ[R] A := def snd : A × B →ₐ[R] B := { RingHom.snd A B with commutes' := fun _r => rfl } -variable {R A B} +variable {A B} + +@[simp] +theorem fst_apply (a) : fst R A B a = a.1 := rfl + +@[simp] +theorem snd_apply (a) : snd R A B a = a.2 := rfl + +variable {R} /-- The `Pi.prod` of two morphisms is a morphism. -/ @[simps!] @@ -81,8 +89,11 @@ theorem fst_prod (f : A →ₐ[R] B) (g : A →ₐ[R] C) : (fst R B C).comp (pro theorem snd_prod (f : A →ₐ[R] B) (g : A →ₐ[R] C) : (snd R B C).comp (prod f g) = g := by ext; rfl @[simp] -theorem prod_fst_snd : prod (fst R A B) (snd R A B) = 1 := - DFunLike.coe_injective Pi.prod_fst_snd +theorem prod_fst_snd : prod (fst R A B) (snd R A B) = AlgHom.id R _ := rfl + +theorem prod_comp {C' : Type*} [Semiring C'] [Algebra R C'] + (f : A →ₐ[R] B) (g : B →ₐ[R] C) (g' : B →ₐ[R] C') : + (g.prod g').comp f = (g.comp f).prod (g'.comp f) := rfl /-- Taking the product of two maps with the same domain is equivalent to taking the product of their codomains. -/ diff --git a/Mathlib/Algebra/Algebra/Quasispectrum.lean b/Mathlib/Algebra/Algebra/Quasispectrum.lean index 5dce203059f77..fa1469c46f1a2 100644 --- a/Mathlib/Algebra/Algebra/Quasispectrum.lean +++ b/Mathlib/Algebra/Algebra/Quasispectrum.lean @@ -586,11 +586,4 @@ lemma quasispectrumRestricts_iff_spectrumRestricts_inr' theorem quasispectrumRestricts_iff_spectrumRestricts {R S A : Type*} [Semifield R] [Semifield S] [Ring A] [Algebra R S] [Algebra R A] [Algebra S A] {a : A} {f : S → R} : - QuasispectrumRestricts a f ↔ SpectrumRestricts a f := by - rw [quasispectrumRestricts_iff, spectrumRestricts_iff, quasispectrum_eq_spectrum_union_zero] - refine and_congr_left fun h ↦ ?_ - refine ⟨(Set.RightInvOn.mono · Set.subset_union_left), fun h' x hx ↦ ?_⟩ - simp only [Set.union_singleton, Set.mem_insert_iff] at hx - obtain (rfl | hx) := hx - · simpa using h 0 - · exact h' hx + QuasispectrumRestricts a f ↔ SpectrumRestricts a f := by rfl diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 15b381c605224..64279f792eac2 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -766,6 +766,43 @@ theorem iSup_toSubsemiring {ι : Sort*} [Nonempty ι] (S : ι → Subalgebra R A (iSup S).toSubsemiring = ⨆ i, (S i).toSubsemiring := by simp only [iSup, Set.range_nonempty, sSup_toSubsemiring, ← Set.range_comp, Function.comp_def] +lemma mem_iSup_of_mem {ι : Sort*} {S : ι → Subalgebra R A} (i : ι) {x : A} (hx : x ∈ S i) : + x ∈ iSup S := + le_iSup S i hx + +@[elab_as_elim] +lemma iSup_induction {ι : Sort*} (S : ι → Subalgebra R A) {motive : A → Prop} + {x : A} (mem : x ∈ ⨆ i, S i) + (basic : ∀ i, ∀ a ∈ S i, motive a) + (zero : motive 0) (one : motive 1) + (add : ∀ a b, motive a → motive b → motive (a + b)) + (mul : ∀ a b, motive a → motive b → motive (a * b)) + (algebraMap : ∀ r, motive (algebraMap R A r)) : motive x := by + let T : Subalgebra R A := + { carrier := {x | motive x} + mul_mem' {a b} := mul a b + one_mem' := one + add_mem' {a b} := add a b + zero_mem' := zero + algebraMap_mem' := algebraMap } + suffices iSup S ≤ T from this mem + rwa [iSup_le_iff] + +/-- A dependent version of `Subalgebra.iSup_induction`. -/ +@[elab_as_elim] +theorem iSup_induction' {ι : Sort*} (S : ι → Subalgebra R A) {motive : ∀ x, (x ∈ ⨆ i, S i) → Prop} + {x : A} (mem : x ∈ ⨆ i, S i) + (basic : ∀ (i) (x) (hx : x ∈ S i), motive x (mem_iSup_of_mem i hx)) + (zero : motive 0 (zero_mem _)) (one : motive 1 (one_mem _)) + (add : ∀ x y hx hy, motive x hx → motive y hy → motive (x + y) (add_mem ‹_› ‹_›)) + (mul : ∀ x y hx hy, motive x hx → motive y hy → motive (x * y) (mul_mem ‹_› ‹_›)) + (algebraMap : ∀ r, motive (algebraMap R A r) (Subalgebra.algebraMap_mem _ ‹_›)) : + motive x mem := by + refine Exists.elim ?_ fun (hx : x ∈ ⨆ i, S i) (hc : motive x hx) ↦ hc + exact iSup_induction S (motive := fun x' ↦ ∃ h, motive x' h) mem + (fun _ _ h ↦ ⟨_, basic _ _ h⟩) ⟨_, zero⟩ ⟨_, one⟩ (fun _ _ h h' ↦ ⟨_, add _ _ _ _ h.2 h'.2⟩) + (fun _ _ h h' ↦ ⟨_, mul _ _ _ _ h.2 h'.2⟩) fun _ ↦ ⟨_, algebraMap _⟩ + instance : Inhabited (Subalgebra R A) := ⟨⊥⟩ theorem mem_bot {x : A} : x ∈ (⊥ : Subalgebra R A) ↔ x ∈ Set.range (algebraMap R A) := Iff.rfl @@ -787,6 +824,9 @@ theorem _root_.AlgHom.range_eq_top (f : A →ₐ[R] B) : @[deprecated (since := "2024-11-11")] alias range_top_iff_surjective := AlgHom.range_eq_top +@[simp] +theorem range_ofId : (Algebra.ofId R A).range = ⊥ := rfl + @[simp] theorem range_id : (AlgHom.id R A).range = ⊤ := SetLike.coe_injective Set.range_id diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Centralizer.lean b/Mathlib/Algebra/Algebra/Subalgebra/Centralizer.lean new file mode 100644 index 0000000000000..500bc456f934d --- /dev/null +++ b/Mathlib/Algebra/Algebra/Subalgebra/Centralizer.lean @@ -0,0 +1,197 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.RingTheory.TensorProduct.Basic +import Mathlib.RingTheory.Adjoin.Basic +import Mathlib.LinearAlgebra.TensorProduct.Basis +import Mathlib.LinearAlgebra.FreeModule.Basic + +/-! +# Properties of centers and centralizers + +This file contains theorems about the center and centralizer of a subalgebra. + +## Main results + +Let `R` be a commutative ring and `A` and `B` two `R`-algebras. +- `Subalgebra.centralizer_sup`: if `S` and `T` are subalgebras of `A`, then the centralizer of + `S ⊔ T` is the intersection of the centralizer of `S` and the centralizer of `T`. +- `Subalgebra.centralizer_range_includeLeft_eq_center_tensorProduct`: if `B` is free as a module, + then the centralizer of `A ⊗ 1` in `A ⊗ B` is `C(A) ⊗ B` where `C(A)` is the center of `A`. +- `Subalgebra.centralizer_range_includeRight_eq_center_tensorProduct`: if `A` is free as a module, + then the centralizer of `1 ⊗ B` in `A ⊗ B` is `A ⊗ C(B)` where `C(B)` is the center of `B`. +-/ + +namespace Subalgebra + +open Algebra.TensorProduct + +section CommSemiring + +variable {R : Type*} [CommSemiring R] +variable {A : Type*} [Semiring A] [Algebra R A] + +lemma le_centralizer_iff (S T : Subalgebra R A) : S ≤ centralizer R T ↔ T ≤ centralizer R S := + ⟨fun h t ht _ hs ↦ (h hs t ht).symm, fun h s hs _ ht ↦ (h ht s hs).symm⟩ + +lemma centralizer_coe_sup (S T : Subalgebra R A) : + centralizer R ((S ⊔ T : Subalgebra R A) : Set A) = centralizer R S ⊓ centralizer R T := + eq_of_forall_le_iff fun K ↦ by + simp_rw [le_centralizer_iff, sup_le_iff, le_inf_iff, K.le_centralizer_iff] + +lemma centralizer_coe_iSup {ι : Sort*} (S : ι → Subalgebra R A) : + centralizer R ((⨆ i, S i : Subalgebra R A) : Set A) = ⨅ i, centralizer R (S i) := + eq_of_forall_le_iff fun K ↦ by + simp_rw [le_centralizer_iff, iSup_le_iff, le_iInf_iff, K.le_centralizer_iff] + +end CommSemiring + +section Free + +variable (R : Type*) [CommSemiring R] +variable (A : Type*) [Semiring A] [Algebra R A] +variable (B : Type*) [Semiring B] [Algebra R B] + +open Finsupp TensorProduct + +/-- +Let `R` be a commutative ring and `A, B` be `R`-algebras where `B` is free as `R`-module. +For any subset `S ⊆ A`, the centralizer of `S ⊗ 1 ⊆ A ⊗ B` is `C_A(S) ⊗ B` where `C_A(S)` is the +centralizer of `S` in `A`. +-/ +lemma centralizer_coe_image_includeLeft_eq_center_tensorProduct + (S : Set A) [Module.Free R B] : + Subalgebra.centralizer R + (Algebra.TensorProduct.includeLeft (S := R) '' S) = + (Algebra.TensorProduct.map (Subalgebra.centralizer R (S : Set A)).val + (AlgHom.id R B)).range := by + classical + ext w + constructor + · intro hw + rw [mem_centralizer_iff] at hw + let ℬ := Module.Free.chooseBasis R B + obtain ⟨b, rfl⟩ := TensorProduct.eq_repr_basis_right ℬ w + refine Subalgebra.sum_mem _ fun j hj => ⟨⟨b j, ?_⟩ ⊗ₜ[R] ℬ j, by simp⟩ + rw [Subalgebra.mem_centralizer_iff] + intro x hx + suffices x • b = b.mapRange (· * x) (by simp) from Finsupp.ext_iff.1 this j + specialize hw (x ⊗ₜ[R] 1) ⟨x, hx, rfl⟩ + simp only [Finsupp.sum, Finset.mul_sum, Algebra.TensorProduct.tmul_mul_tmul, one_mul, + Finset.sum_mul, mul_one] at hw + refine TensorProduct.sum_tmul_basis_right_injective ℬ ?_ + simp only [Finsupp.coe_lsum] + rw [sum_of_support_subset (s := b.support) (hs := Finsupp.support_smul) (h := by aesop), + sum_of_support_subset (s := b.support) (hs := support_mapRange) (h := by aesop)] + simpa only [Finsupp.coe_smul, Pi.smul_apply, smul_eq_mul, LinearMap.flip_apply, + TensorProduct.mk_apply, Finsupp.mapRange_apply] using hw + + · rintro ⟨w, rfl⟩ + rw [Subalgebra.mem_centralizer_iff] + rintro _ ⟨x, hx, rfl⟩ + induction w using TensorProduct.induction_on with + | zero => simp + | tmul b c => + simp [Subalgebra.mem_centralizer_iff _ |>.1 b.2 x hx] + | add y z hy hz => rw [map_add, mul_add, hy, hz, add_mul] + +/-- +Let `R` be a commutative ring and `A, B` be `R`-algebras where `B` is free as `R`-module. +For any subset `S ⊆ B`, the centralizer of `1 ⊗ S ⊆ A ⊗ B` is `A ⊗ C_B(S)` where `C_B(S)` is the +centralizer of `S` in `B`. +-/ +lemma centralizer_coe_image_includeRight_eq_center_tensorProduct + (S : Set B) [Module.Free R A] : + Subalgebra.centralizer R + (Algebra.TensorProduct.includeRight '' S) = + (Algebra.TensorProduct.map (AlgHom.id R A) + (Subalgebra.centralizer R (S : Set B)).val).range := by + have eq1 := centralizer_coe_image_includeLeft_eq_center_tensorProduct R B A S + apply_fun Subalgebra.comap (Algebra.TensorProduct.comm R A B).toAlgHom at eq1 + convert eq1 + · ext x + simpa [mem_centralizer_iff] using + ⟨fun h b hb ↦ (Algebra.TensorProduct.comm R A B).symm.injective <| by aesop, fun h b hb ↦ + (Algebra.TensorProduct.comm R A B).injective <| by aesop⟩ + · ext x + simp only [AlgHom.mem_range, AlgEquiv.toAlgHom_eq_coe, mem_comap, AlgHom.coe_coe] + constructor + · rintro ⟨x, rfl⟩ + exact ⟨(Algebra.TensorProduct.comm R _ _) x, + by rw [Algebra.TensorProduct.comm_comp_map_apply]⟩ + · rintro ⟨y, hy⟩ + refine ⟨(Algebra.TensorProduct.comm R _ _) y, (Algebra.TensorProduct.comm R A B).injective ?_⟩ + rw [← hy, comm_comp_map_apply, ← comm_symm, AlgEquiv.symm_apply_apply] + +/-- +Let `R` be a commutative ring and `A, B` be `R`-algebras where `B` is free as `R`-module. +For any subalgebra `S` of `A`, the centralizer of `S ⊗ 1 ⊆ A ⊗ B` is `C_A(S) ⊗ B` where `C_A(S)` is +the centralizer of `S` in `A`. +-/ +lemma centralizer_coe_map_includeLeft_eq_center_tensorProduct + (S : Subalgebra R A) [Module.Free R B] : + Subalgebra.centralizer R + (S.map (Algebra.TensorProduct.includeLeft (R := R) (B := B))) = + (Algebra.TensorProduct.map (Subalgebra.centralizer R (S : Set A)).val + (AlgHom.id R B)).range := + centralizer_coe_image_includeLeft_eq_center_tensorProduct R A B S + +/-- +Let `R` be a commutative ring and `A, B` be `R`-algebras where `A` is free as `R`-module. +For any subalgebra `S` of `B`, the centralizer of `1 ⊗ S ⊆ A ⊗ B` is `A ⊗ C_B(S)` where `C_B(S)` is +the centralizer of `S` in `B`. +-/ +lemma centralizer_coe_map_includeRight_eq_center_tensorProduct + (S : Subalgebra R B) [Module.Free R A] : + Subalgebra.centralizer R + (S.map (Algebra.TensorProduct.includeRight (R := R) (A := A))) = + (Algebra.TensorProduct.map (AlgHom.id R A) + (Subalgebra.centralizer R (S : Set B)).val).range := + centralizer_coe_image_includeRight_eq_center_tensorProduct R A B S + +/-- +Let `R` be a commutative ring and `A, B` be `R`-algebras where `B` is free as `R`-module. +Then the centralizer of `A ⊗ 1 ⊆ A ⊗ B` is `C(A) ⊗ B` where `C(A)` is the center of `A`. +-/ +lemma centralizer_coe_range_includeLeft_eq_center_tensorProduct [Module.Free R B] : + Subalgebra.centralizer R + (Algebra.TensorProduct.includeLeft : A →ₐ[R] A ⊗[R] B).range = + (Algebra.TensorProduct.map (Subalgebra.center R A).val (AlgHom.id R B)).range := by + rw [← centralizer_univ, ← Algebra.coe_top (R := R) (A := A), + ← centralizer_coe_map_includeLeft_eq_center_tensorProduct R A B ⊤] + ext + simp [includeLeft, includeLeftRingHom, Set.range_comp] + +/-- +Let `R` be a commutative ring and `A, B` be `R`-algebras where `A` is free as `R`-module. +Then the centralizer of `1 ⊗ B ⊆ A ⊗ B` is `A ⊗ C(B)` where `C(B)` is the center of `B`. +-/ +lemma centralizer_range_includeRight_eq_center_tensorProduct [Module.Free R A] : + Subalgebra.centralizer R + (Algebra.TensorProduct.includeRight : B →ₐ[R] A ⊗[R] B).range = + (Algebra.TensorProduct.map (AlgHom.id R A) (center R B).val).range := by + rw [← centralizer_univ, ← Algebra.coe_top (R := R) (A := B), + ← centralizer_coe_map_includeRight_eq_center_tensorProduct R A B ⊤] + ext + simp [includeRight, includeLeftRingHom, Set.range_comp] + +lemma centralizer_tensorProduct_eq_center_tensorProduct_left [Module.Free R B] : + Subalgebra.centralizer R + (Algebra.TensorProduct.map (AlgHom.id R A) (Algebra.ofId R B)).range = + (Algebra.TensorProduct.map (Subalgebra.center R A).val (AlgHom.id R B)).range := by + rw [← centralizer_coe_range_includeLeft_eq_center_tensorProduct] + simp [Algebra.TensorProduct.map_range] + +lemma centralizer_tensorProduct_eq_center_tensorProduct_right [Module.Free R A] : + Subalgebra.centralizer R + (Algebra.TensorProduct.map (Algebra.ofId R A) (AlgHom.id R B)).range = + (Algebra.TensorProduct.map (AlgHom.id R A) (center R B).val).range := by + rw [← centralizer_range_includeRight_eq_center_tensorProduct] + simp [Algebra.TensorProduct.map_range] + +end Free + +end Subalgebra diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Pointwise.lean b/Mathlib/Algebra/Algebra/Subalgebra/Pointwise.lean index 5003f01d48ca5..b68ffd013e5a0 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Pointwise.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Pointwise.lean @@ -31,8 +31,8 @@ theorem mul_toSubmodule_le (S T : Subalgebra R A) : /-- As submodules, subalgebras are idempotent. -/ @[simp] -theorem mul_self (S : Subalgebra R A) : (Subalgebra.toSubmodule S) * (Subalgebra.toSubmodule S) - = (Subalgebra.toSubmodule S) := by +theorem isIdempotentElem_toSubmodule (S : Subalgebra R A) : + IsIdempotentElem S.toSubmodule := by apply le_antisymm · refine (mul_toSubmodule_le _ _).trans_eq ?_ rw [sup_idem] @@ -40,6 +40,8 @@ theorem mul_self (S : Subalgebra R A) : (Subalgebra.toSubmodule S) * (Subalgebra rw [← mul_one x] exact Submodule.mul_mem_mul hx1 (show (1 : A) ∈ S from one_mem S) +@[deprecated (since := "2025-01-12")] alias mul_self := isIdempotentElem_toSubmodule + /-- When `A` is commutative, `Subalgebra.mul_toSubmodule_le` is strict. -/ theorem mul_toSubmodule {R : Type*} {A : Type*} [CommSemiring R] [CommSemiring A] [Algebra R A] (S T : Subalgebra R A) : (Subalgebra.toSubmodule S) * (Subalgebra.toSubmodule T) @@ -58,7 +60,8 @@ theorem mul_toSubmodule {R : Type*} {A : Type*} [CommSemiring R] [CommSemiring A exact Submodule.mul_mem_mul (show (1 : A) ∈ S from one_mem S) (algebraMap_mem T _) have := Submodule.mul_mem_mul hx hy rwa [mul_assoc, mul_comm _ (Subalgebra.toSubmodule T), ← mul_assoc _ _ (Subalgebra.toSubmodule S), - mul_self, mul_comm (Subalgebra.toSubmodule T), ← mul_assoc, mul_self] at this + isIdempotentElem_toSubmodule, mul_comm T.toSubmodule, ← mul_assoc, + isIdempotentElem_toSubmodule] at this variable {R' : Type*} [Semiring R'] [MulSemiringAction R' A] [SMulCommClass R' R A] diff --git a/Mathlib/Algebra/Azumaya/Defs.lean b/Mathlib/Algebra/Azumaya/Defs.lean new file mode 100644 index 0000000000000..08c3a24d68833 --- /dev/null +++ b/Mathlib/Algebra/Azumaya/Defs.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2025 Yunzhou Xie. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yunzhou Xie, Jujian Zhang +-/ +import Mathlib.Algebra.Module.Projective +import Mathlib.RingTheory.Finiteness.Defs +import Mathlib.RingTheory.TensorProduct.Basic + +/-! +# Azumaya Algebras + +An Azumaya algebra over a commutative ring `R` is a finitely generated, projective +and faithful R-algebra where the tensor product `A ⊗[R] Aᵐᵒᵖ` is isomorphic to the +`R`-endomorphisms of A via the map `f : a ⊗ b ↦ (x ↦ a * x * b.unop)`. +TODO : Add the three more definitions and prove they are equivalent: +· There exists an `R`-algebra `B` such that `B ⊗[R] A` is Morita equivalent to `R`; +· `Aᵐᵒᵖ ⊗[R] A` is Morita equivalent to `R`; +· The center of `A` is `R` and `A` is a separable algebra. + +## Reference + +* [Benson Farb , R. Keith Dennis, *Noncommutative Algebra*][bensonfarb1993] + +## Tags + +Azumaya algebra, central simple algebra, noncommutative algebra +-/ + +variable (R A : Type*) [CommSemiring R] [Semiring A] [Algebra R A] + +open TensorProduct MulOpposite + +/-- `A` as a `A ⊗[R] Aᵐᵒᵖ`-module (or equivalently, an `A`-`A` bimodule). -/ +noncomputable abbrev instModuleTensorProductMop : + Module (A ⊗[R] Aᵐᵒᵖ) A := TensorProduct.Algebra.module + +/-- The canonical map from `A ⊗[R] Aᵐᵒᵖ` to `Module.End R A` where + `a ⊗ b` maps to `f : x ↦ a * x * b`. -/ +noncomputable def AlgHom.mulLeftRight : (A ⊗[R] Aᵐᵒᵖ) →ₐ[R] Module.End R A := + letI : Module (A ⊗[R] Aᵐᵒᵖ) A := TensorProduct.Algebra.module + letI : IsScalarTower R (A ⊗[R] Aᵐᵒᵖ) A := { + smul_assoc := fun r ab a ↦ by + change TensorProduct.Algebra.moduleAux _ _ = _ • TensorProduct.Algebra.moduleAux _ _ + simp } + Algebra.lsmul R (A := A ⊗[R] Aᵐᵒᵖ) R A + +@[simp] +lemma AlgHom.mulLeftRight_apply (a : A) (b : Aᵐᵒᵖ) (x : A) : + AlgHom.mulLeftRight R A (a ⊗ₜ b) x = a * x * b.unop := by + simp only [AlgHom.mulLeftRight, Algebra.lsmul_coe] + change TensorProduct.Algebra.moduleAux _ _ = _ + simp [TensorProduct.Algebra.moduleAux, ← mul_assoc] + +/-- An Azumaya algebra is a finitely generated, projective and faithful R-algebra where + `AlgHom.mulLeftRight R A : (A ⊗[R] Aᵐᵒᵖ) →ₐ[R] Module.End R A` is an isomorphism. -/ +class IsAzumaya extends Module.Projective R A, FaithfulSMul R A, Module.Finite R A : Prop where + bij : Function.Bijective <| AlgHom.mulLeftRight R A diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index 045c79a9465a3..4d62dc8fa144a 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -55,7 +55,7 @@ theorem prod_univ_zero [CommMonoid β] (f : Fin 0 → β) : ∏ i, f i = 1 := /-- A product of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the product of `f x`, for some `x : Fin (n + 1)` times the remaining product -/ @[to_additive "A sum of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the sum of -`f x`, for some `x : Fin (n + 1)` plus the remaining product"] +`f x`, for some `x : Fin (n + 1)` plus the remaining sum"] theorem prod_univ_succAbove [CommMonoid β] {n : ℕ} (f : Fin (n + 1) → β) (x : Fin (n + 1)) : ∏ i, f i = f x * ∏ i : Fin n, f (x.succAbove i) := by rw [univ_succAbove, prod_cons, Finset.prod_map _ x.succAboveEmb] @@ -64,7 +64,7 @@ theorem prod_univ_succAbove [CommMonoid β] {n : ℕ} (f : Fin (n + 1) → β) ( /-- A product of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the product of `f 0` plus the remaining product -/ @[to_additive "A sum of a function `f : Fin (n + 1) → β` over all `Fin (n + 1)` is the sum of -`f 0` plus the remaining product"] +`f 0` plus the remaining sum"] theorem prod_univ_succ [CommMonoid β] {n : ℕ} (f : Fin (n + 1) → β) : ∏ i, f i = f 0 * ∏ i : Fin n, f i.succ := prod_univ_succAbove f 0 @@ -215,17 +215,7 @@ theorem partialProd_left_inv {G : Type*} [Group G] (f : Fin (n + 1) → G) : @[to_additive] theorem partialProd_right_inv {G : Type*} [Group G] (f : Fin n → G) (i : Fin n) : (partialProd f (Fin.castSucc i))⁻¹ * partialProd f i.succ = f i := by - obtain ⟨i, hn⟩ := i - induction i with - | zero => simp [-Fin.succ_mk, partialProd_succ] - | succ i hi => - specialize hi (lt_trans (Nat.lt_succ_self i) hn) - simp only [Fin.coe_eq_castSucc, Fin.succ_mk, Fin.castSucc_mk] at hi ⊢ - rw [← Fin.succ_mk _ _ (lt_trans (Nat.lt_succ_self _) hn), ← Fin.succ_mk _ _ hn] - simp only [partialProd_succ, mul_inv_rev, Fin.castSucc_mk] - -- Porting note: was - -- assoc_rw [hi, inv_mul_cancel_left] - rw [← mul_assoc, mul_left_eq_self, mul_assoc, hi, inv_mul_cancel] + rw [partialProd_succ, inv_mul_cancel_left] /-- Let `(g₀, g₁, ..., gₙ)` be a tuple of elements in `Gⁿ⁺¹`. Then if `k < j`, this says `(g₀g₁...gₖ₋₁)⁻¹ * g₀g₁...gₖ = gₖ`. @@ -388,6 +378,48 @@ theorem finPiFinEquiv_single {m : ℕ} {n : Fin m → ℕ} [∀ i, NeZero (n i)] rintro x hx rw [Pi.single_eq_of_ne hx, Fin.val_zero', zero_mul] +/-- Equivalence between the Sigma type `(i : Fin m) × Fin (n i)` and `Fin (∑ i : Fin m, n i)`. -/ +def finSigmaFinEquiv {m : ℕ} {n : Fin m → ℕ} : (i : Fin m) × Fin (n i) ≃ Fin (∑ i : Fin m, n i) := + match m with + | 0 => @Equiv.equivOfIsEmpty _ _ _ (by simp; exact Fin.isEmpty') + | Nat.succ m => + calc _ ≃ _ := (@finSumFinEquiv m 1).sigmaCongrLeft.symm + _ ≃ _ := Equiv.sumSigmaDistrib _ + _ ≃ _ := finSigmaFinEquiv.sumCongr (Equiv.uniqueSigma _) + _ ≃ _ := finSumFinEquiv + _ ≃ _ := finCongr (Fin.sum_univ_castSucc n).symm + +@[simp] +theorem finSigmaFinEquiv_apply {m : ℕ} {n : Fin m → ℕ} (k : (i : Fin m) × Fin (n i)) : + (finSigmaFinEquiv k : ℕ) = ∑ i : Fin k.1, n (Fin.castLE k.1.2.le i) + k.2 := by + induction m + · exact k.fst.elim0 + rename_i m ih + rcases k with ⟨⟨iv,hi⟩,j⟩ + rw [finSigmaFinEquiv] + unfold finSumFinEquiv + simp only [Equiv.coe_fn_mk, Equiv.sigmaCongrLeft, Equiv.coe_fn_symm_mk, Equiv.instTrans_trans, + Equiv.trans_apply, finCongr_apply, Fin.coe_cast] + conv => + enter [1,1,3] + apply Equiv.sumCongr_apply + by_cases him : iv < m + · conv in Sigma.mk _ _ => + equals ⟨Sum.inl ⟨iv, him⟩, j⟩ => simp [Fin.addCases, him] + simpa using ih _ + · replace him := Nat.eq_of_lt_succ_of_not_lt hi him + subst him + conv in Sigma.mk _ _ => + equals ⟨Sum.inr 0, j⟩ => simp [Fin.addCases, Fin.natAdd] + simp + rfl + +/-- `finSigmaFinEquiv` on `Fin 1 × f` is just `f`-/ +theorem finSigmaFinEquiv_one {n : Fin 1 → ℕ} (ij : (i : Fin 1) × Fin (n i)) : + (finSigmaFinEquiv ij : ℕ) = ij.2 := by + rw [finSigmaFinEquiv_apply, add_left_eq_self] + apply @Finset.sum_of_isEmpty _ _ _ _ (by simpa using Fin.isEmpty') + namespace List section CommMonoid diff --git a/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean b/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean index a83b1bce754f2..795c3077f3dde 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset/Basic.lean @@ -1434,6 +1434,50 @@ theorem prod_unique_nonempty {α β : Type*} [CommMonoid β] [Unique α] (s : Fi (h : s.Nonempty) : ∏ x ∈ s, f x = f default := by rw [h.eq_singleton_default, Finset.prod_singleton] +section Image_Overlap + +variable {α β ι : Type*} [DecidableEq ι] [DecidableEq α] + +@[to_additive] +lemma prod_filter_of_pairwise_eq_one [CommMonoid β] {f : ι → α} {g : α → β} {n : ι} {I : Finset ι} + (hn : n ∈ I) (hf : (I : Set ι).Pairwise fun i j ↦ f i = f j → g (f i) = 1) : + ∏ j ∈ filter (fun j ↦ f j = f n) I, g (f j) = g (f n) := by + have h j (hj : j ∈ (filter (fun i ↦ f i = f n) I).erase n) : g (f j) = 1 := by + simp only [mem_erase, mem_filter] at hj + exact hf hj.2.1 hn hj.1 hj.2.2 + rw [← mul_one (g (f n)), ← prod_eq_one h, + ← mul_prod_erase (filter (f · = f n) I) (fun i ↦ g (f i)) <| mem_filter.mpr ⟨hn, by rfl⟩] + +/-- A version of `Finset.prod_map` and `Finset.prod_image`, but we do not assume that `f` is +injective. Rather, we assume that the image of `f` on `I` only overlaps where `g (f i) = 1`. +The conclusion is the same as in `prod_image`.-/ +@[to_additive (attr := simp) +"A version of `Finset.sum_map` and `Finset.sum_image`, but we do not assume that `f` is +injective. Rather, we assume that the image of `f` on `I` only overlaps where `g (f i) = 0`. +The conclusion is the same as in `sum_image`."] +lemma prod_image_of_pairwise_eq_one [CommMonoid β] {f : ι → α} {g : α → β} {I : Finset ι} + (hf : (I : Set ι).Pairwise fun i j ↦ f i = f j → g (f i) = 1) : + ∏ s in I.image f, g s = ∏ i in I, g (f i) := by + rw [prod_image'] + exact fun n hnI => (prod_filter_of_pairwise_eq_one hnI hf).symm + +/-- A version of `Finset.prod_map` and `Finset.prod_image`, but we do not assume that `f` is +injective. Rather, we assume that the images of `f` are disjoint on `I`, and `g ⊥ = 1`. The +conclusion is the same as in `prod_image`.-/ +@[to_additive (attr := simp) +"A version of `Finset.sum_map` and `Finset.sum_image`, but we do not assume that `f` is +injective. Rather, we assume that the images of `f` are disjoint on `I`, and `g ⊥ = 0`. The +conclusion is the same as in `sum_image`." +] +lemma prod_image_of_disjoint [CommMonoid β] [PartialOrder α] [OrderBot α] {f : ι → α} {g : α → β} + (hg_bot : g ⊥ = 1) {I : Finset ι} (hf_disj : (I : Set ι).PairwiseDisjoint f) : + ∏ s in I.image f, g s = ∏ i in I, g (f i) := by + refine prod_image_of_pairwise_eq_one <| hf_disj.imp fun i j (hdisj : Disjoint _ _) hfij ↦ ?_ + rw [← hfij, disjoint_self] at hdisj + rw [hdisj, hg_bot] + +end Image_Overlap + end Finset namespace Fintype diff --git a/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean b/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean index 64683dca211e5..8264c7339228d 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset/Defs.lean @@ -481,6 +481,16 @@ theorem prod_dite_irrel (p : Prop) [Decidable p] (s : Finset α) (f : p → α if h : p then ∏ x ∈ s, f h x else ∏ x ∈ s, g h x := by split_ifs with h <;> rfl +@[to_additive] +theorem ite_prod_one (p : Prop) [Decidable p] (s : Finset α) (f : α → β) : + (if p then (∏ x ∈ s, f x) else 1) = ∏ x ∈ s, if p then f x else 1 := by + simp only [prod_ite_irrel, prod_const_one] + +@[to_additive] +theorem ite_one_prod (p : Prop) [Decidable p] (s : Finset α) (f : α → β) : + (if p then 1 else (∏ x ∈ s, f x)) = ∏ x ∈ s, if p then 1 else f x := by + simp only [prod_ite_irrel, prod_const_one] + @[to_additive] theorem nonempty_of_prod_ne_one (h : ∏ x ∈ s, f x ≠ 1) : s.Nonempty := s.eq_empty_or_nonempty.elim (fun H => False.elim <| h <| H.symm ▸ prod_empty) id diff --git a/Mathlib/Algebra/BigOperators/Pi.lean b/Mathlib/Algebra/BigOperators/Pi.lean index 88511371f4a9e..0db9fbd374673 100644 --- a/Mathlib/Algebra/BigOperators/Pi.lean +++ b/Mathlib/Algebra/BigOperators/Pi.lean @@ -181,3 +181,22 @@ def Pi.monoidHomMulEquiv {ι : Type*} [Fintype ι] [DecidableEq ι] (M : ι → MonoidHom.mul_apply, mul_apply] end MulEquiv + +variable [Finite ι] [DecidableEq ι] {M : Type*} + +-- manually additivized to fix variable names +-- See https://github.com/leanprover-community/mathlib4/issues/11462 +lemma Pi.single_induction [AddCommMonoid M] (p : (ι → M) → Prop) (f : ι → M) + (zero : p 0) (add : ∀ f g, p f → p g → p (f + g)) + (single : ∀ i m, p (Pi.single i m)) : p f := by + cases nonempty_fintype ι + rw [← Finset.univ_sum_single f] + exact Finset.sum_induction _ _ add zero (by simp [single]) + +@[to_additive (attr := elab_as_elim) existing] +lemma Pi.mulSingle_induction [CommMonoid M] (p : (ι → M) → Prop) (f : ι → M) + (one : p 1) (mul : ∀ f g, p f → p g → p (f * g)) + (mulSingle : ∀ i m, p (Pi.mulSingle i m)) : p f := by + cases nonempty_fintype ι + rw [← Finset.univ_prod_mulSingle f] + exact Finset.prod_induction _ _ mul one (by simp [mulSingle]) diff --git a/Mathlib/Algebra/BrauerGroup/Defs.lean b/Mathlib/Algebra/BrauerGroup/Defs.lean new file mode 100644 index 0000000000000..02c990d52f2ac --- /dev/null +++ b/Mathlib/Algebra/BrauerGroup/Defs.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2025 Yunzhou Xie. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yunzhou Xie, Jujian Zhang +-/ +import Mathlib.Algebra.Central.Defs +import Mathlib.LinearAlgebra.FiniteDimensional.Defs +import Mathlib.LinearAlgebra.Matrix.Reindex +import Mathlib.Algebra.Category.AlgebraCat.Basic + +/-! +# Definition of Brauer group of a field K + +We introduce the definition of Brauer group of a field K, which is the quotient of the set of +all finite-dimensional central simple algebras over K modulo the Brauer Equivalence where two +central simple algebras `A` and `B` are Brauer Equivalent if there exist `n, m ∈ ℕ+` such +that `Mₙ(A) ≃ₐ[K] Mₙ(B)`. + +# TODOs +1. Prove that the Brauer group is an abelian group where multiplication is defined as tensorproduct. +2. Prove that the Brauer group is a functor from the category of fields to the category of groups. +3. Prove that over a field, being Brauer equivalent is the same as being Morita equivalent. + +# References +* [Algebraic Number Theory, *J.W.S Cassels*][cassels1967algebraic] + +## Tags +Brauer group, Central simple algebra, Galois Cohomology +-/ + +universe u v + +/-- `CSA` is the set of all finite dimensional central simple algebras over field `K`, for its + generalisation over a `CommRing` please find `IsAzumaya` in `Mathlib.Algebra.Azumaya.Defs`. -/ +structure CSA (K : Type u) [Field K] extends AlgebraCat.{v} K where + /-- Any member of `CSA` is central. -/ + [isCentral : Algebra.IsCentral K carrier] + /-- Any member of `CSA` is simple. -/ + [isSimple : IsSimpleRing carrier] + /-- Any member of `CSA` is finite-dimensional. -/ + [fin_dim : FiniteDimensional K carrier] + +variable {K : Type u} [Field K] + +instance : CoeSort (CSA.{u, v} K) (Type v) := ⟨(·.carrier)⟩ + +attribute [instance] CSA.isCentral CSA.isSimple CSA.fin_dim + +/-- Two finite dimensional central simple algebras `A` and `B` are Brauer Equivalent + if there exist `n, m ∈ ℕ+` such that the `Mₙ(A) ≃ₐ[K] Mₙ(B)`. -/ +abbrev IsBrauerEquivalent (A B : CSA K) : Prop := + ∃n m : ℕ, n ≠ 0 ∧ m ≠ 0 ∧ (Nonempty <| Matrix (Fin n) (Fin n) A ≃ₐ[K] Matrix (Fin m) (Fin m) B) + +namespace IsBrauerEquivalent + +@[refl] +lemma refl (A : CSA K) : IsBrauerEquivalent A A := + ⟨1, 1, one_ne_zero, one_ne_zero, ⟨AlgEquiv.refl⟩⟩ + +@[symm] +lemma symm {A B : CSA K} (h : IsBrauerEquivalent A B) : IsBrauerEquivalent B A := + let ⟨n, m, hn, hm, ⟨iso⟩⟩ := h + ⟨m, n, hm, hn, ⟨iso.symm⟩⟩ + +open Matrix in +@[trans] +lemma trans {A B C : CSA K} (hAB : IsBrauerEquivalent A B) (hBC : IsBrauerEquivalent B C) : + IsBrauerEquivalent A C := by + obtain ⟨n, m, hn, hm, ⟨iso1⟩⟩ := hAB + obtain ⟨p, q, hp, hq, ⟨iso2⟩⟩ := hBC + exact ⟨p * n, m * q, by simp_all, by simp_all, + ⟨reindexAlgEquiv _ _ finProdFinEquiv |>.symm.trans <| compAlgEquiv _ _ _ _|>.symm.trans <| + iso1.mapMatrix (m := Fin p)|>.trans <| compAlgEquiv _ _ _ _|>.trans <| + reindexAlgEquiv K B (.prodComm (Fin p) (Fin m))|>.trans <| compAlgEquiv _ _ _ _|>.symm.trans <| + iso2.mapMatrix.trans <| compAlgEquiv _ _ _ _|>.trans <| reindexAlgEquiv _ _ finProdFinEquiv⟩⟩ + +lemma is_eqv : Equivalence (IsBrauerEquivalent (K := K)) where + refl := refl + symm := symm + trans := trans + +end IsBrauerEquivalent + +variable (K) + +/-- `CSA` equipped with Brauer Equivalence is indeed a setoid. -/ +def Brauer.CSA_Setoid: Setoid (CSA K) where + r := IsBrauerEquivalent + iseqv := IsBrauerEquivalent.is_eqv + +/-- `BrauerGroup` is the set of all finite-dimensional central simple algebras quotient + by Brauer Equivalence. -/ +abbrev BrauerGroup := Quotient (Brauer.CSA_Setoid K) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 87dd242280419..808f200d06da4 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -61,15 +61,38 @@ variable {R} in structure Hom (A B : AlgebraCat.{v} R) where private mk :: /-- The underlying algebra map. -/ - hom : A →ₐ[R] B + hom' : A →ₐ[R] B instance : Category (AlgebraCat.{v} R) where Hom A B := Hom A B id A := ⟨AlgHom.id R A⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {M N : AlgebraCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where - coe f := f.hom +instance : ConcreteCategory (AlgebraCat.{v} R) (· →ₐ[R] ·) where + hom := Hom.hom' + ofHom := Hom.mk + +variable {R} in +/-- Turn a morphism in `AlgebraCat` back into an `AlgHom`. -/ +abbrev Hom.hom {A B : AlgebraCat.{v} R} (f : Hom A B) := + ConcreteCategory.hom (C := AlgebraCat R) f + +variable {R} in +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat`. -/ +abbrev ofHom {A B : Type v} [Ring A] [Ring B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) : + of R A ⟶ of R B := + ConcreteCategory.ofHom (C := AlgebraCat R) f + +variable {R} in +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (A B : AlgebraCat.{v} R) (f : Hom A B) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {A : AlgebraCat.{v} R} : (𝟙 A : A ⟶ A).hom = AlgHom.id R A := rfl @@ -90,11 +113,7 @@ lemma comp_apply {A B C : AlgebraCat.{v} R} (f : A ⟶ B) (g : B ⟶ C) (a : A) lemma hom_ext {A B : AlgebraCat.{v} R} {f g : A ⟶ B} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ -abbrev ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] - (f : X →ₐ[R] Y) : of R X ⟶ of R Y := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) : (ofHom f).hom = f := rfl @@ -127,12 +146,6 @@ lemma hom_inv_apply {A B : AlgebraCat.{v} R} (e : A ≅ B) (x : B) : e.hom (e.in instance : Inhabited (AlgebraCat R) := ⟨of R R⟩ -instance : HasForget.{v} (AlgebraCat.{v} R) where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - lemma forget_obj {A : AlgebraCat.{v} R} : (forget (AlgebraCat.{v} R)).obj A = A := rfl lemma forget_map {A B : AlgebraCat.{v} R} (f : A ⟶ B) : diff --git a/Mathlib/Algebra/Category/Grp/AB.lean b/Mathlib/Algebra/Category/Grp/AB.lean index dde16f6cf9ea2..3038d2cedbcac 100644 --- a/Mathlib/Algebra/Category/Grp/AB.lean +++ b/Mathlib/Algebra/Category/Grp/AB.lean @@ -32,14 +32,15 @@ noncomputable instance : intro x (hx : _ = _) dsimp at hx rcases Concrete.colimit_exists_rep S.X₂ x with ⟨j, y, rfl⟩ - rw [← comp_apply, colimMap_eq, colimit.ι_map, comp_apply, + rw [← CategoryTheory.comp_apply, colimMap_eq, colimit.ι_map, CategoryTheory.comp_apply, ← map_zero (by exact colimit.ι S.X₃ j : (S.X₃).obj j →+ ↑(colimit S.X₃))] at hx rcases Concrete.colimit_exists_of_rep_eq.{u, u, u} S.X₃ _ _ hx with ⟨k, e₁, e₂, hk : _ = S.X₃.map e₂ 0⟩ - rw [map_zero, ← comp_apply, ← NatTrans.naturality, comp_apply] at hk + rw [map_zero, ← CategoryTheory.comp_apply, ← NatTrans.naturality, CategoryTheory.comp_apply] + at hk rcases hS k hk with ⟨t, ht⟩ use colimit.ι S.X₁ k t - erw [← comp_apply, colimit.ι_map, comp_apply, ht] + erw [← CategoryTheory.comp_apply, colimit.ι_map, CategoryTheory.comp_apply, ht] exact colimit.w_apply S.X₂ e₁ y) noncomputable instance : diff --git a/Mathlib/Algebra/Category/Grp/Basic.lean b/Mathlib/Algebra/Category/Grp/Basic.lean index e84967eb2a42b..356b887d22a5c 100644 --- a/Mathlib/Algebra/Category/Grp/Basic.lean +++ b/Mathlib/Algebra/Category/Grp/Basic.lean @@ -132,6 +132,13 @@ theorem ofHom_apply {X Y : Type _} [Group X] [Group Y] (f : X →* Y) (x : X) : (ofHom f) x = f x := rfl +@[to_additive] +lemma ofHom_injective {X Y : Type u} [Group X] [Group Y] : + Function.Injective (fun (f : X →* Y) ↦ ofHom f) := by + intro _ _ h + ext + apply DFunLike.congr_fun h + @[to_additive] instance ofUnique (G : Type*) [Group G] [i : Unique G] : Unique (Grp.of G) := i @@ -276,6 +283,13 @@ theorem ofHom_apply {X Y : Type _} [CommGroup X] [CommGroup Y] (f : X →* Y) (x @DFunLike.coe (X →* Y) X (fun _ ↦ Y) _ (ofHom f) x = f x := rfl +@[to_additive] +lemma ofHom_injective {X Y : Type u} [CommGroup X] [CommGroup Y] : + Function.Injective (fun (f : X →* Y) ↦ ofHom f) := by + intro _ _ h + ext + apply DFunLike.congr_fun h + -- We verify that simp lemmas apply when coercing morphisms to functions. @[to_additive] example {R S : CommGrp} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [h] diff --git a/Mathlib/Algebra/Category/Grp/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean index a7d7824126397..f74b2fe7db828 100644 --- a/Mathlib/Algebra/Category/Grp/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -12,12 +12,12 @@ import Mathlib.GroupTheory.QuotientGroup.Defs /-! # The category of additive commutative groups has all colimits. -This file constructs colimits in the categpry of additive commutative groups, as +This file constructs colimits in the category of additive commutative groups, as quotients of finitely supported functions. -/ -universe w u v +universe u' w u v open CategoryTheory Limits @@ -94,6 +94,87 @@ lemma Quot.map_ι [DecidableEq J] {j j' : J} {f : j ⟶ j'} (x : F.obj j) : simp only [DFinsupp.singleAddHom_apply] exact AddSubgroup.subset_closure ⟨j, j', f, x, rfl⟩ +/-- +The obvious additive map from `Quot F` to `Quot (F ⋙ uliftFunctor.{u'})`. +-/ +def quotToQuotUlift [DecidableEq J] : Quot F →+ Quot (F ⋙ uliftFunctor.{u'}) := by + refine QuotientAddGroup.lift (Relations F) (DFinsupp.sumAddHom (fun j ↦ (Quot.ι _ j).comp + AddEquiv.ulift.symm.toAddMonoidHom)) ?_ + rw [AddSubgroup.closure_le] + intro _ hx + obtain ⟨j, j', u, a, rfl⟩ := hx + rw [SetLike.mem_coe, AddMonoidHom.mem_ker, map_sub, DFinsupp.sumAddHom_single, + DFinsupp.sumAddHom_single] + change Quot.ι (F ⋙ uliftFunctor) j' ((F ⋙ uliftFunctor).map u (AddEquiv.ulift.symm a)) - _ = _ + rw [Quot.map_ι] + dsimp + rw [sub_self] + +lemma quotToQuotUlift_ι [DecidableEq J] (j : J) (x : F.obj j) : + quotToQuotUlift F (Quot.ι F j x) = Quot.ι _ j (ULift.up x) := by + dsimp [quotToQuotUlift, Quot.ι] + conv_lhs => erw [AddMonoidHom.comp_apply (QuotientAddGroup.mk' (Relations F)) + (DFinsupp.singleAddHom _ j), QuotientAddGroup.lift_mk'] + simp only [DFinsupp.singleAddHom_apply, DFinsupp.sumAddHom_single, AddMonoidHom.coe_comp, + AddMonoidHom.coe_coe, Function.comp_apply] + rfl + +/-- +The obvious additive map from `Quot (F ⋙ uliftFunctor.{u'})` to `Quot F`. +-/ +def quotUliftToQuot [DecidableEq J] : Quot (F ⋙ uliftFunctor.{u'}) →+ Quot F := by + refine QuotientAddGroup.lift (Relations (F ⋙ uliftFunctor)) + (DFinsupp.sumAddHom (fun j ↦ (Quot.ι _ j).comp AddEquiv.ulift.toAddMonoidHom)) ?_ + rw [AddSubgroup.closure_le] + intro _ hx + obtain ⟨j, j', u, a, rfl⟩ := hx + rw [SetLike.mem_coe, AddMonoidHom.mem_ker, map_sub, DFinsupp.sumAddHom_single, + DFinsupp.sumAddHom_single] + change Quot.ι F j' (F.map u (AddEquiv.ulift a)) - _ = _ + rw [Quot.map_ι] + dsimp + rw [sub_self] + +lemma quotUliftToQuot_ι [DecidableEq J] (j : J) (x : (F ⋙ uliftFunctor.{u'}).obj j) : + quotUliftToQuot F (Quot.ι _ j x) = Quot.ι F j x.down := by + dsimp [quotUliftToQuot, Quot.ι] + conv_lhs => erw [AddMonoidHom.comp_apply (QuotientAddGroup.mk' (Relations (F ⋙ uliftFunctor))) + (DFinsupp.singleAddHom _ j), QuotientAddGroup.lift_mk'] + simp only [Functor.comp_obj, uliftFunctor_obj, coe_of, DFinsupp.singleAddHom_apply, + DFinsupp.sumAddHom_single, AddMonoidHom.coe_comp, AddMonoidHom.coe_coe, Function.comp_apply] + rfl + +/-- +The additive equivalence between `Quot F` and `Quot (F ⋙ uliftFunctor.{u'})`. +-/ +@[simp] +def quotQuotUliftAddEquiv [DecidableEq J] : Quot F ≃+ Quot (F ⋙ uliftFunctor.{u'}) where + toFun := quotToQuotUlift F + invFun := quotUliftToQuot F + left_inv x := by + conv_rhs => rw [← AddMonoidHom.id_apply _ x] + rw [← AddMonoidHom.comp_apply, Quot.addMonoidHom_ext F (f := (quotUliftToQuot F).comp + (quotToQuotUlift F)) (fun j a ↦ ?_)] + rw [AddMonoidHom.comp_apply, AddMonoidHom.id_apply, quotToQuotUlift_ι, quotUliftToQuot_ι] + right_inv x := by + conv_rhs => rw [← AddMonoidHom.id_apply _ x] + rw [← AddMonoidHom.comp_apply, Quot.addMonoidHom_ext _ (f := (quotToQuotUlift F).comp + (quotUliftToQuot F)) (fun j a ↦ ?_)] + rw [AddMonoidHom.comp_apply, AddMonoidHom.id_apply, quotUliftToQuot_ι, quotToQuotUlift_ι] + rfl + map_add' _ _ := by simp + +lemma Quot.desc_quotQuotUliftAddEquiv [DecidableEq J] (c : Cocone F) : + (Quot.desc (F ⋙ uliftFunctor.{u'}) (uliftFunctor.{u'}.mapCocone c)).comp + (quotQuotUliftAddEquiv F).toAddMonoidHom = + AddEquiv.ulift.symm.toAddMonoidHom.comp (Quot.desc F c) := by + refine Quot.addMonoidHom_ext _ (fun j a ↦ ?_) + dsimp + simp only [quotToQuotUlift_ι, Functor.comp_obj, uliftFunctor_obj, coe_of, ι_desc, + Functor.const_obj_obj] + erw [Quot.ι_desc] + rfl + /-- (implementation detail) A morphism of commutative additive groups `Quot F →+ A` induces a cocone on `F` as long as the universes work out. -/ diff --git a/Mathlib/Algebra/Category/Grp/LargeColimits.lean b/Mathlib/Algebra/Category/Grp/LargeColimits.lean new file mode 100644 index 0000000000000..9d0a40892ca36 --- /dev/null +++ b/Mathlib/Algebra/Category/Grp/LargeColimits.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2025 Sophie Morel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sophie Morel +-/ +import Mathlib.Algebra.Category.Grp.Colimits +import Mathlib.Algebra.Module.CharacterModule +import Mathlib.Algebra.Group.Equiv.Basic + +/-! +# Existence of "big" colimits in the category of additive commutative groups + +If `F : J ⥤ AddCommGrp.{w}` is a functor, we show that `F` admits a colimit if and only +if `Colimits.Quot F` (the quotient of the direct sum of the commutative groups `F.obj j` +by the relations given by the morphisms in the diagram) is `w`-small. + +-/ + +universe w u v + +open CategoryTheory Limits + +namespace AddCommGrp + +variable {J : Type u} [Category.{v} J] {F : J ⥤ AddCommGrp.{w}} (c : Cocone F) + +open Colimits + +/-- +If `c` is a cocone of `F` such that `Quot.desc F c` is bijective, then `c` is a colimit +cocone of `F`. +-/ +lemma isColimit_iff_bijective_desc [DecidableEq J] : + Nonempty (IsColimit c) ↔ Function.Bijective (Quot.desc F c) := by + refine ⟨fun ⟨hc⟩ => ?_, fun h ↦ Nonempty.intro (isColimit_of_bijective_desc F c h)⟩ + change Function.Bijective (Quot.desc F c).toIntLinearMap + rw [← CharacterModule.dual_bijective_iff_bijective] + refine ⟨fun χ ψ eq ↦ ?_, fun χ ↦ ?_⟩ + · apply (AddMonoidHom.postcompEquiv (@AddEquiv.ulift (AddCircle (1 : ℚ)) _).symm _).injective + apply ofHom_injective + refine hc.hom_ext (fun j ↦ ?_) + ext x + rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply, ← Quot.ι_desc _ c j x] + exact ULift.down_injective (DFunLike.congr_fun eq (Quot.ι F j x)) + · set c' : Cocone F := + { pt := AddCommGrp.of (ULift (AddCircle (1 : ℚ))) + ι := + { app j := AddCommGrp.ofHom (((@AddEquiv.ulift _ _).symm.toAddMonoidHom.comp χ).comp + (Quot.ι F j)) + naturality {j j'} u := by + ext + change ofHom ((AddEquiv.ulift.symm.toAddMonoidHom.comp χ).comp _) (F.map u _) = _ + dsimp + rw [Quot.map_ι F (f := u)] + rfl } } + use AddEquiv.ulift.toAddMonoidHom.comp (hc.desc c') + refine Quot.addMonoidHom_ext _ (fun j x ↦ ?_) + dsimp + rw [Quot.ι_desc] + change AddEquiv.ulift ((c.ι.app j ≫ hc.desc c') x) = _ + rw [hc.fac] + dsimp [c'] + rw [AddEquiv.apply_symm_apply] + +/-- +A functor `F : J ⥤ AddCommGrp.{w}` has a colimit if and only if `Colimits.Quot F` is +`w`-small. +-/ +lemma hasColimit_iff_small_quot [DecidableEq J] : HasColimit F ↔ Small.{w} (Quot F) := + ⟨fun _ ↦ Small.mk ⟨_, ⟨(Equiv.ofBijective _ ((isColimit_iff_bijective_desc (colimit.cocone F)).mp + ⟨colimit.isColimit _⟩))⟩⟩, hasColimit_of_small_quot F⟩ + +end AddCommGrp diff --git a/Mathlib/Algebra/Category/Grp/Ulift.lean b/Mathlib/Algebra/Category/Grp/Ulift.lean index 4e58356a103e1..17c4fe820c0c4 100644 --- a/Mathlib/Algebra/Category/Grp/Ulift.lean +++ b/Mathlib/Algebra/Category/Grp/Ulift.lean @@ -3,9 +3,10 @@ Copyright (c) 2024 Sophie Morel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sophie Morel -/ +import Mathlib.Algebra.Category.Grp.LargeColimits import Mathlib.Algebra.Category.Grp.Limits +import Mathlib.Algebra.Module.CharacterModule import Mathlib.CategoryTheory.Limits.Preserves.Ulift -import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor /-! # Properties of the universe lift functor for groups @@ -14,8 +15,56 @@ This file shows that the functors `Grp.uliftFunctor` and `CommGrp.uliftFunctor` (as well as the additive versions) are fully faithful, preserves all limits and create small limits. -It also shows that `AddCommGrp.uliftFunctor` is an additive functor. - +Full faithfulness is pretty obvious. To prove that the functors preserves limits, +we use the fact that the forgetful functor from `Grp` or `CommGrp` into `Type` +creates all limits (because of the way limits are constructed in `Grp` and `CommGrp`), +and that the universe lift functor on `Type` preserves all limits. Once we know +that `Grp.uliftFunctor` preserves all limits and is fully faithful, it will +automatically create all limits that exist, i.e. all small ones. + +We then switch to `AddCommGrp` and show that `AddCommGrp.uliftFunctor` preserves zero morphisms +and is an additive functor, which again is pretty obvious. + +The last result is a proof that `AddCommGrp.uliftFunctor` preserves all colimits +(and hence creates small colimits). This is the only non-formal part of this file, +as we follow the same strategy as for the categories `Type`. + +Suppose that we have a functor `K : J ⥤ AddCommGrp.{u}` (with `J` any category), a +colimit cocone `c` of `K` and a cocone `lc` of `K ⋙ uliftFunctor.{u v}`. We want to +construct a morphism of cocones `uliftFunctor.mapCocone c → lc` witnessing the fact +that `uliftFunctor.mapCocone c` is also a colimit cocone, but we have no direct way +to do this. The idea is to use that `AddCommGrp.{max v u}` has a small cogenerator, +which is just the additive (rational) circle `ℚ / ℤ`, so any abelian group of +any size can be recovered from its morphisms into `ℚ / ℤ`. More precisely, the functor +sending an abelian group `A` to its dual `A →+ ℚ / ℤ` is fully faithful, *if* we consider +the dual as a (right) module over the endomorphism ring of `ℚ / ℤ`. So an abelian +group `C` is totally determined by the restriction of the coyoneda +functor `A ↦ (C →+ A)` to the category of abelian groups at a smaller universe level. +We do not develop this totally here but do a direct construction. Every time we have +a morphism from `lc.pt` into `ℚ / ℤ`, or more generally into any small abelian group +`A`, we can construct a cocone of `K ⋙ uliftFunctor` by sticking `ULift A` at the +cocone point (this is `CategoryTheory.Limits.Cocone.extensions`), and this cocone is +actually made up of `u`-small abelian groups, hence gives a cocone of `K`. Using +the universal property of `c`, we get a morphism `c.pt →+ A`. So we have constructed +a map `(lc.pt →+ A) → (c.pt →+ A)`, and it is easy to prove that it is compatible with +postcomposition of morphisms of small abelian groups. To actually get the +morphism `c.pt →+ lc.pt`, we apply this to the canonical embedding of `lc.pt` into +`Π (_ : lc.pt →+ ℚ / ℤ), ℚ / ℤ` (this group is not small but is a product of small +groups, so our construction extends). To show that the image of `c.pt` in this product +is contained in that of `lc.pt`, we use the compatibility with postcomposition and the +fact that we can detect elements of the image just by applying morphisms from +`Π (_ : lc.pt →+ ℚ / ℤ), ℚ / ℤ` to `ℚ / ℤ`. + +Note that this does *not* work for noncommutative groups, because the existence of +simple groups of arbitrary size implies that a general object `G` of `Grp` is not +determined by the restriction of `coyoneda.obj G` to the category of groups at +a smaller universe level. Indeed, the functor `Grp.uliftFunctor` does not commute +with arbitrary colimits: if we take an increasing family `K` of simple groups in +`Grp.{u}` of unbounded cardinality indexed by a linearly ordered type +(for example finitary alternating groups on a family of types in `u` of unbouded cardinality), +then the colimit of `K` in `Grp.{u}` exists and is the trivial group; meanwhile, the colimit +of `K ⋙ Grp.uliftFunctor.{u + 1}` is nontrivial (it is the "union" of all the `K j`, which is +too big to be in `Grp.{u}`). -/ universe v w w' u @@ -125,11 +174,32 @@ noncomputable instance : CreatesLimitsOfSize.{w, u} uliftFunctor.{v, u} where end CommGrp -namespace AddCommGroup +namespace AddCommGrp /-- The universe lift for commutative additive groups is additive. -/ instance uliftFunctor_additive : AddCommGrp.uliftFunctor.{u,v}.Additive where -end AddCommGroup +open Colimits in +/-- +The functor `uliftFunctor : AddCommGrp.{u} ⥤ AddCommGrp.{max u v}` preserves colimits +of arbitrary size. +-/ +noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} where + preservesColimitsOfShape {J _} := + { preservesColimit := fun {F} ↦ + { preserves := fun {c} hc ↦ by + classical + rw [isColimit_iff_bijective_desc, ← Function.Bijective.of_comp_iff _ + (quotQuotUliftAddEquiv F).bijective, ← AddEquiv.coe_toAddMonoidHom, + ← AddMonoidHom.coe_comp, Quot.desc_quotQuotUliftAddEquiv] + exact ULift.up_bijective.comp ((isColimit_iff_bijective_desc c).mp (Nonempty.intro hc)) } } + +/-- +The functor `uliftFunctor : AddCommGrp.{u} ⥤ AddCommGrp.{max u v}` creates `u`-small colimits. +-/ +noncomputable instance : CreatesColimitsOfSize.{w, u} uliftFunctor.{v, u} where + CreatesColimitsOfShape := { CreatesColimit := fun {_} ↦ createsColimitOfFullyFaithfulOfPreserves } + +end AddCommGrp diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index dbbc854bbc197..d7640873b9f08 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -212,7 +212,6 @@ universe v u we will equip with a category structure where the morphisms are formal `R`-linear combinations of the morphisms in `C`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/pull/5171): Removed has_nonempty_instance nolint; linter not ported yet @[nolint unusedArguments] def Free (_ : Type*) (C : Type u) := C diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 4c63c2f7e9aea..1692cf8873858 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -89,20 +89,40 @@ variable {R} in structure Hom (M N : ModuleCat.{v} R) where private mk :: /-- The underlying linear map. -/ - hom : M →ₗ[R] N + hom' : M →ₗ[R] N instance moduleCategory : Category.{v, max (v+1) u} (ModuleCat.{v} R) where Hom M N := Hom M N id _ := ⟨LinearMap.id⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + 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 +instance : ConcreteCategory (ModuleCat.{v} R) (· →ₗ[R] ·) where + hom := Hom.hom' + ofHom := Hom.mk section variable {R} +/-- Turn a morphism in `ModuleCat` back into a `LinearMap`. -/ +abbrev Hom.hom {A B : ModuleCat.{v} R} (f : Hom A B) := + ConcreteCategory.hom (C := ModuleCat R) f + +/-- Typecheck a `LinearMap` as a morphism in `ModuleCat`. -/ +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 := + ConcreteCategory.ofHom (C := ModuleCat R) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (A B : ModuleCat.{v} R) (f : Hom A B) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ + @[simp] lemma hom_id {M : ModuleCat.{v} R} : (𝟙 M : M ⟶ M).hom = LinearMap.id := rfl @@ -137,15 +157,9 @@ 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 -/- 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. -/ +@[simp] 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 @@ -444,14 +458,17 @@ end variable (M N : ModuleCat.{v} R) -/-- `ModuleCat.Hom.hom` as an isomorphism of monoids. -/ -@[simps] -def endMulEquiv : End M ≃* (M →ₗ[R] M) where +/-- `ModuleCat.Hom.hom` as an isomorphism of rings. -/ +@[simps!] def endRingEquiv : End M ≃+* (M →ₗ[R] M) where toFun := ModuleCat.Hom.hom invFun := ModuleCat.ofHom map_mul' _ _ := rfl left_inv _ := rfl right_inv _ := rfl + map_add' _ _ := rfl + +/-- `ModuleCat.Hom.hom` as an isomorphism of monoids. -/ +@[deprecated (since := "2025-01-23")] alias endMulEquiv := endRingEquiv /-- 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. -/ @@ -537,9 +554,9 @@ a morphism between the underlying objects in `AddCommGrp` and the compatibility with the scalar multiplication. -/ @[simps] def homMk : M ⟶ N where - hom.toFun := φ - hom.map_add' _ _ := φ.map_add _ _ - hom.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 @@ -568,7 +585,7 @@ variable {R : Type*} [CommRing R] namespace ModuleCat /-- Turn a bilinear map into a homomorphism. -/ -@[simps] +@[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 diff --git a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean index 4387064834c7e..cb8a2f9ed669a 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Biproducts.lean @@ -86,7 +86,7 @@ variable {J : Type w} (f : J → ModuleCat.{max w v} R) /-- The map from an arbitrary cone over an indexed family of abelian groups to the cartesian product of those groups. -/ -@[simps] +@[simps!] def lift (s : Fan f) : s.pt ⟶ ModuleCat.of R (∀ j, f j) := ofHom { toFun := fun x j => s.π.app ⟨j⟩ x diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index b922ed2155138..fa3ae76a99db8 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -298,7 +298,7 @@ def map' {M1 M2 : ModuleCat.{v} R} (l : M1 ⟶ M2) : obj' f M1 ⟶ obj' f M2 := theorem map'_id {M : ModuleCat.{v} R} : map' f (𝟙 M) = 𝟙 _ := by ext x - simp [map'] + simp [map', obj'] 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₂₃ := by @@ -415,7 +415,7 @@ instance : CoeFun (obj' f M) fun _ => S → 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] +@[simps!] def map' {M M' : ModuleCat R} (g : M ⟶ M') : obj' f M ⟶ obj' f M' := ofHom { toFun := fun h => g.hom.comp h @@ -610,7 +610,7 @@ 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 hom_apply] +@[simps! hom_apply] def HomEquiv.toRestrictScalars {X Y} (g : (extendScalars f).obj X ⟶ Y) : X ⟶ (restrictScalars f).obj Y := -- TODO: after https://github.com/leanprover-community/mathlib4/pull/19511 we need to hint `(Y := ...)`. @@ -649,7 +649,7 @@ 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 hom_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 @@ -734,7 +734,7 @@ 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 hom_apply] +@[simps! hom_apply] def Counit.map {Y} : (restrictScalars f ⋙ extendScalars f).obj Y ⟶ Y := ofHom { toFun := @@ -873,8 +873,7 @@ lemma homEquiv_extendScalarsId (M : ModuleCat R) : (extendRestrictScalarsAdj (RingHom.id R)).homEquiv _ _ ((extendScalarsId R).hom.app M) = (restrictScalarsId R).inv.app M := by ext m - rw [extendRestrictScalarsAdj_homEquiv_apply, ← extendScalarsId_inv_app_apply] - erw [← comp_apply] + rw [extendRestrictScalarsAdj_homEquiv_apply, ← extendScalarsId_inv_app_apply, ← comp_apply] simp lemma extendScalarsId_hom_app_one_tmul (M : ModuleCat R) (m : M) : diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index da141704256b2..33b8dc3cb3281 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -181,7 +181,8 @@ commutative rings. -/ noncomputable def relativeDifferentials' : PresheafOfModules.{u} (R ⋙ forget₂ _ _) where obj X := CommRingCat.KaehlerDifferential (φ'.app X) - map f := CommRingCat.KaehlerDifferential.map (φ'.naturality f) + -- Have to hint `g' := R.map f` below, or it gets unfolded weirdly. + map f := CommRingCat.KaehlerDifferential.map (g' := R.map f) (φ'.naturality f) -- Without `dsimp`, `ext` doesn't pick up the right lemmas. map_id _ := by dsimp; ext; simp map_comp _ _ := by dsimp; ext; simp @@ -192,7 +193,8 @@ 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).hom (CommRingCat.KaehlerDifferential.d x) = + (ModuleCat.Hom.hom (R := ↑(R.obj X)) ((relativeDifferentials' φ').map f)) + (CommRingCat.KaehlerDifferential.d x) = CommRingCat.KaehlerDifferential.d (R.map f x) := CommRingCat.KaehlerDifferential.map_d (φ'.naturality f) _ diff --git a/Mathlib/Algebra/Category/ModuleCat/ExteriorPower.lean b/Mathlib/Algebra/Category/ModuleCat/ExteriorPower.lean new file mode 100644 index 0000000000000..d3152f601c9e4 --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/ExteriorPower.lean @@ -0,0 +1,106 @@ +/- +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.LinearAlgebra.ExteriorPower.Basic +import Mathlib.Algebra.Category.ModuleCat.Basic + +/-! +# The exterior powers as functors on the category of modules + +In this file, given `M : ModuleCat R` and `n : ℕ`, we define `M.exteriorPower n : ModuleCat R`, +and this extends to a functor `ModuleCat.exteriorPower.functor : ModuleCat R ⥤ ModuleCat R`. + +-/ + +universe v u + +open CategoryTheory + +namespace ModuleCat + +variable {R : Type u} [CommRing R] + +/-- The exterior power of an object in `ModuleCat R`. -/ +def exteriorPower (M : ModuleCat.{v} R) (n : ℕ) : ModuleCat.{max u v} R := + ModuleCat.of R (⋀[R]^n M) + +-- this could be an abbrev, but using a def eases automation +/-- The type of `n`-alternating maps on `M : ModuleCat R` to `N : ModuleCat R`. -/ +def AlternatingMap (M : ModuleCat.{v} R) (N : ModuleCat.{max u v} R) (n : ℕ) := + _root_.AlternatingMap R M N (Fin n) + +instance (M : ModuleCat.{v} R) (N : ModuleCat.{max u v} R) (n : ℕ) : + FunLike (M.AlternatingMap N n) (Fin n → M) N := + inferInstanceAs (FunLike (M [⋀^(Fin n)]→ₗ[R] N) (Fin n → M) N) + +namespace AlternatingMap + +variable {M : ModuleCat.{v} R} {N : ModuleCat.{max u v} R} {n : ℕ} + +@[ext] +lemma ext {φ φ' : M.AlternatingMap N n} (h : ∀ (x : Fin n → M), φ x = φ' x ) : + φ = φ' := + _root_.AlternatingMap.ext h + +variable (φ : M.AlternatingMap N n) {N' : ModuleCat.{max u v} R} (g : N ⟶ N') + +/-- The postcomposition of an alternating map by a linear map. -/ +def postcomp : M.AlternatingMap N' n := + g.hom.compAlternatingMap φ + +@[simp] +lemma postcomp_apply (x : Fin n → M) : + φ.postcomp g x = g (φ x) := rfl + +end AlternatingMap + +namespace exteriorPower + +/-- Constructor for elements in `M.exteriorPower n` when `M` is an object of `ModuleCat R` +and `n : ℕ`. -/ +def mk {M : ModuleCat.{v} R} {n : ℕ} : + M.AlternatingMap (M.exteriorPower n) n := + exteriorPower.ιMulti _ _ + +@[ext] +lemma hom_ext {M : ModuleCat.{v} R} {N : ModuleCat.{max u v} R} {n : ℕ} + {f g : M.exteriorPower n ⟶ N} + (h : mk.postcomp f = mk.postcomp g) : f = g := by + ext : 1 + exact exteriorPower.linearMap_ext h + +/-- The morphism `M.exteriorPower n ⟶ N` induced by an alternating map. -/ +noncomputable def desc {M : ModuleCat.{v} R} {n : ℕ} {N : ModuleCat.{max u v} R} + (φ : M.AlternatingMap N n) : M.exteriorPower n ⟶ N := + ofHom (exteriorPower.alternatingMapLinearEquiv φ) + +@[simp] +lemma desc_mk {M : ModuleCat.{v} R} {n : ℕ} {N : ModuleCat.{max u v} R} + (φ : M.AlternatingMap N n) (x : Fin n → M) : + desc φ (mk x) = φ x := by + apply exteriorPower.alternatingMapLinearEquiv_apply_ιMulti + +/-- The morphism `M.exteriorPower n ⟶ N.exteriorPower n` induced by a morphism `M ⟶ N` +in `ModuleCat R`. -/ +noncomputable def map {M N : ModuleCat.{v} R} (f : M ⟶ N) (n : ℕ) : + M.exteriorPower n ⟶ N.exteriorPower n := + ofHom (_root_.exteriorPower.map n f.hom) + +@[simp] +lemma map_mk {M N : ModuleCat.{v} R} (f : M ⟶ N) {n : ℕ} (x : Fin n → M) : + map f n (mk x) = mk (f ∘ x) := by + apply exteriorPower.map_apply_ιMulti + +variable (R) + +/-- The functor `ModuleCat R ⥤ ModuleCat R` which sends a module to its +`n`th exterior power. -/ +noncomputable def functor (n : ℕ) : ModuleCat.{v} R ⥤ ModuleCat.{max u v} R where + obj M := M.exteriorPower n + map f := map f n + +end exteriorPower + +end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index 431ac7e1734a3..81944fb1cb648 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -249,8 +249,7 @@ def directLimitIsColimit : IsColimit (directLimitCocone G f) where rfl fac s i := by ext - dsimp only [directLimitCocone, CategoryStruct.comp] - rw [LinearMap.comp_apply] + dsimp only [hom_comp, directLimitCocone, hom_ofHom, LinearMap.comp_apply] apply DirectLimit.lift_of uniq s m h := by have : diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean index 89bdba339ff9f..1f38171bfc7b0 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Basic.lean @@ -244,7 +244,7 @@ lemma tensorLift_tmul (m : M₁) (n : M₂) : end -lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f (m ⊗ₜ n) = g (m ⊗ₜ n)) : +lemma tensor_ext {f g : M₁ ⊗ M₂ ⟶ M₃} (h : ∀ m n, f.hom (m ⊗ₜ n) = g.hom (m ⊗ₜ n)) : f = g := hom_ext <| TensorProduct.ext (by ext; apply h) diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index a619b7e79f880..3fd18ca3d8336 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -42,12 +42,12 @@ structure PresheafOfModules where /-- the restriction maps of a presheaf of modules -/ map {X Y : Cᵒᵖ} (f : X ⟶ Y) : obj X ⟶ (ModuleCat.restrictScalars (R.map f).hom).obj (obj Y) map_id (X : Cᵒᵖ) : - map (𝟙 X) = - (ModuleCat.restrictScalarsId' _ (congrArg RingCat.Hom.hom (R.map_id X))).inv.app _ := by + map (𝟙 X) = (ModuleCat.restrictScalarsId' (R.map (𝟙 X)).hom + (congrArg RingCat.Hom.hom (R.map_id X))).inv.app _ := by aesop_cat map_comp {X Y Z : Cᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) = map f ≫ (ModuleCat.restrictScalars _).map (map g) ≫ - (ModuleCat.restrictScalarsComp' _ _ _ + (ModuleCat.restrictScalarsComp' (R.map f).hom (R.map g).hom (R.map (f ≫ g)).hom (congrArg RingCat.Hom.hom <| R.map_comp f g)).inv.app _ := by aesop_cat namespace PresheafOfModules diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean index cd94fb824ae02..869be578d46d6 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Monoidal.lean @@ -53,7 +53,8 @@ noncomputable def tensorObj : PresheafOfModules (R ⋙ forget₂ _ _) where map_id X := ModuleCat.MonoidalCategory.tensor_ext (by intro m₁ m₂ dsimp [tensorObjMap] - simp) + simp + rfl) -- `ModuleCat.restrictScalarsId'App_inv_apply` doesn't get picked up due to type mismatch map_comp f g := ModuleCat.MonoidalCategory.tensor_ext (by intro m₁ m₂ dsimp [tensorObjMap] @@ -65,7 +66,8 @@ variable {M₁ M₂ M₃ M₄} lemma tensorObj_map_tmul {X Y : Cᵒᵖ} (f : X ⟶ Y) (m₁ : M₁.obj X) (m₂ : M₂.obj X) : DFunLike.coe (α := (M₁.obj X ⊗ M₂.obj X :)) (β := fun _ ↦ (ModuleCat.restrictScalars (R.map f).hom).obj (M₁.obj Y ⊗ M₂.obj Y)) - ((tensorObj M₁ M₂).map f).hom (m₁ ⊗ₜ[R.obj X] m₂) = M₁.map f m₁ ⊗ₜ[R.obj Y] M₂.map f m₂ := rfl + (ModuleCat.Hom.hom (R := ↑(R.obj X)) ((tensorObj M₁ M₂).map f)) (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] @@ -73,8 +75,11 @@ 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, ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul, - naturality_apply, naturality_apply]) + rw [tensorObj_map_tmul] + -- Need `erw` because of the type mismatch in `map` and the tensor product. + erw [ModuleCat.MonoidalCategory.tensorHom_tmul, tensorObj_map_tmul] + rw [naturality_apply, naturality_apply] + simp) end Monoidal diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean index 62c75c356bf58..80d229fe13f6c 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Pushforward.lean @@ -28,15 +28,16 @@ namespace PresheafOfModules variable (F : C ⥤ D) +set_option maxHeartbeats 400000 in -- Seems to be a timeout in type checking the definition /-- The pushforward functor on presheaves of modules for a functor `F : C ⥤ D` and `R : Dᵒᵖ ⥤ RingCat`. On the underlying presheaves of abelian groups, it is induced by the precomposition with `F.op`. -/ def pushforward₀ (R : Dᵒᵖ ⥤ RingCat.{u}) : PresheafOfModules.{v} R ⥤ PresheafOfModules.{v} (F.op ⋙ R) where obj M := - { 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 + { obj X := ModuleCat.of _ (M.obj (F.op.obj X)) + map {X Y} f := M.map (F.op.map f) + map_id X := by refine ModuleCat.hom_ext -- Work around an instance diamond for `restrictScalarsId'` (@LinearMap.ext _ _ _ _ _ _ _ _ (_) (_) _ _ _ (fun x => ?_)) @@ -46,7 +47,7 @@ def pushforward₀ (R : Dᵒᵖ ⥤ RingCat.{u}) : -- 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 _ } + map {M₁ M₂} φ := { app X := φ.app _ } /-- The pushforward of presheaves of modules commutes with the forgetful functor to presheaves of abelian groups. -/ diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index 4a8d0a5b70086..a1d35a63c69ed 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -106,7 +106,7 @@ lemma isCompatible_map_smul : ((r₀.smul m₀).map (whiskerRight φ (forget _)) RingCat.comp_apply] have hb₀ : (φ.app (Opposite.op Z)) b₀ = (A.map (f₁.op ≫ g₁.op)) m := by dsimp [b₀] - erw [NatTrans.naturality_apply, hb₁, Functor.map_comp, comp_apply] + erw [NatTrans.naturality_apply, hb₁, Functor.map_comp, CategoryTheory.comp_apply] have ha₀' : (α.app (Opposite.op Z)) a₀ = (R.map (f₂.op ≫ g₂.op)) r := by rw [ha₀, ← op_comp, fac, op_comp] have hb₀' : (φ.app (Opposite.op Z)) b₀ = (A.map (f₂.op ≫ g₂.op)) m := by @@ -161,7 +161,7 @@ def SMulCandidate.mk' (S : Sieve X.unop) (hS : S ∈ J X.unop) apply A.isSeparated _ _ (J.pullback_stable f.unop hS) rintro Z g hg dsimp at hg - erw [← comp_apply, ← A.val.map_comp, ← NatTrans.naturality_apply, M₀.map_smul] + erw [← CategoryTheory.comp_apply, ← A.val.map_comp, ← NatTrans.naturality_apply, M₀.map_smul] refine (ha _ hg).trans (app_eq_of_isLocallyInjective α φ A.isSeparated _ _ _ _ ?_ ?_) · rw [← RingCat.comp_apply, NatTrans.naturality, RingCat.comp_apply, ha₀] apply (hr₀ _ hg).symm.trans @@ -298,9 +298,9 @@ lemma map_smul : apply A.isSeparated _ _ hS rintro Y f ⟨⟨r₀, (hr₀ : (α.app (Opposite.op Y)).hom r₀ = (R.val.map f.op).hom ((R.val.map π).hom r))⟩, ⟨m₀, hm₀⟩⟩ - rw [← comp_apply, ← Functor.map_comp, + rw [← CategoryTheory.comp_apply, ← Functor.map_comp, map_smul_eq α φ r m (π ≫ f.op) r₀ (by rw [hr₀, Functor.map_comp, RingCat.comp_apply]) m₀ - (by rw [hm₀, Functor.map_comp, comp_apply]), + (by rw [hm₀, Functor.map_comp, CategoryTheory.comp_apply]), map_smul_eq α φ (R.val.map π r) (A.val.map π m) f.op r₀ hr₀ m₀ hm₀] end Sheafify diff --git a/Mathlib/Algebra/Category/ModuleCat/Products.lean b/Mathlib/Algebra/Category/ModuleCat/Products.lean index 937317e6ea0a5..d4a7f6a1b8fcb 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Products.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Products.lean @@ -86,7 +86,7 @@ def coproductCoconeIsColimit : IsColimit (coproductCocone Z) where ext : 1 refine DirectSum.linearMap_ext _ fun i ↦ ?_ ext x - simpa only [LinearMap.coe_comp, Function.comp_apply, toModule_lof] using + simpa only [LinearMap.coe_comp, Function.comp_apply, hom_ofHom, toModule_lof] using congr($(h ⟨i⟩) x) variable [HasCoproduct Z] diff --git a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean index 5fb8394a9c465..29c919d23199d 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Subobject.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Subobject.lean @@ -59,7 +59,7 @@ noncomputable def subobjectModule : Subobject M ≃o Submodule R M := ofHom (underlyingIso (ofHom N.subtype)).symm.toLinearEquiv.toLinearMap := by ext x rfl - rw [this, hom_comp, LinearEquiv.range_comp] + rw [this, hom_comp, hom_ofHom, 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⟩ diff --git a/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean b/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean index ce3e674738fab..28ed84e06487d 100644 --- a/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean +++ b/Mathlib/Algebra/Category/MonCat/ForgetCorepresentable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sophie Morel -/ import Mathlib.Algebra.Category.MonCat.Basic -import Mathlib.Data.Nat.Cast.Basic +import Mathlib.Algebra.Group.Nat.Hom /-! # The forget functor is corepresentable diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index 6c70b9d5fdd78..8dd6aac24e41d 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -53,18 +53,37 @@ lemma of_carrier (R : SemiRingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `SemiRingCat`. -/ @[ext] -structure Hom (R S : SemiRingCat) where +structure Hom (R S : SemiRingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category SemiRingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : SemiRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} SemiRingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- Turn a morphism in `SemiRingCat` back into a `RingHom`. -/ +abbrev Hom.hom {R S : SemiRingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := SemiRingCat) f + +/-- Typecheck a `RingHom` as a morphism in `SemiRingCat`. -/ +abbrev ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := SemiRingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : SemiRingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : SemiRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -85,10 +104,7 @@ lemma comp_apply {R S T : SemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : SemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `SemiRingCat`. -/ -abbrev ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) : (ofHom f).hom = f := rfl @[simp] @@ -120,12 +136,6 @@ lemma hom_inv_apply {R S : SemiRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) instance : Inhabited SemiRingCat := ⟨of PUnit⟩ -instance : HasForget.{u} SemiRingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. -/ unif_hint forget_obj_eq_coe (R R' : SemiRingCat) where R ≟ R' ⊢ @@ -154,8 +164,8 @@ instance hasForgetToAddCommMonCat : HasForget₂ SemiRingCat AddCommMonCat where @[simps] def _root_.RingEquiv.toSemiRingCatIso {R S : Type u} [Semiring R] [Semiring S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget SemiRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -166,7 +176,7 @@ instance forgetReflectIsos : (forget SemiRingCat).ReflectsIsomorphisms where end SemiRingCat -/-- The category of semirings. -/ +/-- The category of rings. -/ structure RingCat where private mk :: /-- The underlying type. -/ @@ -197,18 +207,37 @@ lemma of_carrier (R : RingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `RingCat`. -/ @[ext] -structure Hom (R S : RingCat) where +structure Hom (R S : RingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category RingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : RingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} RingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- Turn a morphism in `RingCat` back into a `RingHom`. -/ +abbrev Hom.hom {R S : RingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := RingCat) f + +/-- Typecheck a `RingHom` as a morphism in `RingCat`. -/ +abbrev ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := RingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : RingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : RingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -229,10 +258,7 @@ lemma comp_apply {R S T : RingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : RingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `RingCat`. -/ -abbrev ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [Ring R] [Ring S] (f : R →+* S) : (ofHom f).hom = f := rfl @[simp] @@ -264,12 +290,6 @@ lemma hom_inv_apply {R S : RingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) = s instance : Inhabited RingCat := ⟨of PUnit⟩ -instance : HasForget.{u} RingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. An example where this is needed is in applying @@ -302,8 +322,8 @@ instance hasForgetToAddCommGrp : HasForget₂ RingCat AddCommGrp where @[simps] def _root_.RingEquiv.toRingCatIso {R S : Type u} [Ring R] [Ring S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget RingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -314,7 +334,7 @@ instance forgetReflectIsos : (forget RingCat).ReflectsIsomorphisms where end RingCat -/-- The category of semirings. -/ +/-- The category of commutative semirings. -/ structure CommSemiRingCat where private mk :: /-- The underlying type. -/ @@ -345,18 +365,37 @@ lemma of_carrier (R : CommSemiRingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `CommSemiRingCat`. -/ @[ext] -structure Hom (R S : CommSemiRingCat) where +structure Hom (R S : CommSemiRingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category CommSemiRingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : CommSemiRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} CommSemiRingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- Turn a morphism in `CommSemiRingCat` back into a `RingHom`. -/ +abbrev Hom.hom {R S : CommSemiRingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := CommSemiRingCat) f + +/-- Typecheck a `RingHom` as a morphism in `CommSemiRingCat`. -/ +abbrev ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := CommSemiRingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : CommSemiRingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : CommSemiRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -377,10 +416,7 @@ lemma comp_apply {R S T : CommSemiRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : CommSemiRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `CommSemiRingCat`. -/ -abbrev ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) : (ofHom f).hom = f := rfl @@ -413,12 +449,6 @@ lemma hom_inv_apply {R S : CommSemiRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv instance : Inhabited CommSemiRingCat := ⟨of PUnit⟩ -instance : HasForget.{u} CommSemiRingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. -/ unif_hint forget_obj_eq_coe (R R' : CommSemiRingCat) where R ≟ R' ⊢ @@ -449,8 +479,8 @@ instance hasForgetToCommMonCat : HasForget₂ CommSemiRingCat CommMonCat where def _root_.RingEquiv.toCommSemiRingCatIso {R S : Type u} [CommSemiring R] [CommSemiring S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget CommSemiRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -461,7 +491,7 @@ instance forgetReflectIsos : (forget CommSemiRingCat).ReflectsIsomorphisms where end CommSemiRingCat -/-- The category of semirings. -/ +/-- The category of commutative rings. -/ structure CommRingCat where private mk :: /-- The underlying type. -/ @@ -492,18 +522,37 @@ lemma of_carrier (R : CommRingCat.{u}) : of R = R := rfl variable {R} in /-- The type of morphisms in `CommRingCat`. -/ @[ext] -structure Hom (R S : CommRingCat) where +structure Hom (R S : CommRingCat.{u}) where private mk :: /-- The underlying ring hom. -/ - hom : R →+* S + hom' : R →+* S instance : Category CommRingCat where Hom R S := Hom R S id R := ⟨RingHom.id R⟩ - comp f g := ⟨g.hom.comp f.hom⟩ + comp f g := ⟨g.hom'.comp f.hom'⟩ -instance {R S : CommRingCat.{u}} : CoeFun (R ⟶ S) (fun _ ↦ R → S) where - coe f := f.hom +instance : ConcreteCategory.{u} CommRingCat (fun R S => R →+* S) where + hom := Hom.hom' + ofHom f := ⟨f⟩ + +/-- The underlying ring hom. -/ +abbrev Hom.hom {R S : CommRingCat.{u}} (f : Hom R S) := + ConcreteCategory.hom (C := CommRingCat) f + +/-- Typecheck a `RingHom` as a morphism in `CommRingCat`. -/ +abbrev ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := + ConcreteCategory.ofHom (C := CommRingCat) f + +/-- Use the `ConcreteCategory.hom` projection for `@[simps]` lemmas. -/ +def Hom.Simps.hom (R S : CommRingCat) (f : Hom R S) := + f.hom + +initialize_simps_projections Hom (hom' → hom) + +/-! +The results below duplicate the `ConcreteCategory` simp lemmas, but we can keep them for `dsimp`. +-/ @[simp] lemma hom_id {R : CommRingCat} : (𝟙 R : R ⟶ R).hom = RingHom.id R := rfl @@ -524,10 +573,7 @@ lemma comp_apply {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) (r : R) : lemma hom_ext {R S : CommRingCat} {f g : R ⟶ S} (hf : f.hom = g.hom) : f = g := Hom.ext hf -/-- Typecheck a `RingHom` as a morphism in `CommRingCat`. -/ -abbrev ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : of R ⟶ of S := - ⟨f⟩ - +@[simp] lemma hom_ofHom {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) : (ofHom f).hom = f := rfl @@ -560,12 +606,6 @@ lemma hom_inv_apply {R S : CommRingCat} (e : R ≅ S) (s : S) : e.hom (e.inv s) instance : Inhabited CommRingCat := ⟨of PUnit⟩ -instance : HasForget.{u} CommRingCat where - forget := - { obj := fun R => R - map := fun f => f.hom } - forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ - lemma forget_obj {R : CommRingCat} : (forget CommRingCat).obj R = R := rfl /-- This unification hint helps with problems of the form `(forget ?C).obj R =?= carrier R'`. @@ -606,8 +646,8 @@ instance hasForgetToAddCommMonCat : HasForget₂ CommRingCat CommSemiRingCat whe def _root_.RingEquiv.toCommRingCatIso {R S : Type u} [CommRing R] [CommRing S] (e : R ≃+* S) : of R ≅ of S where - hom := ⟨e⟩ - inv := ⟨e.symm⟩ + hom := ofHom e + inv := ofHom e.symm instance forgetReflectIsos : (forget CommRingCat).ReflectsIsomorphisms where reflects {X Y} f _ := by @@ -669,10 +709,10 @@ abbrev CommRingCatMax.{u1, u2} := CommRingCat.{max u1 u2} lemma RingCat.forget_map_apply {R S : RingCat} (f : R ⟶ S) (x : (CategoryTheory.forget RingCat).obj R) : - @DFunLike.coe _ _ _ HasForget.instFunLike f x = f x := + (forget _).map f x = f x := rfl lemma CommRingCat.forget_map_apply {R S : CommRingCat} (f : R ⟶ S) (x : (CategoryTheory.forget CommRingCat).obj R) : - @DFunLike.coe _ _ _ HasForget.instFunLike f x = f x := + (forget _).map f x = f x := rfl diff --git a/Mathlib/Algebra/Category/Semigrp/Basic.lean b/Mathlib/Algebra/Category/Semigrp/Basic.lean index f020632a9ccb5..7be4ed672212c 100644 --- a/Mathlib/Algebra/Category/Semigrp/Basic.lean +++ b/Mathlib/Algebra/Category/Semigrp/Basic.lean @@ -100,7 +100,7 @@ def ofHom {X Y : Type u} [Mul X] [Mul Y] (f : X →ₙ* Y) : of X ⟶ of Y := f /-- Typecheck an `AddHom` as a morphism in `AddMagmaCat`. -/ add_decl_doc AddMagmaCat.ofHom -@[to_additive] -- Porting note: simp removed, simpNF says LHS simplifies to itself +@[to_additive] theorem ofHom_apply {X Y : Type u} [Mul X] [Mul Y] (f : X →ₙ* Y) (x : X) : ofHom f x = f x := rfl @@ -180,7 +180,7 @@ def ofHom {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) : of X /-- Typecheck an `AddHom` as a morphism in `AddSemigrp`. -/ add_decl_doc AddSemigrp.ofHom -@[to_additive] -- Porting note: simp removed, simpNF says LHS simplifies to itself +@[to_additive] theorem ofHom_apply {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) (x : X) : ofHom f x = f x := rfl diff --git a/Mathlib/Algebra/CharP/Lemmas.lean b/Mathlib/Algebra/CharP/Lemmas.lean index f5182dddfbd03..a1a94ee942c82 100644 --- a/Mathlib/Algebra/CharP/Lemmas.lean +++ b/Mathlib/Algebra/CharP/Lemmas.lean @@ -273,6 +273,10 @@ variable (R) [Ring R] [NoZeroDivisors R] [Nontrivial R] [Finite R] theorem char_is_prime (p : ℕ) [CharP R p] : p.Prime := Or.resolve_right (char_is_prime_or_zero R p) (char_ne_zero_of_finite R p) +lemma prime_ringChar : Nat.Prime (ringChar R) := by + apply CharP.char_prime_of_ne_zero R + exact CharP.ringChar_ne_zero_of_finite R + end Ring end CharP diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean index 079913152e68b..476ff787c924d 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean @@ -5,7 +5,6 @@ Authors: Kevin Kappelmann -/ import Mathlib.Algebra.ContinuedFractions.Determinant import Mathlib.Algebra.ContinuedFractions.Computation.CorrectnessTerminating -import Mathlib.Algebra.Order.Group.Basic import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Data.Nat.Fib.Basic import Mathlib.Tactic.Monotonicity diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean b/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean index 2a78678ad1a4a..b7f727a629aae 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean @@ -213,10 +213,7 @@ theorem coe_of_rat_eq (v_eq_q : v = (↑q : K)) : rcases gcf_v_eq : of v with ⟨h, s⟩; subst v -- Porting note: made coercion target explicit obtain rfl : ↑⌊(q : K)⌋ = h := by injection gcf_v_eq - -- Porting note: was - -- simp [coe_of_h_rat_eq rfl, coe_of_s_rat_eq rfl, gcf_v_eq] - simp only [gcf_v_eq, Int.cast_inj, Rat.floor_cast, of_h_eq_floor, eq_self_iff_true, - Rat.cast_intCast, and_self, coe_of_h_rat_eq rfl, coe_of_s_rat_eq rfl] + simp [coe_of_h_rat_eq rfl, coe_of_s_rat_eq rfl, gcf_v_eq] theorem of_terminates_iff_of_rat_terminates {v : K} {q : ℚ} (v_eq_q : v = (q : K)) : (of v).Terminates ↔ (of q).Terminates := by diff --git a/Mathlib/Algebra/DirectSum/Idempotents.lean b/Mathlib/Algebra/DirectSum/Idempotents.lean index bb76018d54b2c..e4a61553e8176 100644 --- a/Mathlib/Algebra/DirectSum/Idempotents.lean +++ b/Mathlib/Algebra/DirectSum/Idempotents.lean @@ -7,9 +7,9 @@ import Mathlib.RingTheory.Idempotents import Mathlib.Algebra.DirectSum.Decomposition /-! -# Decomposition of the identity of a ring into orthogonal idempotents +# Decomposition of the identity of a semiring into orthogonal idempotents -In this file we show that if a ring `R` can be decomposed into a direct sum +In this file we show that if a semiring `R` can be decomposed into a direct sum of (left) ideals `R = V₁ ⊕ V₂ ⊕ ⬝ ⬝ ⬝ ⊕ Vₙ` then in the corresponding decomposition `1 = e₁ + e₂ + ⬝ ⬝ ⬝ + eₙ` with `eᵢ ∈ Vᵢ`, each `eᵢ` is an idempotent and the `eᵢ`'s form a family of complete orthogonal idempotents. @@ -19,26 +19,25 @@ namespace DirectSum section OrthogonalIdempotents -variable {R I : Type*} [Ring R] [DecidableEq I] (V : I → Ideal R) [Decomposition V] +variable {R I : Type*} [Semiring R] [DecidableEq I] (V : I → Ideal R) [Decomposition V] /-- The decomposition of `(1 : R)` where `1 = e₁ + e₂ + ⬝ ⬝ ⬝ + eₙ` which is induced by - the decomposition of the ring `R = V1 ⊕ V2 ⊕ ⬝ ⬝ ⬝ ⊕ Vn`.-/ + the decomposition of the semiring `R = V1 ⊕ V2 ⊕ ⬝ ⬝ ⬝ ⊕ Vn`.-/ def idempotent (i : I) : R := decompose V 1 i -lemma decompose_eq_mul_idempotent - (x : R) (i : I) : decompose V x i = x * idempotent V i := by +lemma decompose_eq_mul_idempotent (x : R) (i : I) : decompose V x i = x * idempotent V i := by rw [← smul_eq_mul (a := x), idempotent, ← Submodule.coe_smul, ← smul_apply, ← decompose_smul, smul_eq_mul, mul_one] lemma isIdempotentElem_idempotent (i : I) : IsIdempotentElem (idempotent V i : R) := by rw [IsIdempotentElem, ← decompose_eq_mul_idempotent, idempotent, decompose_coe, of_eq_same] -/-- If a ring can be decomposed into direct sum of finite left ideals `Vᵢ` +/-- If a semiring can be decomposed into direct sum of finite left ideals `Vᵢ` where `1 = e₁ + ... + eₙ` and `eᵢ ∈ Vᵢ`, then `eᵢ` is a family of complete orthogonal idempotents.-/ theorem completeOrthogonalIdempotents_idempotent [Fintype I]: - CompleteOrthogonalIdempotents fun i ↦ idempotent V i where + CompleteOrthogonalIdempotents (idempotent V) where idem := isIdempotentElem_idempotent V ortho i j hij := by simp only diff --git a/Mathlib/Algebra/Equiv/TransferInstance.lean b/Mathlib/Algebra/Equiv/TransferInstance.lean index 4555c2270dd82..13b89b97205f4 100644 --- a/Mathlib/Algebra/Equiv/TransferInstance.lean +++ b/Mathlib/Algebra/Equiv/TransferInstance.lean @@ -205,7 +205,6 @@ protected abbrev semigroupWithZero [SemigroupWithZero β] : SemigroupWithZero α let zero := e.zero apply e.injective.semigroupWithZero _ <;> intros <;> exact e.apply_symm_apply _ -@[to_additive] noncomputable instance [Small.{v} α] [SemigroupWithZero α] : SemigroupWithZero (Shrink.{v} α) := (equivShrink α).symm.semigroupWithZero diff --git a/Mathlib/Algebra/Field/Action/ConjAct.lean b/Mathlib/Algebra/Field/Action/ConjAct.lean new file mode 100644 index 0000000000000..f55dc6893ea4d --- /dev/null +++ b/Mathlib/Algebra/Field/Action/ConjAct.lean @@ -0,0 +1,22 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.Field.Defs +import Mathlib.Algebra.GroupWithZero.Action.ConjAct + +/-! +# Conjugation action on a field on itself +-/ + +namespace ConjAct + +variable {K : Type*} [DivisionRing K] + +instance distribMulAction₀ : DistribMulAction (ConjAct K) K := + { ConjAct.mulAction₀ with + smul_zero := by simp [smul_def] + smul_add := by simp [smul_def, mul_add, add_mul] } + +end ConjAct diff --git a/Mathlib/Algebra/Field/Power.lean b/Mathlib/Algebra/Field/Power.lean index 8340559169683..a295fde38a1ad 100644 --- a/Mathlib/Algebra/Field/Power.lean +++ b/Mathlib/Algebra/Field/Power.lean @@ -22,7 +22,7 @@ section DivisionRing variable [DivisionRing α] {n : ℤ} theorem Odd.neg_zpow (h : Odd n) (a : α) : (-a) ^ n = -a ^ n := by - have hn : n ≠ 0 := by rintro rfl; exact Int.not_even_iff_odd.2 h even_zero + have hn : n ≠ 0 := by rintro rfl; exact Int.not_even_iff_odd.2 h .zero obtain ⟨k, rfl⟩ := h simp_rw [zpow_add' (.inr (.inl hn)), zpow_one, zpow_mul, zpow_two, neg_mul_neg, neg_mul_eq_mul_neg] diff --git a/Mathlib/Algebra/Field/Subfield/Defs.lean b/Mathlib/Algebra/Field/Subfield/Defs.lean index 9af131cc79a95..f7955cd9e0f5e 100644 --- a/Mathlib/Algebra/Field/Subfield/Defs.lean +++ b/Mathlib/Algebra/Field/Subfield/Defs.lean @@ -107,7 +107,7 @@ instance instSMulRat (s : S) : SMul ℚ s where smul q x := ⟨q • x, qsmul_me variable (S) /-- A subfield inherits a division ring structure -/ -instance (priority := 75) toDivisionRing (s : S) : DivisionRing s := +instance (priority := 75) toDivisionRing (s : S) : DivisionRing s := fast_instance% Subtype.coe_injective.divisionRing ((↑) : s → K) rfl rfl (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) @@ -117,7 +117,7 @@ instance (priority := 75) toDivisionRing (s : S) : DivisionRing s := -- Prefer subclasses of `Field` over subclasses of `SubfieldClass`. /-- A subfield of a field inherits a field structure -/ instance (priority := 75) toField {K} [Field K] [SetLike S K] [SubfieldClass S K] (s : S) : - Field s := + Field s := fast_instance% Subtype.coe_injective.field ((↑) : s → K) rfl rfl (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) @@ -276,14 +276,14 @@ instance : Pow s ℤ := ⟨fun x z => ⟨x ^ z, s.zpow_mem x.2 z⟩⟩ -- TODO: Those are just special cases of `SubfieldClass.toDivisionRing`/`SubfieldClass.toField` -instance toDivisionRing (s : Subfield K) : DivisionRing s := +instance toDivisionRing (s : Subfield K) : DivisionRing s := fast_instance% Subtype.coe_injective.divisionRing ((↑) : s → K) rfl rfl (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) (fun _ ↦ rfl) fun _ ↦ rfl /-- A subfield inherits a field structure -/ -instance toField {K} [Field K] (s : Subfield K) : Field s := +instance toField {K} [Field K] (s : Subfield K) : Field s := fast_instance% Subtype.coe_injective.field ((↑) : s → K) rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ ↦ rfl) (fun _ => rfl) diff --git a/Mathlib/Algebra/Free.lean b/Mathlib/Algebra/Free.lean index 50961c7a71981..c0663264e177d 100644 --- a/Mathlib/Algebra/Free.lean +++ b/Mathlib/Algebra/Free.lean @@ -28,6 +28,10 @@ import Mathlib.Tactic.AdaptationNote universe u v l +-- Disable generation of `sizeOf_spec` and `injEq`, +-- which are not needed and the `simpNF` linter will complain about. +set_option genSizeOfSpec false in +set_option genInjectivity false in /-- Free nonabelian additive magma over a given alphabet. -/ inductive FreeAddMagma (α : Type u) : Type u | of : α → FreeAddMagma α @@ -35,6 +39,10 @@ inductive FreeAddMagma (α : Type u) : Type u deriving DecidableEq compile_inductive% FreeAddMagma +-- Disable generation of `sizeOf_spec` and `injEq`, +-- which are not needed and the `simpNF` linter will complain about. +set_option genSizeOfSpec false in +set_option genInjectivity false in /-- Free magma over a given alphabet. -/ @[to_additive] inductive FreeMagma (α : Type u) : Type u @@ -59,13 +67,6 @@ instance : Mul (FreeMagma α) := ⟨FreeMagma.mul⟩ @[to_additive (attr := simp)] theorem mul_eq (x y : FreeMagma α) : mul x y = x * y := rfl -/-! These lemmas are autogenerated by the inductive definition and due to -the existence of `mul_eq` not in simp normal form -/ -attribute [nolint simpNF] FreeAddMagma.add.sizeOf_spec -attribute [nolint simpNF] FreeMagma.mul.sizeOf_spec -attribute [nolint simpNF] FreeAddMagma.add.injEq -attribute [nolint simpNF] FreeMagma.mul.injEq - /-- Recursor for `FreeMagma` using `x * y` instead of `FreeMagma.mul x y`. -/ @[to_additive (attr := elab_as_elim, induction_eliminator) "Recursor for `FreeAddMagma` using `x + y` instead of `FreeAddMagma.add x y`."] diff --git a/Mathlib/Algebra/FreeMonoid/Basic.lean b/Mathlib/Algebra/FreeMonoid/Basic.lean index 2c1e153eed82f..7340041857cfc 100644 --- a/Mathlib/Algebra/FreeMonoid/Basic.lean +++ b/Mathlib/Algebra/FreeMonoid/Basic.lean @@ -6,6 +6,7 @@ Authors: Simon Hudon, Yury Kudryashov import Mathlib.Algebra.Group.Action.Defs import Mathlib.Algebra.Group.Units.Defs import Mathlib.Algebra.BigOperators.Group.List.Basic +import Mathlib.Algebra.Group.Equiv.Defs /-! # Free monoid over a given alphabet @@ -121,7 +122,7 @@ variable {a : FreeMonoid α} /-- The length of a free monoid element: 1.length = 0 and (a * b).length = a.length + b.length -/ @[to_additive "The length of an additive free monoid element: 1.length = 0 and (a + b).length = a.length + b.length"] -def length (a : FreeMonoid α) : ℕ := List.length a +def length (a : FreeMonoid α) : ℕ := a.toList.length @[to_additive (attr := simp)] theorem length_one : length (1 : FreeMonoid α) = 0 := rfl @@ -358,6 +359,14 @@ theorem map_comp (g : β → γ) (f : α → β) : map (g ∘ f) = (map g).comp @[to_additive (attr := simp)] theorem map_id : map (@id α) = MonoidHom.id (FreeMonoid α) := hom_eq fun _ ↦ rfl +@[to_additive (attr := simp)] +theorem map_symm_apply_map_eq {x : FreeMonoid α} (e : α ≃ β) : + (map ⇑e.symm) ((map ⇑e) x) = x := by simp [map_map] + +@[to_additive (attr := simp)] +theorem map_apply_map_symm_eq {x : FreeMonoid β} (e : α ≃ β) : + (map ⇑e) ((map ⇑e.symm) x) = x := by simp [map_map] + /-- The only invertible element of the free monoid is 1; this instance enables `units_eq_one`. -/ @[to_additive] instance uniqueUnits : Unique (FreeMonoid α)ˣ where @@ -415,4 +424,27 @@ theorem length_reverse {a : FreeMonoid α} : a.reverse.length = a.length := end Reverse +section IsomorphicTypes + +variable {α β : Type*} + +/-- free monoids over isomorphic types are isomorphic -/ +@[to_additive "if two types are isomorphic, the additive free monoids over those types are +isomorphic"] +def freeMonoidCongr (e : α ≃ β) : FreeMonoid α ≃* FreeMonoid β where + toFun := FreeMonoid.map ⇑e + invFun := FreeMonoid.map ⇑e.symm + left_inv := fun _ => map_symm_apply_map_eq e + right_inv := fun _ => map_apply_map_symm_eq e + map_mul' := (by simp [map_mul]) + +@[to_additive (attr := simp)] +theorem freeMonoidCongr_of (e : α ≃ β) (a : α) : freeMonoidCongr e (of a) = of (e a) := rfl + +@[to_additive (attr := simp)] +theorem freeMonoidCongr_symm_of (e : α ≃ β) (b : β) : + freeMonoidCongr e.symm (of b) = of (e.symm b) := rfl + +end IsomorphicTypes + end FreeMonoid diff --git a/Mathlib/Algebra/FreeMonoid/Count.lean b/Mathlib/Algebra/FreeMonoid/Count.lean index d9add268027ba..82e0442e18495 100644 --- a/Mathlib/Algebra/FreeMonoid/Count.lean +++ b/Mathlib/Algebra/FreeMonoid/Count.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.FreeMonoid.Basic -import Mathlib.Algebra.Group.TypeTags.Hom +import Mathlib.Algebra.Group.TypeTags.Basic /-! # `List.count` as a bundled homomorphism @@ -13,73 +13,76 @@ In this file we define `FreeMonoid.countP`, `FreeMonoid.count`, `FreeAddMonoid.c `FreeAddMonoid.count`. These are `List.countP` and `List.count` bundled as multiplicative and additive homomorphisms from `FreeMonoid` and `FreeAddMonoid`. -We do not use `to_additive` because it can't map `Multiplicative ℕ` to `ℕ`. +We do not use `to_additive` too much because it can't map `Multiplicative ℕ` to `ℕ`. +-/ -## TODO +variable {α : Type*} (p : α → Prop) [DecidablePred p] -There is lots of defeq abuse here of `FreeAddMonoid α = List α`, e.g. in statements like -``` -theorem countP_apply (l : FreeAddMonoid α) : countP p l = List.countP p l := rfl -``` -This needs cleaning up. +namespace FreeMonoid +/-- `List.countP` lifted to free monoids-/ +@[to_additive "`List.countP` lifted to free additive monoids"] +def countP' (l : FreeMonoid α) : ℕ := l.toList.countP p --/ +@[to_additive] +lemma countP'_one : (1 : FreeMonoid α).countP' p = 0 := rfl -variable {α : Type*} (p : α → Prop) [DecidablePred p] +@[to_additive] +lemma countP'_mul (l₁ l₂ : FreeMonoid α) : (l₁ * l₂).countP' p = l₁.countP' p + l₂.countP' p := by + dsimp [countP'] + simp only [List.countP_append] -namespace FreeAddMonoid +/-- `List.countP` as a bundled multiplicative monoid homomorphism. -/ +def countP : FreeMonoid α →* Multiplicative ℕ where + toFun := .ofAdd ∘ FreeMonoid.countP' p + map_one' := by + simp [countP'_one p] + map_mul' x y := by + simp [countP'_mul p] -/-- `List.countP` as a bundled additive monoid homomorphism. -/ -def countP : FreeAddMonoid α →+ ℕ where - toFun := List.countP p - map_zero' := List.countP_nil _ - map_add' := List.countP_append _ +theorem countP_apply (l : FreeMonoid α) : l.countP p = .ofAdd (l.toList.countP p) := rfl -theorem countP_of (x : α) : countP p (of x) = if p x = true then 1 else 0 := by - change List.countP p [x] = _ - simp [List.countP_cons] +lemma countP_of (x : α) : (of x).countP p = + if p x then Multiplicative.ofAdd 1 else Multiplicative.ofAdd 0 := by + rw [countP_apply, toList_of, List.countP_singleton, apply_ite (Multiplicative.ofAdd)] + simp only [decide_eq_true_eq] -theorem countP_apply (l : FreeAddMonoid α) : countP p l = List.countP p l := rfl /-- `List.count` as a bundled additive monoid homomorphism. -/ --- Porting note: was (x = ·) -def count [DecidableEq α] (x : α) : FreeAddMonoid α →+ ℕ := countP (· = x) - -theorem count_of [DecidableEq α] (x y : α) : count x (of y) = (Pi.single x 1 : α → ℕ) y := by - change List.count x [y] = _ - simp [Pi.single, Function.update, List.count_cons] +def count [DecidableEq α] (x : α) : FreeMonoid α →* Multiplicative ℕ := countP (· = x) -theorem count_apply [DecidableEq α] (x : α) (l : FreeAddMonoid α) : count x l = List.count x l := - rfl +theorem count_apply [DecidableEq α] (x : α) (l : FreeAddMonoid α) : + count x l = Multiplicative.ofAdd (l.toList.count x) := rfl -end FreeAddMonoid +theorem count_of [DecidableEq α] (x y : α) : + count x (of y) = Pi.mulSingle (f := fun _ => Multiplicative ℕ) x (Multiplicative.ofAdd 1) y := + by simp only [count, eq_comm, countP_of, ofAdd_zero, Pi.mulSingle_apply] -namespace FreeMonoid +end FreeMonoid -/-- `List.countP` as a bundled multiplicative monoid homomorphism. -/ -def countP : FreeMonoid α →* Multiplicative ℕ := - AddMonoidHom.toMultiplicative (FreeAddMonoid.countP p) +namespace FreeAddMonoid -theorem countP_of' (x : α) : - countP p (of x) = if p x then Multiplicative.ofAdd 1 else Multiplicative.ofAdd 0 := by - erw [FreeAddMonoid.countP_of] - simp only [eq_iff_iff, iff_true, ofAdd_zero]; rfl +/-- `List.countP` as a bundled additive monoid homomorphism. -/ +def countP : FreeAddMonoid α →+ ℕ where + toFun := FreeAddMonoid.countP' p + map_zero' := countP'_zero p + map_add' := countP'_add p -theorem countP_of (x : α) : countP p (of x) = if p x then Multiplicative.ofAdd 1 else 1 := by - rw [countP_of', ofAdd_zero] +theorem countP_apply (l : FreeAddMonoid α) : l.countP p = l.toList.countP p := rfl --- `rfl` is not transitive -theorem countP_apply (l : FreeAddMonoid α) : countP p l = Multiplicative.ofAdd (List.countP p l) := - rfl +theorem countP_of (x : α) : countP p (of x) = if p x then 1 else 0 := by + rw [countP_apply, toList_of, List.countP_singleton] + simp only [decide_eq_true_eq] /-- `List.count` as a bundled additive monoid homomorphism. -/ -def count [DecidableEq α] (x : α) : FreeMonoid α →* Multiplicative ℕ := countP (· = x) +-- Porting note: was (x = ·) +def count [DecidableEq α] (x : α) : FreeAddMonoid α →+ ℕ := countP (· = x) -theorem count_apply [DecidableEq α] (x : α) (l : FreeAddMonoid α) : - count x l = Multiplicative.ofAdd (List.count x l) := rfl +lemma count_of [DecidableEq α] (x y : α) : count x (of y) = (Pi.single x 1 : α → ℕ) y := by + dsimp [count] + rw [countP_of] + simp [Pi.single, Function.update] -theorem count_of [DecidableEq α] (x y : α) : - count x (of y) = @Pi.mulSingle α (fun _ => Multiplicative ℕ) _ _ x (Multiplicative.ofAdd 1) y := - by simp [count, countP_of, Pi.mulSingle_apply, eq_comm, Bool.beq_eq_decide_eq] +theorem count_apply [DecidableEq α] (x : α) (l : FreeAddMonoid α) : l.count x = l.toList.count x := + rfl -end FreeMonoid +end FreeAddMonoid diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 69be1e10b5d9d..cac128232b4fd 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -487,7 +487,7 @@ theorem geom_sum_alternating_of_le_neg_one [StrictOrderedRing α] (hx : x + 1 if Even n then (∑ i ∈ range n, x ^ i) ≤ 0 else 1 ≤ ∑ i ∈ range n, x ^ i := by have hx0 : x ≤ 0 := (le_add_of_nonneg_right zero_le_one).trans hx induction n with - | zero => simp only [range_zero, sum_empty, le_refl, ite_true, even_zero] + | zero => simp only [range_zero, sum_empty, le_refl, ite_true, Even.zero] | succ n ih => simp only [Nat.even_add_one, geom_sum_succ] split_ifs at ih with h diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index dbfa138e3cb38..6e598f103f8a3 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -16,6 +16,11 @@ This file defines square and even elements in a monoid. * `IsSquare a` means that there is some `r` such that `a = r * r` * `Even a` means that there is some `r` such that `a = r + r` +## Note + +* Many lemmas about `Even` / `IsSquare`, including important `simp` lemmas, + are in `Mathlib.Algebra.Ring.Parity`. + ## TODO * Try to generalize `IsSquare/Even` lemmas further. For example, there are still a few lemmas in @@ -26,7 +31,8 @@ This file defines square and even elements in a monoid. ## See also -`Mathlib.Algebra.Ring.Parity` for the definition of odd elements. +`Mathlib.Algebra.Ring.Parity` for the definition of odd elements as well as facts about +`Even` / `IsSquare` in rings. -/ assert_not_exists MonoidWithZero DenselyOrdered @@ -39,19 +45,19 @@ section Mul variable [Mul α] /-- An element `a` of a type `α` with multiplication satisfies `IsSquare a` if `a = r * r`, -for some `r : α`. -/ +for some root `r : α`. -/ @[to_additive "An element `a` of a type `α` with addition satisfies `Even a` if `a = r + r`, for some `r : α`."] def IsSquare (a : α) : Prop := ∃ r, a = r * r -@[to_additive (attr := simp)] lemma IsSquare.mul_self (m : α) : IsSquare (m * m) := ⟨m, rfl⟩ +@[to_additive (attr := simp)] lemma IsSquare.mul_self (r : α) : IsSquare (r * r) := ⟨r, rfl⟩ @[deprecated (since := "2024-08-27")] alias isSquare_mul_self := IsSquare.mul_self @[deprecated (since := "2024-08-27")] alias even_add_self := Even.add_self @[to_additive] lemma isSquare_op_iff {a : α} : IsSquare (op a) ↔ IsSquare a := - ⟨fun ⟨c, hc⟩ ↦ ⟨unop c, congr_arg unop hc⟩, fun ⟨c, hc⟩ ↦ ⟨op c, congr_arg op hc⟩⟩ + ⟨fun ⟨r, hr⟩ ↦ ⟨unop r, congr_arg unop hr⟩, fun ⟨r, hr⟩ ↦ ⟨op r, congr_arg op hr⟩⟩ @[to_additive] lemma isSquare_unop_iff {a : αᵐᵒᵖ} : IsSquare (unop a) ↔ IsSquare a := isSquare_op_iff.symm @@ -89,58 +95,64 @@ end Add @[to_additive (attr := simp)] lemma IsSquare.one [MulOneClass α] : IsSquare (1 : α) := ⟨1, (mul_one _).symm⟩ -@[to_additive, deprecated (since := "2024-12-27")] alias isSquare_one := IsSquare.one +@[deprecated (since := "2024-12-27")] alias isSquare_one := IsSquare.one +@[deprecated (since := "2024-12-27")] alias even_zero := Even.zero + +section MonoidHom +variable [MulOneClass α] [MulOneClass β] [FunLike F α β] [MonoidHomClass F α β] @[to_additive] -lemma IsSquare.map [MulOneClass α] [MulOneClass β] [FunLike F α β] [MonoidHomClass F α β] - {m : α} (f : F) : - IsSquare m → IsSquare (f m) := by - rintro ⟨m, rfl⟩ - exact ⟨f m, by simp⟩ +lemma IsSquare.map {a : α} (f : F) : IsSquare a → IsSquare (f a) := + fun ⟨r, _⟩ => ⟨f r, by simp_all⟩ + +end MonoidHom section Monoid variable [Monoid α] {n : ℕ} {a : α} @[to_additive even_iff_exists_two_nsmul] -lemma isSquare_iff_exists_sq (m : α) : IsSquare m ↔ ∃ c, m = c ^ 2 := by simp [IsSquare, pow_two] +lemma isSquare_iff_exists_sq (a : α) : IsSquare a ↔ ∃ r, a = r ^ 2 := by simp [IsSquare, pow_two] + +@[to_additive Even.exists_two_nsmul + "Alias of the forwards direction of `even_iff_exists_two_nsmul`."] +alias ⟨IsSquare.exists_sq, _⟩ := isSquare_iff_exists_sq -alias ⟨IsSquare.exists_sq, isSquare_of_exists_sq⟩ := isSquare_iff_exists_sq +-- provable by simp in `Algebra.Ring.Parity` +@[to_additive Even.two_nsmul] +lemma IsSquare.sq (r : α) : IsSquare (r ^ 2) := ⟨r, pow_two _⟩ -attribute [to_additive Even.exists_two_nsmul - "Alias of the forwards direction of `even_iff_exists_two_nsmul`."] IsSquare.exists_sq +@[deprecated (since := "2024-12-27")] alias IsSquare_sq := IsSquare.sq +@[deprecated (since := "2024-12-27")] alias even_two_nsmul := Even.two_nsmul -@[to_additive] lemma IsSquare.pow (n : ℕ) : IsSquare a → IsSquare (a ^ n) := by - rintro ⟨a, rfl⟩; exact ⟨a ^ n, (Commute.refl _).mul_pow _⟩ +@[to_additive Even.nsmul_right] lemma IsSquare.pow (n : ℕ) : IsSquare a → IsSquare (a ^ n) := by + rintro ⟨r, rfl⟩; exact ⟨r ^ n, (Commute.refl _).mul_pow _⟩ -@[to_additive Even.nsmul'] lemma Even.isSquare_pow : Even n → ∀ a : α, IsSquare (a ^ n) := by - rintro ⟨n, rfl⟩ a; exact ⟨a ^ n, pow_add _ _ _⟩ +@[deprecated (since := "2025-01-19")] alias Even.nsmul := Even.nsmul_right -@[to_additive Even.two_nsmul] lemma IsSquare.sq (a : α) : IsSquare (a ^ 2) := ⟨a, pow_two _⟩ +@[to_additive (attr := simp) Even.nsmul_left] +lemma Even.isSquare_pow : Even n → ∀ a : α, IsSquare (a ^ n) := by + rintro ⟨m, rfl⟩ a; exact ⟨a ^ m, pow_add _ _ _⟩ -@[deprecated (since := "2024-12-27")] alias IsSquare_sq := IsSquare.sq -@[deprecated (since := "2024-12-27")] alias even_two_nsmul := Even.two_nsmul +@[deprecated (since := "2025-01-19")] alias Even.nsmul' := Even.nsmul_left end Monoid @[to_additive] lemma IsSquare.mul [CommSemigroup α] {a b : α} : IsSquare a → IsSquare b → IsSquare (a * b) := by - rintro ⟨a, rfl⟩ ⟨b, rfl⟩; exact ⟨a * b, mul_mul_mul_comm _ _ _ _⟩ + rintro ⟨r, rfl⟩ ⟨s, rfl⟩; exact ⟨r * s, mul_mul_mul_comm _ _ _ _⟩ section DivisionMonoid variable [DivisionMonoid α] {a : α} @[to_additive (attr := simp)] lemma isSquare_inv : IsSquare a⁻¹ ↔ IsSquare a := by - constructor <;> intro h - · rw [← isSquare_op_iff, ← inv_inv a] - exact h.map (MulEquiv.inv' α) - · exact (isSquare_op_iff.mpr h).map (MulEquiv.inv' α).symm + constructor <;> intro h <;> simpa using (isSquare_op_iff.mpr h).map (MulEquiv.inv' α).symm -alias ⟨_, IsSquare.inv⟩ := isSquare_inv +@[to_additive] alias ⟨_, IsSquare.inv⟩ := isSquare_inv -attribute [to_additive] IsSquare.inv +@[to_additive Even.zsmul_right] lemma IsSquare.zpow (n : ℤ) : IsSquare a → IsSquare (a ^ n) := by + rintro ⟨r, rfl⟩; exact ⟨r ^ n, (Commute.refl _).mul_zpow _⟩ -@[to_additive] lemma IsSquare.zpow (n : ℤ) : IsSquare a → IsSquare (a ^ n) := by - rintro ⟨a, rfl⟩; exact ⟨a ^ n, (Commute.refl _).mul_zpow _⟩ +@[deprecated (since := "2024-01-19")] alias Even.zsmul := Even.zsmul_right end DivisionMonoid @@ -148,6 +160,8 @@ end DivisionMonoid lemma IsSquare.div [DivisionCommMonoid α] {a b : α} (ha : IsSquare a) (hb : IsSquare b) : IsSquare (a / b) := by rw [div_eq_mul_inv]; exact ha.mul hb.inv -@[to_additive (attr := simp) Even.zsmul'] +@[to_additive (attr := simp) Even.zsmul_left] lemma Even.isSquare_zpow [Group α] {n : ℤ} : Even n → ∀ a : α, IsSquare (a ^ n) := by - rintro ⟨n, rfl⟩ a; exact ⟨a ^ n, zpow_add _ _ _⟩ + rintro ⟨m, rfl⟩ a; exact ⟨a ^ m, zpow_add _ _ _⟩ + +@[deprecated (since := "2024-01-07")] alias Even.zsmul' := Even.zsmul_left diff --git a/Mathlib/Algebra/Group/Fin/Basic.lean b/Mathlib/Algebra/Group/Fin/Basic.lean index 5a34b7573582a..c7fb2d5f230d9 100644 --- a/Mathlib/Algebra/Group/Fin/Basic.lean +++ b/Mathlib/Algebra/Group/Fin/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.NeZero import Mathlib.Data.Nat.Cast.Defs import Mathlib.Data.Nat.Defs -import Mathlib.Data.Fin.Basic +import Mathlib.Data.Fin.Rev /-! # Fin is a group diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index 7231ae1d40d90..2ffc6457f0b71 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -884,17 +884,23 @@ namespace Monoid variable (M) [MulOneClass M] /-- The monoid of endomorphisms. -/ +@[to_additive "The monoid of endomorphisms.", to_additive_dont_translate] protected def End := M →* M namespace End +@[to_additive] instance instFunLike : FunLike (Monoid.End M) M M := MonoidHom.instFunLike +@[to_additive] instance instMonoidHomClass : MonoidHomClass (Monoid.End M) M M := MonoidHom.instMonoidHomClass +@[to_additive instOne] instance instOne : One (Monoid.End M) where one := .id _ +@[to_additive instMul] instance instMul : Mul (Monoid.End M) where mul := .comp -instance : Monoid (Monoid.End M) where +@[to_additive instMonoid] +instance instMonoid : Monoid (Monoid.End M) where mul := MonoidHom.comp one := MonoidHom.id M mul_assoc _ _ _ := MonoidHom.comp_assoc _ _ _ @@ -903,53 +909,24 @@ instance : Monoid (Monoid.End M) where npow n f := (npowRec n f).copy f^[n] <| by induction n <;> simp [npowRec, *] <;> rfl npow_succ _ _ := DFunLike.coe_injective <| Function.iterate_succ _ _ +@[to_additive] instance : Inhabited (Monoid.End M) := ⟨1⟩ -@[simp, norm_cast] lemma coe_pow (f : Monoid.End M) (n : ℕ) : (↑(f ^ n) : M → M) = f^[n] := rfl +@[to_additive (attr := simp, norm_cast) coe_pow] +lemma coe_pow (f : Monoid.End M) (n : ℕ) : (↑(f ^ n) : M → M) = f^[n] := rfl -end End - -@[simp] +@[to_additive (attr := simp) coe_one] theorem coe_one : ((1 : Monoid.End M) : M → M) = id := rfl -@[simp] +@[to_additive (attr := simp) coe_mul] theorem coe_mul (f g) : ((f * g : Monoid.End M) : M → M) = f ∘ g := rfl -end Monoid - -namespace AddMonoid - -variable (A : Type*) [AddZeroClass A] - -/-- The monoid of endomorphisms. -/ -protected def End := A →+ A - -namespace End - -instance instFunLike : FunLike (AddMonoid.End A) A A := AddMonoidHom.instFunLike -instance instAddMonoidHomClass : AddMonoidHomClass (AddMonoid.End A) A A := - AddMonoidHom.instAddMonoidHomClass - -instance instOne : One (AddMonoid.End A) where one := .id _ -instance instMul : Mul (AddMonoid.End A) where mul := .comp - -@[simp, norm_cast] lemma coe_one : ((1 : AddMonoid.End A) : A → A) = id := rfl - -@[simp, norm_cast] lemma coe_mul (f g : AddMonoid.End A) : (f * g : A → A) = f ∘ g := rfl - -instance monoid : Monoid (AddMonoid.End A) where - mul_assoc _ _ _ := AddMonoidHom.comp_assoc _ _ _ - mul_one := AddMonoidHom.comp_id - one_mul := AddMonoidHom.id_comp - npow n f := (npowRec n f).copy (Nat.iterate f n) <| by induction n <;> simp [npowRec, *] <;> rfl - npow_succ _ _ := DFunLike.coe_injective <| Function.iterate_succ _ _ - -@[simp, norm_cast] lemma coe_pow (f : AddMonoid.End A) (n : ℕ) : (↑(f ^ n) : A → A) = f^[n] := rfl +end End -instance : Inhabited (AddMonoid.End A) := ⟨1⟩ +@[deprecated (since := "2024-11-20")] protected alias coe_one := End.coe_one +@[deprecated (since := "2024-11-20")] protected alias coe_mul := End.coe_mul -end End -end AddMonoid +end Monoid end End diff --git a/Mathlib/Algebra/Group/Hom/End.lean b/Mathlib/Algebra/Group/Hom/End.lean index f3538c92cd2d9..faba093247ff3 100644 --- a/Mathlib/Algebra/Group/Hom/End.lean +++ b/Mathlib/Algebra/Group/Hom/End.lean @@ -38,7 +38,9 @@ lemma natCast_apply [AddCommMonoid M] (n : ℕ) (m : M) : (↑n : AddMonoid.End (ofNat(n) : AddMonoid.End M) m = n • m := rfl instance instSemiring [AddCommMonoid M] : Semiring (AddMonoid.End M) := - { AddMonoid.End.monoid M, AddMonoidHom.addCommMonoid, AddMonoid.End.instAddMonoidWithOne M with + { AddMonoid.End.instMonoid M, + AddMonoidHom.instAddCommMonoid, + AddMonoid.End.instAddMonoidWithOne M with zero_mul := fun _ => AddMonoidHom.ext fun _ => rfl, mul_zero := fun _ => AddMonoidHom.ext fun _ => AddMonoidHom.map_zero _, left_distrib := fun _ _ _ => AddMonoidHom.ext fun _ => AddMonoidHom.map_add _ _ _, diff --git a/Mathlib/Algebra/Group/Hom/Instances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean index a74b5cfdf1412..eb5b76b157ba5 100644 --- a/Mathlib/Algebra/Group/Hom/Instances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -27,7 +27,7 @@ variable {M : Type uM} {N : Type uN} {P : Type uP} {Q : Type uQ} /-- `(M →* N)` is a `CommMonoid` if `N` is commutative. -/ @[to_additive "`(M →+ N)` is an `AddCommMonoid` if `N` is commutative."] -instance MonoidHom.commMonoid [MulOneClass M] [CommMonoid N] : +instance MonoidHom.instCommMonoid [MulOneClass M] [CommMonoid N] : CommMonoid (M →* N) where mul := (· * ·) mul_assoc := by intros; ext; apply mul_assoc @@ -47,8 +47,8 @@ instance MonoidHom.commMonoid [MulOneClass M] [CommMonoid N] : /-- If `G` is a commutative group, then `M →* G` is a commutative group too. -/ @[to_additive "If `G` is an additive commutative group, then `M →+ G` is an additive commutative group too."] -instance MonoidHom.commGroup {M G} [MulOneClass M] [CommGroup G] : CommGroup (M →* G) := - { MonoidHom.commMonoid with +instance MonoidHom.instCommGroup {M G} [MulOneClass M] [CommGroup G] : CommGroup (M →* G) := + { MonoidHom.instCommMonoid with inv := Inv.inv, div := Div.div, div_eq_mul_inv := by @@ -71,18 +71,18 @@ instance MonoidHom.commGroup {M G} [MulOneClass M] [CommGroup G] : CommGroup (M simp [zpow_natCast, -Int.natCast_add] } instance AddMonoid.End.instAddCommMonoid [AddCommMonoid M] : AddCommMonoid (AddMonoid.End M) := - AddMonoidHom.addCommMonoid + AddMonoidHom.instAddCommMonoid @[simp] theorem AddMonoid.End.zero_apply [AddCommMonoid M] (m : M) : (0 : AddMonoid.End M) m = 0 := rfl --- Note: `@[simp]` omitted because `(1 : AddMonoid.End M) = id` by `AddMonoid.coe_one` +-- Note: `@[simp]` omitted because `(1 : AddMonoid.End M) = id` by `AddMonoid.End.coe_one` theorem AddMonoid.End.one_apply [AddCommMonoid M] (m : M) : (1 : AddMonoid.End M) m = m := rfl instance AddMonoid.End.instAddCommGroup [AddCommGroup M] : AddCommGroup (AddMonoid.End M) := - AddMonoidHom.addCommGroup + AddMonoidHom.instAddCommGroup instance AddMonoid.End.instIntCast [AddCommGroup M] : IntCast (AddMonoid.End M) := { intCast := fun z => z • (1 : AddMonoid.End M) } diff --git a/Mathlib/Algebra/Group/Nat/Even.lean b/Mathlib/Algebra/Group/Nat/Even.lean index 0cedd8a212532..923b311dee4d1 100644 --- a/Mathlib/Algebra/Group/Nat/Even.lean +++ b/Mathlib/Algebra/Group/Nat/Even.lean @@ -72,7 +72,7 @@ lemma even_pow' (h : n ≠ 0) : Even (m ^ n) ↔ Even m := even_pow.trans <| and lemma even_mul_succ_self (n : ℕ) : Even (n * (n + 1)) := by rw [even_mul, even_add_one]; exact em _ lemma even_mul_pred_self : ∀ n : ℕ, Even (n * (n - 1)) - | 0 => even_zero + | 0 => .zero | (n + 1) => mul_comm (n + 1 - 1) (n + 1) ▸ even_mul_succ_self n lemma two_mul_div_two_of_even : Even n → 2 * (n / 2) = n := fun h ↦ diff --git a/Mathlib/Algebra/Group/Nat/Hom.lean b/Mathlib/Algebra/Group/Nat/Hom.lean new file mode 100644 index 0000000000000..d54b5ae8d598a --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/Hom.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Equiv.Defs +import Mathlib.Algebra.Group.Hom.Basic +import Mathlib.Algebra.Group.Nat.Defs +import Mathlib.Algebra.Group.TypeTags.Hom +import Mathlib.Tactic.MinImports + +/-! +# Extensionality of monoid homs from `ℕ` +-/ + +assert_not_exists OrderedCommMonoid MonoidWithZero + +open Additive Multiplicative + +variable {M M : Type*} + +section AddMonoidHomClass + +variable {A B F : Type*} [FunLike F ℕ A] + +lemma ext_nat' [AddMonoid A] [AddMonoidHomClass F ℕ A] (f g : F) (h : f 1 = g 1) : f = g := + DFunLike.ext f g <| by + intro n + induction n with + | zero => simp_rw [map_zero f, map_zero g] + | succ n ihn => + simp [h, ihn] + +@[ext] +lemma AddMonoidHom.ext_nat [AddMonoid A] {f g : ℕ →+ A} : f 1 = g 1 → f = g := + ext_nat' f g + +end AddMonoidHomClass + +section AddMonoid +variable [AddMonoid M] + +variable (M) in +/-- Additive homomorphisms from `ℕ` are defined by the image of `1`. -/ +def multiplesHom : M ≃ (ℕ →+ M) where + toFun x := + { toFun := fun n ↦ n • x + map_zero' := zero_nsmul x + map_add' := fun _ _ ↦ add_nsmul _ _ _ } + invFun f := f 1 + left_inv := one_nsmul + right_inv f := AddMonoidHom.ext_nat <| one_nsmul (f 1) + +@[simp] lemma multiplesHom_apply (x : M) (n : ℕ) : multiplesHom M x n = n • x := rfl + +lemma multiplesHom_symm_apply (f : ℕ →+ M) : (multiplesHom M).symm f = f 1 := rfl + +lemma AddMonoidHom.apply_nat (f : ℕ →+ M) (n : ℕ) : f n = n • f 1 := by + rw [← multiplesHom_symm_apply, ← multiplesHom_apply, Equiv.apply_symm_apply] + +end AddMonoid + +section Monoid +variable [Monoid M] + +variable (M) in +/-- Monoid homomorphisms from `Multiplicative ℕ` are defined by the image +of `Multiplicative.ofAdd 1`. -/ +@[to_additive existing] +def powersHom : M ≃ (Multiplicative ℕ →* M) := + Additive.ofMul.trans <| (multiplesHom _).trans <| AddMonoidHom.toMultiplicative'' + +@[to_additive existing (attr := simp)] +lemma powersHom_apply (x : M) (n : Multiplicative ℕ) : + powersHom M x n = x ^ n.toAdd := rfl + +@[to_additive existing (attr := simp)] +lemma powersHom_symm_apply (f : Multiplicative ℕ →* M) : + (powersHom M).symm f = f (Multiplicative.ofAdd 1) := rfl + +@[to_additive existing AddMonoidHom.apply_nat] +lemma MonoidHom.apply_mnat (f : Multiplicative ℕ →* M) (n : Multiplicative ℕ) : + f n = f (Multiplicative.ofAdd 1) ^ n.toAdd := by + rw [← powersHom_symm_apply, ← powersHom_apply, Equiv.apply_symm_apply] + +@[to_additive existing (attr := ext) AddMonoidHom.ext_nat] +lemma MonoidHom.ext_mnat ⦃f g : Multiplicative ℕ →* M⦄ + (h : f (Multiplicative.ofAdd 1) = g (Multiplicative.ofAdd 1)) : f = g := + MonoidHom.ext fun n ↦ by rw [f.apply_mnat, g.apply_mnat, h] + +end Monoid + +section AddCommMonoid +variable [AddCommMonoid M] + +variable (M) in +/-- If `M` is commutative, `multiplesHom` is an additive equivalence. -/ +def multiplesAddHom : M ≃+ (ℕ →+ M) where + __ := multiplesHom M + map_add' a b := AddMonoidHom.ext fun n ↦ by simp [nsmul_add] + +@[simp] lemma multiplesAddHom_apply (x : M) (n : ℕ) : multiplesAddHom M x n = n • x := rfl + +@[simp] lemma multiplesAddHom_symm_apply (f : ℕ →+ M) : (multiplesAddHom M).symm f = f 1 := rfl + +end AddCommMonoid + +section CommMonoid +variable [CommMonoid M] + +variable (M) in +/-- If `M` is commutative, `powersHom` is a multiplicative equivalence. -/ +@[to_additive existing] +def powersMulHom : M ≃* (Multiplicative ℕ →* M) where + __ := powersHom M + map_mul' a b := MonoidHom.ext fun n ↦ by simp [mul_pow] + +@[to_additive existing (attr := simp)] +lemma powersMulHom_apply (x : M) (n : Multiplicative ℕ) : powersMulHom M x n = x ^ n.toAdd := rfl + +@[to_additive existing (attr := simp)] +lemma powersMulHom_symm_apply (f : Multiplicative ℕ →* M) : (powersMulHom M).symm f = f (ofAdd 1) := + rfl + +end CommMonoid diff --git a/Mathlib/Algebra/Group/Operations.lean b/Mathlib/Algebra/Group/Operations.lean index 72ef50266f6c9..596f447f86c0e 100644 --- a/Mathlib/Algebra/Group/Operations.lean +++ b/Mathlib/Algebra/Group/Operations.lean @@ -140,7 +140,7 @@ variable {G : Type*} /-- Class of types that have an inversion operation. -/ @[to_additive, notation_class] class Inv (α : Type u) where - /-- Invert an element of α. -/ + /-- Invert an element of α, denoted by `a⁻¹`. -/ inv : α → α @[inherit_doc] diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index af7620fdb7069..e1433b9eac116 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -1689,13 +1689,13 @@ theorem smul_finset_subset_smul_finset_iff : a • s ⊆ a • t ↔ s ⊆ t := theorem smul_finset_subset_iff : a • s ⊆ t ↔ s ⊆ a⁻¹ • t := by simp_rw [← coe_subset] push_cast - exact Set.set_smul_subset_iff + exact Set.smul_set_subset_iff_subset_inv_smul_set @[to_additive] theorem subset_smul_finset_iff : s ⊆ a • t ↔ a⁻¹ • s ⊆ t := by simp_rw [← coe_subset] push_cast - exact Set.subset_set_smul_iff + exact Set.subset_smul_set_iff @[to_additive] theorem smul_finset_inter : a • (s ∩ t) = a • s ∩ a • t := diff --git a/Mathlib/Algebra/Group/Subgroup/Defs.lean b/Mathlib/Algebra/Group/Subgroup/Defs.lean index eea53f4a65caf..a97477f0b4657 100644 --- a/Mathlib/Algebra/Group/Subgroup/Defs.lean +++ b/Mathlib/Algebra/Group/Subgroup/Defs.lean @@ -5,6 +5,7 @@ Authors: Kexing Ying -/ import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Tactic.Common +import Mathlib.Tactic.FastInstance /-! # Subgroups @@ -175,7 +176,7 @@ variable (H) -- Prefer subclasses of `Group` over subclasses of `SubgroupClass`. /-- A subgroup of a group inherits a group structure. -/ @[to_additive "An additive subgroup of an `AddGroup` inherits an `AddGroup` structure."] -instance (priority := 75) toGroup : Group H := +instance (priority := 75) toGroup : Group H := fast_instance% Subtype.coe_injective.group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -183,7 +184,7 @@ instance (priority := 75) toGroup : Group H := /-- A subgroup of a `CommGroup` is a `CommGroup`. -/ @[to_additive "An additive subgroup of an `AddCommGroup` is an `AddCommGroup`."] instance (priority := 75) toCommGroup {G : Type*} [CommGroup G] [SetLike S G] [SubgroupClass S G] : - CommGroup H := + CommGroup H := fast_instance% Subtype.coe_injective.commGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -504,13 +505,13 @@ theorem mk_eq_one {g : G} {h} : (⟨g, h⟩ : H) = 1 ↔ g = 1 := Submonoid.mk_e /-- A subgroup of a group inherits a group structure. -/ @[to_additive "An `AddSubgroup` of an `AddGroup` inherits an `AddGroup` structure."] -instance toGroup {G : Type*} [Group G] (H : Subgroup G) : Group H := +instance toGroup {G : Type*} [Group G] (H : Subgroup G) : Group H := fast_instance% Subtype.coe_injective.group _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl /-- A subgroup of a `CommGroup` is a `CommGroup`. -/ @[to_additive "An `AddSubgroup` of an `AddCommGroup` is an `AddCommGroup`."] -instance toCommGroup {G : Type*} [CommGroup G] (H : Subgroup G) : CommGroup H := +instance toCommGroup {G : Type*} [CommGroup G] (H : Subgroup G) : CommGroup H := fast_instance% Subtype.coe_injective.commGroup _ rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index 9d32674ef1a3b..b2d5af26623ba 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -403,13 +403,13 @@ theorem mem_inv_pointwise_smul_iff {a : α} {S : Subgroup G} {x : G} : x ∈ a @[simp] theorem pointwise_smul_le_pointwise_smul_iff {a : α} {S T : Subgroup G} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff + smul_set_subset_smul_set_iff theorem pointwise_smul_subset_iff {a : α} {S T : Subgroup G} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff + smul_set_subset_iff_subset_inv_smul_set theorem subset_pointwise_smul_iff {a : α} {S T : Subgroup G} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff + subset_smul_set_iff @[simp] theorem smul_inf (a : α) (S T : Subgroup G) : a • (S ⊓ T) = a • S ⊓ a • T := by @@ -473,13 +473,13 @@ theorem mem_inv_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) (S : Subgroup G) ( @[simp] theorem pointwise_smul_le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : Subgroup G} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff₀ ha + smul_set_subset_smul_set_iff₀ ha theorem pointwise_smul_le_iff₀ {a : α} (ha : a ≠ 0) {S T : Subgroup G} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff₀ ha + smul_set_subset_iff₀ ha theorem le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : Subgroup G} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff₀ ha + subset_smul_set_iff₀ ha end GroupWithZero @@ -551,13 +551,13 @@ theorem mem_inv_pointwise_smul_iff {a : α} {S : AddSubgroup A} {x : A} : x ∈ @[simp] theorem pointwise_smul_le_pointwise_smul_iff {a : α} {S T : AddSubgroup A} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff + smul_set_subset_smul_set_iff theorem pointwise_smul_le_iff {a : α} {S T : AddSubgroup A} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff + smul_set_subset_iff_subset_inv_smul_set theorem le_pointwise_smul_iff {a : α} {S T : AddSubgroup A} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff + subset_smul_set_iff end Group @@ -583,15 +583,15 @@ theorem mem_inv_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) (S : AddSubgroup A @[simp] theorem pointwise_smul_le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubgroup A} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff₀ ha + smul_set_subset_smul_set_iff₀ ha theorem pointwise_smul_le_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubgroup A} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff₀ ha + smul_set_subset_iff₀ ha theorem le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubgroup A} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff₀ ha + subset_smul_set_iff₀ ha end GroupWithZero diff --git a/Mathlib/Algebra/Group/Submonoid/Defs.lean b/Mathlib/Algebra/Group/Submonoid/Defs.lean index 4f7ad4ad1b945..e4735cec1b93f 100644 --- a/Mathlib/Algebra/Group/Submonoid/Defs.lean +++ b/Mathlib/Algebra/Group/Submonoid/Defs.lean @@ -6,6 +6,7 @@ Amelia Livingston, Yury Kudryashov -/ import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Group.Subsemigroup.Defs +import Mathlib.Tactic.FastInstance /-! # Submonoids: definition @@ -351,21 +352,21 @@ theorem mk_pow {M} [Monoid M] {A : Type*} [SetLike A M] [SubmonoidClass A M] {S @[to_additive "An `AddSubmonoid` of a unital additive magma inherits a unital additive magma structure."] instance (priority := 75) toMulOneClass {M : Type*} [MulOneClass M] {A : Type*} [SetLike A M] - [SubmonoidClass A M] (S : A) : MulOneClass S := - Subtype.coe_injective.mulOneClass Subtype.val rfl (fun _ _ => rfl) + [SubmonoidClass A M] (S : A) : MulOneClass S := fast_instance% + Subtype.coe_injective.mulOneClass Subtype.val rfl (fun _ _ => rfl) -- Prefer subclasses of `Monoid` over subclasses of `SubmonoidClass`. /-- A submonoid of a monoid inherits a monoid structure. -/ @[to_additive "An `AddSubmonoid` of an `AddMonoid` inherits an `AddMonoid` structure."] instance (priority := 75) toMonoid {M : Type*} [Monoid M] {A : Type*} [SetLike A M] - [SubmonoidClass A M] (S : A) : Monoid S := + [SubmonoidClass A M] (S : A) : Monoid S := fast_instance% Subtype.coe_injective.monoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) -- Prefer subclasses of `Monoid` over subclasses of `SubmonoidClass`. /-- A submonoid of a `CommMonoid` is a `CommMonoid`. -/ @[to_additive "An `AddSubmonoid` of an `AddCommMonoid` is an `AddCommMonoid`."] instance (priority := 75) toCommMonoid {M} [CommMonoid M] {A : Type*} [SetLike A M] - [SubmonoidClass A M] (S : A) : CommMonoid S := + [SubmonoidClass A M] (S : A) : CommMonoid S := fast_instance% Subtype.coe_injective.commMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- The natural monoid hom from a submonoid of monoid `M` to `M`. -/ @@ -420,7 +421,8 @@ theorem one_def : (1 : S) = ⟨1, S.one_mem⟩ := /-- A submonoid of a unital magma inherits a unital magma structure. -/ @[to_additive "An `AddSubmonoid` of a unital additive magma inherits a unital additive magma structure."] -instance toMulOneClass {M : Type*} [MulOneClass M] (S : Submonoid M) : MulOneClass S := +instance toMulOneClass {M : Type*} [MulOneClass M] (S : Submonoid M) : + MulOneClass S := fast_instance% Subtype.coe_injective.mulOneClass Subtype.val rfl fun _ _ => rfl @[to_additive] @@ -430,12 +432,12 @@ protected theorem pow_mem {M : Type*} [Monoid M] (S : Submonoid M) {x : M} (hx : /-- A submonoid of a monoid inherits a monoid structure. -/ @[to_additive "An `AddSubmonoid` of an `AddMonoid` inherits an `AddMonoid` structure."] -instance toMonoid {M : Type*} [Monoid M] (S : Submonoid M) : Monoid S := +instance toMonoid {M : Type*} [Monoid M] (S : Submonoid M) : Monoid S := fast_instance% Subtype.coe_injective.monoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- A submonoid of a `CommMonoid` is a `CommMonoid`. -/ @[to_additive "An `AddSubmonoid` of an `AddCommMonoid` is an `AddCommMonoid`."] -instance toCommMonoid {M} [CommMonoid M] (S : Submonoid M) : CommMonoid S := +instance toCommMonoid {M} [CommMonoid M] (S : Submonoid M) : CommMonoid S := fast_instance% Subtype.coe_injective.commMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- The natural monoid hom from a submonoid of monoid `M` to `M`. -/ diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 8121874d5e833..5c401ebbaecf0 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -6,6 +6,7 @@ Amelia Livingston, Yury Kudryashov -/ import Mathlib.Algebra.FreeMonoid.Basic import Mathlib.Algebra.Group.Idempotent +import Mathlib.Algebra.Group.Nat.Hom import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.GroupWithZero.Divisibility diff --git a/Mathlib/Algebra/Group/Submonoid/MulAction.lean b/Mathlib/Algebra/Group/Submonoid/MulAction.lean new file mode 100644 index 0000000000000..7c37bc72c9f08 --- /dev/null +++ b/Mathlib/Algebra/Group/Submonoid/MulAction.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.Group.Submonoid.Defs +import Mathlib.Algebra.Group.Action.Defs + +/-! +# Actions by `Submonoid`s + +These instances transfer the action by an element `m : M` of a monoid `M` written as `m • a` onto +the action by an element `s : S` of a submonoid `S : Submonoid M` such that `s • a = (s : M) • a`. + +These instances work particularly well in conjunction with `Monoid.toMulAction`, enabling +`s • m` as an alias for `↑s * m`. +-/ + +namespace Submonoid + +variable {M' : Type*} {α β : Type*} + +section MulOneClass + +variable [MulOneClass M'] + +@[to_additive] +instance smul [SMul M' α] (S : Submonoid M') : SMul S α := + SMul.comp _ S.subtype + +@[to_additive] +instance smulCommClass_left [SMul M' β] [SMul α β] [SMulCommClass M' α β] + (S : Submonoid M') : SMulCommClass S α β := + ⟨fun a _ _ => smul_comm (a : M') _ _⟩ + +@[to_additive] +instance smulCommClass_right [SMul α β] [SMul M' β] [SMulCommClass α M' β] + (S : Submonoid M') : SMulCommClass α S β := + ⟨fun a s => smul_comm a (s : M')⟩ + +/-- Note that this provides `IsScalarTower S M' M'` which is needed by `SMulMulAssoc`. -/ +instance isScalarTower [SMul α β] [SMul M' α] [SMul M' β] [IsScalarTower M' α β] + (S : Submonoid M') : + IsScalarTower S α β := + ⟨fun a => smul_assoc (a : M')⟩ + +section SMul +variable [SMul M' α] {S : Submonoid M'} + +@[to_additive] lemma smul_def (g : S) (a : α) : g • a = (g : M') • a := rfl + +@[to_additive (attr := simp)] +lemma mk_smul (g : M') (hg : g ∈ S) (a : α) : (⟨g, hg⟩ : S) • a = g • a := rfl + +end SMul +end MulOneClass + +variable [Monoid M'] + +/-- The action by a submonoid is the action by the underlying monoid. -/ +@[to_additive + "The additive action by an `AddSubmonoid` is the action by the underlying `AddMonoid`. "] +instance mulAction [MulAction M' α] (S : Submonoid M') : MulAction S α where + one_smul := one_smul M' + mul_smul m₁ m₂ := mul_smul (m₁ : M') m₂ + +example {S : Submonoid M'} : IsScalarTower S M' M' := by infer_instance + +end Submonoid diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 9037ade39062b..a997c5650f854 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.Group.Nat.Defs import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Submonoid.Basic +import Mathlib.Algebra.Group.Submonoid.MulAction import Mathlib.Algebra.Group.TypeTags.Basic /-! @@ -973,75 +974,10 @@ theorem Submonoid.equivMapOfInjective_coe_mulEquiv (e : M ≃* N) : ext rfl -section Actions - -/-! ### Actions by `Submonoid`s - -These instances transfer the action by an element `m : M` of a monoid `M` written as `m • a` onto -the action by an element `s : S` of a submonoid `S : Submonoid M` such that `s • a = (s : M) • a`. - -These instances work particularly well in conjunction with `Monoid.toMulAction`, enabling -`s • m` as an alias for `↑s * m`. --/ - - -namespace Submonoid - -variable {M' : Type*} {α β : Type*} - -section MulOneClass - -variable [MulOneClass M'] - -@[to_additive] -instance smul [SMul M' α] (S : Submonoid M') : SMul S α := - SMul.comp _ S.subtype - -@[to_additive] -instance smulCommClass_left [SMul M' β] [SMul α β] [SMulCommClass M' α β] - (S : Submonoid M') : SMulCommClass S α β := - ⟨fun a _ _ => smul_comm (a : M') _ _⟩ - -@[to_additive] -instance smulCommClass_right [SMul α β] [SMul M' β] [SMulCommClass α M' β] - (S : Submonoid M') : SMulCommClass α S β := - ⟨fun a s => smul_comm a (s : M')⟩ - -/-- Note that this provides `IsScalarTower S M' M'` which is needed by `SMulMulAssoc`. -/ -instance isScalarTower [SMul α β] [SMul M' α] [SMul M' β] [IsScalarTower M' α β] - (S : Submonoid M') : - IsScalarTower S α β := - ⟨fun a => smul_assoc (a : M')⟩ - -section SMul -variable [SMul M' α] {S : Submonoid M'} - -@[to_additive] lemma smul_def (g : S) (a : α) : g • a = (g : M') • a := rfl - -@[to_additive (attr := simp)] -lemma mk_smul (g : M') (hg : g ∈ S) (a : α) : (⟨g, hg⟩ : S) • a = g • a := rfl - -instance faithfulSMul [FaithfulSMul M' α] : FaithfulSMul S α := +instance Submonoid.faithfulSMul {M' α : Type*} [MulOneClass M'] [SMul M' α] {S : Submonoid M'} + [FaithfulSMul M' α] : FaithfulSMul S α := ⟨fun h => Subtype.ext <| eq_of_smul_eq_smul h⟩ -end SMul -end MulOneClass - -variable [Monoid M'] - -/-- The action by a submonoid is the action by the underlying monoid. -/ -@[to_additive - "The additive action by an `AddSubmonoid` is the action by the underlying `AddMonoid`. "] -instance mulAction [MulAction M' α] (S : Submonoid M') : MulAction S α := - MulAction.compHom _ S.subtype - -example {S : Submonoid M'} : IsScalarTower S M' M' := by infer_instance - - -end Submonoid - -end Actions - section Units namespace Submonoid diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 567936d4f37ef..f698bdf137947 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -275,13 +275,13 @@ theorem mem_inv_pointwise_smul_iff {a : α} {S : Submonoid M} {x : M} : x ∈ a @[simp] theorem pointwise_smul_le_pointwise_smul_iff {a : α} {S T : Submonoid M} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff + smul_set_subset_smul_set_iff theorem pointwise_smul_subset_iff {a : α} {S T : Submonoid M} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff + smul_set_subset_iff_subset_inv_smul_set theorem subset_pointwise_smul_iff {a : α} {S T : Submonoid M} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff + subset_smul_set_iff end Group @@ -305,13 +305,13 @@ theorem mem_inv_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) (S : Submonoid M) @[simp] theorem pointwise_smul_le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : Submonoid M} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff₀ ha + smul_set_subset_smul_set_iff₀ ha theorem pointwise_smul_le_iff₀ {a : α} (ha : a ≠ 0) {S T : Submonoid M} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff₀ ha + smul_set_subset_iff₀ ha theorem le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : Submonoid M} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff₀ ha + subset_smul_set_iff₀ ha end GroupWithZero @@ -389,13 +389,13 @@ theorem mem_inv_pointwise_smul_iff {a : α} {S : AddSubmonoid A} {x : A} : x ∈ @[simp] theorem pointwise_smul_le_pointwise_smul_iff {a : α} {S T : AddSubmonoid A} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff + smul_set_subset_smul_set_iff theorem pointwise_smul_le_iff {a : α} {S T : AddSubmonoid A} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff + smul_set_subset_iff_subset_inv_smul_set theorem le_pointwise_smul_iff {a : α} {S T : AddSubmonoid A} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff + subset_smul_set_iff end Group @@ -419,15 +419,15 @@ theorem mem_inv_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) (S : AddSubmonoid @[simp] theorem pointwise_smul_le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubmonoid A} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff₀ ha + smul_set_subset_smul_set_iff₀ ha theorem pointwise_smul_le_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubmonoid A} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff₀ ha + smul_set_subset_iff₀ ha theorem le_pointwise_smul_iff₀ {a : α} (ha : a ≠ 0) {S T : AddSubmonoid A} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff₀ ha + subset_smul_set_iff₀ ha end GroupWithZero diff --git a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean index 24404d2fb09fc..1523e0b5c57c6 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Defs.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Defs.lean @@ -7,6 +7,7 @@ Amelia Livingston, Yury Kudryashov, Yakov Pechersky import Mathlib.Algebra.Group.Hom.Defs import Mathlib.Algebra.Group.InjSurj import Mathlib.Data.SetLike.Basic +import Mathlib.Tactic.FastInstance /-! # Subsemigroups: definition @@ -254,13 +255,13 @@ theorem mul_def (x y : S') : x * y = ⟨x * y, mul_mem x.2 y.2⟩ := /-- A subsemigroup of a semigroup inherits a semigroup structure. -/ @[to_additive "An `AddSubsemigroup` of an `AddSemigroup` inherits an `AddSemigroup` structure."] instance toSemigroup {M : Type*} [Semigroup M] {A : Type*} [SetLike A M] [MulMemClass A M] - (S : A) : Semigroup S := + (S : A) : Semigroup S := fast_instance% Subtype.coe_injective.semigroup Subtype.val fun _ _ => rfl /-- A subsemigroup of a `CommSemigroup` is a `CommSemigroup`. -/ @[to_additive "An `AddSubsemigroup` of an `AddCommSemigroup` is an `AddCommSemigroup`."] instance toCommSemigroup {M} [CommSemigroup M] {A : Type*} [SetLike A M] [MulMemClass A M] - (S : A) : CommSemigroup S := + (S : A) : CommSemigroup S := fast_instance% Subtype.coe_injective.commSemigroup Subtype.val fun _ _ => rfl /-- The natural semigroup hom from a subsemigroup of semigroup `M` to `M`. -/ diff --git a/Mathlib/Algebra/Group/Units/Defs.lean b/Mathlib/Algebra/Group/Units/Defs.lean index 62cfa10d66c99..4a0c7e074d41d 100644 --- a/Mathlib/Algebra/Group/Units/Defs.lean +++ b/Mathlib/Algebra/Group/Units/Defs.lean @@ -296,7 +296,7 @@ section Monoid variable [Monoid α] {a : α} -/-- Partial division. It is defined when the +/-- Partial division, denoted `a /ₚ u`. It is defined when the second argument is invertible, and unlike the division operator in `DivisionRing` it is not totalized at zero. -/ def divp (a : α) (u : Units α) : α := diff --git a/Mathlib/Algebra/Group/WithOne/Defs.lean b/Mathlib/Algebra/Group/WithOne/Defs.lean index dc7d3fcec8135..135c6d8b35e74 100644 --- a/Mathlib/Algebra/Group/WithOne/Defs.lean +++ b/Mathlib/Algebra/Group/WithOne/Defs.lean @@ -170,7 +170,7 @@ lemma coe_mul [Mul α] (a b : α) : (↑(a * b) : WithOne α) = a * b := rfl @[to_additive] instance monoid [Semigroup α] : Monoid (WithOne α) where __ := mulOneClass - mul_assoc a b c := match a, b, c with + mul_assoc | 1, b, c => by simp | (a : α), 1, c => by simp | (a : α), (b : α), 1 => by simp @@ -178,7 +178,7 @@ instance monoid [Semigroup α] : Monoid (WithOne α) where @[to_additive] instance commMonoid [CommSemigroup α] : CommMonoid (WithOne α) where - mul_comm := fun a b => match a, b with + mul_comm | (a : α), (b : α) => congr_arg some (mul_comm a b) | (_ : α), 1 => rfl | 1, (_ : α) => rfl diff --git a/Mathlib/Algebra/GroupWithZero/Action/ConjAct.lean b/Mathlib/Algebra/GroupWithZero/Action/ConjAct.lean new file mode 100644 index 0000000000000..844e8ba3b3efc --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Action/ConjAct.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.GroupTheory.GroupAction.ConjAct + +/-! +# Conjugation action of a group with zero on itself +-/ + +assert_not_exists Ring + +variable {α G₀ : Type*} + +namespace ConjAct +variable [GroupWithZero G₀] + +instance : GroupWithZero (ConjAct G₀) := ‹GroupWithZero G₀› + +@[simp] lemma ofConjAct_zero : ofConjAct 0 = (0 : G₀) := rfl +@[simp] lemma toConjAct_zero : toConjAct (0 : G₀) = 0 := rfl + +instance mulAction₀ : MulAction (ConjAct G₀) G₀ where + one_smul := by simp [smul_def] + mul_smul := by simp [smul_def, mul_assoc] + +instance smulCommClass₀ [SMul α G₀] [SMulCommClass α G₀ G₀] [IsScalarTower α G₀ G₀] : + SMulCommClass α (ConjAct G₀) G₀ where + smul_comm a ug g := by rw [smul_def, smul_def, mul_smul_comm, smul_mul_assoc] + +instance smulCommClass₀' [SMul α G₀] [SMulCommClass G₀ α G₀] [IsScalarTower α G₀ G₀] : + SMulCommClass (ConjAct G₀) α G₀ := + haveI := SMulCommClass.symm G₀ α G₀ + .symm .. + +end ConjAct diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index 66ad8bce61f90..b2bdd72d7a6b6 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -237,6 +237,16 @@ theorem nonZeroDivisors_le_comap_nonZeroDivisors_of_injective [NoZeroDivisors M' [MonoidWithZeroHomClass F M M'] (f : F) (hf : Function.Injective f) : M⁰ ≤ M'⁰.comap f := Submonoid.le_comap_of_map_le _ (map_le_nonZeroDivisors_of_injective _ hf le_rfl) +/-- If an element maps to a non-zero-divisor via injective homomorphism, +then it is non-zero-divisor. -/ +theorem mem_nonZeroDivisor_of_injective [MonoidWithZeroHomClass F M M'] {f : F} + (hf : Function.Injective f) {a : M} (ha : f a ∈ M'⁰) : a ∈ M⁰ := + fun x hx ↦ hf <| map_zero f ▸ ha (f x) (map_mul f x a ▸ map_zero f ▸ congrArg f hx) + +theorem comap_nonZeroDivisor_le_of_injective [MonoidWithZeroHomClass F M M'] {f : F} + (hf : Function.Injective f) : M'⁰.comap f ≤ M⁰ := + fun _ ha ↦ mem_nonZeroDivisor_of_injective hf (Submonoid.mem_comap.mp ha) + /-- In a finite ring, an element is a unit iff it is a non-zero-divisor. -/ lemma isUnit_iff_mem_nonZeroDivisors_of_finite [Finite R] {a : R} : IsUnit a ↔ a ∈ nonZeroDivisors R := by diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 9c876dd881321..c4108f058926d 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -196,7 +196,6 @@ figure out `p` when using `Units.exists0` from right to left. -/ theorem exists0' {p : ∀ g : G₀, g ≠ 0 → Prop} : (∃ (g : G₀) (hg : g ≠ 0), p g hg) ↔ ∃ g : G₀ˣ, p g g.ne_zero := Iff.trans (by simp_rw [val_mk0]) exists0.symm - -- Porting note: had to add the `rfl` @[simp] theorem exists_iff_ne_zero {p : G₀ → Prop} : (∃ u : G₀ˣ, p u) ↔ ∃ x ≠ 0, p x := by diff --git a/Mathlib/Algebra/GroupWithZero/WithZero.lean b/Mathlib/Algebra/GroupWithZero/WithZero.lean index 8acf0b8bcb620..43ea0028206c0 100644 --- a/Mathlib/Algebra/GroupWithZero/WithZero.lean +++ b/Mathlib/Algebra/GroupWithZero/WithZero.lean @@ -142,7 +142,7 @@ section Pow variable [One α] [Pow α ℕ] instance pow : Pow (WithZero α) ℕ where - pow x n := match x, n with + pow | none, 0 => 1 | none, _ + 1 => 0 | some x, n => ↑(x ^ n) @@ -155,12 +155,12 @@ instance monoidWithZero [Monoid α] : MonoidWithZero (WithZero α) where __ := mulZeroOneClass __ := semigroupWithZero npow n a := a ^ n - npow_zero a := match a with - | none => rfl + npow_zero + | 0 => rfl | some _ => congr_arg some (pow_zero _) - npow_succ n a := match a with - | none => by change 0 ^ (n + 1) = 0 ^ n * 0; simp only [mul_zero]; rfl - | some _ => congr_arg some <| pow_succ _ _ + npow_succ + | n, 0 => by simp only [mul_zero]; rfl + | n, some _ => congr_arg some <| pow_succ _ _ instance commMonoidWithZero [CommMonoid α] : CommMonoidWithZero (WithZero α) := { WithZero.monoidWithZero, WithZero.commSemigroup with } @@ -193,7 +193,7 @@ section ZPow variable [One α] [Pow α ℤ] instance : Pow (WithZero α) ℤ where - pow a n := match a, n with + pow | none, Int.ofNat 0 => 1 | none, Int.ofNat (Nat.succ _) => 0 | none, Int.negSucc _ => 0 @@ -205,20 +205,20 @@ end ZPow instance divInvMonoid [DivInvMonoid α] : DivInvMonoid (WithZero α) where __ := monoidWithZero - div_eq_mul_inv a b := match a, b with + div_eq_mul_inv | none, _ => rfl | some _, none => rfl | some a, some b => congr_arg some (div_eq_mul_inv a b) zpow n a := a ^ n - zpow_zero' a := match a with + zpow_zero' | none => rfl | some _ => congr_arg some (zpow_zero _) - zpow_succ' n a := match a with - | none => by change 0 ^ _ = 0 ^ _ * 0; simp only [mul_zero]; rfl - | some _ => congr_arg some (DivInvMonoid.zpow_succ' _ _) - zpow_neg' _ a := match a with - | none => rfl - | some _ => congr_arg some (DivInvMonoid.zpow_neg' _ _) + zpow_succ' + | n, none => by change 0 ^ _ = 0 ^ _ * 0; simp only [mul_zero]; rfl + | n, some _ => congr_arg some (DivInvMonoid.zpow_succ' _ _) + zpow_neg' + | n, none => rfl + | n, some _ => congr_arg some (DivInvMonoid.zpow_neg' _ _) instance divInvOneMonoid [DivInvOneMonoid α] : DivInvOneMonoid (WithZero α) where __ := divInvMonoid @@ -230,16 +230,14 @@ instance involutiveInv [InvolutiveInv α] : InvolutiveInv (WithZero α) where instance divisionMonoid [DivisionMonoid α] : DivisionMonoid (WithZero α) where __ := divInvMonoid __ := involutiveInv - mul_inv_rev a b := match a, b with + mul_inv_rev | none, none => rfl | none, some _ => rfl | some _, none => rfl | some _, some _ => congr_arg some (mul_inv_rev _ _) - inv_eq_of_mul a b := match a, b with - | none, none => fun _ ↦ rfl - | none, some b => fun _ ↦ by contradiction - | some a, none => fun _ ↦ by contradiction - | some _, some _ => fun h ↦ + inv_eq_of_mul + | none, none, _ => rfl + | some _, some _, h => congr_arg some <| inv_eq_of_mul_eq_one_right <| Option.some_injective _ h instance divisionCommMonoid [DivisionCommMonoid α] : DivisionCommMonoid (WithZero α) where @@ -260,7 +258,6 @@ instance groupWithZero : GroupWithZero (WithZero α) where norm_cast apply mul_inv_cancel - /-- Any group is isomorphic to the units of itself adjoined with `0`. -/ def unitsWithZeroEquiv : (WithZero α)ˣ ≃* α where toFun a := unzero a.ne_zero @@ -277,9 +274,7 @@ def withZeroUnitsEquiv {G : Type*} [GroupWithZero G] invFun a := if h : a = 0 then 0 else (Units.mk0 a h : Gˣ) left_inv := (by induction · <;> simp) right_inv _ := by simp only; split <;> simp_all - map_mul' x y := by - induction x <;> induction y <;> - simp [← WithZero.coe_mul, ← Units.val_mul] + map_mul' := (by induction · <;> induction · <;> simp [← WithZero.coe_mul]) /-- A version of `Equiv.optionCongr` for `WithZero`. -/ noncomputable def _root_.MulEquiv.withZero [Group β] (e : α ≃* β) : @@ -288,9 +283,7 @@ noncomputable def _root_.MulEquiv.withZero [Group β] (e : α ≃* β) : invFun := map' e.symm.toMonoidHom left_inv := (by induction · <;> simp) right_inv := (by induction · <;> simp) - map_mul' x y := by - induction x <;> induction y <;> - simp + map_mul' := (by induction · <;> induction · <;> simp) /-- The inverse of `MulEquiv.withZero`. -/ protected noncomputable def _root_.MulEquiv.unzero [Group β] (e : WithZero α ≃* WithZero β) : @@ -314,11 +307,6 @@ instance commGroupWithZero [CommGroup α] : CommGroupWithZero (WithZero α) := instance addMonoidWithOne [AddMonoidWithOne α] : AddMonoidWithOne (WithZero α) where natCast n := if n = 0 then 0 else (n : α) natCast_zero := rfl - natCast_succ n := by - cases n with - | zero => show (((1 : ℕ) : α) : WithZero α) = 0 + 1; · rw [Nat.cast_one, coe_one, zero_add] - | succ n => - show (((n + 2 : ℕ) : α) : WithZero α) = ((n + 1 : ℕ) : α) + 1 - rw [Nat.cast_succ, coe_add, coe_one] + natCast_succ n := by cases n <;> simp end WithZero diff --git a/Mathlib/Algebra/Homology/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ConcreteCategory.lean index f510e2876a6f3..9c7f9f7bae473 100644 --- a/Mathlib/Algebra/Homology/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ConcreteCategory.lean @@ -80,7 +80,7 @@ lemma δ_apply (x₃ : (forget₂ C Ab).obj (S.X₃.X i)) AddMonoidHom.zero_apply])) := by refine hS.δ_apply' i j hij _ ((forget₂ C Ab).map (S.X₂.pOpcycles i) x₂) _ ?_ ?_ · rw [← forget₂_comp_apply, ← forget₂_comp_apply, - HomologicalComplex.p_opcyclesMap, Functor.map_comp, comp_apply, + HomologicalComplex.p_opcyclesMap, Functor.map_comp, CategoryTheory.comp_apply, HomologicalComplex.homology_π_ι, forget₂_comp_apply, hx₂, HomologicalComplex.i_cyclesMk] · apply (Preadditive.mono_iff_injective (S.X₂.iCycles j)).1 inferInstance conv_lhs => diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean index 805216557b246..908ef5c957030 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean @@ -77,6 +77,10 @@ lemma hasExt_of_hasDerivedCategory [HasDerivedCategory.{w} C] : HasExt.{w} C := rw [hasExt_iff.{w}] infer_instance +lemma HasExt.standard : HasExt.{max u v} C := by + letI := HasDerivedCategory.standard + exact hasExt_of_hasDerivedCategory _ + variable {C} variable [HasExt.{w} C] diff --git a/Mathlib/Algebra/Homology/DifferentialObject.lean b/Mathlib/Algebra/Homology/DifferentialObject.lean index b83a6ba0c7334..c434db0589dc7 100644 --- a/Mathlib/Algebra/Homology/DifferentialObject.lean +++ b/Mathlib/Algebra/Homology/DifferentialObject.lean @@ -61,10 +61,7 @@ namespace HomologicalComplex variable {β : Type*} [AddCommGroup β] (b : β) variable (V : Type*) [Category V] [HasZeroMorphisms V] --- Porting note: this should be moved to an earlier file. --- Porting note: simpNF linter silenced, both `d_eqToHom` and its `_assoc` version --- do not simplify under themselves -@[reassoc (attr := simp, nolint simpNF)] +@[reassoc] theorem d_eqToHom (X : HomologicalComplex V (ComplexShape.up' b)) {x y z : β} (h : y = z) : X.d x y ≫ eqToHom (congr_arg X.X h) = X.d x z := by cases h; simp diff --git a/Mathlib/Algebra/Homology/Embedding/Basic.lean b/Mathlib/Algebra/Homology/Embedding/Basic.lean index 00176a4d60d8d..828c12d3950a3 100644 --- a/Mathlib/Algebra/Homology/Embedding/Basic.lean +++ b/Mathlib/Algebra/Homology/Embedding/Basic.lean @@ -6,6 +6,7 @@ Authors: Joël Riou import Mathlib.Algebra.Homology.ComplexShape import Mathlib.Algebra.Ring.Int.Defs import Mathlib.Algebra.Ring.Nat +import Mathlib.Tactic.ByContra /-! # Embeddings of complex shapes @@ -213,4 +214,28 @@ instance : (embeddingUpIntLE p).IsRelIff := by dsimp [embeddingUpIntLE]; infer_i instance : (embeddingUpIntLE p).IsTruncLE where mem_prev {_ k} h := ⟨k + 1, by dsimp at h ⊢; omega⟩ +lemma not_mem_range_embeddingUpIntLE_iff (n : ℤ) : + (∀ (i : ℕ), (embeddingUpIntLE p).f i ≠ n) ↔ p < n := by + constructor + · intro h + by_contra! + simp only [Int.not_lt] at this + obtain ⟨k, rfl⟩ := Int.le.dest this + exact (h k) (by simp) + · intros + dsimp + omega + +lemma not_mem_range_embeddingUpIntGE_iff (n : ℤ) : + (∀ (i : ℕ), (embeddingUpIntGE p).f i ≠ n) ↔ n < p := by + constructor + · intro h + by_contra! + simp only [Int.not_lt] at this + obtain ⟨k, rfl⟩ := Int.le.dest this + exact (h k) (by simp) + · intros + dsimp + omega + end ComplexShape diff --git a/Mathlib/Algebra/Homology/Embedding/CochainComplex.lean b/Mathlib/Algebra/Homology/Embedding/CochainComplex.lean new file mode 100644 index 0000000000000..54fe12e166c96 --- /dev/null +++ b/Mathlib/Algebra/Homology/Embedding/CochainComplex.lean @@ -0,0 +1,224 @@ +/- +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.Homology.Embedding.TruncLE + +/-! +# Truncations on cochain complexes indexed by the integers. + +In this file, we introduce abbreviations for the canonical truncations +`CochainComplex.truncLE`, `CochainComplex.truncGE` of cochain +complexes indexed by `ℤ`, as well as the conditions +`CochainComplex.IsStrictlyLE`, `CochainComplex.IsStrictlyGE`, +`CochainComplex.IsLE`, and `CochainComplex.IsGE`. + +-/ + +open CategoryTheory Category Limits ComplexShape ZeroObject + +namespace CochainComplex + +variable {C : Type*} [Category C] + +open HomologicalComplex + +section HasZeroMorphisms + +variable [HasZeroMorphisms C] (K L : CochainComplex C ℤ) (φ : K ⟶ L) (e : K ≅ L) + +section + +variable [HasZeroObject C] [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] + +/-- If `K : CochainComplex C ℤ`, this is the canonical truncation `≤ n` of `K`. -/ +noncomputable abbrev truncLE (n : ℤ) : CochainComplex C ℤ := + HomologicalComplex.truncLE K (embeddingUpIntLE n) + +/-- If `K : CochainComplex C ℤ`, this is the canonical truncation `≥ n` of `K`. -/ +noncomputable abbrev truncGE (n : ℤ) : CochainComplex C ℤ := + HomologicalComplex.truncGE K (embeddingUpIntGE n) + +/-- The canonical map `K.truncLE n ⟶ K` for `K : CochainComplex C ℤ`. -/ +noncomputable def ιTruncLE (n : ℤ) : K.truncLE n ⟶ K := + HomologicalComplex.ιTruncLE K (embeddingUpIntLE n) + +/-- The canonical map `K ⟶ K.truncGE n` for `K : CochainComplex C ℤ`. -/ +noncomputable def πTruncGE (n : ℤ) : K ⟶ K.truncGE n := + HomologicalComplex.πTruncGE K (embeddingUpIntGE n) + +section + +variable {K L} + +/-- The morphism `K.truncLE n ⟶ L.truncLE n` induced by a morphism `K ⟶ L`. -/ +noncomputable abbrev truncLEMap (n : ℤ) : K.truncLE n ⟶ L.truncLE n := + HomologicalComplex.truncLEMap φ (embeddingUpIntLE n) + +/-- The morphism `K.truncGE n ⟶ L.truncGE n` induced by a morphism `K ⟶ L`. -/ +noncomputable abbrev truncGEMap (n : ℤ) : K.truncGE n ⟶ L.truncGE n := + HomologicalComplex.truncGEMap φ (embeddingUpIntGE n) + +@[reassoc (attr := simp)] +lemma ιTruncLE_naturality (n : ℤ) : + truncLEMap φ n ≫ L.ιTruncLE n = K.ιTruncLE n ≫ φ := by + apply HomologicalComplex.ιTruncLE_naturality + +@[reassoc (attr := simp)] +lemma πTruncGE_naturality (n : ℤ) : + K.πTruncGE n ≫ truncGEMap φ n = φ ≫ L.πTruncGE n := by + apply HomologicalComplex.πTruncGE_naturality + +end + +end + +/-- The condition that a cochain complex `K` is strictly `≤ n`. -/ +abbrev IsStrictlyGE (n : ℤ) := K.IsStrictlySupported (embeddingUpIntGE n) + +/-- The condition that a cochain complex `K` is strictly `≥ n`. -/ +abbrev IsStrictlyLE (n : ℤ) := K.IsStrictlySupported (embeddingUpIntLE n) + +/-- The condition that a cochain complex `K` is (cohomologically) `≤ n`. -/ +abbrev IsGE (n : ℤ) := K.IsSupported (embeddingUpIntGE n) + +/-- The condition that a cochain complex `K` is (cohomologically) `≥ n`. -/ +abbrev IsLE (n : ℤ) := K.IsSupported (embeddingUpIntLE n) + +lemma isZero_of_isStrictlyGE (n i : ℤ) (hi : i < n) [K.IsStrictlyGE n] : + IsZero (K.X i) := + isZero_X_of_isStrictlySupported K (embeddingUpIntGE n) i + (by simpa only [not_mem_range_embeddingUpIntGE_iff] using hi) + +lemma isZero_of_isStrictlyLE (n i : ℤ) (hi : n < i) [K.IsStrictlyLE n] : + IsZero (K.X i) := + isZero_X_of_isStrictlySupported K (embeddingUpIntLE n) i + (by simpa only [not_mem_range_embeddingUpIntLE_iff] using hi) + +lemma exactAt_of_isGE (n i : ℤ) (hi : i < n) [K.IsGE n] : + K.ExactAt i := + exactAt_of_isSupported K (embeddingUpIntGE n) i + (by simpa only [not_mem_range_embeddingUpIntGE_iff] using hi) + +lemma exactAt_of_isLE (n i : ℤ) (hi : n < i) [K.IsLE n] : + K.ExactAt i := + exactAt_of_isSupported K (embeddingUpIntLE n) i + (by simpa only [not_mem_range_embeddingUpIntLE_iff] using hi) + +lemma isZero_of_isGE (n i : ℤ) (hi : i < n) [K.IsGE n] [K.HasHomology i] : + IsZero (K.homology i) := + (K.exactAt_of_isGE n i hi).isZero_homology + +lemma isZero_of_isLE (n i : ℤ) (hi : n < i) [K.IsLE n] [K.HasHomology i] : + IsZero (K.homology i) := + (K.exactAt_of_isLE n i hi).isZero_homology + +lemma isStrictlyGE_iff (n : ℤ) : + K.IsStrictlyGE n ↔ ∀ (i : ℤ) (_ : i < n), IsZero (K.X i) := by + constructor + · intro _ i hi + exact K.isZero_of_isStrictlyGE n i hi + · intro h + refine IsStrictlySupported.mk (fun i hi ↦ ?_) + rw [not_mem_range_embeddingUpIntGE_iff] at hi + exact h i hi + +lemma isStrictlyLE_iff (n : ℤ) : + K.IsStrictlyLE n ↔ ∀ (i : ℤ) (_ : n < i), IsZero (K.X i) := by + constructor + · intro _ i hi + exact K.isZero_of_isStrictlyLE n i hi + · intro h + refine IsStrictlySupported.mk (fun i hi ↦ ?_) + rw [not_mem_range_embeddingUpIntLE_iff] at hi + exact h i hi + +lemma isGE_iff (n : ℤ) : + K.IsGE n ↔ ∀ (i : ℤ) (_ : i < n), K.ExactAt i := by + constructor + · intro _ i hi + exact K.exactAt_of_isGE n i hi + · intro h + refine IsSupported.mk (fun i hi ↦ ?_) + rw [not_mem_range_embeddingUpIntGE_iff] at hi + exact h i hi + +lemma isLE_iff (n : ℤ) : + K.IsLE n ↔ ∀ (i : ℤ) (_ : n < i), K.ExactAt i := by + constructor + · intro _ i hi + exact K.exactAt_of_isLE n i hi + · intro h + refine IsSupported.mk (fun i hi ↦ ?_) + rw [not_mem_range_embeddingUpIntLE_iff] at hi + exact h i hi + +lemma isStrictlyLE_of_le (p q : ℤ) (hpq : p ≤ q) [K.IsStrictlyLE p] : + K.IsStrictlyLE q := by + rw [isStrictlyLE_iff] + intro i hi + apply K.isZero_of_isStrictlyLE p + omega + +lemma isStrictlyGE_of_ge (p q : ℤ) (hpq : p ≤ q) [K.IsStrictlyGE q] : + K.IsStrictlyGE p := by + rw [isStrictlyGE_iff] + intro i hi + apply K.isZero_of_isStrictlyGE q + omega + +lemma isLE_of_le (p q : ℤ) (hpq : p ≤ q) [K.IsLE p] : + K.IsLE q := by + rw [isLE_iff] + intro i hi + apply K.exactAt_of_isLE p + omega + +lemma isGE_of_ge (p q : ℤ) (hpq : p ≤ q) [K.IsGE q] : + K.IsGE p := by + rw [isGE_iff] + intro i hi + apply K.exactAt_of_isGE q + omega + +section + +variable {K L} + +include e + +lemma isStrictlyLE_of_iso (n : ℤ) [K.IsStrictlyLE n] : L.IsStrictlyLE n := by + apply isStrictlySupported_of_iso e + +lemma isStrictlyGE_of_iso (n : ℤ) [K.IsStrictlyGE n] : L.IsStrictlyGE n := by + apply isStrictlySupported_of_iso e + +lemma isLE_of_iso (n : ℤ) [K.IsLE n] : L.IsLE n := by + apply isSupported_of_iso e + +lemma isGE_of_iso (n : ℤ) [K.IsGE n] : L.IsGE n := by + apply isSupported_of_iso e + +end + +/-- A cochain complex that is both strictly `≤ n` and `≥ n` is isomorphic to +a complex `(single _ _ n).obj M` for some object `M`. -/ +lemma exists_iso_single [HasZeroObject C] (n : ℤ) [K.IsStrictlyGE n] [K.IsStrictlyLE n] : + ∃ (M : C), Nonempty (K ≅ (single _ _ n).obj M) := + ⟨K.X n, ⟨{ + hom := mkHomToSingle (𝟙 _) (fun i (hi : i + 1 = n) ↦ + (K.isZero_of_isStrictlyGE n i (by omega)).eq_of_src _ _) + inv := mkHomFromSingle (𝟙 _) (fun i (hi : n + 1 = i) ↦ + (K.isZero_of_isStrictlyLE n i (by omega)).eq_of_tgt _ _) + hom_inv_id := by + ext i + obtain hi | rfl | hi := lt_trichotomy i n + · apply (K.isZero_of_isStrictlyGE n i (by omega)).eq_of_src + · simp + · apply (K.isZero_of_isStrictlyLE n i (by omega)).eq_of_tgt + inv_hom_id := by aesop }⟩⟩ + +end HasZeroMorphisms + +end CochainComplex diff --git a/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean b/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean index 5962382e8c31d..2d5d00e3d6d6f 100644 --- a/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean +++ b/Mathlib/Algebra/Homology/Embedding/ExtendHomology.lean @@ -241,6 +241,15 @@ noncomputable def rightHomologyData (h : (K.sc' i j k).RightHomologyData) : exact h.wι hι := isLimitKernelFork K e hj' hi hi' hk hk' _ h.hp _ h.hι +/-- Computation of the `g'` field of `extend.rightHomologyData`. -/ +lemma rightHomologyData_g' (h : (K.sc' i j k).RightHomologyData) (hk'' : e.f k = k') : + (rightHomologyData K e hj' hi hi' hk hk' h).g' = h.g' ≫ (K.extendXIso e hk'').inv := by + rw [← cancel_epi h.p, ← cancel_epi (extendXIso K e hj').hom] + have := (rightHomologyData K e hj' hi hi' hk hk' h).p_g' + dsimp at this + rw [assoc] at this + rw [this, K.extend_d_eq e hj' hk'', h.p_g'_assoc, shortComplexFunctor'_obj_g] + /-- The homology data of `(K.extend e).sc' i' j' k'` that is deduced from a homology data of `K.sc' i j k`. -/ @[simps] diff --git a/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean b/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean index fff53bfea4947..d39f0d57c6052 100644 --- a/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean +++ b/Mathlib/Algebra/Homology/Embedding/TruncGEHomology.lean @@ -3,6 +3,7 @@ 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.Homology.Embedding.ExtendHomology import Mathlib.Algebra.Homology.Embedding.TruncGE import Mathlib.Algebra.Homology.Embedding.RestrictionHomology import Mathlib.Algebra.Homology.QuasiIso @@ -11,9 +12,11 @@ import Mathlib.Algebra.Homology.QuasiIso Given an embedding of complex shapes `e : Embedding c c'`, we shall relate the homology of `K : HomologicalComplex C c'` and of -`K.truncGE e : HomologicalComplex C c'` (TODO). +`K.truncGE e : HomologicalComplex C c'`. -So far, we only compute the homology of `K.truncGE' e : HomologicalComplex C c`. +The main result is that `K.πTruncGE e : K ⟶ K.truncGE e` induces a +quasi-isomorphism in degree `e.f i` for all `i`. (Note that the complex +`K.truncGE e` is exact in degrees that are not in the image of `e.f`.) -/ @@ -23,8 +26,8 @@ namespace HomologicalComplex variable {ι ι' : Type*} {c : ComplexShape ι} {c' : ComplexShape ι'} {C : Type*} [Category C] [HasZeroMorphisms C] - (K : HomologicalComplex C c') (e : c.Embedding c') [e.IsTruncGE] - [∀ i', K.HasHomology i'] + (K L : HomologicalComplex C c') (φ : K ⟶ L) (e : c.Embedding c') [e.IsTruncGE] + [∀ i', K.HasHomology i'] [∀ i', L.HasHomology i'] namespace truncGE' @@ -109,4 +112,123 @@ instance truncGE'_hasHomology (i : ι) : (K.truncGE' e).HasHomology i := by end truncGE' +variable [HasZeroObject C] + +namespace truncGE + +instance (i' : ι') : (K.truncGE e).HasHomology i' := by + dsimp [truncGE] + infer_instance + +/-- The right homology data which allows to show that `K.πTruncGE e` +induces an isomorphism in homology in degrees `j'` such that `e.f j = j'` for some `j`. -/ +@[simps] +noncomputable def rightHomologyMapData {i j k : ι} {j' : ι'} (hj' : e.f j = j') + (hi : c.prev j = i) (hk : c.next j = k) (hj : e.BoundaryGE j) : + ShortComplex.RightHomologyMapData ((shortComplexFunctor C c' j').map (K.πTruncGE e)) + (ShortComplex.RightHomologyData.canonical (K.sc j')) + (extend.rightHomologyData (K.truncGE' e) e hj' hi rfl hk rfl + (truncGE'.homologyData K e i j k hk hj' hj).right) where + φQ := (K.truncGE'XIsoOpcycles e hj' hj).inv + φH := 𝟙 _ + commp := by + change K.pOpcycles j' ≫ _ = _ + simp [truncGE'.homologyData, πTruncGE, e.liftExtend_f _ _ hj', + K.restrictionToTruncGE'_f_eq_iso_hom_pOpcycles_iso_inv e hj' hj] + commg' := by + have hk' : e.f k = c'.next j' := by rw [← hj', e.next_f hk] + dsimp + rw [extend.rightHomologyData_g' _ _ _ _ _ _ _ _ hk', πTruncGE, + e.liftExtend_f _ _ hk', truncGE'.homologyData_right_g'] + by_cases hjk : c.Rel j k + · simp [K.truncGE'_d_eq_fromOpcycles e hjk hj' hk' hj, + K.restrictionToTruncGE'_f_eq_iso_hom_iso_inv e hk' (e.not_boundaryGE_next hjk)] + rfl + · obtain rfl : k = j := by rw [← c.next_eq_self j (by simpa only [hk] using hjk), hk] + rw [shape _ _ _ hjk, zero_comp, comp_zero, + K.restrictionToTruncGE'_f_eq_iso_hom_pOpcycles_iso_inv e hk' hj] + simp only [restriction_X, restrictionXIso, eqToIso.inv, eqToIso.hom, assoc, + eqToHom_trans_assoc, eqToHom_refl, id_comp] + change 0 = K.fromOpcycles _ _ ≫ _ + rw [← cancel_epi (K.pOpcycles _), comp_zero, p_fromOpcycles_assoc, + d_pOpcycles_assoc, zero_comp] + +end truncGE + +lemma quasiIsoAt_πTruncGE {j : ι} {j' : ι'} (hj' : e.f j = j') : + QuasiIsoAt (K.πTruncGE e) j' := by + rw [quasiIsoAt_iff] + by_cases hj : e.BoundaryGE j + · rw [(truncGE.rightHomologyMapData K e hj' rfl rfl hj).quasiIso_iff] + dsimp + infer_instance + · let φ := (shortComplexFunctor C c' j').map (K.πTruncGE e) + have : Epi φ.τ₁ := by + by_cases hi : ∃ i, e.f i = c'.prev j' + · obtain ⟨i, hi⟩ := hi + dsimp [φ, πTruncGE] + rw [e.epi_liftExtend_f_iff _ _ hi] + infer_instance + · apply IsZero.epi (isZero_extend_X _ _ _ (by simpa using hi)) + have : IsIso φ.τ₂ := by + dsimp [φ, πTruncGE] + rw [e.isIso_liftExtend_f_iff _ _ hj'] + exact K.isIso_restrictionToTruncGE' e j hj + have : IsIso φ.τ₃ := by + dsimp [φ, πTruncGE] + have : c'.next j' = e.f (c.next j) := by simpa only [← hj'] using e.next_f rfl + rw [e.isIso_liftExtend_f_iff _ _ this.symm] + exact K.isIso_restrictionToTruncGE' e _ (e.not_boundaryGE_next' hj rfl) + exact ShortComplex.quasiIso_of_epi_of_isIso_of_mono φ + +instance (i : ι) : QuasiIsoAt (K.πTruncGE e) (e.f i) := K.quasiIsoAt_πTruncGE e rfl + +lemma quasiIso_πTruncGE_iff_isSupported : + QuasiIso (K.πTruncGE e) ↔ K.IsSupported e := by + constructor + · intro + refine ⟨fun i' hi' => ?_⟩ + rw [exactAt_iff_of_quasiIsoAt (K.πTruncGE e) i'] + exact (K.truncGE e).exactAt_of_isSupported e i' hi' + · intro + rw [quasiIso_iff] + intro i' + by_cases hi' : ∃ i, e.f i = i' + · obtain ⟨i, rfl⟩ := hi' + infer_instance + · rw [quasiIsoAt_iff_exactAt (K.πTruncGE e) i'] + all_goals exact exactAt_of_isSupported _ e i' (by simpa using hi') + +lemma acyclic_truncGE_iff_isSupportedOutside : + (K.truncGE e).Acyclic ↔ K.IsSupportedOutside e := by + constructor + · intro hK + exact ⟨fun i => + by simpa only [exactAt_iff_of_quasiIsoAt (K.πTruncGE e)] using hK (e.f i)⟩ + · intro hK i' + by_cases hi' : ∃ i, e.f i = i' + · obtain ⟨i, rfl⟩ := hi' + simpa only [← exactAt_iff_of_quasiIsoAt (K.πTruncGE e)] using hK.exactAt i + · exact exactAt_of_isSupported _ e i' (by simpa using hi') + +variable {K L} + +lemma quasiIso_truncGEMap_iff : + QuasiIso (truncGEMap φ e) ↔ ∀ (i : ι) (i' : ι') (_ : e.f i = i'), QuasiIsoAt φ i' := by + have : ∀ (i : ι) (i' : ι') (_ : e.f i = i'), + QuasiIsoAt (truncGEMap φ e) i' ↔ QuasiIsoAt φ i' := by + rintro i _ rfl + rw [← quasiIsoAt_iff_comp_left (K.πTruncGE e), πTruncGE_naturality φ e, + quasiIsoAt_iff_comp_right] + rw [quasiIso_iff] + constructor + · intro h i i' hi + simpa only [← this i i' hi] using h i' + · intro h i' + by_cases hi' : ∃ i, e.f i = i' + · obtain ⟨i, hi⟩ := hi' + simpa only [this i i' hi] using h i i' hi + · rw [quasiIsoAt_iff_exactAt] + all_goals exact exactAt_of_isSupported _ e i' (by simpa using hi') + end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/Homotopy.lean b/Mathlib/Algebra/Homology/Homotopy.lean index 2749531ddb90a..bb13f21e29877 100644 --- a/Mathlib/Algebra/Homology/Homotopy.lean +++ b/Mathlib/Algebra/Homology/Homotopy.lean @@ -112,7 +112,6 @@ theorem prevD_nat (C D : CochainComplex V ℕ) (i : ℕ) (f : ∀ i j, C.X i ⟶ not_false_iff, comp_zero, reduceCtorEq] · congr <;> simp --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[has_nonempty_instance] /-- A homotopy `h` between chain maps `f` and `g` consists of components `h i j : C.X i ⟶ D.X j` which are zero unless `c.Rel j i`, satisfying the homotopy condition. -/ diff --git a/Mathlib/Algebra/Homology/QuasiIso.lean b/Mathlib/Algebra/Homology/QuasiIso.lean index 19e311240e6ed..71ef037c0f1b0 100644 --- a/Mathlib/Algebra/Homology/QuasiIso.lean +++ b/Mathlib/Algebra/Homology/QuasiIso.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Joël Riou -/ import Mathlib.Algebra.Homology.Homotopy +import Mathlib.Algebra.Homology.ShortComplex.Retract +import Mathlib.CategoryTheory.MorphismProperty.Composition /-! # Quasi-isomorphisms @@ -51,6 +53,15 @@ lemma quasiIsoAt_iff' (f : K ⟶ L) (i j k : ι) (hi : c.prev j = i) (hk : c.nex exact ShortComplex.quasiIso_iff_of_arrow_mk_iso _ _ (Arrow.isoOfNatIso (natIsoSc' C c i j k hi hk) (Arrow.mk f)) +lemma quasiIsoAt_of_retract {f : K ⟶ L} {f' : K' ⟶ L'} + (h : RetractArrow f f') (i : ι) [K.HasHomology i] [L.HasHomology i] + [K'.HasHomology i] [L'.HasHomology i] [hf' : QuasiIsoAt f' i] : + QuasiIsoAt f i := by + rw [quasiIsoAt_iff] at hf' ⊢ + have : RetractArrow ((shortComplexFunctor C c i).map f) + ((shortComplexFunctor C c i).map f') := h.map (shortComplexFunctor C c i).mapArrow + exact ShortComplex.quasiIso_of_retract this + lemma quasiIsoAt_iff_isIso_homologyMap (f : K ⟶ L) (i : ι) [K.HasHomology i] [L.HasHomology i] : QuasiIsoAt f i ↔ IsIso (homologyMap f i) := by @@ -79,6 +90,12 @@ lemma quasiIsoAt_iff_exactAt' (f : K ⟶ L) (i : ι) [K.HasHomology i] [L.HasHom · intro hK exact ⟨⟨0, IsZero.eq_of_src hK _ _, IsZero.eq_of_tgt hL _ _⟩⟩ +lemma exactAt_iff_of_quasiIsoAt (f : K ⟶ L) (i : ι) + [K.HasHomology i] [L.HasHomology i] [QuasiIsoAt f i] : + K.ExactAt i ↔ L.ExactAt i := + ⟨fun hK => (quasiIsoAt_iff_exactAt f i hK).1 inferInstance, + fun hL => (quasiIsoAt_iff_exactAt' f i hL).1 inferInstance⟩ + instance (f : K ⟶ L) (i : ι) [K.HasHomology i] [L.HasHomology i] [hf : QuasiIsoAt f i] : IsIso (homologyMap f i) := by simpa only [quasiIsoAt_iff, ShortComplex.quasiIso_iff] using hf @@ -217,6 +234,12 @@ lemma quasiIso_of_arrow_mk_iso (φ : K ⟶ L) (φ' : K' ⟶ L') (e : Arrow.mk φ [hφ : QuasiIso φ] : QuasiIso φ' := by simpa only [← quasiIso_iff_of_arrow_mk_iso φ φ' e] +lemma quasiIso_of_retractArrow {f : K ⟶ L} {f' : K' ⟶ L'} + (h : RetractArrow f f') [∀ i, K.HasHomology i] [∀ i, L.HasHomology i] + [∀ i, K'.HasHomology i] [∀ i, L'.HasHomology i] [QuasiIso f'] : + QuasiIso f where + quasiIsoAt i := quasiIsoAt_of_retract h i + namespace HomologicalComplex section PreservesHomology @@ -268,10 +291,35 @@ variable (C c) def quasiIso [CategoryWithHomology C] : MorphismProperty (HomologicalComplex C c) := fun _ _ f => QuasiIso f -variable {C c} +variable {C c} [CategoryWithHomology C] @[simp] -lemma mem_quasiIso_iff [CategoryWithHomology C] (f : K ⟶ L) : quasiIso C c f ↔ QuasiIso f := by rfl +lemma mem_quasiIso_iff (f : K ⟶ L) : quasiIso C c f ↔ QuasiIso f := by rfl + +instance : (quasiIso C c).IsMultiplicative where + id_mem _ := by + rw [mem_quasiIso_iff] + infer_instance + comp_mem _ _ hf hg := by + rw [mem_quasiIso_iff] at hf hg ⊢ + infer_instance + +instance : (quasiIso C c).HasTwoOutOfThreeProperty where + of_postcomp f g hg hfg := by + rw [mem_quasiIso_iff] at hg hfg ⊢ + rwa [← quasiIso_iff_comp_right f g] + of_precomp f g hf hfg := by + rw [mem_quasiIso_iff] at hf hfg ⊢ + rwa [← quasiIso_iff_comp_left f g] + +instance : (quasiIso C c).IsStableUnderRetracts where + of_retract h hg := by + rw [mem_quasiIso_iff] at hg ⊢ + exact quasiIso_of_retractArrow h + +instance : (quasiIso C c).RespectsIso := + MorphismProperty.respectsIso_of_isStableUnderComposition + (fun _ _ _ (_ : IsIso _) ↦ by rw [mem_quasiIso_iff]; infer_instance) end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean index 393b22f24e512..f1b0f96f3b7a0 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean @@ -37,7 +37,7 @@ variable (S : ShortComplex Ab.{u}) @[simp] lemma ab_zero_apply (x : S.X₁) : S.g (S.f x) = 0 := by - rw [← comp_apply, S.zero] + rw [← CategoryTheory.comp_apply, S.zero] rfl /-- The canonical additive morphism `S.X₁ →+ AddMonoidHom.ker S.g` induced by `S.f`. -/ @@ -81,7 +81,7 @@ noncomputable def abCyclesIso : S.cycles ≅ AddCommGrp.of (AddMonoidHom.ker S.g lemma abCyclesIso_inv_apply_iCycles (x : AddMonoidHom.ker S.g) : S.iCycles (S.abCyclesIso.inv x) = x := by dsimp only [abCyclesIso] - rw [← comp_apply, S.abLeftHomologyData.cyclesIso_inv_comp_iCycles] + rw [← CategoryTheory.comp_apply, S.abLeftHomologyData.cyclesIso_inv_comp_iCycles] rfl /-- Given a short complex `S` of abelian groups, this is the isomorphism between @@ -129,7 +129,7 @@ lemma ab_exact_iff_range_eq_ker : S.Exact ↔ S.f.range = S.g.ker := by · intro h refine le_antisymm ?_ h rintro _ ⟨x₁, rfl⟩ - rw [AddMonoidHom.mem_ker, ← comp_apply, S.zero] + rw [AddMonoidHom.mem_ker, ← CategoryTheory.comp_apply, S.zero] rfl · intro h rw [h] diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean index 92d14ea2564f5..610dbd89fe414 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -32,7 +32,7 @@ lemma ShortComplex.zero_apply [Limits.HasZeroMorphisms C] [(forget₂ C Ab).PreservesZeroMorphisms] (S : ShortComplex C) (x : (forget₂ C Ab).obj S.X₁) : ((forget₂ C Ab).map S.g) (((forget₂ C Ab).map S.f) x) = 0 := by - rw [← comp_apply, ← Functor.map_comp, S.zero, Functor.map_zero] + rw [← CategoryTheory.comp_apply, ← Functor.map_comp, S.zero, Functor.map_zero] rfl section preadditive @@ -112,8 +112,8 @@ lemma i_cyclesMk [S.HasHomology] (x₂ : (forget₂ C Ab).obj S.X₂) (hx₂ : ((forget₂ C Ab).map S.g) x₂ = 0) : (forget₂ C Ab).map S.iCycles (S.cyclesMk x₂ hx₂) = x₂ := by dsimp [cyclesMk] - erw [← comp_apply, S.mapCyclesIso_hom_iCycles (forget₂ C Ab), - ← comp_apply, abCyclesIso_inv_apply_iCycles ] + erw [← CategoryTheory.comp_apply, S.mapCyclesIso_hom_iCycles (forget₂ C Ab), + ← CategoryTheory.comp_apply, abCyclesIso_inv_apply_iCycles ] end ShortComplex @@ -172,13 +172,13 @@ lemma δ_apply' (x₃ : (forget₂ C Ab).obj D.L₀.X₃) · refine ((congr_hom (e.hom.naturality D.L₁.g) x₂).symm.trans ?_).trans (congr_hom (e.hom.naturality D.v₀₁.τ₃) x₃) dsimp - rw [comp_apply, comp_apply] + rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply] erw [h₂] rfl · refine ((congr_hom (e.hom.naturality D.L₂.f) x₁).symm.trans ?_).trans (congr_hom (e.hom.naturality D.v₁₂.τ₂) x₂) dsimp - rw [comp_apply, comp_apply] + rw [CategoryTheory.comp_apply, CategoryTheory.comp_apply] erw [h₁] rfl diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index 97bd4dc2b02fa..6f0f54b5b2b70 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -614,6 +614,24 @@ lemma exactAt_iff_isZero_homology [K.HasHomology i] : dsimp [homology] rw [exactAt_iff, ShortComplex.exact_iff_isZero_homology] +variable {K i} in +lemma ExactAt.isZero_homology [K.HasHomology i] (h : K.ExactAt i) : + IsZero (K.homology i) := by + rwa [← exactAt_iff_isZero_homology] + +/-- A homological complex `K` is acyclic if it is exact at `i` for any `i`. -/ +def Acyclic := ∀ i, K.ExactAt i + +lemma acyclic_iff : + K.Acyclic ↔ ∀ i, K.ExactAt i := by rfl + +lemma acyclic_of_isZero (hK : IsZero K) : + K.Acyclic := by + rw [acyclic_iff] + intro i + apply ShortComplex.exact_of_isZero_X₂ + exact (eval _ _ i).map_isZero hK + end HomologicalComplex namespace ChainComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean index 9f3d4faff0bed..8d34059412da3 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean @@ -1173,6 +1173,53 @@ instance epi_homologyMap_of_epi_cyclesMap Epi (homologyMap φ) := epi_homologyMap_of_epi_cyclesMap' φ inferInstance +/-- Given a short complex `S` such that `S.HasHomology`, this is the canonical +left homology data for `S` whose `K` and `H` fields are +respectively `S.cycles` and `S.homology`. -/ +@[simps!] +noncomputable def LeftHomologyData.canonical [S.HasHomology] : S.LeftHomologyData where + K := S.cycles + H := S.homology + i := S.iCycles + π := S.homologyπ + wi := by simp + hi := S.cyclesIsKernel + wπ := S.toCycles_comp_homologyπ + hπ := S.homologyIsCokernel + +/-- Computation of the `f'` field of `LeftHomologyData.canonical`. -/ +@[simp] +lemma LeftHomologyData.canonical_f' [S.HasHomology] : + (LeftHomologyData.canonical S).f' = S.toCycles := rfl + +/-- Given a short complex `S` such that `S.HasHomology`, this is the canonical +right homology data for `S` whose `Q` and `H` fields are +respectively `S.opcycles` and `S.homology`. -/ +@[simps!] +noncomputable def RightHomologyData.canonical [S.HasHomology] : S.RightHomologyData where + Q := S.opcycles + H := S.homology + p := S.pOpcycles + ι := S.homologyι + wp := by simp + hp := S.opcyclesIsCokernel + wι := S.homologyι_comp_fromOpcycles + hι := S.homologyIsKernel + +/-- Computation of the `g'` field of `RightHomologyData.canonical`. -/ +@[simp] +lemma RightHomologyData.canonical_g' [S.HasHomology] : + (RightHomologyData.canonical S).g' = S.fromOpcycles := rfl + +/-- Given a short complex `S` such that `S.HasHomology`, this is the canonical +homology data for `S` whose `left.K`, `left/right.H` and `right.Q` fields are +respectively `S.cycles`, `S.homology` and `S.opcycles`. -/ +@[simps!] +noncomputable def HomologyData.canonical [S.HasHomology] : S.HomologyData where + left := LeftHomologyData.canonical S + right := RightHomologyData.canonical S + iso := Iso.refl _ + end ShortComplex end CategoryTheory diff --git a/Mathlib/Algebra/Homology/ShortComplex/Retract.lean b/Mathlib/Algebra/Homology/ShortComplex/Retract.lean new file mode 100644 index 0000000000000..6a83271823fd1 --- /dev/null +++ b/Mathlib/Algebra/Homology/ShortComplex/Retract.lean @@ -0,0 +1,39 @@ +/- +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.Homology.ShortComplex.QuasiIso +import Mathlib.CategoryTheory.MorphismProperty.Retract + +/-! +# Quasi-isomorphisms of short complexes are stable under retracts + +-/ + +namespace CategoryTheory + +open Limits + +namespace ShortComplex + +variable {C : Type*} [Category C] [HasZeroMorphisms C] + {S₁ T₁ S₂ T₂ : ShortComplex C} + [S₁.HasHomology] [T₁.HasHomology] [S₂.HasHomology] [T₂.HasHomology] + {f₁ : S₁ ⟶ T₁} {f₂ : S₂ ⟶ T₂} + +lemma quasiIso_of_retract (h : RetractArrow f₁ f₂) [hf₂ : QuasiIso f₂] : + QuasiIso f₁ := by + rw [quasiIso_iff] at hf₂ ⊢ + have h : RetractArrow (homologyMap f₁) (homologyMap f₂) := + { i := Arrow.homMk (u := homologyMap (show S₁ ⟶ S₂ from h.i.left)) + (v := homologyMap (show T₁ ⟶ T₂ from h.i.right)) (by simp [← homologyMap_comp]) + r := Arrow.homMk (u := homologyMap (show S₂ ⟶ S₁ from h.r.left)) + (v := homologyMap (show T₂ ⟶ T₁ from h.r.right)) (by simp [← homologyMap_comp]) + retract := by ext <;> simp [← homologyMap_comp] } + exact (MorphismProperty.isomorphisms C).of_retract h hf₂ + +end ShortComplex + +end CategoryTheory diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index 9d53a393cf7e3..49a2a97527997 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -277,7 +277,8 @@ instance Module.Dual.instLieModule : LieModule R L (M →ₗ[R] R) where end BasicProperties -/-- A morphism of Lie algebras is a linear map respecting the bracket operations. -/ +/-- A morphism of Lie algebras (denoted as `L₁ →ₗ⁅R⁆ L₂`) +is a linear map respecting the bracket operations. -/ structure LieHom (R L L' : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] [LieRing L'] [LieAlgebra R L'] extends L →ₗ[R] L' where /-- A morphism of Lie algebras is compatible with brackets. -/ @@ -470,9 +471,10 @@ theorem LieModule.compLieHom [Module R M] [LieModule R L₂ M] : end ModulePullBack -/-- An equivalence of Lie algebras is a morphism which is also a linear equivalence. We could -instead define an equivalence to be a morphism which is also a (plain) equivalence. However it is -more convenient to define via linear equivalence to get `.toLinearEquiv` for free. -/ +/-- An equivalence of Lie algebras (denoted as `L₁ ≃ₗ⁅R⁆ L₂`) is a morphism +which is also a linear equivalence. +We could instead define an equivalence to be a morphism which is also a (plain) equivalence. +However, it is more convenient to define via linear equivalence to get `.toLinearEquiv` for free. -/ structure LieEquiv (R : Type u) (L : Type v) (L' : Type w) [CommRing R] [LieRing L] [LieAlgebra R L] [LieRing L'] [LieAlgebra R L'] extends L →ₗ⁅R⁆ L' where /-- The inverse function of an equivalence of Lie algebras -/ @@ -642,8 +644,8 @@ variable [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] variable [Module R M] [Module R N] [Module R P] variable [LieRingModule L M] [LieRingModule L N] [LieRingModule L P] -/-- A morphism of Lie algebra modules is a linear map which commutes with the action of the Lie -algebra. -/ +/-- A morphism of Lie algebra modules (denoted as `M →ₗ⁅R,L⁆ N`) is a linear map +which commutes with the action of the Lie algebra. -/ structure LieModuleHom extends M →ₗ[R] N where /-- A module of Lie algebra modules is compatible with the action of the Lie algebra on the modules. -/ @@ -857,8 +859,8 @@ instance : Module R (M →ₗ⁅R,L⁆ N) := end LieModuleHom -/-- An equivalence of Lie algebra modules is a linear equivalence which is also a morphism of -Lie algebra modules. -/ +/-- An equivalence of Lie algebra modules (denoted as `M ≃ₗ⁅R,L⁆ N`) is a linear equivalence +which is also a morphism of Lie algebra modules. -/ structure LieModuleEquiv extends M →ₗ⁅R,L⁆ N where /-- The inverse function of an equivalence of Lie modules -/ invFun : N → M diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index 73c1e4f1312b4..2c4915771269b 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -410,7 +410,7 @@ alias rootSystem_toLin_apply := rootSystem_toPerfectPairing_apply @[simp] lemma rootSystem_coroot_apply (α) : (rootSystem H).coroot α = coroot α := rfl instance : (rootSystem H).IsCrystallographic where - exists_int α β := + exists_value α β := ⟨chainBotCoeff β.1 α.1 - chainTopCoeff β.1 α.1, by simp [apply_coroot_eq_cast β.1 α.1]⟩ theorem isReduced_rootSystem : (rootSystem H).IsReduced := by diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index b0228d7304c79..ddbbd94475a75 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -482,6 +482,41 @@ theorem symm_neg : (neg R : M ≃ₗ[R] M).symm = neg R := end Neg +section Semiring + +open LinearMap + +variable {M₂₁ M₂₂ : Type*} [Semiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] + [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] + +/-- A linear isomorphism between the domains and codomains of two spaces of linear maps gives a +additive isomorphism between the two function spaces. + +See also `LinearEquiv.arrowCongr` for the linear version of this isomorphism. -/ +def arrowCongrAddEquiv (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) : + (M₁ →ₗ[R] M₂₁) ≃+ (M₂ →ₗ[R] M₂₂) where + toFun f := e₂.comp (f.comp e₁.symm.toLinearMap) + invFun f := e₂.symm.comp (f.comp e₁.toLinearMap) + left_inv f := by + ext x + simp only [symm_apply_apply, Function.comp_apply, coe_comp, coe_coe] + right_inv f := by + ext x + simp only [Function.comp_apply, apply_symm_apply, coe_comp, coe_coe] + map_add' f g := by + ext x + simp only [map_add, add_apply, Function.comp_apply, coe_comp, coe_coe] + +/-- If `M` and `M₂` are linearly isomorphic then the endomorphism rings of `M` and `M₂` +are isomorphic. + +See `LinearEquiv.conj` for the linear version of this isomorphism. -/ +def conjRingEquiv (e : M₁ ≃ₗ[R] M₂) : Module.End R M₁ ≃+* Module.End R M₂ where + __ := arrowCongrAddEquiv e e + map_mul' _ _ := by ext; simp [arrowCongrAddEquiv] + +end Semiring + section CommSemiring variable [CommSemiring R] [AddCommMonoid M] [AddCommMonoid M₂] [AddCommMonoid M₃] @@ -494,24 +529,15 @@ def smulOfUnit (a : Rˣ) : M ≃ₗ[R] M := DistribMulAction.toLinearEquiv R M a /-- A linear isomorphism between the domains and codomains of two spaces of linear maps gives a -linear isomorphism between the two function spaces. -/ +linear isomorphism between the two function spaces. + +See `LinearEquiv.arrowCongrAddEquiv` for the additive version of this isomorphism that works +over a not necessarily commutative semiring. -/ def arrowCongr {R M₁ M₂ M₂₁ M₂₂ : Sort _} [CommSemiring R] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₂₁] [AddCommMonoid M₂₂] [Module R M₁] [Module R M₂] [Module R M₂₁] [Module R M₂₂] (e₁ : M₁ ≃ₗ[R] M₂) (e₂ : M₂₁ ≃ₗ[R] M₂₂) : (M₁ →ₗ[R] M₂₁) ≃ₗ[R] M₂ →ₗ[R] M₂₂ where - toFun := fun f : M₁ →ₗ[R] M₂₁ ↦ (e₂ : M₂₁ →ₗ[R] M₂₂).comp <| f.comp (e₁.symm : M₂ →ₗ[R] M₁) - invFun f := (e₂.symm : M₂₂ →ₗ[R] M₂₁).comp <| f.comp (e₁ : M₁ →ₗ[R] M₂) - left_inv f := by - ext x - simp only [symm_apply_apply, Function.comp_apply, coe_comp, coe_coe] - right_inv f := by - ext x - simp only [Function.comp_apply, apply_symm_apply, coe_comp, coe_coe] - map_add' f g := by - ext x - simp only [map_add, add_apply, Function.comp_apply, coe_comp, coe_coe] - map_smul' c f := by - ext x - simp only [smul_apply, Function.comp_apply, coe_comp, map_smulₛₗ e₂, coe_coe] + __ := arrowCongrAddEquiv e₁ e₂ + map_smul' c f := by ext; simp [arrowCongrAddEquiv] @[simp] theorem arrowCongr_apply {R M₁ M₂ M₂₁ M₂₂ : Sort _} [CommSemiring R] [AddCommMonoid M₁] @@ -547,7 +573,11 @@ def congrRight (f : M₂ ≃ₗ[R] M₃) : (M →ₗ[R] M₂) ≃ₗ[R] M →ₗ arrowCongr (LinearEquiv.refl R M) f /-- If `M` and `M₂` are linearly isomorphic then the two spaces of linear maps from `M` and `M₂` to -themselves are linearly isomorphic. -/ +themselves are linearly isomorphic. + +See `LinearEquiv.conjRingEquiv` for the isomorphism between endomorphism rings, +which works over a not necessarily commutative semiring. -/ +-- TODO: upgrade to AlgEquiv (but this file currently cannot import AlgEquiv) def conj (e : M ≃ₗ[R] M₂) : Module.End R M ≃ₗ[R] Module.End R M₂ := arrowCongr e e @@ -582,12 +612,8 @@ isomorphism between `M₂ →ₗ[R] M` and `M₃ →ₗ[R] M`, if `M` is both an `S`-module and their actions commute. -/ def congrLeft {R} (S) [Semiring R] [Semiring S] [Module R M₂] [Module R M₃] [Module R M] [Module S M] [SMulCommClass R S M] (e : M₂ ≃ₗ[R] M₃) : (M₂ →ₗ[R] M) ≃ₗ[S] (M₃ →ₗ[R] M) where - toFun f := f.comp e.symm.toLinearMap - invFun f := f.comp e.toLinearMap - map_add' _ _ := rfl + __ := e.arrowCongrAddEquiv (.refl ..) map_smul' _ _ := rfl - left_inv f := by dsimp only; apply DFunLike.ext; exact (congr_arg f <| e.left_inv ·) - right_inv f := by dsimp only; apply DFunLike.ext; exact (congr_arg f <| e.right_inv ·) end CommSemiring diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 678ac96dd6526..e2acc6cc59c16 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -66,11 +66,11 @@ add_decl_doc LinearEquiv.right_inv /-- `LinearEquiv.invFun` is a left inverse to the linear equivalence's underlying function. -/ add_decl_doc LinearEquiv.left_inv -/-- The notation `M ≃ₛₗ[σ] M₂` denotes the type of linear equivalences between `M` and `M₂` over a +/-- `M ≃ₛₗ[σ] M₂` denotes the type of linear equivalences between `M` and `M₂` over a ring homomorphism `σ`. -/ notation:50 M " ≃ₛₗ[" σ "] " M₂ => LinearEquiv σ M M₂ -/-- The notation `M ≃ₗ [R] M₂` denotes the type of linear equivalences between `M` and `M₂` over +/-- `M ≃ₗ [R] M₂` denotes the type of linear equivalences between `M` and `M₂` over a plain linear map `M →ₗ M₂`. -/ notation:50 M " ≃ₗ[" R "] " M₂ => LinearEquiv (RingHom.id R) M M₂ @@ -297,7 +297,7 @@ def trans (e₁₂ : M₁ ≃ₛₗ[σ₁₂] M₂) (e₂₃ : M₂ ≃ₛₗ[σ₂₃] M₃) : M₁ ≃ₛₗ[σ₁₃] M₃ := { e₂₃.toLinearMap.comp e₁₂.toLinearMap, e₁₂.toEquiv.trans e₂₃.toEquiv with } -/-- The notation `e₁ ≪≫ₗ e₂` denotes the composition of the linear equivalences `e₁` and `e₂`. -/ +/-- `e₁ ≪≫ₗ e₂` denotes the composition of the linear equivalences `e₁` and `e₂`. -/ notation3:80 (name := transNotation) e₁:80 " ≪≫ₗ " e₂:81 => @LinearEquiv.trans _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ (RingHom.id _) (RingHom.id _) (RingHom.id _) (RingHom.id _) (RingHom.id _) (RingHom.id _) RingHomCompTriple.ids RingHomCompTriple.ids diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index 4c43da972ed6a..51af8af5b2d8b 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -241,6 +241,35 @@ lemma Module.finitePresentation_of_ker [Module.FinitePresentation R N] LinearMap.ker_eq_bot.mpr (Submodule.injective_subtype (LinearMap.ker l)), Submodule.comap_bot] exact (Module.FinitePresentation.fg_ker f hf).map (Submodule.subtype _) +/-- Given a split exact sequence `0 → M → N → P → 0` with `N` finitely presented, +then `M` is also finitely presented. -/ +lemma Module.finitePresentation_of_split_exact + {P : Type*} [AddCommGroup P] [Module R P] + [Module.FinitePresentation R N] + (f : M →ₗ[R] N) (g : N →ₗ[R] P) (l : P →ₗ[R] N) (hl : g ∘ₗ l = .id) + (hf : Function.Injective f) (H : Function.Exact f g) : + Module.FinitePresentation R M := by + have hg : Function.Surjective g := Function.LeftInverse.surjective (DFunLike.congr_fun hl) + have := Module.Finite.of_surjective g hg + obtain ⟨e, rfl, rfl⟩ := ((Function.Exact.split_tfae' H).out 0 2 rfl rfl).mp + ⟨hf, l, hl⟩ + refine Module.finitePresentation_of_surjective (LinearMap.fst _ _ _ ∘ₗ e.toLinearMap) + (Prod.fst_surjective.comp e.surjective) ?_ + rw [LinearMap.ker_comp, Submodule.comap_equiv_eq_map_symm, + LinearMap.exact_iff.mp Function.Exact.inr_fst, ← Submodule.map_top] + exact .map _ (.map _ (Module.Finite.fg_top)) + +/-- Given an exact sequence `0 → M → N → P → 0` +with `N` finitely presented and `P` projective, then `M` is also finitely presented. -/ +lemma Module.finitePresentation_of_projective_of_exact + {P : Type*} [AddCommGroup P] [Module R P] + [Module.FinitePresentation R N] [Module.Projective R P] + (f : M →ₗ[R] N) (g : N →ₗ[R] P) + (hf : Function.Injective f) (hg : Function.Surjective g) (H : Function.Exact f g) : + Module.FinitePresentation R M := + have ⟨l, hl⟩ := Module.projective_lifting_property g .id hg + Module.finitePresentation_of_split_exact f g l hl hf H + end Ring section CommRing diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index b806dbdd9c38e..4a3b1f34b2439 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -72,7 +72,6 @@ example {R} [CommSemiring R] (S : Submonoid R) : ⇑(Localization.r' S) = Locali /-- If `S` is a multiplicative subset of a ring `R` and `M` an `R`-module, then we can localize `M` by `S`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): @[nolint has_nonempty_instance] def _root_.LocalizedModule : Type max u v := Quotient (r.setoid S M) diff --git a/Mathlib/Algebra/Module/Presentation/Cokernel.lean b/Mathlib/Algebra/Module/Presentation/Cokernel.lean index ceaa3b65c0451..d9f2c84e6cd05 100644 --- a/Mathlib/Algebra/Module/Presentation/Cokernel.lean +++ b/Mathlib/Algebra/Module/Presentation/Cokernel.lean @@ -71,7 +71,7 @@ and one for each generator of `M₁`. -/ def cokernelRelations : Relations A where G := pres₂.G R := Sum pres₂.R ι - relation x := match x with + relation | .inl r => pres₂.relation r | .inr i => data.lift i diff --git a/Mathlib/Algebra/Module/Presentation/Tautological.lean b/Mathlib/Algebra/Module/Presentation/Tautological.lean index 72fa47899f48e..b73d5ed63d3bf 100644 --- a/Mathlib/Algebra/Module/Presentation/Tautological.lean +++ b/Mathlib/Algebra/Module/Presentation/Tautological.lean @@ -33,7 +33,7 @@ inductive tautological.R noncomputable def tautologicalRelations : Relations A where G := M R := tautological.R A M - relation r := match r with + relation | .add m₁ m₂ => Finsupp.single m₁ 1 + Finsupp.single m₂ 1 - Finsupp.single (m₁ + m₂) 1 | .smul a m => a • Finsupp.single m 1 - Finsupp.single (a • m) 1 diff --git a/Mathlib/Algebra/Module/Presentation/Tensor.lean b/Mathlib/Algebra/Module/Presentation/Tensor.lean index f1a1b39c41ac8..c941b59b9f7e6 100644 --- a/Mathlib/Algebra/Module/Presentation/Tensor.lean +++ b/Mathlib/Algebra/Module/Presentation/Tensor.lean @@ -32,7 +32,7 @@ noncomputable def tensor : Relations A where G := relations₁.G × relations₂.G R := Sum (relations₁.R × relations₂.G) (relations₁.G × relations₂.R) - relation r := match r with + relation | .inl ⟨r₁, g₂⟩ => Finsupp.embDomain (Function.Embedding.sectL relations₁.G g₂) (relations₁.relation r₁) | .inr ⟨g₁, r₂⟩ => Finsupp.embDomain (Function.Embedding.sectR g₁ relations₂.G) diff --git a/Mathlib/Algebra/Module/Submodule/Defs.lean b/Mathlib/Algebra/Module/Submodule/Defs.lean index 6430f82524dff..50ffade02ba27 100644 --- a/Mathlib/Algebra/Module/Submodule/Defs.lean +++ b/Mathlib/Algebra/Module/Submodule/Defs.lean @@ -5,6 +5,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Group.Subgroup.Defs import Mathlib.GroupTheory.GroupAction.SubMulAction +import Mathlib.Algebra.Group.Submonoid.Basic /-! @@ -137,7 +138,7 @@ variable [Semiring R] [AddCommMonoid M] [Module R M] {A : Type*} [SetLike A M] -- Prefer subclasses of `Module` over `SMulMemClass`. /-- A submodule of a `Module` is a `Module`. -/ -instance (priority := 75) toModule : Module R S' := +instance (priority := 75) toModule : Module R S' := fast_instance% Subtype.coe_injective.module R (AddSubmonoidClass.subtype S') (SetLike.val_smul S') /-- This can't be an instance because Lean wouldn't know how to find `R`, but we can still use @@ -246,10 +247,11 @@ theorem coe_mem (x : p) : (x : M) ∈ p := variable (p) -instance addCommMonoid : AddCommMonoid p := +instance addCommMonoid : AddCommMonoid p := fast_instance% { p.toAddSubmonoid.toAddCommMonoid with } -instance module' [Semiring S] [SMul S R] [Module S M] [IsScalarTower S R M] : Module S p := +instance module' [Semiring S] [SMul S R] [Module S M] [IsScalarTower S R M] : + Module S p := fast_instance% { (show MulAction S p from p.toSubMulAction.mulAction') with smul_zero := fun a => by ext; simp zero_smul := fun a => by ext; simp @@ -319,7 +321,7 @@ theorem sub_mem_iff_left (hy : y ∈ p) : x - y ∈ p ↔ x ∈ p := by theorem sub_mem_iff_right (hx : x ∈ p) : x - y ∈ p ↔ y ∈ p := by rw [sub_eq_add_neg, p.add_mem_iff_right hx, p.neg_mem_iff] -instance addCommGroup : AddCommGroup p := +instance addCommGroup : AddCommGroup p := fast_instance% { p.toAddSubgroup.toAddCommGroup with } end AddCommGroup diff --git a/Mathlib/Algebra/Module/Submodule/Order.lean b/Mathlib/Algebra/Module/Submodule/Order.lean index 25d64e4e3fe8e..25cb72fc29c8d 100644 --- a/Mathlib/Algebra/Module/Submodule/Order.lean +++ b/Mathlib/Algebra/Module/Submodule/Order.lean @@ -18,24 +18,24 @@ variable [Semiring R] /-- A submodule of an `OrderedAddCommMonoid` is an `OrderedAddCommMonoid`. -/ instance toOrderedAddCommMonoid [OrderedAddCommMonoid M] [Module R M] (S : Submodule R M) : - OrderedAddCommMonoid S := + OrderedAddCommMonoid S := fast_instance% Subtype.coe_injective.orderedAddCommMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- A submodule of a `LinearOrderedAddCommMonoid` is a `LinearOrderedAddCommMonoid`. -/ instance toLinearOrderedAddCommMonoid [LinearOrderedAddCommMonoid M] [Module R M] - (S : Submodule R M) : LinearOrderedAddCommMonoid S := + (S : Submodule R M) : LinearOrderedAddCommMonoid S := fast_instance% Subtype.coe_injective.linearOrderedAddCommMonoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl /-- A submodule of an `OrderedCancelAddCommMonoid` is an `OrderedCancelAddCommMonoid`. -/ instance toOrderedCancelAddCommMonoid [OrderedCancelAddCommMonoid M] [Module R M] - (S : Submodule R M) : OrderedCancelAddCommMonoid S := + (S : Submodule R M) : OrderedCancelAddCommMonoid S := fast_instance% Subtype.coe_injective.orderedCancelAddCommMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- A submodule of a `LinearOrderedCancelAddCommMonoid` is a `LinearOrderedCancelAddCommMonoid`. -/ instance toLinearOrderedCancelAddCommMonoid [LinearOrderedCancelAddCommMonoid M] [Module R M] - (S : Submodule R M) : LinearOrderedCancelAddCommMonoid S := + (S : Submodule R M) : LinearOrderedCancelAddCommMonoid S := fast_instance% Subtype.coe_injective.linearOrderedCancelAddCommMonoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -46,14 +46,14 @@ variable [Ring R] /-- A submodule of an `OrderedAddCommGroup` is an `OrderedAddCommGroup`. -/ instance toOrderedAddCommGroup [OrderedAddCommGroup M] [Module R M] (S : Submodule R M) : - OrderedAddCommGroup S := + OrderedAddCommGroup S := fast_instance% Subtype.coe_injective.orderedAddCommGroup Subtype.val rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl /-- A submodule of a `LinearOrderedAddCommGroup` is a `LinearOrderedAddCommGroup`. -/ instance toLinearOrderedAddCommGroup [LinearOrderedAddCommGroup M] [Module R M] - (S : Submodule R M) : LinearOrderedAddCommGroup S := + (S : Submodule R M) : LinearOrderedAddCommGroup S := fast_instance% Subtype.coe_injective.linearOrderedAddCommGroup Subtype.val rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 7443d4cffa6a7..099a6f0601239 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Module.ZMod import Mathlib.GroupTheory.Torsion import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.RingTheory.Coprime.Ideal +import Mathlib.RingTheory.Ideal.Quotient.Defs /-! # Torsion submodules @@ -865,7 +866,7 @@ namespace AddSubgroup variable (A : Type*) [AddCommGroup A] (n : ℤ) -/-- The additive `n`-torsion subgroup for an integer `n`. -/ +/-- The additive `n`-torsion subgroup for an integer `n`, denoted as `A[n]`. -/ @[reducible] def torsionBy : AddSubgroup A := (Submodule.torsionBy ℤ A n).toAddSubgroup diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index fbea78035339d..4ed491fbb056d 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -139,6 +139,17 @@ theorem single_algebraMap_eq_algebraMap_mul_of {A : Type*} [CommSemiring k] [Sem [Algebra k A] [Monoid G] (a : G) (b : k) : single a (algebraMap k A b) = algebraMap k (MonoidAlgebra A G) b * of A G a := by simp +instance isLocalHom_singleOneAlgHom + {A : Type*} [CommSemiring k] [Semiring A] [Algebra k A] [Monoid G] : + IsLocalHom (singleOneAlgHom : A →ₐ[k] MonoidAlgebra A G) where + map_nonunit := isLocalHom_singleOneRingHom.map_nonunit + +instance isLocalHom_algebraMap + {A : Type*} [CommSemiring k] [Semiring A] [Algebra k A] [Monoid G] + [IsLocalHom (algebraMap k A)] : + IsLocalHom (algebraMap k (MonoidAlgebra A G)) where + map_nonunit _ hx := .of_map _ _ <| isLocalHom_singleOneAlgHom (k := k).map_nonunit _ hx + end Algebra section lift @@ -410,6 +421,15 @@ theorem coe_algebraMap [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] (algebraMap R k[G] : R → k[G]) = single 0 ∘ algebraMap R k := rfl +instance isLocalHom_singleZeroAlgHom [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] : + IsLocalHom (singleZeroAlgHom : k →ₐ[R] k[G]) where + map_nonunit := isLocalHom_singleZeroRingHom.map_nonunit + +instance isLocalHom_algebraMap [CommSemiring R] [Semiring k] [Algebra R k] [AddMonoid G] + [IsLocalHom (algebraMap R k)] : + IsLocalHom (algebraMap R k[G]) where + map_nonunit _ hx := .of_map _ _ <| isLocalHom_singleZeroAlgHom (R := R).map_nonunit _ hx + end Algebra section lift diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 337f590109744..9876076f4945c 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -381,6 +381,11 @@ def comapDistribMulActionSelf [Group G] [Semiring k] : DistribMulAction G (Monoi end DerivedInstances +@[simp] +theorem smul_single [Semiring k] [SMulZeroClass R k] (a : G) (c : R) (b : k) : + c • single a b = single a (c • b) := + Finsupp.smul_single _ _ _ + /-! #### Copies of `ext` lemmas and bundled `single`s from `Finsupp` @@ -535,7 +540,6 @@ def of [MulOneClass G] : G →* MonoidAlgebra k G := end /-- Copy of `Finsupp.smul_single'` that avoids the `MonoidAlgebra = Finsupp` defeq abuse. -/ -@[simp] theorem smul_single' (c : k) (a : G) (b : k) : c • single a b = single a (c * b) := Finsupp.smul_single' c a b @@ -562,22 +566,22 @@ def singleHom [MulOneClass G] : k × G →* MonoidAlgebra k G where map_mul' _a _b := single_mul_single.symm theorem mul_single_apply_aux [Mul G] (f : MonoidAlgebra k G) {r : k} {x y z : G} - (H : ∀ a, a * x = z ↔ a = y) : (f * single x r) z = f y * r := by + (H : ∀ a ∈ f.support, a * x = z ↔ a = y) : (f * single x r) z = f y * r := by classical exact - have A : - ∀ a₁ b₁, - ((single x r).sum fun a₂ b₂ => ite (a₁ * a₂ = z) (b₁ * b₂) 0) = - ite (a₁ * x = z) (b₁ * r) 0 := - fun a₁ b₁ => sum_single_index <| by simp - calc - (HMul.hMul (β := MonoidAlgebra k G) f (single x r)) z = - sum f fun a b => if a = y then b * r else 0 := by simp only [mul_apply, A, H] - _ = if y ∈ f.support then f y * r else 0 := f.support.sum_ite_eq' _ _ - _ = f y * r := by split_ifs with h <;> simp at h <;> simp [h] + calc + (f * single x r) z + _ = sum f fun a b => ite (a * x = z) (b * r) 0 := + (mul_apply _ _ _).trans <| Finsupp.sum_congr fun _ _ => sum_single_index (by simp) + + _ = f.sum fun a b => ite (a = y) (b * r) 0 := Finsupp.sum_congr fun x hx => by + simp only [H _ hx] + _ = if y ∈ f.support then f y * r else 0 := f.support.sum_ite_eq' _ _ + _ = _ := by split_ifs with h <;> simp at h <;> simp [h] + theorem mul_single_one_apply [MulOneClass G] (f : MonoidAlgebra k G) (r : k) (x : G) : (HMul.hMul (β := MonoidAlgebra k G) f (single 1 r)) x = f x * r := - f.mul_single_apply_aux fun a => by rw [mul_one] + f.mul_single_apply_aux fun a ha => by rw [mul_one] theorem mul_single_apply_of_not_exists_mul [Mul G] (r : k) {g g' : G} (x : MonoidAlgebra k G) (h : ¬∃ d, g' = d * g) : (x * single g r) g' = 0 := by @@ -592,20 +596,21 @@ theorem mul_single_apply_of_not_exists_mul [Mul G] (r : k) {g g' : G} (x : Monoi exact h ⟨_, rfl⟩ theorem single_mul_apply_aux [Mul G] (f : MonoidAlgebra k G) {r : k} {x y z : G} - (H : ∀ a, x * a = y ↔ a = z) : (single x r * f) y = r * f z := by + (H : ∀ a ∈ f.support, x * a = y ↔ a = z) : (single x r * f) y = r * f z := by classical exact have : (f.sum fun a b => ite (x * a = y) (0 * b) 0) = 0 := by simp calc - (HMul.hMul (α := MonoidAlgebra k G) (single x r) f) y = - sum f fun a b => ite (x * a = y) (r * b) 0 := + (single x r * f) y + _ = sum f fun a b => ite (x * a = y) (r * b) 0 := (mul_apply _ _ _).trans <| sum_single_index this - _ = f.sum fun a b => ite (a = z) (r * b) 0 := by simp only [H] + _ = f.sum fun a b => ite (a = z) (r * b) 0 := Finsupp.sum_congr fun x hx => by + simp only [H _ hx] _ = if z ∈ f.support then r * f z else 0 := f.support.sum_ite_eq' _ _ _ = _ := by split_ifs with h <;> simp at h <;> simp [h] theorem single_one_mul_apply [MulOneClass G] (f : MonoidAlgebra k G) (r : k) (x : G) : (single (1 : G) r * f) x = r * f x := - f.single_mul_apply_aux fun a => by rw [one_mul] + f.single_mul_apply_aux fun a ha => by rw [one_mul] theorem single_mul_apply_of_not_exists_mul [Mul G] (r : k) {g g' : G} (x : MonoidAlgebra k G) (h : ¬∃ d, g' = g * d) : (single g r * x) g' = 0 := by @@ -625,12 +630,8 @@ theorem liftNC_smul [MulOneClass G] {R : Type*} [Semiring R] (f : k →+* R) (g (AddMonoidHom.mulLeft (f c)).comp (liftNC (↑f) g) from DFunLike.congr_fun this φ ext - -- Porting note: `reducible` cannot be `local` so the proof gets more complex. - unfold MonoidAlgebra - simp only [AddMonoidHom.coe_comp, Function.comp_apply, singleAddHom_apply, smulAddHom_apply, - smul_single, smul_eq_mul, AddMonoidHom.coe_mulLeft, Finsupp.singleAddHom_apply] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [liftNC_single, liftNC_single]; rw [AddMonoidHom.coe_coe, map_mul, mul_assoc] + simp_rw [AddMonoidHom.comp_apply, singleAddHom_apply, smulAddHom_apply, + AddMonoidHom.coe_mulLeft, smul_single', liftNC_single, AddMonoidHom.coe_coe, map_mul, mul_assoc] end MiscTheorems @@ -684,7 +685,8 @@ theorem single_one_comm [CommSemiring k] [MulOneClass G] (r : k) (f : MonoidAlge /-- `Finsupp.single 1` as a `RingHom` -/ @[simps] def singleOneRingHom [Semiring k] [MulOneClass G] : k →+* MonoidAlgebra k G := - { Finsupp.singleAddHom 1 with + { singleAddHom 1 with + toFun := single 1 map_one' := rfl map_mul' := fun x y => by simp } @@ -753,12 +755,12 @@ variable [Semiring k] [Group G] @[simp] theorem mul_single_apply (f : MonoidAlgebra k G) (r : k) (x y : G) : (f * single x r) y = f (y * x⁻¹) * r := - f.mul_single_apply_aux fun _a => eq_mul_inv_iff_mul_eq.symm + f.mul_single_apply_aux fun _a _ => eq_mul_inv_iff_mul_eq.symm @[simp] theorem single_mul_apply (r : k) (x : G) (f : MonoidAlgebra k G) (y : G) : (single x r * f) y = r * f (x⁻¹ * y) := - f.single_mul_apply_aux fun _z => eq_inv_mul_iff_mul_eq.symm + f.single_mul_apply_aux fun _z _ => eq_inv_mul_iff_mul_eq.symm theorem mul_apply_left (f g : MonoidAlgebra k G) (x : G) : (f * g) x = f.sum fun a b => b * g (a⁻¹ * x) := @@ -836,6 +838,18 @@ def submoduleOfSMulMem (W : Submodule k V) (h : ∀ (g : G) (v : V), v ∈ W → end Submodule +instance isLocalHom_singleOneRingHom [Semiring k] [Monoid G] : + IsLocalHom (singleOneRingHom (k := k) (G := G)) where + map_nonunit x hx := by + obtain ⟨⟨x, xi, hx, hxi⟩, rfl⟩ := hx + simp_rw [MonoidAlgebra.ext_iff, singleOneRingHom_apply] at hx hxi ⊢ + specialize hx 1 + specialize hxi 1 + classical + simp_rw [single_one_mul_apply, one_def, single_apply, if_pos] at hx + simp_rw [mul_single_one_apply, one_def, single_apply, if_pos] at hxi + exact ⟨⟨x, xi 1, hx, hxi⟩, rfl⟩ + end MonoidAlgebra /-! ### Additive monoids -/ @@ -845,7 +859,7 @@ section variable [Semiring k] -/-- The monoid algebra over a semiring `k` generated by the additive monoid `G`. +/-- The monoid algebra over a semiring `k` generated by the additive monoid `G`, denoted by `G[k]`. It is the type of finite formal `k`-linear combinations of terms of `G`, endowed with the convolution product. -/ @@ -1171,6 +1185,11 @@ because we've never discussed actions of additive groups. -/ end DerivedInstances +@[simp] +theorem smul_single [Semiring k] [SMulZeroClass R k] (a : G) (c : R) (b : k) : + c • single a b = single a (c • b) := + Finsupp.smul_single _ _ _ + /-! #### Copies of `ext` lemmas and bundled `single`s from `Finsupp` @@ -1333,29 +1352,28 @@ def singleHom [AddZeroClass G] : k × Multiplicative G →* k[G] where map_mul' _a _b := single_mul_single.symm /-- Copy of `Finsupp.smul_single'` that avoids the `AddMonoidAlgebra = Finsupp` defeq abuse. -/ -@[simp] theorem smul_single' (c : k) (a : G) (b : k) : c • single a b = single a (c * b) := Finsupp.smul_single' c a b theorem mul_single_apply_aux [Add G] (f : k[G]) (r : k) (x y z : G) - (H : ∀ a, a + x = z ↔ a = y) : (f * single x r) z = f y * r := + (H : ∀ a ∈ f.support, a + x = z ↔ a = y) : (f * single x r) z = f y * r := @MonoidAlgebra.mul_single_apply_aux k (Multiplicative G) _ _ _ _ _ _ _ H theorem mul_single_zero_apply [AddZeroClass G] (f : k[G]) (r : k) (x : G) : (f * single (0 : G) r) x = f x * r := - f.mul_single_apply_aux r _ _ _ fun a => by rw [add_zero] + f.mul_single_apply_aux r _ _ _ fun a _ => by rw [add_zero] theorem mul_single_apply_of_not_exists_add [Add G] (r : k) {g g' : G} (x : k[G]) (h : ¬∃ d, g' = d + g) : (x * single g r) g' = 0 := @MonoidAlgebra.mul_single_apply_of_not_exists_mul k (Multiplicative G) _ _ _ _ _ _ h theorem single_mul_apply_aux [Add G] (f : k[G]) (r : k) (x y z : G) - (H : ∀ a, x + a = y ↔ a = z) : (single x r * f) y = r * f z := + (H : ∀ a ∈ f.support, x + a = y ↔ a = z) : (single x r * f) y = r * f z := @MonoidAlgebra.single_mul_apply_aux k (Multiplicative G) _ _ _ _ _ _ _ H theorem single_zero_mul_apply [AddZeroClass G] (f : k[G]) (r : k) (x : G) : (single (0 : G) r * f) x = r * f x := - f.single_mul_apply_aux r _ _ _ fun a => by rw [zero_add] + f.single_mul_apply_aux r _ _ _ fun a _ => by rw [zero_add] theorem single_mul_apply_of_not_exists_add [Add G] (r : k) {g g' : G} (x : k[G]) (h : ¬∃ d, g' = g + d) : (single g r * x) g' = 0 := @@ -1471,7 +1489,8 @@ section Algebra /-- `Finsupp.single 0` as a `RingHom` -/ @[simps] def singleZeroRingHom [Semiring k] [AddMonoid G] : k →+* k[G] := - { Finsupp.singleAddHom 0 with + { singleAddHom 0 with + toFun := single 0 map_one' := rfl map_mul' := fun x y => by simp only [Finsupp.singleAddHom, single_mul_single, zero_add] } @@ -1551,6 +1570,10 @@ theorem prod_single [CommSemiring k] [AddCommMonoid G] {s : Finset ι} {a : ι end +instance isLocalHom_singleZeroRingHom [Semiring k] [AddMonoid G] : + IsLocalHom (singleZeroRingHom (k := k) (G := G)) := + MonoidAlgebra.isLocalHom_singleOneRingHom (G := Multiplicative G) + end AddMonoidAlgebra set_option linter.style.longFile 1700 diff --git a/Mathlib/Algebra/MonoidAlgebra/Division.lean b/Mathlib/Algebra/MonoidAlgebra/Division.lean index 7b63bd6d1bbd6..d5b413992f933 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Division.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Division.lean @@ -95,13 +95,13 @@ noncomputable def divOfHom : Multiplicative G →* AddMonoid.End k[G] where theorem of'_mul_divOf (a : G) (x : k[G]) : of' k G a * x /ᵒᶠ a = x := by refine Finsupp.ext fun _ => ?_ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` doesn't work rw [AddMonoidAlgebra.divOf_apply, of'_apply, single_mul_apply_aux, one_mul] - intro c + intro c hc exact add_right_inj _ theorem mul_of'_divOf (x : k[G]) (a : G) : x * of' k G a /ᵒᶠ a = x := by refine Finsupp.ext fun _ => ?_ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` doesn't work rw [AddMonoidAlgebra.divOf_apply, of'_apply, mul_single_apply_aux, mul_one] - intro c + intro c hc rw [add_comm] exact add_right_inj _ @@ -160,7 +160,7 @@ theorem divOf_add_modOf (x : k[G]) (g : G) : zero_add] · rw [modOf_apply_self_add, add_zero] rw [of'_apply, single_mul_apply_aux _ _ _, one_mul, divOf_apply] - intro a + intro a ha exact add_right_inj _ theorem modOf_add_divOf (x : k[G]) (g : G) : x %ᵒᶠ g + of' k G g * (x /ᵒᶠ g) = x := by diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index a9fe286b2e396..61c751a168619 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -617,12 +617,12 @@ theorem coeff_mul [DecidableEq σ] (p q : MvPolynomial σ R) (n : σ →₀ ℕ) @[simp] theorem coeff_mul_monomial (m) (s : σ →₀ ℕ) (r : R) (p : MvPolynomial σ R) : coeff (m + s) (p * monomial s r) = coeff m p * r := - AddMonoidAlgebra.mul_single_apply_aux p _ _ _ _ fun _a => add_left_inj _ + AddMonoidAlgebra.mul_single_apply_aux p _ _ _ _ fun _a _ => add_left_inj _ @[simp] theorem coeff_monomial_mul (m) (s : σ →₀ ℕ) (r : R) (p : MvPolynomial σ R) : coeff (s + m) (monomial s r * p) = r * coeff m p := - AddMonoidAlgebra.single_mul_apply_aux p _ _ _ _ fun _a => add_right_inj _ + AddMonoidAlgebra.single_mul_apply_aux p _ _ _ _ fun _a _ => add_right_inj _ @[simp] theorem coeff_mul_X (m) (s : σ) (p : MvPolynomial σ R) : @@ -869,6 +869,104 @@ theorem as_sum (p : MvPolynomial σ R) : p = ∑ v ∈ p.support, monomial v (co end AsSum +section coeffsIn +variable {R S σ : Type*} [CommSemiring R] [CommSemiring S] + +section Module +variable [Module R S] {M N : Submodule R S} {p : MvPolynomial σ S} {s : σ} {i : σ →₀ ℕ} {x : S} + {n : ℕ} + +variable (σ M) in +/-- The `R`-submodule of multivariate polynomials whose coefficients lie in a `R`-submodule `M`. -/ +@[simps] +def coeffsIn : Submodule R (MvPolynomial σ S) where + carrier := {p | ∀ i, p.coeff i ∈ M} + add_mem' := by simp+contextual [add_mem] + zero_mem' := by simp + smul_mem' := by simp+contextual [Submodule.smul_mem] + +lemma mem_coeffsIn : p ∈ coeffsIn σ M ↔ ∀ i, p.coeff i ∈ M := .rfl + +@[simp] +lemma monomial_mem_coeffsIn : monomial i x ∈ coeffsIn σ M ↔ x ∈ M := by + classical + simp only [mem_coeffsIn, coeff_monomial] + exact ⟨fun h ↦ by simpa using h i, fun hs j ↦ by split <;> simp [hs]⟩ + +@[simp] +lemma C_mem_coeffsIn : C x ∈ coeffsIn σ M ↔ x ∈ M := by simpa using monomial_mem_coeffsIn (i := 0) + +@[simp] +lemma one_coeffsIn : 1 ∈ coeffsIn σ M ↔ 1 ∈ M := by simpa using C_mem_coeffsIn (x := (1 : S)) + +@[simp] +lemma mul_monomial_mem_coeffsIn : p * monomial i 1 ∈ coeffsIn σ M ↔ p ∈ coeffsIn σ M := by + classical + simp only [mem_coeffsIn, coeff_mul_monomial', Finsupp.mem_support_iff] + constructor + · rintro hp j + simpa using hp (j + i) + · rintro hp i + split <;> simp [hp] + +@[simp] +lemma monomial_mul_mem_coeffsIn : monomial i 1 * p ∈ coeffsIn σ M ↔ p ∈ coeffsIn σ M := by + simp [mul_comm] + +@[simp] +lemma mul_X_mem_coeffsIn : p * X s ∈ coeffsIn σ M ↔ p ∈ coeffsIn σ M := by + simpa [-mul_monomial_mem_coeffsIn] using mul_monomial_mem_coeffsIn (i := .single s 1) + +@[simp] +lemma X_mul_mem_coeffsIn : X s * p ∈ coeffsIn σ M ↔ p ∈ coeffsIn σ M := by simp [mul_comm] + +variable (M) in +lemma coeffsIn_eq_span_monomial : coeffsIn σ M = .span R {monomial i m | (m ∈ M) (i : σ →₀ ℕ)} := by + classical + refine le_antisymm ?_ <| Submodule.span_le.2 ?_ + · rintro p hp + rw [p.as_sum] + exact sum_mem fun i hi ↦ Submodule.subset_span ⟨_, hp i, _, rfl⟩ + · rintro _ ⟨m, hm, s, n, rfl⟩ i + simp [coeff_X_pow] + split <;> simp [hm] + +lemma coeffsIn_le {N : Submodule R (MvPolynomial σ S)} : + coeffsIn σ M ≤ N ↔ ∀ m ∈ M, ∀ i, monomial i m ∈ N := by + simp [coeffsIn_eq_span_monomial, Submodule.span_le, Set.subset_def, + forall_swap (α := MvPolynomial σ S)] + +end Module + +section Algebra +variable [Algebra R S] {M : Submodule R S} + +lemma coeffsIn_mul (M N : Submodule R S) : coeffsIn σ (M * N) = coeffsIn σ M * coeffsIn σ N := by + classical + refine le_antisymm (coeffsIn_le.2 ?_) ?_ + · intros r hr s + induction hr using Submodule.mul_induction_on' with + | mem_mul_mem m hm n hn => + rw [← add_zero s, ← monomial_mul] + apply Submodule.mul_mem_mul <;> simpa + | add x _ y _ hx hy => + simpa [map_add] using add_mem hx hy + · rw [Submodule.mul_le] + intros x hx y hy k + rw [MvPolynomial.coeff_mul] + exact sum_mem fun c hc ↦ Submodule.mul_mem_mul (hx _) (hy _) + +lemma coeffsIn_pow : ∀ {n}, n ≠ 0 → ∀ M : Submodule R S, coeffsIn σ (M ^ n) = coeffsIn σ M ^ n + | 1, _, M => by simp + | n + 2, _, M => by rw [pow_succ, coeffsIn_mul, coeffsIn_pow, ← pow_succ]; exact n.succ_ne_zero + +lemma le_coeffsIn_pow : ∀ {n}, coeffsIn σ M ^ n ≤ coeffsIn σ (M ^ n) + | 0 => by simpa using ⟨1, map_one _⟩ + | n + 1 => (coeffsIn_pow n.succ_ne_zero _).ge + +end Algebra +end coeffsIn + end CommSemiring end MvPolynomial diff --git a/Mathlib/Algebra/MvPolynomial/CommRing.lean b/Mathlib/Algebra/MvPolynomial/CommRing.lean index 7d2f5a35c441e..8a253650778fc 100644 --- a/Mathlib/Algebra/MvPolynomial/CommRing.lean +++ b/Mathlib/Algebra/MvPolynomial/CommRing.lean @@ -87,9 +87,11 @@ section Degrees theorem degrees_neg (p : MvPolynomial σ R) : (-p).degrees = p.degrees := by rw [degrees, support_neg]; rfl -theorem degrees_sub [DecidableEq σ] (p q : MvPolynomial σ R) : - (p - q).degrees ≤ p.degrees ⊔ q.degrees := by - simpa only [sub_eq_add_neg] using le_trans (degrees_add p (-q)) (by rw [degrees_neg]) +theorem degrees_sub_le [DecidableEq σ] {p q : MvPolynomial σ R} : + (p - q).degrees ≤ p.degrees ∪ q.degrees := by + simpa [degrees_def] using AddMonoidAlgebra.supDegree_sub_le + +@[deprecated (since := "2024-12-28")] alias degrees_sub := degrees_sub_le end Degrees diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 55b5f391e4aff..cbfaf3b83f88c 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -110,33 +110,38 @@ theorem degrees_zero : degrees (0 : MvPolynomial σ R) = 0 := by theorem degrees_one : degrees (1 : MvPolynomial σ R) = 0 := degrees_C 1 -theorem degrees_add [DecidableEq σ] (p q : MvPolynomial σ R) : +theorem degrees_add_le [DecidableEq σ] {p q : MvPolynomial σ R} : (p + q).degrees ≤ p.degrees ⊔ q.degrees := by simp_rw [degrees_def]; exact supDegree_add_le -theorem degrees_sum {ι : Type*} [DecidableEq σ] (s : Finset ι) (f : ι → MvPolynomial σ R) : +theorem degrees_sum_le {ι : Type*} [DecidableEq σ] (s : Finset ι) (f : ι → MvPolynomial σ R) : (∑ i ∈ s, f i).degrees ≤ s.sup fun i => (f i).degrees := by simp_rw [degrees_def]; exact supDegree_sum_le -theorem degrees_mul (p q : MvPolynomial σ R) : (p * q).degrees ≤ p.degrees + q.degrees := by +theorem degrees_mul_le {p q : MvPolynomial σ R} : (p * q).degrees ≤ p.degrees + q.degrees := by classical simp_rw [degrees_def] exact supDegree_mul_le (map_add _) -theorem degrees_prod {ι : Type*} (s : Finset ι) (f : ι → MvPolynomial σ R) : +theorem degrees_prod_le {ι : Type*} {s : Finset ι} {f : ι → MvPolynomial σ R} : (∏ i ∈ s, f i).degrees ≤ ∑ i ∈ s, (f i).degrees := by classical exact supDegree_prod_le (map_zero _) (map_add _) -theorem degrees_pow (p : MvPolynomial σ R) (n : ℕ) : (p ^ n).degrees ≤ n • p.degrees := by - simpa using degrees_prod (Finset.range n) fun _ ↦ p +theorem degrees_pow_le {p : MvPolynomial σ R} {n : ℕ} : (p ^ n).degrees ≤ n • p.degrees := by + simpa using degrees_prod_le (s := .range n) (f := fun _ ↦ p) + +@[deprecated (since := "2024-12-28")] alias degrees_add := degrees_add_le +@[deprecated (since := "2024-12-28")] alias degrees_sum := degrees_sum_le +@[deprecated (since := "2024-12-28")] alias degrees_mul := degrees_mul_le +@[deprecated (since := "2024-12-28")] alias degrees_prod := degrees_prod_le +@[deprecated (since := "2024-12-28")] alias degrees_pow := degrees_pow_le theorem mem_degrees {p : MvPolynomial σ R} {i : σ} : i ∈ p.degrees ↔ ∃ d, p.coeff d ≠ 0 ∧ i ∈ d.support := by classical simp only [degrees_def, Multiset.mem_sup, ← mem_support_iff, Finsupp.mem_toMultiset, exists_prop] -theorem le_degrees_add {p q : MvPolynomial σ R} (h : Disjoint p.degrees q.degrees) : - p.degrees ≤ (p + q).degrees := by +theorem le_degrees_add_left (h : Disjoint p.degrees q.degrees) : p.degrees ≤ (p + q).degrees := by classical apply Finset.sup_le intro d hd @@ -153,22 +158,19 @@ theorem le_degrees_add {p q : MvPolynomial σ R} (h : Disjoint p.degrees q.degre refine ⟨j, ?_, j, ?_, rfl⟩ all_goals rw [mem_degrees]; refine ⟨d, ?_, hj⟩; assumption -theorem degrees_add_of_disjoint [DecidableEq σ] {p q : MvPolynomial σ R} - (h : Disjoint p.degrees q.degrees) : (p + q).degrees = p.degrees ∪ q.degrees := by - apply le_antisymm - · apply degrees_add - · apply Multiset.union_le - · apply le_degrees_add h - · rw [add_comm] - apply le_degrees_add h.symm - -theorem degrees_map [CommSemiring S] (p : MvPolynomial σ R) (f : R →+* S) : - (map f p).degrees ⊆ p.degrees := by - classical - dsimp only [degrees] - apply Multiset.subset_of_le - apply Finset.sup_mono - apply MvPolynomial.support_map_subset +@[deprecated (since := "2024-12-28")] alias le_degrees_add := le_degrees_add_left + +lemma le_degrees_add_right (h : Disjoint p.degrees q.degrees) : q.degrees ≤ (p + q).degrees := by + simpa [add_comm] using le_degrees_add_left h.symm + +theorem degrees_add_of_disjoint [DecidableEq σ] (h : Disjoint p.degrees q.degrees) : + (p + q).degrees = p.degrees ∪ q.degrees := + degrees_add_le.antisymm <| Multiset.union_le (le_degrees_add_left h) (le_degrees_add_right h) + +lemma degrees_map_le [CommSemiring S] {f : R →+* S} : (map f p).degrees ≤ p.degrees := by + classical exact Finset.sup_mono <| support_map_subset .. + +@[deprecated (since := "2024-12-28")] alias degrees_map := degrees_map_le theorem degrees_rename (f : σ → τ) (φ : MvPolynomial σ R) : (rename f φ).degrees ⊆ φ.degrees.map f := by @@ -262,7 +264,7 @@ theorem degreeOf_mul_le (i : σ) (f g : MvPolynomial σ R) : degreeOf i (f * g) ≤ degreeOf i f + degreeOf i g := by classical simp only [degreeOf] - convert Multiset.count_le_of_le i (degrees_mul f g) + convert Multiset.count_le_of_le i degrees_mul_le rw [Multiset.count_add] theorem degreeOf_sum_le {ι : Type*} (i : σ) (s : Finset ι) (f : ι → MvPolynomial σ R) : @@ -297,7 +299,7 @@ theorem degreeOf_mul_X_self (j : σ) (f : MvPolynomial σ R) : degreeOf j (f * X j) ≤ degreeOf j f + 1 := by classical simp only [degreeOf] - apply (Multiset.count_le_of_le j (degrees_mul f (X j))).trans + apply (Multiset.count_le_of_le j degrees_mul_le).trans simp only [Multiset.count_add, add_le_add_iff_left] convert Multiset.count_le_of_le j <| degrees_X' j rw [Multiset.count_singleton_self] @@ -320,13 +322,13 @@ theorem degreeOf_mul_X_eq_degreeOf_add_one_iff (j : σ) (f : MvPolynomial σ R) theorem degreeOf_C_mul_le (p : MvPolynomial σ R) (i : σ) (c : R) : (C c * p).degreeOf i ≤ p.degreeOf i := by unfold degreeOf - convert Multiset.count_le_of_le i <| degrees_mul (C c) p + convert Multiset.count_le_of_le i degrees_mul_le simp only [degrees_C, zero_add] theorem degreeOf_mul_C_le (p : MvPolynomial σ R) (i : σ) (c : R) : (p * C c).degreeOf i ≤ p.degreeOf i := by unfold degreeOf - convert Multiset.count_le_of_le i <| degrees_mul p (C c) + convert Multiset.count_le_of_le i degrees_mul_le simp only [degrees_C, add_zero] theorem degreeOf_rename_of_injective {p : MvPolynomial σ R} {f : σ → τ} (h : Function.Injective f) @@ -522,6 +524,51 @@ theorem totalDegree_rename_le (f : σ → τ) (p : MvPolynomial σ R) : end TotalDegree +section degreesLE +variable {s t : Multiset σ} + +variable (R σ s) in +/-- The submodule of multivariate polynomials of degrees bounded by a monomial `s`. -/ +def degreesLE : Submodule R (MvPolynomial σ R) where + carrier := {p | p.degrees ≤ s} + add_mem' {a b} ha hb := by classical exact degrees_add_le.trans (sup_le ha hb) + zero_mem' := by simp + smul_mem' c {x} hx := by + dsimp + rw [Algebra.smul_def] + refine degrees_mul_le.trans ?_ + simpa [degrees_C] using hx + +@[simp] lemma mem_degreesLE : p ∈ degreesLE R σ s ↔ p.degrees ≤ s := Iff.rfl + +variable (s t) in +lemma degreesLE_add : degreesLE R σ (s + t) = degreesLE R σ s * degreesLE R σ t := by + classical + rw [le_antisymm_iff, Submodule.mul_le] + refine ⟨fun x hx ↦ x.as_sum ▸ sum_mem fun i hi ↦ ?_, + fun x hx y hy ↦ degrees_mul_le.trans (add_le_add hx hy)⟩ + replace hi : i.toMultiset ≤ s + t := (Finset.le_sup hi).trans hx + let a := (i.toMultiset - t).toFinsupp + let b := (i.toMultiset ⊓ t).toFinsupp + have : a + b = i := Multiset.toFinsupp.symm.injective (by simp [a, b, Multiset.sub_add_inter]) + have ha : a.toMultiset ≤ s := by simpa [a, add_comm (a := t)] using hi + have hb : b.toMultiset ≤ t := by simp [b, Multiset.inter_le_right] + rw [show monomial i (x.coeff i) = monomial a (x.coeff i) * monomial b 1 by simp [this]] + exact Submodule.mul_mem_mul ((degrees_monomial _ _).trans ha) ((degrees_monomial _ _).trans hb) + +@[simp] lemma degreesLE_zero : degreesLE R σ 0 = 1 := by + refine le_antisymm (fun x hx ↦ ?_) (by simp) + simp only [mem_degreesLE, nonpos_iff_eq_zero] at hx + have := (totalDegree_eq_zero_iff_eq_C (p := x)).mp + (Nat.eq_zero_of_le_zero (x.totalDegree_le_degrees_card.trans (by simp [hx]))) + exact ⟨x.coeff 0, by simp [Algebra.smul_def, ← this]⟩ + +variable (s) in +lemma degreesLE_nsmul : ∀ n, degreesLE R σ (n • s) = degreesLE R σ s ^ n + | 0 => by simp + | k + 1 => by simp only [pow_succ, degreesLE_nsmul, degreesLE_add, add_smul, one_smul] + +end degreesLE end CommSemiring end MvPolynomial diff --git a/Mathlib/Algebra/MvPolynomial/Equiv.lean b/Mathlib/Algebra/MvPolynomial/Equiv.lean index 9f7c660990ad0..11ed4077ba4e3 100644 --- a/Mathlib/Algebra/MvPolynomial/Equiv.lean +++ b/Mathlib/Algebra/MvPolynomial/Equiv.lean @@ -184,21 +184,19 @@ theorem iterToSum_X (b : S₁) : iterToSum R S₁ S₂ (X b) = X (Sum.inl b) := theorem iterToSum_C_X (c : S₂) : iterToSum R S₁ S₂ (C (X c)) = X (Sum.inr c) := Eq.trans (eval₂_C _ _ (X c)) (eval₂_X _ _ _) -variable (σ) +section isEmptyRingEquiv +variable [IsEmpty σ] +variable (σ) in /-- The algebra isomorphism between multivariable polynomials in no variables and the ground ring. -/ -@[simps!] -def isEmptyAlgEquiv [he : IsEmpty σ] : MvPolynomial σ R ≃ₐ[R] R := - AlgEquiv.ofAlgHom (aeval (IsEmpty.elim he)) (Algebra.ofId _ _) - (by ext) - (by - ext i m - exact IsEmpty.elim' he i) +@[simps! apply] +def isEmptyAlgEquiv : MvPolynomial σ R ≃ₐ[R] R := + .ofAlgHom (aeval isEmptyElim) (Algebra.ofId _ _) (by ext) (by ext i m; exact isEmptyElim i) -variable {R S₁ σ} in +variable {R S₁} in @[simp] -lemma aeval_injective_iff_of_isEmpty [IsEmpty σ] [CommSemiring S₁] [Algebra R S₁] {f : σ → S₁} : +lemma aeval_injective_iff_of_isEmpty [CommSemiring S₁] [Algebra R S₁] {f : σ → S₁} : Function.Injective (aeval f : MvPolynomial σ R →ₐ[R] S₁) ↔ Function.Injective (algebraMap R S₁) := by have : aeval f = (Algebra.ofId R S₁).comp (@isEmptyAlgEquiv R σ _ _).toAlgHom := by @@ -207,13 +205,20 @@ lemma aeval_injective_iff_of_isEmpty [IsEmpty σ] [CommSemiring S₁] [Algebra R rw [this, ← Injective.of_comp_iff' _ (@isEmptyAlgEquiv R σ _ _).bijective] rfl +variable (σ) in /-- The ring isomorphism between multivariable polynomials in no variables and the ground ring. -/ -@[simps!] -def isEmptyRingEquiv [IsEmpty σ] : MvPolynomial σ R ≃+* R := - (isEmptyAlgEquiv R σ).toRingEquiv +@[simps! apply] +def isEmptyRingEquiv : MvPolynomial σ R ≃+* R := (isEmptyAlgEquiv R σ).toRingEquiv -variable {σ} +lemma isEmptyRingEquiv_symm_toRingHom : (isEmptyRingEquiv R σ).symm.toRingHom = C := rfl +@[simp] lemma isEmptyRingEquiv_symm_apply (r : R) : (isEmptyRingEquiv R σ).symm r = C r := rfl + +lemma isEmptyRingEquiv_eq_coeff_zero {σ R : Type*} [CommSemiring R] [IsEmpty σ] {x} : + isEmptyRingEquiv R σ x = x.coeff 0 := by + obtain ⟨x, rfl⟩ := (isEmptyRingEquiv R σ).symm.surjective x; simp + +end isEmptyRingEquiv /-- A helper function for `sumRingEquiv`. -/ @[simps] @@ -267,6 +272,30 @@ lemma sumAlgEquiv_comp_rename_inl : ext i simp +section commAlgEquiv +variable {R S₁ S₂ : Type*} [CommSemiring R] + +variable (R S₁ S₂) in +/-- The algebra isomorphism between multivariable polynomials in variables `S₁` of multivariable +polynomials in variables `S₂` and multivariable polynomials in variables `S₂` of multivariable +polynomials in variables `S₁`. -/ +noncomputable +def commAlgEquiv : MvPolynomial S₁ (MvPolynomial S₂ R) ≃ₐ[R] MvPolynomial S₂ (MvPolynomial S₁ R) := + (sumAlgEquiv R S₁ S₂).symm.trans <| (renameEquiv _ (.sumComm S₁ S₂)).trans (sumAlgEquiv R S₂ S₁) + +@[simp] lemma commAlgEquiv_C (p) : commAlgEquiv R S₁ S₂ (.C p) = .map C p := by + suffices (commAlgEquiv R S₁ S₂).toAlgHom.comp + (IsScalarTower.toAlgHom R (MvPolynomial S₂ R) _) = mapAlgHom (Algebra.ofId _ _) by + exact DFunLike.congr_fun this p + ext x : 1 + simp [commAlgEquiv] + +lemma commAlgEquiv_C_X (i) : commAlgEquiv R S₁ S₂ (.C (.X i)) = .X i := by simp + +@[simp] lemma commAlgEquiv_X (i) : commAlgEquiv R S₁ S₂ (.X i) = .C (.X i) := by simp [commAlgEquiv] + +end commAlgEquiv + section -- this speeds up typeclass search in the lemma below diff --git a/Mathlib/Algebra/MvPolynomial/Variables.lean b/Mathlib/Algebra/MvPolynomial/Variables.lean index d135fae4e62de..c13b44ae4ffd3 100644 --- a/Mathlib/Algebra/MvPolynomial/Variables.lean +++ b/Mathlib/Algebra/MvPolynomial/Variables.lean @@ -98,7 +98,7 @@ theorem vars_add_subset [DecidableEq σ] (p q : MvPolynomial σ R) : (p + q).vars ⊆ p.vars ∪ q.vars := by intro x hx simp only [vars_def, Finset.mem_union, Multiset.mem_toFinset] at hx ⊢ - simpa using Multiset.mem_of_le (degrees_add _ _) hx + simpa using Multiset.mem_of_le degrees_add_le hx theorem vars_add_of_disjoint [DecidableEq σ] (h : Disjoint p.vars q.vars) : (p + q).vars = p.vars ∪ q.vars := by @@ -110,7 +110,7 @@ section Mul theorem vars_mul [DecidableEq σ] (φ ψ : MvPolynomial σ R) : (φ * ψ).vars ⊆ φ.vars ∪ ψ.vars := by simp_rw [vars_def, ← Multiset.toFinset_add, Multiset.toFinset_subset] - exact Multiset.subset_of_le (degrees_mul φ ψ) + exact Multiset.subset_of_le degrees_mul_le @[simp] theorem vars_one : (1 : MvPolynomial σ R).vars = ∅ := @@ -194,7 +194,8 @@ section Map variable [CommSemiring S] (f : R →+* S) variable (p) -theorem vars_map : (map f p).vars ⊆ p.vars := by classical simp [vars_def, degrees_map] +theorem vars_map : (map f p).vars ⊆ p.vars := by + classical simp [vars_def, Multiset.subset_of_le degrees_map_le] variable {f} diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index a507d68457b98..672e458b039da 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -80,7 +80,7 @@ where simp_rw [Finset.mem_map, Embedding.coeFn_mk] at hti htj obtain ⟨ai, hai, hij'⟩ := hti obtain ⟨aj, haj, rfl⟩ := htj - rw [Fin.cons_eq_cons] at hij' + rw [Fin.cons_inj] at hij' ext · exact hij'.1 · obtain ⟨-, rfl⟩ := hij' diff --git a/Mathlib/Algebra/Order/Archimedean/IndicatorCard.lean b/Mathlib/Algebra/Order/Archimedean/IndicatorCard.lean new file mode 100644 index 0000000000000..18769abc95e63 --- /dev/null +++ b/Mathlib/Algebra/Order/Archimedean/IndicatorCard.lean @@ -0,0 +1,74 @@ +/- +Copyright (c) 2024 Damien Thomine. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damien Thomine +-/ +import Mathlib.Algebra.Order.Archimedean.Basic +import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Group.Indicator +import Mathlib.Order.LiminfLimsup +import Mathlib.SetTheory.Cardinal.Finite + +/-! +# Cardinality and limit of sum of indicators +This file contains results relating the cardinality of subsets of ℕ and limits, +limsups of sums of indicators. + +## Tags +finite, indicator, limsup, tendsto +-/ + +namespace Set + +open Filter Finset + +lemma sum_indicator_eventually_eq_card {α : Type*} [AddCommMonoid α] (a : α) {s : Set ℕ} + (hs : s.Finite) : + ∀ᶠ n in atTop, ∑ k ∈ Finset.range n, s.indicator (fun _ ↦ a) k = (Nat.card s) • a := by + have key : ∀ x ∈ hs.toFinset, s.indicator (fun _ ↦ a) x = a := by + intro x hx + rw [indicator_of_mem (hs.mem_toFinset.1 hx) (fun _ ↦ a)] + rw [Nat.card_eq_card_finite_toFinset hs, ← sum_eq_card_nsmul key, eventually_atTop] + obtain ⟨m, hm⟩ := hs.bddAbove + refine ⟨m + 1, fun n n_m ↦ (sum_subset ?_ ?_).symm⟩ <;> intro x <;> rw [hs.mem_toFinset] + · rw [Finset.mem_range] + exact fun x_s ↦ ((mem_upperBounds.1 hm) x x_s).trans_lt (Nat.lt_of_succ_le n_m) + · exact fun _ x_s ↦ indicator_of_not_mem x_s (fun _ ↦ a) + +lemma infinite_iff_tendsto_sum_indicator_atTop {R : Type*} [OrderedAddCommMonoid R] + [AddLeftStrictMono R] [Archimedean R] {r : R} (h : 0 < r) {s : Set ℕ} : + s.Infinite ↔ atTop.Tendsto (fun n ↦ ∑ k ∈ Finset.range n, s.indicator (fun _ ↦ r) k) atTop := by + constructor + · have h_mono : Monotone fun n ↦ ∑ k ∈ Finset.range n, s.indicator (fun _ ↦ r) k := by + refine (sum_mono_set_of_nonneg ?_).comp range_mono + exact (fun _ ↦ indicator_nonneg (fun _ _ ↦ h.le) _) + rw [h_mono.tendsto_atTop_atTop_iff] + intro hs n + obtain ⟨n', hn'⟩ := exists_lt_nsmul h n + obtain ⟨t, t_s, t_card⟩ := hs.exists_subset_card_eq n' + obtain ⟨m, hm⟩ := t.bddAbove + refine ⟨m + 1, hn'.le.trans ?_⟩ + apply (sum_le_sum fun i _ ↦ (indicator_le_indicator_of_subset t_s (fun _ ↦ h.le)) i).trans_eq' + have h : t ⊆ Finset.range (m + 1) := by + intro i i_t + rw [Finset.mem_range] + exact (hm i_t).trans_lt (lt_add_one m) + rw [sum_indicator_subset (fun _ ↦ r) h, sum_eq_card_nsmul (fun _ _ ↦ rfl), t_card] + · contrapose + intro hs + rw [not_infinite] at hs + rw [tendsto_congr' (sum_indicator_eventually_eq_card r hs), tendsto_atTop_atTop] + push_neg + obtain ⟨m, hm⟩ := exists_lt_nsmul h (Nat.card s • r) + exact ⟨m • r, fun n ↦ ⟨n, le_refl n, not_le_of_lt hm⟩⟩ + +lemma limsup_eq_tendsto_sum_indicator_atTop {α R : Type*} [OrderedAddCommMonoid R] + [AddLeftStrictMono R] [Archimedean R] {r : R} (h : 0 < r) (s : ℕ → Set α) : + atTop.limsup s = { ω | atTop.Tendsto + (fun n ↦ ∑ k ∈ Finset.range n, (s k).indicator (fun _ ↦ r) ω) atTop } := by + nth_rw 1 [← Nat.cofinite_eq_atTop, cofinite.limsup_set_eq] + ext ω + rw [mem_setOf_eq, mem_setOf_eq, infinite_iff_tendsto_sum_indicator_atTop h, iff_eq_eq] + congr + +end Set diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index fb76b07c2edb2..b7366261beeea 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -404,7 +404,7 @@ theorem zero_limZero : LimZero (0 : CauSeq β abv) theorem const_limZero {x : β} : LimZero (const x) ↔ x = 0 := ⟨fun H => (abv_eq_zero abv).1 <| - (eq_of_le_of_forall_le_of_dense (abv_nonneg abv _)) fun _ ε0 => + (eq_of_le_of_forall_lt_imp_le_of_dense (abv_nonneg abv _)) fun _ ε0 => let ⟨_, hi⟩ := H _ ε0 le_of_lt <| hi _ le_rfl, fun e => e.symm ▸ zero_limZero⟩ diff --git a/Mathlib/Algebra/Order/CauSeq/Completion.lean b/Mathlib/Algebra/Order/CauSeq/Completion.lean index 283cb6611892c..2272380559614 100644 --- a/Mathlib/Algebra/Order/CauSeq/Completion.lean +++ b/Mathlib/Algebra/Order/CauSeq/Completion.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Robert Y. Lewis -/ import Mathlib.Algebra.Order.CauSeq.Basic import Mathlib.Data.Rat.Cast.Defs +import Mathlib.Tactic.FastInstance /-! # Cauchy completion @@ -139,7 +140,7 @@ private theorem zero_def : 0 = mk (abv := abv) 0 := private theorem one_def : 1 = mk (abv := abv) 1 := rfl -instance Cauchy.ring : Ring (Cauchy abv) := +instance Cauchy.ring : Ring (Cauchy abv) := fast_instance% Function.Surjective.ring mk Quotient.mk'_surjective zero_def.symm one_def.symm (fun _ _ => (mk_add _ _).symm) (fun _ _ => (mk_mul _ _).symm) (fun _ => (mk_neg _).symm) (fun _ _ => (mk_sub _ _).symm) (fun _ _ => (mk_smul _ _).symm) (fun _ _ => (mk_smul _ _).symm) @@ -164,7 +165,7 @@ section variable {α : Type*} [LinearOrderedField α] variable {β : Type*} [CommRing β] {abv : β → α} [IsAbsoluteValue abv] -instance Cauchy.commRing : CommRing (Cauchy abv) := +instance Cauchy.commRing : CommRing (Cauchy abv) := fast_instance% Function.Surjective.commRing mk Quotient.mk'_surjective zero_def.symm one_def.symm (fun _ _ => (mk_add _ _).symm) (fun _ _ => (mk_mul _ _).symm) (fun _ => (mk_neg _).symm) (fun _ _ => (mk_sub _ _).symm) (fun _ _ => (mk_smul _ _).symm) (fun _ _ => (mk_smul _ _).symm) diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index 5ac0f4d4135a6..7a096783dc245 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -99,7 +99,6 @@ variable {β} @[simp] theorem mem_cutMap_iff : b ∈ cutMap β a ↔ ∃ q : ℚ, (q : α) < a ∧ (q : β) = b := Iff.rfl --- @[simp] -- Porting note: not in simpNF theorem coe_mem_cutMap_iff [CharZero β] : (q : β) ∈ cutMap β a ↔ (q : α) < a := Rat.cast_injective.mem_set_image diff --git a/Mathlib/Algebra/Order/Disjointed.lean b/Mathlib/Algebra/Order/Disjointed.lean new file mode 100644 index 0000000000000..77556150bed67 --- /dev/null +++ b/Mathlib/Algebra/Order/Disjointed.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2025 David Loeffler. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Loeffler +-/ +import Mathlib.Algebra.Order.PartialSups +import Mathlib.Data.Nat.SuccPred +import Mathlib.Order.Disjointed + +/-! +# `Disjointed` for functions on a `SuccAddOrder` + +This file contains material excised from `Mathlib.Order.Disjointed` to avoid import dependencies +from `Mathlib.Algebra.Order` into `Mathlib.Order`. + +## TODO + +Find a useful statement of `disjointedRec_succ`. +-/ + +open Order + +variable {α ι : Type*} [GeneralizedBooleanAlgebra α] + +section SuccAddOrder + +variable [LinearOrder ι] [LocallyFiniteOrderBot ι] [Add ι] [One ι] [SuccAddOrder ι] + +theorem disjointed_add_one [NoMaxOrder ι] (f : ι → α) (i : ι) : + disjointed f (i + 1) = f (i + 1) \ partialSups f i := by + simpa only [succ_eq_add_one] using disjointed_succ f (not_isMax i) + +protected lemma Monotone.disjointed_add_one_sup {f : ι → α} (hf : Monotone f) (i : ι) : + disjointed f (i + 1) ⊔ f i = f (i + 1) := by + simpa only [succ_eq_add_one i] using hf.disjointed_succ_sup i + +protected lemma Monotone.disjointed_add_one [NoMaxOrder ι] {f : ι → α} (hf : Monotone f) (i : ι) : + disjointed f (i + 1) = f (i + 1) \ f i := by + rw [← succ_eq_add_one, hf.disjointed_succ] + exact not_isMax i + +end SuccAddOrder + +section Nat + +/-- A recursion principle for `disjointed`. To construct / define something for `disjointed f i`, +it's enough to construct / define it for `f n` and to able to extend through diffs. + +Note that this version allows an arbitrary `Sort*`, but requires the domain to be `Nat`, while +the root-level `disjointedRec` allows more general domains but requires `p` to be `Prop`-valued. -/ +def Nat.disjointedRec {f : ℕ → α} {p : α → Sort*} (hdiff : ∀ ⦃t i⦄, p t → p (t \ f i)) : + ∀ ⦃n⦄, p (f n) → p (disjointed f n) + | 0 => fun h₀ ↦ disjointed_zero f ▸ h₀ + | n + 1 => fun h => by + suffices H : ∀ k, p (f (n + 1) \ partialSups f k) from disjointed_add_one f n ▸ H n + intro k + induction k with + | zero => exact hdiff h + | succ k ih => simpa only [partialSups_add_one, ← sdiff_sdiff_left] using hdiff ih + +@[simp] +theorem disjointedRec_zero {f : ℕ → α} {p : α → Sort*} + (hdiff : ∀ ⦃t i⦄, p t → p (t \ f i)) (h₀ : p (f 0)) : + Nat.disjointedRec hdiff h₀ = (disjointed_zero f ▸ h₀) := + rfl + +-- TODO: Find a useful statement of `disjointedRec_succ`. + +end Nat diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index d2b9446b43e01..5d80dad8dfcfd 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -414,7 +414,7 @@ theorem le_iff_forall_one_lt_le_mul₀ {α : Type*} [LinearOrderedSemifield α] obtain rfl|hb := hb.eq_or_lt · simp_rw [zero_mul] at h exact h 2 one_lt_two - refine le_of_forall_le_of_dense fun x hbx => ?_ + refine le_of_forall_gt_imp_ge_of_dense fun x hbx => ?_ convert h (x / b) ((one_lt_div hb).mpr hbx) rw [mul_div_cancel₀ _ hb.ne'] @@ -755,14 +755,14 @@ private lemma exists_mul_right_lt₀ {a b c : α} (hc : a * b < c) : ∃ b' > b, simp_rw [mul_comm a] at hc ⊢; exact exists_mul_left_lt₀ hc lemma le_mul_of_forall_lt₀ {a b c : α} (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') : c ≤ a * b := by - refine le_of_forall_le_of_dense fun d hd ↦ ?_ + refine le_of_forall_gt_imp_ge_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_mul_left_lt₀ hd obtain ⟨b', hb', hd⟩ := exists_mul_right_lt₀ hd exact (h a' ha' b' hb').trans hd.le lemma mul_le_of_forall_lt_of_nonneg {a b c : α} (ha : 0 ≤ a) (hc : 0 ≤ c) (h : ∀ a' ≥ 0, a' < a → ∀ b' ≥ 0, b' < b → a' * b' ≤ c) : a * b ≤ c := by - refine le_of_forall_ge_of_dense fun d d_ab ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d d_ab ↦ ?_ rcases lt_or_le d 0 with hd | hd · exact hd.le.trans hc obtain ⟨a', ha', d_ab⟩ := exists_lt_mul_left_of_nonneg ha hd d_ab diff --git a/Mathlib/Algebra/Order/Field/Canonical.lean b/Mathlib/Algebra/Order/Field/Canonical.lean index 512b93cdd1c40..1435a769ed076 100644 --- a/Mathlib/Algebra/Order/Field/Canonical.lean +++ b/Mathlib/Algebra/Order/Field/Canonical.lean @@ -27,9 +27,11 @@ variable {α : Type*} [LinearOrderedSemifield α] [CanonicallyOrderedAdd α] /-- Construct a `LinearOrderedCommGroupWithZero` from a canonically ordered `LinearOrderedSemifield`. -/ abbrev LinearOrderedSemifield.toLinearOrderedCommGroupWithZero : - LinearOrderedCommGroupWithZero α := - { ‹LinearOrderedSemifield α› with - mul_le_mul_left := fun _ _ h _ ↦ mul_le_mul_of_nonneg_left h <| zero_le _ } + LinearOrderedCommGroupWithZero α where + __ := ‹LinearOrderedSemifield α› + bot := 0 + bot_le := zero_le + mul_le_mul_left _ _ h _:= mul_le_mul_of_nonneg_left h <| zero_le _ variable [Sub α] [OrderedSub α] diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index ac82917ae0852..f8c6e2dddb310 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -116,7 +116,7 @@ protected lemma Odd.zpow_nonneg_iff (hn : Odd n) : 0 ≤ a ^ n ↔ 0 ≤ a := theorem Odd.zpow_nonpos_iff (hn : Odd n) : a ^ n ≤ 0 ↔ a ≤ 0 := by rw [le_iff_lt_or_eq, le_iff_lt_or_eq, hn.zpow_neg_iff, zpow_eq_zero_iff] rintro rfl - exact Int.not_even_iff_odd.2 hn even_zero + exact Int.not_even_iff_odd.2 hn .zero lemma Odd.zpow_pos_iff (hn : Odd n) : 0 < a ^ n ↔ 0 < a := lt_iff_lt_of_le_iff_le hn.zpow_nonpos_iff diff --git a/Mathlib/Algebra/Order/Field/Subfield.lean b/Mathlib/Algebra/Order/Field/Subfield.lean index 6556c94ad7517..6b0ef75e83012 100644 --- a/Mathlib/Algebra/Order/Field/Subfield.lean +++ b/Mathlib/Algebra/Order/Field/Subfield.lean @@ -17,7 +17,7 @@ variable {K S : Type*} [SetLike S K] -- Prefer subclasses of `Field` over subclasses of `SubfieldClass`. /-- A subfield of a `LinearOrderedField` is a `LinearOrderedField`. -/ instance (priority := 75) toLinearOrderedField [LinearOrderedField K] - [SubfieldClass S K] (s : S) : LinearOrderedField s := + [SubfieldClass S K] (s : S) : LinearOrderedField s := fast_instance% Subtype.coe_injective.linearOrderedField Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) @@ -30,7 +30,8 @@ namespace Subfield variable {K : Type*} /-- A subfield of a `LinearOrderedField` is a `LinearOrderedField`. -/ -instance toLinearOrderedField [LinearOrderedField K] (s : Subfield K) : LinearOrderedField s := +instance toLinearOrderedField [LinearOrderedField K] (s : Subfield K) : + LinearOrderedField s := fast_instance% Subtype.coe_injective.linearOrderedField Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (by intros; rfl) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index d11212a76f386..94a5729037a50 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -4,13 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kevin Kappelmann -/ import Mathlib.Algebra.CharZero.Lemmas -import Mathlib.Algebra.Order.Interval.Set.Group import Mathlib.Algebra.Group.Int.Even import Mathlib.Algebra.Group.Int.Units import Mathlib.Data.Int.Lemmas import Mathlib.Data.Nat.Cast.Order.Field import Mathlib.Data.Set.Subsingleton -import Mathlib.Order.GaloisConnection import Mathlib.Tactic.Abel import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Linarith @@ -33,7 +31,6 @@ We define the natural- and integer-valued floor and ceil functions on linearly o * `Int.floor a`: Greatest integer `z` such that `z ≤ a`. * `Int.ceil a`: Least integer `z` such that `a ≤ z`. * `Int.fract a`: Fractional part of `a`, defined as `a - floor a`. -* `round a`: Nearest integer to `a`. It rounds halves towards infinity. ## Notations @@ -605,7 +602,7 @@ def floor : α → ℤ := def ceil : α → ℤ := FloorRing.ceil -/-- `Int.fract a`, the fractional part of `a`, is `a` minus its floor. -/ +/-- `Int.fract a` the fractional part of `a`, is `a` minus its floor. -/ def fract (a : α) : α := a - floor a @@ -1324,177 +1321,6 @@ end Int open Int -/-! ### Round -/ - - -section round - -section LinearOrderedRing - -variable [LinearOrderedRing α] [FloorRing α] - -/-- `round` rounds a number to the nearest integer. `round (1 / 2) = 1` -/ -def round (x : α) : ℤ := - if 2 * fract x < 1 then ⌊x⌋ else ⌈x⌉ - -@[simp] -theorem round_zero : round (0 : α) = 0 := by simp [round] - -@[simp] -theorem round_one : round (1 : α) = 1 := by simp [round] - -@[simp] -theorem round_natCast (n : ℕ) : round (n : α) = n := by simp [round] - -@[simp] -theorem round_ofNat (n : ℕ) [n.AtLeastTwo] : round (ofNat(n) : α) = ofNat(n) := - round_natCast n - -@[simp] -theorem round_intCast (n : ℤ) : round (n : α) = n := by simp [round] - -@[simp] -theorem round_add_int (x : α) (y : ℤ) : round (x + y) = round x + y := by - rw [round, round, Int.fract_add_int, Int.floor_add_int, Int.ceil_add_int, ← apply_ite₂, ite_self] - -@[simp] -theorem round_add_one (a : α) : round (a + 1) = round a + 1 := by - -- Porting note: broken `convert round_add_int a 1` - rw [← round_add_int a 1, cast_one] - -@[simp] -theorem round_sub_int (x : α) (y : ℤ) : round (x - y) = round x - y := by - rw [sub_eq_add_neg] - norm_cast - rw [round_add_int, sub_eq_add_neg] - -@[simp] -theorem round_sub_one (a : α) : round (a - 1) = round a - 1 := by - -- Porting note: broken `convert round_sub_int a 1` - rw [← round_sub_int a 1, cast_one] - -@[simp] -theorem round_add_nat (x : α) (y : ℕ) : round (x + y) = round x + y := - mod_cast round_add_int x y - -@[simp] -theorem round_add_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : - round (x + ofNat(n)) = round x + ofNat(n) := - round_add_nat x n - -@[simp] -theorem round_sub_nat (x : α) (y : ℕ) : round (x - y) = round x - y := - mod_cast round_sub_int x y - -@[simp] -theorem round_sub_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : - round (x - ofNat(n)) = round x - ofNat(n) := - round_sub_nat x n - -@[simp] -theorem round_int_add (x : α) (y : ℤ) : round ((y : α) + x) = y + round x := by - rw [add_comm, round_add_int, add_comm] - -@[simp] -theorem round_nat_add (x : α) (y : ℕ) : round ((y : α) + x) = y + round x := by - rw [add_comm, round_add_nat, add_comm] - -@[simp] -theorem round_ofNat_add (n : ℕ) [n.AtLeastTwo] (x : α) : - round (ofNat(n) + x) = ofNat(n) + round x := - round_nat_add x n - -theorem abs_sub_round_eq_min (x : α) : |x - round x| = min (fract x) (1 - fract x) := by - simp_rw [round, min_def_lt, two_mul, ← lt_tsub_iff_left] - cases' lt_or_ge (fract x) (1 - fract x) with hx hx - · rw [if_pos hx, if_pos hx, self_sub_floor, abs_fract] - · have : 0 < fract x := by - replace hx : 0 < fract x + fract x := lt_of_lt_of_le zero_lt_one (tsub_le_iff_left.mp hx) - simpa only [← two_mul, mul_pos_iff_of_pos_left, zero_lt_two] using hx - rw [if_neg (not_lt.mpr hx), if_neg (not_lt.mpr hx), abs_sub_comm, ceil_sub_self_eq this.ne.symm, - abs_one_sub_fract] - -theorem round_le (x : α) (z : ℤ) : |x - round x| ≤ |x - z| := by - rw [abs_sub_round_eq_min, min_le_iff] - rcases le_or_lt (z : α) x with (hx | hx) <;> [left; right] - · conv_rhs => rw [abs_eq_self.mpr (sub_nonneg.mpr hx), ← fract_add_floor x, add_sub_assoc] - simpa only [le_add_iff_nonneg_right, sub_nonneg, cast_le] using le_floor.mpr hx - · rw [abs_eq_neg_self.mpr (sub_neg.mpr hx).le] - conv_rhs => rw [← fract_add_floor x] - rw [add_sub_assoc, add_comm, neg_add, neg_sub, le_add_neg_iff_add_le, sub_add_cancel, - le_sub_comm] - norm_cast - exact floor_le_sub_one_iff.mpr hx - -end LinearOrderedRing - -section LinearOrderedField - -variable [LinearOrderedField α] [FloorRing α] - -theorem round_eq (x : α) : round x = ⌊x + 1 / 2⌋ := by - simp_rw [round, (by simp only [lt_div_iff₀', two_pos] : 2 * fract x < 1 ↔ fract x < 1 / 2)] - cases' lt_or_le (fract x) (1 / 2) with hx hx - · conv_rhs => rw [← fract_add_floor x, add_assoc, add_left_comm, floor_int_add] - rw [if_pos hx, self_eq_add_right, floor_eq_iff, cast_zero, zero_add] - constructor - · linarith [fract_nonneg x] - · linarith - · have : ⌊fract x + 1 / 2⌋ = 1 := by - rw [floor_eq_iff] - constructor - · norm_num - linarith - · norm_num - linarith [fract_lt_one x] - rw [if_neg (not_lt.mpr hx), ← fract_add_floor x, add_assoc, add_left_comm, floor_int_add, - ceil_add_int, add_comm _ ⌊x⌋, add_right_inj, ceil_eq_iff, this, cast_one, sub_self] - constructor - · linarith - · linarith [fract_lt_one x] - -@[simp] -theorem round_two_inv : round (2⁻¹ : α) = 1 := by - simp only [round_eq, ← one_div, add_halves, floor_one] - -@[simp] -theorem round_neg_two_inv : round (-2⁻¹ : α) = 0 := by - simp only [round_eq, ← one_div, neg_add_cancel, floor_zero] - -@[simp] -theorem round_eq_zero_iff {x : α} : round x = 0 ↔ x ∈ Ico (-(1 / 2)) ((1 : α) / 2) := by - rw [round_eq, floor_eq_zero_iff, add_mem_Ico_iff_left] - norm_num - -theorem abs_sub_round (x : α) : |x - round x| ≤ 1 / 2 := by - rw [round_eq, abs_sub_le_iff] - have := floor_le (x + 1 / 2) - have := lt_floor_add_one (x + 1 / 2) - constructor <;> linarith - -theorem abs_sub_round_div_natCast_eq {m n : ℕ} : - |(m : α) / n - round ((m : α) / n)| = ↑(min (m % n) (n - m % n)) / n := by - rcases n.eq_zero_or_pos with (rfl | hn) - · simp - have hn' : 0 < (n : α) := by - norm_cast - rw [abs_sub_round_eq_min, Nat.cast_min, ← min_div_div_right hn'.le, - fract_div_natCast_eq_div_natCast_mod, Nat.cast_sub (m.mod_lt hn).le, sub_div, div_self hn'.ne'] - -@[bound] -theorem sub_half_lt_round (x : α) : x - 1 / 2 < round x := by - rw [round_eq x, show x - 1 / 2 = x + 1 / 2 - 1 by linarith] - exact Int.sub_one_lt_floor (x + 1 / 2) - -@[bound] -theorem round_le_add_half (x : α) : round x ≤ x + 1 / 2 := by - rw [round_eq x] - exact Int.floor_le (x + 1 / 2) - -end LinearOrderedField - -end round - namespace Nat variable [LinearOrderedSemiring α] [LinearOrderedSemiring β] [FloorSemiring α] [FloorSemiring β] @@ -1539,20 +1365,6 @@ theorem map_fract (f : F) (hf : StrictMono f) (a : α) : fract (f a) = f (fract end Int -namespace Int - -variable [LinearOrderedField α] [LinearOrderedField β] [FloorRing α] [FloorRing β] -variable [FunLike F α β] [RingHomClass F α β] {a : α} {b : β} - -theorem map_round (f : F) (hf : StrictMono f) (a : α) : round (f a) = round a := by - have H : f 2 = 2 := map_natCast f 2 - simp_rw [round_eq, ← map_floor _ hf, map_add, one_div, map_inv₀, H] - -- Porting note: was - -- simp_rw [round_eq, ← map_floor _ hf, map_add, one_div, map_inv₀, map_bit0, map_one] - -- Would have thought that `map_natCast` would replace `map_bit0, map_one` but seems not - -end Int - section FloorRingToSemiring variable [LinearOrderedRing α] [FloorRing α] @@ -1673,5 +1485,3 @@ def evalIntCeil : PositivityExt where eval {u α} _zα _pα e := do | _, _, _ => throwError "failed to match on Int.ceil application" end Mathlib.Meta.Positivity - -set_option linter.style.longFile 1800 diff --git a/Mathlib/Algebra/Order/Floor/Div.lean b/Mathlib/Algebra/Order/Floor/Div.lean index 92562cfb116a0..5b20fe4336d57 100644 --- a/Mathlib/Algebra/Order/Floor/Div.lean +++ b/Mathlib/Algebra/Order/Floor/Div.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.GroupWithZero.Action.Pi import Mathlib.Algebra.Order.Module.Defs import Mathlib.Algebra.Order.Pi import Mathlib.Data.Finsupp.Order -import Mathlib.Order.GaloisConnection /-! # Flooring, ceiling division diff --git a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean index a019c1739800a..a1bd1983375e5 100644 --- a/Mathlib/Algebra/Order/Group/DenselyOrdered.lean +++ b/Mathlib/Algebra/Order/Group/DenselyOrdered.lean @@ -68,7 +68,7 @@ private lemma exists_mul_right_lt [CommGroup α] [LT α] [DenselyOrdered α] lemma le_mul_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] [DenselyOrdered α] {a b c : α} (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') : c ≤ a * b := by - refine le_of_forall_le_of_dense fun d hd ↦ ?_ + refine le_of_forall_gt_imp_ge_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_mul_left_lt hd obtain ⟨b', hb', hd⟩ := exists_mul_right_lt hd exact (h a' ha' b' hb').trans hd.le @@ -77,7 +77,7 @@ lemma le_mul_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α lemma mul_le_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] [DenselyOrdered α] {a b c : α} (h : ∀ a' < a, ∀ b' < b, a' * b' ≤ c) : a * b ≤ c := by - refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_lt_mul_left hd obtain ⟨b', hb', hd⟩ := exists_lt_mul_right hd exact hd.le.trans (h a' ha' b' hb') diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean index 56042ed6f46bf..c356f6487f6a8 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean @@ -28,8 +28,9 @@ variable [Lattice α] section Group variable [Group α] {a b : α} -/-- `mabs a` is the absolute value of `a`. -/ -@[to_additive "`abs a` is the absolute value of `a`"] def mabs (a : α) : α := a ⊔ a⁻¹ +/-- `mabs a`, denoted `|a|ₘ`, is the absolute value of `a`. -/ +@[to_additive "`abs a`, denoted `|a|`, is the absolute value of `a`"] +def mabs (a : α) : α := a ⊔ a⁻¹ @[inherit_doc mabs] macro:max atomic("|" noWs) a:term noWs "|ₘ" : term => `(mabs $a) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 2b1886a061185..5262c6b338451 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Algebra.GroupWithZero.WithZero import Mathlib.Algebra.Order.AddGroupWithTop -import Mathlib.Algebra.Order.GroupWithZero.Synonym import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Order.Monoid.OrderDual @@ -31,7 +30,7 @@ variable {α : Type*} /-- A linearly ordered commutative monoid with a zero element. -/ class LinearOrderedCommMonoidWithZero (α : Type*) extends LinearOrderedCommMonoid α, - CommMonoidWithZero α where + CommMonoidWithZero α, OrderBot α where /-- `0 ≤ 1` in any linearly ordered commutative monoid. -/ zero_le_one : (0 : α) ≤ 1 @@ -55,15 +54,17 @@ The following facts are true more generally in a (linearly) ordered commutative -/ /-- Pullback a `LinearOrderedCommMonoidWithZero` under an injective map. See note [reducible non-instances]. -/ -abbrev Function.Injective.linearOrderedCommMonoidWithZero {β : Type*} [Zero β] [One β] [Mul β] - [Pow β ℕ] [Max β] [Min β] (f : β → α) (hf : Function.Injective f) (zero : f 0 = 0) +abbrev Function.Injective.linearOrderedCommMonoidWithZero {β : Type*} [Zero β] [Bot β] [One β] + [Mul β] [Pow β ℕ] [Max β] [Min β] (f : β → α) (hf : Function.Injective f) (zero : f 0 = 0) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - LinearOrderedCommMonoidWithZero β := - { LinearOrder.lift f hf hsup hinf, hf.orderedCommMonoid f one mul npow, - hf.commMonoidWithZero f zero one mul npow with - zero_le_one := - show f 0 ≤ f 1 by simp only [zero, one, LinearOrderedCommMonoidWithZero.zero_le_one] } + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) + (bot : f ⊥ = ⊥) : LinearOrderedCommMonoidWithZero β where + __ := LinearOrder.lift f hf hsup hinf + __ := hf.orderedCommMonoid f one mul npow + __ := hf.commMonoidWithZero f zero one mul npow + zero_le_one := + show f 0 ≤ f 1 by simp only [zero, one, LinearOrderedCommMonoidWithZero.zero_le_one] + bot_le a := show f ⊥ ≤ f a from bot ▸ bot_le @[simp] lemma zero_le' : 0 ≤ a := by simpa only [mul_zero, mul_one] using mul_le_mul_left' (zero_le_one' α) a @@ -81,15 +82,18 @@ theorem zero_lt_iff : 0 < a ↔ a ≠ 0 := theorem ne_zero_of_lt (h : b < a) : a ≠ 0 := fun h1 ↦ not_lt_zero' <| show b < 0 from h1 ▸ h +/-- See also `bot_eq_zero` and `bot_eq_zero'` for canonically ordered monoids. -/ +lemma bot_eq_zero'' : (⊥ : α) = 0 := eq_of_forall_ge_iff fun _ ↦ by simp + instance instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual : LinearOrderedAddCommMonoidWithTop (Additive αᵒᵈ) where - top := (0 : α) - top_add' := fun a ↦ zero_mul a.toMul - le_top := fun _ ↦ zero_le' + top := .ofMul <| .toDual 0 + top_add' a := zero_mul a.toMul.ofDual + le_top _ := zero_le' instance instLinearOrderedAddCommMonoidWithTopOrderDualAdditive : LinearOrderedAddCommMonoidWithTop (Additive α)ᵒᵈ where - top := OrderDual.toDual (Additive.ofMul 0) + top := .toDual <| .ofMul _ top_add' := fun a ↦ zero_mul (Additive.toMul (OrderDual.ofDual a)) le_top := fun a ↦ @zero_le' _ _ (Additive.toMul (OrderDual.ofDual a)) diff --git a/Mathlib/Algebra/Order/Hom/Ring.lean b/Mathlib/Algebra/Order/Hom/Ring.lean index f0d4f29ca26ed..d3bd1d7c4c040 100644 --- a/Mathlib/Algebra/Order/Hom/Ring.lean +++ b/Mathlib/Algebra/Order/Hom/Ring.lean @@ -39,7 +39,8 @@ open Function variable {F α β γ δ : Type*} -/-- `OrderRingHom α β` is the type of monotone semiring homomorphisms from `α` to `β`. +/-- `OrderRingHom α β`, denoted `α →+*o β`, +is the type of monotone semiring homomorphisms from `α` to `β`. When possible, instead of parametrizing results over `(f : OrderRingHom α β)`, you should parametrize over `(F : Type*) [OrderRingHomClass F α β] (f : F)`. @@ -63,7 +64,8 @@ to otherwise the [refl] attribute on `OrderRingIso.refl` complains. TODO: change back when `refl` attribute is fixed, github issue https://github.com/leanprover-community/mathlib4/issues/2505 -/ -/-- `OrderRingHom α β` is the type of order-preserving semiring isomorphisms between `α` and `β`. +/-- `OrderRingIso α β`, denoted as `α ≃+*o β`, +is the type of order-preserving semiring isomorphisms between `α` and `β`. When possible, instead of parametrizing results over `(f : OrderRingIso α β)`, you should parametrize over `(F : Type*) [OrderRingIsoClass F α β] (f : F)`. diff --git a/Mathlib/Algebra/Order/Interval/Set/Instances.lean b/Mathlib/Algebra/Order/Interval/Set/Instances.lean index 06d73380d901f..9a3eed9085dfd 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Instances.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Instances.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Regular import Mathlib.Order.Interval.Set.Basic +import Mathlib.Tactic.FastInstance /-! # Algebraic instances for unit intervals @@ -118,20 +119,20 @@ theorem mul_le_left {x y : Icc (0 : α) 1} : x * y ≤ x := theorem mul_le_right {x y : Icc (0 : α) 1} : x * y ≤ y := (mul_le_mul_of_nonneg_right x.2.2 y.2.1).trans_eq (one_mul _) -instance monoidWithZero : MonoidWithZero (Icc (0 : α) 1) := +instance monoidWithZero : MonoidWithZero (Icc (0 : α) 1) := fast_instance% Subtype.coe_injective.monoidWithZero _ coe_zero coe_one coe_mul coe_pow instance commMonoidWithZero {α : Type*} [OrderedCommSemiring α] : - CommMonoidWithZero (Icc (0 : α) 1) := + CommMonoidWithZero (Icc (0 : α) 1) := fast_instance% Subtype.coe_injective.commMonoidWithZero _ coe_zero coe_one coe_mul coe_pow instance cancelMonoidWithZero {α : Type*} [OrderedRing α] [NoZeroDivisors α] : - CancelMonoidWithZero (Icc (0 : α) 1) := + CancelMonoidWithZero (Icc (0 : α) 1) := fast_instance% @Function.Injective.cancelMonoidWithZero α _ NoZeroDivisors.toCancelMonoidWithZero _ _ _ _ (fun v => v.val) Subtype.coe_injective coe_zero coe_one coe_mul coe_pow instance cancelCommMonoidWithZero {α : Type*} [OrderedCommRing α] [NoZeroDivisors α] : - CancelCommMonoidWithZero (Icc (0 : α) 1) := + CancelCommMonoidWithZero (Icc (0 : α) 1) := fast_instance% @Function.Injective.cancelCommMonoidWithZero α _ NoZeroDivisors.toCancelCommMonoidWithZero _ _ _ _ (fun v => v.val) Subtype.coe_injective coe_zero coe_one coe_mul coe_pow @@ -191,10 +192,11 @@ instance mul : Mul (Ico (0 : α) 1) where theorem coe_mul (x y : Ico (0 : α) 1) : ↑(x * y) = (x * y : α) := rfl -instance semigroup : Semigroup (Ico (0 : α) 1) := +instance semigroup : Semigroup (Ico (0 : α) 1) := fast_instance% Subtype.coe_injective.semigroup _ coe_mul -instance commSemigroup {α : Type*} [OrderedCommSemiring α] : CommSemigroup (Ico (0 : α) 1) := +instance commSemigroup {α : Type*} [OrderedCommSemiring α] : + CommSemigroup (Ico (0 : α) 1) := fast_instance% Subtype.coe_injective.commSemigroup _ coe_mul end Set.Ico @@ -250,17 +252,18 @@ theorem coe_mul (x y : Ioc (0 : α) 1) : ↑(x * y) = (x * y : α) := theorem coe_pow (x : Ioc (0 : α) 1) (n : ℕ) : ↑(x ^ n) = ((x : α) ^ n) := rfl -instance semigroup : Semigroup (Ioc (0 : α) 1) := +instance semigroup : Semigroup (Ioc (0 : α) 1) := fast_instance% Subtype.coe_injective.semigroup _ coe_mul -instance monoid : Monoid (Ioc (0 : α) 1) := +instance monoid : Monoid (Ioc (0 : α) 1) := fast_instance% Subtype.coe_injective.monoid _ coe_one coe_mul coe_pow -instance commSemigroup {α : Type*} [StrictOrderedCommSemiring α] : CommSemigroup (Ioc (0 : α) 1) := +instance commSemigroup {α : Type*} [StrictOrderedCommSemiring α] : + CommSemigroup (Ioc (0 : α) 1) := fast_instance% Subtype.coe_injective.commSemigroup _ coe_mul instance commMonoid {α : Type*} [StrictOrderedCommSemiring α] : - CommMonoid (Ioc (0 : α) 1) := + CommMonoid (Ioc (0 : α) 1) := fast_instance% Subtype.coe_injective.commMonoid _ coe_one coe_mul coe_pow instance cancelMonoid {α : Type*} [StrictOrderedRing α] [IsDomain α] : @@ -296,10 +299,11 @@ instance mul : Mul (Ioo (0 : α) 1) where theorem coe_mul (x y : Ioo (0 : α) 1) : ↑(x * y) = (x * y : α) := rfl -instance semigroup : Semigroup (Ioo (0 : α) 1) := +instance semigroup : Semigroup (Ioo (0 : α) 1) := fast_instance% Subtype.coe_injective.semigroup _ coe_mul -instance commSemigroup {α : Type*} [StrictOrderedCommSemiring α] : CommSemigroup (Ioo (0 : α) 1) := +instance commSemigroup {α : Type*} [StrictOrderedCommSemiring α] : + CommSemigroup (Ioo (0 : α) 1) := fast_instance% Subtype.coe_injective.commSemigroup _ coe_mul variable {β : Type*} [OrderedRing β] diff --git a/Mathlib/Algebra/Order/Monoid/Prod.lean b/Mathlib/Algebra/Order/Monoid/Prod.lean index a245abb7cdf2b..4a35b9544902a 100644 --- a/Mathlib/Algebra/Order/Monoid/Prod.lean +++ b/Mathlib/Algebra/Order/Monoid/Prod.lean @@ -48,19 +48,19 @@ namespace Lex instance orderedCommMonoid [OrderedCommMonoid α] [MulLeftStrictMono α] [OrderedCommMonoid β] : OrderedCommMonoid (α ×ₗ β) where - mul_le_mul_left _ _ hxy z := ((le_iff _ _).1 hxy).elim + mul_le_mul_left _ _ hxy z := (le_iff.1 hxy).elim (fun hxy => left _ _ <| mul_lt_mul_left' hxy _) -- Note: the `congr_arg` used to be `rw [hxy.1]` before https://github.com/leanprover-community/mathlib4/pull/8386 -- but the definition of `Mul.mul` got unfolded differently. - (fun hxy => (le_iff _ _).2 <| Or.inr ⟨congr_arg (z.1 * ·) hxy.1, mul_le_mul_left' hxy.2 _⟩) + (fun hxy => le_iff.2 <| Or.inr ⟨congr_arg (z.1 * ·) hxy.1, mul_le_mul_left' hxy.2 _⟩) @[to_additive] instance orderedCancelCommMonoid [OrderedCancelCommMonoid α] [OrderedCancelCommMonoid β] : OrderedCancelCommMonoid (α ×ₗ β) where mul_le_mul_left _ _ := mul_le_mul_left' - le_of_mul_le_mul_left _ _ _ hxyz := ((le_iff _ _).1 hxyz).elim + le_of_mul_le_mul_left _ _ _ hxyz := (le_iff.1 hxyz).elim (fun hxy => left _ _ <| lt_of_mul_lt_mul_left' hxy) - (fun hxy => (le_iff _ _).2 <| Or.inr ⟨mul_left_cancel hxy.1, le_of_mul_le_mul_left' hxy.2⟩) + (fun hxy => le_iff.2 <| Or.inr ⟨mul_left_cancel hxy.1, le_of_mul_le_mul_left' hxy.2⟩) @[to_additive] instance linearOrderedCancelCommMonoid [LinearOrderedCancelCommMonoid α] diff --git a/Mathlib/Algebra/Order/Monoid/Submonoid.lean b/Mathlib/Algebra/Order/Monoid/Submonoid.lean index f12cf9b2fc99b..ed546cdce5154 100644 --- a/Mathlib/Algebra/Order/Monoid/Submonoid.lean +++ b/Mathlib/Algebra/Order/Monoid/Submonoid.lean @@ -6,6 +6,7 @@ Authors: Damiano Testa import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Order.Interval.Set.Defs +import Mathlib.Tactic.FastInstance /-! # Ordered instances on submonoids @@ -20,7 +21,7 @@ variable {M S : Type*} [SetLike S M] /-- A submonoid of an `OrderedCommMonoid` is an `OrderedCommMonoid`. -/ @[to_additive "An `AddSubmonoid` of an `OrderedAddCommMonoid` is an `OrderedAddCommMonoid`."] instance (priority := 75) toOrderedCommMonoid [OrderedCommMonoid M] - [SubmonoidClass S M] (s : S) : OrderedCommMonoid s := + [SubmonoidClass S M] (s : S) : OrderedCommMonoid s := fast_instance% Subtype.coe_injective.orderedCommMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl -- Prefer subclasses of `Monoid` over subclasses of `SubmonoidClass`. @@ -28,7 +29,7 @@ instance (priority := 75) toOrderedCommMonoid [OrderedCommMonoid M] @[to_additive "An `AddSubmonoid` of a `LinearOrderedAddCommMonoid` is a `LinearOrderedAddCommMonoid`."] instance (priority := 75) toLinearOrderedCommMonoid [LinearOrderedCommMonoid M] - [SubmonoidClass S M] (s : S) : LinearOrderedCommMonoid s := + [SubmonoidClass S M] (s : S) : LinearOrderedCommMonoid s := fast_instance% Subtype.coe_injective.linearOrderedCommMonoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -37,7 +38,7 @@ instance (priority := 75) toLinearOrderedCommMonoid [LinearOrderedCommMonoid M] @[to_additive AddSubmonoidClass.toOrderedCancelAddCommMonoid "An `AddSubmonoid` of an `OrderedCancelAddCommMonoid` is an `OrderedCancelAddCommMonoid`."] instance (priority := 75) toOrderedCancelCommMonoid [OrderedCancelCommMonoid M] - [SubmonoidClass S M] (s : S) : OrderedCancelCommMonoid s := + [SubmonoidClass S M] (s : S) : OrderedCancelCommMonoid s := fast_instance% Subtype.coe_injective.orderedCancelCommMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl -- Prefer subclasses of `Monoid` over subclasses of `SubmonoidClass`. @@ -47,7 +48,7 @@ instance (priority := 75) toOrderedCancelCommMonoid [OrderedCancelCommMonoid M] "An `AddSubmonoid` of a `LinearOrderedCancelAddCommMonoid` is a `LinearOrderedCancelAddCommMonoid`."] instance (priority := 75) toLinearOrderedCancelCommMonoid [LinearOrderedCancelCommMonoid M] - [SubmonoidClass S M] (s : S) : LinearOrderedCancelCommMonoid s := + [SubmonoidClass S M] (s : S) : LinearOrderedCancelCommMonoid s := fast_instance% Subtype.coe_injective.linearOrderedCancelCommMonoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -59,14 +60,15 @@ variable {M : Type*} /-- A submonoid of an `OrderedCommMonoid` is an `OrderedCommMonoid`. -/ @[to_additive "An `AddSubmonoid` of an `OrderedAddCommMonoid` is an `OrderedAddCommMonoid`."] -instance toOrderedCommMonoid [OrderedCommMonoid M] (S : Submonoid M) : OrderedCommMonoid S := +instance toOrderedCommMonoid [OrderedCommMonoid M] (S : Submonoid M) : + OrderedCommMonoid S := fast_instance% Subtype.coe_injective.orderedCommMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- A submonoid of a `LinearOrderedCommMonoid` is a `LinearOrderedCommMonoid`. -/ @[to_additive "An `AddSubmonoid` of a `LinearOrderedAddCommMonoid` is a `LinearOrderedAddCommMonoid`."] instance toLinearOrderedCommMonoid [LinearOrderedCommMonoid M] (S : Submonoid M) : - LinearOrderedCommMonoid S := + LinearOrderedCommMonoid S := fast_instance% Subtype.coe_injective.linearOrderedCommMonoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -74,7 +76,7 @@ instance toLinearOrderedCommMonoid [LinearOrderedCommMonoid M] (S : Submonoid M) @[to_additive AddSubmonoid.toOrderedCancelAddCommMonoid "An `AddSubmonoid` of an `OrderedCancelAddCommMonoid` is an `OrderedCancelAddCommMonoid`."] instance toOrderedCancelCommMonoid [OrderedCancelCommMonoid M] (S : Submonoid M) : - OrderedCancelCommMonoid S := + OrderedCancelCommMonoid S := fast_instance% Subtype.coe_injective.orderedCancelCommMonoid Subtype.val rfl (fun _ _ => rfl) fun _ _ => rfl /-- A submonoid of a `LinearOrderedCancelCommMonoid` is a `LinearOrderedCancelCommMonoid`. @@ -83,7 +85,7 @@ instance toOrderedCancelCommMonoid [OrderedCancelCommMonoid M] (S : Submonoid M) "An `AddSubmonoid` of a `LinearOrderedCancelAddCommMonoid` is a `LinearOrderedCancelAddCommMonoid`."] instance toLinearOrderedCancelCommMonoid [LinearOrderedCancelCommMonoid M] (S : Submonoid M) : - LinearOrderedCancelCommMonoid S := + LinearOrderedCancelCommMonoid S := fast_instance% Subtype.coe_injective.linearOrderedCancelCommMonoid Subtype.val rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index a6577a705c6e3..aa2b6f7455877 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -144,13 +144,21 @@ section Preorder variable [Preorder α] @[to_additive] -lemma mul_left_mono [CovariantClass α α (· * ·) (· ≤ ·)] {a : α} : Monotone (a * ·) := +lemma mul_left_mono [MulLeftMono α] {a : α} : Monotone (a * ·) := fun _ _ h ↦ mul_le_mul_left' h _ @[to_additive] -lemma mul_right_mono [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a : α} : Monotone (· * a) := +lemma mul_right_mono [MulRightMono α] {a : α} : Monotone (· * a) := fun _ _ h ↦ mul_le_mul_right' h _ +@[to_additive] +lemma mul_left_strictMono [MulLeftStrictMono α] {a : α} : StrictMono (a * ·) := + fun _ _ h ↦ mul_lt_mul_left' h _ + +@[to_additive] +lemma mul_right_strictMono [MulRightStrictMono α] {a : α} : StrictMono (· * a) := + fun _ _ h ↦ mul_lt_mul_right' h _ + @[to_additive (attr := gcongr)] theorem mul_lt_mul_of_lt_of_lt [MulLeftStrictMono α] [MulRightStrictMono α] @@ -283,6 +291,24 @@ theorem mul_right_cancel'' [MulRightReflectLE α] {a b c : α} haveI := mulRightMono_of_mulRightStrictMono α rw [le_antisymm_iff, eq_true (mul_le_mul' hac hbd), true_and, mul_le_mul_iff_of_ge hac hbd] +@[to_additive] +lemma mul_left_inj_of_comparable [MulRightStrictMono α] {a b c : α} (h : b ≤ c ∨ c ≤ b) : + c * a = b * a ↔ c = b := by + refine ⟨fun h' => ?_, (· ▸ rfl)⟩ + contrapose h' + obtain h | h := h + · exact mul_lt_mul_right' (h.lt_of_ne' h') a |>.ne' + · exact mul_lt_mul_right' (h.lt_of_ne h') a |>.ne + +@[to_additive] +lemma mul_right_inj_of_comparable [MulLeftStrictMono α] {a b c : α} (h : b ≤ c ∨ c ≤ b) : + a * c = a * b ↔ c = b := by + refine ⟨fun h' => ?_, (· ▸ rfl)⟩ + contrapose h' + obtain h | h := h + · exact mul_lt_mul_left' (h.lt_of_ne' h') a |>.ne' + · exact mul_lt_mul_left' (h.lt_of_ne h') a |>.ne + end PartialOrder section LinearOrder @@ -325,6 +351,16 @@ lemma min_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) : haveI := mulRightMono_of_mulRightStrictMono α Left.min_le_max_of_mul_le_mul h +/-- Not an instance, to avoid loops with `IsLeftCancelMul.mulLeftStrictMono_of_mulLeftMono`. -/ +@[to_additive] +theorem MulLeftStrictMono.toIsLeftCancelMul [MulLeftStrictMono α] : IsLeftCancelMul α where + mul_left_cancel _ _ _ h := mul_left_strictMono.injective h + +/-- Not an instance, to avoid loops with `IsRightCancelMul.mulRightStrictMono_of_mulRightMono`. -/ +@[to_additive] +theorem MulRightStrictMono.toIsRightCancelMul [MulRightStrictMono α] : IsRightCancelMul α where + mul_right_cancel _ _ _ h := mul_right_strictMono.injective h + end LinearOrder section LinearOrder diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean index 1bd716addf327..c2e25592e58cd 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/ExistsOfLE.lean @@ -69,7 +69,7 @@ variable [LinearOrder α] [DenselyOrdered α] [Monoid α] [ExistsMulOfLE α] @[to_additive] theorem le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := - le_of_forall_le_of_dense fun x hxb => by + le_of_forall_gt_imp_ge_of_dense fun x hxb => by obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le exact h _ (one_lt_of_lt_mul_right hxb) diff --git a/Mathlib/Algebra/Order/Nonneg/Basic.lean b/Mathlib/Algebra/Order/Nonneg/Basic.lean index 31863d814ea0c..b802e57fd0526 100644 --- a/Mathlib/Algebra/Order/Nonneg/Basic.lean +++ b/Mathlib/Algebra/Order/Nonneg/Basic.lean @@ -242,13 +242,12 @@ instance commMonoidWithZero : CommMonoidWithZero { x : α // 0 ≤ x } := inferI end CommSemiring -section LinearOrder - -variable [Zero α] [LinearOrder α] +section SemilatticeSup +variable [Zero α] [SemilatticeSup α] /-- The function `a ↦ max a 0` of type `α → {x : α // 0 ≤ x}`. -/ def toNonneg (a : α) : { x : α // 0 ≤ x } := - ⟨max a 0, le_max_right _ _⟩ + ⟨max a 0, le_sup_right⟩ @[simp] theorem coe_toNonneg {a : α} : (toNonneg a : α) = max a 0 := @@ -266,11 +265,6 @@ theorem toNonneg_le {a : α} {b : { x : α // 0 ≤ x }} : toNonneg a ≤ b ↔ cases' b with b hb simp [toNonneg, hb] -@[simp] -theorem toNonneg_lt {a : { x : α // 0 ≤ x }} {b : α} : a < toNonneg b ↔ ↑a < b := by - cases' a with a ha - simp [toNonneg, ha.not_lt] - instance sub [Sub α] : Sub { x : α // 0 ≤ x } := ⟨fun x y => toNonneg (x - y)⟩ @@ -279,6 +273,16 @@ theorem mk_sub_mk [Sub α] {x y : α} (hx : 0 ≤ x) (hy : 0 ≤ y) : (⟨x, hx⟩ : { x : α // 0 ≤ x }) - ⟨y, hy⟩ = toNonneg (x - y) := rfl +end SemilatticeSup + +section LinearOrder +variable [Zero α] [LinearOrder α] + +@[simp] +theorem toNonneg_lt {a : { x : α // 0 ≤ x }} {b : α} : a < toNonneg b ↔ ↑a < b := by + cases' a with a ha + simp [toNonneg, ha.not_lt] + end LinearOrder end Nonneg diff --git a/Mathlib/Algebra/Order/Nonneg/Field.lean b/Mathlib/Algebra/Order/Nonneg/Field.lean index 49c06d2e70cdd..1ffb11a94f50a 100644 --- a/Mathlib/Algebra/Order/Nonneg/Field.lean +++ b/Mathlib/Algebra/Order/Nonneg/Field.lean @@ -94,7 +94,7 @@ instance instNNRatSMul : SMul ℚ≥0 {x : α // 0 ≤ x} where (⟨q • a, by rw [NNRat.smul_def]; exact mul_nonneg q.cast_nonneg ha⟩ : {x : α // 0 ≤ x}) = q • a := rfl -instance linearOrderedSemifield : LinearOrderedSemifield { x : α // 0 ≤ x } := +instance linearOrderedSemifield : LinearOrderedSemifield { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.linearOrderedSemifield _ Nonneg.coe_zero Nonneg.coe_one Nonneg.coe_add Nonneg.coe_mul Nonneg.coe_inv Nonneg.coe_div (fun _ _ => rfl) coe_nnqsmul Nonneg.coe_pow Nonneg.coe_zpow Nonneg.coe_natCast coe_nnratCast (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/Algebra/Order/Nonneg/Ring.lean b/Mathlib/Algebra/Order/Nonneg/Ring.lean index 6c20dd04819c4..742ffb82c8154 100644 --- a/Mathlib/Algebra/Order/Nonneg/Ring.lean +++ b/Mathlib/Algebra/Order/Nonneg/Ring.lean @@ -11,6 +11,7 @@ import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Order.Ring.InjSurj import Mathlib.Order.CompleteLatticeIntervals import Mathlib.Order.LatticeIntervals +import Mathlib.Tactic.FastInstance /-! # Bundled ordered algebra instance on the type of nonnegative elements @@ -43,37 +44,39 @@ variable {α : Type*} namespace Nonneg -instance orderedAddCommMonoid [OrderedAddCommMonoid α] : OrderedAddCommMonoid { x : α // 0 ≤ x } := +instance orderedAddCommMonoid [OrderedAddCommMonoid α] : + OrderedAddCommMonoid { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.orderedAddCommMonoid _ Nonneg.coe_zero (fun _ _ => rfl) fun _ _ => rfl instance linearOrderedAddCommMonoid [LinearOrderedAddCommMonoid α] : - LinearOrderedAddCommMonoid { x : α // 0 ≤ x } := + LinearOrderedAddCommMonoid { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.linearOrderedAddCommMonoid _ Nonneg.coe_zero (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl instance orderedCancelAddCommMonoid [OrderedCancelAddCommMonoid α] : - OrderedCancelAddCommMonoid { x : α // 0 ≤ x } := + OrderedCancelAddCommMonoid { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.orderedCancelAddCommMonoid _ Nonneg.coe_zero (fun _ _ => rfl) fun _ _ => rfl instance linearOrderedCancelAddCommMonoid [LinearOrderedCancelAddCommMonoid α] : - LinearOrderedCancelAddCommMonoid { x : α // 0 ≤ x } := + LinearOrderedCancelAddCommMonoid { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.linearOrderedCancelAddCommMonoid _ Nonneg.coe_zero (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -instance orderedSemiring [OrderedSemiring α] : OrderedSemiring { x : α // 0 ≤ x } := +instance orderedSemiring [OrderedSemiring α] : OrderedSemiring { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.orderedSemiring _ Nonneg.coe_zero Nonneg.coe_one (fun _ _ => rfl) (fun _ _=> rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl instance strictOrderedSemiring [StrictOrderedSemiring α] : - StrictOrderedSemiring { x : α // 0 ≤ x } := + StrictOrderedSemiring { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.strictOrderedSemiring _ Nonneg.coe_zero Nonneg.coe_one (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl -instance orderedCommSemiring [OrderedCommSemiring α] : OrderedCommSemiring { x : α // 0 ≤ x } := +instance orderedCommSemiring [OrderedCommSemiring α] : + OrderedCommSemiring { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.orderedCommSemiring _ Nonneg.coe_zero Nonneg.coe_one (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl @@ -82,7 +85,7 @@ instance orderedCommMonoid [OrderedCommSemiring α] : OrderedCommMonoid { x : α mul_le_mul_left a _ h c := mul_le_mul le_rfl h a.prop c.prop instance strictOrderedCommSemiring [StrictOrderedCommSemiring α] : - StrictOrderedCommSemiring { x : α // 0 ≤ x } := + StrictOrderedCommSemiring { x : α // 0 ≤ x } := fast_instance% Subtype.coe_injective.strictOrderedCommSemiring _ Nonneg.coe_zero Nonneg.coe_one (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl diff --git a/Mathlib/Algebra/Order/Positive/Ring.lean b/Mathlib/Algebra/Order/Positive/Ring.lean index ff395b0df05ef..73dd89bc15ab3 100644 --- a/Mathlib/Algebra/Order/Positive/Ring.lean +++ b/Mathlib/Algebra/Order/Positive/Ring.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.InjSurj +import Mathlib.Tactic.FastInstance /-! # Algebraic structures on the set of positive numbers @@ -32,19 +33,19 @@ instance : Add { x : M // 0 < x } := theorem coe_add (x y : { x : M // 0 < x }) : ↑(x + y) = (x + y : M) := rfl -instance addSemigroup : AddSemigroup { x : M // 0 < x } := +instance addSemigroup : AddSemigroup { x : M // 0 < x } := fast_instance% Subtype.coe_injective.addSemigroup _ coe_add instance addCommSemigroup {M : Type*} [AddCommMonoid M] [Preorder M] - [AddLeftStrictMono M] : AddCommSemigroup { x : M // 0 < x } := + [AddLeftStrictMono M] : AddCommSemigroup { x : M // 0 < x } := fast_instance% Subtype.coe_injective.addCommSemigroup _ coe_add instance addLeftCancelSemigroup {M : Type*} [AddLeftCancelMonoid M] [Preorder M] - [AddLeftStrictMono M] : AddLeftCancelSemigroup { x : M // 0 < x } := + [AddLeftStrictMono M] : AddLeftCancelSemigroup { x : M // 0 < x } := fast_instance% Subtype.coe_injective.addLeftCancelSemigroup _ coe_add instance addRightCancelSemigroup {M : Type*} [AddRightCancelMonoid M] [Preorder M] - [AddLeftStrictMono M] : AddRightCancelSemigroup { x : M // 0 < x } := + [AddLeftStrictMono M] : AddRightCancelSemigroup { x : M // 0 < x } := fast_instance% Subtype.coe_injective.addRightCancelSemigroup _ coe_add instance addLeftStrictMono : AddLeftStrictMono { x : M // 0 < x } := @@ -90,10 +91,10 @@ theorem val_pow (x : { x : R // 0 < x }) (n : ℕ) : ↑(x ^ n) = (x : R) ^ n := rfl -instance : Semigroup { x : R // 0 < x } := +instance : Semigroup { x : R // 0 < x } := fast_instance% Subtype.coe_injective.semigroup Subtype.val val_mul -instance : Distrib { x : R // 0 < x } := +instance : Distrib { x : R // 0 < x } := fast_instance% Subtype.coe_injective.distrib _ coe_add val_mul instance : One { x : R // 0 < x } := @@ -103,7 +104,7 @@ instance : One { x : R // 0 < x } := theorem val_one : ((1 : { x : R // 0 < x }) : R) = 1 := rfl -instance : Monoid { x : R // 0 < x } := +instance : Monoid { x : R // 0 < x } := fast_instance% Subtype.coe_injective.monoid _ val_one val_mul val_pow end Mul @@ -111,7 +112,7 @@ end Mul section mul_comm instance orderedCommMonoid [StrictOrderedCommSemiring R] : - OrderedCommMonoid { x : R // 0 < x } := + OrderedCommMonoid { x : R // 0 < x } := fast_instance% { Subtype.partialOrder _, Subtype.coe_injective.commMonoid (M₂ := R) (Subtype.val) val_one val_mul val_pow with mul_le_mul_left := fun _ _ hxy c => diff --git a/Mathlib/Algebra/Order/Rearrangement.lean b/Mathlib/Algebra/Order/Rearrangement.lean index b44e25f7c4d85..d9ab756af9db1 100644 --- a/Mathlib/Algebra/Order/Rearrangement.lean +++ b/Mathlib/Algebra/Order/Rearrangement.lean @@ -99,12 +99,12 @@ theorem MonovaryOn.sum_smul_comp_perm_le_sum_smul (hfg : MonovaryOn f g s) simp only [hτ, swap_apply_left, Function.comp_apply, Equiv.coe_trans, apply_inv_self] refine add_le_add (smul_add_smul_le_smul_add_smul' ?_ ?_) (sum_congr rfl fun x hx ↦ ?_).le · specialize hamax (σ⁻¹ a) h1s - rw [Prod.Lex.le_iff] at hamax + rw [Prod.Lex.toLex_le_toLex] at hamax cases' hamax with hamax hamax · exact hfg (mem_insert_of_mem h1s) (mem_insert_self _ _) hamax · exact hamax.2 · specialize hamax (σ a) (mem_of_mem_insert_of_ne (hσ <| σ.injective.ne hσa.symm) hσa.symm) - rw [Prod.Lex.le_iff] at hamax + rw [Prod.Lex.toLex_le_toLex] at hamax cases' hamax with hamax hamax · exact hamax.le · exact hamax.1.le diff --git a/Mathlib/Algebra/Order/Ring/WithTop.lean b/Mathlib/Algebra/Order/Ring/WithTop.lean index 5ba1619244795..894e0aabbc46c 100644 --- a/Mathlib/Algebra/Order/Ring/WithTop.lean +++ b/Mathlib/Algebra/Order/Ring/WithTop.lean @@ -24,15 +24,15 @@ variable [MulZeroClass α] {a b : WithTop α} instance instMulZeroClass : MulZeroClass (WithTop α) where zero := 0 - mul a b := match a, b with + mul | (a : α), (b : α) => ↑(a * b) | (a : α), ⊤ => if a = 0 then 0 else ⊤ | ⊤, (b : α) => if b = 0 then 0 else ⊤ | ⊤, ⊤ => ⊤ - mul_zero a := match a with + mul_zero | (a : α) => congr_arg some <| mul_zero _ | ⊤ => if_pos rfl - zero_mul b := match b with + zero_mul | (b : α) => congr_arg some <| zero_mul _ | ⊤ => if_pos rfl @@ -93,10 +93,10 @@ end MulZeroClass /-- `Nontrivial α` is needed here as otherwise we have `1 * ⊤ = ⊤` but also `0 * ⊤ = 0`. -/ instance instMulZeroOneClass [MulZeroOneClass α] [Nontrivial α] : MulZeroOneClass (WithTop α) where __ := instMulZeroClass - one_mul a := match a with + one_mul | ⊤ => mul_top (mt coe_eq_coe.1 one_ne_zero) | (a : α) => by rw [← coe_one, ← coe_mul, one_mul] - mul_one a := match a with + mul_one | ⊤ => top_mul (mt coe_eq_coe.1 one_ne_zero) | (a : α) => by rw [← coe_one, ← coe_mul, mul_one] diff --git a/Mathlib/Algebra/Order/Round.lean b/Mathlib/Algebra/Order/Round.lean new file mode 100644 index 0000000000000..126741c7c3750 --- /dev/null +++ b/Mathlib/Algebra/Order/Round.lean @@ -0,0 +1,214 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Kevin Kappelmann +-/ +import Mathlib.Algebra.Order.Floor +import Mathlib.Algebra.Order.Interval.Set.Group + +/-! +# Rounding + +This file defines the `round` function, which uses the `floor` or `ceil` function to round a number +to the nearest integer. + +## Main Definitions + +* `round a`: Nearest integer to `a`. It rounds halves towards infinity. + +## Tags + +rounding +-/ + +assert_not_exists Finset + +open Set + +variable {F α β : Type*} + +open Int + +/-! ### Round -/ + +section round + +section LinearOrderedRing + +variable [LinearOrderedRing α] [FloorRing α] + +/-- `round` rounds a number to the nearest integer. `round (1 / 2) = 1` -/ +def round (x : α) : ℤ := + if 2 * fract x < 1 then ⌊x⌋ else ⌈x⌉ + +@[simp] +theorem round_zero : round (0 : α) = 0 := by simp [round] + +@[simp] +theorem round_one : round (1 : α) = 1 := by simp [round] + +@[simp] +theorem round_natCast (n : ℕ) : round (n : α) = n := by simp [round] + +@[simp] +theorem round_ofNat (n : ℕ) [n.AtLeastTwo] : round (ofNat(n) : α) = ofNat(n) := + round_natCast n + +@[simp] +theorem round_intCast (n : ℤ) : round (n : α) = n := by simp [round] + +@[simp] +theorem round_add_int (x : α) (y : ℤ) : round (x + y) = round x + y := by + rw [round, round, Int.fract_add_int, Int.floor_add_int, Int.ceil_add_int, ← apply_ite₂, ite_self] + +@[simp] +theorem round_add_one (a : α) : round (a + 1) = round a + 1 := by + -- Porting note: broken `convert round_add_int a 1` + rw [← round_add_int a 1, cast_one] + +@[simp] +theorem round_sub_int (x : α) (y : ℤ) : round (x - y) = round x - y := by + rw [sub_eq_add_neg] + norm_cast + rw [round_add_int, sub_eq_add_neg] + +@[simp] +theorem round_sub_one (a : α) : round (a - 1) = round a - 1 := by + -- Porting note: broken `convert round_sub_int a 1` + rw [← round_sub_int a 1, cast_one] + +@[simp] +theorem round_add_nat (x : α) (y : ℕ) : round (x + y) = round x + y := + mod_cast round_add_int x y + +@[simp] +theorem round_add_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : + round (x + ofNat(n)) = round x + ofNat(n) := + round_add_nat x n + +@[simp] +theorem round_sub_nat (x : α) (y : ℕ) : round (x - y) = round x - y := + mod_cast round_sub_int x y + +@[simp] +theorem round_sub_ofNat (x : α) (n : ℕ) [n.AtLeastTwo] : + round (x - ofNat(n)) = round x - ofNat(n) := + round_sub_nat x n + +@[simp] +theorem round_int_add (x : α) (y : ℤ) : round ((y : α) + x) = y + round x := by + rw [add_comm, round_add_int, add_comm] + +@[simp] +theorem round_nat_add (x : α) (y : ℕ) : round ((y : α) + x) = y + round x := by + rw [add_comm, round_add_nat, add_comm] + +@[simp] +theorem round_ofNat_add (n : ℕ) [n.AtLeastTwo] (x : α) : + round (ofNat(n) + x) = ofNat(n) + round x := + round_nat_add x n + +theorem abs_sub_round_eq_min (x : α) : |x - round x| = min (fract x) (1 - fract x) := by + simp_rw [round, min_def_lt, two_mul, ← lt_tsub_iff_left] + cases' lt_or_ge (fract x) (1 - fract x) with hx hx + · rw [if_pos hx, if_pos hx, self_sub_floor, abs_fract] + · have : 0 < fract x := by + replace hx : 0 < fract x + fract x := lt_of_lt_of_le zero_lt_one (tsub_le_iff_left.mp hx) + simpa only [← two_mul, mul_pos_iff_of_pos_left, zero_lt_two] using hx + rw [if_neg (not_lt.mpr hx), if_neg (not_lt.mpr hx), abs_sub_comm, ceil_sub_self_eq this.ne.symm, + abs_one_sub_fract] + +theorem round_le (x : α) (z : ℤ) : |x - round x| ≤ |x - z| := by + rw [abs_sub_round_eq_min, min_le_iff] + rcases le_or_lt (z : α) x with (hx | hx) <;> [left; right] + · conv_rhs => rw [abs_eq_self.mpr (sub_nonneg.mpr hx), ← fract_add_floor x, add_sub_assoc] + simpa only [le_add_iff_nonneg_right, sub_nonneg, cast_le] using le_floor.mpr hx + · rw [abs_eq_neg_self.mpr (sub_neg.mpr hx).le] + conv_rhs => rw [← fract_add_floor x] + rw [add_sub_assoc, add_comm, neg_add, neg_sub, le_add_neg_iff_add_le, sub_add_cancel, + le_sub_comm] + norm_cast + exact floor_le_sub_one_iff.mpr hx + +end LinearOrderedRing + +section LinearOrderedField + +variable [LinearOrderedField α] [FloorRing α] + +theorem round_eq (x : α) : round x = ⌊x + 1 / 2⌋ := by + simp_rw [round, (by simp only [lt_div_iff₀', two_pos] : 2 * fract x < 1 ↔ fract x < 1 / 2)] + cases' lt_or_le (fract x) (1 / 2) with hx hx + · conv_rhs => rw [← fract_add_floor x, add_assoc, add_left_comm, floor_int_add] + rw [if_pos hx, self_eq_add_right, floor_eq_iff, cast_zero, zero_add] + constructor + · linarith [fract_nonneg x] + · linarith + · have : ⌊fract x + 1 / 2⌋ = 1 := by + rw [floor_eq_iff] + constructor + · norm_num + linarith + · norm_num + linarith [fract_lt_one x] + rw [if_neg (not_lt.mpr hx), ← fract_add_floor x, add_assoc, add_left_comm, floor_int_add, + ceil_add_int, add_comm _ ⌊x⌋, add_right_inj, ceil_eq_iff, this, cast_one, sub_self] + constructor + · linarith + · linarith [fract_lt_one x] + +@[simp] +theorem round_two_inv : round (2⁻¹ : α) = 1 := by + simp only [round_eq, ← one_div, add_halves, floor_one] + +@[simp] +theorem round_neg_two_inv : round (-2⁻¹ : α) = 0 := by + simp only [round_eq, ← one_div, neg_add_cancel, floor_zero] + +@[simp] +theorem round_eq_zero_iff {x : α} : round x = 0 ↔ x ∈ Ico (-(1 / 2)) ((1 : α) / 2) := by + rw [round_eq, floor_eq_zero_iff, add_mem_Ico_iff_left] + norm_num + +theorem abs_sub_round (x : α) : |x - round x| ≤ 1 / 2 := by + rw [round_eq, abs_sub_le_iff] + have := floor_le (x + 1 / 2) + have := lt_floor_add_one (x + 1 / 2) + constructor <;> linarith + +theorem abs_sub_round_div_natCast_eq {m n : ℕ} : + |(m : α) / n - round ((m : α) / n)| = ↑(min (m % n) (n - m % n)) / n := by + rcases n.eq_zero_or_pos with (rfl | hn) + · simp + have hn' : 0 < (n : α) := by + norm_cast + rw [abs_sub_round_eq_min, Nat.cast_min, ← min_div_div_right hn'.le, + fract_div_natCast_eq_div_natCast_mod, Nat.cast_sub (m.mod_lt hn).le, sub_div, div_self hn'.ne'] + +@[bound] +theorem sub_half_lt_round (x : α) : x - 1 / 2 < round x := by + rw [round_eq x, show x - 1 / 2 = x + 1 / 2 - 1 by linarith] + exact Int.sub_one_lt_floor (x + 1 / 2) + +@[bound] +theorem round_le_add_half (x : α) : round x ≤ x + 1 / 2 := by + rw [round_eq x] + exact Int.floor_le (x + 1 / 2) + +end LinearOrderedField + +end round + +namespace Int + +variable [LinearOrderedField α] [LinearOrderedField β] [FloorRing α] [FloorRing β] +variable [FunLike F α β] [RingHomClass F α β] {a : α} {b : β} + +theorem map_round (f : F) (hf : StrictMono f) (a : α) : round (f a) = round a := by + have H : f 2 = 2 := map_natCast f 2 + simp_rw [round_eq, ← map_floor _ hf, map_add, one_div, map_inv₀, H] + -- Porting note: was + -- simp_rw [round_eq, ← map_floor _ hf, map_add, one_div, map_inv₀, map_bit0, map_one] + -- Would have thought that `map_natCast` would replace `map_bit0, map_one` but seems not + +end Int diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 40e975a43fbfe..2129497e1006c 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ import Mathlib.Algebra.Algebra.Pi +import Mathlib.Algebra.Algebra.Prod import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.MonoidAlgebra.Basic @@ -368,6 +369,38 @@ theorem aeval_algebraMap_apply_eq_algebraMap_eval (x : R) (p : R[X]) : aeval (algebraMap R A x) p = algebraMap R A (p.eval x) := aeval_algHom_apply (Algebra.ofId R A) x p +/-- Polynomial evaluation on a pair is a product of the evaluations on the components. -/ +theorem aeval_prod (x : A × B) : aeval (R := R) x = (aeval x.1).prod (aeval x.2) := + aeval_algHom (.fst R A B) x ▸ aeval_algHom (.snd R A B) x ▸ + (aeval x).prod_comp (.fst R A B) (.snd R A B) + +/-- Polynomial evaluation on a pair is a pair of evaluations. -/ +theorem aeval_prod_apply (x : A × B) (p : Polynomial R) : + p.aeval x = (p.aeval x.1, p.aeval x.2) := by simp [aeval_prod] + +section Pi + +variable {I : Type*} {A : I → Type*} [∀ i, Semiring (A i)] [∀ i, Algebra R (A i)] +variable (x : Π i, A i) (p : R[X]) + +/-- Polynomial evaluation on an indexed tuple is the indexed product of the evaluations +on the components. +Generalizes `Polynomial.aeval_prod` to indexed products. -/ +theorem aeval_pi (x : Π i, A i) : aeval (R := R) x = Pi.algHom R A (fun i ↦ aeval (x i)) := + (funext fun i ↦ aeval_algHom (Pi.evalAlgHom R A i) x) ▸ + (Pi.algHom_comp R A (Pi.evalAlgHom R A) (aeval x)) + +theorem aeval_pi_apply₂ (j : I) : p.aeval x j = p.aeval (x j) := + aeval_pi (R := R) x ▸ Pi.algHom_apply R A (fun i ↦ aeval (x i)) p j + +/-- Polynomial evaluation on an indexed tuple is the indexed tuple of the evaluations +on the components. +Generalizes `Polynomial.aeval_prod_apply` to indexed products. -/ +theorem aeval_pi_apply : p.aeval x = fun j ↦ p.aeval (x j) := + funext fun j ↦ aeval_pi_apply₂ x p j + +end Pi + @[simp] theorem coe_aeval_eq_eval (r : R) : (aeval r : R[X] → R) = eval r := rfl diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index d88f59b1d6cca..880597eb13c1f 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -7,6 +7,8 @@ import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.MonoidAlgebra.Defs import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Data.Finset.Sort +import Mathlib.Tactic.FastInstance +import Mathlib.Algebra.Group.Submonoid.Operations /-! # Theory of univariate polynomials @@ -52,7 +54,8 @@ in general not be used once the basic API for polynomials is constructed. noncomputable section -/-- `Polynomial R` is the type of univariate polynomials over `R`. +/-- `Polynomial R` is the type of univariate polynomials over `R`, +denoted as `R[X]` within the `Polynomial` namespace. Polynomials should be seen as (semi-)rings with the additional constructor `X`. The embedding from `R` is called `C`. -/ @@ -266,37 +269,25 @@ instance inhabited : Inhabited R[X] := instance instNatCast : NatCast R[X] where natCast n := ofFinsupp n instance semiring : Semiring R[X] := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { Function.Injective.semiring toFinsupp toFinsupp_injective toFinsupp_zero toFinsupp_one - toFinsupp_add toFinsupp_mul (fun _ _ => toFinsupp_nsmul _ _) toFinsupp_pow fun _ => rfl with - toAdd := Polynomial.add' - toMul := Polynomial.mul' - toZero := Polynomial.zero - toOne := Polynomial.one - nsmul := (· • ·) - npow := fun n x => (x ^ n) } + fast_instance% Function.Injective.semiring toFinsupp toFinsupp_injective toFinsupp_zero + toFinsupp_one toFinsupp_add toFinsupp_mul (fun _ _ => toFinsupp_nsmul _ _) toFinsupp_pow + fun _ => rfl instance distribSMul {S} [DistribSMul S R] : DistribSMul S R[X] := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { Function.Injective.distribSMul ⟨⟨toFinsupp, toFinsupp_zero⟩, toFinsupp_add⟩ toFinsupp_injective - toFinsupp_smul with - toSMulZeroClass := Polynomial.smulZeroClass } + fast_instance% Function.Injective.distribSMul ⟨⟨toFinsupp, toFinsupp_zero⟩, toFinsupp_add⟩ + toFinsupp_injective toFinsupp_smul instance distribMulAction {S} [Monoid S] [DistribMulAction S R] : DistribMulAction S R[X] := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { Function.Injective.distribMulAction ⟨⟨toFinsupp, toFinsupp_zero (R := R)⟩, toFinsupp_add⟩ - toFinsupp_injective toFinsupp_smul with - toSMul := Polynomial.smulZeroClass.toSMul } + fast_instance% Function.Injective.distribMulAction + ⟨⟨toFinsupp, toFinsupp_zero (R := R)⟩, toFinsupp_add⟩ toFinsupp_injective toFinsupp_smul instance faithfulSMul {S} [SMulZeroClass S R] [FaithfulSMul S R] : FaithfulSMul S R[X] where eq_of_smul_eq_smul {_s₁ _s₂} h := eq_of_smul_eq_smul fun a : ℕ →₀ R => congr_arg toFinsupp (h ⟨a⟩) instance module {S} [Semiring S] [Module S R] : Module S R[X] := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { Function.Injective.module _ ⟨⟨toFinsupp, toFinsupp_zero⟩, toFinsupp_add⟩ toFinsupp_injective - toFinsupp_smul with - toDistribMulAction := Polynomial.distribMulAction } + fast_instance% Function.Injective.module _ ⟨⟨toFinsupp, toFinsupp_zero⟩, toFinsupp_add⟩ + toFinsupp_injective toFinsupp_smul instance smulCommClass {S₁ S₂} [SMulZeroClass S₁ R] [SMulZeroClass S₂ R] [SMulCommClass S₁ S₂ R] : SMulCommClass S₁ S₂ R[X] := @@ -430,7 +421,7 @@ theorem monomial_pow (n : ℕ) (r : R) (k : ℕ) : monomial n r ^ k = monomial ( theorem smul_monomial {S} [SMulZeroClass S R] (a : S) (n : ℕ) (b : R) : a • monomial n b = monomial n (a • b) := - toFinsupp_injective <| by simp; rw [smul_single] + toFinsupp_injective <| AddMonoidAlgebra.smul_single _ _ _ theorem monomial_injective (n : ℕ) : Function.Injective (monomial n : R → R[X]) := (toFinsuppIso R).symm.injective.comp (single_injective n) @@ -1051,7 +1042,7 @@ section CommSemiring variable [CommSemiring R] instance commSemiring : CommSemiring R[X] := - { Function.Injective.commSemigroup toFinsupp toFinsupp_injective toFinsupp_mul with + fast_instance% { Function.Injective.commSemigroup toFinsupp toFinsupp_injective toFinsupp_mul with toSemiring := Polynomial.semiring } end CommSemiring @@ -1076,15 +1067,10 @@ theorem toFinsupp_zsmul (a : ℤ) (b : R[X]) : instance instIntCast : IntCast R[X] where intCast n := ofFinsupp n instance ring : Ring R[X] := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { Function.Injective.ring toFinsupp toFinsupp_injective (toFinsupp_zero (R := R)) + fast_instance% Function.Injective.ring toFinsupp toFinsupp_injective (toFinsupp_zero (R := R)) toFinsupp_one toFinsupp_add toFinsupp_mul toFinsupp_neg toFinsupp_sub (fun _ _ => toFinsupp_nsmul _ _) - (fun _ _ => toFinsupp_zsmul _ _) toFinsupp_pow (fun _ => rfl) fun _ => rfl with - toSemiring := Polynomial.semiring, - toNeg := Polynomial.neg' - toSub := Polynomial.sub - zsmul := ((· • ·) : ℤ → R[X] → R[X]) } + (fun _ _ => toFinsupp_zsmul _ _) toFinsupp_pow (fun _ => rfl) fun _ => rfl @[simp] theorem coeff_neg (p : R[X]) (n : ℕ) : coeff (-p) n = -coeff p n := by diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index 48d16d1d4a6f2..6d1a1c3d96070 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -111,12 +111,12 @@ noncomputable def divModByMonicAux : ∀ (_p : R[X]) {q : R[X]}, Monic q → R[X else ⟨0, p⟩ termination_by p => p -/-- `divByMonic` gives the quotient of `p` by a monic polynomial `q`. -/ +/-- `divByMonic`, denoted as `p /ₘ q`, gives the quotient of `p` by a monic polynomial `q`. -/ def divByMonic (p q : R[X]) : R[X] := letI := Classical.decEq R if hq : Monic q then (divModByMonicAux p hq).1 else 0 -/-- `modByMonic` gives the remainder of `p` by a monic polynomial `q`. -/ +/-- `modByMonic`, denoted as `p %ₘ q`, gives the remainder of `p` by a monic polynomial `q`. -/ def modByMonic (p q : R[X]) : R[X] := letI := Classical.decEq R if hq : Monic q then (divModByMonicAux p hq).2 else p diff --git a/Mathlib/Algebra/Polynomial/Eval/Defs.lean b/Mathlib/Algebra/Polynomial/Eval/Defs.lean index 088010ca51334..1cb42ee480336 100644 --- a/Mathlib/Algebra/Polynomial/Eval/Defs.lean +++ b/Mathlib/Algebra/Polynomial/Eval/Defs.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ +import Mathlib.Algebra.Group.Nat.Hom import Mathlib.Algebra.Polynomial.Basic /-! diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 07d250cccea6d..6c59297f72e39 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -187,11 +187,7 @@ theorem isRoot_of_isRoot_of_dvd_derivative_mul [CharZero R] {f g : R[X]} (hf0 : rw [rootMultiplicity_mul hdfg0, derivative_rootMultiplicity_of_root haf, rootMultiplicity_eq_zero hg, add_zero, rootMultiplicity_mul (hr ▸ hdfg0), add_comm, Nat.sub_eq_iff_eq_add (Nat.succ_le_iff.2 ((rootMultiplicity_pos hf0).2 haf))] at hr' - refine lt_irrefl (rootMultiplicity a f) ?_ - refine lt_of_lt_of_le (Nat.lt_succ_self _) - (le_trans (le_add_of_nonneg_left (Nat.zero_le (rootMultiplicity a r))) ?_) - conv_rhs => rw [hr'] - simp [add_assoc] + omega section NormalizationMonoid diff --git a/Mathlib/Algebra/Polynomial/Monomial.lean b/Mathlib/Algebra/Polynomial/Monomial.lean index f994bf2057858..c22c4afba88b1 100644 --- a/Mathlib/Algebra/Polynomial/Monomial.lean +++ b/Mathlib/Algebra/Polynomial/Monomial.lean @@ -3,12 +3,11 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ +import Mathlib.Algebra.Group.Nat.Hom import Mathlib.Algebra.Polynomial.Basic /-! # Univariate monomials - -Preparatory lemmas for degree_basic. -/ diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index fc5bcda77af79..bfc6a75772dcc 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -3,12 +3,9 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Algebra.Algebra.Equiv import Mathlib.Algebra.Star.SelfAdjoint import Mathlib.LinearAlgebra.Dimension.StrongRankCondition -import Mathlib.LinearAlgebra.FreeModule.Basic import Mathlib.LinearAlgebra.FreeModule.Finite.Basic -import Mathlib.SetTheory.Cardinal.Arithmetic /-! # Quaternions @@ -18,15 +15,18 @@ algebraic structures on `ℍ[R]`. ## Main definitions -* `QuaternionAlgebra R a b`, `ℍ[R, a, b]` : - [quaternion algebra](https://en.wikipedia.org/wiki/Quaternion_algebra) with coefficients `a`, `b` -* `Quaternion R`, `ℍ[R]` : the space of quaternions, a.k.a. `QuaternionAlgebra R (-1) (-1)`; +* `QuaternionAlgebra R a b c`, `ℍ[R, a, b, c]` : + [Bourbaki, *Algebra I*][bourbaki1989] with coefficients `a`, `b`, `c` + (Many other references such as Wikipedia assume $\operatorname{char} R ≠ 2$ therefore one can + complete the square and WLOG assume $b = 0$.) +* `Quaternion R`, `ℍ[R]` : the space of quaternions, a.k.a. + `QuaternionAlgebra R (-1) (0) (-1)`; * `Quaternion.normSq` : square of the norm of a quaternion; We also define the following algebraic structures on `ℍ[R]`: -* `Ring ℍ[R, a, b]`, `StarRing ℍ[R, a, b]`, and `Algebra R ℍ[R, a, b]` : for any commutative ring - `R`; +* `Ring ℍ[R, a, b, c]`, `StarRing ℍ[R, a, b, c]`, and `Algebra R ℍ[R, a, b, c]` : + for any commutative ring `R`; * `Ring ℍ[R]`, `StarRing ℍ[R]`, and `Algebra R ℍ[R]` : for any commutative ring `R`; * `IsDomain ℍ[R]` : for a linear ordered commutative ring `R`; * `DivisionRing ℍ[R]` : for a linear ordered field `R`. @@ -35,7 +35,8 @@ We also define the following algebraic structures on `ℍ[R]`: The following notation is available with `open Quaternion` or `open scoped Quaternion`. -* `ℍ[R, c₁, c₂]` : `QuaternionAlgebra R c₁ c₂` +* `ℍ[R, c₁, c₂, c₃]` : `QuaternionAlgebra R c₁ c₂ c₃` +* `ℍ[R, c₁, c₂]` : `QuaternionAlgebra R c₁ 0 c₂` * `ℍ[R]` : quaternions over `R`. ## Implementation notes @@ -50,10 +51,11 @@ quaternion -/ -/-- Quaternion algebra over a type with fixed coefficients $a=i^2$ and $b=j^2$. +/-- Quaternion algebra over a type with fixed coefficients where $i^2 = a + bi$ and $j^2 = c$, +denoted as `ℍ[R,a,b]`. Implemented as a structure with four fields: `re`, `imI`, `imJ`, and `imK`. -/ @[ext] -structure QuaternionAlgebra (R : Type*) (a b : R) where +structure QuaternionAlgebra (R : Type*) (a b c : R) where /-- Real part of a quaternion. -/ re : R /-- First imaginary part (i) of a quaternion. -/ @@ -64,14 +66,18 @@ structure QuaternionAlgebra (R : Type*) (a b : R) where imK : R @[inherit_doc] -scoped[Quaternion] notation "ℍ[" R "," a "," b "]" => QuaternionAlgebra R a b -open Quaternion +scoped[Quaternion] notation "ℍ[" R "," a "," b "," c "]" => + QuaternionAlgebra R a b c + +@[inherit_doc] +scoped[Quaternion] notation "ℍ[" R "," a "," b "]" => QuaternionAlgebra R a 0 b namespace QuaternionAlgebra +open Quaternion /-- The equivalence between a quaternion algebra over `R` and `R × R × R × R`. -/ @[simps] -def equivProd {R : Type*} (c₁ c₂ : R) : ℍ[R,c₁,c₂] ≃ R × R × R × R where +def equivProd {R : Type*} (c₁ c₂ c₃: R) : ℍ[R,c₁,c₂,c₃] ≃ R × R × R × R where toFun a := ⟨a.1, a.2, a.3, a.4⟩ invFun a := ⟨a.1, a.2.1, a.2.2.1, a.2.2.2⟩ left_inv _ := rfl @@ -79,30 +85,34 @@ def equivProd {R : Type*} (c₁ c₂ : R) : ℍ[R,c₁,c₂] ≃ R × R × R × /-- The equivalence between a quaternion algebra over `R` and `Fin 4 → R`. -/ @[simps symm_apply] -def equivTuple {R : Type*} (c₁ c₂ : R) : ℍ[R,c₁,c₂] ≃ (Fin 4 → R) where +def equivTuple {R : Type*} (c₁ c₂ c₃: R) : ℍ[R,c₁,c₂,c₃] ≃ (Fin 4 → R) where toFun a := ![a.1, a.2, a.3, a.4] invFun a := ⟨a 0, a 1, a 2, a 3⟩ left_inv _ := rfl right_inv f := by ext ⟨_, _ | _ | _ | _ | _ | ⟨⟩⟩ <;> rfl @[simp] -theorem equivTuple_apply {R : Type*} (c₁ c₂ : R) (x : ℍ[R,c₁,c₂]) : - equivTuple c₁ c₂ x = ![x.re, x.imI, x.imJ, x.imK] := +theorem equivTuple_apply {R : Type*} (c₁ c₂ c₃: R) (x : ℍ[R,c₁,c₂,c₃]) : + equivTuple c₁ c₂ c₃ x = ![x.re, x.imI, x.imJ, x.imK] := rfl @[simp] -theorem mk.eta {R : Type*} {c₁ c₂} (a : ℍ[R,c₁,c₂]) : mk a.1 a.2 a.3 a.4 = a := rfl +theorem mk.eta {R : Type*} {c₁ c₂ c₃} (a : ℍ[R,c₁,c₂,c₃]) : mk a.1 a.2 a.3 a.4 = a := rfl -variable {S T R : Type*} {c₁ c₂ : R} (r x y : R) (a b : ℍ[R,c₁,c₂]) +variable {S T R : Type*} {c₁ c₂ c₃ : R} (r x y : R) (a b : ℍ[R,c₁,c₂,c₃]) -instance [Subsingleton R] : Subsingleton ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).subsingleton -instance [Nontrivial R] : Nontrivial ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).surjective.nontrivial +instance [Subsingleton R] : Subsingleton ℍ[R, c₁, c₂, c₃] := (equivTuple c₁ c₂ c₃).subsingleton +instance [Nontrivial R] : Nontrivial ℍ[R, c₁, c₂, c₃] := (equivTuple c₁ c₂ c₃).surjective.nontrivial section Zero variable [Zero R] -/-- The imaginary part of a quaternion. -/ -def im (x : ℍ[R,c₁,c₂]) : ℍ[R,c₁,c₂] := +/-- The imaginary part of a quaternion. + +Note that unless `c₂ = 0`, this definition is not particularly well-behaved; +for instance, `QuaternionAlgebra.star_im` only says that the star of an imaginary quaternion +is imaginary under this condition. -/ +def im (x : ℍ[R,c₁,c₂,c₃]) : ℍ[R,c₁,c₂,c₃] := ⟨0, x.imI, x.imJ, x.imK⟩ @[simp] @@ -125,66 +135,66 @@ theorem im_imK : a.im.imK = a.imK := theorem im_idem : a.im.im = a.im := rfl -/-- Coercion `R → ℍ[R,c₁,c₂]`. -/ -@[coe] def coe (x : R) : ℍ[R,c₁,c₂] := ⟨x, 0, 0, 0⟩ +/-- Coercion `R → ℍ[R,c₁,c₂,c₃]`. -/ +@[coe] def coe (x : R) : ℍ[R,c₁,c₂,c₃] := ⟨x, 0, 0, 0⟩ -instance : CoeTC R ℍ[R,c₁,c₂] := ⟨coe⟩ +instance : CoeTC R ℍ[R,c₁,c₂,c₃] := ⟨coe⟩ @[simp, norm_cast] -theorem coe_re : (x : ℍ[R,c₁,c₂]).re = x := rfl +theorem coe_re : (x : ℍ[R,c₁,c₂,c₃]).re = x := rfl @[simp, norm_cast] -theorem coe_imI : (x : ℍ[R,c₁,c₂]).imI = 0 := rfl +theorem coe_imI : (x : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl @[simp, norm_cast] -theorem coe_imJ : (x : ℍ[R,c₁,c₂]).imJ = 0 := rfl +theorem coe_imJ : (x : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl @[simp, norm_cast] -theorem coe_imK : (x : ℍ[R,c₁,c₂]).imK = 0 := rfl +theorem coe_imK : (x : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl -theorem coe_injective : Function.Injective (coe : R → ℍ[R,c₁,c₂]) := fun _ _ h => congr_arg re h +theorem coe_injective : Function.Injective (coe : R → ℍ[R,c₁,c₂,c₃]) := fun _ _ h => congr_arg re h @[simp] -theorem coe_inj {x y : R} : (x : ℍ[R,c₁,c₂]) = y ↔ x = y := +theorem coe_inj {x y : R} : (x : ℍ[R,c₁,c₂,c₃]) = y ↔ x = y := coe_injective.eq_iff -- Porting note: removed `simps`, added simp lemmas manually. -- Should adjust `simps` to name properly, i.e. as `zero_re` rather than `instZero_zero_re`. -instance : Zero ℍ[R,c₁,c₂] := ⟨⟨0, 0, 0, 0⟩⟩ +instance : Zero ℍ[R,c₁,c₂,c₃] := ⟨⟨0, 0, 0, 0⟩⟩ -@[simp] theorem zero_re : (0 : ℍ[R,c₁,c₂]).re = 0 := rfl +@[simp] theorem zero_re : (0 : ℍ[R,c₁,c₂,c₃]).re = 0 := rfl -@[simp] theorem zero_imI : (0 : ℍ[R,c₁,c₂]).imI = 0 := rfl +@[simp] theorem zero_imI : (0 : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl -@[simp] theorem zero_imJ : (0 : ℍ[R,c₁,c₂]).imJ = 0 := rfl +@[simp] theorem zero_imJ : (0 : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl -@[simp] theorem zero_imK : (0 : ℍ[R,c₁,c₂]).imK = 0 := rfl +@[simp] theorem zero_imK : (0 : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl -@[simp] theorem zero_im : (0 : ℍ[R,c₁,c₂]).im = 0 := rfl +@[simp] theorem zero_im : (0 : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl @[simp, norm_cast] -theorem coe_zero : ((0 : R) : ℍ[R,c₁,c₂]) = 0 := rfl +theorem coe_zero : ((0 : R) : ℍ[R,c₁,c₂,c₃]) = 0 := rfl -instance : Inhabited ℍ[R,c₁,c₂] := ⟨0⟩ +instance : Inhabited ℍ[R,c₁,c₂,c₃] := ⟨0⟩ section One variable [One R] -- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly -instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩ +instance : One ℍ[R,c₁,c₂,c₃] := ⟨⟨1, 0, 0, 0⟩⟩ -@[simp] theorem one_re : (1 : ℍ[R,c₁,c₂]).re = 1 := rfl +@[simp] theorem one_re : (1 : ℍ[R,c₁,c₂,c₃]).re = 1 := rfl -@[simp] theorem one_imI : (1 : ℍ[R,c₁,c₂]).imI = 0 := rfl +@[simp] theorem one_imI : (1 : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl -@[simp] theorem one_imJ : (1 : ℍ[R,c₁,c₂]).imJ = 0 := rfl +@[simp] theorem one_imJ : (1 : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl -@[simp] theorem one_imK : (1 : ℍ[R,c₁,c₂]).imK = 0 := rfl +@[simp] theorem one_imK : (1 : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl -@[simp] theorem one_im : (1 : ℍ[R,c₁,c₂]).im = 0 := rfl +@[simp] theorem one_im : (1 : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl @[simp, norm_cast] -theorem coe_one : ((1 : R) : ℍ[R,c₁,c₂]) = 1 := rfl +theorem coe_one : ((1 : R) : ℍ[R,c₁,c₂,c₃]) = 1 := rfl end One end Zero @@ -192,7 +202,7 @@ section Add variable [Add R] -- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly -instance : Add ℍ[R,c₁,c₂] := +instance : Add ℍ[R,c₁,c₂,c₃] := ⟨fun a b => ⟨a.1 + b.1, a.2 + b.2, a.3 + b.3, a.4 + b.4⟩⟩ @[simp] theorem add_re : (a + b).re = a.re + b.re := rfl @@ -205,7 +215,8 @@ instance : Add ℍ[R,c₁,c₂] := @[simp] theorem mk_add_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : - (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) + mk b₁ b₂ b₃ b₄ = mk (a₁ + b₁) (a₂ + b₂) (a₃ + b₃) (a₄ + b₄) := + (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂,c₃]) + mk b₁ b₂ b₃ b₄ = + mk (a₁ + b₁) (a₂ + b₂) (a₃ + b₃) (a₄ + b₄) := rfl end Add @@ -217,7 +228,7 @@ variable [AddZeroClass R] QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl @[simp, norm_cast] -theorem coe_add : ((x + y : R) : ℍ[R,c₁,c₂]) = x + y := by ext <;> simp +theorem coe_add : ((x + y : R) : ℍ[R,c₁,c₂,c₃]) = x + y := by ext <;> simp end AddZeroClass @@ -225,7 +236,7 @@ section Neg variable [Neg R] -- Porting note: removed `simps`, added simp lemmas manually. Should adjust `simps` to name properly -instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ +instance : Neg ℍ[R,c₁,c₂,c₃] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ @[simp] theorem neg_re : (-a).re = -a.re := rfl @@ -236,7 +247,7 @@ instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ @[simp] theorem neg_imK : (-a).imK = -a.imK := rfl @[simp] -theorem neg_mk (a₁ a₂ a₃ a₄ : R) : -(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨-a₁, -a₂, -a₃, -a₄⟩ := +theorem neg_mk (a₁ a₂ a₃ a₄ : R) : -(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂,c₃]) = ⟨-a₁, -a₂, -a₃, -a₄⟩ := rfl end Neg @@ -248,9 +259,9 @@ variable [AddGroup R] QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl @[simp, norm_cast] -theorem coe_neg : ((-x : R) : ℍ[R,c₁,c₂]) = -x := by ext <;> simp +theorem coe_neg : ((-x : R) : ℍ[R,c₁,c₂,c₃]) = -x := by ext <;> simp -instance : Sub ℍ[R,c₁,c₂] := +instance : Sub ℍ[R,c₁,c₂,c₃] := ⟨fun a b => ⟨a.1 - b.1, a.2 - b.2, a.3 - b.3, a.4 - b.4⟩⟩ @[simp] theorem sub_re : (a - b).re = a.re - b.re := rfl @@ -266,11 +277,12 @@ instance : Sub ℍ[R,c₁,c₂] := @[simp] theorem mk_sub_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : - (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) - mk b₁ b₂ b₃ b₄ = mk (a₁ - b₁) (a₂ - b₂) (a₃ - b₃) (a₄ - b₄) := + (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂,c₃]) - mk b₁ b₂ b₃ b₄ = + mk (a₁ - b₁) (a₂ - b₂) (a₃ - b₃) (a₄ - b₄) := rfl @[simp, norm_cast] -theorem coe_im : (x : ℍ[R,c₁,c₂]).im = 0 := +theorem coe_im : (x : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl @[simp] @@ -293,37 +305,43 @@ variable [Ring R] /-- Multiplication is given by * `1 * x = x * 1 = x`; -* `i * i = c₁`; -* `j * j = c₂`; -* `i * j = k`, `j * i = -k`; -* `k * k = -c₁ * c₂`; -* `i * k = c₁ * j`, `k * i = -c₁ * j`; -* `j * k = -c₂ * i`, `k * j = c₂ * i`. -/ -instance : Mul ℍ[R,c₁,c₂] := +* `i * i = c₁ + c₂ * i`; +* `j * j = c₃`; +* `i * j = k`, `j * i = c₂ * j - k`; +* `k * k = - c₁ * c₃`; +* `i * k = c₁ * j + c₂ * k`, `k * i = -c₁ * j`; +* `j * k = c₂ * c₃ - c₃ * i`, `k * j = c₃ * i`. -/ +instance : Mul ℍ[R,c₁,c₂,c₃] := ⟨fun a b => - ⟨a.1 * b.1 + c₁ * a.2 * b.2 + c₂ * a.3 * b.3 - c₁ * c₂ * a.4 * b.4, - a.1 * b.2 + a.2 * b.1 - c₂ * a.3 * b.4 + c₂ * a.4 * b.3, - a.1 * b.3 + c₁ * a.2 * b.4 + a.3 * b.1 - c₁ * a.4 * b.2, - a.1 * b.4 + a.2 * b.3 - a.3 * b.2 + a.4 * b.1⟩⟩ + ⟨a.1 * b.1 + c₁ * a.2 * b.2 + c₃ * a.3 * b.3 + c₂ * c₃ * a.3 * b.4 - c₁ * c₃ * a.4 * b.4, + a.1 * b.2 + a.2 * b.1 + c₂ * a.2 * b.2 - c₃ * a.3 * b.4 + c₃ * a.4 * b.3, + a.1 * b.3 + c₁ * a.2 * b.4 + a.3 * b.1 + c₂ * a.3 * b.2 - c₁ * a.4 * b.2, + a.1 * b.4 + a.2 * b.3 + c₂ * a.2 * b.4 - a.3 * b.2 + a.4 * b.1⟩⟩ @[simp] -theorem mul_re : (a * b).re = a.1 * b.1 + c₁ * a.2 * b.2 + c₂ * a.3 * b.3 - c₁ * c₂ * a.4 * b.4 := - rfl +theorem mul_re : (a * b).re = a.1 * b.1 + c₁ * a.2 * b.2 + c₃ * a.3 * b.3 + + c₂ * c₃ * a.3 * b.4 - c₁ * c₃ * a.4 * b.4 := rfl @[simp] -theorem mul_imI : (a * b).imI = a.1 * b.2 + a.2 * b.1 - c₂ * a.3 * b.4 + c₂ * a.4 * b.3 := rfl +theorem mul_imI : (a * b).imI = a.1 * b.2 + a.2 * b.1 + + c₂ * a.2 * b.2 - c₃ * a.3 * b.4 + c₃ * a.4 * b.3 := rfl @[simp] -theorem mul_imJ : (a * b).imJ = a.1 * b.3 + c₁ * a.2 * b.4 + a.3 * b.1 - c₁ * a.4 * b.2 := rfl +theorem mul_imJ : (a * b).imJ = a.1 * b.3 + c₁ * a.2 * b.4 + a.3 * b.1 + + c₂ * a.3 * b.2 - c₁ * a.4 * b.2 := rfl -@[simp] theorem mul_imK : (a * b).imK = a.1 * b.4 + a.2 * b.3 - a.3 * b.2 + a.4 * b.1 := rfl +@[simp] +theorem mul_imK : (a * b).imK = a.1 * b.4 + a.2 * b.3 + + c₂ * a.2 * b.4 - a.3 * b.2 + a.4 * b.1 := rfl @[simp] theorem mk_mul_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : - (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) * mk b₁ b₂ b₃ b₄ = - ⟨a₁ * b₁ + c₁ * a₂ * b₂ + c₂ * a₃ * b₃ - c₁ * c₂ * a₄ * b₄, - a₁ * b₂ + a₂ * b₁ - c₂ * a₃ * b₄ + c₂ * a₄ * b₃, - a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ - c₁ * a₄ * b₂, a₁ * b₄ + a₂ * b₃ - a₃ * b₂ + a₄ * b₁⟩ := + (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂,c₃]) * mk b₁ b₂ b₃ b₄ = + mk + (a₁ * b₁ + c₁ * a₂ * b₂ + c₃ * a₃ * b₃ + c₂ * c₃ * a₃ * b₄ - c₁ * c₃ * a₄ * b₄) + (a₁ * b₂ + a₂ * b₁ + c₂ * a₂ * b₂ - c₃ * a₃ * b₄ + c₃ * a₄ * b₃) + (a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ + c₂ * a₃ * b₂ - c₁ * a₄ * b₂) + (a₁ * b₄ + a₂ * b₃ + c₂ * a₂ * b₄ - a₃ * b₂ + a₄ * b₁) := rfl end Ring @@ -331,12 +349,12 @@ section SMul variable [SMul S R] [SMul T R] (s : S) -instance : SMul S ℍ[R,c₁,c₂] where smul s a := ⟨s • a.1, s • a.2, s • a.3, s • a.4⟩ +instance : SMul S ℍ[R,c₁,c₂,c₃] where smul s a := ⟨s • a.1, s • a.2, s • a.3, s • a.4⟩ -instance [SMul S T] [IsScalarTower S T R] : IsScalarTower S T ℍ[R,c₁,c₂] where +instance [SMul S T] [IsScalarTower S T R] : IsScalarTower S T ℍ[R,c₁,c₂,c₃] where smul_assoc s t x := by ext <;> exact smul_assoc _ _ _ -instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂] where +instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂,c₃] where smul_comm s t x := by ext <;> exact smul_comm _ _ _ @[simp] theorem smul_re : (s • a).re = s • a.re := rfl @@ -352,112 +370,127 @@ instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂] where @[simp] theorem smul_mk (re im_i im_j im_k : R) : - s • (⟨re, im_i, im_j, im_k⟩ : ℍ[R,c₁,c₂]) = ⟨s • re, s • im_i, s • im_j, s • im_k⟩ := + s • (⟨re, im_i, im_j, im_k⟩ : ℍ[R,c₁,c₂,c₃]) = ⟨s • re, s • im_i, s • im_j, s • im_k⟩ := rfl end SMul @[simp, norm_cast] theorem coe_smul [Zero R] [SMulZeroClass S R] (s : S) (r : R) : - (↑(s • r) : ℍ[R,c₁,c₂]) = s • (r : ℍ[R,c₁,c₂]) := - QuaternionAlgebra.ext rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm + (↑(s • r) : ℍ[R,c₁,c₂,c₃]) = s • (r : ℍ[R,c₁,c₂,c₃]) := + QuaternionAlgebra.ext rfl (smul_zero _).symm (smul_zero _).symm (smul_zero _).symm -instance [AddCommGroup R] : AddCommGroup ℍ[R,c₁,c₂] := - (equivProd c₁ c₂).injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) +instance [AddCommGroup R] : AddCommGroup ℍ[R,c₁,c₂,c₃] := + (equivProd c₁ c₂ c₃).injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) section AddCommGroupWithOne variable [AddCommGroupWithOne R] -instance : AddCommGroupWithOne ℍ[R,c₁,c₂] where - natCast n := ((n : R) : ℍ[R,c₁,c₂]) +instance : AddCommGroupWithOne ℍ[R,c₁,c₂,c₃] where + natCast n := ((n : R) : ℍ[R,c₁,c₂,c₃]) natCast_zero := by simp natCast_succ := by simp - intCast n := ((n : R) : ℍ[R,c₁,c₂]) + intCast n := ((n : R) : ℍ[R,c₁,c₂,c₃]) intCast_ofNat _ := congr_arg coe (Int.cast_natCast _) intCast_negSucc n := by change coe _ = -coe _ rw [Int.cast_negSucc, coe_neg] @[simp, norm_cast] -theorem natCast_re (n : ℕ) : (n : ℍ[R,c₁,c₂]).re = n := +theorem natCast_re (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).re = n := rfl @[deprecated (since := "2024-04-17")] alias nat_cast_re := natCast_re @[simp, norm_cast] -theorem natCast_imI (n : ℕ) : (n : ℍ[R,c₁,c₂]).imI = 0 := +theorem natCast_imI (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl @[deprecated (since := "2024-04-17")] alias nat_cast_imI := natCast_imI @[simp, norm_cast] -theorem natCast_imJ (n : ℕ) : (n : ℍ[R,c₁,c₂]).imJ = 0 := +theorem natCast_imJ (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl @[deprecated (since := "2024-04-17")] alias nat_cast_imJ := natCast_imJ @[simp, norm_cast] -theorem natCast_imK (n : ℕ) : (n : ℍ[R,c₁,c₂]).imK = 0 := +theorem natCast_imK (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl @[deprecated (since := "2024-04-17")] alias nat_cast_imK := natCast_imK @[simp, norm_cast] -theorem natCast_im (n : ℕ) : (n : ℍ[R,c₁,c₂]).im = 0 := +theorem natCast_im (n : ℕ) : (n : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl @[deprecated (since := "2024-04-17")] alias nat_cast_im := natCast_im @[norm_cast] -theorem coe_natCast (n : ℕ) : ↑(n : R) = (n : ℍ[R,c₁,c₂]) := +theorem coe_natCast (n : ℕ) : ↑(n : R) = (n : ℍ[R,c₁,c₂,c₃]) := rfl @[deprecated (since := "2024-04-17")] alias coe_nat_cast := coe_natCast @[simp, norm_cast] -theorem intCast_re (z : ℤ) : (z : ℍ[R,c₁,c₂]).re = z := +theorem intCast_re (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).re = z := rfl +@[simp] +theorem ofNat_re (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]).re = ofNat(n) := rfl + +@[simp] +theorem ofNat_imI (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl + +@[simp] +theorem ofNat_imJ (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl + +@[simp] +theorem ofNat_imK (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl + +@[simp] +theorem ofNat_im (n : ℕ) [n.AtLeastTwo] : (ofNat(n) : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl + @[deprecated (since := "2024-04-17")] alias int_cast_re := intCast_re @[simp, norm_cast] -theorem intCast_imI (z : ℤ) : (z : ℍ[R,c₁,c₂]).imI = 0 := +theorem intCast_imI (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).imI = 0 := rfl @[deprecated (since := "2024-04-17")] alias int_cast_imI := intCast_imI @[simp, norm_cast] -theorem intCast_imJ (z : ℤ) : (z : ℍ[R,c₁,c₂]).imJ = 0 := +theorem intCast_imJ (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).imJ = 0 := rfl @[deprecated (since := "2024-04-17")] alias int_cast_imJ := intCast_imJ @[simp, norm_cast] -theorem intCast_imK (z : ℤ) : (z : ℍ[R,c₁,c₂]).imK = 0 := +theorem intCast_imK (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).imK = 0 := rfl @[deprecated (since := "2024-04-17")] alias int_cast_imK := intCast_imK @[simp, norm_cast] -theorem intCast_im (z : ℤ) : (z : ℍ[R,c₁,c₂]).im = 0 := +theorem intCast_im (z : ℤ) : (z : ℍ[R,c₁,c₂,c₃]).im = 0 := rfl @[deprecated (since := "2024-04-17")] alias int_cast_im := intCast_im @[norm_cast] -theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R,c₁,c₂]) := +theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R,c₁,c₂,c₃]) := rfl @[deprecated (since := "2024-04-17")] @@ -468,8 +501,8 @@ end AddCommGroupWithOne -- For the remainder of the file we assume `CommRing R`. variable [CommRing R] -instance instRing : Ring ℍ[R,c₁,c₂] where - __ := inferInstanceAs (AddCommGroupWithOne ℍ[R,c₁,c₂]) +instance instRing : Ring ℍ[R,c₁,c₂,c₃] where + __ := inferInstanceAs (AddCommGroupWithOne ℍ[R,c₁,c₂,c₃]) left_distrib _ _ _ := by ext <;> simp <;> ring right_distrib _ _ _ := by ext <;> simp <;> ring zero_mul _ := by ext <;> simp @@ -479,11 +512,16 @@ instance instRing : Ring ℍ[R,c₁,c₂] where mul_one _ := by ext <;> simp @[norm_cast, simp] -theorem coe_mul : ((x * y : R) : ℍ[R,c₁,c₂]) = x * y := by ext <;> simp +theorem coe_mul : ((x * y : R) : ℍ[R,c₁,c₂,c₃]) = x * y := by ext <;> simp + +@[norm_cast, simp] +lemma coe_ofNat {n : ℕ} [n.AtLeastTwo]: + ((ofNat(n) : R) : ℍ[R,c₁,c₂,c₃]) = (ofNat(n) : ℍ[R,c₁,c₂,c₃]) := by + rfl -- TODO: add weaker `MulAction`, `DistribMulAction`, and `Module` instances (and repeat them -- for `ℍ[R]`) -instance [CommSemiring S] [Algebra S R] : Algebra S ℍ[R,c₁,c₂] where +instance [CommSemiring S] [Algebra S R] : Algebra S ℍ[R,c₁,c₂,c₃] where smul := (· • ·) algebraMap := { toFun s := coe (algebraMap S R s) @@ -494,13 +532,13 @@ instance [CommSemiring S] [Algebra S R] : Algebra S ℍ[R,c₁,c₂] where smul_def' s x := by ext <;> simp [Algebra.smul_def] commutes' s x := by ext <;> simp [Algebra.commutes] -theorem algebraMap_eq (r : R) : algebraMap R ℍ[R,c₁,c₂] r = ⟨r, 0, 0, 0⟩ := +theorem algebraMap_eq (r : R) : algebraMap R ℍ[R,c₁,c₂,c₃] r = ⟨r, 0, 0, 0⟩ := rfl -theorem algebraMap_injective : (algebraMap R ℍ[R,c₁,c₂] : _ → _).Injective := +theorem algebraMap_injective : (algebraMap R ℍ[R,c₁,c₂,c₃] : _ → _).Injective := fun _ _ ↦ by simp [algebraMap_eq] -instance [NoZeroDivisors R] : NoZeroSMulDivisors R ℍ[R,c₁,c₂] := ⟨by +instance [NoZeroDivisors R] : NoZeroSMulDivisors R ℍ[R,c₁,c₂,c₃] := ⟨by rintro t ⟨a, b, c, d⟩ h rw [or_iff_not_imp_left] intro ht @@ -508,97 +546,94 @@ instance [NoZeroDivisors R] : NoZeroSMulDivisors R ℍ[R,c₁,c₂] := ⟨by section -variable (c₁ c₂) +variable (c₁ c₂ c₃) /-- `QuaternionAlgebra.re` as a `LinearMap`-/ @[simps] -def reₗ : ℍ[R,c₁,c₂] →ₗ[R] R where +def reₗ : ℍ[R,c₁,c₂,c₃] →ₗ[R] R where toFun := re map_add' _ _ := rfl map_smul' _ _ := rfl /-- `QuaternionAlgebra.imI` as a `LinearMap`-/ @[simps] -def imIₗ : ℍ[R,c₁,c₂] →ₗ[R] R where +def imIₗ : ℍ[R,c₁,c₂,c₃] →ₗ[R] R where toFun := imI map_add' _ _ := rfl map_smul' _ _ := rfl /-- `QuaternionAlgebra.imJ` as a `LinearMap`-/ @[simps] -def imJₗ : ℍ[R,c₁,c₂] →ₗ[R] R where +def imJₗ : ℍ[R,c₁,c₂,c₃] →ₗ[R] R where toFun := imJ map_add' _ _ := rfl map_smul' _ _ := rfl /-- `QuaternionAlgebra.imK` as a `LinearMap`-/ @[simps] -def imKₗ : ℍ[R,c₁,c₂] →ₗ[R] R where +def imKₗ : ℍ[R,c₁,c₂,c₃] →ₗ[R] R where toFun := imK map_add' _ _ := rfl map_smul' _ _ := rfl /-- `QuaternionAlgebra.equivTuple` as a linear equivalence. -/ -def linearEquivTuple : ℍ[R,c₁,c₂] ≃ₗ[R] Fin 4 → R := +def linearEquivTuple : ℍ[R,c₁,c₂,c₃] ≃ₗ[R] Fin 4 → R := LinearEquiv.symm -- proofs are not `rfl` in the forward direction - { (equivTuple c₁ c₂).symm with - toFun := (equivTuple c₁ c₂).symm - invFun := equivTuple c₁ c₂ + { (equivTuple c₁ c₂ c₃).symm with + toFun := (equivTuple c₁ c₂ c₃).symm + invFun := equivTuple c₁ c₂ c₃ map_add' := fun _ _ => rfl map_smul' := fun _ _ => rfl } @[simp] -theorem coe_linearEquivTuple : ⇑(linearEquivTuple c₁ c₂) = equivTuple c₁ c₂ := - rfl +theorem coe_linearEquivTuple : + ⇑(linearEquivTuple c₁ c₂ c₃) = equivTuple c₁ c₂ c₃ := rfl @[simp] -theorem coe_linearEquivTuple_symm : ⇑(linearEquivTuple c₁ c₂).symm = (equivTuple c₁ c₂).symm := - rfl +theorem coe_linearEquivTuple_symm : + ⇑(linearEquivTuple c₁ c₂ c₃).symm = (equivTuple c₁ c₂ c₃).symm := rfl -/-- `ℍ[R, c₁, c₂]` has a basis over `R` given by `1`, `i`, `j`, and `k`. -/ -noncomputable def basisOneIJK : Basis (Fin 4) R ℍ[R,c₁,c₂] := - .ofEquivFun <| linearEquivTuple c₁ c₂ +/-- `ℍ[R, c₁, c₂, c₃]` has a basis over `R` given by `1`, `i`, `j`, and `k`. -/ +noncomputable def basisOneIJK : Basis (Fin 4) R ℍ[R,c₁,c₂,c₃] := + .ofEquivFun <| linearEquivTuple c₁ c₂ c₃ @[simp] -theorem coe_basisOneIJK_repr (q : ℍ[R,c₁,c₂]) : - ⇑((basisOneIJK c₁ c₂).repr q) = ![q.re, q.imI, q.imJ, q.imK] := +theorem coe_basisOneIJK_repr (q : ℍ[R,c₁,c₂,c₃]) : + ((basisOneIJK c₁ c₂ c₃).repr q) = ![q.re, q.imI, q.imJ, q.imK] := rfl -instance : Module.Finite R ℍ[R,c₁,c₂] := .of_basis (basisOneIJK c₁ c₂) +instance : Module.Finite R ℍ[R,c₁,c₂,c₃] := .of_basis (basisOneIJK c₁ c₂ c₃) -instance : Module.Free R ℍ[R,c₁,c₂] := .of_basis (basisOneIJK c₁ c₂) +instance : Module.Free R ℍ[R,c₁,c₂,c₃] := .of_basis (basisOneIJK c₁ c₂ c₃) -theorem rank_eq_four [StrongRankCondition R] : Module.rank R ℍ[R,c₁,c₂] = 4 := by - rw [rank_eq_card_basis (basisOneIJK c₁ c₂), Fintype.card_fin] +theorem rank_eq_four [StrongRankCondition R] : Module.rank R ℍ[R,c₁,c₂,c₃] = 4 := by + rw [rank_eq_card_basis (basisOneIJK c₁ c₂ c₃), Fintype.card_fin] norm_num -theorem finrank_eq_four [StrongRankCondition R] : Module.finrank R ℍ[R,c₁,c₂] = 4 := by +theorem finrank_eq_four [StrongRankCondition R] : Module.finrank R ℍ[R,c₁,c₂,c₃] = 4 := by rw [Module.finrank, rank_eq_four, Cardinal.toNat_ofNat] -/-- There is a natural equivalence when swapping the coefficients of a quaternion algebra. -/ +/-- There is a natural equivalence when swapping the first and third coefficients of a + quaternion algebra if `c₂` is 0. -/ @[simps] -def swapEquiv : ℍ[R,c₁,c₂] ≃ₐ[R] ℍ[R, c₂, c₁] where +def swapEquiv : ℍ[R,c₁,0,c₃] ≃ₐ[R] ℍ[R,c₃,0,c₁] where toFun t := ⟨t.1, t.3, t.2, -t.4⟩ invFun t := ⟨t.1, t.3, t.2, -t.4⟩ left_inv _ := by simp right_inv _ := by simp - map_mul' _ _ := by - ext - <;> simp only [mul_re, mul_imJ, mul_imI, add_left_inj, mul_imK, neg_mul, neg_add_rev, - neg_sub, mk_mul_mk, mul_neg, neg_neg, sub_neg_eq_add] - <;> ring + map_mul' _ _ := by ext <;> simp <;> ring map_add' _ _ := by ext <;> simp [add_comm] commutes' _ := by simp [algebraMap_eq] end @[norm_cast, simp] -theorem coe_sub : ((x - y : R) : ℍ[R,c₁,c₂]) = x - y := - (algebraMap R ℍ[R,c₁,c₂]).map_sub x y +theorem coe_sub : ((x - y : R) : ℍ[R,c₁,c₂,c₃]) = x - y := + (algebraMap R ℍ[R,c₁,c₂,c₃]).map_sub x y @[norm_cast, simp] -theorem coe_pow (n : ℕ) : (↑(x ^ n) : ℍ[R,c₁,c₂]) = (x : ℍ[R,c₁,c₂]) ^ n := - (algebraMap R ℍ[R,c₁,c₂]).map_pow x n +theorem coe_pow (n : ℕ) : (↑(x ^ n) : ℍ[R,c₁,c₂,c₃]) = (x : ℍ[R,c₁,c₂,c₃]) ^ n := + (algebraMap R ℍ[R,c₁,c₂,c₃]).map_pow x n theorem coe_commutes : ↑r * a = a * r := Algebra.commutes r a @@ -612,15 +647,16 @@ theorem coe_mul_eq_smul : ↑r * a = r • a := theorem mul_coe_eq_smul : a * r = r • a := by rw [← coe_commutes, coe_mul_eq_smul] @[norm_cast, simp] -theorem coe_algebraMap : ⇑(algebraMap R ℍ[R,c₁,c₂]) = coe := +theorem coe_algebraMap : ⇑(algebraMap R ℍ[R,c₁,c₂,c₃]) = coe := rfl -theorem smul_coe : x • (y : ℍ[R,c₁,c₂]) = ↑(x * y) := by rw [coe_mul, coe_mul_eq_smul] +theorem smul_coe : x • (y : ℍ[R,c₁,c₂,c₃]) = ↑(x * y) := by rw [coe_mul, coe_mul_eq_smul] /-- Quaternion conjugate. -/ -instance instStarQuaternionAlgebra : Star ℍ[R,c₁,c₂] where star a := ⟨a.1, -a.2, -a.3, -a.4⟩ +instance instStarQuaternionAlgebra : Star ℍ[R,c₁,c₂,c₃] where star a := + ⟨a.1 + c₂ * a.2, -a.2, -a.3, -a.4⟩ -@[simp] theorem re_star : (star a).re = a.re := rfl +@[simp] theorem re_star : (star a).re = a.re + c₂ * a.imI := rfl @[simp] theorem imI_star : (star a).imI = -a.imI := @@ -639,44 +675,54 @@ theorem im_star : (star a).im = -a.im := QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl @[simp] -theorem star_mk (a₁ a₂ a₃ a₄ : R) : star (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨a₁, -a₂, -a₃, -a₄⟩ := - rfl +theorem star_mk (a₁ a₂ a₃ a₄ : R) : star (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂,c₃]) = + ⟨a₁ + c₂ * a₂, -a₂, -a₃, -a₄⟩ := rfl -instance instStarRing : StarRing ℍ[R,c₁,c₂] where +instance instStarRing : StarRing ℍ[R,c₁,c₂,c₃] where star_involutive x := by simp [Star.star] - star_add a b := by ext <;> simp [add_comm] + star_add a b := by ext <;> simp [add_comm] ; ring star_mul a b := by ext <;> simp <;> ring -theorem self_add_star' : a + star a = ↑(2 * a.re) := by ext <;> simp [two_mul] +theorem self_add_star' : a + star a = ↑(2 * a.re + c₂ * a.imI) := by ext <;> simp [two_mul]; ring -theorem self_add_star : a + star a = 2 * a.re := by simp only [self_add_star', two_mul, coe_add] +theorem self_add_star : a + star a = 2 * a.re + c₂ * a.imI := by simp [self_add_star'] -theorem star_add_self' : star a + a = ↑(2 * a.re) := by rw [add_comm, self_add_star'] +theorem star_add_self' : star a + a = ↑(2 * a.re + c₂ * a.imI) := by rw [add_comm, self_add_star'] -theorem star_add_self : star a + a = 2 * a.re := by rw [add_comm, self_add_star] +theorem star_add_self : star a + a = 2 * a.re + c₂ * a.imI := by rw [add_comm, self_add_star] -theorem star_eq_two_re_sub : star a = ↑(2 * a.re) - a := +theorem star_eq_two_re_sub : star a = ↑(2 * a.re + c₂ * a.imI) - a := eq_sub_iff_add_eq.2 a.star_add_self' +lemma comm (r : R) (x : ℍ[R, c₁, c₂, c₃]) : r * x = x * r := by + ext <;> simp [mul_comm] + instance : IsStarNormal a := ⟨by - rw [a.star_eq_two_re_sub] - exact (coe_commute (2 * a.re) a).sub_left (Commute.refl a)⟩ + rw [commute_iff_eq, a.star_eq_two_re_sub]; + ext <;> simp <;> ring⟩ @[simp, norm_cast] -theorem star_coe : star (x : ℍ[R,c₁,c₂]) = x := by ext <;> simp +theorem star_coe : star (x : ℍ[R,c₁,c₂,c₃]) = x := by ext <;> simp -@[simp] theorem star_im : star a.im = -a.im := im_star _ +@[simp] theorem star_im : star a.im = -a.im + c₂ * a.imI := by ext <;> simp @[simp] -theorem star_smul [Monoid S] [DistribMulAction S R] (s : S) (a : ℍ[R,c₁,c₂]) : +theorem star_smul [Monoid S] [DistribMulAction S R] [SMulCommClass S R R] + (s : S) (a : ℍ[R,c₁,c₂,c₃]) : + star (s • a) = s • star a := + QuaternionAlgebra.ext + (by simp [mul_smul_comm]) (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm + +/-- A version of `star_smul` for the special case when `c₂ = 0`, without `SMulCommClass S R R`. -/ +theorem star_smul' [Monoid S] [DistribMulAction S R] (s : S) (a : ℍ[R,c₁,0,c₃]) : star (s • a) = s • star a := - QuaternionAlgebra.ext rfl (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm + QuaternionAlgebra.ext (by simp) (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm -theorem eq_re_of_eq_coe {a : ℍ[R,c₁,c₂]} {x : R} (h : a = x) : a = a.re := by rw [h, coe_re] +theorem eq_re_of_eq_coe {a : ℍ[R,c₁,c₂,c₃]} {x : R} (h : a = x) : a = a.re := by rw [h, coe_re] -theorem eq_re_iff_mem_range_coe {a : ℍ[R,c₁,c₂]} : - a = a.re ↔ a ∈ Set.range (coe : R → ℍ[R,c₁,c₂]) := +theorem eq_re_iff_mem_range_coe {a : ℍ[R,c₁,c₂,c₃]} : + a = a.re ↔ a ∈ Set.range (coe : R → ℍ[R,c₁,c₂,c₃]) := ⟨fun h => ⟨a.re, h.symm⟩, fun ⟨_, h⟩ => eq_re_of_eq_coe h.symm⟩ section CharZero @@ -684,10 +730,10 @@ section CharZero variable [NoZeroDivisors R] [CharZero R] @[simp] -theorem star_eq_self {c₁ c₂ : R} {a : ℍ[R,c₁,c₂]} : star a = a ↔ a = a.re := by - simp [QuaternionAlgebra.ext_iff, neg_eq_iff_add_eq_zero, add_self_eq_zero] +theorem star_eq_self {c₁ c₂ : R} {a : ℍ[R,c₁,c₂,c₃]} : star a = a ↔ a = a.re := by + simp_all [QuaternionAlgebra.ext_iff, neg_eq_iff_add_eq_zero, add_self_eq_zero] -theorem star_eq_neg {c₁ c₂ : R} {a : ℍ[R,c₁,c₂]} : star a = -a ↔ a.re = 0 := by +theorem star_eq_neg {c₁ : R} {a : ℍ[R,c₁,0,c₃]} : star a = -a ↔ a.re = 0 := by simp [QuaternionAlgebra.ext_iff, eq_neg_iff_add_eq_zero] end CharZero @@ -702,7 +748,7 @@ theorem mul_star_eq_coe : a * star a = (a * star a).re := by open MulOpposite /-- Quaternion conjugate as an `AlgEquiv` to the opposite ring. -/ -def starAe : ℍ[R,c₁,c₂] ≃ₐ[R] ℍ[R,c₁,c₂]ᵐᵒᵖ := +def starAe : ℍ[R,c₁,c₂,c₃] ≃ₐ[R] ℍ[R,c₁,c₂,c₃]ᵐᵒᵖ := { starAddEquiv.trans opAddEquiv with toFun := op ∘ star invFun := star ∘ unop @@ -710,38 +756,40 @@ def starAe : ℍ[R,c₁,c₂] ≃ₐ[R] ℍ[R,c₁,c₂]ᵐᵒᵖ := commutes' := fun r => by simp } @[simp] -theorem coe_starAe : ⇑(starAe : ℍ[R,c₁,c₂] ≃ₐ[R] _) = op ∘ star := +theorem coe_starAe : ⇑(starAe : ℍ[R,c₁,c₂,c₃] ≃ₐ[R] _) = op ∘ star := rfl end QuaternionAlgebra -/-- Space of quaternions over a type. Implemented as a structure with four fields: -`re`, `im_i`, `im_j`, and `im_k`. -/ -def Quaternion (R : Type*) [One R] [Neg R] := - QuaternionAlgebra R (-1) (-1) +/-- Space of quaternions over a type, denoted as `ℍ[R]`. +Implemented as a structure with four fields: `re`, `im_i`, `im_j`, and `im_k`. -/ +def Quaternion (R : Type*) [Zero R] [One R] [Neg R] := + QuaternionAlgebra R (-1) (0) (-1) @[inherit_doc] scoped[Quaternion] notation "ℍ[" R "]" => Quaternion R +open Quaternion + /-- The equivalence between the quaternions over `R` and `R × R × R × R`. -/ @[simps!] -def Quaternion.equivProd (R : Type*) [One R] [Neg R] : ℍ[R] ≃ R × R × R × R := - QuaternionAlgebra.equivProd _ _ +def Quaternion.equivProd (R : Type*) [Zero R] [One R] [Neg R] : ℍ[R] ≃ R × R × R × R := + QuaternionAlgebra.equivProd _ _ _ /-- The equivalence between the quaternions over `R` and `Fin 4 → R`. -/ @[simps! symm_apply] -def Quaternion.equivTuple (R : Type*) [One R] [Neg R] : ℍ[R] ≃ (Fin 4 → R) := - QuaternionAlgebra.equivTuple _ _ +def Quaternion.equivTuple (R : Type*) [Zero R] [One R] [Neg R] : ℍ[R] ≃ (Fin 4 → R) := + QuaternionAlgebra.equivTuple _ _ _ @[simp] -theorem Quaternion.equivTuple_apply (R : Type*) [One R] [Neg R] (x : ℍ[R]) : +theorem Quaternion.equivTuple_apply (R : Type*) [Zero R] [One R] [Neg R] (x : ℍ[R]) : Quaternion.equivTuple R x = ![x.re, x.imI, x.imJ, x.imK] := rfl -instance {R : Type*} [One R] [Neg R] [Subsingleton R] : Subsingleton ℍ[R] := - inferInstanceAs (Subsingleton <| ℍ[R, -1, -1]) -instance {R : Type*} [One R] [Neg R] [Nontrivial R] : Nontrivial ℍ[R] := - inferInstanceAs (Nontrivial <| ℍ[R, -1, -1]) +instance {R : Type*} [Zero R] [One R] [Neg R] [Subsingleton R] : Subsingleton ℍ[R] := + inferInstanceAs (Subsingleton <| ℍ[R, -1, 0, -1]) +instance {R : Type*} [Zero R] [One R] [Neg R] [Nontrivial R] : Nontrivial ℍ[R] := + inferInstanceAs (Nontrivial <| ℍ[R, -1, 0, -1]) namespace Quaternion @@ -754,23 +802,23 @@ instance : CoeTC R ℍ[R] := ⟨coe⟩ instance instRing : Ring ℍ[R] := QuaternionAlgebra.instRing -instance : Inhabited ℍ[R] := inferInstanceAs <| Inhabited ℍ[R,-1,-1] +instance : Inhabited ℍ[R] := inferInstanceAs <| Inhabited ℍ[R,-1, 0, -1] -instance [SMul S R] : SMul S ℍ[R] := inferInstanceAs <| SMul S ℍ[R,-1,-1] +instance [SMul S R] : SMul S ℍ[R] := inferInstanceAs <| SMul S ℍ[R,-1, 0, -1] instance [SMul S T] [SMul S R] [SMul T R] [IsScalarTower S T R] : IsScalarTower S T ℍ[R] := - inferInstanceAs <| IsScalarTower S T ℍ[R,-1,-1] + inferInstanceAs <| IsScalarTower S T ℍ[R,-1,0,-1] instance [SMul S R] [SMul T R] [SMulCommClass S T R] : SMulCommClass S T ℍ[R] := - inferInstanceAs <| SMulCommClass S T ℍ[R,-1,-1] + inferInstanceAs <| SMulCommClass S T ℍ[R,-1,0,-1] protected instance algebra [CommSemiring S] [Algebra S R] : Algebra S ℍ[R] := - inferInstanceAs <| Algebra S ℍ[R,-1,-1] + inferInstanceAs <| Algebra S ℍ[R,-1,0,-1] -- Porting note: added shortcut instance : Star ℍ[R] := QuaternionAlgebra.instStarQuaternionAlgebra instance : StarRing ℍ[R] := QuaternionAlgebra.instStarRing -instance : IsStarNormal a := inferInstanceAs <| IsStarNormal (R := ℍ[R,-1,-1]) a +instance : IsStarNormal a := inferInstanceAs <| IsStarNormal (R := ℍ[R,-1,0,-1]) a @[ext] theorem ext : a.re = b.re → a.imI = b.imI → a.imJ = b.imJ → a.imK = b.imK → a = b := @@ -880,19 +928,19 @@ theorem coe_sub : ((x - y : R) : ℍ[R]) = x - y := @[simp] theorem mul_re : (a * b).re = a.re * b.re - a.imI * b.imI - a.imJ * b.imJ - a.imK * b.imK := - (QuaternionAlgebra.mul_re a b).trans <| by simp only [one_mul, neg_mul, sub_eq_add_neg, neg_neg] + (QuaternionAlgebra.mul_re a b).trans <| by simp [one_mul, neg_mul, sub_eq_add_neg, neg_neg] @[simp] theorem mul_imI : (a * b).imI = a.re * b.imI + a.imI * b.re + a.imJ * b.imK - a.imK * b.imJ := - (QuaternionAlgebra.mul_imI a b).trans <| by simp only [one_mul, neg_mul, sub_eq_add_neg, neg_neg] + (QuaternionAlgebra.mul_imI a b).trans <| by ring @[simp] theorem mul_imJ : (a * b).imJ = a.re * b.imJ - a.imI * b.imK + a.imJ * b.re + a.imK * b.imI := - (QuaternionAlgebra.mul_imJ a b).trans <| by simp only [one_mul, neg_mul, sub_eq_add_neg, neg_neg] + (QuaternionAlgebra.mul_imJ a b).trans <| by ring @[simp] theorem mul_imK : (a * b).imK = a.re * b.imK + a.imI * b.imJ - a.imJ * b.imI + a.imK * b.re := - (QuaternionAlgebra.mul_imK a b).trans <| by simp only [one_mul, neg_mul, sub_eq_add_neg, neg_neg] + (QuaternionAlgebra.mul_imK a b).trans <| by ring @[simp, norm_cast] theorem coe_mul : ((x * y : R) : ℍ[R]) = x * y := QuaternionAlgebra.coe_mul x y @@ -1020,16 +1068,17 @@ theorem algebraMap_injective : (algebraMap R ℍ[R] : _ → _).Injective := theorem smul_coe : x • (y : ℍ[R]) = ↑(x * y) := QuaternionAlgebra.smul_coe x y -instance : Module.Finite R ℍ[R] := inferInstanceAs <| Module.Finite R ℍ[R,-1,-1] -instance : Module.Free R ℍ[R] := inferInstanceAs <| Module.Free R ℍ[R,-1,-1] +instance : Module.Finite R ℍ[R] := inferInstanceAs <| Module.Finite R ℍ[R,-1,0,-1] +instance : Module.Free R ℍ[R] := inferInstanceAs <| Module.Free R ℍ[R,-1,0,-1] theorem rank_eq_four [StrongRankCondition R] : Module.rank R ℍ[R] = 4 := - QuaternionAlgebra.rank_eq_four _ _ + QuaternionAlgebra.rank_eq_four _ _ _ theorem finrank_eq_four [StrongRankCondition R] : Module.finrank R ℍ[R] = 4 := - QuaternionAlgebra.finrank_eq_four _ _ + QuaternionAlgebra.finrank_eq_four _ _ _ -@[simp] theorem star_re : (star a).re = a.re := rfl +@[simp] theorem star_re : (star a).re = a.re := by + rw [QuaternionAlgebra.re_star, zero_mul, add_zero] @[simp] theorem star_imI : (star a).imI = -a.imI := rfl @@ -1039,33 +1088,31 @@ theorem finrank_eq_four [StrongRankCondition R] : Module.finrank R ℍ[R] = 4 := @[simp] theorem star_im : (star a).im = -a.im := a.im_star -nonrec theorem self_add_star' : a + star a = ↑(2 * a.re) := - a.self_add_star' +nonrec theorem self_add_star' : a + star a = ↑(2 * a.re) := by + simp [a.self_add_star', Quaternion.coe] -nonrec theorem self_add_star : a + star a = 2 * a.re := - a.self_add_star +nonrec theorem self_add_star : a + star a = 2 * a.re := by + simp [a.self_add_star, Quaternion.coe] -nonrec theorem star_add_self' : star a + a = ↑(2 * a.re) := - a.star_add_self' +nonrec theorem star_add_self' : star a + a = ↑(2 * a.re) := by + simp [a.star_add_self', Quaternion.coe] -nonrec theorem star_add_self : star a + a = 2 * a.re := - a.star_add_self +nonrec theorem star_add_self : star a + a = 2 * a.re := by + simp [a.star_add_self, Quaternion.coe] -nonrec theorem star_eq_two_re_sub : star a = ↑(2 * a.re) - a := - a.star_eq_two_re_sub +nonrec theorem star_eq_two_re_sub : star a = ↑(2 * a.re) - a := by + simp [a.star_eq_two_re_sub, Quaternion.coe] @[simp, norm_cast] theorem star_coe : star (x : ℍ[R]) = x := QuaternionAlgebra.star_coe x @[simp] -theorem im_star : star a.im = -a.im := - QuaternionAlgebra.im_star _ +theorem im_star : star a.im = -a.im := by ext <;> simp @[simp] theorem star_smul [Monoid S] [DistribMulAction S R] (s : S) (a : ℍ[R]) : - star (s • a) = s • star a := - QuaternionAlgebra.star_smul _ _ + star (s • a) = s • star a := QuaternionAlgebra.star_smul' s a theorem eq_re_of_eq_coe {a : ℍ[R]} {x : R} (h : a = x) : a = a.re := QuaternionAlgebra.eq_re_of_eq_coe h @@ -1308,34 +1355,34 @@ open Quaternion section QuaternionAlgebra -variable {R : Type*} (c₁ c₂ : R) +variable {R : Type*} (c₁ c₂ c₃ : R) private theorem pow_four [Infinite R] : #R ^ 4 = #R := power_nat_eq (aleph0_le_mk R) <| by decide /-- The cardinality of a quaternion algebra, as a type. -/ -theorem mk_quaternionAlgebra : #(ℍ[R,c₁,c₂]) = #R ^ 4 := by - rw [mk_congr (QuaternionAlgebra.equivProd c₁ c₂)] +theorem mk_quaternionAlgebra : #(ℍ[R,c₁,c₂,c₃]) = #R ^ 4 := by + rw [mk_congr (QuaternionAlgebra.equivProd c₁ c₂ c₃)] simp only [mk_prod, lift_id] ring @[simp] -theorem mk_quaternionAlgebra_of_infinite [Infinite R] : #(ℍ[R,c₁,c₂]) = #R := by +theorem mk_quaternionAlgebra_of_infinite [Infinite R] : #(ℍ[R,c₁,c₂,c₃]) = #R := by rw [mk_quaternionAlgebra, pow_four] /-- The cardinality of a quaternion algebra, as a set. -/ -theorem mk_univ_quaternionAlgebra : #(Set.univ : Set ℍ[R,c₁,c₂]) = #R ^ 4 := by +theorem mk_univ_quaternionAlgebra : #(Set.univ : Set ℍ[R,c₁,c₂,c₃]) = #R ^ 4 := by rw [mk_univ, mk_quaternionAlgebra] theorem mk_univ_quaternionAlgebra_of_infinite [Infinite R] : - #(Set.univ : Set ℍ[R,c₁,c₂]) = #R := by rw [mk_univ_quaternionAlgebra, pow_four] + #(Set.univ : Set ℍ[R,c₁,c₂,c₃]) = #R := by rw [mk_univ_quaternionAlgebra, pow_four] /-- Show the quaternion ⟨w, x, y, z⟩ as a string "{ re := w, imI := x, imJ := y, imK := z }". For the typical case of quaternions over ℝ, each component will show as a Cauchy sequence due to the way Real numbers are represented. -/ -instance [Repr R] {a b : R} : Repr ℍ[R, a, b] where +instance [Repr R] {a b c : R} : Repr ℍ[R, a, b, c] where reprPrec q _ := s!"\{ re := {repr q.re}, imI := {repr q.imI}, imJ := {repr q.imJ}, imK := {repr q.imK} }" @@ -1343,22 +1390,22 @@ end QuaternionAlgebra section Quaternion -variable (R : Type*) [One R] [Neg R] +variable (R : Type*) [Zero R] [One R] [Neg R] /-- The cardinality of the quaternions, as a type. -/ @[simp] theorem mk_quaternion : #(ℍ[R]) = #R ^ 4 := - mk_quaternionAlgebra _ _ + mk_quaternionAlgebra _ _ _ theorem mk_quaternion_of_infinite [Infinite R] : #(ℍ[R]) = #R := - mk_quaternionAlgebra_of_infinite _ _ + mk_quaternionAlgebra_of_infinite _ _ _ /-- The cardinality of the quaternions, as a set. -/ theorem mk_univ_quaternion : #(Set.univ : Set ℍ[R]) = #R ^ 4 := - mk_univ_quaternionAlgebra _ _ + mk_univ_quaternionAlgebra _ _ _ theorem mk_univ_quaternion_of_infinite [Infinite R] : #(Set.univ : Set ℍ[R]) = #R := - mk_univ_quaternionAlgebra_of_infinite _ _ + mk_univ_quaternionAlgebra_of_infinite _ _ _ end Quaternion diff --git a/Mathlib/Algebra/QuaternionBasis.lean b/Mathlib/Algebra/QuaternionBasis.lean index 0bbcb83c8be73..9fcaf7da1c225 100644 --- a/Mathlib/Algebra/QuaternionBasis.lean +++ b/Mathlib/Algebra/QuaternionBasis.lean @@ -11,11 +11,11 @@ import Mathlib.Tactic.Ring ## Main definitions -* `QuaternionAlgebra.Basis A c₁ c₂`: a basis for a subspace of an `R`-algebra `A` that has the - same algebra structure as `ℍ[R,c₁,c₂]`. -* `QuaternionAlgebra.Basis.self R`: the canonical basis for `ℍ[R,c₁,c₂]`. +* `QuaternionAlgebra.Basis A c₁ c₂ c₃`: a basis for a subspace of an `R`-algebra `A` that has the + same algebra structure as `ℍ[R,c₁,c₂,c₃]`. +* `QuaternionAlgebra.Basis.self R`: the canonical basis for `ℍ[R,c₁,c₂,c₃]`. * `QuaternionAlgebra.Basis.compHom b f`: transform a basis `b` by an AlgHom `f`. -* `QuaternionAlgebra.lift`: Define an `AlgHom` out of `ℍ[R,c₁,c₂]` by its action on the basis +* `QuaternionAlgebra.lift`: Define an `AlgHom` out of `ℍ[R,c₁,c₂,c₃]` by its action on the basis elements `i`, `j`, and `k`. In essence, this is a universal property. Analogous to `Complex.lift`, but takes a bundled `QuaternionAlgebra.Basis` instead of just a `Subtype` as the amount of data / proves is non-negligible. @@ -27,26 +27,27 @@ open Quaternion namespace QuaternionAlgebra /-- A quaternion basis contains the information both sufficient and necessary to construct an -`R`-algebra homomorphism from `ℍ[R,c₁,c₂]` to `A`; or equivalently, a surjective -`R`-algebra homomorphism from `ℍ[R,c₁,c₂]` to an `R`-subalgebra of `A`. +`R`-algebra homomorphism from `ℍ[R,c₁,c₂,c₃]` to `A`; or equivalently, a surjective +`R`-algebra homomorphism from `ℍ[R,c₁,c₂,c₃]` to an `R`-subalgebra of `A`. Note that for definitional convenience, `k` is provided as a field even though `i_mul_j` fully determines it. -/ -structure Basis {R : Type*} (A : Type*) [CommRing R] [Ring A] [Algebra R A] (c₁ c₂ : R) where +structure Basis {R : Type*} (A : Type*) [CommRing R] [Ring A] [Algebra R A] (c₁ c₂ c₃ : R) where (i j k : A) - i_mul_i : i * i = c₁ • (1 : A) - j_mul_j : j * j = c₂ • (1 : A) + i_mul_i : i * i = c₁ • (1 : A) + c₂ • i + j_mul_j : j * j = c₃ • (1 : A) i_mul_j : i * j = k - j_mul_i : j * i = -k + j_mul_i : j * i = c₂ • j - k variable {R : Type*} {A B : Type*} [CommRing R] [Ring A] [Ring B] [Algebra R A] [Algebra R B] -variable {c₁ c₂ : R} +variable {c₁ c₂ c₃: R} namespace Basis /-- Since `k` is redundant, it is not necessary to show `q₁.k = q₂.k` when showing `q₁ = q₂`. -/ @[ext] -protected theorem ext ⦃q₁ q₂ : Basis A c₁ c₂⦄ (hi : q₁.i = q₂.i) (hj : q₁.j = q₂.j) : q₁ = q₂ := by +protected theorem ext ⦃q₁ q₂ : Basis A c₁ c₂ c₃⦄ (hi : q₁.i = q₂.i) + (hj : q₁.j = q₂.j) : q₁ = q₂ := by cases q₁; rename_i q₁_i_mul_j _ cases q₂; rename_i q₂_i_mul_j _ congr @@ -57,7 +58,7 @@ variable (R) /-- There is a natural quaternionic basis for the `QuaternionAlgebra`. -/ @[simps i j k] -protected def self : Basis ℍ[R,c₁,c₂] c₁ c₂ where +protected def self : Basis ℍ[R,c₁,c₂,c₃] c₁ c₂ c₃ where i := ⟨0, 1, 0, 0⟩ i_mul_i := by ext <;> simp j := ⟨0, 0, 1, 0⟩ @@ -68,64 +69,70 @@ protected def self : Basis ℍ[R,c₁,c₂] c₁ c₂ where variable {R} -instance : Inhabited (Basis ℍ[R,c₁,c₂] c₁ c₂) := +instance : Inhabited (Basis ℍ[R,c₁,c₂,c₃] c₁ c₂ c₃) := ⟨Basis.self R⟩ -variable (q : Basis A c₁ c₂) +variable (q : Basis A c₁ c₂ c₃) attribute [simp] i_mul_i j_mul_j i_mul_j j_mul_i @[simp] -theorem i_mul_k : q.i * q.k = c₁ • q.j := by - rw [← i_mul_j, ← mul_assoc, i_mul_i, smul_mul_assoc, one_mul] +theorem i_mul_k : q.i * q.k = c₁ • q.j + c₂ • q.k := by + rw [← i_mul_j, ← mul_assoc, i_mul_i, add_mul, smul_mul_assoc, one_mul, smul_mul_assoc] @[simp] theorem k_mul_i : q.k * q.i = -c₁ • q.j := by - rw [← i_mul_j, mul_assoc, j_mul_i, mul_neg, i_mul_k, neg_smul] + rw [← i_mul_j, mul_assoc, j_mul_i, mul_sub, i_mul_k, neg_smul, mul_smul_comm, i_mul_j] + linear_combination (norm := module) @[simp] -theorem k_mul_j : q.k * q.j = c₂ • q.i := by +theorem k_mul_j : q.k * q.j = c₃ • q.i := by rw [← i_mul_j, mul_assoc, j_mul_j, mul_smul_comm, mul_one] @[simp] -theorem j_mul_k : q.j * q.k = -c₂ • q.i := by - rw [← i_mul_j, ← mul_assoc, j_mul_i, neg_mul, k_mul_j, neg_smul] +theorem j_mul_k : q.j * q.k = (c₂ * c₃) • 1 - c₃ • q.i := by + rw [← i_mul_j, ← mul_assoc, j_mul_i, sub_mul, smul_mul_assoc, j_mul_j, ← smul_assoc, k_mul_j] + rfl @[simp] -theorem k_mul_k : q.k * q.k = -((c₁ * c₂) • (1 : A)) := by - rw [← i_mul_j, mul_assoc, ← mul_assoc q.j _ _, j_mul_i, ← i_mul_j, ← mul_assoc, mul_neg, ← - mul_assoc, i_mul_i, smul_mul_assoc, one_mul, neg_mul, smul_mul_assoc, j_mul_j, smul_smul] +theorem k_mul_k : q.k * q.k = -((c₁ * c₃) • (1 : A)) := by + rw [← i_mul_j, mul_assoc, ← mul_assoc q.j _ _, j_mul_i, ← i_mul_j, ← mul_assoc, mul_sub, ← + mul_assoc, i_mul_i, add_mul, smul_mul_assoc, one_mul, sub_mul, smul_mul_assoc, mul_smul_comm, + smul_mul_assoc, mul_assoc, j_mul_j, add_mul, smul_mul_assoc, j_mul_j, smul_smul, + smul_mul_assoc, mul_assoc, j_mul_j] + linear_combination (norm := module) + /-- Intermediate result used to define `QuaternionAlgebra.Basis.liftHom`. -/ -def lift (x : ℍ[R,c₁,c₂]) : A := +def lift (x : ℍ[R,c₁,c₂,c₃]) : A := algebraMap R _ x.re + x.imI • q.i + x.imJ • q.j + x.imK • q.k -theorem lift_zero : q.lift (0 : ℍ[R,c₁,c₂]) = 0 := by simp [lift] +theorem lift_zero : q.lift (0 : ℍ[R,c₁,c₂,c₃]) = 0 := by simp [lift] -theorem lift_one : q.lift (1 : ℍ[R,c₁,c₂]) = 1 := by simp [lift] +theorem lift_one : q.lift (1 : ℍ[R,c₁,c₂,c₃]) = 1 := by simp [lift] -theorem lift_add (x y : ℍ[R,c₁,c₂]) : q.lift (x + y) = q.lift x + q.lift y := by +theorem lift_add (x y : ℍ[R,c₁,c₂,c₃]) : q.lift (x + y) = q.lift x + q.lift y := by simp only [lift, add_re, map_add, add_imI, add_smul, add_imJ, add_imK] abel -theorem lift_mul (x y : ℍ[R,c₁,c₂]) : q.lift (x * y) = q.lift x * q.lift y := by +theorem lift_mul (x y : ℍ[R,c₁,c₂,c₃]) : q.lift (x * y) = q.lift x * q.lift y := by simp only [lift, Algebra.algebraMap_eq_smul_one] simp_rw [add_mul, mul_add, smul_mul_assoc, mul_smul_comm, one_mul, mul_one, smul_smul] simp only [i_mul_i, j_mul_j, i_mul_j, j_mul_i, i_mul_k, k_mul_i, k_mul_j, j_mul_k, k_mul_k] simp only [smul_smul, smul_neg, sub_eq_add_neg, add_smul, ← add_assoc, mul_neg, neg_smul] - simp only [mul_right_comm _ _ (c₁ * c₂), mul_comm _ (c₁ * c₂)] + simp only [mul_right_comm _ _ (c₁ * c₃), mul_comm _ (c₁ * c₃)] simp only [mul_comm _ c₁, mul_right_comm _ _ c₁] - simp only [mul_comm _ c₂, mul_right_comm _ _ c₂] + simp only [mul_comm _ c₂, mul_right_comm _ _ c₃] simp only [← mul_comm c₁ c₂, ← mul_assoc] simp only [mul_re, sub_eq_add_neg, add_smul, neg_smul, mul_imI, ← add_assoc, mul_imJ, mul_imK] - abel + linear_combination (norm := module) -theorem lift_smul (r : R) (x : ℍ[R,c₁,c₂]) : q.lift (r • x) = r • q.lift x := by +theorem lift_smul (r : R) (x : ℍ[R,c₁,c₂,c₃]) : q.lift (r • x) = r • q.lift x := by simp [lift, mul_smul, ← Algebra.smul_def] /-- A `QuaternionAlgebra.Basis` implies an `AlgHom` from the quaternions. -/ @[simps!] -def liftHom : ℍ[R,c₁,c₂] →ₐ[R] A := +def liftHom : ℍ[R,c₁,c₂,c₃] →ₐ[R] A := AlgHom.mk' { toFun := q.lift map_zero' := q.lift_zero @@ -135,20 +142,20 @@ def liftHom : ℍ[R,c₁,c₂] →ₐ[R] A := /-- Transform a `QuaternionAlgebra.Basis` through an `AlgHom`. -/ @[simps i j k] -def compHom (F : A →ₐ[R] B) : Basis B c₁ c₂ where +def compHom (F : A →ₐ[R] B) : Basis B c₁ c₂ c₃ where i := F q.i - i_mul_i := by rw [← map_mul, q.i_mul_i, map_smul, map_one] + i_mul_i := by rw [← map_mul, q.i_mul_i, map_add, map_smul, map_smul, map_one] j := F q.j j_mul_j := by rw [← map_mul, q.j_mul_j, map_smul, map_one] k := F q.k i_mul_j := by rw [← map_mul, q.i_mul_j] - j_mul_i := by rw [← map_mul, q.j_mul_i, map_neg] + j_mul_i := by rw [← map_mul, q.j_mul_i, map_sub, map_smul] end Basis /-- A quaternionic basis on `A` is equivalent to a map from the quaternion algebra to `A`. -/ @[simps] -def lift : Basis A c₁ c₂ ≃ (ℍ[R,c₁,c₂] →ₐ[R] A) where +def lift : Basis A c₁ c₂ c₃ ≃ (ℍ[R,c₁,c₂,c₃] →ₐ[R] A) where toFun := Basis.liftHom invFun := (Basis.self R).compHom left_inv q := by ext <;> simp [Basis.lift] @@ -162,7 +169,7 @@ def lift : Basis A c₁ c₂ ≃ (ℍ[R,c₁,c₂] →ₐ[R] A) where /-- Two `R`-algebra morphisms from a quaternion algebra are equal if they agree on `i` and `j`. -/ @[ext] -theorem hom_ext ⦃f g : ℍ[R,c₁,c₂] →ₐ[R] A⦄ +theorem hom_ext ⦃f g : ℍ[R,c₁,c₂,c₃] →ₐ[R] A⦄ (hi : f (Basis.self R).i = g (Basis.self R).i) (hj : f (Basis.self R).j = g (Basis.self R).j) : f = g := lift.symm.injective <| Basis.ext hi hj diff --git a/Mathlib/Algebra/Quotient.lean b/Mathlib/Algebra/Quotient.lean index ea7436110634b..608df82c9e2de 100644 --- a/Mathlib/Algebra/Quotient.lean +++ b/Mathlib/Algebra/Quotient.lean @@ -46,7 +46,7 @@ class HasQuotient (A : outParam <| Type u) (B : Type v) where quotient' : B → Type max u v -- Will be provided by e.g. `Ideal.Quotient.inhabited` -/-- `HasQuotient.Quotient A b` (with notation `A ⧸ b`) is the quotient +/-- `HasQuotient.Quotient A b` (denoted as `A ⧸ b`) is the quotient of the type `A` by `b`. This differs from `HasQuotient.quotient'` in that the `A` argument is diff --git a/Mathlib/Algebra/Ring/Action/ConjAct.lean b/Mathlib/Algebra/Ring/Action/ConjAct.lean new file mode 100644 index 0000000000000..f6401f3572c32 --- /dev/null +++ b/Mathlib/Algebra/Ring/Action/ConjAct.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Algebra.Ring.Action.Basic +import Mathlib.GroupTheory.GroupAction.ConjAct + +/-! +# Conjugation action of a ring on itself +-/ + +assert_not_exists Field + +namespace ConjAct +variable {R : Type*} [Semiring R] + +instance unitsMulSemiringAction : MulSemiringAction (ConjAct Rˣ) R := + { ConjAct.unitsMulDistribMulAction with + smul_zero := by simp [units_smul_def] + smul_add := by simp [units_smul_def, mul_add, add_mul] } + +end ConjAct diff --git a/Mathlib/Algebra/Ring/BooleanRing.lean b/Mathlib/Algebra/Ring/BooleanRing.lean index a40d3face51cc..bf9d1d82ec74c 100644 --- a/Mathlib/Algebra/Ring/BooleanRing.lean +++ b/Mathlib/Algebra/Ring/BooleanRing.lean @@ -8,6 +8,7 @@ import Mathlib.Tactic.Abel import Mathlib.Tactic.Ring import Mathlib.Order.Hom.Lattice import Mathlib.Algebra.Ring.Equiv +import Mathlib.Algebra.Group.Idempotent /-! # Boolean rings @@ -46,17 +47,18 @@ variable {α β γ : Type*} /-- A Boolean ring is a ring where multiplication is idempotent. -/ class BooleanRing (α) extends Ring α where /-- Multiplication in a boolean ring is idempotent. -/ - mul_self : ∀ a : α, a * a = a + isIdempotentElem (a : α) : IsIdempotentElem a namespace BooleanRing variable [BooleanRing α] (a b : α) +@[scoped simp] +lemma mul_self : a * a = a := IsIdempotentElem.eq (isIdempotentElem a) + instance : Std.IdempotentOp (α := α) (· * ·) := ⟨BooleanRing.mul_self⟩ -attribute [scoped simp] mul_self - @[scoped simp] theorem add_self : a + a = 0 := by have : a + a = a + a + (a + a) := @@ -410,7 +412,7 @@ abbrev BooleanAlgebra.toBooleanRing : BooleanRing α where one := ⊤ one_mul := top_inf_eq mul_one := inf_top_eq - mul_self := inf_idem + isIdempotentElem := inf_idem scoped[BooleanRingOfBooleanAlgebra] attribute [instance] GeneralizedBooleanAlgebra.toNonUnitalCommRing BooleanAlgebra.toBooleanRing @@ -529,7 +531,7 @@ instance : BooleanRing Bool where mul_one := Bool.and_true left_distrib := and_xor_distrib_left right_distrib := and_xor_distrib_right - mul_self := Bool.and_self + isIdempotentElem := Bool.and_self zero_mul _ := rfl mul_zero a := by cases a <;> rfl nsmul := nsmulRec diff --git a/Mathlib/Algebra/Ring/Commute.lean b/Mathlib/Algebra/Ring/Commute.lean index 381b7109ee286..794103b3e5ae2 100644 --- a/Mathlib/Algebra/Ring/Commute.lean +++ b/Mathlib/Algebra/Ring/Commute.lean @@ -133,7 +133,6 @@ lemma neg_pow' (a : R) (n : ℕ) : (-a) ^ n = a ^ n * (-1) ^ n := lemma neg_sq (a : R) : (-a) ^ 2 = a ^ 2 := by simp [sq] --- Porting note: removed the simp attribute to please the simpNF linter lemma neg_one_sq : (-1 : R) ^ 2 = 1 := by simp [neg_sq, one_pow] alias neg_pow_two := neg_sq diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 474e811d144ba..7288db996a5d0 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -443,6 +443,7 @@ theorem ofBijective_apply [NonUnitalRingHomClass F R S] (f : F) (hf : Function.B /-- Product of a singleton family of (non-unital non-associative semi)rings is isomorphic to the only member of this family. -/ +@[simps! (config := .asFn)] def piUnique {ι : Type*} (R : ι → Type*) [Unique ι] [∀ i, NonUnitalNonAssocSemiring (R i)] : (∀ i, R i) ≃+* R default where __ := Equiv.piUnique R diff --git a/Mathlib/Algebra/Ring/Int/Parity.lean b/Mathlib/Algebra/Ring/Int/Parity.lean index 95909efda8ec2..3a8f37de3bdb1 100644 --- a/Mathlib/Algebra/Ring/Int/Parity.lean +++ b/Mathlib/Algebra/Ring/Int/Parity.lean @@ -113,20 +113,7 @@ lemma four_dvd_add_or_sub_of_odd {a b : ℤ} (ha : Odd a) (hb : Odd b) : 4 ∣ a + b ∨ 4 ∣ a - b := by obtain ⟨m, rfl⟩ := ha obtain ⟨n, rfl⟩ := hb - obtain h | h := Int.even_or_odd (m + n) - · right - rw [Int.even_add, ← Int.even_sub] at h - obtain ⟨k, hk⟩ := h - convert dvd_mul_right 4 k using 1 - rw [eq_add_of_sub_eq hk, mul_add, add_assoc, add_sub_cancel_right, ← two_mul, ← mul_assoc] - rfl - · left - obtain ⟨k, hk⟩ := h - convert dvd_mul_right 4 (k + 1) using 1 - rw [eq_sub_of_add_eq hk, add_right_comm, ← add_sub, mul_add, mul_sub, add_assoc, add_assoc, - sub_add, add_assoc, ← sub_sub (2 * n), sub_self, zero_sub, sub_neg_eq_add, ← mul_assoc, - mul_add] - rfl + omega lemma two_mul_ediv_two_add_one_of_odd : Odd n → 2 * (n / 2) + 1 = n := by rintro ⟨c, rfl⟩ diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 0883314e27689..ed97a0457e825 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -54,7 +54,9 @@ lemma Even.neg_one_zpow (h : Even n) : (-1 : α) ^ n = 1 := by rw [h.neg_zpow, o end DivisionMonoid -@[simp] lemma isSquare_zero [MulZeroClass α] : IsSquare (0 : α) := ⟨0, (mul_zero _).symm⟩ +@[simp] lemma IsSquare.zero [MulZeroClass α] : IsSquare (0 : α) := ⟨0, (mul_zero _).symm⟩ + +@[deprecated (since := "2024-01-07")] alias isSquare_zero := IsSquare.zero section Semiring variable [Semiring α] [Semiring β] {a b : α} {m n : ℕ} @@ -255,7 +257,7 @@ lemma even_sub' (h : n ≤ m) : Even (m - n) ↔ (Odd m ↔ Odd n) := by lemma Odd.sub_odd (hm : Odd m) (hn : Odd n) : Even (m - n) := (le_total n m).elim (fun h ↦ by simp only [even_sub' h, *]) fun h ↦ by - simp only [Nat.sub_eq_zero_iff_le.2 h, even_zero] + simp only [Nat.sub_eq_zero_iff_le.2 h, Even.zero] alias _root_.Odd.tsub_odd := Nat.Odd.sub_odd diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 6e38a46a7b4a4..8c113ce263a89 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -384,6 +384,15 @@ theorem center_eq_top (R) [CommRing R] : center R = ⊤ := instance : CommRing (center R) := { inferInstanceAs (CommSemiring (Subsemiring.center R)), (center R).toRing with } +/-- The center of isomorphic (not necessarily associative) rings are isomorphic. -/ +@[simps!] def centerCongr (e : R ≃+* S) : center R ≃+* center S := + NonUnitalSubsemiring.centerCongr e + +/-- The center of a (not necessarily associative) ring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ := + NonUnitalSubsemiring.centerToMulOpposite + end section DivisionRing @@ -873,14 +882,9 @@ theorem ofLeftInverse_symm_apply {g : S → R} {f : R →+* S} (h : Function.Lef /-- Given an equivalence `e : R ≃+* S` of rings and a subring `s` of `R`, `subringMap e s` is the induced equivalence between `s` and `s.map e` -/ -@[simps!] def subringMap (e : R ≃+* S) : s ≃+* s.map e.toRingHom := e.subsemiringMap s.toSubsemiring --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] RingEquiv.subringMap_symm_apply_coe - RingEquiv.subringMap_apply_coe - end RingEquiv namespace Subring diff --git a/Mathlib/Algebra/Ring/Subring/Defs.lean b/Mathlib/Algebra/Ring/Subring/Defs.lean index 4ac441db39d99..69b8b9aab8fb8 100644 --- a/Mathlib/Algebra/Ring/Subring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subring/Defs.lean @@ -97,14 +97,14 @@ instance (priority := 75) toHasIntCast : IntCast s := -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a ring inherits a ring structure -/ -instance (priority := 75) toRing : Ring s := +instance (priority := 75) toRing : Ring s := fast_instance% Subtype.coe_injective.ring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a `CommRing` is a `CommRing`. -/ instance (priority := 75) toCommRing {R} [CommRing R] [SetLike S R] [SubringClass S R] : - CommRing s := + CommRing s := fast_instance% Subtype.coe_injective.commRing Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl diff --git a/Mathlib/Algebra/Ring/Subring/Order.lean b/Mathlib/Algebra/Ring/Subring/Order.lean index e5e5fe0305d5a..ff7a1d9af142d 100644 --- a/Mathlib/Algebra/Ring/Subring/Order.lean +++ b/Mathlib/Algebra/Ring/Subring/Order.lean @@ -26,7 +26,7 @@ variable {R S : Type*} [SetLike S R] (s : S) -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of an `OrderedRing` is an `OrderedRing`. -/ instance (priority := 75) toOrderedRing [OrderedRing R] [SubringClass S R] : - OrderedRing s := + OrderedRing s := fast_instance% Subtype.coe_injective.orderedRing Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl @@ -34,7 +34,7 @@ instance (priority := 75) toOrderedRing [OrderedRing R] [SubringClass S R] : -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of an `OrderedCommRing` is an `OrderedCommRing`. -/ instance (priority := 75) toOrderedCommRing [OrderedCommRing R] [SubringClass S R] : - OrderedCommRing s := + OrderedCommRing s := fast_instance% Subtype.coe_injective.orderedCommRing Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl @@ -42,7 +42,7 @@ instance (priority := 75) toOrderedCommRing [OrderedCommRing R] [SubringClass S -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a `LinearOrderedRing` is a `LinearOrderedRing`. -/ instance (priority := 75) toLinearOrderedRing [LinearOrderedRing R] [SubringClass S R] : - LinearOrderedRing s := + LinearOrderedRing s := fast_instance% Subtype.coe_injective.linearOrderedRing Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -50,7 +50,7 @@ instance (priority := 75) toLinearOrderedRing [LinearOrderedRing R] [SubringClas -- Prefer subclasses of `Ring` over subclasses of `SubringClass`. /-- A subring of a `LinearOrderedCommRing` is a `LinearOrderedCommRing`. -/ instance (priority := 75) toLinearOrderedCommRing [LinearOrderedCommRing R] [SubringClass S R] : - LinearOrderedCommRing s := + LinearOrderedCommRing s := fast_instance% Subtype.coe_injective.linearOrderedCommRing Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/Algebra/Ring/Subring/Pointwise.lean b/Mathlib/Algebra/Ring/Subring/Pointwise.lean index b472910f34900..da145f5a27b74 100644 --- a/Mathlib/Algebra/Ring/Subring/Pointwise.lean +++ b/Mathlib/Algebra/Ring/Subring/Pointwise.lean @@ -108,13 +108,13 @@ theorem mem_inv_pointwise_smul_iff {a : M} {S : Subring R} {x : R} : x ∈ a⁻ @[simp] theorem pointwise_smul_le_pointwise_smul_iff {a : M} {S T : Subring R} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff + smul_set_subset_smul_set_iff theorem pointwise_smul_subset_iff {a : M} {S T : Subring R} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff + smul_set_subset_iff_subset_inv_smul_set theorem subset_pointwise_smul_iff {a : M} {S T : Subring R} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff + subset_smul_set_iff /-! TODO: add `equivSMul` like we have for subgroup. -/ @@ -143,13 +143,13 @@ theorem mem_inv_pointwise_smul_iff₀ {a : M} (ha : a ≠ 0) (S : Subring R) (x @[simp] theorem pointwise_smul_le_pointwise_smul_iff₀ {a : M} (ha : a ≠ 0) {S T : Subring R} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff₀ ha + smul_set_subset_smul_set_iff₀ ha theorem pointwise_smul_le_iff₀ {a : M} (ha : a ≠ 0) {S T : Subring R} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff₀ ha + smul_set_subset_iff₀ ha theorem le_pointwise_smul_iff₀ {a : M} (ha : a ≠ 0) {S T : Subring R} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff₀ ha + subset_smul_set_iff₀ ha end GroupWithZero diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index e38a8720a241a..65b72af50f1f4 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -262,6 +262,17 @@ This is not an instance as it forms a non-defeq diamond with abbrev center.commSemiring' : CommSemiring (center R) := { Submonoid.center.commMonoid', (center R).toNonAssocSemiring with } +variable {R} + +/-- The center of isomorphic (not necessarily associative) semirings are isomorphic. -/ +@[simps!] def centerCongr [NonAssocSemiring S] (e : R ≃+* S) : center R ≃+* center S := + NonUnitalSubsemiring.centerCongr e + +/-- The center of a (not necessarily associative) semiring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ := + NonUnitalSubsemiring.centerToMulOpposite + end NonAssocSemiring section Semiring @@ -285,7 +296,6 @@ instance decidableMemCenter {R} [Semiring R] [DecidableEq R] [Fintype R] : theorem center_eq_top (R) [CommSemiring R] : center R = ⊤ := SetLike.coe_injective (Set.center_eq_univ R) - end Semiring section Centralizer diff --git a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean index 3d0186346b217..f409e13f8b5bc 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Defs.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Defs.lean @@ -71,7 +71,7 @@ namespace SubsemiringClass -- Prefer subclasses of `NonAssocSemiring` over subclasses of `SubsemiringClass`. /-- A subsemiring of a `NonAssocSemiring` inherits a `NonAssocSemiring` structure -/ -instance (priority := 75) toNonAssocSemiring : NonAssocSemiring s := +instance (priority := 75) toNonAssocSemiring : NonAssocSemiring s := fast_instance% Subtype.coe_injective.nonAssocSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl @@ -92,7 +92,7 @@ theorem coe_subtype : (subtype s : s → R) = ((↑) : s → R) := -- Prefer subclasses of `Semiring` over subclasses of `SubsemiringClass`. /-- A subsemiring of a `Semiring` is a `Semiring`. -/ instance (priority := 75) toSemiring {R} [Semiring R] [SetLike S R] [SubsemiringClass S R] : - Semiring s := + Semiring s := fast_instance% Subtype.coe_injective.semiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl @@ -105,7 +105,7 @@ theorem coe_pow {R} [Semiring R] [SetLike S R] [SubsemiringClass S R] (x : s) (n /-- A subsemiring of a `CommSemiring` is a `CommSemiring`. -/ instance toCommSemiring {R} [CommSemiring R] [SetLike S R] [SubsemiringClass S R] : - CommSemiring s := + CommSemiring s := fast_instance% Subtype.coe_injective.commSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl diff --git a/Mathlib/Algebra/Ring/Subsemiring/Order.lean b/Mathlib/Algebra/Ring/Subsemiring/Order.lean index cfe828f062d3f..3b042de898b4b 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Order.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Order.lean @@ -6,6 +6,7 @@ Authors: Damiano Testa import Mathlib.Algebra.Order.Ring.InjSurj import Mathlib.Algebra.Ring.Subsemiring.Defs import Mathlib.Order.Interval.Set.Defs +import Mathlib.Tactic.FastInstance /-! # `Order`ed instances for `SubsemiringClass` and `Subsemiring`. @@ -16,37 +17,37 @@ variable {R S : Type*} [SetLike S R] (s : S) /-- A subsemiring of an `OrderedSemiring` is an `OrderedSemiring`. -/ instance toOrderedSemiring [OrderedSemiring R] [SubsemiringClass S R] : - OrderedSemiring s := + OrderedSemiring s := fast_instance% Subtype.coe_injective.orderedSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of a `StrictOrderedSemiring` is a `StrictOrderedSemiring`. -/ instance toStrictOrderedSemiring [StrictOrderedSemiring R] - [SubsemiringClass S R] : StrictOrderedSemiring s := + [SubsemiringClass S R] : StrictOrderedSemiring s := fast_instance% Subtype.coe_injective.strictOrderedSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of an `OrderedCommSemiring` is an `OrderedCommSemiring`. -/ instance toOrderedCommSemiring [OrderedCommSemiring R] [SubsemiringClass S R] : - OrderedCommSemiring s := + OrderedCommSemiring s := fast_instance% Subtype.coe_injective.orderedCommSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of a `StrictOrderedCommSemiring` is a `StrictOrderedCommSemiring`. -/ instance toStrictOrderedCommSemiring [StrictOrderedCommSemiring R] - [SubsemiringClass S R] : StrictOrderedCommSemiring s := + [SubsemiringClass S R] : StrictOrderedCommSemiring s := fast_instance% Subtype.coe_injective.strictOrderedCommSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of a `LinearOrderedSemiring` is a `LinearOrderedSemiring`. -/ instance toLinearOrderedSemiring [LinearOrderedSemiring R] - [SubsemiringClass S R] : LinearOrderedSemiring s := + [SubsemiringClass S R] : LinearOrderedSemiring s := fast_instance% Subtype.coe_injective.linearOrderedSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) fun _ _ => rfl /-- A subsemiring of a `LinearOrderedCommSemiring` is a `LinearOrderedCommSemiring`. -/ instance toLinearOrderedCommSemiring [LinearOrderedCommSemiring R] - [SubsemiringClass S R] : LinearOrderedCommSemiring s := + [SubsemiringClass S R] : LinearOrderedCommSemiring s := fast_instance% Subtype.coe_injective.linearOrderedCommSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -58,37 +59,38 @@ namespace Subsemiring variable {R : Type*} /-- A subsemiring of an `OrderedSemiring` is an `OrderedSemiring`. -/ -instance toOrderedSemiring [OrderedSemiring R] (s : Subsemiring R) : OrderedSemiring s := +instance toOrderedSemiring [OrderedSemiring R] (s : Subsemiring R) : + OrderedSemiring s := fast_instance% Subtype.coe_injective.orderedSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of a `StrictOrderedSemiring` is a `StrictOrderedSemiring`. -/ instance toStrictOrderedSemiring [StrictOrderedSemiring R] (s : Subsemiring R) : - StrictOrderedSemiring s := + StrictOrderedSemiring s := fast_instance% Subtype.coe_injective.strictOrderedSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of an `OrderedCommSemiring` is an `OrderedCommSemiring`. -/ instance toOrderedCommSemiring [OrderedCommSemiring R] (s : Subsemiring R) : - OrderedCommSemiring s := + OrderedCommSemiring s := fast_instance% Subtype.coe_injective.orderedCommSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of a `StrictOrderedCommSemiring` is a `StrictOrderedCommSemiring`. -/ instance toStrictOrderedCommSemiring [StrictOrderedCommSemiring R] (s : Subsemiring R) : - StrictOrderedCommSemiring s := + StrictOrderedCommSemiring s := fast_instance% Subtype.coe_injective.strictOrderedCommSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl /-- A subsemiring of a `LinearOrderedSemiring` is a `LinearOrderedSemiring`. -/ instance toLinearOrderedSemiring [LinearOrderedSemiring R] (s : Subsemiring R) : - LinearOrderedSemiring s := + LinearOrderedSemiring s := fast_instance% Subtype.coe_injective.linearOrderedSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) fun _ _ => rfl /-- A subsemiring of a `LinearOrderedCommSemiring` is a `LinearOrderedCommSemiring`. -/ instance toLinearOrderedCommSemiring [LinearOrderedCommSemiring R] (s : Subsemiring R) : - LinearOrderedCommSemiring s := + LinearOrderedCommSemiring s := fast_instance% Subtype.coe_injective.linearOrderedCommSemiring Subtype.val rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/Algebra/Ring/Subsemiring/Pointwise.lean b/Mathlib/Algebra/Ring/Subsemiring/Pointwise.lean index 250d388ca8b0a..4e24ba3121d75 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Pointwise.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Pointwise.lean @@ -103,13 +103,13 @@ theorem mem_inv_pointwise_smul_iff {a : M} {S : Subsemiring R} {x : R} : x ∈ a @[simp] theorem pointwise_smul_le_pointwise_smul_iff {a : M} {S T : Subsemiring R} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff + smul_set_subset_smul_set_iff theorem pointwise_smul_subset_iff {a : M} {S T : Subsemiring R} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff + smul_set_subset_iff_subset_inv_smul_set theorem subset_pointwise_smul_iff {a : M} {S T : Subsemiring R} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff + subset_smul_set_iff /-! TODO: add `equiv_smul` like we have for subgroup. -/ @@ -138,15 +138,15 @@ theorem mem_inv_pointwise_smul_iff₀ {a : M} (ha : a ≠ 0) (S : Subsemiring R) @[simp] theorem pointwise_smul_le_pointwise_smul_iff₀ {a : M} (ha : a ≠ 0) {S T : Subsemiring R} : a • S ≤ a • T ↔ S ≤ T := - set_smul_subset_set_smul_iff₀ ha + smul_set_subset_smul_set_iff₀ ha theorem pointwise_smul_le_iff₀ {a : M} (ha : a ≠ 0) {S T : Subsemiring R} : a • S ≤ T ↔ S ≤ a⁻¹ • T := - set_smul_subset_iff₀ ha + smul_set_subset_iff₀ ha theorem le_pointwise_smul_iff₀ {a : M} (ha : a ≠ 0) {S T : Subsemiring R} : S ≤ a • T ↔ a⁻¹ • S ≤ T := - subset_set_smul_iff₀ ha + subset_smul_set_iff₀ ha end GroupWithZero diff --git a/Mathlib/Algebra/Star/CHSH.lean b/Mathlib/Algebra/Star/CHSH.lean index 6b716f7e73fde..45bf5a0c7bc8e 100644 --- a/Mathlib/Algebra/Star/CHSH.lean +++ b/Mathlib/Algebra/Star/CHSH.lean @@ -82,7 +82,6 @@ the `Aᵢ` commute with the `Bⱼ`. The physical interpretation is that `A₀` and `A₁` are a pair of boolean observables which are spacelike separated from another pair `B₀` and `B₁` of boolean observables. -/ ---@[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet structure IsCHSHTuple {R} [Monoid R] [StarMul R] (A₀ A₁ B₀ B₁ : R) : Prop where A₀_inv : A₀ ^ 2 = 1 A₁_inv : A₁ ^ 2 = 1 diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index 5d4e197bf27e2..8a465c5325c76 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -30,8 +30,8 @@ The approach taken here is inspired by `Mathlib/Algebra/Opposites.lean`. We use open Function -/-- The symmetrized algebra has the same underlying space as the original algebra. --/ +/-- The symmetrized algebra (denoted as `αˢʸᵐ`) +has the same underlying space as the original algebra `α`. -/ def SymAlg (α : Type*) : Type _ := α diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index a2a2716bb20af..4772bc08cc2a9 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -46,8 +46,6 @@ namespace AlgebraicGeometry open Spec (structureSheaf) /-- The category of affine schemes -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] def AffineScheme := Scheme.Spec.EssImageSubcategory deriving Category @@ -76,6 +74,16 @@ theorem Scheme.isoSpec_inv_naturality {X Y : Scheme} [IsAffine X] [IsAffine Y] ( rw [Iso.eq_inv_comp, isoSpec, asIso_hom, ← Scheme.toSpecΓ_naturality_assoc, isoSpec, asIso_inv, IsIso.hom_inv_id, Category.comp_id] +@[reassoc (attr := simp)] +lemma Scheme.toSpecΓ_isoSpec_inv (X : Scheme.{u}) [IsAffine X] : + X.toSpecΓ ≫ X.isoSpec.inv = 𝟙 _ := + X.isoSpec.hom_inv_id + +@[reassoc (attr := simp)] +lemma Scheme.isoSpec_inv_toSpecΓ (X : Scheme.{u}) [IsAffine X] : + X.isoSpec.inv ≫ X.toSpecΓ = 𝟙 _ := + X.isoSpec.inv_hom_id + /-- Construct an affine scheme from a scheme and the information that it is affine. Also see `AffineScheme.of` for a typeclass version. -/ @[simps] @@ -283,12 +291,17 @@ def Scheme.Opens.toSpecΓ {X : Scheme.{u}} (U : X.Opens) : lemma Scheme.Opens.toSpecΓ_SpecMap_map {X : Scheme} (U V : X.Opens) (h : U ≤ V) : U.toSpecΓ ≫ Spec.map (X.presheaf.map (homOfLE h).op) = X.homOfLE h ≫ V.toSpecΓ := by delta Scheme.Opens.toSpecΓ - simp [← Spec.map_comp, ← X.presheaf.map_comp] + simp [← Spec.map_comp, ← X.presheaf.map_comp, toSpecΓ_naturality_assoc] @[simp] lemma Scheme.Opens.toSpecΓ_top {X : Scheme} : (⊤ : X.Opens).toSpecΓ = (⊤ : X.Opens).ι ≫ X.toSpecΓ := by - simp [Scheme.Opens.toSpecΓ]; rfl + simp [Scheme.Opens.toSpecΓ, toSpecΓ_naturality]; rfl + +@[reassoc] +lemma Scheme.Opens.toSpecΓ_appTop {X : Scheme.{u}} (U : X.Opens) : + U.toSpecΓ.appTop = (Scheme.ΓSpecIso Γ(X, U)).hom ≫ U.topIso.inv := by + simp [Scheme.Opens.toSpecΓ] namespace IsAffineOpen @@ -302,8 +315,13 @@ def isoSpec : haveI : IsAffine U := hU U.toScheme.isoSpec ≪≫ Scheme.Spec.mapIso U.topIso.symm.op -lemma isoSpec_hom {X : Scheme.{u}} {U : X.Opens} (hU : IsAffineOpen U) : - hU.isoSpec.hom = U.toSpecΓ := rfl +lemma isoSpec_hom : hU.isoSpec.hom = U.toSpecΓ := rfl + +@[reassoc (attr := simp)] +lemma toSpecΓ_isoSpec_inv : U.toSpecΓ ≫ hU.isoSpec.inv = 𝟙 _ := hU.isoSpec.hom_inv_id + +@[reassoc (attr := simp)] +lemma isoSpec_inv_toSpecΓ : hU.isoSpec.inv ≫ U.toSpecΓ = 𝟙 _ := hU.isoSpec.inv_hom_id open IsLocalRing in lemma isoSpec_hom_base_apply (x : U) : @@ -347,6 +365,9 @@ instance isOpenImmersion_fromSpec : @[reassoc (attr := simp)] lemma isoSpec_inv_ι : hU.isoSpec.inv ≫ U.ι = hU.fromSpec := rfl +@[reassoc (attr := simp)] +lemma toSpecΓ_fromSpec : U.toSpecΓ ≫ hU.fromSpec = U.ι := toSpecΓ_isoSpec_inv_assoc _ _ + @[simp] theorem range_fromSpec : Set.range hU.fromSpec.base = (U : Set X) := by @@ -609,7 +630,7 @@ lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (hU : IsA letI := hV.isLocalization_basicOpen (f.appLE U V e r) ext : 1 apply IsLocalization.ringHom_ext (.powers r) - rw [IsLocalization.Away.map, IsLocalization.map_comp, + rw [IsLocalization.Away.map, CommRingCat.hom_ofHom, IsLocalization.map_comp, RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← CommRingCat.hom_comp, ← CommRingCat.hom_comp, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] @@ -625,7 +646,7 @@ lemma app_basicOpen_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} haveI := h.isLocalization_basicOpen (f.app U r) ext : 1 apply IsLocalization.ringHom_ext (.powers r) - rw [IsLocalization.Away.map, CommRingCat.hom_comp, RingHom.comp_assoc, + rw [IsLocalization.Away.map, CommRingCat.hom_comp, RingHom.comp_assoc, CommRingCat.hom_ofHom, IsLocalization.map_comp, RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← RingHom.comp_assoc, ← CommRingCat.hom_comp, ← CommRingCat.hom_comp, ← X.presheaf.map_comp] @@ -796,31 +817,35 @@ def SpecMapRestrictBasicOpenIso {R S : CommRingCat} (f : R ⟶ S) (r : R) : Arrow.mk (Spec.map <| CommRingCat.ofHom (Localization.awayMap f.hom r)) := by letI e₁ : Localization.Away r ≃ₐ[R] Γ(Spec R, basicOpen r) := IsLocalization.algEquiv (Submonoid.powers r) _ _ - letI e₂ : Localization.Away (f r) ≃ₐ[S] Γ(Spec S, basicOpen (f r)) := - IsLocalization.algEquiv (Submonoid.powers (f r)) _ _ + letI e₂ : Localization.Away (f.hom r) ≃ₐ[S] Γ(Spec S, basicOpen (f.hom r)) := + IsLocalization.algEquiv (Submonoid.powers (f.hom r)) _ _ refine Arrow.isoMk ?_ ?_ ?_ · exact (Spec (.of S)).isoOfEq (comap_basicOpen _ _) ≪≫ - (IsAffineOpen.Spec_basicOpen (f r)).isoSpec ≪≫ Scheme.Spec.mapIso e₂.toCommRingCatIso.op + (IsAffineOpen.Spec_basicOpen (f.hom r)).isoSpec ≪≫ Scheme.Spec.mapIso e₂.toCommRingCatIso.op · exact (IsAffineOpen.Spec_basicOpen r).isoSpec ≪≫ Scheme.Spec.mapIso e₁.toCommRingCatIso.op · have := AlgebraicGeometry.IsOpenImmersion.of_isLocalization (S := (Localization.Away r)) r rw [← cancel_mono (Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away r))))] simp only [Arrow.mk_left, Arrow.mk_right, Functor.id_obj, Scheme.isoOfEq_rfl, Iso.refl_trans, - Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Iso.symm_hom, - Scheme.Spec_map, Quiver.Hom.unop_op, Arrow.mk_hom, Category.assoc, - ← Spec.map_comp] - show _ ≫ Spec.map (CommRingCat.ofHom - ((e₂.toRingHom.comp (Localization.awayMap f.hom r)).comp (algebraMap R _))) - = _ ≫ _ ≫ Spec.map (CommRingCat.ofHom (e₁.toRingHom.comp (algebraMap R _))) - rw [RingHom.comp_assoc] - conv => enter [1,2,1,1,2]; tactic => exact IsLocalization.map_comp _ - rw [← RingHom.comp_assoc] - conv => enter [1,2,1,1,1]; tactic => exact e₂.toAlgHom.comp_algebraMap - conv => enter [2,2,2,1,1]; tactic => exact e₁.toAlgHom.comp_algebraMap - show _ ≫ Spec.map (f ≫ (Scheme.ΓSpecIso S).inv ≫ - (Spec S).presheaf.map (homOfLE le_top).op) = - Spec.map f ∣_ basicOpen r ≫ _ ≫ Spec.map ((Scheme.ΓSpecIso R).inv ≫ - (Spec R).presheaf.map (homOfLE le_top).op) + Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, Scheme.Spec_map, Quiver.Hom.unop_op, + Arrow.mk_hom, Category.assoc, ← Spec.map_comp] + conv => + congr + · enter [2, 1]; tactic => + show _ = + (f ≫ (Scheme.ΓSpecIso S).inv ≫ (Spec S).presheaf.map (homOfLE le_top).op) + ext + simp only [Localization.awayMap, IsLocalization.Away.map, AlgEquiv.toRingEquiv_eq_coe, + RingEquiv.toCommRingCatIso_hom, AlgEquiv.toRingEquiv_toRingHom, CommRingCat.hom_comp, + CommRingCat.hom_ofHom, RingHom.comp_apply, IsLocalization.map_eq, RingHom.coe_coe, + AlgEquiv.commutes, IsAffineOpen.algebraMap_Spec_obj] + · enter [2, 2, 1]; tactic => + show _ = (Scheme.ΓSpecIso R).inv ≫ (Spec R).presheaf.map (homOfLE le_top).op + ext + simp only [AlgEquiv.toRingEquiv_eq_coe, RingEquiv.toCommRingCatIso_hom, + AlgEquiv.toRingEquiv_toRingHom, CommRingCat.hom_comp, CommRingCat.hom_ofHom, + RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, AlgEquiv.commutes, + IsAffineOpen.algebraMap_Spec_obj, homOfLE_leOfHom] simp only [IsAffineOpen.isoSpec_hom, homOfLE_leOfHom, Spec.map_comp, Category.assoc, Scheme.Opens.toSpecΓ_SpecMap_map_assoc, Scheme.Opens.toSpecΓ_top, Scheme.homOfLE_ι_assoc, morphismRestrict_ι_assoc] @@ -895,36 +920,136 @@ theorem of_affine_open_cover {X : Scheme} {P : X.affineOpens → Prop} section ZeroLocus -/-- On a locally ringed space `X`, the preimage of the zero locus of the prime spectrum -of `Γ(X, ⊤)` under `toΓSpecFun` agrees with the associated zero locus on `X`. -/ -lemma Scheme.toΓSpec_preimage_zeroLocus_eq {X : Scheme.{u}} (s : Set Γ(X, ⊤)) : +namespace Scheme + +open ConcreteCategory + +variable (X : Scheme.{u}) + +/-- On a scheme `X`, the preimage of the zero locus of the prime spectrum +of `Γ(X, ⊤)` under `X.toSpecΓ : X ⟶ Spec Γ(X, ⊤)` agrees with the associated zero locus on `X`. -/ +lemma toSpecΓ_preimage_zeroLocus (s : Set Γ(X, ⊤)) : X.toSpecΓ.base ⁻¹' PrimeSpectrum.zeroLocus s = X.zeroLocus s := LocallyRingedSpace.toΓSpec_preimage_zeroLocus_eq s -open ConcreteCategory +@[deprecated (since := "2025-01-17")] alias toΓSpec_preimage_zeroLocus_eq := + toSpecΓ_preimage_zeroLocus -/-- If `X` is affine, the image of the zero locus of global sections of `X` under `toΓSpecFun` +/-- If `X` is affine, the image of the zero locus of global sections of `X` under `X.isoSpec` is the zero locus in terms of the prime spectrum of `Γ(X, ⊤)`. -/ -lemma Scheme.toΓSpec_image_zeroLocus_eq_of_isAffine {X : Scheme.{u}} [IsAffine X] +lemma isoSpec_image_zeroLocus [IsAffine X] (s : Set Γ(X, ⊤)) : X.isoSpec.hom.base '' X.zeroLocus s = PrimeSpectrum.zeroLocus s := by - erw [← X.toΓSpec_preimage_zeroLocus_eq, Set.image_preimage_eq] + erw [← X.toSpecΓ_preimage_zeroLocus, Set.image_preimage_eq] exact (bijective_of_isIso X.isoSpec.hom.base).surjective +lemma toSpecΓ_image_zeroLocus [IsAffine X] (s : Set Γ(X, ⊤)) : + X.toSpecΓ.base '' X.zeroLocus s = PrimeSpectrum.zeroLocus s := + X.isoSpec_image_zeroLocus _ + +lemma isoSpec_inv_preimage_zeroLocus [IsAffine X] (s : Set Γ(X, ⊤)) : + X.isoSpec.inv.base ⁻¹' X.zeroLocus s = PrimeSpectrum.zeroLocus s := by + rw [← toSpecΓ_preimage_zeroLocus, ← Set.preimage_comp, ← TopCat.coe_comp, ← Scheme.comp_base, + X.isoSpec_inv_toSpecΓ] + rfl + +lemma isoSpec_inv_image_zeroLocus [IsAffine X] (s : Set Γ(X, ⊤)) : + X.isoSpec.inv.base '' PrimeSpectrum.zeroLocus s = X.zeroLocus s := by + rw [← isoSpec_inv_preimage_zeroLocus, Set.image_preimage_eq] + exact (bijective_of_isIso X.isoSpec.inv.base).surjective + +@[deprecated (since := "2025-01-17")] alias toΓSpec_image_zeroLocus_eq_of_isAffine := + Scheme.isoSpec_image_zeroLocus + /-- If `X` is an affine scheme, every closed set of `X` is the zero locus of a set of global sections. -/ -lemma Scheme.eq_zeroLocus_of_isClosed_of_isAffine (X : Scheme.{u}) [IsAffine X] (s : Set X) : +lemma eq_zeroLocus_of_isClosed_of_isAffine [IsAffine X] (s : Set X) : IsClosed s ↔ ∃ I : Ideal (Γ(X, ⊤)), s = X.zeroLocus (I : Set Γ(X, ⊤)) := by refine ⟨fun hs ↦ ?_, ?_⟩ · let Z : Set (Spec <| Γ(X, ⊤)) := X.toΓSpecFun '' s have hZ : IsClosed Z := (X.isoSpec.hom.homeomorph).isClosedMap _ hs obtain ⟨I, (hI : Z = _)⟩ := (PrimeSpectrum.isClosed_iff_zeroLocus_ideal _).mp hZ use I - simp only [← Scheme.toΓSpec_preimage_zeroLocus_eq, ← hI, Z] + simp only [← Scheme.toSpecΓ_preimage_zeroLocus, ← hI, Z] erw [Set.preimage_image_eq _ (bijective_of_isIso X.isoSpec.hom.base).injective] · rintro ⟨I, rfl⟩ exact zeroLocus_isClosed X I.carrier +open Set.Notation in +lemma Opens.toSpecΓ_preimage_zeroLocus {X : Scheme.{u}} (U : X.Opens) + (s : Set Γ(X, U)) : + U.toSpecΓ.base ⁻¹' PrimeSpectrum.zeroLocus s = U.1 ↓∩ X.zeroLocus s := by + rw [toSpecΓ, Scheme.comp_base, TopCat.coe_comp, Set.preimage_comp, Spec.map_base] + erw [PrimeSpectrum.preimage_comap_zeroLocus] + rw [Scheme.toSpecΓ_preimage_zeroLocus] + show _ = U.ι.base ⁻¹' (X.zeroLocus s) + rw [Scheme.preimage_zeroLocus, U.ι_app_self, ← zeroLocus_map_of_eq _ U.ι_preimage_self, + ← Set.image_comp, ← RingHom.coe_comp, ← CommRingCat.hom_comp] + congr! + simp [← Functor.map_comp] + rfl + +end Scheme + +lemma IsAffineOpen.fromSpec_preimage_zeroLocus {X : Scheme.{u}} {U : X.Opens} + (hU : IsAffineOpen U) (s : Set Γ(X, U)) : + hU.fromSpec.base ⁻¹' X.zeroLocus s = PrimeSpectrum.zeroLocus s := by + ext x + suffices (∀ f ∈ s, ¬¬ f ∈ x.asIdeal) ↔ s ⊆ x.asIdeal by + simpa [← hU.fromSpec_image_basicOpen, -not_not] using this + simp_rw [not_not] + rfl + +lemma IsAffineOpen.fromSpec_image_zeroLocus {X : Scheme.{u}} {U : X.Opens} + (hU : IsAffineOpen U) (s : Set Γ(X, U)) : + hU.fromSpec.base '' PrimeSpectrum.zeroLocus s = X.zeroLocus s ∩ U := by + rw [← hU.fromSpec_preimage_zeroLocus, Set.image_preimage_eq_inter_range, range_fromSpec] + +open Set.Notation in +lemma Scheme.zeroLocus_inf (X : Scheme.{u}) {U : X.Opens} (I J : Ideal Γ(X, U)) : + X.zeroLocus (U := U) ↑(I ⊓ J) = X.zeroLocus (U := U) I ∪ X.zeroLocus (U := U) J := by + suffices U.1 ↓∩ (X.zeroLocus (U := U) ↑(I ⊓ J)) = + U.1 ↓∩ (X.zeroLocus (U := U) I ∪ X.zeroLocus (U := U) J) by + ext x + by_cases hxU : x ∈ U + · simpa [hxU] using congr(⟨x, hxU⟩ ∈ $this) + · simp only [Submodule.inf_coe, Set.mem_union, + codisjoint_iff_compl_le_left.mp (X.codisjoint_zeroLocus (U := U) (I ∩ J)) hxU, + codisjoint_iff_compl_le_left.mp (X.codisjoint_zeroLocus (U := U) I) hxU, true_or] + simp only [← U.toSpecΓ_preimage_zeroLocus, PrimeSpectrum.zeroLocus_inf I J, + Set.preimage_union] + +lemma Scheme.zeroLocus_biInf + {X : Scheme.{u}} {U : X.Opens} {ι : Type*} + (I : ι → Ideal Γ(X, U)) {t : Set ι} (ht : t.Finite) : + X.zeroLocus (U := U) ↑(⨅ i ∈ t, I i) = (⋃ i ∈ t, X.zeroLocus (U := U) (I i)) ∪ (↑U)ᶜ := by + refine ht.induction_on _ (by simp) fun {i t} hit ht IH ↦ ?_ + simp only [Set.mem_insert_iff, Set.iUnion_iUnion_eq_or_left, ← IH, ← zeroLocus_inf, + Submodule.inf_coe, Set.union_assoc] + congr! + simp + +lemma Scheme.zeroLocus_biInf_of_nonempty + {X : Scheme.{u}} {U : X.Opens} {ι : Type*} + (I : ι → Ideal Γ(X, U)) {t : Set ι} (ht : t.Finite) (ht' : t.Nonempty) : + X.zeroLocus (U := U) ↑(⨅ i ∈ t, I i) = ⋃ i ∈ t, X.zeroLocus (U := U) (I i) := by + rw [zeroLocus_biInf I ht, Set.union_eq_left] + obtain ⟨i, hi⟩ := ht' + exact fun x hx ↦ Set.mem_iUnion₂_of_mem hi + (codisjoint_iff_compl_le_left.mp (X.codisjoint_zeroLocus (U := U) (I i)) hx) + +lemma Scheme.zeroLocus_iInf + {X : Scheme.{u}} {U : X.Opens} {ι : Type*} + (I : ι → Ideal Γ(X, U)) [Finite ι] : + X.zeroLocus (U := U) ↑(⨅ i, I i) = (⋃ i, X.zeroLocus (U := U) (I i)) ∪ (↑U)ᶜ := by + simpa using zeroLocus_biInf I Set.finite_univ + +lemma Scheme.zeroLocus_iInf_of_nonempty + {X : Scheme.{u}} {U : X.Opens} {ι : Type*} + (I : ι → Ideal Γ(X, U)) [Finite ι] [Nonempty ι] : + X.zeroLocus (U := U) ↑(⨅ i, I i) = ⋃ i, X.zeroLocus (U := U) (I i) := by + simpa using zeroLocus_biInf_of_nonempty I Set.finite_univ + end ZeroLocus section Factorization diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 20441b03528a1..b8779620b2064 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.Algebra.MvPolynomial.Monad -import Mathlib.AlgebraicGeometry.Limits +import Mathlib.AlgebraicGeometry.Morphisms.Finite +import Mathlib.AlgebraicGeometry.Morphisms.FinitePresentation +import Mathlib.RingTheory.Spectrum.Prime.Polynomial +import Mathlib.AlgebraicGeometry.PullbackCarrier /-! # Affine space @@ -32,12 +35,12 @@ universe v u variable (n : Type v) (S : Scheme.{max u v}) local notation3 "ℤ[" n "]" => CommRingCat.of (MvPolynomial n (ULift ℤ)) -local notation3 "ℤ[" n "].{" u "}" => CommRingCat.of (MvPolynomial n (ULift.{u} ℤ)) +local notation3 "ℤ[" n "].{" u "," v "}" => CommRingCat.of (MvPolynomial n (ULift.{max u v} ℤ)) /-- `𝔸(n; S)` is the affine `n`-space over `S`. Note that `n` is an arbitrary index type (e.g. `Fin m`). -/ def AffineSpace (n : Type v) (S : Scheme.{max u v}) : Scheme.{max u v} := - pullback (terminal.from S) (terminal.from (Spec ℤ[n])) + pullback (terminal.from S) (terminal.from (Spec ℤ[n].{u, v})) namespace AffineSpace @@ -62,7 +65,7 @@ instance over : 𝔸(n; S).CanonicallyOver S where hom := pullback.fst _ _ /-- The map from the affine `n`-space over `S` to the integral model `Spec ℤ[n]`. -/ -def toSpecMvPoly : 𝔸(n; S) ⟶ Spec ℤ[n] := pullback.snd _ _ +def toSpecMvPoly : 𝔸(n; S) ⟶ Spec ℤ[n].{u, v} := pullback.snd _ _ variable {X : Scheme.{max u v}} @@ -86,7 +89,7 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where Iso.cancel_iso_inv_left, ← Iso.eq_inv_comp] apply of_mvPolynomial_int_ext intro i - rw [coe_eval₂Hom, eval₂_X] + rw [ConcreteCategory.hom_ofHom, coe_eval₂Hom, eval₂_X] rfl right_inv v := by ext i @@ -200,7 +203,7 @@ def isoOfIsAffine [IsAffine S] : ext : 1 apply ringHom_ext' · show _ = (CommRingCat.ofHom C ≫ _).hom - rw [CommRingCat.hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, + rw [CommRingCat.hom_comp, RingHom.comp_assoc, CommRingCat.hom_ofHom, eval₂Hom_comp_C, ← CommRingCat.hom_comp, ← CommRingCat.hom_ext_iff, ← cancel_mono (Scheme.ΓSpecIso _).hom] rw [← Scheme.comp_appTop, homOfVector_over, Scheme.comp_appTop] @@ -209,7 +212,7 @@ def isoOfIsAffine [IsAffine S] : rw [← Scheme.comp_appTop_assoc, Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] simp · intro i - rw [CommRingCat.comp_apply, coe_eval₂Hom] + rw [CommRingCat.comp_apply, ConcreteCategory.hom_ofHom, coe_eval₂Hom] simp only [eval₂_X] exact homOfVector_appTop_coord _ _ _ @@ -293,6 +296,13 @@ lemma map_appTop_coord {S T : Scheme.{max u v}} (f : S ⟶ T) (i) : (map n f).appTop (coord T i) = coord S i := homOfVector_appTop_coord _ _ _ +@[reassoc (attr := simp)] +lemma map_toSpecMvPoly {S T : Scheme.{max u v}} (f : S ⟶ T) : + map n f ≫ toSpecMvPoly n T = toSpecMvPoly n S := by + apply (toSpecMvPolyIntEquiv _).injective + ext i + rw [toSpecMvPolyIntEquiv_comp, ← coord, map_appTop_coord, coord] + @[simp] lemma map_id : map n (𝟙 S) = 𝟙 𝔸(n; S) := by ext1 <;> simp @@ -318,14 +328,21 @@ lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : conv_lhs => enter[2]; tactic => exact map_appTop_coord _ _ conv_rhs => enter[2]; tactic => exact SpecIso_inv_appTop_coord _ _ rw [SpecIso_inv_appTop_coord, ← CommRingCat.comp_apply, ← Scheme.ΓSpecIso_inv_naturality, - CommRingCat.comp_apply, map_X] + CommRingCat.comp_apply, ConcreteCategory.hom_ofHom, map_X] /-- The map between affine spaces over affine bases is -isomorphic to the natural map between polynomial rings. -/ +isomorphic to the natural map between polynomial rings. -/ def mapSpecMap {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : Arrow.mk (map n (Spec.map φ)) ≅ Arrow.mk (Spec.map (CommRingCat.ofHom (MvPolynomial.map (σ := n) φ.hom))) := - Arrow.isoMk (SpecIso n S) (SpecIso n R) (by simp [map_Spec_map]) + Arrow.isoMk (SpecIso n S) (SpecIso n R) (by have := (SpecIso n R).inv_hom_id; simp [map_Spec_map]) + +lemma isPullback_map {S T : Scheme.{max u v}} (f : S ⟶ T) : + IsPullback (map n f) (𝔸(n; S) ↘ S) (𝔸(n; T) ↘ T) f := by + refine (IsPullback.paste_horiz_iff (.flip <| .of_hasPullback _ _) (map_over f)).mp ?_ + simp only [terminal.comp_from, ] + convert (IsPullback.of_hasPullback _ _).flip + rw [← toSpecMvPoly, ← toSpecMvPoly, map_toSpecMvPoly] /-- `𝔸(n; S)` is functorial wrt `n`. -/ def reindex {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : 𝔸(n; S) ⟶ 𝔸(m; S) := @@ -369,6 +386,82 @@ def functor : (Type v)ᵒᵖ ⥤ Scheme.{max u v} ⥤ Scheme.{max u v} where map_comp f g := by ext: 2; dsimp; exact reindex_comp _ _ _ end functorial +section instances + +instance : IsAffineHom (𝔸(n; S) ↘ S) := MorphismProperty.pullback_fst _ _ inferInstance + +instance : Surjective (𝔸(n; S) ↘ S) := MorphismProperty.pullback_fst _ _ <| by + have := isIso_of_isTerminal specULiftZIsTerminal terminalIsTerminal (terminal.from _) + rw [← terminal.comp_from (Spec.map (CommRingCat.ofHom C)), + MorphismProperty.cancel_right_of_respectsIso (P := @Surjective)] + exact ⟨MvPolynomial.comap_C_surjective⟩ + +instance [Finite n] : LocallyOfFinitePresentation (𝔸(n; S) ↘ S) := + MorphismProperty.pullback_fst _ _ <| by + have := isIso_of_isTerminal specULiftZIsTerminal.{max u v} terminalIsTerminal (terminal.from _) + rw [← terminal.comp_from (Spec.map (CommRingCat.ofHom C)), + MorphismProperty.cancel_right_of_respectsIso (P := @LocallyOfFinitePresentation), + HasRingHomProperty.Spec_iff (P := @LocallyOfFinitePresentation), RingHom.FinitePresentation] + convert (inferInstanceAs (Algebra.FinitePresentation (ULift ℤ) ℤ[n])) + exact Algebra.algebra_ext _ _ fun _ ↦ rfl + +lemma isOpenMap_over : IsOpenMap (𝔸(n; S) ↘ S).base := by + show topologically @IsOpenMap _ + wlog hS : ∃ R, S = Spec R + · refine (IsLocalAtTarget.iff_of_openCover (P := topologically @IsOpenMap) S.affineCover).mpr ?_ + intro i + have := this (n := n) (S.affineCover.obj i) ⟨_, rfl⟩ + rwa [← (isPullback_map (n := n) (S.affineCover.map i)).isoPullback_hom_snd, + MorphismProperty.cancel_left_of_respectsIso (P := topologically @IsOpenMap)] at this + obtain ⟨R, rfl⟩ := hS + rw [← MorphismProperty.cancel_left_of_respectsIso (P := topologically @IsOpenMap) + (SpecIso n R).inv, SpecIso_inv_over] + exact MvPolynomial.isOpenMap_comap_C + +open MorphismProperty in +instance [IsEmpty n] : IsIso (𝔸(n; S) ↘ S) := pullback_fst + (P := isomorphisms _) _ _ <| by + rw [← terminal.comp_from (Spec.map (CommRingCat.ofHom C))] + apply IsStableUnderComposition.comp_mem + · rw [HasAffineProperty.iff_of_isAffine (P := isomorphisms _), ← isomorphisms, + ← arrow_mk_iso_iff (isomorphisms _) (arrowIsoΓSpecOfIsAffine _)] + exact ⟨inferInstance, (ConcreteCategory.isIso_iff_bijective _).mpr + ⟨C_injective n _, C_surjective _⟩⟩ + · exact isIso_of_isTerminal specULiftZIsTerminal terminalIsTerminal (terminal.from _) + +lemma isIntegralHom_over_iff_isEmpty : IsIntegralHom (𝔸(n; S) ↘ S) ↔ IsEmpty S ∨ IsEmpty n := by + constructor + · intro h + cases isEmpty_or_nonempty S + · exact .inl ‹_› + refine .inr ?_ + wlog hS : ∃ R, S = Spec R + · obtain ⟨x⟩ := ‹Nonempty S› + obtain ⟨y, hy⟩ := S.affineCover.covers x + exact this (S.affineCover.obj x) (MorphismProperty.IsStableUnderBaseChange.of_isPullback + (isPullback_map (S.affineCover.map x)) h) ⟨y⟩ ⟨_, rfl⟩ + obtain ⟨R, rfl⟩ := hS + have : Nontrivial R := (subsingleton_or_nontrivial R).resolve_left fun H ↦ + not_isEmpty_of_nonempty (Spec R) (inferInstanceAs (IsEmpty (PrimeSpectrum R))) + constructor + intro i + have := RingHom.toMorphismProperty_respectsIso_iff.mp RingHom.isIntegral_respectsIso.{max u v} + rw [← MorphismProperty.cancel_left_of_respectsIso @IsIntegralHom (SpecIso n R).inv, + SpecIso_inv_over, HasAffineProperty.iff_of_isAffine (P := @IsIntegralHom)] at h + obtain ⟨p : Polynomial R, hp, hp'⟩ := + (MorphismProperty.arrow_mk_iso_iff (RingHom.toMorphismProperty RingHom.IsIntegral) + (arrowIsoΓSpecOfIsAffine _)).mpr h.2 (X i) + have : (rename fun _ ↦ i).comp (pUnitAlgEquiv.{_, v} _).symm.toAlgHom p = 0 := by + simp [← hp', ← algebraMap_eq] + rw [AlgHom.comp_apply, map_eq_zero_iff _ (rename_injective _ (fun _ _ _ ↦ rfl))] at this + simp only [AlgEquiv.toAlgHom_eq_coe, AlgHom.coe_coe, EmbeddingLike.map_eq_zero_iff] at this + simp [this] at hp + · rintro (_ | _) <;> infer_instance + +lemma not_isIntegralHom [Nonempty S] [Nonempty n] : ¬ IsIntegralHom (𝔸(n; S) ↘ S) := by + simp [isIntegralHom_over_iff_isEmpty] + +end instances end AffineSpace diff --git a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean index 87849e46136ca..042468ac880d9 100644 --- a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean @@ -113,7 +113,7 @@ def Cover.bind [P.IsStableUnderComposition] (f : ∀ x : 𝒰.J, (𝒰.obj x).Co rcases (f (𝒰.f x)).covers y with ⟨z, hz⟩ change x ∈ Set.range ((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).base use z - erw [comp_apply] + erw [CategoryTheory.comp_apply] rw [hz, hy] map_prop _ := P.comp_mem _ _ ((f _).map_prop _) (𝒰.map_prop _) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index d5026ebed2afd..45e995affc9dd 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -401,11 +401,11 @@ lemma Φ_ofNat (n : ℕ) : W.Φ (n + 1) = @[simp] lemma Φ_zero : W.Φ 0 = 1 := by rw [Φ, ΨSq_zero, mul_zero, zero_sub, zero_add, preΨ_one, one_mul, zero_sub, preΨ_neg, preΨ_one, - neg_one_mul, neg_neg, if_pos even_zero] + neg_one_mul, neg_neg, if_pos Even.zero] @[simp] lemma Φ_one : W.Φ 1 = X := by - rw [show 1 = ((0 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_one, one_pow, mul_one, if_pos even_zero, + rw [show 1 = ((0 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_one, one_pow, mul_one, if_pos Even.zero, mul_one, preΨ'_zero, mul_zero, zero_mul, sub_zero] @[simp] @@ -423,8 +423,8 @@ lemma Φ_three : W.Φ 3 = X * W.Ψ₃ ^ 2 - W.preΨ₄ * W.Ψ₂Sq := by @[simp] lemma Φ_four : W.Φ 4 = X * W.preΨ₄ ^ 2 * W.Ψ₂Sq - W.Ψ₃ * (W.preΨ₄ * W.Ψ₂Sq ^ 2 - W.Ψ₃ ^ 3) := by rw [show 4 = ((3 : ℕ) + 1 : ℤ) by rfl, Φ_ofNat, preΨ'_four, if_neg <| by decide, - show 3 + 2 = 2 * 2 + 1 by rfl, preΨ'_odd, preΨ'_four, preΨ'_two, if_pos even_zero, preΨ'_one, - preΨ'_three, if_pos even_zero, if_neg <| by decide] + show 3 + 2 = 2 * 2 + 1 by rfl, preΨ'_odd, preΨ'_four, preΨ'_two, if_pos Even.zero, preΨ'_one, + preΨ'_three, if_pos Even.zero, if_neg <| by decide] ring1 @[simp] diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 82314906f6c4c..2506ebe79df0c 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -88,14 +88,10 @@ theorem toΓSpec_continuous : Continuous X.toΓSpecFun := by /-- The canonical (bundled) continuous map from the underlying topological space of `X` to the prime spectrum of its global sections. -/ -@[simps] def toΓSpecBase : X.toTopCat ⟶ Spec.topObj (Γ.obj (op X)) where toFun := X.toΓSpecFun continuous_toFun := X.toΓSpec_continuous --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.LocallyRingedSpace.toΓSpecBase_apply - variable (r : Γ.obj (op X)) /-- The preimage in `X` of a basic open in `Spec Γ(X)` (as an open set). -/ @@ -143,7 +139,7 @@ theorem toΓSpecCApp_iff -- Porting Note: Type class problem got stuck in `IsLocalization.Away.AwayMap.lift_comp` -- created instance manually. This replaces the `pick_goal` tactics have loc_inst := IsLocalization.to_basicOpen (Γ.obj (op X)) r - refine CommRingCat.hom_ext_iff.trans ?_ + refine ConcreteCategory.ext_iff.trans ?_ rw [← @IsLocalization.Away.lift_comp _ _ _ _ _ _ _ r loc_inst _ (X.isUnit_res_toΓSpecMapBasicOpen r)] --pick_goal 5; exact is_localization.to_basic_open _ r @@ -424,7 +420,7 @@ theorem Scheme.toSpecΓ_base (X : Scheme.{u}) (x) : (Scheme.toSpecΓ X).base x = (Spec.map (X.presheaf.germ ⊤ x trivial)).base (IsLocalRing.closedPoint _) := rfl -@[reassoc (attr := simp)] +@[reassoc] theorem Scheme.toSpecΓ_naturality {X Y : Scheme.{u}} (f : X ⟶ Y) : f ≫ Y.toSpecΓ = X.toSpecΓ ≫ Spec.map (f.appTop) := ΓSpec.adjunction.unit.naturality f diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index b97e3bf450dcf..2d70705beae89 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -79,7 +79,6 @@ such that We can then glue the schemes `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subschemes of the glued space. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): @[nolint has_nonempty_instance]; linter not ported yet structure GlueData extends CategoryTheory.GlueData Scheme where f_open : ∀ i j, IsOpenImmersion (f i j) @@ -228,7 +227,7 @@ theorem ι_eq_iff (i j : D.J) (x : (D.U i).carrier) (y : (D.U j).carrier) : (TopCat.GlueData.ι_eq_iff_rel D.toLocallyRingedSpaceGlueData.toSheafedSpaceGlueData.toPresheafedSpaceGlueData.toTopGlueData i j x y) - rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff, ← comp_apply] + rw [← ((TopCat.mono_iff_injective D.isoCarrier.inv).mp _).eq_iff, ← CategoryTheory.comp_apply] · simp_rw [← D.ι_isoCarrier_inv] rfl -- `rfl` was not needed before https://github.com/leanprover-community/mathlib4/pull/13170 · infer_instance @@ -340,7 +339,7 @@ theorem fromGlued_injective : Function.Injective 𝒰.fromGlued.base := by intro x y h obtain ⟨i, x, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective x obtain ⟨j, y, rfl⟩ := 𝒰.gluedCover.ι_jointly_surjective y - rw [← comp_apply, ← comp_apply] at h + rw [← CategoryTheory.comp_apply, ← CategoryTheory.comp_apply] at h simp_rw [← Scheme.comp_base] at h rw [ι_fromGlued, ι_fromGlued] at h let e := @@ -390,7 +389,7 @@ instance : Epi 𝒰.fromGlued.base := by intro x obtain ⟨y, h⟩ := 𝒰.covers x use (𝒰.gluedCover.ι (𝒰.f x)).base y - rw [← comp_apply] + rw [← CategoryTheory.comp_apply] rw [← 𝒰.ι_fromGlued (𝒰.f x)] at h exact h diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 6ca38222835bb..6983cda062f3b 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -38,6 +38,11 @@ noncomputable def specZIsTerminal : IsTerminal (Spec (CommRingCat.of ℤ)) := @IsTerminal.isTerminalObj _ _ _ _ Scheme.Spec _ inferInstance (terminalOpOfInitial CommRingCat.zIsInitial) +/-- `Spec ℤ` is the terminal object in the category of schemes. -/ +noncomputable def specULiftZIsTerminal : IsTerminal (Spec (.of (ULift.{u} ℤ))) := + @IsTerminal.isTerminalObj _ _ _ _ Scheme.Spec _ inferInstance + (terminalOpOfInitial CommRingCat.isInitial) + instance : HasTerminal Scheme := hasTerminal_of_hasTerminal_of_preservesLimit Scheme.Spec diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index aa13874400ecd..a8e8fd231a3dd 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -143,7 +143,7 @@ See also `Tilde.isLocallyFraction`. def tildeInType : Sheaf (Type u) (PrimeSpectrum.Top R) := subsheafToTypes (Tilde.isLocallyFraction M) -instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) : +noncomputable instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) : AddCommGroup (M.tildeInType.1.obj U) := inferInstanceAs <| AddCommGroup (Tilde.sectionsSubmodule M U) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean index 54716daaae092..63edc5c703b73 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean @@ -56,7 +56,7 @@ lemma affineAnd_respectsIso (hP : RingHom.RespectsIso Q) : simpa [AffineTargetMorphismProperty.toProperty, isAffine_of_isIso e.inv, hP.cancel_left_isIso] /-- `affineAnd P` is local if `P` is local on the (algebraic) source. -/ -lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.LocalizationPreserves Q) +lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.LocalizationAwayPreserves Q) (hQs : RingHom.OfLocalizationSpan Q) : (affineAnd Q).IsLocal where respectsIso := affineAnd_respectsIso hPi to_basicOpen {X Y _} f r := fun ⟨hX, hf⟩ ↦ by @@ -96,6 +96,10 @@ lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.Localizatio rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map _ (isAffineOpen_top X)] at hf rwa [CommRingCat.hom_comp, hPi.cancel_right_isIso] at hf +lemma affineAnd_isLocal_of_propertyIsLocal + (hPi : RingHom.PropertyIsLocal Q) : (affineAnd Q).IsLocal := + affineAnd_isLocal hPi.respectsIso hPi.localizationAwayPreserves hPi.ofLocalizationSpan + /-- If `P` is stable under base change, so is `affineAnd P`. -/ lemma affineAnd_isStableUnderBaseChange (hQi : RingHom.RespectsIso Q) (hQb : RingHom.IsStableUnderBaseChange Q) : @@ -219,7 +223,7 @@ lemma HasAffineProperty.affineAnd_containsIdentities {P : MorphismProperty Schem /-- A convenience constructor for `HasAffineProperty P (affineAnd Q)`. The `IsAffineHom` is bundled, since this goes well with defining morphism properties via `extends IsAffineHom`. -/ lemma HasAffineProperty.affineAnd_iff (P : MorphismProperty Scheme.{u}) - (hQi : RingHom.RespectsIso Q) (hQl : RingHom.LocalizationPreserves Q) + (hQi : RingHom.RespectsIso Q) (hQl : RingHom.LocalizationAwayPreserves Q) (hQs : RingHom.OfLocalizationSpan Q) : HasAffineProperty P (affineAnd Q) ↔ ∀ {X Y : Scheme.{u}} (f : X ⟶ Y), P f ↔ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean index dcef9ad69d7b4..c674ed3fe7590 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean @@ -41,7 +41,7 @@ instance : HasAffineProperty @IsFinite (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.appTop).hom) := by show HasAffineProperty @IsFinite (affineAnd RingHom.Finite) rw [HasAffineProperty.affineAnd_iff _ RingHom.finite_respectsIso - RingHom.finite_localizationPreserves RingHom.finite_ofLocalizationSpan] + RingHom.finite_localizationPreserves.away RingHom.finite_ofLocalizationSpan] simp [isFinite_iff] instance : IsStableUnderComposition @IsFinite := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean b/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean index 4c3998ab096ac..dc98b92581f92 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/FinitePresentation.lean @@ -59,7 +59,7 @@ instance locallyOfFinitePresentation_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g LocallyOfFinitePresentation (f ≫ g) := MorphismProperty.comp_mem _ f g hf hg -lemma locallyOfFinitePresentation_isStableUnderBaseChange : +instance locallyOfFinitePresentation_isStableUnderBaseChange : MorphismProperty.IsStableUnderBaseChange @LocallyOfFinitePresentation := HasRingHomProperty.isStableUnderBaseChange RingHom.finitePresentation_isStableUnderBaseChange diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FormallyUnramified.lean b/Mathlib/AlgebraicGeometry/Morphisms/FormallyUnramified.lean new file mode 100644 index 0000000000000..4182443309f6f --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/FormallyUnramified.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2025 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties +import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion +import Mathlib.RingTheory.RingHom.Unramified + +/-! +# Formally unramified morphisms + +A morphism of schemes `f : X ⟶ Y` is formally unramified if for each affine `U ⊆ Y` and +`V ⊆ f ⁻¹' U`, the induced map `Γ(Y, U) ⟶ Γ(X, V)` is formally unramified. + +We show that these properties are local, and are stable under compositions and base change. + +-/ + + +noncomputable section + +open CategoryTheory CategoryTheory.Limits Opposite TopologicalSpace + +universe v u + +namespace AlgebraicGeometry + +variable {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- A morphism of schemes `f : X ⟶ Y` is formally unramified if for each affine `U ⊆ Y` and +`V ⊆ f ⁻¹' U`, The induced map `Γ(Y, U) ⟶ Γ(X, V)` is formally unramified. -/ +@[mk_iff] +class FormallyUnramified (f : X ⟶ Y) : Prop where + formallyUnramified_of_affine_subset : + ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), + (f.appLE U V e).hom.FormallyUnramified + +namespace FormallyUnramified + +instance : HasRingHomProperty @FormallyUnramified RingHom.FormallyUnramified where + isLocal_ringHomProperty := RingHom.FormallyUnramified.propertyIsLocal + eq_affineLocally' := by + ext X Y f + rw [formallyUnramified_iff, affineLocally_iff_affineOpens_le] + +instance : MorphismProperty.IsStableUnderComposition @FormallyUnramified := + HasRingHomProperty.stableUnderComposition RingHom.FormallyUnramified.stableUnderComposition + +/-- `f : X ⟶ S` is formally unramified if `X ⟶ X ×ₛ X` is an open immersion. +In particular, monomorphisms (e.g. immersions) are formally unramified. +The converse is true if `f` is locally of finite type. -/ +instance (priority := 900) [IsOpenImmersion (pullback.diagonal f)] : FormallyUnramified f := by + wlog hY : ∃ R, Y = Spec R + · rw [IsLocalAtTarget.iff_of_openCover (P := @FormallyUnramified) Y.affineCover] + intro i + have inst : IsOpenImmersion (pullback.diagonal (pullback.snd f (Y.affineCover.map i))) := + MorphismProperty.pullback_snd (P := .diagonal @IsOpenImmersion) _ _ ‹_› + exact this (pullback.snd _ _) ⟨_, rfl⟩ + obtain ⟨R, rfl⟩ := hY + wlog hX : ∃ S, X = Spec S generalizing X + · rw [IsLocalAtSource.iff_of_openCover (P := @FormallyUnramified) X.affineCover] + intro i + have inst : IsOpenImmersion (pullback.diagonal (X.affineCover.map i ≫ f)) := + MorphismProperty.comp_mem (.diagonal @IsOpenImmersion) _ _ + (inferInstanceAs (IsOpenImmersion _)) ‹_› + exact this (_ ≫ _) ⟨_, rfl⟩ + obtain ⟨S, rfl⟩ := hX + obtain ⟨φ, rfl : Spec.map φ = f⟩ := Spec.homEquiv.symm.surjective f + rw [HasRingHomProperty.Spec_iff (P := @FormallyUnramified)] + algebraize [φ.hom] + let F := (Algebra.TensorProduct.lmul' R (S := S)).toRingHom + have hF : Function.Surjective F := fun x ↦ ⟨.mk _ _ _ x 1, by simp [F]⟩ + have : IsOpenImmersion (Spec.map (CommRingCat.ofHom F)) := by + rwa [← MorphismProperty.cancel_right_of_respectsIso (P := @IsOpenImmersion) _ + (pullbackSpecIso R S S).inv, ← AlgebraicGeometry.diagonal_Spec_map R S] + obtain ⟨e, he, he'⟩ := (isOpenImmersion_SpecMap_iff_of_surjective _ hF).mp this + refine ⟨subsingleton_of_forall_eq 0 fun x ↦ ?_⟩ + obtain ⟨⟨x, hx⟩, rfl⟩ := Ideal.toCotangent_surjective _ x + obtain ⟨x, rfl⟩ := Ideal.mem_span_singleton.mp (he'.le hx) + refine (Ideal.toCotangent_eq_zero _ _).mpr ?_ + rw [pow_two, Subtype.coe_mk, ← he, mul_assoc] + exact Ideal.mul_mem_mul (he'.ge (Ideal.mem_span_singleton_self e)) hx + +theorem of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) + [FormallyUnramified (f ≫ g)] : FormallyUnramified f := + HasRingHomProperty.of_comp (fun {R S T _ _ _} f g H ↦ by + algebraize [f, g, g.comp f] + exact Algebra.FormallyUnramified.of_comp R S T) ‹_› + +instance : MorphismProperty.IsMultiplicative @FormallyUnramified where + id_mem _ := inferInstance + +instance : MorphismProperty.IsStableUnderBaseChange @FormallyUnramified := + HasRingHomProperty.isStableUnderBaseChange RingHom.FormallyUnramified.isStableUnderBaseChange + +end FormallyUnramified + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean index c6a2c5393bdfa..72aed7577ed99 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean @@ -42,7 +42,7 @@ instance hasAffineProperty : HasAffineProperty @IsIntegralHom fun X _ f _ ↦ IsAffine X ∧ RingHom.IsIntegral (f.app ⊤).hom := by show HasAffineProperty @IsIntegralHom (affineAnd RingHom.IsIntegral) rw [HasAffineProperty.affineAnd_iff _ RingHom.isIntegral_respectsIso - RingHom.isIntegral_isStableUnderBaseChange.localizationPreserves + RingHom.isIntegral_isStableUnderBaseChange.localizationPreserves.away RingHom.isIntegral_ofLocalizationSpan] simp [isIntegralHom_iff] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean index 3b9e3811ca311..60ffa081d3d76 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean @@ -26,6 +26,39 @@ universe u namespace AlgebraicGeometry +/-- `Spec (R ⧸ I) ⟶ Spec R` is an open immersion iff `I` is generated by an idempotent. -/ +lemma isOpenImmersion_SpecMap_iff_of_surjective {R S : CommRingCat} + (f : R ⟶ S) (hf : Function.Surjective f.hom) : + IsOpenImmersion (Spec.map f) ↔ + ∃ e, IsIdempotentElem e ∧ RingHom.ker f.hom = Ideal.span {e} := by + constructor + · intro H + obtain ⟨e, he, he'⟩ := PrimeSpectrum.isClopen_iff_zeroLocus.mp + ⟨PrimeSpectrum.isClosed_range_comap_of_surjective _ _ hf, + (Spec.map f).isOpenEmbedding.isOpen_range⟩ + refine ⟨e, he, ?_⟩ + let φ : R ⟶ _ := (CommRingCat.ofHom (Ideal.Quotient.mk (.span {e}))) + have : IsOpenImmersion (Spec.map φ) := + have : IsLocalization.Away (1 - e) (↑R ⧸ Ideal.span {e}) := + IsLocalization.away_of_isIdempotentElem he.one_sub (by simp) Ideal.Quotient.mk_surjective + IsOpenImmersion.of_isLocalization (1 - e) + have H : Set.range (Spec.map φ).base = Set.range (Spec.map f).base := + ((PrimeSpectrum.range_comap_of_surjective _ _ + Ideal.Quotient.mk_surjective).trans (by simp [φ])).trans he'.symm + let i : S ≅ .of _ := (Scheme.Spec.preimageIso + (IsOpenImmersion.isoOfRangeEq (Spec.map φ) (Spec.map f) H)).unop + have hi : Function.Injective i.inv.hom := (ConcreteCategory.bijective_of_isIso i.inv).1 + have : f = φ ≫ i.inv := by apply Spec.map_injective; simp [i, ← Scheme.Spec_map] + rw [this, CommRingCat.hom_comp, RingHom.ker_eq_comap_bot, ← Ideal.comap_comap, + ← RingHom.ker_eq_comap_bot, (RingHom.injective_iff_ker_eq_bot i.inv.hom).mp hi, + ← RingHom.ker_eq_comap_bot] + simp [φ] + · rintro ⟨e, he, he'⟩ + letI := f.hom.toAlgebra + have : IsLocalization.Away (1 - e) S := + IsLocalization.away_of_isIdempotentElem he.one_sub (by simpa using he') hf + exact IsOpenImmersion.of_isLocalization (1 - e) + variable {X Y : Scheme.{u}} theorem isOpenImmersion_iff_stalk {f : X ⟶ Y} : IsOpenImmersion f ↔ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 9f07373d33dce..625901a76887e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -158,7 +158,7 @@ theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) simpa using hU rw [← f.appLE_congr _ rfl this (fun f => P f.hom), IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2 _ r] - simp only + simp only [CommRingCat.hom_ofHom] apply (config := { allowSynthFailures := true }) h₂ exact H U · introv hs hs' U @@ -167,7 +167,7 @@ theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) simp_rw [sourceAffineLocally_morphismRestrict] at hs' have := hs' r ⟨X.basicOpen (f.appLE ⊤ U le_top r.1), U.2.basicOpen (f.appLE ⊤ U le_top r.1)⟩ (by simp [Scheme.Hom.appLE]) - rwa [IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2, + rwa [IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2, CommRingCat.hom_ofHom, ← h₁.is_localization_away_iff] at this variable {P} diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 7d1c4fff13747..73a384e72fb39 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -191,9 +191,8 @@ lemma noetherianSpace_of_isAffineOpen (U : X.Opens) (hU : IsAffineOpen U) (Scheme.restrictFunctorΓ.app (op U)).symm.commRingCatIsoToRingEquiv exact @noetherianSpace_of_isAffine _ hU _ -/-- Any open immersion `Z ⟶ X` with `X` locally Noetherian is quasi-compact. - -[Stacks: Lemma 01OX](https://stacks.math.columbia.edu/tag/01OX) -/ +/-- Any open immersion `Z ⟶ X` with `X` locally Noetherian is quasi-compact. -/ +@[stacks 01OX] instance (priority := 100) {Z : Scheme} [IsLocallyNoetherian X] {f : Z ⟶ X} [IsOpenImmersion f] : QuasiCompact f := by apply (quasiCompact_iff_forall_affine f).mpr @@ -206,9 +205,8 @@ instance (priority := 100) {Z : Scheme} [IsLocallyNoetherian X] · exact Set.inter_subset_left · exact Set.inter_subset_right -/-- A locally Noetherian scheme is quasi-separated. - -[Stacks: Lemma 01OY](https://stacks.math.columbia.edu/tag/01OY) -/ +/-- A locally Noetherian scheme is quasi-separated. -/ +@[stacks 01OY] instance (priority := 100) IsLocallyNoetherian.quasiSeparatedSpace [IsLocallyNoetherian X] : QuasiSeparatedSpace X := by apply (quasiSeparatedSpace_iff_affine X).mpr @@ -266,9 +264,8 @@ theorem isNoetherian_iff_of_finite_affine_openCover {𝒰 : Scheme.OpenCover.{v, · exact Scheme.OpenCover.compactSpace 𝒰 open CategoryTheory in -/-- A Noetherian scheme has a Noetherian underlying topological space. - -[Stacks, Lemma 01OZ](https://stacks.math.columbia.edu/tag/01OZ) -/ +/-- A Noetherian scheme has a Noetherian underlying topological space. -/ +@[stacks 01OZ] instance (priority := 100) IsNoetherian.noetherianSpace [IsNoetherian X] : NoetherianSpace X := by apply TopologicalSpace.noetherian_univ_iff.mp @@ -284,9 +281,8 @@ instance (priority := 100) IsNoetherian.noetherianSpace [IsNoetherian X] : convert noetherianSpace_of_isAffineOpen U.1 U.2 apply IsLocallyNoetherian.component_noetherian -/-- Any morphism of schemes `f : X ⟶ Y` with `X` Noetherian is quasi-compact. - -[Stacks, Lemma 01P0](https://stacks.math.columbia.edu/tag/01P0) -/ +/-- Any morphism of schemes `f : X ⟶ Y` with `X` Noetherian is quasi-compact. -/ +@[stacks 01P0] instance (priority := 100) quasiCompact_of_noetherianSpace_source {X Y : Scheme} [NoetherianSpace X] (f : X ⟶ Y) : QuasiCompact f := ⟨fun _ _ _ => NoetherianSpace.isCompact _⟩ @@ -325,9 +321,8 @@ theorem isNoetherian_Spec {R : CommRingCat} : ⟨fun _ => inferInstance, fun _ => inferInstance⟩ -/-- A Noetherian scheme has a finite number of irreducible components. - -[Stacks, Lemma 0BA8](https://stacks.math.columbia.edu/tag/0BA8) -/ +/-- A Noetherian scheme has a finite number of irreducible components. -/ +@[stacks 0BA8] theorem finite_irreducibleComponents_of_isNoetherian [IsNoetherian X] : (irreducibleComponents X).Finite := NoetherianSpace.finite_irreducibleComponents diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 990eb295b3aa0..f78bd0243cceb 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -142,6 +142,11 @@ lemma image_iSup₂ {ι : Sort*} {κ : ι → Sort*} (s : (i : ι) → κ i → ext : 1 simp [Set.image_iUnion₂] +@[simp] +lemma map_mem_image_iff {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] + {U : X.Opens} {x : X} : f.base x ∈ f ''ᵁ U ↔ x ∈ U := + f.isOpenEmbedding.injective.mem_set_image + /-- The isomorphism `Γ(Y, f(U)) ≅ Γ(X, U)` induced by an open immersion `f : X ⟶ Y`. -/ def appIso (U) : Γ(Y, f ''ᵁ U) ≅ Γ(X, U) := (asIso <| LocallyRingedSpace.IsOpenImmersion.invApp f.toLRSHom U).symm @@ -677,8 +682,9 @@ end MorphismProperty namespace Scheme -theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] {U : X.Opens} - (r : Γ(X, U)) : +variable {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] + +theorem image_basicOpen {U : X.Opens} (r : Γ(X, U)) : f ''ᵁ X.basicOpen r = Y.basicOpen ((f.appIso U).inv r) := by have e := Scheme.preimage_basicOpen f ((f.appIso U).inv r) rw [Scheme.Hom.appIso_inv_app_apply', Scheme.basicOpen_res, inf_eq_right.mpr _] at e @@ -686,6 +692,18 @@ theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] refine Set.Subset.trans (Scheme.basicOpen_le _ _) (Set.image_subset_range _ _) · exact (X.basicOpen_le r).trans (f.preimage_image_eq _).ge +lemma image_zeroLocus {U : X.Opens} (s : Set Γ(X, U)) : + f.base '' X.zeroLocus s = + Y.zeroLocus (U := f ''ᵁ U) ((f.appIso U).inv.hom '' s) ∩ Set.range f.base := by + ext x + by_cases hx : x ∈ Set.range f.base + · obtain ⟨x, rfl⟩ := hx + simp only [f.isOpenEmbedding.injective.mem_set_image, Scheme.mem_zeroLocus_iff, + ← SetLike.mem_coe, Set.mem_inter_iff, Set.forall_mem_image, ← Scheme.image_basicOpen, + IsOpenMap.coe_functor_obj, Set.mem_range, exists_apply_eq_apply, and_true] + · simp only [Set.mem_inter_iff, hx, and_false, iff_false] + exact fun H ↦ hx (Set.image_subset_range _ _ H) + end Scheme end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean index 16327672ba49a..8a215ab44a022 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -179,7 +179,7 @@ lemma awayMap_awayToSection : obtain ⟨⟨n, a, ⟨b, hb'⟩, i, rfl : _ = b⟩, rfl⟩ := mk_surjective a simp only [homOfLE_leOfHom, CommRingCat.hom_comp, RingHom.coe_comp, Function.comp_apply] erw [ProjectiveSpectrum.Proj.awayToSection_apply] - rw [val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', + rw [CommRingCat.hom_ofHom, val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', ← Localization.mk_eq_mk'] refine Localization.mk_eq_mk_iff.mpr ?_ rw [Localization.r_iff_exists] diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean index 818f55693d133..7fc0bc4d9218a 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean @@ -174,7 +174,7 @@ instance commRingStructureSheafInTypeObj (U : (Opens (ProjectiveSpectrum.top /-- The structure presheaf, valued in `CommRing`, constructed by dressing up the `Type` valued structure presheaf. -/ -@[simps] +@[simps obj_carrier] def structurePresheafInCommRing : Presheaf CommRingCat (ProjectiveSpectrum.top 𝒜) where obj U := CommRingCat.of ((structureSheafInType 𝒜).1.obj U) map i := CommRingCat.ofHom @@ -184,10 +184,6 @@ def structurePresheafInCommRing : Presheaf CommRingCat (ProjectiveSpectrum.top map_one' := rfl map_mul' := fun _ _ => rfl } --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] - AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf.structurePresheafInCommRing_map_hom_apply - /-- Some glue, verifying that the structure presheaf valued in `CommRing` agrees with the `Type` valued structure presheaf. -/ def structurePresheafCompForget : diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean index b9a5bdf4d9ecb..55c14d9043da3 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean @@ -44,7 +44,6 @@ variable {R A : Type*} variable [CommSemiring R] [CommRing A] [Algebra R A] variable (𝒜 : ℕ → Submodule R A) [GradedAlgebra 𝒜] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The projective spectrum of a graded commutative ring is the subtype of all homogeneous ideals that are prime and do not contain the irrelevant ideal. -/ @[ext] diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index d0f1ae12a2b67..fdd3809d8e81f 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -577,9 +577,21 @@ theorem basicOpen_mul : X.basicOpen (f * g) = X.basicOpen f ⊓ X.basicOpen g := lemma basicOpen_pow {n : ℕ} (h : 0 < n) : X.basicOpen (f ^ n) = X.basicOpen f := RingedSpace.basicOpen_pow _ _ _ h +lemma basicOpen_add_le : + X.basicOpen (f + g) ≤ X.basicOpen f ⊔ X.basicOpen g := by + intro x hx + have hxU : x ∈ U := X.basicOpen_le _ hx + simp only [SetLike.mem_coe, Scheme.mem_basicOpen _ _ _ hxU, map_add, Opens.coe_sup, + Set.mem_union] at hx ⊢ + exact IsLocalRing.isUnit_or_isUnit_of_isUnit_add hx + theorem basicOpen_of_isUnit {f : Γ(X, U)} (hf : IsUnit f) : X.basicOpen f = U := RingedSpace.basicOpen_of_isUnit _ hf +@[simp] +theorem basicOpen_one : X.basicOpen (1 : Γ(X, U)) = U := + X.basicOpen_of_isUnit isUnit_one + instance algebra_section_section_basicOpen {X : Scheme} {U : X.Opens} (f : Γ(X, U)) : Algebra Γ(X, U) Γ(X, X.basicOpen f) := (X.presheaf.map (homOfLE <| X.basicOpen_le f : _ ⟶ U).op).hom.toAlgebra @@ -618,6 +630,51 @@ lemma mem_zeroLocus_iff {U : X.Opens} (s : Set Γ(X, U)) (x : X) : x ∈ X.zeroLocus s ↔ ∀ f ∈ s, x ∉ X.basicOpen f := X.toRingedSpace.mem_zeroLocus_iff s x +lemma codisjoint_zeroLocus {U : X.Opens} + (s : Set Γ(X, U)) : Codisjoint (X.zeroLocus s) U := by + have (x : X) : ∀ f ∈ s, x ∈ X.basicOpen f → x ∈ U := fun _ _ h ↦ X.basicOpen_le _ h + simpa [codisjoint_iff_le_sup, Set.ext_iff, or_iff_not_imp_left] + +lemma zeroLocus_span {U : X.Opens} (s : Set Γ(X, U)) : + X.zeroLocus (U := U) (Ideal.span s) = X.zeroLocus s := by + ext x + simp only [Scheme.mem_zeroLocus_iff, SetLike.mem_coe] + refine ⟨fun H f hfs ↦ H f (Ideal.subset_span hfs), fun H f ↦ Submodule.span_induction H ?_ ?_ ?_⟩ + · simp only [Scheme.basicOpen_zero]; exact not_false + · exact fun a b _ _ ha hb H ↦ (X.basicOpen_add_le a b H).elim ha hb + · simp +contextual + +lemma zeroLocus_map {U V : X.Opens} (i : U ≤ V) (s : Set Γ(X, V)) : + X.zeroLocus ((X.presheaf.map (homOfLE i).op).hom '' s) = X.zeroLocus s ∪ Uᶜ := by + ext x + suffices (∀ f ∈ s, x ∈ U → x ∉ X.basicOpen f) ↔ x ∈ U → (∀ f ∈ s, x ∉ X.basicOpen f) by + simpa [or_iff_not_imp_right] + conv_lhs => enter [i]; rw [forall_comm (β := x ∈ U)] -- why doesn't simp fire on this + rw [forall_comm (β := x ∈ U)] + +lemma zeroLocus_map_of_eq {U V : X.Opens} (i : U = V) (s : Set Γ(X, V)) : + X.zeroLocus ((X.presheaf.map (eqToHom i).op).hom '' s) = X.zeroLocus s := by + ext; simp + +lemma zeroLocus_mono {U : X.Opens} {s t : Set Γ(X, U)} (h : s ⊆ t) : + X.zeroLocus t ⊆ X.zeroLocus s := by + simp only [Set.subset_def, Scheme.mem_zeroLocus_iff] + exact fun x H f hf hxf ↦ H f (h hf) hxf + +lemma preimage_zeroLocus {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (s : Set Γ(Y, U)) : + f.base ⁻¹' Y.zeroLocus s = X.zeroLocus ((f.app U).hom '' s) := by + ext + simp [← Scheme.preimage_basicOpen] + rfl + +@[simp] +lemma zeroLocus_univ {U : X.Opens} : + X.zeroLocus (U := U) Set.univ = (↑U)ᶜ := by + ext x + simp only [Scheme.mem_zeroLocus_iff, Set.mem_univ, forall_const, Set.mem_compl_iff, + SetLike.mem_coe, ← not_exists, not_iff_not] + exact ⟨fun ⟨f, hf⟩ ↦ X.basicOpen_le f hf, fun _ ↦ ⟨1, by rwa [X.basicOpen_of_isUnit isUnit_one]⟩⟩ + end ZeroLocus end Scheme @@ -678,6 +735,19 @@ lemma Scheme.iso_inv_base_hom_base_apply {X Y : Scheme.{u}} (e : X ≅ Y) (y : Y show (e.inv.base ≫ e.hom.base) y = 𝟙 Y.toPresheafedSpace y simp +theorem Spec_zeroLocus_eq_zeroLocus {R : CommRingCat} (s : Set R) : + (Spec R).zeroLocus ((Scheme.ΓSpecIso R).inv '' s) = PrimeSpectrum.zeroLocus s := by + ext x + suffices (∀ a ∈ s, x ∉ PrimeSpectrum.basicOpen a) ↔ x ∈ PrimeSpectrum.zeroLocus s by simpa + simp [Spec_carrier, PrimeSpectrum.mem_zeroLocus, Set.subset_def, + PrimeSpectrum.mem_basicOpen _ x] + +@[simp] +theorem Spec_zeroLocus {R : CommRingCat} (s : Set Γ(Spec R, ⊤)) : + (Spec R).zeroLocus s = PrimeSpectrum.zeroLocus ((Scheme.ΓSpecIso R).inv ⁻¹' s) := by + convert Spec_zeroLocus_eq_zeroLocus ((Scheme.ΓSpecIso R).inv ⁻¹' s) + rw [Set.image_preimage_eq] + exact (ConcreteCategory.bijective_of_isIso (C := CommRingCat) _).2 section Stalks namespace Scheme diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 0769318a82148..26b9f817d1bf5 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -280,13 +280,9 @@ section SpecΓ open AlgebraicGeometry.LocallyRingedSpace /-- The counit morphism `R ⟶ Γ(Spec R)` given by `AlgebraicGeometry.StructureSheaf.toOpen`. -/ -@[simps!] def toSpecΓ (R : CommRingCat.{u}) : R ⟶ Γ.obj (op (Spec.toLocallyRingedSpace.obj (op R))) := StructureSheaf.toOpen R ⊤ --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.toSpecΓ_hom_apply_coe - instance isIso_toSpecΓ (R : CommRingCat.{u}) : IsIso (toSpecΓ R) := by cases R; apply StructureSheaf.isIso_to_global @@ -424,7 +420,7 @@ instance isLocalizedModule_toPushforwardStalkAlgHom : change PrimeSpectrum.basicOpen r ≤ U at hrU apply_fun (Spec.topMap (CommRingCat.ofHom (algebraMap R S)) _* (structureSheaf S).1).map (homOfLE hrU).op at e - simp only [Functor.op_map, map_zero, ← comp_apply, toOpen_res] at e + simp only [Functor.op_map, map_zero, ← CategoryTheory.comp_apply, toOpen_res] at e have : toOpen S (PrimeSpectrum.basicOpen <| algebraMap R S r) x = 0 := by refine Eq.trans ?_ e; rfl have := diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 309e345b93978..9e79bc6c7f8b9 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -223,7 +223,7 @@ open PrimeSpectrum /-- The structure presheaf, valued in `CommRingCat`, constructed by dressing up the `Type` valued structure presheaf. -/ -@[simps] +@[simps obj_carrier] def structurePresheafInCommRing : Presheaf CommRingCat (PrimeSpectrum.Top R) where obj U := CommRingCat.of ((structureSheafInType R).1.obj U) map {_ _} i := CommRingCat.ofHom @@ -233,9 +233,6 @@ def structurePresheafInCommRing : Presheaf CommRingCat (PrimeSpectrum.Top R) whe map_one' := rfl map_mul' := fun _ _ => rfl } --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.structurePresheafInCommRing_map_hom_apply - /-- Some glue, verifying that the structure presheaf valued in `CommRingCat` agrees with the `Type` valued structure presheaf. -/ @@ -929,13 +926,10 @@ instance isIso_to_global : IsIso (toOpen R ⊤) := by infer_instance /-- The ring isomorphism between the ring `R` and the global sections `Γ(X, 𝒪ₓ)`. -/ -@[simps!] +@[simps! inv] def globalSectionsIso : CommRingCat.of R ≅ (structureSheaf R).1.obj (op ⊤) := asIso (toOpen R ⊤) --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] AlgebraicGeometry.StructureSheaf.globalSectionsIso_hom_hom_apply_coe - @[simp] theorem globalSectionsIso_hom (R : CommRingCat) : (globalSectionsIso R).hom = toOpen R ⊤ := rfl @@ -953,9 +947,8 @@ theorem localizationToStalk_stalkSpecializes {R : Type*} [CommRing R] {x y : Pri StructureSheaf.localizationToStalk R x := by ext : 1 apply IsLocalization.ringHom_ext (S := Localization.AtPrime y.asIdeal) y.asIdeal.primeCompl - erw [RingHom.comp_assoc] - conv_rhs => erw [RingHom.comp_assoc] - dsimp [CommRingCat.ofHom, localizationToStalk, PrimeSpectrum.localizationMapOfSpecializes] + rw [CommRingCat.hom_comp, RingHom.comp_assoc, CommRingCat.hom_comp, RingHom.comp_assoc] + dsimp [localizationToStalk, PrimeSpectrum.localizationMapOfSpecializes] rw [IsLocalization.lift_comp, IsLocalization.lift_comp, IsLocalization.lift_comp] exact CommRingCat.hom_ext_iff.mp (toStalk_stalkSpecializes h) diff --git a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean index d1b827f3f5c7c..8e789299821c5 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean @@ -80,7 +80,6 @@ theorem decomposition_Q (n q : ℕ) : variable (X) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The structure `MorphComponents` is an ad hoc structure that is used in the proof that `N₁ : SimplicialObject C ⥤ Karoubi (ChainComplex C ℕ))` reflects isomorphisms. The fields are the data that are needed in order to diff --git a/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean b/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean index b84886418a547..e75d6e00a68dc 100644 --- a/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean +++ b/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.AlgebraicTopology.AlternatingFaceMapComplex -import Mathlib.AlgebraicTopology.SimplicialSet.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.StdSimplex import Mathlib.AlgebraicTopology.CechNerve import Mathlib.Algebra.Homology.Homotopy import Mathlib.Tactic.FinCases @@ -176,10 +176,10 @@ def shift {n : ℕ} {Δ : SimplexCategory} substs hj₁ hj₂ simpa only [shiftFun_succ] using f.toOrderHom.monotone (Fin.succ_le_succ_iff.mp hi) } -open SSet.standardSimplex in +open SSet.stdSimplex in /-- The obvious extra degeneracy on the standard simplex. -/ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) : - SimplicialObject.Augmented.ExtraDegeneracy (standardSimplex.obj Δ) where + SimplicialObject.Augmented.ExtraDegeneracy (stdSimplex.obj Δ) where s' _ := objMk (OrderHom.const _ 0) s _ f := (objEquiv _ _).symm (shift (objEquiv _ _ f)) @@ -198,7 +198,7 @@ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) : apply (objEquiv _ _).injective apply SimplexCategory.Hom.ext ext i : 2 - dsimp [SimplicialObject.δ, SimplexCategory.δ, SSet.standardSimplex, + dsimp [SimplicialObject.δ, SimplexCategory.δ, SSet.stdSimplex, objEquiv, Equiv.ulift, uliftFunctor] simp only [shiftFun_succ] s_comp_δ n i := by @@ -206,7 +206,7 @@ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) : apply (objEquiv _ _).injective apply SimplexCategory.Hom.ext ext j : 2 - dsimp [SimplicialObject.δ, SimplexCategory.δ, SSet.standardSimplex, + dsimp [SimplicialObject.δ, SimplexCategory.δ, SSet.stdSimplex, objEquiv, Equiv.ulift, uliftFunctor] by_cases h : j = 0 · subst h @@ -219,7 +219,7 @@ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) : apply (objEquiv _ _).injective apply SimplexCategory.Hom.ext ext j : 2 - dsimp [SimplicialObject.σ, SimplexCategory.σ, SSet.standardSimplex, + dsimp [SimplicialObject.σ, SimplexCategory.σ, SSet.stdSimplex, objEquiv, Equiv.ulift, uliftFunctor] by_cases h : j = 0 · subst h @@ -227,8 +227,8 @@ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) : · obtain ⟨_, rfl⟩ := Fin.eq_succ_of_ne_zero h simp only [Fin.succ_predAbove_succ, shiftFun_succ, Function.comp_apply] -instance nonempty_extraDegeneracy_standardSimplex (Δ : SimplexCategory) : - Nonempty (SimplicialObject.Augmented.ExtraDegeneracy (standardSimplex.obj Δ)) := +instance nonempty_extraDegeneracy_stdSimplex (Δ : SimplexCategory) : + Nonempty (SimplicialObject.Augmented.ExtraDegeneracy (stdSimplex.obj Δ)) := ⟨StandardSimplex.extraDegeneracy Δ⟩ end StandardSimplex diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean index 99f2a18da0783..4ba7cc3bd0912 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean @@ -136,8 +136,6 @@ def uliftMap : C(TopCat.of (ULift.{u} I × X), Y) := ⟨fun x => H (x.1.down, x.2), H.continuous.comp ((continuous_uLift_down.comp continuous_fst).prod_mk continuous_snd)⟩ --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem ulift_apply (i : ULift.{u} I) (x : X) : H.uliftMap (i, x) = H (i.down, x) := rfl diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean index c322b54925ac3..fcfb94843d7c4 100644 --- a/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean @@ -31,8 +31,8 @@ open CategoryTheory Simplicial /-- A simplicial set `S` is a *quasicategory* if it satisfies the following horn-filling condition: for every `n : ℕ` and `0 < i < n`, every map of simplicial sets `σ₀ : Λ[n, i] → S` can be extended to a map `σ : Δ[n] → S`. - -[Kerodon, 003A] -/ +-/ +@[kerodon 003A] class Quasicategory (S : SSet) : Prop where hornFilling' : ∀ ⦃n : ℕ⦄ ⦃i : Fin (n+3)⦄ (σ₀ : Λ[n+2, i] ⟶ S) (_h0 : 0 < i) (_hn : i < Fin.last (n+2)), @@ -50,9 +50,8 @@ lemma Quasicategory.hornFilling {S : SSet} [Quasicategory S] ⦃n : ℕ⦄ ⦃i simp [hn] at h0 | succ n => exact Quasicategory.hornFilling' σ₀ h0 hn -/-- Every Kan complex is a quasicategory. - -[Kerodon, 003C] -/ +/-- Every Kan complex is a quasicategory. -/ +@[kerodon 003C] instance (S : SSet) [KanComplex S] : Quasicategory S where hornFilling' _ _ σ₀ _ _ := KanComplex.hornFilling σ₀ diff --git a/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean b/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean index 2cecadb7f1d91..f083df9d4a6ab 100644 --- a/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean @@ -39,18 +39,18 @@ instance quasicategory {X : SSet.{u}} [StrictSegal X] : Quasicategory X := by · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] apply congr_arg - simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + simp only [horn, horn.spineId, stdSimplex, uliftFunctor, Functor.comp_obj, yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, - standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + stdSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] rw [mkOfSucc_δ_lt hlt] rfl · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] apply congr_arg - simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + simp only [horn, horn.spineId, stdSimplex, uliftFunctor, Functor.comp_obj, yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, - standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + stdSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] rw [mkOfSucc_δ_gt hgt] rfl @@ -74,7 +74,7 @@ instance quasicategory {X : SSet.{u}} [StrictSegal X] : Quasicategory X := by dsimp [spine_arrow, Path.interval, Path.map] rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] apply congr_arg - simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + simp only [horn, stdSimplex, uliftFunctor, Functor.comp_obj, whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, Nat.reduceAdd, Quiver.Hom.unop_op] @@ -85,7 +85,7 @@ instance quasicategory {X : SSet.{u}} [StrictSegal X] : Quasicategory X := by simp only [spineToDiagonal, diagonal, spineToSimplex_spine] rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] apply congr_arg - simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + simp only [horn, stdSimplex, uliftFunctor, Functor.comp_obj, whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] @@ -94,7 +94,7 @@ instance quasicategory {X : SSet.{u}} [StrictSegal X] : Quasicategory X := by | zero => contradiction | succ _ => fin_cases z <;> - · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, + · simp only [stdSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, Quiver.Hom.unop_op, Equiv.ulift_symm_down] rw [mkOfSucc_δ_eq heq] rfl diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index fc1308dc19570..ba34f389c2b8c 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -82,7 +82,6 @@ theorem mk_len (n : SimplexCategory) : ([n.len] : SimplexCategory) = n := protected def rec {F : SimplexCategory → Sort*} (h : ∀ n : ℕ, F [n]) : ∀ X, F X := fun n => h n.len --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Morphisms in the `SimplexCategory`. -/ protected def Hom (a b : SimplexCategory) := Fin (a.len + 1) →o Fin (b.len + 1) @@ -904,14 +903,7 @@ theorem eq_σ_comp_of_not_injective' {n : ℕ} {Δ' : SimplexCategory} (θ : mk rw [Fin.succAbove_of_le_castSucc i.succ _] simp only [Fin.lt_iff_val_lt_val, Fin.le_iff_val_le_val, Fin.val_succ, Fin.coe_castSucc, Nat.lt_succ_iff, Fin.ext_iff] at h' h'' ⊢ - cases' Nat.le.dest h' with c hc - cases c - · exfalso - simp only [add_zero, len_mk, Fin.coe_pred] at hc - rw [hc] at h'' - exact h'' rfl - · rw [← hc] - simp only [add_le_add_iff_left, Nat.succ_eq_add_one, le_add_iff_nonneg_left, zero_le] + omega theorem eq_σ_comp_of_not_injective {n : ℕ} {Δ' : SimplexCategory} (θ : mk (n + 1) ⟶ Δ') (hθ : ¬Function.Injective θ.toOrderHom) : diff --git a/Mathlib/AlgebraicTopology/SimplicialCategory/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialCategory/Basic.lean index 88f5cb27d8813..3858e6c423419 100644 --- a/Mathlib/AlgebraicTopology/SimplicialCategory/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialCategory/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal -import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Enriched.Ordinary.Basic /-! # Simplicial categories diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index ae05211a3fe1e..4e4e34909f442 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -33,7 +33,6 @@ namespace CategoryTheory variable (C : Type u) [Category.{v} C] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The category of simplicial objects valued in a category `C`. This is the category of contravariant functors from `SimplexCategory` to `C`. -/ def SimplicialObject := @@ -206,7 +205,6 @@ variable (C) def whiskering (D : Type*) [Category D] : (C ⥤ D) ⥤ SimplicialObject C ⥤ SimplicialObject D := whiskeringRight _ _ _ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Truncated simplicial objects. -/ def Truncated (n : ℕ) := (SimplexCategory.Truncated n)ᵒᵖ ⥤ C @@ -360,7 +358,6 @@ variable (C) abbrev const : C ⥤ SimplicialObject C := CategoryTheory.Functor.const _ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The category of augmented simplicial objects, defined as a comma category. -/ def Augmented := Comma (𝟭 (SimplicialObject C)) (const C) @@ -472,7 +469,6 @@ theorem augment_hom_zero (X : SimplicialObject C) (X₀ : C) (f : X _[0] ⟶ X end SimplicialObject --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Cosimplicial objects. -/ def CosimplicialObject := SimplexCategory ⥤ C @@ -640,7 +636,6 @@ variable (C) def whiskering (D : Type*) [Category D] : (C ⥤ D) ⥤ CosimplicialObject C ⥤ CosimplicialObject D := whiskeringRight _ _ _ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Truncated cosimplicial objects. -/ def Truncated (n : ℕ) := SimplexCategory.Truncated n ⥤ C @@ -694,7 +689,6 @@ variable (C) abbrev const : C ⥤ CosimplicialObject C := CategoryTheory.Functor.const _ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Augmented cosimplicial objects. -/ def Augmented := Comma (const C) (𝟭 (CosimplicialObject C)) diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean index 7e6de96ffecde..9061d57109b29 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Split.lean @@ -198,7 +198,6 @@ def cofan' (Δ : SimplexCategoryᵒᵖ) : Cofan (summand N Δ) := end Splitting ---Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- A splitting of a simplicial object `X` consists of the datum of a sequence of objects `N`, a sequence of morphisms `ι : N n ⟶ X _[n]` such that for all `Δ : SimplexCategoryᵒᵖ`, the canonical map `Splitting.map X ι Δ` @@ -287,7 +286,6 @@ end Splitting variable (C) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The category `SimplicialObject.Split C` is the category of simplicial objects in `C` equipped with a splitting, and morphisms are morphisms of simplicial objects which are compatible with the splittings. -/ @@ -308,7 +306,6 @@ of a simplicial object `X`. -/ def mk' {X : SimplicialObject C} (s : Splitting X) : Split C := ⟨X, s⟩ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Morphisms in `SimplicialObject.Split C` are morphisms of simplicial objects that are compatible with the splittings. -/ structure Hom (S₁ S₂ : Split C) where diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index bddd445599fef..628c8c85b00e8 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Kim Morrison, Adam Topaz -/ import Mathlib.AlgebraicTopology.SimplicialObject.Basic -import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.CategoryTheory.Yoneda -import Mathlib.Data.Fin.VecNotation import Mathlib.Tactic.FinCases /-! @@ -19,16 +17,6 @@ i.e. a `Type`-valued presheaf on the simplex category. but this would be unnecessarily confusing given the existing notion of a simplicial type in homotopy type theory.) -We define the standard simplices `Δ[n]` as simplicial sets, -and their boundaries `∂Δ[n]` and horns `Λ[n, i]`. -(The notations are available via `Open Simplicial`.) - -## Future work - -There isn't yet a complete API for simplices, boundaries, and horns. -As an example, we should have a function that constructs -from a non-surjective order preserving function `Fin n → Fin n` -a morphism `Δ[n] ⟶ ∂Δ[n]`. -/ universe v u @@ -69,279 +57,6 @@ lemma comp_app {X Y Z : SSet} (f : X ⟶ Y) (g : Y ⟶ Z) (n : SimplexCategory def uliftFunctor : SSet.{u} ⥤ SSet.{max u v} := (SimplicialObject.whiskering _ _).obj CategoryTheory.uliftFunctor.{v, u} -/-- The `n`-th standard simplex `Δ[n]` associated with a nonempty finite linear order `n` -is the Yoneda embedding of `n`. -/ -def standardSimplex : SimplexCategory ⥤ SSet.{u} := - yoneda ⋙ uliftFunctor - -@[inherit_doc SSet.standardSimplex] -scoped[Simplicial] notation3 "Δ[" n "]" => SSet.standardSimplex.obj (SimplexCategory.mk n) - -instance : Inhabited SSet := - ⟨Δ[0]⟩ - -namespace standardSimplex - -open Finset Opposite SimplexCategory - -@[simp] -lemma map_id (n : SimplexCategory) : - (SSet.standardSimplex.map (SimplexCategory.Hom.mk OrderHom.id : n ⟶ n)) = 𝟙 _ := - CategoryTheory.Functor.map_id _ _ - -/-- Simplices of the standard simplex identify to morphisms in `SimplexCategory`. -/ -def objEquiv (n : SimplexCategory) (m : SimplexCategoryᵒᵖ) : - (standardSimplex.{u}.obj n).obj m ≃ (m.unop ⟶ n) := - Equiv.ulift.{u, 0} - -/-- Constructor for simplices of the standard simplex which takes a `OrderHom` as an input. -/ -abbrev objMk {n : SimplexCategory} {m : SimplexCategoryᵒᵖ} - (f : Fin (len m.unop + 1) →o Fin (n.len + 1)) : - (standardSimplex.{u}.obj n).obj m := - (objEquiv _ _).symm (Hom.mk f) - -lemma map_apply {m₁ m₂ : SimplexCategoryᵒᵖ} (f : m₁ ⟶ m₂) {n : SimplexCategory} - (x : (standardSimplex.{u}.obj n).obj m₁) : - (standardSimplex.{u}.obj n).map f x = (objEquiv _ _).symm (f.unop ≫ (objEquiv _ _) x) := by - rfl - -/-- The canonical bijection `(standardSimplex.obj n ⟶ X) ≃ X.obj (op n)`. -/ -def _root_.SSet.yonedaEquiv (X : SSet.{u}) (n : SimplexCategory) : - (standardSimplex.obj n ⟶ X) ≃ X.obj (op n) := - yonedaCompUliftFunctorEquiv X n - -/-- The unique non-degenerate `n`-simplex in `Δ[n]`. -/ -def id (n : ℕ) : Δ[n] _[n] := yonedaEquiv Δ[n] [n] (𝟙 Δ[n]) - -lemma id_eq_objEquiv_symm (n : ℕ) : id n = (objEquiv _ _).symm (𝟙 _) := rfl - -lemma objEquiv_id (n : ℕ) : objEquiv _ _ (id n) = 𝟙 _ := rfl - -/-- The (degenerate) `m`-simplex in the standard simplex concentrated in vertex `k`. -/ -def const (n : ℕ) (k : Fin (n+1)) (m : SimplexCategoryᵒᵖ) : Δ[n].obj m := - objMk (OrderHom.const _ k ) - -@[simp] -lemma const_down_toOrderHom (n : ℕ) (k : Fin (n+1)) (m : SimplexCategoryᵒᵖ) : - (const n k m).down.toOrderHom = OrderHom.const _ k := - rfl - -/-- The edge of the standard simplex with endpoints `a` and `b`. -/ -def edge (n : ℕ) (a b : Fin (n+1)) (hab : a ≤ b) : Δ[n] _[1] := by - refine objMk ⟨![a, b], ?_⟩ - rw [Fin.monotone_iff_le_succ] - simp only [unop_op, len_mk, Fin.forall_fin_one] - apply Fin.mk_le_mk.mpr hab - -lemma coe_edge_down_toOrderHom (n : ℕ) (a b : Fin (n+1)) (hab : a ≤ b) : - ↑(edge n a b hab).down.toOrderHom = ![a, b] := - rfl - -/-- The triangle in the standard simplex with vertices `a`, `b`, and `c`. -/ -def triangle {n : ℕ} (a b c : Fin (n+1)) (hab : a ≤ b) (hbc : b ≤ c) : Δ[n] _[2] := by - refine objMk ⟨![a, b, c], ?_⟩ - rw [Fin.monotone_iff_le_succ] - simp only [unop_op, len_mk, Fin.forall_fin_two] - dsimp - simp only [*, Matrix.tail_cons, Matrix.head_cons, true_and] - -lemma coe_triangle_down_toOrderHom {n : ℕ} (a b c : Fin (n+1)) (hab : a ≤ b) (hbc : b ≤ c) : - ↑(triangle a b c hab hbc).down.toOrderHom = ![a, b, c] := - rfl - -end standardSimplex - -section - -/-- The `m`-simplices of the `n`-th standard simplex are -the monotone maps from `Fin (m+1)` to `Fin (n+1)`. -/ -def asOrderHom {n} {m} (α : Δ[n].obj m) : OrderHom (Fin (m.unop.len + 1)) (Fin (n + 1)) := - α.down.toOrderHom - -end - -/-- The boundary `∂Δ[n]` of the `n`-th standard simplex consists of -all `m`-simplices of `standardSimplex n` that are not surjective -(when viewed as monotone function `m → n`). -/ -def boundary (n : ℕ) : SSet.{u} where - obj m := { α : Δ[n].obj m // ¬Function.Surjective (asOrderHom α) } - map {m₁ m₂} f α := - ⟨Δ[n].map f α.1, by - intro h - apply α.property - exact Function.Surjective.of_comp h⟩ - -/-- The boundary `∂Δ[n]` of the `n`-th standard simplex -/ -scoped[Simplicial] notation3 "∂Δ[" n "]" => SSet.boundary n - -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ α : Δ[n].obj m // _ }`. --/ -set_option linter.unusedVariables false in -/-- The inclusion of the boundary of the `n`-th standard simplex into that standard simplex. -/ -def boundaryInclusion (n : ℕ) : ∂Δ[n] ⟶ Δ[n] where app m (α : { α : Δ[n].obj m // _ }) := α - -/-- `horn n i` (or `Λ[n, i]`) is the `i`-th horn of the `n`-th standard simplex, where `i : n`. -It consists of all `m`-simplices `α` of `Δ[n]` -for which the union of `{i}` and the range of `α` is not all of `n` -(when viewing `α` as monotone function `m → n`). -/ -def horn (n : ℕ) (i : Fin (n + 1)) : SSet where - obj m := { α : Δ[n].obj m // Set.range (asOrderHom α) ∪ {i} ≠ Set.univ } - map {m₁ m₂} f α := - ⟨Δ[n].map f α.1, by - intro h; apply α.property - rw [Set.eq_univ_iff_forall] at h ⊢; intro j - apply Or.imp _ id (h j) - intro hj - exact Set.range_comp_subset_range _ _ hj⟩ - -/-- The `i`-th horn `Λ[n, i]` of the standard `n`-simplex -/ -scoped[Simplicial] notation3 "Λ[" n ", " i "]" => SSet.horn (n : ℕ) i - -#adaptation_note -/-- -The new unused variable linter in -https://github.com/leanprover/lean4/pull/5338 -flags `{ α : Δ[n].obj m // _ }`. --/ -set_option linter.unusedVariables false in -/-- The inclusion of the `i`-th horn of the `n`-th standard simplex into that standard simplex. -/ -def hornInclusion (n : ℕ) (i : Fin (n + 1)) : Λ[n, i] ⟶ Δ[n] where - app m (α : { α : Δ[n].obj m // _ }) := α - -namespace horn - -open SimplexCategory Finset Opposite - -/-- The (degenerate) subsimplex of `Λ[n+2, i]` concentrated in vertex `k`. -/ -@[simps] -def const (n : ℕ) (i k : Fin (n+3)) (m : SimplexCategoryᵒᵖ) : Λ[n+2, i].obj m := by - refine ⟨standardSimplex.const _ k _, ?_⟩ - suffices ¬ Finset.univ ⊆ {i, k} by - simpa [← Set.univ_subset_iff, Set.subset_def, asOrderHom, not_or, Fin.forall_fin_one, - subset_iff, mem_univ, @eq_comm _ _ k] - intro h - have := (card_le_card h).trans card_le_two - rw [card_fin] at this - omega - -/-- The edge of `Λ[n, i]` with endpoints `a` and `b`. - -This edge only exists if `{i, a, b}` has cardinality less than `n`. -/ -@[simps] -def edge (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : #{i, a, b} ≤ n) : Λ[n, i] _[1] := by - refine ⟨standardSimplex.edge n a b hab, ?range⟩ - case range => - suffices ∃ x, ¬i = x ∧ ¬a = x ∧ ¬b = x by - simpa only [unop_op, len_mk, Nat.reduceAdd, asOrderHom, yoneda_obj_obj, Set.union_singleton, - ne_eq, ← Set.univ_subset_iff, Set.subset_def, Set.mem_univ, Set.mem_insert_iff, - @eq_comm _ _ i, Set.mem_range, forall_const, not_forall, not_or, not_exists, - Fin.forall_fin_two, Fin.isValue] - contrapose! H - replace H : univ ⊆ {i, a, b} := - fun x _ ↦ by simpa [or_iff_not_imp_left, eq_comm] using H x - replace H := card_le_card H - rwa [card_fin] at H - -/-- Alternative constructor for the edge of `Λ[n, i]` with endpoints `a` and `b`, -assuming `3 ≤ n`. -/ -@[simps!] -def edge₃ (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : 3 ≤ n) : - Λ[n, i] _[1] := - horn.edge n i a b hab <| Finset.card_le_three.trans H - -/-- The edge of `Λ[n, i]` with endpoints `j` and `j+1`. - -This constructor assumes `0 < i < n`, -which is the type of horn that occurs in the horn-filling condition of quasicategories. -/ -@[simps!] -def primitiveEdge {n : ℕ} {i : Fin (n+1)} - (h₀ : 0 < i) (hₙ : i < Fin.last n) (j : Fin n) : - Λ[n, i] _[1] := by - refine horn.edge n i j.castSucc j.succ ?_ ?_ - · simp only [← Fin.val_fin_le, Fin.coe_castSucc, Fin.val_succ, le_add_iff_nonneg_right, zero_le] - simp only [← Fin.val_fin_lt, Fin.val_zero, Fin.val_last] at h₀ hₙ - obtain rfl|hn : n = 2 ∨ 2 < n := by - rw [eq_comm, or_comm, ← le_iff_lt_or_eq]; omega - · revert i j; decide - · exact Finset.card_le_three.trans hn - -/-- The triangle in the standard simplex with vertices `k`, `k+1`, and `k+2`. - -This constructor assumes `0 < i < n`, -which is the type of horn that occurs in the horn-filling condition of quasicategories. -/ -@[simps] -def primitiveTriangle {n : ℕ} (i : Fin (n+4)) - (h₀ : 0 < i) (hₙ : i < Fin.last (n+3)) - (k : ℕ) (h : k < n+2) : Λ[n+3, i] _[2] := by - refine ⟨standardSimplex.triangle - (n := n+3) ⟨k, by omega⟩ ⟨k+1, by omega⟩ ⟨k+2, by omega⟩ ?_ ?_, ?_⟩ - · simp only [Fin.mk_le_mk, le_add_iff_nonneg_right, zero_le] - · simp only [Fin.mk_le_mk, add_le_add_iff_left, one_le_two] - simp only [unop_op, SimplexCategory.len_mk, asOrderHom, SimplexCategory.Hom.toOrderHom_mk, - OrderHom.const_coe_coe, Set.union_singleton, ne_eq, ← Set.univ_subset_iff, Set.subset_def, - Set.mem_univ, Set.mem_insert_iff, Set.mem_range, Function.const_apply, exists_const, - forall_true_left, not_forall, not_or, unop_op, not_exists, - standardSimplex.triangle, OrderHom.coe_mk, @eq_comm _ _ i, - standardSimplex.objMk, standardSimplex.objEquiv, Equiv.ulift] - dsimp - by_cases hk0 : k = 0 - · subst hk0 - use Fin.last (n+3) - simp only [hₙ.ne, not_false_eq_true, Fin.zero_eta, zero_add, true_and] - intro j - fin_cases j <;> simp [Fin.ext_iff] - · use 0 - simp only [h₀.ne', not_false_eq_true, true_and] - intro j - fin_cases j <;> simp [Fin.ext_iff, hk0] - -/-- The `j`th subface of the `i`-th horn. -/ -@[simps] -def face {n : ℕ} (i j : Fin (n+2)) (h : j ≠ i) : Λ[n+1, i] _[n] := - ⟨(standardSimplex.objEquiv _ _).symm (SimplexCategory.δ j), by - simpa [← Set.univ_subset_iff, Set.subset_def, asOrderHom, SimplexCategory.δ, not_or, - standardSimplex.objEquiv, asOrderHom, Equiv.ulift]⟩ - -/-- Two morphisms from a horn are equal if they are equal on all suitable faces. -/ -protected -lemma hom_ext {n : ℕ} {i : Fin (n+2)} {S : SSet} (σ₁ σ₂ : Λ[n+1, i] ⟶ S) - (h : ∀ (j) (h : j ≠ i), σ₁.app _ (face i j h) = σ₂.app _ (face i j h)) : - σ₁ = σ₂ := by - apply NatTrans.ext; apply funext; apply Opposite.rec; apply SimplexCategory.rec - intro m; ext f - obtain ⟨f', hf⟩ := (standardSimplex.objEquiv _ _).symm.surjective f.1 - obtain ⟨j, hji, hfj⟩ : ∃ j, ¬j = i ∧ ∀ k, f'.toOrderHom k ≠ j := by - obtain ⟨f, hf'⟩ := f - subst hf - simpa [← Set.univ_subset_iff, Set.subset_def, asOrderHom, not_or] using hf' - have H : f = (Λ[n+1, i].map (factor_δ f' j).op) (face i j hji) := by - apply Subtype.ext - apply (standardSimplex.objEquiv _ _).injective - rw [← hf] - exact (factor_δ_spec f' j hfj).symm - have H₁ := congrFun (σ₁.naturality (factor_δ f' j).op) (face i j hji) - have H₂ := congrFun (σ₂.naturality (factor_δ f' j).op) (face i j hji) - dsimp at H₁ H₂ - rw [H, H₁, H₂, h _ hji] - -end horn - -section Examples - -open Simplicial - -/-- The simplicial circle. -/ -noncomputable def S1 : SSet := - Limits.colimit <| - Limits.parallelPair (standardSimplex.map <| SimplexCategory.δ 0 : Δ[0] ⟶ Δ[1]) - (standardSimplex.map <| SimplexCategory.δ 1) - -end Examples - /-- Truncated simplicial sets. -/ def Truncated (n : ℕ) := SimplicialObject.Truncated (Type u) n @@ -371,10 +86,6 @@ lemma Truncated.hom_ext {n : ℕ} {X Y : Truncated n} {f g : X ⟶ Y} (w : ∀ n /-- The truncation functor on simplicial sets. -/ abbrev truncation (n : ℕ) : SSet ⥤ SSet.Truncated n := SimplicialObject.truncation n -instance {n} : Inhabited (SSet.Truncated n) := - ⟨(truncation n).obj <| Δ[0]⟩ - - open SimplexCategory noncomputable section @@ -448,24 +159,6 @@ augmented simplicial objects. -/ abbrev Augmented := SimplicialObject.Augmented (Type u) -namespace Augmented - --- Porting note: an instance of `Subsingleton (⊤_ (Type u))` was added in --- `CategoryTheory.Limits.Types` to ease the automation in this definition -/-- The functor which sends `[n]` to the simplicial set `Δ[n]` equipped by -the obvious augmentation towards the terminal object of the category of sets. -/ -@[simps] -noncomputable def standardSimplex : SimplexCategory ⥤ SSet.Augmented.{u} where - obj Δ := - { left := SSet.standardSimplex.obj Δ - right := terminal _ - hom := { app := fun _ => terminal.from _ } } - map θ := - { left := SSet.standardSimplex.map θ - right := terminal.from _ } - -end Augmented - section applications variable {S : SSet} diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean new file mode 100644 index 0000000000000..58805409d1ceb --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Boundary.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2021 Kim Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Kim Morrison, Adam Topaz +-/ +import Mathlib.AlgebraicTopology.SimplicialSet.StdSimplex + +/-! +# The boundary of the standard simplex + +We introduce the boundary `∂Δ[n]` of the standard simplex `Δ[n]`. +(These notations become available by doing `open Simplicial`.) + +## Future work + +There isn't yet a complete API for simplices, boundaries, and horns. +As an example, we should have a function that constructs +from a non-surjective order preserving function `Fin n → Fin n` +a morphism `Δ[n] ⟶ ∂Δ[n]`. + + +-/ + +universe u + +open Simplicial + +namespace SSet + +/-- The boundary `∂Δ[n]` of the `n`-th standard simplex consists of +all `m`-simplices of `stdSimplex n` that are not surjective +(when viewed as monotone function `m → n`). -/ +def boundary (n : ℕ) : SSet.{u} where + obj m := { α : Δ[n].obj m // ¬Function.Surjective (asOrderHom α) } + map {m₁ m₂} f α := + ⟨Δ[n].map f α.1, by + intro h + apply α.property + exact Function.Surjective.of_comp h⟩ + +/-- The boundary `∂Δ[n]` of the `n`-th standard simplex -/ +scoped[Simplicial] notation3 "∂Δ[" n "]" => SSet.boundary n + +#adaptation_note +/-- +The new unused variable linter in +https://github.com/leanprover/lean4/pull/5338 +flags `{ α : Δ[n].obj m // _ }`. +-/ +set_option linter.unusedVariables false in +/-- The inclusion of the boundary of the `n`-th standard simplex into that standard simplex. -/ +def boundaryInclusion (n : ℕ) : ∂Δ[n] ⟶ Δ[n] where app m (α : { α : Δ[n].obj m // _ }) := α + +end SSet diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean new file mode 100644 index 0000000000000..ef8c001fd1d79 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Horn.lean @@ -0,0 +1,167 @@ +/- +Copyright (c) 2021 Kim Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Kim Morrison, Adam Topaz +-/ +import Mathlib.AlgebraicTopology.SimplicialSet.StdSimplex + +/-! +# Horns + +This file introduce horns `Λ[n, i]`. + +-/ + +universe u + +open CategoryTheory Simplicial + +namespace SSet + +/-- `horn n i` (or `Λ[n, i]`) is the `i`-th horn of the `n`-th standard simplex, where `i : n`. +It consists of all `m`-simplices `α` of `Δ[n]` +for which the union of `{i}` and the range of `α` is not all of `n` +(when viewing `α` as monotone function `m → n`). -/ +def horn (n : ℕ) (i : Fin (n + 1)) : SSet where + obj m := { α : Δ[n].obj m // Set.range (asOrderHom α) ∪ {i} ≠ Set.univ } + map {m₁ m₂} f α := + ⟨Δ[n].map f α.1, by + intro h; apply α.property + rw [Set.eq_univ_iff_forall] at h ⊢; intro j + apply Or.imp _ id (h j) + intro hj + exact Set.range_comp_subset_range _ _ hj⟩ + +/-- The `i`-th horn `Λ[n, i]` of the standard `n`-simplex -/ +scoped[Simplicial] notation3 "Λ[" n ", " i "]" => SSet.horn (n : ℕ) i + +#adaptation_note +/-- +The new unused variable linter in +https://github.com/leanprover/lean4/pull/5338 +flags `{ α : Δ[n].obj m // _ }`. +-/ +set_option linter.unusedVariables false in +/-- The inclusion of the `i`-th horn of the `n`-th standard simplex into that standard simplex. -/ +def hornInclusion (n : ℕ) (i : Fin (n + 1)) : Λ[n, i] ⟶ Δ[n] where + app m (α : { α : Δ[n].obj m // _ }) := α + +namespace horn + +open SimplexCategory Finset Opposite + +/-- The (degenerate) subsimplex of `Λ[n+2, i]` concentrated in vertex `k`. -/ +@[simps] +def const (n : ℕ) (i k : Fin (n+3)) (m : SimplexCategoryᵒᵖ) : Λ[n+2, i].obj m := by + refine ⟨stdSimplex.const _ k _, ?_⟩ + suffices ¬ Finset.univ ⊆ {i, k} by + simpa [← Set.univ_subset_iff, Set.subset_def, asOrderHom, not_or, Fin.forall_fin_one, + subset_iff, mem_univ, @eq_comm _ _ k] + intro h + have := (card_le_card h).trans card_le_two + rw [card_fin] at this + omega + +/-- The edge of `Λ[n, i]` with endpoints `a` and `b`. + +This edge only exists if `{i, a, b}` has cardinality less than `n`. -/ +@[simps] +def edge (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : #{i, a, b} ≤ n) : Λ[n, i] _[1] := by + refine ⟨stdSimplex.edge n a b hab, ?range⟩ + case range => + suffices ∃ x, ¬i = x ∧ ¬a = x ∧ ¬b = x by + simpa only [unop_op, len_mk, Nat.reduceAdd, asOrderHom, yoneda_obj_obj, Set.union_singleton, + ne_eq, ← Set.univ_subset_iff, Set.subset_def, Set.mem_univ, Set.mem_insert_iff, + @eq_comm _ _ i, Set.mem_range, forall_const, not_forall, not_or, not_exists, + Fin.forall_fin_two, Fin.isValue] + contrapose! H + replace H : univ ⊆ {i, a, b} := + fun x _ ↦ by simpa [or_iff_not_imp_left, eq_comm] using H x + replace H := card_le_card H + rwa [card_fin] at H + +/-- Alternative constructor for the edge of `Λ[n, i]` with endpoints `a` and `b`, +assuming `3 ≤ n`. -/ +@[simps!] +def edge₃ (n : ℕ) (i a b : Fin (n+1)) (hab : a ≤ b) (H : 3 ≤ n) : + Λ[n, i] _[1] := + horn.edge n i a b hab <| Finset.card_le_three.trans H + +/-- The edge of `Λ[n, i]` with endpoints `j` and `j+1`. + +This constructor assumes `0 < i < n`, +which is the type of horn that occurs in the horn-filling condition of quasicategories. -/ +@[simps!] +def primitiveEdge {n : ℕ} {i : Fin (n+1)} + (h₀ : 0 < i) (hₙ : i < Fin.last n) (j : Fin n) : + Λ[n, i] _[1] := by + refine horn.edge n i j.castSucc j.succ ?_ ?_ + · simp only [← Fin.val_fin_le, Fin.coe_castSucc, Fin.val_succ, le_add_iff_nonneg_right, zero_le] + simp only [← Fin.val_fin_lt, Fin.val_zero, Fin.val_last] at h₀ hₙ + obtain rfl|hn : n = 2 ∨ 2 < n := by + rw [eq_comm, or_comm, ← le_iff_lt_or_eq]; omega + · revert i j; decide + · exact Finset.card_le_three.trans hn + +/-- The triangle in the standard simplex with vertices `k`, `k+1`, and `k+2`. + +This constructor assumes `0 < i < n`, +which is the type of horn that occurs in the horn-filling condition of quasicategories. -/ +@[simps] +def primitiveTriangle {n : ℕ} (i : Fin (n+4)) + (h₀ : 0 < i) (hₙ : i < Fin.last (n+3)) + (k : ℕ) (h : k < n+2) : Λ[n+3, i] _[2] := by + refine ⟨stdSimplex.triangle + (n := n+3) ⟨k, by omega⟩ ⟨k+1, by omega⟩ ⟨k+2, by omega⟩ ?_ ?_, ?_⟩ + · simp only [Fin.mk_le_mk, le_add_iff_nonneg_right, zero_le] + · simp only [Fin.mk_le_mk, add_le_add_iff_left, one_le_two] + simp only [unop_op, SimplexCategory.len_mk, asOrderHom, SimplexCategory.Hom.toOrderHom_mk, + OrderHom.const_coe_coe, Set.union_singleton, ne_eq, ← Set.univ_subset_iff, Set.subset_def, + Set.mem_univ, Set.mem_insert_iff, Set.mem_range, Function.const_apply, exists_const, + forall_true_left, not_forall, not_or, unop_op, not_exists, + stdSimplex.triangle, OrderHom.coe_mk, @eq_comm _ _ i, + stdSimplex.objMk, stdSimplex.objEquiv, Equiv.ulift] + dsimp + by_cases hk0 : k = 0 + · subst hk0 + use Fin.last (n+3) + simp only [hₙ.ne, not_false_eq_true, Fin.zero_eta, zero_add, true_and] + intro j + fin_cases j <;> simp [Fin.ext_iff] + · use 0 + simp only [h₀.ne', not_false_eq_true, true_and] + intro j + fin_cases j <;> simp [Fin.ext_iff, hk0] + +/-- The `j`th subface of the `i`-th horn. -/ +@[simps] +def face {n : ℕ} (i j : Fin (n+2)) (h : j ≠ i) : Λ[n+1, i] _[n] := + ⟨(stdSimplex.objEquiv _ _).symm (SimplexCategory.δ j), by + simpa [← Set.univ_subset_iff, Set.subset_def, asOrderHom, SimplexCategory.δ, not_or, + stdSimplex.objEquiv, asOrderHom, Equiv.ulift]⟩ + +/-- Two morphisms from a horn are equal if they are equal on all suitable faces. -/ +protected +lemma hom_ext {n : ℕ} {i : Fin (n+2)} {S : SSet} (σ₁ σ₂ : Λ[n+1, i] ⟶ S) + (h : ∀ (j) (h : j ≠ i), σ₁.app _ (face i j h) = σ₂.app _ (face i j h)) : + σ₁ = σ₂ := by + apply NatTrans.ext; apply funext; apply Opposite.rec; apply SimplexCategory.rec + intro m; ext f + obtain ⟨f', hf⟩ := (stdSimplex.objEquiv _ _).symm.surjective f.1 + obtain ⟨j, hji, hfj⟩ : ∃ j, ¬j = i ∧ ∀ k, f'.toOrderHom k ≠ j := by + obtain ⟨f, hf'⟩ := f + subst hf + simpa [← Set.univ_subset_iff, Set.subset_def, asOrderHom, not_or] using hf' + have H : f = (Λ[n+1, i].map (factor_δ f' j).op) (face i j hji) := by + apply Subtype.ext + apply (stdSimplex.objEquiv _ _).injective + rw [← hf] + exact (factor_δ_spec f' j hfj).symm + have H₁ := congrFun (σ₁.naturality (factor_δ f' j).op) (face i j hji) + have H₂ := congrFun (σ₂.naturality (factor_δ f' j).op) (face i j hji) + dsimp at H₁ H₂ + rw [H, H₁, H₂, h _ hji] + +end horn + +end SSet diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean b/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean index 982bd19ff062b..47ca5bae4df0c 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.AlgebraicTopology.SimplicialSet.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.Horn /-! # Kan complexes diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index ff920be3bfa29..3969919746617 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -3,7 +3,7 @@ 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.SimplicialSet.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.Horn /-! # Paths in simplicial sets @@ -110,8 +110,8 @@ lemma map_interval {X Y : SSet.{u}} {n : ℕ} (f : X.Path n) (σ : X ⟶ Y) (f.map σ).interval j l hjl = (f.interval j l hjl).map σ := rfl /-- The spine of the unique non-degenerate `n`-simplex in `Δ[n]`.-/ -def standardSimplex.spineId (n : ℕ) : Path Δ[n] n := - spine Δ[n] n (standardSimplex.id n) +def stdSimplex.spineId (n : ℕ) : Path Δ[n] n := + spine Δ[n] n (stdSimplex.id n) /-- Any inner horn contains the spine of the unique non-degenerate `n`-simplex in `Δ[n]`.-/ @@ -119,27 +119,27 @@ in `Δ[n]`.-/ def horn.spineId {n : ℕ} (i : Fin (n + 3)) (h₀ : 0 < i) (hₙ : i < Fin.last (n + 2)) : Path Λ[n + 2, i] (n + 2) where - vertex j := ⟨standardSimplex.spineId _ |>.vertex j, (horn.const n i j _).property⟩ - arrow j := ⟨standardSimplex.spineId _ |>.arrow j, by + vertex j := ⟨stdSimplex.spineId _ |>.vertex j, (horn.const n i j _).property⟩ + arrow j := ⟨stdSimplex.spineId _ |>.arrow j, by let edge := horn.primitiveEdge h₀ hₙ j - have ha : (standardSimplex.spineId _).arrow j = edge.val := by - dsimp only [edge, standardSimplex.spineId, standardSimplex.id, spine_arrow, - mkOfSucc, horn.primitiveEdge, horn.edge, standardSimplex.edge, - standardSimplex.map_apply] + have ha : (stdSimplex.spineId _).arrow j = edge.val := by + dsimp only [edge, stdSimplex.spineId, stdSimplex.id, spine_arrow, + mkOfSucc, horn.primitiveEdge, horn.edge, stdSimplex.edge, + stdSimplex.map_apply] aesop rw [ha] exact edge.property⟩ arrow_src := by simp only [horn, SimplicialObject.δ, Subtype.mk.injEq] - exact standardSimplex.spineId _ |>.arrow_src + exact stdSimplex.spineId _ |>.arrow_src arrow_tgt := by simp only [horn, SimplicialObject.δ, Subtype.mk.injEq] - exact standardSimplex.spineId _ |>.arrow_tgt + exact stdSimplex.spineId _ |>.arrow_tgt @[simp] lemma horn.spineId_map_hornInclusion {n : ℕ} (i : Fin (n + 3)) (h₀ : 0 < i) (hₙ : i < Fin.last (n + 2)) : Path.map (horn.spineId i h₀ hₙ) (hornInclusion (n + 2) i) = - standardSimplex.spineId (n + 2) := rfl + stdSimplex.spineId (n + 2) := rfl end SSet diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StdSimplex.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StdSimplex.lean new file mode 100644 index 0000000000000..60fefd38146d0 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StdSimplex.lean @@ -0,0 +1,152 @@ +/- +Copyright (c) 2021 Kim Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Kim Morrison, Adam Topaz +-/ +import Mathlib.AlgebraicTopology.SimplicialSet.Basic +import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.Data.Fin.VecNotation + +/-! +# The standard simplex + +We define the standard simplices `Δ[n]` as simplicial sets. +See files `SimplicialSet.Boundary` and `SimplicialSet.Horn` +for their boundaries`∂Δ[n]` and horns `Λ[n, i]`. +(The notations are available via `open Simplicial`.) + +-/ + +universe u + +open CategoryTheory Limits Simplicial + +namespace SSet + +/-- The functor `SimplexCategory ⥤ SSet` which sends `SimplexCategory.mk n` to +the standard simplex `Δ[n]` is a cosimplicial object in the category of simplicial sets. +(This functor is essentially given by the Yoneda embedding). -/ +def stdSimplex : CosimplicialObject SSet.{u} := + yoneda ⋙ uliftFunctor + +@[deprecated (since := "2025-01-23")] alias standardSimplex := stdSimplex + +@[inherit_doc SSet.stdSimplex] +scoped[Simplicial] notation3 "Δ[" n "]" => SSet.stdSimplex.obj (SimplexCategory.mk n) + +instance : Inhabited SSet := + ⟨Δ[0]⟩ + +instance {n} : Inhabited (SSet.Truncated n) := + ⟨(truncation n).obj <| Δ[0]⟩ + +namespace stdSimplex + +open Finset Opposite SimplexCategory + +@[simp] +lemma map_id (n : SimplexCategory) : + (SSet.stdSimplex.map (SimplexCategory.Hom.mk OrderHom.id : n ⟶ n)) = 𝟙 _ := + CategoryTheory.Functor.map_id _ _ + +/-- Simplices of the standard simplex identify to morphisms in `SimplexCategory`. -/ +def objEquiv (n : SimplexCategory) (m : SimplexCategoryᵒᵖ) : + (stdSimplex.{u}.obj n).obj m ≃ (m.unop ⟶ n) := + Equiv.ulift.{u, 0} + +/-- Constructor for simplices of the standard simplex which takes a `OrderHom` as an input. -/ +abbrev objMk {n : SimplexCategory} {m : SimplexCategoryᵒᵖ} + (f : Fin (len m.unop + 1) →o Fin (n.len + 1)) : + (stdSimplex.{u}.obj n).obj m := + (objEquiv _ _).symm (Hom.mk f) + +lemma map_apply {m₁ m₂ : SimplexCategoryᵒᵖ} (f : m₁ ⟶ m₂) {n : SimplexCategory} + (x : (stdSimplex.{u}.obj n).obj m₁) : + (stdSimplex.{u}.obj n).map f x = (objEquiv _ _).symm (f.unop ≫ (objEquiv _ _) x) := by + rfl + +/-- The canonical bijection `(stdSimplex.obj n ⟶ X) ≃ X.obj (op n)`. -/ +def _root_.SSet.yonedaEquiv (X : SSet.{u}) (n : SimplexCategory) : + (stdSimplex.obj n ⟶ X) ≃ X.obj (op n) := + yonedaCompUliftFunctorEquiv X n + +/-- The unique non-degenerate `n`-simplex in `Δ[n]`. -/ +def id (n : ℕ) : Δ[n] _[n] := yonedaEquiv Δ[n] [n] (𝟙 Δ[n]) + +lemma id_eq_objEquiv_symm (n : ℕ) : id n = (objEquiv _ _).symm (𝟙 _) := rfl + +lemma objEquiv_id (n : ℕ) : objEquiv _ _ (id n) = 𝟙 _ := rfl + +/-- The (degenerate) `m`-simplex in the standard simplex concentrated in vertex `k`. -/ +def const (n : ℕ) (k : Fin (n+1)) (m : SimplexCategoryᵒᵖ) : Δ[n].obj m := + objMk (OrderHom.const _ k ) + +@[simp] +lemma const_down_toOrderHom (n : ℕ) (k : Fin (n+1)) (m : SimplexCategoryᵒᵖ) : + (const n k m).down.toOrderHom = OrderHom.const _ k := + rfl + +/-- The edge of the standard simplex with endpoints `a` and `b`. -/ +def edge (n : ℕ) (a b : Fin (n+1)) (hab : a ≤ b) : Δ[n] _[1] := by + refine objMk ⟨![a, b], ?_⟩ + rw [Fin.monotone_iff_le_succ] + simp only [unop_op, len_mk, Fin.forall_fin_one] + apply Fin.mk_le_mk.mpr hab + +lemma coe_edge_down_toOrderHom (n : ℕ) (a b : Fin (n+1)) (hab : a ≤ b) : + ↑(edge n a b hab).down.toOrderHom = ![a, b] := + rfl + +/-- The triangle in the standard simplex with vertices `a`, `b`, and `c`. -/ +def triangle {n : ℕ} (a b c : Fin (n+1)) (hab : a ≤ b) (hbc : b ≤ c) : Δ[n] _[2] := by + refine objMk ⟨![a, b, c], ?_⟩ + rw [Fin.monotone_iff_le_succ] + simp only [unop_op, len_mk, Fin.forall_fin_two] + dsimp + simp only [*, Matrix.tail_cons, Matrix.head_cons, true_and] + +lemma coe_triangle_down_toOrderHom {n : ℕ} (a b c : Fin (n+1)) (hab : a ≤ b) (hbc : b ≤ c) : + ↑(triangle a b c hab hbc).down.toOrderHom = ![a, b, c] := + rfl + +end stdSimplex + +section + +/-- The `m`-simplices of the `n`-th standard simplex are +the monotone maps from `Fin (m+1)` to `Fin (n+1)`. -/ +def asOrderHom {n} {m} (α : Δ[n].obj m) : OrderHom (Fin (m.unop.len + 1)) (Fin (n + 1)) := + α.down.toOrderHom + +end + +section Examples + +open Simplicial + +/-- The simplicial circle. -/ +noncomputable def S1 : SSet := + Limits.colimit <| + Limits.parallelPair (stdSimplex.δ 0 : Δ[0] ⟶ Δ[1]) (stdSimplex.δ 1) + +end Examples + +namespace Augmented + +-- Porting note: an instance of `Subsingleton (⊤_ (Type u))` was added in +-- `CategoryTheory.Limits.Types` to ease the automation in this definition +/-- The functor which sends `[n]` to the simplicial set `Δ[n]` equipped by +the obvious augmentation towards the terminal object of the category of sets. -/ +@[simps] +noncomputable def stdSimplex : SimplexCategory ⥤ SSet.Augmented.{u} where + obj Δ := + { left := SSet.stdSimplex.obj Δ + right := terminal _ + hom := { app := fun _ => terminal.from _ } } + map θ := + { left := SSet.stdSimplex.map θ + right := terminal.from _ } + +end Augmented + +end SSet diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index d770f727726f8..f607e676809c6 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -442,6 +442,7 @@ variable (𝕜) /-- Given a function `f : E → F`, we say that `f` is analytic at `x` if it admits a convergent power series expansion around `x`. -/ +@[fun_prop] def AnalyticAt (f : E → F) (x : E) := ∃ p : FormalMultilinearSeries 𝕜 E F, HasFPowerSeriesAt f p x @@ -1343,6 +1344,7 @@ protected theorem AnalyticWithinAt.continuousWithinAt (hf : AnalyticWithinAt ContinuousWithinAt f s x := hf.continuousWithinAt_insert.mono (subset_insert x s) +@[fun_prop] protected theorem AnalyticAt.continuousAt (hf : AnalyticAt 𝕜 f x) : ContinuousAt f x := let ⟨_, hp⟩ := hf hp.continuousAt diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index f7188918787a3..ae4386b4831c7 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -860,17 +860,30 @@ alias AnalyticWithinOn.comp := AnalyticOn.comp /-- If two functions `g` and `f` are analytic respectively at `f x` and `x`, then `g ∘ f` is analytic at `x`. -/ +@[fun_prop] theorem AnalyticAt.comp {g : F → G} {f : E → F} {x : E} (hg : AnalyticAt 𝕜 g (f x)) (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (g ∘ f) x := by rw [← analyticWithinAt_univ] at hg hf ⊢ apply hg.comp hf (by simp) +/-- If two functions `g` and `f` are analytic respectively at `f x` and `x`, then `g ∘ f` is +analytic at `x`. -/ +@[fun_prop] +theorem AnalyticAt.comp' {g : F → G} {f : E → F} {x : E} (hg : AnalyticAt 𝕜 g (f x)) + (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (fun z ↦ g (f z)) x := + hg.comp hf + /-- Version of `AnalyticAt.comp` where point equality is a separate hypothesis. -/ theorem AnalyticAt.comp_of_eq {g : F → G} {f : E → F} {y : F} {x : E} (hg : AnalyticAt 𝕜 g y) (hf : AnalyticAt 𝕜 f x) (hy : f x = y) : AnalyticAt 𝕜 (g ∘ f) x := by rw [← hy] at hg exact hg.comp hf +/-- Version of `AnalyticAt.comp` where point equality is a separate hypothesis. -/ +theorem AnalyticAt.comp_of_eq' {g : F → G} {f : E → F} {y : F} {x : E} (hg : AnalyticAt 𝕜 g y) + (hf : AnalyticAt 𝕜 f x) (hy : f x = y) : AnalyticAt 𝕜 (fun z ↦ g (f z)) x := by + apply hg.comp_of_eq hf hy + theorem AnalyticAt.comp_analyticWithinAt {g : F → G} {f : E → F} {x : E} {s : Set E} (hg : AnalyticAt 𝕜 g (f x)) (hf : AnalyticWithinAt 𝕜 f s x) : AnalyticWithinAt 𝕜 (g ∘ f) s x := by diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index f56a2644814f0..a835539121144 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -46,6 +46,7 @@ theorem hasFPowerSeriesAt_const {c : F} {e : E} : HasFPowerSeriesAt (fun _ => c) (constFormalMultilinearSeries 𝕜 E c) e := ⟨⊤, hasFPowerSeriesOnBall_const⟩ +@[fun_prop] theorem analyticAt_const {v : F} {x : E} : AnalyticAt 𝕜 (fun _ => v) x := ⟨constFormalMultilinearSeries 𝕜 E v, hasFPowerSeriesAt_const⟩ @@ -100,11 +101,17 @@ theorem AnalyticWithinAt.add (hf : AnalyticWithinAt 𝕜 f s x) (hg : AnalyticWi let ⟨_, hqf⟩ := hg (hpf.add hqf).analyticWithinAt +@[fun_prop] theorem AnalyticAt.add (hf : AnalyticAt 𝕜 f x) (hg : AnalyticAt 𝕜 g x) : AnalyticAt 𝕜 (f + g) x := let ⟨_, hpf⟩ := hf let ⟨_, hqf⟩ := hg (hpf.add hqf).analyticAt +@[fun_prop] +theorem AnalyticAt.add' (hf : AnalyticAt 𝕜 f x) (hg : AnalyticAt 𝕜 g x) : + AnalyticAt 𝕜 (fun z ↦ f z + g z) x := + hf.add hg + theorem HasFPowerSeriesWithinOnBall.neg (hf : HasFPowerSeriesWithinOnBall f pf s x r) : HasFPowerSeriesWithinOnBall (-f) (-pf) s x r := { r_le := by @@ -134,10 +141,15 @@ theorem AnalyticWithinAt.neg (hf : AnalyticWithinAt 𝕜 f s x) : AnalyticWithin let ⟨_, hpf⟩ := hf hpf.neg.analyticWithinAt +@[fun_prop] theorem AnalyticAt.neg (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (-f) x := let ⟨_, hpf⟩ := hf hpf.neg.analyticAt +@[fun_prop] +theorem AnalyticAt.neg' (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (fun z ↦ -f z) x := + hf.neg + theorem HasFPowerSeriesWithinOnBall.sub (hf : HasFPowerSeriesWithinOnBall f pf s x r) (hg : HasFPowerSeriesWithinOnBall g pg s x r) : HasFPowerSeriesWithinOnBall (f - g) (pf - pg) s x r := by @@ -160,10 +172,16 @@ theorem AnalyticWithinAt.sub (hf : AnalyticWithinAt 𝕜 f s x) (hg : AnalyticWi AnalyticWithinAt 𝕜 (f - g) s x := by simpa only [sub_eq_add_neg] using hf.add hg.neg +@[fun_prop] theorem AnalyticAt.sub (hf : AnalyticAt 𝕜 f x) (hg : AnalyticAt 𝕜 g x) : AnalyticAt 𝕜 (f - g) x := by simpa only [sub_eq_add_neg] using hf.add hg.neg +@[fun_prop] +theorem AnalyticAt.sub' (hf : AnalyticAt 𝕜 f x) (hg : AnalyticAt 𝕜 g x) : + AnalyticAt 𝕜 (fun z ↦ f z - g z) x := + hf.sub hg + theorem HasFPowerSeriesWithinOnBall.const_smul (hf : HasFPowerSeriesWithinOnBall f pf s x r) : HasFPowerSeriesWithinOnBall (c • f) (c • pf) s x r where r_le := le_trans hf.r_le pf.radius_le_smul @@ -191,10 +209,15 @@ theorem AnalyticWithinAt.const_smul (hf : AnalyticWithinAt 𝕜 f s x) : let ⟨_, hpf⟩ := hf hpf.const_smul.analyticWithinAt +@[fun_prop] theorem AnalyticAt.const_smul (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (c • f) x := let ⟨_, hpf⟩ := hf hpf.const_smul.analyticAt +@[fun_prop] +theorem AnalyticAt.const_smul' (hf : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (fun z ↦ c • f z) x := + hf.const_smul + theorem AnalyticOn.add (hf : AnalyticOn 𝕜 f s) (hg : AnalyticOn 𝕜 g s) : AnalyticOn 𝕜 (f + g) s := fun z hz => (hf z hz).add (hg z hz) @@ -307,6 +330,7 @@ lemma AnalyticWithinAt.prod {e : E} {f : E → F} {g : E → G} {s : Set E} exact ⟨_, hf.prod hg⟩ /-- The Cartesian product of analytic functions is analytic. -/ +@[fun_prop] lemma AnalyticAt.prod {e : E} {f : E → F} {g : E → G} (hf : AnalyticAt 𝕜 f e) (hg : AnalyticAt 𝕜 g e) : AnalyticAt 𝕜 (fun x ↦ (f x, g x)) e := by @@ -590,11 +614,13 @@ pedantic to allow towers of field extensions. TODO: can we replace `𝕜'` with a "normed module" in such a way that `analyticAt_mul` is a special case of this? -/ +@[fun_prop] lemma analyticAt_smul [NormedSpace 𝕝 E] [IsScalarTower 𝕜 𝕝 E] (z : 𝕝 × E) : AnalyticAt 𝕜 (fun x : 𝕝 × E ↦ x.1 • x.2) z := (ContinuousLinearMap.lsmul 𝕜 𝕝).analyticAt_bilinear z /-- Multiplication in a normed algebra over `𝕜` is analytic. -/ +@[fun_prop] lemma analyticAt_mul (z : A × A) : AnalyticAt 𝕜 (fun x : A × A ↦ x.1 * x.2) z := (ContinuousLinearMap.mul 𝕜 A).analyticAt_bilinear z @@ -606,11 +632,19 @@ lemma AnalyticWithinAt.smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] (analyticAt_smul _).comp₂_analyticWithinAt hf hg /-- Scalar multiplication of one analytic function by another. -/ +@[fun_prop] lemma AnalyticAt.smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E → 𝕝} {g : E → F} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : - AnalyticAt 𝕜 (fun x ↦ f x • g x) z := + AnalyticAt 𝕜 (f • g) z := (analyticAt_smul _).comp₂ hf hg +/-- Scalar multiplication of one analytic function by another. -/ +@[fun_prop] +lemma AnalyticAt.smul' [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E → 𝕝} {g : E → F} {z : E} + (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : + AnalyticAt 𝕜 (fun x ↦ f x • g x) z := + hf.smul hg + /-- Scalar multiplication of one analytic function by another. -/ lemma AnalyticOn.smul [NormedSpace 𝕝 F] [IsScalarTower 𝕜 𝕝 F] {f : E → 𝕝} {g : E → F} {s : Set E} @@ -634,10 +668,16 @@ lemma AnalyticWithinAt.mul {f g : E → A} {s : Set E} {z : E} (analyticAt_mul _).comp₂_analyticWithinAt hf hg /-- Multiplication of analytic functions (valued in a normed `𝕜`-algebra) is analytic. -/ +@[fun_prop] lemma AnalyticAt.mul {f g : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : - AnalyticAt 𝕜 (fun x ↦ f x * g x) z := + AnalyticAt 𝕜 (f * g) z := (analyticAt_mul _).comp₂ hf hg +@[fun_prop] +lemma AnalyticAt.mul' {f g : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (hg : AnalyticAt 𝕜 g z) : + AnalyticAt 𝕜 (fun x ↦ f x * g x) z := + hf.mul hg + /-- Multiplication of analytic functions (valued in a normed `𝕜`-algebra) is analytic. -/ lemma AnalyticOn.mul {f g : E → A} {s : Set E} (hf : AnalyticOn 𝕜 f s) (hg : AnalyticOn 𝕜 g s) : @@ -665,6 +705,7 @@ lemma AnalyticWithinAt.pow {f : E → A} {z : E} {s : Set E} (hf : AnalyticWithi exact hm.mul hf /-- Powers of analytic functions (into a normed `𝕜`-algebra) are analytic. -/ +@[fun_prop] lemma AnalyticAt.pow {f : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (n : ℕ) : AnalyticAt 𝕜 (fun x ↦ f x ^ n) z := by rw [← analyticWithinAt_univ] at hf ⊢ @@ -791,6 +832,7 @@ lemma hasFPowerSeriesOnBall_inverse_one_sub List.ofFn_const, List.prod_replicate] exact (summable_geometric_of_norm_lt_one hy).hasSum +@[fun_prop] lemma analyticAt_inverse_one_sub (𝕜 : Type*) [NontriviallyNormedField 𝕜] (A : Type*) [NormedRing A] [NormedAlgebra 𝕜 A] [HasSummableGeomSeries A] : AnalyticAt 𝕜 (fun x : A ↦ Ring.inverse (1 - x)) 0 := @@ -798,6 +840,7 @@ lemma analyticAt_inverse_one_sub (𝕜 : Type*) [NontriviallyNormedField 𝕜] /-- If `A` is a normed algebra over `𝕜` with summable geometric series, then inversion on `A` is analytic at any unit. -/ +@[fun_prop] lemma analyticAt_inverse {𝕜 : Type*} [NontriviallyNormedField 𝕜] {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] [HasSummableGeomSeries A] (z : Aˣ) : AnalyticAt 𝕜 Ring.inverse (z : A) := by @@ -837,12 +880,14 @@ lemma hasFPowerSeriesOnBall_inv_one_sub convert hasFPowerSeriesOnBall_inverse_one_sub 𝕜 𝕝 exact Ring.inverse_eq_inv'.symm +@[fun_prop] lemma analyticAt_inv_one_sub (𝕝 : Type*) [NontriviallyNormedField 𝕝] [NormedAlgebra 𝕜 𝕝] : AnalyticAt 𝕜 (fun x : 𝕝 ↦ (1 - x)⁻¹) 0 := ⟨_, ⟨_, hasFPowerSeriesOnBall_inv_one_sub 𝕜 𝕝⟩⟩ /-- If `𝕝` is a normed field extension of `𝕜`, then the inverse map `𝕝 → 𝕝` is `𝕜`-analytic away from 0. -/ +@[fun_prop] lemma analyticAt_inv {z : 𝕝} (hz : z ≠ 0) : AnalyticAt 𝕜 Inv.inv z := by convert analyticAt_inverse (𝕜 := 𝕜) (Units.mk0 _ hz) exact Ring.inverse_eq_inv'.symm @@ -861,10 +906,16 @@ theorem AnalyticWithinAt.inv {f : E → 𝕝} {x : E} {s : Set E} (analyticAt_inv f0).comp_analyticWithinAt fa /-- `(f x)⁻¹` is analytic away from `f x = 0` -/ +@[fun_prop] theorem AnalyticAt.inv {f : E → 𝕝} {x : E} (fa : AnalyticAt 𝕜 f x) (f0 : f x ≠ 0) : - AnalyticAt 𝕜 (fun x ↦ (f x)⁻¹) x := + AnalyticAt 𝕜 f⁻¹ x := (analyticAt_inv f0).comp fa +@[fun_prop] +theorem AnalyticAt.inv' {f : E → 𝕝} {x : E} (fa : AnalyticAt 𝕜 f x) (f0 : f x ≠ 0) : + AnalyticAt 𝕜 (fun x ↦ (f x)⁻¹) x := + fa.inv f0 + /-- `(f x)⁻¹` is analytic away from `f x = 0` -/ theorem AnalyticOn.inv {f : E → 𝕝} {s : Set E} (fa : AnalyticOn 𝕜 f s) (f0 : ∀ x ∈ s, f x ≠ 0) : @@ -887,11 +938,18 @@ theorem AnalyticWithinAt.div {f g : E → 𝕝} {s : Set E} {x : E} simp_rw [div_eq_mul_inv]; exact fa.mul (ga.inv g0) /-- `f x / g x` is analytic away from `g x = 0` -/ +@[fun_prop] theorem AnalyticAt.div {f g : E → 𝕝} {x : E} (fa : AnalyticAt 𝕜 f x) (ga : AnalyticAt 𝕜 g x) (g0 : g x ≠ 0) : - AnalyticAt 𝕜 (fun x ↦ f x / g x) x := by + AnalyticAt 𝕜 (f / g) x := by simp_rw [div_eq_mul_inv]; exact fa.mul (ga.inv g0) +@[fun_prop] +theorem AnalyticAt.div' {f g : E → 𝕝} {x : E} + (fa : AnalyticAt 𝕜 f x) (ga : AnalyticAt 𝕜 g x) (g0 : g x ≠ 0) : + AnalyticAt 𝕜 (fun x ↦ f x / g x) x := + fa.div ga g0 + /-- `f x / g x` is analytic away from `g x = 0` -/ theorem AnalyticOn.div {f g : E → 𝕝} {s : Set E} (fa : AnalyticOn 𝕜 f s) (ga : AnalyticOn 𝕜 g s) (g0 : ∀ x ∈ s, g x ≠ 0) : @@ -924,6 +982,7 @@ theorem Finset.analyticWithinAt_sum {f : α → E → F} {c : E} {s : Set E} exact (h a (Or.inl rfl)).add (hB fun b m ↦ h b (Or.inr m)) /-- Finite sums of analytic functions are analytic -/ +@[fun_prop] theorem Finset.analyticAt_sum {f : α → E → F} {c : E} (N : Finset α) (h : ∀ n ∈ N, AnalyticAt 𝕜 (f n) c) : AnalyticAt 𝕜 (fun z ↦ ∑ n ∈ N, f n z) c := by @@ -958,6 +1017,7 @@ theorem Finset.analyticWithinAt_prod {A : Type*} [NormedCommRing A] [NormedAlgeb exact (h a (Or.inl rfl)).mul (hB fun b m ↦ h b (Or.inr m)) /-- Finite products of analytic functions are analytic -/ +@[fun_prop] theorem Finset.analyticAt_prod {A : Type*} [NormedCommRing A] [NormedAlgebra 𝕜 A] {f : α → E → A} {c : E} (N : Finset α) (h : ∀ n ∈ N, AnalyticAt 𝕜 (f n) c) : AnalyticAt 𝕜 (fun z ↦ ∏ n ∈ N, f n z) c := by diff --git a/Mathlib/Analysis/Analytic/IsolatedZeros.lean b/Mathlib/Analysis/Analytic/IsolatedZeros.lean index 7b5636ce8c665..0c12a9fea2373 100644 --- a/Mathlib/Analysis/Analytic/IsolatedZeros.lean +++ b/Mathlib/Analysis/Analytic/IsolatedZeros.lean @@ -7,6 +7,7 @@ import Mathlib.Analysis.Analytic.Constructions import Mathlib.Analysis.Calculus.DSlope import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Analysis.Analytic.Uniqueness +import Mathlib.Topology.Perfect /-! # Principle of isolated zeros @@ -217,6 +218,67 @@ lemma order_eq_nat_iff (hf : AnalyticAt 𝕜 f z₀) (n : ℕ) : hf.order = ↑n refine ⟨fun hn ↦ (WithTop.coe_inj.mp hn : h.choose = n) ▸ h.choose_spec, fun h' ↦ ?_⟩ rw [unique_eventuallyEq_pow_smul_nonzero h.choose_spec h'] +/- An analytic function `f` has finite order at a point `z₀` iff it locally looks + like `(z - z₀) ^ order • g`, where `g` is analytic and does not vanish at + `z₀`. -/ +lemma order_neq_top_iff (hf : AnalyticAt 𝕜 f z₀) : + hf.order ≠ ⊤ ↔ ∃ (g : 𝕜 → E), AnalyticAt 𝕜 g z₀ ∧ g z₀ ≠ 0 + ∧ f =ᶠ[𝓝 z₀] fun z ↦ (z - z₀) ^ (hf.order.toNat) • g z := by + simp only [← ENat.coe_toNat_eq_self, Eq.comm, EventuallyEq, ← hf.order_eq_nat_iff] + +/- An analytic function has order zero at a point iff it does not vanish there. -/ +lemma order_eq_zero_iff (hf : AnalyticAt 𝕜 f z₀) : + hf.order = 0 ↔ f z₀ ≠ 0 := by + rw [← ENat.coe_zero, order_eq_nat_iff hf 0] + constructor + · intro ⟨g, _, _, hg⟩ + simpa [hg.self_of_nhds] + · exact fun hz ↦ ⟨f, hf, hz, by simp⟩ + +/- An analytic function vanishes at a point if its order is nonzero when converted to ℕ. -/ +lemma apply_eq_zero_of_order_toNat_ne_zero (hf : AnalyticAt 𝕜 f z₀) : + hf.order.toNat ≠ 0 → f z₀ = 0 := by + simp [hf.order_eq_zero_iff] + tauto + +/- Helper lemma for `AnalyticAt.order_mul` -/ +lemma order_mul_of_order_eq_top {f g : 𝕜 → 𝕜} (hf : AnalyticAt 𝕜 f z₀) + (hg : AnalyticAt 𝕜 g z₀) (h'f : hf.order = ⊤) : + (hf.mul hg).order = ⊤ := by + rw [AnalyticAt.order_eq_top_iff, eventually_nhds_iff] at * + obtain ⟨t, h₁t, h₂t, h₃t⟩ := h'f + exact ⟨t, fun y hy ↦ (by simp [h₁t y hy]), h₂t, h₃t⟩ + +/-- The order is additive when multiplying analytic functions. -/ +theorem order_mul {f g : 𝕜 → 𝕜} (hf : AnalyticAt 𝕜 f z₀) (hg : AnalyticAt 𝕜 g z₀) : + (hf.mul hg).order = hf.order + hg.order := by + -- Trivial cases: one of the functions vanishes around z₀ + by_cases h₂f : hf.order = ⊤ + · simp [hf.order_mul_of_order_eq_top hg h₂f, h₂f] + by_cases h₂g : hg.order = ⊤ + · simp [mul_comm f g, hg.order_mul_of_order_eq_top hf h₂g, h₂g] + -- Non-trivial case: both functions do not vanish around z₀ + obtain ⟨g₁, h₁g₁, h₂g₁, h₃g₁⟩ := hf.order_neq_top_iff.1 h₂f + obtain ⟨g₂, h₁g₂, h₂g₂, h₃g₂⟩ := hg.order_neq_top_iff.1 h₂g + rw [← ENat.coe_toNat h₂f, ← ENat.coe_toNat h₂g, ← ENat.coe_add, (hf.mul hg).order_eq_nat_iff] + use g₁ * g₂, by exact h₁g₁.mul h₁g₂ + constructor + · simp + tauto + · obtain ⟨t, h₁t, h₂t, h₃t⟩ := eventually_nhds_iff.1 h₃g₁ + obtain ⟨s, h₁s, h₂s, h₃s⟩ := eventually_nhds_iff.1 h₃g₂ + exact eventually_nhds_iff.2 + ⟨t ∩ s, fun y hy ↦ (by simp [h₁t y hy.1, h₁s y hy.2]; ring), h₂t.inter h₂s, h₃t, h₃s⟩ + +/-- The order multiplies by `n` when taking an analytic function to its `n`th power. -/ +theorem order_pow {f : 𝕜 → 𝕜} (hf : AnalyticAt 𝕜 f z₀) {n : ℕ} : + (hf.pow n).order = n • hf.order := by + induction n + case zero => + simp [AnalyticAt.order_eq_zero_iff] + case succ n hn => + simp [add_mul, pow_add, (hf.pow n).order_mul hf, hn] + end AnalyticAt namespace AnalyticOnNhd @@ -283,4 +345,46 @@ theorem eq_of_frequently_eq [ConnectedSpace 𝕜] (hf : AnalyticOnNhd 𝕜 f uni @[deprecated (since := "2024-09-26")] alias _root_.AnalyticOn.eq_of_frequently_eq := eq_of_frequently_eq +section Mul +/-! +### Vanishing of products of analytic functions +-/ + +variable {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] + {B : Type*} [NormedAddCommGroup B] [NormedSpace 𝕜 B] [Module A B] + +/-- If `f, g` are analytic on a neighbourhood of the preconnected open set `U`, and `f • g = 0` +on `U`, then either `f = 0` on `U` or `g = 0` on `U`. -/ +lemma eq_zero_or_eq_zero_of_smul_eq_zero [NoZeroSMulDivisors A B] + {f : 𝕜 → A} {g : 𝕜 → B} (hf : AnalyticOnNhd 𝕜 f U) (hg : AnalyticOnNhd 𝕜 g U) + (hfg : ∀ z ∈ U, f z • g z = 0) (hU : IsPreconnected U) : + (∀ z ∈ U, f z = 0) ∨ (∀ z ∈ U, g z = 0) := by + -- We want to apply `IsPreconnected.preperfect_of_nontrivial` which requires `U` to have at least + -- two elements. So we need to dispose of the cases `#U = 0` and `#U = 1` first. + by_cases hU' : U = ∅ + · simp [hU'] + obtain ⟨z, hz⟩ : ∃ z, z ∈ U := nonempty_iff_ne_empty.mpr hU' + by_cases hU'' : U = {z} + · simpa [hU''] using hfg z hz + apply (nontrivial_iff_ne_singleton hz).mpr at hU'' + -- Now connectedness implies that `z` is an accumulation point of `U`, so at least one of + -- `f` and `g` must vanish frequently in a neighbourhood of `z`. + have : ∃ᶠ w in 𝓝[≠] z, w ∈ U := + frequently_mem_iff_neBot.mpr <| hU.preperfect_of_nontrivial hU'' z hz + have : ∃ᶠ w in 𝓝[≠] z, f w = 0 ∨ g w = 0 := + this.mp <| by filter_upwards with w hw using smul_eq_zero.mp (hfg w hw) + cases frequently_or_distrib.mp this with + | inl h => exact Or.inl <| hf.eqOn_zero_of_preconnected_of_frequently_eq_zero hU hz h + | inr h => exact Or.inr <| hg.eqOn_zero_of_preconnected_of_frequently_eq_zero hU hz h + +/-- If `f, g` are analytic on a neighbourhood of the preconnected open set `U`, and `f * g = 0` +on `U`, then either `f = 0` on `U` or `g = 0` on `U`. -/ +lemma eq_zero_or_eq_zero_of_mul_eq_zero [NoZeroDivisors A] + {f g : 𝕜 → A} (hf : AnalyticOnNhd 𝕜 f U) (hg : AnalyticOnNhd 𝕜 g U) + (hfg : ∀ z ∈ U, f z * g z = 0) (hU : IsPreconnected U) : + (∀ z ∈ U, f z = 0) ∨ (∀ z ∈ U, g z = 0) := + eq_zero_or_eq_zero_of_smul_eq_zero hf hg hfg hU + +end Mul + end AnalyticOnNhd diff --git a/Mathlib/Analysis/Analytic/Linear.lean b/Mathlib/Analysis/Analytic/Linear.lean index bd405daef5895..5b7595bccbaa9 100644 --- a/Mathlib/Analysis/Analytic/Linear.lean +++ b/Mathlib/Analysis/Analytic/Linear.lean @@ -136,6 +136,7 @@ end ContinuousLinearMap variable {s : Set E} {z : E} {t : Set (E × F)} {p : E × F} +@[fun_prop] lemma analyticAt_id : AnalyticAt 𝕜 (id : E → E) z := (ContinuousLinearMap.id 𝕜 E).analyticAt z diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index 3ed9fadaa6e5c..86a7ea98ddcc0 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -26,9 +26,11 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-- Meromorphy of `f` at `x` (more precisely, on a punctured neighbourhood of `x`; the value at `x` itself is irrelevant). -/ +@[fun_prop] def MeromorphicAt (f : 𝕜 → E) (x : 𝕜) := ∃ (n : ℕ), AnalyticAt 𝕜 (fun z ↦ (z - x) ^ n • f z) x +@[fun_prop] lemma AnalyticAt.meromorphicAt {f : 𝕜 → E} {x : 𝕜} (hf : AnalyticAt 𝕜 f x) : MeromorphicAt f x := ⟨0, by simpa only [pow_zero, one_smul]⟩ @@ -37,9 +39,11 @@ namespace MeromorphicAt lemma id (x : 𝕜) : MeromorphicAt id x := analyticAt_id.meromorphicAt +@[fun_prop] lemma const (e : E) (x : 𝕜) : MeromorphicAt (fun _ ↦ e) x := analyticAt_const.meromorphicAt +@[fun_prop] lemma add {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : MeromorphicAt (f + g) x := by rcases hf with ⟨m, hf⟩ @@ -53,35 +57,63 @@ lemma add {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : Meromorph exact (((analyticAt_id.sub analyticAt_const).pow _).smul hf).add (((analyticAt_id.sub analyticAt_const).pow _).smul hg) +@[fun_prop] +lemma add' {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + MeromorphicAt (fun z ↦ f z + g z) x := + hf.add hg + +@[fun_prop] lemma smul {f : 𝕜 → 𝕜} {g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : MeromorphicAt (f • g) x := by rcases hf with ⟨m, hf⟩ rcases hg with ⟨n, hg⟩ refine ⟨m + n, ?_⟩ - convert hf.smul hg using 2 with z + convert hf.smul' hg using 2 with z rw [Pi.smul_apply', smul_eq_mul] module +@[fun_prop] +lemma smul' {f : 𝕜 → 𝕜} {g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + MeromorphicAt (fun z ↦ f z • g z) x := + hf.smul hg + +@[fun_prop] lemma mul {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : MeromorphicAt (f * g) x := hf.smul hg +@[fun_prop] +lemma mul' {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + MeromorphicAt (fun z ↦ f z * g z) x := + hf.smul hg + +@[fun_prop] lemma neg {f : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) : MeromorphicAt (-f) x := by convert (MeromorphicAt.const (-1 : 𝕜) x).smul hf using 1 ext1 z simp only [Pi.neg_apply, Pi.smul_apply', neg_smul, one_smul] +@[fun_prop] +lemma neg' {f : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) : MeromorphicAt (fun z ↦ -f z) x := + hf.neg + @[simp] lemma neg_iff {f : 𝕜 → E} {x : 𝕜} : MeromorphicAt (-f) x ↔ MeromorphicAt f x := ⟨fun h ↦ by simpa only [neg_neg] using h.neg, MeromorphicAt.neg⟩ +@[fun_prop] lemma sub {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : MeromorphicAt (f - g) x := by convert hf.add hg.neg using 1 ext1 z simp_rw [Pi.sub_apply, Pi.add_apply, Pi.neg_apply, sub_eq_add_neg] +@[fun_prop] +lemma sub' {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + MeromorphicAt (fun z ↦ f z - g z) x := + hf.sub hg + /-- With our definitions, `MeromorphicAt f x` depends only on the values of `f` on a punctured neighbourhood of `x` (not on `f x`) -/ lemma congr {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hfg : f =ᶠ[𝓝[≠] x] g) : @@ -89,13 +121,14 @@ lemma congr {f g : 𝕜 → E} {x : 𝕜} (hf : MeromorphicAt f x) (hfg : f =ᶠ rcases hf with ⟨m, hf⟩ refine ⟨m + 1, ?_⟩ have : AnalyticAt 𝕜 (fun z ↦ z - x) x := analyticAt_id.sub analyticAt_const - refine (this.smul hf).congr ?_ + refine (this.smul' hf).congr ?_ rw [eventuallyEq_nhdsWithin_iff] at hfg filter_upwards [hfg] with z hz rcases eq_or_ne z x with rfl | hn · simp · rw [hz (Set.mem_compl_singleton_iff.mp hn), pow_succ', mul_smul] +@[fun_prop] lemma inv {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) : MeromorphicAt f⁻¹ x := by rcases hf with ⟨m, hf⟩ by_cases h_eq : (fun z ↦ (z - x) ^ m • f z) =ᶠ[𝓝 x] 0 @@ -110,7 +143,7 @@ lemma inv {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) : MeromorphicA have : AnalyticAt 𝕜 (fun z ↦ (z - x) ^ (m + 1)) x := (analyticAt_id.sub analyticAt_const).pow _ -- use `m + 1` rather than `m` to damp out any silly issues with the value at `z = x` - refine ⟨n + 1, (this.smul <| hg_an.inv hg_ne).congr ?_⟩ + refine ⟨n + 1, (this.smul' <| hg_an.inv hg_ne).congr ?_⟩ filter_upwards [hg_eq, hg_an.continuousAt.eventually_ne hg_ne] with z hfg hg_ne' rcases eq_or_ne z x with rfl | hz_ne · simp only [sub_self, pow_succ, mul_zero, zero_smul] @@ -123,25 +156,47 @@ lemma inv {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) : MeromorphicA rw [pow_succ', mul_assoc, hfg] ring +@[fun_prop] +lemma inv' {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) : MeromorphicAt (fun z ↦ (f z)⁻¹) x := + hf.inv + @[simp] lemma inv_iff {f : 𝕜 → 𝕜} {x : 𝕜} : MeromorphicAt f⁻¹ x ↔ MeromorphicAt f x := ⟨fun h ↦ by simpa only [inv_inv] using h.inv, MeromorphicAt.inv⟩ +@[fun_prop] lemma div {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : MeromorphicAt (f / g) x := (div_eq_mul_inv f g).symm ▸ (hf.mul hg.inv) +@[fun_prop] +lemma div' {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : MeromorphicAt g x) : + MeromorphicAt (fun z ↦ f z / g z) x := + hf.div hg + +@[fun_prop] lemma pow {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℕ) : MeromorphicAt (f ^ n) x := by induction n with | zero => simpa only [pow_zero] using MeromorphicAt.const 1 x | succ m hm => simpa only [pow_succ] using hm.mul hf +@[fun_prop] +lemma pow' {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℕ) : + MeromorphicAt (fun z ↦ (f z) ^ n) x := + hf.pow n + +@[fun_prop] lemma zpow {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℤ) : MeromorphicAt (f ^ n) x := by induction n with | ofNat m => simpa only [Int.ofNat_eq_coe, zpow_natCast] using hf.pow m | negSucc m => simpa only [zpow_negSucc, inv_iff] using hf.pow (m + 1) +@[fun_prop] +lemma zpow' {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℤ) : + MeromorphicAt (fun z ↦ (f z) ^ n) x := + hf.zpow n + theorem eventually_analyticAt [CompleteSpace E] {f : 𝕜 → E} {x : 𝕜} (h : MeromorphicAt f x) : ∀ᶠ y in 𝓝[≠] x, AnalyticAt 𝕜 f y := by rw [MeromorphicAt] at h diff --git a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean index aca2bbc1fa356..e444c0049889b 100644 --- a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean +++ b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean @@ -66,8 +66,9 @@ section NormedAddCommGroup variable {α β : Type*} [NormedAddCommGroup β] -/-- Two functions `u` and `v` are said to be asymptotically equivalent along a filter `l` when - `u x - v x = o(v x)` as `x` converges along `l`. -/ +/-- Two functions `u` and `v` are said to be asymptotically equivalent along a filter `l` + (denoted as `u ~[l] v` in the `Asymptotics` namespace) + when `u x - v x = o(v x)` as `x` converges along `l`. -/ def IsEquivalent (l : Filter α) (u v : α → β) := (u - v) =o[l] v diff --git a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean index e1c0e57e890c5..00668c051caf3 100644 --- a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean @@ -3,9 +3,8 @@ 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.Analysis.Normed.Order.Basic -import Mathlib.Analysis.Asymptotics.Lemmas -import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.Analysis.Asymptotics.AsymptoticEquivalent +import Mathlib.Analysis.SpecificLimits.Basic /-! # A collection of specific asymptotic results @@ -143,3 +142,19 @@ theorem Filter.Tendsto.cesaro {u : ℕ → ℝ} {l : ℝ} (h : Tendsto u atTop ( h.cesaro_smul end Real + +section NormedLinearOrderedField + +variable {R : Type*} [NormedLinearOrderedField R] [OrderTopology R] [FloorRing R] + +theorem Asymptotics.isEquivalent_nat_floor : + (fun (x : R) ↦ ↑⌊x⌋₊) ~[atTop] (fun x ↦ x) := by + refine isEquivalent_of_tendsto_one ?_ tendsto_nat_floor_div_atTop + filter_upwards with x hx using by rw [hx, Nat.floor_zero, Nat.cast_eq_zero] + +theorem Asymptotics.isEquivalent_nat_ceil : + (fun (x : R) ↦ ↑⌈x⌉₊) ~[atTop] (fun x ↦ x) := by + refine isEquivalent_of_tendsto_one ?_ tendsto_nat_ceil_div_atTop + filter_upwards with x hx using by rw [hx, Nat.ceil_zero, Nat.cast_eq_zero] + +end NormedLinearOrderedField diff --git a/Mathlib/Analysis/Asymptotics/TVS.lean b/Mathlib/Analysis/Asymptotics/TVS.lean index 5d0db2d9253a3..59720fc3eebdd 100644 --- a/Mathlib/Analysis/Asymptotics/TVS.lean +++ b/Mathlib/Analysis/Asymptotics/TVS.lean @@ -193,6 +193,17 @@ lemma IsLittleOTVS.insert [TopologicalSpace α] {x : α} {s : Set α} lemma IsLittleOTVS.bot : f =o[𝕜;⊥] g := fun u hU => ⟨univ, by simp⟩ +theorem IsLittleOTVS.add [TopologicalAddGroup E] [ContinuousSMul 𝕜 E] + {f₁ f₂ : α → E} {g : α → F} {l : Filter α} + (h₁ : f₁ =o[𝕜; l] g) (h₂ : f₂ =o[𝕜; l] g) : (f₁ + f₂) =o[𝕜; l] g := by + rw [(nhds_basis_balanced 𝕜 E).add_self.isLittleOTVS_iff (basis_sets _)] + rintro U ⟨hU, hUb⟩ + rcases ((h₁.eventually_smallSets U hU).and (h₂.eventually_smallSets U hU)).exists_mem_of_smallSets + with ⟨V, hV, hVf₁, hVf₂⟩ + refine ⟨V, hV, fun ε hε ↦ ?_⟩ + filter_upwards [hVf₁ ε hε, hVf₂ ε hε] with x hx₁ hx₂ + exact (egauge_add_add_le hUb hUb _ _).trans (max_le hx₁ hx₂) + protected lemma IsLittleOTVS.smul_left (h : f =o[𝕜;l] g) (c : α → 𝕜) : (fun x ↦ c x • f x) =o[𝕜;l] (fun x ↦ c x • g x) := by unfold IsLittleOTVS at * diff --git a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean index 0479cfb3716ed..ddb1eca0f9d22 100644 --- a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean +++ b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean @@ -5,7 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order import Mathlib.Analysis.CStarAlgebra.SpecialFunctions.PosPart -import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow.Basic import Mathlib.Topology.ApproximateUnit /-! # Nonnegative contractions in a C⋆-algebra form an approximate unit diff --git a/Mathlib/Analysis/CStarAlgebra/Basic.lean b/Mathlib/Analysis/CStarAlgebra/Basic.lean index cd2cb5b8628a8..4338e91c5e23b 100644 --- a/Mathlib/Analysis/CStarAlgebra/Basic.lean +++ b/Mathlib/Analysis/CStarAlgebra/Basic.lean @@ -177,9 +177,6 @@ section Unital variable [NormedRing E] [StarRing E] [CStarRing E] --- Porting note https://github.com/leanprover-community/mathlib4/issues/10959 --- simp cannot prove this -@[simp, nolint simpNF] theorem norm_one [Nontrivial E] : ‖(1 : E)‖ = 1 := by have : 0 < ‖(1 : E)‖ := norm_pos_iff.mpr one_ne_zero rw [← mul_left_inj' this.ne', ← norm_star_mul_self, mul_one, star_one, one_mul] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean index 080dd2282f617..c32eadf334515 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Isometric.lean @@ -140,6 +140,26 @@ lemma nnnorm_cfc_lt_iff (f : 𝕜 → 𝕜) (a : A) {c : ℝ≥0} (hc : 0 < c) (ha : p a := by cfc_tac) : ‖cfc f a‖₊ < c ↔ ∀ x ∈ σ 𝕜 a, ‖f x‖₊ < c := norm_cfc_lt_iff f a hc +namespace IsometricContinuousFunctionalCalculus + +lemma isGreatest_norm_spectrum [Nontrivial A] (a : A) (ha : p a := by cfc_tac) : + IsGreatest ((‖·‖) '' spectrum 𝕜 a) ‖a‖ := by + simpa only [cfc_id 𝕜 a] using IsGreatest.norm_cfc (id : 𝕜 → 𝕜) a + +lemma norm_spectrum_le (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σ 𝕜 a) (ha : p a := by cfc_tac) : + ‖x‖ ≤ ‖a‖ := by + simpa only [cfc_id 𝕜 a] using norm_apply_le_norm_cfc (id : 𝕜 → 𝕜) a hx + +lemma isGreatest_nnnorm_spectrum [Nontrivial A] (a : A) (ha : p a := by cfc_tac) : + IsGreatest ((‖·‖₊) '' spectrum 𝕜 a) ‖a‖₊ := by + simpa only [cfc_id 𝕜 a] using IsGreatest.nnnorm_cfc (id : 𝕜 → 𝕜) a + +lemma nnnorm_spectrum_le (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σ 𝕜 a) (ha : p a := by cfc_tac) : + ‖x‖₊ ≤ ‖a‖₊ := by + simpa only [cfc_id 𝕜 a] using nnnorm_apply_le_nnnorm_cfc (id : 𝕜 → 𝕜) a hx + +end IsometricContinuousFunctionalCalculus + end NormedRing namespace SpectrumRestricts @@ -307,6 +327,26 @@ lemma nnnorm_cfcₙ_lt_iff (f : 𝕜 → 𝕜) (a : A) (c : ℝ≥0) (ha : p a := by cfc_tac) : ‖cfcₙ f a‖₊ < c ↔ ∀ x ∈ σₙ 𝕜 a, ‖f x‖₊ < c := norm_cfcₙ_lt_iff f a c.1 hf hf₀ ha +namespace NonUnitalIsometricContinuousFunctionalCalculus + +lemma isGreatest_norm_quasispectrum (a : A) (ha : p a := by cfc_tac) : + IsGreatest ((‖·‖) '' σₙ 𝕜 a) ‖a‖ := by + simpa only [cfcₙ_id 𝕜 a] using IsGreatest.norm_cfcₙ (id : 𝕜 → 𝕜) a + +lemma norm_quasispectrum_le (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σₙ 𝕜 a) (ha : p a := by cfc_tac) : + ‖x‖ ≤ ‖a‖ := by + simpa only [cfcₙ_id 𝕜 a] using norm_apply_le_norm_cfcₙ (id : 𝕜 → 𝕜) a hx + +lemma isGreatest_nnnorm_quasispectrum (a : A) (ha : p a := by cfc_tac) : + IsGreatest ((‖·‖₊) '' σₙ 𝕜 a) ‖a‖₊ := by + simpa only [cfcₙ_id 𝕜 a] using IsGreatest.nnnorm_cfcₙ (id : 𝕜 → 𝕜) a + +lemma nnnorm_quasispectrum_le (a : A) ⦃x : 𝕜⦄ (hx : x ∈ σₙ 𝕜 a) (ha : p a := by cfc_tac) : + ‖x‖₊ ≤ ‖a‖₊ := by + simpa only [cfcₙ_id 𝕜 a] using nnnorm_apply_le_nnnorm_cfcₙ (id : 𝕜 → 𝕜) a hx + +end NonUnitalIsometricContinuousFunctionalCalculus + end NormedRing namespace QuasispectrumRestricts @@ -457,6 +497,24 @@ lemma nnnorm_cfc_nnreal_lt_iff (f : ℝ≥0 → ℝ≥0) (a : A) {c : ℝ≥0} ( (ha : 0 ≤ a := by cfc_tac) : ‖cfc f a‖₊ < c ↔ ∀ x ∈ σ ℝ≥0 a, f x < c := ⟨fun h _ hx ↦ apply_le_nnnorm_cfc_nnreal f a hx hf ha |>.trans_lt h, nnnorm_cfc_nnreal_lt hc⟩ +namespace IsometricContinuousFunctionalCalculus + +lemma isGreatest_spectrum [Nontrivial A] (a : A) (ha : 0 ≤ a := by cfc_tac) : + IsGreatest (σ ℝ≥0 a) ‖a‖₊ := by + simpa [cfc_id ℝ≥0 a] using IsGreatest.nnnorm_cfc_nnreal id a + +lemma spectrum_le (a : A) ⦃x : ℝ≥0⦄ (hx : x ∈ σ ℝ≥0 a) (ha : 0 ≤ a := by cfc_tac) : + x ≤ ‖a‖₊ := by + simpa [cfc_id ℝ≥0 a] using apply_le_nnnorm_cfc_nnreal id a hx + +end IsometricContinuousFunctionalCalculus + +open IsometricContinuousFunctionalCalculus in +lemma MonotoneOn.nnnorm_cfc [Nontrivial A] (f : ℝ≥0 → ℝ≥0) (a : A) + (hf : MonotoneOn f (σ ℝ≥0 a)) (hf₂ : ContinuousOn f (σ ℝ≥0 a) := by cfc_cont_tac) + (ha : 0 ≤ a := by cfc_tac) : ‖cfc f a‖₊ = f ‖a‖₊ := + IsGreatest.nnnorm_cfc_nnreal f a |>.unique <| hf.map_isGreatest (isGreatest_spectrum a) + end Unital section NonUnital @@ -511,6 +569,61 @@ lemma nnnorm_cfcₙ_nnreal_lt_iff (f : ℝ≥0 → ℝ≥0) (a : A) (c : ℝ≥0 (ha : 0 ≤ a := by cfc_tac) : ‖cfcₙ f a‖₊ < c ↔ ∀ x ∈ σₙ ℝ≥0 a, f x < c := ⟨fun h _ hx ↦ apply_le_nnnorm_cfcₙ_nnreal f a hx hf hf₀ ha |>.trans_lt h, nnnorm_cfcₙ_nnreal_lt⟩ +namespace NonUnitalIsometricContinuousFunctionalCalculus + +lemma isGreatest_quasispectrum (a : A) (ha : 0 ≤ a := by cfc_tac) : + IsGreatest (σₙ ℝ≥0 a) ‖a‖₊ := by + simpa [cfcₙ_id ℝ≥0 a] using IsGreatest.nnnorm_cfcₙ_nnreal id a + +lemma quasispectrum_le (a : A) ⦃x : ℝ≥0⦄ (hx : x ∈ σₙ ℝ≥0 a) (ha : 0 ≤ a := by cfc_tac) : + x ≤ ‖a‖₊ := by + simpa [cfcₙ_id ℝ≥0 a] using apply_le_nnnorm_cfcₙ_nnreal id a hx + +end NonUnitalIsometricContinuousFunctionalCalculus + +open NonUnitalIsometricContinuousFunctionalCalculus in +lemma MonotoneOn.nnnorm_cfcₙ (f : ℝ≥0 → ℝ≥0) (a : A) + (hf : MonotoneOn f (σₙ ℝ≥0 a)) (hf₂ : ContinuousOn f (σₙ ℝ≥0 a) := by cfc_cont_tac) + (hf0 : f 0 = 0 := by cfc_zero_tac) (ha : 0 ≤ a := by cfc_tac) : + ‖cfcₙ f a‖₊ = f ‖a‖₊ := + IsGreatest.nnnorm_cfcₙ_nnreal f a |>.unique <| hf.map_isGreatest (isGreatest_quasispectrum a) + end NonUnital end NNReal + +/-! ### Non-unital instance for unital algebras -/ + +namespace IsometricContinuousFunctionalCalculus + +variable {𝕜 A : Type*} {p : outParam (A → Prop)} +variable [RCLike 𝕜] [NormedRing A] [StarRing A] [NormedAlgebra 𝕜 A] +variable [IsometricContinuousFunctionalCalculus 𝕜 A p] + +open scoped ContinuousFunctionalCalculus in +/-- An isometric continuous functional calculus on a unital algebra yields to a non-unital one. -/ +instance toNonUnital : NonUnitalIsometricContinuousFunctionalCalculus 𝕜 A p where + isometric a ha := by + have : CompactSpace (σₙ 𝕜 a) := by + have h_cpct : CompactSpace (spectrum 𝕜 a) := inferInstance + simp only [← isCompact_iff_compactSpace, quasispectrum_eq_spectrum_union_zero] at h_cpct ⊢ + exact h_cpct |>.union isCompact_singleton + rw [cfcₙHom_eq_cfcₙHom_of_cfcHom, cfcₙHom_of_cfcHom] + refine isometry_cfcHom a |>.comp ?_ + simp only [MulHom.coe_coe, NonUnitalStarAlgHom.coe_toNonUnitalAlgHom, + NonUnitalStarAlgHom.coe_coe] + refine AddMonoidHomClass.isometry_of_norm _ fun f ↦ ?_ + let ι : C(σ 𝕜 a, σₙ 𝕜 a) := ⟨_, continuous_inclusion <| spectrum_subset_quasispectrum 𝕜 a⟩ + show ‖(f : C(σₙ 𝕜 a, 𝕜)).comp ι‖ = ‖(f : C(σₙ 𝕜 a, 𝕜))‖ + apply le_antisymm (ContinuousMap.norm_le _ (by positivity) |>.mpr ?_) + (ContinuousMap.norm_le _ (by positivity) |>.mpr ?_) + · rintro ⟨x, hx⟩ + exact (f : C(σₙ 𝕜 a, 𝕜)).norm_coe_le_norm ⟨x, spectrum_subset_quasispectrum 𝕜 a hx⟩ + · rintro ⟨x, hx⟩ + obtain (rfl | hx') : x = 0 ∨ x ∈ σ 𝕜 a := by + simpa [quasispectrum_eq_spectrum_union_zero] using hx + · show ‖f 0‖ ≤ _ + simp + · exact (f : C(σₙ 𝕜 a, 𝕜)).comp ι |>.norm_coe_le_norm ⟨x, hx'⟩ + +end IsometricContinuousFunctionalCalculus diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index c829c663a59c8..c093bde67133c 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -6,8 +6,9 @@ Authors: Frédéric Dupuis import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic import Mathlib.Analysis.CStarAlgebra.Unitization -import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow.Basic import Mathlib.Topology.ContinuousMap.StarOrdered +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Isometric /-! # Facts about star-ordered rings that depend on the continuous functional calculus @@ -478,7 +479,8 @@ section Pow namespace CStarAlgebra -variable {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] +variable {A : Type*} {B : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + [NonUnitalCStarAlgebra B] [PartialOrder B] [StarOrderedRing B] lemma pow_nonneg {a : A} (ha : 0 ≤ a := by cfc_tac) (n : ℕ) : 0 ≤ a ^ n := by rw [← cfc_pow_id (R := ℝ≥0) a] diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean index 4dae3f6a8222d..95c7421f62a16 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean @@ -47,8 +47,8 @@ The pattern here is the same one as is used by `Lex` for order structures; it av separate synonym for each type, and allows all the structure-copying code to be shared. -/ -/-- A type synonym for endowing a given type with a `CStarModule` structure. This has the scoped -notation `C⋆ᵐᵒᵈ`. +/-- A type synonym for endowing a given type with a `CStarModule` structure. +This has the scoped notation `C⋆ᵐᵒᵈ` in the WithCStarModule namespace. Note: because the C⋆-algebra `A` over which `E` is a `CStarModule` is listed as an `outParam` in that class, we don't pass it as an unused argument to `WithCStarModule`, unlike the `p` parameter diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean index 922effc6f97d8..2f7721e6ffc3c 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean @@ -74,7 +74,6 @@ add more properties if they are useful and satisfied in the examples of inner pr and finite dimensional vector spaces, notably derivative norm control in terms of `R - 1`. TODO: do we ever need `f x = 1 ↔ ‖x‖ ≤ 1`? -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not yet ported; was @[nolint has_nonempty_instance] structure ContDiffBumpBase (E : Type*) [NormedAddCommGroup E] [NormedSpace ℝ E] where /-- The function underlying this family of bump functions -/ toFun : ℝ → E → ℝ diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 6eecb4ca014a6..d433760da7719 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1818,6 +1818,13 @@ theorem contDiffAt_map_inverse [CompleteSpace E] (e : E ≃L[𝕜] F) : convert contDiffAt_ring_inverse 𝕜 (1 : (E →L[𝕜] E)ˣ) simp [O₂, one_def] +/-- At an invertible map `e : M →L[R] M₂` between Banach spaces, the operation of +inversion is `C^n`, for all `n`. -/ +theorem ContinuousLinearMap.IsInvertible.contDiffAt_map_inverse [CompleteSpace E] {e : E →L[𝕜] F} + (he : e.IsInvertible) : ContDiffAt 𝕜 n inverse e := by + rcases he with ⟨M, rfl⟩ + exact _root_.contDiffAt_map_inverse M + end MapInverse section FunctionInverse diff --git a/Mathlib/Analysis/Calculus/Deriv/Add.lean b/Mathlib/Analysis/Calculus/Deriv/Add.lean index 3d6b994ca666b..d5ed6db3ab4cd 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Add.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Add.lean @@ -52,10 +52,12 @@ nonrec theorem HasDerivAt.add (hf : HasDerivAt f f' x) (hg : HasDerivAt g g' x) HasDerivAt (fun x => f x + g x) (f' + g') x := hf.add hg -theorem derivWithin_add (hxs : UniqueDiffWithinAt 𝕜 s x) (hf : DifferentiableWithinAt 𝕜 f s x) +theorem derivWithin_add (hf : DifferentiableWithinAt 𝕜 f s x) (hg : DifferentiableWithinAt 𝕜 g s x) : - derivWithin (fun y => f y + g y) s x = derivWithin f s x + derivWithin g s x := - (hf.hasDerivWithinAt.add hg.hasDerivWithinAt).derivWithin hxs + derivWithin (fun y => f y + g y) s x = derivWithin f s x + derivWithin g s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hf.hasDerivWithinAt.add hg.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_add (hf : DifferentiableAt 𝕜 f x) (hg : DifferentiableAt 𝕜 g x) : @@ -78,9 +80,11 @@ nonrec theorem HasDerivAt.add_const (hf : HasDerivAt f f' x) (c : F) : HasDerivAt (fun x => f x + c) f' x := hf.add_const c -theorem derivWithin_add_const (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : +theorem derivWithin_add_const (c : F) : derivWithin (fun y => f y + c) s x = derivWithin f s x := by - simp only [derivWithin, fderivWithin_add_const hxs] + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · simp only [derivWithin, fderivWithin_add_const hxs] + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_add_const (c : F) : deriv (fun y => f y + c) x = deriv f x := by simp only [deriv, fderiv_add_const] @@ -105,9 +109,11 @@ nonrec theorem HasDerivAt.const_add (c : F) (hf : HasDerivAt f f' x) : HasDerivAt (fun x => c + f x) f' x := hf.const_add c -theorem derivWithin_const_add (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : +theorem derivWithin_const_add (c : F) : derivWithin (fun y => c + f y) s x = derivWithin f s x := by - simp only [derivWithin, fderivWithin_const_add hxs] + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · simp only [derivWithin, fderivWithin_const_add hxs] + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_const_add (c : F) : deriv (fun y => c + f y) x = deriv f x := by simp only [deriv, fderiv_const_add] @@ -160,10 +166,11 @@ theorem HasDerivAt.sum (h : ∀ i ∈ u, HasDerivAt (A i) (A' i) x) : HasDerivAt (fun y => ∑ i ∈ u, A i y) (∑ i ∈ u, A' i) x := HasDerivAtFilter.sum h -theorem derivWithin_sum (hxs : UniqueDiffWithinAt 𝕜 s x) - (h : ∀ i ∈ u, DifferentiableWithinAt 𝕜 (A i) s x) : - derivWithin (fun y => ∑ i ∈ u, A i y) s x = ∑ i ∈ u, derivWithin (A i) s x := - (HasDerivWithinAt.sum fun i hi => (h i hi).hasDerivWithinAt).derivWithin hxs +theorem derivWithin_sum (h : ∀ i ∈ u, DifferentiableWithinAt 𝕜 (A i) s x) : + derivWithin (fun y => ∑ i ∈ u, A i y) s x = ∑ i ∈ u, derivWithin (A i) s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (HasDerivWithinAt.sum fun i hi => (h i hi).hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_sum (h : ∀ i ∈ u, DifferentiableAt 𝕜 (A i) x) : @@ -189,9 +196,10 @@ nonrec theorem HasDerivAt.neg (h : HasDerivAt f f' x) : HasDerivAt (fun x => -f nonrec theorem HasStrictDerivAt.neg (h : HasStrictDerivAt f f' x) : HasStrictDerivAt (fun x => -f x) (-f') x := by simpa using h.neg.hasStrictDerivAt -theorem derivWithin.neg (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (fun y => -f y) s x = -derivWithin f s x := by - simp only [derivWithin, fderivWithin_neg hxs, ContinuousLinearMap.neg_apply] +theorem derivWithin.neg : derivWithin (fun y => -f y) s x = -derivWithin f s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · simp only [derivWithin, fderivWithin_neg hxs, ContinuousLinearMap.neg_apply] + · simp [derivWithin_zero_of_isolated hxs] theorem deriv.neg : deriv (fun y => -f y) x = -deriv f x := by simp only [deriv, fderiv_neg, ContinuousLinearMap.neg_apply] @@ -276,10 +284,12 @@ theorem HasStrictDerivAt.sub (hf : HasStrictDerivAt f f' x) (hg : HasStrictDeriv HasStrictDerivAt (fun x => f x - g x) (f' - g') x := by simpa only [sub_eq_add_neg] using hf.add hg.neg -theorem derivWithin_sub (hxs : UniqueDiffWithinAt 𝕜 s x) (hf : DifferentiableWithinAt 𝕜 f s x) +theorem derivWithin_sub (hf : DifferentiableWithinAt 𝕜 f s x) (hg : DifferentiableWithinAt 𝕜 g s x) : - derivWithin (fun y => f y - g y) s x = derivWithin f s x - derivWithin g s x := - (hf.hasDerivWithinAt.sub hg.hasDerivWithinAt).derivWithin hxs + derivWithin (fun y => f y - g y) s x = derivWithin f s x - derivWithin g s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hf.hasDerivWithinAt.sub hg.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_sub (hf : DifferentiableAt 𝕜 f x) (hg : DifferentiableAt 𝕜 g x) : @@ -298,9 +308,11 @@ nonrec theorem HasDerivAt.sub_const (hf : HasDerivAt f f' x) (c : F) : HasDerivAt (fun x => f x - c) f' x := hf.sub_const c -theorem derivWithin_sub_const (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : +theorem derivWithin_sub_const (c : F) : derivWithin (fun y => f y - c) s x = derivWithin f s x := by - simp only [derivWithin, fderivWithin_sub_const hxs] + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · simp only [derivWithin, fderivWithin_sub_const hxs] + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_sub_const (c : F) : deriv (fun y => f y - c) x = deriv f x := by simp only [deriv, fderiv_sub_const] @@ -321,13 +333,14 @@ nonrec theorem HasDerivAt.const_sub (c : F) (hf : HasDerivAt f f' x) : HasDerivAt (fun x => c - f x) (-f') x := hf.const_sub c -theorem derivWithin_const_sub (hxs : UniqueDiffWithinAt 𝕜 s x) (c : F) : +theorem derivWithin_const_sub (c : F) : derivWithin (fun y => c - f y) s x = -derivWithin f s x := by - simp [derivWithin, fderivWithin_const_sub hxs] + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · simp [derivWithin, fderivWithin_const_sub hxs] + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_const_sub (c : F) : deriv (fun y => c - f y) x = -deriv f x := by - simp only [← derivWithin_univ, - derivWithin_const_sub (uniqueDiffWithinAt_univ : UniqueDiffWithinAt 𝕜 _ _)] + simp only [← derivWithin_univ, derivWithin_const_sub] lemma differentiableAt_comp_sub_const {a b : 𝕜} : DifferentiableAt 𝕜 (fun x ↦ f (x - b)) a ↔ DifferentiableAt 𝕜 f (a - b) := by diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index 64f66f524e715..9f7a163866129 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -658,6 +658,9 @@ theorem deriv_const' : (deriv fun _ : 𝕜 => c) = fun _ => 0 := theorem derivWithin_const : derivWithin (fun _ => c) s = 0 := by ext; simp [derivWithin] +@[simp] +theorem derivWithin_zero : derivWithin (0 : 𝕜 → F) s = 0 := derivWithin_const _ _ + end Const section Continuous diff --git a/Mathlib/Analysis/Calculus/Deriv/Comp.lean b/Mathlib/Analysis/Calculus/Deriv/Comp.lean index 015bd0a24acb8..b8fe2a15b8dee 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Comp.lean @@ -127,15 +127,17 @@ theorem HasDerivAt.scomp_hasDerivWithinAt_of_eq (hg : HasDerivAt g₁ g₁' y) rw [hy] at hg; exact hg.scomp_hasDerivWithinAt x hh theorem derivWithin.scomp (hg : DifferentiableWithinAt 𝕜' g₁ t' (h x)) - (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s t') (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (g₁ ∘ h) s x = derivWithin h s x • derivWithin g₁ t' (h x) := - (HasDerivWithinAt.scomp x hg.hasDerivWithinAt hh.hasDerivWithinAt hs).derivWithin hxs + (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s t') : + derivWithin (g₁ ∘ h) s x = derivWithin h s x • derivWithin g₁ t' (h x) := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (HasDerivWithinAt.scomp x hg.hasDerivWithinAt hh.hasDerivWithinAt hs).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem derivWithin.scomp_of_eq (hg : DifferentiableWithinAt 𝕜' g₁ t' y) - (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s t') (hxs : UniqueDiffWithinAt 𝕜 s x) + (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s t') (hy : y = h x) : derivWithin (g₁ ∘ h) s x = derivWithin h s x • derivWithin g₁ t' (h x) := by - rw [hy] at hg; exact derivWithin.scomp x hg hh hs hxs + rw [hy] at hg; exact derivWithin.scomp x hg hh hs theorem deriv.scomp (hg : DifferentiableAt 𝕜' g₁ (h x)) (hh : DifferentiableAt 𝕜 h x) : deriv (g₁ ∘ h) x = deriv h x • deriv g₁ (h x) := @@ -265,17 +267,19 @@ theorem HasDerivAt.comp_hasDerivWithinAt_of_eq (hh₂ : HasDerivAt h₂ h₂' y) rw [hy] at hh₂; exact hh₂.comp_hasDerivWithinAt x hh theorem derivWithin_comp (hh₂ : DifferentiableWithinAt 𝕜' h₂ s' (h x)) - (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s s') (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (h₂ ∘ h) s x = derivWithin h₂ s' (h x) * derivWithin h s x := - (hh₂.hasDerivWithinAt.comp x hh.hasDerivWithinAt hs).derivWithin hxs + (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s s') : + derivWithin (h₂ ∘ h) s x = derivWithin h₂ s' (h x) * derivWithin h s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hh₂.hasDerivWithinAt.comp x hh.hasDerivWithinAt hs).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[deprecated (since := "2024-10-31")] alias derivWithin.comp := derivWithin_comp theorem derivWithin_comp_of_eq (hh₂ : DifferentiableWithinAt 𝕜' h₂ s' y) - (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s s') (hxs : UniqueDiffWithinAt 𝕜 s x) + (hh : DifferentiableWithinAt 𝕜 h s x) (hs : MapsTo h s s') (hy : h x = y) : derivWithin (h₂ ∘ h) s x = derivWithin h₂ s' (h x) * derivWithin h s x := by - subst hy; exact derivWithin_comp x hh₂ hh hs hxs + subst hy; exact derivWithin_comp x hh₂ hh hs @[deprecated (since := "2024-10-31")] alias derivWithin.comp_of_eq := derivWithin_comp_of_eq @@ -373,18 +377,19 @@ theorem HasStrictFDerivAt.comp_hasStrictDerivAt_of_eq (hl : HasStrictFDerivAt l rw [hy] at hl; exact hl.comp_hasStrictDerivAt x hf theorem fderivWithin_comp_derivWithin {t : Set F} (hl : DifferentiableWithinAt 𝕜 l t (f x)) - (hf : DifferentiableWithinAt 𝕜 f s x) (hs : MapsTo f s t) (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (l ∘ f) s x = (fderivWithin 𝕜 l t (f x) : F → E) (derivWithin f s x) := - (hl.hasFDerivWithinAt.comp_hasDerivWithinAt x hf.hasDerivWithinAt hs).derivWithin hxs + (hf : DifferentiableWithinAt 𝕜 f s x) (hs : MapsTo f s t) : + derivWithin (l ∘ f) s x = (fderivWithin 𝕜 l t (f x) : F → E) (derivWithin f s x) := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hl.hasFDerivWithinAt.comp_hasDerivWithinAt x hf.hasDerivWithinAt hs).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[deprecated (since := "2024-10-31")] alias fderivWithin.comp_derivWithin := fderivWithin_comp_derivWithin theorem fderivWithin_comp_derivWithin_of_eq {t : Set F} (hl : DifferentiableWithinAt 𝕜 l t y) - (hf : DifferentiableWithinAt 𝕜 f s x) (hs : MapsTo f s t) (hxs : UniqueDiffWithinAt 𝕜 s x) - (hy : y = f x) : + (hf : DifferentiableWithinAt 𝕜 f s x) (hs : MapsTo f s t) (hy : y = f x) : derivWithin (l ∘ f) s x = (fderivWithin 𝕜 l t (f x) : F → E) (derivWithin f s x) := by - rw [hy] at hl; exact fderivWithin_comp_derivWithin x hl hf hs hxs + rw [hy] at hl; exact fderivWithin_comp_derivWithin x hl hf hs @[deprecated (since := "2024-10-31")] alias fderivWithin.comp_derivWithin_of_eq := fderivWithin_comp_derivWithin_of_eq diff --git a/Mathlib/Analysis/Calculus/Deriv/Inv.lean b/Mathlib/Analysis/Calculus/Deriv/Inv.lean index 082c6d0c809ed..605b8d06e2e87 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Inv.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Inv.lean @@ -106,10 +106,11 @@ theorem HasDerivAt.inv (hc : HasDerivAt c c' x) (hx : c x ≠ 0) : rw [← hasDerivWithinAt_univ] at * exact hc.inv hx -theorem derivWithin_inv' (hc : DifferentiableWithinAt 𝕜 c s x) (hx : c x ≠ 0) - (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (fun x => (c x)⁻¹) s x = -derivWithin c s x / c x ^ 2 := - (hc.hasDerivWithinAt.inv hx).derivWithin hxs +theorem derivWithin_inv' (hc : DifferentiableWithinAt 𝕜 c s x) (hx : c x ≠ 0) : + derivWithin (fun x => (c x)⁻¹) s x = -derivWithin c s x / c x ^ 2 := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.inv hx).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_inv'' (hc : DifferentiableAt 𝕜 c x) (hx : c x ≠ 0) : @@ -163,10 +164,12 @@ theorem Differentiable.div (hc : Differentiable 𝕜 c) (hd : Differentiable Differentiable 𝕜 fun x => c x / d x := fun x => (hc x).div (hd x) (hx x) theorem derivWithin_div (hc : DifferentiableWithinAt 𝕜 c s x) (hd : DifferentiableWithinAt 𝕜 d s x) - (hx : d x ≠ 0) (hxs : UniqueDiffWithinAt 𝕜 s x) : + (hx : d x ≠ 0) : derivWithin (fun x => c x / d x) s x = - (derivWithin c s x * d x - c x * derivWithin d s x) / d x ^ 2 := - (hc.hasDerivWithinAt.div hd.hasDerivWithinAt hx).derivWithin hxs + (derivWithin c s x * d x - c x * derivWithin d s x) / d x ^ 2 := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.div hd.hasDerivWithinAt hx).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_div (hc : DifferentiableAt 𝕜 c x) (hd : DifferentiableAt 𝕜 d x) (hx : d x ≠ 0) : diff --git a/Mathlib/Analysis/Calculus/Deriv/Mul.lean b/Mathlib/Analysis/Calculus/Deriv/Mul.lean index b7b58303ed4e3..3da9a8ed34991 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Mul.lean @@ -61,11 +61,13 @@ theorem hasStrictDerivAt_of_bilinear (hu : HasStrictDerivAt u u' x) (hv : HasStr simpa using (B.hasStrictFDerivAt_of_bilinear hu.hasStrictFDerivAt hv.hasStrictFDerivAt).hasStrictDerivAt -theorem derivWithin_of_bilinear (hxs : UniqueDiffWithinAt 𝕜 s x) +theorem derivWithin_of_bilinear (hu : DifferentiableWithinAt 𝕜 u s x) (hv : DifferentiableWithinAt 𝕜 v s x) : derivWithin (fun y => B (u y) (v y)) s x = - B (u x) (derivWithin v s x) + B (derivWithin u s x) (v x) := - (B.hasDerivWithinAt_of_bilinear hu.hasDerivWithinAt hv.hasDerivWithinAt).derivWithin hxs + B (u x) (derivWithin v s x) + B (derivWithin u s x) (v x) := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (B.hasDerivWithinAt_of_bilinear hu.hasDerivWithinAt hv.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_of_bilinear (hu : DifferentiableAt 𝕜 u x) (hv : DifferentiableAt 𝕜 v x) : deriv (fun y => B (u y) (v y)) x = B (u x) (deriv v x) + B (deriv u x) (v x) := @@ -94,10 +96,12 @@ nonrec theorem HasStrictDerivAt.smul (hc : HasStrictDerivAt c c' x) (hf : HasStr HasStrictDerivAt (fun y => c y • f y) (c x • f' + c' • f x) x := by simpa using (hc.smul hf).hasStrictDerivAt -theorem derivWithin_smul (hxs : UniqueDiffWithinAt 𝕜 s x) (hc : DifferentiableWithinAt 𝕜 c s x) +theorem derivWithin_smul (hc : DifferentiableWithinAt 𝕜 c s x) (hf : DifferentiableWithinAt 𝕜 f s x) : - derivWithin (fun y => c y • f y) s x = c x • derivWithin f s x + derivWithin c s x • f x := - (hc.hasDerivWithinAt.smul hf.hasDerivWithinAt).derivWithin hxs + derivWithin (fun y => c y • f y) s x = c x • derivWithin f s x + derivWithin c s x • f x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.smul hf.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_smul (hc : DifferentiableAt 𝕜 c x) (hf : DifferentiableAt 𝕜 f x) : deriv (fun y => c y • f y) x = c x • deriv f x + deriv c x • f x := @@ -118,10 +122,11 @@ theorem HasDerivAt.smul_const (hc : HasDerivAt c c' x) (f : F) : rw [← hasDerivWithinAt_univ] at * exact hc.smul_const f -theorem derivWithin_smul_const (hxs : UniqueDiffWithinAt 𝕜 s x) - (hc : DifferentiableWithinAt 𝕜 c s x) (f : F) : - derivWithin (fun y => c y • f) s x = derivWithin c s x • f := - (hc.hasDerivWithinAt.smul_const f).derivWithin hxs +theorem derivWithin_smul_const (hc : DifferentiableWithinAt 𝕜 c s x) (f : F) : + derivWithin (fun y => c y • f) s x = derivWithin c s x • f := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.smul_const f).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_smul_const (hc : DifferentiableAt 𝕜 c x) (f : F) : deriv (fun y => c y • f) x = deriv c x • f := @@ -149,10 +154,11 @@ nonrec theorem HasDerivAt.const_smul (c : R) (hf : HasDerivAt f f' x) : HasDerivAt (fun y => c • f y) (c • f') x := hf.const_smul c -theorem derivWithin_const_smul (hxs : UniqueDiffWithinAt 𝕜 s x) (c : R) - (hf : DifferentiableWithinAt 𝕜 f s x) : - derivWithin (fun y => c • f y) s x = c • derivWithin f s x := - (hf.hasDerivWithinAt.const_smul c).derivWithin hxs +theorem derivWithin_const_smul (c : R) (hf : DifferentiableWithinAt 𝕜 f s x) : + derivWithin (fun y => c • f y) s x = c • derivWithin f s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hf.hasDerivWithinAt.const_smul c).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_const_smul (c : R) (hf : DifferentiableAt 𝕜 f x) : deriv (fun y => c • f y) x = c • deriv f x := @@ -204,10 +210,12 @@ theorem HasStrictDerivAt.mul (hc : HasStrictDerivAt c c' x) (hd : HasStrictDeriv ContinuousLinearMap.smulRight_apply, ContinuousLinearMap.one_apply, one_smul, one_smul, add_comm] at this -theorem derivWithin_mul (hxs : UniqueDiffWithinAt 𝕜 s x) (hc : DifferentiableWithinAt 𝕜 c s x) +theorem derivWithin_mul (hc : DifferentiableWithinAt 𝕜 c s x) (hd : DifferentiableWithinAt 𝕜 d s x) : - derivWithin (fun y => c y * d y) s x = derivWithin c s x * d x + c x * derivWithin d s x := - (hc.hasDerivWithinAt.mul hd.hasDerivWithinAt).derivWithin hxs + derivWithin (fun y => c y * d y) s x = derivWithin c s x * d x + c x * derivWithin d s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.mul hd.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_mul (hc : DifferentiableAt 𝕜 c x) (hd : DifferentiableAt 𝕜 d x) : @@ -232,9 +240,22 @@ theorem HasStrictDerivAt.mul_const (hc : HasStrictDerivAt c c' x) (d : 𝔸) : convert hc.mul (hasStrictDerivAt_const x d) using 1 rw [mul_zero, add_zero] -theorem derivWithin_mul_const (hxs : UniqueDiffWithinAt 𝕜 s x) (hc : DifferentiableWithinAt 𝕜 c s x) - (d : 𝔸) : derivWithin (fun y => c y * d) s x = derivWithin c s x * d := - (hc.hasDerivWithinAt.mul_const d).derivWithin hxs +theorem derivWithin_mul_const (hc : DifferentiableWithinAt 𝕜 c s x) (d : 𝔸) : + derivWithin (fun y => c y * d) s x = derivWithin c s x * d := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.mul_const d).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] + +lemma derivWithin_mul_const_field (u : 𝕜') : + derivWithin (fun y => v y * u) s x = derivWithin v s x * u := by + by_cases hv : DifferentiableWithinAt 𝕜 v s x + · rw [derivWithin_mul_const hv u] + by_cases hu : u = 0 + · simp [hu] + rw [derivWithin_zero_of_not_differentiableWithinAt hv, zero_mul, + derivWithin_zero_of_not_differentiableWithinAt] + have : v = fun x ↦ (v x * u) * u⁻¹ := by ext; simp [hu] + exact fun h_diff ↦ hv <| this ▸ h_diff.mul_const _ theorem deriv_mul_const (hc : DifferentiableAt 𝕜 c x) (d : 𝔸) : deriv (fun y => c y * d) x = deriv c x * d := @@ -268,10 +289,16 @@ theorem HasStrictDerivAt.const_mul (c : 𝔸) (hd : HasStrictDerivAt d d' x) : convert (hasStrictDerivAt_const _ _).mul hd using 1 rw [zero_mul, zero_add] -theorem derivWithin_const_mul (hxs : UniqueDiffWithinAt 𝕜 s x) (c : 𝔸) - (hd : DifferentiableWithinAt 𝕜 d s x) : - derivWithin (fun y => c * d y) s x = c * derivWithin d s x := - (hd.hasDerivWithinAt.const_mul c).derivWithin hxs +theorem derivWithin_const_mul (c : 𝔸) (hd : DifferentiableWithinAt 𝕜 d s x) : + derivWithin (fun y => c * d y) s x = c * derivWithin d s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hd.hasDerivWithinAt.const_mul c).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] + +lemma derivWithin_const_mul_field (u : 𝕜') : + derivWithin (fun y => u * v y) s x = u * derivWithin v s x := by + simp_rw [mul_comm u] + exact derivWithin_mul_const_field u theorem deriv_const_mul (c : 𝔸) (hd : DifferentiableAt 𝕜 d x) : deriv (fun y => c * d y) x = c * deriv d x := @@ -312,11 +339,13 @@ theorem deriv_finset_prod (hf : ∀ i ∈ u, DifferentiableAt 𝕜 (f i) x) : deriv (∏ i ∈ u, f i ·) x = ∑ i ∈ u, (∏ j ∈ u.erase i, f j x) • deriv (f i) x := (HasDerivAt.finset_prod fun i hi ↦ (hf i hi).hasDerivAt).deriv -theorem derivWithin_finset_prod (hxs : UniqueDiffWithinAt 𝕜 s x) +theorem derivWithin_finset_prod (hf : ∀ i ∈ u, DifferentiableWithinAt 𝕜 (f i) s x) : derivWithin (∏ i ∈ u, f i ·) s x = - ∑ i ∈ u, (∏ j ∈ u.erase i, f j x) • derivWithin (f i) s x := - (HasDerivWithinAt.finset_prod fun i hi ↦ (hf i hi).hasDerivWithinAt).derivWithin hxs + ∑ i ∈ u, (∏ j ∈ u.erase i, f j x) • derivWithin (f i) s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (HasDerivWithinAt.finset_prod fun i hi ↦ (hf i hi).hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] end HasDeriv @@ -383,10 +412,9 @@ theorem DifferentiableOn.div_const (hc : DifferentiableOn 𝕜 c s) (d : 𝕜') theorem Differentiable.div_const (hc : Differentiable 𝕜 c) (d : 𝕜') : Differentiable 𝕜 fun x => c x / d := fun x => (hc x).div_const d -theorem derivWithin_div_const (hc : DifferentiableWithinAt 𝕜 c s x) - (d : 𝕜') (hxs : UniqueDiffWithinAt 𝕜 s x) : +theorem derivWithin_div_const (hc : DifferentiableWithinAt 𝕜 c s x) (d : 𝕜') : derivWithin (fun x => c x / d) s x = derivWithin c s x / d := by - simp [div_eq_inv_mul, derivWithin_const_mul, hc, hxs] + simp [div_eq_inv_mul, derivWithin_const_mul, hc] @[simp] theorem deriv_div_const (d : 𝕜') : deriv (fun x => c x / d) x = deriv c x / d := by @@ -423,10 +451,12 @@ theorem HasDerivAt.clm_comp (hc : HasDerivAt c c' x) (hd : HasDerivAt d d' x) : exact hc.clm_comp hd theorem derivWithin_clm_comp (hc : DifferentiableWithinAt 𝕜 c s x) - (hd : DifferentiableWithinAt 𝕜 d s x) (hxs : UniqueDiffWithinAt 𝕜 s x) : + (hd : DifferentiableWithinAt 𝕜 d s x) : derivWithin (fun y => (c y).comp (d y)) s x = - (derivWithin c s x).comp (d x) + (c x).comp (derivWithin d s x) := - (hc.hasDerivWithinAt.clm_comp hd.hasDerivWithinAt).derivWithin hxs + (derivWithin c s x).comp (d x) + (c x).comp (derivWithin d s x) := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.clm_comp hd.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_clm_comp (hc : DifferentiableAt 𝕜 c x) (hd : DifferentiableAt 𝕜 d x) : deriv (fun y => (c y).comp (d y)) x = (deriv c x).comp (d x) + (c x).comp (deriv d x) := @@ -451,10 +481,12 @@ theorem HasDerivAt.clm_apply (hc : HasDerivAt c c' x) (hu : HasDerivAt u u' x) : rwa [add_apply, comp_apply, flip_apply, smulRight_apply, smulRight_apply, one_apply, one_smul, one_smul, add_comm] at this -theorem derivWithin_clm_apply (hxs : UniqueDiffWithinAt 𝕜 s x) (hc : DifferentiableWithinAt 𝕜 c s x) +theorem derivWithin_clm_apply (hc : DifferentiableWithinAt 𝕜 c s x) (hu : DifferentiableWithinAt 𝕜 u s x) : - derivWithin (fun y => (c y) (u y)) s x = derivWithin c s x (u x) + c x (derivWithin u s x) := - (hc.hasDerivWithinAt.clm_apply hu.hasDerivWithinAt).derivWithin hxs + derivWithin (fun y => (c y) (u y)) s x = derivWithin c s x (u x) + c x (derivWithin u s x) := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.clm_apply hu.hasDerivWithinAt).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] theorem deriv_clm_apply (hc : DifferentiableAt 𝕜 c x) (hu : DifferentiableAt 𝕜 u x) : deriv (fun y => (c y) (u y)) x = deriv c x (u x) + c x (deriv u x) := diff --git a/Mathlib/Analysis/Calculus/Deriv/Pow.lean b/Mathlib/Analysis/Calculus/Deriv/Pow.lean index 07f1c232737a5..e20d8811626b5 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Pow.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Pow.lean @@ -76,9 +76,11 @@ theorem HasDerivAt.pow (hc : HasDerivAt c c' x) : rw [← hasDerivWithinAt_univ] at * exact hc.pow n -theorem derivWithin_pow' (hc : DifferentiableWithinAt 𝕜 c s x) (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (fun x => c x ^ n) s x = (n : 𝕜) * c x ^ (n - 1) * derivWithin c s x := - (hc.hasDerivWithinAt.pow n).derivWithin hxs +theorem derivWithin_pow' (hc : DifferentiableWithinAt 𝕜 c s x) : + derivWithin (fun x => c x ^ n) s x = (n : 𝕜) * c x ^ (n - 1) * derivWithin c s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hc.hasDerivWithinAt.pow n).derivWithin hxs + · simp [derivWithin_zero_of_isolated hxs] @[simp] theorem deriv_pow'' (hc : DifferentiableAt 𝕜 c x) : diff --git a/Mathlib/Analysis/Calculus/Deriv/Prod.lean b/Mathlib/Analysis/Calculus/Deriv/Prod.lean index a6a1adbcdb78d..caa6c11683543 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Prod.lean @@ -78,10 +78,12 @@ theorem hasDerivWithinAt_pi : HasDerivWithinAt φ φ' s x ↔ ∀ i, HasDerivWithinAt (fun x => φ x i) (φ' i) s x := hasDerivAtFilter_pi -theorem derivWithin_pi (h : ∀ i, DifferentiableWithinAt 𝕜 (fun x => φ x i) s x) - (hs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin φ s x = fun i => derivWithin (fun x => φ x i) s x := - (hasDerivWithinAt_pi.2 fun i => (h i).hasDerivWithinAt).derivWithin hs +theorem derivWithin_pi (h : ∀ i, DifferentiableWithinAt 𝕜 (fun x => φ x i) s x) : + derivWithin φ s x = fun i => derivWithin (fun x => φ x i) s x := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact (hasDerivWithinAt_pi.2 fun i => (h i).hasDerivWithinAt).derivWithin hxs + · simp only [derivWithin_zero_of_isolated hxs] + rfl theorem deriv_pi (h : ∀ i, DifferentiableAt 𝕜 (fun x => φ x i) x) : deriv φ x = fun i => deriv (fun x => φ x i) x := diff --git a/Mathlib/Analysis/Calculus/Deriv/Star.lean b/Mathlib/Analysis/Calculus/Deriv/Star.lean index 2b4e197ec1c94..636a850b2f0e1 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Star.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Star.lean @@ -41,9 +41,11 @@ protected nonrec theorem HasDerivAt.star (h : HasDerivAt f f' x) : protected nonrec theorem HasStrictDerivAt.star (h : HasStrictDerivAt f f' x) : HasStrictDerivAt (fun x => star (f x)) (star f') x := by simpa using h.star.hasStrictDerivAt -protected theorem derivWithin.star (hxs : UniqueDiffWithinAt 𝕜 s x) : - derivWithin (fun y => star (f y)) s x = star (derivWithin f s x) := - DFunLike.congr_fun (fderivWithin_star hxs) _ +protected theorem derivWithin.star : + derivWithin (fun y => star (f y)) s x = star (derivWithin f s x) := by + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · exact DFunLike.congr_fun (fderivWithin_star hxs) _ + · simp [derivWithin_zero_of_isolated hxs] protected theorem deriv.star : deriv (fun y => star (f y)) x = star (deriv f x) := DFunLike.congr_fun fderiv_star _ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 59fd0722065d8..09a1501984d89 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -118,6 +118,7 @@ theorem AnalyticWithinAt.differentiableWithinAt (h : AnalyticWithinAt 𝕜 f s x obtain ⟨p, hp⟩ := h exact hp.differentiableWithinAt +@[fun_prop] theorem AnalyticAt.differentiableAt : AnalyticAt 𝕜 f x → DifferentiableAt 𝕜 f x | ⟨_, hp⟩ => hp.differentiableAt @@ -234,6 +235,7 @@ protected theorem HasFPowerSeriesWithinOnBall.fderivWithin_of_mem [CompleteSpace exact this.symm /-- If a function is analytic on a set `s`, so is its Fréchet derivative. -/ +@[fun_prop] protected theorem AnalyticAt.fderiv [CompleteSpace F] (h : AnalyticAt 𝕜 f x) : AnalyticAt 𝕜 (fderiv 𝕜 f) x := by rcases h with ⟨p, r, hp⟩ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index b518f82c09cb7..d9ecb60adbc29 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -485,6 +485,20 @@ lemma le_minSmoothness {n : WithTop ℕ∞} : n ≤ minSmoothness 𝕜 n := by simp only [minSmoothness] split_ifs <;> simp +lemma minSmoothness_add {n m : WithTop ℕ∞} : minSmoothness 𝕜 (n + m) = minSmoothness 𝕜 n + m := by + simp only [minSmoothness] + split_ifs <;> simp + +lemma minSmoothness_monotone : Monotone (minSmoothness 𝕜) := by + intro m n hmn + simp only [minSmoothness] + split_ifs <;> simp [hmn] + +@[simp] lemma minSmoothness_eq_infty {n : WithTop ℕ∞} : + minSmoothness 𝕜 n = ∞ ↔ (n = ∞ ∧ IsRCLikeNormedField 𝕜) := by + simp only [minSmoothness] + split_ifs with h <;> simp [h] + /-- If `minSmoothness 𝕜 m ≤ n` for some (finite) integer `m`, then one can find `n' ∈ [minSmoothness 𝕜 m, n]` which is not `∞`: over `ℝ` or `ℂ`, just take `m`, and otherwise just take `ω`. The interest of this technical lemma is that, if a function is `C^{n'}` at a point diff --git a/Mathlib/Analysis/Calculus/Gradient/Basic.lean b/Mathlib/Analysis/Calculus/Gradient/Basic.lean index 5c6489419a6d0..e67253e3e2ec7 100644 --- a/Mathlib/Analysis/Calculus/Gradient/Basic.lean +++ b/Mathlib/Analysis/Calculus/Gradient/Basic.lean @@ -68,6 +68,7 @@ def gradientWithin (f : F → 𝕜) (s : Set F) (x : F) : F := (toDual 𝕜 F).symm (fderivWithin 𝕜 f s x) /-- Gradient of `f` at the point `x`, if it exists. Zero otherwise. +Denoted as `∇` within the Gradient namespace. If the derivative exists (i.e., `∃ f', HasGradientAt f f' x`), then `f x' = f x + ⟨f', x' - x⟩ + o (x' - x)` where `x'` converges to `x`. -/ diff --git a/Mathlib/Analysis/Calculus/Implicit.lean b/Mathlib/Analysis/Calculus/Implicit.lean index d3b9846464c99..0fafd1d4cd140 100644 --- a/Mathlib/Analysis/Calculus/Implicit.lean +++ b/Mathlib/Analysis/Calculus/Implicit.lean @@ -94,7 +94,6 @@ needs to have a complete control over the choice of the implicit function. * both functions are strictly differentiable at `a`; * the derivatives are surjective; * the kernels of the derivatives are complementary subspaces of `E`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not yet ported @[nolint has_nonempty_instance] structure ImplicitFunctionData (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Type*) [NormedAddCommGroup E] [NormedSpace 𝕜 E] [CompleteSpace E] (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] [CompleteSpace F] (G : Type*) [NormedAddCommGroup G] [NormedSpace 𝕜 G] diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index a013a881fc12f..bd390d2a8ebcb 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -102,9 +102,12 @@ theorem iteratedDerivWithin_zero : iteratedDerivWithin 0 f s = f := by simp [iteratedDerivWithin] @[simp] -theorem iteratedDerivWithin_one {x : 𝕜} (h : UniqueDiffWithinAt 𝕜 s x) : +theorem iteratedDerivWithin_one {x : 𝕜} : iteratedDerivWithin 1 f s x = derivWithin f s x := by - simp only [iteratedDerivWithin, iteratedFDerivWithin_one_apply h]; rfl + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · simp only [iteratedDerivWithin, iteratedFDerivWithin_one_apply hxs]; rfl + · simp [derivWithin_zero_of_isolated hxs, iteratedDerivWithin, iteratedFDerivWithin, + fderivWithin_zero_of_isolated hxs] /-- If the first `n` derivatives within a set of a function are continuous, and its first `n-1` derivatives are differentiable, then the function is `C^n`. This is not an equivalence in general, @@ -170,30 +173,33 @@ theorem contDiffOn_nat_iff_continuousOn_differentiableOn_deriv {n : ℕ} (hs : U /-- The `n+1`-th iterated derivative within a set with unique derivatives can be obtained by differentiating the `n`-th iterated derivative. -/ -theorem iteratedDerivWithin_succ {x : 𝕜} (hxs : UniqueDiffWithinAt 𝕜 s x) : +theorem iteratedDerivWithin_succ {x : 𝕜} : iteratedDerivWithin (n + 1) f s x = derivWithin (iteratedDerivWithin n f s) s x := by - rw [iteratedDerivWithin_eq_iteratedFDerivWithin, iteratedFDerivWithin_succ_apply_left, - iteratedFDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_fderivWithin _ hxs, derivWithin] - change ((ContinuousMultilinearMap.mkPiRing 𝕜 (Fin n) ((fderivWithin 𝕜 - (iteratedDerivWithin n f s) s x : 𝕜 → F) 1) : (Fin n → 𝕜) → F) fun _ : Fin n => 1) = - (fderivWithin 𝕜 (iteratedDerivWithin n f s) s x : 𝕜 → F) 1 - simp + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs + · rw [iteratedDerivWithin_eq_iteratedFDerivWithin, iteratedFDerivWithin_succ_apply_left, + iteratedFDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_fderivWithin _ hxs, derivWithin] + change ((ContinuousMultilinearMap.mkPiRing 𝕜 (Fin n) ((fderivWithin 𝕜 + (iteratedDerivWithin n f s) s x : 𝕜 → F) 1) : (Fin n → 𝕜) → F) fun _ : Fin n => 1) = + (fderivWithin 𝕜 (iteratedDerivWithin n f s) s x : 𝕜 → F) 1 + simp + · simp [derivWithin_zero_of_isolated hxs, iteratedDerivWithin, iteratedFDerivWithin, + fderivWithin_zero_of_isolated hxs] /-- The `n`-th iterated derivative within a set with unique derivatives can be obtained by iterating `n` times the differentiation operation. -/ -theorem iteratedDerivWithin_eq_iterate {x : 𝕜} (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : +theorem iteratedDerivWithin_eq_iterate {x : 𝕜} : iteratedDerivWithin n f s x = (fun g : 𝕜 → F => derivWithin g s)^[n] f x := by induction n generalizing x with | zero => simp | succ n IH => - rw [iteratedDerivWithin_succ (hs x hx), Function.iterate_succ'] - exact derivWithin_congr (fun y hy => IH hy) (IH hx) + rw [iteratedDerivWithin_succ, Function.iterate_succ'] + exact derivWithin_congr (fun y hy => IH) IH /-- The `n+1`-th iterated derivative within a set with unique derivatives can be obtained by taking the `n`-th derivative of the derivative. -/ -theorem iteratedDerivWithin_succ' {x : 𝕜} (hxs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : +theorem iteratedDerivWithin_succ' {x : 𝕜} : iteratedDerivWithin (n + 1) f s x = (iteratedDerivWithin n (derivWithin f s) s) x := by - rw [iteratedDerivWithin_eq_iterate hxs hx, iteratedDerivWithin_eq_iterate hxs hx]; rfl + rw [iteratedDerivWithin_eq_iterate, iteratedDerivWithin_eq_iterate]; rfl /-! ### Properties of the iterated derivative on the whole space -/ @@ -270,14 +276,14 @@ iterated derivative. -/ theorem iteratedDeriv_succ : iteratedDeriv (n + 1) f = deriv (iteratedDeriv n f) := by ext x rw [← iteratedDerivWithin_univ, ← iteratedDerivWithin_univ, ← derivWithin_univ] - exact iteratedDerivWithin_succ uniqueDiffWithinAt_univ + exact iteratedDerivWithin_succ /-- The `n`-th iterated derivative can be obtained by iterating `n` times the differentiation operation. -/ theorem iteratedDeriv_eq_iterate : iteratedDeriv n f = deriv^[n] f := by ext x rw [← iteratedDerivWithin_univ] - convert iteratedDerivWithin_eq_iterate uniqueDiffOn_univ (F := F) (mem_univ x) + convert iteratedDerivWithin_eq_iterate (F := F) simp [derivWithin_univ] /-- The `n+1`-th iterated derivative can be obtained by taking the `n`-th derivative of the diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index 01f21bd7293eb..fb639723208c3 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -24,7 +24,6 @@ variable {n : ℕ} {x : 𝕜} {s : Set 𝕜} (hx : x ∈ s) (h : UniqueDiffOn 𝕜 s) {f g : 𝕜 → F} section -include h theorem iteratedDerivWithin_congr (hfg : Set.EqOn f g s) : Set.EqOn (iteratedDerivWithin n f s) (iteratedDerivWithin n g s) s := by @@ -32,12 +31,10 @@ theorem iteratedDerivWithin_congr (hfg : Set.EqOn f g s) : | zero => rwa [iteratedDerivWithin_zero] | succ n IH => intro y hy - have : UniqueDiffWithinAt 𝕜 s y := h.uniqueDiffWithinAt hy - rw [iteratedDerivWithin_succ this, iteratedDerivWithin_succ this] + rw [iteratedDerivWithin_succ, iteratedDerivWithin_succ] exact derivWithin_congr (IH hfg) (IH hfg hy) -include hx - +include h hx in theorem iteratedDerivWithin_add (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) : iteratedDerivWithin n (f + g) s x = iteratedDerivWithin n f s x + iteratedDerivWithin n g s x := by @@ -47,50 +44,61 @@ theorem iteratedDerivWithin_add (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn theorem iteratedDerivWithin_const_add (hn : 0 < n) (c : F) : iteratedDerivWithin n (fun z => c + f z) s x = iteratedDerivWithin n f s x := by obtain ⟨n, rfl⟩ := n.exists_eq_succ_of_ne_zero hn.ne' - rw [iteratedDerivWithin_succ' h hx, iteratedDerivWithin_succ' h hx] - refine iteratedDerivWithin_congr h ?_ hx - intro y hy - exact derivWithin_const_add (h.uniqueDiffWithinAt hy) _ + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs; swap + · simp [iteratedDerivWithin_succ, derivWithin_zero_of_isolated hxs] + rw [iteratedDerivWithin_succ', iteratedDerivWithin_succ'] + congr with y + exact derivWithin_const_add _ theorem iteratedDerivWithin_const_sub (hn : 0 < n) (c : F) : iteratedDerivWithin n (fun z => c - f z) s x = iteratedDerivWithin n (fun z => -f z) s x := by obtain ⟨n, rfl⟩ := n.exists_eq_succ_of_ne_zero hn.ne' - rw [iteratedDerivWithin_succ' h hx, iteratedDerivWithin_succ' h hx] - refine iteratedDerivWithin_congr h ?_ hx - intro y hy - have : UniqueDiffWithinAt 𝕜 s y := h.uniqueDiffWithinAt hy - rw [derivWithin.neg this] - exact derivWithin_const_sub this _ + rcases uniqueDiffWithinAt_or_nhdsWithin_eq_bot s x with hxs | hxs; swap + · simp [iteratedDerivWithin_succ, derivWithin_zero_of_isolated hxs] + rw [iteratedDerivWithin_succ', iteratedDerivWithin_succ'] + congr with y + rw [derivWithin.neg] + exact derivWithin_const_sub _ @[deprecated (since := "2024-12-10")] alias iteratedDerivWithin_const_neg := iteratedDerivWithin_const_sub +include h hx in theorem iteratedDerivWithin_const_smul (c : R) (hf : ContDiffWithinAt 𝕜 n f s x) : iteratedDerivWithin n (c • f) s x = c • iteratedDerivWithin n f s x := by simp_rw [iteratedDerivWithin] rw [iteratedFDerivWithin_const_smul_apply hf h hx] simp only [ContinuousMultilinearMap.smul_apply] +include h hx in theorem iteratedDerivWithin_const_mul (c : 𝕜) {f : 𝕜 → 𝕜} (hf : ContDiffWithinAt 𝕜 n f s x) : iteratedDerivWithin n (fun z => c * f z) s x = c * iteratedDerivWithin n f s x := by simpa using iteratedDerivWithin_const_smul (F := 𝕜) hx h c hf variable (f) in +omit h hx in theorem iteratedDerivWithin_neg : iteratedDerivWithin n (-f) s x = -iteratedDerivWithin n f s x := by - rw [iteratedDerivWithin, iteratedDerivWithin, iteratedFDerivWithin_neg_apply h hx, - ContinuousMultilinearMap.neg_apply] + induction n generalizing x with + | zero => simp + | succ n IH => + simp only [iteratedDerivWithin_succ, derivWithin_neg] + rw [← derivWithin.neg] + congr with y + exact IH variable (f) in theorem iteratedDerivWithin_neg' : iteratedDerivWithin n (fun z => -f z) s x = -iteratedDerivWithin n f s x := - iteratedDerivWithin_neg hx h f + iteratedDerivWithin_neg f + +include h hx theorem iteratedDerivWithin_sub (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) : iteratedDerivWithin n (f - g) s x = iteratedDerivWithin n f s x - iteratedDerivWithin n g s x := by rw [sub_eq_add_neg, sub_eq_add_neg, Pi.neg_def, iteratedDerivWithin_add hx h hf hg.neg, - iteratedDerivWithin_neg' hx h] + iteratedDerivWithin_neg'] theorem iteratedDerivWithin_comp_const_smul (hf : ContDiffOn 𝕜 n f s) (c : 𝕜) (hs : Set.MapsTo (c * ·) s s) : @@ -111,11 +119,11 @@ theorem iteratedDerivWithin_comp_const_smul (hf : ContDiffOn 𝕜 n f s) (c : · exact hf.differentiableOn_iteratedDerivWithin (Nat.cast_lt.mpr n.lt_succ_self) h _ hcx · exact differentiableWithinAt_id'.const_mul _ · exact hs - rw [iteratedDerivWithin_succ (h _ hx), derivWithin_congr h₀ (ih hx hf.of_succ), - derivWithin_const_smul (h _ hx) (c ^ n) h₂, iteratedDerivWithin_succ (h _ hcx), + rw [iteratedDerivWithin_succ, derivWithin_congr h₀ (ih hx hf.of_succ), + derivWithin_const_smul (c ^ n) h₂, iteratedDerivWithin_succ, ← Function.comp_def, - derivWithin.scomp x h₁ (differentiableWithinAt_id'.const_mul _) hs (h _ hx), - derivWithin_const_mul (h _ hx) _ differentiableWithinAt_id', derivWithin_id' _ _ (h _ hx), + derivWithin.scomp x h₁ (differentiableWithinAt_id'.const_mul _) hs, + derivWithin_const_mul _ differentiableWithinAt_id', derivWithin_id' _ _ (h _ hx), smul_smul, mul_one, pow_succ] end @@ -128,18 +136,15 @@ lemma iteratedDeriv_add (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : theorem iteratedDeriv_const_add (hn : 0 < n) (c : F) : iteratedDeriv n (fun z => c + f z) x = iteratedDeriv n f x := by - simpa only [iteratedDerivWithin_univ] using - iteratedDerivWithin_const_add (Set.mem_univ _) uniqueDiffOn_univ hn c + simpa only [← iteratedDerivWithin_univ] using iteratedDerivWithin_const_add hn c theorem iteratedDeriv_const_sub (hn : 0 < n) (c : F) : iteratedDeriv n (fun z => c - f z) x = iteratedDeriv n (-f) x := by - simpa only [iteratedDerivWithin_univ] using - iteratedDerivWithin_const_sub (Set.mem_univ _) uniqueDiffOn_univ hn c + simpa only [← iteratedDerivWithin_univ] using iteratedDerivWithin_const_sub hn c lemma iteratedDeriv_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : iteratedDeriv n (fun x ↦ -(f x)) a = -(iteratedDeriv n f a) := by - simpa only [iteratedDerivWithin_univ] using - iteratedDerivWithin_neg (Set.mem_univ a) uniqueDiffOn_univ f + simpa only [← iteratedDerivWithin_univ] using iteratedDerivWithin_neg f lemma iteratedDeriv_sub (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) : iteratedDeriv n (f - g) x = iteratedDeriv n f x - iteratedDeriv n g x := by diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index c818f6cfc6e58..e3061bb13fff2 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -5,6 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Convex.Topology import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.Analysis.Seminorm import Mathlib.Analysis.SpecificLimits.Basic /-! @@ -69,15 +70,17 @@ def UniqueDiffOn (s : Set E) : Prop := end TangentCone -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] -variable {𝕜} {x y : E} {s t : Set E} +variable {𝕜} +variable {E F G : Type*} section TangentCone -- This section is devoted to the properties of the tangent cone. + open NormedField +section TVS +variable [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] +variable {x y : E} {s t : Set E} theorem mem_tangentConeAt_of_pow_smul {r : 𝕜} (hr₀ : r ≠ 0) (hr : ‖r‖ < 1) (hs : ∀ᶠ n : ℕ in atTop, x + r ^ n • y ∈ s) : y ∈ tangentConeAt 𝕜 s x := by @@ -95,6 +98,14 @@ theorem tangentCone_mono (h : s ⊆ t) : tangentConeAt 𝕜 s x ⊆ tangentConeA rintro y ⟨c, d, ds, ctop, clim⟩ exact ⟨c, d, mem_of_superset ds fun n hn => h hn, ctop, clim⟩ +end TVS + +section Normed +variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] +variable [NormedAddCommGroup F] [NormedSpace 𝕜 F] +variable [NormedAddCommGroup G] [NormedSpace ℝ G] +variable {x y : E} {s t : Set E} + /-- Auxiliary lemma ensuring that, under the assumptions defining the tangent cone, the sequence `d` tends to 0 at infinity. -/ theorem tangentConeAt.lim_zero {α : Type*} (l : Filter α) {c : α → 𝕜} {d : α → E} @@ -207,6 +218,137 @@ theorem mem_tangentCone_of_segment_subset {s : Set G} {x y : G} (h : segment ℝ y - x ∈ tangentConeAt ℝ s x := mem_tangentCone_of_openSegment_subset ((openSegment_subset_segment ℝ x y).trans h) +/-- The tangent cone at a non-isolated point contains `0`. -/ +theorem zero_mem_tangentCone {s : Set E} {x : E} (hx : (𝓝[s \ {x}] x).NeBot) : + 0 ∈ tangentConeAt 𝕜 s x := by + /- Take a sequence `d n` tending to `0` such that `x + d n ∈ s`. Taking `c n` of the order + of `1 / (d n) ^ (1/2)`, then `c n` tends to infinity, but `c n • d n` tends to `0`. By definition, + this shows that `0` belongs to the tangent cone. -/ + obtain ⟨u, -, u_pos, u_lim⟩ : + ∃ u, StrictAnti u ∧ (∀ (n : ℕ), 0 < u n) ∧ Tendsto u atTop (𝓝 (0 : ℝ)) := + exists_seq_strictAnti_tendsto (0 : ℝ) + have A n : ((s \ {x}) ∩ Metric.ball x (u n * u n)).Nonempty := + NeBot.nonempty_of_mem hx (inter_mem_nhdsWithin _ + (Metric.ball_mem_nhds _ (mul_pos (u_pos n) (u_pos n)))) + choose v hv using A + let d n := v n - x + have M n : x + d n ∈ s \ {x} := by simpa [d] using (hv n).1 + let ⟨r, hr⟩ := exists_one_lt_norm 𝕜 + have W n := rescale_to_shell hr (u_pos n) (x := d n) (by simpa using (M n).2) + choose c c_ne c_le le_c hc using W + have c_lim : Tendsto (fun n ↦ ‖c n‖) atTop atTop := by + suffices Tendsto (fun n ↦ ‖c n‖⁻¹ ⁻¹) atTop atTop by simpa + apply tendsto_inv_nhdsGT_zero.comp + simp only [nhdsWithin, tendsto_inf, tendsto_principal, mem_Ioi, norm_pos_iff, ne_eq, + eventually_atTop, ge_iff_le] + have B (n : ℕ) : ‖c n‖⁻¹ ≤ ‖r‖ * u n := calc + ‖c n‖⁻¹ + _ ≤ (u n)⁻¹ * ‖r‖ * ‖d n‖ := hc n + _ ≤ (u n)⁻¹ * ‖r‖ * (u n * u n) := by + gcongr + · exact mul_nonneg (by simp [(u_pos n).le]) (norm_nonneg _) + · specialize hv n + simp only [mem_inter_iff, mem_diff, mem_singleton_iff, Metric.mem_ball, dist_eq_norm] + at hv + simpa using hv.2.le + _ = ‖r‖ * u n := by field_simp [(u_pos n).ne']; ring + refine ⟨?_, 0, fun n hn ↦ by simpa using c_ne n⟩ + apply squeeze_zero (fun n ↦ by positivity) B + simpa using u_lim.const_mul _ + refine ⟨c, d, Eventually.of_forall (fun n ↦ by simpa [d] using (hv n).1.1), c_lim, ?_⟩ + rw [tendsto_zero_iff_norm_tendsto_zero] + exact squeeze_zero (fun n ↦ by positivity) (fun n ↦ (c_le n).le) u_lim + +/-- In a proper space, the tangent cone at a non-isolated point is nontrivial. -/ +theorem tangentCone_nonempty_of_properSpace [ProperSpace E] + {s : Set E} {x : E} (hx : (𝓝[s \ {x}] x).NeBot) : + (tangentConeAt 𝕜 s x ∩ {0}ᶜ).Nonempty := by + /- Take a sequence `d n` tending to `0` such that `x + d n ∈ s`. Taking `c n` of the order + of `1 / d n`. Then `c n • d n` belongs to a fixed annulus. By compactness, one can extract + a subsequence converging to a limit `l`. Then `l` is nonzero, and by definition it belongs to + the tangent cone. -/ + obtain ⟨u, -, u_pos, u_lim⟩ : + ∃ u, StrictAnti u ∧ (∀ (n : ℕ), 0 < u n) ∧ Tendsto u atTop (𝓝 (0 : ℝ)) := + exists_seq_strictAnti_tendsto (0 : ℝ) + have A n : ((s \ {x}) ∩ Metric.ball x (u n)).Nonempty := by + apply NeBot.nonempty_of_mem hx (inter_mem_nhdsWithin _ (Metric.ball_mem_nhds _ (u_pos n))) + choose v hv using A + let d := fun n ↦ v n - x + have M n : x + d n ∈ s \ {x} := by simpa [d] using (hv n).1 + let ⟨r, hr⟩ := exists_one_lt_norm 𝕜 + have W n := rescale_to_shell hr zero_lt_one (x := d n) (by simpa using (M n).2) + choose c c_ne c_le le_c hc using W + have c_lim : Tendsto (fun n ↦ ‖c n‖) atTop atTop := by + suffices Tendsto (fun n ↦ ‖c n‖⁻¹ ⁻¹ ) atTop atTop by simpa + apply tendsto_inv_nhdsGT_zero.comp + simp only [nhdsWithin, tendsto_inf, tendsto_principal, mem_Ioi, norm_pos_iff, ne_eq, + eventually_atTop, ge_iff_le] + have B (n : ℕ) : ‖c n‖⁻¹ ≤ 1⁻¹ * ‖r‖ * u n := by + apply (hc n).trans + gcongr + specialize hv n + simp only [mem_inter_iff, mem_diff, mem_singleton_iff, Metric.mem_ball, dist_eq_norm] at hv + simpa using hv.2.le + refine ⟨?_, 0, fun n hn ↦ by simpa using c_ne n⟩ + apply squeeze_zero (fun n ↦ by positivity) B + simpa using u_lim.const_mul _ + obtain ⟨l, l_mem, φ, φ_strict, hφ⟩ : + ∃ l ∈ Metric.closedBall (0 : E) 1 \ Metric.ball (0 : E) (1 / ‖r‖), + ∃ (φ : ℕ → ℕ), StrictMono φ ∧ Tendsto ((fun n ↦ c n • d n) ∘ φ) atTop (𝓝 l) := by + apply IsCompact.tendsto_subseq _ (fun n ↦ ?_) + · exact (isCompact_closedBall 0 1).diff Metric.isOpen_ball + simp only [mem_diff, Metric.mem_closedBall, dist_zero_right, (c_le n).le, + Metric.mem_ball, not_lt, true_and, le_c n] + refine ⟨l, ?_, ?_⟩; swap + · simp only [mem_compl_iff, mem_singleton_iff] + contrapose! l_mem + simp only [one_div, l_mem, mem_diff, Metric.mem_closedBall, dist_self, zero_le_one, + Metric.mem_ball, inv_pos, norm_pos_iff, ne_eq, not_not, true_and] + contrapose! hr + simp [hr] + refine ⟨c ∘ φ, d ∘ φ, ?_, ?_, hφ⟩ + · exact Eventually.of_forall (fun n ↦ by simpa [d] using (hv (φ n)).1.1) + · exact c_lim.comp φ_strict.tendsto_atTop + +/-- The tangent cone at a non-isolated point in dimension 1 is the whole space. -/ +theorem tangentCone_eq_univ {s : Set 𝕜} {x : 𝕜} (hx : (𝓝[s \ {x}] x).NeBot) : + tangentConeAt 𝕜 s x = univ := by + apply eq_univ_iff_forall.2 (fun y ↦ ?_) + -- first deal with the case of `0`, which has to be handled separately. + rcases eq_or_ne y 0 with rfl | hy + · exact zero_mem_tangentCone hx + /- Assume now `y` is a fixed nonzero scalar. Take a sequence `d n` tending to `0` such + that `x + d n ∈ s`. Let `c n = y / d n`. Then `‖c n‖` tends to infinity, and `c n • d n` + converges to `y` (as it is equal to `y`). By definition, this shows that `y` belongs to the + tangent cone. -/ + obtain ⟨u, -, u_pos, u_lim⟩ : + ∃ u, StrictAnti u ∧ (∀ (n : ℕ), 0 < u n) ∧ Tendsto u atTop (𝓝 (0 : ℝ)) := + exists_seq_strictAnti_tendsto (0 : ℝ) + have A n : ((s \ {x}) ∩ Metric.ball x (u n)).Nonempty := by + apply NeBot.nonempty_of_mem hx (inter_mem_nhdsWithin _ (Metric.ball_mem_nhds _ (u_pos n))) + choose v hv using A + let d := fun n ↦ v n - x + have d_ne n : d n ≠ 0 := by + simp only [mem_inter_iff, mem_diff, mem_singleton_iff, Metric.mem_ball, d] at hv + simpa [d, sub_ne_zero] using (hv n).1.2 + refine ⟨fun n ↦ y * (d n)⁻¹, d, ?_, ?_, ?_⟩ + · exact Eventually.of_forall (fun n ↦ by simpa [d] using (hv n).1.1) + · simp only [norm_mul, norm_inv] + apply (tendsto_const_mul_atTop_of_pos (by simpa using hy)).2 + apply tendsto_inv_nhdsGT_zero.comp + simp only [nhdsWithin, tendsto_inf, tendsto_principal, mem_Ioi, norm_pos_iff, ne_eq, + eventually_atTop, ge_iff_le] + have B (n : ℕ) : ‖d n‖ ≤ u n := by + specialize hv n + simp only [mem_inter_iff, mem_diff, mem_singleton_iff, Metric.mem_ball, dist_eq_norm] at hv + simpa using hv.2.le + refine ⟨?_, 0, fun n hn ↦ by simpa using d_ne n⟩ + exact squeeze_zero (fun n ↦ by positivity) B u_lim + · convert tendsto_const_nhds (α := ℕ) (x := y) with n + simp [mul_assoc, inv_mul_cancel₀ (d_ne n)] + +end Normed + end TangentCone section UniqueDiff @@ -216,6 +358,9 @@ section UniqueDiff This section is devoted to properties of the predicates `UniqueDiffWithinAt` and `UniqueDiffOn`. -/ +section TVS +variable [AddCommGroup E] [Module 𝕜 E] [TopologicalSpace E] +variable {x y : E} {s t : Set E} theorem UniqueDiffOn.uniqueDiffWithinAt {s : Set E} {x} (hs : UniqueDiffOn 𝕜 s) (h : x ∈ s) : UniqueDiffWithinAt 𝕜 s x := @@ -234,6 +379,13 @@ theorem uniqueDiffOn_empty : UniqueDiffOn 𝕜 (∅ : Set E) := theorem UniqueDiffWithinAt.congr_pt (h : UniqueDiffWithinAt 𝕜 s x) (hy : x = y) : UniqueDiffWithinAt 𝕜 s y := hy ▸ h +end TVS + +section Normed +variable [NormedAddCommGroup E] [NormedSpace 𝕜 E] +variable [NormedAddCommGroup F] [NormedSpace 𝕜 F] +variable {x y : E} {s t : Set E} + theorem UniqueDiffWithinAt.mono_nhds (h : UniqueDiffWithinAt 𝕜 s x) (st : 𝓝[s] x ≤ 𝓝[t] x) : UniqueDiffWithinAt 𝕜 t x := by simp only [uniqueDiffWithinAt_iff] at * @@ -327,6 +479,11 @@ theorem UniqueDiffOn.univ_pi (ι : Type*) [Finite ι] (E : ι → Type*) (h : ∀ i, UniqueDiffOn 𝕜 (s i)) : UniqueDiffOn 𝕜 (Set.pi univ s) := UniqueDiffOn.pi _ _ _ _ fun i _ => h i +end Normed + +section RealNormed +variable [NormedAddCommGroup G] [NormedSpace ℝ G] + /-- In a real vector space, a convex set with nonempty interior is a set of unique differentiability at every point of its closure. -/ theorem uniqueDiffWithinAt_convex {s : Set G} (conv : Convex ℝ s) (hs : (interior s).Nonempty) @@ -349,6 +506,10 @@ theorem uniqueDiffOn_convex {s : Set G} (conv : Convex ℝ s) (hs : (interior s) UniqueDiffOn ℝ s := fun _ xs => uniqueDiffWithinAt_convex conv hs (subset_closure xs) +end RealNormed + +section Real + theorem uniqueDiffOn_Ici (a : ℝ) : UniqueDiffOn ℝ (Ici a) := uniqueDiffOn_convex (convex_Ici a) <| by simp only [interior_Ici, nonempty_Ioi] @@ -391,4 +552,17 @@ theorem uniqueDiffWithinAt_Ioi (a : ℝ) : UniqueDiffWithinAt ℝ (Ioi a) a := theorem uniqueDiffWithinAt_Iio (a : ℝ) : UniqueDiffWithinAt ℝ (Iio a) a := uniqueDiffWithinAt_convex (convex_Iio a) (by simp) (by simp) +/-- In one dimension, every point is either a point of unique differentiability, or isolated. -/ +theorem uniqueDiffWithinAt_or_nhdsWithin_eq_bot (s : Set 𝕜) (x : 𝕜) : + UniqueDiffWithinAt 𝕜 s x ∨ 𝓝[s \ {x}] x = ⊥ := by + rcases eq_or_neBot (𝓝[s \ {x}] x) with h | h + · exact Or.inr h + refine Or.inl ⟨?_, ?_⟩ + · simp [tangentCone_eq_univ h] + · simp only [mem_closure_iff_nhdsWithin_neBot] + apply neBot_of_le (hf := h) + exact nhdsWithin_mono _ diff_subset + +end Real + end UniqueDiff diff --git a/Mathlib/Analysis/Calculus/Taylor.lean b/Mathlib/Analysis/Calculus/Taylor.lean index 7d6b563e1c66a..a0021a92c18a4 100644 --- a/Mathlib/Analysis/Calculus/Taylor.lean +++ b/Mathlib/Analysis/Calculus/Taylor.lean @@ -145,7 +145,7 @@ theorem hasDerivWithinAt_taylor_coeff_within {f : ℝ → E} {x y : ℝ} {k : replace hf : HasDerivWithinAt (iteratedDerivWithin (k + 1) f s) (iteratedDerivWithin (k + 2) f s y) t y := by convert (hf.mono_of_mem_nhdsWithin hs).hasDerivWithinAt using 1 - rw [iteratedDerivWithin_succ (ht.mono_nhds (nhdsWithin_le_iff.mpr hs))] + rw [iteratedDerivWithin_succ] exact (derivWithin_of_mem_nhdsWithin hs ht hf).symm have : HasDerivWithinAt (fun t => ((k + 1 : ℝ) * k !)⁻¹ * (x - t) ^ (k + 1)) (-((k ! : ℝ)⁻¹ * (x - y) ^ k)) t y := by @@ -162,17 +162,19 @@ theorem hasDerivWithinAt_taylor_coeff_within {f : ℝ → E} {x y : ℝ} {k : Version for arbitrary sets -/ theorem hasDerivWithinAt_taylorWithinEval {f : ℝ → E} {x y : ℝ} {n : ℕ} {s s' : Set ℝ} - (hs'_unique : UniqueDiffWithinAt ℝ s' y) (hs_unique : UniqueDiffOn ℝ s) (hs' : s' ∈ 𝓝[s] y) + (hs_unique : UniqueDiffOn ℝ s) (hs' : s' ∈ 𝓝[s] y) (hy : y ∈ s') (h : s' ⊆ s) (hf : ContDiffOn ℝ n f s) (hf' : DifferentiableWithinAt ℝ (iteratedDerivWithin n f s) s y) : HasDerivWithinAt (fun t => taylorWithinEval f n s t x) (((n ! : ℝ)⁻¹ * (x - y) ^ n) • iteratedDerivWithin (n + 1) f s y) s' y := by + have hs'_unique : UniqueDiffWithinAt ℝ s' y := + UniqueDiffWithinAt.mono_nhds (hs_unique _ (h hy)) (nhdsWithin_le_iff.mpr hs') induction n with | zero => simp only [taylor_within_zero_eval, Nat.factorial_zero, Nat.cast_one, inv_one, pow_zero, mul_one, zero_add, one_smul] simp only [iteratedDerivWithin_zero] at hf' - rw [iteratedDerivWithin_one (hs_unique _ (h hy))] + rw [iteratedDerivWithin_one] exact hf'.hasDerivWithinAt.mono h | succ k hk => simp_rw [Nat.add_succ, taylorWithinEval_succ] @@ -195,7 +197,7 @@ theorem taylorWithinEval_hasDerivAt_Ioo {f : ℝ → E} {a b t : ℝ} (x : ℝ) (((n ! : ℝ)⁻¹ * (x - t) ^ n) • iteratedDerivWithin (n + 1) f (Icc a b) t) t := have h_nhds : Ioo a b ∈ 𝓝 t := isOpen_Ioo.mem_nhds ht have h_nhds' : Ioo a b ∈ 𝓝[Icc a b] t := nhdsWithin_le_nhds h_nhds - (hasDerivWithinAt_taylorWithinEval (uniqueDiffWithinAt_Ioo ht) (uniqueDiffOn_Icc hx) h_nhds' ht + (hasDerivWithinAt_taylorWithinEval (uniqueDiffOn_Icc hx) h_nhds' ht Ioo_subset_Icc_self hf <| (hf' t ht).mono_of_mem_nhdsWithin h_nhds').hasDerivAt h_nhds /-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. @@ -206,7 +208,7 @@ theorem hasDerivWithinAt_taylorWithinEval_at_Icc {f : ℝ → E} {a b t : ℝ} ( (hf' : DifferentiableOn ℝ (iteratedDerivWithin n f (Icc a b)) (Icc a b)) : HasDerivWithinAt (fun y => taylorWithinEval f n (Icc a b) y x) (((n ! : ℝ)⁻¹ * (x - t) ^ n) • iteratedDerivWithin (n + 1) f (Icc a b) t) (Icc a b) t := - hasDerivWithinAt_taylorWithinEval (uniqueDiffOn_Icc hx t ht) (uniqueDiffOn_Icc hx) + hasDerivWithinAt_taylorWithinEval (uniqueDiffOn_Icc hx) self_mem_nhdsWithin ht rfl.subset hf (hf' t ht) /-! ### Taylor's theorem with mean value type remainder estimate -/ diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 73c4f7ae5cdfa..af99bc7e8dd7f 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Data.Complex.Module import Mathlib.Data.Complex.Order -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric import Mathlib.Analysis.RCLike.Basic import Mathlib.Topology.Algebra.InfiniteSum.Field import Mathlib.Topology.Algebra.InfiniteSum.Module diff --git a/Mathlib/Analysis/Complex/CauchyIntegral.lean b/Mathlib/Analysis/Complex/CauchyIntegral.lean index 99b3e0d3c22dc..68eaaa612c9ce 100644 --- a/Mathlib/Analysis/Complex/CauchyIntegral.lean +++ b/Mathlib/Analysis/Complex/CauchyIntegral.lean @@ -343,7 +343,7 @@ theorem circleIntegral_sub_center_inv_smul_of_differentiable_on_off_countable_of (hd : ∀ z ∈ (ball c R \ {c}) \ s, DifferentiableAt ℂ f z) (hy : Tendsto f (𝓝[{c}ᶜ] c) (𝓝 y)) : (∮ z in C(c, R), (z - c)⁻¹ • f z) = (2 * π * I : ℂ) • y := by rw [← sub_eq_zero, ← norm_le_zero_iff] - refine le_of_forall_le_of_dense fun ε ε0 => ?_ + refine le_of_forall_gt_imp_ge_of_dense fun ε ε0 => ?_ obtain ⟨δ, δ0, hδ⟩ : ∃ δ > (0 : ℝ), ∀ z ∈ closedBall c δ \ {c}, dist (f z) y < ε / (2 * π) := ((nhdsWithin_hasBasis nhds_basis_closedBall _).tendsto_iff nhds_basis_ball).1 hy _ (div_pos ε0 Real.two_pi_pos) diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index aee36657dab9e..75d46597fd3d1 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -89,7 +89,7 @@ theorem liouville_theorem_aux {f : ℂ → F} (hf : Differentiable ℂ f) (hb : exact ⟨max C 1, lt_max_iff.2 (Or.inr zero_lt_one), fun z => (hC (f z) (mem_range_self _)).trans (le_max_left _ _)⟩ - refine norm_le_zero_iff.1 (le_of_forall_le_of_dense fun ε ε₀ => ?_) + refine norm_le_zero_iff.1 (le_of_forall_gt_imp_ge_of_dense fun ε ε₀ => ?_) calc ‖deriv f c‖ ≤ C / (C / ε) := norm_deriv_le_of_forall_mem_sphere_norm_le (div_pos C₀ ε₀) hf.diffContOnCl fun z _ => hC z diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index ee37b9bc62b36..1d6cb0068727f 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -116,7 +116,7 @@ theorem horizontal_strip (hfd : DiffContOnCl ℂ f (im ⁻¹' Ioo a b)) cases' hza with hza hza; · exact hle_a _ hza.symm cases' hzb with hzb hzb; · exact hle_b _ hzb wlog hC₀ : 0 < C generalizing C - · refine le_of_forall_le_of_dense fun C' hC' => this (fun w hw => ?_) (fun w hw => ?_) ?_ + · refine le_of_forall_gt_imp_ge_of_dense fun C' hC' => this (fun w hw => ?_) (fun w hw => ?_) ?_ · exact (hle_a _ hw).trans hC'.le · exact (hle_b _ hw).trans hC'.le · refine ((norm_nonneg (f (a * I))).trans (hle_a _ ?_)).trans_lt hC' diff --git a/Mathlib/Analysis/Complex/Polynomial/Basic.lean b/Mathlib/Analysis/Complex/Polynomial/Basic.lean index 706b89c25be7f..b3296662c3ab7 100644 --- a/Mathlib/Analysis/Complex/Polynomial/Basic.lean +++ b/Mathlib/Analysis/Complex/Polynomial/Basic.lean @@ -159,15 +159,7 @@ theorem galActionHom_bijective_of_prime_degree' {p : ℚ[X]} (p_irr : Irreducibl MonoidHom.map_one, MonoidHom.map_one]) have key := card_complex_roots_eq_card_real_add_card_not_gal_inv p simp_rw [Set.toFinset_card] at key - rw [key, add_le_add_iff_left] at p_roots1 p_roots2 - rw [key, add_right_inj] - suffices ∀ m : ℕ, 2 ∣ m → 1 ≤ m → m ≤ 3 → m = 2 by exact this n hn p_roots1 p_roots2 - rintro m ⟨k, rfl⟩ h2 h3 - exact le_antisymm - (Nat.lt_succ_iff.mp - (lt_of_le_of_ne h3 (show 2 * k ≠ 2 * 1 + 1 from Nat.two_mul_ne_two_mul_add_one))) - (Nat.succ_le_iff.mpr - (lt_of_le_of_ne h2 (show 2 * 0 + 1 ≠ 2 * k from Nat.two_mul_ne_two_mul_add_one.symm))) + omega end Rationals diff --git a/Mathlib/Analysis/Complex/UnitDisc/Basic.lean b/Mathlib/Analysis/Complex/UnitDisc/Basic.lean index adaa04bfb5b12..d5d3b8b0e4d30 100644 --- a/Mathlib/Analysis/Complex/UnitDisc/Basic.lean +++ b/Mathlib/Analysis/Complex/UnitDisc/Basic.lean @@ -22,7 +22,7 @@ local notation "conj'" => starRingEnd ℂ namespace Complex -/-- Complex unit disc. -/ +/-- The complex unit disc, denoted as `𝔻` withinin the Complex namespace -/ def UnitDisc : Type := ball (0 : ℂ) 1 deriving TopologicalSpace diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 3a538cfd9d8ac..f2a07cb6618ce 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -24,7 +24,7 @@ noncomputable section open Matrix Matrix.SpecialLinearGroup open scoped MatrixGroups -/-- The open upper half plane -/ +/-- The open upper half plane, denoted as `ℍ` within the `UpperHalfPlane` namespace -/ def UpperHalfPlane := { point : ℂ // 0 < point.im } diff --git a/Mathlib/Analysis/Convex/Caratheodory.lean b/Mathlib/Analysis/Convex/Caratheodory.lean index 59b4f09c6a3ec..1bdb62e308505 100644 --- a/Mathlib/Analysis/Convex/Caratheodory.lean +++ b/Mathlib/Analysis/Convex/Caratheodory.lean @@ -100,17 +100,17 @@ variable {s : Set E} {x : E} /-- Given a point `x` in the convex hull of a set `s`, this is a finite subset of `s` of minimum cardinality, whose convex hull contains `x`. -/ noncomputable def minCardFinsetOfMemConvexHull (hx : x ∈ convexHull 𝕜 s) : Finset E := - Function.argminOn Finset.card Nat.lt_wfRel.2 { t | ↑t ⊆ s ∧ x ∈ convexHull 𝕜 (t : Set E) } <| by + Function.argminOn Finset.card { t | ↑t ⊆ s ∧ x ∈ convexHull 𝕜 (t : Set E) } <| by simpa only [convexHull_eq_union_convexHull_finite_subsets s, exists_prop, mem_iUnion] using hx variable (hx : x ∈ convexHull 𝕜 s) theorem minCardFinsetOfMemConvexHull_subseteq : ↑(minCardFinsetOfMemConvexHull hx) ⊆ s := - (Function.argminOn_mem _ _ { t : Finset E | ↑t ⊆ s ∧ x ∈ convexHull 𝕜 (t : Set E) } _).1 + (Function.argminOn_mem _ { t : Finset E | ↑t ⊆ s ∧ x ∈ convexHull 𝕜 (t : Set E) } _).1 theorem mem_minCardFinsetOfMemConvexHull : x ∈ convexHull 𝕜 (minCardFinsetOfMemConvexHull hx : Set E) := - (Function.argminOn_mem _ _ { t : Finset E | ↑t ⊆ s ∧ x ∈ convexHull 𝕜 (t : Set E) } _).2 + (Function.argminOn_mem _ { t : Finset E | ↑t ⊆ s ∧ x ∈ convexHull 𝕜 (t : Set E) } _).2 theorem minCardFinsetOfMemConvexHull_nonempty : (minCardFinsetOfMemConvexHull hx).Nonempty := by rw [← Finset.coe_nonempty, ← @convexHull_nonempty_iff 𝕜] @@ -118,7 +118,7 @@ theorem minCardFinsetOfMemConvexHull_nonempty : (minCardFinsetOfMemConvexHull hx theorem minCardFinsetOfMemConvexHull_card_le_card {t : Finset E} (ht₁ : ↑t ⊆ s) (ht₂ : x ∈ convexHull 𝕜 (t : Set E)) : #(minCardFinsetOfMemConvexHull hx) ≤ #t := - Function.argminOn_le _ _ _ (by exact ⟨ht₁, ht₂⟩) + Function.argminOn_le _ _ (by exact ⟨ht₁, ht₂⟩) theorem affineIndependent_minCardFinsetOfMemConvexHull : AffineIndependent 𝕜 ((↑) : minCardFinsetOfMemConvexHull hx → E) := by diff --git a/Mathlib/Analysis/Convex/Deriv.lean b/Mathlib/Analysis/Convex/Deriv.lean index 65dbae3b210de..40ff12e20d1fe 100644 --- a/Mathlib/Analysis/Convex/Deriv.lean +++ b/Mathlib/Analysis/Convex/Deriv.lean @@ -395,11 +395,14 @@ lemma le_slope_of_hasDerivWithinAt_Ioi (hfc : ConvexOn ℝ S f) exact hfc.1.ordConnected.out hx hy ⟨ht'.le, ht.le⟩ /-- Reformulation of `ConvexOn.le_slope_of_hasDerivWithinAt_Ioi` using `derivWithin`. -/ -lemma right_deriv_le_slope (hfc : ConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma rightDeriv_le_slope (hfc : ConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Ioi x) x) : derivWithin f (Ioi x) x ≤ slope f x y := le_slope_of_hasDerivWithinAt_Ioi hfc hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias right_deriv_le_slope := rightDeriv_le_slope + /-- If `f : ℝ → ℝ` is convex on `S` and differentiable within `S` at `x`, then the slope of any secant line with left endpoint at `x` is bounded below by the derivative of `f` within `S` at `x`. @@ -449,11 +452,14 @@ lemma slope_le_of_hasDerivWithinAt_Iio (hfc : ConvexOn ℝ S f) exact hfc.1.ordConnected.out hx hy ⟨ht.le, ht'.le⟩ /-- Reformulation of `ConvexOn.slope_le_of_hasDerivWithinAt_Iio` using `derivWithin`. -/ -lemma slope_le_left_deriv (hfc : ConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma slope_le_leftDeriv (hfc : ConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Iio y) y) : slope f x y ≤ derivWithin f (Iio y) y := hfc.slope_le_of_hasDerivWithinAt_Iio hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias slope_le_left_deriv := slope_le_leftDeriv + /-- If `f : ℝ → ℝ` is convex on `S` and differentiable within `S` at `y`, then the slope of any secant line with right endpoint at `y` is bounded above by the derivative of `f` within `S` at `y`. @@ -531,11 +537,14 @@ lemma lt_slope_of_hasDerivWithinAt_Ioi (hfc : StrictConvexOn ℝ S f) simp only [← slope_def_field] at this exact (hfc.convexOn.le_slope_of_hasDerivWithinAt_Ioi hx hu hxu hf').trans_lt this -lemma right_deriv_lt_slope (hfc : StrictConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma rightDeriv_lt_slope (hfc : StrictConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Ioi x) x) : derivWithin f (Ioi x) x < slope f x y := hfc.lt_slope_of_hasDerivWithinAt_Ioi hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias right_deriv_lt_slope := rightDeriv_lt_slope + /-- If `f : ℝ → ℝ` is strictly convex on `S` and differentiable within `S` at `x ∈ S`, then the slope of any secant line with left endpoint at `x` is strictly greater than the derivative of `f` within `S` at `x`. @@ -583,11 +592,14 @@ lemma slope_lt_of_hasDerivWithinAt_Iio (hfc : StrictConvexOn ℝ S f) simp_rw [← slope_def_field, slope_comm _ y] at this exact this.trans_le <| hfc.convexOn.slope_le_of_hasDerivWithinAt_Iio hu hy huy hf' -lemma slope_lt_left_deriv (hfc : StrictConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma slope_lt_leftDeriv (hfc : StrictConvexOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Iio y) y) : slope f x y < derivWithin f (Iio y) y := hfc.slope_lt_of_hasDerivWithinAt_Iio hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias slope_lt_left_deriv := slope_lt_leftDeriv + /-- If `f : ℝ → ℝ` is strictly convex on `S` and differentiable within `S` at `y ∈ S`, then the slope of any secant line with right endpoint at `y` is strictly less than the derivative of `f` within `S` at `y`. @@ -657,11 +669,14 @@ lemma slope_le_of_hasDerivWithinAt_Ioi (hfc : ConcaveOn ℝ S f) simpa only [Pi.neg_def, slope_neg, neg_neg] using neg_le_neg (hfc.neg.le_slope_of_hasDerivWithinAt_Ioi hx hy hxy hf'.neg) -lemma slope_le_right_deriv (hfc : ConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma slope_le_rightDeriv (hfc : ConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Ioi x) x) : slope f x y ≤ derivWithin f (Ioi x) x := hfc.slope_le_of_hasDerivWithinAt_Ioi hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias slope_le_right_deriv := slope_le_rightDeriv + lemma slope_le_of_hasDerivWithinAt (hfc : ConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : HasDerivWithinAt f f' S x) : slope f x y ≤ f' := @@ -696,11 +711,14 @@ lemma le_slope_of_hasDerivWithinAt_Iio (hfc : ConcaveOn ℝ S f) simpa only [neg_neg, Pi.neg_def, slope_neg] using neg_le_neg (hfc.neg.slope_le_of_hasDerivWithinAt_Iio hx hy hxy hf'.neg) -lemma left_deriv_le_slope (hfc : ConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma leftDeriv_le_slope (hfc : ConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Iio y) y) : derivWithin f (Iio y) y ≤ slope f x y := hfc.le_slope_of_hasDerivWithinAt_Iio hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias left_deriv_le_slope := leftDeriv_le_slope + lemma le_slope_of_hasDerivWithinAt (hfc : ConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hf' : HasDerivWithinAt f f' S y) : f' ≤ slope f x y := @@ -757,11 +775,14 @@ lemma slope_lt_of_hasDerivWithinAt_Ioi (hfc : StrictConcaveOn ℝ S f) simpa only [Pi.neg_def, slope_neg, neg_neg] using neg_lt_neg (hfc.neg.lt_slope_of_hasDerivWithinAt_Ioi hx hy hxy hf'.neg) -lemma slope_lt_right_deriv (hfc : StrictConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma slope_lt_rightDeriv (hfc : StrictConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Ioi x) x) : slope f x y < derivWithin f (Ioi x) x := hfc.slope_lt_of_hasDerivWithinAt_Ioi hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias slope_lt_right_deriv := slope_lt_rightDeriv + lemma slope_lt_of_hasDerivWithinAt (hfc : StrictConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : HasDerivWithinAt f f' S x) : slope f x y < f' := by @@ -797,11 +818,14 @@ lemma lt_slope_of_hasDerivWithinAt_Iio (hfc : StrictConcaveOn ℝ S f) simpa only [Pi.neg_def, slope_neg, neg_neg] using neg_lt_neg (hfc.neg.slope_lt_of_hasDerivWithinAt_Iio hx hy hxy hf'.neg) -lemma left_deriv_lt_slope (hfc : StrictConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) +lemma leftDeriv_lt_slope (hfc : StrictConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hfd : DifferentiableWithinAt ℝ f (Iio y) y) : derivWithin f (Iio y) y < slope f x y := hfc.lt_slope_of_hasDerivWithinAt_Iio hx hy hxy hfd.hasDerivWithinAt +@[deprecated (since := "2025-01-26")] +alias left_deriv_lt_slope := leftDeriv_lt_slope + lemma lt_slope_of_hasDerivWithinAt (hfc : StrictConcaveOn ℝ S f) (hx : x ∈ S) (hy : y ∈ S) (hxy : x < y) (hf' : HasDerivWithinAt f f' S y) : f' < slope f x y := by diff --git a/Mathlib/Analysis/Convex/EGauge.lean b/Mathlib/Analysis/Convex/EGauge.lean index 5083bf5351560..3680e97a36527 100644 --- a/Mathlib/Analysis/Convex/EGauge.lean +++ b/Mathlib/Analysis/Convex/EGauge.lean @@ -67,7 +67,7 @@ lemma egauge_union (s t : Set E) (x : E) : egauge 𝕜 (s ∪ t) x = egauge 𝕜 lemma le_egauge_inter (s t : Set E) (x : E) : egauge 𝕜 s x ⊔ egauge 𝕜 t x ≤ egauge 𝕜 (s ∩ t) x := - max_le_iff.2 ⟨egauge_anti _ inter_subset_left _, egauge_anti _ inter_subset_right _⟩ + max_le (egauge_anti _ inter_subset_left _) (egauge_anti _ inter_subset_right _) end SMul @@ -168,6 +168,23 @@ lemma egauge_smul_right (h : c = 0 → s.Nonempty) (x : E) : end Module +section VectorSpace + +variable {𝕜 : Type*} [NormedField 𝕜] {E : Type*} [AddCommGroup E] [Module 𝕜 E] + +theorem egauge_add_add_le {U V : Set E} (hU : Balanced 𝕜 U) (hV : Balanced 𝕜 V) (a b : E) : + egauge 𝕜 (U + V) (a + b) ≤ max (egauge 𝕜 U a) (egauge 𝕜 V b) := by + refine le_of_forall_lt' fun c hc ↦ ?_ + simp only [max_lt_iff, egauge_lt_iff] at hc ⊢ + rcases hc with ⟨⟨x, hx, hxc⟩, ⟨y, hy, hyc⟩⟩ + wlog hxy : ‖x‖ ≤ ‖y‖ generalizing a b x y U V + · simpa only [add_comm] using this hV hU b a y hy hyc x hx hxc (le_of_not_le hxy) + refine ⟨y, ?_, hyc⟩ + rw [smul_add] + exact add_mem_add (hU.smul_mono hxy hx) hy + +end VectorSpace + section SeminormedAddCommGroup variable (𝕜 : Type*) [NormedField 𝕜] {E : Type*} [SeminormedAddCommGroup E] [NormedSpace 𝕜 E] diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 2277ac898fcf3..ee456daba0041 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -146,7 +146,7 @@ lemma exists_mem_interior_convexHull_affineBasis (hs : s ∈ 𝓝 x) : wlog hx : x = 0 · obtain ⟨b, hb⟩ := this (s := -x +ᵥ s) (by simpa using vadd_mem_nhds_vadd (-x) hs) rfl use x +ᵥ b - simpa [subset_set_vadd_iff, mem_vadd_set_iff_neg_vadd_mem, convexHull_vadd, interior_vadd, + simpa [subset_vadd_set_iff, mem_vadd_set_iff_neg_vadd_mem, convexHull_vadd, interior_vadd, Pi.vadd_def, -vadd_eq_add, vadd_eq_add (a := -x), ← Set.vadd_set_range] using hb subst hx -- The strategy is now to find an arbitrary maximal spanning simplex (aka an affine basis)... @@ -171,7 +171,7 @@ lemma exists_mem_interior_convexHull_affineBasis (hs : s ∈ 𝓝 x) : set d : AffineBasis (Fin (finrank ℝ E + 1)) ℝ E := Units.mk0 ε' hε'.ne' • c have hε₀ : 0 < ε / 2 := by positivity have hdnorm : (range d : Set E) ⊆ closedBall 0 (ε / 2) := by - simp [d, Set.set_smul_subset_iff₀ hε'.ne', hε₀.le, _root_.smul_closedBall, abs_of_nonneg hε'.le, + simp [d, Set.smul_set_subset_iff₀ hε'.ne', hε₀.le, _root_.smul_closedBall, abs_of_nonneg hε'.le, range_subset_iff, norm_smul] simpa [ε', hε₀.ne', range_subset_iff, ← mul_div_right_comm (ε / 2), div_le_iff₀ hc', mul_le_mul_left hε₀] using hcnorm diff --git a/Mathlib/Analysis/Convex/Segment.lean b/Mathlib/Analysis/Convex/Segment.lean index 1c4e577df9e16..13d7ab381b215 100644 --- a/Mathlib/Analysis/Convex/Segment.lean +++ b/Mathlib/Analysis/Convex/Segment.lean @@ -48,7 +48,8 @@ def segment (x y : E) : Set E := { z : E | ∃ a b : 𝕜, 0 ≤ a ∧ 0 ≤ b ∧ a + b = 1 ∧ a • x + b • y = z } /-- Open segment in a vector space. Note that `openSegment 𝕜 x x = {x}` instead of being `∅` when -the base semiring has some element between `0` and `1`. -/ +the base semiring has some element between `0` and `1`. +Denoted as `[x -[𝕜] y]` within the `Convex` namespace. -/ def openSegment (x y : E) : Set E := { z : E | ∃ a b : 𝕜, 0 < a ∧ 0 < b ∧ a + b = 1 ∧ a • x + b • y = z } diff --git a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean index fb149455f547a..d43a840accbf1 100644 --- a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean +++ b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean @@ -198,8 +198,8 @@ theorem ae_eq_of_integral_contDiff_smul_eq (fun g g_diff g_supp ↦ h g g_diff.contDiff g_supp) /-- If a function `f` locally integrable on an open subset `U` of a finite-dimensional real - manifold has zero integral when multiplied by any smooth function compactly supported - in an open set `U`, then `f` vanishes almost everywhere in `U`. -/ + vector space has zero integral when multiplied by any smooth function compactly supported + in `U`, then `f` vanishes almost everywhere in `U`. -/ theorem IsOpen.ae_eq_zero_of_integral_contDiff_smul_eq_zero {U : Set E} (hU : IsOpen U) (hf : LocallyIntegrableOn f U μ) (h : ∀ (g : E → ℝ), ContDiff ℝ ∞ g → HasCompactSupport g → tsupport g ⊆ U → diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index c75680f9ddcd1..dcde87aa5b436 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -298,7 +298,8 @@ open scoped Real namespace Real -/-- The standard additive character of `ℝ`, given by `fun x ↦ exp (2 * π * x * I)`. -/ +/-- The standard additive character of `ℝ`, given by `fun x ↦ exp (2 * π * x * I)`. +Denoted as `𝐞` within the `Real.FourierTransform` namespace. -/ def fourierChar : AddChar ℝ 𝕊 where toFun z := .exp (2 * π * z) map_zero_eq_one' := by simp only; rw [mul_zero, Circle.exp_zero] @@ -372,12 +373,14 @@ open scoped RealInnerProductSpace variable [FiniteDimensional ℝ V] /-- The Fourier transform of a function on an inner product space, with respect to the standard -additive character `ω ↦ exp (2 i π ω)`. -/ +additive character `ω ↦ exp (2 i π ω)`. +Denoted as `𝓕` within the `Real.FourierTransform` namespace. -/ def fourierIntegral (f : V → E) (w : V) : E := VectorFourier.fourierIntegral 𝐞 volume (innerₗ V) f w /-- The inverse Fourier transform of a function on an inner product space, defined as the Fourier -transform but with opposite sign in the exponential. -/ +transform but with opposite sign in the exponential. +Denoted as `𝓕⁻¹` within the `Real.FourierTransform` namespace. -/ def fourierIntegralInv (f : V → E) (w : V) : E := VectorFourier.fourierIntegral 𝐞 volume (-innerₗ V) f w diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index fe95597e697b1..da429e36827e9 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -806,9 +806,7 @@ theorem iteratedDeriv_fourierIntegral {f : ℝ → E} {N : ℕ∞} {n : ℕ} have A (n : ℕ) (hn : n ≤ N) : Integrable (fun v ↦ ‖v‖^n * ‖f v‖) := by convert (hf n hn).norm with x simp [norm_smul] - have B : AEStronglyMeasurable f := by - convert (hf 0 (zero_le _)).1 with x - simp + have B : AEStronglyMeasurable f := by simpa using (hf 0 (zero_le _)).1 rw [iteratedDeriv, iteratedFDeriv_fourierIntegral A B hn, fourierIntegral_continuousMultilinearMap_apply (integrable_fourierPowSMulRight _ (A n hn) B), fourierIntegral_eq, fourierIntegral_eq] diff --git a/Mathlib/Analysis/Fourier/ZMod.lean b/Mathlib/Analysis/Fourier/ZMod.lean index ae3c4f996bbcc..22825977656de 100644 --- a/Mathlib/Analysis/Fourier/ZMod.lean +++ b/Mathlib/Analysis/Fourier/ZMod.lean @@ -76,7 +76,7 @@ section defs /-- The discrete Fourier transform on `ℤ / N ℤ` (with the counting measure), bundled as a linear -equivalence. +equivalence. Denoted as `𝓕` within the `ZMod` namespace. -/ noncomputable def dft : (ZMod N → E) ≃ₗ[ℂ] (ZMod N → E) where toFun := auxDFT diff --git a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean index 83a790f784a3c..ad17c521ac625 100644 --- a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean +++ b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean @@ -99,11 +99,13 @@ theorem adjointAux_norm (A : E →L[𝕜] F) : ‖adjointAux A‖ = ‖A‖ := b rw [adjointAux_apply, LinearIsometryEquiv.norm_map] exact toSesqForm_apply_norm_le -/-- The adjoint of a bounded operator from Hilbert space `E` to Hilbert space `F`. -/ +/-- The adjoint of a bounded operator `A` from a Hilbert space `E` to another Hilbert space `F`, + denoted as `A†`. -/ def adjoint : (E →L[𝕜] F) ≃ₗᵢ⋆[𝕜] F →L[𝕜] E := LinearIsometryEquiv.ofSurjective { adjointAux with norm_map' := adjointAux_norm } fun A => ⟨adjointAux A, adjointAux_adjointAux A⟩ +@[inherit_doc] scoped[InnerProduct] postfix:1000 "†" => ContinuousLinearMap.adjoint open InnerProduct diff --git a/Mathlib/Analysis/InnerProductSpace/Defs.lean b/Mathlib/Analysis/InnerProductSpace/Defs.lean index c89fcb0886d38..8b3140b2eae8a 100644 --- a/Mathlib/Analysis/InnerProductSpace/Defs.lean +++ b/Mathlib/Analysis/InnerProductSpace/Defs.lean @@ -193,12 +193,16 @@ variable [AddCommGroup F] [Module 𝕜 F] [c : PreInnerProductSpace.Core 𝕜 F] local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y +/-- Local notation for `RCLike.normSq 𝕜` -/ local notation "normSqK" => @RCLike.normSq 𝕜 _ +/-- Local notation for `RCLike.re 𝕜` -/ local notation "reK" => @RCLike.re 𝕜 _ +/-- Local notation for `RCLike.ext_iff 𝕜` -/ local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ +/-- Local notation for `starRingEnd _` -/ local postfix:90 "†" => starRingEnd _ /-- Inner product defined by the `PreInnerProductSpace.Core` structure. We can't reuse @@ -213,6 +217,7 @@ attribute [local instance] toPreInner' def normSq (x : F) := reK ⟪x, x⟫ +/-- The norm squared function for `PreInnerProductSpace.Core` structure. -/ local notation "normSqF" => @normSq 𝕜 F _ _ _ _ theorem inner_conj_symm (x y : F) : ⟪y, x⟫† = ⟪x, y⟫ := diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index e80e7b0443eb7..8c535957fea7e 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -143,11 +143,12 @@ theorem adjointAux_unique (y : T.adjointDomain) {x₀ : E} variable (T) open scoped Classical in -/-- The adjoint operator as a partially defined linear operator. -/ +/-- The adjoint operator as a partially defined linear operator, denoted as `T†`. -/ def adjoint : F →ₗ.[𝕜] E where domain := T.adjointDomain toFun := if hT : Dense (T.domain : Set E) then adjointAux hT else 0 +@[inherit_doc] scoped postfix:1024 "†" => LinearPMap.adjoint theorem mem_adjoint_domain_iff (y : F) : y ∈ T†.domain ↔ Continuous ((innerₛₗ 𝕜 y).comp T.toFun) := diff --git a/Mathlib/Analysis/InnerProductSpace/NormPow.lean b/Mathlib/Analysis/InnerProductSpace/NormPow.lean index 1d53811b6834c..78320b9555e7f 100644 --- a/Mathlib/Analysis/InnerProductSpace/NormPow.lean +++ b/Mathlib/Analysis/InnerProductSpace/NormPow.lean @@ -98,6 +98,12 @@ theorem nnnorm_fderiv_norm_rpow_le {f : F → E} (hf : Differentiable ℝ f) ‖fderiv ℝ (fun x ↦ ‖f x‖ ^ (p : ℝ)) x‖₊ ≤ p * ‖f x‖₊ ^ ((p : ℝ) - 1) * ‖fderiv ℝ f x‖₊ := norm_fderiv_norm_rpow_le hf hp +lemma enorm_fderiv_norm_rpow_le {f : F → E} (hf : Differentiable ℝ f) + {x : F} {p : ℝ≥0} (hp : 1 < p) : + ‖fderiv ℝ (fun x ↦ ‖f x‖ ^ (p : ℝ)) x‖ₑ ≤ p * ‖f x‖ₑ ^ ((p : ℝ) - 1) * ‖fderiv ℝ f x‖ₑ := by + simpa [enorm, ← ENNReal.coe_rpow_of_nonneg _ (sub_nonneg.2 <| NNReal.one_le_coe.2 hp.le), + ← ENNReal.coe_mul] using nnnorm_fderiv_norm_rpow_le hf hp + theorem contDiff_norm_rpow {p : ℝ} (hp : 1 < p) : ContDiff ℝ 1 (fun x : E ↦ ‖x‖ ^ p) := by rw [contDiff_one_iff_fderiv] refine ⟨fun x ↦ hasFDerivAt_norm_rpow x hp |>.differentiableAt, ?_⟩ diff --git a/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean b/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean index 876666070250b..d0bbc998a808f 100644 --- a/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean +++ b/Mathlib/Analysis/InnerProductSpace/Orthogonal.lean @@ -33,7 +33,7 @@ namespace Submodule variable (K : Submodule 𝕜 E) -/-- The subspace of vectors orthogonal to a given subspace. -/ +/-- The subspace of vectors orthogonal to a given subspace, denoted `Kᗮ`. -/ def orthogonal : Submodule 𝕜 E where carrier := { v | ∀ u ∈ K, ⟪u, v⟫ = 0 } zero_mem' _ _ := inner_zero_right _ @@ -199,7 +199,7 @@ theorem bilinFormOfRealInner_orthogonal {E} [NormedAddCommGroup E] [InnerProduct /-! ### Orthogonality of submodules -In this section we define `Submodule.IsOrtho U V`, with notation `U ⟂ V`. +In this section we define `Submodule.IsOrtho U V`, denoted as `U ⟂ V`. The API roughly matches that of `Disjoint`. -/ @@ -207,7 +207,7 @@ The API roughly matches that of `Disjoint`. namespace Submodule -/-- The proposition that two submodules are orthogonal. Has notation `U ⟂ V`. -/ +/-- The proposition that two submodules are orthogonal, denoted as `U ⟂ V`. -/ def IsOrtho (U V : Submodule 𝕜 E) : Prop := U ≤ Vᗮ diff --git a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean index bf2b65c776bf0..f1d9d6b5a0bcf 100644 --- a/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean +++ b/Mathlib/Analysis/LocallyConvex/BalancedCoreHull.lean @@ -196,7 +196,7 @@ theorem subset_balancedCore (ht : (0 : E) ∈ t) (hst : ∀ a : 𝕜, ‖a‖ s ⊆ balancedCore 𝕜 t := by rw [balancedCore_eq_iInter ht] refine subset_iInter₂ fun a ha ↦ ?_ - rw [subset_set_smul_iff₀ (norm_pos_iff.mp <| zero_lt_one.trans_le ha)] + rw [subset_smul_set_iff₀ (norm_pos_iff.mp <| zero_lt_one.trans_le ha)] apply hst rw [norm_inv] exact inv_le_one_of_one_le₀ ha diff --git a/Mathlib/Analysis/LocallyConvex/Polar.lean b/Mathlib/Analysis/LocallyConvex/Polar.lean index af5dc28df843c..d49aedb9e4449 100644 --- a/Mathlib/Analysis/LocallyConvex/Polar.lean +++ b/Mathlib/Analysis/LocallyConvex/Polar.lean @@ -139,7 +139,7 @@ variable (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) theorem polar_univ (h : SeparatingRight B) : B.polar Set.univ = {(0 : F)} := by rw [Set.eq_singleton_iff_unique_mem] refine ⟨by simp only [zero_mem_polar], fun y hy => h _ fun x => ?_⟩ - refine norm_le_zero_iff.mp (le_of_forall_le_of_dense fun ε hε => ?_) + refine norm_le_zero_iff.mp (le_of_forall_gt_imp_ge_of_dense fun ε hε => ?_) rcases NormedField.exists_norm_lt 𝕜 hε with ⟨c, hc, hcε⟩ calc ‖B x y‖ = ‖c‖ * ‖B (c⁻¹ • x) y‖ := by diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index b4f0354b00717..d7caa79c94031 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -46,7 +46,7 @@ variable (𝕜 : Type*) {V V₁ V₁' V₂ V₃ V₄ : Type*} {P₁ P₁' : Type [SeminormedAddCommGroup V₄] [NormedSpace 𝕜 V₄] [PseudoMetricSpace P₄] [NormedAddTorsor V₄ P₄] /-- A `𝕜`-affine isometric embedding of one normed add-torsor over a normed `𝕜`-space into -another. -/ +another, denoted as `f : P →ᵃⁱ[𝕜] P₂`. -/ structure AffineIsometry extends P →ᵃ[𝕜] P₂ where norm_map : ∀ x : V, ‖linear x‖ = ‖x‖ @@ -260,7 +260,8 @@ end AffineSubspace variable (𝕜 P P₂) -/-- An affine isometric equivalence between two normed vector spaces. -/ +/-- An affine isometric equivalence between two normed vector spaces, +denoted `f : P ≃ᵃⁱ[𝕜] P₂`. -/ structure AffineIsometryEquiv extends P ≃ᵃ[𝕜] P₂ where norm_map : ∀ x, ‖linear x‖ = ‖x‖ diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index 9166797fa69d8..3d2e1d2cbdb3b 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.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, Eric Wieser -/ +import Mathlib.Algebra.Ring.Action.ConjAct import Mathlib.Analysis.Analytic.ChangeOrigin import Mathlib.Analysis.Complex.Basic import Mathlib.Data.Nat.Choose.Cast diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 0bc9524cf8635..97a91622891e1 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -31,7 +31,7 @@ assert_not_exists AddChar comap_norm_atTop DilationEquiv Finset.sup_mul_le_mul_s IsOfFinOrder Isometry.norm_map_of_map_one NNReal.isOpen_Ico_zero Rat.norm_cast_real RestrictScalars -variable {α : Type*} {β : Type*} {ι : Type*} +variable {G α β ι : Type*} open Filter open scoped Topology NNReal @@ -167,14 +167,17 @@ export NormOneClass (norm_one) attribute [simp] norm_one -@[simp] -theorem nnnorm_one [SeminormedAddCommGroup α] [One α] [NormOneClass α] : ‖(1 : α)‖₊ = 1 := - NNReal.eq norm_one +section SeminormedAddCommGroup +variable [SeminormedAddCommGroup G] [One G] [NormOneClass G] + +@[simp] lemma nnnorm_one : ‖(1 : G)‖₊ = 1 := NNReal.eq norm_one +@[simp] lemma enorm_one : ‖(1 : G)‖ₑ = 1 := by simp [enorm] -theorem NormOneClass.nontrivial (α : Type*) [SeminormedAddCommGroup α] [One α] [NormOneClass α] : - Nontrivial α := +theorem NormOneClass.nontrivial : Nontrivial G := nontrivial_of_ne 0 1 <| ne_of_apply_ne norm <| by simp +end SeminormedAddCommGroup + -- see Note [lower instance priority] instance (priority := 100) NonUnitalSeminormedCommRing.toNonUnitalCommRing [β : NonUnitalSeminormedCommRing α] : NonUnitalCommRing α := @@ -646,9 +649,8 @@ instance isAbsoluteValue_norm : IsAbsoluteValue (norm : α → ℝ) where abv_add' := norm_add_le abv_mul' := norm_mul -@[simp] -theorem nnnorm_mul (a b : α) : ‖a * b‖₊ = ‖a‖₊ * ‖b‖₊ := - NNReal.eq <| norm_mul a b +@[simp] lemma nnnorm_mul (a b : α) : ‖a * b‖₊ = ‖a‖₊ * ‖b‖₊ := NNReal.eq <| norm_mul a b +@[simp] lemma enorm_mul (a b : α) : ‖a * b‖ₑ = ‖a‖ₑ * ‖b‖ₑ := by simp [enorm] /-- `norm` as a `MonoidWithZeroHom`. -/ @[simps] @@ -674,6 +676,8 @@ theorem norm_pow (a : α) : ∀ n : ℕ, ‖a ^ n‖ = ‖a‖ ^ n := theorem nnnorm_pow (a : α) (n : ℕ) : ‖a ^ n‖₊ = ‖a‖₊ ^ n := (nnnormHom.toMonoidHom : α →* ℝ≥0).map_pow a n +@[simp] lemma enorm_pow (a : α) (n : ℕ) : ‖a ^ n‖ₑ = ‖a‖ₑ ^ n := by simp [enorm] + protected theorem List.norm_prod (l : List α) : ‖l.prod‖ = (l.map norm).prod := map_list_prod (normHom.toMonoidHom : α →* ℝ) _ @@ -696,6 +700,9 @@ theorem norm_inv (a : α) : ‖a⁻¹‖ = ‖a‖⁻¹ := theorem nnnorm_inv (a : α) : ‖a⁻¹‖₊ = ‖a‖₊⁻¹ := NNReal.eq <| by simp +@[simp] +lemma enorm_inv {a : α} (ha : a ≠ 0) : ‖a⁻¹‖ₑ = ‖a‖ₑ⁻¹ := by simp [enorm, ENNReal.coe_inv, ha] + @[simp] theorem norm_zpow : ∀ (a : α) (n : ℤ), ‖a ^ n‖ = ‖a‖ ^ n := map_zpow₀ (normHom : α →*₀ ℝ) @@ -923,20 +930,11 @@ open NNReal theorem norm_eq (x : ℝ≥0) : ‖(x : ℝ)‖ = x := by rw [Real.norm_eq_abs, x.abs_eq] -@[simp] -theorem nnnorm_eq (x : ℝ≥0) : ‖(x : ℝ)‖₊ = x := - NNReal.eq <| Real.norm_of_nonneg x.2 +@[simp] lemma nnnorm_eq (x : ℝ≥0) : ‖(x : ℝ)‖₊ = x := by ext; simp [nnnorm] +@[simp] lemma enorm_eq (x : ℝ≥0) : ‖(x : ℝ)‖ₑ = x := by simp [enorm] end NNReal -@[simp 1001] -- Porting note: increase priority so that the LHS doesn't simplify -theorem norm_norm [SeminormedAddCommGroup α] (x : α) : ‖‖x‖‖ = ‖x‖ := - Real.norm_of_nonneg (norm_nonneg _) - -@[simp] -theorem nnnorm_norm [SeminormedAddCommGroup α] (a : α) : ‖‖a‖‖₊ = ‖a‖₊ := by - rw [Real.nnnorm_of_nonneg (norm_nonneg a)]; rfl - /-- A restatement of `MetricSpace.tendsto_atTop` in terms of the norm. -/ theorem NormedAddCommGroup.tendsto_atTop [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] {β : Type*} [SeminormedAddCommGroup β] {f : α → β} {b : β} : diff --git a/Mathlib/Analysis/Normed/Field/Ultra.lean b/Mathlib/Analysis/Normed/Field/Ultra.lean index 8ef392b9b1afd..79e9e92787437 100644 --- a/Mathlib/Analysis/Normed/Field/Ultra.lean +++ b/Mathlib/Analysis/Normed/Field/Ultra.lean @@ -76,7 +76,7 @@ lemma isUltrametricDist_of_forall_pow_norm_le_nsmul_pow_max_one_norm -- that avoids explicitly mentioning `m`-th roots. -- First note it suffices to show that `‖x + 1‖ ≤ a` for all `a : ℝ` with `max ‖x‖ 1 < a`. rw [max_comm] - refine le_of_forall_le_of_dense fun a ha ↦ ?_ + refine le_of_forall_gt_imp_ge_of_dense fun a ha ↦ ?_ have ha' : 1 < a := (max_lt_iff.mp ha).left -- `max 1 ‖x‖ < a`, so there must be some `m : ℕ` such that `m + 1 < (a / max 1 ‖x‖) ^ m` -- by the virtue of exponential growth being faster than linear growth diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 79417c8b0f215..0b033f8d12618 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Subgroup.Ker import Mathlib.Analysis.Normed.Group.Seminorm +import Mathlib.Topology.Instances.ENNReal.Defs import Mathlib.Topology.Metrizable.Uniformity import Mathlib.Topology.Sequences @@ -77,6 +78,23 @@ export ENorm (enorm) @[inherit_doc] notation "‖" e "‖₊" => nnnorm e @[inherit_doc] notation "‖" e "‖ₑ" => enorm e +section ENorm +variable {E : Type*} [NNNorm E] {x : E} {r : ℝ≥0} + +instance NNNorm.toENorm : ENorm E where enorm := (‖·‖₊ : E → ℝ≥0∞) + +lemma enorm_eq_nnnorm (x : E) : ‖x‖ₑ = ‖x‖₊ := rfl + +@[simp, norm_cast] lemma coe_le_enorm : r ≤ ‖x‖ₑ ↔ r ≤ ‖x‖₊ := by simp [enorm] +@[simp, norm_cast] lemma enorm_le_coe : ‖x‖ₑ ≤ r ↔ ‖x‖₊ ≤ r := by simp [enorm] +@[simp, norm_cast] lemma coe_lt_enorm : r < ‖x‖ₑ ↔ r < ‖x‖₊ := by simp [enorm] +@[simp, norm_cast] lemma enorm_lt_coe : ‖x‖ₑ < r ↔ ‖x‖₊ < r := by simp [enorm] + +@[simp] lemma enorm_ne_top : ‖x‖ₑ ≠ ∞ := by simp [enorm] +@[simp] lemma enorm_lt_top : ‖x‖ₑ < ∞ := by simp [enorm] + +end ENorm + /-- A seminormed group is an additive group endowed with a norm for which `dist x y = ‖x - y‖` defines a pseudometric space structure. -/ class SeminormedAddGroup (E : Type*) extends Norm E, AddGroup E, PseudoMetricSpace E where @@ -651,6 +669,13 @@ theorem coe_comp_nnnorm' : (toReal : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥ theorem norm_toNNReal' : ‖a‖.toNNReal = ‖a‖₊ := @Real.toNNReal_coe ‖a‖₊ +@[to_additive toReal_enorm] +lemma toReal_enorm' (x : E) : ‖x‖ₑ.toReal = ‖x‖ := by simp [enorm] + +@[to_additive ofReal_norm] +lemma ofReal_norm' (x : E) : .ofReal ‖x‖ = ‖x‖ₑ := by + simp [enorm, ENNReal.ofReal, Real.toNNReal, nnnorm] + @[to_additive] theorem nndist_eq_nnnorm_div (a b : E) : nndist a b = ‖a / b‖₊ := NNReal.eq <| dist_eq_norm_div _ _ @@ -677,6 +702,10 @@ theorem ne_one_of_nnnorm_ne_zero {a : E} : ‖a‖₊ ≠ 0 → a ≠ 1 := theorem nnnorm_mul_le' (a b : E) : ‖a * b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := NNReal.coe_le_coe.1 <| norm_mul_le' a b +@[to_additive enorm_add_le] +lemma enorm_mul_le' (a b : E) : ‖a * b‖ₑ ≤ ‖a‖ₑ + ‖b‖ₑ := by + simpa [enorm, ← ENNReal.coe_add] using nnnorm_mul_le' a b + @[to_additive norm_nsmul_le] lemma norm_pow_le_mul_norm : ∀ {n : ℕ}, ‖a ^ n‖ ≤ n * ‖a‖ | 0 => by simp @@ -690,6 +719,9 @@ lemma nnnorm_pow_le_mul_norm {n : ℕ} : ‖a ^ n‖₊ ≤ n * ‖a‖₊ := by theorem nnnorm_inv' (a : E) : ‖a⁻¹‖₊ = ‖a‖₊ := NNReal.eq <| norm_inv' a +@[to_additive (attr := simp) enorm_neg] +lemma enorm_inv' (a : E) : ‖a⁻¹‖ₑ = ‖a‖ₑ := by simp [enorm] + @[to_additive (attr := simp) nnnorm_abs_zsmul] theorem nnnorm_zpow_abs (a : E) (n : ℤ) : ‖a ^ |n|‖₊ = ‖a ^ n‖₊ := NNReal.eq <| norm_zpow_abs a n @@ -723,6 +755,10 @@ theorem nndist_mulIndicator (s t : Set α) (f : α → E) (x : α) : theorem nnnorm_div_le (a b : E) : ‖a / b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := NNReal.coe_le_coe.1 <| norm_div_le _ _ +@[to_additive] +lemma enorm_div_le : ‖a / b‖ₑ ≤ ‖a‖ₑ + ‖b‖ₑ := by + simpa [enorm, ← ENNReal.coe_add] using nnnorm_div_le a b + @[to_additive nndist_nnnorm_nnnorm_le] theorem nndist_nnnorm_nnnorm_le' (a b : E) : nndist ‖a‖₊ ‖b‖₊ ≤ ‖a / b‖₊ := NNReal.coe_le_coe.1 <| dist_norm_norm_le' a b @@ -793,19 +829,16 @@ theorem mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ EMetric.ball (1 : E) r end NNNorm -section ENNNorm - -instance {E : Type*} [NNNorm E] : ENorm E where - enorm := (‖·‖₊ : E → ℝ≥0∞) +section ENorm -lemma enorm_eq_nnnorm {E : Type*} [NNNorm E] {x : E} : ‖x‖ₑ = ‖x‖₊ := rfl +@[to_additive (attr := simp) enorm_zero] lemma enorm_one' : ‖(1 : E)‖ₑ = 0 := by simp [enorm] instance : ENorm ℝ≥0∞ where enorm x := x @[simp] lemma enorm_eq_self (x : ℝ≥0∞) : ‖x‖ₑ = x := rfl -end ENNNorm +end ENorm @[to_additive] theorem tendsto_iff_norm_div_tendsto_zero {f : α → E} {a : Filter α} {b : E} : @@ -870,6 +903,10 @@ theorem continuous_norm' : Continuous fun a : E => ‖a‖ := by theorem continuous_nnnorm' : Continuous fun a : E => ‖a‖₊ := continuous_norm'.subtype_mk _ +@[to_additive (attr := continuity) continuous_enorm] +lemma continuous_enorm' : Continuous fun a : E ↦ ‖a‖ₑ := + ENNReal.isOpenEmbedding_coe.continuous.comp continuous_nnnorm' + set_option linter.docPrime false in @[to_additive Inseparable.norm_eq_norm] theorem Inseparable.norm_eq_norm' {u v : E} (h : Inseparable u v) : ‖u‖ = ‖v‖ := @@ -900,11 +937,15 @@ theorem Filter.Tendsto.norm' (h : Tendsto f l (𝓝 a)) : Tendsto (fun x => ‖f theorem Filter.Tendsto.nnnorm' (h : Tendsto f l (𝓝 a)) : Tendsto (fun x => ‖f x‖₊) l (𝓝 ‖a‖₊) := Tendsto.comp continuous_nnnorm'.continuousAt h +@[to_additive Filter.Tendsto.enorm] +lemma Filter.Tendsto.enorm' (h : Tendsto f l (𝓝 a)) : Tendsto (‖f ·‖ₑ) l (𝓝 ‖a‖ₑ) := + .comp continuous_enorm'.continuousAt h + end section -variable [TopologicalSpace α] {f : α → E} +variable [TopologicalSpace α] {f : α → E} {s : Set α} {a : α} @[to_additive (attr := fun_prop) Continuous.norm] theorem Continuous.norm' : Continuous f → Continuous fun x => ‖f x‖ := @@ -914,6 +955,9 @@ theorem Continuous.norm' : Continuous f → Continuous fun x => ‖f x‖ := theorem Continuous.nnnorm' : Continuous f → Continuous fun x => ‖f x‖₊ := continuous_nnnorm'.comp +@[to_additive (attr := fun_prop) Continuous.enorm] +lemma Continuous.enorm' : Continuous f → Continuous (‖f ·‖ₑ) := continuous_enorm'.comp + @[to_additive (attr := fun_prop) ContinuousAt.norm] theorem ContinuousAt.norm' {a : α} (h : ContinuousAt f a) : ContinuousAt (fun x => ‖f x‖) a := Tendsto.norm' h @@ -922,6 +966,9 @@ theorem ContinuousAt.norm' {a : α} (h : ContinuousAt f a) : ContinuousAt (fun x theorem ContinuousAt.nnnorm' {a : α} (h : ContinuousAt f a) : ContinuousAt (fun x => ‖f x‖₊) a := Tendsto.nnnorm' h +@[to_additive (attr := fun_prop) ContinuousAt.enorm] +lemma ContinuousAt.enorm' (h : ContinuousAt f a) : ContinuousAt (‖f ·‖ₑ) a := Tendsto.enorm' h + @[to_additive ContinuousWithinAt.norm] theorem ContinuousWithinAt.norm' {s : Set α} {a : α} (h : ContinuousWithinAt f s a) : ContinuousWithinAt (fun x => ‖f x‖) s a := @@ -932,6 +979,10 @@ theorem ContinuousWithinAt.nnnorm' {s : Set α} {a : α} (h : ContinuousWithinAt ContinuousWithinAt (fun x => ‖f x‖₊) s a := Tendsto.nnnorm' h +@[to_additive ContinuousWithinAt.enorm] +lemma ContinuousWithinAt.enorm' (h : ContinuousWithinAt f s a) : ContinuousWithinAt (‖f ·‖ₑ) s a := + Tendsto.enorm' h + @[to_additive (attr := fun_prop) ContinuousOn.norm] theorem ContinuousOn.norm' {s : Set α} (h : ContinuousOn f s) : ContinuousOn (fun x => ‖f x‖) s := fun x hx => (h x hx).norm' @@ -940,6 +991,10 @@ theorem ContinuousOn.norm' {s : Set α} (h : ContinuousOn f s) : ContinuousOn (f theorem ContinuousOn.nnnorm' {s : Set α} (h : ContinuousOn f s) : ContinuousOn (fun x => ‖f x‖₊) s := fun x hx => (h x hx).nnnorm' +@[to_additive (attr := fun_prop) ContinuousOn.enorm] +lemma ContinuousOn.enorm' (h : ContinuousOn f s) : ContinuousOn (‖f ·‖ₑ) s := + fun x hx => (h x hx).enorm' + end /-- If `‖y‖ → ∞`, then we can assume `y ≠ x` for any fixed `x`. -/ @@ -1235,6 +1290,7 @@ theorem le_norm_self (r : ℝ) : r ≤ ‖r‖ := @[simp 1100] lemma norm_natCast (n : ℕ) : ‖(n : ℝ)‖ = n := abs_of_nonneg n.cast_nonneg @[simp 1100] lemma nnnorm_natCast (n : ℕ) : ‖(n : ℝ)‖₊ = n := NNReal.eq <| norm_natCast _ +@[simp 1100] lemma enorm_natCast (n : ℕ) : ‖(n : ℝ)‖ₑ = n := by simp [enorm] @[deprecated (since := "2024-04-05")] alias norm_coe_nat := norm_natCast @[deprecated (since := "2024-04-05")] alias nnnorm_coe_nat := nnnorm_natCast @@ -1257,8 +1313,11 @@ lemma nnnorm_nnratCast (q : ℚ≥0) : ‖(q : ℝ)‖₊ = q := by simp [nnnorm theorem nnnorm_of_nonneg (hr : 0 ≤ r) : ‖r‖₊ = ⟨r, hr⟩ := NNReal.eq <| norm_of_nonneg hr -@[simp] -theorem nnnorm_abs (r : ℝ) : ‖|r|‖₊ = ‖r‖₊ := by simp [nnnorm] +lemma enorm_of_nonneg (hr : 0 ≤ r) : ‖r‖ₑ = .ofReal r := by + simp [enorm, nnnorm_of_nonneg hr, ENNReal.ofReal, toNNReal, hr] + +@[simp] lemma nnnorm_abs (r : ℝ) : ‖|r|‖₊ = ‖r‖₊ := by simp [nnnorm] +@[simp] lemma enorm_abs (r : ℝ) : ‖|r|‖ₑ = ‖r‖ₑ := by simp [enorm] theorem ennnorm_eq_ofReal (hr : 0 ≤ r) : (‖r‖₊ : ℝ≥0∞) = ENNReal.ofReal r := by rw [← ofReal_norm_eq_coe_nnnorm, norm_of_nonneg hr] @@ -1290,6 +1349,16 @@ instance : NNNorm ℝ≥0 where end NNReal + -- Porting note: increase priority so that the LHS doesn't simplify +@[to_additive (attr := simp 1001) norm_norm] +lemma norm_norm' (x : E) : ‖‖x‖‖ = ‖x‖ := Real.norm_of_nonneg (norm_nonneg' _) + +@[to_additive (attr := simp) nnnorm_norm] +lemma nnnorm_norm' (x : E) : ‖‖x‖‖₊ = ‖x‖₊ := by simp [nnnorm] + +@[to_additive (attr := simp) enorm_norm] +lemma enorm_norm' (x : E) : ‖‖x‖‖ₑ = ‖x‖ₑ := by simp [enorm] + end SeminormedCommGroup section NormedGroup @@ -1342,13 +1411,22 @@ theorem eq_one_or_nnnorm_pos (a : E) : a = 1 ∨ 0 < ‖a‖₊ := theorem nnnorm_eq_zero' : ‖a‖₊ = 0 ↔ a = 1 := by rw [← NNReal.coe_eq_zero, coe_nnnorm', norm_eq_zero'] +@[to_additive (attr := simp) enorm_eq_zero] +lemma enorm_eq_zero' : ‖a‖ₑ = 0 ↔ a = 1 := by simp [enorm] + @[to_additive nnnorm_ne_zero_iff] theorem nnnorm_ne_zero_iff' : ‖a‖₊ ≠ 0 ↔ a ≠ 1 := nnnorm_eq_zero'.not +@[to_additive enorm_ne_zero] +lemma enorm_ne_zero' : ‖a‖ₑ ≠ 0 ↔ a ≠ 1 := enorm_eq_zero'.ne + @[to_additive (attr := simp) nnnorm_pos] lemma nnnorm_pos' : 0 < ‖a‖₊ ↔ a ≠ 1 := pos_iff_ne_zero.trans nnnorm_ne_zero_iff' +@[to_additive (attr := simp) enorm_pos] +lemma enorm_pos' : 0 < ‖a‖ₑ ↔ a ≠ 1 := pos_iff_ne_zero.trans enorm_ne_zero' + /-- See `tendsto_norm_one` for a version with full neighborhoods. -/ @[to_additive "See `tendsto_norm_zero` for a version with full neighborhoods."] lemma tendsto_norm_nhdsNE_one : Tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := diff --git a/Mathlib/Analysis/Normed/Group/Completion.lean b/Mathlib/Analysis/Normed/Group/Completion.lean index 1bdd796d8f4a9..20db58aa839b3 100644 --- a/Mathlib/Analysis/Normed/Group/Completion.lean +++ b/Mathlib/Analysis/Normed/Group/Completion.lean @@ -44,6 +44,10 @@ instance [SeminormedAddCommGroup E] : NormedAddCommGroup (Completion E) where theorem nnnorm_coe {E} [SeminormedAddCommGroup E] (x : E) : ‖(x : Completion E)‖₊ = ‖x‖₊ := by simp [nnnorm] +@[simp] +lemma enorm_coe {E} [SeminormedAddCommGroup E] (x : E) : ‖(x : Completion E)‖ₑ = ‖x‖ₑ := by + simp [enorm] + end Completion end UniformSpace diff --git a/Mathlib/Analysis/Normed/Group/Constructions.lean b/Mathlib/Analysis/Normed/Group/Constructions.lean index 5a2a7e261e62b..2fc6b8e143742 100644 --- a/Mathlib/Analysis/Normed/Group/Constructions.lean +++ b/Mathlib/Analysis/Normed/Group/Constructions.lean @@ -403,6 +403,9 @@ theorem Pi.nnnorm_single [DecidableEq ι] [∀ i, NormedAddCommGroup (π i)] {i simp simp [Pi.nnnorm_def, H, Pi.single_apply, Finset.sup_ite, Finset.filter_eq'] +lemma Pi.enorm_single [DecidableEq ι] [∀ i, NormedAddCommGroup (π i)] {i : ι} (y : π i) : + ‖Pi.single i y‖ₑ = ‖y‖ₑ := by simp [enorm, Pi.nnnorm_single] + theorem Pi.norm_single [DecidableEq ι] [∀ i, NormedAddCommGroup (π i)] {i : ι} (y : π i) : ‖Pi.single i y‖ = ‖y‖ := congr_arg Subtype.val <| Pi.nnnorm_single y diff --git a/Mathlib/Analysis/Normed/Group/Hom.lean b/Mathlib/Analysis/Normed/Group/Hom.lean index 9a2101245416a..2e71c2d63fefc 100644 --- a/Mathlib/Analysis/Normed/Group/Hom.lean +++ b/Mathlib/Analysis/Normed/Group/Hom.lean @@ -110,7 +110,6 @@ variable (f g) theorem toFun_eq_coe : f.toFun = f := rfl --- Porting note: removed `simp` because `simpNF` complains the LHS doesn't simplify. theorem coe_mk (f) (h₁) (h₂) (h₃) : ⇑(⟨f, h₁, h₂, h₃⟩ : NormedAddGroupHom V₁ V₂) = f := rfl diff --git a/Mathlib/Analysis/Normed/Group/Quotient.lean b/Mathlib/Analysis/Normed/Group/Quotient.lean index d931b9c66849f..0009c2a130e4e 100644 --- a/Mathlib/Analysis/Normed/Group/Quotient.lean +++ b/Mathlib/Analysis/Normed/Group/Quotient.lean @@ -7,7 +7,6 @@ import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Analysis.Normed.Group.Hom import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.Topology.MetricSpace.HausdorffDistance -import Mathlib.Topology.Instances.Real.Lemmas /-! # Quotients of seminormed groups diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean index e0c929f53e5a7..15de9d58c9cde 100644 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean @@ -51,15 +51,11 @@ instance completion_completeSpace {V : SemiNormedGrp} : CompleteSpace (completio Completion.completeSpace _ /-- The canonical morphism from a seminormed group `V` to its completion. -/ -@[simps] def completion.incl {V : SemiNormedGrp} : V ⟶ completion.obj V where toFun v := (v : Completion V) map_add' := Completion.coe_add bound' := ⟨1, fun v => by simp⟩ --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] SemiNormedGrp.completion.incl_apply - theorem completion.norm_incl_eq {V : SemiNormedGrp} {v : V} : ‖completion.incl v‖ = ‖v‖ := UniformSpace.Completion.norm_coe _ diff --git a/Mathlib/Analysis/Normed/Group/Seminorm.lean b/Mathlib/Analysis/Normed/Group/Seminorm.lean index e895bf01b8d3c..810ab65630fd8 100644 --- a/Mathlib/Analysis/Normed/Group/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Group/Seminorm.lean @@ -443,7 +443,6 @@ instance nonarchAddGroupSeminormClass : map_zero f := f.map_zero' map_neg_eq_map' f := f.neg' --- Porting note: `simpNF` said the left hand side simplified to this @[simp] theorem toZeroHom_eq_coe : ⇑p.toZeroHom = p := by rfl @@ -659,7 +658,6 @@ instance groupNormClass : GroupNormClass (GroupNorm E) E ℝ where map_inv_eq_map f := f.inv' eq_one_of_map_eq_zero f := f.eq_one_of_map_eq_zero' _ --- Porting note: `simpNF` told me the left-hand side simplified to this @[to_additive (attr := simp)] theorem toGroupSeminorm_eq_coe : ⇑p.toGroupSeminorm = p := rfl @@ -784,7 +782,6 @@ instance nonarchAddGroupNormClass : NonarchAddGroupNormClass (NonarchAddGroupNor map_neg_eq_map' f := f.neg' eq_zero_of_map_eq_zero f := f.eq_zero_of_map_eq_zero' _ --- Porting note: `simpNF` told me the left-hand side simplified to this @[simp] theorem toNonarchAddGroupSeminorm_eq_coe : ⇑p.toNonarchAddGroupSeminorm = p := rfl diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index b2a36ecac89b5..356ca2b7ac0dc 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -288,7 +288,8 @@ instance : AddCommGroup (PreLp E) := by unfold PreLp; infer_instance instance PreLp.unique [IsEmpty α] : Unique (PreLp E) := Pi.uniqueOfIsEmpty E -/-- lp space -/ +/-- lp space +The `p=∞` case has notation `ℓ^∞(ι, E)` resp. `ℓ^∞(ι)` (for `E = ℝ`) in the `lp` namespace. -/ def lp (E : α → Type*) [∀ i, NormedAddCommGroup (E i)] (p : ℝ≥0∞) : AddSubgroup (PreLp E) where carrier := { f | Memℓp f p } zero_mem' := zero_memℓp diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index 3ca89f37a1d43..1ef16a5d78305 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -651,3 +651,34 @@ abbrev NormedSpace.ofCore {𝕜 : Type*} {E : Type*} [NormedField 𝕜] [Seminor norm_smul_le r x := by rw [core.norm_smul r x] end Core + +variable {G H : Type*} [SeminormedAddCommGroup G] [SeminormedAddCommGroup H] [NormedSpace ℝ H] + {s : Set G} + +/-- A group homomorphism from a normed group to a real normed space, +bounded on a neighborhood of `0`, must be continuous. -/ +lemma AddMonoidHom.continuous_of_isBounded_nhds_zero (f : G →+ H) (hs : s ∈ 𝓝 (0 : G)) + (hbounded : IsBounded (f '' s)) : Continuous f := by + obtain ⟨δ, hδ, hUε⟩ := Metric.mem_nhds_iff.mp hs + obtain ⟨C, hC⟩ := (isBounded_iff_subset_ball 0).1 (hbounded.subset <| image_subset f hUε) + refine continuous_of_continuousAt_zero _ (continuousAt_iff.2 fun ε (hε : _ < _) => ?_) + simp only [dist_zero_right, _root_.map_zero, exists_prop] + simp only [subset_def, mem_image, mem_ball, dist_zero_right, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂] at hC + have hC₀ : 0 < C := (norm_nonneg _).trans_lt <| hC 0 (by simpa) + obtain ⟨n, hn⟩ := exists_nat_gt (C / ε) + have hnpos : 0 < (n : ℝ) := (div_pos hC₀ hε).trans hn + have hn₀ : n ≠ 0 := by rintro rfl; simp at hnpos + refine ⟨δ / n, div_pos hδ hnpos, fun {x} hxδ => ?_⟩ + calc + ‖f x‖ + _ = ‖(n : ℝ)⁻¹ • f (n • x)‖ := by simp [← Nat.cast_smul_eq_nsmul ℝ, hn₀] + _ ≤ ‖(n : ℝ)⁻¹‖ * ‖f (n • x)‖ := norm_smul_le .. + _ < ‖(n : ℝ)⁻¹‖ * C := by + gcongr + · simpa [pos_iff_ne_zero] + · refine hC _ <| norm_nsmul_le.trans_lt ?_ + simpa only [norm_mul, Real.norm_natCast, lt_div_iff₀ hnpos, mul_comm] using hxδ + _ = (n : ℝ)⁻¹ * C := by simp + _ < (C / ε : ℝ)⁻¹ * C := by gcongr + _ = ε := by simp [hC₀.ne'] diff --git a/Mathlib/Analysis/Normed/Module/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean index 8a947e649a0b2..279e098d12e2c 100644 --- a/Mathlib/Analysis/Normed/Module/Dual.lean +++ b/Mathlib/Analysis/Normed/Module/Dual.lean @@ -240,7 +240,7 @@ theorem polar_ball {𝕜 E : Type*} [RCLike 𝕜] [NormedAddCommGroup E] [Normed apply le_antisymm · intro x hx rw [mem_closedBall_zero_iff] - apply le_of_forall_le_of_dense + apply le_of_forall_gt_imp_ge_of_dense intro a ha rw [← mem_closedBall_zero_iff, ← (mul_div_cancel_left₀ a (Ne.symm (ne_of_lt hr)))] rw [← RCLike.norm_of_nonneg (K := 𝕜) (le_trans zero_le_one diff --git a/Mathlib/Analysis/Normed/MulAction.lean b/Mathlib/Analysis/Normed/MulAction.lean index 2c9c706fdb6c1..48d51cde9e44f 100644 --- a/Mathlib/Analysis/Normed/MulAction.lean +++ b/Mathlib/Analysis/Normed/MulAction.lean @@ -23,7 +23,7 @@ variable {α β : Type*} section SeminormedAddGroup variable [SeminormedAddGroup α] [SeminormedAddGroup β] [SMulZeroClass α β] -variable [BoundedSMul α β] +variable [BoundedSMul α β] {r : α} {x : β} @[bound] theorem norm_smul_le (r : α) (x : β) : ‖r • x‖ ≤ ‖r‖ * ‖x‖ := by @@ -33,6 +33,10 @@ theorem norm_smul_le (r : α) (x : β) : ‖r • x‖ ≤ ‖r‖ * ‖x‖ := theorem nnnorm_smul_le (r : α) (x : β) : ‖r • x‖₊ ≤ ‖r‖₊ * ‖x‖₊ := norm_smul_le _ _ +@[bound] +lemma enorm_smul_le : ‖r • x‖ₑ ≤ ‖r‖ₑ * ‖x‖ₑ := by + simpa [enorm, ← ENNReal.coe_mul] using nnnorm_smul_le .. + theorem dist_smul_le (s : α) (x y : β) : dist (s • x) (s • y) ≤ ‖s‖ * dist x y := by simpa only [dist_eq_norm, sub_zero] using dist_smul_pair s x y @@ -91,6 +95,8 @@ theorem norm_smul (r : α) (x : β) : ‖r • x‖ = ‖r‖ * ‖x‖ := by theorem nnnorm_smul (r : α) (x : β) : ‖r • x‖₊ = ‖r‖₊ * ‖x‖₊ := NNReal.eq <| norm_smul r x +lemma enorm_smul (r : α) (x : β) : ‖r • x‖ₑ = ‖r‖ₑ * ‖x‖ₑ := by simp [enorm, nnnorm_smul] + end NormedDivisionRing section NormedDivisionRingModule diff --git a/Mathlib/Analysis/Normed/Operator/Compact.lean b/Mathlib/Analysis/Normed/Operator/Compact.lean index da69e15c03410..972e8b93105c1 100644 --- a/Mathlib/Analysis/Normed/Operator/Compact.lean +++ b/Mathlib/Analysis/Normed/Operator/Compact.lean @@ -326,7 +326,7 @@ theorem IsCompactOperator.continuous {f : M₁ →ₛₗ[σ₁₂] M₂} (hf : I have : IsUnit c⁻¹ := hcnz.isUnit.inv rwa [mem_map, preimage_smul_setₛₗ _ _ _ f this, set_smul_mem_nhds_zero_iff (inv_ne_zero hcnz)] -- Since `σ₁₂ c⁻¹` = `(σ₁₂ c)⁻¹`, we have to prove that `K ⊆ σ₁₂ c • U`. - rw [map_inv₀, ← subset_set_smul_iff₀ ((map_ne_zero σ₁₂).mpr hcnz)] + rw [map_inv₀, ← subset_smul_set_iff₀ ((map_ne_zero σ₁₂).mpr hcnz)] -- But `σ₁₂` is isometric, so `‖σ₁₂ c‖ = ‖c‖ > r`, which concludes the argument since -- `∀ a : 𝕜₂, r ≤ ‖a‖ → K ⊆ a • U`. refine hrU (σ₁₂ c) ?_ diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 03481455bb86d..a59154a0c16c3 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -42,7 +42,8 @@ variable {R R₂ R₃ R₄ E E₂ E₃ E₄ F 𝓕 : Type*} [Semiring R] [Semiri [SeminormedAddCommGroup E₄] [Module R E] [Module R₂ E₂] [Module R₃ E₃] [Module R₄ E₄] [NormedAddCommGroup F] [Module R F] -/-- A `σ₁₂`-semilinear isometric embedding of a normed `R`-module into an `R₂`-module. -/ +/-- A `σ₁₂`-semilinear isometric embedding of a normed `R`-module into an `R₂`-module, +denoted as `f : E →ₛₗᵢ[σ₁₂] E₂`. -/ structure LinearIsometry (σ₁₂ : R →+* R₂) (E E₂ : Type*) [SeminormedAddCommGroup E] [SeminormedAddCommGroup E₂] [Module R E] [Module R₂ E₂] extends E →ₛₗ[σ₁₂] E₂ where norm_map' : ∀ x, ‖toLinearMap x‖ = ‖x‖ @@ -402,7 +403,8 @@ theorem subtypeₗᵢ_toContinuousLinearMap : p.subtypeₗᵢ.toContinuousLinear end Submodule -/-- A semilinear isometric equivalence between two normed vector spaces. -/ +/-- A semilinear isometric equivalence between two normed vector spaces, +denoted as `f : E ≃ₛₗᵢ[σ₁₂] E₂`. -/ structure LinearIsometryEquiv (σ₁₂ : R →+* R₂) {σ₂₁ : R₂ →+* R} [RingHomInvPair σ₁₂ σ₂₁] [RingHomInvPair σ₂₁ σ₁₂] (E E₂ : Type*) [SeminormedAddCommGroup E] [SeminormedAddCommGroup E₂] [Module R E] [Module R₂ E₂] extends E ≃ₛₗ[σ₁₂] E₂ where diff --git a/Mathlib/Analysis/NormedSpace/FunctionSeries.lean b/Mathlib/Analysis/NormedSpace/FunctionSeries.lean index 99a048f2d05da..c53248e341e13 100644 --- a/Mathlib/Analysis/NormedSpace/FunctionSeries.lean +++ b/Mathlib/Analysis/NormedSpace/FunctionSeries.lean @@ -72,6 +72,14 @@ theorem tendstoUniformlyOn_tsum_of_cofinite_eventually {ι : Type*} {f : ι → have : ¬ i ∈ hN.toFinset := fun hg ↦ hi (Finset.union_subset_left hn hg) aesop +theorem tendstoUniformlyOn_tsum_nat_eventually {α F : Type*} [NormedAddCommGroup F] + [CompleteSpace F] {f : ℕ → α → F} {u : ℕ → ℝ} (hu : Summable u) {s : Set α} + (hfu : ∀ᶠ n in atTop, ∀ x ∈ s, ‖f n x‖ ≤ u n) : + TendstoUniformlyOn (fun N x => ∑ n ∈ Finset.range N, f n x) + (fun x => ∑' n, f n x) atTop s := + fun v hv ↦ tendsto_finset_range.eventually <| + tendstoUniformlyOn_tsum_of_cofinite_eventually hu (Nat.cofinite_eq_atTop ▸ hfu) v hv + /-- An infinite sum of functions with summable sup norm is the uniform limit of its partial sums. Version with general index set. -/ theorem tendstoUniformly_tsum {f : α → β → F} (hu : Summable u) (hfu : ∀ n x, ‖f n x‖ ≤ u n) : diff --git a/Mathlib/Analysis/NormedSpace/IndicatorFunction.lean b/Mathlib/Analysis/NormedSpace/IndicatorFunction.lean index 2b4d9f5b4bd2d..22a5c81e03806 100644 --- a/Mathlib/Analysis/NormedSpace/IndicatorFunction.lean +++ b/Mathlib/Analysis/NormedSpace/IndicatorFunction.lean @@ -27,6 +27,10 @@ theorem nnnorm_indicator_eq_indicator_nnnorm : ‖indicator s f a‖₊ = indicator s (fun a => ‖f a‖₊) a := flip congr_fun a (indicator_comp_of_zero nnnorm_zero).symm +lemma enorm_indicator_eq_indicator_enorm : + ‖indicator s f a‖ₑ = indicator s (fun a => ‖f a‖ₑ) a := + flip congr_fun a (indicator_comp_of_zero enorm_zero).symm + theorem norm_indicator_le_of_subset (h : s ⊆ t) (f : α → E) (a : α) : ‖indicator s f a‖ ≤ ‖indicator t f a‖ := by simp only [norm_indicator_eq_indicator_norm] diff --git a/Mathlib/Analysis/NormedSpace/Int.lean b/Mathlib/Analysis/NormedSpace/Int.lean index b08f55039887a..d8a954c1320a1 100644 --- a/Mathlib/Analysis/NormedSpace/Int.lean +++ b/Mathlib/Analysis/NormedSpace/Int.lean @@ -30,6 +30,8 @@ theorem norm_coe_units (e : ℤˣ) : ‖(e : ℤ)‖ = 1 := by theorem nnnorm_natCast (n : ℕ) : ‖(n : ℤ)‖₊ = n := Real.nnnorm_natCast _ +@[simp] lemma enorm_natCast (n : ℕ) : ‖(n : ℤ)‖ₑ = n := Real.enorm_natCast _ + @[deprecated (since := "2024-04-05")] alias nnnorm_coe_nat := nnnorm_natCast @[simp] diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean index 018b01d78582b..58490a8265750 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NNNorm.lean @@ -86,14 +86,16 @@ theorem opNNNorm_le_iff {f : E →SL[σ₁₂] F} {C : ℝ≥0} : ‖f‖₊ ≤ theorem isLeast_opNNNorm : IsLeast {C : ℝ≥0 | ∀ x, ‖f x‖₊ ≤ C * ‖x‖₊} ‖f‖₊ := by simpa only [← opNNNorm_le_iff] using isLeast_Ici - theorem opNNNorm_comp_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₂] F) : ‖h.comp f‖₊ ≤ ‖h‖₊ * ‖f‖₊ := opNorm_comp_le h f +lemma opENorm_comp_le [RingHomIsometric σ₁₃] (f : E →SL[σ₁₂] F) : ‖h.comp f‖ₑ ≤ ‖h‖ₑ * ‖f‖ₑ := by + simpa [enorm, ← ENNReal.coe_mul] using opNNNorm_comp_le h f theorem le_opNNNorm : ‖f x‖₊ ≤ ‖f‖₊ * ‖x‖₊ := f.le_opNorm x +lemma le_opENorm : ‖f x‖ₑ ≤ ‖f‖ₑ * ‖x‖ₑ := by dsimp [enorm]; exact mod_cast le_opNNNorm .. theorem nndist_le_opNNNorm (x y : E) : nndist (f x) (f y) ≤ ‖f‖₊ * nndist x y := dist_le_opNorm f x y diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean index edb1a33efca0f..2f9a18162eca1 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean @@ -98,14 +98,7 @@ noncomputable def projectiveSeminorm : Seminorm 𝕜 (⨂[𝕜] i, E i) := by refine Seminorm.ofSMulLE (fun x ↦ iInf (fun (p : lifts x) ↦ projectiveSeminormAux p.1)) ?_ ?_ ?_ · refine le_antisymm ?_ ?_ · refine ciInf_le_of_le (bddBelow_projectiveSemiNormAux (0 : ⨂[𝕜] i, E i)) ⟨0, lifts_zero⟩ ?_ - simp only [projectiveSeminormAux, Function.comp_apply] - rw [List.sum_eq_zero] - intro _ - simp only [List.mem_map, Prod.exists, forall_exists_index, and_imp] - intro _ _ hxm - rw [← FreeAddMonoid.ofList_nil] at hxm - exfalso - exact List.not_mem_nil _ hxm + rfl · letI : Nonempty (lifts 0) := ⟨0, lifts_zero (R := 𝕜) (s := E)⟩ exact le_ciInf (fun p ↦ projectiveSeminormAux_nonneg p.1) · intro x y diff --git a/Mathlib/Analysis/Quaternion.lean b/Mathlib/Analysis/Quaternion.lean index 9f75be8137503..6c28375ee66a8 100644 --- a/Mathlib/Analysis/Quaternion.lean +++ b/Mathlib/Analysis/Quaternion.lean @@ -163,7 +163,7 @@ theorem norm_piLp_equiv_symm_equivTuple (x : ℍ) : /-- `QuaternionAlgebra.linearEquivTuple` as a `LinearIsometryEquiv`. -/ @[simps apply symm_apply] noncomputable def linearIsometryEquivTuple : ℍ ≃ₗᵢ[ℝ] EuclideanSpace ℝ (Fin 4) := - { (QuaternionAlgebra.linearEquivTuple (-1 : ℝ) (-1 : ℝ)).trans + { (QuaternionAlgebra.linearEquivTuple (-1 : ℝ) (0 : ℝ) (-1 : ℝ)).trans (WithLp.linearEquiv 2 ℝ (Fin 4 → ℝ)).symm with toFun := fun a => !₂[a.1, a.2, a.3, a.4] invFun := fun a => ⟨a 0, a 1, a 2, a 3⟩ @@ -209,13 +209,16 @@ variable {α : Type*} @[simp, norm_cast] theorem hasSum_coe {f : α → ℝ} {r : ℝ} : HasSum (fun a => (f a : ℍ)) (↑r : ℍ) ↔ HasSum f r := - ⟨fun h => by simpa only using h.map (show ℍ →ₗ[ℝ] ℝ from QuaternionAlgebra.reₗ _ _) continuous_re, + ⟨fun h => by + simpa only using + h.map (show ℍ →ₗ[ℝ] ℝ from QuaternionAlgebra.reₗ _ _ _) continuous_re, fun h => by simpa only using h.map (algebraMap ℝ ℍ) (continuous_algebraMap _ _)⟩ @[simp, norm_cast] theorem summable_coe {f : α → ℝ} : (Summable fun a => (f a : ℍ)) ↔ Summable f := by simpa only using - Summable.map_iff_of_leftInverse (algebraMap ℝ ℍ) (show ℍ →ₗ[ℝ] ℝ from QuaternionAlgebra.reₗ _ _) + Summable.map_iff_of_leftInverse (algebraMap ℝ ℍ) (show ℍ →ₗ[ℝ] ℝ from + QuaternionAlgebra.reₗ _ _ _) (continuous_algebraMap _ _) continuous_re coe_re @[norm_cast] diff --git a/Mathlib/Analysis/RCLike/Inner.lean b/Mathlib/Analysis/RCLike/Inner.lean index 4136cd8069e82..31f4ce68c9de2 100644 --- a/Mathlib/Analysis/RCLike/Inner.lean +++ b/Mathlib/Analysis/RCLike/Inner.lean @@ -32,7 +32,7 @@ variable [RCLike 𝕜] section Pi variable [∀ i, SeminormedAddCommGroup (E i)] [∀ i, InnerProductSpace 𝕜 (E i)] {w : ι → ℝ} -/-- Weighted inner product giving rise to the L2 norm. -/ +/-- Weighted inner product giving rise to the L2 norm, denoted as `⟪g, f⟫_[𝕜, w]`. -/ def wInner (w : ι → ℝ) (f g : ∀ i, E i) : 𝕜 := ∑ i, w i • inner (f i) (g i) /-- The weight function making `wInner` into the compact inner product. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean index ef1748b715e87..45424cd36fd7d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Analytic.lean @@ -21,6 +21,7 @@ variable {E : Type} [NormedAddCommGroup E] [NormedSpace ℂ E] variable {f g : E → ℂ} {z : ℂ} {x : E} {s : Set E} /-- `log` is analytic away from nonpositive reals -/ +@[fun_prop] theorem analyticAt_clog (m : z ∈ slitPlane) : AnalyticAt ℂ log z := by rw [analyticAt_iff_eventually_differentiableAt] filter_upwards [isOpen_slitPlane.eventually_mem m] @@ -28,6 +29,7 @@ theorem analyticAt_clog (m : z ∈ slitPlane) : AnalyticAt ℂ log z := by exact differentiableAt_id.clog m /-- `log` is analytic away from nonpositive reals -/ +@[fun_prop] theorem AnalyticAt.clog (fa : AnalyticAt ℂ f x) (m : f x ∈ slitPlane) : AnalyticAt ℂ (fun z ↦ log (f z)) x := (analyticAt_clog m).comp fa @@ -56,6 +58,7 @@ theorem AnalyticWithinAt.cpow (fa : AnalyticWithinAt ℂ f s x) (ga : AnalyticWi exact ((fa.clog m).mul ga).cexp /-- `f z ^ g z` is analytic if `f z` is not a nonpositive real -/ +@[fun_prop] theorem AnalyticAt.cpow (fa : AnalyticAt ℂ f x) (ga : AnalyticAt ℂ g x) (m : f x ∈ slitPlane) : AnalyticAt ℂ (fun z ↦ f z ^ g z) x := by rw [← analyticWithinAt_univ] at fa ga ⊢ diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 1b221e58e8826..e2af86fcb1b81 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -176,8 +176,11 @@ theorem arg_eq_arg_iff {x y : ℂ} (hx : x ≠ 0) (hy : y ≠ 0) : rw [← ofReal_div, arg_real_mul] exact div_pos (abs.pos hy) (abs.pos hx) -@[simp] -theorem arg_one : arg 1 = 0 := by simp [arg, zero_le_one] +@[simp] lemma arg_one : arg 1 = 0 := by simp [arg, zero_le_one] + +/-- This holds true for all `x : ℂ` because of the junk values `0 / 0 = 0` and `arg 0 = 0`. -/ +@[simp] lemma arg_div_self (x : ℂ) : arg (x / x) = 0 := by + obtain rfl | hx := eq_or_ne x 0 <;> simp [*] @[simp] theorem arg_neg_one : arg (-1) = π := by simp [arg, le_refl, not_le.2 (zero_lt_one' ℝ)] diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean index 60934d34bd710..34903e135a74d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Log.lean @@ -93,6 +93,9 @@ theorem log_zero : log 0 = 0 := by simp [log] @[simp] theorem log_one : log 1 = 0 := by simp [log] +/-- This holds true for all `x : ℂ` because of the junk values `0 / 0 = 0` and `log 0 = 0`. -/ +@[simp] lemma log_div_self (x : ℂ) : log (x / x) = 0 := by simp [log] + theorem log_neg_one : log (-1) = π * I := by simp [log] theorem log_I : log I = π / 2 * I := by simp [log] @@ -255,62 +258,34 @@ variable {α ι: Type*} open Real -lemma Real.HasSum_rexp_HasProd (f : ι → α → ℝ) (hfn : ∀ x n, 0 < f n x) - (hf : ∀ x : α, HasSum (fun n => log (f n x)) (∑' i, log (f i x))) (a : α) : - HasProd (fun b ↦ f b a) (∏' n : ι, (f n a)) := by - have : HasProd (fun b ↦ f b a) ((rexp ∘ fun a ↦ ∑' (n : ι), log (f n a)) a) := by - apply ((hf a).rexp).congr - intro _ - congr - exact funext fun x ↦ exp_log (hfn a x) - rwa [HasProd.tprod_eq this] +lemma Real.hasProd_of_hasSum_log {f : ι → ℝ} (hfn : ∀ n, 0 < f n) {a : ℝ} + (hf : HasSum (fun n => log (f n)) a) : HasProd f (rexp a) := + hf.rexp.congr (by simp [exp_log, hfn]) +lemma Real.multipliable_of_summable_log (f : ι → ℝ) (hfn : ∀ n, 0 < f n) + (hf : Summable fun n => log (f n)) : Multipliable f := + ⟨_, Real.hasProd_of_hasSum_log hfn hf.hasSum⟩ -/--The exponential of a infinite sum of real logs (which converges absolutely) is an infinite -product.-/ -lemma Real.rexp_tsum_eq_tprod (f : ι → α → ℝ) (hfn : ∀ x n, 0 < f n x) - (hf : ∀ x : α, Summable fun n => log ((f n x))) : - (rexp ∘ (fun a : α => (∑' n : ι, log (f n a)))) = (fun a : α => ∏' n : ι, (f n a)) := by - ext a - apply (HasProd.tprod_eq ?_).symm - apply ((hf a).hasSum.rexp).congr - intro _ - congr - exact funext fun x ↦ exp_log (hfn a x) - -lemma Real.summable_cexp_multipliable (f : ι → α → ℝ) (hfn : ∀ x n, 0 < f n x) - (hf : ∀ x : α, Summable fun n => log (f n x)) (a : α) : Multipliable fun b ↦ f b a := by - have := (Real.HasSum_rexp_HasProd f hfn fun a => (hf a).hasSum) a - use (∏' n : ι, (f n a)) +/-- The exponential of a infinite sum of real logs (which converges absolutely) is an infinite +product. -/ +lemma Real.rexp_tsum_eq_tprod (f : ι → ℝ) (hfn : ∀ n, 0 < f n) + (hf : Summable fun n => log (f n)) : rexp (∑' n : ι, log (f n)) = ∏' n : ι, f n := + (Real.hasProd_of_hasSum_log hfn hf.hasSum).tprod_eq.symm open Complex -lemma Complex.HasSum_cexp_HasProd (f : ι → α → ℂ) (hfn : ∀ x n, f n x ≠ 0) - (hf : ∀ x : α, HasSum (fun n => log (f n x)) (∑' i, log (f i x))) (a : α) : - HasProd (fun b ↦ f b a) (∏' n : ι, (f n a)) := by - have : HasProd (fun b ↦ f b a) ((cexp ∘ fun a ↦ ∑' (n : ι), log (f n a)) a) := by - apply ((hf a).cexp).congr - intro _ - congr - exact funext fun x ↦ exp_log (hfn a x) - rwa [HasProd.tprod_eq this] - -lemma Complex.summable_cexp_multipliable (f : ι → α → ℂ) (hfn : ∀ x n, f n x ≠ 0) - (hf : ∀ x : α, Summable fun n => log (f n x)) (a : α) : - Multipliable fun b ↦ f b a := by - have := (Complex.HasSum_cexp_HasProd f hfn fun a => (hf a).hasSum) a - use (∏' n : ι, (f n a)) +lemma Complex.hasProd_of_hasSum_log (f : ι → ℂ) (hfn : ∀ n, f n ≠ 0) {a : ℂ} + (hf : HasSum (fun n => log (f n)) a) : HasProd (fun b ↦ f b) (cexp a) := + hf.cexp.congr (by simp [exp_log, hfn]) + +lemma Complex.multipliable_of_summable_log (f : ι → ℂ) (hfn : ∀ n, f n ≠ 0) + (hf : Summable fun n => log (f n)) : Multipliable fun b ↦ f b := + ⟨_, Complex.hasProd_of_hasSum_log _ hfn hf.hasSum⟩ /--The exponential of a infinite sum of comples logs (which converges absolutely) is an infinite product.-/ -lemma Complex.cexp_tsum_eq_tprod (f : ι → α → ℂ) (hfn : ∀ x n, f n x ≠ 0) - (hf : ∀ x : α, Summable fun n => log (f n x)) : - (cexp ∘ (fun a : α => (∑' n : ι, log (f n a)))) = fun a : α => ∏' n : ι, f n a := by - ext a - apply (HasProd.tprod_eq ?_).symm - apply ((hf a).hasSum.cexp).congr - intro _ - congr - exact funext fun x ↦ exp_log (hfn a x) +lemma Complex.cexp_tsum_eq_tprod (f : ι → ℂ) (hfn : ∀ n, f n ≠ 0) + (hf : Summable fun n => log (f n)) : (cexp ((∑' n : ι, log (f n )))) = ∏' n : ι, f n := + (Complex.hasProd_of_hasSum_log _ hfn hf.hasSum).tprod_eq.symm end tsum_tprod diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow/Basic.lean similarity index 99% rename from Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean rename to Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow/Basic.lean index 89d6038a11e76..4da05cf09f25a 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow/Basic.lean @@ -60,6 +60,9 @@ lemma continuous_nnrpow_const (y : ℝ≥0) : Continuous (nnrpow · y) := /- This is a "redeclaration" of the attribute to speed up the proofs in this file. -/ attribute [fun_prop] continuousOn_rpow_const +lemma monotone_nnrpow_const (y : ℝ≥0) : Monotone (nnrpow · y) := + monotone_rpow_of_nonneg zero_le_coe + end NNReal namespace CFC diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow/Isometric.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow/Isometric.lean new file mode 100644 index 0000000000000..55a50094d5046 --- /dev/null +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow/Isometric.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2025 Frédéric Dupuis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Frédéric Dupuis +-/ + +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow.Basic +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Isometric + +/-! # Properties of `rpow` and `sqrt` over an algebra with an isometric CFC + +This file collects results about `CFC.rpow`, `CFC.nnrpow` and `CFC.sqrt` that use facts that +rely on an isometric continuous functional calculus. + +## Main theorems + +* Various versions of `‖a ^ r‖ = ‖a‖ ^ r` and `‖CFC.sqrt a‖ = sqrt ‖a‖`. + +## Tags + +continuous functional calculus, rpow, sqrt +-/ + +open scoped NNReal + +namespace CFC + +section nonunital + +variable {A : Type*} [NonUnitalNormedRing A] [StarRing A] [NormedSpace ℝ A] [IsScalarTower ℝ A A] + [SMulCommClass ℝ A A] [PartialOrder A] [StarOrderedRing A] [NonnegSpectrumClass ℝ A] + [NonUnitalIsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint] + +lemma nnnorm_nnrpow (a : A) {r : ℝ≥0} (hr : 0 < r) (ha : 0 ≤ a := by cfc_tac) : + ‖a ^ r‖₊ = ‖a‖₊ ^ (r : ℝ) := + NNReal.monotone_nnrpow_const r |>.monotoneOn _ |>.nnnorm_cfcₙ _ _ + +lemma norm_nnrpow (a : A) {r : ℝ≥0} (hr : 0 < r) (ha : 0 ≤ a := by cfc_tac) : + ‖a ^ r‖ = ‖a‖ ^ (r : ℝ) := + congr(NNReal.toReal $(nnnorm_nnrpow a hr ha)) + +lemma nnnorm_sqrt (a : A) (ha : 0 ≤ a := by cfc_tac) : ‖sqrt a‖₊ = NNReal.sqrt ‖a‖₊ := by + rw [sqrt_eq_nnrpow, NNReal.sqrt_eq_rpow] + exact nnnorm_nnrpow a (by norm_num) ha + +lemma norm_sqrt (a : A) (ha : 0 ≤ a := by cfc_tac) : ‖sqrt a‖ = Real.sqrt ‖a‖ := by + simpa using congr(NNReal.toReal $(nnnorm_sqrt a ha)) + +end nonunital + +section unital + +variable {A : Type*} [NormedRing A] [StarRing A] [NormedAlgebra ℝ A] + [PartialOrder A] [StarOrderedRing A] [NonnegSpectrumClass ℝ A] + [IsometricContinuousFunctionalCalculus ℝ A IsSelfAdjoint] + +lemma nnnorm_rpow (a : A) {r : ℝ} (hr : 0 < r) (ha : 0 ≤ a := by cfc_tac) : + ‖a ^ r‖₊ = ‖a‖₊ ^ r := by + lift r to ℝ≥0 using hr.le + rw [← nnrpow_eq_rpow, ← nnnorm_nnrpow a] + all_goals simpa + +lemma norm_rpow (a : A) {r : ℝ} (hr : 0 < r) (ha : 0 ≤ a := by cfc_tac) : + ‖a ^ r‖ = ‖a‖ ^ r := + congr(NNReal.toReal $(nnnorm_rpow a hr ha)) + +end unital + +end CFC diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index d13b1341c8a8d..63ef6f5d143ab 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -43,13 +43,20 @@ theorem analyticOnNhd_cexp : AnalyticOnNhd ℂ exp univ := by theorem analyticOn_cexp : AnalyticOn ℂ exp univ := analyticOnNhd_cexp.analyticOn /-- `exp` is analytic at any point -/ +@[fun_prop] theorem analyticAt_cexp : AnalyticAt ℂ exp z := analyticOnNhd_cexp z (mem_univ _) /-- `exp ∘ f` is analytic -/ -theorem AnalyticAt.cexp (fa : AnalyticAt ℂ f x) : AnalyticAt ℂ (fun z ↦ exp (f z)) x := +@[fun_prop] +theorem AnalyticAt.cexp (fa : AnalyticAt ℂ f x) : AnalyticAt ℂ (exp ∘ f) x := analyticAt_cexp.comp fa +/-- `exp ∘ f` is analytic -/ +@[fun_prop] +theorem AnalyticAt.cexp' (fa : AnalyticAt ℂ f x) : AnalyticAt ℂ (fun z ↦ exp (f z)) x := + fa.cexp + theorem AnalyticWithinAt.cexp (fa : AnalyticWithinAt ℂ f s x) : AnalyticWithinAt ℂ (fun z ↦ exp (f z)) s x := analyticAt_cexp.comp_analyticWithinAt fa @@ -204,13 +211,20 @@ theorem analyticOnNhd_rexp : AnalyticOnNhd ℝ exp univ := by theorem analyticOn_rexp : AnalyticOn ℝ exp univ := analyticOnNhd_rexp.analyticOn /-- `exp` is analytic at any point -/ +@[fun_prop] theorem analyticAt_rexp : AnalyticAt ℝ exp x := analyticOnNhd_rexp x (mem_univ _) /-- `exp ∘ f` is analytic -/ -theorem AnalyticAt.rexp {x : E} (fa : AnalyticAt ℝ f x) : AnalyticAt ℝ (fun z ↦ exp (f z)) x := +@[fun_prop] +theorem AnalyticAt.rexp {x : E} (fa : AnalyticAt ℝ f x) : AnalyticAt ℝ (exp ∘ f) x := analyticAt_rexp.comp fa +/-- `exp ∘ f` is analytic -/ +@[fun_prop] +theorem AnalyticAt.rexp' {x : E} (fa : AnalyticAt ℝ f x) : AnalyticAt ℝ (fun z ↦ exp (f z)) x := + fa.rexp + theorem AnalyticWithinAt.rexp {x : E} (fa : AnalyticWithinAt ℝ f s x) : AnalyticWithinAt ℝ (fun z ↦ exp (f z)) s x := analyticAt_rexp.comp_analyticWithinAt fa diff --git a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean index 3ef729e02e9f1..4e8ea696d800a 100644 --- a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean @@ -81,6 +81,14 @@ theorem integrableOn_Ioi_rpow_iff {s t : ℝ} (ht : 0 < t) : exact Real.rpow_le_rpow_of_exponent_le x_one h exact not_IntegrableOn_Ioi_inv this +theorem integrableAtFilter_rpow_atTop_iff {s : ℝ} : + IntegrableAtFilter (fun x : ℝ ↦ x ^ s) atTop ↔ s < -1 := by + refine ⟨fun ⟨t, ht, hint⟩ ↦ ?_, fun h ↦ + ⟨Set.Ioi 1, Ioi_mem_atTop 1, (integrableOn_Ioi_rpow_iff zero_lt_one).mpr h⟩⟩ + obtain ⟨a, ha⟩ := mem_atTop_sets.mp ht + refine (integrableOn_Ioi_rpow_iff (zero_lt_one.trans_le (le_max_right a 1))).mp ?_ + exact hint.mono_set <| fun x hx ↦ ha _ <| (le_max_left a 1).trans hx.le + /-- The real power function with any exponent is not integrable on `(0, +∞)`. -/ theorem not_integrableOn_Ioi_rpow (s : ℝ) : ¬ IntegrableOn (fun x ↦ x ^ s) (Ioi (0 : ℝ)) := by intro h @@ -127,6 +135,24 @@ theorem integrableOn_Ioi_cpow_iff {s : ℂ} {t : ℝ} (ht : 0 < t) : simp [Complex.abs_cpow_eq_rpow_re_of_pos this] rwa [integrableOn_Ioi_rpow_iff ht] at B +theorem integrableOn_Ioi_deriv_ofReal_cpow {s : ℂ} {t : ℝ} (ht : 0 < t) (hs : s.re < 0) : + IntegrableOn (deriv fun x : ℝ ↦ (x : ℂ) ^ s) (Set.Ioi t) := by + have h : IntegrableOn (fun x : ℝ ↦ s * x ^ (s - 1)) (Set.Ioi t) := by + refine (integrableOn_Ioi_cpow_of_lt ?_ ht).const_mul _ + rwa [Complex.sub_re, Complex.one_re, sub_lt_iff_lt_add, neg_add_cancel] + refine h.congr_fun (fun x hx ↦ ?_) measurableSet_Ioi + rw [Complex.deriv_ofReal_cpow_const (ht.trans hx).ne' (fun h ↦ (Complex.zero_re ▸ h ▸ hs).false)] + +theorem integrableOn_Ioi_deriv_norm_ofReal_cpow {s : ℂ} {t : ℝ} (ht : 0 < t) (hs : s.re ≤ 0): + IntegrableOn (deriv fun x : ℝ ↦ ‖(x : ℂ) ^ s‖) (Set.Ioi t) := by + rw [integrableOn_congr_fun (fun x hx ↦ by + rw [deriv_norm_ofReal_cpow _ (ht.trans hx)]) measurableSet_Ioi] + obtain hs | hs := eq_or_lt_of_le hs + · simp_rw [hs, zero_mul] + exact integrableOn_zero + · replace hs : s.re - 1 < - 1 := by rwa [sub_lt_iff_lt_add, neg_add_cancel] + exact (integrableOn_Ioi_rpow_of_lt hs ht).const_mul s.re + /-- The complex power function with any exponent is not integrable on `(0, +∞)`. -/ theorem not_integrableOn_Ioi_cpow (s : ℂ) : ¬ IntegrableOn (fun x : ℝ ↦ (x : ℂ) ^ s) (Ioi (0 : ℝ)) := by diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 23f0796e7d2ac..ee0b68c3a430b 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -353,7 +353,7 @@ theorem integral_cpow {r : ℂ} (h : -1 < r.re ∨ r ≠ -1 ∧ (0 : ℝ) ∉ [[ by_cases hab : (0 : ℝ) ∉ [[a, b]] · apply integral_eq_sub_of_hasDerivAt (fun x hx => ?_) (intervalIntegrable_cpow (r := r) <| Or.inr hab) - refine hasDerivAt_ofReal_cpow (ne_of_mem_of_not_mem hx hab) ?_ + refine hasDerivAt_ofReal_cpow_const' (ne_of_mem_of_not_mem hx hab) ?_ contrapose! hr; rwa [add_eq_zero_iff_eq_neg] replace h : -1 < r.re := by tauto suffices ∀ c : ℝ, (∫ x : ℝ in (0)..c, (x : ℂ) ^ r) = @@ -365,7 +365,7 @@ theorem integral_cpow {r : ℂ} (h : -1 < r.re ∨ r ≠ -1 ∧ (0 : ℝ) ∉ [[ apply integral_eq_sub_of_hasDeriv_right · refine ((Complex.continuous_ofReal_cpow_const ?_).div_const _).continuousOn rwa [Complex.add_re, Complex.one_re, ← neg_lt_iff_pos_add] - · refine fun x hx => (hasDerivAt_ofReal_cpow ?_ ?_).hasDerivWithinAt + · refine fun x hx => (hasDerivAt_ofReal_cpow_const' ?_ ?_).hasDerivWithinAt · rcases le_total c 0 with (hc | hc) · rw [max_eq_left hc] at hx; exact hx.2.ne · rw [min_eq_left hc] at hx; exact hx.1.ne' diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 3091ec96f13d0..3944edf91c60c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -86,6 +86,10 @@ theorem log_zero : log 0 = 0 := theorem log_one : log 1 = 0 := exp_injective <| by rw [exp_log zero_lt_one, exp_zero] +/-- This holds true for all `x : \` because of the junk values `0 / 0 = 0` and `arg 0 = 0`. -/ +@[simp] lemma log_div_self (x : ℝ) : log (x / x) = 0 := by + obtain rfl | hx := eq_or_ne x 0 <;> simp [*] + @[simp] theorem log_abs (x : ℝ) : log |x| = log x := by by_cases h : x = 0 diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean b/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean new file mode 100644 index 0000000000000..3c1969101ea96 --- /dev/null +++ b/Mathlib/Analysis/SpecialFunctions/Log/Summable.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2024 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ + +import Mathlib.Analysis.SpecialFunctions.Complex.LogBounds +import Mathlib.Analysis.NormedSpace.FunctionSeries + +/-! +# Summable logs + +We give conditions under which the logarithms of a summble sequence is summable. We also give some +results about when the sums converge uniformly. + +TODO: Generalise the indexing set from ℕ to some arbitrary type, but this needs +Summable.tendsto_atTop_zero to first be done. Also remove hff from +`Complex.multipliable_one_add_of_summable` once we have vanishing/non-vanishing results for infinite +products. + +-/ + +open Filter Function Complex Real + +open scoped Interval Topology BigOperators Nat Classical Complex + +variable {ι : Type*} + +lemma Complex.summable_log_one_add_of_summable {f : ι → ℂ} (hf : Summable f) : + Summable (fun i : ι => Complex.log (1 + f i)) := by + apply (hf.norm.const_smul (3 / 2 : ℝ)).of_norm_bounded_eventually + filter_upwards [hf.norm.tendsto_cofinite_zero.eventually_le_const one_half_pos] with i hi + exact norm_log_one_add_half_le_self hi + +lemma Real.summable_log_one_add_of_summable {f : ι → ℝ} (hf : Summable f) : + Summable (fun i : ι => log (1 + |f i|)) := by + have : Summable (fun n ↦ Complex.ofRealCLM (log (1 + |f n|))) := by + convert Complex.summable_log_one_add_of_summable (Complex.ofRealCLM.summable hf.norm) with x + rw [ofRealCLM_apply, ofReal_log (by positivity)] + simp only [ofReal_add, ofReal_one, norm_eq_abs, ofRealCLM_apply] + convert Complex.reCLM.summable this + +lemma Complex.multipliable_one_add_of_summable (f : ι → ℂ) (hf : Summable f) + (hff : ∀ n : ι, 1 + f n ≠ 0) : Multipliable (fun n : ι => 1 + f n) := by + refine Complex.multipliable_of_summable_log (fun n => 1 + f n) (by simpa) ?_ + simpa only [forall_const] using Complex.summable_log_one_add_of_summable hf + +lemma Real.multipliable_one_add_of_summable (f : ι → ℝ) (hf : Summable f) : + Multipliable (fun n : ι => 1 + |f n|) := by + refine Real.multipliable_of_summable_log (fun n => 1 + |f n|) (fun _ ↦ by positivity) ?_ + simpa only [forall_const] using Real.summable_log_one_add_of_summable hf + +lemma Complex.tendstoUniformlyOn_tsum_nat_log_one_add {α : Type*} {f : ℕ → α → ℂ} (K : Set α) + {u : ℕ → ℝ} (hu : Summable u) (h : ∀ᶠ n in atTop, ∀ x ∈ K, ‖f n x‖ ≤ u n) : + TendstoUniformlyOn (fun (n : ℕ) (a : α) => ∑ i in Finset.range n, + (Complex.log (1 + f i a))) (fun a => ∑' i : ℕ, Complex.log (1 + f i a)) atTop K := by + apply tendstoUniformlyOn_tsum_nat_eventually (hu.mul_left (3/2)) + obtain ⟨N, hN⟩ := Metric.tendsto_atTop.mp (Summable.tendsto_atTop_zero hu) (1/2) (one_half_pos) + simp only [Complex.norm_eq_abs, eventually_atTop, ge_iff_le] at * + obtain ⟨N2, hN2⟩ := h + refine ⟨max N N2, fun n hn x hx => ?_⟩ + apply le_trans (Complex.norm_log_one_add_half_le_self (z := (f n x)) ?_) + · simp only [Complex.norm_eq_abs, Nat.ofNat_pos, div_pos_iff_of_pos_left, mul_le_mul_left] + exact hN2 n (le_of_max_le_right hn) x hx + · apply le_trans (le_trans (hN2 n (le_of_max_le_right hn) x hx) + (by simpa using Real.le_norm_self (u n))) (hN n (le_of_max_le_left hn)).le diff --git a/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean b/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean index bbbd460e9abe3..557084bf800d2 100644 --- a/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean +++ b/Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean @@ -67,8 +67,8 @@ noncomputable def ordinaryHypergeometricSeries (a b c : 𝕂) : FormalMultilinea variable {𝔸} (a b c : 𝕂) -/-- `ordinaryHypergeometric (a b c : 𝕂) : 𝔸 → 𝔸` is the ordinary hypergeometric map, defined as the -sum of the `FormalMultilinearSeries` `ordinaryHypergeometricSeries 𝔸 a b c`. +/-- `ordinaryHypergeometric (a b c : 𝕂) : 𝔸 → 𝔸`, denoted `₂F₁`, is the ordinary hypergeometric map, +defined as the sum of the `FormalMultilinearSeries` `ordinaryHypergeometricSeries 𝔸 a b c`. Note that this takes the junk value `0` outside the radius of convergence. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index 016678accf824..86233fe6cde8c 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -89,20 +89,24 @@ def polarCoord : PartialHomeomorph (ℝ × ℝ) (ℝ × ℝ) where · exact (Complex.continuousAt_arg hz).continuousWithinAt · exact Complex.equivRealProdCLM.symm.continuous.continuousOn +/-- The derivative of `polarCoord.symm`, see `hasFDerivAt_polarCoord_symm`. -/ +def fderivPolarCoordSymm (p : ℝ × ℝ) : ℝ × ℝ →L[ℝ] ℝ × ℝ := + LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) + (Basis.finTwoProd ℝ) !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2]) + theorem hasFDerivAt_polarCoord_symm (p : ℝ × ℝ) : - HasFDerivAt polarCoord.symm - (LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) - !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2])) p := by + HasFDerivAt polarCoord.symm (fderivPolarCoordSymm p) p := by + unfold fderivPolarCoordSymm rw [Matrix.toLin_finTwoProd_toContinuousLinearMap] convert HasFDerivAt.prod (𝕜 := ℝ) (hasFDerivAt_fst.mul ((hasDerivAt_cos p.2).comp_hasFDerivAt p hasFDerivAt_snd)) (hasFDerivAt_fst.mul ((hasDerivAt_sin p.2).comp_hasFDerivAt p hasFDerivAt_snd)) using 2 <;> simp [smul_smul, add_comm, neg_mul, smul_neg, neg_smul _ (ContinuousLinearMap.snd ℝ ℝ ℝ)] -theorem det_fderiv_polarCoord_symm (p : ℝ × ℝ) : - (LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) - !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2])).det = p.1 := by +theorem det_fderivPolarCoordSymm (p : ℝ × ℝ) : + (fderivPolarCoordSymm p).det = p.1 := by conv_rhs => rw [← one_mul p.1, ← cos_sq_add_sin_sq p.2] + unfold fderivPolarCoordSymm simp only [neg_mul, LinearMap.det_toContinuousLinearMap, LinearMap.det_toLin, Matrix.det_fin_two_of, sub_neg_eq_add] ring @@ -138,7 +142,7 @@ theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup E] [Normed _ = ∫ p in polarCoord.target, |p.1| • f (polarCoord.symm p) := by rw [← PartialHomeomorph.symm_target, integral_target_eq_integral_abs_det_fderiv_smul volume (fun p _ ↦ hasFDerivAt_polarCoord_symm p), PartialHomeomorph.symm_source] - simp_rw [det_fderiv_polarCoord_symm] + simp_rw [det_fderivPolarCoordSymm] _ = ∫ p in polarCoord.target, p.1 • f (polarCoord.symm p) := by apply setIntegral_congr_fun polarCoord.open_target.measurableSet fun x hx => ?_ rw [abs_of_pos hx.1] @@ -154,7 +158,7 @@ theorem lintegral_comp_polarCoord_symm (f : ℝ × ℝ → ℝ≥0∞) : _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, ENNReal.ofReal |p.1| • f (polarCoord.symm p) := by rw [lintegral_image_eq_lintegral_abs_det_fderiv_mul volume _ (fun p _ ↦ (hasFDerivAt_polarCoord_symm p).hasFDerivWithinAt)] - · simp_rw [det_fderiv_polarCoord_symm]; rfl + · simp_rw [det_fderivPolarCoordSymm]; rfl exacts [polarCoord.symm.injOn, measurableSet_Ioi.prod measurableSet_Ioo] _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, ENNReal.ofReal p.1 • f (polarCoord.symm p) := by refine setLIntegral_congr_fun polarCoord.open_target.measurableSet ?_ @@ -211,3 +215,87 @@ protected theorem lintegral_comp_polarCoord_symm (f : ℂ → ℝ≥0∞) : simp_rw [measurableEquivRealProd_symm_polarCoord_symm_apply] end Complex + +section Pi + +open ENNReal MeasureTheory MeasureTheory.Measure + +variable {ι : Type*} + +open ContinuousLinearMap in +/-- The derivative of `polarCoord.symm` on `ι → ℝ × ℝ`, see `hasFDerivAt_pi_polarCoord_symm`. -/ +noncomputable def fderivPiPolarCoordSymm (p : ι → ℝ × ℝ) : (ι → ℝ × ℝ) →L[ℝ] ι → ℝ × ℝ := + pi fun i ↦ (fderivPolarCoordSymm (p i)).comp (proj i) + +theorem injOn_pi_polarCoord_symm : + Set.InjOn (fun p (i : ι) ↦ polarCoord.symm (p i)) (Set.univ.pi fun _ ↦ polarCoord.target) := + fun _ hx _ hy h ↦ funext fun i ↦ polarCoord.symm.injOn (hx i trivial) (hy i trivial) + ((funext_iff.mp h) i) + +theorem abs_fst_of_mem_pi_polarCoord_target {p : ι → ℝ × ℝ} + (hp : p ∈ (Set.univ.pi fun _ : ι ↦ polarCoord.target)) (i : ι) : + |(p i).1| = (p i).1 := + abs_of_pos ((Set.mem_univ_pi.mp hp) i).1 + +variable [Fintype ι] + +theorem hasFDerivAt_pi_polarCoord_symm (p : ι → ℝ × ℝ) : + HasFDerivAt (fun x i ↦ polarCoord.symm (x i)) (fderivPiPolarCoordSymm p) p := by + rw [fderivPiPolarCoordSymm, hasFDerivAt_pi] + exact fun i ↦ HasFDerivAt.comp _ (hasFDerivAt_polarCoord_symm _) (hasFDerivAt_apply i _) + +theorem det_fderivPiPolarCoordSymm (p : ι → ℝ × ℝ) : + (fderivPiPolarCoordSymm p).det = ∏ i, (p i).1 := by + simp_rw [fderivPiPolarCoordSymm, ContinuousLinearMap.det_pi, det_fderivPolarCoordSymm] + +theorem pi_polarCoord_symm_target_ae_eq_univ : + (Pi.map (fun _ : ι ↦ polarCoord.symm) '' Set.univ.pi fun _ ↦ polarCoord.target) + =ᵐ[volume] Set.univ := by + rw [Set.piMap_image_univ_pi, polarCoord.symm_image_target_eq_source, volume_pi, ← Set.pi_univ] + exact ae_eq_set_pi fun _ _ ↦ polarCoord_source_ae_eq_univ + +theorem measurableSet_pi_polarCoord_target : + MeasurableSet (Set.univ.pi fun _ : ι ↦ polarCoord.target) := + MeasurableSet.univ_pi fun _ ↦ polarCoord.open_target.measurableSet + +theorem integral_comp_pi_polarCoord_symm {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + (f : (ι → ℝ × ℝ) → E) : + (∫ p in (Set.univ.pi fun _ : ι ↦ polarCoord.target), + (∏ i, (p i).1) • f (fun i ↦ polarCoord.symm (p i))) = ∫ p, f p := by + rw [← setIntegral_univ (f := f), ← setIntegral_congr_set pi_polarCoord_symm_target_ae_eq_univ] + convert (integral_image_eq_integral_abs_det_fderiv_smul volume measurableSet_pi_polarCoord_target + (fun p _ ↦ (hasFDerivAt_pi_polarCoord_symm p).hasFDerivWithinAt) + injOn_pi_polarCoord_symm f).symm using 1 + refine setIntegral_congr_fun measurableSet_pi_polarCoord_target fun x hx ↦ ?_ + simp_rw [det_fderivPiPolarCoordSymm, Finset.abs_prod, abs_fst_of_mem_pi_polarCoord_target hx] + +protected theorem Complex.integral_comp_pi_polarCoord_symm {E : Type*} [NormedAddCommGroup E] + [NormedSpace ℝ E] (f : (ι → ℂ) → E) : + (∫ p in (Set.univ.pi fun _ : ι ↦ Complex.polarCoord.target), + (∏ i, (p i).1) • f (fun i ↦ Complex.polarCoord.symm (p i))) = ∫ p, f p := by + let e := MeasurableEquiv.piCongrRight (fun _ : ι ↦ measurableEquivRealProd.symm) + have := volume_preserving_pi (fun _ : ι ↦ Complex.volume_preserving_equiv_real_prod.symm) + rw [← MeasurePreserving.integral_comp this e.measurableEmbedding f] + exact integral_comp_pi_polarCoord_symm (f ∘ e) + +theorem lintegral_comp_pi_polarCoord_symm (f : (ι → ℝ × ℝ) → ℝ≥0∞) : + ∫⁻ p in (Set.univ.pi fun _ : ι ↦ polarCoord.target), + (∏ i, .ofReal (p i).1) * f (fun i ↦ polarCoord.symm (p i)) = ∫⁻ p, f p := by + rw [← setLIntegral_univ f, ← setLIntegral_congr pi_polarCoord_symm_target_ae_eq_univ] + convert (lintegral_image_eq_lintegral_abs_det_fderiv_mul volume measurableSet_pi_polarCoord_target + (fun p _ ↦ (hasFDerivAt_pi_polarCoord_symm p).hasFDerivWithinAt) + injOn_pi_polarCoord_symm f).symm using 1 + refine setLIntegral_congr_fun measurableSet_pi_polarCoord_target ?_ + filter_upwards with x hx + simp_rw [det_fderivPiPolarCoordSymm, Finset.abs_prod, ENNReal.ofReal_prod_of_nonneg (fun _ _ ↦ + abs_nonneg _), abs_fst_of_mem_pi_polarCoord_target hx] + +protected theorem Complex.lintegral_comp_pi_polarCoord_symm (f : (ι → ℂ) → ℝ≥0∞) : + ∫⁻ p in (Set.univ.pi fun _ : ι ↦ Complex.polarCoord.target), + (∏ i, .ofReal (p i).1) * f (fun i ↦ Complex.polarCoord.symm (p i)) = ∫⁻ p, f p := by + let e := MeasurableEquiv.piCongrRight (fun _ : ι ↦ measurableEquivRealProd.symm) + have := volume_preserving_pi (fun _ : ι ↦ Complex.volume_preserving_equiv_real_prod.symm) + rw [← MeasurePreserving.lintegral_comp_emb this e.measurableEmbedding] + exact lintegral_comp_pi_polarCoord_symm (f ∘ e) + +end Pi diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 0407d95d6e2cc..30cb5251b049e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -42,6 +42,13 @@ theorem cpow_eq_zero_iff (x y : ℂ) : x ^ y = 0 ↔ x = 0 ∧ y ≠ 0 := by simp only [cpow_def] split_ifs <;> simp [*, exp_ne_zero] +theorem cpow_ne_zero_iff {x y : ℂ} : + x ^ y ≠ 0 ↔ x ≠ 0 ∨ y = 0 := by + rw [ne_eq, cpow_eq_zero_iff, not_and_or, ne_eq, not_not] + +theorem cpow_ne_zero_iff_of_exponent_ne_zero {x y : ℂ} (hy : y ≠ 0) : + x ^ y ≠ 0 ↔ x ≠ 0 := by simp [hy] + @[simp] theorem zero_cpow {x : ℂ} (h : x ≠ 0) : (0 : ℂ) ^ x = 0 := by simp [cpow_def, *] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 5dd15f48d3220..88657855528d1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -105,6 +105,10 @@ theorem DifferentiableAt.const_cpow (hf : DifferentiableAt ℂ f x) (h0 : c ≠ DifferentiableAt ℂ (fun x => c ^ f x) x := (hf.hasFDerivAt.const_cpow h0).differentiableAt +theorem DifferentiableAt.cpow_const (hf : DifferentiableAt ℂ f x) (h0 : f x ∈ slitPlane) : + DifferentiableAt ℂ (fun x => f x ^ c) x := + hf.cpow (differentiableAt_const c) h0 + theorem DifferentiableWithinAt.cpow (hf : DifferentiableWithinAt ℂ f s x) (hg : DifferentiableWithinAt ℂ g s x) (h0 : f x ∈ slitPlane) : DifferentiableWithinAt ℂ (fun x => f x ^ g x) s x := @@ -114,6 +118,11 @@ theorem DifferentiableWithinAt.const_cpow (hf : DifferentiableWithinAt ℂ f s x (h0 : c ≠ 0 ∨ f x ≠ 0) : DifferentiableWithinAt ℂ (fun x => c ^ f x) s x := (hf.hasFDerivWithinAt.const_cpow h0).differentiableWithinAt +theorem DifferentiableWithinAt.cpow_const (hf : DifferentiableWithinAt ℂ f s x) + (h0 : f x ∈ slitPlane) : + DifferentiableWithinAt ℂ (fun x => f x ^ c) s x := + hf.cpow (differentiableWithinAt_const c) h0 + theorem DifferentiableOn.cpow (hf : DifferentiableOn ℂ f s) (hg : DifferentiableOn ℂ g s) (h0 : Set.MapsTo f s slitPlane) : DifferentiableOn ℂ (fun x ↦ f x ^ g x) s := fun x hx ↦ (hf x hx).cpow (hg x hx) (h0 hx) @@ -122,6 +131,11 @@ theorem DifferentiableOn.const_cpow (hf : DifferentiableOn ℂ f s) (h0 : c ≠ 0 ∨ ∀ x ∈ s, f x ≠ 0) : DifferentiableOn ℂ (fun x ↦ c ^ f x) s := fun x hx ↦ (hf x hx).const_cpow (h0.imp_right fun h ↦ h x hx) +theorem DifferentiableOn.cpow_const (hf : DifferentiableOn ℂ f s) + (h0 : ∀ x ∈ s, f x ∈ slitPlane) : + DifferentiableOn ℂ (fun x => f x ^ c) s := + hf.cpow (differentiableOn_const c) h0 + theorem Differentiable.cpow (hf : Differentiable ℂ f) (hg : Differentiable ℂ g) (h0 : ∀ x, f x ∈ slitPlane) : Differentiable ℂ (fun x ↦ f x ^ g x) := fun x ↦ (hf x).cpow (hg x) (h0 x) @@ -204,8 +218,9 @@ theorem HasDerivWithinAt.cpow_const (hf : HasDerivWithinAt f f' s x) (Complex.hasStrictDerivAt_cpow_const h0).hasDerivAt.comp_hasDerivWithinAt x hf /-- Although `fun x => x ^ r` for fixed `r` is *not* complex-differentiable along the negative real -line, it is still real-differentiable, and the derivative is what one would formally expect. -/ -theorem hasDerivAt_ofReal_cpow {x : ℝ} (hx : x ≠ 0) {r : ℂ} (hr : r ≠ -1) : +line, it is still real-differentiable, and the derivative is what one would formally expect. +See `hasDerivAt_ofReal_cpow_const` for an alternate formulation. -/ +theorem hasDerivAt_ofReal_cpow_const' {x : ℝ} (hx : x ≠ 0) {r : ℂ} (hr : r ≠ -1) : HasDerivAt (fun y : ℝ => (y : ℂ) ^ (r + 1) / (r + 1)) (x ^ r) x := by rw [Ne, ← add_eq_zero_iff_eq_neg, ← Ne] at hr rcases lt_or_gt_of_ne hx.symm with (hx | hx) @@ -248,6 +263,51 @@ theorem hasDerivAt_ofReal_cpow {x : ℝ} (hx : x ≠ 0) {r : ℂ} (hr : r ≠ -1 · exact hasDerivAt_id ((-x : ℝ) : ℂ) · simp [hx] +@[deprecated (since := "2024-12-15")] alias hasDerivAt_ofReal_cpow := hasDerivAt_ofReal_cpow_const' + +/-- An alternate formulation of `hasDerivAt_ofReal_cpow_const'`. -/ +theorem hasDerivAt_ofReal_cpow_const {x : ℝ} (hx : x ≠ 0) {r : ℂ} (hr : r ≠ 0) : + HasDerivAt (fun y : ℝ => (y : ℂ) ^ r) (r * x ^ (r - 1)) x := by + have := HasDerivAt.const_mul r <| hasDerivAt_ofReal_cpow_const' hx + (by rwa [ne_eq, sub_eq_neg_self]) + simpa [sub_add_cancel, mul_div_cancel₀ _ hr] using this + +/-- A version of `DifferentiableAt.cpow_const` for a real function. -/ +theorem DifferentiableAt.ofReal_cpow_const {f : ℝ → ℝ} {x : ℝ} (hf : DifferentiableAt ℝ f x) + (h0 : f x ≠ 0) (h1 : c ≠ 0) : + DifferentiableAt ℝ (fun (y : ℝ) => (f y : ℂ) ^ c) x := + (hasDerivAt_ofReal_cpow_const h0 h1).differentiableAt.comp x hf + +theorem Complex.deriv_cpow_const (hx : x ∈ Complex.slitPlane) : + deriv (fun (x : ℂ) ↦ x ^ c) x = c * x ^ (c - 1) := + (hasStrictDerivAt_cpow_const hx).hasDerivAt.deriv + +/-- A version of `Complex.deriv_cpow_const` for a real variable. -/ +theorem Complex.deriv_ofReal_cpow_const {x : ℝ} (hx : x ≠ 0) (hc : c ≠ 0) : + deriv (fun x : ℝ ↦ (x : ℂ) ^ c) x = c * x ^ (c - 1) := + (hasDerivAt_ofReal_cpow_const hx hc).deriv + +theorem deriv_cpow_const (hf : DifferentiableAt ℂ f x) (hx : f x ∈ Complex.slitPlane) : + deriv (fun (x : ℂ) ↦ f x ^ c) x = c * f x ^ (c - 1) * deriv f x := + (hf.hasDerivAt.cpow_const hx).deriv + +theorem isTheta_deriv_ofReal_cpow_const_atTop {c : ℂ} (hc : c ≠ 0) : + deriv (fun (x : ℝ) => (x : ℂ) ^ c) =Θ[atTop] fun x => x ^ (c.re - 1) := by + calc + _ =ᶠ[atTop] fun x : ℝ ↦ c * x ^ (c - 1) := by + filter_upwards [eventually_ne_atTop 0] with x hx using by rw [deriv_ofReal_cpow_const hx hc] + _ =Θ[atTop] fun x : ℝ ↦ ‖(x : ℂ) ^ (c - 1)‖ := + (Asymptotics.IsTheta.of_norm_eventuallyEq EventuallyEq.rfl).const_mul_left hc + _ =ᶠ[atTop] fun x ↦ x ^ (c.re - 1) := by + filter_upwards [eventually_gt_atTop 0] with x hx + rw [norm_eq_abs, abs_cpow_eq_rpow_re_of_pos hx, sub_re, one_re] + +theorem isBigO_deriv_ofReal_cpow_const_atTop (c : ℂ) : + deriv (fun (x : ℝ) => (x : ℂ) ^ c) =O[atTop] fun x => x ^ (c.re - 1) := by + obtain rfl | hc := eq_or_ne c 0 + · simp_rw [cpow_zero, deriv_const', Asymptotics.isBigO_zero] + · exact (isTheta_deriv_ofReal_cpow_const_atTop hc).1 + end deriv namespace Real @@ -563,6 +623,13 @@ theorem deriv_rpow_const (hf : DifferentiableAt ℝ f x) (hx : f x ≠ 0 ∨ 1 deriv (fun x => f x ^ p) x = deriv f x * p * f x ^ (p - 1) := (hf.hasDerivAt.rpow_const hx).deriv +theorem deriv_norm_ofReal_cpow (c : ℂ) {t : ℝ} (ht : 0 < t) : + (deriv fun x : ℝ ↦ ‖(x : ℂ) ^ c‖) t = c.re * t ^ (c.re - 1) := by + rw [EventuallyEq.deriv_eq (f := fun x ↦ x ^ c.re)] + · rw [Real.deriv_rpow_const (Or.inl ht.ne')] + · filter_upwards [eventually_gt_nhds ht] with x hx + rw [Complex.norm_eq_abs, Complex.abs_cpow_eq_rpow_re_of_pos hx] + lemma isTheta_deriv_rpow_const_atTop {p : ℝ} (hp : p ≠ 0) : deriv (fun (x : ℝ) => x ^ p) =Θ[atTop] fun x => x ^ (p-1) := by calc deriv (fun (x : ℝ) => x ^ p) =ᶠ[atTop] fun x => p * x ^ (p - 1) := by diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index cbe2f010c0703..3356da74c63dd 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -963,6 +963,9 @@ theorem rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : Function.Surjective fun theorem rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : Function.Bijective fun y : ℝ≥0∞ => y ^ x := ⟨rpow_left_injective hx, rpow_left_surjective hx⟩ +lemma _root_.Real.enorm_rpow_of_nonneg {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : + ‖x ^ y‖ₑ = ‖x‖ₑ ^ y := by simp [enorm, nnnorm_rpow_of_nonneg hx, coe_rpow_of_nonneg _ hy] + end ENNReal -- Porting note(https://github.com/leanprover-community/mathlib4/issues/6038): restore diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 7781b43786177..fabc68692ced8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -318,6 +318,12 @@ theorem abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : r abs (x ^ y) = x ^ re y := by rw [abs_cpow_of_imp] <;> simp [*, arg_ofReal_of_nonneg, _root_.abs_of_nonneg] +open Filter in +lemma norm_ofReal_cpow_eventually_eq_atTop (c : ℂ) : + (fun t : ℝ ↦ ‖(t : ℂ) ^ c‖) =ᶠ[atTop] fun t ↦ t ^ c.re := by + filter_upwards [eventually_gt_atTop 0] with t ht + rw [Complex.norm_eq_abs, abs_cpow_eq_rpow_re_of_pos ht] + lemma norm_natCast_cpow_of_re_ne_zero (n : ℕ) {s : ℂ} (hs : s.re ≠ 0) : ‖(n : ℂ) ^ s‖ = (n : ℝ) ^ (s.re) := by rw [norm_eq_abs, ← ofReal_natCast, abs_cpow_eq_rpow_re_of_nonneg n.cast_nonneg hs] @@ -719,6 +725,12 @@ theorem rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ x ^ y ≤ x ^ z := rpow_le_rpow_of_exponent_ge_of_imp hx0 hx1 hyz fun _ hy ↦ le_antisymm (hyz.trans_eq hy) hz +lemma rpow_max {x y p : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) (hp : 0 ≤ p) : + (max x y) ^ p = max (x ^ p) (y ^ p) := by + rcases le_total x y with hxy | hxy + · rw [max_eq_right hxy, max_eq_right (rpow_le_rpow hx hxy hp)] + · rw [max_eq_left hxy, max_eq_left (rpow_le_rpow hy hxy hp)] + 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) @@ -956,6 +968,35 @@ lemma norm_log_natCast_le_rpow_div (n : ℕ) {ε : ℝ} (hε : 0 < ε) : ‖log end Complex +namespace Asymptotics + +open Filter + +variable {E : Type*} [SeminormedRing E] (a b c : ℝ) + +theorem IsBigO.mul_atTop_rpow_of_isBigO_rpow {f g : ℝ → E} + (hf : f =O[atTop] fun t ↦ (t : ℝ) ^ a) (hg : g =O[atTop] fun t ↦ (t : ℝ) ^ b) + (h : a + b ≤ c) : + (f * g) =O[atTop] fun t ↦ (t : ℝ) ^ c := by + refine (hf.mul hg).trans (Eventually.isBigO ?_) + filter_upwards [eventually_ge_atTop 1] with t ht + rw [← Real.rpow_add (zero_lt_one.trans_le ht), Real.norm_of_nonneg (Real.rpow_nonneg + (zero_le_one.trans ht) (a + b))] + exact Real.rpow_le_rpow_of_exponent_le ht h + +theorem IsBigO.mul_atTop_rpow_natCast_of_isBigO_rpow {f g : ℕ → E} + (hf : f =O[atTop] fun n ↦ (n : ℝ) ^ a) (hg : g =O[atTop] fun n ↦ (n : ℝ) ^ b) + (h : a + b ≤ c) : + (f * g) =O[atTop] fun n ↦ (n : ℝ) ^ c := by + refine (hf.mul hg).trans (Eventually.isBigO ?_) + filter_upwards [eventually_ge_atTop 1] with t ht + replace ht : 1 ≤ (t : ℝ) := Nat.one_le_cast.mpr ht + rw [← Real.rpow_add (zero_lt_one.trans_le ht), Real.norm_of_nonneg (Real.rpow_nonneg + (zero_le_one.trans ht) (a + b))] + exact Real.rpow_le_rpow_of_exponent_le ht h + +end Asymptotics + /-! ## Square roots of reals -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 9caf00dfe93d5..e2816cc24eaff 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -3,10 +3,11 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Benjamin Davidson -/ +import Mathlib.Algebra.Periodic import Mathlib.Algebra.QuadraticDiscriminant +import Mathlib.Algebra.Ring.NegOnePow import Mathlib.Analysis.SpecialFunctions.Exp import Mathlib.Tactic.Positivity.Core -import Mathlib.Algebra.Ring.NegOnePow /-! # Trigonometric functions @@ -115,7 +116,9 @@ theorem exists_cos_eq_zero : 0 ∈ cos '' Icc (1 : ℝ) 2 := ⟨le_of_lt cos_two_neg, le_of_lt cos_one_pos⟩ /-- The number π = 3.14159265... Defined here using choice as twice a zero of cos in [1,2], from -which one can derive all its properties. For explicit bounds on π, see `Data.Real.Pi.Bounds`. -/ +which one can derive all its properties. For explicit bounds on π, see `Data.Real.Pi.Bounds`. + +Denoted `π`, once the `Real` namespace is opened. -/ protected noncomputable def pi : ℝ := 2 * Classical.choose exists_cos_eq_zero diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean index 049879b836902..e56f20010f8cf 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Chebyshev.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric import Mathlib.Data.Complex.Module import Mathlib.RingTheory.Polynomial.Chebyshev diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index ae93fb202eb70..10783716eb430 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -15,7 +15,6 @@ import Mathlib.Data.Nat.Choose.Bounds import Mathlib.Order.Filter.AtTopBot.ModEq import Mathlib.RingTheory.Polynomial.Pochhammer import Mathlib.Tactic.NoncommRing -import Mathlib.Topology.Instances.Real.Lemmas /-! # A collection of specific limit computations diff --git a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean index 4bf4e8678ab11..ef14d27e5fae7 100644 --- a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean +++ b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean @@ -95,7 +95,6 @@ instance instSubringClass : SubringClass (VonNeumannAlgebra H) (H →L[ℂ] H) w theorem mem_carrier {S : VonNeumannAlgebra H} {x : H →L[ℂ] H} : x ∈ S.toStarSubalgebra ↔ x ∈ (S : Set (H →L[ℂ] H)) := Iff.rfl --- Porting note: changed the declaration because `simpNF` indicated the LHS simplifies to this. @[simp] theorem coe_toStarSubalgebra (S : VonNeumannAlgebra H) : diff --git a/Mathlib/CategoryTheory/Abelian/Basic.lean b/Mathlib/CategoryTheory/Abelian/Basic.lean index 34d8b479eb65e..7071943140dbb 100644 --- a/Mathlib/CategoryTheory/Abelian/Basic.lean +++ b/Mathlib/CategoryTheory/Abelian/Basic.lean @@ -242,11 +242,9 @@ attribute [local instance] OfCoimageImageComparisonIsIso.isNormalEpiCategory /-- A preadditive category with kernels, cokernels, and finite products, in which the coimage-image comparison morphism is always an isomorphism, -is an abelian category. - -The Stacks project uses this characterisation at the definition of an abelian category. -See . --/ +is an abelian category. -/ +@[stacks 0109 +"The Stacks project uses this characterisation at the definition of an abelian category."] def ofCoimageImageComparisonIsIso : Abelian C where end CategoryTheory.Abelian diff --git a/Mathlib/CategoryTheory/Abelian/Images.lean b/Mathlib/CategoryTheory/Abelian/Images.lean index 1bed70213ca90..8c3bf7213ffe1 100644 --- a/Mathlib/CategoryTheory/Abelian/Images.lean +++ b/Mathlib/CategoryTheory/Abelian/Images.lean @@ -85,10 +85,8 @@ end Coimage In any abelian category this is an isomorphism. Conversely, any additive category with kernels and cokernels and -in which this is always an isomorphism, is abelian. - -See --/ +in which this is always an isomorphism, is abelian. -/ +@[stacks 0107] def coimageImageComparison : Abelian.coimage f ⟶ Abelian.image f := cokernel.desc (kernel.ι f) (kernel.lift (cokernel.π f) f (by simp)) (by ext; simp) @@ -106,4 +104,12 @@ theorem coimageImageComparison_eq_coimageImageComparison' : theorem coimage_image_factorisation : coimage.π f ≫ coimageImageComparison f ≫ image.ι f = f := by simp [coimageImageComparison] +/-- The coimage-image comparison morphism is functorial. -/ +@[simps!] +def coimageImageComparisonFunctor : Arrow C ⥤ Arrow C where + obj f := Arrow.mk (coimageImageComparison f.hom) + map {f g} η := Arrow.homMk + (cokernel.map _ _ (kernel.map _ _ η.left η.right (by simp)) η.left (by simp)) + (kernel.map _ _ η.right (cokernel.map _ _ η.left η.right (by simp)) (by simp)) (by aesop_cat) + end CategoryTheory.Abelian diff --git a/Mathlib/CategoryTheory/Abelian/ProjectiveDimension.lean b/Mathlib/CategoryTheory/Abelian/ProjectiveDimension.lean new file mode 100644 index 0000000000000..d9a81801da948 --- /dev/null +++ b/Mathlib/CategoryTheory/Abelian/ProjectiveDimension.lean @@ -0,0 +1,189 @@ +/- +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.Homology.DerivedCategory.Ext.ExactSequences + +/-! +# Projective dimension + +In an abelian category `C`, we shall say that `X : C` has projective dimension `< n` +if all `Ext X Y i` vanish when `n ≤ i`. This defines a type class +`HasProjectiveDimensionLT X n`. We also define a type class +`HasProjectiveDimensionLE X n` as an abbreviation for +`HasProjectiveDimensionLT X (n + 1)`. +(Note that the fact that `X` is a zero object is equivalent to the condition +`HasProjectiveDimensionLT X 0`, but this cannot be expressed in terms of +`HasProjectiveDimensionLE`.) + +-/ + +universe w v u + +namespace CategoryTheory + +open Abelian Limits ZeroObject + +variable {C : Type u} [Category.{v} C] [Abelian C] + +/-- An object `X` in an abelian category has projective dimension `< n` if +all `Ext X Y i` vanish when `n ≤ i`. See also `HasProjectiveDimensionLE`. +(Do not use the `subsingleton'` field directly. Use the constructor +`HasProjectiveDimensionLT.mk`, and the lemmas `hasProjectiveDimensionLT_iff` and +`Ext.eq_zero_of_hasProjectiveDimensionLT`.) -/ +class HasProjectiveDimensionLT (X : C) (n : ℕ) : Prop where mk' :: + subsingleton' (i : ℕ) (hi : n ≤ i) ⦃Y : C⦄ : + letI := HasExt.standard C + Subsingleton (Ext.{max u v} X Y i) + +/-- An object `X` in an abelian category has projective dimension `≤ n` if +all `Ext X Y i` vanish when `n + 1 ≤ i` -/ +abbrev HasProjectiveDimensionLE (X : C) (n : ℕ) : Prop := + HasProjectiveDimensionLT X (n + 1) + +namespace HasProjectiveDimensionLT + +variable [HasExt.{w} C] (X : C) (n : ℕ) + +lemma subsingleton [hX : HasProjectiveDimensionLT X n] (i : ℕ) (hi : n ≤ i) (Y : C) : + Subsingleton (Ext.{w} X Y i) := by + letI := HasExt.standard C + have := hX.subsingleton' i hi + exact Ext.chgUniv.{w, max u v}.symm.subsingleton + +variable {X n} in +lemma mk (hX : ∀ (i : ℕ) (_ : n ≤ i) ⦃Y : C⦄, ∀ (e : Ext X Y i), e = 0) : + HasProjectiveDimensionLT X n where + subsingleton' i hi Y := by + have : Subsingleton (Ext X Y i) := ⟨fun e₁ e₂ ↦ by simp only [hX i hi]⟩ + letI := HasExt.standard C + exact Ext.chgUniv.{max u v, w}.symm.subsingleton + +end HasProjectiveDimensionLT + +lemma Abelian.Ext.eq_zero_of_hasProjectiveDimensionLT [HasExt.{w} C] + {X Y : C} {i : ℕ} (e : Ext X Y i) (n : ℕ) [HasProjectiveDimensionLT X n] + (hi : n ≤ i) : e = 0 := + (HasProjectiveDimensionLT.subsingleton X n i hi Y).elim _ _ + +section + +variable (X : C) (n : ℕ) + +lemma hasProjectiveDimensionLT_iff [HasExt.{w} C] : + HasProjectiveDimensionLT X n ↔ + ∀ (i : ℕ) (_ : n ≤ i) ⦃Y : C⦄, ∀ (e : Ext X Y i), e = 0 := + ⟨fun _ _ hi _ e ↦ e.eq_zero_of_hasProjectiveDimensionLT n hi, + HasProjectiveDimensionLT.mk⟩ + +variable {X} in +lemma Limits.IsZero.hasProjectiveDimensionLT_zero (hX : IsZero X) : + HasProjectiveDimensionLT X 0 := by + letI := HasExt.standard C + rw [hasProjectiveDimensionLT_iff] + intro i hi Y e + rw [← e.mk₀_id_comp, hX.eq_of_src (𝟙 X) 0, Ext.mk₀_zero, Ext.zero_comp] + +instance : HasProjectiveDimensionLT (0 : C) 0 := + (isZero_zero C).hasProjectiveDimensionLT_zero + +lemma hasProjectiveDimensionLT_of_ge (m : ℕ) (h : n ≤ m) + [HasProjectiveDimensionLT X n] : + HasProjectiveDimensionLT X m := by + letI := HasExt.standard C + rw [hasProjectiveDimensionLT_iff] + intro i hi Y e + exact e.eq_zero_of_hasProjectiveDimensionLT n (by omega) + +instance [HasProjectiveDimensionLT X n] (k : ℕ) : + HasProjectiveDimensionLT X (n + k) := + hasProjectiveDimensionLT_of_ge X n (n + k) (by omega) + +instance [HasProjectiveDimensionLT X n] (k : ℕ) : + HasProjectiveDimensionLT X (k + n) := + hasProjectiveDimensionLT_of_ge X n (k + n) (by omega) + +instance [HasProjectiveDimensionLT X n] : + HasProjectiveDimensionLT X n.succ := + inferInstanceAs (HasProjectiveDimensionLT X (n + 1)) + +end + +lemma Retract.hasProjectiveDimensionLT {X Y : C} (h : Retract X Y) (n : ℕ) + [HasProjectiveDimensionLT Y n] : + HasProjectiveDimensionLT X n := by + letI := HasExt.standard C + rw [hasProjectiveDimensionLT_iff] + intro i hi T x + rw [← x.mk₀_id_comp, ← h.retract, ← Ext.mk₀_comp_mk₀, + Ext.comp_assoc_of_second_deg_zero, + ((Ext.mk₀ h.r).comp x (zero_add i)).eq_zero_of_hasProjectiveDimensionLT n hi, + Ext.comp_zero] + +lemma hasProjectiveDimensionLT_of_iso {X X' : C} (e : X ≅ X') (n : ℕ) + [HasProjectiveDimensionLT X n] : + HasProjectiveDimensionLT X' n := + e.symm.retract.hasProjectiveDimensionLT n + +namespace ShortComplex + +namespace ShortExact + +variable {S : ShortComplex C} (hS : S.ShortExact) (n : ℕ) +include hS + +-- In the following lemmas, the parameters `HasProjectiveDimensionLT` are +-- explicit as it is unlikely we may infer them, unless the short complex `S` +-- was declared reducible + +lemma hasProjectiveDimensionLT_X₂ (h₁ : HasProjectiveDimensionLT S.X₁ n) + (h₃ : HasProjectiveDimensionLT S.X₃ n) : + HasProjectiveDimensionLT S.X₂ n := by + letI := HasExt.standard C + rw [hasProjectiveDimensionLT_iff] + intro i hi Y x₂ + obtain ⟨x₃, rfl⟩ := Ext.contravariant_sequence_exact₂ hS _ x₂ + (Ext.eq_zero_of_hasProjectiveDimensionLT _ n hi) + rw [x₃.eq_zero_of_hasProjectiveDimensionLT n hi, Ext.comp_zero] + +lemma hasProjectiveDimensionLT_X₃ (h₁ : HasProjectiveDimensionLT S.X₁ n) + (h₂ : HasProjectiveDimensionLT S.X₂ (n + 1)) : + HasProjectiveDimensionLT S.X₃ (n + 1) := by + letI := HasExt.standard C + rw [hasProjectiveDimensionLT_iff] + rintro (_ | i) hi Y x₃ + · simp at hi + · obtain ⟨x₁, rfl⟩ := Ext.contravariant_sequence_exact₃ hS _ x₃ + (Ext.eq_zero_of_hasProjectiveDimensionLT _ (n + 1) hi) (add_comm _ _) + rw [x₁.eq_zero_of_hasProjectiveDimensionLT n (by omega), Ext.comp_zero] + +lemma hasProjectiveDimensionLT_X₁ (h₂ : HasProjectiveDimensionLT S.X₂ n) + (h₃ : HasProjectiveDimensionLT S.X₃ (n + 1)) : + HasProjectiveDimensionLT S.X₁ n := by + letI := HasExt.standard C + rw [hasProjectiveDimensionLT_iff] + intro i hi Y x₁ + obtain ⟨x₂, rfl⟩ := Ext.contravariant_sequence_exact₁ hS _ x₁ (add_comm _ _) + (Ext.eq_zero_of_hasProjectiveDimensionLT _ (n + 1) (by omega)) + rw [x₂.eq_zero_of_hasProjectiveDimensionLT n (by omega), Ext.comp_zero] + +-- When we know `HasProjectiveDimensionLT S.X₂ 1` is equivalent to `Projective S.X₂`, +-- the assumption `h₂` can be changed to `h₂ : Projective S.X₂`. +lemma hasProjectiveDimensionLT_X₃_iff (n : ℕ) (h₂ : HasProjectiveDimensionLT S.X₂ 1) : + HasProjectiveDimensionLT S.X₃ (n + 2) ↔ HasProjectiveDimensionLT S.X₁ (n + 1) := + ⟨fun _ ↦ hS.hasProjectiveDimensionLT_X₁ (n + 1) inferInstance inferInstance, + fun _ ↦ hS.hasProjectiveDimensionLT_X₃ (n + 1) inferInstance inferInstance⟩ + +end ShortExact + +end ShortComplex + +instance (X Y : C) (n : ℕ) [HasProjectiveDimensionLT X n] + [HasProjectiveDimensionLT Y n] : + HasProjectiveDimensionLT (X ⊞ Y) n := + (ShortComplex.Splitting.ofHasBinaryBiproduct X Y).shortExact.hasProjectiveDimensionLT_X₂ n + (by assumption) (by assumption) + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Abelian/Transfer.lean b/Mathlib/CategoryTheory/Abelian/Transfer.lean index 61b48e84e435b..20ca075bae757 100644 --- a/Mathlib/CategoryTheory/Abelian/Transfer.lean +++ b/Mathlib/CategoryTheory/Abelian/Transfer.lean @@ -154,10 +154,8 @@ open AbelianOfAdjunction we have `F : C ⥤ D` `G : D ⥤ C` (both preserving zero morphisms), `G` is left exact (that is, preserves finite limits), and further we have `adj : G ⊣ F` and `i : F ⋙ G ≅ 𝟭 C`, -then `C` is also abelian. - -See --/ +then `C` is also abelian. -/ +@[stacks 03A3] def abelianOfAdjunction {C : Type u₁} [Category.{v₁} C] [Preadditive C] [HasFiniteProducts C] {D : Type u₂} [Category.{v₂} D] [Abelian D] (F : C ⥤ D) [Functor.PreservesZeroMorphisms F] (G : D ⥤ C) [Functor.PreservesZeroMorphisms G] [PreservesFiniteLimits G] (i : F ⋙ G ≅ 𝟭 C) diff --git a/Mathlib/CategoryTheory/Action/Monoidal.lean b/Mathlib/CategoryTheory/Action/Monoidal.lean index cc31bc6edd902..8bb4bce035d12 100644 --- a/Mathlib/CategoryTheory/Action/Monoidal.lean +++ b/Mathlib/CategoryTheory/Action/Monoidal.lean @@ -183,13 +183,9 @@ theorem rightDual_v [RightRigidCategory V] : Xᘁ.V = X.Vᘁ := theorem leftDual_v [LeftRigidCategory V] : (ᘁX).V = ᘁX.V := rfl --- This lemma was always bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644 -@[simp, nolint simpNF] theorem rightDual_ρ [RightRigidCategory V] (h : H) : Xᘁ.ρ h = (X.ρ (h⁻¹ : H))ᘁ := by rw [← SingleObj.inv_as_inv]; rfl --- This lemma was always bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644 -@[simp, nolint simpNF] theorem leftDual_ρ [LeftRigidCategory V] (h : H) : (ᘁX).ρ h = ᘁX.ρ (h⁻¹ : H) := by rw [← SingleObj.inv_as_inv]; rfl diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 6c4183721f9f5..edb9845c4258c 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -97,10 +97,8 @@ There is also a constructor `Adjunction.mkOfHomEquiv` which constructs an adjunc hom set equivalence. To construct adjoints to a given functor, there are constructors `leftAdjointOfEquiv` and -`adjunctionOfEquivLeft` (as well as their duals). - -See . --/ +`adjunctionOfEquivLeft` (as well as their duals). -/ +@[stacks 0037] structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where /-- The unit of an adjunction -/ unit : 𝟭 C ⟶ F.comp G @@ -328,8 +326,6 @@ structure CoreHomEquivUnitCounit (F : C ⥤ D) (G : D ⥤ C) where See `Adjunction.mkOfHomEquiv`. This structure won't typically be used anywhere else. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): `has_nonempty_instance` linter not ported yet --- @[nolint has_nonempty_instance] structure CoreHomEquiv (F : C ⥤ D) (G : D ⥤ C) where /-- The equivalence between `Hom (F X) Y` and `Hom X (G Y)` -/ homEquiv : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G.obj Y) @@ -364,8 +360,6 @@ end CoreHomEquiv See `Adjunction.mkOfUnitCounit`. This structure won't typically be used anywhere else. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): `has_nonempty_instance` linter not ported yet --- @[nolint has_nonempty_instance] structure CoreUnitCounit (F : C ⥤ D) (G : D ⥤ C) where /-- The unit of an adjunction between `F` and `G` -/ unit : 𝟭 C ⟶ F.comp G @@ -512,11 +506,8 @@ section variable {E : Type u₃} [ℰ : Category.{v₃} E] {H : D ⥤ E} {I : E ⥤ D} (adj₁ : F ⊣ G) (adj₂ : H ⊣ I) -/-- Composition of adjunctions. - -See . --/ -@[simps! (config := .lemmasOnly) unit counit] +/-- Composition of adjunctions. -/ +@[simps! (config := .lemmasOnly) unit counit, stacks 0DV0] def comp : F ⋙ H ⊣ I ⋙ G := mk' { homEquiv := fun _ _ ↦ Equiv.trans (adj₂.homEquiv _ _) (adj₁.homEquiv _ _) diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index 2db2c795ef2c2..2706ad1af5434 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -79,10 +79,8 @@ def functorialityAdjunction : Cocones.functoriality K F ⊣ functorialityRightAd counit := functorialityCounit adj K include adj in -/-- A left adjoint preserves colimits. - -See . --/ +/-- A left adjoint preserves colimits. -/ +@[stacks 0038] lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where preservesColimitsOfShape := { preservesColimit := @@ -203,10 +201,8 @@ def functorialityAdjunction' : functorialityLeftAdjoint adj K ⊣ Cones.functori counit := functorialityCounit' adj K include adj in -/-- A right adjoint preserves limits. - -See . --/ +/-- A right adjoint preserves limits. -/ +@[stacks 0038] lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where preservesLimitsOfShape := { preservesLimit := diff --git a/Mathlib/CategoryTheory/Bicategory/Adjunction.lean b/Mathlib/CategoryTheory/Bicategory/Adjunction/Basic.lean similarity index 100% rename from Mathlib/CategoryTheory/Bicategory/Adjunction.lean rename to Mathlib/CategoryTheory/Bicategory/Adjunction/Basic.lean diff --git a/Mathlib/CategoryTheory/Bicategory/Adjunction/Mate.lean b/Mathlib/CategoryTheory/Bicategory/Adjunction/Mate.lean new file mode 100644 index 0000000000000..1e93d955a592b --- /dev/null +++ b/Mathlib/CategoryTheory/Bicategory/Adjunction/Mate.lean @@ -0,0 +1,219 @@ +/- +Copyright (c) 2025 Yuma Mizuno. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuma Mizuno +-/ +import Mathlib.CategoryTheory.Bicategory.Adjunction.Basic + +/-! +# Mates in bicategories + +This file establishes the bijection between the 2-cells + +``` + l₁ r₁ + c --→ d c ←-- d + g ↓ ↗ ↓ h g ↓ ↘ ↓ h + e --→ f e ←-- f + l₂ r₂ +``` + +where `l₁ ⊣ r₁` and `l₂ ⊣ r₂`. The corresponding natural transformations are called mates. + +For the bicategory `Cat`, the definitions in this file are provided in +`Mathlib/CategoryTheory/Adjunction/Mates.lean`, where you can find more detailed documentation +about mates. + +## Remarks + +To be precise, the definitions in `Mathlib/CategoryTheory/Adjunction/Mates.lean` are universe +polymorphic, so they are not simple specializations of the definitions in this file. + +-/ + +universe w v u + +namespace CategoryTheory + +namespace Bicategory + +open Bicategory + +variable {B : Type u} [Bicategory.{w, v} B] + +section mateEquiv + +variable {c d e f : B} {g : c ⟶ e} {h : d ⟶ f} {l₁ : c ⟶ d} {r₁ : d ⟶ c} {l₂ : e ⟶ f} {r₂ : f ⟶ e} +variable (adj₁ : l₁ ⊣ r₁) (adj₂ : l₂ ⊣ r₂) + +/-- Suppose we have a square of 1-morphisms (where the top and bottom are adjunctions `l₁ ⊣ r₁` +and `l₂ ⊣ r₂` respectively). + + c ↔ d + g ↓ ↓ h + e ↔ f + +Then we have a bijection between natural transformations `g ≫ l₂ ⟶ l₁ ≫ h` and +`r₁ ≫ g ⟶ h ≫ r₂`. This can be seen as a bijection of the 2-cells: + + l₁ r₁ + c --→ d c ←-- d + g ↓ ↗ ↓ h g ↓ ↘ ↓ h + e --→ f e ←-- f + L₂ R₂ + +Note that if one of the transformations is an iso, it does not imply the other is an iso. +-/ +@[simps] +def mateEquiv : (g ≫ l₂ ⟶ l₁ ≫ h) ≃ (r₁ ≫ g ⟶ h ≫ r₂) where + toFun α := 𝟙 _ ⊗≫ r₁ ◁ g ◁ adj₂.unit ⊗≫ r₁ ◁ α ▷ r₂ ⊗≫ adj₁.counit ▷ h ▷ r₂ ⊗≫ 𝟙 _ + invFun β := 𝟙 _ ⊗≫ adj₁.unit ▷ g ▷ l₂ ⊗≫ l₁ ◁ β ▷ l₂ ⊗≫ l₁ ◁ h ◁ adj₂.counit ⊗≫ 𝟙 _ + left_inv α := + calc + _ = 𝟙 _ ⊗≫ (adj₁.unit ▷ (g ≫ 𝟙 e) ≫ (l₁ ≫ r₁) ◁ g ◁ adj₂.unit) ▷ l₂ ⊗≫ + l₁ ◁ r₁ ◁ α ▷ r₂ ▷ l₂ ⊗≫ + l₁ ◁ (adj₁.counit ▷ h ▷ (r₂ ≫ l₂) ≫ (𝟙 d ≫ h) ◁ adj₂.counit) ⊗≫ 𝟙 _ := by + bicategory + _ = 𝟙 _ ⊗≫ g ◁ adj₂.unit ▷ l₂ ⊗≫ + (adj₁.unit ▷ (g ≫ l₂) ≫ (l₁ ≫ r₁) ◁ α) ▷ (r₂ ≫ l₂) ⊗≫ + l₁ ◁ (((r₁ ≫ l₁) ≫ h) ◁ adj₂.counit ≫ adj₁.counit ▷ h ▷ 𝟙 f) ⊗≫ 𝟙 _ := by + rw [← whisker_exchange, ← whisker_exchange] + bicategory + _ = 𝟙 _ ⊗≫ g ◁ adj₂.unit ▷ l₂ ⊗≫ α ▷ r₂ ▷ l₂ ⊗≫ + leftZigzag adj₁.unit adj₁.counit ▷ h ▷ r₂ ▷ l₂ ⊗≫ l₁ ◁ h ◁ adj₂.counit ⊗≫ 𝟙 _ := by + rw [← whisker_exchange, whisker_exchange _ adj₂.counit] + bicategory + _ = 𝟙 _ ⊗≫ g ◁ adj₂.unit ▷ l₂ ⊗≫ (α ▷ (r₂ ≫ l₂) ≫ (l₁ ≫ h) ◁ adj₂.counit) ⊗≫ 𝟙 _ := by + rw [adj₁.left_triangle] + bicategory + _ = 𝟙 _ ⊗≫ g ◁ (leftZigzag adj₂.unit adj₂.counit) ⊗≫ α ⊗≫ 𝟙 _ := by + rw [← whisker_exchange] + bicategory + _ = α := by + rw [adj₂.left_triangle] + bicategory + right_inv β := + calc + _ = 𝟙 _ ⊗≫ r₁ ◁ ((𝟙 c ≫ g) ◁ adj₂.unit ≫ adj₁.unit ▷ g ▷ (l₂ ≫ r₂)) ⊗≫ + r₁ ◁ l₁ ◁ β ▷ l₂ ▷ r₂ ⊗≫ + ((r₁ ≫ l₁) ◁ h ◁ adj₂.counit ≫ adj₁.counit ▷ (h ≫ 𝟙 f)) ▷ r₂ ⊗≫ 𝟙 _ := by + bicategory + _ = 𝟙 _ ⊗≫ r₁ ◁ (adj₁.unit ▷ g ▷ 𝟙 e ≫ ((l₁ ≫ r₁) ≫ g) ◁ adj₂.unit) ⊗≫ + ((r₁ ≫ l₁) ◁ β ≫ adj₁.counit ▷ (h ≫ r₂)) ▷ l₂ ▷ r₂ ⊗≫ + h ◁ adj₂.counit ▷ r₂ ⊗≫ 𝟙 _ := by + rw [whisker_exchange, whisker_exchange] + bicategory + _ = 𝟙 _ ⊗≫ r₁ ◁ g ◁ adj₂.unit ⊗≫ rightZigzag adj₁.unit adj₁.counit ▷ g ▷ l₂ ▷ r₂ ⊗≫ + β ▷ l₂ ▷ r₂ ⊗≫ h ◁ adj₂.counit ▷ r₂ ⊗≫ 𝟙 _ := by + rw [whisker_exchange, ← whisker_exchange _ adj₂.unit] + bicategory + _ = 𝟙 _ ⊗≫ ((r₁ ≫ g) ◁ adj₂.unit ≫ β ▷ (l₂ ≫ r₂)) ⊗≫ h ◁ adj₂.counit ▷ r₂ ⊗≫ 𝟙 _ := by + rw [adj₁.right_triangle] + bicategory + _ = 𝟙 _ ⊗≫ β ⊗≫ h ◁ rightZigzag adj₂.unit adj₂.counit ⊗≫ 𝟙 _ := by + rw [whisker_exchange] + bicategory + _ = β := by + rw [adj₂.right_triangle] + bicategory + +end mateEquiv + +section mateEquivVComp + +variable {a b c d e f : B} {g₁ : a ⟶ c} {g₂ : c ⟶ e} {h₁ : b ⟶ d} {h₂ : d ⟶ f} +variable {l₁ : a ⟶ b} {r₁ : b ⟶ a} {l₂ : c ⟶ d} {r₂ : d ⟶ c} {l₃ : e ⟶ f} {r₃ : f ⟶ e} +variable (adj₁ : l₁ ⊣ r₁) (adj₂ : l₂ ⊣ r₂) (adj₃ : l₃ ⊣ r₃) + +/-- Squares between left adjoints can be composed "vertically" by pasting. -/ +def leftAdjointSquare.vcomp (α : g₁ ≫ l₂ ⟶ l₁ ≫ h₁) (β : g₂ ≫ l₃ ⟶ l₂ ≫ h₂) : + (g₁ ≫ g₂) ≫ l₃ ⟶ l₁ ≫ (h₁ ≫ h₂) := + (α_ _ _ _).hom ≫ g₁ ◁ β ≫ (α_ _ _ _).inv ≫ α ▷ h₂ ≫ (α_ _ _ _).hom + +/-- Squares between right adjoints can be composed "vertically" by pasting. -/ +def rightAdjointSquare.vcomp (α : r₁ ≫ g₁ ⟶ h₁ ≫ r₂) (β : r₂ ≫ g₂ ⟶ h₂ ≫ r₃) : + r₁ ≫ (g₁ ≫ g₂) ⟶ (h₁ ≫ h₂) ≫ r₃ := + (α_ _ _ _).inv ≫ α ▷ g₂ ≫ (α_ _ _ _).hom ≫ h₁ ◁ β ≫ (α_ _ _ _).inv + +/-- The mates equivalence commutes with vertical composition. -/ +theorem mateEquiv_vcomp (α : g₁ ≫ l₂ ⟶ l₁ ≫ h₁) (β : g₂ ≫ l₃ ⟶ l₂ ≫ h₂) : + mateEquiv adj₁ adj₃ (leftAdjointSquare.vcomp α β) = + rightAdjointSquare.vcomp (mateEquiv adj₁ adj₂ α) (mateEquiv adj₂ adj₃ β) := by + dsimp only [leftAdjointSquare.vcomp, mateEquiv_apply, rightAdjointSquare.vcomp] + symm + calc + _ = 𝟙 _ ⊗≫ r₁ ◁ g₁ ◁ adj₂.unit ▷ g₂ ⊗≫ r₁ ◁ α ▷ r₂ ▷ g₂ ⊗≫ + ((adj₁.counit ▷ (h₁ ≫ r₂ ≫ g₂ ≫ 𝟙 e)) ≫ 𝟙 b ◁ (h₁ ◁ r₂ ◁ g₂ ◁ adj₃.unit)) ⊗≫ + h₁ ◁ r₂ ◁ β ▷ r₃ ⊗≫ h₁ ◁ adj₂.counit ▷ h₂ ▷ r₃ ⊗≫ 𝟙 _ := by + bicategory + _ = 𝟙 _ ⊗≫ r₁ ◁ g₁ ◁ adj₂.unit ▷ g₂ ⊗≫ + (r₁ ◁ (α ▷ (r₂ ≫ g₂ ≫ 𝟙 e) ≫ (l₁ ≫ h₁) ◁ r₂ ◁ g₂ ◁ adj₃.unit)) ⊗≫ + ((adj₁.counit ▷ (h₁ ≫ r₂) ▷ (g₂ ≫ l₃) ≫ (𝟙 b ≫ h₁ ≫ r₂) ◁ β) ▷ r₃) ⊗≫ + h₁ ◁ adj₂.counit ▷ h₂ ▷ r₃ ⊗≫ 𝟙 _ := by + rw [← whisker_exchange] + bicategory + _ = 𝟙 _ ⊗≫ r₁ ◁ g₁ ◁ (adj₂.unit ▷ (g₂ ≫ 𝟙 e) ≫ (l₂ ≫ r₂) ◁ g₂ ◁ adj₃.unit) ⊗≫ + (r₁ ◁ (α ▷ (r₂ ≫ g₂ ≫ l₃) ≫ (l₁ ≫ h₁) ◁ r₂ ◁ β) ▷ r₃) ⊗≫ + (adj₁.counit ▷ h₁ ▷ (r₂ ≫ l₂) ≫ (𝟙 b ≫ h₁) ◁ adj₂.counit) ▷ h₂ ▷ r₃ ⊗≫ 𝟙 _ := by + rw [← whisker_exchange, ← whisker_exchange] + bicategory + _ = 𝟙 _ ⊗≫ r₁ ◁ g₁ ◁ g₂ ◁ adj₃.unit ⊗≫ + r₁ ◁ g₁ ◁ (adj₂.unit ▷ (g₂ ≫ l₃) ≫ (l₂ ≫ r₂) ◁ β) ▷ r₃ ⊗≫ + r₁ ◁ (α ▷ (r₂ ≫ l₂) ≫ (l₁ ≫ h₁) ◁ adj₂.counit) ▷ h₂ ▷ r₃ ⊗≫ + adj₁.counit ▷ h₁ ▷ h₂ ▷ r₃ ⊗≫ 𝟙 _ := by + rw [← whisker_exchange, ← whisker_exchange, ← whisker_exchange] + bicategory + _ = 𝟙 _ ⊗≫ r₁ ◁ g₁ ◁ g₂ ◁ adj₃.unit ⊗≫ r₁ ◁ g₁ ◁ β ▷ r₃ ⊗≫ + ((r₁ ≫ g₁) ◁ leftZigzag adj₂.unit adj₂.counit ▷ (h₂ ≫ r₃)) ⊗≫ + r₁ ◁ α ▷ h₂ ▷ r₃ ⊗≫ adj₁.counit ▷ h₁ ▷ h₂ ▷ r₃ ⊗≫ 𝟙 _ := by + rw [← whisker_exchange, ← whisker_exchange] + bicategory + _ = _ := by + rw [adj₂.left_triangle] + bicategory + +end mateEquivVComp + +section mateEquivHComp + +variable {a : B} {b : B} {c : B} {d : B} {e : B} {f : B} +variable {g : a ⟶ d} {h : b ⟶ e} {k : c ⟶ f} +variable {l₁ : a ⟶ b} {r₁ : b ⟶ a} {l₂ : d ⟶ e} {r₂ : e ⟶ d} +variable {l₃ : b ⟶ c} {r₃ : c ⟶ b} {l₄ : e ⟶ f} {r₄ : f ⟶ e} +variable (adj₁ : l₁ ⊣ r₁) (adj₂ : l₂ ⊣ r₂) (adj₃ : l₃ ⊣ r₃) (adj₄ : l₄ ⊣ r₄) + +/-- Squares between left adjoints can be composed "horizontally" by pasting. -/ +def leftAdjointSquare.hcomp (α : g ≫ l₂ ⟶ l₁ ≫ h) (β : h ≫ l₄ ⟶ l₃ ≫ k) : + g ≫ (l₂ ≫ l₄) ⟶ (l₁ ≫ l₃) ≫ k := + (α_ _ _ _).inv ≫ α ▷ l₄ ≫ (α_ _ _ _).hom ≫ l₁ ◁ β ≫ (α_ _ _ _).inv + +/-- Squares between right adjoints can be composed "horizontally" by pasting. -/ +def rightAdjointSquare.hcomp (α : r₁ ≫ g ⟶ h ≫ r₂) (β : r₃ ≫ h ⟶ k ≫ r₄) : + (r₃ ≫ r₁) ≫ g ⟶ k ≫ (r₄ ≫ r₂) := + (α_ _ _ _).hom ≫ r₃ ◁ α ≫ (α_ _ _ _).inv ≫ β ▷ r₂ ≫ (α_ _ _ _).hom + +/-- The mates equivalence commutes with horizontal composition of squares. -/ +theorem mateEquiv_hcomp (α : g ≫ l₂ ⟶ l₁ ≫ h) (β : h ≫ l₄ ⟶ l₃ ≫ k) : + (mateEquiv (adj₁.comp adj₃) (adj₂.comp adj₄)) (leftAdjointSquare.hcomp α β) = + rightAdjointSquare.hcomp (mateEquiv adj₁ adj₂ α) (mateEquiv adj₃ adj₄ β) := by + dsimp [mateEquiv, leftAdjointSquare.hcomp, rightAdjointSquare.hcomp] + calc + _ = 𝟙 _ ⊗≫ r₃ ◁ r₁ ◁ g ◁ adj₂.unit ⊗≫ + r₃ ◁ r₁ ◁ ((g ≫ l₂) ◁ adj₄.unit ≫ α ▷ (l₄ ≫ r₄)) ▷ r₂ ⊗≫ + r₃ ◁ ((r₁ ≫ l₁) ◁ β ≫ adj₁.counit ▷ (l₃ ≫ k)) ▷ r₄ ▷ r₂ ⊗≫ + adj₃.counit ▷ k ▷ r₄ ▷ r₂ ⊗≫ 𝟙 _ := by + bicategory + _ = 𝟙 _ ⊗≫ r₃ ◁ r₁ ◁ g ◁ adj₂.unit ⊗≫ r₃ ◁ r₁ ◁ α ▷ r₂ ⊗≫ + r₃ ◁ ((r₁ ≫ l₁) ◁ h ◁ adj₄.unit ≫ adj₁.counit ▷ (h ≫ l₄ ≫ r₄)) ▷ r₂ ⊗≫ + r₃ ◁ β ▷ r₄ ▷ r₂ ⊗≫ adj₃.counit ▷ k ▷ r₄ ▷ r₂ ⊗≫ 𝟙 _ := by + rw [whisker_exchange, whisker_exchange] + bicategory + _ = _ := by + rw [whisker_exchange] + bicategory + +end mateEquivHComp + +end Bicategory + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Bicategory/Free.lean b/Mathlib/CategoryTheory/Bicategory/Free.lean index 08cee6ecf1d42..4029841250cc9 100644 --- a/Mathlib/CategoryTheory/Bicategory/Free.lean +++ b/Mathlib/CategoryTheory/Bicategory/Free.lean @@ -60,8 +60,6 @@ instance categoryStruct : CategoryStruct.{max u v} (FreeBicategory B) where comp := @fun _ _ _ => Hom.comp /-- Representatives of 2-morphisms in the free bicategory. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] inductive Hom₂ : ∀ {a b : FreeBicategory B}, (a ⟶ b) → (a ⟶ b) → Type max u v | id {a b} (f : a ⟶ b) : Hom₂ f f | vcomp {a b} {f g h : a ⟶ b} (η : Hom₂ f g) (θ : Hom₂ g h) : Hom₂ f h diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean index b80c8b92adcd7..e96f46366f895 100644 --- a/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean @@ -154,8 +154,6 @@ def comp (F : OplaxFunctor B C) (G : OplaxFunctor C D) : OplaxFunctor B D where /-- A structure on an oplax functor that promotes an oplax functor to a pseudofunctor. See `Pseudofunctor.mkOfOplax`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] structure PseudoCore (F : OplaxFunctor B C) where /-- The isomorphism giving rise to the oplax unity constraint -/ mapIdIso (a : B) : F.map (𝟙 a) ≅ 𝟙 (F.obj a) diff --git a/Mathlib/CategoryTheory/Bicategory/Kan/Adjunction.lean b/Mathlib/CategoryTheory/Bicategory/Kan/Adjunction.lean index ac3e11be6fb37..99316716f3e62 100644 --- a/Mathlib/CategoryTheory/Bicategory/Kan/Adjunction.lean +++ b/Mathlib/CategoryTheory/Bicategory/Kan/Adjunction.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ import Mathlib.CategoryTheory.Bicategory.Kan.HasKan -import Mathlib.CategoryTheory.Bicategory.Adjunction +import Mathlib.CategoryTheory.Bicategory.Adjunction.Basic import Mathlib.Tactic.TFAE /-! diff --git a/Mathlib/CategoryTheory/Category/Basic.lean b/Mathlib/CategoryTheory/Category/Basic.lean index d314808670795..247f3da889691 100644 --- a/Mathlib/CategoryTheory/Category/Basic.lean +++ b/Mathlib/CategoryTheory/Category/Basic.lean @@ -23,7 +23,7 @@ Introduces notations in the `CategoryTheory` scope Users may like to add `g ⊚ f` for composition in the standard convention, using ```lean -local notation g ` ⊚ `:80 f:80 := category.comp f g -- type as \oo +local notation:80 g " ⊚ " f:80 => CategoryTheory.CategoryStruct.comp f g -- type as \oo ``` ## Porting note @@ -142,11 +142,8 @@ attribute [aesop safe (rule_sets := [CategoryTheory])] Subsingleton.elim /-- The typeclass `Category C` describes morphisms associated to objects of type `C`. The universe levels of the objects and morphisms are unconstrained, and will often need to be -specified explicitly, as `Category.{v} C`. (See also `LargeCategory` and `SmallCategory`.) - -See . --/ -@[pp_with_univ] +specified explicitly, as `Category.{v} C`. (See also `LargeCategory` and `SmallCategory`.) -/ +@[pp_with_univ, stacks 0014] class Category (obj : Type u) extends CategoryStruct.{v} obj : Type max u (v + 1) where /-- Identity morphisms are left identities for composition. -/ id_comp : ∀ {X Y : obj} (f : X ⟶ Y), 𝟙 X ≫ f = f := by aesop_cat @@ -235,19 +232,15 @@ theorem dite_comp {P : Prop} [Decidable P] (if h : P then f h else f' h) ≫ g = if h : P then f h ≫ g else f' h ≫ g := by aesop /-- A morphism `f` is an epimorphism if it can be cancelled when precomposed: -`f ≫ g = f ≫ h` implies `g = h`. - -See . --/ +`f ≫ g = f ≫ h` implies `g = h`. -/ +@[stacks 003B] class Epi (f : X ⟶ Y) : Prop where /-- A morphism `f` is an epimorphism if it can be cancelled when precomposed. -/ left_cancellation : ∀ {Z : C} (g h : Y ⟶ Z), f ≫ g = f ≫ h → g = h /-- A morphism `f` is a monomorphism if it can be cancelled when postcomposed: -`g ≫ f = h ≫ f` implies `g = h`. - -See . --/ +`g ≫ f = h ≫ f` implies `g = h`. -/ +@[stacks 003B] class Mono (f : X ⟶ Y) : Prop where /-- A morphism `f` is a monomorphism if it can be cancelled when postcomposed. -/ right_cancellation : ∀ {Z : C} (g h : Z ⟶ X), g ≫ f = h ≫ f → g = h diff --git a/Mathlib/CategoryTheory/Category/GaloisConnection.lean b/Mathlib/CategoryTheory/Category/GaloisConnection.lean index a79e29a39e4c2..df642c01f6f1e 100644 --- a/Mathlib/CategoryTheory/Category/GaloisConnection.lean +++ b/Mathlib/CategoryTheory/Category/GaloisConnection.lean @@ -5,7 +5,7 @@ Authors: Stephen Morgan, Kim Morrison, Johannes Hölzl, Reid Barton -/ import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.Adjunction.Basic -import Mathlib.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Defs /-! diff --git a/Mathlib/CategoryTheory/Category/PartialFun.lean b/Mathlib/CategoryTheory/Category/PartialFun.lean index 6daedc19bcadb..c5d65f076306a 100644 --- a/Mathlib/CategoryTheory/Category/PartialFun.lean +++ b/Mathlib/CategoryTheory/Category/PartialFun.lean @@ -39,7 +39,6 @@ namespace PartialFun instance : CoeSort PartialFun Type* := ⟨id⟩ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]`; linter not ported yet /-- Turns a type into a `PartialFun`. -/ def of : Type* → PartialFun := id diff --git a/Mathlib/CategoryTheory/Category/Preorder.lean b/Mathlib/CategoryTheory/Category/Preorder.lean index f2935460a2529..fee0af705c592 100644 --- a/Mathlib/CategoryTheory/Category/Preorder.lean +++ b/Mathlib/CategoryTheory/Category/Preorder.lean @@ -39,10 +39,8 @@ The category structure coming from a preorder. There is a morphism `X ⟶ Y` if Because we don't allow morphisms to live in `Prop`, we have to define `X ⟶ Y` as `ULift (PLift (X ≤ Y))`. -See `CategoryTheory.homOfLE` and `CategoryTheory.leOfHom`. - -See . --/ +See `CategoryTheory.homOfLE` and `CategoryTheory.leOfHom`. -/ +@[stacks 00D3] instance (priority := 100) smallCategory (α : Type u) [Preorder α] : SmallCategory α where Hom U V := ULift (PLift (U ≤ V)) id X := ⟨⟨le_refl X⟩⟩ diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index c79932f744135..22eb522f83727 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -163,13 +163,9 @@ def curry : (A ⊗ Y ⟶ X) → (Y ⟶ A ⟹ X) := def uncurry : (Y ⟶ A ⟹ X) → (A ⊗ Y ⟶ X) := ((exp.adjunction A).homEquiv _ _).symm --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem homEquiv_apply_eq (f : A ⊗ Y ⟶ X) : (exp.adjunction A).homEquiv _ _ f = curry f := rfl --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem homEquiv_symm_apply_eq (f : Y ⟶ A ⟹ X) : ((exp.adjunction A).homEquiv _ _).symm f = uncurry f := rfl @@ -375,6 +371,4 @@ def cartesianClosedOfEquiv (e : C ≌ D) [CartesianClosed C] : CartesianClosed D end Functor -attribute [nolint simpNF] CategoryTheory.CartesianClosed.homEquiv_apply_eq - CategoryTheory.CartesianClosed.homEquiv_symm_apply_eq end CategoryTheory diff --git a/Mathlib/CategoryTheory/Closed/FunctorCategory/Groupoid.lean b/Mathlib/CategoryTheory/Closed/FunctorCategory/Groupoid.lean index a304a8071b6d1..16828b2e30c69 100644 --- a/Mathlib/CategoryTheory/Closed/FunctorCategory/Groupoid.lean +++ b/Mathlib/CategoryTheory/Closed/FunctorCategory/Groupoid.lean @@ -57,8 +57,6 @@ def closedCounit (F : D ⥤ C) : closedIhom F ⋙ tensorLeft F ⟶ 𝟭 (D ⥤ C /-- If `C` is a monoidal closed category and `D` is a groupoid, then every functor `F : D ⥤ C` is closed in the functor category `F : D ⥤ C` with the pointwise monoidal structure. -/ --- Porting note: removed `@[simps]`, as some of the generated lemmas were failing the simpNF linter, --- and none of the generated lemmas was actually used in mathlib3. instance closed (F : D ⥤ C) : Closed F where rightAdj := closedIhom F adj := diff --git a/Mathlib/CategoryTheory/CofilteredSystem.lean b/Mathlib/CategoryTheory/CofilteredSystem.lean index fba99256f7078..a3088857c8cf8 100644 --- a/Mathlib/CategoryTheory/CofilteredSystem.lean +++ b/Mathlib/CategoryTheory/CofilteredSystem.lean @@ -344,7 +344,7 @@ theorem eventually_injective [Nonempty J] [Finite F.sections] : have card_le : ∀ j, Fintype.card (F.obj j) ≤ Fintype.card F.sections := fun j => Fintype.card_le_of_surjective _ (F.eval_section_surjective_of_surjective Fsur j) let fn j := Fintype.card F.sections - Fintype.card (F.obj j) - refine ⟨fn.argmin Nat.lt_wfRel.wf, + refine ⟨fn.argmin, fun i f => ((Fintype.bijective_iff_surjective_and_card _).2 ⟨Fsur f, le_antisymm ?_ (Fintype.card_le_of_surjective _ <| Fsur f)⟩).1⟩ rw [← Nat.sub_le_sub_iff_left (card_le i)] diff --git a/Mathlib/CategoryTheory/CommSq.lean b/Mathlib/CategoryTheory/CommSq.lean index 3d76ea7106835..22ac6660feea9 100644 --- a/Mathlib/CategoryTheory/CommSq.lean +++ b/Mathlib/CategoryTheory/CommSq.lean @@ -153,7 +153,6 @@ variable {A B X Y : C} {f : A ⟶ X} {i : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} The datum of a lift in a commutative square, i.e. an up-right-diagonal morphism which makes both triangles commute. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] @[ext] structure LiftStruct (sq : CommSq f i p g) where /-- The lift. -/ diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index 6a161770de191..ea02bd9201175 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -277,6 +277,16 @@ def rightFunc : Arrow C ⥤ C := @[simps] def leftToRight : (leftFunc : Arrow C ⥤ C) ⟶ rightFunc where app f := f.hom +lemma ext {f g : Arrow C} + (h₁ : f.left = g.left) (h₂ : f.right = g.right) + (h₃ : f.hom = eqToHom h₁ ≫ g.hom ≫ eqToHom h₂.symm) : f = g := by + obtain ⟨X, Y, f⟩ := f + obtain ⟨X', Y', g⟩ := g + obtain rfl : X = X' := h₁ + obtain rfl : Y = Y' := h₂ + obtain rfl : f = g := by simpa using h₃ + rfl + end Arrow namespace Functor @@ -334,4 +344,25 @@ def Arrow.isoOfNatIso {C D : Type*} [Category C] [Category D] {F G : C ⥤ D} (e (f : Arrow C) : F.mapArrow.obj f ≅ G.mapArrow.obj f := Arrow.isoMk (e.app f.left) (e.app f.right) +variable (T) + +/-- `Arrow T` is equivalent to a sigma type. -/ +@[simps!] +def Arrow.equivSigma : + Arrow T ≃ Σ (X Y : T), X ⟶ Y where + toFun f := ⟨_, _, f.hom⟩ + invFun x := Arrow.mk x.2.2 + left_inv _ := rfl + right_inv _ := rfl + +/-- The equivalence `Arrow (Discrete S) ≃ S`. -/ +def Arrow.discreteEquiv (S : Type u) : Arrow (Discrete S) ≃ S where + toFun f := f.left.as + invFun s := Arrow.mk (𝟙 (Discrete.mk s)) + left_inv := by + rintro ⟨⟨_⟩, ⟨_⟩, f⟩ + obtain rfl := Discrete.eq_of_hom f + rfl + right_inv _ := rfl + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Comma/CardinalArrow.lean b/Mathlib/CategoryTheory/Comma/CardinalArrow.lean new file mode 100644 index 0000000000000..14443de419c69 --- /dev/null +++ b/Mathlib/CategoryTheory/Comma/CardinalArrow.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Comma.Arrow +import Mathlib.CategoryTheory.FinCategory.Basic +import Mathlib.CategoryTheory.EssentiallySmall +import Mathlib.Data.Set.Finite.Basic +import Mathlib.SetTheory.Cardinal.HasCardinalLT + +/-! +# Cardinal of Arrow + +We obtain various results about the cardinality of `Arrow C`. For example, +If `A` is a (small) category, `Arrow C` is finite iff `FinCategory C` holds. + +-/ + +universe w w' v u + +namespace CategoryTheory + +lemma Arrow.finite_iff (C : Type u) [SmallCategory C] : + Finite (Arrow C) ↔ Nonempty (FinCategory C) := by + constructor + · intro + refine ⟨?_, fun a b ↦ ?_⟩ + · have := Finite.of_injective (fun (a : C) ↦ Arrow.mk (𝟙 a)) + (fun _ _ ↦ congr_arg Comma.left) + apply Fintype.ofFinite + · have := Finite.of_injective (fun (f : a ⟶ b) ↦ Arrow.mk f) + (fun f g h ↦ by + change (Arrow.mk f).hom = (Arrow.mk g).hom + congr) + apply Fintype.ofFinite + · rintro ⟨_⟩ + have := Fintype.ofEquiv _ (Arrow.equivSigma C).symm + infer_instance + +instance Arrow.finite {C : Type u} [SmallCategory C] [FinCategory C] : + Finite (Arrow C) := by + rw [Arrow.finite_iff] + exact ⟨inferInstance⟩ + +/-- The bijection `Arrow Cᵒᵖ ≃ Arrow C`. -/ +def Arrow.opEquiv (C : Type u) [Category.{v} C] : Arrow Cᵒᵖ ≃ Arrow C where + toFun f := Arrow.mk f.hom.unop + invFun g := Arrow.mk g.hom.op + left_inv _ := rfl + right_inv _ := rfl + +@[simp] +lemma hasCardinalLT_arrow_op_iff (C : Type u) [Category.{v} C] (κ : Cardinal.{w}) : + HasCardinalLT (Arrow Cᵒᵖ) κ ↔ HasCardinalLT (Arrow C) κ := + hasCardinalLT_iff_of_equiv (Arrow.opEquiv C) κ + +@[simp] +lemma hasCardinalLT_arrow_discrete_iff {X : Type u} (κ : Cardinal.{w}) : + HasCardinalLT (Arrow (Discrete X)) κ ↔ HasCardinalLT X κ := + hasCardinalLT_iff_of_equiv (Arrow.discreteEquiv X) κ + +lemma small_of_small_arrow (C : Type u) [Category.{v} C] [Small.{w} (Arrow C)] : + Small.{w} C := + small_of_injective (f := fun X ↦ Arrow.mk (𝟙 X)) (fun _ _ h ↦ congr_arg Comma.left h) + +lemma locallySmall_of_small_arrow (C : Type u) [Category.{v} C] [Small.{w} (Arrow C)] : + LocallySmall.{w} C where + hom_small X Y := + small_of_injective (f := fun f ↦ Arrow.mk f) (fun f g h ↦ by + change (Arrow.mk f).hom = (Arrow.mk g).hom + congr) + +/-- The bijection `Arrow.{w} (ShrinkHoms C) ≃ Arrow C`. -/ +noncomputable def Arrow.shrinkHomsEquiv (C : Type u) [Category.{v} C] [LocallySmall.{w} C] : + Arrow.{w} (ShrinkHoms C) ≃ Arrow C where + toFun := (ShrinkHoms.equivalence C).inverse.mapArrow.obj + invFun := (ShrinkHoms.equivalence C).functor.mapArrow.obj + left_inv _ := by simp [Functor.mapArrow]; rfl + right_inv _ := by simp [Functor.mapArrow]; rfl + +/-- The bijection `Arrow (Shrink C) ≃ Arrow C`. -/ +noncomputable def Arrow.shrinkEquiv (C : Type u) [Category.{v} C] [Small.{w} C] : + Arrow (Shrink.{w} C) ≃ Arrow C where + toFun := (Shrink.equivalence C).inverse.mapArrow.obj + invFun := (Shrink.equivalence C).functor.mapArrow.obj + left_inv f := by + refine Arrow.ext (Equiv.apply_symm_apply _ _) + ((Equiv.apply_symm_apply _ _)) (by simp; rfl) + right_inv _ := Arrow.ext (by simp [Shrink.equivalence]) + (by simp [Shrink.equivalence]) (by simp [Shrink.equivalence]) + +@[simp] +lemma hasCardinalLT_arrow_shrinkHoms_iff (C : Type u) [Category.{v} C] [LocallySmall.{w'} C] + (κ : Cardinal.{w}) : + HasCardinalLT (Arrow.{w'} (ShrinkHoms C)) κ ↔ HasCardinalLT (Arrow C) κ := + hasCardinalLT_iff_of_equiv (Arrow.shrinkHomsEquiv C) κ + +@[simp] +lemma hasCardinalLT_arrow_shrink_iff (C : Type u) [Category.{v} C] [Small.{w'} C] + (κ : Cardinal.{w}) : + HasCardinalLT (Arrow (Shrink.{w'} C)) κ ↔ HasCardinalLT (Arrow C) κ := + hasCardinalLT_iff_of_equiv (Arrow.shrinkEquiv C) κ + +lemma hasCardinalLT_of_hasCardinalLT_arrow + {C : Type u} [Category.{v} C] {κ : Cardinal.{w}} (h : HasCardinalLT (Arrow C) κ) : + HasCardinalLT C κ := + h.of_injective (fun X ↦ Arrow.mk (𝟙 X)) (fun _ _ h ↦ congr_arg Comma.left h) + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Comma/Final.lean b/Mathlib/CategoryTheory/Comma/Final.lean index 6c77ce36c3afb..23e5a76639305 100644 --- a/Mathlib/CategoryTheory/Comma/Final.lean +++ b/Mathlib/CategoryTheory/Comma/Final.lean @@ -5,7 +5,9 @@ Authors: Jakob von Raumer -/ import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Limits.IsConnected +import Mathlib.CategoryTheory.Limits.Sifted import Mathlib.CategoryTheory.Filtered.Final +import Mathlib.CategoryTheory.Filtered.Flat import Mathlib.CategoryTheory.Grothendieck import Mathlib.CategoryTheory.Comma.StructuredArrow.CommaMap @@ -19,9 +21,12 @@ final and `A` is connected. We then use this in a proof that derives finality of `map` between two comma categories on a quasi-commutative diagram of functors, some of which need to be final. +Finally we prove filteredness of a `Comma L R` and finality of `snd L R`, given that `R` is final +and `A` and `B` are filtered. + ## References -* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Lemma 3.4.3 & 3.4.4 +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Lemma 3.4.3 -- 3.4.5 -/ universe v₁ v₂ v₃ v₄ v₅ v₆ u₁ u₂ u₃ u₄ u₅ u₆ @@ -80,9 +85,9 @@ instance final_fst [R.Final] : (fst L R).Final := by exact (Functor.associator _ _ _).symm.trans (Iso.compInverseIso (mapFst _ _)) instance initial_snd [L.Initial] : (snd L R).Initial := by - haveI : ((opFunctor L R).leftOp ⋙ fst R.op L.op).Final := + have : ((opFunctor L R).leftOp ⋙ fst R.op L.op).Final := final_equivalence_comp (opEquiv L R).functor.leftOp (fst R.op L.op) - haveI : (snd L R).op.Final := final_of_natIso (opFunctorCompFst _ _) + have : (snd L R).op.Final := final_of_natIso (opFunctorCompFst _ _) apply initial_of_final_op /-- `Comma L R` with `L : A ⥤ T` and `R : B ⥤ T` is connected if `R` is final and `A` is @@ -109,9 +114,9 @@ end NonSmall Let `F`, `G`, `R` and `R'` be final and `B` be filtered. Then, the induced functor between the comma categories of the first and second row of the diagram is final. -/ -lemma map_final {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{u₂} B] {T : Type u₃} +lemma map_final {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{v₂} B] {T : Type u₃} [Category.{v₃} T] {L : A ⥤ T} {R : B ⥤ T} {A' : Type u₄} [Category.{v₄} A'] {B' : Type u₅} - [Category.{v₅} B'] {T' : Type u₂} [Category.{u₂} T'] {L' : A' ⥤ T'} {R' : B' ⥤ T'} {F : A ⥤ A'} + [Category.{v₅} B'] {T' : Type u₆} [Category.{v₆} T'] {L' : A' ⥤ T'} {R' : B' ⥤ T'} {F : A ⥤ A'} {G : B ⥤ B'} {H : T ⥤ T'} (iL : F ⋙ L' ≅ L ⋙ H) (iR : G ⋙ R' ≅ R ⋙ H) [IsFiltered B] [R.Final] [R'.Final] [F.Final] [G.Final] : (Comma.map iL.hom iR.inv).Final := ⟨fun ⟨i₂, j₂, u₂⟩ => by @@ -132,6 +137,57 @@ lemma map_final {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{u rw [IsIso.Iso.inv_inv] infer_instance⟩ +section Filtered + +variable {A : Type u₁} [Category.{v₁} A] +variable {B : Type u₂} [Category.{v₂} B] +variable {T : Type u₃} [Category.{v₃} T] +variable (L : A ⥤ T) (R : B ⥤ T) + +attribute [local instance] map_final in +/-- Let `A` and `B` be filtered categories, `R : B ⥤ T` be final and `L : A ⥤ T`. Then, the +comma category `Comma L R` is filtered. -/ +instance isFiltered_of_final [IsFiltered A] [IsFiltered B] [R.Final] : IsFiltered (Comma L R) := by + haveI (a : A) : IsFiltered (Comma (fromPUnit (L.obj a)) R) := + R.final_iff_isFiltered_structuredArrow.mp inferInstance (L.obj a) + have (a : A) : (fromPUnit (Over.mk (𝟙 a))).Final := final_const_of_isTerminal Over.mkIdTerminal + let η (a : A) : fromPUnit (Over.mk (𝟙 a)) ⋙ Over.forget a ⋙ L ≅ fromPUnit (L.obj a) := + NatIso.ofComponents (fun _ => Iso.refl _) + have (a : A) := IsFiltered.of_final (map (L := fromPUnit (L.obj a)) (F := 𝟭 T) (η a).hom + ((Iso.refl (𝟭 B ⋙ R)).inv)) + have : RepresentablyCoflat (fst L R) := + ⟨fun a => IsFiltered.of_equivalence (CostructuredArrow.ofCommaFstEquivalence L R a).symm⟩ + apply isFiltered_of_representablyCoflat (fst L R) + +attribute [local instance] isFiltered_of_final in +/-- Let `A` and `B` be cofiltered categories, `L : A ⥤ T` be initial and `R : B ⥤ T`. Then, the +comma category `Comma L R` is cofiltered. -/ +lemma isCofiltered_of_initial [IsCofiltered A] [IsCofiltered B] [L.Initial] : + IsCofiltered (Comma L R) := + IsCofiltered.of_equivalence (Comma.opEquiv _ _).symm + +attribute [local instance] final_of_isFiltered_of_pUnit in +/-- Let `A` and `B` be filtered categories, `R : B ⥤ T` be final and `R : A ⥤ T`. Then, the +projection `snd L R : Comma L R ⥤ B` is final. -/ +instance final_snd [IsFiltered A] [IsFiltered B] [R.Final] : (snd L R).Final := by + let iL : star.{1} A ⋙ 𝟭 _ ≅ L ⋙ star _ := Iso.refl _ + let iR : 𝟭 B ⋙ star.{1} B ≅ R ⋙ star _ := Iso.refl _ + have := map_final iL iR + let s := (equivProd (𝟭 _) (star B)).trans <| prod.leftUnitorEquivalence B + let iS : map iL.hom iR.inv ⋙ s.functor ≅ snd L R := + NatIso.ofComponents (fun _ => Iso.refl _) (fun f => by simp [iL, iR, s]) + apply final_of_natIso iS + +/-- Let `A` and `B` be cofiltered categories, `L : A ⥤ T` be initial and `R : B ⥤ T`. Then, the +projection `fst L R : Comma L R ⥤ A` is initial. -/ +instance initial_fst [IsCofiltered A] [IsCofiltered B] [L.Initial] : (fst L R).Initial := by + have : ((opFunctor L R).leftOp ⋙ snd R.op L.op).Final := + final_equivalence_comp (opEquiv L R).functor.leftOp _ + have : (fst L R).op.Final := final_of_natIso <| opFunctorCompSnd _ _ + apply initial_of_final_op + +end Filtered + end Comma end CategoryTheory diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 1ecfb385e2edb..2a293365b9ad6 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -30,10 +30,8 @@ variable {T : Type u₁} [Category.{v₁} T] variable {D : Type u₂} [Category.{v₂} D] /-- The over category has as objects arrows in `T` with codomain `X` and as morphisms commutative -triangles. - -See . --/ +triangles. -/ +@[stacks 001G] def Over (X : T) := CostructuredArrow (𝟭 T) X @@ -92,24 +90,18 @@ end /-- To give a morphism in the over category, it suffices to give an arrow fitting in a commutative triangle. -/ -@[simps!] +@[simps! left] def homMk {U V : Over X} (f : U.left ⟶ V.left) (w : f ≫ V.hom = U.hom := by aesop_cat) : U ⟶ V := CostructuredArrow.homMk f w --- Porting note: simp solves this; simpNF still sees them after `-simp` (?) -attribute [-simp, nolint simpNF] homMk_right_down_down - /-- Construct an isomorphism in the over category given isomorphisms of the objects whose forward direction gives a commutative triangle. -/ -@[simps!] +@[simps! hom_left inv_left] def isoMk {f g : Over X} (hl : f.left ≅ g.left) (hw : hl.hom ≫ g.hom = f.hom := by aesop_cat) : f ≅ g := CostructuredArrow.isoMk hl hw --- Porting note: simp solves this; simpNF still sees them after `-simp` (?) -attribute [-simp, nolint simpNF] isoMk_hom_right_down_down isoMk_inv_right_down_down - @[reassoc (attr := simp)] lemma hom_left_inv_left {f g : Over X} (e : f ≅ g) : e.hom.left ≫ e.inv.left = 𝟙 f.left := by @@ -124,10 +116,8 @@ section variable (X) -/-- The forgetful functor mapping an arrow to its domain. - -See . --/ +/-- The forgetful functor mapping an arrow to its domain. -/ +@[stacks 001G] def forget : Over X ⥤ T := Comma.fst _ _ @@ -147,10 +137,8 @@ def forgetCocone (X : T) : Limits.Cocone (forget X) := { pt := X ι := { app := Comma.hom } } -/-- A morphism `f : X ⟶ Y` induces a functor `Over X ⥤ Over Y` in the obvious way. - -See . --/ +/-- A morphism `f : X ⟶ Y` induces a functor `Over X ⥤ Over Y` in the obvious way. -/ +@[stacks 001G] def map {Y : T} (f : X ⟶ Y) : Over X ⥤ Over Y := Comma.mapRight _ <| Discrete.natTrans fun _ => f @@ -464,13 +452,10 @@ def mk {X Y : T} (f : X ⟶ Y) : Under X := /-- To give a morphism in the under category, it suffices to give a morphism fitting in a commutative triangle. -/ -@[simps!] +@[simps! right] def homMk {U V : Under X} (f : U.right ⟶ V.right) (w : U.hom ≫ f = V.hom := by aesop_cat) : U ⟶ V := StructuredArrow.homMk f w --- Porting note: simp solves this; simpNF still sees them after `-simp` (?) -attribute [-simp, nolint simpNF] homMk_left_down_down - /-- Construct an isomorphism in the over category given isomorphisms of the objects whose forward direction gives a commutative triangle. -/ diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index 6b90ae8f3ff45..427319b4e9619 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -32,8 +32,6 @@ and morphisms `C`-morphisms `Y ⟶ Y'` making the obvious triangle commute. -/ -- We explicitly come from `PUnit.{1}` here to obtain the correct universe for morphisms of -- structured arrows. --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] def StructuredArrow (S : D) (T : C ⥤ D) := Comma (Functor.fromPUnit.{0} S) T @@ -97,7 +95,7 @@ theorem left_eq_id {X Y : StructuredArrow S T} (f : X ⟶ Y) : f.left = 𝟙 X.l we need a morphism of the objects underlying the target, and to check that the triangle commutes. -/ -@[simps] +@[simps right] def homMk {f f' : StructuredArrow S T} (g : f.right ⟶ f'.right) (w : f.hom ≫ T.map g = f'.hom := by aesop_cat) : f ⟶ f' where left := 𝟙 f.left @@ -106,10 +104,6 @@ def homMk {f f' : StructuredArrow S T} (g : f.right ⟶ f'.right) dsimp simpa using w.symm -/- Porting note: it appears the simp lemma is not getting generated but the linter -picks up on it (seems like a bug). Either way simp solves it. -/ -attribute [-simp, nolint simpNF] homMk_left - theorem homMk_surjective {f f' : StructuredArrow S T} (φ : f ⟶ f') : ∃ (ψ : f.right ⟶ f'.right) (hψ : f.hom ≫ T.map ψ = f'.hom), φ = StructuredArrow.homMk ψ hψ := @@ -153,16 +147,12 @@ lemma mkPostcomp_comp (f : S ⟶ T.obj Y) (g : Y ⟶ Y') (g' : Y' ⟶ Y'') : we need an isomorphism of the objects underlying the target, and to check that the triangle commutes. -/ -@[simps!] +@[simps! hom_right inv_right] def isoMk {f f' : StructuredArrow S T} (g : f.right ≅ f'.right) (w : f.hom ≫ T.map g.hom = f'.hom := by aesop_cat) : f ≅ f' := Comma.isoMk (eqToIso (by ext)) g (by simpa using w.symm) -/- Porting note: it appears the simp lemma is not getting generated but the linter -picks up on it. Either way simp solves these. -/ -attribute [-simp, nolint simpNF] isoMk_hom_left_down_down isoMk_inv_left_down_down - theorem ext {A B : StructuredArrow S T} (f g : A ⟶ B) : f.right = g.right → f = g := CommaMorphism.ext (Subsingleton.elim _ _) @@ -193,14 +183,10 @@ theorem eq_mk (f : StructuredArrow S T) : f = mk f.hom := rfl /-- Eta rule for structured arrows. -/ -@[simps!] +@[simps! hom_right inv_right] def eta (f : StructuredArrow S T) : f ≅ mk f.hom := isoMk (Iso.refl _) -/- Porting note: it appears the simp lemma is not getting generated but the linter -picks up on it. Either way simp solves these. -/ -attribute [-simp, nolint simpNF] eta_hom_left_down_down eta_inv_left_down_down - lemma mk_surjective (f : StructuredArrow S T) : ∃ (Y : C) (g : S ⟶ T.obj Y), f = mk g := ⟨_, _, eq_mk f⟩ @@ -398,7 +384,6 @@ and morphisms `C`-morphisms `Y ⟶ Y'` making the obvious triangle commute. -/ -- We explicitly come from `PUnit.{1}` here to obtain the correct universe for morphisms of -- costructured arrows. --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet def CostructuredArrow (S : C ⥤ D) (T : D) := Comma S (Functor.fromPUnit.{0} T) @@ -461,16 +446,12 @@ theorem right_eq_id {X Y : CostructuredArrow S T} (f : X ⟶ Y) : f.right = 𝟙 we need a morphism of the objects underlying the source, and to check that the triangle commutes. -/ -@[simps!] +@[simps! left] def homMk {f f' : CostructuredArrow S T} (g : f.left ⟶ f'.left) (w : S.map g ≫ f'.hom = f.hom := by aesop_cat) : f ⟶ f' where left := g right := 𝟙 f.right -/- Porting note: it appears the simp lemma is not getting generated but the linter -picks up on it. Either way simp can prove this -/ -attribute [-simp, nolint simpNF] homMk_right_down_down - theorem homMk_surjective {f f' : CostructuredArrow S T} (φ : f ⟶ f') : ∃ (ψ : f.left ⟶ f'.left) (hψ : S.map ψ ≫ f'.hom = f.hom), φ = CostructuredArrow.homMk ψ hψ := @@ -514,15 +495,11 @@ lemma mkPrecomp_comp (f : S.obj Y ⟶ T) (g : Y' ⟶ Y) (g' : Y'' ⟶ Y') : we need an isomorphism of the objects underlying the source, and to check that the triangle commutes. -/ -@[simps!] +@[simps! hom_left inv_left] def isoMk {f f' : CostructuredArrow S T} (g : f.left ≅ f'.left) (w : S.map g.hom ≫ f'.hom = f.hom := by aesop_cat) : f ≅ f' := Comma.isoMk g (eqToIso (by ext)) (by simpa using w) -/- Porting note: it appears the simp lemma is not getting generated but the linter -picks up on it. Either way simp solves these. -/ -attribute [-simp, nolint simpNF] isoMk_hom_right_down_down isoMk_inv_right_down_down - theorem ext {A B : CostructuredArrow S T} (f g : A ⟶ B) (h : f.left = g.left) : f = g := CommaMorphism.ext h (Subsingleton.elim _ _) @@ -552,14 +529,10 @@ theorem eq_mk (f : CostructuredArrow S T) : f = mk f.hom := rfl /-- Eta rule for costructured arrows. -/ -@[simps!] +@[simps! hom_left inv_left] def eta (f : CostructuredArrow S T) : f ≅ mk f.hom := isoMk (Iso.refl _) -/- Porting note: it appears the simp lemma is not getting generated but the linter -picks up on it. Either way simp solves these. -/ -attribute [-simp, nolint simpNF] eta_hom_right_down_down eta_inv_right_down_down - lemma mk_surjective (f : CostructuredArrow S T) : ∃ (Y : C) (g : S.obj Y ⟶ T), f = mk g := ⟨_, _, eq_mk f⟩ @@ -975,4 +948,92 @@ def CostructuredArrow.map₂IsoPreEquivalenceInverseCompProj (T : C ⥤ D) (S : end Pre +section Prod + +section + +variable {C' : Type u₃} [Category.{v₃} C'] {D' : Type u₄} [Category.{v₄} D'] + (S : D) (S' : D') (T : C ⥤ D) (T' : C' ⥤ D') + +@[reassoc (attr := simp)] +theorem StructuredArrow.w_prod_fst {X Y : StructuredArrow (S, S') (T.prod T')} + (f : X ⟶ Y) : X.hom.1 ≫ T.map f.right.1 = Y.hom.1 := + congr_arg _root_.Prod.fst (StructuredArrow.w f) + +@[reassoc (attr := simp)] +theorem StructuredArrow.w_prod_snd {X Y : StructuredArrow (S, S') (T.prod T')} + (f : X ⟶ Y) : X.hom.2 ≫ T'.map f.right.2 = Y.hom.2 := + congr_arg _root_.Prod.snd (StructuredArrow.w f) + +/-- Implementation; see `StructuredArrow.prodEquivalence`. -/ +@[simps] +def StructuredArrow.prodFunctor : + StructuredArrow (S, S') (T.prod T') ⥤ StructuredArrow S T × StructuredArrow S' T' where + obj f := ⟨.mk f.hom.1, .mk f.hom.2⟩ + map η := ⟨StructuredArrow.homMk η.right.1 (by simp), + StructuredArrow.homMk η.right.2 (by simp)⟩ + +/-- Implementation; see `StructuredArrow.prodEquivalence`. -/ +@[simps] +def StructuredArrow.prodInverse : + StructuredArrow S T × StructuredArrow S' T' ⥤ StructuredArrow (S, S') (T.prod T') where + obj f := .mk (Y := (f.1.right, f.2.right)) ⟨f.1.hom, f.2.hom⟩ + map η := StructuredArrow.homMk ⟨η.1.right, η.2.right⟩ (by simp) + +/-- The natural equivalence +`StructuredArrow (S, S') (T.prod T') ≌ StructuredArrow S T × StructuredArrow S' T'`. -/ +@[simps] +def StructuredArrow.prodEquivalence : + StructuredArrow (S, S') (T.prod T') ≌ StructuredArrow S T × StructuredArrow S' T' where + functor := StructuredArrow.prodFunctor S S' T T' + inverse := StructuredArrow.prodInverse S S' T T' + unitIso := NatIso.ofComponents (fun f => Iso.refl _) (by simp) + counitIso := NatIso.ofComponents (fun f => Iso.refl _) (by simp) + +end + +section + +variable {C' : Type u₃} [Category.{v₃} C'] {D' : Type u₄} [Category.{v₄} D'] + (S : C ⥤ D) (S' : C' ⥤ D') (T : D) (T' : D') + +@[reassoc (attr := simp)] +theorem CostructuredArrow.w_prod_fst {A B : CostructuredArrow (S.prod S') (T, T')} (f : A ⟶ B) : + S.map f.left.1 ≫ B.hom.1 = A.hom.1 := + congr_arg _root_.Prod.fst (CostructuredArrow.w f) + +@[reassoc (attr := simp)] +theorem CostructuredArrow.w_prod_snd {A B : CostructuredArrow (S.prod S') (T, T')} (f : A ⟶ B) : + S'.map f.left.2 ≫ B.hom.2 = A.hom.2 := + congr_arg _root_.Prod.snd (CostructuredArrow.w f) + +/-- Implementation; see `CostructuredArrow.prodEquivalence`. -/ +@[simps] +def CostructuredArrow.prodFunctor : + CostructuredArrow (S.prod S') (T, T') ⥤ CostructuredArrow S T × CostructuredArrow S' T' where + obj f := ⟨.mk f.hom.1, .mk f.hom.2⟩ + map η := ⟨CostructuredArrow.homMk η.left.1 (by simp), + CostructuredArrow.homMk η.left.2 (by simp)⟩ + +/-- Implementation; see `CostructuredArrow.prodEquivalence`. -/ +@[simps] +def CostructuredArrow.prodInverse : + CostructuredArrow S T × CostructuredArrow S' T' ⥤ CostructuredArrow (S.prod S') (T, T') where + obj f := .mk (Y := (f.1.left, f.2.left)) ⟨f.1.hom, f.2.hom⟩ + map η := CostructuredArrow.homMk ⟨η.1.left, η.2.left⟩ (by simp) + +/-- The natural equivalence +`CostructuredArrow (S.prod S') (T, T') ≌ CostructuredArrow S T × CostructuredArrow S' T'`. -/ +@[simps] +def CostructuredArrow.prodEquivalence : + CostructuredArrow (S.prod S') (T, T') ≌ CostructuredArrow S T × CostructuredArrow S' T' where + functor := CostructuredArrow.prodFunctor S S' T T' + inverse := CostructuredArrow.prodInverse S S' T T' + unitIso := NatIso.ofComponents (fun f => Iso.refl _) (by simp) + counitIso := NatIso.ofComponents (fun f => Iso.refl _) (by simp) + +end + +end Prod + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean index 5a2d0d2c79798..9a0d5b34cc35b 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/CommaMap.lean @@ -23,7 +23,7 @@ noncomputable section variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {T : Type u₃} [Category.{v₃} T] {L : C ⥤ T} {R : D ⥤ T} {C' : Type u₄} [Category.{v₄} C'] {D' : Type u₅} [Category.{v₅} D'] {T' : Type u₆} - [Category.{u₆} T'] {L' : C' ⥤ T'} {R' : D' ⥤ T'} {F₁ : C ⥤ C'} {F₂ : D ⥤ D'} {F : T ⥤ T'} + [Category.{v₆} T'] {L' : C' ⥤ T'} {R' : D' ⥤ T'} {F₁ : C ⥤ C'} {F₂ : D ⥤ D'} {F : T ⥤ T'} (α : F₁ ⋙ L' ⟶ L ⋙ F) (β : R ⋙ F ⟶ F₂ ⋙ R') /-- The functor establishing the equivalence `StructuredArrow.commaMapEquivalence`. -/ diff --git a/Mathlib/CategoryTheory/Core.lean b/Mathlib/CategoryTheory/Core.lean index be75572d3d4ab..5eeb69325f060 100644 --- a/Mathlib/CategoryTheory/Core.lean +++ b/Mathlib/CategoryTheory/Core.lean @@ -28,9 +28,6 @@ universe v₁ v₂ u₁ u₂ -- morphism levels before object levels. See note [CategoryTheory universes]. /-- The core of a category C is the groupoid whose morphisms are all the isomorphisms of C. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not yet ported --- @[nolint has_nonempty_instance] - def Core (C : Type u₁) := C variable {C : Type u₁} [Category.{v₁} C] diff --git a/Mathlib/CategoryTheory/DifferentialObject.lean b/Mathlib/CategoryTheory/DifferentialObject.lean index 067ddda0ea0c1..863323c9203e5 100644 --- a/Mathlib/CategoryTheory/DifferentialObject.lean +++ b/Mathlib/CategoryTheory/DifferentialObject.lean @@ -33,7 +33,6 @@ variable [HasZeroMorphisms C] [HasShift C S] /-- A differential object in a category with zero morphisms and a shift is an object `obj` equipped with a morphism `d : obj ⟶ obj⟦1⟧`, such that `d^2 = 0`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` structure DifferentialObject where /-- The underlying object of a differential object. -/ obj : C @@ -49,7 +48,7 @@ variable {S C} namespace DifferentialObject /-- A morphism of differential objects is a morphism commuting with the differentials. -/ -@[ext] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `nolint has_nonempty_instance` +@[ext] structure Hom (X Y : DifferentialObject S C) where /-- The morphism between underlying objects of the two differentiable objects. -/ f : X.obj ⟶ Y.obj diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index 4505a8e1b9584..f97365afa177a 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -67,10 +67,8 @@ instance {α : Type u₁} [DecidableEq α] : DecidableEq (Discrete α) := /-- The "Discrete" category on a type, whose morphisms are equalities. Because we do not allow morphisms in `Prop` (only in `Type`), -somewhat annoyingly we have to define `X ⟶ Y` as `ULift (PLift (X = Y))`. - -See --/ +somewhat annoyingly we have to define `X ⟶ Y` as `ULift (PLift (X = Y))`. -/ +@[stacks 001A] instance discreteCategory (α : Type u₁) : SmallCategory (Discrete α) where Hom X Y := ULift (PLift (X.as = Y.as)) id _ := ULift.up (PLift.up rfl) diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index fe740cc36ff77..aaf7e3bd09385 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -17,7 +17,6 @@ coalgebras over `G`. ## TODO -* Prove the dual result about the structure map of the terminal coalgebra of an endofunctor. * Prove that if the countable infinite product over the powers of the endofunctor exists, then algebras over the endofunctor coincide with algebras over the free monad on the endofunctor. -/ @@ -382,6 +381,35 @@ def equivOfNatIso {F G : C ⥤ C} (α : F ≅ G) : Coalgebra F ≌ Coalgebra G w counitIso := (functorOfNatTransComp _ _).symm ≪≫ functorOfNatTransEq (by simp) ≪≫ functorOfNatTransId +namespace Terminal + +variable {A} (h : @Limits.IsTerminal (Coalgebra F) _ A) + +/-- The inverse of the structure map of an terminal coalgebra -/ +@[simp] +def strInv : F.obj A.1 ⟶ A.1 := + (h.from ⟨F.obj A.V, F.map A.str⟩).f + +theorem right_inv' : + ⟨A.str ≫ strInv h, by rw [Category.assoc, F.map_comp, strInv, ← Hom.h] ⟩ = 𝟙 A := + Limits.IsTerminal.hom_ext h _ (𝟙 A) + +theorem right_inv : A.str ≫ strInv h = 𝟙 _ := + congr_arg Hom.f (right_inv' h) + +theorem left_inv : strInv h ≫ A.str = 𝟙 _ := by + rw [strInv, ← (h.from ⟨F.obj A.V, F.map A.str⟩).h, ← F.map_id, ← F.map_comp] + congr + exact right_inv h + +/-- The structure map of the terminal coalgebra is an isomorphism, +hence endofunctors preserve their terminal coalgebras +-/ +theorem str_isIso (h : Limits.IsTerminal A) : IsIso A.str := + { out := ⟨strInv h, right_inv _, left_inv _⟩ } + +end Terminal + end Coalgebra namespace Adjunction diff --git a/Mathlib/CategoryTheory/Enriched/Basic.lean b/Mathlib/CategoryTheory/Enriched/Basic.lean index 0654c73c33d3c..41a5c83cc50e2 100644 --- a/Mathlib/CategoryTheory/Enriched/Basic.lean +++ b/Mathlib/CategoryTheory/Enriched/Basic.lean @@ -169,7 +169,6 @@ section variable {W : Type v} [Category.{w} W] [MonoidalCategory W] [EnrichedCategory W C] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` /-- A type synonym for `C`, which should come equipped with a `V`-enriched category structure. In a moment we will equip this with the (honest) category structure so that `X ⟶ Y` is `(𝟙_ W) ⟶ (X ⟶[W] Y)`. @@ -380,7 +379,6 @@ coming from the ambient braiding on `V`.) -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` /-- The type of `A`-graded natural transformations between `V`-functors `F` and `G`. This is the type of morphisms in `V` from `A` to the `V`-object of natural transformations. -/ diff --git a/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean b/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean index b798428bc9175..4636b05731fa4 100644 --- a/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Enriched/FunctorCategory.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Enriched.Ordinary.Basic import Mathlib.CategoryTheory.Functor.Category import Mathlib.CategoryTheory.Limits.Shapes.End diff --git a/Mathlib/CategoryTheory/Enriched/HomCongr.lean b/Mathlib/CategoryTheory/Enriched/HomCongr.lean index 1d4ab2b6588ea..9536a0d158544 100644 --- a/Mathlib/CategoryTheory/Enriched/HomCongr.lean +++ b/Mathlib/CategoryTheory/Enriched/HomCongr.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Nick Ward. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nick Ward -/ -import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Enriched.Ordinary.Basic /-! # Congruence of enriched homs diff --git a/Mathlib/CategoryTheory/Enriched/Opposite.lean b/Mathlib/CategoryTheory/Enriched/Opposite.lean index d25af18fc5ba9..abb743c68e6b0 100644 --- a/Mathlib/CategoryTheory/Enriched/Opposite.lean +++ b/Mathlib/CategoryTheory/Enriched/Opposite.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Daniel Carranza. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Daniel Carranza -/ -import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Enriched.Ordinary.Basic import Mathlib.CategoryTheory.Monoidal.Braided.Basic /-! diff --git a/Mathlib/CategoryTheory/Enriched/Ordinary.lean b/Mathlib/CategoryTheory/Enriched/Ordinary/Basic.lean similarity index 100% rename from Mathlib/CategoryTheory/Enriched/Ordinary.lean rename to Mathlib/CategoryTheory/Enriched/Ordinary/Basic.lean diff --git a/Mathlib/CategoryTheory/EpiMono.lean b/Mathlib/CategoryTheory/EpiMono.lean index a30f9c2501b46..590fa61289cdb 100644 --- a/Mathlib/CategoryTheory/EpiMono.lean +++ b/Mathlib/CategoryTheory/EpiMono.lean @@ -37,7 +37,6 @@ such that `f ≫ retraction f = 𝟙 X`. Every split monomorphism is a monomorphism. -/ -/- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] -/ @[ext, aesop apply safe (rule_sets := [CategoryTheory])] structure SplitMono {X Y : C} (f : X ⟶ Y) where /-- The map splitting `f` -/ @@ -68,7 +67,6 @@ such that `section_ f ≫ f = 𝟙 Y`. Every split epimorphism is an epimorphism. -/ -/- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] -/ @[ext, aesop apply safe (rule_sets := [CategoryTheory])] structure SplitEpi {X Y : C} (f : X ⟶ Y) where /-- The map splitting `f` -/ diff --git a/Mathlib/CategoryTheory/EqToHom.lean b/Mathlib/CategoryTheory/EqToHom.lean index 1a0622b180b50..e9ede8e6dc5e4 100644 --- a/Mathlib/CategoryTheory/EqToHom.lean +++ b/Mathlib/CategoryTheory/EqToHom.lean @@ -338,4 +338,26 @@ theorem dcongr_arg {ι : Type*} {F G : ι → C} (α : ∀ i, F i ⟶ G i) {i j subst h simp +/-- If `T ≃ D` is a bijection and `D` is a category, then +`InducedCategory D e` is equivalent to `D`. -/ +@[simps] +def Equivalence.induced {T : Type*} (e : T ≃ D) : + InducedCategory D e ≌ D where + functor := inducedFunctor e + inverse := + { obj := e.symm + map {X Y} f := show e (e.symm X) ⟶ e (e.symm Y) from + eqToHom (e.apply_symm_apply X) ≫ f ≫ + eqToHom (e.apply_symm_apply Y).symm + map_comp {X Y Z} f g := by + dsimp + erw [Category.assoc, Category.assoc, Category.assoc] + rw [eqToHom_trans_assoc, eqToHom_refl, Category.id_comp] } + unitIso := NatIso.ofComponents (fun _ ↦ eqToIso (by simp)) (fun {X Y} f ↦ by + dsimp + erw [eqToHom_trans_assoc _ (by simp), eqToHom_refl, Category.id_comp] + rfl ) + counitIso := NatIso.ofComponents (fun _ ↦ eqToIso (by simp)) + functor_unitIso_comp X := eqToHom_trans (by simp) (by simp) + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 85a5fc5503a8c..073a622465f51 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -71,11 +71,8 @@ universe v₁ v₂ v₃ u₁ u₂ u₃ The triangle equation is written as a family of equalities between morphisms, it is more complicated if we write it as an equality of natural transformations, because then we would have - to insert natural transformations like `F ⟶ F1`. - -See --/ -@[ext] + to insert natural transformations like `F ⟶ F1`. -/ +@[ext, stacks 001J] structure Equivalence (C : Type u₁) (D : Type u₂) [Category.{v₁} C] [Category.{v₂} D] where mk' :: /-- A functor in one direction -/ functor : C ⥤ D @@ -425,10 +422,8 @@ theorem pow_neg_one (e : C ≌ C) : e ^ (-1 : ℤ) = e.symm := -- Note: the better formulation of this would involve `HasShift`. end -/-- The functor of an equivalence of categories is essentially surjective. - -See . --/ +/-- The functor of an equivalence of categories is essentially surjective. -/ +@[stacks 02C3] instance essSurj_functor (e : C ≌ E) : e.functor.EssSurj := ⟨fun Y => ⟨e.inverse.obj Y, ⟨e.counitIso.app Y⟩⟩⟩ @@ -443,20 +438,16 @@ def fullyFaithfulFunctor (e : C ≌ E) : e.functor.FullyFaithful where def fullyFaithfulInverse (e : C ≌ E) : e.inverse.FullyFaithful where preimage {X Y} f := e.counitIso.inv.app X ≫ e.functor.map f ≫ e.counitIso.hom.app Y -/-- The functor of an equivalence of categories is faithful. - -See . --/ +/-- The functor of an equivalence of categories is faithful. -/ +@[stacks 02C3] instance faithful_functor (e : C ≌ E) : e.functor.Faithful := e.fullyFaithfulFunctor.faithful instance faithful_inverse (e : C ≌ E) : e.inverse.Faithful := e.fullyFaithfulInverse.faithful -/-- The functor of an equivalence of categories is full. - -See . --/ +/-- The functor of an equivalence of categories is full. -/ +@[stacks 02C3] instance full_functor (e : C ≌ E) : e.functor.Full := e.fullyFaithfulFunctor.full @@ -527,10 +518,8 @@ noncomputable def inv (F : C ⥤ D) [F.IsEquivalence] : D ⥤ C where map_id X := by apply F.map_injective; simp map_comp {X Y Z} f g := by apply F.map_injective; simp -/-- Interpret a functor that is an equivalence as an equivalence. - -See . -/ -@[simps functor] +/-- Interpret a functor that is an equivalence as an equivalence. -/ +@[simps functor, stacks 02C3] noncomputable def asEquivalence (F : C ⥤ D) [F.IsEquivalence] : C ≌ D where functor := F inverse := F.inv diff --git a/Mathlib/CategoryTheory/EssentialImage.lean b/Mathlib/CategoryTheory/EssentialImage.lean index a9a14f0935a2b..f7fd17a62ab59 100644 --- a/Mathlib/CategoryTheory/EssentialImage.lean +++ b/Mathlib/CategoryTheory/EssentialImage.lean @@ -109,10 +109,8 @@ def toEssImageCompEssentialImageInclusion (F : C ⥤ D) : F.toEssImage ⋙ F.ess FullSubcategory.lift_comp_inclusion _ _ _ /-- A functor `F : C ⥤ D` is essentially surjective if every object of `D` is in the essential -image of `F`. In other words, for every `Y : D`, there is some `X : C` with `F.obj X ≅ Y`. - -See . --/ +image of `F`. In other words, for every `Y : D`, there is some `X : C` with `F.obj X ≅ Y`. -/ +@[stacks 001C] class EssSurj (F : C ⥤ D) : Prop where /-- All the objects of the target category are in the essential image. -/ mem_essImage (Y : D) : Y ∈ F.essImage diff --git a/Mathlib/CategoryTheory/EssentiallySmall.lean b/Mathlib/CategoryTheory/EssentiallySmall.lean index febf03357f7fd..1f2d4856e72bb 100644 --- a/Mathlib/CategoryTheory/EssentiallySmall.lean +++ b/Mathlib/CategoryTheory/EssentiallySmall.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.CategoryTheory.Category.ULift +import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Skeletal import Mathlib.Logic.UnivLE import Mathlib.Logic.Small.Basic @@ -22,7 +23,7 @@ the type `Skeleton C` is `w`-small, and `C` is `w`-locally small. -/ -universe w v v' u u' +universe w w' v v' u u' open CategoryTheory @@ -44,7 +45,6 @@ theorem EssentiallySmall.mk' {C : Type u} [Category.{v} C] {S : Type w} [SmallCa /-- An arbitrarily chosen small model for an essentially small category. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171) removed @[nolint has_nonempty_instance] @[pp_with_univ] def SmallModel (C : Type u) [Category.{v} C] [EssentiallySmall.{w} C] : Type w := Classical.choose (@EssentiallySmall.equiv_smallCategory C _ _) @@ -118,7 +118,6 @@ instance (priority := 100) locallySmall_of_essentiallySmall (C : Type u) [Catego /-- We define a type alias `ShrinkHoms C` for `C`. When we have `LocallySmall.{w} C`, we'll put a `Category.{w}` instance on `ShrinkHoms C`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] @[pp_with_univ] def ShrinkHoms (C : Type u) := C @@ -189,7 +188,11 @@ noncomputable instance [Small.{w} C] : Category.{v} (Shrink.{w} C) := /-- The categorical equivalence between `C` and `Shrink C`, when `C` is small. -/ noncomputable def equivalence [Small.{w} C] : C ≌ Shrink.{w} C := - (inducedFunctor (equivShrink C).symm).asEquivalence.symm + (Equivalence.induced _).symm + +instance [Small.{w'} C] [LocallySmall.{w} C] : + LocallySmall.{w} (Shrink.{w'} C) := + locallySmall_of_faithful.{w} (equivalence.{w'} C).inverse end Shrink diff --git a/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean b/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean index b6f222bcedcbc..08a933d5053c9 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/Cartesian.lean @@ -65,9 +65,8 @@ a' a --φ--> b v v v R' --g--> R --f--> S ``` -such that `φ'` lifts `g ≫ f`, there exists a lift `χ` of `g` such that `φ' = χ ≫ φ`. - -See . -/ +such that `φ'` lifts `g ≫ f`, there exists a lift `χ` of `g` such that `φ' = χ ≫ φ`. -/ +@[stacks 02XK] class IsStronglyCartesian extends IsHomLift p f φ : Prop where universal_property' {a' : 𝒳} (g : p.obj a' ⟶ R) (φ' : a' ⟶ b) [IsHomLift p (g ≫ f) φ'] : ∃! χ : a' ⟶ a, IsHomLift p g χ ∧ χ ≫ φ = φ' diff --git a/Mathlib/CategoryTheory/FiberedCategory/Cocartesian.lean b/Mathlib/CategoryTheory/FiberedCategory/Cocartesian.lean index d2e2b408959c9..48278cc831bc2 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/Cocartesian.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/Cocartesian.lean @@ -60,9 +60,8 @@ a --φ--> b b' v v v R --f--> S --g--> S' ``` -such that `φ'` lifts `f ≫ g`, there exists a lift `χ` of `g` such that `φ' = χ ≫ φ`. - -See . -/ +such that `φ'` lifts `f ≫ g`, there exists a lift `χ` of `g` such that `φ' = χ ≫ φ`. -/ +@[stacks 02XK] class IsStronglyCocartesian extends IsHomLift p f φ : Prop where universal_property' {b' : 𝒳} (g : S ⟶ p.obj b') (φ' : a ⟶ b') [IsHomLift p (f ≫ g) φ'] : ∃! χ : b ⟶ b', IsHomLift p g χ ∧ φ ≫ χ = φ' diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index e356e98a41ca6..3e22decd81bac 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -76,10 +76,8 @@ class IsFilteredOrEmpty : Prop where 1. for every pair of objects there exists another object "to the right", 2. for every pair of parallel morphisms there exists a morphism to the right so the compositions are equal, and -3. there exists some object. - -See . (They also define a diagram being filtered.) --/ +3. there exists some object. -/ +@[stacks 002V "They also define a diagram being filtered."] class IsFiltered extends IsFilteredOrEmpty C : Prop where /-- a filtered category must be non empty -/ [nonempty : Nonempty C] @@ -504,10 +502,8 @@ class IsCofilteredOrEmpty : Prop where 1. for every pair of objects there exists another object "to the left", 2. for every pair of parallel morphisms there exists a morphism to the left so the compositions are equal, and -3. there exists some object. - -See . --/ +3. there exists some object. -/ +@[stacks 04AZ] class IsCofiltered extends IsCofilteredOrEmpty C : Prop where /-- a cofiltered category must be non empty -/ [nonempty : Nonempty C] diff --git a/Mathlib/CategoryTheory/Filtered/CostructuredArrow.lean b/Mathlib/CategoryTheory/Filtered/CostructuredArrow.lean new file mode 100644 index 0000000000000..2e51870130933 --- /dev/null +++ b/Mathlib/CategoryTheory/Filtered/CostructuredArrow.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2024 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer +-/ +import Mathlib.CategoryTheory.Filtered.OfColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Limits.Preserves.Grothendieck +import Mathlib.CategoryTheory.Limits.Final + +/-! +# Inferring Filteredness from Filteredness of Costructured Arrow Categories + +# References + +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Proposition 3.1.8 + +-/ + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory + +open Limits Functor + +section Small + +variable {A : Type u₁} [SmallCategory A] {B : Type u₁} [SmallCategory B] +variable {T : Type u₁} [SmallCategory T] + +private lemma isFiltered_of_isFiltered_costructuredArrow_small (L : A ⥤ T) (R : B ⥤ T) + [IsFiltered B] [Final R] [∀ b, IsFiltered (CostructuredArrow L (R.obj b))] : IsFiltered A := + isFiltered_of_nonempty_limit_colimit_to_colimit_limit fun J {_ _} F => ⟨by + let R' := Grothendieck.pre (CostructuredArrow.functor L) R + haveI : ∀ b, PreservesLimitsOfShape J + (colim (J := (R ⋙ CostructuredArrow.functor L).obj b) (C := Type u₁)) := fun b => by + simp only [comp_obj, CostructuredArrow.functor_obj, Cat.of_α] + exact filtered_colim_preservesFiniteLimits + refine lim.map ((colimitIsoColimitGrothendieck L F.flip).hom ≫ + (inv (colimit.pre (CostructuredArrow.grothendieckProj L ⋙ F.flip) R'))) ≫ + (colimitLimitIso (R' ⋙ CostructuredArrow.grothendieckProj L ⋙ F.flip).flip).inv ≫ + colim.map ?_ ≫ + colimit.pre _ R' ≫ + (colimitIsoColimitGrothendieck L (limit F)).inv + exact (limitCompWhiskeringLeftIsoCompLimit F (R' ⋙ CostructuredArrow.grothendieckProj L)).hom⟩ + +end Small + +variable {A : Type u₁} [Category.{v₁} A] {B : Type u₂} [Category.{v₂} B] +variable {T : Type u₃} [Category.{v₃} T] + +/-- Given functors `L : A ⥤ T` and `R : B ⥤ T` with a common codomain we can conclude that `A` +is filtered given that `R` is final, `B` is filtered and each costructured arrow category +`CostructuredArrow L (R.obj b)` is filtered. -/ +theorem isFiltered_of_isFiltered_costructuredArrow (L : A ⥤ T) (R : B ⥤ T) + [IsFiltered B] [Final R] [∀ b, IsFiltered (CostructuredArrow L (R.obj b))] : IsFiltered A := by + let sA : A ≌ AsSmall.{max u₁ u₂ u₃ v₁ v₂ v₃} A := AsSmall.equiv + let sB : B ≌ AsSmall.{max u₁ u₂ u₃ v₁ v₂ v₃} B := AsSmall.equiv + let sT : T ≌ AsSmall.{max u₁ u₂ u₃ v₁ v₂ v₃} T := AsSmall.equiv + let sC : ∀ b, CostructuredArrow (sA.inverse ⋙ L ⋙ sT.functor) + ((sB.inverse ⋙ R ⋙ sT.functor).obj ⟨b⟩) ≌ CostructuredArrow L (R.obj b) := fun b => + (CostructuredArrow.pre sA.inverse (L ⋙ sT.functor) _).asEquivalence.trans + (CostructuredArrow.post L sT.functor _).asEquivalence.symm + haveI : ∀ b, IsFiltered (CostructuredArrow _ ((sB.inverse ⋙ R ⋙ sT.functor).obj b)) := + fun b => IsFiltered.of_equivalence (sC b.1).symm + haveI := isFiltered_of_isFiltered_costructuredArrow_small + (sA.inverse ⋙ L ⋙ sT.functor) (sB.inverse ⋙ R ⋙ sT.functor) + exact IsFiltered.of_equivalence sA.symm + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Filtered/Final.lean b/Mathlib/CategoryTheory/Filtered/Final.lean index dc5eaabbc2b35..24e561db7288e 100644 --- a/Mathlib/CategoryTheory/Filtered/Final.lean +++ b/Mathlib/CategoryTheory/Filtered/Final.lean @@ -5,7 +5,7 @@ Authors: Markus Himmel -/ import Mathlib.CategoryTheory.Filtered.Connected import Mathlib.CategoryTheory.Limits.TypesFiltered -import Mathlib.CategoryTheory.Limits.Final +import Mathlib.CategoryTheory.Limits.Sifted /-! # Final functors with filtered (co)domain @@ -39,8 +39,6 @@ namespace CategoryTheory open CategoryTheory.Limits CategoryTheory.Functor Opposite -section ArbitraryUniverses - variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] (F : C ⥤ D) /-- If `StructuredArrow d F` is filtered for any `d : D`, then `F : C ⥤ D` is final. This is @@ -224,12 +222,24 @@ instance Over.initial_forget [IsCofilteredOrEmpty C] (c : C) : Initial (Over.for simp only [forget_obj, mk_left, forget_map, homMk_left] rw [IsCofiltered.eq_condition]) -end ArbitraryUniverses - section LocallySmall variable {C : Type v₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₁} D] (F : C ⥤ D) +/-- Implementation; use `Functor.Final.exists_coeq instead`. -/ +theorem Functor.Final.exists_coeq_of_locally_small [IsFilteredOrEmpty C] [Final F] {d : D} {c : C} + (s s' : d ⟶ F.obj c) : ∃ (c' : C) (t : c ⟶ c'), s ≫ F.map t = s' ≫ F.map t := by + have : colimit.ι (F ⋙ coyoneda.obj (op d)) c s = colimit.ι (F ⋙ coyoneda.obj (op d)) c s' := by + apply (Final.colimitCompCoyonedaIso F d).toEquiv.injective + subsingleton + obtain ⟨c', t₁, t₂, h⟩ := (Types.FilteredColimit.colimit_eq_iff.{v₁, v₁, v₁} _).mp this + refine ⟨IsFiltered.coeq t₁ t₂, t₁ ≫ IsFiltered.coeqHom t₁ t₂, ?_⟩ + conv_rhs => rw [IsFiltered.coeq_condition t₁ t₂] + dsimp only [comp_obj, coyoneda_obj_obj, unop_op, Functor.comp_map, coyoneda_obj_map] at h + simp [reassoc_of% h] + +end LocallySmall + /-- If `C` is filtered, then we can give an explicit condition for a functor `F : C ⥤ D` to be final. -/ theorem Functor.final_iff_of_isFiltered [IsFilteredOrEmpty C] : @@ -239,15 +249,13 @@ theorem Functor.final_iff_of_isFiltered [IsFilteredOrEmpty C] : · intro d obtain ⟨f⟩ : Nonempty (StructuredArrow d F) := IsConnected.is_nonempty exact ⟨_, ⟨f.hom⟩⟩ - · intro d c s s' - have : colimit.ι (F ⋙ coyoneda.obj (op d)) c s = colimit.ι (F ⋙ coyoneda.obj (op d)) c s' := by - apply (Final.colimitCompCoyonedaIso F d).toEquiv.injective - subsingleton - obtain ⟨c', t₁, t₂, h⟩ := (Types.FilteredColimit.colimit_eq_iff.{v₁, v₁, v₁} _).mp this - refine ⟨IsFiltered.coeq t₁ t₂, t₁ ≫ IsFiltered.coeqHom t₁ t₂, ?_⟩ - conv_rhs => rw [IsFiltered.coeq_condition t₁ t₂] - dsimp only [comp_obj, coyoneda_obj_obj, unop_op, Functor.comp_map, coyoneda_obj_map] at h - simp [reassoc_of% h] + · let s₁ : C ≌ AsSmall.{max u₁ v₁ u₂ v₂} C := AsSmall.equiv + let s₂ : D ≌ AsSmall.{max u₁ v₁ u₂ v₂} D := AsSmall.equiv + have : IsFilteredOrEmpty (AsSmall.{max u₁ v₁ u₂ v₂} C) := .of_equivalence s₁ + intro d c s s' + obtain ⟨c', t, ht⟩ := Functor.Final.exists_coeq_of_locally_small (s₁.inverse ⋙ F ⋙ s₂.functor) + (AsSmall.up.map s) (AsSmall.up.map s') + exact ⟨AsSmall.down.obj c', AsSmall.down.map t, s₂.functor.map_injective (by simp_all [s₁, s₂])⟩ /-- If `C` is cofiltered, then we can give an explicit condition for a functor `F : C ⥤ D` to be initial. -/ @@ -290,32 +298,36 @@ theorem Functor.initial_iff_isCofiltered_costructuredArrow [IsCofilteredOrEmpty /-- If `C` is filtered, then the structured arrow category on the diagonal functor `C ⥤ C × C` is filtered as well. -/ -instance [IsFiltered C] (X : C × C) : IsFiltered (StructuredArrow X (diag C)) := by +instance [IsFilteredOrEmpty C] (X : C × C) : IsFiltered (StructuredArrow X (diag C)) := by haveI : ∀ Y, IsFiltered (StructuredArrow Y (Under.forget X.1)) := by rw [← final_iff_isFiltered_structuredArrow (Under.forget X.1)] infer_instance apply IsFiltered.of_equivalence (StructuredArrow.ofDiagEquivalence X).symm /-- The diagonal functor on any filtered category is final. -/ -instance Functor.final_diag_of_isFiltered [IsFiltered C] : Final (Functor.diag C) := +instance Functor.final_diag_of_isFiltered [IsFilteredOrEmpty C] : Final (Functor.diag C) := final_of_isFiltered_structuredArrow _ +-- Adding this instance causes performance problems elsewhere, even with low priority +theorem IsFilteredOrEmpty.isSiftedOrEmpty [IsFilteredOrEmpty C] : IsSiftedOrEmpty C := + Functor.final_diag_of_isFiltered + +-- Adding this instance causes performance problems elsewhere, even with low priority +attribute [local instance] IsFiltered.nonempty in +theorem IsFiltered.isSifted [IsFiltered C] : IsSifted C where + /-- If `C` is cofiltered, then the costructured arrow category on the diagonal functor `C ⥤ C × C` is cofiltered as well. -/ -instance [IsCofiltered C] (X : C × C) : IsCofiltered (CostructuredArrow (diag C) X) := by +instance [IsCofilteredOrEmpty C] (X : C × C) : IsCofiltered (CostructuredArrow (diag C) X) := by haveI : ∀ Y, IsCofiltered (CostructuredArrow (Over.forget X.1) Y) := by rw [← initial_iff_isCofiltered_costructuredArrow (Over.forget X.1)] infer_instance apply IsCofiltered.of_equivalence (CostructuredArrow.ofDiagEquivalence X).symm /-- The diagonal functor on any cofiltered category is initial. -/ -instance Functor.initial_diag_of_isFiltered [IsCofiltered C] : Initial (Functor.diag C) := +instance Functor.initial_diag_of_isFiltered [IsCofilteredOrEmpty C] : Initial (Functor.diag C) := initial_of_isCofiltered_costructuredArrow _ -end LocallySmall - -variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] - /-- If `C` is filtered, then every functor `F : C ⥤ Discrete PUnit` is final. -/ theorem Functor.final_of_isFiltered_of_pUnit [IsFiltered C] (F : C ⥤ Discrete PUnit) : Final F := by @@ -353,8 +365,7 @@ instance CostructuredArrow.initial_proj_of_isCofiltered [IsCofilteredOrEmpty C] /-- The functor `StructuredArrow d T ⥤ StructuredArrow e (T ⋙ S)` that `u : e ⟶ S.obj d` induces via `StructuredArrow.map₂` is final, if `T` and `S` are final and the domain of `T` is filtered. -/ -instance StructuredArrow.final_map₂_id {D : Type u₂} [Category.{v₂} D] - {C : Type v₁} [Category.{v₁} C] [IsFiltered C] {E : Type u₃} [Category.{v₁} E] +instance StructuredArrow.final_map₂_id [IsFiltered C] {E : Type u₃} [Category.{v₃} E] {T : C ⥤ D} [T.Final] {S : D ⥤ E} [S.Final] {T' : C ⥤ E} {d : D} {e : E} (u : e ⟶ S.obj d) (α : T ⋙ S ⟶ T') [IsIso α] : Final (map₂ (F := 𝟭 _) u α) := by @@ -363,8 +374,7 @@ instance StructuredArrow.final_map₂_id {D : Type u₂} [Category.{v₂} D] apply final_of_natIso (map₂IsoPreEquivalenceInverseCompProj d e u α).symm /-- `StructuredArrow.map` is final if the functor `T` is final` and its domain is filtered. -/ -instance StructuredArrow.final_map {C : Type v₁} [Category.{v₁} C] [IsFiltered C] - {D : Type v₁} [Category.{v₁} D] {S S' : D} (f : S ⟶ S') (T : C ⥤ D) [T.Final] : +instance StructuredArrow.final_map [IsFiltered C] {S S' : D} (f : S ⟶ S') (T : C ⥤ D) [T.Final] : Final (map (T := T) f) := by haveI := NatIso.isIso_of_isIso_app (𝟙 T) have : (map₂ (F := 𝟭 C) (G := 𝟭 D) f (𝟙 T)).Final := by @@ -373,24 +383,23 @@ instance StructuredArrow.final_map {C : Type v₁} [Category.{v₁} C] [IsFilter /-- `StructuredArrow.post X T S` is final if `T` and `S` are final and the domain of `T` is filtered. -/ -instance StructuredArrow.final_post {C : Type v₁} [Category.{v₁} C] [IsFiltered C] {E : Type u₃} - [Category.{v₁} E] (X : D) (T : C ⥤ D) [T.Final] (S : D ⥤ E) [S.Final] : Final (post X T S) := by +instance StructuredArrow.final_post [IsFiltered C] {E : Type u₃} [Category.{v₃} E] (X : D) + (T : C ⥤ D) [T.Final] (S : D ⥤ E) [S.Final] : Final (post X T S) := by apply final_of_natIso (postIsoMap₂ X T S).symm /-- The functor `CostructuredArrow T d ⥤ CostructuredArrow (T ⋙ S) e` that `u : S.obj d ⟶ e` induces via `CostructuredArrow.map₂` is initial, if `T` and `S` are initial and the domain of `T` is filtered. -/ -instance CostructuredArrow.initial_map₂_id {C : Type v₁} [Category.{v₁} C] [IsCofiltered C] - {E : Type u₃} [Category.{v₁} E] (T : C ⥤ D) [T.Initial] (S : D ⥤ E) [S.Initial] (d : D) (e : E) +instance CostructuredArrow.initial_map₂_id [IsCofiltered C] {E : Type u₃} [Category.{v₃} E] + (T : C ⥤ D) [T.Initial] (S : D ⥤ E) [S.Initial] (d : D) (e : E) (u : S.obj d ⟶ e) : Initial (map₂ (F := 𝟭 _) (U := T ⋙ S) (𝟙 (T ⋙ S)) u) := by have := (T ⋙ S).initial_iff_isCofiltered_costructuredArrow.mp inferInstance e apply initial_of_natIso (map₂IsoPreEquivalenceInverseCompProj T S d e u).symm /-- `CostructuredArrow.post T S X` is initial if `T` and `S` are initial and the domain of `T` is cofiltered. -/ -instance CostructuredArrow.initial_post {C : Type v₁} [Category.{v₁} C] [IsCofiltered C] - {E : Type u₃} [Category.{v₁} E] (X : D) (T : C ⥤ D) [T.Initial] (S : D ⥤ E) [S.Initial] : - Initial (post T S X) := by +instance CostructuredArrow.initial_post [IsCofiltered C] {E : Type u₃} [Category.{v₃} E] (X : D) + (T : C ⥤ D) [T.Initial] (S : D ⥤ E) [S.Initial] : Initial (post T S X) := by apply initial_of_natIso (postIsoMap₂ X T S).symm section Pi diff --git a/Mathlib/CategoryTheory/Filtered/Flat.lean b/Mathlib/CategoryTheory/Filtered/Flat.lean new file mode 100644 index 0000000000000..f4ca7ef2b7bb2 --- /dev/null +++ b/Mathlib/CategoryTheory/Filtered/Flat.lean @@ -0,0 +1,39 @@ +/- +Copyright (c) 2025 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer +-/ +import Mathlib.CategoryTheory.Filtered.CostructuredArrow +import Mathlib.CategoryTheory.Functor.Flat + +/-! +# Pulling back filteredness along representably flat functors + +We show that if `F : C ⥤ D` is a representably coflat functor between two categories, +filteredness of `D` implies filteredness of `C`. Dually, if `F` is representably flat, +cofilteredness of `D` implies cofilteredness of `C`. + +Transferring (co)filteredness *along* representably (co)flat functors is given by +`IsFiltered.of_final` and its dual, since every representably flat functor is final and every +representably coflat functor is initial. +-/ + +universe v₁ v₂ u₁ u₂ + +namespace CategoryTheory + +open Limits + +variable {C : Type u₁} [Category.{v₁} C] +variable {D : Type u₂} [Category.{v₂} D] +variable (F : C ⥤ D) + +lemma isFiltered_of_representablyCoflat [IsFiltered D] [RepresentablyCoflat F] : IsFiltered C := + isFiltered_of_isFiltered_costructuredArrow F (𝟭 _) + +lemma isCofiltered_of_representablyFlat [IsCofiltered D] [RepresentablyFlat F] : + IsCofiltered C := by + have := isFiltered_of_representablyCoflat F.op + exact isCofiltered_of_isFiltered_op C + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Filtered/OfColimitCommutesFiniteLimit.lean b/Mathlib/CategoryTheory/Filtered/OfColimitCommutesFiniteLimit.lean index 45baf017ad8b0..0a993d2e227e7 100644 --- a/Mathlib/CategoryTheory/Filtered/OfColimitCommutesFiniteLimit.lean +++ b/Mathlib/CategoryTheory/Filtered/OfColimitCommutesFiniteLimit.lean @@ -18,7 +18,7 @@ variable {K : Type v} [SmallCategory K] open Limits -/-- A converse to `colimitLimitIsoLimitColimit`: if colimits of shape `K` commute with finite +/-- A converse to `colimitLimitIso`: if colimits of shape `K` commute with finite limits, then `K` is filtered. -/ theorem isFiltered_of_nonempty_limit_colimit_to_colimit_limit (h : ∀ {J : Type v} [SmallCategory J] [FinCategory J] (F : J ⥤ K ⥤ Type v), diff --git a/Mathlib/CategoryTheory/FullSubcategory.lean b/Mathlib/CategoryTheory/FullSubcategory.lean index 7e3a337eddf7e..3e85d68da2273 100644 --- a/Mathlib/CategoryTheory/FullSubcategory.lean +++ b/Mathlib/CategoryTheory/FullSubcategory.lean @@ -47,7 +47,6 @@ variable (F : C → D) which provides a category structure so that the morphisms `X ⟶ Y` are the morphisms in `D` from `F X` to `F Y`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] @[nolint unusedArguments] def InducedCategory (_F : C → D) : Type u₁ := C @@ -100,11 +99,8 @@ variable (Z : C → Prop) /-- A subtype-like structure for full subcategories. Morphisms just ignore the property. We don't use actual subtypes since the simp-normal form `↑X` of `X.val` does not work well for full -subcategories. - -See . We do not define 'strictly full' subcategories. --/ -@[ext] +subcategories. -/ +@[ext, stacks 001D "We do not define 'strictly full' subcategories."] structure FullSubcategory where /-- The category of which this is a full subcategory -/ obj : C diff --git a/Mathlib/CategoryTheory/Functor/Basic.lean b/Mathlib/CategoryTheory/Functor/Basic.lean index 8219f68054377..13a3440fd5fb9 100644 --- a/Mathlib/CategoryTheory/Functor/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/Basic.lean @@ -30,10 +30,8 @@ section To apply a functor `F` to an object use `F.obj X`, and to a morphism use `F.map f`. The axiom `map_id` expresses preservation of identities, and -`map_comp` expresses functoriality. - -See . --/ +`map_comp` expresses functoriality. -/ +@[stacks 001B] structure Functor (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] extends Prefunctor C D : Type max v₁ v₂ u₁ u₂ where /-- A functor preserves identity morphisms. -/ diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index c828dd56ec407..b521f581f772a 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ +import Mathlib.CategoryTheory.Filtered.Connected import Mathlib.CategoryTheory.Limits.ConeCategory import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit import Mathlib.CategoryTheory.Limits.Preserves.Filtered @@ -142,6 +143,12 @@ instance RepresentablyCoflat.comp (G : D ⥤ E) [RepresentablyCoflat F] [Represe RepresentablyCoflat (F ⋙ G) := (representablyFlat_op_iff _).1 <| inferInstanceAs <| RepresentablyFlat (F.op ⋙ G.op) +lemma final_of_representablyFlat [h : RepresentablyFlat F] : F.Final where + out _ := IsCofiltered.isConnected _ + +lemma initial_of_representablyCoflat [h : RepresentablyCoflat F] : F.Initial where + out _ := IsFiltered.isConnected _ + end RepresentablyFlat section HasLimit diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index c550e1890088a..98607af7c86c4 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -36,17 +36,13 @@ variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] namespace Functor -/-- A functor `F : C ⥤ D` is full if for each `X Y : C`, `F.map` is surjective. - -See . --/ +/-- A functor `F : C ⥤ D` is full if for each `X Y : C`, `F.map` is surjective. -/ +@[stacks 001C] class Full (F : C ⥤ D) : Prop where map_surjective {X Y : C} : Function.Surjective (F.map (X := X) (Y := Y)) -/-- A functor `F : C ⥤ D` is faithful if for each `X Y : C`, `F.map` is injective. - -See . --/ +/-- A functor `F : C ⥤ D` is faithful if for each `X Y : C`, `F.map` is injective. -/ +@[stacks 001C] class Faithful (F : C ⥤ D) : Prop where /-- `F.map` is injective for each `X Y : C`. -/ map_injective : ∀ {X Y : C}, Function.Injective (F.map : (X ⟶ Y) → (F.obj X ⟶ F.obj Y)) := by diff --git a/Mathlib/CategoryTheory/Galois/Topology.lean b/Mathlib/CategoryTheory/Galois/Topology.lean index 17d6329fb58cd..7020a117b13ee 100644 --- a/Mathlib/CategoryTheory/Galois/Topology.lean +++ b/Mathlib/CategoryTheory/Galois/Topology.lean @@ -17,7 +17,7 @@ embedding of `Aut F` into `∀ X, Aut (F.obj X)` where ## References -- Stacks Project: Tag 0BMQ +- [Stacks 0BMQ](https://stacks.math.columbia.edu/tag/0BMQ) -/ diff --git a/Mathlib/CategoryTheory/Generator/Basic.lean b/Mathlib/CategoryTheory/Generator/Basic.lean index 049514c52cef9..70168c6d7626d 100644 --- a/Mathlib/CategoryTheory/Generator/Basic.lean +++ b/Mathlib/CategoryTheory/Generator/Basic.lean @@ -527,15 +527,21 @@ theorem isSeparator_coprod_of_isSeparator_right (G H : C) [HasBinaryCoproduct G (hH : IsSeparator H) : IsSeparator (G ⨿ H) := (isSeparator_coprod _ _).2 <| IsSeparating.mono hH <| by simp +lemma isSeparator_of_isColimit_cofan {β : Type w} {f : β → C} + (hf : IsSeparating (Set.range f)) {c : Cofan f} (hc : IsColimit c) : IsSeparator c.pt := by + refine (isSeparator_def _).2 fun X Y u v huv => hf _ _ fun Z hZ g => ?_ + obtain ⟨b, rfl⟩ := Set.mem_range.1 hZ + classical simpa using c.ι.app ⟨b⟩ ≫= huv (hc.desc (Cofan.mk _ (Pi.single b g))) + theorem isSeparator_sigma {β : Type w} (f : β → C) [HasCoproduct f] : IsSeparator (∐ f) ↔ IsSeparating (Set.range f) := by - refine - ⟨fun h X Y u v huv => ?_, fun h => - (isSeparator_def _).2 fun X Y u v huv => h _ _ fun Z hZ g => ?_⟩ - · refine h.def _ _ fun g => colimit.hom_ext fun b => ?_ - simpa using huv (f b.as) (by simp) (colimit.ι (Discrete.functor f) _ ≫ g) - · obtain ⟨b, rfl⟩ := Set.mem_range.1 hZ - classical simpa using Sigma.ι f b ≫= huv (Sigma.desc (Pi.single b g)) + refine ⟨fun h X Y u v huv => ?_, fun h => isSeparator_of_isColimit_cofan h (colimit.isColimit _)⟩ + refine h.def _ _ fun g => colimit.hom_ext fun b => ?_ + simpa using huv (f b.as) (by simp) (colimit.ι (Discrete.functor f) _ ≫ g) + +theorem IsSeparating.isSeparator_coproduct {β : Type w} {f : β → C} [HasCoproduct f] + (hS : IsSeparating (Set.range f)) : IsSeparator (∐ f) := + (isSeparator_sigma _).2 hS theorem isSeparator_sigma_of_isSeparator {β : Type w} (f : β → C) [HasCoproduct f] (b : β) (hb : IsSeparator (f b)) : IsSeparator (∐ f) := diff --git a/Mathlib/CategoryTheory/Generator/Coproduct.lean b/Mathlib/CategoryTheory/Generator/Coproduct.lean deleted file mode 100644 index 0e948301d4819..0000000000000 --- a/Mathlib/CategoryTheory/Generator/Coproduct.lean +++ /dev/null @@ -1,43 +0,0 @@ -/- -Copyright (c) 2024 Joël Riou. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joël Riou --/ - -import Mathlib.CategoryTheory.Generator.Basic -import Mathlib.CategoryTheory.Limits.Shapes.Products - -/-! -# The coproduct of a separating family of objects is separating - -If a family of objects `S : ι → C` in a category with zero morphisms -is separating, then the coproduct of `S` is a separator in `C`. - --/ - -universe w v u - -namespace CategoryTheory.IsSeparating - -open Limits - -variable {C : Type u} [Category.{v} C] [HasZeroMorphisms C] - {ι : Type w} {S : ι → C} - -open Classical in -lemma isSeparator_of_isColimit_cofan - (hS : IsSeparating (Set.range S)) {c : Cofan S} (hc : IsColimit c) : - IsSeparator c.pt := by - intro X Y f g h - apply hS - rintro _ ⟨i, rfl⟩ α - let β : c.pt ⟶ X := Cofan.IsColimit.desc hc - (fun j ↦ if hij : i = j then eqToHom (by rw [hij]) ≫ α else 0) - have hβ : c.inj i ≫ β = α := by simp [β] - simp only [← hβ, Category.assoc, h c.pt (by simp) β] - -lemma isSeparator_coproduct (hS : IsSeparating (Set.range S)) [HasCoproduct S] : - IsSeparator (∐ S) := - isSeparator_of_isColimit_cofan hS (colimit.isColimit _) - -end CategoryTheory.IsSeparating diff --git a/Mathlib/CategoryTheory/Generator/Indization.lean b/Mathlib/CategoryTheory/Generator/Indization.lean new file mode 100644 index 0000000000000..4c79e7d781802 --- /dev/null +++ b/Mathlib/CategoryTheory/Generator/Indization.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2025 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Generator.Basic +import Mathlib.CategoryTheory.Limits.Indization.Category + +/-! +# Separating set in the category of ind-objects + +We construct a separating set in the category of ind-objects. + +## Future work + +Once we have constructed zero morphisms in the category of ind-objects, we will be able to show +that under sufficient conditions, the category of ind-objects has a separating object. + +-/ + +universe v u + +namespace CategoryTheory + +open Limits + +section + +variable {C : Type u} [Category.{v} C] + +theorem Ind.isSeparating_range_yoneda : IsSeparating (Set.range (Ind.yoneda : C ⥤ _).obj) := by + refine fun X Y f g h => (cancel_epi (Ind.colimitPresentationCompYoneda X).hom).1 ?_ + exact colimit.hom_ext (fun i => by simp [← Category.assoc, h]) + +end + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Generator/Presheaf.lean b/Mathlib/CategoryTheory/Generator/Presheaf.lean index d12ab50d4882a..61042589d44a2 100644 --- a/Mathlib/CategoryTheory/Generator/Presheaf.lean +++ b/Mathlib/CategoryTheory/Generator/Presheaf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Generator.Coproduct +import Mathlib.CategoryTheory.Generator.Basic import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic /-! diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index fcea852442c19..7ce222e6e0218 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -43,8 +43,6 @@ such that `t' : V i j ×[U i] V i k ⟶ V j k ×[U j] V j i`. 10. `t' i j k ≫ t' j k i ≫ t' k i j = 𝟙 _`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] structure GlueData where J : Type v U : J → C diff --git a/Mathlib/CategoryTheory/Grothendieck.lean b/Mathlib/CategoryTheory/Grothendieck.lean index c987c94e3af65..799069e98d5d9 100644 --- a/Mathlib/CategoryTheory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Grothendieck.lean @@ -54,8 +54,6 @@ gives a category whose `base : X.base ⟶ Y.base` and `f.fiber : (F.map base).obj X.fiber ⟶ Y.fiber` -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): no such linter yet --- @[nolint has_nonempty_instance] structure Grothendieck where /-- The underlying object in `C` -/ base : C diff --git a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean index 7f4c9c681866e..9534f673817a0 100644 --- a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean @@ -8,7 +8,6 @@ import Mathlib.CategoryTheory.Groupoid.VertexGroup import Mathlib.CategoryTheory.Groupoid.Basic import Mathlib.CategoryTheory.Groupoid import Mathlib.Data.Set.Lattice -import Mathlib.Order.GaloisConnection /-! # Subgroupoid diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index aa8ebd6d96a7d..f022f13162297 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -31,7 +31,6 @@ variable (C : Type*) [Category C] namespace Idempotents --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- In a preadditive category `C`, when an object `X` decomposes as `X ≅ P ⨿ Q`, one may consider `P` as a direct factor of `X` and up to unique isomorphism, it is determined by the obvious idempotent `X ⟶ P ⟶ X` which is the projection onto `P` with kernel `Q`. More generally, diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 62109a61f38ae..e1ff3d18f906e 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -67,10 +67,8 @@ NB. Some authors include the empty category as connected, we do not. We instead are interested in categories with exactly one 'connected component'. -This allows us to show that the functor X ⨯ - preserves connected limits. - -See --/ +This allows us to show that the functor X ⨯ - preserves connected limits. -/ +@[stacks 002S] class IsConnected (J : Type u₁) [Category.{v₁} J] extends IsPreconnected J : Prop where [is_nonempty : Nonempty J] @@ -119,6 +117,13 @@ theorem IsPreconnected.of_any_functor_const_on_obj IsPreconnected J where iso_constant := fun F j' => ⟨NatIso.ofComponents fun j => eqToIso (h F j j')⟩ +instance IsPreconnected.prod [IsPreconnected J] [IsPreconnected K] : IsPreconnected (J × K) := by + refine .of_any_functor_const_on_obj (fun {a} F ⟨j, k⟩ ⟨j', k'⟩ => ?_) + exact (any_functor_const_on_obj (Prod.sectL J k ⋙ F) j j').trans + (any_functor_const_on_obj (Prod.sectR j' K ⋙ F) k k') + +instance IsConnected.prod [IsConnected J] [IsConnected K] : IsConnected (J × K) where + /-- If any functor to a discrete category is constant on objects, J is connected. The converse of `any_functor_const_on_obj`. -/ diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index 6e0dd598bc818..781198c903a31 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -42,10 +42,8 @@ open Category The inverse morphism is bundled. See also `CategoryTheory.Core` for the category with the same objects and isomorphisms playing -the role of morphisms. - -See . --/ +the role of morphisms. -/ +@[stacks 0017] structure Iso {C : Type u} [Category.{v} C] (X Y : C) where /-- The forward direction of an isomorphism. -/ hom : X ⟶ Y diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean index d5e68a50703d7..1b1231d4bd3c3 100644 --- a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean +++ b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean @@ -74,7 +74,7 @@ lemma colimit_no_zero_smul_divisor 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), this, 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) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean index 84ff0c9c14b61..43c2773df6cc1 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean @@ -123,19 +123,15 @@ instance limitSubobjectProduct_mono [HasLimitsOfSize.{w, w} C] (F : J ⥤ C) : Mono (limitSubobjectProduct F) := mono_comp _ _ -/-- Any category with products and equalizers has all limits. - -See . --/ +/-- Any category with products and equalizers has all limits. -/ +@[stacks 002N] theorem has_limits_of_hasEqualizers_and_products [HasProducts.{w} C] [HasEqualizers C] : HasLimitsOfSize.{w, w} C := { has_limits_of_shape := fun _ _ => { has_limit := fun F => hasLimit_of_equalizer_and_product F } } -/-- Any category with finite products and equalizers has all finite limits. - -See . --/ +/-- Any category with finite products and equalizers has all finite limits. -/ +@[stacks 002O] theorem hasFiniteLimits_of_hasEqualizers_and_finite_products [HasFiniteProducts C] [HasEqualizers C] : HasFiniteLimits C where out _ := { has_limit := fun F => hasLimit_of_equalizer_and_product F } @@ -343,19 +339,15 @@ instance colimitQuotientCoproduct_epi [HasColimitsOfSize.{w, w} C] (F : J ⥤ C) Epi (colimitQuotientCoproduct F) := epi_comp _ _ -/-- Any category with coproducts and coequalizers has all colimits. - -See . --/ +/-- Any category with coproducts and coequalizers has all colimits. -/ +@[stacks 002P] theorem has_colimits_of_hasCoequalizers_and_coproducts [HasCoproducts.{w} C] [HasCoequalizers C] : HasColimitsOfSize.{w, w} C where has_colimits_of_shape := fun _ _ => { has_colimit := fun F => hasColimit_of_coequalizer_and_coproduct F } -/-- Any category with finite coproducts and coequalizers has all finite colimits. - -See . --/ +/-- Any category with finite coproducts and coequalizers has all finite colimits. -/ +@[stacks 002Q] theorem hasFiniteColimits_of_hasCoequalizers_and_finite_coproducts [HasFiniteCoproducts C] [HasCoequalizers C] : HasFiniteColimits C where out _ := { has_colimit := fun F => hasColimit_of_coequalizer_and_coproduct F } diff --git a/Mathlib/CategoryTheory/Limits/ExactFunctor.lean b/Mathlib/CategoryTheory/Limits/ExactFunctor.lean index 811252d88fc96..c39bc0d0cfc33 100644 --- a/Mathlib/CategoryTheory/Limits/ExactFunctor.lean +++ b/Mathlib/CategoryTheory/Limits/ExactFunctor.lean @@ -28,7 +28,6 @@ section variable (C) (D) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled left-exact functors. -/ def LeftExactFunctor := FullSubcategory fun F : C ⥤ D => PreservesFiniteLimits F @@ -49,7 +48,6 @@ instance : (LeftExactFunctor.forget C D).Full := instance : (LeftExactFunctor.forget C D).Faithful := FullSubcategory.faithful _ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled right-exact functors. -/ def RightExactFunctor := FullSubcategory fun F : C ⥤ D => PreservesFiniteColimits F @@ -70,7 +68,6 @@ instance : (RightExactFunctor.forget C D).Full := instance : (RightExactFunctor.forget C D).Faithful := FullSubcategory.faithful _ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled exact functors. -/ def ExactFunctor := FullSubcategory fun F : C ⥤ D => diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 5124dc554a97f..3431f905c7954 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -82,10 +82,8 @@ variable {D : Type u₂} [Category.{v₂} D] /-- A functor `F : C ⥤ D` is final if for every `d : D`, the comma category of morphisms `d ⟶ F.obj c` -is connected. - -See --/ +is connected. -/ +@[stacks 04E6] class Final (F : C ⥤ D) : Prop where out (d : D) : IsConnected (StructuredArrow d F) @@ -337,11 +335,8 @@ section variable (G) /-- When `F : C ⥤ D` is final, and `G : D ⥤ E` has a colimit, then `F ⋙ G` has a colimit also and -`colimit (F ⋙ G) ≅ colimit G` - -https://stacks.math.columbia.edu/tag/04E7 --/ -@[simps! (config := .lemmasOnly)] +`colimit (F ⋙ G) ≅ colimit G`. -/ +@[simps! (config := .lemmasOnly), stacks 04E7] def colimitIso [HasColimit G] : colimit (F ⋙ G) ≅ colimit G := asIso (colimit.pre G F) @@ -686,11 +681,8 @@ section variable (G) /-- When `F : C ⥤ D` is initial, and `G : D ⥤ E` has a limit, then `F ⋙ G` has a limit also and -`limit (F ⋙ G) ≅ limit G` - -https://stacks.math.columbia.edu/tag/04E7 --/ -@[simps! (config := .lemmasOnly)] +`limit (F ⋙ G) ≅ limit G`. -/ +@[simps! (config := .lemmasOnly), stacks 04E7] def limitIso [HasLimit G] : limit (F ⋙ G) ≅ limit G := (asIso (limit.pre G F)).symm @@ -1027,4 +1019,17 @@ instance Grothendieck.final_pre [hG : Final G] : (Grothendieck.pre F G).Final := end Grothendieck +section Prod + +variable {C : Type u₁} [Category.{v₁} C] +variable {D : Type u₂} [Category.{v₂} D] +variable {C' : Type u₃} [Category.{v₃} C'] +variable {D' : Type u₄} [Category.{v₄} D'] +variable (F : C ⥤ D) (G : C' ⥤ D') + +instance [F.Final] [G.Final] : (F.prod G).Final where + out := fun ⟨d, d'⟩ => isConnected_of_equivalent (StructuredArrow.prodEquivalence d d' F G).symm + +end Prod + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index 8ebca957d879b..f5172396cb6f6 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -71,7 +71,6 @@ variable {F : J ⥤ C} section Limit /-- `LimitCone F` contains a cone over `F` together with the information that it is a limit. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed; linter not ported yet structure LimitCone (F : J ⥤ C) where /-- The cone itself -/ cone : Cone F @@ -583,7 +582,6 @@ section Colimit /-- `ColimitCocone F` contains a cocone over `F` together with the information that it is a colimit. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed; linter not ported yet structure ColimitCocone (F : J ⥤ C) where /-- The cocone itself -/ cocone : Cocone F diff --git a/Mathlib/CategoryTheory/Limits/Indization/Category.lean b/Mathlib/CategoryTheory/Limits/Indization/Category.lean index 4de4c45fb1b0f..c1cf76d2a04fe 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Category.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Category.lean @@ -7,6 +7,7 @@ import Mathlib.CategoryTheory.Functor.Flat import Mathlib.CategoryTheory.Limits.Constructions.Filtered import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.Indization.LocallySmall +import Mathlib.CategoryTheory.Limits.Indization.ParallelPair import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.Products @@ -16,6 +17,8 @@ import Mathlib.CategoryTheory.Limits.Indization.Products We define the `v`-category of Ind-objects of a category `C`, called `Ind C`, as well as the functors `Ind.yoneda : C ⥤ Ind C` and `Ind.inclusion C : Ind C ⥤ Cᵒᵖ ⥤ Type v`. +For a small filtered category `I`, we also define `Ind.lim I : (I ⥤ C) ⥤ Ind C`. + This file will mainly collect results about ind-objects (stated in terms of `IsIndObject`) and reinterpret them in terms of `Ind C`. @@ -28,7 +31,8 @@ Adopting the theorem numbering of [Kashiwara2006], we show that * if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)), * if `C` has products indexed by `α`, then `Ind C` has products indexed by `α`, and the functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` creates such products (6.1.17) -* the functor `C ⥤ Ind C` preserves small limits. +* the functor `C ⥤ Ind C` preserves small limits, +* if `C` has coequalizers, then `Ind C` has coequalizers (6.1.18(i)). More limit-colimit properties will follow. @@ -41,7 +45,7 @@ Note that: * [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Chapter 6 -/ -universe v u +universe w v u namespace CategoryTheory @@ -157,6 +161,39 @@ instance : RepresentablyCoflat (Ind.yoneda (C := C)) := by noncomputable instance : PreservesFiniteColimits (Ind.yoneda (C := C)) := preservesFiniteColimits_of_coflat _ +/-- This is the functor `(I ⥤ C) ⥤ Ind C` that sends a functor `F` to `colim (Y ∘ F)`, where `Y` +is the Yoneda embedding. It is known as "ind-lim" and denoted `“colim”` in [Kashiwara2006]. -/ +protected noncomputable def Ind.lim (I : Type v) [SmallCategory I] [IsFiltered I] : + (I ⥤ C) ⥤ Ind C := + (whiskeringRight _ _ _).obj Ind.yoneda ⋙ colim + +/-- Computing ind-lims in `Ind C` is the same as computing them in `Cᵒᵖ ⥤ Type v`. -/ +noncomputable def Ind.limCompInclusion {I : Type v} [SmallCategory I] [IsFiltered I] : + Ind.lim I ⋙ Ind.inclusion C ≅ (whiskeringRight _ _ _).obj yoneda ⋙ colim := calc + Ind.lim I ⋙ Ind.inclusion C + ≅ (whiskeringRight _ _ _).obj Ind.yoneda ⋙ colim ⋙ Ind.inclusion C := Functor.associator _ _ _ + _ ≅ (whiskeringRight _ _ _).obj Ind.yoneda ⋙ + (whiskeringRight _ _ _).obj (Ind.inclusion C) ⋙ colim := + isoWhiskerLeft _ (preservesColimitNatIso _) + _ ≅ ((whiskeringRight _ _ _).obj Ind.yoneda ⋙ + (whiskeringRight _ _ _).obj (Ind.inclusion C)) ⋙ colim := (Functor.associator _ _ _).symm + _ ≅ (whiskeringRight _ _ _).obj (Ind.yoneda ⋙ Ind.inclusion C) ⋙ colim := + isoWhiskerRight (whiskeringRightObjCompIso _ _) colim + _ ≅ (whiskeringRight _ _ _).obj yoneda ⋙ colim := + isoWhiskerRight ((whiskeringRight _ _ _).mapIso (Ind.yonedaCompInclusion)) colim + +instance {α : Type w} [SmallCategory α] [FinCategory α] [HasLimitsOfShape α C] {I : Type v} + [SmallCategory I] [IsFiltered I] : + PreservesLimitsOfShape α (Ind.lim I : (I ⥤ C) ⥤ _) := + haveI : PreservesLimitsOfShape α (Ind.lim I ⋙ Ind.inclusion C) := + preservesLimitsOfShape_of_natIso Ind.limCompInclusion.symm + preservesLimitsOfShape_of_reflects_of_preserves _ (Ind.inclusion C) + +instance {α : Type w} [SmallCategory α] [FinCategory α] [HasColimitsOfShape α C] {I : Type v} + [SmallCategory I] [IsFiltered I] : + PreservesColimitsOfShape α (Ind.lim I : (I ⥤ C) ⥤ _) := + inferInstanceAs (PreservesColimitsOfShape α (_ ⋙ colim)) + instance {α : Type v} [Finite α] [HasColimitsOfShape (Discrete α) C] : HasColimitsOfShape (Discrete α) (Ind C) := by refine ⟨fun F => ?_⟩ @@ -174,11 +211,43 @@ instance {α : Type v} [Finite α] [HasColimitsOfShape (Discrete α) C] : -- ``` -- from the fact that finite limits commute with filtered colimits and from the fact that -- `Ind.yoneda` preserves finite colimits. - apply hasColimitOfIso iso.symm + exact hasColimitOfIso iso.symm instance [HasFiniteCoproducts C] : HasCoproducts.{v} (Ind C) := have : HasFiniteCoproducts (Ind C) := ⟨fun _ => hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift)⟩ hasCoproducts_of_finite_and_filtered +/-- Given an `IndParallelPairPresentation f g`, we can understand the parallel pair `(f, g)` as +the colimit of `(P.φ, P.ψ)` in `Ind C`. -/ +noncomputable def IndParallelPairPresentation.parallelPairIsoParallelPairCompIndYoneda + {A B : Ind C} {f g : A ⟶ B} + (P : IndParallelPairPresentation ((Ind.inclusion _).map f) ((Ind.inclusion _).map g)) : + parallelPair f g ≅ parallelPair P.φ P.ψ ⋙ Ind.lim P.I := + ((whiskeringRight WalkingParallelPair _ _).obj (Ind.inclusion C)).preimageIso <| + diagramIsoParallelPair _ ≪≫ + P.parallelPairIsoParallelPairCompYoneda ≪≫ + isoWhiskerLeft (parallelPair _ _) Ind.limCompInclusion.symm + +instance [HasColimitsOfShape WalkingParallelPair C] : + HasColimitsOfShape WalkingParallelPair (Ind C) := by + refine ⟨fun F => ?_⟩ + obtain ⟨P⟩ := nonempty_indParallelPairPresentation (F.obj WalkingParallelPair.zero).2 + (F.obj WalkingParallelPair.one).2 (Ind.inclusion _ |>.map <| F.map WalkingParallelPairHom.left) + (Ind.inclusion _ |>.map <| F.map WalkingParallelPairHom.right) + exact hasColimitOfIso (diagramIsoParallelPair _ ≪≫ P.parallelPairIsoParallelPairCompIndYoneda) + +/-- A way to understand morphisms in `Ind C`: every morphism is induced by a natural transformation +of diagrams. -/ +theorem Ind.exists_nonempty_arrow_mk_iso_ind_lim {A B : Ind C} {f : A ⟶ B} : + ∃ (I : Type v) (_ : SmallCategory I) (_ : IsFiltered I) (F G : I ⥤ C) (φ : F ⟶ G), + Nonempty (Arrow.mk f ≅ Arrow.mk ((Ind.lim _).map φ)) := by + obtain ⟨P⟩ := nonempty_indParallelPairPresentation A.2 B.2 + (Ind.inclusion _ |>.map f) (Ind.inclusion _ |>.map f) + refine ⟨P.I, inferInstance, inferInstance, P.F₁, P.F₂, P.φ, ⟨Arrow.isoMk ?_ ?_ ?_⟩⟩ + · exact P.parallelPairIsoParallelPairCompIndYoneda.app WalkingParallelPair.zero + · exact P.parallelPairIsoParallelPairCompIndYoneda.app WalkingParallelPair.one + · simpa using + (P.parallelPairIsoParallelPairCompIndYoneda.hom.naturality WalkingParallelPairHom.left).symm + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean index ef9fd1afa688d..880fad0485062 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Equalizers.lean @@ -4,16 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits +import Mathlib.CategoryTheory.Limits.FullSubcategory +import Mathlib.CategoryTheory.Limits.Indization.ParallelPair /-! # Equalizers of ind-objects -The eventual goal of this file is to show that if a category `C` has equalizers, then ind-objects -are closed under equalizers. For now, it contains one of the main prerequisites, namely we show -that under sufficient conditions the limit of a diagram in `I ⥤ C` taken in `Cᵒᵖ ⥤ Type v` is an -ind-object. +We show that if a category `C` has equalizers, then ind-objects are closed under equalizers. -## References +# References * [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Section 6.1 -/ @@ -60,4 +59,21 @@ theorem isIndObject_limit_comp_yoneda_comp_colim end +/-- If `C` has equalizers. then ind-objects are closed under equalizers. + +This is Proposition 6.1.17(i) in [Kashiwara2006]. +-/ +theorem closedUnderLimitsOfShape_walkingParallelPair_isIndObject [HasEqualizers C] : + ClosedUnderLimitsOfShape WalkingParallelPair (IsIndObject (C := C)) := by + apply closedUnderLimitsOfShape_of_limit + intro F hF h + obtain ⟨P⟩ := nonempty_indParallelPairPresentation (h WalkingParallelPair.zero) + (h WalkingParallelPair.one) (F.map WalkingParallelPairHom.left) + (F.map WalkingParallelPairHom.right) + exact IsIndObject.map + (HasLimit.isoOfNatIso (P.parallelPairIsoParallelPairCompYoneda.symm ≪≫ + (diagramIsoParallelPair _).symm)).hom + (isIndObject_limit_comp_yoneda_comp_colim (parallelPair P.φ P.ψ) + (fun i => isIndObject_limit_comp_yoneda _)) + end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Indization/ParallelPair.lean b/Mathlib/CategoryTheory/Limits/Indization/ParallelPair.lean new file mode 100644 index 0000000000000..68c898bc32634 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Indization/ParallelPair.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2025 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Comma.Final +import Mathlib.CategoryTheory.Limits.Indization.IndObject + +/-! +# Parallel pairs of natural transformations between ind-objects + +We show that if `A` and `B` are ind-objects and `f` and `g` are natural transformations between +`A` and `B`, then there is a small filtered category `I` such that `A`, `B`, `f` and `g` are +commonly presented by diagrams and natural transformations in `I ⥤ C`. + + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Proposition 6.1.15 (though + our proof is more direct). +-/ + +universe v₁ v₂ v₃ u₁ u₂ u₃ + +namespace CategoryTheory + +open Limits + +variable {C : Type u₁} [Category.{v₁} C] + +/-- Structure containing data exhibiting two parallel natural transformations `f` and `g` between +presheaves `A` and `B` as induced by a natural transformation in a functor category exhibiting +`A` and `B` as ind-objects. -/ +structure IndParallelPairPresentation {A B : Cᵒᵖ ⥤ Type v₁} (f g : A ⟶ B) where + /-- The indexing category. -/ + I : Type v₁ + /-- Category instance on the indexing category. -/ + [ℐ : SmallCategory I] + [hI : IsFiltered I] + /-- The diagram presenting `A`. -/ + F₁ : I ⥤ C + /-- The diagram presenting `B`. -/ + F₂ : I ⥤ C + /-- The cocone on `F₁` with apex `A`. -/ + ι₁ : F₁ ⋙ yoneda ⟶ (Functor.const I).obj A + /-- The cocone on `F₁` with apex `A` is a colimit cocone. -/ + isColimit₁ : IsColimit (Cocone.mk A ι₁) + /-- The cocone on `F₂` with apex `B`. -/ + ι₂ : F₂ ⋙ yoneda ⟶ (Functor.const I).obj B + /-- The cocone on `F₂` with apex `B` is a colimit cocone. -/ + isColimit₂ : IsColimit (Cocone.mk B ι₂) + /-- The natural transformation presenting `f`. -/ + φ : F₁ ⟶ F₂ + /-- The natural transformation presenting `g`. -/ + ψ : F₁ ⟶ F₂ + /-- `f` is in fact presented by `φ`. -/ + hf : f = IsColimit.map isColimit₁ (Cocone.mk B ι₂) (whiskerRight φ yoneda) + /-- `g` is in fact presented by `ψ`. -/ + hg : g = IsColimit.map isColimit₁ (Cocone.mk B ι₂) (whiskerRight ψ yoneda) + +instance {A B : Cᵒᵖ ⥤ Type v₁} {f g : A ⟶ B} (P : IndParallelPairPresentation f g) : + SmallCategory P.I := P.ℐ +instance {A B : Cᵒᵖ ⥤ Type v₁} {f g : A ⟶ B} (P : IndParallelPairPresentation f g) : + IsFiltered P.I := P.hI + +namespace NonemptyParallelPairPresentationAux + +variable {A B : Cᵒᵖ ⥤ Type v₁} (f g : A ⟶ B) (P₁ : IndObjectPresentation A) + (P₂ : IndObjectPresentation B) + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev K : Type v₁ := + Comma ((P₁.toCostructuredArrow ⋙ CostructuredArrow.map f).prod' + (P₁.toCostructuredArrow ⋙ CostructuredArrow.map g)) + (P₂.toCostructuredArrow.prod' P₂.toCostructuredArrow) + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev F₁ : K f g P₁ P₂ ⥤ C := Comma.fst _ _ ⋙ P₁.F +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev F₂ : K f g P₁ P₂ ⥤ C := Comma.snd _ _ ⋙ P₂.F + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev ι₁ : F₁ f g P₁ P₂ ⋙ yoneda ⟶ (Functor.const (K f g P₁ P₂)).obj A := + whiskerLeft (Comma.fst _ _) P₁.ι + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +noncomputable abbrev isColimit₁ : IsColimit (Cocone.mk A (ι₁ f g P₁ P₂)) := + (Functor.Final.isColimitWhiskerEquiv _ _).symm P₁.isColimit + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +abbrev ι₂ : F₂ f g P₁ P₂ ⋙ yoneda ⟶ (Functor.const (K f g P₁ P₂)).obj B := + whiskerLeft (Comma.snd _ _) P₂.ι + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +noncomputable abbrev isColimit₂ : IsColimit (Cocone.mk B (ι₂ f g P₁ P₂)) := + (Functor.Final.isColimitWhiskerEquiv _ _).symm P₂.isColimit + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +def ϕ : F₁ f g P₁ P₂ ⟶ F₂ f g P₁ P₂ where + app h := h.hom.1.left + naturality _ _ h := by + have := h.w + simp only [Functor.prod'_obj, Functor.comp_obj, prod_Hom, Functor.prod'_map, Functor.comp_map, + prod_comp, Prod.mk.injEq, CostructuredArrow.hom_eq_iff, CostructuredArrow.map_obj_left, + IndObjectPresentation.toCostructuredArrow_obj_left, CostructuredArrow.comp_left, + CostructuredArrow.map_map_left, IndObjectPresentation.toCostructuredArrow_map_left] at this + exact this.1 + +theorem hf : f = IsColimit.map (isColimit₁ f g P₁ P₂) + (Cocone.mk B (ι₂ f g P₁ P₂)) (whiskerRight (ϕ f g P₁ P₂) yoneda) := by + refine (isColimit₁ f g P₁ P₂).hom_ext (fun i => ?_) + rw [IsColimit.ι_map] + simpa using i.hom.1.w.symm + +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +def ψ : F₁ f g P₁ P₂ ⟶ F₂ f g P₁ P₂ where + app h := h.hom.2.left + naturality _ _ h := by + have := h.w + simp only [Functor.prod'_obj, Functor.comp_obj, prod_Hom, Functor.prod'_map, Functor.comp_map, + prod_comp, Prod.mk.injEq, CostructuredArrow.hom_eq_iff, CostructuredArrow.map_obj_left, + IndObjectPresentation.toCostructuredArrow_obj_left, CostructuredArrow.comp_left, + CostructuredArrow.map_map_left, IndObjectPresentation.toCostructuredArrow_map_left] at this + exact this.2 + +theorem hg : g = IsColimit.map (isColimit₁ f g P₁ P₂) + (Cocone.mk B (ι₂ f g P₁ P₂)) (whiskerRight (ψ f g P₁ P₂) yoneda) := by + refine (isColimit₁ f g P₁ P₂).hom_ext (fun i => ?_) + rw [IsColimit.ι_map] + simpa using i.hom.2.w.symm + +attribute [local instance] Comma.isFiltered_of_final in +/-- Implementation; see `nonempty_indParallelPairPresentation`. -/ +noncomputable def presentation : IndParallelPairPresentation f g where + I := K f g P₁ P₂ + F₁ := F₁ f g P₁ P₂ + F₂ := F₂ f g P₁ P₂ + ι₁ := ι₁ f g P₁ P₂ + isColimit₁ := isColimit₁ f g P₁ P₂ + ι₂ := ι₂ f g P₁ P₂ + isColimit₂ := isColimit₂ f g P₁ P₂ + φ := ϕ f g P₁ P₂ + ψ := ψ f g P₁ P₂ + hf := hf f g P₁ P₂ + hg := hg f g P₁ P₂ + +end NonemptyParallelPairPresentationAux + +theorem nonempty_indParallelPairPresentation {A B : Cᵒᵖ ⥤ Type v₁} (hA : IsIndObject A) + (hB : IsIndObject B) (f g : A ⟶ B) : Nonempty (IndParallelPairPresentation f g) := + ⟨NonemptyParallelPairPresentationAux.presentation f g hA.presentation hB.presentation⟩ + +namespace IndParallelPairPresentation + +/-- Given an `IndParallelPairPresentation f g`, we can understand the parallel pair `(f, g)` +as the colimit of `(P.φ, P.ψ)` in `Cᵒᵖ ⥤ Type v`. -/ +noncomputable def parallelPairIsoParallelPairCompYoneda {A B : Cᵒᵖ ⥤ Type v₁} {f g : A ⟶ B} + (P : IndParallelPairPresentation f g) : + parallelPair f g ≅ parallelPair P.φ P.ψ ⋙ (whiskeringRight _ _ _).obj yoneda ⋙ colim := + parallelPair.ext + (P.isColimit₁.coconePointUniqueUpToIso (colimit.isColimit _)) + (P.isColimit₂.coconePointUniqueUpToIso (colimit.isColimit _)) + (P.isColimit₁.hom_ext (fun j => by + simp [P.hf, P.isColimit₁.ι_map_assoc, P.isColimit₁.comp_coconePointUniqueUpToIso_hom_assoc, + P.isColimit₂.comp_coconePointUniqueUpToIso_hom])) + (P.isColimit₁.hom_ext (fun j => by + simp [P.hg, P.isColimit₁.ι_map_assoc, P.isColimit₁.comp_coconePointUniqueUpToIso_hom_assoc, + P.isColimit₂.comp_coconePointUniqueUpToIso_hom])) + +end IndParallelPairPresentation + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/IsLimit.lean b/Mathlib/CategoryTheory/Limits/IsLimit.lean index a15e2af7ecdaa..aead84cff4f8a 100644 --- a/Mathlib/CategoryTheory/Limits/IsLimit.lean +++ b/Mathlib/CategoryTheory/Limits/IsLimit.lean @@ -46,11 +46,8 @@ variable {C : Type u₃} [Category.{v₃} C] variable {F : J ⥤ C} /-- A cone `t` on `F` is a limit cone if each cone on `F` admits a unique -cone morphism to `t`. - -See . - -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] +cone morphism to `t`. -/ +@[stacks 002E] structure IsLimit (t : Cone F) where /-- There is a morphism from any cone point to `t.pt` -/ lift : ∀ s : Cone F, s.pt ⟶ t.pt @@ -496,11 +493,8 @@ end end IsLimit /-- A cocone `t` on `F` is a colimit cocone if each cocone on `F` admits a unique -cocone morphism from `t`. - -See . --/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] +cocone morphism from `t`. -/ +@[stacks 002F] structure IsColimit (t : Cocone F) where /-- `t.pt` maps to all other cocone covertices -/ desc : ∀ s : Cocone F, t.pt ⟶ s.pt diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Grothendieck.lean b/Mathlib/CategoryTheory/Limits/Preserves/Grothendieck.lean new file mode 100644 index 0000000000000..c3504b261f99f --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Preserves/Grothendieck.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2024 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer +-/ +import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic +import Mathlib.CategoryTheory.Limits.Shapes.Grothendieck + +/-! +# Colimits on Grothendieck constructions preserving limits + +We characterize the condition in which colimits on Grothendieck constructions preserve limits: By +preserving limits on the Grothendieck construction's base category as well as on each of its fibers. +-/ + + +universe v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ + +namespace CategoryTheory + +namespace Limits + +noncomputable section + +variable {C : Type u₁} [Category.{v₁} C] +variable {H : Type u₂} [Category.{v₂} H] +variable {J : Type u₃} [Category.{v₃} J] +variable {F : C ⥤ Cat.{v₄, u₄}} + +/-- If `colim` on each fiber `F.obj c` of a functor `F : C ⥤ Cat` preserves limits of shape `J`, +then the fiberwise colimit of the limit of a functor `K : J ⥤ Grothendieck F ⥤ H` is naturally +isomorphic to taking the limit of the composition `K ⋙ fiberwiseColim F H`. -/ +@[simps!] +def fiberwiseColimitLimitIso (K : J ⥤ Grothendieck F ⥤ H) + [∀ (c : C), HasColimitsOfShape (↑(F.obj c)) H] [HasLimitsOfShape J H] + [∀ c, PreservesLimitsOfShape J (colim (J := F.obj c) (C := H))] : + fiberwiseColimit (limit K) ≅ limit (K ⋙ fiberwiseColim F H) := + NatIso.ofComponents + (fun c => HasColimit.isoOfNatIso + (limitCompWhiskeringLeftIsoCompLimit K (Grothendieck.ι F c)).symm ≪≫ + preservesLimitIso colim _ ≪≫ + HasLimit.isoOfNatIso + (Functor.associator _ _ _ ≪≫ + isoWhiskerLeft _ (fiberwiseColimCompEvaluationIso _).symm ≪≫ + (Functor.associator _ _ _).symm) ≪≫ + (limitObjIsoLimitCompEvaluation _ c).symm) + fun {c₁ c₂} f => by + simp only [fiberwiseColimit_obj, fiberwiseColimit_map, Iso.trans_hom, Iso.symm_hom, + Category.assoc, limitObjIsoLimitCompEvaluation_inv_limit_map] + apply colimit.hom_ext + intro d + simp only [← Category.assoc] + congr 1 + apply limit.hom_ext + intro e + simp [← NatTrans.comp_app_assoc] + +variable (C) (F) in +/-- If `colim` on a category `C` preserves limits of shape `J` and if it does so for `colim` on +every `F.obj c` for a functor `F : C ⥤ Cat`, then `colim` on `Grothendieck F` also preserves limits +of shape `J`. -/ +instance preservesLimitsOfShape_colim_grothendieck [HasColimitsOfShape C H] [HasLimitsOfShape J H] + [∀ c, HasColimitsOfShape (↑(F.obj c)) H] [PreservesLimitsOfShape J (colim (J := C) (C := H))] + [∀ c, PreservesLimitsOfShape J (colim (J := F.obj c) (C := H))] : + PreservesLimitsOfShape J (colim (J := Grothendieck F) (C := H)) := by + constructor + intro K + let i₂ := calc colimit (limit K) + _ ≅ colimit (fiberwiseColimit (limit K)) := (colimitFiberwiseColimitIso _).symm + _ ≅ colimit (limit (K ⋙ fiberwiseColim _ _)) := + HasColimit.isoOfNatIso (fiberwiseColimitLimitIso _) + _ ≅ limit ((K ⋙ fiberwiseColim _ _) ⋙ colim) := + preservesLimitIso colim (K ⋙ fiberwiseColim _ _) + _ ≅ limit (K ⋙ colim) := + HasLimit.isoOfNatIso + (Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ fiberwiseColimCompColimIso) + haveI : IsIso (limit.post K colim) := by + convert Iso.isIso_hom i₂ + ext + simp only [colim_obj, Functor.comp_obj, limit.post_π, colim_map, Iso.instTransIso_trans, + Iso.trans_assoc, Iso.trans_hom, Category.assoc, HasLimit.isoOfNatIso_hom_π, + fiberwiseColim_obj, isoWhiskerLeft_hom, NatTrans.comp_app, Functor.associator_hom_app, + whiskerLeft_app, fiberwiseColimCompColimIso_hom_app, Category.id_comp, + preservesLimitIso_hom_π_assoc, i₂] + ext + simp only [ι_colimMap, Trans.trans, Iso.symm_hom, ι_colimitFiberwiseColimitIso_inv_assoc, + HasColimit.isoOfNatIso_ι_hom_assoc, fiberwiseColimit_obj, fiberwiseColimitLimitIso_hom_app, + ι_colimMap_assoc, Category.assoc, limitObjIsoLimitCompEvaluation_inv_π_app_assoc, + Functor.comp_obj, fiberwiseColim_obj, HasLimit.isoOfNatIso_hom_π_assoc, + whiskeringLeft_obj_obj, colim_obj, evaluation_obj_obj, Iso.trans_hom, isoWhiskerLeft_hom, + NatTrans.comp_app, Functor.associator_hom_app, whiskerLeft_app, + fiberwiseColimCompEvaluationIso_inv_app, Functor.associator_inv_app, Category.comp_id, + Category.id_comp, preservesLimitIso_hom_π_assoc, colim_map, Grothendieck.ι_obj, + ι_colimitFiberwiseColimitIso_hom] + simp [← Category.assoc, ← NatTrans.comp_app] + apply preservesLimit_of_isIso_post + +end + +end Limits + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean index 5f3c1866a623d..3ee93b6a24866 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/BinaryProducts.lean @@ -46,14 +46,14 @@ open WalkingPair /-- The equivalence swapping left and right. -/ def WalkingPair.swap : WalkingPair ≃ WalkingPair where - toFun j := match j with + toFun | left => right | right => left - invFun j := match j with + invFun | left => right | right => left - left_inv j := by cases j; repeat rfl - right_inv j := by cases j; repeat rfl + left_inv j := by cases j <;> rfl + right_inv j := by cases j <;> rfl @[simp] theorem WalkingPair.swap_apply_left : WalkingPair.swap left = right := @@ -74,13 +74,13 @@ theorem WalkingPair.swap_symm_apply_ff : WalkingPair.swap.symm right = left := /-- An equivalence from `WalkingPair` to `Bool`, sometimes useful when reindexing limits. -/ def WalkingPair.equivBool : WalkingPair ≃ Bool where - toFun j := match j with + toFun | left => true | right => false -- to match equiv.sum_equiv_sigma_bool invFun b := Bool.recOn b right left - left_inv j := by cases j; repeat rfl - right_inv b := by cases b; repeat rfl + left_inv j := by cases j <;> rfl + right_inv b := by cases b <;> rfl @[simp] theorem WalkingPair.equivBool_apply_left : WalkingPair.equivBool left = true := @@ -136,7 +136,7 @@ attribute [local aesop safe tactic (rule_sets := [CategoryTheory])] /-- The natural transformation between two functors out of the walking pair, specified by its components. -/ def mapPair : F ⟶ G where - app j := match j with + app | ⟨left⟩ => f | ⟨right⟩ => g naturality := fun ⟨X⟩ ⟨Y⟩ ⟨⟨u⟩⟩ => by aesop_cat @@ -837,17 +837,13 @@ end CoprodLemmas variable (C) -/-- `HasBinaryProducts` represents a choice of product for every pair of objects. - -See . --/ +/-- `HasBinaryProducts` represents a choice of product for every pair of objects. -/ +@[stacks 001T] abbrev HasBinaryProducts := HasLimitsOfShape (Discrete WalkingPair) C -/-- `HasBinaryCoproducts` represents a choice of coproduct for every pair of objects. - -See . --/ +/-- `HasBinaryCoproducts` represents a choice of coproduct for every pair of objects. -/ +@[stacks 04AP] abbrev HasBinaryCoproducts := HasColimitsOfShape (Discrete WalkingPair) C diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index 29d7bf37aa103..ef037726a2d00 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -67,7 +67,6 @@ open scoped Classical in * morphisms `π j : pt ⟶ F j` and `ι j : F j ⟶ pt` for each `j`, * such that `ι j ≫ π j'` is the identity when `j = j'` and zero otherwise. -/ --- @[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed structure Bicone (F : J → C) where pt : C π : ∀ j, pt ⟶ F j @@ -245,7 +244,6 @@ theorem π_of_isColimit {f : J → C} {t : Bicone f} (ht : IsColimit t.toCocone) simp [t.ι_π] /-- Structure witnessing that a bicone is both a limit cone and a colimit cocone. -/ --- @[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed structure IsBilimit {F : J → C} (B : Bicone F) where isLimit : IsLimit B.toCone isColimit : IsColimit B.toCocone @@ -313,7 +311,6 @@ end Bicone /-- A bicone over `F : J → C`, which is both a limit cone and a colimit cocone. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed; linter not ported yet structure LimitBicone (F : J → C) where bicone : Bicone F isBilimit : bicone.IsBilimit @@ -1103,7 +1100,6 @@ variable {C} maps from `X` to both `P` and `Q`, and maps from both `P` and `Q` to `X`, so that `inl ≫ fst = 𝟙 P`, `inl ≫ snd = 0`, `inr ≫ fst = 0`, and `inr ≫ snd = 𝟙 Q` -/ --- @[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed structure BinaryBicone (P Q : C) where pt : C fst : pt ⟶ P @@ -1353,7 +1349,6 @@ def toBinaryBiconeIsColimit {X Y : C} (b : Bicone (pairFunction X Y)) : end Bicone /-- Structure witnessing that a binary bicone is a limit cone and a limit cocone. -/ --- @[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed structure BinaryBicone.IsBilimit {P Q : C} (b : BinaryBicone P Q) where isLimit : IsLimit b.toCone isColimit : IsColimit b.toCocone @@ -1380,7 +1375,6 @@ def Bicone.toBinaryBiconeIsBilimit {X Y : C} (b : Bicone (pairFunction X Y)) : /-- A bicone over `P Q : C`, which is both a limit cone and a colimit cocone. -/ --- @[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed structure BinaryBiproductData (P Q : C) where bicone : BinaryBicone P Q isBilimit : bicone.IsBilimit diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean index 6f0db0a360ddb..9064ae11e1cb9 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean @@ -219,9 +219,7 @@ instance sequentialFunctor_initial : (sequentialFunctor J).Initial where · right exact ⟨CostructuredArrow.homMk (homOfLE h).op rfl⟩ -/-- -This is proved in https://stacks.math.columbia.edu/tag/0032 --/ +@[stacks 0032] proof_wanted preorder_of_cofiltered (J : Type*) [Category J] [IsCofiltered J] : ∃ (I : Type*) (_ : Preorder I) (_ : IsCofiltered I) (F : I ⥤ J), F.Initial diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean index fdf4259edc924..20e34c229d380 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean @@ -63,6 +63,8 @@ inductive WalkingParallelPair : Type open WalkingParallelPair +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- The type family of morphisms for the diagram indexing a (co)equalizer. -/ inductive WalkingParallelPairHom : WalkingParallelPair → WalkingParallelPair → Type | left : WalkingParallelPairHom zero one @@ -70,10 +72,6 @@ inductive WalkingParallelPairHom : WalkingParallelPair → WalkingParallelPair | id (X : WalkingParallelPair) : WalkingParallelPairHom X X deriving DecidableEq -/- Porting note: this simplifies using walkingParallelPairHom_id; replacement is below; -simpNF still complains of striking this from the simp list -/ -attribute [-simp, nolint simpNF] WalkingParallelPairHom.id.sizeOf_spec - /-- Satisfying the inhabited linter -/ instance : Inhabited (WalkingParallelPairHom zero one) where default := WalkingParallelPairHom.left @@ -114,11 +112,6 @@ instance walkingParallelPairHomCategory : SmallCategory WalkingParallelPair wher theorem walkingParallelPairHom_id (X : WalkingParallelPair) : WalkingParallelPairHom.id X = 𝟙 X := rfl --- Porting note: simpNF asked me to do this because the LHS of the non-primed version reduced -@[simp] -theorem WalkingParallelPairHom.id.sizeOf_spec' (X : WalkingParallelPair) : - (WalkingParallelPairHom._sizeOf_inst X X).sizeOf (𝟙 X) = 1 + sizeOf X := by cases X <;> rfl - /-- The functor `WalkingParallelPair ⥤ WalkingParallelPairᵒᵖ` sending left to left and right to right. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Grothendieck.lean b/Mathlib/CategoryTheory/Limits/Shapes/Grothendieck.lean index ff2c149902ac8..6e478aa734361 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Grothendieck.lean @@ -75,6 +75,18 @@ def fiberwiseColimit : C ⥤ H where conv_rhs => enter [2, 1]; rw [eqToHom_map (F.map (𝟙 Z))] conv_rhs => rw [eqToHom_trans, eqToHom_trans] +variable (H) (F) in +/-- Similar to `colimit` and `colim`, taking fiberwise colimits is a functor +`(Grothendieck F ⥤ H) ⥤ (C ⥤ H)` between functor categories. -/ +@[simps] +def fiberwiseColim [∀ c, HasColimitsOfShape (F.obj c) H] : (Grothendieck F ⥤ H) ⥤ (C ⥤ H) where + obj G := fiberwiseColimit G + map α := + { app := fun c => colim.map (whiskerLeft _ α) + naturality := fun c₁ c₂ f => by apply colimit.hom_ext; simp } + map_id G := by ext; simp; apply Functor.map_id colim + map_comp α β := by ext; simp; apply Functor.map_comp colim + /-- Every functor `G : Grothendieck F ⥤ H` induces a natural transformation from `G` to the composition of the forgetful functor on `Grothendieck F` with the fiberwise colimit on `G`. -/ @[simps] @@ -111,9 +123,9 @@ def isColimitCoconeFiberwiseColimitOfCocone {c : Cocone G} (hc : IsColimit c) : uniq s m hm := hc.hom_ext fun X => by have := hm X.base simp only [Functor.const_obj_obj, IsColimit.fac, NatTrans.comp_app, Functor.comp_obj, - Grothendieck.forget_obj, fiberwiseColimit_obj, natTransIntoForgetCompFiberwiseColimit_app, + Grothendieck.forget_obj, fiberwiseColim_obj, natTransIntoForgetCompFiberwiseColimit_app, whiskerLeft_app] - simp only [fiberwiseColimit_obj, coconeFiberwiseColimitOfCocone_pt, Functor.const_obj_obj, + simp only [fiberwiseColim_obj, coconeFiberwiseColimitOfCocone_pt, Functor.const_obj_obj, coconeFiberwiseColimitOfCocone_ι_app] at this simp [← this] @@ -131,9 +143,9 @@ def coconeOfCoconeFiberwiseColimit (c : Cocone (fiberwiseColimit G)) : Cocone G naturality := fun {X Y} ⟨f, g⟩ => by simp only [Functor.const_obj_obj, Functor.const_obj_map, Category.comp_id] rw [← Category.assoc, ← c.w f, ← Category.assoc] - simp only [fiberwiseColimit_obj, fiberwiseColimit_map, ι_colimMap_assoc, Functor.comp_obj, - Grothendieck.ι_obj, NatTrans.comp_app, whiskerRight_app, Functor.associator_hom_app, - Category.comp_id, colimit.ι_pre] + simp only [fiberwiseColimit_obj, fiberwiseColimit_map, ι_colimMap_assoc, + Functor.comp_obj, Grothendieck.ι_obj, NatTrans.comp_app, whiskerRight_app, + Functor.associator_hom_app, Category.comp_id, colimit.ι_pre] rw [← colimit.w _ g, ← Category.assoc, Functor.comp_map, ← G.map_comp] congr <;> simp } @@ -180,7 +192,8 @@ lemma ι_colimitFiberwiseColimitIso_hom (X : C) (d : F.obj X) : @[reassoc (attr := simp)] lemma ι_colimitFiberwiseColimitIso_inv (X : Grothendieck F) : colimit.ι G X ≫ (colimitFiberwiseColimitIso G).inv = - colimit.ι (Grothendieck.ι F X.base ⋙ G) X.fiber ≫ colimit.ι (fiberwiseColimit G) X.base := by + colimit.ι (Grothendieck.ι F X.base ⋙ G) X.fiber ≫ + colimit.ι (fiberwiseColimit G) X.base := by rw [Iso.comp_inv_eq] simp @@ -191,6 +204,27 @@ theorem hasColimitsOfShape_grothendieck [∀ X, HasColimitsOfShape (F.obj X) H] [HasColimitsOfShape C H] : HasColimitsOfShape (Grothendieck F) H where has_colimit _ := hasColimit_of_hasColimit_fiberwiseColimit_of_hasColimit _ +noncomputable section FiberwiseColim + +variable [∀ (c : C), HasColimitsOfShape (↑(F.obj c)) H] [HasColimitsOfShape C H] + +/-- The isomorphism `colimitFiberwiseColimitIso` induces an isomorphism of functors `(J ⥤ C) ⥤ C` +between `fiberwiseColim F H ⋙ colim` and `colim`. -/ +@[simps!] +def fiberwiseColimCompColimIso : fiberwiseColim F H ⋙ colim ≅ colim := + NatIso.ofComponents (fun G => colimitFiberwiseColimitIso G) + fun _ => by (iterate 2 apply colimit.hom_ext; intro); simp + +/-- Composing `fiberwiseColim F H` with the evaluation functor `(evaluation C H).obj c` is +naturally isomorphic to precomposing the Grothendieck inclusion `Grothendieck.ι` to `colim`. -/ +@[simps!] +def fiberwiseColimCompEvaluationIso (c : C) : + fiberwiseColim F H ⋙ (evaluation C H).obj c ≅ + (whiskeringLeft _ _ _).obj (Grothendieck.ι F c) ⋙ colim := + Iso.refl _ + +end FiberwiseColim + end Limits end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index f11792f5eaace..3db2bd31165a7 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -24,7 +24,7 @@ so that `m` factors through the `m'` in any other such factorisation. morphism `e`. * `HasImages C` means that every morphism in `C` has an image. * Let `f : X ⟶ Y` and `g : P ⟶ Q` be morphisms in `C`, which we will represent as objects of the - arrow category `arrow C`. Then `sq : f ⟶ g` is a commutative square in `C`. If `f` and `g` have + arrow category `Arrow C`. Then `sq : f ⟶ g` is a commutative square in `C`. If `f` and `g` have images, then `HasImageMap sq` represents the fact that there is a morphism `i : image f ⟶ image g` making the diagram @@ -241,7 +241,7 @@ def ofArrowIso {f g : Arrow C} (F : ImageFactorisation f.hom) (sq : f ⟶ g) [Is end ImageFactorisation -/-- `has_image f` means that there exists an image factorisation of `f`. -/ +/-- `HasImage f` means that there exists an image factorisation of `f`. -/ class HasImage (f : X ⟶ Y) : Prop where mk' :: exists_image : Nonempty (ImageFactorisation f) @@ -601,6 +601,8 @@ end section HasImageMap +-- Don't generate unnecessary injectivity lemmas which the `simpNF` linter will complain about. +set_option genInjectivity false in /-- An image map is a morphism `image f → image g` fitting into a commutative square and satisfying the obvious commutativity conditions. -/ structure ImageMap {f g : Arrow C} [HasImage f.hom] [HasImage g.hom] (sq : f ⟶ g) where @@ -609,9 +611,6 @@ structure ImageMap {f g : Arrow C} [HasImage f.hom] [HasImage g.hom] (sq : f ⟶ attribute [inherit_doc ImageMap] ImageMap.map ImageMap.map_ι --- Porting note: LHS of this simplifies, simpNF still complains after blacklisting -attribute [-simp, nolint simpNF] ImageMap.mk.injEq - instance inhabitedImageMap {f : Arrow C} [HasImage f.hom] : Inhabited (ImageMap (𝟙 f)) := ⟨⟨𝟙 _, by simp⟩⟩ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean index b61c4579338d9..f73849f24a405 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean @@ -52,16 +52,14 @@ variable {L : Type w} {R : Type w'} {fst snd : R → L} instance [Inhabited L] : Inhabited (WalkingMulticospan fst snd) := ⟨left default⟩ +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- Morphisms for `WalkingMulticospan`. -/ inductive Hom : ∀ _ _ : WalkingMulticospan fst snd, Type max w w' | id (A) : Hom A A | fst (b) : Hom (left (fst b)) (right b) | snd (b) : Hom (left (snd b)) (right b) -/- Porting note: simpNF says the LHS of this internal identifier simplifies -(which it does, using Hom.id_eq_id) -/ -attribute [-simp, nolint simpNF] WalkingMulticospan.Hom.id.sizeOf_spec - instance {a : WalkingMulticospan fst snd} : Inhabited (Hom a a) := ⟨Hom.id _⟩ @@ -99,16 +97,14 @@ variable {L : Type w} {R : Type w'} {fst snd : L → R} instance [Inhabited L] : Inhabited (WalkingMultispan fst snd) := ⟨left default⟩ +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- Morphisms for `WalkingMultispan`. -/ inductive Hom : ∀ _ _ : WalkingMultispan fst snd, Type max w w' | id (A) : Hom A A | fst (a) : Hom (left a) (right (fst a)) | snd (a) : Hom (left a) (right (snd a)) -/- Porting note: simpNF says the LHS of this internal identifier simplifies -(which it does, using Hom.id_eq_id) -/ -attribute [-simp, nolint simpNF] WalkingMultispan.Hom.id.sizeOf_spec - instance {a : WalkingMultispan fst snd} : Inhabited (Hom a a) := ⟨Hom.id _⟩ @@ -139,7 +135,6 @@ lemma Hom.comp_eq_comp {X Y Z : WalkingMultispan fst snd} end WalkingMultispan /-- This is a structure encapsulating the data necessary to define a `Multicospan`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): has_nonempty_instance linter not ported yet @[nolint checkUnivs] structure MulticospanIndex (C : Type u) [Category.{v} C] where (L : Type w) @@ -151,7 +146,6 @@ structure MulticospanIndex (C : Type u) [Category.{v} C] where snd : ∀ b, left (sndTo b) ⟶ right b /-- This is a structure encapsulating the data necessary to define a `Multispan`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): has_nonempty_instance linter not ported yet @[nolint checkUnivs] structure MultispanIndex (C : Type u) [Category.{v} C] where (L : Type w) @@ -276,14 +270,10 @@ end MultispanIndex variable {C : Type u} [Category.{v} C] /-- A multifork is a cone over a multicospan. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] abbrev Multifork (I : MulticospanIndex.{w, w'} C) := Cone I.multicospan /-- A multicofork is a cocone over a multispan. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] abbrev Multicofork (I : MultispanIndex.{w, w'} C) := Cocone I.multispan diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean new file mode 100644 index 0000000000000..fd2ca7692b49c --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/Basic.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Limits.Shapes.Terminal + +/-! +# Limits and colimits indexed by preorders + +In this file, we obtain the following very basic results +about limits and colimits indexed by a preordered type `J`: +* a least element in `J` implies the existence of all limits indexed by `J` +* a greatest element in `J` implies the existence of all colimits indexed by `J` + +-/ + +universe v v' u u' w + +open CategoryTheory Limits + +variable (J : Type w) [Preorder J] (C : Type u) [Category.{v} C] + +namespace Preorder + +section OrderBot + +variable [OrderBot J] + +/-- The least element in a preordered type is initial in the category +associated to this preorder. -/ +def isInitialBot : IsInitial (⊥ : J) := IsInitial.ofUnique _ + +instance : HasInitial J := hasInitial_of_unique ⊥ + +instance : HasLimitsOfShape J C := ⟨fun _ ↦ by infer_instance⟩ + +end OrderBot + +section OrderTop + +variable [OrderTop J] + +/-- The greatest element of a preordered type is terminal in the category +associated to this preorder. -/ +def isTerminalBot : IsTerminal (⊤ : J) := IsTerminal.ofUnique _ + +instance : HasTerminal J := hasTerminal_of_unique ⊤ + +instance : HasColimitsOfShape J C := ⟨fun _ ↦ by infer_instance⟩ + +end OrderTop + +end Preorder diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean new file mode 100644 index 0000000000000..b255543b53a02 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/HasIterationOfShape.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Limits.FunctorCategory.Basic +import Mathlib.CategoryTheory.Limits.Comma +import Mathlib.Order.SuccPred.Limit + +/-! +# An assumption for constructions by transfinite induction + +In this file, we introduce the typeclass `HasIterationOfShape J C` which is +an assumption in order to do constructions by transfinite induction indexed by +a well-ordered type `J` in a category `C` (see `CategoryTheory.SmallObject`). + +-/ + +universe w v v' u u' + +namespace CategoryTheory.Limits + +variable (J : Type w) [Preorder J] (C : Type u) [Category.{v} C] + (K : Type u') [Category.{v'} K] + +/-- A category `C` has iterations of shape a preordered type `J` +when certain specific shapes of colimits exists: colimits indexed by `J`, +and by `Set.Iio j` for `j : J`. -/ +class HasIterationOfShape : Prop where + hasColimitsOfShape_of_isSuccLimit (j : J) (hj : Order.IsSuccLimit j) : + HasColimitsOfShape (Set.Iio j) C := by infer_instance + hasColimitsOfShape : HasColimitsOfShape J C := by infer_instance + +attribute [instance] HasIterationOfShape.hasColimitsOfShape + +variable [HasIterationOfShape J C] + +variable {J} in +lemma hasColimitsOfShape_of_isSuccLimit (j : J) + (hj : Order.IsSuccLimit j) : + HasColimitsOfShape (Set.Iio j) C := + HasIterationOfShape.hasColimitsOfShape_of_isSuccLimit j hj + +instance : HasIterationOfShape J (Arrow C) where + hasColimitsOfShape_of_isSuccLimit j hj := by + have := hasColimitsOfShape_of_isSuccLimit C j hj + infer_instance + +instance : HasIterationOfShape J (K ⥤ C) where + hasColimitsOfShape_of_isSuccLimit j hj := by + have := hasColimitsOfShape_of_isSuccLimit C j hj + infer_instance + +end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Preorder/WellOrderContinuous.lean b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/WellOrderContinuous.lean new file mode 100644 index 0000000000000..d32739ec06f2a --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/Preorder/WellOrderContinuous.lean @@ -0,0 +1,82 @@ +/- +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.Data.Nat.SuccPred +import Mathlib.Order.SuccPred.Limit +import Mathlib.CategoryTheory.Limits.IsLimit +import Mathlib.CategoryTheory.Category.Preorder + +/-! +# Continuity of functors from well ordered types + +Let `F : J ⥤ C` be functor from a well ordered type `J`. +We introduce the typeclass `F.IsWellOrderContinuous` +to say that if `m` is a limit element, then `F.obj m` +is the colimit of the `F.obj j` for `j < m`. + +## TODO +* use the API for initial segments in order to generalize some +definitions in this file +* given a morphism `f` in `C`, introduce a structure +`TransfiniteCompositionOfShape J f` which contains the data +of a continuous functor `F : J ⥤ C` and an identification +of `f` to the map from `F.obj ⊥` to the colimit of `F` +* redefine `MorphismProperty.transfiniteCompositionsOfShape` +in terms of this structure `TransfiniteCompositionOfShape` + +-/ + +universe w v u + +namespace CategoryTheory.Functor + +open Category Limits + +variable {C : Type u} [Category.{v} C] {J : Type w} [Preorder J] + +/-- Given a functor `F : J ⥤ C` and `m : J`, this is the induced +functor `Set.Iio j ⥤ C`. -/ +@[simps!] +def restrictionLT (F : J ⥤ C) (j : J) : Set.Iio j ⥤ C := + Monotone.functor (f := fun k ↦ k.1) (fun _ _ ↦ id) ⋙ F + +/-- Given a functor `F : J ⥤ C` and `m : J`, this is the cocone with point `F.obj m` +for the restriction of `F` to `Set.Iio m`. -/ +@[simps] +def coconeLT (F : J ⥤ C) (m : J) : + Cocone (F.restrictionLT m) where + pt := F.obj m + ι := + { app := fun ⟨i, hi⟩ ↦ F.map (homOfLE hi.le) + naturality := fun ⟨i₁, hi₁⟩ ⟨i₂, hi₂⟩ f ↦ by + dsimp + rw [← F.map_comp, comp_id] + rfl } + +/-- A functor `F : J ⥤ C` is well-order-continuous if for any limit element `m : J`, +`F.obj m` identifies to the colimit of the `F.obj j` for `j < m`. -/ +class IsWellOrderContinuous (F : J ⥤ C) : Prop where + nonempty_isColimit (m : J) (hm : Order.IsSuccLimit m) : + Nonempty (IsColimit (F.coconeLT m)) + +/-- If `F : J ⥤ C` is well-order-continuous and `m : J` is a limit element, then +the cocone `F.coconeLT m` is colimit, i.e. `F.obj m` identifies to the colimit +of the `F.obj j` for `j < m`. -/ +noncomputable def isColimitOfIsWellOrderContinuous (F : J ⥤ C) [F.IsWellOrderContinuous] + (m : J) (hm : Order.IsSuccLimit m) : + IsColimit (F.coconeLT m) := (IsWellOrderContinuous.nonempty_isColimit m hm).some + +instance (F : ℕ ⥤ C) : F.IsWellOrderContinuous where + nonempty_isColimit m hm := by simp at hm + +lemma isWellOrderContinuous_of_iso {F G : J ⥤ C} (e : F ≅ G) [F.IsWellOrderContinuous] : + G.IsWellOrderContinuous where + nonempty_isColimit (m : J) (hm : Order.IsSuccLimit m) := + ⟨(IsColimit.precomposeHomEquiv (isoWhiskerLeft _ e) _).1 + (IsColimit.ofIsoColimit (F.isColimitOfIsWellOrderContinuous m hm) + (Cocones.ext (e.app _)))⟩ + +end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean index 7a97bde56f6a0..4283ca314dade 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/HasPullback.lean @@ -502,10 +502,8 @@ end PushoutSymmetry variable (C) -/-- `HasPullbacks` represents a choice of pullback for every pair of morphisms - -See --/ +/-- `HasPullbacks` represents a choice of pullback for every pair of morphisms. -/ +@[stacks 001W] abbrev HasPullbacks := HasLimitsOfShape WalkingCospan C diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean index d1c5d29b44ebf..9287d4c3696b8 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Reflexive.lean @@ -189,6 +189,8 @@ open WalkingReflexivePair namespace WalkingReflexivePair +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- The type of morphisms for the diagram indexing reflexive (co)equalizers -/ inductive Hom : (WalkingReflexivePair → WalkingReflexivePair → Type) | left : Hom one zero @@ -199,10 +201,6 @@ inductive Hom : (WalkingReflexivePair → WalkingReflexivePair → Type) | id (X : WalkingReflexivePair) : Hom X X deriving DecidableEq -attribute [-simp, nolint simpNF] Hom.id.sizeOf_spec -attribute [-simp, nolint simpNF] Hom.leftCompReflexion.sizeOf_spec -attribute [-simp, nolint simpNF] Hom.rightCompReflexion.sizeOf_spec - /-- Composition of morphisms in the diagram indexing reflexive (co)equalizers -/ def Hom.comp : ∀ { X Y Z : WalkingReflexivePair } (_ : Hom X Y) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index 5706484d58e29..5d87af2751e5f 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -585,7 +585,6 @@ instance : HasPushouts.{u} (Type u) := variable {X Y Z : Type u} {X' Y' Z' : Type v} variable (f : X ⟶ Z) (g : Y ⟶ Z) (f' : X' ⟶ Z') (g' : Y' ⟶ Z') --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The usual explicit pullback in the category of types, as a subtype of the product. The full `LimitCone` data is bundled as `pullbackLimitCone f g`. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean index e8a53e2075d20..349e66c0e30f5 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean @@ -68,6 +68,8 @@ instance : DecidableEq (WalkingParallelFamily J) instance : Inhabited (WalkingParallelFamily J) := ⟨zero⟩ +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- The type family of morphisms for the diagram indexing a wide (co)equalizer. -/ inductive WalkingParallelFamily.Hom (J : Type w) : WalkingParallelFamily J → WalkingParallelFamily J → Type w @@ -101,6 +103,22 @@ theorem WalkingParallelFamily.hom_id (X : WalkingParallelFamily J) : WalkingParallelFamily.Hom.id X = 𝟙 X := rfl +variable (J) in +/-- `Arrow (WalkingParallelFamily J)` identifies to the type obtained +by adding two elements to `T`. -/ +def WalkingParallelFamily.arrowEquiv : + Arrow (WalkingParallelFamily J) ≃ Option (Option J) where + toFun f := match f.left, f.right, f.hom with + | zero, _, .id _ => none + | one, _, .id _ => some none + | zero, one, .line t => some (some t) + invFun x := match x with + | none => Arrow.mk (𝟙 zero) + | some none => Arrow.mk (𝟙 one) + | some (some t) => Arrow.mk (.line t) + left_inv := by rintro ⟨(_ | _), _, (_ | _)⟩ <;> rfl + right_inv := by rintro (_ | (_ | _)) <;> rfl + variable {C : Type u} [Category.{v} C] variable {X Y : C} (f : J → (X ⟶ Y)) @@ -680,5 +698,3 @@ instance (priority := 10) hasCoequalizers_of_hasWideCoequalizers [HasWideCoequal hasColimitsOfShape_of_equivalence.{w} walkingParallelFamilyEquivWalkingParallelPair end CategoryTheory.Limits - -attribute [nolint simpNF] CategoryTheory.Limits.WalkingParallelFamily.Hom.id.sizeOf_spec diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean index 06d17c50c7d91..13aac3489474c 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean @@ -49,6 +49,8 @@ namespace WidePullbackShape variable {J} +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- The type of arrows for the shape indexing a wide pullback. -/ inductive Hom : WidePullbackShape J → WidePullbackShape J → Type w | id : ∀ X, Hom X X @@ -97,11 +99,6 @@ instance category : SmallCategory (WidePullbackShape J) := theorem hom_id (X : WidePullbackShape J) : Hom.id X = 𝟙 X := rfl -/- Porting note: we get a warning that we should change LHS to `sizeOf (𝟙 X)` but Lean cannot -find the category instance on `WidePullbackShape J` in that case. Once supplied in the proof, -the proposed proof of `simp [only WidePullbackShape.hom_id]` does not work -/ -attribute [nolint simpNF] Hom.id.sizeOf_spec - variable {C : Type u} [Category.{v} C] /-- Construct a functor out of the wide pullback shape given a J-indexed collection of arrows to a @@ -153,6 +150,8 @@ namespace WidePushoutShape variable {J} +-- Don't generate unnecessary `sizeOf_spec` lemma which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- The type of arrows for the shape indexing a wide pushout. -/ inductive Hom : WidePushoutShape J → WidePushoutShape J → Type w | id : ∀ X, Hom X X @@ -198,10 +197,6 @@ instance category : SmallCategory (WidePushoutShape J) := theorem hom_id (X : WidePushoutShape J) : Hom.id X = 𝟙 X := rfl -/- Porting note: we get a warning that we should change LHS to `sizeOf (𝟙 X)` but Lean cannot -find the category instance on `WidePushoutShape J` in that case. Once supplied in the proof, -the proposed proof of `simp [only WidePushoutShape.hom_id]` does not work -/ -attribute [nolint simpNF] Hom.id.sizeOf_spec variable {C : Type u} [Category.{v} C] /-- Construct a functor out of the wide pushout shape given a J-indexed collection of arrows from a @@ -359,13 +354,10 @@ noncomputable abbrev ι (j : J) : objs j ⟶ widePushout _ _ arrows := noncomputable abbrev head : B ⟶ widePushout B objs arrows := colimit.ι (WidePushoutShape.wideSpan _ _ _) Option.none -@[reassoc (attr := simp)] +@[reassoc, simp] theorem arrow_ι (j : J) : arrows j ≫ ι arrows j = head arrows := by apply colimit.w (WidePushoutShape.wideSpan _ _ _) (WidePushoutShape.Hom.init j) --- Porting note: this can simplify itself -attribute [nolint simpNF] WidePushout.arrow_ι WidePushout.arrow_ι_assoc - variable {arrows} /-- Descend a collection of morphisms to a morphism from the pushout. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean index 8b225b7e7ee7d..5239662237b56 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean @@ -420,7 +420,6 @@ variable [HasZeroMorphisms C] /-- A zero morphism `0 : X ⟶ Y` is an isomorphism if and only if the identities on both `X` and `Y` are zero. -/ -@[simps] def isIsoZeroEquiv (X Y : C) : IsIso (0 : X ⟶ Y) ≃ 𝟙 X = 0 ∧ 𝟙 Y = 0 where toFun := by intro i @@ -431,9 +430,6 @@ def isIsoZeroEquiv (X Y : C) : IsIso (0 : X ⟶ Y) ≃ 𝟙 X = 0 ∧ 𝟙 Y = 0 left_inv := by aesop_cat right_inv := by aesop_cat --- Porting note: simp solves these -attribute [-simp, nolint simpNF] isIsoZeroEquiv_apply isIsoZeroEquiv_symm_apply - /-- A zero morphism `0 : X ⟶ X` is an isomorphism if and only if the identity on `X` is zero. -/ diff --git a/Mathlib/CategoryTheory/Limits/Sifted.lean b/Mathlib/CategoryTheory/Limits/Sifted.lean index 8f5721d4f8dae..7baf9a4e2a34b 100644 --- a/Mathlib/CategoryTheory/Limits/Sifted.lean +++ b/Mathlib/CategoryTheory/Limits/Sifted.lean @@ -20,7 +20,7 @@ preserves finite products. - [*Algebraic Theories*, Chapter 2.][Adamek_Rosicky_Vitale_2010] -/ -universe w v v₁ u u₁ +universe w v v₁ v₂ u u₁ u₂ namespace CategoryTheory @@ -101,4 +101,14 @@ end IsSifted end +section + +variable {C : Type u} [Category.{v} C] [IsSiftedOrEmpty C] {D : Type u₁} [Category.{v₁} D] + {D' : Type u₂} [Category.{v₂} D'] (F : C ⥤ D) (G : C ⥤ D') + +instance [F.Final] [G.Final] : (F.prod' G).Final := + show (diag C ⋙ F.prod G).Final from final_comp _ _ + +end + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index fd38583178553..462571b6714a7 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -192,10 +192,8 @@ instance hasLimitsOfShape [Small.{u} J] : HasLimitsOfShape J (Type u) where /-- The category of types has all limits. -More specifically, when `UnivLE.{v, u}`, the category `Type u` has all `v`-small limits. - -See . --/ +More specifically, when `UnivLE.{v, u}`, the category `Type u` has all `v`-small limits. -/ +@[stacks 002U] instance (priority := 1300) hasLimitsOfSize [UnivLE.{v, u}] : HasLimitsOfSize.{w, v} (Type u) where has_limits_of_shape _ := { } @@ -321,7 +319,6 @@ See `CategoryTheory.Limits.Types.Quot`. def Quot.Rel (F : J ⥤ Type u) : (Σ j, F.obj j) → (Σ j, F.obj j) → Prop := fun p p' => ∃ f : p.1 ⟶ p'.1, p'.2 = F.map f p.2 --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- A quotient type implementing the colimit of a functor `F : J ⥤ Type u`, as pairs `⟨j, x⟩` where `x : F.obj j`, modulo the equivalence relation generated by `⟨j, x⟩ ~ ⟨j', x'⟩` whenever there is a morphism `f : j ⟶ j'` so `F.map f x = x'`. @@ -485,10 +482,8 @@ instance hasColimit [Small.{u} J] (F : J ⥤ Type u) : HasColimit F := instance hasColimitsOfShape [Small.{u} J] : HasColimitsOfShape J (Type u) where -/-- The category of types has all colimits. - -See . --/ +/-- The category of types has all colimits. -/ +@[stacks 002U] instance (priority := 1300) hasColimitsOfSize [UnivLE.{v, u}] : HasColimitsOfSize.{w, v} (Type u) where diff --git a/Mathlib/CategoryTheory/Localization/Construction.lean b/Mathlib/CategoryTheory/Localization/Construction.lean index 5bcd82a33e3e7..9cc4362686b03 100644 --- a/Mathlib/CategoryTheory/Localization/Construction.lean +++ b/Mathlib/CategoryTheory/Localization/Construction.lean @@ -50,7 +50,6 @@ namespace Localization namespace Construction --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- If `W : MorphismProperty C`, `LocQuiver W` is a quiver with the same objects as `C`, and whose morphisms are those in `C` and placeholders for formal inverses of the morphisms in `W`. -/ @@ -90,7 +89,6 @@ namespace MorphismProperty open Localization.Construction --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The localized category obtained by formally inverting the morphisms in `W : MorphismProperty C` -/ def Localization := diff --git a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean index d3be845da56a6..a5fc9f3df1d9d 100644 --- a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean +++ b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean @@ -38,7 +38,7 @@ include hW lemma inverts : (W.functorCategory (Discrete J)).IsInvertedBy (lim ⋙ L) := - fun _ _ f hf => Localization.inverts L W _ (hW.lim_map f hf) + fun _ _ f hf => Localization.inverts L W _ (hW.limMap f hf) variable [W.ContainsIdentities] [Finite J] diff --git a/Mathlib/CategoryTheory/Localization/LocallySmall.lean b/Mathlib/CategoryTheory/Localization/LocallySmall.lean new file mode 100644 index 0000000000000..12c04e06860ef --- /dev/null +++ b/Mathlib/CategoryTheory/Localization/LocallySmall.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Localization.HasLocalization +import Mathlib.CategoryTheory.EssentiallySmall + +/-! +# Locally small localizations + +In this file, given `W : MorphismProperty C` and a universe `w`, we show +that there exists a term in `HasLocalization.{w} W` if and only if +there exists (or for all) localization functors `L : C ⥤ D` for `W`, +the category `D` is locally `w`-small. + +-/ + +universe w v₁ v₂ u₁ u₂ + +namespace CategoryTheory.MorphismProperty + +variable {C : Type u₁} [Category.{v₁} C] (W : MorphismProperty C) + +/-- If `L : C ⥤ D` is a localization functor for a class of morphisms +`W : MorphismProperty C`, and `D` is locally `w`-small, we may obtain +a `HasLocalization.{w} W` instance by shrinking the morphisms in `D`. +(This version assumes that the types of objects of the categories +`C` and `D` are in the same universe.) -/ +noncomputable def hasLocalizationOfLocallySmall + {D : Type u₁} [Category.{v₂} D] [LocallySmall.{w} D] + (L : C ⥤ D) [L.IsLocalization W] : + HasLocalization.{w} W where + D := ShrinkHoms D + L := L ⋙ (ShrinkHoms.equivalence D).functor + +/-- If `L : C ⥤ D` is a localization functor for a class of morphisms +`W : MorphismProperty C`, and `D` is locally `w`-small, we may obtain +a `HasLocalization.{w} W` instance. This should be used only in the +unlikely situation where the types of objects of `C` and `D` are in +different universes. Otherwise, use `hasLocalizationOfLocallySmall`. -/ +noncomputable irreducible_def hasLocalizationOfLocallySmall' + {D : Type u₂} [Category.{v₂} D] [LocallySmall.{w} D] + (L : C ⥤ D) [L.IsLocalization W] : + HasLocalization.{w} W := by + have : LocallySmall.{w} (InducedCategory _ L.obj) := + ⟨fun X Y ↦ inferInstanceAs (Small.{w} (L.obj X ⟶ L.obj Y))⟩ + let L' : C ⥤ (InducedCategory _ L.obj) := + { obj X := X + map f := L.map f } + have := Localization.essSurj L W + have : (inducedFunctor L.obj).EssSurj := ⟨fun Y ↦ ⟨_, ⟨L.objObjPreimageIso Y⟩⟩⟩ + have : (inducedFunctor L.obj).IsEquivalence := { } + let e := (inducedFunctor L.obj).asEquivalence + let e' : (L' ⋙ e.functor) ⋙ e.inverse ≅ L' := + Functor.associator _ _ _ ≪≫ isoWhiskerLeft L' e.unitIso.symm ≪≫ L'.rightUnitor + have : L'.IsLocalization W := + Functor.IsLocalization.of_iso W (L₁ := L ⋙ e.inverse) e' + exact hasLocalizationOfLocallySmall.{w} W L' + +/-- If a class of morphisms `W : MorphismProperty C` satisfies `HasLocalization.{w} W`, +then any localized category for `W` (i.e. any target of a localization functor +`L : C ⥤ D` for `W`) is locally `w`-small. -/ +lemma locallySmall_of_hasLocalization {D : Type u₂} [Category.{v₂} D] + (L : C ⥤ D) [L.IsLocalization W] [HasLocalization.{w} W] : + LocallySmall.{w} D where + hom_small _ _ := small_of_injective (fun _ _ h ↦ + (Localization.uniq L W.Q' W).functor.map_injective h) + +end CategoryTheory.MorphismProperty diff --git a/Mathlib/CategoryTheory/Localization/Trifunctor.lean b/Mathlib/CategoryTheory/Localization/Trifunctor.lean new file mode 100644 index 0000000000000..782b08c753808 --- /dev/null +++ b/Mathlib/CategoryTheory/Localization/Trifunctor.lean @@ -0,0 +1,225 @@ +/- +Copyright (c) 2025 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Localization.Bifunctor +import Mathlib.CategoryTheory.Functor.CurryingThree +import Mathlib.CategoryTheory.Products.Associator + +/-! +# Lifting of trifunctors + +In this file, in the context of the localization of categories, we extend the notion +of lifting of functors to the case of trifunctors +(see also the file `Localization.Bifunctor` for the case of bifunctors). +The main result in this file is that we can localize "associator" isomorphisms +(see the definition `Localization.associator`). + +-/ + +namespace CategoryTheory + +variable {C₁ C₂ C₃ C₁₂ C₂₃ D₁ D₂ D₃ D₁₂ D₂₃ C D E : Type*} + [Category C₁] [Category C₂] [Category C₃] [Category D₁] [Category D₂] [Category D₃] + [Category C₁₂] [Category C₂₃] [Category D₁₂] [Category D₂₃] + [Category C] [Category D] [Category E] + +namespace MorphismProperty + +/-- Classes of morphisms `W₁ : MorphismProperty C₁`, `W₂ : MorphismProperty C₂` and +`W₃ : MorphismProperty C₃` are said to be inverted by `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E` if +`W₁.prod (W₂.prod W₃)` is inverted by the +functor `currying₃.functor.obj F : C₁ × C₂ × C₃ ⥤ E`. -/ +def IsInvertedBy₃ (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) + (W₃ : MorphismProperty C₃) (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) : Prop := + (W₁.prod (W₂.prod W₃)).IsInvertedBy (currying₃.functor.obj F) + +end MorphismProperty + +namespace Localization + +section + +variable (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) (L₃ : C₃ ⥤ D₃) + +/-- Given functors `L₁ : C₁ ⥤ D₁`, `L₂ : C₂ ⥤ D₂`, `L₃ : C₃ ⥤ D₃`, +morphisms properties `W₁` on `C₁`, `W₂` on `C₂`, `W₃` on `C₃`, and +functors `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E` and `F' : D₁ ⥤ D₂ ⥤ D₃ ⥤ E`, we say +`Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F F'` holds if `F` is induced by `F'`, up to an isomorphism. -/ +class Lifting₃ (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) + (W₃ : MorphismProperty C₃) + (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) (F' : D₁ ⥤ D₂ ⥤ D₃ ⥤ E) where + /-- the isomorphism `((((whiskeringLeft₃ E).obj L₁).obj L₂).obj L₃).obj F' ≅ F` expressing + that `F` is induced by `F'` up to an isomorphism -/ + iso' : ((((whiskeringLeft₃ E).obj L₁).obj L₂).obj L₃).obj F' ≅ F + +variable (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) (W₃ : MorphismProperty C₃) + (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) (F' : D₁ ⥤ D₂ ⥤ D₃ ⥤ E) [Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F F'] + +/-- The isomorphism `((((whiskeringLeft₃ E).obj L₁).obj L₂).obj L₃).obj F' ≅ F` +when `Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F F'` holds. -/ +noncomputable def Lifting₃.iso : + ((((whiskeringLeft₃ E).obj L₁).obj L₂).obj L₃).obj F' ≅ F := + Lifting₃.iso' W₁ W₂ W₃ + +variable (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) (F' : D₁ ⥤ D₂ ⥤ D₃ ⥤ E) + +noncomputable instance Lifting₃.uncurry [Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F F'] : + Lifting (L₁.prod (L₂.prod L₃)) (W₁.prod (W₂.prod W₃)) + (uncurry₃.obj F) (uncurry₃.obj F') where + iso' := uncurry₃.mapIso (Lifting₃.iso L₁ L₂ L₃ W₁ W₂ W₃ F F') + +end + +section + +variable (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) {W₁ : MorphismProperty C₁} {W₂ : MorphismProperty C₂} + {W₃ : MorphismProperty C₃} + (hF : MorphismProperty.IsInvertedBy₃ W₁ W₂ W₃ F) + (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) (L₃ : C₃ ⥤ D₃) + [L₁.IsLocalization W₁] [L₂.IsLocalization W₂] [L₃.IsLocalization W₃] + [W₁.ContainsIdentities] [W₂.ContainsIdentities] [W₃.ContainsIdentities] + +/-- Given localization functor `L₁ : C₁ ⥤ D₁`, `L₂ : C₂ ⥤ D₂` and `L₃ : C₃ ⥤ D₃` +with respect to `W₁ : MorphismProperty C₁`, `W₂ : MorphismProperty C₂` and +`W₃ : MorphismProperty C₃` respectively, and a trifunctor `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ E` +which inverts `W₁`, `W₂` and `W₃`, this is the induced localized +trifunctor `D₁ ⥤ D₂ ⥤ D₃ ⥤ E`. -/ +noncomputable def lift₃ : D₁ ⥤ D₂ ⥤ D₃ ⥤ E := + curry₃.obj (lift (uncurry₃.obj F) hF (L₁.prod (L₂.prod L₃))) + +noncomputable instance : Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F (lift₃ F hF L₁ L₂ L₃) where + iso' := + (curry₃ObjProdComp L₁ L₂ L₃ _).symm ≪≫ + curry₃.mapIso (fac (uncurry₃.obj F) hF (L₁.prod (L₂.prod L₃))) ≪≫ + currying₃.unitIso.symm.app F + +end + +section + +variable (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) (L₃ : C₃ ⥤ D₃) + (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) (W₃ : MorphismProperty C₃) + [L₁.IsLocalization W₁] [L₂.IsLocalization W₂] [L₃.IsLocalization W₃] + [W₁.ContainsIdentities] [W₂.ContainsIdentities] [W₃.ContainsIdentities] + (F₁ F₂ : C₁ ⥤ C₂ ⥤ C₃ ⥤ E) (F₁' F₂' : D₁ ⥤ D₂ ⥤ D₃ ⥤ E) + [Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F₁ F₁'] [Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F₂ F₂'] (τ : F₁ ⟶ F₂) + (e : F₁ ≅ F₂) + +/-- The natural transformation `F₁' ⟶ F₂'` of trifunctors induced by a +natural transformation `τ : F₁ ⟶ F₂` when `Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F₁ F₁'` +and `Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F₂ F₂'` hold. -/ +noncomputable def lift₃NatTrans : F₁' ⟶ F₂' := + fullyFaithfulUncurry₃.preimage + (liftNatTrans (L₁.prod (L₂.prod L₃)) (W₁.prod (W₂.prod W₃)) (uncurry₃.obj F₁) + (uncurry₃.obj F₂) (uncurry₃.obj F₁') (uncurry₃.obj F₂') (uncurry₃.map τ)) + +@[simp] +theorem lift₃NatTrans_app_app_app (X₁ : C₁) (X₂ : C₂) (X₃ : C₃) : + (((lift₃NatTrans L₁ L₂ L₃ W₁ W₂ W₃ F₁ F₂ F₁' F₂' τ).app + (L₁.obj X₁)).app (L₂.obj X₂)).app (L₃.obj X₃) = + (((Lifting₃.iso L₁ L₂ L₃ W₁ W₂ W₃ F₁ F₁').hom.app X₁).app X₂).app X₃ ≫ + ((τ.app X₁).app X₂).app X₃ ≫ + (((Lifting₃.iso L₁ L₂ L₃ W₁ W₂ W₃ F₂ F₂').inv.app X₁).app X₂).app X₃ := by + dsimp [lift₃NatTrans, fullyFaithfulUncurry₃, Equivalence.fullyFaithfulFunctor] + simp only [currying₃_unitIso_hom_app_app_app_app, Functor.id_obj, + currying₃_unitIso_inv_app_app_app_app, Functor.comp_obj, + Category.comp_id, Category.id_comp] + exact liftNatTrans_app _ _ _ _ (uncurry₃.obj F₁') (uncurry₃.obj F₂') (uncurry₃.map τ) ⟨X₁, X₂, X₃⟩ + +variable {F₁' F₂'} in +include W₁ W₂ W₃ in +theorem natTrans₃_ext {τ τ' : F₁' ⟶ F₂'} + (h : ∀ (X₁ : C₁) (X₂ : C₂) (X₃ : C₃), ((τ.app (L₁.obj X₁)).app (L₂.obj X₂)).app (L₃.obj X₃) = + ((τ'.app (L₁.obj X₁)).app (L₂.obj X₂)).app (L₃.obj X₃)) : τ = τ' := + uncurry₃.map_injective (natTrans_ext (L₁.prod (L₂.prod L₃)) (W₁.prod (W₂.prod W₃)) + (fun _ ↦ h _ _ _)) + +/-- The natural isomorphism `F₁' ≅ F₂'` of trifunctors induced by a +natural isomorphism `e : F₁ ≅ F₂` when `Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F₁ F₁'` +and `Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ F₂ F₂'` hold. -/ +@[simps] +noncomputable def lift₃NatIso : F₁' ≅ F₂' where + hom := lift₃NatTrans L₁ L₂ L₃ W₁ W₂ W₃ F₁ F₂ F₁' F₂' e.hom + inv := lift₃NatTrans L₁ L₂ L₃ W₁ W₂ W₃ F₂ F₁ F₂' F₁' e.inv + hom_inv_id := natTrans₃_ext L₁ L₂ L₃ W₁ W₂ W₃ (by aesop_cat) + inv_hom_id := natTrans₃_ext L₁ L₂ L₃ W₁ W₂ W₃ (by aesop_cat) + +end + +section + +variable + (L₁ : C₁ ⥤ D₁) (L₂ : C₂ ⥤ D₂) (L₃ : C₃ ⥤ D₃) (L₁₂ : C₁₂ ⥤ D₁₂) (L₂₃ : C₂₃ ⥤ D₂₃) (L : C ⥤ D) + (W₁ : MorphismProperty C₁) (W₂ : MorphismProperty C₂) (W₃ : MorphismProperty C₃) + (W₁₂ : MorphismProperty C₁₂) (W₂₃ : MorphismProperty C₂₃) (W : MorphismProperty C) + [W₁.ContainsIdentities] [W₂.ContainsIdentities] [W₃.ContainsIdentities] + [L₁.IsLocalization W₁] [L₂.IsLocalization W₂] [L₃.IsLocalization W₃] [L.IsLocalization W] + (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ C) + (F : C₁ ⥤ C₂₃ ⥤ C) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃) + (iso : bifunctorComp₁₂ F₁₂ G ≅ bifunctorComp₂₃ F G₂₃) + (F₁₂' : D₁ ⥤ D₂ ⥤ D₁₂) (G' : D₁₂ ⥤ D₃ ⥤ D) + (F' : D₁ ⥤ D₂₃ ⥤ D) (G₂₃' : D₂ ⥤ D₃ ⥤ D₂₃) + [Lifting₂ L₁ L₂ W₁ W₂ (F₁₂ ⋙ (whiskeringRight _ _ _).obj L₁₂) F₁₂'] + [Lifting₂ L₁₂ L₃ W₁₂ W₃ (G ⋙ (whiskeringRight _ _ _).obj L) G'] + [Lifting₂ L₁ L₂₃ W₁ W₂₃ (F ⋙ (whiskeringRight _ _ _).obj L) F'] + [Lifting₂ L₂ L₃ W₂ W₃ (G₂₃ ⋙ (whiskeringRight _ _ _).obj L₂₃) G₂₃'] + +/-- The construction `bifunctorComp₁₂` of a trifunctor by composition of bifunctors +is compatible with localization. -/ +noncomputable def Lifting₃.bifunctorComp₁₂ : + Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ + ((Functor.postcompose₃.obj L).obj (bifunctorComp₁₂ F₁₂ G)) + (bifunctorComp₁₂ F₁₂' G') where + iso' := + ((whiskeringRight C₁ _ _).obj + ((whiskeringRight C₂ _ _).obj ((whiskeringLeft _ _ D).obj L₃))).mapIso + ((bifunctorComp₁₂Functor.mapIso + (Lifting₂.iso L₁ L₂ W₁ W₂ (F₁₂ ⋙ (whiskeringRight _ _ _).obj L₁₂) F₁₂')).app G') ≪≫ + (bifunctorComp₁₂Functor.obj F₁₂).mapIso + (Lifting₂.iso L₁₂ L₃ W₁₂ W₃ (G ⋙ (whiskeringRight _ _ _).obj L) G') + +/-- The construction `bifunctorComp₂₃` of a trifunctor by composition of bifunctors +is compatible with localization. -/ +noncomputable def Lifting₃.bifunctorComp₂₃ : + Lifting₃ L₁ L₂ L₃ W₁ W₂ W₃ + ((Functor.postcompose₃.obj L).obj (bifunctorComp₂₃ F G₂₃)) + (bifunctorComp₂₃ F' G₂₃') where + iso' := + ((whiskeringLeft _ _ _).obj L₁).mapIso ((bifunctorComp₂₃Functor.obj F').mapIso + (Lifting₂.iso L₂ L₃ W₂ W₃ (G₂₃ ⋙ (whiskeringRight _ _ _).obj L₂₃) G₂₃')) ≪≫ + (bifunctorComp₂₃Functor.mapIso + (Lifting₂.iso L₁ L₂₃ W₁ W₂₃ (F ⋙ (whiskeringRight _ _ _).obj L) F')).app G₂₃ + +variable {F₁₂ G F G₂₃} + +/-- The associator isomorphism obtained by localization. -/ +noncomputable def associator : bifunctorComp₁₂ F₁₂' G' ≅ bifunctorComp₂₃ F' G₂₃' := + letI := Lifting₃.bifunctorComp₁₂ L₁ L₂ L₃ L₁₂ L W₁ W₂ W₃ W₁₂ F₁₂ G F₁₂' G' + letI := Lifting₃.bifunctorComp₂₃ L₁ L₂ L₃ L₂₃ L W₁ W₂ W₃ W₂₃ F G₂₃ F' G₂₃' + lift₃NatIso L₁ L₂ L₃ W₁ W₂ W₃ _ _ _ _ ((Functor.postcompose₃.obj L).mapIso iso) + +lemma associator_hom_app_app_app (X₁ : C₁) (X₂ : C₂) (X₃ : C₃) : + (((associator L₁ L₂ L₃ L₁₂ L₂₃ L W₁ W₂ W₃ W₁₂ W₂₃ iso F₁₂' G' F' G₂₃').hom.app (L₁.obj X₁)).app + (L₂.obj X₂)).app (L₃.obj X₃) = + (G'.map (((Lifting₂.iso L₁ L₂ W₁ W₂ + (F₁₂ ⋙ (whiskeringRight C₂ C₁₂ D₁₂).obj L₁₂) F₁₂').hom.app X₁).app X₂)).app (L₃.obj X₃) ≫ + ((Lifting₂.iso L₁₂ L₃ W₁₂ W₃ (G ⋙ (whiskeringRight C₃ C D).obj L) G').hom.app + ((F₁₂.obj X₁).obj X₂)).app X₃ ≫ + L.map (((iso.hom.app X₁).app X₂).app X₃) ≫ + ((Lifting₂.iso L₁ L₂₃ W₁ W₂₃ + (F ⋙ (whiskeringRight _ _ _).obj L) F').inv.app X₁).app ((G₂₃.obj X₂).obj X₃) ≫ + (F'.obj (L₁.obj X₁)).map + (((Lifting₂.iso L₂ L₃ W₂ W₃ + (G₂₃ ⋙ (whiskeringRight _ _ _).obj L₂₃) G₂₃').inv.app X₂).app X₃) := by + dsimp [associator] + rw [lift₃NatTrans_app_app_app] + dsimp [Lifting₃.iso, Lifting₃.bifunctorComp₁₂, Lifting₃.bifunctorComp₂₃] + simp only [Category.assoc] + +end + +end Localization + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monad/Algebra.lean b/Mathlib/CategoryTheory/Monad/Algebra.lean index b9f87cb3658a5..9e6a0c7fb3331 100644 --- a/Mathlib/CategoryTheory/Monad/Algebra.lean +++ b/Mathlib/CategoryTheory/Monad/Algebra.lean @@ -262,8 +262,6 @@ end Monad namespace Comonad /-- An Eilenberg-Moore coalgebra for a comonad `T`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] structure Coalgebra (G : Comonad C) : Type max u₁ v₁ where /-- The underlying object associated to a coalgebra. -/ A : C @@ -288,8 +286,6 @@ namespace Coalgebra variable {G : Comonad C} /-- A morphism of Eilenberg-Moore coalgebras for the comonad `G`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet ---@[ext, nolint has_nonempty_instance] @[ext] structure Hom (A B : Coalgebra G) where /-- The underlying morphism associated to a morphism of coalgebras. -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index 4f1f7e9e08cb8..77c4c2e4eb1e1 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -342,10 +342,8 @@ theorem braiding_inv_tensorUnit_right (X : C) : (β_ X (𝟙_ C)).inv = (λ_ X). end /-- -A symmetric monoidal category is a braided monoidal category for which the braiding is symmetric. - -See . --/ +A symmetric monoidal category is a braided monoidal category for which the braiding is symmetric. -/ +@[stacks 0FFW] class SymmetricCategory (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] extends BraidedCategory.{v} C where -- braiding symmetric: diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index f91baa59cfa43..2451bdd366d10 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -149,10 +149,8 @@ In a monoidal category, we can take the tensor product of objects, `X ⊗ Y` and Tensor product does not need to be strictly associative on objects, but there is a specified associator, `α_ X Y Z : (X ⊗ Y) ⊗ Z ≅ X ⊗ (Y ⊗ Z)`. There is a tensor unit `𝟙_ C`, with specified left and right unitor isomorphisms `λ_ X : 𝟙_ C ⊗ X ≅ X` and `ρ_ X : X ⊗ 𝟙_ C ≅ X`. -These associators and unitors satisfy the pentagon and triangle equations. - -See . --/ +These associators and unitors satisfy the pentagon and triangle equations. -/ +@[stacks 0FFK] -- Porting note: The Mathport did not translate the temporary notation class MonoidalCategory (C : Type u) [𝒞 : Category.{v} C] extends MonoidalCategoryStruct C where tensorHom_def {X₁ Y₁ X₂ Y₂ : C} (f : X₁ ⟶ Y₁) (g : X₂ ⟶ Y₂) : diff --git a/Mathlib/CategoryTheory/Monoidal/Center.lean b/Mathlib/CategoryTheory/Monoidal/Center.lean index d670e07b11bf3..d0aed0f699e35 100644 --- a/Mathlib/CategoryTheory/Monoidal/Center.lean +++ b/Mathlib/CategoryTheory/Monoidal/Center.lean @@ -47,7 +47,6 @@ Thinking of `C` as a 2-category with a single `0`-morphism, these are the same a transformations (in the pseudo- sense) of the identity 2-functor on `C`, which send the unique `0`-morphism to `X`. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. structure HalfBraiding (X : C) where β : ∀ U, X ⊗ U ≅ U ⊗ X monoidal : ∀ U U', (β (U ⊗ U')).hom = @@ -66,7 +65,6 @@ variable (C) /-- The Drinfeld center of a monoidal category `C` has as objects pairs `⟨X, b⟩`, where `X : C` and `b` is a half-braiding on `X`. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. def Center := Σ X : C, HalfBraiding X @@ -75,8 +73,7 @@ namespace Center variable {C} /-- A morphism in the Drinfeld center of `C`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet -@[ext] -- @[nolint has_nonempty_instance] +@[ext] structure Hom (X Y : Center C) where f : X.1 ⟶ Y.1 comm : ∀ U, (f ▷ U) ≫ (Y.2.β U).hom = (X.2.β U).hom ≫ (U ◁ f) := by aesop_cat diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 04bf2a16a1b33..94bea212340b0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -35,6 +35,10 @@ section variable (C) +-- Don't generate unnecessary `sizeOf_spec` or `injEq` lemmas +-- which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in +set_option genInjectivity false in /-- Given a type `C`, the free monoidal category over `C` has as objects formal expressions built from (formal) tensor products of terms of `C` and a formal unit. Its morphisms are compositions and @@ -52,13 +56,9 @@ local notation "F" => FreeMonoidalCategory namespace FreeMonoidalCategory -attribute [nolint simpNF] unit.sizeOf_spec tensor.injEq tensor.sizeOf_spec - /-- Formal compositions and tensor products of identities, unitors and associators. The morphisms of the free monoidal category are obtained as a quotient of these formal morphisms by the relations defining a monoidal category. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet --- @[nolint has_nonempty_instance] inductive Hom : F C → F C → Type u | id (X) : Hom X X | α_hom (X Y Z : F C) : Hom ((X.tensor Y).tensor Z) (X.tensor (Y.tensor Z)) diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index 910625a02dc20..3b648ca7dd9f5 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -50,7 +50,6 @@ variable (C) /-- We say an object in the free monoidal category is in normal form if it is of the form `(((𝟙_ C) ⊗ X₁) ⊗ X₂) ⊗ ⋯`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] inductive NormalMonoidalObject : Type u | unit : NormalMonoidalObject | tensor : NormalMonoidalObject → C → NormalMonoidalObject diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index 75e6de20c0015..1f421e496a054 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -55,7 +55,7 @@ namespace Functor -- "weigh more" than structural maps. -- (However by this argument `associativity` is currently stated backwards!) /-- A functor `F : C ⥤ D` between monoidal categories is lax monoidal if it is -equipped with morphisms `ε : 𝟙 _D ⟶ F.obj (𝟙_ C)` and `μ X Y : F.obj X ⊗ F.obj Y ⟶ F.obj (X ⊗ Y)`, +equipped with morphisms `ε : 𝟙_ D ⟶ F.obj (𝟙_ C)` and `μ X Y : F.obj X ⊗ F.obj Y ⟶ F.obj (X ⊗ Y)`, satisfying the appropriate coherences. -/ class LaxMonoidal where /-- unit morphism -/ diff --git a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean index ad627f3e283cf..045a1d884c4b3 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean @@ -311,8 +311,6 @@ a fixed choice of limit data for the empty functor, and for `pair X Y` for every This is an implementation detail for `SymmetricOfChosenFiniteProducts`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter `has_nonempty_instance` not ported yet --- @[nolint has_nonempty_instance] @[nolint unusedArguments] def MonoidalOfChosenFiniteProductsSynonym (_𝒯 : LimitCone (Functor.empty.{0} C)) (_ℬ : ∀ X Y : C, LimitCone (pair X Y)) := diff --git a/Mathlib/CategoryTheory/Monoidal/Opposite.lean b/Mathlib/CategoryTheory/Monoidal/Opposite.lean index a9d80869adb38..ed2e80caacecd 100644 --- a/Mathlib/CategoryTheory/Monoidal/Opposite.lean +++ b/Mathlib/CategoryTheory/Monoidal/Opposite.lean @@ -22,7 +22,6 @@ open CategoryTheory.MonoidalCategory /-- The type of objects of the opposite (or "reverse") monoidal category. Use the notation `Cᴹᵒᵖ`. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. structure MonoidalOpposite (C : Type u₁) where /-- The object of `MonoidalOpposite C` that represents `x : C`. -/ mop :: /-- The object of `C` represented by `x : MonoidalOpposite C`. -/ unmop : C diff --git a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean index 0d34626b7468f..6d548e45cfabe 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean @@ -61,6 +61,26 @@ lemma ext (W W' : MorphismProperty C) (h : ∀ ⦃X Y : C⦄ (f : X ⟶ Y), W f lemma top_apply {X Y : C} (f : X ⟶ Y) : (⊤ : MorphismProperty C) f := by simp only [top_eq] +@[simp] +lemma sSup_iff (S : Set (MorphismProperty C)) {X Y : C} (f : X ⟶ Y) : + sSup S f ↔ ∃ (W : S), W.1 f := by + dsimp [sSup, iSup] + constructor + · rintro ⟨_, ⟨⟨_, ⟨⟨_, ⟨_, h⟩, rfl⟩, rfl⟩⟩, rfl⟩, hf⟩ + exact ⟨⟨_, h⟩, hf⟩ + · rintro ⟨⟨W, hW⟩, hf⟩ + exact ⟨_, ⟨⟨_, ⟨_, ⟨⟨W, hW⟩, rfl⟩⟩, rfl⟩, rfl⟩, hf⟩ + +@[simp] +lemma iSup_iff {ι : Type*} (W : ι → MorphismProperty C) {X Y : C} (f : X ⟶ Y) : + iSup W f ↔ ∃ i, W i f := by + apply (sSup_iff (Set.range W) f).trans + constructor + · rintro ⟨⟨_, i, rfl⟩, hf⟩ + exact ⟨i, hf⟩ + · rintro ⟨i, hf⟩ + exact ⟨⟨_, i, rfl⟩, hf⟩ + /-- The morphism property in `Cᵒᵖ` associated to a morphism property in `C` -/ @[simp] def op (P : MorphismProperty C) : MorphismProperty Cᵒᵖ := fun _ _ f => P f.unop diff --git a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean index 16ed860f8bcb4..68929bd30839c 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean @@ -88,7 +88,6 @@ lemma pi {J : Type w} {C : J → Type u} {D : J → Type u'} end IsInvertedBy --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The full subcategory of `C ⥤ D` consisting of functors inverting morphisms in `W` -/ def FunctorsInverting (W : MorphismProperty C) (D : Type*) [Category D] := FullSubcategory fun F : C ⥤ D => W.IsInvertedBy F diff --git a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean index d9280218e658f..48d927c184b04 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Limits.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Limits.lean @@ -23,35 +23,103 @@ We also introduce properties `IsStableUnderProductsOfShape`, `IsStableUnderLimit -/ -universe v u +universe w v u namespace CategoryTheory -open Limits +open Category Limits namespace MorphismProperty variable {C : Type u} [Category.{v} C] +section + +variable (P : MorphismProperty C) + +/-- Given a class of morphisms `P`, this is the class of pullbacks +of morphisms in `P`. -/ +def pullbacks : MorphismProperty C := fun A B q ↦ + ∃ (X Y : C) (p : X ⟶ Y) (f : A ⟶ X) (g : B ⟶ Y) (_ : P p), + IsPullback f q p g + +lemma pullbacks_mk {A B X Y : C} {f : A ⟶ X} {q : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} + (sq : IsPullback f q p g) (hp : P p) : + P.pullbacks q := + ⟨_, _, _, _, _, hp, sq⟩ + +lemma le_pullbacks : P ≤ P.pullbacks := by + intro A B q hq + exact P.pullbacks_mk IsPullback.of_id_fst hq + +/-- Given a class of morphisms `P`, this is the class of pushouts +of morphisms in `P`. -/ +def pushouts : MorphismProperty C := fun X Y q ↦ + ∃ (A B : C) (p : A ⟶ B) (f : A ⟶ X) (g : B ⟶ Y) (_ : P p), + IsPushout f p q g + +lemma pushouts_mk {A B X Y : C} {f : A ⟶ X} {q : A ⟶ B} {p : X ⟶ Y} {g : B ⟶ Y} + (sq : IsPushout f q p g) (hq : P q) : + P.pushouts p := + ⟨_, _, _, _, _, hq, sq⟩ + +lemma le_pushouts : P ≤ P.pushouts := by + intro X Y p hp + exact P.pushouts_mk IsPushout.of_id_fst hp + +instance : P.pushouts.RespectsIso := + RespectsIso.of_respects_arrow_iso _ (by + rintro q q' e ⟨A, B, p, f, g, hp, h⟩ + exact ⟨A, B, p, f ≫ e.hom.left, g ≫ e.hom.right, hp, + IsPushout.paste_horiz h (IsPushout.of_horiz_isIso ⟨e.hom.w⟩)⟩) + +instance : P.pullbacks.RespectsIso := + RespectsIso.of_respects_arrow_iso _ (by + rintro q q' e ⟨X, Y, p, f, g, hp, h⟩ + exact ⟨X, Y, p, e.inv.left ≫ f, e.inv.right ≫ g, hp, + IsPullback.paste_horiz (IsPullback.of_horiz_isIso ⟨e.inv.w⟩) h⟩) + /-- A morphism property is `IsStableUnderBaseChange` if the base change of such a morphism still falls in the class. -/ -class IsStableUnderBaseChange (P : MorphismProperty C) : Prop where +class IsStableUnderBaseChange : Prop where of_isPullback {X Y Y' S : C} {f : X ⟶ S} {g : Y ⟶ S} {f' : Y' ⟶ Y} {g' : Y' ⟶ X} (sq : IsPullback f' g' g f) (hg : P g) : P g' +instance : P.pullbacks.IsStableUnderBaseChange where + of_isPullback := by + rintro _ _ _ _ _ _ _ _ h ⟨_, _, _, _, _, hp, hq⟩ + exact P.pullbacks_mk (h.paste_horiz hq) hp + /-- A morphism property is `IsStableUnderCobaseChange` if the cobase change of such a morphism still falls in the class. -/ -class IsStableUnderCobaseChange (P : MorphismProperty C) : Prop where +class IsStableUnderCobaseChange : Prop where of_isPushout {A A' B B' : C} {f : A ⟶ A'} {g : A ⟶ B} {f' : B ⟶ B'} {g' : A' ⟶ B'} (sq : IsPushout g f f' g') (hf : P f) : P f' -lemma of_isPullback {P : MorphismProperty C} [P.IsStableUnderBaseChange] +instance : P.pushouts.IsStableUnderCobaseChange where + of_isPushout := by + rintro _ _ _ _ _ _ _ _ h ⟨_, _, _, _, _, hp, hq⟩ + exact P.pushouts_mk (hq.paste_horiz h) hp + +variable {P} in +lemma of_isPullback [P.IsStableUnderBaseChange] {X Y Y' S : C} {f : X ⟶ S} {g : Y ⟶ S} {f' : Y' ⟶ Y} {g' : Y' ⟶ X} (sq : IsPullback f' g' g f) (hg : P g) : P g' := IsStableUnderBaseChange.of_isPullback sq hg +lemma isStableUnderBaseChange_iff_pullbacks_le : + P.IsStableUnderBaseChange ↔ P.pullbacks ≤ P := by + constructor + · intro h _ _ _ ⟨_, _, _, _, _, h₁, h₂⟩ + exact of_isPullback h₂ h₁ + · intro h + constructor + intro _ _ _ _ _ _ _ _ h₁ h₂ + exact h _ ⟨_, _, _, _, _, h₂, h₁⟩ + +variable {P} in /-- Alternative constructor for `IsStableUnderBaseChange`. -/ -theorem IsStableUnderBaseChange.mk' {P : MorphismProperty C} [RespectsIso P] +theorem IsStableUnderBaseChange.mk' [RespectsIso P] (hP₂ : ∀ (X Y S : C) (f : X ⟶ S) (g : Y ⟶ S) [HasPullback f g] (_ : P g), P (pullback.fst f g)) : IsStableUnderBaseChange P where @@ -61,6 +129,8 @@ theorem IsStableUnderBaseChange.mk' {P : MorphismProperty C} [RespectsIso P] rw [← P.cancel_left_of_respectsIso e.inv, sq.flip.isoPullback_inv_fst] exact hP₂ _ _ _ f g hg +variable (C) + instance IsStableUnderBaseChange.isomorphisms : (isomorphisms C).IsStableUnderBaseChange where of_isPullback {_ _ _ _ f g _ _} h hg := @@ -68,7 +138,6 @@ instance IsStableUnderBaseChange.isomorphisms : have := hasPullback_of_left_iso g f h.isoPullback_hom_snd ▸ inferInstanceAs (IsIso _) -variable (C) in instance IsStableUnderBaseChange.monomorphisms : (monomorphisms C).IsStableUnderBaseChange where of_isPullback {X Y Y' S f g f' g'} h hg := by @@ -81,34 +150,36 @@ instance IsStableUnderBaseChange.monomorphisms : simp only [Category.assoc, h.w, reassoc_of% h₁₂] · exact h₁₂ -instance (priority := 900) IsStableUnderBaseChange.respectsIso {P : MorphismProperty C} +variable {C P} + +instance (priority := 900) IsStableUnderBaseChange.respectsIso [IsStableUnderBaseChange P] : RespectsIso P := by apply RespectsIso.of_respects_arrow_iso intro f g e exact of_isPullback (IsPullback.of_horiz_isIso (CommSq.mk e.inv.w)) -theorem pullback_fst {P : MorphismProperty C} [IsStableUnderBaseChange P] +theorem pullback_fst [IsStableUnderBaseChange P] {X Y S : C} (f : X ⟶ S) (g : Y ⟶ S) [HasPullback f g] (H : P g) : P (pullback.fst f g) := of_isPullback (IsPullback.of_hasPullback f g).flip H @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.fst := pullback_fst -theorem pullback_snd {P : MorphismProperty C} [IsStableUnderBaseChange P] +theorem pullback_snd [IsStableUnderBaseChange P] {X Y S : C} (f : X ⟶ S) (g : Y ⟶ S) [HasPullback f g] (H : P f) : P (pullback.snd f g) := of_isPullback (IsPullback.of_hasPullback f g) H @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.snd := pullback_snd -theorem baseChange_obj [HasPullbacks C] {P : MorphismProperty C} +theorem baseChange_obj [HasPullbacks C] [IsStableUnderBaseChange P] {S S' : C} (f : S' ⟶ S) (X : Over S) (H : P X.hom) : P ((Over.pullback f).obj X).hom := pullback_snd X.hom f H @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.baseChange_obj := baseChange_obj -theorem baseChange_map [HasPullbacks C] {P : MorphismProperty C} +theorem baseChange_map [HasPullbacks C] [IsStableUnderBaseChange P] {S S' : C} (f : S' ⟶ S) {X Y : Over S} (g : X ⟶ Y) (H : P g.left) : P ((Over.pullback f).map g).left := by let e := @@ -121,7 +192,7 @@ theorem baseChange_map [HasPullbacks C] {P : MorphismProperty C} @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.baseChange_map := baseChange_map -theorem pullback_map [HasPullbacks C] {P : MorphismProperty C} +theorem pullback_map [HasPullbacks C] [IsStableUnderBaseChange P] [P.IsStableUnderComposition] {S X X' Y Y' : C} {f : X ⟶ S} {g : Y ⟶ S} {f' : X' ⟶ S} {g' : Y' ⟶ S} {i₁ : X ⟶ X'} {i₂ : Y ⟶ Y'} (h₁ : P i₁) (h₂ : P i₂) (e₁ : f = i₁ ≫ f') (e₂ : g = i₂ ≫ g') : @@ -142,13 +213,23 @@ theorem pullback_map [HasPullbacks C] {P : MorphismProperty C} @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.pullback_map := pullback_map -lemma of_isPushout {P : MorphismProperty C} [P.IsStableUnderCobaseChange] +lemma of_isPushout [P.IsStableUnderCobaseChange] {A A' B B' : C} {f : A ⟶ A'} {g : A ⟶ B} {f' : B ⟶ B'} {g' : A' ⟶ B'} (sq : IsPushout g f f' g') (hf : P f) : P f' := IsStableUnderCobaseChange.of_isPushout sq hf +lemma isStableUnderCobaseChange_iff_pushouts_le : + P.IsStableUnderCobaseChange ↔ P.pushouts ≤ P := by + constructor + · intro h _ _ _ ⟨_, _, _, _, _, h₁, h₂⟩ + exact of_isPushout h₂ h₁ + · intro h + constructor + intro _ _ _ _ _ _ _ _ h₁ h₂ + exact h _ ⟨_, _, _, _, _, h₂, h₁⟩ + /-- An alternative constructor for `IsStableUnderCobaseChange`. -/ -theorem IsStableUnderCobaseChange.mk' {P : MorphismProperty C} [RespectsIso P] +theorem IsStableUnderCobaseChange.mk' [RespectsIso P] (hP₂ : ∀ (A B A' : C) (f : A ⟶ A') (g : A ⟶ B) [HasPushout f g] (_ : P f), P (pushout.inr f g)) : IsStableUnderCobaseChange P where @@ -177,25 +258,25 @@ instance IsStableUnderCobaseChange.epimorphisms : dsimp simp only [← reassoc_of% h.w, h₁₂] -instance IsStableUnderCobaseChange.respectsIso {P : MorphismProperty C} +instance IsStableUnderCobaseChange.respectsIso [IsStableUnderCobaseChange P] : RespectsIso P := RespectsIso.of_respects_arrow_iso _ fun _ _ e ↦ of_isPushout (IsPushout.of_horiz_isIso (CommSq.mk e.hom.w)) -theorem pushout_inl {P : MorphismProperty C} [IsStableUnderCobaseChange P] +theorem pushout_inl [IsStableUnderCobaseChange P] {A B A' : C} (f : A ⟶ A') (g : A ⟶ B) [HasPushout f g] (H : P g) : P (pushout.inl f g) := of_isPushout (IsPushout.of_hasPushout f g) H @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.inl := pushout_inl -theorem pushout_inr {P : MorphismProperty C} [IsStableUnderCobaseChange P] +theorem pushout_inr [IsStableUnderCobaseChange P] {A B A' : C} (f : A ⟶ A') (g : A ⟶ B) [HasPushout f g] (H : P f) : P (pushout.inr f g) := of_isPushout (IsPushout.of_hasPushout f g).flip H @[deprecated (since := "2024-11-06")] alias IsStableUnderBaseChange.inr := pushout_inr -instance IsStableUnderCobaseChange.op {P : MorphismProperty C} [IsStableUnderCobaseChange P] : +instance IsStableUnderCobaseChange.op [IsStableUnderCobaseChange P] : IsStableUnderBaseChange P.op where of_isPullback sq hg := P.of_isPushout sq.unop hg @@ -203,7 +284,7 @@ instance IsStableUnderCobaseChange.unop {P : MorphismProperty Cᵒᵖ} [IsStable IsStableUnderBaseChange P.unop where of_isPullback sq hg := P.of_isPushout sq.op hg -instance IsStableUnderBaseChange.op {P : MorphismProperty C} [IsStableUnderBaseChange P] : +instance IsStableUnderBaseChange.op [IsStableUnderBaseChange P] : IsStableUnderCobaseChange P.op where of_isPushout sq hf := P.of_isPullback sq.unop hf @@ -221,39 +302,156 @@ instance IsStableUnderCobaseChange.inf {P Q : MorphismProperty C} [IsStableUnder IsStableUnderCobaseChange (P ⊓ Q) where of_isPushout hp hg := ⟨of_isPushout hp hg.left, of_isPushout hp hg.right⟩ -section +end -variable (W : MorphismProperty C) +section LimitsOfShape + +variable (W : MorphismProperty C) (J : Type*) [Category J] + +/-- The class of morphisms in `C` that are limits of shape `J` of +natural transformations involving morphisms in `W`. -/ +inductive limitsOfShape : MorphismProperty C + | mk (X₁ X₂ : J ⥤ C) (c₁ : Cone X₁) (c₂ : Cone X₂) + (_ : IsLimit c₁) (h₂ : IsLimit c₂) (f : X₁ ⟶ X₂) (_ : W.functorCategory J f) : + limitsOfShape (h₂.lift (Cone.mk _ (c₁.π ≫ f))) + +instance : (W.limitsOfShape J).RespectsIso := + RespectsIso.of_respects_arrow_iso _ (by + rintro ⟨_, _, f⟩ ⟨Y₁, Y₂, g⟩ e ⟨X₁, X₂, c₁, c₂, h₁, h₂, f, hf⟩ + let e₁ := Arrow.leftFunc.mapIso e + let e₂ := Arrow.rightFunc.mapIso e + have fac : g ≫ e₂.inv = e₁.inv ≫ h₂.lift (Cone.mk _ (c₁.π ≫ f)) := + e.inv.w.symm + let c₁' : Cone X₁ := { pt := Y₁, π := (Functor.const _).map e₁.inv ≫ c₁.π } + let c₂' : Cone X₂ := { pt := Y₂, π := (Functor.const _).map e₂.inv ≫ c₂.π } + have h₁' : IsLimit c₁' := IsLimit.ofIsoLimit h₁ (Cones.ext e₁) + have h₂' : IsLimit c₂' := IsLimit.ofIsoLimit h₂ (Cones.ext e₂) + obtain hg : h₂'.lift (Cone.mk _ (c₁'.π ≫ f)) = g := + h₂'.hom_ext (fun j ↦ by + rw [h₂'.fac] + simp [reassoc_of% fac, c₁', c₂']) + rw [← hg] + exact ⟨_, _, _, _, h₁', _, _, hf⟩) + +variable {J} in +lemma limitsOfShape_limMap {X Y : J ⥤ C} + (f : X ⟶ Y) [HasLimit X] [HasLimit Y] (hf : W.functorCategory _ f) : + W.limitsOfShape J (limMap f) := + ⟨_, _, _, _, limit.isLimit X, _, _, hf⟩ /-- The property that a morphism property `W` is stable under limits indexed by a category `J`. -/ -def IsStableUnderLimitsOfShape (J : Type*) [Category J] : Prop := +def IsStableUnderLimitsOfShape : Prop := ∀ (X₁ X₂ : J ⥤ C) (c₁ : Cone X₁) (c₂ : Cone X₂) (_ : IsLimit c₁) (h₂ : IsLimit c₂) (f : X₁ ⟶ X₂) (_ : W.functorCategory J f), W (h₂.lift (Cone.mk _ (c₁.π ≫ f))) +lemma isStableUnderLimitsOfShape_iff_limitsOfShape_le : + W.IsStableUnderLimitsOfShape J ↔ W.limitsOfShape J ≤ W := by + constructor + · rintro h _ _ _ ⟨_, _, _, _, h₁, h₂, f, hf⟩ + exact h _ _ _ _ h₁ h₂ f hf + · intro h _ _ _ _ h₁ h₂ f hf + exact h _ ⟨_, _, _, _, h₁, _, _, hf⟩ + +variable {W J} + +lemma IsStableUnderLimitsOfShape.limitsOfShape_le + (hW : W.IsStableUnderLimitsOfShape J) : W.limitsOfShape J ≤ W := + (W.isStableUnderLimitsOfShape_iff_limitsOfShape_le J).1 hW + +lemma IsStableUnderLimitsOfShape.limMap + (hW : W.IsStableUnderLimitsOfShape J) {X Y : J ⥤ C} + (f : X ⟶ Y) [HasLimit X] [HasLimit Y] (hf : W.functorCategory _ f) : + W (limMap f) := + hW.limitsOfShape_le _ (limitsOfShape_limMap _ _ hf) + +end LimitsOfShape + +section ColimitsOfShape + +variable (W : MorphismProperty C) (J : Type*) [Category J] + +/-- The class of morphisms in `C` that are colimits of shape `J` of +natural transformations involving morphisms in `W`. -/ +inductive colimitsOfShape : MorphismProperty C + | mk (X₁ X₂ : J ⥤ C) (c₁ : Cocone X₁) (c₂ : Cocone X₂) + (h₁ : IsColimit c₁) (h₂ : IsColimit c₂) (f : X₁ ⟶ X₂) (_ : W.functorCategory J f) : + colimitsOfShape (h₁.desc (Cocone.mk _ (f ≫ c₂.ι))) + +instance : (W.colimitsOfShape J).RespectsIso := + RespectsIso.of_respects_arrow_iso _ (by + rintro ⟨_, _, f⟩ ⟨Y₁, Y₂, g⟩ e ⟨X₁, X₂, c₁, c₂, h₁, h₂, f, hf⟩ + let e₁ := Arrow.leftFunc.mapIso e + let e₂ := Arrow.rightFunc.mapIso e + have fac : e₁.hom ≫ g = h₁.desc (Cocone.mk _ (f ≫ c₂.ι)) ≫ e₂.hom := e.hom.w + let c₁' : Cocone X₁ := { pt := Y₁, ι := c₁.ι ≫ (Functor.const _).map e₁.hom} + let c₂' : Cocone X₂ := { pt := Y₂, ι := c₂.ι ≫ (Functor.const _).map e₂.hom} + have h₁' : IsColimit c₁' := IsColimit.ofIsoColimit h₁ (Cocones.ext e₁) + have h₂' : IsColimit c₂' := IsColimit.ofIsoColimit h₂ (Cocones.ext e₂) + obtain hg : h₁'.desc (Cocone.mk _ (f ≫ c₂'.ι)) = g := + h₁'.hom_ext (fun j ↦ by + rw [h₁'.fac] + simp [fac, c₁', c₂']) + rw [← hg] + exact ⟨_, _, _, _, _, h₂', _, hf⟩) + +variable {J} in +lemma colimitsOfShape_colimMap {X Y : J ⥤ C} + (f : X ⟶ Y) [HasColimit X] [HasColimit Y] (hf : W.functorCategory _ f) : + W.colimitsOfShape J (colimMap f) := + ⟨_, _, _, _, _, colimit.isColimit Y, _, hf⟩ + /-- The property that a morphism property `W` is stable under colimits indexed by a category `J`. -/ -def IsStableUnderColimitsOfShape (J : Type*) [Category J] : Prop := +def IsStableUnderColimitsOfShape : Prop := ∀ (X₁ X₂ : J ⥤ C) (c₁ : Cocone X₁) (c₂ : Cocone X₂) (h₁ : IsColimit c₁) (_ : IsColimit c₂) (f : X₁ ⟶ X₂) (_ : W.functorCategory J f), W (h₁.desc (Cocone.mk _ (f ≫ c₂.ι))) -variable {W} +lemma isStableUnderColimitsOfShape_iff_colimitsOfShape_le : + W.IsStableUnderColimitsOfShape J ↔ W.colimitsOfShape J ≤ W := by + constructor + · rintro h _ _ _ ⟨_, _, _, _, h₁, h₂, f, hf⟩ + exact h _ _ _ _ h₁ h₂ f hf + · intro h _ _ _ _ h₁ h₂ f hf + exact h _ ⟨_, _, _, _, _, h₂, _, hf⟩ -lemma IsStableUnderLimitsOfShape.lim_map {J : Type*} [Category J] - (hW : W.IsStableUnderLimitsOfShape J) {X Y : J ⥤ C} - (f : X ⟶ Y) [HasLimitsOfShape J C] (hf : W.functorCategory _ f) : - W (lim.map f) := - hW X Y _ _ (limit.isLimit X) (limit.isLimit Y) f hf +variable {W J} -lemma IsStableUnderColimitsOfShape.colim_map {J : Type*} [Category J] +lemma IsStableUnderColimitsOfShape.colimitsOfShape_le + (hW : W.IsStableUnderColimitsOfShape J) : W.colimitsOfShape J ≤ W := + (W.isStableUnderColimitsOfShape_iff_colimitsOfShape_le J).1 hW + +lemma IsStableUnderColimitsOfShape.colimMap (hW : W.IsStableUnderColimitsOfShape J) {X Y : J ⥤ C} - (f : X ⟶ Y) [HasColimitsOfShape J C] (hf : W.functorCategory _ f) : - W (colim.map f) := - hW X Y _ _ (colimit.isColimit X) (colimit.isColimit Y) f hf + (f : X ⟶ Y) [HasColimit X] [HasColimit Y] (hf : W.functorCategory _ f) : + W (colimMap f) := + hW.colimitsOfShape_le _ (colimitsOfShape_colimMap _ _ hf) + +end ColimitsOfShape + +section Coproducts + +variable (W : MorphismProperty C) + +/-- Given `W : MorphismProperty C`, this is class of morphisms that are +isomorphic to a coproduct of a family (indexed by some `J : Type w`) of maps in `W`. -/ +def coproducts : MorphismProperty C := ⨆ (J : Type w), W.colimitsOfShape (Discrete J) + +lemma colimitsOfShape_le_coproducts (J : Type w) : + W.colimitsOfShape (Discrete J) ≤ coproducts.{w} W := + le_iSup (f := fun (J : Type w) ↦ W.colimitsOfShape (Discrete J)) J + +lemma coproducts_iff {X Y : C} (f : X ⟶ Y) : + coproducts.{w} W f ↔ ∃ (J : Type w), W.colimitsOfShape (Discrete J) f := by + simp only [coproducts, iSup_iff] -variable (W) +end Coproducts + +section Products + +variable (W : MorphismProperty C) /-- The property that a morphism property `W` is stable under products indexed by a type `J`. -/ abbrev IsStableUnderProductsOfShape (J : Type*) := W.IsStableUnderLimitsOfShape (Discrete J) @@ -321,7 +519,7 @@ lemma isStableUnderCoproductsOfShape_of_isStableUnderFiniteCoproducts W.IsStableUnderCoproductsOfShape J := IsStableUnderFiniteCoproducts.isStableUnderCoproductsOfShape J -end +end Products section Diagonal diff --git a/Mathlib/CategoryTheory/MorphismProperty/TransfiniteComposition.lean b/Mathlib/CategoryTheory/MorphismProperty/TransfiniteComposition.lean index a795a1546a370..d6082ef5b309b 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/TransfiniteComposition.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/TransfiniteComposition.lean @@ -3,11 +3,8 @@ 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.Data.Nat.SuccPred -import Mathlib.Order.SuccPred.Limit -import Mathlib.CategoryTheory.Category.Preorder -import Mathlib.CategoryTheory.Limits.IsLimit import Mathlib.CategoryTheory.MorphismProperty.Composition +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.WellOrderContinuous /-! # Classes of morphisms that are stable under transfinite composition @@ -47,59 +44,10 @@ open Category Limits variable {C : Type u} [Category.{v} C] -namespace Functor - -variable {J : Type w} [Preorder J] - -/-- Given a functor `F : J ⥤ C` and `m : J`, this is the induced -functor `Set.Iio j ⥤ C`. -/ -@[simps!] -def restrictionLT (F : J ⥤ C) (j : J) : Set.Iio j ⥤ C := - Monotone.functor (f := fun k ↦ k.1) (fun _ _ ↦ id) ⋙ F - -/-- Given a functor `F : J ⥤ C` and `m : J`, this is the cocone with point `F.obj m` -for the restriction of `F` to `Set.Iio m`. -/ -@[simps] -def coconeLT (F : J ⥤ C) (m : J) : - Cocone (F.restrictionLT m) where - pt := F.obj m - ι := - { app := fun ⟨i, hi⟩ ↦ F.map (homOfLE hi.le) - naturality := fun ⟨i₁, hi₁⟩ ⟨i₂, hi₂⟩ f ↦ by - dsimp - rw [← F.map_comp, comp_id] - rfl } - -/-- A functor `F : J ⥤ C` is well-order-continuous if for any limit element `m : J`, -`F.obj m` identifies to the colimit of the `F.obj j` for `j < m`. -/ -class IsWellOrderContinuous (F : J ⥤ C) : Prop where - nonempty_isColimit (m : J) (hm : Order.IsSuccLimit m) : - Nonempty (IsColimit (F.coconeLT m)) - -/-- If `F : J ⥤ C` is well-order-continuous and `m : J` is a limit element, then -the cocone `F.coconeLT m` is colimit, i.e. `F.obj m` identifies to the colimit -of the `F.obj j` for `j < m`. -/ -noncomputable def isColimitOfIsWellOrderContinuous (F : J ⥤ C) [F.IsWellOrderContinuous] - (m : J) (hm : Order.IsSuccLimit m) : - IsColimit (F.coconeLT m) := (IsWellOrderContinuous.nonempty_isColimit m hm).some - -instance (F : ℕ ⥤ C) : F.IsWellOrderContinuous where - nonempty_isColimit m hm := by simp at hm - -lemma isWellOrderContinuous_of_iso {F G : J ⥤ C} (e : F ≅ G) [F.IsWellOrderContinuous] : - G.IsWellOrderContinuous where - nonempty_isColimit (m : J) (hm : Order.IsSuccLimit m) := - ⟨(IsColimit.precomposeHomEquiv (isoWhiskerLeft _ e) _).1 - (IsColimit.ofIsoColimit (F.isColimitOfIsWellOrderContinuous m hm) - (Cocones.ext (e.app _)))⟩ - -end Functor - namespace MorphismProperty variable (W : MorphismProperty C) --- for Basic.lean section variable (J : Type w) [LinearOrder J] [SuccOrder J] [OrderBot J] diff --git a/Mathlib/CategoryTheory/Noetherian.lean b/Mathlib/CategoryTheory/Noetherian.lean index 6b7a78e563ca6..8c80a261f546d 100644 --- a/Mathlib/CategoryTheory/Noetherian.lean +++ b/Mathlib/CategoryTheory/Noetherian.lean @@ -30,10 +30,8 @@ open CategoryTheory.Limits variable {C : Type*} [Category C] /-- A noetherian object is an object -which does not have infinite increasing sequences of subobjects. - -See https://stacks.math.columbia.edu/tag/0FCG --/ +which does not have infinite increasing sequences of subobjects. -/ +@[stacks 0FCG] class NoetherianObject (X : C) : Prop where subobject_gt_wellFounded' : WellFounded ((· > ·) : Subobject X → Subobject X → Prop) @@ -42,10 +40,8 @@ lemma NoetherianObject.subobject_gt_wellFounded (X : C) [NoetherianObject X] : NoetherianObject.subobject_gt_wellFounded' /-- An artinian object is an object -which does not have infinite decreasing sequences of subobjects. - -See https://stacks.math.columbia.edu/tag/0FCF --/ +which does not have infinite decreasing sequences of subobjects. -/ +@[stacks 0FCF] class ArtinianObject (X : C) : Prop where subobject_lt_wellFounded' : WellFounded ((· < ·) : Subobject X → Subobject X → Prop) diff --git a/Mathlib/CategoryTheory/Opposites.lean b/Mathlib/CategoryTheory/Opposites.lean index 914d995a8d4cd..8a6a9972931b0 100644 --- a/Mathlib/CategoryTheory/Opposites.lean +++ b/Mathlib/CategoryTheory/Opposites.lean @@ -60,10 +60,8 @@ namespace CategoryTheory variable [Category.{v₁} C] -/-- The opposite category. - -See . --/ +/-- The opposite category. -/ +@[stacks 001M] instance Category.opposite : Category.{v₁} Cᵒᵖ where comp f g := (g.unop ≫ f.unop).op id X := (𝟙 (unop X)).op diff --git a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean index 0c7bc7c4e3cc3..28da24fb31033 100644 --- a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean @@ -193,7 +193,6 @@ section variable (C D : Type*) [Category C] [Category D] [Preadditive C] [Preadditive D] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled additive functors. -/ def AdditiveFunctor := FullSubcategory fun F : C ⥤ D => F.Additive diff --git a/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean b/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean index 1e95960d69dc2..0416818a8fec6 100644 --- a/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Preadditive/InjectiveResolution.lean @@ -38,8 +38,6 @@ variable {C : Type u} [Category.{v} C] [HasZeroObject C] [HasZeroMorphisms C] An `InjectiveResolution Z` consists of a bundled `ℕ`-indexed cochain complex of injective objects, along with a quasi-isomorphism from the complex consisting of just `Z` supported in degree `0`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure InjectiveResolution (Z : C) where /-- the cochain complex involved in the resolution -/ cocomplex : CochainComplex C ℕ diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 17cee933b9716..f323817080a5e 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -51,8 +51,6 @@ Ideally this would conveniently interact with both `Mat_` and `Matrix`. open CategoryTheory CategoryTheory.Preadditive -open scoped Classical - noncomputable section namespace CategoryTheory @@ -74,13 +72,13 @@ namespace Mat_ variable {C} --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- A morphism in `Mat_ C` is a dependently typed matrix of morphisms. -/ def Hom (M N : Mat_ C) : Type v₁ := DMatrix M.ι N.ι fun i j => M.X i ⟶ N.X j namespace Hom +open scoped Classical in /-- The identity matrix consists of identity morphisms on the diagonal, and zeros elsewhere. -/ def id (M : Mat_ C) : Hom M M := fun i j => if h : i = j then eqToHom (congr_arg M.X h) else 0 @@ -98,8 +96,12 @@ instance : Category.{v₁} (Mat_ C) where Hom := Hom id := Hom.id comp f g := f.comp g - id_comp f := by simp (config := { unfoldPartialApp := true }) [dite_comp] - comp_id f := by simp (config := { unfoldPartialApp := true }) [comp_dite] + id_comp f := by + classical + simp (config := { unfoldPartialApp := true }) [dite_comp] + comp_id f := by + classical + simp (config := { unfoldPartialApp := true }) [comp_dite] assoc f g h := by apply DMatrix.ext intros @@ -110,10 +112,12 @@ instance : Category.{v₁} (Mat_ C) where theorem hom_ext {M N : Mat_ C} (f g : M ⟶ N) (H : ∀ i j, f i j = g i j) : f = g := DMatrix.ext_iff.mp H +open scoped Classical in theorem id_def (M : Mat_ C) : (𝟙 M : Hom M M) = fun i j => if h : i = j then eqToHom (congr_arg M.X h) else 0 := rfl +open scoped Classical in theorem id_apply (M : Mat_ C) (i j : M.ι) : (𝟙 M : Hom M M) i j = if h : i = j then eqToHom (congr_arg M.X h) else 0 := rfl @@ -155,6 +159,7 @@ instance : Preadditive (Mat_ C) where open CategoryTheory.Limits +open scoped Classical in /-- We now prove that `Mat_ C` has finite biproducts. Be warned, however, that `Mat_ C` is not necessarily Krull-Schmidt, @@ -249,6 +254,7 @@ def mapMat_ (F : C ⥤ D) [Functor.Additive F] : Mat_ C ⥤ Mat_ D where @[simps!] def mapMatId : (𝟭 C).mapMat_ ≅ 𝟭 (Mat_ C) := NatIso.ofComponents (fun M => eqToIso (by cases M; rfl)) fun {M N} f => by + classical ext cases M; cases N simp [comp_dite, dite_comp] @@ -259,6 +265,7 @@ def mapMatId : (𝟭 C).mapMat_ ≅ 𝟭 (Mat_ C) := def mapMatComp {E : Type*} [Category.{v₁} E] [Preadditive E] (F : C ⥤ D) [Functor.Additive F] (G : D ⥤ E) [Functor.Additive G] : (F ⋙ G).mapMat_ ≅ F.mapMat_ ⋙ G.mapMat_ := NatIso.ofComponents (fun M => eqToIso (by cases M; rfl)) fun {M N} f => by + classical ext cases M; cases N simp [comp_dite, dite_comp] @@ -294,6 +301,7 @@ open CategoryTheory.Limits variable {C} +open scoped Classical in /-- Every object in `Mat_ C` is isomorphic to the biproduct of its summands. -/ @[simps] @@ -365,6 +373,7 @@ theorem additiveObjIsoBiproduct_naturality (F : Mat_ C ⥤ D) [Functor.Additive F.map f ≫ (additiveObjIsoBiproduct F N).hom = (additiveObjIsoBiproduct F M).hom ≫ biproduct.matrix fun i j => F.map ((embedding C).map (f i j)) := by + classical ext i : 1 simp only [Category.assoc, additiveObjIsoBiproduct_hom_π, isoBiproductEmbedding_hom, embedding_obj_ι, embedding_obj_X, biproduct.lift_π, biproduct.matrix_π, @@ -492,6 +501,7 @@ instance (R : Type u) : CoeSort (Mat R) (Type u) := open Matrix +open scoped Classical in instance (R : Type u) [Semiring R] : Category (Mat R) where Hom X Y := Matrix X Y R id X := (1 : Matrix X X R) @@ -510,9 +520,11 @@ theorem hom_ext {X Y : Mat R} (f g : X ⟶ Y) (h : ∀ i j, f i j = g i j) : f = variable (R) +open scoped Classical in theorem id_def (M : Mat R) : 𝟙 M = fun i j => if i = j then 1 else 0 := rfl +open scoped Classical in theorem id_apply (M : Mat R) (i j : M) : (𝟙 M : Matrix M M R) i j = if i = j then 1 else 0 := rfl diff --git a/Mathlib/CategoryTheory/Preadditive/Projective.lean b/Mathlib/CategoryTheory/Preadditive/Projective.lean index b576bee9c0e8a..8568390a4fb5b 100644 --- a/Mathlib/CategoryTheory/Preadditive/Projective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Projective.lean @@ -52,7 +52,6 @@ section /-- A projective presentation of an object `X` consists of an epimorphism `f : P ⟶ X` from some projective object `P`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure ProjectivePresentation (X : C) where p : C [projective : Projective p] diff --git a/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean b/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean index c743e88b1c70d..26b561b060a32 100644 --- a/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Preadditive/ProjectiveResolution.lean @@ -29,7 +29,6 @@ open Projective variable [HasZeroObject C] [HasZeroMorphisms C] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- A `ProjectiveResolution Z` consists of a bundled `ℕ`-indexed chain complex of projective objects, along with a quasi-isomorphism to the complex consisting of just `Z` supported in degree `0`. diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean index beedbdb8d0a67..82e861ee09f3f 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean @@ -48,7 +48,7 @@ def preadditiveYonedaObj (Y : C) : Cᵒᵖ ⥤ ModuleCat.{v} (End Y) where object `X` to the group of morphisms `X ⟶ Y`. At each point, we get an additional `End Y`-module structure, see `preadditiveYonedaObj`. -/ -@[simps] +@[simps obj] def preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGrp.{v} where obj Y := preadditiveYonedaObj Y ⋙ forget₂ _ _ map f := @@ -75,7 +75,7 @@ def preadditiveCoyonedaObj (X : Cᵒᵖ) : C ⥤ ModuleCat.{v} (End X) where object `Y` to the group of morphisms `X ⟶ Y`. At each point, we get an additional `End X`-module structure, see `preadditiveCoyonedaObj`. -/ -@[simps] +@[simps obj] def preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGrp.{v} where obj X := preadditiveCoyonedaObj X ⋙ forget₂ _ _ map f := @@ -88,10 +88,6 @@ def preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGrp.{v} where map_id _ := by ext; dsimp; simp map_comp f g := by ext; dsimp; simp --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] CategoryTheory.preadditiveYoneda_map_app_apply - CategoryTheory.preadditiveCoyoneda_map_app_apply - instance additive_yonedaObj (X : C) : Functor.Additive (preadditiveYonedaObj X) where instance additive_yonedaObj' (X : C) : Functor.Additive (preadditiveYoneda.obj X) where diff --git a/Mathlib/CategoryTheory/Presentable/Basic.lean b/Mathlib/CategoryTheory/Presentable/Basic.lean new file mode 100644 index 0000000000000..0ffacc86a36fc --- /dev/null +++ b/Mathlib/CategoryTheory/Presentable/Basic.lean @@ -0,0 +1,126 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Filtered.Basic +import Mathlib.CategoryTheory.Limits.Preserves.Basic +import Mathlib.CategoryTheory.Presentable.IsCardinalFiltered +import Mathlib.SetTheory.Cardinal.Cofinality +import Mathlib.SetTheory.Cardinal.HasCardinalLT + +/-! # Presentable objects + +A functor `F : C ⥤ D` is `κ`-accessible (`Functor.IsCardinalAccessible`) +if it commutes with colimits of shape `J` where `J` is any `κ`-filtered category +(that is essentially small relative to the universe `w` such that `κ : Cardinal.{w}`.). +We also introduce another typeclass `Functor.IsAccessible` saying that there exists +a regular cardinal `κ` such that `Functor.IsCardinalAccessible`. + +An object `X` of a category is `κ`-presentable (`IsCardinalPresentable`) +if the functor `Hom(X, _)` (i.e. `coyoneda.obj (op X)`) is `κ`-accessible. +Similar as for accessible functors, we define a type class `IsAccessible`. + +## References +* [Adámek, J. and Rosický, J., *Locally presentable and accessible categories*][Adamek_Rosicky_1994] + +-/ + +universe w w' v'' v' v u'' u' u + +namespace CategoryTheory + +open Limits Opposite + +variable {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D] + +namespace Functor + +variable (F G : C ⥤ D) (e : F ≅ G) (κ : Cardinal.{w}) [Fact κ.IsRegular] + +/-- A functor `F : C ⥤ D` is `κ`-accessible (with `κ` a regular cardinal) +if it preserves colimits of shape `J` where `J` is any `κ`-filtered category. +In the mathematical literature, some assumptions are often made on the +categories `C` or `D` (e.g. the existence of `κ`-filtered colimits, +see `HasCardinalFilteredColimits` below), but here we do not +make such assumptions. -/ +class IsCardinalAccessible : Prop where + preservesColimitOfShape (J : Type w) [SmallCategory J] [IsCardinalFiltered J κ] : + PreservesColimitsOfShape J F + +lemma preservesColimitsOfShape_of_isCardinalAccessible [F.IsCardinalAccessible κ] + (J : Type w) [SmallCategory J] [IsCardinalFiltered J κ] : + PreservesColimitsOfShape J F := + IsCardinalAccessible.preservesColimitOfShape κ _ + +lemma preservesColimitsOfShape_of_isCardinalAccessible_of_essentiallySmall + [F.IsCardinalAccessible κ] + (J : Type u'') [Category.{v''} J] [EssentiallySmall.{w} J] [IsCardinalFiltered J κ] : + PreservesColimitsOfShape J F := by + have := IsCardinalFiltered.of_equivalence κ (equivSmallModel.{w} J) + have := F.preservesColimitsOfShape_of_isCardinalAccessible κ (SmallModel.{w} J) + exact preservesColimitsOfShape_of_equiv (equivSmallModel.{w} J).symm F + +variable {κ} in +lemma isCardinalAccessible_of_le + [F.IsCardinalAccessible κ] {κ' : Cardinal.{w}} [Fact κ'.IsRegular] (h : κ ≤ κ') : + F.IsCardinalAccessible κ' where + preservesColimitOfShape {J _ _} := by + have := IsCardinalFiltered.of_le J h + exact F.preservesColimitsOfShape_of_isCardinalAccessible κ J + +include e in +variable {F G} in +lemma isCardinalAccessible_of_natIso [F.IsCardinalAccessible κ] : G.IsCardinalAccessible κ where + preservesColimitOfShape J _ hκ := by + have := F.preservesColimitsOfShape_of_isCardinalAccessible κ J + exact preservesColimitsOfShape_of_natIso e + +end Functor + +variable (X : C) (Y : C) (e : X ≅ Y) (κ : Cardinal.{w}) [Fact κ.IsRegular] + +/-- An object `X` in a category is `κ`-presentable (for `κ` a regular cardinal) +when the functor `Hom(X, _)` preserves colimits indexed by +`κ`-filtered categories. -/ +abbrev IsCardinalPresentable : Prop := (coyoneda.obj (op X)).IsCardinalAccessible κ + +lemma preservesColimitsOfShape_of_isCardinalPresentable [IsCardinalPresentable X κ] + (J : Type w) [SmallCategory.{w} J] [IsCardinalFiltered J κ] : + PreservesColimitsOfShape J (coyoneda.obj (op X)) := + (coyoneda.obj (op X)).preservesColimitsOfShape_of_isCardinalAccessible κ J + +lemma preservesColimitsOfShape_of_isCardinalPresentable_of_essentiallySmall + [IsCardinalPresentable X κ] + (J : Type u'') [Category.{v''} J] [EssentiallySmall.{w} J] [IsCardinalFiltered J κ] : + PreservesColimitsOfShape J (coyoneda.obj (op X)) := + (coyoneda.obj (op X)).preservesColimitsOfShape_of_isCardinalAccessible_of_essentiallySmall κ J + +variable {κ} in +lemma isCardinalPresentable_of_le [IsCardinalPresentable X κ] + {κ' : Cardinal.{w}} [Fact κ'.IsRegular] (h : κ ≤ κ') : + IsCardinalPresentable X κ' := + (coyoneda.obj (op X)).isCardinalAccessible_of_le h + +include e in +variable {X Y} in +lemma isCardinalPresentable_of_iso [IsCardinalPresentable X κ] : IsCardinalPresentable Y κ := + Functor.isCardinalAccessible_of_natIso (coyoneda.mapIso e.symm.op) κ + +section + +variable (C) (κ : Cardinal.{w}) [Fact κ.IsRegular] + +/-- A category has `κ`-filtered colimits if it has colimits of shape `J` +for any `κ`-filtered category `J`. -/ +class HasCardinalFilteredColimits : Prop where + hasColimitsOfShape (J : Type w) [SmallCategory J] [IsCardinalFiltered J κ] : + HasColimitsOfShape J C := by intros; infer_instance + +attribute [instance] HasCardinalFilteredColimits.hasColimitsOfShape + +instance [HasColimitsOfSize.{w, w} C] : HasCardinalFilteredColimits.{w} C κ where + +end + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Presentable/IsCardinalFiltered.lean b/Mathlib/CategoryTheory/Presentable/IsCardinalFiltered.lean new file mode 100644 index 0000000000000..ae10f8303d505 --- /dev/null +++ b/Mathlib/CategoryTheory/Presentable/IsCardinalFiltered.lean @@ -0,0 +1,168 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.CategoryTheory.Filtered.Basic +import Mathlib.CategoryTheory.Limits.Shapes.WideEqualizers +import Mathlib.CategoryTheory.Comma.CardinalArrow +import Mathlib.SetTheory.Cardinal.Cofinality +import Mathlib.SetTheory.Cardinal.HasCardinalLT +import Mathlib.SetTheory.Cardinal.Arithmetic + +/-! # κ-filtered category + +If `κ` is a regular cardinal, we introduce the notion of `κ`-filtered +category `J`: it means that any functor `A ⥤ J` from a small category such +that `Arrow A` is of cardinality `< κ` admits a cocone. +This generalizes the notion of filtered category. +Indeed, we obtain the equivalence `IsCardinalFiltered J ℵ₀ ↔ IsFiltered J`. +The API is mostly parallel to that of filtered categories. + +A preordered type `J` is a `κ`-filtered category (i.e. `κ`-directed set) +if any subset of `J` of cardinality `< κ` has an upper bound. + +## References +* [Adámek, J. and Rosický, J., *Locally presentable and accessible categories*][Adamek_Rosicky_1994] + +-/ + +universe w v' v u' u + +namespace CategoryTheory + +open Limits Opposite + +/-- A category `J` is `κ`-filtered (for a regular cardinal `κ`) if +any functor `F : A ⥤ J` from a category `A` such that `HasCardinalLT (Arrow A) κ` +admits a cocone. -/ +class IsCardinalFiltered (J : Type u) [Category.{v} J] + (κ : Cardinal.{w}) [Fact κ.IsRegular] : Prop where + nonempty_cocone {A : Type w} [SmallCategory A] (F : A ⥤ J) + (hA : HasCardinalLT (Arrow A) κ) : Nonempty (Cocone F) + +lemma hasCardinalLT_arrow_walkingParallelFamily {T : Type u} + {κ : Cardinal.{w}} (hT : HasCardinalLT T κ) (hκ : Cardinal.aleph0 ≤ κ) : + HasCardinalLT (Arrow (WalkingParallelFamily T)) κ := by + simpa only [hasCardinalLT_iff_of_equiv (WalkingParallelFamily.arrowEquiv T), + hasCardinalLT_option_iff _ _ hκ] using hT + +namespace IsCardinalFiltered + +variable {J : Type u} [Category.{v} J] {κ : Cardinal.{w}} [hκ : Fact κ.IsRegular] + [IsCardinalFiltered J κ] + +/-- A choice of cocone for a functor `F : A ⥤ J` such that `HasCardinatLT (Arrow A) κ` +when `J` is a `κ`-filtered category, and `Arrow A` has cardinality `< κ`. -/ +noncomputable def cocone {A : Type v'} [Category.{u'} A] + (F : A ⥤ J) (hA : HasCardinalLT (Arrow A) κ) : + Cocone F := by + have := hA.small + have := small_of_small_arrow.{w} A + have := locallySmall_of_small_arrow.{w} A + let e := (Shrink.equivalence.{w} A).trans (ShrinkHoms.equivalence.{w} (Shrink.{w} A)) + exact (Cocones.equivalenceOfReindexing e.symm (Iso.refl _)).inverse.obj + (nonempty_cocone (κ := κ) (e.inverse ⋙ F) (by simpa)).some + +variable (J) in +lemma of_le {κ' : Cardinal.{w}} [Fact κ'.IsRegular] (h : κ' ≤ κ) : + IsCardinalFiltered J κ' where + nonempty_cocone F hA := ⟨cocone F (hA.of_le h)⟩ + +variable (κ) in +lemma of_equivalence {J' : Type u'} [Category.{v'} J'] (e : J ≌ J') : + IsCardinalFiltered J' κ where + nonempty_cocone F hA := ⟨e.inverse.mapCoconeInv (cocone (F ⋙ e.inverse) hA)⟩ + +section max + +variable {K : Type u'} (S : K → J) (hS : HasCardinalLT K κ) + +/-- If `S : K → J` is a family of objects of cardinality `< κ` in a `κ`-filtered category, +this is a choice of objects in `J` which is the target of a map from any of +the objects `S k`. -/ +noncomputable def max : J := + (cocone (Discrete.functor S) (by simpa using hS)).pt + +/-- If `S : K → J` is a family of objects of cardinality `< κ` in a `κ`-filtered category, +this is a choice of map `S k ⟶ max S hS` for any `k : K`. -/ +noncomputable def toMax (k : K) : + S k ⟶ max S hS := + (cocone (Discrete.functor S) (by simpa using hS)).ι.app ⟨k⟩ + +end max + +section coeq + +variable {K : Type v'} {j j' : J} (f : K → (j ⟶ j')) (hK : HasCardinalLT K κ) + +/-- Given a family of maps `f : K → (j ⟶ j')` in a `κ`-filtered category `J`, +with `HasCardinalLT K κ`, this is an object of `J` where these morphisms +shall be equalized. -/ +noncomputable def coeq : J := + (cocone (parallelFamily f) + (hasCardinalLT_arrow_walkingParallelFamily hK hκ.out.aleph0_le)).pt + +/-- Given a family of maps `f : K → (j ⟶ j')` in a `κ`-filtered category `J`, +with `HasCardinalLT K κ`, and `k : K`, this is a choice of morphism `j' ⟶ coeq f hK`. -/ +noncomputable def coeqHom : j' ⟶ coeq f hK := + (cocone (parallelFamily f) + (hasCardinalLT_arrow_walkingParallelFamily hK hκ.out.aleph0_le)).ι.app .one + +/-- Given a family of maps `f : K → (j ⟶ j')` in a `κ`-filtered category `J`, +with `HasCardinalLT K κ`, this is a morphism `j ⟶ coeq f hK` which is equal +to all compositions `f k ≫ coeqHom f hK` for `k : K`. -/ +noncomputable def toCoeq : j ⟶ coeq f hK := + (cocone (parallelFamily f) + (hasCardinalLT_arrow_walkingParallelFamily hK hκ.out.aleph0_le)).ι.app .zero + +@[reassoc] +lemma coeq_condition (k : K) : f k ≫ coeqHom f hK = toCoeq f hK := + (cocone (parallelFamily f) + (hasCardinalLT_arrow_walkingParallelFamily hK hκ.out.aleph0_le)).w + (.line k) + +end coeq + +end IsCardinalFiltered + +open IsCardinalFiltered in +lemma isFiltered_of_isCardinalDirected (J : Type u) [Category.{v} J] + (κ : Cardinal.{w}) [hκ : Fact κ.IsRegular] [IsCardinalFiltered J κ] : + IsFiltered J := by + rw [IsFiltered.iff_cocone_nonempty.{w}] + intro A _ _ F + have hA : HasCardinalLT (Arrow A) κ := by + refine HasCardinalLT.of_le ?_ hκ.out.aleph0_le + simp only [hasCardinalLT_aleph0_iff] + infer_instance + exact ⟨cocone F hA⟩ + +attribute [local instance] Cardinal.fact_isRegular_aleph0 + +lemma isCardinalFiltered_aleph0_iff (J : Type u) [Category.{v} J] : + IsCardinalFiltered J Cardinal.aleph0.{w} ↔ IsFiltered J := by + constructor + · intro + exact isFiltered_of_isCardinalDirected J Cardinal.aleph0 + · intro + constructor + intro A _ F hA + rw [hasCardinalLT_aleph0_iff] at hA + have := ((Arrow.finite_iff A).1 hA).some + exact ⟨IsFiltered.cocone F⟩ + +lemma isCardinalFiltered_preorder (J : Type w) [Preorder J] + (κ : Cardinal.{w}) [Fact κ.IsRegular] + (h : ∀ ⦃K : Type w⦄ (s : K → J) (_ : Cardinal.mk K < κ), + ∃ (j : J), ∀ (k : K), s k ≤ j) : + IsCardinalFiltered J κ where + nonempty_cocone {A _ F hA} := by + obtain ⟨j, hj⟩ := h F.obj (by simpa only [hasCardinalLT_iff_cardinal_mk_lt] using + hasCardinalLT_of_hasCardinalLT_arrow hA) + exact ⟨Cocone.mk j + { app a := homOfLE (hj a) + naturality _ _ _ := rfl }⟩ + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Products/Basic.lean b/Mathlib/CategoryTheory/Products/Basic.lean index 0894a2421f869..c32a60a69ef2c 100644 --- a/Mathlib/CategoryTheory/Products/Basic.lean +++ b/Mathlib/CategoryTheory/Products/Basic.lean @@ -35,11 +35,8 @@ section variable (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] -- the generates simp lemmas like `id_fst` and `comp_snd` -/-- `prod C D` gives the cartesian product of two categories. - -See . --/ -@[simps (config := { notRecursive := [] }) Hom id_fst id_snd comp_fst comp_snd] +/-- `prod C D` gives the cartesian product of two categories. -/ +@[simps (config := { notRecursive := [] }) Hom id_fst id_snd comp_fst comp_snd, stacks 001K] instance prod : Category.{max v₁ v₂} (C × D) where Hom X Y := (X.1 ⟶ Y.1) × (X.2 ⟶ Y.2) id X := ⟨𝟙 X.1, 𝟙 X.2⟩ diff --git a/Mathlib/CategoryTheory/Retract.lean b/Mathlib/CategoryTheory/Retract.lean index 952554b768d5a..992d4761e7632 100644 --- a/Mathlib/CategoryTheory/Retract.lean +++ b/Mathlib/CategoryTheory/Retract.lean @@ -53,6 +53,19 @@ instance : IsSplitEpi h.r := ⟨⟨h.splitEpi⟩⟩ instance : IsSplitMono h.i := ⟨⟨h.splitMono⟩⟩ +variable (X) in +/-- Any object is a retract of itself. -/ +@[simps] +def refl : Retract X X where + i := 𝟙 X + r := 𝟙 X + +/-- A retract of a retract is a retract. -/ +@[simps] +def trans {Z : C} (h' : Retract Y Z) : Retract X Z where + i := h.i ≫ h'.i + r := h'.r ≫ h.r + end Retract /-- @@ -103,4 +116,14 @@ instance : IsSplitMono h.i.right := ⟨⟨h.right.splitMono⟩⟩ end RetractArrow +namespace Iso + +/-- If `X` is isomorphic to `Y`, then `X` is a retract of `Y`. -/ +@[simps] +def retract {X Y : C} (e : X ≅ Y) : Retract X Y where + i := e.hom + r := e.inv + +end Iso + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index d08d8b67edc4f..48ff0cc129d93 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -61,7 +61,6 @@ class HasShift (C : Type u) (A : Type*) [Category.{v} C] [AddMonoid A] where /-- `shift` is monoidal -/ shiftMonoidal : shift.Monoidal := by infer_instance --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- A helper structure to construct the shift functor `(Discrete A) ⥤ (C ⥤ C)`. -/ structure ShiftMkCore where /-- the family of shift functors -/ diff --git a/Mathlib/CategoryTheory/SingleObj.lean b/Mathlib/CategoryTheory/SingleObj.lean index 48fed64ab7d7d..b4844a38d181a 100644 --- a/Mathlib/CategoryTheory/SingleObj.lean +++ b/Mathlib/CategoryTheory/SingleObj.lean @@ -74,10 +74,8 @@ theorem comp_as_mul {x y z : SingleObj M} (f : x ⟶ y) (g : y ⟶ z) : f ≫ g /-- If `M` is finite and in universe zero, then `SingleObj M` is a `FinCategory`. -/ instance finCategoryOfFintype (M : Type) [Fintype M] [Monoid M] : FinCategory (SingleObj M) where -/-- Groupoid structure on `SingleObj M`. - -See . --/ +/-- Groupoid structure on `SingleObj M`. -/ +@[stacks 0019] instance groupoid : Groupoid (SingleObj G) where inv x := x⁻¹ inv_comp := mul_inv_cancel @@ -104,12 +102,8 @@ theorem toEnd_def (x : M) : toEnd M x = x := variable (N : Type v) [Monoid N] /-- There is a 1-1 correspondence between monoid homomorphisms `M → N` and functors between the - corresponding single-object categories. It means that `SingleObj` is a fully faithful - functor. - -See -- -although we do not characterize when the functor is full or faithful. --/ +corresponding single-object categories. It means that `SingleObj` is a fully faithful functor. -/ +@[stacks 001F "We do not characterize when the functor is full or faithful."] def mapHom : (M →* N) ≃ SingleObj M ⥤ SingleObj N where toFun f := { obj := id diff --git a/Mathlib/CategoryTheory/Sites/Canonical.lean b/Mathlib/CategoryTheory/Sites/Canonical.lean index 1c9d6de317649..be99c646d4406 100644 --- a/Mathlib/CategoryTheory/Sites/Canonical.lean +++ b/Mathlib/CategoryTheory/Sites/Canonical.lean @@ -139,11 +139,9 @@ theorem isSheafFor_trans (P : Cᵒᵖ ⥤ Type v) (R S : Sieve X) rw [this] apply hR' hf -/-- Construct the finest (largest) Grothendieck topology for which the given presheaf is a sheaf. - -This is a special case of https://stacks.math.columbia.edu/tag/00Z9, but following a different -proof (see the comments there). --/ +/-- Construct the finest (largest) Grothendieck topology for which the given presheaf is a sheaf. -/ +@[stacks 00Z9 "This is a special case of the Stacks entry, but following a different +proof (see the Stacks comments)."] def finestTopologySingle (P : Cᵒᵖ ⥤ Type v) : GrothendieckTopology C where sieves X S := ∀ (Y) (f : Y ⟶ X), Presieve.IsSheafFor P (S.pullback f : Presieve Y) top_mem' X Y f := by @@ -164,11 +162,9 @@ def finestTopologySingle (P : Cᵒᵖ ⥤ Type v) : GrothendieckTopology C where rw [pullback_id, pullback_comp] at this apply this -/-- -Construct the finest (largest) Grothendieck topology for which all the given presheaves are sheaves. - -This is equal to the construction of . --/ +/-- Construct the finest (largest) Grothendieck topology for which all the given presheaves are +sheaves. -/ +@[stacks 00Z9 "Equal to that Stacks construction"] def finestTopology (Ps : Set (Cᵒᵖ ⥤ Type v)) : GrothendieckTopology C := sInf (finestTopologySingle '' Ps) @@ -188,10 +184,8 @@ theorem le_finestTopology (Ps : Set (Cᵒᵖ ⥤ Type v)) (J : GrothendieckTopol exact hJ P hP (S.pullback f) (J.pullback_stable f hS) /-- The `canonicalTopology` on a category is the finest (largest) topology for which every -representable presheaf is a sheaf. - -See --/ +representable presheaf is a sheaf. -/ +@[stacks 00ZA] def canonicalTopology (C : Type u) [Category.{v} C] : GrothendieckTopology C := finestTopology (Set.range yoneda.obj) diff --git a/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean b/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean index 2997401067b07..dad92afd416e4 100644 --- a/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean +++ b/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean @@ -34,7 +34,6 @@ variable [HasForget.{max v u} D] attribute [local instance] HasForget.hasCoeToSort HasForget.instFunLike --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- A concrete version of the multiequalizer, to be used below. -/ def Meq {X : C} (P : Cᵒᵖ ⥤ D) (S : J.Cover X) := { x : ∀ I : S.Arrow, P.obj (op I.Y) // diff --git a/Mathlib/CategoryTheory/Sites/CoverLifting.lean b/Mathlib/CategoryTheory/Sites/CoverLifting.lean index 631707e2b37e5..db34551ea0323 100644 --- a/Mathlib/CategoryTheory/Sites/CoverLifting.lean +++ b/Mathlib/CategoryTheory/Sites/CoverLifting.lean @@ -68,7 +68,6 @@ variable {L : GrothendieckTopology E} /-- A functor `G : (C, J) ⥤ (D, K)` between sites is called cocontinuous (SGA 4 III 2.1) if for all covering sieves `R` in `D`, `R.pullback G` is a covering sieve in `C`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` class Functor.IsCocontinuous : Prop where cover_lift : ∀ {U : C} {S : Sieve (G.obj U)} (_ : S ∈ K (G.obj U)), S.functorPullback G ∈ J U @@ -212,11 +211,9 @@ variable [∀ (F : Cᵒᵖ ⥤ A), G.op.HasPointwiseRightKanExtension F] /-- If `G` is cocontinuous, then `G.op.ran` pushes sheaves to sheaves. -This is SGA 4 III 2.2. An alternative reference is -https://stacks.math.columbia.edu/tag/00XK (where results -are obtained under the additional assumption that -`C` and `D` have pullbacks). --/ +This is SGA 4 III 2.2. -/ +@[stacks 00XK "Alternative reference. There, results are obtained under the additional assumption +that `C` and `D` have pullbacks."] theorem ran_isSheaf_of_isCocontinuous (ℱ : Sheaf J A) : Presheaf.IsSheaf K (G.op.ran.obj ℱ.val) := by rw [Presheaf.isSheaf_iff_multifork] diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 08f728dd0bf0e..f7d3acf8d8072 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -51,7 +51,6 @@ variable {L : GrothendieckTopology A} /-- A functor `G : (C, J) ⥤ (D, K)` between sites is *cover-preserving* if for all covering sieves `R` in `C`, `R.functorPushforward G` is a covering sieve in `D`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` structure CoverPreserving (G : C ⥤ D) : Prop where cover_preserve : ∀ {U : C} {S : Sieve U} (_ : S ∈ J U), S.functorPushforward G ∈ K (G.obj U) @@ -72,7 +71,6 @@ compatible family of elements at `C` and valued in `G.op ⋙ ℱ`, and each comm This is actually stronger than merely preserving compatible families because of the definition of `functorPushforward` used. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet @[nolint has_nonempty_instance] structure CompatiblePreserving (K : GrothendieckTopology D) (G : C ⥤ D) : Prop where compatible : ∀ (ℱ : Sheaf K (Type w)) {Z} {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} @@ -156,11 +154,8 @@ theorem compatiblePreservingOfDownwardsClosed (F : C ⥤ D) [F.Full] [F.Faithful variable {F J K} -/-- If `F` is cover-preserving and compatible-preserving, -then `F` is a continuous functor. - -This result is basically . --/ +/-- If `F` is cover-preserving and compatible-preserving, then `F` is a continuous functor. -/ +@[stacks 00WW "This is basically this Stacks entry."] lemma Functor.isContinuous_of_coverPreserving (hF₁ : CompatiblePreserving.{w} K F) (hF₂ : CoverPreserving J K F) : Functor.IsContinuous.{w} F J K where op_comp_isSheaf_of_types G X S hS x hx := by diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean index e857f880260dd..4d65746a7c821 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean @@ -48,7 +48,6 @@ variable {L : GrothendieckTopology E} /-- An auxiliary structure that witnesses the fact that `f` factors through an image object of `G`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` structure Presieve.CoverByImageStructure (G : C ⥤ D) {V U : D} (f : V ⟶ U) where obj : C lift : V ⟶ G.obj obj diff --git a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean index 29b141d4386aa..c52235dba062f 100644 --- a/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean +++ b/Mathlib/CategoryTheory/Sites/EqualizerSheafCondition.lean @@ -45,8 +45,9 @@ noncomputable section /-- The middle object of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . +of the Stacks entry. -/ +@[stacks 00VM "This is the middle object of the fork diagram there."] def FirstObj : Type max v u := ∏ᶜ fun f : ΣY, { f : Y ⟶ X // R f } => P.obj (op f.1) @@ -78,8 +79,9 @@ instance : Inhabited (FirstObj P ((⊥ : Sieve X) : Presieve X)) := /-- The left morphism of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . +of the Stacks entry. -/ +@[stacks 00VM "This is the left morphism of the fork diagram there."] def forkMap : P.obj (op X) ⟶ FirstObj P R := Pi.lift fun f => P.map f.2.1.op @@ -176,15 +178,17 @@ namespace Presieve variable [R.hasPullbacks] /-- -The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM, which +The rightmost object of the fork diagram of the Stacks entry, which contains the data used to check a family of elements for a presieve is compatible. -/ -@[simp] def SecondObj : Type max v u := +@[simp, stacks 00VM "This is the rightmost object of the fork diagram there."] +def SecondObj : Type max v u := ∏ᶜ fun fg : (ΣY, { f : Y ⟶ X // R f }) × ΣZ, { g : Z ⟶ X // R g } => haveI := Presieve.hasPullbacks.has_pullbacks fg.1.2.2 fg.2.2.2 P.obj (op (pullback fg.1.2.1 fg.2.2.1)) -/-- The map `pr₀*` of . -/ +/-- The map `pr₀*` of the Stacks entry. -/ +@[stacks 00VM "This is the map `pr₀*` there."] def firstMap : FirstObj P R ⟶ SecondObj P R := Pi.lift fun fg => haveI := Presieve.hasPullbacks.has_pullbacks fg.1.2.2 fg.2.2.2 @@ -193,7 +197,8 @@ def firstMap : FirstObj P R ⟶ SecondObj P R := instance [HasPullbacks C] : Inhabited (SecondObj P (⊥ : Presieve X)) := ⟨firstMap _ _ default⟩ -/-- The map `pr₁*` of . -/ +/-- The map `pr₁*` of the Stacks entry. -/ +@[stacks 00VM "This is the map `pr₁*` there."] def secondMap : FirstObj P R ⟶ SecondObj P R := Pi.lift fun fg => haveI := Presieve.hasPullbacks.has_pullbacks fg.1.2.2 fg.2.2.2 @@ -224,9 +229,8 @@ theorem compatible_iff (x : FirstObj P R) : rw [Types.limit_ext_iff'] at t simpa [firstMap, secondMap] using t ⟨⟨⟨Y, f, hf⟩, Z, g, hg⟩⟩ -/-- `P` is a sheaf for `R`, iff the fork given by `w` is an equalizer. -See . --/ +/-- `P` is a sheaf for `R`, iff the fork given by `w` is an equalizer. -/ +@[stacks 00VM] theorem sheaf_condition : R.IsSheafFor P ↔ Nonempty (IsLimit (Fork.ofι _ (w P R))) := by rw [Types.type_equalizer_iff_unique, ← Equiv.forall_congr_right (firstObjEqFamily P R).toEquiv.symm] @@ -255,10 +259,11 @@ variable {B : C} {I : Type} (X : I → C) (π : (i : I) → X i ⟶ B) -- TODO: allow `I : Type w` /-- -The middle object of the fork diagram of . +The middle object of the fork diagram of the Stacks entry. The difference between this and `Equalizer.FirstObj P (ofArrows X π)` arises if the family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. -/ +@[stacks 00VM "The middle object of the fork diagram there."] def FirstObj : Type w := ∏ᶜ (fun i ↦ P.obj (op (X i))) @[ext] @@ -269,10 +274,11 @@ lemma FirstObj.ext (z₁ z₂ : FirstObj P X) (h : ∀ i, (Pi.π _ i : FirstObj exact h i /-- -The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM. +The rightmost object of the fork diagram of the Stacks entry. The difference between this and `Equalizer.Presieve.SecondObj P (ofArrows X π)` arises if the family of arrows `π` contains duplicates. The `Presieve.ofArrows` doesn't see those. -/ +@[stacks 00VM "The rightmost object of the fork diagram there."] def SecondObj : Type w := ∏ᶜ (fun (ij : I × I) ↦ P.obj (op (pullback (π ij.1) (π ij.2)))) @@ -322,10 +328,8 @@ theorem compatible_iff (x : FirstObj P X) : (Arrows.Compatible P π ((Types.prod apply_fun Pi.π (fun (ij : I × I) ↦ P.obj (op (pullback (π ij.1) (π ij.2)))) ⟨i, j⟩ at t simpa [firstMap, secondMap] using t -/-- -`P` is a sheaf for `Presieve.ofArrows X π`, iff the fork given by `w` is an equalizer. -See . --/ +/-- `P` is a sheaf for `Presieve.ofArrows X π`, iff the fork given by `w` is an equalizer. -/ +@[stacks 00VM] theorem sheaf_condition : (Presieve.ofArrows X π).IsSheafFor P ↔ Nonempty (IsLimit (Fork.ofι (forkMap P X π) (w P X π))) := by rw [Types.type_equalizer_iff_unique, isSheafFor_arrows_iff] diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 5570569b02e2b..816a9f82f51ce 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -66,9 +66,8 @@ three axioms: A sieve `S` on `X` is referred to as `J`-covering, (or just covering), if `S ∈ J X`. -See , or [nlab], or [MM92][] Chapter III, Section 2, -Definition 1. --/ +See also [nlab] or [MM92] Chapter III, Section 2, Definition 1. -/ +@[stacks 00Z4] structure GrothendieckTopology where /-- A Grothendieck topology on `C` consists of a set of sieves for each object `X`, which satisfy some axioms. -/ @@ -132,9 +131,8 @@ theorem covering_of_eq_top : S = ⊤ → S ∈ J X := fun h => h.symm ▸ J.top_ /-- If `S` is a subset of `R`, and `S` is covering, then `R` is covering as well. -See (2), or discussion after [MM92] Chapter III, -Section 2, Definition 1. --/ +See also discussion after [MM92] Chapter III, Section 2, Definition 1. -/ +@[stacks 00Z5 "(2)"] theorem superset_covering (Hss : S ≤ R) (sjx : S ∈ J X) : R ∈ J X := by apply J.transitive sjx R fun Y f hf => _ intros Y f hf @@ -144,9 +142,8 @@ theorem superset_covering (Hss : S ≤ R) (sjx : S ∈ J X) : R ∈ J X := by /-- The intersection of two covering sieves is covering. -See (1), or [MM92] Chapter III, -Section 2, Definition 1 (iv). --/ +See also [MM92] Chapter III, Section 2, Definition 1 (iv). -/ +@[stacks 00Z5 "(1)"] theorem intersection_covering (rj : R ∈ J X) (sj : S ∈ J X) : R ⊓ S ∈ J X := by apply J.transitive rj _ fun Y f Hf => _ intros Y f hf @@ -231,21 +228,21 @@ variable {C} theorem trivial_covering : S ∈ trivial C X ↔ S = ⊤ := Set.mem_singleton_iff -/-- See -/ +@[stacks 00Z6] instance instLEGrothendieckTopology : LE (GrothendieckTopology C) where le J₁ J₂ := (J₁ : ∀ X : C, Set (Sieve X)) ≤ (J₂ : ∀ X : C, Set (Sieve X)) theorem le_def {J₁ J₂ : GrothendieckTopology C} : J₁ ≤ J₂ ↔ (J₁ : ∀ X : C, Set (Sieve X)) ≤ J₂ := Iff.rfl -/-- See -/ +@[stacks 00Z6] instance : PartialOrder (GrothendieckTopology C) := { instLEGrothendieckTopology with le_refl := fun _ => le_def.mpr le_rfl le_trans := fun _ _ _ h₁₂ h₂₃ => le_def.mpr (le_trans h₁₂ h₂₃) le_antisymm := fun _ _ h₁₂ h₂₁ => GrothendieckTopology.ext (le_antisymm h₁₂ h₂₁) } -/-- See -/ +@[stacks 00Z7] instance : InfSet (GrothendieckTopology C) where sInf T := { sieves := sInf (sieves '' T) @@ -265,7 +262,7 @@ lemma mem_sInf (s : Set (GrothendieckTopology C)) {X : C} (S : Sieve X) : show S ∈ sInf (sieves '' s) X ↔ _ simp -/-- See -/ +@[stacks 00Z7] theorem isGLB_sInf (s : Set (GrothendieckTopology C)) : IsGLB s (sInf s) := by refine @IsGLB.of_image _ _ _ _ sieves ?_ _ _ ?_ · #adaptation_note @@ -413,8 +410,6 @@ instance : Inhabited (J.Cover X) := ⟨⊤⟩ /-- An auxiliary structure, used to define `S.index`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] @[ext] structure Arrow (S : J.Cover X) where /-- The source of the arrow. -/ @@ -551,8 +546,6 @@ theorem Arrow.middle_spec {X : C} {S : J.Cover X} {T : ∀ I : S.Arrow, J.Cover I.hf.choose_spec.choose_spec.choose_spec.choose_spec.2 /-- An auxiliary structure, used to define `S.index`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance, ext] @[ext] structure Relation (S : J.Cover X) where /-- The first arrow. -/ diff --git a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean index 0d7effb4b905f..e229cdcd97a5f 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean @@ -83,9 +83,10 @@ A presheaf is a sheaf (resp, separated) if every *compatible* family of elements (resp, at most one) amalgamation. This data is referred to as a `family` in [MM92], Chapter III, Section 4. It is also a concrete -version of the elements of the middle object in https://stacks.math.columbia.edu/tag/00VM which is +version of the elements of the middle object in the Stacks entry which is more useful for direct calculations. It is also used implicitly in Definition C2.1.2 in [Elephant]. -/ +@[stacks 00VM "This is a concrete version of the elements of the middle object there."] def FamilyOfElements (P : Cᵒᵖ ⥤ Type w) (R : Presieve X) := ∀ ⦃Y : C⦄ (f : Y ⟶ X), R f → P.obj (op Y) @@ -416,8 +417,8 @@ This version is also useful to establish that being a sheaf is preserved under i presheaves. See the discussion before Equation (3) of [MM92], Chapter III, Section 4. See also C2.1.4 of -[Elephant]. This is also a direct reformulation of . --/ +[Elephant]. -/ +@[stacks 00Z8 "Direct reformulation"] def YonedaSheafCondition (P : Cᵒᵖ ⥤ Type v₁) (S : Sieve X) : Prop := ∀ f : S.functor ⟶ P, ∃! g, S.functorInclusion ≫ g = f diff --git a/Mathlib/CategoryTheory/Sites/LocallyInjective.lean b/Mathlib/CategoryTheory/Sites/LocallyInjective.lean index 64b1d4140a669..9f60ac73c6f1a 100644 --- a/Mathlib/CategoryTheory/Sites/LocallyInjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallyInjective.lean @@ -243,8 +243,8 @@ lemma mono_of_isLocallyInjective [IsLocallyInjective φ] : Mono φ := by infer_instance instance {F G : Sheaf J (Type w)} (f : F ⟶ G) : - IsLocallyInjective (imageSheafι f) := by - dsimp [imageSheafι] + IsLocallyInjective (Sheaf.imageι f) := by + dsimp [Sheaf.imageι] infer_instance end Sheaf diff --git a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean index e3c5a1393050f..a4fe2b6272f33 100644 --- a/Mathlib/CategoryTheory/Sites/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallySurjective.lean @@ -50,7 +50,7 @@ def imageSieve {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : rw [op_comp, G.map_comp, comp_apply, ← ht, elementwise_of% f.naturality] theorem imageSieve_eq_sieveOfSection {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : - imageSieve f s = (imagePresheaf (whiskerRight f (forget A))).sieveOfSection s := + imageSieve f s = (Subpresheaf.range (whiskerRight f (forget A))).sieveOfSection s := rfl theorem imageSieve_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : @@ -92,13 +92,13 @@ instance {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) [IsLocallySurjective J f] : imageSieve_mem s := imageSieve_mem J f s theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : - IsLocallySurjective J f ↔ (imagePresheaf (whiskerRight f (forget A))).sheafify J = ⊤ := by + IsLocallySurjective J f ↔ (Subpresheaf.range (whiskerRight f (forget A))).sheafify J = ⊤ := by simp only [Subpresheaf.ext_iff, funext_iff, Set.ext_iff, top_subpresheaf_obj, Set.top_eq_univ, Set.mem_univ, iff_true] exact ⟨fun H _ => H.imageSieve_mem, fun H => ⟨H _⟩⟩ theorem isLocallySurjective_iff_imagePresheaf_sheafify_eq_top' {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) : - IsLocallySurjective J f ↔ (imagePresheaf f).sheafify J = ⊤ := by + IsLocallySurjective J f ↔ (Subpresheaf.range f).sheafify J = ⊤ := by apply isLocallySurjective_iff_imagePresheaf_sheafify_eq_top theorem isLocallySurjective_iff_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : @@ -248,7 +248,7 @@ lemma isLocallySurjective_comp_iff infer_instance instance {F₁ F₂ : Cᵒᵖ ⥤ Type w} (f : F₁ ⟶ F₂) : - IsLocallySurjective J (toImagePresheafSheafify J f) where + IsLocallySurjective J (Subpresheaf.toRangeSheafify J f) where imageSieve_mem {X} := by rintro ⟨s, hs⟩ refine J.superset_covering ?_ hs @@ -257,20 +257,20 @@ instance {F₁ F₂ : Cᵒᵖ ⥤ Type w} (f : F₁ ⟶ F₂) : /-- The image of `F` in `J.sheafify F` is isomorphic to the sheafification. -/ noncomputable def sheafificationIsoImagePresheaf (F : Cᵒᵖ ⥤ Type max u v) : - J.sheafify F ≅ ((imagePresheaf (J.toSheafify F)).sheafify J).toPresheaf where + J.sheafify F ≅ ((Subpresheaf.range (J.toSheafify F)).sheafify J).toPresheaf where hom := - J.sheafifyLift (toImagePresheafSheafify J _) + J.sheafifyLift (Subpresheaf.toRangeSheafify J _) ((isSheaf_iff_isSheaf_of_type J _).mpr <| Subpresheaf.sheafify_isSheaf _ <| (isSheaf_iff_isSheaf_of_type J _).mp <| GrothendieckTopology.sheafify_isSheaf J _) inv := Subpresheaf.ι _ hom_inv_id := - J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [toImagePresheafSheafify]) + J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [Subpresheaf.toRangeSheafify]) inv_hom_id := by rw [← cancel_mono (Subpresheaf.ι _), Category.id_comp, Category.assoc] refine Eq.trans ?_ (Category.comp_id _) congr 1 - exact J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [toImagePresheafSheafify]) + exact J.sheafify_hom_ext _ _ (J.sheafify_isSheaf _) (by simp [Subpresheaf.toRangeSheafify]) section @@ -328,8 +328,8 @@ instance isLocallySurjective_of_iso [IsIso φ] : IsLocallySurjective φ := by infer_instance instance {F G : Sheaf J (Type w)} (f : F ⟶ G) : - IsLocallySurjective (toImageSheaf f) := by - dsimp [toImageSheaf] + IsLocallySurjective (Sheaf.toImage f) := by + dsimp [Sheaf.toImage] infer_instance variable [J.HasSheafCompose (forget A)] @@ -339,11 +339,11 @@ instance [IsLocallySurjective φ] : (Presheaf.isLocallySurjective_iff_whisker_forget J φ.val).1 inferInstance theorem isLocallySurjective_iff_isIso {F G : Sheaf J (Type w)} (f : F ⟶ G) : - IsLocallySurjective f ↔ IsIso (imageSheafι f) := by + IsLocallySurjective f ↔ IsIso (Sheaf.imageι f) := by dsimp only [IsLocallySurjective] - rw [imageSheafι, Presheaf.isLocallySurjective_iff_imagePresheaf_sheafify_eq_top', + rw [Sheaf.imageι, Presheaf.isLocallySurjective_iff_imagePresheaf_sheafify_eq_top', Subpresheaf.eq_top_iff_isIso] - exact isIso_iff_of_reflects_iso (f := imageSheafι f) (F := sheafToPresheaf J (Type w)) + exact isIso_iff_of_reflects_iso (f := Sheaf.imageι f) (F := sheafToPresheaf J (Type w)) instance epi_of_isLocallySurjective' {F₁ F₂ : Sheaf J (Type w)} (φ : F₁ ⟶ F₂) [IsLocallySurjective φ] : Epi φ where @@ -369,7 +369,7 @@ lemma isLocallySurjective_iff_epi {F G : Sheaf J (Type w)} (φ : F ⟶ G) · intro infer_instance · intro - have := epi_of_epi_fac (toImageSheaf_ι φ) + have := epi_of_epi_fac (Sheaf.toImage_ι φ) rw [isLocallySurjective_iff_isIso φ] apply isIso_of_mono_of_epi diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index bde2d0fb7bd77..1732a166ae6f3 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -52,12 +52,10 @@ three axioms: In some sense, a pretopology can be seen as Grothendieck topology with weaker saturation conditions, in that each covering is not necessarily downward closed. -See: https://ncatlab.org/nlab/show/Grothendieck+pretopology, or -https://stacks.math.columbia.edu/tag/00VH, or [MM92] Chapter III, Section 2, Definition 2. -Note that Stacks calls a category together with a pretopology a site, and [MM92] calls this -a basis for a topology. --/ -@[ext] +See: https://ncatlab.org/nlab/show/Grothendieck+pretopology or [MM92] Chapter III, +Section 2, Definition 2. -/ +@[ext, stacks 00VH "Note that Stacks calls a category together with a pretopology a site, +and [MM92] calls this a basis for a topology."] structure Pretopology where coverings : ∀ X : C, Set (Presieve X) has_isos : ∀ ⦃X Y⦄ (f : Y ⟶ X) [IsIso f], Presieve.singleton f ∈ coverings X @@ -101,8 +99,9 @@ instance : Inhabited (Pretopology C) := /-- A pretopology `K` can be completed to a Grothendieck topology `J` by declaring a sieve to be `J`-covering if it contains a family in `K`. -See , or [MM92] Chapter III, Section 2, Equation (2). +See also [MM92] Chapter III, Section 2, Equation (2). -/ +@[stacks 00ZC] def toGrothendieck (K : Pretopology C) : GrothendieckTopology C where sieves X S := ∃ R ∈ K X, R ≤ (S : Presieve _) top_mem' _ := ⟨Presieve.singleton (𝟙 _), K.has_isos _, fun _ _ _ => ⟨⟩⟩ @@ -162,10 +161,8 @@ lemma mem_ofGrothendieck (t : GrothendieckTopology C) {X : C} (S : Presieve X) : /-- The trivial pretopology, in which the coverings are exactly singleton isomorphisms. This topology is -also known as the indiscrete, coarse, or chaotic topology. - -See --/ +also known as the indiscrete, coarse, or chaotic topology. -/ +@[stacks 07GE] def trivial : Pretopology C where coverings X S := ∃ (Y : _) (f : Y ⟶ X) (_ : IsIso f), S = Presieve.singleton f has_isos _ _ _ i := ⟨_, _, i, rfl⟩ diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 6f528a0d3f6a1..b213f408e9c07 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -68,10 +68,8 @@ variable (J : GrothendieckTopology C) -- We follow https://stacks.math.columbia.edu/tag/00VL definition 00VR /-- A sheaf of A is a presheaf P : Cᵒᵖ => A such that for every E : A, the -presheaf of types given by sending U : C to Hom_{A}(E, P U) is a sheaf of types. - -https://stacks.math.columbia.edu/tag/00VR --/ +presheaf of types given by sending U : C to Hom_{A}(E, P U) is a sheaf of types. -/ +@[stacks 00VR] def IsSheaf (P : Cᵒᵖ ⥤ A) : Prop := ∀ E : A, Presieve.IsSheaf J (P ⋙ coyoneda.obj (op E)) @@ -92,7 +90,6 @@ variable (P : Cᵒᵖ ⥤ A) {X : C} (S : Sieve X) (R : Presieve X) (E : Aᵒᵖ the cones over the natural diagram `S.arrows.diagram.op ⋙ P` associated to `S` and `P` with cone point `E` are in 1-1 correspondence with sieve_compatible family of elements for the sieve `S` and the presheaf of types `Hom (E, P -)`. -/ -@[simps] def conesEquivSieveCompatibleFamily : (S.arrows.diagram.op ⋙ P).cones.obj E ≃ { x : FamilyOfElements (P ⋙ coyoneda.obj E) (S : Presieve X) // x.SieveCompatible } where @@ -112,10 +109,6 @@ def conesEquivSieveCompatibleFamily : left_inv _ := rfl right_inv _ := rfl --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] CategoryTheory.Presheaf.conesEquivSieveCompatibleFamily_apply_coe - CategoryTheory.Presheaf.conesEquivSieveCompatibleFamily_symm_apply_app - variable {P S E} variable {x : FamilyOfElements (P ⋙ coyoneda.obj E) S.arrows} (hx : SieveCompatible x) @@ -591,34 +584,35 @@ section variable [HasProducts.{max u₁ v₁} A] variable [HasProducts.{max u₁ v₁} A'] -/-- -The middle object of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . --/ +/-- The middle object of the fork diagram given in Equation (3) of [MM92], as well as the fork +diagram of the Stacks entry. -/ +@[stacks 00VM "The middle object of the fork diagram there."] def firstObj : A := ∏ᶜ fun f : ΣV, { f : V ⟶ U // R f } => P.obj (op f.1) -/-- -The left morphism of the fork diagram given in Equation (3) of [MM92], as well as the fork diagram -of . --/ +/-- The left morphism of the fork diagram given in Equation (3) of [MM92], as well as the fork +diagram of the Stacks entry. -/ +@[stacks 00VM "The left morphism the fork diagram there."] def forkMap : P.obj (op U) ⟶ firstObj R P := Pi.lift fun f => P.map f.2.1.op variable [HasPullbacks C] -/-- The rightmost object of the fork diagram of https://stacks.math.columbia.edu/tag/00VM, which +/-- The rightmost object of the fork diagram of the Stacks entry, which contains the data used to check a family of elements for a presieve is compatible. -/ +@[stacks 00VM "The rightmost object of the fork diagram there."] def secondObj : A := ∏ᶜ fun fg : (ΣV, { f : V ⟶ U // R f }) × ΣW, { g : W ⟶ U // R g } => P.obj (op (pullback fg.1.2.1 fg.2.2.1)) -/-- The map `pr₀*` of . -/ +/-- The map `pr₀*` of the Stacks entry. -/ +@[stacks 00VM "The map `pr₀*` there."] def firstMap : firstObj R P ⟶ secondObj R P := Pi.lift fun _ => Pi.π _ _ ≫ P.map (pullback.fst _ _).op -/-- The map `pr₁*` of . -/ +/-- The map `pr₁*` of the Stacks entry. -/ +@[stacks 00VM "The map `pr₁*` there."] def secondMap : firstObj R P ⟶ secondObj R P := Pi.lift fun _ => Pi.π _ _ ≫ P.map (pullback.snd _ _).op diff --git a/Mathlib/CategoryTheory/Sites/Subsheaf.lean b/Mathlib/CategoryTheory/Sites/Subsheaf.lean index 3e31f82d383ff..03d5262b3fe05 100644 --- a/Mathlib/CategoryTheory/Sites/Subsheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Subsheaf.lean @@ -8,7 +8,8 @@ import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono import Mathlib.Tactic.CategoryTheory.Elementwise import Mathlib.CategoryTheory.Adhesive import Mathlib.CategoryTheory.Sites.ConcreteSheafification -import Mathlib.CategoryTheory.Subpresheaf.Basic +import Mathlib.CategoryTheory.Subpresheaf.Image +import Mathlib.CategoryTheory.Subpresheaf.Sieves /-! @@ -193,16 +194,17 @@ variable (J) /-- A morphism factors through the sheafification of the image presheaf. -/ @[simps!] -def toImagePresheafSheafify (f : F' ⟶ F) : F' ⟶ ((imagePresheaf f).sheafify J).toPresheaf := - toImagePresheaf f ≫ Subpresheaf.homOfLe ((imagePresheaf f).le_sheafify J) +def Subpresheaf.toRangeSheafify (f : F' ⟶ F) : F' ⟶ ((Subpresheaf.range f).sheafify J).toPresheaf := + toRange f ≫ Subpresheaf.homOfLe ((range f).le_sheafify J) + variable {J} /-- The image sheaf of a morphism between sheaves, defined to be the sheafification of `image_presheaf`. -/ @[simps] -def imageSheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := - ⟨((imagePresheaf f.1).sheafify J).toPresheaf, by +def Sheaf.image {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := + ⟨((Subpresheaf.range f.1).sheafify J).toPresheaf, by rw [isSheaf_iff_isSheaf_of_type] apply Subpresheaf.sheafify_isSheaf rw [← isSheaf_iff_isSheaf_of_type] @@ -210,34 +212,35 @@ def imageSheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := /-- A morphism factors through the image sheaf. -/ @[simps] -def toImageSheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : F ⟶ imageSheaf f := - ⟨toImagePresheafSheafify J f.1⟩ +def Sheaf.toImage {F F' : Sheaf J (Type w)} (f : F ⟶ F') : F ⟶ Sheaf.image f := + ⟨Subpresheaf.toRangeSheafify J f.1⟩ /-- The inclusion of the image sheaf to the target. -/ @[simps] -def imageSheafι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : imageSheaf f ⟶ F' := +def Sheaf.imageι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf.image f ⟶ F' := ⟨Subpresheaf.ι _⟩ + @[reassoc (attr := simp)] -theorem toImageSheaf_ι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : - toImageSheaf f ≫ imageSheafι f = f := by +theorem Sheaf.toImage_ι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : + toImage f ≫ imageι f = f := by ext1 - simp [toImagePresheafSheafify] + simp [Subpresheaf.toRangeSheafify] -instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Mono (imageSheafι f) := +instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Mono (Sheaf.imageι f) := (sheafToPresheaf J _).mono_of_mono_map (by dsimp infer_instance) -instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Epi (toImageSheaf f) := by +instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Epi (Sheaf.toImage f) := by refine ⟨@fun G' g₁ g₂ e => ?_⟩ ext U ⟨s, hx⟩ apply ((isSheaf_iff_isSheaf_of_type J _).mp G'.2 _ hx).isSeparatedFor.ext rintro V i ⟨y, e'⟩ change (g₁.val.app _ ≫ G'.val.map _) _ = (g₂.val.app _ ≫ G'.val.map _) _ rw [← NatTrans.naturality, ← NatTrans.naturality] - have E : (toImageSheaf f).val.app (op V) y = (imageSheaf f).val.map i.op ⟨s, hx⟩ := + have E : (Sheaf.toImage f).val.app (op V) y = (Sheaf.image f).val.map i.op ⟨s, hx⟩ := Subtype.ext e' have := congr_arg (fun f : F ⟶ G' => (Sheaf.Hom.val f).app _ y) e dsimp at this ⊢ @@ -245,9 +248,9 @@ instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Epi (toImageSheaf f) := by /-- The mono factorization given by `image_sheaf` for a morphism. -/ def imageMonoFactorization {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Limits.MonoFactorisation f where - I := imageSheaf f - m := imageSheafι f - e := toImageSheaf f + I := Sheaf.image f + m := Sheaf.imageι f + e := Sheaf.toImage f /-- The mono factorization given by `image_sheaf` for a morphism is an image. -/ noncomputable def imageFactorization {F F' : Sheaf J (Type (max v u))} (f : F ⟶ F') : @@ -256,14 +259,13 @@ noncomputable def imageFactorization {F F' : Sheaf J (Type (max v u))} (f : F isImage := { lift := fun I => by haveI M := (Sheaf.Hom.mono_iff_presheaf_mono J (Type (max v u)) _).mp I.m_mono - haveI := isIso_toImagePresheaf I.m.1 - refine ⟨Subpresheaf.homOfLe ?_ ≫ inv (toImagePresheaf I.m.1)⟩ + refine ⟨Subpresheaf.homOfLe ?_ ≫ inv (Subpresheaf.toRange I.m.1)⟩ apply Subpresheaf.sheafify_le · conv_lhs => rw [← I.fac] - apply imagePresheaf_comp_le + apply Subpresheaf.range_comp_le · rw [← isSheaf_iff_isSheaf_of_type] exact F'.2 - · apply Presieve.isSheaf_iso J (asIso <| toImagePresheaf I.m.1) + · apply Presieve.isSheaf_iso J (asIso <| Subpresheaf.toRange I.m.1) rw [← isSheaf_iff_isSheaf_of_type] exact I.I.2 lift_fac := fun I => by @@ -272,10 +274,17 @@ noncomputable def imageFactorization {F F' : Sheaf J (Type (max v u))} (f : F generalize_proofs h rw [← Subpresheaf.homOfLe_ι h, Category.assoc] congr 1 - rw [IsIso.inv_comp_eq, toImagePresheaf_ι] } + rw [IsIso.inv_comp_eq, Subpresheaf.toRange_ι] } instance : Limits.HasImages (Sheaf J (Type max v u)) := - ⟨@fun _ _ f => ⟨⟨imageFactorization f⟩⟩⟩ + ⟨fun f => ⟨⟨imageFactorization f⟩⟩⟩ + +@[deprecated (since := "2025-01-25")] alias toImagePresheafSheafify := + Subpresheaf.toRangeSheafify +@[deprecated (since := "2025-01-25")] alias imageSheaf := Sheaf.image +@[deprecated (since := "2025-01-25")] alias toImageSheaf := Sheaf.toImage +@[deprecated (since := "2025-01-25")] alias imageSheafι := Sheaf.imageι +@[deprecated (since := "2025-01-25")] alias toImageSheaf_ι := Sheaf.toImage_ι end Image diff --git a/Mathlib/CategoryTheory/SmallObject/TransfiniteCompositionLifting.lean b/Mathlib/CategoryTheory/SmallObject/TransfiniteCompositionLifting.lean index 7b7f5cdd144f9..715f2e390e058 100644 --- a/Mathlib/CategoryTheory/SmallObject/TransfiniteCompositionLifting.lean +++ b/Mathlib/CategoryTheory/SmallObject/TransfiniteCompositionLifting.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.CategoryTheory.SmallObject.WellOrderInductionData -import Mathlib.CategoryTheory.MorphismProperty.TransfiniteComposition -import Mathlib.CategoryTheory.LiftingProperties.Basic +import Mathlib.CategoryTheory.Limits.Shapes.Preorder.WellOrderContinuous /-! # The left lifting property is stable under transfinite composition @@ -56,8 +55,7 @@ variable {C : Type u} [Category.{v} C] namespace HasLiftingProperty -variable {W : MorphismProperty C} - {J : Type w} [LinearOrder J] [OrderBot J] +variable {J : Type w} [LinearOrder J] [OrderBot J] namespace transfiniteComposition diff --git a/Mathlib/CategoryTheory/Subobject/Comma.lean b/Mathlib/CategoryTheory/Subobject/Comma.lean index 31c43bb091ef1..0c552180fb4a3 100644 --- a/Mathlib/CategoryTheory/Subobject/Comma.lean +++ b/Mathlib/CategoryTheory/Subobject/Comma.lean @@ -93,7 +93,6 @@ theorem lift_projectSubobject [HasFiniteLimits C] [PreservesFiniteLimits T] /-- If `A : S → T.obj B` is a structured arrow for `S : D` and `T : C ⥤ D`, then we can explicitly describe the subobjects of `A` as the subobjects `P` of `B` in `C` for which `A.hom` factors through the image of `P` under `T`. -/ -@[simps!] def subobjectEquiv [HasFiniteLimits C] [PreservesFiniteLimits T] (A : StructuredArrow S T) : Subobject A ≃o { P : Subobject A.right // ∃ q, q ≫ T.map P.arrow = A.hom } where toFun P := ⟨projectSubobject P, projectSubobject_factors P⟩ @@ -111,10 +110,6 @@ def subobjectEquiv [HasFiniteLimits C] [PreservesFiniteLimits T] (A : Structured · refine Subobject.mk_le_mk_of_comm (Subobject.ofMkLEMk _ _ h).right ?_ exact congr_arg CommaMorphism.right (Subobject.ofMkLEMk_comp h) --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] CategoryTheory.StructuredArrow.subobjectEquiv_symm_apply - CategoryTheory.StructuredArrow.subobjectEquiv_apply_coe - /-- If `C` is well-powered and complete and `T` preserves limits, then `StructuredArrow S T` is well-powered. -/ instance wellPowered_structuredArrow [LocallySmall.{w} C] diff --git a/Mathlib/CategoryTheory/Subpresheaf/Basic.lean b/Mathlib/CategoryTheory/Subpresheaf/Basic.lean index e272dd86854b4..2a751fadcfb62 100644 --- a/Mathlib/CategoryTheory/Subpresheaf/Basic.lean +++ b/Mathlib/CategoryTheory/Subpresheaf/Basic.lean @@ -4,10 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.CategoryTheory.Elementwise -import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono -import Mathlib.Tactic.CategoryTheory.Elementwise -import Mathlib.CategoryTheory.Sites.IsSheafFor - +import Mathlib.Data.Set.Basic /-! @@ -49,6 +46,8 @@ variable {F F' F'' : Cᵒᵖ ⥤ Type w} (G G' : Subpresheaf F) instance : PartialOrder (Subpresheaf F) := PartialOrder.lift Subpresheaf.obj (fun _ _ => Subpresheaf.ext) +lemma Subpresheaf.le_def (S T : Subpresheaf F) : S ≤ T ↔ ∀ U, S.obj U ≤ T.obj U := Iff.rfl + instance : Top (Subpresheaf F) := ⟨⟨fun _ => ⊤, @fun U _ _ x _ => by simp⟩⟩ @@ -115,94 +114,12 @@ theorem Subpresheaf.eq_top_iff_isIso : G = ⊤ ↔ IsIso G.ι := by rw [← IsIso.inv_hom_id_apply (G.ι.app U) x] exact ((inv (G.ι.app U)) x).2 -/-- If the image of a morphism falls in a subpresheaf, then the morphism factors through it. -/ -@[simps!] -def Subpresheaf.lift (f : F' ⟶ F) (hf : ∀ U x, f.app U x ∈ G.obj U) : F' ⟶ G.toPresheaf where - app U x := ⟨f.app U x, hf U x⟩ - naturality := by - have := elementwise_of% f.naturality - intros - refine funext fun x => Subtype.ext ?_ - simp only [toPresheaf_obj, types_comp_apply] - exact this _ _ - -@[reassoc (attr := simp)] -theorem Subpresheaf.lift_ι (f : F' ⟶ F) (hf : ∀ U x, f.app U x ∈ G.obj U) : - G.lift f hf ≫ G.ι = f := by - ext - rfl - -/-- Given a subpresheaf `G` of `F`, an `F`-section `s` on `U`, we may define a sieve of `U` -consisting of all `f : V ⟶ U` such that the restriction of `s` along `f` is in `G`. -/ -@[simps] -def Subpresheaf.sieveOfSection {U : Cᵒᵖ} (s : F.obj U) : Sieve (unop U) where - arrows V f := F.map f.op s ∈ G.obj (op V) - downward_closed := @fun V W i hi j => by - simp only [op_unop, op_comp, FunctorToTypes.map_comp_apply] - exact G.map _ hi - -/-- Given an `F`-section `s` on `U` and a subpresheaf `G`, we may define a family of elements in -`G` consisting of the restrictions of `s` -/ -def Subpresheaf.familyOfElementsOfSection {U : Cᵒᵖ} (s : F.obj U) : - (G.sieveOfSection s).1.FamilyOfElements G.toPresheaf := fun _ i hi => ⟨F.map i.op s, hi⟩ - -theorem Subpresheaf.family_of_elements_compatible {U : Cᵒᵖ} (s : F.obj U) : - (G.familyOfElementsOfSection s).Compatible := by - intro Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ e - refine Subtype.ext ?_ -- Porting note: `ext1` does not work here - change F.map g₁.op (F.map f₁.op s) = F.map g₂.op (F.map f₂.op s) - rw [← FunctorToTypes.map_comp_apply, ← FunctorToTypes.map_comp_apply, ← op_comp, ← op_comp, e] - theorem Subpresheaf.nat_trans_naturality (f : F' ⟶ G.toPresheaf) {U V : Cᵒᵖ} (i : U ⟶ V) (x : F'.obj U) : (f.app V (F'.map i x)).1 = F.map i (f.app U x).1 := congr_arg Subtype.val (FunctorToTypes.naturality _ _ f i x) -section Image - -/-- The image presheaf of a morphism, whose components are the set-theoretic images. -/ -@[simps] -def imagePresheaf (f : F' ⟶ F) : Subpresheaf F where - obj U := Set.range (f.app U) - map := by - rintro U V i _ ⟨x, rfl⟩ - have := elementwise_of% f.naturality - exact ⟨_, this i x⟩ - @[simp] theorem top_subpresheaf_obj (U) : (⊤ : Subpresheaf F).obj U = ⊤ := rfl -@[simp] -theorem imagePresheaf_id : imagePresheaf (𝟙 F) = ⊤ := by - ext - simp - -/-- A morphism factors through the image presheaf. -/ -@[simps!] -def toImagePresheaf (f : F' ⟶ F) : F' ⟶ (imagePresheaf f).toPresheaf := - (imagePresheaf f).lift f fun _ _ => Set.mem_range_self _ - -@[reassoc (attr := simp)] -theorem toImagePresheaf_ι (f : F' ⟶ F) : toImagePresheaf f ≫ (imagePresheaf f).ι = f := - (imagePresheaf f).lift_ι _ _ - -theorem imagePresheaf_comp_le (f₁ : F ⟶ F') (f₂ : F' ⟶ F'') : - imagePresheaf (f₁ ≫ f₂) ≤ imagePresheaf f₂ := fun U _ hx => ⟨f₁.app U hx.choose, hx.choose_spec⟩ - -instance isIso_toImagePresheaf {F F' : Cᵒᵖ ⥤ Type w} (f : F ⟶ F') [hf : Mono f] : - IsIso (toImagePresheaf f) := by - have : ∀ (X : Cᵒᵖ), IsIso ((toImagePresheaf f).app X) := by - intro X - rw [isIso_iff_bijective] - constructor - · intro x y e - have := (NatTrans.mono_iff_mono_app f).mp hf X - rw [mono_iff_injective] at this - exact this (congr_arg Subtype.val e :) - · rintro ⟨_, ⟨x, rfl⟩⟩ - exact ⟨x, rfl⟩ - apply NatIso.isIso_of_isIso_app - -end Image - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Subpresheaf/Image.lean b/Mathlib/CategoryTheory/Subpresheaf/Image.lean new file mode 100644 index 0000000000000..4247b888fb3af --- /dev/null +++ b/Mathlib/CategoryTheory/Subpresheaf/Image.lean @@ -0,0 +1,190 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang, Joël Riou +-/ +import Mathlib.CategoryTheory.Subpresheaf.Basic +import Mathlib.CategoryTheory.Limits.FunctorCategory.EpiMono + +/-! +# The image of a subpresheaf + +Given a morphism of presheaves of types `p : F' ⟶ F`, we define its range +`Subpresheaf.range p`. More generally, if `G' : Subpresheaf F'`, we +define `G'.image p : Subpresheaf F` as the image of `G'` by `f`, and +if `G : Subpresheaf F`, we define its preimage `G.preimage f : Subpresheaf F'`. + +-/ + +universe w v u + +namespace CategoryTheory + +variable {C : Type u} [Category.{v} C] {F F' F'' : Cᵒᵖ ⥤ Type w} + +namespace Subpresheaf + +section range + +/-- The range of a morphism of presheaves of types, as a subpresheaf of the target. -/ +@[simps] +def range (p : F' ⟶ F) : Subpresheaf F where + obj U := Set.range (p.app U) + map := by + rintro U V i _ ⟨x, rfl⟩ + exact ⟨_, FunctorToTypes.naturality _ _ p i x⟩ + +variable (F) in +lemma range_id : range (𝟙 F) = ⊤ := by aesop + +@[simp] +lemma range_ι (G : Subpresheaf F) : range G.ι = G := by aesop + +end range + +section lift + +variable (f : F' ⟶ F) {G : Subpresheaf F} (hf : range f ≤ G) + +/-- If the image of a morphism falls in a subpresheaf, then the morphism factors through it. -/ +@[simps!] +def lift : F' ⟶ G.toPresheaf where + app U x := ⟨f.app U x, hf U (by simp)⟩ + naturality _ _ g := by + ext x + simpa [Subtype.ext_iff] using FunctorToTypes.naturality _ _ f g x + +@[reassoc (attr := simp)] +theorem lift_ι : lift f hf ≫ G.ι = f := rfl + +end lift + +section range + +variable (p : F' ⟶ F) + +/-- Given a morphism `p : F' ⟶ F` of presheaves of types, this is the morphism +from `F'` to its range. -/ +def toRange : + F' ⟶ (range p).toPresheaf := + lift p (by rfl) + +@[reassoc (attr := simp)] +lemma toRange_ι : toRange p ≫ (range p).ι = p := rfl + +lemma toRange_app_val {i : Cᵒᵖ} (x : F'.obj i) : + ((toRange p).app i x).val = p.app i x := by + simp [toRange] + +@[simp] +lemma range_toRange : range (toRange p) = ⊤ := by + ext i ⟨x, hx⟩ + dsimp at hx ⊢ + simp only [Set.mem_range, Set.mem_univ, iff_true] + simp only [Set.mem_range] at hx + obtain ⟨y, rfl⟩ := hx + exact ⟨y, rfl⟩ + +lemma epi_iff_range_eq_top : + Epi p ↔ range p = ⊤ := by + simp [NatTrans.epi_iff_epi_app, epi_iff_surjective, Subpresheaf.ext_iff, funext_iff, + Set.range_eq_univ] + +instance : Epi (toRange p) := by simp [epi_iff_range_eq_top] + +instance [Mono p] : IsIso (toRange p) := by + have := mono_of_mono_fac (toRange_ι p) + rw [NatTrans.isIso_iff_isIso_app] + intro i + rw [isIso_iff_bijective] + constructor + · rw [← mono_iff_injective] + infer_instance + · rw [← epi_iff_surjective] + infer_instance + +lemma range_comp_le (f : F ⟶ F') (g : F' ⟶ F'') : + range (f ≫ g) ≤ range g := fun _ _ _ ↦ by aesop + +end range + +section image + +variable (G : Subpresheaf F) (f : F ⟶ F') + +/-- The image of a subpresheaf by a morphism of presheaves of types. -/ +@[simps] +def image : Subpresheaf F' where + obj i := (f.app i) '' (G.obj i) + map := by + rintro Δ Δ' φ _ ⟨x, hx, rfl⟩ + exact ⟨F.map φ x, G.map φ hx, by apply FunctorToTypes.naturality⟩ + +lemma image_top : (⊤ : Subpresheaf F).image f = range f := by aesop + +lemma image_comp (g : F' ⟶ F'') : + G.image (f ≫ g) = (G.image f).image g := by aesop + +lemma range_comp (g : F' ⟶ F'') : + range (f ≫ g) = (range f).image g := by aesop + +end image + +section preimage + +/-- The preimage of a subpresheaf by a morphism of presheaves of types. -/ +@[simps] +def preimage (G : Subpresheaf F) (p : F' ⟶ F) : Subpresheaf F' where + obj n := p.app n ⁻¹' (G.obj n) + map f := (Set.preimage_mono (G.map f)).trans (by + simp only [Set.preimage_preimage, FunctorToTypes.naturality _ _ p f] + rfl) + +@[simp] +lemma preimage_id (G : Subpresheaf F) : + G.preimage (𝟙 F) = G := by aesop + +lemma preimage_comp (G : Subpresheaf F) (f : F'' ⟶ F') (g : F' ⟶ F) : + G.preimage (f ≫ g) = (G.preimage g).preimage f := by aesop + +lemma image_le_iff (G : Subpresheaf F) (f : F ⟶ F') (G' : Subpresheaf F') : + G.image f ≤ G' ↔ G ≤ G'.preimage f := by + simp [Subpresheaf.le_def] + +/-- Given a morphism `p : F' ⟶ F` of presheaves of types and `G : Subpresheaf F`, +this is the morphism from the preimage of `G` by `p` to `G`. -/ +def fromPreimage (G : Subpresheaf F) (p : F' ⟶ F) : + (G.preimage p).toPresheaf ⟶ G.toPresheaf := + lift ((G.preimage p).ι ≫ p) (by + rw [range_comp, range_ι, image_le_iff]) + +@[reassoc] +lemma fromPreimage_ι (G : Subpresheaf F) (p : F' ⟶ F) : + G.fromPreimage p ≫ G.ι = (G.preimage p).ι ≫ p := rfl + +lemma preimage_eq_top_iff (G : Subpresheaf F) (p : F' ⟶ F) : + G.preimage p = ⊤ ↔ range p ≤ G := by + rw [← image_top, image_le_iff] + aesop + +@[simp] +lemma preimage_image_of_epi (G : Subpresheaf F) (p : F' ⟶ F) [hp : Epi p] : + (G.preimage p).image p = G := by + apply le_antisymm + · rw [image_le_iff] + · intro i x hx + simp only [NatTrans.epi_iff_epi_app, epi_iff_surjective] at hp + obtain ⟨y, rfl⟩ := hp _ x + exact ⟨y, hx, rfl⟩ + +end preimage + +end Subpresheaf + +@[deprecated (since := "2025-01-25")] alias imagePresheaf := Subpresheaf.range +@[deprecated (since := "2025-01-25")] alias imagePresheaf_id := Subpresheaf.range_id +@[deprecated (since := "2025-01-25")] alias toImagePresheaf := Subpresheaf.toRange +@[deprecated (since := "2025-01-25")] alias toImagePresheaf_ι := Subpresheaf.toRange_ι +@[deprecated (since := "2025-01-25")] alias imagePresheaf_comp_le := Subpresheaf.range_comp_le + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean b/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean new file mode 100644 index 0000000000000..7c84f551edcd1 --- /dev/null +++ b/Mathlib/CategoryTheory/Subpresheaf/Sieves.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.CategoryTheory.Subpresheaf.Basic +import Mathlib.CategoryTheory.Sites.IsSheafFor + +/-! +# Sieves attached to subpresheaves + +Given a subpresheaf `G` of a presheaf of types `F : Cᵒᵖ ⥤ Type w` and +a section `s : F.obj U`, we define a sieve `G.sieveOfSection s : Sieve (unop U)` +and the associated compatible family of elements with values in `G.toPresheaf`. + +-/ + +universe w v u + +namespace CategoryTheory.Subpresheaf + +open Opposite + +variable {C : Type u} [Category.{v} C] {F : Cᵒᵖ ⥤ Type w} (G : Subpresheaf F) + +/-- Given a subpresheaf `G` of `F`, an `F`-section `s` on `U`, we may define a sieve of `U` +consisting of all `f : V ⟶ U` such that the restriction of `s` along `f` is in `G`. -/ +@[simps] +def sieveOfSection {U : Cᵒᵖ} (s : F.obj U) : Sieve (unop U) where + arrows V f := F.map f.op s ∈ G.obj (op V) + downward_closed := @fun V W i hi j => by + simp only [op_unop, op_comp, FunctorToTypes.map_comp_apply] + exact G.map _ hi + +/-- Given an `F`-section `s` on `U` and a subpresheaf `G`, we may define a family of elements in +`G` consisting of the restrictions of `s` -/ +def familyOfElementsOfSection {U : Cᵒᵖ} (s : F.obj U) : + (G.sieveOfSection s).1.FamilyOfElements G.toPresheaf := fun _ i hi => ⟨F.map i.op s, hi⟩ + +theorem family_of_elements_compatible {U : Cᵒᵖ} (s : F.obj U) : + (G.familyOfElementsOfSection s).Compatible := by + intro Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ e + refine Subtype.ext ?_ -- Porting note: `ext1` does not work here + change F.map g₁.op (F.map f₁.op s) = F.map g₂.op (F.map f₂.op s) + rw [← FunctorToTypes.map_comp_apply, ← FunctorToTypes.map_comp_apply, ← op_comp, ← op_comp, e] + +end CategoryTheory.Subpresheaf diff --git a/Mathlib/CategoryTheory/Triangulated/Basic.lean b/Mathlib/CategoryTheory/Triangulated/Basic.lean index b8e3b05949917..588667615973e 100644 --- a/Mathlib/CategoryTheory/Triangulated/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/Basic.lean @@ -34,9 +34,8 @@ We work in a category `C` equipped with a shift. variable (C : Type u) [Category.{v} C] [HasShift C ℤ] /-- A triangle in `C` is a sextuple `(X,Y,Z,f,g,h)` where `X,Y,Z` are objects of `C`, -and `f : X ⟶ Y`, `g : Y ⟶ Z`, `h : Z ⟶ X⟦1⟧` are morphisms in `C`. -See . --/ +and `f : X ⟶ Y`, `g : Y ⟶ Z`, `h : Z ⟶ X⟦1⟧` are morphisms in `C`. -/ +@[stacks 0144] structure Triangle where mk' :: /-- the first object of a triangle -/ obj₁ : C @@ -95,9 +94,8 @@ In other words, we have a commutative diagram: X' ───> Y' ───> Z' ───> X'⟦1⟧ f' g' h' ``` -See . -/ -@[ext] +@[ext, stacks 0144] structure TriangleMorphism (T₁ : Triangle C) (T₂ : Triangle C) where /-- the first morphism in a triangle morphism -/ hom₁ : T₁.obj₁ ⟶ T₂.obj₁ diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean b/Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean index 94ae9013b1c47..f336e41e98c31 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite/Basic.lean @@ -99,9 +99,6 @@ lemma shiftFunctorAdd'_op_hom_app (X : Cᵒᵖ) (a₁ a₂ a₃ : ℤ) (h : a₁ erw [@pullbackShiftFunctorAdd'_hom_app (OppositeShift C ℤ) _ _ _ _ _ _ _ X a₁ a₂ a₃ h b₁ b₂ b₃ (by dsimp; omega) (by dsimp; omega) (by dsimp; omega)] rw [oppositeShiftFunctorAdd'_hom_app] - obtain rfl : b₁ = -a₁ := by omega - obtain rfl : b₂ = -a₂ := by omega - obtain rfl : b₃ = -a₃ := by omega rfl lemma shiftFunctorAdd'_op_inv_app (X : Cᵒᵖ) (a₁ a₂ a₃ : ℤ) (h : a₁ + a₂ = a₃) diff --git a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean index bdaf65cf13c90..553fd66efc6ca 100644 --- a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean @@ -56,9 +56,8 @@ relative to that shift is called pretriangulated if the following hold: ``` where the left square commutes, and whose rows are distinguished triangles, there exists a morphism `c : Z ⟶ Z'` such that `(a,b,c)` is a triangle morphism. - -See -/ +@[stacks 0145] class Pretriangulated [∀ n : ℤ, Functor.Additive (shiftFunctor C n)] where /-- a class of triangle which are called `distinguished` -/ distinguishedTriangles : Set (Triangle C) @@ -115,10 +114,8 @@ theorem inv_rot_of_distTriang (T : Triangle C) (H : T ∈ distTriang C) : f g h X ───> Y ───> Z ───> X⟦1⟧ ``` -the composition `f ≫ g = 0`. -See --/ -@[reassoc] +the composition `f ≫ g = 0`. -/ +@[reassoc, stacks 0146] theorem comp_distTriang_mor_zero₁₂ (T) (H : T ∈ (distTriang C)) : T.mor₁ ≫ T.mor₂ = 0 := by obtain ⟨c, hc⟩ := complete_distinguished_triangle_morphism _ _ (contractible_distinguished T.obj₁) H (𝟙 T.obj₁) @@ -130,10 +127,8 @@ theorem comp_distTriang_mor_zero₁₂ (T) (H : T ∈ (distTriang C)) : T.mor₁ f g h X ───> Y ───> Z ───> X⟦1⟧ ``` -the composition `g ≫ h = 0`. -See --/ -@[reassoc] +the composition `g ≫ h = 0`. -/ +@[reassoc, stacks 0146] theorem comp_distTriang_mor_zero₂₃ (T : Triangle C) (H : T ∈ distTriang C) : T.mor₂ ≫ T.mor₃ = 0 := comp_distTriang_mor_zero₁₂ T.rotate (rot_of_distTriang T H) @@ -143,10 +138,8 @@ theorem comp_distTriang_mor_zero₂₃ (T : Triangle C) (H : T ∈ distTriang C) f g h X ───> Y ───> Z ───> X⟦1⟧ ``` -the composition `h ≫ f⟦1⟧ = 0`. -See --/ -@[reassoc] +the composition `h ≫ f⟦1⟧ = 0`. -/ +@[reassoc, stacks 0146] theorem comp_distTriang_mor_zero₃₁ (T : Triangle C) (H : T ∈ distTriang C) : T.mor₃ ≫ T.mor₁⟦1⟧' = 0 := by have H₂ := rot_of_distTriang T.rotate (rot_of_distTriang T H) diff --git a/Mathlib/CategoryTheory/Triangulated/Triangulated.lean b/Mathlib/CategoryTheory/Triangulated/Triangulated.lean index d626b0eb72e4d..bb65bf364d393 100644 --- a/Mathlib/CategoryTheory/Triangulated/Triangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Triangulated.lean @@ -32,8 +32,8 @@ variable {C} -- Porting note: see https://github.com/leanprover/lean4/issues/2188 set_option genInjectivity false in -/-- An octahedron is a type of datum whose existence is asserted by -the octahedron axiom (TR 4), see https://stacks.math.columbia.edu/tag/05QK -/ +/-- An octahedron is a type of datum whose existence is asserted by the octahedron axiom (TR 4). -/ +@[stacks 05QK] structure Octahedron {X₁ X₂ X₃ Z₁₂ Z₂₃ Z₁₃ : C} {u₁₂ : X₁ ⟶ X₂} {u₂₃ : X₂ ⟶ X₃} {u₁₃ : X₁ ⟶ X₃} (comm : u₁₂ ≫ u₂₃ = u₁₃) @@ -160,7 +160,8 @@ end Triangulated open Triangulated /-- A triangulated category is a pretriangulated category which satisfies -the octahedron axiom (TR 4), see https://stacks.math.columbia.edu/tag/05QK -/ +the octahedron axiom (TR 4). -/ +@[stacks 05QK] class IsTriangulated : Prop where /-- the octahedron axiom (TR 4) -/ octahedron_axiom : diff --git a/Mathlib/CategoryTheory/Types.lean b/Mathlib/CategoryTheory/Types.lean index 12a318305da27..9eeb5f3d051ae 100644 --- a/Mathlib/CategoryTheory/Types.lean +++ b/Mathlib/CategoryTheory/Types.lean @@ -224,10 +224,8 @@ def homOfElement {X : Type u} (x : X) : PUnit ⟶ X := fun _ => x theorem homOfElement_eq_iff {X : Type u} (x y : X) : homOfElement x = homOfElement y ↔ x = y := ⟨fun H => congr_fun H PUnit.unit, by aesop⟩ -/-- A morphism in `Type` is a monomorphism if and only if it is injective. - -See . --/ +/-- A morphism in `Type` is a monomorphism if and only if it is injective. -/ +@[stacks 003C] theorem mono_iff_injective {X Y : Type u} (f : X ⟶ Y) : Mono f ↔ Function.Injective f := by constructor · intro H x x' h @@ -238,10 +236,8 @@ theorem mono_iff_injective {X Y : Type u} (f : X ⟶ Y) : Mono f ↔ Function.In theorem injective_of_mono {X Y : Type u} (f : X ⟶ Y) [hf : Mono f] : Function.Injective f := (mono_iff_injective f).1 hf -/-- A morphism in `Type` is an epimorphism if and only if it is surjective. - -See . --/ +/-- A morphism in `Type` is an epimorphism if and only if it is surjective. -/ +@[stacks 003C] theorem epi_iff_surjective {X Y : Type u} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := by constructor · rintro ⟨H⟩ diff --git a/Mathlib/CategoryTheory/Widesubcategory.lean b/Mathlib/CategoryTheory/Widesubcategory.lean index 8c70e208ae0c9..4a9779b92eb3b 100644 --- a/Mathlib/CategoryTheory/Widesubcategory.lean +++ b/Mathlib/CategoryTheory/Widesubcategory.lean @@ -39,7 +39,6 @@ variable (F : C → D) (P : MorphismProperty D) [P.IsMultiplicative] which provides a category structure so that the morphisms `X ⟶ Y` are the morphisms in `D` from `F X` to `F Y` which satisfy a property `P : MorphismProperty D` that is multiplicative. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] @[nolint unusedArguments] def InducedWideCategory (_F : C → D) (_P : MorphismProperty D) [IsMultiplicative _P] := C diff --git a/Mathlib/CategoryTheory/WithTerminal.lean b/Mathlib/CategoryTheory/WithTerminal.lean index 3fbad067a0b3d..e8309c595d1f7 100644 --- a/Mathlib/CategoryTheory/WithTerminal.lean +++ b/Mathlib/CategoryTheory/WithTerminal.lean @@ -64,7 +64,6 @@ namespace WithTerminal variable {C} /-- Morphisms for `WithTerminal C`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `nolint has_nonempty_instance` @[simp] def Hom : WithTerminal C → WithTerminal C → Type v | of X, of Y => X ⟶ Y @@ -368,7 +367,6 @@ namespace WithInitial variable {C} /-- Morphisms for `WithInitial C`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `nolint has_nonempty_instance` @[simp] def Hom : WithInitial C → WithInitial C → Type v | of X, of Y => X ⟶ Y diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index e362440ade3cf..70282b07cb005 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -29,11 +29,8 @@ universe v v₁ v₂ u₁ u₂ -- morphism levels before object levels. See note [CategoryTheory universes]. variable {C : Type u₁} [Category.{v₁} C] -/-- The Yoneda embedding, as a functor from `C` into presheaves on `C`. - -See . --/ -@[simps] +/-- The Yoneda embedding, as a functor from `C` into presheaves on `C`. -/ +@[simps, stacks 001O] def yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁ where obj X := { obj := fun Y => unop Y ⟶ X @@ -70,17 +67,13 @@ def fullyFaithful : (yoneda (C := C)).FullyFaithful where lemma fullyFaithful_preimage {X Y : C} (f : yoneda.obj X ⟶ yoneda.obj Y) : fullyFaithful.preimage f = f.app (op X) (𝟙 X) := rfl -/-- The Yoneda embedding is full. - -See . --/ +/-- The Yoneda embedding is full. -/ +@[stacks 001P] instance yoneda_full : (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁).Full := fullyFaithful.full -/-- The Yoneda embedding is faithful. - -See . --/ +/-- The Yoneda embedding is faithful. -/ +@[stacks 001P] instance yoneda_faithful : (yoneda : C ⥤ Cᵒᵖ ⥤ Type v₁).Faithful := fullyFaithful.faithful @@ -254,10 +247,8 @@ def CorepresentableBy.toIso {F : C ⥤ Type v₁} {X : C} (e : F.Corepresentable /-- A functor `F : Cᵒᵖ ⥤ Type v` is representable if there is an object `Y` with a structure `F.RepresentableBy Y`, i.e. there is a natural bijection `(X ⟶ Y) ≃ F.obj (op X)`, -which may also be rephrased as a natural isomorphism `yoneda.obj X ≅ F` when `Category.{v} C`. - -See . --/ +which may also be rephrased as a natural isomorphism `yoneda.obj X ≅ F` when `Category.{v} C`. -/ +@[stacks 001Q] class IsRepresentable (F : Cᵒᵖ ⥤ Type v) : Prop where has_representation : ∃ (Y : C), Nonempty (F.RepresentableBy Y) @@ -277,9 +268,8 @@ instance {X : C} : IsRepresentable (yoneda.obj X) := IsRepresentable.mk' (Iso.refl _) /-- A functor `F : C ⥤ Type v₁` is corepresentable if there is object `X` so `F ≅ coyoneda.obj X`. - -See . -/ +@[stacks 001Q] class IsCorepresentable (F : C ⥤ Type v) : Prop where has_corepresentation : ∃ (X : C), Nonempty (F.CorepresentableBy X) @@ -527,10 +517,8 @@ def yonedaCompUliftFunctorEquiv (F : Cᵒᵖ ⥤ Type max v₁ w) (X : C) : /-- The Yoneda lemma asserts that the Yoneda pairing `(X : Cᵒᵖ, F : Cᵒᵖ ⥤ Type) ↦ (yoneda.obj (unop X) ⟶ F)` -is naturally isomorphic to the evaluation `(X, F) ↦ F.obj X`. - -See . --/ +is naturally isomorphic to the evaluation `(X, F) ↦ F.obj X`. -/ +@[stacks 001P] def yonedaLemma : yonedaPairing C ≅ yonedaEvaluation C := NatIso.ofComponents (fun _ ↦ Equiv.toIso (yonedaEquiv.trans Equiv.ulift.symm)) @@ -709,10 +697,8 @@ def coyonedaCompUliftFunctorEquiv (F : C ⥤ Type max v₁ w) (X : Cᵒᵖ) : /-- The Coyoneda lemma asserts that the Coyoneda pairing `(X : C, F : C ⥤ Type) ↦ (coyoneda.obj X ⟶ F)` -is naturally isomorphic to the evaluation `(X, F) ↦ F.obj X`. - -See . --/ +is naturally isomorphic to the evaluation `(X, F) ↦ F.obj X`. -/ +@[stacks 001P] def coyonedaLemma : coyonedaPairing C ≅ coyonedaEvaluation C := NatIso.ofComponents (fun _ ↦ Equiv.toIso (coyonedaEquiv.trans Equiv.ulift.symm)) diff --git a/Mathlib/Combinatorics/Additive/ApproximateSubgroup.lean b/Mathlib/Combinatorics/Additive/ApproximateSubgroup.lean index 208904c043257..9fdc80e1a7677 100644 --- a/Mathlib/Combinatorics/Additive/ApproximateSubgroup.lean +++ b/Mathlib/Combinatorics/Additive/ApproximateSubgroup.lean @@ -126,7 +126,7 @@ lemma of_small_tripling [DecidableEq G] {A : Finset G} (hA₁ : 1 ∈ A) (hAsymm sq_covBySMul := by replace hA := calc (#(A ^ 4 * A) : ℝ) _ = #(A ^ 5) := by rw [← pow_succ] - _ ≤ K ^ 3 * #A := small_pow_of_small_tripling' (by omega) hA hAsymm + _ ≤ K ^ 3 * #A := small_pow_of_small_tripling (by omega) hA hAsymm have hA₀ : A.Nonempty := ⟨1, hA₁⟩ obtain ⟨F, -, hF, hAF⟩ := ruzsa_covering_mul hA₀ hA exact ⟨F, hF, by norm_cast; simpa [div_eq_mul_inv, pow_succ, mul_assoc, hAsymm] using hAF⟩ @@ -197,7 +197,7 @@ lemma isApproximateSubgroup_one {A : Set G} : · simp [hA.nonempty.ne_empty] at hKA · rw [Finset.coe_singleton, singleton_smul, sq] at hKA use x - have hx' : x ⁻¹ • (A * A) ⊆ A := by rwa [← subset_set_smul_iff] + have hx' : x ⁻¹ • (A * A) ⊆ A := by rwa [← subset_smul_set_iff] have hx_inv : x⁻¹ ∈ A := by simpa using hx' (smul_mem_smul_set (mul_mem_mul hA.one_mem hA.one_mem)) have hx_sq : x * x ∈ A := by diff --git a/Mathlib/Combinatorics/Additive/SmallTripling.lean b/Mathlib/Combinatorics/Additive/SmallTripling.lean index 783cfbe6ceb6f..997ee229b26f0 100644 --- a/Mathlib/Combinatorics/Additive/SmallTripling.lean +++ b/Mathlib/Combinatorics/Additive/SmallTripling.lean @@ -130,7 +130,7 @@ terms in the product. When `A` is symmetric (`-A = A`), the base of the exponential can be lowered from `K ^ 3` to `K`, where `K` is the tripling constant. See `Finset.small_nsmul_of_small_tripling`."] -lemma small_alternating_pow_of_small_tripling' (hm : 3 ≤ m) (hA : #(A ^ 3) ≤ K * #A) (ε : Fin m → ℤ) +lemma small_alternating_pow_of_small_tripling (hm : 3 ≤ m) (hA : #(A ^ 3) ≤ K * #A) (ε : Fin m → ℤ) (hε : ∀ i, |ε i| = 1) : #((finRange m).map fun i ↦ A ^ ε i).prod ≤ K ^ (3 * (m - 2)) * #A := by have hm₀ : m ≠ 0 := by positivity @@ -174,7 +174,7 @@ in the sense that `|m • A|` is at most `|A|` times a constant exponential in ` See also `Finset.small_alternating_nsmul_of_small_tripling` for a version with a weaker constant but which encompasses non-symmetric sets."] -lemma small_pow_of_small_tripling' (hm : 3 ≤ m) (hA : #(A ^ 3) ≤ K * #A) (hAsymm : A⁻¹ = A) : +lemma small_pow_of_small_tripling (hm : 3 ≤ m) (hA : #(A ^ 3) ≤ K * #A) (hAsymm : A⁻¹ = A) : #(A ^ m) ≤ K ^ (m - 2) * #A := by have (ε : ℤ) (hε : |ε| = 1) : A ^ ε = A := by obtain rfl | rfl := eq_or_eq_neg_of_abs_eq hε <;> simp [hAsymm] diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index b7b6c98651639..46dc3c0b92821 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -747,9 +747,7 @@ def compositionAsSetEquiv (n : ℕ) : CompositionAsSet n ≃ Finset (Fin (n - 1) exact Nat.succ_pred_eq_of_pos (pos_iff_ne_zero.mpr i_ne_zero) refine ⟨⟨i - 1, ?_⟩, ?_, ?_⟩ · have : (i : ℕ) < n + 1 := i.2 - simp? [Nat.lt_succ_iff_lt_or_eq, i_ne_last] at this says - simp only [Nat.succ_eq_add_one, Nat.lt_succ_iff_lt_or_eq, i_ne_last, or_false] at this - exact Nat.pred_lt_pred i_ne_zero this + omega · convert i_mem simp only rwa [add_comm] diff --git a/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean b/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean index 2ba5461d3a0a6..8156e1b2c8422 100644 --- a/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean +++ b/Mathlib/Combinatorics/Enumerative/IncidenceAlgebra.lean @@ -402,6 +402,10 @@ lemma mu_apply (a b : α) : mu 𝕜 a b = if a = b then 1 else -∑ x ∈ Ico a lemma mu_eq_neg_sum_Ico_of_ne (hab : a ≠ b) : mu 𝕜 a b = -∑ x ∈ Ico a b, mu 𝕜 a x := by rw [mu_apply, if_neg hab] +variable (𝕜 α) +/-- The Euler characteristic of a finite bounded order. -/ +def eulerChar [BoundedOrder α] : 𝕜 := mu 𝕜 (⊥ : α) ⊤ + end Mu section MuSpec @@ -537,6 +541,10 @@ lemma mu_toDual (a b : α) : mu 𝕜 (toDual a) (toDual b) = mu 𝕜 b a := by @[simp] lemma mu_ofDual (a b : αᵒᵈ) : mu 𝕜 (ofDual a) (ofDual b) = mu 𝕜 b a := (mu_toDual ..).symm +@[simp] +lemma eulerChar_orderDual [BoundedOrder α] : eulerChar 𝕜 αᵒᵈ = eulerChar 𝕜 α := by + simp [eulerChar, ← mu_toDual 𝕜 (α := α)] + end OrderDual section InversionTop @@ -657,6 +665,10 @@ lemma mu_prod_mu : (mu 𝕜).prod (mu 𝕜) = (mu 𝕜 : IncidenceAlgebra 𝕜 ( rw [← zeta_prod_zeta, prod_mul_prod', mu_mul_zeta, mu_mul_zeta, one_prod_one] exact fun _ _ _ _ _ _ ↦ Commute.mul_mul_mul_comm (by simp : _ = _) _ _ +@[simp] +lemma eulerChar_prod [BoundedOrder α] [BoundedOrder β] : + eulerChar 𝕜 (α × β) = eulerChar 𝕜 α * eulerChar 𝕜 β := by simp [eulerChar, ← mu_prod_mu] + end PartialOrder end Prod end IncidenceAlgebra diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index ed4441ac1f8be..a9ab68044b012 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -69,8 +69,6 @@ def Hom.opEquiv {V} [Quiver V] {X Y : V} : right_inv _ := rfl /-- A type synonym for a quiver with no arrows. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] def Empty (V : Type u) : Type u := V instance emptyQuiver (V : Type u) : Quiver.{u} (Empty V) := ⟨fun _ _ => PEmpty⟩ diff --git a/Mathlib/Combinatorics/Quiver/Subquiver.lean b/Mathlib/Combinatorics/Quiver/Subquiver.lean index 365c0628c73d1..c200a9d3a9d2a 100644 --- a/Mathlib/Combinatorics/Quiver/Subquiver.lean +++ b/Mathlib/Combinatorics/Quiver/Subquiver.lean @@ -25,7 +25,6 @@ def WideSubquiver (V) [Quiver.{v + 1} V] := /-- A type synonym for `V`, when thought of as a quiver having only the arrows from some `WideSubquiver`. -/ --- Porting note: no hasNonemptyInstance linter yet https://github.com/leanprover-community/mathlib4/issues/5171 @[nolint unusedArguments] def WideSubquiver.toType (V) [Quiver V] (_ : WideSubquiver V) : Type u := V diff --git a/Mathlib/Combinatorics/Quiver/Symmetric.lean b/Mathlib/Combinatorics/Quiver/Symmetric.lean index 39e75756c695f..9609bd98e12aa 100644 --- a/Mathlib/Combinatorics/Quiver/Symmetric.lean +++ b/Mathlib/Combinatorics/Quiver/Symmetric.lean @@ -26,7 +26,6 @@ namespace Quiver /-- A type synonym for the symmetrized quiver (with an arrow both ways for each original arrow). NB: this does not work for `Prop`-valued quivers. It requires `[Quiver.{v+1} V]`. -/ --- Porting note: no hasNonemptyInstance linter yet https://github.com/leanprover-community/mathlib4/issues/5171 def Symmetrify (V : Type*) := V instance symmetrifyQuiver (V : Type u) [Quiver V] : Quiver (Symmetrify V) := @@ -160,7 +159,7 @@ variable {V' : Type*} [Quiver.{v' + 1} V'] def lift [HasReverse V'] (φ : Prefunctor V V') : Prefunctor (Symmetrify V) V' where obj := φ.obj - map f := match f with + map | Sum.inl g => φ.map g | Sum.inr g => reverse (φ.map g) diff --git a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean index 99106b3256d7b..d4c83ed115ba7 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean @@ -133,4 +133,33 @@ theorem cycleGraph_preconnected {n : ℕ} : (cycleGraph n).Preconnected := theorem cycleGraph_connected {n : ℕ} : (cycleGraph (n + 1)).Connected := (pathGraph_connected n).mono pathGraph_le_cycleGraph +private def cycleGraph_EulerianCircuit_cons (n : ℕ) : + ∀ m : Fin (n + 3), (cycleGraph (n + 3)).Walk m 0 + | ⟨0, h⟩ => Walk.nil + | ⟨m + 1, h⟩ => + have hAdj : (cycleGraph (n + 3)).Adj ⟨m + 1, h⟩ ⟨m, Nat.lt_of_succ_lt h⟩ := by + simp [cycleGraph_adj, Fin.ext_iff, Fin.sub_val_of_le] + Walk.cons hAdj (cycleGraph_EulerianCircuit_cons n ⟨m, Nat.lt_of_succ_lt h⟩) + +/-- Eulerian trail of `cycleGraph (n + 3)` -/ +def cycleGraph_EulerianCircuit (n : ℕ) : (cycleGraph (n + 3)).Walk 0 0 := + have hAdj : (cycleGraph (n + 3)).Adj 0 (Fin.last (n + 2)) := by + simp [cycleGraph_adj] + Walk.cons hAdj (cycleGraph_EulerianCircuit_cons n (Fin.last (n + 2))) + +private theorem cycleGraph_EulerianCircuit_cons_length (n : ℕ) : ∀ m : Fin (n + 3), + (cycleGraph_EulerianCircuit_cons n m).length = m.val + | ⟨0, h⟩ => by + unfold cycleGraph_EulerianCircuit_cons + rfl + | ⟨m + 1, h⟩ => by + unfold cycleGraph_EulerianCircuit_cons + simp only [Walk.length_cons] + rw [cycleGraph_EulerianCircuit_cons_length n] + +theorem cycleGraph_EulerianCircuit_length {n : ℕ} : + (cycleGraph_EulerianCircuit n).length = n + 3 := by + unfold cycleGraph_EulerianCircuit + simp [cycleGraph_EulerianCircuit_cons_length] + end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean b/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean index c3b78c334c5d1..b6c01c361ac73 100644 --- a/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean +++ b/Mathlib/Combinatorics/SimpleGraph/ConcreteColorings.lean @@ -3,8 +3,10 @@ Copyright (c) 2023 Iván Renison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Iván Renison -/ +import Mathlib.Combinatorics.SimpleGraph.Circulant import Mathlib.Combinatorics.SimpleGraph.Coloring import Mathlib.Combinatorics.SimpleGraph.Hasse +import Mathlib.Data.Fin.Parity /-! # Concrete colorings of common graphs @@ -83,4 +85,75 @@ theorem Walk.three_le_chromaticNumber_of_odd_loop {α} {G : SimpleGraph α} {u : have : ¬c' u ↔ c' u := (c'.odd_length_iff_not_congr p).mp hOdd simp_all +/-- Bicoloring of a cycle graph of even size -/ +def cycleGraph.bicoloring_of_even (n : ℕ) (h : Even n) : Coloring (cycleGraph n) Bool := + Coloring.mk (fun u ↦ u.val % 2 = 0) <| by + intro u v hAdj + match n with + | 0 => exact u.elim0 + | 1 => simp at h + | n + 2 => + simp only [ne_eq, decide_eq_decide] + simp only [cycleGraph_adj] at hAdj + cases hAdj with + | inl huv | inr huv => + rw [← add_eq_of_eq_sub' huv.symm, ← Fin.even_iff_mod_of_even h, + ← Fin.even_iff_mod_of_even h, Fin.even_add_one_iff_odd] + apply Classical.not_iff.mpr + simp [Fin.not_odd_iff_even_of_even h, Fin.not_even_iff_odd_of_even h] + +theorem chromaticNumber_cycleGraph_of_even (n : ℕ) (h : 2 ≤ n) (hEven : Even n) : + (cycleGraph n).chromaticNumber = 2 := by + have hc := (cycleGraph.bicoloring_of_even n hEven).colorable + apply le_antisymm + · apply hc.chromaticNumber_le + · have hAdj : (cycleGraph n).Adj ⟨0, Nat.zero_lt_of_lt h⟩ ⟨1, h⟩ := by + simp [cycleGraph_adj', Fin.sub_val_of_le] + exact two_le_chromaticNumber_of_adj hAdj + +/-- Tricoloring of a cycle graph -/ +def cycleGraph.tricoloring (n : ℕ) (h : 2 ≤ n) : Coloring (cycleGraph n) + (Fin 3) := Coloring.mk (fun u ↦ if u.val = n - 1 then 2 else ⟨u.val % 2, by fin_omega⟩) <| by + intro u v hAdj + match n with + | 0 => exact u.elim0 + | 1 => simp at h + | n + 2 => + simp only + simp [cycleGraph_adj] at hAdj + split_ifs with hu hv + · simp [Fin.eq_mk_iff_val_eq.mpr hu, Fin.eq_mk_iff_val_eq.mpr hv] at hAdj + · refine (Fin.ne_of_lt (Fin.mk_lt_of_lt_val (?_))).symm + exact v.val.mod_lt Nat.zero_lt_two + · refine (Fin.ne_of_lt (Fin.mk_lt_of_lt_val ?_)) + exact u.val.mod_lt Nat.zero_lt_two + · simp [Fin.ext_iff] + have hu' : u.val + (1 : Fin (n + 2)) < n + 2 := by fin_omega + have hv' : v.val + (1 : Fin (n + 2)) < n + 2 := by fin_omega + cases hAdj with + | inl huv | inr huv => + rw [← add_eq_of_eq_sub' huv.symm] + simp only [Fin.val_add_eq_of_add_lt hv', Fin.val_add_eq_of_add_lt hu', Fin.val_one] + rw [show ∀x y : ℕ, x % 2 = y % 2 ↔ (Even x ↔ Even y) by simp [Nat.even_iff]; omega, + Nat.even_add] + simp only [Nat.not_even_one, iff_false, not_iff_self, iff_not_self] + exact id + +theorem chromaticNumber_cycleGraph_of_odd (n : ℕ) (h : 2 ≤ n) (hOdd : Odd n) : + (cycleGraph n).chromaticNumber = 3 := by + have hc := (cycleGraph.tricoloring n h).colorable + apply le_antisymm + · apply hc.chromaticNumber_le + · have hn3 : n - 3 + 3 = n := by + refine Nat.sub_add_cancel (Nat.succ_le_of_lt (Nat.lt_of_le_of_ne h ?_)) + intro h2 + rw [← h2] at hOdd + exact (Nat.not_odd_iff.mpr rfl) hOdd + let w : (cycleGraph (n - 3 + 3)).Walk 0 0 := cycleGraph_EulerianCircuit (n - 3) + have hOdd' : Odd w.length := by + rw [cycleGraph_EulerianCircuit_length, hn3] + exact hOdd + rw [← hn3] + exact Walk.three_le_chromaticNumber_of_odd_loop w hOdd' + end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Sum.lean b/Mathlib/Combinatorics/SimpleGraph/Sum.lean index fb0baffa19abc..aaf460a2f7c93 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Sum.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Sum.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Iván Renison -/ import Mathlib.Combinatorics.SimpleGraph.Basic +import Mathlib.Combinatorics.SimpleGraph.Coloring import Mathlib.Combinatorics.SimpleGraph.Maps /-! @@ -29,11 +30,11 @@ namespace SimpleGraph /-- Disjoint sum of `G` and `H`. -/ @[simps!] protected def sum (G : SimpleGraph α) (H : SimpleGraph β) : SimpleGraph (α ⊕ β) where - Adj u v := match u, v with + Adj | Sum.inl u, Sum.inl v => G.Adj u v | Sum.inr u, Sum.inr v => H.Adj u v | _, _ => false - symm u v := match u, v with + symm | Sum.inl u, Sum.inl v => G.adj_symm | Sum.inr u, Sum.inr v => H.adj_symm | Sum.inl _, Sum.inr _ | Sum.inr _, Sum.inl _ => id @@ -46,18 +47,12 @@ variable {G : SimpleGraph α} {H : SimpleGraph β} /-- The disjoint sum is commutative up to isomorphism. `Iso.sumComm` as a graph isomorphism. -/ @[simps!] def Iso.sumComm : G ⊕g H ≃g H ⊕g G := ⟨Equiv.sumComm α β, by - intro u v - cases u <;> cases v <;> simp⟩ + rintro (u | u) (v | v) <;> simp⟩ /-- The disjoint sum is associative up to isomorphism. `Iso.sumAssoc` as a graph isomorphism. -/ @[simps!] def Iso.sumAssoc {I : SimpleGraph γ} : (G ⊕g H) ⊕g I ≃g G ⊕g (H ⊕g I) := ⟨Equiv.sumAssoc α β γ, by - intro u v - cases u <;> cases v <;> rename_i u v - · cases u <;> cases v <;> simp - · cases u <;> simp - · cases v <;> simp - · simp⟩ + rintro ((u | u) | u) ((v | v) | v) <;> simp⟩ /-- The embedding of `G` into `G ⊕g H`. -/ @[simps] @@ -73,4 +68,71 @@ def Embedding.sumInr : H ↪g G ⊕g H where inj' u v := by simp map_rel_iff' := by simp +/-- Color `G ⊕g H` with colorings of `G` and `H` -/ +def Coloring.sum (cG : G.Coloring γ) (cH : H.Coloring γ) : (G ⊕g H).Coloring γ where + toFun := Sum.elim cG cH + map_rel' {u v} huv := by cases u <;> cases v <;> simp_all [cG.valid, cH.valid] + +/-- Get coloring of `G` from coloring of `G ⊕g H` -/ +def Coloring.sumLeft (c : (G ⊕g H).Coloring γ) : G.Coloring γ := c.comp Embedding.sumInl.toHom + +/-- Get coloring of `H` from coloring of `G ⊕g H` -/ +def Coloring.sumRight (c : (G ⊕g H).Coloring γ) : H.Coloring γ := c.comp Embedding.sumInr.toHom + +@[simp] +theorem Coloring.sumLeft_sum (cG : G.Coloring γ) (cH : H.Coloring γ) : (cG.sum cH).sumLeft = cG := + rfl + +@[simp] +theorem Coloring.sumRight_sum (cG : G.Coloring γ) (cH : H.Coloring γ) : (cG.sum cH).sumRight = cH := + rfl + +@[simp] +theorem Coloring.sum_sumLeft_sumRight (c : (G ⊕g H).Coloring γ) : c.sumLeft.sum c.sumRight = c := by + ext (u | u) <;> rfl + +/-- Bijection between `(G ⊕g H).Coloring γ` and `G.Coloring γ × H.Coloring γ` -/ +def Coloring.sumEquiv : (G ⊕g H).Coloring γ ≃ G.Coloring γ × H.Coloring γ where + toFun c := ⟨c.sumLeft, c.sumRight⟩ + invFun p := p.1.sum p.2 + left_inv c := by simp [sum_sumLeft_sumRight c] + right_inv p := rfl + +/-- Color `G ⊕g H` with `Fin (n + m)` given a coloring of `G` with `Fin n` and a coloring of `H` +with `Fin m` -/ +def Coloring.sumFin {n m : ℕ} (cG : G.Coloring (Fin n)) (cH : H.Coloring (Fin m)) : + (G ⊕g H).Coloring (Fin (max n m)) := sum + (G.recolorOfEmbedding (Fin.castLEEmb (n.le_max_left m)) cG) + (H.recolorOfEmbedding (Fin.castLEEmb (n.le_max_right m)) cH) + +theorem Colorable.sum_max {n m : ℕ} (hG : G.Colorable n) (hH : H.Colorable m) : + (G ⊕g H).Colorable (max n m) := Nonempty.intro (hG.some.sumFin hH.some) + +theorem Colorable.of_sum_left {n : ℕ} (h : (G ⊕g H).Colorable n) : G.Colorable n := + Nonempty.intro (h.some.sumLeft) + +theorem Colorable.of_sum_right {n : ℕ} (h : (G ⊕g H).Colorable n) : H.Colorable n := + Nonempty.intro (h.some.sumRight) + +@[simp] +theorem colorable_sum {n : ℕ} : (G ⊕g H).Colorable n ↔ G.Colorable n ∧ H.Colorable n := + ⟨fun cGH => ⟨cGH.of_sum_left, cGH.of_sum_right⟩, + fun ⟨cG, cH⟩ => by rw [← n.max_self]; exact cG.sum_max cH⟩ + +theorem chromaticNumber_le_sum_left : G.chromaticNumber ≤ (G ⊕g H).chromaticNumber := + chromaticNumber_le_of_forall_imp (fun _ h ↦ h.of_sum_left) + +theorem chromaticNumber_le_sum_right : H.chromaticNumber ≤ (G ⊕g H).chromaticNumber := + chromaticNumber_le_of_forall_imp (fun _ h ↦ h.of_sum_right) + +@[simp] +theorem chromaticNumber_sum : + (G ⊕g H).chromaticNumber = max G.chromaticNumber H.chromaticNumber := by + refine eq_max chromaticNumber_le_sum_left chromaticNumber_le_sum_right ?_ + rintro (n | n) hG hH + · simp [show (none : ℕ∞) = (⊤ : ℕ∞) from rfl] + · let cG : G.Coloring (Fin n) := (chromaticNumber_le_iff_colorable.mp hG).some + let cH : H.Coloring (Fin n) := (chromaticNumber_le_iff_colorable.mp hH).some + exact chromaticNumber_le_iff_colorable.mpr (Nonempty.intro (cG.sum cH)) + end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index 06aaabee0f997..86e1e71cf4a52 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -1120,15 +1120,10 @@ theorem exists_boundary_dart {u v : V} (p : G.Walk u v) (S : Set V) (uS : u ∈ exact ⟨d, List.Mem.tail _ hd, hcd⟩ · exact ⟨⟨(x, y), a⟩, List.Mem.head _, uS, h⟩ -@[simp] lemma getVert_copy {u v w x : V} (p : G.Walk u v) (i : ℕ) (h : u = w) (h' : v = x) : +@[simp] lemma getVert_copy {u v w x : V} (p : G.Walk u v) (i : ℕ) (h : u = w) (h' : v = x) : (p.copy h h').getVert i = p.getVert i := by subst_vars - match p, i with - | .nil, _ => - rw [getVert_of_length_le _ (by simp only [length_nil, Nat.zero_le] : nil.length ≤ _)] - rw [getVert_of_length_le _ (by simp only [length_copy, length_nil, Nat.zero_le])] - | .cons hadj q, 0 => simp only [copy_rfl_rfl, getVert_zero] - | .cons hadj q, (n + 1) => simp only [copy_cons, getVert_cons_succ]; rfl + rfl @[simp] lemma getVert_tail {u v n} (p : G.Walk u v) : p.tail.getVert n = p.getVert (n + 1) := by diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 92b4e6f1c3a1a..fc5ffcfc01658 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -83,7 +83,6 @@ theorem ack_two (n : ℕ) : ack 2 n = 2 * n + 3 := by · simp · simpa [mul_succ] --- Porting note: re-written to get rid of ack_three_aux @[simp] theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by induction' n with n IH diff --git a/Mathlib/Computability/ContextFreeGrammar.lean b/Mathlib/Computability/ContextFreeGrammar.lean index 62502385e7d1f..aae94be8a4dc5 100644 --- a/Mathlib/Computability/ContextFreeGrammar.lean +++ b/Mathlib/Computability/ContextFreeGrammar.lean @@ -223,10 +223,10 @@ lemma reverse_bijective : Bijective (reverse : ContextFreeRule T N → ContextFr reverse_involutive.bijective lemma reverse_injective : Injective (reverse : ContextFreeRule T N → ContextFreeRule T N) := - reverse_involutive.injective + reverse_bijective.injective lemma reverse_surjective : Surjective (reverse : ContextFreeRule T N → ContextFreeRule T N) := - reverse_involutive.surjective + reverse_bijective.surjective protected lemma Rewrites.reverse : ∀ {u v}, r.Rewrites u v → r.reverse.Rewrites u.reverse v.reverse | _, _, head s => by simpa using .append_left .input_output _ @@ -257,10 +257,10 @@ lemma reverse_bijective : Bijective (reverse : ContextFreeGrammar T → ContextF reverse_involutive.bijective lemma reverse_injective : Injective (reverse : ContextFreeGrammar T → ContextFreeGrammar T) := - reverse_involutive.injective + reverse_bijective.injective lemma reverse_surjective : Surjective (reverse : ContextFreeGrammar T → ContextFreeGrammar T) := - reverse_involutive.surjective + reverse_bijective.surjective lemma produces_reverse : g.reverse.Produces u.reverse v.reverse ↔ g.Produces u v := (Equiv.ofBijective _ ContextFreeRule.reverse_bijective).exists_congr diff --git a/Mathlib/Computability/Encoding.lean b/Mathlib/Computability/Encoding.lean index 67db006d7f41b..2f5c7c2632dda 100644 --- a/Mathlib/Computability/Encoding.lean +++ b/Mathlib/Computability/Encoding.lean @@ -70,11 +70,16 @@ def sectionΓ'Bool : Γ' → Bool | Γ'.bit b => b | _ => Inhabited.default +@[simp] +theorem sectionΓ'Bool_inclusionBoolΓ' {b} : sectionΓ'Bool (inclusionBoolΓ' b) = b := by + cases b <;> rfl + +@[deprecated sectionΓ'Bool_inclusionBoolΓ' (since := "2025-01-21")] theorem leftInverse_section_inclusion : Function.LeftInverse sectionΓ'Bool inclusionBoolΓ' := fun x => Bool.casesOn x rfl rfl theorem inclusionBoolΓ'_injective : Function.Injective inclusionBoolΓ' := - Function.HasLeftInverse.injective (Exists.intro sectionΓ'Bool leftInverse_section_inclusion) + Function.HasLeftInverse.injective ⟨_, (fun _ => sectionΓ'Bool_inclusionBoolΓ')⟩ /-- An encoding function of the positive binary numbers in bool. -/ def encodePosNum : PosNum → List Bool @@ -107,7 +112,7 @@ theorem encodePosNum_nonempty (n : PosNum) : encodePosNum n ≠ [] := PosNum.casesOn n (List.cons_ne_nil _ _) (fun _m => List.cons_ne_nil _ _) fun _m => List.cons_ne_nil _ _ -theorem decode_encodePosNum : ∀ n, decodePosNum (encodePosNum n) = n := by +@[simp] theorem decode_encodePosNum : ∀ n, decodePosNum (encodePosNum n) = n := by intro n induction' n with m hm m hm <;> unfold encodePosNum decodePosNum · rfl @@ -115,7 +120,7 @@ theorem decode_encodePosNum : ∀ n, decodePosNum (encodePosNum n) = n := by exact if_neg (encodePosNum_nonempty m) · exact congr_arg PosNum.bit0 hm -theorem decode_encodeNum : ∀ n, decodeNum (encodeNum n) = n := by +@[simp] theorem decode_encodeNum : ∀ n, decodeNum (encodeNum n) = n := by intro n cases' n with n <;> unfold encodeNum decodeNum · rfl @@ -123,7 +128,7 @@ theorem decode_encodeNum : ∀ n, decodeNum (encodeNum n) = n := by rw [PosNum.cast_to_num] exact if_neg (encodePosNum_nonempty n) -theorem decode_encodeNat : ∀ n, decodeNat (encodeNat n) = n := by +@[simp] theorem decode_encodeNat : ∀ n, decodeNat (encodeNat n) = n := by intro n conv_rhs => rw [← Num.to_of_nat n] exact congr_arg ((↑) : Num → ℕ) (decode_encodeNum n) @@ -144,11 +149,7 @@ def encodingNatΓ' : Encoding ℕ where Γ := Γ' encode x := List.map inclusionBoolΓ' (encodeNat x) decode x := some (decodeNat (List.map sectionΓ'Bool x)) - decode_encode x := - congr_arg _ <| by - -- Porting note: `rw` can't unify `g ∘ f` with `fun x => g (f x)`, used `LeftInverse.id` - -- instead. - rw [List.map_map, leftInverse_section_inclusion.id, List.map_id, decode_encodeNat] + decode_encode x := congr_arg _ <| by simp [Function.comp_def] /-- A binary FinEncoding of ℕ in Γ'. -/ def finEncodingNatΓ' : FinEncoding ℕ := @@ -163,7 +164,7 @@ def unaryEncodeNat : Nat → List Bool def unaryDecodeNat : List Bool → Nat := List.length -theorem unary_decode_encode_nat : ∀ n, unaryDecodeNat (unaryEncodeNat n) = n := fun n => +@[simp] theorem unary_decode_encode_nat : ∀ n, unaryDecodeNat (unaryEncodeNat n) = n := fun n => Nat.rec rfl (fun (_m : ℕ) hm => (congr_arg Nat.succ hm.symm).symm) n /-- A unary fin_encoding of ℕ. -/ @@ -182,7 +183,7 @@ def decodeBool : List Bool → Bool | b :: _ => b | _ => Inhabited.default -theorem decode_encodeBool (b : Bool) : decodeBool (encodeBool b) = b := rfl +@[simp] theorem decode_encodeBool (b : Bool) : decodeBool (encodeBool b) = b := rfl /-- A fin_encoding of bool in bool. -/ def finEncodingBoolBool : FinEncoding Bool where diff --git a/Mathlib/Computability/Language.lean b/Mathlib/Computability/Language.lean index 06251e189c2c0..a8ec5be41e453 100644 --- a/Mathlib/Computability/Language.lean +++ b/Mathlib/Computability/Language.lean @@ -41,9 +41,6 @@ instance instCompleteAtomicBooleanAlgebra : CompleteAtomicBooleanAlgebra (Langua variable {l m : Language α} {a b x : List α} --- Porting note: `reducible` attribute cannot be local. --- attribute [local reducible] Language - /-- Zero language has no elements. -/ instance : Zero (Language α) := ⟨(∅ : Set _)⟩ @@ -82,8 +79,6 @@ instance : KStar (Language α) := ⟨fun l ↦ {x | ∃ L : List (List α), x = lemma kstar_def (l : Language α) : l∗ = {x | ∃ L : List (List α), x = L.flatten ∧ ∀ y ∈ L, y ∈ l} := rfl --- Porting note: `reducible` attribute cannot be local, --- so this new theorem is required in place of `Set.ext`. @[ext] theorem ext {l m : Language α} (h : ∀ (x : List α), x ∈ l ↔ x ∈ m) : l = m := Set.ext h @@ -162,9 +157,7 @@ lemma mem_kstar_iff_exists_nonempty {x : List α} : · rintro ⟨S, rfl, h⟩ refine ⟨S.filter fun l ↦ !List.isEmpty l, by simp [List.flatten_filter_not_isEmpty], fun y hy ↦ ?_⟩ - -- Porting note: The previous code was: - -- rw [mem_filter, empty_iff_eq_nil] at hy - rw [mem_filter, Bool.not_eq_true', ← Bool.bool_iff_false, List.isEmpty_iff] at hy + simp only [mem_filter, Bool.not_eq_eq_eq_not, Bool.not_true, isEmpty_eq_false, ne_eq] at hy exact ⟨h y hy.1, hy.2⟩ · rintro ⟨S, hx, h⟩ exact ⟨S, hx, fun y hy ↦ (h y hy).1⟩ diff --git a/Mathlib/Computability/MyhillNerode.lean b/Mathlib/Computability/MyhillNerode.lean new file mode 100644 index 0000000000000..0ab520eb65238 --- /dev/null +++ b/Mathlib/Computability/MyhillNerode.lean @@ -0,0 +1,103 @@ +/- +Copyright (c) 2024 Google. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Wong +-/ +import Mathlib.Computability.DFA +import Mathlib.Data.Set.Finite.Basic + +/-! +# Myhill–Nerode theorem + +This file proves the Myhill–Nerode theorem using left quotients. + +Given a language `L` and a word `x`, the *left quotient* of `L` by `x` is the set of suffixes `y` +such that `x ++ y` is in `L`. The *Myhill–Nerode theorem* shows that each left quotient, in fact, +corresponds to the state of an automaton that matches `L`, and that `L` is regular if and only if +there are finitely many such states. + +## References + +* +-/ + +universe u v +variable {α : Type u} {σ : Type v} {L : Language α} + +namespace Language + +variable (L) in +/-- The *left quotient* of `x` is the set of suffixes `y` such that `x ++ y` is in `L`. -/ +def leftQuotient (x : List α) : Language α := { y | x ++ y ∈ L } + +variable (L) in +@[simp] +theorem leftQuotient_nil : L.leftQuotient [] = L := rfl + +variable (L) in +theorem leftQuotient_append (x y : List α) : + L.leftQuotient (x ++ y) = (L.leftQuotient x).leftQuotient y := by + simp [leftQuotient, Language] + +@[simp] +theorem mem_leftQuotient (x y : List α) : y ∈ L.leftQuotient x ↔ x ++ y ∈ L := Iff.rfl + +theorem leftQuotient_accepts_apply (M : DFA α σ) (x : List α) : + leftQuotient M.accepts x = M.acceptsFrom (M.eval x) := by + ext y + simp [DFA.mem_accepts, DFA.mem_acceptsFrom, DFA.eval, DFA.evalFrom_of_append] + +theorem leftQuotient_accepts (M : DFA α σ) : leftQuotient M.accepts = M.acceptsFrom ∘ M.eval := + funext <| leftQuotient_accepts_apply M + +theorem IsRegular.finite_range_leftQuotient (h : L.IsRegular) : + (Set.range L.leftQuotient).Finite := by + have ⟨σ, x, M, hM⟩ := h + rw [← hM, leftQuotient_accepts] + exact Set.finite_of_finite_preimage (Set.toFinite _) + (Set.range_comp_subset_range M.eval M.acceptsFrom) + +variable (L) in +/-- The left quotients of a language are the states of an automaton that accepts the language. -/ +def toDFA : DFA α (Set.range L.leftQuotient) where + step s a := by + refine ⟨s.val.leftQuotient [a], ?_⟩ + obtain ⟨y, hy⟩ := s.prop + exists y ++ [a] + rw [← hy, leftQuotient_append] + start := ⟨L, by exists []⟩ + accept := { s | [] ∈ s.val } + +@[simp] +theorem mem_accept_toDFA (s : Set.range L.leftQuotient) : s ∈ L.toDFA.accept ↔ [] ∈ s.val := Iff.rfl + +@[simp] +theorem step_toDFA (s : Set.range L.leftQuotient) (a : α) : + (L.toDFA.step s a).val = s.val.leftQuotient [a] := rfl + +variable (L) in +@[simp] +theorem start_toDFA : L.toDFA.start.val = L := rfl + +variable (L) in +@[simp] +theorem accepts_toDFA : L.toDFA.accepts = L := by + ext x + rw [DFA.mem_accepts] + suffices L.toDFA.eval x = L.leftQuotient x by simp [this] + induction x using List.list_reverse_induction with + | base => simp + | ind x a ih => simp [ih, leftQuotient_append] + +theorem IsRegular.of_finite_range_leftQuotient (h : Set.Finite (Set.range L.leftQuotient)) : + L.IsRegular := + Language.isRegular_iff.mpr ⟨_, h.fintype, L.toDFA, by simp⟩ + +/-- +**Myhill–Nerode theorem**. A language is regular if and only if the set of left quotients is finite. +-/ +theorem isRegular_iff_finite_range_leftQuotient : + L.IsRegular ↔ (Set.range L.leftQuotient).Finite := + ⟨IsRegular.finite_range_leftQuotient, .of_finite_range_leftQuotient⟩ + +end Language diff --git a/Mathlib/Computability/Partrec.lean b/Mathlib/Computability/Partrec.lean index 1068bd7cf6bf7..ebdc557e4cb3a 100644 --- a/Mathlib/Computability/Partrec.lean +++ b/Mathlib/Computability/Partrec.lean @@ -387,7 +387,7 @@ protected theorem bind {f : α →. β} {g : α → β →. σ} (hf : Partrec f) theorem map {f : α →. β} {g : α → β → σ} (hf : Partrec f) (hg : Computable₂ g) : Partrec fun a => (f a).map (g a) := by - simpa [bind_some_eq_map] using @Partrec.bind _ _ _ _ _ _ _ (fun a => Part.some ∘ (g a)) hf hg + simpa [bind_some_eq_map] using Partrec.bind (g := fun a x => some (g a x)) hf hg theorem to₂ {f : α × β →. σ} (hf : Partrec f) : Partrec₂ fun a b => f (a, b) := hf.of_eq fun ⟨_, _⟩ => rfl @@ -656,11 +656,7 @@ variable [Primcodable α] [Primcodable β] [Primcodable γ] [Primcodable σ] open Computable theorem option_some_iff {f : α →. σ} : (Partrec fun a => (f a).map Option.some) ↔ Partrec f := - ⟨fun h => (Nat.Partrec.ppred.comp h).of_eq fun n => by - -- Porting note: needed to help with applying bind_some_eq_map because `Function.comp` got - -- less reducible. - simp [Part.bind_assoc, ← Function.comp_apply (f := Part.some) (g := encode), bind_some_eq_map, - -Function.comp_apply], + ⟨fun h => (Nat.Partrec.ppred.comp h).of_eq fun n => by simp [Part.bind_assoc, bind_some_eq_map], fun hf => hf.map (option_some.comp snd).to₂⟩ theorem option_casesOn_right {o : α → Option β} {f : α → σ} {g : α → β →. σ} (ho : Computable o) diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index f86ef6f073202..4a04dbd781f9d 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -198,16 +198,10 @@ theorem encode_lt_prec (cf cg) : theorem encode_lt_rfind' (cf) : encode cf < encode (rfind' cf) := by simp only [encodeCode_eq, encodeCode] - have := Nat.mul_le_mul_right cf.encodeCode (by decide : 1 ≤ 2 * 2) - rw [one_mul, mul_assoc] at this - refine lt_of_le_of_lt (le_trans this ?_) (lt_add_of_pos_right _ (by decide : 0 < 4)) - exact le_of_lt (Nat.lt_succ_of_le <| Nat.mul_le_mul_left _ <| le_of_lt <| - Nat.lt_succ_of_le <| Nat.mul_le_mul_left _ <| le_rfl) + omega end Nat.Partrec.Code --- Porting note: Opening `Primrec` inside `namespace Nat.Partrec.Code` causes it to resolve --- to `Nat.Partrec`. Needs `open _root_.Partrec` support section open Primrec namespace Nat.Partrec.Code @@ -557,7 +551,6 @@ theorem exists_code {f : ℕ →. ℕ} : Nat.Partrec f ↔ ∃ c : Code, eval c | prec cf cg pf pg => exact pf.prec pg | rfind' cf pf => exact pf.rfind' --- Porting note: `>>`s in `evaln` are now `>>=` because `>>`s are not elaborated well in Lean4. /-- A modified evaluation for the code which returns an `Option ℕ` instead of a `Part ℕ`. To avoid undecidability, `evaln` takes a parameter `k` and fails if it encounters a number ≥ k in the course of its execution. Other than this, the semantics are the same as in `Nat.Partrec.Code.eval`. diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 722792093ba58..20b710c3ba844 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -32,26 +32,6 @@ open Denumerable Encodable Function namespace Nat --- Porting note: elim is no longer required because lean 4 is better --- at inferring motive types (I think this is the reason) --- and worst case, we can always explicitly write (motive := fun _ => C) --- without having to then add all the other underscores - --- /-- The non-dependent recursor on naturals. -/ --- def elim {C : Sort*} : C → (ℕ → C → C) → ℕ → C := --- @Nat.rec fun _ => C --- example {C : Sort*} (base : C) (succ : ℕ → C → C) (a : ℕ) : --- a.elim base succ = a.rec base succ := rfl - --- Porting note: cases is no longer required because lean 4 is better --- at inferring motive types (I think this is the reason) - --- /-- Cases on whether the input is 0 or a successor. -/ --- def cases {C : Sort*} (a : C) (f : ℕ → C) : ℕ → C := --- Nat.elim a fun n _ => f n --- example {C : Sort*} (a : C) (f : ℕ → C) (n : ℕ) : --- n.cases a f = n.casesOn a f := rfl - /-- Calls the given function on a pair of entries `n`, encoded via the pairing function. -/ @[simp, reducible] def unpaired {α} (f : ℕ → ℕ → α) (n : ℕ) : α := @@ -611,10 +591,6 @@ theorem _root_.PrimrecPred.or {p q : α → Prop} [DecidablePred p] [DecidablePr (hp : PrimrecPred p) (hq : PrimrecPred q) : PrimrecPred fun a => p a ∨ q a := (Primrec.or.comp hp hq).of_eq fun n => by simp --- Porting note: It is unclear whether we want to boolean versions --- of these lemmas, just the prop versions, or both --- The boolean versions are often actually easier to use --- but did not exist in Lean 3 protected theorem beq [DecidableEq α] : Primrec₂ (@BEq.beq α _) := have : PrimrecRel fun a b : ℕ => a = b := (PrimrecPred.and nat_le nat_le.swap).of_eq fun a => by simp [le_antisymm_iff] @@ -1107,9 +1083,6 @@ instance vector {n} : Primcodable (List.Vector α n) := instance finArrow {n} : Primcodable (Fin n → α) := ofEquiv _ (Equiv.vectorEquivFin _ _).symm --- Porting note: Equiv.arrayEquivFin is not ported yet --- instance array {n} : Primcodable (Array' n α) := --- ofEquiv _ (Equiv.arrayEquivFin _ _) section ULower diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index 8082aa2d00f52..4e10a5bf7e78c 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -30,6 +30,9 @@ universe u variable {α β γ : Type*} +-- Disable generation of unneeded lemmas which the simpNF linter would complain about. +set_option genSizeOfSpec false in +set_option genInjectivity false in /-- This is the definition of regular expressions. The names used here is to mirror the definition of a Kleene algebra (https://en.wikipedia.org/wiki/Kleene_algebra). * `0` (`zero`) matches nothing @@ -47,15 +50,6 @@ inductive RegularExpression (α : Type u) : Type u | comp : RegularExpression α → RegularExpression α → RegularExpression α | star : RegularExpression α → RegularExpression α - --- Porting note: `simpNF` gets grumpy about how the `foo_def`s below can simplify these.. -attribute [nolint simpNF] RegularExpression.zero.sizeOf_spec -attribute [nolint simpNF] RegularExpression.epsilon.sizeOf_spec -attribute [nolint simpNF] RegularExpression.plus.sizeOf_spec -attribute [nolint simpNF] RegularExpression.plus.injEq -attribute [nolint simpNF] RegularExpression.comp.injEq -attribute [nolint simpNF] RegularExpression.comp.sizeOf_spec - namespace RegularExpression variable {a b : α} @@ -79,7 +73,7 @@ instance : Pow (RegularExpression α) ℕ := ⟨fun n r => npowRec r n⟩ -- Porting note: declaration in an imported module ---attribute [match_pattern] Mul.mul +-- attribute [match_pattern] Mul.mul @[simp] theorem zero_def : (zero : RegularExpression α) = 0 := @@ -97,7 +91,7 @@ theorem plus_def (P Q : RegularExpression α) : plus P Q = P + Q := theorem comp_def (P Q : RegularExpression α) : comp P Q = P * Q := rfl --- Porting note: `matches` is reserved, moved to `matches'` +-- This was renamed to `matches'` during the port of Lean 4 as `matches` is a reserved word. #adaptation_note /-- around nightly-2024-02-25, we need to write `comp x y` in the pattern `comp P Q`, instead of `x * y`. -/ /-- `matches' P` provides a language which contains all strings that `P` matches -/ @@ -401,13 +395,8 @@ theorem matches'_map (f : α → β) : | char a => by rw [eq_comm] exact image_singleton - -- Porting note: the following close with last `rw` but not with `simp`? - | R + S => by simp only [matches'_map, map, matches'_add]; rw [map_add] + | R + S => by simp only [matches'_map, map, matches'_add, map_add] | comp R S => by simp [matches'_map] - | star R => by - simp_rw [map, matches', matches'_map] - rw [Language.kstar_eq_iSup_pow, Language.kstar_eq_iSup_pow] - simp_rw [← map_pow] - exact image_iUnion.symm + | star R => by simp [matches'_map] end RegularExpression diff --git a/Mathlib/Computability/TMComputable.lean b/Mathlib/Computability/TMComputable.lean index 773ba2fe6a6f2..025046b60c925 100644 --- a/Mathlib/Computability/TMComputable.lean +++ b/Mathlib/Computability/TMComputable.lean @@ -82,7 +82,6 @@ instance inhabitedσ : Inhabited tm.σ := def Stmt : Type := Turing.TM2.Stmt tm.Γ tm.Λ tm.σ --- Porting note: The `deriving Inhabited` handler couldn't derive this. instance inhabitedStmt : Inhabited (Stmt tm) := inferInstanceAs (Inhabited (Turing.TM2.Stmt tm.Γ tm.Λ tm.σ)) @@ -135,8 +134,7 @@ structure EvalsToInTime {σ : Type*} (f : σ → Option σ) (a : σ) (b : Option steps_le_m : steps ≤ m /-- Reflexivity of `EvalsTo` in 0 steps. -/ --- @[refl] -- Porting note: `@[refl]` attribute only applies to lemmas proving `x ∼ x` in Lean4. -def EvalsTo.refl {σ : Type*} (f : σ → Option σ) (a : σ) : EvalsTo f a a := +def EvalsTo.refl {σ : Type*} (f : σ → Option σ) (a : σ) : EvalsTo f a (some a) := ⟨0, rfl⟩ /-- Transitivity of `EvalsTo` in the sum of the numbers of steps. -/ @@ -146,8 +144,7 @@ def EvalsTo.trans {σ : Type*} (f : σ → Option σ) (a : σ) (b : σ) (c : Opt ⟨h₂.steps + h₁.steps, by rw [Function.iterate_add_apply, h₁.evals_in_steps, h₂.evals_in_steps]⟩ /-- Reflexivity of `EvalsToInTime` in 0 steps. -/ --- @[refl] -- Porting note: `@[refl]` attribute only applies to lemmas proving `x ∼ x` in Lean4. -def EvalsToInTime.refl {σ : Type*} (f : σ → Option σ) (a : σ) : EvalsToInTime f a a 0 := +def EvalsToInTime.refl {σ : Type*} (f : σ → Option σ) (a : σ) : EvalsToInTime f a (some a) 0 := ⟨EvalsTo.refl f a, le_refl 0⟩ /-- Transitivity of `EvalsToInTime` in the sum of the numbers of steps. -/ diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index e0c093182ad2c..88be54ec886b6 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -122,10 +122,6 @@ def Code.eval : Code → List ℕ →. List ℕ namespace Code -/- Porting note: The equation lemma of `eval` is too strong; it simplifies terms like the LHS of -`pred_eval`. Even `eqns` can't fix this. We removed `simp` attr from `eval` and prepare new simp -lemmas for `eval`. -/ - @[simp] theorem zero'_eval : zero'.eval = fun v => pure (0 :: v) := by simp [eval] diff --git a/Mathlib/Computability/Tape.lean b/Mathlib/Computability/Tape.lean index 64bf4d02f6936..cf293a33770e0 100644 --- a/Mathlib/Computability/Tape.lean +++ b/Mathlib/Computability/Tape.lean @@ -584,9 +584,7 @@ theorem Tape.map_write {Γ Γ'} [Inhabited Γ] [Inhabited Γ'] (f : PointedMap ∀ T : Tape Γ, (T.write b).map f = (T.map f).write (f b) := by rintro ⟨⟩; rfl --- Porting note: `simpNF` complains about LHS does not simplify when using the simp lemma on --- itself, but it does indeed. -@[simp, nolint simpNF] +@[simp] theorem Tape.write_move_right_n {Γ} [Inhabited Γ] (f : Γ → Γ) (L R : ListBlank Γ) (n : ℕ) : ((Tape.move Dir.right)^[n] (Tape.mk' L R)).write (f (R.nth n)) = (Tape.move Dir.right)^[n] (Tape.mk' L (R.modifyNth f n)) := by diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 19a73e31325fb..423c7c3acaf78 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -147,8 +147,7 @@ def evalInduction {σ} {f : σ → Option σ} {b : σ} {C : σ → Sort*} {a : theorem mem_eval {σ} {f : σ → Option σ} {a b} : b ∈ eval f a ↔ Reaches f a b ∧ f b = none := by refine ⟨fun h ↦ ?_, fun ⟨h₁, h₂⟩ ↦ ?_⟩ - · -- Porting note: Explicitly specify `c`. - refine @evalInduction _ _ _ (fun a ↦ Reaches f a b ∧ f b = none) _ h fun a h IH ↦ ?_ + · refine evalInduction h fun a h IH ↦ ?_ cases' e : f a with a' · rw [Part.mem_unique h (PFun.mem_fix_iff.2 <| Or.inl <| Part.mem_some_iff.2 <| by rw [e]; rfl)] @@ -580,8 +579,7 @@ def SupportsStmt (S : Finset Λ) : Stmt₁ → Prop | goto l => ∀ a v, l a v ∈ S | halt => True -open scoped Classical - +open scoped Classical in /-- The subterm closure of a statement. -/ noncomputable def stmts₁ : Stmt₁ → Finset Stmt₁ | Q@(move _ q) => insert Q (stmts₁ q) @@ -594,6 +592,7 @@ theorem stmts₁_self {q : Stmt₁} : q ∈ stmts₁ q := by cases q <;> simp only [stmts₁, Finset.mem_insert_self, Finset.mem_singleton_self] theorem stmts₁_trans {q₁ q₂ : Stmt₁} : q₁ ∈ stmts₁ q₂ → stmts₁ q₁ ⊆ stmts₁ q₂ := by + classical intro h₁₂ q₀ h₀₁ induction q₂ with ( simp only [stmts₁] at h₁₂ ⊢ @@ -621,6 +620,7 @@ theorem stmts₁_supportsStmt_mono {S : Finset Λ} {q₁ q₂ : Stmt₁} (h : q | halt => subst h; trivial | _ _ q IH => rcases h with (rfl | h) <;> [exact hs; exact IH h hs] +open scoped Classical in /-- The set of all statements in a Turing machine, plus one extra value `none` representing the halt state. This is used in the TM1 to TM0 reduction. -/ noncomputable def stmts (M : Λ → Stmt₁) (S : Finset Λ) : Finset (Option Stmt₁) := @@ -781,12 +781,11 @@ machine states in the target (even though the type `Λ'` is infinite). -/ noncomputable def trStmts (S : Finset Λ) : Finset Λ'₁₀ := (TM1.stmts M S) ×ˢ Finset.univ -open scoped Classical - attribute [local simp] TM1.stmts₁_self theorem tr_supports {S : Finset Λ} (ss : TM1.Supports M S) : TM0.Supports (tr M) ↑(trStmts M S) := by + classical constructor · apply Finset.mem_product.2 constructor @@ -1113,10 +1112,10 @@ theorem tr_respects : end -open scoped Classical variable [Fintype Γ] +open scoped Classical in /-- The set of accessible `Λ'.write` machine states. -/ noncomputable def writes : Stmt₁ → Finset Λ'₁ | Stmt.move _ q => writes q @@ -1126,11 +1125,13 @@ noncomputable def writes : Stmt₁ → Finset Λ'₁ | Stmt.goto _ => ∅ | Stmt.halt => ∅ +open scoped Classical in /-- The set of accessible machine states, assuming that the input machine is supported on `S`, are the normal states embedded from `S`, plus all write states accessible from these states. -/ noncomputable def trSupp (S : Finset Λ) : Finset Λ'₁ := S.biUnion fun l ↦ insert (Λ'.normal l) (writes (M l)) +open scoped Classical in theorem tr_supports [Inhabited Λ] {S : Finset Λ} (ss : Supports M S) : Supports (tr enc dec M) (trSupp M S) := ⟨Finset.mem_biUnion.2 ⟨_, ss.1, Finset.mem_insert_self _ _⟩, fun q h ↦ by @@ -1378,8 +1379,8 @@ def SupportsStmt (S : Finset Λ) : Stmt₂ → Prop | halt => True section -open scoped Classical +open scoped Classical in /-- The set of subtree statements in a statement. -/ noncomputable def stmts₁ : Stmt₂ → Finset Stmt₂ | Q@(push _ _ q) => insert Q (stmts₁ q) @@ -1394,6 +1395,7 @@ theorem stmts₁_self {q : Stmt₂} : q ∈ stmts₁ q := by cases q <;> simp only [Finset.mem_insert_self, Finset.mem_singleton_self, stmts₁] theorem stmts₁_trans {q₁ q₂ : Stmt₂} : q₁ ∈ stmts₁ q₂ → stmts₁ q₁ ⊆ stmts₁ q₂ := by + classical intro h₁₂ q₀ h₀₁ induction q₂ with ( simp only [stmts₁] at h₁₂ ⊢ @@ -1422,6 +1424,7 @@ theorem stmts₁_supportsStmt_mono {S : Finset Λ} {q₁ q₂ : Stmt₂} (h : q | halt => subst h; trivial | load _ _ IH | _ _ _ _ IH => rcases h with (rfl | h) <;> [exact hs; exact IH h hs] +open scoped Classical in /-- The set of statements accessible from initial set `S` of labels. -/ noncomputable def stmts (M : Λ → Stmt₂) (S : Finset Λ) : Finset (Option Stmt₂) := Finset.insertNone (S.biUnion fun q ↦ stmts₁ (M q)) @@ -1698,8 +1701,8 @@ theorem trNormal_run {k : K} (s : StAct₂ k) (q : Stmt₂) : cases s <;> rfl section -open scoped Classical +open scoped Classical in /-- The set of machine states accessible from an initial TM2 statement. -/ noncomputable def trStmts₁ : Stmt₂ → Finset Λ'₂₁ | TM2.Stmt.push k f q => {go k (StAct.push f) q, ret q} ∪ trStmts₁ q @@ -1709,6 +1712,7 @@ noncomputable def trStmts₁ : Stmt₂ → Finset Λ'₂₁ | TM2.Stmt.branch _ q₁ q₂ => trStmts₁ q₁ ∪ trStmts₁ q₂ | _ => ∅ +open scoped Classical in theorem trStmts₁_run {k : K} {s : StAct₂ k} {q : Stmt₂} : trStmts₁ (stRun s q) = {go k s q, ret q} ∪ trStmts₁ q := by cases s <;> simp only [trStmts₁, stRun] @@ -1927,13 +1931,15 @@ theorem tr_eval (k) (L : List (Γ k)) {L₁ L₂} (H₁ : L₁ ∈ TM1.eval (tr end section -open scoped Classical + variable [Inhabited Λ] +open scoped Classical in /-- The support of a set of TM2 states in the TM2 emulator. -/ noncomputable def trSupp (S : Finset Λ) : Finset Λ'₂₁ := S.biUnion fun l ↦ insert (normal l) (trStmts₁ (M l)) +open scoped Classical in theorem tr_supports {S} (ss : TM2.Supports M S) : TM1.Supports (tr M) (trSupp M S) := ⟨Finset.mem_biUnion.2 ⟨_, ss.1, Finset.mem_insert.2 <| Or.inl rfl⟩, fun l' h ↦ by suffices ∀ (q) (_ : TM2.SupportsStmt S q) (_ : ∀ x ∈ trStmts₁ q, x ∈ trSupp M S), diff --git a/Mathlib/Control/LawfulFix.lean b/Mathlib/Control/LawfulFix.lean index f22207cffe5c1..9c9a30c5b82a3 100644 --- a/Mathlib/Control/LawfulFix.lean +++ b/Mathlib/Control/LawfulFix.lean @@ -22,8 +22,6 @@ omega complete partial orders (ωCPO). Proofs of the lawfulness of all `Fix` ins universe u v -open scoped Classical - variable {α : Type*} {β : α → Type*} open OmegaCompletePartialOrder @@ -63,6 +61,7 @@ theorem approx_mono ⦃i j : ℕ⦄ (hij : i ≤ j) : approx f i ≤ approx f j exact le_trans (ih ‹_›) (approx_mono' f) theorem mem_iff (a : α) (b : β a) : b ∈ Part.fix f a ↔ ∃ i, b ∈ approx f i a := by + classical by_cases h₀ : ∃ i : ℕ, (approx f i a).Dom · simp only [Part.fix_def f h₀] constructor <;> intro hh diff --git a/Mathlib/Data/Bool/Count.lean b/Mathlib/Data/Bool/Count.lean index 8f71b21123ee0..a3bf1706905da 100644 --- a/Mathlib/Data/Bool/Count.lean +++ b/Mathlib/Data/Bool/Count.lean @@ -96,7 +96,7 @@ theorem two_mul_count_bool_eq_ite (hl : Chain' (· ≠ ·) l) (b : Bool) : by_cases h2 : Even (length l) · rw [if_pos h2, hl.two_mul_count_bool_of_even h2] · cases' l with x l - · exact (h2 even_zero).elim + · exact (h2 .zero).elim simp only [if_neg h2, count_cons, mul_add, head?, Option.mem_some_iff, @eq_comm _ x] rw [length_cons, Nat.even_add_one, not_not] at h2 replace hl : l.Chain' (· ≠ ·) := hl.tail diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 7531c38a54ff3..92b673fb0b86a 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -291,10 +291,16 @@ theorem lim_abs (f : CauSeq ℂ Complex.abs) : lim (cauSeqAbs f) = Complex.abs ( let ⟨i, hi⟩ := equiv_lim f ε ε0 ⟨i, fun j hj => lt_of_le_of_lt (Complex.abs.abs_abv_sub_le_abv_sub _ _) (hi j hj)⟩ +lemma ne_zero_of_re_pos {s : ℂ} (hs : 0 < s.re) : s ≠ 0 := + fun h ↦ (zero_re ▸ h ▸ hs).false + lemma ne_zero_of_one_lt_re {s : ℂ} (hs : 1 < s.re) : s ≠ 0 := - fun h ↦ ((zero_re ▸ h ▸ hs).trans zero_lt_one).false + ne_zero_of_re_pos <| zero_lt_one.trans hs + +lemma re_neg_ne_zero_of_re_pos {s : ℂ} (hs : 0 < s.re) : (-s).re ≠ 0 := + ne_iff_lt_or_gt.mpr <| Or.inl <| neg_re s ▸ (neg_lt_zero.mpr hs) lemma re_neg_ne_zero_of_one_lt_re {s : ℂ} (hs : 1 < s.re) : (-s).re ≠ 0 := - ne_iff_lt_or_gt.mpr <| Or.inl <| neg_re s ▸ by linarith + re_neg_ne_zero_of_re_pos <| zero_lt_one.trans hs end Complex diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index eb7b8802a1a40..fd4dfde3db5b6 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -12,10 +12,9 @@ import Mathlib.Tactic.Bound.Attribute import Mathlib.Algebra.CharP.Defs /-! -# Exponential, trigonometric and hyperbolic trigonometric functions +# Exponential Function -This file contains the definitions of the real and complex exponential, sine, cosine, tangent, -hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. +This file contains the definitions of the real and complex exponential function. -/ @@ -51,40 +50,6 @@ def exp' (z : ℂ) : CauSeq ℂ Complex.abs := def exp (z : ℂ) : ℂ := CauSeq.lim (exp' z) -/-- The complex sine function, defined via `exp` -/ -@[pp_nodot] -def sin (z : ℂ) : ℂ := - (exp (-z * I) - exp (z * I)) * I / 2 - -/-- The complex cosine function, defined via `exp` -/ -@[pp_nodot] -def cos (z : ℂ) : ℂ := - (exp (z * I) + exp (-z * I)) / 2 - -/-- The complex tangent function, defined as `sin z / cos z` -/ -@[pp_nodot] -def tan (z : ℂ) : ℂ := - sin z / cos z - -/-- The complex cotangent function, defined as `cos z / sin z` -/ -def cot (z : ℂ) : ℂ := - cos z / sin z - -/-- The complex hyperbolic sine function, defined via `exp` -/ -@[pp_nodot] -def sinh (z : ℂ) : ℂ := - (exp z - exp (-z)) / 2 - -/-- The complex hyperbolic cosine function, defined via `exp` -/ -@[pp_nodot] -def cosh (z : ℂ) : ℂ := - (exp z + exp (-z)) / 2 - -/-- The complex hyperbolic tangent function, defined as `sinh z / cosh z` -/ -@[pp_nodot] -def tanh (z : ℂ) : ℂ := - sinh z / cosh z - /-- scoped notation for the complex exponential function -/ scoped notation "cexp" => Complex.exp @@ -103,41 +68,6 @@ noncomputable section nonrec def exp (x : ℝ) : ℝ := (exp x).re -/-- The real sine function, defined as the real part of the complex sine -/ -@[pp_nodot] -nonrec def sin (x : ℝ) : ℝ := - (sin x).re - -/-- The real cosine function, defined as the real part of the complex cosine -/ -@[pp_nodot] -nonrec def cos (x : ℝ) : ℝ := - (cos x).re - -/-- The real tangent function, defined as the real part of the complex tangent -/ -@[pp_nodot] -nonrec def tan (x : ℝ) : ℝ := - (tan x).re - -/-- The real cotangent function, defined as the real part of the complex cotangent -/ -nonrec def cot (x : ℝ) : ℝ := - (cot x).re - -/-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ -@[pp_nodot] -nonrec def sinh (x : ℝ) : ℝ := - (sinh x).re - -/-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ -@[pp_nodot] -nonrec def cosh (x : ℝ) : ℝ := - (cosh x).re - -/-- The real hypebolic tangent function, defined as the real part of -the complex hyperbolic tangent -/ -@[pp_nodot] -nonrec def tanh (x : ℝ) : ℝ := - (tanh x).re - /-- scoped notation for the real exponential function -/ scoped notation "rexp" => Real.exp @@ -248,430 +178,6 @@ theorem exp_ofReal_im (x : ℝ) : (exp x).im = 0 := by rw [← ofReal_exp_ofReal theorem exp_ofReal_re (x : ℝ) : (exp x).re = Real.exp x := rfl -theorem two_sinh : 2 * sinh x = exp x - exp (-x) := - mul_div_cancel₀ _ two_ne_zero - -theorem two_cosh : 2 * cosh x = exp x + exp (-x) := - mul_div_cancel₀ _ two_ne_zero - -@[simp] -theorem sinh_zero : sinh 0 = 0 := by simp [sinh] - -@[simp] -theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] - -private theorem sinh_add_aux {a b c d : ℂ} : - (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring - -theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← - mul_assoc, two_sinh, mul_left_comm, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, - mul_left_comm, two_cosh, ← mul_assoc, two_cosh] - exact sinh_add_aux - -@[simp] -theorem cosh_zero : cosh 0 = 1 := by simp [cosh] - -@[simp] -theorem cosh_neg : cosh (-x) = cosh x := by simp [add_comm, cosh, exp_neg] - -private theorem cosh_add_aux {a b c d : ℂ} : - (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring - -theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← - mul_assoc, two_cosh, ← mul_assoc, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, - mul_left_comm, two_cosh, mul_left_comm, two_sinh] - exact cosh_add_aux - -theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by - simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] - -theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by - simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] - -theorem sinh_conj : sinh (conj x) = conj (sinh x) := by - rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀] - -- Porting note: not nice - simp [← one_add_one_eq_two] - -@[simp] -theorem ofReal_sinh_ofReal_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := - conj_eq_iff_re.1 <| by rw [← sinh_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := - ofReal_sinh_ofReal_re _ - -@[simp] -theorem sinh_ofReal_im (x : ℝ) : (sinh x).im = 0 := by rw [← ofReal_sinh_ofReal_re, ofReal_im] - -theorem sinh_ofReal_re (x : ℝ) : (sinh x).re = Real.sinh x := - rfl - -theorem cosh_conj : cosh (conj x) = conj (cosh x) := by - rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀] - -- Porting note: not nice - simp [← one_add_one_eq_two] - -theorem ofReal_cosh_ofReal_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := - conj_eq_iff_re.1 <| by rw [← cosh_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := - ofReal_cosh_ofReal_re _ - -@[simp] -theorem cosh_ofReal_im (x : ℝ) : (cosh x).im = 0 := by rw [← ofReal_cosh_ofReal_re, ofReal_im] - -@[simp] -theorem cosh_ofReal_re (x : ℝ) : (cosh x).re = Real.cosh x := - rfl - -theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := - rfl - -@[simp] -theorem tanh_zero : tanh 0 = 0 := by simp [tanh] - -@[simp] -theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] - -theorem tanh_conj : tanh (conj x) = conj (tanh x) := by - rw [tanh, sinh_conj, cosh_conj, ← map_div₀, tanh] - -@[simp] -theorem ofReal_tanh_ofReal_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := - conj_eq_iff_re.1 <| by rw [← tanh_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := - ofReal_tanh_ofReal_re _ - -@[simp] -theorem tanh_ofReal_im (x : ℝ) : (tanh x).im = 0 := by rw [← ofReal_tanh_ofReal_re, ofReal_im] - -theorem tanh_ofReal_re (x : ℝ) : (tanh x).re = Real.tanh x := - rfl - -@[simp] -theorem cosh_add_sinh : cosh x + sinh x = exp x := by - rw [← mul_right_inj' (two_ne_zero' ℂ), mul_add, two_cosh, two_sinh, add_add_sub_cancel, two_mul] - -@[simp] -theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] - -@[simp] -theorem exp_sub_cosh : exp x - cosh x = sinh x := - sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm - -@[simp] -theorem exp_sub_sinh : exp x - sinh x = cosh x := - sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm - -@[simp] -theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by - rw [← mul_right_inj' (two_ne_zero' ℂ), mul_sub, two_cosh, two_sinh, add_sub_sub_cancel, two_mul] - -@[simp] -theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] - -@[simp] -theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by - rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_cancel, exp_zero] - -theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by - rw [← cosh_sq_sub_sinh_sq x] - ring - -theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by - rw [← cosh_sq_sub_sinh_sq x] - ring - -theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by rw [two_mul, cosh_add, sq, sq] - -theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by - rw [two_mul, sinh_add] - ring - -theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, cosh_add x (2 * x)] - simp only [cosh_two_mul, sinh_two_mul] - have h2 : sinh x * (2 * sinh x * cosh x) = 2 * cosh x * sinh x ^ 2 := by ring - rw [h2, sinh_sq] - ring - -theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, sinh_add x (2 * x)] - simp only [cosh_two_mul, sinh_two_mul] - have h2 : cosh x * (2 * sinh x * cosh x) = 2 * sinh x * cosh x ^ 2 := by ring - rw [h2, cosh_sq] - ring - -@[simp] -theorem sin_zero : sin 0 = 0 := by simp [sin] - -@[simp] -theorem sin_neg : sin (-x) = -sin x := by - simp [sin, sub_eq_add_neg, exp_neg, (neg_div _ _).symm, add_mul] - -theorem two_sin : 2 * sin x = (exp (-x * I) - exp (x * I)) * I := - mul_div_cancel₀ _ two_ne_zero - -theorem two_cos : 2 * cos x = exp (x * I) + exp (-x * I) := - mul_div_cancel₀ _ two_ne_zero - -theorem sinh_mul_I : sinh (x * I) = sin x * I := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, - mul_neg_one, neg_sub, neg_mul_eq_neg_mul] - -theorem cosh_mul_I : cosh (x * I) = cos x := by - rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] - -theorem tanh_mul_I : tanh (x * I) = tan x * I := by - rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] - -theorem cos_mul_I : cos (x * I) = cosh x := by rw [← cosh_mul_I]; ring_nf; simp - -theorem sin_mul_I : sin (x * I) = sinh x * I := by - have h : I * sin (x * I) = -sinh x := by - rw [mul_comm, ← sinh_mul_I] - ring_nf - simp - rw [← neg_neg (sinh x), ← h] - apply Complex.ext <;> simp - -theorem tan_mul_I : tan (x * I) = tanh x * I := by - rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] - -theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by - rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, add_mul, add_mul, mul_right_comm, ← sinh_mul_I, - mul_assoc, ← sinh_mul_I, ← cosh_mul_I, ← cosh_mul_I, sinh_add] - -@[simp] -theorem cos_zero : cos 0 = 1 := by simp [cos] - -@[simp] -theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] - -theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by - rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, - mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] - -theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by - simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] - -theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by - simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] - -theorem sin_add_mul_I (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by - rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] - -theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * I := by - convert sin_add_mul_I z.re z.im; exact (re_add_im z).symm - -theorem cos_add_mul_I (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by - rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] - -theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * I := by - convert cos_add_mul_I z.re z.im; exact (re_add_im z).symm - -theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by - have s1 := sin_add ((x + y) / 2) ((x - y) / 2) - have s2 := sin_sub ((x + y) / 2) ((x - y) / 2) - rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 - rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 - rw [s1, s2] - ring - -theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by - have s1 := cos_add ((x + y) / 2) ((x - y) / 2) - have s2 := cos_sub ((x + y) / 2) ((x - y) / 2) - rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 - rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 - rw [s1, s2] - ring - -theorem sin_add_sin : sin x + sin y = 2 * sin ((x + y) / 2) * cos ((x - y) / 2) := by - simpa using sin_sub_sin x (-y) - -theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by - calc - cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := ?_ - _ = - cos ((x + y) / 2) * cos ((x - y) / 2) - sin ((x + y) / 2) * sin ((x - y) / 2) + - (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := - ?_ - _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := ?_ - - · congr <;> field_simp - · rw [cos_add, cos_sub] - ring - -theorem sin_conj : sin (conj x) = conj (sin x) := by - rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← RingHom.map_mul, - sinh_conj, mul_neg, sinh_neg, sinh_mul_I, mul_neg] - -@[simp] -theorem ofReal_sin_ofReal_re (x : ℝ) : ((sin x).re : ℂ) = sin x := - conj_eq_iff_re.1 <| by rw [← sin_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := - ofReal_sin_ofReal_re _ - -@[simp] -theorem sin_ofReal_im (x : ℝ) : (sin x).im = 0 := by rw [← ofReal_sin_ofReal_re, ofReal_im] - -theorem sin_ofReal_re (x : ℝ) : (sin x).re = Real.sin x := - rfl - -theorem cos_conj : cos (conj x) = conj (cos x) := by - rw [← cosh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] - -@[simp] -theorem ofReal_cos_ofReal_re (x : ℝ) : ((cos x).re : ℂ) = cos x := - conj_eq_iff_re.1 <| by rw [← cos_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := - ofReal_cos_ofReal_re _ - -@[simp] -theorem cos_ofReal_im (x : ℝ) : (cos x).im = 0 := by rw [← ofReal_cos_ofReal_re, ofReal_im] - -theorem cos_ofReal_re (x : ℝ) : (cos x).re = Real.cos x := - rfl - -@[simp] -theorem tan_zero : tan 0 = 0 := by simp [tan] - -theorem tan_eq_sin_div_cos : tan x = sin x / cos x := - rfl - -theorem cot_eq_cos_div_sin : cot x = cos x / sin x := - rfl - -theorem tan_mul_cos {x : ℂ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by - rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] - -@[simp] -theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] - -theorem tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj, ← map_div₀, tan] - -theorem cot_conj : cot (conj x) = conj (cot x) := by rw [cot, sin_conj, cos_conj, ← map_div₀, cot] - -@[simp] -theorem ofReal_tan_ofReal_re (x : ℝ) : ((tan x).re : ℂ) = tan x := - conj_eq_iff_re.1 <| by rw [← tan_conj, conj_ofReal] - -@[simp] -theorem ofReal_cot_ofReal_re (x : ℝ) : ((cot x).re : ℂ) = cot x := - conj_eq_iff_re.1 <| by rw [← cot_conj, conj_ofReal] - -@[simp, norm_cast] -theorem ofReal_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := - ofReal_tan_ofReal_re _ - -@[simp, norm_cast] -theorem ofReal_cot (x : ℝ) : (Real.cot x : ℂ) = cot x := - ofReal_cot_ofReal_re _ - -@[simp] -theorem tan_ofReal_im (x : ℝ) : (tan x).im = 0 := by rw [← ofReal_tan_ofReal_re, ofReal_im] - -theorem tan_ofReal_re (x : ℝ) : (tan x).re = Real.tan x := - rfl - -theorem cos_add_sin_I : cos x + sin x * I = exp (x * I) := by - rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] - -theorem cos_sub_sin_I : cos x - sin x * I = exp (-x * I) := by - rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] - -@[simp] -theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := - Eq.trans (by rw [cosh_mul_I, sinh_mul_I, mul_pow, I_sq, mul_neg_one, sub_neg_eq_add, add_comm]) - (cosh_sq_sub_sinh_sq (x * I)) - -@[simp] -theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] - -theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by rw [two_mul, cos_add, ← sq, ← sq] - -theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by - rw [cos_two_mul', eq_sub_iff_add_eq.2 (sin_sq_add_cos_sq x), ← sub_add, sub_add_eq_add_sub, - two_mul] - -theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by - rw [two_mul, sin_add, two_mul, add_mul, mul_comm] - -theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := by - simp [cos_two_mul, div_add_div_same, mul_div_cancel_left₀, two_ne_zero, -one_div] - -theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] - -theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_right] - -theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := by - rw [tan_eq_sin_div_cos, div_pow] - field_simp - -theorem tan_sq_div_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : - tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by - simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] - -theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, cos_add x (2 * x)] - simp only [cos_two_mul, sin_two_mul, mul_add, mul_sub, mul_one, sq] - have h2 : 4 * cos x ^ 3 = 2 * cos x * cos x * cos x + 2 * cos x * cos x ^ 2 := by ring - rw [h2, cos_sq'] - ring - -theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by - have h1 : x + 2 * x = 3 * x := by ring - rw [← h1, sin_add x (2 * x)] - simp only [cos_two_mul, sin_two_mul, cos_sq'] - have h2 : cos x * (2 * sin x * cos x) = 2 * sin x * cos x ^ 2 := by ring - rw [h2, cos_sq'] - ring - -theorem exp_mul_I : exp (x * I) = cos x + sin x * I := - (cos_add_sin_I _).symm - -theorem exp_add_mul_I : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] - -theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * I) := by - rw [← exp_add_mul_I, re_add_im] - -theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := by - rw [exp_eq_exp_re_mul_sin_add_cos] - simp [exp_ofReal_re, cos_ofReal_re] - -theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by - rw [exp_eq_exp_re_mul_sin_add_cos] - simp [exp_ofReal_re, sin_ofReal_re] - -@[simp] -theorem exp_ofReal_mul_I_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by - simp [exp_mul_I, cos_ofReal_re] - -@[simp] -theorem exp_ofReal_mul_I_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by - simp [exp_mul_I, sin_ofReal_re] - -/-- **De Moivre's formula** -/ -theorem cos_add_sin_mul_I_pow (n : ℕ) (z : ℂ) : - (cos z + sin z * I) ^ n = cos (↑n * z) + sin (↑n * z) * I := by - rw [← exp_mul_I, ← exp_mul_I] - induction' n with n ih - · rw [pow_zero, Nat.cast_zero, zero_mul, zero_mul, exp_zero] - · rw [pow_succ, ih, Nat.cast_succ, add_mul, add_mul, one_mul, exp_add] - end Complex namespace Real @@ -719,241 +225,6 @@ nonrec theorem exp_neg : exp (-x) = (exp x)⁻¹ := theorem exp_sub : exp (x - y) = exp x / exp y := by simp [sub_eq_add_neg, exp_add, exp_neg, div_eq_mul_inv] -@[simp] -theorem sin_zero : sin 0 = 0 := by simp [sin] - -@[simp] -theorem sin_neg : sin (-x) = -sin x := by simp [sin, exp_neg, (neg_div _ _).symm, add_mul] - -nonrec theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := - ofReal_injective <| by simp [sin_add] - -@[simp] -theorem cos_zero : cos 0 = 1 := by simp [cos] - -@[simp] -theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] - -@[simp] -theorem cos_abs : cos |x| = cos x := by - cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] - -nonrec theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := - ofReal_injective <| by simp [cos_add] - -theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by - simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] - -theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by - simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] - -nonrec theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := - ofReal_injective <| by simp [sin_sub_sin] - -nonrec theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := - ofReal_injective <| by simp [cos_sub_cos] - -nonrec theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := - ofReal_injective <| by simp [cos_add_cos] - -theorem two_mul_sin_mul_sin (x y : ℝ) : 2 * sin x * sin y = cos (x - y) - cos (x + y) := by - simp [cos_add, cos_sub] - ring - -theorem two_mul_cos_mul_cos (x y : ℝ) : 2 * cos x * cos y = cos (x - y) + cos (x + y) := by - simp [cos_add, cos_sub] - ring - -theorem two_mul_sin_mul_cos (x y : ℝ) : 2 * sin x * cos y = sin (x - y) + sin (x + y) := by - simp [sin_add, sin_sub] - ring - -nonrec theorem tan_eq_sin_div_cos : tan x = sin x / cos x := - ofReal_injective <| by simp only [ofReal_tan, tan_eq_sin_div_cos, ofReal_div, ofReal_sin, - ofReal_cos] - -nonrec theorem cot_eq_cos_div_sin : cot x = cos x / sin x := - ofReal_injective <| by simp [cot_eq_cos_div_sin] - -theorem tan_mul_cos {x : ℝ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by - rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] - -@[simp] -theorem tan_zero : tan 0 = 0 := by simp [tan] - -@[simp] -theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] - -@[simp] -nonrec theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := - ofReal_injective (by simp [sin_sq_add_cos_sq]) - -@[simp] -theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] - -theorem sin_sq_le_one : sin x ^ 2 ≤ 1 := by - rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_right (sq_nonneg _) - -theorem cos_sq_le_one : cos x ^ 2 ≤ 1 := by - rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_left (sq_nonneg _) - -theorem abs_sin_le_one : |sin x| ≤ 1 := - abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, sin_sq_le_one] - -theorem abs_cos_le_one : |cos x| ≤ 1 := - abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, cos_sq_le_one] - -theorem sin_le_one : sin x ≤ 1 := - (abs_le.1 (abs_sin_le_one _)).2 - -theorem cos_le_one : cos x ≤ 1 := - (abs_le.1 (abs_cos_le_one _)).2 - -theorem neg_one_le_sin : -1 ≤ sin x := - (abs_le.1 (abs_sin_le_one _)).1 - -theorem neg_one_le_cos : -1 ≤ cos x := - (abs_le.1 (abs_cos_le_one _)).1 - -nonrec theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := - ofReal_injective <| by simp [cos_two_mul] - -nonrec theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := - ofReal_injective <| by simp [cos_two_mul'] - -nonrec theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := - ofReal_injective <| by simp [sin_two_mul] - -nonrec theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := - ofReal_injective <| by simp [cos_sq] - -theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] - -theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := - eq_sub_iff_add_eq.2 <| sin_sq_add_cos_sq _ - -lemma sin_sq_eq_half_sub : sin x ^ 2 = 1 / 2 - cos (2 * x) / 2 := by - rw [sin_sq, cos_sq, ← sub_sub, sub_half] - -theorem abs_sin_eq_sqrt_one_sub_cos_sq (x : ℝ) : |sin x| = √(1 - cos x ^ 2) := by - rw [← sin_sq, sqrt_sq_eq_abs] - -theorem abs_cos_eq_sqrt_one_sub_sin_sq (x : ℝ) : |cos x| = √(1 - sin x ^ 2) := by - rw [← cos_sq', sqrt_sq_eq_abs] - -theorem inv_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := - have : Complex.cos x ≠ 0 := mt (congr_arg re) hx - ofReal_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this - -theorem tan_sq_div_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : - tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by - simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] - -theorem inv_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : (√(1 + tan x ^ 2))⁻¹ = cos x := by - rw [← sqrt_sq hx.le, ← sqrt_inv, inv_one_add_tan_sq hx.ne'] - -theorem tan_div_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : - tan x / √(1 + tan x ^ 2) = sin x := by - rw [← tan_mul_cos hx.ne', ← inv_sqrt_one_add_tan_sq hx, div_eq_mul_inv] - -nonrec theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by - rw [← ofReal_inj]; simp [cos_three_mul] - -nonrec theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by - rw [← ofReal_inj]; simp [sin_three_mul] - -/-- The definition of `sinh` in terms of `exp`. -/ -nonrec theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := - ofReal_injective <| by simp [Complex.sinh] - -@[simp] -theorem sinh_zero : sinh 0 = 0 := by simp [sinh] - -@[simp] -theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] - -nonrec theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by - rw [← ofReal_inj]; simp [sinh_add] - -/-- The definition of `cosh` in terms of `exp`. -/ -theorem cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 := - eq_div_of_mul_eq two_ne_zero <| by - rw [cosh, exp, exp, Complex.ofReal_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, - div_mul_cancel₀ _ (two_ne_zero' ℂ), Complex.add_re] - -@[simp] -theorem cosh_zero : cosh 0 = 1 := by simp [cosh] - -@[simp] -theorem cosh_neg : cosh (-x) = cosh x := - ofReal_inj.1 <| by simp - -@[simp] -theorem cosh_abs : cosh |x| = cosh x := by - cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] - -nonrec theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by - rw [← ofReal_inj]; simp [cosh_add] - -theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by - simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] - -theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by - simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] - -nonrec theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := - ofReal_inj.1 <| by simp [tanh_eq_sinh_div_cosh] - -@[simp] -theorem tanh_zero : tanh 0 = 0 := by simp [tanh] - -@[simp] -theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] - -@[simp] -theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← ofReal_inj]; simp - -@[simp] -theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] - -@[simp] -theorem exp_sub_cosh : exp x - cosh x = sinh x := - sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm - -@[simp] -theorem exp_sub_sinh : exp x - sinh x = cosh x := - sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm - -@[simp] -theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by - rw [← ofReal_inj] - simp - -@[simp] -theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] - -@[simp] -theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← ofReal_inj]; simp - -nonrec theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← ofReal_inj]; simp [cosh_sq] - -theorem cosh_sq' : cosh x ^ 2 = 1 + sinh x ^ 2 := - (cosh_sq x).trans (add_comm _ _) - -nonrec theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← ofReal_inj]; simp [sinh_sq] - -nonrec theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by - rw [← ofReal_inj]; simp [cosh_two_mul] - -nonrec theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by - rw [← ofReal_inj]; simp [sinh_two_mul] - -nonrec theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by - rw [← ofReal_inj]; simp [cosh_three_mul] - -nonrec theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by - rw [← ofReal_inj]; simp [sinh_three_mul] - open IsAbsoluteValue Nat theorem sum_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : ∑ i ∈ range n, x ^ i / i ! ≤ exp x := @@ -1058,13 +329,6 @@ theorem exp_le_one_iff {x : ℝ} : exp x ≤ 1 ↔ x ≤ 0 := theorem one_le_exp_iff {x : ℝ} : 1 ≤ exp x ↔ 0 ≤ x := exp_zero ▸ exp_le_exp -/-- `Real.cosh` is always positive -/ -theorem cosh_pos (x : ℝ) : 0 < Real.cosh x := - (cosh_eq x).symm ▸ half_pos (add_pos (exp_pos x) (exp_pos (-x))) - -theorem sinh_lt_cosh : sinh x < cosh x := - lt_of_pow_lt_pow_left₀ 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ - end Real namespace Complex @@ -1175,6 +439,71 @@ theorem abs_exp_sub_one_sub_id_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1 _ ≤ abs x ^ 2 * 1 := by gcongr; norm_num [Nat.factorial] _ = abs x ^ 2 := by rw [mul_one] +lemma abs_exp_sub_sum_le_exp_abs_sub_sum (x : ℂ) (n : ℕ) : + abs (exp x - ∑ m ∈ range n, x ^ m / m.factorial) + ≤ Real.exp (abs x) - ∑ m ∈ range n, (abs x) ^ m / m.factorial := by + rw [← CauSeq.lim_const (abv := Complex.abs) (∑ m ∈ range n, _), Complex.exp, sub_eq_add_neg, + ← CauSeq.lim_neg, CauSeq.lim_add, ← lim_abs] + refine CauSeq.lim_le (CauSeq.le_of_exists ⟨n, fun j hj => ?_⟩) + simp_rw [← sub_eq_add_neg] + calc abs ((∑ m ∈ range j, x ^ m / m.factorial) - ∑ m ∈ range n, x ^ m / m.factorial) + _ ≤ (∑ m ∈ range j, abs x ^ m / m.factorial) - ∑ m ∈ range n, abs x ^ m / m.factorial := by + rw [sum_range_sub_sum_range hj, sum_range_sub_sum_range hj] + refine (IsAbsoluteValue.abv_sum Complex.abs ..).trans_eq ?_ + congr with i + simp + _ ≤ Real.exp (abs x) - ∑ m ∈ range n, (abs x) ^ m / m.factorial := by + gcongr + exact Real.sum_le_exp_of_nonneg (by exact AbsoluteValue.nonneg abs x) _ + +lemma abs_exp_le_exp_abs (x : ℂ) : abs (exp x) ≤ Real.exp (abs x) := by + convert abs_exp_sub_sum_le_exp_abs_sub_sum x 0 using 1 <;> simp + +lemma abs_exp_sub_sum_le_abs_mul_exp (x : ℂ) (n : ℕ) : + abs (exp x - ∑ m ∈ range n, x ^ m / m.factorial) ≤ abs x ^ n * Real.exp (abs x) := by + rw [← CauSeq.lim_const (abv := Complex.abs) (∑ m ∈ range n, _), Complex.exp, sub_eq_add_neg, + ← CauSeq.lim_neg, CauSeq.lim_add, ← lim_abs] + refine CauSeq.lim_le (CauSeq.le_of_exists ⟨n, fun j hj => ?_⟩) + simp_rw [← sub_eq_add_neg] + show abs ((∑ m ∈ range j, x ^ m / m.factorial) - ∑ m ∈ range n, x ^ m / m.factorial) ≤ _ + rw [sum_range_sub_sum_range hj] + calc + abs (∑ m ∈ range j with n ≤ m, (x ^ m / m.factorial : ℂ)) + = abs (∑ m ∈ range j with n ≤ m, (x ^ n * (x ^ (m - n) / m.factorial) : ℂ)) := by + refine congr_arg abs (sum_congr rfl fun m hm => ?_) + rw [mem_filter, mem_range] at hm + rw [← mul_div_assoc, ← pow_add, add_tsub_cancel_of_le hm.2] + _ ≤ ∑ m ∈ range j with n ≤ m, abs (x ^ n * (x ^ (m - n) / m.factorial)) := + IsAbsoluteValue.abv_sum Complex.abs .. + _ ≤ ∑ m ∈ range j with n ≤ m, abs x ^ n * (abs x ^ (m - n) / (m - n).factorial) := by + simp_rw [map_mul, map_pow, map_div₀, abs_natCast] + gcongr with i hi + · rw [IsAbsoluteValue.abv_pow abs] + · simp + _ = abs x ^ n * ∑ m ∈ range j with n ≤ m, (abs x ^ (m - n) / (m - n).factorial) := by + rw [← mul_sum] + _ = abs x ^ n * ∑ m ∈ range (j - n), (abs x ^ m / m.factorial) := by + congr 1 + refine (sum_bij (fun m hm ↦ m + n) ?_ ?_ ?_ ?_).symm + · intro a ha + simp only [mem_filter, mem_range, le_add_iff_nonneg_left, zero_le, and_true] + simp only [mem_range] at ha + rwa [← lt_tsub_iff_right] + · intro a ha b hb hab + simpa using hab + · intro b hb + simp only [mem_range, exists_prop] + simp only [mem_filter, mem_range] at hb + refine ⟨b - n, ?_, ?_⟩ + · rw [tsub_lt_tsub_iff_right hb.2] + exact hb.1 + · rw [tsub_add_cancel_of_le hb.2] + · simp + _ ≤ abs x ^ n * Real.exp (abs x) := by + gcongr + refine Real.sum_le_exp_of_nonneg ?_ _ + exact AbsoluteValue.nonneg abs x + end Complex namespace Real @@ -1272,125 +601,6 @@ theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {r theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 / Nat.factorial 0 * b) : |exp x - a| ≤ b := by simpa using h -theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := - calc - |cos x - (1 - x ^ 2 / 2)| = Complex.abs (Complex.cos x - (1 - (x : ℂ) ^ 2 / 2)) := by - rw [← abs_ofReal]; simp - _ = Complex.abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - (x : ℂ) ^ 2)) / 2) := by - simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] - _ = abs - (((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) + - (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial)) / 2) := - (congr_arg Complex.abs - (congr_arg (fun x : ℂ => x / 2) - (by - simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, range_zero, sum_empty, - Nat.factorial, Nat.cast_one, ne_eq, one_ne_zero, not_false_eq_true, div_self, - zero_add, div_one, Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat, mul_neg, - neg_neg] - apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring_nf - ))) - _ ≤ abs ((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2) + - abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2) := by - rw [add_div]; exact Complex.abs.add_le _ _ - _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + - abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by - simp [map_div₀] - _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 := by - gcongr - · exact Complex.exp_bound (by simpa) (by decide) - · exact Complex.exp_bound (by simpa) (by decide) - _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] - -theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := - calc - |sin x - (x - x ^ 3 / 6)| = Complex.abs (Complex.sin x - (x - x ^ 3 / 6 : ℝ)) := by - rw [← abs_ofReal]; simp - _ = Complex.abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - - (2 * x - x ^ 3 / 3 : ℝ)) / 2) := by - simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left₀ _ (two_ne_zero' ℂ), - div_div, show (3 : ℂ) * 2 = 6 by norm_num] - _ = Complex.abs (((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) - - (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial)) * I / 2) := - (congr_arg Complex.abs - (congr_arg (fun x : ℂ => x / 2) - (by - simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, ofReal_sub, ofReal_mul, - ofReal_ofNat, ofReal_div, range_zero, sum_empty, Nat.factorial, Nat.cast_one, ne_eq, - one_ne_zero, not_false_eq_true, div_self, zero_add, div_one, mul_neg, neg_neg, - Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat] - apply Complex.ext <;> simp [div_eq_mul_inv, normSq]; ring))) - _ ≤ abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) * I / 2) + - abs (-((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) * I) / 2) := by - rw [sub_mul, sub_eq_add_neg, add_div]; exact Complex.abs.add_le _ _ - _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + - abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by - simp [add_comm, map_div₀] - _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 := by - gcongr - · exact Complex.exp_bound (by simpa) (by decide) - · exact Complex.exp_bound (by simpa) (by decide) - _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] - -theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := - calc 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := - sub_pos.2 <| - lt_sub_iff_add_lt.2 - (calc - |x| ^ 4 * (5 / 96) + x ^ 2 / 2 ≤ 1 * (5 / 96) + 1 / 2 := by - gcongr - · exact pow_le_one₀ (abs_nonneg _) hx - · rw [sq, ← abs_mul_self, abs_mul] - exact mul_le_one₀ hx (abs_nonneg _) hx - _ < 1 := by norm_num) - _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 - -theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := - calc 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := - sub_pos.2 <| lt_sub_iff_add_lt.2 - (calc - |x| ^ 4 * (5 / 96) + x ^ 3 / 6 ≤ x * (5 / 96) + x / 6 := by - gcongr - · calc - |x| ^ 4 ≤ |x| ^ 1 := - pow_le_pow_of_le_one (abs_nonneg _) - (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) - _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] - · calc - x ^ 3 ≤ x ^ 1 := pow_le_pow_of_le_one (le_of_lt hx0) hx (by decide) - _ = x := pow_one _ - _ < x := by linarith) - _ ≤ sin x := - sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 - -theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := - have : x / 2 ≤ 1 := (div_le_iff₀ (by norm_num)).mpr (by simpa) - calc - 0 < 2 * sin (x / 2) * cos (x / 2) := - mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) - (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) - _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] - -theorem cos_one_le : cos 1 ≤ 2 / 3 := - calc - cos 1 ≤ |(1 : ℝ)| ^ 4 * (5 / 96) + (1 - 1 ^ 2 / 2) := - sub_le_iff_le_add.1 (abs_sub_le_iff.1 (cos_bound (by simp))).1 - _ ≤ 2 / 3 := by norm_num - -theorem cos_one_pos : 0 < cos 1 := - cos_pos_of_le_one (le_of_eq abs_one) - -theorem cos_two_neg : cos 2 < 0 := - calc cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm - _ = _ := Real.cos_two_mul 1 - _ ≤ 2 * (2 / 3) ^ 2 - 1 := by - gcongr - · exact cos_one_pos.le - · apply cos_one_le - _ < 0 := by norm_num - theorem exp_bound_div_one_sub_of_interval' {x : ℝ} (h1 : 0 < x) (h2 : x < 1) : Real.exp x < 1 / (1 - x) := by have H : 0 < 1 - (1 + x + x ^ 2) * (1 - x) := calc @@ -1445,6 +655,12 @@ theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t · exact one_sub_le_exp_neg _ _ = rexp (-t) := by rw [← Real.exp_nat_mul, mul_neg, mul_comm, div_mul_cancel₀]; positivity +lemma le_inv_mul_exp (x : ℝ) {c : ℝ} (hc : 0 < c) : x ≤ c⁻¹ * exp (c * x) := by + rw [le_inv_mul_iff₀ hc] + calc c * x + _ ≤ c * x + 1 := le_add_of_nonneg_right zero_le_one + _ ≤ _ := Real.add_one_le_exp (c * x) + end Real namespace Mathlib.Meta.Positivity @@ -1459,39 +675,13 @@ def evalExp : PositivityExt where eval {u α} _ _ e := do pure (.positive q(Real.exp_pos $a)) | _, _, _ => throwError "not Real.exp" -/-- Extension for the `positivity` tactic: `Real.cosh` is always positive. -/ -@[positivity Real.cosh _] -def evalCosh : PositivityExt where eval {u α} _ _ e := do - match u, α, e with - | 0, ~q(ℝ), ~q(Real.cosh $a) => - assertInstancesCommute - return .positive q(Real.cosh_pos $a) - | _, _, _ => throwError "not Real.cosh" - -example (x : ℝ) : 0 < x.cosh := by positivity - end Mathlib.Meta.Positivity namespace Complex -@[simp] -theorem abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := by - have := Real.sin_sq_add_cos_sq x - simp_all [add_comm, abs, normSq, sq, sin_ofReal_re, cos_ofReal_re, mul_re] - @[simp] theorem abs_exp_ofReal (x : ℝ) : abs (exp x) = Real.exp x := by rw [← ofReal_exp] exact abs_of_nonneg (le_of_lt (Real.exp_pos _)) -@[simp] -theorem abs_exp_ofReal_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by - rw [exp_mul_I, abs_cos_add_sin_mul_I] - -theorem abs_exp (z : ℂ) : abs (exp z) = Real.exp z.re := by - rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_ofReal, abs_cos_add_sin_mul_I, mul_one] - -theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by - rw [abs_exp, abs_exp, Real.exp_eq_exp] - end Complex diff --git a/Mathlib/Data/Complex/Trigonometric.lean b/Mathlib/Data/Complex/Trigonometric.lean new file mode 100644 index 0000000000000..ea6b53a511038 --- /dev/null +++ b/Mathlib/Data/Complex/Trigonometric.lean @@ -0,0 +1,951 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Abhimanyu Pallavi Sudhir +-/ +import Mathlib.Data.Complex.Exponential + +/-! +# Trigonometric and hyperbolic trigonometric functions + +This file contains the definitions of the sine, cosine, tangent, +hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. + +-/ + +open CauSeq Finset IsAbsoluteValue +open scoped ComplexConjugate + +namespace Complex + +noncomputable section + +/-- The complex sine function, defined via `exp` -/ +@[pp_nodot] +def sin (z : ℂ) : ℂ := + (exp (-z * I) - exp (z * I)) * I / 2 + +/-- The complex cosine function, defined via `exp` -/ +@[pp_nodot] +def cos (z : ℂ) : ℂ := + (exp (z * I) + exp (-z * I)) / 2 + +/-- The complex tangent function, defined as `sin z / cos z` -/ +@[pp_nodot] +def tan (z : ℂ) : ℂ := + sin z / cos z + +/-- The complex cotangent function, defined as `cos z / sin z` -/ +def cot (z : ℂ) : ℂ := + cos z / sin z + +/-- The complex hyperbolic sine function, defined via `exp` -/ +@[pp_nodot] +def sinh (z : ℂ) : ℂ := + (exp z - exp (-z)) / 2 + +/-- The complex hyperbolic cosine function, defined via `exp` -/ +@[pp_nodot] +def cosh (z : ℂ) : ℂ := + (exp z + exp (-z)) / 2 + +/-- The complex hyperbolic tangent function, defined as `sinh z / cosh z` -/ +@[pp_nodot] +def tanh (z : ℂ) : ℂ := + sinh z / cosh z + +end + +end Complex + +namespace Real + +open Complex + +noncomputable section + +/-- The real sine function, defined as the real part of the complex sine -/ +@[pp_nodot] +nonrec def sin (x : ℝ) : ℝ := + (sin x).re + +/-- The real cosine function, defined as the real part of the complex cosine -/ +@[pp_nodot] +nonrec def cos (x : ℝ) : ℝ := + (cos x).re + +/-- The real tangent function, defined as the real part of the complex tangent -/ +@[pp_nodot] +nonrec def tan (x : ℝ) : ℝ := + (tan x).re + +/-- The real cotangent function, defined as the real part of the complex cotangent -/ +nonrec def cot (x : ℝ) : ℝ := + (cot x).re + +/-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ +@[pp_nodot] +nonrec def sinh (x : ℝ) : ℝ := + (sinh x).re + +/-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ +@[pp_nodot] +nonrec def cosh (x : ℝ) : ℝ := + (cosh x).re + +/-- The real hypebolic tangent function, defined as the real part of +the complex hyperbolic tangent -/ +@[pp_nodot] +nonrec def tanh (x : ℝ) : ℝ := + (tanh x).re + +end + +end Real + +namespace Complex + +variable (x y : ℂ) + +theorem two_sinh : 2 * sinh x = exp x - exp (-x) := + mul_div_cancel₀ _ two_ne_zero + +theorem two_cosh : 2 * cosh x = exp x + exp (-x) := + mul_div_cancel₀ _ two_ne_zero + +@[simp] +theorem sinh_zero : sinh 0 = 0 := by simp [sinh] + +@[simp] +theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] + +private theorem sinh_add_aux {a b c d : ℂ} : + (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring + +theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← + mul_assoc, two_sinh, mul_left_comm, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, + mul_left_comm, two_cosh, ← mul_assoc, two_cosh] + exact sinh_add_aux + +@[simp] +theorem cosh_zero : cosh 0 = 1 := by simp [cosh] + +@[simp] +theorem cosh_neg : cosh (-x) = cosh x := by simp [add_comm, cosh, exp_neg] + +private theorem cosh_add_aux {a b c d : ℂ} : + (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring + +theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← + mul_assoc, two_cosh, ← mul_assoc, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, + mul_left_comm, two_cosh, mul_left_comm, two_sinh] + exact cosh_add_aux + +theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by + simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] + +theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by + simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] + +theorem sinh_conj : sinh (conj x) = conj (sinh x) := by + rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀] + -- Porting note: not nice + simp [← one_add_one_eq_two] + +@[simp] +theorem ofReal_sinh_ofReal_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := + conj_eq_iff_re.1 <| by rw [← sinh_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := + ofReal_sinh_ofReal_re _ + +@[simp] +theorem sinh_ofReal_im (x : ℝ) : (sinh x).im = 0 := by rw [← ofReal_sinh_ofReal_re, ofReal_im] + +theorem sinh_ofReal_re (x : ℝ) : (sinh x).re = Real.sinh x := + rfl + +theorem cosh_conj : cosh (conj x) = conj (cosh x) := by + rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀] + -- Porting note: not nice + simp [← one_add_one_eq_two] + +theorem ofReal_cosh_ofReal_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := + conj_eq_iff_re.1 <| by rw [← cosh_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := + ofReal_cosh_ofReal_re _ + +@[simp] +theorem cosh_ofReal_im (x : ℝ) : (cosh x).im = 0 := by rw [← ofReal_cosh_ofReal_re, ofReal_im] + +@[simp] +theorem cosh_ofReal_re (x : ℝ) : (cosh x).re = Real.cosh x := + rfl + +theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + rfl + +@[simp] +theorem tanh_zero : tanh 0 = 0 := by simp [tanh] + +@[simp] +theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] + +theorem tanh_conj : tanh (conj x) = conj (tanh x) := by + rw [tanh, sinh_conj, cosh_conj, ← map_div₀, tanh] + +@[simp] +theorem ofReal_tanh_ofReal_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := + conj_eq_iff_re.1 <| by rw [← tanh_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := + ofReal_tanh_ofReal_re _ + +@[simp] +theorem tanh_ofReal_im (x : ℝ) : (tanh x).im = 0 := by rw [← ofReal_tanh_ofReal_re, ofReal_im] + +theorem tanh_ofReal_re (x : ℝ) : (tanh x).re = Real.tanh x := + rfl + +@[simp] +theorem cosh_add_sinh : cosh x + sinh x = exp x := by + rw [← mul_right_inj' (two_ne_zero' ℂ), mul_add, two_cosh, two_sinh, add_add_sub_cancel, two_mul] + +@[simp] +theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] + +@[simp] +theorem exp_sub_cosh : exp x - cosh x = sinh x := + sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm + +@[simp] +theorem exp_sub_sinh : exp x - sinh x = cosh x := + sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm + +@[simp] +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by + rw [← mul_right_inj' (two_ne_zero' ℂ), mul_sub, two_cosh, two_sinh, add_sub_sub_cancel, two_mul] + +@[simp] +theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] + +@[simp] +theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by + rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_cancel, exp_zero] + +theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by + rw [← cosh_sq_sub_sinh_sq x] + ring + +theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by + rw [← cosh_sq_sub_sinh_sq x] + ring + +theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by rw [two_mul, cosh_add, sq, sq] + +theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by + rw [two_mul, sinh_add] + ring + +theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, cosh_add x (2 * x)] + simp only [cosh_two_mul, sinh_two_mul] + have h2 : sinh x * (2 * sinh x * cosh x) = 2 * cosh x * sinh x ^ 2 := by ring + rw [h2, sinh_sq] + ring + +theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, sinh_add x (2 * x)] + simp only [cosh_two_mul, sinh_two_mul] + have h2 : cosh x * (2 * sinh x * cosh x) = 2 * sinh x * cosh x ^ 2 := by ring + rw [h2, cosh_sq] + ring + +@[simp] +theorem sin_zero : sin 0 = 0 := by simp [sin] + +@[simp] +theorem sin_neg : sin (-x) = -sin x := by + simp [sin, sub_eq_add_neg, exp_neg, (neg_div _ _).symm, add_mul] + +theorem two_sin : 2 * sin x = (exp (-x * I) - exp (x * I)) * I := + mul_div_cancel₀ _ two_ne_zero + +theorem two_cos : 2 * cos x = exp (x * I) + exp (-x * I) := + mul_div_cancel₀ _ two_ne_zero + +theorem sinh_mul_I : sinh (x * I) = sin x * I := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, + mul_neg_one, neg_sub, neg_mul_eq_neg_mul] + +theorem cosh_mul_I : cosh (x * I) = cos x := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] + +theorem tanh_mul_I : tanh (x * I) = tan x * I := by + rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] + +theorem cos_mul_I : cos (x * I) = cosh x := by rw [← cosh_mul_I]; ring_nf; simp + +theorem sin_mul_I : sin (x * I) = sinh x * I := by + have h : I * sin (x * I) = -sinh x := by + rw [mul_comm, ← sinh_mul_I] + ring_nf + simp + rw [← neg_neg (sinh x), ← h] + apply Complex.ext <;> simp + +theorem tan_mul_I : tan (x * I) = tanh x * I := by + rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] + +theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by + rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, add_mul, add_mul, mul_right_comm, ← sinh_mul_I, + mul_assoc, ← sinh_mul_I, ← cosh_mul_I, ← cosh_mul_I, sinh_add] + +@[simp] +theorem cos_zero : cos 0 = 1 := by simp [cos] + +@[simp] +theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] + +theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by + rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, + mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] + +theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by + simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] + +theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by + simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] + +theorem sin_add_mul_I (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by + rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] + +theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * I := by + convert sin_add_mul_I z.re z.im; exact (re_add_im z).symm + +theorem cos_add_mul_I (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by + rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] + +theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * I := by + convert cos_add_mul_I z.re z.im; exact (re_add_im z).symm + +theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by + have s1 := sin_add ((x + y) / 2) ((x - y) / 2) + have s2 := sin_sub ((x + y) / 2) ((x - y) / 2) + rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 + rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 + rw [s1, s2] + ring + +theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by + have s1 := cos_add ((x + y) / 2) ((x - y) / 2) + have s2 := cos_sub ((x + y) / 2) ((x - y) / 2) + rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel_right, add_self_div_two] at s1 + rw [div_sub_div_same, ← sub_add, add_sub_cancel_left, add_self_div_two] at s2 + rw [s1, s2] + ring + +theorem sin_add_sin : sin x + sin y = 2 * sin ((x + y) / 2) * cos ((x - y) / 2) := by + simpa using sin_sub_sin x (-y) + +theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by + calc + cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := ?_ + _ = + cos ((x + y) / 2) * cos ((x - y) / 2) - sin ((x + y) / 2) * sin ((x - y) / 2) + + (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := + ?_ + _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := ?_ + + · congr <;> field_simp + · rw [cos_add, cos_sub] + ring + +theorem sin_conj : sin (conj x) = conj (sin x) := by + rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← RingHom.map_mul, + sinh_conj, mul_neg, sinh_neg, sinh_mul_I, mul_neg] + +@[simp] +theorem ofReal_sin_ofReal_re (x : ℝ) : ((sin x).re : ℂ) = sin x := + conj_eq_iff_re.1 <| by rw [← sin_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := + ofReal_sin_ofReal_re _ + +@[simp] +theorem sin_ofReal_im (x : ℝ) : (sin x).im = 0 := by rw [← ofReal_sin_ofReal_re, ofReal_im] + +theorem sin_ofReal_re (x : ℝ) : (sin x).re = Real.sin x := + rfl + +theorem cos_conj : cos (conj x) = conj (cos x) := by + rw [← cosh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] + +@[simp] +theorem ofReal_cos_ofReal_re (x : ℝ) : ((cos x).re : ℂ) = cos x := + conj_eq_iff_re.1 <| by rw [← cos_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := + ofReal_cos_ofReal_re _ + +@[simp] +theorem cos_ofReal_im (x : ℝ) : (cos x).im = 0 := by rw [← ofReal_cos_ofReal_re, ofReal_im] + +theorem cos_ofReal_re (x : ℝ) : (cos x).re = Real.cos x := + rfl + +@[simp] +theorem tan_zero : tan 0 = 0 := by simp [tan] + +theorem tan_eq_sin_div_cos : tan x = sin x / cos x := + rfl + +theorem cot_eq_cos_div_sin : cot x = cos x / sin x := + rfl + +theorem tan_mul_cos {x : ℂ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by + rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] + +@[simp] +theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] + +theorem tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj, ← map_div₀, tan] + +theorem cot_conj : cot (conj x) = conj (cot x) := by rw [cot, sin_conj, cos_conj, ← map_div₀, cot] + +@[simp] +theorem ofReal_tan_ofReal_re (x : ℝ) : ((tan x).re : ℂ) = tan x := + conj_eq_iff_re.1 <| by rw [← tan_conj, conj_ofReal] + +@[simp] +theorem ofReal_cot_ofReal_re (x : ℝ) : ((cot x).re : ℂ) = cot x := + conj_eq_iff_re.1 <| by rw [← cot_conj, conj_ofReal] + +@[simp, norm_cast] +theorem ofReal_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := + ofReal_tan_ofReal_re _ + +@[simp, norm_cast] +theorem ofReal_cot (x : ℝ) : (Real.cot x : ℂ) = cot x := + ofReal_cot_ofReal_re _ + +@[simp] +theorem tan_ofReal_im (x : ℝ) : (tan x).im = 0 := by rw [← ofReal_tan_ofReal_re, ofReal_im] + +theorem tan_ofReal_re (x : ℝ) : (tan x).re = Real.tan x := + rfl + +theorem cos_add_sin_I : cos x + sin x * I = exp (x * I) := by + rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] + +theorem cos_sub_sin_I : cos x - sin x * I = exp (-x * I) := by + rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] + +@[simp] +theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + Eq.trans (by rw [cosh_mul_I, sinh_mul_I, mul_pow, I_sq, mul_neg_one, sub_neg_eq_add, add_comm]) + (cosh_sq_sub_sinh_sq (x * I)) + +@[simp] +theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] + +theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by rw [two_mul, cos_add, ← sq, ← sq] + +theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by + rw [cos_two_mul', eq_sub_iff_add_eq.2 (sin_sq_add_cos_sq x), ← sub_add, sub_add_eq_add_sub, + two_mul] + +theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by + rw [two_mul, sin_add, two_mul, add_mul, mul_comm] + +theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := by + simp [cos_two_mul, div_add_div_same, mul_div_cancel_left₀, two_ne_zero, -one_div] + +theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] + +theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_right] + +theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := by + rw [tan_eq_sin_div_cos, div_pow] + field_simp + +theorem tan_sq_div_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : + tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by + simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] + +theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, cos_add x (2 * x)] + simp only [cos_two_mul, sin_two_mul, mul_add, mul_sub, mul_one, sq] + have h2 : 4 * cos x ^ 3 = 2 * cos x * cos x * cos x + 2 * cos x * cos x ^ 2 := by ring + rw [h2, cos_sq'] + ring + +theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, sin_add x (2 * x)] + simp only [cos_two_mul, sin_two_mul, cos_sq'] + have h2 : cos x * (2 * sin x * cos x) = 2 * sin x * cos x ^ 2 := by ring + rw [h2, cos_sq'] + ring + +theorem exp_mul_I : exp (x * I) = cos x + sin x * I := + (cos_add_sin_I _).symm + +theorem exp_add_mul_I : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] + +theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * I) := by + rw [← exp_add_mul_I, re_add_im] + +theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := by + rw [exp_eq_exp_re_mul_sin_add_cos] + simp [exp_ofReal_re, cos_ofReal_re] + +theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by + rw [exp_eq_exp_re_mul_sin_add_cos] + simp [exp_ofReal_re, sin_ofReal_re] + +@[simp] +theorem exp_ofReal_mul_I_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by + simp [exp_mul_I, cos_ofReal_re] + +@[simp] +theorem exp_ofReal_mul_I_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by + simp [exp_mul_I, sin_ofReal_re] + +/-- **De Moivre's formula** -/ +theorem cos_add_sin_mul_I_pow (n : ℕ) (z : ℂ) : + (cos z + sin z * I) ^ n = cos (↑n * z) + sin (↑n * z) * I := by + rw [← exp_mul_I, ← exp_mul_I, ← exp_nat_mul, mul_assoc] + +end Complex + +namespace Real + +open Complex + +variable (x y : ℝ) + +@[simp] +theorem sin_zero : sin 0 = 0 := by simp [sin] + +@[simp] +theorem sin_neg : sin (-x) = -sin x := by simp [sin, exp_neg, (neg_div _ _).symm, add_mul] + +nonrec theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := + ofReal_injective <| by simp [sin_add] + +@[simp] +theorem cos_zero : cos 0 = 1 := by simp [cos] + +@[simp] +theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] + +@[simp] +theorem cos_abs : cos |x| = cos x := by + cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] + +nonrec theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := + ofReal_injective <| by simp [cos_add] + +theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by + simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] + +theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by + simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] + +nonrec theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := + ofReal_injective <| by simp [sin_sub_sin] + +nonrec theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := + ofReal_injective <| by simp [cos_sub_cos] + +nonrec theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := + ofReal_injective <| by simp [cos_add_cos] + +theorem two_mul_sin_mul_sin (x y : ℝ) : 2 * sin x * sin y = cos (x - y) - cos (x + y) := by + simp [cos_add, cos_sub] + ring + +theorem two_mul_cos_mul_cos (x y : ℝ) : 2 * cos x * cos y = cos (x - y) + cos (x + y) := by + simp [cos_add, cos_sub] + ring + +theorem two_mul_sin_mul_cos (x y : ℝ) : 2 * sin x * cos y = sin (x - y) + sin (x + y) := by + simp [sin_add, sin_sub] + ring + +nonrec theorem tan_eq_sin_div_cos : tan x = sin x / cos x := + ofReal_injective <| by simp only [ofReal_tan, tan_eq_sin_div_cos, ofReal_div, ofReal_sin, + ofReal_cos] + +nonrec theorem cot_eq_cos_div_sin : cot x = cos x / sin x := + ofReal_injective <| by simp [cot_eq_cos_div_sin] + +theorem tan_mul_cos {x : ℝ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by + rw [tan_eq_sin_div_cos, div_mul_cancel₀ _ hx] + +@[simp] +theorem tan_zero : tan 0 = 0 := by simp [tan] + +@[simp] +theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] + +@[simp] +nonrec theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + ofReal_injective (by simp [sin_sq_add_cos_sq]) + +@[simp] +theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] + +theorem sin_sq_le_one : sin x ^ 2 ≤ 1 := by + rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_right (sq_nonneg _) + +theorem cos_sq_le_one : cos x ^ 2 ≤ 1 := by + rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_left (sq_nonneg _) + +theorem abs_sin_le_one : |sin x| ≤ 1 := + abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, sin_sq_le_one] + +theorem abs_cos_le_one : |cos x| ≤ 1 := + abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, cos_sq_le_one] + +theorem sin_le_one : sin x ≤ 1 := + (abs_le.1 (abs_sin_le_one _)).2 + +theorem cos_le_one : cos x ≤ 1 := + (abs_le.1 (abs_cos_le_one _)).2 + +theorem neg_one_le_sin : -1 ≤ sin x := + (abs_le.1 (abs_sin_le_one _)).1 + +theorem neg_one_le_cos : -1 ≤ cos x := + (abs_le.1 (abs_cos_le_one _)).1 + +nonrec theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := + ofReal_injective <| by simp [cos_two_mul] + +nonrec theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := + ofReal_injective <| by simp [cos_two_mul'] + +nonrec theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := + ofReal_injective <| by simp [sin_two_mul] + +nonrec theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := + ofReal_injective <| by simp [cos_sq] + +theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel_left] + +theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := + eq_sub_iff_add_eq.2 <| sin_sq_add_cos_sq _ + +lemma sin_sq_eq_half_sub : sin x ^ 2 = 1 / 2 - cos (2 * x) / 2 := by + rw [sin_sq, cos_sq, ← sub_sub, sub_half] + +theorem abs_sin_eq_sqrt_one_sub_cos_sq (x : ℝ) : |sin x| = √(1 - cos x ^ 2) := by + rw [← sin_sq, sqrt_sq_eq_abs] + +theorem abs_cos_eq_sqrt_one_sub_sin_sq (x : ℝ) : |cos x| = √(1 - sin x ^ 2) := by + rw [← cos_sq', sqrt_sq_eq_abs] + +theorem inv_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := + have : Complex.cos x ≠ 0 := mt (congr_arg re) hx + ofReal_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this + +theorem tan_sq_div_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : + tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by + simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] + +theorem inv_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : (√(1 + tan x ^ 2))⁻¹ = cos x := by + rw [← sqrt_sq hx.le, ← sqrt_inv, inv_one_add_tan_sq hx.ne'] + +theorem tan_div_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : + tan x / √(1 + tan x ^ 2) = sin x := by + rw [← tan_mul_cos hx.ne', ← inv_sqrt_one_add_tan_sq hx, div_eq_mul_inv] + +nonrec theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by + rw [← ofReal_inj]; simp [cos_three_mul] + +nonrec theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by + rw [← ofReal_inj]; simp [sin_three_mul] + +/-- The definition of `sinh` in terms of `exp`. -/ +nonrec theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := + ofReal_injective <| by simp [Complex.sinh] + +@[simp] +theorem sinh_zero : sinh 0 = 0 := by simp [sinh] + +@[simp] +theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] + +nonrec theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by + rw [← ofReal_inj]; simp [sinh_add] + +/-- The definition of `cosh` in terms of `exp`. -/ +theorem cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 := + eq_div_of_mul_eq two_ne_zero <| by + rw [cosh, exp, exp, Complex.ofReal_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, + div_mul_cancel₀ _ (two_ne_zero' ℂ), Complex.add_re] + +@[simp] +theorem cosh_zero : cosh 0 = 1 := by simp [cosh] + +@[simp] +theorem cosh_neg : cosh (-x) = cosh x := + ofReal_inj.1 <| by simp + +@[simp] +theorem cosh_abs : cosh |x| = cosh x := by + cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] + +nonrec theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by + rw [← ofReal_inj]; simp [cosh_add] + +theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by + simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] + +theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by + simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] + +nonrec theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + ofReal_inj.1 <| by simp [tanh_eq_sinh_div_cosh] + +@[simp] +theorem tanh_zero : tanh 0 = 0 := by simp [tanh] + +@[simp] +theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] + +@[simp] +theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← ofReal_inj]; simp + +@[simp] +theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] + +@[simp] +theorem exp_sub_cosh : exp x - cosh x = sinh x := + sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm + +@[simp] +theorem exp_sub_sinh : exp x - sinh x = cosh x := + sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm + +@[simp] +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by + rw [← ofReal_inj] + simp + +@[simp] +theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] + +@[simp] +theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← ofReal_inj]; simp + +nonrec theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← ofReal_inj]; simp [cosh_sq] + +theorem cosh_sq' : cosh x ^ 2 = 1 + sinh x ^ 2 := + (cosh_sq x).trans (add_comm _ _) + +nonrec theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← ofReal_inj]; simp [sinh_sq] + +nonrec theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by + rw [← ofReal_inj]; simp [cosh_two_mul] + +nonrec theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by + rw [← ofReal_inj]; simp [sinh_two_mul] + +nonrec theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by + rw [← ofReal_inj]; simp [cosh_three_mul] + +nonrec theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by + rw [← ofReal_inj]; simp [sinh_three_mul] + +open IsAbsoluteValue Nat + +private theorem add_one_lt_exp_of_pos {x : ℝ} (hx : 0 < x) : x + 1 < exp x := + (by nlinarith : x + 1 < 1 + x + x ^ 2 / 2).trans_le (quadratic_le_exp_of_nonneg hx.le) + +private theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := by + rcases eq_or_lt_of_le hx with (rfl | h) + · simp + exact (add_one_lt_exp_of_pos h).le + +/-- `Real.cosh` is always positive -/ +theorem cosh_pos (x : ℝ) : 0 < Real.cosh x := + (cosh_eq x).symm ▸ half_pos (add_pos (exp_pos x) (exp_pos (-x))) + +theorem sinh_lt_cosh : sinh x < cosh x := + lt_of_pow_lt_pow_left₀ 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ + +end Real + +namespace Real + +open Complex Finset + +theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := + calc + |cos x - (1 - x ^ 2 / 2)| = Complex.abs (Complex.cos x - (1 - (x : ℂ) ^ 2 / 2)) := by + rw [← abs_ofReal]; simp + _ = Complex.abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - (x : ℂ) ^ 2)) / 2) := by + simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] + _ = abs + (((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) + + (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial)) / 2) := + (congr_arg Complex.abs + (congr_arg (fun x : ℂ => x / 2) + (by + simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, range_zero, sum_empty, + Nat.factorial, Nat.cast_one, ne_eq, one_ne_zero, not_false_eq_true, div_self, + zero_add, div_one, Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat, mul_neg, + neg_neg] + apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring_nf + ))) + _ ≤ abs ((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2) + + abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2) := by + rw [add_div]; exact Complex.abs.add_le _ _ + _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + + abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by + simp [map_div₀] + _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 := by + gcongr + · exact Complex.exp_bound (by simpa) (by decide) + · exact Complex.exp_bound (by simpa) (by decide) + _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] + +theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := + calc + |sin x - (x - x ^ 3 / 6)| = Complex.abs (Complex.sin x - (x - x ^ 3 / 6 : ℝ)) := by + rw [← abs_ofReal]; simp + _ = Complex.abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - + (2 * x - x ^ 3 / 3 : ℝ)) / 2) := by + simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left₀ _ (two_ne_zero' ℂ), + div_div, show (3 : ℂ) * 2 = 6 by norm_num] + _ = Complex.abs (((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) - + (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial)) * I / 2) := + (congr_arg Complex.abs + (congr_arg (fun x : ℂ => x / 2) + (by + simp only [sum_range_succ, neg_mul, pow_succ, pow_zero, mul_one, ofReal_sub, ofReal_mul, + ofReal_ofNat, ofReal_div, range_zero, sum_empty, Nat.factorial, Nat.cast_one, ne_eq, + one_ne_zero, not_false_eq_true, div_self, zero_add, div_one, mul_neg, neg_neg, + Nat.mul_one, Nat.cast_succ, Nat.cast_mul, Nat.cast_ofNat] + apply Complex.ext <;> simp [div_eq_mul_inv, normSq]; ring))) + _ ≤ abs ((Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) * I / 2) + + abs (-((Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) * I) / 2) := by + rw [sub_mul, sub_eq_add_neg, add_div]; exact Complex.abs.add_le _ _ + _ = abs (Complex.exp (x * I) - ∑ m ∈ range 4, (x * I) ^ m / m.factorial) / 2 + + abs (Complex.exp (-x * I) - ∑ m ∈ range 4, (-x * I) ^ m / m.factorial) / 2 := by + simp [add_comm, map_div₀] + _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 := by + gcongr + · exact Complex.exp_bound (by simpa) (by decide) + · exact Complex.exp_bound (by simpa) (by decide) + _ ≤ |x| ^ 4 * (5 / 96) := by norm_num [Nat.factorial] + +theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := + calc 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| + lt_sub_iff_add_lt.2 + (calc + |x| ^ 4 * (5 / 96) + x ^ 2 / 2 ≤ 1 * (5 / 96) + 1 / 2 := by + gcongr + · exact pow_le_one₀ (abs_nonneg _) hx + · rw [sq, ← abs_mul_self, abs_mul] + exact mul_le_one₀ hx (abs_nonneg _) hx + _ < 1 := by norm_num) + _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 + +theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := + calc 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| lt_sub_iff_add_lt.2 + (calc + |x| ^ 4 * (5 / 96) + x ^ 3 / 6 ≤ x * (5 / 96) + x / 6 := by + gcongr + · calc + |x| ^ 4 ≤ |x| ^ 1 := + pow_le_pow_of_le_one (abs_nonneg _) + (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) + _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] + · calc + x ^ 3 ≤ x ^ 1 := pow_le_pow_of_le_one (le_of_lt hx0) hx (by decide) + _ = x := pow_one _ + _ < x := by linarith) + _ ≤ sin x := + sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 + +theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := + have : x / 2 ≤ 1 := (div_le_iff₀ (by norm_num)).mpr (by simpa) + calc + 0 < 2 * sin (x / 2) * cos (x / 2) := + mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) + (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) + _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] + +theorem cos_one_le : cos 1 ≤ 2 / 3 := + calc + cos 1 ≤ |(1 : ℝ)| ^ 4 * (5 / 96) + (1 - 1 ^ 2 / 2) := + sub_le_iff_le_add.1 (abs_sub_le_iff.1 (cos_bound (by simp))).1 + _ ≤ 2 / 3 := by norm_num + +theorem cos_one_pos : 0 < cos 1 := + cos_pos_of_le_one (le_of_eq abs_one) + +theorem cos_two_neg : cos 2 < 0 := + calc cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm + _ = _ := Real.cos_two_mul 1 + _ ≤ 2 * (2 / 3) ^ 2 - 1 := by + gcongr + · exact cos_one_pos.le + · apply cos_one_le + _ < 0 := by norm_num + +end Real + +namespace Mathlib.Meta.Positivity +open Lean.Meta Qq + +/-- Extension for the `positivity` tactic: `Real.cosh` is always positive. -/ +@[positivity Real.cosh _] +def evalCosh : PositivityExt where eval {u α} _ _ e := do + match u, α, e with + | 0, ~q(ℝ), ~q(Real.cosh $a) => + assertInstancesCommute + return .positive q(Real.cosh_pos $a) + | _, _, _ => throwError "not Real.cosh" + +example (x : ℝ) : 0 < x.cosh := by positivity + +end Mathlib.Meta.Positivity + +namespace Complex + +@[simp] +theorem abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := by + have := Real.sin_sq_add_cos_sq x + simp_all [add_comm, abs, normSq, sq, sin_ofReal_re, cos_ofReal_re, mul_re] + +@[simp] +theorem abs_exp_ofReal_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by + rw [exp_mul_I, abs_cos_add_sin_mul_I] + +theorem abs_exp (z : ℂ) : abs (exp z) = Real.exp z.re := by + rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_ofReal, abs_cos_add_sin_mul_I, mul_one] + +theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by + rw [abs_exp, abs_exp, Real.exp_eq_exp] + +end Complex diff --git a/Mathlib/Data/DFinsupp/BigOperators.lean b/Mathlib/Data/DFinsupp/BigOperators.lean index 20adad039c4ac..32fe109404721 100644 --- a/Mathlib/Data/DFinsupp/BigOperators.lean +++ b/Mathlib/Data/DFinsupp/BigOperators.lean @@ -307,7 +307,6 @@ def liftAddHom [∀ i, AddZeroClass (β i)] [AddCommMonoid γ] : -- This applies to roughly the remainder of the file. /-- The `DFinsupp` version of `Finsupp.liftAddHom_singleAddHom`,-/ -@[simp, nolint simpNF] -- simpNF correctly complains that the LHS simplifies already theorem liftAddHom_singleAddHom [∀ i, AddCommMonoid (β i)] : liftAddHom (β := β) (singleAddHom β) = AddMonoidHom.id (Π₀ i, β i) := (liftAddHom (β := β)).toEquiv.apply_eq_iff_eq_symm_apply.2 rfl diff --git a/Mathlib/Data/DFinsupp/Defs.lean b/Mathlib/Data/DFinsupp/Defs.lean index a8b854e30a101..1715e3d3fb437 100644 --- a/Mathlib/Data/DFinsupp/Defs.lean +++ b/Mathlib/Data/DFinsupp/Defs.lean @@ -988,7 +988,7 @@ theorem subtypeDomain_def (f : Π₀ i, β i) : f.subtypeDomain p = mk (f.support.subtype p) fun i => f i := by ext i; by_cases h2 : f i ≠ 0 <;> try simp at h2; dsimp; simp [h2] -@[simp, nolint simpNF] -- Porting note: simpNF claims that LHS does not simplify, but it does +@[simp] theorem support_subtypeDomain {f : Π₀ i, β i} : (subtypeDomain p f).support = f.support.subtype p := by ext i diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 4ce8bf4d46529..c09ecafc6dfa1 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -94,7 +94,7 @@ variable {α : Type*} /-- The extended nonnegative real numbers. This is usually denoted [0, ∞], and is relevant as the codomain of a measure. -/ def ENNReal := WithTop ℝ≥0 - deriving Zero, AddCommMonoidWithOne, SemilatticeSup, DistribLattice, Nontrivial + deriving Zero, Top, AddCommMonoidWithOne, SemilatticeSup, DistribLattice, Nontrivial @[inherit_doc] scoped[ENNReal] notation "ℝ≥0∞" => ENNReal @@ -105,6 +105,7 @@ scoped[ENNReal] notation "∞" => (⊤ : ENNReal) namespace ENNReal instance : OrderBot ℝ≥0∞ := inferInstanceAs (OrderBot (WithTop ℝ≥0)) +instance : OrderTop ℝ≥0∞ := inferInstanceAs (OrderTop (WithTop ℝ≥0)) instance : BoundedOrder ℝ≥0∞ := inferInstanceAs (BoundedOrder (WithTop ℝ≥0)) instance : CharZero ℝ≥0∞ := inferInstanceAs (CharZero (WithTop ℝ≥0)) instance : Min ℝ≥0∞ := SemilatticeInf.toMin @@ -127,8 +128,8 @@ instance : DenselyOrdered ℝ≥0∞ := inferInstanceAs (DenselyOrdered (WithTop noncomputable instance : LinearOrderedAddCommMonoid ℝ≥0∞ := inferInstanceAs (LinearOrderedAddCommMonoid (WithTop ℝ≥0)) -noncomputable instance instSub : Sub ℝ≥0∞ := inferInstanceAs (Sub (WithTop ℝ≥0)) -noncomputable instance : OrderedSub ℝ≥0∞ := inferInstanceAs (OrderedSub (WithTop ℝ≥0)) +instance instSub : Sub ℝ≥0∞ := inferInstanceAs (Sub (WithTop ℝ≥0)) +instance : OrderedSub ℝ≥0∞ := inferInstanceAs (OrderedSub (WithTop ℝ≥0)) noncomputable instance : LinearOrderedAddCommMonoidWithTop ℝ≥0∞ := inferInstanceAs (LinearOrderedAddCommMonoidWithTop (WithTop ℝ≥0)) @@ -152,7 +153,7 @@ noncomputable instance : LinearOrderedCommMonoidWithZero ℝ≥0∞ := mul_le_mul_left := fun _ _ => mul_le_mul_left' zero_le_one := zero_le 1 } -noncomputable instance : Unique (AddUnits ℝ≥0∞) where +instance : Unique (AddUnits ℝ≥0∞) where default := 0 uniq a := AddUnits.ext <| le_zero_iff.1 <| by rw [← a.add_neg]; exact le_self_add @@ -185,6 +186,12 @@ lemma coe_ne_coe : (p : ℝ≥0∞) ≠ q ↔ p ≠ q := coe_inj.not theorem range_coe' : range ofNNReal = Iio ∞ := WithTop.range_coe theorem range_coe : range ofNNReal = {∞}ᶜ := (isCompl_range_some_none ℝ≥0).symm.compl_eq.symm +instance : NNRatCast ℝ≥0∞ where + nnratCast r := ofNNReal r + +@[norm_cast] +theorem coe_nnratCast (q : ℚ≥0) : ↑(q : ℝ≥0) = (q : ℝ≥0∞) := rfl + /-- `toNNReal x` returns `x` if it is real, otherwise 0. -/ protected def toNNReal : ℝ≥0∞ → ℝ≥0 := WithTop.untop' 0 @@ -192,7 +199,7 @@ protected def toNNReal : ℝ≥0∞ → ℝ≥0 := WithTop.untop' 0 protected def toReal (a : ℝ≥0∞) : Real := a.toNNReal /-- `ofReal x` returns `x` if it is nonnegative, `0` otherwise. -/ -protected noncomputable def ofReal (r : Real) : ℝ≥0∞ := r.toNNReal +protected def ofReal (r : Real) : ℝ≥0∞ := r.toNNReal @[simp, norm_cast] lemma toNNReal_coe (r : ℝ≥0) : (r : ℝ≥0∞).toNNReal = r := rfl @@ -686,6 +693,12 @@ end OrdConnected end Set +/-- While not very useful, this instance uses the same representation as `Real.instRepr`. -/ +unsafe instance : Repr ℝ≥0∞ where + reprPrec + | (r : ℝ≥0), p => Repr.addAppParen f!"ENNReal.ofReal ({repr r.val})" p + | ∞, _ => "∞" + namespace Mathlib.Meta.Positivity open Lean Meta Qq diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 9b9e842c4e348..63249a8693464 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -44,7 +44,8 @@ protected theorem div_eq_inv_mul : a / b = b⁻¹ * a := by rw [div_eq_mul_inv, show sInf { b : ℝ≥0∞ | 1 ≤ 0 * b } = ∞ by simp @[simp] theorem inv_top : ∞⁻¹ = 0 := - bot_unique <| le_of_forall_le_of_dense fun a (h : 0 < a) => sInf_le <| by simp [*, h.ne', top_mul] + bot_unique <| le_of_forall_gt_imp_ge_of_dense fun a (h : 0 < a) => sInf_le <| by + simp [*, h.ne', top_mul] theorem coe_inv_le : (↑r⁻¹ : ℝ≥0∞) ≤ (↑r)⁻¹ := le_sInf fun b (hb : 1 ≤ ↑r * b) => @@ -444,7 +445,7 @@ instance : SMulPosMono ℝ≥0 ℝ≥0∞ where elim _r _ _a _b hab := mul_le_mul_right' (coe_le_coe.2 hab) _ theorem le_of_forall_nnreal_lt {x y : ℝ≥0∞} (h : ∀ r : ℝ≥0, ↑r < x → ↑r ≤ y) : x ≤ y := by - refine le_of_forall_ge_of_dense fun r hr => ?_ + refine le_of_forall_lt_imp_le_of_dense fun r hr => ?_ lift r to ℝ≥0 using ne_top_of_lt hr exact h r hr @@ -533,7 +534,7 @@ private lemma exists_lt_mul_right {a b c : ℝ≥0∞} (hc : c < a * b) : ∃ b' simp_rw [mul_comm a] at hc ⊢; exact exists_lt_mul_left hc lemma mul_le_of_forall_lt {a b c : ℝ≥0∞} (h : ∀ a' < a, ∀ b' < b, a' * b' ≤ c) : a * b ≤ c := by - refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_lt_mul_left hd obtain ⟨b', hb', hd⟩ := exists_lt_mul_right hd exact le_trans hd.le <| h _ ha' _ hb' diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 24b26caf10c98..df6744aa82316 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -68,12 +68,18 @@ theorem mul_left_strictMono (h0 : a ≠ 0) (hinf : a ≠ ∞) : StrictMono (a * mul_comm b a ▸ mul_comm c a ▸ ENNReal.mul_left_strictMono h0 hinf bc -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: generalize to `WithTop` -theorem mul_eq_mul_left (h0 : a ≠ 0) (hinf : a ≠ ∞) : a * b = a * c ↔ b = c := +protected theorem mul_right_inj (h0 : a ≠ 0) (hinf : a ≠ ∞) : a * b = a * c ↔ b = c := (mul_left_strictMono h0 hinf).injective.eq_iff +@[deprecated (since := "2025-01-20")] +alias mul_eq_mul_left := ENNReal.mul_right_inj + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: generalize to `WithTop` -theorem mul_eq_mul_right : c ≠ 0 → c ≠ ∞ → (a * c = b * c ↔ a = b) := - mul_comm c a ▸ mul_comm c b ▸ mul_eq_mul_left +protected theorem mul_left_inj (h0 : c ≠ 0) (hinf : c ≠ ∞) : a * c = b * c ↔ a = b := + mul_comm c a ▸ mul_comm c b ▸ ENNReal.mul_right_inj h0 hinf + +@[deprecated (since := "2025-01-20")] +alias mul_eq_mul_right := ENNReal.mul_left_inj -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: generalize to `WithTop` theorem mul_le_mul_left (h0 : a ≠ 0) (hinf : a ≠ ∞) : a * b ≤ a * c ↔ b ≤ c := @@ -92,10 +98,10 @@ theorem mul_lt_mul_right : c ≠ 0 → c ≠ ∞ → (a * c < b * c ↔ a < b) : mul_comm c a ▸ mul_comm c b ▸ mul_lt_mul_left protected lemma mul_eq_left (ha₀ : a ≠ 0) (ha : a ≠ ∞) : a * b = a ↔ b = 1 := by - simpa using ENNReal.mul_eq_mul_left ha₀ ha (c := 1) + simpa using ENNReal.mul_right_inj ha₀ ha (c := 1) protected lemma mul_eq_right (hb₀ : b ≠ 0) (hb : b ≠ ∞) : a * b = b ↔ a = 1 := by - simpa using ENNReal.mul_eq_mul_right hb₀ hb (b := 1) + simpa using ENNReal.mul_left_inj hb₀ hb (b := 1) end Mul diff --git a/Mathlib/Data/Erased.lean b/Mathlib/Data/Erased.lean index 6460d80645efa..b1b6a212314a3 100644 --- a/Mathlib/Data/Erased.lean +++ b/Mathlib/Data/Erased.lean @@ -21,7 +21,7 @@ universe u and proofs. This can be used to track data without storing it literally. -/ def Erased (α : Sort u) : Sort max 1 u := - Σ's : α → Prop, ∃ a, (fun b => a = b) = s + { s : α → Prop // ∃ a, (a = ·) = s } namespace Erased diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index c4f67ec7c5759..7dc978fddb102 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -58,10 +58,6 @@ This file expands on the development in the core library. * `Fin.divNat i` : divides `i : Fin (m * n)` by `n`; * `Fin.modNat i` : takes the mod of `i : Fin (m * n)` by `n`; -### Misc definitions - -* `Fin.revPerm : Equiv.Perm (Fin n)` : `Fin.rev` as an `Equiv.Perm`, the antitone involution given - by `i ↦ n-(i+1)` -/ @@ -267,52 +263,6 @@ theorem pos_iff_ne_zero' [NeZero n] (a : Fin n) : 0 < a ↔ a ≠ 0 := by lemma cast_injective {k l : ℕ} (h : k = l) : Injective (Fin.cast h) := fun a b hab ↦ by simpa [← val_eq_val] using hab -theorem rev_involutive : Involutive (rev : Fin n → Fin n) := rev_rev - -/-- `Fin.rev` as an `Equiv.Perm`, the antitone involution `Fin n → Fin n` given by -`i ↦ n-(i+1)`. -/ -@[simps! apply symm_apply] -def revPerm : Equiv.Perm (Fin n) := - Involutive.toPerm rev rev_involutive - -theorem rev_injective : Injective (@rev n) := - rev_involutive.injective - -theorem rev_surjective : Surjective (@rev n) := - rev_involutive.surjective - -theorem rev_bijective : Bijective (@rev n) := - rev_involutive.bijective - -@[simp] -theorem revPerm_symm : (@revPerm n).symm = revPerm := - rfl - -theorem cast_rev (i : Fin n) (h : n = m) : - 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 - rw [← rev_inj, rev_rev] - -theorem rev_ne_iff {i j : Fin n} : rev i ≠ j ↔ i ≠ rev j := rev_eq_iff.not - -theorem rev_lt_iff {i j : Fin n} : rev i < j ↔ rev j < i := by - rw [← rev_lt_rev, rev_rev] - -theorem rev_le_iff {i j : Fin n} : rev i ≤ j ↔ rev j ≤ i := by - rw [← rev_le_rev, rev_rev] - -theorem lt_rev_iff {i j : Fin n} : i < rev j ↔ j < rev i := by - rw [← rev_lt_rev, rev_rev] - -theorem le_rev_iff {i j : Fin n} : i ≤ rev j ↔ j ≤ rev i := by - rw [← rev_le_rev, rev_rev] - --- Porting note: this is now syntactically equal to `val_last` - -@[simp] theorem val_rev_zero [NeZero n] : ((rev 0 : Fin n) : ℕ) = n.pred := rfl - theorem last_pos' [NeZero n] : 0 < last n := n.pos_of_neZero theorem one_lt_last [NeZero n] : 1 < last (n + 1) := by @@ -908,15 +858,6 @@ theorem castPred_one [NeZero n] (h := Fin.ext_iff.not.2 one_lt_last.ne) : · exact subsingleton_one.elim _ 1 · rfl -theorem rev_pred {i : Fin (n + 1)} (h : i ≠ 0) (h' := rev_ne_iff.mpr ((rev_last _).symm ▸ h)) : - rev (pred i h) = castPred (rev i) h' := by - rw [← castSucc_inj, castSucc_castPred, ← rev_succ, succ_pred] - -theorem rev_castPred {i : Fin (n + 1)} - (h : i ≠ last n) (h' := rev_ne_iff.mpr ((rev_zero _).symm ▸ h)) : - rev (castPred i h) = pred (rev i) h' := by - rw [← succ_inj, succ_pred, ← rev_castSucc, castSucc_castPred] - theorem succ_castPred_eq_castPred_succ {a : Fin (n + 1)} (ha : a ≠ last n) (ha' := a.succ_ne_last_iff.mpr ha) : (a.castPred ha).succ = (succ a).castPred ha' := rfl @@ -1026,17 +967,6 @@ lemma succAbove_castPred_of_le (p i : Fin (n + 1)) (h : p ≤ i) (hi : i ≠ las lemma succAbove_castPred_self (p : Fin (n + 1)) (h : p ≠ last n) : succAbove p (p.castPred h) = (p.castPred h).succ := succAbove_castPred_of_le _ _ Fin.le_rfl h -lemma succAbove_rev_left (p : Fin (n + 1)) (i : Fin n) : - p.rev.succAbove i = (p.succAbove i.rev).rev := by - obtain h | h := (rev p).succ_le_or_le_castSucc i - · rw [succAbove_of_succ_le _ _ h, - succAbove_of_le_castSucc _ _ (rev_succ _ ▸ (le_rev_iff.mpr h)), rev_succ, rev_rev] - · rw [succAbove_of_le_castSucc _ _ h, - succAbove_of_succ_le _ _ (rev_castSucc _ ▸ (rev_le_iff.mpr h)), rev_castSucc, rev_rev] - -lemma succAbove_rev_right (p : Fin (n + 1)) (i : Fin n) : - p.succAbove i.rev = (p.rev.succAbove i).rev := by rw [succAbove_rev_left, rev_rev] - /-- Embedding `i : Fin n` into `Fin (n + 1)` with a hole around `p : Fin (n + 1)` never results in `p` itself -/ lemma succAbove_ne (p : Fin (n + 1)) (i : Fin n) : p.succAbove i ≠ p := by @@ -1204,11 +1134,6 @@ lemma castPred_succAbove_castPred {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a simp_rw [← castSucc_inj (b := (a.succAbove b).castPred hk), ← castSucc_succAbove_castSucc, castSucc_castPred] -/-- `rev` commutes with `succAbove`. -/ -lemma rev_succAbove (p : Fin (n + 1)) (i : Fin n) : - rev (succAbove p i) = succAbove (rev p) (rev i) := by - rw [succAbove_rev_left, rev_rev] - --@[simp] -- Porting note: can be proved by `simp` lemma one_succAbove_zero {n : ℕ} : (1 : Fin (n + 2)).succAbove 0 = 0 := by rfl @@ -1295,17 +1220,6 @@ lemma predAbove_castPred_of_le (p i : Fin (n + 1)) (h : i ≤ p) (hp : p ≠ las lemma predAbove_castPred_self (p : Fin (n + 1)) (hp : p ≠ last n) : (castPred p hp).predAbove p = castPred p hp := predAbove_castPred_of_le _ _ Fin.le_rfl hp -lemma predAbove_rev_left (p : Fin n) (i : Fin (n + 1)) : - p.rev.predAbove i = (p.predAbove i.rev).rev := by - obtain h | h := (rev i).succ_le_or_le_castSucc p - · rw [predAbove_of_succ_le _ _ h, rev_pred, - predAbove_of_le_castSucc _ _ (rev_succ _ ▸ (le_rev_iff.mpr h)), castPred_inj, rev_rev] - · rw [predAbove_of_le_castSucc _ _ h, rev_castPred, - predAbove_of_succ_le _ _ (rev_castSucc _ ▸ (rev_le_iff.mpr h)), pred_inj, rev_rev] - -lemma predAbove_rev_right (p : Fin n) (i : Fin (n + 1)) : - p.predAbove i.rev = (p.rev.predAbove i).rev := by rw [predAbove_rev_left, rev_rev] - @[simp] lemma predAbove_right_zero [NeZero n] {i : Fin n} : predAbove (i : Fin n) 0 = 0 := by cases n · exact i.elim0 @@ -1379,10 +1293,6 @@ lemma predAbove_succAbove (p : Fin n) (i : Fin n) : p.predAbove ((castSucc p).su castSucc_pred_eq_pred_castSucc] · rw [predAbove_of_le_castSucc _ _ h, predAbove_castSucc_of_le _ _ h, castSucc_castPred] -/-- `rev` commutes with `predAbove`. -/ -lemma rev_predAbove {n : ℕ} (p : Fin n) (i : Fin (n + 1)) : - (predAbove p i).rev = predAbove p.rev i.rev := by rw [predAbove_rev_left, rev_rev] - end PredAbove section DivMod @@ -1531,5 +1441,3 @@ protected theorem zero_mul' [NeZero n] (k : Fin n) : (0 : Fin n) * k = 0 := by end Mul end Fin - -set_option linter.style.longFile 1700 diff --git a/Mathlib/Data/Fin/Rev.lean b/Mathlib/Data/Fin/Rev.lean new file mode 100644 index 0000000000000..929b6cdbd6fde --- /dev/null +++ b/Mathlib/Data/Fin/Rev.lean @@ -0,0 +1,112 @@ +/- +Copyright (c) 2022 Eric Rodriguez. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Rodriguez, Joel Riou, Yury Kudryashov +-/ +import Mathlib.Data.Fin.Basic +/-! +# Reverse on `Fin n` + +This file contains lemmas about `Fin.rev : Fin n → Fin n` which maps `i` to `n - 1 - i`. + +## Definitions + +* `Fin.revPerm : Equiv.Perm (Fin n)` : `Fin.rev` as an `Equiv.Perm`, the antitone involution given + by `i ↦ n-(i+1)` +-/ + +assert_not_exists Monoid Fintype + +open Fin Nat Function + +namespace Fin + +variable {n m : ℕ} + +theorem rev_involutive : Involutive (rev : Fin n → Fin n) := rev_rev + +/-- `Fin.rev` as an `Equiv.Perm`, the antitone involution `Fin n → Fin n` given by +`i ↦ n-(i+1)`. -/ +@[simps! apply symm_apply] +def revPerm : Equiv.Perm (Fin n) := + Involutive.toPerm rev rev_involutive + +theorem rev_injective : Injective (@rev n) := + rev_involutive.injective + +theorem rev_surjective : Surjective (@rev n) := + rev_involutive.surjective + +theorem rev_bijective : Bijective (@rev n) := + rev_involutive.bijective + +@[simp] +theorem revPerm_symm : (@revPerm n).symm = revPerm := + rfl + +theorem cast_rev (i : Fin n) (h : n = m) : + 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 + rw [← rev_inj, rev_rev] + +theorem rev_ne_iff {i j : Fin n} : rev i ≠ j ↔ i ≠ rev j := rev_eq_iff.not + +theorem rev_lt_iff {i j : Fin n} : rev i < j ↔ rev j < i := by + rw [← rev_lt_rev, rev_rev] + +theorem rev_le_iff {i j : Fin n} : rev i ≤ j ↔ rev j ≤ i := by + rw [← rev_le_rev, rev_rev] + +theorem lt_rev_iff {i j : Fin n} : i < rev j ↔ j < rev i := by + rw [← rev_lt_rev, rev_rev] + +theorem le_rev_iff {i j : Fin n} : i ≤ rev j ↔ j ≤ rev i := by + rw [← rev_le_rev, rev_rev] + +-- Porting note: this is now syntactically equal to `val_last` + +@[simp] theorem val_rev_zero [NeZero n] : ((rev 0 : Fin n) : ℕ) = n.pred := rfl + +theorem rev_pred {i : Fin (n + 1)} (h : i ≠ 0) (h' := rev_ne_iff.mpr ((rev_last _).symm ▸ h)) : + rev (pred i h) = castPred (rev i) h' := by + rw [← castSucc_inj, castSucc_castPred, ← rev_succ, succ_pred] + +theorem rev_castPred {i : Fin (n + 1)} + (h : i ≠ last n) (h' := rev_ne_iff.mpr ((rev_zero _).symm ▸ h)) : + rev (castPred i h) = pred (rev i) h' := by + rw [← succ_inj, succ_pred, ← rev_castSucc, castSucc_castPred] + +lemma succAbove_rev_left (p : Fin (n + 1)) (i : Fin n) : + p.rev.succAbove i = (p.succAbove i.rev).rev := by + obtain h | h := (rev p).succ_le_or_le_castSucc i + · rw [succAbove_of_succ_le _ _ h, + succAbove_of_le_castSucc _ _ (rev_succ _ ▸ (le_rev_iff.mpr h)), rev_succ, rev_rev] + · rw [succAbove_of_le_castSucc _ _ h, + succAbove_of_succ_le _ _ (rev_castSucc _ ▸ (rev_le_iff.mpr h)), rev_castSucc, rev_rev] + +lemma succAbove_rev_right (p : Fin (n + 1)) (i : Fin n) : + p.succAbove i.rev = (p.rev.succAbove i).rev := by rw [succAbove_rev_left, rev_rev] + +/-- `rev` commutes with `succAbove`. -/ +lemma rev_succAbove (p : Fin (n + 1)) (i : Fin n) : + rev (succAbove p i) = succAbove (rev p) (rev i) := by + rw [succAbove_rev_left, rev_rev] + +lemma predAbove_rev_left (p : Fin n) (i : Fin (n + 1)) : + p.rev.predAbove i = (p.predAbove i.rev).rev := by + obtain h | h := (rev i).succ_le_or_le_castSucc p + · rw [predAbove_of_succ_le _ _ h, rev_pred, + predAbove_of_le_castSucc _ _ (rev_succ _ ▸ (le_rev_iff.mpr h)), castPred_inj, rev_rev] + · rw [predAbove_of_le_castSucc _ _ h, rev_castPred, + predAbove_of_succ_le _ _ (rev_castSucc _ ▸ (rev_le_iff.mpr h)), pred_inj, rev_rev] + +lemma predAbove_rev_right (p : Fin n) (i : Fin (n + 1)) : + p.predAbove i.rev = (p.rev.predAbove i).rev := by rw [predAbove_rev_left, rev_rev] + +/-- `rev` commutes with `predAbove`. -/ +lemma rev_predAbove {n : ℕ} (p : Fin n) (i : Fin (n + 1)) : + (predAbove p i).rev = predAbove p.rev i.rev := by rw [predAbove_rev_left, rev_rev] + +end Fin diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 3a31ba890f95d..49034b164fdbd 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yury Kudryashov, Sébastien Gouëzel, Chris Hughes -/ -import Mathlib.Data.Fin.Basic +import Mathlib.Data.Fin.Rev import Mathlib.Data.Nat.Find /-! @@ -142,7 +142,7 @@ theorem cons_injective2 : Function.Injective2 (@cons n α) := fun x₀ y₀ x y ⟨congr_fun h 0, funext fun i ↦ by simpa using congr_fun h (Fin.succ i)⟩ @[simp] -theorem cons_eq_cons {x₀ y₀ : α 0} {x y : ∀ i : Fin n, α i.succ} : +theorem cons_inj {x₀ y₀ : α 0} {x y : ∀ i : Fin n, α i.succ} : cons x₀ x = cons y₀ y ↔ x₀ = y₀ ∧ x = y := cons_injective2.eq_iff @@ -165,7 +165,7 @@ theorem update_cons_zero : update (cons x p) 0 z = cons z p := by rw [← this, cons_succ, cons_succ] /-- Concatenating the first element of a tuple with its tail gives back the original tuple -/ -@[simp, nolint simpNF] -- Porting note: linter claims LHS doesn't simplify +@[simp] theorem cons_self_tail : cons (q 0) (tail q) = q := by ext j by_cases h : j = 0 @@ -586,6 +586,22 @@ theorem update_snoc_last : update (snoc p x) (last n) z = snoc p z := by · rw [eq_last_of_not_lt h] simp +/-- As a binary function, `Fin.snoc` is injective. -/ +theorem snoc_injective2 : Function.Injective2 (@snoc n α) := fun x y xₙ yₙ h ↦ + ⟨funext fun i ↦ by simpa using congr_fun h (castSucc i), by simpa using congr_fun h (last n)⟩ + +@[simp] +theorem snoc_inj {x y : ∀ i : Fin n, α i.castSucc} {xₙ yₙ : α (last n)} : + snoc x xₙ = snoc y yₙ ↔ x = y ∧ xₙ = yₙ := + snoc_injective2.eq_iff + +theorem snoc_right_injective (x : ∀ i : Fin n, α i.castSucc) : + Function.Injective (snoc x) := + snoc_injective2.right _ + +theorem snoc_left_injective (xₙ : α (last n)) : Function.Injective (snoc · xₙ) := + snoc_injective2.left _ + /-- Concatenating the first element of a tuple with its tail gives back the original tuple -/ @[simp] theorem snoc_init_self : snoc (init q) (q (last n)) = q := by @@ -791,6 +807,10 @@ lemma exists_iff_succAbove {P : Fin (n + 1) → Prop} (p : Fin (n + 1)) : · exact .inr ⟨_, hi⟩ mpr := by rintro (h | ⟨i, hi⟩) <;> exact ⟨_, ‹_›⟩ +/-- Analogue of `Fin.eq_zero_or_eq_succ` for `succAbove`. -/ +theorem eq_self_or_eq_succAbove (p i : Fin (n + 1)) : i = p ∨ ∃ j, i = p.succAbove j := + succAboveCases p (.inl rfl) (fun j => .inr ⟨j, rfl⟩) i + /-- Remove the `p`-th entry of a tuple. -/ def removeNth (p : Fin (n + 1)) (f : ∀ i, α i) : ∀ i, α (p.succAbove i) := fun i ↦ f (p.succAbove i) @@ -847,6 +867,24 @@ theorem eq_insertNth_iff {p : Fin (n + 1)} {a : α p} {f : ∀ i, α (p.succAbov g = insertNth p a f ↔ g p = a ∧ removeNth p g = f := by simpa [eq_comm] using insertNth_eq_iff +/-- As a binary function, `Fin.insertNth` is injective. -/ +theorem insertNth_injective2 {p : Fin (n + 1)} : + Function.Injective2 (@insertNth n α p) := fun xₚ yₚ x y h ↦ + ⟨by simpa using congr_fun h p, funext fun i ↦ by simpa using congr_fun h (succAbove p i)⟩ + +@[simp] +theorem insertNth_inj {p : Fin (n + 1)} {x y : ∀ i, α (succAbove p i)} {xₚ yₚ : α p} : + insertNth p xₚ x = insertNth p yₚ y ↔ xₚ = yₚ ∧ x = y := + insertNth_injective2.eq_iff + +theorem insertNth_left_injective {p : Fin (n + 1)} (x : ∀ i, α (succAbove p i)) : + Function.Injective (insertNth p · x) := + insertNth_injective2.left _ + +theorem insertNth_right_injective {p : Fin (n + 1)} (x : α p) : + Function.Injective (insertNth p x) := + insertNth_injective2.right _ + /- Porting note: Once again, Lean told me `(fun x x_1 ↦ α x)` was an invalid motive, but disabling automatic insertion and specifying that motive seems to work. -/ theorem insertNth_apply_below {i j : Fin (n + 1)} (h : j < i) (x : α i) diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 51f860e725e71..c78a2338915fd 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -80,7 +80,7 @@ theorem mem_antidiagonalTuple {n : ℕ} {k : ℕ} {x : Fin k → ℕ} : · simp [eq_comm] | h x₀ x ih => simp_rw [Fin.sum_cons, antidiagonalTuple, List.mem_flatMap, List.mem_map, - List.Nat.mem_antidiagonal, Fin.cons_eq_cons, exists_eq_right_right, ih, + List.Nat.mem_antidiagonal, Fin.cons_inj, exists_eq_right_right, ih, @eq_comm _ _ (Prod.snd _), and_comm (a := Prod.snd _ = _), ← Prod.mk.inj_iff (a₁ := Prod.fst _), exists_eq_right] @@ -100,12 +100,12 @@ theorem nodup_antidiagonalTuple (k n : ℕ) : List.Nodup (antidiagonalTuple k n) refine List.Pairwise.cons (fun a ha x hx₁ hx₂ => ?_) (n_ih.map _ fun a b h x hx₁ hx₂ => ?_) · rw [List.mem_map] at hx₁ hx₂ ha obtain ⟨⟨a, -, rfl⟩, ⟨x₁, -, rfl⟩, ⟨x₂, -, h⟩⟩ := ha, hx₁, hx₂ - rw [Fin.cons_eq_cons] at h + rw [Fin.cons_inj] at h injection h.1 · rw [List.mem_map] at hx₁ hx₂ obtain ⟨⟨x₁, hx₁, rfl⟩, ⟨x₂, hx₂, h₁₂⟩⟩ := hx₁, hx₂ dsimp at h₁₂ - rw [Fin.cons_eq_cons, Nat.succ_inj'] at h₁₂ + rw [Fin.cons_inj, Nat.succ_inj'] at h₁₂ obtain ⟨h₁₂, rfl⟩ := h₁₂ rw [Function.onFun, h₁₂] at h exact h (List.mem_map_of_mem _ hx₁) (List.mem_map_of_mem _ hx₂) diff --git a/Mathlib/Data/Fin/Tuple/Sort.lean b/Mathlib/Data/Fin/Tuple/Sort.lean index b8073b76ebc30..b34388877eee2 100644 --- a/Mathlib/Data/Fin/Tuple/Sort.lean +++ b/Mathlib/Data/Fin/Tuple/Sort.lean @@ -163,8 +163,8 @@ theorem eq_sort_iff : σ = sort f ↔ Monotone (f ∘ σ) ∧ ∀ i j, i < j → f (σ i) = f (σ j) → σ i < σ j := by rw [eq_sort_iff'] refine ⟨fun h => ⟨(monotone_proj f).comp h.monotone, fun i j hij hfij => ?_⟩, fun h i j hij => ?_⟩ - · exact (((Prod.Lex.lt_iff _ _).1 <| h hij).resolve_left hfij.not_lt).2 - · obtain he | hl := (h.1 hij.le).eq_or_lt <;> apply (Prod.Lex.lt_iff _ _).2 + · exact ((Prod.Lex.toLex_lt_toLex.1 <| h hij).resolve_left hfij.not_lt).2 + · obtain he | hl := (h.1 hij.le).eq_or_lt <;> apply Prod.Lex.toLex_lt_toLex.2 exacts [Or.inr ⟨he, h.2 i j hij he⟩, Or.inl hl] /-- The permutation that sorts `f` is the identity if and only if `f` is monotone. -/ diff --git a/Mathlib/Data/FinEnum.lean b/Mathlib/Data/FinEnum.lean index 3d499baf8ef5b..d928b5d459a79 100644 --- a/Mathlib/Data/FinEnum.lean +++ b/Mathlib/Data/FinEnum.lean @@ -46,7 +46,7 @@ def ofNodupList [DecidableEq α] (xs : List α) (h : ∀ x : α, x ∈ xs) (h' : FinEnum α where card := xs.length equiv := - ⟨fun x => ⟨xs.indexOf x, by rw [List.indexOf_lt_length]; apply h⟩, xs.get, fun x => by simp, + ⟨fun x => ⟨xs.indexOf x, by rw [List.indexOf_lt_length_iff]; apply h⟩, xs.get, fun x => by simp, fun i => by ext; simp [List.indexOf_getElem h']⟩ /-- create a `FinEnum` instance from an exhaustive list; duplicates are removed -/ diff --git a/Mathlib/Data/Finite/Sum.lean b/Mathlib/Data/Finite/Sum.lean index 177f8e038134a..4c3ecc2999ef5 100644 --- a/Mathlib/Data/Finite/Sum.lean +++ b/Mathlib/Data/Finite/Sum.lean @@ -24,4 +24,13 @@ theorem sum_left (β) [Finite (α ⊕ β)] : Finite α := theorem sum_right (α) [Finite (α ⊕ β)] : Finite β := of_injective (Sum.inr : β → α ⊕ β) Sum.inr_injective +instance {α β : Sort*} [Finite α] [Finite β] : Finite (α ⊕' β) := + of_equiv _ ((Equiv.psumEquivSum _ _).symm.trans (Equiv.plift.psumCongr Equiv.plift)) + +theorem psum_left {α β : Sort*} [Finite (α ⊕' β)] : Finite α := + of_injective (PSum.inl : α → α ⊕' β) PSum.inl_injective + +theorem psum_right {α β : Sort*} [Finite (α ⊕' β)] : Finite β := + of_injective (PSum.inr : β → α ⊕' β) PSum.inr_injective + end Finite diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 84a285a152f6a..afb347e53972a 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -92,6 +92,11 @@ theorem disjoint_of_subset_iff_left_eq_empty (h : s ⊆ t) : Disjoint s t ↔ s = ∅ := disjoint_of_le_iff_left_eq_bot h +lemma pairwiseDisjoint_iff {ι : Type*} {s : Set ι} {f : ι → Finset α} : + s.PairwiseDisjoint f ↔ ∀ ⦃i⦄, i ∈ s → ∀ ⦃j⦄, j ∈ s → (f i ∩ f j).Nonempty → i = j := by + simp [Set.PairwiseDisjoint, Set.Pairwise, Function.onFun, not_imp_comm (a := _ = _), + not_disjoint_iff_nonempty_inter] + end Lattice instance isDirected_le : IsDirected (Finset α) (· ≤ ·) := by classical infer_instance diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 9c9758dc56612..2648f2ebda7ac 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -411,6 +411,18 @@ lemma card_le_card_of_injOn (f : α → β) (hf : ∀ a ∈ s, f a ∈ t) (f_inj _ ≤ #t := card_le_card <| image_subset_iff.2 hf @[deprecated (since := "2024-06-01")] alias card_le_card_of_inj_on := card_le_card_of_injOn +lemma card_le_card_of_injective {f : s → t} (hf : f.Injective) : #s ≤ #t := by + rcases s.eq_empty_or_nonempty with rfl | ⟨a₀, ha₀⟩ + · simp + · classical + let f' : α → β := fun a => f (if ha : a ∈ s then ⟨a, ha⟩ else ⟨a₀, ha₀⟩) + apply card_le_card_of_injOn f' + · aesop + · intro a₁ ha₁ a₂ ha₂ haa + rw [mem_coe] at ha₁ ha₂ + simp only [f', ha₁, ha₂, ← Subtype.ext_iff] at haa + exact Subtype.ext_iff.mp (hf haa) + lemma card_le_card_of_surjOn (f : α → β) (hf : Set.SurjOn f s t) : #t ≤ #s := by classical unfold Set.SurjOn at hf; exact (card_le_card (mod_cast hf)).trans card_image_le diff --git a/Mathlib/Data/Finset/Defs.lean b/Mathlib/Data/Finset/Defs.lean index 93ee27d0d49a1..334d0c01cd155 100644 --- a/Mathlib/Data/Finset/Defs.lean +++ b/Mathlib/Data/Finset/Defs.lean @@ -253,6 +253,8 @@ theorem mem_of_subset {s₁ s₂ : Finset α} {a : α} : s₁ ⊆ s₂ → a ∈ theorem not_mem_mono {s t : Finset α} (h : s ⊆ t) {a : α} : a ∉ t → a ∉ s := mt <| @h _ +alias not_mem_subset := not_mem_mono + theorem Subset.antisymm {s₁ s₂ : Finset α} (H₁ : s₁ ⊆ s₂) (H₂ : s₂ ⊆ s₁) : s₁ = s₂ := ext fun a => ⟨@H₁ a, @H₂ a⟩ diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index c05c5afa72d5d..7491192b997e4 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -768,4 +768,26 @@ theorem finsetCongr_toEmbedding (e : α ≃ β) : e.finsetCongr.toEmbedding = (Finset.mapEmbedding e.toEmbedding).toEmbedding := rfl +/-- Given a predicate `p : α → Prop`, produces an equivalence between + `Finset {a : α // p a}` and `{s : Finset α // ∀ a ∈ s, p a}`. -/ + +@[simps] +protected def finsetSubtypeComm (p : α → Prop) : + Finset {a : α // p a} ≃ {s : Finset α // ∀ a ∈ s, p a} where + toFun s := ⟨s.map ⟨fun a ↦ a.val, Subtype.val_injective⟩, fun _ h ↦ + have ⟨v, _, h⟩ := Embedding.coeFn_mk _ _ ▸ mem_map.mp h; h ▸ v.property⟩ + invFun s := s.val.attach.map (Subtype.impEmbedding _ _ s.property) + left_inv s := by + ext a; constructor <;> intro h <;> + simp only [Finset.mem_map, Finset.mem_attach, true_and, Subtype.exists, Embedding.coeFn_mk, + exists_and_right, exists_eq_right, Subtype.impEmbedding, Subtype.mk.injEq] at * + · rcases h with ⟨_, ⟨_, h₁⟩, h₂⟩; exact h₂ ▸ h₁ + · use a, ⟨a.property, h⟩ + right_inv s := by + ext a; constructor <;> intro h <;> + simp only [Finset.mem_map, Finset.mem_attach, true_and, Subtype.exists, Embedding.coeFn_mk, + exists_and_right, exists_eq_right, Subtype.impEmbedding, Subtype.mk.injEq] at * + · rcases h with ⟨_, _, h₁, h₂⟩; exact h₂ ▸ h₁ + · use s.property _ h, a + end Equiv diff --git a/Mathlib/Data/Finset/Insert.lean b/Mathlib/Data/Finset/Insert.lean index 3c72a9bb66c17..5c2e92723c7e6 100644 --- a/Mathlib/Data/Finset/Insert.lean +++ b/Mathlib/Data/Finset/Insert.lean @@ -197,6 +197,13 @@ theorem nontrivial_iff_ne_singleton (ha : a ∈ s) : s.Nontrivial ↔ s ≠ {a} theorem Nonempty.exists_eq_singleton_or_nontrivial : s.Nonempty → (∃ a, s = {a}) ∨ s.Nontrivial := fun ⟨a, ha⟩ => (eq_singleton_or_nontrivial ha).imp_left <| Exists.intro a +@[simp, norm_cast] lemma nontrivial_coe : (s : Set α).Nontrivial ↔ s.Nontrivial := .rfl + +alias ⟨Nontrivial.of_coe, Nontrivial.coe⟩ := nontrivial_coe + +lemma Nontrivial.not_subset_singleton (hs : s.Nontrivial) : ¬s ⊆ {a} := + mod_cast hs.coe.not_subset_singleton + instance instNontrivial [Nonempty α] : Nontrivial (Finset α) := ‹Nonempty α›.elim fun a => ⟨⟨{a}, ∅, singleton_ne_empty _⟩⟩ diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index ff436a9d458d3..0a04809d8a4e6 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -143,7 +143,7 @@ protected theorem sup_comm (s : Finset β) (t : Finset γ) (f : β → γ → α (s.sup fun b => t.sup (f b)) = t.sup fun c => s.sup fun b => f b c := eq_of_forall_ge_iff fun a => by simpa using forall₂_swap -@[simp, nolint simpNF] -- Porting note: linter claims that LHS does not simplify +@[simp] theorem sup_attach (s : Finset β) (f : β → α) : (s.attach.sup fun x => f x) = s.sup f := (s.attach.sup_map (Function.Embedding.subtype _) f).symm.trans <| congr_arg _ attach_map_val diff --git a/Mathlib/Data/Finset/Powerset.lean b/Mathlib/Data/Finset/Powerset.lean index dc12dc7eccee3..8c387f738be55 100644 --- a/Mathlib/Data/Finset/Powerset.lean +++ b/Mathlib/Data/Finset/Powerset.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Mathlib.Data.Finset.Lattice.Fold import Mathlib.Data.Multiset.Powerset +import Mathlib.Data.Set.Pairwise.Lattice /-! # The powerset of a finset @@ -93,6 +94,16 @@ theorem powerset_insert [DecidableEq α] (s : Finset α) (a : α) : · have : ¬∃ u : Finset α, u ⊆ s ∧ insert a u = t := by simp [Ne.symm (ne_insert_of_not_mem _ _ h)] simp [Finset.erase_eq_of_not_mem h, this] +lemma pairwiseDisjoint_pair_insert [DecidableEq α] {a : α} (ha : a ∉ s) : + (s.powerset : Set (Finset α)).PairwiseDisjoint fun t ↦ ({t, insert a t} : Set (Finset α)) := by + simp_rw [Set.pairwiseDisjoint_iff, mem_coe, mem_powerset] + rintro i hi j hj + simp only [Set.Nonempty, Set.mem_inter_iff, Set.mem_insert_iff, Set.mem_singleton_iff, + exists_eq_or_imp, exists_eq_left, or_imp, imp_self, true_and] + refine ⟨?_, ?_, insert_erase_invOn.2.injOn (not_mem_mono hi ha) (not_mem_mono hj ha)⟩ <;> + rintro rfl <;> + cases Finset.not_mem_mono ‹_› ha (Finset.mem_insert_self _ _) + /-- For predicate `p` decidable on subsets, it is decidable whether `p` holds for any subset. -/ instance decidableExistsOfDecidableSubsets {s : Finset α} {p : ∀ t ⊆ s, Prop} [∀ (t) (h : t ⊆ s), Decidable (p t h)] : Decidable (∃ (t : _) (h : t ⊆ s), p t h) := diff --git a/Mathlib/Data/Finset/Sups.lean b/Mathlib/Data/Finset/Sups.lean index 04324a2636c5c..f627f7e9a839c 100644 --- a/Mathlib/Data/Finset/Sups.lean +++ b/Mathlib/Data/Finset/Sups.lean @@ -169,7 +169,6 @@ lemma biUnion_image_sup_left : s.biUnion (fun a ↦ t.image (a ⊔ ·)) = s ⊻ lemma biUnion_image_sup_right : t.biUnion (fun b ↦ s.image (· ⊔ b)) = s ⊻ t := biUnion_image_right --- Porting note: simpNF linter doesn't like @[simp] theorem image_sup_product (s t : Finset α) : (s ×ˢ t).image (uncurry (· ⊔ ·)) = s ⊻ t := image_uncurry_product _ _ _ @@ -314,7 +313,6 @@ lemma biUnion_image_inf_left : s.biUnion (fun a ↦ t.image (a ⊓ ·)) = s ⊼ lemma biUnion_image_inf_right : t.biUnion (fun b ↦ s.image (· ⊓ b)) = s ⊼ t := biUnion_image_right --- Porting note: simpNF linter doesn't like @[simp] theorem image_inf_product (s t : Finset α) : (s ×ˢ t).image (uncurry (· ⊓ ·)) = s ⊼ t := image_uncurry_product _ _ _ diff --git a/Mathlib/Data/Finset/Sym.lean b/Mathlib/Data/Finset/Sym.lean index 29ac2bd2208d5..9f35d53d8cc67 100644 --- a/Mathlib/Data/Finset/Sym.lean +++ b/Mathlib/Data/Finset/Sym.lean @@ -156,8 +156,6 @@ section Sym2 variable {m : Sym2 α} --- Porting note: add this lemma and remove simp in the next lemma since simpNF lint --- warns that its LHS is not in normal form @[simp] theorem diag_mem_sym2_mem_iff : (∀ b, b ∈ Sym2.diag a → b ∈ s) ↔ a ∈ s := by rw [← mem_sym2_iff] diff --git a/Mathlib/Data/Finset/Union.lean b/Mathlib/Data/Finset/Union.lean index 42efbe5a93ae2..e457938f87ae0 100644 --- a/Mathlib/Data/Finset/Union.lean +++ b/Mathlib/Data/Finset/Union.lean @@ -81,14 +81,20 @@ lemma disjiUnion_disjiUnion (s : Finset α) (f : α → Finset β) (g : β → F exact disjoint_left.mp (h1 a.prop b.prop <| Subtype.coe_injective.ne hab) hfa hfb := eq_of_veq <| Multiset.bind_assoc.trans (Multiset.attach_bind_coe _ _).symm +lemma sUnion_disjiUnion {f : α → Finset (Set β)} (I : Finset α) + (hf : (I : Set α).PairwiseDisjoint f) : + ⋃₀ (I.disjiUnion f hf : Set (Set β)) = ⋃ a ∈ I, ⋃₀ ↑(f a) := by + ext + simp only [coe_disjiUnion, Set.mem_sUnion, Set.mem_iUnion, mem_coe, exists_prop] + tauto + variable [DecidableEq β] {s : Finset α} {t : Finset β} {f : α → β} private lemma pairwiseDisjoint_fibers : Set.PairwiseDisjoint ↑t fun a ↦ s.filter (f · = a) := fun x' hx y' hy hne ↦ by simp_rw [disjoint_left, mem_filter]; rintro i ⟨_, rfl⟩ ⟨_, rfl⟩; exact hne rfl --- `simpNF` claims that the statement can't simplify itself, but it can (as of 2024-02-14) -@[simp, nolint simpNF] lemma disjiUnion_filter_eq (s : Finset α) (t : Finset β) (f : α → β) : +@[simp] lemma disjiUnion_filter_eq (s : Finset α) (t : Finset β) (f : α → β) : t.disjiUnion (fun a ↦ s.filter (f · = a)) pairwiseDisjoint_fibers = s.filter fun c ↦ f c ∈ t := ext fun b => by simpa using and_comm diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 1ce490ed52759..6abe2e3b9c254 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -1213,7 +1213,7 @@ section variable [Zero M] [MonoidWithZero R] [MulActionWithZero R M] -@[simp, nolint simpNF] -- `simpNF` incorrectly complains the LHS doesn't simplify. +@[simp] theorem single_smul (a b : α) (f : α → M) (r : R) : single a r b • f a = single a (r • f b) b := by by_cases h : a = b <;> simp [h] diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 56eb68dc8f217..8b53485f90f37 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Kim Morrison -/ import Mathlib.Algebra.Group.Indicator import Mathlib.Data.Set.Finite.Basic +import Mathlib.Tactic.FastInstance /-! # Type of functions with finite support @@ -487,7 +488,7 @@ theorem support_add_eq [DecidableEq α] {g₁ g₂ : α →₀ M} (h : Disjoint simp only [mem_support_iff, not_not] at *; simpa only [add_apply, this, zero_add] instance instAddZeroClass : AddZeroClass (α →₀ M) := - DFunLike.coe_injective.addZeroClass _ coe_zero coe_add + fast_instance% DFunLike.coe_injective.addZeroClass _ coe_zero coe_add instance instIsLeftCancelAdd [IsLeftCancelAdd M] : IsLeftCancelAdd (α →₀ M) where add_left_cancel _ _ _ h := ext fun x => add_left_cancel <| DFunLike.congr_fun h x @@ -566,14 +567,13 @@ instance instNatSMul : SMul ℕ (α →₀ M) := ⟨fun n v => v.mapRange (n • ·) (nsmul_zero _)⟩ instance instAddMonoid : AddMonoid (α →₀ M) := - DFunLike.coe_injective.addMonoid _ coe_zero coe_add fun _ _ => rfl + fast_instance% DFunLike.coe_injective.addMonoid _ coe_zero coe_add fun _ _ => rfl end AddMonoid instance instAddCommMonoid [AddCommMonoid M] : AddCommMonoid (α →₀ M) := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { DFunLike.coe_injective.addCommMonoid DFunLike.coe coe_zero coe_add (fun _ _ => rfl) with - toAddMonoid := Finsupp.instAddMonoid } + fast_instance% DFunLike.coe_injective.addCommMonoid + DFunLike.coe coe_zero coe_add (fun _ _ => rfl) instance instNeg [NegZeroClass G] : Neg (α →₀ G) := ⟨mapRange Neg.neg neg_zero⟩ @@ -616,16 +616,12 @@ instance instIntSMul [AddGroup G] : SMul ℤ (α →₀ G) := ⟨fun n v => v.mapRange (n • ·) (zsmul_zero _)⟩ instance instAddGroup [AddGroup G] : AddGroup (α →₀ G) := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { DFunLike.coe_injective.addGroup DFunLike.coe coe_zero coe_add coe_neg coe_sub (fun _ _ => rfl) - fun _ _ => rfl with - toAddMonoid := Finsupp.instAddMonoid } + fast_instance% DFunLike.coe_injective.addGroup DFunLike.coe coe_zero coe_add coe_neg coe_sub + (fun _ _ => rfl) fun _ _ => rfl instance instAddCommGroup [AddCommGroup G] : AddCommGroup (α →₀ G) := - --TODO: add reference to library note in PR https://github.com/leanprover-community/mathlib4/pull/7432 - { DFunLike.coe_injective.addCommGroup DFunLike.coe coe_zero coe_add coe_neg coe_sub - (fun _ _ => rfl) fun _ _ => rfl with - toAddGroup := Finsupp.instAddGroup } + fast_instance% DFunLike.coe_injective.addCommGroup DFunLike.coe coe_zero coe_add coe_neg coe_sub + (fun _ _ => rfl) fun _ _ => rfl @[simp] theorem support_neg [AddGroup G] (f : α →₀ G) : support (-f) = support f := diff --git a/Mathlib/Data/Finsupp/Fin.lean b/Mathlib/Data/Finsupp/Fin.lean index 049744c5fa6a9..027350c69ca70 100644 --- a/Mathlib/Data/Finsupp/Fin.lean +++ b/Mathlib/Data/Finsupp/Fin.lean @@ -42,7 +42,6 @@ theorem cons_zero : cons y s 0 = y := @[simp] theorem cons_succ : cons y s i.succ = s i := - -- Porting note: was Fin.cons_succ _ _ _ rfl @[simp] diff --git a/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean index fc03f56398bf2..d5a68b16f016c 100644 --- a/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean +++ b/Mathlib/Data/Finsupp/MonomialOrder/DegLex.lean @@ -107,7 +107,7 @@ theorem lt_def [LT α] {a b : DegLex (α →₀ ℕ)} : 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] + simp [lt_def, Prod.Lex.toLex_lt_toLex] variable [LinearOrder α] diff --git a/Mathlib/Data/Finsupp/ToDFinsupp.lean b/Mathlib/Data/Finsupp/ToDFinsupp.lean index 3776ed103ea6c..5174a0b7e5313 100644 --- a/Mathlib/Data/Finsupp/ToDFinsupp.lean +++ b/Mathlib/Data/Finsupp/ToDFinsupp.lean @@ -220,8 +220,6 @@ variable (R) /-- The additive version of `Finsupp.toFinsupp`. Note that this is `noncomputable` because `Finsupp.add` is noncomputable. -/ --- Porting note: `simps` generated lemmas that did not pass `simpNF` lints, manually added below ---@[simps? (config := .asFn)] def finsuppLequivDFinsupp [DecidableEq ι] [Semiring R] [AddCommMonoid M] [∀ m : M, Decidable (m ≠ 0)] [Module R M] : (ι →₀ M) ≃ₗ[R] Π₀ _ : ι, M := { finsuppEquivDFinsupp with @@ -230,7 +228,6 @@ def finsuppLequivDFinsupp [DecidableEq ι] [Semiring R] [AddCommMonoid M] map_smul' := Finsupp.toDFinsupp_smul map_add' := Finsupp.toDFinsupp_add } --- Porting note: `simps` generated as `↑(finsuppLequivDFinsupp R).toLinearMap = Finsupp.toDFinsupp` @[simp] theorem finsuppLequivDFinsupp_apply_apply [DecidableEq ι] [Semiring R] [AddCommMonoid M] [∀ m : M, Decidable (m ≠ 0)] [Module R M] : diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index ff4691dc54789..3207bf60506dd 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -1060,11 +1060,7 @@ variable [Fintype α] [DecidableEq β] {f : α → β} /-- `bijInv f` is the unique inverse to a bijection `f`. This acts as a computable alternative to `Function.invFun`. -/ def bijInv (f_bij : Bijective f) (b : β) : α := - Fintype.choose (fun a => f a = b) - (by - rcases f_bij.right b with ⟨a', fa_eq_b⟩ - rw [← fa_eq_b] - exact ⟨a', ⟨rfl, fun a h => f_bij.left h⟩⟩) + Fintype.choose (fun a => f a = b) (f_bij.existsUnique b) theorem leftInverse_bijInv (f_bij : Bijective f) : LeftInverse (bijInv f_bij) f := fun a => f_bij.left (choose_spec (fun a' => f a' = f a) _) diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index a9229d8aa827b..628317d302d5e 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Algebra.Group.TypeTags.Hom import Mathlib.Algebra.Ring.Hom.Basic import Mathlib.Algebra.Ring.Int.Defs import Mathlib.Algebra.Ring.Parity diff --git a/Mathlib/Data/Int/ConditionallyCompleteOrder.lean b/Mathlib/Data/Int/ConditionallyCompleteOrder.lean index 077cb7a5b47c8..03ec19f1799b6 100644 --- a/Mathlib/Data/Int/ConditionallyCompleteOrder.lean +++ b/Mathlib/Data/Int/ConditionallyCompleteOrder.lean @@ -17,8 +17,8 @@ open Int noncomputable section -open scoped Classical +open scoped Classical in instance instConditionallyCompleteLinearOrder : ConditionallyCompleteLinearOrder ℤ where __ := instLinearOrder __ := LinearOrder.toLattice diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 2b92ec1456668..74a86d6296cf7 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -228,8 +228,7 @@ where | 0 => Hp _ Int.le_rfl H0 | n+1 => by refine cast (by rw [Int.add_sub_assoc]; rfl) (Hp _ (Int.le_of_lt ?_) (neg n)) - conv => rhs; exact b.add_zero.symm - rw [Int.add_lt_add_iff_left]; apply negSucc_lt_zero + omega variable {z b H0 Hs Hp} diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 8955bf6337362..d51295c269ce1 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -704,10 +704,12 @@ theorem indexOf_le_length {a : α} {l : List α} : indexOf a l ≤ length l := b · rw [if_pos h]; exact Nat.zero_le _ · rw [if_neg h]; exact succ_le_succ ih -theorem indexOf_lt_length {a} {l : List α} : indexOf a l < length l ↔ a ∈ l := +theorem indexOf_lt_length_iff {a} {l : List α} : indexOf a l < length l ↔ a ∈ l := ⟨fun h => Decidable.byContradiction fun al => Nat.ne_of_lt h <| indexOf_eq_length.2 al, fun al => (lt_of_le_of_ne indexOf_le_length) fun h => indexOf_eq_length.1 h al⟩ +@[deprecated (since := "2025-01-22")] alias indexOf_lt_length := indexOf_lt_length_iff + theorem indexOf_append_of_mem {a : α} (h : a ∈ l₁) : indexOf a (l₁ ++ l₂) = indexOf a l₁ := by induction' l₁ with d₁ t₁ ih · exfalso @@ -799,7 +801,8 @@ theorem indexOf_get [DecidableEq α] {a : α} {l : List α} (h) : get l ⟨index @[simp] theorem getElem?_indexOf [DecidableEq α] {a : α} {l : List α} (h : a ∈ l) : - l[indexOf a l]? = some a := by rw [getElem?_eq_getElem, getElem_indexOf (indexOf_lt_length.2 h)] + l[indexOf a l]? = some a := by + rw [getElem?_eq_getElem, getElem_indexOf (indexOf_lt_length_iff.2 h)] -- This is incorrectly named and should be `get?_indexOf`; -- this already exists, so will require a deprecation dance. @@ -810,8 +813,8 @@ theorem indexOf_inj [DecidableEq α] {l : List α} {x y : α} (hx : x ∈ l) (hy indexOf x l = indexOf y l ↔ x = y := ⟨fun h => by have x_eq_y : - get l ⟨indexOf x l, indexOf_lt_length.2 hx⟩ = - get l ⟨indexOf y l, indexOf_lt_length.2 hy⟩ := by + get l ⟨indexOf x l, indexOf_lt_length_iff.2 hx⟩ = + get l ⟨indexOf y l, indexOf_lt_length_iff.2 hy⟩ := by simp only [h] simp only [indexOf_get] at x_eq_y; exact x_eq_y, fun h => by subst h; rfl⟩ @@ -965,21 +968,6 @@ end map_bijectivity theorem eq_of_mem_map_const {b₁ b₂ : β} {l : List α} (h : b₁ ∈ map (const α b₂) l) : b₁ = b₂ := by rw [map_const] at h; exact eq_of_mem_replicate h -/-! ### zipWith -/ - -theorem nil_zipWith (f : α → β → γ) (l : List β) : zipWith f [] l = [] := by cases l <;> rfl - -theorem zipWith_nil (f : α → β → γ) (l : List α) : zipWith f l [] = [] := by cases l <;> rfl - -@[simp] -theorem zipWith_flip (f : α → β → γ) : ∀ as bs, zipWith (flip f) bs as = zipWith f as bs - | [], [] => rfl - | [], _ :: _ => rfl - | _ :: _, [] => rfl - | a :: as, b :: bs => by - simp! [zipWith_flip] - rfl - /-! ### take, drop -/ @[simp] lemma take_eq_self_iff (x : List α) {n : ℕ} : x.take n = x ↔ x.length ≤ n := @@ -1633,12 +1621,7 @@ theorem length_eq_length_filter_add {l : List (α)} (f : α → Bool) : /-! ### filterMap -/ --- Later porting note (at time of this lemma moving to Batteries): --- removing attribute `nolint simpNF` attribute [simp 1100] filterMap_cons_none - --- Later porting note (at time of this lemma moving to Batteries): --- removing attribute `nolint simpNF` attribute [simp 1100] filterMap_cons_some theorem filterMap_eq_flatMap_toList (f : α → Option β) (l : List α) : @@ -1966,225 +1949,6 @@ theorem choose_property (hp : ∃ a, a ∈ l ∧ p a) : p (choose p l hp) := end Choose -/-! ### map₂Left' -/ - -section Map₂Left' - --- The definitional equalities for `map₂Left'` can already be used by the --- simplifier because `map₂Left'` is marked `@[simp]`. -@[simp] -theorem map₂Left'_nil_right (f : α → Option β → γ) (as) : - map₂Left' f as [] = (as.map fun a => f a none, []) := by cases as <;> rfl - -end Map₂Left' - -/-! ### map₂Right' -/ - -section Map₂Right' - -variable (f : Option α → β → γ) (a : α) (as : List α) (b : β) (bs : List β) - -@[simp] -theorem map₂Right'_nil_left : map₂Right' f [] bs = (bs.map (f none), []) := by cases bs <;> rfl - -@[simp] -theorem map₂Right'_nil_right : map₂Right' f as [] = ([], as) := - rfl - -@[simp] -theorem map₂Right'_nil_cons : map₂Right' f [] (b :: bs) = (f none b :: bs.map (f none), []) := - rfl - -@[simp] -theorem map₂Right'_cons_cons : - map₂Right' f (a :: as) (b :: bs) = - let r := map₂Right' f as bs - (f (some a) b :: r.fst, r.snd) := - rfl - -end Map₂Right' - -/-! ### zipLeft' -/ - -section ZipLeft' - -variable (a : α) (as : List α) (b : β) (bs : List β) - -@[simp] -theorem zipLeft'_nil_right : zipLeft' as ([] : List β) = (as.map fun a => (a, none), []) := by - cases as <;> rfl - -@[simp] -theorem zipLeft'_nil_left : zipLeft' ([] : List α) bs = ([], bs) := - rfl - -@[simp] -theorem zipLeft'_cons_nil : - zipLeft' (a :: as) ([] : List β) = ((a, none) :: as.map fun a => (a, none), []) := - rfl - -@[simp] -theorem zipLeft'_cons_cons : - zipLeft' (a :: as) (b :: bs) = - let r := zipLeft' as bs - ((a, some b) :: r.fst, r.snd) := - rfl - -end ZipLeft' - -/-! ### zipRight' -/ - -section ZipRight' - -variable (a : α) (as : List α) (b : β) (bs : List β) - -@[simp] -theorem zipRight'_nil_left : zipRight' ([] : List α) bs = (bs.map fun b => (none, b), []) := by - cases bs <;> rfl - -@[simp] -theorem zipRight'_nil_right : zipRight' as ([] : List β) = ([], as) := - rfl - -@[simp] -theorem zipRight'_nil_cons : - zipRight' ([] : List α) (b :: bs) = ((none, b) :: bs.map fun b => (none, b), []) := - rfl - -@[simp] -theorem zipRight'_cons_cons : - zipRight' (a :: as) (b :: bs) = - let r := zipRight' as bs - ((some a, b) :: r.fst, r.snd) := - rfl - -end ZipRight' - -/-! ### map₂Left -/ - -section Map₂Left - -variable (f : α → Option β → γ) (as : List α) - --- The definitional equalities for `map₂Left` can already be used by the --- simplifier because `map₂Left` is marked `@[simp]`. -@[simp] -theorem map₂Left_nil_right : map₂Left f as [] = as.map fun a => f a none := by cases as <;> rfl - -theorem map₂Left_eq_map₂Left' : ∀ as bs, map₂Left f as bs = (map₂Left' f as bs).fst - | [], _ => by simp - | a :: as, [] => by simp - | a :: as, b :: bs => by simp [map₂Left_eq_map₂Left'] - -theorem map₂Left_eq_zipWith : - ∀ as bs, length as ≤ length bs → map₂Left f as bs = zipWith (fun a b => f a (some b)) as bs - | [], [], _ => by simp - | [], _ :: _, _ => by simp - | a :: as, [], h => by - simp at h - | a :: as, b :: bs, h => by - simp only [length_cons, succ_le_succ_iff] at h - simp [h, map₂Left_eq_zipWith] - -end Map₂Left - -/-! ### map₂Right -/ - -section Map₂Right - -variable (f : Option α → β → γ) (a : α) (as : List α) (b : β) (bs : List β) - -@[simp] -theorem map₂Right_nil_left : map₂Right f [] bs = bs.map (f none) := by cases bs <;> rfl - -@[simp] -theorem map₂Right_nil_right : map₂Right f as [] = [] := - rfl - -@[simp] -theorem map₂Right_nil_cons : map₂Right f [] (b :: bs) = f none b :: bs.map (f none) := - rfl - -@[simp] -theorem map₂Right_cons_cons : - map₂Right f (a :: as) (b :: bs) = f (some a) b :: map₂Right f as bs := - rfl - -theorem map₂Right_eq_map₂Right' : map₂Right f as bs = (map₂Right' f as bs).fst := by - simp only [map₂Right, map₂Right', map₂Left_eq_map₂Left'] - -theorem map₂Right_eq_zipWith (h : length bs ≤ length as) : - map₂Right f as bs = zipWith (fun a b => f (some a) b) as bs := by - have : (fun a b => flip f a (some b)) = flip fun a b => f (some a) b := rfl - simp only [map₂Right, map₂Left_eq_zipWith, zipWith_flip, *] - -end Map₂Right - -/-! ### zipLeft -/ - -section ZipLeft - -variable (a : α) (as : List α) (b : β) (bs : List β) - -@[simp] -theorem zipLeft_nil_right : zipLeft as ([] : List β) = as.map fun a => (a, none) := by - cases as <;> rfl - -@[simp] -theorem zipLeft_nil_left : zipLeft ([] : List α) bs = [] := - rfl - -@[simp] -theorem zipLeft_cons_nil : - zipLeft (a :: as) ([] : List β) = (a, none) :: as.map fun a => (a, none) := - rfl - -@[simp] -theorem zipLeft_cons_cons : zipLeft (a :: as) (b :: bs) = (a, some b) :: zipLeft as bs := - rfl - --- Porting note: arguments explicit for recursion -theorem zipLeft_eq_zipLeft' (as : List α) (bs : List β) : zipLeft as bs = (zipLeft' as bs).fst := by - rw [zipLeft, zipLeft'] - cases as with - | nil => rfl - | cons _ atl => - cases bs with - | nil => rfl - | cons _ btl => - rw [zipWithLeft, zipWithLeft', cons_inj_right] - exact @zipLeft_eq_zipLeft' atl btl - -end ZipLeft - -/-! ### zipRight -/ - -section ZipRight - -variable (a : α) (as : List α) (b : β) (bs : List β) - -@[simp] -theorem zipRight_nil_left : zipRight ([] : List α) bs = bs.map fun b => (none, b) := by - cases bs <;> rfl - -@[simp] -theorem zipRight_nil_right : zipRight as ([] : List β) = [] := - rfl - -@[simp] -theorem zipRight_nil_cons : - zipRight ([] : List α) (b :: bs) = (none, b) :: bs.map fun b => (none, b) := - rfl - -@[simp] -theorem zipRight_cons_cons : zipRight (a :: as) (b :: bs) = (some a, b) :: zipRight as bs := - rfl - -theorem zipRight_eq_zipRight' : zipRight as bs = (zipRight' as bs).fst := by - induction as generalizing bs <;> cases bs <;> simp [*] - -end ZipRight - /-! ### Forall -/ section Forall @@ -2321,4 +2085,4 @@ end lookup end List -set_option linter.style.longFile 2400 +set_option linter.style.longFile 2200 diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index 8281f5ce92052..de54055232145 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -61,7 +61,7 @@ theorem finRange_map_get (l : List α) : (finRange l.length).map l.get = l := List.ext_get (by simp) (by simp) @[simp] theorem indexOf_finRange {k : ℕ} (i : Fin k) : (finRange k).indexOf i = i := by - have : (finRange k).indexOf i < (finRange k).length := indexOf_lt_length.mpr (by simp) + have : (finRange k).indexOf i < (finRange k).length := indexOf_lt_length_iff.mpr (by simp) have h₁ : (finRange k).get ⟨(finRange k).indexOf i, this⟩ = i := indexOf_get this have h₂ : (finRange k).get ⟨i, by simp⟩ = i := get_finRange _ simpa using (Nodup.get_inj_iff (nodup_finRange k)).mp (Eq.trans h₁ h₂.symm) diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index c9eee93aee149..0ea92237b0822 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -56,12 +56,7 @@ theorem nodup (n m : ℕ) : Nodup (Ico n m) := by @[simp] theorem mem {n m l : ℕ} : l ∈ Ico n m ↔ n ≤ l ∧ l < m := by suffices n ≤ l ∧ l < n + (m - n) ↔ n ≤ l ∧ l < m by simp [Ico, this] - rcases le_total n m with hnm | hmn - · rw [Nat.add_sub_cancel' hnm] - · rw [Nat.sub_eq_zero_iff_le.mpr hmn, Nat.add_zero] - exact - and_congr_right fun hnl => - Iff.intro (fun hln => (not_le_of_gt hln hnl).elim) fun hlm => lt_of_lt_of_le hlm hmn + omega theorem eq_nil_of_le {n m : ℕ} (h : m ≤ n) : Ico n m = [] := by simp [Ico, Nat.sub_eq_zero_iff_le.mpr h] diff --git a/Mathlib/Data/List/Map2.lean b/Mathlib/Data/List/Map2.lean new file mode 100644 index 0000000000000..bea17d98be9f5 --- /dev/null +++ b/Mathlib/Data/List/Map2.lean @@ -0,0 +1,274 @@ +/- +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.Defs +import Mathlib.Tactic.Common + +/-! +# Map₂ Lemmas + +This file contains additional lemmas about a number of list functions related to combining zipping +Lists together. In particular, we include lemmas about: + +* `map₂Left'` +* `map₂Right'` +* `zipWith` +* `zipLeft'` +* `zipRight'` + +-/ + +assert_not_exists GroupWithZero +assert_not_exists Lattice +assert_not_exists Prod.swap_eq_iff_eq_swap +assert_not_exists Ring +assert_not_exists Set.range + +open Function + +open Nat hiding one_pos + +namespace List + +universe u v w + +variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} + +/-! ### map₂Left' -/ + +section Map₂Left' + +-- The definitional equalities for `map₂Left'` can already be used by the +-- simplifier because `map₂Left'` is marked `@[simp]`. +@[simp] +theorem map₂Left'_nil_right (f : α → Option β → γ) (as) : + map₂Left' f as [] = (as.map fun a => f a none, []) := by cases as <;> rfl + +end Map₂Left' + +/-! ### map₂Right' -/ + +section Map₂Right' + +variable (f : Option α → β → γ) (a : α) (as : List α) (b : β) (bs : List β) + +@[simp] +theorem map₂Right'_nil_left : map₂Right' f [] bs = (bs.map (f none), []) := by cases bs <;> rfl + +@[simp] +theorem map₂Right'_nil_right : map₂Right' f as [] = ([], as) := + rfl + +@[simp] +theorem map₂Right'_nil_cons : map₂Right' f [] (b :: bs) = (f none b :: bs.map (f none), []) := + rfl + +@[simp] +theorem map₂Right'_cons_cons : + map₂Right' f (a :: as) (b :: bs) = + let r := map₂Right' f as bs + (f (some a) b :: r.fst, r.snd) := + rfl + +end Map₂Right' + +/-! ### zipWith -/ + +theorem nil_zipWith (f : α → β → γ) (l : List β) : zipWith f [] l = [] := by cases l <;> rfl + +theorem zipWith_nil (f : α → β → γ) (l : List α) : zipWith f l [] = [] := by cases l <;> rfl + +@[simp] +theorem zipWith_flip (f : α → β → γ) : ∀ as bs, zipWith (flip f) bs as = zipWith f as bs + | [], [] => rfl + | [], _ :: _ => rfl + | _ :: _, [] => rfl + | a :: as, b :: bs => by + simp! [zipWith_flip] + rfl + + +/-! ### zipLeft' -/ + +section ZipLeft' + +variable (a : α) (as : List α) (b : β) (bs : List β) + +@[simp] +theorem zipLeft'_nil_right : zipLeft' as ([] : List β) = (as.map fun a => (a, none), []) := by + cases as <;> rfl + +@[simp] +theorem zipLeft'_nil_left : zipLeft' ([] : List α) bs = ([], bs) := + rfl + +@[simp] +theorem zipLeft'_cons_nil : + zipLeft' (a :: as) ([] : List β) = ((a, none) :: as.map fun a => (a, none), []) := + rfl + +@[simp] +theorem zipLeft'_cons_cons : + zipLeft' (a :: as) (b :: bs) = + let r := zipLeft' as bs + ((a, some b) :: r.fst, r.snd) := + rfl + +end ZipLeft' + +/-! ### zipRight' -/ + +section ZipRight' + +variable (a : α) (as : List α) (b : β) (bs : List β) + +@[simp] +theorem zipRight'_nil_left : zipRight' ([] : List α) bs = (bs.map fun b => (none, b), []) := by + cases bs <;> rfl + +@[simp] +theorem zipRight'_nil_right : zipRight' as ([] : List β) = ([], as) := + rfl + +@[simp] +theorem zipRight'_nil_cons : + zipRight' ([] : List α) (b :: bs) = ((none, b) :: bs.map fun b => (none, b), []) := + rfl + +@[simp] +theorem zipRight'_cons_cons : + zipRight' (a :: as) (b :: bs) = + let r := zipRight' as bs + ((some a, b) :: r.fst, r.snd) := + rfl + +end ZipRight' + +/-! ### map₂Left -/ + +section Map₂Left + +variable (f : α → Option β → γ) (as : List α) + +-- The definitional equalities for `map₂Left` can already be used by the +-- simplifier because `map₂Left` is marked `@[simp]`. +@[simp] +theorem map₂Left_nil_right : map₂Left f as [] = as.map fun a => f a none := by cases as <;> rfl + +theorem map₂Left_eq_map₂Left' : ∀ as bs, map₂Left f as bs = (map₂Left' f as bs).fst + | [], _ => by simp + | a :: as, [] => by simp + | a :: as, b :: bs => by simp [map₂Left_eq_map₂Left'] + +theorem map₂Left_eq_zipWith : + ∀ as bs, length as ≤ length bs → map₂Left f as bs = zipWith (fun a b => f a (some b)) as bs + | [], [], _ => by simp + | [], _ :: _, _ => by simp + | a :: as, [], h => by + simp at h + | a :: as, b :: bs, h => by + simp only [length_cons, succ_le_succ_iff] at h + simp [h, map₂Left_eq_zipWith] + +end Map₂Left + +/-! ### map₂Right -/ + +section Map₂Right + +variable (f : Option α → β → γ) (a : α) (as : List α) (b : β) (bs : List β) + +@[simp] +theorem map₂Right_nil_left : map₂Right f [] bs = bs.map (f none) := by cases bs <;> rfl + +@[simp] +theorem map₂Right_nil_right : map₂Right f as [] = [] := + rfl + +@[simp] +theorem map₂Right_nil_cons : map₂Right f [] (b :: bs) = f none b :: bs.map (f none) := + rfl + +@[simp] +theorem map₂Right_cons_cons : + map₂Right f (a :: as) (b :: bs) = f (some a) b :: map₂Right f as bs := + rfl + +theorem map₂Right_eq_map₂Right' : map₂Right f as bs = (map₂Right' f as bs).fst := by + simp only [map₂Right, map₂Right', map₂Left_eq_map₂Left'] + +theorem map₂Right_eq_zipWith (h : length bs ≤ length as) : + map₂Right f as bs = zipWith (fun a b => f (some a) b) as bs := by + have : (fun a b => flip f a (some b)) = flip fun a b => f (some a) b := rfl + simp only [map₂Right, map₂Left_eq_zipWith, zipWith_flip, *] + +end Map₂Right + +/-! ### zipLeft -/ + +section ZipLeft + +variable (a : α) (as : List α) (b : β) (bs : List β) + +@[simp] +theorem zipLeft_nil_right : zipLeft as ([] : List β) = as.map fun a => (a, none) := by + cases as <;> rfl + +@[simp] +theorem zipLeft_nil_left : zipLeft ([] : List α) bs = [] := + rfl + +@[simp] +theorem zipLeft_cons_nil : + zipLeft (a :: as) ([] : List β) = (a, none) :: as.map fun a => (a, none) := + rfl + +@[simp] +theorem zipLeft_cons_cons : zipLeft (a :: as) (b :: bs) = (a, some b) :: zipLeft as bs := + rfl + +-- Porting note: arguments explicit for recursion +theorem zipLeft_eq_zipLeft' (as : List α) (bs : List β) : zipLeft as bs = (zipLeft' as bs).fst := by + rw [zipLeft, zipLeft'] + cases as with + | nil => rfl + | cons _ atl => + cases bs with + | nil => rfl + | cons _ btl => + rw [zipWithLeft, zipWithLeft', cons_inj_right] + exact @zipLeft_eq_zipLeft' atl btl + +end ZipLeft + +/-! ### zipRight -/ + +section ZipRight + +variable (a : α) (as : List α) (b : β) (bs : List β) + +@[simp] +theorem zipRight_nil_left : zipRight ([] : List α) bs = bs.map fun b => (none, b) := by + cases bs <;> rfl + +@[simp] +theorem zipRight_nil_right : zipRight as ([] : List β) = [] := + rfl + +@[simp] +theorem zipRight_nil_cons : + zipRight ([] : List α) (b :: bs) = (none, b) :: bs.map fun b => (none, b) := + rfl + +@[simp] +theorem zipRight_cons_cons : zipRight (a :: as) (b :: bs) = (some a, b) :: zipRight as bs := + rfl + +theorem zipRight_eq_zipRight' : zipRight as bs = (zipRight' as bs).fst := by + induction as generalizing bs <;> cases bs <;> simp [*] + +end ZipRight + +end List diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index dfa426586cded..084e0e51bba63 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -122,7 +122,7 @@ theorem not_nodup_of_get_eq_of_ne (xs : List α) (n m : Fin xs.length) theorem indexOf_getElem [DecidableEq α] {l : List α} (H : Nodup l) (i : Nat) (h : i < l.length) : indexOf l[i] l = i := - suffices (⟨indexOf l[i] l, indexOf_lt_length.2 (getElem_mem _)⟩ : Fin l.length) = ⟨i, h⟩ + suffices (⟨indexOf l[i] l, indexOf_lt_length_iff.2 (getElem_mem _)⟩ : Fin l.length) = ⟨i, h⟩ from Fin.val_eq_of_eq this nodup_iff_injective_get.1 H (by simp) diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 9b381e748861c..0edcc7b2aba6c 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -50,7 +50,7 @@ the set of elements of `l`. -/ @[simps] def getEquiv (l : List α) (H : Nodup l) : Fin (length l) ≃ { x // x ∈ l } where toFun i := ⟨get l i, get_mem _ _⟩ - invFun x := ⟨indexOf (↑x) l, indexOf_lt_length.2 x.2⟩ + invFun x := ⟨indexOf (↑x) l, indexOf_lt_length_iff.2 x.2⟩ left_inv i := by simp only [List.get_indexOf, eq_self_iff_true, Fin.eta, Subtype.coe_mk, H] right_inv x := by simp @@ -63,7 +63,7 @@ decidable equality. -/ def getEquivOfForallMemList (l : List α) (nd : l.Nodup) (h : ∀ x : α, x ∈ l) : Fin l.length ≃ α where toFun i := l.get i - invFun a := ⟨_, indexOf_lt_length.2 (h a)⟩ + invFun a := ⟨_, indexOf_lt_length_iff.2 (h a)⟩ left_inv i := by simp [List.indexOf_getElem, nd] right_inv a := by simp diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 69638ec2666c4..d7eeeed30a977 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -49,6 +49,17 @@ theorem stdBasisMatrix_zero (i : m) (j : n) : stdBasisMatrix i j (0 : α) = 0 := ext simp +@[simp] +lemma transpose_stdBasisMatrix (i : m) (j : n) (a : α) : + (stdBasisMatrix i j a)ᵀ = stdBasisMatrix j i a := by + aesop (add unsafe unfold stdBasisMatrix) + +@[simp] +lemma map_stdBasisMatrix (i : m) (j : n) (a : α) {β : Type*} [Zero β] + {F : Type*} [FunLike F α β] [ZeroHomClass F α β] (f : F) : + (stdBasisMatrix i j a).map f = stdBasisMatrix i j (f a) := by + aesop (add unsafe unfold stdBasisMatrix) + end Zero theorem stdBasisMatrix_add [AddZeroClass α] (i : m) (j : n) (a b : α) : diff --git a/Mathlib/Data/Matrix/ConjTranspose.lean b/Mathlib/Data/Matrix/ConjTranspose.lean index 4b856340799ca..f43539add8aae 100644 --- a/Mathlib/Data/Matrix/ConjTranspose.lean +++ b/Mathlib/Data/Matrix/ConjTranspose.lean @@ -11,6 +11,7 @@ import Mathlib.Algebra.Star.BigOperators import Mathlib.Algebra.Star.Module import Mathlib.Algebra.Star.Pi import Mathlib.Data.Fintype.BigOperators +import Mathlib.Data.Matrix.Basis import Mathlib.Data.Matrix.Mul /-! @@ -40,6 +41,13 @@ def conjTranspose [Star α] (M : Matrix m n α) : Matrix n m α := @[inherit_doc] scoped postfix:1024 "ᴴ" => Matrix.conjTranspose +@[simp] +lemma conjTranspose_stdBasisMatrix [DecidableEq n] [DecidableEq m] [AddMonoid α] + [StarAddMonoid α] (i : m) (j : n) (a : α) : + (stdBasisMatrix i j a)ᴴ = stdBasisMatrix j i (star a) := by + show (stdBasisMatrix i j a).transpose.map starAddEquiv = stdBasisMatrix j i (star a) + simp + section Diagonal variable [DecidableEq n] diff --git a/Mathlib/Data/Matrix/DMatrix.lean b/Mathlib/Data/Matrix/DMatrix.lean index cf00fbd771964..94cc3254c7718 100644 --- a/Mathlib/Data/Matrix/DMatrix.lean +++ b/Mathlib/Data/Matrix/DMatrix.lean @@ -105,12 +105,7 @@ instance [∀ i j, Unique (α i j)] : Unique (DMatrix m n α) := instance [∀ i j, Subsingleton (α i j)] : Subsingleton (DMatrix m n α) := inferInstanceAs <| Subsingleton <| ∀ i j, α i j -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4481 -the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. --/ -@[simp, nolint simpNF] +@[simp] theorem zero_apply [∀ i j, Zero (α i j)] (i j) : (0 : DMatrix m n α) i j = 0 := rfl @[simp] diff --git a/Mathlib/Data/Matrix/Defs.lean b/Mathlib/Data/Matrix/Defs.lean index cee907c0593df..fb49c651c7630 100644 --- a/Mathlib/Data/Matrix/Defs.lean +++ b/Mathlib/Data/Matrix/Defs.lean @@ -212,12 +212,7 @@ instance module [Semiring R] [AddCommMonoid α] [Module R α] : Module R (Matrix section -#adaptation_note -/-- -After https://github.com/leanprover/lean4/pull/4481 -the `simpNF` linter incorrectly claims this lemma can't be applied by `simp`. --/ -@[simp, nolint simpNF] +@[simp] theorem zero_apply [Zero α] (i : m) (j : n) : (0 : Matrix m n α) i j = 0 := rfl @[simp] diff --git a/Mathlib/Data/Matrix/Diagonal.lean b/Mathlib/Data/Matrix/Diagonal.lean index 3f6042398a5d2..8eb3bcd89d1df 100644 --- a/Mathlib/Data/Matrix/Diagonal.lean +++ b/Mathlib/Data/Matrix/Diagonal.lean @@ -225,12 +225,9 @@ end Diagonal section Diag /-- The diagonal of a square matrix. -/ --- @[simp] -- Porting note: simpNF does not like this. def diag (A : Matrix n n α) (i : n) : α := A i i --- Porting note: new, because of removed `simp` above. --- TODO: set as an equation lemma for `diag`, see https://github.com/leanprover-community/mathlib4/pull/3024 @[simp] theorem diag_apply (A : Matrix n n α) (i) : diag A i = A i i := rfl diff --git a/Mathlib/Data/Matrix/Hadamard.lean b/Mathlib/Data/Matrix/Hadamard.lean index cdaa2360b95c5..4b8efd38c185e 100644 --- a/Mathlib/Data/Matrix/Hadamard.lean +++ b/Mathlib/Data/Matrix/Hadamard.lean @@ -34,7 +34,7 @@ variable {α m n R : Type*} namespace Matrix -/-- `Matrix.hadamard` defines the Hadamard product, +/-- `Matrix.hadamard` (denoted as `⊙` within the Matrix namespace) defines the Hadamard product, which is the pointwise product of two matrices of the same size. -/ def hadamard [Mul α] (A : Matrix m n α) (B : Matrix m n α) : Matrix m n α := of fun i j => A i j * B i j @@ -45,7 +45,7 @@ theorem hadamard_apply [Mul α] (A : Matrix m n α) (B : Matrix m n α) (i j) : hadamard A B i j = A i j * B i j := rfl -scoped infixl:100 " ⊙ " => Matrix.hadamard +@[inherit_doc] scoped infixl:100 " ⊙ " => Matrix.hadamard section BasicProperties diff --git a/Mathlib/Data/Matrix/Kronecker.lean b/Mathlib/Data/Matrix/Kronecker.lean index 525e8243ba190..31a238dac9c94 100644 --- a/Mathlib/Data/Matrix/Kronecker.lean +++ b/Mathlib/Data/Matrix/Kronecker.lean @@ -36,8 +36,8 @@ This defines the [Kronecker product](https://en.wikipedia.org/wiki/Kronecker_pro These require `open Kronecker`: * `A ⊗ₖ B` for `kroneckerMap (*) A B`. Lemmas about this notation use the token `kronecker`. -* `A ⊗ₖₜ B` and `A ⊗ₖₜ[R] B` for `kroneckerMap (⊗ₜ) A B`. Lemmas about this notation use the token - `kroneckerTMul`. +* `A ⊗ₖₜ B` and `A ⊗ₖₜ[R] B` for `kroneckerMap (⊗ₜ) A B`. + Lemmas about this notation use the token `kroneckerTMul`. -/ @@ -239,6 +239,7 @@ open Matrix def kronecker [Mul α] : Matrix l m α → Matrix n p α → Matrix (l × n) (m × p) α := kroneckerMap (· * ·) +@[inherit_doc Matrix.kroneckerMap] scoped[Kronecker] infixl:100 " ⊗ₖ " => Matrix.kroneckerMap (· * ·) open Kronecker @@ -417,10 +418,11 @@ Prefer the notation `⊗ₖₜ` rather than this definition. -/ def kroneckerTMul : Matrix l m α → Matrix n p β → Matrix (l × n) (m × p) (α ⊗[R] β) := kroneckerMap (· ⊗ₜ ·) +@[inherit_doc kroneckerTMul] scoped[Kronecker] infixl:100 " ⊗ₖₜ " => Matrix.kroneckerMap (· ⊗ₜ ·) -scoped[Kronecker] - notation:100 x " ⊗ₖₜ[" R "] " y:100 => Matrix.kroneckerMap (TensorProduct.tmul R) x y +@[inherit_doc kroneckerTMul] scoped[Kronecker] notation:100 x " ⊗ₖₜ[" R "] " y:100 => + Matrix.kroneckerMap (TensorProduct.tmul R) x y open Kronecker diff --git a/Mathlib/Data/Matrix/Mul.lean b/Mathlib/Data/Matrix/Mul.lean index 4f7140b712651..c074c0ef47fff 100644 --- a/Mathlib/Data/Matrix/Mul.lean +++ b/Mathlib/Data/Matrix/Mul.lean @@ -642,6 +642,13 @@ theorem mul_apply_eq_vecMul [Fintype n] (A : Matrix m n α) (B : Matrix n o α) (A * B) i = A i ᵥ* B := rfl +theorem vecMul_eq_sum [Fintype m] (v : m → α) (M : Matrix m n α) : v ᵥ* M = ∑ i, v i • M i := + (Finset.sum_fn ..).symm + +theorem mulVec_eq_sum [Fintype n] (v : n → α) (M : Matrix m n α) : + M *ᵥ v = ∑ i, MulOpposite.op (v i) • Mᵀ i := + (Finset.sum_fn ..).symm + theorem mulVec_diagonal [Fintype m] [DecidableEq m] (v w : m → α) (x : m) : (diagonal v *ᵥ w) x = v x * w x := diagonal_dotProduct v w x @@ -715,12 +722,12 @@ theorem mulVec_smul [Fintype n] [Monoid R] [NonUnitalNonAssocSemiring S] [Distri @[simp] theorem mulVec_single [Fintype n] [DecidableEq n] [NonUnitalNonAssocSemiring R] (M : Matrix m n R) - (j : n) (x : R) : M *ᵥ Pi.single j x = fun i => M i j * x := + (j : n) (x : R) : M *ᵥ Pi.single j x = MulOpposite.op x • Mᵀ j := funext fun _ => dotProduct_single _ _ _ @[simp] theorem single_vecMul [Fintype m] [DecidableEq m] [NonUnitalNonAssocSemiring R] (M : Matrix m n R) - (i : m) (x : R) : Pi.single i x ᵥ* M = fun j => x * M i j := + (i : m) (x : R) : Pi.single i x ᵥ* M = x • M i := funext fun _ => single_dotProduct _ _ _ theorem mulVec_single_one [Fintype n] [DecidableEq n] [NonAssocSemiring R] @@ -729,16 +736,14 @@ theorem mulVec_single_one [Fintype n] [DecidableEq n] [NonAssocSemiring R] theorem single_one_vecMul [Fintype m] [DecidableEq m] [NonAssocSemiring R] (i : m) (M : Matrix m n R) : - Pi.single i 1 ᵥ* M = M i := by simp + Pi.single i 1 ᵥ* M = M i := by ext; simp --- @[simp] -- Porting note: not in simpNF theorem diagonal_mulVec_single [Fintype n] [DecidableEq n] [NonUnitalNonAssocSemiring R] (v : n → R) (j : n) (x : R) : diagonal v *ᵥ Pi.single j x = Pi.single j (v j * x) := by ext i rw [mulVec_diagonal] exact Pi.apply_single (fun i x => v i * x) (fun i => mul_zero _) j x i --- @[simp] -- Porting note: not in simpNF theorem single_vecMul_diagonal [Fintype n] [DecidableEq n] [NonUnitalNonAssocSemiring R] (v : n → R) (j : n) (x : R) : (Pi.single j x) ᵥ* (diagonal v) = Pi.single j (x * v j) := by ext i @@ -776,12 +781,28 @@ section NonAssocSemiring variable [NonAssocSemiring α] -theorem mulVec_one [Fintype n] (A : Matrix m n α) : A *ᵥ 1 = fun i => ∑ j, A i j := by +theorem mulVec_one [Fintype n] (A : Matrix m n α) : A *ᵥ 1 = ∑ j, Aᵀ j := by ext; simp [mulVec, dotProduct] -theorem vec_one_mul [Fintype m] (A : Matrix m n α) : 1 ᵥ* A = fun j => ∑ i, A i j := by +theorem one_vecMul [Fintype m] (A : Matrix m n α) : 1 ᵥ* A = ∑ i, A i := by ext; simp [vecMul, dotProduct] +@[deprecated (since := "2025-01-26")] alias vec_one_mul := one_vecMul + +lemma ext_of_mulVec_single [DecidableEq n] [Fintype n] {M N : Matrix m n α} + (h : ∀ i, M *ᵥ Pi.single i 1 = N *ᵥ Pi.single i 1) : + M = N := by + ext i j + simp_rw [mulVec_single_one] at h + exact congrFun (h j) i + +lemma ext_of_single_vecMul [DecidableEq m] [Fintype m] {M N : Matrix m n α} + (h : ∀ i, Pi.single i 1 ᵥ* M = Pi.single i 1 ᵥ* N) : + M = N := by + ext i j + simp_rw [single_one_vecMul] at h + exact congrFun (h i) j + variable [Fintype m] [Fintype n] [DecidableEq m] @[simp] diff --git a/Mathlib/Data/Matrix/PEquiv.lean b/Mathlib/Data/Matrix/PEquiv.lean index eb70fede2430c..98cb355cc674a 100644 --- a/Mathlib/Data/Matrix/PEquiv.lean +++ b/Mathlib/Data/Matrix/PEquiv.lean @@ -40,7 +40,7 @@ open Matrix universe u v variable {k l m n : Type*} -variable {α : Type v} +variable {α β : Type*} open Matrix @@ -55,13 +55,28 @@ theorem toMatrix_apply [DecidableEq n] [Zero α] [One α] (f : m ≃. n) (i j) : toMatrix f i j = if j ∈ f i then (1 : α) else 0 := rfl -theorem mul_matrix_apply [Fintype m] [DecidableEq m] [Semiring α] (f : l ≃. m) (M : Matrix m n α) +theorem toMatrix_mul_apply [Fintype m] [DecidableEq m] [Semiring α] (f : l ≃. m) (M : Matrix m n α) (i j) : (f.toMatrix * M :) i j = Option.casesOn (f i) 0 fun fi => M fi j := by dsimp [toMatrix, Matrix.mul_apply] cases' h : f i with fi · simp [h] · rw [Finset.sum_eq_single fi] <;> simp +contextual [h, eq_comm] +@[deprecated (since := "2025-01-27")] alias mul_matrix_apply := toMatrix_mul_apply + +theorem mul_toMatrix_apply [Fintype m] [Semiring α] [DecidableEq n] (M : Matrix l m α) (f : m ≃. n) + (i j) : (M * f.toMatrix :) i j = Option.casesOn (f.symm j) 0 (M i) := by + dsimp [Matrix.mul_apply, toMatrix_apply] + cases' h : f.symm j with fj + · simp [h, ← f.eq_some_iff] + · rw [Finset.sum_eq_single fj] + · simp [h, ← f.eq_some_iff] + · rintro b - n + simp [h, ← f.eq_some_iff, n.symm] + · simp + +@[deprecated (since := "2025-01-27")] alias matrix_mul_apply := mul_toMatrix_apply + theorem toMatrix_symm [DecidableEq m] [DecidableEq n] [Zero α] [One α] (f : m ≃. n) : (f.symm.toMatrix : Matrix n m α) = f.toMatrixᵀ := by ext @@ -74,32 +89,53 @@ theorem toMatrix_refl [DecidableEq n] [Zero α] [One α] : ext simp [toMatrix_apply, one_apply] -theorem matrix_mul_apply [Fintype m] [Semiring α] [DecidableEq n] (M : Matrix l m α) (f : m ≃. n) - (i j) : (M * f.toMatrix :) i j = Option.casesOn (f.symm j) 0 fun fj => M i fj := by - dsimp [toMatrix, Matrix.mul_apply] - cases' h : f.symm j with fj - · simp [h, ← f.eq_some_iff] - · rw [Finset.sum_eq_single fj] - · simp [h, ← f.eq_some_iff] - · rintro b - n - simp [h, ← f.eq_some_iff, n.symm] - · simp +@[simp] +theorem toMatrix_toPEquiv_apply [DecidableEq n] [Zero α] [One α] (f : m ≃ n) (i) : + f.toPEquiv.toMatrix i = Pi.single (f i) (1 : α) := by + ext + simp [toMatrix_apply, Pi.single_apply, eq_comm] + +@[simp] +theorem transpose_toMatrix_toPEquiv_apply + [DecidableEq m] [DecidableEq n] [Zero α] [One α] (f : m ≃ n) (j) : + f.toPEquiv.toMatrixᵀ j = Pi.single (f.symm j) (1 : α) := by + ext + simp [toMatrix_apply, Pi.single_apply, eq_comm, ← Equiv.apply_eq_iff_eq_symm_apply] -theorem toPEquiv_mul_matrix [Fintype m] [DecidableEq m] [Semiring α] (f : m ≃ m) - (M : Matrix m n α) : f.toPEquiv.toMatrix * M = M.submatrix f id := by +theorem toMatrix_toPEquiv_mul [Fintype m] [DecidableEq m] + [Semiring α] (f : l ≃ m) (M : Matrix m n α) : + f.toPEquiv.toMatrix * M = M.submatrix f id := by ext i j - rw [mul_matrix_apply, Equiv.toPEquiv_apply, submatrix_apply, id] + rw [toMatrix_mul_apply, Equiv.toPEquiv_apply, submatrix_apply, id] -theorem mul_toPEquiv_toMatrix {m n α : Type*} [Fintype n] [DecidableEq n] [Semiring α] (f : n ≃ n) - (M : Matrix m n α) : M * f.toPEquiv.toMatrix = M.submatrix id f.symm := +@[deprecated (since := "2025-01-27")] alias toPEquiv_mul_matrix := toMatrix_toPEquiv_mul + +theorem mul_toMatrix_toPEquiv [Fintype m] [DecidableEq n] + [Semiring α] (M : Matrix l m α) (f : m ≃ n) : + (M * f.toPEquiv.toMatrix) = M.submatrix id f.symm := Matrix.ext fun i j => by - rw [PEquiv.matrix_mul_apply, ← Equiv.toPEquiv_symm, Equiv.toPEquiv_apply, + rw [PEquiv.mul_toMatrix_apply, ← Equiv.toPEquiv_symm, Equiv.toPEquiv_apply, Matrix.submatrix_apply, id] +@[deprecated (since := "2025-01-27")] alias mul_toPEquiv_toMatrix := mul_toMatrix_toPEquiv + +lemma toMatrix_toPEquiv_mulVec [DecidableEq n] [Fintype n] + [NonAssocSemiring α] (σ : m ≃ n) (a : n → α) : + σ.toPEquiv.toMatrix *ᵥ a = a ∘ σ := by + ext j + simp [toMatrix, mulVec, dotProduct] + +lemma vecMul_toMatrix_toPEquiv [DecidableEq n] [Fintype m] + [NonAssocSemiring α] (σ : m ≃ n) (a : m → α) : + a ᵥ* σ.toPEquiv.toMatrix = a ∘ σ.symm := by + classical + ext j + simp [toMatrix, σ.apply_eq_iff_eq_symm_apply, vecMul, dotProduct] + theorem toMatrix_trans [Fintype m] [DecidableEq m] [DecidableEq n] [Semiring α] (f : l ≃. m) (g : m ≃. n) : ((f.trans g).toMatrix : Matrix l n α) = f.toMatrix * g.toMatrix := by ext i j - rw [mul_matrix_apply] + rw [toMatrix_mul_apply] dsimp [toMatrix, PEquiv.trans] cases f i <;> simp @@ -110,19 +146,18 @@ theorem toMatrix_bot [DecidableEq n] [Zero α] [One α] : theorem toMatrix_injective [DecidableEq n] [MonoidWithZero α] [Nontrivial α] : Function.Injective (@toMatrix m n α _ _ _) := by - classical - intro f g - refine not_imp_not.1 ?_ - simp only [Matrix.ext_iff.symm, toMatrix_apply, PEquiv.ext_iff, not_forall, exists_imp] - intro i hi - use i - cases' hf : f i with fi - · cases' hg : g i with gi - · rw [hf, hg] at hi; exact (hi rfl).elim - · use gi - simp - · use fi - simp [hf.symm, Ne.symm hi] + intro f g + refine not_imp_not.1 ?_ + simp only [Matrix.ext_iff.symm, toMatrix_apply, PEquiv.ext_iff, not_forall, exists_imp] + intro i hi + use i + cases' hf : f i with fi + · cases' hg : g i with gi + · rw [hf, hg] at hi; exact (hi rfl).elim + · use gi + simp + · use fi + simp [hf.symm, Ne.symm hi] theorem toMatrix_swap [DecidableEq n] [Ring α] (i j : n) : (Equiv.swap i j).toPEquiv.toMatrix = @@ -152,8 +187,16 @@ theorem single_mul_single_right [Fintype n] [Fintype k] [DecidableEq n] [Decidab rw [← Matrix.mul_assoc, single_mul_single] /-- We can also define permutation matrices by permuting the rows of the identity matrix. -/ -theorem equiv_toPEquiv_toMatrix [DecidableEq n] [Zero α] [One α] (σ : Equiv n n) (i j : n) : - σ.toPEquiv.toMatrix i j = (1 : Matrix n n α) (σ i) j := - if_congr Option.some_inj rfl rfl +theorem toMatrix_toPEquiv_eq [DecidableEq n] [Zero α] [One α] (σ : Equiv.Perm n) : + σ.toPEquiv.toMatrix = (1 : Matrix n n α).submatrix σ id := + Matrix.ext fun _ _ => if_congr Option.some_inj rfl rfl + +@[deprecated (since := "2025-01-27")] alias equiv_toPEquiv_toMatrix := toMatrix_toPEquiv_eq + +@[simp] +lemma map_toMatrix [DecidableEq n] [NonAssocSemiring α] [NonAssocSemiring β] + (f : α →+* β) (σ : m ≃. n) : σ.toMatrix.map f = σ.toMatrix := by + ext i j + simp end PEquiv diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index 54558d4e57700..4bef987a38c32 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -217,8 +217,11 @@ namespace Matroid variable {α : Type*} {M : Matroid α} +instance (M : Matroid α) : Nonempty {B // M.Base B} := + nonempty_subtype.2 M.exists_base + /-- Typeclass for a matroid having finite ground set. Just a wrapper for `M.E.Finite`-/ -protected class Finite (M : Matroid α) : Prop where +@[mk_iff] protected class Finite (M : Matroid α) : Prop where /-- The ground set is finite -/ (ground_finite : M.E.Finite) @@ -233,6 +236,9 @@ theorem ground_nonempty (M : Matroid α) [M.Nonempty] : M.E.Nonempty := theorem ground_nonempty_iff (M : Matroid α) : M.E.Nonempty ↔ M.Nonempty := ⟨fun h ↦ ⟨h⟩, fun ⟨h⟩ ↦ h⟩ +lemma nonempty_type (M : Matroid α) [h : M.Nonempty] : Nonempty α := + ⟨M.ground_nonempty.some⟩ + theorem ground_finite (M : Matroid α) [M.Finite] : M.E.Finite := Finite.ground_finite @@ -243,7 +249,7 @@ instance finite_of_finite [Finite α] {M : Matroid α} : M.Finite := ⟨Set.toFinite _⟩ /-- A `FiniteRk` matroid is one whose bases are finite -/ -class FiniteRk (M : Matroid α) : Prop where +@[mk_iff] class FiniteRk (M : Matroid α) : Prop where /-- There is a finite base -/ exists_finite_base : ∃ B, M.Base B ∧ B.Finite @@ -251,17 +257,22 @@ instance finiteRk_of_finite (M : Matroid α) [M.Finite] : FiniteRk M := ⟨M.exists_base.imp (fun B hB ↦ ⟨hB, M.set_finite B (M.subset_ground _ hB)⟩)⟩ /-- An `InfiniteRk` matroid is one whose bases are infinite. -/ -class InfiniteRk (M : Matroid α) : Prop where +@[mk_iff] class InfiniteRk (M : Matroid α) : Prop where /-- There is an infinite base -/ exists_infinite_base : ∃ B, M.Base B ∧ B.Infinite /-- A `RkPos` matroid is one whose bases are nonempty. -/ -class RkPos (M : Matroid α) : Prop where +@[mk_iff] class RkPos (M : Matroid α) : Prop where /-- The empty set isn't a base -/ empty_not_base : ¬M.Base ∅ -theorem rkPos_iff_empty_not_base : M.RkPos ↔ ¬M.Base ∅ := - ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ +instance rkPos_nonempty {M : Matroid α} [M.RkPos] : M.Nonempty := by + obtain ⟨B, hB⟩ := M.exists_base + obtain rfl | ⟨e, heB⟩ := B.eq_empty_or_nonempty + · exact False.elim <| RkPos.empty_not_base hB + exact ⟨e, M.subset_ground B hB heB ⟩ + +@[deprecated (since := "2025-01-20")] alias rkPos_iff_empty_not_base := rkPos_iff section exchange namespace ExchangeProperty @@ -434,7 +445,7 @@ theorem Base.nonempty [RkPos M] (hB : M.Base B) : B.Nonempty := by rw [nonempty_iff_ne_empty]; rintro rfl; exact M.empty_not_base hB theorem Base.rkPos_of_nonempty (hB : M.Base B) (h : B.Nonempty) : M.RkPos := by - rw [rkPos_iff_empty_not_base] + rw [rkPos_iff] intro he obtain rfl := he.eq_of_subset_base hB (empty_subset B) simp at h @@ -636,6 +647,14 @@ theorem Base.exchange_base_of_indep' (hB : M.Base B) (he : e ∈ B) (hf : f ∉ rw [← insert_diff_singleton_comm hfe] at * exact hB.exchange_base_of_indep hf hI +lemma insert_base_of_insert_indep {M : Matroid α} {I : Set α} {e f : α} + (he : e ∉ I) (hf : f ∉ I) (heI : M.Base (insert e I)) (hfI : M.Indep (insert f I)) : + M.Base (insert f I) := by + obtain rfl | hef := eq_or_ne e f + · assumption + simpa [diff_singleton_eq_self he, hfI] + using heI.exchange_base_of_indep (e := e) (f := f) (by simp [hef.symm, hf]) + theorem Base.insert_dep (hB : M.Base B) (h : e ∈ M.E \ B) : M.Dep (insert e B) := by rw [← not_indep_iff (insert_subset h.1 hB.subset_ground)] exact h.2 ∘ (fun hi ↦ insert_eq_self.mp (hB.eq_of_subset_indep hi (subset_insert e B)).symm) @@ -705,7 +724,7 @@ lemma ext_base_indep {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) (hM₁ : /-- A `Finitary` matroid is one where a set is independent if and only if it all its finite subsets are independent, or equivalently a matroid whose circuits are finite. -/ -class Finitary (M : Matroid α) : Prop where +@[mk_iff] class Finitary (M : Matroid α) : Prop where /-- `I` is independent if all its finite subsets are independent. -/ indep_of_forall_finite : ∀ I, (∀ J, J ⊆ I → J.Finite → M.Indep J) → M.Indep I @@ -734,6 +753,43 @@ theorem existsMaximalSubsetProperty_indep (M : Matroid α) : end dep_indep +section copy + +/-- create a copy of `M : Matroid α` with independence and base predicates and ground set defeq +to supplied arguments that are provably equal to those of `M`. -/ +@[simps] def copy (M : Matroid α) (E : Set α) (Base Indep : Set α → Prop) + (hE : E = M.E) (hB : ∀ B, Base B ↔ M.Base B) (hI : ∀ I, Indep I ↔ M.Indep I) : Matroid α where + E := E + Base := Base + Indep := Indep + indep_iff' _ := by simp_rw [hI, hB, M.indep_iff] + exists_base := by + simp_rw [hB] + exact M.exists_base + base_exchange := by + simp_rw [show Base = M.Base from funext (by simp [hB])] + exact M.base_exchange + maximality := by + simp_rw [hE, show Indep = M.Indep from funext (by simp [hI])] + exact M.maximality + subset_ground := by + simp_rw [hE, hB] + exact M.subset_ground + +/-- create a copy of `M : Matroid α` with an independence predicate and ground set defeq +to supplied arguments that are provably equal to those of `M`. -/ +@[simps!] def copyIndep (M : Matroid α) (E : Set α) (Indep : Set α → Prop) + (hE : E = M.E) (h : ∀ I, Indep I ↔ M.Indep I) : Matroid α := + M.copy E M.Base Indep hE (fun _ ↦ Iff.rfl) h + +/-- create a copy of `M : Matroid α` with a base predicate and ground set defeq +to supplied arguments that are provably equal to those of `M`. -/ +@[simps!] def copyBase (M : Matroid α) (E : Set α) (Base : Set α → Prop) + (hE : E = M.E) (h : ∀ B, Base B ↔ M.Base B) : Matroid α := + M.copy E Base M.Indep hE h (fun _ ↦ Iff.rfl) + +end copy + section Basis /-- A Basis for a set `X ⊆ M.E` is a maximal independent subset of `X` diff --git a/Mathlib/Data/Matroid/Circuit.lean b/Mathlib/Data/Matroid/Circuit.lean new file mode 100644 index 0000000000000..cd068114bce61 --- /dev/null +++ b/Mathlib/Data/Matroid/Circuit.lean @@ -0,0 +1,322 @@ +/- +Copyright (c) 2025 Peter Nelson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Peter Nelson +-/ +import Mathlib.Data.Matroid.Closure + +/-! +# Matroid Circuits + +A `Circuit` of a matroid `M` is a minimal set `C` that is dependent in `M`. +A matroid is determined by its set of circuits, and often the circuits +offer a more compact description of a matroid than the collection of independent sets or bases. +In matroids arising from graphs, circuits correspond to graphical cycles. + +# Main Declarations + +* `Matroid.Circuit M C` means that `C` is minimally dependent in `M`. +* For an `Indep`endent set `I` whose closure contains an element `e ∉ I`, + `Matroid.fundCircuit M e I` is the unique circuit contained in `insert e I`. +* `Matroid.Indep.fundCircuit_circuit` states that `Matroid.fundCircuit M e I` is indeed a circuit. +* `Circuit.eq_fundCircuit_of_subset` states that `Matroid.fundCircuit M e I` is the + unique circuit contained in `insert e I`. +* `Matroid.dep_iff_superset_circuit` states that the dependent subsets of the ground set + are precisely those that contain a circuit. +* `Matroid.ext_circuit` : a matroid is determined by its collection of circuits. + +# Implementation Details + +Since `Matroid.fundCircuit M e I` is only sensible if `I` is independent and `e ∈ M.closure I \ I`, +to avoid hypotheses being explicitly included in the definition, +junk values need to be chosen if either hypothesis fails. +The definition is chosen so that the junk values satisfy +`M.fundCircuit e I = {e}` for `e ∈ I` and +`M.fundCircuit e I = insert e I` if `e ∉ M.closure I`. +These make the useful statement `e ∈ M.fundCircuit e I ⊆ insert e I` true unconditionally. +-/ + +variable {α : Type*} {M : Matroid α} {C C' I X Y R : Set α} {e f x y : α} + +open Set + +namespace Matroid + +/-- A `Circuit` of `M` is a minimal dependent set in `M` -/ +def Circuit (M : Matroid α) := Minimal M.Dep + +lemma circuit_def : M.Circuit C ↔ Minimal M.Dep C := Iff.rfl + +lemma Circuit.dep (hC : M.Circuit C) : M.Dep C := + hC.prop + +lemma Circuit.not_indep (hC : M.Circuit C) : ¬ M.Indep C := + hC.dep.not_indep + +lemma Circuit.minimal (hC : M.Circuit C) : Minimal M.Dep C := + hC + +@[aesop unsafe 20% (rule_sets := [Matroid])] +lemma Circuit.subset_ground (hC : M.Circuit C) : C ⊆ M.E := + hC.dep.subset_ground + +lemma Circuit.nonempty (hC : M.Circuit C) : C.Nonempty := + hC.dep.nonempty + +lemma empty_not_circuit (M : Matroid α) : ¬M.Circuit ∅ := + fun h ↦ by simpa using h.nonempty + +lemma circuit_iff : M.Circuit C ↔ M.Dep C ∧ ∀ ⦃D⦄, M.Dep D → D ⊆ C → D = C := by + simp_rw [circuit_def, minimal_subset_iff, eq_comm (a := C)] + +lemma Circuit.ssubset_indep (hC : M.Circuit C) (hXC : X ⊂ C) : M.Indep X := by + rw [← not_dep_iff (hXC.subset.trans hC.subset_ground)] + exact fun h ↦ hXC.ne ((circuit_iff.1 hC).2 h hXC.subset) + +lemma Circuit.minimal_not_indep (hC : M.Circuit C) : Minimal (¬ M.Indep ·) C := by + simp_rw [minimal_iff_forall_ssubset, and_iff_right hC.not_indep, not_not] + exact fun ⦃t⦄ a ↦ ssubset_indep hC a + +lemma circuit_iff_minimal_not_indep (hCE : C ⊆ M.E) : M.Circuit C ↔ Minimal (¬ M.Indep ·) C := + ⟨Circuit.minimal_not_indep, fun h ↦ ⟨(not_indep_iff hCE).1 h.prop, + fun _ hJ hJC ↦ (h.eq_of_superset hJ.not_indep hJC).le⟩⟩ + +lemma Circuit.diff_singleton_indep (hC : M.Circuit C) (he : e ∈ C) : M.Indep (C \ {e}) := + hC.ssubset_indep (diff_singleton_sSubset.2 he) + +lemma circuit_iff_forall_ssubset : M.Circuit C ↔ M.Dep C ∧ ∀ ⦃I⦄, I ⊂ C → M.Indep I := by + rw [Circuit, minimal_iff_forall_ssubset, and_congr_right_iff] + exact fun h ↦ ⟨fun h' I hIC ↦ ((not_dep_iff (hIC.subset.trans h.subset_ground)).1 (h' hIC)), + fun h I hIC ↦ (h hIC).not_dep⟩ + +lemma circuit_antichain : IsAntichain (· ⊆ ·) (setOf M.Circuit) := + fun _ hC _ hC' hne hss ↦ hne <| (Circuit.minimal hC').eq_of_subset hC.dep hss + +lemma Circuit.eq_of_not_indep_subset (hC : M.Circuit C) (hX : ¬ M.Indep X) (hXC : X ⊆ C) : + X = C := + eq_of_le_of_not_lt hXC (hX ∘ hC.ssubset_indep) + +lemma Circuit.eq_of_dep_subset (hC : M.Circuit C) (hX : M.Dep X) (hXC : X ⊆ C) : X = C := + hC.eq_of_not_indep_subset hX.not_indep hXC + +lemma Circuit.not_ssubset (hC : M.Circuit C) (hC' : M.Circuit C') : ¬C' ⊂ C := + fun h' ↦ h'.ne (hC.eq_of_dep_subset hC'.dep h'.subset) + +lemma Circuit.eq_of_subset_circuit (hC : M.Circuit C) (hC' : M.Circuit C') (h : C ⊆ C') : C = C' := + hC'.eq_of_dep_subset hC.dep h + +lemma Circuit.eq_of_superset_circuit (hC : M.Circuit C) (hC' : M.Circuit C') (h : C' ⊆ C) : + C = C' := + (hC'.eq_of_subset_circuit hC h).symm + +lemma circuit_iff_dep_forall_diff_singleton_indep : + M.Circuit C ↔ M.Dep C ∧ ∀ e ∈ C, M.Indep (C \ {e}) := by + wlog hCE : C ⊆ M.E + · exact iff_of_false (hCE ∘ Circuit.subset_ground) (fun h ↦ hCE h.1.subset_ground) + simp [circuit_iff_minimal_not_indep hCE, ← not_indep_iff hCE, + minimal_iff_forall_diff_singleton (P := (¬ M.Indep ·)) + (fun _ _ hY hYX hX ↦ hY <| hX.subset hYX)] + +/-! ### Independence and bases -/ + +lemma Indep.insert_circuit_of_forall (hI : M.Indep I) (heI : e ∉ I) (he : e ∈ M.closure I) + (h : ∀ f ∈ I, e ∉ M.closure (I \ {f})) : M.Circuit (insert e I) := by + rw [circuit_iff_dep_forall_diff_singleton_indep, hI.insert_dep_iff, and_iff_right ⟨he, heI⟩] + rintro f (rfl | hfI) + · simpa [heI] + rw [← insert_diff_singleton_comm (by rintro rfl; contradiction), + (hI.diff _).insert_indep_iff_of_not_mem (by simp [heI])] + exact ⟨mem_ground_of_mem_closure he, h f hfI⟩ + +lemma Indep.insert_circuit_of_forall_of_nontrivial (hI : M.Indep I) (hInt : I.Nontrivial) + (he : e ∈ M.closure I) (h : ∀ f ∈ I, e ∉ M.closure (I \ {f})) : M.Circuit (insert e I) := by + refine hI.insert_circuit_of_forall (fun heI ↦ ?_) he h + obtain ⟨f, hf, hne⟩ := hInt.exists_ne e + exact h f hf (mem_closure_of_mem' _ (by simp [heI, hne.symm])) + +lemma Circuit.diff_singleton_basis (hC : M.Circuit C) (he : e ∈ C) : M.Basis (C \ {e}) C := by + nth_rw 2 [← insert_eq_of_mem he] + rw [← insert_diff_singleton, (hC.diff_singleton_indep he).basis_insert_iff, + insert_diff_singleton, insert_eq_of_mem he] + exact Or.inl hC.dep + +lemma Circuit.basis_iff_eq_diff_singleton (hC : M.Circuit C) : + M.Basis I C ↔ ∃ e ∈ C, I = C \ {e} := by + refine ⟨fun h ↦ ?_, ?_⟩ + · obtain ⟨e, he⟩ := exists_of_ssubset + (h.subset.ssubset_of_ne (by rintro rfl; exact hC.dep.not_indep h.indep)) + exact ⟨e, he.1, h.eq_of_subset_indep (hC.diff_singleton_indep he.1) + (subset_diff_singleton h.subset he.2) diff_subset⟩ + rintro ⟨e, he, rfl⟩ + exact hC.diff_singleton_basis he + +lemma Circuit.basis_iff_insert_eq (hC : M.Circuit C) : + M.Basis I C ↔ ∃ e ∈ C \ I, C = insert e I := by + rw [hC.basis_iff_eq_diff_singleton] + refine ⟨fun ⟨e, he, hI⟩ ↦ ⟨e, ⟨he, fun heI ↦ (hI.subset heI).2 rfl⟩, ?_⟩, + fun ⟨e, he, hC⟩ ↦ ⟨e, he.1, ?_⟩⟩ + · rw [hI, insert_diff_singleton, insert_eq_of_mem he] + rw [hC, insert_diff_self_of_not_mem he.2] + +/-! ### Closure -/ + +lemma Circuit.closure_diff_singleton_eq (hC : M.Circuit C) (e : α) : + M.closure (C \ {e}) = M.closure C := + (em (e ∈ C)).elim + (fun he ↦ by rw [(hC.diff_singleton_basis he).closure_eq_closure]) + (fun he ↦ by rw [diff_singleton_eq_self he]) + +lemma Circuit.subset_closure_diff_singleton (hC : M.Circuit C) (e : α) : + C ⊆ M.closure (C \ {e}) := by + rw [hC.closure_diff_singleton_eq] + exact M.subset_closure _ hC.subset_ground + +lemma Circuit.mem_closure_diff_singleton_of_mem (hC : M.Circuit C) (heC : e ∈ C) : + e ∈ M.closure (C \ {e}) := + (hC.subset_closure_diff_singleton e) heC + +/-! ### Restriction -/ + +lemma Circuit.circuit_restrict_of_subset (hC : M.Circuit C) (hCR : C ⊆ R) : (M ↾ R).Circuit C := by + simp_rw [circuit_iff, restrict_dep_iff, dep_iff, and_imp] at * + exact ⟨⟨hC.1.1, hCR⟩, fun I hI _ hIC ↦ hC.2 hI (hIC.trans hC.1.2) hIC⟩ + +lemma restrict_circuit_iff (hR : R ⊆ M.E := by aesop_mat) : + (M ↾ R).Circuit C ↔ M.Circuit C ∧ C ⊆ R := by + refine ⟨?_, fun h ↦ h.1.circuit_restrict_of_subset h.2⟩ + simp_rw [circuit_iff, restrict_dep_iff, and_imp, dep_iff] + exact fun hC hCR h ↦ ⟨⟨⟨hC,hCR.trans hR⟩,fun I hI hIC ↦ h hI.1 (hIC.trans hCR) hIC⟩,hCR⟩ + +/-! ### Fundamental Circuits -/ + +/-- For an independent set `I` and some `e ∈ M.closure I \ I`, +`M.fundCircuit e I` is the unique circuit contained in `insert e I`. +For the fact that this is a circuit, see `Matroid.Indep.fundCircuit_circuit`, +and the fact that it is unique, see `Matroid.Circuit.eq_fundCircuit_of_subset`. +Has the junk value `{e}` if `e ∈ I` and `insert e I` if `e ∉ M.closure I`. -/ +def fundCircuit (M : Matroid α) (e : α) (I : Set α) : Set α := + insert e (I ∩ (⋂₀ {J | J ⊆ I ∧ e ∈ M.closure J})) + +lemma fundCircuit_eq_sInter (he : e ∈ M.closure I) : + M.fundCircuit e I = insert e (⋂₀ {J | J ⊆ I ∧ e ∈ M.closure J}) := by + rw [fundCircuit, inter_eq_self_of_subset_right] + exact sInter_subset_of_mem ⟨Subset.rfl, he⟩ + +lemma fundCircuit_subset_insert (M : Matroid α) (e : α) (I : Set α) : + M.fundCircuit e I ⊆ insert e I := + insert_subset_insert inter_subset_left + +lemma fundCircuit_subset_ground (he : e ∈ M.E) (hI : I ⊆ M.E := by aesop_mat) : + M.fundCircuit e I ⊆ M.E := + (M.fundCircuit_subset_insert e I).trans (insert_subset he hI) + +lemma mem_fundCircuit (M : Matroid α) (e : α) (I : Set α) : e ∈ fundCircuit M e I := + mem_insert .. + +/-- The fundamental circuit of `e` and `X` has the junk value `{e}` if `e ∈ X` -/ +lemma fundCircuit_eq_of_mem (heX : e ∈ X) (heE : e ∈ M.E := by aesop_mat) : + M.fundCircuit e X = {e} := by + suffices h : ∀ a ∈ X, (∀ I ⊆ X, e ∈ M.closure I → a ∈ I) → a = e by + simpa [subset_antisymm_iff, fundCircuit] + exact fun f hfX h ↦ h {e} (by simpa) (mem_closure_of_mem' _ rfl) + +/-- A version of `Matroid.fundCircuit_eq_of_mem` that applies when `X ⊆ M.E` instead of `e ∈ X`.-/ +lemma fundCircuit_eq_of_mem' (heX : e ∈ X) (hX : X ⊆ M.E := by aesop_mat) : + M.fundCircuit e X = {e} := by + rwa [fundCircuit_eq_of_mem] + +lemma Indep.fundCircuit_circuit (hI : M.Indep I) (hecl : e ∈ M.closure I) (heI : e ∉ I) : + M.Circuit (M.fundCircuit e I) := by + apply (hI.inter_right _).insert_circuit_of_forall (by simp [heI]) + · rw [(hI.subset _).closure_inter_eq_inter_closure, mem_inter_iff, and_iff_right hecl, + hI.closure_sInter_eq_biInter_closure_of_forall_subset _ (by simp +contextual)] + · simp + · exact ⟨I, rfl.subset, hecl⟩ + exact union_subset rfl.subset (sInter_subset_of_mem ⟨rfl.subset, hecl⟩) + simp only [mem_inter_iff, mem_sInter, mem_setOf_eq, and_imp] + exact fun f hfI hf hecl ↦ (hf _ (diff_subset.trans inter_subset_left) hecl).2 rfl + +lemma Indep.mem_fundCircuit_iff (hI : M.Indep I) (hecl : e ∈ M.closure I) (heI : e ∉ I) : + x ∈ M.fundCircuit e I ↔ M.Indep (insert e I \ {x}) := by + obtain rfl | hne := eq_or_ne x e + · simp [hI.diff, mem_fundCircuit] + suffices (∀ t ⊆ I, e ∈ M.closure t → x ∈ t) ↔ e ∉ M.closure (I \ {x}) by + simpa [fundCircuit_eq_sInter hecl, hne, ← insert_diff_singleton_comm hne.symm, + (hI.diff _).insert_indep_iff, mem_ground_of_mem_closure hecl, heI] + refine ⟨fun h hecl ↦ (h _ diff_subset hecl).2 rfl, fun h J hJ heJ ↦ by_contra fun hxJ ↦ h ?_⟩ + exact M.closure_subset_closure (subset_diff_singleton hJ hxJ) heJ + +lemma Base.fundCircuit_circuit {B : Set α} (hB : M.Base B) (hxE : x ∈ M.E) (hxB : x ∉ B) : + M.Circuit (M.fundCircuit x B) := + hB.indep.fundCircuit_circuit (by rwa [hB.closure_eq]) hxB + +/-- For `I` independent, `M.fundCircuit e I` is the only circuit contained in `insert e I`. -/ +lemma Circuit.eq_fundCircuit_of_subset (hC : M.Circuit C) (hI : M.Indep I) (hCss : C ⊆ insert e I) : + C = M.fundCircuit e I := by + obtain hCI | ⟨heC, hCeI⟩ := subset_insert_iff.1 hCss + · exact (hC.not_indep (hI.subset hCI)).elim + suffices hss : M.fundCircuit e I ⊆ C by + refine hC.eq_of_superset_circuit (hI.fundCircuit_circuit ?_ fun heI ↦ ?_) hss + · rw [hI.mem_closure_iff] + exact .inl (hC.dep.superset hCss (insert_subset (hC.subset_ground heC) hI.subset_ground)) + exact hC.not_indep (hI.subset (hCss.trans (by simp [heI]))) + refine insert_subset heC (inter_subset_right.trans ?_) + refine (sInter_subset_of_mem (t := C \ {e}) ?_).trans diff_subset + simp [hCss, hC.mem_closure_diff_singleton_of_mem heC] + +/-! ### Dependence -/ + +lemma Dep.exists_circuit_subset (hX : M.Dep X) : ∃ C, C ⊆ X ∧ M.Circuit C := by + obtain ⟨I, hI⟩ := M.exists_basis X + obtain ⟨e, heX, heI⟩ := exists_of_ssubset + (hI.subset.ssubset_of_ne (by rintro rfl; exact hI.indep.not_dep hX)) + exact ⟨M.fundCircuit e I, (M.fundCircuit_subset_insert e I).trans (insert_subset heX hI.subset), + hI.indep.fundCircuit_circuit (hI.subset_closure heX) heI⟩ + +lemma dep_iff_superset_circuit (hX : X ⊆ M.E := by aesop_mat) : + M.Dep X ↔ ∃ C, C ⊆ X ∧ M.Circuit C := + ⟨Dep.exists_circuit_subset, fun ⟨C, hCX, hC⟩ ↦ hC.dep.superset hCX⟩ + +/-- A version of `Matroid.dep_iff_superset_circuit` that has the supportedness hypothesis +as part of the equivalence, rather than a hypothesis. -/ +lemma dep_iff_superset_circuit' : M.Dep X ↔ (∃ C, C ⊆ X ∧ M.Circuit C) ∧ X ⊆ M.E := + ⟨fun h ↦ ⟨h.exists_circuit_subset, h.subset_ground⟩, fun ⟨⟨C, hCX, hC⟩, h⟩ ↦ hC.dep.superset hCX⟩ + +/-- A version of `Matroid.indep_iff_forall_subset_not_circuit` that has the supportedness hypothesis +as part of the equivalence, rather than a hypothesis. -/ +lemma indep_iff_forall_subset_not_circuit' : + M.Indep I ↔ (∀ C, C ⊆ I → ¬M.Circuit C) ∧ I ⊆ M.E := by + simp_rw [indep_iff_not_dep, dep_iff_superset_circuit'] + aesop + +lemma indep_iff_forall_subset_not_circuit (hI : I ⊆ M.E := by aesop_mat) : + M.Indep I ↔ ∀ C, C ⊆ I → ¬M.Circuit C := by + rw [indep_iff_forall_subset_not_circuit', and_iff_left hI] + +/-! ### Extensionality -/ + +lemma ext_circuit {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) + (h : ∀ ⦃C⦄, C ⊆ M₁.E → (M₁.Circuit C ↔ M₂.Circuit C)) : M₁ = M₂ := by + have h' {C} : M₁.Circuit C ↔ M₂.Circuit C := + (em (C ⊆ M₁.E)).elim (h (C := C)) (fun hC ↦ iff_of_false (mt Circuit.subset_ground hC) + (mt Circuit.subset_ground fun hss ↦ hC (hss.trans_eq hE.symm))) + refine ext_indep hE fun I hI ↦ ?_ + simp_rw [indep_iff_forall_subset_not_circuit hI, h', + indep_iff_forall_subset_not_circuit (hI.trans_eq hE)] + +/-- A stronger version of `Matroid.ext_circuit`: +two matroids on the same ground set are equal if no circuit of one is independent in the other. -/ +lemma ext_circuit_not_indep {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) + (h₁ : ∀ C, M₁.Circuit C → ¬ M₂.Indep C) (h₂ : ∀ C, M₂.Circuit C → ¬ M₁.Indep C) : + M₁ = M₂ := by + refine ext_circuit hE fun C hCE ↦ ⟨fun hC ↦ ?_, fun hC ↦ ?_⟩ + · obtain ⟨C', hC'C, hC'⟩ := ((not_indep_iff (by rwa [← hE])).1 (h₁ C hC)).exists_circuit_subset + rwa [← hC.eq_of_not_indep_subset (h₂ C' hC') hC'C] + obtain ⟨C', hC'C, hC'⟩ := ((not_indep_iff hCE).1 (h₂ C hC)).exists_circuit_subset + rwa [← hC.eq_of_not_indep_subset (h₁ C' hC') hC'C] + +lemma ext_iff_circuit {M₁ M₂ : Matroid α} : + M₁ = M₂ ↔ M₁.E = M₂.E ∧ ∀ C, M₁.Circuit C ↔ M₂.Circuit C := + ⟨fun h ↦ by simp [h], fun h ↦ ext_circuit h.1 fun C hC ↦ h.2 (C := C)⟩ + +end Matroid diff --git a/Mathlib/Data/Matroid/Closure.lean b/Mathlib/Data/Matroid/Closure.lean index 4dec4fadae6e9..f00a1faa7dc23 100644 --- a/Mathlib/Data/Matroid/Closure.lean +++ b/Mathlib/Data/Matroid/Closure.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ -import Mathlib.Data.Matroid.Restrict +import Mathlib.Data.Matroid.Map import Mathlib.Order.Closure import Mathlib.Order.CompleteLatticeIntervals @@ -270,6 +270,18 @@ lemma closure_closure_union_closure_eq_closure_union (M : Matroid α) (X Y : Set M.closure (insert e (M.closure X)) = M.closure (insert e X) := by simp_rw [← singleton_union, closure_union_closure_right_eq] +lemma closure_union_congr_left {X' : Set α} (h : M.closure X = M.closure X') : + M.closure (X ∪ Y) = M.closure (X' ∪ Y) := by + rw [← M.closure_union_closure_left_eq, h, M.closure_union_closure_left_eq] + +lemma closure_union_congr_right {Y' : Set α} (h : M.closure Y = M.closure Y') : + M.closure (X ∪ Y) = M.closure (X ∪ Y') := by + rw [← M.closure_union_closure_right_eq, h, M.closure_union_closure_right_eq] + +lemma closure_insert_congr_right (h : M.closure X = M.closure Y) : + M.closure (insert e X) = M.closure (insert e Y) := by + simp [← union_singleton, closure_union_congr_left h] + @[simp] lemma closure_union_closure_empty_eq (M : Matroid α) (X : Set α) : M.closure X ∪ M.closure ∅ = M.closure X := union_eq_self_of_subset_right (M.closure_subset_closure (empty_subset _)) @@ -485,6 +497,27 @@ lemma Indep.closure_inter_eq_inter_closure (h : M.Indep (I ∪ J)) : · exact iInter_congr (by simp) rwa [← union_eq_iUnion] +lemma Indep.inter_Basis_biInter {ι : Type*} (hI : M.Indep I) {X : ι → Set α} {A : Set ι} + (hA : A.Nonempty) (h : ∀ i ∈ A, M.Basis ((X i) ∩ I) (X i)) : + M.Basis ((⋂ i ∈ A, X i) ∩ I) (⋂ i ∈ A, X i) := by + refine (hI.inter_left _).basis_of_subset_of_subset_closure inter_subset_left ?_ + simp_rw [← biInter_inter hA, + closure_biInter_eq_biInter_closure_of_biUnion_indep hA (I := fun i ↦ (X i) ∩ I) + (hI.subset (by simp)), subset_iInter_iff] + exact fun i hiA ↦ (biInter_subset_of_mem hiA).trans (h i hiA).subset_closure + +lemma Indep.inter_Basis_iInter [Nonempty ι] {X : ι → Set α} (hI : M.Indep I) + (h : ∀ i, M.Basis ((X i) ∩ I) (X i)) : M.Basis ((⋂ i, X i) ∩ I) (⋂ i, X i) := by + convert hI.inter_Basis_biInter (ι := PLift ι) univ_nonempty (X := fun i ↦ X i.down) + (by simpa using fun (i : PLift ι) ↦ h i.down) <;> + · simp only [mem_univ, iInter_true] + exact (iInter_plift_down X).symm + +lemma Indep.inter_Basis_sInter {Xs : Set (Set α)} (hI : M.Indep I) (hXs : Xs.Nonempty) + (h : ∀ X ∈ Xs, M.Basis (X ∩ I) X) : M.Basis (⋂₀ Xs ∩ I) (⋂₀ Xs) := by + rw [sInter_eq_biInter] + exact hI.inter_Basis_biInter hXs h + lemma basis_iff_basis_closure_of_subset (hIX : I ⊆ X) (hX : X ⊆ M.E := by aesop_mat) : M.Basis I X ↔ M.Basis I (M.closure X) := ⟨fun h ↦ h.basis_closure_right, fun h ↦ h.basis_subset hIX (M.subset_closure X hX)⟩ @@ -519,11 +552,25 @@ lemma basis_iff_indep_closure : M.Basis I X ↔ M.Indep I ∧ X ⊆ M.closure I ⟨fun h ↦ ⟨h.indep, h.subset_closure, h.subset⟩, fun h ↦ (basis_union_iff_indep_closure.mpr ⟨h.1, h.2.1⟩).basis_subset h.2.2 subset_union_right⟩ +lemma Indep.inter_basis_closure_iff_subset_closure_inter {X : Set α} (hI : M.Indep I) : + M.Basis (X ∩ I) X ↔ X ⊆ M.closure (X ∩ I) := + ⟨Basis.subset_closure, (hI.inter_left X).basis_of_subset_of_subset_closure inter_subset_left⟩ + +lemma Basis.closure_inter_basis_closure (h : M.Basis (X ∩ I) X) (hI : M.Indep I) : + M.Basis (M.closure X ∩ I) (M.closure X) := by + rw [hI.inter_basis_closure_iff_subset_closure_inter] at h ⊢ + exact (M.closure_subset_closure_of_subset_closure h).trans (M.closure_subset_closure + (inter_subset_inter_left _ (h.trans (M.closure_subset_closure inter_subset_left)))) + lemma Basis.eq_of_closure_subset (hI : M.Basis I X) (hJI : J ⊆ I) (hJ : X ⊆ M.closure J) : J = I := by rw [← hI.indep.closure_inter_eq_self_of_subset hJI, inter_eq_self_of_subset_right] exact hI.subset.trans hJ +lemma Basis.insert_basis_insert_of_not_mem_closure (hIX : M.Basis I X) (heI : e ∉ M.closure I) + (heE : e ∈ M.E := by aesop_mat) : M.Basis (insert e I) (insert e X) := + hIX.insert_basis_insert <| hIX.indep.insert_indep_iff.2 <| .inl ⟨heE, heI⟩ + @[simp] lemma empty_basis_iff : M.Basis ∅ X ↔ X ⊆ M.closure ∅ := by rw [basis_iff_indep_closure, and_iff_right M.empty_indep, and_iff_left (empty_subset _)] @@ -557,6 +604,26 @@ lemma indep_iff_forall_closure_diff_ne : nth_rw 2 [show I = insert e (I \ {e}) by simp [heI]] rw [← closure_insert_closure_eq_closure_insert, insert_eq_of_mem hin, closure_closure] +lemma Indep.union_indep_iff_forall_not_mem_closure_right (hI : M.Indep I) (hJ : M.Indep J) : + M.Indep (I ∪ J) ↔ ∀ e ∈ J \ I, e ∉ M.closure (I ∪ (J \ {e})) := by + refine ⟨fun h e heJ hecl ↦ h.not_mem_closure_diff_of_mem (.inr heJ.1) ?_, fun h ↦ ?_⟩ + · rwa [union_diff_distrib, diff_singleton_eq_self heJ.2] + obtain ⟨K, hKIJ, hK⟩ := hI.subset_basis_of_subset (show I ⊆ I ∪ J from subset_union_left) + obtain rfl | hssu := hKIJ.subset.eq_or_ssubset + · exact hKIJ.indep + exfalso + obtain ⟨e, heI, heK⟩ := exists_of_ssubset hssu + have heJI : e ∈ J \ I := by + rw [← union_diff_right, union_comm] + exact ⟨heI, not_mem_subset hK heK⟩ + refine h _ heJI ?_ + rw [← diff_singleton_eq_self heJI.2, ← union_diff_distrib] + exact M.closure_subset_closure (subset_diff_singleton hKIJ.subset heK) <| hKIJ.subset_closure heI + +lemma Indep.union_indep_iff_forall_not_mem_closure_left (hI : M.Indep I) (hJ : M.Indep J) : + M.Indep (I ∪ J) ↔ ∀ e ∈ I \ J, e ∉ M.closure ((I \ {e}) ∪ J) := by + simp_rw [union_comm I J, hJ.union_indep_iff_forall_not_mem_closure_right hI, union_comm] + lemma Indep.closure_ssubset_closure (hI : M.Indep I) (hJI : J ⊂ I) : M.closure J ⊂ M.closure I := by obtain ⟨e, heI, heJ⟩ := exists_of_ssubset hJI exact (M.closure_subset_closure hJI.subset).ssubset_of_not_subset fun hss ↦ heJ <| @@ -740,6 +807,35 @@ lemma Indep.base_of_spanning (hI : M.Indep I) (hIs : M.Spanning I) : M.Base I := lemma Spanning.base_of_indep (hIs : M.Spanning I) (hI : M.Indep I) : M.Base I := hI.base_of_spanning hIs +lemma Indep.eq_of_spanning_subset (hI : M.Indep I) (hS : M.Spanning S) (hSI : S ⊆ I) : S = I := + ((hI.subset hSI).base_of_spanning hS).eq_of_subset_indep hI hSI + +lemma Basis.spanning_iff_spanning (hIX : M.Basis I X) : M.Spanning I ↔ M.Spanning X := by + rw [spanning_iff_closure_eq, spanning_iff_closure_eq, hIX.closure_eq_closure] + +lemma Spanning.base_restrict_iff (hS : M.Spanning S) : (M ↾ S).Base B ↔ M.Base B ∧ B ⊆ S := by + rw [base_restrict_iff', basis'_iff_basis] + refine ⟨fun h ↦ ⟨?_, h.subset⟩, fun h ↦ h.1.indep.basis_of_subset_of_subset_closure h.2 ?_⟩ + · exact h.indep.base_of_spanning <| by rwa [h.spanning_iff_spanning] + rw [h.1.closure_eq] + exact hS.subset_ground + +lemma Spanning.compl_coindep (hS : M.Spanning S) : M.Coindep (M.E \ S) := by + rwa [← spanning_iff_compl_coindep] + +lemma Basis.base_of_spanning (hIX : M.Basis I X) (hX : M.Spanning X) : M.Base I := + hIX.indep.base_of_spanning <| by rwa [hIX.spanning_iff_spanning] + +lemma Indep.exists_base_subset_spanning (hI : M.Indep I) (hS : M.Spanning S) (hIS : I ⊆ S) : + ∃ B, M.Base B ∧ I ⊆ B ∧ B ⊆ S := by + obtain ⟨B, hB⟩ := hI.subset_basis_of_subset hIS + exact ⟨B, hB.1.base_of_spanning hS, hB.2, hB.1.subset⟩ + +lemma Restriction.base_iff_of_spanning {N : Matroid α} (hR : N ≤r M) (hN : M.Spanning N.E) : + N.Base B ↔ (M.Base B ∧ B ⊆ N.E) := by + obtain ⟨R, hR : R ⊆ M.E, rfl⟩ := hR + rw [Spanning.base_restrict_iff (show M.Spanning R from hN), restrict_ground_eq] + lemma ext_spanning {M M' : Matroid α} (h : M.E = M'.E) (hsp : ∀ S, S ⊆ M.E → (M.Spanning S ↔ M'.Spanning S )) : M = M' := by have hsp' : M.Spanning = M'.Spanning := by @@ -753,5 +849,98 @@ lemma ext_spanning {M M' : Matroid α} (h : M.E = M'.E) end Spanning +section Constructions + +variable {R S : Set α} + +@[simp] lemma restrict_closure_eq' (M : Matroid α) (X R : Set α) : + (M ↾ R).closure X = (M.closure (X ∩ R) ∩ R) ∪ (R \ M.E) := by + obtain ⟨I, hI⟩ := (M ↾ R).exists_basis' X + obtain ⟨hI', hIR⟩ := basis'_restrict_iff.1 hI + ext e + rw [← hI.closure_eq_closure, ← hI'.closure_eq_closure, hI.indep.mem_closure_iff', mem_union, + mem_inter_iff, hI'.indep.mem_closure_iff', restrict_ground_eq, restrict_indep_iff, mem_diff] + by_cases he : M.Indep (insert e I) + · simp [he, and_comm, insert_subset_iff, hIR, (he.subset_ground (mem_insert ..)), imp_or] + tauto + +lemma restrict_closure_eq (M : Matroid α) (hXR : X ⊆ R) (hR : R ⊆ M.E := by aesop_mat) : + (M ↾ R).closure X = M.closure X ∩ R := by + rw [restrict_closure_eq', diff_eq_empty.mpr hR, union_empty, inter_eq_self_of_subset_left hXR] + +@[simp] lemma emptyOn_closure_eq (X : Set α) : (emptyOn α).closure X = ∅ := + (closure_subset_ground ..).antisymm <| empty_subset _ + +@[simp] lemma loopyOn_closure_eq (E X : Set α) : (loopyOn E).closure X = E := by + simp [loopyOn, restrict_closure_eq'] + +@[simp] lemma loopyOn_spanning_iff {E : Set α} : (loopyOn E).Spanning X ↔ X ⊆ E := by + rw [spanning_iff, loopyOn_closure_eq, loopyOn_ground, and_iff_right rfl] + +@[simp] lemma freeOn_closure_eq (E X : Set α) : (freeOn E).closure X = X ∩ E := by + simp (config := {contextual := true}) [← closure_inter_ground _ X, Set.ext_iff, and_comm, + insert_subset_iff, freeOn_indep_iff, (freeOn_indep inter_subset_right).mem_closure_iff'] + +@[simp] lemma uniqueBaseOn_closure_eq (I E X : Set α) : + (uniqueBaseOn I E).closure X = (X ∩ I ∩ E) ∪ (E \ I) := by + rw [uniqueBaseOn, restrict_closure_eq', freeOn_closure_eq, inter_right_comm, + inter_assoc (c := E), inter_self, inter_right_comm, freeOn_ground] + +lemma closure_empty_eq_ground_iff : M.closure ∅ = M.E ↔ M = loopyOn M.E := by + refine ⟨fun h ↦ ext_closure ?_, fun h ↦ by rw [h, loopyOn_closure_eq, loopyOn_ground]⟩ + refine fun X ↦ subset_antisymm (by simp [closure_subset_ground]) ?_ + rw [loopyOn_closure_eq, ← h] + exact M.closure_mono (empty_subset _) + +@[simp] lemma comap_closure_eq {β : Type*} (M : Matroid β) (f : α → β) (X : Set α) : + (M.comap f).closure X = f ⁻¹' M.closure (f '' X) := by + -- Use a choice of basis and extensionality to change the goal to a statement about independence. + obtain ⟨I, hI⟩ := (M.comap f).exists_basis' X + obtain ⟨hI', hIinj, -⟩ := comap_basis'_iff.1 hI + simp_rw [← hI.closure_eq_closure, ← hI'.closure_eq_closure, Set.ext_iff, + hI.indep.mem_closure_iff', comap_ground_eq, mem_preimage, hI'.indep.mem_closure_iff', + comap_indep_iff, and_imp, mem_image, and_congr_right_iff, ← image_insert_eq] + -- the lemma now easily follows by considering elements/non-elements of `I` separately. + intro x hxE + by_cases hxI : x ∈ I + · simp [hxI, show ∃ y ∈ I, f y = f x from ⟨x, hxI, rfl⟩] + simp [hxI, injOn_insert hxI, hIinj] + +@[simp] lemma map_closure_eq {β : Type*} (M : Matroid α) (f : α → β) (hf) (X : Set β) : + (M.map f hf).closure X = f '' M.closure (f ⁻¹' X) := by + -- It is enough to prove that `map` and `closure` commute for `M`-independent sets. + suffices aux : ∀ ⦃I⦄, M.Indep I → (M.map f hf).closure (f '' I) = f '' (M.closure I) by + obtain ⟨I, hI⟩ := M.exists_basis (f ⁻¹' X ∩ M.E) + rw [← closure_inter_ground, map_ground, ← M.closure_inter_ground, ← hI.closure_eq_closure, + ← aux hI.indep, ← image_preimage_inter, ← (hI.map hf).closure_eq_closure] + -- Let `I` be independent, and transform the goal using closure/independence lemmas + refine fun I hI ↦ Set.ext fun e ↦ ?_ + simp only [(hI.map f hf).mem_closure_iff', map_ground, mem_image, map_indep_iff, + forall_exists_index, and_imp, hI.mem_closure_iff'] + -- The goal now easily follows from the invariance of independence under maps. + constructor + · rintro ⟨⟨x, hxE, rfl⟩, h2⟩ + refine ⟨x, ⟨hxE, fun hI' ↦ ?_⟩, rfl⟩ + obtain ⟨y, hyI, hfy⟩ := h2 _ hI' image_insert_eq.symm + rw [hf.eq_iff (hI.subset_ground hyI) hxE] at hfy + rwa [← hfy] + rintro ⟨x, ⟨hxE, hxi⟩, rfl⟩ + refine ⟨⟨x, hxE, rfl⟩, fun J hJ hJI ↦ ⟨x, hxi ?_, rfl⟩⟩ + replace hJ := hJ.map f hf + have hrw := image_insert_eq ▸ hJI + rwa [← hrw, map_image_indep_iff (insert_subset hxE hI.subset_ground)] at hJ + +lemma restrict_spanning_iff (hSR : S ⊆ R) (hR : R ⊆ M.E := by aesop_mat) : + (M ↾ R).Spanning S ↔ R ⊆ M.closure S := by + rw [spanning_iff, restrict_ground_eq, and_iff_left hSR, restrict_closure_eq _ hSR, inter_eq_right] + +lemma restrict_spanning_iff' : (M ↾ R).Spanning S ↔ R ∩ M.E ⊆ M.closure S ∧ S ⊆ R := by + rw [spanning_iff, restrict_closure_eq', restrict_ground_eq, and_congr_left_iff, + diff_eq_compl_inter, ← union_inter_distrib_right, inter_eq_right, union_comm, + ← diff_subset_iff, diff_compl] + intro hSR + rw [inter_eq_self_of_subset_left hSR] + +end Constructions end Matroid diff --git a/Mathlib/Data/Matroid/Constructions.lean b/Mathlib/Data/Matroid/Constructions.lean index 2d46aa252305c..a6e15c479d63c 100644 --- a/Mathlib/Data/Matroid/Constructions.lean +++ b/Mathlib/Data/Matroid/Constructions.lean @@ -123,10 +123,13 @@ theorem empty_base_iff : M.Base ∅ ↔ M = loopyOn M.E := by exact ⟨fun h I _ ↦ ⟨@h _, fun hI ↦ by simp [hI]⟩, fun h I hI ↦ (h hI.subset_ground).1 hI⟩ theorem eq_loopyOn_or_rkPos (M : Matroid α) : M = loopyOn M.E ∨ RkPos M := by - rw [← empty_base_iff, rkPos_iff_empty_not_base]; apply em + rw [← empty_base_iff, rkPos_iff]; apply em theorem not_rkPos_iff : ¬RkPos M ↔ M = loopyOn M.E := by - rw [rkPos_iff_empty_not_base, not_iff_comm, empty_base_iff] + rw [rkPos_iff, not_iff_comm, empty_base_iff] + +instance loopyOn_finiteRk : FiniteRk (loopyOn E) := + ⟨∅, by simp⟩ end LoopyOn @@ -183,6 +186,13 @@ theorem restrict_eq_freeOn_iff : M ↾ I = freeOn I ↔ M.Indep I := by theorem Indep.restrict_eq_freeOn (hI : M.Indep I) : M ↾ I = freeOn I := by rwa [restrict_eq_freeOn_iff] +instance freeOn_finitary : Finitary (freeOn E) := by + simp only [finitary_iff, freeOn_indep_iff] + exact fun I h e heI ↦ by simpa using h {e} (by simpa) + +lemma freeOn_rkPos (hE : E.Nonempty) : RkPos (freeOn E) := by + simp [rkPos_iff, hE.ne_empty.symm] + end FreeOn section uniqueBaseOn @@ -243,6 +253,20 @@ theorem uniqueBaseOn_restrict (h : I ⊆ E) (R : Set α) : (uniqueBaseOn I E) ↾ R = uniqueBaseOn (I ∩ R) R := by rw [uniqueBaseOn_restrict', inter_right_comm, inter_eq_self_of_subset_left h] +lemma uniqueBaseOn_finiteRk (hI : I.Finite) : FiniteRk (uniqueBaseOn I E) := by + rw [← uniqueBaseOn_inter_ground_eq] + refine ⟨I ∩ E, ?_⟩ + rw [uniqueBaseOn_base_iff inter_subset_right, and_iff_right rfl] + exact hI.subset inter_subset_left + +instance uniqueBaseOn_finitary : Finitary (uniqueBaseOn I E) := by + refine ⟨fun K hK ↦ ?_⟩ + simp only [uniqueBaseOn_indep_iff'] at hK ⊢ + exact fun e heK ↦ singleton_subset_iff.1 <| hK _ (by simpa) (by simp) + +lemma uniqueBaseOn_rkPos (hIE : I ⊆ E) (hI : I.Nonempty) : RkPos (uniqueBaseOn I E) where + empty_not_base := by simpa [uniqueBaseOn_base_iff hIE] using Ne.symm <| hI.ne_empty + end uniqueBaseOn end Matroid diff --git a/Mathlib/Data/Matroid/Dual.lean b/Mathlib/Data/Matroid/Dual.lean index a6b523bb31bc1..4061f8c33178f 100644 --- a/Mathlib/Data/Matroid/Dual.lean +++ b/Mathlib/Data/Matroid/Dual.lean @@ -206,7 +206,7 @@ theorem base_iff_dual_base_compl (hB : B ⊆ M.E := by aesop_mat) : rw [dual_base_iff, diff_diff_cancel_left hB] theorem ground_not_base (M : Matroid α) [h : RkPos M✶] : ¬M.Base M.E := by - rwa [rkPos_iff_empty_not_base, dual_base_iff, diff_empty] at h + rwa [rkPos_iff, dual_base_iff, diff_empty] at h theorem Base.ssubset_ground [h : RkPos M✶] (hB : M.Base B) : B ⊂ M.E := hB.subset_ground.ssubset_of_ne (by rintro rfl; exact M.ground_not_base hB) diff --git a/Mathlib/Data/Matroid/IndepAxioms.lean b/Mathlib/Data/Matroid/IndepAxioms.lean index 562ae11604dbd..89808d55cb1a6 100644 --- a/Mathlib/Data/Matroid/IndepAxioms.lean +++ b/Mathlib/Data/Matroid/IndepAxioms.lean @@ -61,6 +61,9 @@ for the inverse of `e`). in the special case where independence of a set is determined only by that of its finite subsets. This construction uses Zorn's lemma. +* `IndepMatroid.ofFinitaryCardAugment` is a variant of `IndepMatroid.ofFinitary` where the + augmentation axiom resembles the finite augmentation axiom. + * `IndepMatroid.ofBdd` constructs an `IndepMatroid` in the case where there is some known absolute upper bound on the size of an independent set. This uses the infinite version of the augmentation axiom; the corresponding `Matroid` is `FiniteRk`. @@ -139,88 +142,23 @@ namespace IndepMatroid @[simp] theorem matroid_indep_iff {M : IndepMatroid α} {I : Set α} : M.matroid.Indep I ↔ M.Indep I := Iff.rfl -/-- An independence predicate satisfying the finite matroid axioms determines a matroid, - provided independence is determined by its behaviour on finite sets. - This fundamentally needs choice, since it can be used to prove that every vector space - has a basis. -/ +/-- If `Indep` has the 'compactness' property that each set `I` satisfies `Indep I` if and only if +`Indep J` for every finite subset `J` of `I`, +then an `IndepMatroid` can be constructed without proving the maximality axiom. +This needs choice, since it can be used to prove that every vector space has a basis. -/ @[simps E] protected def ofFinitary (E : Set α) (Indep : Set α → Prop) (indep_empty : Indep ∅) (indep_subset : ∀ ⦃I J⦄, Indep J → I ⊆ J → Indep I) - (indep_aug : ∀ ⦃I J⦄, Indep I → I.Finite → Indep J → J.Finite → I.ncard < J.ncard → - ∃ e ∈ J, e ∉ I ∧ Indep (insert e I)) + (indep_aug : ∀ ⦃I B⦄, Indep I → ¬ Maximal Indep I → Maximal Indep B → + ∃ x ∈ B \ I, Indep (insert x I)) (indep_compact : ∀ I, (∀ J, J ⊆ I → J.Finite → Indep J) → Indep I) - (subset_ground : ∀ I, Indep I → I ⊆ E) : IndepMatroid α := - have htofin : ∀ I e, Indep I → ¬ Indep (insert e I) → - ∃ I₀, I₀ ⊆ I ∧ I₀.Finite ∧ ¬ Indep (insert e I₀) := by - by_contra h; push_neg at h - obtain ⟨I, e, -, hIe, h⟩ := h - refine hIe <| indep_compact _ fun J hJss hJfin ↦ ?_ - exact indep_subset (h (J \ {e}) (by rwa [diff_subset_iff]) hJfin.diff) (by simp) - IndepMatroid.mk - (E := E) - (Indep := Indep) - (indep_empty := indep_empty) - (indep_subset := indep_subset) - (indep_aug := by - intro I B hI hImax hBmax - obtain ⟨e, heI, hins⟩ := exists_insert_of_not_maximal indep_subset hI hImax - by_cases heB : e ∈ B - · exact ⟨e, ⟨heB, heI⟩, hins⟩ - by_contra hcon; push_neg at hcon - - have heBdep := hBmax.not_prop_of_ssuperset (ssubset_insert heB) - - -- There is a finite subset `B₀` of `B` so that `B₀ + e` is dependent - obtain ⟨B₀, hB₀B, hB₀fin, hB₀e⟩ := htofin B e hBmax.1 heBdep - have hB₀ := indep_subset hBmax.1 hB₀B - - -- `I` has a finite subset `I₀` that doesn't extend into `B₀` - have hexI₀ : ∃ I₀, I₀ ⊆ I ∧ I₀.Finite ∧ ∀ x, x ∈ B₀ \ I₀ → ¬Indep (insert x I₀) := by - have hchoose : ∀ (b : ↑(B₀ \ I)), ∃ Ib, Ib ⊆ I ∧ Ib.Finite ∧ ¬Indep (insert (b : α) Ib) := by - rintro ⟨b, hb⟩; exact htofin I b hI (hcon b ⟨hB₀B hb.1, hb.2⟩) - choose! f hf using hchoose - have : Finite ↑(B₀ \ I) := hB₀fin.diff.to_subtype - refine ⟨iUnion f ∪ (B₀ ∩ I), - union_subset (iUnion_subset (fun i ↦ (hf i).1)) inter_subset_right, - (finite_iUnion fun i ↦ (hf i).2.1).union (hB₀fin.subset inter_subset_left), - fun x ⟨hxB₀, hxn⟩ hi ↦ ?_⟩ - have hxI : x ∉ I := fun hxI ↦ hxn <| Or.inr ⟨hxB₀, hxI⟩ - refine (hf ⟨x, ⟨hxB₀, hxI⟩⟩).2.2 (indep_subset hi <| insert_subset_insert ?_) - apply subset_union_of_subset_left - apply subset_iUnion - - obtain ⟨I₀, hI₀I, hI₀fin, hI₀⟩ := hexI₀ - - set E₀ := insert e (I₀ ∪ B₀) - have hE₀fin : E₀.Finite := (hI₀fin.union hB₀fin).insert e - - -- Extend `B₀` to a maximal independent subset of `I₀ ∪ B₀ + e` - obtain ⟨J, ⟨hB₀J, hJ, hJss⟩, hJmax⟩ := Finite.exists_maximal_wrt (f := id) - (s := {J | B₀ ⊆ J ∧ Indep J ∧ J ⊆ E₀}) - (hE₀fin.finite_subsets.subset (by simp)) - ⟨B₀, Subset.rfl, hB₀, subset_union_right.trans (subset_insert _ _)⟩ - - have heI₀ : e ∉ I₀ := not_mem_subset hI₀I heI - have heI₀i : Indep (insert e I₀) := indep_subset hins (insert_subset_insert hI₀I) - - have heJ : e ∉ J := fun heJ ↦ hB₀e (indep_subset hJ <| insert_subset heJ hB₀J) - - have hJfin := hE₀fin.subset hJss - - -- We have `|I₀ + e| ≤ |J|`, since otherwise we could extend the maximal set `J` - have hcard : (insert e I₀).ncard ≤ J.ncard := by - refine not_lt.1 fun hlt ↦ ?_ - obtain ⟨f, hfI, hfJ, hfi⟩ := indep_aug hJ hJfin heI₀i (hI₀fin.insert e) hlt - have hfE₀ : f ∈ E₀ := mem_of_mem_of_subset hfI (insert_subset_insert subset_union_left) - refine hfJ (insert_eq_self.1 <| Eq.symm (hJmax _ - ⟨hB₀J.trans <| subset_insert _ _,hfi,insert_subset hfE₀ hJss⟩ (subset_insert _ _))) - - -- But this means `|I₀| < |J|`, and extending `I₀` into `J` gives a contradiction - rw [ncard_insert_of_not_mem heI₀ hI₀fin, ← Nat.lt_iff_add_one_le] at hcard - - obtain ⟨f, hfJ, hfI₀, hfi⟩ := indep_aug (indep_subset hI hI₀I) hI₀fin hJ hJfin hcard - exact hI₀ f ⟨Or.elim (hJss hfJ) (fun hfe ↦ (heJ <| hfe ▸ hfJ).elim) (by aesop), hfI₀⟩ hfi) - (indep_maximal := by + (subset_ground : ∀ I, Indep I → I ⊆ E) : IndepMatroid α where + E := E + Indep := Indep + indep_empty := indep_empty + indep_subset := indep_subset + indep_aug := indep_aug + indep_maximal := by refine fun X _ I hI hIX ↦ zorn_subset_nonempty {Y | Indep Y ∧ Y ⊆ X} ?_ I ⟨hI, hIX⟩ refine fun Is hIs hchain _ ↦ ⟨⋃₀ Is, ⟨?_, sUnion_subset fun Y hY ↦ (hIs hY).2⟩, fun _ ↦ subset_sUnion_of_mem⟩ @@ -232,12 +170,12 @@ namespace IndepMatroid refine indep_subset (hIs (hf x hxJ).1).1 fun y hyJ ↦ ?_ obtain (hle | hle) := hchain.total (hf _ hxJ).1 (hf _ hyJ).1 · rw [hxmax _ hyJ hle]; exact (hf _ hyJ).2 - exact hle (hf _ hyJ).2) - - (subset_ground := subset_ground) + exact hle (hf _ hyJ).2 + subset_ground := subset_ground @[simp] theorem ofFinitary_indep (E : Set α) (Indep : Set α → Prop) - indep_empty indep_subset indep_aug indep_compact subset_ground : (IndepMatroid.ofFinitary + indep_empty indep_subset indep_aug indep_compact subset_ground : + (IndepMatroid.ofFinitary E Indep indep_empty indep_subset indep_aug indep_compact subset_ground).Indep = Indep := rfl instance ofFinitary_finitary (E : Set α) (Indep : Set α → Prop) @@ -246,6 +184,100 @@ instance ofFinitary_finitary (E : Set α) (Indep : Set α → Prop) E Indep indep_empty indep_subset indep_aug indep_compact subset_ground).matroid := ⟨by simpa⟩ +/-- An independence predicate satisfying the finite matroid axioms determines a matroid, +provided independence is determined by its behaviour on finite sets. -/ +@[simps! E] protected def ofFinitaryCardAugment (E : Set α) (Indep : Set α → Prop) + (indep_empty : Indep ∅) + (indep_subset : ∀ ⦃I J⦄, Indep J → I ⊆ J → Indep I) + (indep_aug : ∀ ⦃I J⦄, Indep I → I.Finite → Indep J → J.Finite → I.ncard < J.ncard → + ∃ e ∈ J, e ∉ I ∧ Indep (insert e I)) + (indep_compact : ∀ I, (∀ J, J ⊆ I → J.Finite → Indep J) → Indep I) + (subset_ground : ∀ I, Indep I → I ⊆ E) : IndepMatroid α := + IndepMatroid.ofFinitary + (E := E) + (Indep := Indep) + (indep_empty := indep_empty) + (indep_subset := indep_subset) + (indep_compact := indep_compact) + (indep_aug := by + have htofin : ∀ I e, Indep I → ¬ Indep (insert e I) → + ∃ I₀, I₀ ⊆ I ∧ I₀.Finite ∧ ¬ Indep (insert e I₀) := by + by_contra h; push_neg at h + obtain ⟨I, e, -, hIe, h⟩ := h + refine hIe <| indep_compact _ fun J hJss hJfin ↦ ?_ + exact indep_subset (h (J \ {e}) (by rwa [diff_subset_iff]) hJfin.diff) (by simp) + + intro I B hI hImax hBmax + obtain ⟨e, heI, hins⟩ := exists_insert_of_not_maximal indep_subset hI hImax + by_cases heB : e ∈ B + · exact ⟨e, ⟨heB, heI⟩, hins⟩ + by_contra hcon; push_neg at hcon + + have heBdep := hBmax.not_prop_of_ssuperset (ssubset_insert heB) + + -- There is a finite subset `B₀` of `B` so that `B₀ + e` is dependent + obtain ⟨B₀, hB₀B, hB₀fin, hB₀e⟩ := htofin B e hBmax.1 heBdep + have hB₀ := indep_subset hBmax.1 hB₀B + + -- `I` has a finite subset `I₀` that doesn't extend into `B₀` + have hexI₀ : ∃ I₀, I₀ ⊆ I ∧ I₀.Finite ∧ ∀ x, x ∈ B₀ \ I₀ → ¬Indep (insert x I₀) := by + have hch : ∀ (b : ↑(B₀ \ I)), ∃ Ib, Ib ⊆ I ∧ Ib.Finite ∧ ¬Indep (insert (b : α) Ib) := by + rintro ⟨b, hb⟩; exact htofin I b hI (hcon b ⟨hB₀B hb.1, hb.2⟩) + choose! f hf using hch + have : Finite ↑(B₀ \ I) := hB₀fin.diff.to_subtype + refine ⟨iUnion f ∪ (B₀ ∩ I), + union_subset (iUnion_subset (fun i ↦ (hf i).1)) inter_subset_right, + (finite_iUnion fun i ↦ (hf i).2.1).union (hB₀fin.subset inter_subset_left), + fun x ⟨hxB₀, hxn⟩ hi ↦ ?_⟩ + have hxI : x ∉ I := fun hxI ↦ hxn <| Or.inr ⟨hxB₀, hxI⟩ + refine (hf ⟨x, ⟨hxB₀, hxI⟩⟩).2.2 (indep_subset hi <| insert_subset_insert ?_) + apply subset_union_of_subset_left + apply subset_iUnion + + obtain ⟨I₀, hI₀I, hI₀fin, hI₀⟩ := hexI₀ + + set E₀ := insert e (I₀ ∪ B₀) + have hE₀fin : E₀.Finite := (hI₀fin.union hB₀fin).insert e + + -- Extend `B₀` to a maximal independent subset of `I₀ ∪ B₀ + e` + obtain ⟨J, ⟨hB₀J, hJ, hJss⟩, hJmax⟩ := Finite.exists_maximal_wrt (f := id) + (s := {J | B₀ ⊆ J ∧ Indep J ∧ J ⊆ E₀}) + (hE₀fin.finite_subsets.subset (by simp)) + ⟨B₀, Subset.rfl, hB₀, subset_union_right.trans (subset_insert _ _)⟩ + + have heI₀ : e ∉ I₀ := not_mem_subset hI₀I heI + have heI₀i : Indep (insert e I₀) := indep_subset hins (insert_subset_insert hI₀I) + + have heJ : e ∉ J := fun heJ ↦ hB₀e (indep_subset hJ <| insert_subset heJ hB₀J) + + have hJfin := hE₀fin.subset hJss + + -- We have `|I₀ + e| ≤ |J|`, since otherwise we could extend the maximal set `J` + have hcard : (insert e I₀).ncard ≤ J.ncard := by + refine not_lt.1 fun hlt ↦ ?_ + obtain ⟨f, hfI, hfJ, hfi⟩ := indep_aug hJ hJfin heI₀i (hI₀fin.insert e) hlt + have hfE₀ : f ∈ E₀ := mem_of_mem_of_subset hfI (insert_subset_insert subset_union_left) + refine hfJ (insert_eq_self.1 <| Eq.symm (hJmax _ + ⟨hB₀J.trans <| subset_insert _ _,hfi,insert_subset hfE₀ hJss⟩ (subset_insert _ _))) + + -- But this means `|I₀| < |J|`, and extending `I₀` into `J` gives a contradiction + rw [ncard_insert_of_not_mem heI₀ hI₀fin, ← Nat.lt_iff_add_one_le] at hcard + + obtain ⟨f, hfJ, hfI₀, hfi⟩ := indep_aug (indep_subset hI hI₀I) hI₀fin hJ hJfin hcard + exact hI₀ f ⟨Or.elim (hJss hfJ) (fun hfe ↦ (heJ <| hfe ▸ hfJ).elim) (by aesop), hfI₀⟩ hfi ) + (subset_ground := subset_ground) + +@[simp] theorem ofFinitaryCardAugment_indep (E : Set α) (Indep : Set α → Prop) + indep_empty indep_subset indep_aug indep_compact subset_ground : + (IndepMatroid.ofFinitaryCardAugment + E Indep indep_empty indep_subset indep_aug indep_compact subset_ground).Indep = Indep := rfl + +instance ofFinitaryCardAugment_finitary (E : Set α) (Indep : Set α → Prop) + indep_empty indep_subset indep_aug indep_compact subset_ground : Finitary + (IndepMatroid.ofFinitaryCardAugment + E Indep indep_empty indep_subset indep_aug indep_compact subset_ground).matroid := + ⟨by simpa⟩ + /-- If there is an absolute upper bound on the size of a set satisfying `P`, then the maximal subset property always holds. -/ theorem _root_.Matroid.existsMaximalSubsetProperty_of_bdd {P : Set α → Prop} @@ -380,7 +412,7 @@ protected def ofFinset [DecidableEq α] (E : Set α) (Indep : Finset α → Prop (indep_subset : ∀ ⦃I J⦄, Indep J → I ⊆ J → Indep I) (indep_aug : ∀ ⦃I J⦄, Indep I → Indep J → I.card < J.card → ∃ e ∈ J, e ∉ I ∧ Indep (insert e I)) (subset_ground : ∀ ⦃I⦄, Indep I → (I : Set α) ⊆ E) : IndepMatroid α := - IndepMatroid.ofFinitary + IndepMatroid.ofFinitaryCardAugment (E := E) (Indep := (fun I ↦ (∀ (J : Finset α), (J : Set α) ⊆ I → Indep J))) (indep_empty := by simpa [subset_empty_iff]) @@ -402,7 +434,7 @@ protected def ofFinset [DecidableEq α] (E : Set α) (Indep : Finset α → Prop @[simp] theorem ofFinset_indep [DecidableEq α] (E : Set α) Indep indep_empty indep_subset indep_aug subset_ground {I : Finset α} : (IndepMatroid.ofFinset E Indep indep_empty indep_subset indep_aug subset_ground).Indep I ↔ Indep I := by - simp only [IndepMatroid.ofFinset, ofFinitary_indep, Finset.coe_subset] + simp only [IndepMatroid.ofFinset, ofFinitaryCardAugment_indep, Finset.coe_subset] exact ⟨fun h ↦ h _ Subset.rfl, fun h J hJI ↦ indep_subset h hJI⟩ /-- This can't be `@[simp]`, because it would cause the more useful @@ -411,7 +443,7 @@ theorem ofFinset_indep' [DecidableEq α] (E : Set α) Indep indep_empty indep_su subset_ground {I : Set α} : (IndepMatroid.ofFinset E Indep indep_empty indep_subset indep_aug subset_ground).Indep I ↔ ∀ (J : Finset α), (J : Set α) ⊆ I → Indep J := by - simp only [IndepMatroid.ofFinset, ofFinitary_indep] + simp only [IndepMatroid.ofFinset, ofFinitaryCardAugment_indep] end IndepMatroid diff --git a/Mathlib/Data/Matroid/Rank/Cardinal.lean b/Mathlib/Data/Matroid/Rank/Cardinal.lean new file mode 100644 index 0000000000000..0d10845213b06 --- /dev/null +++ b/Mathlib/Data/Matroid/Rank/Cardinal.lean @@ -0,0 +1,345 @@ +/- +Copyright (c) 2025 Peter Nelson and Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Peter Nelson, Junyan Xu +-/ +import Mathlib.Data.Matroid.Closure +import Mathlib.Data.Matroid.Map +import Mathlib.SetTheory.Cardinal.Arithmetic + +/-! +# Cardinal-valued rank + +In a finitary matroid, all bases have the same cardinality. +In fact, something stronger holds: if `I` and `J` are both bases for a set `X`, +then `#(I \ J) = #(J \ I)` and (consequently) `#I = #J`. +This file introduces a typeclass `InvariantCardinalRank` that applies to any matroid +such that this property holds for all `I`, `J` and `X`. + +A matroid satisfying this condition has a well-defined cardinality-valued rank function, +both for itself and all its minors. + +# Main Declarations + +* `Matroid.InvariantCardinalRank` : a typeclass capturing the idea that a matroid and all its minors + have a well-behaved cardinal-valued rank function. +* `Matroid.cRank M` is the supremum of the cardinalities of the bases of matroid `M`. +* `Matroid.cRk M X` is the supremum of the cardinalities of the bases of a set `X` in a matroid `M`. +* `invariantCardinalRank_of_finitary` is the instance + showing that `Finitary` matroids are `InvariantCardinalRank`. +* `cRk_inter_add_cRk_union_le` states that cardinal rank is submodular. + +# Notes + +It is not the case that all matroids are `InvariantCardinalRank`, +since the equicardinality of bases in general matroids is independent of ZFC +(see the module docstring of `Mathlib.Data.Matroid.Basic`). +Lemmas like `Matroid.Base.cardinalMk_diff_comm` become true for all matroids +only if they are weakened by replacing `Cardinal.mk` +with the cruder `ℕ∞`-valued `Set.encard`; see, for example, `Matroid.Base.encard_diff_comm`. + +# Implementation Details + +Since the functions `cRank` and `cRk` are defined as suprema, +independently of the `Matroid.InvariantCardinalRank` typeclass, +they are well-defined for all matroids. +However, for matroids that do not satisfy `InvariantCardinalRank`, they are badly behaved. +For instance, in general `cRk` is not submodular, +and its value may differ on a set `X` and the closure of `X`. +We state and prove theorems without `InvariantCardinalRank` whenever possible, +which sometime makes their proofs longer than they would be with the instance. + +# TODO + +* Higgs' theorem : if the generalized continuum hypothesis holds, + then every matroid is `InvariantCardinalRank`. + +-/ + +universe u v + +variable {α : Type u} {β : Type v} {f : α → β} {M : Matroid α} {I J B B' X Y : Set α} + +open Cardinal Set + +namespace Matroid + +section Rank + +variable {κ : Cardinal} + +/-- The rank (supremum of the cardinalities of bases) of a matroid `M` as a `Cardinal`. -/ +noncomputable def cRank (M : Matroid α) := ⨆ (B : {B // M.Base B}), #B + +/-- The rank (supremum of the cardinalities of bases) of a set `X` in a matroid `M`, +as a `Cardinal`. -/ +noncomputable def cRk (M : Matroid α) (X : Set α) := (M ↾ X).cRank + +theorem Base.cardinalMk_le_cRank (hB : M.Base B) : #B ≤ M.cRank := + le_ciSup (f := fun (B : {B // M.Base B}) ↦ #(B.1)) (bddAbove_range _) ⟨B, hB⟩ + +theorem Basis'.cardinalMk_le_cRk (hIX : M.Basis' I X) : #I ≤ M.cRk X := + (base_restrict_iff'.2 hIX).cardinalMk_le_cRank + +theorem Basis.cardinalMk_le_cRk (hIX : M.Basis I X) : #I ≤ M.cRk X := + hIX.basis'.cardinalMk_le_cRk + +theorem cRank_le_iff : M.cRank ≤ κ ↔ ∀ ⦃B⦄, M.Base B → #B ≤ κ := + ⟨fun h _ hB ↦ (hB.cardinalMk_le_cRank.trans h), fun h ↦ ciSup_le fun ⟨_, hB⟩ ↦ h hB⟩ + +theorem cRk_le_iff : M.cRk X ≤ κ ↔ ∀ ⦃I⦄, M.Basis' I X → #I ≤ κ := by + simp_rw [cRk, cRank_le_iff, base_restrict_iff'] + +theorem Indep.cardinalMk_le_cRk_of_subset (hI : M.Indep I) (hIX : I ⊆ X) : #I ≤ M.cRk X := + let ⟨_, hJ, hIJ⟩ := hI.subset_basis'_of_subset hIX + (mk_le_mk_of_subset hIJ).trans hJ.cardinalMk_le_cRk + +theorem cRk_le_cardinalMk (M : Matroid α) (X : Set α) : M.cRk X ≤ #X := + ciSup_le fun ⟨_, hI⟩ ↦ mk_le_mk_of_subset hI.subset_ground + +@[simp] theorem cRk_ground (M : Matroid α) : M.cRk M.E = M.cRank := by + rw [cRk, restrict_ground_eq_self] + +@[simp] theorem cRank_restrict (M : Matroid α) (X : Set α) : (M ↾ X).cRank = M.cRk X := rfl + +theorem cRk_mono (M : Matroid α) : Monotone M.cRk := by + simp only [Monotone, le_eq_subset, cRk_le_iff] + intro X Y hXY I hIX + obtain ⟨J, hJ, hIJ⟩ := hIX.indep.subset_basis'_of_subset (hIX.subset.trans hXY) + exact (mk_le_mk_of_subset hIJ).trans hJ.cardinalMk_le_cRk + +theorem cRk_le_of_subset (M : Matroid α) (hXY : X ⊆ Y) : M.cRk X ≤ M.cRk Y := + M.cRk_mono hXY + +@[simp] theorem cRk_inter_ground (M : Matroid α) (X : Set α) : M.cRk (X ∩ M.E) = M.cRk X := + (M.cRk_le_of_subset inter_subset_left).antisymm <| cRk_le_iff.2 + fun _ h ↦ h.basis_inter_ground.cardinalMk_le_cRk + +theorem cRk_restrict_subset (M : Matroid α) (hYX : Y ⊆ X) : (M ↾ X).cRk Y = M.cRk Y := by + have aux : ∀ ⦃I⦄, M.Basis' I Y ↔ (M ↾ X).Basis' I Y := by + simp_rw [basis'_restrict_iff, inter_eq_self_of_subset_left hYX, iff_self_and] + exact fun I h ↦ h.subset.trans hYX + simp_rw [le_antisymm_iff, cRk_le_iff] + exact ⟨fun I hI ↦ (aux.2 hI).cardinalMk_le_cRk, fun I hI ↦ (aux.1 hI).cardinalMk_le_cRk⟩ + +theorem cRk_restrict (M : Matroid α) (X Y : Set α) : (M ↾ X).cRk Y = M.cRk (X ∩ Y) := by + rw [← cRk_inter_ground, restrict_ground_eq, cRk_restrict_subset _ inter_subset_right, + inter_comm] + +theorem Indep.cRk_eq_cardinalMk (hI : M.Indep I) : #I = M.cRk I := + (M.cRk_le_cardinalMk I).antisymm' (hI.basis_self.cardinalMk_le_cRk) + +@[simp] theorem cRk_map_image_lift (M : Matroid α) (hf : InjOn f M.E) (X : Set α) + (hX : X ⊆ M.E := by aesop_mat) : lift.{u,v} ((M.map f hf).cRk (f '' X)) = lift (M.cRk X) := by + nth_rw 1 [cRk, cRank, le_antisymm_iff, lift_iSup (bddAbove_range _), cRk, cRank, cRk, cRank] + nth_rw 2 [lift_iSup (bddAbove_range _)] + simp only [ciSup_le_iff (bddAbove_range _), ge_iff_le, Subtype.forall, base_restrict_iff', + basis'_iff_basis hX, basis'_iff_basis (show f '' X ⊆ (M.map f hf).E from image_mono hX)] + refine ⟨fun I hI ↦ ?_, fun I hI ↦ ?_⟩ + · obtain ⟨I, X', hIX, rfl, hXX'⟩ := map_basis_iff'.1 hI + rw [mk_image_eq_of_injOn_lift _ _ (hf.mono hIX.indep.subset_ground), lift_le] + obtain rfl : X = X' := by rwa [hf.image_eq_image_iff hX hIX.subset_ground] at hXX' + exact hIX.cardinalMk_le_cRk + rw [← mk_image_eq_of_injOn_lift _ _ (hf.mono hI.indep.subset_ground), lift_le] + exact (hI.map hf).cardinalMk_le_cRk + +@[simp] theorem cRk_map_image {β : Type u} {f : α → β} (M : Matroid α) (hf : InjOn f M.E) + (X : Set α) (hX : X ⊆ M.E := by aesop_mat) : (M.map f hf).cRk (f '' X) = M.cRk X := + lift_inj.1 <| M.cRk_map_image_lift .. + +theorem cRk_map_eq {β : Type u} {f : α → β} {X : Set β} (M : Matroid α) (hf : InjOn f M.E) : + (M.map f hf).cRk X = M.cRk (f ⁻¹' X) := by + rw [← M.cRk_inter_ground, ← M.cRk_map_image hf _, image_preimage_inter, ← map_ground _ _ hf, + cRk_inter_ground] + +@[simp] theorem cRk_comap_lift (M : Matroid β) (f : α → β) (X : Set α) : + lift.{v,u} ((M.comap f).cRk X) = lift (M.cRk (f '' X)) := by + nth_rw 1 [cRk, cRank, le_antisymm_iff, lift_iSup (bddAbove_range _), cRk, cRank, cRk, cRank] + nth_rw 2 [lift_iSup (bddAbove_range _)] + simp only [ciSup_le_iff (bddAbove_range _), ge_iff_le, Subtype.forall, base_restrict_iff', + comap_basis'_iff, and_imp] + refine ⟨fun I hI hfI hIX ↦ ?_, fun I hIX ↦ ?_⟩ + · rw [← mk_image_eq_of_injOn_lift _ _ hfI, lift_le] + exact hI.cardinalMk_le_cRk + obtain ⟨I₀, hI₀X, rfl, hfI₀⟩ := show ∃ I₀ ⊆ X, f '' I₀ = I ∧ InjOn f I₀ by + obtain ⟨I₀, hI₀ss, hbij⟩ := exists_subset_bijOn (f ⁻¹' I ∩ X) f + refine ⟨I₀, hI₀ss.trans inter_subset_right, ?_, hbij.injOn⟩ + rw [hbij.image_eq, image_preimage_inter, inter_eq_self_of_subset_left hIX.subset] + rw [mk_image_eq_of_injOn_lift _ _ hfI₀, lift_le] + exact Basis'.cardinalMk_le_cRk <| comap_basis'_iff.2 ⟨hIX, hfI₀, hI₀X⟩ + +@[simp] theorem cRk_comap {β : Type u} (M : Matroid β) (f : α → β) (X : Set α) : + (M.comap f).cRk X = M.cRk (f '' X) := + lift_inj.1 <| M.cRk_comap_lift .. + +end Rank + +section Invariant + +/-- A class stating that cardinality-valued rank is well-defined +(i.e. all bases are equicardinal) for a matroid `M` and its minors. +Notably, this holds for `Finitary` matroids; see `Matroid.invarCardinalRank_of_finitary`. -/ +@[mk_iff] +class InvariantCardinalRank (M : Matroid α) : Prop where + forall_card_basis_diff : + ∀ ⦃I J X⦄, M.Basis I X → M.Basis J X → #(I \ J : Set α) = #(J \ I : Set α) + +variable [InvariantCardinalRank M] + +theorem Basis.cardinalMk_diff_comm (hIX : M.Basis I X) (hJX : M.Basis J X) : + #(I \ J : Set α) = #(J \ I : Set α) := + InvariantCardinalRank.forall_card_basis_diff hIX hJX + +theorem Basis'.cardinalMk_diff_comm (hIX : M.Basis' I X) (hJX : M.Basis' J X) : + #(I \ J : Set α) = #(J \ I : Set α) := + hIX.basis_inter_ground.cardinalMk_diff_comm hJX.basis_inter_ground + +theorem Base.cardinalMk_diff_comm (hB : M.Base B) (hB' : M.Base B') : + #(B \ B' : Set α) = #(B' \ B : Set α) := + hB.basis_ground.cardinalMk_diff_comm hB'.basis_ground + +theorem Basis.cardinalMk_eq (hIX : M.Basis I X) (hJX : M.Basis J X) : #I = #J := by + rw [← diff_union_inter I J, + mk_union_of_disjoint (disjoint_sdiff_left.mono_right inter_subset_right), + hIX.cardinalMk_diff_comm hJX, + ← mk_union_of_disjoint (disjoint_sdiff_left.mono_right inter_subset_left), + inter_comm, diff_union_inter] + +theorem Basis'.cardinalMk_eq (hIX : M.Basis' I X) (hJX : M.Basis' J X) : #I = #J := + hIX.basis_inter_ground.cardinalMk_eq hJX.basis_inter_ground + +theorem Base.cardinalMk_eq (hB : M.Base B) (hB' : M.Base B') : #B = #B' := + hB.basis_ground.cardinalMk_eq hB'.basis_ground + +theorem Indep.cardinalMk_le_base (hI : M.Indep I) (hB : M.Base B) : #I ≤ #B := + have ⟨_B', hB', hIB'⟩ := hI.exists_base_superset + hB'.cardinalMk_eq hB ▸ mk_le_mk_of_subset hIB' + +theorem Indep.cardinalMk_le_basis' (hI : M.Indep I) (hJ : M.Basis' J X) (hIX : I ⊆ X) : + #I ≤ #J := + have ⟨_J', hJ', hIJ'⟩ := hI.subset_basis'_of_subset hIX + hJ'.cardinalMk_eq hJ ▸ mk_le_mk_of_subset hIJ' + +theorem Indep.cardinalMk_le_basis (hI : M.Indep I) (hJ : M.Basis J X) (hIX : I ⊆ X) : + #I ≤ #J := + hI.cardinalMk_le_basis' hJ.basis' hIX + +theorem Base.cardinalMk_eq_cRank (hB : M.Base B) : #B = M.cRank := by + have hrw : ∀ B' : {B : Set α // M.Base B}, #B' = #B := fun B' ↦ B'.2.cardinalMk_eq hB + simp [cRank, hrw] + +/-- Restrictions of matroids with cardinal rank functions have cardinal rank functions- -/ +instance invariantCardinalRank_restrict [InvariantCardinalRank M] : + InvariantCardinalRank (M ↾ X) := by + refine ⟨fun I J Y hI hJ ↦ ?_⟩ + rw [basis_restrict_iff'] at hI hJ + exact hI.1.cardinalMk_diff_comm hJ.1 + +theorem Basis'.cardinalMk_eq_cRk (hIX : M.Basis' I X) : #I = M.cRk X := by + rw [cRk, (base_restrict_iff'.2 hIX).cardinalMk_eq_cRank] + +theorem Basis.cardinalMk_eq_cRk (hIX : M.Basis I X) : #I = M.cRk X := + hIX.basis'.cardinalMk_eq_cRk + +@[simp] theorem cRk_closure (M : Matroid α) [InvariantCardinalRank M] (X : Set α) : + M.cRk (M.closure X) = M.cRk X := by + obtain ⟨I, hI⟩ := M.exists_basis' X + rw [← hI.basis_closure_right.cardinalMk_eq_cRk, ← hI.cardinalMk_eq_cRk] + +theorem cRk_closure_congr (hXY : M.closure X = M.closure Y) : M.cRk X = M.cRk Y := by + rw [← cRk_closure, hXY, cRk_closure] + +variable (M : Matroid α) [InvariantCardinalRank M] (e : α) (X Y : Set α) + +@[simp] theorem cRk_union_closure_right_eq : M.cRk (X ∪ M.closure Y) = M.cRk (X ∪ Y) := + M.cRk_closure_congr (M.closure_union_closure_right_eq _ _) + +@[simp] theorem cRk_union_closure_left_eq : M.cRk (M.closure X ∪ Y) = M.cRk (X ∪ Y) := + M.cRk_closure_congr (M.closure_union_closure_left_eq _ _) + +@[simp] theorem cRk_insert_closure_eq : M.cRk (insert e (M.closure X)) = M.cRk (insert e X) := by + rw [← union_singleton, cRk_union_closure_left_eq, union_singleton] + +theorem cRk_union_closure_eq : M.cRk (M.closure X ∪ M.closure Y) = M.cRk (X ∪ Y) := by + simp + +/-- The `Cardinal` rank function is submodular. -/ +theorem cRk_inter_add_cRk_union_le : M.cRk (X ∩ Y) + M.cRk (X ∪ Y) ≤ M.cRk X + M.cRk Y := by + obtain ⟨Ii, hIi⟩ := M.exists_basis' (X ∩ Y) + obtain ⟨IX, hIX, hIX'⟩ := + hIi.indep.subset_basis'_of_subset (hIi.subset.trans inter_subset_left) + obtain ⟨IY, hIY, hIY'⟩ := + hIi.indep.subset_basis'_of_subset (hIi.subset.trans inter_subset_right) + rw [← cRk_union_closure_eq, ← hIX.closure_eq_closure, ← hIY.closure_eq_closure, + cRk_union_closure_eq, ← hIi.cardinalMk_eq_cRk, ← hIX.cardinalMk_eq_cRk, + ← hIY.cardinalMk_eq_cRk, ← mk_union_add_mk_inter, add_comm] + exact add_le_add (M.cRk_le_cardinalMk _) (mk_le_mk_of_subset (subset_inter hIX' hIY')) + +end Invariant + +section Instances + +/-- `Finitary` matroids have a cardinality-valued rank function. -/ +instance invariantCardinalRank_of_finitary [Finitary M] : InvariantCardinalRank M := by + suffices aux : ∀ ⦃B B'⦄ ⦃N : Matroid α⦄, Finitary N → N.Base B → N.Base B' → + #(B \ B' : Set α) ≤ #(B' \ B : Set α) from + ⟨fun I J X hI hJ ↦ (aux (restrict_finitary X) hI.base_restrict hJ.base_restrict).antisymm + (aux (restrict_finitary X) hJ.base_restrict hI.base_restrict)⟩ + intro B B' N hfin hB hB' + by_cases h : (B' \ B).Finite + · rw [← cast_ncard h, ← cast_ncard, hB.ncard_diff_comm hB'] + exact (hB'.diff_finite_comm hB).mp h + rw [← Set.Infinite, ← infinite_coe_iff] at h + have (a : α) (ha : a ∈ B' \ B) : ∃ S : Set α, Finite S ∧ S ⊆ B ∧ ¬ N.Indep (insert a S) := by + have := (hB.insert_dep ⟨hB'.subset_ground ha.1, ha.2⟩).1 + contrapose! this + exact Finitary.indep_of_forall_finite _ fun J hJ fin ↦ (this (J \ {a}) fin.diff.to_subtype <| + diff_singleton_subset_iff.mpr hJ).subset (subset_insert_diff_singleton ..) + choose S S_fin hSB dep using this + let U := ⋃ a : ↥(B' \ B), S a a.2 + suffices B \ B' ⊆ U by + refine (mk_le_mk_of_subset this).trans <| (mk_iUnion_le ..).trans + <| (mul_le_max_of_aleph0_le_left (by simp)).trans ?_ + simp only [sup_le_iff, le_refl, true_and] + exact ciSup_le' fun e ↦ (lt_aleph0_of_finite _).le.trans <| by simp + rw [← diff_inter_self_eq_diff, diff_subset_iff, inter_comm] + have hUB : (B ∩ B') ∪ U ⊆ B := + union_subset inter_subset_left (iUnion_subset fun e ↦ (hSB e.1 e.2)) + by_contra hBU + have ⟨a, ha, ind⟩ := hB.exists_insert_of_ssubset ⟨hUB, hBU⟩ hB' + have : a ∈ B' \ B := ⟨ha.1, fun haB ↦ ha.2 (.inl ⟨haB, ha.1⟩)⟩ + refine dep a this (ind.subset <| insert_subset_insert <| .trans ?_ subset_union_right) + exact subset_iUnion_of_subset ⟨a, this⟩ subset_rfl + +instance invariantCardinalRank_map (M : Matroid α) [InvariantCardinalRank M] (hf : InjOn f M.E) : + InvariantCardinalRank (M.map f hf) := by + refine ⟨fun I J X hI hJ ↦ ?_⟩ + obtain ⟨I, X, hIX, rfl, rfl⟩ := map_basis_iff'.1 hI + obtain ⟨J, X', hJX, rfl, h'⟩ := map_basis_iff'.1 hJ + obtain rfl : X = X' := by + rwa [InjOn.image_eq_image_iff hf hIX.subset_ground hJX.subset_ground] at h' + have hcard := hIX.cardinalMk_diff_comm hJX + rwa [← lift_inj.{u,v}, + ← mk_image_eq_of_injOn_lift _ _ (hf.mono ((hIX.indep.diff _).subset_ground)), + ← mk_image_eq_of_injOn_lift _ _ (hf.mono ((hJX.indep.diff _).subset_ground)), + lift_inj, (hf.mono hIX.indep.subset_ground).image_diff, + (hf.mono hJX.indep.subset_ground).image_diff, inter_comm, + hf.image_inter hJX.indep.subset_ground hIX.indep.subset_ground, + diff_inter_self_eq_diff, diff_self_inter] at hcard + +instance invariantCardinalRank_comap (M : Matroid β) [InvariantCardinalRank M] (f : α → β) : + InvariantCardinalRank (M.comap f) := by + refine ⟨fun I J X hI hJ ↦ ?_⟩ + obtain ⟨hI, hfI, hIX⟩ := comap_basis_iff.1 hI + obtain ⟨hJ, hfJ, hJX⟩ := comap_basis_iff.1 hJ + rw [← lift_inj.{u,v}, ← mk_image_eq_of_injOn_lift _ _ (hfI.mono diff_subset), + ← mk_image_eq_of_injOn_lift _ _ (hfJ.mono diff_subset), lift_inj, hfI.image_diff, + hfJ.image_diff, ← diff_union_diff_cancel inter_subset_left (image_inter_subset f I J), + inter_comm, diff_inter_self_eq_diff, mk_union_of_disjoint, hI.cardinalMk_diff_comm hJ, + ← diff_union_diff_cancel inter_subset_left (image_inter_subset f J I), inter_comm, + diff_inter_self_eq_diff, mk_union_of_disjoint, inter_comm J I] <;> + exact disjoint_sdiff_left.mono_right (diff_subset.trans inter_subset_left) + +end Instances + +end Matroid diff --git a/Mathlib/Data/Matroid/Restrict.lean b/Mathlib/Data/Matroid/Restrict.lean index 1eed5c8f2b783..fbb0716dcb9e3 100644 --- a/Mathlib/Data/Matroid/Restrict.lean +++ b/Mathlib/Data/Matroid/Restrict.lean @@ -195,6 +195,9 @@ theorem basis_restrict_iff (hR : R ⊆ M.E := by aesop_mat) : intro hXR rw [← basis'_iff_basis_inter_ground, basis'_iff_basis] +lemma basis'_iff_basis_restrict_univ : M.Basis' I X ↔ (M ↾ univ).Basis I X := by + rw [basis_restrict_iff', basis'_iff_basis_inter_ground, and_iff_left (subset_univ _)] + theorem restrict_eq_restrict_iff (M M' : Matroid α) (X : Set α) : M ↾ X = M' ↾ X ↔ ∀ I, I ⊆ X → (M.Indep I ↔ M'.Indep I) := by refine ⟨fun h I hIX ↦ ?_, fun h ↦ ext_indep rfl fun I (hI : I ⊆ X) ↦ ?_⟩ @@ -354,6 +357,9 @@ theorem Indep.of_restriction (hI : N.Indep I) (hNM : N ≤r M) : M.Indep I := by theorem Indep.indep_restriction (hI : M.Indep I) (hNM : N ≤r M) (hIN : I ⊆ N.E) : N.Indep I := by obtain ⟨R, -, rfl⟩ := hNM; simpa [hI] +theorem Restriction.indep_iff (hMN : N ≤r M) : N.Indep I ↔ M.Indep I ∧ I ⊆ N.E := + ⟨fun h ↦ ⟨h.of_restriction hMN, h.subset_ground⟩, fun h ↦ h.1.indep_restriction hMN h.2⟩ + theorem Basis.basis_restriction (hI : M.Basis I X) (hNM : N ≤r M) (hX : X ⊆ N.E) : N.Basis I X := by obtain ⟨R, hR, rfl⟩ := hNM; rwa [basis_restrict_iff, and_iff_left (show X ⊆ R from hX)] @@ -363,6 +369,13 @@ theorem Basis.of_restriction (hI : N.Basis I X) (hNM : N ≤r M) : M.Basis I X : theorem Base.basis_of_restriction (hI : N.Base I) (hNM : N ≤r M) : M.Basis I N.E := by obtain ⟨R, hR, rfl⟩ := hNM; rwa [base_restrict_iff] at hI +theorem Restriction.base_iff (hMN : N ≤r M) {B : Set α} : N.Base B ↔ M.Basis B N.E := + ⟨fun h ↦ Base.basis_of_restriction h hMN, + fun h ↦ by simpa [hMN.eq_restrict] using h.restrict_base⟩ + +theorem Restriction.basis_iff (hMN : N ≤r M) : N.Basis I X ↔ M.Basis I X ∧ X ⊆ N.E := + ⟨fun h ↦ ⟨h.of_restriction hMN, h.subset_ground⟩, fun h ↦ h.1.basis_restriction hMN h.2⟩ + theorem Dep.of_restriction (hX : N.Dep X) (hNM : N ≤r M) : M.Dep X := by obtain ⟨R, hR, rfl⟩ := hNM rw [restrict_dep_iff] at hX @@ -372,6 +385,9 @@ theorem Dep.dep_restriction (hX : M.Dep X) (hNM : N ≤r M) (hXE : X ⊆ N.E := N.Dep X := by obtain ⟨R, -, rfl⟩ := hNM; simpa [hX.not_indep] +theorem Restriction.dep_iff (hMN : N ≤r M) : N.Dep X ↔ M.Dep X ∧ X ⊆ N.E := + ⟨fun h ↦ ⟨h.of_restriction hMN, h.subset_ground⟩, fun h ↦ h.1.dep_restriction hMN h.2⟩ + end Restriction /-! @@ -440,6 +456,12 @@ theorem Indep.augment (hI : M.Indep I) (hJ : M.Indep J) (hIJ : I.encard < J.enca rw [← hJ'.encard_eq_encard hb] at hIJ exact hIJ.not_le (encard_mono hJJ') +lemma Indep.augment_finset {I J : Finset α} (hI : M.Indep I) (hJ : M.Indep J) + (hIJ : I.card < J.card) : ∃ e ∈ J, e ∉ I ∧ M.Indep (insert e I) := by + obtain ⟨x, hx, hxI⟩ := hI.augment hJ (by simpa [encard_eq_coe_toFinset_card] ) + simp only [mem_diff, Finset.mem_coe] at hx + exact ⟨x, hx.1, hx.2, hxI⟩ + end Basis end Matroid diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index ed70e899db4db..794a59feb7c3f 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -33,11 +33,13 @@ multiset enumeration -/ -variable {α : Type*} [DecidableEq α] {m : Multiset α} +variable {α β : Type*} [DecidableEq α] [DecidableEq β] {m : Multiset α} + +namespace Multiset /-- Auxiliary definition for the `CoeSort` instance. This prevents the `CoeOut m α` instance from inadvertently applying to other sigma types. -/ -def Multiset.ToType (m : Multiset α) : Type _ := (x : α) × Fin (m.count x) +def ToType (m : Multiset α) : Type _ := (x : α) × Fin (m.count x) /-- Create a type that has the same number of elements as the multiset. Terms of this type are triples `⟨x, ⟨i, h⟩⟩` where `x : α`, `i : ℕ`, and `h : i < m.count x`. @@ -51,7 +53,7 @@ example : DecidableEq m := inferInstanceAs <| DecidableEq ((x : α) × Fin (m.co /-- Constructor for terms of the coercion of `m` to a type. This helps Lean pick up the correct instances. -/ @[reducible, match_pattern] -def Multiset.mkToType (m : Multiset α) (x : α) (i : Fin (m.count x)) : m := +def mkToType (m : Multiset α) (x : α) (i : Fin (m.count x)) : m := ⟨x, i⟩ /-- As a convenience, there is a coercion from `m : Type*` to `α` by projecting onto the first @@ -64,18 +66,18 @@ instance instCoeSortMultisetType.instCoeOutToType : CoeOut m α := -- Syntactic equality -- @[simp] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10685): dsimp can prove this -theorem Multiset.coe_mk {x : α} {i : Fin (m.count x)} : ↑(m.mkToType x i) = x := +theorem coe_mk {x : α} {i : Fin (m.count x)} : ↑(m.mkToType x i) = x := rfl -@[simp] lemma Multiset.coe_mem {x : m} : ↑x ∈ m := Multiset.count_pos.mp (by have := x.2.2; omega) +@[simp] lemma coe_mem {x : m} : ↑x ∈ m := Multiset.count_pos.mp (by have := x.2.2; omega) @[simp] -protected theorem Multiset.forall_coe (p : m → Prop) : +protected theorem forall_coe (p : m → Prop) : (∀ x : m, p x) ↔ ∀ (x : α) (i : Fin (m.count x)), p ⟨x, i⟩ := Sigma.forall @[simp] -protected theorem Multiset.exists_coe (p : m → Prop) : +protected theorem exists_coe (p : m → Prop) : (∃ x : m, p x) ↔ ∃ (x : α) (i : Fin (m.count x)), p ⟨x, i⟩ := Sigma.exists @@ -92,19 +94,17 @@ instance : Fintype { p : α × ℕ | p.2 < m.count p.1 } := /-- Construct a finset whose elements enumerate the elements of the multiset `m`. The `ℕ` component is used to differentiate between equal elements: if `x` appears `n` times then `(x, 0)`, ..., and `(x, n-1)` appear in the `Finset`. -/ -def Multiset.toEnumFinset (m : Multiset α) : Finset (α × ℕ) := +def toEnumFinset (m : Multiset α) : Finset (α × ℕ) := { p : α × ℕ | p.2 < m.count p.1 }.toFinset @[simp] -theorem Multiset.mem_toEnumFinset (m : Multiset α) (p : α × ℕ) : +theorem mem_toEnumFinset (m : Multiset α) (p : α × ℕ) : p ∈ m.toEnumFinset ↔ p.2 < m.count p.1 := Set.mem_toFinset -theorem Multiset.mem_of_mem_toEnumFinset {p : α × ℕ} (h : p ∈ m.toEnumFinset) : p.1 ∈ m := +theorem mem_of_mem_toEnumFinset {p : α × ℕ} (h : p ∈ m.toEnumFinset) : p.1 ∈ m := have := (m.mem_toEnumFinset p).mp h; Multiset.count_pos.mp (by omega) -namespace Multiset - @[simp] lemma toEnumFinset_filter_eq (m : Multiset α) (a : α) : m.toEnumFinset.filter (·.1 = a) = {a} ×ˢ Finset.range (m.count a) := by aesop @@ -132,24 +132,22 @@ namespace Multiset omega exact Nat.le_of_pred_lt (han.trans_lt <| by simpa using hsm hn) -end Multiset - @[mono] -theorem Multiset.toEnumFinset_mono {m₁ m₂ : Multiset α} (h : m₁ ≤ m₂) : +theorem toEnumFinset_mono {m₁ m₂ : Multiset α} (h : m₁ ≤ m₂) : m₁.toEnumFinset ⊆ m₂.toEnumFinset := by intro p simp only [Multiset.mem_toEnumFinset] exact gt_of_ge_of_gt (Multiset.le_iff_count.mp h p.1) @[simp] -theorem Multiset.toEnumFinset_subset_iff {m₁ m₂ : Multiset α} : +theorem toEnumFinset_subset_iff {m₁ m₂ : Multiset α} : m₁.toEnumFinset ⊆ m₂.toEnumFinset ↔ m₁ ≤ m₂ := ⟨fun h ↦ by simpa using map_fst_le_of_subset_toEnumFinset h, Multiset.toEnumFinset_mono⟩ /-- The embedding from a multiset into `α × ℕ` where the second coordinate enumerates repeats. If you are looking for the function `m → α`, that would be plain `(↑)`. -/ @[simps] -def Multiset.coeEmbedding (m : Multiset α) : m ↪ α × ℕ where +def coeEmbedding (m : Multiset α) : m ↪ α × ℕ where toFun x := (x, x.2) inj' := by intro ⟨x, i, hi⟩ ⟨y, j, hj⟩ @@ -159,7 +157,7 @@ def Multiset.coeEmbedding (m : Multiset α) : m ↪ α × ℕ where /-- Another way to coerce a `Multiset` to a type is to go through `m.toEnumFinset` and coerce that `Finset` to a type. -/ @[simps] -def Multiset.coeEquiv (m : Multiset α) : m ≃ m.toEnumFinset where +def coeEquiv (m : Multiset α) : m ≃ m.toEnumFinset where toFun x := ⟨m.coeEmbedding x, by rw [Multiset.mem_toEnumFinset] @@ -176,14 +174,14 @@ def Multiset.coeEquiv (m : Multiset α) : m ≃ m.toEnumFinset where rfl @[simp] -theorem Multiset.toEmbedding_coeEquiv_trans (m : Multiset α) : +theorem toEmbedding_coeEquiv_trans (m : Multiset α) : m.coeEquiv.toEmbedding.trans (Function.Embedding.subtype _) = m.coeEmbedding := by ext <;> rfl @[irreducible] -instance Multiset.fintypeCoe : Fintype m := +instance fintypeCoe : Fintype m := Fintype.ofEquiv m.toEnumFinset m.coeEquiv.symm -theorem Multiset.map_univ_coeEmbedding (m : Multiset α) : +theorem map_univ_coeEmbedding (m : Multiset α) : (Finset.univ : Finset m).map m.coeEmbedding = m.toEnumFinset := by ext ⟨x, i⟩ simp only [Fin.exists_iff, Finset.mem_map, Finset.mem_univ, Multiset.coeEmbedding_apply, @@ -191,7 +189,7 @@ theorem Multiset.map_univ_coeEmbedding (m : Multiset α) : exists_prop, exists_eq_right_right, exists_eq_right, Multiset.mem_toEnumFinset, true_and] @[simp] -theorem Multiset.map_univ_coe (m : Multiset α) : +theorem map_univ_coe (m : Multiset α) : (Finset.univ : Finset m).val.map (fun x : m ↦ (x : α)) = m := by have := m.map_toEnumFinset_fst rw [← m.map_univ_coeEmbedding] at this @@ -199,36 +197,129 @@ theorem Multiset.map_univ_coe (m : Multiset α) : Function.comp_apply] using this @[simp] -theorem Multiset.map_univ {β : Type*} (m : Multiset α) (f : α → β) : +theorem map_univ {β : Type*} (m : Multiset α) (f : α → β) : ((Finset.univ : Finset m).val.map fun (x : m) ↦ f (x : α)) = m.map f := by erw [← Multiset.map_map, Multiset.map_univ_coe] @[simp] -theorem Multiset.card_toEnumFinset (m : Multiset α) : m.toEnumFinset.card = Multiset.card m := by +theorem card_toEnumFinset (m : Multiset α) : m.toEnumFinset.card = Multiset.card m := by rw [Finset.card, ← Multiset.card_map Prod.fst m.toEnumFinset.val] congr exact m.map_toEnumFinset_fst @[simp] -theorem Multiset.card_coe (m : Multiset α) : Fintype.card m = Multiset.card m := by +theorem card_coe (m : Multiset α) : Fintype.card m = Multiset.card m := by rw [Fintype.card_congr m.coeEquiv] simp only [Fintype.card_coe, card_toEnumFinset] @[to_additive] -theorem Multiset.prod_eq_prod_coe [CommMonoid α] (m : Multiset α) : m.prod = ∏ x : m, (x : α) := by +theorem prod_eq_prod_coe [CommMonoid α] (m : Multiset α) : m.prod = ∏ x : m, (x : α) := by congr simp @[to_additive] -theorem Multiset.prod_eq_prod_toEnumFinset [CommMonoid α] (m : Multiset α) : +theorem prod_eq_prod_toEnumFinset [CommMonoid α] (m : Multiset α) : m.prod = ∏ x ∈ m.toEnumFinset, x.1 := by congr simp @[to_additive] -theorem Multiset.prod_toEnumFinset {β : Type*} [CommMonoid β] (m : Multiset α) (f : α → ℕ → β) : +theorem prod_toEnumFinset {β : Type*} [CommMonoid β] (m : Multiset α) (f : α → ℕ → β) : ∏ x ∈ m.toEnumFinset, f x.1 x.2 = ∏ x : m, f x x.2 := by rw [Fintype.prod_equiv m.coeEquiv (fun x ↦ f x x.2) fun x ↦ f x.1.1 x.1.2] · rw [← m.toEnumFinset.prod_coe_sort fun x ↦ f x.1 x.2] · intro x rfl + +/-- +If `s = t` then there's an equivalence between the appropriate types. +-/ +@[simps] +def cast {s t : Multiset α} (h : s = t) : s ≃ t where + toFun x := ⟨x.1, x.2.cast (by simp [h])⟩ + invFun x := ⟨x.1, x.2.cast (by simp [h])⟩ + left_inv x := rfl + right_inv x := rfl + +instance : IsEmpty (0 : Multiset α) := Fintype.card_eq_zero_iff.mp (by simp) + +instance : IsEmpty (∅ : Multiset α) := Fintype.card_eq_zero_iff.mp (by simp) + +/-- +`v ::ₘ m` is equivalent to `Option m` by mapping one `v` to `none` and everything else to `m`. +-/ +def consEquiv {v : α} : v ::ₘ m ≃ Option m where + toFun x := if h : x.1 = v ∧ x.2.val = m.count v then none else some ⟨x.1, ⟨x.2, by + by_cases hv : x.1 = v + · simp only [hv, true_and] at h ⊢ + apply lt_of_le_of_ne (Nat.le_of_lt_add_one _) h + convert x.2.2 using 1 + simp [hv] + · convert x.2.2 using 1 + exact (count_cons_of_ne hv _).symm + ⟩⟩ + invFun x := x.elim ⟨v, ⟨m.count v, by simp⟩⟩ (fun x ↦ ⟨x.1, x.2.castLE (count_le_count_cons ..)⟩) + left_inv := by + rintro ⟨x, hx⟩ + dsimp only + split + · rename_i h + obtain ⟨rfl, h2⟩ := h + simp [← h2] + · simp + right_inv := by + rintro (_ | x) + · simp + · simp only [Option.elim_some, Nat.zero_eq, Fin.coe_castLE, Fin.eta, Sigma.eta, dite_eq_ite, + ite_eq_right_iff, reduceCtorEq, imp_false, not_and] + rintro rfl + exact x.2.2.ne + +@[simp] +lemma consEquiv_symm_none {v : α} : + (consEquiv (m := m) (v := v)).symm none = + ⟨v, ⟨m.count v, (count_cons_self v m) ▸ (Nat.lt_add_one _)⟩⟩ := + rfl + +@[simp] +lemma consEquiv_symm_some {v : α} {x : m} : + (consEquiv (v := v)).symm (some x) = + ⟨x, x.2.castLE (count_le_count_cons ..)⟩ := + rfl + +lemma coe_consEquiv_of_ne {v : α} (x : v ::ₘ m) (hx : ↑x ≠ v) : + consEquiv x = some ⟨x.1, x.2.cast (by simp [hx])⟩ := by + simp [consEquiv, hx] + rfl + +lemma coe_consEquiv_of_eq_of_eq {v : α} (x : v ::ₘ m) (hx : ↑x = v) (hx2 : x.2 = m.count v) : + consEquiv x = none := by simp [consEquiv, hx, hx2] + +lemma coe_consEquiv_of_eq_of_lt {v : α} (x : v ::ₘ m) (hx : ↑x = v) (hx2 : x.2 < m.count v) : + consEquiv x = some ⟨x.1, ⟨x.2, by simpa [hx]⟩⟩ := by simp [consEquiv, hx, hx2.ne] + +/-- +There is some equivalence between `m` and `m.map f` which respects `f`. +-/ +def mapEquiv_aux (m : Multiset α) (f : α → β) : + Squash { v : m ≃ m.map f // ∀ a : m, v a = f a} := + Quotient.recOnSubsingleton m fun l ↦ .mk <| + List.recOn l + ⟨@Equiv.equivOfIsEmpty _ _ (by dsimp; infer_instance) (by dsimp; infer_instance), by simp⟩ + fun a s ⟨v, hv⟩ ↦ ⟨Multiset.consEquiv.trans v.optionCongr |>.trans + Multiset.consEquiv.symm |>.trans (Multiset.cast (map_cons f a s)).symm, fun x ↦ by + simp only [consEquiv, Equiv.trans_apply, Equiv.coe_fn_mk, Equiv.optionCongr_apply, + Equiv.coe_fn_symm_mk, cast_symm_apply_fst] + split <;> simp_all⟩ + +/-- +One of the possible equivalences from `Multiset.mapEquiv_aux`, selected using choice. +-/ +noncomputable def mapEquiv (s : Multiset α) (f : α → β) : s ≃ s.map f := + (Multiset.mapEquiv_aux s f).out.1 + +@[simp] +theorem mapEquiv_apply (s : Multiset α) (f : α → β) (v : s) : s.mapEquiv f v = f v := + (Multiset.mapEquiv_aux s f).out.2 v + +end Multiset diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 1847740ee58e0..053a63b88459a 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -7,7 +7,9 @@ import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Order.Nonneg.Basic import Mathlib.Algebra.Order.Ring.Unbundled.Rat import Mathlib.Algebra.Ring.Rat -import Mathlib.Order.GaloisConnection +import Mathlib.Data.Set.Operations +import Mathlib.Order.Bounds.Defs +import Mathlib.Order.GaloisConnection.Defs /-! # Nonnegative rationals @@ -33,7 +35,7 @@ Whenever you state a lemma about the coercion `ℚ≥0 → ℚ`, check that Lean `Subtype.val`. Else your lemma will never apply. -/ -assert_not_exists OrderedCommMonoid +assert_not_exists CompleteLattice OrderedCommMonoid library_note "specialised high priority simp lemma" /-- It sometimes happens that a `@[simp]` lemma declared early in the library can be proved by `simp` @@ -156,8 +158,7 @@ theorem coe_le_coe : (p : ℚ) ≤ q ↔ p ≤ q := theorem coe_lt_coe : (p : ℚ) < q ↔ p < q := Iff.rfl --- `cast_pos`, defined in a later file, makes this lemma redundant -@[simp, norm_cast, nolint simpNF] +@[norm_cast] theorem coe_pos : (0 : ℚ) < q ↔ 0 < q := Iff.rfl @@ -212,13 +213,11 @@ theorem bddAbove_coe {s : Set ℚ≥0} : BddAbove ((↑) '' s : Set ℚ) ↔ Bdd theorem bddBelow_coe (s : Set ℚ≥0) : BddBelow (((↑) : ℚ≥0 → ℚ) '' s) := ⟨0, fun _ ⟨q, _, h⟩ ↦ h ▸ q.2⟩ --- `cast_max`, defined in a later file, makes this lemma redundant -@[simp, norm_cast, nolint simpNF] +@[norm_cast] theorem coe_max (x y : ℚ≥0) : ((max x y : ℚ≥0) : ℚ) = max (x : ℚ) (y : ℚ) := coe_mono.map_max --- `cast_max`, defined in a later file, makes this lemma redundant -@[simp, norm_cast, nolint simpNF] +@[norm_cast] theorem coe_min (x y : ℚ≥0) : ((min x y : ℚ≥0) : ℚ) = min (x : ℚ) (y : ℚ) := coe_mono.map_min diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index 68c6a36bd742b..be4890f48ea14 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -50,7 +50,7 @@ assert_not_exists Star open Function -- to ensure these instances are computable -/-- Nonnegative real numbers. -/ +/-- Nonnegative real numbers, denoted as `ℝ≥0` withinin the NNReal namespace -/ def NNReal := { r : ℝ // 0 ≤ r } deriving Zero, One, Semiring, StrictOrderedSemiring, CommMonoidWithZero, CommSemiring, PartialOrder, SemilatticeInf, SemilatticeSup, DistribLattice, OrderedCommSemiring, @@ -58,7 +58,7 @@ def NNReal := { r : ℝ // 0 ≤ r } deriving namespace NNReal -scoped notation "ℝ≥0" => NNReal +@[inherit_doc] scoped notation "ℝ≥0" => NNReal instance : CanonicallyOrderedAdd ℝ≥0 := Nonneg.canonicallyOrderedAdd instance : NoZeroDivisors ℝ≥0 := Nonneg.noZeroDivisors @@ -68,8 +68,11 @@ instance instArchimedean : Archimedean ℝ≥0 := Nonneg.instArchimedean instance instMulArchimedean : MulArchimedean ℝ≥0 := Nonneg.instMulArchimedean instance : Min ℝ≥0 := SemilatticeInf.toMin instance : Max ℝ≥0 := SemilatticeSup.toMax -noncomputable instance : Sub ℝ≥0 := Nonneg.sub -noncomputable instance : OrderedSub ℝ≥0 := Nonneg.orderedSub +instance : Sub ℝ≥0 := Nonneg.sub +instance : OrderedSub ℝ≥0 := Nonneg.orderedSub + +-- a computable copy of `Nonneg.instNNRatCast` +instance : NNRatCast ℝ≥0 where nnratCast r := ⟨r, r.cast_nonneg⟩ noncomputable instance : LinearOrderedSemifield ℝ≥0 := Nonneg.linearOrderedSemifield @@ -105,7 +108,7 @@ protected theorem «exists» {p : ℝ≥0 → Prop} : Subtype.exists /-- Reinterpret a real number `r` as a non-negative real number. Returns `0` if `r < 0`. -/ -noncomputable def _root_.Real.toNNReal (r : ℝ) : ℝ≥0 := +def _root_.Real.toNNReal (r : ℝ) : ℝ≥0 := ⟨max r 0, le_max_right _ _⟩ theorem _root_.Real.coe_toNNReal (r : ℝ) (hr : 0 ≤ r) : (Real.toNNReal r : ℝ) = r := @@ -127,7 +130,7 @@ example : One ℝ≥0 := by infer_instance example : Add ℝ≥0 := by infer_instance -noncomputable example : Sub ℝ≥0 := by infer_instance +example : Sub ℝ≥0 := by infer_instance example : Mul ℝ≥0 := by infer_instance @@ -325,7 +328,7 @@ theorem _root_.Real.toNNReal_ofNat (n : ℕ) [n.AtLeastTwo] : toNNReal_coe_nat n /-- `Real.toNNReal` and `NNReal.toReal : ℝ≥0 → ℝ` form a Galois insertion. -/ -noncomputable def gi : GaloisInsertion Real.toNNReal (↑) := +def gi : GaloisInsertion Real.toNNReal (↑) := GaloisInsertion.monotoneIntro NNReal.coe_mono Real.toNNReal_mono Real.le_coe_toNNReal fun _ => Real.toNNReal_coe @@ -797,7 +800,7 @@ nonrec theorem div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : div_le_div_iff_of_pos_left a0 b0 c0 theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := - le_of_forall_ge_of_dense fun a ha => by + le_of_forall_lt_imp_le_of_dense fun a ha => by have hx : x ≠ 0 := pos_iff_ne_zero.1 (lt_of_le_of_lt (zero_le _) ha) have hx' : x⁻¹ ≠ 0 := by rwa [Ne, inv_eq_zero] have : a * x⁻¹ < 1 := by rwa [← lt_inv_iff_mul_lt hx', inv_inv] @@ -977,6 +980,10 @@ theorem Real.exists_lt_of_strictMono [h : Nontrivial Γ₀ˣ] {f : Γ₀ →*₀ end StrictMono +/-- While not very useful, this instance uses the same representation as `Real.instRepr`. -/ +unsafe instance : Repr ℝ≥0 where + reprPrec r _ := f!"({repr r.val}).toNNReal" + namespace Mathlib.Meta.Positivity open Lean Meta Qq Function diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index d190cb22f4599..617f4537a4023 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro -/ import Mathlib.Algebra.Divisibility.Hom import Mathlib.Algebra.Group.Even -import Mathlib.Algebra.Group.TypeTags.Hom +import Mathlib.Algebra.Group.Nat.Hom import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.Nat @@ -92,21 +92,7 @@ end Nat section AddMonoidHomClass -variable {A B F : Type*} [AddMonoidWithOne B] [FunLike F ℕ A] - -theorem ext_nat' [AddMonoid A] [AddMonoidHomClass F ℕ A] (f g : F) (h : f 1 = g 1) : f = g := - DFunLike.ext f g <| by - intro n - induction n with - | zero => simp_rw [map_zero f, map_zero g] - | succ n ihn => - simp [h, ihn] - -@[ext] -theorem AddMonoidHom.ext_nat [AddMonoid A] {f g : ℕ →+ A} : f 1 = g 1 → f = g := - ext_nat' f g - -variable [AddMonoidWithOne A] +variable {A B F : Type*} [AddMonoidWithOne B] [FunLike F ℕ A] [AddMonoidWithOne A] -- these versions are primed so that the `RingHomClass` versions aren't theorem eq_natCast' [AddMonoidHomClass F ℕ A] (f : F) (h1 : f 1 = 1) : ∀ n : ℕ, f n = n @@ -195,78 +181,6 @@ instance Nat.uniqueRingHom {R : Type*} [NonAssocSemiring R] : Unique (ℕ →+* default := Nat.castRingHom R uniq := RingHom.eq_natCast' -section Monoid -variable (α) [Monoid α] (β) [AddMonoid β] - -/-- Additive homomorphisms from `ℕ` are defined by the image of `1`. -/ -def multiplesHom : β ≃ (ℕ →+ β) where - toFun x := - { toFun := fun n ↦ n • x - map_zero' := zero_nsmul x - map_add' := fun _ _ ↦ add_nsmul _ _ _ } - invFun f := f 1 - left_inv := one_nsmul - right_inv f := AddMonoidHom.ext_nat <| one_nsmul (f 1) - -/-- Monoid homomorphisms from `Multiplicative ℕ` are defined by the image -of `Multiplicative.ofAdd 1`. -/ -@[to_additive existing] -def powersHom : α ≃ (Multiplicative ℕ →* α) := - Additive.ofMul.trans <| (multiplesHom _).trans <| AddMonoidHom.toMultiplicative'' - -variable {α} - --- TODO: can `to_additive` generate the following lemmas automatically? - -lemma multiplesHom_apply (x : β) (n : ℕ) : multiplesHom β x n = n • x := rfl - -@[to_additive existing (attr := simp)] -lemma powersHom_apply (x : α) (n : Multiplicative ℕ) : - powersHom α x n = x ^ n.toAdd := rfl - -lemma multiplesHom_symm_apply (f : ℕ →+ β) : (multiplesHom β).symm f = f 1 := rfl - -@[to_additive existing (attr := simp)] -lemma powersHom_symm_apply (f : Multiplicative ℕ →* α) : - (powersHom α).symm f = f (Multiplicative.ofAdd 1) := rfl - -lemma MonoidHom.apply_mnat (f : Multiplicative ℕ →* α) (n : Multiplicative ℕ) : - f n = f (Multiplicative.ofAdd 1) ^ n.toAdd := by - rw [← powersHom_symm_apply, ← powersHom_apply, Equiv.apply_symm_apply] - -@[ext] -lemma MonoidHom.ext_mnat ⦃f g : Multiplicative ℕ →* α⦄ - (h : f (Multiplicative.ofAdd 1) = g (Multiplicative.ofAdd 1)) : f = g := - MonoidHom.ext fun n ↦ by rw [f.apply_mnat, g.apply_mnat, h] - -lemma AddMonoidHom.apply_nat (f : ℕ →+ β) (n : ℕ) : f n = n • f 1 := by - rw [← multiplesHom_symm_apply, ← multiplesHom_apply, Equiv.apply_symm_apply] - -end Monoid - -section CommMonoid -variable (α) [CommMonoid α] (β) [AddCommMonoid β] - -/-- If `α` is commutative, `multiplesHom` is an additive equivalence. -/ -def multiplesAddHom : β ≃+ (ℕ →+ β) := - { multiplesHom β with map_add' := fun a b ↦ AddMonoidHom.ext fun n ↦ by simp [nsmul_add] } - -/-- If `α` is commutative, `powersHom` is a multiplicative equivalence. -/ -def powersMulHom : α ≃* (Multiplicative ℕ →* α) := - { powersHom α with map_mul' := fun a b ↦ MonoidHom.ext fun n ↦ by simp [mul_pow] } - -@[simp] lemma multiplesAddHom_apply (x : β) (n : ℕ) : multiplesAddHom β x n = n • x := rfl - -@[simp] -lemma powersMulHom_apply (x : α) (n : Multiplicative ℕ) : powersMulHom α x n = x ^ n.toAdd := rfl - -@[simp] lemma multiplesAddHom_symm_apply (f : ℕ →+ β) : (multiplesAddHom β).symm f = f 1 := rfl - -@[simp] lemma powersMulHom_symm_apply (f : Multiplicative ℕ →* α) : - (powersMulHom α).symm f = f (ofAdd 1) := rfl - -end CommMonoid - namespace Pi variable {π : α → Type*} [∀ a, NatCast (π a)] @@ -286,6 +200,8 @@ theorem natCast_def (n : ℕ) : (n : ∀ a, π a) = fun _ ↦ ↑n := @[simp] theorem ofNat_apply (n : ℕ) [n.AtLeastTwo] (a : α) : (OfNat.ofNat n : ∀ a, π a) a = n := rfl +lemma ofNat_def (n : ℕ) [n.AtLeastTwo] : (OfNat.ofNat n : ∀ a, π a) = fun _ ↦ OfNat.ofNat n := rfl + end Pi theorem Sum.elim_natCast_natCast {α β γ : Type*} [NatCast γ] (n : ℕ) : diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index c98ee34c8d002..b75fe2b7a54fa 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -104,11 +104,7 @@ theorem sum_range_choose_halfway (m : ℕ) : (∑ i ∈ range (m + 1), (2 * m + ∑ i ∈ Ico (m + 1) (2 * m + 2), (2 * m + 1).choose i := by rw [range_eq_Ico, sum_Ico_reflect _ _ (by omega)] congr - have A : m + 1 ≤ 2 * m + 1 := by omega - rw [add_comm, add_tsub_assoc_of_le A, ← add_comm] - congr - rw [tsub_eq_iff_eq_add_of_le A] - ring + omega _ = ∑ i ∈ range (2 * m + 2), (2 * m + 1).choose i := sum_range_add_sum_Ico _ (by omega) _ = 2 ^ (2 * m + 1) := sum_range_choose (2 * m + 1) _ = 2 * 4 ^ m := by rw [pow_succ, pow_mul, mul_comm]; rfl @@ -234,9 +230,9 @@ theorem sum_antidiagonal_choose_succ_mul (f : ℕ → ℕ → R) (n : ℕ) : simpa only [nsmul_eq_mul] using sum_antidiagonal_choose_succ_nsmul f n theorem sum_antidiagonal_choose_add (d n : ℕ) : - (∑ ij ∈ antidiagonal n, (d + ij.2).choose d) = (d + n).choose d + (d + n).choose (d + 1) := by + (∑ ij ∈ antidiagonal n, (d + ij.2).choose d) = (d + n + 1).choose (d + 1) := by induction n with | zero => simp - | succ n hn => simpa [Nat.sum_antidiagonal_succ] using hn + | succ n hn => rw [Nat.sum_antidiagonal_succ, hn, Nat.choose_succ_succ (d + (n + 1)), ← add_assoc] end Finset diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 415504a2a854d..89a1dbbbbec52 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -493,16 +493,10 @@ lemma div_mul_div_le_div (a b c : ℕ) : a / c * b / a ≤ b / c := by lemma eq_zero_of_le_half (h : n ≤ n / 2) : n = 0 := eq_zero_of_le_div (Nat.le_refl _) h lemma le_half_of_half_lt_sub (h : a / 2 < a - b) : b ≤ a / 2 := by - rw [Nat.le_div_iff_mul_le Nat.two_pos] - rw [Nat.div_lt_iff_lt_mul Nat.two_pos, Nat.sub_mul, Nat.lt_sub_iff_add_lt, - Nat.mul_two a] at h - exact Nat.le_of_lt (Nat.lt_of_add_lt_add_left h) + omega lemma half_le_of_sub_le_half (h : a - b ≤ a / 2) : a / 2 ≤ b := by - rw [Nat.le_div_iff_mul_le Nat.two_pos, Nat.sub_mul, Nat.sub_le_iff_le_add, - Nat.mul_two, Nat.add_le_add_iff_left] at h - rw [← Nat.mul_div_left b Nat.two_pos] - exact Nat.div_le_div_right h + omega protected lemma div_le_of_le_mul' (h : m ≤ k * n) : m / k ≤ n := by obtain rfl | hk := k.eq_zero_or_pos diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index a56075aac773f..15357904895cd 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -556,10 +556,7 @@ theorem eq_iff_prime_padicValNat_eq (a b : ℕ) (ha : a ≠ 0) (hb : b ≠ 0) : theorem prod_pow_prime_padicValNat (n : Nat) (hn : n ≠ 0) (m : Nat) (pr : n < m) : ∏ p ∈ range m with p.Prime, p ^ padicValNat p n = n := by - -- Porting note: was `nth_rw_rhs` - conv => - rhs - rw [← factorization_prod_pow_eq_self hn] + nth_rw 2 [← factorization_prod_pow_eq_self hn] rw [eq_comm] apply Finset.prod_subset_one_on_sdiff · exact fun p hp => Finset.mem_filter.mpr ⟨Finset.mem_range.2 <| pr.trans_le' <| diff --git a/Mathlib/Data/Nat/Hyperoperation.lean b/Mathlib/Data/Nat/Hyperoperation.lean index ea0cc84e74cc6..5fadfa52f931a 100644 --- a/Mathlib/Data/Nat/Hyperoperation.lean +++ b/Mathlib/Data/Nat/Hyperoperation.lean @@ -104,7 +104,7 @@ theorem hyperoperation_ge_four_zero (n k : ℕ) : hyperoperation (n + 4) 0 k = if Even k then 1 else 0 := by induction' k with kk kih · rw [hyperoperation_ge_three_eq_one] - simp only [even_zero, if_true] + simp only [Even.zero, if_true] · rw [hyperoperation_recursion] rw [kih] simp_rw [Nat.even_add_one] diff --git a/Mathlib/Data/Nat/Lattice.lean b/Mathlib/Data/Nat/Lattice.lean index cd3c50fdd53c6..1ede706e2e0fb 100644 --- a/Mathlib/Data/Nat/Lattice.lean +++ b/Mathlib/Data/Nat/Lattice.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Floris van Doorn, Gabriel Ebner, Yury Kudryashov -/ +import Mathlib.Data.Set.Accumulate import Mathlib.Order.ConditionallyCompleteLattice.Finset import Mathlib.Order.Interval.Finset.Nat @@ -21,17 +22,19 @@ open Set namespace Nat -open scoped Classical - +open scoped Classical in noncomputable instance : InfSet ℕ := ⟨fun s ↦ if h : ∃ n, n ∈ s then @Nat.find (fun n ↦ n ∈ s) _ h else 0⟩ +open scoped Classical in noncomputable instance : SupSet ℕ := ⟨fun s ↦ if h : ∃ n, ∀ a ∈ s, a ≤ n then @Nat.find (fun n ↦ ∀ a ∈ s, a ≤ n) _ h else 0⟩ +open scoped Classical in theorem sInf_def {s : Set ℕ} (h : s.Nonempty) : sInf s = @Nat.find (fun n ↦ n ∈ s) _ h := dif_pos _ +open scoped Classical in theorem sSup_def {s : Set ℕ} (h : ∃ n, ∀ a ∈ s, a ≤ n) : sSup s = @Nat.find (fun n ↦ ∀ a ∈ s, a ≤ n) _ h := dif_pos _ @@ -65,15 +68,18 @@ lemma iInf_const_zero {ι : Sort*} : ⨅ _ : ι, 0 = 0 := (isEmpty_or_nonempty ι).elim (fun h ↦ by simp) fun h ↦ sInf_eq_zero.2 <| by simp theorem sInf_mem {s : Set ℕ} (h : s.Nonempty) : sInf s ∈ s := by + classical rw [Nat.sInf_def h] exact Nat.find_spec h theorem not_mem_of_lt_sInf {s : Set ℕ} {m : ℕ} (hm : m < sInf s) : m ∉ s := by + classical cases eq_empty_or_nonempty s with | inl h => subst h; apply not_mem_empty | inr h => rw [Nat.sInf_def h] at hm; exact Nat.find_min h hm protected theorem sInf_le {s : Set ℕ} {m : ℕ} (hm : m ∈ s) : sInf s ≤ m := by + classical rw [Nat.sInf_def ⟨m, hm⟩] exact Nat.find_min' ⟨m, hm⟩ hm @@ -95,6 +101,7 @@ theorem eq_Ici_of_nonempty_of_upward_closed {s : Set ℕ} (hs : s.Nonempty) theorem sInf_upward_closed_eq_succ_iff {s : Set ℕ} (hs : ∀ k₁ k₂ : ℕ, k₁ ≤ k₂ → k₁ ∈ s → k₂ ∈ s) (k : ℕ) : sInf s = k + 1 ↔ k + 1 ∈ s ∧ k ∉ s := by + classical constructor · intro H rw [eq_Ici_of_nonempty_of_upward_closed (nonempty_of_sInf_eq_succ _) hs, H, mem_Ici, mem_Ici] @@ -110,6 +117,7 @@ theorem sInf_upward_closed_eq_succ_iff {s : Set ℕ} (hs : ∀ k₁ k₂ : ℕ, instance : Lattice ℕ := LinearOrder.toLattice +open scoped Classical in noncomputable instance : ConditionallyCompleteLinearOrderBot ℕ := { (inferInstance : OrderBot ℕ), (LinearOrder.toLattice : Lattice ℕ), (inferInstance : LinearOrder ℕ) with @@ -140,6 +148,7 @@ theorem sSup_mem {s : Set ℕ} (h₁ : s.Nonempty) (h₂ : BddAbove s) : sSup s theorem sInf_add {n : ℕ} {p : ℕ → Prop} (hn : n ≤ sInf { m | p m }) : sInf { m | p (m + n) } + n = sInf { m | p m } := by + classical obtain h | ⟨m, hm⟩ := { m | p (m + n) }.eq_empty_or_nonempty · rw [h, Nat.sInf_empty, zero_add] obtain hnp | hnp := hn.eq_or_lt @@ -227,4 +236,7 @@ theorem biInter_le_succ (u : ℕ → Set α) (n : ℕ) : ⋂ k ≤ n + 1, u k = theorem biInter_le_succ' (u : ℕ → Set α) (n : ℕ) : ⋂ k ≤ n + 1, u k = u 0 ∩ ⋂ k ≤ n, u (k + 1) := Nat.iInf_le_succ' u n +theorem accumulate_succ (u : ℕ → Set α) (n : ℕ) : + Accumulate u (n + 1) = Accumulate u n ∪ u (n + 1) := biUnion_le_succ u n + end Set diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index b2e55d35fe599..2d4196b2fe33a 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -330,4 +330,117 @@ theorem log_le_clog (b n : ℕ) : log b n ≤ clog b n := by exact (Nat.pow_le_pow_iff_right hb).1 ((pow_log_le_self b n.succ_ne_zero).trans <| le_pow_clog hb _) +/-! ### Computating the logarithm efficiently -/ +section computation + +private lemma logC_aux {m b : ℕ} (hb : 1 < b) (hbm : b ≤ m) : m / (b * b) < m / b := by + have hb' : 0 < b := zero_lt_of_lt hb + rw [div_lt_iff_lt_mul (Nat.mul_pos hb' hb'), ← Nat.mul_assoc, ← div_lt_iff_lt_mul hb'] + exact (Nat.lt_mul_iff_one_lt_right (Nat.div_pos hbm hb')).2 hb + +-- This option is necessary because of lean4#2920 +set_option linter.unusedVariables false in +/-- +An alternate definition for `Nat.log` which computes more efficiently. For mathematical purposes, +use `Nat.log` instead, and see `Nat.log_eq_logC`. + +Note a tail-recursive version of `Nat.log` is also possible: +``` +def logTR (b n : ℕ) : ℕ := + let rec go : ℕ → ℕ → ℕ | n, acc => if h : b ≤ n ∧ 1 < b then go (n / b) (acc + 1) else acc + decreasing_by + have : n / b < n := Nat.div_lt_self (by omega) h.2 + decreasing_trivial + go n 0 +``` +but performs worse for large numbers than `Nat.logC`: +``` +#eval Nat.logTR 2 (2 ^ 1000000) +#eval Nat.logC 2 (2 ^ 1000000) +``` + +The definition `Nat.logC` is not tail-recursive, however, but the stack limit will only be reached +if the output size is around 2^10000, meaning the input will be around 2^(2^10000), which will +take far too long to compute in the first place. + +Adapted from https://downloads.haskell.org/~ghc/9.0.1/docs/html/libraries/ghc-bignum-1.0/GHC-Num-BigNat.html#v:bigNatLogBase-35- +-/ +@[pp_nodot] def logC (b m : ℕ) : ℕ := + if h : 1 < b then let (_, e) := step b h; e else 0 where + /-- + An auxiliary definition for `Nat.logC`, where the base of the logarithm is _squared_ in each + loop. This allows significantly faster computation of the logarithm: it takes logarithmic time + in the size of the output, rather than linear time. + -/ + step (pw : ℕ) (hpw : 1 < pw) : ℕ × ℕ := + if h : m < pw then (m, 0) + else + let (q, e) := step (pw * pw) (Nat.mul_lt_mul_of_lt_of_lt hpw hpw) + if q < pw then (q, 2 * e) else (q / pw, 2 * e + 1) + termination_by m / pw + decreasing_by + have : m / (pw * pw) < m / pw := logC_aux hpw (le_of_not_lt h) + decreasing_trivial + +private lemma logC_step {m pw q e : ℕ} (hpw : 1 < pw) (hqe : logC.step m pw hpw = (q, e)) : + pw ^ e * q ≤ m ∧ q < pw ∧ (m < pw ^ e * (q + 1)) ∧ (0 < m → 0 < q) := by + induction pw, hpw using logC.step.induct m generalizing q e with + | case1 pw hpw hmpw => + rw [logC.step, dif_pos hmpw] at hqe + cases hqe + simpa + | case2 pw hpw hmpw q' e' hqe' hqpw ih => + simp only [logC.step, dif_neg hmpw, hqe', if_pos hqpw] at hqe + cases hqe + rw [Nat.pow_mul, Nat.pow_two] + exact ⟨(ih hqe').1, hqpw, (ih hqe').2.2⟩ + | case3 pw hpw hmpw q' e' hqe' hqpw ih => + simp only [Nat.logC.step, dif_neg hmpw, hqe', if_neg hqpw] at hqe + cases hqe + rw [Nat.pow_succ, Nat.mul_assoc, Nat.pow_mul, Nat.pow_two, Nat.mul_assoc] + refine ⟨(ih hqe').1.trans' (Nat.mul_le_mul_left _ (Nat.mul_div_le _ _)), + Nat.div_lt_of_lt_mul (ih hqe').2.1, (ih hqe').2.2.1.trans_le ?_, + fun _ => Nat.div_pos (le_of_not_lt hqpw) (by omega)⟩ + exact Nat.mul_le_mul_left _ (Nat.lt_mul_div_succ _ (zero_lt_of_lt hpw)) + +private lemma logC_spec {b m : ℕ} (hb : 1 < b) (hm : 0 < m) : + b ^ logC b m ≤ m ∧ m < b ^ (logC b m + 1) := by + rw [logC, dif_pos hb] + split + next q e heq => + obtain ⟨h₁, h₂, h₃, h₄⟩ := logC_step hb heq + exact ⟨h₁.trans' (Nat.le_mul_of_pos_right _ (h₄ hm)), h₃.trans_le (Nat.mul_le_mul_left _ h₂)⟩ + +private lemma logC_of_left_le_one {b m : ℕ} (hb : b ≤ 1) : logC b m = 0 := by + rw [logC, dif_neg hb.not_lt] + +private lemma logC_zero {b : ℕ} : + logC b 0 = 0 := by + rcases le_or_lt b 1 with hb | hb + case inl => exact logC_of_left_le_one hb + case inr => + rw [logC, dif_pos hb] + split + next q e heq => + rw [logC.step, dif_pos (zero_lt_of_lt hb)] at heq + rw [(Prod.mk.inj heq).2] + +/-- +The result of `Nat.log` agrees with the result of `Nat.logC`. The latter will be computed more +efficiently, but the former is easier to prove things about and has more lemmas. +This lemma is tagged @[csimp] so that the code generated for `Nat.log` uses `Nat.logC` instead. +-/ +@[csimp] theorem log_eq_logC : log = logC := by + ext b m + rcases le_or_lt b 1 with hb | hb + case inl => rw [logC_of_left_le_one hb, Nat.log_of_left_le_one hb] + case inr => + rcases eq_or_ne m 0 with rfl | hm + case inl => rw [Nat.log_zero_right, logC_zero] + case inr => + rw [Nat.log_eq_iff (Or.inr ⟨hb, hm⟩)] + exact logC_spec hb (zero_lt_of_ne_zero hm) + +end computation + end Nat diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index 5b4dd68288d83..0350d9d4b3cb9 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -640,8 +640,7 @@ theorem ofENat_lt {x y : ℕ∞} : ofENat x < ofENat y ↔ x < y := by section WithTopEquiv -open scoped Classical - +open scoped Classical in @[simp] theorem toWithTop_add {x y : PartENat} : toWithTop (x + y) = toWithTop x + toWithTop y := by refine PartENat.casesOn y ?_ ?_ <;> refine PartENat.casesOn x ?_ ?_ @@ -651,6 +650,7 @@ theorem toWithTop_add {x y : PartENat} : toWithTop (x + y) = toWithTop x + toWit · simp only [top_add, toWithTop_top', toWithTop_natCast', _root_.top_add, forall_const] · simp_rw [toWithTop_natCast', ← Nat.cast_add, toWithTop_natCast', forall_const] +open scoped Classical in /-- `Equiv` between `PartENat` and `ℕ∞` (for the order isomorphism see `withTopOrderIso`). -/ @[simps] diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index 8a275196668f5..d6f5815d92c9e 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -289,8 +289,6 @@ theorem orElse_eq_none' (o o' : Option α) : o.orElse (fun _ ↦ o') = none ↔ section -open scoped Classical - theorem choice_eq_none (α : Type*) [IsEmpty α] : choice α = none := dif_neg (not_nonempty_iff_imp_false.mpr isEmptyElim) diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index 007830364ceb2..fa6cbb13d14d7 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -432,6 +432,7 @@ theorem toPEquiv_trans (f : α ≃ β) (g : β ≃ γ) : theorem toPEquiv_symm (f : α ≃ β) : f.symm.toPEquiv = f.toPEquiv.symm := rfl +@[simp] theorem toPEquiv_apply (f : α ≃ β) (x : α) : f.toPEquiv x = some (f x) := rfl diff --git a/Mathlib/Data/PFunctor/Multivariate/W.lean b/Mathlib/Data/PFunctor/Multivariate/W.lean index d5250ef8adc16..9fce44601bb01 100644 --- a/Mathlib/Data/PFunctor/Multivariate/W.lean +++ b/Mathlib/Data/PFunctor/Multivariate/W.lean @@ -108,7 +108,6 @@ def wp : MvPFunctor n where B := P.WPath /-- W-type of `P` -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): used to have @[nolint has_nonempty_instance] def W (α : TypeVec n) : Type _ := P.wp α diff --git a/Mathlib/Data/PFunctor/Univariate/Basic.lean b/Mathlib/Data/PFunctor/Univariate/Basic.lean index 9f6015ea64417..d9421dae365a4 100644 --- a/Mathlib/Data/PFunctor/Univariate/Basic.lean +++ b/Mathlib/Data/PFunctor/Univariate/Basic.lean @@ -85,8 +85,6 @@ def W := /- inhabitants of W types is awkward to encode as an instance assumption because there needs to be a value `a : P.A` such that `P.B a` is empty to yield a finite tree -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- attribute [nolint has_nonempty_instance] W variable {P} diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 98ebf824ed437..5a508e4709c66 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -436,7 +436,7 @@ theorem bind_some (a : α) (f : α → Part β) : (some a).bind f = f a := theorem bind_of_mem {o : Part α} {a : α} (h : a ∈ o) (f : α → Part β) : o.bind f = f a := by rw [eq_some_iff.2 h, bind_some] -theorem bind_some_eq_map (f : α → β) (x : Part α) : x.bind (some ∘ f) = map f x := +theorem bind_some_eq_map (f : α → β) (x : Part α) : x.bind (fun y => some (f y)) = map f x := ext <| by simp [eq_comm] theorem bind_toOption (f : α → Part β) (o : Part α) [Decidable o.Dom] [∀ a, Decidable (f a).Dom] diff --git a/Mathlib/Data/Prod/Lex.lean b/Mathlib/Data/Prod/Lex.lean index 2aafe976d991a..10da7294c8f5b 100644 --- a/Mathlib/Data/Prod/Lex.lean +++ b/Mathlib/Data/Prod/Lex.lean @@ -6,6 +6,7 @@ Authors: Kim Morrison, Minchao Wu import Mathlib.Data.Prod.Basic import Mathlib.Order.Lattice import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Tactic.Tauto /-! # Lexicographic order @@ -44,16 +45,24 @@ instance instLE (α β : Type*) [LT α] [LE β] : LE (α ×ₗ β) where le := P instance instLT (α β : Type*) [LT α] [LT β] : LT (α ×ₗ β) where lt := Prod.Lex (· < ·) (· < ·) -theorem le_iff [LT α] [LE β] (a b : α × β) : - toLex a ≤ toLex b ↔ a.1 < b.1 ∨ a.1 = b.1 ∧ a.2 ≤ b.2 := +theorem toLex_le_toLex [LT α] [LE β] {x y : α × β} : + toLex x ≤ toLex y ↔ x.1 < y.1 ∨ x.1 = y.1 ∧ x.2 ≤ y.2 := Prod.lex_def -theorem lt_iff [LT α] [LT β] (a b : α × β) : - toLex a < toLex b ↔ a.1 < b.1 ∨ a.1 = b.1 ∧ a.2 < b.2 := +theorem toLex_lt_toLex [LT α] [LT β] {x y : α × β} : + toLex x < toLex y ↔ x.1 < y.1 ∨ x.1 = y.1 ∧ x.2 < y.2 := + Prod.lex_def + +lemma le_iff [LT α] [LE β] {x y : α ×ₗ β} : + x ≤ y ↔ (ofLex x).1 < (ofLex y).1 ∨ (ofLex x).1 = (ofLex y).1 ∧ (ofLex x).2 ≤ (ofLex y).2 := + Prod.lex_def + +lemma lt_iff [LT α] [LT β] {x y : α ×ₗ β} : + x < y ↔ (ofLex x).1 < (ofLex y).1 ∨ (ofLex x).1 = (ofLex y).1 ∧ (ofLex x).2 < (ofLex y).2 := Prod.lex_def instance [LT α] [LT β] [WellFoundedLT α] [WellFoundedLT β] : WellFoundedLT (α ×ₗ β) := - ⟨WellFounded.prod_lex wellFounded_lt wellFounded_lt⟩ + instIsWellFounded instance [LT α] [LT β] [WellFoundedLT α] [WellFoundedLT β] : WellFoundedRelation (α ×ₗ β) := ⟨(· < ·), wellFounded_lt⟩ @@ -97,19 +106,37 @@ instance preorder (α β : Type*) [Preorder α] [Preorder β] : Preorder (α × theorem monotone_fst [Preorder α] [LE β] (t c : α ×ₗ β) (h : t ≤ c) : (ofLex t).1 ≤ (ofLex c).1 := by - cases (Prod.Lex.le_iff t c).mp h with + cases toLex_le_toLex.mp h with | inl h' => exact h'.le | inr h' => exact h'.1.le section Preorder -variable [PartialOrder α] [Preorder β] +variable [PartialOrder α] [Preorder β] {x y : α × β} -theorem toLex_mono : Monotone (toLex : α × β → α ×ₗ β) := by - rintro ⟨a₁, b₁⟩ ⟨a₂, b₂⟩ ⟨ha, hb⟩ - obtain rfl | ha : a₁ = a₂ ∨ _ := ha.eq_or_lt - · exact right _ hb - · exact left _ _ ha +/-- Variant of `Prod.Lex.toLex_le_toLex` for partial orders. -/ +lemma toLex_le_toLex' : toLex x ≤ toLex y ↔ x.1 ≤ y.1 ∧ (x.1 = y.1 → x.2 ≤ y.2) := by + simp only [toLex_le_toLex, lt_iff_le_not_le, le_antisymm_iff] + tauto + +/-- Variant of `Prod.Lex.toLex_lt_toLex` for partial orders. -/ +lemma toLex_lt_toLex' : toLex x < toLex y ↔ x.1 ≤ y.1 ∧ (x.1 = y.1 → x.2 < y.2) := by + rw [toLex_lt_toLex] + simp only [lt_iff_le_not_le, le_antisymm_iff] + tauto + +/-- Variant of `Prod.Lex.le_iff` for partial orders. -/ +lemma le_iff' {x y : α ×ₗ β} : + x ≤ y ↔ (ofLex x).1 ≤ (ofLex y).1 ∧ ((ofLex x).1 = (ofLex y).1 → (ofLex x).2 ≤ (ofLex y).2) := + toLex_le_toLex' + +/-- Variant of `Prod.Lex.lt_iff` for partial orders. -/ +lemma lt_iff' {x y : α ×ₗ β} : + x < y ↔ (ofLex x).1 ≤ (ofLex y).1 ∧ ((ofLex x).1 = (ofLex y).1 → (ofLex x).2 < (ofLex y).2) := + toLex_lt_toLex' + +theorem toLex_mono : Monotone (toLex : α × β → α ×ₗ β) := + fun _x _y hxy ↦ toLex_le_toLex'.2 ⟨hxy.1, fun _ ↦ hxy.2⟩ theorem toLex_strictMono : StrictMono (toLex : α × β → α ×ₗ β) := by rintro ⟨a₁, b₁⟩ ⟨a₂, b₂⟩ h @@ -153,7 +180,7 @@ instance linearOrder (α β : Type*) [LinearOrder α] [LinearOrder β] : LinearO have : BEqOrd (α ×ₗ β) := ⟨by simp [compare_def, compareLex, compareOn, Ordering.then_eq_eq, compare_eq_iff_eq]⟩ have : LTOrd (α ×ₗ β) := ⟨by - simp [compare_def, compareLex, compareOn, Ordering.then_eq_lt, lt_iff, + simp [compare_def, compareLex, compareOn, Ordering.then_eq_lt, toLex_lt_toLex, compare_lt_iff_lt, compare_eq_iff_eq]⟩ convert LTCmp.eq_compareOfLessAndEq (cmp := compare) a b } diff --git a/Mathlib/Data/QPF/Univariate/Basic.lean b/Mathlib/Data/QPF/Univariate/Basic.lean index ebb6886d06e51..445a0779eb6e1 100644 --- a/Mathlib/Data/QPF/Univariate/Basic.lean +++ b/Mathlib/Data/QPF/Univariate/Basic.lean @@ -218,8 +218,6 @@ def Wsetoid : Setoid q.P.W := attribute [local instance] Wsetoid /-- inductive type defined as initial algebra of a Quotient of Polynomial Functor -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] def Fix (F : Type u → Type u) [q : QPF F] := Quotient (Wsetoid : Setoid q.P.W) diff --git a/Mathlib/Data/Rat/Floor.lean b/Mathlib/Data/Rat/Floor.lean index fa01f1ac3ee5f..42464ab2589e4 100644 --- a/Mathlib/Data/Rat/Floor.lean +++ b/Mathlib/Data/Rat/Floor.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Kappelmann -/ -import Mathlib.Algebra.Order.Floor +import Mathlib.Algebra.Order.Round import Mathlib.Data.Rat.Cast.Order import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Ring diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 039acaddf3405..3e9bf433eb681 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -42,7 +42,7 @@ theorem of_near (f : ℕ → ℚ) (x : ℝ) (h : ∀ ε > 0, ∃ i, ∀ j ≥ i, ⟨isCauSeq_iff_lift.2 (CauSeq.of_near _ (const abs x) h), sub_eq_zero.1 <| abs_eq_zero.1 <| - (eq_of_le_of_forall_le_of_dense (abs_nonneg _)) fun _ε ε0 => + (eq_of_le_of_forall_lt_imp_le_of_dense (abs_nonneg _)) fun _ε ε0 => mk_near_of_forall_near <| (h _ ε0).imp fun _i h j ij => le_of_lt (h j ij)⟩ theorem exists_floor (x : ℝ) : ∃ ub : ℤ, (ub : ℝ) ≤ x ∧ ∀ z : ℤ, (z : ℝ) ≤ x → z ≤ ub := @@ -88,7 +88,7 @@ theorem exists_isLUB (hne : s.Nonempty) (hbdd : BddAbove s) : ∃ x, IsLUB s x : simpa using sub_lt_iff_lt_add'.2 (lt_of_le_of_lt hy <| sub_lt_iff_lt_add.1 <| hf₂ _ k0 _ yS) let g : CauSeq ℚ abs := ⟨fun n => f n / n, hg⟩ refine ⟨mk g, ⟨fun x xS => ?_, fun y h => ?_⟩⟩ - · refine le_of_forall_ge_of_dense fun z xz => ?_ + · refine le_of_forall_lt_imp_le_of_dense fun z xz => ?_ cases' exists_nat_gt (x - z)⁻¹ with K hK refine le_mk_of_forall_le ⟨K, fun n nK => ?_⟩ replace xz := sub_pos.2 xz diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 65b39b8a38b14..efd505490dce2 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -543,7 +543,8 @@ noncomputable instance decidableEq (a b : ℝ) : Decidable (a = b) := by infer_i The representative chosen is the one passed in the VM to `Quot.mk`, so two cauchy sequences converging to the same number may be printed differently. -/ -unsafe instance : Repr ℝ where reprPrec r _ := "Real.ofCauchy " ++ repr r.cauchy +unsafe instance : Repr ℝ where + reprPrec r p := Repr.addAppParen ("Real.ofCauchy " ++ repr r.cauchy) p theorem le_mk_of_forall_le {f : CauSeq ℚ abs} : (∃ i, ∀ j ≥ i, x ≤ f j) → x ≤ mk f := by intro h diff --git a/Mathlib/Data/Real/EReal.lean b/Mathlib/Data/Real/EReal.lean index eb40ab8dbcdd0..8e308ba21e9aa 100644 --- a/Mathlib/Data/Real/EReal.lean +++ b/Mathlib/Data/Real/EReal.lean @@ -1291,7 +1291,7 @@ private lemma exists_lt_add_right {a b c : EReal} (hc : c < a + b) : ∃ b' < b, simp_rw [add_comm a] at hc ⊢; exact exists_lt_add_left hc lemma add_le_of_forall_lt {a b c : EReal} (h : ∀ a' < a, ∀ b' < b, a' + b' ≤ c) : a + b ≤ c := by - refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun d hd ↦ ?_ obtain ⟨a', ha', hd⟩ := exists_lt_add_left hd obtain ⟨b', hb', hd⟩ := exists_lt_add_right hd exact hd.le.trans (h _ ha' _ hb') @@ -1863,24 +1863,36 @@ lemma sign_mul_inv_abs' (a : EReal) : (sign a) * ((a.abs⁻¹ : ℝ≥0∞) : ER /-! #### Inversion and Positivity -/ +lemma bot_lt_inv (x : EReal) : ⊥ < x⁻¹ := by + cases x with + | h_bot => exact inv_bot ▸ bot_lt_zero + | h_top => exact EReal.inv_top ▸ bot_lt_zero + | h_real x => exact (coe_inv x).symm ▸ bot_lt_coe (x⁻¹) + +lemma inv_lt_top (x : EReal) : x⁻¹ < ⊤ := by + cases x with + | h_bot => exact inv_bot ▸ zero_lt_top + | h_top => exact EReal.inv_top ▸ zero_lt_top + | h_real x => exact (coe_inv x).symm ▸ coe_lt_top (x⁻¹) + lemma inv_nonneg_of_nonneg {a : EReal} (h : 0 ≤ a) : 0 ≤ a⁻¹ := by - induction a with + cases a with | h_bot | h_top => simp | h_real a => rw [← coe_inv a, EReal.coe_nonneg, inv_nonneg]; exact EReal.coe_nonneg.1 h lemma inv_nonpos_of_nonpos {a : EReal} (h : a ≤ 0) : a⁻¹ ≤ 0 := by - induction a with + cases a with | h_bot | h_top => simp | h_real a => rw [← coe_inv a, EReal.coe_nonpos, inv_nonpos]; exact EReal.coe_nonpos.1 h lemma inv_pos_of_pos_ne_top {a : EReal} (h : 0 < a) (h' : a ≠ ⊤) : 0 < a⁻¹ := by - induction a with + cases a with | h_bot => exact (not_lt_bot h).rec | h_real a => rw [← coe_inv a]; norm_cast at *; exact inv_pos_of_pos h | h_top => exact (h' (Eq.refl ⊤)).rec lemma inv_neg_of_neg_ne_bot {a : EReal} (h : a < 0) (h' : a ≠ ⊥) : a⁻¹ < 0 := by - induction a with + cases a with | h_bot => exact (h' (Eq.refl ⊥)).rec | h_real a => rw [← coe_inv a]; norm_cast at *; exact inv_lt_zero.2 h | h_top => exact (not_top_lt h).rec @@ -1956,11 +1968,11 @@ lemma mul_div_mul_cancel {a b c : EReal} (h₁ : c ≠ ⊥) (h₂ : c ≠ ⊤) ( congr exact mul_div_cancel h₁ h₂ h₃ -/-! #### Division Distributivity -/ - -lemma div_right_distrib_of_nonneg {a b c : EReal} (h : 0 ≤ a) (h' : 0 ≤ b) : - (a + b) / c = (a / c) + (b / c) := - EReal.right_distrib_of_nonneg h h' +lemma div_eq_iff {a b c : EReal} (hbot : b ≠ ⊥) (htop : b ≠ ⊤) (hzero : b ≠ 0) : + c / b = a ↔ c = a * b := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [← @mul_div_cancel c b hbot htop hzero, h, mul_comm a b] + · rw [h, mul_comm a b, ← mul_div b a b, @mul_div_cancel a b hbot htop hzero] /-! #### Division and Order -/ @@ -2022,6 +2034,18 @@ lemma div_le_iff_le_mul {a b c : EReal} (h : 0 < b) (h' : b ≠ ⊤) : rw [mul_div b c b, mul_comm b] exact StrictMono.le_iff_le (strictMono_div_right_of_pos h h') +lemma lt_div_iff {a b c : EReal} (h : 0 < b) (h' : b ≠ ⊤) : + a < c / b ↔ a * b < c := by + nth_rw 1 [← @mul_div_cancel a b (ne_bot_of_gt h) h' (ne_of_gt h)] + rw [EReal.mul_div b a b, mul_comm a b] + exact (strictMono_div_right_of_pos h h').lt_iff_lt + +lemma div_lt_iff {a b c : EReal} (h : 0 < b) (h' : b ≠ ⊤) : + c / b < a ↔ c < a * b := by + nth_rw 1 [← @mul_div_cancel a b (ne_bot_of_gt h) h' (ne_of_gt h)] + rw [EReal.mul_div b a b, mul_comm a b] + exact (strictMono_div_right_of_pos h h').lt_iff_lt + lemma div_nonneg {a b : EReal} (h : 0 ≤ a) (h' : 0 ≤ b) : 0 ≤ a / b := mul_nonneg h (inv_nonneg_of_nonneg h') @@ -2034,6 +2058,16 @@ lemma div_nonpos_of_nonneg_of_nonpos {a b : EReal} (h : 0 ≤ a) (h' : b ≤ 0) lemma div_nonneg_of_nonpos_of_nonpos {a b : EReal} (h : a ≤ 0) (h' : b ≤ 0) : 0 ≤ a / b := le_of_eq_of_le (Eq.symm zero_div) (div_le_div_right_of_nonpos h' h) +/-! #### Division Distributivity -/ + +lemma div_right_distrib_of_nonneg {a b c : EReal} (h : 0 ≤ a) (h' : 0 ≤ b) : + (a + b) / c = (a / c) + (b / c) := + EReal.right_distrib_of_nonneg h h' + +lemma add_div_of_nonneg_right {a b c : EReal} (h : 0 ≤ c) : + (a + b) / c = a / c + b / c := by + apply right_distrib_of_nonneg_of_ne_top (inv_nonneg_of_nonneg h) (inv_lt_top c).ne + end EReal -- Porting note(https://github.com/leanprover-community/mathlib4/issues/6038): restore diff --git a/Mathlib/Data/Rel.lean b/Mathlib/Data/Rel.lean index 9717d93d288df..ced1fb359df56 100644 --- a/Mathlib/Data/Rel.lean +++ b/Mathlib/Data/Rel.lean @@ -3,9 +3,8 @@ Copyright (c) 2018 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ -import Mathlib.Order.CompleteLattice -import Mathlib.Order.GaloisConnection import Mathlib.Data.Set.Lattice +import Mathlib.Order.CompleteLattice import Mathlib.Tactic.AdaptationNote /-! diff --git a/Mathlib/Data/Seq/Parallel.lean b/Mathlib/Data/Seq/Parallel.lean index ee4bdac0f33db..16ea3f436d55d 100644 --- a/Mathlib/Data/Seq/Parallel.lean +++ b/Mathlib/Data/Seq/Parallel.lean @@ -157,9 +157,7 @@ theorem terminates_parallel {S : WSeq (Computation α)} {c} (h : c ∈ S) [T : T have TT : ∀ l', Terminates (corec parallel.aux1 (l', S.tail)) := by intro apply IH _ _ _ (Or.inr _) T - rw [a] - cases' S with f al - rfl + rw [a, Seq.get?_tail] induction' e : Seq.get? S 0 with o · have D : Seq.destruct S = none := by dsimp [Seq.destruct] diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index efcc4dfa696f6..172c4fc696654 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -230,14 +230,16 @@ theorem get?_tail (s : Seq α) (n) : get? (tail s) n = get? s (n + 1) := rfl /-- Recursion principle for sequences, compare with `List.recOn`. -/ -def recOn {C : Seq α → Sort v} (s : Seq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) : - C s := by +@[cases_eliminator] +def recOn {motive : Seq α → Sort v} (s : Seq α) (nil : motive nil) + (cons : ∀ x s, motive (cons x s)) : + motive s := by cases' H : destruct s with v · rw [destruct_eq_nil H] - apply h1 + apply nil · cases' v with a s' rw [destruct_eq_cons H] - apply h2 + apply cons theorem mem_rec_on {C : Seq α → Prop} {a s} (M : a ∈ s) (h1 : ∀ b s', a = b ∨ C s' → C (cons b s')) : C s := by @@ -251,10 +253,9 @@ theorem mem_rec_on {C : Seq α → Prop} {a s} (M : a ∈ s) rw [TH] apply h1 _ _ (Or.inl rfl) -- Porting note: had to reshuffle `intro` - revert e; apply s.recOn _ fun b s' => _ - · intro e; injection e - · intro b s' e - have h_eq : (cons b s').val (Nat.succ k) = s'.val k := by cases s'; rfl + cases' s with b s' + · injection e + · have h_eq : (cons b s').val (Nat.succ k) = s'.val k := by cases s' using Subtype.recOn; rfl rw [h_eq] at e apply h1 _ _ (Or.inr (IH e)) @@ -331,20 +332,21 @@ theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s match t₁, t₂, e with | _, _, ⟨s, s', rfl, rfl, r⟩ => by suffices head s = head s' ∧ R (tail s) (tail s') from - And.imp id (fun r => ⟨tail s, tail s', by cases s; rfl, by cases s'; rfl, r⟩) this + And.imp id (fun r => ⟨tail s, tail s', by cases s using Subtype.recOn; rfl, + by cases s' using Subtype.recOn; rfl, r⟩) this have := bisim r; revert r this - apply recOn s _ _ <;> apply recOn s' _ _ + cases' s with x s <;> cases' s' with x' s' · intro r _ constructor · rfl · assumption - · intro x s _ this + · intro _ this rw [destruct_nil, destruct_cons] at this exact False.elim this - · intro x s _ this + · intro _ this rw [destruct_nil, destruct_cons] at this exact False.elim this - · intro x s x' s' _ this + · intro _ this rw [destruct_cons, destruct_cons] at this rw [head_cons, head_cons, tail_cons, tail_cons] cases' this with h1 h2 @@ -568,10 +570,10 @@ def toListOrStream (s : Seq α) [Decidable s.Terminates] : List α ⊕ Stream' theorem nil_append (s : Seq α) : append nil s = s := by apply coinduction2; intro s dsimp [append]; rw [corec_eq] - dsimp [append]; apply recOn s _ _ + dsimp [append] + cases' s with x s · trivial - · intro x s - rw [destruct_cons] + · rw [destruct_cons] dsimp exact ⟨rfl, s, rfl, rfl⟩ @@ -718,10 +720,9 @@ theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) : @[simp] theorem append_nil (s : Seq α) : append s nil = s := by apply coinduction2 s; intro s - apply recOn s _ _ + cases' s with x s · trivial - · intro x s - rw [cons_append, destruct_cons, destruct_cons] + · rw [cons_append, destruct_cons, destruct_cons] dsimp exact ⟨rfl, s, rfl, rfl⟩ @@ -732,15 +733,12 @@ theorem append_assoc (s t u : Seq α) : append (append s t) u = append s (append exact match s1, s2, h with | _, _, ⟨s, t, u, rfl, rfl⟩ => by - apply recOn s <;> simp - · apply recOn t <;> simp - · apply recOn u <;> simp - · intro _ u - refine ⟨nil, nil, u, ?_, ?_⟩ <;> simp - · intro _ t - refine ⟨nil, t, u, ?_, ?_⟩ <;> simp - · intro _ s - exact ⟨s, t, u, rfl, rfl⟩ + cases' s with _ s <;> simp + · cases' t with _ t <;> simp + · cases' u with _ u <;> simp + · refine ⟨nil, nil, u, ?_, ?_⟩ <;> simp + · refine ⟨nil, t, u, ?_, ?_⟩ <;> simp + · exact ⟨s, t, u, rfl, rfl⟩ · exact ⟨s, t, u, rfl, rfl⟩ @[simp] @@ -776,12 +774,10 @@ theorem map_append (f : α → β) (s t) : map f (append s t) = append (map f s) exact match s1, s2, h with | _, _, ⟨s, t, rfl, rfl⟩ => by - apply recOn s <;> simp - · apply recOn t <;> simp - · intro _ t - refine ⟨nil, t, ?_, ?_⟩ <;> simp - · intro _ s - exact ⟨s, t, rfl, rfl⟩ + cases' s with _ s <;> simp + · cases' t with _ t <;> simp + · refine ⟨nil, t, ?_, ?_⟩ <;> simp + · exact ⟨s, t, rfl, rfl⟩ @[simp] theorem map_get? (f : α → β) : ∀ s n, get? (map f s) n = (get? s n).map f @@ -835,15 +831,13 @@ theorem join_cons (a : α) (s S) : join (cons (a, s) S) = cons a (append s (join exact match s1, s2, h with | s, _, Or.inl <| Eq.refl s => by - apply recOn s; · trivial - · intro x s - rw [destruct_cons] + cases' s with x s; · trivial + · rw [destruct_cons] exact ⟨rfl, Or.inl rfl⟩ | _, _, Or.inr ⟨a, s, S, rfl, rfl⟩ => by - apply recOn s + cases' s with x s · simp [join_cons_cons, join_cons_nil] - · intro x s - simpa [join_cons_cons, join_cons_nil] using Or.inr ⟨x, s, S, rfl, rfl⟩ + · simpa [join_cons_cons, join_cons_nil] using Or.inr ⟨x, s, S, rfl, rfl⟩ @[simp] theorem join_append (S T : Seq (Seq1 α)) : join (append S T) = append (join S) (join T) := by @@ -854,18 +848,15 @@ theorem join_append (S T : Seq (Seq1 α)) : join (append S T) = append (join S) exact match s1, s2, h with | _, _, ⟨s, S, T, rfl, rfl⟩ => by - apply recOn s <;> simp - · apply recOn S <;> simp - · apply recOn T + cases' s with _ s <;> simp + · cases' S with s S <;> simp + · cases' T with s T · simp - · intro s T - cases' s with a s; simp only [join_cons, destruct_cons, true_and] + · cases' s with a s; simp only [join_cons, destruct_cons, true_and] refine ⟨s, nil, T, ?_, ?_⟩ <;> simp - · intro s S - cases' s with a s + · cases' s with a s simpa using ⟨s, S, T, rfl, rfl⟩ - · intro _ s - exact ⟨s, S, T, rfl, rfl⟩ + · exact ⟨s, S, T, rfl, rfl⟩ · refine ⟨nil, S, T, ?_, ?_⟩ <;> simp @[simp] @@ -920,11 +911,11 @@ theorem of_mem_append {s₁ s₂ : Seq α} {a : α} (h : a ∈ append s₁ s₂) generalize e : append s₁ s₂ = ss; intro h; revert s₁ apply mem_rec_on h _ intro b s' o s₁ - apply s₁.recOn _ fun c t₁ => _ + cases' s₁ with c t₁ · intro m _ apply Or.inr simpa using m - · intro c t₁ m e + · intro m e have this := congr_arg destruct e cases' show a = c ∨ a ∈ append t₁ s₂ by simpa using m with e' m · rw [e'] @@ -1002,7 +993,7 @@ def bind (s : Seq1 α) (f : α → Seq1 β) : Seq1 β := @[simp] theorem join_map_ret (s : Seq α) : Seq.join (Seq.map ret s) = s := by - apply coinduction2 s; intro s; apply recOn s <;> simp [ret] + apply coinduction2 s; intro s; cases s <;> simp [ret] @[simp] theorem bind_ret (f : α → β) : ∀ s, bind s (ret ∘ f) = map f s @@ -1016,7 +1007,7 @@ theorem bind_ret (f : α → β) : ∀ s, bind s (ret ∘ f) = map f s theorem ret_bind (a : α) (f : α → Seq1 β) : bind (ret a) f = f a := by simp only [bind, map, ret.eq_1, map_nil] cases' f a with a s - apply recOn s <;> intros <;> simp + cases s <;> simp @[simp] theorem map_join' (f : α → β) (S) : Seq.map f (Seq.join S) = Seq.join (Seq.map (map f) S) := by @@ -1028,18 +1019,16 @@ theorem map_join' (f : α → β) (S) : Seq.map f (Seq.join S) = Seq.join (Seq.m exact match s1, s2, h with | _, _, ⟨s, S, rfl, rfl⟩ => by - apply recOn s <;> simp - · apply recOn S <;> simp - · intro x S - cases' x with a s + cases' s with _ s <;> simp + · cases' S with x S <;> simp + · cases' x with a s simpa [map] using ⟨_, _, rfl, rfl⟩ - · intro _ s - exact ⟨s, S, rfl, rfl⟩ + · exact ⟨s, S, rfl, rfl⟩ · refine ⟨nil, S, ?_, ?_⟩ <;> simp @[simp] theorem map_join (f : α → β) : ∀ S, map f (join S) = join (map (map f) S) - | ((a, s), S) => by apply recOn s <;> intros <;> simp [map] + | ((a, s), S) => by cases s <;> simp [map] @[simp] theorem join_join (SS : Seq (Seq1 (Seq1 α))) : @@ -1052,17 +1041,14 @@ theorem join_join (SS : Seq (Seq1 (Seq1 α))) : exact match s1, s2, h with | _, _, ⟨s, SS, rfl, rfl⟩ => by - apply recOn s <;> simp - · apply recOn SS <;> simp - · intro S SS - cases' S with s S; cases' s with x s + cases' s with _ s <;> simp + · cases' SS with S SS <;> simp + · cases' S with s S; cases' s with x s simp only [Seq.join_cons, join_append, destruct_cons] - apply recOn s <;> simp + cases' s with x s <;> simp · exact ⟨_, _, rfl, rfl⟩ - · intro x s - refine ⟨Seq.cons x (append s (Seq.join S)), SS, ?_, ?_⟩ <;> simp - · intro _ s - exact ⟨s, SS, rfl, rfl⟩ + · refine ⟨Seq.cons x (append s (Seq.join S)), SS, ?_, ?_⟩ <;> simp + · exact ⟨s, SS, rfl, rfl⟩ · refine ⟨nil, SS, ?_, ?_⟩ <;> simp @[simp] @@ -1080,7 +1066,7 @@ theorem bind_assoc (s : Seq1 α) (f : α → Seq1 β) (g : β → Seq1 γ) : -- give names to variables. induction' s using recOn with x s_1 <;> induction' S using recOn with x_1 s_2 <;> simp · cases' x_1 with x t - apply recOn t <;> intros <;> simp + cases t <;> simp · cases' x_1 with y t; simp instance monad : Monad Seq1 where diff --git a/Mathlib/Data/Set/Accumulate.lean b/Mathlib/Data/Set/Accumulate.lean index ffa9845a9ad40..72726156c5014 100644 --- a/Mathlib/Data/Set/Accumulate.lean +++ b/Mathlib/Data/Set/Accumulate.lean @@ -52,4 +52,16 @@ theorem iUnion_accumulate [Preorder α] : ⋃ x, Accumulate s x = ⋃ x, s x := exact ⟨x', hz⟩ · exact iUnion_mono fun i => subset_accumulate +@[simp] +lemma accumulate_zero_nat (s : ℕ → Set β) : Accumulate s 0 = s 0 := by + simp [accumulate_def] + +open Function in +theorem disjoint_accumulate [Preorder α] (hs : Pairwise (Disjoint on s)) {i j : α} (hij : i < j) : + Disjoint (Accumulate s i) (s j) := by + apply disjoint_left.2 (fun x hx ↦ ?_) + simp only [Accumulate, mem_iUnion, exists_prop] at hx + rcases hx with ⟨k, hk, hx⟩ + exact disjoint_left.1 (hs (hk.trans_lt hij).ne) hx + end Set diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 3ca8906f367ae..edefcf869588c 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -956,10 +956,6 @@ theorem insert_inter_distrib (a : α) (s t : Set α) : insert a (s ∩ t) = inse theorem insert_union_distrib (a : α) (s t : Set α) : insert a (s ∪ t) = insert a s ∪ insert a t := ext fun _ => or_or_distrib_left -theorem insert_inj (ha : a ∉ s) : insert a s = insert b s ↔ a = b := - ⟨fun h => eq_of_not_mem_of_mem_insert (h ▸ mem_insert a s) ha, - congr_arg (fun x => insert x s)⟩ - -- useful in proofs by induction theorem forall_of_forall_insert {P : α → Prop} {a : α} {s : Set α} (H : ∀ x, x ∈ insert a s → P x) (x) (h : x ∈ s) : P x := @@ -1060,6 +1056,8 @@ theorem singleton_subset_iff {a : α} {s : Set α} : {a} ⊆ s ↔ a ∈ s := theorem singleton_subset_singleton : ({a} : Set α) ⊆ {b} ↔ a = b := by simp +@[gcongr] protected alias ⟨_, GCongr.singleton_subset_singleton⟩ := singleton_subset_singleton + theorem set_compr_eq_eq_singleton {a : α} : { b | b = a } = {a} := rfl @@ -1355,6 +1353,11 @@ theorem compl_univ_iff {s : Set α} : sᶜ = univ ↔ s = ∅ := theorem compl_ne_univ : sᶜ ≠ univ ↔ s.Nonempty := compl_univ_iff.not.trans nonempty_iff_ne_empty.symm +lemma inl_compl_union_inr_compl {α β : Type*} {s : Set α} {t : Set β} : + Sum.inl '' sᶜ ∪ Sum.inr '' tᶜ = (Sum.inl '' s ∪ Sum.inr '' t)ᶜ := by + rw [compl_union] + aesop + theorem nonempty_compl : sᶜ.Nonempty ↔ s ≠ univ := (ne_univ_iff_exists_not_mem s).symm @@ -1591,6 +1594,17 @@ theorem insert_diff_self_of_not_mem {a : α} {s : Set α} (h : a ∉ s) : insert ext x simp [and_iff_left_of_imp (ne_of_mem_of_not_mem · h)] +lemma insert_diff_self_of_mem (ha : a ∈ s) : insert a (s \ {a}) = s := by + ext; simp +contextual [or_and_left, em, ha] + +lemma insert_erase_invOn : + InvOn (insert a) (fun s ↦ s \ {a}) {s : Set α | a ∈ s} {s : Set α | a ∉ s} := + ⟨fun _s ha ↦ insert_diff_self_of_mem ha, fun _s ↦ insert_diff_self_of_not_mem⟩ + +theorem insert_inj (ha : a ∉ s) : insert a s = insert b s ↔ a = b := + ⟨fun h => eq_of_not_mem_of_mem_insert (h ▸ mem_insert a s) ha, + congr_arg (fun x => insert x s)⟩ + @[simp] theorem insert_diff_eq_singleton {a : α} {s : Set α} (h : a ∉ s) : insert a s \ s = {a} := by ext @@ -2160,4 +2174,27 @@ end Disjoint @[simp] theorem Prop.compl_singleton (p : Prop) : ({p}ᶜ : Set Prop) = {¬p} := ext fun q ↦ by simpa [@Iff.comm q] using not_iff +namespace Equiv + +/-- Given a predicate `p : α → Prop`, produces an equivalence between + `Set {a : α // p a}` and `{s : Set α // ∀ a ∈ s, p a}`. -/ +protected def setSubtypeComm (p : α → Prop) : + Set {a : α // p a} ≃ {s : Set α // ∀ a ∈ s, p a} where + toFun s := ⟨{a | ∃ h : p a, s ⟨a, h⟩}, fun _ h ↦ h.1⟩ + invFun s := {a | a.val ∈ s.val} + left_inv s := by ext a; exact ⟨fun h ↦ h.2, fun h ↦ ⟨a.property, h⟩⟩ + right_inv s := by ext; exact ⟨fun h ↦ h.2, fun h ↦ ⟨s.property _ h, h⟩⟩ + +@[simp] +protected lemma setSubtypeComm_apply (p : α → Prop) (s : Set {a // p a}) : + (Equiv.setSubtypeComm p) s = ⟨{a | ∃ h : p a, ⟨a, h⟩ ∈ s}, fun _ h ↦ h.1⟩ := + rfl + +@[simp] +protected lemma setSubtypeComm_symm_apply (p : α → Prop) (s : {s // ∀ a ∈ s, p a}) : + (Equiv.setSubtypeComm p).symm s = {a | a.val ∈ s.val} := + rfl + +end Equiv + set_option linter.style.longFile 2300 diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index a0711752d0b0d..138d503bcedad 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -1335,7 +1335,6 @@ theorem piecewise_univ [∀ i : α, Decidable (i ∈ (Set.univ : Set α))] : ext i simp [piecewise] ---@[simp] -- Porting note: simpNF linter complains theorem piecewise_insert_self {j : α} [∀ i, Decidable (i ∈ insert j s)] : (insert j s).piecewise f g j = f j := by simp [piecewise] diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index a36c4d6171444..ac85ff8a86b20 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -974,6 +974,13 @@ theorem range_some_union_none (α : Type*) : range (some : α → Option α) ∪ theorem insert_none_range_some (α : Type*) : insert none (range (some : α → Option α)) = univ := (isCompl_range_some_none α).symm.sup_eq_top +lemma image_of_range_union_range_eq_univ {α β γ γ' δ δ' : Type*} + {h : β → α} {f : γ → β} {f₁ : γ' → α} {f₂ : γ → γ'} {g : δ → β} {g₁ : δ' → α} {g₂ : δ → δ'} + (hf : h ∘ f = f₁ ∘ f₂) (hg : h ∘ g = g₁ ∘ g₂) (hfg : range f ∪ range g = univ) (s : Set β) : + h '' s = f₁ '' (f₂ '' (f ⁻¹' s)) ∪ g₁ '' (g₂ '' (g ⁻¹' s)) := by + rw [← image_comp, ← image_comp, ← hf, ← hg, image_comp, image_comp, image_preimage_eq_inter_range, + image_preimage_eq_inter_range, ← image_union, ← inter_union_distrib_left, hfg, inter_univ] + end Range section Subsingleton diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 0a1edf84d64c4..69708f8597b2d 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Johannes Hölzl, Mario Carneiro import Mathlib.Logic.Pairwise import Mathlib.Order.CompleteBooleanAlgebra import Mathlib.Order.Directed -import Mathlib.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Basic import Mathlib.Tactic.Cases /-! diff --git a/Mathlib/Data/Set/MulAntidiagonal.lean b/Mathlib/Data/Set/MulAntidiagonal.lean index 8bddb878d7cc2..bad4cd3613dbe 100644 --- a/Mathlib/Data/Set/MulAntidiagonal.lean +++ b/Mathlib/Data/Set/MulAntidiagonal.lean @@ -38,7 +38,7 @@ theorem mulAntidiagonal_mono_right (h : t₁ ⊆ t₂) : end Mul --- Porting note: Removed simp attribute, simpnf linter can simplify lhs. Added aux version below +-- The left hand side is not in simp normal form, see variant below. @[to_additive] theorem swap_mem_mulAntidiagonal [CommSemigroup α] {s t : Set α} {a : α} {x : α × α} : x.swap ∈ Set.mulAntidiagonal s t a ↔ x ∈ Set.mulAntidiagonal t s a := by diff --git a/Mathlib/Data/Set/Pairwise/Lattice.lean b/Mathlib/Data/Set/Pairwise/Lattice.lean index 644704f5ff134..018037bee624d 100644 --- a/Mathlib/Data/Set/Pairwise/Lattice.lean +++ b/Mathlib/Data/Set/Pairwise/Lattice.lean @@ -133,6 +133,18 @@ section variable {f : ι → Set α} {s t : Set ι} +lemma Set.pairwiseDisjoint_iff : + s.PairwiseDisjoint f ↔ ∀ ⦃i⦄, i ∈ s → ∀ ⦃j⦄, j ∈ s → (f i ∩ f j).Nonempty → i = j := by + simp [Set.PairwiseDisjoint, Set.Pairwise, Function.onFun, not_imp_comm (a := _ = _), + not_disjoint_iff_nonempty_inter] + +lemma Set.pairwiseDisjoint_pair_insert {s : Set α} {a : α} (ha : a ∉ s) : + s.powerset.PairwiseDisjoint fun t ↦ ({t, insert a t} : Set (Set α)) := by + rw [pairwiseDisjoint_iff] + rintro i hi j hj + have := insert_erase_invOn.2.injOn (not_mem_subset hi ha) (not_mem_subset hj ha) + aesop (add simp [Set.Nonempty, Set.subset_def]) + theorem Set.PairwiseDisjoint.subset_of_biUnion_subset_biUnion (h₀ : (s ∪ t).PairwiseDisjoint f) (h₁ : ∀ i ∈ s, (f i).Nonempty) (h : ⋃ i ∈ s, f i ⊆ ⋃ i ∈ t, f i) : s ⊆ t := by rintro i hi diff --git a/Mathlib/Data/Set/Pointwise/Iterate.lean b/Mathlib/Data/Set/Pointwise/Iterate.lean index 3aa53135e8e06..acffc678dbd7e 100644 --- a/Mathlib/Data/Set/Pointwise/Iterate.lean +++ b/Mathlib/Data/Set/Pointwise/Iterate.lean @@ -32,7 +32,7 @@ theorem smul_eq_self_of_preimage_zpow_eq_self {G : Type*} [CommGroup G] {n : ℤ refine le_antisymm (this hg) ?_ conv_lhs => rw [← smul_inv_smul g s] replace hg : g⁻¹ ^ n ^ j = 1 := by rw [inv_zpow, hg, inv_one] - simpa only [le_eq_subset, set_smul_subset_set_smul_iff] using this hg + simpa only [le_eq_subset, smul_set_subset_smul_set_iff] using this hg rw [(IsFixedPt.preimage_iterate hs j : (zpowGroupHom n)^[j] ⁻¹' s = s).symm] rintro g' hg' - ⟨y, hy, rfl⟩ change (zpowGroupHom n)^[j] (g' * y) ∈ s diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 84372884b6f19..343a74fb58eb3 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -330,20 +330,28 @@ theorem preimage_smul_inv (a : α) (t : Set β) : (fun x ↦ a⁻¹ • x) ⁻¹ preimage_smul (toUnits a)⁻¹ t @[to_additive (attr := simp)] -theorem set_smul_subset_set_smul_iff : a • A ⊆ a • B ↔ A ⊆ B := +theorem smul_set_subset_smul_set_iff : a • A ⊆ a • B ↔ A ⊆ B := image_subset_image_iff <| MulAction.injective _ +@[deprecated (since := "2024-12-28")] +alias set_smul_subset_set_smul_iff := smul_set_subset_smul_set_iff + @[to_additive] -theorem set_smul_subset_iff : a • A ⊆ B ↔ A ⊆ a⁻¹ • B := +theorem smul_set_subset_iff_subset_inv_smul_set : a • A ⊆ B ↔ A ⊆ a⁻¹ • B := image_subset_iff.trans <| iff_of_eq <| congr_arg _ <| preimage_equiv_eq_image_symm _ <| MulAction.toPerm _ +@[deprecated (since := "2024-12-28")] +alias set_smul_subset_iff := smul_set_subset_iff_subset_inv_smul_set + @[to_additive] -theorem subset_set_smul_iff : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := +theorem subset_smul_set_iff : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := Iff.symm <| image_subset_iff.trans <| Iff.symm <| iff_of_eq <| congr_arg _ <| image_equiv_eq_preimage_symm _ <| MulAction.toPerm _ +@[deprecated (since := "2024-12-28")] alias subset_set_smul_iff := subset_smul_set_iff + @[to_additive] theorem smul_set_inter : a • (s ∩ t) = a • s ∩ a • t := image_inter <| MulAction.injective a @@ -515,14 +523,21 @@ theorem preimage_smul_inv₀ (ha : a ≠ 0) (t : Set β) : (fun x ↦ a⁻¹ • preimage_smul (Units.mk0 a ha)⁻¹ t @[simp] -theorem set_smul_subset_set_smul_iff₀ (ha : a ≠ 0) {A B : Set β} : a • A ⊆ a • B ↔ A ⊆ B := - show Units.mk0 a ha • _ ⊆ _ ↔ _ from set_smul_subset_set_smul_iff +theorem smul_set_subset_smul_set_iff₀ (ha : a ≠ 0) {A B : Set β} : a • A ⊆ a • B ↔ A ⊆ B := + show Units.mk0 a ha • _ ⊆ _ ↔ _ from smul_set_subset_smul_set_iff + +@[deprecated (since := "2024-12-28")] +alias set_smul_subset_set_smul_iff₀ := smul_set_subset_smul_set_iff₀ + +theorem smul_set_subset_iff₀ (ha : a ≠ 0) {A B : Set β} : a • A ⊆ B ↔ A ⊆ a⁻¹ • B := + show Units.mk0 a ha • _ ⊆ _ ↔ _ from smul_set_subset_iff_subset_inv_smul_set + +@[deprecated (since := "2024-12-28")] alias set_smul_subset_iff₀ := smul_set_subset_iff₀ -theorem set_smul_subset_iff₀ (ha : a ≠ 0) {A B : Set β} : a • A ⊆ B ↔ A ⊆ a⁻¹ • B := - show Units.mk0 a ha • _ ⊆ _ ↔ _ from set_smul_subset_iff +theorem subset_smul_set_iff₀ (ha : a ≠ 0) {A B : Set β} : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := + show _ ⊆ Units.mk0 a ha • _ ↔ _ from subset_smul_set_iff -theorem subset_set_smul_iff₀ (ha : a ≠ 0) {A B : Set β} : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := - show _ ⊆ Units.mk0 a ha • _ ↔ _ from subset_set_smul_iff +@[deprecated (since := "2024-12-28")] alias subset_set_smul_iff₀ := subset_smul_set_iff₀ theorem smul_set_inter₀ (ha : a ≠ 0) : a • (s ∩ t) = a • s ∩ a • t := show Units.mk0 a ha • _ = _ from smul_set_inter diff --git a/Mathlib/Data/SetLike/Basic.lean b/Mathlib/Data/SetLike/Basic.lean index fb74683ca034a..9f99e745a2ce2 100644 --- a/Mathlib/Data/SetLike/Basic.lean +++ b/Mathlib/Data/SetLike/Basic.lean @@ -181,7 +181,7 @@ theorem coe_mem (x : p) : (x : B) ∈ p := @[aesop 5% apply (rule_sets := [SetLike])] lemma mem_of_subset {s : Set B} (hp : s ⊆ p) {x : B} (hx : x ∈ s) : x ∈ p := hp hx --- Porting note: removed `@[simp]` because `simpNF` linter complained +@[simp] protected theorem eta (x : p) (hx : (x : B) ∈ p) : (⟨x, hx⟩ : p) = x := rfl @[simp] lemma setOf_mem_eq (a : A) : {b | b ∈ a} = a := rfl @@ -214,4 +214,14 @@ theorem exists_of_lt : p < q → ∃ x ∈ q, x ∉ p := theorem lt_iff_le_and_exists : p < q ↔ p ≤ q ∧ ∃ x ∈ q, x ∉ p := by rw [lt_iff_le_not_le, not_le_iff_exists] +/-- membership is inherited from `Set X` -/ +abbrev instSubtypeSet {X} {p : Set X → Prop} : SetLike {s // p s} X where + coe := (↑) + coe_injective' := Subtype.val_injective + +/-- membership is inherited from `S` -/ +abbrev instSubtype {X S} [SetLike S X] {p : S → Prop} : SetLike {s // p s} X where + coe := (↑) + coe_injective' := SetLike.coe_injective.comp Subtype.val_injective + end SetLike diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 48dc6951f8870..dac555b84fad5 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Bryan Gin-ge Chen -/ import Mathlib.Logic.Relation -import Mathlib.Order.GaloisConnection +import Mathlib.Order.CompleteLattice +import Mathlib.Order.GaloisConnection.Defs /-! # Equivalence relations diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index c0909554e4007..7539329317c93 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -336,8 +336,7 @@ instance [Unique ι] [Inhabited α] : Inhabited (IndexedPartition fun _i : ι => index := default mem_index := Set.mem_univ }⟩ --- Porting note: `simpNF` complains about `mem_index` -attribute [simp] some_mem --mem_index +attribute [simp] some_mem variable (hs : IndexedPartition s) diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index 888d19b2ed8d8..e9e037d20b88d 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -51,7 +51,7 @@ instance instDecidableEqSigma [h₁ : DecidableEq α] [h₂ : ∀ a, DecidableEq | _, _, _, _, isFalse n => isFalse fun h ↦ Sigma.noConfusion h fun e₁ _ ↦ n e₁ -- sometimes the built-in injectivity support does not work -@[simp] -- @[nolint simpNF] +@[simp] theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : Sigma.mk a₁ b₁ = ⟨a₂, b₂⟩ ↔ a₁ = a₂ ∧ HEq b₁ b₂ := ⟨fun h ↦ by cases h; simp, diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index b9ca1bae2f936..4607e88ca0d3e 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -16,6 +16,8 @@ This file defines the sign function for types with zero and a decidable less-tha proves some basic theorems about it. -/ +-- Don't generate unnecessary `sizeOf_spec` lemmas which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in /-- The type of signs. -/ inductive SignType | zero @@ -23,12 +25,6 @@ inductive SignType | pos deriving DecidableEq, Inhabited, Fintype --- Porting note: these lemmas are autogenerated by the inductive definition and are not --- in simple form due to the below `x_eq_x` lemmas -attribute [nolint simpNF] SignType.zero.sizeOf_spec -attribute [nolint simpNF] SignType.neg.sizeOf_spec -attribute [nolint simpNF] SignType.pos.sizeOf_spec - namespace SignType instance : Zero SignType := diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index 79accfa97efb3..ea0164bca8d41 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -101,16 +101,12 @@ lemma coe_ne_coe {a b : Subtype p} : (a : α) ≠ b ↔ a ≠ b := coe_injective @[deprecated (since := "2024-04-04")] alias ⟨ne_of_val_ne, _⟩ := coe_ne_coe --- Porting note: it is unclear why the linter doesn't like this. --- If you understand why, please replace this comment with an explanation, or resolve. -@[simp, nolint simpNF] +@[simp] theorem _root_.exists_eq_subtype_mk_iff {a : Subtype p} {b : α} : (∃ h : p b, a = Subtype.mk b h) ↔ ↑a = b := coe_eq_iff.symm --- Porting note: it is unclear why the linter doesn't like this. --- If you understand why, please replace this comment with an explanation, or resolve. -@[simp, nolint simpNF] +@[simp] theorem _root_.exists_subtype_mk_eq_iff {a : Subtype p} {b : α} : (∃ h : p b, Subtype.mk b h = a) ↔ b = a := by simp only [@eq_comm _ b, exists_eq_subtype_mk_iff, @eq_comm _ _ a] diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 5dd68ca13d171..cf572757b919f 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -17,6 +17,8 @@ universe u v w x variable {α : Type u} {α' : Type w} {β : Type v} {β' : Type x} {γ δ : Type*} +lemma not_isLeft_and_isRight {x : α ⊕ β} : ¬(x.isLeft ∧ x.isRight) := by simp + namespace Sum -- Lean has removed the `@[simp]` attribute on these. For now Mathlib adds it back. @@ -267,3 +269,17 @@ def in₂ (c : γ) : α ⊕ (β ⊕ γ) := inr <| inr c end Sum3 + +/-! +### PSum +-/ + +namespace PSum + +variable {α β : Sort*} + +theorem inl_injective : Function.Injective (PSum.inl : α → α ⊕' β) := fun _ _ ↦ inl.inj + +theorem inr_injective : Function.Injective (PSum.inr : β → α ⊕' β) := fun _ _ ↦ inr.inj + +end PSum diff --git a/Mathlib/Data/Sum/Lattice.lean b/Mathlib/Data/Sum/Lattice.lean index 72e63dc5fa589..cf1ae58eec0b1 100644 --- a/Mathlib/Data/Sum/Lattice.lean +++ b/Mathlib/Data/Sum/Lattice.lean @@ -24,22 +24,22 @@ variable [SemilatticeSup α] [SemilatticeSup β] -- The linter significantly hinders readability here. set_option linter.unusedVariables false in instance instSemilatticeSup : SemilatticeSup (α ⊕ₗ β) where - sup x y := match x, y with + sup | inlₗ a₁, inlₗ a₂ => inl (a₁ ⊔ a₂) | inlₗ a₁, inrₗ b₂ => inr b₂ | inrₗ b₁, inlₗ a₂ => inr b₁ | inrₗ b₁, inrₗ b₂ => inr (b₁ ⊔ b₂) - le_sup_left x y := match x, y with + le_sup_left | inlₗ a₁, inlₗ a₂ => inl_le_inl_iff.2 le_sup_left | inlₗ a₁, inrₗ b₂ => inl_le_inr _ _ | inrₗ b₁, inlₗ a₂ => le_rfl | inrₗ b₁, inrₗ b₂ => inr_le_inr_iff.2 le_sup_left - le_sup_right x y := match x, y with + le_sup_right | inlₗ a₁, inlₗ a₂ => inl_le_inl_iff.2 le_sup_right | inlₗ a₁, inrₗ b₂ => le_rfl | inrₗ b₁, inlₗ a₂ => inl_le_inr _ _ | inrₗ b₁, inrₗ b₂ => inr_le_inr_iff.2 le_sup_right - sup_le x y z hxz hyz := match x, y, z, hxz, hyz with + sup_le | inlₗ a₁, inlₗ a₂, inlₗ a₃, Lex.inl h₁₃, Lex.inl h₂₃ => inl_le_inl_iff.2 <| sup_le h₁₃ h₂₃ | inlₗ a₁, inlₗ a₂, inrₗ b₃, Lex.sep _ _, Lex.sep _ _ => Lex.sep _ _ | inlₗ a₁, inrₗ b₂, inrₗ b₃, Lex.sep _ _, Lex.inr h₂₃ => inr_le_inr_iff.2 h₂₃ @@ -57,22 +57,22 @@ variable [SemilatticeInf α] [SemilatticeInf β] -- The linter significantly hinders readability here. set_option linter.unusedVariables false in instance instSemilatticeInf : SemilatticeInf (α ⊕ₗ β) where - inf x y := match x, y with + inf | inlₗ a₁, inlₗ a₂ => inl (a₁ ⊓ a₂) | inlₗ a₁, inrₗ b₂ => inl a₁ | inrₗ b₁, inlₗ a₂ => inl a₂ | inrₗ b₁, inrₗ b₂ => inr (b₁ ⊓ b₂) - inf_le_left x y := match x, y with + inf_le_left | inlₗ a₁, inlₗ a₂ => inl_le_inl_iff.2 inf_le_left | inlₗ a₁, inrₗ b₂ => le_rfl | inrₗ b₁, inlₗ a₂ => inl_le_inr _ _ | inrₗ b₁, inrₗ b₂ => inr_le_inr_iff.2 inf_le_left - inf_le_right x y := match x, y with + inf_le_right | inlₗ a₁, inlₗ a₂ => inl_le_inl_iff.2 inf_le_right | inlₗ a₁, inrₗ b₂ => inl_le_inr _ _ | inrₗ b₁, inlₗ a₂ => le_rfl | inrₗ b₁, inrₗ b₂ => inr_le_inr_iff.2 inf_le_right - le_inf x y z hzx hzy := match x, y, z, hzx, hzy with + le_inf | inlₗ a₁, inlₗ a₂, inlₗ a₃, Lex.inl h₁₃, Lex.inl h₂₃ => inl_le_inl_iff.2 <| le_inf h₁₃ h₂₃ | inlₗ a₁, inlₗ a₂, inrₗ b₃, Lex.inl h₁₃, Lex.sep _ _ => inl_le_inl_iff.2 h₁₃ | inlₗ a₁, inrₗ b₂, inlₗ a₃, Lex.sep _ _, Lex.inl h₂₃ => inl_le_inl_iff.2 h₂₃ diff --git a/Mathlib/Data/Tree/Basic.lean b/Mathlib/Data/Tree/Basic.lean index 75cba8e0b74ae..4c3e05fa8a410 100644 --- a/Mathlib/Data/Tree/Basic.lean +++ b/Mathlib/Data/Tree/Basic.lean @@ -18,10 +18,6 @@ We also specialize for `Tree Unit`, which is a binary tree without any additional data. We provide the notation `a △ b` for making a `Tree Unit` with children `a` and `b`. -## TODO - -Implement a `Traversable` instance for `Tree`. - ## References @@ -45,12 +41,39 @@ variable {α : Type u} instance : Inhabited (Tree α) := ⟨nil⟩ +/-- +Do an action for every node of the tree. +Actions are taken in node -> left subtree -> right subtree recursive order. +This function is the `traverse` function for the `Traversable Tree` instance. +-/ +def traverse {m : Type* → Type*} [Applicative m] {α β} (f : α → m β) : Tree α → m (Tree β) + | .nil => pure nil + | .node a l r => .node <$> f a <*> traverse f l <*> traverse f r + /-- Apply a function to each value in the tree. This is the `map` function for the `Tree` functor. -/ def map {β} (f : α → β) : Tree α → Tree β | nil => nil | node a l r => node (f a) (map f l) (map f r) +theorem id_map (t : Tree α) : t.map id = t := by + induction t with + | nil => rw [map] + | node v l r hl hr => rw [map, hl, hr, id_eq] + +theorem comp_map {β γ : Type*} (f : α → β) (g : β → γ) (t : Tree α) : + t.map (g ∘ f) = (t.map f).map g := by + induction t with + | nil => rw [map, map, map] + | node v l r hl hr => rw [map, map, map, hl, hr, Function.comp_apply] + +theorem traverse_pure (t : Tree α) {m : Type u → Type*} [Applicative m] [LawfulApplicative m] : + t.traverse (pure : α → m α) = pure t := by + induction t with + | nil => rw [traverse] + | node v l r hl hr => + rw [traverse, hl, hr, map_pure, pure_seq, seq_pure, map_pure, map_pure] + /-- The number of internal nodes (i.e. not including leaves) of a binary tree -/ @[simp] def numNodes : Tree α → ℕ diff --git a/Mathlib/Data/Tree/Traversable.lean b/Mathlib/Data/Tree/Traversable.lean new file mode 100644 index 0000000000000..9d1bd920b7724 --- /dev/null +++ b/Mathlib/Data/Tree/Traversable.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2025 Edward van de Meent. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Edward van de Meent +-/ + +import Mathlib.Data.Tree.Basic +import Mathlib.Control.Applicative +import Mathlib.Control.Traversable.Basic + +/-! +# Traversable Binary Tree + +Provides a `Traversable` instance for the `Tree` type. +-/ + +universe u v w + +namespace Tree +section Traverse +variable {α β : Type*} + +instance : Traversable Tree where + map := map + traverse := traverse + +lemma comp_traverse + {F : Type u → Type v} {G : Type v → Type w} [Applicative F] [Applicative G] + [LawfulApplicative G] {β : Type v} {γ : Type u} (f : β → F γ) (g : α → G β) + (t : Tree α) : t.traverse (Functor.Comp.mk ∘ (f <$> ·) ∘ g) = + Functor.Comp.mk ((·.traverse f) <$> (t.traverse g)) := by + induction t with + | nil => rw [traverse, traverse, map_pure, traverse]; rfl + | node v l r hl hr => + rw [traverse, hl, hr, traverse] + simp only [Function.comp_def, Function.comp_apply, Functor.Comp.map_mk, Functor.map_map, + Comp.seq_mk, seq_map_assoc, map_seq] + rfl + +lemma traverse_eq_map_id (f : α → β) (t : Tree α) : + t.traverse ((pure : β → Id β) ∘ f) = t.map f := by + rw [← Id.pure_eq (t.map f)] + induction t with + | nil => rw [traverse, map] + | node v l r hl hr => + rw [traverse, map, hl, hr, Function.comp_apply, map_pure, pure_seq, map_pure, pure_seq, + map_pure] + +lemma naturality {F G : Type u → Type*} [Applicative F] [Applicative G] [LawfulApplicative F] + [LawfulApplicative G] (η : ApplicativeTransformation F G) {β : Type u} (f : α → F β) + (t : Tree α) : η (t.traverse f) = t.traverse (η.app β ∘ f : α → G β) := by + induction t with + | nil => rw [traverse, traverse, η.preserves_pure] + | node v l r hl hr => + rw [traverse, traverse, η.preserves_seq, η.preserves_seq, η.preserves_map, hl, hr, + Function.comp_apply] + +instance : LawfulTraversable Tree where + map_const := rfl + id_map := id_map + comp_map := comp_map + id_traverse t := traverse_pure t + comp_traverse := comp_traverse + traverse_eq_map_id := traverse_eq_map_id + naturality η := naturality η + +end Traverse + +end Tree diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index fec363551b146..ef6d34d2aedac 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -705,6 +705,9 @@ theorem val_injective (n : ℕ) [NeZero n] : Function.Injective (val : ZMod n theorem val_one_eq_one_mod (n : ℕ) : (1 : ZMod n).val = 1 % n := by rw [← Nat.cast_one, val_natCast] +theorem val_two_eq_two_mod (n : ℕ) : (2 : ZMod n).val = 2 % n := by + rw [← Nat.cast_two, val_natCast] + theorem val_one (n : ℕ) [Fact (1 < n)] : (1 : ZMod n).val = 1 := by rw [val_one_eq_one_mod] exact Nat.mod_eq_of_lt Fact.out diff --git a/Mathlib/Data/ZMod/Coprime.lean b/Mathlib/Data/ZMod/Coprime.lean index 91030c12cfdc1..2923900a3f990 100644 --- a/Mathlib/Data/ZMod/Coprime.lean +++ b/Mathlib/Data/ZMod/Coprime.lean @@ -3,8 +3,9 @@ Copyright (c) 2022 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ +import Mathlib.Algebra.EuclideanDomain.Int +import Mathlib.Data.Nat.Prime.Int import Mathlib.Data.ZMod.Basic -import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.PrincipalIdealDomain /-! @@ -22,7 +23,7 @@ namespace ZMod `gcd a p ≠ 1`. -/ theorem eq_zero_iff_gcd_ne_one {a : ℤ} {p : ℕ} [pp : Fact p.Prime] : (a : ZMod p) = 0 ↔ a.gcd p ≠ 1 := by - rw [Ne, Int.gcd_comm, Int.gcd_eq_one_iff_coprime, + rw [Ne, Int.gcd_comm, ← Int.isCoprime_iff_gcd_eq_one, (Nat.prime_iff_prime_int.1 pp.1).coprime_iff_not_dvd, Classical.not_not, intCast_zmod_eq_zero_iff_dvd] diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/QuotientGroup.lean similarity index 80% rename from Mathlib/Data/ZMod/Quotient.lean rename to Mathlib/Data/ZMod/QuotientGroup.lean index 8f05b3b2992d5..10496556935c0 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/QuotientGroup.lean @@ -3,32 +3,28 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import Mathlib.RingTheory.Ideal.Quotient.Operations +import Mathlib.Data.ZMod.Basic import Mathlib.RingTheory.Int.Basic -import Mathlib.RingTheory.ZMod /-! # `ZMod n` and quotient groups / rings -This file relates `ZMod n` to the quotient group -`ℤ / AddSubgroup.zmultiples (n : ℤ)` and to the quotient ring -`ℤ ⧸ Ideal.span {(n : ℤ)}`. +This file relates `ZMod n` to the quotient group `ℤ / AddSubgroup.zmultiples (n : ℤ)`. ## Main definitions - `ZMod.quotientZMultiplesNatEquivZMod` and `ZMod.quotientZMultiplesEquivZMod`: `ZMod n` is the group quotient of `ℤ` by `n ℤ := AddSubgroup.zmultiples (n)`, (where `n : ℕ` and `n : ℤ` respectively) - - `ZMod.quotient_span_nat_equiv_zmod` and `ZMod.quotientSpanEquivZMod `: - `ZMod n` is the ring quotient of `ℤ` by `n ℤ : Ideal.span {n}` - (where `n : ℕ` and `n : ℤ` respectively) - `ZMod.lift n f` is the map from `ZMod n` induced by `f : ℤ →+ A` that maps `n` to `0`. ## Tags -zmod, quotient group, quotient ring, ideal quotient +zmod, quotient group -/ +assert_not_exists TwoSidedIdeal + open QuotientAddGroup Set ZMod variable (n : ℕ) {A R : Type*} [AddGroup A] [Ring R] @@ -44,38 +40,12 @@ def quotientZMultiplesNatEquivZMod : ℤ ⧸ AddSubgroup.zmultiples (n : ℤ) def quotientZMultiplesEquivZMod (a : ℤ) : ℤ ⧸ AddSubgroup.zmultiples a ≃+ ZMod a.natAbs := (quotientAddEquivOfEq (zmultiples_natAbs a)).symm.trans (quotientZMultiplesNatEquivZMod a.natAbs) -/-- `ℤ` modulo the ideal generated by `n : ℕ` is `ZMod n`. -/ -def quotientSpanNatEquivZMod : ℤ ⧸ Ideal.span {(n : ℤ)} ≃+* ZMod n := - (Ideal.quotEquivOfEq (ZMod.ker_intCastRingHom _)).symm.trans <| - RingHom.quotientKerEquivOfRightInverse <| - show Function.RightInverse ZMod.cast (Int.castRingHom (ZMod n)) from intCast_zmod_cast - -/-- `ℤ` modulo the ideal generated by `a : ℤ` is `ZMod a.natAbs`. -/ -def quotientSpanEquivZMod (a : ℤ) : ℤ ⧸ Ideal.span ({a} : Set ℤ) ≃+* ZMod a.natAbs := - (Ideal.quotEquivOfEq (span_natAbs a)).symm.trans (quotientSpanNatEquivZMod a.natAbs) - @[simp] lemma index_zmultiples (a : ℤ) : (AddSubgroup.zmultiples a).index = a.natAbs := by rw [AddSubgroup.index, Nat.card_congr (quotientZMultiplesEquivZMod a).toEquiv, Nat.card_zmod] end Int -noncomputable section ChineseRemainder -open Ideal - -open scoped Function in -- required for scoped `on` notation -/-- The **Chinese remainder theorem**, elementary version for `ZMod`. See also -`Mathlib.Data.ZMod.Basic` for versions involving only two numbers. -/ -def ZMod.prodEquivPi {ι : Type*} [Fintype ι] (a : ι → ℕ) - (coprime : Pairwise (Nat.Coprime on a)) : ZMod (∏ i, a i) ≃+* Π i, ZMod (a i) := - have : Pairwise fun i j => IsCoprime (span {(a i : ℤ)}) (span {(a j : ℤ)}) := - fun _i _j h ↦ (isCoprime_span_singleton_iff _ _).mpr ((coprime h).cast (R := ℤ)) - Int.quotientSpanNatEquivZMod _ |>.symm.trans <| - quotEquivOfEq (iInf_span_singleton_natCast (R := ℤ) coprime) |>.symm.trans <| - quotientInfRingEquivPiQuotient _ this |>.trans <| - RingEquiv.piCongrRight fun i ↦ Int.quotientSpanNatEquivZMod (a i) - -end ChineseRemainder namespace AddAction diff --git a/Mathlib/Data/ZMod/QuotientRing.lean b/Mathlib/Data/ZMod/QuotientRing.lean new file mode 100644 index 0000000000000..131a259756d32 --- /dev/null +++ b/Mathlib/Data/ZMod/QuotientRing.lean @@ -0,0 +1,59 @@ +/- +Copyright (c) 2021 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import Mathlib.RingTheory.Ideal.Quotient.Operations +import Mathlib.RingTheory.Int.Basic +import Mathlib.RingTheory.ZMod + +/-! +# `ZMod n` and quotient groups / rings + +This file relates `ZMod n` to the quotient ring `ℤ ⧸ Ideal.span {(n : ℤ)}`. + +## Main definitions + + - `ZMod.quotient_span_nat_equiv_zmod` and `ZMod.quotientSpanEquivZMod `: + `ZMod n` is the ring quotient of `ℤ` by `n ℤ : Ideal.span {n}` + (where `n : ℕ` and `n : ℤ` respectively) + +## Tags + +zmod, quotient ring, ideal quotient +-/ + +open QuotientAddGroup Set ZMod + +variable (n : ℕ) {A R : Type*} [AddGroup A] [Ring R] + +namespace Int + +/-- `ℤ` modulo the ideal generated by `n : ℕ` is `ZMod n`. -/ +def quotientSpanNatEquivZMod : ℤ ⧸ Ideal.span {(n : ℤ)} ≃+* ZMod n := + (Ideal.quotEquivOfEq (ZMod.ker_intCastRingHom _)).symm.trans <| + RingHom.quotientKerEquivOfRightInverse <| + show Function.RightInverse ZMod.cast (Int.castRingHom (ZMod n)) from intCast_zmod_cast + +/-- `ℤ` modulo the ideal generated by `a : ℤ` is `ZMod a.natAbs`. -/ +def quotientSpanEquivZMod (a : ℤ) : ℤ ⧸ Ideal.span ({a} : Set ℤ) ≃+* ZMod a.natAbs := + (Ideal.quotEquivOfEq (span_natAbs a)).symm.trans (quotientSpanNatEquivZMod a.natAbs) + +end Int + +noncomputable section ChineseRemainder +open Ideal + +open scoped Function in -- required for scoped `on` notation +/-- The **Chinese remainder theorem**, elementary version for `ZMod`. See also +`Mathlib.Data.ZMod.Basic` for versions involving only two numbers. -/ +def ZMod.prodEquivPi {ι : Type*} [Fintype ι] (a : ι → ℕ) + (coprime : Pairwise (Nat.Coprime on a)) : ZMod (∏ i, a i) ≃+* Π i, ZMod (a i) := + have : Pairwise fun i j => IsCoprime (span {(a i : ℤ)}) (span {(a j : ℤ)}) := + fun _i _j h ↦ (isCoprime_span_singleton_iff _ _).mpr ((coprime h).cast (R := ℤ)) + Int.quotientSpanNatEquivZMod _ |>.symm.trans <| + quotEquivOfEq (iInf_span_singleton_natCast (R := ℤ) coprime) |>.symm.trans <| + quotientInfRingEquivPiQuotient _ this |>.trans <| + RingEquiv.piCongrRight fun i ↦ Int.quotientSpanNatEquivZMod (a i) + +end ChineseRemainder diff --git a/Mathlib/Dynamics/Ergodic/Ergodic.lean b/Mathlib/Dynamics/Ergodic/Ergodic.lean index cc9abb51e7198..29c03db776487 100644 --- a/Mathlib/Dynamics/Ergodic/Ergodic.lean +++ b/Mathlib/Dynamics/Ergodic/Ergodic.lean @@ -42,13 +42,11 @@ structure PreErgodic (f : α → α) (μ : Measure α := by volume_tac) : Prop w /-- A map `f : α → α` is said to be ergodic with respect to a measure `μ` if it is measure preserving and pre-ergodic. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] structure Ergodic (f : α → α) (μ : Measure α := by volume_tac) extends MeasurePreserving f μ μ, PreErgodic f μ : Prop /-- A map `f : α → α` is said to be quasi ergodic with respect to a measure `μ` if it is quasi measure preserving and pre-ergodic. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] structure QuasiErgodic (f : α → α) (μ : Measure α := by volume_tac) extends QuasiMeasurePreserving f μ μ, PreErgodic f μ : Prop diff --git a/Mathlib/FieldTheory/AlgebraicClosure.lean b/Mathlib/FieldTheory/AlgebraicClosure.lean index 056881d31cde2..559bd7ea774b1 100644 --- a/Mathlib/FieldTheory/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/AlgebraicClosure.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Jiedong Jiang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu, Jiedong Jiang -/ -import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.Normal.Closure import Mathlib.FieldTheory.IsAlgClosed.Basic import Mathlib.FieldTheory.IntermediateField.Algebraic diff --git a/Mathlib/FieldTheory/CardinalEmb.lean b/Mathlib/FieldTheory/CardinalEmb.lean index 0bcfea73f5a13..3f51c66e061a4 100644 --- a/Mathlib/FieldTheory/CardinalEmb.lean +++ b/Mathlib/FieldTheory/CardinalEmb.lean @@ -122,6 +122,7 @@ def leastExt : ι → ι := rw [adjoin_basis_eq_top, ← eq_top_iff] at this apply_fun Module.rank F at this refine ne_of_lt ?_ this + let _ : AddCommMonoid (⊤ : IntermediateField F E) := inferInstance conv_rhs => rw [topEquiv.toLinearEquiv.rank_eq] have := mk_Iio_ord_toType i obtain eq | lt := rank_inf.out.eq_or_lt diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 134c9b94c12f6..9f98c44c3ceac 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -249,6 +249,10 @@ theorem card' : ∃ (p : ℕ) (n : ℕ+), Nat.Prime p ∧ Fintype.card K = p ^ ( let ⟨p, hc⟩ := CharP.exists K ⟨p, @FiniteField.card K _ _ p hc⟩ +lemma isPrimePow_card : IsPrimePow (Fintype.card K) := by + obtain ⟨p, n, hp, hn⟩ := card' K + exact ⟨p, n, Nat.prime_iff.mp hp, n.prop, hn.symm⟩ + -- Porting note: this was a `simp` lemma with a 5 lines proof. theorem cast_card_eq_zero : (q : K) = 0 := by simp diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index d93f2b2225060..22d9e4abf7175 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -71,13 +71,12 @@ theorem eval_indicator_apply_eq_one (a : σ → K) : eval a (indicator a) = 1 := theorem degrees_indicator (c : σ → K) : degrees (indicator c) ≤ ∑ s : σ, (Fintype.card K - 1) • {s} := by rw [indicator] - refine le_trans (degrees_prod _ _) (Finset.sum_le_sum fun s _ => ?_) classical - refine le_trans (degrees_sub _ _) ?_ - rw [degrees_one, ← bot_eq_zero, bot_sup_eq] - refine le_trans (degrees_pow _ _) (nsmul_le_nsmul_right ?_ _) - refine le_trans (degrees_sub _ _) ?_ - rw [degrees_C, ← bot_eq_zero, sup_bot_eq] + refine degrees_prod_le.trans <| Finset.sum_le_sum fun s _ ↦ degrees_sub_le.trans ?_ + rw [degrees_one, Multiset.zero_union] + refine le_trans degrees_pow_le (nsmul_le_nsmul_right ?_ _) + refine degrees_sub_le.trans ?_ + rw [degrees_C, Multiset.union_zero] exact degrees_X' _ theorem indicator_mem_restrictDegree (c : σ → K) : diff --git a/Mathlib/FieldTheory/Fixed.lean b/Mathlib/FieldTheory/Fixed.lean index 973277f311115..6c4eee03e3080 100644 --- a/Mathlib/FieldTheory/Fixed.lean +++ b/Mathlib/FieldTheory/Fixed.lean @@ -6,8 +6,9 @@ Authors: Kenny Lau import Mathlib.Algebra.Polynomial.GroupRingAction import Mathlib.Algebra.Ring.Action.Field import Mathlib.Algebra.Ring.Action.Invariant -import Mathlib.FieldTheory.Normal +import Mathlib.FieldTheory.Normal.Defs import Mathlib.FieldTheory.Separable +import Mathlib.LinearAlgebra.Dual import Mathlib.LinearAlgebra.FreeModule.Finite.Matrix /-! diff --git a/Mathlib/FieldTheory/Galois/Basic.lean b/Mathlib/FieldTheory/Galois/Basic.lean index e5dd455a5a331..212915fa47195 100644 --- a/Mathlib/FieldTheory/Galois/Basic.lean +++ b/Mathlib/FieldTheory/Galois/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Patrick Lutz, Yongle Hu, Jingting Wang -/ import Mathlib.FieldTheory.Fixed -import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.Normal.Closure import Mathlib.FieldTheory.PrimitiveElement import Mathlib.GroupTheory.GroupAction.FixingSubgroup diff --git a/Mathlib/FieldTheory/Galois/GaloisClosure.lean b/Mathlib/FieldTheory/Galois/GaloisClosure.lean index 3a5effad525c8..bb80f0aad565f 100644 --- a/Mathlib/FieldTheory/Galois/GaloisClosure.lean +++ b/Mathlib/FieldTheory/Galois/GaloisClosure.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nailin Guan, Yuyang Zhao -/ -import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.Normal.Closure import Mathlib.FieldTheory.SeparableClosure /-! diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index a80409efa5208..7f59bd51abc1d 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -460,16 +460,10 @@ between `E` and `E.map e` -/ def intermediateFieldMap (e : L ≃ₐ[K] L') (E : IntermediateField K L) : E ≃ₐ[K] E.map e.toAlgHom := e.subalgebraMap E.toSubalgebra -/- We manually add these two simp lemmas because `@[simps]` before `intermediate_field_map` - led to a timeout. -/ --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem intermediateFieldMap_apply_coe (e : L ≃ₐ[K] L') (E : IntermediateField K L) (a : E) : ↑(intermediateFieldMap e E a) = e a := rfl --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem intermediateFieldMap_symm_apply_coe (e : L ≃ₐ[K] L') (E : IntermediateField K L) (a : E.map e.toAlgHom) : ↑((intermediateFieldMap e E).symm a) = e.symm a := rfl diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index 156df9e238ca5..6c1c121b87fb7 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.FieldTheory.Normal +import Mathlib.FieldTheory.Extension +import Mathlib.FieldTheory.Normal.Defs import Mathlib.FieldTheory.Perfect import Mathlib.RingTheory.Localization.Integral diff --git a/Mathlib/FieldTheory/Isaacs.lean b/Mathlib/FieldTheory/Isaacs.lean index 2219c66b3c027..a5f3fdb5637b1 100644 --- a/Mathlib/FieldTheory/Isaacs.lean +++ b/Mathlib/FieldTheory/Isaacs.lean @@ -3,6 +3,7 @@ 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.FieldTheory.Normal.Basic import Mathlib.FieldTheory.PrimitiveElement import Mathlib.GroupTheory.CosetCover diff --git a/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean b/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean index fe5e7910abb47..e6491b2c7bb02 100644 --- a/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean +++ b/Mathlib/FieldTheory/Minpoly/IsConjRoot.lean @@ -3,9 +3,10 @@ Copyright (c) 2024 Jiedong Jiang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jiedong Jiang -/ -import Mathlib.FieldTheory.Minpoly.Basic +import Mathlib.FieldTheory.Extension import Mathlib.FieldTheory.IntermediateField.Adjoin.Basic -import Mathlib.FieldTheory.Normal +import Mathlib.FieldTheory.Minpoly.Basic +import Mathlib.FieldTheory.Normal.Defs /-! # Conjugate roots diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal/Basic.lean similarity index 59% rename from Mathlib/FieldTheory/Normal.lean rename to Mathlib/FieldTheory/Normal/Basic.lean index c209741701297..2ce3b4b8935ec 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal/Basic.lean @@ -4,19 +4,17 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Thomas Browning, Patrick Lutz -/ import Mathlib.FieldTheory.Extension -import Mathlib.FieldTheory.SplittingField.Construction +import Mathlib.FieldTheory.Normal.Defs import Mathlib.GroupTheory.Solvable +import Mathlib.FieldTheory.SplittingField.Construction /-! # Normal field extensions -In this file we define normal field extensions and prove that for a finite extension, being normal +In this file we prove that for a finite extension, being normal is the same as being a splitting field (`Normal.of_isSplittingField` and `Normal.exists_isSplittingField`). -## Main Definitions - -- `Normal F K` where `K` is a field extension of `F`. -/ @@ -26,34 +24,6 @@ open Polynomial IsScalarTower variable (F K : Type*) [Field F] [Field K] [Algebra F K] -/-- Typeclass for normal field extension: `K` is a normal extension of `F` iff the minimal -polynomial of every element `x` in `K` splits in `K`, i.e. every conjugate of `x` is in `K`. -/ -@[stacks 09HM] -class Normal extends Algebra.IsAlgebraic F K : Prop where - splits' (x : K) : Splits (algebraMap F K) (minpoly F x) - -variable {F K} - -theorem Normal.isIntegral (_ : Normal F K) (x : K) : IsIntegral F x := - Algebra.IsIntegral.isIntegral x - -theorem Normal.splits (_ : Normal F K) (x : K) : Splits (algebraMap F K) (minpoly F x) := - Normal.splits' x - -theorem normal_iff : Normal F K ↔ ∀ x : K, IsIntegral F x ∧ Splits (algebraMap F K) (minpoly F x) := - ⟨fun h x => ⟨h.isIntegral x, h.splits x⟩, fun h => - { isAlgebraic := fun x => (h x).1.isAlgebraic - splits' := fun x => (h x).2 }⟩ - -theorem Normal.out : Normal F K → ∀ x : K, IsIntegral F x ∧ Splits (algebraMap F K) (minpoly F x) := - normal_iff.1 - -variable (F K) - -instance normal_self : Normal F F where - isAlgebraic := fun _ => isIntegral_algebraMap.isAlgebraic - splits' := fun x => (minpoly.eq_X_sub_C' x).symm ▸ splits_X_sub_C _ - theorem Normal.exists_isSplittingField [h : Normal F K] [FiniteDimensional F K] : ∃ p : F[X], IsSplittingField F K p := by classical @@ -76,32 +46,7 @@ section NormalTower variable (E : Type*) [Field E] [Algebra F E] [Algebra K E] [IsScalarTower F K E] -@[stacks 09HN] -theorem Normal.tower_top_of_normal [h : Normal F E] : Normal K E := - normal_iff.2 fun x => by - cases' h.out x with hx hhx - rw [algebraMap_eq F K E] at hhx - exact - ⟨hx.tower_top, - Polynomial.splits_of_splits_of_dvd (algebraMap K E) - (Polynomial.map_ne_zero (minpoly.ne_zero hx)) - ((Polynomial.splits_map_iff (algebraMap F K) (algebraMap K E)).mpr hhx) - (minpoly.dvd_map_of_isScalarTower F K x)⟩ - -theorem AlgHom.normal_bijective [h : Normal F E] (ϕ : E →ₐ[F] K) : Function.Bijective ϕ := - h.toIsAlgebraic.bijective_of_isScalarTower' ϕ - variable {E F} -variable {E' : Type*} [Field E'] [Algebra F E'] - -theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' := by - rw [normal_iff] at h ⊢ - intro x; specialize h (f.symm x) - rw [← f.apply_symm_apply x, minpoly.algEquiv_eq, ← f.toAlgHom.comp_algebraMap] - exact ⟨h.1.map f, splits_comp_of_splits _ _ h.2⟩ - -theorem AlgEquiv.transfer_normal (f : E ≃ₐ[F] E') : Normal F E ↔ Normal F E' := - ⟨fun _ ↦ Normal.of_algEquiv f, fun _ ↦ Normal.of_algEquiv f.symm⟩ open IntermediateField @@ -200,14 +145,6 @@ instance normal_inf Normal F (E ⊓ E' : IntermediateField F K) := iInf_bool_eq (f := Bool.rec E' E) ▸ normal_iInf (h := by rintro (_|_) <;> infer_instance) -variable {F K} -variable {L : Type*} [Field L] [Algebra F L] [Algebra K L] [IsScalarTower F K L] - -@[simp] -theorem restrictScalars_normal {E : IntermediateField K L} : - Normal F (E.restrictScalars F) ↔ Normal F E := - Iff.rfl - end IntermediateField variable {F} {K} @@ -220,50 +157,6 @@ section Restrict variable (E : Type*) [Field E] [Algebra F E] [Algebra E K₁] [Algebra E K₂] [Algebra E K₃] [IsScalarTower F E K₁] [IsScalarTower F E K₂] [IsScalarTower F E K₃] -/-- Restrict algebra homomorphism to image of normal subfield -/ -def AlgHom.restrictNormalAux [h : Normal F E] : - (toAlgHom F E K₁).range →ₐ[F] (toAlgHom F E K₂).range where - toFun x := - ⟨ϕ x, by - suffices (toAlgHom F E K₁).range.map ϕ ≤ _ by exact this ⟨x, Subtype.mem x, rfl⟩ - rintro x ⟨y, ⟨z, hy⟩, hx⟩ - rw [← hx, ← hy] - apply minpoly.mem_range_of_degree_eq_one E - refine - Or.resolve_left (h.splits z).def (minpoly.ne_zero (h.isIntegral z)) (minpoly.irreducible ?_) - (minpoly.dvd E _ (by simp [aeval_algHom_apply])) - simp only [AlgHom.toRingHom_eq_coe, AlgHom.coe_toRingHom] - suffices IsIntegral F _ by exact this.tower_top - exact ((h.isIntegral z).map <| toAlgHom F E K₁).map ϕ⟩ - map_zero' := Subtype.ext (map_zero _) - map_one' := Subtype.ext (map_one _) - map_add' x y := Subtype.ext <| by simp - map_mul' x y := Subtype.ext <| by simp - commutes' x := Subtype.ext (ϕ.commutes x) - -/-- Restrict algebra homomorphism to normal subfield. -/ -@[stacks 0BME "Part 1"] -def AlgHom.restrictNormal [Normal F E] : E →ₐ[F] E := - ((AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K₂)).symm.toAlgHom.comp - (ϕ.restrictNormalAux E)).comp - (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K₁)).toAlgHom - -/-- Restrict algebra homomorphism to normal subfield (`AlgEquiv` version) -/ -def AlgHom.restrictNormal' [Normal F E] : E ≃ₐ[F] E := - AlgEquiv.ofBijective (AlgHom.restrictNormal ϕ E) (AlgHom.normal_bijective F E E _) - -@[simp] -theorem AlgHom.restrictNormal_commutes [Normal F E] (x : E) : - algebraMap E K₂ (ϕ.restrictNormal E x) = ϕ (algebraMap E K₁ x) := - Subtype.ext_iff.mp - (AlgEquiv.apply_symm_apply (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K₂)) - (ϕ.restrictNormalAux E ⟨IsScalarTower.toAlgHom F E K₁ x, x, rfl⟩)) - -theorem AlgHom.restrictNormal_comp [Normal F E] : - (ψ.restrictNormal E).comp (ϕ.restrictNormal E) = (ψ.comp ϕ).restrictNormal E := - AlgHom.ext fun _ => - (algebraMap E K₃).injective (by simp only [AlgHom.comp_apply, AlgHom.restrictNormal_commutes]) - theorem AlgHom.fieldRange_of_normal {E : IntermediateField F K} [Normal F E] (f : E →ₐ[F] K) : f.fieldRange = E := by let g := f.restrictNormal' E @@ -271,47 +164,6 @@ theorem AlgHom.fieldRange_of_normal {E : IntermediateField F K} [Normal F E] ← AlgHom.map_fieldRange, AlgEquiv.fieldRange_eq_top g, ← AlgHom.fieldRange_eq_map, IntermediateField.fieldRange_val] -/-- Restrict algebra isomorphism to a normal subfield -/ -def AlgEquiv.restrictNormal [Normal F E] : E ≃ₐ[F] E := - AlgHom.restrictNormal' χ.toAlgHom E - -@[simp] -theorem AlgEquiv.restrictNormal_commutes [Normal F E] (x : E) : - algebraMap E K₂ (χ.restrictNormal E x) = χ (algebraMap E K₁ x) := - χ.toAlgHom.restrictNormal_commutes E x - -theorem AlgEquiv.restrictNormal_trans [Normal F E] : - (χ.trans ω).restrictNormal E = (χ.restrictNormal E).trans (ω.restrictNormal E) := - AlgEquiv.ext fun _ => - (algebraMap E K₃).injective - (by simp only [AlgEquiv.trans_apply, AlgEquiv.restrictNormal_commutes]) - -/-- Restriction to a normal subfield as a group homomorphism -/ -def AlgEquiv.restrictNormalHom [Normal F E] : (K₁ ≃ₐ[F] K₁) →* E ≃ₐ[F] E := - MonoidHom.mk' (fun χ => χ.restrictNormal E) fun ω χ => χ.restrictNormal_trans ω E - -lemma AlgEquiv.restrictNormalHom_apply (L : IntermediateField F K₁) [Normal F L] - (σ : (K₁ ≃ₐ[F] K₁)) (x : L) : restrictNormalHom L σ x = σ x := - AlgEquiv.restrictNormal_commutes σ L x - -variable (F K₁) - -/-- If `K₁/E/F` is a tower of fields with `E/F` normal then `AlgHom.restrictNormal'` is an - equivalence. -/ -@[simps, stacks 0BR4] -def Normal.algHomEquivAut [Normal F E] : (E →ₐ[F] K₁) ≃ E ≃ₐ[F] E where - toFun σ := AlgHom.restrictNormal' σ E - invFun σ := (IsScalarTower.toAlgHom F E K₁).comp σ.toAlgHom - left_inv σ := by - ext - simp [AlgHom.restrictNormal'] - right_inv σ := by - ext - simp only [AlgHom.restrictNormal', AlgEquiv.toAlgHom_eq_coe, AlgEquiv.coe_ofBijective] - apply NoZeroSMulDivisors.algebraMap_injective E K₁ - rw [AlgHom.restrictNormal_commutes] - simp - end Restrict section lift @@ -373,49 +225,6 @@ theorem AlgEquiv.restrictNormalHom_surjective [Normal F K₁] [Normal F E] : Function.Surjective (AlgEquiv.restrictNormalHom K₁ : (E ≃ₐ[F] E) → K₁ ≃ₐ[F] K₁) := fun χ => ⟨χ.liftNormal E, χ.restrict_liftNormal E⟩ -/-- The group homomorphism given by restricting an algebra isomorphism to itself -is the identity map. -/ -@[simp] -theorem AlgEquiv.restrictNormalHom_id (F K : Type*) - [Field F] [Field K] [Algebra F K] [Normal F K] : - AlgEquiv.restrictNormalHom K = MonoidHom.id (K ≃ₐ[F] K) := by - ext f x - dsimp only [restrictNormalHom, MonoidHom.mk'_apply, MonoidHom.id_apply] - apply (algebraMap K K).injective - rw [AlgEquiv.restrictNormal_commutes] - simp only [Algebra.id.map_eq_id, RingHom.id_apply] - -namespace IsScalarTower - -/-- In a scalar tower `K₃/K₂/K₁/F` with `K₁` and `K₂` are normal over `F`, the group homomorphism -given by the restriction of algebra isomorphisms of `K₃` to `K₁` is equal to the composition of -the group homomorphism given by the restricting an algebra isomorphism of `K₃` to `K₂` and -the group homomorphism given by the restricting an algebra isomorphism of `K₂` to `K₁` -/ -theorem AlgEquiv.restrictNormalHom_comp (F K₁ K₂ K₃ : Type*) - [Field F] [Field K₁] [Field K₂] [Field K₃] - [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] - [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] - [Normal F K₁] [Normal F K₂] : - AlgEquiv.restrictNormalHom K₁ = - (AlgEquiv.restrictNormalHom K₁).comp - (AlgEquiv.restrictNormalHom (F := F) (K₁ := K₃) K₂) := by - ext f x - apply (algebraMap K₁ K₃).injective - rw [IsScalarTower.algebraMap_eq K₁ K₂ K₃] - simp only [AlgEquiv.restrictNormalHom, MonoidHom.mk'_apply, RingHom.coe_comp, Function.comp_apply, - ← algebraMap_apply, AlgEquiv.restrictNormal_commutes, MonoidHom.coe_comp] - -theorem AlgEquiv.restrictNormalHom_comp_apply (K₁ K₂ : Type*) {F K₃ : Type*} - [Field F] [Field K₁] [Field K₂] [Field K₃] - [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] - [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] - [Normal F K₁] [Normal F K₂] (f : K₃ ≃ₐ[F] K₃) : - AlgEquiv.restrictNormalHom K₁ f = - (AlgEquiv.restrictNormalHom K₁) (AlgEquiv.restrictNormalHom K₂ f) := by - rw [IsScalarTower.AlgEquiv.restrictNormalHom_comp F K₁ K₂ K₃, MonoidHom.comp_apply] - -end IsScalarTower - open IntermediateField in theorem Normal.minpoly_eq_iff_mem_orbit [h : Normal F E] {x y : E} : minpoly F x = minpoly F y ↔ x ∈ MulAction.orbit (E ≃ₐ[F] E) y := by diff --git a/Mathlib/FieldTheory/NormalClosure.lean b/Mathlib/FieldTheory/Normal/Closure.lean similarity index 99% rename from Mathlib/FieldTheory/NormalClosure.lean rename to Mathlib/FieldTheory/Normal/Closure.lean index f71e7bfc0ab78..a2a41471b6957 100644 --- a/Mathlib/FieldTheory/NormalClosure.lean +++ b/Mathlib/FieldTheory/Normal/Closure.lean @@ -5,7 +5,7 @@ Authors: Thomas Browning -/ import Mathlib.RingTheory.SimpleRing.Basic -import Mathlib.FieldTheory.Normal +import Mathlib.FieldTheory.Normal.Basic import Mathlib.Order.Closure import Mathlib.LinearAlgebra.FreeModule.Finite.Matrix /-! diff --git a/Mathlib/FieldTheory/Normal/Defs.lean b/Mathlib/FieldTheory/Normal/Defs.lean new file mode 100644 index 0000000000000..a8977634afdf3 --- /dev/null +++ b/Mathlib/FieldTheory/Normal/Defs.lean @@ -0,0 +1,242 @@ +/- +Copyright (c) 2020 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau, Thomas Browning, Patrick Lutz +-/ +import Mathlib.Algebra.Polynomial.Splits +import Mathlib.FieldTheory.IntermediateField.Basic +import Mathlib.FieldTheory.Minpoly.Field + +/-! +# Normal field extensions + +In this file we define normal field extensions. + +## Main Definitions + +- `Normal F K` where `K` is a field extension of `F`. +-/ + + +noncomputable section + +open Polynomial IsScalarTower + +variable (F K : Type*) [Field F] [Field K] [Algebra F K] + +/-- Typeclass for normal field extension: `K` is a normal extension of `F` iff the minimal +polynomial of every element `x` in `K` splits in `K`, i.e. every conjugate of `x` is in `K`. -/ +@[stacks 09HM] +class Normal extends Algebra.IsAlgebraic F K : Prop where + splits' (x : K) : Splits (algebraMap F K) (minpoly F x) + +variable {F K} + +theorem Normal.isIntegral (_ : Normal F K) (x : K) : IsIntegral F x := + Algebra.IsIntegral.isIntegral x + +theorem Normal.splits (_ : Normal F K) (x : K) : Splits (algebraMap F K) (minpoly F x) := + Normal.splits' x + +theorem normal_iff : Normal F K ↔ ∀ x : K, IsIntegral F x ∧ Splits (algebraMap F K) (minpoly F x) := + ⟨fun h x => ⟨h.isIntegral x, h.splits x⟩, fun h => + { isAlgebraic := fun x => (h x).1.isAlgebraic + splits' := fun x => (h x).2 }⟩ + +theorem Normal.out : Normal F K → ∀ x : K, IsIntegral F x ∧ Splits (algebraMap F K) (minpoly F x) := + normal_iff.1 + +variable (F K) + +instance normal_self : Normal F F where + isAlgebraic := fun _ => isIntegral_algebraMap.isAlgebraic + splits' := fun x => (minpoly.eq_X_sub_C' x).symm ▸ splits_X_sub_C _ + +section NormalTower + +variable (E : Type*) [Field E] [Algebra F E] [Algebra K E] [IsScalarTower F K E] + +@[stacks 09HN] +theorem Normal.tower_top_of_normal [h : Normal F E] : Normal K E := + normal_iff.2 fun x => by + cases' h.out x with hx hhx + rw [algebraMap_eq F K E] at hhx + exact + ⟨hx.tower_top, + Polynomial.splits_of_splits_of_dvd (algebraMap K E) + (Polynomial.map_ne_zero (minpoly.ne_zero hx)) + ((Polynomial.splits_map_iff (algebraMap F K) (algebraMap K E)).mpr hhx) + (minpoly.dvd_map_of_isScalarTower F K x)⟩ + +theorem AlgHom.normal_bijective [h : Normal F E] (ϕ : E →ₐ[F] K) : Function.Bijective ϕ := + h.toIsAlgebraic.bijective_of_isScalarTower' ϕ + +variable {E F} +variable {E' : Type*} [Field E'] [Algebra F E'] + +theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' := by + rw [normal_iff] at h ⊢ + intro x; specialize h (f.symm x) + rw [← f.apply_symm_apply x, minpoly.algEquiv_eq, ← f.toAlgHom.comp_algebraMap] + exact ⟨h.1.map f, splits_comp_of_splits _ _ h.2⟩ + +theorem AlgEquiv.transfer_normal (f : E ≃ₐ[F] E') : Normal F E ↔ Normal F E' := + ⟨fun _ ↦ Normal.of_algEquiv f, fun _ ↦ Normal.of_algEquiv f.symm⟩ + +end NormalTower + +namespace IntermediateField + +variable {F K} +variable {L : Type*} [Field L] [Algebra F L] [Algebra K L] [IsScalarTower F K L] + +@[simp] +theorem restrictScalars_normal {E : IntermediateField K L} : + Normal F (E.restrictScalars F) ↔ Normal F E := + Iff.rfl + +end IntermediateField + +variable {F} {K} +variable {K₁ K₂ K₃ : Type*} [Field K₁] [Field K₂] [Field K₃] [Algebra F K₁] + [Algebra F K₂] [Algebra F K₃] (ϕ : K₁ →ₐ[F] K₂) (χ : K₁ ≃ₐ[F] K₂) (ψ : K₂ →ₐ[F] K₃) + (ω : K₂ ≃ₐ[F] K₃) + +section Restrict + +variable (E : Type*) [Field E] [Algebra F E] [Algebra E K₁] [Algebra E K₂] [Algebra E K₃] + [IsScalarTower F E K₁] [IsScalarTower F E K₂] [IsScalarTower F E K₃] + +/-- Restrict algebra homomorphism to image of normal subfield -/ +def AlgHom.restrictNormalAux [h : Normal F E] : + (toAlgHom F E K₁).range →ₐ[F] (toAlgHom F E K₂).range where + toFun x := + ⟨ϕ x, by + suffices (toAlgHom F E K₁).range.map ϕ ≤ _ by exact this ⟨x, Subtype.mem x, rfl⟩ + rintro x ⟨y, ⟨z, hy⟩, hx⟩ + rw [← hx, ← hy] + apply minpoly.mem_range_of_degree_eq_one E + refine + Or.resolve_left (h.splits z).def (minpoly.ne_zero (h.isIntegral z)) (minpoly.irreducible ?_) + (minpoly.dvd E _ (by simp [aeval_algHom_apply])) + simp only [AlgHom.toRingHom_eq_coe, AlgHom.coe_toRingHom] + suffices IsIntegral F _ by exact this.tower_top + exact ((h.isIntegral z).map <| toAlgHom F E K₁).map ϕ⟩ + map_zero' := Subtype.ext (map_zero _) + map_one' := Subtype.ext (map_one _) + map_add' x y := Subtype.ext <| by simp + map_mul' x y := Subtype.ext <| by simp + commutes' x := Subtype.ext (ϕ.commutes x) + +/-- Restrict algebra homomorphism to normal subfield. -/ +@[stacks 0BME "Part 1"] +def AlgHom.restrictNormal [Normal F E] : E →ₐ[F] E := + ((AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K₂)).symm.toAlgHom.comp + (ϕ.restrictNormalAux E)).comp + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K₁)).toAlgHom + +/-- Restrict algebra homomorphism to normal subfield (`AlgEquiv` version) -/ +def AlgHom.restrictNormal' [Normal F E] : E ≃ₐ[F] E := + AlgEquiv.ofBijective (AlgHom.restrictNormal ϕ E) (AlgHom.normal_bijective F E E _) + +@[simp] +theorem AlgHom.restrictNormal_commutes [Normal F E] (x : E) : + algebraMap E K₂ (ϕ.restrictNormal E x) = ϕ (algebraMap E K₁ x) := + Subtype.ext_iff.mp + (AlgEquiv.apply_symm_apply (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F E K₂)) + (ϕ.restrictNormalAux E ⟨IsScalarTower.toAlgHom F E K₁ x, x, rfl⟩)) + +theorem AlgHom.restrictNormal_comp [Normal F E] : + (ψ.restrictNormal E).comp (ϕ.restrictNormal E) = (ψ.comp ϕ).restrictNormal E := + AlgHom.ext fun _ => + (algebraMap E K₃).injective (by simp only [AlgHom.comp_apply, AlgHom.restrictNormal_commutes]) + +/-- Restrict algebra isomorphism to a normal subfield -/ +def AlgEquiv.restrictNormal [Normal F E] : E ≃ₐ[F] E := + AlgHom.restrictNormal' χ.toAlgHom E + +@[simp] +theorem AlgEquiv.restrictNormal_commutes [Normal F E] (x : E) : + algebraMap E K₂ (χ.restrictNormal E x) = χ (algebraMap E K₁ x) := + χ.toAlgHom.restrictNormal_commutes E x + +theorem AlgEquiv.restrictNormal_trans [Normal F E] : + (χ.trans ω).restrictNormal E = (χ.restrictNormal E).trans (ω.restrictNormal E) := + AlgEquiv.ext fun _ => + (algebraMap E K₃).injective + (by simp only [AlgEquiv.trans_apply, AlgEquiv.restrictNormal_commutes]) + +/-- Restriction to a normal subfield as a group homomorphism -/ +def AlgEquiv.restrictNormalHom [Normal F E] : (K₁ ≃ₐ[F] K₁) →* E ≃ₐ[F] E := + MonoidHom.mk' (fun χ => χ.restrictNormal E) fun ω χ => χ.restrictNormal_trans ω E + +lemma AlgEquiv.restrictNormalHom_apply (L : IntermediateField F K₁) [Normal F L] + (σ : (K₁ ≃ₐ[F] K₁)) (x : L) : restrictNormalHom L σ x = σ x := + AlgEquiv.restrictNormal_commutes σ L x + +variable (F K₁) + +/-- If `K₁/E/F` is a tower of fields with `E/F` normal then `AlgHom.restrictNormal'` is an + equivalence. -/ +@[simps, stacks 0BR4] +def Normal.algHomEquivAut [Normal F E] : (E →ₐ[F] K₁) ≃ E ≃ₐ[F] E where + toFun σ := AlgHom.restrictNormal' σ E + invFun σ := (IsScalarTower.toAlgHom F E K₁).comp σ.toAlgHom + left_inv σ := by + ext + simp [AlgHom.restrictNormal'] + right_inv σ := by + ext + simp only [AlgHom.restrictNormal', AlgEquiv.toAlgHom_eq_coe, AlgEquiv.coe_ofBijective] + apply NoZeroSMulDivisors.algebraMap_injective E K₁ + rw [AlgHom.restrictNormal_commutes] + simp + +end Restrict + +section lift + +/-- The group homomorphism given by restricting an algebra isomorphism to itself +is the identity map. -/ +@[simp] +theorem AlgEquiv.restrictNormalHom_id (F K : Type*) + [Field F] [Field K] [Algebra F K] [Normal F K] : + AlgEquiv.restrictNormalHom K = MonoidHom.id (K ≃ₐ[F] K) := by + ext f x + dsimp only [restrictNormalHom, MonoidHom.mk'_apply, MonoidHom.id_apply] + apply (algebraMap K K).injective + rw [AlgEquiv.restrictNormal_commutes] + simp only [Algebra.id.map_eq_id, RingHom.id_apply] + +namespace IsScalarTower + +/-- In a scalar tower `K₃/K₂/K₁/F` with `K₁` and `K₂` are normal over `F`, the group homomorphism +given by the restriction of algebra isomorphisms of `K₃` to `K₁` is equal to the composition of +the group homomorphism given by the restricting an algebra isomorphism of `K₃` to `K₂` and +the group homomorphism given by the restricting an algebra isomorphism of `K₂` to `K₁` -/ +theorem AlgEquiv.restrictNormalHom_comp (F K₁ K₂ K₃ : Type*) + [Field F] [Field K₁] [Field K₂] [Field K₃] + [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] + [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] + [Normal F K₁] [Normal F K₂] : + AlgEquiv.restrictNormalHom K₁ = + (AlgEquiv.restrictNormalHom K₁).comp + (AlgEquiv.restrictNormalHom (F := F) (K₁ := K₃) K₂) := by + ext f x + apply (algebraMap K₁ K₃).injective + rw [IsScalarTower.algebraMap_eq K₁ K₂ K₃] + simp only [AlgEquiv.restrictNormalHom, MonoidHom.mk'_apply, RingHom.coe_comp, Function.comp_apply, + ← algebraMap_apply, AlgEquiv.restrictNormal_commutes, MonoidHom.coe_comp] + +theorem AlgEquiv.restrictNormalHom_comp_apply (K₁ K₂ : Type*) {F K₃ : Type*} + [Field F] [Field K₁] [Field K₂] [Field K₃] + [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] [Algebra K₁ K₂] [Algebra K₁ K₃] [Algebra K₂ K₃] + [IsScalarTower F K₁ K₃] [IsScalarTower F K₁ K₂] [IsScalarTower F K₂ K₃] [IsScalarTower K₁ K₂ K₃] + [Normal F K₁] [Normal F K₂] (f : K₃ ≃ₐ[F] K₃) : + AlgEquiv.restrictNormalHom K₁ f = + (AlgEquiv.restrictNormalHom K₁) (AlgEquiv.restrictNormalHom K₂ f) := by + rw [IsScalarTower.AlgEquiv.restrictNormalHom_comp F K₁ K₂ K₃, MonoidHom.comp_apply] + +end IsScalarTower + +end lift diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index b8174c3c41d64..239f457cd3474 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning, Patrick Lutz -/ import Mathlib.FieldTheory.IsAlgClosed.Basic +import Mathlib.FieldTheory.SplittingField.Construction import Mathlib.RingTheory.IntegralDomain import Mathlib.RingTheory.Polynomial.UniqueFactorization diff --git a/Mathlib/FieldTheory/Relrank.lean b/Mathlib/FieldTheory/Relrank.lean index 6b32cbc7d88e9..2ed666bbdae0a 100644 --- a/Mathlib/FieldTheory/Relrank.lean +++ b/Mathlib/FieldTheory/Relrank.lean @@ -125,6 +125,7 @@ theorem relfinrank_top_left : relfinrank ⊤ A = 1 := relfinrank_eq_one_of_le le set_option synthInstance.maxHeartbeats 400000 in @[simp] theorem relrank_top_right : relrank A ⊤ = Module.rank A E := by + let _ : AddCommMonoid (⊤ : IntermediateField A E) := inferInstance rw [relrank_eq_rank_of_le (show A ≤ ⊤ from le_top), extendScalars_top, IntermediateField.topEquiv.toLinearEquiv.rank_eq] diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 8cc00db3d0c2b..2762762a786c4 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -6,7 +6,7 @@ Authors: Jz Pan import Mathlib.FieldTheory.SplittingField.Construction import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure import Mathlib.FieldTheory.Separable -import Mathlib.FieldTheory.NormalClosure +import Mathlib.FieldTheory.Normal.Closure import Mathlib.RingTheory.AlgebraicIndependent.Adjoin import Mathlib.RingTheory.AlgebraicIndependent.TranscendenceBasis import Mathlib.RingTheory.Polynomial.SeparableDegree @@ -598,7 +598,7 @@ theorem eq_X_pow_char_pow_sub_C_pow_of_natSepDegree_eq_one (q : ℕ) [ExpChar F refine ⟨_, n, y, hne, H, ?_⟩ obtain ⟨c, hf, H⟩ := hF.exists_eq_pow_mul_and_not_dvd rw [hf, natSepDegree_mul_of_isCoprime _ c <| IsCoprime.pow_left <| - (hI.coprime_or_dvd c).resolve_right H, natSepDegree_pow_of_ne_zero _ hne, hD, + (hI.isCoprime_or_dvd c).resolve_right H, natSepDegree_pow_of_ne_zero _ hne, hD, add_right_eq_self, natSepDegree_eq_zero_iff] at h simpa only [eq_one_of_monic_natDegree_zero ((hM.pow _).of_mul_monic_left (hf ▸ hm)) h, mul_one, ← hp] using hf diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 038fc99fd056f..90119d260be35 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -301,8 +301,7 @@ theorem eq_mongePoint_of_forall_mem_mongePlane {n : ℕ} {s : Simplex ℝ P (n + · rintro ⟨i, rfl⟩ use i, ⟨Set.mem_univ _, i.property.symm⟩ · rintro ⟨i, ⟨-, hi⟩, rfl⟩ - -- Porting note: was `use ⟨i, hi.symm⟩, rfl` - exact ⟨⟨i, hi.symm⟩, rfl⟩ + use ⟨i, hi.symm⟩ rw [hu, ← vectorSpan_image_eq_span_vsub_set_left_ne ℝ _ (Set.mem_univ _), Set.image_univ] at hi have hv : p -ᵥ s.mongePoint ∈ vectorSpan ℝ (Set.range s.points) := by let s₁ : Finset (Fin (n + 3)) := univ.erase i₁ diff --git a/Mathlib/Geometry/Group/Growth/LinearLowerBound.lean b/Mathlib/Geometry/Group/Growth/LinearLowerBound.lean new file mode 100644 index 0000000000000..153cde694cf7f --- /dev/null +++ b/Mathlib/Geometry/Group/Growth/LinearLowerBound.lean @@ -0,0 +1,90 @@ +/- +Copyright (c) 2024 Yaël Dillies, Patrick Luo, Eric Rodriguez. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Patrick Luo, Eric Rodriguez +-/ +import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Data.Nat.SuccPred + +/-! +# Linear lower bound on the growth of a generating set + +This file proves that the growth of a set generating an infinite group is at least linear. +-/ + +open Subgroup +open scoped Pointwise + +namespace Finset +variable {G : Type*} [Group G] [DecidableEq G] {X : Finset G} {n : ℕ} + +@[to_additive] +lemma pow_ssubset_pow_succ_of_pow_ne_closure (hX₁ : (1 : G) ∈ X) (hX : X.Nontrivial) + (hXclosure : (X ^ n : Set G) ≠ closure (X : Set G)) : X ^ n ⊂ X ^ (n + 1) := by + obtain rfl | hn := eq_or_ne n 0 + · simpa [ssubset_iff_subset_not_subset, hX₁, -Finset.subset_singleton_iff] + using hX.not_subset_singleton + refine (pow_subset_pow_right hX₁ <| n.le_add_right _).ssubset_of_ne ?_ + contrapose! hXclosure with hXn + rw [← closure_pow (mod_cast hX₁) hn] + wlog hn₁ : n = 1 + · simp +contextual only [pow_one] at this + replace hXn d : X ^ (n + d) = X ^ n := by + induction' d with d hd + · rw [add_zero] + · rw [pow_add, pow_one] at hXn + rw [← add_assoc, pow_add, pow_one, hd, ← hXn] + exact mod_cast this (one_mem_pow hX₁) (hX.pow hn) one_ne_zero + (by simp [hXn, ← pow_mul, mul_two]) (by simp [hXn, ← pow_mul, mul_two]) + subst hn₁ + simp only [ne_eq, one_ne_zero, not_false_eq_true, Nat.reduceAdd, pow_one] at * + let Xgp : Subgroup G := + { carrier := X + mul_mem' := fun {x y} hx hy ↦ by + norm_cast at * + simpa [← hXn, ← sq] using mul_mem_mul hx hy + one_mem' := hX₁ + inv_mem' := fun {x} hx ↦ by + dsimp at * + norm_cast at * + have : x • X ⊆ X := by + simpa [← hXn, add_assoc, ← sq] using smul_finset_subset_mul (t := X) hx + have : x • X = X := eq_of_subset_of_card_le this (card_smul_finset ..).ge + rw [← eq_inv_smul_iff] at this + rw [this] + simpa [mem_inv_smul_finset_iff] } + exact subset_closure.antisymm <| (closure_le Xgp).2 subset_rfl + +@[to_additive] +lemma pow_right_strictMonoOn (hX₁ : 1 ∈ X) (hX : X.Nontrivial) : + StrictMonoOn (fun n ↦ X ^ n) {n | (X ^ (n - 1) : Set G) ≠ closure (X : Set G)} := by + refine strictMonoOn_of_lt_add_one ⟨?_⟩ fun n _ _ hn ↦ + pow_ssubset_pow_succ_of_pow_ne_closure hX₁ hX hn + rintro - - n hn m ⟨-, hmn⟩ hm + apply hn + obtain rfl | hm₀ := m.eq_zero_or_pos + · simp [eq_comm (a := (1 : Set _)), coe_set_eq_one, -Set.subset_singleton_iff, + hX.coe.not_subset_singleton] at hm + · calc (X : Set G) ^ (n - 1) + _ = X ^ (n - m) * X ^ (m - 1) := by rw [← pow_add]; congr 1; omega + _ = closure (X : Set G) := by rw [hm, Set.pow_mul_subgroupClosure hX.nonempty.to_set] + +@[to_additive] +lemma pow_right_strictMono (hX₁ : 1 ∈ X) (hXclosure : (closure (X : Set G) : Set G).Infinite) : + StrictMono fun n ↦ X ^ n := by + obtain rfl | hX := eq_singleton_or_nontrivial hX₁ + · simp [closure_singleton_one] at hXclosure + have h n : (X ^ (n - 1) : Set G) ≠ closure (X : Set G) := + fun h ↦ by simp [← h, ← coe_pow] at hXclosure + simpa [h] using pow_right_strictMonoOn hX₁ hX + +/-- The growth of a set generating an infinite group is at least linear. -/ +@[to_additive "The growth of a set generating an infinite group is at least linear."] +lemma add_one_le_card_pow (hX₁ : 1 ∈ X) (hXclosure : (closure (X : Set G) : Set G).Infinite) : + ∀ n, n + 1 ≤ #(X ^ n) + | 0 => by simp + | n + 1 => (add_one_le_card_pow hX₁ hXclosure _).trans_lt <| card_lt_card <| + pow_right_strictMono hX₁ (by simp [hXclosure, Set.infinite_univ]) n.lt_succ_self + +end Finset diff --git a/Mathlib/Geometry/Group/Growth/README.md b/Mathlib/Geometry/Group/Growth/README.md index 6f07dca284151..41906654ee16b 100644 --- a/Mathlib/Geometry/Group/Growth/README.md +++ b/Mathlib/Geometry/Group/Growth/README.md @@ -7,4 +7,5 @@ If $G$ is a finitely generated group, say $G = \langle S\rangle$ for some finite ## Topics The results on growth of groups currently covered are: +* The growth of a group is at least linear * The growth of a group $G$ is roughly the growth of a normal subgroup $H \le G$ times the growth of $G / H$. diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index a862ca47f3f78..7bcb7a3778950 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -120,9 +120,9 @@ variable {H : Type u} {H' : Type*} {M : Type*} {M' : Type*} {M'' : Type*} `PartialHomeomorph.trans` and `PartialEquiv.trans`. Note that, as is usual for equivs, the composition is from left to right, hence the direction of the arrow. -/ -scoped[Manifold] infixr:100 " ≫ₕ " => PartialHomeomorph.trans +@[inherit_doc] scoped[Manifold] infixr:100 " ≫ₕ " => PartialHomeomorph.trans -scoped[Manifold] infixr:100 " ≫ " => PartialEquiv.trans +@[inherit_doc] scoped[Manifold] infixr:100 " ≫ " => PartialEquiv.trans open Set PartialHomeomorph Manifold -- Porting note: Added `Manifold` @@ -777,9 +777,8 @@ instance (H : Type*) [TopologicalSpace H] (H' : Type*) [TopologicalSpace H'] : TopologicalSpace (ModelProd H H') := instTopologicalSpaceProd --- Porting note: simpNF false positive --- Next lemma shows up often when dealing with derivatives, register it as simp. -@[simp, mfld_simps, nolint simpNF] +-- Next lemma shows up often when dealing with derivatives, so we register it as simp lemma. +@[simp, mfld_simps] theorem modelProd_range_prod_id {H : Type*} {H' : Type*} {α : Type*} (f : H → α) : (range fun p : ModelProd H H' ↦ (f p.1, p.2)) = range f ×ˢ (univ : Set H') := by rw [prod_range_univ_eq] @@ -911,8 +910,6 @@ end ChartedSpace have a topological structure, where the topology would come from the charts. For this, one needs charts that are only partial equivalences, and continuity properties for their composition. This is formalised in `ChartedSpaceCore`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure ChartedSpaceCore (H : Type*) [TopologicalSpace H] (M : Type*) where /-- An atlas of charts, which are only `PartialEquiv`s -/ atlas : Set (PartialEquiv M H) @@ -1283,8 +1280,6 @@ lemma StructureGroupoid.restriction_in_maximalAtlas {e : PartialHomeomorph M H} /-- A `G`-diffeomorphism between two charted spaces is a homeomorphism which, when read in the charts, belongs to `G`. We avoid the word diffeomorph as it is too related to the smooth category, and use structomorph instead. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure Structomorph (G : StructureGroupoid H) (M : Type*) (M' : Type*) [TopologicalSpace M] [TopologicalSpace M'] [ChartedSpace H M] [ChartedSpace H M'] extends Homeomorph M M' where mem_groupoid : ∀ c : PartialHomeomorph M H, ∀ c' : PartialHomeomorph M' H, c ∈ atlas H M → diff --git a/Mathlib/Geometry/Manifold/ContMDiffMap.lean b/Mathlib/Geometry/Manifold/ContMDiffMap.lean index c4d2dbbb08733..07becbcbb3db7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMap.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMap.lean @@ -25,7 +25,9 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] (n : WithTop ℕ∞) variable (I I') in -/-- Bundled `n` times continuously differentiable maps. -/ +/-- Bundled `n` times continuously differentiable maps, +denoted as `C^n(I, M; I', M')` and `C^n(I, M; k)` (when the target is a normed space `k` with +the trivial model) in the `Manifold` namespace. -/ def ContMDiffMap := { f : M → M' // ContMDiff I I' n f } diff --git a/Mathlib/Geometry/Manifold/DerivationBundle.lean b/Mathlib/Geometry/Manifold/DerivationBundle.lean index 0cbb33ee8efac..4cca168870307 100644 --- a/Mathlib/Geometry/Manifold/DerivationBundle.lean +++ b/Mathlib/Geometry/Manifold/DerivationBundle.lean @@ -32,7 +32,8 @@ instance smoothFunctionsAlgebra : Algebra 𝕜 C^∞⟮I, M; 𝕜⟯ := by infer instance smooth_functions_tower : IsScalarTower 𝕜 C^∞⟮I, M; 𝕜⟯ C^∞⟮I, M; 𝕜⟯ := by infer_instance /-- Type synonym, introduced to put a different `SMul` action on `C^n⟮I, M; 𝕜⟯` -which is defined as `f • r = f(x) * r`. -/ +which is defined as `f • r = f(x) * r`. +Denoted as `C^n⟮I, M; 𝕜⟯⟨x⟩` within the `Derivation` namespace. -/ @[nolint unusedArguments] def PointedContMDiffMap (_ : M) := C^n⟮I, M; 𝕜⟯ @@ -123,7 +124,8 @@ variable {I} {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Ty [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] -/-- The heterogeneous differential as a linear map. Instead of taking a function as an argument this +/-- The heterogeneous differential as a linear map, denoted as `𝒅ₕ` within the `Manifold` namespace. +Instead of taking a function as an argument this differential takes `h : f x = y`. It is particularly handy to deal with situations where the points on where it has to be evaluated are equal but not definitionally equal. -/ def hfdifferential {f : C^∞⟮I, M; I', M'⟯} {x : M} {y : M'} (h : f x = y) : @@ -143,16 +145,16 @@ def hfdifferential {f : C^∞⟮I, M; I', M'⟯} {x : M} {y : M'} (h : f x = y) map_smul' _ _ := rfl map_add' _ _ := rfl -/-- The homogeneous differential as a linear map. -/ +/-- The homogeneous differential as a linear map, denoted as `𝒅` within the `Manifold` namespace. -/ def fdifferential (f : C^∞⟮I, M; I', M'⟯) (x : M) : PointDerivation I x →ₗ[𝕜] PointDerivation I' (f x) := hfdifferential (rfl : f x = f x) -- Standard notation for the differential. The abbreviation is `MId`. -scoped[Manifold] notation "𝒅" => fdifferential +@[inherit_doc] scoped[Manifold] notation "𝒅" => fdifferential -- Standard notation for the differential. The abbreviation is `MId`. -scoped[Manifold] notation "𝒅ₕ" => hfdifferential +@[inherit_doc] scoped[Manifold] notation "𝒅ₕ" => hfdifferential @[simp] theorem fdifferential_apply (f : C^∞⟮I, M; I', M'⟯) {x : M} (v : PointDerivation I x) diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 39b1fb3e9f0b4..8c50a4f7055ec 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -64,8 +64,7 @@ section Defs variable (I I' M M' n) /-- `n`-times continuously differentiable diffeomorphism between `M` and `M'` with respect to `I` -and `I'`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] +and `I'`, denoted as `M ≃ₘ^n⟮I, I'⟯ M'` (in the `Manifold` namespace). -/ structure Diffeomorph extends M ≃ M' where protected contMDiff_toFun : ContMDiff I I' n toEquiv protected contMDiff_invFun : ContMDiff I' I n toEquiv.symm diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 55635af781f7f..78db01d88e9e2 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.IsManifold import Mathlib.Analysis.InnerProductSpace.PiL2 @@ -11,7 +12,9 @@ import Mathlib.Analysis.InnerProductSpace.PiL2 We introduce the necessary bits to be able to define manifolds modelled over `ℝ^n`, boundaryless or with boundary or with corners. As a concrete example, we construct explicitly the manifold with -boundary structure on the real interval `[x, y]`. +boundary structure on the real interval `[x, y]`, and prove that its boundary is indeed `{x,y}` +whenever `x < y`. As a corollary, a product `M × [x, y]` with a manifold `M` without boundary +has boundary `M × {x, y}`. More specifically, we introduce * `modelWithCornersEuclideanHalfSpace n : @@ -77,11 +80,9 @@ instance [NeZero n] : Inhabited (EuclideanHalfSpace n) := instance : Inhabited (EuclideanQuadrant n) := ⟨⟨0, fun _ => le_rfl⟩⟩ -instance {n : ℕ} [NeZero n] : Zero (EuclideanHalfSpace n) where - zero := ⟨fun _ ↦ 0, by norm_num⟩ +instance {n : ℕ} [NeZero n] : Zero (EuclideanHalfSpace n) := ⟨⟨fun _ ↦ 0, by norm_num⟩⟩ -instance {n : ℕ} : Zero (EuclideanQuadrant n) where - zero := ⟨fun _ ↦ 0, by norm_num⟩ +instance {n : ℕ} : Zero (EuclideanQuadrant n) := ⟨⟨fun _ ↦ 0, by norm_num⟩⟩ @[ext] theorem EuclideanQuadrant.ext (x y : EuclideanQuadrant n) (h : x.1 = y.1) : x = y := @@ -382,15 +383,18 @@ lemma IccRightChart_extend_top : norm_num [IccRightChart, modelWithCornersEuclideanHalfSpace_zero] congr -lemma IccRightChart_extend_right_mem_frontier : +lemma IccRightChart_extend_top_mem_frontier : (IccRightChart x y).extend (𝓡∂ 1) ⊤ ∈ frontier (range (𝓡∂ 1)) := by rw [IccRightChart_extend_top, frontier_range_modelWithCornersEuclideanHalfSpace, mem_setOf, PiLp.zero_apply] +@[deprecated (since := "2025-01-25")] +alias IccRightChart_extend_right_mem_frontier := IccRightChart_extend_top_mem_frontier + /-- Charted space structure on `[x, y]`, using only two charts taking values in `EuclideanHalfSpace 1`. -/ -instance IccChartedSpace (x y : ℝ) [h : Fact (x < y)] : +instance instIccChartedSpace (x y : ℝ) [h : Fact (x < y)] : ChartedSpace (EuclideanHalfSpace 1) (Icc x y) where atlas := {IccLeftChart x y, IccRightChart x y} chartAt z := if z.val < y then IccLeftChart x y else IccRightChart x y @@ -403,8 +407,61 @@ instance IccChartedSpace (x y : ℝ) [h : Fact (x < y)] : simpa only [not_lt] using h' chart_mem_atlas z := by by_cases h' : (z : ℝ) < y <;> simp [h'] -/-- The manifold structure on `[x, y]` is smooth. --/ +@[simp] +lemma Icc_chartedSpaceChartAt {z : Set.Icc x y} : + chartAt _ z = if z.val < y then IccLeftChart x y else IccRightChart x y := rfl + +lemma Icc_chartedSpaceChartAt_of_le_top {z : Set.Icc x y} (h : z.val < y) : + chartAt _ z = IccLeftChart x y := by + simp [Icc_chartedSpaceChartAt, h] + +lemma Icc_chartedSpaceChartAt_of_top_le {z : Set.Icc x y} (h : y ≤ z.val) : + chartAt _ z = IccRightChart x y := by + simp [Icc_chartedSpaceChartAt, reduceIte, not_lt.mpr h] + +lemma Icc_isBoundaryPoint_bot : (𝓡∂ 1).IsBoundaryPoint (⊥ : Set.Icc x y) := by + rw [ModelWithCorners.isBoundaryPoint_iff, extChartAt, + Icc_chartedSpaceChartAt_of_le_top (by norm_num [hxy.out])] + exact IccLeftChart_extend_bot_mem_frontier + +lemma Icc_isBoundaryPoint_top : (𝓡∂ 1).IsBoundaryPoint (⊤ : Set.Icc x y) := by + rw [ModelWithCorners.isBoundaryPoint_iff, extChartAt, + Icc_chartedSpaceChartAt_of_top_le (by norm_num)] + exact IccRightChart_extend_top_mem_frontier + +lemma Icc_isInteriorPoint_interior {p : Set.Icc x y} (hp : x < p.val ∧ p.val < y) : + (𝓡∂ 1).IsInteriorPoint p := by + rw [ModelWithCorners.IsInteriorPoint, extChartAt, Icc_chartedSpaceChartAt_of_le_top hp.2, + interior_range_modelWithCornersEuclideanHalfSpace] + exact IccLeftChart_extend_interior_pos hp + +lemma boundary_Icc : (𝓡∂ 1).boundary (Icc x y) = {⊥, ⊤} := by + ext p + rcases Set.eq_endpoints_or_mem_Ioo_of_mem_Icc p.2 with (hp | hp | hp) + · have : p = ⊥ := SetCoe.ext hp + rw [this] + apply iff_of_true Icc_isBoundaryPoint_bot (mem_insert ⊥ {⊤}) + · have : p = ⊤ := SetCoe.ext hp + rw [this] + apply iff_of_true Icc_isBoundaryPoint_top (mem_insert_of_mem ⊥ rfl) + · apply iff_of_false + · simpa [← mem_compl_iff, ModelWithCorners.compl_boundary] using + Icc_isInteriorPoint_interior hp + · rw [mem_insert_iff, mem_singleton_iff] + push_neg + constructor <;> by_contra h <;> rw [congrArg Subtype.val h] at hp + exacts [left_mem_Ioo.mp hp, right_mem_Ioo.mp hp] + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + {H : Type*} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + +/-- A product `M × [x,y]` for `M` boundaryless has boundary `M × {x, y}`. -/ +lemma boundary_product [I.Boundaryless] : + (I.prod (𝓡∂ 1)).boundary (M × Icc x y) = Set.prod univ {⊥, ⊤} := by + rw [I.boundary_of_boundaryless_left, boundary_Icc] + +/-- The manifold structure on `[x, y]` is smooth. -/ instance instIsManifoldIcc (x y : ℝ) [Fact (x < y)] {n : WithTop ℕ∞} : IsManifold (𝓡∂ 1) n (Icc x y) := by have M : ContDiff ℝ n (show EuclideanSpace ℝ (Fin 1) → EuclideanSpace ℝ (Fin 1) @@ -444,7 +501,7 @@ instance instIsManifoldIcc (x y : ℝ) [Fact (x < y)] {n : WithTop ℕ∞} : exact (mem_groupoid_of_pregroupoid.mpr (symm_trans_mem_contDiffGroupoid _)).1 /-! Register the manifold structure on `Icc 0 1`. These are merely special cases of -`IccChartedSpace` and `instIsManifoldIcc`. -/ +`instIccChartedSpace` and `instIsManifoldIcc`. -/ section diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index d3f2825cc0551..7f6c2a06f8852 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -28,6 +28,13 @@ of `M` and `N`. - `ModelWithCorners.boundary_prod`: the boundary of `M × N` is `∂M × N ∪ (M × ∂N)`. - `ModelWithCorners.BoundarylessManifold.prod`: if `M` and `N` are boundaryless, so is `M × N` +- `ModelWithCorners.interior_disjointUnion`: the interior of a disjoint union `M ⊔ M'` + is the union of the interior of `M` and `M'` +- `ModelWithCorners.boundary_disjointUnion`: the boundary of a disjoint union `M ⊔ M'` + is the union of the boundaries of `M` and `M'` +- `ModelWithCorners.boundaryless_disjointUnion`: if `M` and `M'` are boundaryless, + so is their disjoint union `M ⊔ M'` + ## Tags manifold, interior, boundary @@ -85,7 +92,7 @@ lemma isBoundaryPoint_iff {x : M} : I.IsBoundaryPoint x ↔ extChartAt I x x ∈ lemma isInteriorPoint_or_isBoundaryPoint (x : M) : I.IsInteriorPoint x ∨ I.IsBoundaryPoint x := by rw [IsInteriorPoint, or_iff_not_imp_left, I.isBoundaryPoint_iff, ← closure_diff_interior, I.isClosed_range.closure_eq, mem_diff] - exact fun h => ⟨mem_range_self _, h⟩ + exact fun h ↦ ⟨mem_range_self _, h⟩ /-- A manifold decomposes into interior and boundary. -/ lemma interior_union_boundary_eq_univ : (I.interior M) ∪ (I.boundary M) = (univ : Set M) := @@ -97,13 +104,22 @@ lemma disjoint_interior_boundary : Disjoint (I.interior M) (I.boundary M) := by -- Choose some x in the intersection of interior and boundary. obtain ⟨x, h1, h2⟩ := not_disjoint_iff.mp h rw [← mem_empty_iff_false (extChartAt I x x), - ← disjoint_iff_inter_eq_empty.mp (disjoint_interior_frontier (s := range I)), mem_inter_iff] + ← disjoint_iff_inter_eq_empty.mp disjoint_interior_frontier, mem_inter_iff] exact ⟨h1, h2⟩ +lemma isInteriorPoint_iff_not_isBoundaryPoint (x : M) : + I.IsInteriorPoint x ↔ ¬I.IsBoundaryPoint x := by + refine ⟨?_, + by simpa only [or_iff_not_imp_right] using isInteriorPoint_or_isBoundaryPoint x (I := I)⟩ + by_contra! h + rw [← mem_empty_iff_false (extChartAt I x x), + ← disjoint_iff_inter_eq_empty.mp disjoint_interior_frontier, mem_inter_iff] + exact h + /-- The boundary is the complement of the interior. -/ lemma compl_interior : (I.interior M)ᶜ = I.boundary M:= by apply compl_unique ?_ I.interior_union_boundary_eq_univ - exact disjoint_iff_inter_eq_empty.mp (I.disjoint_interior_boundary) + exact disjoint_iff_inter_eq_empty.mp I.disjoint_interior_boundary /-- The interior is the complement of the boundary. -/ lemma compl_boundary : (I.boundary M)ᶜ = I.interior M:= by @@ -231,4 +247,115 @@ instance BoundarylessManifold.prod [BoundarylessManifold I M] [BoundarylessManif end prod +section disjointUnion + +variable {M' : Type*} [TopologicalSpace M'] [ChartedSpace H M'] {n : WithTop ℕ∞} + [hM : IsManifold I n M] [hM' : IsManifold I n M'] + [Nonempty M] [Nonempty M'] [Nonempty H] + {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] + {J : Type*} {J : ModelWithCorners 𝕜 E' H'} + {N N' : Type*} [TopologicalSpace N] [TopologicalSpace N'] [ChartedSpace H' N] [ChartedSpace H' N'] + [IsManifold J n N] [IsManifold J n N'] [Nonempty N] [Nonempty N'] [Nonempty H'] + +open Topology + +lemma interiorPoint_inl (x : M) (hx : I.IsInteriorPoint x) : + I.IsInteriorPoint (.inl x: M ⊕ M') := by + rw [I.isInteriorPoint_iff, extChartAt, ChartedSpace.sum_chartAt_inl] + dsimp only [PartialHomeomorph.extend.eq_1, PartialEquiv.trans_target, toPartialEquiv_coe_symm, + PartialHomeomorph.lift_openEmbedding_target, PartialEquiv.coe_trans, toPartialEquiv_coe, + PartialHomeomorph.toFun_eq_coe, PartialHomeomorph.lift_openEmbedding_toFun, Function.comp_apply] + rw [Sum.inl_injective.extend_apply (chartAt H x)] + simpa [I.isInteriorPoint_iff, extChartAt] using hx + +lemma boundaryPoint_inl (x : M) (hx : I.IsBoundaryPoint x) : + I.IsBoundaryPoint (.inl x: M ⊕ M') := by + rw [I.isBoundaryPoint_iff, extChartAt, ChartedSpace.sum_chartAt_inl] + dsimp + rw [Sum.inl_injective.extend_apply (chartAt H x)] + simpa [I.isBoundaryPoint_iff, extChartAt] using hx + +lemma interiorPoint_inr (x : M') (hx : I.IsInteriorPoint x) : + I.IsInteriorPoint (.inr x : M ⊕ M') := by + rw [I.isInteriorPoint_iff, extChartAt, ChartedSpace.sum_chartAt_inr] + dsimp + rw [Sum.inr_injective.extend_apply (chartAt H x)] + simpa [I.isInteriorPoint_iff, extChartAt] using hx + +lemma boundaryPoint_inr (x : M') (hx : I.IsBoundaryPoint x) : + I.IsBoundaryPoint (.inr x : M ⊕ M') := by + rw [I.isBoundaryPoint_iff, extChartAt, ChartedSpace.sum_chartAt_inr] + dsimp + rw [Sum.inr_injective.extend_apply (chartAt H x)] + simpa [I.isBoundaryPoint_iff, extChartAt] using hx + +-- Converse to the previous direction: if `x` were not an interior point, +-- it had to be a boundary point, hence `p` were a boundary point also, contradiction. +lemma isInteriorPoint_disjointUnion_left {p : M ⊕ M'} (hp : I.IsInteriorPoint p) + (hleft : Sum.isLeft p) : I.IsInteriorPoint (Sum.getLeft p hleft) := by + by_contra h + set x := Sum.getLeft p hleft + rw [isInteriorPoint_iff_not_isBoundaryPoint x, not_not] at h + rw [isInteriorPoint_iff_not_isBoundaryPoint p] at hp + have := boundaryPoint_inl (M' := M') x (by tauto) + rw [← Sum.eq_left_getLeft_of_isLeft hleft] at this + exact hp this + +lemma isInteriorPoint_disjointUnion_right {p : M ⊕ M'} (hp : I.IsInteriorPoint p) + (hright : Sum.isRight p) : I.IsInteriorPoint (Sum.getRight p hright) := by + by_contra h + set x := Sum.getRight p hright + rw [← mem_empty_iff_false p, ← (disjoint_interior_boundary (I := I)).inter_eq] + constructor + · rw [ModelWithCorners.interior, mem_setOf]; exact hp + · rw [ModelWithCorners.boundary, mem_setOf, Sum.eq_right_getRight_of_isRight hright] + have := isInteriorPoint_or_isBoundaryPoint (I := I) x + exact boundaryPoint_inr (M' := M') x (by tauto) + +lemma interior_disjointUnion : + ModelWithCorners.interior (I := I) (M ⊕ M') = + Sum.inl '' (ModelWithCorners.interior (I := I) M) + ∪ Sum.inr '' (ModelWithCorners.interior (I := I) M') := by + ext p + constructor + · intro hp + by_cases h : Sum.isLeft p + · left + exact ⟨Sum.getLeft p h, isInteriorPoint_disjointUnion_left hp h, Sum.inl_getLeft p h⟩ + · replace h := Sum.not_isLeft.mp h + right + exact ⟨Sum.getRight p h, isInteriorPoint_disjointUnion_right hp h, Sum.inr_getRight p h⟩ + · intro hp + by_cases h : Sum.isLeft p + · set x := Sum.getLeft p h with x_eq + rw [Sum.eq_left_getLeft_of_isLeft h] + apply interiorPoint_inl x + have hp : p ∈ Sum.inl '' (ModelWithCorners.interior (I := I) M) := by + obtain (good | ⟨y, hy, hxy⟩) := hp + exacts [good, (not_isLeft_and_isRight ⟨h, by rw [← hxy]; exact rfl⟩).elim] + obtain ⟨x', hx', hx'p⟩ := hp + simpa [x_eq, ← hx'p, Sum.getLeft_inl] + · set x := Sum.getRight p (Sum.not_isLeft.mp h) with x_eq + rw [Sum.eq_right_getRight_of_isRight (Sum.not_isLeft.mp h)] + apply interiorPoint_inr x + have hp : p ∈ Sum.inr '' (ModelWithCorners.interior (I := I) M') := by + obtain (⟨y, hy, hxy⟩ | good) := hp + exacts [(not_isLeft_and_isRight ⟨by rw [← hxy]; exact rfl, Sum.not_isLeft.mp h⟩).elim, good] + obtain ⟨x', hx', hx'p⟩ := hp + simpa [x_eq, ← hx'p, Sum.getRight_inr] + +lemma boundary_disjointUnion : ModelWithCorners.boundary (I := I) (M ⊕ M') = + Sum.inl '' (ModelWithCorners.boundary (I := I) M) + ∪ Sum.inr '' (ModelWithCorners.boundary (I := I) M') := by + simp only [← ModelWithCorners.compl_interior, interior_disjointUnion, inl_compl_union_inr_compl] + +/-- If `M` and `M'` are boundaryless, so is their disjoint union `M ⊔ M'`. -/ +instance boundaryless_disjointUnion + [hM: BoundarylessManifold I M] [hM': BoundarylessManifold I M'] : + BoundarylessManifold I (M ⊕ M') := by + rw [← Boundaryless.iff_boundary_eq_empty] at hM hM' ⊢ + simp [boundary_disjointUnion, hM, hM'] + +end disjointUnion + end ModelWithCorners diff --git a/Mathlib/Geometry/Manifold/IsManifold.lean b/Mathlib/Geometry/Manifold/IsManifold.lean index 29e30318e722e..69d14ee8012e6 100644 --- a/Mathlib/Geometry/Manifold/IsManifold.lean +++ b/Mathlib/Geometry/Manifold/IsManifold.lean @@ -46,6 +46,11 @@ but add these assumptions later as needed. (Quite a few results still do not req we register them as `PartialEquiv`s. `extChartAt I x` is the canonical such partial equiv around `x`. +We define a few constructions of smooth manifolds: +* every empty type is a smooth manifold +* the product of two smooth manifolds +* the disjoint union of two manifolds (over the same charted space) + As specific examples of models with corners, we define (in `Geometry.Manifold.Instances.Real`) * `modelWithCornersEuclideanHalfSpace n : ModelWithCorners ℝ (EuclideanSpace ℝ (Fin n)) (EuclideanHalfSpace n)` for the model space used to @@ -149,7 +154,7 @@ defined, the latter ensures that for `C^2` maps the second derivatives are symme on the boundary, as these are limit points of interior points where symmetry holds. If further conditions turn out to be useful, they can be added here. -/ -@[ext] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was nolint has_nonempty_instance +@[ext] structure ModelWithCorners (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Type*) [NormedAddCommGroup E] [NormedSpace 𝕜 E] (H : Type*) [TopologicalSpace H] extends PartialEquiv H E where @@ -161,7 +166,7 @@ structure ModelWithCorners (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Ty attribute [simp, mfld_simps] ModelWithCorners.source_eq -/-- A vector space is a model with corners. -/ +/-- A vector space is a model with corners, denoted as `𝓘(𝕜, E)` within the `Manifold` namespace. -/ def modelWithCornersSelf (𝕜 : Type*) [NontriviallyNormedField 𝕜] (E : Type*) [NormedAddCommGroup E] [NormedSpace 𝕜 E] : ModelWithCorners 𝕜 E E where toPartialEquiv := PartialEquiv.refl E @@ -598,6 +603,26 @@ theorem contDiffGroupoid_zero_eq : contDiffGroupoid 0 I = continuousGroupoid H : · refine I.continuous.comp_continuousOn (u.symm.continuousOn.comp I.continuousOn_symm ?_) exact (mapsTo_preimage _ _).mono_left inter_subset_left +-- FIXME: does this generalise to other groupoids? The argument is not specific +-- to C^n functions, but uses something about the groupoid's property that is not easy to abstract. +/-- Any change of coordinates with empty source belongs to `contDiffGroupoid`. -/ +lemma ContDiffGroupoid.mem_of_source_eq_empty (f : PartialHomeomorph H H) + (hf : f.source = ∅) : f ∈ contDiffGroupoid n I := by + constructor + · intro x ⟨hx, _⟩ + rw [mem_preimage] at hx + simp_all only [mem_empty_iff_false] + · intro x ⟨hx, _⟩ + have : f.target = ∅ := by simp [← f.image_source_eq_target, hf] + simp_all [hx] + +include I in +/-- Any change of coordinates with empty source belongs to `continuousGroupoid`. -/ +lemma ContinuousGroupoid.mem_of_source_eq_empty (f : PartialHomeomorph H H) + (hf : f.source = ∅) : f ∈ continuousGroupoid H := by + rw [← contDiffGroupoid_zero_eq (I := I)] + exact ContDiffGroupoid.mem_of_source_eq_empty f hf + /-- An identity partial homeomorphism belongs to the `C^n` groupoid. -/ theorem ofSet_mem_contDiffGroupoid {s : Set H} (hs : IsOpen s) : PartialHomeomorph.ofSet s hs ∈ contDiffGroupoid n I := by @@ -797,6 +822,34 @@ instance prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedA have h2 := (contDiffGroupoid n I').compatible hf2 hg2 exact contDiffGroupoid_prod h1 h2 +section DisjointUnion + +variable {M' : Type*} [TopologicalSpace M'] [ChartedSpace H M'] + [hM : IsManifold I n M] [hM' : IsManifold I n M'] + +/-- The disjoint union of two `C^n` manifolds modelled on `(E, H)` +is a `C^n` manifold modeled on `(E, H)`. -/ +instance disjointUnion [Nonempty M] [Nonempty M'] [Nonempty H] : IsManifold I n (M ⊕ M') where + compatible {e} e' he he' := by + obtain (⟨f, hf, hef⟩ | ⟨f, hf, hef⟩) := ChartedSpace.mem_atlas_sum he + · obtain (⟨f', hf', he'f'⟩ | ⟨f', hf', he'f'⟩) := ChartedSpace.mem_atlas_sum he' + · rw [hef, he'f', f.lift_openEmbedding_trans f' IsOpenEmbedding.inl] + exact hM.compatible hf hf' + · rw [hef, he'f'] + apply ContDiffGroupoid.mem_of_source_eq_empty + ext x + exact ⟨fun ⟨hx₁, hx₂⟩ ↦ by simp_all [hx₂], fun hx ↦ hx.elim⟩ + · -- Analogous argument to the first case: is there a way to deduplicate? + obtain (⟨f', hf', he'f'⟩ | ⟨f', hf', he'f'⟩) := ChartedSpace.mem_atlas_sum he' + · rw [hef, he'f'] + apply ContDiffGroupoid.mem_of_source_eq_empty + ext x + exact ⟨fun ⟨hx₁, hx₂⟩ ↦ by simp_all [hx₂], fun hx ↦ hx.elim⟩ + · rw [hef, he'f', f.lift_openEmbedding_trans f' IsOpenEmbedding.inr] + exact hM'.compatible hf hf' + +end DisjointUnion + end IsManifold theorem PartialHomeomorph.isManifold_singleton @@ -1662,7 +1715,6 @@ variable (M) in -- is empty if the base manifold is empty /-- The tangent bundle to a manifold, as a Sigma type. Defined in terms of `Bundle.TotalSpace` to be able to put a suitable topology on it. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was nolint has_nonempty_instance abbrev TangentBundle := Bundle.TotalSpace E (TangentSpace I : M → Type _) end TangentSpace @@ -1676,4 +1728,4 @@ instance : PathConnectedSpace (TangentSpace I x) := inferInstanceAs (PathConnect end Real -set_option linter.style.longFile 1700 +set_option linter.style.longFile 1900 diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index ef5099d8835ea..9b52a1a42507b 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -95,7 +95,6 @@ subordinate to `U`, see `SmoothBumpCovering.exists_isSubordinate`. This covering can be used, e.g., to construct a partition of unity and to prove the weak Whitney embedding theorem. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure SmoothBumpCovering [FiniteDimensional ℝ E] (s : Set M := univ) where /-- The center point of each bump in the smooth covering. -/ c : ι → M diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index fa5a1a2036d04..fc9c216accc73 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -31,7 +31,8 @@ variable (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] -- `V` vector bundle [∀ x : M, TopologicalSpace (V x)] [FiberBundle F V] -/-- Bundled `n` times continuously differentiable sections of a vector bundle. -/ +/-- Bundled `n` times continuously differentiable sections of a vector bundle. +Denoted as `Cₛ^n⟮I; F, V⟯` within the `Manifold` namespace. -/ structure ContMDiffSection where /-- the underlying function of this section -/ protected toFun : ∀ x, V x diff --git a/Mathlib/Geometry/Manifold/VectorField.lean b/Mathlib/Geometry/Manifold/VectorField.lean new file mode 100644 index 0000000000000..cd573eaaa75b7 --- /dev/null +++ b/Mathlib/Geometry/Manifold/VectorField.lean @@ -0,0 +1,341 @@ +/- +Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Analysis.Calculus.VectorField +import Mathlib.Geometry.Manifold.ContMDiffMFDeriv +import Mathlib.Geometry.Manifold.MFDeriv.NormedSpace +import Mathlib.Geometry.Manifold.VectorBundle.MDifferentiable + +/-! +# Vector fields in manifolds + +We study functions of the form `V : Π (x : M) → TangentSpace I x` on a manifold, i.e., +vector fields. + +We define the pullback of a vector field under a map, as +`VectorField.mpullback I I' f V x := (mfderiv I I' f x).inverse (V (f x))` +(together with the same notion within a set). Note that the pullback uses the junk-value pattern: +if the derivative of the map is not invertible, then pullback is given the junk value zero. + +We start developing API around this notion. + +All these are given in the `VectorField` namespace because pullbacks, Lie brackets, and so on, +are notions that make sense in a variety of contexts. We also prefix the notions with `m` to +distinguish the manifold notions from the vector space notions. + +For notions that come naturally in other namespaces for dot notation, we specify `vectorField` in +the name to lift ambiguities. For instance, the fact that the Lie bracket of two smooth vector +fields is smooth will be `ContMDiffAt.mlieBracket_vectorField`. + +Note that a smoothness assumption for a vector field is written by seeing the vector field as +a function from `M` to its tangent bundle through a coercion, as in: +`MDifferentiableWithinAt I I.tangent (fun y ↦ (V y : TangentBundle I M)) s x`. +-/ + +open Set Function Filter +open scoped Topology Manifold ContDiff + +noncomputable section + +/- We work in the `VectorField` namespace because pullbacks, Lie brackets, and so on, are notions +that make sense in a variety of contexts. We also prefix the notions with `m` to distinguish the +manifold notions from the vector spaces notions. For instance, the Lie bracket of two vector +fields in a manifold is denoted with `VectorField.mlieBracket I V W x`, where `I` is the relevant +model with corners, `V W : Π (x : M), TangentSpace I x` are the vector fields, and `x : M` is +the basepoint. +-/ + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {H : Type*} [TopologicalSpace H] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {I : ModelWithCorners 𝕜 E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] + {H' : Type*} [TopologicalSpace H'] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] + {I' : ModelWithCorners 𝕜 E' H'} + {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + {H'' : Type*} [TopologicalSpace H''] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] + {I'' : ModelWithCorners 𝕜 E'' H''} + {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + {f : M → M'} {s t : Set M} {x x₀ : M} + +instance {n : ℕ} [n.AtLeastTwo] [IsManifold I (minSmoothness 𝕜 (ofNat(n))) M] : + IsManifold I (ofNat(n)) M := + IsManifold.of_le (n := minSmoothness 𝕜 n) le_minSmoothness + +instance [IsManifold I (minSmoothness 𝕜 1) M] : + IsManifold I 1 M := + IsManifold.of_le (n := minSmoothness 𝕜 1) le_minSmoothness + +instance [IsManifold I (minSmoothness 𝕜 3) M] : + IsManifold I (minSmoothness 𝕜 2) M := + IsManifold.of_le (n := minSmoothness 𝕜 3) (minSmoothness_monotone (by norm_cast)) + +instance [IsManifold I (minSmoothness 𝕜 2) M] : + IsManifold I (minSmoothness 𝕜 1) M := + IsManifold.of_le (n := minSmoothness 𝕜 2) (minSmoothness_monotone (by norm_cast)) + +namespace VectorField + +section Pullback + +/-! ### Pullback of vector fields in manifolds -/ + +open ContinuousLinearMap + +variable {V W V₁ W₁ : Π (x : M'), TangentSpace I' x} +variable {c : 𝕜} {m n : WithTop ℕ∞} {t : Set M'} {y₀ : M'} + +variable (I I') in +/-- The pullback of a vector field under a map between manifolds, within a set `s`. If the +derivative of the map within `s` is not invertible, then pullback is given the junk value zero.-/ +def mpullbackWithin (f : M → M') (V : Π (x : M'), TangentSpace I' x) (s : Set M) (x : M) : + TangentSpace I x := + (mfderivWithin I I' f s x).inverse (V (f x)) + +variable (I I') in +/-- The pullback of a vector field under a map between manifolds. If the derivative of the map is +not invertible, then pullback is given the junk value zero. -/ +def mpullback (f : M → M') (V : Π (x : M'), TangentSpace I' x) (x : M) : + TangentSpace I x := + (mfderiv I I' f x).inverse (V (f x)) + +lemma mpullbackWithin_apply : + mpullbackWithin I I' f V s x = (mfderivWithin I I' f s x).inverse (V (f x)) := rfl + +lemma mpullbackWithin_smul_apply : + mpullbackWithin I I' f (c • V) s x = c • mpullbackWithin I I' f V s x := by + simp [mpullbackWithin_apply] + +lemma mpullbackWithin_smul : + mpullbackWithin I I' f (c • V) s = c • mpullbackWithin I I' f V s := by + ext x + simp [mpullbackWithin_apply] + +lemma mpullbackWithin_add_apply : + mpullbackWithin I I' f (V + V₁) s x = + mpullbackWithin I I' f V s x + mpullbackWithin I I' f V₁ s x := by + simp [mpullbackWithin_apply] + +lemma mpullbackWithin_add : + mpullbackWithin I I' f (V + V₁) s = + mpullbackWithin I I' f V s + mpullbackWithin I I' f V₁ s := by + ext x + simp [mpullbackWithin_apply] + +lemma mpullbackWithin_neg_apply : + mpullbackWithin I I' f (-V) s x = - mpullbackWithin I I' f V s x := by + simp [mpullbackWithin_apply] + +lemma mpullbackWithin_neg : + mpullbackWithin I I' f (-V) s = - mpullbackWithin I I' f V s := by + ext x + simp [mpullbackWithin_apply] + +lemma mpullbackWithin_id {V : Π (x : M), TangentSpace I x} (h : UniqueMDiffWithinAt I s x) : + mpullbackWithin I I id V s x = V x := by + simp [mpullbackWithin_apply, mfderivWithin_id h] + +lemma mpullback_apply : + mpullback I I' f V x = (mfderiv I I' f x).inverse (V (f x)) := rfl + +lemma mpullback_smul_apply : + mpullback I I' f (c • V) x = c • mpullback I I' f V x := by + simp [mpullback] + +lemma mpullback_smul : + mpullback I I' f (c • V) = c • mpullback I I' f V := by + ext x + simp [mpullback_apply] + +lemma mpullback_add_apply : + mpullback I I' f (V + V₁) x = mpullback I I' f V x + mpullback I I' f V₁ x := by + simp [mpullback_apply] + +lemma mpullback_add : + mpullback I I' f (V + V₁) = mpullback I I' f V + mpullback I I' f V₁ := by + ext x + simp [mpullback_apply] + +lemma mpullback_neg_apply : + mpullback I I' f (-V) x = - mpullback I I' f V x := by + simp [mpullback_apply] + +lemma mpullback_neg : + mpullback I I' f (-V) = - mpullback I I' f V := by + ext x + simp [mpullback_apply] + +@[simp] lemma mpullbackWithin_univ : mpullbackWithin I I' f V univ = mpullback I I' f V := by + ext x + simp [mpullback_apply, mpullbackWithin_apply] + +lemma mpullbackWithin_eq_pullbackWithin {f : E → E'} {V : E' → E'} {s : Set E} : + mpullbackWithin 𝓘(𝕜, E) 𝓘(𝕜, E') f V s = pullbackWithin 𝕜 f V s := by + ext x + simp only [mpullbackWithin, mfderivWithin_eq_fderivWithin, pullbackWithin] + rfl + +lemma mpullback_eq_pullback {f : E → E'} {V : E' → E'} : + mpullback 𝓘(𝕜, E) 𝓘(𝕜, E') f V = pullback 𝕜 f V := by + simp only [← mpullbackWithin_univ, ← pullbackWithin_univ, mpullbackWithin_eq_pullbackWithin] + +@[simp] lemma mpullback_id {V : Π (x : M), TangentSpace I x} : mpullback I I id V = V := by + ext x + simp [mpullback] + +/-! ### Regularity of pullback of vector fields + +In this paragraph, we assume that the model space is complete, to ensure that the set of invertible +linear maps is open and that inversion is a smooth map there. Otherwise, the pullback of vector +fields could behave wildly, even at points where the derivative of the map is invertible. +-/ + +section MDifferentiability + +variable [IsManifold I 2 M] [IsManifold I' 2 M'] [CompleteSpace E] + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. Version within a set at a point. -/ +protected lemma _root_.MDifferentiableWithinAt.mpullbackWithin_vectorField_inter + (hV : MDifferentiableWithinAt I' I'.tangent + (fun (y : M') ↦ (V y : TangentBundle I' M')) t (f x₀)) + (hf : ContMDiffWithinAt I I' n f s x₀) (hf' : (mfderivWithin I I' f s x₀).IsInvertible) + (hx₀ : x₀ ∈ s) (hs : UniqueMDiffOn I s) (hmn : 2 ≤ n) : + MDifferentiableWithinAt I I.tangent + (fun (y : M) ↦ (mpullbackWithin I I' f V s y : TangentBundle I M)) (s ∩ f ⁻¹' t) x₀ := by + /- We want to apply the theorem `MDifferentiableWithinAt.clm_apply_of_inCoordinates`, + stating that applying linear maps to vector fields gives a smooth result when the linear map and + the vector field are smooth. This theorem is general, we will apply it to + `b₁ = f`, `b₂ = id`, `v = V ∘ f`, `ϕ = fun x ↦ (mfderivWithin I I' f s x).inverse`-/ + let b₁ := f + let b₂ : M → M := id + let v : Π (x : M), TangentSpace I' (f x) := V ∘ f + let ϕ : Π (x : M), TangentSpace I' (f x) →L[𝕜] TangentSpace I x := + fun x ↦ (mfderivWithin I I' f s x).inverse + have hv : MDifferentiableWithinAt I I'.tangent + (fun x ↦ (v x : TangentBundle I' M')) (s ∩ f ⁻¹' t) x₀ := by + apply hV.comp x₀ ((hf.mdifferentiableWithinAt (one_le_two.trans hmn)).mono inter_subset_left) + exact MapsTo.mono_left (mapsTo_preimage _ _) inter_subset_right + /- The only nontrivial fact, from which the conclusion follows, is + that `ϕ` depends smoothly on `x`. -/ + suffices hϕ : MDifferentiableWithinAt I 𝓘(𝕜, E' →L[𝕜] E) + (fun (x : M) ↦ ContinuousLinearMap.inCoordinates + E' (TangentSpace I' (M := M')) E (TangentSpace I (M := M)) + (b₁ x₀) (b₁ x) (b₂ x₀) (b₂ x) (ϕ x)) s x₀ from + MDifferentiableWithinAt.clm_apply_of_inCoordinates (hϕ.mono inter_subset_left) + hv mdifferentiableWithinAt_id + /- To prove that `ϕ` depends smoothly on `x`, we use that the derivative depends smoothly on `x` + (this is `ContMDiffWithinAt.mfderivWithin_const`), and that taking the inverse is a smooth + operation at an invertible map. -/ + -- the derivative in coordinates depends smoothly on the point + have : MDifferentiableWithinAt I 𝓘(𝕜, E →L[𝕜] E') + (fun (x : M) ↦ ContinuousLinearMap.inCoordinates + E (TangentSpace I (M := M)) E' (TangentSpace I' (M := M')) + x₀ x (f x₀) (f x) (mfderivWithin I I' f s x)) s x₀ := + ((hf.of_le hmn).mfderivWithin_const le_rfl hx₀ hs).mdifferentiableWithinAt le_rfl + -- therefore, its inverse in coordinates also depends smoothly on the point + have : MDifferentiableWithinAt I 𝓘(𝕜, E' →L[𝕜] E) + (ContinuousLinearMap.inverse ∘ (fun (x : M) ↦ ContinuousLinearMap.inCoordinates + E (TangentSpace I (M := M)) E' (TangentSpace I' (M := M')) + x₀ x (f x₀) (f x) (mfderivWithin I I' f s x))) s x₀ := by + apply MDifferentiableAt.comp_mdifferentiableWithinAt _ _ this + apply ContMDiffAt.mdifferentiableAt _ le_rfl + apply ContDiffAt.contMDiffAt + apply IsInvertible.contDiffAt_map_inverse + rw [inCoordinates_eq (FiberBundle.mem_baseSet_trivializationAt' x₀) + (FiberBundle.mem_baseSet_trivializationAt' (f x₀))] + exact isInvertible_equiv.comp (hf'.comp isInvertible_equiv) + -- the inverse in coordinates coincides with the in-coordinate version of the inverse, + -- therefore the previous point gives the conclusion + apply this.congr_of_eventuallyEq_of_mem _ hx₀ + have A : (trivializationAt E (TangentSpace I) x₀).baseSet ∈ 𝓝[s] x₀ := by + apply nhdsWithin_le_nhds + apply (trivializationAt _ _ _).open_baseSet.mem_nhds + exact FiberBundle.mem_baseSet_trivializationAt' _ + have B : f ⁻¹' (trivializationAt E' (TangentSpace I') (f x₀)).baseSet ∈ 𝓝[s] x₀ := by + apply hf.continuousWithinAt.preimage_mem_nhdsWithin + apply (trivializationAt _ _ _).open_baseSet.mem_nhds + exact FiberBundle.mem_baseSet_trivializationAt' _ + filter_upwards [A, B] with x hx h'x + simp only [Function.comp_apply] + rw [inCoordinates_eq hx h'x, inCoordinates_eq h'x (by exact hx)] + simp only [inverse_equiv_comp, inverse_comp_equiv, ContinuousLinearEquiv.symm_symm, ϕ] + rfl + +lemma _root_.MDifferentiableWithinAt.mpullbackWithin_vectorField_inter_of_eq + (hV : MDifferentiableWithinAt I' I'.tangent + (fun (y : M') ↦ (V y : TangentBundle I' M')) t y₀) + (hf : ContMDiffWithinAt I I' n f s x₀) (hf' : (mfderivWithin I I' f s x₀).IsInvertible) + (hx₀ : x₀ ∈ s) (hs : UniqueMDiffOn I s) (hmn : 2 ≤ n) (h : y₀ = f x₀) : + MDifferentiableWithinAt I I.tangent + (fun (y : M) ↦ (mpullbackWithin I I' f V s y : TangentBundle I M)) (s ∩ f⁻¹' t) x₀ := by + subst h + exact hV.mpullbackWithin_vectorField_inter hf hf' hx₀ hs hmn + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. Version on a set. -/ +protected lemma _root_.MDifferentiableOn.mpullbackWithin_vectorField_inter + (hV : MDifferentiableOn I' I'.tangent (fun (y : M') ↦ (V y : TangentBundle I' M')) t) + (hf : ContMDiffOn I I' n f s) (hf' : ∀ x ∈ s ∩ f ⁻¹' t, (mfderivWithin I I' f s x).IsInvertible) + (hs : UniqueMDiffOn I s) (hmn : 2 ≤ n) : + MDifferentiableOn I I.tangent + (fun (y : M) ↦ (mpullbackWithin I I' f V s y : TangentBundle I M)) (s ∩ f ⁻¹' t) := + fun _ hx₀ ↦ MDifferentiableWithinAt.mpullbackWithin_vectorField_inter + (hV _ hx₀.2) (hf _ hx₀.1) (hf' _ hx₀) hx₀.1 hs hmn + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. Version within a set at a point, but with full pullback. -/ +protected lemma _root_.MDifferentiableWithinAt.mpullback_vectorField_preimage + (hV : MDifferentiableWithinAt I' I'.tangent + (fun (y : M') ↦ (V y : TangentBundle I' M')) t (f x₀)) + (hf : ContMDiffAt I I' n f x₀) (hf' : (mfderiv I I' f x₀).IsInvertible) (hmn : 2 ≤ n) : + MDifferentiableWithinAt I I.tangent + (fun (y : M) ↦ (mpullback I I' f V y : TangentBundle I M)) (f ⁻¹' t) x₀ := by + simp only [← contMDiffWithinAt_univ, ← mfderivWithin_univ, ← mpullbackWithin_univ] at hV hf hf' ⊢ + simpa using hV.mpullbackWithin_vectorField_inter hf hf' (mem_univ _) uniqueMDiffOn_univ hmn + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. Version within a set at a point, but with full pullback. -/ +protected lemma _root_.MDifferentiableWithinAt.mpullback_vectorField_preimage_of_eq + (hV : MDifferentiableWithinAt I' I'.tangent (fun (y : M') ↦ (V y : TangentBundle I' M')) t y₀) + (hf : ContMDiffAt I I' n f x₀) (hf' : (mfderiv I I' f x₀).IsInvertible) (hmn : 2 ≤ n) + (hy₀ : y₀ = f x₀) : + MDifferentiableWithinAt I I.tangent + (fun (y : M) ↦ (mpullback I I' f V y : TangentBundle I M)) (f ⁻¹' t) x₀ := by + subst hy₀ + exact hV.mpullback_vectorField_preimage hf hf' hmn + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. Version on a set, but with full pullback -/ +protected lemma _root_.MDifferentiableOn.mpullback_vectorField_preimage + (hV : MDifferentiableOn I' I'.tangent (fun (y : M') ↦ (V y : TangentBundle I' M')) t) + (hf : ContMDiff I I' n f) (hf' : ∀ x ∈ f ⁻¹' t, (mfderiv I I' f x).IsInvertible) + (hmn : 2 ≤ n) : + MDifferentiableOn I I.tangent + (fun (y : M) ↦ (mpullback I I' f V y : TangentBundle I M)) (f ⁻¹' t) := + fun x₀ hx₀ ↦ MDifferentiableWithinAt.mpullback_vectorField_preimage + (hV _ hx₀) (hf x₀) (hf' _ hx₀) hmn + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. Version at a point. -/ +protected lemma _root_.MDifferentiableAt.mpullback_vectorField + (hV : MDifferentiableAt I' I'.tangent (fun (y : M') ↦ (V y : TangentBundle I' M')) (f x₀)) + (hf : ContMDiffAt I I' n f x₀) (hf' : (mfderiv I I' f x₀).IsInvertible) (hmn : 2 ≤ n) : + MDifferentiableAt I I.tangent + (fun (y : M) ↦ (mpullback I I' f V y : TangentBundle I M)) x₀ := by + simpa using MDifferentiableWithinAt.mpullback_vectorField_preimage hV hf hf' hmn + +/-- The pullback of a differentiable vector field by a `C^n` function with `2 ≤ n` is +differentiable. -/ +protected lemma _root_.MDifferentiable.mpullback_vectorField + (hV : MDifferentiable I' I'.tangent (fun (y : M') ↦ (V y : TangentBundle I' M'))) + (hf : ContMDiff I I' n f) (hf' : ∀ x, (mfderiv I I' f x).IsInvertible) (hmn : 2 ≤ n) : + MDifferentiable I I.tangent (fun (y : M) ↦ (mpullback I I' f V y : TangentBundle I M)) := + fun x ↦ MDifferentiableAt.mpullback_vectorField (hV (f x)) (hf x) (hf' x) hmn + +end MDifferentiability + +end Pullback + +end VectorField diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index e05dffc7df37d..a4c26869dcdfe 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -227,9 +227,10 @@ theorem coequalizer_π_stalk_isLocalHom (x : Y) : constructor rintro a ha rcases TopCat.Presheaf.germ_exist _ _ a with ⟨U, hU, s, rfl⟩ - rw [← CommRingCat.forget_map_apply, PresheafedSpace.stalkMap_germ_apply - (coequalizer.π (C := SheafedSpace _) f.toShHom g.toShHom) U _ hU] at ha - rw [CommRingCat.forget_map_apply] + rw [← CommRingCat.forget_map_apply, forget_map_eq_coe, + PresheafedSpace.stalkMap_germ_apply + (coequalizer.π (C := SheafedSpace _) f.toShHom g.toShHom) U _ hU] at ha + rw [coe_toHasForget_instFunLike] let V := imageBasicOpen f g U s have hV : (coequalizer.π f.toShHom g.toShHom).base ⁻¹' ((coequalizer.π f.toShHom g.toShHom).base '' V.1) = V.1 := diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean index e96b86ffdeceb..5120aff056674 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/ResidueField.lean @@ -113,7 +113,7 @@ lemma residueFieldMap_comp {Z : LocallyRingedSpace.{u}} (g : Y ⟶ Z) (x : X) : simp only [comp_toShHom, SheafedSpace.comp_base, Function.comp_apply, residueFieldMap, CommRingCat.hom_comp, TopCat.comp_app] simp_rw [stalkMap_comp] - apply IsLocalRing.ResidueField.map_comp + apply IsLocalRing.ResidueField.map_comp (Hom.stalkMap g (f.base x)).hom (Hom.stalkMap f x).hom @[reassoc] lemma evaluation_naturality {V : Opens Y} (x : (Opens.map f.base).obj V) : diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 867c314932f6a..912cb224dd41a 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -90,8 +90,6 @@ such that We can then glue the spaces `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subspaces of the glued space. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure GlueData extends CategoryTheory.GlueData (PresheafedSpace.{u, v, v} C) where f_open : ∀ i j, IsOpenImmersion (f i j) @@ -194,14 +192,16 @@ theorem snd_invApp_t_app' (i j k : D.J) (U : Opens (pullback (D.f i j) (D.f i k) rw [this, Set.image_comp, Set.image_comp, Set.preimage_image_eq] swap · refine Function.HasLeftInverse.injective ⟨(D.t i k).base, fun x => ?_⟩ - rw [← comp_apply, ← comp_base, D.t_inv, id_base, id_apply] + rw [← CategoryTheory.comp_apply, ← comp_base, D.t_inv, id_base, CategoryTheory.id_apply] refine congr_arg (_ '' ·) ?_ refine congr_fun ?_ _ refine Set.image_eq_preimage_of_inverse ?_ ?_ · intro x - rw [← comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, id_apply] + rw [← CategoryTheory.comp_apply, ← comp_base, IsIso.inv_hom_id, id_base, + CategoryTheory.id_apply] · intro x - rw [← comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, id_apply] + rw [← CategoryTheory.comp_apply, ← comp_base, IsIso.hom_inv_id, id_base, + CategoryTheory.id_apply] · rw [← IsIso.eq_inv_comp, IsOpenImmersion.inv_invApp, Category.assoc, (D.t' k i j).c.naturality_assoc] simp_rw [← Category.assoc] @@ -392,8 +392,7 @@ def ιInvApp {i : D.J} (U : Opens (D.U i).carrier) : rw [← (D.V (j, k)).presheaf.map_comp] erw [← (D.V (j, k)).presheaf.map_comp] repeat rw [← (D.V (j, k)).presheaf.map_comp] - -- Porting note: was just `congr` - exact congr_arg ((D.V (j, k)).presheaf.map ·) rfl } } + rfl } } /-- `ιInvApp` is the left inverse of `D.ι i` on `U`. -/ theorem ιInvApp_π {i : D.J} (U : Opens (D.U i).carrier) : @@ -552,8 +551,6 @@ such that We can then glue the spaces `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subspaces of the glued space. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure GlueData extends CategoryTheory.GlueData (SheafedSpace.{u, v, v} C) where f_open : ∀ i j, SheafedSpace.IsOpenImmersion (f i j) @@ -625,8 +622,6 @@ such that We can then glue the spaces `U i` together by identifying `V i j` with `V j i`, such that the `U i`'s are open subspaces of the glued space. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure GlueData extends CategoryTheory.GlueData LocallyRingedSpace where f_open : ∀ i j, LocallyRingedSpace.IsOpenImmersion (f i j) diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index afafd4588a723..259e4a2312471 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -522,8 +522,6 @@ theorem equiv_fst_eq_one_of_mem_of_one_mem {g : G} (h1 : 1 ∈ S) (hg : g ∈ T) ext rw [equiv_fst_eq_mul_inv, equiv_snd_eq_self_of_mem_of_one_mem _ h1 hg, mul_inv_cancel] --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem equiv_mul_right (g : G) (k : K) : hSK.equiv (g * k) = ((hSK.equiv g).fst, (hSK.equiv g).snd * k) := by have : (hSK.equiv (g * k)).fst = (hSK.equiv g).fst := @@ -537,8 +535,6 @@ theorem equiv_mul_right_of_mem {g k : G} (h : k ∈ K) : hSK.equiv (g * k) = ((hSK.equiv g).fst, (hSK.equiv g).snd * ⟨k, h⟩) := equiv_mul_right _ g ⟨k, h⟩ --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[simp, nolint simpNF] theorem equiv_mul_left (h : H) (g : G) : hHT.equiv (h * g) = (h * (hHT.equiv g).fst, (hHT.equiv g).snd) := by have : (hHT.equiv (h * g)).2 = (hHT.equiv g).2 := hHT.equiv_snd_eq_iff_rightCosetEquivalence.2 ?_ diff --git a/Mathlib/GroupTheory/Congruence/Defs.lean b/Mathlib/GroupTheory/Congruence/Defs.lean index f3b58fe980782..43eb71a59dbeb 100644 --- a/Mathlib/GroupTheory/Congruence/Defs.lean +++ b/Mathlib/GroupTheory/Congruence/Defs.lean @@ -6,6 +6,7 @@ Authors: Amelia Livingston import Mathlib.Algebra.Group.InjSurj import Mathlib.Algebra.Group.Units.Defs import Mathlib.Data.Setoid.Basic +import Mathlib.Tactic.FastInstance /-! # Congruence relations @@ -638,48 +639,27 @@ instance {M : Type*} [Monoid M] (c : Con M) : Pow c.Quotient ℕ where /-- The quotient of a semigroup by a congruence relation is a semigroup. -/ @[to_additive "The quotient of an `AddSemigroup` by an additive congruence relation is an `AddSemigroup`."] -instance semigroup {M : Type*} [Semigroup M] (c : Con M) : Semigroup c.Quotient := - { (Function.Surjective.semigroup _ - Quotient.mk''_surjective fun _ _ => rfl : - Semigroup c.Quotient) with - /- The `toMul` field is given explicitly for performance reasons. - This avoids any need to unfold `Function.Surjective.semigroup` when the type checker is checking - that instance diagrams commute -/ - toMul := Con.hasMul _ } +instance semigroup {M : Type*} [Semigroup M] (c : Con M) : Semigroup c.Quotient := fast_instance% + Function.Surjective.semigroup _ Quotient.mk''_surjective fun _ _ => rfl /-- The quotient of a commutative semigroup by a congruence relation is a semigroup. -/ @[to_additive "The quotient of an `AddCommSemigroup` by an additive congruence relation is an `AddCommSemigroup`."] instance commSemigroup {M : Type*} [CommSemigroup M] (c : Con M) : CommSemigroup c.Quotient := - { (Function.Surjective.commSemigroup _ Quotient.mk''_surjective fun _ _ => rfl : - CommSemigroup c.Quotient) with - /- The `toSemigroup` field is given explicitly for performance reasons. - This avoids any need to unfold `Function.Surjective.commSemigroup` when the type checker is - checking that instance diagrams commute -/ - toSemigroup := Con.semigroup _ } + Function.Surjective.commSemigroup _ Quotient.mk''_surjective fun _ _ => rfl /-- The quotient of a monoid by a congruence relation is a monoid. -/ @[to_additive "The quotient of an `AddMonoid` by an additive congruence relation is an `AddMonoid`."] -instance monoid {M : Type*} [Monoid M] (c : Con M) : Monoid c.Quotient := - { (Function.Surjective.monoid _ Quotient.mk''_surjective rfl - (fun _ _ => rfl) fun _ _ => rfl : Monoid c.Quotient) with - /- The `toSemigroup` and `toOne` fields are given explicitly for performance reasons. - This avoids any need to unfold `Function.Surjective.monoid` when the type checker is - checking that instance diagrams commute -/ - toSemigroup := Con.semigroup _ - toOne := Con.one _ } +instance monoid {M : Type*} [Monoid M] (c : Con M) : Monoid c.Quotient := fast_instance% + Function.Surjective.monoid _ Quotient.mk''_surjective rfl (fun _ _ => rfl) fun _ _ => rfl /-- The quotient of a `CommMonoid` by a congruence relation is a `CommMonoid`. -/ @[to_additive "The quotient of an `AddCommMonoid` by an additive congruence relation is an `AddCommMonoid`."] -instance commMonoid {M : Type*} [CommMonoid M] (c : Con M) : CommMonoid c.Quotient := - { (Function.Surjective.commMonoid _ Quotient.mk''_surjective rfl - (fun _ _ => rfl) fun _ _ => rfl : CommMonoid c.Quotient) with - /- The `toMonoid` field is given explicitly for performance reasons. - This avoids any need to unfold `Function.Surjective.commMonoid` when the type checker is - checking that instance diagrams commute -/ - toMonoid := Con.monoid _ } +instance commMonoid {M : Type*} [CommMonoid M] (c : Con M) : CommMonoid c.Quotient := fast_instance% + fast_instance% Function.Surjective.commMonoid _ Quotient.mk''_surjective rfl + (fun _ _ => rfl) fun _ _ => rfl /-- Sometimes, a group is defined as a quotient of a monoid by a congruence relation. Usually, the inverse operation is defined as `Setoid.map f _` for some `f`. @@ -754,24 +734,16 @@ instance zpowinst : Pow c.Quotient ℤ := /-- The quotient of a group by a congruence relation is a group. -/ @[to_additive "The quotient of an `AddGroup` by an additive congruence relation is an `AddGroup`."] -instance group : Group c.Quotient := - { (Function.Surjective.group Quotient.mk'' - Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ => rfl) - (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl : Group c.Quotient) with - toMonoid := Con.monoid _ - toInv := Con.hasInv _ - toDiv := Con.hasDiv _ } +instance group : Group c.Quotient := fast_instance% + Function.Surjective.group Quotient.mk'' Quotient.mk''_surjective + rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl /-- The quotient of a `CommGroup` by a congruence relation is a `CommGroup`. -/ @[to_additive "The quotient of an `AddCommGroup` by an additive congruence relation is an `AddCommGroup`."] -instance commGroup {M : Type*} [CommGroup M] (c : Con M) : CommGroup c.Quotient := - { (Function.Surjective.commGroup _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ => rfl) - (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) : CommGroup c.Quotient) with - /- The `toGroup` field is given explicitly for performance reasons. - This avoids any need to unfold `Function.Surjective.commGroup` when the type checker is - checking that instance diagrams commute -/ - toGroup := Con.group _ } +instance commGroup {M : Type*} [CommGroup M] (c : Con M) : CommGroup c.Quotient := fast_instance% + Function.Surjective.commGroup _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ => rfl) + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) end Groups diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 09d71fbf32674..a67918662b655 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -621,7 +621,6 @@ variable (M) /-- A `NeWord M i j` is a representation of a non-empty reduced words where the first letter comes from `M i` and the last letter comes from `M j`. It can be constructed from singletons and via concatenation, and thus provides a useful induction principle. -/ ---@[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): commented out inductive NeWord : ι → ι → Type _ | singleton : ∀ {i : ι} (x : M i), x ≠ 1 → NeWord i i | append : ∀ {i j k l} (_w₁ : NeWord i j) (_hne : j ≠ k) (_w₂ : NeWord k l), NeWord i l @@ -836,7 +835,7 @@ theorem lift_word_ping_pong {i j k} (w : NeWord H i j) (hk : j ≠ k) : · calc lift f (NeWord.append w₁ hne w₂).prod • X k = lift f w₁.prod • lift f w₂.prod • X k := by simp [MulAction.mul_smul] - _ ⊆ lift f w₁.prod • X _ := set_smul_subset_set_smul_iff.mpr (hIw₂ hk) + _ ⊆ lift f w₁.prod • X _ := smul_set_subset_smul_set_iff.mpr (hIw₂ hk) _ ⊆ X i := hIw₁ hne include hXnonempty hXdisj diff --git a/Mathlib/GroupTheory/Coxeter/Basic.lean b/Mathlib/GroupTheory/Coxeter/Basic.lean index c33678eadac13..f19ef6d87ef99 100644 --- a/Mathlib/GroupTheory/Coxeter/Basic.lean +++ b/Mathlib/GroupTheory/Coxeter/Basic.lean @@ -458,7 +458,7 @@ lemma listTake_alternatingWord (i j : B) (p k : ℕ) (h : k < 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] + simp only [take_zero, Even.zero, ↓reduceIte, alternatingWord] | succ k h' => have hk : k < 2 * p := by omega apply h' at hk diff --git a/Mathlib/GroupTheory/FiniteAbelian/Basic.lean b/Mathlib/GroupTheory/FiniteAbelian/Basic.lean index 38fa544799b25..45360f7df2277 100644 --- a/Mathlib/GroupTheory/FiniteAbelian/Basic.lean +++ b/Mathlib/GroupTheory/FiniteAbelian/Basic.lean @@ -5,7 +5,7 @@ Authors: Pierre-Alexandre Bazin -/ import Mathlib.Algebra.Module.PID import Mathlib.Algebra.Group.TypeTags.Finite -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientRing /-! # Structure of finite(ly generated) abelian groups diff --git a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean index 5fc605e18a73b..2343d928c711c 100644 --- a/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean +++ b/Mathlib/GroupTheory/FreeGroup/NielsenSchreier.lean @@ -61,7 +61,6 @@ open CategoryTheory CategoryTheory.ActionCategory CategoryTheory.SingleObj Quive /-- `IsFreeGroupoid.Generators G` is a type synonym for `G`. We think of this as the vertices of the generating quiver of `G` when `G` is free. We can't use `G` directly, since `G` already has a quiver instance from being a groupoid. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): @[nolint has_nonempty_instance] @[nolint unusedArguments] def IsFreeGroupoid.Generators (G) [Groupoid G] := G diff --git a/Mathlib/GroupTheory/GroupAction/Blocks.lean b/Mathlib/GroupTheory/GroupAction/Blocks.lean index 5655e4bb9cc96..7bc87f229f0a2 100644 --- a/Mathlib/GroupTheory/GroupAction/Blocks.lean +++ b/Mathlib/GroupTheory/GroupAction/Blocks.lean @@ -229,7 +229,7 @@ lemma IsBlock.subgroup {H : Subgroup G} (hB : IsBlock G B) : IsBlock H B := fun /-- A block of a group action is invariant iff it is fixed. -/ @[to_additive "A block of a group action is invariant iff it is fixed."] lemma isInvariantBlock_iff_isFixedBlock : IsInvariantBlock G B ↔ IsFixedBlock G B := - ⟨fun hB g ↦ (hB g).antisymm <| subset_set_smul_iff.2 <| hB _, IsFixedBlock.isInvariantBlock⟩ + ⟨fun hB g ↦ (hB g).antisymm <| subset_smul_set_iff.2 <| hB _, IsFixedBlock.isInvariantBlock⟩ /-- An invariant block of a group action is a fixed block. -/ @[to_additive "An invariant block of a group action is a fixed block."] diff --git a/Mathlib/GroupTheory/GroupAction/ConjAct.lean b/Mathlib/GroupTheory/GroupAction/ConjAct.lean index 970e9def54589..00274cf0874cf 100644 --- a/Mathlib/GroupTheory/GroupAction/ConjAct.lean +++ b/Mathlib/GroupTheory/GroupAction/ConjAct.lean @@ -1,15 +1,12 @@ /- -Copyright (c) 2021 . All rights reserved. +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.Algebra.Field.Defs -import Mathlib.Algebra.GroupWithZero.Action.Basic -import Mathlib.Algebra.Ring.Action.Basic +import Mathlib.Algebra.Group.Subgroup.ZPowers.Basic import Mathlib.Data.Fintype.Card import Mathlib.GroupTheory.GroupAction.Defs import Mathlib.GroupTheory.Subgroup.Centralizer -import Mathlib.Algebra.Group.Subgroup.ZPowers.Basic /-! # Conjugation action of a group on itself @@ -35,8 +32,11 @@ is that some theorems about the group actions will not apply when since this -/ +-- TODO +-- assert_not_exists GroupWithZero +assert_not_exists Ring -variable (α M G G₀ R K : Type*) +variable (α M G : Type*) /-- A type alias for a group `G`. `ConjAct G` acts on `G` by conjugation -/ def ConjAct : Type _ := @@ -46,14 +46,12 @@ namespace ConjAct open MulAction Subgroup -variable {M G G₀ R K} +variable {M G} instance [Group G] : Group (ConjAct G) := ‹Group G› instance [DivInvMonoid G] : DivInvMonoid (ConjAct G) := ‹DivInvMonoid G› -instance [GroupWithZero G] : GroupWithZero (ConjAct G) := ‹GroupWithZero G› - instance [Fintype G] : Fintype (ConjAct G) := ‹Fintype G› @[simp] @@ -167,57 +165,8 @@ instance unitsSMulCommClass' [SMul α M] [SMulCommClass M α M] [IsScalarTower end Monoid -section Semiring - -variable [Semiring R] - -instance unitsMulSemiringAction : MulSemiringAction (ConjAct Rˣ) R := - { ConjAct.unitsMulDistribMulAction with - smul_zero := by simp [units_smul_def] - smul_add := by simp [units_smul_def, mul_add, add_mul] } - -end Semiring - end Units -section GroupWithZero - -variable [GroupWithZero G₀] - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it -theorem ofConjAct_zero : ofConjAct (0 : ConjAct G₀) = 0 := - rfl - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it -theorem toConjAct_zero : toConjAct (0 : G₀) = 0 := - rfl - -instance mulAction₀ : MulAction (ConjAct G₀) G₀ where - one_smul := by simp [smul_def] - mul_smul := by simp [smul_def, mul_assoc] - -instance smulCommClass₀ [SMul α G₀] [SMulCommClass α G₀ G₀] [IsScalarTower α G₀ G₀] : - SMulCommClass α (ConjAct G₀) G₀ where - smul_comm a ug g := by rw [smul_def, smul_def, mul_smul_comm, smul_mul_assoc] - -instance smulCommClass₀' [SMul α G₀] [SMulCommClass G₀ α G₀] [IsScalarTower α G₀ G₀] : - SMulCommClass (ConjAct G₀) α G₀ := - haveI := SMulCommClass.symm G₀ α G₀ - SMulCommClass.symm _ _ _ - -end GroupWithZero - -section DivisionRing - -variable [DivisionRing K] - -instance distribMulAction₀ : DistribMulAction (ConjAct K) K := - { ConjAct.mulAction₀ with - smul_zero := by simp [smul_def] - smul_add := by simp [smul_def, mul_add, add_mul] } - -end DivisionRing - variable [Group G] -- todo: this file is not in good order; I will refactor this after the PR diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index d2cba75e644f6..fe04a1f0932df 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -5,8 +5,8 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Group.Subgroup.Defs -import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.Group.Submonoid.MulAction /-! # Definition of `orbit`, `fixedPoints` and `stabilizer` diff --git a/Mathlib/GroupTheory/GroupAction/FixedPoints.lean b/Mathlib/GroupTheory/GroupAction/FixedPoints.lean index ec44881e43a15..2a40a937b2799 100644 --- a/Mathlib/GroupTheory/GroupAction/FixedPoints.lean +++ b/Mathlib/GroupTheory/GroupAction/FixedPoints.lean @@ -149,7 +149,7 @@ theorem set_mem_fixedBy_of_subset_fixedBy {s : Set α} {g : G} (s_ss_fixedBy : s theorem smul_subset_of_set_mem_fixedBy {s t : Set α} {g : G} (t_ss_s : t ⊆ s) (s_in_fixedBy : s ∈ fixedBy (Set α) g) : g • t ⊆ s := - (Set.set_smul_subset_set_smul_iff.mpr t_ss_s).trans s_in_fixedBy.subset + (Set.smul_set_subset_smul_set_iff.mpr t_ss_s).trans s_in_fixedBy.subset /-! If a set `s : Set α` is a superset of `(MulAction.fixedBy α g)ᶜ` (resp. `(AddAction.fixedBy α g)ᶜ`), diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index f10d91924ee1c..c958a472e60c7 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -75,8 +75,6 @@ structure AddActionHom {M N : Type*} (φ: M → N) (X : Type*) [VAdd M X] (Y : T /-- Equivariant functions : When `φ : M → N` is a function, and types `X` and `Y` are endowed with actions of `M` and `N`, a function `f : X → Y` is `φ`-equivariant if `f (m • x) = (φ m) • (f x)`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] @[to_additive] structure MulActionHom where /-- The underlying function. -/ @@ -591,8 +589,6 @@ variable (T : Type*) [Semiring T] [MulSemiringAction P T] -- variable [AddMonoid N'] [DistribMulAction S N'] /-- Equivariant ring homomorphisms. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure MulSemiringActionHom extends R →ₑ+[φ] S, R →+* S /- diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 70d089f4ff23f..c1d5b153c65e4 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -6,6 +6,7 @@ Authors: Amelia Livingston import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.GroupTheory.Congruence.Hom import Mathlib.GroupTheory.OreLocalization.Basic +import Mathlib.Algebra.Group.Submonoid.Operations /-! # Localizations of commutative monoids @@ -82,8 +83,6 @@ variable {M : Type*} [AddCommMonoid M] (S : AddSubmonoid M) (N : Type*) [AddComm /-- The type of AddMonoid homomorphisms satisfying the characteristic predicate: if `f : M →+ N` satisfies this predicate, then `N` is isomorphic to the localization of `M` at `S`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure LocalizationMap extends AddMonoidHom M N where map_add_units' : ∀ y : S, IsAddUnit (toFun y) surj' : ∀ z : N, ∃ x : M × S, z + toFun x.2 = toFun x.1 @@ -107,8 +106,6 @@ namespace Submonoid /-- The type of monoid homomorphisms satisfying the characteristic predicate: if `f : M →* N` satisfies this predicate, then `N` is isomorphic to the localization of `M` at `S`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure LocalizationMap extends MonoidHom M N where map_units' : ∀ y : S, IsUnit (toFun y) surj' : ∀ z : N, ∃ x : M × S, z * toFun x.2 = toFun x.1 diff --git a/Mathlib/GroupTheory/MonoidLocalization/MonoidWithZero.lean b/Mathlib/GroupTheory/MonoidLocalization/MonoidWithZero.lean index ac8015e1578b9..f7e0e187cbf70 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/MonoidWithZero.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/MonoidWithZero.lean @@ -53,8 +53,6 @@ theorem LocalizationMap.subsingleton (f : Submonoid.LocalizationMap S N) (h : 0 /-- The type of homomorphisms between monoids with zero satisfying the characteristic predicate: if `f : M →*₀ N` satisfies this predicate, then `N` is isomorphic to the localization of `M` at `S`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure LocalizationWithZeroMap extends LocalizationMap S N where map_zero' : toFun 0 = 0 diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index c89a47cd28aa4..4a207ce12d1ef 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -10,6 +10,7 @@ import Mathlib.Algebra.Order.Group.Action import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Dynamics.PeriodicPts.Lemmas import Mathlib.GroupTheory.Index +import Mathlib.NumberTheory.Divisors import Mathlib.Order.Interval.Set.Infinite import Mathlib.Tactic.Positivity @@ -660,13 +661,11 @@ noncomputable def finEquivZPowers (x : G) (hx : IsOfFinOrder x) : Fin (orderOf x) ≃ (zpowers x : Set G) := (finEquivPowers x hx).trans <| Equiv.Set.ofEq hx.powers_eq_zpowers --- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[to_additive (attr := simp, nolint simpNF)] +@[to_additive] lemma finEquivZPowers_apply (hx) {n : Fin (orderOf x)} : finEquivZPowers x hx n = ⟨x ^ (n : ℕ), n, zpow_natCast x n⟩ := rfl - -- This lemma has always been bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644. -@[to_additive (attr := simp, nolint simpNF)] +@[to_additive] lemma finEquivZPowers_symm_apply (x : G) (hx) (n : ℕ) : (finEquivZPowers x hx).symm ⟨x ^ n, ⟨n, by simp⟩⟩ = ⟨n % orderOf x, Nat.mod_lt _ hx.orderOf_pos⟩ := by @@ -691,28 +690,12 @@ variable [Monoid G] {x : G} {n : ℕ} @[to_additive] theorem sum_card_orderOf_eq_card_pow_eq_one [Fintype G] [DecidableEq G] (hn : n ≠ 0) : - (∑ m ∈ (Finset.range n.succ).filter (· ∣ n), + (∑ m ∈ divisors n, (Finset.univ.filter fun x : G => orderOf x = m).card) = - (Finset.univ.filter fun x : G => x ^ n = 1).card := - calc - (∑ m ∈ (Finset.range n.succ).filter (· ∣ n), - (Finset.univ.filter fun x : G => orderOf x = m).card) = _ := - (Finset.card_biUnion - (by - intros - apply Finset.disjoint_filter.2 - rintro _ _ rfl; assumption)).symm - _ = _ := - congr_arg Finset.card - (Finset.ext - (by - intro x - suffices orderOf x ≤ n ∧ orderOf x ∣ n ↔ x ^ n = 1 by simpa [Nat.lt_succ_iff] - exact - ⟨fun h => by - let ⟨m, hm⟩ := h.2 - rw [hm, pow_mul, pow_orderOf_eq_one, one_pow], fun h => - ⟨orderOf_le_of_pow_eq_one hn.bot_lt h, orderOf_dvd_of_pow_eq_one h⟩⟩)) + (Finset.univ.filter fun x : G => x ^ n = 1).card := by + refine (Finset.card_biUnion ?_).symm.trans ?_ + · simp +contextual [disjoint_iff, Finset.ext_iff] + · congr; ext; simp [hn, orderOf_dvd_iff_pow_eq_one] @[to_additive] theorem orderOf_le_card_univ [Fintype G] : orderOf x ≤ Fintype.card G := diff --git a/Mathlib/GroupTheory/OreLocalization/Basic.lean b/Mathlib/GroupTheory/OreLocalization/Basic.lean index 24b43521f550b..356e809eb79d9 100644 --- a/Mathlib/GroupTheory/OreLocalization/Basic.lean +++ b/Mathlib/GroupTheory/OreLocalization/Basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2022 Jakob von Raumer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob von Raumer, Kevin Klinge, Andrew Yang -/ -import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.GroupTheory.OreLocalization.OreSet import Mathlib.Tactic.Common +import Mathlib.Algebra.Group.Submonoid.MulAction +import Mathlib.Algebra.Group.Units.Defs /-! diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 7a9275a5282d0..dd80ddc889ffc 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -5,7 +5,6 @@ Authors: Chris Hughes, Yaël Dillies -/ import Mathlib.Algebra.Module.BigOperators -import Mathlib.Data.List.Iterate import Mathlib.GroupTheory.Perm.Finite import Mathlib.GroupTheory.Perm.List @@ -193,10 +192,6 @@ theorem SameCycle.exists_pow_eq'' [Finite α] (h : SameCycle f x y) : rw [pow_orderOf_eq_one, pow_zero] · exact ⟨i.succ, i.zero_lt_succ, hi.le, by rfl⟩ -instance (f : Perm α) [DecidableRel (SameCycle f⁻¹)] : - DecidableRel (SameCycle f) := fun x y => - decidable_of_iff (f⁻¹.SameCycle x y) (sameCycle_inv) - instance (f : Perm α) [DecidableRel (SameCycle f)] : DecidableRel (SameCycle f⁻¹) := fun x y => decidable_of_iff (f.SameCycle x y) (sameCycle_inv).symm @@ -204,18 +199,6 @@ instance (f : Perm α) [DecidableRel (SameCycle f)] : instance (priority := 100) [DecidableEq α] : DecidableRel (SameCycle (1 : Perm α)) := fun x y => decidable_of_iff (x = y) sameCycle_one.symm -instance [Fintype α] [DecidableEq α] (f : Perm α) : DecidableRel (SameCycle f) := fun x y => - decidable_of_iff (y ∈ List.iterate f x (Fintype.card (Perm α))) <| by - simp only [List.mem_iterate, iterate_eq_pow, eq_comm (a := y)] - exact ⟨fun ⟨n, _, hn⟩ => ⟨n, hn⟩, fun ⟨i, hi⟩ => ⟨(i % orderOf f).natAbs, - Int.ofNat_lt.1 <| by - rw [Int.natAbs_of_nonneg (Int.emod_nonneg _ <| Int.natCast_ne_zero.2 (orderOf_pos _).ne')] - refine (Int.emod_lt _ <| Int.natCast_ne_zero_iff_pos.2 <| orderOf_pos _).trans_le ?_ - simp [orderOf_le_card_univ], - by - rw [← zpow_natCast, Int.natAbs_of_nonneg (Int.emod_nonneg _ <| - Int.natCast_ne_zero_iff_pos.2 <| orderOf_pos _), zpow_mod_orderOf, hi]⟩⟩ - end SameCycle /-! @@ -867,7 +850,7 @@ variable [DecidableEq α] {l : List α} theorem Nodup.isCycleOn_formPerm (h : l.Nodup) : l.formPerm.IsCycleOn { a | a ∈ l } := by refine ⟨l.formPerm.bijOn fun _ => List.formPerm_mem_iff_mem, fun a ha b hb => ?_⟩ - rw [Set.mem_setOf, ← List.indexOf_lt_length] at ha hb + rw [Set.mem_setOf, ← List.indexOf_lt_length_iff] at ha hb rw [← List.getElem_indexOf ha, ← List.getElem_indexOf hb] refine ⟨l.indexOf b - l.indexOf a, ?_⟩ simp only [sub_eq_neg_add, zpow_add, zpow_neg, Equiv.Perm.inv_eq_iff_eq, zpow_natCast, diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index a3fe94c15f780..516cc85380d9d 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yaël Dillies -/ +import Mathlib.Data.List.Iterate import Mathlib.GroupTheory.Perm.Cycle.Basic /-! @@ -156,24 +157,6 @@ theorem isCycle_cycleOf (f : Perm α) [DecidableRel f.SameCycle] (hx : f x ≠ x rw [cycleOf_apply_of_not_sameCycle hxy] at h exact (h rfl).elim⟩ -@[simp] -theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : - 2 ≤ #(cycleOf f x).support ↔ f x ≠ x := by - refine ⟨fun h => ?_, fun h => by simpa using (isCycle_cycleOf _ h).two_le_card_support⟩ - contrapose! h - rw [← cycleOf_eq_one_iff] at h - simp [h] - -@[simp] lemma support_cycleOf_nonempty [DecidableEq α] [Fintype α] : - (cycleOf f x).support.Nonempty ↔ f x ≠ x := by - rw [← two_le_card_support_cycleOf_iff, ← card_pos, ← Nat.succ_le_iff] - exact ⟨fun h => Or.resolve_left h.eq_or_lt (card_support_ne_one _).symm, zero_lt_two.trans_le⟩ - -@[deprecated support_cycleOf_nonempty (since := "2024-06-16")] -theorem card_support_cycleOf_pos_iff [DecidableEq α] [Fintype α] : - 0 < #(cycleOf f x).support ↔ f x ≠ x := by - rw [card_pos, support_cycleOf_nonempty] - theorem pow_mod_orderOf_cycleOf_apply (f : Perm α) [DecidableRel f.SameCycle] (n : ℕ) (x : α) : (f ^ (n % orderOf (cycleOf f x))) x = (f ^ n) x := by rw [← cycleOf_pow_apply_self f, ← cycleOf_pow_apply_self f, pow_mod_orderOf] @@ -199,18 +182,7 @@ theorem Disjoint.cycleOf_mul_distrib [DecidableRel f.SameCycle] [DecidableRel g. · simp [h.commute.eq, cycleOf_mul_of_apply_right_eq_self h.symm.commute, hfx] · simp [cycleOf_mul_of_apply_right_eq_self h.commute, hgx] -theorem support_cycleOf_eq_nil_iff [DecidableEq α] [Fintype α] : - (f.cycleOf x).support = ∅ ↔ x ∉ f.support := by simp - -theorem support_cycleOf_le [DecidableEq α] [Fintype α] (f : Perm α) (x : α) : - support (f.cycleOf x) ≤ support f := by - intro y hy - rw [mem_support, cycleOf_apply] at hy - split_ifs at hy - · exact mem_support.mpr hy - · exact absurd rfl hy - -theorem mem_support_cycleOf_iff [DecidableEq α] [Fintype α] : +private theorem mem_support_cycleOf_iff_aux [DecidableRel f.SameCycle] [DecidableEq α] [Fintype α] : y ∈ support (f.cycleOf x) ↔ SameCycle f x y ∧ x ∈ support f := by by_cases hx : f x = x · rw [(cycleOf_eq_one_iff _).mpr hx] @@ -223,21 +195,10 @@ theorem mem_support_cycleOf_iff [DecidableEq α] [Fintype α] : simpa using hx · simpa [hx] using hy -theorem mem_support_cycleOf_iff' (hx : f x ≠ x) [DecidableEq α] [Fintype α] : +private theorem mem_support_cycleOf_iff'_aux (hx : f x ≠ x) + [DecidableRel f.SameCycle] [DecidableEq α] [Fintype α] : y ∈ support (f.cycleOf x) ↔ SameCycle f x y := by - rw [mem_support_cycleOf_iff, and_iff_left (mem_support.2 hx)] - -theorem SameCycle.mem_support_iff {f} [DecidableEq α] [Fintype α] (h : SameCycle f x y) : - x ∈ support f ↔ y ∈ support f := - ⟨fun hx => support_cycleOf_le f x (mem_support_cycleOf_iff.mpr ⟨h, hx⟩), fun hy => - support_cycleOf_le f y (mem_support_cycleOf_iff.mpr ⟨h.symm, hy⟩)⟩ - -theorem pow_mod_card_support_cycleOf_self_apply [DecidableEq α] [Fintype α] - (f : Perm α) (n : ℕ) (x : α) : (f ^ (n % #(f.cycleOf x).support)) x = (f ^ n) x := by - by_cases hx : f x = x - · rw [pow_apply_eq_self_of_apply_eq_self hx, pow_apply_eq_self_of_apply_eq_self hx] - · rw [← cycleOf_pow_apply_self, ← cycleOf_pow_apply_self f, ← (isCycle_cycleOf f hx).orderOf, - pow_mod_orderOf] + rw [mem_support_cycleOf_iff_aux, and_iff_left (mem_support.2 hx)] /-- `x` is in the support of `f` iff `Equiv.Perm.cycle_of f x` is a cycle. -/ theorem isCycle_cycleOf_iff (f : Perm α) [DecidableRel f.SameCycle] : @@ -246,24 +207,104 @@ theorem isCycle_cycleOf_iff (f : Perm α) [DecidableRel f.SameCycle] : rw [Ne, ← cycleOf_eq_one_iff f] exact hx.ne_one -theorem isCycleOn_support_cycleOf [DecidableEq α] [Fintype α] (f : Perm α) (x : α) : - f.IsCycleOn (f.cycleOf x).support := +private theorem isCycleOn_support_cycleOf_aux [DecidableEq α] [Fintype α] (f : Perm α) + [DecidableRel f.SameCycle] (x : α) : f.IsCycleOn (f.cycleOf x).support := ⟨f.bijOn <| by - refine fun _ ↦ ⟨fun h ↦ mem_support_cycleOf_iff.2 ?_, fun h ↦ mem_support_cycleOf_iff.2 ?_⟩ - · exact ⟨sameCycle_apply_right.1 (mem_support_cycleOf_iff.1 h).1, - (mem_support_cycleOf_iff.1 h).2⟩ - · exact ⟨sameCycle_apply_right.2 (mem_support_cycleOf_iff.1 h).1, - (mem_support_cycleOf_iff.1 h).2⟩ + refine fun _ ↦ + ⟨fun h ↦ mem_support_cycleOf_iff_aux.2 ?_, fun h ↦ mem_support_cycleOf_iff_aux.2 ?_⟩ + · exact ⟨sameCycle_apply_right.1 (mem_support_cycleOf_iff_aux.1 h).1, + (mem_support_cycleOf_iff_aux.1 h).2⟩ + · exact ⟨sameCycle_apply_right.2 (mem_support_cycleOf_iff_aux.1 h).1, + (mem_support_cycleOf_iff_aux.1 h).2⟩ , fun a ha b hb => by - rw [mem_coe, mem_support_cycleOf_iff] at ha hb + rw [mem_coe, mem_support_cycleOf_iff_aux] at ha hb exact ha.1.symm.trans hb.1⟩ -theorem SameCycle.exists_pow_eq_of_mem_support {f} [DecidableEq α] [Fintype α] (h : SameCycle f x y) - (hx : x ∈ f.support) : ∃ i < #(f.cycleOf x).support, (f ^ i) x = y := by +private theorem SameCycle.exists_pow_eq_of_mem_support_aux {f} [DecidableEq α] [Fintype α] + [DecidableRel f.SameCycle] (h : SameCycle f x y) (hx : x ∈ f.support) : + ∃ i < #(f.cycleOf x).support, (f ^ i) x = y := by rw [mem_support] at hx - exact Equiv.Perm.IsCycleOn.exists_pow_eq (b := y) (f.isCycleOn_support_cycleOf x) - (by rw [mem_support_cycleOf_iff' hx]) (by rwa [mem_support_cycleOf_iff' hx]) + exact Equiv.Perm.IsCycleOn.exists_pow_eq (b := y) (f.isCycleOn_support_cycleOf_aux x) + (by rw [mem_support_cycleOf_iff'_aux hx]) (by rwa [mem_support_cycleOf_iff'_aux hx]) + +instance instDecidableRelSameCycle [DecidableEq α] [Fintype α] (f : Perm α) : + DecidableRel (SameCycle f) := fun x y => + decidable_of_iff (y ∈ List.iterate f x (Fintype.card α)) <| by + simp only [List.mem_iterate, iterate_eq_pow, eq_comm (a := y)] + constructor + · rintro ⟨n, _, hn⟩ + exact ⟨n, hn⟩ + · intro hxy + by_cases hx : x ∈ f.support + case pos => + -- we can't invoke the aux lemmas above without obtaining the decidable instance we are + -- already building; but now we've left the data, so we can do this non-constructively + -- without sacrificing computability. + let _inst (f : Perm α) : DecidableRel (SameCycle f) := Classical.decRel _ + rcases hxy.exists_pow_eq_of_mem_support_aux hx with ⟨i, hixy, hi⟩ + refine ⟨i, lt_of_lt_of_le hixy (card_le_univ _), hi⟩ + case neg => + haveI : Nonempty α := ⟨x⟩ + rw [not_mem_support] at hx + exact ⟨0, Fintype.card_pos, hxy.eq_of_left hx⟩ + +@[simp] +theorem two_le_card_support_cycleOf_iff [DecidableEq α] [Fintype α] : + 2 ≤ #(cycleOf f x).support ↔ f x ≠ x := by + refine ⟨fun h => ?_, fun h => by simpa using (isCycle_cycleOf _ h).two_le_card_support⟩ + contrapose! h + rw [← cycleOf_eq_one_iff] at h + simp [h] + +@[simp] lemma support_cycleOf_nonempty [DecidableEq α] [Fintype α] : + (cycleOf f x).support.Nonempty ↔ f x ≠ x := by + rw [← two_le_card_support_cycleOf_iff, ← card_pos, ← Nat.succ_le_iff] + exact ⟨fun h => Or.resolve_left h.eq_or_lt (card_support_ne_one _).symm, zero_lt_two.trans_le⟩ + +@[deprecated support_cycleOf_nonempty (since := "2024-06-16")] +theorem card_support_cycleOf_pos_iff [DecidableEq α] [Fintype α] : + 0 < #(cycleOf f x).support ↔ f x ≠ x := by + rw [card_pos, support_cycleOf_nonempty] + +theorem mem_support_cycleOf_iff [DecidableEq α] [Fintype α] : + y ∈ support (f.cycleOf x) ↔ SameCycle f x y ∧ x ∈ support f := + mem_support_cycleOf_iff_aux + +theorem mem_support_cycleOf_iff' (hx : f x ≠ x) [DecidableEq α] [Fintype α] : + y ∈ support (f.cycleOf x) ↔ SameCycle f x y := + mem_support_cycleOf_iff'_aux hx + +theorem support_cycleOf_eq_nil_iff [DecidableEq α] [Fintype α] : + (f.cycleOf x).support = ∅ ↔ x ∉ f.support := by simp + +theorem isCycleOn_support_cycleOf [DecidableEq α] [Fintype α] (f : Perm α) (x : α) : + f.IsCycleOn (f.cycleOf x).support := + isCycleOn_support_cycleOf_aux f x + +theorem SameCycle.exists_pow_eq_of_mem_support {f} [DecidableEq α] [Fintype α] (h : SameCycle f x y) + (hx : x ∈ f.support) : ∃ i < #(f.cycleOf x).support, (f ^ i) x = y := + h.exists_pow_eq_of_mem_support_aux hx + +theorem support_cycleOf_le [DecidableEq α] [Fintype α] (f : Perm α) (x : α) : + support (f.cycleOf x) ≤ support f := by + intro y hy + rw [mem_support, cycleOf_apply] at hy + split_ifs at hy + · exact mem_support.mpr hy + · exact absurd rfl hy + +theorem SameCycle.mem_support_iff {f} [DecidableEq α] [Fintype α] (h : SameCycle f x y) : + x ∈ support f ↔ y ∈ support f := + ⟨fun hx => support_cycleOf_le f x (mem_support_cycleOf_iff.mpr ⟨h, hx⟩), fun hy => + support_cycleOf_le f y (mem_support_cycleOf_iff.mpr ⟨h.symm, hy⟩)⟩ + +theorem pow_mod_card_support_cycleOf_self_apply [DecidableEq α] [Fintype α] + (f : Perm α) (n : ℕ) (x : α) : (f ^ (n % #(f.cycleOf x).support)) x = (f ^ n) x := by + by_cases hx : f x = x + · rw [pow_apply_eq_self_of_apply_eq_self hx, pow_apply_eq_self_of_apply_eq_self hx] + · rw [← cycleOf_pow_apply_self, ← cycleOf_pow_apply_self f, ← (isCycle_cycleOf f hx).orderOf, + pow_mod_orderOf] theorem SameCycle.exists_pow_eq [DecidableEq α] [Fintype α] (f : Perm α) (h : SameCycle f x y) : ∃ i : ℕ, 0 < i ∧ i ≤ #(f.cycleOf x).support + 1 ∧ (f ^ i) x = y := by @@ -307,50 +348,75 @@ section cycleFactors open scoped List in /-- Given a list `l : List α` and a permutation `f : Perm α` whose nonfixed points are all in `l`, recursively factors `f` into cycles. -/ -def cycleFactorsAux [DecidableEq α] [Fintype α] (l : List α) (f : Perm α) - (h : ∀ {x}, f x ≠ x → x ∈ l) : - { l : List (Perm α) // l.prod = f ∧ (∀ g ∈ l, IsCycle g) ∧ l.Pairwise Disjoint } := +def cycleFactorsAux [DecidableEq α] [Fintype α] + (l : List α) (f : Perm α) (h : ∀ {x}, f x ≠ x → x ∈ l) : + { pl : List (Perm α) // pl.prod = f ∧ (∀ g ∈ pl, IsCycle g) ∧ pl.Pairwise Disjoint } := + go l f h (fun _ => rfl) +where + /-- The auxiliary of `cycleFactorsAux`. This functions separates cycles from `f` instead of `g` + to prevent the process of a cycle gets complex. -/ + go (l : List α) (g : Perm α) (hg : ∀ {x}, g x ≠ x → x ∈ l) + (hfg : ∀ {x}, g x ≠ x → cycleOf f x = cycleOf g x) : + { pl : List (Perm α) // pl.prod = g ∧ (∀ g' ∈ pl, IsCycle g') ∧ pl.Pairwise Disjoint } := match l with | [] => ⟨[], by { simp only [imp_false, List.Pairwise.nil, List.not_mem_nil, forall_const, and_true, forall_prop_of_false, Classical.not_not, not_false_iff, List.prod_nil] at * ext simp [*]}⟩ - | x::l => - if hx : f x = x then cycleFactorsAux l f (by - intro y hy; exact List.mem_of_ne_of_mem (fun h => hy (by rwa [h])) (h hy)) + | x :: l => + if hx : g x = x then go l g (by + intro y hy; exact List.mem_of_ne_of_mem (fun h => hy (by rwa [h])) (hg hy)) hfg else - let ⟨m, hm⟩ := - cycleFactorsAux l ((cycleOf f x)⁻¹ * f) (by - intro y hy - exact List.mem_of_ne_of_mem - (fun h : y = x => by - rw [h, mul_apply, Ne, inv_eq_iff_eq, cycleOf_apply_self] at hy - exact hy rfl) - (h fun h : f y = y => by - rw [mul_apply, h, Ne, inv_eq_iff_eq, cycleOf_apply] at hy - split_ifs at hy <;> tauto)) - ⟨cycleOf f x :: m, by simp [List.prod_cons, hm.1], - fun g hg ↦ ((List.mem_cons).1 hg).elim (fun hg => hg ▸ isCycle_cycleOf _ hx) (hm.2.1 g), - List.pairwise_cons.2 - ⟨fun g hg y => - or_iff_not_imp_left.2 fun hfy => - have hxy : SameCycle f x y := - Classical.not_not.1 (mt cycleOf_apply_of_not_sameCycle hfy) - have hgm : (g::m.erase g) ~ m := - List.cons_perm_iff_perm_erase.2 ⟨hg, List.Perm.refl _⟩ - have : ∀ h ∈ m.erase g, Disjoint g h := - (List.pairwise_cons.1 ((hgm.pairwise_iff Disjoint.symm).2 hm.2.2)).1 - by_cases id fun hgy : g y ≠ y => - (disjoint_prod_right _ this y).resolve_right <| by - have hsc : SameCycle f⁻¹ x (f y) := by - rwa [sameCycle_inv, sameCycle_apply_right] - have hm₁ := hm.1 - rw [disjoint_prod_perm hm.2.2 hgm.symm, List.prod_cons, - ← eq_inv_mul_iff_mul_eq] at hm₁ - rwa [hm₁, mul_apply, mul_apply, cycleOf_inv, hsc.cycleOf_apply, inv_apply_self, - inv_eq_iff_eq, eq_comm], - hm.2.2⟩⟩ + let ⟨m, hm₁, hm₂, hm₃⟩ := + go l ((cycleOf f x)⁻¹ * g) (by + rw [hfg hx] + intro y hy + exact List.mem_of_ne_of_mem + (fun h : y = x => by + rw [h, mul_apply, Ne, inv_eq_iff_eq, cycleOf_apply_self] at hy + exact hy rfl) + (hg fun h : g y = y => by + rw [mul_apply, h, Ne, inv_eq_iff_eq, cycleOf_apply] at hy + split_ifs at hy <;> tauto)) + (by + rw [hfg hx] + intro y hy + simp [inv_eq_iff_eq, cycleOf_apply, eq_comm (a := g y)] at hy + rw [hfg (Ne.symm hy.right), ← mul_inv_eq_one (a := g.cycleOf y), cycleOf_inv] + simp_rw [mul_inv_rev] + rw [inv_inv, cycleOf_mul_of_apply_right_eq_self, ← cycleOf_inv, mul_inv_eq_one] + · rw [Commute.inv_left_iff, commute_iff_eq] + ext z; by_cases hz : SameCycle g x z + · simp [cycleOf_apply, hz] + · simp [cycleOf_apply_of_not_sameCycle, hz] + · exact cycleOf_apply_of_not_sameCycle hy.left) + ⟨cycleOf f x :: m, by + rw [hfg hx] at hm₁ ⊢ + constructor + · rw [List.prod_cons, hm₁] + simp + · exact + ⟨fun g' hg' => + ((List.mem_cons).1 hg').elim (fun hg' => hg'.symm ▸ isCycle_cycleOf _ hx) (hm₂ g'), + List.pairwise_cons.2 + ⟨fun g' hg' y => + or_iff_not_imp_left.2 fun hgy => + have hxy : SameCycle g x y := + Classical.not_not.1 (mt cycleOf_apply_of_not_sameCycle hgy) + have hg'm : (g' :: m.erase g') ~ m := + List.cons_perm_iff_perm_erase.2 ⟨hg', List.Perm.refl _⟩ + have : ∀ h ∈ m.erase g', Disjoint g' h := + (List.pairwise_cons.1 ((hg'm.pairwise_iff Disjoint.symm).2 hm₃)).1 + by_cases id fun hg'y : g' y ≠ y => + (disjoint_prod_right _ this y).resolve_right <| by + have hsc : SameCycle g⁻¹ x (g y) := by + rwa [sameCycle_inv, sameCycle_apply_right] + rw [disjoint_prod_perm hm₃ hg'm.symm, List.prod_cons, + ← eq_inv_mul_iff_mul_eq] at hm₁ + rwa [hm₁, mul_apply, mul_apply, cycleOf_inv, hsc.cycleOf_apply, + inv_apply_self, inv_eq_iff_eq, eq_comm], + hm₃⟩⟩⟩ theorem mem_list_cycles_iff {α : Type*} [Finite α] {l : List (Perm α)} (h1 : ∀ σ : Perm α, σ ∈ l → σ.IsCycle) (h2 : l.Pairwise Disjoint) {σ : Perm α} : diff --git a/Mathlib/GroupTheory/SemidirectProduct.lean b/Mathlib/GroupTheory/SemidirectProduct.lean index a003d225cc7e2..fb34e1cb43a1a 100644 --- a/Mathlib/GroupTheory/SemidirectProduct.lean +++ b/Mathlib/GroupTheory/SemidirectProduct.lean @@ -35,6 +35,9 @@ open Subgroup variable (N : Type*) (G : Type*) {H : Type*} [Group N] [Group G] [Group H] +-- Don't generate sizeOf and injectivity lemmas, which the `simpNF` linter will complain about. +set_option genSizeOfSpec false in +set_option genInjectivity false in /-- The semidirect product of groups `N` and `G`, given a map `φ` from `G` to the automorphism group of `N`. It the product of sets with the group operation `⟨n₁, g₁⟩ * ⟨n₂, g₂⟩ = ⟨n₁ * φ g₁ n₂, g₁ * g₂⟩` -/ @@ -46,11 +49,6 @@ structure SemidirectProduct (φ : G →* MulAut N) where right : G deriving DecidableEq --- Porting note: these lemmas are autogenerated by the inductive definition and are not --- in simple form due to the existence of mk_eq_inl_mul_inr -attribute [nolint simpNF] SemidirectProduct.mk.injEq -attribute [nolint simpNF] SemidirectProduct.mk.sizeOf_spec - -- Porting note: unknown attribute -- attribute [pp_using_anonymous_constructor] SemidirectProduct @@ -179,12 +177,18 @@ theorem range_inl_eq_ker_rightHom : (inl : N →* N ⋊[φ] G).range = rightHom. fun x hx ↦ ⟨x.left, by ext <;> simp_all [MonoidHom.mem_ker]⟩ /-- The bijection between the semidirect product and the product. -/ -@[simps!] -def equivProd : N ⋊[φ] G ≃ N × G := - { toFun := fun ⟨n, g⟩ ↦ ⟨n, g⟩ - invFun := fun ⟨n, g⟩ ↦ ⟨n, g⟩ - left_inv := fun _ ↦ rfl - right_inv := fun _ ↦ rfl } +@[simps] +def equivProd : N ⋊[φ] G ≃ N × G where + toFun x := ⟨x.1, x.2⟩ + invFun x := ⟨x.1, x.2⟩ + left_inv _ := rfl + right_inv _ := rfl + +/-- The group isomorphism between a semidirect product with respect to the trivial map + and the product. -/ +@[simps (config := {rhsMd := .default})] +def mulEquivProd : N ⋊[1] G ≃* N × G := + { equivProd with map_mul' _ _ := rfl } section lift @@ -305,4 +309,8 @@ def congr' : end Congr +@[simp] +lemma card : Nat.card (N ⋊[φ] G) = Nat.card N * Nat.card G := + Nat.card_prod _ _ ▸ Nat.card_congr equivProd + end SemidirectProduct diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 23b0bc5cfb952..d000af6ee0f64 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -6,8 +6,9 @@ Authors: Johannes Hölzl import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Data.Nat.Totient import Mathlib.Data.ZMod.Aut -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientGroup import Mathlib.GroupTheory.OrderOfElement +import Mathlib.GroupTheory.SpecificGroups.Dihedral import Mathlib.GroupTheory.Subgroup.Simple import Mathlib.Tactic.Group import Mathlib.GroupTheory.Exponent @@ -39,6 +40,7 @@ For the concrete cyclic group of order `n`, see `Data.ZMod.Basic`. cyclic group -/ +assert_not_exists TwoSidedIdeal variable {α G G' : Type*} {a : α} @@ -73,21 +75,25 @@ theorem isAddCyclic_additive_iff [Group α] : IsAddCyclic (Additive α) ↔ IsCy instance isAddCyclic_additive [Group α] [IsCyclic α] : IsAddCyclic (Additive α) := isAddCyclic_additive_iff.mpr inferInstance +@[to_additive] +instance IsCyclic.commutative [Group α] [IsCyclic α] : + Std.Commutative (· * · : α → α → α) where + comm x y := + let ⟨_, hg⟩ := IsCyclic.exists_generator (α := α) + let ⟨_, hx⟩ := hg x + let ⟨_, hy⟩ := hg y + hy ▸ hx ▸ zpow_mul_comm _ _ _ + /-- A cyclic group is always commutative. This is not an `instance` because often we have a better proof of `CommGroup`. -/ @[to_additive "A cyclic group is always commutative. This is not an `instance` because often we have a better proof of `AddCommGroup`."] def IsCyclic.commGroup [hg : Group α] [IsCyclic α] : CommGroup α := - { hg with - mul_comm := fun x y => - let ⟨_, hg⟩ := IsCyclic.exists_generator (α := α) - let ⟨_, hn⟩ := hg x - let ⟨_, hm⟩ := hg y - hm ▸ hn ▸ zpow_mul_comm _ _ _ } + { hg with mul_comm := commutative.comm } instance [Group G] (H : Subgroup G) [IsCyclic H] : H.IsCommutative := - ⟨⟨IsCyclic.commGroup.mul_comm⟩⟩ + ⟨IsCyclic.commutative⟩ variable [Group α] [Group G] [Group G'] @@ -283,8 +289,7 @@ open Finset Nat section Classical -open scoped Classical - +open scoped Classical in @[to_additive IsAddCyclic.card_nsmul_eq_zero_le] theorem IsCyclic.card_pow_eq_one_le [DecidableEq α] [Fintype α] [IsCyclic α] {n : ℕ} (hn0 : 0 < n) : #{a : α | a ^ n = 1} ≤ n := @@ -451,8 +456,8 @@ private theorem card_orderOf_eq_totient_aux₁ {d : ℕ} (hd : d ∣ Fintype.car have h2 : (∑ m ∈ d.divisors, #{a : α | orderOf a = m}) = ∑ m ∈ d.divisors, φ m := by - rw [← filter_dvd_eq_divisors hd0, sum_card_orderOf_eq_card_pow_eq_one hd0, - filter_dvd_eq_divisors hd0, sum_totient, ← ha, card_pow_eq_one_eq_orderOf_aux hn a] + rw [sum_card_orderOf_eq_card_pow_eq_one hd0, sum_totient, + ← ha, card_pow_eq_one_eq_orderOf_aux hn a] simpa [← cons_self_properDivisors hd0, ← h1] using h2 @[to_additive] @@ -467,7 +472,7 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : apply lt_irrefl c calc c = ∑ m ∈ c.divisors, #{a : α | orderOf a = m} := by - simp only [← filter_dvd_eq_divisors hc0.ne', sum_card_orderOf_eq_card_pow_eq_one hc0.ne'] + simp only [sum_card_orderOf_eq_card_pow_eq_one hc0.ne'] apply congr_arg card simp [c] _ = ∑ m ∈ c.divisors.erase d, #{a : α | orderOf a = m} := by @@ -858,3 +863,13 @@ lemma mulEquivOfOrderOfEq_symm_apply_gen : (mulEquivOfOrderOfEq hg hg' h).symm g end mulEquiv end generator + +lemma DihedralGroup.not_isCyclic {n : ℕ} (h1 : n ≠ 1) : ¬IsCyclic (DihedralGroup n) := fun h' => by + by_cases h2 : n = 2 + · simpa [exponent, card, h2] using h'.exponent_eq_card + · exact not_commutative h1 h2 h'.commutative + +lemma DihedralGroup.isCyclic_iff {n : ℕ} : + IsCyclic (DihedralGroup n) ↔ n = 1 where + mp := by contrapose; exact not_isCyclic + mpr := by rintro rfl; exact isCyclic_of_prime_card (p := 2) nat_card diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index ba8587d7527f9..49a2df71269d5 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Shing Tak Lam. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Shing Tak Lam -/ +import Mathlib.Data.Finite.Sum import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.Exponent import Mathlib.GroupTheory.GroupAction.CardCommute -import Mathlib.Data.Finite.Sum /-! # Dihedral Groups @@ -103,12 +103,12 @@ theorem one_def : (1 : DihedralGroup n) = r 0 := rfl private def fintypeHelper : (ZMod n) ⊕ (ZMod n) ≃ DihedralGroup n where - invFun i := match i with - | r j => Sum.inl j - | sr j => Sum.inr j - toFun i := match i with - | Sum.inl j => r j - | Sum.inr j => sr j + invFun + | r j => .inl j + | sr j => .inr j + toFun + | .inl j => r j + | .inr j => sr j left_inv := by rintro (x | x) <;> rfl right_inv := by rintro (x | x) <;> rfl @@ -208,6 +208,21 @@ theorem exponent : Monoid.exponent (DihedralGroup n) = lcm n 2 := by · convert Monoid.order_dvd_exponent (sr (0 : ZMod n)) exact (orderOf_sr 0).symm +lemma not_commutative : ∀ {n : ℕ}, n ≠ 1 → n ≠ 2 → + ¬Std.Commutative fun (x y : DihedralGroup n) => x * y + | 0, _, _ => fun ⟨h'⟩ ↦ by simpa using h' (r 1) (sr 0) + | n + 3, _, _ => by + rintro ⟨h'⟩ + specialize h' (r 1) (sr 0) + rw [r_mul_sr, zero_sub, sr_mul_r, zero_add, sr.injEq, neg_eq_iff_add_eq_zero, + one_add_one_eq_two, ← ZMod.val_eq_zero, ZMod.val_two_eq_two_mod] at h' + simpa using Nat.le_of_dvd Nat.zero_lt_two <| Nat.dvd_of_mod_eq_zero h' + +lemma commutative_iff {n : ℕ} : + Std.Commutative (fun x y : DihedralGroup n ↦ x * y) ↔ n = 1 ∨ n = 2 where + mp := by contrapose!; rintro ⟨h1, h2⟩; exact not_commutative h1 h2 + mpr := by rintro (rfl | rfl) <;> exact ⟨by decide⟩ + /-- If n is odd, then the Dihedral group of order $2n$ has $n(n+3)$ pairs (represented as $n + n + n + n*n$) of commuting elements. -/ @[simps] diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index 2ad3ed38cd7f4..542ab4e01b5a7 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -90,8 +90,8 @@ open Finset variable {G : Type*} [Group G] [IsKleinFour G] @[to_additive] -lemma not_isCyclic : ¬ IsCyclic G := - fun h ↦ by let _inst := Fintype.ofFinite G; simpa using h.exponent_eq_card +lemma not_isCyclic : ¬IsCyclic G := + fun h ↦ by simpa using h.exponent_eq_card @[to_additive] lemma inv_eq_self (x : G) : x⁻¹ = x := inv_eq_self_of_exponent_two (by simp) x diff --git a/Mathlib/GroupTheory/Subgroup/Center.lean b/Mathlib/GroupTheory/Subgroup/Center.lean index 6dd4dbfa41447..b8dc94177cf81 100644 --- a/Mathlib/GroupTheory/Subgroup/Center.lean +++ b/Mathlib/GroupTheory/Subgroup/Center.lean @@ -40,6 +40,16 @@ theorem center_toSubmonoid : (center G).toSubmonoid = Submonoid.center G := instance center.isCommutative : (center G).IsCommutative := ⟨⟨fun a b => Subtype.ext (b.2.comm a).symm⟩⟩ +variable {G} in +/-- The center of isomorphic groups are isomorphic. -/ +@[to_additive (attr := simps!) "The center of isomorphic additive groups are isomorphic."] +def centerCongr {H} [Group H] (e : G ≃* H) : center G ≃* center H := Submonoid.centerCongr e + +/-- The center of a group is isomorphic to the center of its opposite. -/ +@[to_additive (attr := simps!) +"The center of an additive group is isomorphic to the center of its opposite."] +def centerToMulOpposite : center G ≃* center Gᵐᵒᵖ := Submonoid.centerToMulOpposite + /-- For a group with zero, the center of the units is the same as the units of the center. -/ @[simps! apply_val_coe symm_apply_coe_val] def centerUnitsEquivUnitsCenter (G₀ : Type*) [GroupWithZero G₀] : diff --git a/Mathlib/GroupTheory/Submonoid/Center.lean b/Mathlib/GroupTheory/Submonoid/Center.lean index 44028df37b868..addceff424890 100644 --- a/Mathlib/GroupTheory/Submonoid/Center.lean +++ b/Mathlib/GroupTheory/Submonoid/Center.lean @@ -126,3 +126,61 @@ def unitsCenterToCenterUnits [Monoid M] : (Submonoid.center M)ˣ →* Submonoid. theorem unitsCenterToCenterUnits_injective [Monoid M] : Function.Injective (unitsCenterToCenterUnits M) := fun _a _b h => Units.ext <| Subtype.ext <| congr_arg (Units.val ∘ Subtype.val) h + +section congr + +variable {M} {N : Type*} + +@[to_additive] theorem _root_.MulEquivClass.apply_mem_center {F} [EquivLike F M N] [Mul M] [Mul N] + [MulEquivClass F M N] (e : F) {x : M} (hx : x ∈ Set.center M) : e x ∈ Set.center N := by + let e := MulEquivClass.toMulEquiv e + show e x ∈ Set.center N + constructor <;> + (intros; apply e.symm.injective; + simp only [map_mul, e.symm_apply_apply, (isMulCentral_iff _).mp hx]) + +@[to_additive] theorem _root_.MulEquivClass.apply_mem_center_iff {F} [EquivLike F M N] + [Mul M] [Mul N] [MulEquivClass F M N] (e : F) {x : M} : + e x ∈ Set.center N ↔ x ∈ Set.center M := + ⟨(by simpa using MulEquivClass.apply_mem_center (MulEquivClass.toMulEquiv e).symm ·), + MulEquivClass.apply_mem_center e⟩ + +/-- The center of isomorphic magmas are isomorphic. -/ +@[to_additive (attr := simps) "The center of isomorphic additive magmas are isomorphic."] +def Subsemigroup.centerCongr [Mul M] [Mul N] (e : M ≃* N) : center M ≃* center N where + toFun r := ⟨e r, MulEquivClass.apply_mem_center e r.2⟩ + invFun s := ⟨e.symm s, MulEquivClass.apply_mem_center e.symm s.2⟩ + left_inv _ := Subtype.ext (e.left_inv _) + right_inv _ := Subtype.ext (e.right_inv _) + map_mul' _ _ := Subtype.ext (map_mul ..) + +/-- The center of isomorphic monoids are isomorphic. -/ +@[to_additive (attr := simps!) "The center of isomorphic additive monoids are isomorphic."] +def Submonoid.centerCongr [MulOneClass M] [MulOneClass N] (e : M ≃* N) : center M ≃* center N := + Subsemigroup.centerCongr e + +@[to_additive] theorem MulOpposite.op_mem_center_iff [Mul M] {x : M} : + op x ∈ Set.center Mᵐᵒᵖ ↔ x ∈ Set.center M := by + simp_rw [Set.mem_center_iff, isMulCentral_iff, MulOpposite.forall, ← op_mul, op_inj]; aesop + +@[to_additive] theorem MulOpposite.unop_mem_center_iff [Mul M] {x : Mᵐᵒᵖ} : + unop x ∈ Set.center M ↔ x ∈ Set.center Mᵐᵒᵖ := + op_mem_center_iff.symm + +/-- The center of a magma is isomorphic to the center of its opposite. -/ +@[to_additive (attr := simps) +"The center of an additive magma is isomorphic to the center of its opposite."] +def Subsemigroup.centerToMulOpposite [Mul M] : center M ≃* center Mᵐᵒᵖ where + toFun r := ⟨_, MulOpposite.op_mem_center_iff.mpr r.2⟩ + invFun r := ⟨_, MulOpposite.unop_mem_center_iff.mpr r.2⟩ + left_inv _ := rfl + right_inv _ := rfl + map_mul' r _ := Subtype.ext (congr_arg MulOpposite.op <| r.2.1 _) + +/-- The center of a monoid is isomorphic to the center of its opposite. -/ +@[to_additive (attr := simps!) +"The center of an additive monoid is isomorphic to the center of its opposite. "] +def Submonoid.centerToMulOpposite [MulOneClass M] : center M ≃* center Mᵐᵒᵖ := + Subsemigroup.centerToMulOpposite + +end congr diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index f33f7fc8e19a5..02bb04970dc68 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -325,10 +325,7 @@ theorem normalizer_le_centralizer (hP : IsCyclic P) : P.normalizer ≤ centraliz apply Nat.coprime_one_right rw [hP.card_mulAut, hk, Nat.totient_prime_pow Fact.out h0] refine (Nat.Coprime.pow_right _ ?_).mul_right ?_ - · replace key : P.IsCommutative := by - let h := hP.commGroup - exact ⟨⟨CommGroup.mul_comm⟩⟩ - apply Nat.Coprime.coprime_dvd_left (relindex_dvd_of_le_left P.normalizer P.le_centralizer) + · apply Nat.Coprime.coprime_dvd_left (relindex_dvd_of_le_left P.normalizer P.le_centralizer) apply Nat.Coprime.coprime_dvd_left (relindex_dvd_index_of_le P.le_normalizer) rw [Nat.coprime_comm, Nat.Prime.coprime_iff_not_dvd Fact.out] exact P.not_dvd_index diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index 8ad9b91993a9b..47a11f3d84615 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -244,6 +244,8 @@ instance [Zero α] [∀ i, Zero (β i)] [∀ i, SMulWithZero α (β i)] : SMulWi instance [∀ i, AddMonoid (β i)] : AddMonoid (Hamming β) := Pi.addMonoid +instance [∀ i, AddGroup (β i)] : AddGroup (Hamming β) := Pi.addGroup + instance [∀ i, AddCommMonoid (β i)] : AddCommMonoid (Hamming β) := Pi.addCommMonoid @@ -396,15 +398,15 @@ instance [∀ i, Zero (β i)] : Norm (Hamming β) := theorem norm_eq_hammingNorm [∀ i, Zero (β i)] (x : Hamming β) : ‖x‖ = hammingNorm (ofHamming x) := rfl --- Porting note: merged `SeminormedAddCommGroup` and `NormedAddCommGroup` instances +instance [∀ i, AddGroup (β i)] : NormedAddGroup (Hamming β) where + dist_eq := by push_cast; exact mod_cast hammingDist_eq_hammingNorm instance [∀ i, AddCommGroup (β i)] : NormedAddCommGroup (Hamming β) where dist_eq := by push_cast; exact mod_cast hammingDist_eq_hammingNorm @[simp, push_cast] -theorem nnnorm_eq_hammingNorm [∀ i, AddCommGroup (β i)] (x : Hamming β) : - ‖x‖₊ = hammingNorm (ofHamming x) := - rfl +theorem nnnorm_eq_hammingNorm [∀ i, AddGroup (β i)] (x : Hamming β) : + ‖x‖₊ = hammingNorm (ofHamming x) := rfl end diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 1ac5e28cebf7a..1f0bbd7643e0b 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -40,8 +40,6 @@ such that both forward and inverse maps are affine. We define it using an `Equiv` for the map and a `LinearEquiv` for the linear part in order to allow affine equivalences with good definitional equalities. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure AffineEquiv (k P₁ P₂ : Type*) {V₁ V₂ : Type*} [Ring k] [AddCommGroup V₁] [Module k V₁] [AddTorsor V₁ P₁] [AddCommGroup V₂] [Module k V₂] [AddTorsor V₂ P₂] extends P₁ ≃ P₂ where linear : V₁ ≃ₗ[k] V₂ diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index 74dd02ec163ae..f3eee53f08930 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -748,7 +748,7 @@ This is the affine version of `Submodule.topEquiv`. -/ def topEquiv : (⊤ : AffineSubspace k P) ≃ᵃ[k] P where toEquiv := Equiv.Set.univ P linear := .ofEq _ _ (direction_top _ _ _) ≪≫ₗ Submodule.topEquiv - map_vadd' _p _v := rfl + map_vadd' _ _ := rfl variable {P} @@ -1471,7 +1471,7 @@ This is the affine version of `Submodule.inclusion`. -/ def inclusion (h : S₁ ≤ S₂) : S₁ →ᵃ[k] S₂ where toFun := Set.inclusion h linear := Submodule.inclusion <| AffineSubspace.direction_le h - map_vadd' _ _ := rfl + map_vadd' := fun ⟨_,_⟩ ⟨_,_⟩ => rfl @[simp] theorem coe_inclusion_apply (h : S₁ ≤ S₂) (x : S₁) : (inclusion h x : P₁) = x := @@ -1511,14 +1511,15 @@ This is the affine version of `LinearEquiv.ofEq`. -/ def ofEq (h : S₁ = S₂) : S₁ ≃ᵃ[k] S₂ where toEquiv := Equiv.Set.ofEq <| congr_arg _ h linear := .ofEq _ _ <| congr_arg _ h - map_vadd' _ _ := rfl + map_vadd' := fun ⟨_,_⟩ ⟨_,_⟩ => rfl @[simp] theorem coe_ofEq_apply (h : S₁ = S₂) (x : S₁) : (ofEq S₁ S₂ h x : P₁) = x := rfl @[simp] -theorem ofEq_symm (h : S₁ = S₂) : (ofEq S₁ S₂ h).symm = ofEq S₂ S₁ h.symm := +theorem ofEq_symm (h : S₁ = S₂) : (ofEq S₁ S₂ h).symm = ofEq S₂ S₁ h.symm := by + ext rfl @[simp] diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index d2945cbb493aa..dfdec73ed3c15 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -311,6 +311,32 @@ theorem repr_isUnitSMul {v : Basis ι R₂ M} {w : ι → R₂} (hw : ∀ i, IsU (v.isUnitSMul hw).repr x i = (hw i).unit⁻¹ • v.repr x i := repr_unitsSMul _ _ _ _ +/-- Any basis is a maximal linear independent set. +-/ +theorem maximal [Nontrivial R] (b : Basis ι R M) : b.linearIndependent.Maximal := fun w hi h => by + -- If `w` is strictly bigger than `range b`, + apply le_antisymm h + -- then choose some `x ∈ w \ range b`, + intro x p + by_contra q + -- and write it in terms of the basis. + have e := b.linearCombination_repr x + -- This then expresses `x` as a linear combination + -- of elements of `w` which are in the range of `b`, + let u : ι ↪ w := + ⟨fun i => ⟨b i, h ⟨i, rfl⟩⟩, fun i i' r => + b.injective (by simpa only [Subtype.mk_eq_mk] using r)⟩ + simp_rw [Finsupp.linearCombination_apply] at e + change ((b.repr x).sum fun (i : ι) (a : R) ↦ a • (u i : M)) = ((⟨x, p⟩ : w) : M) at e + rw [← Finsupp.sum_embDomain (f := u) (g := fun x r ↦ r • (x : M)), + ← Finsupp.linearCombination_apply] at e + -- Now we can contradict the linear independence of `hi` + refine hi.linearCombination_ne_of_not_mem_support _ ?_ e + simp only [Finset.mem_map, Finsupp.support_embDomain] + rintro ⟨j, -, W⟩ + simp only [u, Embedding.coeFn_mk, Subtype.mk_eq_mk] at W + apply q ⟨j, W⟩ + end Basis end Module @@ -341,32 +367,6 @@ theorem Basis.eq_bot_of_rank_eq_zero [NoZeroDivisors R] (b : Basis ι R M) (N : namespace Basis -/-- Any basis is a maximal linear independent set. --/ -theorem maximal [Nontrivial R] (b : Basis ι R M) : b.linearIndependent.Maximal := fun w hi h => by - -- If `w` is strictly bigger than `range b`, - apply le_antisymm h - -- then choose some `x ∈ w \ range b`, - intro x p - by_contra q - -- and write it in terms of the basis. - have e := b.linearCombination_repr x - -- This then expresses `x` as a linear combination - -- of elements of `w` which are in the range of `b`, - let u : ι ↪ w := - ⟨fun i => ⟨b i, h ⟨i, rfl⟩⟩, fun i i' r => - b.injective (by simpa only [Subtype.mk_eq_mk] using r)⟩ - simp_rw [Finsupp.linearCombination_apply] at e - change ((b.repr x).sum fun (i : ι) (a : R) ↦ a • (u i : M)) = ((⟨x, p⟩ : w) : M) at e - rw [← Finsupp.sum_embDomain (f := u) (g := fun x r ↦ r • (x : M)), - ← Finsupp.linearCombination_apply] at e - -- Now we can contradict the linear independence of `hi` - refine hi.linearCombination_ne_of_not_mem_support _ ?_ e - simp only [Finset.mem_map, Finsupp.support_embDomain] - rintro ⟨j, -, W⟩ - simp only [u, Embedding.coeFn_mk, Subtype.mk_eq_mk] at W - apply q ⟨j, W⟩ - 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/Cardinality.lean b/Mathlib/LinearAlgebra/Basis/Cardinality.lean index 64108a9f48908..406b7974e64e3 100644 --- a/Mathlib/LinearAlgebra/Basis/Cardinality.lean +++ b/Mathlib/LinearAlgebra/Basis/Cardinality.lean @@ -73,7 +73,7 @@ end Semiring section Ring -variable [Ring R] [AddCommGroup M] [Nontrivial R] [Module R M] +variable [Semiring R] [AddCommMonoid M] [Nontrivial R] [Module R M] -- From [Les familles libres maximales d'un module ont-elles le meme cardinal?][lazarus1973] /-- Over any ring `R`, if `b` is a basis for a module `M`, @@ -81,63 +81,43 @@ and `s` is a maximal linearly independent set, then the union of the supports of `x ∈ s` (when written out in the basis `b`) is all of `b`. -/ theorem union_support_maximal_linearIndependent_eq_range_basis {ι : Type w} (b : Basis ι R M) - {κ : Type w'} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : + {κ : Type w'} (v : κ → M) (ind : LinearIndependent R v) (m : ind.Maximal) : ⋃ k, ((b.repr (v k)).support : Set ι) = Set.univ := by -- If that's not the case, by_contra h simp only [← Ne.eq_def, ne_univ_iff_exists_not_mem, mem_iUnion, not_exists_not, Finsupp.mem_support_iff, Finset.mem_coe] at h - -- We have some basis element `b b'` which is not in the support of any of the `v i`. - obtain ⟨b', w⟩ := h + -- We have some basis element `b i` which is not in the support of any of the `v k`. + obtain ⟨i, w⟩ := h + have repr_eq_zero (l) : b.repr (linearCombination R v l) i = 0 := by + simp [linearCombination_apply, Finsupp.sum, w] -- Using this, we'll construct a linearly independent family strictly larger than `v`, - -- by also using this `b b'`. - let v' : Option κ → M := fun o => o.elim (b b') v - have r : range v ⊆ range v' := by - rintro - ⟨k, rfl⟩ - use some k - simp only [v', Option.elim_some] - have r' : b b' ∉ range v := by - rintro ⟨k, p⟩ - simpa [w] using congr_arg (fun m => (b.repr m) b') p - have r'' : range v ≠ range v' := by - intro e - have p : b b' ∈ range v' := by - use none - simp only [v', Option.elim_none] - rw [← e] at p - exact r' p + -- by also using this `b i`. + let v' (o : Option κ) : M := o.elim (b i) v + have r : range v ⊆ range v' := by rintro - ⟨k, rfl⟩; exact ⟨some k, rfl⟩ + have r' : b i ∉ range v := fun ⟨k, p⟩ ↦ by simpa [w] using congr(b.repr $p i) + have r'' : range v ≠ range v' := (r' <| · ▸ ⟨none, rfl⟩) -- The key step in the proof is checking that this strictly larger family is linearly independent. have i' : LinearIndependent R ((↑) : range v' → M) := by apply LinearIndependent.to_subtype_range - rw [linearIndependent_iff] - intro l z - rw [Finsupp.linearCombination_option] at z - simp only [v', Option.elim'] at z - change _ + Finsupp.linearCombination R v l.some = 0 at z - -- We have some linear combination of `b b'` and the `v i`, which we want to show is trivial. - -- We'll first show the coefficient of `b b'` is zero, - -- by expressing the `v i` in the basis `b`, and using that the `v i` have no `b b'` term. - have l₀ : l none = 0 := by - rw [← eq_neg_iff_add_eq_zero] at z - replace z := neg_eq_iff_eq_neg.mpr z - apply_fun fun x => b.repr x b' at z - simp only [repr_self, map_smul, mul_one, Finsupp.single_eq_same, Pi.neg_apply, - Finsupp.smul_single', map_neg, Finsupp.coe_neg] at z - erw [DFunLike.congr_fun (apply_linearCombination R (b.repr : M →ₗ[R] ι →₀ R) v l.some) b'] - at z - simpa [Finsupp.linearCombination_apply, w] using z - -- Then all the other coefficients are zero, because `v` is linear independent. - have l₁ : l.some = 0 := by - rw [l₀, zero_smul, zero_add] at z - exact linearIndependent_iff.mp i _ z - -- Finally we put those facts together to show the linear combination is trivial. + rw [linearIndependent_iffₛ] + intro l l' z + simp_rw [linearCombination_option, v', Option.elim'] at z + change _ + linearCombination R v l.some = _ + linearCombination R v l'.some at z + -- We have some equality between linear combinations of `b i` and the `v k`, + -- and want to show the coefficients are equal. ext (_ | a) - · simp only [l₀, Finsupp.coe_zero, Pi.zero_apply] - · erw [DFunLike.congr_fun l₁ a] - simp only [Finsupp.coe_zero, Pi.zero_apply] - rw [LinearIndependent.Maximal] at m - specialize m (range v') i' r - exact r'' m + -- We'll first show the coefficient of `b i` is zero, + -- by expressing the `v k` in the basis `b`, and using that the `v k` have no `b i` term. + · simpa [repr_eq_zero] using congr(b.repr $z i) + -- All the other coefficients are also equal, because `v` is linear independent, + -- by comparing the coefficients in the basis `b`. + have l₁ : l.some = l'.some := ind <| b.repr.injective <| ext fun j ↦ by + obtain rfl | ne := eq_or_ne i j + · simp_rw [repr_eq_zero] + classical simpa [single_apply, ne] using congr(b.repr $z j) + exact DFunLike.congr_fun l₁ a + exact r'' (m (range v') i' r) /-- Over any ring `R`, if `b` is an infinite basis for a module `M`, and `s` is a maximal linearly independent set, diff --git a/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean b/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean index 6d38ee83cd9d5..ebb2c333b7633 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean @@ -70,6 +70,16 @@ protected abbrev tmul (B₁ : BilinMap A M₁ N₁) (B₂ : BilinMap R M₂ N₂ BilinMap A (M₁ ⊗[R] M₂) (N₁ ⊗[R] N₂) := tensorDistrib R A (B₁ ⊗ₜ[R] B₂) +attribute [local ext] TensorProduct.ext in +/-- A tensor product of symmetric bilinear maps is symmetric. -/ +lemma tmul_isSymm {B₁ : BilinMap A M₁ N₁} {B₂ : BilinMap R M₂ N₂} + (hB₁ : ∀ x y, B₁ x y = B₁ y x) (hB₂ : ∀ x y, B₂ x y = B₂ y x) + (x y : M₁ ⊗[R] M₂) : + B₁.tmul B₂ x y = B₁.tmul B₂ y x := by + revert x y + rw [isSymm_iff_eq_flip] + aesop + variable (A) in /-- The base change of a bilinear map (also known as "extension of scalars"). -/ protected def baseChange (B : BilinMap R M₂ N₂) : BilinMap A (A ⊗[R] M₂) (A ⊗[R] N₂) := @@ -81,6 +91,10 @@ theorem baseChange_tmul (B₂ : BilinMap R M₂ N₂) (a : A) (m₂ : M₂) B₂.baseChange A (a ⊗ₜ m₂) (a' ⊗ₜ m₂') = (a * a') ⊗ₜ (B₂ m₂ m₂') := rfl +lemma baseChange_isSymm {B₂ : BilinMap R M₂ N₂} (hB₂ : ∀ x y, B₂ x y = B₂ y x) (x y : A ⊗[R] M₂) : + B₂.baseChange A x y = B₂.baseChange A y x := + tmul_isSymm mul_comm hB₂ x y + end BilinMap namespace BilinForm diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean index 9abbbe4782efb..22f87694b197f 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean @@ -242,7 +242,7 @@ theorem Q_apply (v : R × R) : Q c₁ c₂ v = c₁ * (v.1 * v.1) + c₂ * (v.2 /-- The quaternion basis vectors within the algebra. -/ @[simps i j k] -def quaternionBasis : QuaternionAlgebra.Basis (CliffordAlgebra (Q c₁ c₂)) c₁ c₂ where +def quaternionBasis : QuaternionAlgebra.Basis (CliffordAlgebra (Q c₁ c₂)) c₁ 0 c₂ where i := ι (Q c₁ c₂) (1, 0) j := ι (Q c₁ c₂) (0, 1) k := ι (Q c₁ c₂) (1, 0) * ι (Q c₁ c₂) (0, 1) @@ -254,16 +254,16 @@ def quaternionBasis : QuaternionAlgebra.Basis (CliffordAlgebra (Q c₁ c₂)) c simp i_mul_j := rfl j_mul_i := by - rw [eq_neg_iff_add_eq_zero, ι_mul_ι_add_swap, QuadraticMap.polar] + rw [zero_smul, zero_sub, eq_neg_iff_add_eq_zero, ι_mul_ι_add_swap, QuadraticMap.polar] simp variable {c₁ c₂} /-- Intermediate result of `CliffordAlgebraQuaternion.equiv`: clifford algebras over `CliffordAlgebraQuaternion.Q` can be converted to `ℍ[R,c₁,c₂]`. -/ -def toQuaternion : CliffordAlgebra (Q c₁ c₂) →ₐ[R] ℍ[R,c₁,c₂] := +def toQuaternion : CliffordAlgebra (Q c₁ c₂) →ₐ[R] ℍ[R,c₁,0,c₂] := CliffordAlgebra.lift (Q c₁ c₂) - ⟨{ toFun := fun v => (⟨0, v.1, v.2, 0⟩ : ℍ[R,c₁,c₂]) + ⟨{ toFun := fun v => (⟨0, v.1, v.2, 0⟩ : ℍ[R,c₁,0,c₂]) map_add' := fun v₁ v₂ => by simp map_smul' := fun r v => by dsimp; rw [mul_zero] }, fun v => by dsimp @@ -272,7 +272,7 @@ def toQuaternion : CliffordAlgebra (Q c₁ c₂) →ₐ[R] ℍ[R,c₁,c₂] := @[simp] theorem toQuaternion_ι (v : R × R) : - toQuaternion (ι (Q c₁ c₂) v) = (⟨0, v.1, v.2, 0⟩ : ℍ[R,c₁,c₂]) := + toQuaternion (ι (Q c₁ c₂) v) = (⟨0, v.1, v.2, 0⟩ : ℍ[R,c₁,0,c₂]) := CliffordAlgebra.lift_ι_apply _ _ v /-- The "clifford conjugate" maps to the quaternion conjugate. -/ @@ -280,22 +280,18 @@ theorem toQuaternion_star (c : CliffordAlgebra (Q c₁ c₂)) : toQuaternion (star c) = star (toQuaternion c) := by simp only [CliffordAlgebra.star_def'] induction c using CliffordAlgebra.induction with - | algebraMap r => - simp only [reverse.commutes, AlgHom.commutes, QuaternionAlgebra.coe_algebraMap, - QuaternionAlgebra.star_coe] - | ι x => - rw [reverse_ι, involute_ι, toQuaternion_ι, map_neg, toQuaternion_ι, - QuaternionAlgebra.neg_mk, star_mk, neg_zero] - | mul x₁ x₂ hx₁ hx₂ => simp only [reverse.map_mul, map_mul, hx₁, hx₂, star_mul] - | add x₁ x₂ hx₁ hx₂ => simp only [reverse.map_add, map_add, hx₁, hx₂, star_add] + | algebraMap r => simp + | ι x => simp + | mul x₁ x₂ hx₁ hx₂ => simp [hx₁, hx₂] + | add x₁ x₂ hx₁ hx₂ => simp [hx₁, hx₂] /-- Map a quaternion into the clifford algebra. -/ -def ofQuaternion : ℍ[R,c₁,c₂] →ₐ[R] CliffordAlgebra (Q c₁ c₂) := +def ofQuaternion : ℍ[R,c₁,0,c₂] →ₐ[R] CliffordAlgebra (Q c₁ c₂) := (quaternionBasis c₁ c₂).liftHom @[simp] theorem ofQuaternion_mk (a₁ a₂ a₃ a₄ : R) : - ofQuaternion (⟨a₁, a₂, a₃, a₄⟩ : ℍ[R,c₁,c₂]) = + ofQuaternion (⟨a₁, a₂, a₃, a₄⟩ : ℍ[R,c₁,0,c₂]) = algebraMap R _ a₁ + a₂ • ι (Q c₁ c₂) (1, 0) + a₃ • ι (Q c₁ c₂) (0, 1) + a₄ • (ι (Q c₁ c₂) (1, 0) * ι (Q c₁ c₂) (0, 1)) := rfl @@ -319,23 +315,23 @@ theorem ofQuaternion_toQuaternion (c : CliffordAlgebra (Q c₁ c₂)) : @[simp] theorem toQuaternion_comp_ofQuaternion : - toQuaternion.comp ofQuaternion = AlgHom.id R ℍ[R,c₁,c₂] := by + toQuaternion.comp ofQuaternion = AlgHom.id R ℍ[R,c₁,0,c₂] := by ext : 1 <;> simp @[simp] -theorem toQuaternion_ofQuaternion (q : ℍ[R,c₁,c₂]) : toQuaternion (ofQuaternion q) = q := +theorem toQuaternion_ofQuaternion (q : ℍ[R,c₁,0,c₂]) : toQuaternion (ofQuaternion q) = q := AlgHom.congr_fun toQuaternion_comp_ofQuaternion q /-- The clifford algebra over `CliffordAlgebraQuaternion.Q c₁ c₂` is isomorphic as an `R`-algebra to `ℍ[R,c₁,c₂]`. -/ @[simps!] -protected def equiv : CliffordAlgebra (Q c₁ c₂) ≃ₐ[R] ℍ[R,c₁,c₂] := +protected def equiv : CliffordAlgebra (Q c₁ c₂) ≃ₐ[R] ℍ[R,c₁,0,c₂] := AlgEquiv.ofAlgHom toQuaternion ofQuaternion toQuaternion_comp_ofQuaternion ofQuaternion_comp_toQuaternion /-- The quaternion conjugate maps to the "clifford conjugate" (aka `star`). -/ @[simp] -theorem ofQuaternion_star (q : ℍ[R,c₁,c₂]) : ofQuaternion (star q) = star (ofQuaternion q) := +theorem ofQuaternion_star (q : ℍ[R,c₁,0,c₂]) : ofQuaternion (star q) = star (ofQuaternion q) := CliffordAlgebraQuaternion.equiv.injective <| by rw [equiv_apply, equiv_apply, toQuaternion_star, toQuaternion_ofQuaternion, toQuaternion_ofQuaternion] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/SpinGroup.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/SpinGroup.lean index 06dfb81ec8f83..8bac15419c249 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/SpinGroup.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/SpinGroup.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Jiale Miao. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jiale Miao, Utensil Song, Eric Wieser -/ +import Mathlib.Algebra.Ring.Action.ConjAct import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.Algebra.Star.Unitary import Mathlib.LinearAlgebra.CliffordAlgebra.Star diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index 6ab5b47a0e4ee..7851cb879dbf9 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -336,6 +336,28 @@ theorem bot_lt_ker_of_det_eq_zero {𝕜 : Type*} [Field 𝕜] [Module 𝕜 M] {f lemma det_mulLeft (a : R) : (mulLeft R a).det = a := by simp lemma det_mulRight (a : R) : (mulRight R a).det = a := by simp +omit [DecidableEq ι] in +theorem det_pi [Module.Free R M] [Module.Finite R M] (f : ι → M →ₗ[R] M) : + (LinearMap.pi (fun i ↦ (f i).comp (LinearMap.proj i))).det = ∏ i, (f i).det := by + classical + let b := Module.Free.chooseBasis R M + let B := (Pi.basis (fun _ : ι ↦ b)).reindex <| + (Equiv.sigmaEquivProd _ _).trans (Equiv.prodComm _ _) + simp_rw [← LinearMap.det_toMatrix B, ← LinearMap.det_toMatrix b] + have : ((LinearMap.toMatrix B B) (LinearMap.pi fun i ↦ f i ∘ₗ LinearMap.proj i)) = + Matrix.blockDiagonal (fun i ↦ LinearMap.toMatrix b b (f i)) := by + ext ⟨i₁, i₂⟩ ⟨j₁, j₂⟩ + unfold B + simp_rw [LinearMap.toMatrix_apply', Matrix.blockDiagonal_apply, Basis.coe_reindex, + Function.comp_apply, Basis.repr_reindex_apply, Equiv.symm_trans_apply, Equiv.prodComm_symm, + Equiv.prodComm_apply, Equiv.sigmaEquivProd_symm_apply, Prod.swap_prod_mk, Pi.basis_apply, + Pi.basis_repr, LinearMap.pi_apply, LinearMap.coe_comp, Function.comp_apply, + LinearMap.toMatrix_apply', LinearMap.coe_proj, Function.eval, Pi.single_apply] + split_ifs with h + · rw [h] + · simp only [map_zero, Finsupp.coe_zero, Pi.zero_apply] + rw [this, Matrix.det_blockDiagonal] + end LinearMap diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 124960fd5f2ea..403f3c596d79d 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -48,6 +48,7 @@ variable (R M) /-- The rank of a module, defined as a term of type `Cardinal`. We define this as the supremum of the cardinalities of linearly independent subsets. +The supremum may not be attained, see https://mathoverflow.net/a/263053. For a free module over any ring satisfying the strong rank condition (e.g. left-noetherian rings, commutative rings, and in particular division rings and fields), @@ -95,81 +96,107 @@ end LinearIndependent section SurjectiveInjective -section Module -variable [Ring R] [AddCommGroup M] [Module R M] [Ring R'] +section Semiring +variable [Semiring R] [AddCommMonoid M] [Module R M] [Semiring R'] section -variable [AddCommGroup M'] [Module R' M'] +variable [AddCommMonoid M'] [Module R' M'] -/-- If `M / R` and `M' / R'` are modules, `i : R' → R` is a map which sends non-zero elements to -non-zero elements, `j : M →+ M'` is an injective group homomorphism, such that the scalar +/-- If `M / R` and `M' / R'` are modules, `i : R' → R` is an injective map +non-zero elements, `j : M →+ M'` is an injective monoid homomorphism, such that the scalar multiplications on `M` and `M'` are compatible, then the rank of `M / R` is smaller than or equal to the rank of `M' / R'`. As a special case, taking `R = R'` it is `LinearMap.lift_rank_le_of_injective`. -/ -theorem lift_rank_le_of_injective_injective (i : R' → R) (j : M →+ M') - (hi : ∀ r, i r = 0 → r = 0) (hj : Injective j) +theorem lift_rank_le_of_injective_injectiveₛ (i : R' → R) (j : M →+ M') + (hi : Injective i) (hj : Injective j) (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : lift.{v'} (Module.rank R M) ≤ lift.{v} (Module.rank R' M') := by simp_rw [Module.rank, lift_iSup (bddAbove_range _)] exact ciSup_mono' (bddAbove_range _) fun ⟨s, h⟩ ↦ ⟨⟨j '' s, - (h.map_of_injective_injective i j hi (fun _ _ ↦ hj <| by rwa [j.map_zero]) hc).image⟩, + (h.map_of_injective_injectiveₛ i j hi hj hc).image⟩, lift_mk_le'.mpr ⟨(Equiv.Set.image j s hj).toEmbedding⟩⟩ -/-- If `M / R` and `M' / R'` are modules, `i : R → R'` is a surjective map which maps zero to zero, -`j : M →+ M'` is an injective group homomorphism, such that the scalar multiplications on `M` and +/-- If `M / R` and `M' / R'` are modules, `i : R → R'` is a surjective map, and +`j : M →+ M'` is an injective monoid homomorphism, such that the scalar multiplications on `M` and `M'` are compatible, then the rank of `M / R` is smaller than or equal to the rank of `M' / R'`. As a special case, taking `R = R'` it is `LinearMap.lift_rank_le_of_injective`. -/ -theorem lift_rank_le_of_surjective_injective (i : ZeroHom R R') (j : M →+ M') +theorem lift_rank_le_of_surjective_injective (i : R → R') (j : M →+ M') (hi : Surjective i) (hj : Injective j) (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : lift.{v'} (Module.rank R M) ≤ lift.{v} (Module.rank R' M') := by obtain ⟨i', hi'⟩ := hi.hasRightInverse - refine lift_rank_le_of_injective_injective i' j (fun _ h ↦ ?_) hj fun r m ↦ ?_ + refine lift_rank_le_of_injective_injectiveₛ i' j (fun _ _ h ↦ ?_) hj fun r m ↦ ?_ · apply_fun i at h - rwa [hi', i.map_zero] at h + rwa [hi', hi'] at h rw [hc (i' r) m, hi'] /-- If `M / R` and `M' / R'` are modules, `i : R → R'` is a bijective map which maps zero to zero, `j : M ≃+ M'` is a group isomorphism, such that the scalar multiplications on `M` and `M'` are compatible, then the rank of `M / R` is equal to the rank of `M' / R'`. As a special case, taking `R = R'` it is `LinearEquiv.lift_rank_eq`. -/ -theorem lift_rank_eq_of_equiv_equiv (i : ZeroHom R R') (j : M ≃+ M') +theorem lift_rank_eq_of_equiv_equiv (i : R → R') (j : M ≃+ M') (hi : Bijective i) (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : lift.{v'} (Module.rank R M) = lift.{v} (Module.rank R' M') := (lift_rank_le_of_surjective_injective i j hi.2 j.injective hc).antisymm <| - lift_rank_le_of_injective_injective i j.symm (fun _ _ ↦ hi.1 <| by rwa [i.map_zero]) + lift_rank_le_of_injective_injectiveₛ i j.symm hi.1 j.symm.injective fun _ _ ↦ j.symm_apply_eq.2 <| by erw [hc, j.apply_symm_apply] end section -variable [AddCommGroup M₁] [Module R' M₁] +variable [AddCommMonoid M₁] [Module R' M₁] /-- The same-universe version of `lift_rank_le_of_injective_injective`. -/ -theorem rank_le_of_injective_injective (i : R' → R) (j : M →+ M₁) - (hi : ∀ r, i r = 0 → r = 0) (hj : Injective j) +theorem rank_le_of_injective_injectiveₛ (i : R' → R) (j : M →+ M₁) + (hi : Injective i) (hj : Injective j) (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : Module.rank R M ≤ Module.rank R' M₁ := by - simpa only [lift_id] using lift_rank_le_of_injective_injective i j hi hj hc + simpa only [lift_id] using lift_rank_le_of_injective_injectiveₛ i j hi hj hc /-- The same-universe version of `lift_rank_le_of_surjective_injective`. -/ -theorem rank_le_of_surjective_injective (i : ZeroHom R R') (j : M →+ M₁) +theorem rank_le_of_surjective_injective (i : R → R') (j : M →+ M₁) (hi : Surjective i) (hj : Injective j) (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : Module.rank R M ≤ Module.rank R' M₁ := by simpa only [lift_id] using lift_rank_le_of_surjective_injective i j hi hj hc /-- The same-universe version of `lift_rank_eq_of_equiv_equiv`. -/ -theorem rank_eq_of_equiv_equiv (i : ZeroHom R R') (j : M ≃+ M₁) +theorem rank_eq_of_equiv_equiv (i : R → R') (j : M ≃+ M₁) (hi : Bijective i) (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : Module.rank R M = Module.rank R' M₁ := by simpa only [lift_id] using lift_rank_eq_of_equiv_equiv i j hi hc end -end Module +end Semiring + +section Ring +variable [Ring R] [AddCommGroup M] [Module R M] [Ring R'] + +/-- If `M / R` and `M' / R'` are modules, `i : R' → R` is a map which sends non-zero elements to +non-zero elements, `j : M →+ M'` is an injective group homomorphism, such that the scalar +multiplications on `M` and `M'` are compatible, then the rank of `M / R` is smaller than or equal to +the rank of `M' / R'`. As a special case, taking `R = R'` it is +`LinearMap.lift_rank_le_of_injective`. -/ +theorem lift_rank_le_of_injective_injective [AddCommGroup M'] [Module R' M'] + (i : R' → R) (j : M →+ M') (hi : ∀ r, i r = 0 → r = 0) (hj : Injective j) + (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : + lift.{v'} (Module.rank R M) ≤ lift.{v} (Module.rank R' M') := by + simp_rw [Module.rank, lift_iSup (bddAbove_range _)] + exact ciSup_mono' (bddAbove_range _) fun ⟨s, h⟩ ↦ ⟨⟨j '' s, + (h.map_of_injective_injective i j hi (fun _ _ ↦ hj <| by rwa [j.map_zero]) hc).image⟩, + lift_mk_le'.mpr ⟨(Equiv.Set.image j s hj).toEmbedding⟩⟩ + +/-- The same-universe version of `lift_rank_le_of_injective_injective`. -/ +theorem rank_le_of_injective_injective [AddCommGroup M₁] [Module R' M₁] + (i : R' → R) (j : M →+ M₁) (hi : ∀ r, i r = 0 → r = 0) (hj : Injective j) + (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : + Module.rank R M ≤ Module.rank R' M₁ := by + simpa only [lift_id] using lift_rank_le_of_injective_injective i j hi hj hc + +end Ring namespace Algebra -variable {R : Type w} {S : Type v} [CommRing R] [Ring S] [Algebra R S] - {R' : Type w'} {S' : Type v'} [CommRing R'] [Ring S'] [Algebra R' S'] +variable {R : Type w} {S : Type v} [CommSemiring R] [Semiring S] [Algebra R S] + {R' : Type w'} {S' : Type v'} [CommSemiring R'] [Semiring S'] [Algebra R' S'] /-- If `S / R` and `S' / R'` are algebras, `i : R' →+* R` and `j : S →+* S'` are injective ring homomorphisms, such that `R' → R → S → S'` and `R' → S'` commute, then the rank of `S / R` is @@ -178,8 +205,7 @@ theorem lift_rank_le_of_injective_injective (i : R' →+* R) (j : S →+* S') (hi : Injective i) (hj : Injective j) (hc : (j.comp (algebraMap R S)).comp i = algebraMap R' S') : lift.{v'} (Module.rank R S) ≤ lift.{v} (Module.rank R' S') := by - refine _root_.lift_rank_le_of_injective_injective i j - (fun _ _ ↦ hi <| by rwa [i.map_zero]) hj fun r _ ↦ ?_ + refine _root_.lift_rank_le_of_injective_injectiveₛ i j hi hj fun r _ ↦ ?_ have := congr($hc r) simp only [RingHom.coe_comp, comp_apply] at this simp_rw [smul_def, AddMonoidHom.coe_coe, map_mul, this] @@ -207,7 +233,7 @@ theorem lift_rank_eq_of_equiv_equiv (i : R ≃+* R') (j : S ≃+* S') simp only [RingEquiv.toRingHom_eq_coe, RingHom.coe_comp, RingHom.coe_coe, comp_apply] at this simp only [smul_def, RingEquiv.coe_toAddEquiv, map_mul, ZeroHom.coe_coe, this] -variable {S' : Type v} [Ring S'] [Algebra R' S'] +variable {S' : Type v} [Semiring S'] [Algebra R' S'] /-- The same-universe version of `Algebra.lift_rank_le_of_injective_injective`. -/ theorem rank_le_of_injective_injective @@ -233,16 +259,15 @@ end Algebra end SurjectiveInjective -variable [Ring R] [AddCommGroup M] [Module R M] - [Ring R'] - [AddCommGroup M'] [AddCommGroup M₁] +variable [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring R'] [AddCommMonoid M'] [AddCommMonoid M₁] [Module R M'] [Module R M₁] [Module R' M'] [Module R' M₁] section theorem LinearMap.lift_rank_le_of_injective (f : M →ₗ[R] M') (i : Injective f) : Cardinal.lift.{v'} (Module.rank R M) ≤ Cardinal.lift.{v} (Module.rank R M') := - lift_rank_le_of_injective_injective (RingHom.id R) f (fun _ h ↦ h) i f.map_smul + lift_rank_le_of_injective_injectiveₛ (RingHom.id R) f (fun _ _ h ↦ h) i f.map_smul theorem LinearMap.rank_le_of_injective (f : M →ₗ[R] M₁) (i : Injective f) : Module.rank R M ≤ Module.rank R M₁ := @@ -350,7 +375,7 @@ theorem rank_subsingleton [Subsingleton R] : Module.rank R M = 1 := by apply subsingleton_of_subsingleton intro w hw refine ⟨⟨{0}, ?_⟩, ?_⟩ - · rw [linearIndependent_iff'] + · rw [linearIndependent_iff'ₛ] subsingleton · exact hw.trans_eq (Cardinal.mk_singleton _).symm diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index 427174ca24c09..5cdfb7e8f22f7 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro, Johannes Hölzl, Sander Dahmen, Kim Morrison, Chris Hug -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.LinearAlgebra.Dimension.Free -import Mathlib.LinearAlgebra.Isomorphisms /-! # Rank of various constructions @@ -38,15 +37,15 @@ variable {ι : Type w} {ι' : Type w'} {η : Type u₁'} {φ : η → Type*} open Basis Cardinal DirectSum Function Module Set Submodule +section Quotient + variable [Ring R] [CommRing S] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M₁] variable [Module R M] -section Quotient - theorem LinearIndependent.sum_elim_of_quotient {M' : Submodule R M} {ι₁ ι₂} {f : ι₁ → M'} (hf : LinearIndependent R f) (g : ι₂ → M) (hg : LinearIndependent R (Submodule.Quotient.mk (p := M') ∘ g)) : - LinearIndependent R (Sum.elim (f · : ι₁ → M) g) := by + LinearIndependent R (Sum.elim (f · : ι₁ → M) g) := by refine .sum_type (hf.map' M'.subtype M'.ker_subtype) (.of_comp M'.mkQ hg) ?_ refine disjoint_def.mpr fun x h₁ h₂ ↦ ?_ have : x ∈ M' := span_le.mpr (Set.range_subset_iff.mpr fun i ↦ (f i).prop) h₁ @@ -54,9 +53,9 @@ theorem LinearIndependent.sum_elim_of_quotient simp_rw [← Quotient.mk_eq_zero, ← mkQ_apply, map_finsupp_sum, map_smul, mkQ_apply] at this rw [linearIndependent_iff.mp hg _ this, Finsupp.sum_zero_index] -theorem LinearIndependent.union_of_quotient - {M' : Submodule R M} {s : Set M} (hs : s ⊆ M') (hs' : LinearIndependent (ι := s) R Subtype.val) - {t : Set M} (ht : LinearIndependent (ι := t) R (Submodule.Quotient.mk (p := M') ∘ Subtype.val)) : +theorem LinearIndependent.union_of_quotient {M' : Submodule R M} + {s : Set M} (hs : s ⊆ M') (hs' : LinearIndependent (ι := s) R Subtype.val) {t : Set M} + (ht : LinearIndependent (ι := t) R (Submodule.Quotient.mk (p := M') ∘ Subtype.val)) : LinearIndependent (ι := (s ∪ t :)) R Subtype.val := by refine (LinearIndependent.sum_elim_of_quotient (f := Set.embeddingOfSubset s M' hs) (of_comp M'.subtype (by simpa using hs')) Subtype.val ht).to_subtype_range' ?_ @@ -76,8 +75,17 @@ theorem rank_quotient_add_rank_le [Nontrivial R] (M' : Submodule R M) : theorem rank_quotient_le (p : Submodule R M) : Module.rank R (M ⧸ p) ≤ Module.rank R M := (mkQ p).rank_le_of_surjective Quot.mk_surjective +/-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ +theorem Submodule.finrank_quotient_le [StrongRankCondition R] [Module.Finite R M] + (s : Submodule R M) : finrank R (M ⧸ s) ≤ finrank R M := + toNat_le_toNat ((Submodule.mkQ s).rank_le_of_surjective Quot.mk_surjective) + (rank_lt_aleph0 _ _) + end Quotient +variable [Semiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid M'] [AddCommMonoid M₁] +variable [Module R M] + section ULift @[simp] @@ -95,20 +103,20 @@ section Prod variable (R M M') variable [Module R M₁] [Module R M'] -open LinearMap in -theorem lift_rank_add_lift_rank_le_rank_prod [Nontrivial R] : - lift.{v'} (Module.rank R M) + lift.{v} (Module.rank R M') ≤ Module.rank R (M × M') := by - convert rank_quotient_add_rank_le (ker <| LinearMap.fst R M M') - · refine Eq.trans ?_ (lift_id'.{v, v'} _) - rw [(quotKerEquivRange _).lift_rank_eq, - rank_range_of_surjective _ fst_surjective, lift_umax.{v, v'}] - · refine Eq.trans ?_ (lift_id'.{v', v} _) - rw [ker_fst, ← (LinearEquiv.ofInjective _ <| inr_injective (M := M) (M₂ := M')).lift_rank_eq, - lift_umax.{v', v}] - theorem rank_add_rank_le_rank_prod [Nontrivial R] : Module.rank R M + Module.rank R M₁ ≤ Module.rank R (M × M₁) := by - convert ← lift_rank_add_lift_rank_le_rank_prod R M M₁ <;> apply lift_id + conv_lhs => simp only [Module.rank_def] + have := nonempty_linearIndependent_set R M + have := nonempty_linearIndependent_set R M₁ + rw [Cardinal.ciSup_add_ciSup _ (bddAbove_range _) _ (bddAbove_range _)] + exact ciSup_le fun ⟨s, hs⟩ ↦ ciSup_le fun ⟨t, ht⟩ ↦ + (linearIndependent_inl_union_inr' hs ht).cardinal_le_rank + +theorem lift_rank_add_lift_rank_le_rank_prod [Nontrivial R] : + lift.{v'} (Module.rank R M) + lift.{v} (Module.rank R M') ≤ Module.rank R (M × M') := by + rw [← rank_ulift, ← rank_ulift] + exact (rank_add_rank_le_rank_prod R _).trans_eq + (ULift.moduleEquiv.prod ULift.moduleEquiv).rank_eq variable {R M M'} variable [StrongRankCondition R] [Module.Free R M] [Module.Free R M'] [Module.Free R M₁] @@ -161,7 +169,7 @@ theorem rank_finsupp_self' {ι : Type u} : Module.rank R (ι →₀ R) = #ι := /-- The rank of the direct sum is the sum of the ranks. -/ @[simp] -theorem rank_directSum {ι : Type v} (M : ι → Type w) [∀ i : ι, AddCommGroup (M i)] +theorem rank_directSum {ι : Type v} (M : ι → Type w) [∀ i : ι, AddCommMonoid (M i)] [∀ i : ι, Module R (M i)] [∀ i : ι, Module.Free R (M i)] : Module.rank R (⨁ i, M i) = Cardinal.sum fun i => Module.rank R (M i) := by let B i := chooseBasis R (M i) @@ -222,7 +230,7 @@ theorem finrank_finsupp_self {ι : Type v} [Fintype ι] : finrank R (ι →₀ R /-- The finrank of the direct sum is the sum of the finranks. -/ @[simp] -theorem finrank_directSum {ι : Type v} [Fintype ι] (M : ι → Type w) [∀ i : ι, AddCommGroup (M i)] +theorem finrank_directSum {ι : Type v} [Fintype ι] (M : ι → Type w) [∀ i : ι, AddCommMonoid (M i)] [∀ i : ι, Module R (M i)] [∀ i : ι, Module.Free R (M i)] [∀ i : ι, Module.Finite R (M i)] : finrank R (⨁ i, M i) = ∑ i, finrank R (M i) := by letI := nontrivial_of_invariantBasisNumber R @@ -241,7 +249,7 @@ end Finsupp section Pi variable [StrongRankCondition R] [Module.Free R M] -variable [∀ i, AddCommGroup (φ i)] [∀ i, Module R (φ i)] [∀ i, Module.Free R (φ i)] +variable [∀ i, AddCommMonoid (φ i)] [∀ i, Module R (φ i)] [∀ i, Module.Free R (φ i)] open Module.Free @@ -267,7 +275,7 @@ theorem Module.finrank_pi {ι : Type v} [Fintype ι] : --TODO: this should follow from `LinearEquiv.finrank_eq`, that is over a field. /-- The finrank of a finite product is the sum of the finranks. -/ theorem Module.finrank_pi_fintype - {ι : Type v} [Fintype ι] {M : ι → Type w} [∀ i : ι, AddCommGroup (M i)] + {ι : Type v} [Fintype ι] {M : ι → Type w} [∀ i : ι, AddCommMonoid (M i)] [∀ i : ι, Module R (M i)] [∀ i : ι, Module.Free R (M i)] [∀ i : ι, Module.Finite R (M i)] : finrank R (∀ i, M i) = ∑ i, finrank R (M i) := by letI := nontrivial_of_invariantBasisNumber R @@ -277,7 +285,7 @@ theorem Module.finrank_pi_fintype variable {R} variable [Fintype η] -theorem rank_fun {M η : Type u} [Fintype η] [AddCommGroup M] [Module R M] [Module.Free R M] : +theorem rank_fun {M η : Type u} [Fintype η] [AddCommMonoid M] [Module R M] [Module.Free R M] : Module.rank R (η → M) = Fintype.card η * Module.rank R M := by rw [rank_pi, Cardinal.sum_const', Cardinal.mk_fintype] @@ -378,12 +386,6 @@ theorem Submodule.finrank_le [Module.Finite R M] (s : Submodule R M) : finrank R s ≤ finrank R M := toNat_le_toNat (Submodule.rank_le s) (rank_lt_aleph0 _ _) -/-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ -theorem Submodule.finrank_quotient_le [Module.Finite R M] (s : Submodule R M) : - finrank R (M ⧸ s) ≤ finrank R M := - toNat_le_toNat ((Submodule.mkQ s).rank_le_of_surjective Quot.mk_surjective) - (rank_lt_aleph0 _ _) - /-- Pushforwards of finite submodules have a smaller finrank. -/ theorem Submodule.finrank_map_le [Module R M'] (f : M →ₗ[R] M') (p : Submodule R M) [Module.Finite R p] : @@ -485,7 +487,9 @@ section SubalgebraRank open Module -variable {F E : Type*} [CommRing F] [Ring E] [Algebra F E] +section Semiring + +variable {F E : Type*} [CommSemiring F] [Semiring E] [Algebra F E] @[simp] theorem Subalgebra.rank_toSubmodule (S : Subalgebra F E) : @@ -511,8 +515,11 @@ theorem Subalgebra.rank_top : Module.rank F (⊤ : Subalgebra F E) = Module.rank rw [subalgebra_top_rank_eq_submodule_top_rank] exact _root_.rank_top F E -section +end Semiring +section Ring + +variable {F E : Type*} [CommRing F] [Ring E] [Algebra F E] variable [StrongRankCondition F] [NoZeroSMulDivisors F E] [Nontrivial E] @[simp] @@ -526,6 +533,6 @@ theorem Subalgebra.rank_bot : Module.rank F (⊥ : Subalgebra F E) = 1 := theorem Subalgebra.finrank_bot : finrank F (⊥ : Subalgebra F E) = 1 := finrank_eq_of_rank_eq (by simp) -end +end Ring end SubalgebraRank diff --git a/Mathlib/LinearAlgebra/Dimension/Finrank.lean b/Mathlib/LinearAlgebra/Dimension/Finrank.lean index 90041e059c7a2..229fc22554cce 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finrank.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finrank.lean @@ -36,11 +36,11 @@ universe u v w open Cardinal Submodule Module Function variable {R : Type u} {M : Type v} {N : Type w} -variable [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] +variable [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] [Module R N] namespace Module -section Ring +section Semiring /-- The rank of a module as a natural number. @@ -51,7 +51,7 @@ of `M` over `R`. Note that it is possible to have `M` with `¬(Module.Finite R M)` but `finrank R M ≠ 0`, for example `ℤ × ℚ/ℤ` has `finrank` equal to `1`. -/ -noncomputable def finrank (R M : Type*) [Semiring R] [AddCommGroup M] [Module R M] : ℕ := +noncomputable def finrank (R M : Type*) [Semiring R] [AddCommMonoid M] [Module R M] : ℕ := Cardinal.toNat (Module.rank R M) @[deprecated (since := "2024-10-01")] protected alias _root_.FiniteDimensional.finrank := finrank @@ -94,7 +94,7 @@ theorem finrank_le_finrank_of_rank_le_rank (h' : Module.rank R N < ℵ₀) : finrank R M ≤ finrank R N := by simpa only [toNat_lift] using toNat_le_toNat h (lift_lt_aleph0.mpr h') -end Ring +end Semiring end Module @@ -102,7 +102,7 @@ open Module namespace LinearEquiv -variable {R M M₂ : Type*} [Ring R] [AddCommGroup M] [AddCommGroup M₂] +variable {R M M₂ : Type*} [Semiring R] [AddCommMonoid M] [AddCommMonoid M₂] variable [Module R M] [Module R M₂] /-- The dimension of a finite dimensional space is preserved under linear equivalence. -/ diff --git a/Mathlib/LinearAlgebra/Dimension/Free.lean b/Mathlib/LinearAlgebra/Dimension/Free.lean index a22d2a62026a5..2b06cd9f409f4 100644 --- a/Mathlib/LinearAlgebra/Dimension/Free.lean +++ b/Mathlib/LinearAlgebra/Dimension/Free.lean @@ -30,7 +30,7 @@ open Cardinal Basis Submodule Function Set DirectSum Module section Tower variable (F : Type u) (K : Type v) (A : Type w) -variable [Ring F] [Ring K] [AddCommGroup A] +variable [Semiring F] [Semiring K] [AddCommMonoid A] variable [Module F K] [Module K A] [Module F A] [IsScalarTower F K A] variable [StrongRankCondition F] [StrongRankCondition K] [Module.Free F K] [Module.Free K A] @@ -67,10 +67,10 @@ theorem Module.finrank_mul_finrank : finrank F K * finrank K A = finrank F A := end Tower variable {R : Type u} {M M₁ : Type v} {M' : Type v'} -variable [Ring R] [StrongRankCondition R] -variable [AddCommGroup M] [Module R M] [Module.Free R M] -variable [AddCommGroup M'] [Module R M'] [Module.Free R M'] -variable [AddCommGroup M₁] [Module R M₁] [Module.Free R M₁] +variable [Semiring R] [StrongRankCondition R] +variable [AddCommMonoid M] [Module R M] [Module.Free R M] +variable [AddCommMonoid M'] [Module R M'] [Module.Free R M'] +variable [AddCommMonoid M₁] [Module R M₁] [Module.Free R M₁] namespace Module.Free diff --git a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean index 48fd4277e7c4f..a8da7666e5135 100644 --- a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean +++ b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.LinearAlgebra.Dimension.Constructions import Mathlib.LinearAlgebra.Dimension.Finite +import Mathlib.LinearAlgebra.Isomorphisms /-! diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index 71be346b20eaa..55bb014a486dd 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -41,7 +41,7 @@ noncomputable section universe u v w w' -variable {R : Type u} {M : Type v} [Ring R] [AddCommGroup M] [Module R M] +variable {R : Type u} {M : Type v} [Semiring R] [AddCommMonoid M] [Module R M] variable {ι : Type w} {ι' : Type w'} open Cardinal Basis Submodule Function Set @@ -181,8 +181,7 @@ theorem linearIndependent_le_span_aux' {ι : Type*} [Fintype ι] (v : ι → M) apply_fun linearCombination R ((↑) : w → M) at h simp only [linearCombination_linearCombination, Submodule.coe_mk, Span.finsupp_linearCombination_repr] at h - rw [← sub_eq_zero, ← LinearMap.map_sub] at h - exact sub_eq_zero.mp (linearIndependent_iff.mp i _ h) + exact i h /-- If `R` satisfies the strong rank condition, then any linearly independent family `v : ι → M` @@ -320,12 +319,13 @@ theorem Basis.card_le_card_of_linearIndependent {ι : Type*} [Fintype ι] (b : B theorem Basis.card_le_card_of_submodule (N : Submodule R M) [Fintype ι] (b : Basis ι R M) [Fintype ι'] (b' : Basis ι' R N) : Fintype.card ι' ≤ Fintype.card ι := - b.card_le_card_of_linearIndependent (b'.linearIndependent.map' N.subtype N.ker_subtype) + b.card_le_card_of_linearIndependent + (b'.linearIndependent.map_injOn N.subtype N.injective_subtype.injOn) theorem Basis.card_le_card_of_le {N O : Submodule R M} (hNO : N ≤ O) [Fintype ι] (b : Basis ι R O) [Fintype ι'] (b' : Basis ι' R N) : Fintype.card ι' ≤ Fintype.card ι := b.card_le_card_of_linearIndependent - (b'.linearIndependent.map' (Submodule.inclusion hNO) (N.ker_inclusion O _)) + (b'.linearIndependent.map_injOn (inclusion hNO) (N.inclusion_injective _).injOn) theorem Basis.mk_eq_rank (v : Basis ι R M) : Cardinal.lift.{v} #ι = Cardinal.lift.{w} (Module.rank R M) := by @@ -351,9 +351,10 @@ theorem rank_span_set {s : Set M} (hs : LinearIndependent R (fun x => x : s → finite free module `M`. A property is true for all submodules of `M` if it satisfies the following "inductive step": the property is true for a submodule `N` if it's true for all submodules `N'` of `N` with the property that there exists `0 ≠ x ∈ N` such that the sum `N' + Rx` is direct. -/ -def Submodule.inductionOnRank [IsDomain R] [Finite ι] (b : Basis ι R M) - (P : Submodule R M → Sort*) (ih : ∀ N : Submodule R M, - (∀ N' ≤ N, ∀ x ∈ N, (∀ (c : R), ∀ y ∈ N', c • x + y = (0 : M) → c = 0) → P N') → P N) +def Submodule.inductionOnRank {R M} [Ring R] [StrongRankCondition R] [AddCommGroup M] [Module R M] + [IsDomain R] [Finite ι] (b : Basis ι R M) (P : Submodule R M → Sort*) + (ih : ∀ N : Submodule R M, + (∀ N' ≤ N, ∀ x ∈ N, (∀ (c : R), ∀ y ∈ N', c • x + y = (0 : M) → c = 0) → P N') → P N) (N : Submodule R M) : P N := letI := Fintype.ofFinite ι Submodule.inductionOnRankAux b P ih (Fintype.card ι) N fun hs hli => by @@ -455,7 +456,7 @@ end Module open Module -variable {M'} [AddCommGroup M'] [Module R M'] +variable {M'} [AddCommMonoid M'] [Module R M'] theorem LinearMap.finrank_le_finrank_of_injective [Module.Finite R M'] {f : M →ₗ[R] M'} (hf : Function.Injective f) : finrank R M ≤ finrank R M' := diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index b8ca669106157..a6c69a3ba7efa 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -824,7 +824,6 @@ def evalUseFiniteInstance : TacticM Unit := do elab "use_finite_instance" : tactic => evalUseFiniteInstance /-- `e` and `ε` have characteristic properties of a basis and its dual -/ --- @[nolint has_nonempty_instance] Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed structure Module.DualBases (e : ι → M) (ε : ι → Dual R M) : Prop where eval_same : ∀ i, ε i (e i) = 1 eval_of_ne : Pairwise fun i j ↦ ε i (e j) = 0 diff --git a/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean b/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean index e02ccb29fa8ea..fce6e1013a043 100644 --- a/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorPower/Basic.lean @@ -17,8 +17,12 @@ We study the exterior powers of a module `M` over a commutative ring `R`. * `exteriorPower.presentation R n M` is the standard presentation of the `R`-module `⋀[R]^n M`. +* `exteriorPower.map n f : ⋀[R]^n M →ₗ[R] ⋀[R]^n N` is the linear map on `nth` exterior powers +induced by a linear map `f : M →ₗ[R] N`. (See the file `Algebra.Category.ModuleCat.ExteriorPower` +for the corresponding functor `ModuleCat R ⥤ ModuleCat R`.) + ## Theorems -* `ιMulti_span`: The image of `exteriorPower.ιMulti` spans `⋀[R]^n M`. +* `exteriorPower.ι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 @@ -29,7 +33,7 @@ alternating maps and linear maps from the exterior power. open scoped TensorProduct -universe u v uM uN uN' uN'' uE uF +universe u variable (R : Type u) [CommRing R] (n : ℕ) {M N N' : Type*} [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] @@ -82,7 +86,7 @@ noncomputable def relations (ι : Type*) [DecidableEq ι] (M : Type*) Module.Relations R where G := ι → M R := Rels R ι M - relation r := match r with + relation | .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 @@ -198,4 +202,36 @@ lemma alternatingMapLinearEquiv_comp (g : N →ₗ[R] N') (f : M [⋀^Fin n]→ simp only [alternatingMapLinearEquiv_comp_ιMulti, LinearMap.compAlternatingMap_apply, LinearMap.coe_comp, comp_apply, alternatingMapLinearEquiv_apply_ιMulti] +/-! Functoriality of the exterior powers. -/ + +variable (n) in +/-- The linear map between `n`th exterior powers induced by a linear map between the modules. -/ +noncomputable def map (f : M →ₗ[R] N) : ⋀[R]^n M →ₗ[R] ⋀[R]^n N := + alternatingMapLinearEquiv ((ιMulti R n).compLinearMap f) + +@[simp] lemma alternatingMapLinearEquiv_symm_map (f : M →ₗ[R] N) : + alternatingMapLinearEquiv.symm (map n f) = (ιMulti R n).compLinearMap f := by + simp only [map, LinearEquiv.symm_apply_apply] + +@[simp] +theorem map_comp_ιMulti (f : M →ₗ[R] N) : + (map n f).compAlternatingMap (ιMulti R n) = (ιMulti R n).compLinearMap f := by + simp only [map, alternatingMapLinearEquiv_comp_ιMulti] + +@[simp] +theorem map_apply_ιMulti (f : M →ₗ[R] N) (m : Fin n → M) : + map n f (ιMulti R n m) = ιMulti R n (f ∘ m) := by + simp only [map, alternatingMapLinearEquiv_apply_ιMulti, AlternatingMap.compLinearMap_apply] + rfl + +@[simp] +theorem map_id : + map n (LinearMap.id (R := R) (M := M)) = LinearMap.id := by + aesop + +@[simp] +theorem map_comp (f : M →ₗ[R] N) (g : N →ₗ[R] N') : + map n (g ∘ₗ f) = map n g ∘ₗ map n f := by + aesop + end exteriorPower diff --git a/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean b/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean index 7e5fc7690e990..c083bf2d95447 100644 --- a/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean +++ b/Mathlib/LinearAlgebra/FreeModule/IdealQuotient.lean @@ -9,7 +9,7 @@ import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.LinearAlgebra.Quotient.Pi import Mathlib.RingTheory.Ideal.Basis import Mathlib.LinearAlgebra.Dimension.Constructions -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientRing /-! # Ideals in free modules over PIDs diff --git a/Mathlib/LinearAlgebra/FreeModule/Int.lean b/Mathlib/LinearAlgebra/FreeModule/Int.lean index c1b47b07a0ef7..c9a94cdfd1bce 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Int.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Int.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientGroup import Mathlib.GroupTheory.Index import Mathlib.LinearAlgebra.FreeModule.PID diff --git a/Mathlib/LinearAlgebra/FreeModule/PID.lean b/Mathlib/LinearAlgebra/FreeModule/PID.lean index ee2c18e4201b6..85fa3318e2a59 100644 --- a/Mathlib/LinearAlgebra/FreeModule/PID.lean +++ b/Mathlib/LinearAlgebra/FreeModule/PID.lean @@ -412,7 +412,6 @@ section SmithNormal /-- A Smith normal form basis for a submodule `N` of a module `M` consists of bases for `M` and `N` such that the inclusion map `N → M` can be written as a (rectangular) matrix with `a` along the diagonal: in Smith normal form. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): @[nolint has_nonempty_instance] structure Basis.SmithNormalForm (N : Submodule R M) (ι : Type*) (n : ℕ) where /-- The basis of M. -/ bM : Basis ι R M diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index b6bd8094f1851..b4156be52862b 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -100,8 +100,6 @@ theorem coe_quotientInfToSupQuotient (p p' : Submodule R M) : ⇑(quotientInfToSupQuotient p p') = quotientInfEquivSupQuotient p p' := rfl --- This lemma was always bad, but the linter only noticed after https://github.com/leanprover/lean4/pull/2644 -@[simp, nolint simpNF] theorem quotientInfEquivSupQuotient_apply_mk (p p' : Submodule R M) (x : p) : let map := inclusion (le_sup_left : p ≤ p ⊔ p') quotientInfEquivSupQuotient p p' (Submodule.Quotient.mk x) = diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 81a9e0d48fe9e..1b216c6e08dc5 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -297,16 +297,31 @@ theorem Submodule.range_ker_disjoint {f : M →ₗ[R] M'} such that they are both injective, and compatible with the scalar multiplications on `M` and `M'`, then `j` sends linearly independent families of vectors to linearly independent families of vectors. As a special case, taking `R = R'` -it is `LinearIndependent.map'`. -/ +it is `LinearIndependent.map_injOn`. -/ theorem LinearIndependent.map_of_injective_injectiveₛ {R' M' : Type*} [Semiring R'] [AddCommMonoid M'] [Module R' M'] (hv : LinearIndependent R v) - (i : R' → R) (j : M →+ M') (hi : Function.Injective i) (hj : Function.Injective j) + (i : R' → R) (j : M →+ M') (hi : Injective i) (hj : Injective j) (hc : ∀ (r : R') (m : M), j (i r • m) = r • j m) : LinearIndependent R' (j ∘ v) := by rw [linearIndependent_iff'ₛ] at hv ⊢ intro S r₁ r₂ H s hs simp_rw [comp_apply, ← hc, ← map_sum] at H exact hi <| hv _ _ _ (hj H) s hs +/-- If `M / R` and `M' / R'` are modules, `i : R → R'` is a surjective map, +and `j : M →+ M'` is an injective monoid map, such that the scalar multiplications +on `M` and `M'` are compatible, then `j` sends linearly independent families +of vectors to linearly independent families of vectors. As a special case, taking `R = R'` +it is `LinearIndependent.map_injOn`. -/ +theorem LinearIndependent.map_of_surjective_injectiveₛ {R' M' : Type*} + [Semiring R'] [AddCommMonoid M'] [Module R' M'] (hv : LinearIndependent R v) + (i : R → R') (j : M →+ M') (hi : Surjective i) (hj : Injective j) + (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : LinearIndependent R' (j ∘ v) := by + obtain ⟨i', hi'⟩ := hi.hasRightInverse + refine hv.map_of_injective_injectiveₛ i' j (fun _ _ h ↦ ?_) hj fun r m ↦ ?_ + · apply_fun i at h + rwa [hi', hi'] at h + rw [hc (i' r) m, hi'] + /-- If the image of a family of vectors under a linear map is linearly independent, then so is the original family. -/ theorem LinearIndependent.of_comp (f : M →ₗ[R] M') (hfv : LinearIndependent R (f ∘ v)) : @@ -925,15 +940,11 @@ theorem LinearIndependent.map_of_injective_injective {R' M' : Type*} scalar multiplications on `M` and `M'` are compatible, then `j` sends linearly independent families of vectors to linearly independent families of vectors. As a special case, taking `R = R'` it is `LinearIndependent.map'`. -/ -theorem LinearIndependent.map_of_surjective_injective {R' : Type*} {M' : Type*} +theorem LinearIndependent.map_of_surjective_injective {R' M' : Type*} [Ring R'] [AddCommGroup M'] [Module R' M'] (hv : LinearIndependent R v) - (i : ZeroHom R R') (j : M →+ M') (hi : Surjective i) (hj : ∀ m, j m = 0 → m = 0) - (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : LinearIndependent R' (j ∘ v) := by - obtain ⟨i', hi'⟩ := hi.hasRightInverse - refine hv.map_of_injective_injective i' j (fun _ h ↦ ?_) hj fun r m ↦ ?_ - · apply_fun i at h - rwa [hi', i.map_zero] at h - rw [hc (i' r) m, hi'] + (i : R → R') (j : M →+ M') (hi : Surjective i) (hj : ∀ m, j m = 0 → m = 0) + (hc : ∀ (r : R) (m : M), j (r • m) = i r • j m) : LinearIndependent R' (j ∘ v) := + hv.map_of_surjective_injectiveₛ i _ hi ((injective_iff_map_eq_zero _).mpr hj) hc /-- If `f` is an injective linear map, then the family `f ∘ v` is linearly independent if and only if the family `v` is linearly independent. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index 8befb3f75b370..8ae8358bfca55 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -39,7 +39,7 @@ theorem PiToModule.fromMatrix_apply [DecidableEq ι] (A : Matrix ι ι R) (w : theorem PiToModule.fromMatrix_apply_single_one [DecidableEq ι] (A : Matrix ι ι R) (j : ι) : PiToModule.fromMatrix R b A (Pi.single j 1) = ∑ i : ι, A i j • b i := by rw [PiToModule.fromMatrix_apply, Fintype.linearCombination_apply, Matrix.mulVec_single] - simp_rw [mul_one] + simp_rw [MulOpposite.op_one, one_smul, transpose_apply] /-- The endomorphisms of `M` acts on `(ι → R) →ₗ[R] M`, and takes the projection to a `(ι → R) →ₗ[R] M`. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 6a59c9e32abf2..e61efb9e17587 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -130,6 +130,7 @@ end CoeLemmas variable {S T : Type*} [CommRing S] [CommRing T] /-- A ring homomorphism ``f : R →+* S`` induces a homomorphism ``GLₙ(f) : GLₙ(R) →* GLₙ(S)``. -/ +@[simps! apply_val] def map (f : R →+* S) : GL n R →* GL n S := Units.map <| (RingHom.mapMatrix f).toMonoidHom @[simp] diff --git a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean index 89a34b681a4a6..c7e39b7cc0e1d 100644 --- a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean +++ b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean @@ -279,7 +279,7 @@ theorem isHermitian_iff_isSymmetric [Fintype n] [DecidableEq n] {A : Matrix n n · intro h ext i j simpa only [(Pi.single_star i 1).symm, ← star_mulVec, mul_one, dotProduct_single, - single_vecMul, star_one, one_mul] using h (Pi.single i 1) (Pi.single j 1) + single_one_vecMul, star_one] using h (Pi.single i 1) (Pi.single j 1) end RCLike diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index b279eda8cbc9c..4e0b53338d713 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -666,11 +666,9 @@ def invertibleOfSubmatrixEquivInvertible (A : Matrix m m α) (e₁ e₂ : n ≃ [Invertible (A.submatrix e₁ e₂)] : Invertible A := invertibleOfRightInverse _ ((⅟ (A.submatrix e₁ e₂)).submatrix e₂.symm e₁.symm) <| by have : A = (A.submatrix e₁ e₂).submatrix e₁.symm e₂.symm := by simp - -- Porting note: was - -- conv in _ * _ => - -- congr - -- rw [this] - rw [congr_arg₂ (· * ·) this rfl] + conv in _ * _ => + congr + rw [this] rw [Matrix.submatrix_mul_equiv, mul_invOf_self, submatrix_one_equiv] theorem invOf_submatrix_equiv_eq (A : Matrix m m α) (e₁ e₂ : n ≃ m) [Invertible A] diff --git a/Mathlib/LinearAlgebra/Matrix/Permutation.lean b/Mathlib/LinearAlgebra/Matrix/Permutation.lean index 96c0d8f5fc771..990b92bc1ee05 100644 --- a/Mathlib/LinearAlgebra/Matrix/Permutation.lean +++ b/Mathlib/LinearAlgebra/Matrix/Permutation.lean @@ -39,7 +39,7 @@ namespace Matrix /-- The determinant of a permutation matrix equals its sign. -/ @[simp] theorem det_permutation [CommRing R] : det (σ.permMatrix R) = Perm.sign σ := by - rw [← Matrix.mul_one (σ.permMatrix R), PEquiv.toPEquiv_mul_matrix, + rw [← Matrix.mul_one (σ.permMatrix R), PEquiv.toMatrix_toPEquiv_mul, det_permute, det_one, mul_one] /-- The trace of a permutation matrix equals the number of fixed points. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 5e46c1956c392..7776201ccffd1 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -363,8 +363,8 @@ theorem _root_.Matrix.posDef_diagonal_iff PosDef (diagonal d) ↔ ∀ i, 0 < d i := by refine ⟨fun h i => ?_, .diagonal⟩ have := h.2 (Pi.single i 1) - simp only [mulVec_single, mul_one, dotProduct_diagonal', Pi.star_apply, Pi.single_eq_same, - star_one, one_mul, Function.ne_iff, Pi.zero_apply] at this + simp_rw [mulVec_single_one, ← Pi.single_star, star_one, single_dotProduct, one_mul, + transpose_apply, diagonal_apply_eq, Function.ne_iff] at this exact this ⟨i, by simp⟩ protected theorem one [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] : diff --git a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean index 5bd189f1e2b73..eb134172c1ac3 100644 --- a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean +++ b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean @@ -445,6 +445,7 @@ section StarOrderedRing variable {𝕜 : Type*} [CommRing 𝕜] [StarRing 𝕜] +/-- Notation for `Sum.elim`, scoped within the `Matrix` namespace. -/ scoped infixl:65 " ⊕ᵥ " => Sum.elim theorem schur_complement_eq₁₁ [Fintype m] [DecidableEq m] [Fintype n] {A : Matrix m m 𝕜} diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index 1bc8dac9cd047..04b4156f29cde 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Wen Yang -/ import Mathlib.LinearAlgebra.Matrix.Adjugate +import Mathlib.LinearAlgebra.Matrix.ToLin import Mathlib.LinearAlgebra.Matrix.Transvection import Mathlib.RingTheory.RootsOfUnity.Basic diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index 3367f16ef2d73..ef6287892c69e 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -74,6 +74,11 @@ lemma eigenvectorUnitary_coe {𝕜 : Type*} [RCLike 𝕜] {n : Type*} [Fintype n (EuclideanSpace.basisFun n 𝕜).toBasis.toMatrix (hA.eigenvectorBasis).toBasis := rfl +@[simp] +theorem eigenvectorUnitary_transpose_apply (j : n) : + (eigenvectorUnitary hA)ᵀ j = ⇑(hA.eigenvectorBasis j) := + rfl + @[simp] theorem eigenvectorUnitary_apply (i j : n) : eigenvectorUnitary hA i j = ⇑(hA.eigenvectorBasis j) i := @@ -81,7 +86,7 @@ theorem eigenvectorUnitary_apply (i j : n) : theorem eigenvectorUnitary_mulVec (j : n) : eigenvectorUnitary hA *ᵥ Pi.single j 1 = ⇑(hA.eigenvectorBasis j) := by - simp only [mulVec_single, eigenvectorUnitary_apply, mul_one] + simp_rw [mulVec_single_one, eigenvectorUnitary_transpose_apply] theorem star_eigenvectorUnitary_mulVec (j : n) : (star (eigenvectorUnitary hA : Matrix n n 𝕜)) *ᵥ ⇑(hA.eigenvectorBasis j) = Pi.single j 1 := by diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index 942c5135c5ac8..506f38bc302e3 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -249,7 +249,7 @@ variable [Fintype n] @[simp] theorem Matrix.mulVecLin_one [DecidableEq n] : Matrix.mulVecLin (1 : Matrix n n R) = LinearMap.id := by - ext; simp [Matrix.one_apply, Pi.single_apply] + ext; simp [Matrix.one_apply, Pi.single_apply, eq_comm] @[simp] theorem Matrix.mulVecLin_mul [Fintype m] (M : Matrix l m R) (N : Matrix m n R) : @@ -469,9 +469,6 @@ theorem Matrix.toLinAlgEquiv'_apply (M : Matrix n n R) (v : n → R) : Matrix.toLinAlgEquiv' M v = M *ᵥ v := rfl --- Porting note: the simpNF linter rejects this, as `simp` already simplifies the lhs --- to `(1 : (n → R) →ₗ[R] n → R)`. --- @[simp] theorem Matrix.toLinAlgEquiv'_one : Matrix.toLinAlgEquiv' (1 : Matrix n n R) = LinearMap.id := Matrix.toLin'_one @@ -750,9 +747,6 @@ theorem Matrix.toLinAlgEquiv_self (M : Matrix n n R) (i : n) : theorem LinearMap.toMatrixAlgEquiv_id : LinearMap.toMatrixAlgEquiv v₁ id = 1 := by simp_rw [LinearMap.toMatrixAlgEquiv, AlgEquiv.ofLinearEquiv_apply, LinearMap.toMatrix_id] --- Porting note: the simpNF linter rejects this, as `simp` already simplifies the lhs --- to `(1 : M₁ →ₗ[R] M₁)`. --- @[simp] theorem Matrix.toLinAlgEquiv_one : Matrix.toLinAlgEquiv v₁ 1 = LinearMap.id := by rw [← LinearMap.toMatrixAlgEquiv_id v₁, Matrix.toLinAlgEquiv_toMatrixAlgEquiv] diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 9b9593fe097ac..8dc27fc0a5f62 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -136,7 +136,6 @@ variable (R n) /-- A structure containing all the information from which one can build a nontrivial transvection. This structure is easier to manipulate than transvections as one has a direct access to all the relevant fields. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] structure TransvectionStruct where (i j : n) hij : i ≠ j diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index a63088b6181d5..91124505b00ab 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -82,6 +82,8 @@ universe uR uS uι v v' v₁ v₂ v₃ variable {R : Type uR} {S : Type uS} {ι : Type uι} {n : ℕ} {M : Fin n.succ → Type v} {M₁ : ι → Type v₁} {M₂ : Type v₂} {M₃ : Type v₃} {M' : Type v'} +-- Don't generate injectivity lemmas, which the `simpNF` linter will time out on. +set_option genInjectivity false in /-- Multilinear maps over the ring `R`, from `∀ i, M₁ i` to `M₂` where `M₁ i` and `M₂` are modules over `R`. -/ structure MultilinearMap (R : Type uR) {ι : Type uι} (M₁ : ι → Type v₁) (M₂ : Type v₂) [Semiring R] @@ -97,9 +99,6 @@ structure MultilinearMap (R : Type uR) {ι : Type uι} (M₁ : ι → Type v₁) ∀ [DecidableEq ι] (m : ∀ i, M₁ i) (i : ι) (c : R) (x : M₁ i), toFun (update m i (c • x)) = c • toFun (update m i x) --- Porting note: added to avoid a linter timeout. -attribute [nolint simpNF] MultilinearMap.mk.injEq - namespace MultilinearMap section Semiring @@ -1156,13 +1155,13 @@ protected def mkPiAlgebraFin : MultilinearMap R (fun _ : Fin n => A) A where map_update_add' {dec} m i x y := by rw [Subsingleton.elim dec (by infer_instance)] have : (List.finRange n).indexOf i < n := by - simpa using List.indexOf_lt_length.2 (List.mem_finRange i) + simpa using List.indexOf_lt_length_iff.2 (List.mem_finRange i) simp [List.ofFn_eq_map, (List.nodup_finRange n).map_update, List.prod_set, add_mul, this, mul_add, add_mul] map_update_smul' {dec} m i c x := by rw [Subsingleton.elim dec (by infer_instance)] have : (List.finRange n).indexOf i < n := by - simpa using List.indexOf_lt_length.2 (List.mem_finRange i) + simpa using List.indexOf_lt_length_iff.2 (List.mem_finRange i) simp [List.ofFn_eq_map, (List.nodup_finRange n).map_update, List.prod_set, this] variable {R A n} diff --git a/Mathlib/LinearAlgebra/PerfectPairing/Basic.lean b/Mathlib/LinearAlgebra/PerfectPairing/Basic.lean index 50ffa39ec68c9..6aaa7e342862c 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing/Basic.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing/Basic.lean @@ -204,6 +204,27 @@ protected lemma flip (h : p.IsPerfectCompl U V) : isCompl_left := h.isCompl_right isCompl_right := h.isCompl_left +@[simp] +protected lemma flip_iff : + p.flip.IsPerfectCompl V U ↔ p.IsPerfectCompl U V := + ⟨fun h ↦ h.flip, fun h ↦ h.flip⟩ + +@[simp] +lemma left_top_iff : + p.IsPerfectCompl ⊤ V ↔ V = ⊤ := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · exact eq_top_of_isCompl_bot <| by simpa using h.isCompl_right + · rw [h] + exact + { isCompl_left := by simpa using isCompl_top_bot + isCompl_right := by simpa using isCompl_top_bot } + +@[simp] +lemma right_top_iff : + p.IsPerfectCompl U ⊤ ↔ U = ⊤ := by + rw [← IsPerfectCompl.flip_iff] + exact left_top_iff + end IsPerfectCompl end PerfectPairing diff --git a/Mathlib/LinearAlgebra/Projection.lean b/Mathlib/LinearAlgebra/Projection.lean index c025f3b525548..1a1fee69243a9 100644 --- a/Mathlib/LinearAlgebra/Projection.lean +++ b/Mathlib/LinearAlgebra/Projection.lean @@ -367,12 +367,11 @@ structure IsProj {F : Type*} [FunLike F M M] (f : F) : Prop where map_mem : ∀ x, f x ∈ m map_id : ∀ x ∈ m, f x = x -theorem isProj_iff_idempotent (f : M →ₗ[S] M) : (∃ p : Submodule S M, IsProj p f) ↔ f ∘ₗ f = f := by +theorem isProj_iff_isIdempotentElem (f : M →ₗ[S] M) : + (∃ p : Submodule S M, IsProj p f) ↔ IsIdempotentElem f := by constructor - · intro h - obtain ⟨p, hp⟩ := h + · intro ⟨p, hp⟩ ext x - rw [comp_apply] exact hp.map_id (f x) (hp.map_mem x) · intro h use range f @@ -381,7 +380,9 @@ theorem isProj_iff_idempotent (f : M →ₗ[S] M) : (∃ p : Submodule S M, IsPr exact mem_range_self f x · intro x hx obtain ⟨y, hy⟩ := mem_range.1 hx - rw [← hy, ← comp_apply, h] + rw [← hy, ← mul_apply, h] + +@[deprecated (since := "2025-01-12")] alias isProj_iff_idempotent := isProj_iff_isIdempotentElem namespace IsProj diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 9778e86170c13..893c99b52d4fa 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -896,9 +896,13 @@ theorem associated_apply (x y : M) : rw [← LinearMap.smul_apply, nsmul_eq_mul, Nat.cast_ofNat, mul_invOf_self', LinearMap.one_apply, polar] -theorem associated_isSymm (Q : QuadraticForm R M) [Invertible (2 : R)] : - (associatedHom S Q).IsSymm := fun x y ↦ by - simp only [associated_apply, sub_eq_add_neg, add_assoc, RingHom.id_apply, add_comm, add_left_comm] +theorem associated_isSymm (Q : QuadraticMap R M N) (x y : M) : + associatedHom S Q x y = associatedHom S Q y x := by + simp only [associated_apply, sub_eq_add_neg, add_assoc, add_comm, add_left_comm] + +theorem _root_.QuadraticForm.associated_isSymm (Q : QuadraticForm R M) [Invertible (2 : R)] : + (associatedHom S Q).IsSymm := + QuadraticMap.associated_isSymm S Q /-- A version of `QuadraticMap.associated_isSymm` for general targets (using `flip` because `IsSymm` does not apply here). -/ @@ -919,11 +923,10 @@ theorem associated_toQuadraticMap (B : BilinMap R M N) (x y : M) : LinearMap.smul_def, _root_.map_sub] abel_nf -theorem associated_left_inverse [Invertible (2 : R)] {B₁ : BilinMap R M R} (h : B₁.IsSymm) : +theorem associated_left_inverse {B₁ : BilinMap R M N} (h : ∀ x y, B₁ x y = B₁ y x) : associatedHom S B₁.toQuadraticMap = B₁ := LinearMap.ext₂ fun x y ↦ by - rw [associated_toQuadraticMap, ← h.eq x y, RingHom.id_apply, ← two_mul, ← smul_eq_mul, - invOf_smul_eq_iff, two_smul, two_smul] + rw [associated_toQuadraticMap, ← h x y, ← two_smul R, invOf_smul_eq_iff, two_smul, two_smul] /-- A version of `QuadraticMap.associated_left_inverse` for general targets. -/ lemma associated_left_inverse' {B₁ : BilinMap R M N} (hB₁ : B₁.flip = B₁) : @@ -1190,10 +1193,10 @@ theorem QuadraticMap.toMatrix'_smul (a : R) (Q : QuadraticMap R (n → R) R) : (a • Q).toMatrix' = a • Q.toMatrix' := by simp only [toMatrix', LinearEquiv.map_smul, LinearMap.map_smul] -theorem QuadraticMap.isSymm_toMatrix' (Q : QuadraticMap R (n → R) R) : Q.toMatrix'.IsSymm := by +theorem QuadraticMap.isSymm_toMatrix' (Q : QuadraticForm R (n → R)) : Q.toMatrix'.IsSymm := by ext i j rw [toMatrix', Matrix.transpose_apply, LinearMap.toMatrix₂'_apply, LinearMap.toMatrix₂'_apply, - ← associated_isSymm, RingHom.id_apply, associated_apply] + ← associated_isSymm] end diff --git a/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean b/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean index 24854dd6d048c..da4ad9b49c356 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean @@ -34,7 +34,6 @@ variable [Module R M] [Module R M₁] [Module R M₂] [Module R M₃] [Module R /-- An isometric equivalence between two quadratic spaces `M₁, Q₁` and `M₂, Q₂` over a ring `R`, is a linear equivalence between `M₁` and `M₂` that commutes with the quadratic forms. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet @[nolint has_nonempty_instance] structure IsometryEquiv (Q₁ : QuadraticMap R M₁ N) (Q₂ : QuadraticMap R M₂ N) extends M₁ ≃ₗ[R] M₂ where map_app' : ∀ m, Q₂ (toFun m) = Q₁ m diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index c9ef8e47cf38d..561f9f6950e0b 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -280,7 +280,7 @@ def Isometry.proj [Fintype ι] [DecidableEq ι] (i : ι) (Q : QuadraticMap R (M rw [Pi.single_eq_of_ne hij, zero_apply] /-- Note that `QuadraticMap.Isometry.id` would not be well-typed as the RHS. -/ -@[simp, nolint simpNF] -- ignore the bogus "Left-hand side does not simplify" lint error +@[simp] theorem Isometry.proj_comp_single_of_same [Fintype ι] [DecidableEq ι] (i : ι) (Q : QuadraticMap R (Mᵢ i) P) : (proj i Q).comp (single _ i) = .ofEq (Pi.single_eq_same _ _) := diff --git a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean index d873b9641653c..7a26465a9f79c 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean @@ -66,6 +66,15 @@ protected noncomputable abbrev tmul (Q₁ : QuadraticMap A M₁ N₁) (Q₂ : QuadraticMap R M₂ N₂) : QuadraticMap A (M₁ ⊗[R] M₂) (N₁ ⊗[R] N₂) := tensorDistrib R A (Q₁ ⊗ₜ[R] Q₂) +theorem associated_tmul [Invertible (2 : A)] + (Q₁ : QuadraticMap A M₁ N₁) (Q₂ : QuadraticMap R M₂ N₂) : + (Q₁.tmul Q₂).associated = Q₁.associated.tmul Q₂.associated := by + letI : Invertible (2 : A) := (Invertible.map (algebraMap R A) 2).copy 2 (map_ofNat _ _).symm + rw [QuadraticMap.tmul, BilinMap.tmul] + have : Subsingleton (Invertible (2 : A)) := inferInstance + convert associated_left_inverse A (LinearMap.BilinMap.tmul_isSymm + (QuadraticMap.associated_isSymm A Q₁) (QuadraticMap.associated_isSymm R Q₂)) + end QuadraticMap namespace QuadraticForm @@ -96,26 +105,10 @@ protected noncomputable abbrev tmul (Q₁ : QuadraticForm A M₁) (Q₂ : Quadra tensorDistrib R A (Q₁ ⊗ₜ[R] Q₂) theorem associated_tmul [Invertible (2 : A)] (Q₁ : QuadraticForm A M₁) (Q₂ : QuadraticForm R M₂) : - associated (R := A) (Q₁.tmul Q₂) - = BilinForm.tmul ((associated (R := A) Q₁)) (associated (R := R) Q₂) := by - letI : Invertible (2 : A) := (Invertible.map (algebraMap R A) 2).copy 2 (map_ofNat _ _).symm - /- Previously `QuadraticForm.tensorDistrib` was defined in a similar way to - `QuadraticMap.tensorDistrib`. We now obtain the definition of `QuadraticForm.tensorDistrib` - from `QuadraticMap.tensorDistrib` using `A ⊗[R] R ≃ₗ[A] A`. Hypothesis `e1` below shows that the - new definition is equal to the old, and allows us to reuse the old proof. - - TODO: Define `IsSymm` for bilinear maps and generalise this result to Quadratic Maps. - -/ - have e1: (BilinMap.toQuadraticMapLinearMap A A (M₁ ⊗[R] M₂) ∘ - BilinForm.tensorDistrib R A (M₁ := M₁) (M₂ := M₂) ∘ - AlgebraTensorModule.map - (QuadraticMap.associated : QuadraticForm A M₁ →ₗ[A] BilinForm A M₁) - (QuadraticMap.associated : QuadraticForm R M₂ →ₗ[R] BilinForm R M₂)) = - tensorDistrib R A := rfl - rw [QuadraticForm.tmul, ← e1, BilinForm.tmul] - dsimp - have : Subsingleton (Invertible (2 : A)) := inferInstance - convert associated_left_inverse A ((associated_isSymm A Q₁).tmul (associated_isSymm R Q₂)) + (Q₁.tmul Q₂).associated = BilinForm.tmul Q₁.associated Q₂.associated := by + rw [BilinForm.tmul, BilinForm.tensorDistrib, LinearMap.comp_apply, ← BilinMap.tmul, + ← QuadraticMap.associated_tmul Q₁ Q₂] + aesop theorem polarBilin_tmul [Invertible (2 : A)] (Q₁ : QuadraticForm A M₁) (Q₂ : QuadraticForm R M₂) : polarBilin (Q₁.tmul Q₂) = ⅟(2 : A) • BilinForm.tmul (polarBilin Q₁) (polarBilin Q₂) := by diff --git a/Mathlib/LinearAlgebra/Quotient/Defs.lean b/Mathlib/LinearAlgebra/Quotient/Defs.lean index 810d76fdc6985..fc58414515648 100644 --- a/Mathlib/LinearAlgebra/Quotient/Defs.lean +++ b/Mathlib/LinearAlgebra/Quotient/Defs.lean @@ -143,12 +143,9 @@ section Module variable {S : Type*} --- Performance of `Function.Surjective.mulAction` is worse since it has to unify data to apply --- TODO: https://github.com/leanprover-community/mathlib4/pull/7432 instance mulAction' [Monoid S] [SMul S R] [MulAction S M] [IsScalarTower S R M] - (P : Submodule R M) : MulAction S (M ⧸ P) := - { Function.Surjective.mulAction mk Quot.mk_surjective <| Submodule.Quotient.mk_smul P with - toSMul := instSMul' _ } + (P : Submodule R M) : MulAction S (M ⧸ P) := fast_instance% + Function.Surjective.mulAction mk Quot.mk_surjective <| Submodule.Quotient.mk_smul P -- Porting note: should this be marked as a `@[default_instance]`? instance mulAction (P : Submodule R M) : MulAction R (M ⧸ P) := @@ -162,37 +159,28 @@ instance smulZeroClass' [SMul S R] [SMulZeroClass S M] [IsScalarTower S R M] (P instance smulZeroClass (P : Submodule R M) : SMulZeroClass R (M ⧸ P) := Quotient.smulZeroClass' P --- Performance of `Function.Surjective.distribSMul` is worse since it has to unify data to apply --- TODO: https://github.com/leanprover-community/mathlib4/pull/7432 instance distribSMul' [SMul S R] [DistribSMul S M] [IsScalarTower S R M] (P : Submodule R M) : - DistribSMul S (M ⧸ P) := - { Function.Surjective.distribSMul {toFun := mk, map_zero' := rfl, map_add' := fun _ _ => rfl} - Quot.mk_surjective (Submodule.Quotient.mk_smul P) with - toSMulZeroClass := smulZeroClass' _ } + DistribSMul S (M ⧸ P) := fast_instance% + Function.Surjective.distribSMul {toFun := mk, map_zero' := rfl, map_add' := fun _ _ => rfl} + Quot.mk_surjective (Submodule.Quotient.mk_smul P) -- Porting note: should this be marked as a `@[default_instance]`? instance distribSMul (P : Submodule R M) : DistribSMul R (M ⧸ P) := Quotient.distribSMul' P --- Performance of `Function.Surjective.distribMulAction` is worse since it has to unify data --- TODO: https://github.com/leanprover-community/mathlib4/pull/7432 instance distribMulAction' [Monoid S] [SMul S R] [DistribMulAction S M] [IsScalarTower S R M] - (P : Submodule R M) : DistribMulAction S (M ⧸ P) := - { Function.Surjective.distribMulAction {toFun := mk, map_zero' := rfl, map_add' := fun _ _ => rfl} - Quot.mk_surjective (Submodule.Quotient.mk_smul P) with - toMulAction := mulAction' _ } + (P : Submodule R M) : DistribMulAction S (M ⧸ P) := fast_instance% + Function.Surjective.distribMulAction {toFun := mk, map_zero' := rfl, map_add' := fun _ _ => rfl} + Quot.mk_surjective (Submodule.Quotient.mk_smul P) -- Porting note: should this be marked as a `@[default_instance]`? instance distribMulAction (P : Submodule R M) : DistribMulAction R (M ⧸ P) := Quotient.distribMulAction' P --- Performance of `Function.Surjective.module` is worse since it has to unify data to apply --- TODO: https://github.com/leanprover-community/mathlib4/pull/7432 instance module' [Semiring S] [SMul S R] [Module S M] [IsScalarTower S R M] (P : Submodule R M) : - Module S (M ⧸ P) := - { Function.Surjective.module _ {toFun := mk, map_zero' := by rfl, map_add' := fun _ _ => by rfl} - Quot.mk_surjective (Submodule.Quotient.mk_smul P) with - toDistribMulAction := distribMulAction' _ } + Module S (M ⧸ P) := fast_instance% + Function.Surjective.module _ {toFun := mk, map_zero' := by rfl, map_add' := fun _ _ => by rfl} + Quot.mk_surjective (Submodule.Quotient.mk_smul P) -- Porting note: should this be marked as a `@[default_instance]`? instance module (P : Submodule R M) : Module R (M ⧸ P) := diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index c460672f898e7..b545070466409 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -180,7 +180,6 @@ theorem add_right (hy : SameRay R x y) (hz : SameRay R x z) : SameRay R x (y + z end SameRay --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed has_nonempty_instance nolint, no such linter set_option linter.unusedVariables false in /-- Nonzero vectors, as used to define rays. This type depends on an unused argument `R` so that `RayVector.Setoid` can be an instance. -/ @@ -205,7 +204,6 @@ instance RayVector.Setoid : Setoid (RayVector R M) where exact hxy.trans hyz fun hy => (y.2 hy).elim⟩ /-- A ray (equivalence class of nonzero vectors with common positive multiples) in a module. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed has_nonempty_instance nolint, no such linter def Module.Ray := Quotient (RayVector.Setoid R M) diff --git a/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean b/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean new file mode 100644 index 0000000000000..60e98abdd2d75 --- /dev/null +++ b/Mathlib/LinearAlgebra/RootSystem/BaseChange.lean @@ -0,0 +1,128 @@ +/- +Copyright (c) 2025 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.Algebra.Algebra.Rat +import Mathlib.LinearAlgebra.PerfectPairing.Restrict +import Mathlib.LinearAlgebra.RootSystem.Defs + +/-! +# Base change for root pairings + +When the coefficients are a field, root pairings behave well with respect to restriction and +extension of scalars. + +## Main results: + * `RootPairing.restrict`: if `RootPairing.pairing` takes values in a subfield, we may restrict to + get a root _system_ with coefficients in the subfield. Of particular interest is the case when + the pairing takes values in its prime subfield (which happens for crystallographic pairings). + +## TODO + + * Extension of scalars + +-/ + +noncomputable section + +open Set Function +open Submodule (span injective_subtype span subset_span span_setOf_mem_eq_top) + +namespace RootPairing + +/-- We say a root pairing is balanced if the root span and coroot span are perfectly +complementary. + +All root systems are balanced and all finite root pairings over a field are balanced. -/ +class IsBalanced {ι R M N : Type*} [AddCommGroup M] [AddCommGroup N] + [CommRing R] [Module R M] [Module R N] (P : RootPairing ι R M N) : Prop where + isPerfectCompl : P.toPerfectPairing.IsPerfectCompl P.rootSpan P.corootSpan + +instance {ι R M N : Type*} [AddCommGroup M] [AddCommGroup N] + [CommRing R] [Module R M] [Module R N] (P : RootSystem ι R M N) : + P.IsBalanced where + isPerfectCompl := by simp + +variable {ι L M N : Type*} + [Field L] [AddCommGroup M] [AddCommGroup N] [Module L M] [Module L N] + (P : RootPairing ι L M N) + +section restrictScalars + +variable (K : Type*) [Field K] [Algebra K L] + [Module K M] [Module K N] [IsScalarTower K L M] [IsScalarTower K L N] + [P.IsBalanced] + +section SubfieldValued + +variable (hP : ∀ i j, P.pairing i j ∈ (algebraMap K L).range) + +/-- Restriction of scalars for a root pairing taking values in a subfield. + +Note that we obtain a root system (not just a root pairing). See also +`RootPairing.restrictScalars`. -/ +def restrictScalars' : + RootSystem ι K (span K (range P.root)) (span K (range P.coroot)) := + { toPerfectPairing := (P.toPerfectPairing.restrictScalarsField _ _ + (injective_subtype _) (injective_subtype _) (by simpa using IsBalanced.isPerfectCompl) + (fun x y ↦ LinearMap.BilinMap.apply_apply_mem_of_mem_span + (LinearMap.range (Algebra.linearMap K L)) (range P.root) (range P.coroot) + ((LinearMap.restrictScalarsₗ K L _ _ _) ∘ₗ (P.toPerfectPairing.toLin.restrictScalars K)) + (by rintro - ⟨i, rfl⟩ - ⟨j, rfl⟩; exact hP i j) _ _ x.property y.property)) + root := ⟨fun i ↦ ⟨_, subset_span (mem_range_self i)⟩, fun i j h ↦ by simpa using h⟩ + coroot := ⟨fun i ↦ ⟨_, subset_span (mem_range_self i)⟩, fun i j h ↦ by simpa using h⟩ + root_coroot_two i := by + have : algebraMap K L 2 = 2 := by + rw [← Int.cast_two (R := K), ← Int.cast_two (R := L), map_intCast] + exact NoZeroSMulDivisors.algebraMap_injective K L <| by simp [this] + reflection_perm := P.reflection_perm + reflection_perm_root i j := by + ext; simpa [algebra_compatible_smul L] using P.reflection_perm_root i j + reflection_perm_coroot i j := by + ext; simpa [algebra_compatible_smul L] using P.reflection_perm_coroot i j + span_root_eq_top := by + rw [← span_setOf_mem_eq_top] + congr + ext ⟨x, hx⟩ + simp + span_coroot_eq_top := by + rw [← span_setOf_mem_eq_top] + congr + ext ⟨x, hx⟩ + simp } + +@[simp] lemma restrictScalars_toPerfectPairing_apply_apply + (x : span K (range P.root)) (y : span K (range P.coroot)) : + algebraMap K L ((P.restrictScalars' K hP).toPerfectPairing x y) = P.toPerfectPairing x y := by + simp [restrictScalars'] + +@[simp] lemma restrictScalars_coe_root (i : ι) : + (P.restrictScalars' K hP).root i = P.root i := + rfl + +@[simp] lemma restrictScalars_coe_coroot (i : ι) : + (P.restrictScalars' K hP).coroot i = P.coroot i := + rfl + +@[simp] lemma restrictScalars_pairing (i j : ι) : + algebraMap K L ((P.restrictScalars' K hP).pairing i j) = P.pairing i j := by + simp only [pairing, restrictScalars_toPerfectPairing_apply_apply, restrictScalars_coe_root, + restrictScalars_coe_coroot] + +end SubfieldValued + +/-- Restriction of scalars for a crystallographic root pairing. -/ +abbrev restrictScalars [P.IsCrystallographic] : + RootSystem ι K (span K (range P.root)) (span K (range P.coroot)) := + P.restrictScalars' K (IsValuedIn.trans P K ℤ).exists_value + +/-- Restriction of scalars to `ℚ` for a crystallographic root pairing in characteristic zero. -/ +abbrev restrictScalarsRat [CharZero L] [P.IsCrystallographic] := + let _i : Module ℚ M := Module.compHom M (algebraMap ℚ L) + let _i : Module ℚ N := Module.compHom N (algebraMap ℚ L) + P.restrictScalars ℚ + +end restrictScalars + +end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean b/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean new file mode 100644 index 0000000000000..3c23104e4e2bb --- /dev/null +++ b/Mathlib/LinearAlgebra/RootSystem/CartanMatrix.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2025 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.LinearAlgebra.RootSystem.Base + +/-! +# Cartan matrices for root systems + +This file contains definitions and basic results about Cartan matrices of root pairings / systems. + +## Main definitions: + * `RootPairing.cartanMatrix`: the Cartan matrix of a crystallographic root pairing, with respect to + a base `b`. + +-/ + +noncomputable section + +open Function Set + +variable {ι R M N : Type*} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] + +namespace RootPairing.Base + +variable (S : Type*) [CommRing S] [Algebra S R] + {P : RootPairing ι R M N} [P.IsValuedIn S] (b : P.Base) + +/-- The Cartan matrix of a root pairing, taking values in `S`, with respect to a base `b`. + +See also `RootPairing.Base.cartanMatrix`. -/ +def cartanMatrixIn : + Matrix b.support b.support S := + .of fun i j ↦ P.pairingIn S i j + +/-- The Cartan matrix of a crystallographic root pairing, with respect to a base `b`. -/ +abbrev cartanMatrix [P.IsCrystallographic] : + Matrix b.support b.support ℤ := + b.cartanMatrixIn ℤ + +lemma cartanMatrixIn_def (i j : b.support) : + b.cartanMatrixIn S i j = P.pairingIn S i j := + rfl + +variable [Nontrivial R] [NoZeroSMulDivisors S R] + +@[simp] +lemma cartanMatrixIn_apply_same (i : b.support) : + b.cartanMatrixIn S i i = 2 := + NoZeroSMulDivisors.algebraMap_injective S R <| by simp [cartanMatrixIn_def, map_ofNat] + +end RootPairing.Base diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 8b01626b5d947..7ca876bc7654e 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -380,22 +380,60 @@ lemma pairing_reflection_perm_self_right (i j : ι) : sub_add_cancel_left, ← toLin_toPerfectPairing, map_neg, toLin_toPerfectPairing, root_coroot_eq_pairing] +/-- If `R` is an `S`-algebra, a root pairing over `R` is said to be valued in `S` if the pairing +between a root and coroot always belongs to `S`. + +Of particular interest is the case `S = ℤ`. See `RootPairing.IsCrystallographic`. -/ +@[mk_iff] +class IsValuedIn (S : Type*) [CommRing S] [Algebra S R] : Prop where + exists_value : ∀ i j, ∃ s, algebraMap S R s = P.pairing i j + +protected alias exists_value := IsValuedIn.exists_value + /-- A root pairing is said to be crystallographic if the pairing between a root and coroot is always an integer. -/ -class IsCrystallographic : Prop where - exists_int : ∀ i j, ∃ z : ℤ, z = P.pairing i j +abbrev IsCrystallographic := P.IsValuedIn ℤ + +section IsValuedIn + +instance : P.IsValuedIn R where + exists_value i j := by simp + +variable (S : Type*) [CommRing S] [Algebra S R] + +variable {S} in +lemma isValuedIn_iff_mem_range : + P.IsValuedIn S ↔ ∀ i j, P.pairing i j ∈ range (algebraMap S R) := by + simp only [isValuedIn_iff, mem_range] -protected lemma exists_int [P.IsCrystallographic] (i j : ι) : - ∃ z : ℤ, z = P.pairing i j := - IsCrystallographic.exists_int i j +instance : P.IsValuedIn R where + exists_value := by simp -lemma isCrystallographic_iff : - P.IsCrystallographic ↔ ∀ i j, ∃ z : ℤ, z = P.pairing i j := - ⟨fun ⟨h⟩ ↦ h, fun h ↦ ⟨h⟩⟩ +instance [P.IsValuedIn S] : P.flip.IsValuedIn S := by + rw [isValuedIn_iff, forall_comm] + exact P.exists_value -instance [P.IsCrystallographic] : P.flip.IsCrystallographic := by - rw [isCrystallographic_iff, forall_comm] - exact P.exists_int +/-- A variant of `RootPairing.pairing` for root pairings which are valued in a smaller set of +coefficients. + +Note that it is uniquely-defined only when the map `S → R` is injective, i.e., when we have +`[NoZeroSMulDivisors S R]`. -/ +def pairingIn [P.IsValuedIn S] (i j : ι) : S := + (P.exists_value i j).choose + +@[simp] +lemma algebraMap_pairingIn [P.IsValuedIn S] (i j : ι) : + algebraMap S R (P.pairingIn S i j) = P.pairing i j := + (P.exists_value i j).choose_spec + +lemma IsValuedIn.trans (T : Type*) [CommRing T] [Algebra T S] [Algebra T R] [IsScalarTower T S R] + [P.IsValuedIn T] : + P.IsValuedIn S where + exists_value i j := by + use algebraMap T S (P.pairingIn T i j) + simp [← RingHom.comp_apply, ← IsScalarTower.algebraMap_eq T S R] + +end IsValuedIn /-- A root pairing is said to be reduced if any linearly dependent pair of roots is related by a sign. -/ @@ -540,11 +578,18 @@ 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 [P.IsCrystallographic] (i j : ι) : - ∃ z : ℤ, P.coxeterWeight i j = z := by - obtain ⟨a, ha⟩ := P.exists_int i j - obtain ⟨b, hb⟩ := P.exists_int j i - exact ⟨a * b, by simp [coxeterWeight, ha, hb]⟩ +/-- A variant of `RootPairing.coxeterWeight` for root pairings which are valued in a smaller set of +coefficients. + +Note that it is uniquely-defined only when the map `S → R` is injective, i.e., when we have +`[NoZeroSMulDivisors S R]`. -/ +def coxeterWeightIn (S : Type*) [CommRing S] [Algebra S R] [P.IsValuedIn S] (i j : ι) : S := + P.pairingIn S i j * P.pairingIn S j i + +@[simp] lemma algebraMap_coxeterWeightIn (S : Type*) [CommRing S] [Algebra S R] [P.IsValuedIn S] + (i j : ι) : + algebraMap S R (P.coxeterWeightIn S i j) = P.coxeterWeight i j := by + simp [coxeterWeightIn, coxeterWeight] /-- 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 39b0b06f3214c..e2e30db0f2b7c 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -268,10 +268,10 @@ instance instIsRootPositiveRootForm : IsRootPositive P P.RootForm where lemma coxeterWeight_mem_set_of_isCrystallographic (i j : ι) [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 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 : 0 ≤ P.coxeterWeightIn ℤ i j := by + simpa [← P.algebraMap_coxeterWeightIn ℤ] using P.coxeterWeight_non_neg P.RootForm i j + obtain ⟨n, hn⟩ := Int.eq_ofNat_of_zero_le this + exact ⟨n, by simp [← P.algebraMap_coxeterWeightIn ℤ, hn]⟩ 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 ⊢ diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 1d04a0c2f32ce..aa86853fa0777 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -7,6 +7,7 @@ import Mathlib.LinearAlgebra.BilinearForm.Basic import Mathlib.LinearAlgebra.BilinearForm.Orthogonal import Mathlib.LinearAlgebra.Dimension.Localization import Mathlib.LinearAlgebra.QuadraticForm.Basic +import Mathlib.LinearAlgebra.RootSystem.BaseChange import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear /-! @@ -76,7 +77,8 @@ instance [P.IsAnisotropic] : P.flip.IsAnisotropic where /-- An auxiliary lemma en route to `RootPairing.instIsAnisotropicOfIsCrystallographic`. -/ private lemma rootForm_root_ne_zero_aux [CharZero R] [P.IsCrystallographic] (i : ι) : P.RootForm (P.root i) (P.root i) ≠ 0 := by - choose z hz using P.exists_int i + choose z hz using P.exists_value (S := ℤ) i + simp_rw [algebraMap_int_eq, Int.coe_castRingHom] at hz simp only [rootForm_apply_apply, PerfectPairing.flip_apply_apply, root_coroot_eq_pairing, ← hz] suffices 0 < ∑ i, z i * z i by norm_cast; exact this.ne' refine Finset.sum_pos' (fun i _ ↦ mul_self_nonneg (z i)) ⟨i, Finset.mem_univ i, ?_⟩ @@ -159,6 +161,30 @@ lemma isCompl_corootSpan_ker_corootForm : IsCompl P.corootSpan (LinearMap.ker P.CorootForm) := P.flip.isCompl_rootSpan_ker_rootForm +lemma ker_rootForm_eq_dualAnnihilator : + LinearMap.ker P.RootForm = P.corootSpan.dualAnnihilator.map P.toDualLeft.symm := by + have _iM : IsReflexive R M := PerfectPairing.reflexive_left P.toPerfectPairing + have _iN : IsReflexive R N := PerfectPairing.reflexive_right P.toPerfectPairing + suffices finrank R (LinearMap.ker P.RootForm) = finrank R P.corootSpan.dualAnnihilator by + refine (Submodule.eq_of_le_of_finrank_eq P.corootSpan_dualAnnihilator_le_ker_rootForm ?_).symm + rw [this] + apply LinearEquiv.finrank_map_eq + have aux0 := Subspace.finrank_add_finrank_dualAnnihilator_eq P.corootSpan + have aux1 := Submodule.finrank_add_eq_of_isCompl P.isCompl_rootSpan_ker_rootForm + rw [← P.finrank_corootSpan_eq, P.toPerfectPairing.finrank_eq] at aux1 + omega + +lemma ker_corootForm_eq_dualAnnihilator : + LinearMap.ker P.CorootForm = P.rootSpan.dualAnnihilator.map P.toDualRight.symm := + P.flip.ker_rootForm_eq_dualAnnihilator + +instance : P.IsBalanced where + isPerfectCompl := + { isCompl_left := by + simpa only [ker_rootForm_eq_dualAnnihilator] using P.isCompl_rootSpan_ker_rootForm + isCompl_right := by + simpa only [ker_corootForm_eq_dualAnnihilator] using P.isCompl_corootSpan_ker_corootForm } + /-- See also `RootPairing.rootForm_restrict_nondegenerate_of_ordered`. Note that this applies to crystallographic root systems in characteristic zero via diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 504d2acfe7e8e..8657bbed12d1d 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -79,11 +79,7 @@ theorem isOrthoᵢ_def {B : M₁ →ₛₗ[I₁] M₁ →ₛₗ[I₁'] M} {v : n theorem isOrthoᵢ_flip (B : M₁ →ₛₗ[I₁] M₁ →ₛₗ[I₁'] M) {v : n → M₁} : B.IsOrthoᵢ v ↔ B.flip.IsOrthoᵢ v := by simp_rw [isOrthoᵢ_def] - constructor <;> intro h i j hij - · rw [flip_apply] - exact h j i (Ne.symm hij) - simp_rw [flip_apply] at h - exact h j i (Ne.symm hij) + constructor <;> exact fun h i j hij ↦ h j i hij.symm end CommRing @@ -217,13 +213,12 @@ end IsSymm @[simp] theorem isSymm_zero : (0 : M →ₛₗ[I] M →ₗ[R] R).IsSymm := fun _ _ => map_zero _ -theorem isSymm_iff_eq_flip {B : LinearMap.BilinForm R M} : B.IsSymm ↔ B = B.flip := by - constructor <;> intro h - · ext - rw [← h, flip_apply, RingHom.id_apply] - intro x y - conv_lhs => rw [h] - rfl +theorem BilinMap.isSymm_iff_eq_flip {N : Type*} [AddCommMonoid N] [Module R N] + {B : LinearMap.BilinMap R M N} : (∀ x y, B x y = B y x) ↔ B = B.flip := by + simp [LinearMap.ext_iff₂] + +theorem isSymm_iff_eq_flip {B : LinearMap.BilinForm R M} : B.IsSymm ↔ B = B.flip := + BilinMap.isSymm_iff_eq_flip end Symmetric diff --git a/Mathlib/LinearAlgebra/Span/Basic.lean b/Mathlib/LinearAlgebra/Span/Basic.lean index 145579823242b..ea92a62c9cc47 100644 --- a/Mathlib/LinearAlgebra/Span/Basic.lean +++ b/Mathlib/LinearAlgebra/Span/Basic.lean @@ -191,6 +191,9 @@ theorem span_span_of_tower : span S (span R s : Set M) = span S s := le_antisymm (span_le.2 <| span_subset_span R S s) (span_mono subset_span) +theorem span_eq_top_of_span_eq_top (s : Set M) (hs : span R s = ⊤) : span S s = ⊤ := + le_top.antisymm (hs.ge.trans (span_le_restrictScalars R S s)) + variable {R S} in lemma span_range_inclusion_eq_top (p : Submodule R M) (q : Submodule S M) (h₁ : p ≤ q.restrictScalars R) (h₂ : q ≤ span S p) : @@ -584,11 +587,17 @@ theorem span_singleton_eq_range (x : M) : (R ∙ x) = range (toSpanSingleton R M theorem toSpanSingleton_one (x : M) : toSpanSingleton R M x 1 = x := one_smul _ _ +theorem toSpanSingleton_injective : Function.Injective (toSpanSingleton R M) := + fun _ _ eq ↦ by simpa using congr($eq 1) + @[simp] theorem toSpanSingleton_zero : toSpanSingleton R M 0 = 0 := by ext simp +theorem toSpanSingleton_eq_zero_iff {x : M} : toSpanSingleton R M x = 0 ↔ x = 0 := by + rw [← toSpanSingleton_zero, (toSpanSingleton_injective R M).eq_iff] + variable {R M} theorem toSpanSingleton_isIdempotentElem_iff {e : R} : diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 56e15a7a87dc5..a9d2a671964dd 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -426,6 +426,14 @@ theorem rec_heq_iff_heq {α β : Sort _} {a b : α} {C : α → Sort*} {x : C a} theorem heq_rec_iff_heq {α β : Sort _} {a b : α} {C : α → Sort*} {x : β} {y : C a} {e : a = b} : HEq x (e ▸ y) ↔ HEq x y := by subst e; rfl +@[simp] +theorem cast_heq_iff_heq {α β γ : Sort _} (e : α = β) (a : α) (c : γ) : + HEq (cast e a) c ↔ HEq a c := by subst e; rfl + +@[simp] +theorem heq_cast_iff_heq {α β γ : Sort _} (e : β = γ) (a : α) (b : β) : + HEq a (cast e b) ↔ HEq a b := by subst e; rfl + universe u variable {α β : Sort u} {e : β = α} {a : α} {b : β} diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index a77e11e473f68..4e516587e5f5a 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -15,8 +15,6 @@ universe u v w x namespace Function --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] /-- `α ↪ β` is a bundled injective function. -/ structure Embedding (α : Sort*) (β : Sort*) where /-- An embedding as a function. Use coercion instead. -/ diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 52c2e153e76e9..42d71d763a324 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -215,6 +215,23 @@ theorem sigmaUnique_symm_apply {α} {β : α → Type*} [∀ a, Unique (β a)] ( (sigmaUnique α β).symm x = ⟨x, default⟩ := rfl +/-- Any `Unique` type is a left identity for type sigma up to equivalence. Compare with `uniqueProd` +which is non-dependent. -/ +def uniqueSigma {α} (β : α → Type*) [Unique α] : (i:α) × β i ≃ β default := + ⟨fun p ↦ (Unique.eq_default _).rec p.2, + fun b ↦ ⟨default, b⟩, + fun _ ↦ Sigma.ext (Unique.default_eq _) (eqRec_heq _ _), + fun _ ↦ rfl⟩ + +theorem uniqueSigma_apply {α} {β : α → Type*} [Unique α] (x : (a : α) × β a) : + uniqueSigma β x = (Unique.eq_default _).rec x.2 := + rfl + +@[simp] +theorem uniqueSigma_symm_apply {α} {β : α → Type*} [Unique α] (y : β default) : + (uniqueSigma β).symm y = ⟨default, y⟩ := + rfl + /-- `Empty` type is a right absorbing element for type product up to an equivalence. -/ def prodEmpty (α) : α × Empty ≃ Empty := equivEmpty _ @@ -286,12 +303,12 @@ theorem sumCongr_refl {α β} : /-- A subtype of a sum is equivalent to a sum of subtypes. -/ def subtypeSum {α β} {p : α ⊕ β → Prop} : {c // p c} ≃ {a // p (Sum.inl a)} ⊕ {b // p (Sum.inr b)} where - toFun c := match h : c.1 with - | Sum.inl a => Sum.inl ⟨a, h ▸ c.2⟩ - | Sum.inr b => Sum.inr ⟨b, h ▸ c.2⟩ - invFun c := match c with - | Sum.inl a => ⟨Sum.inl a, a.2⟩ - | Sum.inr b => ⟨Sum.inr b, b.2⟩ + toFun + | ⟨.inl a, h⟩ => .inl ⟨a, h⟩ + | ⟨.inr b, h⟩ => .inr ⟨b, h⟩ + invFun + | .inl a => ⟨.inl a, a.2⟩ + | .inr b => ⟨.inr b, b.2⟩ left_inv := by rintro ⟨a | b, h⟩ <;> rfl right_inv := by rintro (a | b) <;> rfl @@ -919,7 +936,8 @@ theorem prodSumDistrib_symm_apply_right {α β γ} (a : α × γ) : (prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) := rfl -/-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. -/ +/-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. Compare +with `Equiv.sumSigmaDistrib` which is indexed by sums. -/ @[simps] def sigmaSumDistrib {ι} (α β : ι → Type*) : (Σ i, α i ⊕ β i) ≃ (Σ i, α i) ⊕ (Σ i, β i) := @@ -927,6 +945,18 @@ def sigmaSumDistrib {ι} (α β : ι → Type*) : Sum.elim (Sigma.map id fun _ => Sum.inl) (Sigma.map id fun _ => Sum.inr), fun p => by rcases p with ⟨i, a | b⟩ <;> rfl, fun p => by rcases p with (⟨i, a⟩ | ⟨i, b⟩) <;> rfl⟩ +/-- A type indexed by disjoint sums of types is equivalent to the sum of the sums. Compare with +`Equiv.sigmaSumDistrib` which has the sums as the output type. -/ +@[simps] +def sumSigmaDistrib {α β} (t : α ⊕ β → Type*) : + (Σ i, t i) ≃ (Σ i, t (.inl i)) ⊕ (Σ i, t (.inr i)) := + ⟨(match · with + | .mk (.inl x) y => .inl ⟨x, y⟩ + | .mk (.inr x) y => .inr ⟨x, y⟩), + Sum.elim (fun a ↦ ⟨.inl a.1, a.2⟩) (fun b ↦ ⟨.inr b.1, b.2⟩), + by rintro ⟨x|x,y⟩ <;> simp, + by rintro (⟨x,y⟩|⟨x,y⟩) <;> simp⟩ + /-- The product of an indexed sum of types (formally, a `Sigma`-type `Σ i, α i`) by a type `β` is equivalent to the sum of products `Σ i, (α i × β)`. -/ @[simps apply symm_apply] diff --git a/Mathlib/Logic/Equiv/PartialEquiv.lean b/Mathlib/Logic/Equiv/PartialEquiv.lean index ff661bd6d8b3e..ca658fa24c549 100644 --- a/Mathlib/Logic/Equiv/PartialEquiv.lean +++ b/Mathlib/Logic/Equiv/PartialEquiv.lean @@ -574,7 +574,9 @@ protected def trans' (e' : PartialEquiv β γ) (h : e.target = e'.source) : Part right_inv' y hy := by simp [hy, h] /-- Composing two partial equivs, by restricting to the maximal domain where their composition -is well defined. -/ +is well defined. +Within the `Manifold` namespace, there is the notation `e ≫ f` for this. +-/ @[trans] protected def trans : PartialEquiv α γ := PartialEquiv.trans' (e.symm.restr e'.source).symm (e'.restr e.target) (inter_comm _ _) diff --git a/Mathlib/Logic/IsEmpty.lean b/Mathlib/Logic/IsEmpty.lean index fa48bd9fc85c3..d6c0c75306607 100644 --- a/Mathlib/Logic/IsEmpty.lean +++ b/Mathlib/Logic/IsEmpty.lean @@ -213,10 +213,22 @@ variable {α β : Type*} (R : α → β → Prop) theorem leftTotal_empty [IsEmpty α] : LeftTotal R := by simp only [LeftTotal, IsEmpty.forall_iff] +theorem leftTotal_iff_isEmpty_left [IsEmpty β] : LeftTotal R ↔ IsEmpty α := by + simp only [LeftTotal, IsEmpty.exists_iff, isEmpty_iff] + @[simp] theorem rightTotal_empty [IsEmpty β] : RightTotal R := by simp only [RightTotal, IsEmpty.forall_iff] +theorem rightTotal_iff_isEmpty_right [IsEmpty α] : RightTotal R ↔ IsEmpty β := by + simp only [RightTotal, IsEmpty.exists_iff, isEmpty_iff, imp_self] + @[simp] theorem biTotal_empty [IsEmpty α] [IsEmpty β] : BiTotal R := ⟨leftTotal_empty R, rightTotal_empty R⟩ + +theorem biTotal_iff_isEmpty_right [IsEmpty α] : BiTotal R ↔ IsEmpty β := by + simp only [BiTotal, leftTotal_empty, rightTotal_iff_isEmpty_right, true_and] + +theorem biTotal_iff_isEmpty_left [IsEmpty β] : BiTotal R ↔ IsEmpty α := by + simp only [BiTotal, leftTotal_iff_isEmpty_left, rightTotal_empty, and_true] diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index e49b8327144ca..e2b92731ce5a2 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -41,6 +41,8 @@ for good definitional properties of the default term. universe u v w +-- Don't generate injectivity lemmas, which the `simpNF` linter will complain about. +set_option genInjectivity false in /-- `Unique α` expresses that `α` is a type with a unique term `default`. This is implemented as a type, rather than a `Prop`-valued predicate, @@ -51,8 +53,6 @@ structure Unique (α : Sort u) extends Inhabited α where uniq : ∀ a : α, a = default attribute [class] Unique --- The simplifier can already prove this using `eq_iff_true_of_subsingleton` -attribute [nolint simpNF] Unique.mk.injEq theorem unique_iff_existsUnique (α : Sort u) : Nonempty (Unique α) ↔ ∃! _ : α, True := ⟨fun ⟨u⟩ ↦ ⟨u.default, trivial, fun a _ ↦ u.uniq a⟩, @@ -84,11 +84,7 @@ instance PUnit.instUnique : Unique PUnit.{u} where default := PUnit.unit uniq x := subsingleton x _ --- Porting note: --- This should not require a nolint, --- but it is currently failing due to a problem in the linter discussed at --- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.60simpNF.60.20error.20.22unknown.20metavariable.22 -@[simp, nolint simpNF] +@[simp] theorem PUnit.default_eq_unit : (default : PUnit) = PUnit.unit := rfl diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 5b2934b6db6d8..61f42f43718a6 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -310,7 +310,7 @@ theorem exists_disjoint_covering_ae linarith [Idist_v b ⟨bu, hbx⟩] -- we will show that, in `ball x (R x)`, almost all `s` is covered by the family `u`. refine ⟨_ ∩ ball x (R x), inter_mem_nhdsWithin _ (ball_mem_nhds _ (hR0 _)), - nonpos_iff_eq_zero.mp (le_of_forall_le_of_dense fun ε εpos => ?_)⟩ + nonpos_iff_eq_zero.mp (le_of_forall_gt_imp_ge_of_dense fun ε εpos => ?_)⟩ -- the elements of `v` are disjoint and all contained in a finite volume ball, hence the sum -- of their measures is finite. have I : (∑' a : v, μ (B a)) < ∞ := by diff --git a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean index 7671daaa30640..6d40a1d80bc52 100644 --- a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean +++ b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean @@ -61,8 +61,6 @@ Vitali families are provided by covering theorems such as the Besicovitch coveri Vitali covering theorem. They make it possible to formulate general versions of theorems on differentiations of measure that apply in both contexts. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure VitaliFamily {m : MeasurableSpace X} (μ : Measure X) where /-- Sets of the family "centered" at a given point. -/ setsAt : X → Set (Set X) diff --git a/Mathlib/MeasureTheory/Decomposition/IntegralRNDeriv.lean b/Mathlib/MeasureTheory/Decomposition/IntegralRNDeriv.lean new file mode 100644 index 0000000000000..e2ddf8630cf97 --- /dev/null +++ b/Mathlib/MeasureTheory/Decomposition/IntegralRNDeriv.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2025 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne, Lorenzo Luccioli +-/ +import Mathlib.Analysis.Convex.Continuous +import Mathlib.Analysis.Convex.Integral +import Mathlib.MeasureTheory.Decomposition.RadonNikodym + +/-! +# Integrals of functions of Radon-Nikodym derivatives + +## Main statements + +* `mul_le_integral_rnDeriv_of_ac`: for a convex continuous function `f` on `[0, ∞)`, if `μ` + is absolutely continuous with respect to `ν`, then + `(ν univ).toReal * f ((μ univ).toReal / (ν univ).toReal) ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν`. + +-/ + + +open Set + +namespace MeasureTheory + +variable {α : Type*} {mα : MeasurableSpace α} {μ ν : Measure α} {f : ℝ → ℝ} + +/-- For a convex continuous function `f` on `[0, ∞)`, if `μ` is absolutely continuous +with respect to a probability measure `ν`, then +`f (μ univ).toReal ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν`. -/ +lemma le_integral_rnDeriv_of_ac [IsFiniteMeasure μ] [IsProbabilityMeasure ν] + (hf_cvx : ConvexOn ℝ (Ici 0) f) (hf_cont : ContinuousWithinAt f (Ici 0) 0) + (hf_int : Integrable (fun x ↦ f (μ.rnDeriv ν x).toReal) ν) (hμν : μ ≪ ν) : + f (μ univ).toReal ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν := by + have hf_cont' : ContinuousOn f (Ici 0) := by + intro x hx + rcases eq_or_lt_of_le (α := ℝ) (hx : 0 ≤ x) with rfl | hx_pos + · exact hf_cont + · have h := hf_cvx.continuousOn_interior x + simp only [nonempty_Iio, interior_Ici', mem_Ioi] at h + rw [continuousWithinAt_iff_continuousAt (Ioi_mem_nhds hx_pos)] at h + exact (h hx_pos).continuousWithinAt + calc f (μ univ).toReal + = f (∫ x, (μ.rnDeriv ν x).toReal ∂ν) := by rw [Measure.integral_toReal_rnDeriv hμν] + _ ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν := by + rw [← average_eq_integral, ← average_eq_integral] + exact ConvexOn.map_average_le hf_cvx hf_cont' isClosed_Ici (by simp) + Measure.integrable_toReal_rnDeriv hf_int + +/-- For a convex continuous function `f` on `[0, ∞)`, if `μ` is absolutely continuous +with respect to `ν`, then +`(ν univ).toReal * f ((μ univ).toReal / (ν univ).toReal) ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν`. -/ +lemma mul_le_integral_rnDeriv_of_ac [IsFiniteMeasure μ] [IsFiniteMeasure ν] + (hf_cvx : ConvexOn ℝ (Ici 0) f) (hf_cont : ContinuousWithinAt f (Ici 0) 0) + (hf_int : Integrable (fun x ↦ f (μ.rnDeriv ν x).toReal) ν) (hμν : μ ≪ ν) : + (ν univ).toReal * f ((μ univ).toReal / (ν univ).toReal) + ≤ ∫ x, f (μ.rnDeriv ν x).toReal ∂ν := by + by_cases hν : ν = 0 + · simp [hν] + have : NeZero ν := ⟨hν⟩ + let μ' := (ν univ)⁻¹ • μ + let ν' := (ν univ)⁻¹ • ν + have : IsFiniteMeasure μ' := μ.smul_finite (by simp [hν]) + have hμν' : μ' ≪ ν' := hμν.smul _ + have h_rnDeriv_eq : μ'.rnDeriv ν' =ᵐ[ν] μ.rnDeriv ν := by + have h1' : μ'.rnDeriv ν' =ᵐ[ν'] (ν univ)⁻¹ • μ.rnDeriv ν' := + Measure.rnDeriv_smul_left_of_ne_top' (μ := ν') (ν := μ) (by simp [hν]) + have h1 : μ'.rnDeriv ν' =ᵐ[ν] (ν univ)⁻¹ • μ.rnDeriv ν' := by + rwa [Measure.ae_smul_measure_eq] at h1' + simp + have h2 : μ.rnDeriv ν' =ᵐ[ν] (ν univ)⁻¹⁻¹ • μ.rnDeriv ν := + Measure.rnDeriv_smul_right_of_ne_top' (μ := ν) (ν := μ) (by simp) (by simp [hν]) + filter_upwards [h1, h2] with x h1 h2 + rw [h1, Pi.smul_apply, smul_eq_mul, h2] + simp only [inv_inv, Pi.smul_apply, smul_eq_mul] + rw [← mul_assoc, ENNReal.inv_mul_cancel, one_mul] + · simp [hν] + · simp + have h_eq : ∫ x, f (μ'.rnDeriv ν' x).toReal ∂ν' + = (ν univ).toReal⁻¹ * ∫ x, f ((μ.rnDeriv ν x).toReal) ∂ν := by + rw [integral_smul_measure, smul_eq_mul, ENNReal.toReal_inv] + congr 1 + refine integral_congr_ae ?_ + filter_upwards [h_rnDeriv_eq] with x hx + rw [hx] + have h : f (μ' univ).toReal ≤ ∫ x, f (μ'.rnDeriv ν' x).toReal ∂ν' := + le_integral_rnDeriv_of_ac hf_cvx hf_cont ?_ hμν' + swap + · refine Integrable.smul_measure ?_ (by simp [hν]) + refine (integrable_congr ?_).mpr hf_int + filter_upwards [h_rnDeriv_eq] with x hx + rw [hx] + rw [h_eq, mul_comm, ← div_le_iff₀, div_eq_inv_mul, inv_inv] at h + · convert h + · simp only [div_eq_inv_mul, Measure.smul_apply, smul_eq_mul, ENNReal.toReal_mul, + ENNReal.toReal_inv, μ'] + · simp [ENNReal.toReal_pos_iff, hν] + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean index 95319082eb996..164391648cfb7 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean @@ -3,21 +3,19 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lp -import Mathlib.MeasureTheory.Integral.Bochner -import Mathlib.Order.Filter.IndicatorFunction -import Mathlib.MeasureTheory.Function.StronglyMeasurable.Inner import Mathlib.MeasureTheory.Function.LpSeminorm.Trim +import Mathlib.MeasureTheory.Function.StronglyMeasurable.Inner +import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lp /-! # Functions a.e. measurable with respect to a sub-σ-algebra -A function `f` verifies `AEStronglyMeasurable' m f μ` if it is `μ`-a.e. equal to +A function `f` verifies `AEStronglyMeasurable[m] f μ` if it is `μ`-a.e. equal to an `m`-strongly measurable function. This is similar to `AEStronglyMeasurable`, but the `MeasurableSpace` structures used for the measurability statement and for the measure are different. We define `lpMeas F 𝕜 m p μ`, the subspace of `Lp F p μ` containing functions `f` verifying -`AEStronglyMeasurable' m f μ`, i.e. functions which are `μ`-a.e. equal to an `m`-strongly +`AEStronglyMeasurable[m] f μ`, i.e. functions which are `μ`-a.e. equal to an `m`-strongly measurable function. ## Main statements @@ -42,104 +40,97 @@ open scoped ENNReal MeasureTheory namespace MeasureTheory -/-- A function `f` verifies `AEStronglyMeasurable' m f μ` if it is `μ`-a.e. equal to +/-- A function `f` verifies `AEStronglyMeasurable[m] f μ` if it is `μ`-a.e. equal to an `m`-strongly measurable function. This is similar to `AEStronglyMeasurable`, but the `MeasurableSpace` structures used for the measurability statement and for the measure are different. -/ +@[deprecated AEStronglyMeasurable (since := "2025-01-23")] def AEStronglyMeasurable' {α β} [TopologicalSpace β] (m : MeasurableSpace α) - {_ : MeasurableSpace α} (f : α → β) (μ : Measure α) : Prop := - ∃ g : α → β, StronglyMeasurable[m] g ∧ f =ᵐ[μ] g + {_ : MeasurableSpace α} (f : α → β) (μ : Measure α) : Prop := AEStronglyMeasurable[m] f μ + +set_option linter.deprecated false namespace AEStronglyMeasurable' variable {α β 𝕜 : Type*} {m m0 : MeasurableSpace α} {μ : Measure α} [TopologicalSpace β] {f g : α → β} -theorem congr (hf : AEStronglyMeasurable' m f μ) (hfg : f =ᵐ[μ] g) : - AEStronglyMeasurable' m g μ := by - obtain ⟨f', hf'_meas, hff'⟩ := hf; exact ⟨f', hf'_meas, hfg.symm.trans hff'⟩ - -theorem mono {m'} (hf : AEStronglyMeasurable' m f μ) (hm : m ≤ m') : - AEStronglyMeasurable' m' f μ := - let ⟨f', hf'_meas, hff'⟩ := hf; ⟨f', hf'_meas.mono hm, hff'⟩ - -theorem add [Add β] [ContinuousAdd β] (hf : AEStronglyMeasurable' m f μ) - (hg : AEStronglyMeasurable' m g μ) : AEStronglyMeasurable' m (f + g) μ := by - rcases hf with ⟨f', h_f'_meas, hff'⟩ - rcases hg with ⟨g', h_g'_meas, hgg'⟩ - exact ⟨f' + g', h_f'_meas.add h_g'_meas, hff'.add hgg'⟩ - -theorem neg [AddGroup β] [TopologicalAddGroup β] {f : α → β} (hfm : AEStronglyMeasurable' m f μ) : - AEStronglyMeasurable' m (-f) μ := by - rcases hfm with ⟨f', hf'_meas, hf_ae⟩ - refine ⟨-f', hf'_meas.neg, hf_ae.mono fun x hx => ?_⟩ - simp_rw [Pi.neg_apply] - rw [hx] - -theorem sub [AddGroup β] [TopologicalAddGroup β] {f g : α → β} (hfm : AEStronglyMeasurable' m f μ) - (hgm : AEStronglyMeasurable' m g μ) : AEStronglyMeasurable' m (f - g) μ := by - rcases hfm with ⟨f', hf'_meas, hf_ae⟩ - rcases hgm with ⟨g', hg'_meas, hg_ae⟩ - refine ⟨f' - g', hf'_meas.sub hg'_meas, hf_ae.mp (hg_ae.mono fun x hx1 hx2 => ?_)⟩ - simp_rw [Pi.sub_apply] - rw [hx1, hx2] - -theorem const_smul [SMul 𝕜 β] [ContinuousConstSMul 𝕜 β] (c : 𝕜) (hf : AEStronglyMeasurable' m f μ) : - AEStronglyMeasurable' m (c • f) μ := by - rcases hf with ⟨f', h_f'_meas, hff'⟩ - refine ⟨c • f', h_f'_meas.const_smul c, ?_⟩ - exact EventuallyEq.fun_comp hff' fun x => c • x +@[deprecated AEStronglyMeasurable.congr (since := "2025-01-23")] +theorem congr (hf : AEStronglyMeasurable[m] f μ) (hfg : f =ᵐ[μ] g) : + AEStronglyMeasurable[m] g μ := AEStronglyMeasurable.congr hf hfg + +@[deprecated AEStronglyMeasurable.mono (since := "2025-01-23")] +theorem mono {m'} (hf : AEStronglyMeasurable[m] f μ) (hm : m ≤ m') : + AEStronglyMeasurable' m' f μ := AEStronglyMeasurable.mono hm hf + +@[deprecated AEStronglyMeasurable.add (since := "2025-01-23")] +theorem add [Add β] [ContinuousAdd β] (hf : AEStronglyMeasurable[m] f μ) + (hg : AEStronglyMeasurable[m] g μ) : AEStronglyMeasurable[m] (f + g) μ := + AEStronglyMeasurable.add hf hg + +@[deprecated AEStronglyMeasurable.neg (since := "2025-01-23")] +theorem neg [AddGroup β] [TopologicalAddGroup β] {f : α → β} (hfm : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] (-f) μ := + AEStronglyMeasurable.neg hfm +@[deprecated AEStronglyMeasurable.sub (since := "2025-01-23")] +theorem sub [AddGroup β] [TopologicalAddGroup β] {f g : α → β} (hfm : AEStronglyMeasurable[m] f μ) + (hgm : AEStronglyMeasurable[m] g μ) : AEStronglyMeasurable[m] (f - g) μ := + AEStronglyMeasurable.sub hfm hgm + +@[deprecated AEStronglyMeasurable.const_smul (since := "2025-01-23")] +theorem const_smul [SMul 𝕜 β] [ContinuousConstSMul 𝕜 β] (c : 𝕜) (hf : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] (c • f) μ := + AEStronglyMeasurable.const_smul hf _ + +@[deprecated AEStronglyMeasurable.const_inner (since := "2025-01-23")] theorem const_inner {𝕜 β} [RCLike 𝕜] [NormedAddCommGroup β] [InnerProductSpace 𝕜 β] {f : α → β} - (hfm : AEStronglyMeasurable' m f μ) (c : β) : - AEStronglyMeasurable' m (fun x => (inner c (f x) : 𝕜)) μ := by - rcases hfm with ⟨f', hf'_meas, hf_ae⟩ - refine - ⟨fun x => (inner c (f' x) : 𝕜), (@stronglyMeasurable_const _ _ m _ c).inner hf'_meas, - hf_ae.mono fun x hx => ?_⟩ - dsimp only - rw [hx] + (hfm : AEStronglyMeasurable[m] f μ) (c : β) : + AEStronglyMeasurable[m] (fun x => (inner c (f x) : 𝕜)) μ := + AEStronglyMeasurable.const_inner hfm -@[simp] theorem of_subsingleton [Subsingleton β] : AEStronglyMeasurable' m f μ := - ⟨f, by simp, by simp⟩ +@[deprecated AEStronglyMeasurable.of_subsingleton_cod (since := "2025-01-23")] +theorem of_subsingleton [Subsingleton β] : AEStronglyMeasurable[m] f μ := .of_subsingleton_cod -@[simp] theorem of_subsingleton' [Subsingleton α] : AEStronglyMeasurable' m f μ := - ⟨f, by simp, by simp⟩ +@[deprecated AEStronglyMeasurable.of_subsingleton_dom (since := "2025-01-23")] +theorem of_subsingleton' [Subsingleton α] : AEStronglyMeasurable[m] f μ := .of_subsingleton_dom /-- An `m`-strongly measurable function almost everywhere equal to `f`. -/ -noncomputable def mk (f : α → β) (hfm : AEStronglyMeasurable' m f μ) : α → β := - hfm.choose +@[deprecated AEStronglyMeasurable.mk (since := "2025-01-23")] +noncomputable def mk (f : α → β) (hfm : AEStronglyMeasurable[m] f μ) : α → β := + AEStronglyMeasurable.mk f hfm -theorem stronglyMeasurable_mk {f : α → β} (hfm : AEStronglyMeasurable' m f μ) : +@[deprecated AEStronglyMeasurable.stronglyMeasurable_mk (since := "2025-01-23")] +theorem stronglyMeasurable_mk {f : α → β} (hfm : AEStronglyMeasurable[m] f μ) : StronglyMeasurable[m] (hfm.mk f) := - hfm.choose_spec.1 + AEStronglyMeasurable.stronglyMeasurable_mk hfm -theorem ae_eq_mk {f : α → β} (hfm : AEStronglyMeasurable' m f μ) : f =ᵐ[μ] hfm.mk f := - hfm.choose_spec.2 +@[deprecated AEStronglyMeasurable.ae_eq_mk (since := "2025-01-23")] +theorem ae_eq_mk {f : α → β} (hfm : AEStronglyMeasurable[m] f μ) : f =ᵐ[μ] hfm.mk f := + AEStronglyMeasurable.ae_eq_mk hfm +@[deprecated Continuous.comp_aestronglyMeasurable (since := "2025-01-23")] theorem continuous_comp {γ} [TopologicalSpace γ] {f : α → β} {g : β → γ} (hg : Continuous g) - (hf : AEStronglyMeasurable' m f μ) : AEStronglyMeasurable' m (g ∘ f) μ := - ⟨fun x => g (hf.mk _ x), - @Continuous.comp_stronglyMeasurable _ _ _ m _ _ _ _ hg hf.stronglyMeasurable_mk, - hf.ae_eq_mk.mono fun x hx => by rw [Function.comp_apply, hx]⟩ + (hf : AEStronglyMeasurable[m] f μ) : AEStronglyMeasurable[m] (g ∘ f) μ := + hg.comp_aestronglyMeasurable hf end AEStronglyMeasurable' +@[deprecated AEStronglyMeasurable.of_trim (since := "2025-01-23")] theorem aeStronglyMeasurable'_of_aeStronglyMeasurable'_trim {α β} {m m0 m0' : MeasurableSpace α} [TopologicalSpace β] (hm0 : m0 ≤ m0') {μ : Measure α} {f : α → β} - (hf : AEStronglyMeasurable' m f (μ.trim hm0)) : AEStronglyMeasurable' m f μ := by - obtain ⟨g, hg_meas, hfg⟩ := hf; exact ⟨g, hg_meas, ae_eq_of_ae_eq_trim hfg⟩ + (hf : AEStronglyMeasurable[m] f (μ.trim hm0)) : AEStronglyMeasurable[m] f μ := .of_trim hm0 hf +@[deprecated StronglyMeasurable.aestronglyMeasurable (since := "2025-01-23")] theorem StronglyMeasurable.aeStronglyMeasurable' {α β} {m _ : MeasurableSpace α} [TopologicalSpace β] {μ : Measure α} {f : α → β} (hf : StronglyMeasurable[m] f) : - AEStronglyMeasurable' m f μ := - ⟨f, hf, ae_eq_refl _⟩ + AEStronglyMeasurable[m] f μ := hf.aestronglyMeasurable theorem ae_eq_trim_iff_of_aeStronglyMeasurable' {α β} [TopologicalSpace β] [MetrizableSpace β] {m m0 : MeasurableSpace α} {μ : Measure α} {f g : α → β} (hm : m ≤ m0) - (hfm : AEStronglyMeasurable' m f μ) (hgm : AEStronglyMeasurable' m g μ) : + (hfm : AEStronglyMeasurable[m] f μ) (hgm : AEStronglyMeasurable[m] g μ) : hfm.mk f =ᵐ[μ.trim hm] hgm.mk g ↔ f =ᵐ[μ] g := - (ae_eq_trim_iff hm hfm.stronglyMeasurable_mk hgm.stronglyMeasurable_mk).trans + (hfm.stronglyMeasurable_mk.ae_eq_trim_iff hm hgm.stronglyMeasurable_mk).trans ⟨fun h => hfm.ae_eq_mk.trans (h.trans hgm.ae_eq_mk.symm), fun h => hfm.ae_eq_mk.symm.trans (h.trans hgm.ae_eq_mk)⟩ @@ -154,26 +145,14 @@ theorem AEStronglyMeasurable.comp_ae_measurable' {α β γ : Type*} [Topological another σ-algebra `m₂` (hypothesis `hs`), the set `s` is `m` measurable and a function `f` almost everywhere supported on `s` is `m`-ae-strongly-measurable, then `f` is also `m₂`-ae-strongly-measurable. -/ +@[deprecated AEStronglyMeasurable.of_measurableSpace_le_on (since := "2025-01-23")] theorem AEStronglyMeasurable'.aeStronglyMeasurable'_of_measurableSpace_le_on {α E} {m m₂ m0 : MeasurableSpace α} {μ : Measure α} [TopologicalSpace E] [Zero E] (hm : m ≤ m0) {s : Set α} {f : α → E} (hs_m : MeasurableSet[m] s) (hs : ∀ t, MeasurableSet[m] (s ∩ t) → MeasurableSet[m₂] (s ∩ t)) - (hf : AEStronglyMeasurable' m f μ) (hf_zero : f =ᵐ[μ.restrict sᶜ] 0) : - AEStronglyMeasurable' m₂ f μ := by - have h_ind_eq : s.indicator (hf.mk f) =ᵐ[μ] f := by - refine Filter.EventuallyEq.trans ?_ <| - indicator_ae_eq_of_restrict_compl_ae_eq_zero (hm _ hs_m) hf_zero - filter_upwards [hf.ae_eq_mk] with x hx - by_cases hxs : x ∈ s - · simp [hxs, hx] - · simp [hxs] - suffices StronglyMeasurable[m₂] (s.indicator (hf.mk f)) from - AEStronglyMeasurable'.congr this.aeStronglyMeasurable' h_ind_eq - have hf_ind : StronglyMeasurable[m] (s.indicator (hf.mk f)) := - hf.stronglyMeasurable_mk.indicator hs_m - exact - hf_ind.stronglyMeasurable_of_measurableSpace_le_on hs_m hs fun x hxs => - Set.indicator_of_not_mem hxs _ + (hf : AEStronglyMeasurable[m] f μ) (hf_zero : f =ᵐ[μ.restrict sᶜ] 0) : + AEStronglyMeasurable' m₂ f μ := + .of_measurableSpace_le_on hm hs_m hs hf hf_zero variable {α F 𝕜 : Type*} {p : ℝ≥0∞} [RCLike 𝕜] -- 𝕜 for ℝ or ℂ @@ -188,44 +167,53 @@ section LpMeas variable (F) /-- `lpMeasSubgroup F m p μ` is the subspace of `Lp F p μ` containing functions `f` verifying -`AEStronglyMeasurable' m f μ`, i.e. functions which are `μ`-a.e. equal to +`AEStronglyMeasurable[m] f μ`, i.e. functions which are `μ`-a.e. equal to an `m`-strongly measurable function. -/ def lpMeasSubgroup (m : MeasurableSpace α) [MeasurableSpace α] (p : ℝ≥0∞) (μ : Measure α) : AddSubgroup (Lp F p μ) where - carrier := {f : Lp F p μ | AEStronglyMeasurable' m f μ} + carrier := {f : Lp F p μ | AEStronglyMeasurable[m] f μ} zero_mem' := ⟨(0 : α → F), @stronglyMeasurable_zero _ _ m _ _, Lp.coeFn_zero _ _ _⟩ add_mem' {f g} hf hg := (hf.add hg).congr (Lp.coeFn_add f g).symm - neg_mem' {f} hf := AEStronglyMeasurable'.congr hf.neg (Lp.coeFn_neg f).symm + neg_mem' {f} hf := AEStronglyMeasurable.congr hf.neg (Lp.coeFn_neg f).symm variable (𝕜) /-- `lpMeas F 𝕜 m p μ` is the subspace of `Lp F p μ` containing functions `f` verifying -`AEStronglyMeasurable' m f μ`, i.e. functions which are `μ`-a.e. equal to +`AEStronglyMeasurable[m] f μ`, i.e. functions which are `μ`-a.e. equal to an `m`-strongly measurable function. -/ def lpMeas (m : MeasurableSpace α) [MeasurableSpace α] (p : ℝ≥0∞) (μ : Measure α) : Submodule 𝕜 (Lp F p μ) where - carrier := {f : Lp F p μ | AEStronglyMeasurable' m f μ} + carrier := {f : Lp F p μ | AEStronglyMeasurable[m] f μ} zero_mem' := ⟨(0 : α → F), @stronglyMeasurable_zero _ _ m _ _, Lp.coeFn_zero _ _ _⟩ add_mem' {f g} hf hg := (hf.add hg).congr (Lp.coeFn_add f g).symm smul_mem' c f hf := (hf.const_smul c).congr (Lp.coeFn_smul c f).symm variable {F 𝕜} -theorem mem_lpMeasSubgroup_iff_aeStronglyMeasurable' {m m0 : MeasurableSpace α} {μ : Measure α} - {f : Lp F p μ} : f ∈ lpMeasSubgroup F m p μ ↔ AEStronglyMeasurable' m f μ := by +theorem mem_lpMeasSubgroup_iff_aeStronglyMeasurable {m m0 : MeasurableSpace α} {μ : Measure α} + {f : Lp F p μ} : f ∈ lpMeasSubgroup F m p μ ↔ AEStronglyMeasurable[m] f μ := by rw [← AddSubgroup.mem_carrier, lpMeasSubgroup, Set.mem_setOf_eq] -theorem mem_lpMeas_iff_aeStronglyMeasurable' {m m0 : MeasurableSpace α} {μ : Measure α} - {f : Lp F p μ} : f ∈ lpMeas F 𝕜 m p μ ↔ AEStronglyMeasurable' m f μ := by +@[deprecated (since := "2025-01-24")] +alias mem_lpMeasSubgroup_iff_aeStronglyMeasurable' := mem_lpMeasSubgroup_iff_aeStronglyMeasurable + +theorem mem_lpMeas_iff_aeStronglyMeasurable {m m0 : MeasurableSpace α} {μ : Measure α} + {f : Lp F p μ} : f ∈ lpMeas F 𝕜 m p μ ↔ AEStronglyMeasurable[m] f μ := by rw [← SetLike.mem_coe, ← Submodule.mem_carrier, lpMeas, Set.mem_setOf_eq] -theorem lpMeas.aeStronglyMeasurable' {m _ : MeasurableSpace α} {μ : Measure α} - (f : lpMeas F 𝕜 m p μ) : AEStronglyMeasurable' (β := F) m f μ := - mem_lpMeas_iff_aeStronglyMeasurable'.mp f.mem +@[deprecated (since := "2025-01-24")] +alias mem_lpMeas_iff_aeStronglyMeasurable' := mem_lpMeas_iff_aeStronglyMeasurable + +theorem lpMeas.aeStronglyMeasurable {m _ : MeasurableSpace α} {μ : Measure α} + (f : lpMeas F 𝕜 m p μ) : AEStronglyMeasurable[m] (f : α → F) μ := + mem_lpMeas_iff_aeStronglyMeasurable.mp f.mem + +@[deprecated (since := "2025-01-24")] +alias lpMeas.aeStronglyMeasurable' := lpMeas.aeStronglyMeasurable theorem mem_lpMeas_self {m0 : MeasurableSpace α} (μ : Measure α) (f : Lp F p μ) : f ∈ lpMeas F 𝕜 m0 p μ := - mem_lpMeas_iff_aeStronglyMeasurable'.mpr (Lp.aestronglyMeasurable f) + mem_lpMeas_iff_aeStronglyMeasurable.mpr (Lp.aestronglyMeasurable f) theorem lpMeasSubgroup_coe {m _ : MeasurableSpace α} {μ : Measure α} {f : lpMeasSubgroup F m p μ} : (f : _ → _) = (f : Lp F p μ) := @@ -256,9 +244,9 @@ variable {m m0 : MeasurableSpace α} {μ : Measure α} everywhere equal to (given by `AEMeasurable.mk`) belongs to `ℒp` for the measure `μ.trim hm`. -/ theorem memℒp_trim_of_mem_lpMeasSubgroup (hm : m ≤ m0) (f : Lp F p μ) (hf_meas : f ∈ lpMeasSubgroup F m p μ) : - Memℒp (mem_lpMeasSubgroup_iff_aeStronglyMeasurable'.mp hf_meas).choose p (μ.trim hm) := by - have hf : AEStronglyMeasurable' m f μ := - mem_lpMeasSubgroup_iff_aeStronglyMeasurable'.mp hf_meas + Memℒp (mem_lpMeasSubgroup_iff_aeStronglyMeasurable.mp hf_meas).choose p (μ.trim hm) := by + have hf : AEStronglyMeasurable[m] f μ := + mem_lpMeasSubgroup_iff_aeStronglyMeasurable.mp hf_meas let g := hf.choose obtain ⟨hg, hfg⟩ := hf.choose_spec change Memℒp g p (μ.trim hm) @@ -274,7 +262,7 @@ theorem memℒp_trim_of_mem_lpMeasSubgroup (hm : m ≤ m0) (f : Lp F p μ) theorem mem_lpMeasSubgroup_toLp_of_trim (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : (memℒp_of_memℒp_trim hm (Lp.memℒp f)).toLp f ∈ lpMeasSubgroup F m p μ := by let hf_mem_ℒp := memℒp_of_memℒp_trim hm (Lp.memℒp f) - rw [mem_lpMeasSubgroup_iff_aeStronglyMeasurable'] + rw [mem_lpMeasSubgroup_iff_aeStronglyMeasurable] refine AEStronglyMeasurable'.congr ?_ (Memℒp.coeFn_toLp hf_mem_ℒp).symm refine aeStronglyMeasurable'_of_aeStronglyMeasurable'_trim hm ?_ exact Lp.aestronglyMeasurable f @@ -284,7 +272,7 @@ variable (F p μ) /-- Map from `lpMeasSubgroup` to `Lp F p (μ.trim hm)`. -/ noncomputable def lpMeasSubgroupToLpTrim (hm : m ≤ m0) (f : lpMeasSubgroup F m p μ) : Lp F p (μ.trim hm) := - Memℒp.toLp (mem_lpMeasSubgroup_iff_aeStronglyMeasurable'.mp f.mem).choose + Memℒp.toLp (mem_lpMeasSubgroup_iff_aeStronglyMeasurable.mp f.mem).choose -- Porting note: had to replace `f` with `f.1` here. (memℒp_trim_of_mem_lpMeasSubgroup hm f.1 f.mem) @@ -292,7 +280,7 @@ variable (𝕜) /-- Map from `lpMeas` to `Lp F p (μ.trim hm)`. -/ noncomputable def lpMeasToLpTrim (hm : m ≤ m0) (f : lpMeas F 𝕜 m p μ) : Lp F p (μ.trim hm) := - Memℒp.toLp (mem_lpMeas_iff_aeStronglyMeasurable'.mp f.mem).choose + Memℒp.toLp (mem_lpMeas_iff_aeStronglyMeasurable.mp f.mem).choose -- Porting note: had to replace `f` with `f.1` here. (memℒp_trim_of_mem_lpMeasSubgroup hm f.1 f.mem) @@ -316,7 +304,7 @@ theorem lpMeasSubgroupToLpTrim_ae_eq (hm : m ≤ m0) (f : lpMeasSubgroup F m p lpMeasSubgroupToLpTrim F p μ hm f =ᵐ[μ] f := -- Porting note: replaced `(↑f)` with `f.1` here. (ae_eq_of_ae_eq_trim (Memℒp.coeFn_toLp (memℒp_trim_of_mem_lpMeasSubgroup hm f.1 f.mem))).trans - (mem_lpMeasSubgroup_iff_aeStronglyMeasurable'.mp f.mem).choose_spec.2.symm + (mem_lpMeasSubgroup_iff_aeStronglyMeasurable.mp f.mem).choose_spec.2.symm theorem lpTrimToLpMeasSubgroup_ae_eq (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : lpTrimToLpMeasSubgroup F p μ hm f =ᵐ[μ] f := @@ -327,7 +315,7 @@ theorem lpMeasToLpTrim_ae_eq (hm : m ≤ m0) (f : lpMeas F 𝕜 m p μ) : lpMeasToLpTrim F 𝕜 p μ hm f =ᵐ[μ] f := -- Porting note: replaced `(↑f)` with `f.1` here. (ae_eq_of_ae_eq_trim (Memℒp.coeFn_toLp (memℒp_trim_of_mem_lpMeasSubgroup hm f.1 f.mem))).trans - (mem_lpMeasSubgroup_iff_aeStronglyMeasurable'.mp f.mem).choose_spec.2.symm + (mem_lpMeasSubgroup_iff_aeStronglyMeasurable.mp f.mem).choose_spec.2.symm theorem lpTrimToLpMeas_ae_eq (hm : m ≤ m0) (f : Lp F p (μ.trim hm)) : lpTrimToLpMeas F 𝕜 p μ hm f =ᵐ[μ] f := @@ -340,7 +328,7 @@ theorem lpMeasSubgroupToLpTrim_right_inv (hm : m ≤ m0) : intro f ext1 refine - ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _) (Lp.stronglyMeasurable _) ?_ + (Lp.stronglyMeasurable _).ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _) ?_ exact (lpMeasSubgroupToLpTrim_ae_eq hm _).trans (lpTrimToLpMeasSubgroup_ae_eq hm _) /-- `lpTrimToLpMeasSubgroup` is a left inverse of `lpMeasSubgroupToLpTrim`. -/ @@ -357,7 +345,7 @@ theorem lpMeasSubgroupToLpTrim_add (hm : m ≤ m0) (f g : lpMeasSubgroup F m p lpMeasSubgroupToLpTrim F p μ hm f + lpMeasSubgroupToLpTrim F p μ hm g := by ext1 refine EventuallyEq.trans ?_ (Lp.coeFn_add _ _).symm - refine ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _) ?_ ?_ + refine (Lp.stronglyMeasurable _).ae_eq_trim_of_stronglyMeasurable hm ?_ ?_ · exact (Lp.stronglyMeasurable _).add (Lp.stronglyMeasurable _) refine (lpMeasSubgroupToLpTrim_ae_eq hm _).trans ?_ refine @@ -372,11 +360,9 @@ theorem lpMeasSubgroupToLpTrim_neg (hm : m ≤ m0) (f : lpMeasSubgroup F m p μ) lpMeasSubgroupToLpTrim F p μ hm (-f) = -lpMeasSubgroupToLpTrim F p μ hm f := by ext1 refine EventuallyEq.trans ?_ (Lp.coeFn_neg _).symm - refine ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _) ?_ ?_ - · exact @StronglyMeasurable.neg _ _ _ m _ _ _ (Lp.stronglyMeasurable _) - refine (lpMeasSubgroupToLpTrim_ae_eq hm _).trans ?_ - refine EventuallyEq.trans ?_ (EventuallyEq.neg (lpMeasSubgroupToLpTrim_ae_eq hm f).symm) - refine (Lp.coeFn_neg _).trans ?_ + refine (Lp.stronglyMeasurable _).ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _).neg + <| (lpMeasSubgroupToLpTrim_ae_eq hm _).trans <| + ((Lp.coeFn_neg _).trans ?_).trans (lpMeasSubgroupToLpTrim_ae_eq hm f).symm.neg simp_rw [lpMeasSubgroup_coe] exact Eventually.of_forall fun x => by rfl @@ -390,7 +376,7 @@ theorem lpMeasToLpTrim_smul (hm : m ≤ m0) (c : 𝕜) (f : lpMeas F 𝕜 m p μ lpMeasToLpTrim F 𝕜 p μ hm (c • f) = c • lpMeasToLpTrim F 𝕜 p μ hm f := by ext1 refine EventuallyEq.trans ?_ (Lp.coeFn_smul _ _).symm - refine ae_eq_trim_of_stronglyMeasurable hm (Lp.stronglyMeasurable _) ?_ ?_ + refine (Lp.stronglyMeasurable _).ae_eq_trim_of_stronglyMeasurable hm ?_ ?_ · exact (Lp.stronglyMeasurable _).const_smul c refine (lpMeasToLpTrim_ae_eq hm _).trans ?_ refine (Lp.coeFn_smul _ _).trans ?_ @@ -454,14 +440,14 @@ instance [hm : Fact (m ≤ m0)] [CompleteSpace F] [hp : Fact (1 ≤ p)] : rw [(lpMeasSubgroupToLpMeasIso F 𝕜 p μ).symm.completeSpace_iff]; infer_instance theorem isComplete_aeStronglyMeasurable' [hp : Fact (1 ≤ p)] [CompleteSpace F] (hm : m ≤ m0) : - IsComplete {f : Lp F p μ | AEStronglyMeasurable' m f μ} := by + IsComplete {f : Lp F p μ | AEStronglyMeasurable[m] f μ} := by rw [← completeSpace_coe_iff_isComplete] haveI : Fact (m ≤ m0) := ⟨hm⟩ change CompleteSpace (lpMeasSubgroup F m p μ) infer_instance theorem isClosed_aeStronglyMeasurable' [Fact (1 ≤ p)] [CompleteSpace F] (hm : m ≤ m0) : - IsClosed {f : Lp F p μ | AEStronglyMeasurable' m f μ} := + IsClosed {f : Lp F p μ | AEStronglyMeasurable[m] f μ} := IsComplete.isClosed (isComplete_aeStronglyMeasurable' hm) end CompleteSubspace @@ -517,11 +503,11 @@ variable {m m0 : MeasurableSpace α} {μ : Measure α} [Fact (1 ≤ p)] [NormedS theorem Lp.induction_stronglyMeasurable_aux (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) (P : Lp F p μ → Prop) (h_ind : ∀ (c : F) {s : Set α} (hs : MeasurableSet[m] s) (hμs : μ s < ∞), P (Lp.simpleFunc.indicatorConst p (hm s hs) hμs.ne c)) - (h_add : ∀ ⦃f g⦄, ∀ hf : Memℒp f p μ, ∀ hg : Memℒp g p μ, AEStronglyMeasurable' m f μ → - AEStronglyMeasurable' m g μ → Disjoint (Function.support f) (Function.support g) → + (h_add : ∀ ⦃f g⦄, ∀ hf : Memℒp f p μ, ∀ hg : Memℒp g p μ, AEStronglyMeasurable[m] f μ → + AEStronglyMeasurable[m] g μ → Disjoint (Function.support f) (Function.support g) → P (hf.toLp f) → P (hg.toLp g) → P (hf.toLp f + hg.toLp g)) (h_closed : IsClosed {f : lpMeas F ℝ m p μ | P f}) : - ∀ f : Lp F p μ, AEStronglyMeasurable' m f μ → P f := by + ∀ f : Lp F p μ, AEStronglyMeasurable[m] f μ → P f := by intro f hf let f' := (⟨f, hf⟩ : lpMeas F ℝ m p μ) let g := lpMeasToLpTrimLie F ℝ p μ hm f' @@ -573,11 +559,11 @@ theorem Lp.induction_stronglyMeasurable (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) StronglyMeasurable[m] g → Disjoint (Function.support f) (Function.support g) → P (hf.toLp f) → P (hg.toLp g) → P (hf.toLp f + hg.toLp g)) (h_closed : IsClosed {f : lpMeas F ℝ m p μ | P f}) : - ∀ f : Lp F p μ, AEStronglyMeasurable' m f μ → P f := by + ∀ f : Lp F p μ, AEStronglyMeasurable[m] f μ → P f := by intro f hf suffices h_add_ae : - ∀ ⦃f g⦄, ∀ hf : Memℒp f p μ, ∀ hg : Memℒp g p μ, AEStronglyMeasurable' m f μ → - AEStronglyMeasurable' m g μ → Disjoint (Function.support f) (Function.support g) → + ∀ ⦃f g⦄, ∀ hf : Memℒp f p μ, ∀ hg : Memℒp g p μ, AEStronglyMeasurable[m] f μ → + AEStronglyMeasurable[m] g μ → Disjoint (Function.support f) (Function.support g) → P (hf.toLp f) → P (hg.toLp g) → P (hf.toLp f + hg.toLp g) from Lp.induction_stronglyMeasurable_aux hm hp_ne_top _ h_ind h_add_ae h_closed f hf intro f g hf hg hfm hgm h_disj hPf hPg @@ -636,10 +622,10 @@ theorem Memℒp.induction_stronglyMeasurable (hm : m ≤ m0) (hp_ne_top : p ≠ P f → P g → P (f + g)) (h_closed : IsClosed {f : lpMeas F ℝ m p μ | P f}) (h_ae : ∀ ⦃f g⦄, f =ᵐ[μ] g → Memℒp f p μ → P f → P g) : - ∀ ⦃f : α → F⦄, Memℒp f p μ → AEStronglyMeasurable' m f μ → P f := by + ∀ ⦃f : α → F⦄, Memℒp f p μ → AEStronglyMeasurable[m] f μ → P f := by intro f hf hfm let f_Lp := hf.toLp f - have hfm_Lp : AEStronglyMeasurable' m f_Lp μ := hfm.congr hf.coeFn_toLp.symm + have hfm_Lp : AEStronglyMeasurable[m] f_Lp μ := hfm.congr hf.coeFn_toLp.symm refine h_ae hf.coeFn_toLp (Lp.memℒp _) ?_ change P f_Lp refine Lp.induction_stronglyMeasurable hm hp_ne_top (fun f => P f) ?_ ?_ h_closed f_Lp hfm_Lp diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index 8b2d1dfe66dd5..e5fb77b18b62e 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -8,8 +8,8 @@ import Mathlib.MeasureTheory.Function.ConditionalExpectation.CondexpL1 /-! # Conditional expectation We build the conditional expectation of an integrable function `f` with value in a Banach space -with respect to a measure `μ` (defined on a measurable space structure `m0`) and a measurable space -structure `m` with `hm : m ≤ m0` (a sub-sigma-algebra). This is an `m`-strongly measurable +with respect to a measure `μ` (defined on a measurable space structure `m₀`) and a measurable space +structure `m` with `hm : m ≤ m₀` (a sub-sigma-algebra). This is an `m`-strongly measurable function `μ[f|hm]` which is integrable and verifies `∫ x in s, μ[f|hm] x ∂μ = ∫ x in s, f x ∂μ` for all `m`-measurable sets `s`. It is unique as an element of `L¹`. @@ -20,11 +20,11 @@ The construction is done in four steps: is integrable and define a map `Set α → (E →L[ℝ] (α →₁[μ] E))` which to a set associates a linear map. That linear map sends `x ∈ E` to the conditional expectation of the indicator of the set with value `x`. -* Extend that map to `condexpL1CLM : (α →₁[μ] E) →L[ℝ] (α →₁[μ] E)`. This is done using the same +* Extend that map to `condExpL1CLM : (α →₁[μ] E) →L[ℝ] (α →₁[μ] E)`. This is done using the same construction as the Bochner integral (see the file `MeasureTheory/Integral/SetToL1`). * Define the conditional expectation of a function `f : α → E`, which is an integrable function `α → E` equal to 0 if `f` is not integrable, and equal to an `m`-measurable representative of - `condexpL1CLM` applied to `[f]`, the equivalence class of `f` in `L¹`. + `condExpL1CLM` applied to `[f]`, the equivalence class of `f` in `L¹`. The first step is done in `MeasureTheory.Function.ConditionalExpectation.CondexpL2`, the two next steps in `MeasureTheory.Function.ConditionalExpectation.CondexpL1` and the final step is @@ -34,27 +34,33 @@ performed in this file. The conditional expectation and its properties -* `condexp (m : MeasurableSpace α) (μ : Measure α) (f : α → E)`: conditional expectation of `f` +* `condExp (m : MeasurableSpace α) (μ : Measure α) (f : α → E)`: conditional expectation of `f` with respect to `m`. -* `integrable_condexp` : `condexp` is integrable. -* `stronglyMeasurable_condexp` : `condexp` is `m`-strongly-measurable. -* `setIntegral_condexp (hf : Integrable f μ) (hs : MeasurableSet[m] s)` : if `m ≤ m0` (the +* `integrable_condExp` : `condExp` is integrable. +* `stronglyMeasurable_condExp` : `condExp` is `m`-strongly-measurable. +* `setIntegral_condExp (hf : Integrable f μ) (hs : MeasurableSet[m] s)` : if `m ≤ m₀` (the σ-algebra over which the measure is defined), then the conditional expectation verifies - `∫ x in s, condexp m μ f x ∂μ = ∫ x in s, f x ∂μ` for any `m`-measurable set `s`. + `∫ x in s, condExp m μ f x ∂μ = ∫ x in s, f x ∂μ` for any `m`-measurable set `s`. -While `condexp` is function-valued, we also define `condexpL1` with value in `L1` and a continuous -linear map `condexpL1CLM` from `L1` to `L1`. `condexp` should be used in most cases. +While `condExp` is function-valued, we also define `condExpL1` with value in `L1` and a continuous +linear map `condExpL1CLM` from `L1` to `L1`. `condExp` should be used in most cases. Uniqueness of the conditional expectation -* `ae_eq_condexp_of_forall_setIntegral_eq`: an a.e. `m`-measurable function which verifies the - equality of integrals is a.e. equal to `condexp`. +* `ae_eq_condExp_of_forall_setIntegral_eq`: an a.e. `m`-measurable function which verifies the + equality of integrals is a.e. equal to `condExp`. ## Notations -For a measure `μ` defined on a measurable space structure `m0`, another measurable space structure -`m` with `hm : m ≤ m0` (a sub-σ-algebra) and a function `f`, we define the notation -* `μ[f|m] = condexp m μ f`. +For a measure `μ` defined on a measurable space structure `m₀`, another measurable space structure +`m` with `hm : m ≤ m₀` (a sub-σ-algebra) and a function `f`, we define the notation +* `μ[f|m] = condExp m μ f`. + +## TODO + +See https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/Conditional.20expectation.20of.20product +for how to prove that we can pull `m`-measurable continuous linear maps out of the `m`-conditional +expectation. This would generalise `MeasureTheory.condExp_mul_of_stronglyMeasurable_left`. ## Tags @@ -62,306 +68,394 @@ conditional expectation, conditional expected value -/ - open TopologicalSpace MeasureTheory.Lp Filter - -open scoped ENNReal Topology MeasureTheory +open scoped Classical ENNReal Topology MeasureTheory namespace MeasureTheory - -variable {α F F' 𝕜 : Type*} [RCLike 𝕜] -- 𝕜 for ℝ or ℂ - -- F' for integrals on a Lp submodule - [NormedAddCommGroup F'] - [NormedSpace 𝕜 F'] [NormedSpace ℝ F'] [CompleteSpace F'] - -open scoped Classical + -- E for integrals on a Lp submodule +variable {α β E 𝕜 : Type*} [RCLike 𝕜] {m m₀ : MeasurableSpace α} {μ : Measure α} {f g : α → E} + {s : Set α} -variable {m m0 : MeasurableSpace α} {μ : Measure α} {f g : α → F'} {s : Set α} +section NormedAddCommGroup +variable [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable (m) in /-- Conditional expectation of a function. It is defined as 0 if any one of the following conditions is true: -- `m` is not a sub-σ-algebra of `m0`, +- `m` is not a sub-σ-algebra of `m₀`, - `μ` is not σ-finite with respect to `m`, - `f` is not integrable. -/ -noncomputable irreducible_def condexp (m : MeasurableSpace α) {m0 : MeasurableSpace α} - (μ : Measure α) (f : α → F') : α → F' := - if hm : m ≤ m0 then +noncomputable irreducible_def condExp (μ : Measure[m₀] α) (f : α → E) : α → E := + if hm : m ≤ m₀ then if h : SigmaFinite (μ.trim hm) ∧ Integrable f μ then if StronglyMeasurable[m] f then f - else (@aestronglyMeasurable'_condexpL1 _ _ _ _ _ m m0 μ hm h.1 _).mk - (@condexpL1 _ _ _ _ _ _ _ hm μ h.1 f) + else have := h.1; aestronglyMeasurable_condExpL1.mk (condExpL1 hm μ f) else 0 else 0 +@[deprecated (since := "2025-01-21")] alias condexp := condExp + -- We define notation `μ[f|m]` for the conditional expectation of `f` with respect to `m`. -scoped notation μ "[" f "|" m "]" => MeasureTheory.condexp m μ f +scoped notation μ "[" f "|" m "]" => MeasureTheory.condExp m μ f -theorem condexp_of_not_le (hm_not : ¬m ≤ m0) : μ[f|m] = 0 := by rw [condexp, dif_neg hm_not] +theorem condExp_of_not_le (hm_not : ¬m ≤ m₀) : μ[f|m] = 0 := by rw [condExp, dif_neg hm_not] -theorem condexp_of_not_sigmaFinite (hm : m ≤ m0) (hμm_not : ¬SigmaFinite (μ.trim hm)) : - μ[f|m] = 0 := by rw [condexp, dif_pos hm, dif_neg]; push_neg; exact fun h => absurd h hμm_not +@[deprecated (since := "2025-01-21")] alias condexp_of_not_le := condExp_of_not_le -theorem condexp_of_sigmaFinite (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm)] : +theorem condExp_of_not_sigmaFinite (hm : m ≤ m₀) (hμm_not : ¬SigmaFinite (μ.trim hm)) : + μ[f|m] = 0 := by rw [condExp, dif_pos hm, dif_neg]; push_neg; exact fun h => absurd h hμm_not + +@[deprecated (since := "2025-01-21")] alias condexp_of_not_sigmaFinite := condExp_of_not_sigmaFinite + +theorem condExp_of_sigmaFinite (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] : μ[f|m] = if Integrable f μ then if StronglyMeasurable[m] f then f - else aestronglyMeasurable'_condexpL1.mk (condexpL1 hm μ f) + else aestronglyMeasurable_condExpL1.mk (condExpL1 hm μ f) else 0 := by - rw [condexp, dif_pos hm] + rw [condExp, dif_pos hm] simp only [hμm, Ne, true_and] by_cases hf : Integrable f μ · rw [dif_pos hf, if_pos hf] · rw [dif_neg hf, if_neg hf] -theorem condexp_of_stronglyMeasurable (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm)] {f : α → F'} +@[deprecated (since := "2025-01-21")] alias condexp_of_sigmaFinite := condExp_of_sigmaFinite + +theorem condExp_of_stronglyMeasurable (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] {f : α → E} (hf : StronglyMeasurable[m] f) (hfi : Integrable f μ) : μ[f|m] = f := by - rw [condexp_of_sigmaFinite hm, if_pos hfi, if_pos hf] + rw [condExp_of_sigmaFinite hm, if_pos hfi, if_pos hf] -theorem condexp_const (hm : m ≤ m0) (c : F') [IsFiniteMeasure μ] : - μ[fun _ : α => c|m] = fun _ => c := - condexp_of_stronglyMeasurable hm (@stronglyMeasurable_const _ _ m _ _) (integrable_const c) +@[deprecated (since := "2025-01-21")] +alias condexp_of_stronglyMeasurable := condExp_of_stronglyMeasurable -theorem condexp_ae_eq_condexpL1 (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm)] (f : α → F') : - μ[f|m] =ᵐ[μ] condexpL1 hm μ f := by - rw [condexp_of_sigmaFinite hm] +@[simp] +theorem condExp_const (hm : m ≤ m₀) (c : E) [IsFiniteMeasure μ] : μ[fun _ : α ↦ c|m] = fun _ ↦ c := + condExp_of_stronglyMeasurable hm stronglyMeasurable_const (integrable_const c) + +@[deprecated (since := "2025-01-21")] alias condexp_const := condExp_const + +theorem condExp_ae_eq_condExpL1 (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] (f : α → E) : + μ[f|m] =ᵐ[μ] condExpL1 hm μ f := by + rw [condExp_of_sigmaFinite hm] by_cases hfi : Integrable f μ · rw [if_pos hfi] by_cases hfm : StronglyMeasurable[m] f · rw [if_pos hfm] - exact (condexpL1_of_aestronglyMeasurable' (StronglyMeasurable.aeStronglyMeasurable' hfm) - hfi).symm + exact (condExpL1_of_aestronglyMeasurable' hfm.aestronglyMeasurable hfi).symm · rw [if_neg hfm] - exact (AEStronglyMeasurable'.ae_eq_mk aestronglyMeasurable'_condexpL1).symm - rw [if_neg hfi, condexpL1_undef hfi] + exact aestronglyMeasurable_condExpL1.ae_eq_mk.symm + rw [if_neg hfi, condExpL1_undef hfi] exact (coeFn_zero _ _ _).symm -theorem condexp_ae_eq_condexpL1CLM (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hf : Integrable f μ) : - μ[f|m] =ᵐ[μ] condexpL1CLM F' hm μ (hf.toL1 f) := by - refine (condexp_ae_eq_condexpL1 hm f).trans (Eventually.of_forall fun x => ?_) - rw [condexpL1_eq hf] +@[deprecated (since := "2025-01-21")] alias condexp_ae_eq_condexpL1 := condExp_ae_eq_condExpL1 + +theorem condExp_ae_eq_condExpL1CLM (hm : m ≤ m₀) [SigmaFinite (μ.trim hm)] (hf : Integrable f μ) : + μ[f|m] =ᵐ[μ] condExpL1CLM E hm μ (hf.toL1 f) := by + refine (condExp_ae_eq_condExpL1 hm f).trans (Eventually.of_forall fun x => ?_) + rw [condExpL1_eq hf] -theorem condexp_undef (hf : ¬Integrable f μ) : μ[f|m] = 0 := by - by_cases hm : m ≤ m0 - swap; · rw [condexp_of_not_le hm] +@[deprecated (since := "2025-01-21")] alias condexp_ae_eq_condexpL1CLM := condExp_ae_eq_condExpL1CLM + +theorem condExp_of_not_integrable (hf : ¬Integrable f μ) : μ[f|m] = 0 := by + by_cases hm : m ≤ m₀ + swap; · rw [condExp_of_not_le hm] by_cases hμm : SigmaFinite (μ.trim hm) - swap; · rw [condexp_of_not_sigmaFinite hm hμm] - haveI : SigmaFinite (μ.trim hm) := hμm - rw [condexp_of_sigmaFinite, if_neg hf] + swap; · rw [condExp_of_not_sigmaFinite hm hμm] + rw [condExp_of_sigmaFinite, if_neg hf] + +@[deprecated (since := "2025-01-21")] alias condexp_undef := condExp_of_not_integrable +@[deprecated (since := "2025-01-21")] alias condExp_undef := condExp_of_not_integrable @[simp] -theorem condexp_zero : μ[(0 : α → F')|m] = 0 := by - by_cases hm : m ≤ m0 - swap; · rw [condexp_of_not_le hm] +theorem condExp_zero : μ[(0 : α → E)|m] = 0 := by + by_cases hm : m ≤ m₀ + swap; · rw [condExp_of_not_le hm] by_cases hμm : SigmaFinite (μ.trim hm) - swap; · rw [condexp_of_not_sigmaFinite hm hμm] - haveI : SigmaFinite (μ.trim hm) := hμm - exact - condexp_of_stronglyMeasurable hm (@stronglyMeasurable_zero _ _ m _ _) (integrable_zero _ _ _) - -theorem stronglyMeasurable_condexp : StronglyMeasurable[m] (μ[f|m]) := by - by_cases hm : m ≤ m0 - swap; · rw [condexp_of_not_le hm]; exact stronglyMeasurable_zero + swap; · rw [condExp_of_not_sigmaFinite hm hμm] + exact condExp_of_stronglyMeasurable hm stronglyMeasurable_zero (integrable_zero _ _ _) + +@[deprecated (since := "2025-01-21")] alias condexp_zero := condExp_zero + +theorem stronglyMeasurable_condExp : StronglyMeasurable[m] (μ[f|m]) := by + by_cases hm : m ≤ m₀ + swap; · rw [condExp_of_not_le hm]; exact stronglyMeasurable_zero by_cases hμm : SigmaFinite (μ.trim hm) - swap; · rw [condexp_of_not_sigmaFinite hm hμm]; exact stronglyMeasurable_zero - haveI : SigmaFinite (μ.trim hm) := hμm - rw [condexp_of_sigmaFinite hm] + swap; · rw [condExp_of_not_sigmaFinite hm hμm]; exact stronglyMeasurable_zero + rw [condExp_of_sigmaFinite hm] split_ifs with hfi hfm · exact hfm - · exact AEStronglyMeasurable'.stronglyMeasurable_mk _ + · exact aestronglyMeasurable_condExpL1.stronglyMeasurable_mk · exact stronglyMeasurable_zero -theorem condexp_congr_ae (h : f =ᵐ[μ] g) : μ[f|m] =ᵐ[μ] μ[g|m] := by - by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm]; rfl +@[deprecated (since := "2025-01-21")] alias stronglyMeasurable_condexp := stronglyMeasurable_condExp + +theorem condExp_congr_ae (h : f =ᵐ[μ] g) : μ[f|m] =ᵐ[μ] μ[g|m] := by + by_cases hm : m ≤ m₀ + swap; · simp_rw [condExp_of_not_le hm]; rfl by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; rfl - haveI : SigmaFinite (μ.trim hm) := hμm - exact (condexp_ae_eq_condexpL1 hm f).trans - (Filter.EventuallyEq.trans (by rw [condexpL1_congr_ae hm h]) - (condexp_ae_eq_condexpL1 hm g).symm) - -theorem condexp_of_aestronglyMeasurable' (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm)] {f : α → F'} - (hf : AEStronglyMeasurable' m f μ) (hfi : Integrable f μ) : μ[f|m] =ᵐ[μ] f := by - refine ((condexp_congr_ae hf.ae_eq_mk).trans ?_).trans hf.ae_eq_mk.symm - rw [condexp_of_stronglyMeasurable hm hf.stronglyMeasurable_mk + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; rfl + exact (condExp_ae_eq_condExpL1 hm f).trans + (Filter.EventuallyEq.trans (by rw [condExpL1_congr_ae hm h]) + (condExp_ae_eq_condExpL1 hm g).symm) + +@[deprecated (since := "2025-01-21")] alias condexp_congr_ae := condExp_congr_ae + +theorem condExp_of_aestronglyMeasurable' (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] {f : α → E} + (hf : AEStronglyMeasurable[m] f μ) (hfi : Integrable f μ) : μ[f|m] =ᵐ[μ] f := by + refine ((condExp_congr_ae hf.ae_eq_mk).trans ?_).trans hf.ae_eq_mk.symm + rw [condExp_of_stronglyMeasurable hm hf.stronglyMeasurable_mk ((integrable_congr hf.ae_eq_mk).mp hfi)] -theorem integrable_condexp : Integrable (μ[f|m]) μ := by - by_cases hm : m ≤ m0 - swap; · rw [condexp_of_not_le hm]; exact integrable_zero _ _ _ +@[deprecated (since := "2025-01-21")] +alias condexp_of_aestronglyMeasurable' := condExp_of_aestronglyMeasurable' + +@[fun_prop] +theorem integrable_condExp : Integrable (μ[f|m]) μ := by + by_cases hm : m ≤ m₀ + swap; · rw [condExp_of_not_le hm]; exact integrable_zero _ _ _ by_cases hμm : SigmaFinite (μ.trim hm) - swap; · rw [condexp_of_not_sigmaFinite hm hμm]; exact integrable_zero _ _ _ - haveI : SigmaFinite (μ.trim hm) := hμm - exact (integrable_condexpL1 f).congr (condexp_ae_eq_condexpL1 hm f).symm + swap; · rw [condExp_of_not_sigmaFinite hm hμm]; exact integrable_zero _ _ _ + exact (integrable_condExpL1 f).congr (condExp_ae_eq_condExpL1 hm f).symm + +@[deprecated (since := "2025-01-21")] alias integrable_condexp := integrable_condExp /-- The integral of the conditional expectation `μ[f|hm]` over an `m`-measurable set is equal to the integral of `f` on that set. -/ -theorem setIntegral_condexp (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hf : Integrable f μ) +theorem setIntegral_condExp (hm : m ≤ m₀) [SigmaFinite (μ.trim hm)] (hf : Integrable f μ) (hs : MeasurableSet[m] s) : ∫ x in s, (μ[f|m]) x ∂μ = ∫ x in s, f x ∂μ := by - rw [setIntegral_congr_ae (hm s hs) ((condexp_ae_eq_condexpL1 hm f).mono fun x hx _ => hx)] - exact setIntegral_condexpL1 hf hs + rw [setIntegral_congr_ae (hm s hs) ((condExp_ae_eq_condExpL1 hm f).mono fun x hx _ => hx)] + exact setIntegral_condExpL1 hf hs + +@[deprecated (since := "2025-01-21")] alias setIntegral_condexp := setIntegral_condExp -@[deprecated (since := "2024-04-17")] alias set_integral_condexp := setIntegral_condexp +@[deprecated (since := "2024-04-17")] alias set_integral_condexp := setIntegral_condExp -theorem integral_condexp (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm)] : +theorem integral_condExp (hm : m ≤ m₀) [hμm : SigmaFinite (μ.trim hm)] : ∫ x, (μ[f|m]) x ∂μ = ∫ x, f x ∂μ := by by_cases hf : Integrable f μ · suffices ∫ x in Set.univ, (μ[f|m]) x ∂μ = ∫ x in Set.univ, f x ∂μ by simp_rw [setIntegral_univ] at this; exact this - exact setIntegral_condexp hm hf (@MeasurableSet.univ _ m) - simp only [condexp_undef hf, Pi.zero_apply, integral_zero, integral_undef hf] + exact setIntegral_condExp hm hf .univ + simp only [condExp_of_not_integrable hf, Pi.zero_apply, integral_zero, integral_undef hf] -/-- Total probability law using `condexp` as conditional probability. -/ -theorem integral_condexp_indicator [mF : MeasurableSpace F] {Y : α → F} (hY : Measurable Y) +@[deprecated (since := "2025-01-21")] alias integral_condexp := integral_condExp + +/-- **Law of total probability** using `condExp` as conditional probability. -/ +theorem integral_condExp_indicator [mβ : MeasurableSpace β] {Y : α → β} (hY : Measurable Y) [SigmaFinite (μ.trim hY.comap_le)] {A : Set α} (hA : MeasurableSet A) : - ∫ x, (μ[(A.indicator fun _ ↦ (1 : ℝ)) | mF.comap Y]) x ∂μ = (μ A).toReal := by - rw [integral_condexp, integral_indicator hA, setIntegral_const, smul_eq_mul, mul_one] + ∫ x, (μ[(A.indicator fun _ ↦ (1 : ℝ)) | mβ.comap Y]) x ∂μ = (μ A).toReal := by + rw [integral_condExp, integral_indicator hA, setIntegral_const, smul_eq_mul, mul_one] + +@[deprecated (since := "2025-01-21")] alias integral_condexp_indicator := integral_condExp_indicator /-- **Uniqueness of the conditional expectation** If a function is a.e. `m`-measurable, verifies an integrability condition and has same integral as `f` on all `m`-measurable sets, then it is a.e. equal to `μ[f|hm]`. -/ -theorem ae_eq_condexp_of_forall_setIntegral_eq (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] - {f g : α → F'} (hf : Integrable f μ) +theorem ae_eq_condExp_of_forall_setIntegral_eq (hm : m ≤ m₀) [SigmaFinite (μ.trim hm)] + {f g : α → E} (hf : Integrable f μ) (hg_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn g s μ) (hg_eq : ∀ s : Set α, MeasurableSet[m] s → μ s < ∞ → ∫ x in s, g x ∂μ = ∫ x in s, f x ∂μ) - (hgm : AEStronglyMeasurable' m g μ) : g =ᵐ[μ] μ[f|m] := by + (hgm : AEStronglyMeasurable[m] g μ) : g =ᵐ[μ] μ[f|m] := by refine ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' hm hg_int_finite - (fun s _ _ => integrable_condexp.integrableOn) (fun s hs hμs => ?_) hgm - (StronglyMeasurable.aeStronglyMeasurable' stronglyMeasurable_condexp) - rw [hg_eq s hs hμs, setIntegral_condexp hm hf hs] + (fun s _ _ => integrable_condExp.integrableOn) (fun s hs hμs => ?_) hgm + (StronglyMeasurable.aestronglyMeasurable stronglyMeasurable_condExp) + rw [hg_eq s hs hμs, setIntegral_condExp hm hf hs] + +@[deprecated (since := "2025-01-21")] +alias ae_eq_condexp_of_forall_setIntegral_eq := ae_eq_condExp_of_forall_setIntegral_eq @[deprecated (since := "2024-04-17")] -alias ae_eq_condexp_of_forall_set_integral_eq := ae_eq_condexp_of_forall_setIntegral_eq +alias ae_eq_condExp_of_forall_set_integral_eq := ae_eq_condExp_of_forall_setIntegral_eq -theorem condexp_bot' [hμ : NeZero μ] (f : α → F') : +@[deprecated (since := "2025-01-21")] +alias ae_eq_condexp_of_forall_set_integral_eq := ae_eq_condExp_of_forall_set_integral_eq + +theorem condExp_bot' [hμ : NeZero μ] (f : α → E) : μ[f|⊥] = fun _ => (μ Set.univ).toReal⁻¹ • ∫ x, f x ∂μ := by by_cases hμ_finite : IsFiniteMeasure μ swap · have h : ¬SigmaFinite (μ.trim bot_le) := by rwa [sigmaFinite_trim_bot_iff] rw [not_isFiniteMeasure_iff] at hμ_finite - rw [condexp_of_not_sigmaFinite bot_le h] + rw [condExp_of_not_sigmaFinite bot_le h] simp only [hμ_finite, ENNReal.top_toReal, inv_zero, zero_smul] rfl - have h_meas : StronglyMeasurable[⊥] (μ[f|⊥]) := stronglyMeasurable_condexp + have h_meas : StronglyMeasurable[⊥] (μ[f|⊥]) := stronglyMeasurable_condExp obtain ⟨c, h_eq⟩ := stronglyMeasurable_bot_iff.mp h_meas rw [h_eq] - have h_integral : ∫ x, (μ[f|⊥]) x ∂μ = ∫ x, f x ∂μ := integral_condexp bot_le + have h_integral : ∫ x, (μ[f|⊥]) x ∂μ = ∫ x, f x ∂μ := integral_condExp bot_le simp_rw [h_eq, integral_const] at h_integral rw [← h_integral, ← smul_assoc, smul_eq_mul, inv_mul_cancel₀, one_smul] rw [Ne, ENNReal.toReal_eq_zero_iff, not_or] exact ⟨NeZero.ne _, measure_ne_top μ Set.univ⟩ -theorem condexp_bot_ae_eq (f : α → F') : +@[deprecated (since := "2025-01-21")] alias condexp_bot' := condExp_bot' + +theorem condExp_bot_ae_eq (f : α → E) : μ[f|⊥] =ᵐ[μ] fun _ => (μ Set.univ).toReal⁻¹ • ∫ x, f x ∂μ := by rcases eq_zero_or_neZero μ with rfl | hμ · rw [ae_zero]; exact eventually_bot - · exact Eventually.of_forall <| congr_fun (condexp_bot' f) + · exact Eventually.of_forall <| congr_fun (condExp_bot' f) + +@[deprecated (since := "2025-01-21")] alias condexp_bot_ae_eq := condExp_bot_ae_eq + +theorem condExp_bot [IsProbabilityMeasure μ] (f : α → E) : μ[f|⊥] = fun _ => ∫ x, f x ∂μ := by + refine (condExp_bot' f).trans ?_; rw [measure_univ, ENNReal.one_toReal, inv_one, one_smul] -theorem condexp_bot [IsProbabilityMeasure μ] (f : α → F') : μ[f|⊥] = fun _ => ∫ x, f x ∂μ := by - refine (condexp_bot' f).trans ?_; rw [measure_univ, ENNReal.one_toReal, inv_one, one_smul] +@[deprecated (since := "2025-01-21")] alias condexp_bot := condExp_bot -theorem condexp_add (hf : Integrable f μ) (hg : Integrable g μ) : +theorem condExp_add (hf : Integrable f μ) (hg : Integrable g μ) (m : MeasurableSpace α) : μ[f + g|m] =ᵐ[μ] μ[f|m] + μ[g|m] := by - by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm]; simp + by_cases hm : m ≤ m₀ + swap; · simp_rw [condExp_of_not_le hm]; simp by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; simp - haveI : SigmaFinite (μ.trim hm) := hμm - refine (condexp_ae_eq_condexpL1 hm _).trans ?_ - rw [condexpL1_add hf hg] + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; simp + refine (condExp_ae_eq_condExpL1 hm _).trans ?_ + rw [condExpL1_add hf hg] exact (coeFn_add _ _).trans - ((condexp_ae_eq_condexpL1 hm _).symm.add (condexp_ae_eq_condexpL1 hm _).symm) + ((condExp_ae_eq_condExpL1 hm _).symm.add (condExp_ae_eq_condExpL1 hm _).symm) -theorem condexp_finset_sum {ι : Type*} {s : Finset ι} {f : ι → α → F'} - (hf : ∀ i ∈ s, Integrable (f i) μ) : μ[∑ i ∈ s, f i|m] =ᵐ[μ] ∑ i ∈ s, μ[f i|m] := by +@[deprecated (since := "2025-01-21")] alias condexp_add := condExp_add + +theorem condExp_finset_sum {ι : Type*} {s : Finset ι} {f : ι → α → E} + (hf : ∀ i ∈ s, Integrable (f i) μ) (m : MeasurableSpace α) : + μ[∑ i ∈ s, f i|m] =ᵐ[μ] ∑ i ∈ s, μ[f i|m] := by induction' s using Finset.induction_on with i s his heq hf - · rw [Finset.sum_empty, Finset.sum_empty, condexp_zero] + · rw [Finset.sum_empty, Finset.sum_empty, condExp_zero] · rw [Finset.sum_insert his, Finset.sum_insert his] - exact (condexp_add (hf i <| Finset.mem_insert_self i s) <| - integrable_finset_sum' _ fun j hmem => hf j <| Finset.mem_insert_of_mem hmem).trans - ((EventuallyEq.refl _ _).add (heq fun j hmem => hf j <| Finset.mem_insert_of_mem hmem)) + exact (condExp_add (hf i <| Finset.mem_insert_self i s) + (integrable_finset_sum' _ <| Finset.forall_of_forall_insert hf) _).trans + ((EventuallyEq.refl _ _).add <| heq <| Finset.forall_of_forall_insert hf) + +@[deprecated (since := "2025-01-21")] alias condexp_finset_sum := condExp_finset_sum -theorem condexp_smul (c : 𝕜) (f : α → F') : μ[c • f|m] =ᵐ[μ] c • μ[f|m] := by - by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm]; simp +theorem condExp_smul [NormedSpace 𝕜 E] (c : 𝕜) (f : α → E) (m : MeasurableSpace α) : + μ[c • f|m] =ᵐ[μ] c • μ[f|m] := by + by_cases hm : m ≤ m₀ + swap; · simp_rw [condExp_of_not_le hm]; simp by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; simp - haveI : SigmaFinite (μ.trim hm) := hμm - refine (condexp_ae_eq_condexpL1 hm _).trans ?_ - rw [condexpL1_smul c f] - refine (@condexp_ae_eq_condexpL1 _ _ _ _ _ m _ _ hm _ f).mp ?_ - refine (coeFn_smul c (condexpL1 hm μ f)).mono fun x hx1 hx2 => ?_ + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; simp + refine (condExp_ae_eq_condExpL1 hm _).trans ?_ + rw [condExpL1_smul c f] + refine (condExp_ae_eq_condExpL1 hm f).mp ?_ + refine (coeFn_smul c (condExpL1 hm μ f)).mono fun x hx1 hx2 => ?_ simp only [hx1, hx2, Pi.smul_apply] -theorem condexp_neg (f : α → F') : μ[-f|m] =ᵐ[μ] -μ[f|m] := by - letI : Module ℝ (α → F') := @Pi.module α (fun _ => F') ℝ _ _ fun _ => inferInstance +@[deprecated (since := "2025-01-21")] alias condexp_smul := condExp_smul + +theorem condExp_neg (f : α → E) (m : MeasurableSpace α) : μ[-f|m] =ᵐ[μ] -μ[f|m] := by calc μ[-f|m] = μ[(-1 : ℝ) • f|m] := by rw [neg_one_smul ℝ f] - _ =ᵐ[μ] (-1 : ℝ) • μ[f|m] := condexp_smul (-1) f + _ =ᵐ[μ] (-1 : ℝ) • μ[f|m] := condExp_smul .. _ = -μ[f|m] := neg_one_smul ℝ (μ[f|m]) -theorem condexp_sub (hf : Integrable f μ) (hg : Integrable g μ) : +@[deprecated (since := "2025-01-21")] alias condexp_neg := condExp_neg + +theorem condExp_sub (hf : Integrable f μ) (hg : Integrable g μ) (m : MeasurableSpace α) : μ[f - g|m] =ᵐ[μ] μ[f|m] - μ[g|m] := by simp_rw [sub_eq_add_neg] - exact (condexp_add hf hg.neg).trans (EventuallyEq.rfl.add (condexp_neg g)) + exact (condExp_add hf hg.neg _).trans (EventuallyEq.rfl.add (condExp_neg ..)) + +@[deprecated (since := "2025-01-21")] alias condexp_sub := condExp_sub + +/-- **Tower property of the conditional expectation**. -theorem condexp_condexp_of_le {m₁ m₂ m0 : MeasurableSpace α} {μ : Measure α} (hm₁₂ : m₁ ≤ m₂) - (hm₂ : m₂ ≤ m0) [SigmaFinite (μ.trim hm₂)] : μ[μ[f|m₂]|m₁] =ᵐ[μ] μ[f|m₁] := by +Taking the `m₂`-conditional expectation then the `m₁`-conditional expectation, where `m₁` is a +smaller σ-algebra, is the same as taking the `m₁`-conditional expectation directly. -/ +theorem condExp_condExp_of_le {m₁ m₂ m₀ : MeasurableSpace α} {μ : Measure α} (hm₁₂ : m₁ ≤ m₂) + (hm₂ : m₂ ≤ m₀) [SigmaFinite (μ.trim hm₂)] : μ[μ[f|m₂]|m₁] =ᵐ[μ] μ[f|m₁] := by by_cases hμm₁ : SigmaFinite (μ.trim (hm₁₂.trans hm₂)) - swap; · simp_rw [condexp_of_not_sigmaFinite (hm₁₂.trans hm₂) hμm₁]; rfl - haveI : SigmaFinite (μ.trim (hm₁₂.trans hm₂)) := hμm₁ + swap; · simp_rw [condExp_of_not_sigmaFinite (hm₁₂.trans hm₂) hμm₁]; rfl by_cases hf : Integrable f μ - swap; · simp_rw [condexp_undef hf, condexp_zero]; rfl + swap; · simp_rw [condExp_of_not_integrable hf, condExp_zero]; rfl refine ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' (hm₁₂.trans hm₂) - (fun s _ _ => integrable_condexp.integrableOn) - (fun s _ _ => integrable_condexp.integrableOn) ?_ - (StronglyMeasurable.aeStronglyMeasurable' stronglyMeasurable_condexp) - (StronglyMeasurable.aeStronglyMeasurable' stronglyMeasurable_condexp) + (fun s _ _ => integrable_condExp.integrableOn) (fun s _ _ => integrable_condExp.integrableOn) ?_ + stronglyMeasurable_condExp.aestronglyMeasurable + stronglyMeasurable_condExp.aestronglyMeasurable intro s hs _ - rw [setIntegral_condexp (hm₁₂.trans hm₂) integrable_condexp hs] - rw [setIntegral_condexp (hm₁₂.trans hm₂) hf hs, setIntegral_condexp hm₂ hf (hm₁₂ s hs)] + rw [setIntegral_condExp (hm₁₂.trans hm₂) integrable_condExp hs] + rw [setIntegral_condExp (hm₁₂.trans hm₂) hf hs, setIntegral_condExp hm₂ hf (hm₁₂ s hs)] + +@[deprecated (since := "2025-01-21")] alias condexp_condexp_of_le := condExp_condExp_of_le + +section RCLike +variable [InnerProductSpace 𝕜 E] + +lemma Memℒp.condExpL2_ae_eq_condExp' (hm : m ≤ m₀) (hf1 : Integrable f μ) (hf2 : Memℒp f 2 μ) + [SigmaFinite (μ.trim hm)] : condExpL2 E 𝕜 hm hf2.toLp =ᵐ[μ] μ[f | m] := by + refine ae_eq_condExp_of_forall_setIntegral_eq hm hf1 + (fun s hs htop ↦ integrableOn_condExpL2_of_measure_ne_top hm htop.ne _) (fun s hs htop ↦ ?_) + (aestronglyMeasurable_condExpL2 hm _) + rw [integral_condExpL2_eq hm (hf2.toLp _) hs htop.ne] + refine setIntegral_congr_ae (hm _ hs) ?_ + filter_upwards [hf2.coeFn_toLp] with ω hω _ using hω + +lemma Memℒp.condExpL2_ae_eq_condExp (hm : m ≤ m₀) (hf : Memℒp f 2 μ) [IsFiniteMeasure μ] : + condExpL2 E 𝕜 hm hf.toLp =ᵐ[μ] μ[f | m] := + hf.condExpL2_ae_eq_condExp' hm (memℒp_one_iff_integrable.1 <| hf.mono_exponent one_le_two) + +end RCLike + +section Real +variable [InnerProductSpace ℝ E] + +-- TODO: Generalize via the conditional Jensen inequality +lemma eLpNorm_condExp_le : eLpNorm (μ[f | m]) 2 μ ≤ eLpNorm f 2 μ := by + by_cases hm : m ≤ m₀; swap + · simp [condExp_of_not_le hm] + by_cases hfμ : SigmaFinite (μ.trim hm); swap + · rw [condExp_of_not_sigmaFinite hm hfμ] + simp + by_cases hfi : Integrable f μ; swap + · rw [condExp_of_not_integrable hfi] + simp + obtain hf | hf := eq_or_ne (eLpNorm f 2 μ) ∞ + · simp [hf] + replace hf : Memℒp f 2 μ := ⟨hfi.1, Ne.lt_top' fun a ↦ hf (id (Eq.symm a))⟩ + rw [← eLpNorm_congr_ae (hf.condExpL2_ae_eq_condExp' (𝕜 := ℝ) hm hfi)] + refine le_trans (eLpNorm_condExpL2_le hm _) ?_ + rw [eLpNorm_congr_ae hf.coeFn_toLp] + +protected lemma Memℒp.condExp (hf : Memℒp f 2 μ) : Memℒp (μ[f | m]) 2 μ := by + by_cases hm : m ≤ m₀ + · exact ⟨(stronglyMeasurable_condExp.mono hm).aestronglyMeasurable, + eLpNorm_condExp_le.trans_lt hf.eLpNorm_lt_top⟩ + · simp [condExp_of_not_le hm] + +end Real +end NormedAddCommGroup + +section NormedRing +variable {R : Type*} [NormedRing R] [NormedSpace ℝ R] [CompleteSpace R] -theorem condexp_mono {E} [NormedLatticeAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] - [OrderedSMul ℝ E] {f g : α → E} (hf : Integrable f μ) (hg : Integrable g μ) (hfg : f ≤ᵐ[μ] g) : - μ[f|m] ≤ᵐ[μ] μ[g|m] := by - by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm]; rfl - by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; rfl - haveI : SigmaFinite (μ.trim hm) := hμm - exact (condexp_ae_eq_condexpL1 hm _).trans_le - ((condexpL1_mono hf hg hfg).trans_eq (condexp_ae_eq_condexpL1 hm _).symm) +@[simp] +lemma condExp_ofNat (n : ℕ) [n.AtLeastTwo] (f : α → R) : + μ[ofNat(n) * f|m] =ᵐ[μ] ofNat(n) * μ[f|m] := by + simpa [Nat.cast_smul_eq_nsmul] using condExp_smul (μ := μ) (m := m) (n : ℝ) f -theorem condexp_nonneg {E} [NormedLatticeAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] - [OrderedSMul ℝ E] {f : α → E} (hf : 0 ≤ᵐ[μ] f) : 0 ≤ᵐ[μ] μ[f|m] := by - by_cases hfint : Integrable f μ - · rw [(condexp_zero.symm : (0 : α → E) = μ[0|m])] - exact condexp_mono (integrable_zero _ _ _) hfint hf - · rw [condexp_undef hfint] +end NormedRing -theorem condexp_nonpos {E} [NormedLatticeAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] - [OrderedSMul ℝ E] {f : α → E} (hf : f ≤ᵐ[μ] 0) : μ[f|m] ≤ᵐ[μ] 0 := by - by_cases hfint : Integrable f μ - · rw [(condexp_zero.symm : (0 : α → E) = μ[0|m])] - exact condexp_mono hfint (integrable_zero _ _ _) hf - · rw [condexp_undef hfint] +section NormedLatticeAddCommGroup +variable [NormedLatticeAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] /-- **Lebesgue dominated convergence theorem**: sufficient conditions under which almost everywhere convergence of a sequence of functions implies the convergence of their image by - `condexpL1`. -/ -theorem tendsto_condexpL1_of_dominated_convergence (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] - {fs : ℕ → α → F'} {f : α → F'} (bound_fs : α → ℝ) + `condExpL1`. -/ +theorem tendsto_condExpL1_of_dominated_convergence (hm : m ≤ m₀) [SigmaFinite (μ.trim hm)] + {fs : ℕ → α → E} {f : α → E} (bound_fs : α → ℝ) (hfs_meas : ∀ n, AEStronglyMeasurable (fs n) μ) (h_int_bound_fs : Integrable bound_fs μ) (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ‖fs n x‖ ≤ bound_fs x) (hfs : ∀ᵐ x ∂μ, Tendsto (fun n => fs n x) atTop (𝓝 (f x))) : - Tendsto (fun n => condexpL1 hm μ (fs n)) atTop (𝓝 (condexpL1 hm μ f)) := + Tendsto (fun n => condExpL1 hm μ (fs n)) atTop (𝓝 (condExpL1 hm μ f)) := tendsto_setToFun_of_dominated_convergence _ bound_fs hfs_meas h_int_bound_fs hfs_bound hfs +@[deprecated (since := "2025-01-21")] +alias tendsto_condexpL1_of_dominated_convergence := tendsto_condExpL1_of_dominated_convergence + /-- If two sequences of functions have a.e. equal conditional expectations at each step, converge and verify dominated convergence hypotheses, then the conditional expectations of their limits are a.e. equal. -/ -theorem tendsto_condexp_unique (fs gs : ℕ → α → F') (f g : α → F') +theorem tendsto_condExp_unique (fs gs : ℕ → α → E) (f g : α → E) (hfs_int : ∀ n, Integrable (fs n) μ) (hgs_int : ∀ n, Integrable (gs n) μ) (hfs : ∀ᵐ x ∂μ, Tendsto (fun n => fs n x) atTop (𝓝 (f x))) (hgs : ∀ᵐ x ∂μ, Tendsto (fun n => gs n x) atTop (𝓝 (g x))) (bound_fs : α → ℝ) @@ -369,22 +463,51 @@ theorem tendsto_condexp_unique (fs gs : ℕ → α → F') (f g : α → F') (h_int_bound_gs : Integrable bound_gs μ) (hfs_bound : ∀ n, ∀ᵐ x ∂μ, ‖fs n x‖ ≤ bound_fs x) (hgs_bound : ∀ n, ∀ᵐ x ∂μ, ‖gs n x‖ ≤ bound_gs x) (hfg : ∀ n, μ[fs n|m] =ᵐ[μ] μ[gs n|m]) : μ[f|m] =ᵐ[μ] μ[g|m] := by - by_cases hm : m ≤ m0; swap; · simp_rw [condexp_of_not_le hm]; rfl - by_cases hμm : SigmaFinite (μ.trim hm); swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; rfl - haveI : SigmaFinite (μ.trim hm) := hμm - refine (condexp_ae_eq_condexpL1 hm f).trans ((condexp_ae_eq_condexpL1 hm g).trans ?_).symm + by_cases hm : m ≤ m₀; swap; · simp_rw [condExp_of_not_le hm]; rfl + by_cases hμm : SigmaFinite (μ.trim hm); swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; rfl + refine (condExp_ae_eq_condExpL1 hm f).trans ((condExp_ae_eq_condExpL1 hm g).trans ?_).symm rw [← Lp.ext_iff] - have hn_eq : ∀ n, condexpL1 hm μ (gs n) = condexpL1 hm μ (fs n) := by + have hn_eq : ∀ n, condExpL1 hm μ (gs n) = condExpL1 hm μ (fs n) := by intro n ext1 - refine (condexp_ae_eq_condexpL1 hm (gs n)).symm.trans ((hfg n).symm.trans ?_) - exact condexp_ae_eq_condexpL1 hm (fs n) - have hcond_fs : Tendsto (fun n => condexpL1 hm μ (fs n)) atTop (𝓝 (condexpL1 hm μ f)) := - tendsto_condexpL1_of_dominated_convergence hm _ (fun n => (hfs_int n).1) h_int_bound_fs + refine (condExp_ae_eq_condExpL1 hm (gs n)).symm.trans ((hfg n).symm.trans ?_) + exact condExp_ae_eq_condExpL1 hm (fs n) + have hcond_fs : Tendsto (fun n => condExpL1 hm μ (fs n)) atTop (𝓝 (condExpL1 hm μ f)) := + tendsto_condExpL1_of_dominated_convergence hm _ (fun n => (hfs_int n).1) h_int_bound_fs hfs_bound hfs - have hcond_gs : Tendsto (fun n => condexpL1 hm μ (gs n)) atTop (𝓝 (condexpL1 hm μ g)) := - tendsto_condexpL1_of_dominated_convergence hm _ (fun n => (hgs_int n).1) h_int_bound_gs + have hcond_gs : Tendsto (fun n => condExpL1 hm μ (gs n)) atTop (𝓝 (condExpL1 hm μ g)) := + tendsto_condExpL1_of_dominated_convergence hm _ (fun n => (hgs_int n).1) h_int_bound_gs hgs_bound hgs exact tendsto_nhds_unique_of_eventuallyEq hcond_gs hcond_fs (Eventually.of_forall hn_eq) +@[deprecated (since := "2025-01-21")] alias tendsto_condexp_unique := tendsto_condExp_unique + +variable [OrderedSMul ℝ E] + +lemma condExp_mono (hf : Integrable f μ) (hg : Integrable g μ) (hfg : f ≤ᵐ[μ] g) : + μ[f|m] ≤ᵐ[μ] μ[g|m] := by + by_cases hm : m ≤ m₀ + swap; · simp_rw [condExp_of_not_le hm]; rfl + by_cases hμm : SigmaFinite (μ.trim hm) + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; rfl + exact (condExp_ae_eq_condExpL1 hm _).trans_le + ((condExpL1_mono hf hg hfg).trans_eq (condExp_ae_eq_condExpL1 hm _).symm) + +lemma condExp_nonneg (hf : 0 ≤ᵐ[μ] f) : 0 ≤ᵐ[μ] μ[f|m] := by + by_cases hfint : Integrable f μ + · rw [(condExp_zero.symm : (0 : α → E) = μ[0|m])] + exact condExp_mono (integrable_zero _ _ _) hfint hf + · rw [condExp_of_not_integrable hfint] + +lemma condExp_nonpos (hf : f ≤ᵐ[μ] 0) : μ[f|m] ≤ᵐ[μ] 0 := by + by_cases hfint : Integrable f μ + · rw [(condExp_zero.symm : (0 : α → E) = μ[0|m])] + exact condExp_mono hfint (integrable_zero _ _ _) hf + · rw [condExp_of_not_integrable hfint] + +@[deprecated (since := "2025-01-21")] alias condexp_mono := condExp_mono +@[deprecated (since := "2025-01-21")] alias condexp_nonneg := condExp_nonneg +@[deprecated (since := "2025-01-21")] alias condexp_nonpos := condExp_nonpos + +end NormedLatticeAddCommGroup end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index b90f7b167be62..1871e8db5807f 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -17,12 +17,12 @@ The conditional expectation of an `L²` function is defined in is integrable and define a map `Set α → (E →L[ℝ] (α →₁[μ] E))` which to a set associates a linear map. That linear map sends `x ∈ E` to the conditional expectation of the indicator of the set with value `x`. -* Extend that map to `condexpL1CLM : (α →₁[μ] E) →L[ℝ] (α →₁[μ] E)`. This is done using the same +* Extend that map to `condExpL1CLM : (α →₁[μ] E) →L[ℝ] (α →₁[μ] E)`. This is done using the same construction as the Bochner integral (see the file `MeasureTheory/Integral/SetToL1`). ## Main definitions -* `condexpL1`: Conditional expectation of a function as a linear map from `L1` to itself. +* `condExpL1`: Conditional expectation of a function as a linear map from `L1` to itself. -/ @@ -54,7 +54,7 @@ section CondexpInd /-! ## Conditional expectation of an indicator as a continuous linear map. The goal of this section is to build -`condexpInd (hm : m ≤ m0) (μ : Measure α) (s : Set s) : G →L[ℝ] α →₁[μ] G`, which +`condExpInd (hm : m ≤ m0) (μ : Measure α) (s : Set s) : G →L[ℝ] α →₁[μ] G`, which takes `x : G` to the conditional expectation of the indicator of the set `s` with value `x`, seen as an element of `α →₁[μ] G`. -/ @@ -67,91 +67,107 @@ section CondexpIndL1Fin /-- Conditional expectation of the indicator of a measurable set with finite measure, as a function in L1. -/ -def condexpIndL1Fin (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) +def condExpIndL1Fin (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : α →₁[μ] G := - (integrable_condexpIndSMul hm hs hμs x).toL1 _ + (integrable_condExpIndSMul hm hs hμs x).toL1 _ -theorem condexpIndL1Fin_ae_eq_condexpIndSMul (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] +@[deprecated (since := "2025-01-21")] noncomputable alias condexpIndL1Fin := condExpIndL1Fin + +theorem condExpIndL1Fin_ae_eq_condExpIndSMul (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : - condexpIndL1Fin hm hs hμs x =ᵐ[μ] condexpIndSMul hm hs hμs x := - (integrable_condexpIndSMul hm hs hμs x).coeFn_toL1 + condExpIndL1Fin hm hs hμs x =ᵐ[μ] condExpIndSMul hm hs hμs x := + (integrable_condExpIndSMul hm hs hμs x).coeFn_toL1 + +@[deprecated (since := "2025-01-21")] +alias condexpIndL1Fin_ae_eq_condexpIndSMul := condExpIndL1Fin_ae_eq_condExpIndSMul variable {hm : m ≤ m0} [SigmaFinite (μ.trim hm)] -- Porting note: this lemma fills the hole in `refine' (Memℒp.coeFn_toLp _) ...` -- which is not automatically filled in Lean 4 private theorem q {hs : MeasurableSet s} {hμs : μ s ≠ ∞} {x : G} : - Memℒp (condexpIndSMul hm hs hμs x) 1 μ := by - rw [memℒp_one_iff_integrable]; apply integrable_condexpIndSMul + Memℒp (condExpIndSMul hm hs hμs x) 1 μ := by + rw [memℒp_one_iff_integrable]; apply integrable_condExpIndSMul -theorem condexpIndL1Fin_add (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x y : G) : - condexpIndL1Fin hm hs hμs (x + y) = - condexpIndL1Fin hm hs hμs x + condexpIndL1Fin hm hs hμs y := by +theorem condExpIndL1Fin_add (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x y : G) : + condExpIndL1Fin hm hs hμs (x + y) = + condExpIndL1Fin hm hs hμs x + condExpIndL1Fin hm hs hμs y := by ext1 refine (Memℒp.coeFn_toLp q).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_add _ _).symm refine EventuallyEq.trans ?_ (EventuallyEq.add (Memℒp.coeFn_toLp q).symm (Memℒp.coeFn_toLp q).symm) - rw [condexpIndSMul_add] + rw [condExpIndSMul_add] refine (Lp.coeFn_add _ _).trans (Eventually.of_forall fun a => ?_) rfl -theorem condexpIndL1Fin_smul (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : - condexpIndL1Fin hm hs hμs (c • x) = c • condexpIndL1Fin hm hs hμs x := by +@[deprecated (since := "2025-01-21")] alias condexpIndL1Fin_add := condExpIndL1Fin_add + +theorem condExpIndL1Fin_smul (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : + condExpIndL1Fin hm hs hμs (c • x) = c • condExpIndL1Fin hm hs hμs x := by ext1 refine (Memℒp.coeFn_toLp q).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_smul _ _).symm - rw [condexpIndSMul_smul hs hμs c x] + rw [condExpIndSMul_smul hs hμs c x] refine (Lp.coeFn_smul _ _).trans ?_ - refine (condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x).mono fun y hy => ?_ + refine (condExpIndL1Fin_ae_eq_condExpIndSMul hm hs hμs x).mono fun y hy => ?_ simp only [Pi.smul_apply, hy] -theorem condexpIndL1Fin_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (hs : MeasurableSet s) +@[deprecated (since := "2025-01-21")] alias condexpIndL1Fin_smul := condExpIndL1Fin_smul + +theorem condExpIndL1Fin_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : - condexpIndL1Fin hm hs hμs (c • x) = c • condexpIndL1Fin hm hs hμs x := by + condExpIndL1Fin hm hs hμs (c • x) = c • condExpIndL1Fin hm hs hμs x := by ext1 refine (Memℒp.coeFn_toLp q).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_smul _ _).symm - rw [condexpIndSMul_smul' hs hμs c x] + rw [condExpIndSMul_smul' hs hμs c x] refine (Lp.coeFn_smul _ _).trans ?_ - refine (condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x).mono fun y hy => ?_ + refine (condExpIndL1Fin_ae_eq_condExpIndSMul hm hs hμs x).mono fun y hy => ?_ simp only [Pi.smul_apply, hy] -theorem norm_condexpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : - ‖condexpIndL1Fin hm hs hμs x‖ ≤ (μ s).toReal * ‖x‖ := by +@[deprecated (since := "2025-01-21")] alias condexpIndL1Fin_smul' := condExpIndL1Fin_smul' + +theorem norm_condExpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : + ‖condExpIndL1Fin hm hs hμs x‖ ≤ (μ s).toReal * ‖x‖ := by rw [L1.norm_eq_integral_norm, ← ENNReal.toReal_ofReal (norm_nonneg x), ← ENNReal.toReal_mul, ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.mul_ne_top hμs ENNReal.ofReal_ne_top), ofReal_integral_norm_eq_lintegral_nnnorm] swap; · rw [← memℒp_one_iff_integrable]; exact Lp.memℒp _ have h_eq : - ∫⁻ a, ‖condexpIndL1Fin hm hs hμs x a‖₊ ∂μ = ∫⁻ a, ‖condexpIndSMul hm hs hμs x a‖₊ ∂μ := by + ∫⁻ a, ‖condExpIndL1Fin hm hs hμs x a‖₊ ∂μ = ∫⁻ a, ‖condExpIndSMul hm hs hμs x a‖₊ ∂μ := by refine lintegral_congr_ae ?_ - refine (condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x).mono fun z hz => ?_ + refine (condExpIndL1Fin_ae_eq_condExpIndSMul hm hs hμs x).mono fun z hz => ?_ dsimp only rw [hz] rw [h_eq, ofReal_norm_eq_coe_nnnorm] - exact lintegral_nnnorm_condexpIndSMul_le hm hs hμs x + exact lintegral_nnnorm_condExpIndSMul_le hm hs hμs x + +@[deprecated (since := "2025-01-21")] alias norm_condexpIndL1Fin_le := norm_condExpIndL1Fin_le -theorem condexpIndL1Fin_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) +theorem condExpIndL1Fin_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : - condexpIndL1Fin hm (hs.union ht) ((measure_union_le s t).trans_lt + condExpIndL1Fin hm (hs.union ht) ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x = - condexpIndL1Fin hm hs hμs x + condexpIndL1Fin hm ht hμt x := by + condExpIndL1Fin hm hs hμs x + condExpIndL1Fin hm ht hμt x := by ext1 have hμst := measure_union_ne_top hμs hμt - refine (condexpIndL1Fin_ae_eq_condexpIndSMul hm (hs.union ht) hμst x).trans ?_ + refine (condExpIndL1Fin_ae_eq_condExpIndSMul hm (hs.union ht) hμst x).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_add _ _).symm - have hs_eq := condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x - have ht_eq := condexpIndL1Fin_ae_eq_condexpIndSMul hm ht hμt x + have hs_eq := condExpIndL1Fin_ae_eq_condExpIndSMul hm hs hμs x + have ht_eq := condExpIndL1Fin_ae_eq_condExpIndSMul hm ht hμt x refine EventuallyEq.trans ?_ (EventuallyEq.add hs_eq.symm ht_eq.symm) - rw [condexpIndSMul] + rw [condExpIndSMul] rw [indicatorConstLp_disjoint_union hs ht hμs hμt hst (1 : ℝ)] - rw [(condexpL2 ℝ ℝ hm).map_add] + rw [(condExpL2 ℝ ℝ hm).map_add] push_cast rw [((toSpanSingleton ℝ x).compLpL 2 μ).map_add] refine (Lp.coeFn_add _ _).trans ?_ filter_upwards with y using rfl +@[deprecated (since := "2025-01-21")] +alias condexpIndL1Fin_disjoint_union := condExpIndL1Fin_disjoint_union + end CondexpIndL1Fin open scoped Classical @@ -161,171 +177,231 @@ section CondexpIndL1 /-- Conditional expectation of the indicator of a set, as a function in L1. Its value for sets which are not both measurable and of finite measure is not used: we set it to 0. -/ -def condexpIndL1 {m m0 : MeasurableSpace α} (hm : m ≤ m0) (μ : Measure α) (s : Set α) +def condExpIndL1 {m m0 : MeasurableSpace α} (hm : m ≤ m0) (μ : Measure α) (s : Set α) [SigmaFinite (μ.trim hm)] (x : G) : α →₁[μ] G := - if hs : MeasurableSet s ∧ μ s ≠ ∞ then condexpIndL1Fin hm hs.1 hs.2 x else 0 + if hs : MeasurableSet s ∧ μ s ≠ ∞ then condExpIndL1Fin hm hs.1 hs.2 x else 0 + +@[deprecated (since := "2025-01-21")] noncomputable alias condexpIndL1 := condExpIndL1 variable {hm : m ≤ m0} [SigmaFinite (μ.trim hm)] -theorem condexpIndL1_of_measurableSet_of_measure_ne_top (hs : MeasurableSet s) (hμs : μ s ≠ ∞) - (x : G) : condexpIndL1 hm μ s x = condexpIndL1Fin hm hs hμs x := by - simp only [condexpIndL1, And.intro hs hμs, dif_pos, Ne, not_false_iff, and_self_iff] +theorem condExpIndL1_of_measurableSet_of_measure_ne_top (hs : MeasurableSet s) (hμs : μ s ≠ ∞) + (x : G) : condExpIndL1 hm μ s x = condExpIndL1Fin hm hs hμs x := by + simp only [condExpIndL1, And.intro hs hμs, dif_pos, Ne, not_false_iff, and_self_iff] -theorem condexpIndL1_of_measure_eq_top (hμs : μ s = ∞) (x : G) : condexpIndL1 hm μ s x = 0 := by - simp only [condexpIndL1, hμs, eq_self_iff_true, not_true, Ne, dif_neg, not_false_iff, +@[deprecated (since := "2025-01-21")] +alias condexpIndL1_of_measurableSet_of_measure_ne_top := + condExpIndL1_of_measurableSet_of_measure_ne_top + +theorem condExpIndL1_of_measure_eq_top (hμs : μ s = ∞) (x : G) : condExpIndL1 hm μ s x = 0 := by + simp only [condExpIndL1, hμs, eq_self_iff_true, not_true, Ne, dif_neg, not_false_iff, and_false] -theorem condexpIndL1_of_not_measurableSet (hs : ¬MeasurableSet s) (x : G) : - condexpIndL1 hm μ s x = 0 := by - simp only [condexpIndL1, hs, dif_neg, not_false_iff, false_and] +@[deprecated (since := "2025-01-21")] +alias condexpIndL1_of_measure_eq_top := condExpIndL1_of_measure_eq_top + +theorem condExpIndL1_of_not_measurableSet (hs : ¬MeasurableSet s) (x : G) : + condExpIndL1 hm μ s x = 0 := by + simp only [condExpIndL1, hs, dif_neg, not_false_iff, false_and] -theorem condexpIndL1_add (x y : G) : - condexpIndL1 hm μ s (x + y) = condexpIndL1 hm μ s x + condexpIndL1 hm μ s y := by +@[deprecated (since := "2025-01-21")] +alias condexpIndL1_of_not_measurableSet := condExpIndL1_of_not_measurableSet + +theorem condExpIndL1_add (x y : G) : + condExpIndL1 hm μ s (x + y) = condExpIndL1 hm μ s x + condExpIndL1 hm μ s y := by by_cases hs : MeasurableSet s - swap; · simp_rw [condexpIndL1_of_not_measurableSet hs]; rw [zero_add] + swap; · simp_rw [condExpIndL1_of_not_measurableSet hs]; rw [zero_add] by_cases hμs : μ s = ∞ - · simp_rw [condexpIndL1_of_measure_eq_top hμs]; rw [zero_add] - · simp_rw [condexpIndL1_of_measurableSet_of_measure_ne_top hs hμs] - exact condexpIndL1Fin_add hs hμs x y + · simp_rw [condExpIndL1_of_measure_eq_top hμs]; rw [zero_add] + · simp_rw [condExpIndL1_of_measurableSet_of_measure_ne_top hs hμs] + exact condExpIndL1Fin_add hs hμs x y + +@[deprecated (since := "2025-01-21")] alias condexpIndL1_add := condExpIndL1_add -theorem condexpIndL1_smul (c : ℝ) (x : G) : - condexpIndL1 hm μ s (c • x) = c • condexpIndL1 hm μ s x := by +theorem condExpIndL1_smul (c : ℝ) (x : G) : + condExpIndL1 hm μ s (c • x) = c • condExpIndL1 hm μ s x := by by_cases hs : MeasurableSet s - swap; · simp_rw [condexpIndL1_of_not_measurableSet hs]; rw [smul_zero] + swap; · simp_rw [condExpIndL1_of_not_measurableSet hs]; rw [smul_zero] by_cases hμs : μ s = ∞ - · simp_rw [condexpIndL1_of_measure_eq_top hμs]; rw [smul_zero] - · simp_rw [condexpIndL1_of_measurableSet_of_measure_ne_top hs hμs] - exact condexpIndL1Fin_smul hs hμs c x + · simp_rw [condExpIndL1_of_measure_eq_top hμs]; rw [smul_zero] + · simp_rw [condExpIndL1_of_measurableSet_of_measure_ne_top hs hμs] + exact condExpIndL1Fin_smul hs hμs c x -theorem condexpIndL1_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (c : 𝕜) (x : F) : - condexpIndL1 hm μ s (c • x) = c • condexpIndL1 hm μ s x := by +@[deprecated (since := "2025-01-21")] alias condexpIndL1_smul := condExpIndL1_smul + +theorem condExpIndL1_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (c : 𝕜) (x : F) : + condExpIndL1 hm μ s (c • x) = c • condExpIndL1 hm μ s x := by by_cases hs : MeasurableSet s - swap; · simp_rw [condexpIndL1_of_not_measurableSet hs]; rw [smul_zero] + swap; · simp_rw [condExpIndL1_of_not_measurableSet hs]; rw [smul_zero] by_cases hμs : μ s = ∞ - · simp_rw [condexpIndL1_of_measure_eq_top hμs]; rw [smul_zero] - · simp_rw [condexpIndL1_of_measurableSet_of_measure_ne_top hs hμs] - exact condexpIndL1Fin_smul' hs hμs c x + · simp_rw [condExpIndL1_of_measure_eq_top hμs]; rw [smul_zero] + · simp_rw [condExpIndL1_of_measurableSet_of_measure_ne_top hs hμs] + exact condExpIndL1Fin_smul' hs hμs c x + +@[deprecated (since := "2025-01-21")] alias condexpIndL1_smul' := condExpIndL1_smul' -theorem norm_condexpIndL1_le (x : G) : ‖condexpIndL1 hm μ s x‖ ≤ (μ s).toReal * ‖x‖ := by +theorem norm_condExpIndL1_le (x : G) : ‖condExpIndL1 hm μ s x‖ ≤ (μ s).toReal * ‖x‖ := by by_cases hs : MeasurableSet s swap - · simp_rw [condexpIndL1_of_not_measurableSet hs]; rw [Lp.norm_zero] + · simp_rw [condExpIndL1_of_not_measurableSet hs]; rw [Lp.norm_zero] exact mul_nonneg ENNReal.toReal_nonneg (norm_nonneg _) by_cases hμs : μ s = ∞ - · rw [condexpIndL1_of_measure_eq_top hμs x, Lp.norm_zero] + · rw [condExpIndL1_of_measure_eq_top hμs x, Lp.norm_zero] exact mul_nonneg ENNReal.toReal_nonneg (norm_nonneg _) - · rw [condexpIndL1_of_measurableSet_of_measure_ne_top hs hμs x] - exact norm_condexpIndL1Fin_le hs hμs x + · rw [condExpIndL1_of_measurableSet_of_measure_ne_top hs hμs x] + exact norm_condExpIndL1Fin_le hs hμs x + +@[deprecated (since := "2025-01-21")] alias norm_condexpIndL1_le := norm_condExpIndL1_le + +theorem continuous_condExpIndL1 : Continuous fun x : G => condExpIndL1 hm μ s x := + continuous_of_linear_of_bound condExpIndL1_add condExpIndL1_smul norm_condExpIndL1_le -theorem continuous_condexpIndL1 : Continuous fun x : G => condexpIndL1 hm μ s x := - continuous_of_linear_of_bound condexpIndL1_add condexpIndL1_smul norm_condexpIndL1_le +@[deprecated (since := "2025-01-21")] alias continuous_condexpIndL1 := continuous_condExpIndL1 -theorem condexpIndL1_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) +theorem condExpIndL1_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : - condexpIndL1 hm μ (s ∪ t) x = condexpIndL1 hm μ s x + condexpIndL1 hm μ t x := by + condExpIndL1 hm μ (s ∪ t) x = condExpIndL1 hm μ s x + condExpIndL1 hm μ t x := by have hμst : μ (s ∪ t) ≠ ∞ := ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne - rw [condexpIndL1_of_measurableSet_of_measure_ne_top hs hμs x, - condexpIndL1_of_measurableSet_of_measure_ne_top ht hμt x, - condexpIndL1_of_measurableSet_of_measure_ne_top (hs.union ht) hμst x] - exact condexpIndL1Fin_disjoint_union hs ht hμs hμt hst x + rw [condExpIndL1_of_measurableSet_of_measure_ne_top hs hμs x, + condExpIndL1_of_measurableSet_of_measure_ne_top ht hμt x, + condExpIndL1_of_measurableSet_of_measure_ne_top (hs.union ht) hμst x] + exact condExpIndL1Fin_disjoint_union hs ht hμs hμt hst x + +@[deprecated (since := "2025-01-21")] +alias condexpIndL1_disjoint_union := condExpIndL1_disjoint_union end CondexpIndL1 --- Porting note: `G` is not automatically inferred in `condexpInd` in Lean 4; +-- Porting note: `G` is not automatically inferred in `condExpInd` in Lean 4; -- to avoid repeatedly typing `(G := ...)` it is made explicit. variable (G) /-- Conditional expectation of the indicator of a set, as a linear map from `G` to L1. -/ -def condexpInd {m m0 : MeasurableSpace α} (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] +def condExpInd {m m0 : MeasurableSpace α} (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] (s : Set α) : G →L[ℝ] α →₁[μ] G where - toFun := condexpIndL1 hm μ s - map_add' := condexpIndL1_add - map_smul' := condexpIndL1_smul - cont := continuous_condexpIndL1 + toFun := condExpIndL1 hm μ s + map_add' := condExpIndL1_add + map_smul' := condExpIndL1_smul + cont := continuous_condExpIndL1 + +@[deprecated (since := "2025-01-21")] noncomputable alias condexpInd := condExpInd variable {G} -theorem condexpInd_ae_eq_condexpIndSMul (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] +theorem condExpInd_ae_eq_condExpIndSMul (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : - condexpInd G hm μ s x =ᵐ[μ] condexpIndSMul hm hs hμs x := by - refine EventuallyEq.trans ?_ (condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x) - simp [condexpInd, condexpIndL1, hs, hμs] + condExpInd G hm μ s x =ᵐ[μ] condExpIndSMul hm hs hμs x := by + refine EventuallyEq.trans ?_ (condExpIndL1Fin_ae_eq_condExpIndSMul hm hs hμs x) + simp [condExpInd, condExpIndL1, hs, hμs] + +@[deprecated (since := "2025-01-21")] +alias condexpInd_ae_eq_condexpIndSMul := condExpInd_ae_eq_condExpIndSMul variable {hm : m ≤ m0} [SigmaFinite (μ.trim hm)] -theorem aestronglyMeasurable'_condexpInd (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : - AEStronglyMeasurable' m (condexpInd G hm μ s x) μ := - AEStronglyMeasurable'.congr (aeStronglyMeasurable'_condexpIndSMul hm hs hμs x) - (condexpInd_ae_eq_condexpIndSMul hm hs hμs x).symm +theorem aestronglyMeasurable_condExpInd (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : + AEStronglyMeasurable[m] (condExpInd G hm μ s x) μ := + (aestronglyMeasurable_condExpIndSMul hm hs hμs x).congr + (condExpInd_ae_eq_condExpIndSMul hm hs hμs x).symm + +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_condExpInd := aestronglyMeasurable_condExpInd + +@[deprecated (since := "2025-01-21")] +alias aestronglyMeasurable'_condexpInd := aestronglyMeasurable_condExpInd @[simp] -theorem condexpInd_empty : condexpInd G hm μ ∅ = (0 : G →L[ℝ] α →₁[μ] G) := by +theorem condExpInd_empty : condExpInd G hm μ ∅ = (0 : G →L[ℝ] α →₁[μ] G) := by ext1 x ext1 - refine (condexpInd_ae_eq_condexpIndSMul hm MeasurableSet.empty (by simp) x).trans ?_ - rw [condexpIndSMul_empty] + refine (condExpInd_ae_eq_condExpIndSMul hm MeasurableSet.empty (by simp) x).trans ?_ + rw [condExpIndSMul_empty] refine (Lp.coeFn_zero G 2 μ).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_zero G 1 μ).symm rfl -theorem condexpInd_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (c : 𝕜) (x : F) : - condexpInd F hm μ s (c • x) = c • condexpInd F hm μ s x := - condexpIndL1_smul' c x +@[deprecated (since := "2025-01-21")] alias condexpInd_empty := condExpInd_empty + +theorem condExpInd_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (c : 𝕜) (x : F) : + condExpInd F hm μ s (c • x) = c • condExpInd F hm μ s x := + condExpIndL1_smul' c x + +@[deprecated (since := "2025-01-21")] alias condexpInd_smul' := condExpInd_smul' + +theorem norm_condExpInd_apply_le (x : G) : ‖condExpInd G hm μ s x‖ ≤ (μ s).toReal * ‖x‖ := + norm_condExpIndL1_le x -theorem norm_condexpInd_apply_le (x : G) : ‖condexpInd G hm μ s x‖ ≤ (μ s).toReal * ‖x‖ := - norm_condexpIndL1_le x +@[deprecated (since := "2025-01-21")] alias norm_condexpInd_apply_le := norm_condExpInd_apply_le -theorem norm_condexpInd_le : ‖(condexpInd G hm μ s : G →L[ℝ] α →₁[μ] G)‖ ≤ (μ s).toReal := - ContinuousLinearMap.opNorm_le_bound _ ENNReal.toReal_nonneg norm_condexpInd_apply_le +theorem norm_condExpInd_le : ‖(condExpInd G hm μ s : G →L[ℝ] α →₁[μ] G)‖ ≤ (μ s).toReal := + ContinuousLinearMap.opNorm_le_bound _ ENNReal.toReal_nonneg norm_condExpInd_apply_le -theorem condexpInd_disjoint_union_apply (hs : MeasurableSet s) (ht : MeasurableSet t) +@[deprecated (since := "2025-01-21")] alias norm_condexpInd_le := norm_condExpInd_le + +theorem condExpInd_disjoint_union_apply (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (hst : Disjoint s t) (x : G) : - condexpInd G hm μ (s ∪ t) x = condexpInd G hm μ s x + condexpInd G hm μ t x := - condexpIndL1_disjoint_union hs ht hμs hμt hst x + condExpInd G hm μ (s ∪ t) x = condExpInd G hm μ s x + condExpInd G hm μ t x := + condExpIndL1_disjoint_union hs ht hμs hμt hst x + +@[deprecated (since := "2025-01-21")] +alias condexpInd_disjoint_union_apply := condExpInd_disjoint_union_apply -theorem condexpInd_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (hst : Disjoint s t) : (condexpInd G hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = - condexpInd G hm μ s + condexpInd G hm μ t := by - ext1 x; push_cast; exact condexpInd_disjoint_union_apply hs ht hμs hμt hst x +theorem condExpInd_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) + (hμt : μ t ≠ ∞) (hst : Disjoint s t) : (condExpInd G hm μ (s ∪ t) : G →L[ℝ] α →₁[μ] G) = + condExpInd G hm μ s + condExpInd G hm μ t := by + ext1 x; push_cast; exact condExpInd_disjoint_union_apply hs ht hμs hμt hst x + +@[deprecated (since := "2025-01-21")] alias condexpInd_disjoint_union := condExpInd_disjoint_union variable (G) -theorem dominatedFinMeasAdditive_condexpInd (hm : m ≤ m0) (μ : Measure α) +theorem dominatedFinMeasAdditive_condExpInd (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] : - DominatedFinMeasAdditive μ (condexpInd G hm μ : Set α → G →L[ℝ] α →₁[μ] G) 1 := - ⟨fun _ _ => condexpInd_disjoint_union, fun _ _ _ => norm_condexpInd_le.trans (one_mul _).symm.le⟩ + DominatedFinMeasAdditive μ (condExpInd G hm μ : Set α → G →L[ℝ] α →₁[μ] G) 1 := + ⟨fun _ _ => condExpInd_disjoint_union, fun _ _ _ => norm_condExpInd_le.trans (one_mul _).symm.le⟩ + +@[deprecated (since := "2025-01-21")] +alias dominatedFinMeasAdditive_condexpInd := dominatedFinMeasAdditive_condExpInd variable {G} -theorem setIntegral_condexpInd (hs : MeasurableSet[m] s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) - (hμt : μ t ≠ ∞) (x : G') : ∫ a in s, condexpInd G' hm μ t x a ∂μ = (μ (t ∩ s)).toReal • x := +theorem setIntegral_condExpInd (hs : MeasurableSet[m] s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) + (hμt : μ t ≠ ∞) (x : G') : ∫ a in s, condExpInd G' hm μ t x a ∂μ = (μ (t ∩ s)).toReal • x := calc - ∫ a in s, condexpInd G' hm μ t x a ∂μ = ∫ a in s, condexpIndSMul hm ht hμt x a ∂μ := + ∫ a in s, condExpInd G' hm μ t x a ∂μ = ∫ a in s, condExpIndSMul hm ht hμt x a ∂μ := setIntegral_congr_ae (hm s hs) - ((condexpInd_ae_eq_condexpIndSMul hm ht hμt x).mono fun _ hx _ => hx) - _ = (μ (t ∩ s)).toReal • x := setIntegral_condexpIndSMul hs ht hμs hμt x + ((condExpInd_ae_eq_condExpIndSMul hm ht hμt x).mono fun _ hx _ => hx) + _ = (μ (t ∩ s)).toReal • x := setIntegral_condExpIndSMul hs ht hμs hμt x + +@[deprecated (since := "2025-01-21")] alias setIntegral_condexpInd := setIntegral_condExpInd @[deprecated (since := "2024-04-17")] -alias set_integral_condexpInd := setIntegral_condexpInd +alias set_integral_condExpInd := setIntegral_condExpInd + +@[deprecated (since := "2025-01-21")] alias set_integral_condexpInd := set_integral_condExpInd -theorem condexpInd_of_measurable (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) (c : G) : - condexpInd G hm μ s c = indicatorConstLp 1 (hm s hs) hμs c := by +theorem condExpInd_of_measurable (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) (c : G) : + condExpInd G hm μ s c = indicatorConstLp 1 (hm s hs) hμs c := by ext1 refine EventuallyEq.trans ?_ indicatorConstLp_coeFn.symm - refine (condexpInd_ae_eq_condexpIndSMul hm (hm s hs) hμs c).trans ?_ - refine (condexpIndSMul_ae_eq_smul hm (hm s hs) hμs c).trans ?_ - rw [lpMeas_coe, condexpL2_indicator_of_measurable hm hs hμs (1 : ℝ)] + refine (condExpInd_ae_eq_condExpIndSMul hm (hm s hs) hμs c).trans ?_ + refine (condExpIndSMul_ae_eq_smul hm (hm s hs) hμs c).trans ?_ + rw [lpMeas_coe, condExpL2_indicator_of_measurable hm hs hμs (1 : ℝ)] refine (@indicatorConstLp_coeFn α _ _ 2 μ _ s (hm s hs) hμs (1 : ℝ)).mono fun x hx => ?_ dsimp only rw [hx] by_cases hx_mem : x ∈ s <;> simp [hx_mem] -theorem condexpInd_nonneg {E} [NormedLatticeAddCommGroup E] [NormedSpace ℝ E] [OrderedSMul ℝ E] - (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : 0 ≤ condexpInd E hm μ s x := by +@[deprecated (since := "2025-01-21")] alias condexpInd_of_measurable := condExpInd_of_measurable + +theorem condExpInd_nonneg {E} [NormedLatticeAddCommGroup E] [NormedSpace ℝ E] [OrderedSMul ℝ E] + (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : 0 ≤ condExpInd E hm μ s x := by rw [← coeFn_le] - refine EventuallyLE.trans_eq ?_ (condexpInd_ae_eq_condexpIndSMul hm hs hμs x).symm - exact (coeFn_zero E 1 μ).trans_le (condexpIndSMul_nonneg hs hμs x hx) + refine EventuallyLE.trans_eq ?_ (condExpInd_ae_eq_condExpIndSMul hm hs hμs x).symm + exact (coeFn_zero E 1 μ).trans_le (condExpIndSMul_nonneg hs hμs x hx) + +@[deprecated (since := "2025-01-21")] alias condexpInd_nonneg := condExpInd_nonneg end CondexpInd @@ -335,62 +411,78 @@ section CondexpL1 variable {m m0 : MeasurableSpace α} {μ : Measure α} {hm : m ≤ m0} [SigmaFinite (μ.trim hm)] {f g : α → F'} {s : Set α} --- Porting note: `F'` is not automatically inferred in `condexpL1CLM` in Lean 4; +-- Porting note: `F'` is not automatically inferred in `condExpL1CLM` in Lean 4; -- to avoid repeatedly typing `(F' := ...)` it is made explicit. variable (F') /-- Conditional expectation of a function as a linear map from `α →₁[μ] F'` to itself. -/ -def condexpL1CLM (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] : +def condExpL1CLM (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] : (α →₁[μ] F') →L[ℝ] α →₁[μ] F' := - L1.setToL1 (dominatedFinMeasAdditive_condexpInd F' hm μ) + L1.setToL1 (dominatedFinMeasAdditive_condExpInd F' hm μ) + +@[deprecated (since := "2025-01-21")] noncomputable alias condexpL1CLM := condExpL1CLM variable {F'} -theorem condexpL1CLM_smul (c : 𝕜) (f : α →₁[μ] F') : - condexpL1CLM F' hm μ (c • f) = c • condexpL1CLM F' hm μ f := by - refine L1.setToL1_smul (dominatedFinMeasAdditive_condexpInd F' hm μ) ?_ c f - exact fun c s x => condexpInd_smul' c x +theorem condExpL1CLM_smul (c : 𝕜) (f : α →₁[μ] F') : + condExpL1CLM F' hm μ (c • f) = c • condExpL1CLM F' hm μ f := by + refine L1.setToL1_smul (dominatedFinMeasAdditive_condExpInd F' hm μ) ?_ c f + exact fun c s x => condExpInd_smul' c x + +@[deprecated (since := "2025-01-21")] alias condexpL1CLM_smul := condExpL1CLM_smul + +theorem condExpL1CLM_indicatorConstLp (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : F') : + (condExpL1CLM F' hm μ) (indicatorConstLp 1 hs hμs x) = condExpInd F' hm μ s x := + L1.setToL1_indicatorConstLp (dominatedFinMeasAdditive_condExpInd F' hm μ) hs hμs x + +@[deprecated (since := "2025-01-21")] +alias condexpL1CLM_indicatorConstLp := condExpL1CLM_indicatorConstLp -theorem condexpL1CLM_indicatorConstLp (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : F') : - (condexpL1CLM F' hm μ) (indicatorConstLp 1 hs hμs x) = condexpInd F' hm μ s x := - L1.setToL1_indicatorConstLp (dominatedFinMeasAdditive_condexpInd F' hm μ) hs hμs x +theorem condExpL1CLM_indicatorConst (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : F') : + (condExpL1CLM F' hm μ) ↑(simpleFunc.indicatorConst 1 hs hμs x) = condExpInd F' hm μ s x := by + rw [Lp.simpleFunc.coe_indicatorConst]; exact condExpL1CLM_indicatorConstLp hs hμs x -theorem condexpL1CLM_indicatorConst (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : F') : - (condexpL1CLM F' hm μ) ↑(simpleFunc.indicatorConst 1 hs hμs x) = condexpInd F' hm μ s x := by - rw [Lp.simpleFunc.coe_indicatorConst]; exact condexpL1CLM_indicatorConstLp hs hμs x +@[deprecated (since := "2025-01-21")] +alias condexpL1CLM_indicatorConst := condExpL1CLM_indicatorConst -/-- Auxiliary lemma used in the proof of `setIntegral_condexpL1CLM`. -/ -theorem setIntegral_condexpL1CLM_of_measure_ne_top (f : α →₁[μ] F') (hs : MeasurableSet[m] s) - (hμs : μ s ≠ ∞) : ∫ x in s, condexpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ := by +/-- Auxiliary lemma used in the proof of `setIntegral_condExpL1CLM`. -/ +theorem setIntegral_condExpL1CLM_of_measure_ne_top (f : α →₁[μ] F') (hs : MeasurableSet[m] s) + (hμs : μ s ≠ ∞) : ∫ x in s, condExpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ := by refine @Lp.induction _ _ _ _ _ _ _ ENNReal.one_ne_top - (fun f : α →₁[μ] F' => ∫ x in s, condexpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ) ?_ ?_ + (fun f : α →₁[μ] F' => ∫ x in s, condExpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ) ?_ ?_ (isClosed_eq ?_ ?_) f · intro x t ht hμt - simp_rw [condexpL1CLM_indicatorConst ht hμt.ne x] + simp_rw [condExpL1CLM_indicatorConst ht hμt.ne x] rw [Lp.simpleFunc.coe_indicatorConst, setIntegral_indicatorConstLp (hm _ hs)] - exact setIntegral_condexpInd hs ht hμs hμt.ne x + exact setIntegral_condExpInd hs ht hμs hμt.ne x · intro f g hf_Lp hg_Lp _ hf hg - simp_rw [(condexpL1CLM F' hm μ).map_add] - rw [setIntegral_congr_ae (hm s hs) ((Lp.coeFn_add (condexpL1CLM F' hm μ (hf_Lp.toLp f)) - (condexpL1CLM F' hm μ (hg_Lp.toLp g))).mono fun x hx _ => hx)] + simp_rw [(condExpL1CLM F' hm μ).map_add] + rw [setIntegral_congr_ae (hm s hs) ((Lp.coeFn_add (condExpL1CLM F' hm μ (hf_Lp.toLp f)) + (condExpL1CLM F' hm μ (hg_Lp.toLp g))).mono fun x hx _ => hx)] rw [setIntegral_congr_ae (hm s hs) ((Lp.coeFn_add (hf_Lp.toLp f) (hg_Lp.toLp g)).mono fun x hx _ => hx)] simp_rw [Pi.add_apply] rw [integral_add (L1.integrable_coeFn _).integrableOn (L1.integrable_coeFn _).integrableOn, integral_add (L1.integrable_coeFn _).integrableOn (L1.integrable_coeFn _).integrableOn, hf, hg] - · exact (continuous_setIntegral s).comp (condexpL1CLM F' hm μ).continuous + · exact (continuous_setIntegral s).comp (condExpL1CLM F' hm μ).continuous · exact continuous_setIntegral s +@[deprecated (since := "2025-01-21")] +alias setIntegral_condexpL1CLM_of_measure_ne_top := setIntegral_condExpL1CLM_of_measure_ne_top + @[deprecated (since := "2024-04-17")] alias set_integral_condexpL1CLM_of_measure_ne_top := - setIntegral_condexpL1CLM_of_measure_ne_top + setIntegral_condExpL1CLM_of_measure_ne_top + +@[deprecated (since := "2025-01-21")] +alias setIntegral_condexpL1CLM := set_integral_condexpL1CLM_of_measure_ne_top -/-- The integral of the conditional expectation `condexpL1CLM` over an `m`-measurable set is equal -to the integral of `f` on that set. See also `setIntegral_condexp`, the similar statement for -`condexp`. -/ -theorem setIntegral_condexpL1CLM (f : α →₁[μ] F') (hs : MeasurableSet[m] s) : - ∫ x in s, condexpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ := by +/-- The integral of the conditional expectation `condExpL1CLM` over an `m`-measurable set is equal +to the integral of `f` on that set. See also `setIntegral_condExp`, the similar statement for +`condExp`. -/ +theorem setIntegral_condExpL1CLM (f : α →₁[μ] F') (hs : MeasurableSet[m] s) : + ∫ x in s, condExpL1CLM F' hm μ f x ∂μ = ∫ x in s, f x ∂μ := by let S := spanningSets (μ.trim hm) have hS_meas : ∀ i, MeasurableSet[m] (S i) := measurableSet_spanningSets (μ.trim hm) have hS_meas0 : ∀ i, MeasurableSet (S i) := fun i => hm _ (hS_meas i) @@ -406,146 +498,200 @@ theorem setIntegral_condexpL1CLM (f : α →₁[μ] F') (hs : MeasurableSet[m] s simp_rw [Set.mem_inter_iff] exact fun h => ⟨monotone_spanningSets (μ.trim hm) hij h.1, h.2⟩ have h_eq_forall : - (fun i => ∫ x in S i ∩ s, condexpL1CLM F' hm μ f x ∂μ) = fun i => ∫ x in S i ∩ s, f x ∂μ := + (fun i => ∫ x in S i ∩ s, condExpL1CLM F' hm μ f x ∂μ) = fun i => ∫ x in S i ∩ s, f x ∂μ := funext fun i => - setIntegral_condexpL1CLM_of_measure_ne_top f (@MeasurableSet.inter α m _ _ (hS_meas i) hs) + setIntegral_condExpL1CLM_of_measure_ne_top f (@MeasurableSet.inter α m _ _ (hS_meas i) hs) (hS_finite i).ne have h_right : Tendsto (fun i => ∫ x in S i ∩ s, f x ∂μ) atTop (𝓝 (∫ x in s, f x ∂μ)) := by have h := tendsto_setIntegral_of_monotone (fun i => (hS_meas0 i).inter (hm s hs)) h_mono (L1.integrable_coeFn f).integrableOn rwa [← hs_eq] at h - have h_left : Tendsto (fun i => ∫ x in S i ∩ s, condexpL1CLM F' hm μ f x ∂μ) atTop - (𝓝 (∫ x in s, condexpL1CLM F' hm μ f x ∂μ)) := by + have h_left : Tendsto (fun i => ∫ x in S i ∩ s, condExpL1CLM F' hm μ f x ∂μ) atTop + (𝓝 (∫ x in s, condExpL1CLM F' hm μ f x ∂μ)) := by have h := tendsto_setIntegral_of_monotone (fun i => (hS_meas0 i).inter (hm s hs)) h_mono - (L1.integrable_coeFn (condexpL1CLM F' hm μ f)).integrableOn + (L1.integrable_coeFn (condExpL1CLM F' hm μ f)).integrableOn rwa [← hs_eq] at h rw [h_eq_forall] at h_left exact tendsto_nhds_unique h_left h_right @[deprecated (since := "2024-04-17")] -alias set_integral_condexpL1CLM := setIntegral_condexpL1CLM +alias set_integral_condExpL1CLM := setIntegral_condExpL1CLM -theorem aestronglyMeasurable'_condexpL1CLM (f : α →₁[μ] F') : - AEStronglyMeasurable' m (condexpL1CLM F' hm μ f) μ := by +@[deprecated (since := "2025-01-21")] alias set_integral_condexpL1CLM := set_integral_condExpL1CLM + +theorem aestronglyMeasurable_condExpL1CLM (f : α →₁[μ] F') : + AEStronglyMeasurable[m] (condExpL1CLM F' hm μ f) μ := by refine @Lp.induction _ _ _ _ _ _ _ ENNReal.one_ne_top - (fun f : α →₁[μ] F' => AEStronglyMeasurable' m (condexpL1CLM F' hm μ f) μ) ?_ ?_ ?_ f + (fun f : α →₁[μ] F' => AEStronglyMeasurable[m] (condExpL1CLM F' hm μ f) μ) ?_ ?_ ?_ f · intro c s hs hμs - rw [condexpL1CLM_indicatorConst hs hμs.ne c] - exact aestronglyMeasurable'_condexpInd hs hμs.ne c + rw [condExpL1CLM_indicatorConst hs hμs.ne c] + exact aestronglyMeasurable_condExpInd hs hμs.ne c · intro f g hf hg _ hfm hgm - rw [(condexpL1CLM F' hm μ).map_add] - refine AEStronglyMeasurable'.congr ?_ (coeFn_add _ _).symm - exact AEStronglyMeasurable'.add hfm hgm - · have : {f : Lp F' 1 μ | AEStronglyMeasurable' m (condexpL1CLM F' hm μ f) μ} = - condexpL1CLM F' hm μ ⁻¹' {f | AEStronglyMeasurable' m f μ} := rfl + rw [(condExpL1CLM F' hm μ).map_add] + exact (hfm.add hgm).congr (coeFn_add ..).symm + · have : {f : Lp F' 1 μ | AEStronglyMeasurable[m] (condExpL1CLM F' hm μ f) μ} = + condExpL1CLM F' hm μ ⁻¹' {f | AEStronglyMeasurable[m] f μ} := rfl rw [this] - refine IsClosed.preimage (condexpL1CLM F' hm μ).continuous ?_ + refine IsClosed.preimage (condExpL1CLM F' hm μ).continuous ?_ exact isClosed_aeStronglyMeasurable' hm -theorem condexpL1CLM_lpMeas (f : lpMeas F' ℝ m 1 μ) : - condexpL1CLM F' hm μ (f : α →₁[μ] F') = ↑f := by +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_condExpL1CLM := aestronglyMeasurable_condExpL1CLM + +@[deprecated (since := "2025-01-21")] +alias aestronglyMeasurable_condexpL1CLM := aestronglyMeasurable_condExpL1CLM + +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_condexpL1CLM := aestronglyMeasurable_condexpL1CLM + +theorem condExpL1CLM_lpMeas (f : lpMeas F' ℝ m 1 μ) : + condExpL1CLM F' hm μ (f : α →₁[μ] F') = ↑f := by let g := lpMeasToLpTrimLie F' ℝ 1 μ hm f have hfg : f = (lpMeasToLpTrimLie F' ℝ 1 μ hm).symm g := by simp only [g, LinearIsometryEquiv.symm_apply_apply] rw [hfg] refine @Lp.induction α F' m _ 1 (μ.trim hm) _ ENNReal.coe_ne_top (fun g : α →₁[μ.trim hm] F' => - condexpL1CLM F' hm μ ((lpMeasToLpTrimLie F' ℝ 1 μ hm).symm g : α →₁[μ] F') = + condExpL1CLM F' hm μ ((lpMeasToLpTrimLie F' ℝ 1 μ hm).symm g : α →₁[μ] F') = ↑((lpMeasToLpTrimLie F' ℝ 1 μ hm).symm g)) ?_ ?_ ?_ g · intro c s hs hμs rw [@Lp.simpleFunc.coe_indicatorConst _ _ m, lpMeasToLpTrimLie_symm_indicator hs hμs.ne c, - condexpL1CLM_indicatorConstLp] - exact condexpInd_of_measurable hs ((le_trim hm).trans_lt hμs).ne c + condExpL1CLM_indicatorConstLp] + exact condExpInd_of_measurable hs ((le_trim hm).trans_lt hμs).ne c · intro f g hf hg _ hf_eq hg_eq rw [LinearIsometryEquiv.map_add] push_cast rw [map_add, hf_eq, hg_eq] · refine isClosed_eq ?_ ?_ - · refine (condexpL1CLM F' hm μ).continuous.comp (continuous_induced_dom.comp ?_) + · refine (condExpL1CLM F' hm μ).continuous.comp (continuous_induced_dom.comp ?_) exact LinearIsometryEquiv.continuous _ · refine continuous_induced_dom.comp ?_ exact LinearIsometryEquiv.continuous _ -theorem condexpL1CLM_of_aestronglyMeasurable' (f : α →₁[μ] F') (hfm : AEStronglyMeasurable' m f μ) : - condexpL1CLM F' hm μ f = f := - condexpL1CLM_lpMeas (⟨f, hfm⟩ : lpMeas F' ℝ m 1 μ) +@[deprecated (since := "2025-01-21")] alias condexpL1CLM_lpMeas := condExpL1CLM_lpMeas + +theorem condExpL1CLM_of_aestronglyMeasurable' (f : α →₁[μ] F') (hfm : AEStronglyMeasurable[m] f μ) : + condExpL1CLM F' hm μ f = f := + condExpL1CLM_lpMeas (⟨f, hfm⟩ : lpMeas F' ℝ m 1 μ) + +@[deprecated (since := "2025-01-21")] +alias condexpL1CLM_of_aestronglyMeasurable' := condExpL1CLM_of_aestronglyMeasurable' /-- Conditional expectation of a function, in L1. Its value is 0 if the function is not -integrable. The function-valued `condexp` should be used instead in most cases. -/ -def condexpL1 (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] (f : α → F') : α →₁[μ] F' := - setToFun μ (condexpInd F' hm μ) (dominatedFinMeasAdditive_condexpInd F' hm μ) f +integrable. The function-valued `condExp` should be used instead in most cases. -/ +def condExpL1 (hm : m ≤ m0) (μ : Measure α) [SigmaFinite (μ.trim hm)] (f : α → F') : α →₁[μ] F' := + setToFun μ (condExpInd F' hm μ) (dominatedFinMeasAdditive_condExpInd F' hm μ) f -theorem condexpL1_undef (hf : ¬Integrable f μ) : condexpL1 hm μ f = 0 := - setToFun_undef (dominatedFinMeasAdditive_condexpInd F' hm μ) hf +@[deprecated (since := "2025-01-21")] noncomputable alias condexpL1 := condExpL1 -theorem condexpL1_eq (hf : Integrable f μ) : condexpL1 hm μ f = condexpL1CLM F' hm μ (hf.toL1 f) := - setToFun_eq (dominatedFinMeasAdditive_condexpInd F' hm μ) hf +theorem condExpL1_undef (hf : ¬Integrable f μ) : condExpL1 hm μ f = 0 := + setToFun_undef (dominatedFinMeasAdditive_condExpInd F' hm μ) hf + +@[deprecated (since := "2025-01-21")] alias condexpL1_undef := condExpL1_undef + +theorem condExpL1_eq (hf : Integrable f μ) : condExpL1 hm μ f = condExpL1CLM F' hm μ (hf.toL1 f) := + setToFun_eq (dominatedFinMeasAdditive_condExpInd F' hm μ) hf + +@[deprecated (since := "2025-01-21")] alias condexpL1_eq := condExpL1_eq @[simp] -theorem condexpL1_zero : condexpL1 hm μ (0 : α → F') = 0 := +theorem condExpL1_zero : condExpL1 hm μ (0 : α → F') = 0 := setToFun_zero _ +@[deprecated (since := "2025-01-21")] alias condexpL1_zero := condExpL1_zero + @[simp] -theorem condexpL1_measure_zero (hm : m ≤ m0) : condexpL1 hm (0 : Measure α) f = 0 := +theorem condExpL1_measure_zero (hm : m ≤ m0) : condExpL1 hm (0 : Measure α) f = 0 := setToFun_measure_zero _ rfl -theorem aestronglyMeasurable'_condexpL1 {f : α → F'} : - AEStronglyMeasurable' m (condexpL1 hm μ f) μ := by +@[deprecated (since := "2025-01-21")] alias condexpL1_measure_zero := condExpL1_measure_zero + +theorem aestronglyMeasurable_condExpL1 {f : α → F'} : + AEStronglyMeasurable[m] (condExpL1 hm μ f) μ := by by_cases hf : Integrable f μ - · rw [condexpL1_eq hf] - exact aestronglyMeasurable'_condexpL1CLM _ - · rw [condexpL1_undef hf] - refine AEStronglyMeasurable'.congr ?_ (coeFn_zero _ _ _).symm - exact StronglyMeasurable.aeStronglyMeasurable' (@stronglyMeasurable_zero _ _ m _ _) - -theorem condexpL1_congr_ae (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (h : f =ᵐ[μ] g) : - condexpL1 hm μ f = condexpL1 hm μ g := + · rw [condExpL1_eq hf] + exact aestronglyMeasurable_condExpL1CLM _ + · rw [condExpL1_undef hf] + exact stronglyMeasurable_zero.aestronglyMeasurable.congr (coeFn_zero ..).symm + +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_condExpL1 := aestronglyMeasurable_condExpL1 + +@[deprecated (since := "2025-01-21")] +alias aestronglyMeasurable_condexpL1 := aestronglyMeasurable_condExpL1 + +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_condexpL1 := aestronglyMeasurable_condexpL1 + +theorem condExpL1_congr_ae (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (h : f =ᵐ[μ] g) : + condExpL1 hm μ f = condExpL1 hm μ g := setToFun_congr_ae _ h -theorem integrable_condexpL1 (f : α → F') : Integrable (condexpL1 hm μ f) μ := +@[deprecated (since := "2025-01-21")] alias condexpL1_congr_ae := condExpL1_congr_ae + +theorem integrable_condExpL1 (f : α → F') : Integrable (condExpL1 hm μ f) μ := L1.integrable_coeFn _ -/-- The integral of the conditional expectation `condexpL1` over an `m`-measurable set is equal to -the integral of `f` on that set. See also `setIntegral_condexp`, the similar statement for -`condexp`. -/ -theorem setIntegral_condexpL1 (hf : Integrable f μ) (hs : MeasurableSet[m] s) : - ∫ x in s, condexpL1 hm μ f x ∂μ = ∫ x in s, f x ∂μ := by - simp_rw [condexpL1_eq hf] - rw [setIntegral_condexpL1CLM (hf.toL1 f) hs] +@[deprecated (since := "2025-01-21")] alias integrable_condexpL1 := integrable_condExpL1 + +/-- The integral of the conditional expectation `condExpL1` over an `m`-measurable set is equal to +the integral of `f` on that set. See also `setIntegral_condExp`, the similar statement for +`condExp`. -/ +theorem setIntegral_condExpL1 (hf : Integrable f μ) (hs : MeasurableSet[m] s) : + ∫ x in s, condExpL1 hm μ f x ∂μ = ∫ x in s, f x ∂μ := by + simp_rw [condExpL1_eq hf] + rw [setIntegral_condExpL1CLM (hf.toL1 f) hs] exact setIntegral_congr_ae (hm s hs) (hf.coeFn_toL1.mono fun x hx _ => hx) +@[deprecated (since := "2025-01-21")] alias setIntegral_condexpL1 := setIntegral_condExpL1 + @[deprecated (since := "2024-04-17")] -alias set_integral_condexpL1 := setIntegral_condexpL1 +alias set_integral_condExpL1 := setIntegral_condExpL1 -theorem condexpL1_add (hf : Integrable f μ) (hg : Integrable g μ) : - condexpL1 hm μ (f + g) = condexpL1 hm μ f + condexpL1 hm μ g := +@[deprecated (since := "2025-01-21")] alias set_integral_condexpL1 := set_integral_condExpL1 + +theorem condExpL1_add (hf : Integrable f μ) (hg : Integrable g μ) : + condExpL1 hm μ (f + g) = condExpL1 hm μ f + condExpL1 hm μ g := setToFun_add _ hf hg -theorem condexpL1_neg (f : α → F') : condexpL1 hm μ (-f) = -condexpL1 hm μ f := +@[deprecated (since := "2025-01-21")] alias condexpL1_add := condExpL1_add + +theorem condExpL1_neg (f : α → F') : condExpL1 hm μ (-f) = -condExpL1 hm μ f := setToFun_neg _ f -theorem condexpL1_smul (c : 𝕜) (f : α → F') : condexpL1 hm μ (c • f) = c • condexpL1 hm μ f := by +@[deprecated (since := "2025-01-21")] alias condexpL1_neg := condExpL1_neg + +theorem condExpL1_smul (c : 𝕜) (f : α → F') : condExpL1 hm μ (c • f) = c • condExpL1 hm μ f := by refine setToFun_smul _ ?_ c f - exact fun c _ x => condexpInd_smul' c x + exact fun c _ x => condExpInd_smul' c x + +@[deprecated (since := "2025-01-21")] alias condexpL1_smul := condExpL1_smul -theorem condexpL1_sub (hf : Integrable f μ) (hg : Integrable g μ) : - condexpL1 hm μ (f - g) = condexpL1 hm μ f - condexpL1 hm μ g := +theorem condExpL1_sub (hf : Integrable f μ) (hg : Integrable g μ) : + condExpL1 hm μ (f - g) = condExpL1 hm μ f - condExpL1 hm μ g := setToFun_sub _ hf hg -theorem condexpL1_of_aestronglyMeasurable' (hfm : AEStronglyMeasurable' m f μ) - (hfi : Integrable f μ) : condexpL1 hm μ f =ᵐ[μ] f := by - rw [condexpL1_eq hfi] +@[deprecated (since := "2025-01-21")] alias condexpL1_sub := condExpL1_sub + +theorem condExpL1_of_aestronglyMeasurable' (hfm : AEStronglyMeasurable[m] f μ) + (hfi : Integrable f μ) : condExpL1 hm μ f =ᵐ[μ] f := by + rw [condExpL1_eq hfi] refine EventuallyEq.trans ?_ (Integrable.coeFn_toL1 hfi) - rw [condexpL1CLM_of_aestronglyMeasurable'] - exact AEStronglyMeasurable'.congr hfm (Integrable.coeFn_toL1 hfi).symm + rw [condExpL1CLM_of_aestronglyMeasurable'] + exact hfm.congr hfi.coeFn_toL1.symm + +@[deprecated (since := "2025-01-21")] +alias condexpL1_of_aestronglyMeasurable' := condExpL1_of_aestronglyMeasurable' -theorem condexpL1_mono {E} [NormedLatticeAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] +theorem condExpL1_mono {E} [NormedLatticeAddCommGroup E] [CompleteSpace E] [NormedSpace ℝ E] [OrderedSMul ℝ E] {f g : α → E} (hf : Integrable f μ) (hg : Integrable g μ) (hfg : f ≤ᵐ[μ] g) : - condexpL1 hm μ f ≤ᵐ[μ] condexpL1 hm μ g := by + condExpL1 hm μ f ≤ᵐ[μ] condExpL1 hm μ g := by rw [coeFn_le] - have h_nonneg : ∀ s, MeasurableSet s → μ s < ∞ → ∀ x : E, 0 ≤ x → 0 ≤ condexpInd E hm μ s x := - fun s hs hμs x hx => condexpInd_nonneg hs hμs.ne x hx - exact setToFun_mono (dominatedFinMeasAdditive_condexpInd E hm μ) h_nonneg hf hg hfg + have h_nonneg : ∀ s, MeasurableSet s → μ s < ∞ → ∀ x : E, 0 ≤ x → 0 ≤ condExpInd E hm μ s x := + fun s hs hμs x hx => condExpInd_nonneg hs hμs.ne x hx + exact setToFun_mono (dominatedFinMeasAdditive_condExpInd E hm μ) h_nonneg hf hg hfg + +@[deprecated (since := "2025-01-21")] alias condexpL1_mono := condExpL1_mono end CondexpL1 diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index 77abe61e2911f..b6524435ba77e 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -18,14 +18,14 @@ orthogonal projection on the subspace of almost everywhere `m`-measurable functi ## Main definitions -* `condexpL2`: Conditional expectation of a function in L2 with respect to a sigma-algebra: it is +* `condExpL2`: Conditional expectation of a function in L2 with respect to a sigma-algebra: it is the orthogonal projection on the subspace `lpMeas`. ## Implementation notes Most of the results in this file are valid for a complete real normed space `F`. However, some lemmas also use `𝕜 : RCLike`: -* `condexpL2` is defined only for an `InnerProductSpace` for now, and we use `𝕜` for its field. +* `condExpL2` is defined only for an `InnerProductSpace` for now, and we use `𝕜` for its field. * results about scalar multiplication are stated not only for `ℝ` but also for `𝕜` if we happen to have `NormedSpace 𝕜 F`. @@ -61,63 +61,90 @@ local notation "⟪" x ", " y "⟫" => @inner 𝕜 E _ x y local notation "⟪" x ", " y "⟫₂" => @inner 𝕜 (α →₂[μ] E) _ x y --- Porting note: the argument `E` of `condexpL2` is not automatically filled in Lean 4. +-- Porting note: the argument `E` of `condExpL2` is not automatically filled in Lean 4. -- To avoid typing `(E := _)` every time it is made explicit. variable (E 𝕜) /-- Conditional expectation of a function in L2 with respect to a sigma-algebra -/ -noncomputable def condexpL2 (hm : m ≤ m0) : (α →₂[μ] E) →L[𝕜] lpMeas E 𝕜 m 2 μ := +noncomputable def condExpL2 (hm : m ≤ m0) : (α →₂[μ] E) →L[𝕜] lpMeas E 𝕜 m 2 μ := @orthogonalProjection 𝕜 (α →₂[μ] E) _ _ _ (lpMeas E 𝕜 m 2 μ) haveI : Fact (m ≤ m0) := ⟨hm⟩ inferInstance +@[deprecated (since := "2025-01-21")] alias condexpL2 := condExpL2 + variable {E 𝕜} -theorem aeStronglyMeasurable'_condexpL2 (hm : m ≤ m0) (f : α →₂[μ] E) : - AEStronglyMeasurable' (β := E) m (condexpL2 E 𝕜 hm f) μ := - lpMeas.aeStronglyMeasurable' _ +theorem aestronglyMeasurable_condExpL2 (hm : m ≤ m0) (f : α →₂[μ] E) : + AEStronglyMeasurable[m] (condExpL2 E 𝕜 hm f : α → E) μ := + lpMeas.aeStronglyMeasurable _ + +@[deprecated (since := "2025-01-24")] +alias aeStronglyMeasurable'_condExpL2 := aestronglyMeasurable_condExpL2 + +@[deprecated (since := "2025-01-24")] +alias aeStronglyMeasurable'_condexpL2 := aestronglyMeasurable_condExpL2 -theorem integrableOn_condexpL2_of_measure_ne_top (hm : m ≤ m0) (hμs : μ s ≠ ∞) (f : α →₂[μ] E) : - IntegrableOn (ε := E) (condexpL2 E 𝕜 hm f) s μ := - integrableOn_Lp_of_measure_ne_top (condexpL2 E 𝕜 hm f : α →₂[μ] E) fact_one_le_two_ennreal.elim +theorem integrableOn_condExpL2_of_measure_ne_top (hm : m ≤ m0) (hμs : μ s ≠ ∞) (f : α →₂[μ] E) : + IntegrableOn (ε := E) (condExpL2 E 𝕜 hm f) s μ := + integrableOn_Lp_of_measure_ne_top (condExpL2 E 𝕜 hm f : α →₂[μ] E) fact_one_le_two_ennreal.elim hμs -theorem integrable_condexpL2_of_isFiniteMeasure (hm : m ≤ m0) [IsFiniteMeasure μ] {f : α →₂[μ] E} : - Integrable (ε := E) (condexpL2 E 𝕜 hm f) μ := - integrableOn_univ.mp <| integrableOn_condexpL2_of_measure_ne_top hm (measure_ne_top _ _) f +@[deprecated (since := "2025-01-21")] +alias integrableOn_condexpL2_of_measure_ne_top := integrableOn_condExpL2_of_measure_ne_top -theorem norm_condexpL2_le_one (hm : m ≤ m0) : ‖@condexpL2 α E 𝕜 _ _ _ _ _ _ μ hm‖ ≤ 1 := +theorem integrable_condExpL2_of_isFiniteMeasure (hm : m ≤ m0) [IsFiniteMeasure μ] {f : α →₂[μ] E} : + Integrable (ε := E) (condExpL2 E 𝕜 hm f) μ := + integrableOn_univ.mp <| integrableOn_condExpL2_of_measure_ne_top hm (measure_ne_top _ _) f + +@[deprecated (since := "2025-01-21")] +alias integrable_condexpL2_of_isFiniteMeasure := integrable_condExpL2_of_isFiniteMeasure + +theorem norm_condExpL2_le_one (hm : m ≤ m0) : ‖@condExpL2 α E 𝕜 _ _ _ _ _ _ μ hm‖ ≤ 1 := haveI : Fact (m ≤ m0) := ⟨hm⟩ orthogonalProjection_norm_le _ -theorem norm_condexpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condexpL2 E 𝕜 hm f‖ ≤ ‖f‖ := - ((@condexpL2 _ E 𝕜 _ _ _ _ _ _ μ hm).le_opNorm f).trans - (mul_le_of_le_one_left (norm_nonneg _) (norm_condexpL2_le_one hm)) +@[deprecated (since := "2025-01-21")] alias norm_condexpL2_le_one := norm_condExpL2_le_one + +theorem norm_condExpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖condExpL2 E 𝕜 hm f‖ ≤ ‖f‖ := + ((@condExpL2 _ E 𝕜 _ _ _ _ _ _ μ hm).le_opNorm f).trans + (mul_le_of_le_one_left (norm_nonneg _) (norm_condExpL2_le_one hm)) -theorem eLpNorm_condexpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) : - eLpNorm (ε := E) (condexpL2 E 𝕜 hm f) 2 μ ≤ eLpNorm f 2 μ := by +@[deprecated (since := "2025-01-21")] alias norm_condexpL2_le := norm_condExpL2_le + +theorem eLpNorm_condExpL2_le (hm : m ≤ m0) (f : α →₂[μ] E) : + eLpNorm (ε := E) (condExpL2 E 𝕜 hm f) 2 μ ≤ eLpNorm f 2 μ := by rw [lpMeas_coe, ← ENNReal.toReal_le_toReal (Lp.eLpNorm_ne_top _) (Lp.eLpNorm_ne_top _), ← Lp.norm_def, ← Lp.norm_def, Submodule.norm_coe] - exact norm_condexpL2_le hm f + exact norm_condExpL2_le hm f + +@[deprecated (since := "2025-01-21")] alias eLpNorm_condexpL2_le := eLpNorm_condExpL2_le @[deprecated (since := "2024-07-27")] -alias snorm_condexpL2_le := eLpNorm_condexpL2_le +alias snorm_condExpL2_le := eLpNorm_condExpL2_le + +@[deprecated (since := "2025-01-21")] alias snorm_condexpL2_le := snorm_condExpL2_le -theorem norm_condexpL2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : - ‖(condexpL2 E 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := by +theorem norm_condExpL2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : + ‖(condExpL2 E 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := by rw [Lp.norm_def, Lp.norm_def, ← lpMeas_coe] - exact ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) (eLpNorm_condexpL2_le hm f) + exact ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) (eLpNorm_condExpL2_le hm f) -theorem inner_condexpL2_left_eq_right (hm : m ≤ m0) {f g : α →₂[μ] E} : - ⟪(condexpL2 E 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, (condexpL2 E 𝕜 hm g : α →₂[μ] E)⟫₂ := +@[deprecated (since := "2025-01-21")] alias norm_condexpL2_coe_le := norm_condExpL2_coe_le + +theorem inner_condExpL2_left_eq_right (hm : m ≤ m0) {f g : α →₂[μ] E} : + ⟪(condExpL2 E 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, (condExpL2 E 𝕜 hm g : α →₂[μ] E)⟫₂ := haveI : Fact (m ≤ m0) := ⟨hm⟩ inner_orthogonalProjection_left_eq_right _ f g -theorem condexpL2_indicator_of_measurable (hm : m ≤ m0) (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) +@[deprecated (since := "2025-01-21")] +alias inner_condexpL2_left_eq_right := inner_condExpL2_left_eq_right + +theorem condExpL2_indicator_of_measurable (hm : m ≤ m0) (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) (c : E) : - (condexpL2 E 𝕜 hm (indicatorConstLp 2 (hm s hs) hμs c) : α →₂[μ] E) = + (condExpL2 E 𝕜 hm (indicatorConstLp 2 (hm s hs) hμs c) : α →₂[μ] E) = indicatorConstLp 2 (hm s hs) hμs c := by - rw [condexpL2] + rw [condExpL2] haveI : Fact (m ≤ m0) := ⟨hm⟩ have h_mem : indicatorConstLp 2 (hm s hs) hμs c ∈ lpMeas E 𝕜 m 2 μ := mem_lpMeas_indicatorConstLp hm hs hμs @@ -126,34 +153,43 @@ theorem condexpL2_indicator_of_measurable (hm : m ≤ m0) (hs : MeasurableSet[m] have h_orth_mem := orthogonalProjection_mem_subspace_eq_self ind rw [← h_coe_ind, h_orth_mem] -theorem inner_condexpL2_eq_inner_fun (hm : m ≤ m0) (f g : α →₂[μ] E) - (hg : AEStronglyMeasurable' m g μ) : - ⟪(condexpL2 E 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, g⟫₂ := by +@[deprecated (since := "2025-01-21")] +alias condexpL2_indicator_of_measurable := condExpL2_indicator_of_measurable + +theorem inner_condExpL2_eq_inner_fun (hm : m ≤ m0) (f g : α →₂[μ] E) + (hg : AEStronglyMeasurable[m] g μ) : + ⟪(condExpL2 E 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, g⟫₂ := by symm - rw [← sub_eq_zero, ← inner_sub_left, condexpL2] - simp only [mem_lpMeas_iff_aeStronglyMeasurable'.mpr hg, orthogonalProjection_inner_eq_zero f g] + rw [← sub_eq_zero, ← inner_sub_left, condExpL2] + simp only [mem_lpMeas_iff_aeStronglyMeasurable.mpr hg, orthogonalProjection_inner_eq_zero f g] + +@[deprecated (since := "2025-01-21")] +alias inner_condexpL2_eq_inner_fun := inner_condExpL2_eq_inner_fun section Real variable {hm : m ≤ m0} -theorem integral_condexpL2_eq_of_fin_meas_real (f : Lp 𝕜 2 μ) (hs : MeasurableSet[m] s) - (hμs : μ s ≠ ∞) : ∫ x in s, (condexpL2 𝕜 𝕜 hm f : α → 𝕜) x ∂μ = ∫ x in s, f x ∂μ := by +theorem integral_condExpL2_eq_of_fin_meas_real (f : Lp 𝕜 2 μ) (hs : MeasurableSet[m] s) + (hμs : μ s ≠ ∞) : ∫ x in s, (condExpL2 𝕜 𝕜 hm f : α → 𝕜) x ∂μ = ∫ x in s, f x ∂μ := by rw [← L2.inner_indicatorConstLp_one (𝕜 := 𝕜) (hm s hs) hμs f] - have h_eq_inner : ∫ x in s, (condexpL2 𝕜 𝕜 hm f : α → 𝕜) x ∂μ = - inner (indicatorConstLp 2 (hm s hs) hμs (1 : 𝕜)) (condexpL2 𝕜 𝕜 hm f) := by + have h_eq_inner : ∫ x in s, (condExpL2 𝕜 𝕜 hm f : α → 𝕜) x ∂μ = + inner (indicatorConstLp 2 (hm s hs) hμs (1 : 𝕜)) (condExpL2 𝕜 𝕜 hm f) := by rw [L2.inner_indicatorConstLp_one (hm s hs) hμs] - rw [h_eq_inner, ← inner_condexpL2_left_eq_right, condexpL2_indicator_of_measurable hm hs hμs] + rw [h_eq_inner, ← inner_condExpL2_left_eq_right, condExpL2_indicator_of_measurable hm hs hμs] + +@[deprecated (since := "2025-01-21")] +alias integral_condexpL2_eq_of_fin_meas_real := integral_condExpL2_eq_of_fin_meas_real -theorem lintegral_nnnorm_condexpL2_le (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) (f : Lp ℝ 2 μ) : - ∫⁻ x in s, ‖(condexpL2 ℝ ℝ hm f : α → ℝ) x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := by - let h_meas := lpMeas.aeStronglyMeasurable' (condexpL2 ℝ ℝ hm f) +theorem lintegral_nnnorm_condExpL2_le (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) (f : Lp ℝ 2 μ) : + ∫⁻ x in s, ‖(condExpL2 ℝ ℝ hm f : α → ℝ) x‖₊ ∂μ ≤ ∫⁻ x in s, ‖f x‖₊ ∂μ := by + let h_meas := lpMeas.aeStronglyMeasurable (condExpL2 ℝ ℝ hm f) let g := h_meas.choose have hg_meas : StronglyMeasurable[m] g := h_meas.choose_spec.1 - have hg_eq : g =ᵐ[μ] condexpL2 ℝ ℝ hm f := h_meas.choose_spec.2.symm - have hg_eq_restrict : g =ᵐ[μ.restrict s] condexpL2 ℝ ℝ hm f := ae_restrict_of_ae hg_eq + have hg_eq : g =ᵐ[μ] condExpL2 ℝ ℝ hm f := h_meas.choose_spec.2.symm + have hg_eq_restrict : g =ᵐ[μ.restrict s] condExpL2 ℝ ℝ hm f := ae_restrict_of_ae hg_eq have hg_nnnorm_eq : (fun x => (‖g x‖₊ : ℝ≥0∞)) =ᵐ[μ.restrict s] fun x => - (‖(condexpL2 ℝ ℝ hm f : α → ℝ) x‖₊ : ℝ≥0∞) := by + (‖(condExpL2 ℝ ℝ hm f : α → ℝ) x‖₊ : ℝ≥0∞) := by refine hg_eq_restrict.mono fun x hx => ?_ dsimp only simp_rw [hx] @@ -163,14 +199,17 @@ theorem lintegral_nnnorm_condexpL2_le (hs : MeasurableSet[m] s) (hμs : μ s ≠ · exact integrableOn_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs · exact hg_meas · rw [IntegrableOn, integrable_congr hg_eq_restrict] - exact integrableOn_condexpL2_of_measure_ne_top hm hμs f + exact integrableOn_condExpL2_of_measure_ne_top hm hμs f · intro t ht hμt - rw [← integral_condexpL2_eq_of_fin_meas_real f ht hμt.ne] + rw [← integral_condExpL2_eq_of_fin_meas_real f ht hμt.ne] exact setIntegral_congr_ae (hm t ht) (hg_eq.mono fun x hx _ => hx) -theorem condexpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) {f : Lp ℝ 2 μ} - (hf : f =ᵐ[μ.restrict s] 0) : condexpL2 ℝ ℝ hm f =ᵐ[μ.restrict s] (0 : α → ℝ) := by - suffices h_nnnorm_eq_zero : ∫⁻ x in s, ‖(condexpL2 ℝ ℝ hm f : α → ℝ) x‖₊ ∂μ = 0 by +@[deprecated (since := "2025-01-21")] +alias lintegral_nnnorm_condexpL2_le := lintegral_nnnorm_condExpL2_le + +theorem condExpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ s ≠ ∞) {f : Lp ℝ 2 μ} + (hf : f =ᵐ[μ.restrict s] 0) : condExpL2 ℝ ℝ hm f =ᵐ[μ.restrict s] (0 : α → ℝ) := by + suffices h_nnnorm_eq_zero : ∫⁻ x in s, ‖(condExpL2 ℝ ℝ hm f : α → ℝ) x‖₊ ∂μ = 0 by rw [lintegral_eq_zero_iff] at h_nnnorm_eq_zero · refine h_nnnorm_eq_zero.mono fun x hx => ?_ dsimp only at hx @@ -180,7 +219,7 @@ theorem condexpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ rw [lpMeas_coe] exact (Lp.stronglyMeasurable _).measurable refine le_antisymm ?_ (zero_le _) - refine (lintegral_nnnorm_condexpL2_le hs hμs f).trans (le_of_eq ?_) + refine (lintegral_nnnorm_condExpL2_le hs hμs f).trans (le_of_eq ?_) rw [lintegral_eq_zero_iff] · refine hf.mono fun x hx => ?_ dsimp only @@ -188,10 +227,13 @@ theorem condexpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ simp · exact (Lp.stronglyMeasurable _).ennnorm -theorem lintegral_nnnorm_condexpL2_indicator_le_real (hs : MeasurableSet s) (hμs : μ s ≠ ∞) +@[deprecated (since := "2025-01-21")] +alias condexpL2_ae_eq_zero_of_ae_eq_zero := condExpL2_ae_eq_zero_of_ae_eq_zero + +theorem lintegral_nnnorm_condExpL2_indicator_le_real (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (ht : MeasurableSet[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ‖(condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ ≤ μ (s ∩ t) := by - refine (lintegral_nnnorm_condexpL2_le ht hμt _).trans (le_of_eq ?_) + ∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ ≤ μ (s ∩ t) := by + refine (lintegral_nnnorm_condExpL2_le ht hμt _).trans (le_of_eq ?_) have h_eq : ∫⁻ x in t, ‖(indicatorConstLp 2 hs hμs (1 : ℝ)) x‖₊ ∂μ = ∫⁻ x in t, s.indicator (fun _ => (1 : ℝ≥0∞)) x ∂μ := by @@ -205,45 +247,50 @@ theorem lintegral_nnnorm_condexpL2_indicator_le_real (hs : MeasurableSet s) (hμ rw [h_eq, lintegral_indicator hs, lintegral_const, Measure.restrict_restrict hs] simp only [one_mul, Set.univ_inter, MeasurableSet.univ, Measure.restrict_apply] +@[deprecated (since := "2025-01-21")] +alias lintegral_nnnorm_condexpL2_indicator_le_real := lintegral_nnnorm_condExpL2_indicator_le_real + end Real -/-- `condexpL2` commutes with taking inner products with constants. See the lemma -`condexpL2_comp_continuousLinearMap` for a more general result about commuting with continuous +/-- `condExpL2` commutes with taking inner products with constants. See the lemma +`condExpL2_comp_continuousLinearMap` for a more general result about commuting with continuous linear maps. -/ -theorem condexpL2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) : - condexpL2 𝕜 𝕜 hm (((Lp.memℒp f).const_inner c).toLp fun a => ⟪c, f a⟫) =ᵐ[μ] - fun a => ⟪c, (condexpL2 E 𝕜 hm f : α → E) a⟫ := by +theorem condExpL2_const_inner (hm : m ≤ m0) (f : Lp E 2 μ) (c : E) : + condExpL2 𝕜 𝕜 hm (((Lp.memℒp f).const_inner c).toLp fun a => ⟪c, f a⟫) =ᵐ[μ] + fun a => ⟪c, (condExpL2 E 𝕜 hm f : α → E) a⟫ := by rw [lpMeas_coe] - have h_mem_Lp : Memℒp (fun a => ⟪c, (condexpL2 E 𝕜 hm f : α → E) a⟫) 2 μ := by + have h_mem_Lp : Memℒp (fun a => ⟪c, (condExpL2 E 𝕜 hm f : α → E) a⟫) 2 μ := by refine Memℒp.const_inner _ ?_; rw [lpMeas_coe]; exact Lp.memℒp _ - have h_eq : h_mem_Lp.toLp _ =ᵐ[μ] fun a => ⟪c, (condexpL2 E 𝕜 hm f : α → E) a⟫ := + have h_eq : h_mem_Lp.toLp _ =ᵐ[μ] fun a => ⟪c, (condExpL2 E 𝕜 hm f : α → E) a⟫ := h_mem_Lp.coeFn_toLp refine EventuallyEq.trans ?_ h_eq refine Lp.ae_eq_of_forall_setIntegral_eq' 𝕜 hm _ _ two_ne_zero ENNReal.coe_ne_top - (fun s _ hμs => integrableOn_condexpL2_of_measure_ne_top hm hμs.ne _) ?_ ?_ ?_ ?_ + (fun s _ hμs => integrableOn_condExpL2_of_measure_ne_top hm hμs.ne _) ?_ ?_ ?_ ?_ · intro s _ hμs rw [IntegrableOn, integrable_congr (ae_restrict_of_ae h_eq)] - exact (integrableOn_condexpL2_of_measure_ne_top hm hμs.ne _).const_inner _ + exact (integrableOn_condExpL2_of_measure_ne_top hm hμs.ne _).const_inner _ · intro s hs hμs - rw [← lpMeas_coe, integral_condexpL2_eq_of_fin_meas_real _ hs hμs.ne, + rw [← lpMeas_coe, integral_condExpL2_eq_of_fin_meas_real _ hs hμs.ne, integral_congr_ae (ae_restrict_of_ae h_eq), lpMeas_coe, ← - L2.inner_indicatorConstLp_eq_setIntegral_inner 𝕜 (↑(condexpL2 E 𝕜 hm f)) (hm s hs) c hμs.ne, - ← inner_condexpL2_left_eq_right, condexpL2_indicator_of_measurable _ hs, + L2.inner_indicatorConstLp_eq_setIntegral_inner 𝕜 (↑(condExpL2 E 𝕜 hm f)) (hm s hs) c hμs.ne, + ← inner_condExpL2_left_eq_right, condExpL2_indicator_of_measurable _ hs, L2.inner_indicatorConstLp_eq_setIntegral_inner 𝕜 f (hm s hs) c hμs.ne, setIntegral_congr_ae (hm s hs) ((Memℒp.coeFn_toLp ((Lp.memℒp f).const_inner c)).mono fun x hx _ => hx)] - · rw [← lpMeas_coe]; exact lpMeas.aeStronglyMeasurable' _ - · refine AEStronglyMeasurable'.congr ?_ h_eq.symm - exact (lpMeas.aeStronglyMeasurable' _).const_inner _ + · rw [← lpMeas_coe]; exact lpMeas.aeStronglyMeasurable _ + · refine AEStronglyMeasurable.congr ?_ h_eq.symm + exact (lpMeas.aeStronglyMeasurable _).const_inner + +@[deprecated (since := "2025-01-21")] alias condexpL2_const_inner := condExpL2_const_inner -/-- `condexpL2` verifies the equality of integrals defining the conditional expectation. -/ -theorem integral_condexpL2_eq (hm : m ≤ m0) (f : Lp E' 2 μ) (hs : MeasurableSet[m] s) - (hμs : μ s ≠ ∞) : ∫ x in s, (condexpL2 E' 𝕜 hm f : α → E') x ∂μ = ∫ x in s, f x ∂μ := by +/-- `condExpL2` verifies the equality of integrals defining the conditional expectation. -/ +theorem integral_condExpL2_eq (hm : m ≤ m0) (f : Lp E' 2 μ) (hs : MeasurableSet[m] s) + (hμs : μ s ≠ ∞) : ∫ x in s, (condExpL2 E' 𝕜 hm f : α → E') x ∂μ = ∫ x in s, f x ∂μ := by rw [← sub_eq_zero, lpMeas_coe, ← integral_sub' (integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs) (integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs)] refine integral_eq_zero_of_forall_integral_inner_eq_zero 𝕜 _ ?_ ?_ - · rw [integrable_congr (ae_restrict_of_ae (Lp.coeFn_sub (↑(condexpL2 E' 𝕜 hm f)) f).symm)] + · rw [integrable_congr (ae_restrict_of_ae (Lp.coeFn_sub (↑(condExpL2 E' 𝕜 hm f)) f).symm)] exact integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs intro c simp_rw [Pi.sub_apply, inner_sub_right] @@ -252,34 +299,39 @@ theorem integral_condexpL2_eq (hm : m ≤ m0) (f : Lp E' 2 μ) (hs : MeasurableS ((integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs).const_inner c)] have h_ae_eq_f := Memℒp.coeFn_toLp (E := 𝕜) ((Lp.memℒp f).const_inner c) rw [← lpMeas_coe, sub_eq_zero, ← - setIntegral_congr_ae (hm s hs) ((condexpL2_const_inner hm f c).mono fun x hx _ => hx), ← + setIntegral_congr_ae (hm s hs) ((condExpL2_const_inner hm f c).mono fun x hx _ => hx), ← setIntegral_congr_ae (hm s hs) (h_ae_eq_f.mono fun x hx _ => hx)] - exact integral_condexpL2_eq_of_fin_meas_real _ hs hμs + exact integral_condExpL2_eq_of_fin_meas_real _ hs hμs + +@[deprecated (since := "2025-01-21")] alias integral_condexpL2_eq := integral_condExpL2_eq variable {E'' 𝕜' : Type*} [RCLike 𝕜'] [NormedAddCommGroup E''] [InnerProductSpace 𝕜' E''] [CompleteSpace E''] [NormedSpace ℝ E''] variable (𝕜 𝕜') -theorem condexpL2_comp_continuousLinearMap (hm : m ≤ m0) (T : E' →L[ℝ] E'') (f : α →₂[μ] E') : - (condexpL2 E'' 𝕜' hm (T.compLp f) : α →₂[μ] E'') =ᵐ[μ] - T.compLp (condexpL2 E' 𝕜 hm f : α →₂[μ] E') := by +theorem condExpL2_comp_continuousLinearMap (hm : m ≤ m0) (T : E' →L[ℝ] E'') (f : α →₂[μ] E') : + (condExpL2 E'' 𝕜' hm (T.compLp f) : α →₂[μ] E'') =ᵐ[μ] + T.compLp (condExpL2 E' 𝕜 hm f : α →₂[μ] E') := by refine Lp.ae_eq_of_forall_setIntegral_eq' 𝕜' hm _ _ two_ne_zero ENNReal.coe_ne_top - (fun s _ hμs => integrableOn_condexpL2_of_measure_ne_top hm hμs.ne _) (fun s _ hμs => + (fun s _ hμs => integrableOn_condExpL2_of_measure_ne_top hm hμs.ne _) (fun s _ hμs => integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne) ?_ ?_ ?_ · intro s hs hμs rw [T.setIntegral_compLp _ (hm s hs), T.integral_comp_comm (integrableOn_Lp_of_measure_ne_top _ fact_one_le_two_ennreal.elim hμs.ne), - ← lpMeas_coe, ← lpMeas_coe, integral_condexpL2_eq hm f hs hμs.ne, - integral_condexpL2_eq hm (T.compLp f) hs hμs.ne, T.setIntegral_compLp _ (hm s hs), + ← lpMeas_coe, ← lpMeas_coe, integral_condExpL2_eq hm f hs hμs.ne, + integral_condExpL2_eq hm (T.compLp f) hs hμs.ne, T.setIntegral_compLp _ (hm s hs), T.integral_comp_comm (integrableOn_Lp_of_measure_ne_top f fact_one_le_two_ennreal.elim hμs.ne)] - · rw [← lpMeas_coe]; exact lpMeas.aeStronglyMeasurable' _ - · have h_coe := T.coeFn_compLp (condexpL2 E' 𝕜 hm f : α →₂[μ] E') + · rw [← lpMeas_coe]; exact lpMeas.aeStronglyMeasurable _ + · have h_coe := T.coeFn_compLp (condExpL2 E' 𝕜 hm f : α →₂[μ] E') rw [← EventuallyEq] at h_coe - refine AEStronglyMeasurable'.congr ?_ h_coe.symm - exact (lpMeas.aeStronglyMeasurable' (condexpL2 E' 𝕜 hm f)).continuous_comp T.continuous + refine AEStronglyMeasurable.congr ?_ h_coe.symm + exact T.continuous.comp_aestronglyMeasurable (lpMeas.aeStronglyMeasurable (condExpL2 E' 𝕜 hm f)) + +@[deprecated (since := "2025-01-21")] +alias condexpL2_comp_continuousLinearMap := condExpL2_comp_continuousLinearMap variable {𝕜 𝕜'} @@ -287,72 +339,90 @@ section CondexpL2Indicator variable (𝕜) -theorem condexpL2_indicator_ae_eq_smul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) +theorem condExpL2_indicator_ae_eq_smul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E') : - condexpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) =ᵐ[μ] fun a => - (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs (1 : ℝ)) : α → ℝ) a • x := by + condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) =ᵐ[μ] fun a => + (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs (1 : ℝ)) : α → ℝ) a • x := by rw [indicatorConstLp_eq_toSpanSingleton_compLp hs hμs x] have h_comp := - condexpL2_comp_continuousLinearMap ℝ 𝕜 hm (toSpanSingleton ℝ x) + condExpL2_comp_continuousLinearMap ℝ 𝕜 hm (toSpanSingleton ℝ x) (indicatorConstLp 2 hs hμs (1 : ℝ)) rw [← lpMeas_coe] at h_comp refine h_comp.trans ?_ exact (toSpanSingleton ℝ x).coeFn_compLp _ -theorem condexpL2_indicator_eq_toSpanSingleton_comp (hm : m ≤ m0) (hs : MeasurableSet s) - (hμs : μ s ≠ ∞) (x : E') : (condexpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α →₂[μ] E') = - (toSpanSingleton ℝ x).compLp (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1)) := by +@[deprecated (since := "2025-01-21")] +alias condexpL2_indicator_ae_eq_smul := condExpL2_indicator_ae_eq_smul + +theorem condExpL2_indicator_eq_toSpanSingleton_comp (hm : m ≤ m0) (hs : MeasurableSet s) + (hμs : μ s ≠ ∞) (x : E') : (condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α →₂[μ] E') = + (toSpanSingleton ℝ x).compLp (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1)) := by ext1 rw [← lpMeas_coe] - refine (condexpL2_indicator_ae_eq_smul 𝕜 hm hs hμs x).trans ?_ + refine (condExpL2_indicator_ae_eq_smul 𝕜 hm hs hμs x).trans ?_ have h_comp := (toSpanSingleton ℝ x).coeFn_compLp - (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α →₂[μ] ℝ) + (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α →₂[μ] ℝ) rw [← EventuallyEq] at h_comp refine EventuallyEq.trans ?_ h_comp.symm filter_upwards with y using rfl +@[deprecated (since := "2025-01-21")] +alias condexpL2_indicator_eq_toSpanSingleton_comp := condExpL2_indicator_eq_toSpanSingleton_comp + variable {𝕜} -theorem setLIntegral_nnnorm_condexpL2_indicator_le (hm : m ≤ m0) (hs : MeasurableSet s) +theorem setLIntegral_nnnorm_condExpL2_indicator_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E') {t : Set α} (ht : MeasurableSet[m] t) (hμt : μ t ≠ ∞) : - ∫⁻ a in t, ‖(condexpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ ≤ + ∫⁻ a in t, ‖(condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ ≤ μ (s ∩ t) * ‖x‖₊ := calc - ∫⁻ a in t, ‖(condexpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ = - ∫⁻ a in t, ‖(condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a • x‖₊ ∂μ := + ∫⁻ a in t, ‖(condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ = + ∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a • x‖₊ ∂μ := setLIntegral_congr_fun (hm t ht) - ((condexpL2_indicator_ae_eq_smul 𝕜 hm hs hμs x).mono fun a ha _ => by rw [ha]) - _ = (∫⁻ a in t, ‖(condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by + ((condExpL2_indicator_ae_eq_smul 𝕜 hm hs hμs x).mono fun a ha _ => by rw [ha]) + _ = (∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by simp_rw [nnnorm_smul, ENNReal.coe_mul] rw [lintegral_mul_const, lpMeas_coe] exact (Lp.stronglyMeasurable _).ennnorm _ ≤ μ (s ∩ t) * ‖x‖₊ := - mul_le_mul_right' (lintegral_nnnorm_condexpL2_indicator_le_real hs hμs ht hμt) _ + mul_le_mul_right' (lintegral_nnnorm_condExpL2_indicator_le_real hs hμs ht hμt) _ + +@[deprecated (since := "2025-01-21")] +alias setLIntegral_nnnorm_condexpL2_indicator_le := setLIntegral_nnnorm_condExpL2_indicator_le @[deprecated (since := "2024-06-29")] -alias set_lintegral_nnnorm_condexpL2_indicator_le := setLIntegral_nnnorm_condexpL2_indicator_le +alias set_lintegral_nnnorm_condExpL2_indicator_le := setLIntegral_nnnorm_condExpL2_indicator_le + +@[deprecated (since := "2025-01-21")] +alias set_lintegral_nnnorm_condexpL2_indicator_le := set_lintegral_nnnorm_condExpL2_indicator_le -theorem lintegral_nnnorm_condexpL2_indicator_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) +theorem lintegral_nnnorm_condExpL2_indicator_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E') [SigmaFinite (μ.trim hm)] : - ∫⁻ a, ‖(condexpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ ≤ μ s * ‖x‖₊ := by + ∫⁻ a, ‖(condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x) : α → E') a‖₊ ∂μ ≤ μ s * ‖x‖₊ := by refine lintegral_le_of_forall_fin_meas_trim_le hm (μ s * ‖x‖₊) fun t ht hμt => ?_ - refine (setLIntegral_nnnorm_condexpL2_indicator_le hm hs hμs x ht hμt).trans ?_ + refine (setLIntegral_nnnorm_condExpL2_indicator_le hm hs hμs x ht hμt).trans ?_ gcongr apply Set.inter_subset_left +@[deprecated (since := "2025-01-21")] +alias lintegral_nnnorm_condexpL2_indicator_le := lintegral_nnnorm_condExpL2_indicator_le + /-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set with finite measure is integrable. -/ -theorem integrable_condexpL2_indicator (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] +theorem integrable_condExpL2_indicator (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E') : - Integrable (ε := E') (condexpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x)) μ := by + Integrable (ε := E') (condExpL2 E' 𝕜 hm (indicatorConstLp 2 hs hμs x)) μ := by refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) (ENNReal.mul_lt_top hμs.lt_top ENNReal.coe_lt_top) ?_ ?_ · rw [lpMeas_coe]; exact Lp.aestronglyMeasurable _ · refine fun t ht hμt => - (setLIntegral_nnnorm_condexpL2_indicator_le hm hs hμs x ht hμt).trans ?_ + (setLIntegral_nnnorm_condExpL2_indicator_le hm hs hμs x ht hμt).trans ?_ gcongr apply Set.inter_subset_left +@[deprecated (since := "2025-01-21")] +alias integrable_condexpL2_indicator := integrable_condExpL2_indicator + end CondexpL2Indicator section CondexpIndSMul @@ -360,115 +430,151 @@ section CondexpIndSMul variable [NormedSpace ℝ G] {hm : m ≤ m0} /-- Conditional expectation of the indicator of a measurable set with finite measure, in L2. -/ -noncomputable def condexpIndSMul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : +noncomputable def condExpIndSMul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : Lp G 2 μ := - (toSpanSingleton ℝ x).compLpL 2 μ (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs (1 : ℝ))) - -theorem aeStronglyMeasurable'_condexpIndSMul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) - (x : G) : AEStronglyMeasurable' m (condexpIndSMul hm hs hμs x) μ := by - have h : AEStronglyMeasurable' m (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) μ := - aeStronglyMeasurable'_condexpL2 _ _ - rw [condexpIndSMul] - suffices AEStronglyMeasurable' m - (toSpanSingleton ℝ x ∘ condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1)) μ by - refine AEStronglyMeasurable'.congr this ?_ - refine EventuallyEq.trans ?_ (coeFn_compLpL _ _).symm - rfl - exact AEStronglyMeasurable'.continuous_comp (toSpanSingleton ℝ x).continuous h - -theorem condexpIndSMul_add (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x y : G) : - condexpIndSMul hm hs hμs (x + y) = condexpIndSMul hm hs hμs x + condexpIndSMul hm hs hμs y := by - simp_rw [condexpIndSMul]; rw [toSpanSingleton_add, add_compLpL, add_apply] - -theorem condexpIndSMul_smul (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : - condexpIndSMul hm hs hμs (c • x) = c • condexpIndSMul hm hs hμs x := by - simp_rw [condexpIndSMul]; rw [toSpanSingleton_smul, smul_compLpL, smul_apply] - -theorem condexpIndSMul_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (hs : MeasurableSet s) + (toSpanSingleton ℝ x).compLpL 2 μ (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs (1 : ℝ))) + +@[deprecated (since := "2025-01-21")] alias condexpIndSMul := condExpIndSMul + +theorem aestronglyMeasurable_condExpIndSMul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) + (x : G) : AEStronglyMeasurable[m] (condExpIndSMul hm hs hμs x) μ := by + have h : AEStronglyMeasurable[m] (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) μ := + aestronglyMeasurable_condExpL2 _ _ + rw [condExpIndSMul] + exact ((toSpanSingleton ℝ x).continuous.comp_aestronglyMeasurable h).congr + (coeFn_compLpL _ _).symm + +@[deprecated (since := "2025-01-24")] +alias aeStronglyMeasurable'_condExpIndSMul := aestronglyMeasurable_condExpIndSMul + +@[deprecated (since := "2025-01-21")] +alias aestronglyMeasurable'_condexpIndSMul := aestronglyMeasurable_condExpIndSMul + +theorem condExpIndSMul_add (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x y : G) : + condExpIndSMul hm hs hμs (x + y) = condExpIndSMul hm hs hμs x + condExpIndSMul hm hs hμs y := by + simp_rw [condExpIndSMul]; rw [toSpanSingleton_add, add_compLpL, add_apply] + +@[deprecated (since := "2025-01-21")] alias condexpIndSMul_add := condExpIndSMul_add + +theorem condExpIndSMul_smul (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : + condExpIndSMul hm hs hμs (c • x) = c • condExpIndSMul hm hs hμs x := by + simp_rw [condExpIndSMul]; rw [toSpanSingleton_smul, smul_compLpL, smul_apply] + +@[deprecated (since := "2025-01-21")] alias condexpIndSMul_smul := condExpIndSMul_smul + +theorem condExpIndSMul_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : 𝕜) (x : F) : - condexpIndSMul hm hs hμs (c • x) = c • condexpIndSMul hm hs hμs x := by - rw [condexpIndSMul, condexpIndSMul, toSpanSingleton_smul', + condExpIndSMul hm hs hμs (c • x) = c • condExpIndSMul hm hs hμs x := by + rw [condExpIndSMul, condExpIndSMul, toSpanSingleton_smul', (toSpanSingleton ℝ x).smul_compLpL c, smul_apply] -theorem condexpIndSMul_ae_eq_smul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : - condexpIndSMul hm hs hμs x =ᵐ[μ] fun a => - (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a • x := +@[deprecated (since := "2025-01-21")] alias condexpIndSMul_smul' := condExpIndSMul_smul' + +theorem condExpIndSMul_ae_eq_smul (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : + condExpIndSMul hm hs hμs x =ᵐ[μ] fun a => + (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a • x := (toSpanSingleton ℝ x).coeFn_compLpL _ -theorem setLIntegral_nnnorm_condexpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) +@[deprecated (since := "2025-01-21")] alias condexpIndSMul_ae_eq_smul := condExpIndSMul_ae_eq_smul + +theorem setLIntegral_nnnorm_condExpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) {t : Set α} (ht : MeasurableSet[m] t) (hμt : μ t ≠ ∞) : - (∫⁻ a in t, ‖condexpIndSMul hm hs hμs x a‖₊ ∂μ) ≤ μ (s ∩ t) * ‖x‖₊ := + (∫⁻ a in t, ‖condExpIndSMul hm hs hμs x a‖₊ ∂μ) ≤ μ (s ∩ t) * ‖x‖₊ := calc - ∫⁻ a in t, ‖condexpIndSMul hm hs hμs x a‖₊ ∂μ = - ∫⁻ a in t, ‖(condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a • x‖₊ ∂μ := + ∫⁻ a in t, ‖condExpIndSMul hm hs hμs x a‖₊ ∂μ = + ∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a • x‖₊ ∂μ := setLIntegral_congr_fun (hm t ht) - ((condexpIndSMul_ae_eq_smul hm hs hμs x).mono fun a ha _ => by rw [ha]) - _ = (∫⁻ a in t, ‖(condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by + ((condExpIndSMul_ae_eq_smul hm hs hμs x).mono fun a ha _ => by rw [ha]) + _ = (∫⁻ a in t, ‖(condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) a‖₊ ∂μ) * ‖x‖₊ := by simp_rw [nnnorm_smul, ENNReal.coe_mul] rw [lintegral_mul_const, lpMeas_coe] exact (Lp.stronglyMeasurable _).ennnorm _ ≤ μ (s ∩ t) * ‖x‖₊ := - mul_le_mul_right' (lintegral_nnnorm_condexpL2_indicator_le_real hs hμs ht hμt) _ + mul_le_mul_right' (lintegral_nnnorm_condExpL2_indicator_le_real hs hμs ht hμt) _ + +@[deprecated (since := "2025-01-21")] +alias setLIntegral_nnnorm_condexpIndSMul_le := setLIntegral_nnnorm_condExpIndSMul_le @[deprecated (since := "2024-06-29")] -alias set_lintegral_nnnorm_condexpIndSMul_le := setLIntegral_nnnorm_condexpIndSMul_le +alias set_lintegral_nnnorm_condExpIndSMul_le := setLIntegral_nnnorm_condExpIndSMul_le + +@[deprecated (since := "2025-01-21")] +alias set_lintegral_nnnorm_condexpIndSMul_le := set_lintegral_nnnorm_condExpIndSMul_le -theorem lintegral_nnnorm_condexpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) - (x : G) [SigmaFinite (μ.trim hm)] : ∫⁻ a, ‖condexpIndSMul hm hs hμs x a‖₊ ∂μ ≤ μ s * ‖x‖₊ := by +theorem lintegral_nnnorm_condExpIndSMul_le (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) + (x : G) [SigmaFinite (μ.trim hm)] : ∫⁻ a, ‖condExpIndSMul hm hs hμs x a‖₊ ∂μ ≤ μ s * ‖x‖₊ := by refine lintegral_le_of_forall_fin_meas_trim_le hm (μ s * ‖x‖₊) fun t ht hμt => ?_ - refine (setLIntegral_nnnorm_condexpIndSMul_le hm hs hμs x ht hμt).trans ?_ + refine (setLIntegral_nnnorm_condExpIndSMul_le hm hs hμs x ht hμt).trans ?_ gcongr apply Set.inter_subset_left +@[deprecated (since := "2025-01-21")] +alias lintegral_nnnorm_condexpIndSMul_le := lintegral_nnnorm_condExpIndSMul_le + /-- If the measure `μ.trim hm` is sigma-finite, then the conditional expectation of a measurable set with finite measure is integrable. -/ -theorem integrable_condexpIndSMul (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) - (hμs : μ s ≠ ∞) (x : G) : Integrable (condexpIndSMul hm hs hμs x) μ := by +theorem integrable_condExpIndSMul (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) + (hμs : μ s ≠ ∞) (x : G) : Integrable (condExpIndSMul hm hs hμs x) μ := by refine integrable_of_forall_fin_meas_le' hm (μ s * ‖x‖₊) (ENNReal.mul_lt_top hμs.lt_top ENNReal.coe_lt_top) ?_ ?_ · exact Lp.aestronglyMeasurable _ - · refine fun t ht hμt => (setLIntegral_nnnorm_condexpIndSMul_le hm hs hμs x ht hμt).trans ?_ + · refine fun t ht hμt => (setLIntegral_nnnorm_condExpIndSMul_le hm hs hμs x ht hμt).trans ?_ gcongr apply Set.inter_subset_left -theorem condexpIndSMul_empty {x : G} : condexpIndSMul hm MeasurableSet.empty +@[deprecated (since := "2025-01-21")] alias integrable_condexpIndSMul := integrable_condExpIndSMul + +theorem condExpIndSMul_empty {x : G} : condExpIndSMul hm MeasurableSet.empty ((measure_empty (μ := μ)).le.trans_lt ENNReal.coe_lt_top).ne x = 0 := by - rw [condexpIndSMul, indicatorConstLp_empty] + rw [condExpIndSMul, indicatorConstLp_empty] simp only [Submodule.coe_zero, ContinuousLinearMap.map_zero] -theorem setIntegral_condexpL2_indicator (hs : MeasurableSet[m] s) (ht : MeasurableSet t) +@[deprecated (since := "2025-01-21")] alias condexpIndSMul_empty := condExpIndSMul_empty + +theorem setIntegral_condExpL2_indicator (hs : MeasurableSet[m] s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) : - ∫ x in s, (condexpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) x ∂μ = (μ (t ∩ s)).toReal := + ∫ x in s, (condExpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) x ∂μ = (μ (t ∩ s)).toReal := calc - ∫ x in s, (condexpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) x ∂μ = + ∫ x in s, (condExpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) x ∂μ = ∫ x in s, indicatorConstLp 2 ht hμt (1 : ℝ) x ∂μ := - @integral_condexpL2_eq α _ ℝ _ _ _ _ _ _ _ _ _ hm (indicatorConstLp 2 ht hμt (1 : ℝ)) hs hμs + @integral_condExpL2_eq α _ ℝ _ _ _ _ _ _ _ _ _ hm (indicatorConstLp 2 ht hμt (1 : ℝ)) hs hμs _ = (μ (t ∩ s)).toReal • (1 : ℝ) := setIntegral_indicatorConstLp (hm s hs) ht hμt 1 _ = (μ (t ∩ s)).toReal := by rw [smul_eq_mul, mul_one] +@[deprecated (since := "2025-01-21")] +alias setIntegral_condexpL2_indicator := setIntegral_condExpL2_indicator + @[deprecated (since := "2024-04-17")] -alias set_integral_condexpL2_indicator := setIntegral_condexpL2_indicator +alias set_integral_condExpL2_indicator := setIntegral_condExpL2_indicator + +@[deprecated (since := "2025-01-21")] +alias set_integral_condexpL2_indicator := set_integral_condExpL2_indicator -theorem setIntegral_condexpIndSMul (hs : MeasurableSet[m] s) (ht : MeasurableSet t) +theorem setIntegral_condExpIndSMul (hs : MeasurableSet[m] s) (ht : MeasurableSet t) (hμs : μ s ≠ ∞) (hμt : μ t ≠ ∞) (x : G') : - ∫ a in s, (condexpIndSMul hm ht hμt x) a ∂μ = (μ (t ∩ s)).toReal • x := + ∫ a in s, (condExpIndSMul hm ht hμt x) a ∂μ = (μ (t ∩ s)).toReal • x := calc - ∫ a in s, (condexpIndSMul hm ht hμt x) a ∂μ = - ∫ a in s, (condexpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) a • x ∂μ := + ∫ a in s, (condExpIndSMul hm ht hμt x) a ∂μ = + ∫ a in s, (condExpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) a • x ∂μ := setIntegral_congr_ae (hm s hs) - ((condexpIndSMul_ae_eq_smul hm ht hμt x).mono fun _ hx _ => hx) - _ = (∫ a in s, (condexpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) a ∂μ) • x := + ((condExpIndSMul_ae_eq_smul hm ht hμt x).mono fun _ hx _ => hx) + _ = (∫ a in s, (condExpL2 ℝ ℝ hm (indicatorConstLp 2 ht hμt 1) : α → ℝ) a ∂μ) • x := (integral_smul_const _ x) - _ = (μ (t ∩ s)).toReal • x := by rw [setIntegral_condexpL2_indicator hs ht hμs hμt] + _ = (μ (t ∩ s)).toReal • x := by rw [setIntegral_condExpL2_indicator hs ht hμs hμt] + +@[deprecated (since := "2025-01-21")] alias setIntegral_condexpIndSMul := setIntegral_condExpIndSMul @[deprecated (since := "2024-04-17")] -alias set_integral_condexpIndSMul := setIntegral_condexpIndSMul +alias set_integral_condExpIndSMul := setIntegral_condExpIndSMul -theorem condexpL2_indicator_nonneg (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) +@[deprecated (since := "2025-01-21")] +alias set_integral_condexpIndSMul := set_integral_condExpIndSMul + +theorem condExpL2_indicator_nonneg (hm : m ≤ m0) (hs : MeasurableSet s) (hμs : μ s ≠ ∞) [SigmaFinite (μ.trim hm)] : (0 : α → ℝ) ≤ᵐ[μ] - condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) := by - have h : AEStronglyMeasurable' m (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) μ := - aeStronglyMeasurable'_condexpL2 _ _ + condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) := by + have h : AEStronglyMeasurable[m] (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) μ := + aestronglyMeasurable_condExpL2 _ _ refine EventuallyLE.trans_eq ?_ h.ae_eq_mk.symm refine @ae_le_of_ae_le_trim _ _ _ _ _ _ hm (0 : α → ℝ) _ ?_ refine ae_nonneg_of_forall_setIntegral_nonneg_of_sigmaFinite ?_ ?_ @@ -476,25 +582,29 @@ theorem condexpL2_indicator_nonneg (hm : m ≤ m0) (hs : MeasurableSet s) (hμs refine @Integrable.integrableOn _ _ m _ _ _ _ ?_ refine Integrable.trim hm ?_ ?_ · rw [integrable_congr h.ae_eq_mk.symm] - exact integrable_condexpL2_indicator hm hs hμs _ + exact integrable_condExpL2_indicator hm hs hμs _ · exact h.stronglyMeasurable_mk · intro t ht hμt rw [← setIntegral_trim hm h.stronglyMeasurable_mk ht] have h_ae : - ∀ᵐ x ∂μ, x ∈ t → h.mk _ x = (condexpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) x := by + ∀ᵐ x ∂μ, x ∈ t → h.mk _ x = (condExpL2 ℝ ℝ hm (indicatorConstLp 2 hs hμs 1) : α → ℝ) x := by filter_upwards [h.ae_eq_mk] with x hx exact fun _ => hx.symm rw [setIntegral_congr_ae (hm t ht) h_ae, - setIntegral_condexpL2_indicator ht hs ((le_trim hm).trans_lt hμt).ne hμs] + setIntegral_condExpL2_indicator ht hs ((le_trim hm).trans_lt hμt).ne hμs] exact ENNReal.toReal_nonneg -theorem condexpIndSMul_nonneg {E} [NormedLatticeAddCommGroup E] [NormedSpace ℝ E] [OrderedSMul ℝ E] +@[deprecated (since := "2025-01-21")] alias condexpL2_indicator_nonneg := condExpL2_indicator_nonneg + +theorem condExpIndSMul_nonneg {E} [NormedLatticeAddCommGroup E] [NormedSpace ℝ E] [OrderedSMul ℝ E] [SigmaFinite (μ.trim hm)] (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : E) (hx : 0 ≤ x) : - (0 : α → E) ≤ᵐ[μ] condexpIndSMul hm hs hμs x := by - refine EventuallyLE.trans_eq ?_ (condexpIndSMul_ae_eq_smul hm hs hμs x).symm - filter_upwards [condexpL2_indicator_nonneg hm hs hμs] with a ha + (0 : α → E) ≤ᵐ[μ] condExpIndSMul hm hs hμs x := by + refine EventuallyLE.trans_eq ?_ (condExpIndSMul_ae_eq_smul hm hs hμs x).symm + filter_upwards [condExpL2_indicator_nonneg hm hs hμs] with a ha exact smul_nonneg ha hx +@[deprecated (since := "2025-01-21")] alias condexpIndSMul_nonneg := condExpIndSMul_nonneg + end CondexpIndSMul end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean index 260205e45cac9..8610e0a9cbe52 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean @@ -15,7 +15,7 @@ a restricted measure. ## Main result -* `MeasureTheory.condexp_indicator`: If `s` is an `m`-measurable set, then the conditional +* `MeasureTheory.condExp_indicator`: If `s` is an `m`-measurable set, then the conditional expectation of the indicator function of `s` is almost everywhere equal to the indicator of `s` of the conditional expectation. Namely, `𝔼[s.indicator f | m] = s.indicator 𝔼[f | m]` a.e. @@ -33,47 +33,52 @@ namespace MeasureTheory variable {α E : Type*} {m m0 : MeasurableSpace α} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] {μ : Measure α} {f : α → E} {s : Set α} -theorem condexp_ae_eq_restrict_zero (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.restrict s] 0) : +theorem condExp_ae_eq_restrict_zero (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.restrict s] 0) : μ[f|m] =ᵐ[μ.restrict s] 0 := by by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm]; rfl + swap; · simp_rw [condExp_of_not_le hm]; rfl by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; rfl + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; rfl haveI : SigmaFinite (μ.trim hm) := hμm have : SigmaFinite ((μ.restrict s).trim hm) := by rw [← restrict_trim hm _ hs] exact Restrict.sigmaFinite _ s by_cases hf_int : Integrable f μ - swap; · rw [condexp_undef hf_int] + swap; · rw [condExp_of_not_integrable hf_int] refine ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' hm ?_ ?_ ?_ ?_ ?_ - · exact fun t _ _ => integrable_condexp.integrableOn.integrableOn + · exact fun t _ _ => integrable_condExp.integrableOn.integrableOn · exact fun t _ _ => (integrable_zero _ _ _).integrableOn · intro t ht _ - rw [Measure.restrict_restrict (hm _ ht), setIntegral_condexp hm hf_int (ht.inter hs), ← + rw [Measure.restrict_restrict (hm _ ht), setIntegral_condExp hm hf_int (ht.inter hs), ← Measure.restrict_restrict (hm _ ht)] refine setIntegral_congr_ae (hm _ ht) ?_ filter_upwards [hf] with x hx _ using hx - · exact stronglyMeasurable_condexp.aeStronglyMeasurable' - · exact stronglyMeasurable_zero.aeStronglyMeasurable' + · exact stronglyMeasurable_condExp.aestronglyMeasurable + · exact stronglyMeasurable_zero.aestronglyMeasurable -/-- Auxiliary lemma for `condexp_indicator`. -/ -theorem condexp_indicator_aux (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.restrict sᶜ] 0) : +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_restrict_zero := condExp_ae_eq_restrict_zero + +/-- Auxiliary lemma for `condExp_indicator`. -/ +theorem condExp_indicator_aux (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.restrict sᶜ] 0) : μ[s.indicator f|m] =ᵐ[μ] s.indicator (μ[f|m]) := by by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm, Set.indicator_zero']; rfl + swap; · simp_rw [condExp_of_not_le hm, Set.indicator_zero']; rfl have hsf_zero : ∀ g : α → E, g =ᵐ[μ.restrict sᶜ] 0 → s.indicator g =ᵐ[μ] g := fun g => indicator_ae_eq_of_restrict_compl_ae_eq_zero (hm _ hs) - refine ((hsf_zero (μ[f|m]) (condexp_ae_eq_restrict_zero hs.compl hf)).trans ?_).symm - exact condexp_congr_ae (hsf_zero f hf).symm + refine ((hsf_zero (μ[f|m]) (condExp_ae_eq_restrict_zero hs.compl hf)).trans ?_).symm + exact condExp_congr_ae (hsf_zero f hf).symm + +@[deprecated (since := "2025-01-21")] alias condexp_indicator_aux := condExp_indicator_aux /-- The conditional expectation of the indicator of a function over an `m`-measurable set with respect to the σ-algebra `m` is a.e. equal to the indicator of the conditional expectation. -/ -theorem condexp_indicator (hf_int : Integrable f μ) (hs : MeasurableSet[m] s) : +theorem condExp_indicator (hf_int : Integrable f μ) (hs : MeasurableSet[m] s) : μ[s.indicator f|m] =ᵐ[μ] s.indicator (μ[f|m]) := by by_cases hm : m ≤ m0 - swap; · simp_rw [condexp_of_not_le hm, Set.indicator_zero']; rfl + swap; · simp_rw [condExp_of_not_le hm, Set.indicator_zero']; rfl by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm, Set.indicator_zero']; rfl + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm, Set.indicator_zero']; rfl haveI : SigmaFinite (μ.trim hm) := hμm -- use `have` to perform what should be the first calc step because of an error I don't -- understand @@ -83,9 +88,8 @@ theorem condexp_indicator (hf_int : Integrable f μ) (hs : MeasurableSet[m] s) : calc s.indicator (μ[s.indicator f + sᶜ.indicator f|m]) =ᵐ[μ] s.indicator (μ[s.indicator f|m] + μ[sᶜ.indicator f|m]) := by - have : μ[s.indicator f + sᶜ.indicator f|m] =ᵐ[μ] μ[s.indicator f|m] + μ[sᶜ.indicator f|m] := - condexp_add (hf_int.indicator (hm _ hs)) (hf_int.indicator (hm _ hs.compl)) - filter_upwards [this] with x hx + filter_upwards [condExp_add (hf_int.indicator (hm _ hs)) (hf_int.indicator (hm _ hs.compl)) m] + with x hx classical rw [Set.indicator_apply, Set.indicator_apply, hx] _ = s.indicator (μ[s.indicator f|m]) + s.indicator (μ[sᶜ.indicator f|m]) := (s.indicator_add' _ _) @@ -93,7 +97,7 @@ theorem condexp_indicator (hf_int : Integrable f μ) (hs : MeasurableSet[m] s) : s.indicator (sᶜ.indicator (μ[sᶜ.indicator f|m])) := by refine Filter.EventuallyEq.rfl.add ?_ have : sᶜ.indicator (μ[sᶜ.indicator f|m]) =ᵐ[μ] μ[sᶜ.indicator f|m] := by - refine (condexp_indicator_aux hs.compl ?_).symm.trans ?_ + refine (condExp_indicator_aux hs.compl ?_).symm.trans ?_ · exact indicator_ae_eq_restrict_compl (hm _ hs.compl) · rw [Set.indicator_indicator, Set.inter_self] filter_upwards [this] with x hx @@ -103,40 +107,44 @@ theorem condexp_indicator (hf_int : Integrable f μ) (hs : MeasurableSet[m] s) : _ =ᵐ[μ] s.indicator (μ[s.indicator f|m]) := by rw [Set.indicator_indicator, Set.inter_compl_self, Set.indicator_empty', add_zero] _ =ᵐ[μ] μ[s.indicator f|m] := by - refine (condexp_indicator_aux hs ?_).symm.trans ?_ + refine (condExp_indicator_aux hs ?_).symm.trans ?_ · exact indicator_ae_eq_restrict_compl (hm _ hs) · rw [Set.indicator_indicator, Set.inter_self] -theorem condexp_restrict_ae_eq_restrict (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] +@[deprecated (since := "2025-01-21")] alias condexp_indicator := condExp_indicator + +theorem condExp_restrict_ae_eq_restrict (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hs_m : MeasurableSet[m] s) (hf_int : Integrable f μ) : (μ.restrict s)[f|m] =ᵐ[μ.restrict s] μ[f|m] := by have : SigmaFinite ((μ.restrict s).trim hm) := by rw [← restrict_trim hm _ hs_m]; infer_instance rw [ae_eq_restrict_iff_indicator_ae_eq (hm _ hs_m)] - refine EventuallyEq.trans ?_ (condexp_indicator hf_int hs_m) - refine ae_eq_condexp_of_forall_setIntegral_eq hm (hf_int.indicator (hm _ hs_m)) ?_ ?_ ?_ + refine EventuallyEq.trans ?_ (condExp_indicator hf_int hs_m) + refine ae_eq_condExp_of_forall_setIntegral_eq hm (hf_int.indicator (hm _ hs_m)) ?_ ?_ ?_ · intro t ht _ rw [← integrable_indicator_iff (hm _ ht), Set.indicator_indicator, Set.inter_comm, ← Set.indicator_indicator] suffices h_int_restrict : Integrable (t.indicator ((μ.restrict s)[f|m])) (μ.restrict s) by rw [integrable_indicator_iff (hm _ hs_m), IntegrableOn] - rw [integrable_indicator_iff (hm _ ht), IntegrableOn] at h_int_restrict ⊢ exact h_int_restrict - exact integrable_condexp.indicator (hm _ ht) + exact integrable_condExp.indicator (hm _ ht) · intro t ht _ calc ∫ x in t, s.indicator ((μ.restrict s)[f|m]) x ∂μ = ∫ x in t, ((μ.restrict s)[f|m]) x ∂μ.restrict s := by rw [integral_indicator (hm _ hs_m), Measure.restrict_restrict (hm _ hs_m), Measure.restrict_restrict (hm _ ht), Set.inter_comm] - _ = ∫ x in t, f x ∂μ.restrict s := setIntegral_condexp hm hf_int.integrableOn ht + _ = ∫ x in t, f x ∂μ.restrict s := setIntegral_condExp hm hf_int.integrableOn ht _ = ∫ x in t, s.indicator f x ∂μ := by rw [integral_indicator (hm _ hs_m), Measure.restrict_restrict (hm _ hs_m), Measure.restrict_restrict (hm _ ht), Set.inter_comm] - · exact (stronglyMeasurable_condexp.indicator hs_m).aeStronglyMeasurable' + · exact (stronglyMeasurable_condExp.indicator hs_m).aestronglyMeasurable + +@[deprecated (since := "2025-01-21")] +alias condexp_restrict_ae_eq_restrict := condExp_restrict_ae_eq_restrict /-- If the restriction to an `m`-measurable set `s` of a σ-algebra `m` is equal to the restriction to `s` of another σ-algebra `m₂` (hypothesis `hs`), then `μ[f | m] =ᵐ[μ.restrict s] μ[f | m₂]`. -/ -theorem condexp_ae_eq_restrict_of_measurableSpace_eq_on {m m₂ m0 : MeasurableSpace α} +theorem condExp_ae_eq_restrict_of_measurableSpace_eq_on {m m₂ m0 : MeasurableSpace α} {μ : Measure α} (hm : m ≤ m0) (hm₂ : m₂ ≤ m0) [SigmaFinite (μ.trim hm)] [SigmaFinite (μ.trim hm₂)] (hs_m : MeasurableSet[m] s) (hs : ∀ t, MeasurableSet[m] (s ∩ t) ↔ MeasurableSet[m₂] (s ∩ t)) : @@ -144,20 +152,19 @@ theorem condexp_ae_eq_restrict_of_measurableSpace_eq_on {m m₂ m0 : MeasurableS rw [ae_eq_restrict_iff_indicator_ae_eq (hm _ hs_m)] have hs_m₂ : MeasurableSet[m₂] s := by rwa [← Set.inter_univ s, ← hs Set.univ, Set.inter_univ] by_cases hf_int : Integrable f μ - swap; · simp_rw [condexp_undef hf_int]; rfl - refine ((condexp_indicator hf_int hs_m).symm.trans ?_).trans (condexp_indicator hf_int hs_m₂) + swap; · simp_rw [condExp_of_not_integrable hf_int]; rfl + refine ((condExp_indicator hf_int hs_m).symm.trans ?_).trans (condExp_indicator hf_int hs_m₂) refine ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' hm₂ - (fun s _ _ => integrable_condexp.integrableOn) - (fun s _ _ => integrable_condexp.integrableOn) ?_ ?_ - stronglyMeasurable_condexp.aeStronglyMeasurable' + (fun s _ _ => integrable_condExp.integrableOn) + (fun s _ _ => integrable_condExp.integrableOn) ?_ ?_ + stronglyMeasurable_condExp.aestronglyMeasurable swap - · have : StronglyMeasurable[m] (μ[s.indicator f|m]) := stronglyMeasurable_condexp - refine this.aeStronglyMeasurable'.aeStronglyMeasurable'_of_measurableSpace_le_on hm hs_m - (fun t => (hs t).mp) ?_ - exact condexp_ae_eq_restrict_zero hs_m.compl (indicator_ae_eq_restrict_compl (hm _ hs_m)) + · have : StronglyMeasurable[m] (μ[s.indicator f|m]) := stronglyMeasurable_condExp + refine this.aestronglyMeasurable.of_measurableSpace_le_on hm hs_m (fun t => (hs t).mp) ?_ + exact condExp_ae_eq_restrict_zero hs_m.compl (indicator_ae_eq_restrict_compl (hm _ hs_m)) intro t ht _ have : ∫ x in t, (μ[s.indicator f|m]) x ∂μ = ∫ x in s ∩ t, (μ[s.indicator f|m]) x ∂μ := by - rw [← integral_add_compl (hm _ hs_m) integrable_condexp.integrableOn] + rw [← integral_add_compl (hm _ hs_m) integrable_condExp.integrableOn] suffices ∫ x in sᶜ, (μ[s.indicator f|m]) x ∂μ.restrict t = 0 by rw [this, add_zero, Measure.restrict_restrict (hm _ hs_m)] rw [Measure.restrict_restrict (MeasurableSet.compl (hm _ hs_m))] @@ -169,11 +176,15 @@ theorem condexp_ae_eq_restrict_of_measurableSpace_eq_on {m m₂ m0 : MeasurableS refine setIntegral_congr_ae (hm₂ _ ht) ?_ filter_upwards [this] with x hx _ using hx _ = 0 := integral_zero _ _ - refine condexp_ae_eq_restrict_zero hs_m.compl ?_ + refine condExp_ae_eq_restrict_zero hs_m.compl ?_ exact indicator_ae_eq_restrict_compl (hm _ hs_m) have hst_m : MeasurableSet[m] (s ∩ t) := (hs _).mpr (hs_m₂.inter ht) - simp_rw [this, setIntegral_condexp hm₂ (hf_int.indicator (hm _ hs_m)) ht, - setIntegral_condexp hm (hf_int.indicator (hm _ hs_m)) hst_m, integral_indicator (hm _ hs_m), + simp_rw [this, setIntegral_condExp hm₂ (hf_int.indicator (hm _ hs_m)) ht, + setIntegral_condExp hm (hf_int.indicator (hm _ hs_m)) hst_m, integral_indicator (hm _ hs_m), Measure.restrict_restrict (hm _ hs_m), ← Set.inter_assoc, Set.inter_self] +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_restrict_of_measurableSpace_eq_on := + condExp_ae_eq_restrict_of_measurableSpace_eq_on + end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index ffbdbbdff17f6..41c370cc5b4b6 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -15,11 +15,11 @@ This file proves some results regarding the conditional expectation of real-valu ## Main results -* `MeasureTheory.rnDeriv_ae_eq_condexp`: the conditional expectation `μ[f | m]` is equal to the +* `MeasureTheory.rnDeriv_ae_eq_condExp`: the conditional expectation `μ[f | m]` is equal to the Radon-Nikodym derivative of `fμ` restricted on `m` with respect to `μ` restricted on `m`. -* `MeasureTheory.Integrable.uniformIntegrable_condexp`: the conditional expectation of a function +* `MeasureTheory.Integrable.uniformIntegrable_condExp`: the conditional expectation of a function form a uniformly integrable class. -* `MeasureTheory.condexp_stronglyMeasurable_mul`: the pull-out property of the conditional +* `MeasureTheory.condExp_mul_of_stronglyMeasurable_left`: the pull-out property of the conditional expectation. -/ @@ -35,10 +35,10 @@ namespace MeasureTheory variable {α : Type*} {m m0 : MeasurableSpace α} {μ : Measure α} -theorem rnDeriv_ae_eq_condexp {hm : m ≤ m0} [hμm : SigmaFinite (μ.trim hm)] {f : α → ℝ} +theorem rnDeriv_ae_eq_condExp {hm : m ≤ m0} [hμm : SigmaFinite (μ.trim hm)] {f : α → ℝ} (hf : Integrable f μ) : SignedMeasure.rnDeriv ((μ.withDensityᵥ f).trim hm) (μ.trim hm) =ᵐ[μ] μ[f|m] := by - refine ae_eq_condexp_of_forall_setIntegral_eq hm hf ?_ ?_ ?_ + refine ae_eq_condExp_of_forall_setIntegral_eq hm hf ?_ ?_ ?_ · exact fun _ _ _ => (integrable_of_integrable_trim hm (SignedMeasure.integrable_rnDeriv ((μ.withDensityᵥ f).trim hm) (μ.trim hm))).integrableOn · intro s hs _ @@ -49,106 +49,119 @@ theorem rnDeriv_ae_eq_condexp {hm : m ≤ m0} [hμm : SigmaFinite (μ.trim hm)] (SignedMeasure.integrable_rnDeriv ((μ.withDensityᵥ f).trim hm) (μ.trim hm)) hs, ← setIntegral_trim hm _ hs] exact (SignedMeasure.measurable_rnDeriv _ _).stronglyMeasurable - · exact (SignedMeasure.measurable_rnDeriv _ _).stronglyMeasurable.aeStronglyMeasurable' + · exact (SignedMeasure.measurable_rnDeriv _ _).stronglyMeasurable.aestronglyMeasurable + +@[deprecated (since := "2025-01-21")] alias rnDeriv_ae_eq_condexp := rnDeriv_ae_eq_condExp -- TODO: the following couple of lemmas should be generalized and proved using Jensen's inequality -- for the conditional expectation (not in mathlib yet) . -theorem eLpNorm_one_condexp_le_eLpNorm (f : α → ℝ) : eLpNorm (μ[f|m]) 1 μ ≤ eLpNorm f 1 μ := by +theorem eLpNorm_one_condExp_le_eLpNorm (f : α → ℝ) : eLpNorm (μ[f|m]) 1 μ ≤ eLpNorm f 1 μ := by by_cases hf : Integrable f μ - swap; · rw [condexp_undef hf, eLpNorm_zero]; exact zero_le _ + swap; · rw [condExp_of_not_integrable hf, eLpNorm_zero]; exact zero_le _ by_cases hm : m ≤ m0 - swap; · rw [condexp_of_not_le hm, eLpNorm_zero]; exact zero_le _ + swap; · rw [condExp_of_not_le hm, eLpNorm_zero]; exact zero_le _ by_cases hsig : SigmaFinite (μ.trim hm) - swap; · rw [condexp_of_not_sigmaFinite hm hsig, eLpNorm_zero]; exact zero_le _ + swap; · rw [condExp_of_not_sigmaFinite hm hsig, eLpNorm_zero]; exact zero_le _ calc eLpNorm (μ[f|m]) 1 μ ≤ eLpNorm (μ[(|f|)|m]) 1 μ := by refine eLpNorm_mono_ae ?_ - filter_upwards [condexp_mono hf hf.abs + filter_upwards [condExp_mono hf hf.abs (ae_of_all μ (fun x => le_abs_self (f x) : ∀ x, f x ≤ |f x|)), - EventuallyLE.trans (condexp_neg f).symm.le - (condexp_mono hf.neg hf.abs + (condExp_neg ..).symm.le.trans (condExp_mono hf.neg hf.abs (ae_of_all μ (fun x => neg_le_abs (f x) : ∀ x, -f x ≤ |f x|)))] with x hx₁ hx₂ exact abs_le_abs hx₁ hx₂ _ = eLpNorm f 1 μ := by rw [eLpNorm_one_eq_lintegral_nnnorm, eLpNorm_one_eq_lintegral_nnnorm, - ← ENNReal.toReal_eq_toReal (hasFiniteIntegral_iff_nnnorm.mp integrable_condexp.2).ne + ← ENNReal.toReal_eq_toReal (hasFiniteIntegral_iff_nnnorm.mp integrable_condExp.2).ne (hasFiniteIntegral_iff_nnnorm.mp hf.2).ne, ← integral_norm_eq_lintegral_nnnorm - (stronglyMeasurable_condexp.mono hm).aestronglyMeasurable, + (stronglyMeasurable_condExp.mono hm).aestronglyMeasurable, ← integral_norm_eq_lintegral_nnnorm hf.1] simp_rw [Real.norm_eq_abs] - rw (config := {occs := .pos [2]}) [← integral_condexp hm] + rw (config := {occs := .pos [2]}) [← integral_condExp hm] refine integral_congr_ae ?_ have : 0 ≤ᵐ[μ] μ[(|f|)|m] := by - rw [← condexp_zero] - exact condexp_mono (integrable_zero _ _ _) hf.abs + rw [← condExp_zero] + exact condExp_mono (integrable_zero _ _ _) hf.abs (ae_of_all μ (fun x => abs_nonneg (f x) : ∀ x, 0 ≤ |f x|)) filter_upwards [this] with x hx exact abs_eq_self.2 hx +@[deprecated (since := "2025-01-21")] +alias eLpNorm_one_condexp_le_eLpNorm := eLpNorm_one_condExp_le_eLpNorm + @[deprecated (since := "2024-07-27")] -alias snorm_one_condexp_le_snorm := eLpNorm_one_condexp_le_eLpNorm +alias snorm_one_condExp_le_snorm := eLpNorm_one_condExp_le_eLpNorm + +@[deprecated (since := "2025-01-21")] alias snorm_one_condexp_le_snorm := snorm_one_condExp_le_snorm -theorem integral_abs_condexp_le (f : α → ℝ) : ∫ x, |(μ[f|m]) x| ∂μ ≤ ∫ x, |f x| ∂μ := by +theorem integral_abs_condExp_le (f : α → ℝ) : ∫ x, |(μ[f|m]) x| ∂μ ≤ ∫ x, |f x| ∂μ := by by_cases hm : m ≤ m0 swap - · simp_rw [condexp_of_not_le hm, Pi.zero_apply, abs_zero, integral_zero] + · simp_rw [condExp_of_not_le hm, Pi.zero_apply, abs_zero, integral_zero] positivity by_cases hfint : Integrable f μ swap - · simp only [condexp_undef hfint, Pi.zero_apply, abs_zero, integral_const, Algebra.id.smul_eq_mul, - mul_zero] + · simp only [condExp_of_not_integrable hfint, Pi.zero_apply, abs_zero, integral_const, + Algebra.id.smul_eq_mul, mul_zero] positivity rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] · apply ENNReal.toReal_mono <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_coe_nnnorm] · exact hfint.2.ne · rw [← eLpNorm_one_eq_lintegral_nnnorm, ← eLpNorm_one_eq_lintegral_nnnorm] - exact eLpNorm_one_condexp_le_eLpNorm _ + exact eLpNorm_one_condExp_le_eLpNorm _ · filter_upwards with x using abs_nonneg _ · simp_rw [← Real.norm_eq_abs] exact hfint.1.norm · filter_upwards with x using abs_nonneg _ · simp_rw [← Real.norm_eq_abs] - exact (stronglyMeasurable_condexp.mono hm).aestronglyMeasurable.norm + exact (stronglyMeasurable_condExp.mono hm).aestronglyMeasurable.norm -theorem setIntegral_abs_condexp_le {s : Set α} (hs : MeasurableSet[m] s) (f : α → ℝ) : +@[deprecated (since := "2025-01-21")] alias integral_abs_condexp_le := integral_abs_condExp_le + +theorem setIntegral_abs_condExp_le {s : Set α} (hs : MeasurableSet[m] s) (f : α → ℝ) : ∫ x in s, |(μ[f|m]) x| ∂μ ≤ ∫ x in s, |f x| ∂μ := by by_cases hnm : m ≤ m0 swap - · simp_rw [condexp_of_not_le hnm, Pi.zero_apply, abs_zero, integral_zero] + · simp_rw [condExp_of_not_le hnm, Pi.zero_apply, abs_zero, integral_zero] positivity by_cases hfint : Integrable f μ swap - · simp only [condexp_undef hfint, Pi.zero_apply, abs_zero, integral_const, Algebra.id.smul_eq_mul, - mul_zero] + · simp only [condExp_of_not_integrable hfint, Pi.zero_apply, abs_zero, integral_const, + Algebra.id.smul_eq_mul, mul_zero] positivity have : ∫ x in s, |(μ[f|m]) x| ∂μ = ∫ x, |(μ[s.indicator f|m]) x| ∂μ := by rw [← integral_indicator (hnm _ hs)] refine integral_congr_ae ?_ have : (fun x => |(μ[s.indicator f|m]) x|) =ᵐ[μ] fun x => |s.indicator (μ[f|m]) x| := - (condexp_indicator hfint hs).fun_comp abs + (condExp_indicator hfint hs).fun_comp abs refine EventuallyEq.trans (Eventually.of_forall fun x => ?_) this.symm rw [← Real.norm_eq_abs, norm_indicator_eq_indicator_norm] simp only [Real.norm_eq_abs] rw [this, ← integral_indicator (hnm _ hs)] - refine (integral_abs_condexp_le _).trans + refine (integral_abs_condExp_le _).trans (le_of_eq <| integral_congr_ae <| Eventually.of_forall fun x => ?_) simp_rw [← Real.norm_eq_abs, norm_indicator_eq_indicator_norm] +@[deprecated (since := "2025-01-21")] alias setIntegral_abs_condexp_le := setIntegral_abs_condExp_le + @[deprecated (since := "2024-04-17")] -alias set_integral_abs_condexp_le := setIntegral_abs_condexp_le +alias set_integral_abs_condExp_le := setIntegral_abs_condExp_le + +@[deprecated (since := "2025-01-21")] +alias set_integral_abs_condexp_le := set_integral_abs_condExp_le /-- If the real valued function `f` is bounded almost everywhere by `R`, then so is its conditional expectation. -/ -theorem ae_bdd_condexp_of_ae_bdd {R : ℝ≥0} {f : α → ℝ} (hbdd : ∀ᵐ x ∂μ, |f x| ≤ R) : +theorem ae_bdd_condExp_of_ae_bdd {R : ℝ≥0} {f : α → ℝ} (hbdd : ∀ᵐ x ∂μ, |f x| ≤ R) : ∀ᵐ x ∂μ, |(μ[f|m]) x| ≤ R := by by_cases hnm : m ≤ m0 swap - · simp_rw [condexp_of_not_le hnm, Pi.zero_apply, abs_zero] + · simp_rw [condExp_of_not_le hnm, Pi.zero_apply, abs_zero] exact Eventually.of_forall fun _ => R.coe_nonneg by_cases hfint : Integrable f μ swap - · simp_rw [condexp_undef hfint] + · simp_rw [condExp_of_not_integrable hfint] filter_upwards [hbdd] with x hx rw [Pi.zero_apply, abs_zero] exact (abs_nonneg _).trans hx @@ -158,46 +171,48 @@ theorem ae_bdd_condexp_of_ae_bdd {R : ℝ≥0} {f : α → ℝ} (hbdd : ∀ᵐ x suffices (μ {x | ↑R < |(μ[f|m]) x|}).toReal * ↑R < (μ {x | ↑R < |(μ[f|m]) x|}).toReal * ↑R by exact this.ne rfl refine lt_of_lt_of_le (setIntegral_gt_gt R.coe_nonneg ?_ h.ne') ?_ - · exact integrable_condexp.abs.integrableOn - refine (setIntegral_abs_condexp_le ?_ _).trans ?_ + · exact integrable_condExp.abs.integrableOn + refine (setIntegral_abs_condExp_le ?_ _).trans ?_ · simp_rw [← Real.norm_eq_abs] exact @measurableSet_lt _ _ _ _ _ m _ _ _ _ _ measurable_const - stronglyMeasurable_condexp.norm.measurable + stronglyMeasurable_condExp.norm.measurable simp only [← smul_eq_mul, ← setIntegral_const, NNReal.val_eq_coe, RCLike.ofReal_real_eq_id, _root_.id] refine setIntegral_mono_ae hfint.abs.integrableOn ?_ hbdd refine ⟨aestronglyMeasurable_const, lt_of_le_of_lt ?_ - (integrable_condexp.integrableOn : IntegrableOn (μ[f|m]) {x | ↑R < |(μ[f|m]) x|} μ).2⟩ + (integrable_condExp.integrableOn : IntegrableOn (μ[f|m]) {x | ↑R < |(μ[f|m]) x|} μ).2⟩ refine setLIntegral_mono - (stronglyMeasurable_condexp.mono hnm).measurable.nnnorm.coe_nnreal_ennreal fun x hx => ?_ + (stronglyMeasurable_condExp.mono hnm).measurable.nnnorm.coe_nnreal_ennreal fun x hx => ?_ rw [enorm_eq_nnnorm, enorm_eq_nnnorm, ENNReal.coe_le_coe, Real.nnnorm_of_nonneg R.coe_nonneg] exact Subtype.mk_le_mk.2 (le_of_lt hx) +@[deprecated (since := "2025-01-21")] alias ae_bdd_condexp_of_ae_bdd := ae_bdd_condExp_of_ae_bdd + /-- Given an integrable function `g`, the conditional expectations of `g` with respect to a sequence of sub-σ-algebras is uniformly integrable. -/ -theorem Integrable.uniformIntegrable_condexp {ι : Type*} [IsFiniteMeasure μ] {g : α → ℝ} +theorem Integrable.uniformIntegrable_condExp {ι : Type*} [IsFiniteMeasure μ] {g : α → ℝ} (hint : Integrable g μ) {ℱ : ι → MeasurableSpace α} (hℱ : ∀ i, ℱ i ≤ m0) : UniformIntegrable (fun i => μ[g|ℱ i]) 1 μ := by let A : MeasurableSpace α := m0 have hmeas : ∀ n, ∀ C, MeasurableSet {x | C ≤ ‖(μ[g|ℱ n]) x‖₊} := fun n C => - measurableSet_le measurable_const (stronglyMeasurable_condexp.mono (hℱ n)).measurable.nnnorm + measurableSet_le measurable_const (stronglyMeasurable_condExp.mono (hℱ n)).measurable.nnnorm have hg : Memℒp g 1 μ := memℒp_one_iff_integrable.2 hint refine uniformIntegrable_of le_rfl ENNReal.one_ne_top - (fun n => (stronglyMeasurable_condexp.mono (hℱ n)).aestronglyMeasurable) fun ε hε => ?_ + (fun n => (stronglyMeasurable_condExp.mono (hℱ n)).aestronglyMeasurable) fun ε hε => ?_ by_cases hne : eLpNorm g 1 μ = 0 · rw [eLpNorm_eq_zero_iff hg.1 one_ne_zero] at hne refine ⟨0, fun n => (le_of_eq <| - (eLpNorm_eq_zero_iff ((stronglyMeasurable_condexp.mono (hℱ n)).aestronglyMeasurable.indicator + (eLpNorm_eq_zero_iff ((stronglyMeasurable_condExp.mono (hℱ n)).aestronglyMeasurable.indicator (hmeas n 0)) one_ne_zero).2 ?_).trans (zero_le _)⟩ - filter_upwards [condexp_congr_ae (m := ℱ n) hne] with x hx - simp only [zero_le', Set.setOf_true, Set.indicator_univ, Pi.zero_apply, hx, condexp_zero] + filter_upwards [condExp_congr_ae (m := ℱ n) hne] with x hx + simp only [zero_le', Set.setOf_true, Set.indicator_univ, Pi.zero_apply, hx, condExp_zero] obtain ⟨δ, hδ, h⟩ := hg.eLpNorm_indicator_le le_rfl ENNReal.one_ne_top hε set C : ℝ≥0 := ⟨δ, hδ.le⟩⁻¹ * (eLpNorm g 1 μ).toNNReal with hC have hCpos : 0 < C := mul_pos (inv_pos.2 hδ) (ENNReal.toNNReal_pos hne hg.eLpNorm_lt_top.ne) have : ∀ n, μ {x : α | C ≤ ‖(μ[g|ℱ n]) x‖₊} ≤ ENNReal.ofReal δ := by intro n have := mul_meas_ge_le_pow_eLpNorm' μ one_ne_zero ENNReal.one_ne_top - ((stronglyMeasurable_condexp (m := ℱ n) (μ := μ) (f := g)).mono (hℱ n)).aestronglyMeasurable C + ((stronglyMeasurable_condExp (m := ℱ n) (μ := μ) (f := g)).mono (hℱ n)).aestronglyMeasurable C rw [ENNReal.one_toReal, ENNReal.rpow_one, ENNReal.rpow_one, mul_comm, ← ENNReal.le_div_iff_mul_le (Or.inl (ENNReal.coe_ne_zero.2 hCpos.ne')) (Or.inl ENNReal.coe_lt_top.ne)] at this @@ -208,19 +223,22 @@ theorem Integrable.uniformIntegrable_condexp {ι : Type*} [IsFiniteMeasure μ] { hC, Nonneg.inv_mk, ENNReal.coe_mul, ENNReal.coe_toNNReal hg.eLpNorm_lt_top.ne, ← mul_assoc, ← ENNReal.ofReal_eq_coe_nnreal, ← ENNReal.ofReal_mul hδ.le, mul_inv_cancel₀ hδ.ne', ENNReal.ofReal_one, one_mul] - exact eLpNorm_one_condexp_le_eLpNorm _ + exact eLpNorm_one_condExp_le_eLpNorm _ refine ⟨C, fun n => le_trans ?_ (h {x : α | C ≤ ‖(μ[g|ℱ n]) x‖₊} (hmeas n C) (this n))⟩ have hmeasℱ : MeasurableSet[ℱ n] {x : α | C ≤ ‖(μ[g|ℱ n]) x‖₊} := @measurableSet_le _ _ _ _ _ (ℱ n) _ _ _ _ _ measurable_const - (@Measurable.nnnorm _ _ _ _ _ (ℱ n) _ stronglyMeasurable_condexp.measurable) - rw [← eLpNorm_congr_ae (condexp_indicator hint hmeasℱ)] - exact eLpNorm_one_condexp_le_eLpNorm _ + (@Measurable.nnnorm _ _ _ _ _ (ℱ n) _ stronglyMeasurable_condExp.measurable) + rw [← eLpNorm_congr_ae (condExp_indicator hint hmeasℱ)] + exact eLpNorm_one_condExp_le_eLpNorm _ + +@[deprecated (since := "2025-01-21")] +alias Integrable.uniformIntegrable_condexp := Integrable.uniformIntegrable_condExp section PullOut -- TODO: this section could be generalized beyond multiplication, to any bounded bilinear map. -/-- Auxiliary lemma for `condexp_stronglyMeasurable_mul`. -/ -theorem condexp_stronglyMeasurable_simpleFunc_mul (hm : m ≤ m0) (f : @SimpleFunc α m ℝ) {g : α → ℝ} +/-- Auxiliary lemma for `condexp_mul_of_stronglyMeasurable_left`. -/ +theorem condExp_stronglyMeasurable_simpleFunc_mul (hm : m ≤ m0) (f : @SimpleFunc α m ℝ) {g : α → ℝ} (hg : Integrable g μ) : μ[(f * g : α → ℝ)|m] =ᵐ[μ] f * μ[g|m] := by have : ∀ (s c) (f : α → ℝ), Set.indicator s (Function.const α c) * f = s.indicator (c • f) := by intro s c f @@ -236,19 +254,22 @@ theorem condexp_stronglyMeasurable_simpleFunc_mul (hm : m ≤ m0) (f : @SimpleFu classical simp only [@SimpleFunc.const_zero _ _ m, @SimpleFunc.coe_piecewise _ _ m, @SimpleFunc.coe_const _ _ m, @SimpleFunc.coe_zero _ _ m, Set.piecewise_eq_indicator] rw [this, this] - refine (condexp_indicator (hg.smul c) hs).trans ?_ - filter_upwards [condexp_smul (m := m) (m0 := m0) c g] with x hx + refine (condExp_indicator (hg.smul c) hs).trans ?_ + filter_upwards [condExp_smul c g m] with x hx classical simp_rw [Set.indicator_apply, hx] · have h_add := @SimpleFunc.coe_add _ _ m _ g₁ g₂ calc μ[⇑(g₁ + g₂) * g|m] =ᵐ[μ] μ[(⇑g₁ + ⇑g₂) * g|m] := by - refine condexp_congr_ae (EventuallyEq.mul ?_ EventuallyEq.rfl); rw [h_add] + refine condExp_congr_ae (EventuallyEq.mul ?_ EventuallyEq.rfl); rw [h_add] _ =ᵐ[μ] μ[⇑g₁ * g|m] + μ[⇑g₂ * g|m] := by - rw [add_mul]; exact condexp_add (hg.simpleFunc_mul' hm _) (hg.simpleFunc_mul' hm _) + rw [add_mul]; exact condExp_add (hg.simpleFunc_mul' hm _) (hg.simpleFunc_mul' hm _) _ _ =ᵐ[μ] ⇑g₁ * μ[g|m] + ⇑g₂ * μ[g|m] := EventuallyEq.add h_eq₁ h_eq₂ _ =ᵐ[μ] ⇑(g₁ + g₂) * μ[g|m] := by rw [h_add, add_mul] -theorem condexp_stronglyMeasurable_mul_of_bound (hm : m ≤ m0) [IsFiniteMeasure μ] {f g : α → ℝ} +@[deprecated (since := "2025-01-21")] +alias condexp_stronglyMeasurable_simpleFunc_mul := condExp_stronglyMeasurable_simpleFunc_mul + +theorem condExp_stronglyMeasurable_mul_of_bound (hm : m ≤ m0) [IsFiniteMeasure μ] {f g : α → ℝ} (hf : StronglyMeasurable[m] f) (hg : Integrable g μ) (c : ℝ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : μ[f * g|m] =ᵐ[μ] f * μ[g|m] := by let fs := hf.approxBounded c @@ -262,14 +283,14 @@ theorem condexp_stronglyMeasurable_mul_of_bound (hm : m ≤ m0) [IsFiniteMeasure exact (norm_nonneg _).trans hx have hfs_bound : ∀ n x, ‖fs n x‖ ≤ c := hf.norm_approxBounded_le hc have : μ[f * μ[g|m]|m] = f * μ[g|m] := by - refine condexp_of_stronglyMeasurable hm (hf.mul stronglyMeasurable_condexp) ?_ - exact integrable_condexp.bdd_mul' (hf.mono hm).aestronglyMeasurable hf_bound + refine condExp_of_stronglyMeasurable hm (hf.mul stronglyMeasurable_condExp) ?_ + exact integrable_condExp.bdd_mul' (hf.mono hm).aestronglyMeasurable hf_bound rw [← this] - refine tendsto_condexp_unique (fun n x => fs n x * g x) (fun n x => fs n x * (μ[g|m]) x) (f * g) + refine tendsto_condExp_unique (fun n x => fs n x * g x) (fun n x => fs n x * (μ[g|m]) x) (f * g) (f * μ[g|m]) ?_ ?_ ?_ ?_ (c * ‖g ·‖) ?_ (c * ‖(μ[g|m]) ·‖) ?_ ?_ ?_ ?_ · exact fun n => hg.bdd_mul' ((SimpleFunc.stronglyMeasurable (fs n)).mono hm).aestronglyMeasurable (Eventually.of_forall (hfs_bound n)) - · exact fun n => integrable_condexp.bdd_mul' + · exact fun n => integrable_condExp.bdd_mul' ((SimpleFunc.stronglyMeasurable (fs n)).mono hm).aestronglyMeasurable (Eventually.of_forall (hfs_bound n)) · filter_upwards [hfs_tendsto] with x hx @@ -277,38 +298,44 @@ theorem condexp_stronglyMeasurable_mul_of_bound (hm : m ≤ m0) [IsFiniteMeasure · filter_upwards [hfs_tendsto] with x hx exact hx.mul tendsto_const_nhds · exact hg.norm.const_mul c - · exact integrable_condexp.norm.const_mul c + · fun_prop · refine fun n => Eventually.of_forall fun x => ?_ exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right (hfs_bound n x) (norm_nonneg _)) · refine fun n => Eventually.of_forall fun x => ?_ exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right (hfs_bound n x) (norm_nonneg _)) · intro n simp_rw [← Pi.mul_apply] - refine (condexp_stronglyMeasurable_simpleFunc_mul hm _ hg).trans ?_ - rw [condexp_of_stronglyMeasurable hm - ((SimpleFunc.stronglyMeasurable _).mul stronglyMeasurable_condexp) _] - exact integrable_condexp.bdd_mul' + refine (condExp_stronglyMeasurable_simpleFunc_mul hm _ hg).trans ?_ + rw [condExp_of_stronglyMeasurable hm + ((SimpleFunc.stronglyMeasurable _).mul stronglyMeasurable_condExp) _] + exact integrable_condExp.bdd_mul' ((SimpleFunc.stronglyMeasurable (fs n)).mono hm).aestronglyMeasurable (Eventually.of_forall (hfs_bound n)) -theorem condexp_stronglyMeasurable_mul_of_bound₀ (hm : m ≤ m0) [IsFiniteMeasure μ] {f g : α → ℝ} - (hf : AEStronglyMeasurable' m f μ) (hg : Integrable g μ) (c : ℝ) +@[deprecated (since := "2025-01-21")] +alias condexp_stronglyMeasurable_mul_of_bound := condExp_stronglyMeasurable_mul_of_bound + +theorem condExp_stronglyMeasurable_mul_of_bound₀ (hm : m ≤ m0) [IsFiniteMeasure μ] {f g : α → ℝ} + (hf : AEStronglyMeasurable[m] f μ) (hg : Integrable g μ) (c : ℝ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) : μ[f * g|m] =ᵐ[μ] f * μ[g|m] := by have : μ[f * g|m] =ᵐ[μ] μ[hf.mk f * g|m] := - condexp_congr_ae (EventuallyEq.mul hf.ae_eq_mk EventuallyEq.rfl) + condExp_congr_ae (EventuallyEq.mul hf.ae_eq_mk EventuallyEq.rfl) refine this.trans ?_ have : f * μ[g|m] =ᵐ[μ] hf.mk f * μ[g|m] := EventuallyEq.mul hf.ae_eq_mk EventuallyEq.rfl refine EventuallyEq.trans ?_ this.symm - refine condexp_stronglyMeasurable_mul_of_bound hm hf.stronglyMeasurable_mk hg c ?_ + refine condExp_stronglyMeasurable_mul_of_bound hm hf.stronglyMeasurable_mk hg c ?_ filter_upwards [hf_bound, hf.ae_eq_mk] with x hxc hx_eq rwa [← hx_eq] +@[deprecated (since := "2025-01-21")] +alias condexp_stronglyMeasurable_mul_of_bound₀ := condExp_stronglyMeasurable_mul_of_bound₀ + /-- Pull-out property of the conditional expectation. -/ -theorem condexp_stronglyMeasurable_mul {f g : α → ℝ} (hf : StronglyMeasurable[m] f) +theorem condExp_mul_of_stronglyMeasurable_left {f g : α → ℝ} (hf : StronglyMeasurable[m] f) (hfg : Integrable (f * g) μ) (hg : Integrable g μ) : μ[f * g|m] =ᵐ[μ] f * μ[g|m] := by - by_cases hm : m ≤ m0; swap; · simp_rw [condexp_of_not_le hm]; rw [mul_zero] + by_cases hm : m ≤ m0; swap; · simp_rw [condExp_of_not_le hm]; rw [mul_zero] by_cases hμm : SigmaFinite (μ.trim hm) - swap; · simp_rw [condexp_of_not_sigmaFinite hm hμm]; rw [mul_zero] + swap; · simp_rw [condExp_of_not_sigmaFinite hm hμm]; rw [mul_zero] haveI : SigmaFinite (μ.trim hm) := hμm obtain ⟨sets, sets_prop, h_univ⟩ := hf.exists_spanning_measurableSet_norm_le hm μ simp_rw [forall_and] at sets_prop @@ -322,35 +349,51 @@ theorem condexp_stronglyMeasurable_mul {f g : α → ℝ} (hf : StronglyMeasurab exact hx i hi refine fun n => ae_imp_of_ae_restrict ?_ suffices (μ.restrict (sets n))[f * g|m] =ᵐ[μ.restrict (sets n)] f * (μ.restrict (sets n))[g|m] by - refine (condexp_restrict_ae_eq_restrict hm (h_meas n) hfg).symm.trans ?_ - exact this.trans (EventuallyEq.rfl.mul (condexp_restrict_ae_eq_restrict hm (h_meas n) hg)) + refine (condExp_restrict_ae_eq_restrict hm (h_meas n) hfg).symm.trans ?_ + exact this.trans (EventuallyEq.rfl.mul (condExp_restrict_ae_eq_restrict hm (h_meas n) hg)) suffices (μ.restrict (sets n))[(sets n).indicator f * g|m] =ᵐ[μ.restrict (sets n)] (sets n).indicator f * (μ.restrict (sets n))[g|m] by refine EventuallyEq.trans ?_ (this.trans ?_) · exact - condexp_congr_ae ((indicator_ae_eq_restrict <| hm _ <| h_meas n).symm.mul EventuallyEq.rfl) + condExp_congr_ae ((indicator_ae_eq_restrict <| hm _ <| h_meas n).symm.mul EventuallyEq.rfl) · exact (indicator_ae_eq_restrict <| hm _ <| h_meas n).mul EventuallyEq.rfl have : IsFiniteMeasure (μ.restrict (sets n)) := by constructor rw [Measure.restrict_apply_univ] exact h_finite n - refine condexp_stronglyMeasurable_mul_of_bound hm (hf.indicator (h_meas n)) hg.integrableOn n ?_ + refine condExp_stronglyMeasurable_mul_of_bound hm (hf.indicator (h_meas n)) hg.integrableOn n ?_ filter_upwards with x by_cases hxs : x ∈ sets n · simpa only [hxs, Set.indicator_of_mem] using h_norm n x hxs · simp only [hxs, Set.indicator_of_not_mem, not_false_iff, _root_.norm_zero, Nat.cast_nonneg] +@[deprecated (since := "2025-01-22")] +alias condexp_stronglyMeasurable_mul := condExp_mul_of_stronglyMeasurable_left + +/-- Pull-out property of the conditional expectation. -/ +lemma condExp_mul_of_stronglyMeasurable_right {f g : α → ℝ} (hg : StronglyMeasurable[m] g) + (hfg : Integrable (f * g) μ) (hf : Integrable f μ) : μ[f * g | m] =ᵐ[μ] μ[f | m] * g := by + simpa [mul_comm] using condExp_mul_of_stronglyMeasurable_left hg (mul_comm f g ▸ hfg) hf + /-- Pull-out property of the conditional expectation. -/ -theorem condexp_stronglyMeasurable_mul₀ {f g : α → ℝ} (hf : AEStronglyMeasurable' m f μ) +theorem condExp_mul_of_aestronglyMeasurable_left {f g : α → ℝ} (hf : AEStronglyMeasurable[m] f μ) (hfg : Integrable (f * g) μ) (hg : Integrable g μ) : μ[f * g|m] =ᵐ[μ] f * μ[g|m] := by have : μ[f * g|m] =ᵐ[μ] μ[hf.mk f * g|m] := - condexp_congr_ae (hf.ae_eq_mk.mul EventuallyEq.rfl) + condExp_congr_ae (hf.ae_eq_mk.mul EventuallyEq.rfl) refine this.trans ?_ have : f * μ[g|m] =ᵐ[μ] hf.mk f * μ[g|m] := hf.ae_eq_mk.mul EventuallyEq.rfl - refine (condexp_stronglyMeasurable_mul hf.stronglyMeasurable_mk ?_ hg).trans this.symm + refine (condExp_mul_of_stronglyMeasurable_left hf.stronglyMeasurable_mk ?_ hg).trans this.symm refine (integrable_congr ?_).mp hfg exact hf.ae_eq_mk.mul EventuallyEq.rfl +@[deprecated (since := "2025-01-22")] +alias condexp_stronglyMeasurable_mul₀ := condExp_mul_of_aestronglyMeasurable_left + +/-- Pull-out property of the conditional expectation. -/ +lemma condExp_mul_of_aestronglyMeasurable_right {f g : α → ℝ} (hg : AEStronglyMeasurable[m] g μ) + (hfg : Integrable (f * g) μ) (hf : Integrable f μ) : μ[f * g | m] =ᵐ[μ] μ[f | m] * g := by + simpa [mul_comm] using condExp_mul_of_aestronglyMeasurable_left hg (mul_comm f g ▸ hfg) hf + end PullOut end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean index 99a5b211c1158..033ce021ee4be 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean @@ -74,7 +74,7 @@ theorem Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' (hm : m ≤ m0) (f : Lp E' (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn f s μ) (hf_zero : ∀ s : Set α, MeasurableSet[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = 0) - (hf_meas : AEStronglyMeasurable' m f μ) : f =ᵐ[μ] 0 := by + (hf_meas : AEStronglyMeasurable[m] f μ) : f =ᵐ[μ] 0 := by let f_meas : lpMeas E' 𝕜 m p μ := ⟨f, hf_meas⟩ -- Porting note: `simp only` does not call `rfl` to try to close the goal. See https://github.com/leanprover-community/mathlib4/issues/5025 have hf_f_meas : f =ᵐ[μ] f_meas := by simp only [f_meas, Subtype.coe_mk]; rfl @@ -99,7 +99,7 @@ theorem Lp.ae_eq_of_forall_setIntegral_eq' (hm : m ≤ m0) (f g : Lp E' p μ) (h (hp_ne_top : p ≠ ∞) (hf_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn f s μ) (hg_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn g s μ) (hfg : ∀ s : Set α, MeasurableSet[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) - (hf_meas : AEStronglyMeasurable' m f μ) (hg_meas : AEStronglyMeasurable' m g μ) : + (hf_meas : AEStronglyMeasurable[m] f μ) (hg_meas : AEStronglyMeasurable[m] g μ) : f =ᵐ[μ] g := by suffices h_sub : ⇑(f - g) =ᵐ[μ] 0 by rw [← sub_ae_eq_zero]; exact (Lp.coeFn_sub f g).symm.trans h_sub @@ -112,11 +112,8 @@ theorem Lp.ae_eq_of_forall_setIntegral_eq' (hm : m ≤ m0) (f g : Lp E' p μ) (h intro s hs hμs rw [IntegrableOn, integrable_congr (ae_restrict_of_ae (Lp.coeFn_sub f g))] exact (hf_int_finite s hs hμs).sub (hg_int_finite s hs hμs) - have hfg_meas : AEStronglyMeasurable' m (⇑(f - g)) μ := - AEStronglyMeasurable'.congr (hf_meas.sub hg_meas) (Lp.coeFn_sub f g).symm - exact - Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' 𝕜 hm (f - g) hp_ne_zero hp_ne_top hfg_int hfg' - hfg_meas + exact Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' 𝕜 hm (f - g) hp_ne_zero hp_ne_top hfg_int hfg' + <| (hf_meas.sub hg_meas).congr (Lp.coeFn_sub f g).symm @[deprecated (since := "2024-04-17")] alias Lp.ae_eq_of_forall_set_integral_eq' := Lp.ae_eq_of_forall_setIntegral_eq' @@ -127,7 +124,7 @@ theorem ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' (hm : m ≤ m0) [SigmaFin {f g : α → F'} (hf_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn f s μ) (hg_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn g s μ) (hfg_eq : ∀ s : Set α, MeasurableSet[m] s → μ s < ∞ → ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ) - (hfm : AEStronglyMeasurable' m f μ) (hgm : AEStronglyMeasurable' m g μ) : f =ᵐ[μ] g := by + (hfm : AEStronglyMeasurable[m] f μ) (hgm : AEStronglyMeasurable[m] g μ) : f =ᵐ[μ] g := by rw [← ae_eq_trim_iff_of_aeStronglyMeasurable' hm hfm hgm] have hf_mk_int_finite (s) : MeasurableSet[m] s → μ.trim hm s < ∞ → @IntegrableOn _ _ m _ _ (hfm.mk f) s (μ.trim hm) := by diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 61370ba5ba955..2624f72f454f2 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -150,13 +150,17 @@ theorem hasFiniteIntegral_congr {f g : α → β} (h : f =ᵐ[μ] g) : hasFiniteIntegral_congr' <| h.fun_comp norm theorem hasFiniteIntegral_const_iff {c : β} : - HasFiniteIntegral (fun _ : α => c) μ ↔ c = 0 ∨ μ univ < ∞ := by + HasFiniteIntegral (fun _ : α => c) μ ↔ c = 0 ∨ IsFiniteMeasure μ := by simp [hasFiniteIntegral_iff_nnnorm, lintegral_const, lt_top_iff_ne_top, ENNReal.mul_eq_top, - or_iff_not_imp_left] + or_iff_not_imp_left, isFiniteMeasure_iff] + +lemma hasFiniteIntegral_const_iff_isFiniteMeasure {c : β} (hc : c ≠ 0) : + HasFiniteIntegral (fun _ ↦ c) μ ↔ IsFiniteMeasure μ := by + simp [hasFiniteIntegral_const_iff, hc, isFiniteMeasure_iff] theorem hasFiniteIntegral_const [IsFiniteMeasure μ] (c : β) : HasFiniteIntegral (fun _ : α => c) μ := - hasFiniteIntegral_const_iff.2 (Or.inr <| measure_lt_top _ _) + hasFiniteIntegral_const_iff.2 <| .inr ‹_› theorem HasFiniteIntegral.of_mem_Icc [IsFiniteMeasure μ] (a b : ℝ) {X : α → ℝ} (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) : @@ -405,6 +409,7 @@ variable [TopologicalSpace ε] /-- `Integrable f μ` means that `f` is measurable and that the integral `∫⁻ a, ‖f a‖ ∂μ` is finite. `Integrable f` means `Integrable f volume`. -/ +@[fun_prop] def Integrable {α} {_ : MeasurableSpace α} (f : α → ε) (μ : Measure α := by volume_tac) : Prop := AEStronglyMeasurable f μ ∧ HasFiniteIntegral f μ @@ -449,18 +454,22 @@ theorem Integrable.congr {f g : α → β} (hf : Integrable f μ) (h : f =ᵐ[μ theorem integrable_congr {f g : α → β} (h : f =ᵐ[μ] g) : Integrable f μ ↔ Integrable g μ := ⟨fun hf => hf.congr h, fun hg => hg.congr h.symm⟩ -theorem integrable_const_iff {c : β} : Integrable (fun _ : α => c) μ ↔ c = 0 ∨ μ univ < ∞ := by +lemma integrable_const_iff {c : β} : Integrable (fun _ : α => c) μ ↔ c = 0 ∨ IsFiniteMeasure μ := by have : AEStronglyMeasurable (fun _ : α => c) μ := aestronglyMeasurable_const rw [Integrable, and_iff_right this, hasFiniteIntegral_const_iff] +lemma integrable_const_iff_isFiniteMeasure {c : β} (hc : c ≠ 0) : + Integrable (fun _ ↦ c) μ ↔ IsFiniteMeasure μ := by + simp [integrable_const_iff, hc, isFiniteMeasure_iff] + theorem Integrable.of_mem_Icc [IsFiniteMeasure μ] (a b : ℝ) {X : α → ℝ} (hX : AEMeasurable X μ) (h : ∀ᵐ ω ∂μ, X ω ∈ Set.Icc a b) : Integrable X μ := ⟨hX.aestronglyMeasurable, .of_mem_Icc a b h⟩ -@[simp] +@[simp, fun_prop] theorem integrable_const [IsFiniteMeasure μ] (c : β) : Integrable (fun _ : α => c) μ := - integrable_const_iff.2 <| Or.inr <| measure_lt_top _ _ + integrable_const_iff.2 <| .inr ‹_› @[simp] lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMeasure μ] {f : α → β} : @@ -485,6 +494,12 @@ theorem Memℒp.integrable_norm_rpow' [IsFiniteMeasure μ] {f : α → β} {p : · simp [h_top, integrable_const] exact hf.integrable_norm_rpow h_zero h_top +lemma integrable_norm_rpow_iff {f : α → β} {p : ℝ≥0∞} + (hf : AEStronglyMeasurable f μ) (p_zero : p ≠ 0) (p_top : p ≠ ∞) : + Integrable (fun x : α => ‖f x‖ ^ p.toReal) μ ↔ Memℒp f p μ := by + rw [← memℒp_norm_rpow_iff (q := p) hf p_zero p_top, ← memℒp_one_iff_integrable, + ENNReal.div_self p_zero p_top] + theorem Integrable.mono_measure {f : α → β} (h : Integrable f ν) (hμ : μ ≤ ν) : Integrable f μ := ⟨h.aestronglyMeasurable.mono_measure hμ, h.hasFiniteIntegral.mono_measure hμ⟩ @@ -493,6 +508,7 @@ theorem Integrable.of_measure_le_smul {μ' : Measure α} (c : ℝ≥0∞) (hc : rw [← memℒp_one_iff_integrable] at hf ⊢ exact hf.of_measure_le_smul c hc hμ'_le +@[fun_prop] theorem Integrable.add_measure {f : α → β} (hμ : Integrable f μ) (hν : Integrable f ν) : Integrable f (μ + ν) := by simp_rw [← memℒp_one_iff_integrable] at hμ hν ⊢ @@ -528,6 +544,7 @@ theorem Integrable.smul_measure {f : α → β} (h : Integrable f μ) {c : ℝ rw [← memℒp_one_iff_integrable] at h ⊢ exact h.smul_measure hc +@[fun_prop] theorem Integrable.smul_measure_nnreal {f : α → β} (h : Integrable f μ) {c : ℝ≥0} : Integrable f (c • μ) := by apply h.smul_measure @@ -616,26 +633,35 @@ theorem Integrable.add' {f g : α → β} (hf : Integrable f μ) (hg : Integrabl _ = _ := lintegral_nnnorm_add_left hf.aestronglyMeasurable _ _ < ∞ := add_lt_top.2 ⟨hf.hasFiniteIntegral, hg.hasFiniteIntegral⟩ +@[fun_prop] theorem Integrable.add {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : Integrable (f + g) μ := ⟨hf.aestronglyMeasurable.add hg.aestronglyMeasurable, hf.add' hg⟩ +@[fun_prop] +theorem Integrable.add'' {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : + Integrable (fun x ↦ f x + g x) μ := hf.add hg + +@[fun_prop] theorem integrable_finset_sum' {ι} (s : Finset ι) {f : ι → α → β} (hf : ∀ i ∈ s, Integrable (f i) μ) : Integrable (∑ i ∈ s, f i) μ := Finset.sum_induction f (fun g => Integrable g μ) (fun _ _ => Integrable.add) (integrable_zero _ _ _) hf +@[fun_prop] theorem integrable_finset_sum {ι} (s : Finset ι) {f : ι → α → β} (hf : ∀ i ∈ s, Integrable (f i) μ) : Integrable (fun a => ∑ i ∈ s, f i a) μ := by simpa only [← Finset.sum_apply] using integrable_finset_sum' s hf /-- If `f` is integrable, then so is `-f`. See `Integrable.neg'` for the same statement, but formulated with `x ↦ - f x` instead of `-f`. -/ +@[fun_prop] theorem Integrable.neg {f : α → β} (hf : Integrable f μ) : Integrable (-f) μ := ⟨hf.aestronglyMeasurable.neg, hf.hasFiniteIntegral.neg⟩ /-- If `f` is integrable, then so is `fun x ↦ - f x`. See `Integrable.neg` for the same statement, but formulated with `-f` instead of `fun x ↦ - f x`. -/ +@[fun_prop] theorem Integrable.neg' {f : α → β} (hf : Integrable f μ) : Integrable (fun x ↦ - f x) μ := ⟨hf.aestronglyMeasurable.neg, hf.hasFiniteIntegral.neg⟩ @@ -711,22 +737,31 @@ lemma integrable_const_add_iff [IsFiniteMeasure μ] {f : α → β} {c : β} : Integrable (fun x ↦ c + f x) μ ↔ Integrable f μ := integrable_add_iff_integrable_right (integrable_const _) +@[fun_prop] theorem Integrable.sub {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : Integrable (f - g) μ := by simpa only [sub_eq_add_neg] using hf.add hg.neg +@[fun_prop] +theorem Integrable.sub' {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : + Integrable (fun a ↦ f a - g a) μ := by simpa only [sub_eq_add_neg] using hf.add hg.neg + +@[fun_prop] theorem Integrable.norm {f : α → β} (hf : Integrable f μ) : Integrable (fun a => ‖f a‖) μ := ⟨hf.aestronglyMeasurable.norm, hf.hasFiniteIntegral.norm⟩ +@[fun_prop] theorem Integrable.inf {β} [NormedLatticeAddCommGroup β] {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : Integrable (f ⊓ g) μ := by rw [← memℒp_one_iff_integrable] at hf hg ⊢ exact hf.inf hg +@[fun_prop] theorem Integrable.sup {β} [NormedLatticeAddCommGroup β] {f g : α → β} (hf : Integrable f μ) (hg : Integrable g μ) : Integrable (f ⊔ g) μ := by rw [← memℒp_one_iff_integrable] at hf hg ⊢ exact hf.sup hg +@[fun_prop] theorem Integrable.abs {β} [NormedLatticeAddCommGroup β] {f : α → β} (hf : Integrable f μ) : Integrable (fun a => |f a|) μ := by rw [← memℒp_one_iff_integrable] at hf ⊢ @@ -811,6 +846,7 @@ lemma integrable_of_le_of_le {f g₁ g₂ : α → ℝ} (hf : AEStronglyMeasurab exact max_le_add_of_nonneg (norm_nonneg _) (norm_nonneg _) exact Integrable.mono (h_int₁.norm.add h_int₂.norm) hf h_le_add +@[fun_prop] theorem Integrable.prod_mk {f : α → β} {g : α → γ} (hf : Integrable f μ) (hg : Integrable g μ) : Integrable (fun x => (f x, g x)) μ := ⟨hf.aestronglyMeasurable.prod_mk hg.aestronglyMeasurable, @@ -881,6 +917,7 @@ theorem LipschitzWith.integrable_comp_iff_of_antilipschitz {K K'} {f : α → β Integrable (g ∘ f) μ ↔ Integrable f μ := by simp [← memℒp_one_iff_integrable, hg.memℒp_comp_iff_of_antilipschitz hg' g0] +@[fun_prop] theorem Integrable.real_toNNReal {f : α → ℝ} (hf : Integrable f μ) : Integrable (fun x => ((f x).toNNReal : ℝ)) μ := by refine @@ -1099,11 +1136,13 @@ section PosPart /-! ### Lemmas used for defining the positive part of an `L¹` function -/ +@[fun_prop] theorem Integrable.pos_part {f : α → ℝ} (hf : Integrable f μ) : Integrable (fun a => max (f a) 0) μ := ⟨(hf.aestronglyMeasurable.aemeasurable.max aemeasurable_const).aestronglyMeasurable, hf.hasFiniteIntegral.max_zero⟩ +@[fun_prop] theorem Integrable.neg_part {f : α → ℝ} (hf : Integrable f μ) : Integrable (fun a => max (-f a) 0) μ := hf.neg.pos_part @@ -1114,6 +1153,7 @@ section BoundedSMul variable {𝕜 : Type*} +@[fun_prop] theorem Integrable.smul [NormedAddCommGroup 𝕜] [SMulZeroClass 𝕜 β] [BoundedSMul 𝕜 β] (c : 𝕜) {f : α → β} (hf : Integrable f μ) : Integrable (c • f) μ := ⟨hf.aestronglyMeasurable.const_smul c, hf.hasFiniteIntegral.smul c⟩ @@ -1138,6 +1178,7 @@ theorem Integrable.smul_of_top_left {f : α → β} {φ : α → 𝕜} (hφ : In rw [← memℒp_one_iff_integrable] at hφ ⊢ exact Memℒp.smul_of_top_left hf hφ +@[fun_prop] theorem Integrable.smul_const {f : α → 𝕜} (hf : Integrable f μ) (c : β) : Integrable (fun x => f x • c) μ := hf.smul_of_top_left (memℒp_top_const c) @@ -1163,6 +1204,7 @@ section NormedRing variable {𝕜 : Type*} [NormedRing 𝕜] {f : α → 𝕜} +@[fun_prop] theorem Integrable.const_mul {f : α → 𝕜} (h : Integrable f μ) (c : 𝕜) : Integrable (fun x => c * f x) μ := h.smul c @@ -1171,6 +1213,7 @@ theorem Integrable.const_mul' {f : α → 𝕜} (h : Integrable f μ) (c : 𝕜) Integrable ((fun _ : α => c) * f) μ := Integrable.const_mul h c +@[fun_prop] theorem Integrable.mul_const {f : α → 𝕜} (h : Integrable f μ) (c : 𝕜) : Integrable (fun x => f x * c) μ := h.smul (MulOpposite.op c) @@ -1209,6 +1252,7 @@ section NormedDivisionRing variable {𝕜 : Type*} [NormedDivisionRing 𝕜] {f : α → 𝕜} +@[fun_prop] theorem Integrable.div_const {f : α → 𝕜} (h : Integrable f μ) (c : 𝕜) : Integrable (fun x => f x / c) μ := by simp_rw [div_eq_mul_inv, h.mul_const] @@ -1218,6 +1262,7 @@ section RCLike variable {𝕜 : Type*} [RCLike 𝕜] {f : α → 𝕜} +@[fun_prop] theorem Integrable.ofReal {f : α → ℝ} (hf : Integrable f μ) : Integrable (fun x => (f x : 𝕜)) μ := by rw [← memℒp_one_iff_integrable] at hf ⊢ @@ -1229,10 +1274,12 @@ theorem Integrable.re_im_iff : simp_rw [← memℒp_one_iff_integrable] exact memℒp_re_im_iff +@[fun_prop] theorem Integrable.re (hf : Integrable f μ) : Integrable (fun x => RCLike.re (f x)) μ := by rw [← memℒp_one_iff_integrable] at hf ⊢ exact hf.re +@[fun_prop] theorem Integrable.im (hf : Integrable f μ) : Integrable (fun x => RCLike.im (f x)) μ := by rw [← memℒp_one_iff_integrable] at hf ⊢ exact hf.im @@ -1270,7 +1317,7 @@ theorem integrable_of_forall_fin_meas_le' {μ : Measure α} (hm : m ≤ m0) [Sig ⟨hf_meas, (lintegral_le_of_forall_fin_meas_trim_le hm C hf).trans_lt hC⟩ theorem integrable_of_forall_fin_meas_le [SigmaFinite μ] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} - (hf_meas : AEStronglyMeasurable f μ) + (hf_meas : AEStronglyMeasurable[m] f μ) (hf : ∀ s : Set α, MeasurableSet[m] s → μ s ≠ ∞ → (∫⁻ x in s, ‖f x‖₊ ∂μ) ≤ C) : Integrable f μ := have : SigmaFinite (μ.trim le_rfl) := by rwa [@trim_eq_self _ m] @@ -1503,6 +1550,7 @@ open MeasureTheory variable {E : Type*} [NormedAddCommGroup E] {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedSpace 𝕜 E] {H : Type*} [NormedAddCommGroup H] [NormedSpace 𝕜 H] +@[fun_prop] theorem ContinuousLinearMap.integrable_comp {φ : α → H} (L : H →L[𝕜] E) (φ_int : Integrable φ μ) : Integrable (fun a : α => L (φ a)) μ := ((Integrable.norm φ_int).const_mul ‖L‖).mono' @@ -1531,9 +1579,11 @@ namespace MeasureTheory variable {E F : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [NormedAddCommGroup F] [NormedSpace ℝ F] +@[fun_prop] lemma Integrable.fst {f : α → E × F} (hf : Integrable f μ) : Integrable (fun x ↦ (f x).1) μ := (ContinuousLinearMap.fst ℝ E F).integrable_comp hf +@[fun_prop] lemma Integrable.snd {f : α → E × F} (hf : Integrable f μ) : Integrable (fun x ↦ (f x).2) μ := (ContinuousLinearMap.snd ℝ E F).integrable_comp hf diff --git a/Mathlib/MeasureTheory/Function/L2Space.lean b/Mathlib/MeasureTheory/Function/L2Space.lean index 6877529dd2328..3594c9e9b9d38 100644 --- a/Mathlib/MeasureTheory/Function/L2Space.lean +++ b/Mathlib/MeasureTheory/Function/L2Space.lean @@ -70,10 +70,12 @@ theorem Memℒp.inner_const {f : α → E} (hf : Memℒp f p μ) (c : E) : Mem variable {f : α → E} +@[fun_prop] theorem Integrable.const_inner (c : E) (hf : Integrable f μ) : Integrable (fun x => ⟪c, f x⟫) μ := by rw [← memℒp_one_iff_integrable] at hf ⊢; exact hf.const_inner c +@[fun_prop] theorem Integrable.inner_const (hf : Integrable f μ) (c : E) : Integrable (fun x => ⟪f x, c⟫) μ := by rw [← memℒp_one_iff_integrable] at hf ⊢; exact hf.inner_const c diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 783a4b1169548..a2019826e26f5 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -265,12 +265,13 @@ theorem eLpNorm_zero' : eLpNorm (fun _ : α => (0 : F)) p μ = 0 := by convert e @[deprecated (since := "2024-07-27")] alias snorm_zero' := eLpNorm_zero' -theorem zero_memℒp : Memℒp (0 : α → E) p μ := - ⟨aestronglyMeasurable_zero, by - rw [eLpNorm_zero] - exact ENNReal.coe_lt_top⟩ +@[simp] lemma Memℒp.zero : Memℒp (0 : α → E) p μ := + ⟨aestronglyMeasurable_zero, by rw [eLpNorm_zero]; exact ENNReal.coe_lt_top⟩ -theorem zero_mem_ℒp' : Memℒp (fun _ : α => (0 : E)) p μ := zero_memℒp (E := E) +@[simp] lemma Memℒp.zero' : Memℒp (fun _ : α => (0 : E)) p μ := Memℒp.zero + +@[deprecated (since := "2025-01-21")] alias zero_memℒp := Memℒp.zero +@[deprecated (since := "2025-01-21")] alias zero_mem_ℒp := Memℒp.zero' variable [MeasurableSpace α] @@ -886,7 +887,7 @@ lemma memℒp_indicator_const (p : ℝ≥0∞) (hs : MeasurableSet s) (c : E) (h Memℒp (s.indicator fun _ => c) p μ := by rw [memℒp_indicator_iff_restrict hs] obtain rfl | hμ := hμsc - · exact zero_memℒp + · exact Memℒp.zero · have := Fact.mk hμ.lt_top apply memℒp_const diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index 9661035b3d111..649bc7c74f6ed 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -411,4 +411,26 @@ theorem Memℒp.mul_of_top_left' (hf : Memℒp f ∞ μ) (hφ : Memℒp φ p μ) end Mul +section Prod +variable {ι α 𝕜 : Type*} {_ : MeasurableSpace α} [NormedCommRing 𝕜] {μ : Measure α} {f : ι → α → 𝕜} + {p : ι → ℝ≥0∞} {s : Finset ι} + +open Finset in +/-- See `Memℒp.prod'` for the applied version. -/ +protected lemma Memℒp.prod (hf : ∀ i ∈ s, Memℒp (f i) (p i) μ) : + Memℒp (∏ i ∈ s, f i) (∑ i ∈ s, (p i)⁻¹)⁻¹ μ := by + induction s using cons_induction with + | empty => + by_cases hμ : μ = 0 <;> + simp [Memℒp, eLpNormEssSup_const, hμ, aestronglyMeasurable_const, Pi.one_def] + | cons i s hi ih => + rw [prod_cons] + exact (ih <| forall_of_forall_cons hf).mul (hf i <| mem_cons_self ..) (by simp) + +/-- See `Memℒp.prod` for the unapplied version. -/ +protected lemma Memℒp.prod' (hf : ∀ i ∈ s, Memℒp (f i) (p i) μ) : + Memℒp (fun ω ↦ ∏ i ∈ s, f i ω) (∑ i ∈ s, (p i)⁻¹)⁻¹ μ := by + simpa [Finset.prod_fn] using Memℒp.prod hf + +end Prod end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean index a77dc026f492f..bde7165f033fb 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean @@ -191,7 +191,7 @@ theorem memℒp_finset_sum {ι} (s : Finset ι) {f : ι → α → E} (hf : ∀ haveI : DecidableEq ι := Classical.decEq _ revert hf refine Finset.induction_on s ?_ ?_ - · simp only [zero_mem_ℒp', Finset.sum_empty, imp_true_iff] + · simp only [Memℒp.zero', Finset.sum_empty, imp_true_iff] · intro i s his ih hf simp only [his, Finset.sum_insert, not_false_iff] exact (hf i (s.mem_insert_self i)).add (ih fun j hj => hf j (Finset.mem_insert_of_mem hj)) diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean index 3136274c2a507..b6776ff95e87a 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Trim.lean @@ -67,7 +67,7 @@ theorem eLpNorm_trim (hm : m ≤ m0) {f : α → E} (hf : StronglyMeasurable[m] @[deprecated (since := "2024-07-27")] alias snorm_trim := eLpNorm_trim -theorem eLpNorm_trim_ae (hm : m ≤ m0) {f : α → E} (hf : AEStronglyMeasurable f (μ.trim hm)) : +theorem eLpNorm_trim_ae (hm : m ≤ m0) {f : α → E} (hf : AEStronglyMeasurable[m] f (μ.trim hm)) : eLpNorm f p (μ.trim hm) = eLpNorm f p μ := by rw [eLpNorm_congr_ae hf.ae_eq_mk, eLpNorm_congr_ae (ae_eq_of_ae_eq_trim hf.ae_eq_mk)] exact eLpNorm_trim hm hf.stronglyMeasurable_mk diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index ea88132fbc69c..7ff9c55c9cd7d 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -289,7 +289,7 @@ theorem edist_toLp_toLp (f g : α → E) (hf : Memℒp f p μ) (hg : Memℒp g p @[simp] theorem edist_toLp_zero (f : α → E) (hf : Memℒp f p μ) : edist (hf.toLp f) 0 = eLpNorm f p μ := by - convert edist_toLp_toLp f 0 hf zero_memℒp + convert edist_toLp_toLp f 0 hf Memℒp.zero simp @[simp] @@ -322,7 +322,7 @@ theorem norm_eq_zero_iff {f : Lp E p μ} (hp : 0 < p) : ‖f‖ = 0 ↔ f = 0 := NNReal.coe_eq_zero.trans (nnnorm_eq_zero_iff hp) theorem eq_zero_iff_ae_eq_zero {f : Lp E p μ} : f = 0 ↔ f =ᵐ[μ] 0 := by - rw [← (Lp.memℒp f).toLp_eq_toLp_iff zero_memℒp, Memℒp.toLp_zero, toLp_coeFn] + rw [← (Lp.memℒp f).toLp_eq_toLp_iff Memℒp.zero, Memℒp.toLp_zero, toLp_coeFn] @[simp] theorem nnnorm_neg (f : Lp E p μ) : ‖-f‖₊ = ‖f‖₊ := by diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index d64f4c733873a..7ad4405faa9c1 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -175,7 +175,7 @@ alias tendsto_approxOn_range_Lp_snorm := tendsto_approxOn_range_Lp_eLpNorm theorem memℒp_approxOn_range [BorelSpace E] {f : β → E} {μ : Measure β} (fmeas : Measurable f) [SeparableSpace (range f ∪ {0} : Set E)] (hf : Memℒp f p μ) (n : ℕ) : Memℒp (approxOn f fmeas (range f ∪ {0}) 0 (by simp) n) p μ := - memℒp_approxOn fmeas hf (y₀ := 0) (by simp) zero_memℒp n + memℒp_approxOn fmeas hf (y₀ := 0) (by simp) Memℒp.zero n theorem tendsto_approxOn_range_Lp [BorelSpace E] {f : β → E} [hp : Fact (1 ≤ p)] (hp_ne_top : p ≠ ∞) {μ : Measure β} (fmeas : Measurable f) [SeparableSpace (range f ∪ {0} : Set E)] @@ -359,6 +359,7 @@ theorem memℒp_of_isFiniteMeasure (f : α →ₛ E) (p : ℝ≥0∞) (μ : Meas let ⟨C, hfC⟩ := f.exists_forall_norm_le Memℒp.of_bound f.aestronglyMeasurable C <| Eventually.of_forall hfC +@[fun_prop] theorem integrable_of_isFiniteMeasure [IsFiniteMeasure μ] (f : α →ₛ E) : Integrable f μ := memℒp_one_iff_integrable.mp (f.memℒp_of_isFiniteMeasure 1 μ) @@ -496,7 +497,7 @@ theorem toLp_eq_mk (f : α →ₛ E) (hf : Memℒp f p μ) : (toLp f hf : α →ₘ[μ] E) = AEEqFun.mk f f.aestronglyMeasurable := rfl -theorem toLp_zero : toLp (0 : α →ₛ E) zero_memℒp = (0 : Lp.simpleFunc E p μ) := +theorem toLp_zero : toLp (0 : α →ₛ E) Memℒp.zero = (0 : Lp.simpleFunc E p μ) := rfl theorem toLp_add (f g : α →ₛ E) (hf : Memℒp f p μ) (hg : Memℒp g p μ) : @@ -737,12 +738,7 @@ theorem coeFn_le (f g : Lp.simpleFunc G p μ) : (f : α → G) ≤ᵐ[μ] g ↔ instance instAddLeftMono : AddLeftMono (Lp.simpleFunc G p μ) := by refine ⟨fun f g₁ g₂ hg₁₂ => ?_⟩ - rw [← Lp.simpleFunc.coeFn_le] at hg₁₂ ⊢ - have h_add_1 : ((f + g₁ : Lp.simpleFunc G p μ) : α → G) =ᵐ[μ] (f : α → G) + g₁ := Lp.coeFn_add _ _ - have h_add_2 : ((f + g₂ : Lp.simpleFunc G p μ) : α → G) =ᵐ[μ] (f : α → G) + g₂ := Lp.coeFn_add _ _ - filter_upwards [h_add_1, h_add_2, hg₁₂] with _ h1 h2 h3 - rw [h1, h2, Pi.add_apply, Pi.add_apply] - exact add_le_add le_rfl h3 + exact add_le_add_left hg₁₂ f variable (p μ G) @@ -950,6 +946,7 @@ theorem L1.SimpleFunc.toLp_one_eq_toL1 (f : α →ₛ E) (hf : Integrable f μ) (Lp.simpleFunc.toLp f (memℒp_one_iff_integrable.2 hf) : α →₁[μ] E) = hf.toL1 f := rfl +@[fun_prop] protected theorem L1.SimpleFunc.integrable (f : α →₁ₛ[μ] E) : Integrable (Lp.simpleFunc.toSimpleFunc f) μ := by rw [← memℒp_one_iff_integrable]; exact Lp.simpleFunc.memℒp f diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean index 5cc12c4593425..ebd7459e13141 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean @@ -139,27 +139,27 @@ open Real variable {α : Type*} {m : MeasurableSpace α} {f : α → ℝ} (hf : Measurable f) include hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.exp : Measurable fun x => Real.exp (f x) := Real.measurable_exp.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.log : Measurable fun x => log (f x) := measurable_log.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.cos : Measurable fun x ↦ cos (f x) := measurable_cos.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.sin : Measurable fun x ↦ sin (f x) := measurable_sin.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.cosh : Measurable fun x ↦ cosh (f x) := measurable_cosh.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.sinh : Measurable fun x ↦ sinh (f x) := measurable_sinh.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.sqrt : Measurable fun x => √(f x) := continuous_sqrt.measurable.comp hf end RealComposition @@ -208,31 +208,31 @@ open Complex variable {α : Type*} {m : MeasurableSpace α} {f : α → ℂ} (hf : Measurable f) include hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.cexp : Measurable fun x => Complex.exp (f x) := Complex.measurable_exp.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.ccos : Measurable fun x => Complex.cos (f x) := Complex.measurable_cos.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.csin : Measurable fun x => Complex.sin (f x) := Complex.measurable_sin.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.ccosh : Measurable fun x => Complex.cosh (f x) := Complex.measurable_cosh.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.csinh : Measurable fun x => Complex.sinh (f x) := Complex.measurable_sinh.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.carg : Measurable fun x => arg (f x) := measurable_arg.comp hf -@[measurability] +@[measurability, fun_prop] protected theorem Measurable.clog : Measurable fun x => Complex.log (f x) := measurable_log.comp hf @@ -275,6 +275,18 @@ protected lemma AEMeasurable.clog : AEMeasurable (fun x ↦ log (f x)) μ := end ComplexComposition +@[measurability, fun_prop] +protected theorem Measurable.complex_ofReal {α : Type*} {m : MeasurableSpace α} {f : α → ℝ} + (hf : Measurable f) : + Measurable fun x ↦ (f x : ℂ) := + Complex.measurable_ofReal.comp hf + +@[measurability, fun_prop] +protected theorem AEMeasurable.complex_ofReal {α : Type*} {m : MeasurableSpace α} {μ : Measure α} + {f : α → ℝ} (hf : AEMeasurable f μ) : + AEMeasurable (fun x ↦ (f x : ℂ)) μ := + Complex.measurable_ofReal.comp_aemeasurable hf + section PowInstances instance Complex.hasMeasurablePow : MeasurablePow ℂ ℂ := diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 3fcd3f10b57d4..3b54bf1d8f7db 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -85,11 +85,18 @@ def FinStronglyMeasurable [Zero β] ∃ fs : ℕ → α →ₛ β, (∀ n, μ (support (fs n)) < ∞) ∧ ∀ x, Tendsto (fun n => fs n x) atTop (𝓝 (f x)) /-- A function is `AEStronglyMeasurable` with respect to a measure `μ` if it is almost everywhere -equal to the limit of a sequence of simple functions. -/ +equal to the limit of a sequence of simple functions. + +One can specify the sigma-algebra according to which simple functions are taken using the +`AEStronglyMeasurable[m]` notation in the `MeasureTheory` scope. -/ @[fun_prop] -def AEStronglyMeasurable - {_ : MeasurableSpace α} (f : α → β) (μ : Measure α := by volume_tac) : Prop := - ∃ g, StronglyMeasurable g ∧ f =ᵐ[μ] g +def AEStronglyMeasurable [m : MeasurableSpace α] {m₀ : MeasurableSpace α} (f : α → β) + (μ : Measure[m₀] α := by volume_tac) : Prop := + ∃ g : α → β, StronglyMeasurable[m] g ∧ f =ᵐ[μ] g + +/-- A function is `m`-`AEStronglyMeasurable` with respect to a measure `μ` if it is almost +everywhere equal to the limit of a sequence of `m`-simple functions. -/ +scoped notation "AEStronglyMeasurable[" m "]" => @MeasureTheory.AEStronglyMeasurable _ _ _ m /-- A function is `AEFinStronglyMeasurable` with respect to a measure if it is almost everywhere equal to the limit of a sequence of simple functions with support with finite measure. -/ @@ -103,15 +110,11 @@ open MeasureTheory /-! ## Strongly measurable functions -/ -@[aesop 30% apply (rule_sets := [Measurable])] -protected theorem StronglyMeasurable.aestronglyMeasurable {α β} {_ : MeasurableSpace α} - [TopologicalSpace β] {f : α → β} {μ : Measure α} (hf : StronglyMeasurable f) : - AEStronglyMeasurable f μ := - ⟨f, hf, EventuallyEq.refl _ _⟩ +section StronglyMeasurable +variable {_ : MeasurableSpace α} [TopologicalSpace β] {f : α → β} {μ : Measure α} @[simp] -theorem Subsingleton.stronglyMeasurable {α β} [MeasurableSpace α] [TopologicalSpace β] - [Subsingleton β] (f : α → β) : StronglyMeasurable f := by +theorem Subsingleton.stronglyMeasurable [Subsingleton β] (f : α → β) : StronglyMeasurable f := by let f_sf : α →ₛ β := ⟨f, fun x => ?_, Set.Subsingleton.finite Set.subsingleton_of_subsingleton⟩ · exact ⟨fun _ => f_sf, fun x => tendsto_const_nhds⟩ · have h_univ : f ⁻¹' {x} = Set.univ := by @@ -120,29 +123,23 @@ theorem Subsingleton.stronglyMeasurable {α β} [MeasurableSpace α] [Topologica rw [h_univ] exact MeasurableSet.univ -theorem SimpleFunc.stronglyMeasurable {α β} {_ : MeasurableSpace α} [TopologicalSpace β] - (f : α →ₛ β) : StronglyMeasurable f := +theorem SimpleFunc.stronglyMeasurable (f : α →ₛ β) : StronglyMeasurable f := ⟨fun _ => f, fun _ => tendsto_const_nhds⟩ @[nontriviality] -theorem StronglyMeasurable.of_finite [Finite α] {_ : MeasurableSpace α} - [MeasurableSingletonClass α] [TopologicalSpace β] +theorem StronglyMeasurable.of_finite [Finite α] [MeasurableSingletonClass α] {f : α → β} : StronglyMeasurable f := ⟨fun _ => SimpleFunc.ofFinite f, fun _ => tendsto_const_nhds⟩ -theorem stronglyMeasurable_const {α β} {_ : MeasurableSpace α} [TopologicalSpace β] {b : β} : - StronglyMeasurable fun _ : α => b := +theorem stronglyMeasurable_const {b : β} : StronglyMeasurable fun _ : α => b := ⟨fun _ => SimpleFunc.const α b, fun _ => tendsto_const_nhds⟩ @[to_additive] -theorem stronglyMeasurable_one {α β} {_ : MeasurableSpace α} [TopologicalSpace β] [One β] : - StronglyMeasurable (1 : α → β) := - stronglyMeasurable_const +theorem stronglyMeasurable_one [One β] : StronglyMeasurable (1 : α → β) := stronglyMeasurable_const /-- A version of `stronglyMeasurable_const` that assumes `f x = f y` for all `x, y`. This version works for functions between empty types. -/ -theorem stronglyMeasurable_const' {α β} {m : MeasurableSpace α} [TopologicalSpace β] {f : α → β} - (hf : ∀ x y, f x = f y) : StronglyMeasurable f := by +theorem stronglyMeasurable_const' (hf : ∀ x y, f x = f y) : StronglyMeasurable f := by nontriviality α inhabit α convert stronglyMeasurable_const (β := β) using 1 @@ -150,10 +147,11 @@ theorem stronglyMeasurable_const' {α β} {m : MeasurableSpace α} [TopologicalS -- Porting note: changed binding type of `MeasurableSpace α`. @[simp] -theorem Subsingleton.stronglyMeasurable' {α β} [MeasurableSpace α] [TopologicalSpace β] - [Subsingleton α] (f : α → β) : StronglyMeasurable f := +theorem Subsingleton.stronglyMeasurable' [Subsingleton α] (f : α → β) : StronglyMeasurable f := stronglyMeasurable_const' fun x y => by rw [Subsingleton.elim x y] +end StronglyMeasurable + namespace StronglyMeasurable variable {f g : α → β} @@ -823,23 +821,49 @@ protected theorem real_toNNReal {_ : MeasurableSpace α} {f : α → ℝ} (hf : StronglyMeasurable fun x => (f x).toNNReal := continuous_real_toNNReal.comp_stronglyMeasurable hf -theorem measurableSet_eq_fun {m : MeasurableSpace α} {E} [TopologicalSpace E] [MetrizableSpace E] - {f g : α → E} (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) : - MeasurableSet { x | f x = g x } := by +section PseudoMetrizableSpace +variable {E : Type*} {m m₀ : MeasurableSpace α} {μ : Measure[m₀] α} {f g : α → E} + [TopologicalSpace E] [Preorder E] [OrderClosedTopology E] [PseudoMetrizableSpace E] + +lemma measurableSet_le (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) : + MeasurableSet[m] {a | f a ≤ g a} := by + borelize (E × E) + exact (hf.prod_mk hg).measurable isClosed_le_prod.measurableSet + +lemma measurableSet_lt (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) : + MeasurableSet[m] {a | f a < g a} := by + simpa only [lt_iff_le_not_le] using (hf.measurableSet_le hg).inter (hg.measurableSet_le hf).compl + +lemma ae_le_trim_of_stronglyMeasurable (hm : m ≤ m₀) (hf : StronglyMeasurable[m] f) + (hg : StronglyMeasurable[m] g) (hfg : f ≤ᵐ[μ] g) : f ≤ᵐ[μ.trim hm] g := by + rwa [EventuallyLE, ae_iff, trim_measurableSet_eq hm] + exact (hf.measurableSet_le hg).compl + +lemma ae_le_trim_iff (hm : m ≤ m₀) (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) : + f ≤ᵐ[μ.trim hm] g ↔ f ≤ᵐ[μ] g := + ⟨ae_le_of_ae_le_trim, ae_le_trim_of_stronglyMeasurable hm hf hg⟩ + +end PseudoMetrizableSpace + +section MetrizableSpace +variable {E : Type*} {m m₀ : MeasurableSpace α} {μ : Measure[m₀] α} {f g : α → E} + [TopologicalSpace E] [MetrizableSpace E] + +lemma measurableSet_eq_fun (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) : + MeasurableSet[m] {a | f a = g a} := by borelize (E × E) exact (hf.prod_mk hg).measurable isClosed_diagonal.measurableSet -theorem measurableSet_lt {m : MeasurableSpace α} [TopologicalSpace β] [LinearOrder β] - [OrderClosedTopology β] [PseudoMetrizableSpace β] {f g : α → β} (hf : StronglyMeasurable f) - (hg : StronglyMeasurable g) : MeasurableSet { a | f a < g a } := by - borelize (β × β) - exact (hf.prod_mk hg).measurable isOpen_lt_prod.measurableSet +lemma ae_eq_trim_of_stronglyMeasurable (hm : m ≤ m₀) (hf : StronglyMeasurable[m] f) + (hg : StronglyMeasurable[m] g) (hfg : f =ᵐ[μ] g) : f =ᵐ[μ.trim hm] g := by + rwa [EventuallyEq, ae_iff, trim_measurableSet_eq hm] + exact (hf.measurableSet_eq_fun hg).compl -theorem measurableSet_le {m : MeasurableSpace α} [TopologicalSpace β] [Preorder β] - [OrderClosedTopology β] [PseudoMetrizableSpace β] {f g : α → β} (hf : StronglyMeasurable f) - (hg : StronglyMeasurable g) : MeasurableSet { a | f a ≤ g a } := by - borelize (β × β) - exact (hf.prod_mk hg).measurable isClosed_le_prod.measurableSet +lemma ae_eq_trim_iff (hm : m ≤ m₀) (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) : + f =ᵐ[μ.trim hm] g ↔ f =ᵐ[μ] g := + ⟨ae_eq_of_ae_eq_trim, ae_eq_trim_of_stronglyMeasurable hm hf hg⟩ + +end MetrizableSpace theorem stronglyMeasurable_in_set {m : MeasurableSpace α} [TopologicalSpace β] [Zero β] {s : Set α} {f : α → β} (hs : MeasurableSet s) (hf : StronglyMeasurable f) @@ -1075,43 +1099,49 @@ theorem aefinStronglyMeasurable_zero {α β} {_ : MeasurableSpace α} (μ : Meas /-! ## Almost everywhere strongly measurable functions -/ +section AEStronglyMeasurable +variable [TopologicalSpace β] [TopologicalSpace γ] {m m₀ : MeasurableSpace α} {μ ν : Measure[m₀] α} + {f g : α → β} + +@[aesop 30% apply (rule_sets := [Measurable])] +protected theorem StronglyMeasurable.aestronglyMeasurable (hf : StronglyMeasurable[m] f) : + AEStronglyMeasurable[m] f μ := ⟨f, hf, EventuallyEq.refl _ _⟩ + @[measurability] -theorem aestronglyMeasurable_const {α β} {_ : MeasurableSpace α} {μ : Measure α} - [TopologicalSpace β] {b : β} : AEStronglyMeasurable (fun _ : α => b) μ := +theorem aestronglyMeasurable_const {b : β} : AEStronglyMeasurable[m] (fun _ : α => b) μ := stronglyMeasurable_const.aestronglyMeasurable @[to_additive (attr := measurability)] -theorem aestronglyMeasurable_one {α β} {_ : MeasurableSpace α} {μ : Measure α} [TopologicalSpace β] - [One β] : AEStronglyMeasurable (1 : α → β) μ := +theorem aestronglyMeasurable_one [One β] : AEStronglyMeasurable[m] (1 : α → β) μ := stronglyMeasurable_one.aestronglyMeasurable @[simp] -theorem Subsingleton.aestronglyMeasurable {_ : MeasurableSpace α} [TopologicalSpace β] - [Subsingleton β] {μ : Measure α} (f : α → β) : AEStronglyMeasurable f μ := - (Subsingleton.stronglyMeasurable f).aestronglyMeasurable +lemma AEStronglyMeasurable.of_subsingleton_dom [Subsingleton α] : AEStronglyMeasurable[m] f μ := + (Subsingleton.stronglyMeasurable' f).aestronglyMeasurable @[simp] -theorem Subsingleton.aestronglyMeasurable' {_ : MeasurableSpace α} [TopologicalSpace β] - [Subsingleton α] {μ : Measure α} (f : α → β) : AEStronglyMeasurable f μ := - (Subsingleton.stronglyMeasurable' f).aestronglyMeasurable +lemma AEStronglyMeasurable.of_subsingleton_cod [Subsingleton β] : AEStronglyMeasurable[m] f μ := + (Subsingleton.stronglyMeasurable f).aestronglyMeasurable + +theorem Subsingleton.aestronglyMeasurable [Subsingleton β] (f : α → β) : AEStronglyMeasurable f μ := + .of_subsingleton_cod + +lemma Subsingleton.aestronglyMeasurable' [Subsingleton α] (f : α → β) : AEStronglyMeasurable f μ := + .of_subsingleton_dom @[simp] -theorem aestronglyMeasurable_zero_measure [MeasurableSpace α] [TopologicalSpace β] (f : α → β) : - AEStronglyMeasurable f (0 : Measure α) := by +theorem aestronglyMeasurable_zero_measure (f : α → β) : + AEStronglyMeasurable[m] f (0 : Measure[m₀] α) := by nontriviality α inhabit α exact ⟨fun _ => f default, stronglyMeasurable_const, rfl⟩ @[measurability] -theorem SimpleFunc.aestronglyMeasurable {_ : MeasurableSpace α} {μ : Measure α} [TopologicalSpace β] - (f : α →ₛ β) : AEStronglyMeasurable f μ := +theorem SimpleFunc.aestronglyMeasurable (f : α →ₛ β) : AEStronglyMeasurable f μ := f.stronglyMeasurable.aestronglyMeasurable namespace AEStronglyMeasurable -variable {m : MeasurableSpace α} {μ ν : Measure α} [TopologicalSpace β] [TopologicalSpace γ] - {f g : α → β} - lemma of_finite [DiscreteMeasurableSpace α] [Finite α] : AEStronglyMeasurable f μ := ⟨_, .of_finite, ae_eq_rfl⟩ @@ -1119,17 +1149,17 @@ section Mk /-- A `StronglyMeasurable` function such that `f =ᵐ[μ] hf.mk f`. See lemmas `stronglyMeasurable_mk` and `ae_eq_mk`. -/ -protected noncomputable def mk (f : α → β) (hf : AEStronglyMeasurable f μ) : α → β := +protected noncomputable def mk (f : α → β) (hf : AEStronglyMeasurable[m] f μ) : α → β := hf.choose -theorem stronglyMeasurable_mk (hf : AEStronglyMeasurable f μ) : StronglyMeasurable (hf.mk f) := +lemma stronglyMeasurable_mk (hf : AEStronglyMeasurable[m] f μ) : StronglyMeasurable[m] (hf.mk f) := hf.choose_spec.1 theorem measurable_mk [PseudoMetrizableSpace β] [MeasurableSpace β] [BorelSpace β] - (hf : AEStronglyMeasurable f μ) : Measurable (hf.mk f) := + (hf : AEStronglyMeasurable[m] f μ) : Measurable[m] (hf.mk f) := hf.stronglyMeasurable_mk.measurable -theorem ae_eq_mk (hf : AEStronglyMeasurable f μ) : f =ᵐ[μ] hf.mk f := +theorem ae_eq_mk (hf : AEStronglyMeasurable[m] f μ) : f =ᵐ[μ] hf.mk f := hf.choose_spec.2 @[aesop 5% apply (rule_sets := [Measurable])] @@ -1140,37 +1170,43 @@ protected theorem aemeasurable {β} [MeasurableSpace β] [TopologicalSpace β] end Mk -theorem congr (hf : AEStronglyMeasurable f μ) (h : f =ᵐ[μ] g) : AEStronglyMeasurable g μ := +theorem congr (hf : AEStronglyMeasurable[m] f μ) (h : f =ᵐ[μ] g) : AEStronglyMeasurable[m] g μ := ⟨hf.mk f, hf.stronglyMeasurable_mk, h.symm.trans hf.ae_eq_mk⟩ theorem _root_.aestronglyMeasurable_congr (h : f =ᵐ[μ] g) : - AEStronglyMeasurable f μ ↔ AEStronglyMeasurable g μ := + AEStronglyMeasurable[m] f μ ↔ AEStronglyMeasurable[m] g μ := ⟨fun hf => hf.congr h, fun hg => hg.congr h.symm⟩ -theorem mono_measure {ν : Measure α} (hf : AEStronglyMeasurable f μ) (h : ν ≤ μ) : - AEStronglyMeasurable f ν := +theorem mono_measure {ν : Measure α} (hf : AEStronglyMeasurable[m] f μ) (h : ν ≤ μ) : + AEStronglyMeasurable[m] f ν := ⟨hf.mk f, hf.stronglyMeasurable_mk, Eventually.filter_mono (ae_mono h) hf.ae_eq_mk⟩ -protected lemma mono_ac (h : ν ≪ μ) (hμ : AEStronglyMeasurable f μ) : AEStronglyMeasurable f ν := - let ⟨g, hg, hg'⟩ := hμ; ⟨g, hg, h.ae_eq hg'⟩ +protected lemma mono_ac (h : ν ≪ μ) (hμ : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] f ν := let ⟨g, hg, hg'⟩ := hμ; ⟨g, hg, h.ae_eq hg'⟩ - -theorem mono_set {s t} (h : s ⊆ t) (ht : AEStronglyMeasurable f (μ.restrict t)) : - AEStronglyMeasurable f (μ.restrict s) := +theorem mono_set {s t} (h : s ⊆ t) (ht : AEStronglyMeasurable[m] f (μ.restrict t)) : + AEStronglyMeasurable[m] f (μ.restrict s) := ht.mono_measure (restrict_mono h le_rfl) -protected theorem restrict (hfm : AEStronglyMeasurable f μ) {s} : - AEStronglyMeasurable f (μ.restrict s) := +lemma mono {m'} (hm : m ≤ m') (hf : AEStronglyMeasurable[m] f μ) : AEStronglyMeasurable[m'] f μ := + let ⟨f', hf'_meas, hff'⟩ := hf; ⟨f', hf'_meas.mono hm, hff'⟩ + +lemma of_trim {m₀' : MeasurableSpace α} (hm₀ : m₀' ≤ m₀) + (hf : AEStronglyMeasurable[m] f (μ.trim hm₀)) : AEStronglyMeasurable[m] f μ := by + obtain ⟨g, hg_meas, hfg⟩ := hf; exact ⟨g, hg_meas, ae_eq_of_ae_eq_trim hfg⟩ + +protected theorem restrict (hfm : AEStronglyMeasurable[m] f μ) {s} : + AEStronglyMeasurable[m] f (μ.restrict s) := hfm.mono_measure Measure.restrict_le_self -theorem ae_mem_imp_eq_mk {s} (h : AEStronglyMeasurable f (μ.restrict s)) : +theorem ae_mem_imp_eq_mk {s} (h : AEStronglyMeasurable[m] f (μ.restrict s)) : ∀ᵐ x ∂μ, x ∈ s → f x = h.mk f x := ae_imp_of_ae_restrict h.ae_eq_mk /-- The composition of a continuous function and an ae strongly measurable function is ae strongly measurable. -/ theorem _root_.Continuous.comp_aestronglyMeasurable {g : β → γ} {f : α → β} (hg : Continuous g) - (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (fun x => g (f x)) μ := + (hf : AEStronglyMeasurable[m] f μ) : AEStronglyMeasurable[m] (fun x => g (f x)) μ := ⟨_, hg.comp_stronglyMeasurable hf.stronglyMeasurable_mk, EventuallyEq.fun_comp hf.ae_eq_mk g⟩ /-- A continuous function from `α` to `β` is ae strongly measurable when one of the two spaces is @@ -1180,8 +1216,8 @@ theorem _root_.Continuous.aestronglyMeasurable [TopologicalSpace α] [OpensMeasu AEStronglyMeasurable f μ := hf.stronglyMeasurable.aestronglyMeasurable -protected theorem prod_mk {f : α → β} {g : α → γ} (hf : AEStronglyMeasurable f μ) - (hg : AEStronglyMeasurable g μ) : AEStronglyMeasurable (fun x => (f x, g x)) μ := +protected theorem prod_mk {f : α → β} {g : α → γ} (hf : AEStronglyMeasurable[m] f μ) + (hg : AEStronglyMeasurable[m] g μ) : AEStronglyMeasurable[m] (fun x => (f x, g x)) μ := ⟨fun x => (hf.mk f x, hg.mk g x), hf.stronglyMeasurable_mk.prod_mk hg.stronglyMeasurable_mk, hf.ae_eq_mk.prod_mk hg.ae_eq_mk⟩ @@ -1190,81 +1226,102 @@ functions is ae strongly measurable. -/ theorem _root_.Continuous.comp_aestronglyMeasurable₂ {β' : Type*} [TopologicalSpace β'] {g : β → β' → γ} {f : α → β} {f' : α → β'} (hg : Continuous g.uncurry) - (hf : AEStronglyMeasurable f μ) (h'f : AEStronglyMeasurable f' μ) : - AEStronglyMeasurable (fun x => g (f x) (f' x)) μ := + (hf : AEStronglyMeasurable[m] f μ) (h'f : AEStronglyMeasurable[m] f' μ) : + AEStronglyMeasurable[m] (fun x => g (f x) (f' x)) μ := hg.comp_aestronglyMeasurable (hf.prod_mk h'f) /-- In a space with second countable topology, measurable implies ae strongly measurable. -/ @[fun_prop, aesop unsafe 30% apply (rule_sets := [Measurable])] -theorem _root_.Measurable.aestronglyMeasurable {_ : MeasurableSpace α} {μ : Measure α} +theorem _root_.Measurable.aestronglyMeasurable [MeasurableSpace β] [PseudoMetrizableSpace β] [SecondCountableTopology β] - [OpensMeasurableSpace β] (hf : Measurable f) : AEStronglyMeasurable f μ := + [OpensMeasurableSpace β] (hf : Measurable[m] f) : AEStronglyMeasurable[m] f μ := hf.stronglyMeasurable.aestronglyMeasurable +/-- If the restriction to a set `s` of a σ-algebra `m` is included in the restriction to `s` of +another σ-algebra `m₂` (hypothesis `hs`), the set `s` is `m` measurable and a function `f` almost +everywhere supported on `s` is `m`-ae-strongly-measurable, then `f` is also +`m₂`-ae-strongly-measurable. -/ +lemma of_measurableSpace_le_on {m' m₀ : MeasurableSpace α} {μ : Measure[m₀] α} [Zero β] + (hm : m ≤ m₀) {s : Set α} (hs_m : MeasurableSet[m] s) + (hs : ∀ t, MeasurableSet[m] (s ∩ t) → MeasurableSet[m'] (s ∩ t)) + (hf : AEStronglyMeasurable[m] f μ) (hf_zero : f =ᵐ[μ.restrict sᶜ] 0) : + AEStronglyMeasurable[m'] f μ := by + have h_ind_eq : s.indicator (hf.mk f) =ᵐ[μ] f := by + refine Filter.EventuallyEq.trans ?_ <| + indicator_ae_eq_of_restrict_compl_ae_eq_zero (hm _ hs_m) hf_zero + filter_upwards [hf.ae_eq_mk] with x hx + by_cases hxs : x ∈ s + · simp [hxs, hx] + · simp [hxs] + suffices StronglyMeasurable[m'] (s.indicator (hf.mk f)) from + this.aestronglyMeasurable.congr h_ind_eq + exact (hf.stronglyMeasurable_mk.indicator hs_m).stronglyMeasurable_of_measurableSpace_le_on hs_m + hs fun x hxs => Set.indicator_of_not_mem hxs _ + section Arithmetic @[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] -protected theorem mul [Mul β] [ContinuousMul β] (hf : AEStronglyMeasurable f μ) - (hg : AEStronglyMeasurable g μ) : AEStronglyMeasurable (f * g) μ := +protected theorem mul [Mul β] [ContinuousMul β] (hf : AEStronglyMeasurable[m] f μ) + (hg : AEStronglyMeasurable[m] g μ) : AEStronglyMeasurable[m] (f * g) μ := ⟨hf.mk f * hg.mk g, hf.stronglyMeasurable_mk.mul hg.stronglyMeasurable_mk, hf.ae_eq_mk.mul hg.ae_eq_mk⟩ @[to_additive (attr := measurability)] -protected theorem mul_const [Mul β] [ContinuousMul β] (hf : AEStronglyMeasurable f μ) (c : β) : - AEStronglyMeasurable (fun x => f x * c) μ := +protected theorem mul_const [Mul β] [ContinuousMul β] (hf : AEStronglyMeasurable[m] f μ) (c : β) : + AEStronglyMeasurable[m] (fun x => f x * c) μ := hf.mul aestronglyMeasurable_const @[to_additive (attr := measurability)] -protected theorem const_mul [Mul β] [ContinuousMul β] (hf : AEStronglyMeasurable f μ) (c : β) : - AEStronglyMeasurable (fun x => c * f x) μ := +protected theorem const_mul [Mul β] [ContinuousMul β] (hf : AEStronglyMeasurable[m] f μ) (c : β) : + AEStronglyMeasurable[m] (fun x => c * f x) μ := aestronglyMeasurable_const.mul hf @[to_additive (attr := measurability)] -protected theorem inv [Inv β] [ContinuousInv β] (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable f⁻¹ μ := +protected theorem inv [Inv β] [ContinuousInv β] (hf : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] f⁻¹ μ := ⟨(hf.mk f)⁻¹, hf.stronglyMeasurable_mk.inv, hf.ae_eq_mk.inv⟩ @[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] -protected theorem div [Group β] [TopologicalGroup β] (hf : AEStronglyMeasurable f μ) - (hg : AEStronglyMeasurable g μ) : AEStronglyMeasurable (f / g) μ := +protected theorem div [Group β] [TopologicalGroup β] (hf : AEStronglyMeasurable[m] f μ) + (hg : AEStronglyMeasurable[m] g μ) : AEStronglyMeasurable[m] (f / g) μ := ⟨hf.mk f / hg.mk g, hf.stronglyMeasurable_mk.div hg.stronglyMeasurable_mk, hf.ae_eq_mk.div hg.ae_eq_mk⟩ @[to_additive] -theorem mul_iff_right [CommGroup β] [TopologicalGroup β] (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable (f * g) μ ↔ AEStronglyMeasurable g μ := +theorem mul_iff_right [CommGroup β] [TopologicalGroup β] (hf : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] (f * g) μ ↔ AEStronglyMeasurable[m] g μ := ⟨fun h ↦ show g = f * g * f⁻¹ by simp only [mul_inv_cancel_comm] ▸ h.mul hf.inv, fun h ↦ hf.mul h⟩ @[to_additive] -theorem mul_iff_left [CommGroup β] [TopologicalGroup β] (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable (g * f) μ ↔ AEStronglyMeasurable g μ := +theorem mul_iff_left [CommGroup β] [TopologicalGroup β] (hf : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] (g * f) μ ↔ AEStronglyMeasurable[m] g μ := mul_comm g f ▸ AEStronglyMeasurable.mul_iff_right hf @[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable]))] protected theorem smul {𝕜} [TopologicalSpace 𝕜] [SMul 𝕜 β] [ContinuousSMul 𝕜 β] {f : α → 𝕜} - {g : α → β} (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) : - AEStronglyMeasurable (fun x => f x • g x) μ := + {g : α → β} (hf : AEStronglyMeasurable[m] f μ) (hg : AEStronglyMeasurable[m] g μ) : + AEStronglyMeasurable[m] (fun x => f x • g x) μ := continuous_smul.comp_aestronglyMeasurable (hf.prod_mk hg) @[to_additive (attr := aesop safe 20 apply (rule_sets := [Measurable])) const_nsmul] -protected theorem pow [Monoid β] [ContinuousMul β] (hf : AEStronglyMeasurable f μ) (n : ℕ) : - AEStronglyMeasurable (f ^ n) μ := +protected theorem pow [Monoid β] [ContinuousMul β] (hf : AEStronglyMeasurable[m] f μ) (n : ℕ) : + AEStronglyMeasurable[m] (f ^ n) μ := ⟨hf.mk f ^ n, hf.stronglyMeasurable_mk.pow _, hf.ae_eq_mk.pow_const _⟩ @[to_additive (attr := measurability)] protected theorem const_smul {𝕜} [SMul 𝕜 β] [ContinuousConstSMul 𝕜 β] - (hf : AEStronglyMeasurable f μ) (c : 𝕜) : AEStronglyMeasurable (c • f) μ := + (hf : AEStronglyMeasurable[m] f μ) (c : 𝕜) : AEStronglyMeasurable[m] (c • f) μ := ⟨c • hf.mk f, hf.stronglyMeasurable_mk.const_smul c, hf.ae_eq_mk.const_smul c⟩ @[to_additive (attr := measurability)] protected theorem const_smul' {𝕜} [SMul 𝕜 β] [ContinuousConstSMul 𝕜 β] - (hf : AEStronglyMeasurable f μ) (c : 𝕜) : AEStronglyMeasurable (fun x => c • f x) μ := + (hf : AEStronglyMeasurable[m] f μ) (c : 𝕜) : AEStronglyMeasurable[m] (fun x => c • f x) μ := hf.const_smul c @[to_additive (attr := measurability)] protected theorem smul_const {𝕜} [TopologicalSpace 𝕜] [SMul 𝕜 β] [ContinuousSMul 𝕜 β] {f : α → 𝕜} - (hf : AEStronglyMeasurable f μ) (c : β) : AEStronglyMeasurable (fun x => f x • c) μ := + (hf : AEStronglyMeasurable[m] f μ) (c : β) : AEStronglyMeasurable[m] (fun x => f x • c) μ := continuous_smul.comp_aestronglyMeasurable (hf.prod_mk aestronglyMeasurable_const) end Arithmetic @@ -1448,7 +1505,7 @@ theorem nullMeasurableSet_le [Preorder β] [OrderClosedTopology β] [PseudoMetri simp only [hfx, hgx] theorem _root_.aestronglyMeasurable_of_aestronglyMeasurable_trim {α} {m m0 : MeasurableSpace α} - {μ : Measure α} (hm : m ≤ m0) {f : α → β} (hf : AEStronglyMeasurable f (μ.trim hm)) : + {μ : Measure α} (hm : m ≤ m0) {f : α → β} (hf : AEStronglyMeasurable[m] f (μ.trim hm)) : AEStronglyMeasurable f μ := ⟨hf.mk f, StronglyMeasurable.mono hf.stronglyMeasurable_mk hm, ae_eq_of_ae_eq_trim hf.ae_eq_mk⟩ @@ -1669,6 +1726,7 @@ theorem _root_.aestronglyMeasurable_const_smul_iff₀ {c : G₀} (hc : c ≠ 0) end MulAction +end AEStronglyMeasurable end AEStronglyMeasurable /-! ## Almost everywhere finitely strongly measurable functions -/ diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Inner.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Inner.lean index 3cebf19542574..eecceb92a2a2d 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Inner.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Inner.lean @@ -11,8 +11,7 @@ import Mathlib.Analysis.InnerProductSpace.Basic -/ - -variable {α : Type*} +variable {α 𝕜 E : Type*} [RCLike 𝕜] [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] namespace MeasureTheory @@ -21,33 +20,36 @@ namespace MeasureTheory namespace StronglyMeasurable -protected theorem inner {𝕜 : Type*} {E : Type*} [RCLike 𝕜] [NormedAddCommGroup E] - [InnerProductSpace 𝕜 E] {_ : MeasurableSpace α} {f g : α → E} (hf : StronglyMeasurable f) +protected theorem inner {_ : MeasurableSpace α} {f g : α → E} (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) : StronglyMeasurable fun t => @inner 𝕜 _ _ (f t) (g t) := Continuous.comp_stronglyMeasurable continuous_inner (hf.prod_mk hg) end StronglyMeasurable namespace AEStronglyMeasurable - -variable {m : MeasurableSpace α} {μ : Measure α} {𝕜 : Type*} {E : Type*} [RCLike 𝕜] - [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable {m m₀ : MeasurableSpace α} {μ : Measure[m₀] α} {f g : α → E} {c : E} local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y -protected theorem re {f : α → 𝕜} (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable (fun x => RCLike.re (f x)) μ := +protected theorem re {f : α → 𝕜} (hf : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] (fun x => RCLike.re (f x)) μ := RCLike.continuous_re.comp_aestronglyMeasurable hf -protected theorem im {f : α → 𝕜} (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable (fun x => RCLike.im (f x)) μ := +protected theorem im {f : α → 𝕜} (hf : AEStronglyMeasurable[m] f μ) : + AEStronglyMeasurable[m] (fun x => RCLike.im (f x)) μ := RCLike.continuous_im.comp_aestronglyMeasurable hf protected theorem inner {_ : MeasurableSpace α} {μ : Measure α} {f g : α → E} - (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) : - AEStronglyMeasurable (fun x => ⟪f x, g x⟫) μ := + (hf : AEStronglyMeasurable[m] f μ) (hg : AEStronglyMeasurable[m] g μ) : + AEStronglyMeasurable[m] (fun x => ⟪f x, g x⟫) μ := continuous_inner.comp_aestronglyMeasurable (hf.prod_mk hg) +lemma inner_const (hf : AEStronglyMeasurable[m] f μ) : AEStronglyMeasurable[m] (⟪f ·, c⟫) μ := + hf.inner aestronglyMeasurable_const + +lemma const_inner (hg : AEStronglyMeasurable[m] g μ) : AEStronglyMeasurable[m] (⟪c, g ·⟫) μ := + aestronglyMeasurable_const.inner hg + end AEStronglyMeasurable end MeasureTheory diff --git a/Mathlib/MeasureTheory/Group/AddCircle.lean b/Mathlib/MeasureTheory/Group/AddCircle.lean index 3f20417353a69..1361c8e0de0cc 100644 --- a/Mathlib/MeasureTheory/Group/AddCircle.lean +++ b/Mathlib/MeasureTheory/Group/AddCircle.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.MeasureTheory.Integral.Periodic -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientGroup import Mathlib.MeasureTheory.Group.AEStabilizer /-! diff --git a/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean b/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean index 95bd76b8d5572..1cc3a54fa2bda 100644 --- a/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean +++ b/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean @@ -34,6 +34,8 @@ We also deduce that the corresponding maps are measurable embeddings. measurable, equivalence, group action -/ +open scoped Pointwise + namespace MeasurableEquiv variable {G G₀ α : Type*} [MeasurableSpace G] [MeasurableSpace G₀] [MeasurableSpace α] [Group G] @@ -201,3 +203,30 @@ def divLeft [MeasurableMul G] [MeasurableInv G] (g : G) : G ≃ᵐ G where measurable_invFun := measurable_inv.mul_const g end MeasurableEquiv + +namespace MeasureTheory.Measure +variable {G A : Type*} [Group G] [AddCommGroup A] [DistribMulAction G A] [MeasurableSpace A] + -- We only need `MeasurableConstSMul G A` but we don't have this class. So we erroneously must + -- assume `MeasurableSpace G` + `MeasurableSMul G A` + [MeasurableSpace G] [MeasurableSMul G A] +variable {μ ν : Measure A} {g : G} + +noncomputable instance : DistribMulAction Gᵈᵐᵃ (Measure A) where + smul g μ := μ.map (DomMulAct.mk.symm g⁻¹ • ·) + one_smul μ := show μ.map _ = _ by simp + mul_smul g g' μ := show μ.map _ = ((μ.map _).map _) by + rw [map_map] + · simp [Function.comp_def, mul_smul] + · exact measurable_const_smul .. + · exact measurable_const_smul .. + smul_zero g := show (0 : Measure A).map _ = 0 by simp + smul_add g μ ν := show (μ + ν).map _ = μ.map _ + ν.map _ by + rw [Measure.map_add]; exact measurable_const_smul .. + +lemma dmaSMul_apply (μ : Measure A) (g : Gᵈᵐᵃ) (s : Set A) : + (g • μ) s = μ (DomMulAct.mk.symm g • s) := by + refine ((MeasurableEquiv.smul ((DomMulAct.mk.symm g : G)⁻¹)).map_apply _).trans ?_ + congr 1 + exact Set.preimage_smul_inv (DomMulAct.mk.symm g) s + +end MeasureTheory.Measure diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 2720940a760a3..51627a057494c 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1388,16 +1388,14 @@ theorem SimpleFunc.integral_eq_sum (f : α →ₛ E) (hfi : Integrable f μ) : @[simp] theorem integral_const (c : E) : ∫ _ : α, c ∂μ = (μ univ).toReal • c := by - cases' (@le_top _ _ _ (μ univ)).lt_or_eq with hμ hμ - · haveI : IsFiniteMeasure μ := ⟨hμ⟩ - simp only [integral, hE, L1.integral] + by_cases hμ : IsFiniteMeasure μ + · simp only [integral, hE, L1.integral] exact setToFun_const (dominatedFinMeasAdditive_weightedSMul _) _ - · by_cases hc : c = 0 - · simp [hc, integral_zero] - · have : ¬Integrable (fun _ : α => c) μ := by - simp only [integrable_const_iff, not_or] - exact ⟨hc, hμ.not_lt⟩ - simp [integral_undef, *] + by_cases hc : c = 0 + · simp [hc, integral_zero] + · simp [integral_undef, (integrable_const_iff_isFiniteMeasure hc).not.2 hμ, *] + simp only [isFiniteMeasure_iff, not_lt, top_le_iff] at hμ + simp [hμ] theorem norm_integral_le_of_norm_le_const [IsFiniteMeasure μ] {f : α → G} {C : ℝ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ C) : ‖∫ x, f x ∂μ‖ ≤ C * (μ univ).toReal := @@ -1849,34 +1847,11 @@ theorem integral_trim (hm : m ≤ m0) {f : β → G} (hf : StronglyMeasurable[m] (hf_int.trim hm hf) exact tendsto_nhds_unique h_lim_1 h_lim_2 -theorem integral_trim_ae (hm : m ≤ m0) {f : β → G} (hf : AEStronglyMeasurable f (μ.trim hm)) : +theorem integral_trim_ae (hm : m ≤ m0) {f : β → G} (hf : AEStronglyMeasurable[m] f (μ.trim hm)) : ∫ x, f x ∂μ = ∫ x, f x ∂μ.trim hm := by rw [integral_congr_ae (ae_eq_of_ae_eq_trim hf.ae_eq_mk), integral_congr_ae hf.ae_eq_mk] exact integral_trim hm hf.stronglyMeasurable_mk -theorem ae_eq_trim_of_stronglyMeasurable [TopologicalSpace γ] [MetrizableSpace γ] (hm : m ≤ m0) - {f g : β → γ} (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) - (hfg : f =ᵐ[μ] g) : f =ᵐ[μ.trim hm] g := by - rwa [EventuallyEq, ae_iff, trim_measurableSet_eq hm] - exact (hf.measurableSet_eq_fun hg).compl - -theorem ae_eq_trim_iff [TopologicalSpace γ] [MetrizableSpace γ] (hm : m ≤ m0) {f g : β → γ} - (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) : - f =ᵐ[μ.trim hm] g ↔ f =ᵐ[μ] g := - ⟨ae_eq_of_ae_eq_trim, ae_eq_trim_of_stronglyMeasurable hm hf hg⟩ - -theorem ae_le_trim_of_stronglyMeasurable [LinearOrder γ] [TopologicalSpace γ] - [OrderClosedTopology γ] [PseudoMetrizableSpace γ] (hm : m ≤ m0) {f g : β → γ} - (hf : StronglyMeasurable[m] f) (hg : StronglyMeasurable[m] g) (hfg : f ≤ᵐ[μ] g) : - f ≤ᵐ[μ.trim hm] g := by - rwa [EventuallyLE, ae_iff, trim_measurableSet_eq hm] - exact (hf.measurableSet_le hg).compl - -theorem ae_le_trim_iff [LinearOrder γ] [TopologicalSpace γ] [OrderClosedTopology γ] - [PseudoMetrizableSpace γ] (hm : m ≤ m0) {f g : β → γ} (hf : StronglyMeasurable[m] f) - (hg : StronglyMeasurable[m] g) : f ≤ᵐ[μ.trim hm] g ↔ f ≤ᵐ[μ] g := - ⟨ae_le_of_ae_le_trim, ae_le_trim_of_stronglyMeasurable hm hf hg⟩ - end IntegralTrim section SnormBound diff --git a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean index 4fd2b90860e81..346ad13c7ca91 100644 --- a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean @@ -300,10 +300,11 @@ theorem circleIntegrable_sub_inv_iff {c w : ℂ} {R : ℝ} : variable [NormedSpace ℂ E] -/-- Definition for $\oint_{|z-c|=R} f(z)\,dz$. -/ +/-- Definition for $\oint_{|z-c|=R} f(z)\,dz$ -/ def circleIntegral (f : ℂ → E) (c : ℂ) (R : ℝ) : E := ∫ θ : ℝ in (0)..2 * π, deriv (circleMap c R) θ • f (circleMap c R θ) +/-- `∮ z in C(c, R), f z` is the circle integral $\oint_{|z-c|=R} f(z)\,dz$. -/ notation3 "∮ "(...)" in ""C("c", "R")"", "r:(scoped f => circleIntegral f c R) => r theorem circleIntegral_def_Icc (f : ℂ → E) (c : ℂ) (R : ℝ) : diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index 4da7bce1104cf..9093ea80babe2 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -95,7 +95,7 @@ theorem integrableOn_zero : IntegrableOn (fun _ => (0 : E)) s μ := @[simp] theorem integrableOn_const {C : E} : IntegrableOn (fun _ => C) s μ ↔ C = 0 ∨ μ s < ∞ := - integrable_const_iff.trans <| by rw [Measure.restrict_apply_univ] + integrable_const_iff.trans <| by rw [isFiniteMeasure_restrict, lt_top_iff_ne_top] theorem IntegrableOn.mono (h : IntegrableOn f t ν) (hs : s ⊆ t) (hμ : μ ≤ ν) : IntegrableOn f s μ := h.mono_measure <| Measure.restrict_mono hs hμ @@ -166,8 +166,8 @@ theorem integrableOn_singleton_iff {x : α} [MeasurableSingletonClass α] : have : f =ᵐ[μ.restrict {x}] fun _ => f x := by filter_upwards [ae_restrict_mem (measurableSet_singleton x)] with _ ha simp only [mem_singleton_iff.1 ha] - rw [IntegrableOn, integrable_congr this, integrable_const_iff] - simp + rw [IntegrableOn, integrable_congr this, integrable_const_iff, isFiniteMeasure_restrict, + lt_top_iff_ne_top] @[simp] theorem integrableOn_finite_biUnion {s : Set β} (hs : s.Finite) {t : β → Set α} : @@ -240,6 +240,7 @@ theorem IntegrableOn.integrable_indicator (h : IntegrableOn f s μ) (hs : Measur Integrable (indicator s f) μ := (integrable_indicator_iff hs).2 h +@[fun_prop] theorem Integrable.indicator (h : Integrable f μ) (hs : MeasurableSet s) : Integrable (indicator s f) μ := h.integrableOn.integrable_indicator hs @@ -252,9 +253,8 @@ theorem integrable_indicatorConstLp {E} [NormedAddCommGroup E] {p : ℝ≥0∞} (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : E) : Integrable (indicatorConstLp p hs hμs c) μ := by rw [integrable_congr indicatorConstLp_coeFn, integrable_indicator_iff hs, IntegrableOn, - integrable_const_iff, lt_top_iff_ne_top] - right - simpa only [Set.univ_inter, MeasurableSet.univ, Measure.restrict_apply] using hμs + integrable_const_iff, isFiniteMeasure_restrict] + exact .inr hμs /-- If a function is integrable on a set `s` and nonzero there, then the measurable hull of `s` is well behaved: the restriction of the measure to `toMeasurable μ s` coincides with its restriction @@ -376,6 +376,18 @@ protected theorem IntegrableAtFilter.eventually (h : IntegrableAtFilter f l μ) ∀ᶠ s in l.smallSets, IntegrableOn f s μ := Iff.mpr (eventually_smallSets' fun _s _t hst ht => ht.mono_set hst) h +theorem integrableAtFilter_atBot_iff [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 ≥ x2] + [Nonempty α] : + IntegrableAtFilter f atBot μ ↔ ∃ a, IntegrableOn f (Iic a) μ := by + refine ⟨fun ⟨s, hs, hi⟩ ↦ ?_, fun ⟨a, ha⟩ ↦ ⟨Iic a, Iic_mem_atBot a, ha⟩⟩ + obtain ⟨t, ht⟩ := mem_atBot_sets.mp hs + exact ⟨t, hi.mono_set fun _ hx ↦ ht _ hx⟩ + +theorem integrableAtFilter_atTop_iff [Preorder α] [IsDirected α fun (x1 x2 : α) => x1 ≤ x2] + [Nonempty α] : + IntegrableAtFilter f atTop μ ↔ ∃ a, IntegrableOn f (Ici a) μ := + integrableAtFilter_atBot_iff (α := αᵒᵈ) + protected theorem IntegrableAtFilter.add {f g : α → E} (hf : IntegrableAtFilter f l μ) (hg : IntegrableAtFilter g l μ) : IntegrableAtFilter (f + g) l μ := by diff --git a/Mathlib/MeasureTheory/Integral/IntervalAverage.lean b/Mathlib/MeasureTheory/Integral/IntervalAverage.lean index 0ecbef306508c..0f592cd4f7158 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalAverage.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalAverage.lean @@ -31,6 +31,7 @@ open scoped Interval variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] +/-- `⨍ x in a..b, f x` is the average of `f` over the interval `Ι a w.r.t. the Lebesgue measure. -/ notation3 "⨍ "(...)" in "a".."b", "r:60:(scoped f => average (Measure.restrict volume (uIoc a b)) f) => r diff --git a/Mathlib/MeasureTheory/Integral/Periodic.lean b/Mathlib/MeasureTheory/Integral/Periodic.lean index 90abdbcb2baaf..0061e3099ce37 100644 --- a/Mathlib/MeasureTheory/Integral/Periodic.lean +++ b/Mathlib/MeasureTheory/Integral/Periodic.lean @@ -27,6 +27,10 @@ open Set Function MeasureTheory MeasureTheory.Measure TopologicalSpace AddSubgro open scoped MeasureTheory NNReal ENNReal +/-! +## Measures and integrability on ℝ and on the circle +-/ + @[measurability] protected theorem AddCircle.measurable_mk' {a : ℝ} : Measurable (β := AddCircle a) ((↑) : ℝ → AddCircle a) := @@ -221,14 +225,61 @@ protected theorem intervalIntegral_preimage (t : ℝ) (f : UnitAddCircle → E) end UnitAddCircle -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] - +/-! +## Interval integrability of periodic functions +-/ namespace Function namespace Periodic +variable {E : Type*} [NormedAddCommGroup E] + variable {f : ℝ → E} {T : ℝ} +/-- A periodic function is interval integrable over every interval if it is interval integrable +over one period. -/ +theorem intervalIntegrable {t : ℝ} (h₁f : Function.Periodic f T) (hT : 0 < T) + (h₂f : IntervalIntegrable f MeasureTheory.volume t (t + T)) (a₁ a₂ : ℝ) : + IntervalIntegrable f MeasureTheory.volume a₁ a₂ := by + -- Replace [a₁, a₂] by [t - n₁ * T, t + n₂ * T], where n₁ and n₂ are natural numbers + obtain ⟨n₁, hn₁⟩ := exists_nat_ge ((t - min a₁ a₂) / T) + obtain ⟨n₂, hn₂⟩ := exists_nat_ge ((max a₁ a₂ - t) / T) + have : Set.uIcc a₁ a₂ ⊆ Set.uIcc (t - n₁ * T) (t + n₂ * T) := by + rw [Set.uIcc_subset_uIcc_iff_le] + constructor + · calc min (t - n₁ * T) (t + n₂ * T) + _ ≤ (t - n₁ * T) := by apply min_le_left + _ ≤ min a₁ a₂ := by linarith [(div_le_iff₀ hT).1 hn₁] + · calc max a₁ a₂ + _ ≤ t + n₂ * T := by linarith [(div_le_iff₀ hT).1 hn₂] + _ ≤ max (t - n₁ * T) (t + n₂ * T) := by apply le_max_right + apply IntervalIntegrable.mono_set _ this + -- Suffices to show integrability over shifted periods + let a : ℕ → ℝ := fun n ↦ t + (n - n₁) * T + rw [(by ring : t - n₁ * T = a 0), (by simp [a] : t + n₂ * T = a (n₁ + n₂))] + apply IntervalIntegrable.trans_iterate + -- Show integrability over a shifted period + intro k hk + convert (IntervalIntegrable.comp_sub_right h₂f ((k - n₁) * T)) using 1 + · funext x + simpa using (h₁f.sub_int_mul_eq (k - n₁)).symm + · simp [a, Nat.cast_add] + ring + +/-- Special case of Function.Periodic.intervalIntegrable: A periodic function is interval integrable +over every interval if it is interval integrable over the period starting from zero. -/ +theorem intervalIntegrable₀ (h₁f : Function.Periodic f T) (hT : 0 < T) + (h₂f : IntervalIntegrable f MeasureTheory.volume 0 T) (a₁ a₂ : ℝ) : + IntervalIntegrable f MeasureTheory.volume a₁ a₂ := by + apply h₁f.intervalIntegrable hT (t := 0) + simpa + +/-! +## Interval integrals of periodic functions +-/ + +variable [NormedSpace ℝ E] + /-- An auxiliary lemma for a more general `Function.Periodic.intervalIntegral_add_eq`. -/ theorem intervalIntegral_add_eq_of_pos (hf : Periodic f T) (hT : 0 < T) (t s : ℝ) : ∫ x in t..t + T, f x = ∫ x in s..s + T, f x := by @@ -284,43 +335,48 @@ section RealValued open Filter variable {g : ℝ → ℝ} -variable (hg : Periodic g T) (h_int : ∀ t₁ t₂, IntervalIntegrable g MeasureSpace.volume t₁ t₂) -include hg h_int +variable (hg : Periodic g T) +include hg /-- If `g : ℝ → ℝ` is periodic with period `T > 0`, then for any `t : ℝ`, the function `t ↦ ∫ x in 0..t, g x` is bounded below by `t ↦ X + ⌊t/T⌋ • Y` for appropriate constants `X` and `Y`. -/ -theorem sInf_add_zsmul_le_integral_of_pos (hT : 0 < T) (t : ℝ) : +theorem sInf_add_zsmul_le_integral_of_pos (h_int : IntervalIntegrable g MeasureSpace.volume 0 T) + (hT : 0 < T) (t : ℝ) : (sInf ((fun t => ∫ x in (0)..t, g x) '' Icc 0 T) + ⌊t / T⌋ • ∫ x in (0)..T, g x) ≤ ∫ x in (0)..t, g x := by + let h'_int := hg.intervalIntegrable₀ hT h_int let ε := Int.fract (t / T) * T conv_rhs => - rw [← Int.fract_div_mul_self_add_zsmul_eq T t (by linarith), ← - integral_add_adjacent_intervals (h_int 0 ε) (h_int _ _)] - rw [hg.intervalIntegral_add_zsmul_eq ⌊t / T⌋ ε h_int, hg.intervalIntegral_add_eq ε 0, zero_add, - add_le_add_iff_right] - exact (continuous_primitive h_int 0).continuousOn.sInf_image_Icc_le <| + rw [← Int.fract_div_mul_self_add_zsmul_eq T t (by linarith), + ← integral_add_adjacent_intervals (h'_int 0 ε) (h'_int _ _)] + rw [hg.intervalIntegral_add_zsmul_eq ⌊t / T⌋ ε (hg.intervalIntegrable₀ hT h_int), + hg.intervalIntegral_add_eq ε 0, zero_add, add_le_add_iff_right] + exact (continuous_primitive h'_int 0).continuousOn.sInf_image_Icc_le <| mem_Icc_of_Ico (Int.fract_div_mul_self_mem_Ico T t hT) /-- If `g : ℝ → ℝ` is periodic with period `T > 0`, then for any `t : ℝ`, the function `t ↦ ∫ x in 0..t, g x` is bounded above by `t ↦ X + ⌊t/T⌋ • Y` for appropriate constants `X` and `Y`. -/ -theorem integral_le_sSup_add_zsmul_of_pos (hT : 0 < T) (t : ℝ) : +theorem integral_le_sSup_add_zsmul_of_pos (h_int : IntervalIntegrable g MeasureSpace.volume 0 T) + (hT : 0 < T) (t : ℝ) : (∫ x in (0)..t, g x) ≤ sSup ((fun t => ∫ x in (0)..t, g x) '' Icc 0 T) + ⌊t / T⌋ • ∫ x in (0)..T, g x := by + let h'_int := hg.intervalIntegrable₀ hT h_int let ε := Int.fract (t / T) * T conv_lhs => rw [← Int.fract_div_mul_self_add_zsmul_eq T t (by linarith), ← - integral_add_adjacent_intervals (h_int 0 ε) (h_int _ _)] - rw [hg.intervalIntegral_add_zsmul_eq ⌊t / T⌋ ε h_int, hg.intervalIntegral_add_eq ε 0, zero_add, + integral_add_adjacent_intervals (h'_int 0 ε) (h'_int _ _)] + rw [hg.intervalIntegral_add_zsmul_eq ⌊t / T⌋ ε h'_int, hg.intervalIntegral_add_eq ε 0, zero_add, add_le_add_iff_right] - exact (continuous_primitive h_int 0).continuousOn.le_sSup_image_Icc + exact (continuous_primitive h'_int 0).continuousOn.le_sSup_image_Icc (mem_Icc_of_Ico (Int.fract_div_mul_self_mem_Ico T t hT)) /-- If `g : ℝ → ℝ` is periodic with period `T > 0` and `0 < ∫ x in 0..T, g x`, then `t ↦ ∫ x in 0..t, g x` tends to `∞` as `t` tends to `∞`. -/ theorem tendsto_atTop_intervalIntegral_of_pos (h₀ : 0 < ∫ x in (0)..T, g x) (hT : 0 < T) : Tendsto (fun t => ∫ x in (0)..t, g x) atTop atTop := by + have h_int := intervalIntegrable_of_integral_ne_zero h₀.ne' apply tendsto_atTop_mono (hg.sInf_add_zsmul_le_integral_of_pos h_int hT) apply atTop.tendsto_atTop_add_const_left (sInf <| (fun t => ∫ x in (0)..t, g x) '' Icc 0 T) apply Tendsto.atTop_zsmul_const h₀ @@ -330,6 +386,7 @@ theorem tendsto_atTop_intervalIntegral_of_pos (h₀ : 0 < ∫ x in (0)..T, g x) `t ↦ ∫ x in 0..t, g x` tends to `-∞` as `t` tends to `-∞`. -/ theorem tendsto_atBot_intervalIntegral_of_pos (h₀ : 0 < ∫ x in (0)..T, g x) (hT : 0 < T) : Tendsto (fun t => ∫ x in (0)..t, g x) atBot atBot := by + have h_int := intervalIntegrable_of_integral_ne_zero h₀.ne' apply tendsto_atBot_mono (hg.integral_le_sSup_add_zsmul_of_pos h_int hT) apply atBot.tendsto_atBot_add_const_left (sSup <| (fun t => ∫ x in (0)..t, g x) '' Icc 0 T) apply Tendsto.atBot_zsmul_const h₀ @@ -337,15 +394,17 @@ theorem tendsto_atBot_intervalIntegral_of_pos (h₀ : 0 < ∫ x in (0)..T, g x) /-- If `g : ℝ → ℝ` is periodic with period `T > 0` and `∀ x, 0 < g x`, then `t ↦ ∫ x in 0..t, g x` tends to `∞` as `t` tends to `∞`. -/ -theorem tendsto_atTop_intervalIntegral_of_pos' (h₀ : ∀ x, 0 < g x) (hT : 0 < T) : +theorem tendsto_atTop_intervalIntegral_of_pos' + (h_int : IntervalIntegrable g MeasureSpace.volume 0 T) (h₀ : ∀ x, 0 < g x) (hT : 0 < T) : Tendsto (fun t => ∫ x in (0)..t, g x) atTop atTop := - hg.tendsto_atTop_intervalIntegral_of_pos h_int (intervalIntegral_pos_of_pos (h_int 0 T) h₀ hT) hT + hg.tendsto_atTop_intervalIntegral_of_pos (intervalIntegral_pos_of_pos h_int h₀ hT) hT /-- If `g : ℝ → ℝ` is periodic with period `T > 0` and `∀ x, 0 < g x`, then `t ↦ ∫ x in 0..t, g x` tends to `-∞` as `t` tends to `-∞`. -/ -theorem tendsto_atBot_intervalIntegral_of_pos' (h₀ : ∀ x, 0 < g x) (hT : 0 < T) : - Tendsto (fun t => ∫ x in (0)..t, g x) atBot atBot := - hg.tendsto_atBot_intervalIntegral_of_pos h_int (intervalIntegral_pos_of_pos (h_int 0 T) h₀ hT) hT +theorem tendsto_atBot_intervalIntegral_of_pos' + (h_int : IntervalIntegrable g MeasureSpace.volume 0 T) (h₀ : ∀ x, 0 < g x) (hT : 0 < T) : + Tendsto (fun t => ∫ x in (0)..t, g x) atBot atBot := by + exact hg.tendsto_atBot_intervalIntegral_of_pos (intervalIntegral_pos_of_pos h_int h₀ hT) hT end RealValued diff --git a/Mathlib/MeasureTheory/Integral/Pi.lean b/Mathlib/MeasureTheory/Integral/Pi.lean index 2bf9b9af0c309..9e6ba84a0ce90 100644 --- a/Mathlib/MeasureTheory/Integral/Pi.lean +++ b/Mathlib/MeasureTheory/Integral/Pi.lean @@ -28,7 +28,7 @@ theorem Integrable.fin_nat_prod {n : ℕ} {E : Fin n → Type*} {f : (i : Fin n) → E i → 𝕜} (hf : ∀ i, Integrable (f i)) : Integrable (fun (x : (i : Fin n) → E i) ↦ ∏ i, f i (x i)) := by induction n with - | zero => simp only [Finset.univ_eq_empty, Finset.prod_empty, volume_pi, + | zero => simp only [Finset.univ_eq_empty, Finset.prod_empty, volume_pi, isFiniteMeasure_iff, integrable_const_iff, one_ne_zero, pi_empty_univ, ENNReal.one_lt_top, or_true] | succ n n_ih => have := ((measurePreserving_piFinSuccAbove (fun i => (volume : Measure (E i))) 0).symm) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 5d1379bf8604a..8216fb8b065b7 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -27,10 +27,10 @@ the space such that the measurable space is generated by the union of all partit ## Main statements -* `MeasurableSpace.measurable_equiv_nat_bool_of_countablyGenerated`: if a measurable space is +* `MeasurableSpace.measurableEquiv_nat_bool_of_countablyGenerated`: if a measurable space is countably generated and separates points, it is measure equivalent to a subset of the Cantor Space `ℕ → Bool` (equipped with the product sigma algebra). -* `MeasurableSpace.measurable_injection_nat_bool_of_countablyGenerated`: If a measurable space +* `MeasurableSpace.measurable_injection_nat_bool_of_countablySeparated`: If a measurable space admits a countable sequence of measurable sets separating points, it admits a measurable injection into the Cantor space `ℕ → Bool` `ℕ → Bool` (equipped with the product sigma algebra). diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index 40d6422e75b95..2332cc948c41d 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -3,9 +3,9 @@ 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, Peter Pfaffelhuber -/ -import Mathlib.Algebra.BigOperators.Group.Finset.Basic -import Mathlib.Data.ENNReal.Basic import Mathlib.MeasureTheory.SetSemiring +import Mathlib.Topology.Algebra.InfiniteSum.Defs +import Mathlib.Topology.Instances.ENNReal.Lemmas /-! # Additive Contents @@ -37,12 +37,16 @@ If `C` is a set ring (`MeasureTheory.IsSetRing C`), we have, for `s, t ∈ C`, * `MeasureTheory.addContent_union_le`: `m (s ∪ t) ≤ m s + m t` * `MeasureTheory.addContent_le_diff`: `m s - m t ≤ m (s \ t)` +* `IsSetRing.addContent_of_union`: a function on a ring of sets which is additive on pairs of +disjoint sets defines an additive content +* `addContent_iUnion_eq_sum_of_tendsto_zero`: if an additive content is continuous at `∅`, then +its value on a countable disjoint union is the sum of the values -/ -open Set Finset +open Set Finset Function Filter -open scoped ENNReal +open scoped ENNReal Topology namespace MeasureTheory @@ -157,9 +161,10 @@ lemma addContent_biUnion_le {ι : Type*} (hC : IsSetRing C) {s : ι → Set α} {S : Finset ι} (hs : ∀ n ∈ S, s n ∈ C) : m (⋃ i ∈ S, s i) ≤ ∑ i ∈ S, m (s i) := by classical - induction' S using Finset.induction with i S hiS h hs - · simp - · rw [Finset.sum_insert hiS] + induction S using Finset.induction with + | empty => simp + | @insert i S hiS h => + rw [Finset.sum_insert hiS] simp_rw [← Finset.mem_coe, Finset.coe_insert, Set.biUnion_insert] simp only [Finset.mem_insert, forall_eq_or_imp] at hs refine (addContent_union_le hC hs.1 (hC.biUnion_mem S hs.2)).trans ?_ @@ -173,6 +178,82 @@ lemma le_addContent_diff (m : AddContent C) (hC : IsSetRing C) (hs : s ∈ C) (h rw [tsub_eq_zero_of_le (addContent_mono hC.isSetSemiring (hC.inter_mem hs ht) ht inter_subset_right), add_zero] +lemma addContent_diff_of_ne_top (m : AddContent C) (hC : IsSetRing C) + (hm_ne_top : ∀ s ∈ C, m s ≠ ∞) + {s t : Set α} (hs : s ∈ C) (ht : t ∈ C) (hts : t ⊆ s) : + m (s \ t) = m s - m t := by + have h_union : m (t ∪ s \ t) = m t + m (s \ t) := + addContent_union hC ht (hC.diff_mem hs ht) disjoint_sdiff_self_right + simp_rw [Set.union_diff_self, Set.union_eq_right.mpr hts] at h_union + rw [h_union, ENNReal.add_sub_cancel_left (hm_ne_top _ ht)] + +lemma addContent_accumulate (m : AddContent C) (hC : IsSetRing C) + {s : ℕ → Set α} (hs_disj : Pairwise (Disjoint on s)) (hsC : ∀ i, s i ∈ C) (n : ℕ) : + m (Set.Accumulate s n) = ∑ i in Finset.range (n + 1), m (s i) := by + induction n with + | zero => simp + | succ n hn => + rw [Finset.sum_range_succ, ← hn, Set.accumulate_succ, addContent_union hC _ (hsC _)] + · exact Set.disjoint_accumulate hs_disj (Nat.lt_succ_self n) + · exact hC.accumulate_mem hsC n + +/-- A function which is additive on disjoint elements in a ring of sets `C` defines an +additive content on `C`. -/ +def IsSetRing.addContent_of_union (m : Set α → ℝ≥0∞) (hC : IsSetRing C) (m_empty : m ∅ = 0) + (m_add : ∀ {s t : Set α} (_hs : s ∈ C) (_ht : t ∈ C), Disjoint s t → m (s ∪ t) = m s + m t) : + AddContent C where + toFun := m + empty' := m_empty + sUnion' I h_ss h_dis h_mem := by + classical + induction I using Finset.induction with + | empty => simp only [Finset.coe_empty, Set.sUnion_empty, Finset.sum_empty, m_empty] + | @insert s I hsI h => + rw [Finset.coe_insert] at * + rw [Set.insert_subset_iff] at h_ss + rw [Set.pairwiseDisjoint_insert_of_not_mem] at h_dis + swap; · exact hsI + have h_sUnion_mem : ⋃₀ ↑I ∈ C := by + rw [Set.sUnion_eq_biUnion] + apply hC.biUnion_mem + intro n hn + exact h_ss.2 hn + rw [Set.sUnion_insert, m_add h_ss.1 h_sUnion_mem (Set.disjoint_sUnion_right.mpr h_dis.2), + Finset.sum_insert hsI, h h_ss.2 h_dis.1] + rwa [Set.sUnion_insert] at h_mem + +/-- In a ring of sets, continuity of an additive content at `∅` implies σ-additivity. +This is not true in general in semirings, or without the hypothesis that `m` is finite. See the +examples 7 and 8 in Halmos' book Measure Theory (1974), page 40. -/ +theorem addContent_iUnion_eq_sum_of_tendsto_zero (hC : IsSetRing C) (m : AddContent C) + (hm_ne_top : ∀ s ∈ C, m s ≠ ∞) + (hm_tendsto : ∀ ⦃s : ℕ → Set α⦄ (_ : ∀ n, s n ∈ C), + Antitone s → (⋂ n, s n) = ∅ → Tendsto (fun n ↦ m (s n)) atTop (𝓝 0)) + ⦃f : ℕ → Set α⦄ (hf : ∀ i, f i ∈ C) (hUf : (⋃ i, f i) ∈ C) + (h_disj : Pairwise (Disjoint on f)) : + m (⋃ i, f i) = ∑' i, m (f i) := by + -- We use the continuity of `m` at `∅` on the sequence `n ↦ (⋃ i, f i) \ (set.accumulate f n)` + let s : ℕ → Set α := fun n ↦ (⋃ i, f i) \ Set.Accumulate f n + have hCs n : s n ∈ C := hC.diff_mem hUf (hC.accumulate_mem hf n) + have h_tendsto : Tendsto (fun n ↦ m (s n)) atTop (𝓝 0) := by + refine hm_tendsto hCs ?_ ?_ + · intro i j hij x hxj + rw [Set.mem_diff] at hxj ⊢ + exact ⟨hxj.1, fun hxi ↦ hxj.2 (Set.monotone_accumulate hij hxi)⟩ + · simp_rw [s, Set.diff_eq] + rw [Set.iInter_inter_distrib, Set.iInter_const, ← Set.compl_iUnion, Set.iUnion_accumulate] + exact Set.inter_compl_self _ + have hmsn n : m (s n) = m (⋃ i, f i) - ∑ i in Finset.range (n + 1), m (f i) := by + rw [addContent_diff_of_ne_top m hC hm_ne_top hUf (hC.accumulate_mem hf n) + (Set.accumulate_subset_iUnion _), addContent_accumulate m hC h_disj hf n] + simp_rw [hmsn] at h_tendsto + refine tendsto_nhds_unique ?_ (ENNReal.tendsto_nat_tsum fun i ↦ m (f i)) + refine (Filter.tendsto_add_atTop_iff_nat 1).mp ?_ + rwa [ENNReal.tendsto_const_sub_nhds_zero_iff (hm_ne_top _ hUf) (fun n ↦ ?_)] at h_tendsto + rw [← addContent_accumulate m hC h_disj hf] + exact addContent_mono hC.isSetSemiring (hC.accumulate_mem hf n) hUf + (Set.accumulate_subset_iUnion _) + end IsSetRing end MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/Comap.lean b/Mathlib/MeasureTheory/Measure/Comap.lean index 457a53c721731..e9f8da9f066b3 100644 --- a/Mathlib/MeasureTheory/Measure/Comap.lean +++ b/Mathlib/MeasureTheory/Measure/Comap.lean @@ -173,5 +173,5 @@ lemma map_symm {μ : Measure α} (e : β ≃ᵐ α) : μ.map e.symm = μ.comap e end MeasurableEquiv -lemma comap_swap (μ : Measure (α × β)) : μ.comap Prod.swap = μ.map Prod.swap := +lemma MeasureTheory.Measure.comap_swap (μ : Measure (α × β)) : μ.comap .swap = μ.map .swap := (MeasurableEquiv.prodComm ..).comap_symm diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index babc5786d359e..847ad1bcd24bf 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -209,7 +209,6 @@ instance instSMul : SMul R (FiniteMeasure Ω) where @[simp, norm_cast] theorem toMeasure_zero : ((↑) : FiniteMeasure Ω → Measure Ω) 0 = 0 := rfl --- Porting note: with `simp` here the `coeFn` lemmas below fall prey to `simpNF`: the LHS simplifies @[norm_cast] theorem toMeasure_add (μ ν : FiniteMeasure Ω) : ↑(μ + ν) = (↑μ + ↑ν : Measure Ω) := rfl diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index c52acf767b115..1e44f801378f7 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -46,7 +46,7 @@ variable {G H : Type*} [MeasurableSpace G] [Group G] [TopologicalSpace G] [TopologicalGroup G] [BorelSpace G] [LocallyCompactSpace G] [MeasurableSpace H] [SeminormedGroup H] [OpensMeasurableSpace H] --- TODO: This could be streamlined by proving that inner regular always exist +-- TODO: This could be streamlined by proving that inner regular measures always exist open Metric Bornology in @[to_additive] lemma _root_.MonoidHom.exists_nhds_isBounded (f : G →* H) (hf : Measurable f) (x : G) : @@ -69,6 +69,14 @@ lemma _root_.MonoidHom.exists_nhds_isBounded (f : G →* H) (hf : Measurable f) end SeminormedGroup +/-- A Borel-measurable group hom from a locally compact normed group to a real normed space is +continuous. -/ +lemma AddMonoidHom.continuous_of_measurable {G H : Type*} + [SeminormedAddCommGroup G] [MeasurableSpace G] [BorelSpace G] [LocallyCompactSpace G] + [SeminormedAddCommGroup H] [MeasurableSpace H] [OpensMeasurableSpace H] [NormedSpace ℝ H] + (f : G →+ H) (hf : Measurable f) : Continuous f := + let ⟨_s, hs, hbdd⟩ := f.exists_nhds_isBounded hf 0; f.continuous_of_isBounded_nhds_zero hs hbdd + variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [MeasurableSpace E] [BorelSpace E] [FiniteDimensional ℝ E] (μ : Measure E) [IsAddHaarMeasure μ] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index a4ed925b7ac08..3a14731891763 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -921,7 +921,7 @@ theorem _root_.MonoidHom.measurePreserving have : C * ν univ = 1 * ν univ := by rw [one_mul, ← smul_eq_mul, ← ENNReal.smul_def, ← smul_apply, ← hC, map_apply hcont.measurable .univ, preimage_univ, huniv] - rwa [ENNReal.mul_eq_mul_right (NeZero.ne _) (measure_ne_top _ _), ENNReal.coe_eq_one] at this + rwa [ENNReal.mul_left_inj (NeZero.ne _) (measure_ne_top _ _), ENNReal.coe_eq_one] at this end Group @@ -950,7 +950,7 @@ instance (priority := 100) IsHaarMeasure.isInvInvariant_of_regular conv_rhs => rw [μeq] simp have : c ^ 2 = 1 ^ 2 := - (ENNReal.mul_eq_mul_right (measure_pos_of_nonempty_interior _ K.interior_nonempty).ne' + (ENNReal.mul_left_inj (measure_pos_of_nonempty_interior _ K.interior_nonempty).ne' K.isCompact.measure_lt_top.ne).1 this have : c = 1 := (ENNReal.pow_right_strictMono two_ne_zero).injective this rw [hc, this, one_smul] @@ -975,7 +975,7 @@ instance (priority := 100) IsHaarMeasure.isInvInvariant_of_innerRegular conv_rhs => rw [μeq] simp have : c ^ 2 = 1 ^ 2 := - (ENNReal.mul_eq_mul_right (measure_pos_of_nonempty_interior _ K.interior_nonempty).ne' + (ENNReal.mul_left_inj (measure_pos_of_nonempty_interior _ K.interior_nonempty).ne' K.isCompact.measure_lt_top.ne).1 this have : c = 1 := (ENNReal.pow_right_strictMono two_ne_zero).injective this rw [hc, this, one_smul] diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 2f8b288b9314b..5bd12b8bc5683 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -501,7 +501,7 @@ theorem mkMetric_le_liminf_tsum {β : Type*} {ι : β → Type*} [∀ n, Countab haveI : ∀ n, Encodable (ι n) := fun n => Encodable.ofCountable _ simp only [mkMetric_apply] refine iSup₂_le fun ε hε => ?_ - refine le_of_forall_le_of_dense fun c hc => ?_ + refine le_of_forall_gt_imp_ge_of_dense fun c hc => ?_ rcases ((frequently_lt_of_liminf_lt (by isBoundedDefault) hc).and_eventually ((hr.eventually (gt_mem_nhds hε)).and (ht.and hst))).exists with ⟨n, hn, hrn, htn, hstn⟩ diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 71c507d192908..cb4284f411de1 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -34,13 +34,14 @@ variable {m0 : MeasurableSpace α} [MeasurableSpace β] {μ ν ν₁ ν₂ : Mea section IsFiniteMeasure /-- A measure `μ` is called finite if `μ univ < ∞`. -/ +@[mk_iff] class IsFiniteMeasure (μ : Measure α) : Prop where measure_univ_lt_top : μ univ < ∞ -theorem not_isFiniteMeasure_iff : ¬IsFiniteMeasure μ ↔ μ Set.univ = ∞ := by - refine ⟨fun h => ?_, fun h => fun h' => h'.measure_univ_lt_top.ne h⟩ - by_contra h' - exact h ⟨lt_top_iff_ne_top.mpr h'⟩ +lemma not_isFiniteMeasure_iff : ¬IsFiniteMeasure μ ↔ μ univ = ∞ := by simp [isFiniteMeasure_iff] + +lemma isFiniteMeasure_restrict : IsFiniteMeasure (μ.restrict s) ↔ μ s ≠ ∞ := by + simp [isFiniteMeasure_iff, lt_top_iff_ne_top] instance Restrict.isFiniteMeasure (μ : Measure α) [hs : Fact (μ s < ∞)] : IsFiniteMeasure (μ.restrict s) := @@ -540,8 +541,6 @@ theorem finiteAtBot {m0 : MeasurableSpace α} (μ : Measure α) : μ.FiniteAtFil about the sets, such as that they are monotone. `SigmaFinite` is defined in terms of this: `μ` is σ-finite if there exists a sequence of finite spanning sets in the collection of all measurable sets. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure FiniteSpanningSetsIn {m0 : MeasurableSpace α} (μ : Measure α) (C : Set (Set α)) where protected set : ℕ → Set α protected set_mem : ∀ i, set i ∈ C diff --git a/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean b/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean index be505ddb8561a..36c1a59159274 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean @@ -72,6 +72,18 @@ theorem isCaratheodory_union (h₁ : IsCaratheodory m s₁) (h₂ : IsCaratheodo union_diff_left, h₂ (t ∩ s₁)] simp [diff_eq, add_assoc] +variable {m} in +lemma IsCaratheodory.biUnion_of_finite {ι : Type*} {s : ι → Set α} {t : Set ι} (ht : t.Finite) + (h : ∀ i ∈ t, m.IsCaratheodory (s i)) : + m.IsCaratheodory (⋃ i ∈ t, s i) := by + classical + lift t to Finset ι using ht + induction t using Finset.induction_on with + | empty => simp + | @insert i t hi IH => + simp only [Finset.mem_coe, Finset.mem_insert, iUnion_iUnion_eq_or_left] at h ⊢ + exact m.isCaratheodory_union (h _ <| Or.inl rfl) (IH fun _ hj ↦ h _ <| Or.inr hj) + theorem measure_inter_union (h : s₁ ∩ s₂ ⊆ ∅) (h₁ : IsCaratheodory m s₁) {t : Set α} : m (t ∩ (s₁ ∪ s₂)) = m (t ∩ s₁) + m (t ∩ s₂) := by rw [h₁, Set.inter_assoc, Set.union_inter_cancel_left, inter_diff_assoc, union_diff_cancel_left h] @@ -93,17 +105,16 @@ theorem isCaratheodory_inter (h₁ : IsCaratheodory m s₁) (h₂ : IsCaratheodo lemma isCaratheodory_diff (h₁ : IsCaratheodory m s₁) (h₂ : IsCaratheodory m s₂) : IsCaratheodory m (s₁ \ s₂) := m.isCaratheodory_inter h₁ (m.isCaratheodory_compl h₂) -lemma isCaratheodory_partialSups {s : ℕ → Set α} (h : ∀ i, m.IsCaratheodory (s i)) (i : ℕ) : +lemma isCaratheodory_partialSups {ι : Type*} [PartialOrder ι] [LocallyFiniteOrderBot ι] + {s : ι → Set α} (h : ∀ i, m.IsCaratheodory (s i)) (i : ι) : m.IsCaratheodory (partialSups s i) := by - induction i with - | zero => exact h 0 - | succ i hi => exact partialSups_add_one s i ▸ m.isCaratheodory_union hi (h (i + 1)) - -lemma isCaratheodory_disjointed {s : ℕ → Set α} (h : ∀ i, m.IsCaratheodory (s i)) (i : ℕ) : - m.IsCaratheodory (disjointed s i) := by - induction i with - | zero => exact h 0 - | succ i _ => exact m.isCaratheodory_diff (h (i + 1)) (m.isCaratheodory_partialSups h i) + simpa only [partialSups_apply, Finset.sup'_eq_sup, Finset.sup_set_eq_biUnion, ← Finset.mem_coe, + Finset.coe_Iic] using .biUnion_of_finite (finite_Iic _) (fun j _ ↦ h j) + +lemma isCaratheodory_disjointed {ι : Type*} [PartialOrder ι] [LocallyFiniteOrderBot ι] + {s : ι → Set α} (h : ∀ i, m.IsCaratheodory (s i)) (i : ι) : + m.IsCaratheodory (disjointed s i) := + disjointedRec (fun _ j ht ↦ m.isCaratheodory_diff ht <| h j) (h i) theorem isCaratheodory_sum {s : ℕ → Set α} (h : ∀ i, IsCaratheodory m (s i)) (hd : Pairwise (Disjoint on s)) {t : Set α} : diff --git a/Mathlib/MeasureTheory/OuterMeasure/Operations.lean b/Mathlib/MeasureTheory/OuterMeasure/Operations.lean index 9d73c991fc3ea..ff3e6271b01bd 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Operations.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Operations.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.Order.Group.Indicator import Mathlib.MeasureTheory.OuterMeasure.Basic /-! diff --git a/Mathlib/MeasureTheory/SetSemiring.lean b/Mathlib/MeasureTheory/SetSemiring.lean index 9df3fa00df048..503f9ca6b3e8d 100644 --- a/Mathlib/MeasureTheory/SetSemiring.lean +++ b/Mathlib/MeasureTheory/SetSemiring.lean @@ -3,6 +3,8 @@ 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, Peter Pfaffelhuber -/ +import Mathlib.Data.Nat.Lattice +import Mathlib.Data.Set.Accumulate import Mathlib.Data.Set.Pairwise.Lattice import Mathlib.MeasureTheory.PiSystem @@ -303,9 +305,10 @@ lemma biUnion_mem {ι : Type*} (hC : IsSetRing C) {s : ι → Set α} (S : Finset ι) (hs : ∀ n ∈ S, s n ∈ C) : ⋃ i ∈ S, s i ∈ C := by classical - induction' S using Finset.induction with i S _ h hs - · simp [hC.empty_mem] - · simp_rw [← Finset.mem_coe, Finset.coe_insert, Set.biUnion_insert] + induction S using Finset.induction with + | empty => simp [hC.empty_mem] + | @insert i S _ h => + simp_rw [← Finset.mem_coe, Finset.coe_insert, Set.biUnion_insert] refine hC.union_mem (hs i (mem_insert_self i S)) ?_ exact h (fun n hnS ↦ hs n (mem_insert_of_mem hnS)) @@ -321,16 +324,43 @@ lemma biInter_mem {ι : Type*} (hC : IsSetRing C) {s : ι → Set α} refine hC.inter_mem hs.1 ?_ exact h (fun n hnS ↦ hs.2 n hnS) -lemma partialSups_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ n, s n ∈ C) (n : ℕ) : +lemma finsetSup_mem (hC : IsSetRing C) {ι : Type*} {s : ι → Set α} {t : Finset ι} + (hs : ∀ i ∈ t, s i ∈ C) : + t.sup s ∈ C := by + classical + induction t using Finset.induction_on with + | empty => exact hC.empty_mem + | @insert m t hm ih => + simpa only [sup_insert] using + hC.union_mem (hs m <| mem_insert_self m t) (ih <| fun i hi ↦ hs _ <| mem_insert_of_mem hi) + +lemma partialSups_mem {ι : Type*} [Preorder ι] [LocallyFiniteOrderBot ι] + (hC : IsSetRing C) {s : ι → Set α} (hs : ∀ n, s n ∈ C) (n : ι) : partialSups s n ∈ C := by - rw [partialSups_eq_biUnion_range] - exact hC.biUnion_mem _ (fun n _ ↦ hs n) - -lemma disjointed_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ n, s n ∈ C) (n : ℕ) : - disjointed s n ∈ C := by - cases n with - | zero => exact hs 0 - | succ n => exact hC.diff_mem (hs n.succ) (hC.partialSups_mem hs n) + simpa only [partialSups_apply, sup'_eq_sup] using hC.finsetSup_mem (fun i hi ↦ hs i) + +lemma disjointed_mem {ι : Type*} [Preorder ι] [LocallyFiniteOrderBot ι] + (hC : IsSetRing C) {s : ι → Set α} (hs : ∀ j, s j ∈ C) (i : ι) : + disjointed s i ∈ C := + disjointedRec (fun _ j ht ↦ hC.diff_mem ht <| hs j) (hs i) + +theorem iUnion_le_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ n, s n ∈ C) (n : ℕ) : + (⋃ i ≤ n, s i) ∈ C := by + induction n with + | zero => simp [hs 0] + | succ n hn => rw [biUnion_le_succ]; exact hC.union_mem hn (hs _) + +theorem iInter_le_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ n, s n ∈ C) (n : ℕ) : + (⋂ i ≤ n, s i) ∈ C := by + induction n with + | zero => simp [hs 0] + | succ n hn => rw [biInter_le_succ]; exact hC.inter_mem hn (hs _) + +theorem accumulate_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ i, s i ∈ C) (n : ℕ) : + Accumulate s n ∈ C := by + induction n with + | zero => simp [hs 0] + | succ n hn => rw [accumulate_succ]; exact hC.union_mem hn (hs _) end IsSetRing diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index 6e23ef16e9acd..6de4fea493467 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -81,14 +81,10 @@ protected def sum (L' : Language.{u', v'}) : Language := ⟨fun n => L.Functions n ⊕ L'.Functions n, fun n => L.Relations n ⊕ L'.Relations n⟩ /-- The type of constants in a given language. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] protected abbrev Constants := L.Functions 0 /-- The type of symbols in a given language. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] abbrev Symbols := (Σ l, L.Functions l) ⊕ (Σ l, L.Relations l) diff --git a/Mathlib/ModelTheory/Satisfiability.lean b/Mathlib/ModelTheory/Satisfiability.lean index 779378c10f0df..acf91d57f6bab 100644 --- a/Mathlib/ModelTheory/Satisfiability.lean +++ b/Mathlib/ModelTheory/Satisfiability.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import Mathlib.ModelTheory.Ultraproducts import Mathlib.ModelTheory.Bundled import Mathlib.ModelTheory.Skolem +import Mathlib.Order.Filter.AtTopBot /-! # First-Order Satisfiability diff --git a/Mathlib/ModelTheory/Ultraproducts.lean b/Mathlib/ModelTheory/Ultraproducts.lean index dbe1d5d34724f..338cdcea73860 100644 --- a/Mathlib/ModelTheory/Ultraproducts.lean +++ b/Mathlib/ModelTheory/Ultraproducts.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.ModelTheory.Quotients +import Mathlib.Order.Filter.Finite import Mathlib.Order.Filter.Germ.Basic -import Mathlib.Order.Filter.Ultrafilter +import Mathlib.Order.Filter.Ultrafilter.Defs /-! # Ultraproducts and Łoś's Theorem diff --git a/Mathlib/NumberTheory/AbelSummation.lean b/Mathlib/NumberTheory/AbelSummation.lean index 202136f5c40ad..e31f094ec8fa5 100644 --- a/Mathlib/NumberTheory/AbelSummation.lean +++ b/Mathlib/NumberTheory/AbelSummation.lean @@ -51,8 +51,8 @@ namespace abelSummationProof open intervalIntegral IntervalIntegrable -private theorem sumlocc (n : ℕ) : - ∀ᵐ t, t ∈ Set.Icc (n : ℝ) (n + 1) → ∑ k ∈ Icc 0 ⌊t⌋₊, c k = ∑ k ∈ Icc 0 n, c k := by +private theorem sumlocc {m : ℕ} (n : ℕ) : + ∀ᵐ t, t ∈ Set.Icc (n : ℝ) (n + 1) → ∑ k ∈ Icc m ⌊t⌋₊, c k = ∑ k ∈ Icc m n, c k := by filter_upwards [Ico_ae_eq_Icc] with t h ht rw [Nat.floor_eq_on_Ico _ _ (h.mpr ht)] @@ -83,11 +83,12 @@ private theorem ineqofmemIco' {k : ℕ} (hk : k ∈ Ico (⌊a⌋₊ + 1) ⌊b⌋ a ≤ k ∧ k + 1 ≤ b := ineqofmemIco (by rwa [← Finset.coe_Ico]) -private theorem integrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg_int : IntegrableOn g (Set.Icc a b)) : - IntegrableOn (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) (Set.Icc a b) := by +theorem _root_.integrableOn_mul_sum_Icc {m : ℕ} (ha : 0 ≤ a) {g : ℝ → 𝕜} + (hg_int : IntegrableOn g (Set.Icc a b)) : + IntegrableOn (fun t ↦ g t * ∑ k ∈ Icc m ⌊t⌋₊, c k) (Set.Icc a b) := by obtain hab | hab := le_or_gt a b · obtain hb | hb := eq_or_lt_of_le (Nat.floor_le_floor hab) - · have : ∀ᵐ t, t ∈ Set.Icc a b → ∑ k ∈ Icc 0 ⌊a⌋₊, c k = ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by + · have : ∀ᵐ t, t ∈ Set.Icc a b → ∑ k ∈ Icc m ⌊a⌋₊, c k = ∑ k ∈ Icc m ⌊t⌋₊, c k := by filter_upwards [sumlocc c ⌊a⌋₊] with t ht₁ ht₂ rw [ht₁ ⟨(Nat.floor_le ha).trans ht₂.1, hb ▸ ht₂.2.trans (Nat.lt_floor_add_one b).le⟩] rw [← ae_restrict_iff' measurableSet_Icc] at this @@ -95,7 +96,7 @@ private theorem integrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg_int : Int (hg_int.mul_const _) ((Filter.EventuallyEq.refl _ g).mul this) · have h_locint {t₁ t₂ : ℝ} {n : ℕ} (h : t₁ ≤ t₂) (h₁ : n ≤ t₁) (h₂ : t₂ ≤ n + 1) (h₃ : a ≤ t₁) (h₄ : t₂ ≤ b) : - IntervalIntegrable (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) volume t₁ t₂ := by + IntervalIntegrable (fun t ↦ g t * ∑ k ∈ Icc m ⌊t⌋₊, c k) volume t₁ t₂ := by rw [intervalIntegrable_iff_integrableOn_Icc_of_le h] exact (IntegrableOn.mono_set (hg_int.mul_const _) (Set.Icc_subset_Icc h₃ h₄)).congr <| ae_restrict_of_ae_restrict_of_subset (Set.Icc_subset_Icc h₁ h₂) @@ -154,15 +155,15 @@ theorem _root_.sum_mul_eq_sub_sub_integral_mul (ha : 0 ≤ a) (hab : a ≤ b) -- (Note we have 5 goals, but the 1st and 3rd are identical. TODO: find a non-hacky way of dealing -- with both at once.) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux5] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_left aux6) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_left aux6) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux6] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux5) · rw [intervalIntegrable_iff_integrableOn_Icc_of_le aux3] - exact (integrablemulsum c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux4) + exact (integrableOn_mul_sum_Icc c ha hf_int).mono_set (Set.Icc_subset_Icc_right aux4) · exact fun k hk ↦ (intervalIntegrable_iff_integrableOn_Icc_of_le (mod_cast k.le_succ)).mpr - <| (integrablemulsum c ha hf_int).mono_set + <| (integrableOn_mul_sum_Icc c ha hf_int).mono_set <| (Set.Icc_subset_Icc_iff (mod_cast k.le_succ)).mpr <| mod_cast (ineqofmemIco hk) /-- A version of `sum_mul_eq_sub_sub_integral_mul` where the endpoints are `Nat`. -/ @@ -209,7 +210,7 @@ theorem sum_mul_eq_sub_integral_mul₀ (hc : c 0 = 0) (b : ℝ) f b * (∑ k ∈ Icc 0 ⌊b⌋₊, c k) - ∫ t in Set.Ioc 1 b, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k := by obtain hb | hb := le_or_gt 1 b · have : 1 ≤ ⌊b⌋₊ := (Nat.one_le_floor_iff _).mpr hb - nth_rewrite 1 [Icc_eq_cons_Ioc (by omega), sum_cons, ← Nat.Icc_succ_left, + nth_rewrite 1 [Icc_eq_cons_Ioc (Nat.zero_le _), sum_cons, ← Nat.Icc_succ_left, Icc_eq_cons_Ioc (by omega), sum_cons] rw [Nat.succ_eq_add_one, zero_add, ← Nat.floor_one (α := ℝ), sum_mul_eq_sub_sub_integral_mul c zero_le_one hb hf_diff hf_int, Nat.floor_one, Nat.cast_one, @@ -235,13 +236,14 @@ section limit open Filter Topology abelSummationProof intervalIntegral -private theorem locallyintegrablemulsum (ha : 0 ≤ a) {g : ℝ → 𝕜} (hg : IntegrableOn g (Set.Ici a)) : - LocallyIntegrableOn (fun t ↦ g t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k) (Set.Ici a) := by +theorem locallyIntegrableOn_mul_sum_Icc {m : ℕ} (ha : 0 ≤ a) {g : ℝ → 𝕜} + (hg : IntegrableOn g (Set.Ici a)) : + LocallyIntegrableOn (fun t ↦ g t * ∑ k ∈ Icc m ⌊t⌋₊, c k) (Set.Ici a) := by refine (locallyIntegrableOn_iff isLocallyClosed_Ici).mpr fun K hK₁ hK₂ ↦ ?_ by_cases hK₃ : K.Nonempty · have h_inf : a ≤ sInf K := (hK₁ (hK₂.sInf_mem hK₃)) refine IntegrableOn.mono_set ?_ (Bornology.IsBounded.subset_Icc_sInf_sSup hK₂.isBounded) - refine integrablemulsum _ (ha.trans h_inf) (hg.mono_set ?_) + refine integrableOn_mul_sum_Icc _ (ha.trans h_inf) (hg.mono_set ?_) exact (Set.Icc_subset_Ici_iff (Real.sInf_le_sSup _ hK₂.bddBelow hK₂.bddAbove)).mpr h_inf · rw [Set.not_nonempty_iff_eq_empty.mp hK₃] exact integrableOn_empty @@ -259,7 +261,8 @@ theorem tendsto_sum_mul_atTop_nhds_one_sub_integral refine Tendsto.congr (fun _ ↦ by rw [← integral_of_le (Nat.cast_nonneg _)]) ?_ refine intervalIntegral_tendsto_integral_Ioi _ ?_ tendsto_natCast_atTop_atTop exact integrableOn_Ici_iff_integrableOn_Ioi.mp - <| (locallyintegrablemulsum c le_rfl hf_int).integrableOn_of_isBigO_atTop hg_dom hg_int + <| (locallyIntegrableOn_mul_sum_Icc c le_rfl hf_int).integrableOn_of_isBigO_atTop + hg_dom hg_int refine (h_lim.sub h_lim').congr (fun _ ↦ ?_) rw [sum_mul_eq_sub_integral_mul' _ _ (fun t ht ↦ hf_diff _ ht.1) (hf_int.mono_set Set.Icc_subset_Ici_self)] @@ -280,7 +283,8 @@ theorem tendsto_sum_mul_atTop_nhds_one_sub_integral₀ (hc : c 0 = 0) atTop (𝓝 (∫ t in Set.Ioi 1, deriv f t * ∑ k ∈ Icc 0 ⌊t⌋₊, c k)) := by refine Tendsto.congr' h (intervalIntegral_tendsto_integral_Ioi _ ?_ tendsto_natCast_atTop_atTop) exact integrableOn_Ici_iff_integrableOn_Ioi.mp - <| (locallyintegrablemulsum c zero_le_one hf_int).integrableOn_of_isBigO_atTop hg_dom hg_int + <| (locallyIntegrableOn_mul_sum_Icc c zero_le_one hf_int).integrableOn_of_isBigO_atTop + hg_dom hg_int refine (h_lim.sub h_lim').congr (fun _ ↦ ?_) rw [sum_mul_eq_sub_integral_mul₀' _ hc _ (fun t ht ↦ hf_diff _ ht.1) (hf_int.mono_set Set.Icc_subset_Ici_self)] @@ -325,9 +329,10 @@ private theorem summable_mul_of_bigO_atTop_aux (m : ℕ) · exact add_le_add_left (le_trans (neg_le_abs _) (Real.norm_eq_abs _ ▸ norm_integral_le_integral_norm _)) _ · refine add_le_add_left (setIntegral_mono_set ?_ ?_ Set.Ioc_subset_Ioi_self.eventuallyLE) C₁ - · refine integrableOn_Ici_iff_integrableOn_Ioi.mp <| + · exact integrableOn_Ici_iff_integrableOn_Ioi.mp <| (integrable_norm_iff h_mes.aestronglyMeasurable).mpr <| - (locallyintegrablemulsum _ m.cast_nonneg hf_int).integrableOn_of_isBigO_atTop hg₁ hg₂ + (locallyIntegrableOn_mul_sum_Icc _ m.cast_nonneg hf_int).integrableOn_of_isBigO_atTop + hg₁ hg₂ · filter_upwards with t using norm_nonneg _ theorem summable_mul_of_bigO_atTop diff --git a/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean b/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean index dde29ae2b1750..3ecf1367d697a 100644 --- a/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean +++ b/Mathlib/NumberTheory/Cyclotomic/CyclotomicCharacter.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Hanneke Wiersema. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Hanneke Wiersema -/ +import Mathlib.Algebra.Ring.Aut import Mathlib.RingTheory.RootsOfUnity.PrimitiveRoots /-! diff --git a/Mathlib/NumberTheory/DiophantineApproximation/Basic.lean b/Mathlib/NumberTheory/DiophantineApproximation/Basic.lean index 2a02d29488dbd..66f093b9e5047 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation/Basic.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation/Basic.lean @@ -550,7 +550,7 @@ This version uses `Real.convergent`. -/ theorem exists_rat_eq_convergent {q : ℚ} (h : |ξ - q| < 1 / (2 * (q.den : ℝ) ^ 2)) : ∃ n, q = ξ.convergent n := by refine q.num_div_den ▸ exists_rat_eq_convergent' ⟨?_, fun hd => ?_, ?_⟩ - · exact coprime_iff_nat_coprime.mpr (natAbs_ofNat q.den ▸ q.reduced) + · exact isCoprime_iff_nat_coprime.mpr (natAbs_ofNat q.den ▸ q.reduced) · rw [← q.den_eq_one_iff.mp (Nat.cast_eq_one.mp hd)] at h simpa only [Rat.den_intCast, Nat.cast_one, one_pow, mul_one] using (abs_lt.mp h).1 · obtain ⟨hq₀, hq₁⟩ := aux₀ (Nat.cast_pos.mpr q.pos) diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 0a26a00f55e98..4ef71ae103a18 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -240,17 +240,11 @@ theorem divisorsAntidiagonal_one : divisorsAntidiagonal 1 = {(1, 1)} := by ext simp [mul_eq_one, Prod.ext_iff] -/- Porting note: simpnf linter; added aux lemma below -Left-hand side simplifies from - Prod.swap x ∈ Nat.divisorsAntidiagonal n -to - x.snd * x.fst = n ∧ ¬n = 0-/ --- @[simp] +-- The left hand side is not in simp normal form, see the variant below. theorem swap_mem_divisorsAntidiagonal {x : ℕ × ℕ} : x.swap ∈ divisorsAntidiagonal n ↔ x ∈ divisorsAntidiagonal n := by rw [mem_divisorsAntidiagonal, mem_divisorsAntidiagonal, mul_comm, Prod.swap] --- Porting note: added below thm to replace the simp from the previous thm @[simp] theorem swap_mem_divisorsAntidiagonal_aux {x : ℕ × ℕ} : x.snd * x.fst = n ∧ ¬n = 0 ↔ x ∈ divisorsAntidiagonal n := by diff --git a/Mathlib/NumberTheory/FLT/Four.lean b/Mathlib/NumberTheory/FLT/Four.lean index 5803ad9288031..2ce906d75b4c6 100644 --- a/Mathlib/NumberTheory/FLT/Four.lean +++ b/Mathlib/NumberTheory/FLT/Four.lean @@ -80,7 +80,7 @@ theorem exists_minimal {a b c : ℤ} (h : Fermat42 a b c) : ∃ a0 b0 c0, Minima /-- a minimal solution to `a ^ 4 + b ^ 4 = c ^ 2` must have `a` and `b` coprime. -/ theorem coprime_of_minimal {a b c : ℤ} (h : Minimal a b c) : IsCoprime a b := by - apply Int.gcd_eq_one_iff_coprime.mp + apply Int.isCoprime_iff_gcd_eq_one.mpr by_contra hab obtain ⟨p, hp, hpa, hpb⟩ := Nat.Prime.not_coprime_iff_dvd.mp hab obtain ⟨a1, rfl⟩ := Int.natCast_dvd.mpr hpa @@ -119,7 +119,7 @@ theorem exists_odd_minimal {a b c : ℤ} (h : Fermat42 a b c) : · exfalso have h1 : 2 ∣ (Int.gcd a0 b0 : ℤ) := Int.dvd_gcd (Int.dvd_of_emod_eq_zero hap) (Int.dvd_of_emod_eq_zero hbp) - rw [Int.gcd_eq_one_iff_coprime.mpr (coprime_of_minimal hf)] at h1 + rw [Int.isCoprime_iff_gcd_eq_one.mp (coprime_of_minimal hf)] at h1 revert h1 decide · exact ⟨b0, ⟨a0, ⟨c0, minimal_comm hf, hbp⟩⟩⟩ @@ -139,14 +139,18 @@ theorem exists_pos_odd_minimal {a b c : ℤ} (h : Fermat42 a b c) : end Fermat42 -theorem Int.coprime_of_sq_sum {r s : ℤ} (h2 : IsCoprime s r) : IsCoprime (r ^ 2 + s ^ 2) r := by +theorem Int.isCoprime_of_sq_sum {r s : ℤ} (h2 : IsCoprime s r) : IsCoprime (r ^ 2 + s ^ 2) r := by rw [sq, sq] exact (IsCoprime.mul_left h2 h2).mul_add_left_left r -theorem Int.coprime_of_sq_sum' {r s : ℤ} (h : IsCoprime r s) : +@[deprecated (since := "2025-01-23")] alias Int.coprime_of_sq_sum := Int.isCoprime_of_sq_sum + +theorem Int.isCoprime_of_sq_sum' {r s : ℤ} (h : IsCoprime r s) : IsCoprime (r ^ 2 + s ^ 2) (r * s) := by - apply IsCoprime.mul_right (Int.coprime_of_sq_sum (isCoprime_comm.mp h)) - rw [add_comm]; apply Int.coprime_of_sq_sum h + apply IsCoprime.mul_right (Int.isCoprime_of_sq_sum (isCoprime_comm.mp h)) + rw [add_comm]; apply Int.isCoprime_of_sq_sum h + +@[deprecated (since := "2025-01-23")] alias Int.coprime_of_sq_sum' := Int.isCoprime_of_sq_sum' namespace Fermat42 @@ -160,7 +164,8 @@ theorem not_minimal {a b c : ℤ} (h : Minimal a b c) (ha2 : a % 2 = 1) (hc : 0 delta PythagoreanTriple linear_combination h.1.2.2 -- coprime requirement: - have h2 : Int.gcd (a ^ 2) (b ^ 2) = 1 := Int.gcd_eq_one_iff_coprime.mpr (coprime_of_minimal h).pow + have h2 : Int.gcd (a ^ 2) (b ^ 2) = 1 := + Int.isCoprime_iff_gcd_eq_one.mp (coprime_of_minimal h).pow -- in order to reduce the possibilities we get from the classification of pythagorean triples -- it helps if we know the parity of a ^ 2 (and the sign of c): have ha22 : a ^ 2 % 2 = 1 := by @@ -175,10 +180,10 @@ theorem not_minimal {a b c : ℤ} (h : Minimal a b c) (ha2 : a % 2 = 1) (hc : 0 linear_combination ht1 -- a and n are coprime, because a ^ 2 = m ^ 2 - n ^ 2 and m and n are coprime. have h3 : Int.gcd a n = 1 := by - apply Int.gcd_eq_one_iff_coprime.mpr + apply Int.isCoprime_iff_gcd_eq_one.mp apply @IsCoprime.of_mul_left_left _ _ _ a rw [← sq, ht1, (by ring : m ^ 2 - n ^ 2 = m ^ 2 + -n * n)] - exact (Int.gcd_eq_one_iff_coprime.mp ht4).pow_left.add_mul_right_left (-n) + exact (Int.isCoprime_iff_gcd_eq_one.mpr ht4).pow_left.add_mul_right_left (-n) -- m is positive because b is non-zero and b ^ 2 = 2 * m * n and we already have 0 ≤ m. have hb20 : b ^ 2 ≠ 0 := mt pow_eq_zero h.1.2.1 have h4 : 0 < m := by @@ -193,8 +198,8 @@ theorem not_minimal {a b c : ℤ} (h : Minimal a b c) (ha2 : a % 2 = 1) (hc : 0 -- m and r * s are coprime because m = r ^ 2 + s ^ 2 and r and s are coprime. have hcp : Int.gcd m (r * s) = 1 := by rw [htt3] - exact - Int.gcd_eq_one_iff_coprime.mpr (Int.coprime_of_sq_sum' (Int.gcd_eq_one_iff_coprime.mp htt4)) + exact Int.isCoprime_iff_gcd_eq_one.mp + (Int.isCoprime_of_sq_sum' (Int.isCoprime_iff_gcd_eq_one.mpr htt4)) -- b is even because b ^ 2 = 2 * m * n. have hb2 : 2 ∣ b := by apply @Int.Prime.dvd_pow' _ 2 _ Nat.prime_two diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index 12a27a0b7613d..5c593d0213ec5 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -306,10 +306,7 @@ private theorem psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_ · exact tsub_le_tsub_left (one_le_of_lt p_gt_two) ((b ^ 2) ^ (p - 1) * b ^ 2) · have : p ≤ p * b ^ 2 := Nat.le_mul_of_pos_right _ (show 0 < b ^ 2 by positivity) exact tsub_lt_tsub_right_of_le this h - suffices h : p < (b ^ 2) ^ (p - 1) by - have : 4 ≤ b ^ 2 := by nlinarith - have : 0 < b ^ 2 := by omega - exact mul_lt_mul_of_pos_right h this + suffices h : p < (b ^ 2) ^ (p - 1) by gcongr rw [← pow_mul, Nat.mul_sub_left_distrib, mul_one] have : 2 ≤ 2 * p - 2 := le_tsub_of_add_le_left (show 4 ≤ 2 * p by omega) have : 2 + p ≤ 2 * p := by omega diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index 3faf700efb403..3f0a95887411d 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -68,6 +68,12 @@ lemma term_def (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : term f s n = if n = 0 then 0 else f n / n ^ s := rfl +/-- An alternate spelling of `term_def` for the case `f 0 = 0`. -/ +lemma term_def₀ {f : ℕ → ℂ} (hf : f 0 = 0) (s : ℂ) (n : ℕ) : + LSeries.term f s n = f n * (n : ℂ) ^ (- s) := by + rw [LSeries.term] + split_ifs with h <;> simp [h, hf, cpow_neg, div_eq_inv_mul, mul_comm] + @[simp] lemma term_zero (f : ℕ → ℂ) (s : ℂ) : term f s 0 = 0 := rfl diff --git a/Mathlib/NumberTheory/LSeries/SumCoeff.lean b/Mathlib/NumberTheory/LSeries/SumCoeff.lean new file mode 100644 index 0000000000000..35a11908fb0fe --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/SumCoeff.lean @@ -0,0 +1,154 @@ +/- +Copyright (c) 2025 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.Analysis.Asymptotics.SpecificAsymptotics +import Mathlib.Analysis.InnerProductSpace.Calculus +import Mathlib.Analysis.SpecialFunctions.ImproperIntegrals +import Mathlib.NumberTheory.AbelSummation +import Mathlib.NumberTheory.LSeries.Basic + +/-! +# Partial sums of coefficients of L-series + +We prove several results involving partial sums of coefficients (or norm of coefficients) of +L-series. + +## Main results + +* `LSeriesSummable_of_sum_norm_bigO`: for `f : ℕ → ℂ`, if the partial sums + `∑ k ∈ Icc 1 n, ‖f k‖` are `O(n ^ r)` for some real `0 ≤ r`, then the L-series `LSeries f` + converges at `s : ℂ` for all `s` such that `r < s.re`. + +* `LSeries_eq_mul_integral` : for `f : ℕ → ℂ`, if the partial sums `∑ k ∈ Icc 1 n, f k` are + `O(n ^ r)` for some real `0 ≤ r` and the L-series `LSeries f` converges at `s : ℂ` with + `r < s.re`, then `LSeries f s = s * ∫ t in Set.Ioi 1, (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1))`. + +-/ + +open Finset Filter MeasureTheory Topology Complex Asymptotics + +section summable + +variable {f : ℕ → ℂ} {r : ℝ} {s : ℂ} + +private theorem LSeriesSummable_of_sum_norm_bigO_aux (hf : f 0 = 0) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, ‖f k‖) =O[atTop] fun n ↦ (n : ℝ) ^ r) + (hr : 0 ≤ r) (hs : r < s.re) : + LSeriesSummable f s := by + have h₁ : -s ≠ 0 := neg_ne_zero.mpr <| ne_zero_of_re_pos (hr.trans_lt hs) + have h₂ : (-s).re + r ≤ 0 := by + rw [neg_re, neg_add_nonpos_iff] + exact hs.le + have h₃ (t : ℝ) (ht : t ∈ Set.Ici 1) : DifferentiableAt ℝ (fun x : ℝ ↦ ‖(x : ℂ) ^ (-s)‖) t := + have ht' : t ≠ 0 := (zero_lt_one.trans_le ht).ne' + (differentiableAt_id.ofReal_cpow_const ht' h₁).norm ℝ <| + (cpow_ne_zero_iff_of_exponent_ne_zero h₁).mpr <| ofReal_ne_zero.mpr ht' + have h₄ : (deriv fun t : ℝ ↦ ‖(t : ℂ) ^ (-s)‖) =ᶠ[atTop] fun t ↦ -s.re * t ^ (-(s.re +1)) := by + filter_upwards [eventually_gt_atTop 0] with t ht + rw [deriv_norm_ofReal_cpow _ ht, neg_re, neg_add'] + simp_rw [LSeriesSummable, funext (LSeries.term_def₀ hf s), mul_comm (f _)] + refine summable_mul_of_bigO_atTop' (f := fun t ↦ (t : ℂ) ^ (-s)) + (g := fun t ↦ t ^ (-(s.re + 1) + r)) _ h₃ ?_ ?_ ?_ ?_ + · refine integrableOn_Ici_iff_integrableOn_Ioi.mpr + (integrableOn_Ioi_deriv_norm_ofReal_cpow zero_lt_one ?_) + exact neg_re _ ▸ neg_nonpos.mpr <| hr.trans hs.le + · refine (IsBigO.mul_atTop_rpow_natCast_of_isBigO_rpow _ _ _ ?_ hO h₂).congr_right (by simp) + exact (norm_ofReal_cpow_eventually_eq_atTop _).isBigO.natCast_atTop + · refine h₄.isBigO.of_const_mul_right.mul_atTop_rpow_of_isBigO_rpow _ r _ ?_ le_rfl + exact (hO.comp_tendsto tendsto_nat_floor_atTop).trans <| + isEquivalent_nat_floor.isBigO.rpow hr (eventually_ge_atTop 0) + · rwa [integrableAtFilter_rpow_atTop_iff, neg_add_lt_iff_lt_add, add_neg_cancel_right] + +/-- If the partial sums `∑ k ∈ Icc 1 n, ‖f k‖` are `O(n ^ r)` for some real `0 ≤ r`, then the +L-series `LSeries f` converges at `s : ℂ` for all `s` such that `r < s.re`. -/ +theorem LSeriesSummable_of_sum_norm_bigO + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, ‖f k‖) =O[atTop] fun n ↦ (n : ℝ) ^ r) + (hr : 0 ≤ r) (hs : r < s.re) : + LSeriesSummable f s := by + have h₁ : (fun n ↦ if n = 0 then 0 else f n) =ᶠ[atTop] f := by + filter_upwards [eventually_ne_atTop 0] with n hn using by simp_rw [if_neg hn] + refine (LSeriesSummable_of_sum_norm_bigO_aux (if_pos rfl) ?_ hr hs).congr' _ h₁ + refine hO.congr' (Eventually.of_forall fun _ ↦ Finset.sum_congr rfl fun _ h ↦ ?_) EventuallyEq.rfl + rw [if_neg (zero_lt_one.trans_le (mem_Icc.mp h).1).ne'] + +/-- If `f` takes nonnegative real values and the partial sums `∑ k ∈ Icc 1 n, f k` are `O(n ^ r)` +for some real `0 ≤ r`, then the L-series `LSeries f` converges at `s : ℂ` for all `s` +such that `r < s.re`. -/ +theorem LSeriesSummable_of_sum_norm_bigO_and_nonneg + {f : ℕ → ℝ} (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) + (hf : ∀ n, 0 ≤ f n) (hr : 0 ≤ r) (hs : r < s.re) : + LSeriesSummable (fun n ↦ f n) s := + LSeriesSummable_of_sum_norm_bigO (by simpa [_root_.abs_of_nonneg (hf _)]) hr hs + +end summable + +section integralrepresentation + +private theorem LSeries_eq_mul_integral_aux {f : ℕ → ℂ} (hf : f 0 = 0) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} + (hs : r < s.re) (hS : LSeriesSummable f s) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) : + LSeries f s = s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (- (s + 1)) := by + have h₁ : (-s - 1).re + r < -1 := by + rwa [sub_re, one_re, neg_re, neg_sub_left, neg_add_lt_iff_lt_add, add_neg_cancel_comm] + have h₂ : s ≠ 0 := ne_zero_of_re_pos (hr.trans_lt hs) + have h₃ (t : ℝ) (ht : t ∈ Set.Ici 1) : DifferentiableAt ℝ (fun x : ℝ ↦ (x : ℂ) ^ (-s)) t := + differentiableAt_id.ofReal_cpow_const (zero_lt_one.trans_le ht).ne' (neg_ne_zero.mpr h₂) + have h₄ : ∀ n, ∑ k ∈ Icc 0 n, f k = ∑ k ∈ Icc 1 n, f k := fun n ↦ by + rw [← Nat.Icc_insert_succ_left n.zero_le, sum_insert (by aesop), hf, zero_add, zero_add] + simp_rw [← h₄] at hO + rw [← integral_mul_left] + refine tendsto_nhds_unique ((tendsto_add_atTop_iff_nat 1).mpr hS.hasSum.tendsto_sum_nat) ?_ + simp_rw [Nat.range_succ_eq_Icc_zero, LSeries.term_def₀ hf, mul_comm (f _)] + convert tendsto_sum_mul_atTop_nhds_one_sub_integral₀ (f := fun x ↦ (x : ℂ) ^ (-s)) (l := 0) + ?_ hf h₃ ?_ ?_ ?_ (integrableAtFilter_rpow_atTop_iff.mpr h₁) + · rw [zero_sub, ← integral_neg] + refine setIntegral_congr_fun measurableSet_Ioi fun t ht ↦ ?_ + rw [deriv_ofReal_cpow_const (zero_lt_one.trans ht).ne', h₄] + · ring_nf + · exact neg_ne_zero.mpr <| ne_zero_of_re_pos (hr.trans_lt hs) + · refine integrableOn_Ici_iff_integrableOn_Ioi.mpr <| + integrableOn_Ioi_deriv_ofReal_cpow zero_lt_one (by simpa using hr.trans_lt hs) + · have hlim : Tendsto (fun n : ℕ ↦ (n : ℝ) ^ (-(s.re - r))) atTop (𝓝 0) := + (tendsto_rpow_neg_atTop (by rwa [sub_pos])).comp tendsto_natCast_atTop_atTop + refine (IsBigO.mul_atTop_rpow_natCast_of_isBigO_rpow (-s.re) _ _ ?_ hO ?_).trans_tendsto hlim + · exact isBigO_norm_left.mp <| (norm_ofReal_cpow_eventually_eq_atTop _).isBigO.natCast_atTop + · linarith + · refine .mul_atTop_rpow_of_isBigO_rpow (-(s + 1).re) r _ ?_ ?_ (by rw [← neg_re, neg_add']) + · simpa [- neg_add_rev, neg_add'] using isBigO_deriv_ofReal_cpow_const_atTop _ + · exact (hO.comp_tendsto tendsto_nat_floor_atTop).trans <| + isEquivalent_nat_floor.isBigO.rpow hr (eventually_ge_atTop 0) + +/-- If the partial sums `∑ k ∈ Icc 1 n, f k` are `O(n ^ r)` for some real `0 ≤ r` and the +L-series `LSeries f` converges at `s : ℂ` with `r < s.re`, then +`LSeries f s = s * ∫ t in Set.Ioi 1, (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1))`. -/ +theorem LSeries_eq_mul_integral (f : ℕ → ℂ) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} (hs : r < s.re) + (hS : LSeriesSummable f s) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) : + LSeries f s = s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1)) := by + rw [← LSeriesSummable_congr' s (f := fun n ↦ if n = 0 then 0 else f n) + (by filter_upwards [eventually_ne_atTop 0] with n h using if_neg h)] at hS + have (n) : ∑ k ∈ Icc 1 n, (if k = 0 then 0 else f k) = ∑ k ∈ Icc 1 n, f k := + Finset.sum_congr rfl fun k hk ↦ by rw [if_neg (zero_lt_one.trans_le (mem_Icc.mp hk).1).ne'] + rw [← LSeries_congr _ (fun _ ↦ if_neg _), LSeries_eq_mul_integral_aux (if_pos rfl) hr hs hS] <;> + simp_all + +/-- A version of `LSeries_eq_mul_integral` where we use the stronger condition that the partial sums +`∑ k ∈ Icc 1 n, ‖f k‖` are `O(n ^ r)` to deduce the integral representation. -/ +theorem LSeries_eq_mul_integral' (f : ℕ → ℂ) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} (hs : r < s.re) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, ‖f k‖) =O[atTop] fun n ↦ (n : ℝ) ^ r) : + LSeries f s = s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1)) := + LSeries_eq_mul_integral _ hr hs (LSeriesSummable_of_sum_norm_bigO hO hr hs) <| + (isBigO_of_le _ fun _ ↦ (norm_sum_le _ _).trans <| Real.le_norm_self _).trans hO + +/-- If `f` takes nonnegative real values and the partial sums `∑ k ∈ Icc 1 n, f k` are `O(n ^ r)` +for some real `0 ≤ r`, then for `s : ℂ` with `r < s.re`, we have +`LSeries f s = s * ∫ t in Set.Ioi 1, (∑ k ∈ Icc 1 ⌊t⌋₊, f k) * t ^ (-(s + 1))`. -/ +theorem LSeries_eq_mul_integral_of_nonneg (f : ℕ → ℝ) {r : ℝ} (hr : 0 ≤ r) {s : ℂ} (hs : r < s.re) + (hO : (fun n ↦ ∑ k ∈ Icc 1 n, f k) =O[atTop] fun n ↦ (n : ℝ) ^ r) (hf : ∀ n, 0 ≤ f n) : + LSeries (fun n ↦ f n) s = + s * ∫ t in Set.Ioi (1 : ℝ), (∑ k ∈ Icc 1 ⌊t⌋₊, (f k : ℂ)) * t ^ (-(s + 1)) := + LSeries_eq_mul_integral' _ hr hs <| hO.congr_left fun _ ↦ by simp [_root_.abs_of_nonneg (hf _)] + +end integralrepresentation diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index 5fd3caaed9a9c..acf857f2bf3e1 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -92,9 +92,6 @@ lemma not_isPrimitive_mulShift [Finite R] (e : AddChar R R') {r : R} /-- Definition for a primitive additive character on a finite ring `R` into a cyclotomic extension of a field `R'`. It records which cyclotomic extension it is, the character, and the fact that the character is primitive. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- can't prove that they always exist (referring to providing an `Inhabited` instance) --- @[nolint has_nonempty_instance] structure PrimitiveAddChar (R : Type u) [CommRing R] (R' : Type v) [Field R'] where /-- The first projection from `PrimitiveAddChar`, giving the cyclotomic field. -/ n : ℕ+ diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 030cc20bab5bb..8368544665f34 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -160,7 +160,7 @@ theorem quadraticChar_eq_neg_one_iff_not_one {a : F} (ha : a ≠ 0) : /-- For `a : F`, `quadraticChar F a = -1 ↔ ¬ IsSquare a`. -/ theorem quadraticChar_neg_one_iff_not_isSquare {a : F} : quadraticChar F a = -1 ↔ ¬IsSquare a := by by_cases ha : a = 0 - · simp only [ha, MulChar.map_zero, zero_eq_neg, one_ne_zero, isSquare_zero, not_true] + · simp only [ha, MulChar.map_zero, zero_eq_neg, one_ne_zero, IsSquare.zero, not_true] · rw [quadraticChar_eq_neg_one_iff_not_one ha, quadraticChar_one_iff_isSquare ha] /-- If `F` has odd characteristic, then `quadraticChar F` takes the value `-1`. -/ diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean index b10a0d651815b..29fa07a6854fe 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean @@ -102,7 +102,7 @@ theorem FiniteField.isSquare_odd_prime_iff (hF : ringChar F ≠ 2) {p : ℕ} [Fa classical by_cases hFp : ringChar F = p · rw [show (p : F) = 0 by rw [← hFp]; exact ringChar.Nat.cast_ringChar] - simp only [isSquare_zero, Ne, true_iff, map_mul] + simp only [IsSquare.zero, Ne, true_iff, map_mul] obtain ⟨n, _, hc⟩ := FiniteField.card F (ringChar F) have hchar : ringChar F = ringChar (ZMod p) := by rw [hFp]; exact (ringChar_zmod_n p).symm conv => enter [1, 1, 2]; rw [hc, Nat.cast_pow, map_pow, hchar, map_ringChar] diff --git a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean index ee91db811cdfa..cf81fe88e9ba2 100644 --- a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean +++ b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean @@ -67,17 +67,7 @@ theorem Gamma_one_top : Gamma 1 = ⊤ := by lemma mem_Gamma_one (γ : SL(2, ℤ)) : γ ∈ Γ(1) := by simp only [Gamma_one_top, Subgroup.mem_top] -theorem Gamma_zero_bot : Gamma 0 = ⊥ := by - ext - simp only [Gamma_mem, coe_matrix_coe, Int.coe_castRingHom, map_apply, Int.cast_id, - Subgroup.mem_bot] - constructor - · intro h - ext i j - fin_cases i <;> fin_cases j <;> simp only [h] - exacts [h.1, h.2.1, h.2.2.1, h.2.2.2] - · intro h - simp [h] +theorem Gamma_zero_bot : Gamma 0 = ⊥ := rfl lemma ModularGroup_T_pow_mem_Gamma (N M : ℤ) (hNM : N ∣ M) : (ModularGroup.T ^ M) ∈ Gamma (Int.natAbs N) := by diff --git a/Mathlib/NumberTheory/ModularForms/QExpansion.lean b/Mathlib/NumberTheory/ModularForms/QExpansion.lean index 27d4e3389e4bb..43be8d61e52f2 100644 --- a/Mathlib/NumberTheory/ModularForms/QExpansion.lean +++ b/Mathlib/NumberTheory/ModularForms/QExpansion.lean @@ -149,7 +149,7 @@ lemma qExpansionFormalMultilinearSeries_apply_norm (m : ℕ) : lemma qExpansionFormalMultilinearSeries_radius [NeZero n] [ModularFormClass F Γ(n) k] : 1 ≤ (qExpansionFormalMultilinearSeries n f).radius := by - refine le_of_forall_ge_of_dense fun r hr ↦ ?_ + refine le_of_forall_lt_imp_le_of_dense fun r hr ↦ ?_ lift r to NNReal using hr.ne_top apply FormalMultilinearSeries.le_radius_of_summable simp only [qExpansionFormalMultilinearSeries_apply_norm] diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index c057b1aa09be8..b14ba38413d3d 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -881,7 +881,7 @@ lemma even_nat_card_aut_of_not_isUnramified [IsGalois k K] (hw : ¬ IsUnramified · cases nonempty_fintype (K ≃ₐ[k] K) rw [even_iff_two_dvd, ← not_isUnramified_iff_card_stabilizer_eq_two.mp hw] exact Subgroup.card_subgroup_dvd_card (Stab w) - · convert even_zero + · convert Even.zero by_contra e exact H (Nat.finite_of_card_ne_zero e) @@ -894,7 +894,7 @@ lemma even_finrank_of_not_isUnramified [IsGalois k K] (hw : ¬ IsUnramified k w) : Even (finrank k K) := by by_cases FiniteDimensional k K · exact IsGalois.card_aut_eq_finrank k K ▸ even_card_aut_of_not_isUnramified hw - · exact finrank_of_not_finite ‹_› ▸ even_zero + · exact finrank_of_not_finite ‹_› ▸ Even.zero lemma isUnramified_smul_iff : IsUnramified k (σ • w) ↔ IsUnramified k w := by diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index b3c0223113d11..11aa143aeb098 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -3,8 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Order.Ring.Basic -import Mathlib.Data.Nat.ModEq +import Mathlib.Data.Nat.Prime.Basic import Mathlib.NumberTheory.Zsqrtd.Basic /-! diff --git a/Mathlib/NumberTheory/SumPrimeReciprocals.lean b/Mathlib/NumberTheory/SumPrimeReciprocals.lean index b7c071799368c..cc8ba145bd698 100644 --- a/Mathlib/NumberTheory/SumPrimeReciprocals.lean +++ b/Mathlib/NumberTheory/SumPrimeReciprocals.lean @@ -3,8 +3,9 @@ Copyright (c) 2023 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ -import Mathlib.NumberTheory.SmoothNumbers +import Mathlib.Algebra.Order.Group.Indicator import Mathlib.Analysis.PSeries +import Mathlib.NumberTheory.SmoothNumbers /-! # The sum of the reciprocals of the primes diverges diff --git a/Mathlib/NumberTheory/SumTwoSquares.lean b/Mathlib/NumberTheory/SumTwoSquares.lean index ae1099ea7d72e..dcc58f8d748e3 100644 --- a/Mathlib/NumberTheory/SumTwoSquares.lean +++ b/Mathlib/NumberTheory/SumTwoSquares.lean @@ -200,7 +200,7 @@ the right hand side holds, since `padicValNat q 0 = 0` by definition.) -/ theorem Nat.eq_sq_add_sq_iff {n : ℕ} : (∃ x y : ℕ, n = x ^ 2 + y ^ 2) ↔ ∀ {q : ℕ}, q.Prime → q % 4 = 3 → Even (padicValNat q n) := by rcases n.eq_zero_or_pos with (rfl | hn₀) - · exact ⟨fun _ q _ _ => (@padicValNat.zero q).symm ▸ even_zero, fun _ => ⟨0, 0, rfl⟩⟩ + · exact ⟨fun _ q _ _ => (@padicValNat.zero q).symm ▸ Even.zero, fun _ => ⟨0, 0, rfl⟩⟩ -- now `0 < n` rw [Nat.eq_sq_add_sq_iff_eq_sq_mul] refine ⟨fun H q hq h => ?_, fun H => ?_⟩ diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 1babdaa7537b8..00bc9cbbccdaa 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -5,9 +5,9 @@ Authors: Mario Carneiro -/ import Mathlib.Algebra.Ring.Associated import Mathlib.Algebra.Star.Unitary -import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Tactic.Ring +import Mathlib.Algebra.EuclideanDomain.Int /-! # ℤ[√d] @@ -326,7 +326,7 @@ theorem gcd_eq_zero_iff (a : ℤ√d) : Int.gcd a.re a.im = 0 ↔ a = 0 := by theorem gcd_pos_iff (a : ℤ√d) : 0 < Int.gcd a.re a.im ↔ a ≠ 0 := pos_iff_ne_zero.trans <| not_congr a.gcd_eq_zero_iff -theorem coprime_of_dvd_coprime {a b : ℤ√d} (hcoprime : IsCoprime a.re a.im) (hdvd : b ∣ a) : +theorem isCoprime_of_dvd_isCoprime {a b : ℤ√d} (hcoprime : IsCoprime a.re a.im) (hdvd : b ∣ a) : IsCoprime b.re b.im := by apply isCoprime_of_dvd · rintro ⟨hre, him⟩ @@ -342,13 +342,15 @@ theorem coprime_of_dvd_coprime {a b : ℤ√d} (hcoprime : IsCoprime a.re a.im) exact ⟨hzdvdu, hzdvdv⟩ exact hcoprime.isUnit_of_dvd' ha hb +@[deprecated (since := "2025-01-23")] alias coprime_of_dvd_coprime := isCoprime_of_dvd_isCoprime + theorem exists_coprime_of_gcd_pos {a : ℤ√d} (hgcd : 0 < Int.gcd a.re a.im) : ∃ b : ℤ√d, a = ((Int.gcd a.re a.im : ℤ) : ℤ√d) * b ∧ IsCoprime b.re b.im := by obtain ⟨re, im, H1, Hre, Him⟩ := Int.exists_gcd_one hgcd rw [mul_comm] at Hre Him refine ⟨⟨re, im⟩, ?_, ?_⟩ · rw [smul_val, ← Hre, ← Him] - · rw [← Int.gcd_eq_one_iff_coprime, H1] + · rw [Int.isCoprime_iff_gcd_eq_one, H1] end Gcd diff --git a/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean b/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean index a2da9ffabe073..b3ef1877a2913 100644 --- a/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean +++ b/Mathlib/NumberTheory/Zsqrtd/GaussianInt.lean @@ -3,10 +3,10 @@ Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.NumberTheory.Zsqrtd.Basic -import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Data.Complex.Basic +import Mathlib.Data.Nat.Prime.Basic import Mathlib.Data.Real.Archimedean +import Mathlib.NumberTheory.Zsqrtd.Basic /-! # Gaussian integers diff --git a/Mathlib/Order/Antichain.lean b/Mathlib/Order/Antichain.lean index df46cd2e7a882..307d08ef8455c 100644 --- a/Mathlib/Order/Antichain.lean +++ b/Mathlib/Order/Antichain.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Set.Pairwise.Basic import Mathlib.Order.Bounds.Basic import Mathlib.Order.Directed import Mathlib.Order.Hom.Set +import Mathlib.Order.Chain /-! # Antichains @@ -172,6 +173,28 @@ theorem isAntichain_singleton (a : α) (r : α → α → Prop) : IsAntichain r theorem Set.Subsingleton.isAntichain (hs : s.Subsingleton) (r : α → α → Prop) : IsAntichain r s := hs.pairwise _ +/-- A set which is simultaneously a chain and antichain is subsingleton. -/ +lemma subsingleton_of_isChain_of_isAntichain (hs : IsChain r s) (ht : IsAntichain r s) : + s.Subsingleton := by + intro x hx y hy + by_contra! hne + cases hs hx hy hne with + | inl h => exact ht hx hy hne h + | inr h => exact ht hy hx hne.symm h + +lemma isChain_and_isAntichain_iff_subsingleton : IsChain r s ∧ IsAntichain r s ↔ s.Subsingleton := + ⟨fun h ↦ subsingleton_of_isChain_of_isAntichain h.1 h.2, fun h ↦ ⟨h.isChain, h.isAntichain _⟩⟩ + +/-- The intersection of a chain and an antichain is subsingleton. -/ +lemma inter_subsingleton_of_isChain_of_isAntichain (hs : IsChain r s) (ht : IsAntichain r t) : + (s ∩ t).Subsingleton := + subsingleton_of_isChain_of_isAntichain (hs.mono (by simp)) (ht.subset (by simp)) + +/-- The intersection of an antichain and a chain is subsingleton. -/ +lemma inter_subsingleton_of_isAntichain_of_isChain (hs : IsAntichain r s) (ht : IsChain r t) : + (s ∩ t).Subsingleton := + inter_comm _ _ ▸ inter_subsingleton_of_isChain_of_isAntichain ht hs + section Preorder variable [Preorder α] diff --git a/Mathlib/Order/Antisymmetrization.lean b/Mathlib/Order/Antisymmetrization.lean index 2a5cced208b85..1ec2f0a774711 100644 --- a/Mathlib/Order/Antisymmetrization.lean +++ b/Mathlib/Order/Antisymmetrization.lean @@ -3,8 +3,9 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Order.Hom.Basic import Mathlib.Logic.Relation +import Mathlib.Order.Hom.Basic +import Mathlib.Tactic.Tauto /-! # Turning a preorder into a partial order @@ -27,43 +28,62 @@ such that `a ≤ b` and `b ≤ a`. open Function OrderDual -variable {α β : Type*} +variable {α β : Type*} {a b c d : α} section Relation variable (r : α → α → Prop) -/-- The antisymmetrization relation. -/ +/-- The antisymmetrization relation `AntisymmRel r` is defined so that +`AntisymmRel r a b ↔ r a b ∧ r b a`. -/ def AntisymmRel (a b : α) : Prop := r a b ∧ r b a theorem antisymmRel_swap : AntisymmRel (swap r) = AntisymmRel r := - funext fun _ => funext fun _ => propext and_comm + funext₂ fun _ _ ↦ propext and_comm @[refl] theorem antisymmRel_refl [IsRefl α r] (a : α) : AntisymmRel r a a := ⟨refl _, refl _⟩ +variable {r} in +lemma AntisymmRel.rfl [IsRefl α r] (a : α) : AntisymmRel r a a := antisymmRel_refl .. + +instance [IsRefl α r] : IsRefl α (AntisymmRel r) where + refl := antisymmRel_refl r + variable {r} @[symm] -theorem AntisymmRel.symm {a b : α} : AntisymmRel r a b → AntisymmRel r b a := +theorem AntisymmRel.symm : AntisymmRel r a b → AntisymmRel r b a := And.symm +instance : IsSymm α (AntisymmRel r) where + symm _ _ := AntisymmRel.symm + +theorem antisymmRel_comm : AntisymmRel r a b ↔ AntisymmRel r b a := + And.comm + @[trans] -theorem AntisymmRel.trans [IsTrans α r] {a b c : α} (hab : AntisymmRel r a b) - (hbc : AntisymmRel r b c) : AntisymmRel r a c := +theorem AntisymmRel.trans [IsTrans α r] (hab : AntisymmRel r a b) (hbc : AntisymmRel r b c) : + AntisymmRel r a c := ⟨_root_.trans hab.1 hbc.1, _root_.trans hbc.2 hab.2⟩ -instance AntisymmRel.decidableRel [DecidableRel r] : DecidableRel (AntisymmRel r) := fun _ _ => - instDecidableAnd +instance [IsTrans α r] : IsTrans α (AntisymmRel r) where + trans _ _ _ := .trans + +instance AntisymmRel.decidableRel [DecidableRel r] : DecidableRel (AntisymmRel r) := + fun _ _ ↦ instDecidableAnd @[simp] -theorem antisymmRel_iff_eq [IsRefl α r] [IsAntisymm α r] {a b : α} : AntisymmRel r a b ↔ a = b := +theorem antisymmRel_iff_eq [IsRefl α r] [IsAntisymm α r] : AntisymmRel r a b ↔ a = b := antisymm_iff alias ⟨AntisymmRel.eq, _⟩ := antisymmRel_iff_eq +theorem AntisymmRel.le [LE α] (h : AntisymmRel (· ≤ ·) a b) : a ≤ b := h.1 +theorem AntisymmRel.ge [LE α] (h : AntisymmRel (· ≤ ·) a b) : b ≤ a := h.2 + end Relation section IsPreorder @@ -112,9 +132,82 @@ end IsPreorder section Preorder -variable [Preorder α] [Preorder β] {a b : α} +variable [Preorder α] [Preorder β] + +theorem le_iff_lt_or_antisymmRel : a ≤ b ↔ a < b ∨ AntisymmRel (· ≤ ·) a b := by + rw [lt_iff_le_not_le, AntisymmRel] + tauto + +@[trans] +theorem le_of_le_of_antisymmRel (h₁ : a ≤ b) (h₂ : AntisymmRel (· ≤ ·) b c) : a ≤ c := + h₁.trans h₂.le + +@[trans] +theorem le_of_antisymmRel_of_le (h₁ : AntisymmRel (· ≤ ·) a b) (h₂ : b ≤ c) : a ≤ c := + h₁.le.trans h₂ + +@[trans] +theorem lt_of_lt_of_antisymmRel (h₁ : a < b) (h₂ : AntisymmRel (· ≤ ·) b c) : a < c := + h₁.trans_le h₂.le + +@[trans] +theorem lt_of_antisymmRel_of_lt (h₁ : AntisymmRel (· ≤ ·) a b) (h₂ : b < c) : a < c := + h₁.le.trans_lt h₂ + +alias ⟨LE.le.lt_or_antisymmRel, _⟩ := le_iff_lt_or_antisymmRel +alias LE.le.trans_antisymmRel := le_of_le_of_antisymmRel +alias AntisymmRel.trans_le := le_of_antisymmRel_of_le +alias LT.lt.trans_antisymmRel := lt_of_lt_of_antisymmRel +alias AntisymmRel.trans_lt := lt_of_antisymmRel_of_lt + +instance : @Trans α α α (· ≤ ·) (AntisymmRel (· ≤ ·)) (· ≤ ·) where + trans := le_of_le_of_antisymmRel + +instance : @Trans α α α (AntisymmRel (· ≤ ·)) (· ≤ ·) (· ≤ ·) where + trans := le_of_antisymmRel_of_le + +instance : @Trans α α α (· < ·) (AntisymmRel (· ≤ ·)) (· < ·) where + trans := lt_of_lt_of_antisymmRel + +instance : @Trans α α α (AntisymmRel (· ≤ ·)) (· < ·) (· < ·) where + trans := lt_of_antisymmRel_of_lt + +theorem AntisymmRel.le_congr (h₁ : AntisymmRel (· ≤ ·) a b) (h₂ : AntisymmRel (· ≤ ·) c d) : + a ≤ c ↔ b ≤ d where + mp h := (h₁.symm.trans_le h).trans_antisymmRel h₂ + mpr h := (h₁.trans_le h).trans_antisymmRel h₂.symm + +theorem AntisymmRel.le_congr_left (h : AntisymmRel (· ≤ ·) a b) : a ≤ c ↔ b ≤ c := + h.le_congr (antisymmRel_refl _ c) + +theorem AntisymmRel.le_congr_right (h : AntisymmRel (· ≤ ·) b c) : a ≤ b ↔ a ≤ c := + (antisymmRel_refl _ a).le_congr h + +theorem AntisymmRel.lt_congr (h₁ : AntisymmRel (· ≤ ·) a b) (h₂ : AntisymmRel (· ≤ ·) c d) : + a < c ↔ b < d where + mp h := (h₁.symm.trans_lt h).trans_antisymmRel h₂ + mpr h := (h₁.trans_lt h).trans_antisymmRel h₂.symm + +theorem AntisymmRel.lt_congr_left (h : AntisymmRel (· ≤ ·) a b) : a < c ↔ b < c := + h.lt_congr (antisymmRel_refl _ c) + +theorem AntisymmRel.lt_congr_right (h : AntisymmRel (· ≤ ·) b c) : a < b ↔ a < c := + (antisymmRel_refl _ a).lt_congr h + +theorem AntisymmRel.antisymmRel_congr + (h₁ : AntisymmRel (· ≤ ·) a b) (h₂ : AntisymmRel (· ≤ ·) c d) : + AntisymmRel (· ≤ ·) a c ↔ AntisymmRel (· ≤ ·) b d := + rel_congr h₁ h₂ + +theorem AntisymmRel.antisymmRel_congr_left (h : AntisymmRel (· ≤ ·) a b) : + AntisymmRel (· ≤ ·) a c ↔ AntisymmRel (· ≤ ·) b c := + rel_congr_left h + +theorem AntisymmRel.antisymmRel_congr_right (h : AntisymmRel (· ≤ ·) b c) : + AntisymmRel (· ≤ ·) a b ↔ AntisymmRel (· ≤ ·) a c := + rel_congr_right h -theorem AntisymmRel.image {a b : α} (h : AntisymmRel (· ≤ ·) a b) {f : α → β} (hf : Monotone f) : +theorem AntisymmRel.image (h : AntisymmRel (· ≤ ·) a b) {f : α → β} (hf : Monotone f) : AntisymmRel (· ≤ ·) (f a) (f b) := ⟨hf h.1, hf h.2⟩ diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 0f96648eca68c..11f4a94f61cf1 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -79,6 +79,9 @@ theorem lt_of_le_of_lt' : b ≤ c → a < b → a < c := theorem lt_of_lt_of_le' : b < c → a ≤ b → a < c := flip lt_of_le_of_lt +theorem not_lt_iff_not_le_or_ge : ¬a < b ↔ ¬a ≤ b ∨ b ≤ a := by + rw [lt_iff_le_not_le, Classical.not_and_iff_or_not_not, Classical.not_not] + end Preorder section PartialOrder @@ -1260,32 +1263,47 @@ instance [∀ i, Preorder (π i)] [∀ i, DenselyOrdered (π i)] : ⟨le_update_iff.2 ⟨ha.le, fun _ _ ↦ le_rfl⟩, i, by rwa [update_self]⟩, update_le_iff.2 ⟨hb.le, fun _ _ ↦ hab _⟩, i, by rwa [update_self]⟩⟩ -theorem le_of_forall_le_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} - (h : ∀ a, a₂ < a → a₁ ≤ a) : a₁ ≤ a₂ := +section LinearOrder +variable [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} + +theorem le_of_forall_gt_imp_ge_of_dense (h : ∀ a, a₂ < a → a₁ ≤ a) : a₁ ≤ a₂ := le_of_not_gt fun ha ↦ let ⟨a, ha₁, ha₂⟩ := exists_between ha lt_irrefl a <| lt_of_lt_of_le ‹a < a₁› (h _ ‹a₂ < a›) -theorem eq_of_le_of_forall_le_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} (h₁ : a₂ ≤ a₁) - (h₂ : ∀ a, a₂ < a → a₁ ≤ a) : a₁ = a₂ := - le_antisymm (le_of_forall_le_of_dense h₂) h₁ +lemma forall_gt_imp_ge_iff_le_of_dense : (∀ a, a₂ < a → a₁ ≤ a) ↔ a₁ ≤ a₂ := + ⟨le_of_forall_gt_imp_ge_of_dense, fun ha _a ha₂ ↦ ha.trans ha₂.le⟩ -theorem le_of_forall_ge_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} - (h : ∀ a₃ < a₁, a₃ ≤ a₂) : a₁ ≤ a₂ := +lemma eq_of_le_of_forall_lt_imp_le_of_dense (h₁ : a₂ ≤ a₁) (h₂ : ∀ a, a₂ < a → a₁ ≤ a) : a₁ = a₂ := + le_antisymm (le_of_forall_gt_imp_ge_of_dense h₂) h₁ + +theorem le_of_forall_lt_imp_le_of_dense (h : ∀ a < a₁, a ≤ a₂) : a₁ ≤ a₂ := le_of_not_gt fun ha ↦ let ⟨a, ha₁, ha₂⟩ := exists_between ha lt_irrefl a <| lt_of_le_of_lt (h _ ‹a < a₁›) ‹a₂ < a› -theorem eq_of_le_of_forall_ge_of_dense [LinearOrder α] [DenselyOrdered α] {a₁ a₂ : α} (h₁ : a₂ ≤ a₁) - (h₂ : ∀ a₃ < a₁, a₃ ≤ a₂) : a₁ = a₂ := - (le_of_forall_ge_of_dense h₂).antisymm h₁ +lemma forall_lt_imp_le_iff_le_of_dense : (∀ a < a₁, a ≤ a₂) ↔ a₁ ≤ a₂ := + ⟨le_of_forall_lt_imp_le_of_dense, fun ha _a ha₁ ↦ ha₁.le.trans ha⟩ + +theorem eq_of_le_of_forall_gt_imp_ge_of_dense (h₁ : a₂ ≤ a₁) (h₂ : ∀ a < a₁, a ≤ a₂) : a₁ = a₂ := + (le_of_forall_lt_imp_le_of_dense h₂).antisymm h₁ + +@[deprecated (since := "2025-01-21")] +alias le_of_forall_le_of_dense := le_of_forall_gt_imp_ge_of_dense + +@[deprecated (since := "2025-01-21")] +alias le_of_forall_ge_of_dense := le_of_forall_lt_imp_le_of_dense + +@[deprecated (since := "2025-01-21")] alias forall_lt_le_iff := forall_lt_imp_le_iff_le_of_dense +@[deprecated (since := "2025-01-21")] alias forall_gt_ge_iff := forall_gt_imp_ge_iff_le_of_dense + +@[deprecated (since := "2025-01-21")] +alias eq_of_le_of_forall_le_of_dense := eq_of_le_of_forall_lt_imp_le_of_dense -theorem forall_lt_le_iff [LinearOrder α] [DenselyOrdered α] {a b : α} : (∀ c < a, c ≤ b) ↔ a ≤ b := - ⟨le_of_forall_ge_of_dense, fun hab _c hca ↦ hca.le.trans hab⟩ +@[deprecated (since := "2025-01-21")] +alias eq_of_le_of_forall_ge_of_dense := eq_of_le_of_forall_gt_imp_ge_of_dense -theorem forall_gt_ge_iff [LinearOrder α] [DenselyOrdered α] {a b : α} : - (∀ c, a < c → b ≤ c) ↔ b ≤ a := - forall_lt_le_iff (α := αᵒᵈ) +end LinearOrder theorem dense_or_discrete [LinearOrder α] (a₁ a₂ : α) : (∃ a, a₁ < a ∧ a < a₂) ∨ (∀ a, a₁ < a → a₂ ≤ a) ∧ ∀ a < a₂, a ≤ a₁ := diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index ea8924624d601..e46f2bb3b9b91 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -697,6 +697,12 @@ theorem himp_le : x ⇨ y ≤ z ↔ y ≤ z ∧ Codisjoint x z := lemma himp_ne_right : x ⇨ y ≠ x ↔ x ≠ ⊤ ∨ y ≠ ⊤ := himp_eq_left.not.trans not_and_or +lemma codisjoint_iff_compl_le_left : Codisjoint x y ↔ yᶜ ≤ x := + hnot_le_iff_codisjoint_left.symm + +lemma codisjoint_iff_compl_le_right : Codisjoint x y ↔ xᶜ ≤ y := + hnot_le_iff_codisjoint_right.symm + end BooleanAlgebra instance Prop.instBooleanAlgebra : BooleanAlgebra Prop where diff --git a/Mathlib/Order/Booleanisation.lean b/Mathlib/Order/Booleanisation.lean index 6d5ee593160b5..736b9fbe34360 100644 --- a/Mathlib/Order/Booleanisation.lean +++ b/Mathlib/Order/Booleanisation.lean @@ -51,7 +51,7 @@ algebra. -/ /-- The complement operator on `Booleanisation α` sends `a` to `aᶜ` and `aᶜ` to `a`, for `a : α`. -/ instance instCompl : HasCompl (Booleanisation α) where - compl x := match x with + compl | lift a => comp a | comp a => lift a @@ -94,7 +94,7 @@ instance instLT : LT (Booleanisation α) where * `aᶜ ⊔ b` is `(a \ b)ᶜ` * `aᶜ ⊔ bᶜ` is `(a ⊓ b)ᶜ` -/ instance instSup : Max (Booleanisation α) where - max x y := match x, y with + max | lift a, lift b => lift (a ⊔ b) | lift a, comp b => comp (b \ a) | comp a, lift b => comp (a \ b) @@ -106,7 +106,7 @@ instance instSup : Max (Booleanisation α) where * `aᶜ ⊓ b` is `b \ a` * `aᶜ ⊓ bᶜ` is `(a ⊔ b)ᶜ` -/ instance instInf : Min (Booleanisation α) where - min x y := match x, y with + min | lift a, lift b => lift (a ⊓ b) | lift a, comp b => lift (a \ b) | comp a, lift b => lift (b \ a) @@ -126,7 +126,7 @@ instance instTop : Top (Booleanisation α) where * `aᶜ \ b` is `(a ⊔ b)ᶜ` * `aᶜ \ bᶜ` is `b \ a` -/ instance instSDiff : SDiff (Booleanisation α) where - sdiff x y := match x, y with + sdiff | lift a, lift b => lift (a \ b) | lift a, comp b => lift (a ⊓ b) | comp a, lift b => comp (a ⊔ b) @@ -162,22 +162,22 @@ instance instSDiff : SDiff (Booleanisation α) where instance instPreorder : Preorder (Booleanisation α) where lt := (· < ·) - lt_iff_le_not_le x y := match x, y with + lt_iff_le_not_le | lift a, lift b => by simp [lt_iff_le_not_le] | lift a, comp b => by simp | comp a, lift b => by simp | comp a, comp b => by simp [lt_iff_le_not_le] - le_refl x := match x with + le_refl | lift _ => LE.lift le_rfl | comp _ => LE.comp le_rfl - le_trans x y z hxy hyz := match x, y, z, hxy, hyz with + le_trans | lift _, lift _, lift _, LE.lift hab, LE.lift hbc => LE.lift <| hab.trans hbc | lift _, lift _, comp _, LE.lift hab, LE.sep hbc => LE.sep <| hbc.mono_left hab | lift _, comp _, comp _, LE.sep hab, LE.comp hcb => LE.sep <| hab.mono_right hcb | comp _, comp _, comp _, LE.comp hba, LE.comp hcb => LE.comp <| hcb.trans hba instance instPartialOrder : PartialOrder (Booleanisation α) where - le_antisymm x y hxy hyx := match x, y, hxy, hyx with + le_antisymm | lift a, lift b, LE.lift hab, LE.lift hba => by rw [hab.antisymm hba] | comp a, comp b, LE.comp hab, LE.comp hba => by rw [hab.antisymm hba] @@ -185,17 +185,17 @@ instance instPartialOrder : PartialOrder (Booleanisation α) where set_option linter.unusedVariables false in instance instSemilatticeSup : SemilatticeSup (Booleanisation α) where sup x y := max x y - le_sup_left x y := match x, y with + le_sup_left | lift a, lift b => LE.lift le_sup_left | lift a, comp b => LE.sep disjoint_sdiff_self_right | comp a, lift b => LE.comp sdiff_le | comp a, comp b => LE.comp inf_le_left - le_sup_right x y := match x, y with + le_sup_right | lift a, lift b => LE.lift le_sup_right | lift a, comp b => LE.comp sdiff_le | comp a, lift b => LE.sep disjoint_sdiff_self_right | comp a, comp b => LE.comp inf_le_right - sup_le x y z hxz hyz := match x, y, z, hxz, hyz with + sup_le | lift a, lift b, lift c, LE.lift hac, LE.lift hbc => LE.lift <| sup_le hac hbc | lift a, lift b, comp c, LE.sep hac, LE.sep hbc => LE.sep <| hac.sup_left hbc | lift a, comp b, comp c, LE.sep hac, LE.comp hcb => LE.comp <| le_sdiff.2 ⟨hcb, hac.symm⟩ @@ -206,17 +206,17 @@ instance instSemilatticeSup : SemilatticeSup (Booleanisation α) where set_option linter.unusedVariables false in instance instSemilatticeInf : SemilatticeInf (Booleanisation α) where inf x y := min x y - inf_le_left x y := match x, y with + inf_le_left | lift a, lift b => LE.lift inf_le_left | lift a, comp b => LE.lift sdiff_le | comp a, lift b => LE.sep disjoint_sdiff_self_left | comp a, comp b => LE.comp le_sup_left - inf_le_right x y := match x, y with + inf_le_right | lift a, lift b => LE.lift inf_le_right | lift a, comp b => LE.sep disjoint_sdiff_self_left | comp a, lift b => LE.lift sdiff_le | comp a, comp b => LE.comp le_sup_right - le_inf x y z hxz hyz := match x, y, z, hxz, hyz with + le_inf | lift a, lift b, lift c, LE.lift hab, LE.lift hac => LE.lift <| le_inf hab hac | lift a, lift b, comp c, LE.lift hab, LE.sep hac => LE.lift <| le_sdiff.2 ⟨hab, hac⟩ | lift a, comp b, lift c, LE.sep hab, LE.lift hac => LE.lift <| le_sdiff.2 ⟨hac, hab⟩ @@ -227,7 +227,7 @@ instance instDistribLattice : DistribLattice (Booleanisation α) where inf_le_left _ _ := inf_le_left inf_le_right _ _ := inf_le_right le_inf _ _ _ := le_inf - le_sup_inf x y z := match x, y, z with + le_sup_inf | lift _, lift _, lift _ => LE.lift le_sup_inf | lift a, lift b, comp c => LE.lift <| by simp [sup_left_comm, sup_comm, sup_assoc] | lift a, comp b, lift c => LE.lift <| by @@ -241,23 +241,23 @@ instance instDistribLattice : DistribLattice (Booleanisation α) where -- The linter significantly hinders readability here. set_option linter.unusedVariables false in instance instBoundedOrder : BoundedOrder (Booleanisation α) where - le_top x := match x with + le_top | lift a => LE.sep disjoint_bot_right | comp a => LE.comp bot_le - bot_le x := match x with + bot_le | lift a => LE.lift bot_le | comp a => LE.sep disjoint_bot_left instance instBooleanAlgebra : BooleanAlgebra (Booleanisation α) where le_top _ := le_top bot_le _ := bot_le - inf_compl_le_bot x := match x with + inf_compl_le_bot | lift a => by simp | comp a => by simp - top_le_sup_compl x := match x with + top_le_sup_compl | lift a => by simp | comp a => by simp - sdiff_eq x y := match x, y with + sdiff_eq | lift a, lift b => by simp | lift a, comp b => by simp | comp a, lift b => by simp diff --git a/Mathlib/Order/BoundedOrder/Basic.lean b/Mathlib/Order/BoundedOrder/Basic.lean index e767d449e62e6..2ad6d7f4daaf3 100644 --- a/Mathlib/Order/BoundedOrder/Basic.lean +++ b/Mathlib/Order/BoundedOrder/Basic.lean @@ -517,10 +517,10 @@ instance instTop [Top α] [Top β] : Top (α × β) := instance instBot [Bot α] [Bot β] : Bot (α × β) := ⟨⟨⊥, ⊥⟩⟩ -theorem fst_top [Top α] [Top β] : (⊤ : α × β).fst = ⊤ := rfl -theorem snd_top [Top α] [Top β] : (⊤ : α × β).snd = ⊤ := rfl -theorem fst_bot [Bot α] [Bot β] : (⊥ : α × β).fst = ⊥ := rfl -theorem snd_bot [Bot α] [Bot β] : (⊥ : α × β).snd = ⊥ := rfl +@[simp] lemma fst_top [Top α] [Top β] : (⊤ : α × β).fst = ⊤ := rfl +@[simp] lemma snd_top [Top α] [Top β] : (⊤ : α × β).snd = ⊤ := rfl +@[simp] lemma fst_bot [Bot α] [Bot β] : (⊥ : α × β).fst = ⊥ := rfl +@[simp] lemma snd_bot [Bot α] [Bot β] : (⊥ : α × β).snd = ⊥ := rfl instance instOrderTop [LE α] [LE β] [OrderTop α] [OrderTop β] : OrderTop (α × β) where __ := inferInstanceAs (Top (α × β)) diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index ad2107194dde3..a161ef7fab20c 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -119,6 +119,12 @@ abbrev IsGreatest.orderTop (h : IsGreatest s a) : top := ⟨a, h.1⟩ le_top := Subtype.forall.2 h.2 +theorem isLUB_congr (h : upperBounds s = upperBounds t) : IsLUB s a ↔ IsLUB t a := by + rw [IsLUB, IsLUB, h] + +theorem isGLB_congr (h : lowerBounds s = lowerBounds t) : IsGLB s a ↔ IsGLB t a := by + rw [IsGLB, IsGLB, h] + /-! ### Monotonicity -/ @@ -452,7 +458,7 @@ theorem exists_glb_Ioi (i : γ) : ∃ j, IsGLB (Ioi i) j := variable [DenselyOrdered γ] theorem isLUB_Iio {a : γ} : IsLUB (Iio a) a := - ⟨fun _ hx => le_of_lt hx, fun _ hy => le_of_forall_ge_of_dense hy⟩ + ⟨fun _ hx => le_of_lt hx, fun _ hy => le_of_forall_lt_imp_le_of_dense hy⟩ theorem isGLB_Ioi {a : γ} : IsGLB (Ioi a) a := @isLUB_Iio γᵒᵈ _ _ a diff --git a/Mathlib/Order/Bounds/Lattice.lean b/Mathlib/Order/Bounds/Lattice.lean new file mode 100644 index 0000000000000..4debbe2ec02e8 --- /dev/null +++ b/Mathlib/Order/Bounds/Lattice.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2024 Christopher Hoskin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christopher Hoskin +-/ +import Mathlib.Data.Set.Lattice + +/-! +# Unions and intersections of bounds + +Some results about upper and lower bounds over collections of sets. + +## Implementation notes + +In a separate file as we need to import `Mathlib.Data.Set.Lattice`. + +-/ + +variable {α : Type*} [Preorder α] {ι : Sort*} {s : ι → Set α} + +open Set + +theorem gc_upperBounds_lowerBounds : GaloisConnection + (OrderDual.toDual ∘ upperBounds : Set α → (Set α)ᵒᵈ) + (lowerBounds ∘ OrderDual.ofDual : (Set α)ᵒᵈ → Set α) := by + simpa [GaloisConnection, subset_def, mem_upperBounds, mem_lowerBounds] + using fun S T ↦ forall₂_swap + +@[simp] +theorem upperBounds_iUnion : + upperBounds (⋃ i, s i) = ⋂ i, upperBounds (s i) := + gc_upperBounds_lowerBounds.l_iSup + +@[simp] +theorem lowerBounds_iUnion : + lowerBounds (⋃ i, s i) = ⋂ i, lowerBounds (s i) := + gc_upperBounds_lowerBounds.u_iInf + +theorem isLUB_iUnion_iff_of_isLUB {u : ι → α} (hs : ∀ i, IsLUB (s i) (u i)) (c : α) : + IsLUB (Set.range u) c ↔ IsLUB (⋃ i, s i) c := by + refine isLUB_congr ?_ + simp_rw [range_eq_iUnion, upperBounds_iUnion, upperBounds_singleton, (hs _).upperBounds_eq] + +theorem isGLB_iUnion_iff_of_isLUB {u : ι → α} (hs : ∀ i, IsGLB (s i) (u i)) (c : α) : + IsGLB (Set.range u) c ↔ IsGLB (⋃ i, s i) c := by + refine isGLB_congr ?_ + simp_rw [range_eq_iUnion, lowerBounds_iUnion, lowerBounds_singleton, (hs _).lowerBounds_eq] diff --git a/Mathlib/Order/Category/BddLat.lean b/Mathlib/Order/Category/BddLat.lean index 994a24252c635..531f0ef0124de 100644 --- a/Mathlib/Order/Category/BddLat.lean +++ b/Mathlib/Order/Category/BddLat.lean @@ -199,15 +199,6 @@ def latToBddLatForgetAdjunction : latToBddLat.{u} ⊣ forget₂ BddLat Lat := homEquiv_naturality_right := fun _ _ => LatticeHom.ext fun _ => rfl } /-- `latToBddLat` and `OrderDual` commute. -/ --- Porting note: the `simpNF` linter is not happy as it simplifies something that does not --- have prettyprinting effects. --- It seems like it is simplifying for example the first type --- `(↑(BddLat.dualEquiv.functor.obj (latToBddLat.obj X.op.unop)).toLat)` --- to --- `(↑(latToBddLat.obj X).toLat)ᵒᵈ` --- Interestingly, the linter is silent, if the proof is `sorry`-ed out... --- see https://github.com/leanprover-community/mathlib4/issues/5049 --- @[simps!] def latToBddLatCompDualIsoDualCompLatToBddLat : latToBddLat.{u} ⋙ BddLat.dual ≅ Lat.dual ⋙ latToBddLat := Adjunction.leftAdjointUniq (latToBddLatForgetAdjunction.comp BddLat.dualEquiv.toAdjunction) diff --git a/Mathlib/Order/Chain.lean b/Mathlib/Order/Chain.lean index 4b93268a6ec0b..60ab154fab092 100644 --- a/Mathlib/Order/Chain.lean +++ b/Mathlib/Order/Chain.lean @@ -137,6 +137,24 @@ theorem IsChain.exists3 (hchain : IsChain r s) [IsTrans α r] {a b c} (mem1 : a end Total +lemma IsChain.le_of_not_lt [Preorder α] (hs : IsChain (· ≤ ·) s) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) (h : ¬ x < y) : y ≤ x := by + cases hs.total hx hy with + | inr h' => exact h' + | inl h' => simpa [lt_iff_le_not_le, h'] using h + +lemma IsChain.not_lt [Preorder α] (hs : IsChain (· ≤ ·) s) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) : ¬ x < y ↔ y ≤ x := + ⟨(hs.le_of_not_lt hx hy ·), fun h h' ↦ h'.not_le h⟩ + +lemma IsChain.lt_of_not_le [Preorder α] (hs : IsChain (· ≤ ·) s) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) (h : ¬ x ≤ y) : y < x := + (hs.total hx hy).elim (h · |>.elim) (lt_of_le_not_le · h) + +lemma IsChain.not_le [Preorder α] (hs : IsChain (· ≤ ·) s) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) : ¬ x ≤ y ↔ y < x := + ⟨(hs.lt_of_not_le hx hy ·), fun h h' ↦ h'.not_lt h⟩ + theorem IsMaxChain.isChain (h : IsMaxChain r s) : IsChain r s := h.1 diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 8fb5b69f0275c..6353db8a7b166 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -5,7 +5,6 @@ Authors: Bhavik Mehta, Yaël Dillies -/ import Mathlib.Data.Set.Lattice import Mathlib.Data.SetLike.Basic -import Mathlib.Order.GaloisConnection import Mathlib.Order.Hom.Basic /-! diff --git a/Mathlib/Order/Cofinal.lean b/Mathlib/Order/Cofinal.lean index 532481a8f4cbe..fb434c2720926 100644 --- a/Mathlib/Order/Cofinal.lean +++ b/Mathlib/Order/Cofinal.lean @@ -3,7 +3,9 @@ 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.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Defs +import Mathlib.Order.Interval.Set.Basic +import Mathlib.Order.WellFounded /-! # Cofinal sets diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index e68234ef30fbb..6dfb80d1a72f2 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -113,6 +113,7 @@ theorem sSup_le_sSup_of_forall_exists_le (h : ∀ x ∈ s, ∃ y ∈ t, x ≤ y) hac.trans (hb hct) -- We will generalize this to conditionally complete lattices in `csSup_singleton`. +@[simp] theorem sSup_singleton {a : α} : sSup {a} = a := isLUB_singleton.sSup_eq @@ -168,6 +169,7 @@ theorem sInf_le_sInf_of_forall_exists_le (h : ∀ x ∈ s, ∃ y ∈ t, y ≤ x) le_sInf fun x hx ↦ let ⟨_y, hyt, hyx⟩ := h x hx; sInf_le_of_le hyt hyx -- We will generalize this to conditionally complete lattices in `csInf_singleton`. +@[simp] theorem sInf_singleton {a : α} : sInf {a} = a := isGLB_singleton.sInf_eq diff --git a/Mathlib/Order/CompleteLattice/SetLike.lean b/Mathlib/Order/CompleteLattice/SetLike.lean new file mode 100644 index 0000000000000..2e5426d136ae5 --- /dev/null +++ b/Mathlib/Order/CompleteLattice/SetLike.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2024 Sven Manthe. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sven Manthe +-/ +import Mathlib.Order.CompleteSublattice + +/-! +# `SetLike` instance for elements of `CompleteSublattice (Set X)` + +This file provides lemmas for the `SetLike` instance for elements of `CompleteSublattice (Set X)` +-/ + +attribute [local instance] SetLike.instSubtypeSet + +namespace Sublattice + +variable {X : Type*} {L : Sublattice (Set X)} + +variable {S T : L} {x : X} + +@[ext] lemma ext_mem (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := SetLike.ext h + +lemma mem_subtype : x ∈ L.subtype T ↔ x ∈ T := Iff.rfl + +@[simp] lemma setLike_mem_inf : x ∈ S ⊓ T ↔ x ∈ S ∧ x ∈ T := by simp [← mem_subtype] +@[simp] lemma setLike_mem_sup : x ∈ S ⊔ T ↔ x ∈ S ∨ x ∈ T := by simp [← mem_subtype] + +@[simp] lemma setLike_mem_coe : x ∈ T.val ↔ x ∈ T := Iff.rfl +@[simp] lemma setLike_mem_mk (U : Set X) (h : U ∈ L) : x ∈ (⟨U, h⟩ : L) ↔ x ∈ U := Iff.rfl + +end Sublattice + +namespace CompleteSublattice + +variable {X : Type*} {L : CompleteSublattice (Set X)} + +variable {S T : L} {𝒮 : Set L} {I : Sort*} {f : I → L} {x : X} + +@[ext] lemma ext (h : ∀ x, x ∈ S ↔ x ∈ T) : S = T := SetLike.ext h + +lemma mem_subtype : x ∈ L.subtype T ↔ x ∈ T := Iff.rfl + +@[simp] lemma mem_inf : x ∈ S ⊓ T ↔ x ∈ S ∧ x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_sInf : x ∈ sInf 𝒮 ↔ ∀ T ∈ 𝒮, x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_iInf : x ∈ ⨅ i : I, f i ↔ ∀ i : I, x ∈ f i := by simp [← mem_subtype] +@[simp] lemma mem_top : x ∈ (⊤ : L) := by simp [← mem_subtype] + +@[simp] lemma mem_sup : x ∈ S ⊔ T ↔ x ∈ S ∨ x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_sSup : x ∈ sSup 𝒮 ↔ ∃ T ∈ 𝒮, x ∈ T := by simp [← mem_subtype] +@[simp] lemma mem_iSup : x ∈ ⨆ i : I, f i ↔ ∃ i : I, x ∈ f i := by simp [← mem_subtype] +@[simp] lemma not_mem_bot : ¬ x ∈ (⊥ : L) := by simp [← mem_subtype] + +@[simp] lemma mem_mk (U : Set X) (h : U ∈ L) : x ∈ (⟨U, h⟩ : L) ↔ x ∈ U := Iff.rfl + +end CompleteSublattice diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 0773cfe1e11e4..84fa0aae1c71a 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -317,11 +317,9 @@ theorem csSup_singleton (a : α) : sSup {a} = a := theorem csInf_singleton (a : α) : sInf {a} = a := isLeast_singleton.csInf_eq -@[simp] theorem csSup_pair (a b : α) : sSup {a, b} = a ⊔ b := (@isLUB_pair _ _ a b).csSup_eq (insert_nonempty _ _) -@[simp] theorem csInf_pair (a b : α) : sInf {a, b} = a ⊓ b := (@isGLB_pair _ _ a b).csInf_eq (insert_nonempty _ _) @@ -356,11 +354,13 @@ theorem le_csInf_inter : /-- The supremum of `insert a s` is the maximum of `a` and the supremum of `s`, if `s` is nonempty and bounded above. -/ +@[simp] theorem csSup_insert (hs : BddAbove s) (sne : s.Nonempty) : sSup (insert a s) = a ⊔ sSup s := ((isLUB_csSup sne hs).insert a).csSup_eq (insert_nonempty a s) /-- The infimum of `insert a s` is the minimum of `a` and the infimum of `s`, if `s` is nonempty and bounded below. -/ +@[simp] theorem csInf_insert (hs : BddBelow s) (sne : s.Nonempty) : sInf (insert a s) = a ⊓ sInf s := csSup_insert (α := αᵒᵈ) hs sne @@ -538,12 +538,12 @@ open Function variable [WellFoundedLT α] -theorem sInf_eq_argmin_on (hs : s.Nonempty) : sInf s = argminOn id wellFounded_lt s hs := - IsLeast.csInf_eq ⟨argminOn_mem _ _ _ _, fun _ ha => argminOn_le id _ _ ha⟩ +theorem sInf_eq_argmin_on (hs : s.Nonempty) : sInf s = argminOn id s hs := + IsLeast.csInf_eq ⟨argminOn_mem _ _ _, fun _ ha => argminOn_le id _ ha⟩ theorem isLeast_csInf (hs : s.Nonempty) : IsLeast s (sInf s) := by rw [sInf_eq_argmin_on hs] - exact ⟨argminOn_mem _ _ _ _, fun a ha => argminOn_le id _ _ ha⟩ + exact ⟨argminOn_mem _ _ _, fun a ha => argminOn_le id _ ha⟩ theorem le_csInf_iff' (hs : s.Nonempty) : b ≤ sInf s ↔ b ∈ lowerBounds s := le_isGLB_iff (isLeast_csInf hs).isGLB diff --git a/Mathlib/Order/Disjointed.lean b/Mathlib/Order/Disjointed.lean index 62ea4568a5fa6..22ee30f0339ee 100644 --- a/Mathlib/Order/Disjointed.lean +++ b/Mathlib/Order/Disjointed.lean @@ -1,42 +1,43 @@ /- 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, Yaël Dillies +Authors: Johannes Hölzl, Yaël Dillies, David Loeffler -/ -import Mathlib.Algebra.Order.PartialSups -import Mathlib.Data.Nat.SuccPred +import Mathlib.Order.PartialSups +import Mathlib.Order.Interval.Finset.Fin /-! -# Consecutive differences of sets +# Making a sequence disjoint -This file defines the way to make a sequence of elements into a sequence of disjoint elements with +This file defines the way to make a sequence of sets - or, more generally, a map from a partially +ordered type `ι` into a (generalized) Boolean algebra `α` - into a *pairwise disjoint* sequence with the same partial sups. -For a sequence `f : ℕ → α`, this new sequence will be `f 0`, `f 1 \ f 0`, `f 2 \ (f 0 ⊔ f 1)`. +For a sequence `f : ℕ → α`, this new sequence will be `f 0`, `f 1 \ f 0`, `f 2 \ (f 0 ⊔ f 1) ⋯`. It is actually unique, as `disjointed_unique` shows. ## Main declarations -* `disjointed f`: The sequence `f 0`, `f 1 \ f 0`, `f 2 \ (f 0 ⊔ f 1)`, .... +* `disjointed f`: The map sending `i` to `f i \ (⨆ j < i, f j)`. We require the index type to be a + `LocallyFiniteOrderBot` to ensure that the supremum is well defined. * `partialSups_disjointed`: `disjointed f` has the same partial sups as `f`. * `disjoint_disjointed`: The elements of `disjointed f` are pairwise disjoint. * `disjointed_unique`: `disjointed f` is the only pairwise disjoint sequence having the same partial sups as `f`. -* `iSup_disjointed`: `disjointed f` has the same supremum as `f`. Limiting case of - `partialSups_disjointed`. +* `Fintype.sup_disjointed` (for finite `ι`) or `iSup_disjointed` (for complete `α`): + `disjointed f` has the same supremum as `f`. Limiting case of `partialSups_disjointed`. +* `Fintype.exists_disjointed_le`: for any finite family `f : ι → α`, there exists a pairwise + disjoint family `g : ι → α` which is bounded above by `f` and has the same supremum. This is + an analogue of `disjointed` for arbitrary finite index types (but without any uniqueness). We also provide set notation variants of some lemmas. - -## TODO - -Find a useful statement of `disjointedRec_succ`. - -One could generalize `disjointed` to any locally finite bot preorder domain, in place of `ℕ`. -Related to the TODO in the module docstring of `Mathlib.Order.PartialSups`. -/ +assert_not_exists SuccAddOrder -variable {α : Type*} +open Finset Order + +variable {α ι : Type*} open scoped Function -- required for scoped `on` notation @@ -44,124 +45,261 @@ section GeneralizedBooleanAlgebra variable [GeneralizedBooleanAlgebra α] -/-- If `f : ℕ → α` is a sequence of elements, then `disjointed f` is the sequence formed by -subtracting each element from the nexts. This is the unique disjoint sequence whose partial sups -are the same as the original sequence. -/ -def disjointed (f : ℕ → α) : ℕ → α - | 0 => f 0 - | n + 1 => f (n + 1) \ partialSups f n +section Preorder -- the *index type* is a preorder -@[simp] -theorem disjointed_zero (f : ℕ → α) : disjointed f 0 = f 0 := - rfl +variable [Preorder ι] [LocallyFiniteOrderBot ι] -theorem disjointed_succ (f : ℕ → α) (n : ℕ) : disjointed f (n + 1) = f (n + 1) \ partialSups f n := - rfl +/-- The function mapping `i` to `f i \ (⨆ j < i, f j)`. When `ι` is a partial order, this is the +unique function `g` having the same `partialSups` as `f` and such that `g i` and `g j` are +disjoint whenever `i < j`. -/ +def disjointed (f : ι → α) (i : ι) : α := f i \ (Iio i).sup f -theorem disjointed_le_id : disjointed ≤ (id : (ℕ → α) → ℕ → α) := by - rintro f n - cases n - · rfl - · exact sdiff_le +lemma disjointed_apply (f : ι → α) (i : ι) : disjointed f i = f i \ (Iio i).sup f := rfl -theorem disjointed_le (f : ℕ → α) : disjointed f ≤ f := +lemma disjointed_of_isMin (f : ι → α) {i : ι} (hn : IsMin i) : + disjointed f i = f i := by + have : Iio i = ∅ := by rwa [← Finset.coe_eq_empty, coe_Iio, Set.Iio_eq_empty_iff] + simp only [disjointed_apply, this, sup_empty, sdiff_bot] + +@[simp] lemma disjointed_bot [OrderBot ι] (f : ι → α) : disjointed f ⊥ = f ⊥ := + disjointed_of_isMin _ isMin_bot + +theorem disjointed_le_id : disjointed ≤ (id : (ι → α) → ι → α) := + fun _ _ ↦ sdiff_le + +theorem disjointed_le (f : ι → α) : disjointed f ≤ f := disjointed_le_id f -theorem disjoint_disjointed (f : ℕ → α) : Pairwise (Disjoint on disjointed f) := by - refine (Symmetric.pairwise_on Disjoint.symm _).2 fun m n h => ?_ - cases n - · exact (Nat.not_lt_zero _ h).elim - exact - disjoint_sdiff_self_right.mono_left - ((disjointed_le f m).trans (le_partialSups_of_le f (Nat.lt_add_one_iff.1 h))) - --- Porting note: `disjointedRec` had a change in universe level. -/-- An induction principle for `disjointed`. To define/prove something on `disjointed f n`, it's -enough to define/prove it for `f n` and being able to extend through diffs. -/ -def disjointedRec {f : ℕ → α} {p : α → Sort*} (hdiff : ∀ ⦃t i⦄, p t → p (t \ f i)) : - ∀ ⦃n⦄, p (f n) → p (disjointed f n) - | 0 => id - | n + 1 => fun h => by - suffices H : ∀ k, p (f (n + 1) \ partialSups f k) from H n - rintro k - induction' k with k ih - · exact hdiff h - rw [partialSups_add_one, ← sdiff_sdiff_left] - exact hdiff ih +theorem disjoint_disjointed_of_lt (f : ι → α) {i j : ι} (h : i < j) : + Disjoint (disjointed f i) (disjointed f j) := + (disjoint_sdiff_self_right.mono_left <| le_sup (mem_Iio.mpr h)).mono_left (disjointed_le f i) + +lemma disjointed_eq_self {f : ι → α} {i : ι} (hf : ∀ j < i, Disjoint (f j) (f i)) : + disjointed f i = f i := by + rw [disjointed_apply, sdiff_eq_left, disjoint_iff, sup_inf_distrib_left, + sup_congr rfl <| fun j hj ↦ disjoint_iff.mp <| (hf _ (mem_Iio.mp hj)).symm] + exact sup_bot _ + +/- NB: The original statement for `ι = ℕ` was a `def` and worked for `p : α → Sort*`. I couldn't +prove the `Sort*` version for general `ι`, but all instances of `disjointedRec` in the library are +for Prop anyway. -/ +/-- +An induction principle for `disjointed`. To prove something about `disjointed f i`, it's +enough to prove it for `f i` and being able to extend through diffs. +-/ +lemma disjointedRec {f : ι → α} {p : α → Prop} (hdiff : ∀ ⦃t i⦄, p t → p (t \ f i)) : + ∀ ⦃i⦄, p (f i) → p (disjointed f i) := by + classical + intro i hpi + rw [disjointed] + suffices ∀ (s : Finset ι), p (f i \ s.sup f) from this _ + intro s + induction s using Finset.induction with + | empty => simpa only [sup_empty, sdiff_bot] using hpi + | insert ht IH => + rw [sup_insert, sup_comm, ← sdiff_sdiff] + exact hdiff IH -@[simp] -theorem disjointedRec_zero {f : ℕ → α} {p : α → Sort*} (hdiff : ∀ ⦃t i⦄, p t → p (t \ f i)) - (h₀ : p (f 0)) : disjointedRec hdiff h₀ = h₀ := - rfl +end Preorder --- TODO: Find a useful statement of `disjointedRec_succ`. -protected lemma Monotone.disjointed_succ {f : ℕ → α} (hf : Monotone f) (n : ℕ) : - disjointed f (n + 1) = f (n + 1) \ f n := by rw [disjointed_succ, hf.partialSups_eq] +section PartialOrder -- the index type is a partial order -protected lemma Monotone.disjointed_succ_sup {f : ℕ → α} (hf : Monotone f) (n : ℕ) : - disjointed f (n + 1) ⊔ f n = f (n + 1) := by - rw [hf.disjointed_succ, sdiff_sup_cancel]; exact hf n.le_succ +variable [PartialOrder ι] [LocallyFiniteOrderBot ι] @[simp] -theorem partialSups_disjointed (f : ℕ → α) : partialSups (disjointed f) = partialSups f := by - ext n - induction' n with k ih - · rw [partialSups_zero, partialSups_zero, disjointed_zero] - · rw [partialSups_add_one, partialSups_add_one, disjointed_succ, ih, sup_sdiff_self_right] +theorem partialSups_disjointed (f : ι → α) : + partialSups (disjointed f) = partialSups f := by + -- This seems to be much more awkward than the case of linear orders, because the supremum + -- in the definition of `disjointed` can involve multiple "paths" through the poset. + classical + -- We argue by induction on the size of `Iio i`. + suffices ∀ r i (hi : #(Iio i) ≤ r), partialSups (disjointed f) i = partialSups f i from + OrderHom.ext _ _ (funext fun i ↦ this _ i le_rfl) + intro r i hi + induction r generalizing i with + | zero => + -- Base case: `n` is minimal, so `partialSups f i = partialSups (disjointed f) n = f i`. + simp only [Nat.le_zero, card_eq_zero] at hi + simp only [partialSups_apply, Iic_eq_cons_Iio, hi, disjointed_apply, sup'_eq_sup, sup_cons, + sup_empty, sdiff_bot] + | succ n ih => + -- Induction step: first WLOG arrange that `#(Iio i) = r + 1` + rcases lt_or_eq_of_le hi with hn | hn + · exact ih _ <| Nat.le_of_lt_succ hn + simp only [partialSups_apply (disjointed f), Iic_eq_cons_Iio, sup'_eq_sup, sup_cons] + -- Key claim: we can write `Iio i` as a union of (finitely many) `Ici` intervals. + have hun : (Iio i).biUnion Iic = Iio i := by + ext r; simpa using ⟨fun ⟨a, ha⟩ ↦ ha.2.trans_lt ha.1, fun hr ↦ ⟨r, hr, le_rfl⟩⟩ + -- Use claim and `sup_biUnion` to rewrite the supremum in the definition of `disjointed f` + -- in terms of suprema over `Iic`'s. Then the RHS is a `sup` over `partialSups`, which we + -- can rewrite via the induction hypothesis. + rw [← hun, sup_biUnion, sup_congr rfl (g := partialSups f)] + · simp only [funext (partialSups_apply f), sup'_eq_sup, ← sup_biUnion, hun] + simp only [disjointed, sdiff_sup_self, Iic_eq_cons_Iio, sup_cons] + · simp only [partialSups, sup'_eq_sup, OrderHom.coe_mk] at ih ⊢ + refine fun x hx ↦ ih x ?_ + -- Remains to show `∀ x in Iio i, #(Iio x) ≤ r`. + rw [← Nat.lt_add_one_iff, ← hn] + apply lt_of_lt_of_le (b := #(Iic x)) + · simpa only [Iic_eq_cons_Iio, card_cons] using Nat.lt_succ_self _ + · refine card_le_card (fun r hr ↦ ?_) + simp only [mem_Iic, mem_Iio] at hx hr ⊢ + exact hr.trans_lt hx + +lemma Fintype.sup_disjointed [Fintype ι] (f : ι → α) : + univ.sup (disjointed f) = univ.sup f := by + classical + have hun : univ.biUnion Iic = (univ : Finset ι) := by + ext r; simpa only [mem_biUnion, mem_univ, mem_Iic, true_and, iff_true] using ⟨r, le_rfl⟩ + rw [← hun, sup_biUnion, sup_biUnion, sup_congr rfl (fun i _ ↦ ?_)] + rw [← sup'_eq_sup nonempty_Iic, ← sup'_eq_sup nonempty_Iic, + ← partialSups_apply, ← partialSups_apply, partialSups_disjointed] + +lemma disjointed_partialSups (f : ι → α) : + disjointed (partialSups f) = disjointed f := by + classical + ext i + have step1 : f i \ (Iio i).sup f = partialSups f i \ (Iio i).sup f := by + rw [sdiff_eq_symm (sdiff_le.trans (le_partialSups f i))] + simp only [funext (partialSups_apply f), sup'_eq_sup] + rw [sdiff_sdiff_eq_sdiff_sup (sup_mono Iio_subset_Iic_self), sup_eq_right] + simp only [Iic_eq_cons_Iio, sup_cons, sup_sdiff_left_self, sdiff_le_iff, le_sup_right] + simp only [disjointed_apply, step1, funext (partialSups_apply f), sup'_eq_sup, ← sup_biUnion] + congr 2 with r + simpa only [mem_biUnion, mem_Iio, mem_Iic] using + ⟨fun ⟨a, ha⟩ ↦ ha.2.trans_lt ha.1, fun hr ↦ ⟨r, hr, le_rfl⟩⟩ + +/-- `disjointed f` is the unique map `d : ι → α` such that `d` has the same partial sups as `f`, +and `d i` and `d j` are disjoint whenever `i < j`. -/ +theorem disjointed_unique {f d : ι → α} (hdisj : ∀ {i j : ι} (_ : i < j), Disjoint (d i) (d j)) + (hsups : partialSups d = partialSups f) : + d = disjointed f := by + rw [← disjointed_partialSups, ← hsups, disjointed_partialSups] + exact funext fun _ ↦ (disjointed_eq_self (fun _ hj ↦ hdisj hj)).symm + +end PartialOrder + +section LinearOrder -- the index type is a linear order + +/-! +### Linear orders +-/ + +variable [LinearOrder ι] [LocallyFiniteOrderBot ι] + +theorem disjoint_disjointed (f : ι → α) : Pairwise (Disjoint on disjointed f) := + (pairwise_disjoint_on _).mpr fun _ _ ↦ disjoint_disjointed_of_lt f /-- `disjointed f` is the unique sequence that is pairwise disjoint and has the same partial sups as `f`. -/ -theorem disjointed_unique {f d : ℕ → α} (hdisj : Pairwise (Disjoint on d)) - (hsups : partialSups d = partialSups f) : d = disjointed f := by - ext n - cases' n with n - · rw [← partialSups_zero d, hsups, partialSups_zero, disjointed_zero] - suffices h : d (n + 1) = partialSups d (n + 1) \ partialSups d n by - rw [h, hsups, partialSups_add_one, disjointed_succ, sup_sdiff, sdiff_self, bot_sup_eq] - rw [partialSups_add_one, sup_sdiff, sdiff_self, bot_sup_eq, eq_comm, sdiff_eq_self_iff_disjoint] - suffices h : ∀ m ≤ n, Disjoint (partialSups d m) (d n.succ) from h n le_rfl - rintro m hm - induction' m with m ih - · exact hdisj (Nat.succ_ne_zero _).symm - rw [partialSups_add_one, disjoint_iff, inf_sup_right, sup_eq_bot_iff, - ← disjoint_iff, ← disjoint_iff] - exact ⟨ih (Nat.le_of_succ_le hm), hdisj (Nat.lt_succ_of_le hm).ne⟩ +theorem disjointed_unique' {f d : ι → α} (hdisj : Pairwise (Disjoint on d)) + (hsups : partialSups d = partialSups f) : d = disjointed f := + disjointed_unique (fun hij ↦ hdisj hij.ne) hsups + +section SuccOrder + +variable [SuccOrder ι] + +lemma disjointed_succ (f : ι → α) {i : ι} (hi : ¬IsMax i) : + disjointed f (succ i) = f (succ i) \ partialSups f i := by + rw [disjointed_apply, partialSups_apply, sup'_eq_sup] + congr 2 with m + simpa only [mem_Iio, mem_Iic] using lt_succ_iff_of_not_isMax hi + +protected lemma Monotone.disjointed_succ {f : ι → α} (hf : Monotone f) {i : ι} (hn : ¬IsMax i) : + disjointed f (succ i) = f (succ i) \ f i := by + rwa [disjointed_succ, hf.partialSups_eq] + +/-- Note this lemma does not require `¬IsMax i`, unlike `disjointed_succ`. -/ +lemma Monotone.disjointed_succ_sup {f : ι → α} (hf : Monotone f) (i : ι) : + disjointed f (succ i) ⊔ f i = f (succ i) := by + by_cases h : IsMax i + · simpa only [succ_eq_iff_isMax.mpr h, sup_eq_right] using disjointed_le f i + · rw [disjointed_apply] + have : Iio (succ i) = Iic i := by + ext + simp only [mem_Iio, lt_succ_iff_eq_or_lt_of_not_isMax h, mem_Iic, le_iff_lt_or_eq, Or.comm] + rw [this, ← sup'_eq_sup, ← partialSups_apply, hf.partialSups_eq, + sdiff_sup_cancel <| hf <| le_succ i] + +end SuccOrder + +end LinearOrder + +/-! +### Functions on an arbitrary fintype +-/ + +/-- For any finite family of elements `f : ι → α`, we can find a pairwise-disjoint family `g` +bounded above by `f` and having the same supremum. This is non-canonical, depending on an arbitrary +choice of ordering of `ι`. -/ +lemma Fintype.exists_disjointed_le {ι : Type*} [Fintype ι] {f : ι → α} : + ∃ g, g ≤ f ∧ univ.sup g = univ.sup f ∧ Pairwise (Disjoint on g) := by + rcases isEmpty_or_nonempty ι with hι | hι + · -- do `ι = ∅` separately since `⊤ : Fin n` isn't defined for `n = 0` + exact ⟨f, le_rfl, rfl, Subsingleton.pairwise⟩ + let R : ι ≃ Fin _ := equivFin ι + let f' : Fin _ → α := f ∘ R.symm + have hf' : f = f' ∘ R := by ext; simp only [Function.comp_apply, Equiv.symm_apply_apply, f'] + refine ⟨disjointed f' ∘ R, ?_, ?_, ?_⟩ + · intro n + simpa only [hf'] using disjointed_le f' (R n) + · simpa only [← sup_image, image_univ_equiv, hf'] using sup_disjointed f' + · exact fun i j hij ↦ disjoint_disjointed f' (R.injective.ne hij) end GeneralizedBooleanAlgebra section CompleteBooleanAlgebra +/-! ### Complete Boolean algebras -/ + variable [CompleteBooleanAlgebra α] -theorem iSup_disjointed (f : ℕ → α) : ⨆ n, disjointed f n = ⨆ n, f n := +theorem iSup_disjointed [PartialOrder ι] [LocallyFiniteOrderBot ι] (f : ι → α) : + ⨆ i, disjointed f i = ⨆ i, f i := iSup_eq_iSup_of_partialSups_eq_partialSups (partialSups_disjointed f) -theorem disjointed_eq_inf_compl (f : ℕ → α) (n : ℕ) : disjointed f n = f n ⊓ ⨅ i < n, (f i)ᶜ := by - cases n - · rw [disjointed_zero, eq_comm, inf_eq_left] - simp_rw [le_iInf_iff] - exact fun i hi => (i.not_lt_zero hi).elim - simp_rw [disjointed_succ, partialSups_eq_biSup, sdiff_eq, compl_iSup] - congr - ext i - rw [Nat.lt_succ_iff] +theorem disjointed_eq_inf_compl [Preorder ι] [LocallyFiniteOrderBot ι] (f : ι → α) (i : ι) : + disjointed f i = f i ⊓ ⨅ j < i, (f j)ᶜ := by + simp only [disjointed_apply, Finset.sup_eq_iSup, mem_Iio, sdiff_eq, compl_iSup] end CompleteBooleanAlgebra -/-! ### Set notation variants of lemmas -/ +section Set +/-! ### Lemmas specific to set-valued functions -/ -theorem disjointed_subset (f : ℕ → Set α) (n : ℕ) : disjointed f n ⊆ f n := - disjointed_le f n +theorem disjointed_subset [Preorder ι] [LocallyFiniteOrderBot ι] (f : ι → Set α) (i : ι) : + disjointed f i ⊆ f i := + disjointed_le f i -theorem iUnion_disjointed {f : ℕ → Set α} : ⋃ n, disjointed f n = ⋃ n, f n := +theorem iUnion_disjointed [PartialOrder ι] [LocallyFiniteOrderBot ι] {f : ι → Set α} : + ⋃ i, disjointed f i = ⋃ i, f i := iSup_disjointed f -theorem disjointed_eq_inter_compl (f : ℕ → Set α) (n : ℕ) : - disjointed f n = f n ∩ ⋂ i < n, (f i)ᶜ := - disjointed_eq_inf_compl f n +theorem disjointed_eq_inter_compl [Preorder ι] [LocallyFiniteOrderBot ι] (f : ι → Set α) (i : ι) : + disjointed f i = f i ∩ ⋂ j < i, (f j)ᶜ := + disjointed_eq_inf_compl f i theorem preimage_find_eq_disjointed (s : ℕ → Set α) (H : ∀ x, ∃ n, x ∈ s n) [∀ x n, Decidable (x ∈ s n)] (n : ℕ) : (fun x => Nat.find (H x)) ⁻¹' {n} = disjointed s n := by ext x simp [Nat.find_eq_iff, disjointed_eq_inter_compl] + +end Set + +section Nat + +/-! +### Functions on `ℕ` + +(See also `Mathlib.Algebra.Order.Disjointed` for results with more algebra pre-requsisites.) +-/ + +variable [GeneralizedBooleanAlgebra α] + +@[simp] +theorem disjointed_zero (f : ℕ → α) : disjointed f 0 = f 0 := + disjointed_bot f + +end Nat diff --git a/Mathlib/Order/Filter/CardinalInter.lean b/Mathlib/Order/Filter/CardinalInter.lean index 9a1e5bb744169..50c3f20937b94 100644 --- a/Mathlib/Order/Filter/CardinalInter.lean +++ b/Mathlib/Order/Filter/CardinalInter.lean @@ -6,8 +6,8 @@ Authors: Josha Dekker import Mathlib.Order.Filter.Tendsto import Mathlib.Order.Filter.Finite import Mathlib.Order.Filter.CountableInter -import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Cofinality +import Mathlib.Tactic.Linarith /-! # Filters with a cardinal intersection property diff --git a/Mathlib/Order/Filter/FilterProduct.lean b/Mathlib/Order/Filter/FilterProduct.lean index c36a347ed094f..c7c04367c9495 100644 --- a/Mathlib/Order/Filter/FilterProduct.lean +++ b/Mathlib/Order/Filter/FilterProduct.lean @@ -6,7 +6,7 @@ Authors: Abhimanyu Pallavi Sudhir, Yury Kudryashov import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Order.Filter.Ring -import Mathlib.Order.Filter.Ultrafilter +import Mathlib.Order.Filter.Ultrafilter.Defs /-! # Ultraproducts diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index 05f1a4c1cf4e0..fcd4bd41d71c3 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -6,8 +6,9 @@ Authors: Zhouhang Zhou, Yaël Dillies import Mathlib.Algebra.Order.Group.Defs import Mathlib.Algebra.Order.Group.OrderIso import Mathlib.Data.Set.Pointwise.SMul +import Mathlib.Order.Filter.AtTopBot import Mathlib.Order.Filter.NAry -import Mathlib.Order.Filter.Ultrafilter +import Mathlib.Order.Filter.Ultrafilter.Defs /-! # Pointwise operations on filters diff --git a/Mathlib/Order/Filter/SmallSets.lean b/Mathlib/Order/Filter/SmallSets.lean index 1e40a2da092eb..28c02d2e8dc15 100644 --- a/Mathlib/Order/Filter/SmallSets.lean +++ b/Mathlib/Order/Filter/SmallSets.lean @@ -54,6 +54,21 @@ theorem hasBasis_smallSets (l : Filter α) : HasBasis l.smallSets (fun t : Set α => t ∈ l) powerset := l.basis_sets.smallSets +theorem Eventually.exists_mem_basis_of_smallSets {p : ι → Prop} {s : ι → Set α} {P : Set α → Prop} + (h₁ : ∀ᶠ t in l.smallSets, P t) (h₂ : HasBasis l p s) : ∃ i, p i ∧ P (s i) := + (h₂.smallSets.eventually_iff.mp h₁).imp fun _i ⟨hpi, hi⟩ ↦ ⟨hpi, hi Subset.rfl⟩ + +theorem Frequently.smallSets_of_forall_mem_basis {p : ι → Prop} {s : ι → Set α} {P : Set α → Prop} + (h₁ : ∀ i, p i → P (s i)) (h₂ : HasBasis l p s) : ∃ᶠ t in l.smallSets, P t := + h₂.smallSets.frequently_iff.mpr fun _ hi => ⟨_, Subset.rfl, h₁ _ hi⟩ + +theorem Eventually.exists_mem_of_smallSets {p : Set α → Prop} + (h : ∀ᶠ t in l.smallSets, p t) : ∃ s ∈ l, p s := + h.exists_mem_basis_of_smallSets l.basis_sets + +/-! No `Frequently.smallSets_of_forall_mem (h : ∀ s ∈ l, p s) : ∃ᶠ t in l.smallSets, p t` as +`Filter.frequently_smallSets_mem : ∃ᶠ t in l.smallSets, t ∈ l` is preferred. -/ + /-- `g` converges to `f.smallSets` if for all `s ∈ f`, eventually we have `g x ⊆ s`. -/ theorem tendsto_smallSets_iff {f : α → Set β} : Tendsto f la lb.smallSets ↔ ∀ t ∈ lb, ∀ᶠ x in la, f x ⊆ t := diff --git a/Mathlib/Order/Filter/Subsingleton.lean b/Mathlib/Order/Filter/Subsingleton.lean index c8d0816c3f17b..7d82c9902f6a3 100644 --- a/Mathlib/Order/Filter/Subsingleton.lean +++ b/Mathlib/Order/Filter/Subsingleton.lean @@ -3,7 +3,9 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Order.Filter.Ultrafilter +import Mathlib.Order.Filter.CountablyGenerated +import Mathlib.Order.Filter.Prod +import Mathlib.Order.Filter.Ultrafilter.Defs /-! # Subsingleton filters diff --git a/Mathlib/Order/Filter/Ultrafilter/Basic.lean b/Mathlib/Order/Filter/Ultrafilter/Basic.lean new file mode 100644 index 0000000000000..4ba9a72dd8feb --- /dev/null +++ b/Mathlib/Order/Filter/Ultrafilter/Basic.lean @@ -0,0 +1,126 @@ +/- +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 +-/ +import Mathlib.Order.Filter.Ultrafilter.Defs +import Mathlib.Order.Filter.Cofinite +import Mathlib.Order.ZornAtoms + +/-! +# Ultrafilters + +An ultrafilter is a minimal (maximal in the set order) proper filter. +In this file we define + +* `hyperfilter`: the ultrafilter extending the cofinite filter. +-/ + +universe u v + +variable {α : Type u} {β : Type v} + +open Set Filter + +namespace Ultrafilter + +variable {f : Ultrafilter α} {s : Set α} + +theorem finite_sUnion_mem_iff {s : Set (Set α)} (hs : s.Finite) : ⋃₀ s ∈ f ↔ ∃ t ∈ s, t ∈ f := + Finite.induction_on _ hs (by simp) fun _ _ his => by + simp [union_mem_iff, his, or_and_right, exists_or] + +theorem finite_biUnion_mem_iff {is : Set β} {s : β → Set α} (his : is.Finite) : + (⋃ i ∈ is, s i) ∈ f ↔ ∃ i ∈ is, s i ∈ f := by + simp only [← sUnion_image, finite_sUnion_mem_iff (his.image s), exists_mem_image] + +lemma eventually_exists_mem_iff {is : Set β} {P : β → α → Prop} (his : is.Finite) : + (∀ᶠ i in f, ∃ a ∈ is, P a i) ↔ ∃ a ∈ is, ∀ᶠ i in f, P a i := by + simp only [Filter.Eventually, Ultrafilter.mem_coe] + convert f.finite_biUnion_mem_iff his (s := P) with i + aesop + +lemma eventually_exists_iff [Finite β] {P : β → α → Prop} : + (∀ᶠ i in f, ∃ a, P a i) ↔ ∃ a, ∀ᶠ i in f, P a i := by + simpa using eventually_exists_mem_iff (f := f) (P := P) Set.finite_univ + +theorem eq_pure_of_finite_mem (h : s.Finite) (h' : s ∈ f) : ∃ x ∈ s, f = pure x := by + rw [← biUnion_of_singleton s] at h' + rcases (Ultrafilter.finite_biUnion_mem_iff h).mp h' with ⟨a, has, haf⟩ + exact ⟨a, has, eq_of_le (Filter.le_pure_iff.2 haf)⟩ + +theorem eq_pure_of_finite [Finite α] (f : Ultrafilter α) : ∃ a, f = pure a := + (eq_pure_of_finite_mem finite_univ univ_mem).imp fun _ ⟨_, ha⟩ => ha + +theorem le_cofinite_or_eq_pure (f : Ultrafilter α) : (f : Filter α) ≤ cofinite ∨ ∃ a, f = pure a := + or_iff_not_imp_left.2 fun h => + let ⟨_, hs, hfin⟩ := Filter.disjoint_cofinite_right.1 (disjoint_iff_not_le.2 h) + let ⟨a, _, hf⟩ := eq_pure_of_finite_mem hfin hs + ⟨a, hf⟩ + +theorem exists_ultrafilter_of_finite_inter_nonempty (S : Set (Set α)) + (cond : ∀ T : Finset (Set α), (↑T : Set (Set α)) ⊆ S → (⋂₀ (↑T : Set (Set α))).Nonempty) : + ∃ F : Ultrafilter α, S ⊆ F.sets := + haveI : NeBot (generate S) := + generate_neBot_iff.2 fun _ hts ht => + ht.coe_toFinset ▸ cond ht.toFinset (ht.coe_toFinset.symm ▸ hts) + ⟨of (generate S), fun _ ht => (of_le <| generate S) <| GenerateSets.basic ht⟩ + +end Ultrafilter + +namespace Filter + +open Ultrafilter + +lemma atTop_eq_pure_of_isTop [LinearOrder α] {x : α} (hx : IsTop x) : + (atTop : Filter α) = pure x := by + have : Nonempty α := ⟨x⟩ + apply atTop_neBot.eq_pure_iff.2 + convert Ici_mem_atTop x using 1 + exact (Ici_eq_singleton_iff_isTop.2 hx).symm + +lemma atBot_eq_pure_of_isBot [LinearOrder α] {x : α} (hx : IsBot x) : + (atBot : Filter α) = pure x := + @atTop_eq_pure_of_isTop αᵒᵈ _ _ hx + +/-- The `tendsto` relation can be checked on ultrafilters. -/ +theorem tendsto_iff_ultrafilter (f : α → β) (l₁ : Filter α) (l₂ : Filter β) : + Tendsto f l₁ l₂ ↔ ∀ g : Ultrafilter α, ↑g ≤ l₁ → Tendsto f g l₂ := by + simpa only [tendsto_iff_comap] using le_iff_ultrafilter + +section Hyperfilter + +variable (α) [Infinite α] + +/-- The ultrafilter extending the cofinite filter. -/ +noncomputable def hyperfilter : Ultrafilter α := + Ultrafilter.of cofinite + +variable {α} + +theorem hyperfilter_le_cofinite : ↑(hyperfilter α) ≤ @cofinite α := + Ultrafilter.of_le cofinite + +theorem _root_.Nat.hyperfilter_le_atTop : (hyperfilter ℕ).toFilter ≤ atTop := + hyperfilter_le_cofinite.trans_eq Nat.cofinite_eq_atTop + +@[simp] +theorem bot_ne_hyperfilter : (⊥ : Filter α) ≠ hyperfilter α := + (NeBot.ne inferInstance).symm + +theorem nmem_hyperfilter_of_finite {s : Set α} (hf : s.Finite) : s ∉ hyperfilter α := fun hy => + compl_not_mem hy <| hyperfilter_le_cofinite hf.compl_mem_cofinite + +alias _root_.Set.Finite.nmem_hyperfilter := nmem_hyperfilter_of_finite + +theorem compl_mem_hyperfilter_of_finite {s : Set α} (hf : Set.Finite s) : sᶜ ∈ hyperfilter α := + compl_mem_iff_not_mem.2 hf.nmem_hyperfilter + +alias _root_.Set.Finite.compl_mem_hyperfilter := compl_mem_hyperfilter_of_finite + +theorem mem_hyperfilter_of_finite_compl {s : Set α} (hf : Set.Finite sᶜ) : s ∈ hyperfilter α := + compl_compl s ▸ hf.compl_mem_hyperfilter + +end Hyperfilter + +end Filter diff --git a/Mathlib/Order/Filter/Ultrafilter.lean b/Mathlib/Order/Filter/Ultrafilter/Defs.lean similarity index 78% rename from Mathlib/Order/Filter/Ultrafilter.lean rename to Mathlib/Order/Filter/Ultrafilter/Defs.lean index 2e3a26c64e31b..4b0710720dc94 100644 --- a/Mathlib/Order/Filter/Ultrafilter.lean +++ b/Mathlib/Order/Filter/Ultrafilter/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad, Yury Kudryashov -/ -import Mathlib.Order.Filter.Cofinite +import Mathlib.Order.Filter.Basic import Mathlib.Order.ZornAtoms /-! @@ -16,9 +16,10 @@ In this file we define * `Ultrafilter`: subtype of ultrafilters; * `pure x : Ultrafilter α`: `pure x` as an `Ultrafilter`; * `Ultrafilter.map`, `Ultrafilter.bind`, `Ultrafilter.comap` : operations on ultrafilters; -* `hyperfilter`: the ultrafilter extending the cofinite filter. -/ +assert_not_exists Set.Finite + universe u v variable {α : Type u} {β : Type v} {γ : Type*} @@ -160,24 +161,6 @@ theorem eventually_not : (∀ᶠ x in f, ¬p x) ↔ ¬∀ᶠ x in f, p x := theorem eventually_imp : (∀ᶠ x in f, p x → q x) ↔ (∀ᶠ x in f, p x) → ∀ᶠ x in f, q x := by simp only [imp_iff_not_or, eventually_or, eventually_not] -theorem finite_sUnion_mem_iff {s : Set (Set α)} (hs : s.Finite) : ⋃₀ s ∈ f ↔ ∃ t ∈ s, t ∈ f := - Finite.induction_on _ hs (by simp) fun _ _ his => by - simp [union_mem_iff, his, or_and_right, exists_or] - -theorem finite_biUnion_mem_iff {is : Set β} {s : β → Set α} (his : is.Finite) : - (⋃ i ∈ is, s i) ∈ f ↔ ∃ i ∈ is, s i ∈ f := by - simp only [← sUnion_image, finite_sUnion_mem_iff (his.image s), exists_mem_image] - -lemma eventually_exists_mem_iff {is : Set β} {P : β → α → Prop} (his : is.Finite) : - (∀ᶠ i in f, ∃ a ∈ is, P a i) ↔ ∃ a ∈ is, ∀ᶠ i in f, P a i := by - simp only [Filter.Eventually, Ultrafilter.mem_coe] - convert f.finite_biUnion_mem_iff his (s := P) with i - aesop - -lemma eventually_exists_iff [Finite β] {P : β → α → Prop} : - (∀ᶠ i in f, ∃ a, P a i) ↔ ∃ a, ∀ᶠ i in f, P a i := by - simpa using eventually_exists_mem_iff (f := f) (P := P) Set.finite_univ - /-- Pushforward for ultrafilters. -/ nonrec def map (m : α → β) (f : Ultrafilter α) : Ultrafilter β := ofComplNotMemIff (map m f) fun s => @compl_not_mem_iff _ f (m ⁻¹' s) @@ -269,20 +252,6 @@ instance [Inhabited α] : Inhabited (Ultrafilter α) := instance [Nonempty α] : Nonempty (Ultrafilter α) := Nonempty.map pure inferInstance -theorem eq_pure_of_finite_mem (h : s.Finite) (h' : s ∈ f) : ∃ x ∈ s, f = pure x := by - rw [← biUnion_of_singleton s] at h' - rcases (Ultrafilter.finite_biUnion_mem_iff h).mp h' with ⟨a, has, haf⟩ - exact ⟨a, has, eq_of_le (Filter.le_pure_iff.2 haf)⟩ - -theorem eq_pure_of_finite [Finite α] (f : Ultrafilter α) : ∃ a, f = pure a := - (eq_pure_of_finite_mem finite_univ univ_mem).imp fun _ ⟨_, ha⟩ => ha - -theorem le_cofinite_or_eq_pure (f : Ultrafilter α) : (f : Filter α) ≤ cofinite ∨ ∃ a, f = pure a := - or_iff_not_imp_left.2 fun h => - let ⟨_, hs, hfin⟩ := Filter.disjoint_cofinite_right.1 (disjoint_iff_not_le.2 h) - let ⟨a, _, hf⟩ := eq_pure_of_finite_mem hfin hs - ⟨a, hf⟩ - /-- Monadic bind for ultrafilters, coming from the one on filters defined in terms of map and join. -/ def bind (f : Ultrafilter α) (m : α → Ultrafilter β) : Ultrafilter β := @@ -332,14 +301,6 @@ theorem of_le (f : Filter α) [NeBot f] : ↑(of f) ≤ f := theorem of_coe (f : Ultrafilter α) : of ↑f = f := coe_inj.1 <| f.unique (of_le f.toFilter) -theorem exists_ultrafilter_of_finite_inter_nonempty (S : Set (Set α)) - (cond : ∀ T : Finset (Set α), (↑T : Set (Set α)) ⊆ S → (⋂₀ (↑T : Set (Set α))).Nonempty) : - ∃ F : Ultrafilter α, S ⊆ F.sets := - haveI : NeBot (generate S) := - generate_neBot_iff.2 fun _ hts ht => - ht.coe_toFinset ▸ cond ht.toFinset (ht.coe_toFinset.symm ▸ hts) - ⟨of (generate S), fun _ ht => (of_le <| generate S) <| GenerateSets.basic ht⟩ - end Ultrafilter namespace Filter @@ -358,17 +319,6 @@ protected theorem NeBot.eq_pure_iff (hf : f.NeBot) {x : α} : f = pure x ↔ {x} ∈ f := by rw [← hf.le_pure_iff, le_pure_iff] -lemma atTop_eq_pure_of_isTop [LinearOrder α] {x : α} (hx : IsTop x) : - (atTop : Filter α) = pure x := by - have : Nonempty α := ⟨x⟩ - apply atTop_neBot.eq_pure_iff.2 - convert Ici_mem_atTop x using 1 - exact (Ici_eq_singleton_iff_isTop.2 hx).symm - -lemma atBot_eq_pure_of_isBot [LinearOrder α] {x : α} (hx : IsBot x) : - (atBot : Filter α) = pure x := - @atTop_eq_pure_of_isTop αᵒᵈ _ _ hx - @[simp] theorem lt_pure_iff : f < pure a ↔ f = ⊥ := isAtom_pure.lt_iff @@ -394,11 +344,6 @@ theorem iSup_ultrafilter_le_eq (f : Filter α) : ⨆ (g : Ultrafilter α) (_ : g ≤ f), (g : Filter α) = f := eq_of_forall_ge_iff fun f' => by simp only [iSup_le_iff, ← le_iff_ultrafilter] -/-- The `tendsto` relation can be checked on ultrafilters. -/ -theorem tendsto_iff_ultrafilter (f : α → β) (l₁ : Filter α) (l₂ : Filter β) : - Tendsto f l₁ l₂ ↔ ∀ g : Ultrafilter α, ↑g ≤ l₁ → Tendsto f g l₂ := by - simpa only [tendsto_iff_comap] using le_iff_ultrafilter - theorem exists_ultrafilter_iff {f : Filter α} : (∃ u : Ultrafilter α, ↑u ≤ f) ↔ NeBot f := ⟨fun ⟨_, uf⟩ => neBot_of_le uf, fun h => @exists_ultrafilter_le _ _ h⟩ @@ -408,41 +353,6 @@ theorem forall_neBot_le_iff {g : Filter α} {p : Filter α → Prop} (hp : Monot intro H f hf hfg exact hp (of_le f) (H _ ((of_le f).trans hfg)) -section Hyperfilter - -variable (α) [Infinite α] - -/-- The ultrafilter extending the cofinite filter. -/ -noncomputable def hyperfilter : Ultrafilter α := - Ultrafilter.of cofinite - -variable {α} - -theorem hyperfilter_le_cofinite : ↑(hyperfilter α) ≤ @cofinite α := - Ultrafilter.of_le cofinite - -theorem _root_.Nat.hyperfilter_le_atTop : (hyperfilter ℕ).toFilter ≤ atTop := - hyperfilter_le_cofinite.trans_eq Nat.cofinite_eq_atTop - -@[simp] -theorem bot_ne_hyperfilter : (⊥ : Filter α) ≠ hyperfilter α := - (NeBot.ne inferInstance).symm - -theorem nmem_hyperfilter_of_finite {s : Set α} (hf : s.Finite) : s ∉ hyperfilter α := fun hy => - compl_not_mem hy <| hyperfilter_le_cofinite hf.compl_mem_cofinite - -alias _root_.Set.Finite.nmem_hyperfilter := nmem_hyperfilter_of_finite - -theorem compl_mem_hyperfilter_of_finite {s : Set α} (hf : Set.Finite s) : sᶜ ∈ hyperfilter α := - compl_mem_iff_not_mem.2 hf.nmem_hyperfilter - -alias _root_.Set.Finite.compl_mem_hyperfilter := compl_mem_hyperfilter_of_finite - -theorem mem_hyperfilter_of_finite_compl {s : Set α} (hf : Set.Finite sᶜ) : s ∈ hyperfilter α := - compl_compl s ▸ hf.compl_mem_hyperfilter - -end Hyperfilter - end Filter namespace Ultrafilter diff --git a/Mathlib/Order/Fin/Basic.lean b/Mathlib/Order/Fin/Basic.lean index 4e47785db09e5..174adbcde047a 100644 --- a/Mathlib/Order/Fin/Basic.lean +++ b/Mathlib/Order/Fin/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis, Keeley Hoek -/ -import Mathlib.Data.Fin.Basic +import Mathlib.Data.Fin.Rev import Mathlib.Order.Hom.Set /-! diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection/Basic.lean similarity index 57% rename from Mathlib/Order/GaloisConnection.lean rename to Mathlib/Order/GaloisConnection/Basic.lean index c5ad8900514d8..e27bddfa66024 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection/Basic.lean @@ -3,24 +3,15 @@ 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 -/ -import Mathlib.Order.CompleteLattice -import Mathlib.Order.Synonym -import Mathlib.Order.Hom.Set import Mathlib.Order.Bounds.Image +import Mathlib.Order.CompleteLattice +import Mathlib.Order.GaloisConnection.Defs /-! # Galois connections, insertions and coinsertions -Galois connections are order theoretic adjoints, i.e. a pair of functions `u` and `l`, -such that `∀ a b, l a ≤ b ↔ a ≤ u b`. - -## Main definitions - -* `GaloisConnection`: A Galois connection is a pair of functions `l` and `u` satisfying - `l a ≤ b ↔ a ≤ u b`. They are special cases of adjoint functors in category theory, - but do not depend on the category theory library in mathlib. -* `GaloisInsertion`: A Galois insertion is a Galois connection where `l ∘ u = id` -* `GaloisCoinsertion`: A Galois coinsertion is a Galois connection where `u ∘ l = id` +This file contains basic results on Galois connections, insertions and coinsertions in various +order structures, and provides constructions that lift order structures from one type to another. ## Implementation details @@ -53,54 +44,15 @@ universe u v w x variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a₁ a₂ : α} {b₁ b₂ : β} -/-- A Galois connection is a pair of functions `l` and `u` satisfying -`l a ≤ b ↔ a ≤ u b`. They are special cases of adjoint functors in category theory, -but do not depend on the category theory library in mathlib. -/ -def GaloisConnection [Preorder α] [Preorder β] (l : α → β) (u : β → α) := - ∀ a b, l a ≤ b ↔ a ≤ u b - -/-- Makes a Galois connection from an order-preserving bijection. -/ -theorem OrderIso.to_galoisConnection [Preorder α] [Preorder β] (oi : α ≃o β) : - GaloisConnection oi oi.symm := fun _ _ => oi.rel_symm_apply.symm - namespace GaloisConnection section variable [Preorder α] [Preorder β] {l : α → β} {u : β → α} -theorem monotone_intro (hu : Monotone u) (hl : Monotone l) (hul : ∀ a, a ≤ u (l a)) - (hlu : ∀ a, l (u a) ≤ a) : GaloisConnection l u := fun _ _ => - ⟨fun h => (hul _).trans (hu h), fun h => (hl h).trans (hlu _)⟩ - -protected theorem dual {l : α → β} {u : β → α} (gc : GaloisConnection l u) : - GaloisConnection (OrderDual.toDual ∘ u ∘ OrderDual.ofDual) - (OrderDual.toDual ∘ l ∘ OrderDual.ofDual) := - fun a b => (gc b a).symm - variable (gc : GaloisConnection l u) include gc -theorem le_iff_le {a : α} {b : β} : l a ≤ b ↔ a ≤ u b := - gc _ _ - -theorem l_le {a : α} {b : β} : a ≤ u b → l a ≤ b := - (gc _ _).mpr - -theorem le_u {a : α} {b : β} : l a ≤ b → a ≤ u b := - (gc _ _).mp - -theorem le_u_l (a) : a ≤ u (l a) := - gc.le_u <| le_rfl - -theorem l_u_le (a) : l (u a) ≤ a := - gc.l_le <| le_rfl - -theorem monotone_u : Monotone u := fun a _ H => gc.le_u ((gc.l_u_le a).trans H) - -theorem monotone_l : Monotone l := - gc.dual.monotone_u.dual - theorem upperBounds_l_image (s : Set α) : upperBounds (l '' s) = u ⁻¹' upperBounds s := Set.ext fun b => by simp [upperBounds, gc _ _] @@ -134,96 +86,8 @@ theorem isGLB_l {a : α} : IsGLB { b | a ≤ u b } (l a) := theorem isLUB_u {b : β} : IsLUB { a | l a ≤ b } (u b) := gc.isGreatest_u.isLUB -/-- If `(l, u)` is a Galois connection, then the relation `x ≤ u (l y)` is a transitive relation. -If `l` is a closure operator (`Submodule.span`, `Subgroup.closure`, ...) and `u` is the coercion to -`Set`, this reads as "if `U` is in the closure of `V` and `V` is in the closure of `W` then `U` is -in the closure of `W`". -/ -theorem le_u_l_trans {x y z : α} (hxy : x ≤ u (l y)) (hyz : y ≤ u (l z)) : x ≤ u (l z) := - hxy.trans (gc.monotone_u <| gc.l_le hyz) - -theorem l_u_le_trans {x y z : β} (hxy : l (u x) ≤ y) (hyz : l (u y) ≤ z) : l (u x) ≤ z := - (gc.monotone_l <| gc.le_u hxy).trans hyz - end -section PartialOrder - -variable [PartialOrder α] [Preorder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) -include gc - -theorem u_l_u_eq_u (b : β) : u (l (u b)) = u b := - (gc.monotone_u (gc.l_u_le _)).antisymm (gc.le_u_l _) - -theorem u_l_u_eq_u' : u ∘ l ∘ u = u := - funext gc.u_l_u_eq_u - -theorem u_unique {l' : α → β} {u' : β → α} (gc' : GaloisConnection l' u') (hl : ∀ a, l a = l' a) - {b : β} : u b = u' b := - le_antisymm (gc'.le_u <| hl (u b) ▸ gc.l_u_le _) (gc.le_u <| (hl (u' b)).symm ▸ gc'.l_u_le _) - -/-- If there exists a `b` such that `a = u a`, then `b = l a` is one such element. -/ -theorem exists_eq_u (a : α) : (∃ b : β, a = u b) ↔ a = u (l a) := - ⟨fun ⟨_, hS⟩ => hS.symm ▸ (gc.u_l_u_eq_u _).symm, fun HI => ⟨_, HI⟩⟩ - -theorem u_eq {z : α} {y : β} : u y = z ↔ ∀ x, x ≤ z ↔ l x ≤ y := by - constructor - · rintro rfl x - exact (gc x y).symm - · intro H - exact ((H <| u y).mpr (gc.l_u_le y)).antisymm ((gc _ _).mp <| (H z).mp le_rfl) - -end PartialOrder - -section PartialOrder - -variable [Preorder α] [PartialOrder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) -include gc - -theorem l_u_l_eq_l (a : α) : l (u (l a)) = l a := gc.dual.u_l_u_eq_u _ - -theorem l_u_l_eq_l' : l ∘ u ∘ l = l := funext gc.l_u_l_eq_l - -theorem l_unique {l' : α → β} {u' : β → α} (gc' : GaloisConnection l' u') (hu : ∀ b, u b = u' b) - {a : α} : l a = l' a := - gc.dual.u_unique gc'.dual hu - -/-- If there exists an `a` such that `b = l a`, then `a = u b` is one such element. -/ -theorem exists_eq_l (b : β) : (∃ a : α, b = l a) ↔ b = l (u b) := gc.dual.exists_eq_u _ - -theorem l_eq {x : α} {z : β} : l x = z ↔ ∀ y, z ≤ y ↔ x ≤ u y := gc.dual.u_eq - -end PartialOrder - -section OrderTop - -variable [PartialOrder α] [Preorder β] [OrderTop α] - -theorem u_eq_top {l : α → β} {u : β → α} (gc : GaloisConnection l u) {x} : u x = ⊤ ↔ l ⊤ ≤ x := - top_le_iff.symm.trans gc.le_iff_le.symm - -theorem u_top [OrderTop β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : u ⊤ = ⊤ := - gc.u_eq_top.2 le_top - -theorem u_l_top {l : α → β} {u : β → α} (gc : GaloisConnection l u) : u (l ⊤) = ⊤ := - gc.u_eq_top.mpr le_rfl - -end OrderTop - -section OrderBot - -variable [Preorder α] [PartialOrder β] [OrderBot β] - -theorem l_eq_bot {l : α → β} {u : β → α} (gc : GaloisConnection l u) {x} : l x = ⊥ ↔ x ≤ u ⊥ := - gc.dual.u_eq_top - -theorem l_bot [OrderBot α] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : l ⊥ = ⊥ := - gc.dual.u_top - -theorem l_u_bot {l : α → β} {u : β → α} (gc : GaloisConnection l u) : l (u ⊥) = ⊥ := - gc.l_eq_bot.mpr le_rfl - -end OrderBot - section SemilatticeSup variable [SemilatticeSup α] [SemilatticeSup β] {l : α → β} {u : β → α} @@ -272,31 +136,9 @@ theorem u_sInf {s : Set β} : u (sInf s) = ⨅ a ∈ s, u a := end CompleteLattice -section LinearOrder - -variable [LinearOrder α] [LinearOrder β] {l : α → β} {u : β → α} - -theorem lt_iff_lt (gc : GaloisConnection l u) {a : α} {b : β} : b < l a ↔ u b < a := - lt_iff_lt_of_le_iff_le (gc a b) - -end LinearOrder - -- Constructing Galois connections section Constructions -protected theorem id [pα : Preorder α] : @GaloisConnection α α pα pα id id := fun _ _ => - Iff.intro (fun x => x) fun x => x - -protected theorem compose [Preorder α] [Preorder β] [Preorder γ] {l1 : α → β} {u1 : β → α} - {l2 : β → γ} {u2 : γ → β} (gc1 : GaloisConnection l1 u1) (gc2 : GaloisConnection l2 u2) : - GaloisConnection (l2 ∘ l1) (u1 ∘ u2) := fun _ _ ↦ (gc2 _ _).trans (gc1 _ _) - -protected theorem dfun {ι : Type u} {α : ι → Type v} {β : ι → Type w} [∀ i, Preorder (α i)] - [∀ i, Preorder (β i)] (l : ∀ i, α i → β i) (u : ∀ i, β i → α i) - (gc : ∀ i, GaloisConnection (l i) (u i)) : - GaloisConnection (fun (a : ∀ i, α i) i => l i (a i)) fun b i => u i (b i) := fun a b => - forall_congr' fun i => gc i (a i) (b i) - protected theorem compl [BooleanAlgebra α] [BooleanAlgebra β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : GaloisConnection (compl ∘ u ∘ compl) (compl ∘ l ∘ compl) := fun a b ↦ by @@ -305,30 +147,6 @@ protected theorem compl [BooleanAlgebra α] [BooleanAlgebra β] {l : α → β} end Constructions -theorem l_comm_of_u_comm {X : Type*} [Preorder X] {Y : Type*} [Preorder Y] {Z : Type*} - [Preorder Z] {W : Type*} [PartialOrder W] {lYX : X → Y} {uXY : Y → X} - (hXY : GaloisConnection lYX uXY) {lWZ : Z → W} {uZW : W → Z} (hZW : GaloisConnection lWZ uZW) - {lWY : Y → W} {uYW : W → Y} (hWY : GaloisConnection lWY uYW) {lZX : X → Z} {uXZ : Z → X} - (hXZ : GaloisConnection lZX uXZ) (h : ∀ w, uXZ (uZW w) = uXY (uYW w)) {x : X} : - lWZ (lZX x) = lWY (lYX x) := - (hXZ.compose hZW).l_unique (hXY.compose hWY) h - -theorem u_comm_of_l_comm {X : Type*} [PartialOrder X] {Y : Type*} [Preorder Y] {Z : Type*} - [Preorder Z] {W : Type*} [Preorder W] {lYX : X → Y} {uXY : Y → X} - (hXY : GaloisConnection lYX uXY) {lWZ : Z → W} {uZW : W → Z} (hZW : GaloisConnection lWZ uZW) - {lWY : Y → W} {uYW : W → Y} (hWY : GaloisConnection lWY uYW) {lZX : X → Z} {uXZ : Z → X} - (hXZ : GaloisConnection lZX uXZ) (h : ∀ x, lWZ (lZX x) = lWY (lYX x)) {w : W} : - uXZ (uZW w) = uXY (uYW w) := - (hXZ.compose hZW).u_unique (hXY.compose hWY) h - -theorem l_comm_iff_u_comm {X : Type*} [PartialOrder X] {Y : Type*} [Preorder Y] {Z : Type*} - [Preorder Z] {W : Type*} [PartialOrder W] {lYX : X → Y} {uXY : Y → X} - (hXY : GaloisConnection lYX uXY) {lWZ : Z → W} {uZW : W → Z} (hZW : GaloisConnection lWZ uZW) - {lWY : Y → W} {uYW : W → Y} (hWY : GaloisConnection lWY uYW) {lZX : X → Z} {uXZ : Z → X} - (hXZ : GaloisConnection lZX uXZ) : - (∀ w : W, uXZ (uZW w) = uXY (uYW w)) ↔ ∀ x : X, lWZ (lZX x) = lWY (lYX x) := - ⟨hXY.l_comm_of_u_comm hZW hWY hXZ, hXY.u_comm_of_l_comm hZW hWY hXZ⟩ - end GaloisConnection section @@ -413,73 +231,10 @@ theorem galoisConnection_mul_div {k : ℕ} (h : 0 < k) : end Nat --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this used to have a `@[nolint has_nonempty_instance]` -/-- A Galois insertion is a Galois connection where `l ∘ u = id`. It also contains a constructive -choice function, to give better definitional equalities when lifting order structures. Dual -to `GaloisCoinsertion` -/ -structure GaloisInsertion {α β : Type*} [Preorder α] [Preorder β] (l : α → β) (u : β → α) where - /-- A contructive choice function for images of `l`. -/ - choice : ∀ x : α, u (l x) ≤ x → β - /-- The Galois connection associated to a Galois insertion. -/ - gc : GaloisConnection l u - /-- Main property of a Galois insertion. -/ - le_l_u : ∀ x, x ≤ l (u x) - /-- Property of the choice function. -/ - choice_eq : ∀ a h, choice a h = l a - -/-- A constructor for a Galois insertion with the trivial `choice` function. -/ -def GaloisInsertion.monotoneIntro {α β : Type*} [Preorder α] [Preorder β] {l : α → β} {u : β → α} - (hu : Monotone u) (hl : Monotone l) (hul : ∀ a, a ≤ u (l a)) (hlu : ∀ b, l (u b) = b) : - GaloisInsertion l u where - choice x _ := l x - gc := GaloisConnection.monotone_intro hu hl hul fun b => le_of_eq (hlu b) - le_l_u b := le_of_eq <| (hlu b).symm - choice_eq _ _ := rfl - -/-- Makes a Galois insertion from an order-preserving bijection. -/ -protected def OrderIso.toGaloisInsertion [Preorder α] [Preorder β] (oi : α ≃o β) : - GaloisInsertion oi oi.symm where - choice b _ := oi b - gc := oi.to_galoisConnection - le_l_u g := le_of_eq (oi.right_inv g).symm - choice_eq _ _ := rfl - -/-- Make a `GaloisInsertion l u` from a `GaloisConnection l u` such that `∀ b, b ≤ l (u b)` -/ -def GaloisConnection.toGaloisInsertion {α β : Type*} [Preorder α] [Preorder β] {l : α → β} - {u : β → α} (gc : GaloisConnection l u) (h : ∀ b, b ≤ l (u b)) : GaloisInsertion l u := - { choice := fun x _ => l x - gc - le_l_u := h - choice_eq := fun _ _ => rfl } - -/-- Lift the bottom along a Galois connection -/ -def GaloisConnection.liftOrderBot {α β : Type*} [Preorder α] [OrderBot α] [PartialOrder β] - {l : α → β} {u : β → α} (gc : GaloisConnection l u) : - OrderBot β where - bot := l ⊥ - bot_le _ := gc.l_le <| bot_le - namespace GaloisInsertion variable {l : α → β} {u : β → α} -theorem l_u_eq [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) (b : β) : l (u b) = b := - (gi.gc.l_u_le _).antisymm (gi.le_l_u _) - -theorem leftInverse_l_u [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) : - LeftInverse l u := - gi.l_u_eq - -theorem l_top [Preorder α] [PartialOrder β] [OrderTop α] [OrderTop β] - (gi : GaloisInsertion l u) : l ⊤ = ⊤ := - top_unique <| (gi.le_l_u _).trans <| gi.gc.monotone_l le_top - -theorem l_surjective [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) : Surjective l := - gi.leftInverse_l_u.surjective - -theorem u_injective [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) : Injective u := - gi.leftInverse_l_u.injective - theorem l_sup_u [SemilatticeSup α] [SemilatticeSup β] (gi : GaloisInsertion l u) (a b : β) : l (u a ⊔ u b) = a ⊔ b := calc @@ -530,12 +285,6 @@ theorem l_biInf_of_ul_eq_self [CompleteLattice α] [CompleteLattice β] (gi : Ga rw [iInf_subtype', iInf_subtype'] exact gi.l_iInf_of_ul_eq_self _ fun _ => hf _ _ -theorem u_le_u_iff [Preorder α] [Preorder β] (gi : GaloisInsertion l u) {a b} : u a ≤ u b ↔ a ≤ b := - ⟨fun h => (gi.le_l_u _).trans (gi.gc.l_le h), fun h => gi.gc.monotone_u h⟩ - -theorem strictMono_u [Preorder α] [Preorder β] (gi : GaloisInsertion l u) : StrictMono u := - strictMono_of_le_iff_le fun _ _ => gi.u_le_u_iff.symm - theorem isLUB_of_u_image [Preorder α] [Preorder β] (gi : GaloisInsertion l u) {s : Set β} {a : α} (hs : IsLUB (u '' s) a) : IsLUB s (l a) := ⟨fun x hx => (gi.le_l_u x).trans <| gi.gc.monotone_l <| hs.1 <| mem_image_of_mem _ hx, fun _ hx => @@ -612,94 +361,10 @@ end lift end GaloisInsertion --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this used to have a `@[nolint has_nonempty_instance]` -/-- A Galois coinsertion is a Galois connection where `u ∘ l = id`. It also contains a constructive -choice function, to give better definitional equalities when lifting order structures. Dual to -`GaloisInsertion` -/ -structure GaloisCoinsertion [Preorder α] [Preorder β] (l : α → β) (u : β → α) where - /-- A contructive choice function for images of `u`. -/ - choice : ∀ x : β, x ≤ l (u x) → α - /-- The Galois connection associated to a Galois coinsertion. -/ - gc : GaloisConnection l u - /-- Main property of a Galois coinsertion. -/ - u_l_le : ∀ x, u (l x) ≤ x - /-- Property of the choice function. -/ - choice_eq : ∀ a h, choice a h = u a - -/-- Make a `GaloisInsertion` between `αᵒᵈ` and `βᵒᵈ` from a `GaloisCoinsertion` between `α` and -`β`. -/ -def GaloisCoinsertion.dual [Preorder α] [Preorder β] {l : α → β} {u : β → α} : - GaloisCoinsertion l u → GaloisInsertion (toDual ∘ u ∘ ofDual) (toDual ∘ l ∘ ofDual) := - fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ - -/-- Make a `GaloisCoinsertion` between `αᵒᵈ` and `βᵒᵈ` from a `GaloisInsertion` between `α` and -`β`. -/ -def GaloisInsertion.dual [Preorder α] [Preorder β] {l : α → β} {u : β → α} : - GaloisInsertion l u → GaloisCoinsertion (toDual ∘ u ∘ ofDual) (toDual ∘ l ∘ ofDual) := - fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ - -/-- Make a `GaloisInsertion` between `α` and `β` from a `GaloisCoinsertion` between `αᵒᵈ` and -`βᵒᵈ`. -/ -def GaloisCoinsertion.ofDual [Preorder α] [Preorder β] {l : αᵒᵈ → βᵒᵈ} {u : βᵒᵈ → αᵒᵈ} : - GaloisCoinsertion l u → GaloisInsertion (ofDual ∘ u ∘ toDual) (ofDual ∘ l ∘ toDual) := - fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ - -/-- Make a `GaloisCoinsertion` between `α` and `β` from a `GaloisInsertion` between `αᵒᵈ` and -`βᵒᵈ`. -/ -def GaloisInsertion.ofDual [Preorder α] [Preorder β] {l : αᵒᵈ → βᵒᵈ} {u : βᵒᵈ → αᵒᵈ} : - GaloisInsertion l u → GaloisCoinsertion (ofDual ∘ u ∘ toDual) (ofDual ∘ l ∘ toDual) := - fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ - -/-- Makes a Galois coinsertion from an order-preserving bijection. -/ -protected def OrderIso.toGaloisCoinsertion [Preorder α] [Preorder β] (oi : α ≃o β) : - GaloisCoinsertion oi oi.symm where - choice b _ := oi.symm b - gc := oi.to_galoisConnection - u_l_le g := le_of_eq (oi.left_inv g) - choice_eq _ _ := rfl - -/-- A constructor for a Galois coinsertion with the trivial `choice` function. -/ -def GaloisCoinsertion.monotoneIntro [Preorder α] [Preorder β] {l : α → β} {u : β → α} - (hu : Monotone u) (hl : Monotone l) (hlu : ∀ b, l (u b) ≤ b) (hul : ∀ a, u (l a) = a) : - GaloisCoinsertion l u := - (GaloisInsertion.monotoneIntro hl.dual hu.dual hlu hul).ofDual - -/-- Make a `GaloisCoinsertion l u` from a `GaloisConnection l u` such that `∀ a, u (l a) ≤ a` -/ -def GaloisConnection.toGaloisCoinsertion {α β : Type*} [Preorder α] [Preorder β] {l : α → β} - {u : β → α} (gc : GaloisConnection l u) (h : ∀ a, u (l a) ≤ a) : GaloisCoinsertion l u := - { choice := fun x _ => u x - gc - u_l_le := h - choice_eq := fun _ _ => rfl } - -/-- Lift the top along a Galois connection -/ -def GaloisConnection.liftOrderTop {α β : Type*} [PartialOrder α] [Preorder β] [OrderTop β] - {l : α → β} {u : β → α} (gc : GaloisConnection l u) : - OrderTop α where - top := u ⊤ - le_top _ := gc.le_u <| le_top - namespace GaloisCoinsertion variable {l : α → β} {u : β → α} -theorem u_l_eq [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) (a : α) : u (l a) = a := - gi.dual.l_u_eq a - -theorem u_l_leftInverse [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) : - LeftInverse u l := - gi.u_l_eq - -theorem u_bot [PartialOrder α] [Preorder β] [OrderBot α] [OrderBot β] (gi : GaloisCoinsertion l u) : - u ⊥ = ⊥ := - gi.dual.l_top - -theorem u_surjective [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) : Surjective u := - gi.dual.l_surjective - -theorem l_injective [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) : Injective l := - gi.dual.u_injective - theorem u_inf_l [SemilatticeInf α] [SemilatticeInf β] (gi : GaloisCoinsertion l u) (a b : α) : u (l a ⊓ l b) = a ⊓ b := gi.dual.l_sup_u a b @@ -737,13 +402,6 @@ theorem u_biSup_of_lu_eq_self [CompleteLattice α] [CompleteLattice β] (gi : Ga u (⨆ (i) (hi), f i hi) = ⨆ (i) (hi), u (f i hi) := gi.dual.l_biInf_of_ul_eq_self _ hf -theorem l_le_l_iff [Preorder α] [Preorder β] (gi : GaloisCoinsertion l u) {a b} : - l a ≤ l b ↔ a ≤ b := - gi.dual.u_le_u_iff - -theorem strictMono_l [Preorder α] [Preorder β] (gi : GaloisCoinsertion l u) : StrictMono l := - fun _ _ h => gi.dual.strictMono_u h - theorem isGLB_of_l_image [Preorder α] [Preorder β] (gi : GaloisCoinsertion l u) {s : Set α} {a : β} (hs : IsGLB (l '' s) a) : IsGLB s (u a) := gi.dual.isLUB_of_u_image hs diff --git a/Mathlib/Order/GaloisConnection/Defs.lean b/Mathlib/Order/GaloisConnection/Defs.lean new file mode 100644 index 0000000000000..42e6d95b14602 --- /dev/null +++ b/Mathlib/Order/GaloisConnection/Defs.lean @@ -0,0 +1,395 @@ +/- +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 +-/ +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.Hom.Basic +import Mathlib.Order.Monotone.Basic + +/-! +# Galois connections, insertions and coinsertions + +Galois connections are order theoretic adjoints, i.e. a pair of functions `u` and `l`, +such that `∀ a b, l a ≤ b ↔ a ≤ u b`. + +## Main definitions + +* `GaloisConnection`: A Galois connection is a pair of functions `l` and `u` satisfying + `l a ≤ b ↔ a ≤ u b`. They are special cases of adjoint functors in category theory, + but do not depend on the category theory library in mathlib. +* `GaloisInsertion`: A Galois insertion is a Galois connection where `l ∘ u = id` +* `GaloisCoinsertion`: A Galois coinsertion is a Galois connection where `u ∘ l = id` +-/ + +assert_not_exists CompleteLattice + +open Function OrderDual Set + +universe u v w x + +variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a₁ a₂ : α} + {b₁ b₂ : β} + +/-- A Galois connection is a pair of functions `l` and `u` satisfying +`l a ≤ b ↔ a ≤ u b`. They are special cases of adjoint functors in category theory, +but do not depend on the category theory library in mathlib. -/ +def GaloisConnection [Preorder α] [Preorder β] (l : α → β) (u : β → α) := + ∀ a b, l a ≤ b ↔ a ≤ u b + +/-- Makes a Galois connection from an order-preserving bijection. -/ +theorem OrderIso.to_galoisConnection [Preorder α] [Preorder β] (oi : α ≃o β) : + GaloisConnection oi oi.symm := fun _ _ => oi.rel_symm_apply.symm + +namespace GaloisConnection + +section + +variable [Preorder α] [Preorder β] {l : α → β} {u : β → α} + +theorem monotone_intro (hu : Monotone u) (hl : Monotone l) (hul : ∀ a, a ≤ u (l a)) + (hlu : ∀ a, l (u a) ≤ a) : GaloisConnection l u := fun _ _ => + ⟨fun h => (hul _).trans (hu h), fun h => (hl h).trans (hlu _)⟩ + +protected theorem dual {l : α → β} {u : β → α} (gc : GaloisConnection l u) : + GaloisConnection (OrderDual.toDual ∘ u ∘ OrderDual.ofDual) + (OrderDual.toDual ∘ l ∘ OrderDual.ofDual) := + fun a b => (gc b a).symm + +variable (gc : GaloisConnection l u) +include gc + +theorem le_iff_le {a : α} {b : β} : l a ≤ b ↔ a ≤ u b := + gc _ _ + +theorem l_le {a : α} {b : β} : a ≤ u b → l a ≤ b := + (gc _ _).mpr + +theorem le_u {a : α} {b : β} : l a ≤ b → a ≤ u b := + (gc _ _).mp + +theorem le_u_l (a) : a ≤ u (l a) := + gc.le_u <| le_rfl + +theorem l_u_le (a) : l (u a) ≤ a := + gc.l_le <| le_rfl + +theorem monotone_u : Monotone u := fun a _ H => gc.le_u ((gc.l_u_le a).trans H) + +theorem monotone_l : Monotone l := + gc.dual.monotone_u.dual + +/-- If `(l, u)` is a Galois connection, then the relation `x ≤ u (l y)` is a transitive relation. +If `l` is a closure operator (`Submodule.span`, `Subgroup.closure`, ...) and `u` is the coercion to +`Set`, this reads as "if `U` is in the closure of `V` and `V` is in the closure of `W` then `U` is +in the closure of `W`". -/ +theorem le_u_l_trans {x y z : α} (hxy : x ≤ u (l y)) (hyz : y ≤ u (l z)) : x ≤ u (l z) := + hxy.trans (gc.monotone_u <| gc.l_le hyz) + +theorem l_u_le_trans {x y z : β} (hxy : l (u x) ≤ y) (hyz : l (u y) ≤ z) : l (u x) ≤ z := + (gc.monotone_l <| gc.le_u hxy).trans hyz + +end + +section PartialOrder + +variable [PartialOrder α] [Preorder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +include gc + +theorem u_l_u_eq_u (b : β) : u (l (u b)) = u b := + (gc.monotone_u (gc.l_u_le _)).antisymm (gc.le_u_l _) + +theorem u_l_u_eq_u' : u ∘ l ∘ u = u := + funext gc.u_l_u_eq_u + +theorem u_unique {l' : α → β} {u' : β → α} (gc' : GaloisConnection l' u') (hl : ∀ a, l a = l' a) + {b : β} : u b = u' b := + le_antisymm (gc'.le_u <| hl (u b) ▸ gc.l_u_le _) (gc.le_u <| (hl (u' b)).symm ▸ gc'.l_u_le _) + +/-- If there exists a `b` such that `a = u a`, then `b = l a` is one such element. -/ +theorem exists_eq_u (a : α) : (∃ b : β, a = u b) ↔ a = u (l a) := + ⟨fun ⟨_, hS⟩ => hS.symm ▸ (gc.u_l_u_eq_u _).symm, fun HI => ⟨_, HI⟩⟩ + +theorem u_eq {z : α} {y : β} : u y = z ↔ ∀ x, x ≤ z ↔ l x ≤ y := by + constructor + · rintro rfl x + exact (gc x y).symm + · intro H + exact ((H <| u y).mpr (gc.l_u_le y)).antisymm ((gc _ _).mp <| (H z).mp le_rfl) + +end PartialOrder + +section PartialOrder + +variable [Preorder α] [PartialOrder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +include gc + +theorem l_u_l_eq_l (a : α) : l (u (l a)) = l a := gc.dual.u_l_u_eq_u _ + +theorem l_u_l_eq_l' : l ∘ u ∘ l = l := funext gc.l_u_l_eq_l + +theorem l_unique {l' : α → β} {u' : β → α} (gc' : GaloisConnection l' u') (hu : ∀ b, u b = u' b) + {a : α} : l a = l' a := + gc.dual.u_unique gc'.dual hu + +/-- If there exists an `a` such that `b = l a`, then `a = u b` is one such element. -/ +theorem exists_eq_l (b : β) : (∃ a : α, b = l a) ↔ b = l (u b) := gc.dual.exists_eq_u _ + +theorem l_eq {x : α} {z : β} : l x = z ↔ ∀ y, z ≤ y ↔ x ≤ u y := gc.dual.u_eq + +end PartialOrder + +section OrderTop + +variable [PartialOrder α] [Preorder β] [OrderTop α] + +theorem u_eq_top {l : α → β} {u : β → α} (gc : GaloisConnection l u) {x} : u x = ⊤ ↔ l ⊤ ≤ x := + top_le_iff.symm.trans gc.le_iff_le.symm + +theorem u_top [OrderTop β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : u ⊤ = ⊤ := + gc.u_eq_top.2 le_top + +theorem u_l_top {l : α → β} {u : β → α} (gc : GaloisConnection l u) : u (l ⊤) = ⊤ := + gc.u_eq_top.mpr le_rfl + +end OrderTop + +section OrderBot + +variable [Preorder α] [PartialOrder β] [OrderBot β] + +theorem l_eq_bot {l : α → β} {u : β → α} (gc : GaloisConnection l u) {x} : l x = ⊥ ↔ x ≤ u ⊥ := + gc.dual.u_eq_top + +theorem l_bot [OrderBot α] {l : α → β} {u : β → α} (gc : GaloisConnection l u) : l ⊥ = ⊥ := + gc.dual.u_top + +theorem l_u_bot {l : α → β} {u : β → α} (gc : GaloisConnection l u) : l (u ⊥) = ⊥ := + gc.l_eq_bot.mpr le_rfl + +end OrderBot + +section LinearOrder + +variable [LinearOrder α] [LinearOrder β] {l : α → β} {u : β → α} + +theorem lt_iff_lt (gc : GaloisConnection l u) {a : α} {b : β} : b < l a ↔ u b < a := + lt_iff_lt_of_le_iff_le (gc a b) + +end LinearOrder + +-- Constructing Galois connections +section Constructions + +protected theorem id [pα : Preorder α] : @GaloisConnection α α pα pα id id := fun _ _ => + Iff.intro (fun x => x) fun x => x + +protected theorem compose [Preorder α] [Preorder β] [Preorder γ] {l1 : α → β} {u1 : β → α} + {l2 : β → γ} {u2 : γ → β} (gc1 : GaloisConnection l1 u1) (gc2 : GaloisConnection l2 u2) : + GaloisConnection (l2 ∘ l1) (u1 ∘ u2) := fun _ _ ↦ (gc2 _ _).trans (gc1 _ _) + +protected theorem dfun {ι : Type u} {α : ι → Type v} {β : ι → Type w} [∀ i, Preorder (α i)] + [∀ i, Preorder (β i)] (l : ∀ i, α i → β i) (u : ∀ i, β i → α i) + (gc : ∀ i, GaloisConnection (l i) (u i)) : + GaloisConnection (fun (a : ∀ i, α i) i => l i (a i)) fun b i => u i (b i) := fun a b => + forall_congr' fun i => gc i (a i) (b i) + +end Constructions + +theorem l_comm_of_u_comm {X : Type*} [Preorder X] {Y : Type*} [Preorder Y] {Z : Type*} + [Preorder Z] {W : Type*} [PartialOrder W] {lYX : X → Y} {uXY : Y → X} + (hXY : GaloisConnection lYX uXY) {lWZ : Z → W} {uZW : W → Z} (hZW : GaloisConnection lWZ uZW) + {lWY : Y → W} {uYW : W → Y} (hWY : GaloisConnection lWY uYW) {lZX : X → Z} {uXZ : Z → X} + (hXZ : GaloisConnection lZX uXZ) (h : ∀ w, uXZ (uZW w) = uXY (uYW w)) {x : X} : + lWZ (lZX x) = lWY (lYX x) := + (hXZ.compose hZW).l_unique (hXY.compose hWY) h + +theorem u_comm_of_l_comm {X : Type*} [PartialOrder X] {Y : Type*} [Preorder Y] {Z : Type*} + [Preorder Z] {W : Type*} [Preorder W] {lYX : X → Y} {uXY : Y → X} + (hXY : GaloisConnection lYX uXY) {lWZ : Z → W} {uZW : W → Z} (hZW : GaloisConnection lWZ uZW) + {lWY : Y → W} {uYW : W → Y} (hWY : GaloisConnection lWY uYW) {lZX : X → Z} {uXZ : Z → X} + (hXZ : GaloisConnection lZX uXZ) (h : ∀ x, lWZ (lZX x) = lWY (lYX x)) {w : W} : + uXZ (uZW w) = uXY (uYW w) := + (hXZ.compose hZW).u_unique (hXY.compose hWY) h + +theorem l_comm_iff_u_comm {X : Type*} [PartialOrder X] {Y : Type*} [Preorder Y] {Z : Type*} + [Preorder Z] {W : Type*} [PartialOrder W] {lYX : X → Y} {uXY : Y → X} + (hXY : GaloisConnection lYX uXY) {lWZ : Z → W} {uZW : W → Z} (hZW : GaloisConnection lWZ uZW) + {lWY : Y → W} {uYW : W → Y} (hWY : GaloisConnection lWY uYW) {lZX : X → Z} {uXZ : Z → X} + (hXZ : GaloisConnection lZX uXZ) : + (∀ w : W, uXZ (uZW w) = uXY (uYW w)) ↔ ∀ x : X, lWZ (lZX x) = lWY (lYX x) := + ⟨hXY.l_comm_of_u_comm hZW hWY hXZ, hXY.u_comm_of_l_comm hZW hWY hXZ⟩ + +end GaloisConnection + +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this used to have a `@[nolint has_nonempty_instance]` +/-- A Galois insertion is a Galois connection where `l ∘ u = id`. It also contains a constructive +choice function, to give better definitional equalities when lifting order structures. Dual +to `GaloisCoinsertion` -/ +structure GaloisInsertion {α β : Type*} [Preorder α] [Preorder β] (l : α → β) (u : β → α) where + /-- A contructive choice function for images of `l`. -/ + choice : ∀ x : α, u (l x) ≤ x → β + /-- The Galois connection associated to a Galois insertion. -/ + gc : GaloisConnection l u + /-- Main property of a Galois insertion. -/ + le_l_u : ∀ x, x ≤ l (u x) + /-- Property of the choice function. -/ + choice_eq : ∀ a h, choice a h = l a + +/-- Makes a Galois insertion from an order-preserving bijection. -/ +protected def OrderIso.toGaloisInsertion [Preorder α] [Preorder β] (oi : α ≃o β) : + GaloisInsertion oi oi.symm where + choice b _ := oi b + gc := oi.to_galoisConnection + le_l_u g := le_of_eq (oi.right_inv g).symm + choice_eq _ _ := rfl + +/-- A constructor for a Galois insertion with the trivial `choice` function. -/ +def GaloisInsertion.monotoneIntro {α β : Type*} [Preorder α] [Preorder β] {l : α → β} {u : β → α} + (hu : Monotone u) (hl : Monotone l) (hul : ∀ a, a ≤ u (l a)) (hlu : ∀ b, l (u b) = b) : + GaloisInsertion l u where + choice x _ := l x + gc := GaloisConnection.monotone_intro hu hl hul fun b => le_of_eq (hlu b) + le_l_u b := le_of_eq <| (hlu b).symm + choice_eq _ _ := rfl + +/-- Make a `GaloisInsertion l u` from a `GaloisConnection l u` such that `∀ b, b ≤ l (u b)` -/ +def GaloisConnection.toGaloisInsertion {α β : Type*} [Preorder α] [Preorder β] {l : α → β} + {u : β → α} (gc : GaloisConnection l u) (h : ∀ b, b ≤ l (u b)) : GaloisInsertion l u := + { choice := fun x _ => l x + gc + le_l_u := h + choice_eq := fun _ _ => rfl } + +/-- Lift the bottom along a Galois connection -/ +def GaloisConnection.liftOrderBot {α β : Type*} [Preorder α] [OrderBot α] [PartialOrder β] + {l : α → β} {u : β → α} (gc : GaloisConnection l u) : + OrderBot β where + bot := l ⊥ + bot_le _ := gc.l_le <| bot_le + +namespace GaloisInsertion + +variable {l : α → β} {u : β → α} + +theorem l_u_eq [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) (b : β) : l (u b) = b := + (gi.gc.l_u_le _).antisymm (gi.le_l_u _) + +theorem leftInverse_l_u [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) : + LeftInverse l u := + gi.l_u_eq + +theorem l_top [Preorder α] [PartialOrder β] [OrderTop α] [OrderTop β] + (gi : GaloisInsertion l u) : l ⊤ = ⊤ := + top_unique <| (gi.le_l_u _).trans <| gi.gc.monotone_l le_top + +theorem l_surjective [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) : Surjective l := + gi.leftInverse_l_u.surjective + +theorem u_injective [Preorder α] [PartialOrder β] (gi : GaloisInsertion l u) : Injective u := + gi.leftInverse_l_u.injective + +theorem u_le_u_iff [Preorder α] [Preorder β] (gi : GaloisInsertion l u) {a b} : u a ≤ u b ↔ a ≤ b := + ⟨fun h => (gi.le_l_u _).trans (gi.gc.l_le h), fun h => gi.gc.monotone_u h⟩ + +theorem strictMono_u [Preorder α] [Preorder β] (gi : GaloisInsertion l u) : StrictMono u := + strictMono_of_le_iff_le fun _ _ => gi.u_le_u_iff.symm + +end GaloisInsertion + +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this used to have a `@[nolint has_nonempty_instance]` +/-- A Galois coinsertion is a Galois connection where `u ∘ l = id`. It also contains a constructive +choice function, to give better definitional equalities when lifting order structures. Dual to +`GaloisInsertion` -/ +structure GaloisCoinsertion [Preorder α] [Preorder β] (l : α → β) (u : β → α) where + /-- A contructive choice function for images of `u`. -/ + choice : ∀ x : β, x ≤ l (u x) → α + /-- The Galois connection associated to a Galois coinsertion. -/ + gc : GaloisConnection l u + /-- Main property of a Galois coinsertion. -/ + u_l_le : ∀ x, u (l x) ≤ x + /-- Property of the choice function. -/ + choice_eq : ∀ a h, choice a h = u a + +/-- Makes a Galois coinsertion from an order-preserving bijection. -/ +protected def OrderIso.toGaloisCoinsertion [Preorder α] [Preorder β] (oi : α ≃o β) : + GaloisCoinsertion oi oi.symm where + choice b _ := oi.symm b + gc := oi.to_galoisConnection + u_l_le g := le_of_eq (oi.left_inv g) + choice_eq _ _ := rfl + +/-- Make a `GaloisInsertion` between `αᵒᵈ` and `βᵒᵈ` from a `GaloisCoinsertion` between `α` and +`β`. -/ +def GaloisCoinsertion.dual [Preorder α] [Preorder β] {l : α → β} {u : β → α} : + GaloisCoinsertion l u → GaloisInsertion (toDual ∘ u ∘ ofDual) (toDual ∘ l ∘ ofDual) := + fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ + +/-- Make a `GaloisCoinsertion` between `αᵒᵈ` and `βᵒᵈ` from a `GaloisInsertion` between `α` and +`β`. -/ +def GaloisInsertion.dual [Preorder α] [Preorder β] {l : α → β} {u : β → α} : + GaloisInsertion l u → GaloisCoinsertion (toDual ∘ u ∘ ofDual) (toDual ∘ l ∘ ofDual) := + fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ + +/-- Make a `GaloisInsertion` between `α` and `β` from a `GaloisCoinsertion` between `αᵒᵈ` and +`βᵒᵈ`. -/ +def GaloisCoinsertion.ofDual [Preorder α] [Preorder β] {l : αᵒᵈ → βᵒᵈ} {u : βᵒᵈ → αᵒᵈ} : + GaloisCoinsertion l u → GaloisInsertion (ofDual ∘ u ∘ toDual) (ofDual ∘ l ∘ toDual) := + fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ + +/-- Make a `GaloisCoinsertion` between `α` and `β` from a `GaloisInsertion` between `αᵒᵈ` and +`βᵒᵈ`. -/ +def GaloisInsertion.ofDual [Preorder α] [Preorder β] {l : αᵒᵈ → βᵒᵈ} {u : βᵒᵈ → αᵒᵈ} : + GaloisInsertion l u → GaloisCoinsertion (ofDual ∘ u ∘ toDual) (ofDual ∘ l ∘ toDual) := + fun x => ⟨x.1, x.2.dual, x.3, x.4⟩ + +/-- A constructor for a Galois coinsertion with the trivial `choice` function. -/ +def GaloisCoinsertion.monotoneIntro [Preorder α] [Preorder β] {l : α → β} {u : β → α} + (hu : Monotone u) (hl : Monotone l) (hlu : ∀ b, l (u b) ≤ b) (hul : ∀ a, u (l a) = a) : + GaloisCoinsertion l u := + (GaloisInsertion.monotoneIntro hl.dual hu.dual hlu hul).ofDual + +/-- Make a `GaloisCoinsertion l u` from a `GaloisConnection l u` such that `∀ a, u (l a) ≤ a` -/ +def GaloisConnection.toGaloisCoinsertion {α β : Type*} [Preorder α] [Preorder β] {l : α → β} + {u : β → α} (gc : GaloisConnection l u) (h : ∀ a, u (l a) ≤ a) : GaloisCoinsertion l u := + { choice := fun x _ => u x + gc + u_l_le := h + choice_eq := fun _ _ => rfl } + +/-- Lift the top along a Galois connection -/ +def GaloisConnection.liftOrderTop {α β : Type*} [PartialOrder α] [Preorder β] [OrderTop β] + {l : α → β} {u : β → α} (gc : GaloisConnection l u) : + OrderTop α where + top := u ⊤ + le_top _ := gc.le_u <| le_top + +namespace GaloisCoinsertion + +variable {l : α → β} {u : β → α} + +theorem u_l_eq [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) (a : α) : u (l a) = a := + gi.dual.l_u_eq a + +theorem u_l_leftInverse [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) : + LeftInverse u l := + gi.u_l_eq + +theorem u_bot [PartialOrder α] [Preorder β] [OrderBot α] [OrderBot β] (gi : GaloisCoinsertion l u) : + u ⊥ = ⊥ := + gi.dual.l_top + +theorem u_surjective [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) : Surjective u := + gi.dual.l_surjective + +theorem l_injective [PartialOrder α] [Preorder β] (gi : GaloisCoinsertion l u) : Injective l := + gi.dual.u_injective + +theorem l_le_l_iff [Preorder α] [Preorder β] (gi : GaloisCoinsertion l u) {a b} : + l a ≤ l b ↔ a ≤ b := + gi.dual.u_le_u_iff + +theorem strictMono_l [Preorder α] [Preorder β] (gi : GaloisCoinsertion l u) : StrictMono l := + fun _ _ h => gi.dual.strictMono_u h + +end GaloisCoinsertion diff --git a/Mathlib/Order/Heyting/Regular.lean b/Mathlib/Order/Heyting/Regular.lean index e58ab2d79ed47..6af155458a241 100644 --- a/Mathlib/Order/Heyting/Regular.lean +++ b/Mathlib/Order/Heyting/Regular.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Basic /-! # Heyting regular elements diff --git a/Mathlib/Order/Hom/Lattice.lean b/Mathlib/Order/Hom/Lattice.lean index 8d090e9e4b22a..306587a3fd1f8 100644 --- a/Mathlib/Order/Hom/Lattice.lean +++ b/Mathlib/Order/Hom/Lattice.lean @@ -44,39 +44,59 @@ variable {F ι α β γ δ : Type*} /-- The type of `⊔`-preserving functions from `α` to `β`. -/ structure SupHom (α β : Type*) [Max α] [Max β] where - /-- The underlying function of a `SupHom` -/ + /-- The underlying function of a `SupHom`. + + Do not use this function directly. Instead use the coercion coming from the `FunLike` + instance. -/ toFun : α → β - /-- A `SupHom` preserves suprema. -/ + /-- A `SupHom` preserves suprema. + + Do not use this directly. Use `map_sup` instead. -/ map_sup' (a b : α) : toFun (a ⊔ b) = toFun a ⊔ toFun b /-- The type of `⊓`-preserving functions from `α` to `β`. -/ structure InfHom (α β : Type*) [Min α] [Min β] where - /-- The underlying function of an `InfHom` -/ + /-- The underlying function of an `InfHom`. + + Do not use this function directly. Instead use the coercion coming from the `FunLike` + instance. -/ toFun : α → β - /-- An `InfHom` preserves infima. -/ + /-- An `InfHom` preserves infima. + + Do not use this directly. Use `map_inf` instead. -/ map_inf' (a b : α) : toFun (a ⊓ b) = toFun a ⊓ toFun b /-- The type of finitary supremum-preserving homomorphisms from `α` to `β`. -/ structure SupBotHom (α β : Type*) [Max α] [Max β] [Bot α] [Bot β] extends SupHom α β where - /-- A `SupBotHom` preserves the bottom element. -/ + /-- A `SupBotHom` preserves the bottom element. + + Do not use this directly. Use `map_bot` instead. -/ map_bot' : toFun ⊥ = ⊥ /-- The type of finitary infimum-preserving homomorphisms from `α` to `β`. -/ structure InfTopHom (α β : Type*) [Min α] [Min β] [Top α] [Top β] extends InfHom α β where - /-- An `InfTopHom` preserves the top element. -/ + /-- An `InfTopHom` preserves the top element. + + Do not use this directly. Use `map_top` instead. -/ map_top' : toFun ⊤ = ⊤ /-- The type of lattice homomorphisms from `α` to `β`. -/ structure LatticeHom (α β : Type*) [Lattice α] [Lattice β] extends SupHom α β where - /-- A `LatticeHom` preserves infima. -/ + /-- A `LatticeHom` preserves infima. + + Do not use this directly. Use `map_inf` instead. -/ map_inf' (a b : α) : toFun (a ⊓ b) = toFun a ⊓ toFun b /-- The type of bounded lattice homomorphisms from `α` to `β`. -/ structure BoundedLatticeHom (α β : Type*) [Lattice α] [Lattice β] [BoundedOrder α] [BoundedOrder β] extends LatticeHom α β where - /-- A `BoundedLatticeHom` preserves the top element. -/ + /-- A `BoundedLatticeHom` preserves the top element. + + Do not use this directly. Use `map_top` instead. -/ map_top' : toFun ⊤ = ⊤ - /-- A `BoundedLatticeHom` preserves the bottom element. -/ + /-- A `BoundedLatticeHom` preserves the bottom element. + + Do not use this directly. Use `map_bot` instead. -/ map_bot' : toFun ⊥ = ⊥ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: remove this configuration and use the default configuration. diff --git a/Mathlib/Order/Hom/Order.lean b/Mathlib/Order/Hom/Order.lean index 730ef64eb8a36..e543b1d530b51 100644 --- a/Mathlib/Order/Hom/Order.lean +++ b/Mathlib/Order/Hom/Order.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Anne Baanen -/ import Mathlib.Logic.Function.Iterate -import Mathlib.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Basic import Mathlib.Order.Hom.Basic /-! diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 291d014cb4e11..e684d9af2792b 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -200,11 +200,11 @@ instance setLike : SetLike (NonemptyInterval α) α where coe s := Icc s.fst s.snd coe_injective' := coeHom.injective -@[norm_cast] -- @[simp, norm_cast] -- Porting note: not in simpNF +@[norm_cast] theorem coe_subset_coe : (s : Set α) ⊆ t ↔ (s : NonemptyInterval α) ≤ t := (@coeHom α _).le_iff_le -@[norm_cast] -- @[simp, norm_cast] -- Porting note: not in simpNF +@[norm_cast] theorem coe_ssubset_coe : (s : Set α) ⊂ t ↔ s < t := (@coeHom α _).lt_iff_lt @@ -292,7 +292,7 @@ def recBotCoe {C : Interval α → Sort*} (bot : C ⊥) (coe : ∀ a : NonemptyI theorem coe_injective : Injective ((↑) : NonemptyInterval α → Interval α) := WithBot.coe_injective -@[norm_cast] -- @[simp, norm_cast] -- Porting note: not in simpNF +@[norm_cast] theorem coe_inj {s t : NonemptyInterval α} : (s : Interval α) = t ↔ s = t := WithBot.coe_inj @@ -400,11 +400,11 @@ instance setLike : SetLike (Interval α) α where coe := coeHom coe_injective' := coeHom.injective -@[norm_cast] -- @[simp, norm_cast] -- Porting note: not in simpNF +@[norm_cast] theorem coe_subset_coe : (s : Set α) ⊆ t ↔ s ≤ t := (@coeHom α _).le_iff_le -@[norm_cast] -- @[simp, norm_cast] -- Porting note: not in simpNF +@[norm_cast] theorem coe_sSubset_coe : (s : Set α) ⊂ t ↔ s < t := (@coeHom α _).lt_iff_lt @@ -681,7 +681,6 @@ theorem coe_sInf [DecidableRel (α := α) (· ≤ ·)] (S : Set (Interval α)) : theorem coe_iInf [DecidableRel (α := α) (· ≤ ·)] (f : ι → Interval α) : ↑(⨅ i, f i) = ⋂ i, (f i : Set α) := by simp [iInf] --- @[simp] -- Porting note: not in simpNF @[norm_cast] theorem coe_iInf₂ [DecidableRel (α := α) (· ≤ ·)] (f : ∀ i, κ i → Interval α) : ↑(⨅ (i) (j), f i j) = ⋂ (i) (j), (f i j : Set α) := by simp_rw [coe_iInf] diff --git a/Mathlib/Order/Interval/Finset/Box.lean b/Mathlib/Order/Interval/Finset/Box.lean index c84087a315416..9519e4a5286d7 100644 --- a/Mathlib/Order/Interval/Finset/Box.lean +++ b/Mathlib/Order/Interval/Finset/Box.lean @@ -3,10 +3,9 @@ Copyright (c) 2024 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.Disjointed import Mathlib.Algebra.Order.Ring.Prod import Mathlib.Data.Int.Interval -import Mathlib.Order.Disjointed -import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Ring import Mathlib.Tactic.Zify @@ -37,13 +36,16 @@ def box : ℕ → Finset α := disjointed fun n ↦ Icc (-n : α) n @[simp] lemma box_zero : (box 0 : Finset α) = {0} := by simp [box] lemma box_succ_eq_sdiff (n : ℕ) : - box (n + 1) = Icc (-n.succ : α) n.succ \ Icc (-n) n := Icc_neg_mono.disjointed_succ _ + box (n + 1) = Icc (-n.succ : α) n.succ \ Icc (-n) n := by + rw [box, Icc_neg_mono.disjointed_add_one] + simp only [Nat.cast_add_one, Nat.succ_eq_add_one] lemma disjoint_box_succ_prod (n : ℕ) : Disjoint (box (n + 1)) (Icc (-n : α) n) := by rw [box_succ_eq_sdiff]; exact disjoint_sdiff_self_left @[simp] lemma box_succ_union_prod (n : ℕ) : - box (n + 1) ∪ Icc (-n : α) n = Icc (-n.succ : α) n.succ := Icc_neg_mono.disjointed_succ_sup _ + box (n + 1) ∪ Icc (-n : α) n = Icc (-n.succ : α) n.succ := + Icc_neg_mono.disjointed_add_one_sup _ lemma box_succ_disjUnion (n : ℕ) : (box (n + 1)).disjUnion (Icc (-n : α) n) (disjoint_box_succ_prod _) = diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index d446eb5f35743..56f30d50f8236 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -65,6 +65,9 @@ lemma range_eq_Icc_zero_sub_one (n : ℕ) (hn : n ≠ 0) : range n = Icc 0 (n - theorem _root_.Finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm +theorem range_succ_eq_Icc_zero (n : ℕ) : range (n + 1) = Icc 0 n := by + rw [range_eq_Icc_zero_sub_one _ (Nat.add_one_ne_zero _), Nat.add_sub_cancel_right] + @[simp] lemma card_Icc : #(Icc a b) = b + 1 - a := List.length_range' .. @[simp] lemma card_Ico : #(Ico a b) = b - a := List.length_range' .. @[simp] lemma card_Ioc : #(Ioc a b) = b - a := List.length_range' .. diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index aecd91615f063..f55cfb93fbc83 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -76,14 +76,18 @@ noncomputable def krullDim (α : Type*) [Preorder α] : WithBot ℕ∞ := /-- The **height** of an element `a` in a preorder `α` is the supremum of the rightmost index of all -relation series of `α` ordered by `<` and ending below or at `a`. +relation series of `α` ordered by `<` and ending below or at `a`. In other words, it is +the largest `n` such that there's a series `a₀ < a₁ < ... < aₙ = a` (or `∞` if there is +no largest `n`). -/ noncomputable def height {α : Type*} [Preorder α] (a : α) : ℕ∞ := ⨆ (p : LTSeries α) (_ : p.last ≤ a), p.length /-- The **coheight** of an element `a` in a preorder `α` is the supremum of the rightmost index of all -relation series of `α` ordered by `<` and beginning with `a`. +relation series of `α` ordered by `<` and beginning with `a`. In other words, it is +the largest `n` such that there's a series `a = a₀ < a₁ < ... < aₙ` (or `∞` if there is +no largest `n`). The definition of `coheight` is via the `height` in the dual order, in order to easily transfer theorems between `height` and `coheight`. See `coheight_eq` for the definition with a @@ -275,11 +279,37 @@ private lemma height_add_const (a : α) (n : ℕ∞) : have := length_le_height_last (p := p.snoc y (by simp [*])) simpa using this +lemma height_add_one_le {a b : α} (hab : a < b) : height a + 1 ≤ height b := by + cases hfin : height a with + | top => + have : ⊤ ≤ height b := by + rw [← hfin] + gcongr + simp [this] + | coe n => + apply Order.add_one_le_of_lt + rw [← hfin] + gcongr + simp [hfin] + /- For elements of finite height, `coheight` is strictly antitone. -/ @[gcongr] lemma coheight_strictAnti {x y : α} (hyx : y < x) (hfin : coheight x < ⊤) : coheight x < coheight y := height_strictMono (α := αᵒᵈ) hyx hfin +lemma coheight_add_one_le {a b : α} (hab : b < a) : coheight a + 1 ≤ coheight b := by + cases hfin : coheight a with + | top => + have : ⊤ ≤ coheight b := by + rw [← hfin] + gcongr + simp [this] + | coe n => + apply Order.add_one_le_of_lt + rw [← hfin] + gcongr + simp [hfin] + lemma height_le_height_apply_of_strictMono (f : α → β) (hf : StrictMono f) (x : α) : height x ≤ height (f x) := by simp only [height_eq_iSup_last_eq] diff --git a/Mathlib/Order/ModularLattice.lean b/Mathlib/Order/ModularLattice.lean index 7b2b7f77ae7f8..0f8ca712abcc2 100644 --- a/Mathlib/Order/ModularLattice.lean +++ b/Mathlib/Order/ModularLattice.lean @@ -5,7 +5,7 @@ Authors: Aaron Anderson, Yaël Dillies -/ import Mathlib.Order.Cover import Mathlib.Order.LatticeIntervals -import Mathlib.Order.GaloisConnection +import Mathlib.Order.GaloisConnection.Defs /-! # Modular Lattices diff --git a/Mathlib/Order/Nucleus.lean b/Mathlib/Order/Nucleus.lean new file mode 100644 index 0000000000000..e867d9f21930b --- /dev/null +++ b/Mathlib/Order/Nucleus.lean @@ -0,0 +1,113 @@ +/- +Copyright (c) 2024 Christian Krause. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chriara Cimino, Christian Krause +-/ +import Mathlib.Order.Closure +import Mathlib.Order.CompleteLattice +import Mathlib.Order.Hom.Lattice + +/-! +# Nucleus + +Locales are the dual concept to frames. Locale theory is a branch of point-free topology, where +intuitively locales are like topological spaces which may or may not have enough points. +Sublocales of a locale generalize the concept of subspaces in topology to the point-free setting. + +A nucleus is an endomorphism of a frame which corresponds to a sublocale. + +## References +https://ncatlab.org/nlab/show/sublocale +https://ncatlab.org/nlab/show/nucleus +-/ + +open Order InfHom + +variable {X : Type*} [CompleteLattice X] + +/-- A nucleus is an inflationary idempotent `inf`-preserving endomorphism of a semilattice. +In a frame, nuclei correspond to sublocales. -/ -- TODO: Formalise that claim +structure Nucleus (X : Type*) [SemilatticeInf X] extends InfHom X X where + /-- A `Nucleus` is idempotent. + + Do not use this directly. Instead use `NucleusClass.idempotent`. -/ + idempotent' (x : X) : toFun (toFun x) ≤ toFun x + /-- A `Nucleus` is increasing. + + Do not use this directly. Instead use `NucleusClass.le_apply`. -/ + le_apply' (x : X) : x ≤ toFun x + +/-- `NucleusClass F X` states that F is a type of nuclei. -/ +class NucleusClass (F X : Type*) [SemilatticeInf X] [FunLike F X X] extends InfHomClass F X X : + Prop where + /-- A nucleus is idempotent. -/ + idempotent (x : X) (f : F) : f (f x) ≤ f x + /-- A nucleus is increasing. -/ + le_apply (x : X) (f : F) : x ≤ f x + +namespace Nucleus + +variable {n m : Nucleus X} {x y : X} + +instance : FunLike (Nucleus X) X X where + coe x := x.toFun + coe_injective' f g h := by obtain ⟨⟨_, _⟩, _⟩ := f; congr! + +lemma toFun_eq_coe (n : Nucleus X) : n.toFun = n := rfl +@[simp] lemma coe_toInfHom (n : Nucleus X) : ⇑n.toInfHom = n := rfl + +instance : NucleusClass (Nucleus X) X where + idempotent _ _ := idempotent' .. + le_apply _ _ := le_apply' .. + map_inf _ _ _ := map_inf' .. + +/-- Every `Nucleus` is a `ClosureOperator`. -/ +def toClosureOperator (n : Nucleus X) : ClosureOperator X := + ClosureOperator.mk' n (OrderHomClass.mono n) n.le_apply' n.idempotent' + +lemma idempotent : n (n x) = n x := + n.toClosureOperator.idempotent x + +lemma le_apply : x ≤ n x := + n.toClosureOperator.le_closure x + +lemma map_inf : n (x ⊓ y) = n x ⊓ n y := + InfHomClass.map_inf n x y + +@[ext] lemma ext {m n : Nucleus X} (h : ∀ a, m a = n a) : m = n := + DFunLike.ext m n h + +/-- A `Nucleus` preserves ⊤. -/ +instance : TopHomClass (Nucleus X) X X where + map_top _ := eq_top_iff.mpr le_apply + +instance : PartialOrder (Nucleus X) := .lift (⇑) DFunLike.coe_injective + +@[simp, norm_cast] lemma coe_le_coe : ⇑m ≤ n ↔ m ≤ n := .rfl +@[simp, norm_cast] lemma coe_lt_coe : ⇑m < n ↔ m < n := .rfl + +/-- The smallest `Nucleus` is the identity. -/ +instance instBot : Bot (Nucleus X) where + bot.toFun x := x + bot.idempotent' := by simp + bot.le_apply' := by simp + bot.map_inf' := by simp + +/-- The biggest `Nucleus` sends everything to `⊤`. -/ +instance instTop : Top (Nucleus X) where + top.toFun := ⊤ + top.idempotent' := by simp + top.le_apply' := by simp + top.map_inf' := by simp + +@[simp, norm_cast] lemma coe_bot : ⇑(⊥ : Nucleus X) = id := rfl +@[simp, norm_cast] lemma coe_top : ⇑(⊤ : Nucleus X) = ⊤ := rfl + +@[simp] lemma bot_apply (x : X) : (⊥ : Nucleus X) x = x := rfl +@[simp] lemma top_apply (x : X) : (⊤ : Nucleus X) x = ⊤ := rfl + +instance : BoundedOrder (Nucleus X) where + bot_le _ _ := le_apply + le_top _ _ := by simp + +end Nucleus diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 9e76dd904111b..c1fe4b32ab082 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -134,7 +134,7 @@ def zip (c₀ : Chain α) (c₁ : Chain β) : Chain (α × β) := /-- An example of a `Chain` constructed from an ordered pair. -/ def pair (a b : α) (hab : a ≤ b) : Chain α where - toFun n := match n with + toFun | 0 => a | _ => b monotone' _ _ _ := by aesop diff --git a/Mathlib/Order/OrderIsoNat.lean b/Mathlib/Order/OrderIsoNat.lean index da0fabb452ad3..891efab009d6f 100644 --- a/Mathlib/Order/OrderIsoNat.lean +++ b/Mathlib/Order/OrderIsoNat.lean @@ -77,7 +77,9 @@ theorem not_acc_of_decreasing_seq (f : ((· > ·) : ℕ → ℕ → Prop) ↪r r rw [acc_iff_no_decreasing_seq, not_isEmpty_iff] exact ⟨⟨f, k, rfl⟩⟩ -/-- A relation is well-founded iff it doesn't have any infinite decreasing sequence. -/ +/-- A strict order relation is well-founded iff it doesn't have any infinite decreasing sequence. + +See `wellFounded_iff_no_descending_seq` for a version which works on any relation. -/ theorem wellFounded_iff_no_descending_seq : WellFounded r ↔ IsEmpty (((· > ·) : ℕ → ℕ → Prop) ↪r r) := by constructor @@ -92,6 +94,14 @@ theorem not_wellFounded_of_decreasing_seq (f : ((· > ·) : ℕ → ℕ → Prop end RelEmbedding +theorem not_strictAnti_of_wellFoundedLT [Preorder α] [WellFoundedLT α] (f : ℕ → α) : + ¬ StrictAnti f := fun hf ↦ + (RelEmbedding.natGT f (fun n ↦ hf (by simp))).not_wellFounded_of_decreasing_seq wellFounded_lt + +theorem not_strictMono_of_wellFoundedGT [Preorder α] [WellFoundedGT α] (f : ℕ → α) : + ¬ StrictMono f := + not_strictAnti_of_wellFoundedLT (α := αᵒᵈ) f + namespace Nat variable (s : Set ℕ) [Infinite s] diff --git a/Mathlib/Order/PrimeIdeal.lean b/Mathlib/Order/PrimeIdeal.lean index e672d2ef9634f..71262162b9097 100644 --- a/Mathlib/Order/PrimeIdeal.lean +++ b/Mathlib/Order/PrimeIdeal.lean @@ -41,8 +41,6 @@ namespace Ideal /-- A pair of an `Order.Ideal` and an `Order.PFilter` which form a partition of `P`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure PrimePair (P : Type*) [Preorder P] where I : Ideal P F : PFilter P diff --git a/Mathlib/Order/Rel/GaloisConnection.lean b/Mathlib/Order/Rel/GaloisConnection.lean index 9666eb999ee6b..a6899c904cc0a 100644 --- a/Mathlib/Order/Rel/GaloisConnection.lean +++ b/Mathlib/Order/Rel/GaloisConnection.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anthony Bordg -/ import Mathlib.Data.Rel -import Mathlib.Order.GaloisConnection /-! # The Galois Connection Induced by a Relation diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 2655560ce6d55..3ca9b0f2679b2 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -137,6 +137,16 @@ theorem trans_trichotomous_right [IsTrans α r] [IsTrichotomous α r] {a b c : theorem transitive_of_trans (r : α → α → Prop) [IsTrans α r] : Transitive r := IsTrans.trans +theorem rel_congr_left [IsSymm α r] [IsTrans α r] {a b c : α} (h : r a b) : r a c ↔ r b c := + ⟨trans_of r (symm_of r h), trans_of r h⟩ + +theorem rel_congr_right [IsSymm α r] [IsTrans α r] {a b c : α} (h : r b c) : r a b ↔ r a c := + ⟨(trans_of r · h), (trans_of r · (symm_of r h))⟩ + +theorem rel_congr [IsSymm α r] [IsTrans α r] {a b c d : α} (h₁ : r a b) (h₂ : r c d) : + r a c ↔ r b d := by + rw [rel_congr_left h₁, rel_congr_right h₂] + /-- In a trichotomous irreflexive order, every element is determined by the set of predecessors. -/ theorem extensional_of_trichotomous_of_irrefl (r : α → α → Prop) [IsTrichotomous α r] [IsIrrefl α r] {a b : α} (H : ∀ x, r x a ↔ r x b) : a = b := @@ -446,7 +456,8 @@ instance (priority := 100) [IsEmpty α] (r : α → α → Prop) : IsWellOrder trans := isEmptyElim wf := wellFounded_of_isEmpty r -instance [IsWellFounded α r] [IsWellFounded β s] : IsWellFounded (α × β) (Prod.Lex r s) := +instance Prod.Lex.instIsWellFounded [IsWellFounded α r] [IsWellFounded β s] : + IsWellFounded (α × β) (Prod.Lex r s) := ⟨IsWellFounded.wf.prod_lex IsWellFounded.wf⟩ instance [IsWellOrder α r] [IsWellOrder β s] : IsWellOrder (α × β) (Prod.Lex r s) where diff --git a/Mathlib/Order/RelIso/Set.lean b/Mathlib/Order/RelIso/Set.lean index 73c00fb2a9a04..90cb227e3edb9 100644 --- a/Mathlib/Order/RelIso/Set.lean +++ b/Mathlib/Order/RelIso/Set.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Order.Directed import Mathlib.Order.RelIso.Basic import Mathlib.Logic.Embedding.Set import Mathlib.Logic.Equiv.Set @@ -35,6 +36,14 @@ theorem map_sup [SemilatticeSup α] [LinearOrder β] [FunLike F β α] a (m ⊔ n) = a m ⊔ a n := map_inf (α := αᵒᵈ) (β := βᵒᵈ) _ _ _ +theorem directed [FunLike F α β] [RelHomClass F r s] {ι : Sort*} {a : ι → α} {f : F} + (ha : Directed r a) : Directed s (f ∘ a) := + ha.mono_comp _ fun _ _ h ↦ map_rel f h + +theorem directedOn [FunLike F α β] [RelHomClass F r s] {f : F} + {t : Set α} (hs : DirectedOn r t) : DirectedOn s (f '' t) := + hs.mono_comp fun _ _ h ↦ map_rel f h + end RelHomClass namespace RelIso @@ -78,9 +87,6 @@ theorem coe_inclusionEmbedding (r : α → α → Prop) {s t : Set α} (h : s (Subrel.inclusionEmbedding r h : s → t) = Set.inclusion h := rfl -instance (r : α → α → Prop) [IsWellOrder α r] (p : α → Prop) : IsWellOrder _ (Subrel r p) := - RelEmbedding.isWellOrder (Subrel.relEmbedding r p) - instance (r : α → α → Prop) [IsRefl α r] (p : α → Prop) : IsRefl _ (Subrel r p) := ⟨fun x => @IsRefl.refl α r _ x⟩ @@ -96,6 +102,16 @@ instance (r : α → α → Prop) [IsTrans α r] (p : α → Prop) : IsTrans _ ( instance (r : α → α → Prop) [IsIrrefl α r] (p : α → Prop) : IsIrrefl _ (Subrel r p) := ⟨fun x => @IsIrrefl.irrefl α r _ x⟩ +instance (r : α → α → Prop) [IsTrichotomous α r] (p : α → Prop) : IsTrichotomous _ (Subrel r p) := + ⟨fun x y => by rw [Subtype.eq_iff]; exact @IsTrichotomous.trichotomous α r _ x y⟩ + +instance (r : α → α → Prop) [IsWellFounded α r] (p : α → Prop) : IsWellFounded _ (Subrel r p) := + (Subrel.relEmbedding r p).isWellFounded + +instance (r : α → α → Prop) [IsPreorder α r] (p : α → Prop) : IsPreorder _ (Subrel r p) where +instance (r : α → α → Prop) [IsStrictOrder α r] (p : α → Prop) : IsStrictOrder _ (Subrel r p) where +instance (r : α → α → Prop) [IsWellOrder α r] (p : α → Prop) : IsWellOrder _ (Subrel r p) where + end Subrel /-- Restrict the codomain of a relation embedding. -/ diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 68a3a6a1f2487..3642adc8ce776 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -5,7 +5,6 @@ Authors: Aaron Anderson, Kevin Buzzard, Yaël Dillies, Eric Wieser -/ import Mathlib.Data.Finset.Sigma import Mathlib.Data.Finset.Pairwise -import Mathlib.Data.Finset.Powerset import Mathlib.Data.Fintype.Basic import Mathlib.Order.CompleteLatticeIntervals @@ -58,13 +57,27 @@ def SupIndep (s : Finset ι) (f : ι → α) : Prop := variable {s t : Finset ι} {f : ι → α} {i : ι} -instance [DecidableEq ι] [DecidableEq α] : Decidable (SupIndep s f) := by - refine @Finset.decidableForallOfDecidableSubsets _ _ _ (?_) - rintro t - - refine @Finset.decidableDforallFinset _ _ _ (?_) - rintro i - - have : Decidable (Disjoint (f i) (sup t f)) := decidable_of_iff' (_ = ⊥) disjoint_iff - infer_instance +/-- The RHS looks like the definition of `iSupIndep`. -/ +theorem supIndep_iff_disjoint_erase [DecidableEq ι] : + s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) := + ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => + (hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩ + +/-- If both the index type and the lattice have decidable equality, +then the `SupIndep` predicate is decidable. + +TODO: speedup the definition and drop the `[DecidableEq ι]` assumption +by iterating over the pairs `(a, t)` such that `s = Finset.cons a t _` +using something like `List.eraseIdx` +or by generating both `f i` and `(s.erase i).sup f` in one loop over `s`. +Yet another possible optimization is to precompute partial suprema of `f` +over the inits and tails of the list representing `s`, +store them in 2 `Array`s, +then compute each `sup` in 1 operation. -/ +instance [DecidableEq ι] [DecidableEq α] : Decidable (SupIndep s f) := + have : ∀ i, Decidable (Disjoint (f i) ((s.erase i).sup f)) := fun _ ↦ + decidable_of_iff _ disjoint_iff.symm + decidable_of_iff _ supIndep_iff_disjoint_erase.symm theorem SupIndep.subset (ht : t.SupIndep f) (h : s ⊆ t) : s.SupIndep f := fun _ hu _ hi => ht (hu.trans h) (h hi) @@ -88,12 +101,6 @@ theorem SupIndep.le_sup_iff (hs : s.SupIndep f) (hts : t ⊆ s) (hi : i ∈ s) ( by_contra hit exact hf i (disjoint_self.1 <| (hs hts hi hit).mono_right h) -/-- The RHS looks like the definition of `iSupIndep`. -/ -theorem supIndep_iff_disjoint_erase [DecidableEq ι] : - s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) := - ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => - (hs i hi).mono_right (sup_mono fun _ hj => mem_erase.2 ⟨ne_of_mem_of_not_mem hj hit, ht hj⟩)⟩ - theorem supIndep_antimono_fun {g : ι → α} (h : ∀ x ∈ s, f x ≤ g x) (h : s.SupIndep g) : s.SupIndep f := by classical diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 5aa9a2eaa166b..4123fe4543d71 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -78,6 +78,26 @@ theorem wellFounded_iff_has_min {r : α → α → Prop} : by_contra hy' exact hm' y hy' hy +/-- A relation is well-founded iff it doesn't have any infinite decreasing sequence. + +See `RelEmbedding.wellFounded_iff_no_descending_seq` for a version on strict orders. -/ +theorem wellFounded_iff_no_descending_seq : + WellFounded r ↔ IsEmpty { f : ℕ → α // ∀ n, r (f (n + 1)) (f n) } := by + rw [WellFounded.wellFounded_iff_has_min] + refine ⟨fun hr ↦ ⟨fun ⟨f, hf⟩ ↦ ?_⟩, ?_⟩ + · obtain ⟨_, ⟨n, rfl⟩, hn⟩ := hr _ (Set.range_nonempty f) + exact hn _ (Set.mem_range_self (n + 1)) (hf n) + · contrapose! + rw [not_isEmpty_iff] + rintro ⟨s, hs, hs'⟩ + let f : ℕ → s := Nat.rec (Classical.indefiniteDescription _ hs) fun n IH ↦ + ⟨(hs' _ IH.2).choose, (hs' _ IH.2).choose_spec.1⟩ + exact ⟨⟨Subtype.val ∘ f, fun n ↦ (hs' _ (f n).2).choose_spec.2⟩⟩ + +theorem not_rel_apply_succ [h : IsWellFounded α r] (f : ℕ → α) : ∃ n, ¬ r (f (n + 1)) (f n) := by + by_contra! hf + exact (wellFounded_iff_no_descending_seq.1 h.wf).elim ⟨f, hf⟩ + open Set /-- The supremum of a bounded, well-founded order -/ @@ -215,45 +235,45 @@ variable (f : α → β) section LT -variable [LT β] (h : WellFounded ((· < ·) : β → β → Prop)) +variable [LT β] [h : WellFoundedLT β] /-- Given a function `f : α → β` where `β` carries a well-founded `<`, this is an element of `α` whose image under `f` is minimal in the sense of `Function.not_lt_argmin`. -/ noncomputable def argmin [Nonempty α] : α := - WellFounded.min (InvImage.wf f h) Set.univ Set.univ_nonempty + WellFounded.min (InvImage.wf f h.wf) Set.univ Set.univ_nonempty -theorem not_lt_argmin [Nonempty α] (a : α) : ¬f a < f (argmin f h) := - WellFounded.not_lt_min (InvImage.wf f h) _ _ (Set.mem_univ a) +theorem not_lt_argmin [Nonempty α] (a : α) : ¬f a < f (argmin f) := + WellFounded.not_lt_min (InvImage.wf f h.wf) _ _ (Set.mem_univ a) /-- Given a function `f : α → β` where `β` carries a well-founded `<`, and a non-empty subset `s` of `α`, this is an element of `s` whose image under `f` is minimal in the sense of `Function.not_lt_argminOn`. -/ noncomputable def argminOn (s : Set α) (hs : s.Nonempty) : α := - WellFounded.min (InvImage.wf f h) s hs + WellFounded.min (InvImage.wf f h.wf) s hs @[simp] -theorem argminOn_mem (s : Set α) (hs : s.Nonempty) : argminOn f h s hs ∈ s := +theorem argminOn_mem (s : Set α) (hs : s.Nonempty) : argminOn f s hs ∈ s := WellFounded.min_mem _ _ _ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): @[simp] removed as it will never apply theorem not_lt_argminOn (s : Set α) {a : α} (ha : a ∈ s) - (hs : s.Nonempty := Set.nonempty_of_mem ha) : ¬f a < f (argminOn f h s hs) := - WellFounded.not_lt_min (InvImage.wf f h) s hs ha + (hs : s.Nonempty := Set.nonempty_of_mem ha) : ¬f a < f (argminOn f s hs) := + WellFounded.not_lt_min (InvImage.wf f h.wf) s hs ha end LT section LinearOrder -variable [LinearOrder β] (h : WellFounded ((· < ·) : β → β → Prop)) +variable [LinearOrder β] [WellFoundedLT β] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): @[simp] removed as it will never apply -theorem argmin_le (a : α) [Nonempty α] : f (argmin f h) ≤ f a := - not_lt.mp <| not_lt_argmin f h a +theorem argmin_le (a : α) [Nonempty α] : f (argmin f) ≤ f a := + not_lt.mp <| not_lt_argmin f a -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): @[simp] removed as it will never apply theorem argminOn_le (s : Set α) {a : α} (ha : a ∈ s) (hs : s.Nonempty := Set.nonempty_of_mem ha) : - f (argminOn f h s hs) ≤ f a := - not_lt.mp <| not_lt_argminOn f h s ha hs + f (argminOn f s hs) ≤ f a := + not_lt.mp <| not_lt_argminOn f s ha hs end LinearOrder diff --git a/Mathlib/Order/WellFoundedSet.lean b/Mathlib/Order/WellFoundedSet.lean index e0c2c0ea13063..99677f208a00d 100644 --- a/Mathlib/Order/WellFoundedSet.lean +++ b/Mathlib/Order/WellFoundedSet.lean @@ -505,9 +505,13 @@ protected theorem IsWF.isPWO (hs : s.IsWF) : s.IsPWO := by simp only [forall_mem_range, not_lt] at hm exact ⟨m, m + 1, by omega, hm _⟩ -/-- In a linear order, the predicates `Set.IsWF` and `Set.IsPWO` are equivalent. -/ +/-- In a linear order, the predicates `Set.IsPWO` and `Set.IsWF` are equivalent. -/ +theorem isPWO_iff_isWF : s.IsPWO ↔ s.IsWF := + ⟨IsPWO.isWF, IsWF.isPWO⟩ + +@[deprecated isPWO_iff_isWF (since := "2025-01-21")] theorem isWF_iff_isPWO : s.IsWF ↔ s.IsPWO := - ⟨IsWF.isPWO, IsPWO.isWF⟩ + isPWO_iff_isWF.symm /-- If `α` is a linear order with well-founded `<`, then any set in it is a partially well-ordered set. @@ -668,6 +672,9 @@ theorem BddBelow.wellFoundedOn_lt : BddBelow s → s.WellFoundedOn (· < ·) := theorem BddAbove.wellFoundedOn_gt : BddAbove s → s.WellFoundedOn (· > ·) := fun h => h.dual.wellFoundedOn_lt +theorem BddBelow.isWF : BddBelow s → IsWF s := + BddBelow.wellFoundedOn_lt + end LocallyFiniteOrder namespace Set.PartiallyWellOrderedOn @@ -789,7 +796,7 @@ theorem subsetProdLex [PartialOrder α] [Preorder β] {s : Set (α ×ₗ β)} constructor · by_contra hx simp_all - · exact (Prod.Lex.le_iff (f (g 0)) _).mpr <| Or.inl hn + · exact Prod.Lex.toLex_le_toLex.mpr <| .inl hn · have hhc : ∀ n, (ofLex f (g 0)).1 = (ofLex f (g n)).1 := by intro n rw [not_exists] at hc @@ -802,9 +809,7 @@ theorem subsetProdLex [PartialOrder α] [Preorder β] {s : Set (α ×ₗ β)} simpa using hf _ use (g (g' 0)), (g (g' 1)) suffices (f (g (g' 0))) ≤ (f (g (g' 1))) by simpa - · refine (Prod.Lex.le_iff (f (g (g' 0))) (f (g (g' 1)))).mpr ?_ - right - constructor + · refine Prod.Lex.toLex_le_toLex.mpr <| .inr ⟨?_, ?_⟩ · exact (hhc (g' 0)).symm.trans (hhc (g' 1)) · exact hg' (Nat.zero_le 1) @@ -823,7 +828,7 @@ theorem fiberProdLex [PartialOrder α] [Preorder β] {s : Set (α ×ₗ β)} rintro b ⟨-, hb⟩ c ⟨-, hc⟩ hbc simp only [mem_preimage, mem_singleton_iff] at hb hc have : (ofLex b).1 < (ofLex c).1 ∨ (ofLex b).1 = (ofLex c).1 ∧ f b ≤ f c := - (Prod.Lex.le_iff b c).mp hbc + Prod.Lex.toLex_le_toLex.mp hbc simp_all only [lt_self_iff_false, true_and, false_or] theorem ProdLex_iff [PartialOrder α] [Preorder β] {s : Set (α ×ₗ β)} : diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 63a1594ad82a0..8cbe3710d0429 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -956,7 +956,7 @@ lemma le_of_forall_lt_iff_le [LinearOrder α] [DenselyOrdered α] [NoMinOrder α | coe x => rw [le_coe_iff] rintro y rfl - exact le_of_forall_le_of_dense (by exact_mod_cast h) + exact le_of_forall_gt_imp_ge_of_dense (by exact_mod_cast h) lemma ge_of_forall_gt_iff_ge [LinearOrder α] [DenselyOrdered α] [NoMinOrder α] {x y : WithBot α} : (∀ z : α, z < x → z ≤ y) ↔ x ≤ y := by @@ -966,7 +966,7 @@ lemma ge_of_forall_gt_iff_ge [LinearOrder α] [DenselyOrdered α] [NoMinOrder α | coe y => rw [le_coe_iff] rintro h x rfl - exact le_of_forall_ge_of_dense (by exact_mod_cast h) + exact le_of_forall_lt_imp_le_of_dense (by exact_mod_cast h) section LE diff --git a/Mathlib/Probability/BorelCantelli.lean b/Mathlib/Probability/BorelCantelli.lean index c94f855abd159..68d93fc17b615 100644 --- a/Mathlib/Probability/BorelCantelli.lean +++ b/Mathlib/Probability/BorelCantelli.lean @@ -43,21 +43,28 @@ theorem iIndepFun.indep_comap_natural_of_lt (hf : ∀ i, StronglyMeasurable (f i (⨆ k ∈ {k | k ≤ i}, MeasurableSpace.comap (f k) mβ) μ by rwa [iSup_singleton] at this exact indep_iSup_of_disjoint (fun k => (hf k).measurable.comap_le) hfi (by simpa) -theorem iIndepFun.condexp_natural_ae_eq_of_lt [SecondCountableTopology β] [CompleteSpace β] +theorem iIndepFun.condExp_natural_ae_eq_of_lt [SecondCountableTopology β] [CompleteSpace β] [NormedSpace ℝ β] (hf : ∀ i, StronglyMeasurable (f i)) (hfi : iIndepFun (fun _ => mβ) f μ) (hij : i < j) : μ[f j|Filtration.natural f hf i] =ᵐ[μ] fun _ => μ[f j] := by have : IsProbabilityMeasure μ := hfi.isProbabilityMeasure - exact condexp_indep_eq (hf j).measurable.comap_le (Filtration.le _ _) + exact condExp_indep_eq (hf j).measurable.comap_le (Filtration.le _ _) (comap_measurable <| f j).stronglyMeasurable (hfi.indep_comap_natural_of_lt hf hij) -theorem iIndepSet.condexp_indicator_filtrationOfSet_ae_eq (hsm : ∀ n, MeasurableSet (s n)) +@[deprecated (since := "2025-01-21")] +alias iIndepFun.condexp_natural_ae_eq_of_lt := iIndepFun.condExp_natural_ae_eq_of_lt + +theorem iIndepSet.condExp_indicator_filtrationOfSet_ae_eq (hsm : ∀ n, MeasurableSet (s n)) (hs : iIndepSet s μ) (hij : i < j) : μ[(s j).indicator (fun _ => 1 : Ω → ℝ)|filtrationOfSet hsm i] =ᵐ[μ] fun _ => (μ (s j)).toReal := by rw [Filtration.filtrationOfSet_eq_natural (β := ℝ) hsm] - refine (iIndepFun.condexp_natural_ae_eq_of_lt _ hs.iIndepFun_indicator hij).trans ?_ + refine (iIndepFun.condExp_natural_ae_eq_of_lt _ hs.iIndepFun_indicator hij).trans ?_ simp only [integral_indicator_const _ (hsm _), Algebra.id.smul_eq_mul, mul_one]; rfl +@[deprecated (since := "2025-01-21")] +alias iIndepSet.condexp_indicator_filtrationOfSet_ae_eq := + iIndepSet.condExp_indicator_filtrationOfSet_ae_eq + open Filter /-- **The second Borel-Cantelli lemma**: Given a sequence of independent sets `(sₙ)` such that @@ -73,7 +80,7 @@ theorem measure_limsup_eq_one {s : ℕ → Set Ω} (hsm : ∀ n, MeasurableSet ( (μ[(s (k + 1)).indicator (1 : Ω → ℝ)|filtrationOfSet hsm k]) ω) atTop atTop} =ᵐ[μ] Set.univ by rw [measure_congr this, measure_univ] have : ∀ᵐ ω ∂μ, ∀ n, (μ[(s (n + 1)).indicator (1 : Ω → ℝ)|filtrationOfSet hsm n]) ω = _ := - ae_all_iff.2 fun n => hs.condexp_indicator_filtrationOfSet_ae_eq hsm n.lt_succ_self + ae_all_iff.2 fun n => hs.condExp_indicator_filtrationOfSet_ae_eq hsm n.lt_succ_self filter_upwards [this] with ω hω refine eq_true (?_ : Tendsto _ _ _) simp_rw [hω] diff --git a/Mathlib/Probability/ConditionalExpectation.lean b/Mathlib/Probability/ConditionalExpectation.lean index 18a5ab72ec167..2cb4e7ad41acd 100644 --- a/Mathlib/Probability/ConditionalExpectation.lean +++ b/Mathlib/Probability/ConditionalExpectation.lean @@ -16,7 +16,7 @@ the main conditional expectation file. ## Main result -* `MeasureTheory.condexp_indep_eq`: If `m₁, m₂` are independent σ-algebras and `f` is an +* `MeasureTheory.condExp_indep_eq`: If `m₁, m₂` are independent σ-algebras and `f` is an `m₁`-measurable function, then `𝔼[f | m₂] = 𝔼[f]` almost everywhere. -/ @@ -35,13 +35,13 @@ variable {Ω E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpac /-- If `m₁, m₂` are independent σ-algebras and `f` is `m₁`-measurable, then `𝔼[f | m₂] = 𝔼[f]` almost everywhere. -/ -theorem condexp_indep_eq (hle₁ : m₁ ≤ m) (hle₂ : m₂ ≤ m) [SigmaFinite (μ.trim hle₂)] +theorem condExp_indep_eq (hle₁ : m₁ ≤ m) (hle₂ : m₂ ≤ m) [SigmaFinite (μ.trim hle₂)] (hf : StronglyMeasurable[m₁] f) (hindp : Indep m₁ m₂ μ) : μ[f|m₂] =ᵐ[μ] fun _ => μ[f] := by by_cases hfint : Integrable f μ - swap; · rw [condexp_undef hfint, integral_undef hfint]; rfl - refine (ae_eq_condexp_of_forall_setIntegral_eq hle₂ hfint + swap; · rw [condExp_of_not_integrable hfint, integral_undef hfint]; rfl + refine (ae_eq_condExp_of_forall_setIntegral_eq hle₂ hfint (fun s _ hs => integrableOn_const.2 (Or.inr hs)) (fun s hms hs => ?_) - stronglyMeasurable_const.aeStronglyMeasurable').symm + stronglyMeasurable_const.aestronglyMeasurable).symm rw [setIntegral_const] rw [← memℒp_one_iff_integrable] at hfint refine Memℒp.induction_stronglyMeasurable hle₁ ENNReal.one_ne_top _ ?_ ?_ ?_ ?_ hfint ?_ @@ -74,4 +74,6 @@ theorem condexp_indep_eq (hle₁ : m₁ ≤ m) (hle₂ : m₂ ≤ m) [SigmaFinit (setIntegral_congr_ae (hle₂ _ hms) _ : ∫ x in s, u x ∂μ = ∫ x in s, v x ∂μ)] filter_upwards [huv] with x hx _ using hx +@[deprecated (since := "2025-01-21")] alias condexp_indep_eq := condExp_indep_eq + end MeasureTheory diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index 4e4d89fc5ac91..2bc58c9893945 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -138,6 +138,12 @@ lemma absolutelyContinuous_cond_univ [IsFiniteMeasure μ] : μ ≪ μ[|univ] := refine absolutelyContinuous_smul ?_ simp [measure_ne_top] +lemma ae_cond_mem₀ (hs : NullMeasurableSet s μ) : ∀ᵐ x ∂μ[|s], x ∈ s := + ae_smul_measure (ae_restrict_mem₀ hs) _ + +lemma ae_cond_mem (hs : MeasurableSet s) : ∀ᵐ x ∂μ[|s], x ∈ s := + ae_smul_measure (ae_restrict_mem hs) _ + section Bayes variable (μ) in diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index 8100dd885dedf..a89d87c31a792 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -50,7 +50,7 @@ which we currently do not have. -/ -open scoped Classical MeasureTheory NNReal ENNReal +open scoped MeasureTheory NNReal ENNReal open TopologicalSpace MeasureTheory.Measure diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 41df4ffe5d082..8b3c99c7d763c 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Lorenzo Luccioli, Rémy Degenne, Alexander Bentkamp -/ import Mathlib.Analysis.SpecialFunctions.Gaussian.GaussianIntegral -import Mathlib.Probability.Moments +import Mathlib.Probability.Moments.Basic /-! # Gaussian distributions over ℝ @@ -71,6 +71,7 @@ lemma stronglyMeasurable_gaussianPDFReal (μ : ℝ) (v : ℝ≥0) : StronglyMeasurable (gaussianPDFReal μ v) := (measurable_gaussianPDFReal μ v).stronglyMeasurable +@[fun_prop] lemma integrable_gaussianPDFReal (μ : ℝ) (v : ℝ≥0) : Integrable (gaussianPDFReal μ v) := by rw [gaussianPDFReal_def] diff --git a/Mathlib/Probability/Distributions/Uniform.lean b/Mathlib/Probability/Distributions/Uniform.lean index 77f3a13bb4a8c..3efd4efda464d 100644 --- a/Mathlib/Probability/Distributions/Uniform.lean +++ b/Mathlib/Probability/Distributions/Uniform.lean @@ -39,7 +39,7 @@ This file defines a number of uniform `PMF` distributions from various inputs, * Refactor the `PMF` definitions to come from a `uniformMeasure` on a `Finset`/`Fintype`/`Multiset`. -/ -open scoped Classical MeasureTheory NNReal ENNReal +open scoped MeasureTheory NNReal ENNReal -- TODO: We can't `open ProbabilityTheory` without opening the `ProbabilityTheory` locale :( open TopologicalSpace MeasureTheory.Measure PMF @@ -193,6 +193,7 @@ lemma uniformPDF_eq_pdf {s : Set E} (hs : MeasurableSet s) (hu : pdf.IsUniform X unfold uniformPDF exact Filter.EventuallyEq.trans (pdf.IsUniform.pdf_eq hs hu).symm (ae_eq_refl _) +open scoped Classical in /-- Alternative way of writing the uniformPDF. -/ lemma uniformPDF_ite {s : Set E} {x : E} : uniformPDF s x μ = if x ∈ s then (μ s)⁻¹ else 0 := by diff --git a/Mathlib/Probability/Independence/Conditional.lean b/Mathlib/Probability/Independence/Conditional.lean index 6456d37f340cf..9fa9a790247a1 100644 --- a/Mathlib/Probability/Independence/Conditional.lean +++ b/Mathlib/Probability/Independence/Conditional.lean @@ -16,8 +16,8 @@ Two σ-algebras `m₁` and `m₂` are conditionally independent given a third σ `μ⟦t₁ ∩ t₂ | m'⟧ =ᵐ[μ] μ⟦t₁ | m'⟧ * μ⟦t₂ | m'⟧`. On standard Borel spaces, the conditional expectation with respect to `m'` defines a kernel -`ProbabilityTheory.condexpKernel`, and the definition above is equivalent to -`∀ᵐ ω ∂μ, condexpKernel μ m' ω (t₁ ∩ t₂) = condexpKernel μ m' ω t₁ * condexpKernel μ m' ω t₂`. +`ProbabilityTheory.condExpKernel`, and the definition above is equivalent to +`∀ᵐ ω ∂μ, condExpKernel μ m' ω (t₁ ∩ t₂) = condExpKernel μ m' ω t₁ * condExpKernel μ m' ω t₂`. We use this property as the definition of conditional independence. ## Main definitions @@ -46,7 +46,7 @@ as for a family, but without the starting `i`, for example `CondIndepFun` is the The definitions of conditional independence in this file are a particular case of independence with respect to a kernel and a measure, as defined in the file `Probability/Independence/Kernel.lean`. -The kernel used is `ProbabilityTheory.condexpKernel`. +The kernel used is `ProbabilityTheory.condExpKernel`. -/ @@ -71,14 +71,14 @@ See `ProbabilityTheory.iCondIndepSets_iff`. It will be used for families of pi_systems. -/ def iCondIndepSets (π : ι → Set (Set Ω)) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.iIndepSets π (condexpKernel μ m') (μ.trim hm') + Kernel.iIndepSets π (condExpKernel μ m') (μ.trim hm') /-- Two sets of sets `s₁, s₂` are conditionally independent given `m'` with respect to a measure `μ` if for any sets `t₁ ∈ s₁, t₂ ∈ s₂`, then `μ⟦t₁ ∩ t₂ | m'⟧ =ᵐ[μ] μ⟦t₁ | m'⟧ * μ⟦t₂ | m'⟧`. See `ProbabilityTheory.condIndepSets_iff`. -/ def CondIndepSets (s1 s2 : Set (Set Ω)) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.IndepSets s1 s2 (condexpKernel μ m') (μ.trim hm') + Kernel.IndepSets s1 s2 (condExpKernel μ m') (μ.trim hm') /-- A family of measurable space structures (i.e. of σ-algebras) is conditionally independent given `m'` with respect to a measure `μ` (typically defined on a finer σ-algebra) if the family of sets of @@ -89,7 +89,7 @@ any sets `f i_1 ∈ m i_1, ..., f i_n ∈ m i_n`, then See `ProbabilityTheory.iCondIndep_iff`. -/ def iCondIndep (m : ι → MeasurableSpace Ω) (μ : @Measure Ω mΩ := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.iIndep m (condexpKernel (mΩ := mΩ) μ m') (μ.trim hm') + Kernel.iIndep m (condExpKernel (mΩ := mΩ) μ m') (μ.trim hm') end @@ -100,7 +100,7 @@ See `ProbabilityTheory.condIndep_iff`. -/ def CondIndep (m' m₁ m₂ : MeasurableSpace Ω) {mΩ : MeasurableSpace Ω} [StandardBorelSpace Ω] (hm' : m' ≤ mΩ) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.Indep m₁ m₂ (condexpKernel μ m') (μ.trim hm') + Kernel.Indep m₁ m₂ (condExpKernel μ m') (μ.trim hm') section @@ -112,14 +112,14 @@ generate is conditionally independent. For a set `s`, the generated measurable s sets `∅, s, sᶜ, univ`. See `ProbabilityTheory.iCondIndepSet_iff`. -/ def iCondIndepSet (s : ι → Set Ω) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.iIndepSet s (condexpKernel μ m') (μ.trim hm') + Kernel.iIndepSet s (condExpKernel μ m') (μ.trim hm') /-- Two sets are conditionally independent if the two measurable space structures they generate are conditionally independent. For a set `s`, the generated measurable space structure has measurable sets `∅, s, sᶜ, univ`. See `ProbabilityTheory.condIndepSet_iff`. -/ def CondIndepSet (s t : Set Ω) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.IndepSet s t (condexpKernel μ m') (μ.trim hm') + Kernel.IndepSet s t (condExpKernel μ m') (μ.trim hm') /-- A family of functions defined on the same space `Ω` and taking values in possibly different spaces, each with a measurable space structure, is conditionally independent if the family of @@ -129,7 +129,7 @@ with codomain having measurable space structure `m`, the generated measurable sp See `ProbabilityTheory.iCondIndepFun_iff`. -/ def iCondIndepFun {β : ι → Type*} (m : ∀ x : ι, MeasurableSpace (β x)) (f : ∀ x : ι, Ω → β x) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.iIndepFun m f (condexpKernel μ m') (μ.trim hm') + Kernel.iIndepFun m f (condExpKernel μ m') (μ.trim hm') /-- Two functions are conditionally independent if the two measurable space structures they generate are conditionally independent. For a function `f` with codomain having measurable space structure @@ -137,7 +137,7 @@ are conditionally independent. For a function `f` with codomain having measurabl See `ProbabilityTheory.condIndepFun_iff`. -/ def CondIndepFun {β γ : Type*} [MeasurableSpace β] [MeasurableSpace γ] (f : Ω → β) (g : Ω → γ) (μ : Measure Ω := by volume_tac) [IsFiniteMeasure μ] : Prop := - Kernel.IndepFun f g (condexpKernel μ m') (μ.trim hm') + Kernel.IndepFun f g (condExpKernel μ m') (μ.trim hm') end @@ -154,34 +154,33 @@ lemma iCondIndepSets_iff (π : ι → Set (Set Ω)) (hπ : ∀ i s (_hs : s ∈ μ⟦⋂ i ∈ s, f i | m'⟧ =ᵐ[μ] ∏ i ∈ s, (μ⟦f i | m'⟧) := by simp only [iCondIndepSets, Kernel.iIndepSets] have h_eq' : ∀ (s : Finset ι) (f : ι → Set Ω) (_H : ∀ i, i ∈ s → f i ∈ π i) i (_hi : i ∈ s), - (fun ω ↦ ENNReal.toReal (condexpKernel μ m' ω (f i))) =ᵐ[μ] μ⟦f i | m'⟧ := - fun s f H i hi ↦ condexpKernel_ae_eq_condexp hm' (hπ i (f i) (H i hi)) + (fun ω ↦ ENNReal.toReal (condExpKernel μ m' ω (f i))) =ᵐ[μ] μ⟦f i | m'⟧ := + fun s f H i hi ↦ condExpKernel_ae_eq_condExp hm' (hπ i (f i) (H i hi)) have h_eq : ∀ (s : Finset ι) (f : ι → Set Ω) (_H : ∀ i, i ∈ s → f i ∈ π i), ∀ᵐ ω ∂μ, - ∀ i ∈ s, ENNReal.toReal (condexpKernel μ m' ω (f i)) = (μ⟦f i | m'⟧) ω := by + ∀ i ∈ s, ENNReal.toReal (condExpKernel μ m' ω (f i)) = (μ⟦f i | m'⟧) ω := by intros s f H simp_rw [← Finset.mem_coe] rw [ae_ball_iff (Finset.countable_toSet s)] exact h_eq' s f H have h_inter_eq : ∀ (s : Finset ι) (f : ι → Set Ω) (_H : ∀ i, i ∈ s → f i ∈ π i), - (fun ω ↦ ENNReal.toReal (condexpKernel μ m' ω (⋂ i ∈ s, f i))) + (fun ω ↦ ENNReal.toReal (condExpKernel μ m' ω (⋂ i ∈ s, f i))) =ᵐ[μ] μ⟦⋂ i ∈ s, f i | m'⟧ := by - refine fun s f H ↦ condexpKernel_ae_eq_condexp hm' ?_ + refine fun s f H ↦ condExpKernel_ae_eq_condExp hm' ?_ exact MeasurableSet.biInter (Finset.countable_toSet _) (fun i hi ↦ hπ i _ (H i hi)) refine ⟨fun h s f hf ↦ ?_, fun h s f hf ↦ ?_⟩ <;> specialize h s hf · have h' := ae_eq_of_ae_eq_trim h filter_upwards [h_eq s f hf, h_inter_eq s f hf, h'] with ω h_eq h_inter_eq h' rw [← h_inter_eq, h', ENNReal.toReal_prod, Finset.prod_apply] exact Finset.prod_congr rfl h_eq - · refine (ae_eq_trim_iff hm' ?_ ?_).mpr ?_ - · refine stronglyMeasurable_condexpKernel ?_ - exact MeasurableSet.biInter (Finset.countable_toSet _) (fun i hi ↦ hπ i _ (hf i hi)) + · refine ((stronglyMeasurable_condExpKernel ?_).ae_eq_trim_iff hm' ?_).mpr ?_ + · exact .biInter (Finset.countable_toSet _) (fun i hi ↦ hπ i _ (hf i hi)) · refine Measurable.stronglyMeasurable ?_ - exact Finset.measurable_prod s (fun i hi ↦ measurable_condexpKernel (hπ i _ (hf i hi))) + exact Finset.measurable_prod s (fun i hi ↦ measurable_condExpKernel (hπ i _ (hf i hi))) filter_upwards [h_eq s f hf, h_inter_eq s f hf, h] with ω h_eq h_inter_eq h - have h_ne_top : condexpKernel μ m' ω (⋂ i ∈ s, f i) ≠ ∞ := - (measure_ne_top (condexpKernel μ m' ω) _) - have : (∏ i ∈ s, condexpKernel μ m' ω (f i)) ≠ ∞ := - ENNReal.prod_ne_top fun _ _ ↦ measure_ne_top (condexpKernel μ m' ω) _ + have h_ne_top : condExpKernel μ m' ω (⋂ i ∈ s, f i) ≠ ∞ := + (measure_ne_top (condExpKernel μ m' ω) _) + have : (∏ i ∈ s, condExpKernel μ m' ω (f i)) ≠ ∞ := + ENNReal.prod_ne_top fun _ _ ↦ measure_ne_top (condExpKernel μ m' ω) _ rw [← ENNReal.ofReal_toReal h_ne_top, h_inter_eq, h, Finset.prod_apply, ← ENNReal.ofReal_toReal this, ENNReal.toReal_prod] congr 1 @@ -192,28 +191,26 @@ lemma condIndepSets_iff (s1 s2 : Set (Set Ω)) (hs1 : ∀ s ∈ s1, MeasurableSe CondIndepSets m' hm' s1 s2 μ ↔ ∀ (t1 t2 : Set Ω) (_ : t1 ∈ s1) (_ : t2 ∈ s2), (μ⟦t1 ∩ t2 | m'⟧) =ᵐ[μ] (μ⟦t1 | m'⟧) * (μ⟦t2 | m'⟧) := by simp only [CondIndepSets, Kernel.IndepSets] - have hs1_eq : ∀ s ∈ s1, (fun ω ↦ ENNReal.toReal (condexpKernel μ m' ω s)) =ᵐ[μ] μ⟦s | m'⟧ := - fun s hs ↦ condexpKernel_ae_eq_condexp hm' (hs1 s hs) - have hs2_eq : ∀ s ∈ s2, (fun ω ↦ ENNReal.toReal (condexpKernel μ m' ω s)) =ᵐ[μ] μ⟦s | m'⟧ := - fun s hs ↦ condexpKernel_ae_eq_condexp hm' (hs2 s hs) - have hs12_eq : ∀ s ∈ s1, ∀ t ∈ s2, (fun ω ↦ ENNReal.toReal (condexpKernel μ m' ω (s ∩ t))) + have hs1_eq : ∀ s ∈ s1, (fun ω ↦ ENNReal.toReal (condExpKernel μ m' ω s)) =ᵐ[μ] μ⟦s | m'⟧ := + fun s hs ↦ condExpKernel_ae_eq_condExp hm' (hs1 s hs) + have hs2_eq : ∀ s ∈ s2, (fun ω ↦ ENNReal.toReal (condExpKernel μ m' ω s)) =ᵐ[μ] μ⟦s | m'⟧ := + fun s hs ↦ condExpKernel_ae_eq_condExp hm' (hs2 s hs) + have hs12_eq : ∀ s ∈ s1, ∀ t ∈ s2, (fun ω ↦ ENNReal.toReal (condExpKernel μ m' ω (s ∩ t))) =ᵐ[μ] μ⟦s ∩ t | m'⟧ := - fun s hs t ht ↦ condexpKernel_ae_eq_condexp hm' ((hs1 s hs).inter ((hs2 t ht))) + fun s hs t ht ↦ condExpKernel_ae_eq_condExp hm' ((hs1 s hs).inter ((hs2 t ht))) refine ⟨fun h s t hs ht ↦ ?_, fun h s t hs ht ↦ ?_⟩ <;> specialize h s t hs ht · have h' := ae_eq_of_ae_eq_trim h filter_upwards [hs1_eq s hs, hs2_eq t ht, hs12_eq s hs t ht, h'] with ω hs_eq ht_eq hst_eq h' rw [← hst_eq, Pi.mul_apply, ← hs_eq, ← ht_eq, h', ENNReal.toReal_mul] - · refine (ae_eq_trim_iff hm' ?_ ?_).mpr ?_ - · exact stronglyMeasurable_condexpKernel ((hs1 s hs).inter ((hs2 t ht))) - · refine Measurable.stronglyMeasurable (Measurable.mul ?_ ?_) - · exact measurable_condexpKernel (hs1 s hs) - · exact measurable_condexpKernel (hs2 t ht) + · refine ((stronglyMeasurable_condExpKernel ((hs1 s hs).inter (hs2 t ht))).ae_eq_trim_iff hm' + ((measurable_condExpKernel (hs1 s hs)).mul + (measurable_condExpKernel (hs2 t ht))).stronglyMeasurable).mpr ?_ filter_upwards [hs1_eq s hs, hs2_eq t ht, hs12_eq s hs t ht, h] with ω hs_eq ht_eq hst_eq h - have h_ne_top : condexpKernel μ m' ω (s ∩ t) ≠ ∞ := measure_ne_top (condexpKernel μ m' ω) _ + have h_ne_top : condExpKernel μ m' ω (s ∩ t) ≠ ∞ := measure_ne_top (condExpKernel μ m' ω) _ rw [← ENNReal.ofReal_toReal h_ne_top, hst_eq, h, Pi.mul_apply, ← hs_eq, ← ht_eq, ← ENNReal.toReal_mul, ENNReal.ofReal_toReal] - exact ENNReal.mul_ne_top (measure_ne_top (condexpKernel μ m' ω) s) - (measure_ne_top (condexpKernel μ m' ω) t) + exact ENNReal.mul_ne_top (measure_ne_top (condExpKernel μ m' ω) s) + (measure_ne_top (condExpKernel μ m' ω) t) lemma iCondIndepSets_singleton_iff (s : ι → Set Ω) (hπ : ∀ i, MeasurableSet (s i)) (μ : Measure Ω) [IsFiniteMeasure μ] : @@ -639,7 +636,7 @@ variable {β β' : Type*} {m' : MeasurableSpace Ω} {hm' : m' ≤ mΩ} {μ : Measure Ω} [IsFiniteMeasure μ] {f : Ω → β} {g : Ω → β'} -theorem condIndepFun_iff_condexp_inter_preimage_eq_mul {mβ : MeasurableSpace β} +theorem condIndepFun_iff_condExp_inter_preimage_eq_mul {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} (hf : Measurable f) (hg : Measurable g) : CondIndepFun m' hm' f g μ ↔ ∀ s t, MeasurableSet s → MeasurableSet t @@ -650,7 +647,11 @@ theorem condIndepFun_iff_condexp_inter_preimage_eq_mul {mβ : MeasurableSpace β · rintro ⟨s, hs, rfl⟩ ⟨t, ht, rfl⟩ exact h s t hs ht -theorem iCondIndepFun_iff_condexp_inter_preimage_eq_mul {β : ι → Type*} +@[deprecated (since := "2025-01-21")] +alias condIndepFun_iff_condexp_inter_preimage_eq_mul := + condIndepFun_iff_condExp_inter_preimage_eq_mul + +theorem iCondIndepFun_iff_condExp_inter_preimage_eq_mul {β : ι → Type*} (m : ∀ x, MeasurableSpace (β x)) (f : ∀ i, Ω → β i) (hf : ∀ i, Measurable (f i)) : iCondIndepFun m' hm' m f μ ↔ ∀ (S : Finset ι) {sets : ∀ i : ι, Set (β i)} (_H : ∀ i, i ∈ S → MeasurableSet[m i] (sets i)), @@ -672,6 +673,10 @@ theorem iCondIndepFun_iff_condexp_inter_preimage_eq_mul {β : ι → Type*} simp only [g, dif_pos hi] convert h with i hi i hi <;> exact hg i hi +@[deprecated (since := "2025-01-21")] +alias iCondIndepFun_iff_condexp_inter_preimage_eq_mul := + iCondIndepFun_iff_condExp_inter_preimage_eq_mul + theorem condIndepFun_iff_condIndepSet_preimage {mβ : MeasurableSpace β} {mβ' : MeasurableSpace β'} (hf : Measurable f) (hg : Measurable g) : CondIndepFun m' hm' f g μ ↔ diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 1b8fc5fb8bb20..9984c68a77708 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -17,7 +17,7 @@ for any sets `f i_1 ∈ π i_1, ..., f i_n ∈ π i_n`, then for `μ`-almost eve `κ a (⋂ i in s, f i) = ∏ i ∈ s, κ a (f i)`. This notion of independence is a generalization of both independence and conditional independence. -For conditional independence, `κ` is the conditional kernel `ProbabilityTheory.condexpKernel` and +For conditional independence, `κ` is the conditional kernel `ProbabilityTheory.condExpKernel` and `μ` is the ambient measure. For (non-conditional) independence, `κ = Kernel.const Unit μ` and the measure is the Dirac measure on `Unit`. diff --git a/Mathlib/Probability/Independence/ZeroOne.lean b/Mathlib/Probability/Independence/ZeroOne.lean index 06403152bcee2..4d8d758321b71 100644 --- a/Mathlib/Probability/Independence/ZeroOne.lean +++ b/Mathlib/Probability/Independence/ZeroOne.lean @@ -38,7 +38,7 @@ theorem Kernel.measure_eq_zero_or_one_or_top_of_indepSet_self {t : Set Ω} · exact Or.inl h0 by_cases h_top : κ a t = ∞ · exact Or.inr (Or.inr h_top) - rw [← one_mul (κ a (t ∩ t)), Set.inter_self, ENNReal.mul_eq_mul_right h0 h_top] at ha + rw [← one_mul (κ a (t ∩ t)), Set.inter_self, ENNReal.mul_left_inj h0 h_top] at ha exact Or.inr (Or.inl ha.symm) theorem measure_eq_zero_or_one_or_top_of_indepSet_self {t : Set Ω} @@ -62,20 +62,23 @@ theorem measure_eq_zero_or_one_of_indepSet_self [IsFiniteMeasure μ] {t : Set Ω simpa only [ae_dirac_eq, Filter.eventually_pure] using Kernel.measure_eq_zero_or_one_of_indepSet_self h_indep -theorem condexp_eq_zero_or_one_of_condIndepSet_self +theorem condExp_eq_zero_or_one_of_condIndepSet_self [StandardBorelSpace Ω] (hm : m ≤ m0) [hμ : IsFiniteMeasure μ] {t : Set Ω} (ht : MeasurableSet t) (h_indep : CondIndepSet m hm t t μ) : ∀ᵐ ω ∂μ, (μ⟦t | m⟧) ω = 0 ∨ (μ⟦t | m⟧) ω = 1 := by -- TODO: Why is not inferred? - have (a) : IsFiniteMeasure (condexpKernel μ m a) := inferInstance + have (a) : IsFiniteMeasure (condExpKernel μ m a) := inferInstance have h := ae_of_ae_trim hm (Kernel.measure_eq_zero_or_one_of_indepSet_self h_indep) - filter_upwards [condexpKernel_ae_eq_condexp hm ht, h] with ω hω_eq hω + filter_upwards [condExpKernel_ae_eq_condExp hm ht, h] with ω hω_eq hω rw [← hω_eq, ENNReal.toReal_eq_zero_iff, ENNReal.toReal_eq_one_iff] cases hω with | inl h => exact Or.inl (Or.inl h) | inr h => exact Or.inr h +@[deprecated (since := "2025-01-21")] +alias condexp_eq_zero_or_one_of_condIndepSet_self := condExp_eq_zero_or_one_of_condIndepSet_self + open Filter theorem Kernel.indep_biSup_compl (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s κ μα) (t : Set ι) : @@ -223,7 +226,7 @@ theorem measure_zero_or_one_of_measurableSet_limsup using Kernel.measure_zero_or_one_of_measurableSet_limsup h_le h_indep hf hns hnsp hns_univ ht_tail -theorem condexp_zero_or_one_of_measurableSet_limsup [StandardBorelSpace Ω] +theorem condExp_zero_or_one_of_measurableSet_limsup [StandardBorelSpace Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) @@ -232,12 +235,15 @@ theorem condexp_zero_or_one_of_measurableSet_limsup [StandardBorelSpace Ω] have h := ae_of_ae_trim hm (Kernel.measure_zero_or_one_of_measurableSet_limsup h_le h_indep hf hns hnsp hns_univ ht_tail) have ht : MeasurableSet t := limsup_le_iSup.trans (iSup_le h_le) t ht_tail - filter_upwards [condexpKernel_ae_eq_condexp hm ht, h] with ω hω_eq hω + filter_upwards [condExpKernel_ae_eq_condExp hm ht, h] with ω hω_eq hω rw [← hω_eq, ENNReal.toReal_eq_zero_iff, ENNReal.toReal_eq_one_iff] cases hω with | inl h => exact Or.inl (Or.inl h) | inr h => exact Or.inr h +@[deprecated (since := "2025-01-21")] +alias condexp_zero_or_one_of_measurableSet_limsup := condExp_zero_or_one_of_measurableSet_limsup + end Abstract section AtTop @@ -285,13 +291,17 @@ theorem measure_zero_or_one_of_measurableSet_limsup_atTop simpa only [ae_dirac_eq, Filter.eventually_pure] using Kernel.measure_zero_or_one_of_measurableSet_limsup_atTop h_le h_indep ht_tail -theorem condexp_zero_or_one_of_measurableSet_limsup_atTop [StandardBorelSpace Ω] +theorem condExp_zero_or_one_of_measurableSet_limsup_atTop [StandardBorelSpace Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) {t : Set Ω} (ht_tail : MeasurableSet[limsup s atTop] t) : ∀ᵐ ω ∂μ, (μ⟦t | m⟧) ω = 0 ∨ (μ⟦t | m⟧) ω = 1 := - condexp_eq_zero_or_one_of_condIndepSet_self hm (limsup_le_iSup.trans (iSup_le h_le) t ht_tail) + condExp_eq_zero_or_one_of_condIndepSet_self hm (limsup_le_iSup.trans (iSup_le h_le) t ht_tail) ((condIndep_limsup_atTop_self hm h_le h_indep).condIndepSet_of_measurableSet ht_tail ht_tail) +@[deprecated (since := "2025-01-21")] +alias condexp_zero_or_one_of_measurableSet_limsup_atTop := + condExp_zero_or_one_of_measurableSet_limsup_atTop + end AtTop section AtBot @@ -343,13 +353,17 @@ theorem measure_zero_or_one_of_measurableSet_limsup_atBot /-- **Kolmogorov's 0-1 law**, conditional version: any event in the tail σ-algebra of a conditionally independent sequence of sub-σ-algebras has conditional probability 0 or 1. -/ -theorem condexp_zero_or_one_of_measurableSet_limsup_atBot [StandardBorelSpace Ω] +theorem condExp_zero_or_one_of_measurableSet_limsup_atBot [StandardBorelSpace Ω] (hm : m ≤ m0) [IsFiniteMeasure μ] (h_le : ∀ n, s n ≤ m0) (h_indep : iCondIndep m hm s μ) {t : Set Ω} (ht_tail : MeasurableSet[limsup s atBot] t) : ∀ᵐ ω ∂μ, (μ⟦t | m⟧) ω = 0 ∨ (μ⟦t | m⟧) ω = 1 := - condexp_eq_zero_or_one_of_condIndepSet_self hm (limsup_le_iSup.trans (iSup_le h_le) t ht_tail) + condExp_eq_zero_or_one_of_condIndepSet_self hm (limsup_le_iSup.trans (iSup_le h_le) t ht_tail) ((condIndep_limsup_atBot_self hm h_le h_indep).condIndepSet_of_measurableSet ht_tail ht_tail) +@[deprecated (since := "2025-01-21")] +alias condexp_zero_or_one_of_measurableSet_limsup_atBot := + condExp_zero_or_one_of_measurableSet_limsup_atBot + end AtBot end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index 56e7c43530b0e..8868754b1be92 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -32,9 +32,9 @@ to `m`. ## Main statements -* `condDistrib_ae_eq_condexp`: for almost all `a`, `condDistrib` evaluated at `X a` and a +* `condDistrib_ae_eq_condExp`: for almost all `a`, `condDistrib` evaluated at `X a` and a measurable set `s` is equal to the conditional expectation `μ⟦Y ⁻¹' s | mβ.comap X⟧ a`. -* `condexp_prod_ae_eq_integral_condDistrib`: the conditional expectation +* `condExp_prod_ae_eq_integral_condDistrib`: the conditional expectation `μ[(fun a => f (X a, Y a)) | X; mβ]` is almost everywhere equal to the integral `∫ y, f (X a, y) ∂(condDistrib Y X μ (X a))`. @@ -100,11 +100,14 @@ theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condDistrib (hX : AEM AEStronglyMeasurable (fun a => ∫ y, f (X a, y) ∂condDistrib Y X μ (X a)) μ := (hf.integral_condDistrib_map hY).comp_aemeasurable hX -theorem aestronglyMeasurable'_integral_condDistrib (hX : AEMeasurable X μ) (hY : AEMeasurable Y μ) +theorem aestronglyMeasurable_integral_condDistrib (hX : AEMeasurable X μ) (hY : AEMeasurable Y μ) (hf : AEStronglyMeasurable f (μ.map fun a => (X a, Y a))) : - AEStronglyMeasurable' (mβ.comap X) (fun a => ∫ y, f (X a, y) ∂condDistrib Y X μ (X a)) μ := + AEStronglyMeasurable[mβ.comap X] (fun a => ∫ y, f (X a, y) ∂condDistrib Y X μ (X a)) μ := (hf.integral_condDistrib_map hY).comp_ae_measurable' hX +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_integral_condDistrib := aestronglyMeasurable_integral_condDistrib + end Measurability /-- `condDistrib` is a.e. uniquely defined as the kernel satisfying the defining property of @@ -202,9 +205,9 @@ alias set_lintegral_condDistrib_of_measurableSet := setLIntegral_condDistrib_of_ /-- For almost every `a : α`, the `condDistrib Y X μ` kernel applied to `X a` and a measurable set `s` is equal to the conditional expectation of the indicator of `Y ⁻¹' s`. -/ -theorem condDistrib_ae_eq_condexp (hX : Measurable X) (hY : Measurable Y) (hs : MeasurableSet s) : +theorem condDistrib_ae_eq_condExp (hX : Measurable X) (hY : Measurable Y) (hs : MeasurableSet s) : (fun a => (condDistrib Y X μ (X a) s).toReal) =ᵐ[μ] μ⟦Y ⁻¹' s|mβ.comap X⟧ := by - refine ae_eq_condexp_of_forall_setIntegral_eq hX.comap_le ?_ ?_ ?_ ?_ + refine ae_eq_condExp_of_forall_setIntegral_eq hX.comap_le ?_ ?_ ?_ ?_ · exact (integrable_const _).indicator (hY hs) · exact fun t _ _ => (integrable_toReal_condDistrib hX.aemeasurable hs).integrableOn · intro t ht _ @@ -212,18 +215,19 @@ theorem condDistrib_ae_eq_condexp (hX : Measurable X) (hY : Measurable Y) (hs : (Eventually.of_forall fun ω => measure_lt_top (condDistrib Y X μ (X ω)) _), integral_indicator_const _ (hY hs), Measure.restrict_apply (hY hs), smul_eq_mul, mul_one, inter_comm, setLIntegral_condDistrib_of_measurableSet hX hY.aemeasurable hs ht] - · refine (Measurable.stronglyMeasurable ?_).aeStronglyMeasurable' - exact @Measurable.ennreal_toReal _ (mβ.comap X) _ (measurable_condDistrib hs) + · exact (measurable_condDistrib hs).ennreal_toReal.aestronglyMeasurable + +@[deprecated (since := "2025-01-21")] alias condDistrib_ae_eq_condexp := condDistrib_ae_eq_condExp /-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal to the integral of `y ↦ f(X, y)` against the `condDistrib` kernel. -/ -theorem condexp_prod_ae_eq_integral_condDistrib' [NormedSpace ℝ F] [CompleteSpace F] +theorem condExp_prod_ae_eq_integral_condDistrib' [NormedSpace ℝ F] [CompleteSpace F] (hX : Measurable X) (hY : AEMeasurable Y μ) (hf_int : Integrable f (μ.map fun a => (X a, Y a))) : μ[fun a => f (X a, Y a)|mβ.comap X] =ᵐ[μ] fun a => ∫ y, f (X a,y) ∂condDistrib Y X μ (X a) := by have hf_int' : Integrable (fun a => f (X a, Y a)) μ := (integrable_map_measure hf_int.1 (hX.aemeasurable.prod_mk hY)).mp hf_int - refine (ae_eq_condexp_of_forall_setIntegral_eq hX.comap_le hf_int' (fun s _ _ => ?_) ?_ ?_).symm + refine (ae_eq_condExp_of_forall_setIntegral_eq hX.comap_le hf_int' (fun s _ _ => ?_) ?_ ?_).symm · exact (hf_int.integral_condDistrib hX.aemeasurable hY).integrableOn · rintro s ⟨t, ht, rfl⟩ _ change ∫ a in X ⁻¹' t, ((fun x' => ∫ y, f (x', y) ∂(condDistrib Y X μ) x') ∘ X) a ∂μ = @@ -237,42 +241,57 @@ theorem condexp_prod_ae_eq_integral_condDistrib' [NormedSpace ℝ F] [CompleteSp Measure.setIntegral_condKernel_univ_right ht hf_int.integrableOn, setIntegral_map (ht.prod MeasurableSet.univ) hf_int.1 (hX.aemeasurable.prod_mk hY), mk_preimage_prod, preimage_univ, inter_univ] - · exact aestronglyMeasurable'_integral_condDistrib hX.aemeasurable hY hf_int.1 + · exact aestronglyMeasurable_integral_condDistrib hX.aemeasurable hY hf_int.1 + +@[deprecated (since := "2025-01-21")] +alias condexp_prod_ae_eq_integral_condDistrib' := condExp_prod_ae_eq_integral_condDistrib' /-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal to the integral of `y ↦ f(X, y)` against the `condDistrib` kernel. -/ -theorem condexp_prod_ae_eq_integral_condDistrib₀ [NormedSpace ℝ F] [CompleteSpace F] +theorem condExp_prod_ae_eq_integral_condDistrib₀ [NormedSpace ℝ F] [CompleteSpace F] (hX : Measurable X) (hY : AEMeasurable Y μ) (hf : AEStronglyMeasurable f (μ.map fun a => (X a, Y a))) (hf_int : Integrable (fun a => f (X a, Y a)) μ) : μ[fun a => f (X a, Y a)|mβ.comap X] =ᵐ[μ] fun a => ∫ y, f (X a, y) ∂condDistrib Y X μ (X a) := haveI hf_int' : Integrable f (μ.map fun a => (X a, Y a)) := by rwa [integrable_map_measure hf (hX.aemeasurable.prod_mk hY)] - condexp_prod_ae_eq_integral_condDistrib' hX hY hf_int' + condExp_prod_ae_eq_integral_condDistrib' hX hY hf_int' + +@[deprecated (since := "2025-01-21")] +alias condexp_prod_ae_eq_integral_condDistrib₀ := condExp_prod_ae_eq_integral_condDistrib₀ /-- The conditional expectation of a function `f` of the product `(X, Y)` is almost everywhere equal to the integral of `y ↦ f(X, y)` against the `condDistrib` kernel. -/ -theorem condexp_prod_ae_eq_integral_condDistrib [NormedSpace ℝ F] [CompleteSpace F] +theorem condExp_prod_ae_eq_integral_condDistrib [NormedSpace ℝ F] [CompleteSpace F] (hX : Measurable X) (hY : AEMeasurable Y μ) (hf : StronglyMeasurable f) (hf_int : Integrable (fun a => f (X a, Y a)) μ) : μ[fun a => f (X a, Y a)|mβ.comap X] =ᵐ[μ] fun a => ∫ y, f (X a, y) ∂condDistrib Y X μ (X a) := haveI hf_int' : Integrable f (μ.map fun a => (X a, Y a)) := by rwa [integrable_map_measure hf.aestronglyMeasurable (hX.aemeasurable.prod_mk hY)] - condexp_prod_ae_eq_integral_condDistrib' hX hY hf_int' + condExp_prod_ae_eq_integral_condDistrib' hX hY hf_int' -theorem condexp_ae_eq_integral_condDistrib [NormedSpace ℝ F] [CompleteSpace F] (hX : Measurable X) +@[deprecated (since := "2025-01-21")] +alias condexp_prod_ae_eq_integral_condDistrib := condExp_prod_ae_eq_integral_condDistrib + +theorem condExp_ae_eq_integral_condDistrib [NormedSpace ℝ F] [CompleteSpace F] (hX : Measurable X) (hY : AEMeasurable Y μ) {f : Ω → F} (hf : StronglyMeasurable f) (hf_int : Integrable (fun a => f (Y a)) μ) : μ[fun a => f (Y a)|mβ.comap X] =ᵐ[μ] fun a => ∫ y, f y ∂condDistrib Y X μ (X a) := - condexp_prod_ae_eq_integral_condDistrib hX hY (hf.comp_measurable measurable_snd) hf_int + condExp_prod_ae_eq_integral_condDistrib hX hY (hf.comp_measurable measurable_snd) hf_int + +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_integral_condDistrib := condExp_ae_eq_integral_condDistrib /-- The conditional expectation of `Y` given `X` is almost everywhere equal to the integral `∫ y, y ∂(condDistrib Y X μ (X a))`. -/ -theorem condexp_ae_eq_integral_condDistrib' {Ω} [NormedAddCommGroup Ω] [NormedSpace ℝ Ω] +theorem condExp_ae_eq_integral_condDistrib' {Ω} [NormedAddCommGroup Ω] [NormedSpace ℝ Ω] [CompleteSpace Ω] [MeasurableSpace Ω] [BorelSpace Ω] [SecondCountableTopology Ω] {Y : α → Ω} (hX : Measurable X) (hY_int : Integrable Y μ) : μ[Y|mβ.comap X] =ᵐ[μ] fun a => ∫ y, y ∂condDistrib Y X μ (X a) := - condexp_ae_eq_integral_condDistrib hX hY_int.1.aemeasurable stronglyMeasurable_id hY_int + condExp_ae_eq_integral_condDistrib hX hY_int.1.aemeasurable stronglyMeasurable_id hY_int + +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_integral_condDistrib' := condExp_ae_eq_integral_condDistrib' open MeasureTheory @@ -319,9 +338,12 @@ theorem integrable_comp_snd_map_prod_mk_iff {Ω} {_ : MeasurableSpace Ω} {X : Integrable (fun x : β × Ω => f x.2) (μ.map fun ω => (X ω, ω)) ↔ Integrable f μ := ⟨fun h => h.comp_measurable (hX.prod_mk measurable_id), fun h => h.comp_snd_map_prod_mk X⟩ -theorem condexp_ae_eq_integral_condDistrib_id [NormedSpace ℝ F] [CompleteSpace F] {X : Ω → β} +theorem condExp_ae_eq_integral_condDistrib_id [NormedSpace ℝ F] [CompleteSpace F] {X : Ω → β} {μ : Measure Ω} [IsFiniteMeasure μ] (hX : Measurable X) {f : Ω → F} (hf_int : Integrable f μ) : μ[f|mβ.comap X] =ᵐ[μ] fun a => ∫ y, f y ∂condDistrib id X μ (X a) := - condexp_prod_ae_eq_integral_condDistrib' hX aemeasurable_id (hf_int.comp_snd_map_prod_mk X) + condExp_prod_ae_eq_integral_condDistrib' hX aemeasurable_id (hf_int.comp_snd_map_prod_mk X) + +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_integral_condDistrib_id := condExp_ae_eq_integral_condDistrib_id end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Condexp.lean b/Mathlib/Probability/Kernel/Condexp.lean index 56675cbb11d1e..d12727bf44629 100644 --- a/Mathlib/Probability/Kernel/Condexp.lean +++ b/Mathlib/Probability/Kernel/Condexp.lean @@ -9,8 +9,8 @@ import Mathlib.Probability.ConditionalProbability /-! # Kernel associated with a conditional expectation -We define `condexpKernel μ m`, a kernel from `Ω` to `Ω` such that for all integrable functions `f`, -`μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condexpKernel μ m ω)`. +We define `condExpKernel μ m`, a kernel from `Ω` to `Ω` such that for all integrable functions `f`, +`μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condExpKernel μ m ω)`. This kernel is defined if `Ω` is a standard Borel space. In general, `μ⟦s | m⟧` maps a measurable set `s` to a function `Ω → ℝ≥0∞`, and for all `s` that map is unique up to a `μ`-null set. For all @@ -21,11 +21,11 @@ on `Ω` allows us to do so. ## Main definitions -* `condexpKernel μ m`: kernel such that `μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condexpKernel μ m ω)`. +* `condExpKernel μ m`: kernel such that `μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condExpKernel μ m ω)`. ## Main statements -* `condexp_ae_eq_integral_condexpKernel`: `μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condexpKernel μ m ω)`. +* `condExp_ae_eq_integral_condExpKernel`: `μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condExpKernel μ m ω)`. -/ @@ -41,7 +41,8 @@ section AuxLemmas variable {Ω F : Type*} {m mΩ : MeasurableSpace Ω} {μ : Measure Ω} {f : Ω → F} theorem _root_.MeasureTheory.AEStronglyMeasurable.comp_snd_map_prod_id [TopologicalSpace F] - (hm : m ≤ mΩ) (hf : AEStronglyMeasurable f μ) : AEStronglyMeasurable (fun x : Ω × Ω => f x.2) + (hm : m ≤ mΩ) (hf : AEStronglyMeasurable f μ) : + AEStronglyMeasurable[m.prod mΩ] (fun x : Ω × Ω => f x.2) (@Measure.map Ω (Ω × Ω) mΩ (m.prod mΩ) (fun ω => (id ω, id ω)) μ) := by rw [← aestronglyMeasurable_comp_snd_map_prod_mk_iff (measurable_id'' hm)] at hf simp_rw [id] at hf ⊢ @@ -61,68 +62,90 @@ variable {Ω F : Type*} {m : MeasurableSpace Ω} [mΩ : MeasurableSpace Ω] open Classical in /-- Kernel associated with the conditional expectation with respect to a σ-algebra. It satisfies -`μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condexpKernel μ m ω)`. +`μ[f | m] =ᵐ[μ] fun ω => ∫ y, f y ∂(condExpKernel μ m ω)`. It is defined as the conditional distribution of the identity given the identity, where the second identity is understood as a map from `Ω` with the σ-algebra `mΩ` to `Ω` with σ-algebra `m ⊓ mΩ`. We use `m ⊓ mΩ` instead of `m` to ensure that it is a sub-σ-algebra of `mΩ`. We then use `Kernel.comap` to get a kernel from `m` to `mΩ` instead of from `m ⊓ mΩ` to `mΩ`. -/ -noncomputable irreducible_def condexpKernel (μ : Measure Ω) [IsFiniteMeasure μ] +noncomputable irreducible_def condExpKernel (μ : Measure Ω) [IsFiniteMeasure μ] (m : MeasurableSpace Ω) : @Kernel Ω Ω m mΩ := if _h : Nonempty Ω then Kernel.comap (@condDistrib Ω Ω Ω mΩ _ _ mΩ (m ⊓ mΩ) id id μ _) id (measurable_id'' (inf_le_left : m ⊓ mΩ ≤ m)) else 0 -lemma condexpKernel_eq (μ : Measure Ω) [IsFiniteMeasure μ] [h : Nonempty Ω] +@[deprecated (since := "2025-01-21")] alias condexpKernel := condExpKernel + +lemma condExpKernel_eq (μ : Measure Ω) [IsFiniteMeasure μ] [h : Nonempty Ω] (m : MeasurableSpace Ω) : - condexpKernel (mΩ := mΩ) μ m = Kernel.comap (@condDistrib Ω Ω Ω mΩ _ _ mΩ (m ⊓ mΩ) id id μ _) id + condExpKernel (mΩ := mΩ) μ m = Kernel.comap (@condDistrib Ω Ω Ω mΩ _ _ mΩ (m ⊓ mΩ) id id μ _) id (measurable_id'' (inf_le_left : m ⊓ mΩ ≤ m)) := by - simp [condexpKernel, h] + simp [condExpKernel, h] + +@[deprecated (since := "2025-01-21")] alias condexpKernel_eq := condExpKernel_eq + +lemma condExpKernel_apply_eq_condDistrib [Nonempty Ω] {ω : Ω} : + condExpKernel μ m ω = @condDistrib Ω Ω Ω mΩ _ _ mΩ (m ⊓ mΩ) id id μ _ (id ω) := by + simp [condExpKernel_eq, Kernel.comap_apply] -lemma condexpKernel_apply_eq_condDistrib [Nonempty Ω] {ω : Ω} : - condexpKernel μ m ω = @condDistrib Ω Ω Ω mΩ _ _ mΩ (m ⊓ mΩ) id id μ _ (id ω) := by - simp [condexpKernel_eq, Kernel.comap_apply] +@[deprecated (since := "2025-01-21")] +alias condexpKernel_apply_eq_condDistrib := condExpKernel_apply_eq_condDistrib -instance : IsMarkovKernel (condexpKernel μ m) := by +instance : IsMarkovKernel (condExpKernel μ m) := by rcases isEmpty_or_nonempty Ω with h | h · exact ⟨fun a ↦ (IsEmpty.false a).elim⟩ - · simp [condexpKernel, h]; infer_instance + · simp [condExpKernel, h]; infer_instance section Measurability variable [NormedAddCommGroup F] {f : Ω → F} -theorem measurable_condexpKernel {s : Set Ω} (hs : MeasurableSet s) : - Measurable[m] fun ω => condexpKernel μ m ω s := by +theorem measurable_condExpKernel {s : Set Ω} (hs : MeasurableSet s) : + Measurable[m] fun ω => condExpKernel μ m ω s := by nontriviality Ω - simp_rw [condexpKernel_apply_eq_condDistrib] + simp_rw [condExpKernel_apply_eq_condDistrib] refine Measurable.mono ?_ (inf_le_left : m ⊓ mΩ ≤ m) le_rfl convert measurable_condDistrib (μ := μ) hs rw [MeasurableSpace.comap_id] -theorem stronglyMeasurable_condexpKernel {s : Set Ω} (hs : MeasurableSet s) : - StronglyMeasurable[m] fun ω => condexpKernel μ m ω s := - Measurable.stronglyMeasurable (measurable_condexpKernel hs) +@[deprecated (since := "2025-01-21")] alias measurable_condexpKernel := measurable_condExpKernel -theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condexpKernel [NormedSpace ℝ F] +theorem stronglyMeasurable_condExpKernel {s : Set Ω} (hs : MeasurableSet s) : + StronglyMeasurable[m] fun ω => condExpKernel μ m ω s := + Measurable.stronglyMeasurable (measurable_condExpKernel hs) + +@[deprecated (since := "2025-01-21")] +alias stronglyMeasurable_condexpKernel := stronglyMeasurable_condExpKernel + +theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condExpKernel [NormedSpace ℝ F] (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable (fun ω => ∫ y, f y ∂condexpKernel μ m ω) μ := by + AEStronglyMeasurable (fun ω => ∫ y, f y ∂condExpKernel μ m ω) μ := by nontriviality Ω - simp_rw [condexpKernel_apply_eq_condDistrib] + simp_rw [condExpKernel_apply_eq_condDistrib] exact AEStronglyMeasurable.integral_condDistrib (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) aemeasurable_id (hf.comp_snd_map_prod_id inf_le_right) -theorem aestronglyMeasurable'_integral_condexpKernel [NormedSpace ℝ F] +@[deprecated (since := "2025-01-21")] +alias _root_.MeasureTheory.AEStronglyMeasurable.integral_condexpKernel := + _root_.MeasureTheory.AEStronglyMeasurable.integral_condExpKernel + +theorem aestronglyMeasurable_integral_condExpKernel [NormedSpace ℝ F] (hf : AEStronglyMeasurable f μ) : - AEStronglyMeasurable' m (fun ω => ∫ y, f y ∂condexpKernel μ m ω) μ := by + AEStronglyMeasurable[m] (fun ω => ∫ y, f y ∂condExpKernel μ m ω) μ := by nontriviality Ω - rw [condexpKernel_eq] - have h := aestronglyMeasurable'_integral_condDistrib + rw [condExpKernel_eq] + have h := aestronglyMeasurable_integral_condDistrib (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) aemeasurable_id (hf.comp_snd_map_prod_id (inf_le_right : m ⊓ mΩ ≤ mΩ)) rw [MeasurableSpace.comap_id] at h - exact AEStronglyMeasurable'.mono h inf_le_left + exact h.mono inf_le_left + +@[deprecated (since := "2025-01-24")] +alias aestronglyMeasurable'_integral_condExpKernel := aestronglyMeasurable_integral_condExpKernel + +@[deprecated (since := "2025-01-21")] +alias aestronglyMeasurable'_integral_condexpKernel := aestronglyMeasurable_integral_condExpKernel end Measurability @@ -130,89 +153,122 @@ section Integrability variable [NormedAddCommGroup F] {f : Ω → F} -theorem _root_.MeasureTheory.Integrable.condexpKernel_ae (hf_int : Integrable f μ) : - ∀ᵐ ω ∂μ, Integrable f (condexpKernel μ m ω) := by +theorem _root_.MeasureTheory.Integrable.condExpKernel_ae (hf_int : Integrable f μ) : + ∀ᵐ ω ∂μ, Integrable f (condExpKernel μ m ω) := by nontriviality Ω - rw [condexpKernel_eq] + rw [condExpKernel_eq] convert Integrable.condDistrib_ae (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) aemeasurable_id (hf_int.comp_snd_map_prod_id (inf_le_right : m ⊓ mΩ ≤ mΩ)) using 1 -theorem _root_.MeasureTheory.Integrable.integral_norm_condexpKernel (hf_int : Integrable f μ) : - Integrable (fun ω => ∫ y, ‖f y‖ ∂condexpKernel μ m ω) μ := by +@[deprecated (since := "2025-01-21")] +alias _root_.MeasureTheory.Integrable.condexpKernel_ae := + _root_.MeasureTheory.Integrable.condExpKernel_ae + +theorem _root_.MeasureTheory.Integrable.integral_norm_condExpKernel (hf_int : Integrable f μ) : + Integrable (fun ω => ∫ y, ‖f y‖ ∂condExpKernel μ m ω) μ := by nontriviality Ω - rw [condexpKernel_eq] + rw [condExpKernel_eq] convert Integrable.integral_norm_condDistrib (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) aemeasurable_id (hf_int.comp_snd_map_prod_id (inf_le_right : m ⊓ mΩ ≤ mΩ)) using 1 -theorem _root_.MeasureTheory.Integrable.norm_integral_condexpKernel [NormedSpace ℝ F] +@[deprecated (since := "2025-01-21")] +alias _root_.MeasureTheory.Integrable.integral_norm_condexpKernel := + _root_.MeasureTheory.Integrable.integral_norm_condExpKernel + +theorem _root_.MeasureTheory.Integrable.norm_integral_condExpKernel [NormedSpace ℝ F] (hf_int : Integrable f μ) : - Integrable (fun ω => ‖∫ y, f y ∂condexpKernel μ m ω‖) μ := by + Integrable (fun ω => ‖∫ y, f y ∂condExpKernel μ m ω‖) μ := by nontriviality Ω - rw [condexpKernel_eq] + rw [condExpKernel_eq] convert Integrable.norm_integral_condDistrib (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) aemeasurable_id (hf_int.comp_snd_map_prod_id (inf_le_right : m ⊓ mΩ ≤ mΩ)) using 1 -theorem _root_.MeasureTheory.Integrable.integral_condexpKernel [NormedSpace ℝ F] +@[deprecated (since := "2025-01-21")] +alias _root_.MeasureTheory.Integrable.norm_integral_condexpKernel := + _root_.MeasureTheory.Integrable.norm_integral_condExpKernel + +theorem _root_.MeasureTheory.Integrable.integral_condExpKernel [NormedSpace ℝ F] (hf_int : Integrable f μ) : - Integrable (fun ω => ∫ y, f y ∂condexpKernel μ m ω) μ := by + Integrable (fun ω => ∫ y, f y ∂condExpKernel μ m ω) μ := by nontriviality Ω - rw [condexpKernel_eq] + rw [condExpKernel_eq] convert Integrable.integral_condDistrib (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) aemeasurable_id (hf_int.comp_snd_map_prod_id (inf_le_right : m ⊓ mΩ ≤ mΩ)) using 1 -theorem integrable_toReal_condexpKernel {s : Set Ω} (hs : MeasurableSet s) : - Integrable (fun ω => (condexpKernel μ m ω s).toReal) μ := by +@[deprecated (since := "2025-01-21")] +alias _root_.MeasureTheory.Integrable.integral_condexpKernel := + _root_.MeasureTheory.Integrable.integral_condExpKernel + +theorem integrable_toReal_condExpKernel {s : Set Ω} (hs : MeasurableSet s) : + Integrable (fun ω => (condExpKernel μ m ω s).toReal) μ := by nontriviality Ω - rw [condexpKernel_eq] + rw [condExpKernel_eq] exact integrable_toReal_condDistrib (aemeasurable_id'' μ (inf_le_right : m ⊓ mΩ ≤ mΩ)) hs +@[deprecated (since := "2025-01-21")] +alias integrable_toReal_condexpKernel := integrable_toReal_condExpKernel + end Integrability -lemma condexpKernel_ae_eq_condexp' {s : Set Ω} (hs : MeasurableSet s) : - (fun ω ↦ (condexpKernel μ m ω s).toReal) =ᵐ[μ] μ⟦s | m ⊓ mΩ⟧ := by +lemma condExpKernel_ae_eq_condExp' {s : Set Ω} (hs : MeasurableSet s) : + (fun ω ↦ (condExpKernel μ m ω s).toReal) =ᵐ[μ] μ⟦s | m ⊓ mΩ⟧ := by rcases isEmpty_or_nonempty Ω with h | h · have : μ = 0 := Measure.eq_zero_of_isEmpty μ simpa [this] using trivial - have h := condDistrib_ae_eq_condexp (μ := μ) + have h := condDistrib_ae_eq_condExp (μ := μ) (measurable_id'' (inf_le_right : m ⊓ mΩ ≤ mΩ)) measurable_id hs simp only [id_eq, MeasurableSpace.comap_id, preimage_id_eq] at h - simp_rw [condexpKernel_apply_eq_condDistrib] + simp_rw [condExpKernel_apply_eq_condDistrib] exact h -lemma condexpKernel_ae_eq_condexp +@[deprecated (since := "2025-01-21")] +alias condexpKernel_ae_eq_condexp' := condExpKernel_ae_eq_condExp' + +lemma condExpKernel_ae_eq_condExp (hm : m ≤ mΩ) {s : Set Ω} (hs : MeasurableSet s) : - (fun ω ↦ (condexpKernel μ m ω s).toReal) =ᵐ[μ] μ⟦s | m⟧ := - (condexpKernel_ae_eq_condexp' hs).trans (by rw [inf_of_le_left hm]) + (fun ω ↦ (condExpKernel μ m ω s).toReal) =ᵐ[μ] μ⟦s | m⟧ := + (condExpKernel_ae_eq_condExp' hs).trans (by rw [inf_of_le_left hm]) -lemma condexpKernel_ae_eq_trim_condexp +@[deprecated (since := "2025-01-21")] +alias condexpKernel_ae_eq_condexp := condExpKernel_ae_eq_condExp + +lemma condExpKernel_ae_eq_trim_condExp (hm : m ≤ mΩ) {s : Set Ω} (hs : MeasurableSet s) : - (fun ω ↦ (condexpKernel μ m ω s).toReal) =ᵐ[μ.trim hm] μ⟦s | m⟧ := by - rw [ae_eq_trim_iff hm _ stronglyMeasurable_condexp] - · exact condexpKernel_ae_eq_condexp hm hs - · refine Measurable.stronglyMeasurable ?_ - exact @Measurable.ennreal_toReal _ m _ (measurable_condexpKernel hs) + (fun ω ↦ (condExpKernel μ m ω s).toReal) =ᵐ[μ.trim hm] μ⟦s | m⟧ := by + rw [(measurable_condExpKernel hs).ennreal_toReal.stronglyMeasurable.ae_eq_trim_iff hm + stronglyMeasurable_condExp] + exact condExpKernel_ae_eq_condExp hm hs + +@[deprecated (since := "2025-01-21")] +alias condexpKernel_ae_eq_trim_condexp := condExpKernel_ae_eq_trim_condExp -theorem condexp_ae_eq_integral_condexpKernel' [NormedAddCommGroup F] {f : Ω → F} +theorem condExp_ae_eq_integral_condExpKernel' [NormedAddCommGroup F] {f : Ω → F} [NormedSpace ℝ F] [CompleteSpace F] (hf_int : Integrable f μ) : - μ[f|m ⊓ mΩ] =ᵐ[μ] fun ω => ∫ y, f y ∂condexpKernel μ m ω := by + μ[f|m ⊓ mΩ] =ᵐ[μ] fun ω => ∫ y, f y ∂condExpKernel μ m ω := by rcases isEmpty_or_nonempty Ω with h | h · have : μ = 0 := Measure.eq_zero_of_isEmpty μ simpa [this] using trivial have hX : @Measurable Ω Ω mΩ (m ⊓ mΩ) id := measurable_id.mono le_rfl (inf_le_right : m ⊓ mΩ ≤ mΩ) - simp_rw [condexpKernel_apply_eq_condDistrib] - have h := condexp_ae_eq_integral_condDistrib_id hX hf_int + simp_rw [condExpKernel_apply_eq_condDistrib] + have h := condExp_ae_eq_integral_condDistrib_id hX hf_int simpa only [MeasurableSpace.comap_id, id_eq] using h +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_integral_condexpKernel' := condExp_ae_eq_integral_condExpKernel' + /-- The conditional expectation of `f` with respect to a σ-algebra `m` is almost everywhere equal to -the integral `∫ y, f y ∂(condexpKernel μ m ω)`. -/ -theorem condexp_ae_eq_integral_condexpKernel [NormedAddCommGroup F] {f : Ω → F} +the integral `∫ y, f y ∂(condExpKernel μ m ω)`. -/ +theorem condExp_ae_eq_integral_condExpKernel [NormedAddCommGroup F] {f : Ω → F} [NormedSpace ℝ F] [CompleteSpace F] (hm : m ≤ mΩ) (hf_int : Integrable f μ) : - μ[f|m] =ᵐ[μ] fun ω => ∫ y, f y ∂condexpKernel μ m ω := - ((condexp_ae_eq_integral_condexpKernel' hf_int).symm.trans (by rw [inf_of_le_left hm])).symm + μ[f|m] =ᵐ[μ] fun ω => ∫ y, f y ∂condExpKernel μ m ω := + ((condExp_ae_eq_integral_condExpKernel' hf_int).symm.trans (by rw [inf_of_le_left hm])).symm + +@[deprecated (since := "2025-01-21")] +alias condexp_ae_eq_integral_condexpKernel := condExp_ae_eq_integral_condExpKernel section Cond @@ -224,17 +280,16 @@ variable {s t : Set Ω} [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpac omit [StandardBorelSpace Ω] -lemma condexp_generateFrom_singleton (hs : MeasurableSet s) {f : Ω → F} (hf : Integrable f μ) : +lemma condExp_generateFrom_singleton (hs : MeasurableSet s) {f : Ω → F} (hf : Integrable f μ) : μ[f | generateFrom {s}] =ᵐ[μ.restrict s] fun _ ↦ ∫ x, f x ∂μ[|s] := by by_cases hμs : μ s = 0 · rw [Measure.restrict_eq_zero.2 hμs] rfl - refine ae_eq_trans (condexp_restrict_ae_eq_restrict + refine ae_eq_trans (condExp_restrict_ae_eq_restrict (generateFrom_singleton_le hs) (measurableSet_generateFrom rfl) hf).symm ?_ - · refine (ae_eq_condexp_of_forall_setIntegral_eq - (generateFrom_singleton_le hs) hf.restrict ?_ ?_ - stronglyMeasurable_const.aeStronglyMeasurable').symm + · refine (ae_eq_condExp_of_forall_setIntegral_eq (generateFrom_singleton_le hs) hf.restrict ?_ ?_ + stronglyMeasurable_const.aestronglyMeasurable).symm · rintro t - - rw [integrableOn_const] exact Or.inr <| measure_lt_top (μ.restrict s) t @@ -252,22 +307,31 @@ lemma condexp_generateFrom_singleton (hs : MeasurableSet s) {f : Ω → F} (hf : integral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter, smul_inv_smul₀ <| ENNReal.toReal_ne_zero.2 ⟨hμs, measure_ne_top _ _⟩] -lemma condexp_set_generateFrom_singleton (hs : MeasurableSet s) (ht : MeasurableSet t) : +@[deprecated (since := "2025-01-21")] +alias condexp_generateFrom_singleton := condExp_generateFrom_singleton + +lemma condExp_set_generateFrom_singleton (hs : MeasurableSet s) (ht : MeasurableSet t) : μ⟦t | generateFrom {s}⟧ =ᵐ[μ.restrict s] fun _ ↦ (μ[t|s]).toReal := by rw [← integral_indicator_one ht] - exact condexp_generateFrom_singleton hs <| Integrable.indicator (integrable_const 1) ht + exact condExp_generateFrom_singleton hs <| Integrable.indicator (integrable_const 1) ht -lemma condexpKernel_singleton_ae_eq_cond [StandardBorelSpace Ω] (hs : MeasurableSet s) +@[deprecated (since := "2025-01-21")] +alias condexp_set_generateFrom_singleton := condExp_set_generateFrom_singleton + +lemma condExpKernel_singleton_ae_eq_cond [StandardBorelSpace Ω] (hs : MeasurableSet s) (ht : MeasurableSet t) : ∀ᵐ ω ∂μ.restrict s, - condexpKernel μ (generateFrom {s}) ω t = μ[t|s] := by - have : (fun ω ↦ (condexpKernel μ (generateFrom {s}) ω t).toReal) =ᵐ[μ.restrict s] + condExpKernel μ (generateFrom {s}) ω t = μ[t|s] := by + have : (fun ω ↦ (condExpKernel μ (generateFrom {s}) ω t).toReal) =ᵐ[μ.restrict s] μ⟦t | generateFrom {s}⟧ := - ae_restrict_le <| condexpKernel_ae_eq_condexp + ae_restrict_le <| condExpKernel_ae_eq_condExp (generateFrom_singleton_le hs) ht - filter_upwards [condexp_set_generateFrom_singleton hs ht, this] with ω hω₁ hω₂ + filter_upwards [condExp_set_generateFrom_singleton hs ht, this] with ω hω₁ hω₂ rwa [hω₁, ENNReal.toReal_eq_toReal (measure_ne_top _ t) (measure_ne_top _ t)] at hω₂ +@[deprecated (since := "2025-01-21")] +alias condexpKernel_singleton_ae_eq_cond := condExpKernel_singleton_ae_eq_cond + end Cond end ProbabilityTheory diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index 9dd84c89fa298..1ac6c0d62bdfe 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -288,23 +288,25 @@ lemma setIntegral_densityProcess_of_le (hκν : fst κ ≤ ν) @[deprecated (since := "2024-04-17")] alias set_integral_densityProcess_of_le := setIntegral_densityProcess_of_le -lemma condexp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] +lemma condExp_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] {i j : ℕ} (hij : i ≤ j) (a : α) {s : Set β} (hs : MeasurableSet s) : (ν a)[fun x ↦ densityProcess κ ν j a x s | countableFiltration γ i] =ᵐ[ν a] fun x ↦ densityProcess κ ν i a x s := by - refine (ae_eq_condexp_of_forall_setIntegral_eq ?_ ?_ ?_ ?_ ?_).symm + refine (ae_eq_condExp_of_forall_setIntegral_eq ?_ ?_ ?_ ?_ ?_).symm · exact integrable_densityProcess hκν j a hs · exact fun _ _ _ ↦ (integrable_densityProcess hκν _ _ hs).integrableOn · intro x hx _ rw [setIntegral_densityProcess hκν i a hs hx, setIntegral_densityProcess_of_le hκν hij a hs hx] - · exact StronglyMeasurable.aeStronglyMeasurable' + · exact StronglyMeasurable.aestronglyMeasurable (stronglyMeasurable_countableFiltration_densityProcess κ ν i a hs) +@[deprecated (since := "2025-01-21")] alias condexp_densityProcess := condExp_densityProcess + lemma martingale_densityProcess (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) {s : Set β} (hs : MeasurableSet s) : Martingale (fun n x ↦ densityProcess κ ν n a x s) (countableFiltration γ) (ν a) := - ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condexp_densityProcess hκν h a hs⟩ + ⟨adapted_densityProcess κ ν a hs, fun _ _ h ↦ condExp_densityProcess hκν h a hs⟩ lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) {s s' : Set β} (h : s ⊆ s') : diff --git a/Mathlib/Probability/Martingale/Basic.lean b/Mathlib/Probability/Martingale/Basic.lean index 1693e5117c3b3..9b7789b528d40 100644 --- a/Mathlib/Probability/Martingale/Basic.lean +++ b/Mathlib/Probability/Martingale/Basic.lean @@ -30,7 +30,7 @@ The definitions of filtration and adapted can be found in `Probability.Process.S ### Results -* `MeasureTheory.martingale_condexp f ℱ μ`: the sequence `fun i => μ[f | ℱ i, ℱ.le i])` is a +* `MeasureTheory.martingale_condExp f ℱ μ`: the sequence `fun i => μ[f | ℱ i, ℱ.le i])` is a martingale with respect to `ℱ` and `μ`. -/ @@ -64,18 +64,18 @@ def Submartingale [LE E] (f : ι → Ω → E) (ℱ : Filtration ι m0) (μ : Me theorem martingale_const (ℱ : Filtration ι m0) (μ : Measure Ω) [IsFiniteMeasure μ] (x : E) : Martingale (fun _ _ => x) ℱ μ := - ⟨adapted_const ℱ _, fun i j _ => by rw [condexp_const (ℱ.le _)]⟩ + ⟨adapted_const ℱ _, fun i j _ => by rw [condExp_const (ℱ.le _)]⟩ theorem martingale_const_fun [OrderBot ι] (ℱ : Filtration ι m0) (μ : Measure Ω) [IsFiniteMeasure μ] {f : Ω → E} (hf : StronglyMeasurable[ℱ ⊥] f) (hfint : Integrable f μ) : Martingale (fun _ => f) ℱ μ := by refine ⟨fun i => hf.mono <| ℱ.mono bot_le, fun i j _ => ?_⟩ - rw [condexp_of_stronglyMeasurable (ℱ.le _) (hf.mono <| ℱ.mono bot_le) hfint] + rw [condExp_of_stronglyMeasurable (ℱ.le _) (hf.mono <| ℱ.mono bot_le) hfint] variable (E) theorem martingale_zero (ℱ : Filtration ι m0) (μ : Measure Ω) : Martingale (0 : ι → Ω → E) ℱ μ := - ⟨adapted_zero E ℱ, fun i j _ => by rw [Pi.zero_apply, condexp_zero]; simp⟩ + ⟨adapted_zero E ℱ, fun i j _ => by rw [Pi.zero_apply, condExp_zero]; simp⟩ variable {E} @@ -88,15 +88,17 @@ protected theorem stronglyMeasurable (hf : Martingale f ℱ μ) (i : ι) : StronglyMeasurable[ℱ i] (f i) := hf.adapted i -theorem condexp_ae_eq (hf : Martingale f ℱ μ) {i j : ι} (hij : i ≤ j) : μ[f j|ℱ i] =ᵐ[μ] f i := +theorem condExp_ae_eq (hf : Martingale f ℱ μ) {i j : ι} (hij : i ≤ j) : μ[f j|ℱ i] =ᵐ[μ] f i := hf.2 i j hij +@[deprecated (since := "2025-01-21")] alias condexp_ae_eq := condExp_ae_eq + protected theorem integrable (hf : Martingale f ℱ μ) (i : ι) : Integrable (f i) μ := - integrable_condexp.congr (hf.condexp_ae_eq (le_refl i)) + integrable_condExp.congr (hf.condExp_ae_eq (le_refl i)) theorem setIntegral_eq [SigmaFiniteFiltration μ ℱ] (hf : Martingale f ℱ μ) {i j : ι} (hij : i ≤ j) {s : Set Ω} (hs : MeasurableSet[ℱ i] s) : ∫ ω in s, f i ω ∂μ = ∫ ω in s, f j ω ∂μ := by - rw [← @setIntegral_condexp _ _ _ _ _ (ℱ i) m0 _ _ _ (ℱ.le i) _ (hf.integrable j) hs] + rw [← setIntegral_condExp (ℱ.le i) (hf.integrable j) hs] refine setIntegral_congr_ae (ℱ.le i s hs) ?_ filter_upwards [hf.2 i j hij] with _ heq _ using heq.symm @@ -105,17 +107,18 @@ alias set_integral_eq := setIntegral_eq theorem add (hf : Martingale f ℱ μ) (hg : Martingale g ℱ μ) : Martingale (f + g) ℱ μ := by refine ⟨hf.adapted.add hg.adapted, fun i j hij => ?_⟩ - exact (condexp_add (hf.integrable j) (hg.integrable j)).trans ((hf.2 i j hij).add (hg.2 i j hij)) + exact (condExp_add (hf.integrable j) (hg.integrable j) _).trans + ((hf.2 i j hij).add (hg.2 i j hij)) theorem neg (hf : Martingale f ℱ μ) : Martingale (-f) ℱ μ := - ⟨hf.adapted.neg, fun i j hij => (condexp_neg (f j)).trans (hf.2 i j hij).neg⟩ + ⟨hf.adapted.neg, fun i j hij => (condExp_neg ..).trans (hf.2 i j hij).neg⟩ theorem sub (hf : Martingale f ℱ μ) (hg : Martingale g ℱ μ) : Martingale (f - g) ℱ μ := by rw [sub_eq_add_neg]; exact hf.add hg.neg theorem smul (c : ℝ) (hf : Martingale f ℱ μ) : Martingale (c • f) ℱ μ := by refine ⟨hf.adapted.smul c, fun i j hij => ?_⟩ - refine (condexp_smul c (f j)).trans ((hf.2 i j hij).mono fun x hx => ?_) + refine (condExp_smul ..).trans ((hf.2 i j hij).mono fun x hx => ?_) simp only [Pi.smul_apply, hx] theorem supermartingale [Preorder E] (hf : Martingale f ℱ μ) : Supermartingale f ℱ μ := @@ -131,9 +134,11 @@ theorem martingale_iff [PartialOrder E] : ⟨fun hf => ⟨hf.supermartingale, hf.submartingale⟩, fun ⟨hf₁, hf₂⟩ => ⟨hf₁.1, fun i j hij => (hf₁.2.1 i j hij).antisymm (hf₂.2.1 i j hij)⟩⟩ -theorem martingale_condexp (f : Ω → E) (ℱ : Filtration ι m0) (μ : Measure Ω) +theorem martingale_condExp (f : Ω → E) (ℱ : Filtration ι m0) (μ : Measure Ω) [SigmaFiniteFiltration μ ℱ] : Martingale (fun i => μ[f|ℱ i]) ℱ μ := - ⟨fun _ => stronglyMeasurable_condexp, fun _ j hij => condexp_condexp_of_le (ℱ.mono hij) (ℱ.le j)⟩ + ⟨fun _ => stronglyMeasurable_condExp, fun _ j hij => condExp_condExp_of_le (ℱ.mono hij) (ℱ.le j)⟩ + +@[deprecated (since := "2025-01-21")] alias martingale_condexp := martingale_condExp namespace Supermartingale @@ -147,15 +152,17 @@ protected theorem stronglyMeasurable [LE E] (hf : Supermartingale f ℱ μ) (i : protected theorem integrable [LE E] (hf : Supermartingale f ℱ μ) (i : ι) : Integrable (f i) μ := hf.2.2 i -theorem condexp_ae_le [LE E] (hf : Supermartingale f ℱ μ) {i j : ι} (hij : i ≤ j) : +theorem condExp_ae_le [LE E] (hf : Supermartingale f ℱ μ) {i j : ι} (hij : i ≤ j) : μ[f j|ℱ i] ≤ᵐ[μ] f i := hf.2.1 i j hij +@[deprecated (since := "2025-01-21")] alias condexp_ae_le := condExp_ae_le + theorem setIntegral_le [SigmaFiniteFiltration μ ℱ] {f : ι → Ω → ℝ} (hf : Supermartingale f ℱ μ) {i j : ι} (hij : i ≤ j) {s : Set Ω} (hs : MeasurableSet[ℱ i] s) : ∫ ω in s, f j ω ∂μ ≤ ∫ ω in s, f i ω ∂μ := by - rw [← setIntegral_condexp (ℱ.le i) (hf.integrable j) hs] - refine setIntegral_mono_ae integrable_condexp.integrableOn (hf.integrable i).integrableOn ?_ + rw [← setIntegral_condExp (ℱ.le i) (hf.integrable j) hs] + refine setIntegral_mono_ae integrable_condExp.integrableOn (hf.integrable i).integrableOn ?_ filter_upwards [hf.2.1 i j hij] with _ heq using heq @[deprecated (since := "2024-04-17")] @@ -164,7 +171,7 @@ alias set_integral_le := setIntegral_le theorem add [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) (hg : Supermartingale g ℱ μ) : Supermartingale (f + g) ℱ μ := by refine ⟨hf.1.add hg.1, fun i j hij => ?_, fun i => (hf.2.2 i).add (hg.2.2 i)⟩ - refine (condexp_add (hf.integrable j) (hg.integrable j)).le.trans ?_ + refine (condExp_add (hf.integrable j) (hg.integrable j) _).le.trans ?_ filter_upwards [hf.2.1 i j hij, hg.2.1 i j hij] intros refine add_le_add ?_ ?_ <;> assumption @@ -176,7 +183,7 @@ theorem add_martingale [Preorder E] [AddLeftMono E] theorem neg [Preorder E] [AddLeftMono E] (hf : Supermartingale f ℱ μ) : Submartingale (-f) ℱ μ := by refine ⟨hf.1.neg, fun i j hij => ?_, fun i => (hf.2.2 i).neg⟩ - refine EventuallyLE.trans ?_ (condexp_neg (f j)).symm.le + refine EventuallyLE.trans ?_ (condExp_neg ..).symm.le filter_upwards [hf.2.1 i j hij] with _ _ simpa @@ -194,14 +201,16 @@ protected theorem stronglyMeasurable [LE E] (hf : Submartingale f ℱ μ) (i : protected theorem integrable [LE E] (hf : Submartingale f ℱ μ) (i : ι) : Integrable (f i) μ := hf.2.2 i -theorem ae_le_condexp [LE E] (hf : Submartingale f ℱ μ) {i j : ι} (hij : i ≤ j) : +theorem ae_le_condExp [LE E] (hf : Submartingale f ℱ μ) {i j : ι} (hij : i ≤ j) : f i ≤ᵐ[μ] μ[f j|ℱ i] := hf.2.1 i j hij +@[deprecated (since := "2025-01-21")] alias ae_le_condexp := ae_le_condExp + theorem add [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) (hg : Submartingale g ℱ μ) : Submartingale (f + g) ℱ μ := by refine ⟨hf.1.add hg.1, fun i j hij => ?_, fun i => (hf.2.2 i).add (hg.2.2 i)⟩ - refine EventuallyLE.trans ?_ (condexp_add (hf.integrable j) (hg.integrable j)).symm.le + refine EventuallyLE.trans ?_ (condExp_add (hf.integrable j) (hg.integrable j) _).symm.le filter_upwards [hf.2.1 i j hij, hg.2.1 i j hij] intros refine add_le_add ?_ ?_ <;> assumption @@ -212,7 +221,7 @@ theorem add_martingale [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ theorem neg [Preorder E] [AddLeftMono E] (hf : Submartingale f ℱ μ) : Supermartingale (-f) ℱ μ := by - refine ⟨hf.1.neg, fun i j hij => (condexp_neg (f j)).le.trans ?_, fun i => (hf.2.2 i).neg⟩ + refine ⟨hf.1.neg, fun i j hij => (condExp_neg ..).le.trans ?_, fun i => (hf.2.2 i).neg⟩ filter_upwards [hf.2.1 i j hij] with _ _ simpa @@ -240,10 +249,10 @@ protected theorem sup {f g : ι → Ω → ℝ} (hf : Submartingale f ℱ μ) (h fun i j hij => ?_, fun i => Integrable.sup (hf.integrable _) (hg.integrable _)⟩ refine EventuallyLE.sup_le ?_ ?_ · exact EventuallyLE.trans (hf.2.1 i j hij) - (condexp_mono (hf.integrable _) (Integrable.sup (hf.integrable j) (hg.integrable j)) + (condExp_mono (hf.integrable _) (Integrable.sup (hf.integrable j) (hg.integrable j)) (Eventually.of_forall fun x => le_max_left _ _)) · exact EventuallyLE.trans (hg.2.1 i j hij) - (condexp_mono (hg.integrable _) (Integrable.sup (hf.integrable j) (hg.integrable j)) + (condExp_mono (hg.integrable _) (Integrable.sup (hf.integrable j) (hg.integrable j)) (Eventually.of_forall fun x => le_max_right _ _)) protected theorem pos {f : ι → Ω → ℝ} (hf : Submartingale f ℱ μ) : Submartingale (f⁺) ℱ μ := @@ -263,37 +272,46 @@ theorem submartingale_of_setIntegral_le [IsFiniteMeasure μ] {f : ι → Ω → filter_upwards [this] with x hx rwa [← sub_nonneg] refine ae_nonneg_of_forall_setIntegral_nonneg - ((integrable_condexp.sub (hint i)).trim _ (stronglyMeasurable_condexp.sub <| hadp i)) + ((integrable_condExp.sub (hint i)).trim _ (stronglyMeasurable_condExp.sub <| hadp i)) fun s hs _ => ?_ specialize hf i j hij s hs - rwa [← setIntegral_trim _ (stronglyMeasurable_condexp.sub <| hadp i) hs, - integral_sub' integrable_condexp.integrableOn (hint i).integrableOn, sub_nonneg, - setIntegral_condexp (ℱ.le i) (hint j) hs] + rwa [← setIntegral_trim _ (stronglyMeasurable_condExp.sub <| hadp i) hs, + integral_sub' integrable_condExp.integrableOn (hint i).integrableOn, sub_nonneg, + setIntegral_condExp (ℱ.le i) (hint j) hs] @[deprecated (since := "2024-04-17")] alias submartingale_of_set_integral_le := submartingale_of_setIntegral_le -theorem submartingale_of_condexp_sub_nonneg [IsFiniteMeasure μ] {f : ι → Ω → ℝ} (hadp : Adapted ℱ f) +theorem submartingale_of_condExp_sub_nonneg [IsFiniteMeasure μ] {f : ι → Ω → ℝ} (hadp : Adapted ℱ f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i j, i ≤ j → 0 ≤ᵐ[μ] μ[f j - f i|ℱ i]) : Submartingale f ℱ μ := by refine ⟨hadp, fun i j hij => ?_, hint⟩ - rw [← condexp_of_stronglyMeasurable (ℱ.le _) (hadp _) (hint _), ← eventually_sub_nonneg] - exact EventuallyLE.trans (hf i j hij) (condexp_sub (hint _) (hint _)).le + rw [← condExp_of_stronglyMeasurable (ℱ.le _) (hadp _) (hint _), ← eventually_sub_nonneg] + exact EventuallyLE.trans (hf i j hij) (condExp_sub (hint _) (hint _) _).le + +@[deprecated (since := "2025-01-21")] +alias submartingale_of_condexp_sub_nonneg := submartingale_of_condExp_sub_nonneg -theorem Submartingale.condexp_sub_nonneg {f : ι → Ω → ℝ} (hf : Submartingale f ℱ μ) {i j : ι} +theorem Submartingale.condExp_sub_nonneg {f : ι → Ω → ℝ} (hf : Submartingale f ℱ μ) {i j : ι} (hij : i ≤ j) : 0 ≤ᵐ[μ] μ[f j - f i|ℱ i] := by by_cases h : SigmaFinite (μ.trim (ℱ.le i)) - swap; · rw [condexp_of_not_sigmaFinite (ℱ.le i) h] - refine EventuallyLE.trans ?_ (condexp_sub (hf.integrable _) (hf.integrable _)).symm.le + swap; · rw [condExp_of_not_sigmaFinite (ℱ.le i) h] + refine EventuallyLE.trans ?_ (condExp_sub (hf.integrable _) (hf.integrable _) _).symm.le rw [eventually_sub_nonneg, - condexp_of_stronglyMeasurable (ℱ.le _) (hf.adapted _) (hf.integrable _)] + condExp_of_stronglyMeasurable (ℱ.le _) (hf.adapted _) (hf.integrable _)] exact hf.2.1 i j hij -theorem submartingale_iff_condexp_sub_nonneg [IsFiniteMeasure μ] {f : ι → Ω → ℝ} : +@[deprecated (since := "2025-01-21")] +alias Submartingale.condexp_sub_nonneg := Submartingale.condExp_sub_nonneg + +theorem submartingale_iff_condExp_sub_nonneg [IsFiniteMeasure μ] {f : ι → Ω → ℝ} : Submartingale f ℱ μ ↔ Adapted ℱ f ∧ (∀ i, Integrable (f i) μ) ∧ ∀ i j, i ≤ j → 0 ≤ᵐ[μ] μ[f j - f i|ℱ i] := - ⟨fun h => ⟨h.adapted, h.integrable, fun _ _ => h.condexp_sub_nonneg⟩, fun ⟨hadp, hint, h⟩ => - submartingale_of_condexp_sub_nonneg hadp hint h⟩ + ⟨fun h => ⟨h.adapted, h.integrable, fun _ _ => h.condExp_sub_nonneg⟩, fun ⟨hadp, hint, h⟩ => + submartingale_of_condExp_sub_nonneg hadp hint h⟩ + +@[deprecated (since := "2025-01-21")] +alias submartingale_iff_condexp_sub_nonneg := submartingale_iff_condExp_sub_nonneg end Submartingale @@ -315,10 +333,8 @@ variable {F : Type*} [NormedLatticeAddCommGroup F] [NormedSpace ℝ F] [Complete theorem smul_nonneg {f : ι → Ω → F} {c : ℝ} (hc : 0 ≤ c) (hf : Supermartingale f ℱ μ) : Supermartingale (c • f) ℱ μ := by refine ⟨hf.1.smul c, fun i j hij => ?_, fun i => (hf.2.2 i).smul c⟩ - refine (condexp_smul c (f j)).le.trans ?_ - filter_upwards [hf.2.1 i j hij] with _ hle - simp_rw [Pi.smul_apply] - exact smul_le_smul_of_nonneg_left hle hc + filter_upwards [condExp_smul c (f j) (ℱ i), hf.2.1 i j hij] with ω hω hle + simpa only [hω, Pi.smul_apply] using smul_le_smul_of_nonneg_left hle hc theorem smul_nonpos {f : ι → Ω → F} {c : ℝ} (hc : c ≤ 0) (hf : Supermartingale f ℱ μ) : Submartingale (c • f) ℱ μ := by @@ -392,16 +408,16 @@ theorem submartingale_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Submartingale f 𝒢 μ := by refine submartingale_of_setIntegral_le_succ hadp hint fun i s hs => ?_ have : ∫ ω in s, f (i + 1) ω ∂μ = ∫ ω in s, (μ[f (i + 1)|𝒢 i]) ω ∂μ := - (setIntegral_condexp (𝒢.le i) (hint _) hs).symm + (setIntegral_condExp (𝒢.le i) (hint _) hs).symm rw [this] - exact setIntegral_mono_ae (hint i).integrableOn integrable_condexp.integrableOn (hf i) + exact setIntegral_mono_ae (hint i).integrableOn integrable_condExp.integrableOn (hf i) theorem supermartingale_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, μ[f (i + 1)|𝒢 i] ≤ᵐ[μ] f i) : Supermartingale f 𝒢 μ := by rw [← neg_neg f] refine (submartingale_nat hadp.neg (fun i => (hint i).neg) fun i => - EventuallyLE.trans ?_ (condexp_neg _).symm.le).neg + EventuallyLE.trans ?_ (condExp_neg ..).symm.le).neg filter_upwards [hf i] with x hx using neg_le_neg hx theorem martingale_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) @@ -409,30 +425,39 @@ theorem martingale_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Ada martingale_iff.2 ⟨supermartingale_nat hadp hint fun i => (hf i).symm.le, submartingale_nat hadp hint fun i => (hf i).le⟩ -theorem submartingale_of_condexp_sub_nonneg_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} +theorem submartingale_of_condExp_sub_nonneg_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, 0 ≤ᵐ[μ] μ[f (i + 1) - f i|𝒢 i]) : Submartingale f 𝒢 μ := by refine submartingale_nat hadp hint fun i => ?_ - rw [← condexp_of_stronglyMeasurable (𝒢.le _) (hadp _) (hint _), ← eventually_sub_nonneg] - exact EventuallyLE.trans (hf i) (condexp_sub (hint _) (hint _)).le + rw [← condExp_of_stronglyMeasurable (𝒢.le _) (hadp _) (hint _), ← eventually_sub_nonneg] + exact EventuallyLE.trans (hf i) (condExp_sub (hint _) (hint _) _).le + +@[deprecated (since := "2025-01-21")] +alias submartingale_of_condexp_sub_nonneg_nat := submartingale_of_condExp_sub_nonneg_nat -theorem supermartingale_of_condexp_sub_nonneg_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} +theorem supermartingale_of_condExp_sub_nonneg_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, 0 ≤ᵐ[μ] μ[f i - f (i + 1)|𝒢 i]) : Supermartingale f 𝒢 μ := by rw [← neg_neg f] - refine (submartingale_of_condexp_sub_nonneg_nat hadp.neg (fun i => (hint i).neg) ?_).neg + refine (submartingale_of_condExp_sub_nonneg_nat hadp.neg (fun i => (hint i).neg) ?_).neg simpa only [Pi.zero_apply, Pi.neg_apply, neg_sub_neg] -theorem martingale_of_condexp_sub_eq_zero_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} +@[deprecated (since := "2025-01-21")] +alias supermartingale_of_condexp_sub_nonneg_nat := supermartingale_of_condExp_sub_nonneg_nat + +theorem martingale_of_condExp_sub_eq_zero_nat [IsFiniteMeasure μ] {f : ℕ → Ω → ℝ} (hadp : Adapted 𝒢 f) (hint : ∀ i, Integrable (f i) μ) (hf : ∀ i, μ[f (i + 1) - f i|𝒢 i] =ᵐ[μ] 0) : Martingale f 𝒢 μ := by - refine martingale_iff.2 ⟨supermartingale_of_condexp_sub_nonneg_nat hadp hint fun i => ?_, - submartingale_of_condexp_sub_nonneg_nat hadp hint fun i => (hf i).symm.le⟩ + refine martingale_iff.2 ⟨supermartingale_of_condExp_sub_nonneg_nat hadp hint fun i => ?_, + submartingale_of_condExp_sub_nonneg_nat hadp hint fun i => (hf i).symm.le⟩ rw [← neg_sub] - refine (EventuallyEq.trans ?_ (condexp_neg _).symm).le + refine (EventuallyEq.trans ?_ (condExp_neg ..).symm).le filter_upwards [hf i] with x hx simpa only [Pi.zero_apply, Pi.neg_apply, zero_eq_neg] +@[deprecated (since := "2025-01-21")] +alias martingale_of_condexp_sub_eq_zero_nat := martingale_of_condExp_sub_eq_zero_nat + -- Note that one cannot use `Submartingale.zero_le_of_predictable` to prove the other two -- corresponding lemmas without imposing more restrictions to the ordering of `E` /-- A predictable submartingale is a.e. greater equal than its initial state. -/ @@ -442,7 +467,7 @@ theorem Submartingale.zero_le_of_predictable [Preorder E] [SigmaFiniteFiltration induction' n with k ih · rfl · exact ih.trans ((hfmgle.2.1 k (k + 1) k.le_succ).trans_eq <| Germ.coe_eq.mp <| - congr_arg Germ.ofFun <| condexp_of_stronglyMeasurable (𝒢.le _) (hfadp _) <| hfmgle.integrable _) + congr_arg Germ.ofFun <| condExp_of_stronglyMeasurable (𝒢.le _) (hfadp _) <| hfmgle.integrable _) /-- A predictable supermartingale is a.e. less equal than its initial state. -/ theorem Supermartingale.le_zero_of_predictable [Preorder E] [SigmaFiniteFiltration μ 𝒢] @@ -450,7 +475,7 @@ theorem Supermartingale.le_zero_of_predictable [Preorder E] [SigmaFiniteFiltrati (n : ℕ) : f n ≤ᵐ[μ] f 0 := by induction' n with k ih · rfl - · exact ((Germ.coe_eq.mp <| congr_arg Germ.ofFun <| condexp_of_stronglyMeasurable (𝒢.le _) + · exact ((Germ.coe_eq.mp <| congr_arg Germ.ofFun <| condExp_of_stronglyMeasurable (𝒢.le _) (hfadp _) <| hfmgle.integrable _).symm.trans_le (hfmgle.2.1 k (k + 1) k.le_succ)).trans ih /-- A predictable martingale is a.e. equal to its initial state. -/ @@ -458,7 +483,7 @@ theorem Martingale.eq_zero_of_predictable [SigmaFiniteFiltration μ 𝒢] {f : (hfmgle : Martingale f 𝒢 μ) (hfadp : Adapted 𝒢 fun n => f (n + 1)) (n : ℕ) : f n =ᵐ[μ] f 0 := by induction' n with k ih · rfl - · exact ((Germ.coe_eq.mp (congr_arg Germ.ofFun <| condexp_of_stronglyMeasurable (𝒢.le _) (hfadp _) + · exact ((Germ.coe_eq.mp (congr_arg Germ.ofFun <| condExp_of_stronglyMeasurable (𝒢.le _) (hfadp _) (hfmgle.integrable _))).symm.trans (hfmgle.2 k (k + 1) k.le_succ)).trans ih namespace Submartingale @@ -486,11 +511,11 @@ theorem Submartingale.sum_mul_sub [IsFiniteMeasure μ] {R : ℝ} {ξ f : ℕ → exact (hξ.stronglyMeasurable_le hi.le).mul ((hf.adapted.stronglyMeasurable_le (Nat.succ_le_of_lt hi)).sub (hf.adapted.stronglyMeasurable_le hi.le)) - refine submartingale_of_condexp_sub_nonneg_nat hadp hint fun i => ?_ + refine submartingale_of_condExp_sub_nonneg_nat hadp hint fun i => ?_ simp only [← Finset.sum_Ico_eq_sub _ (Nat.le_succ _), Finset.sum_apply, Pi.mul_apply, Pi.sub_apply, Nat.Ico_succ_singleton, Finset.sum_singleton] exact EventuallyLE.trans (EventuallyLE.mul_nonneg (Eventually.of_forall (hnonneg _)) - (hf.condexp_sub_nonneg (Nat.le_succ _))) (condexp_stronglyMeasurable_mul (hξ _) + (hf.condExp_sub_nonneg (Nat.le_succ _))) (condExp_mul_of_stronglyMeasurable_left (hξ _) (((hf.integrable _).sub (hf.integrable _)).bdd_mul hξ.stronglyMeasurable.aestronglyMeasurable (hξbdd _)) ((hf.integrable _).sub (hf.integrable _))).symm.le diff --git a/Mathlib/Probability/Martingale/BorelCantelli.lean b/Mathlib/Probability/Martingale/BorelCantelli.lean index 0f483a8596a0b..7edfbb36da6b5 100644 --- a/Mathlib/Probability/Martingale/BorelCantelli.lean +++ b/Mathlib/Probability/Martingale/BorelCantelli.lean @@ -3,9 +3,10 @@ Copyright (c) 2022 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ +import Mathlib.Algebra.Order.Archimedean.IndicatorCard +import Mathlib.Probability.Martingale.Centering import Mathlib.Probability.Martingale.Convergence import Mathlib.Probability.Martingale.OptionalStopping -import Mathlib.Probability.Martingale.Centering /-! @@ -320,7 +321,7 @@ theorem tendsto_sum_indicator_atTop_iff [IsFiniteMeasure μ] have h₂ := (martingale_martingalePart hf hint).ae_not_tendsto_atTop_atBot (martingalePart_bdd_difference ℱ hbdd) have h₃ : ∀ᵐ ω ∂μ, ∀ n, 0 ≤ (μ[f (n + 1) - f n|ℱ n]) ω := by - refine ae_all_iff.2 fun n => condexp_nonneg ?_ + refine ae_all_iff.2 fun n => condExp_nonneg ?_ filter_upwards [ae_all_iff.1 hfmono n] with ω hω using sub_nonneg.2 hω filter_upwards [h₁, h₂, h₃, hfmono] with ω hω₁ hω₂ hω₃ hω₄ constructor <;> intro ht @@ -361,7 +362,9 @@ everywhere equal to the set for which `∑ k, ℙ(s (k + 1) | ℱ k) = ∞`. -/ theorem ae_mem_limsup_atTop_iff (μ : Measure Ω) [IsFiniteMeasure μ] {s : ℕ → Set Ω} (hs : ∀ n, MeasurableSet[ℱ n] (s n)) : ∀ᵐ ω ∂μ, ω ∈ limsup s atTop ↔ Tendsto (fun n => ∑ k ∈ Finset.range n, - (μ[(s (k + 1)).indicator (1 : Ω → ℝ)|ℱ k]) ω) atTop atTop := - (limsup_eq_tendsto_sum_indicator_atTop ℝ s).symm ▸ tendsto_sum_indicator_atTop_iff' hs + (μ[(s (k + 1)).indicator (1 : Ω → ℝ)|ℱ k]) ω) atTop atTop := by + rw [← limsup_nat_add s 1, + Set.limsup_eq_tendsto_sum_indicator_atTop (zero_lt_one (α := ℝ)) (fun n ↦ s (n + 1))] + exact tendsto_sum_indicator_atTop_iff' hs end MeasureTheory diff --git a/Mathlib/Probability/Martingale/Centering.lean b/Mathlib/Probability/Martingale/Centering.lean index 2cfee94899d77..002c97c9f5f23 100644 --- a/Mathlib/Probability/Martingale/Centering.lean +++ b/Mathlib/Probability/Martingale/Centering.lean @@ -49,11 +49,11 @@ theorem predictablePart_zero : predictablePart f ℱ μ 0 = 0 := by theorem adapted_predictablePart : Adapted ℱ fun n => predictablePart f ℱ μ (n + 1) := fun _ => Finset.stronglyMeasurable_sum' _ fun _ hin => - stronglyMeasurable_condexp.mono (ℱ.mono (Finset.mem_range_succ_iff.mp hin)) + stronglyMeasurable_condExp.mono (ℱ.mono (Finset.mem_range_succ_iff.mp hin)) theorem adapted_predictablePart' : Adapted ℱ fun n => predictablePart f ℱ μ n := fun _ => Finset.stronglyMeasurable_sum' _ fun _ hin => - stronglyMeasurable_condexp.mono (ℱ.mono (Finset.mem_range_le hin)) + stronglyMeasurable_condExp.mono (ℱ.mono (Finset.mem_range_le hin)) /-- Any `ℕ`-indexed stochastic process can be written as the sum of a martingale and a predictable process. This is the martingale. See `predictablePart` for the predictable process. -/ @@ -76,8 +76,7 @@ theorem adapted_martingalePart (hf : Adapted ℱ f) : Adapted ℱ (martingalePar theorem integrable_martingalePart (hf_int : ∀ n, Integrable (f n) μ) (n : ℕ) : Integrable (martingalePart f ℱ μ n) μ := by rw [martingalePart_eq_sum] - exact (hf_int 0).add - (integrable_finset_sum' _ fun i _ => ((hf_int _).sub (hf_int _)).sub integrable_condexp) + fun_prop theorem martingale_martingalePart (hf : Adapted ℱ f) (hf_int : ∀ n, Integrable (f n) μ) [SigmaFiniteFiltration μ ℱ] : Martingale (martingalePart f ℱ μ) ℱ μ := by @@ -86,30 +85,28 @@ theorem martingale_martingalePart (hf : Adapted ℱ f) (hf_int : ∀ n, Integrab have h_eq_sum : μ[martingalePart f ℱ μ j|ℱ i] =ᵐ[μ] f 0 + ∑ k ∈ Finset.range j, (μ[f (k + 1) - f k|ℱ i] - μ[μ[f (k + 1) - f k|ℱ k]|ℱ i]) := by rw [martingalePart_eq_sum] - refine (condexp_add (hf_int 0) ?_).trans ?_ - · exact integrable_finset_sum' _ fun i _ => ((hf_int _).sub (hf_int _)).sub integrable_condexp - refine (EventuallyEq.add EventuallyEq.rfl (condexp_finset_sum fun i _ => ?_)).trans ?_ - · exact ((hf_int _).sub (hf_int _)).sub integrable_condexp + refine (condExp_add (hf_int 0) (by fun_prop) _).trans ?_ + refine (EventuallyEq.rfl.add (condExp_finset_sum (fun i _ => by fun_prop) _)).trans ?_ refine EventuallyEq.add ?_ ?_ - · rw [condexp_of_stronglyMeasurable (ℱ.le _) _ (hf_int 0)] + · rw [condExp_of_stronglyMeasurable (ℱ.le _) _ (hf_int 0)] · exact (hf 0).mono (ℱ.mono (zero_le i)) - · exact eventuallyEq_sum fun k _ => condexp_sub ((hf_int _).sub (hf_int _)) integrable_condexp + · exact eventuallyEq_sum fun k _ => condExp_sub (by fun_prop) integrable_condExp _ refine h_eq_sum.trans ?_ have h_ge : ∀ k, i ≤ k → μ[f (k + 1) - f k|ℱ i] - μ[μ[f (k + 1) - f k|ℱ k]|ℱ i] =ᵐ[μ] 0 := by intro k hk have : μ[μ[f (k + 1) - f k|ℱ k]|ℱ i] =ᵐ[μ] μ[f (k + 1) - f k|ℱ i] := - condexp_condexp_of_le (ℱ.mono hk) (ℱ.le k) + condExp_condExp_of_le (ℱ.mono hk) (ℱ.le k) filter_upwards [this] with x hx rw [Pi.sub_apply, Pi.zero_apply, hx, sub_self] have h_lt : ∀ k, k < i → μ[f (k + 1) - f k|ℱ i] - μ[μ[f (k + 1) - f k|ℱ k]|ℱ i] =ᵐ[μ] f (k + 1) - f k - μ[f (k + 1) - f k|ℱ k] := by refine fun k hk => EventuallyEq.sub ?_ ?_ - · rw [condexp_of_stronglyMeasurable] + · rw [condExp_of_stronglyMeasurable] · exact ((hf (k + 1)).mono (ℱ.mono (Nat.succ_le_of_lt hk))).sub ((hf k).mono (ℱ.mono hk.le)) · exact (hf_int _).sub (hf_int _) - · rw [condexp_of_stronglyMeasurable] - · exact stronglyMeasurable_condexp.mono (ℱ.mono hk.le) - · exact integrable_condexp + · rw [condExp_of_stronglyMeasurable] + · exact stronglyMeasurable_condExp.mono (ℱ.mono hk.le) + · exact integrable_condExp rw [martingalePart_eq_sum] refine EventuallyEq.add EventuallyEq.rfl ?_ rw [← Finset.sum_range_add_sum_Ico _ hij, ← @@ -154,7 +151,7 @@ theorem predictablePart_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} ( (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : ∀ᵐ ω ∂μ, ∀ i, |predictablePart f ℱ μ (i + 1) ω - predictablePart f ℱ μ i ω| ≤ R := by simp_rw [predictablePart, Finset.sum_apply, Finset.sum_range_succ_sub_sum] - exact ae_all_iff.2 fun i => ae_bdd_condexp_of_ae_bdd <| ae_all_iff.1 hbdd i + exact ae_all_iff.2 fun i => ae_bdd_condExp_of_ae_bdd <| ae_all_iff.1 hbdd i theorem martingalePart_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} (ℱ : Filtration ℕ m0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 11b1183ad6b5c..78e152f66e770 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -29,13 +29,13 @@ theorems. convergence theorem: a uniformly integrable submartingale adapted to the filtration `ℱ` converges almost everywhere and in L¹ to an integrable function which is measurable with respect to the σ-algebra `⨆ n, ℱ n`. -* `MeasureTheory.Martingale.ae_eq_condexp_limitProcess`: part b the L¹ martingale convergence +* `MeasureTheory.Martingale.ae_eq_condExp_limitProcess`: part b the L¹ martingale convergence theorem: if `f` is a uniformly integrable martingale adapted to the filtration `ℱ`, then `f n` equals `𝔼[g | ℱ n]` almost everywhere where `g` is the limiting process of `f`. -* `MeasureTheory.Integrable.tendsto_ae_condexp`: part c the L¹ martingale convergence theorem: +* `MeasureTheory.Integrable.tendsto_ae_condExp`: part c the L¹ martingale convergence theorem: given a `⨆ n, ℱ n`-measurable function `g` where `ℱ` is a filtration, `𝔼[g | ℱ n]` converges almost everywhere to `g`. -* `MeasureTheory.Integrable.tendsto_eLpNorm_condexp`: part c the L¹ martingale convergence theorem: +* `MeasureTheory.Integrable.tendsto_eLpNorm_condExp`: part c the L¹ martingale convergence theorem: given a `⨆ n, ℱ n`-measurable function `g` where `ℱ` is a filtration, `𝔼[g | ℱ n]` converges in L¹ to `g`. @@ -215,7 +215,7 @@ theorem Submartingale.ae_tendsto_limitProcess [IsFiniteMeasure μ] (hf : Submart filter_upwards [hf.exists_ae_trim_tendsto_of_bdd hbdd] with ω hω simp_rw [g', dif_pos hω] exact hω.choose_spec - have hg'm : @AEStronglyMeasurable _ _ _ (⨆ n, ℱ n) g' (μ.trim hle) := + have hg'm : AEStronglyMeasurable[⨆ n, ℱ n] g' (μ.trim hle) := (@aemeasurable_of_tendsto_metrizable_ae' _ _ (⨆ n, ℱ n) _ _ _ _ _ _ _ (fun n => ((hf.stronglyMeasurable n).measurable.mono (le_sSup ⟨n, rfl⟩ : ℱ n ≤ ⨆ n, ℱ n) le_rfl).aemeasurable) hg').aestronglyMeasurable @@ -325,62 +325,71 @@ theorem Submartingale.ae_tendsto_limitProcess_of_uniformIntegrable (hf : Submart /-- If a martingale `f` adapted to `ℱ` converges in L¹ to `g`, then for all `n`, `f n` is almost everywhere equal to `𝔼[g | ℱ n]`. -/ -theorem Martingale.eq_condexp_of_tendsto_eLpNorm {μ : Measure Ω} (hf : Martingale f ℱ μ) +theorem Martingale.eq_condExp_of_tendsto_eLpNorm {μ : Measure Ω} (hf : Martingale f ℱ μ) (hg : Integrable g μ) (hgtends : Tendsto (fun n => eLpNorm (f n - g) 1 μ) atTop (𝓝 0)) (n : ℕ) : f n =ᵐ[μ] μ[g|ℱ n] := by rw [← sub_ae_eq_zero, ← eLpNorm_eq_zero_iff (((hf.stronglyMeasurable n).mono (ℱ.le _)).sub - (stronglyMeasurable_condexp.mono (ℱ.le _))).aestronglyMeasurable one_ne_zero] + (stronglyMeasurable_condExp.mono (ℱ.le _))).aestronglyMeasurable one_ne_zero] have ht : Tendsto (fun m => eLpNorm (μ[f m - g|ℱ n]) 1 μ) atTop (𝓝 0) := haveI hint : ∀ m, Integrable (f m - g) μ := fun m => (hf.integrable m).sub hg tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds hgtends (fun m => zero_le _) - fun m => eLpNorm_one_condexp_le_eLpNorm _ + fun m => eLpNorm_one_condExp_le_eLpNorm _ have hev : ∀ m ≥ n, eLpNorm (μ[f m - g|ℱ n]) 1 μ = eLpNorm (f n - μ[g|ℱ n]) 1 μ := by - refine fun m hm => eLpNorm_congr_ae ((condexp_sub (hf.integrable m) hg).trans ?_) + refine fun m hm => eLpNorm_congr_ae ((condExp_sub (hf.integrable m) hg _).trans ?_) filter_upwards [hf.2 n m hm] with x hx simp only [hx, Pi.sub_apply] exact tendsto_nhds_unique (tendsto_atTop_of_eventually_const hev) ht +@[deprecated (since := "2025-01-21")] +alias Martingale.eq_condexp_of_tendsto_eLpNorm := Martingale.eq_condExp_of_tendsto_eLpNorm + @[deprecated (since := "2024-07-27")] -alias Martingale.eq_condexp_of_tendsto_snorm := Martingale.eq_condexp_of_tendsto_eLpNorm +alias Martingale.eq_condExp_of_tendsto_snorm := Martingale.eq_condExp_of_tendsto_eLpNorm + +@[deprecated (since := "2025-01-21")] +alias Martingale.eq_condexp_of_tendsto_snorm := Martingale.eq_condExp_of_tendsto_snorm /-- Part b of the **L¹ martingale convergence theorem**: if `f` is a uniformly integrable martingale adapted to the filtration `ℱ`, then for all `n`, `f n` is almost everywhere equal to the conditional expectation of its limiting process wrt. `ℱ n`. -/ -theorem Martingale.ae_eq_condexp_limitProcess (hf : Martingale f ℱ μ) +theorem Martingale.ae_eq_condExp_limitProcess (hf : Martingale f ℱ μ) (hbdd : UniformIntegrable f 1 μ) (n : ℕ) : f n =ᵐ[μ] μ[ℱ.limitProcess f μ|ℱ n] := let ⟨_, hR⟩ := hbdd.2.2 - hf.eq_condexp_of_tendsto_eLpNorm ((memℒp_limitProcess_of_eLpNorm_bdd hbdd.1 hR).integrable le_rfl) + hf.eq_condExp_of_tendsto_eLpNorm ((memℒp_limitProcess_of_eLpNorm_bdd hbdd.1 hR).integrable le_rfl) (hf.submartingale.tendsto_eLpNorm_one_limitProcess hbdd) n +@[deprecated (since := "2025-01-21")] +alias Martingale.ae_eq_condexp_limitProcess := Martingale.ae_eq_condExp_limitProcess + /-- Part c of the **L¹ martingale convergence theorem**: Given an integrable function `g` which is measurable with respect to `⨆ n, ℱ n` where `ℱ` is a filtration, the martingale defined by `𝔼[g | ℱ n]` converges almost everywhere to `g`. This martingale also converges to `g` in L¹ and this result is provided by -`MeasureTheory.Integrable.tendsto_eLpNorm_condexp` -/ -theorem Integrable.tendsto_ae_condexp (hg : Integrable g μ) +`MeasureTheory.Integrable.tendsto_eLpNorm_condExp` -/ +theorem Integrable.tendsto_ae_condExp (hg : Integrable g μ) (hgmeas : StronglyMeasurable[⨆ n, ℱ n] g) : ∀ᵐ x ∂μ, Tendsto (fun n => (μ[g|ℱ n]) x) atTop (𝓝 (g x)) := by have hle : ⨆ n, ℱ n ≤ m0 := sSup_le fun m ⟨n, hn⟩ => hn ▸ ℱ.le _ have hunif : UniformIntegrable (fun n => μ[g|ℱ n]) 1 μ := - hg.uniformIntegrable_condexp_filtration + hg.uniformIntegrable_condExp_filtration obtain ⟨R, hR⟩ := hunif.2.2 have hlimint : Integrable (ℱ.limitProcess (fun n => μ[g|ℱ n]) μ) μ := (memℒp_limitProcess_of_eLpNorm_bdd hunif.1 hR).integrable le_rfl suffices g =ᵐ[μ] ℱ.limitProcess (fun n x => (μ[g|ℱ n]) x) μ by - filter_upwards [this, (martingale_condexp g ℱ μ).submartingale.ae_tendsto_limitProcess hR] with + filter_upwards [this, (martingale_condExp g ℱ μ).submartingale.ae_tendsto_limitProcess hR] with x heq ht rwa [heq] have : ∀ n s, MeasurableSet[ℱ n] s → ∫ x in s, g x ∂μ = ∫ x in s, ℱ.limitProcess (fun n x => (μ[g|ℱ n]) x) μ x ∂μ := by intro n s hs - rw [← setIntegral_condexp (ℱ.le n) hg hs, ← setIntegral_condexp (ℱ.le n) hlimint hs] + rw [← setIntegral_condExp (ℱ.le n) hg hs, ← setIntegral_condExp (ℱ.le n) hlimint hs] refine setIntegral_congr_ae (ℱ.le _ _ hs) ?_ - filter_upwards [(martingale_condexp g ℱ μ).ae_eq_condexp_limitProcess hunif n] with x hx _ + filter_upwards [(martingale_condExp g ℱ μ).ae_eq_condExp_limitProcess hunif n] with x hx _ rw [hx] refine ae_eq_of_forall_setIntegral_eq_of_sigmaFinite' hle (fun s _ _ => hg.integrableOn) - (fun s _ _ => hlimint.integrableOn) (fun s hs _ => ?_) hgmeas.aeStronglyMeasurable' - stronglyMeasurable_limitProcess.aeStronglyMeasurable' + (fun s _ _ => hlimint.integrableOn) (fun s hs _ => ?_) hgmeas.aestronglyMeasurable + stronglyMeasurable_limitProcess.aestronglyMeasurable have hpi : IsPiSystem {s | ∃ n, MeasurableSet[ℱ n] s} := by rw [Set.setOf_exists] exact isPiSystem_iUnion_of_monotone _ (fun n ↦ (ℱ n).isPiSystem_measurableSet) fun _ _ ↦ ℱ.mono @@ -406,51 +415,66 @@ theorem Integrable.tendsto_ae_condexp (hg : Integrable g μ) integral_iUnion (fun n => hle _ (hfmeas n)) hf hlimint.integrableOn] exact tsum_congr fun n => heq _ (measure_lt_top _ _) +@[deprecated (since := "2025-01-21")] +alias Integrable.tendsto_ae_condexp := Integrable.tendsto_ae_condExp + /-- Part c of the **L¹ martingale convergence theorem**: Given an integrable function `g` which is measurable with respect to `⨆ n, ℱ n` where `ℱ` is a filtration, the martingale defined by `𝔼[g | ℱ n]` converges in L¹ to `g`. This martingale also converges to `g` almost everywhere and this result is provided by -`MeasureTheory.Integrable.tendsto_ae_condexp` -/ -theorem Integrable.tendsto_eLpNorm_condexp (hg : Integrable g μ) +`MeasureTheory.Integrable.tendsto_ae_condExp` -/ +theorem Integrable.tendsto_eLpNorm_condExp (hg : Integrable g μ) (hgmeas : StronglyMeasurable[⨆ n, ℱ n] g) : Tendsto (fun n => eLpNorm (μ[g|ℱ n] - g) 1 μ) atTop (𝓝 0) := tendsto_Lp_finite_of_tendstoInMeasure le_rfl ENNReal.one_ne_top - (fun n => (stronglyMeasurable_condexp.mono (ℱ.le n)).aestronglyMeasurable) - (memℒp_one_iff_integrable.2 hg) hg.uniformIntegrable_condexp_filtration.2.1 + (fun n => (stronglyMeasurable_condExp.mono (ℱ.le n)).aestronglyMeasurable) + (memℒp_one_iff_integrable.2 hg) hg.uniformIntegrable_condExp_filtration.2.1 (tendstoInMeasure_of_tendsto_ae - (fun n => (stronglyMeasurable_condexp.mono (ℱ.le n)).aestronglyMeasurable) - (hg.tendsto_ae_condexp hgmeas)) + (fun n => (stronglyMeasurable_condExp.mono (ℱ.le n)).aestronglyMeasurable) + (hg.tendsto_ae_condExp hgmeas)) + +@[deprecated (since := "2025-01-21")] +alias Integrable.tendsto_eLpNorm_condexp := Integrable.tendsto_eLpNorm_condExp @[deprecated (since := "2024-07-27")] -alias Integrable.tendsto_snorm_condexp := Integrable.tendsto_eLpNorm_condexp +alias Integrable.tendsto_snorm_condExp := Integrable.tendsto_eLpNorm_condExp + +@[deprecated (since := "2025-01-21")] +alias Integrable.tendsto_snorm_condexp := Integrable.tendsto_snorm_condExp /-- **Lévy's upward theorem**, almost everywhere version: given a function `g` and a filtration `ℱ`, the sequence defined by `𝔼[g | ℱ n]` converges almost everywhere to `𝔼[g | ⨆ n, ℱ n]`. -/ -theorem tendsto_ae_condexp (g : Ω → ℝ) : +theorem tendsto_ae_condExp (g : Ω → ℝ) : ∀ᵐ x ∂μ, Tendsto (fun n => (μ[g|ℱ n]) x) atTop (𝓝 ((μ[g|⨆ n, ℱ n]) x)) := by have ht : ∀ᵐ x ∂μ, Tendsto (fun n => (μ[μ[g|⨆ n, ℱ n]|ℱ n]) x) atTop (𝓝 ((μ[g|⨆ n, ℱ n]) x)) := - integrable_condexp.tendsto_ae_condexp stronglyMeasurable_condexp + integrable_condExp.tendsto_ae_condExp stronglyMeasurable_condExp have heq : ∀ n, ∀ᵐ x ∂μ, (μ[μ[g|⨆ n, ℱ n]|ℱ n]) x = (μ[g|ℱ n]) x := fun n => - condexp_condexp_of_le (le_iSup _ n) (iSup_le fun n => ℱ.le n) + condExp_condExp_of_le (le_iSup _ n) (iSup_le fun n => ℱ.le n) rw [← ae_all_iff] at heq filter_upwards [heq, ht] with x hxeq hxt exact hxt.congr hxeq +@[deprecated (since := "2025-01-21")] alias tendsto_ae_condexp := tendsto_ae_condExp + /-- **Lévy's upward theorem**, L¹ version: given a function `g` and a filtration `ℱ`, the sequence defined by `𝔼[g | ℱ n]` converges in L¹ to `𝔼[g | ⨆ n, ℱ n]`. -/ -theorem tendsto_eLpNorm_condexp (g : Ω → ℝ) : +theorem tendsto_eLpNorm_condExp (g : Ω → ℝ) : Tendsto (fun n => eLpNorm (μ[g|ℱ n] - μ[g|⨆ n, ℱ n]) 1 μ) atTop (𝓝 0) := by have ht : Tendsto (fun n => eLpNorm (μ[μ[g|⨆ n, ℱ n]|ℱ n] - μ[g|⨆ n, ℱ n]) 1 μ) atTop (𝓝 0) := - integrable_condexp.tendsto_eLpNorm_condexp stronglyMeasurable_condexp + integrable_condExp.tendsto_eLpNorm_condExp stronglyMeasurable_condExp have heq : ∀ n, ∀ᵐ x ∂μ, (μ[μ[g|⨆ n, ℱ n]|ℱ n]) x = (μ[g|ℱ n]) x := fun n => - condexp_condexp_of_le (le_iSup _ n) (iSup_le fun n => ℱ.le n) + condExp_condExp_of_le (le_iSup _ n) (iSup_le fun n => ℱ.le n) refine ht.congr fun n => eLpNorm_congr_ae ?_ filter_upwards [heq n] with x hxeq simp only [hxeq, Pi.sub_apply] +@[deprecated (since := "2025-01-21")] alias tendsto_eLpNorm_condexp := tendsto_eLpNorm_condExp + @[deprecated (since := "2024-07-27")] -alias tendsto_snorm_condexp := tendsto_eLpNorm_condexp +alias tendsto_snorm_condExp := tendsto_eLpNorm_condExp + +@[deprecated (since := "2025-01-21")] alias tendsto_snorm_condexp := tendsto_snorm_condExp end L1Convergence diff --git a/Mathlib/Probability/Martingale/OptionalSampling.lean b/Mathlib/Probability/Martingale/OptionalSampling.lean index 90b9489aaa39c..bf95b0773b2a2 100644 --- a/Mathlib/Probability/Martingale/OptionalSampling.lean +++ b/Mathlib/Probability/Martingale/OptionalSampling.lean @@ -15,13 +15,13 @@ If `τ` is a bounded stopping time and `σ` is another stopping time, then the v ## Main results -* `stoppedValue_ae_eq_condexp_of_le_const`: the value of a martingale `f` at a stopping time `τ` +* `stoppedValue_ae_eq_condExp_of_le_const`: the value of a martingale `f` at a stopping time `τ` bounded by `n` is the conditional expectation of `f n` with respect to the σ-algebra generated by `τ`. -* `stoppedValue_ae_eq_condexp_of_le`: if `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is +* `stoppedValue_ae_eq_condExp_of_le`: if `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is bounded, then the value of a martingale `f` at `σ` is the conditional expectation of its value at `τ` with respect to the σ-algebra generated by `σ`. -* `stoppedValue_min_ae_eq_condexp`: the optional sampling theorem. If `τ` is a bounded stopping +* `stoppedValue_min_ae_eq_condExp`: the optional sampling theorem. If `τ` is a bounded stopping time and `σ` is another stopping time, then the value of a martingale `f` at the stopping time `min τ σ` is almost everywhere equal to the conditional expectation of `f` stopped at `τ` with respect to the σ-algebra generated by `σ`. @@ -46,22 +46,25 @@ variable {ι : Type*} [LinearOrder ι] [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] {ℱ : Filtration ι m} [SigmaFiniteFiltration μ ℱ] {τ σ : Ω → ι} {f : ι → Ω → E} {i n : ι} -theorem condexp_stopping_time_ae_eq_restrict_eq_const +theorem condExp_stopping_time_ae_eq_restrict_eq_const [(Filter.atTop : Filter ι).IsCountablyGenerated] (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) [SigmaFinite (μ.trim hτ.measurableSpace_le)] (hin : i ≤ n) : μ[f n|hτ.measurableSpace] =ᵐ[μ.restrict {x | τ x = i}] f i := by - refine Filter.EventuallyEq.trans ?_ (ae_restrict_of_ae (h.condexp_ae_eq hin)) - refine condexp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le (ℱ.le i) + refine Filter.EventuallyEq.trans ?_ (ae_restrict_of_ae (h.condExp_ae_eq hin)) + refine condExp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le (ℱ.le i) (hτ.measurableSet_eq' i) fun t => ?_ rw [Set.inter_comm _ t, IsStoppingTime.measurableSet_inter_eq_iff] -theorem condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const (h : Martingale f ℱ μ) +@[deprecated (since := "2025-01-21")] +alias condexp_stopping_time_ae_eq_restrict_eq_const := condExp_stopping_time_ae_eq_restrict_eq_const + +theorem condExp_stopping_time_ae_eq_restrict_eq_const_of_le_const (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hτ_le : ∀ x, τ x ≤ n) [SigmaFinite (μ.trim (hτ.measurableSpace_le_of_le hτ_le))] (i : ι) : μ[f n|hτ.measurableSpace] =ᵐ[μ.restrict {x | τ x = i}] f i := by by_cases hin : i ≤ n - · refine Filter.EventuallyEq.trans ?_ (ae_restrict_of_ae (h.condexp_ae_eq hin)) - refine condexp_ae_eq_restrict_of_measurableSpace_eq_on (hτ.measurableSpace_le_of_le hτ_le) + · refine Filter.EventuallyEq.trans ?_ (ae_restrict_of_ae (h.condExp_ae_eq hin)) + refine condExp_ae_eq_restrict_of_measurableSpace_eq_on (hτ.measurableSpace_le_of_le hτ_le) (ℱ.le i) (hτ.measurableSet_eq' i) fun t => ?_ rw [Set.inter_comm _ t, IsStoppingTime.measurableSet_inter_eq_iff] · suffices {x : Ω | τ x = i} = ∅ by simp [this]; norm_cast @@ -70,11 +73,15 @@ theorem condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const (h : Martingal rintro rfl exact hin (hτ_le x) +@[deprecated (since := "2025-01-21")] +alias condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const := + condExp_stopping_time_ae_eq_restrict_eq_const_of_le_const + theorem stoppedValue_ae_eq_restrict_eq (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hτ_le : ∀ x, τ x ≤ n) [SigmaFinite (μ.trim (hτ.measurableSpace_le_of_le hτ_le))] (i : ι) : stoppedValue f τ =ᵐ[μ.restrict {x | τ x = i}] μ[f n|hτ.measurableSpace] := by refine Filter.EventuallyEq.trans ?_ - (condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const h hτ hτ_le i).symm + (condExp_stopping_time_ae_eq_restrict_eq_const_of_le_const h hτ hτ_le i).symm rw [Filter.EventuallyEq, ae_restrict_iff' (ℱ.le _ _ (hτ.measurableSet_eq i))] refine Filter.Eventually.of_forall fun x hx => ?_ rw [Set.mem_setOf_eq] at hx @@ -82,7 +89,7 @@ theorem stoppedValue_ae_eq_restrict_eq (h : Martingale f ℱ μ) (hτ : IsStoppi /-- The value of a martingale `f` at a stopping time `τ` bounded by `n` is the conditional expectation of `f n` with respect to the σ-algebra generated by `τ`. -/ -theorem stoppedValue_ae_eq_condexp_of_le_const_of_countable_range (h : Martingale f ℱ μ) +theorem stoppedValue_ae_eq_condExp_of_le_const_of_countable_range (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hτ_le : ∀ x, τ x ≤ n) (h_countable_range : (Set.range τ).Countable) [SigmaFinite (μ.trim (hτ.measurableSpace_le_of_le hτ_le))] : stoppedValue f τ =ᵐ[μ] μ[f n|hτ.measurableSpace] := by @@ -94,18 +101,25 @@ theorem stoppedValue_ae_eq_condexp_of_le_const_of_countable_range (h : Martingal rw [this, ae_eq_restrict_biUnion_iff _ h_countable_range] exact fun i _ => stoppedValue_ae_eq_restrict_eq h _ hτ_le i +@[deprecated (since := "2025-01-21")] +alias stoppedValue_ae_eq_condexp_of_le_const_of_countable_range := + stoppedValue_ae_eq_condExp_of_le_const_of_countable_range + /-- The value of a martingale `f` at a stopping time `τ` bounded by `n` is the conditional expectation of `f n` with respect to the σ-algebra generated by `τ`. -/ -theorem stoppedValue_ae_eq_condexp_of_le_const [Countable ι] (h : Martingale f ℱ μ) +theorem stoppedValue_ae_eq_condExp_of_le_const [Countable ι] (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hτ_le : ∀ x, τ x ≤ n) [SigmaFinite (μ.trim (hτ.measurableSpace_le_of_le hτ_le))] : stoppedValue f τ =ᵐ[μ] μ[f n|hτ.measurableSpace] := - h.stoppedValue_ae_eq_condexp_of_le_const_of_countable_range hτ hτ_le (Set.to_countable _) + h.stoppedValue_ae_eq_condExp_of_le_const_of_countable_range hτ hτ_le (Set.to_countable _) + +@[deprecated (since := "2025-01-21")] +alias stoppedValue_ae_eq_condexp_of_le_const := stoppedValue_ae_eq_condExp_of_le_const /-- If `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is bounded, then the value of a martingale `f` at `σ` is the conditional expectation of its value at `τ` with respect to the σ-algebra generated by `σ`. -/ -theorem stoppedValue_ae_eq_condexp_of_le_of_countable_range (h : Martingale f ℱ μ) +theorem stoppedValue_ae_eq_condExp_of_le_of_countable_range (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) (hσ_le_τ : σ ≤ τ) (hτ_le : ∀ x, τ x ≤ n) (hτ_countable_range : (Set.range τ).Countable) (hσ_countable_range : (Set.range σ).Countable) [SigmaFinite (μ.trim (hσ.measurableSpace_le_of_le fun x => (hσ_le_τ x).trans (hτ_le x)))] : @@ -113,24 +127,31 @@ theorem stoppedValue_ae_eq_condexp_of_le_of_countable_range (h : Martingale f have : SigmaFinite (μ.trim (hτ.measurableSpace_le_of_le hτ_le)) := sigmaFiniteTrim_mono _ (IsStoppingTime.measurableSpace_mono hσ hτ hσ_le_τ) have : μ[stoppedValue f τ|hσ.measurableSpace] =ᵐ[μ] - μ[μ[f n|hτ.measurableSpace]|hσ.measurableSpace] := condexp_congr_ae - (h.stoppedValue_ae_eq_condexp_of_le_const_of_countable_range hτ hτ_le hτ_countable_range) + μ[μ[f n|hτ.measurableSpace]|hσ.measurableSpace] := condExp_congr_ae + (h.stoppedValue_ae_eq_condExp_of_le_const_of_countable_range hτ hτ_le hτ_countable_range) refine (Filter.EventuallyEq.trans ?_ - (condexp_condexp_of_le ?_ (hτ.measurableSpace_le_of_le hτ_le)).symm).trans this.symm - · exact h.stoppedValue_ae_eq_condexp_of_le_const_of_countable_range hσ + (condExp_condExp_of_le ?_ (hτ.measurableSpace_le_of_le hτ_le)).symm).trans this.symm + · exact h.stoppedValue_ae_eq_condExp_of_le_const_of_countable_range hσ (fun x => (hσ_le_τ x).trans (hτ_le x)) hσ_countable_range · exact hσ.measurableSpace_mono hτ hσ_le_τ +@[deprecated (since := "2025-01-21")] +alias stoppedValue_ae_eq_condexp_of_le_of_countable_range := + stoppedValue_ae_eq_condExp_of_le_of_countable_range + /-- If `τ` and `σ` are two stopping times with `σ ≤ τ` and `τ` is bounded, then the value of a martingale `f` at `σ` is the conditional expectation of its value at `τ` with respect to the σ-algebra generated by `σ`. -/ -theorem stoppedValue_ae_eq_condexp_of_le [Countable ι] (h : Martingale f ℱ μ) +theorem stoppedValue_ae_eq_condExp_of_le [Countable ι] (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) (hσ_le_τ : σ ≤ τ) (hτ_le : ∀ x, τ x ≤ n) [SigmaFinite (μ.trim hσ.measurableSpace_le)] : stoppedValue f σ =ᵐ[μ] μ[stoppedValue f τ|hσ.measurableSpace] := - h.stoppedValue_ae_eq_condexp_of_le_of_countable_range hτ hσ hσ_le_τ hτ_le (Set.to_countable _) + h.stoppedValue_ae_eq_condExp_of_le_of_countable_range hτ hσ hσ_le_τ hτ_le (Set.to_countable _) (Set.to_countable _) +@[deprecated (since := "2025-01-21")] +alias stoppedValue_ae_eq_condexp_of_le := stoppedValue_ae_eq_condExp_of_le + end FirstCountableTopology section SubsetOfNat @@ -145,21 +166,21 @@ variable {ι : Type*} [LinearOrder ι] [LocallyFiniteOrder ι] [OrderBot ι] [To [DiscreteTopology ι] [MeasurableSpace ι] [BorelSpace ι] [MeasurableSpace E] [BorelSpace E] [SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i : ι} -theorem condexp_stoppedValue_stopping_time_ae_eq_restrict_le (h : Martingale f ℱ μ) +theorem condExp_stoppedValue_stopping_time_ae_eq_restrict_le (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) [SigmaFinite (μ.trim hσ.measurableSpace_le)] (hτ_le : ∀ x, τ x ≤ i) : μ[stoppedValue f τ|hσ.measurableSpace] =ᵐ[μ.restrict {x : Ω | τ x ≤ σ x}] stoppedValue f τ := by rw [ae_eq_restrict_iff_indicator_ae_eq (hτ.measurableSpace_le _ (hτ.measurableSet_le_stopping_time hσ))] - refine (condexp_indicator (integrable_stoppedValue ι hτ h.integrable hτ_le) + refine (condExp_indicator (integrable_stoppedValue ι hτ h.integrable hτ_le) (hτ.measurableSet_stopping_time_le hσ)).symm.trans ?_ have h_int : Integrable ({ω : Ω | τ ω ≤ σ ω}.indicator (stoppedValue (fun n : ι => f n) τ)) μ := by refine (integrable_stoppedValue ι hτ h.integrable hτ_le).indicator ?_ exact hτ.measurableSpace_le _ (hτ.measurableSet_le_stopping_time hσ) - have h_meas : AEStronglyMeasurable' hσ.measurableSpace + have h_meas : AEStronglyMeasurable[hσ.measurableSpace] ({ω : Ω | τ ω ≤ σ ω}.indicator (stoppedValue (fun n : ι => f n) τ)) μ := by - refine StronglyMeasurable.aeStronglyMeasurable' ?_ + refine StronglyMeasurable.aestronglyMeasurable ?_ refine StronglyMeasurable.stronglyMeasurable_of_measurableSpace_le_on (hτ.measurableSet_le_stopping_time hσ) ?_ ?_ ?_ · intro t ht @@ -171,20 +192,24 @@ theorem condexp_stoppedValue_stopping_time_ae_eq_restrict_le (h : Martingale f exact measurable_stoppedValue h.adapted.progMeasurable_of_discrete hτ · intro x hx simp only [hx, Set.indicator_of_not_mem, not_false_iff] - exact condexp_of_aestronglyMeasurable' hσ.measurableSpace_le h_meas h_int + exact condExp_of_aestronglyMeasurable' hσ.measurableSpace_le h_meas h_int + +@[deprecated (since := "2025-01-21")] +alias condexp_stoppedValue_stopping_time_ae_eq_restrict_le := + condExp_stoppedValue_stopping_time_ae_eq_restrict_le /-- **Optional Sampling theorem**. If `τ` is a bounded stopping time and `σ` is another stopping time, then the value of a martingale `f` at the stopping time `min τ σ` is almost everywhere equal to the conditional expectation of `f` stopped at `τ` with respect to the σ-algebra generated by `σ`. -/ -theorem stoppedValue_min_ae_eq_condexp [SigmaFiniteFiltration μ ℱ] (h : Martingale f ℱ μ) +theorem stoppedValue_min_ae_eq_condExp [SigmaFiniteFiltration μ ℱ] (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) {n : ι} (hτ_le : ∀ x, τ x ≤ n) [h_sf_min : SigmaFinite (μ.trim (hτ.min hσ).measurableSpace_le)] : (stoppedValue f fun x => min (σ x) (τ x)) =ᵐ[μ] μ[stoppedValue f τ|hσ.measurableSpace] := by refine - (h.stoppedValue_ae_eq_condexp_of_le hτ (hσ.min hτ) (fun x => min_le_right _ _) hτ_le).trans ?_ + (h.stoppedValue_ae_eq_condExp_of_le hτ (hσ.min hτ) (fun x => min_le_right _ _) hτ_le).trans ?_ refine ae_of_ae_restrict_of_ae_restrict_compl {x | σ x ≤ τ x} ?_ ?_ - · exact condexp_min_stopping_time_ae_eq_restrict_le hσ hτ + · exact condExp_min_stopping_time_ae_eq_restrict_le hσ hτ · suffices μ[stoppedValue f τ|(hσ.min hτ).measurableSpace] =ᵐ[μ.restrict {x | τ x ≤ σ x}] μ[stoppedValue f τ|hσ.measurableSpace] by rw [ae_restrict_iff' (hσ.measurableSpace_le _ (hσ.measurableSet_le_stopping_time hτ).compl)] @@ -193,16 +218,19 @@ theorem stoppedValue_min_ae_eq_condexp [SigmaFiniteFiltration μ ℱ] (h : Marti filter_upwards [this] with x hx hx_mem simp only [Set.mem_compl_iff, Set.mem_setOf_eq, not_le] at hx_mem exact hx hx_mem.le - apply Filter.EventuallyEq.trans _ ((condexp_min_stopping_time_ae_eq_restrict_le hτ hσ).trans _) + apply Filter.EventuallyEq.trans _ ((condExp_min_stopping_time_ae_eq_restrict_le hτ hσ).trans _) · exact stoppedValue f τ · rw [IsStoppingTime.measurableSpace_min hσ, IsStoppingTime.measurableSpace_min hτ, inf_comm] · have h1 : μ[stoppedValue f τ|hτ.measurableSpace] = stoppedValue f τ := by - apply condexp_of_stronglyMeasurable hτ.measurableSpace_le + apply condExp_of_stronglyMeasurable hτ.measurableSpace_le · exact Measurable.stronglyMeasurable <| measurable_stoppedValue h.adapted.progMeasurable_of_discrete hτ · exact integrable_stoppedValue ι hτ h.integrable hτ_le rw [h1] - exact (condexp_stoppedValue_stopping_time_ae_eq_restrict_le h hτ hσ hτ_le).symm + exact (condExp_stoppedValue_stopping_time_ae_eq_restrict_le h hτ hσ hτ_le).symm + +@[deprecated (since := "2025-01-21")] +alias stoppedValue_min_ae_eq_condexp := stoppedValue_min_ae_eq_condExp end SubsetOfNat diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments/Basic.lean similarity index 95% rename from Mathlib/Probability/Moments.lean rename to Mathlib/Probability/Moments/Basic.lean index 7e8db5fd848ce..ec9aaa1d55422 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Probability.IdentDistrib -import Mathlib.Probability.Variance /-! # Moments and moment generating function @@ -89,7 +88,7 @@ theorem centralMoment_one [IsZeroOrProbabilityMeasure μ] : centralMoment X 1 μ refine fun h_sub => h_int ?_ have h_add : X = (fun x => X x - integral μ X) + fun _ => integral μ X := by ext1 x; simp rw [h_add] - exact h_sub.add (integrable_const _) + fun_prop rw [integral_undef this] theorem centralMoment_two_eq_variance [IsFiniteMeasure μ] (hX : Memℒp X 2 μ) : @@ -213,6 +212,26 @@ theorem mgf_const_add (α : ℝ) : mgf (fun ω => α + X ω) μ t = exp (t * α) theorem mgf_add_const (α : ℝ) : mgf (fun ω => X ω + α) μ t = mgf X μ t * exp (t * α) := by simp only [add_comm, mgf_const_add, mul_comm] +/-- The moment generating function is monotone in the random variable for `t ≥ 0`. -/ +lemma mgf_mono_of_nonneg {Y : Ω → ℝ} (hXY : X ≤ᵐ[μ] Y) (ht : 0 ≤ t) + (htY : Integrable (fun ω ↦ exp (t * Y ω)) μ) : + mgf X μ t ≤ mgf Y μ t := by + by_cases htX : Integrable (fun ω ↦ exp (t * X ω)) μ + · refine integral_mono_ae htX htY ?_ + filter_upwards [hXY] with ω hω using by gcongr + · rw [mgf_undef htX] + exact mgf_nonneg + +/-- The moment generating function is antitone in the random variable for `t ≤ 0`. -/ +lemma mgf_anti_of_nonpos {Y : Ω → ℝ} (hXY : X ≤ᵐ[μ] Y) (ht : t ≤ 0) + (htX : Integrable (fun ω ↦ exp (t * X ω)) μ) : + mgf Y μ t ≤ mgf X μ t := by + by_cases htY : Integrable (fun ω ↦ exp (t * Y ω)) μ + · refine integral_mono_ae htY htX ?_ + filter_upwards [hXY] with ω hω using exp_monotone <| mul_le_mul_of_nonpos_left hω ht + · rw [mgf_undef htY] + exact mgf_nonneg + section IndepFun /-- This is a trivial application of `IndepFun.comp` but it will come up frequently. -/ diff --git a/Mathlib/Probability/Moments/IntegrableExpMul.lean b/Mathlib/Probability/Moments/IntegrableExpMul.lean new file mode 100644 index 0000000000000..f2135d63a157e --- /dev/null +++ b/Mathlib/Probability/Moments/IntegrableExpMul.lean @@ -0,0 +1,546 @@ +/- +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.MeasureTheory.Function.L1Space +import Mathlib.MeasureTheory.Order.Group.Lattice + +/-! +# Domain of the moment generating function + +For `X` a real random variable and `μ` a finite measure, the set +`{t | Integrable (fun ω ↦ exp (t * X ω)) μ}` is an interval containing zero. This is the set of +points for which the moment generating function `mgf X μ t` is well defined. +We denote that set by `integrableExpSet X μ`. + +We prove the integrability of other functions for `t` in the interior of that interval. + +## Main definitions + +* `ProbabilityTheory.IntegrableExpSet`: the interval of reals for which `exp (t * X)` is integrable. + +## Main results + +* `ProbabilityTheory.integrable_exp_mul_of_le_of_le`: if `exp (u * X)` is integrable for `u = a` and + `u = b` then it is integrable on `[a, b]`. +* `ProbabilityTheory.convex_integrableExpSet`: `integrableExpSet X μ` is a convex set. +* `ProbabilityTheory.integrable_exp_mul_of_nonpos_of_ge`: if `exp (u * X)` is integrable for `u ≤ 0` + then it is integrable on `[u, 0]`. +* `ProbabilityTheory.integrable_rpow_abs_mul_exp_of_mem_interior`: for `v` in the interior of the + interval in which `exp (t * X)` is integrable, for all nonnegative `p : ℝ`, + `|X| ^ p * exp (v * X)` is integrable. +* `ProbabilityTheory.memℒp_of_mem_interior_integrableExpSet`: if 0 belongs to the interior of + `integrableExpSet X μ`, then `X` is in `ℒp` for all finite `p`. + +-/ + + +open MeasureTheory Filter Finset Real + +open scoped MeasureTheory ProbabilityTheory ENNReal NNReal Topology + +namespace ProbabilityTheory + +variable {Ω ι : Type*} {m : MeasurableSpace Ω} {X : Ω → ℝ} {μ : Measure Ω} {t u v : ℝ} + +section Interval + +lemma integrable_exp_mul_of_le_of_le {a b : ℝ} + (ha : Integrable (fun ω ↦ exp (a * X ω)) μ) (hb : Integrable (fun ω ↦ exp (b * X ω)) μ) + (hat : a ≤ t) (htb : t ≤ b) : + Integrable (fun ω ↦ exp (t * X ω)) μ := by + refine Integrable.mono (ha.add hb) ?_ (ae_of_all _ fun ω ↦ ?_) + · by_cases hab : a = b + · have ha_eq_t : a = t := le_antisymm hat (hab ▸ htb) + rw [← ha_eq_t] + exact ha.1 + · refine AEMeasurable.aestronglyMeasurable ?_ + refine measurable_exp.comp_aemeasurable (AEMeasurable.const_mul ?_ _) + by_cases ha_zero : a = 0 + · refine aemeasurable_of_aemeasurable_exp_mul ?_ hb.1.aemeasurable + rw [ha_zero] at hab + exact Ne.symm hab + · exact aemeasurable_of_aemeasurable_exp_mul ha_zero ha.1.aemeasurable + · simp only [norm_eq_abs, abs_exp, Pi.add_apply] + conv_rhs => rw [abs_of_nonneg (by positivity)] + rcases le_total 0 (X ω) with h | h + · calc exp (t * X ω) + _ ≤ exp (b * X ω) := exp_le_exp.mpr (mul_le_mul_of_nonneg_right htb h) + _ ≤ exp (a * X ω) + exp (b * X ω) := le_add_of_nonneg_left (exp_nonneg _) + · calc exp (t * X ω) + _ ≤ exp (a * X ω) := exp_le_exp.mpr (mul_le_mul_of_nonpos_right hat h) + _ ≤ exp (a * X ω) + exp (b * X ω) := le_add_of_nonneg_right (exp_nonneg _) + +/-- If `ω ↦ exp (u * X ω)` is integrable at `u` and `-u`, then it is integrable on `[-u, u]`. -/ +lemma integrable_exp_mul_of_abs_le + (hu_int_pos : Integrable (fun ω ↦ exp (u * X ω)) μ) + (hu_int_neg : Integrable (fun ω ↦ exp (- u * X ω)) μ) + (htu : |t| ≤ |u|) : + Integrable (fun ω ↦ exp (t * X ω)) μ := by + refine integrable_exp_mul_of_le_of_le (a := -|u|) (b := |u|) ?_ ?_ ?_ ?_ + · rcases le_total 0 u with hu | hu + · rwa [abs_of_nonneg hu] + · simpa [abs_of_nonpos hu] + · rcases le_total 0 u with hu | hu + · rwa [abs_of_nonneg hu] + · rwa [abs_of_nonpos hu] + · rw [neg_le] + exact (neg_le_abs t).trans htu + · exact (le_abs_self t).trans htu + +/-- If `ω ↦ exp (u * X ω)` is integrable at `u ≥ 0`, then it is integrable on `[0, u]`. -/ +lemma integrable_exp_mul_of_nonneg_of_le [IsFiniteMeasure μ] + (hu : Integrable (fun ω ↦ exp (u * X ω)) μ) (h_nonneg : 0 ≤ t) (htu : t ≤ u) : + Integrable (fun ω ↦ exp (t * X ω)) μ := + integrable_exp_mul_of_le_of_le (by simp) hu h_nonneg htu + +/-- If `ω ↦ exp (u * X ω)` is integrable at `u ≤ 0`, then it is integrable on `[u, 0]`. -/ +lemma integrable_exp_mul_of_nonpos_of_ge [IsFiniteMeasure μ] + (hu : Integrable (fun ω ↦ exp (u * X ω)) μ) (h_nonpos : t ≤ 0) (htu : u ≤ t) : + Integrable (fun ω ↦ exp (t * X ω)) μ := + integrable_exp_mul_of_le_of_le hu (by simp) htu h_nonpos + +end Interval + +section IntegrableExpSet + +/-- The interval of reals `t` for which `exp (t * X)` is integrable. -/ +def integrableExpSet (X : Ω → ℝ) (μ : Measure Ω) : Set ℝ := + {t | Integrable (fun ω ↦ exp (t * X ω)) μ} + +lemma integrable_of_mem_integrableExpSet (h : t ∈ integrableExpSet X μ) : + Integrable (fun ω ↦ exp (t * X ω)) μ := h + +/-- `integrableExpSet X μ` is a convex subset of `ℝ` (it is an interval). -/ +lemma convex_integrableExpSet : Convex ℝ (integrableExpSet X μ) := by + rintro t₁ ht₁ t₂ ht₂ a b ha hb hab + wlog h_le : t₁ ≤ t₂ + · rw [add_comm] at hab ⊢ + exact this ht₂ ht₁ hb ha hab (not_le.mp h_le).le + refine integrable_exp_mul_of_le_of_le ht₁ ht₂ ?_ ?_ + · simp only [smul_eq_mul] + calc t₁ + _ = a * t₁ + b * t₁ := by rw [← add_mul, hab, one_mul] + _ ≤ a * t₁ + b * t₂ := by gcongr + · simp only [smul_eq_mul] + calc a * t₁ + b * t₂ + _ ≤ a * t₂ + b * t₂ := by gcongr + _ = t₂ := by rw [← add_mul, hab, one_mul] + +end IntegrableExpSet + +section FiniteMoments + +lemma aemeasurable_of_integrable_exp_mul (huv : u ≠ v) + (hu_int : Integrable (fun ω ↦ exp (u * X ω)) μ) + (hv_int : Integrable (fun ω ↦ exp (v * X ω)) μ) : + AEMeasurable X μ := by + by_cases hu : u = 0 + · have hv : v ≠ 0 := fun h_eq ↦ huv (h_eq ▸ hu) + exact aemeasurable_of_aemeasurable_exp_mul hv hv_int.aemeasurable + · exact aemeasurable_of_aemeasurable_exp_mul hu hu_int.aemeasurable + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable, then +`ω ↦ exp (t * |X| + v * X)` is integrable. -/ +lemma integrable_exp_mul_abs_add (ht_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) : + Integrable (fun ω ↦ exp (t * |X ω| + v * X ω)) μ := by + have h_int_add : Integrable (fun a ↦ exp ((v + t) * X a) + exp ((v - t) * X a)) μ := + ht_int_pos.add <| by simpa using ht_int_neg + refine Integrable.mono h_int_add ?_ (ae_of_all _ fun ω ↦ ?_) + · by_cases ht : t = 0 + · simp only [ht, zero_mul, zero_add, add_zero] at ht_int_pos ⊢ + exact ht_int_pos.1 + have hX : AEMeasurable X μ := aemeasurable_of_integrable_exp_mul ?_ ht_int_pos ht_int_neg + swap; · rw [← sub_ne_zero]; simp [ht] + refine AEMeasurable.aestronglyMeasurable ?_ + exact measurable_exp.comp_aemeasurable ((hX.abs.const_mul _).add (hX.const_mul _)) + · simp only [norm_eq_abs, abs_exp] + conv_rhs => rw [abs_of_nonneg (by positivity)] + -- ⊢ exp (t * |X ω| + v * X ω) ≤ exp ((v + t) * X ω) + exp ((v - t) * X ω) + rcases le_total 0 (X ω) with h_nonneg | h_nonpos + · rw [abs_of_nonneg h_nonneg, ← add_mul, add_comm, le_add_iff_nonneg_right] + positivity + · rw [abs_of_nonpos h_nonpos, mul_neg, mul_comm, ← mul_neg, mul_comm, ← add_mul, add_comm, + ← sub_eq_add_neg, le_add_iff_nonneg_left] + positivity + +/-- If `ω ↦ exp (t * X ω)` is integrable at `t` and `-t`, then `ω ↦ exp (t * |X ω|)` is +integrable. -/ +lemma integrable_exp_mul_abs (ht_int_pos : Integrable (fun ω ↦ exp (t * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp (- t * X ω)) μ) : + Integrable (fun ω ↦ exp (t * |X ω|)) μ := by + have h := integrable_exp_mul_abs_add (t := t) (μ := μ) (X := X) (v := 0) ?_ ?_ + · simpa using h + · simpa using ht_int_pos + · simpa using ht_int_neg + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable, then +`ω ↦ exp (|t| * |X| + v * X)` is integrable. -/ +lemma integrable_exp_abs_mul_abs_add (ht_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) : + Integrable (fun ω ↦ exp (|t| * |X ω| + v * X ω)) μ := by + rcases le_total 0 t with ht_nonneg | ht_nonpos + · simp_rw [abs_of_nonneg ht_nonneg] + exact integrable_exp_mul_abs_add ht_int_pos ht_int_neg + · simp_rw [abs_of_nonpos ht_nonpos] + exact integrable_exp_mul_abs_add ht_int_neg (by simpa using ht_int_pos) + +/-- If `ω ↦ exp (t * X ω)` is integrable at `t` and `-t`, then `ω ↦ exp (|t| * |X ω|)` is +integrable. -/ +lemma integrable_exp_abs_mul_abs (ht_int_pos : Integrable (fun ω ↦ exp (t * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp (- t * X ω)) μ) : + Integrable (fun ω ↦ exp (|t| * |X ω|)) μ := by + rcases le_total 0 t with ht_nonneg | ht_nonpos + · simp_rw [abs_of_nonneg ht_nonneg] + exact integrable_exp_mul_abs ht_int_pos ht_int_neg + · simp_rw [abs_of_nonpos ht_nonpos] + exact integrable_exp_mul_abs ht_int_neg (by simpa using ht_int_pos) + +/-- Auxiliary lemma for `rpow_abs_le_mul_max_exp`. -/ +lemma rpow_abs_le_mul_max_exp_of_pos (x : ℝ) {t p : ℝ} (hp : 0 ≤ p) (ht : 0 < t) : + |x| ^ p ≤ (p / t) ^ p * max (exp (t * x)) (exp (- t * x)) := by + by_cases hp_zero : p = 0 + · simp only [hp_zero, rpow_zero, zero_div, neg_mul, one_mul, le_sup_iff, one_le_exp_iff, + Left.nonneg_neg_iff] + exact le_total 0 (t * x) + have h_x_le c (hc : 0 < c) : x ≤ c⁻¹ * exp (c * x) := le_inv_mul_exp x hc + have h_neg_x_le c (hc : 0 < c) : -x ≤ c⁻¹ * exp (- c * x) := by simpa using le_inv_mul_exp (-x) hc + have h_abs_le c (hc : 0 < c) : |x| ≤ c⁻¹ * max (exp (c * x)) (exp (- c * x)) := by + refine abs_le.mpr ⟨?_, ?_⟩ + · rw [neg_le] + refine (h_neg_x_le c hc).trans ?_ + gcongr + exact le_max_right _ _ + · refine (h_x_le c hc).trans ?_ + gcongr + exact le_max_left _ _ + calc |x| ^ p + _ ≤ ((t / p)⁻¹ * max (exp (t / p * x)) (exp (- t / p * x))) ^ p := by + refine rpow_le_rpow (abs_nonneg _) ?_ hp + convert h_abs_le (t / p) (div_pos ht (hp.lt_of_ne' hp_zero)) using 5 + rw [neg_div] + _ = (p / t) ^ p * max (exp (t * x)) (exp (- t * x)) := by + rw [mul_rpow (by positivity) (by positivity)] + congr + · field_simp + · rw [rpow_max (by positivity) (by positivity) hp, ← exp_mul, ← exp_mul] + ring_nf + congr <;> rw [mul_assoc, mul_inv_cancel₀ hp_zero, mul_one] + +lemma rpow_abs_le_mul_max_exp (x : ℝ) {t p : ℝ} (hp : 0 ≤ p) (ht : t ≠ 0) : + |x| ^ p ≤ (p / |t|) ^ p * max (exp (t * x)) (exp (- t * x)) := by + rcases lt_or_gt_of_ne ht with ht_neg | ht_pos + · rw [abs_of_nonpos ht_neg.le, sup_comm] + convert rpow_abs_le_mul_max_exp_of_pos x hp (t := -t) (by simp [ht_neg]) + simp + · rw [abs_of_nonneg ht_pos.le] + exact rpow_abs_le_mul_max_exp_of_pos x hp ht_pos + +lemma rpow_abs_le_mul_exp_abs (x : ℝ) {t p : ℝ} (hp : 0 ≤ p) (ht : t ≠ 0) : + |x| ^ p ≤ (p / |t|) ^ p * exp (|t| * |x|) := by + refine (rpow_abs_le_mul_max_exp_of_pos x hp (t := |t|) ?_).trans_eq ?_ + · simp [ht] + · congr + rcases le_total 0 x with hx | hx + · rw [abs_of_nonneg hx] + simp only [neg_mul, sup_eq_left, exp_le_exp, neg_le_self_iff] + positivity + · rw [abs_of_nonpos hx] + simp only [neg_mul, mul_neg, sup_eq_right, exp_le_exp, le_neg_self_iff] + exact mul_nonpos_of_nonneg_of_nonpos (abs_nonneg _) hx + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable +then for nonnegative `p : ℝ` and any `x ∈ [0, |t|)`, +`|X| ^ p * exp (v * X + x * |X|)` is integrable. -/ +lemma integrable_rpow_abs_mul_exp_add_of_integrable_exp_mul {x : ℝ} + (h_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (h_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) (h_nonneg : 0 ≤ x) (hx : x < |t|) + {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun a ↦ |X a| ^ p * exp (v * X a + x * |X a|)) μ := by + have ht : t ≠ 0 := by + suffices |t| ≠ 0 by simpa + exact (h_nonneg.trans_lt hx).ne' + have hX : AEMeasurable X μ := aemeasurable_of_integrable_exp_mul ?_ h_int_pos h_int_neg + swap; · rw [← sub_ne_zero]; simp [ht] + rw [← integrable_norm_iff] + swap + · exact AEMeasurable.aestronglyMeasurable <| by fun_prop + simp only [norm_mul, norm_pow, norm_eq_abs, sq_abs, abs_exp] + have h_le a : |X a| ^ p * exp (v * X a + x * |X a|) + ≤ (p / (|t| - x)) ^ p * exp (v * X a + |t| * |X a|) := by + simp_rw [exp_add, mul_comm (exp (v * X a)), ← mul_assoc] + gcongr ?_ * _ + have : |t| = |t| - x + x := by simp + nth_rw 2 [this] + rw [add_mul, exp_add, ← mul_assoc] + gcongr ?_ * _ + convert rpow_abs_le_mul_exp_abs (X a) hp (t := |t| - x) _ using 4 + · nth_rw 2 [abs_of_nonneg] + simp [hx.le] + · nth_rw 2 [abs_of_nonneg] + simp [hx.le] + · rw [sub_ne_zero] + exact hx.ne' + refine Integrable.mono (g := fun a ↦ (p / (|t| - x)) ^ p * exp (v * X a + |t| * |X a|)) + ?_ ?_ <| ae_of_all _ fun ω ↦ ?_ + · refine Integrable.const_mul ?_ _ + simp_rw [add_comm (v * X _)] + exact integrable_exp_abs_mul_abs_add h_int_pos h_int_neg + · refine AEMeasurable.aestronglyMeasurable ?_ + exact AEMeasurable.mul (by fun_prop) (measurable_exp.comp_aemeasurable (by fun_prop)) + · simp only [sq_abs, norm_mul, norm_pow, norm_eq_abs, abs_exp, norm_div, norm_ofNat] + simp_rw [abs_rpow_of_nonneg (abs_nonneg _), abs_abs] + refine (h_le ω).trans_eq ?_ + congr + symm + simp only [abs_eq_self, sub_nonneg] + exact rpow_nonneg (div_nonneg hp (sub_nonneg_of_le hx.le)) _ + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable +then for any `n : ℕ` and any `x ∈ [0, |t|)`, +`|X| ^ n * exp (v * X + x * |X|)` is integrable. -/ +lemma integrable_pow_abs_mul_exp_add_of_integrable_exp_mul {x : ℝ} + (h_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (h_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) (h_nonneg : 0 ≤ x) (hx : x < |t|) + (n : ℕ) : + Integrable (fun a ↦ |X a| ^ n * exp (v * X a + x * |X a|)) μ := by + convert integrable_rpow_abs_mul_exp_add_of_integrable_exp_mul h_int_pos h_int_neg h_nonneg hx + n.cast_nonneg + simp + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable +then for nonnegative `p : ℝ`, `|X| ^ p * exp (v * X)` is integrable. -/ +lemma integrable_rpow_abs_mul_exp_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ |X ω| ^ p * exp (v * X ω)) μ := by + convert integrable_rpow_abs_mul_exp_add_of_integrable_exp_mul ht_int_pos ht_int_neg le_rfl _ hp + using 4 + · simp + · simp [ht] + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable, then for all `n : ℕ`, +`|X| ^ n * exp (v * X)` is integrable. -/ +lemma integrable_pow_abs_mul_exp_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) (n : ℕ) : + Integrable (fun ω ↦ |X ω| ^ n * exp (v * X ω)) μ := by + convert integrable_rpow_abs_mul_exp_of_integrable_exp_mul ht ht_int_pos ht_int_neg + (by positivity : 0 ≤ (n : ℝ)) with ω + simp + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable, then for all nonnegative `p : ℝ`, +`X ^ p * exp (v * X)` is integrable. -/ +lemma integrable_rpow_mul_exp_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ X ω ^ p * exp (v * X ω)) μ := by + have hX : AEMeasurable X μ := aemeasurable_of_integrable_exp_mul ?_ ht_int_pos ht_int_neg + swap; · rw [← sub_ne_zero]; simp [ht] + rw [← integrable_norm_iff] + · simp_rw [norm_eq_abs, abs_mul, abs_exp] + have h := integrable_rpow_abs_mul_exp_of_integrable_exp_mul ht ht_int_pos ht_int_neg hp + refine h.mono' ?_ ?_ + · exact ((hX.pow_const _).abs.mul + (measurable_exp.comp_aemeasurable (hX.const_mul _))).aestronglyMeasurable + · refine ae_of_all _ fun ω ↦ ?_ + simp only [norm_mul, norm_eq_abs, abs_abs, abs_exp] + gcongr + exact abs_rpow_le_abs_rpow _ _ + · exact ((hX.pow_const _).mul + (measurable_exp.comp_aemeasurable (hX.const_mul _))).aestronglyMeasurable + +/-- If `exp ((v + t) * X)` and `exp ((v - t) * X)` are integrable, then for all `n : ℕ`, +`X ^ n * exp (v * X)` is integrable. -/ +lemma integrable_pow_mul_exp_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp ((v + t) * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp ((v - t) * X ω)) μ) (n : ℕ) : + Integrable (fun ω ↦ X ω ^ n * exp (v * X ω)) μ := by + convert integrable_rpow_mul_exp_of_integrable_exp_mul ht ht_int_pos ht_int_neg + (by positivity : 0 ≤ (n : ℝ)) with ω + simp + +/-- If `ω ↦ exp (t * X ω)` is integrable at `t` and `-t` for `t ≠ 0`, then `ω ↦ |X ω| ^ p` is +integrable for all nonnegative `p : ℝ`. -/ +lemma integrable_rpow_abs_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp (t * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp (- t * X ω)) μ) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ |X ω| ^ p) μ := by + have h := integrable_rpow_abs_mul_exp_of_integrable_exp_mul (μ := μ) (X := X) ht (v := 0) ?_ ?_ hp + · simpa using h + · simpa using ht_int_pos + · simpa using ht_int_neg + +/-- If `ω ↦ exp (t * X ω)` is integrable at `t` and `-t` for `t ≠ 0`, then `ω ↦ |X ω| ^ n` is +integrable for all `n : ℕ`. That is, all moments of `X` are finite. -/ +lemma integrable_pow_abs_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp (t * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp (- t * X ω)) μ) (n : ℕ) : + Integrable (fun ω ↦ |X ω| ^ n) μ := by + convert integrable_rpow_abs_of_integrable_exp_mul ht ht_int_pos ht_int_neg + (by positivity : 0 ≤ (n : ℝ)) with ω + simp + +/-- If `ω ↦ exp (t * X ω)` is integrable at `t` and `-t` for `t ≠ 0`, then `ω ↦ X ω ^ p` is +integrable for all nonnegative `p : ℝ`. -/ +lemma integrable_rpow_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp (t * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp (- t * X ω)) μ) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ X ω ^ p) μ := by + have h := integrable_rpow_mul_exp_of_integrable_exp_mul (μ := μ) (X := X) ht (v := 0) ?_ ?_ hp + · simpa using h + · simpa using ht_int_pos + · simpa using ht_int_neg + +/-- If `ω ↦ exp (t * X ω)` is integrable at `t` and `-t` for `t ≠ 0`, then `ω ↦ X ω ^ n` is +integrable for all `n : ℕ`. -/ +lemma integrable_pow_of_integrable_exp_mul (ht : t ≠ 0) + (ht_int_pos : Integrable (fun ω ↦ exp (t * X ω)) μ) + (ht_int_neg : Integrable (fun ω ↦ exp (- t * X ω)) μ) (n : ℕ) : + Integrable (fun ω ↦ X ω ^ n) μ := by + convert integrable_rpow_of_integrable_exp_mul ht ht_int_pos ht_int_neg + (by positivity : 0 ≤ (n : ℝ)) with ω + simp + +section IntegrableExpSet + +lemma add_half_inf_sub_mem_Ioo {l u v : ℝ} (hv : v ∈ Set.Ioo l u) : + v + ((v - l) ⊓ (u - v)) / 2 ∈ Set.Ioo l u := by + have h_pos : 0 < (v - l) ⊓ (u - v) := by simp [hv.1, hv.2] + constructor + · calc l < v := hv.1 + _ ≤ v + ((v - l) ⊓ (u - v)) / 2 := le_add_of_nonneg_right (by positivity) + · calc v + ((v - l) ⊓ (u - v)) / 2 + _ < v + ((v - l) ⊓ (u - v)) := by gcongr; exact half_lt_self (by positivity) + _ ≤ v + (u - v) := by gcongr; exact inf_le_right + _ = u := by abel + +lemma sub_half_inf_sub_mem_Ioo {l u v : ℝ} (hv : v ∈ Set.Ioo l u) : + v - ((v - l) ⊓ (u - v)) / 2 ∈ Set.Ioo l u := by + have h_pos : 0 < (v - l) ⊓ (u - v) := by simp [hv.1, hv.2] + constructor + · calc l = v - (v - l) := by abel + _ ≤ v - ((v - l) ⊓ (u - v)) := by gcongr; exact inf_le_left + _ < v - ((v - l) ⊓ (u - v)) / 2 := by gcongr; exact half_lt_self (by positivity) + · calc v - ((v - l) ⊓ (u - v)) / 2 + _ ≤ v := by + rw [sub_le_iff_le_add] + exact le_add_of_nonneg_right (by positivity) + _ < u := hv.2 + +/-- If the interior of the interval `integrableExpSet X μ` is nonempty, +then `X` is a.e. measurable. -/ +lemma aemeasurable_of_mem_interior_integrableExpSet (hv : v ∈ interior (integrableExpSet X μ)) : + AEMeasurable X μ := by + rw [mem_interior_iff_mem_nhds, mem_nhds_iff_exists_Ioo_subset] at hv + obtain ⟨l, u, hvlu, h_subset⟩ := hv + let t := ((v - l) ⊓ (u - v)) / 2 + have h_pos : 0 < (v - l) ⊓ (u - v) := by simp [hvlu.1, hvlu.2] + have ht : 0 < t := half_pos h_pos + by_cases hvt : v + t = 0 + · have hvt' : v - t ≠ 0 := by + rw [sub_ne_zero] + refine fun h_eq ↦ ht.ne' ?_ + simpa [h_eq] using hvt + exact aemeasurable_of_aemeasurable_exp_mul hvt' + (h_subset (sub_half_inf_sub_mem_Ioo hvlu)).aemeasurable + · exact aemeasurable_of_aemeasurable_exp_mul hvt + (h_subset (add_half_inf_sub_mem_Ioo hvlu)).aemeasurable + +/-- If `v` belongs to the interior of the interval `integrableExpSet X μ`, +then `|X| ^ p * exp (v * X)` is integrable for all nonnegative `p : ℝ`. -/ +lemma integrable_rpow_abs_mul_exp_of_mem_interior_integrableExpSet + (hv : v ∈ interior (integrableExpSet X μ)) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ |X ω| ^ p * exp (v * X ω)) μ := by + rw [mem_interior_iff_mem_nhds, mem_nhds_iff_exists_Ioo_subset] at hv + obtain ⟨l, u, hvlu, h_subset⟩ := hv + have h_pos : 0 < (v - l) ⊓ (u - v) := by simp [hvlu.1, hvlu.2] + refine integrable_rpow_abs_mul_exp_of_integrable_exp_mul + (t := min (v - l) (u - v) / 2) ?_ ?_ ?_ hp + · positivity + · exact h_subset (add_half_inf_sub_mem_Ioo hvlu) + · exact h_subset (sub_half_inf_sub_mem_Ioo hvlu) + +/-- If `v` belongs to the interior of the interval `integrableExpSet X μ`, +then `|X| ^ n * exp (v * X)` is integrable for all `n : ℕ`. -/ +lemma integrable_pow_abs_mul_exp_of_mem_interior_integrableExpSet + (hv : v ∈ interior (integrableExpSet X μ)) (n : ℕ) : + Integrable (fun ω ↦ |X ω| ^ n * exp (v * X ω)) μ := by + convert integrable_rpow_abs_mul_exp_of_mem_interior_integrableExpSet hv + (by positivity : 0 ≤ (n : ℝ)) with ω + simp + +/-- If `v` belongs to the interior of the interval `integrableExpSet X μ`, +then `X ^ p * exp (v * X)` is integrable for all nonnegative `p : ℝ`. -/ +lemma integrable_rpow_mul_exp_of_mem_interior_integrableExpSet + (hv : v ∈ interior (integrableExpSet X μ)) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ X ω ^ p * exp (v * X ω)) μ := by + rw [mem_interior_iff_mem_nhds, mem_nhds_iff_exists_Ioo_subset] at hv + obtain ⟨l, u, hvlu, h_subset⟩ := hv + have h_pos : 0 < (v - l) ⊓ (u - v) := by simp [hvlu.1, hvlu.2] + refine integrable_rpow_mul_exp_of_integrable_exp_mul + (t := min (v - l) (u - v) / 2) ?_ ?_ ?_ hp + · positivity + · exact h_subset (add_half_inf_sub_mem_Ioo hvlu) + · exact h_subset (sub_half_inf_sub_mem_Ioo hvlu) + +/-- If `v` belongs to the interior of the interval `integrableExpSet X μ`, +then `X ^ n * exp (v * X)` is integrable for all `n : ℕ`. -/ +lemma integrable_pow_mul_exp_of_mem_interior_integrableExpSet + (hv : v ∈ interior (integrableExpSet X μ)) (n : ℕ) : + Integrable (fun ω ↦ X ω ^ n * exp (v * X ω)) μ := by + convert integrable_rpow_mul_exp_of_mem_interior_integrableExpSet hv + (by positivity : 0 ≤ (n : ℝ)) with ω + simp + +/-- If 0 belongs to the interior of the interval `integrableExpSet X μ`, +then `|X| ^ n` is integrable for all nonnegative `p : ℝ`. -/ +lemma integrable_rpow_abs_of_mem_interior_integrableExpSet + (h : 0 ∈ interior (integrableExpSet X μ)) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ |X ω| ^ p) μ := by + convert integrable_rpow_abs_mul_exp_of_mem_interior_integrableExpSet h hp using 1 + simp + +/-- If 0 belongs to the interior of the interval `integrableExpSet X μ`, +then `|X| ^ n` is integrable for all `n : ℕ`. -/ +lemma integrable_pow_abs_of_mem_interior_integrableExpSet + (h : 0 ∈ interior (integrableExpSet X μ)) (n : ℕ) : + Integrable (fun ω ↦ |X ω| ^ n) μ := by + convert integrable_pow_abs_mul_exp_of_mem_interior_integrableExpSet h n + simp + +/-- If 0 belongs to the interior of the interval `integrableExpSet X μ`, +then `X ^ n` is integrable for all nonnegative `p : ℝ`. -/ +lemma integrable_rpow_of_mem_interior_integrableExpSet + (h : 0 ∈ interior (integrableExpSet X μ)) {p : ℝ} (hp : 0 ≤ p) : + Integrable (fun ω ↦ X ω ^ p) μ := by + convert integrable_rpow_mul_exp_of_mem_interior_integrableExpSet h hp using 1 + simp + +/-- If 0 belongs to the interior of the interval `integrableExpSet X μ`, +then `X ^ n` is integrable for all `n : ℕ`. -/ +lemma integrable_pow_of_mem_interior_integrableExpSet + (h : 0 ∈ interior (integrableExpSet X μ)) (n : ℕ) : + Integrable (fun ω ↦ X ω ^ n) μ := by + convert integrable_pow_mul_exp_of_mem_interior_integrableExpSet h n + simp + +/-- If 0 belongs to the interior of `integrableExpSet X μ`, then `X` is in `ℒp` for all +finite `p`. -/ +lemma memℒp_of_mem_interior_integrableExpSet (h : 0 ∈ interior (integrableExpSet X μ)) (p : ℝ≥0) : + Memℒp X p μ := by + have hX : AEMeasurable X μ := aemeasurable_of_mem_interior_integrableExpSet h + by_cases hp_zero : p = 0 + · simp only [hp_zero, ENNReal.coe_zero, memℒp_zero_iff_aestronglyMeasurable] + exact hX.aestronglyMeasurable + rw [← integrable_norm_rpow_iff hX.aestronglyMeasurable (mod_cast hp_zero) (by simp)] + simp only [norm_eq_abs, ENNReal.coe_toReal] + exact integrable_rpow_abs_of_mem_interior_integrableExpSet h p.2 + +end IntegrableExpSet + +end FiniteMoments + +end ProbabilityTheory diff --git a/Mathlib/Probability/Notation.lean b/Mathlib/Probability/Notation.lean index aaf254c1f9096..1d29018bd6cb6 100644 --- a/Mathlib/Probability/Notation.lean +++ b/Mathlib/Probability/Notation.lean @@ -36,7 +36,7 @@ open scoped MeasureTheory -- We define notations `𝔼[f|m]` for the conditional expectation of `f` with respect to `m`. scoped[ProbabilityTheory] notation "𝔼[" X "|" m "]" => - MeasureTheory.condexp m MeasureTheory.MeasureSpace.volume X + MeasureTheory.condExp m MeasureTheory.MeasureSpace.volume X -- `scoped[ProbabilityTheory]` isn't legal for `macro`s. namespace ProbabilityTheory @@ -51,7 +51,7 @@ end ProbabilityTheory scoped[ProbabilityTheory] notation "𝔼[" X "]" => ∫ a, (X : _ → _) a scoped[ProbabilityTheory] notation P "⟦" s "|" m "⟧" => - MeasureTheory.condexp m P (Set.indicator s fun ω => (1 : ℝ)) + MeasureTheory.condExp m P (Set.indicator s fun ω => (1 : ℝ)) scoped[ProbabilityTheory] notation:50 X " =ₐₛ " Y:50 => X =ᵐ[MeasureTheory.MeasureSpace.volume] Y diff --git a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean index f5124ea768ad5..cf45350093d48 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean @@ -34,7 +34,6 @@ noncomputable section variable {α : Type*} -open scoped Classical open NNReal ENNReal MeasureTheory /-- A probability mass function, or discrete probability measures is a function `α → ℝ≥0∞` such @@ -99,6 +98,7 @@ theorem apply_eq_one_iff (p : PMF α) (a : α) : p a = 1 ↔ p.support = {a} := fun h => _root_.trans (symm <| tsum_eq_single a fun a' ha' => (p.apply_eq_zero_iff a').2 (h.symm ▸ ha')) p.tsum_coe⟩ suffices 1 < ∑' a, p a from ne_of_lt this p.tsum_coe.symm + classical have : 0 < ∑' b, ite (b = a) 0 (p b) := lt_of_le_of_ne' zero_le' ((tsum_ne_zero_iff ENNReal.summable).2 ⟨a', ite_ne_left_iff.2 ⟨ha, Ne.symm <| (p.mem_support_iff a').2 ha'⟩⟩) @@ -114,6 +114,7 @@ theorem apply_eq_one_iff (p : PMF α) (a : α) : p a = 1 ↔ p.support = {a} := _ = ∑' b, p b := tsum_congr fun b => by split_ifs <;> simp only [zero_add, add_zero, le_rfl] theorem coe_le_one (p : PMF α) (a : α) : p a ≤ 1 := by + classical refine hasSum_le (fun b => ?_) (hasSum_ite_eq a (p a)) (hasSum_coe_one p) split_ifs with h <;> simp only [h, zero_le', le_rfl] @@ -152,8 +153,8 @@ theorem toOuterMeasure_apply_finset (s : Finset α) : p.toOuterMeasure s = ∑ x theorem toOuterMeasure_apply_singleton (a : α) : p.toOuterMeasure {a} = p a := by refine (p.toOuterMeasure_apply {a}).trans ((tsum_eq_single a fun b hb => ?_).trans ?_) - · exact ite_eq_right_iff.2 fun hb' => False.elim <| hb hb' - · exact ite_eq_left_iff.2 fun ha' => False.elim <| ha' rfl + · classical exact ite_eq_right_iff.2 fun hb' => False.elim <| hb hb' + · classical exact ite_eq_left_iff.2 fun ha' => False.elim <| ha' rfl theorem toOuterMeasure_injective : (toOuterMeasure : PMF α → OuterMeasure α).Injective := fun p q h => PMF.ext fun x => (p.toOuterMeasure_apply_singleton x).symm.trans @@ -174,7 +175,7 @@ theorem toOuterMeasure_apply_eq_one_iff : p.toOuterMeasure s = 1 ↔ p.support have hsa : s.indicator p a < p a := hs'.symm ▸ (p.apply_pos_iff a).2 hap exact ENNReal.tsum_lt_tsum (p.tsum_coe_indicator_ne_top s) (fun x => Set.indicator_apply_le fun _ => le_rfl) hsa - · suffices ∀ (x) (_ : x ∉ s), p x = 0 from + · classical suffices ∀ (x) (_ : x ∉ s), p x = 0 from _root_.trans (tsum_congr fun a => (Set.indicator_apply s p a).trans (ite_eq_left_iff.2 <| symm ∘ this a)) p.tsum_coe diff --git a/Mathlib/Probability/ProbabilityMassFunction/Constructions.lean b/Mathlib/Probability/ProbabilityMassFunction/Constructions.lean index 6ef283f1de7dc..95ab35f0e5b9f 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Constructions.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Constructions.lean @@ -32,7 +32,6 @@ noncomputable section variable {α β γ : Type*} -open scoped Classical open NNReal ENNReal Finset MeasureTheory section Map @@ -45,6 +44,7 @@ variable (f : α → β) (p : PMF α) (b : β) theorem monad_map_eq_map {α β : Type u} (f : α → β) (p : PMF α) : f <$> p = p.map f := rfl +open scoped Classical in @[simp] theorem map_apply : (map f p) b = ∑' a, if b = f a then p a else 0 := by simp [map] @@ -82,6 +82,7 @@ variable (s : Set β) @[simp] theorem toOuterMeasure_map_apply : (p.map f).toOuterMeasure s = p.toOuterMeasure (f ⁻¹' s) := by simp [map, Set.indicator, toOuterMeasure_apply p (f ⁻¹' s)] + rfl variable {mα : MeasurableSpace α} {mβ : MeasurableSpace β} @@ -110,6 +111,7 @@ variable (q : PMF (α → β)) (p : PMF α) (b : β) theorem monad_seq_eq_seq {α β : Type u} (q : PMF (α → β)) (p : PMF α) : q <*> p = q.seq p := rfl +open scoped Classical in @[simp] theorem seq_apply : (seq q p) b = ∑' (f : α → β) (a : α), if b = f a then q f * p a else 0 := by simp only [seq, mul_boole, bind_apply, pure_apply] @@ -207,6 +209,7 @@ theorem support_ofFintype : (ofFintype f h).support = Function.support f := rfl theorem mem_support_ofFintype_iff (a : α) : a ∈ (ofFintype f h).support ↔ f a ≠ 0 := Iff.rfl +open scoped Classical in @[simp] lemma map_ofFintype [Fintype β] (f : α → ℝ≥0∞) (h : ∑ a, f a = 1) (g : α → β) : (ofFintype f h).map g = ofFintype (fun b ↦ ∑ a with g a = b, f a) diff --git a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean index acf07b214fb38..ee267a4c971de 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean @@ -23,7 +23,6 @@ noncomputable section variable {α β γ : Type*} -open scoped Classical open NNReal ENNReal open MeasureTheory @@ -32,6 +31,7 @@ namespace PMF section Pure +open scoped Classical in /-- The pure `PMF` is the `PMF` where all the mass lies in one point. The value of `pure a` is `1` at `a` and `0` elsewhere. -/ def pure (a : α) : PMF α := @@ -39,6 +39,7 @@ def pure (a : α) : PMF α := variable (a a' : α) +open scoped Classical in @[simp] theorem pure_apply : pure a a' = if a' = a then 1 else 0 := rfl @@ -61,6 +62,7 @@ section Measure variable (s : Set α) +open scoped Classical in @[simp] theorem toOuterMeasure_pure_apply : (pure a).toOuterMeasure s = if a ∈ s then 1 else 0 := by refine (toOuterMeasure_apply (pure a) s).trans ?_ @@ -74,6 +76,7 @@ theorem toOuterMeasure_pure_apply : (pure a).toOuterMeasure s = if a ∈ s then variable [MeasurableSpace α] +open scoped Classical in /-- The measure of a set under `pure a` is `1` for sets containing `a` and `0` otherwise. -/ @[simp] theorem toMeasure_pure_apply (hs : MeasurableSet s) : @@ -115,6 +118,7 @@ theorem mem_support_bind_iff (b : β) : @[simp] theorem pure_bind (a : α) (f : α → PMF β) : (pure a).bind f = f a := by + classical have : ∀ b a', ite (a' = a) (f a' b) 0 = ite (a' = a) (f a b) 0 := fun b a' => by split_ifs with h <;> simp [h] ext b @@ -148,7 +152,8 @@ variable (s : Set β) @[simp] theorem toOuterMeasure_bind_apply : - (p.bind f).toOuterMeasure s = ∑' a, p a * (f a).toOuterMeasure s := + (p.bind f).toOuterMeasure s = ∑' a, p a * (f a).toOuterMeasure s := by + classical calc (p.bind f).toOuterMeasure s = ∑' b, if b ∈ s then ∑' a, p a * f a b else 0 := by simp [toOuterMeasure_apply, Set.indicator_apply] @@ -236,6 +241,7 @@ theorem pure_bindOnSupport (a : α) (f : ∀ (a' : α) (_ : a' ∈ (pure a).supp (pure a).bindOnSupport f = f a ((mem_support_pure_iff a a).mpr rfl) := by refine PMF.ext fun b => ?_ simp only [bindOnSupport_apply, pure_apply] + classical refine _root_.trans (tsum_congr fun a' => ?_) (tsum_ite_eq a _) by_cases h : a' = a <;> simp [h] @@ -252,6 +258,7 @@ theorem bindOnSupport_bindOnSupport (p : PMF α) (f : ∀ a ∈ p.support, PMF refine PMF.ext fun a => ?_ dsimp only [bindOnSupport_apply] simp only [← tsum_dite_right, ENNReal.tsum_mul_left.symm, ENNReal.tsum_mul_right.symm] + classical simp only [ENNReal.tsum_eq_zero, dite_eq_left_iff] refine ENNReal.tsum_comm.trans (tsum_congr fun a' => tsum_congr fun b => ?_) split_ifs with h _ h_1 _ h_2 @@ -279,6 +286,7 @@ theorem toOuterMeasure_bindOnSupport_apply : (p.bindOnSupport f).toOuterMeasure s = ∑' a, p a * if h : p a = 0 then 0 else (f a h).toOuterMeasure s := by simp only [toOuterMeasure_apply, Set.indicator_apply, bindOnSupport_apply] + classical calc (∑' b, ite (b ∈ s) (∑' a, p a * dite (p a = 0) (fun h => 0) fun h => f a h b) 0) = ∑' (b) (a), ite (b ∈ s) (p a * dite (p a = 0) (fun h => 0) fun h => f a h b) 0 := diff --git a/Mathlib/Probability/Process/Adapted.lean b/Mathlib/Probability/Process/Adapted.lean index 101965bb8f539..eb30f8b442411 100644 --- a/Mathlib/Probability/Process/Adapted.lean +++ b/Mathlib/Probability/Process/Adapted.lean @@ -35,7 +35,7 @@ adapted, progressively measurable open Filter Order TopologicalSpace -open scoped Classical MeasureTheory NNReal ENNReal Topology +open scoped MeasureTheory NNReal ENNReal Topology namespace MeasureTheory diff --git a/Mathlib/Probability/Process/Filtration.lean b/Mathlib/Probability/Process/Filtration.lean index 89f83a4bcd145..82b1ace233b96 100644 --- a/Mathlib/Probability/Process/Filtration.lean +++ b/Mathlib/Probability/Process/Filtration.lean @@ -212,10 +212,14 @@ instance (priority := 100) IsFiniteMeasure.sigmaFiniteFiltration [Preorder ι] ( /-- Given an integrable function `g`, the conditional expectations of `g` with respect to a filtration is uniformly integrable. -/ -theorem Integrable.uniformIntegrable_condexp_filtration [Preorder ι] {μ : Measure Ω} +theorem Integrable.uniformIntegrable_condExp_filtration [Preorder ι] {μ : Measure Ω} [IsFiniteMeasure μ] {f : Filtration ι m} {g : Ω → ℝ} (hg : Integrable g μ) : UniformIntegrable (fun i => μ[g|f i]) 1 μ := - hg.uniformIntegrable_condexp f.le + hg.uniformIntegrable_condExp f.le + +@[deprecated (since := "2025-01-21")] +alias Integrable.uniformIntegrable_condexp_filtration := + Integrable.uniformIntegrable_condExp_filtration section OfSet @@ -322,7 +326,7 @@ theorem memℒp_limitProcess_of_eLpNorm_bdd {R : ℝ≥0} {p : ℝ≥0∞} {F : (lt_of_le_of_lt ?_ (ENNReal.coe_lt_top : ↑R < ∞))⟩ simp_rw [liminf_eq, eventually_atTop] exact sSup_le fun b ⟨a, ha⟩ => (ha a le_rfl).trans (hbdd _) - · exact zero_memℒp + · exact Memℒp.zero @[deprecated (since := "2024-07-27")] alias memℒp_limitProcess_of_snorm_bdd := memℒp_limitProcess_of_eLpNorm_bdd diff --git a/Mathlib/Probability/Process/HittingTime.lean b/Mathlib/Probability/Process/HittingTime.lean index b5a22e608a3fe..5ac75975d6483 100644 --- a/Mathlib/Probability/Process/HittingTime.lean +++ b/Mathlib/Probability/Process/HittingTime.lean @@ -36,21 +36,25 @@ hitting times indexed by the natural numbers or the reals. By taking the bounds open Filter Order TopologicalSpace -open scoped Classical MeasureTheory NNReal ENNReal Topology +open scoped MeasureTheory NNReal ENNReal Topology namespace MeasureTheory variable {Ω β ι : Type*} {m : MeasurableSpace Ω} +open scoped Classical in /-- Hitting time: given a stochastic process `u` and a set `s`, `hitting u s n m` is the first time `u` is in `s` after time `n` and before time `m` (if `u` does not hit `s` after time `n` and before `m` then the hitting time is simply `m`). The hitting time is a stopping time if the process is adapted and discrete. -/ -noncomputable def hitting [Preorder ι] [InfSet ι] (u : ι → Ω → β) (s : Set β) (n m : ι) : Ω → ι := - fun x => if ∃ j ∈ Set.Icc n m, u j x ∈ s then sInf (Set.Icc n m ∩ {i : ι | u i x ∈ s}) else m +noncomputable def hitting [Preorder ι] [InfSet ι] (u : ι → Ω → β) + (s : Set β) (n m : ι) : Ω → ι := + fun x => if ∃ j ∈ Set.Icc n m, u j x ∈ s + then sInf (Set.Icc n m ∩ {i : ι | u i x ∈ s}) else m #adaptation_note /-- nightly-2024-03-16: added to replace simp [hitting] -/ +open scoped Classical in theorem hitting_def [Preorder ι] [InfSet ι] (u : ι → Ω → β) (s : Set β) (n m : ι) : hitting u s n m = fun x => if ∃ j ∈ Set.Icc n m, u j x ∈ s then sInf (Set.Icc n m ∩ {i : ι | u i x ∈ s}) else m := @@ -90,11 +94,13 @@ theorem not_mem_of_lt_hitting {m k : ι} (hk₁ : k < hitting u s n m ω) (hk₂ theorem hitting_eq_end_iff {m : ι} : hitting u s n m ω = m ↔ (∃ j ∈ Set.Icc n m, u j ω ∈ s) → sInf (Set.Icc n m ∩ {i : ι | u i ω ∈ s}) = m := by + classical rw [hitting, ite_eq_right_iff] theorem hitting_of_le {m : ι} (hmn : m ≤ n) : hitting u s n m ω = m := by obtain rfl | h := le_iff_eq_or_lt.1 hmn - · rw [hitting, ite_eq_right_iff, forall_exists_index] + · classical + rw [hitting, ite_eq_right_iff, forall_exists_index] conv => intro; rw [Set.mem_Icc, Set.Icc_self, and_imp, and_imp] intro i hi₁ hi₂ hi rw [Set.inter_eq_left.2, csInf_singleton] diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index a9b107e3d88e4..f4be7f5441b4e 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -33,7 +33,7 @@ stopping time, stochastic process open Filter Order TopologicalSpace -open scoped Classical MeasureTheory NNReal ENNReal Topology +open scoped MeasureTheory NNReal ENNReal Topology namespace MeasureTheory @@ -96,6 +96,7 @@ protected theorem measurableSet_eq_of_countable_range (hτ : IsStoppingTime f τ rw [this] refine (hτ.measurableSet_le i).diff ?_ refine MeasurableSet.biUnion h_countable fun j _ => ?_ + classical rw [Set.iUnion_eq_if] split_ifs with hji · exact f.mono hji.le _ (hτ.measurableSet_le j) @@ -803,6 +804,7 @@ variable {μ : Measure Ω} {τ : Ω → ι} {E : Type*} {p : ℝ≥0∞} {u : ι theorem stoppedValue_eq_of_mem_finset [AddCommMonoid E] {s : Finset ι} (hbdd : ∀ ω, τ ω ∈ s) : stoppedValue u τ = ∑ i ∈ s, Set.indicator {ω | τ ω = i} (u i) := by ext y + classical rw [stoppedValue, Finset.sum_apply, Finset.sum_indicator_eq_sum_filter] suffices Finset.filter (fun i => y ∈ {ω : Ω | τ ω = i}) s = ({τ y} : Finset ι) by rw [this, Finset.sum_singleton] @@ -1046,24 +1048,32 @@ section Condexp variable [LinearOrder ι] {μ : Measure Ω} {ℱ : Filtration ι m} {τ σ : Ω → ι} {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] {f : Ω → E} -theorem condexp_stopping_time_ae_eq_restrict_eq_of_countable_range [SigmaFiniteFiltration μ ℱ] +theorem condExp_stopping_time_ae_eq_restrict_eq_of_countable_range [SigmaFiniteFiltration μ ℱ] (hτ : IsStoppingTime ℱ τ) (h_countable : (Set.range τ).Countable) [SigmaFinite (μ.trim (hτ.measurableSpace_le_of_countable_range h_countable))] (i : ι) : μ[f|hτ.measurableSpace] =ᵐ[μ.restrict {x | τ x = i}] μ[f|ℱ i] := by - refine condexp_ae_eq_restrict_of_measurableSpace_eq_on + refine condExp_ae_eq_restrict_of_measurableSpace_eq_on (hτ.measurableSpace_le_of_countable_range h_countable) (ℱ.le i) (hτ.measurableSet_eq_of_countable_range' h_countable i) fun t => ?_ rw [Set.inter_comm _ t, IsStoppingTime.measurableSet_inter_eq_iff] -theorem condexp_stopping_time_ae_eq_restrict_eq_of_countable [Countable ι] +@[deprecated (since := "2025-01-21")] +alias condexp_stopping_time_ae_eq_restrict_eq_of_countable_range := + condExp_stopping_time_ae_eq_restrict_eq_of_countable_range + +theorem condExp_stopping_time_ae_eq_restrict_eq_of_countable [Countable ι] [SigmaFiniteFiltration μ ℱ] (hτ : IsStoppingTime ℱ τ) [SigmaFinite (μ.trim hτ.measurableSpace_le_of_countable)] (i : ι) : μ[f|hτ.measurableSpace] =ᵐ[μ.restrict {x | τ x = i}] μ[f|ℱ i] := - condexp_stopping_time_ae_eq_restrict_eq_of_countable_range hτ (Set.to_countable _) i + condExp_stopping_time_ae_eq_restrict_eq_of_countable_range hτ (Set.to_countable _) i + +@[deprecated (since := "2025-01-21")] +alias condexp_stopping_time_ae_eq_restrict_eq_of_countable := + condExp_stopping_time_ae_eq_restrict_eq_of_countable variable [(Filter.atTop : Filter ι).IsCountablyGenerated] -theorem condexp_min_stopping_time_ae_eq_restrict_le_const (hτ : IsStoppingTime ℱ τ) (i : ι) +theorem condExp_min_stopping_time_ae_eq_restrict_le_const (hτ : IsStoppingTime ℱ τ) (i : ι) [SigmaFinite (μ.trim (hτ.min_const i).measurableSpace_le)] : μ[f|(hτ.min_const i).measurableSpace] =ᵐ[μ.restrict {x | τ x ≤ i}] μ[f|hτ.measurableSpace] := by have : SigmaFinite (μ.trim hτ.measurableSpace_le) := @@ -1071,21 +1081,28 @@ theorem condexp_min_stopping_time_ae_eq_restrict_le_const (hτ : IsStoppingTime rw [IsStoppingTime.measurableSpace_min_const] exact inf_le_left sigmaFiniteTrim_mono _ h_le - refine (condexp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le + refine (condExp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le (hτ.min_const i).measurableSpace_le (hτ.measurableSet_le' i) fun t => ?_).symm rw [Set.inter_comm _ t, hτ.measurableSet_inter_le_const_iff] +@[deprecated (since := "2025-01-21")] +alias condexp_min_stopping_time_ae_eq_restrict_le_const := + condExp_min_stopping_time_ae_eq_restrict_le_const + variable [TopologicalSpace ι] [OrderTopology ι] -theorem condexp_stopping_time_ae_eq_restrict_eq [FirstCountableTopology ι] +theorem condExp_stopping_time_ae_eq_restrict_eq [FirstCountableTopology ι] [SigmaFiniteFiltration μ ℱ] (hτ : IsStoppingTime ℱ τ) [SigmaFinite (μ.trim hτ.measurableSpace_le)] (i : ι) : μ[f|hτ.measurableSpace] =ᵐ[μ.restrict {x | τ x = i}] μ[f|ℱ i] := by - refine condexp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le (ℱ.le i) + refine condExp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le (ℱ.le i) (hτ.measurableSet_eq' i) fun t => ?_ rw [Set.inter_comm _ t, IsStoppingTime.measurableSet_inter_eq_iff] -theorem condexp_min_stopping_time_ae_eq_restrict_le [MeasurableSpace ι] [SecondCountableTopology ι] +@[deprecated (since := "2025-01-21")] +alias condexp_stopping_time_ae_eq_restrict_eq := condExp_stopping_time_ae_eq_restrict_eq + +theorem condExp_min_stopping_time_ae_eq_restrict_le [MeasurableSpace ι] [SecondCountableTopology ι] [BorelSpace ι] (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) [SigmaFinite (μ.trim (hτ.min hσ).measurableSpace_le)] : μ[f|(hτ.min hσ).measurableSpace] =ᵐ[μ.restrict {x | τ x ≤ σ x}] μ[f|hτ.measurableSpace] := by @@ -1095,10 +1112,13 @@ theorem condexp_min_stopping_time_ae_eq_restrict_le [MeasurableSpace ι] [Second · exact inf_le_left · simp_all only sigmaFiniteTrim_mono _ h_le - refine (condexp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le + refine (condExp_ae_eq_restrict_of_measurableSpace_eq_on hτ.measurableSpace_le (hτ.min hσ).measurableSpace_le (hτ.measurableSet_le_stopping_time hσ) fun t => ?_).symm rw [Set.inter_comm _ t, IsStoppingTime.measurableSet_inter_le_iff]; simp_all only +@[deprecated (since := "2025-01-21")] +alias condexp_min_stopping_time_ae_eq_restrict_le := condExp_min_stopping_time_ae_eq_restrict_le + end Condexp end MeasureTheory diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index b5cc11cbd0014..5c3c9f6c42be4 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -56,6 +56,27 @@ to `evariance`. -/ def variance {Ω : Type*} {_ : MeasurableSpace Ω} (X : Ω → ℝ) (μ : Measure Ω) : ℝ := (evariance X μ).toReal +/-- The `ℝ≥0∞`-valued variance of the real-valued random variable `X` according to the measure `μ`. + +This is defined as the Lebesgue integral of `(X - 𝔼[X])^2`. -/ +scoped notation "eVar[" X " ; " μ "]" => ProbabilityTheory.evariance X μ + +/-- The `ℝ≥0∞`-valued variance of the real-valued random variable `X` according to the volume +measure. + +This is defined as the Lebesgue integral of `(X - 𝔼[X])^2`. -/ +scoped notation "eVar[" X "]" => eVar[X ; MeasureTheory.MeasureSpace.volume] + +/-- The `ℝ`-valued variance of the real-valued random variable `X` according to the measure `μ`. + +It is set to `0` if `X` has infinite variance. -/ +scoped notation "Var[" X " ; " μ "]" => ProbabilityTheory.variance X μ + +/-- The `ℝ`-valued variance of the real-valued random variable `X` according to the volume measure. + +It is set to `0` if `X` has infinite variance. -/ +scoped notation "Var[" X "]" => Var[X ; MeasureTheory.MeasureSpace.volume] + variable {Ω : Type*} {m : MeasurableSpace Ω} {X : Ω → ℝ} {μ : Measure Ω} theorem _root_.MeasureTheory.Memℒp.evariance_lt_top [IsFiniteMeasure μ] (hX : Memℒp X 2 μ) : @@ -151,8 +172,6 @@ theorem evariance_mul (c : ℝ) (X : Ω → ℝ) (μ : Measure Ω) : rw [mul_comm] simp_rw [← smul_eq_mul, ← integral_smul_const, smul_eq_mul, mul_comm] -scoped notation "eVar[" X "]" => ProbabilityTheory.evariance X MeasureTheory.MeasureSpace.volume - @[simp] theorem variance_zero (μ : Measure Ω) : variance 0 μ = 0 := by simp only [variance, evariance_zero, ENNReal.zero_toReal] @@ -175,8 +194,6 @@ theorem variance_smul' {A : Type*} [CommSemiring A] [Algebra A ℝ] (c : A) (X : · congr; simp only [algebraMap_smul] · simp only [Algebra.smul_def, map_pow] -scoped notation "Var[" X "]" => ProbabilityTheory.variance X MeasureTheory.MeasureSpace.volume - theorem variance_def' [IsProbabilityMeasure μ] {X : Ω → ℝ} (hX : Memℒp X 2 μ) : variance X μ = μ[X ^ 2] - μ[X] ^ 2 := by rw [hX.variance_eq, sub_sq', integral_sub', integral_add']; rotate_left @@ -363,7 +380,7 @@ lemma variance_le_sub_mul_sub [IsProbabilityMeasure μ] {a b : ℝ} {X : Ω → _ = ∫ ω, - X ω ^ 2 + (a + b) * X ω - a * b ∂μ := integral_congr_ae <| ae_of_all μ fun ω ↦ by ring _ = ∫ ω, - X ω ^ 2 + (a + b) * X ω ∂μ - ∫ _, a * b ∂μ := - integral_sub (hX_int₂.add hX_int₁) (integrable_const (a * b)) + integral_sub (by fun_prop) (integrable_const (a * b)) _ = ∫ ω, - X ω ^ 2 + (a + b) * X ω ∂μ - a * b := by simp _ = - μ[X ^ 2] + (a + b) * μ[X] - a * b := by simp [← integral_neg, ← integral_mul_left, integral_add hX_int₂ hX_int₁] diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index cb98ce35c607a..6dc6591865ab8 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -98,10 +98,9 @@ section Orthogonality variable {G : Grp.{u}} [IsAlgClosed k] -open scoped Classical - variable [Fintype G] [Invertible (Fintype.card G : k)] +open scoped Classical in /-- Orthogonality of characters for irreducible representations of finite group over an algebraically closed field whose characteristic doesn't divide the order of the group. -/ theorem char_orthonormal (V W : FDRep k G) [Simple V] [Simple W] : @@ -116,7 +115,7 @@ 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, ← FDRep.endMulEquiv_comp_ρ (of _), + rw [average_char_eq_finrank_invariants, ← FDRep.endRingEquiv_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)`. diff --git a/Mathlib/RepresentationTheory/FDRep.lean b/Mathlib/RepresentationTheory/FDRep.lean index 8eaedce3ed27a..dae66e0f2aa47 100644 --- a/Mathlib/RepresentationTheory/FDRep.lean +++ b/Mathlib/RepresentationTheory/FDRep.lean @@ -82,15 +82,16 @@ 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 := - (ModuleCat.endMulEquiv _).toMonoidHom.comp (Action.ρ V) + (ModuleCat.endRingEquiv _).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 +lemma endRingEquiv_symm_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endRingEquiv 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 +lemma endRingEquiv_comp_ρ (V : FDRep k G) : + (MonoidHomClass.toMonoidHom (ModuleCat.endRingEquiv 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 @@ -111,7 +112,7 @@ theorem Iso.conj_ρ {V W : FDRep k G} (i : V ≅ W) (g : G) : @[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, ρ ≫ MonCat.ofHom (ModuleCat.endMulEquiv _).symm.toMonoidHom⟩ + ⟨FGModuleCat.of k V, ρ ≫ MonCat.ofHom (ModuleCat.endRingEquiv _).symm.toMonoidHom⟩ instance : HasForget₂ (FDRep k G) (Rep k G) where forget₂ := (forget₂ (FGModuleCat k) (ModuleCat k)).mapAction (MonCat.of G) @@ -128,12 +129,11 @@ example : MonoidalLinear k (FDRep k G) := by infer_instance open Module -open scoped Classical - -- We need to provide this instance explicitly as otherwise `finrank_hom_simple_simple` gives a -- deterministic timeout. instance : HasKernels (FDRep k G) := by infer_instance +open scoped Classical in /-- Schur's Lemma: the dimension of the `Hom`-space between two irreducible representation is `0` if they are not isomorphic, and `1` if they are. -/ theorem finrank_hom_simple_simple [IsAlgClosed k] (V W : FDRep k G) [Simple V] [Simple W] : diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean index 756771dc9da34..c620c7d2c03c8 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Basic.lean @@ -146,8 +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.hom_comp, LinearMap.comp_apply, - resolution.d_eq] + rw [diagonalHomEquiv_apply, Action.comp_hom, ConcreteCategory.comp_apply, resolution.d_eq] erw [resolution.d_of (Fin.partialProd g)] simp only [map_sum, ← Finsupp.smul_single_one _ ((-1 : k) ^ _)] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 @@ -188,8 +187,9 @@ noncomputable abbrev inhomogeneousCochains : CochainComplex (ModuleCat k) ℕ := ((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_hom, LinearEquiv.toModuleIso_hom_hom, - ModuleCat.hom_comp, LinearMap.comp_apply, LinearEquiv.coe_coe, ModuleCat.hom_zero] + simp only [d_eq, LinearEquiv.toModuleIso_inv, LinearEquiv.toModuleIso_hom, + ModuleCat.hom_comp, LinearMap.comp_apply, LinearEquiv.coe_coe, ModuleCat.hom_zero, + ModuleCat.ofHom_comp] /- 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`... -/ @@ -213,11 +213,11 @@ def inhomogeneousCochainsIso : inhomogeneousCochains A ≅ linearYonedaObjResolu 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] + LinearEquiv.toModuleIso_inv, ChainComplex.linearYonedaObj_d, ModuleCat.hom_comp, + ModuleCat.hom_ofHom, LinearMap.coe_comp, Function.comp_apply, Linear.leftComp_apply, + inhomogeneousCochains.d_def, d_eq, LinearEquiv.toModuleIso_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 diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 1c8844cd09b48..ccb8715cbcaa6 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -205,7 +205,7 @@ theorem dTwo_comp_dOne : dTwo A ∘ₗ dOne A = 0 := by 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 [ModuleCat.ofHom_comp, ModuleCat.ofHom_comp, ← LinearEquiv.toModuleIso_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] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index b53ab4e4d4058..ad038c15186b4 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -494,7 +494,7 @@ theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = /- 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] + simp only [ModuleCat.hom_smul, SimplexCategory.len_mk, ModuleCat.hom_ofHom] erw [LinearMap.smul_apply] rw [Finsupp.lmapDomain_apply, Finsupp.mapDomain_single, Finsupp.smul_single', mul_one] rfl @@ -560,7 +560,8 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : LinearMap.id fun i => rfl, LinearMap.id_comp] rfl - · congr + · rw [ModuleCat.hom_comp] + congr · ext x dsimp (config := { unfoldPartialApp := true }) [HomotopyEquiv.ofIso, Finsupp.LinearEquiv.finsuppUnique] diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 68e82a8670b12..5e1201f82dfad 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.ρ` - (ModuleCat.endMulEquiv V.V).toMonoidHom.comp (Action.ρ V) + (ModuleCat.endRingEquiv 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, MonCat.ofHom ((ModuleCat.endMulEquiv _).symm.toMonoidHom.comp ρ) ⟩ + ⟨ModuleCat.of k V, MonCat.ofHom ((ModuleCat.endRingEquiv _).symm.toMonoidHom.comp ρ) ⟩ @[simp] theorem coe_of {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) : @@ -77,7 +77,7 @@ theorem of_ρ {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k rfl theorem Action_ρ_eq_ρ {A : Rep k G} : - Action.ρ A = (ModuleCat.endMulEquiv _).symm.toMonoidHom.comp A.ρ := + Action.ρ A = (ModuleCat.endRingEquiv _).symm.toMonoidHom.comp A.ρ := rfl @[simp] @@ -296,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_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, + erw [leftRegularHom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul, A.ρ.map_one, LinearMap.one_apply] rw [zero_smul] @@ -314,9 +314,10 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ 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 (ModuleCat.hom_ext_iff.mp (f.comm x)) (Finsupp.single 1 1) - simp only [leftRegularHom_hom_hom, LinearMap.comp_apply, Finsupp.lsingle_apply, + 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] + smul_eq_mul, mul_one, zero_smul, Finsupp.sum_single_index, one_smul, + ConcreteCategory.hom_ofHom] -- Mismatched `Zero k` instances rfl right_inv x := leftRegularHom_apply x @@ -324,7 +325,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_hom, Finsupp.lift_apply, + erw [leftRegularHomEquiv_symm_apply, leftRegularHom_hom, Finsupp.lift_apply, Finsupp.sum_single_index, one_smul] rw [zero_smul] diff --git a/Mathlib/RingTheory/AdicCompletion/Algebra.lean b/Mathlib/RingTheory/AdicCompletion/Algebra.lean index 88851779d0e30..732a04c01ad0d 100644 --- a/Mathlib/RingTheory/AdicCompletion/Algebra.lean +++ b/Mathlib/RingTheory/AdicCompletion/Algebra.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Judith Ludwig, Christian Merten. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Judith Ludwig, Christian Merten -/ -import Mathlib.RingTheory.AdicCompletion.Basic import Mathlib.Algebra.Module.Torsion +import Mathlib.Algebra.Algebra.Pi +import Mathlib.RingTheory.AdicCompletion.Basic /-! # Algebra instance on adic completion diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index 2c56bfd1a747e..94d41ee47a20a 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -256,10 +256,10 @@ private def secondRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := include hf private lemma firstRow_exact : (firstRow I M f).Exact where - zero k _ := match k with - | 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) + zero + | 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 @@ -268,10 +268,10 @@ private lemma firstRow_exact : (firstRow I M f).Exact where | 2 => intro _ _; exact ⟨0, rfl⟩ private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M f).Exact where - zero k _ := match k with - | 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) + zero + | 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 diff --git a/Mathlib/RingTheory/Adjoin/FG.lean b/Mathlib/RingTheory/Adjoin/FG.lean index b98e44e79d735..1ea80e3f8b9f8 100644 --- a/Mathlib/RingTheory/Adjoin/FG.lean +++ b/Mathlib/RingTheory/Adjoin/FG.lean @@ -129,11 +129,10 @@ theorem FG.prod {S : Subalgebra R A} {T : Subalgebra R B} (hS : S.FG) (hT : T.FG section -open scoped Classical - -theorem FG.map {S : Subalgebra R A} (f : A →ₐ[R] B) (hs : S.FG) : (S.map f).FG := +theorem FG.map {S : Subalgebra R A} (f : A →ₐ[R] B) (hs : S.FG) : (S.map f).FG := by let ⟨s, hs⟩ := hs - ⟨s.image f, by rw [Finset.coe_image, Algebra.adjoin_image, hs]⟩ + classical + exact ⟨s.image f, by rw [Finset.coe_image, Algebra.adjoin_image, hs]⟩ end diff --git a/Mathlib/RingTheory/Adjoin/Tower.lean b/Mathlib/RingTheory/Adjoin/Tower.lean index c916d29572648..f0bec2c20ba93 100644 --- a/Mathlib/RingTheory/Adjoin/Tower.lean +++ b/Mathlib/RingTheory/Adjoin/Tower.lean @@ -58,17 +58,19 @@ end Algebra section -open scoped Classical - theorem Algebra.fg_trans' {R S A : Type*} [CommSemiring R] [CommSemiring S] [Semiring A] [Algebra R S] [Algebra S A] [Algebra R A] [IsScalarTower R S A] (hRS : (⊤ : Subalgebra R S).FG) - (hSA : (⊤ : Subalgebra S A).FG) : (⊤ : Subalgebra R A).FG := - let ⟨s, hs⟩ := hRS - let ⟨t, ht⟩ := hSA - ⟨s.image (algebraMap S A) ∪ t, by - rw [Finset.coe_union, Finset.coe_image, Algebra.adjoin_algebraMap_image_union_eq_adjoin_adjoin, - hs, Algebra.adjoin_top, ht, Subalgebra.restrictScalars_top, Subalgebra.restrictScalars_top]⟩ - + (hSA : (⊤ : Subalgebra S A).FG) : (⊤ : Subalgebra R A).FG := by + classical + rcases hRS with ⟨s, hs⟩ + rcases hSA with ⟨t, ht⟩ + exact ⟨s.image (algebraMap S A) ∪ t, by + rw [Finset.coe_union, Finset.coe_image, + Algebra.adjoin_algebraMap_image_union_eq_adjoin_adjoin, + hs, Algebra.adjoin_top, ht, Subalgebra.restrictScalars_top, + Subalgebra.restrictScalars_top + ] + ⟩ end section ArtinTate @@ -82,8 +84,6 @@ variable [Algebra A B] [Algebra B C] [Algebra A C] [IsScalarTower A B C] open Finset Submodule -open scoped Classical - theorem exists_subalgebra_of_fg (hAC : (⊤ : Subalgebra A C).FG) (hBC : (⊤ : Submodule B C).FG) : ∃ B₀ : Subalgebra A B, B₀.FG ∧ (⊤ : Submodule B₀ C).FG := by cases' hAC with x hx @@ -91,6 +91,7 @@ theorem exists_subalgebra_of_fg (hAC : (⊤ : Subalgebra A C).FG) (hBC : (⊤ : have := hy simp_rw [eq_top_iff', mem_span_finset] at this choose f hf using this + classical let s : Finset B := Finset.image₂ f (x ∪ y * y) y have hxy : ∀ xi ∈ x, xi ∈ span (Algebra.adjoin A (↑s : Set B)) (↑(insert 1 y : Finset C) : Set C) := @@ -140,7 +141,8 @@ variable [Algebra A B] [Algebra B C] [Algebra A C] [IsScalarTower A B C] A is noetherian, and C is algebra-finite over A, and C is module-finite over B, then B is algebra-finite over A. -References: Atiyah--Macdonald Proposition 7.8; Stacks 00IS; Altman--Kleiman 16.17. -/ +References: Atiyah--Macdonald Proposition 7.8; Altman--Kleiman 16.17. -/ +@[stacks 00IS] theorem fg_of_fg_of_fg [IsNoetherianRing A] (hAC : (⊤ : Subalgebra A C).FG) (hBC : (⊤ : Submodule B C).FG) (hBCi : Function.Injective (algebraMap B C)) : (⊤ : Subalgebra A B).FG := diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index 6d04333d192e0..b861c1ecbe216 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -48,8 +48,6 @@ The main definitions are in the `AdjoinRoot` namespace. noncomputable section -open scoped Classical - open Polynomial universe u v w diff --git a/Mathlib/RingTheory/AlgebraTower.lean b/Mathlib/RingTheory/AlgebraTower.lean index cd4e432e660f0..ef70b261a1645 100644 --- a/Mathlib/RingTheory/AlgebraTower.lean +++ b/Mathlib/RingTheory/AlgebraTower.lean @@ -85,8 +85,6 @@ section Semiring open Finsupp -open scoped Classical - variable {R S A} variable [Semiring R] [Semiring S] [AddCommMonoid A] variable [Module R S] [Module S A] [Module R A] [IsScalarTower R S A] @@ -132,6 +130,7 @@ theorem Basis.smulTower_repr_mk (x i j) : (b.smulTower c).repr x (i, j) = b.repr @[simp] theorem Basis.smulTower_apply (ij) : (b.smulTower c) ij = b ij.1 • c ij.2 := by + classical obtain ⟨i, j⟩ := ij rw [Basis.apply_eq_iff] ext ⟨i', j'⟩ diff --git a/Mathlib/RingTheory/Algebraic/Integral.lean b/Mathlib/RingTheory/Algebraic/Integral.lean index 5a9c3fa6e04f6..625d736514f0f 100644 --- a/Mathlib/RingTheory/Algebraic/Integral.lean +++ b/Mathlib/RingTheory/Algebraic/Integral.lean @@ -211,6 +211,16 @@ section trans variable (R) [NoZeroDivisors S] (inj : Function.Injective (algebraMap S A)) include inj +/-! +The next theorem may fail if only `R` is assumed to be a domain but `S` is not: for example, let +`S = R[X] ⧸ (X² - X)` and let `A` be the subalgebra of `S[Y]` generated by `XY`. +`A` is algebraic over `S` because any element `∑ᵢ sᵢ(XY)ⁱ` is a root of the polynomial +`(X - 1)(Z - s₀)` in `S[Z]`, because `X(X - 1) = X² - X = 0` in `S`. +However, `XY` is a transcendental element in `A` over `R`, because `∑ᵢ rᵢ(XY)ⁱ = 0` in `S[Y]` +implies all `rᵢXⁱ = 0` (i.e., `r₀ = 0` and `rᵢX = 0` for `i > 0`) in `S`, +which implies `rᵢ = 0` in `R`. This example is inspired by the comment +https://mathoverflow.net/questions/482944/when-do-algebraic-elements-form-a-subalgebra#comment1257632_482944. -/ + 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) @@ -358,6 +368,9 @@ theorem Algebra.isAlgebraic_adjoin_of_nonempty [NoZeroDivisors R] {s : Set S} (h have := (isDomain_iff_noZeroDivisors_and_nontrivial _).mpr ⟨‹_›, (h x hx).nontrivial⟩ isAlgebraic_adjoin_iff.mpr h⟩ +/-- In an algebra generated by a single algebraic element over a domain `R`, every element is +algebraic. This may fail when `R` is not a domain: see https://mathoverflow.net/a/132192/ for +an example. -/ 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 diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean b/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean index 46cb51be48adf..66a7c7a829c69 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/Adjoin.lean @@ -25,8 +25,6 @@ noncomputable section open Function Set Subalgebra MvPolynomial Algebra -open scoped Classical - namespace AlgebraicIndependent variable {ι : Type*} diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean b/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean index 0660623b33d9f..6f7f1983917b2 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/Basic.lean @@ -33,8 +33,6 @@ noncomputable section open Function Set Subalgebra MvPolynomial Algebra -open scoped Classical - variable {ι ι' R K A A' : Type*} {x : ι → A} variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] diff --git a/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean b/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean index c3d9b6df75314..83bd22ac7b31c 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent/Defs.lean @@ -41,8 +41,6 @@ noncomputable section open Function Set Subalgebra MvPolynomial Algebra -open scoped Classical - variable {ι ι' : Type*} (R : Type*) {K A A' : Type*} (x : ι → A) variable [CommRing R] [CommRing A] [CommRing A'] [Algebra R A] [Algebra R A'] diff --git a/Mathlib/RingTheory/Artinian/Instances.lean b/Mathlib/RingTheory/Artinian/Instances.lean index a5b60aa963ede..6cd3c5974ed97 100644 --- a/Mathlib/RingTheory/Artinian/Instances.lean +++ b/Mathlib/RingTheory/Artinian/Instances.lean @@ -19,7 +19,7 @@ namespace IsArtinianRing variable (R : Type*) [CommRing R] [IsArtinianRing R] [IsReduced R] -attribute [local instance] subtype_isMaximal_finite fieldOfSubtypeIsMaximal +attribute [local instance] fieldOfSubtypeIsMaximal instance : DecompositionMonoid R := MulEquiv.decompositionMonoid (equivPi R) diff --git a/Mathlib/RingTheory/Artinian/Module.lean b/Mathlib/RingTheory/Artinian/Module.lean index 60a200a8eda6d..17341d196f991 100644 --- a/Mathlib/RingTheory/Artinian/Module.lean +++ b/Mathlib/RingTheory/Artinian/Module.lean @@ -9,6 +9,8 @@ import Mathlib.RingTheory.Ideal.Prod import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Nilpotent.Lemmas import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Spectrum.Maximal.Basic +import Mathlib.RingTheory.Spectrum.Prime.Basic /-! # Artinian rings and modules @@ -452,7 +454,7 @@ section CommSemiring variable (R : Type*) [CommSemiring R] [IsArtinianRing R] @[stacks 00J7] -lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by +lemma setOf_isMaximal_finite : {I : Ideal R | I.IsMaximal}.Finite := by set Spec := {I : Ideal R | I.IsMaximal} obtain ⟨_, ⟨s, rfl⟩, H⟩ := IsArtinian.set_has_minimal (range (Finset.inf · Subtype.val : Finset Spec → Ideal R)) ⟨⊤, ∅, by simp⟩ @@ -462,8 +464,9 @@ lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by inf_le_right.eq_of_not_lt (H (p ⊓ s.inf Subtype.val) ⟨insert p s, by simp⟩) rwa [← Subtype.ext <| q.2.eq_of_le p.2.ne_top hq2] -lemma subtype_isMaximal_finite : Finite {I : Ideal R | I.IsMaximal} := - (maximal_ideals_finite R).to_subtype +instance : Finite (MaximalSpectrum R) := + haveI : Finite {I : Ideal R // I.IsMaximal} := (setOf_isMaximal_finite R).to_subtype + .of_equiv _ (MaximalSpectrum.equivSubtype _).symm end CommSemiring @@ -488,30 +491,56 @@ instance isMaximal_of_isPrime (p : Ideal R) [p.IsPrime] : p.IsMaximal := lemma isPrime_iff_isMaximal (p : Ideal R) : p.IsPrime ↔ p.IsMaximal := ⟨fun _ ↦ isMaximal_of_isPrime p, fun h ↦ h.isPrime⟩ +/-- The prime spectrum is in bijection with the maximal spectrum. -/ +@[simps] +def primeSpectrumEquivMaximalSpectrum : PrimeSpectrum R ≃ MaximalSpectrum R where + toFun I := ⟨I.asIdeal, isPrime_iff_isMaximal I.asIdeal |>.mp I.isPrime⟩ + invFun I := ⟨I.asIdeal, isPrime_iff_isMaximal I.asIdeal |>.mpr I.isMaximal⟩ + left_inv _ := rfl + right_inv _ := rfl + +lemma primeSpectrumEquivMaximalSpectrum_comp_asIdeal : + MaximalSpectrum.asIdeal ∘ primeSpectrumEquivMaximalSpectrum = + PrimeSpectrum.asIdeal (R := R) := rfl + +lemma primeSpectrumEquivMaximalSpectrum_symm_comp_asIdeal : + PrimeSpectrum.asIdeal ∘ primeSpectrumEquivMaximalSpectrum.symm = + MaximalSpectrum.asIdeal (R := R) := rfl + +lemma primeSpectrum_asIdeal_range_eq : + range PrimeSpectrum.asIdeal = (range <| MaximalSpectrum.asIdeal (R := R)) := by + simp only [PrimeSpectrum.range_asIdeal, MaximalSpectrum.range_asIdeal, + isPrime_iff_isMaximal] + variable (R) -lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by - simpa only [isPrime_iff_isMaximal] using maximal_ideals_finite R +lemma setOf_isPrime_finite : {I : Ideal R | I.IsPrime}.Finite := by + simpa only [isPrime_iff_isMaximal] using setOf_isMaximal_finite R -attribute [local instance] subtype_isMaximal_finite +instance : Finite (PrimeSpectrum R) := + haveI : Finite {I : Ideal R // I.IsPrime} := (setOf_isPrime_finite R).to_subtype + .of_equiv _ (PrimeSpectrum.equivSubtype _).symm /-- A temporary field instance on the quotients by maximal ideals. -/ @[local instance] noncomputable def fieldOfSubtypeIsMaximal - (I : {I : Ideal R | I.IsMaximal}) : Field (R ⧸ I.1) := - have := mem_setOf.mp I.2; Ideal.Quotient.field I.1 + (I : MaximalSpectrum R) : Field (R ⧸ I.asIdeal) := + Ideal.Quotient.field I.asIdeal /-- The quotient of a commutative artinian ring by its nilradical is isomorphic to a finite product of fields, namely the quotients by the maximal ideals. -/ noncomputable def quotNilradicalEquivPi : - R ⧸ nilradical R ≃+* ∀ I : {I : Ideal R | I.IsMaximal}, R ⧸ I.1 := - .trans (Ideal.quotEquivOfEq <| ext fun x ↦ by simp_rw [mem_nilradical, - nilpotent_iff_mem_prime, Submodule.mem_iInf, Subtype.forall, isPrime_iff_isMaximal, mem_setOf]) - (Ideal.quotientInfRingEquivPiQuotient _ fun I J h ↦ - Ideal.isCoprime_iff_sup_eq.mpr <| I.2.coprime_of_ne J.2 <| by rwa [Ne, Subtype.coe_inj]) + R ⧸ nilradical R ≃+* ∀ I : MaximalSpectrum R, R ⧸ I.asIdeal := + let f := MaximalSpectrum.asIdeal (R := R) + .trans + (Ideal.quotEquivOfEq <| ext fun x ↦ by + rw [PrimeSpectrum.nilradical_eq_iInf, iInf, primeSpectrum_asIdeal_range_eq]; rfl) + (Ideal.quotientInfRingEquivPiQuotient f <| fun I J h ↦ + Ideal.isCoprime_iff_sup_eq.mpr <| I.2.coprime_of_ne J.2 <| + fun hIJ ↦ h <| MaximalSpectrum.ext hIJ) /-- A reduced commutative artinian ring is isomorphic to a finite product of fields, namely the quotients by the maximal ideals. -/ -noncomputable def equivPi [IsReduced R] : R ≃+* ∀ I : {I : Ideal R | I.IsMaximal}, R ⧸ I.1 := +noncomputable def equivPi [IsReduced R] : R ≃+* ∀ I : MaximalSpectrum R, R ⧸ I.asIdeal := .trans (.symm <| .quotientBot R) <| .trans (Ideal.quotEquivOfEq (nilradical_eq_zero R).symm) (quotNilradicalEquivPi R) diff --git a/Mathlib/RingTheory/Artinian/Ring.lean b/Mathlib/RingTheory/Artinian/Ring.lean index a508c61431e11..455b601e23f61 100644 --- a/Mathlib/RingTheory/Artinian/Ring.lean +++ b/Mathlib/RingTheory/Artinian/Ring.lean @@ -3,8 +3,10 @@ Copyright (c) 2021 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Junyan Xu, Jujian Zhang -/ +import Mathlib.Algebra.Field.Equiv import Mathlib.RingTheory.Artinian.Module import Mathlib.RingTheory.Localization.Defs +import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic import Mathlib.RingTheory.Nakayama /-! @@ -83,6 +85,32 @@ theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) refine this (mul_mem_mul (mem_span_singleton_self x) ?_) rwa [← hn (n + 1) (Nat.le_succ _)] +variable (R) in +/-- Commutative artinian reduced local ring is a field. -/ +theorem isField_of_isReduced_of_isLocalRing [IsReduced R] [IsLocalRing R] : IsField R := + (IsArtinianRing.equivPi R).trans (RingEquiv.piUnique _) |>.toMulEquiv.isField + _ (Ideal.Quotient.field _).toIsField + +section IsUnit + +open nonZeroDivisors + +/-- If an element of an artinian ring is not a zero divisor then it is a unit. -/ +theorem isUnit_of_mem_nonZeroDivisors {a : R} (ha : a ∈ R⁰) : IsUnit a := + IsUnit.isUnit_iff_mulLeft_bijective.mpr <| + IsArtinian.bijective_of_injective_endomorphism (LinearMap.mulLeft R a) + fun _ _ ↦ (mul_cancel_left_mem_nonZeroDivisors ha).mp + +/-- In an artinian ring, an element is a unit iff it is a non-zero-divisor. +See also `isUnit_iff_mem_nonZeroDivisors_of_finite`.-/ +theorem isUnit_iff_mem_nonZeroDivisors {a : R} : IsUnit a ↔ a ∈ R⁰ := + ⟨IsUnit.mem_nonZeroDivisors, isUnit_of_mem_nonZeroDivisors⟩ + +theorem isUnit_submonoid_eq : IsUnit.submonoid R = R⁰ := by + ext; simp [IsUnit.mem_submonoid_iff, isUnit_iff_mem_nonZeroDivisors] + +end IsUnit + section Localization variable (S : Submonoid R) (L : Type*) [CommSemiring L] [Algebra R L] [IsLocalization S L] diff --git a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean index dafc8270e1007..e8efc2bbe44ae 100644 --- a/Mathlib/RingTheory/Coalgebra/TensorProduct.lean +++ b/Mathlib/RingTheory/Coalgebra/TensorProduct.lean @@ -57,7 +57,7 @@ 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_hom, + ModuleCat.MonoidalCategory.instMonoidalCategoryStruct_tensorHom, ModuleCat.hom_comp, ModuleCat.of, ModuleCat.ofHom, ModuleCat.MonoidalCategory.tensorμ_eq_tensorTensorTensorComm] } diff --git a/Mathlib/RingTheory/Congruence/Defs.lean b/Mathlib/RingTheory/Congruence/Defs.lean index 106c5bf38657e..79d50bcfd7bb2 100644 --- a/Mathlib/RingTheory/Congruence/Defs.lean +++ b/Mathlib/RingTheory/Congruence/Defs.lean @@ -6,6 +6,7 @@ Authors: Eric Wieser import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.InjSurj import Mathlib.GroupTheory.Congruence.Defs +import Mathlib.Tactic.FastInstance /-! # Congruence relations on rings @@ -320,45 +321,47 @@ The operations above on the quotient by `c : RingCon R` preserve the algebraic s section Algebraic -instance [NonUnitalNonAssocSemiring R] (c : RingCon R) : NonUnitalNonAssocSemiring c.Quotient := +instance [NonUnitalNonAssocSemiring R] (c : RingCon R) : + NonUnitalNonAssocSemiring c.Quotient := fast_instance% Function.Surjective.nonUnitalNonAssocSemiring _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -instance [NonAssocSemiring R] (c : RingCon R) : NonAssocSemiring c.Quotient := +instance [NonAssocSemiring R] (c : RingCon R) : NonAssocSemiring c.Quotient := fast_instance% Function.Surjective.nonAssocSemiring _ Quotient.mk''_surjective rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl -instance [NonUnitalSemiring R] (c : RingCon R) : NonUnitalSemiring c.Quotient := +instance [NonUnitalSemiring R] (c : RingCon R) : NonUnitalSemiring c.Quotient := fast_instance% Function.Surjective.nonUnitalSemiring _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -instance [Semiring R] (c : RingCon R) : Semiring c.Quotient := +instance [Semiring R] (c : RingCon R) : Semiring c.Quotient := fast_instance% Function.Surjective.semiring _ Quotient.mk''_surjective rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl -instance [CommSemiring R] (c : RingCon R) : CommSemiring c.Quotient := +instance [CommSemiring R] (c : RingCon R) : CommSemiring c.Quotient := fast_instance% Function.Surjective.commSemiring _ Quotient.mk''_surjective rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ => rfl -instance [NonUnitalNonAssocRing R] (c : RingCon R) : NonUnitalNonAssocRing c.Quotient := +instance [NonUnitalNonAssocRing R] (c : RingCon R) : + NonUnitalNonAssocRing c.Quotient := fast_instance% Function.Surjective.nonUnitalNonAssocRing _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -instance [NonAssocRing R] (c : RingCon R) : NonAssocRing c.Quotient := +instance [NonAssocRing R] (c : RingCon R) : NonAssocRing c.Quotient := fast_instance% Function.Surjective.nonAssocRing _ Quotient.mk''_surjective rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl -instance [NonUnitalRing R] (c : RingCon R) : NonUnitalRing c.Quotient := +instance [NonUnitalRing R] (c : RingCon R) : NonUnitalRing c.Quotient := fast_instance% Function.Surjective.nonUnitalRing _ Quotient.mk''_surjective rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -instance [Ring R] (c : RingCon R) : Ring c.Quotient := +instance [Ring R] (c : RingCon R) : Ring c.Quotient := fast_instance% Function.Surjective.ring _ Quotient.mk''_surjective rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl -instance [CommRing R] (c : RingCon R) : CommRing c.Quotient := +instance [CommRing R] (c : RingCon R) : CommRing c.Quotient := fast_instance% Function.Surjective.commRing _ Quotient.mk''_surjective rfl rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) fun _ => rfl diff --git a/Mathlib/RingTheory/Coprime/Lemmas.lean b/Mathlib/RingTheory/Coprime/Lemmas.lean index 4ccc44de66126..b4e4211ac4ea8 100644 --- a/Mathlib/RingTheory/Coprime/Lemmas.lean +++ b/Mathlib/RingTheory/Coprime/Lemmas.lean @@ -33,8 +33,8 @@ section theorem Int.isCoprime_iff_gcd_eq_one {m n : ℤ} : IsCoprime m n ↔ Int.gcd m n = 1 := by constructor · rintro ⟨a, b, h⟩ - have : 1 = m * a + n * b := by rwa [mul_comm m, mul_comm n, eq_comm] - exact Nat.dvd_one.mp (Int.gcd_dvd_iff.mpr ⟨a, b, this⟩) + refine Nat.dvd_one.mp (Int.gcd_dvd_iff.mpr ⟨a, b, ?_⟩) + rwa [mul_comm m, mul_comm n, eq_comm] · rw [← Int.ofNat_inj, IsCoprime, Int.gcd_eq_gcd_ab, mul_comm m, mul_comm n, Nat.cast_one] intro h exact ⟨_, _, h⟩ diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index 110ec9cd323fb..d7d5139cc0b3e 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -403,6 +403,9 @@ def adicCompletion := instance : Field (v.adicCompletion K) := inferInstanceAs <| Field (@UniformSpace.Completion K v.adicValued.toUniformSpace) +instance : Algebra K (v.adicCompletion K) := + RingHom.toAlgebra (@UniformSpace.Completion.coeRingHom K _ v.adicValued.toUniformSpace _ _) + instance : Inhabited (v.adicCompletion K) := ⟨0⟩ diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 68c9890c4641f..6092f3c84aea4 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -5,7 +5,7 @@ Authors: Kenji Nakagawa, Anne Baanen, Filippo A. E. Nuccio -/ import Mathlib.Algebra.Algebra.Subalgebra.Pointwise import Mathlib.Algebra.Polynomial.FieldDivision -import Mathlib.RingTheory.Spectrum.Maximal.Basic +import Mathlib.RingTheory.Spectrum.Maximal.Localization import Mathlib.RingTheory.ChainOfDivisors import Mathlib.RingTheory.DedekindDomain.Basic import Mathlib.RingTheory.FractionalIdeal.Operations @@ -255,7 +255,10 @@ theorem isDedekindDomainInv_iff [Algebra A K] [IsFractionRing A K] : theorem FractionalIdeal.adjoinIntegral_eq_one_of_isUnit [Algebra A K] [IsFractionRing A K] (x : K) (hx : IsIntegral A x) (hI : IsUnit (adjoinIntegral A⁰ x hx)) : adjoinIntegral A⁰ x hx = 1 := by set I := adjoinIntegral A⁰ x hx - have mul_self : I * I = I := by apply coeToSubmodule_injective; simp [I] + have mul_self : IsIdempotentElem I := by + apply coeToSubmodule_injective + simp only [coe_mul, adjoinIntegral_coe, I] + rw [(Algebra.adjoin A {x}).isIdempotentElem_toSubmodule] convert congr_arg (· * I⁻¹) mul_self <;> simp only [(mul_inv_cancel_iff_isUnit K).mpr hI, mul_assoc, mul_one] @@ -947,7 +950,6 @@ variable [IsDedekindDomain R] /-- The height one prime spectrum of a Dedekind domain `R` is the type of nonzero prime ideals of `R`. Note that this equals the maximal spectrum if `R` has Krull dimension 1. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `has_nonempty_instance`, linter doesn't exist yet @[ext, nolint unusedArguments] structure HeightOneSpectrum where asIdeal : Ideal R diff --git a/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean b/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean index 7dc77adaeda72..95dc12c2a9341 100644 --- a/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean +++ b/Mathlib/RingTheory/DedekindDomain/SelmerGroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ import Mathlib.Algebra.Group.Equiv.TypeTags -import Mathlib.Data.ZMod.Quotient +import Mathlib.Data.ZMod.QuotientGroup import Mathlib.RingTheory.DedekindDomain.AdicValuation import Mathlib.Algebra.Group.Int.TypeTags diff --git a/Mathlib/RingTheory/Etale/Basic.lean b/Mathlib/RingTheory/Etale/Basic.lean index ca33968558ff3..1b85886b995e0 100644 --- a/Mathlib/RingTheory/Etale/Basic.lean +++ b/Mathlib/RingTheory/Etale/Basic.lean @@ -38,10 +38,8 @@ variable (R : Type u) [CommRing R] variable (A : Type u) [CommRing A] [Algebra R A] /-- An `R` algebra `A` is formally étale if for every `R`-algebra, every square-zero ideal -`I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists exactly one lift `A →ₐ[R] B`. - -See -/ -@[mk_iff] +`I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists exactly one lift `A →ₐ[R] B`. -/ +@[mk_iff, stacks 00UQ] class FormallyEtale : Prop where comp_bijective : ∀ ⦃B : Type u⦄ [CommRing B], @@ -169,11 +167,9 @@ section variable (R : Type u) [CommRing R] variable (A : Type u) [CommRing A] [Algebra R A] -/-- An `R`-algebra `A` is étale if it is formally étale and of finite presentation. - -Note that the definition in the stacks project is -different, but shows that it is equivalent -to the definition here. -/ +/-- An `R`-algebra `A` is étale if it is formally étale and of finite presentation. -/ +@[stacks 00U1 "Note that this is a different definition from this Stacks entry, but + shows that it is equivalent to the definition here."] class Etale : Prop where formallyEtale : FormallyEtale R A := by infer_instance finitePresentation : FinitePresentation R A := by infer_instance @@ -214,3 +210,24 @@ theorem of_isLocalization_Away (r : R) [IsLocalization.Away r A] : Etale R A whe end Etale end Algebra + +namespace RingHom + +variable {R S : Type u} [CommRing R] [CommRing S] + +/-- +A ring homomorphism `R →+* A` is formally etale if it is formally unramified and formally smooth. +See `Algebra.FormallyEtale`. +-/ +@[algebraize Algebra.FormallyEtale] +def FormallyEtale (f : R →+* S) : Prop := + letI := f.toAlgebra + Algebra.FormallyEtale R S + +lemma formallyEtale_algebraMap [Algebra R S] : + (algebraMap R S).FormallyEtale ↔ Algebra.FormallyEtale R S := by + delta FormallyEtale + congr! + exact Algebra.algebra_ext _ _ fun _ ↦ rfl + +end RingHom diff --git a/Mathlib/RingTheory/Etale/Field.lean b/Mathlib/RingTheory/Etale/Field.lean index be84c58cf9418..da08f62a5267c 100644 --- a/Mathlib/RingTheory/Etale/Field.lean +++ b/Mathlib/RingTheory/Etale/Field.lean @@ -153,8 +153,7 @@ theorem iff_isSeparable [EssFiniteType K L] : FormallyEtale K L ↔ Algebra.IsSeparable K L := ⟨fun _ ↦ FormallyUnramified.isSeparable K L, fun _ ↦ of_isSeparable K L⟩ -attribute [local instance] - IsArtinianRing.subtype_isMaximal_finite IsArtinianRing.fieldOfSubtypeIsMaximal in +attribute [local instance] IsArtinianRing.fieldOfSubtypeIsMaximal in /-- If `A` is an essentially of finite type algebra over a field `K`, then `A` is formally étale over `K` if and only if `A` is a finite product of separable field extensions. @@ -170,8 +169,8 @@ theorem iff_exists_algEquiv_prod [EssFiniteType K A] : have := FormallyUnramified.finite_of_free K A have := FormallyUnramified.isReduced_of_field K A have : IsArtinianRing A := isArtinian_of_tower K inferInstance - letI : Fintype {I : Ideal A | I.IsMaximal} := (nonempty_fintype _).some - let v (i : {I : Ideal A | I.IsMaximal}) : A := (IsArtinianRing.equivPi A).symm (Pi.single i 1) + letI : Fintype (MaximalSpectrum A) := (nonempty_fintype _).some + let v (i : MaximalSpectrum A) : A := (IsArtinianRing.equivPi A).symm (Pi.single i 1) let e : A ≃ₐ[K] _ := { __ := IsArtinianRing.equivPi A, commutes' := fun r ↦ rfl } have := (FormallyEtale.iff_of_equiv e).mp inferInstance rw [FormallyEtale.pi_iff] at this @@ -189,8 +188,6 @@ theorem iff_exists_algEquiv_prod [EssFiniteType K A] : end Algebra.FormallyEtale -attribute [local instance] - IsArtinianRing.subtype_isMaximal_finite IsArtinianRing.fieldOfSubtypeIsMaximal in /-- `A` is étale over a field `K` if and only if `A` is a finite product of finite separable field extensions. diff --git a/Mathlib/RingTheory/Etale/Kaehler.lean b/Mathlib/RingTheory/Etale/Kaehler.lean index 509af402e838f..58bb20c889400 100644 --- a/Mathlib/RingTheory/Etale/Kaehler.lean +++ b/Mathlib/RingTheory/Etale/Kaehler.lean @@ -7,13 +7,16 @@ import Mathlib.RingTheory.Etale.Basic import Mathlib.RingTheory.Kaehler.JacobiZariski import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.Smooth.Kaehler +import Mathlib.RingTheory.Flat.Localization /-! # The differential module and etale algebras ## Main results -`KaehlerDifferential.tensorKaehlerEquivOfFormallyEtale`: -The canonical isomorphism `T ⊗[S] Ω[S⁄R] ≃ₗ[T] Ω[T⁄R]` for `T` a formally etale `S`-algebra. +- `KaehlerDifferential.tensorKaehlerEquivOfFormallyEtale`: + The canonical isomorphism `T ⊗[S] Ω[S⁄R] ≃ₗ[T] Ω[T⁄R]` for `T` a formally etale `S`-algebra. +- `Algebra.tensorH1CotangentOfIsLocalization`: + The canonical isomorphism `T ⊗[S] H¹(L_{S⁄R}) ≃ₗ[T] H¹(L_{T⁄R})` for `T` a localization of `S`. -/ universe u @@ -37,6 +40,12 @@ def KaehlerDifferential.tensorKaehlerEquivOfFormallyEtale [Algebra.FormallyEtale obtain ⟨x, rfl⟩ := (Algebra.H1Cotangent.exact_δ_mapBaseChange R S T x).mp hx rw [Subsingleton.elim x 0, map_zero] +lemma KaehlerDifferential.tensorKaehlerEquivOfFormallyEtale_symm_D_algebraMap + [Algebra.FormallyEtale S T] (s : S) : + (tensorKaehlerEquivOfFormallyEtale R S T).symm (D R T (algebraMap S T s)) = 1 ⊗ₜ D R S s := by + rw [LinearEquiv.symm_apply_eq, tensorKaehlerEquivOfFormallyEtale_apply, mapBaseChange_tmul, + one_smul, map_D] + lemma KaehlerDifferential.isBaseChange_of_formallyEtale [Algebra.FormallyEtale S T] : IsBaseChange T (map R R S T) := by show Function.Bijective _ @@ -51,3 +60,318 @@ instance KaehlerDifferential.isLocalizedModule_map (M : Submonoid S) [IsLocaliza IsLocalizedModule M (map R R S T) := have := Algebra.FormallyEtale.of_isLocalization (Rₘ := T) M (isLocalizedModule_iff_isBaseChange M T _).mpr (isBaseChange_of_formallyEtale R S T) + +namespace Algebra.Extension + +open KaehlerDifferential + +attribute [local instance] SMulCommClass.of_commMonoid + +variable {R S T} + +/-! +Suppose we have a morphism of extensions of `R`-algebras +``` +0 → J → Q → T → 0 + ↑ ↑ ↑ +0 → I → P → S → 0 +``` +-/ +variable {P : Extension.{u} R S} {Q : Extension.{u} R T} (f : P.Hom Q) + +/-- If `P → Q` is formally etale, then `T ⊗ₛ (S ⊗ₚ Ω[P/R]) ≃ T ⊗_Q Ω[Q/R]`. -/ +noncomputable +def tensorCotangentSpace + (H : f.toRingHom.FormallyEtale) : + T ⊗[S] P.CotangentSpace ≃ₗ[T] Q.CotangentSpace := + letI := f.toRingHom.toAlgebra + haveI : IsScalarTower R P.Ring Q.Ring := + .of_algebraMap_eq fun r ↦ (f.toRingHom_algebraMap r).symm + letI := ((algebraMap S T).comp (algebraMap P.Ring S)).toAlgebra + haveI : IsScalarTower P.Ring S T := .of_algebraMap_eq' rfl + haveI : IsScalarTower P.Ring Q.Ring T := + .of_algebraMap_eq fun r ↦ (f.algebraMap_toRingHom r).symm + haveI : FormallyEtale P.Ring Q.Ring := ‹_› + { __ := (CotangentSpace.map f).liftBaseChange T + invFun := LinearMap.liftBaseChange T (by + refine LinearMap.liftBaseChange _ ?_ ∘ₗ + (tensorKaehlerEquivOfFormallyEtale R P.Ring Q.Ring).symm.toLinearMap + exact (TensorProduct.mk _ _ _ 1).restrictScalars P.Ring ∘ₗ + (TensorProduct.mk _ _ _ 1).restrictScalars P.Ring) + left_inv x := by + show (LinearMap.liftBaseChange _ _ ∘ₗ LinearMap.liftBaseChange _ _) x = + LinearMap.id (R := T) x + congr 1 + ext : 4 + refine Derivation.liftKaehlerDifferential_unique + (R := R) (S := P.Ring) (M := T ⊗[S] P.CotangentSpace) _ _ ?_ + ext a + have : (tensorKaehlerEquivOfFormallyEtale R P.Ring Q.Ring).symm + ((D R Q.Ring) (f.toRingHom a)) = 1 ⊗ₜ D _ _ a := + tensorKaehlerEquivOfFormallyEtale_symm_D_algebraMap R P.Ring Q.Ring a + simp [this] + right_inv x := by + show (LinearMap.liftBaseChange _ _ ∘ₗ LinearMap.liftBaseChange _ _) x = + LinearMap.id (R := T) x + congr 1 + ext a + dsimp + obtain ⟨x, hx⟩ := (tensorKaehlerEquivOfFormallyEtale R P.Ring _).surjective (D R Q.Ring a) + simp only [one_smul, ← hx, LinearEquiv.symm_apply_apply] + show (((CotangentSpace.map f).liftBaseChange T).restrictScalars Q.Ring ∘ₗ + LinearMap.liftBaseChange _ _) x = ((TensorProduct.mk _ _ _ 1) ∘ₗ + (tensorKaehlerEquivOfFormallyEtale R P.Ring Q.Ring).toLinearMap) x + congr 1 + ext a + simp; rfl } + +/-- (Implementation) +If `J ≃ Q ⊗ₚ I` (e.g. when `T = Q ⊗ₚ S` and `P → Q` is flat), then `T ⊗ₛ I/I² ≃ J/J²`. +This is the inverse. -/ +noncomputable +def tensorCotangentInvFun + [alg : Algebra P.Ring Q.Ring] (halg : algebraMap P.Ring Q.Ring = f.toRingHom) + (H : Function.Bijective ((f.mapKer halg).liftBaseChange Q.Ring)) : + Q.Cotangent →+ T ⊗[S] P.Cotangent := + letI := ((algebraMap S T).comp (algebraMap P.Ring S)).toAlgebra + haveI : IsScalarTower P.Ring S T := .of_algebraMap_eq' rfl + haveI : IsScalarTower P.Ring Q.Ring T := + .of_algebraMap_eq fun r ↦ halg ▸ (f.algebraMap_toRingHom r).symm + letI e := LinearEquiv.ofBijective _ H + letI f' : Q.ker →ₗ[Q.Ring] T ⊗[S] P.Cotangent := + (LinearMap.liftBaseChange _ + ((TensorProduct.mk _ _ _ 1).restrictScalars _ ∘ₗ Cotangent.mk)) ∘ₗ e.symm.toLinearMap + QuotientAddGroup.lift _ f' <| by + intro x hx + refine Submodule.smul_induction_on hx ?_ fun _ _ ↦ add_mem + clear x hx + rintro a ha b - + obtain ⟨x, hx⟩ := e.surjective ⟨a, ha⟩ + obtain rfl : (e x).1 = a := congr_arg Subtype.val hx + obtain ⟨y, rfl⟩ := e.surjective b + simp only [AddMonoidHom.mem_ker, AddMonoidHom.coe_coe, map_smul, + LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply, + LinearEquiv.symm_apply_apply, f'] + clear hx ha + induction x with + | zero => simp only [LinearEquiv.map_zero, ZeroMemClass.coe_zero, zero_smul] + | add x y _ _ => + simp only [LinearEquiv.map_add, Submodule.coe_add, add_smul, zero_add, *] + | tmul a b => + induction y with + | zero => simp only [LinearEquiv.map_zero, LinearMap.map_zero, smul_zero] + | add x y hx hy => simp only [LinearMap.map_add, smul_add, hx, hy, zero_add] + | tmul c d => + simp only [LinearMap.liftBaseChange_tmul, LinearMap.coe_comp, SetLike.val_smul, + LinearMap.coe_restrictScalars, Function.comp_apply, mk_apply, smul_eq_mul, e, + LinearMap.liftBaseChange_tmul, LinearEquiv.ofBijective_apply] + have h₂ : b.1 • Cotangent.mk d = 0 := by ext; simp [Cotangent.smul_eq_zero_of_mem _ b.2] + rw [TensorProduct.smul_tmul', mul_smul, f.mapKer_apply_coe, ← halg, + algebraMap_smul, ← TensorProduct.tmul_smul, h₂, tmul_zero, smul_zero] + +omit [IsScalarTower R S T] in +lemma tensorCotangentInvFun_smul_mk + [alg : Algebra P.Ring Q.Ring] (halg : algebraMap P.Ring Q.Ring = f.toRingHom) + (H : Function.Bijective ((f.mapKer halg).liftBaseChange Q.Ring)) (x : Q.Ring) (y : P.ker) : + tensorCotangentInvFun f halg H (x • .mk ⟨f.toRingHom y, (f.mapKer halg y).2⟩) = + x • 1 ⊗ₜ .mk y := by + letI := ((algebraMap S T).comp (algebraMap P.Ring S)).toAlgebra + haveI : IsScalarTower P.Ring S T := .of_algebraMap_eq' rfl + haveI : IsScalarTower P.Ring Q.Ring T := + .of_algebraMap_eq fun r ↦ halg ▸ (f.algebraMap_toRingHom r).symm + letI e := LinearEquiv.ofBijective _ H + trans tensorCotangentInvFun f halg H (.mk ((f.mapKer halg).liftBaseChange Q.Ring (x ⊗ₜ y))) + · simp; rfl + show ((TensorProduct.mk _ _ _ 1).restrictScalars _ ∘ₗ Cotangent.mk).liftBaseChange _ + (e.symm (e (x ⊗ₜ y))) = _ + rw [e.symm_apply_apply] + simp + +/-- If `J ≃ Q ⊗ₚ I` (e.g. when `T = Q ⊗ₚ S` and `P → Q` is flat), then `T ⊗ₛ I/I² ≃ J/J²`. -/ +noncomputable +def tensorCotangent [alg : Algebra P.Ring Q.Ring] (halg : algebraMap P.Ring Q.Ring = f.toRingHom) + (H : Function.Bijective ((f.mapKer halg).liftBaseChange Q.Ring)) : + T ⊗[S] P.Cotangent ≃ₗ[T] Q.Cotangent := + { __ := (Cotangent.map f).liftBaseChange T + invFun := tensorCotangentInvFun f halg H + left_inv x := by + simp only [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom] + induction x with + | zero => simp only [map_zero] + | add x y _ _ => simp only [map_add, *] + | tmul a b => + obtain ⟨b, rfl⟩ := Cotangent.mk_surjective b + obtain ⟨a, rfl⟩ := Q.algebraMap_surjective a + simp only [LinearMap.liftBaseChange_tmul, Cotangent.map_mk, Hom.toAlgHom_apply, + algebraMap_smul, map_smul] + refine (tensorCotangentInvFun_smul_mk f halg H a b).trans ?_ + simp [algebraMap_eq_smul_one, TensorProduct.smul_tmul'] + right_inv x := by + obtain ⟨x, rfl⟩ := Cotangent.mk_surjective x + obtain ⟨x, rfl⟩ := H.surjective x + simp only [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom] + induction x with + | zero => simp only [map_zero] + | add x y _ _ => simp only [map_add, *] + | tmul a b => + simp only [LinearMap.liftBaseChange_tmul, map_smul] + erw [tensorCotangentInvFun_smul_mk] + simp + rfl } + +/-- If `J ≃ Q ⊗ₚ I`, `S → T` is flat and `P → Q` is formally etale, then `T ⊗ H¹(L_P) ≃ H¹(L_Q)`. -/ +noncomputable +def tensorH1Cotangent [alg : Algebra P.Ring Q.Ring] (halg : algebraMap P.Ring Q.Ring = f.toRingHom) + [Module.Flat S T] + (H₁ : f.toRingHom.FormallyEtale) + (H₂ : Function.Bijective ((f.mapKer halg).liftBaseChange Q.Ring)) : + T ⊗[S] P.H1Cotangent ≃ₗ[T] Q.H1Cotangent := by + refine .ofBijective ((H1Cotangent.map f).liftBaseChange T) ?_ + constructor + · rw [injective_iff_map_eq_zero] + intro x hx + apply Module.Flat.lTensor_preserves_injective_linearMap _ h1Cotangentι_injective + apply (Extension.tensorCotangent f halg H₂).injective + simp only [map_zero] + rw [← h1Cotangentι.map_zero, ← hx] + show ((Cotangent.map f).liftBaseChange T ∘ₗ h1Cotangentι.baseChange T) x = + (h1Cotangentι ∘ₗ _) x + congr 1 + ext x + simp + · intro x + have : Function.Exact (h1Cotangentι.baseChange T) (P.cotangentComplex.baseChange T) := + Module.Flat.lTensor_exact T (LinearMap.exact_subtype_ker_map _) + obtain ⟨a, ha⟩ := (this ((Extension.tensorCotangent f halg H₂).symm x.1)).mp (by + apply (Extension.tensorCotangentSpace f H₁).injective + rw [map_zero, ← x.2] + have : (CotangentSpace.map f).liftBaseChange T ∘ₗ P.cotangentComplex.baseChange T = + Q.cotangentComplex ∘ₗ (Cotangent.map f).liftBaseChange T := by + ext x; obtain ⟨x, rfl⟩ := Cotangent.mk_surjective x; dsimp + simp only [CotangentSpace.map_tmul, + map_one, Hom.toAlgHom_apply, one_smul, cotangentComplex_mk] + exact (DFunLike.congr_fun this _).trans (DFunLike.congr_arg Q.cotangentComplex + ((tensorCotangent f halg H₂).apply_symm_apply x.1))) + refine ⟨a, Subtype.ext (.trans ?_ ((LinearEquiv.eq_symm_apply _).mp ha))⟩ + show (h1Cotangentι ∘ₗ (H1Cotangent.map f).liftBaseChange T) _ = + ((Cotangent.map f).liftBaseChange T ∘ₗ h1Cotangentι.baseChange T) _ + congr 1 + ext; simp + +end Extension + +variable {S} + +/-- let `p` be a submonoid of an `R`-algebra `S`. Then `Sₚ ⊗ H¹(L_{S/R}) ≃ H¹(L_{Sₚ/R})`. -/ +noncomputable +def tensorH1CotangentOfIsLocalization (M : Submonoid S) [IsLocalization M T] : + T ⊗[S] H1Cotangent R S ≃ₗ[T] H1Cotangent R T := by + letI P : Extension R S := (Generators.self R S).toExtension + letI M' := M.comap (algebraMap P.Ring S) + letI fQ : Localization M' →ₐ[R] T := IsLocalization.liftAlgHom (M := M') + (f := (IsScalarTower.toAlgHom R S T).comp (IsScalarTower.toAlgHom R P.Ring S)) (fun ⟨y, hy⟩ ↦ + by simpa using IsLocalization.map_units T ⟨algebraMap P.Ring S y, hy⟩) + letI Q : Extension R T := .ofSurjective fQ (by + intro x + obtain ⟨x, ⟨s, hs⟩, rfl⟩ := IsLocalization.mk'_surjective M x + obtain ⟨x, rfl⟩ := P.algebraMap_surjective x + obtain ⟨s, rfl⟩ := P.algebraMap_surjective s + refine ⟨IsLocalization.mk' _ x ⟨s, show s ∈ M' from hs⟩, ?_⟩ + simp only [fQ, IsLocalization.coe_liftAlgHom, AlgHom.toRingHom_eq_coe] + rw [IsLocalization.lift_mk'_spec] + simp) + letI f : P.Hom Q := + { toRingHom := algebraMap P.Ring (Localization M') + toRingHom_algebraMap x := (IsScalarTower.algebraMap_apply R P.Ring (Localization M') _).symm + algebraMap_toRingHom x := @IsLocalization.lift_eq .. } + haveI : FormallySmooth R P.Ring := inferInstanceAs (FormallySmooth R (MvPolynomial _ _)) + haveI : FormallySmooth P.Ring (Localization M') := .of_isLocalization M' + haveI : FormallySmooth R Q.Ring := .comp R P.Ring (Localization M') + haveI : Module.Flat S T := IsLocalization.flat T M + letI : Algebra P.Ring Q.Ring := inferInstanceAs (Algebra P.Ring (Localization M')) + letI := ((algebraMap S T).comp (algebraMap P.Ring S)).toAlgebra + letI := fQ.toRingHom.toAlgebra + haveI : IsScalarTower P.Ring S T := .of_algebraMap_eq' rfl + haveI : IsScalarTower P.Ring (Localization M') T := + .of_algebraMap_eq fun r ↦ (f.algebraMap_toRingHom r).symm + haveI : IsLocalizedModule M' (IsScalarTower.toAlgHom P.Ring S T).toLinearMap := by + rw [isLocalizedModule_iff_isLocalization] + convert ‹IsLocalization M T› using 1 + exact Submonoid.map_comap_eq_of_surjective P.algebraMap_surjective _ + refine Extension.tensorH1Cotangent f rfl ?_ ?_ ≪≫ₗ Extension.equivH1CotangentOfFormallySmooth _ + · exact RingHom.formallyEtale_algebraMap.mpr + (FormallyEtale.of_isLocalization (M := M') (Rₘ := Localization M')) + · let F : P.ker →ₗ[P.Ring] RingHom.ker fQ := f.mapKer rfl + refine (isLocalizedModule_iff_isBaseChange M' (Localization M') F).mp ?_ + have : (Algebra.linearMap P.Ring S).ker.localized' (Localization M') M' + (Algebra.linearMap P.Ring (Localization M')) = RingHom.ker fQ := by + rw [LinearMap.localized'_ker_eq_ker_localizedMap (Localization M') M' + (Algebra.linearMap P.Ring (Localization M')) + (f' := (IsScalarTower.toAlgHom P.Ring S T).toLinearMap)] + ext x + obtain ⟨x, ⟨s, hs⟩, rfl⟩ := IsLocalization.mk'_surjective M' x + simp only [LinearMap.mem_ker, LinearMap.extendScalarsOfIsLocalization_apply', RingHom.mem_ker, + IsLocalization.coe_liftAlgHom, AlgHom.toRingHom_eq_coe, IsLocalization.lift_mk'_spec, + RingHom.coe_coe, AlgHom.coe_comp, IsScalarTower.coe_toAlgHom', Function.comp_apply, + mul_zero, fQ] + have : IsLocalization.mk' (Localization M') x ⟨s, hs⟩ = + IsLocalizedModule.mk' (Algebra.linearMap P.Ring (Localization M')) x ⟨s, hs⟩ := by + rw [IsLocalization.mk'_eq_iff_eq_mul, mul_comm, ← Algebra.smul_def, ← Submonoid.smul_def, + IsLocalizedModule.mk'_cancel'] + rfl + simp [this, ← IsScalarTower.algebraMap_apply] + have : F = ((LinearEquiv.ofEq _ _ this).restrictScalars P.Ring).toLinearMap ∘ₗ + P.ker.toLocalized' (Localization M') M' (Algebra.linearMap P.Ring (Localization M')) := by + ext; rfl + rw [this] + exact IsLocalizedModule.of_linearEquiv _ _ _ + +lemma tensorH1CotangentOfIsLocalization_toLinearMap + (M : Submonoid S) [IsLocalization M T] : + (tensorH1CotangentOfIsLocalization R T M).toLinearMap = + (Algebra.H1Cotangent.map R R S T).liftBaseChange T := by + ext x : 3 + simp only [AlgebraTensorModule.curry_apply, curry_apply, LinearMap.coe_restrictScalars, + LinearEquiv.coe_coe, LinearMap.liftBaseChange_tmul, one_smul] + simp only [tensorH1CotangentOfIsLocalization, Generators.toExtension_Ring, + Generators.toExtension_commRing, Generators.self_vars, Generators.toExtension_algebra₂, + Generators.self_algebra, AlgHom.toRingHom_eq_coe, Extension.tensorH1Cotangent, + LinearEquiv.ofBijective_apply, LinearMap.liftBaseChange_tmul, one_smul, + Extension.equivH1CotangentOfFormallySmooth, LinearEquiv.trans_apply] + letI P : Extension R S := (Generators.self R S).toExtension + letI M' := M.comap (algebraMap P.Ring S) + letI fQ : Localization M' →ₐ[R] T := IsLocalization.liftAlgHom (M := M') + (f := (IsScalarTower.toAlgHom R S T).comp (IsScalarTower.toAlgHom R P.Ring S)) (fun ⟨y, hy⟩ ↦ + by simpa using IsLocalization.map_units T ⟨algebraMap P.Ring S y, hy⟩) + letI Q : Extension R T := .ofSurjective fQ (by + intro x + obtain ⟨x, ⟨s, hs⟩, rfl⟩ := IsLocalization.mk'_surjective M x + obtain ⟨x, rfl⟩ := P.algebraMap_surjective x + obtain ⟨s, rfl⟩ := P.algebraMap_surjective s + refine ⟨IsLocalization.mk' _ x ⟨s, show s ∈ M' from hs⟩, ?_⟩ + simp only [fQ, IsLocalization.coe_liftAlgHom, AlgHom.toRingHom_eq_coe] + rw [IsLocalization.lift_mk'_spec] + simp) + letI f : (Generators.self R T).toExtension.Hom Q := + { toRingHom := (MvPolynomial.aeval Q.σ).toRingHom + toRingHom_algebraMap := (MvPolynomial.aeval Q.σ).commutes + algebraMap_toRingHom := by + have : (IsScalarTower.toAlgHom R Q.Ring T).comp (MvPolynomial.aeval Q.σ) = + IsScalarTower.toAlgHom _ (Generators.self R T).toExtension.Ring _ := by + ext i + show _ = algebraMap (Generators.self R T).Ring _ (.X i) + simp + exact DFunLike.congr_fun this } + rw [← Extension.H1Cotangent.equivOfFormallySmooth_symm, LinearEquiv.symm_apply_eq, + @Extension.H1Cotangent.equivOfFormallySmooth_apply (f := f), + Algebra.H1Cotangent.map, ← (Extension.H1Cotangent.map f).coe_restrictScalars S, + ← LinearMap.comp_apply, ← Extension.H1Cotangent.map_comp, Extension.H1Cotangent.map_eq] + +instance H1Cotangent.isLocalizedModule (M : Submonoid S) [IsLocalization M T] : + IsLocalizedModule M (Algebra.H1Cotangent.map R R S T) := by + rw [isLocalizedModule_iff_isBaseChange M T] + show Function.Bijective ((Algebra.H1Cotangent.map R R S T).liftBaseChange T) + rw [← tensorH1CotangentOfIsLocalization_toLinearMap R T M] + exact (tensorH1CotangentOfIsLocalization R T M).bijective + +end Algebra diff --git a/Mathlib/RingTheory/EuclideanDomain.lean b/Mathlib/RingTheory/EuclideanDomain.lean index a0cf45fae7e77..2c36ca8f2fc47 100644 --- a/Mathlib/RingTheory/EuclideanDomain.lean +++ b/Mathlib/RingTheory/EuclideanDomain.lean @@ -96,7 +96,7 @@ theorem dvd_or_coprime (x y : α) (h : Irreducible x) : x ∣ y ∨ IsCoprime x y := letI := Classical.decEq α letI := EuclideanDomain.gcdMonoid α - _root_.dvd_or_coprime x y h + _root_.dvd_or_isCoprime x y h end EuclideanDomain diff --git a/Mathlib/RingTheory/Extension.lean b/Mathlib/RingTheory/Extension.lean index b64aada4929d9..1de194d71dedc 100644 --- a/Mathlib/RingTheory/Extension.lean +++ b/Mathlib/RingTheory/Extension.lean @@ -227,6 +227,15 @@ lemma Hom.comp_id (f : Hom P P') : f.comp (Hom.id P) = f := by ext; simp lemma Hom.id_comp (f : Hom P P') : (Hom.id P').comp f = f := by ext; simp [Hom.id, aeval_X_left] +/-- A map between extensions induce a map between kernels. -/ +@[simps] +def Hom.mapKer (f : P.Hom P') + [alg : Algebra P.Ring P'.Ring] (halg : algebraMap P.Ring P'.Ring = f.toRingHom) : + P.ker →ₗ[P.Ring] P'.ker where + toFun x := ⟨f.toRingHom x, by simp [show algebraMap P.Ring S x = 0 from x.2]⟩ + map_add' _ _ := Subtype.ext (map_add _ _ _) + map_smul' := by simp [Algebra.smul_def, ← halg] + end end Hom @@ -407,6 +416,13 @@ lemma Cotangent.map_comp (f : Hom P P') (g : Hom P' P'') : simp only [map_mk, Hom.toAlgHom_apply, Hom.comp_toRingHom, RingHom.coe_comp, Function.comp_apply, val_mk, LinearMap.coe_comp, LinearMap.coe_restrictScalars] +lemma Cotangent.finite (hP : P.ker.FG) : + Module.Finite S P.Cotangent := by + refine ⟨.of_restrictScalars (R := P.Ring) _ ?_⟩ + rw [Submodule.restrictScalars_top, ← LinearMap.range_eq_top.mpr Extension.Cotangent.mk_surjective, + ← Submodule.map_top] + exact (P.ker.fg_top.mpr hP).map _ + end Cotangent end Algebra.Extension diff --git a/Mathlib/RingTheory/FinitePresentation.lean b/Mathlib/RingTheory/FinitePresentation.lean index 777df2c59cb87..dbc3a551101b0 100644 --- a/Mathlib/RingTheory/FinitePresentation.lean +++ b/Mathlib/RingTheory/FinitePresentation.lean @@ -212,8 +212,8 @@ theorem trans [Algebra A B] [IsScalarTower R A B] [FinitePresentation R A] open MvPolynomial --- We follow the proof of https://stacks.math.columbia.edu/tag/0561 -- TODO: extract out helper lemmas and tidy proof. +@[stacks 0561] theorem of_restrict_scalars_finitePresentation [Algebra A B] [IsScalarTower R A B] [FinitePresentation.{w₁, w₃} R B] [FiniteType R A] : FinitePresentation.{w₂, w₃} A B := by @@ -450,6 +450,63 @@ end FinitePresentation end RingHom +namespace RingHom.FinitePresentation +universe u v + +open Polynomial + +/-- Induction principle for finitely presented ring homomorphisms. + +For a property to hold for all finitely presented ring homs, it suffices for it to hold for +`Polynomial.C : R → R[X]`, surjective ring homs with finitely generated kernels, and to be closed +under composition. + +Note that to state this conveniently for ring homs between rings of different universes, we carry +around two predicates `P` and `Q`, which should be "the same" apart from universes: +* `P`, for ring homs `(R : Type u) → (S : Type u)`. +* `Q`, for ring homs `(R : Type u) → (S : Type v)`. +-/ +lemma polynomial_induction + (P : ∀ (R : Type u) [CommRing R] (S : Type u) [CommRing S], (R →+* S) → Prop) + (Q : ∀ (R : Type u) [CommRing R] (S : Type v) [CommRing S], (R →+* S) → Prop) + (polynomial : ∀ (R) [CommRing R], P R R[X] C) + (fg_ker : ∀ (R : Type u) [CommRing R] (S : Type v) [CommRing S] (f : R →+* S), + Surjective f → (ker f).FG → Q R S f) + (comp : ∀ (R) [CommRing R] (S) [CommRing S] (T) [CommRing T] (f : R →+* S) (g : S →+* T), + P R S f → Q S T g → Q R T (g.comp f)) + {R : Type u} {S : Type v} [CommRing R] [CommRing S] (f : R →+* S) (hf : f.FinitePresentation) : + Q R S f := by + letI := f.toAlgebra + obtain ⟨n, g, hg, hg'⟩ := hf + let g' := g.toRingHom + change Surjective g' at hg + change (ker g').FG at hg' + have : g'.comp MvPolynomial.C = f := g.comp_algebraMap + clear_value g' + subst this + clear g + induction n generalizing R S with + | zero => + refine fg_ker _ _ _ (hg.comp (MvPolynomial.C_surjective (Fin 0))) ?_ + rw [← comap_ker] + convert hg'.map (MvPolynomial.isEmptyRingEquiv R (Fin 0)).toRingHom using 1 + simp only [RingEquiv.toRingHom_eq_coe] + exact Ideal.comap_symm (MvPolynomial.isEmptyRingEquiv R (Fin 0)) + | succ n IH => + let e : MvPolynomial (Fin (n + 1)) R ≃ₐ[R] MvPolynomial (Fin n) R[X] := + (MvPolynomial.renameEquiv R (finSuccEquiv n)).trans (MvPolynomial.optionEquivRight R (Fin n)) + have he : (ker (g'.comp <| RingHomClass.toRingHom e.symm)).FG := by + rw [← RingHom.comap_ker] + convert hg'.map e.toAlgHom.toRingHom using 1 + exact Ideal.comap_symm e.toRingEquiv + have := IH (R := R[X]) (S := S) (g'.comp e.symm) (hg.comp e.symm.surjective) he + convert comp _ _ _ _ _ (polynomial _) this using 1 + rw [comp_assoc, comp_assoc] + congr 1 with r + simp [e] + +end RingHom.FinitePresentation + namespace AlgHom variable {R A B C : Type*} [CommRing R] diff --git a/Mathlib/RingTheory/FiniteType.lean b/Mathlib/RingTheory/FiniteType.lean index 40f8bd11de976..c8b20cf8bab7c 100644 --- a/Mathlib/RingTheory/FiniteType.lean +++ b/Mathlib/RingTheory/FiniteType.lean @@ -75,10 +75,10 @@ protected theorem polynomial : FiniteType R R[X] := rw [Finset.coe_singleton] exact Polynomial.adjoin_X⟩⟩ -open scoped Classical protected theorem freeAlgebra (ι : Type*) [Finite ι] : FiniteType R (FreeAlgebra R ι) := by cases nonempty_fintype ι + classical exact ⟨⟨Finset.univ.image (FreeAlgebra.ι R), by rw [Finset.coe_image, Finset.coe_univ, Set.image_univ] @@ -86,6 +86,7 @@ protected theorem freeAlgebra (ι : Type*) [Finite ι] : FiniteType R (FreeAlgeb protected theorem mvPolynomial (ι : Type*) [Finite ι] : FiniteType R (MvPolynomial ι R) := by cases nonempty_fintype ι + classical exact ⟨⟨Finset.univ.image MvPolynomial.X, by rw [Finset.coe_image, Finset.coe_univ, Set.image_univ] diff --git a/Mathlib/RingTheory/Finiteness/Nakayama.lean b/Mathlib/RingTheory/Finiteness/Nakayama.lean index a5ae263188b4d..d17b83e9cc888 100644 --- a/Mathlib/RingTheory/Finiteness/Nakayama.lean +++ b/Mathlib/RingTheory/Finiteness/Nakayama.lean @@ -26,8 +26,8 @@ variable {R : Type*} {M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] open Set -/-- **Nakayama's Lemma**. Atiyah-Macdonald 2.5, Eisenbud 4.7, Matsumura 2.2, -[Stacks 00DV](https://stacks.math.columbia.edu/tag/00DV) -/ +/-- **Nakayama's Lemma**. Atiyah-Macdonald 2.5, Eisenbud 4.7, Matsumura 2.2. -/ +@[stacks 00DV] theorem exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul {R : Type*} [CommRing R] {M : Type*} [AddCommGroup M] [Module R M] (I : Ideal R) (N : Submodule R M) (hn : N.FG) (hin : N ≤ I • N) : ∃ r : R, r - 1 ∈ I ∧ ∀ n ∈ N, r • n = (0 : M) := by diff --git a/Mathlib/RingTheory/Fintype.lean b/Mathlib/RingTheory/Fintype.lean index 9ee3bc2a315f2..e13486bc2013a 100644 --- a/Mathlib/RingTheory/Fintype.lean +++ b/Mathlib/RingTheory/Fintype.lean @@ -51,8 +51,7 @@ lemma Finset.univ_of_card_le_three (h : Fintype.card R ≤ 3) : · exact zero_ne_one h · exact zero_ne_one h.symm -open scoped Classical - +open scoped Classical in theorem card_units_lt (M₀ : Type*) [MonoidWithZero M₀] [Nontrivial M₀] [Fintype M₀] : Fintype.card M₀ˣ < Fintype.card M₀ := Fintype.card_lt_of_injective_of_not_mem Units.val Units.ext not_isUnit_zero diff --git a/Mathlib/RingTheory/FreeCommRing.lean b/Mathlib/RingTheory/FreeCommRing.lean index af33998c3ea6e..a09a60ff00f16 100644 --- a/Mathlib/RingTheory/FreeCommRing.lean +++ b/Mathlib/RingTheory/FreeCommRing.lean @@ -50,7 +50,6 @@ free commutative ring, free ring noncomputable section -open scoped Classical open Polynomial universe u v @@ -224,6 +223,7 @@ end Restriction theorem isSupported_of {p} {s : Set α} : IsSupported (of p) s ↔ p ∈ s := suffices IsSupported (of p) s → p ∈ s from ⟨this, fun hps => Subring.subset_closure ⟨p, hps, rfl⟩⟩ fun hps : IsSupported (of p) s => by + classical haveI := Classical.decPred s have : ∀ x, IsSupported x s → ∃ n : ℤ, lift (fun a => if a ∈ s then (0 : ℤ[X]) else Polynomial.X) x = n := by diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index d3a7fc8a87136..5db79f945e9dd 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -88,8 +88,6 @@ section /-- Let `x` be a submonoid of `A`, then `NumDenSameDeg 𝒜 x` is a structure with a numerator and a denominator with same grading such that the denominator is contained in `x`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure NumDenSameDeg where deg : ι (num den : 𝒜 deg) @@ -271,8 +269,6 @@ end HomogeneousLocalization kernel of `embedding 𝒜 x`. This is essentially the subring of `Aₓ` where the numerator and denominator share the same grading. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] def HomogeneousLocalization : Type _ := Quotient (Setoid.ker <| HomogeneousLocalization.NumDenSameDeg.embedding 𝒜 x) diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index b074a86e1aa3b..e29ff4fcd2b4d 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -476,10 +476,9 @@ theorem BddBelow_zero [Nonempty Γ] : BddBelow (Function.support (0 : Γ → R)) variable [LocallyFiniteOrder Γ] -theorem suppBddBelow_supp_PWO (f : Γ → R) - (hf : BddBelow (Function.support f)) : +theorem suppBddBelow_supp_PWO (f : Γ → R) (hf : BddBelow (Function.support f)) : (Function.support f).IsPWO := - Set.isWF_iff_isPWO.mp hf.wellFoundedOn_lt + hf.isWF.isPWO /-- Construct a Hahn series from any function whose support is bounded below. -/ @[simps] diff --git a/Mathlib/RingTheory/Ideal/AssociatedPrime.lean b/Mathlib/RingTheory/Ideal/AssociatedPrime.lean index 0a316e3c547c5..e274dedcc24aa 100644 --- a/Mathlib/RingTheory/Ideal/AssociatedPrime.lean +++ b/Mathlib/RingTheory/Ideal/AssociatedPrime.lean @@ -34,9 +34,11 @@ Generalize this to a non-commutative setting once there are annihilator for non- variable {R : Type*} [CommRing R] (I J : Ideal R) (M : Type*) [AddCommGroup M] [Module R M] +open LinearMap + /-- `IsAssociatedPrime I M` if the prime ideal `I` is the annihilator of some `x : M`. -/ def IsAssociatedPrime : Prop := - I.IsPrime ∧ ∃ x : M, I = (R ∙ x).annihilator + I.IsPrime ∧ ∃ x : M, I = ker (toSpanSingleton R M x) variable (R) @@ -55,8 +57,7 @@ theorem IsAssociatedPrime.map_of_injective (h : IsAssociatedPrime I M) (hf : Fun obtain ⟨x, rfl⟩ := h.2 refine ⟨h.1, ⟨f x, ?_⟩⟩ ext r - rw [Submodule.mem_annihilator_span_singleton, Submodule.mem_annihilator_span_singleton, ← - map_smul, ← f.map_zero, hf.eq_iff] + simp_rw [mem_ker, toSpanSingleton_apply, ← map_smul, ← f.map_zero, hf.eq_iff] theorem LinearEquiv.isAssociatedPrime_iff (l : M ≃ₗ[R] M') : IsAssociatedPrime I M ↔ IsAssociatedPrime I M' := @@ -66,32 +67,31 @@ theorem LinearEquiv.isAssociatedPrime_iff (l : M ≃ₗ[R] M') : theorem not_isAssociatedPrime_of_subsingleton [Subsingleton M] : ¬IsAssociatedPrime I M := by rintro ⟨hI, x, hx⟩ apply hI.ne_top - rwa [Subsingleton.elim x 0, Submodule.span_singleton_eq_bot.mpr rfl, Submodule.annihilator_bot] - at hx + simpa [Subsingleton.elim x 0] using hx variable (R) theorem exists_le_isAssociatedPrime_of_isNoetherianRing [H : IsNoetherianRing R] (x : M) - (hx : x ≠ 0) : ∃ P : Ideal R, IsAssociatedPrime P M ∧ (R ∙ x).annihilator ≤ P := by - have : (R ∙ x).annihilator ≠ ⊤ := by - rwa [Ne, Ideal.eq_top_iff_one, Submodule.mem_annihilator_span_singleton, one_smul] + (hx : x ≠ 0) : ∃ P : Ideal R, IsAssociatedPrime P M ∧ ker (toSpanSingleton R M x) ≤ P := by + have : ker (toSpanSingleton R M x) ≠ ⊤ := by + rwa [Ne, Ideal.eq_top_iff_one, mem_ker, toSpanSingleton_apply, one_smul] obtain ⟨P, ⟨l, h₁, y, rfl⟩, h₃⟩ := set_has_maximal_iff_noetherian.mpr H - { P | (R ∙ x).annihilator ≤ P ∧ P ≠ ⊤ ∧ ∃ y : M, P = (R ∙ y).annihilator } - ⟨(R ∙ x).annihilator, rfl.le, this, x, rfl⟩ + { P | ker (toSpanSingleton R M x) ≤ P ∧ P ≠ ⊤ ∧ ∃ y : M, P = ker (toSpanSingleton R M y) } + ⟨_, rfl.le, this, x, rfl⟩ refine ⟨_, ⟨⟨h₁, ?_⟩, y, rfl⟩, l⟩ intro a b hab rw [or_iff_not_imp_left] intro ha - rw [Submodule.mem_annihilator_span_singleton] at ha hab - have H₁ : (R ∙ y).annihilator ≤ (R ∙ a • y).annihilator := by + rw [mem_ker, toSpanSingleton_apply] at ha hab + have H₁ : ker (toSpanSingleton R M y) ≤ ker (toSpanSingleton R M (a • y)) := by intro c hc - rw [Submodule.mem_annihilator_span_singleton] at hc ⊢ + rw [mem_ker, toSpanSingleton_apply] at hc ⊢ rw [smul_comm, hc, smul_zero] - have H₂ : (Submodule.span R {a • y}).annihilator ≠ ⊤ := by - rwa [Ne, Submodule.annihilator_eq_top_iff, Submodule.span_singleton_eq_bot] - rwa [H₁.eq_of_not_lt (h₃ (R ∙ a • y).annihilator ⟨l.trans H₁, H₂, _, rfl⟩), - Submodule.mem_annihilator_span_singleton, smul_comm, smul_smul] + have H₂ : ker (toSpanSingleton R M (a • y)) ≠ ⊤ := by + rwa [Ne, ker_eq_top, toSpanSingleton_eq_zero_iff] + rwa [H₁.eq_of_not_lt (h₃ _ ⟨l.trans H₁, H₂, _, rfl⟩), + mem_ker, toSpanSingleton_apply, smul_comm, smul_smul] variable {R} @@ -117,11 +117,9 @@ theorem associatedPrimes.nonempty [IsNoetherianRing R] [Nontrivial M] : theorem biUnion_associatedPrimes_eq_zero_divisors [IsNoetherianRing R] : ⋃ p ∈ associatedPrimes R M, p = { r : R | ∃ x : M, x ≠ 0 ∧ r • x = 0 } := by - simp_rw [← Submodule.mem_annihilator_span_singleton] refine subset_antisymm (Set.iUnion₂_subset ?_) ?_ · rintro _ ⟨h, x, ⟨⟩⟩ r h' refine ⟨x, ne_of_eq_of_ne (one_smul R x).symm ?_, h'⟩ - refine mt (Submodule.mem_annihilator_span_singleton _ _).mpr ?_ exact (Ideal.ne_top_iff_one _).mp h.ne_top · intro r ⟨x, h, h'⟩ obtain ⟨P, hP, hx⟩ := exists_le_isAssociatedPrime_of_isNoetherianRing R x h @@ -132,6 +130,7 @@ variable {R M} theorem IsAssociatedPrime.annihilator_le (h : IsAssociatedPrime I M) : (⊤ : Submodule R M).annihilator ≤ I := by obtain ⟨hI, x, rfl⟩ := h + rw [← Submodule.annihilator_span_singleton] exact Submodule.annihilator_mono le_top theorem IsAssociatedPrime.eq_radical (hI : I.IsPrimary) (h : IsAssociatedPrime J (R ⧸ I)) : @@ -140,11 +139,11 @@ theorem IsAssociatedPrime.eq_radical (hI : I.IsPrimary) (h : IsAssociatedPrime J have : x ≠ 0 := by rintro rfl apply hJ.1 - rwa [Submodule.span_singleton_eq_bot.mpr rfl, Submodule.annihilator_bot] at e + rwa [toSpanSingleton_zero, ker_zero] at e obtain ⟨x, rfl⟩ := Ideal.Quotient.mkₐ_surjective R _ x replace e : ∀ {y}, y ∈ J ↔ x * y ∈ I := by intro y - rw [e, Submodule.mem_annihilator_span_singleton, ← map_smul, smul_eq_mul, mul_comm, + rw [e, mem_ker, toSpanSingleton_apply, ← map_smul, smul_eq_mul, mul_comm, Ideal.Quotient.mkₐ_eq_mk, ← Ideal.Quotient.mk_eq_mk, Submodule.Quotient.mk_eq_zero] apply le_antisymm · intro y hy diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 864e9f84072c4..bb6c5745b3faf 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -53,6 +53,31 @@ theorem mem_pi (x : Π i, α i) : x ∈ pi I ↔ ∀ i, x i ∈ I i := end Pi +section Commute + +variable {α : Type*} [Semiring α] (I : Ideal α) {a b : α} + +theorem add_pow_mem_of_pow_mem_of_le_of_commute {m n k : ℕ} + (ha : a ^ m ∈ I) (hb : b ^ n ∈ I) (hk : m + n ≤ k + 1) + (hab : Commute a b) : + (a + b) ^ k ∈ I := by + simp_rw [hab.add_pow, ← Nat.cast_comm] + apply I.sum_mem + intro c _ + apply mul_mem_left + by_cases h : m ≤ c + · rw [hab.pow_pow] + exact I.mul_mem_left _ (I.pow_mem_of_pow_mem ha h) + · refine I.mul_mem_left _ (I.pow_mem_of_pow_mem hb ?_) + omega + +theorem add_pow_add_pred_mem_of_pow_mem_of_commute {m n : ℕ} + (ha : a ^ m ∈ I) (hb : b ^ n ∈ I) (hab : Commute a b) : + (a + b) ^ (m + n - 1) ∈ I := + I.add_pow_mem_of_pow_mem_of_le_of_commute ha hb (by rw [← Nat.sub_le_iff_le_add]) hab + +end Commute + end Ideal end Semiring @@ -69,25 +94,13 @@ variable [CommSemiring α] (I : Ideal α) theorem add_pow_mem_of_pow_mem_of_le {m n k : ℕ} (ha : a ^ m ∈ I) (hb : b ^ n ∈ I) (hk : m + n ≤ k + 1) : - (a + b) ^ k ∈ I := by - rw [add_pow] - apply I.sum_mem - intro c _ - apply mul_mem_right - by_cases h : m ≤ c - · exact I.mul_mem_right _ (I.pow_mem_of_pow_mem ha h) - · refine I.mul_mem_left _ (I.pow_mem_of_pow_mem hb ?_) - simp only [not_le, Nat.lt_iff_add_one_le] at h - have hck : c ≤ k := by - rw [← add_le_add_iff_right 1] - exact le_trans h (le_trans (Nat.le_add_right _ _) hk) - rw [Nat.le_sub_iff_add_le hck, ← add_le_add_iff_right 1] - exact le_trans (by rwa [add_comm _ n, add_assoc, add_le_add_iff_left]) hk + (a + b) ^ k ∈ I := + I.add_pow_mem_of_pow_mem_of_le_of_commute ha hb hk (Commute.all ..) theorem add_pow_add_pred_mem_of_pow_mem {m n : ℕ} (ha : a ^ m ∈ I) (hb : b ^ n ∈ I) : (a + b) ^ (m + n - 1) ∈ I := - I.add_pow_mem_of_pow_mem_of_le ha hb <| by rw [← Nat.sub_le_iff_le_add] + I.add_pow_add_pred_mem_of_pow_mem_of_commute ha hb (Commute.all ..) theorem pow_multiset_sum_mem_span_pow [DecidableEq α] (s : Multiset α) (n : ℕ) : s.sum ^ (Multiset.card s * n + 1) ∈ diff --git a/Mathlib/RingTheory/Ideal/Defs.lean b/Mathlib/RingTheory/Ideal/Defs.lean index ba78f8a4b205b..ca2e6072c593a 100644 --- a/Mathlib/RingTheory/Ideal/Defs.lean +++ b/Mathlib/RingTheory/Ideal/Defs.lean @@ -66,6 +66,14 @@ theorem unit_mul_mem_iff_mem {x y : α} (hy : IsUnit y) : y * x ∈ I ↔ x ∈ have := I.mul_mem_left y' h rwa [← mul_assoc, hy', one_mul] at this +theorem pow_mem_of_mem (ha : a ∈ I) (n : ℕ) (hn : 0 < n) : a ^ n ∈ I := + Nat.casesOn n (Not.elim (by decide)) + (fun m _hm => (pow_succ a m).symm ▸ I.mul_mem_left (a ^ m) ha) hn + +theorem pow_mem_of_pow_mem {m n : ℕ} (ha : a ^ m ∈ I) (h : m ≤ n) : a ^ n ∈ I := by + rw [← Nat.add_sub_of_le h, add_comm, pow_add] + exact I.mul_mem_left _ ha + end Ideal /-- For two elements `m` and `m'` in an `R`-module `M`, the set of elements `r : R` with @@ -103,14 +111,6 @@ variable {b} lemma mem_of_dvd (hab : a ∣ b) (ha : a ∈ I) : b ∈ I := by obtain ⟨c, rfl⟩ := hab; exact I.mul_mem_right _ ha -theorem pow_mem_of_mem (ha : a ∈ I) (n : ℕ) (hn : 0 < n) : a ^ n ∈ I := - Nat.casesOn n (Not.elim (by decide)) - (fun m _hm => (pow_succ a m).symm ▸ I.mul_mem_left (a ^ m) ha) hn - -theorem pow_mem_of_pow_mem {m n : ℕ} (ha : a ^ m ∈ I) (h : m ≤ n) : a ^ n ∈ I := by - rw [← Nat.add_sub_of_le h, pow_add] - exact I.mul_mem_right _ ha - end Ideal end CommSemiring diff --git a/Mathlib/RingTheory/Ideal/Height.lean b/Mathlib/RingTheory/Ideal/Height.lean new file mode 100644 index 0000000000000..965faae6a18cf --- /dev/null +++ b/Mathlib/RingTheory/Ideal/Height.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2025 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jiedong Jiang, Jingting Wang, Andrew Yang +-/ +import Mathlib.Order.KrullDimension +import Mathlib.RingTheory.Ideal.MinimalPrime +import Mathlib.RingTheory.Spectrum.Prime.Basic +/-! +# The Height of an Ideal + +In this file, we define the height of a prime ideal and the height of an ideal. + +# Main definitions + +* `Ideal.primeHeight` : The height of a prime ideal $\mathfrak{p}$. We define it as the supremum of + the lengths of strictly decreasing chains of prime ideals below it. This definition is implemented + via `Order.height`. + +* `Ideal.height` : The height of an ideal. We defined it as the infimum of the `primeHeight` of the +minimal prime ideals of I. + +-/ + + +variable {R : Type*} [CommRing R] (I : Ideal R) + +open Ideal + +/-- The height of a prime ideal is defined as the supremum of the lengths of strictly decreasing +chains of prime ideals below it. -/ +noncomputable def Ideal.primeHeight [hI : I.IsPrime] : ℕ∞ := + Order.height (⟨I, hI⟩ : PrimeSpectrum R) + +/-- The height of an ideal is defined as the infimum of the heights of its minimal prime ideals. -/ +noncomputable def Ideal.height : ℕ∞ := + ⨅ J ∈ I.minimalPrimes, @Ideal.primeHeight _ _ J (minimalPrimes_isPrime ‹_›) + +/-- For a prime ideal, its height equals its prime height. -/ +lemma Ideal.height_eq_primeHeight [I.IsPrime] : I.height = I.primeHeight := by + unfold height primeHeight + simp_rw [Ideal.minimalPrimes_eq_subsingleton_self] + simp + +/-- An ideal has finite height if it is either the unit ideal or its height is finite. +We include the unit ideal in order to have the instance `IsNoetherianRing R → FiniteHeight I`. -/ +class Ideal.FiniteHeight : Prop where + eq_top_or_height_ne_top : I = ⊤ ∨ I.height ≠ ⊤ + +lemma Ideal.finiteHeight_iff_lt {I : Ideal R} : + Ideal.FiniteHeight I ↔ I = ⊤ ∨ I.height < ⊤ := by + constructor + · intro h + cases h.eq_top_or_height_ne_top with + | inl h => exact Or.inl h + | inr h => exact Or.inr (lt_of_le_of_ne le_top h) + · intro h + constructor + cases h with + | inl h => exact Or.inl h + | inr h => exact Or.inr (ne_top_of_lt h) + +lemma Ideal.height_ne_top {I : Ideal R} (hI : I ≠ ⊤) [h : I.FiniteHeight] : + I.height ≠ ⊤ := + (h.eq_top_or_height_ne_top).resolve_left hI + +lemma Ideal.height_lt_top {I : Ideal R} (hI : I ≠ ⊤) [h : I.FiniteHeight] : + I.height < ⊤ := + lt_of_le_of_ne le_top (Ideal.height_ne_top hI) + +lemma Ideal.primeHeight_ne_top (I : Ideal R) [I.FiniteHeight] [h : I.IsPrime] : + I.primeHeight ≠ ⊤ := by + rw [← I.height_eq_primeHeight] + exact Ideal.height_ne_top (by exact h.ne_top) + +lemma Ideal.primeHeight_lt_top (I : Ideal R) [I.FiniteHeight] [h : I.IsPrime] : + I.primeHeight < ⊤ := by + rw [← I.height_eq_primeHeight] + exact Ideal.height_lt_top (by exact h.ne_top) + +@[gcongr] +lemma Ideal.primeHeight_mono {I J : Ideal R} [I.IsPrime] [J.IsPrime] (h : I ≤ J) : + I.primeHeight ≤ J.primeHeight := by + unfold primeHeight + gcongr + exact h + +lemma Ideal.primeHeight_add_one_le_of_lt {I J : Ideal R} [I.IsPrime] [J.IsPrime] (h : I < J) : + I.primeHeight + 1 ≤ J.primeHeight := by + unfold primeHeight + exact Order.height_add_one_le h + +@[gcongr] +lemma Ideal.primeHeight_strict_mono {I J : Ideal R} [I.IsPrime] [J.IsPrime] + (h : I < J) [I.FiniteHeight] : + I.primeHeight < J.primeHeight := by + unfold primeHeight + gcongr + · exact I.primeHeight_ne_top.lt_top + · exact h + +@[simp] +theorem Ideal.height_top : (⊤ : Ideal R).height = ⊤ := by + simp only [height, minimalPrimes_top] + rw [iInf₂_eq_top]; intro i hi; exact False.elim hi + +@[gcongr] +theorem Ideal.height_mono {I J : Ideal R} (h : I ≤ J) : I.height ≤ J.height := by + simp only [height] + apply le_iInf₂; intro p hp; haveI := hp.1.1 + obtain ⟨q, hq, e⟩ := Ideal.exists_minimalPrimes_le (h.trans hp.1.2) + haveI := hq.1.1 + exact (iInf₂_le q hq).trans (Ideal.primeHeight_mono e) + +@[gcongr] +lemma Ideal.height_strict_mono_of_is_prime {I J : Ideal R} [I.IsPrime] + (h : I < J) [I.FiniteHeight] : I.height < J.height := by + rw [Ideal.height_eq_primeHeight I] + by_cases hJ : J = ⊤ + · rw [hJ, height_top]; exact I.primeHeight_lt_top + · rw [← ENat.add_one_le_iff I.primeHeight_ne_top, Ideal.height] + apply le_iInf₂; intro K hK; haveI := hK.1.1 + have : I < K := lt_of_lt_of_le h hK.1.2 + exact Ideal.primeHeight_add_one_le_of_lt this diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 8191b555882cb..67dad3b20e060 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -807,6 +807,16 @@ theorem mem_annihilator_span (s : Set M) (r : R) : theorem mem_annihilator_span_singleton (g : M) (r : R) : r ∈ (Submodule.span R ({g} : Set M)).annihilator ↔ r • g = 0 := by simp [mem_annihilator_span] +open LinearMap in +theorem annihilator_span (s : Set M) : + (Submodule.span R s).annihilator = ⨅ g : s, ker (toSpanSingleton R M g.1) := by + ext; simp [mem_annihilator_span] + +open LinearMap in +theorem annihilator_span_singleton (g : M) : + (Submodule.span R {g}).annihilator = ker (toSpanSingleton R M g) := by + simp [annihilator_span] + @[simp] theorem mul_annihilator (I : Ideal R) : I * annihilator I = ⊥ := by rw [mul_comm, annihilator_mul] diff --git a/Mathlib/RingTheory/Ideal/MinimalPrime.lean b/Mathlib/RingTheory/Ideal/MinimalPrime.lean index 3db1a63fcb0df..018eb06f8338f 100644 --- a/Mathlib/RingTheory/Ideal/MinimalPrime.lean +++ b/Mathlib/RingTheory/Ideal/MinimalPrime.lean @@ -49,6 +49,9 @@ lemma minimalPrimes_eq_minimals : minimalPrimes R = {x | Minimal Ideal.IsPrime x variable {I J} +theorem Ideal.minimalPrimes_isPrime {p : Ideal R} (h : p ∈ I.minimalPrimes) : p.IsPrime := + h.1.1 + theorem Ideal.exists_minimalPrimes_le [J.IsPrime] (e : I ≤ J) : ∃ p ∈ I.minimalPrimes, p ≤ J := by set S := { p : (Ideal R)ᵒᵈ | Ideal.IsPrime p ∧ I ≤ OrderDual.ofDual p } suffices h : ∃ m, OrderDual.toDual J ≤ m ∧ Maximal (· ∈ S) m by @@ -67,6 +70,16 @@ theorem Ideal.exists_minimalPrimes_le [J.IsPrime] (e : I ≤ J) : ∃ p ∈ I.mi rw [OrderDual.le_toDual] exact sInf_le hz +theorem Ideal.nonempty_minimalPrimes (h : I ≠ ⊤) : Nonempty I.minimalPrimes := by + obtain ⟨m, hm, hle⟩ := Ideal.exists_le_maximal I h + obtain ⟨p, hp, -⟩ := Ideal.exists_minimalPrimes_le hle + exact ⟨p, hp⟩ + +theorem Ideal.eq_bot_of_minimalPrimes_eq_empty (h : I.minimalPrimes = ∅) : I = ⊤ := by + by_contra hI + obtain ⟨p, hp⟩ := Ideal.nonempty_minimalPrimes hI + exact Set.not_mem_empty p (h ▸ hp) + @[simp] theorem Ideal.radical_minimalPrimes : I.radical.minimalPrimes = I.minimalPrimes := by rw [Ideal.minimalPrimes, Ideal.minimalPrimes] @@ -92,6 +105,59 @@ theorem Ideal.sInf_minimalPrimes : sInf I.minimalPrimes = I.radical := by intro I hI exact hI.1.symm +theorem Ideal.iUnion_minimalPrimes : + ⋃ p ∈ I.minimalPrimes, p = { x | ∃ y ∉ I.radical, x * y ∈ I.radical } := by + classical + ext x + simp only [Set.mem_iUnion, SetLike.mem_coe, exists_prop, Set.mem_setOf_eq] + constructor + · rintro ⟨p, ⟨⟨hp₁, hp₂⟩, hp₃⟩, hxp⟩ + have : p.map (algebraMap R (Localization.AtPrime p)) ≤ (I.map (algebraMap _ _)).radical := by + rw [Ideal.radical_eq_sInf, le_sInf_iff] + rintro q ⟨hq', hq⟩ + obtain ⟨h₁, h₂⟩ := ((IsLocalization.AtPrime.orderIsoOfPrime _ p) ⟨q, hq⟩).2 + rw [Ideal.map_le_iff_le_comap] at hq' ⊢ + exact hp₃ ⟨h₁, hq'⟩ h₂ + obtain ⟨n, hn⟩ := this (Ideal.mem_map_of_mem _ hxp) + rw [IsLocalization.mem_map_algebraMap_iff (M := p.primeCompl)] at hn + obtain ⟨⟨a, b⟩, hn⟩ := hn + rw [← map_pow, ← _root_.map_mul, IsLocalization.eq_iff_exists p.primeCompl] at hn + obtain ⟨t, ht⟩ := hn + refine ⟨t * b, fun h ↦ (t * b).2 (hp₁.radical_le_iff.mpr hp₂ h), n + 1, ?_⟩ + simp only at ht + have : (x * (t.1 * b.1)) ^ (n + 1) = (t.1 ^ n * b.1 ^ n * x * t.1) * a := by + rw [mul_assoc, ← ht]; ring + rw [this] + exact I.mul_mem_left _ a.2 + · rintro ⟨y, hy, hx⟩ + obtain ⟨p, hp, hyp⟩ : ∃ p ∈ I.minimalPrimes, y ∉ p := by + simpa [← Ideal.sInf_minimalPrimes] using hy + refine ⟨p, hp, (hp.1.1.mem_or_mem ?_).resolve_right hyp⟩ + exact hp.1.1.radical_le_iff.mpr hp.1.2 hx + +theorem Ideal.exists_mul_mem_of_mem_minimalPrimes + {p : Ideal R} (hp : p ∈ I.minimalPrimes) {x : R} (hx : x ∈ p) : + ∃ y ∉ I, x * y ∈ I := by + classical + obtain ⟨y, hy, n, hx⟩ := Ideal.iUnion_minimalPrimes.subset (Set.mem_biUnion hp hx) + have H : ∃ m, x ^ m * y ^ n ∈ I := ⟨n, mul_pow x y n ▸ hx⟩ + have : Nat.find H ≠ 0 := + fun h ↦ hy ⟨n, by simpa only [h, pow_zero, one_mul] using Nat.find_spec H⟩ + refine ⟨x ^ (Nat.find H - 1) * y ^ n, Nat.find_min H (Nat.sub_one_lt this), ?_⟩ + rw [← mul_assoc, ← pow_succ', tsub_add_cancel_of_le (Nat.one_le_iff_ne_zero.mpr this)] + exact Nat.find_spec H + +/-- minimal primes are contained in zero divisors. -/ +lemma Ideal.disjoint_nonZeroDivisors_of_mem_minimalPrimes {p : Ideal R} (hp : p ∈ minimalPrimes R) : + Disjoint (p : Set R) (nonZeroDivisors R) := by + classical + rw [← Set.subset_compl_iff_disjoint_right, Set.subset_def] + simp only [SetLike.mem_coe, Set.mem_compl_iff, mem_nonZeroDivisors_iff, not_forall, + Classical.not_imp] + intro x hxp + simp_rw [exists_prop, @and_comm (_ * _ = _), ← mul_comm x] + exact Ideal.exists_mul_mem_of_mem_minimalPrimes hp hxp + theorem Ideal.exists_comap_eq_of_mem_minimalPrimes_of_injective {f : R →+* S} (hf : Function.Injective f) (p) (H : p ∈ minimalPrimes R) : ∃ p' : Ideal S, p'.IsPrime ∧ p'.comap f = p := by @@ -235,3 +301,29 @@ theorem nilpotent_iff_not_unit_of_minimal {x : Localization I.primeCompl} : simpa only [← IsLocalRing.mem_maximalIdeal] using nilpotent_iff_mem_maximal_of_minimal hMin end Localization.AtPrime + +section + +variable {R : Type*} [CommSemiring R] + +theorem Ideal.minimalPrimes_top : (⊤ : Ideal R).minimalPrimes = ∅ := by + ext p + constructor + · intro h + exact False.elim (h.1.1.ne_top (top_le_iff.mp h.1.2)) + · intro h + exact False.elim (Set.not_mem_empty p h) + +theorem Ideal.minimalPrimes_eq_empty_iff (I : Ideal R) : + I.minimalPrimes = ∅ ↔ I = ⊤ := by + constructor + · intro e + by_contra h + have ⟨M, hM, hM'⟩ := Ideal.exists_le_maximal I h + have ⟨p, hp⟩ := Ideal.exists_minimalPrimes_le hM' + show p ∈ (∅ : Set (Ideal R)) + rw [← e]; exact hp.1 + · intro h; rw [h] + exact Ideal.minimalPrimes_top + +end diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index ee0467668b791..1c18bee75de0d 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -1035,7 +1035,8 @@ theorem subset_union_prime' {R : Type u} [CommRing R] {s : Finset ι} {f : ι exact hs <| Or.inr <| Set.mem_biUnion hjt <| add_sub_cancel_left r s ▸ (f j).sub_mem hj <| hr j hjt -/-- Prime avoidance. Atiyah-Macdonald 1.11, Eisenbud 3.3, Stacks 00DS, Matsumura Ex.1.6. -/ +/-- Prime avoidance. Atiyah-Macdonald 1.11, Eisenbud 3.3, Matsumura Ex.1.6. -/ +@[stacks 00DS] theorem subset_union_prime {R : Type u} [CommRing R] {s : Finset ι} {f : ι → Ideal R} (a b : ι) (hp : ∀ i ∈ s, i ≠ a → i ≠ b → IsPrime (f i)) {I : Ideal R} : ((I : Set R) ⊆ ⋃ i ∈ (↑s : Set ι), f i) ↔ ∃ i ∈ s, I ≤ f i := diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index c3fd963f5f2d3..48ef1022ba6d0 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -37,6 +37,18 @@ section CommRing variable {S : Type*} [CommRing S] {f : R →+* S} {I J : Ideal S} +/-- For a prime ideal `p` of `R`, `p` extended to `S` and +restricted back to `R` is `p` if and only if `p` is the restriction of a prime in `S`. -/ +lemma comap_map_eq_self_iff_of_isPrime (p : Ideal R) [p.IsPrime] : + (p.map f).comap f = p ↔ (∃ (q : Ideal S), q.IsPrime ∧ q.comap f = p) := by + refine ⟨fun hp ↦ ?_, ?_⟩ + · obtain ⟨q, hq₁, hq₂, hq₃⟩ := Ideal.exists_le_prime_disjoint _ _ + (disjoint_map_primeCompl_iff_comap_le.mpr hp.le) + exact ⟨q, hq₁, le_antisymm (disjoint_map_primeCompl_iff_comap_le.mp hq₃) + (map_le_iff_le_comap.mp hq₂)⟩ + · rintro ⟨q, hq, rfl⟩ + simp + theorem coeff_zero_mem_comap_of_root_mem_of_eval_mem {r : S} (hr : r ∈ I) {p : R[X]} (hp : p.eval₂ f r ∈ I) : p.coeff 0 ∈ I.comap f := by rw [← p.divX_mul_X_add, eval₂_add, eval₂_C, eval₂_mul, eval₂_X] at hp diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index c7b091ec93c93..9d42bb3accc80 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -35,74 +35,10 @@ In this file we provide various results regarding idempotent elements in rings. If `R` is commutative, then a family `{ eᵢ }` of complete orthogonal idempotent elements induces a ring isomorphism `R ≃ ∏ R ⧸ ⟨1 - eᵢ⟩`. -/ -section Ring -variable {R S : Type*} [Ring R] [Ring S] (f : R →+* S) - -theorem isIdempotentElem_one_sub_one_sub_pow_pow - (x : R) (n : ℕ) (hx : (x - x ^ 2) ^ n = 0) : - IsIdempotentElem (1 - (1 - x ^ n) ^ n) := by - let P : Polynomial ℤ := 1 - (1 - .X ^ n) ^ n - have : (.X - .X ^ 2) ^ n ∣ P - P ^ 2 := by - have H₁ : .X ^ n ∣ P := by - have := sub_dvd_pow_sub_pow 1 ((1 : Polynomial ℤ) - Polynomial.X ^ n) n - rwa [sub_sub_cancel, one_pow] at this - have H₂ : (1 - .X) ^ n ∣ 1 - P := by - simp only [sub_sub_cancel, P] - simpa using pow_dvd_pow_of_dvd (sub_dvd_pow_sub_pow (α := Polynomial ℤ) 1 Polynomial.X n) n - have := mul_dvd_mul H₁ H₂ - simpa only [← mul_pow, mul_sub, mul_one, ← pow_two] using this - have := map_dvd (Polynomial.aeval x) this - simp only [map_pow, map_sub, Polynomial.aeval_X, hx, map_one, zero_dvd_iff, P] at this - rwa [sub_eq_zero, eq_comm, pow_two] at this - -theorem exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent_aux - (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) - (e₁ : S) (he : e₁ ∈ f.range) (he₁ : IsIdempotentElem e₁) - (e₂ : R) (he₂ : IsIdempotentElem e₂) (he₁e₂ : e₁ * f e₂ = 0) : - ∃ e' : R, IsIdempotentElem e' ∧ f e' = e₁ ∧ e' * e₂ = 0 := by - obtain ⟨e₁, rfl⟩ := he - cases subsingleton_or_nontrivial R - · exact ⟨_, Subsingleton.elim _ _, rfl, Subsingleton.elim _ _⟩ - let a := e₁ - e₁ * e₂ - have ha : f a = f e₁ := by rw [map_sub, map_mul, he₁e₂, sub_zero] - have ha' : a * e₂ = 0 := by rw [sub_mul, mul_assoc, he₂.eq, sub_self] - have hx' : a - a ^ 2 ∈ RingHom.ker f := by - simp [RingHom.mem_ker, mul_sub, pow_two, ha, he₁.eq] - obtain ⟨n, hn⟩ := h _ hx' - refine ⟨_, isIdempotentElem_one_sub_one_sub_pow_pow _ _ hn, ?_, ?_⟩ - · cases' n with n - · simp at hn - simp only [map_sub, map_one, map_pow, ha, he₁.pow_succ_eq, - he₁.one_sub.pow_succ_eq, sub_sub_cancel] - · obtain ⟨k, hk⟩ := (Commute.one_left (MulOpposite.op <| 1 - a ^ n)).sub_dvd_pow_sub_pow n - apply_fun MulOpposite.unop at hk - have : 1 - (1 - a ^ n) ^ n = MulOpposite.unop k * a ^ n := by simpa using hk - rw [this, mul_assoc] - cases' n with n - · simp at hn - rw [pow_succ, mul_assoc, ha', mul_zero, mul_zero] - -/-- Orthogonal idempotents lift along nil ideals. -/ -theorem exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent - (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) - (e₁ : S) (he : e₁ ∈ f.range) (he₁ : IsIdempotentElem e₁) - (e₂ : R) (he₂ : IsIdempotentElem e₂) (he₁e₂ : e₁ * f e₂ = 0) (he₂e₁ : f e₂ * e₁ = 0) : - ∃ e' : R, IsIdempotentElem e' ∧ f e' = e₁ ∧ e' * e₂ = 0 ∧ e₂ * e' = 0 := by - obtain ⟨e', h₁, rfl, h₂⟩ := exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent_aux - f h e₁ he he₁ e₂ he₂ he₁e₂ - refine ⟨(1 - e₂) * e', ?_, ?_, ?_, ?_⟩ - · rw [IsIdempotentElem, mul_assoc, ← mul_assoc e', mul_sub, mul_one, h₂, sub_zero, h₁.eq] - · rw [map_mul, map_sub, map_one, sub_mul, one_mul, he₂e₁, sub_zero] - · rw [mul_assoc, h₂, mul_zero] - · rw [← mul_assoc, mul_sub, mul_one, he₂.eq, sub_self, zero_mul] - -/-- Idempotents lift along nil ideals. -/ -theorem exists_isIdempotentElem_eq_of_ker_isNilpotent (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) - (e : S) (he : e ∈ f.range) (he' : IsIdempotentElem e) : - ∃ e' : R, IsIdempotentElem e' ∧ f e' = e := by - simpa using exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent f h e he he' 0 .zero (by simp) +section Semiring +variable {R S : Type*} [Semiring R] [Semiring S] (f : R →+* S) variable {I : Type*} (e : I → R) /-- A family `{ eᵢ }` of idempotent elements is orthogonal if `eᵢ * eⱼ = 0` for all `i ≠ j`. -/ @@ -176,33 +112,6 @@ lemma OrthogonalIdempotents.option (he : OrthogonalIdempotents e) [Fintype I] (x Finset.sum_ite_eq, Finset.mem_univ, ↓reduceIte, mul_zero] using congr_arg (e i * ·) hx₂ · exact he.ortho (Option.some_inj.ne.mp ne) -lemma OrthogonalIdempotents.lift_of_isNilpotent_ker_aux - (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) - {n} {e : Fin n → S} (he : OrthogonalIdempotents e) (he' : ∀ i, e i ∈ f.range) : - ∃ e' : Fin n → R, OrthogonalIdempotents e' ∧ f ∘ e' = e := by - induction' n with n IH - · refine ⟨0, ⟨finZeroElim, finZeroElim⟩, funext finZeroElim⟩ - · obtain ⟨e', h₁, h₂⟩ := IH (he.embedding (Fin.succEmb n)) (fun i ↦ he' _) - have h₂' (i) : f (e' i) = e i.succ := congr_fun h₂ i - obtain ⟨e₀, h₃, h₄, h₅, h₆⟩ := - exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent f h _ (he' 0) (he.idem 0) _ - h₁.isIdempotentElem_sum - (by simp [Finset.mul_sum, h₂', he.mul_eq, Fin.succ_ne_zero, eq_comm]) - (by simp [Finset.sum_mul, h₂', he.mul_eq, Fin.succ_ne_zero]) - refine ⟨_, (h₁.option _ h₃ h₅ h₆).embedding (finSuccEquiv n).toEmbedding, funext fun i ↦ ?_⟩ - obtain ⟨_ | i, rfl⟩ := (finSuccEquiv n).symm.surjective i <;> simp [*] - -/-- A family of orthogonal idempotents lift along nil ideals. -/ -lemma OrthogonalIdempotents.lift_of_isNilpotent_ker [Finite I] - (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) - {e : I → S} (he : OrthogonalIdempotents e) (he' : ∀ i, e i ∈ f.range) : - ∃ e' : I → R, OrthogonalIdempotents e' ∧ f ∘ e' = e := by - cases nonempty_fintype I - obtain ⟨e', h₁, h₂⟩ := lift_of_isNilpotent_ker_aux f h - (he.embedding (Fintype.equivFin I).symm.toEmbedding) (fun _ ↦ he' _) - refine ⟨_, h₁.embedding (Fintype.equivFin I).toEmbedding, - by ext x; simpa using congr_fun h₂ (Fintype.equivFin I x)⟩ - variable [Fintype I] /-- @@ -214,31 +123,31 @@ A family `{ eᵢ }` of idempotent elements is complete orthogonal if structure CompleteOrthogonalIdempotents (e : I → R) extends OrthogonalIdempotents e : Prop where complete : ∑ i, e i = 1 +/-- If a family is complete orthogonal, it consists of idempotents. -/ +lemma CompleteOrthogonalIdempotents.iff_ortho_complete : + CompleteOrthogonalIdempotents e ↔ Pairwise (e · * e · = 0) ∧ ∑ i, e i = 1 := by + rw [completeOrthogonalIdempotents_iff, orthogonalIdempotents_iff, and_assoc, and_iff_right_of_imp] + intro ⟨ortho, complete⟩ i + apply_fun (e i * ·) at complete + rwa [Finset.mul_sum, Finset.sum_eq_single i (fun _ _ ne ↦ ortho ne.symm) (by simp at ·), mul_one] + at complete + +lemma CompleteOrthogonalIdempotents.pair_iff'ₛ {x y : R} : + CompleteOrthogonalIdempotents ![x, y] ↔ x * y = 0 ∧ y * x = 0 ∧ x + y = 1 := by + simp [iff_ortho_complete, Pairwise, Fin.forall_fin_two, and_assoc] + +lemma CompleteOrthogonalIdempotents.pair_iffₛ {R} [CommSemiring R] {x y : R} : + CompleteOrthogonalIdempotents ![x, y] ↔ x * y = 0 ∧ x + y = 1 := by + rw [pair_iff'ₛ, and_left_comm, and_iff_right_of_imp]; exact (mul_comm x y ▸ ·.1) + lemma CompleteOrthogonalIdempotents.unique_iff [Unique I] : CompleteOrthogonalIdempotents e ↔ e default = 1 := by rw [completeOrthogonalIdempotents_iff, OrthogonalIdempotents.unique, Fintype.sum_unique, and_iff_right_iff_imp] exact (· ▸ IsIdempotentElem.one) -lemma CompleteOrthogonalIdempotents.pair_iff {x y : R} : - CompleteOrthogonalIdempotents ![x, y] ↔ IsIdempotentElem x ∧ y = 1 - x := by - rw [completeOrthogonalIdempotents_iff, orthogonalIdempotents_iff, and_assoc, Pairwise] - simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.forall_fin_two, Fin.isValue, - Matrix.cons_val_zero, Matrix.cons_val_one, Matrix.head_cons, ne_eq, not_true_eq_false, - false_implies, zero_ne_one, not_false_eq_true, true_implies, true_and, one_ne_zero, - and_true, and_self, Fin.sum_univ_two, eq_sub_iff_add_eq, @add_comm _ _ y] - constructor - · exact fun h ↦ ⟨h.1.1, h.2.2⟩ - · rintro ⟨h₁, h₂⟩ - obtain rfl := eq_sub_iff_add_eq'.mpr h₂ - exact ⟨⟨h₁, h₁.one_sub⟩, ⟨by simp [mul_sub, h₁.eq], by simp [sub_mul, h₁.eq]⟩, h₂⟩ - -lemma CompleteOrthogonalIdempotents.of_isIdempotentElem {e : R} (he : IsIdempotentElem e) : - CompleteOrthogonalIdempotents ![e, 1 - e] := - pair_iff.mpr ⟨he, rfl⟩ - lemma CompleteOrthogonalIdempotents.single {I : Type*} [Fintype I] [DecidableEq I] - (R : I → Type*) [∀ i, Ring (R i)] : + (R : I → Type*) [∀ i, Semiring (R i)] : CompleteOrthogonalIdempotents (Pi.single (f := R) · 1) := by refine ⟨⟨by simp [IsIdempotentElem, ← Pi.single_mul], ?_⟩, Finset.univ_sum_single 1⟩ intros i j hij @@ -262,6 +171,122 @@ lemma CompleteOrthogonalIdempotents.equiv {J} [Fintype J] (i : J ≃ I) : simp only [completeOrthogonalIdempotents_iff, OrthogonalIdempotents.equiv, Function.comp_apply, and_congr_right_iff, Fintype.sum_equiv i _ e (fun _ ↦ rfl)] +@[nontriviality] +lemma CompleteOrthogonalIdempotents.of_subsingleton [Subsingleton R] : + CompleteOrthogonalIdempotents e := + ⟨⟨fun _ ↦ Subsingleton.elim _ _, fun _ _ _ ↦ Subsingleton.elim _ _⟩, Subsingleton.elim _ _⟩ + +end Semiring + +section Ring + +variable {R S : Type*} [Ring R] [Ring S] (f : R →+* S) + +theorem isIdempotentElem_one_sub_one_sub_pow_pow + (x : R) (n : ℕ) (hx : (x - x ^ 2) ^ n = 0) : + IsIdempotentElem (1 - (1 - x ^ n) ^ n) := by + let P : Polynomial ℤ := 1 - (1 - .X ^ n) ^ n + have : (.X - .X ^ 2) ^ n ∣ P - P ^ 2 := by + have H₁ : .X ^ n ∣ P := by + have := sub_dvd_pow_sub_pow 1 ((1 : Polynomial ℤ) - Polynomial.X ^ n) n + rwa [sub_sub_cancel, one_pow] at this + have H₂ : (1 - .X) ^ n ∣ 1 - P := by + simp only [sub_sub_cancel, P] + simpa using pow_dvd_pow_of_dvd (sub_dvd_pow_sub_pow (α := Polynomial ℤ) 1 Polynomial.X n) n + have := mul_dvd_mul H₁ H₂ + simpa only [← mul_pow, mul_sub, mul_one, ← pow_two] using this + have := map_dvd (Polynomial.aeval x) this + simp only [map_pow, map_sub, Polynomial.aeval_X, hx, map_one, zero_dvd_iff, P] at this + rwa [sub_eq_zero, eq_comm, pow_two] at this + +theorem exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent_aux + (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + (e₁ : S) (he : e₁ ∈ f.range) (he₁ : IsIdempotentElem e₁) + (e₂ : R) (he₂ : IsIdempotentElem e₂) (he₁e₂ : e₁ * f e₂ = 0) : + ∃ e' : R, IsIdempotentElem e' ∧ f e' = e₁ ∧ e' * e₂ = 0 := by + obtain ⟨e₁, rfl⟩ := he + cases subsingleton_or_nontrivial R + · exact ⟨_, Subsingleton.elim _ _, rfl, Subsingleton.elim _ _⟩ + let a := e₁ - e₁ * e₂ + have ha : f a = f e₁ := by rw [map_sub, map_mul, he₁e₂, sub_zero] + have ha' : a * e₂ = 0 := by rw [sub_mul, mul_assoc, he₂.eq, sub_self] + have hx' : a - a ^ 2 ∈ RingHom.ker f := by + simp [RingHom.mem_ker, mul_sub, pow_two, ha, he₁.eq] + obtain ⟨n, hn⟩ := h _ hx' + refine ⟨_, isIdempotentElem_one_sub_one_sub_pow_pow _ _ hn, ?_, ?_⟩ + · cases' n with n + · simp at hn + simp only [map_sub, map_one, map_pow, ha, he₁.pow_succ_eq, + he₁.one_sub.pow_succ_eq, sub_sub_cancel] + · obtain ⟨k, hk⟩ := (Commute.one_left (MulOpposite.op <| 1 - a ^ n)).sub_dvd_pow_sub_pow n + apply_fun MulOpposite.unop at hk + have : 1 - (1 - a ^ n) ^ n = MulOpposite.unop k * a ^ n := by simpa using hk + rw [this, mul_assoc] + cases' n with n + · simp at hn + rw [pow_succ, mul_assoc, ha', mul_zero, mul_zero] + +/-- Orthogonal idempotents lift along nil ideals. -/ +theorem exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent + (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + (e₁ : S) (he : e₁ ∈ f.range) (he₁ : IsIdempotentElem e₁) + (e₂ : R) (he₂ : IsIdempotentElem e₂) (he₁e₂ : e₁ * f e₂ = 0) (he₂e₁ : f e₂ * e₁ = 0) : + ∃ e' : R, IsIdempotentElem e' ∧ f e' = e₁ ∧ e' * e₂ = 0 ∧ e₂ * e' = 0 := by + obtain ⟨e', h₁, rfl, h₂⟩ := exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent_aux + f h e₁ he he₁ e₂ he₂ he₁e₂ + refine ⟨(1 - e₂) * e', ?_, ?_, ?_, ?_⟩ + · rw [IsIdempotentElem, mul_assoc, ← mul_assoc e', mul_sub, mul_one, h₂, sub_zero, h₁.eq] + · rw [map_mul, map_sub, map_one, sub_mul, one_mul, he₂e₁, sub_zero] + · rw [mul_assoc, h₂, mul_zero] + · rw [← mul_assoc, mul_sub, mul_one, he₂.eq, sub_self, zero_mul] + +/-- Idempotents lift along nil ideals. -/ +theorem exists_isIdempotentElem_eq_of_ker_isNilpotent (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + (e : S) (he : e ∈ f.range) (he' : IsIdempotentElem e) : + ∃ e' : R, IsIdempotentElem e' ∧ f e' = e := by + simpa using exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent f h e he he' 0 .zero (by simp) + +lemma OrthogonalIdempotents.lift_of_isNilpotent_ker_aux + (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + {n} {e : Fin n → S} (he : OrthogonalIdempotents e) (he' : ∀ i, e i ∈ f.range) : + ∃ e' : Fin n → R, OrthogonalIdempotents e' ∧ f ∘ e' = e := by + induction' n with n IH + · refine ⟨0, ⟨finZeroElim, finZeroElim⟩, funext finZeroElim⟩ + · obtain ⟨e', h₁, h₂⟩ := IH (he.embedding (Fin.succEmb n)) (fun i ↦ he' _) + have h₂' (i) : f (e' i) = e i.succ := congr_fun h₂ i + obtain ⟨e₀, h₃, h₄, h₅, h₆⟩ := + exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent f h _ (he' 0) (he.idem 0) _ + h₁.isIdempotentElem_sum + (by simp [Finset.mul_sum, h₂', he.mul_eq, Fin.succ_ne_zero, eq_comm]) + (by simp [Finset.sum_mul, h₂', he.mul_eq, Fin.succ_ne_zero]) + refine ⟨_, (h₁.option _ h₃ h₅ h₆).embedding (finSuccEquiv n).toEmbedding, funext fun i ↦ ?_⟩ + obtain ⟨_ | i, rfl⟩ := (finSuccEquiv n).symm.surjective i <;> simp [*] + +variable {I : Type*} {e : I → R} + +/-- A family of orthogonal idempotents lift along nil ideals. -/ +lemma OrthogonalIdempotents.lift_of_isNilpotent_ker [Finite I] + (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + {e : I → S} (he : OrthogonalIdempotents e) (he' : ∀ i, e i ∈ f.range) : + ∃ e' : I → R, OrthogonalIdempotents e' ∧ f ∘ e' = e := by + cases nonempty_fintype I + obtain ⟨e', h₁, h₂⟩ := lift_of_isNilpotent_ker_aux f h + (he.embedding (Fintype.equivFin I).symm.toEmbedding) (fun _ ↦ he' _) + refine ⟨_, h₁.embedding (Fintype.equivFin I).toEmbedding, + by ext x; simpa using congr_fun h₂ (Fintype.equivFin I x)⟩ + +lemma CompleteOrthogonalIdempotents.pair_iff {x y : R} : + CompleteOrthogonalIdempotents ![x, y] ↔ IsIdempotentElem x ∧ y = 1 - x := by + rw [pair_iff'ₛ, ← eq_sub_iff_add_eq', ← and_assoc, and_congr_left_iff] + rintro rfl + simp [mul_sub, sub_mul, IsIdempotentElem, sub_eq_zero, eq_comm] + +lemma CompleteOrthogonalIdempotents.of_isIdempotentElem {e : R} (he : IsIdempotentElem e) : + CompleteOrthogonalIdempotents ![e, 1 - e] := + pair_iff.mpr ⟨he, rfl⟩ + +variable [Fintype I] + lemma CompleteOrthogonalIdempotents.option (he : OrthogonalIdempotents e) : CompleteOrthogonalIdempotents (Option.elim · (1 - ∑ i, e i) e) where __ := he.option _ he.isIdempotentElem_sum.one_sub @@ -270,11 +295,6 @@ lemma CompleteOrthogonalIdempotents.option (he : OrthogonalIdempotents e) : rw [Fintype.sum_option] exact sub_add_cancel _ _ -@[nontriviality] -lemma CompleteOrthogonalIdempotents.of_subsingleton [Subsingleton R] : - CompleteOrthogonalIdempotents e := - ⟨⟨fun _ ↦ Subsingleton.elim _ _, fun _ _ _ ↦ Subsingleton.elim _ _⟩, Subsingleton.elim _ _⟩ - lemma CompleteOrthogonalIdempotents.lift_of_isNilpotent_ker_aux (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) {n} {e : Fin n → S} (he : CompleteOrthogonalIdempotents e) (he' : ∀ i, e i ∈ f.range) : @@ -426,3 +446,110 @@ lemma bijective_pi_of_isIdempotentElem (e : I → R) ⟨fun i ↦ (he i).one_sub, he₁⟩ (by simpa using he₂)).bijective_pi' end CommRing + +section corner + +variable {R : Type*} (e : R) + +namespace Subsemigroup + +variable [Semigroup R] + +/-- The corner associated to an element `e` in a semigroup +is the subsemigroup of all elements of the form `e * r * e`. -/ +def corner : Subsemigroup R where + carrier := Set.range (e * · * e) + mul_mem' := by rintro _ _ ⟨a, rfl⟩ ⟨b, rfl⟩; exact ⟨a * e * e * b, by simp_rw [mul_assoc]⟩ + +variable {e} (idem : IsIdempotentElem e) +include idem + +lemma mem_corner_iff {r : R} : r ∈ corner e ↔ e * r = r ∧ r * e = r := + ⟨by rintro ⟨r, rfl⟩; simp_rw [← mul_assoc, idem.eq, mul_assoc, idem.eq, true_and], + (⟨r, by simp_rw [·]⟩)⟩ + +lemma mem_corner_iff_mul_left (hc : IsMulCentral e) {r : R} : r ∈ corner e ↔ e * r = r := by + rw [mem_corner_iff idem, and_iff_left_of_imp]; intro; rwa [← hc.comm] + +lemma mem_corner_iff_mul_right (hc : IsMulCentral e) {r : R} : r ∈ corner e ↔ r * e = r := by + rw [mem_corner_iff_mul_left idem hc, hc.comm] + +lemma mem_corner_iff_mem_range_mul_left (hc : IsMulCentral e) {r : R} : + r ∈ corner e ↔ r ∈ Set.range (e * ·) := by + simp_rw [corner, mem_mk, Set.mem_range, ← hc.comm, ← mul_assoc, idem.eq] + +lemma mem_corner_iff_mem_range_mul_right (hc : IsMulCentral e) {r : R} : + r ∈ corner e ↔ r ∈ Set.range (· * e) := by + simp_rw [mem_corner_iff_mem_range_mul_left idem hc, hc.comm] + +/-- The corner associated to an idempotent `e` in a semiring without 1 +is the semiring with `e` as 1 consisting of all element of the form `e * r * e`. -/ +@[nolint unusedArguments] +def _root_.IsIdempotentElem.Corner (_ : IsIdempotentElem e) : Type _ := Subsemigroup.corner e + +end Subsemigroup + +/-- The corner associated to an element `e` in a semiring without 1 +is the subsemiring without 1 of all elements of the form `e * r * e`. -/ +def NonUnitalSubsemiring.corner [NonUnitalSemiring R] : NonUnitalSubsemiring R where + __ := Subsemigroup.corner e + add_mem' := by rintro _ _ ⟨a, rfl⟩ ⟨b, rfl⟩; exact ⟨a + b, by simp_rw [mul_add, add_mul]⟩ + zero_mem' := ⟨0, by simp_rw [mul_zero, zero_mul]⟩ + +/-- The corner associated to an element `e` in a ring without ` +is the subring without 1 of all elements of the form `e * r * e`. -/ +def NonUnitalRing.corner [NonUnitalRing R] : NonUnitalSubring R where + __ := NonUnitalSubsemiring.corner e + neg_mem' := by rintro _ ⟨a, rfl⟩; exact ⟨-a, by simp_rw [mul_neg, neg_mul]⟩ + +instance [NonUnitalSemiring R] (idem : IsIdempotentElem e) : Semiring idem.Corner where + __ : NonUnitalSemiring (NonUnitalSubsemiring.corner e) := inferInstance + one := ⟨e, e, by simp_rw [idem.eq]⟩ + one_mul r := Subtype.ext ((Subsemigroup.mem_corner_iff idem).mp r.2).1 + mul_one r := Subtype.ext ((Subsemigroup.mem_corner_iff idem).mp r.2).2 + +instance [NonUnitalCommSemiring R] (idem : IsIdempotentElem e) : CommSemiring idem.Corner where + __ : NonUnitalCommSemiring (NonUnitalSubsemiring.corner e) := inferInstance + __ : Semiring idem.Corner := inferInstance + +instance [NonUnitalRing R] (idem : IsIdempotentElem e) : Ring idem.Corner where + __ : NonUnitalRing (NonUnitalRing.corner e) := inferInstance + __ : Semiring idem.Corner := inferInstance + +instance [NonUnitalCommRing R] (idem : IsIdempotentElem e) : CommRing idem.Corner where + __ : NonUnitalCommRing (NonUnitalRing.corner e) := inferInstance + __ : Semiring idem.Corner := inferInstance + +variable {I : Type*} [Fintype I] {e : I → R} + +/-- A complete orthogonal family of central idempotents in a semiring +give rise to a direct product decomposition. -/ +def CompleteOrthogonalIdempotents.mulEquivOfIsMulCentral [Semiring R] + (he : CompleteOrthogonalIdempotents e) (hc : ∀ i, IsMulCentral (e i)) : + R ≃+* Π i, (he.idem i).Corner where + toFun r i := ⟨_, r, rfl⟩ + invFun r := ∑ i, (r i).1 + left_inv r := by + simp_rw [(hc _).comm, mul_assoc, (he.idem _).eq, ← Finset.mul_sum, he.complete, mul_one] + right_inv r := funext fun i ↦ Subtype.ext <| by + simp_rw [Finset.mul_sum, Finset.sum_mul] + rw [Finset.sum_eq_single i _ (by simp at ·)] + · have ⟨r', eq⟩ := (r i).2 + rw [← eq]; simp_rw [← mul_assoc, (he.idem i).eq, mul_assoc, (he.idem i).eq] + · intro j _ ne; have ⟨r', eq⟩ := (r j).2 + rw [← eq]; simp_rw [← mul_assoc, he.ortho ne.symm, zero_mul] + map_mul' r₁ r₂ := funext fun i ↦ Subtype.ext <| + calc e i * (r₁ * r₂) * e i + _ = e i * (r₁ * e i * r₂) * e i := by simp_rw [← (hc i).comm r₁, ← mul_assoc, (he.idem i).eq] + _ = e i * r₁ * e i * (e i * r₂ * e i) := by + conv in (r₁ * _ * r₂) => rw [← (he.idem i).eq] + simp_rw [mul_assoc] + map_add' r₁ r₂ := funext fun i ↦ Subtype.ext <| by simpa [mul_add] using add_mul .. + +/-- A complete orthogonal family of idempotents in a commutative semiring +give rise to a direct product decomposition. -/ +def CompleteOrthogonalIdempotents.mulEquivOfComm [CommSemiring R] + (he : CompleteOrthogonalIdempotents e) : R ≃+* Π i, (he.idem i).Corner := + he.mulEquivOfIsMulCentral fun _ ↦ Semigroup.mem_center_iff.mpr fun _ ↦ mul_comm .. + +end corner diff --git a/Mathlib/RingTheory/Int/Basic.lean b/Mathlib/RingTheory/Int/Basic.lean index be4be78dc6d35..6222efbb76ead 100644 --- a/Mathlib/RingTheory/Int/Basic.lean +++ b/Mathlib/RingTheory/Int/Basic.lean @@ -30,28 +30,20 @@ unique factors namespace Int -theorem gcd_eq_one_iff_coprime {a b : ℤ} : Int.gcd a b = 1 ↔ IsCoprime a b := by - constructor - · intro hg - obtain ⟨ua, -, ha⟩ := exists_unit_of_abs a - obtain ⟨ub, -, hb⟩ := exists_unit_of_abs b - use Nat.gcdA (Int.natAbs a) (Int.natAbs b) * ua, Nat.gcdB (Int.natAbs a) (Int.natAbs b) * ub - rw [mul_assoc, ← ha, mul_assoc, ← hb, mul_comm, mul_comm _ (Int.natAbs b : ℤ), ← - Nat.gcd_eq_gcd_ab, ← gcd_eq_natAbs, hg, Int.ofNat_one] - · rintro ⟨r, s, h⟩ - by_contra hg - obtain ⟨p, ⟨hp, ha, hb⟩⟩ := Nat.Prime.not_coprime_iff_dvd.mp hg - apply Nat.Prime.not_dvd_one hp - rw [← natCast_dvd_natCast, Int.ofNat_one, ← h] - exact dvd_add ((natCast_dvd.mpr ha).mul_left _) ((natCast_dvd.mpr hb).mul_left _) - -theorem coprime_iff_nat_coprime {a b : ℤ} : IsCoprime a b ↔ Nat.Coprime a.natAbs b.natAbs := by - rw [← gcd_eq_one_iff_coprime, Nat.coprime_iff_gcd_eq_one, gcd_eq_natAbs] +@[deprecated "use `isCoprime_iff_gcd_eq_one.symm` instead" (since := "2025-01-23")] +theorem gcd_eq_one_iff_coprime {a b : ℤ} : Int.gcd a b = 1 ↔ IsCoprime a b := + isCoprime_iff_gcd_eq_one.symm + + +theorem isCoprime_iff_nat_coprime {a b : ℤ} : IsCoprime a b ↔ Nat.Coprime a.natAbs b.natAbs := by + rw [isCoprime_iff_gcd_eq_one, Nat.coprime_iff_gcd_eq_one, gcd_eq_natAbs] + +@[deprecated (since := "2025-01-23")] alias coprime_iff_nat_coprime := isCoprime_iff_nat_coprime /-- If `gcd a (m * n) ≠ 1`, then `gcd a m ≠ 1` or `gcd a n ≠ 1`. -/ theorem gcd_ne_one_iff_gcd_mul_right_ne_one {a : ℤ} {m n : ℕ} : a.gcd (m * n) ≠ 1 ↔ a.gcd m ≠ 1 ∨ a.gcd n ≠ 1 := by - simp only [gcd_eq_one_iff_coprime, ← not_and_or, not_iff_not, IsCoprime.mul_right_iff] + simp only [← isCoprime_iff_gcd_eq_one, ← not_and_or, not_iff_not, IsCoprime.mul_right_iff] theorem sq_of_gcd_eq_one {a b c : ℤ} (h : Int.gcd a b = 1) (heq : a * b = c ^ 2) : ∃ a0 : ℤ, a = a0 ^ 2 ∨ a = -a0 ^ 2 := by @@ -65,9 +57,11 @@ theorem sq_of_gcd_eq_one {a b c : ℤ} (h : Int.gcd a b = 1) (heq : a * b = c ^ · rw [hu'] simp -theorem sq_of_coprime {a b c : ℤ} (h : IsCoprime a b) (heq : a * b = c ^ 2) : +theorem sq_of_isCoprime {a b c : ℤ} (h : IsCoprime a b) (heq : a * b = c ^ 2) : ∃ a0 : ℤ, a = a0 ^ 2 ∨ a = -a0 ^ 2 := - sq_of_gcd_eq_one (gcd_eq_one_iff_coprime.mpr h) heq + sq_of_gcd_eq_one (isCoprime_iff_gcd_eq_one.mp h) heq + +@[deprecated (since := "2025-01-23")] alias sq_of_coprime := sq_of_isCoprime theorem natAbs_euclideanDomain_gcd (a b : ℤ) : Int.natAbs (EuclideanDomain.gcd a b) = Int.gcd a b := by diff --git a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean index 0f9874f3e5ae8..69833b5a71fed 100644 --- a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean @@ -41,6 +41,11 @@ theorem Algebra.IsIntegral.of_injective (f : A →ₐ[R] B) (hf : Function.Injec [Algebra.IsIntegral R B] : Algebra.IsIntegral R A := ⟨fun _ ↦ (isIntegral_algHom_iff f hf).mp (isIntegral _)⟩ +/-- Homomorphic image of an integral algebra is an integral algebra. -/ +theorem Algebra.IsIntegral.of_surjective [Algebra.IsIntegral R A] + (f : A →ₐ[R] B) (hf : Function.Surjective f) : Algebra.IsIntegral R B := + isIntegral_def.mpr fun b ↦ let ⟨a, ha⟩ := hf b; ha ▸ (isIntegral_def.mp ‹_› a).map f + theorem AlgEquiv.isIntegral_iff (e : A ≃ₐ[R] B) : Algebra.IsIntegral R A ↔ Algebra.IsIntegral R B := ⟨fun h ↦ h.of_injective e.symm e.symm.injective, fun h ↦ h.of_injective e e.injective⟩ diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean index 361784bfb1bc2..c40c5bce9c177 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean @@ -201,6 +201,14 @@ theorem fg_adjoin_of_finite {s : Set A} (hfs : s.Finite) (his : ∀ x ∈ s, IsI (his a <| Set.mem_insert a s).fg_adjoin_singleton) his +theorem Algebra.finite_adjoin_of_finite_of_isIntegral {s : Set A} (hf : s.Finite) + (hi : ∀ x ∈ s, IsIntegral R x) : Module.Finite R (adjoin R s) := + Module.Finite.iff_fg.mpr <| fg_adjoin_of_finite hf hi + +theorem Algebra.finite_adjoin_simple_of_isIntegral {x : A} (hi : IsIntegral R x) : + Module.Finite R (adjoin R {x}) := + Module.Finite.iff_fg.mpr hi.fg_adjoin_singleton + theorem isNoetherian_adjoin_finset [IsNoetherianRing R] (s : Finset A) (hs : ∀ x ∈ s, IsIntegral R x) : IsNoetherian R (Algebra.adjoin R (s : Set A)) := isNoetherian_of_fg_of_noetherian _ (fg_adjoin_of_finite s.finite_toSet hs) diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 5b810da90a937..ff4bf59739f94 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -56,7 +56,7 @@ theorem IsIntegral.inv_mem_adjoin (int : IsIntegral R x) : x⁻¹ ∈ adjoin R { /-- The inverse of an integral element in a subalgebra of a division ring over a field also lies in that subalgebra. -/ theorem IsIntegral.inv_mem (int : IsIntegral R x) (hx : x ∈ A) : x⁻¹ ∈ A := - adjoin_le_iff.mpr (Set.singleton_subset_iff.mpr hx) int.inv_mem_adjoin + adjoin_le (Set.singleton_subset_iff.mpr hx) int.inv_mem_adjoin /-- An integral subalgebra of a division ring over a field is closed under inverses. -/ theorem Algebra.IsIntegral.inv_mem [Algebra.IsIntegral R A] (hx : x ∈ A) : x⁻¹ ∈ A := @@ -132,8 +132,8 @@ theorem le_integralClosure_iff_isIntegral {S : Subalgebra R A} : Algebra.isIntegral_def.symm theorem Algebra.IsIntegral.adjoin {S : Set A} (hS : ∀ x ∈ S, IsIntegral R x) : - Algebra.IsIntegral R (Algebra.adjoin R S) := - le_integralClosure_iff_isIntegral.mp <| adjoin_le_iff.mpr hS + Algebra.IsIntegral R (adjoin R S) := + le_integralClosure_iff_isIntegral.mp <| adjoin_le hS theorem integralClosure_eq_top_iff : integralClosure R A = ⊤ ↔ Algebra.IsIntegral R A := by rw [← top_le_iff, le_integralClosure_iff_isIntegral, diff --git a/Mathlib/RingTheory/Invariant.lean b/Mathlib/RingTheory/Invariant.lean index 075f78efeb749..fe8623572f750 100644 --- a/Mathlib/RingTheory/Invariant.lean +++ b/Mathlib/RingTheory/Invariant.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.FieldTheory.Fixed +import Mathlib.FieldTheory.Galois.Basic import Mathlib.RingTheory.Ideal.Over +import Mathlib.RingTheory.IntegralClosure.IntegralRestrict /-! # Invariant Extensions of Rings and Frobenius Elements @@ -46,6 +47,40 @@ by `G` is `smul_algebraMap` (assuming `SMulCommClass A B G`). -/ end Algebra +section Galois + +variable (A K L B : Type*) [CommRing A] [CommRing B] [Field K] [Field L] + [Algebra A K] [Algebra B L] [IsFractionRing A K] [IsFractionRing B L] + [Algebra A B] [Algebra K L] [Algebra A L] [IsScalarTower A K L] [IsScalarTower A B L] + [IsIntegrallyClosed A] [IsIntegralClosure B A L] + +/-- In the AKLB setup, the Galois group of `L/K` acts on `B`. -/ +noncomputable def IsIntegralClosure.MulSemiringAction [Algebra.IsAlgebraic K L] : + MulSemiringAction (L ≃ₐ[K] L) B := + MulSemiringAction.compHom B (galRestrict A K L B).toMonoidHom + +/-- In the AKLB setup, every fixed point of `B` lies in the image of `A`. -/ +theorem Algebra.isInvariant_of_isGalois [FiniteDimensional K L] [h : IsGalois K L] : + letI := IsIntegralClosure.MulSemiringAction A K L B + Algebra.IsInvariant A B (L ≃ₐ[K] L) := by + replace h := ((IsGalois.tfae (F := K) (E := L)).out 0 1).mp h + letI := IsIntegralClosure.MulSemiringAction A K L B + refine ⟨fun b hb ↦ ?_⟩ + replace hb : algebraMap B L b ∈ IntermediateField.fixedField (⊤ : Subgroup (L ≃ₐ[K] L)) := by + rintro ⟨g, -⟩ + exact (algebraMap_galRestrict_apply A g b).symm.trans (congrArg (algebraMap B L) (hb g)) + rw [h, IntermediateField.mem_bot] at hb + obtain ⟨k, hk⟩ := hb + have hb : IsIntegral A b := IsIntegralClosure.isIntegral A L b + rw [← isIntegral_algebraMap_iff (NoZeroSMulDivisors.algebraMap_injective B L), ← hk, + isIntegral_algebraMap_iff (NoZeroSMulDivisors.algebraMap_injective K L)] at hb + obtain ⟨a, rfl⟩ := IsIntegrallyClosed.algebraMap_eq_of_integral hb + rw [← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply A B L, + (NoZeroSMulDivisors.algebraMap_injective B L).eq_iff] at hk + exact ⟨a, hk⟩ + +end Galois + section transitivity variable (A B G : Type*) [CommRing A] [CommRing B] [Algebra A B] [Group G] [MulSemiringAction G B] diff --git a/Mathlib/RingTheory/IsAdjoinRoot.lean b/Mathlib/RingTheory/IsAdjoinRoot.lean index c607d73842aef..cd89996babce3 100644 --- a/Mathlib/RingTheory/IsAdjoinRoot.lean +++ b/Mathlib/RingTheory/IsAdjoinRoot.lean @@ -80,8 +80,6 @@ and `AdjoinRoot` which constructs a new type. This is not a typeclass because the choice of root given `S` and `f` is not unique. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure IsAdjoinRoot {R : Type u} (S : Type v) [CommSemiring R] [Semiring S] [Algebra R S] (f : R[X]) : Type max u v where map : R[X] →+* S diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 2ca02e4c274e9..3c8dd440d4251 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.RingTheory.Kaehler.Polynomial -import Mathlib.RingTheory.Generators +import Mathlib.Algebra.Module.FinitePresentation +import Mathlib.RingTheory.Presentation /-! @@ -394,6 +395,9 @@ lemma cotangentSpaceBasis_apply (i) : P.cotangentSpaceBasis i = ((1 : S) ⊗ₜ[P.Ring] D R P.Ring (.X i) :) := by simp [cotangentSpaceBasis, toExtension] +instance (P : Generators R S) : Module.Free S P.toExtension.CotangentSpace := + .of_basis P.cotangentSpaceBasis + universe w' u' v' variable {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] @@ -433,6 +437,16 @@ lemma toKaehler_cotangentSpaceBasis (i) : end Generators +-- TODO: generalize to essentially of finite presentation algebras +open KaehlerDifferential in +attribute [local instance] Module.finitePresentation_of_projective in +instance [Algebra.FinitePresentation R S] : Module.FinitePresentation S (Ω[S⁄R]) := by + let P := Algebra.Presentation.ofFinitePresentation R S + have : Algebra.FiniteType R P.toExtension.Ring := .mvPolynomial _ _ + refine Module.finitePresentation_of_surjective _ P.toExtension.toKaehler_surjective ?_ + rw [LinearMap.exact_iff.mp P.toExtension.exact_cotangentComplex_toKaehler, ← Submodule.map_top] + exact (Extension.Cotangent.finite P.ideal_fg_of_isFinite).1.map P.toExtension.cotangentComplex + variable {P : Generators R S} open Extension.H1Cotangent in @@ -458,6 +472,28 @@ noncomputable def H1Cotangent.map : H1Cotangent R S' →ₗ[S'] H1Cotangent S T := Extension.H1Cotangent.map (Generators.defaultHom _ _).toExtensionHom +/-- Isomorphic algebras induce isomorphic `H¹(L_{S/R})`. -/ +noncomputable +def H1Cotangent.mapEquiv (e : S ≃ₐ[R] S') : + H1Cotangent R S ≃ₗ[R] H1Cotangent R S' := + -- we are constructing data, so we do not use `algebraize` + letI := e.toRingHom.toAlgebra + letI := e.symm.toRingHom.toAlgebra + have : IsScalarTower R S S' := .of_algebraMap_eq' e.toAlgHom.comp_algebraMap.symm + have : IsScalarTower R S' S := .of_algebraMap_eq' e.symm.toAlgHom.comp_algebraMap.symm + have : IsScalarTower S S' S := .of_algebraMap_eq fun _ ↦ (e.symm_apply_apply _).symm + have : IsScalarTower S' S S' := .of_algebraMap_eq fun _ ↦ (e.apply_symm_apply _).symm + { __ := map R R S S' + invFun := map R R S' S + left_inv x := by + show ((map R R S' S).restrictScalars S ∘ₗ map R R S S') x = x + rw [map, map, ← Extension.H1Cotangent.map_comp, Extension.H1Cotangent.map_eq, + Extension.H1Cotangent.map_id, LinearMap.id_apply] + right_inv x := by + show ((map R R S S').restrictScalars S' ∘ₗ map R R S' S) x = x + rw [map, map, ← Extension.H1Cotangent.map_comp, Extension.H1Cotangent.map_eq, + Extension.H1Cotangent.map_id, LinearMap.id_apply] } + variable {R S S' T} /-- `H¹(L_{S/R})` is independent of the presentation chosen. -/ @@ -466,4 +502,21 @@ abbrev Generators.equivH1Cotangent (P : Generators.{w} R S) : P.toExtension.H1Cotangent ≃ₗ[S] H1Cotangent R S := Generators.H1Cotangent.equiv _ _ +attribute [local instance] Module.finitePresentation_of_projective in +instance [FinitePresentation R S] [Module.Projective S (Ω[S⁄R])] : + Module.Finite S (H1Cotangent R S) := by + let P := Algebra.Presentation.ofFinitePresentation R S + have : Algebra.FiniteType R P.toExtension.Ring := FiniteType.mvPolynomial R P.vars + suffices Module.Finite S P.toExtension.H1Cotangent from + .of_surjective P.equivH1Cotangent.toLinearMap P.equivH1Cotangent.surjective + rw [Module.finite_def, Submodule.fg_top, ← LinearMap.ker_rangeRestrict] + have := Extension.Cotangent.finite P.ideal_fg_of_isFinite + have : Module.FinitePresentation S (LinearMap.range P.toExtension.cotangentComplex) := by + rw [← LinearMap.exact_iff.mp P.toExtension.exact_cotangentComplex_toKaehler] + exact Module.finitePresentation_of_projective_of_exact + _ _ (Subtype.val_injective) P.toExtension.toKaehler_surjective + (LinearMap.exact_subtype_ker_map _) + exact Module.FinitePresentation.fg_ker (N := LinearMap.range P.toExtension.cotangentComplex) + _ P.toExtension.cotangentComplex.surjective_rangeRestrict + end Algebra diff --git a/Mathlib/RingTheory/KrullDimension/Basic.lean b/Mathlib/RingTheory/KrullDimension/Basic.lean index b368cc66ae325..13baff3915c07 100644 --- a/Mathlib/RingTheory/KrullDimension/Basic.lean +++ b/Mathlib/RingTheory/KrullDimension/Basic.lean @@ -5,6 +5,7 @@ Authors: Fangming Li, Jujian Zhang -/ import Mathlib.Algebra.MvPolynomial.CommRing import Mathlib.Algebra.Polynomial.Basic +import Mathlib.RingTheory.Ideal.Quotient.Defs import Mathlib.RingTheory.Spectrum.Prime.Basic import Mathlib.Order.KrullDimension @@ -21,7 +22,7 @@ open Order /-- The ring theoretic Krull dimension is the Krull dimension of its spectrum ordered by inclusion. -/ -noncomputable def ringKrullDim (R : Type*) [CommRing R] : WithBot (WithTop ℕ) := +noncomputable def ringKrullDim (R : Type*) [CommRing R] : WithBot ℕ∞ := krullDim (PrimeSpectrum R) variable {R S : Type*} [CommRing R] [CommRing S] diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index dbbef917b461d..27e5f145a98c1 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -14,6 +14,7 @@ import Mathlib.RingTheory.Localization.FractionRing import Mathlib.Topology.UniformSpace.Cauchy import Mathlib.Algebra.Group.Int.TypeTags + /-! # Laurent Series @@ -24,7 +25,7 @@ type with a zero. They are denoted `R⸨X⸩`. * Defines `LaurentSeries` as an abbreviation for `HahnSeries ℤ`. * Defines `hasseDeriv` of a Laurent series with coefficients in a module over a ring. -* Provides a coercion `from power series `R⟦X⟧` into `R⸨X⸩` given by `HahnSeries.ofPowerSeries`. +* Provides a coercion from power series `R⟦X⟧` into `R⸨X⸩` given by `HahnSeries.ofPowerSeries`. * Defines `LaurentSeries.powerSeriesPart` * Defines the localization map `LaurentSeries.of_powerSeries_localization` which evaluates to `HahnSeries.ofPowerSeries`. @@ -60,8 +61,8 @@ such that the `X`-adic valuation `v` satisfies `v (f - Q) < γ`. `LaurentSeries.exists_powerSeries_of_memIntegers` show that an element in the completion of `RatFunc K` is in the unit ball if and only if it comes from a power series through the isomorphism `LaurentSeriesRingEquiv`. -* `LaurentSeries.powerSeriesRingEquiv` is the ring isomorphism between `K⟦X⟧` and the unit ball -inside the `X`-adic completion of `RatFunc K`. +* `LaurentSeries.powerSeriesAlgEquiv` is the `K`-algebra isomorphism between `K⟦X⟧` +and the unit ball inside the `X`-adic completion of `RatFunc K`. ## Implementation details @@ -75,21 +76,17 @@ that it is complete and contains `RatFunc K` as a dense subspace. The isomorphis equivalence, expressing the mathematical idea that the completion "is unique". It is `LaurentSeries.comparePkg`. * For applications to `K⟦X⟧` it is actually more handy to use the *inverse* of the above -equivalence: `LaurentSeries.LaurentSeriesRingEquiv` is the *topological, ring equivalence* -`K⸨X⸩ ≃+* RatFuncAdicCompl K`. +equivalence: `LaurentSeries.LaurentSeriesAlgEquiv` is the *topological, algebra equivalence* +`K⸨X⸩ ≃ₐ[K] RatFuncAdicCompl K`. * In order to compare `K⟦X⟧` with the valuation subring in the `X`-adic completion of `RatFunc K` we consider its alias `LaurentSeries.powerSeries_as_subring` as a subring of `K⸨X⸩`, -that is itself clearly isomorphic (via `LaurentSeries.powerSeriesEquivSubring.symm`) to `K⟦X⟧`. +that is itself clearly isomorphic (via the inverse of `LaurentSeries.powerSeriesEquivSubring`) +to `K⟦X⟧`. -## To Do -* The `AdicCompletion` construction is currently done for ideals in rings and does not take into -account the relation with algebra structures on the ring, hence it does not yield a `K`-algebra -structure on the `X`-adic completion of `K⸨X⸩`. Once this will be available, we should update -`LaurentSeries.LaurentSeriesRingEquiv` to an algebra equivalence. -/ universe u -open scoped Classical PowerSeries +open scoped PowerSeries open HahnSeries Polynomial noncomputable section @@ -541,6 +538,7 @@ theorem intValuation_eq_of_coe (P : K[X]) : (Ideal.span {↑P} : Ideal K⟦X⟧) ≠ 0 ∧ ((idealX K).asIdeal : Ideal K⟦X⟧) ≠ 0 := by simp only [Ideal.zero_eq_bot, ne_eq, Ideal.span_singleton_eq_bot, coe_eq_zero_iff, hP, not_false_eq_true, true_and, (idealX K).3] + classical rw [count_associates_factors_eq (span_ne_zero).1 (Ideal.span_singleton_prime Polynomial.X_ne_zero|>.mpr prime_X) (span_ne_zero).2, count_associates_factors_eq] @@ -1059,6 +1057,11 @@ equivalence: it goes from `K⸨X⸩` to `RatFuncAdicCompl K` -/ abbrev LaurentSeriesRingEquiv : K⸨X⸩ ≃+* RatFuncAdicCompl K := (ratfuncAdicComplRingEquiv K).symm +@[simp] +lemma LaurentSeriesRingEquiv_def (f : K⟦X⟧) : + (LaurentSeriesRingEquiv K) f = (LaurentSeriesPkg K).compare ratfuncAdicComplPkg (f : K⸨X⸩) := + rfl + @[simp] theorem ratfuncAdicComplRingEquiv_apply (x : RatFuncAdicCompl K) : ratfuncAdicComplRingEquiv K x = ratfuncAdicComplPkg.compare (LaurentSeriesPkg K) x := rfl @@ -1069,6 +1072,17 @@ theorem coe_X_compare : rw [PowerSeries.coe_X, ← RatFunc.coe_X, ← LaurentSeries_coe, ← compare_coe] rfl +theorem algebraMap_apply (a : K) : algebraMap K K⸨X⸩ a = HahnSeries.C a := by + simp [RingHom.algebraMap_toAlgebra] + +instance : Algebra K (RatFuncAdicCompl K) := + RingHom.toAlgebra ((LaurentSeriesRingEquiv K).toRingHom.comp HahnSeries.C) + +/-- The algebra equivalence between `K⸨X⸩` and the `X`-adic completion of `RatFunc X` -/ +def LaurentSeriesAlgEquiv : K⸨X⸩ ≃ₐ[K] RatFuncAdicCompl K := + AlgEquiv.ofRingEquiv (f := LaurentSeriesRingEquiv K) + (fun a ↦ by simp [RingHom.algebraMap_toAlgebra]) + open Filter WithZero open scoped WithZeroTopology Topology Multiplicative @@ -1130,22 +1144,30 @@ section PowerSeries /-- In order to compare `K⟦X⟧` with the valuation subring in the `X`-adic completion of `RatFunc K` we consider its alias as a subring of `K⸨X⸩`. -/ abbrev powerSeries_as_subring : Subring K⸨X⸩ := - RingHom.range (HahnSeries.ofPowerSeries ℤ K) +-- RingHom.range (HahnSeries.ofPowerSeries ℤ K) + Subring.map (HahnSeries.ofPowerSeries ℤ K) ⊤ /-- The ring `K⟦X⟧` is isomorphic to the subring `powerSeries_as_subring K` -/ -abbrev powerSeriesEquivSubring : K⟦X⟧ ≃+* powerSeries_as_subring K := by - rw [powerSeries_as_subring, RingHom.range_eq_map] - exact ((Subring.topEquiv).symm).trans (Subring.equivMapOfInjective ⊤ (ofPowerSeries ℤ K) +abbrev powerSeriesEquivSubring : K⟦X⟧ ≃+* powerSeries_as_subring K := + ((Subring.topEquiv).symm).trans (Subring.equivMapOfInjective ⊤ (ofPowerSeries ℤ K) ofPowerSeries_injective) +lemma powerSeriesEquivSubring_apply (f : K⟦X⟧) : + powerSeriesEquivSubring K f = + ⟨HahnSeries.ofPowerSeries ℤ K f, Subring.mem_map.mpr ⟨f, trivial, rfl⟩⟩ := + rfl + +lemma powerSeriesEquivSubring_coe_apply (f : K⟦X⟧) : + (powerSeriesEquivSubring K f : K⸨X⸩) = ofPowerSeries ℤ K f := + rfl + /- Through the isomorphism `LaurentSeriesRingEquiv`, power series land in the unit ball inside the completion of `RatFunc K`. -/ theorem mem_integers_of_powerSeries (F : K⟦X⟧) : (LaurentSeriesRingEquiv K) F ∈ (idealX K).adicCompletionIntegers (RatFunc K) := by - have : (LaurentSeriesRingEquiv K) F = - (LaurentSeriesPkg K).compare ratfuncAdicComplPkg (F : K⸨X⸩) := rfl simp only [Subring.mem_map, exists_prop, ValuationSubring.mem_toSubring, - mem_adicCompletionIntegers, this, valuation_compare, val_le_one_iff_eq_coe] + mem_adicCompletionIntegers, LaurentSeriesRingEquiv_def, + valuation_compare, val_le_one_iff_eq_coe] exact ⟨F, rfl⟩ /- Conversely, all elements in the unit ball inside the completion of `RatFunc K` come from a power @@ -1164,13 +1186,13 @@ theorem powerSeries_ext_subring : Subring.map (LaurentSeriesRingEquiv K).toRingHom (powerSeries_as_subring K) = ((idealX K).adicCompletionIntegers (RatFunc K)).toSubring := by ext x - refine ⟨fun ⟨f, ⟨F, coe_F⟩, hF⟩ ↦ ?_, fun H ↦ ?_⟩ + refine ⟨fun ⟨f, ⟨F, _, coe_F⟩, hF⟩ ↦ ?_, fun H ↦ ?_⟩ · simp only [ValuationSubring.mem_toSubring, ← hF, ← coe_F] apply mem_integers_of_powerSeries · obtain ⟨F, hF⟩ := exists_powerSeries_of_memIntegers K H simp only [Equiv.toFun_as_coe, UniformEquiv.coe_toEquiv, exists_exists_eq_and, UniformEquiv.coe_symm_toEquiv, Subring.mem_map, Equiv.invFun_as_coe] - exact ⟨F, ⟨F, rfl⟩, hF⟩ + exact ⟨F, ⟨F, trivial, rfl⟩, hF⟩ /-- The ring isomorphism between `K⟦X⟧` and the unit ball inside the `X`-adic completion of `RatFunc K`. -/ @@ -1178,6 +1200,43 @@ abbrev powerSeriesRingEquiv : K⟦X⟧ ≃+* (idealX K).adicCompletionIntegers ( ((powerSeriesEquivSubring K).trans (LaurentSeriesRingEquiv K).subringMap).trans <| RingEquiv.subringCongr (powerSeries_ext_subring K) +lemma powerSeriesRingEquiv_coe_apply (f : K⟦X⟧) : + powerSeriesRingEquiv K f = LaurentSeriesRingEquiv K (f : K⸨X⸩) := + rfl + +lemma LaurentSeriesRingEquiv_mem_valuationSubring (f : K⟦X⟧) : + LaurentSeriesRingEquiv K f ∈ Valued.v.valuationSubring := by + simp only [Valuation.mem_valuationSubring_iff] + rw [LaurentSeriesRingEquiv_def, valuation_compare, val_le_one_iff_eq_coe] + use f + +lemma algebraMap_C_mem_adicCompletionIntegers (x : K) : + ((LaurentSeriesRingEquiv K).toRingHom.comp HahnSeries.C) x ∈ + adicCompletionIntegers (RatFunc K) (idealX K) := by + have : HahnSeries.C x = ofPowerSeries ℤ K (PowerSeries.C K x) := by + simp [C_apply, ofPowerSeries_C] + simp only [RingHom.comp_apply, RingEquiv.toRingHom_eq_coe, RingHom.coe_coe, this] + apply LaurentSeriesRingEquiv_mem_valuationSubring + +instance : Algebra K ((idealX K).adicCompletionIntegers (RatFunc K)) := + RingHom.toAlgebra <| + ((LaurentSeriesRingEquiv K).toRingHom.comp HahnSeries.C).codRestrict _ + (algebraMap_C_mem_adicCompletionIntegers K) + +instance : IsScalarTower K ((idealX K).adicCompletionIntegers (RatFunc K)) + ((idealX K).adicCompletion (RatFunc K)) := + IsScalarTower.of_algebraMap_eq (fun _ ↦ by rfl) + +/-- The algebra isomorphism between `K⟦X⟧` and the unit ball inside the `X`-adic completion of +`RatFunc K`. -/ +def powerSeriesAlgEquiv : K⟦X⟧ ≃ₐ[K] (idealX K).adicCompletionIntegers (RatFunc K) := by + apply AlgEquiv.ofRingEquiv (f := powerSeriesRingEquiv K) + intro a + rw [PowerSeries.algebraMap_eq, RingHom.algebraMap_toAlgebra, ← Subtype.coe_inj, + powerSeriesRingEquiv_coe_apply, + RingHom.codRestrict_apply _ _ (algebraMap_C_mem_adicCompletionIntegers K)] + simp + end PowerSeries end Comparison diff --git a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean index 2c5a8113251ff..e2b6096cd016d 100644 --- a/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean +++ b/Mathlib/RingTheory/LocalRing/RingHom/Basic.lean @@ -86,8 +86,8 @@ section variable [CommSemiring R] [IsLocalRing R] [CommSemiring S] [IsLocalRing S] /-- A ring homomorphism between local rings is a local ring hom iff it reflects units, -i.e. any preimage of a unit is still a unit. https://stacks.math.columbia.edu/tag/07BJ --/ +i.e. any preimage of a unit is still a unit. -/ +@[stacks 07BJ] theorem local_hom_TFAE (f : R →+* S) : List.TFAE [IsLocalHom f, f '' (maximalIdeal R).1 ⊆ maximalIdeal S, diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index e8d94b851a13e..aa913923a7e8f 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -49,6 +49,12 @@ def primeCompl : Submonoid R where theorem primeCompl_le_nonZeroDivisors [NoZeroDivisors R] : P.primeCompl ≤ nonZeroDivisors R := le_nonZeroDivisors_of_noZeroDivisors <| not_not_intro P.zero_mem +lemma disjoint_map_primeCompl_iff_comap_le {S : Type*} [Semiring S] {f : R →+* S} + {p : Ideal R} {I : Ideal S} [p.IsPrime] : + Disjoint (I : Set S) (p.primeCompl.map f) ↔ I.comap f ≤ p := by + rw [disjoint_comm] + simp [Set.disjoint_iff, Set.ext_iff, Ideal.primeCompl, not_imp_not, SetLike.le_def] + end Ideal /-- Given a prime ideal `P`, the typeclass `IsLocalization.AtPrime S P` states that `S` is diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 22ac6c09fc61e..8b0b1e1417be3 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -296,6 +296,11 @@ theorem isLocalization_iff_of_isLocalization [IsLocalization M S] [IsLocalizatio ⟨fun _ ↦ isLocalization_of_algEquiv N (algEquiv M S P), fun _ ↦ isLocalization_of_algEquiv M (algEquiv N S P)⟩ +theorem iff_of_le_of_exists_dvd (N : Submonoid R) (h₁ : M ≤ N) (h₂ : ∀ n ∈ N, ∃ m ∈ M, n ∣ m) : + IsLocalization M S ↔ IsLocalization N S := + have : IsLocalization N (Localization M) := of_le_of_exists_dvd _ _ h₁ h₂ + isLocalization_iff_of_isLocalization _ _ (Localization M) + end variable (M) diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index fc1af5344f892..5238e642ff74e 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -80,6 +80,17 @@ lemma mk'_mem_map_algebraMap_iff (I : Ideal R) (x : R) (s : M) : exact ⟨fun ⟨⟨y, t⟩, c, h⟩ ↦ ⟨_, (c * t).2, h ▸ I.mul_mem_left c.1 y.2⟩, fun ⟨s, hs, h⟩ ↦ ⟨⟨⟨_, h⟩, ⟨s, hs⟩⟩, 1, by simp⟩⟩ +lemma algebraMap_mem_map_algebraMap_iff (I : Ideal R) (x : R) : + algebraMap R S x ∈ I.map (algebraMap R S) ↔ + ∃ m ∈ M, m * x ∈ I := by + rw [← IsLocalization.mk'_one (M := M), mk'_mem_map_algebraMap_iff] + +lemma map_algebraMap_ne_top_iff_disjoint (I : Ideal R) : + I.map (algebraMap R S) ≠ ⊤ ↔ Disjoint (M : Set R) (I : Set R) := by + simp only [ne_eq, Ideal.eq_top_iff_one, ← map_one (algebraMap R S), not_iff_comm, + IsLocalization.algebraMap_mem_map_algebraMap_iff M] + simp [Set.disjoint_left] + include M in theorem map_comap (J : Ideal S) : Ideal.map (algebraMap R S) (Ideal.comap (algebraMap R S) J) = J := diff --git a/Mathlib/RingTheory/Localization/Pi.lean b/Mathlib/RingTheory/Localization/Pi.lean new file mode 100644 index 0000000000000..28a3911c43cd7 --- /dev/null +++ b/Mathlib/RingTheory/Localization/Pi.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2024 Madison Crim. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Madison Crim +-/ +import Mathlib.Algebra.Algebra.Pi +import Mathlib.Algebra.BigOperators.Pi +import Mathlib.Algebra.Divisibility.Prod +import Mathlib.Algebra.Group.Submonoid.BigOperators +import Mathlib.Algebra.Group.Subgroup.Basic +import Mathlib.RingTheory.Localization.Basic +import Mathlib.Algebra.Group.Pi.Units + +/-! +# Localizing a product of commutative rings + +## Main Result + + * `bijective_lift_piRingHom_algebraMap_comp_piEvalRingHom`: the canonical map from a + localization of a finite product of rings `R i `at a monoid `M` to the direct product of + localizations `R i` at the projection of `M` onto each corresponding factor is bijective. + +## Implementation notes + +See `Mathlib/RingTheory/Localization/Defs.lean` for a design overview. + +## Tags +localization, commutative ring +-/ + +namespace IsLocalization + +variable {ι : Type*} (R S : ι → Type*) + [Π i, CommSemiring (R i)] [Π i, CommSemiring (S i)] [Π i, Algebra (R i) (S i)] + +/-- If `S i` is a localization of `R i` at the submonoid `M i` for each `i`, +then `Π i, S i` is a localization of `Π i, R i` at the product submonoid. -/ +instance (M : Π i, Submonoid (R i)) [∀ i, IsLocalization (M i) (S i)] : + IsLocalization (.pi .univ M) (Π i, S i) where + map_units' m := Pi.isUnit_iff.mpr fun i ↦ map_units _ ⟨m.1 i, m.2 i ⟨⟩⟩ + surj' z := by + choose rm h using fun i ↦ surj (M := M i) (z i) + exact ⟨(fun i ↦ (rm i).1, ⟨_, fun i _ ↦ (rm i).2.2⟩), funext h⟩ + exists_of_eq {x y} eq := by + choose c hc using fun i ↦ exists_of_eq (M := M i) (congr_fun eq i) + exact ⟨⟨_, fun i _ ↦ (c i).2⟩, funext hc⟩ + +variable (S' : Type*) [CommSemiring S'] [Algebra (Π i, R i) S'] (M : Submonoid (Π i, R i)) + +theorem iff_map_piEvalRingHom [Finite ι] : + IsLocalization M S' ↔ IsLocalization (.pi .univ fun i ↦ M.map (Pi.evalRingHom R i)) S' := + iff_of_le_of_exists_dvd M _ (fun m hm i _ ↦ ⟨m, hm, rfl⟩) fun n hn ↦ by + choose m mem eq using hn + have := Fintype.ofFinite ι + refine ⟨∏ i, m i ⟨⟩, prod_mem fun i _ ↦ mem i _, pi_dvd_iff.mpr fun i ↦ ?_⟩ + rw [Fintype.prod_apply] + exact (eq i ⟨⟩).symm.dvd.trans (Finset.dvd_prod_of_mem _ <| Finset.mem_univ _) + +variable [∀ i, IsLocalization (M.map (Pi.evalRingHom R i)) (S i)] + +/-- Let `M` be a submonoid of a direct product of commutative rings `R i`, and let `M' i` denote +the projection of `M` onto each corresponding factor. Given a ring homomorphism from the direct +product `Π i, R i` to the product of the localizations of each `R i` at `M' i`, every `y : M` +maps to a unit under this homomorphism. -/ +lemma isUnit_piRingHom_algebraMap_comp_piEvalRingHom (y : M) : + IsUnit ((Pi.ringHom fun i ↦ (algebraMap (R i) (S i)).comp (Pi.evalRingHom R i)) y) := + Pi.isUnit_iff.mpr fun i ↦ map_units _ (⟨y.1 i, y, y.2, rfl⟩ : M.map (Pi.evalRingHom R i)) + +/-- Let `M` be a submonoid of a direct product of commutative rings `R i`, and let `M' i` denote +the projection of `M` onto each factor. Then the canonical map from the localization of the direct +product `Π i, R i` at `M` to the direct product of the localizations of each `R i` at `M' i` +is bijective. -/ +theorem bijective_lift_piRingHom_algebraMap_comp_piEvalRingHom [IsLocalization M S'] [Finite ι] : + Function.Bijective (lift (S := S') (isUnit_piRingHom_algebraMap_comp_piEvalRingHom R S M)) := + have := (iff_map_piEvalRingHom R (Π i, S i) M).mpr inferInstance + (ringEquivOfRingEquiv (M := M) (T := M) _ _ (.refl _) <| + Submonoid.map_equiv_eq_comap_symm _ _).bijective + +end IsLocalization diff --git a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean index 35474da03eabb..ecf3da187d103 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Basic.lean @@ -342,6 +342,17 @@ instance center.instNonUnitalCommRing : NonUnitalCommRing (center R) := { NonUnitalSubsemiring.center.instNonUnitalCommSemiring R, inferInstanceAs <| NonUnitalNonAssocRing (center R) with } +variable {R} + +/-- The center of isomorphic (not necessarily unital or associative) rings are isomorphic. -/ +@[simps!] def centerCongr {S} [NonUnitalNonAssocRing S] (e : R ≃+* S) : center R ≃+* center S := + NonUnitalSubsemiring.centerCongr e + +/-- The center of a (not necessarily uintal or associative) ring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ := + NonUnitalSubsemiring.centerToMulOpposite + end NonUnitalNonAssocRing section NonUnitalRing diff --git a/Mathlib/RingTheory/NonUnitalSubring/Defs.lean b/Mathlib/RingTheory/NonUnitalSubring/Defs.lean index b0d51fc219306..b408644c2fc0e 100644 --- a/Mathlib/RingTheory/NonUnitalSubring/Defs.lean +++ b/Mathlib/RingTheory/NonUnitalSubring/Defs.lean @@ -5,6 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Algebra.Group.Subgroup.Defs import Mathlib.RingTheory.NonUnitalSubsemiring.Defs +import Mathlib.Tactic.FastInstance /-! # `NonUnitalSubring`s @@ -60,21 +61,21 @@ namespace NonUnitalSubringClass -- Prefer subclasses of `NonUnitalRing` over subclasses of `NonUnitalSubringClass`. /-- A non-unital subring of a non-unital ring inherits a non-unital ring structure -/ -instance (priority := 75) toNonUnitalNonAssocRing : NonUnitalNonAssocRing s := +instance (priority := 75) toNonUnitalNonAssocRing : NonUnitalNonAssocRing s := fast_instance% Subtype.val_injective.nonUnitalNonAssocRing _ rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -- Prefer subclasses of `NonUnitalRing` over subclasses of `NonUnitalSubringClass`. /-- A non-unital subring of a non-unital ring inherits a non-unital ring structure -/ instance (priority := 75) toNonUnitalRing {R : Type*} [NonUnitalRing R] [SetLike S R] - [NonUnitalSubringClass S R] (s : S) : NonUnitalRing s := + [NonUnitalSubringClass S R] (s : S) : NonUnitalRing s := fast_instance% Subtype.val_injective.nonUnitalRing _ rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl -- Prefer subclasses of `NonUnitalRing` over subclasses of `NonUnitalSubringClass`. /-- A non-unital subring of a `NonUnitalCommRing` is a `NonUnitalCommRing`. -/ instance (priority := 75) toNonUnitalCommRing {R} [NonUnitalCommRing R] [SetLike S R] - [NonUnitalSubringClass S R] : NonUnitalCommRing s := + [NonUnitalSubringClass S R] : NonUnitalCommRing s := fast_instance% Subtype.val_injective.nonUnitalCommRing _ rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -249,7 +250,7 @@ protected theorem sub_mem {x y : R} (hx : x ∈ s) (hy : y ∈ s) : x - y ∈ s /-- A non-unital subring of a non-unital ring inherits a non-unital ring structure -/ instance toNonUnitalRing {R : Type*} [NonUnitalRing R] (s : NonUnitalSubring R) : - NonUnitalRing s := + NonUnitalRing s := fast_instance% Subtype.coe_injective.nonUnitalRing _ rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl @@ -277,7 +278,7 @@ theorem coe_eq_zero_iff {x : s} : (x : R) = 0 ↔ x = 0 := by /-- A non-unital subring of a `NonUnitalCommRing` is a `NonUnitalCommRing`. -/ instance toNonUnitalCommRing {R} [NonUnitalCommRing R] (s : NonUnitalSubring R) : - NonUnitalCommRing s := + NonUnitalCommRing s := fast_instance% Subtype.coe_injective.nonUnitalCommRing _ rfl (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index 61c5c358fc66c..76b1cd54c6879 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -10,8 +10,10 @@ import Mathlib.Algebra.Group.Subsemigroup.Operations import Mathlib.Algebra.GroupWithZero.Center import Mathlib.Algebra.Ring.Center import Mathlib.Algebra.Ring.Centralizer +import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.Ring.Prod import Mathlib.Data.Set.Finite.Range +import Mathlib.GroupTheory.Submonoid.Center import Mathlib.GroupTheory.Subsemigroup.Centralizer import Mathlib.RingTheory.NonUnitalSubsemiring.Defs @@ -245,6 +247,19 @@ lemma _root_.Set.mem_center_iff_addMonoidHom (a : R) : rw [Set.mem_center_iff, isMulCentral_iff] simp [DFunLike.ext_iff] +variable {R} + +/-- The center of isomorphic (not necessarily unital or associative) semirings are isomorphic. -/ +@[simps!] def centerCongr [NonUnitalNonAssocSemiring S] (e : R ≃+* S) : center R ≃+* center S where + __ := Subsemigroup.centerCongr e + map_add' _ _ := Subtype.ext <| by exact map_add e .. + +/-- The center of a (not necessarily unital or associative) semiring +is isomorphic to the center of its opposite. -/ +@[simps!] def centerToMulOpposite : center R ≃+* center Rᵐᵒᵖ where + __ := Subsemigroup.centerToMulOpposite + map_add' _ _ := rfl + end NonUnitalNonAssocSemiring section NonUnitalSemiring diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean index 366fb8a79c8b3..7104d467ac071 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Defs.lean @@ -6,6 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.InjSurj import Mathlib.Algebra.Group.Submonoid.Defs +import Mathlib.Tactic.FastInstance /-! # Bundled non-unital subsemirings @@ -41,7 +42,8 @@ open AddSubmonoidClass `NonUnitalSubsemiringClass`. -/ /-- A non-unital subsemiring of a `NonUnitalNonAssocSemiring` inherits a `NonUnitalNonAssocSemiring` structure -/ -instance (priority := 75) toNonUnitalNonAssocSemiring : NonUnitalNonAssocSemiring s := +instance (priority := 75) toNonUnitalNonAssocSemiring : + NonUnitalNonAssocSemiring s := fast_instance% Subtype.coe_injective.nonUnitalNonAssocSemiring Subtype.val rfl (by simp) (fun _ _ => rfl) fun _ _ => rfl @@ -59,12 +61,12 @@ theorem coeSubtype : (subtype s : s → R) = ((↑) : s → R) := /-- A non-unital subsemiring of a `NonUnitalSemiring` is a `NonUnitalSemiring`. -/ instance toNonUnitalSemiring {R} [NonUnitalSemiring R] [SetLike S R] - [NonUnitalSubsemiringClass S R] : NonUnitalSemiring s := + [NonUnitalSubsemiringClass S R] : NonUnitalSemiring s := fast_instance% Subtype.coe_injective.nonUnitalSemiring Subtype.val rfl (by simp) (fun _ _ => rfl) fun _ _ => rfl /-- A non-unital subsemiring of a `NonUnitalCommSemiring` is a `NonUnitalCommSemiring`. -/ instance toNonUnitalCommSemiring {R} [NonUnitalCommSemiring R] [SetLike S R] - [NonUnitalSubsemiringClass S R] : NonUnitalCommSemiring s := + [NonUnitalSubsemiringClass S R] : NonUnitalCommSemiring s := fast_instance% Subtype.coe_injective.nonUnitalCommSemiring Subtype.val rfl (by simp) (fun _ _ => rfl) fun _ _ => rfl diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 6dc12681fc5bd..a6b3949cd74a8 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -199,7 +199,6 @@ end Perfection /-- A perfection map to a ring of characteristic `p` is a map that is isomorphic to its perfection. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. structure PerfectionMap (p : ℕ) [Fact p.Prime] {R : Type u₁} [CommSemiring R] [CharP R p] {P : Type u₂} [CommSemiring P] [CharP P p] [PerfectRing P p] (π : P →+* R) : Prop where injective : ∀ ⦃x y : P⦄, @@ -326,7 +325,7 @@ variable (p : ℕ) -- Porting note: Specified all arguments explicitly /-- `O/(p)` for `O`, ring of integers of `K`. -/ -@[nolint unusedArguments] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `nolint has_nonempty_instance` +@[nolint unusedArguments] def ModP (K : Type u₁) [Field K] (v : Valuation K ℝ≥0) (O : Type u₂) [CommRing O] [Algebra O K] (_ : v.Integers O) (p : ℕ) := O ⧸ (Ideal.span {(p : O)} : Ideal O) @@ -438,7 +437,6 @@ end Classical end ModP /-- Perfection of `O/(p)` where `O` is the ring of integers of `K`. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. def PreTilt := Ring.Perfection (ModP K v O hv p) p @@ -454,10 +452,9 @@ instance : CharP (PreTilt K v O hv p) p := section Classical -open scoped Classical - open Perfection +open scoped Classical in /-- The valuation `Perfection(O/(p)) → ℝ≥0` as a function. Given `f ∈ Perfection(O/(p))`, if `f = 0` then output `0`; otherwise output `preVal(f(n))^(p^n)` for any `n` such that `f(n) ≠ 0`. -/ @@ -468,6 +465,7 @@ noncomputable def valAux (f : PreTilt K v O hv p) : ℝ≥0 := variable {K v O hv p} +open scoped Classical in theorem coeff_nat_find_add_ne_zero {f : PreTilt K v O hv p} {h : ∃ n, coeff _ _ n f ≠ 0} (k : ℕ) : coeff _ _ (Nat.find h + k) f ≠ 0 := coeff_add_ne_zero (Nat.find_spec h) k @@ -476,6 +474,7 @@ theorem valAux_eq {f : PreTilt K v O hv p} {n : ℕ} (hfn : coeff _ _ n f ≠ 0) valAux K v O hv p f = ModP.preVal K v O hv p (coeff _ _ n f) ^ p ^ n := by have h : ∃ n, coeff _ _ n f ≠ 0 := ⟨n, hfn⟩ rw [valAux, dif_pos h] + classical obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le (Nat.find_min' h hfn) induction' k with k ih · rfl @@ -574,7 +573,6 @@ end PreTilt /-- The tilt of a field, as defined in Perfectoid Spaces by Peter Scholze, as in [scholze2011perfectoid]. Given a field `K` with valuation `K → ℝ≥0` and ring of integers `O`, this is implemented as the fraction field of the perfection of `O/(p)`. -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. def Tilt [Fact p.Prime] [Fact (v p ≠ 1)] := FractionRing (PreTilt K v O hv p) diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index e6caeb702f390..682fe26b642a9 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -723,12 +723,7 @@ theorem isPrime_map_C_iff_isPrime (P : Ideal R) : rw [Finset.mem_erase, Finset.mem_antidiagonal] at hij simp only [Ne, Prod.mk.inj_iff, not_and_or] at hij obtain hi | hj : i < m ∨ j < n := by - rw [or_iff_not_imp_left, not_lt, le_iff_lt_or_eq] - rintro (hmi | rfl) - · rw [← not_le] - intro hnj - exact (add_lt_add_of_lt_of_le hmi hnj).ne hij.2.symm - · simp only [eq_self_iff_true, not_true, false_or, add_right_inj, not_and_self_iff] at hij + omega · rw [mul_comm] apply P.mul_mem_left exact Classical.not_not.1 (Nat.find_min hf hi) @@ -950,7 +945,7 @@ theorem linearIndependent_powers_iff_aeval (f : M →ₗ[R] M) (v : M) : support, coeff, ofFinsupp_eq_zero] exact Iff.rfl -theorem disjoint_ker_aeval_of_coprime (f : M →ₗ[R] M) {p q : R[X]} (hpq : IsCoprime p q) : +theorem disjoint_ker_aeval_of_isCoprime (f : M →ₗ[R] M) {p q : R[X]} (hpq : IsCoprime p q) : Disjoint (LinearMap.ker (aeval f p)) (LinearMap.ker (aeval f q)) := by rw [disjoint_iff_inf_le] intro v hv @@ -959,7 +954,10 @@ theorem disjoint_ker_aeval_of_coprime (f : M →ₗ[R] M) {p q : R[X]} (hpq : Is LinearMap.mem_ker.1 (Submodule.mem_inf.1 hv).2] using congr_arg (fun p : R[X] => aeval f p v) hpq'.symm -theorem sup_aeval_range_eq_top_of_coprime (f : M →ₗ[R] M) {p q : R[X]} (hpq : IsCoprime p q) : +@[deprecated (since := "2025-01-23")] +alias disjoint_ker_aeval_of_coprime := disjoint_ker_aeval_of_isCoprime + +theorem sup_aeval_range_eq_top_of_isCoprime (f : M →ₗ[R] M) {p q : R[X]} (hpq : IsCoprime p q) : LinearMap.range (aeval f p) ⊔ LinearMap.range (aeval f q) = ⊤ := by rw [eq_top_iff] intro v _ @@ -972,6 +970,9 @@ theorem sup_aeval_range_eq_top_of_coprime (f : M →ₗ[R] M) {p q : R[X]} (hpq simpa only [mul_comm p p', mul_comm q q', aeval_one, aeval_add] using congr_arg (fun p : R[X] => aeval f p v) hpq' +@[deprecated (since := "2025-01-23")] +alias sup_aeval_range_eq_top_of_coprime := sup_aeval_range_eq_top_of_isCoprime + theorem sup_ker_aeval_le_ker_aeval_mul {f : M →ₗ[R] M} {p q : R[X]} : LinearMap.ker (aeval f p) ⊔ LinearMap.ker (aeval f q) ≤ LinearMap.ker (aeval f (p * q)) := by intro v hv diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index fc58485528703..c256a2901825b 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -54,8 +54,6 @@ This is a structure, not a class, since the same algebra can have many power bas For the common case where `S` is defined by adjoining an integral element to `R`, the canonical power basis is given by `{Algebra,IntermediateField}.adjoin.powerBasis`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. --- @[nolint has_nonempty_instance] structure PowerBasis (R S : Type*) [CommRing R] [Ring S] [Algebra R S] where gen : S dim : ℕ diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 783e6348c6aa9..0b795197df1f1 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -93,8 +93,7 @@ theorem mk_one_pow_eq_mk_choose_add : rw [pow_add, hd, pow_one, mul_comm, coeff_mul] simp_rw [coeff_mk, Pi.one_apply, one_mul] norm_cast - rw [Finset.sum_antidiagonal_choose_add, ← Nat.choose_succ_succ, Nat.succ_eq_add_one, - add_right_comm] + rw [Finset.sum_antidiagonal_choose_add, add_right_comm] /-- Given a natural number `d : ℕ` and a commutative ring `S`, `PowerSeries.invOneSubPow S d` is the diff --git a/Mathlib/RingTheory/Presentation.lean b/Mathlib/RingTheory/Presentation.lean index 398f0c3b2efbc..598bfa0dbf673 100644 --- a/Mathlib/RingTheory/Presentation.lean +++ b/Mathlib/RingTheory/Presentation.lean @@ -126,6 +126,28 @@ lemma finitePresentation_of_isFinite [P.IsFinite] : FinitePresentation R S := FinitePresentation.equiv (P.quotientEquiv.restrictScalars R) +variable (R S) in +/-- An arbitrary choice of a finite presentation of a finitely presented algebra. -/ +noncomputable +def ofFinitePresentation [FinitePresentation R S] : Presentation.{0, 0} R S := + letI H := FinitePresentation.out (R := R) (A := S) + letI n : ℕ := H.choose + letI f : MvPolynomial (Fin n) R →ₐ[R] S := H.choose_spec.choose + haveI hf : Function.Surjective f := H.choose_spec.choose_spec.1 + haveI hf' : (RingHom.ker f).FG := H.choose_spec.choose_spec.2 + letI H' := Submodule.fg_iff_exists_fin_generating_family.mp hf' + let m : ℕ := H'.choose + let v : Fin m → MvPolynomial (Fin n) R := H'.choose_spec.choose + let hv : Ideal.span (Set.range v) = RingHom.ker f := H'.choose_spec.choose_spec + { __ := Generators.ofSurjective (fun x ↦ f (.X x)) (by convert hf; ext; simp) + rels := Fin m + relation := v + span_range_relation_eq_ker := hv.trans (by congr; ext; simp) } + +instance [FinitePresentation R S] : (ofFinitePresentation R S).IsFinite where + finite_vars := Finite.of_fintype (Fin _) + finite_rels := Finite.of_fintype (Fin _) + section Construction /-- If `algebraMap R S` is bijective, the empty generators are a presentation with no relations. -/ diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 77a6f3b5770e5..f9de36718ac59 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -307,8 +307,7 @@ variable [CommRing R] [IsDomain R] [IsPrincipalIdealRing R] section -open scoped Classical - +open scoped Classical in /-- `factors a` is a multiset of irreducible elements whose product is `a`, up to units -/ noncomputable def factors (a : R) : Multiset R := if h : a = 0 then ∅ else Classical.choose (WfDvdMonoid.exists_factors a h) @@ -393,23 +392,31 @@ theorem isCoprime_of_dvd (x y : R) (nonzero : ¬(x = 0 ∧ y = 0)) (H : ∀ z ∈ nonunits R, z ≠ 0 → z ∣ x → ¬z ∣ y) : IsCoprime x y := (isRelPrime_of_no_nonunits_factors nonzero H).isCoprime -theorem dvd_or_coprime (x y : R) (h : Irreducible x) : x ∣ y ∨ IsCoprime x y := +theorem dvd_or_isCoprime (x y : R) (h : Irreducible x) : x ∣ y ∨ IsCoprime x y := h.dvd_or_isRelPrime.imp_right IsRelPrime.isCoprime +@[deprecated (since := "2025-01-23")] alias dvd_or_coprime := dvd_or_isCoprime + /-- See also `Irreducible.isRelPrime_iff_not_dvd`. -/ theorem Irreducible.coprime_iff_not_dvd {p n : R} (hp : Irreducible p) : IsCoprime p n ↔ ¬p ∣ n := by rw [← isRelPrime_iff_isCoprime, hp.isRelPrime_iff_not_dvd] /-- See also `Irreducible.coprime_iff_not_dvd'`. -/ -theorem Irreducible.dvd_iff_not_coprime {p n : R} (hp : Irreducible p) : p ∣ n ↔ ¬IsCoprime p n := +theorem Irreducible.dvd_iff_not_isCoprime {p n : R} (hp : Irreducible p) : p ∣ n ↔ ¬IsCoprime p n := iff_not_comm.2 hp.coprime_iff_not_dvd +@[deprecated (since := "2025-01-23")] +alias Irreducible.dvd_iff_not_coprime := Irreducible.dvd_iff_not_isCoprime + theorem Irreducible.coprime_pow_of_not_dvd {p a : R} (m : ℕ) (hp : Irreducible p) (h : ¬p ∣ a) : IsCoprime a (p ^ m) := (hp.coprime_iff_not_dvd.2 h).symm.pow_right -theorem Irreducible.coprime_or_dvd {p : R} (hp : Irreducible p) (i : R) : IsCoprime p i ∨ p ∣ i := - (_root_.em _).imp_right hp.dvd_iff_not_coprime.2 +theorem Irreducible.isCoprime_or_dvd {p : R} (hp : Irreducible p) (i : R) : IsCoprime p i ∨ p ∣ i := + (_root_.em _).imp_right hp.dvd_iff_not_isCoprime.2 + +@[deprecated (since := "2025-01-23")] +alias Irreducible.coprime_or_dvd := Irreducible.isCoprime_or_dvd variable [IsDomain R] diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index a213f7b39bfec..50e56b9aa02d5 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -47,14 +47,13 @@ This is different from the radical of an ideal. noncomputable section -open scoped Classical - namespace UniqueFactorizationMonoid -- `CancelCommMonoidWithZero` is required by `UniqueFactorizationMonoid` variable {M : Type*} [CancelCommMonoidWithZero M] [NormalizationMonoid M] [UniqueFactorizationMonoid M] +open scoped Classical in /-- The finite set of prime factors of an element in a unique factorization monoid. -/ def primeFactors (a : M) : Finset M := (normalizedFactors a).toFinset @@ -74,10 +73,12 @@ def radical (a : M) : M := @[simp] theorem radical_zero_eq : radical (0 : M) = 1 := by + classical rw [radical, primeFactors, normalizedFactors_zero, Multiset.toFinset_zero, Finset.prod_empty] @[simp] theorem radical_one_eq : radical (1 : M) = 1 := by + classical rw [radical, primeFactors, normalizedFactors_one, Multiset.toFinset_zero, Finset.prod_empty] theorem radical_eq_of_associated {a b : M} (h : Associated a b) : radical a = radical b := by @@ -100,6 +101,7 @@ theorem radical_pow (a : M) {n : Nat} (hn : 0 < n) : radical (a ^ n) = radical a simp_rw [radical, primeFactors_pow a hn] theorem radical_dvd_self (a : M) : radical a ∣ a := by + classical by_cases ha : a = 0 · rw [ha] apply dvd_zero @@ -147,13 +149,15 @@ theorem disjoint_normalizedFactors {a b : R} (hc : IsCoprime a b) : /-- Coprime elements have disjoint prime factors (as finsets). -/ theorem disjoint_primeFactors {a b : R} (hc : IsCoprime a b) : - Disjoint (primeFactors a) (primeFactors b) := - Multiset.disjoint_toFinset.mpr (disjoint_normalizedFactors hc) + Disjoint (primeFactors a) (primeFactors b) := by + classical + exact Multiset.disjoint_toFinset.mpr (disjoint_normalizedFactors hc) theorem mul_primeFactors_disjUnion {a b : R} (ha : a ≠ 0) (hb : b ≠ 0) (hc : IsCoprime a b) : primeFactors (a * b) = (primeFactors a).disjUnion (primeFactors b) (disjoint_primeFactors hc) := by + classical rw [Finset.disjUnion_eq_union] simp_rw [primeFactors] rw [normalizedFactors_mul ha hb, Multiset.toFinset_add] diff --git a/Mathlib/RingTheory/RingHomProperties.lean b/Mathlib/RingTheory/RingHomProperties.lean index c2ac311deb389..ae523c83039b3 100644 --- a/Mathlib/RingTheory/RingHomProperties.lean +++ b/Mathlib/RingTheory/RingHomProperties.lean @@ -48,14 +48,13 @@ theorem RespectsIso.cancel_left_isIso (hP : RespectsIso @P) {R S T : CommRingCat (g : S ⟶ T) [IsIso f] : P (g.hom.comp f.hom) ↔ P g.hom := ⟨fun H => by convert hP.2 (f ≫ g).hom (asIso f).symm.commRingCatIsoToRingEquiv H - exact (IsIso.inv_hom_id_assoc _ _).symm, hP.2 g.hom (asIso f).commRingCatIsoToRingEquiv⟩ + simp [← CommRingCat.hom_comp], hP.2 g.hom (asIso f).commRingCatIsoToRingEquiv⟩ theorem RespectsIso.cancel_right_isIso (hP : RespectsIso @P) {R S T : CommRingCat} (f : R ⟶ S) (g : S ⟶ T) [IsIso g] : P (g.hom.comp f.hom) ↔ P f.hom := ⟨fun H => by convert hP.1 (f ≫ g).hom (asIso g).symm.commRingCatIsoToRingEquiv H - change f = f ≫ g ≫ inv g - simp, hP.1 f.hom (asIso g).commRingCatIsoToRingEquiv⟩ + simp [← CommRingCat.hom_comp], hP.1 f.hom (asIso g).commRingCatIsoToRingEquiv⟩ theorem RespectsIso.is_localization_away_iff (hP : RingHom.RespectsIso @P) {R S : Type u} (R' S' : Type u) [CommRing R] [CommRing S] [CommRing R'] [CommRing S'] [Algebra R R'] diff --git a/Mathlib/RingTheory/RootsOfUnity/Complex.lean b/Mathlib/RingTheory/RootsOfUnity/Complex.lean index d2a6eb1145140..0b20c1181f990 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Complex.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Complex.lean @@ -130,17 +130,7 @@ theorem IsPrimitiveRoot.arg {n : ℕ} {ζ : ℂ} (h : IsPrimitiveRoot ζ n) (hn · exact hin · convert hin.add_mul_left_left (-1) using 1 rw [mul_neg_one, sub_eq_add_neg] - on_goal 2 => - split_ifs with h₂ - · exact mod_cast h - suffices (i - n : ℤ).natAbs = n - i by - rw [this] - apply tsub_lt_self hn.bot_lt - contrapose! h₂ - rw [Nat.eq_zero_of_le_zero h₂, zero_mul] - exact zero_le _ - rw [← Int.natAbs_neg, neg_sub, Int.natAbs_eq_iff] - exact Or.inl (Int.ofNat_sub h.le).symm + on_goal 2 => omega split_ifs with h₂ · convert Complex.arg_cos_add_sin_mul_I _ · push_cast; rfl diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index 2da9825d08c25..856a442cc406a 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -127,7 +127,7 @@ theorem minpoly_eq_pow {p : ℕ} [hprime : Fact p.Prime] (hdiv : ¬p ∣ n) : Polynomial.map_mul] refine IsCoprime.mul_dvd ?_ ?_ ?_ · have aux := IsPrimitive.Int.irreducible_iff_irreducible_map_cast Pmonic.isPrimitive - refine (dvd_or_coprime _ _ (aux.1 Pirr)).resolve_left ?_ + refine (dvd_or_isCoprime _ _ (aux.1 Pirr)).resolve_left ?_ rw [map_dvd_map (Int.castRingHom ℚ) Int.cast_injective Pmonic] intro hdiv refine hdiff (eq_of_monic_of_associated Pmonic Qmonic ?_) diff --git a/Mathlib/RingTheory/Smooth/Basic.lean b/Mathlib/RingTheory/Smooth/Basic.lean index 0d42a4a3c9af6..9ccb330ef6a71 100644 --- a/Mathlib/RingTheory/Smooth/Basic.lean +++ b/Mathlib/RingTheory/Smooth/Basic.lean @@ -41,11 +41,8 @@ variable (R : Type u) [CommSemiring R] variable (A : Type u) [Semiring A] [Algebra R A] /-- An `R` algebra `A` is formally smooth if for every `R`-algebra, every square-zero ideal -`I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists at least one lift `A →ₐ[R] B`. - -See . --/ -@[mk_iff] +`I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists at least one lift `A →ₐ[R] B`. -/ +@[mk_iff, stacks 00TI] class FormallySmooth : Prop where comp_surjective : ∀ ⦃B : Type u⦄ [CommRing B], @@ -343,12 +340,9 @@ section variable (R : Type u) [CommSemiring R] variable (A : Type u) [Semiring A] [Algebra R A] -/-- An `R` algebra `A` is smooth if it is formally smooth and of finite presentation. - -In the stacks project, the definition of smooth is completely different, and tag - proves that their definition is equivalent -to this. --/ +/-- An `R` algebra `A` is smooth if it is formally smooth and of finite presentation. -/ +@[stacks 00T2 "In the stacks project, the definition of smooth is completely different, and tag + proves that their definition is equivalent to this."] class Smooth [CommSemiring R] (A : Type u) [Semiring A] [Algebra R A] : Prop where formallySmooth : FormallySmooth R A := by infer_instance finitePresentation : FinitePresentation R A := by infer_instance diff --git a/Mathlib/RingTheory/Smooth/Kaehler.lean b/Mathlib/RingTheory/Smooth/Kaehler.lean index 44d8ef22e6985..e0f903639dfed 100644 --- a/Mathlib/RingTheory/Smooth/Kaehler.lean +++ b/Mathlib/RingTheory/Smooth/Kaehler.lean @@ -561,6 +561,24 @@ def H1Cotangent.equivOfFormallySmooth (P₁ P₂ : Extension R S) H1Cotangent.equiv (Extension.homInfinitesimal _ _) (Extension.homInfinitesimal _ _) ≪≫ₗ .symm (.ofBijective _ (H1Cotangent.map_toInfinitesimal_bijective P₂)) +lemma H1Cotangent.equivOfFormallySmooth_toLinearMap {P₁ P₂ : Extension R S} (f : P₁.Hom P₂) + [FormallySmooth R P₁.Ring] [FormallySmooth R P₂.Ring] : + (H1Cotangent.equivOfFormallySmooth P₁ P₂).toLinearMap = map f := by + ext1 x + refine (LinearEquiv.symm_apply_eq _).mpr ?_ + show ((map (P₁.homInfinitesimal P₂)).restrictScalars S ∘ₗ map P₁.toInfinitesimal) x = + ((map P₂.toInfinitesimal).restrictScalars S ∘ₗ map f) x + rw [← map_comp, ← map_comp, map_eq] + +lemma H1Cotangent.equivOfFormallySmooth_apply {P₁ P₂ : Extension R S} (f : P₁.Hom P₂) + [FormallySmooth R P₁.Ring] [FormallySmooth R P₂.Ring] (x) : + H1Cotangent.equivOfFormallySmooth P₁ P₂ x = map f x := by + rw [← equivOfFormallySmooth_toLinearMap]; rfl + +lemma H1Cotangent.equivOfFormallySmooth_symm (P₁ P₂ : Extension R S) + [FormallySmooth R P₁.Ring] [FormallySmooth R P₂.Ring] : + (equivOfFormallySmooth P₁ P₂).symm = equivOfFormallySmooth P₂ P₁ := rfl + /-- Any formally smooth extension can be used to calculate `H¹(L_{S/R})`. -/ noncomputable def equivH1CotangentOfFormallySmooth (P : Extension R S) [FormallySmooth R P.Ring] : diff --git a/Mathlib/RingTheory/Smooth/Locus.lean b/Mathlib/RingTheory/Smooth/Locus.lean new file mode 100644 index 0000000000000..7e4b969ef0b7e --- /dev/null +++ b/Mathlib/RingTheory/Smooth/Locus.lean @@ -0,0 +1,142 @@ +/- +Copyright (c) 2025 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Etale.Kaehler +import Mathlib.RingTheory.Spectrum.Prime.FreeLocus +import Mathlib.RingTheory.Support + +/-! +# Smooth locus of an algebra + +Most results in this file are proved for algebras of finite presentations. +Some of them are true for arbitrary algebras but the proof is substantially harder. + +## Main results +- `Algebra.smoothLocus` : The set of primes that are smooth over the base. +- `Algebra.basicOpen_subset_smoothLocus_iff` : + `D(f)` is contained in the smooth locus if and only if `A_f` is smooth over `R`. +- `Algebra.smoothLocus_eq_univ_iff` : + The smooth locus is the whole spectrum if and only if `A` is smooth over `R`. +- `Algebra.isOpen_smoothLocus` : The smooth locus is open. +-/ + +universe u + +variable (R A : Type u) [CommRing R] [CommRing A] [Algebra R A] + +namespace Algebra + +variable {A} in +/-- +An `R`-algebra `A` is smooth at a prime `p` of `A` if `Aₚ` is formally smooth over `R`. + +This does not imply `Aₚ` is smooth over `R` under the mathlib definition +even if `A` is finitely presented, +but it can be shown that this is equivalent to the stacks project definition that `A` is smooth +at `p` if and only if there exists `f ∉ p` such that `A_f` is smooth over `R`. +See `Algebra.basicOpen_subset_smoothLocus_iff_smooth` and `Algebra.isOpen_smoothLocus`. +-/ +@[stacks 00TB] +abbrev IsSmoothAt (p : Ideal A) [p.IsPrime] : Prop := + Algebra.FormallySmooth R (Localization.AtPrime p) + +/-- `Algebra.smoothLocus R A` is the set of primes `p` of `A` +such that `Aₚ` is formally smooth over `R`. -/ +def smoothLocus : Set (PrimeSpectrum A) := { p | IsSmoothAt R p.asIdeal } + +variable {R A} + +attribute [local instance] Module.finitePresentation_of_projective in +lemma smoothLocus_eq_compl_support_inter [EssFiniteType R A] : + smoothLocus R A = (Module.support A (H1Cotangent R A))ᶜ ∩ Module.freeLocus A (Ω[A⁄R]) := by + ext p + simp only [Set.mem_inter_iff, Set.mem_compl_iff, Module.not_mem_support_iff, + Module.mem_freeLocus] + refine Algebra.FormallySmooth.iff_subsingleton_and_projective.trans ?_ + congr! 1 + · have := IsLocalizedModule.iso p.asIdeal.primeCompl + (H1Cotangent.map R R A (Localization.AtPrime p.asIdeal)) + exact this.subsingleton_congr.symm + · trans Module.Free (Localization.AtPrime p.asIdeal) (Ω[Localization.AtPrime p.asIdeal⁄R]) + · have : EssFiniteType A (Localization.AtPrime p.asIdeal) := + .of_isLocalization _ p.asIdeal.primeCompl + have : EssFiniteType R (Localization.AtPrime p.asIdeal) := .comp _ A _ + exact ⟨fun _ ↦ Module.free_of_flat_of_isLocalRing, fun _ ↦ inferInstance⟩ + · have := IsLocalizedModule.iso p.asIdeal.primeCompl + (KaehlerDifferential.map R R A (Localization.AtPrime p.asIdeal)) + have := LinearEquiv.ofBijective (this.extendScalarsOfIsLocalization + p.asIdeal.primeCompl (Localization.AtPrime p.asIdeal)) this.bijective + exact ⟨fun H ↦ H.of_equiv' this.symm, fun H ↦ H.of_equiv' this⟩ + +lemma basicOpen_subset_smoothLocus_iff [FinitePresentation R A] {f : A} : + ↑(PrimeSpectrum.basicOpen f) ⊆ smoothLocus R A ↔ + Algebra.FormallySmooth R (Localization.Away f) := by + rw [smoothLocus_eq_compl_support_inter, Set.subset_inter_iff, Set.subset_compl_comm, + PrimeSpectrum.basicOpen_eq_zeroLocus_compl, compl_compl, + ← LocalizedModule.subsingleton_iff_support_subset, + Algebra.FormallySmooth.iff_subsingleton_and_projective] + congr! 1 + · have := IsLocalizedModule.iso (.powers f) (H1Cotangent.map R R A (Localization.Away f)) + rw [this.subsingleton_congr] + · rw [← PrimeSpectrum.basicOpen_eq_zeroLocus_compl, Module.basicOpen_subset_freeLocus_iff] + have := IsLocalizedModule.iso (.powers f) + (KaehlerDifferential.map R R A (Localization.Away f)) + have := LinearEquiv.ofBijective (this.extendScalarsOfIsLocalization + (.powers f) (Localization.Away f)) this.bijective + exact ⟨fun _ ↦ .of_equiv this, fun _ ↦ .of_equiv this.symm⟩ + +lemma basicOpen_subset_smoothLocus_iff_smooth [FinitePresentation R A] {f : A} : + ↑(PrimeSpectrum.basicOpen f) ⊆ smoothLocus R A ↔ + Algebra.Smooth R (Localization.Away f) := by + have : FinitePresentation A (Localization.Away f) := IsLocalization.Away.finitePresentation f + rw [basicOpen_subset_smoothLocus_iff] + exact ⟨fun H ↦ ⟨H, .trans _ A _⟩, fun H ↦ H.1⟩ + +lemma smoothLocus_eq_univ_iff [FinitePresentation R A] : + smoothLocus R A = Set.univ ↔ Algebra.FormallySmooth R A := by + have := IsLocalization.atUnits A (.powers 1) (S := Localization.Away (1 : A)) (by simp) + rw [Algebra.FormallySmooth.iff_of_equiv (this.restrictScalars R), + ← basicOpen_subset_smoothLocus_iff] + simp + +lemma smoothLocus_comap_of_isLocalization {Af : Type u} [CommRing Af] [Algebra A Af] [Algebra R Af] + [IsScalarTower R A Af] (f : A) [IsLocalization.Away f Af] : + PrimeSpectrum.comap (algebraMap A Af) ⁻¹' smoothLocus R A = smoothLocus R Af := by + ext p + let q := PrimeSpectrum.comap (algebraMap A Af) p + have : IsLocalization.AtPrime (Localization.AtPrime p.asIdeal) q.asIdeal := + IsLocalization.isLocalization_isLocalization_atPrime_isLocalization (.powers f) _ p.asIdeal + refine Algebra.FormallySmooth.iff_of_equiv ?_ + exact (IsLocalization.algEquiv q.asIdeal.primeCompl _ _).restrictScalars R + +-- Note that this does not follow directly from `smoothLocus_eq_compl_support_inter` because +-- `H¹(L_{S/R})` is not necessarily finitely generated. +open PrimeSpectrum in +lemma isOpen_smoothLocus [FinitePresentation R A] : IsOpen (smoothLocus R A) := by + rw [isOpen_iff_forall_mem_open] + intro x hx + obtain ⟨_, ⟨_, ⟨f, rfl⟩, rfl⟩, hxf, hf⟩ := + isBasis_basic_opens.exists_subset_of_mem_open + (smoothLocus_eq_compl_support_inter.le hx).2 Module.isOpen_freeLocus + rw [Module.basicOpen_subset_freeLocus_iff] at hf + let Af := Localization.Away f + have : Algebra.FinitePresentation A (Localization.Away f) := + IsLocalization.Away.finitePresentation f + have : Algebra.FinitePresentation R (Localization.Away f) := + .trans _ A _ + have : IsOpen (smoothLocus R Af) := by + have := IsLocalizedModule.iso (.powers f) + (KaehlerDifferential.map R R A (Localization.Away f)) + have := LinearEquiv.ofBijective (this.extendScalarsOfIsLocalization + (.powers f) (Localization.Away f)) this.bijective + have := Module.Projective.of_equiv this + rw [smoothLocus_eq_compl_support_inter, Module.support_eq_zeroLocus] + exact (isClosed_zeroLocus _).isOpen_compl.inter Module.isOpen_freeLocus + rw [← smoothLocus_comap_of_isLocalization f] at this + replace this := (PrimeSpectrum.localization_away_isOpenEmbedding Af f).isOpenMap _ this + rw [Set.image_preimage_eq_inter_range, localization_away_comap_range Af f] at this + exact ⟨_, Set.inter_subset_left, this, hx, hxf⟩ + +end Algebra diff --git a/Mathlib/RingTheory/Spectrum/Maximal/Basic.lean b/Mathlib/RingTheory/Spectrum/Maximal/Basic.lean index 608db0b4cbfb6..dc91ef72f56f7 100644 --- a/Mathlib/RingTheory/Spectrum/Maximal/Basic.lean +++ b/Mathlib/RingTheory/Spectrum/Maximal/Basic.lean @@ -3,9 +3,8 @@ 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.Localization.AsSubring import Mathlib.RingTheory.Spectrum.Maximal.Defs -import Mathlib.RingTheory.Spectrum.Prime.Basic +import Mathlib.RingTheory.Spectrum.Prime.Defs /-! # Maximal spectrum of a commutative (semi)ring @@ -15,12 +14,23 @@ Basic properties the maximal spectrum of a ring. noncomputable section -open scoped Classical - variable (R S P : Type*) [CommSemiring R] [CommSemiring S] [CommSemiring P] namespace MaximalSpectrum +/-- The prime spectrum is in bijection with the set of prime ideals. -/ +@[simps] +def equivSubtype : MaximalSpectrum R ≃ {I : Ideal R // I.IsMaximal} where + toFun I := ⟨I.asIdeal, I.2⟩ + invFun I := ⟨I, I.2⟩ + left_inv _ := rfl + right_inv _ := rfl + +theorem range_asIdeal : Set.range MaximalSpectrum.asIdeal = {J : Ideal R | J.IsMaximal} := + Set.ext fun J ↦ + ⟨fun hJ ↦ let ⟨j, hj⟩ := Set.mem_range.mp hJ; Set.mem_setOf.mpr <| hj ▸ j.isMaximal, + fun hJ ↦ Set.mem_range.mpr ⟨⟨J, Set.mem_setOf.mp hJ⟩, rfl⟩⟩ + variable {R} instance [Nontrivial R] : Nonempty <| MaximalSpectrum R := @@ -34,245 +44,4 @@ def toPrimeSpectrum (x : MaximalSpectrum R) : PrimeSpectrum R := theorem toPrimeSpectrum_injective : (@toPrimeSpectrum R _).Injective := fun ⟨_, _⟩ ⟨_, _⟩ h => by simpa only [MaximalSpectrum.mk.injEq] using PrimeSpectrum.ext_iff.mp h -open PrimeSpectrum Set - -variable (R : Type*) -variable [CommRing R] [IsDomain R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] - -/-- An integral domain is equal to the intersection of its localizations at all its maximal ideals -viewed as subalgebras of its field of fractions. -/ -theorem iInf_localization_eq_bot : (⨅ v : MaximalSpectrum R, - Localization.subalgebra.ofField K _ v.asIdeal.primeCompl_le_nonZeroDivisors) = ⊥ := by - ext x - rw [Algebra.mem_bot, Algebra.mem_iInf] - constructor - · contrapose - intro hrange hlocal - 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⟩ - 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]⟩ - -end MaximalSpectrum - -namespace PrimeSpectrum - -variable (R : Type*) -variable [CommRing R] [IsDomain R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] - -/-- An integral domain is equal to the intersection of its localizations at all its prime ideals -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 - 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 - -namespace MaximalSpectrum - -/-- The product of localizations at all maximal ideals of a commutative semiring. -/ -abbrev PiLocalization : Type _ := Π I : MaximalSpectrum R, Localization.AtPrime I.1 - -/-- The canonical ring homomorphism from a commutative semiring to the product of its -localizations at all maximal ideals. It is always injective. -/ -def toPiLocalization : R →+* PiLocalization R := algebraMap R _ - -theorem toPiLocalization_injective : Function.Injective (toPiLocalization R) := fun r r' eq ↦ by - rw [← one_mul r, ← one_mul r'] - by_contra ne - have ⟨I, mI, hI⟩ := (Module.eqIdeal R r r').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) - have ⟨s, hs⟩ := (IsLocalization.eq_iff_exists I.primeCompl _).mp (congr_fun eq ⟨I, mI⟩) - exact s.2 (hI hs) - -theorem toPiLocalization_apply_apply {r I} : toPiLocalization R r I = algebraMap R _ r := rfl - -variable {R S} (f : R →+* S) (g : S →+* P) (hf : Function.Bijective f) (hg : Function.Bijective g) - -/-- Functoriality of `PiLocalization` but restricted to bijective ring homs. -If R and S are commutative rings, surjectivity would be enough. -/ -noncomputable def mapPiLocalization : PiLocalization R →+* PiLocalization S := - Pi.ringHom fun I ↦ (Localization.localRingHom _ _ f rfl).comp <| - Pi.evalRingHom _ (⟨_, I.2.comap_bijective f hf⟩ : MaximalSpectrum R) - -theorem mapPiLocalization_naturality : - (mapPiLocalization f hf).comp (toPiLocalization R) = - (toPiLocalization S).comp f := by - ext r I - show Localization.localRingHom _ _ _ rfl (algebraMap _ _ r) = algebraMap _ _ (f r) - simp_rw [← IsLocalization.mk'_one (M := (I.1.comap f).primeCompl), Localization.localRingHom_mk', - ← IsLocalization.mk'_one (M := I.1.primeCompl), Submonoid.coe_one, map_one f] - rfl - -theorem mapPiLocalization_id : mapPiLocalization (.id R) Function.bijective_id = .id _ := - RingHom.ext fun _ ↦ funext fun _ ↦ congr($(Localization.localRingHom_id _) _) - -theorem mapPiLocalization_comp : - mapPiLocalization (g.comp f) (hg.comp hf) = - (mapPiLocalization g hg).comp (mapPiLocalization f hf) := - RingHom.ext fun _ ↦ funext fun _ ↦ congr($(Localization.localRingHom_comp _ _ _ _ rfl _ rfl) _) - -theorem mapPiLocalization_bijective : Function.Bijective (mapPiLocalization f hf) := by - let f := RingEquiv.ofBijective f hf - let e := RingEquiv.ofRingHom (mapPiLocalization f hf) - (mapPiLocalization (f.symm : S →+* R) f.symm.bijective) ?_ ?_ - · exact e.bijective - · rw [← mapPiLocalization_comp] - simp_rw [RingEquiv.comp_symm, mapPiLocalization_id] - · rw [← mapPiLocalization_comp] - simp_rw [RingEquiv.symm_comp, mapPiLocalization_id] - -section Pi - -variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial (R i)] - -theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : - ¬ Function.Surjective (toPiLocalization (Π i, R i)) := fun surj ↦ by - have ⟨J, max, nmem⟩ := PrimeSpectrum.exists_maximal_nmem_range_sigmaToPi_of_infinite R - obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max⟩ 1) - have : r = 0 := funext fun i ↦ toPiLocalization_injective _ <| funext fun I ↦ by - replace hr := congr_fun hr ⟨_, I.2.comap_piEvalRingHom⟩ - dsimp only [toPiLocalization_apply_apply, Subtype.coe_mk] at hr - simp_rw [toPiLocalization_apply_apply, - ← Localization.AtPrime.mapPiEvalRingHom_algebraMap_apply, hr] - rw [Function.update_of_ne]; · simp_rw [Pi.zero_apply, map_zero] - exact fun h ↦ nmem ⟨⟨i, I.1, I.2.isPrime⟩, PrimeSpectrum.ext congr($h.1)⟩ - replace hr := congr_fun hr ⟨J, max⟩ - rw [this, map_zero, Function.update_self] at hr - exact zero_ne_one hr - -variable {R} - -theorem finite_of_toPiLocalization_pi_surjective - (h : Function.Surjective (toPiLocalization (Π i, R i))) : - Finite ι := by - contrapose h; rw [not_finite_iff_infinite] at h - exact toPiLocalization_not_surjective_of_infinite _ - -end Pi - -theorem finite_of_toPiLocalization_surjective - (surj : Function.Surjective (toPiLocalization R)) : - Finite (MaximalSpectrum R) := by - replace surj := mapPiLocalization_bijective _ ⟨toPiLocalization_injective R, surj⟩ - |>.2.comp surj - rw [← RingHom.coe_comp, mapPiLocalization_naturality, RingHom.coe_comp] at surj - exact finite_of_toPiLocalization_pi_surjective surj.of_comp - end MaximalSpectrum - -namespace PrimeSpectrum - -/-- The product of localizations at all prime ideals of a commutative semiring. -/ -abbrev PiLocalization : Type _ := Π p : PrimeSpectrum R, Localization p.asIdeal.primeCompl - -/-- The canonical ring homomorphism from a commutative semiring to the product of its -localizations at all prime ideals. It is always injective. -/ -def toPiLocalization : R →+* PiLocalization R := algebraMap R _ - -theorem toPiLocalization_injective : Function.Injective (toPiLocalization R) := - fun _ _ eq ↦ MaximalSpectrum.toPiLocalization_injective R <| - funext fun I ↦ congr_fun eq I.toPrimeSpectrum - -/-- The projection from the product of localizations at primes to the product of -localizations at maximal ideals. -/ -def piLocalizationToMaximal : PiLocalization R →+* MaximalSpectrum.PiLocalization R := - Pi.ringHom fun I ↦ Pi.evalRingHom _ I.toPrimeSpectrum - -theorem piLocalizationToMaximal_surjective : Function.Surjective (piLocalizationToMaximal R) := - fun r ↦ ⟨fun I ↦ if h : I.1.IsMaximal then r ⟨_, h⟩ else 0, funext fun _ ↦ dif_pos _⟩ - -variable {R} - -/-- If R has Krull dimension ≤ 0, then `piLocalizationToIsMaximal R` is an isomorphism. -/ -def piLocalizationToMaximalEquiv (h : ∀ I : Ideal R, I.IsPrime → I.IsMaximal) : - PiLocalization R ≃+* MaximalSpectrum.PiLocalization R where - __ := piLocalizationToMaximal R - invFun := Pi.ringHom fun I ↦ Pi.evalRingHom _ (⟨_, h _ I.2⟩ : MaximalSpectrum R) - left_inv _ := rfl - right_inv _ := rfl - -theorem piLocalizationToMaximal_bijective (h : ∀ I : Ideal R, I.IsPrime → I.IsMaximal) : - Function.Bijective (piLocalizationToMaximal R) := - (piLocalizationToMaximalEquiv h).bijective - -theorem piLocalizationToMaximal_comp_toPiLocalization : - (piLocalizationToMaximal R).comp (toPiLocalization R) = MaximalSpectrum.toPiLocalization R := - rfl - -variable {S} - -theorem isMaximal_of_toPiLocalization_surjective (surj : Function.Surjective (toPiLocalization R)) - (I : PrimeSpectrum R) : I.1.IsMaximal := by - have ⟨J, max, le⟩ := I.1.exists_le_maximal I.2.ne_top - obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max.isPrime⟩ 1) - by_contra h - have hJ : algebraMap _ _ r = _ := (congr_fun hr _).trans (Function.update_self ..) - have hI : algebraMap _ _ r = _ := congr_fun hr I - rw [← IsLocalization.lift_eq (M := J.primeCompl) (S := Localization J.primeCompl), hJ, map_one, - Function.update_of_ne] at hI - · exact one_ne_zero hI - · intro eq; have : I.1 = J := congr_arg (·.1) eq; exact h (this ▸ max) - · exact fun ⟨s, hs⟩ ↦ IsLocalization.map_units (M := I.1.primeCompl) _ ⟨s, fun h ↦ hs (le h)⟩ - -variable (f : R →+* S) - -/-- A ring homomorphism induces a homomorphism between the products of localizations at primes. -/ -noncomputable def mapPiLocalization : PiLocalization R →+* PiLocalization S := - Pi.ringHom fun I ↦ (Localization.localRingHom _ I.1 f rfl).comp (Pi.evalRingHom _ (f.specComap I)) - -theorem mapPiLocalization_naturality : - (mapPiLocalization f).comp (toPiLocalization R) = (toPiLocalization S).comp f := by - ext r I - show Localization.localRingHom _ _ _ rfl (algebraMap _ _ r) = algebraMap _ _ (f r) - simp_rw [← IsLocalization.mk'_one (M := (I.1.comap f).primeCompl), Localization.localRingHom_mk', - ← IsLocalization.mk'_one (M := I.1.primeCompl), Submonoid.coe_one, map_one f] - rfl - -theorem mapPiLocalization_id : mapPiLocalization (.id R) = .id _ := by - ext; exact congr($(Localization.localRingHom_id _) _) - -theorem mapPiLocalization_comp (g : S →+* P) : - mapPiLocalization (g.comp f) = (mapPiLocalization g).comp (mapPiLocalization f) := by - ext; exact congr($(Localization.localRingHom_comp _ _ _ _ rfl _ rfl) _) - -theorem mapPiLocalization_bijective (hf : Function.Bijective f) : - Function.Bijective (mapPiLocalization f) := by - let f := RingEquiv.ofBijective f hf - let e := RingEquiv.ofRingHom (mapPiLocalization (f : R →+* S)) (mapPiLocalization f.symm) ?_ ?_ - · exact e.bijective - · rw [← mapPiLocalization_comp, RingEquiv.comp_symm, mapPiLocalization_id] - · rw [← mapPiLocalization_comp, RingEquiv.symm_comp, mapPiLocalization_id] - -section Pi - -variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial (R i)] - -theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : - ¬ Function.Surjective (toPiLocalization (Π i, R i)) := - fun surj ↦ MaximalSpectrum.toPiLocalization_not_surjective_of_infinite R <| by - rw [← piLocalizationToMaximal_comp_toPiLocalization] - exact (piLocalizationToMaximal_surjective _).comp surj - -variable {R} - -theorem finite_of_toPiLocalization_pi_surjective - (h : Function.Surjective (toPiLocalization (Π i, R i))) : - Finite ι := by - contrapose h; rw [not_finite_iff_infinite] at h - exact toPiLocalization_not_surjective_of_infinite _ - -end Pi - -theorem finite_of_toPiLocalization_surjective - (surj : Function.Surjective (toPiLocalization R)) : - Finite (PrimeSpectrum R) := by - replace surj := (mapPiLocalization_bijective _ ⟨toPiLocalization_injective R, surj⟩).2.comp surj - rw [← RingHom.coe_comp, mapPiLocalization_naturality, RingHom.coe_comp] at surj - exact finite_of_toPiLocalization_pi_surjective surj.of_comp - -end PrimeSpectrum diff --git a/Mathlib/RingTheory/Spectrum/Maximal/Defs.lean b/Mathlib/RingTheory/Spectrum/Maximal/Defs.lean index ab8109fe0112f..0a798a9410e85 100644 --- a/Mathlib/RingTheory/Spectrum/Maximal/Defs.lean +++ b/Mathlib/RingTheory/Spectrum/Maximal/Defs.lean @@ -26,4 +26,4 @@ structure MaximalSpectrum (R : Type*) [CommSemiring R] where @[deprecated (since := "2025-01-16")] alias MaximalSpectrum.IsMaximal := MaximalSpectrum.isMaximal -attribute [instance] MaximalSpectrum.IsMaximal +attribute [instance] MaximalSpectrum.isMaximal diff --git a/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean b/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean new file mode 100644 index 0000000000000..d0723e00ec9b5 --- /dev/null +++ b/Mathlib/RingTheory/Spectrum/Maximal/Localization.lean @@ -0,0 +1,267 @@ +/- +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.Localization.AsSubring +import Mathlib.RingTheory.Spectrum.Maximal.Basic +import Mathlib.RingTheory.Spectrum.Prime.Basic + +/-! +# Maximal spectrum of a commutative (semi)ring + +Localization results. +-/ + +noncomputable section + +open scoped Classical + +variable (R S P : Type*) [CommSemiring R] [CommSemiring S] [CommSemiring P] + +namespace MaximalSpectrum + +variable {R} + +open PrimeSpectrum Set + +variable (R : Type*) +variable [CommRing R] [IsDomain R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] + +/-- An integral domain is equal to the intersection of its localizations at all its maximal ideals +viewed as subalgebras of its field of fractions. -/ +theorem iInf_localization_eq_bot : (⨅ v : MaximalSpectrum R, + Localization.subalgebra.ofField K _ v.asIdeal.primeCompl_le_nonZeroDivisors) = ⊥ := by + ext x + rw [Algebra.mem_bot, Algebra.mem_iInf] + constructor + · contrapose + intro hrange hlocal + 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⟩ + 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]⟩ + +end MaximalSpectrum + +namespace PrimeSpectrum + +variable (R : Type*) +variable [CommRing R] [IsDomain R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] + +/-- An integral domain is equal to the intersection of its localizations at all its prime ideals +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 + 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 + +namespace MaximalSpectrum + +/-- The product of localizations at all maximal ideals of a commutative semiring. -/ +abbrev PiLocalization : Type _ := Π I : MaximalSpectrum R, Localization.AtPrime I.1 + +/-- The canonical ring homomorphism from a commutative semiring to the product of its +localizations at all maximal ideals. It is always injective. -/ +def toPiLocalization : R →+* PiLocalization R := algebraMap R _ + +theorem toPiLocalization_injective : Function.Injective (toPiLocalization R) := fun r r' eq ↦ by + rw [← one_mul r, ← one_mul r'] + by_contra ne + have ⟨I, mI, hI⟩ := (Module.eqIdeal R r r').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) + have ⟨s, hs⟩ := (IsLocalization.eq_iff_exists I.primeCompl _).mp (congr_fun eq ⟨I, mI⟩) + exact s.2 (hI hs) + +theorem toPiLocalization_apply_apply {r I} : toPiLocalization R r I = algebraMap R _ r := rfl + +variable {R S} (f : R →+* S) (g : S →+* P) (hf : Function.Bijective f) (hg : Function.Bijective g) + +/-- Functoriality of `PiLocalization` but restricted to bijective ring homs. +If R and S are commutative rings, surjectivity would be enough. -/ +noncomputable def mapPiLocalization : PiLocalization R →+* PiLocalization S := + Pi.ringHom fun I ↦ (Localization.localRingHom _ _ f rfl).comp <| + Pi.evalRingHom _ (⟨_, I.2.comap_bijective f hf⟩ : MaximalSpectrum R) + +theorem mapPiLocalization_naturality : + (mapPiLocalization f hf).comp (toPiLocalization R) = + (toPiLocalization S).comp f := by + ext r I + show Localization.localRingHom _ _ _ rfl (algebraMap _ _ r) = algebraMap _ _ (f r) + simp_rw [← IsLocalization.mk'_one (M := (I.1.comap f).primeCompl), Localization.localRingHom_mk', + ← IsLocalization.mk'_one (M := I.1.primeCompl), Submonoid.coe_one, map_one f] + rfl + +theorem mapPiLocalization_id : mapPiLocalization (.id R) Function.bijective_id = .id _ := + RingHom.ext fun _ ↦ funext fun _ ↦ congr($(Localization.localRingHom_id _) _) + +theorem mapPiLocalization_comp : + mapPiLocalization (g.comp f) (hg.comp hf) = + (mapPiLocalization g hg).comp (mapPiLocalization f hf) := + RingHom.ext fun _ ↦ funext fun _ ↦ congr($(Localization.localRingHom_comp _ _ _ _ rfl _ rfl) _) + +theorem mapPiLocalization_bijective : Function.Bijective (mapPiLocalization f hf) := by + let f := RingEquiv.ofBijective f hf + let e := RingEquiv.ofRingHom (mapPiLocalization f hf) + (mapPiLocalization (f.symm : S →+* R) f.symm.bijective) ?_ ?_ + · exact e.bijective + · rw [← mapPiLocalization_comp] + simp_rw [RingEquiv.comp_symm, mapPiLocalization_id] + · rw [← mapPiLocalization_comp] + simp_rw [RingEquiv.symm_comp, mapPiLocalization_id] + +section Pi + +variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial (R i)] + +theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : + ¬ Function.Surjective (toPiLocalization (Π i, R i)) := fun surj ↦ by + have ⟨J, max, nmem⟩ := PrimeSpectrum.exists_maximal_nmem_range_sigmaToPi_of_infinite R + obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max⟩ 1) + have : r = 0 := funext fun i ↦ toPiLocalization_injective _ <| funext fun I ↦ by + replace hr := congr_fun hr ⟨_, I.2.comap_piEvalRingHom⟩ + dsimp only [toPiLocalization_apply_apply, Subtype.coe_mk] at hr + simp_rw [toPiLocalization_apply_apply, + ← Localization.AtPrime.mapPiEvalRingHom_algebraMap_apply, hr] + rw [Function.update_of_ne]; · simp_rw [Pi.zero_apply, map_zero] + exact fun h ↦ nmem ⟨⟨i, I.1, I.2.isPrime⟩, PrimeSpectrum.ext congr($h.1)⟩ + replace hr := congr_fun hr ⟨J, max⟩ + rw [this, map_zero, Function.update_self] at hr + exact zero_ne_one hr + +variable {R} + +theorem finite_of_toPiLocalization_pi_surjective + (h : Function.Surjective (toPiLocalization (Π i, R i))) : + Finite ι := by + contrapose h; rw [not_finite_iff_infinite] at h + exact toPiLocalization_not_surjective_of_infinite _ + +end Pi + +theorem finite_of_toPiLocalization_surjective + (surj : Function.Surjective (toPiLocalization R)) : + Finite (MaximalSpectrum R) := by + replace surj := mapPiLocalization_bijective _ ⟨toPiLocalization_injective R, surj⟩ + |>.2.comp surj + rw [← RingHom.coe_comp, mapPiLocalization_naturality, RingHom.coe_comp] at surj + exact finite_of_toPiLocalization_pi_surjective surj.of_comp + +end MaximalSpectrum + +namespace PrimeSpectrum + +/-- The product of localizations at all prime ideals of a commutative semiring. -/ +abbrev PiLocalization : Type _ := Π p : PrimeSpectrum R, Localization p.asIdeal.primeCompl + +/-- The canonical ring homomorphism from a commutative semiring to the product of its +localizations at all prime ideals. It is always injective. -/ +def toPiLocalization : R →+* PiLocalization R := algebraMap R _ + +theorem toPiLocalization_injective : Function.Injective (toPiLocalization R) := + fun _ _ eq ↦ MaximalSpectrum.toPiLocalization_injective R <| + funext fun I ↦ congr_fun eq I.toPrimeSpectrum + +/-- The projection from the product of localizations at primes to the product of +localizations at maximal ideals. -/ +def piLocalizationToMaximal : PiLocalization R →+* MaximalSpectrum.PiLocalization R := + Pi.ringHom fun I ↦ Pi.evalRingHom _ I.toPrimeSpectrum + +theorem piLocalizationToMaximal_surjective : Function.Surjective (piLocalizationToMaximal R) := + fun r ↦ ⟨fun I ↦ if h : I.1.IsMaximal then r ⟨_, h⟩ else 0, funext fun _ ↦ dif_pos _⟩ + +variable {R} + +/-- If R has Krull dimension ≤ 0, then `piLocalizationToIsMaximal R` is an isomorphism. -/ +def piLocalizationToMaximalEquiv (h : ∀ I : Ideal R, I.IsPrime → I.IsMaximal) : + PiLocalization R ≃+* MaximalSpectrum.PiLocalization R where + __ := piLocalizationToMaximal R + invFun := Pi.ringHom fun I ↦ Pi.evalRingHom _ (⟨_, h _ I.2⟩ : MaximalSpectrum R) + left_inv _ := rfl + right_inv _ := rfl + +theorem piLocalizationToMaximal_bijective (h : ∀ I : Ideal R, I.IsPrime → I.IsMaximal) : + Function.Bijective (piLocalizationToMaximal R) := + (piLocalizationToMaximalEquiv h).bijective + +theorem piLocalizationToMaximal_comp_toPiLocalization : + (piLocalizationToMaximal R).comp (toPiLocalization R) = MaximalSpectrum.toPiLocalization R := + rfl + +variable {S} + +theorem isMaximal_of_toPiLocalization_surjective (surj : Function.Surjective (toPiLocalization R)) + (I : PrimeSpectrum R) : I.1.IsMaximal := by + have ⟨J, max, le⟩ := I.1.exists_le_maximal I.2.ne_top + obtain ⟨r, hr⟩ := surj (Function.update 0 ⟨J, max.isPrime⟩ 1) + by_contra h + have hJ : algebraMap _ _ r = _ := (congr_fun hr _).trans (Function.update_self ..) + have hI : algebraMap _ _ r = _ := congr_fun hr I + rw [← IsLocalization.lift_eq (M := J.primeCompl) (S := Localization J.primeCompl), hJ, map_one, + Function.update_of_ne] at hI + · exact one_ne_zero hI + · intro eq; have : I.1 = J := congr_arg (·.1) eq; exact h (this ▸ max) + · exact fun ⟨s, hs⟩ ↦ IsLocalization.map_units (M := I.1.primeCompl) _ ⟨s, fun h ↦ hs (le h)⟩ + +variable (f : R →+* S) + +/-- A ring homomorphism induces a homomorphism between the products of localizations at primes. -/ +noncomputable def mapPiLocalization : PiLocalization R →+* PiLocalization S := + Pi.ringHom fun I ↦ (Localization.localRingHom _ I.1 f rfl).comp (Pi.evalRingHom _ (f.specComap I)) + +theorem mapPiLocalization_naturality : + (mapPiLocalization f).comp (toPiLocalization R) = (toPiLocalization S).comp f := by + ext r I + show Localization.localRingHom _ _ _ rfl (algebraMap _ _ r) = algebraMap _ _ (f r) + simp_rw [← IsLocalization.mk'_one (M := (I.1.comap f).primeCompl), Localization.localRingHom_mk', + ← IsLocalization.mk'_one (M := I.1.primeCompl), Submonoid.coe_one, map_one f] + rfl + +theorem mapPiLocalization_id : mapPiLocalization (.id R) = .id _ := by + ext; exact congr($(Localization.localRingHom_id _) _) + +theorem mapPiLocalization_comp (g : S →+* P) : + mapPiLocalization (g.comp f) = (mapPiLocalization g).comp (mapPiLocalization f) := by + ext; exact congr($(Localization.localRingHom_comp _ _ _ _ rfl _ rfl) _) + +theorem mapPiLocalization_bijective (hf : Function.Bijective f) : + Function.Bijective (mapPiLocalization f) := by + let f := RingEquiv.ofBijective f hf + let e := RingEquiv.ofRingHom (mapPiLocalization (f : R →+* S)) (mapPiLocalization f.symm) ?_ ?_ + · exact e.bijective + · rw [← mapPiLocalization_comp, RingEquiv.comp_symm, mapPiLocalization_id] + · rw [← mapPiLocalization_comp, RingEquiv.symm_comp, mapPiLocalization_id] + +section Pi + +variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] [∀ i, Nontrivial (R i)] + +theorem toPiLocalization_not_surjective_of_infinite [Infinite ι] : + ¬ Function.Surjective (toPiLocalization (Π i, R i)) := + fun surj ↦ MaximalSpectrum.toPiLocalization_not_surjective_of_infinite R <| by + rw [← piLocalizationToMaximal_comp_toPiLocalization] + exact (piLocalizationToMaximal_surjective _).comp surj + +variable {R} + +theorem finite_of_toPiLocalization_pi_surjective + (h : Function.Surjective (toPiLocalization (Π i, R i))) : + Finite ι := by + contrapose h; rw [not_finite_iff_infinite] at h + exact toPiLocalization_not_surjective_of_infinite _ + +end Pi + +theorem finite_of_toPiLocalization_surjective + (surj : Function.Surjective (toPiLocalization R)) : + Finite (PrimeSpectrum R) := by + replace surj := (mapPiLocalization_bijective _ ⟨toPiLocalization_injective R, surj⟩).2.comp surj + rw [← RingHom.coe_comp, mapPiLocalization_naturality, RingHom.coe_comp] at surj + exact finite_of_toPiLocalization_pi_surjective surj.of_comp + +end PrimeSpectrum diff --git a/Mathlib/RingTheory/Spectrum/Prime/Basic.lean b/Mathlib/RingTheory/Spectrum/Prime/Basic.lean index eb955307cdc44..92132833e5f2f 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Basic.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Basic.lean @@ -5,7 +5,6 @@ Authors: Johan Commelin, Filippo A. E. Nuccio, Andrew Yang -/ import Mathlib.LinearAlgebra.Finsupp.SumProd import Mathlib.RingTheory.Ideal.Prod -import Mathlib.RingTheory.Localization.Ideal import Mathlib.RingTheory.Nilpotent.Lemmas import Mathlib.RingTheory.Noetherian.Basic import Mathlib.RingTheory.Spectrum.Prime.Defs @@ -43,7 +42,7 @@ assert_not_exists TopologicalSpace noncomputable section -open scoped Classical Pointwise +open scoped Pointwise universe u v @@ -74,6 +73,11 @@ def equivSubtype : PrimeSpectrum R ≃ {I : Ideal R // I.IsPrime} where left_inv _ := rfl right_inv _ := rfl +theorem range_asIdeal : Set.range PrimeSpectrum.asIdeal = {J : Ideal R | J.IsPrime} := + Set.ext fun J ↦ + ⟨fun hJ ↦ let ⟨j, hj⟩ := Set.mem_range.mp hJ; Set.mem_setOf.mpr <| hj ▸ j.isPrime, + fun hJ ↦ Set.mem_range.mpr ⟨⟨J, Set.mem_setOf.mp hJ⟩, rfl⟩⟩ + /-- The map from the direct sum of prime spectra to the prime spectrum of a direct product. -/ @[simp] def primeSpectrumProdOfSum : PrimeSpectrum R ⊕ PrimeSpectrum S → PrimeSpectrum (R × S) @@ -199,6 +203,9 @@ theorem vanishingIdeal_zeroLocus_eq_radical (I : Ideal R) : rw [mem_vanishingIdeal, Ideal.radical_eq_sInf, Submodule.mem_sInf] exact ⟨fun h x hx => h ⟨x, hx.2⟩ hx.1, fun h x hx => h x.1 ⟨hx, x.2⟩⟩ +theorem nilradical_eq_iInf : nilradical R = iInf asIdeal := by + apply range_asIdeal R ▸ nilradical_eq_sInf R + @[simp] theorem zeroLocus_radical (I : Ideal R) : zeroLocus (I.radical : Set R) = zeroLocus I := vanishingIdeal_zeroLocus_eq_radical I ▸ (gc R).l_u_l_eq_l I @@ -549,29 +556,6 @@ def comapEquiv (e : R ≃+* S) : PrimeSpectrum R ≃ PrimeSpectrum S where RingEquiv.toRingHom_eq_coe, RingEquiv.comp_symm] rfl -variable (S) - -theorem localization_specComap_injective [Algebra R S] (M : Submonoid R) [IsLocalization M S] : - Function.Injective (algebraMap R S).specComap := by - intro p q h - replace h := _root_.congr_arg (fun x : PrimeSpectrum R => Ideal.map (algebraMap R S) x.asIdeal) h - dsimp only [specComap] at h - rw [IsLocalization.map_comap M S, IsLocalization.map_comap M S] at h - ext1 - exact h - -theorem localization_specComap_range [Algebra R S] (M : Submonoid R) [IsLocalization M S] : - Set.range (algebraMap R S).specComap = { p | Disjoint (M : Set R) p.asIdeal } := by - ext x - constructor - · simp_rw [disjoint_iff_inf_le] - rintro ⟨p, rfl⟩ x ⟨hx₁, hx₂⟩ - exact (p.2.1 : ¬_) (p.asIdeal.eq_top_of_isUnit_mem hx₂ (IsLocalization.map_units S ⟨x, hx₁⟩)) - · intro h - use ⟨x.asIdeal.map (algebraMap R S), IsLocalization.isPrime_of_isPrime_disjoint M S _ x.2 h⟩ - ext1 - exact IsLocalization.comap_map_of_isPrime_disjoint M S _ x.2 h - section Pi variable {ι} (R : ι → Type*) [∀ i, CommSemiring (R i)] @@ -584,6 +568,7 @@ and is a homeomorphism when ι is finite. -/ | ⟨i, p⟩ => (Pi.evalRingHom R i).specComap p theorem sigmaToPi_injective : (sigmaToPi R).Injective := fun ⟨i, p⟩ ⟨j, q⟩ eq ↦ by + classical obtain rfl | ne := eq_or_ne i j · congr; ext x simpa using congr_arg (Function.update (0 : ∀ i, R i) i x ∈ ·.asIdeal) eq @@ -599,6 +584,7 @@ range of `sigmaToPi`, i.e. is not of the form `πᵢ⁻¹(𝔭)` for some prime see https://math.stackexchange.com/a/1563190. -/ theorem exists_maximal_nmem_range_sigmaToPi_of_infinite : ∃ (I : Ideal (Π i, R i)) (_ : I.IsMaximal), ⟨I, inferInstance⟩ ∉ Set.range (sigmaToPi R) := by + classical let J : Ideal (Π i, R i) := -- `J := Π₀ i, R i` is an ideal in `Π i, R i` { __ := AddMonoidHom.mrange DFinsupp.coeFnAddMonoidHom smul_mem' := by diff --git a/Mathlib/RingTheory/Spectrum/Prime/ConstructibleSet.lean b/Mathlib/RingTheory/Spectrum/Prime/ConstructibleSet.lean new file mode 100644 index 0000000000000..47a1ca362df0a --- /dev/null +++ b/Mathlib/RingTheory/Spectrum/Prime/ConstructibleSet.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2024 Yaël Dillies, Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Andrew Yang +-/ +import Mathlib.Order.SuccPred.WithBot +import Mathlib.RingTheory.Spectrum.Prime.Topology + +/-! +# Constructible sets in the prime spectrum + +This file provides tooling for manipulating constructible sets in the prime spectrum of a ring. + +## TODO + +Link to constructible sets in a topological space. +-/ + +open Finset +open scoped Polynomial + +namespace PrimeSpectrum +variable {R S T : Type*} [CommSemiring R] [CommSemiring S] [CommSemiring T] + +variable (R) in +/-- The data of a constructible set `s` is finitely many tuples `(f, g₁, ..., gₙ)` such that +`s = ⋃ (f, g₁, ..., gₙ), V(g₁, ..., gₙ) \ V(f)`. + +To obtain `s` from its data, use `PrimeSpectrum.ConstructibleSetData.toSet`. -/ +abbrev ConstructibleSetData := Finset (R × Σ n, Fin n → R) + +namespace ConstructibleSetData + +/-- Given the data of the constructible set `s`, build the data of the constructible set +`{I | {x | f x ∈ I} ∈ s}`. -/ +def map [DecidableEq S] (f : R →+* S) (s : ConstructibleSetData R) : ConstructibleSetData S := + s.image fun ⟨r, n, g⟩ ↦ ⟨f r, n, f ∘ g⟩ + +@[simp] +lemma map_id [DecidableEq R] (s : ConstructibleSetData R) : s.map (.id _) = s := by simp [map] + +lemma map_comp [DecidableEq S] [DecidableEq T] (f : S →+* T) (g : R →+* S) + (s : ConstructibleSetData R) : s.map (f.comp g) = (s.map g).map f := by + simp [map, image_image, Function.comp_def] + +/-- Given the data of a constructible set `s`, namely finitely many tuples `(f, g₁, ..., gₙ)` such +that `s = ⋃ (f, g₁, ..., gₙ), V(g₁, ..., gₙ) \ V(f)`, return `s`. -/ +def toSet (S : ConstructibleSetData R) : Set (PrimeSpectrum R) := + ⋃ x ∈ S, zeroLocus (Set.range x.2.2) \ zeroLocus {x.1} + +@[simp] +lemma toSet_map [DecidableEq S] (f : R →+* S) (s : ConstructibleSetData R) : + (s.map f).toSet = comap f ⁻¹' s.toSet := by + unfold toSet map + rw [set_biUnion_finset_image] + simp [← Set.range_comp] + +/-- The degree bound on a constructible set for Chevalley's theorem for the inclusion `R ↪ R[X]`. -/ +def degBound (S : ConstructibleSetData R[X]) : ℕ := S.sup fun e ↦ ∑ i, (e.2.2 i).degree.succ + +end PrimeSpectrum.ConstructibleSetData diff --git a/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean b/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean index 4a6c542cd5693..2117b0ff26773 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/IsOpenComapC.lean @@ -59,11 +59,8 @@ theorem imageOfDf_eq_comap_C_compl_zeroLocus : · rintro ⟨xli, complement, rfl⟩ exact comap_C_mem_imageOfDf complement -/-- The morphism `C⁺ : Spec R[x] → Spec R` is open. -Stacks Project "Lemma 00FB", first part. - -https://stacks.math.columbia.edu/tag/00FB --/ +/-- The morphism `C⁺ : Spec R[x] → Spec R` is open. -/ +@[stacks 00FB "First part"] theorem isOpenMap_comap_C : IsOpenMap (PrimeSpectrum.comap (C : R →+* R[X])) := by rintro U ⟨s, z⟩ rw [← compl_compl U, ← z, ← iUnion_of_singleton_coe s, zeroLocus_iUnion, compl_iInter, diff --git a/Mathlib/RingTheory/Spectrum/Prime/Noetherian.lean b/Mathlib/RingTheory/Spectrum/Prime/Noetherian.lean index 4e38232fb5a13..e36ce5d245d25 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Noetherian.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Noetherian.lean @@ -43,10 +43,6 @@ section IsArtinianRing variable (R : Type u) [CommRing R] [IsArtinianRing R] -instance : Finite (PrimeSpectrum R) := - have : Finite {I : Ideal R // I.IsPrime} := IsArtinianRing.primeSpectrum_finite R - Finite.of_injective _ (equivSubtype R).injective - instance : DiscreteTopology (PrimeSpectrum R) := discreteTopology_iff_finite_and_isPrime_imp_isMaximal.mpr ⟨inferInstance, fun _ _ ↦ inferInstance⟩ diff --git a/Mathlib/RingTheory/Spectrum/Prime/Polynomial.lean b/Mathlib/RingTheory/Spectrum/Prime/Polynomial.lean index ef44d600253f3..7dfebf9a5be4b 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Polynomial.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Polynomial.lean @@ -161,6 +161,12 @@ lemma isOpenMap_comap_C : IsOpenMap (comap (R := R) C) := by simp only [image_comap_C_basicOpen] exact (isClosed_zeroLocus _).isOpen_compl +lemma comap_C_surjective : Function.Surjective (comap (R := R) C) := by + intro x + refine ⟨comap (evalRingHom 0) x, ?_⟩ + rw [← comap_comp_apply, (show (evalRingHom 0).comp C = .id R by ext; simp), + comap_id, ContinuousMap.id_apply] + lemma exists_image_comap_of_monic (f g : R[X]) (hg : g.Monic) : ∃ t : Finset R, comap C '' (zeroLocus {g} \ zeroLocus {f}) = (zeroLocus t)ᶜ := by apply (config := { allowSynthFailures := true }) exists_image_comap_of_finite_of_free @@ -218,4 +224,10 @@ lemma isOpenMap_comap_C : IsOpenMap (comap (R := R) (C (σ := σ))) := by simp only [image_comap_C_basicOpen] exact (isClosed_zeroLocus _).isOpen_compl +lemma comap_C_surjective : Function.Surjective (comap (R := R) (C (σ := σ))) := by + intro x + refine ⟨comap (eval₂Hom (.id _) 0) x, ?_⟩ + rw [← comap_comp_apply, (show (eval₂Hom (.id _) 0).comp C = .id R by ext; simp), + comap_id, ContinuousMap.id_apply] + end MvPolynomial diff --git a/Mathlib/RingTheory/Spectrum/Prime/Topology.lean b/Mathlib/RingTheory/Spectrum/Prime/Topology.lean index 36eabf32f2950..c846ca364238f 100644 --- a/Mathlib/RingTheory/Spectrum/Prime/Topology.lean +++ b/Mathlib/RingTheory/Spectrum/Prime/Topology.lean @@ -9,8 +9,10 @@ import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic +import Mathlib.RingTheory.Localization.Algebra import Mathlib.RingTheory.Localization.Away.Basic -import Mathlib.RingTheory.Spectrum.Maximal.Basic +import Mathlib.RingTheory.Localization.Ideal +import Mathlib.RingTheory.Spectrum.Maximal.Localization import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober @@ -40,7 +42,7 @@ variable (R : Type u) (S : Type v) namespace PrimeSpectrum -section CommSemiRing +section CommSemiring variable [CommSemiring R] [CommSemiring S] variable {R S} @@ -257,6 +259,10 @@ def comap (f : R →+* S) : C(PrimeSpectrum S, PrimeSpectrum R) where rintro _ ⟨s, rfl⟩ exact ⟨_, preimage_specComap_zeroLocus_aux f s⟩ +lemma coe_comap (f : R →+* S) : comap f = f.specComap := rfl + +lemma comap_apply (f : R →+* S) (x) : comap f x = f.specComap x := rfl + variable (f : R →+* S) @[simp] @@ -285,6 +291,27 @@ theorem comap_injective_of_surjective (f : R →+* S) (hf : Function.Surjective variable (S) +theorem localization_specComap_injective [Algebra R S] (M : Submonoid R) [IsLocalization M S] : + Function.Injective (algebraMap R S).specComap := by + intro p q h + replace h := _root_.congr_arg (fun x : PrimeSpectrum R => Ideal.map (algebraMap R S) x.asIdeal) h + dsimp only [RingHom.specComap] at h + rw [IsLocalization.map_comap M S, IsLocalization.map_comap M S] at h + ext1 + exact h + +theorem localization_specComap_range [Algebra R S] (M : Submonoid R) [IsLocalization M S] : + Set.range (algebraMap R S).specComap = { p | Disjoint (M : Set R) p.asIdeal } := by + ext x + constructor + · simp_rw [disjoint_iff_inf_le] + rintro ⟨p, rfl⟩ x ⟨hx₁, hx₂⟩ + exact (p.2.1 : ¬_) (p.asIdeal.eq_top_of_isUnit_mem hx₂ (IsLocalization.map_units S ⟨x, hx₁⟩)) + · intro h + use ⟨x.asIdeal.map (algebraMap R S), IsLocalization.isPrime_of_isPrime_disjoint M S _ x.2 h⟩ + ext1 + exact IsLocalization.comap_map_of_isPrime_disjoint M S _ x.2 h + theorem localization_comap_isInducing [Algebra R S] (M : Submonoid R) [IsLocalization M S] : IsInducing (comap (algebraMap R S)) := by refine ⟨TopologicalSpace.ext_isClosed fun Z ↦ ?_⟩ @@ -332,9 +359,8 @@ theorem comap_isInducing_of_surjective (hf : Surjective f) : IsInducing (comap f @[deprecated (since := "2024-10-28")] alias comap_inducing_of_surjective := comap_isInducing_of_surjective - end Comap -end CommSemiRing +end CommSemiring section SpecOfSurjective @@ -410,7 +436,7 @@ def primeSpectrumProdHomeo : end SpecProd -section CommSemiRing +section CommSemiring variable [CommSemiring R] [CommSemiring S] variable {R S} @@ -549,6 +575,20 @@ theorem isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton [Algebra R S] exact not_not.mpr (q.span_singleton_le_iff_mem.mp le) IsLocalization.isLocalization_iff_of_isLocalization _ _ (Localization.Away f) +open Localization Polynomial Set in +lemma range_comap_algebraMap_localization_compl_eq_range_comap_quotientMk + {R : Type*} [CommRing R] (c : R) : + letI := (mapRingHom (algebraMap R (Away c))).toAlgebra + (range (comap (algebraMap R[X] (Away c)[X])))ᶜ + = range (comap (mapRingHom (Ideal.Quotient.mk (.span {c})))) := by + letI := (mapRingHom (algebraMap R (Away c))).toAlgebra + have := Polynomial.isLocalization (.powers c) (Away c) + rw [Submonoid.map_powers] at this + have surj : Function.Surjective (mapRingHom (Ideal.Quotient.mk (.span {c}))) := + Polynomial.map_surjective _ Ideal.Quotient.mk_surjective + rw [range_comap_of_surjective _ _ surj, localization_away_comap_range _ (C c)] + simp [Polynomial.ker_mapRingHom, Ideal.map_span] + end BasicOpen section DiscreteTopology @@ -928,7 +968,7 @@ def primeSpectrum_unique_of_localization_at_minimal (h : I ∈ minimalPrimes R) end LocalizationAtMinimal -end CommSemiRing +end CommSemiring end PrimeSpectrum @@ -936,11 +976,9 @@ section CommSemiring variable [CommSemiring R] open PrimeSpectrum in -/-- -[Stacks: Lemma 00ES (3)](https://stacks.math.columbia.edu/tag/00ES) -Zero loci of minimal prime ideals of `R` are irreducible components in `Spec R` and any -irreducible component is a zero locus of some minimal prime ideal. --/ +/-- Zero loci of minimal prime ideals of `R` are irreducible components in `Spec R` and any +irreducible component is a zero locus of some minimal prime ideal. -/ +@[stacks 00ES] protected def minimalPrimes.equivIrreducibleComponents : minimalPrimes R ≃o (irreducibleComponents <| PrimeSpectrum R)ᵒᵈ := by let e : {p : Ideal R | p.IsPrime ∧ ⊥ ≤ p} ≃o PrimeSpectrum R := diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index 06ad42f8fd9ac..23244caec6a64 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -883,6 +883,14 @@ theorem comm_symm : (TensorProduct.comm R A B).symm = TensorProduct.comm R B A := by ext; rfl +@[simp] +lemma comm_comp_includeLeft : + (TensorProduct.comm R A B : A ⊗[R] B →ₐ[R] B ⊗[R] A).comp includeLeft = includeRight := rfl + +@[simp] +lemma comm_comp_includeRight : + (TensorProduct.comm R A B : A ⊗[R] B →ₐ[R] B ⊗[R] A).comp includeRight = includeLeft := rfl + theorem adjoin_tmul_eq_top : adjoin R { t : A ⊗[R] B | ∃ a b, a ⊗ₜ[R] b = t } = ⊤ := top_le_iff.mp <| (top_le_iff.mpr <| span_tmul_eq_top R A B).trans (span_le_adjoin R _) @@ -983,6 +991,16 @@ theorem map_range (f : A →ₐ[R] B) (g : C →ₐ[R] D) : · rw [← map_comp_includeLeft f g, ← map_comp_includeRight f g] exact sup_le (AlgHom.range_comp_le_range _ _) (AlgHom.range_comp_le_range _ _) +lemma comm_comp_map (f : A →ₐ[R] C) (g : B →ₐ[R] D) : + (TensorProduct.comm R C D : C ⊗[R] D →ₐ[R] D ⊗[R] C).comp (Algebra.TensorProduct.map f g) = + (Algebra.TensorProduct.map g f).comp (TensorProduct.comm R A B).toAlgHom := by + ext <;> rfl + +lemma comm_comp_map_apply (f : A →ₐ[R] C) (g : B →ₐ[R] D) (x) : + TensorProduct.comm R C D (Algebra.TensorProduct.map f g x) = + (Algebra.TensorProduct.map g f) (TensorProduct.comm R A B x) := + congr($(comm_comp_map f g) x) + /-- Construct an isomorphism between tensor products of an S-algebra with an R-algebra from S- and R- isomorphisms between the tensor factors. -/ diff --git a/Mathlib/RingTheory/TensorProduct/Quotient.lean b/Mathlib/RingTheory/TensorProduct/Quotient.lean new file mode 100644 index 0000000000000..1995ac993264e --- /dev/null +++ b/Mathlib/RingTheory/TensorProduct/Quotient.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2025 Christian Merten, Yi Song, Sihan Su. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten, Yi Song, Sihan Su +-/ +import Mathlib.LinearAlgebra.TensorProduct.Quotient +import Mathlib.RingTheory.Ideal.Quotient.Operations +import Mathlib.RingTheory.TensorProduct.Basic + +/-! +# Interaction between quotients and tensor products for algebras + +This files proves algebra analogs of the isomorphisms in +`Mathlib.LinearAlgebra.TensorProduct.Quotient`. + +## Main results + +- `Algebra.TensorProduct.quotIdealMapEquivTensorQuot`: + `B ⧸ (I.map <| algebraMap A B) ≃ₐ[B] B ⊗[A] (A ⧸ I)` +-/ + +open TensorProduct + +namespace Algebra.TensorProduct + +variable {A : Type*} (B : Type*) [CommRing A] [CommRing B] [Algebra A B] (I : Ideal A) + +private noncomputable def quotIdealMapEquivTensorQuotAux : + (B ⧸ (I.map <| algebraMap A B)) ≃ₗ[B] B ⊗[A] (A ⧸ I) := + AddEquiv.toLinearEquiv (TensorProduct.tensorQuotEquivQuotSMul B I ≪≫ₗ + Submodule.quotEquivOfEq _ _ (Ideal.smul_top_eq_map I) ≪≫ₗ + Submodule.Quotient.restrictScalarsEquiv A (I.map <| algebraMap A B)).symm <| by + intro c x + obtain ⟨u, rfl⟩ := Ideal.Quotient.mk_surjective x + rfl + +private lemma quotIdealMapEquivTensorQuotAux_mk (b : B) : + (quotIdealMapEquivTensorQuotAux B I) b = b ⊗ₜ[A] 1 := + rfl + +/-- `B ⊗[A] (A ⧸ I)` is isomorphic as an `A`-algebra to `B ⧸ I B`. -/ +noncomputable def quotIdealMapEquivTensorQuot : + (B ⧸ (I.map <| algebraMap A B)) ≃ₐ[B] B ⊗[A] (A ⧸ I) := + AlgEquiv.ofLinearEquiv (quotIdealMapEquivTensorQuotAux B I) rfl + (fun x y ↦ by + obtain ⟨u, rfl⟩ := Ideal.Quotient.mk_surjective x + obtain ⟨v, rfl⟩ := Ideal.Quotient.mk_surjective y + simp_rw [← map_mul, quotIdealMapEquivTensorQuotAux_mk] + simp) + +@[simp] +lemma quotIdealMapEquivTensorQuot_mk (b : B) : + quotIdealMapEquivTensorQuot B I b = b ⊗ₜ[A] 1 := + rfl + +@[simp] +lemma quotIdealMapEquivTensorQuot_symm_tmul (b : B) (a : A) : + (quotIdealMapEquivTensorQuot B I).symm (b ⊗ₜ[A] a) = Submodule.Quotient.mk (a • b) := + rfl + +end Algebra.TensorProduct diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index dbe13c0d621bf..5b322e66a5a7b 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -119,11 +119,22 @@ open IntermediateField variable (K) theorem trace_eq_trace_adjoin [FiniteDimensional K L] (x : L) : - Algebra.trace K L x = finrank K⟮x⟯ L • trace K K⟮x⟯ (AdjoinSimple.gen K x) := by + trace K L x = finrank K⟮x⟯ L • trace K K⟮x⟯ (AdjoinSimple.gen K x) := by rw [← trace_trace (S := K⟮x⟯)] - conv in x => rw [← IntermediateField.AdjoinSimple.algebraMap_gen K x] + conv in x => rw [← AdjoinSimple.algebraMap_gen K x] rw [trace_algebraMap, LinearMap.map_smul_of_tower] +variable {K} in +/-- Trace of the generator of a simple adjoin equals negative of the next coefficient of +its minimal polynomial coefficient. -/ +theorem trace_adjoinSimpleGen {x : L} (hx : IsIntegral K x) : + trace K K⟮x⟯ (AdjoinSimple.gen K x) = -(minpoly K x).nextCoeff := by + simpa [minpoly_gen K x] using PowerBasis.trace_gen_eq_nextCoeff_minpoly <| adjoin.powerBasis hx + +theorem trace_eq_finrank_mul_minpoly_nextCoeff [FiniteDimensional K L] (x : L) : + trace K L x = finrank K⟮x⟯ L * -(minpoly K x).nextCoeff := by + rw [trace_eq_trace_adjoin, trace_adjoinSimpleGen (.of_finite K x), Algebra.smul_def]; rfl + variable {K} theorem trace_eq_sum_roots [FiniteDimensional K L] {x : L} diff --git a/Mathlib/RingTheory/Trace/Defs.lean b/Mathlib/RingTheory/Trace/Defs.lean index e351aa7d74b4a..6a3eecd5d778a 100644 --- a/Mathlib/RingTheory/Trace/Defs.lean +++ b/Mathlib/RingTheory/Trace/Defs.lean @@ -92,6 +92,13 @@ theorem trace_algebraMap_of_basis (b : Basis ι R S) (x : R) : convert Finset.sum_const x simp [-coe_lmul_eq_mul] + +/-- The trace map from `R` to itself is the identity map. -/ +@[simp] theorem trace_self : trace R R = LinearMap.id := by + ext; simpa using trace_algebraMap_of_basis (.singleton (Fin 1) R) 1 + +theorem trace_self_apply (a) : trace R R a = a := by simp + /-- If `x` is in the base field `K`, then the trace is `[L : K] * x`. (If `L` is not finite-dimensional over `K`, then `trace` and `finrank` return `0`.) diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index 9e91002b6692d..8d6c532010819 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -46,10 +46,8 @@ An `R`-algebra `A` is formally unramified if `Ω[A⁄R]` is trivial. This is equivalent to "for every `R`-algebra, every square-zero ideal `I : Ideal B` and `f : A →ₐ[R] B ⧸ I`, there exists at most one lift `A →ₐ[R] B`". -See `Algebra.FormallyUnramified.iff_comp_injective`. - -See . -/ -@[mk_iff] +See `Algebra.FormallyUnramified.iff_comp_injective`. -/ +@[mk_iff, stacks 00UM] class FormallyUnramified : Prop where subsingleton_kaehlerDifferential : Subsingleton (Ω[A⁄R]) @@ -284,12 +282,9 @@ section variable (R : Type*) [CommRing R] variable (A : Type*) [CommRing A] [Algebra R A] -/-- An `R`-algebra `A` is unramified if it is formally unramified and of finite type. - -Note that the Stacks project has a different definition of unramified, and tag - shows that their definition is the -same as this one. --/ +/-- An `R`-algebra `A` is unramified if it is formally unramified and of finite type. -/ +@[stacks 00UT "Note that the Stacks project has a different definition of unramified, and tag + shows that their definition is the same as this one."] class Unramified : Prop where formallyUnramified : FormallyUnramified R A := by infer_instance finiteType : FiniteType R A := by infer_instance diff --git a/Mathlib/RingTheory/Valuation/Archimedean.lean b/Mathlib/RingTheory/Valuation/Archimedean.lean index 2f6d1e71bf789..42b7d56d295ea 100644 --- a/Mathlib/RingTheory/Valuation/Archimedean.lean +++ b/Mathlib/RingTheory/Valuation/Archimedean.lean @@ -24,6 +24,8 @@ variable {F Γ₀ O : Type*} [Field F] [LinearOrderedCommGroupWithZero Γ₀] instance : LinearOrderedCommGroupWithZero (MonoidHom.mrange v) where __ : CommGroupWithZero (MonoidHom.mrange v) := inferInstance __ : LinearOrder (MonoidHom.mrange v) := inferInstance + bot := 0 + bot_le a := show (0 : Γ₀) ≤ _ from zero_le' zero_le_one := Subtype.coe_le_coe.mp zero_le_one mul_le_mul_left := by simp only [Subtype.forall, MonoidHom.mem_mrange, forall_exists_index, Submonoid.mk_mul_mk, diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index 5ff8726949782..2d1fe63a3a6d8 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Johan Commelin, Patrick Massot -/ +import Mathlib.Algebra.Order.Hom.Monoid import Mathlib.Algebra.Order.Ring.Basic import Mathlib.RingTheory.Ideal.Maps import Mathlib.Tactic.TFAE @@ -57,7 +58,6 @@ If ever someone extends `Valuation`, we should fully comply to the `DFunLike` by boilerplate lemmas to `ValuationClass`. -/ -open scoped Classical open Function Ideal noncomputable section @@ -68,7 +68,6 @@ section variable (F R) (Γ₀ : Type*) [LinearOrderedCommMonoidWithZero Γ₀] [Ring R] ---Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- The type of `Γ₀`-valued valuations on `R`. When you extend this structure, make sure to extend `ValuationClass`. -/ @@ -168,6 +167,7 @@ theorem map_add_lt {x y g} (hx : v x < g) (hy : v y < g) : v (x + y) < g := theorem map_sum_le {ι : Type*} {s : Finset ι} {f : ι → R} {g : Γ₀} (hf : ∀ i ∈ s, v (f i) ≤ g) : v (∑ i ∈ s, f i) ≤ g := by + classical refine Finset.induction_on s (fun _ => v.map_zero ▸ zero_le') (fun a s has ih hf => ?_) hf @@ -176,6 +176,7 @@ theorem map_sum_le {ι : Type*} {s : Finset ι} {f : ι → R} {g : Γ₀} (hf : theorem map_sum_lt {ι : Type*} {s : Finset ι} {f : ι → R} {g : Γ₀} (hg : g ≠ 0) (hf : ∀ i ∈ s, v (f i) < g) : v (∑ i ∈ s, f i) < g := by + classical refine Finset.induction_on s (fun _ => v.map_zero ▸ (zero_lt_iff.2 hg)) (fun a s has ih hf => ?_) hf @@ -252,6 +253,14 @@ lemma map_apply (f : Γ₀ →*₀ Γ'₀) (hf : Monotone f) (v : Valuation R Γ def IsEquiv (v₁ : Valuation R Γ₀) (v₂ : Valuation R Γ'₀) : Prop := ∀ r s, v₁ r ≤ v₁ s ↔ v₂ r ≤ v₂ s +/-- An ordered monoid isomorphism `Γ₀ ≃ Γ'₀` induces an equivalence +`Valuation R Γ₀ ≃ Valuation R Γ'₀`. -/ +def congr (f : Γ₀ ≃*o Γ'₀) : Valuation R Γ₀ ≃ Valuation R Γ'₀ where + toFun := map f f.toOrderIso.monotone + invFun := map f.symm f.toOrderIso.symm.monotone + left_inv ν := by ext; simp + right_inv ν := by ext; simp + end Monoid section Group @@ -305,6 +314,7 @@ theorem map_sub_eq_of_lt_right (h : v x < v y) : v (x - y) = v y := by rw [sub_eq_add_neg, map_add_eq_of_lt_right, map_neg] rwa [map_neg] +open scoped Classical in theorem map_sum_eq_of_lt {ι : Type*} {s : Finset ι} {f : ι → R} {j : ι} (hj : j ∈ s) (h0 : v (f j) ≠ 0) (hf : ∀ i ∈ s \ {j}, v (f i) < v (f j)) : v (∑ i ∈ s, f i) = v (f j) := by @@ -592,7 +602,6 @@ section AddMonoid variable (R) [Ring R] (Γ₀ : Type*) [LinearOrderedAddCommMonoidWithTop Γ₀] /-- The type of `Γ₀`-valued additive valuations on `R`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] def AddValuation := Valuation R (Multiplicative Γ₀ᵒᵈ) @@ -896,12 +905,24 @@ variable {K Γ₀ : Type*} [Ring R] [LinearOrderedCommMonoidWithZero Γ₀] /-- The `AddValuation` associated to a `Valuation`. -/ def toAddValuation : Valuation R Γ₀ ≃ AddValuation R (Additive Γ₀)ᵒᵈ := - AddValuation.ofValuation (R := R) (Γ₀ := (Additive Γ₀)ᵒᵈ) + .trans (congr + { toFun := fun x ↦ .ofAdd <| .toDual <| .toDual <| .ofMul x + invFun := fun x ↦ x.toAdd.ofDual.ofDual.toMul + left_inv := fun _x ↦ rfl + right_inv := fun _x ↦ rfl + map_mul' := fun _x _y ↦ rfl + map_le_map_iff' := .rfl }) (AddValuation.ofValuation (R := R) (Γ₀ := (Additive Γ₀)ᵒᵈ)) /-- The `Valuation` associated to a `AddValuation`. -/ def ofAddValuation : AddValuation R (Additive Γ₀)ᵒᵈ ≃ Valuation R Γ₀ := - AddValuation.toValuation + AddValuation.toValuation.trans <| congr <| + { toFun := fun x ↦ x.toAdd.ofDual.ofDual.toMul + invFun := fun x ↦ .ofAdd <| .toDual <| .toDual <| .ofMul x + left_inv := fun _x ↦ rfl + right_inv := fun _x ↦ rfl + map_mul' := fun _x _y ↦ rfl + map_le_map_iff' := .rfl } @[simp] lemma ofAddValuation_symm_eq : ofAddValuation.symm = toAddValuation (R := R) (Γ₀ := Γ₀) := rfl diff --git a/Mathlib/RingTheory/Valuation/ValuationRing.lean b/Mathlib/RingTheory/Valuation/ValuationRing.lean index d0ff48e43af4d..65e0b27c7dd14 100644 --- a/Mathlib/RingTheory/Valuation/ValuationRing.lean +++ b/Mathlib/RingTheory/Valuation/ValuationRing.lean @@ -155,34 +155,36 @@ noncomputable instance linearOrder : LinearOrder (ValueGroup A K) where decidableLE := by classical infer_instance noncomputable instance linearOrderedCommGroupWithZero : - LinearOrderedCommGroupWithZero (ValueGroup A K) := - { linearOrder .. with - mul_assoc := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; apply Quotient.sound'; rw [mul_assoc] - one_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [one_mul] - mul_one := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_one] - mul_comm := by rintro ⟨a⟩ ⟨b⟩; apply Quotient.sound'; rw [mul_comm] - mul_le_mul_left := by - rintro ⟨a⟩ ⟨b⟩ ⟨c, rfl⟩ ⟨d⟩ - use c; simp only [Algebra.smul_def]; ring - zero_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [zero_mul] - mul_zero := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_zero] - zero_le_one := ⟨0, by rw [zero_smul]⟩ - exists_pair_ne := by - use 0, 1 - intro c; obtain ⟨d, hd⟩ := Quotient.exact' c - apply_fun fun t => d⁻¹ • t at hd - simp only [inv_smul_smul, smul_zero, one_ne_zero] at hd - inv_zero := by apply Quotient.sound'; rw [inv_zero] - mul_inv_cancel := by - rintro ⟨a⟩ ha - apply Quotient.sound' - use 1 - simp only [one_smul, ne_eq] - apply (mul_inv_cancel₀ _).symm - contrapose ha - simp only [Classical.not_not] at ha ⊢ - rw [ha] - rfl } + LinearOrderedCommGroupWithZero (ValueGroup A K) where + __ := linearOrder .. + mul_assoc := by rintro ⟨a⟩ ⟨b⟩ ⟨c⟩; apply Quotient.sound'; rw [mul_assoc] + one_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [one_mul] + mul_one := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_one] + mul_comm := by rintro ⟨a⟩ ⟨b⟩; apply Quotient.sound'; rw [mul_comm] + mul_le_mul_left := by + rintro ⟨a⟩ ⟨b⟩ ⟨c, rfl⟩ ⟨d⟩ + use c; simp only [Algebra.smul_def]; ring + zero_mul := by rintro ⟨a⟩; apply Quotient.sound'; rw [zero_mul] + mul_zero := by rintro ⟨a⟩; apply Quotient.sound'; rw [mul_zero] + zero_le_one := ⟨0, by rw [zero_smul]⟩ + exists_pair_ne := by + use 0, 1 + intro c; obtain ⟨d, hd⟩ := Quotient.exact' c + apply_fun fun t => d⁻¹ • t at hd + simp only [inv_smul_smul, smul_zero, one_ne_zero] at hd + inv_zero := by apply Quotient.sound'; rw [inv_zero] + mul_inv_cancel := by + rintro ⟨a⟩ ha + apply Quotient.sound' + use 1 + simp only [one_smul, ne_eq] + apply (mul_inv_cancel₀ _).symm + contrapose ha + simp only [Classical.not_not] at ha ⊢ + rw [ha] + rfl + bot := 0 + bot_le := by rintro ⟨a⟩; exact ⟨0, zero_smul ..⟩ /-- Any valuation ring induces a valuation on its fraction field. -/ def valuation : Valuation K (ValueGroup A K) where diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index d1f3ec7693c78..bf3b19d8bb5aa 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -754,13 +754,13 @@ theorem mem_inv_pointwise_smul_iff {g : G} {S : ValuationSubring K} {x : K} : @[simp] theorem pointwise_smul_le_pointwise_smul_iff {g : G} {S T : ValuationSubring K} : - g • S ≤ g • T ↔ S ≤ T := Set.set_smul_subset_set_smul_iff + g • S ≤ g • T ↔ S ≤ T := Set.smul_set_subset_smul_set_iff theorem pointwise_smul_subset_iff {g : G} {S T : ValuationSubring K} : g • S ≤ T ↔ S ≤ g⁻¹ • T := - Set.set_smul_subset_iff + Set.smul_set_subset_iff_subset_inv_smul_set theorem subset_pointwise_smul_iff {g : G} {S T : ValuationSubring K} : S ≤ g • T ↔ g⁻¹ • S ≤ T := - Set.subset_set_smul_iff + Set.subset_smul_set_iff end PointwiseActions diff --git a/Mathlib/RingTheory/WittVector/Isocrystal.lean b/Mathlib/RingTheory/WittVector/Isocrystal.lean index 6686048e8f6fd..12a87d8bc7fe6 100644 --- a/Mathlib/RingTheory/WittVector/Isocrystal.lean +++ b/Mathlib/RingTheory/WittVector/Isocrystal.lean @@ -129,14 +129,12 @@ variable (V) /-- A homomorphism between isocrystals respects the Frobenius map. Notation `M →ᶠⁱ [p, k]` in the `Isocrystal` namespace. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. @[nolint has_nonempty_instance] structure IsocrystalHom extends V →ₗ[K(p, k)] V₂ where frob_equivariant : ∀ x : V, Φ(p, k) (toLinearMap x) = toLinearMap (Φ(p, k) x) /-- An isomorphism between isocrystals respects the Frobenius map. Notation `M ≃ᶠⁱ [p, k]` in the `Isocrystal` namespace. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. @[nolint has_nonempty_instance] structure IsocrystalEquiv extends V ≃ₗ[K(p, k)] V₂ where frob_equivariant : ∀ x : V, Φ(p, k) (toLinearEquiv x) = toLinearEquiv (Φ(p, k) x) @@ -156,7 +154,6 @@ open scoped Isocrystal of slope `m : ℤ`. -/ @[nolint unusedArguments] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. @[nolint has_nonempty_instance] def StandardOneDimIsocrystal (_m : ℤ) : Type _ := K(p, k) diff --git a/Mathlib/SetTheory/Cardinal/Arithmetic.lean b/Mathlib/SetTheory/Cardinal/Arithmetic.lean index b3bf57c0f1e5a..ab0045568dde8 100644 --- a/Mathlib/SetTheory/Cardinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Cardinal/Arithmetic.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ import Mathlib.SetTheory.Cardinal.Aleph import Mathlib.SetTheory.Ordinal.Principal -import Mathlib.Tactic.Linarith /-! # Cardinal arithmetic @@ -220,7 +219,7 @@ theorem mul_eq_left_iff {a b : Cardinal} : a * b = a ↔ max ℵ₀ b ≤ a ∧ rw [← not_lt] apply fun h2b => ne_of_gt _ h conv_rhs => left; rw [← mul_one n] - rw [mul_lt_mul_left] + rw [Nat.mul_lt_mul_left] · exact id apply Nat.lt_of_succ_le h2a · rintro (⟨⟨ha, hab⟩, hb⟩ | rfl | rfl) @@ -432,14 +431,6 @@ section aleph theorem aleph_add_aleph (o₁ o₂ : Ordinal) : ℵ_ o₁ + ℵ_ o₂ = ℵ_ (max o₁ o₂) := by rw [Cardinal.add_eq_max (aleph0_le_aleph o₁), aleph_max] -theorem principal_add_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Ordinal.Principal (· + ·) c.ord := - fun a b ha hb => by - rw [lt_ord, Ordinal.card_add] at * - exact add_lt_of_lt hc ha hb - -theorem principal_add_aleph (o : Ordinal) : Ordinal.Principal (· + ·) (ℵ_ o).ord := - principal_add_ord <| aleph0_le_aleph o - theorem add_right_inj_of_lt_aleph0 {α β γ : Cardinal} (γ₀ : γ < aleph0) : α + γ = β + γ ↔ α = β := ⟨fun h => Cardinal.eq_of_add_eq_add_right h γ₀, fun h => congr_arg (· + γ) h⟩ @@ -528,6 +519,9 @@ theorem power_nat_le_max {c : Cardinal.{u}} {n : ℕ} : c ^ (n : Cardinal.{u}) · exact le_max_of_le_left (power_nat_le hc) · exact le_max_of_le_right (power_lt_aleph0 hc (nat_lt_aleph0 _)).le +lemma power_le_aleph0 {a b : Cardinal.{u}} (ha : a ≤ ℵ₀) (hb : b < ℵ₀) : a ^ b ≤ ℵ₀ := by + lift b to ℕ using hb; simpa [ha] using power_nat_le_max (c := a) + theorem powerlt_aleph0 {c : Cardinal} (h : ℵ₀ ≤ c) : c ^< ℵ₀ = c := by apply le_antisymm · rw [powerlt_le] @@ -955,4 +949,38 @@ theorem principal_opow_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Principal (· ^ apply (isInitial_ord c).principal_opow rwa [omega0_le_ord] +/-! ### Initial ordinals are principal -/ + +theorem principal_add_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Principal (· + ·) c.ord := by + intro a b ha hb + rw [lt_ord, card_add] at * + exact add_lt_of_lt hc ha hb + +theorem IsInitial.principal_add {o : Ordinal} (h : IsInitial o) (ho : ω ≤ o) : + Principal (· + ·) o := by + rw [← h.ord_card] + apply principal_add_ord + rwa [aleph0_le_card] + +theorem principal_add_omega (o : Ordinal) : Principal (· + ·) (ω_ o) := + (isInitial_omega o).principal_add (omega0_le_omega o) + +theorem principal_mul_ord {c : Cardinal} (hc : ℵ₀ ≤ c) : Principal (· * ·) c.ord := by + intro a b ha hb + rw [lt_ord, card_mul] at * + exact mul_lt_of_lt hc ha hb + +theorem IsInitial.principal_mul {o : Ordinal} (h : IsInitial o) (ho : ω ≤ o) : + Principal (· * ·) o := by + rw [← h.ord_card] + apply principal_mul_ord + rwa [aleph0_le_card] + +theorem principal_mul_omega (o : Ordinal) : Principal (· * ·) (ω_ o) := + (isInitial_omega o).principal_mul (omega0_le_omega o) + +@[deprecated principal_add_omega (since := "2024-11-08")] +theorem _root_.Cardinal.principal_add_aleph (o : Ordinal) : Principal (· + ·) (ℵ_ o).ord := + principal_add_ord <| aleph0_le_aleph o + end Ordinal diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 7d7201ca375ba..809f5a5ecbdc4 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -922,6 +922,9 @@ theorem isRegular_cof {o : Ordinal} (h : o.IsLimit) : IsRegular o.cof := theorem isRegular_aleph0 : IsRegular ℵ₀ := ⟨le_rfl, by simp⟩ +lemma fact_isRegular_aleph0 : Fact Cardinal.aleph0.IsRegular where + out := Cardinal.isRegular_aleph0 + theorem isRegular_succ {c : Cardinal.{u}} (h : ℵ₀ ≤ c) : IsRegular (succ c) := ⟨h.trans (le_succ c), succ_le_of_lt diff --git a/Mathlib/SetTheory/Cardinal/HasCardinalLT.lean b/Mathlib/SetTheory/Cardinal/HasCardinalLT.lean index b14bed334c1f0..2f920d70d1b02 100644 --- a/Mathlib/SetTheory/Cardinal/HasCardinalLT.lean +++ b/Mathlib/SetTheory/Cardinal/HasCardinalLT.lean @@ -3,7 +3,7 @@ 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.SetTheory.Ordinal.Basic +import Mathlib.SetTheory.Cardinal.Arithmetic /-! # The property of being of cardinality less than a cardinal @@ -74,3 +74,15 @@ lemma hasCardinalLT_iff_of_equiv {X : Type u} {Y : Type u'} (e : X ≃ Y) (κ : lemma hasCardinalLT_aleph0_iff (X : Type u) : HasCardinalLT X Cardinal.aleph0.{v} ↔ Finite X := by simpa [HasCardinalLT] using Cardinal.mk_lt_aleph0_iff + +lemma hasCardinalLT_option_iff (X : Type u) (κ : Cardinal.{w}) + (hκ : Cardinal.aleph0 ≤ κ) : + HasCardinalLT (Option X) κ ↔ HasCardinalLT X κ := by + constructor + · intro h + exact h.of_injective _ (Option.some_injective _) + · intro h + dsimp [HasCardinalLT] at h ⊢ + simp only [Cardinal.mk_option, Cardinal.lift_add, Cardinal.lift_one] + exact Cardinal.add_lt_of_lt (by simpa using hκ) h + (lt_of_lt_of_le Cardinal.one_lt_aleph0 (by simpa using hκ)) diff --git a/Mathlib/SetTheory/Descriptive/Tree.lean b/Mathlib/SetTheory/Descriptive/Tree.lean new file mode 100644 index 0000000000000..43a7f2f810ce2 --- /dev/null +++ b/Mathlib/SetTheory/Descriptive/Tree.lean @@ -0,0 +1,46 @@ +/- +Copyright (c) 2024 Sven Manthe. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sven Manthe +-/ +import Mathlib.Order.CompleteLattice.SetLike + +/-! +# Trees in the sense of descriptive set theory + +This file defines trees of depth `ω` in the sense of descriptive set theory as sets of finite +sequences that are stable under taking prefixes. + +## Main declarations + +* `tree A`: a (possibly infinite) tree of depth at most `ω` with nodes in `A` +-/ + +namespace Descriptive + +/-- A tree is a set of finite sequences, implemented as `List A`, that is stable under + taking prefixes. For the definition we use the equivalent property `x ++ [a] ∈ T → x ∈ T`, + which is more convenient to check. We define `tree A` as a complete sublattice of + `Set (List A)`, which coerces to the type of trees on `A`. -/ +def tree (A : Type*) : CompleteSublattice (Set (List A)) := + CompleteSublattice.mk' {T | ∀ ⦃x : List A⦄ ⦃a : A⦄, x ++ [a] ∈ T → x ∈ T} + (by rintro S hS x a ⟨t, ht, hx⟩; use t, ht, hS ht hx) + (by rintro S hS x a h T hT; exact hS hT <| h T hT) + +instance (A : Type*) : SetLike (tree A) (List A) := SetLike.instSubtypeSet + +variable {A : Type*} {T : tree A} + +lemma mem_of_append {x y : List A} (h : x ++ y ∈ T) : x ∈ T := by + induction' y with y ys ih generalizing x + · simpa using h + · exact T.prop (ih (by simpa)) + +lemma mem_of_prefix {x y : List A} (h' : x <+: y) (h : y ∈ T) : x ∈ T := by + obtain ⟨_, rfl⟩ := h'; exact mem_of_append h + +@[simp] lemma tree_eq_bot : T = ⊥ ↔ [] ∉ T where + mp := by rintro rfl; simp + mpr h := by ext x; simpa using fun h' ↦ h <| mem_of_prefix x.nil_prefix h' + +end Descriptive diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index bc7e47200b09a..4e69a754b14ff 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Kim Morrison, Apurva Nakade +Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Kim Morrison, Apurva Nakade, Yuyang Zhao -/ import Mathlib.Algebra.Order.Group.Defs import Mathlib.SetTheory.Game.PGame @@ -27,8 +27,6 @@ namespace SetTheory open Function PGame -open PGame - universe u -- Porting note: moved the setoid instance to PGame.lean @@ -340,25 +338,25 @@ theorem mul_moveRight_inr {x y : PGame} {i j} : cases y rfl --- @[simp] -- Porting note: simpNF linter complains +@[simp] theorem neg_mk_mul_moveLeft_inl {xl xr yl yr} {xL xR yL yR} {i j} : (-(mk xl xr xL xR * mk yl yr yL yR)).moveLeft (Sum.inl (i, j)) = -(xL i * mk yl yr yL yR + mk xl xr xL xR * yR j - xL i * yR j) := rfl --- @[simp] -- Porting note: simpNF linter complains +@[simp] theorem neg_mk_mul_moveLeft_inr {xl xr yl yr} {xL xR yL yR} {i j} : (-(mk xl xr xL xR * mk yl yr yL yR)).moveLeft (Sum.inr (i, j)) = -(xR i * mk yl yr yL yR + mk xl xr xL xR * yL j - xR i * yL j) := rfl --- @[simp] -- Porting note: simpNF linter complains +@[simp] theorem neg_mk_mul_moveRight_inl {xl xr yl yr} {xL xR yL yR} {i j} : (-(mk xl xr xL xR * mk yl yr yL yR)).moveRight (Sum.inl (i, j)) = -(xL i * mk yl yr yL yR + mk xl xr xL xR * yL j - xL i * yL j) := rfl --- @[simp] -- Porting note: simpNF linter complains +@[simp] theorem neg_mk_mul_moveRight_inr {xl xr yl yr} {xL xR yL yR} {i j} : (-(mk xl xr xL xR * mk yl yr yL yR)).moveRight (Sum.inr (i, j)) = -(xR i * mk yl yr yL yR + mk xl xr xL xR * yR j - xR i * yR j) := @@ -380,6 +378,17 @@ theorem rightMoves_mul_cases {x y : PGame} (k) {P : (x * y).RightMoves → Prop} · apply hl · apply hr +/-- `x * y` and `y * x` have the same moves. -/ +protected lemma mul_comm (x y : PGame) : x * y ≡ y * x := + match x, y with + | ⟨xl, xr, xL, xR⟩, ⟨yl, yr, yL, yR⟩ => by + refine Identical.of_equiv ((Equiv.prodComm _ _).sumCongr (Equiv.prodComm _ _)) + ((Equiv.sumComm _ _).trans ((Equiv.prodComm _ _).sumCongr (Equiv.prodComm _ _))) ?_ ?_ <;> + · rintro (⟨_, _⟩ | ⟨_, _⟩) <;> + exact ((((PGame.mul_comm _ (mk _ _ _ _)).add (PGame.mul_comm (mk _ _ _ _) _)).trans + (PGame.add_comm _ _)).sub (PGame.mul_comm _ _)) + termination_by (x, y) + /-- `x * y` and `y * x` have the same moves. -/ def mulCommRelabelling (x y : PGame.{u}) : x * y ≡r y * x := match x, y with @@ -415,6 +424,9 @@ instance isEmpty_rightMoves_mul (x y : PGame.{u}) cases y assumption +/-- `x * 0` has exactly the same moves as `0`. -/ +protected lemma mul_zero (x : PGame) : x * 0 ≡ 0 := identical_zero _ + /-- `x * 0` has exactly the same moves as `0`. -/ def mulZeroRelabelling (x : PGame) : x * 0 ≡r 0 := Relabelling.isEmpty _ @@ -427,6 +439,9 @@ theorem mul_zero_equiv (x : PGame) : x * 0 ≈ 0 := theorem quot_mul_zero (x : PGame) : (⟦x * 0⟧ : Game) = 0 := game_eq x.mul_zero_equiv +/-- `0 * x` has exactly the same moves as `0`. -/ +protected lemma zero_mul (x : PGame) : 0 * x ≡ 0 := identical_zero _ + /-- `0 * x` has exactly the same moves as `0`. -/ def zeroMulRelabelling (x : PGame) : 0 * x ≡r 0 := Relabelling.isEmpty _ @@ -458,6 +473,23 @@ def negMulRelabelling (x y : PGame.{u}) : -x * y ≡r -(x * y) := exact (negMulRelabelling _ _).symm termination_by (x, y) +/-- `x * -y` and `-(x * y)` have the same moves. -/ +@[simp] +lemma mul_neg (x y : PGame) : x * -y = -(x * y) := + match x, y with + | mk xl xr xL xR, mk yl yr yL yR => by + refine ext rfl rfl ?_ ?_ <;> rintro (⟨i, j⟩ | ⟨i, j⟩) _ ⟨rfl⟩ + all_goals + dsimp + rw [PGame.neg_sub', PGame.neg_add] + congr + exacts [mul_neg _ (mk ..), mul_neg .., mul_neg ..] +termination_by (x, y) + +/-- `-x * y` and `-(x * y)` have the same moves. -/ +lemma neg_mul (x y : PGame) : -x * y ≡ -(x * y) := + ((PGame.mul_comm _ _).trans (of_eq (mul_neg _ _))).trans (PGame.mul_comm _ _).neg + @[simp] theorem quot_neg_mul (x y : PGame) : (⟦-x * y⟧ : Game) = -⟦x * y⟧ := game_eq (negMulRelabelling x y).equiv @@ -466,7 +498,6 @@ theorem quot_neg_mul (x y : PGame) : (⟦-x * y⟧ : Game) = -⟦x * y⟧ := def mulNegRelabelling (x y : PGame) : x * -y ≡r -(x * y) := (mulCommRelabelling x _).trans <| (negMulRelabelling _ x).trans (mulCommRelabelling y x).negCongr -@[simp] theorem quot_mul_neg (x y : PGame) : ⟦x * -y⟧ = (-⟦x * y⟧ : Game) := game_eq (mulNegRelabelling x y).equiv @@ -608,6 +639,18 @@ def mulOneRelabelling : ∀ x : PGame.{u}, x * 1 ≡r x exact (addZeroRelabelling _).trans <| (((mulOneRelabelling _).addCongr (mulZeroRelabelling _)).trans <| addZeroRelabelling _) } +/-- `1 * x` has the same moves as `x`. -/ +protected lemma one_mul : ∀ (x : PGame), 1 * x ≡ x + | ⟨xl, xr, xL, xR⟩ => by + refine Identical.of_equiv ((Equiv.sumEmpty _ _).trans (Equiv.punitProd _)) + ((Equiv.sumEmpty _ _).trans (Equiv.punitProd _)) ?_ ?_ <;> + · rintro (⟨⟨⟩, _⟩ | ⟨⟨⟩, _⟩) + exact ((((PGame.zero_mul (mk _ _ _ _)).add (PGame.one_mul _)).trans (PGame.zero_add _)).sub + (PGame.zero_mul _)).trans (PGame.sub_zero _) + +/-- `x * 1` has the same moves as `x`. -/ +protected lemma mul_one (x : PGame) : x * 1 ≡ x := (x.mul_comm _).trans x.one_mul + @[simp] theorem quot_mul_one (x : PGame) : (⟦x * 1⟧ : Game) = ⟦x⟧ := game_eq <| PGame.Relabelling.equiv <| mulOneRelabelling x @@ -922,6 +965,13 @@ def inv'Zero : inv' 0 ≡r 1 := by theorem inv'_zero_equiv : inv' 0 ≈ 1 := inv'Zero.equiv +/-- `inv' 1` has exactly the same moves as `1`. -/ +lemma inv'_one : inv' 1 ≡ 1 := by + rw [Identical.ext_iff] + constructor + · simp [memₗ_def, inv', isEmpty_subtype] + · simp [memᵣ_def, inv', isEmpty_subtype] + /-- `inv' 1` has exactly the same moves as `1`. -/ def inv'One : inv' 1 ≡r (1 : PGame.{u}) := by change Relabelling (mk _ _ _ _) 1 @@ -957,6 +1007,11 @@ theorem inv_eq_of_pos {x : PGame} (h : 0 < x) : x⁻¹ = inv' x := by theorem inv_eq_of_lf_zero {x : PGame} (h : x ⧏ 0) : x⁻¹ = -inv' (-x) := by classical exact (if_neg h.not_equiv).trans (if_neg h.not_gt) +/-- `1⁻¹` has exactly the same moves as `1`. -/ +lemma inv_one : 1⁻¹ ≡ 1 := by + rw [inv_eq_of_pos PGame.zero_lt_one] + exact inv'_one + /-- `1⁻¹` has exactly the same moves as `1`. -/ def invOne : 1⁻¹ ≡r 1 := by rw [inv_eq_of_pos PGame.zero_lt_one] diff --git a/Mathlib/SetTheory/Game/Impartial.lean b/Mathlib/SetTheory/Game/Impartial.lean index e06e5beb85b35..2d6cde0fec9aa 100644 --- a/Mathlib/SetTheory/Game/Impartial.lean +++ b/Mathlib/SetTheory/Game/Impartial.lean @@ -103,16 +103,11 @@ termination_by G variable (G : PGame) [Impartial G] theorem nonpos : ¬0 < G := by - intro h - have h' := neg_lt_neg_iff.2 h - rw [neg_zero, lt_congr_left (Equiv.symm (neg_equiv_self G))] at h' - exact (h.trans h').false + apply (lt_asymm · ?_) + rwa [← neg_lt_neg_iff, neg_zero, ← lt_congr_right (neg_equiv_self G)] theorem nonneg : ¬G < 0 := by - intro h - have h' := neg_lt_neg_iff.2 h - rw [neg_zero, lt_congr_right (Equiv.symm (neg_equiv_self G))] at h' - exact (h.trans h').false + simpa using nonpos (-G) /-- In an impartial game, either the first player always wins, or the second player always wins. -/ theorem equiv_or_fuzzy_zero : (G ≈ 0) ∨ G ‖ 0 := by diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 3ae1499b5e7eb..ac6caee7a3166 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -25,7 +25,8 @@ game after making a move). We may denote a game as $\{L | R\}$, where $L$ and $R$ stand for the collections of left and right moves. This notation is not currently used in Mathlib. -Combinatorial games themselves, as a quotient of pregames, are constructed in `Game.lean`. +Combinatorial games themselves, as a quotient of pregames, are constructed in +`SetTheory.Game.Basic`. ## Conway induction @@ -150,6 +151,16 @@ theorem rightMoves_mk {xl xr xL xR} : (⟨xl, xr, xL, xR⟩ : PGame).RightMoves theorem moveRight_mk {xl xr xL xR} : (⟨xl, xr, xL, xR⟩ : PGame).moveRight = xR := rfl +lemma ext {x y : PGame} (hl : x.LeftMoves = y.LeftMoves) (hr : x.RightMoves = y.RightMoves) + (hL : ∀ i j, HEq i j → x.moveLeft i = y.moveLeft j) + (hR : ∀ i j, HEq i j → x.moveRight i = y.moveRight j) : + x = y := by + cases x + cases y + subst hl hr + simp only [leftMoves_mk, rightMoves_mk, heq_eq_eq, forall_eq', mk.injEq, true_and] at * + exact ⟨funext hL, funext hR⟩ + -- TODO define this at the level of games, as well, and perhaps also for finsets of games. /-- Construct a pre-game from list of pre-games describing the available moves for Left and Right. -/ @@ -635,6 +646,8 @@ theorem lf_irrefl (x : PGame) : ¬x ⧏ x := instance : IsIrrefl _ (· ⧏ ·) := ⟨lf_irrefl⟩ +protected theorem not_lt {x y : PGame} : ¬ x < y ↔ y ⧏ x ∨ y ≤ x := not_lt_iff_not_le_or_ge + @[trans] theorem lf_of_le_of_lf {x y z : PGame} (h₁ : x ≤ y) (h₂ : y ⧏ z) : x ⧏ z := by rw [← PGame.not_le] at h₂ ⊢ @@ -1240,7 +1253,7 @@ instance : Neg PGame := ⟨neg⟩ @[simp] -theorem neg_def {xl xr xL xR} : -mk xl xr xL xR = mk xr xl (fun j => -xR j) fun i => -xL i := +theorem neg_def {xl xr xL xR} : -mk xl xr xL xR = mk xr xl (-xR ·) (-xL ·) := rfl instance : InvolutiveNeg PGame := @@ -1350,20 +1363,58 @@ theorem moveRight_neg_symm {x : PGame} (i) : theorem moveRight_neg_symm' {x : PGame} (i) : x.moveRight i = -(-x).moveLeft (toLeftMovesNeg i) := by simp -@[simp] theorem neg_identical_neg_iff : ∀ {x y : PGame.{u}}, -x ≡ -y ↔ x ≡ y - | mk xl xr xL xR, mk yl yr yL yR => by - rw [neg_def, identical_iff, identical_iff, ← neg_def, and_comm] - simp only [neg_def, rightMoves_mk, moveRight_mk, leftMoves_mk, moveLeft_mk] - apply and_congr <;> - · constructor - · conv in (_ ≡ _) => rw [neg_identical_neg_iff] - simp only [imp_self] - · conv in (_ ≡ _) => rw [← neg_identical_neg_iff] - simp only [imp_self] -termination_by x y => (x, y) +@[simp] +theorem forall_leftMoves_neg {x : PGame} {p : (-x).LeftMoves → Prop} : + (∀ i : (-x).LeftMoves, p i) ↔ (∀ i : x.RightMoves, p (toLeftMovesNeg i)) := + toLeftMovesNeg.forall_congr_right.symm + +@[simp] +theorem exists_leftMoves_neg {x : PGame} {p : (-x).LeftMoves → Prop} : + (∃ i : (-x).LeftMoves, p i) ↔ (∃ i : x.RightMoves, p (toLeftMovesNeg i)) := + toLeftMovesNeg.exists_congr_right.symm -theorem Identical.neg {x y : PGame} : x ≡ y ↔ -x ≡ -y := - neg_identical_neg_iff.symm +@[simp] +theorem forall_rightMoves_neg {x : PGame} {p : (-x).RightMoves → Prop} : + (∀ i : (-x).RightMoves, p i) ↔ (∀ i : x.LeftMoves, p (toRightMovesNeg i)) := + toRightMovesNeg.forall_congr_right.symm + +@[simp] +theorem exists_rightMoves_neg {x : PGame} {p : (-x).RightMoves → Prop} : + (∃ i : (-x).RightMoves, p i) ↔ (∃ i : x.LeftMoves, p (toRightMovesNeg i)) := + toRightMovesNeg.exists_congr_right.symm + +theorem leftMoves_neg_cases {x : PGame} (k) {P : (-x).LeftMoves → Prop} + (h : ∀ i, P <| toLeftMovesNeg i) : + P k := by + rw [← toLeftMovesNeg.apply_symm_apply k] + exact h _ + +theorem rightMoves_neg_cases {x : PGame} (k) {P : (-x).RightMoves → Prop} + (h : ∀ i, P <| toRightMovesNeg i) : + P k := by + rw [← toRightMovesNeg.apply_symm_apply k] + exact h _ + +/-- If `x` has the same moves as `y`, then `-x` has the sames moves as `-y`. -/ +lemma Identical.neg : ∀ {x₁ x₂ : PGame}, x₁ ≡ x₂ → -x₁ ≡ -x₂ + | mk _ _ _ _, mk _ _ _ _, ⟨⟨hL₁, hL₂⟩, ⟨hR₁, hR₂⟩⟩ => + ⟨⟨fun i ↦ (hR₁ i).imp (fun _ ↦ Identical.neg), fun j ↦ (hR₂ j).imp (fun _ ↦ Identical.neg)⟩, + ⟨fun i ↦ (hL₁ i).imp (fun _ ↦ Identical.neg), fun j ↦ (hL₂ j).imp (fun _ ↦ Identical.neg)⟩⟩ + +/-- If `-x` has the same moves as `-y`, then `x` has the sames moves as `y`. -/ +lemma Identical.of_neg : ∀ {x₁ x₂ : PGame}, -x₁ ≡ -x₂ → x₁ ≡ x₂ + | mk x₁l x₁r x₁L x₁R, mk x₂l x₂r x₂L x₂R => by + simpa using Identical.neg (x₁ := mk _ _ (-x₁R ·) (-x₁L ·)) (x₂ := mk _ _ (-x₂R ·) (-x₂L ·)) + +lemma memₗ_neg_iff : ∀ {x y : PGame}, + x ∈ₗ -y ↔ ∃ z ∈ᵣ y, x ≡ -z + | mk _ _ _ _, mk _ _ _ _ => + ⟨fun ⟨_i, hi⟩ ↦ ⟨_, ⟨_, refl _⟩, hi⟩, fun ⟨_, ⟨i, hi⟩, h⟩ ↦ ⟨i, h.trans hi.neg⟩⟩ + +lemma memᵣ_neg_iff : ∀ {x y : PGame}, + x ∈ᵣ -y ↔ ∃ z ∈ₗ y, x ≡ -z + | mk _ _ _ _, mk _ _ _ _ => + ⟨fun ⟨_i, hi⟩ ↦ ⟨_, ⟨_, refl _⟩, hi⟩, fun ⟨_, ⟨i, hi⟩, h⟩ ↦ ⟨i, h.trans hi.neg⟩⟩ /-- If `x` has the same moves as `y`, then `-x` has the same moves as `-y`. -/ def Relabelling.negCongr : ∀ {x y : PGame}, x ≡r y → -x ≡r -y @@ -1393,7 +1444,11 @@ theorem neg_lt_neg_iff {x y : PGame} : -y < -x ↔ x < y := by rw [lt_iff_le_and_lf, lt_iff_le_and_lf, neg_le_neg_iff, neg_lf_neg_iff] @[simp] -theorem neg_equiv_neg_iff {x y : PGame} : (-x ≈ -y) ↔ (x ≈ y) := by +theorem neg_identical_neg {x y : PGame} : -x ≡ -y ↔ x ≡ y := + ⟨Identical.of_neg, Identical.neg⟩ + +@[simp] +theorem neg_equiv_neg_iff {x y : PGame} : -x ≈ -y ↔ x ≈ y := by show Equiv (-x) (-y) ↔ Equiv x y rw [Equiv, Equiv, neg_le_neg_iff, neg_le_neg_iff, and_comm] @@ -1463,6 +1518,16 @@ instance : Add PGame.{u} := · exact fun i => IHxr i y · exact IHyr⟩ +theorem mk_add_moveLeft {xl xr yl yr} {xL xR yL yR} {i} : + (mk xl xr xL xR + mk yl yr yL yR).moveLeft i = + i.rec (xL · + mk yl yr yL yR) (mk xl xr xL xR + yL ·) := + rfl + +theorem mk_add_moveRight {xl xr yl yr} {xL xR yL yR} {i} : + (mk xl xr xL xR + mk yl yr yL yR).moveRight i = + i.rec (xR · + mk yl yr yL yR) (mk xl xr xL xR + yR ·) := + rfl + /-- The pre-game `((0 + 1) + ⋯) + 1`. Note that this is **not** the usual recursive definition `n = {0, 1, … | }`. For instance, @@ -1610,6 +1675,110 @@ instance isEmpty_nat_rightMoves : ∀ n : ℕ, IsEmpty (RightMoves n) rw [PGame.nat_succ, rightMoves_add] infer_instance +/-- `x + y` has exactly the same moves as `y + x`. -/ +protected lemma add_comm (x y : PGame) : x + y ≡ y + x := + match x, y with + | mk xl xr xL xR, mk yl yr yL yR => by + refine Identical.of_equiv (Equiv.sumComm _ _) (Equiv.sumComm _ _) ?_ ?_ <;> + · rintro (_ | _) <;> + · dsimp; exact PGame.add_comm _ _ + termination_by (x, y) + +/-- `(x + y) + z` has exactly the same moves as `x + (y + z)`. -/ +protected lemma add_assoc (x y z : PGame) : x + y + z ≡ x + (y + z) := + match x, y, z with + | mk xl xr xL xR, mk yl yr yL yR, mk zl zr zL zR => by + refine Identical.of_equiv (Equiv.sumAssoc _ _ _) (Equiv.sumAssoc _ _ _) ?_ ?_ <;> + · rintro ((_ | _) | _) + · exact PGame.add_assoc _ _ _ + · exact PGame.add_assoc (mk _ _ _ _) _ _ + · exact PGame.add_assoc (mk _ _ _ _) (mk _ _ _ _) _ + termination_by (x, y, z) + +/-- `x + 0` has exactly the same moves as `x`. -/ +protected lemma add_zero : ∀ (x : PGame), x + 0 ≡ x + | mk xl xr xL xR => by + refine Identical.of_equiv (Equiv.sumEmpty _ _) (Equiv.sumEmpty _ _) ?_ ?_ <;> + · rintro (_ | ⟨⟨⟩⟩) + exact PGame.add_zero _ + +/-- `0 + x` has exactly the same moves as `x`. -/ +protected lemma zero_add (x : PGame) : 0 + x ≡ x := + (PGame.add_comm _ _).trans x.add_zero + +/-- `-(x + y)` has exactly the same moves as `-x + -y`. -/ +protected lemma neg_add (x y : PGame) : -(x + y) = -x + -y := + match x, y with + | mk xl xr xL xR, mk yl yr yL yR => by + refine ext rfl rfl ?_ ?_ <;> + · rintro (i | i) _ ⟨rfl⟩ + · exact PGame.neg_add _ _ + · simpa [Equiv.refl, mk_add_moveLeft, mk_add_moveRight] using PGame.neg_add _ _ + termination_by (x, y) + +/-- `-(x + y)` has exactly the same moves as `-y + -x`. -/ +protected lemma neg_add_rev (x y : PGame) : -(x + y) ≡ -y + -x := + Identical.trans (of_eq (x.neg_add y)) (PGame.add_comm _ _) + +lemma identical_zero_iff : ∀ (x : PGame), + x ≡ 0 ↔ IsEmpty x.LeftMoves ∧ IsEmpty x.RightMoves + | mk xl xr xL xR => by + constructor + · rintro ⟨h₁, h₂⟩ + dsimp [Relator.BiTotal, Relator.LeftTotal, Relator.RightTotal] at h₁ h₂ + simp_rw [IsEmpty.forall_iff, and_true, IsEmpty.exists_iff] at h₁ h₂ + exact ⟨⟨h₁⟩, ⟨h₂⟩⟩ + · rintro ⟨h₁, h₂⟩ + exact identical_of_isEmpty _ _ + +/-- Any game without left or right moves is identical to 0. -/ +lemma identical_zero (x : PGame) [IsEmpty x.LeftMoves] [IsEmpty x.RightMoves] : x ≡ 0 := + x.identical_zero_iff.mpr ⟨by infer_instance, by infer_instance⟩ + +protected lemma add_eq_zero : ∀ {x y : PGame}, x + y ≡ 0 ↔ x ≡ 0 ∧ y ≡ 0 + | mk xl xr xL xR, mk yl yr yL yR => by + simp_rw [identical_zero_iff, leftMoves_add, rightMoves_add, isEmpty_sum] + tauto + +lemma Identical.add_right {x₁ x₂ y} : x₁ ≡ x₂ → x₁ + y ≡ x₂ + y := + match x₁, x₂, y with + | mk x₁l x₁r x₁L x₁R, mk x₂l x₂r x₂L x₂R, mk yl yr yL yR => by + intro h + refine ⟨⟨?_, ?_⟩, ⟨?_, ?_⟩⟩ <;> rintro (_ | _) <;> try exact ⟨.inr _, h.add_right⟩ + · exact (h.1.1 _).elim (⟨.inl ·, ·.add_right⟩) + · exact (h.1.2 _).elim (⟨.inl ·, ·.add_right⟩) + · exact (h.2.1 _).elim (⟨.inl ·, ·.add_right⟩) + · exact (h.2.2 _).elim (⟨.inl ·, ·.add_right⟩) + termination_by (x₁, x₂, y) + +lemma Identical.add_left {x y₁ y₂} (hy : y₁ ≡ y₂) : x + y₁ ≡ x + y₂ := + (x.add_comm y₁).trans (hy.add_right.trans (y₂.add_comm x)) + +/-- If `w` has the same moves as `x` and `y` has the same moves as `z`, +then `w + y` has the same moves as `x + z`. -/ +lemma Identical.add {x₁ x₂ y₁ y₂ : PGame.{u}} (hx : x₁ ≡ x₂) (hy : y₁ ≡ y₂) : x₁ + y₁ ≡ x₂ + y₂ := + hx.add_right.trans hy.add_left + +lemma memₗ_add_iff {x y₁ y₂ : PGame} : + x ∈ₗ y₁ + y₂ ↔ (∃ z ∈ₗ y₁, x ≡ z + y₂) ∨ (∃ z ∈ₗ y₂, x ≡ y₁ + z) := by + cases' y₁ with y₁l y₁r y₁L y₁R + cases' y₂ with y₂l y₂r y₂L y₂R + constructor + · rintro ⟨(i | i), hi⟩ + exacts [.inl ⟨y₁L i, moveLeft_memₗ _ _, hi⟩, .inr ⟨y₂L i, moveLeft_memₗ _ _, hi⟩] + · rintro (⟨_, ⟨i, hi⟩, h⟩ | ⟨_, ⟨i, hi⟩, h⟩) + exacts [⟨.inl i, h.trans hi.add_right⟩, ⟨.inr i, h.trans hi.add_left⟩] + +lemma memᵣ_add_iff {x y₁ y₂ : PGame} : + x ∈ᵣ y₁ + y₂ ↔ (∃ z ∈ᵣ y₁, x ≡ z + y₂) ∨ (∃ z ∈ᵣ y₂, x ≡ y₁ + z) := by + cases' y₁ with y₁l y₁r y₁L y₁R + cases' y₂ with y₂l y₂r y₂L y₂R + constructor + · rintro ⟨(i | i), hi⟩ + exacts [.inl ⟨y₁R i, moveRight_memᵣ _ _, hi⟩, .inr ⟨y₂R i, moveRight_memᵣ _ _, hi⟩] + · rintro (⟨_, ⟨i, hi⟩, h⟩ | ⟨_, ⟨i, hi⟩, h⟩) + exacts [⟨.inl i, h.trans hi.add_right⟩, ⟨.inr i, h.trans hi.add_left⟩] + /-- If `w` has the same moves as `x` and `y` has the same moves as `z`, then `w + y` has the same moves as `x + z`. -/ def Relabelling.addCongr : ∀ {w x y z : PGame.{u}}, w ≡r x → y ≡r z → w + y ≡r x + z @@ -1631,7 +1800,16 @@ instance : Sub PGame := theorem sub_zero_eq_add_zero (x : PGame) : x - 0 = x + 0 := show x + -0 = x + 0 by rw [neg_zero] -@[deprecated (since := "2024-09-26")] alias sub_zero := sub_zero_eq_add_zero +protected lemma sub_zero (x : PGame) : x - 0 ≡ x := + _root_.trans (of_eq x.sub_zero_eq_add_zero) x.add_zero + +/-- This lemma is named to match `neg_sub'`. -/ +protected lemma neg_sub' (x y : PGame) : -(x - y) = -x - -y := PGame.neg_add _ _ + +/-- If `w` has the same moves as `x` and `y` has the same moves as `z`, +then `w - y` has the same moves as `x - z`. -/ +lemma Identical.sub {x₁ x₂ y₁ y₂ : PGame.{u}} (hx : x₁ ≡ x₂) (hy : y₁ ≡ y₂) : x₁ - y₁ ≡ x₂ - y₂ := + hx.add hy.neg /-- If `w` has the same moves as `x` and `y` has the same moves as `z`, then `w - y` has the same moves as `x - z`. -/ @@ -2027,4 +2205,4 @@ end PGame end SetTheory -set_option linter.style.longFile 2100 +set_option linter.style.longFile 2300 diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 1500096b7b3a1..f09bda8e413e2 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -198,9 +198,6 @@ theorem principal_add_iff_add_lt_ne_self : Principal (· + ·) a ↔ ∀ b < a, theorem principal_add_omega0 : Principal (· + ·) ω := principal_add_iff_add_left_eq_self.2 fun _ => add_omega0 -@[deprecated (since := "2024-09-30")] -alias principal_add_omega := principal_add_omega0 - theorem add_omega0_opow (h : a < ω ^ b) : a + ω ^ b = ω ^ b := by refine le_antisymm ?_ (le_add_left _ a) induction' b using limitRecOn with b _ b l IH @@ -352,9 +349,6 @@ theorem principal_mul_omega0 : Principal (· * ·) ω := fun a b ha hb => dsimp only; rw [← natCast_mul] apply nat_lt_omega0 -@[deprecated (since := "2024-09-30")] -alias principal_mul_omega := principal_mul_omega0 - theorem mul_omega0 (a0 : 0 < a) (ha : a < ω) : a * ω = ω := principal_mul_iff_mul_left_eq.1 principal_mul_omega0 a a0 ha diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index c30277fabaaaf..fda897921c1fe 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -80,6 +80,7 @@ import Mathlib.Tactic.ExtractGoal import Mathlib.Tactic.ExtractLets import Mathlib.Tactic.FBinop import Mathlib.Tactic.FailIfNoProgress +import Mathlib.Tactic.FastInstance import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.FinCases import Mathlib.Tactic.Find diff --git a/Mathlib/Tactic/CasesM.lean b/Mathlib/Tactic/CasesM.lean index a1b012c5b4baf..190ae12430667 100644 --- a/Mathlib/Tactic/CasesM.lean +++ b/Mathlib/Tactic/CasesM.lean @@ -40,17 +40,21 @@ partial def casesMatching (matcher : Expr → MetaM Bool) (recursive := false) ( g.cases ldecl.fvarId else let s ← saveState - let subgoals ← g.cases ldecl.fvarId + let subgoals ← g.cases ldecl.fvarId (givenNames := #[⟨true, [ldecl.userName]⟩]) if subgoals.size > 1 then s.restore continue else pure subgoals for subgoal in subgoals do + -- If only one new hypothesis is generated, rename it to the original name. + let g ← match subgoal.fields with + | #[.fvar fvarId] => subgoal.mvarId.rename fvarId ldecl.userName + | _ => pure subgoal.mvarId if recursive then - acc ← go subgoal.mvarId acc + acc ← go g acc else - acc := acc.push subgoal.mvarId + acc := acc.push g return acc return (acc.push g) @@ -77,6 +81,11 @@ def matchPatterns (pats : Array AbstractMVarsResult) (e : Expr) : MetaM Bool := let e ← instantiateMVars e pats.anyM fun p ↦ return (← Conv.matchPattern? p e) matches some (_, #[]) +/-- Common implementation of `casesm` and `casesm!`. -/ +def elabCasesM (pats : Array Term) (recursive allowSplit : Bool) : TacticM Unit := do + let pats ← elabPatterns pats + liftMetaTactic (casesMatching (matchPatterns pats) recursive allowSplit) + /-- * `casesm p` applies the `cases` tactic to a hypothesis `h : type` if `type` matches the pattern `p`. @@ -84,6 +93,7 @@ def matchPatterns (pats : Array AbstractMVarsResult) (e : Expr) : MetaM Bool := if `type` matches one of the given patterns. * `casesm* p` is a more efficient and compact version of `· repeat casesm p`. It is more efficient because the pattern is compiled once. +* `casesm! p` only applies `cases` if the number of resulting subgoals is <= 1. Example: The following tactic destructs all conjunctions and disjunctions in the current context. ``` @@ -91,8 +101,11 @@ casesm* _ ∨ _, _ ∧ _ ``` -/ elab (name := casesM) "casesm" recursive:"*"? ppSpace pats:term,+ : tactic => do - let pats ← elabPatterns pats.getElems - liftMetaTactic (casesMatching (matchPatterns pats) recursive.isSome) + elabCasesM pats recursive.isSome true + +@[inherit_doc casesM] +elab (name := casesm!) "casesm!" recursive:"*"? ppSpace pats:term,+ : tactic => do + elabCasesM pats recursive.isSome false /-- Common implementation of `cases_type` and `cases_type!`. -/ def elabCasesType (heads : Array Ident) diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index f94993bf7de95..f82b8e0189765 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -440,14 +440,14 @@ def miscomputedDegree? (deg : Expr) : List Expr → List MessageData * `degree f ≤ d`, * `coeff f d = r`, if `d` is the degree of `f`. -The tactic may leave goals of the form `d' = d` `d' ≤ d`, or `r ≠ 0`, where `d'` in `ℕ` or +The tactic may leave goals of the form `d' = d`, `d' ≤ d`, or `r ≠ 0`, where `d'` in `ℕ` or `WithBot ℕ` is the tactic's guess of the degree, and `r` is the coefficient's guess of the leading coefficient of `f`. `compute_degree` applies `norm_num` to the left-hand side of all side goals, trying to close them. The variant `compute_degree!` first applies `compute_degree`. -Then it uses `norm_num` on all the whole remaining goals and tries `assumption`. +Then it uses `norm_num` on all the remaining goals and tries `assumption`. -/ syntax (name := computeDegree) "compute_degree" "!"? : tactic diff --git a/Mathlib/Tactic/FastInstance.lean b/Mathlib/Tactic/FastInstance.lean new file mode 100644 index 0000000000000..45ae5c3925ff1 --- /dev/null +++ b/Mathlib/Tactic/FastInstance.lean @@ -0,0 +1,132 @@ +/- +Copyright (c) 2024 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Kyle Miller +-/ + +import Lean.Elab.SyntheticMVars +import Lean +import Mathlib.Init +/-! +# The `fast_instance%` term elaborator +-/ + +namespace Mathlib.Elab.FastInstance + +open Lean Meta Elab Term + +initialize registerTraceClass `Elab.fast_instance + +/-- +Throw an error for `makeFastInstance`. The trace is a list of fields. +Note: with the current implementation, this might not be accurate for multi-structure types, +since `makeFastInstance` just uses `ClassName.paramName` for the trace. +-/ +private def error {α : Type _} (trace : Array Name) (m : MessageData) : MetaM α := + throwError "\ + {m}\n\n\ + Use `set_option trace.Elab.fast_instance true` to analyze the error.\n\n\ + Trace of fields visited: {trace}" + +/-- +Core algorithm for normalizing instances. +* Ideally, the term is replaced with a synthesized instance. +* If not, it is reduced to a constructor + and each instance implicit field is given the same treatment. + If the type is a structure, the algorithm throws an error; + we're more lenient with non-structure classes. + +Many reductions for typeclasses are done with reducible transparency, so the entire body +is `withReducible` with some exceptions. +-/ +private partial def makeFastInstance (provided : Expr) (trace : Array Name := #[]) : + MetaM Expr := withReducible do + let ty ← inferType provided + withTraceNode `Elab.fast_instance (fun e => return m!"{exceptEmoji e} type: {ty}") do + let .some className ← isClass? ty + | error trace m!"Can only be used for classes, but term has type{indentExpr ty}" + trace[Elab.fast_instance] "class is {className}" + if ← withDefault <| Meta.isProp ty then + error trace m!"\ + Provided instance{indentExpr provided}\n\ + is a proof, which does not need normalization." + + -- Try to synthesize a total replacement for this term: + if let .some new ← trySynthInstance ty then + if ← withReducibleAndInstances <| isDefEq provided new then + trace[Elab.fast_instance] "replaced with synthesized instance" + return new + else + if ← withDefault <| isDefEq provided new then + error trace m!"\ + Provided instance{indentExpr provided}\n\ + is defeq only at default transparency to inferred instance{indentExpr new}" + else + error trace m!"\ + Provided instance{indentExpr provided}\n\ + is not defeq to inferred instance{indentExpr new}" + -- Otherwise, try to reduce it to a constructor. + else + -- Telescope since it might be a family of instances. + forallTelescopeReducing ty fun tyArgs _ => do + let provided' ← withReducibleAndInstances <| whnf <| mkAppN provided tyArgs + let error' (m : MessageData) : MetaM Expr := do + if isStructure (← getEnv) className then + error trace m + else + error trace m!"{m}\n\n\ + This instance is not a structure and not canonical. \ + Use a separate 'instance' command to define it." + let .const c .. := provided'.getAppFn + | error' m!"\ + Provided instance does not reduce to a constructor application{indentExpr provided}" + let some (.ctorInfo ci) := (← getEnv).find? c + | error' m!"\ + Provided instance does not reduce to a constructor application{indentExpr provided}\n\ + Reduces to an application of {c}." + let mut args := provided'.getAppArgs + let params ← withDefault <| forallTelescopeReducing ci.type fun args _ => + args.mapM fun arg => do + let recurse ← (return (← arg.fvarId!.getBinderInfo).isInstImplicit) + <&&> not <$> Meta.isProof arg + return (recurse, ← arg.fvarId!.getUserName) + unless args.size == params.size do + -- This is an invalid term. + throwError "Incorrect number of arguments for constructor application{indentExpr provided'}" + for i in [ci.numParams:args.size] do + let (recurse, binderName) := params[i]! + if recurse then + let trace' := trace.push (className ++ binderName) + args := args.set! i (← makeFastInstance args[i]! (trace := trace')) + let provided' := mkAppN provided'.getAppFn args + mkLambdaFVars tyArgs provided' + +/-- +`fast_instance% inst` takes an expression for a typeclass instance `inst`, and unfolds it into +constructor applications that leverage existing instances. + +For instance, when used as +```lean +instance instSemiring : Semiring X := sorry +instance instRing : Ring X := fast_instance% Function.Injective.ring .. +``` +this will define `instRing` as a nested constructor application that refers to `instSemiring` +rather than applications of `Function.Injective.ring` or other non-canonical constructors. +The advantage is then that `instRing.toSemiring` unifies almost immediately with `instSemiring`, +rather than having to break it down into smaller pieces. +-/ +syntax (name := fastInstance) "fast_instance%" term : term + +@[term_elab fastInstance, inherit_doc fastInstance] +def elabFastInstance : TermElab + | `(term| fast_instance%%$tk $arg), expectedType => do + let provided ← withSynthesize <| elabTerm arg expectedType + withRef tk do + try + makeFastInstance provided + catch e => + logException e + return provided + | _, _ => Elab.throwUnsupportedSyntax + +end Mathlib.Elab.FastInstance diff --git a/Mathlib/Tactic/FunProp.lean b/Mathlib/Tactic/FunProp.lean index 1e9f65989489a..a52e272d5b956 100644 --- a/Mathlib/Tactic/FunProp.lean +++ b/Mathlib/Tactic/FunProp.lean @@ -37,7 +37,7 @@ example (y : ℝ) (hy : y ≠ 0) : ContinuousAt (fun x : ℝ => 1/x) y := by fun **Basic debugging:** The most common issue is that a function is missing the appropriate theorem. For example: ```lean -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric example : Continuous (fun x : ℝ => x * Real.sin x) := by fun_prop ``` Fails with the error: diff --git a/Mathlib/Tactic/Linter/MinImports.lean b/Mathlib/Tactic/Linter/MinImports.lean index 0aad206e77cf8..9f72cfba03c25 100644 --- a/Mathlib/Tactic/Linter/MinImports.lean +++ b/Mathlib/Tactic/Linter/MinImports.lean @@ -61,6 +61,10 @@ Unlike the related `#min_imports` command, the linter takes into account notatio information. It also works incrementally, providing information that is better suited, for instance, to split files. + +Another important difference is that the `minImports` *linter* starts counting imports from +where the option is set to `true` *downwards*, whereas the `#min_imports` *command* looks at the +imports needed from the command *upwards*. -/ register_option linter.minImports : Bool := { defValue := false @@ -87,13 +91,25 @@ It returns the modules that are transitively imported by `ms`, using the data in def importsBelow (tc : NameMap NameSet) (ms : NameSet) : NameSet := ms.fold (·.append <| tc.findD · default) ms +@[inherit_doc Mathlib.Linter.linter.minImports] +macro "#import_bumps" : command => `( + -- We emit a message to prevent the `#`-command linter from flagging `#import_bumps`. + run_cmd logInfo "Counting imports from here." + set_option Elab.async false + set_option linter.minImports true) + @[inherit_doc Mathlib.Linter.linter.minImports] def minImportsLinter : Linter where run := withSetOptionIn fun stx ↦ do unless Linter.getLinterValue linter.minImports (← getOptions) do return if (← get).messages.hasErrors then return - if stx == (← `(command| set_option $(mkIdent `linter.minImports) true)) then return + if stx == (← `(command| #import_bumps)) then return + if stx == (← `(command| set_option $(mkIdent `linter.minImports) true)) then + logInfo "Try using '#import_bumps', instead of manually setting the linter option: \ + the linter works best with linear parsing of the file and '#import_bumps' \ + also sets the `Elab.async` option to `false`." + return let env ← getEnv -- the first time `minImportsRef` is read, it has `transClosure = none`; -- in this case, we set it to be the `transClosure` for the file. diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index bbe2f16f58a98..efae3b7795163 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -17,7 +17,7 @@ but do not affect correctness nor global coherence of mathlib. Historically, some of these were ported from the `lint-style.py` Python script. This file defines the following linters: -- the `set_option` linter checks for the presence of `set_option` commands activating +- the `setOption` linter checks for the presence of `set_option` commands activating options disallowed in mathlib: these are meant to be temporary, and not for polished code - the `missingEnd` linter checks for sections or namespaces which are not closed by the end of the file: enforcing this invariant makes minimising files or moving code between files easier @@ -26,12 +26,10 @@ this is allowed Lean syntax, but it is nicer to be uniform - the `dollarSyntax` linter checks for use of the dollar sign `$` instead of the `<|` pipe operator: similarly, both symbols have the same meaning, but mathlib prefers `<|` for the symmetry with the `|>` symbol -- the `lambdaSyntax` linter checks for uses of the `λ` symbol for ananomous functions, +- the `lambdaSyntax` linter checks for uses of the `λ` symbol for anonymous functions, instead of the `fun` keyword: mathlib prefers the latter for reasons of readability -(This linter is still under review in PR #15896.) +- the `longFile` linter checks for files which have more than 1500 lines - the `longLine` linter checks for lines which have more than 100 characters -- the `longFile` linter checks for files which have more than 1500 lines: -this linter is still under development in PR #15610. All of these linters are enabled in mathlib by default, but disabled globally since they enforce conventions which are inherently subjective. diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 91b368722375e..dd9c5a039957b 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Dhruv Bhatia. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Dhruv Bhatia, Eric Wieser, Mario Carneiro +Authors: Dhruv Bhatia, Eric Wieser, Mario Carneiro, Thomas Zhu -/ import Mathlib.Tactic.LinearCombination @@ -15,8 +15,8 @@ field by using multivariable polynomial hypotheses/proof terms over the same fie Used as is, the tactic makes use of those hypotheses in the local context that are over the same field as the target. However, the user can also specify which hypotheses from the local context to use, along with proof terms that might not already be in the -local context. Note: since this tactic uses SageMath via an API call done in Python, -it can only be used with a working internet connection, and with a local installation of Python. +local context. Note: since this tactic uses SageMath via an API call, +it can only be used with a working internet connection. ## Implementation Notes @@ -27,21 +27,16 @@ tactic succeeds. In other words, `linear_combination` is a certificate checker, to the user to find a collection of good coefficients. The `polyrith` tactic automates this process using the theory of Groebner bases. -Polyrith does this by first parsing the relevant hypotheses into a form that Python can understand. -It then calls a Python file that uses the SageMath API to compute the coefficients. These -coefficients are then sent back to Lean, which parses them into pexprs. The information is then -given to the `linear_combination` tactic, which completes the process by checking the certificate. +Polyrith does this by first parsing the relevant hypotheses into a form that SageMath can +understand. It then calls the SageMath API to compute the coefficients. These coefficients are +then sent back to Lean, which parses them into pexprs. The information is then given to the +`linear_combination` tactic, which completes the process by checking the certificate. In fact, `polyrith` uses Sage to test for membership in the *radical* of the ideal. This means it searches for a linear combination of hypotheses that add up to a *power* of the goal. When this power is not 1, it uses the `(exp := n)` feature of `linear_combination` to report the certificate. -`polyrith` calls an external python script `scripts/polyrith_sage.py`. Because this is not a Lean -file, changes to this script may not be noticed during Lean compilation if you have already -generated olean files. If you are modifying this python script, you likely know what you're doing; -remember to force recompilation of any files that call `polyrith`. - ## TODO * Give Sage more information about the specific ring being used for the coefficients. For now, @@ -86,12 +81,12 @@ inductive Poly This converts a poly object into a string representing it. The string maintains the semantic structure of the poly object. -The output of this function must be valid Python syntax, and it assumes the variables `varN` from -`scripts/polyrith.py.` +The output of this function must be valid Python syntax, and it assumes the variables `vars` +(see `sageCreateQuery`). -/ def Poly.format : Poly → Lean.Format | .const z => toString z - | .var n => s!"var{n}" + | .var n => s!"vars[{n}]" -- this references variable `vars`, which need to be bounded (below) | .hyp e => s!"hyp{e}" -- this one can't be used by python | .add p q => s!"({p.format} + {q.format})" | .sub p q => s!"({p.format} - {q.format})" @@ -165,17 +160,17 @@ inductive Source where | fvar : FVarId → Source /-- The first half of `polyrith` produces a list of arguments to be sent to Sage. -/ -def parseContext (only : Bool) (hyps : Array Expr) (tgt : Expr) : +def parseContext (only : Bool) (hyps : Array Expr) (target : Expr) : AtomM (Expr × Array (Source × Poly) × Poly) := do let fail {α} : AtomM α := throwError "polyrith failed: target is not an equality in semirings" - let some (α, e₁, e₂) := (← whnfR <|← instantiateMVars tgt).eq? | fail + let some (α, e₁, e₂) := (← whnfR <|← instantiateMVars target).eq? | fail let .sort u ← instantiateMVars (← whnf (← inferType α)) | unreachable! let some v := u.dec | throwError "not a type{indentExpr α}" have α : Q(Type v) := α have e₁ : Q($α) := e₁; have e₂ : Q($α) := e₂ let sα ← synthInstanceQ (q(CommSemiring $α) : Q(Type v)) let c ← mkCache sα - let tgt := (← parse sα c e₁).sub (← parse sα c e₂) + let target := (← parse sα c e₁).sub (← parse sα c e₂) let rec /-- Parses a hypothesis and adds it to the `out` list. -/ processHyp src ty out := do @@ -189,15 +184,9 @@ def parseContext (only : Bool) (hyps : Array Expr) (tgt : Expr) : out ← processHyp (.fvar ldecl.fvarId) ldecl.type out for hyp in hyps, i in [:hyps.size] do out ← processHyp (.input i) (← inferType hyp) out - pure (α, out, tgt) - -/-- Constructs the list of arguments to pass to the external sage script `polyrith_sage.py`. -/ -def createSageArgs (trace : Bool) (α : Expr) (atoms : Nat) - (hyps : Array (Source × Poly)) (tgt : Poly) : Array String := - let hyps := hyps.map (toString ·.2) |>.toList.toString - #[toString trace, toString α, toString atoms, hyps, toString tgt] + pure (α, out, target) -/-- A JSON parser for `ℚ` specific to the return value of `polyrith_sage.py`. -/ +/-- A JSON parser for `ℚ` specific to the return value of Sage. -/ local instance : FromJson ℚ where fromJson? | .arr #[a, b] => return (← fromJson? a : ℤ) / (← fromJson? b : ℕ) | _ => .error "expected array with two elements" @@ -284,28 +273,100 @@ structure SageError where /-- The result of a sage call. -/ def SageResult := Except SageError SageSuccess + +section SageMathApi + +/-! # Interaction with SageMath -/ + +/-- +These are Sage functions that test membership in the radical and format the output. See +https://github.com/sagemath/sage/blob/f8df80820dc7321dc9b18c9644c3b8315999670b/src/sage/rings/polynomial/multi_polynomial_libsingular.pyx#L4472-L4518 +for a description of `MPolynomial_libsingular.lift`. +-/ +def sageHelperFunctions : String := + include_str ".."/".."/"scripts"/"polyrith_sage_helper.py" + +/-- +The Sage type to use, given a base type of the target. Currently always rational numbers (`QQ`). +Future extensions may change behavior depending on the base type. +-/ +@[nolint unusedArguments] +def sageTypeStr (_ : Expr) : String := "QQ" + +/-- +Create a Sage script to send to SageMath API. +-/ +def sageCreateQuery (α : Expr) (atoms : Nat) (hyps : Array (Source × Poly)) (target : Poly) : + IO String := do + -- since `hyps` and `target` reference the variable list `vars` (see `Poly.format`), + -- `hyps` and `target` are wrapped as functions where `vars` is bounded using `lambda` + let hypsListPython := "[" ++ ",".intercalate (hyps.map (toString ·.2) |>.toList) ++ "]" + let mkHypsListPython := s!"lambda vars: {hypsListPython}" + let mkTargetPython := s!"lambda vars: {target}" + let command := + s!"main({sageTypeStr α}, {atoms}, {mkHypsListPython}, {mkTargetPython})" + return sageHelperFunctions ++ "\n" ++ command + +/-- Parse a `SageResult` from the raw SageMath API output. -/ instance : FromJson SageResult where fromJson? j := do - if let .ok true := fromJson? <| j.getObjValD "success" then - return .ok (← fromJson? j) + -- we expect the output has either "success": true and contains "stdout", + -- or has "success": false and error information under "execute_reply" + if let .ok true := j.getObjValAs? Bool "success" then + -- parse SageSuccess from stdout, which is formatted as SageCoeffAndPower + -- (see sageCreateQuery for the format of stdout) + let stdout ← j.getObjValAs? String "stdout" + let coeffAndPower ← Json.parse stdout >>= fromJson? + let sageSuccess := { data := some coeffAndPower } + return .ok sageSuccess else - return .error (← fromJson? j) + -- parse SageError from execute_reply + let executeReply ← j.getObjVal? "execute_reply" + let errorName ← executeReply.getObjValAs? String "ename" + let errorValue ← executeReply.getObjValAs? String "evalue" + return .error { name := errorName, value := errorValue } + +/-- The User-Agent header value for HTTP calls to SageMath API -/ +register_option polyrith.sageUserAgent : String := + { defValue := "LeanProver (https://leanprover-community.github.io/)" + group := "polyrith" + descr := "The User-Agent header value for HTTP calls to SageMath API" } /-- -This tactic calls python from the command line with the args in `arg_list`. -The output printed to the console is parsed as a `Json`. -It assumes that `python3` is available on the path. +This function calls the Sage API at . +The output is parsed as `SageResult`. -/ -def sageOutput (args : Array String) : IO SageResult := do - let path := (← getMathlibDir) / "scripts" / "polyrith_sage.py" - unless ← path.pathExists do - throw <| IO.userError "could not find python script scripts/polyrith_sage.py" - let out ← IO.Process.output { cmd := "python3", args := #[path.toString] ++ args } +def runSage (trace : Bool) (α : Expr) (atoms : Nat) (hyps : Array (Source × Poly)) (target : Poly) : + CoreM SageResult := do + let query ← sageCreateQuery α atoms hyps target + if trace then + -- dry run enabled + return .ok { trace := query } + + -- send query to SageMath API + let apiUrl := "https://sagecell.sagemath.org/service" + let data := s!"code={System.Uri.escapeUri query}" + let curlArgs := #[ + "-X", "POST", + "--user-agent", polyrith.sageUserAgent.get (← getOptions), + "--data-raw", data, + apiUrl + ] + let out ← IO.Process.output { cmd := "curl", args := curlArgs } if out.exitCode != 0 then - throw <| IO.userError <| - s!"scripts/polyrith_sage.py exited with code {out.exitCode}:\n\n{out.stderr}" + IO.throwServerError <| + "Could not send API request to SageMath. " ++ + s!"curl exited with code {out.exitCode}:\n{out.stderr}" + + -- parse results match Json.parse out.stdout >>= fromJson? with - | .ok v => return v - | .error e => throw <| .userError e + | .ok result => return result + | .error e => IO.throwServerError <| + s!"Could not parse SageMath output (error: {e})\nSageMath output:\n{out.stdout}" + +end SageMathApi + + +/-! # Main function -/ /-- This is the main body of the `polyrith` tactic. It takes in the following inputs: @@ -319,11 +380,11 @@ collects all the relevant hypotheses/proof terms from the context, and from thos selected by the user, taking into account whether `only` is true. (The list of atoms is updated accordingly as well). -This information is used to create a list of args that get used in a call to -the appropriate python file that executes a grobner basis computation. The -output of this computation is a `String` representing the certificate. This -string is parsed into a list of `Poly` objects that are then converted into -`Expr`s (using the updated list of atoms). +This information is used to create an appropriate SageMath script that executes a +Groebner basis computation, which is sent to SageMath's API server. +The output of this computation is a JSON representing the certificate. +This JSON is parsed into the power of the goal and a list of `Poly` objects +that are then converted into `Expr`s (using the updated list of atoms). the names of the hypotheses, along with the corresponding coefficients are given to `linear_combination`. If that tactic succeeds, the user is prompted @@ -336,7 +397,7 @@ def polyrith (g : MVarId) (only : Bool) (hyps : Array Expr) (traceOnly := false) : MetaM (Except MVarId (TSyntax `tactic)) := do IO.sleep 10 -- otherwise can lead to weird errors when actively editing code with polyrith calls g.withContext <| AtomM.run .reducible do - let (α, hyps', tgt) ← parseContext only hyps (← g.getType) + let (α, hyps', target) ← parseContext only hyps (← g.getType) let rec /-- Try to prove the goal by `ring` and fail with the given message otherwise. -/ byRing msg := do @@ -348,7 +409,7 @@ def polyrith (g : MVarId) (only : Bool) (hyps : Array Expr) if hyps'.isEmpty then return ← byRing "polyrith did not find any relevant hypotheses" let vars := (← get).atoms.size - match ← sageOutput (createSageArgs traceOnly α vars hyps' tgt) with + match ← runSage traceOnly α vars hyps' target with | .ok { trace, data } => if let some trace := trace then logInfo trace if let some {coeffs := polys, power := pow} := data then @@ -390,8 +451,7 @@ Notes: * This tactic only works with a working internet connection, since it calls Sage using the SageCell web API at . Many thanks to the Sage team and organization for allowing this use. -* This tactic assumes that the user has `python3` installed and available on the path. - (Test by opening a terminal and executing `python3 --version`.) +* This tactic assumes that the user has `curl` available on path. Examples: @@ -405,7 +465,7 @@ example (x y z w : ℚ) (hzw : z = w) : x*z + 2*y*z = x*w + 2*y*w := by polyrith -- Try this: linear_combination (2 * y + x) * hzw -constant scary : ∀ a b : ℚ, a + b = 0 +axiom scary : ∀ a b : ℚ, a + b = 0 example (a b c d : ℚ) (h : a + b = 0) (h2: b + c = 0) : a + b + c + d = 0 := by polyrith only [scary c d, h] diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index c014fccd8e079..0fbeb2410e869 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -52,6 +52,10 @@ theorem continuousSMul_of_algebraMap [TopologicalSemiring A] (h : Continuous (al ContinuousSMul R A := ⟨(continuous_algebraMap_iff_smul R A).1 h⟩ +instance Subalgebra.continuousSMul (S : Subalgebra R A) (X) [TopologicalSpace X] [MulAction A X] + [ContinuousSMul A X] : ContinuousSMul S X := + Subsemiring.continuousSMul S.toSubsemiring X + section variable [ContinuousSMul R A] diff --git a/Mathlib/Topology/Algebra/Algebra/Equiv.lean b/Mathlib/Topology/Algebra/Algebra/Equiv.lean new file mode 100644 index 0000000000000..0a5464a1c04ae --- /dev/null +++ b/Mathlib/Topology/Algebra/Algebra/Equiv.lean @@ -0,0 +1,296 @@ +/- +Copyright (c) 2024 Salvatore Mercuri. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Salvatore Mercuri +-/ +import Mathlib.Topology.Algebra.Algebra + +/-! +# Isomorphisms of topological algebras + +This file contains an API for `ContinuousAlgEquiv R A B`, the type of +continuous `R`-algebra isomorphisms with continuous inverses. Here `R` is a +commutative (semi)ring, and `A` and `B` are `R`-algebras with topologies. + +## Main definitions + +Let `R` be a commutative semiring and let `A` and `B` be `R`-algebras which +are also topological spaces. + +* `ContinuousAlgEquiv R A B`: the type of continuous `R`-algebra isomorphisms + from `A` to `B` with continuous inverses. + +## Notations + +`A ≃A[R] B` : notation for `ContinuousAlgEquiv R A B`. + +## Tags + +* continuous, isomorphism, algebra +-/ + +open scoped Topology + + +/-- +`ContinuousAlgEquiv R A B`, with notation `A ≃A[R] B`, is the type of bijections +between the topological `R`-algebras `A` and `B` which are both homeomorphisms +and `R`-algebra isomorphisms. +-/ +structure ContinuousAlgEquiv (R A B : Type*) [CommSemiring R] + [Semiring A] [TopologicalSpace A] [Semiring B] [TopologicalSpace B] [Algebra R A] + [Algebra R B] extends A ≃ₐ[R] B, A ≃ₜ B + +@[inherit_doc] +notation:50 A " ≃A[" R "]" B => ContinuousAlgEquiv R A B + +attribute [nolint docBlame] ContinuousAlgEquiv.toHomeomorph + +/-- +`ContinuousAlgEquivClass F R A B` states that `F` is a type of topological algebra + structure-preserving equivalences. You should extend this class when you + extend `ContinuousAlgEquiv`. +-/ +class ContinuousAlgEquivClass (F : Type*) (R A B : outParam Type*) [CommSemiring R] + [Semiring A] [TopologicalSpace A] [Semiring B] [TopologicalSpace B] + [Algebra R A] [Algebra R B] [EquivLike F A B] + extends AlgEquivClass F R A B, HomeomorphClass F A B : Prop + +namespace ContinuousAlgEquiv + +variable {R A B C : Type*} + [CommSemiring R] [Semiring A] [TopologicalSpace A] [Semiring B] + [TopologicalSpace B] [Semiring C] [TopologicalSpace C] [Algebra R A] [Algebra R B] + [Algebra R C] + +/-- The natural coercion from a continuous algebra isomorphism to a continuous +algebra morphism. -/ +@[coe] +def toContinuousAlgHom (e : A ≃A[R] B) : A →A[R] B where + __ := e.toAlgHom + cont := e.continuous_toFun + +instance coe : Coe (A ≃A[R] B) (A →A[R] B) := ⟨toContinuousAlgHom⟩ + +instance equivLike : EquivLike (A ≃A[R] B) A B where + coe f := f.toFun + inv f := f.invFun + coe_injective' f g h₁ h₂ := by + cases' f with f' _ + cases' g with g' _ + rcases f' with ⟨⟨_, _⟩, _⟩ + rcases g' with ⟨⟨_, _⟩, _⟩ + congr + left_inv f := f.left_inv + right_inv f := f.right_inv + +instance continuousAlgEquivClass : ContinuousAlgEquivClass (A ≃A[R] B) R A B where + map_add f := f.map_add' + map_mul f := f.map_mul' + commutes f := f.commutes' + map_continuous := continuous_toFun + inv_continuous := continuous_invFun + +theorem coe_apply (e : A ≃A[R] B) (a : A) : (e : A →A[R] B) a = e a := rfl + +@[simp] +theorem coe_coe (e : A ≃A[R] B) : ⇑(e : A →A[R] B) = e := rfl + +theorem toAlgEquiv_injective : Function.Injective (toAlgEquiv : (A ≃A[R] B) → A ≃ₐ[R] B) := by + rintro ⟨e, _, _⟩ ⟨e', _, _⟩ rfl + rfl + +@[ext] +theorem ext {f g : A ≃A[R] B} (h : ⇑f = ⇑g) : f = g := + toAlgEquiv_injective <| AlgEquiv.ext <| congr_fun h + +theorem coe_injective : Function.Injective ((↑) : (A ≃A[R] B) → A →A[R] B) := + fun _ _ h => ext <| funext <| ContinuousAlgHom.ext_iff.1 h + +@[simp] +theorem coe_inj {f g : A ≃A[R] B} : (f : A →A[R] B) = g ↔ f = g := + coe_injective.eq_iff + +@[simp] +theorem coe_toAlgEquiv (e : A ≃A[R] B) : ⇑e.toAlgEquiv = e := rfl + +theorem isOpenMap (e : A ≃A[R] B) : IsOpenMap e := + e.toHomeomorph.isOpenMap + +theorem image_closure (e : A ≃A[R] B) (S : Set A) : e '' closure S = closure (e '' S) := + e.toHomeomorph.image_closure S + +theorem preimage_closure (e : A ≃A[R] B) (S : Set B) : e ⁻¹' closure S = closure (e ⁻¹' S) := + e.toHomeomorph.preimage_closure S + +@[simp] +theorem isClosed_image (e : A ≃A[R] B) {S : Set A} : IsClosed (e '' S) ↔ IsClosed S := + e.toHomeomorph.isClosed_image + +theorem map_nhds_eq (e : A ≃A[R] B) (a : A) : Filter.map e (𝓝 a) = 𝓝 (e a) := + e.toHomeomorph.map_nhds_eq a + +theorem map_eq_zero_iff (e : A ≃A[R] B) {a : A} : e a = 0 ↔ a = 0 := + e.toAlgEquiv.toLinearEquiv.map_eq_zero_iff + +attribute [continuity] + ContinuousAlgEquiv.continuous_invFun ContinuousAlgEquiv.continuous_toFun + +@[fun_prop] +theorem continuous (e : A ≃A[R] B) : Continuous e := e.continuous_toFun + +theorem continuousOn (e : A ≃A[R] B) {S : Set A} : ContinuousOn e S := + e.continuous.continuousOn + +theorem continuousAt (e : A ≃A[R] B) {a : A} : ContinuousAt e a := + e.continuous.continuousAt + +theorem continuousWithinAt (e : A ≃A[R] B) {S : Set A} {a : A} : + ContinuousWithinAt e S a := + e.continuous.continuousWithinAt + +theorem comp_continuous_iff {α : Type*} [TopologicalSpace α] (e : A ≃A[R] B) {f : α → A} : + Continuous (e ∘ f) ↔ Continuous f := + e.toHomeomorph.comp_continuous_iff + +theorem comp_continuous_iff' {β : Type*} [TopologicalSpace β] (e : A ≃A[R] B) {g : B → β} : + Continuous (g ∘ e) ↔ Continuous g := + e.toHomeomorph.comp_continuous_iff' + +variable (R A) + +/-- The identity isomorphism as a continuous `R`-algebra equivalence. -/ +@[refl] +def refl : A ≃A[R] A where + __ := AlgEquiv.refl + continuous_toFun := continuous_id + continuous_invFun := continuous_id + +@[simp] +theorem refl_apply (a : A) : refl R A a = a := rfl + +@[simp] +theorem coe_refl : refl R A = ContinuousAlgHom.id R A := rfl + +@[simp] +theorem coe_refl' : ⇑(refl R A) = id := rfl + +variable {R A} + +/-- The inverse of a continuous algebra equivalence. -/ +@[symm] +def symm (e : A ≃A[R] B) : B ≃A[R] A where + __ := e.toAlgEquiv.symm + continuous_toFun := e.continuous_invFun + continuous_invFun := e.continuous_toFun + +@[simp] +theorem apply_symm_apply (e : A ≃A[R] B) (b : B) : e (e.symm b) = b := + e.1.right_inv b + +@[simp] +theorem symm_apply_apply (e : A ≃A[R] B) (a : A) : e.symm (e a) = a := + e.1.left_inv a + +@[simp] +theorem symm_image_image (e : A ≃A[R] B) (S : Set A) : e.symm '' (e '' S) = S := + e.toEquiv.symm_image_image S + +@[simp] +theorem image_symm_image (e : A ≃A[R] B) (S : Set B) : e '' (e.symm '' S) = S := + e.symm.symm_image_image S + +@[simp] +theorem symm_toAlgEquiv (e : A ≃A[R] B) : e.symm.toAlgEquiv = e.toAlgEquiv.symm := rfl + +@[simp] +theorem symm_toHomeomorph (e : A ≃A[R] B) : e.symm.toHomeomorph = e.toHomeomorph.symm := rfl + +theorem symm_map_nhds_eq (e : A ≃A[R] B) (a : A) : Filter.map e.symm (𝓝 (e a)) = 𝓝 a := + e.toHomeomorph.symm_map_nhds_eq a + +/-- The composition of two continuous algebra equivalences. -/ +@[trans] +def trans (e₁ : A ≃A[R] B) (e₂ : B ≃A[R] C) : A ≃A[R] C where + __ := e₁.toAlgEquiv.trans e₂.toAlgEquiv + continuous_toFun := e₂.continuous_toFun.comp e₁.continuous_toFun + continuous_invFun := e₁.continuous_invFun.comp e₂.continuous_invFun + +@[simp] +theorem trans_toAlgEquiv (e₁ : A ≃A[R] B) (e₂ : B ≃A[R] C) : + (e₁.trans e₂).toAlgEquiv = e₁.toAlgEquiv.trans e₂.toAlgEquiv := + rfl + +@[simp] +theorem trans_apply (e₁ : A ≃A[R] B) (e₂ : B ≃A[R] C) (a : A) : + (e₁.trans e₂) a = e₂ (e₁ a) := + rfl + +@[simp] +theorem symm_trans_apply (e₁ : B ≃A[R] A) (e₂ : C ≃A[R] B) (a : A) : + (e₂.trans e₁).symm a = e₂.symm (e₁.symm a) := + rfl + +theorem comp_coe (e₁ : A ≃A[R] B) (e₂ : B ≃A[R] C) : + e₂.toAlgHom.comp e₁.toAlgHom = e₁.trans e₂ := by + rfl + +@[simp high] +theorem coe_comp_coe_symm (e : A ≃A[R] B) : + e.toContinuousAlgHom.comp e.symm = ContinuousAlgHom.id R B := + ContinuousAlgHom.ext e.apply_symm_apply + +@[simp high] +theorem coe_symm_comp_coe (e : A ≃A[R] B) : + e.symm.toContinuousAlgHom.comp e = ContinuousAlgHom.id R A := + ContinuousAlgHom.ext e.symm_apply_apply + +@[simp] +theorem symm_comp_self (e : A ≃A[R] B) : (e.symm : B → A) ∘ e = id := by + exact funext <| e.symm_apply_apply + +@[simp] +theorem self_comp_symm (e : A ≃A[R] B) : (e : A → B) ∘ e.symm = id := + funext <| e.apply_symm_apply + +@[simp] +theorem symm_symm (e : A ≃A[R] B) : e.symm.symm = e := rfl + +@[simp] +theorem refl_symm : (refl R A).symm = refl R A := rfl + +theorem symm_symm_apply (e : A ≃A[R] B) (a : A) : e.symm.symm a = e a := rfl + +theorem symm_apply_eq (e : A ≃A[R] B) {a : A} {b : B} : e.symm b = a ↔ b = e a := + e.toEquiv.symm_apply_eq + +theorem eq_symm_apply (e : A ≃A[R] B) {a : A} {b : B} : a = e.symm b ↔ e a = b := + e.toEquiv.eq_symm_apply + +theorem image_eq_preimage (e : A ≃A[R] B) (S : Set A) : e '' S = e.symm ⁻¹' S := + e.toEquiv.image_eq_preimage S + +theorem image_symm_eq_preimage (e : A ≃A[R] B) (S : Set B) : e.symm '' S = e ⁻¹' S := by + rw [e.symm.image_eq_preimage, e.symm_symm] + +@[simp] +theorem symm_preimage_preimage (e : A ≃A[R] B) (S : Set B) : e.symm ⁻¹' (e ⁻¹' S) = S := + e.toEquiv.symm_preimage_preimage S + +@[simp] +theorem preimage_symm_preimage (e : A ≃A[R] B) (S : Set A) : e ⁻¹' (e.symm ⁻¹' S) = S := + e.symm.symm_preimage_preimage S + +theorem isUniformEmbedding {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] [Ring E₁] + [UniformAddGroup E₁] [Algebra R E₁] [Ring E₂] [UniformAddGroup E₂] [Algebra R E₂] + (e : E₁ ≃A[R] E₂) : IsUniformEmbedding e := + e.toAlgEquiv.isUniformEmbedding e.toContinuousAlgHom.uniformContinuous + e.symm.toContinuousAlgHom.uniformContinuous + +theorem _root_.AlgEquiv.isUniformEmbedding {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] + [Ring E₁] [UniformAddGroup E₁] [Algebra R E₁] [Ring E₂] [UniformAddGroup E₂] [Algebra R E₂] + (e : E₁ ≃ₐ[R] E₂) (h₁ : Continuous e) (h₂ : Continuous e.symm) : + IsUniformEmbedding e := + ContinuousAlgEquiv.isUniformEmbedding { e with continuous_toFun := h₁ } + +end ContinuousAlgEquiv diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 8c4ddb97a123d..3809969b08829 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -108,6 +108,48 @@ instance instMonoidHomClass : MonoidHomClass (ContinuousMonoidHom A B) A B where instance instContinuousMapClass : ContinuousMapClass (ContinuousMonoidHom A B) A B where map_continuous f := f.continuous_toFun +@[to_additive (attr := simp)] +lemma coe_toMonoidHom (f : ContinuousMonoidHom A B) : f.toMonoidHom = f := rfl + +@[to_additive (attr := simp)] +lemma coe_toContinuousMap (f : ContinuousMonoidHom A B) : f.toContinuousMap = f := rfl + +section + +variable {A B : Type*} [Monoid A] [Monoid B] [TopologicalSpace A] [TopologicalSpace B] + {F : Type*} [FunLike F A B] + +/-- Turn an element of a type `F` satisfying `MonoidHomClass F A B` and `ContinuousMapClass F A B` +into a`ContinuousMonoidHom`. This is declared as the default coercion from `F` to +`ContinuousMonoidHom A B`. -/ +@[to_additive (attr := coe) "Turn an element of a type `F` satisfying +`AddMonoidHomClass F A B` and `ContinuousMapClass F A B` into a`ContinuousAddMonoidHom`. +This is declared as the default coercion from `F` to `ContinuousAddMonoidHom A B`."] +def toContinuousMonoidHom [MonoidHomClass F A B] [ContinuousMapClass F A B] (f : F) : + ContinuousMonoidHom A B := + { MonoidHomClass.toMonoidHom f with } + +/-- Any type satisfying `MonoidHomClass` and `ContinuousMapClass` can be cast into +`ContinuousMnoidHom` via `ContinuousMonoidHom.toContinuousMonoidHom`. -/ +@[to_additive "Any type satisfying `AddMonoidHomClass` and `ContinuousMapClass` can be cast into +`ContinuousAddMonoidHom` via `ContinuousAddMonoidHom.toContinuousAddMonoidHom`."] +instance [MonoidHomClass F A B] [ContinuousMapClass F A B] : CoeOut F (ContinuousMonoidHom A B) := + ⟨ContinuousMonoidHom.toContinuousMonoidHom⟩ + +@[to_additive (attr := simp)] +lemma coe_coe [MonoidHomClass F A B] [ContinuousMapClass F A B] (f : F) : + ⇑(f : ContinuousMonoidHom A B) = f := rfl + +@[to_additive (attr := simp, norm_cast)] +lemma toMonoidHom_toContinuousMonoidHom [MonoidHomClass F A B] [ContinuousMapClass F A B] (f : F) : + ((f : ContinuousMonoidHom A B) : A →* B) = f := rfl + +@[to_additive (attr := simp, norm_cast)] +lemma toContinuousMap_toContinuousMonoidHom [MonoidHomClass F A B] [ContinuousMapClass F A B] + (f : F) : ((f : ContinuousMonoidHom A B) : C(A, B)) = f := rfl + +end + @[to_additive (attr := ext)] theorem ext {f g : ContinuousMonoidHom A B} (h : ∀ x, f x = g x) : f = g := DFunLike.ext _ _ h diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index cce773d29a6f3..0181efb391afb 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -178,3 +178,13 @@ theorem IsPreconnected.eq_of_sq_eq [Field 𝕜] [HasContinuousInv₀ 𝕜] [Cont (iff_of_eq (iff_false _)).2 (hg_ne _)] at hy' ⊢ <;> assumption end Preconnected + +section ContinuousSMul + +variable {F : Type*} [DivisionRing F] [TopologicalSpace F] [TopologicalRing F] + (X : Type*) [TopologicalSpace X] [MulAction F X] [ContinuousSMul F X] + +instance Subfield.continuousSMul (M : Subfield F) : ContinuousSMul M X := + Subring.continuousSMul M.toSubring X + +end ContinuousSMul diff --git a/Mathlib/Topology/Algebra/Indicator.lean b/Mathlib/Topology/Algebra/Indicator.lean index 189c9d6f29038..87a704835566a 100644 --- a/Mathlib/Topology/Algebra/Indicator.lean +++ b/Mathlib/Topology/Algebra/Indicator.lean @@ -5,6 +5,7 @@ Authors: PFR contributors -/ import Mathlib.Algebra.Group.Indicator import Mathlib.Topology.ContinuousOn +import Mathlib.Topology.Clopen /-! # Continuity of indicator functions @@ -36,3 +37,8 @@ theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s) ⟨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⟩ + +@[to_additive] +lemma IsClopen.continuous_mulIndicator (hs : IsClopen s) (hf : Continuous f) : + Continuous (s.mulIndicator f) := + hf.mulIndicator (by simp [isClopen_iff_frontier_eq_empty.mp hs]) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean index 662314f0bb691..83fcac52c0419 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean @@ -426,11 +426,7 @@ theorem HasProd.nat_mul_neg {f : ℤ → M} (hf : HasProd f m) : · intro x hx simp only [u1, u2, mem_inter, mem_image, exists_prop] at hx suffices x = 0 by simp only [this, eq_self_iff_true, if_true] - apply le_antisymm - · rcases hx.2 with ⟨a, _, rfl⟩ - simp only [Right.neg_nonpos_iff, Nat.cast_nonneg] - · rcases hx.1 with ⟨a, _, rfl⟩ - simp only [Nat.cast_nonneg] + omega _ = (∏ x ∈ u1, f x) * ∏ x ∈ u2, f x := prod_union_inter _ = (∏ b ∈ v', f b) * ∏ b ∈ v', f (-b) := by simp only [u1, u2, Nat.cast_inj, imp_self, implies_true, forall_const, prod_image, neg_inj] diff --git a/Mathlib/Topology/Algebra/IntermediateField.lean b/Mathlib/Topology/Algebra/IntermediateField.lean new file mode 100644 index 0000000000000..17651f34142f9 --- /dev/null +++ b/Mathlib/Topology/Algebra/IntermediateField.lean @@ -0,0 +1,29 @@ +/- +Copyright (c) 2024 Jiedong Jiang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jiedong Jiang +-/ +import Mathlib.FieldTheory.IntermediateField.Adjoin.Defs +import Mathlib.Topology.Algebra.Field + +/-! +# Continuous actions related to intermediate fields + +In this file we define the instances related to continuous actions of +intermediate fields. The topology on intermediate fields is already defined +in earlier file `Mathlib.Topology.Algebra.Field` as the subspace topology. +-/ + +variable {K L : Type*} [Field K] [Field L] [Algebra K L] + [TopologicalSpace L] [TopologicalRing L] + +variable (X : Type*) [TopologicalSpace X] [MulAction L X] [ContinuousSMul L X] +variable (M : IntermediateField K L) + +instance IntermediateField.continuousSMul (M : IntermediateField K L) : ContinuousSMul M X := + M.toSubfield.continuousSMul X + +instance IntermediateField.botContinuousSMul (M : IntermediateField K L) : + ContinuousSMul (⊥ : IntermediateField K L) M := + Topology.IsInducing.continuousSMul (X := L) (N := (⊥ : IntermediateField K L)) (Y := M) + (M := (⊥ : IntermediateField K L)) Topology.IsInducing.subtypeVal continuous_id rfl diff --git a/Mathlib/Topology/Algebra/Module/Determinant.lean b/Mathlib/Topology/Algebra/Module/Determinant.lean index cf06e39304938..7ebf8ee44e47a 100644 --- a/Mathlib/Topology/Algebra/Module/Determinant.lean +++ b/Mathlib/Topology/Algebra/Module/Determinant.lean @@ -20,6 +20,12 @@ noncomputable abbrev det {R : Type*} [CommRing R] {M : Type*} [TopologicalSpace [Module R M] (A : M →L[R] M) : R := LinearMap.det (A : M →ₗ[R] M) +theorem det_pi {ι R M : Type*} [Fintype ι] [CommRing R] [AddCommGroup M] + [TopologicalSpace M] [Module R M] [Module.Free R M] [Module.Finite R M] + (f : ι → M →L[R] M) : + (pi (fun i ↦ (f i).comp (proj i))).det = ∏ i, (f i).det := + LinearMap.det_pi _ + end ContinuousLinearMap namespace ContinuousLinearEquiv diff --git a/Mathlib/Topology/Algebra/Module/Equiv.lean b/Mathlib/Topology/Algebra/Module/Equiv.lean index 8363cb3582442..9df3e155b17c7 100644 --- a/Mathlib/Topology/Algebra/Module/Equiv.lean +++ b/Mathlib/Topology/Algebra/Module/Equiv.lean @@ -30,7 +30,6 @@ variable (R M) /-- Continuous linear equivalences between modules. We only put the type classes that are necessary for the definition, although in applications `M` and `M₂` will be topological modules over the topological semiring `R`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet; was @[nolint has_nonempty_instance] structure ContinuousLinearEquiv {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type*) [TopologicalSpace M] [AddCommMonoid M] (M₂ : Type*) [TopologicalSpace M₂] [AddCommMonoid M₂] [Module R M] diff --git a/Mathlib/Topology/Algebra/Module/Star.lean b/Mathlib/Topology/Algebra/Module/Star.lean index 11d8032d76fb5..43a4ccb76dd08 100644 --- a/Mathlib/Topology/Algebra/Module/Star.lean +++ b/Mathlib/Topology/Algebra/Module/Star.lean @@ -65,16 +65,12 @@ theorem continuous_decomposeProdAdjoint_symm [TopologicalAddGroup A] : (continuous_subtype_val.comp continuous_fst).add (continuous_subtype_val.comp continuous_snd) /-- The self-adjoint part of an element of a star module, as a continuous linear map. -/ -@[simps!] +@[simps! -isSimp] def selfAdjointPartL [ContinuousAdd A] [ContinuousStar A] [ContinuousConstSMul R A] : A →L[R] selfAdjoint A where toLinearMap := selfAdjointPart R cont := continuous_selfAdjointPart _ _ --- This lemma is generated by `simps` but we don't actually want the `@[simp]` attribute. --- https://github.com/leanprover-community/mathlib4/issues/18942 -attribute [nolint simpNF] selfAdjointPartL_apply_coe - /-- The skew-adjoint part of an element of a star module, as a continuous linear map. -/ @[simps!] def skewAdjointPartL [ContinuousSub A] [ContinuousStar A] [ContinuousConstSMul R A] : diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index f136b5fd39568..025c3e373cc41 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -9,6 +9,7 @@ import Mathlib.Topology.Algebra.MulAction import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.Group.ULift import Mathlib.Topology.ContinuousMap.Defs +import Mathlib.Algebra.Group.Submonoid.Basic /-! # Theory of topological monoids @@ -567,6 +568,12 @@ theorem exists_open_nhds_one_mul_subset {U : Set M} (hU : U ∈ 𝓝 (1 : M)) : ∃ V : Set M, IsOpen V ∧ (1 : M) ∈ V ∧ V * V ⊆ U := by simpa only [mul_subset_iff] using exists_open_nhds_one_split hU +@[to_additive] +theorem Filter.HasBasis.mul_self {p : ι → Prop} {s : ι → Set M} (h : (𝓝 1).HasBasis p s) : + (𝓝 1).HasBasis p fun i => s i * s i := by + rw [← nhds_mul_nhds_one, ← map₂_mul, ← map_uncurry_prod] + simpa only [← image_mul_prod] using h.prod_self.map _ + end MulOneClass section ContinuousMul diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index 9f2b202d56789..5be53f4ed4e3c 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -3,12 +3,8 @@ 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, Yury Kudryashov, Yaël Dillies -/ -import Mathlib.Algebra.BigOperators.Intervals -import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.Group.DenselyOrdered -import Mathlib.Algebra.Order.Group.Indicator import Mathlib.Data.Real.Archimedean -import Mathlib.Order.Filter.AtTopBot.Archimedean import Mathlib.Order.LiminfLimsup import Mathlib.Topology.Algebra.Group.Basic import Mathlib.Topology.Order.Monotone @@ -409,78 +405,6 @@ theorem Monotone.map_liminf_of_continuousAt {f : R → S} (f_incr : Monotone f) end Monotone -section Indicator - -theorem limsup_eq_tendsto_sum_indicator_nat_atTop (s : ℕ → Set α) : - limsup s atTop = { ω | Tendsto - (fun n ↦ ∑ k ∈ Finset.range n, (s (k + 1)).indicator (1 : α → ℕ) ω) atTop atTop } := by - ext ω - simp only [limsup_eq_iInf_iSup_of_nat, Set.iSup_eq_iUnion, Set.iInf_eq_iInter, - Set.mem_iInter, Set.mem_iUnion, exists_prop] - constructor - · intro hω - refine tendsto_atTop_atTop_of_monotone' (fun n m hnm ↦ Finset.sum_mono_set_of_nonneg - (fun i ↦ Set.indicator_nonneg (fun _ _ ↦ zero_le_one) _) (Finset.range_mono hnm)) ?_ - rintro ⟨i, h⟩ - simp only [mem_upperBounds, Set.mem_range, forall_exists_index, forall_apply_eq_imp_iff] at h - induction' i with k hk - · obtain ⟨j, hj₁, hj₂⟩ := hω 1 - refine not_lt.2 (h <| j + 1) - (lt_of_le_of_lt (Finset.sum_const_zero.symm : 0 = ∑ k ∈ Finset.range (j + 1), 0).le ?_) - refine Finset.sum_lt_sum (fun m _ ↦ Set.indicator_nonneg (fun _ _ ↦ zero_le_one) _) - ⟨j - 1, Finset.mem_range.2 (lt_of_le_of_lt (Nat.sub_le _ _) j.lt_succ_self), ?_⟩ - rw [Nat.sub_add_cancel hj₁, Set.indicator_of_mem hj₂] - exact zero_lt_one - · rw [imp_false] at hk - push_neg at hk - obtain ⟨i, hi⟩ := hk - obtain ⟨j, hj₁, hj₂⟩ := hω (i + 1) - replace hi : (∑ k ∈ Finset.range i, (s (k + 1)).indicator 1 ω) = k + 1 := - le_antisymm (h i) hi - refine not_lt.2 (h <| j + 1) ?_ - rw [← Finset.sum_range_add_sum_Ico _ (i.le_succ.trans (hj₁.trans j.le_succ)), hi] - refine lt_add_of_pos_right _ ?_ - rw [(Finset.sum_const_zero.symm : 0 = ∑ k ∈ Finset.Ico i (j + 1), 0)] - refine Finset.sum_lt_sum (fun m _ ↦ Set.indicator_nonneg (fun _ _ ↦ zero_le_one) _) - ⟨j - 1, Finset.mem_Ico.2 ⟨(Nat.le_sub_iff_add_le (le_trans ((le_add_iff_nonneg_left _).2 - zero_le') hj₁)).2 hj₁, lt_of_le_of_lt (Nat.sub_le _ _) j.lt_succ_self⟩, ?_⟩ - rw [Nat.sub_add_cancel (le_trans ((le_add_iff_nonneg_left _).2 zero_le') hj₁), - Set.indicator_of_mem hj₂] - exact zero_lt_one - · rintro hω i - rw [Set.mem_setOf_eq, tendsto_atTop_atTop] at hω - by_contra! hcon - obtain ⟨j, h⟩ := hω (i + 1) - have : (∑ k ∈ Finset.range j, (s (k + 1)).indicator 1 ω) ≤ i := by - have hle : ∀ j ≤ i, (∑ k ∈ Finset.range j, (s (k + 1)).indicator 1 ω) ≤ i := by - refine fun j hij ↦ - (Finset.sum_le_card_nsmul _ _ _ ?_ : _ ≤ (Finset.range j).card • 1).trans ?_ - · exact fun m _ ↦ Set.indicator_apply_le' (fun _ ↦ le_rfl) fun _ ↦ zero_le_one - · simpa only [Finset.card_range, smul_eq_mul, mul_one] - by_cases hij : j < i - · exact hle _ hij.le - · rw [← Finset.sum_range_add_sum_Ico _ (not_lt.1 hij)] - suffices (∑ k ∈ Finset.Ico i j, (s (k + 1)).indicator 1 ω) = 0 by - rw [this, add_zero] - exact hle _ le_rfl - refine Finset.sum_eq_zero fun m hm ↦ ?_ - exact Set.indicator_of_not_mem (hcon _ <| (Finset.mem_Ico.1 hm).1.trans m.le_succ) _ - exact not_le.2 (lt_of_lt_of_le i.lt_succ_self <| h _ le_rfl) this - -theorem limsup_eq_tendsto_sum_indicator_atTop (R : Type*) [StrictOrderedSemiring R] [Archimedean R] - (s : ℕ → Set α) : limsup s atTop = { ω | Tendsto - (fun n ↦ ∑ k ∈ Finset.range n, (s (k + 1)).indicator (1 : α → R) ω) atTop atTop } := by - rw [limsup_eq_tendsto_sum_indicator_nat_atTop s] - ext ω - simp only [Set.mem_setOf_eq] - rw [(_ : (fun n ↦ ∑ k ∈ Finset.range n, (s (k + 1)).indicator (1 : α → R) ω) = fun n ↦ - ↑(∑ k ∈ Finset.range n, (s (k + 1)).indicator (1 : α → ℕ) ω))] - · exact tendsto_natCast_atTop_iff.symm - · ext n - simp only [Set.indicator, Pi.one_apply, Finset.sum_boole, Nat.cast_id] - -end Indicator - section LiminfLimsupAdd variable [AddCommGroup α] [ConditionallyCompleteLinearOrder α] [DenselyOrdered α] diff --git a/Mathlib/Topology/Algebra/Ring/Basic.lean b/Mathlib/Topology/Algebra/Ring/Basic.lean index f565a4012f388..613f1181d73b9 100644 --- a/Mathlib/Topology/Algebra/Ring/Basic.lean +++ b/Mathlib/Topology/Algebra/Ring/Basic.lean @@ -130,6 +130,10 @@ namespace Subsemiring instance topologicalSemiring (S : Subsemiring α) : TopologicalSemiring S := { S.toSubmonoid.continuousMul, S.toAddSubmonoid.continuousAdd with } +instance continuousSMul (s : Subsemiring α) (X) [TopologicalSpace X] [MulAction α X] + [ContinuousSMul α X] : ContinuousSMul s X := + Submonoid.continuousSMul + end Subsemiring /-- The (topological-space) closure of a subsemiring of a topological semiring is @@ -308,6 +312,10 @@ variable [Ring α] [TopologicalRing α] instance Subring.instTopologicalRing (S : Subring α) : TopologicalRing S := { S.toSubmonoid.continuousMul, inferInstanceAs (TopologicalAddGroup S.toAddSubgroup) with } +instance Subring.continuousSMul (s : Subring α) (X) [TopologicalSpace X] [MulAction α X] + [ContinuousSMul α X] : ContinuousSMul s X := + Subsemiring.continuousSMul s.toSubsemiring X + /-- The (topological-space) closure of a subring of a topological ring is itself a subring. -/ def Subring.topologicalClosure (S : Subring α) : Subring α := diff --git a/Mathlib/Topology/Algebra/TopologicallyNilpotent.lean b/Mathlib/Topology/Algebra/TopologicallyNilpotent.lean new file mode 100644 index 0000000000000..e9a9dea659cfa --- /dev/null +++ b/Mathlib/Topology/Algebra/TopologicallyNilpotent.lean @@ -0,0 +1,130 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández +-/ +import Mathlib.Topology.Algebra.LinearTopology +import Mathlib.RingTheory.Ideal.Basic + +/-! # Topologically nilpotent elements + +Let `M` be a monoid with zero `M`, endowed with a topology. + +* `IsTopologicallyNilpotent a` says that `a : M` is *topologically nilpotent*, +ie, its powers converge to zero. + +* `IsTopologicallyNilpotent.map`: +The image of a topologically nilpotent element under a continuous morphism of +monoids with zero endowed with a topology is topologically nilpotent. + +* `IsTopologicallyNilpotent.zero`: `0` is topologically nilpotent. + +Let `R` be a commutative ring with a linear topology. + +* `IsTopologicallyNilpotent.mul_left`: if `a : R` is topologically nilpotent, +then `a*b` is topologically nilpotent. + +* `IsTopologicallyNilpotent.mul_right`: if `a : R` is topologically nilpotent, +then `a * b` is topologically nilpotent. + +* `IsTopologicallyNilpotent.add`: if `a b : R` are topologically nilpotent, +then `a + b` is topologically nilpotent. + +These lemmas are actually deduced from their analogues for commuting elements of rings. + +-/ + +open Filter + +open scoped Topology + +/-- An element is topologically nilpotent if its powers converge to `0`. -/ +def IsTopologicallyNilpotent + {R : Type*} [MonoidWithZero R] [TopologicalSpace R] (a : R) : Prop := + Tendsto (a ^ ·) atTop (𝓝 0) + +namespace IsTopologicallyNilpotent + +section MonoidWithZero + +variable {R S : Type*} [TopologicalSpace R] [MonoidWithZero R] + [MonoidWithZero S] [TopologicalSpace S] + +/-- The image of a topologically nilpotent element under a continuous morphism + is topologically nilpotent -/ +theorem map {F : Type*} [FunLike F R S] [MonoidWithZeroHomClass F R S] + {φ : F} (hφ : Continuous φ) {a : R} (ha : IsTopologicallyNilpotent a) : + IsTopologicallyNilpotent (φ a) := by + unfold IsTopologicallyNilpotent at ha ⊢ + simp_rw [← map_pow] + exact (map_zero φ ▸ hφ.tendsto 0).comp ha + +/-- `0` is topologically nilpotent -/ +theorem zero : IsTopologicallyNilpotent (0 : R) := + tendsto_atTop_of_eventually_const (i₀ := 1) + (fun _ hi => by rw [zero_pow (Nat.ne_zero_iff_zero_lt.mpr hi)]) + +theorem exists_pow_mem_of_mem_nhds {a : R} (ha : IsTopologicallyNilpotent a) + {v : Set R} (hv : v ∈ 𝓝 0) : + ∃ n, a ^ n ∈ v := + (ha.eventually_mem hv).exists + +end MonoidWithZero + +section Ring + +variable {R : Type*} [TopologicalSpace R] [Ring R] + +/-- If `a` and `b` commute and `a` is topologically nilpotent, + then `a * b` is topologically nilpotent. -/ +theorem mul_right_of_commute [IsLinearTopology Rᵐᵒᵖ R] + {a b : R} (ha : IsTopologicallyNilpotent a) (hab : Commute a b) : + IsTopologicallyNilpotent (a * b) := by + simp_rw [IsTopologicallyNilpotent, hab.mul_pow] + exact IsLinearTopology.tendsto_mul_zero_of_left _ _ ha + +/-- If `a` and `b` commute and `b` is topologically nilpotent, + then `a * b` is topologically nilpotent. -/ + theorem mul_left_of_commute [IsLinearTopology R R] {a b : R} + (hb : IsTopologicallyNilpotent b) (hab : Commute a b) : + IsTopologicallyNilpotent (a * b) := by + simp_rw [IsTopologicallyNilpotent, hab.mul_pow] + exact IsLinearTopology.tendsto_mul_zero_of_right _ _ hb + +/-- If `a` and `b` are topologically nilpotent and commute, + then `a + b` is topologically nilpotent. -/ +theorem add_of_commute [IsLinearTopology R R] {a b : R} + (ha : IsTopologicallyNilpotent a) (hb : IsTopologicallyNilpotent b) (h : Commute a b) : + IsTopologicallyNilpotent (a + b) := by + simp only [IsTopologicallyNilpotent, atTop_basis.tendsto_iff IsLinearTopology.hasBasis_ideal, + true_and] + intro I I_mem_nhds + obtain ⟨na, ha⟩ := ha.exists_pow_mem_of_mem_nhds I_mem_nhds + obtain ⟨nb, hb⟩ := hb.exists_pow_mem_of_mem_nhds I_mem_nhds + exact ⟨na + nb, fun m hm ↦ + I.add_pow_mem_of_pow_mem_of_le_of_commute ha hb (le_trans hm (Nat.le_add_right _ _)) h⟩ + +end Ring + +section CommRing + +variable {R : Type*} [TopologicalSpace R] [CommRing R] [IsLinearTopology R R] + +/-- If `a` is topologically nilpotent, then `a * b` is topologically nilpotent. -/ +theorem mul_right {a : R} (ha : IsTopologicallyNilpotent a) (b : R) : + IsTopologicallyNilpotent (a * b) := + ha.mul_right_of_commute (Commute.all ..) + +/-- If `b` is topologically nilpotent, then `a * b` is topologically nilpotent. -/ + theorem mul_left (a : R) {b : R} (hb : IsTopologicallyNilpotent b) : + IsTopologicallyNilpotent (a * b) := + hb.mul_left_of_commute (Commute.all ..) + +/-- If `a` and `b` are topologically nilpotent, then `a + b` is topologically nilpotent. -/ +theorem add {a b : R} (ha : IsTopologicallyNilpotent a) (hb : IsTopologicallyNilpotent b) : + IsTopologicallyNilpotent (a + b) := + ha.add_of_commute hb (Commute.all ..) + +end CommRing + +end IsTopologicallyNilpotent diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index d294240517a43..7644318842d13 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -100,15 +100,11 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `LightProfinite`, endowing a finite type with the discrete topology. -/ -@[simps! map_apply] +@[simps! -isSimp map_apply] def FintypeCat.toLightProfinite : FintypeCat ⥤ LightProfinite where obj A := LightProfinite.of A map f := ⟨f, by continuity⟩ --- This lemma is generated by `simps` but we don't actually want the `@[simp]` attribute. --- https://github.com/leanprover-community/mathlib4/issues/18942 -attribute [nolint simpNF] FintypeCat.toLightProfinite_map_apply - /-- `FintypeCat.toLightProfinite` is fully faithful. -/ def FintypeCat.toLightProfiniteFullyFaithful : toLightProfinite.FullyFaithful where preimage f := (f : _ → _) diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index fe5f7d2a47ee4..e0da94050bcc9 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -96,9 +96,8 @@ section Profinite -- unhelpfully defines a function `CompHaus.{max u₁ u₂} → Profinite.{max u₁ u₂}`. /-- (Implementation) The object part of the connected_components functor from compact Hausdorff spaces -to Profinite spaces, given by quotienting a space by its connected components. -See: https://stacks.math.columbia.edu/tag/0900 --/ +to Profinite spaces, given by quotienting a space by its connected components. -/ +@[stacks 0900] def CompHaus.toProfiniteObj (X : CompHaus.{u}) : Profinite.{u} where toTop := TopCat.of (ConnectedComponents X) is_compact := Quotient.compactSpace @@ -141,15 +140,11 @@ attribute [local instance] FintypeCat.discreteTopology /-- The natural functor from `Fintype` to `Profinite`, endowing a finite type with the discrete topology. -/ -@[simps map_apply] +@[simps -isSimp map_apply] def FintypeCat.toProfinite : FintypeCat ⥤ Profinite where obj A := Profinite.of A map f := ⟨f, by continuity⟩ --- This lemma is generated by `simps` but we don't actually want the `@[simp]` attribute. --- https://github.com/leanprover-community/mathlib4/issues/18942 -attribute [nolint simpNF] FintypeCat.toProfinite_map_apply - /-- `FintypeCat.toLightProfinite` is fully faithful. -/ def FintypeCat.toProfiniteFullyFaithful : toProfinite.FullyFaithful where preimage f := (f : _ → _) diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 47f59cf4a2bd6..ca9630572492f 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -67,7 +67,6 @@ variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {x y z : X} {ι /-! ### Paths -/ /-- Continuous path connecting two points `x` and `y` in a topological space -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] structure Path (x y : X) extends C(I, X) where /-- The start point of a `Path`. -/ source' : toFun 0 = x diff --git a/Mathlib/Topology/Constructible.lean b/Mathlib/Topology/Constructible.lean new file mode 100644 index 0000000000000..4ad63357cdd71 --- /dev/null +++ b/Mathlib/Topology/Constructible.lean @@ -0,0 +1,433 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Order.BooleanSubalgebra +import Mathlib.Topology.Spectral.Hom + +/-! +# Constructible sets + +This file defines constructible sets, which are morally sets in a topological space which we can +make out of finite unions and intersections of open and closed sets. + +Precisely, constructible sets are the boolean subalgebra generated by open retrocompact sets, +where a set is retrocompact if its intersection with every compact open set is compact. +In a locally noetherian space, all sets are retrocompact, in which case this boolean subalgebra is +simply the one generated by the open sets. + +Constructible sets are useful because the image of a constructible set under a finitely presented +morphism of schemes is a constructible set (and this is *not* true at the level of varieties). + +## Main declarations + +* `IsRetrocompact`: Predicate for a set to be retrocompact, namely to have its intersection with + every compact open be compact. +* `IsConstructible`: Predicate for a set to be constructible, namely to belong to the boolean + subalgebra generated by open retrocompact sets. +* `IsLocallyConstructible`: Predicate for a set to be locally constructible, namely to be + partitionable along an open cover such that each of its parts is constructible in the + respective open subspace. +-/ + +open Set TopologicalSpace Topology +open scoped Set.Notation + +variable {ι : Sort*} {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : X → Y} + {s t U : Set X} {a : X} + +/-! ### retrocompact sets -/ + +/-- A retrocompact set is a set whose intersection with every compact open is compact. -/ +@[stacks 005A] +def IsRetrocompact (s : Set X) : Prop := ∀ ⦃U⦄, IsCompact U → IsOpen U → IsCompact (s ∩ U) + +@[simp] lemma IsRetrocompact.empty : IsRetrocompact (∅ : Set X) := by simp [IsRetrocompact] +@[simp] lemma IsRetrocompact.univ : IsRetrocompact (univ : Set X) := by + simp +contextual [IsRetrocompact] + +@[simp] lemma IsRetrocompact.singleton : IsRetrocompact {a} := + fun _ _ _ ↦ Subsingleton.singleton_inter.isCompact + +lemma IsRetrocompact.union (hs : IsRetrocompact s) (ht : IsRetrocompact t) : + IsRetrocompact (s ∪ t : Set X) := + fun _U hUcomp hUopen ↦ union_inter_distrib_right .. ▸ (hs hUcomp hUopen).union (ht hUcomp hUopen) + +private lemma supClosed_isRetrocompact : SupClosed {s : Set X | IsRetrocompact s} := + fun _s hs _t ht ↦ hs.union ht + +lemma IsRetrocompact.finsetSup {ι : Type*} {s : Finset ι} {t : ι → Set X} + (ht : ∀ i ∈ s, IsRetrocompact (t i)) : IsRetrocompact (s.sup t) := by + induction' s using Finset.cons_induction with i s ih hi + · simp + · rw [Finset.sup_cons] + exact (ht _ <| by simp).union <| hi <| Finset.forall_of_forall_cons ht + +set_option linter.docPrime false in +lemma IsRetrocompact.finsetSup' {ι : Type*} {s : Finset ι} {hs} {t : ι → Set X} + (ht : ∀ i ∈ s, IsRetrocompact (t i)) : IsRetrocompact (s.sup' hs t) := by + rw [Finset.sup'_eq_sup]; exact .finsetSup ht + +lemma IsRetrocompact.iUnion [Finite ι] {f : ι → Set X} (hf : ∀ i, IsRetrocompact (f i)) : + IsRetrocompact (⋃ i, f i) := supClosed_isRetrocompact.iSup_mem .empty hf + +lemma IsRetrocompact.sUnion {S : Set (Set X)} (hS : S.Finite) (hS' : ∀ s ∈ S, IsRetrocompact s) : + IsRetrocompact (⋃₀ S) := supClosed_isRetrocompact.sSup_mem hS .empty hS' + +lemma IsRetrocompact.biUnion {ι : Type*} {f : ι → Set X} {t : Set ι} (ht : t.Finite) + (hf : ∀ i ∈ t, IsRetrocompact (f i)) : IsRetrocompact (⋃ i ∈ t, f i) := + supClosed_isRetrocompact.biSup_mem ht .empty hf + +section T2Space +variable [T2Space X] + +lemma IsRetrocompact.inter (hs : IsRetrocompact s) (ht : IsRetrocompact t) : + IsRetrocompact (s ∩ t : Set X) := + fun _U hUcomp hUopen ↦ inter_inter_distrib_right .. ▸ (hs hUcomp hUopen).inter (ht hUcomp hUopen) + +private lemma infClosed_isRetrocompact : InfClosed {s : Set X | IsRetrocompact s} := + fun _s hs _t ht ↦ hs.inter ht + +lemma IsRetrocompact.finsetInf {ι : Type*} {s : Finset ι} {t : ι → Set X} + (ht : ∀ i ∈ s, IsRetrocompact (t i)) : IsRetrocompact (s.inf t) := by + induction' s using Finset.cons_induction with i s ih hi + · simp + · rw [Finset.inf_cons] + exact (ht _ <| by simp).inter <| hi <| Finset.forall_of_forall_cons ht + +set_option linter.docPrime false in +lemma IsRetrocompact.finsetInf' {ι : Type*} {s : Finset ι} {hs} {t : ι → Set X} + (ht : ∀ i ∈ s, IsRetrocompact (t i)) : IsRetrocompact (s.inf' hs t) := by + rw [Finset.inf'_eq_inf]; exact .finsetInf ht + +lemma IsRetrocompact.iInter [Finite ι] {f : ι → Set X} (hf : ∀ i, IsRetrocompact (f i)) : + IsRetrocompact (⋂ i, f i) := infClosed_isRetrocompact.iInf_mem .univ hf + +lemma IsRetrocompact.sInter {S : Set (Set X)} (hS : S.Finite) (hS' : ∀ s ∈ S, IsRetrocompact s) : + IsRetrocompact (⋂₀ S) := infClosed_isRetrocompact.sInf_mem hS .univ hS' + +lemma IsRetrocompact.biInter {ι : Type*} {f : ι → Set X} {t : Set ι} (ht : t.Finite) + (hf : ∀ i ∈ t, IsRetrocompact (f i)) : IsRetrocompact (⋂ i ∈ t, f i) := + infClosed_isRetrocompact.biInf_mem ht .univ hf + +end T2Space + +lemma IsRetrocompact.inter_isOpen (hs : IsRetrocompact s) (ht : IsRetrocompact t) + (htopen : IsOpen t) : IsRetrocompact (s ∩ t : Set X) := + fun _U hUcomp hUopen ↦ inter_assoc .. ▸ hs (ht hUcomp hUopen) (htopen.inter hUopen) + +lemma IsRetrocompact.isOpen_inter (hs : IsRetrocompact s) (ht : IsRetrocompact t) + (hsopen : IsOpen s) : IsRetrocompact (s ∩ t : Set X) := + inter_comm .. ▸ ht.inter_isOpen hs hsopen + +lemma IsRetrocompact_iff_isSpectralMap_subtypeVal : + IsRetrocompact s ↔ IsSpectralMap (Subtype.val : s → X) := by + refine ⟨fun hs ↦ ⟨continuous_subtype_val, fun t htopen htcomp ↦ ?_⟩, fun hs t htcomp htopen ↦ ?_⟩ + · rw [IsEmbedding.subtypeVal.isCompact_iff, image_preimage_eq_inter_range, + Subtype.range_coe_subtype, setOf_mem_eq, inter_comm] + exact hs htcomp htopen + · simpa using (hs.isCompact_preimage_of_isOpen htopen htcomp).image continuous_subtype_val + +@[stacks 005B] +lemma IsRetrocompact.image_of_isEmbedding (hs : IsRetrocompact s) (hfemb : IsEmbedding f) + (hfcomp : IsRetrocompact (range f)) : IsRetrocompact (f '' s) := by + rintro U hUcomp hUopen + rw [← image_inter_preimage, ← hfemb.isCompact_iff] + refine hs ?_ <| hUopen.preimage hfemb.continuous + rw [hfemb.isCompact_iff, image_preimage_eq_inter_range, inter_comm] + exact hfcomp hUcomp hUopen + +@[stacks 005J "Extracted from the proof"] +lemma IsRetrocompact.preimage_of_isOpenEmbedding {s : Set Y} (hf : IsOpenEmbedding f) + (hs : IsRetrocompact s) : IsRetrocompact (f ⁻¹' s) := by + rintro U hUcomp hUopen + rw [hf.isCompact_iff, image_preimage_inter] + exact hs (hUcomp.image hf.continuous) <| hf.isOpenMap _ hUopen + +@[stacks 09YE "Extracted from the proof"] +lemma IsRetrocompact.preimage_of_isClosedEmbedding {s : Set Y} (hf : IsClosedEmbedding f) + (hf' : IsCompact (range f)ᶜ) (hs : IsRetrocompact s) : IsRetrocompact (f ⁻¹' s) := by + rintro U hUcomp hUopen + have hfUopen : IsOpen (f '' U ∪ (range f)ᶜ) := by + simpa [← range_diff_image hf.injective, sdiff_eq, compl_inter, union_comm] + using (hf.isClosedMap _ hUopen.isClosed_compl).isOpen_compl + have hfUcomp : IsCompact (f '' U ∪ (range f)ᶜ) := (hUcomp.image hf.continuous).union hf' + simpa [inter_union_distrib_left, inter_left_comm, inter_eq_right.2 (image_subset_range ..), + hf.isCompact_iff, image_preimage_inter] using (hs hfUcomp hfUopen).inter_left hf.isClosed_range + +/-! ### Constructible sets -/ + +namespace Topology + +/-- A constructible set is a set that can be written as the +finite union/finite intersection/complement of open retrocompact sets. + +In other words, constructible sets form the boolean subalgebra generated by open retrocompact sets. +-/ +def IsConstructible (s : Set X) : Prop := + s ∈ BooleanSubalgebra.closure {U | IsOpen U ∧ IsRetrocompact U} + +@[simp] +protected lemma IsConstructible.empty : IsConstructible (∅ : Set X) := BooleanSubalgebra.bot_mem + +@[simp] +protected lemma IsConstructible.univ : IsConstructible (univ : Set X) := BooleanSubalgebra.top_mem + +lemma IsConstructible.union : IsConstructible s → IsConstructible t → IsConstructible (s ∪ t) := + BooleanSubalgebra.sup_mem + +lemma IsConstructible.inter : IsConstructible s → IsConstructible t → IsConstructible (s ∩ t) := + BooleanSubalgebra.inf_mem + +lemma IsConstructible.sdiff : IsConstructible s → IsConstructible t → IsConstructible (s \ t) := + BooleanSubalgebra.sdiff_mem + +lemma IsConstructible.himp : IsConstructible s → IsConstructible t → IsConstructible (s ⇨ t) := + BooleanSubalgebra.himp_mem + +@[simp] lemma isConstructible_compl : IsConstructible sᶜ ↔ IsConstructible s := + BooleanSubalgebra.compl_mem_iff + +alias ⟨IsConstructible.of_compl, IsConstructible.compl⟩ := isConstructible_compl + +lemma IsConstructible.iUnion [Finite ι] {f : ι → Set X} (hf : ∀ i, IsConstructible (f i)) : + IsConstructible (⋃ i, f i) := BooleanSubalgebra.iSup_mem hf + +lemma IsConstructible.iInter [Finite ι] {f : ι → Set X} (hf : ∀ i, IsConstructible (f i)) : + IsConstructible (⋂ i, f i) := BooleanSubalgebra.iInf_mem hf + +lemma IsConstructible.sUnion {S : Set (Set X)} (hS : S.Finite) (hS' : ∀ s ∈ S, IsConstructible s) : + IsConstructible (⋃₀ S) := BooleanSubalgebra.sSup_mem hS hS' + +lemma IsConstructible.sInter {S : Set (Set X)} (hS : S.Finite) (hS' : ∀ s ∈ S, IsConstructible s) : + IsConstructible (⋂₀ S) := BooleanSubalgebra.sInf_mem hS hS' + +lemma IsConstructible.biUnion {ι : Type*} {f : ι → Set X} {t : Set ι} (ht : t.Finite) + (hf : ∀ i ∈ t, IsConstructible (f i)) : IsConstructible (⋃ i ∈ t, f i) := + BooleanSubalgebra.biSup_mem ht hf + +lemma IsConstructible.biInter {ι : Type*} {f : ι → Set X} {t : Set ι} (ht : t.Finite) + (hf : ∀ i ∈ t, IsConstructible (f i)) : IsConstructible (⋂ i ∈ t, f i) := + BooleanSubalgebra.biInf_mem ht hf + +lemma _root_.IsRetrocompact.isConstructible (hUopen : IsOpen U) (hUcomp : IsRetrocompact U) : + IsConstructible U := BooleanSubalgebra.subset_closure ⟨hUopen, hUcomp⟩ + +/-- An induction principle for constructible sets. If `p` holds for all open retrocompact +sets, and is preserved under union and complement, then `p` holds for all constructible sets. -/ +@[elab_as_elim] +lemma IsConstructible.empty_union_induction {p : ∀ s : Set X, IsConstructible s → Prop} + (open_retrocompact : ∀ U (hUopen : IsOpen U) (hUcomp : IsRetrocompact U), + p U (BooleanSubalgebra.subset_closure ⟨hUopen, hUcomp⟩)) + (union : ∀ s hs t ht, p s hs → p t ht → p (s ∪ t) (hs.union ht)) + (compl : ∀ s hs, p s hs → p sᶜ hs.compl) {s} (hs : IsConstructible s) : p s hs := by + induction hs using BooleanSubalgebra.closure_bot_sup_induction with + | mem U hU => exact open_retrocompact _ hU.1 hU.2 + | bot => exact open_retrocompact _ isOpen_empty .empty + | sup s hs t ht hs' ht' => exact union _ _ _ _ hs' ht' + | compl s hs hs' => exact compl _ _ hs' + +/-- If `f` is continuous and is such that preimages of retrocompact sets are retrocompact, then +preimages of constructible sets are constructible. -/ +@[stacks 005I] +lemma IsConstructible.preimage {s : Set Y} (hfcont : Continuous f) + (hf : ∀ s, IsRetrocompact s → IsRetrocompact (f ⁻¹' s)) (hs : IsConstructible s) : + IsConstructible (f ⁻¹' s) := by + induction hs using IsConstructible.empty_union_induction with + | open_retrocompact U hUopen hUcomp => + exact (hf _ hUcomp).isConstructible <| hUopen.preimage hfcont + | union s hs t ht hs' ht' => rw [preimage_union]; exact hs'.union ht' + | compl s hs hs' => rw [preimage_compl]; exact hs'.compl + +@[stacks 005J] +lemma IsConstructible.preimage_of_isOpenEmbedding {s : Set Y} (hf : IsOpenEmbedding f) + (hs : IsConstructible s) : IsConstructible (f ⁻¹' s) := + hs.preimage hf.continuous fun _t ht ↦ ht.preimage_of_isOpenEmbedding hf + +@[stacks 09YE] +lemma IsConstructible.preimage_of_isClosedEmbedding {s : Set Y} (hf : IsClosedEmbedding f) + (hf' : IsCompact (range f)ᶜ) (hs : IsConstructible s) : IsConstructible (f ⁻¹' s) := + hs.preimage hf.continuous fun _t ht ↦ ht.preimage_of_isClosedEmbedding hf hf' + +@[stacks 09YD] +lemma IsConstructible.image_of_isOpenEmbedding (hfopen : IsOpenEmbedding f) + (hfcomp : IsRetrocompact (range f)) (hs : IsConstructible s) : IsConstructible (f '' s) := by + induction hs using IsConstructible.empty_union_induction with + | open_retrocompact U hUopen hUcomp => + exact (hUcomp.image_of_isEmbedding hfopen.isEmbedding hfcomp).isConstructible <| + hfopen.isOpenMap _ hUopen + | union s hs t ht hs' ht' => rw [image_union]; exact hs'.union ht' + | compl s hs hs' => + rw [← range_diff_image hfopen.injective] + exact (hfcomp.isConstructible hfopen.isOpen_range).sdiff hs' + +@[stacks 09YG] +lemma IsConstructible.image_of_isClosedEmbedding (hf : IsClosedEmbedding f) + (hfcomp : IsRetrocompact (range f)ᶜ) (hs : IsConstructible s) : IsConstructible (f '' s) := by + induction hs using IsConstructible.empty_union_induction with + | open_retrocompact U hUopen hUcomp => + have hfU : IsOpen (f '' U ∪ (range f)ᶜ) := by + simpa [← range_diff_image hf.injective, sdiff_eq, compl_inter, union_comm] + using (hf.isClosedMap _ hUopen.isClosed_compl).isOpen_compl + suffices h : IsRetrocompact (f '' U ∪ (range f)ᶜ) by + simpa [union_inter_distrib_right, inter_eq_left.2 (image_subset_range ..)] + using (h.isConstructible hfU).sdiff (hfcomp.isConstructible hf.isClosed_range.isOpen_compl) + rintro V hVcomp hVopen + rw [union_inter_distrib_right, ← image_inter_preimage] + exact ((hUcomp (hf.isCompact_preimage hVcomp) (hVopen.preimage hf.continuous)).image + hf.continuous).union <| hfcomp hVcomp hVopen + | union s hs t ht hs' ht' => rw [image_union]; exact hs'.union ht' + | compl s hs hs' => + rw [← range_diff_image hf.injective] + exact (hfcomp.isConstructible hf.isClosed_range.isOpen_compl).of_compl.sdiff hs' + +lemma isConstructible_preimage_iff_of_isOpenEmbedding {s : Set Y} (hf : IsOpenEmbedding f) + (hfcomp : IsRetrocompact (range f)) (hsf : s ⊆ range f) : + IsConstructible (f ⁻¹' s) ↔ IsConstructible s where + mp hs := by simpa [image_preimage_eq_range_inter, inter_eq_right.2 hsf] + using hs.image_of_isOpenEmbedding hf hfcomp + mpr := .preimage_of_isOpenEmbedding hf + +section CompactSpace +variable [CompactSpace X] {P : ∀ s : Set X, IsConstructible s → Prop} {B : Set (Set X)} + {b : ι → Set X} + +lemma _root_.IsRetrocompact.isCompact (hs : IsRetrocompact s) : IsCompact s := by + simpa using hs CompactSpace.isCompact_univ + +/-- Variant of `TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact` for a +non-indexed topological basis. -/ +lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact' + (basis : IsTopologicalBasis B) (compact_inter : ∀ U ∈ B, ∀ V ∈ B, IsCompact (U ∩ V)) + (hU : IsOpen U) : IsRetrocompact U ↔ IsCompact U := by + refine ⟨IsRetrocompact.isCompact, fun hU' {V} hV' hV ↦ ?_⟩ + obtain ⟨s, rfl⟩ := eq_sUnion_finset_of_isTopologicalBasis_of_isCompact_open _ basis _ hU' hU + obtain ⟨t, rfl⟩ := eq_sUnion_finset_of_isTopologicalBasis_of_isCompact_open _ basis _ hV' hV + rw [Set.sUnion_inter_sUnion] + refine ((s.finite_toSet.image _).prod (t.finite_toSet.image _)).isCompact_biUnion ?_ + simp only [mem_prod, mem_image, Finset.mem_coe, Subtype.exists, exists_and_right, exists_eq_right, + and_imp, forall_exists_index, Prod.forall] + exact fun u v hu _ hv _ ↦ compact_inter _ hu _ hv + +lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact_iff_isCompact + (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) + (hU : IsOpen U) : IsRetrocompact U ↔ IsCompact U := + basis.isRetrocompact_iff_isCompact' (by simpa using compact_inter) hU + +/-- Variant of `TopologicalSpace.IsTopologicalBasis.isRetrocompact` for a non-indexed topological +basis. -/ +lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact' (basis : IsTopologicalBasis B) + (compact_inter : ∀ U ∈ B, ∀ V ∈ B, IsCompact (U ∩ V)) (hU : U ∈ B) : IsRetrocompact U := + (basis.isRetrocompact_iff_isCompact' compact_inter <| basis.isOpen hU).2 <| by + simpa using compact_inter _ hU _ hU + +lemma _root_.TopologicalSpace.IsTopologicalBasis.isRetrocompact + (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) (i : ι) : + IsRetrocompact (b i) := + (basis.isRetrocompact_iff_isCompact compact_inter <| basis.isOpen <| mem_range_self _).2 <| by + simpa using compact_inter i i + +/-- Variant of `TopologicalSpace.IsTopologicalBasis.isConstructible` for a non-indexed topological +basis. -/ +lemma _root_.TopologicalSpace.IsTopologicalBasis.isConstructible' (basis : IsTopologicalBasis B) + (compact_inter : ∀ U ∈ B, ∀ V ∈ B, IsCompact (U ∩ V)) (hU : U ∈ B) : IsConstructible U := + (basis.isRetrocompact' compact_inter hU).isConstructible <| basis.isOpen hU + +lemma _root_.TopologicalSpace.IsTopologicalBasis.isConstructible + (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) (i : ι) : + IsConstructible (b i) := + (basis.isRetrocompact compact_inter _).isConstructible <| basis.isOpen <| mem_range_self _ + +@[elab_as_elim] +lemma IsConstructible.induction_of_isTopologicalBasis {ι : Type*} [Nonempty ι] (b : ι → Set X) + (basis : IsTopologicalBasis (range b)) (compact_inter : ∀ i j, IsCompact (b i ∩ b j)) + (sdiff : ∀ i s (hs : Set.Finite s), P (b i \ ⋃ j ∈ s, b j) + ((basis.isConstructible compact_inter _).sdiff <| .biUnion hs fun _ _ ↦ + basis.isConstructible compact_inter _)) + (union : ∀ s hs t ht, P s hs → P t ht → P (s ∪ t) (hs.union ht)) + (s : Set X) (hs : IsConstructible s) : P s hs := by + induction s, hs using BooleanSubalgebra.closure_sdiff_sup_induction with + | isSublattice => + exact ⟨fun s hs t ht ↦ ⟨hs.1.union ht.1, hs.2.union ht.2⟩, + fun s hs t ht ↦ ⟨hs.1.inter ht.1, hs.2.inter_isOpen ht.2 ht.1⟩⟩ + | bot_mem => exact ⟨isOpen_empty, .empty⟩ + | top_mem => exact ⟨isOpen_univ, .univ⟩ + | sdiff U hU V hV => + have := isCompact_open_iff_eq_finite_iUnion_of_isTopologicalBasis _ basis + fun i ↦ by simpa using compact_inter i i + obtain ⟨s, hs, rfl⟩ := (this _).1 ⟨hU.2.isCompact, hU.1⟩ + obtain ⟨t, ht, rfl⟩ := (this _).1 ⟨hV.2.isCompact, hV.1⟩ + simp_rw [iUnion_diff] + induction s, hs using Set.Finite.induction_on with + | empty => simpa using sdiff (Classical.arbitrary _) {Classical.arbitrary _} + | @insert i s hi hs ih => + simp_rw [biUnion_insert] + exact union _ _ _ + (.biUnion hs fun i _ ↦ (basis.isConstructible compact_inter _).sdiff <| + .biUnion ht fun j _ ↦ basis.isConstructible compact_inter _) + (sdiff _ _ ht) + (ih ⟨isOpen_biUnion fun _ _ ↦ basis.isOpen ⟨_, rfl⟩, .biUnion hs + fun i _ ↦ basis.isRetrocompact compact_inter _⟩) + | sup s _ t _ hs' ht' => exact union _ _ _ _ hs' ht' + +end CompactSpace + +/-! ### Locally constructible sets -/ + +/-- A set in a topological space is locally constructible, if every point has a neighborhood on +which the set is constructible. -/ +@[stacks 005G] +def IsLocallyConstructible (s : Set X) : Prop := ∀ x, ∃ U ∈ 𝓝 x, IsOpen U ∧ IsConstructible (U ↓∩ s) + +lemma IsConstructible.isLocallyConstructible (hs : IsConstructible s) : IsLocallyConstructible s := + fun _ ↦ ⟨univ, by simp, by simp, + (isConstructible_preimage_iff_of_isOpenEmbedding isOpen_univ.isOpenEmbedding_subtypeVal + (by simp) (by simp)).2 hs⟩ + +lemma _root_.IsRetrocompact.isLocallyConstructible (hUopen : IsOpen U) (hUcomp : IsRetrocompact U) : + IsLocallyConstructible U := (hUcomp.isConstructible hUopen).isLocallyConstructible + +@[simp] protected lemma IsLocallyConstructible.empty : IsLocallyConstructible (∅ : Set X) := + IsConstructible.empty.isLocallyConstructible + +@[simp] protected lemma IsLocallyConstructible.univ : IsLocallyConstructible (univ : Set X) := + IsConstructible.univ.isLocallyConstructible + +lemma IsLocallyConstructible.inter (hs : IsLocallyConstructible s) (ht : IsLocallyConstructible t) : + IsLocallyConstructible (s ∩ t) := by + rintro x + obtain ⟨U, hxU, hU, hsU⟩ := hs x + obtain ⟨V, hxV, hV, htV⟩ := ht x + refine ⟨U ∩ V, Filter.inter_mem hxU hxV, hU.inter hV, ?_⟩ + change IsConstructible + (inclusion inter_subset_left ⁻¹' (U ↓∩ s) ∩ inclusion inter_subset_right ⁻¹' (V ↓∩ t)) + exact .inter (hsU.preimage_of_isOpenEmbedding <| .inclusion _ <| + .preimage continuous_subtype_val <| hU.inter hV) + (htV.preimage_of_isOpenEmbedding <| .inclusion _ <| + .preimage continuous_subtype_val <| hU.inter hV ) + +lemma IsLocallyConstructible.finsetInf {ι : Type*} {s : Finset ι} {t : ι → Set X} + (ht : ∀ i ∈ s, IsLocallyConstructible (t i)) : IsLocallyConstructible (s.inf t) := by + induction' s using Finset.cons_induction with i s ih hi + · simp + · rw [Finset.inf_cons] + exact (ht _ <| by simp).inter <| hi <| Finset.forall_of_forall_cons ht + +set_option linter.docPrime false in +lemma IsLocallyConstructible.finsetInf' {ι : Type*} {s : Finset ι} {hs} {t : ι → Set X} + (ht : ∀ i ∈ s, IsLocallyConstructible (t i)) : IsLocallyConstructible (s.inf' hs t) := by + rw [Finset.inf'_eq_inf]; exact .finsetInf ht + +private lemma infClosed_isLocallyConstructible : InfClosed {s : Set X | IsLocallyConstructible s} := + fun _s hs _t ht ↦ hs.inter ht + +lemma IsLocallyConstructible.iInter [Finite ι] {f : ι → Set X} + (hf : ∀ i, IsLocallyConstructible (f i)) : IsLocallyConstructible (⋂ i, f i) := + infClosed_isLocallyConstructible.iInf_mem .univ hf + +lemma IsLocallyConstructible.sInter {S : Set (Set X)} (hS : S.Finite) + (hS' : ∀ s ∈ S, IsLocallyConstructible s) : IsLocallyConstructible (⋂₀ S) := + infClosed_isLocallyConstructible.sInf_mem hS .univ hS' + +end Topology diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 3eac03c81156e..468b25f3f1616 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -37,6 +37,7 @@ product, sum, disjoint union, subspace, quotient space noncomputable section open Topology TopologicalSpace Set Filter Function +open scoped Set.Notation universe u v @@ -1153,9 +1154,21 @@ protected lemma Topology.IsEmbedding.codRestrict {e : X → Y} (he : IsEmbedding @[deprecated (since := "2024-10-26")] alias Embedding.codRestrict := IsEmbedding.codRestrict -protected lemma Topology.IsEmbedding.inclusion {s t : Set X} (h : s ⊆ t) : +variable {s t : Set X} + +protected lemma Topology.IsEmbedding.inclusion (h : s ⊆ t) : IsEmbedding (inclusion h) := IsEmbedding.subtypeVal.codRestrict _ _ +protected lemma Topology.IsOpenEmbedding.inclusion (hst : s ⊆ t) (hs : IsOpen (t ↓∩ s)) : + IsOpenEmbedding (inclusion hst) where + toIsEmbedding := .inclusion _ + isOpen_range := by rwa [range_inclusion] + +protected lemma Topology.IsClosedEmbedding.inclusion (hst : s ⊆ t) (hs : IsClosed (t ↓∩ s)) : + IsClosedEmbedding (inclusion hst) where + toIsEmbedding := .inclusion _ + isClosed_range := by rwa [range_inclusion] + @[deprecated (since := "2024-10-26")] alias embedding_inclusion := IsEmbedding.inclusion diff --git a/Mathlib/Topology/ContinuousMap/Units.lean b/Mathlib/Topology/ContinuousMap/Units.lean index e64dbd15575a5..d355156e8246e 100644 --- a/Mathlib/Topology/ContinuousMap/Units.lean +++ b/Mathlib/Topology/ContinuousMap/Units.lean @@ -27,7 +27,6 @@ variable [Monoid M] [TopologicalSpace M] [ContinuousMul M] and the units of the monoid of continuous maps. -/ -- `simps` generates some lemmas here with LHS not in simp normal form, -- so we write them out manually below. --- https://github.com/leanprover-community/mathlib4/issues/18942 @[to_additive (attr := simps apply_val_apply symm_apply_apply_val) "Equivalence between continuous maps into the additive units of an additive monoid with continuous addition and the additive units of the additive monoid of continuous maps."] diff --git a/Mathlib/Topology/Defs/Ultrafilter.lean b/Mathlib/Topology/Defs/Ultrafilter.lean index 894eb001af96f..312adb41a2af2 100644 --- a/Mathlib/Topology/Defs/Ultrafilter.lean +++ b/Mathlib/Topology/Defs/Ultrafilter.lean @@ -3,9 +3,9 @@ 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, Jeremy Avigad -/ -import Mathlib.Topology.Defs.Basic -import Mathlib.Order.Filter.Ultrafilter import Mathlib.Data.Set.Lattice +import Mathlib.Order.Filter.Ultrafilter.Defs +import Mathlib.Topology.Defs.Basic import Mathlib.Topology.Defs.Filter /-! diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index ac53d77bf26a1..9a65a096c076d 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -39,6 +39,10 @@ structure IsDenseInducing [TopologicalSpace α] [TopologicalSpace β] (i : α namespace IsDenseInducing variable [TopologicalSpace α] [TopologicalSpace β] + +theorem _root_.Dense.isDenseInducing_val {s : Set α} (hs : Dense s) : + IsDenseInducing (@Subtype.val α s) := ⟨IsInducing.subtypeVal, hs.denseRange_val⟩ + variable {i : α → β} lemma isInducing (di : IsDenseInducing i) : IsInducing i := di.toIsInducing @@ -205,6 +209,51 @@ theorem mk' (i : α → β) (c : Continuous i) (dense : ∀ x, x ∈ closure (ra end IsDenseInducing +namespace Dense + +variable [TopologicalSpace α] [TopologicalSpace β] {s : Set α} + +/-- This is a shortcut for `hs.isDenseInducing_val.extend f`. It is useful because if `s : Set α` +is dense then the coercion `(↑) : s → α` automatically satisfies `IsUniformInducing` and +`IsDenseInducing` so this gives access to the theorems satisfied by a uniform extension by simply +mentioning the density hypothesis. -/ +noncomputable def extend (hs : Dense s) (f : s → β) : α → β := + hs.isDenseInducing_val.extend f + +variable {f : s → β} + +theorem extend_eq_of_tendsto [T2Space β] (hs : Dense s) {a : α} {b : β} + (hf : Tendsto f (comap (↑) (𝓝 a)) (𝓝 b)) : hs.extend f a = b := + hs.isDenseInducing_val.extend_eq_of_tendsto hf + +theorem extend_eq_at [T2Space β] (hs : Dense s) {f : s → β} {x : s} + (hf : ContinuousAt f x) : hs.extend f x = f x := + hs.isDenseInducing_val.extend_eq_at hf + +theorem extend_eq [T2Space β] (hs : Dense s) (hf : Continuous f) (x : s) : + hs.extend f x = f x := + hs.extend_eq_at hf.continuousAt + +theorem extend_unique_at [T2Space β] {a : α} {g : α → β} (hs : Dense s) + (hf : ∀ᶠ x : s in comap (↑) (𝓝 a), g x = f x) (hg : ContinuousAt g a) : + hs.extend f a = g a := + hs.isDenseInducing_val.extend_unique_at hf hg + +theorem extend_unique [T2Space β] {g : α → β} (hs : Dense s) + (hf : ∀ x : s, g x = f x) (hg : Continuous g) : hs.extend f = g := + hs.isDenseInducing_val.extend_unique hf hg + +theorem continuousAt_extend [T3Space β] {a : α} (hs : Dense s) + (hf : ∀ᶠ x in 𝓝 a, ∃ b, Tendsto f (comap (↑) <| 𝓝 x) (𝓝 b)) : + ContinuousAt (hs.extend f) a := + hs.isDenseInducing_val.continuousAt_extend hf + +theorem continuous_extend [T3Space β] (hs : Dense s) + (hf : ∀ a : α, ∃ b, Tendsto f (comap (↑) (𝓝 a)) (𝓝 b)) : Continuous (hs.extend f) := + hs.isDenseInducing_val.continuous_extend hf + +end Dense + /-- A dense embedding is an embedding with dense image. -/ structure IsDenseEmbedding [TopologicalSpace α] [TopologicalSpace β] (e : α → β) extends IsDenseInducing e : Prop where diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index e3582222d3610..0c9f840cb9455 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -596,7 +596,7 @@ theorem edist_pos {x y : γ} : 0 < edist x y ↔ x ≠ y := by simp [← not_le] /-- Two points coincide if their distance is `< ε` for all positive ε -/ theorem eq_of_forall_edist_le {x y : γ} (h : ∀ ε > 0, edist x y ≤ ε) : x = y := - eq_of_edist_eq_zero (eq_of_le_of_forall_le_of_dense bot_le h) + eq_of_edist_eq_zero (eq_of_le_of_forall_lt_imp_le_of_dense bot_le h) /-- Auxiliary function to replace the uniformity on an emetric space with a uniformity which is equal to the original one, but maybe not defeq. diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 09a6da952e60f..d90ef77d7c2dd 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -391,7 +391,6 @@ Trivialization changes from `i` to `j` are given by continuous maps `coordChange `baseSet i ∩ baseSet j` to the set of homeomorphisms of `F`, but we express them as maps `B → F → F` and require continuity on `(baseSet i ∩ baseSet j) × F` to avoid the topology on the space of continuous maps on `F`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure FiberBundleCore (ι : Type*) (B : Type*) [TopologicalSpace B] (F : Type*) [TopologicalSpace F] where baseSet : ι → Set B @@ -410,7 +409,7 @@ namespace FiberBundleCore variable [TopologicalSpace B] [TopologicalSpace F] (Z : FiberBundleCore ι B F) /-- The index set of a fiber bundle core, as a convenience function for dot notation -/ -@[nolint unusedArguments] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was has_nonempty_instance +@[nolint unusedArguments] def Index (_Z : FiberBundleCore ι B F) := ι /-- The base space of a fiber bundle core, as a convenience function for dot notation -/ @@ -419,7 +418,7 @@ def Base (_Z : FiberBundleCore ι B F) := B /-- The fiber of a fiber bundle core, as a convenience function for dot notation and typeclass inference -/ -@[nolint unusedArguments] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was has_nonempty_instance +@[nolint unusedArguments] def Fiber (_ : FiberBundleCore ι B F) (_x : B) := F instance topologicalSpaceFiber (x : B) : TopologicalSpace (Z.Fiber x) := ‹_› @@ -720,7 +719,6 @@ variable (E : B → Type*) [TopologicalSpace B] [TopologicalSpace F] equivalences but there is not yet a topology on the total space. The total space is hence given a topology in such a way that there is a fiber bundle structure for which the partial equivalences are also partial homeomorphisms and hence local trivializations. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure FiberPrebundle where pretrivializationAtlas : Set (Pretrivialization F (π F E)) pretrivializationAt : B → Pretrivialization F (π F E) diff --git a/Mathlib/Topology/FiberBundle/Trivialization.lean b/Mathlib/Topology/FiberBundle/Trivialization.lean index 2652c7590a667..f15bd8790c419 100644 --- a/Mathlib/Topology/FiberBundle/Trivialization.lean +++ b/Mathlib/Topology/FiberBundle/Trivialization.lean @@ -262,7 +262,6 @@ variable [TopologicalSpace Z] [TopologicalSpace (TotalSpace F E)] `proj : Z → B` with fiber `F`, as a partial homeomorphism between `Z` and `B × F` defined between two sets of the form `proj ⁻¹' baseSet` and `baseSet × F`, acting trivially on the first coordinate. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure Trivialization (proj : Z → B) extends PartialHomeomorph Z (B × F) where baseSet : Set B open_baseSet : IsOpen baseSet diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 681b44bc53a60..bf2384b9dc371 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -82,7 +82,6 @@ that the `U i`'s are open subspaces of the glued space. Most of the times it would be easier to use the constructor `TopCat.GlueData.mk'` where the conditions are stated in a less categorical way. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] structure GlueData extends CategoryTheory.GlueData TopCat where f_open : ∀ i j, IsOpenEmbedding (f i j) f_mono i j := (TopCat.mono_iff_injective _).mpr (f_open i j).isEmbedding.injective @@ -101,11 +100,7 @@ theorem isOpen_iff (U : Set 𝖣.glued) : IsOpen U ↔ ∀ i, IsOpen (𝖣.ι i simp_rw [← Multicoequalizer.ι_sigmaπ 𝖣.diagram] rw [← (homeoOfIso (Multicoequalizer.isoCoequalizer 𝖣.diagram).symm).isOpen_preimage] rw [coequalizer_isOpen_iff, colimit_isOpen_iff.{u}] - dsimp only [GlueData.diagram_l, GlueData.diagram_left, GlueData.diagram_r, GlueData.diagram_right, - parallelPair_obj_one] - constructor - · intro h j; exact h ⟨j⟩ - · intro h j; cases j; apply h + tauto theorem ι_jointly_surjective (x : 𝖣.glued) : ∃ (i : _) (y : D.U i), 𝖣.ι i y = x := 𝖣.ι_jointly_surjective (forget TopCat) x @@ -304,7 +299,6 @@ such that We can then glue the topological spaces `U i` together by identifying `V i j` with `V j i`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed `@[nolint has_nonempty_instance]` structure MkCore where {J : Type u} U : J → TopCat.{u} diff --git a/Mathlib/Topology/Homotopy/HomotopyGroup.lean b/Mathlib/Topology/Homotopy/HomotopyGroup.lean index 4f7eabdf8af64..5b89b085f0379 100644 --- a/Mathlib/Topology/Homotopy/HomotopyGroup.lean +++ b/Mathlib/Topology/Homotopy/HomotopyGroup.lean @@ -46,6 +46,8 @@ open Homeomorph noncomputable section +/-- `I^N` is notation (in the Topology namespace) for `N → I`, +i.e. the unit cube indexed by a type `N`. -/ scoped[Topology] notation "I^" N => N → I namespace Cube @@ -76,11 +78,12 @@ end Cube variable (N X : Type*) [TopologicalSpace X] (x : X) -/-- The space of paths with both endpoints equal to a specified point `x : X`. -/ +/-- The space of paths with both endpoints equal to a specified point `x : X`. +Denoted as `Ω`, within the `Topology.Homotopy` namespace. -/ abbrev LoopSpace := Path x x -scoped[Topology.Homotopy] notation "Ω" => LoopSpace +@[inherit_doc] scoped[Topology.Homotopy] notation "Ω" => LoopSpace instance LoopSpace.inhabited : Inhabited (Path x x) := ⟨Path.refl x⟩ @@ -364,11 +367,11 @@ def homotopyGroupEquivFundamentalGroup (i : N) : apply Quotient.congr (loopHomeo i).toEquiv exact fun p q => ⟨homotopicTo i, homotopicFrom i⟩ -/-- Homotopy group of finite index. -/ +/-- Homotopy group of finite index, denoted as `π_n` within the Topology namespace. -/ abbrev HomotopyGroup.Pi (n) (X : Type*) [TopologicalSpace X] (x : X) := HomotopyGroup (Fin n) _ x -scoped[Topology] notation "π_" => HomotopyGroup.Pi +@[inherit_doc] scoped[Topology] notation "π_" => HomotopyGroup.Pi /-- The 0-dimensional generalized loops based at `x` are in bijection with `X`. -/ def genLoopHomeoOfIsEmpty (N x) [IsEmpty N] : Ω^ N X x ≃ₜ X where diff --git a/Mathlib/Topology/Instances/ENNReal/Lemmas.lean b/Mathlib/Topology/Instances/ENNReal/Lemmas.lean index 769d4aa27d05a..4cb57c22494fc 100644 --- a/Mathlib/Topology/Instances/ENNReal/Lemmas.lean +++ b/Mathlib/Topology/Instances/ENNReal/Lemmas.lean @@ -3,15 +3,14 @@ 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 -/ -import Mathlib.Topology.Order.MonotoneContinuity +import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Topology.Algebra.Order.LiminfLimsup import Mathlib.Topology.Instances.ENNReal.Defs import Mathlib.Topology.Instances.NNReal.Lemmas -import Mathlib.Topology.EMetricSpace.Lipschitz -import Mathlib.Topology.Metrizable.Basic -import Mathlib.Topology.Order.T5 import Mathlib.Topology.MetricSpace.Pseudo.Real import Mathlib.Topology.Metrizable.Uniformity +import Mathlib.Topology.Order.MonotoneContinuity +import Mathlib.Topology.Order.T5 /-! # Topology on extended non-negative reals @@ -238,6 +237,20 @@ protected theorem tendsto_nhds_zero {f : Filter α} {u : α → ℝ≥0∞} : Tendsto u f (𝓝 0) ↔ ∀ ε > 0, ∀ᶠ x in f, u x ≤ ε := nhds_zero_basis_Iic.tendsto_right_iff +theorem tendsto_const_sub_nhds_zero_iff {l : Filter α} {f : α → ℝ≥0∞} {a : ℝ≥0∞} (ha : a ≠ ∞) + (hfa : ∀ n, f n ≤ a) : + Tendsto (fun n ↦ a - f n) l (𝓝 0) ↔ Tendsto (fun n ↦ f n) l (𝓝 a) := by + rw [ENNReal.tendsto_nhds_zero, ENNReal.tendsto_nhds ha] + refine ⟨fun h ε hε ↦ ?_, fun h ε hε ↦ ?_⟩ + · filter_upwards [h ε hε] with n hn + refine ⟨?_, (hfa n).trans (le_add_right le_rfl)⟩ + rw [tsub_le_iff_right] at hn ⊢ + rwa [add_comm] + · filter_upwards [h ε hε] with n hn + have hN_left := hn.1 + rw [tsub_le_iff_right] at hN_left ⊢ + rwa [add_comm] + protected theorem tendsto_atTop [Nonempty β] [SemilatticeSup β] {f : β → ℝ≥0∞} {a : ℝ≥0∞} (ha : a ≠ ∞) : Tendsto f atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, f n ∈ Icc (a - ε) (a + ε) := .trans (atTop_basis.tendsto_iff (hasBasis_nhds_of_ne_top ha)) (by simp only [true_and]; rfl) diff --git a/Mathlib/Topology/Instances/Real/Defs.lean b/Mathlib/Topology/Instances/Real/Defs.lean index 93400ed393606..6271ff4cf3eb8 100644 --- a/Mathlib/Topology/Instances/Real/Defs.lean +++ b/Mathlib/Topology/Instances/Real/Defs.lean @@ -6,14 +6,15 @@ Authors: Johannes Hölzl, Mario Carneiro import Mathlib.Data.Real.Star import Mathlib.Topology.Algebra.Order.Field import Mathlib.Topology.Algebra.Star +import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.Instances.Int import Mathlib.Topology.Order.Bornology -import Mathlib.Topology.Algebra.UniformGroup.Defs /-! # Topological properties of ℝ -/ -assert_not_exists UniformOnFun + +assert_not_exists UniformContinuousConstSMul UniformOnFun noncomputable section @@ -40,6 +41,10 @@ theorem Real.uniformContinuous_neg : UniformContinuous (@Neg.neg ℝ _) := instance : UniformAddGroup ℝ := UniformAddGroup.mk' Real.uniformContinuous_add Real.uniformContinuous_neg +theorem Real.uniformContinuous_const_mul {x : ℝ} : UniformContinuous (x * ·) := + uniformContinuous_of_continuousAt_zero (DistribMulAction.toAddMonoidHom ℝ x) + (continuous_const_smul x).continuousAt + -- short-circuit type class inference instance : TopologicalAddGroup ℝ := by infer_instance instance : TopologicalRing ℝ := inferInstance diff --git a/Mathlib/Topology/Instances/Real/Lemmas.lean b/Mathlib/Topology/Instances/Real/Lemmas.lean index a1f50ca973a4c..3ba8704780b0b 100644 --- a/Mathlib/Topology/Instances/Real/Lemmas.lean +++ b/Mathlib/Topology/Instances/Real/Lemmas.lean @@ -4,19 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Periodic -import Mathlib.Data.Real.Star import Mathlib.Topology.Algebra.Order.Archimedean -import Mathlib.Topology.Algebra.Order.Field -import Mathlib.Topology.Algebra.Star -import Mathlib.Topology.Algebra.UniformMulAction -import Mathlib.Topology.Instances.Int -import Mathlib.Topology.Order.Bornology -import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.Instances.Real.Defs /-! # Topological properties of ℝ -/ + assert_not_exists UniformOnFun noncomputable section @@ -67,9 +61,6 @@ theorem Real.uniformContinuous_abs : UniformContinuous (abs : ℝ → ℝ) := theorem Real.continuous_inv : Continuous fun a : { r : ℝ // r ≠ 0 } => a.val⁻¹ := continuousOn_inv₀.restrict -theorem Real.uniformContinuous_const_mul {x : ℝ} : UniformContinuous (x * ·) := - uniformContinuous_const_smul x - theorem Real.uniformContinuous_mul (s : Set (ℝ × ℝ)) {r₁ r₂ : ℝ} (H : ∀ x ∈ s, |(x : ℝ × ℝ).1| < r₁ ∧ |x.2| < r₂) : UniformContinuous fun p : s => p.1.1 * p.1.2 := diff --git a/Mathlib/Topology/MetricSpace/Defs.lean b/Mathlib/Topology/MetricSpace/Defs.lean index e678bdcfedb6b..0ada8be44573b 100644 --- a/Mathlib/Topology/MetricSpace/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Defs.lean @@ -78,7 +78,7 @@ theorem dist_pos {x y : γ} : 0 < dist x y ↔ x ≠ y := by simpa only [not_le] using not_congr dist_le_zero theorem eq_of_forall_dist_le {x y : γ} (h : ∀ ε > 0, dist x y ≤ ε) : x = y := - eq_of_dist_eq_zero (eq_of_le_of_forall_le_of_dense dist_nonneg h) + eq_of_dist_eq_zero (eq_of_le_of_forall_lt_imp_le_of_dense dist_nonneg h) /-- Deduce the equality of points from the vanishing of the nonnegative distance -/ theorem eq_of_nndist_eq_zero {x y : γ} : nndist x y = 0 → x = y := by diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index a04bf9dd698c2..8f1af1685861c 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -170,7 +170,7 @@ theorem dist_eq {α β F : Type*} [PseudoMetricSpace α] [PseudoMetricSpace β] `dist` and `nndist` versions below -/ theorem ratio_unique [DilationClass F α β] {f : F} {x y : α} {r : ℝ≥0} (h₀ : edist x y ≠ 0) (htop : edist x y ≠ ⊤) (hr : edist (f x) (f y) = r * edist x y) : r = ratio f := by - simpa only [hr, ENNReal.mul_eq_mul_right h₀ htop, ENNReal.coe_inj] using edist_eq f x y + simpa only [hr, ENNReal.mul_left_inj h₀ htop, ENNReal.coe_inj] using edist_eq f x y /-- The `ratio` is equal to the distance ratio for any two points with nonzero finite distance; `nndist` version -/ @@ -305,7 +305,7 @@ theorem ratio_comp' {g : β →ᵈ γ} {f : α →ᵈ β} (hne : ∃ x y : α, edist x y ≠ 0 ∧ edist x y ≠ ⊤) : ratio (g.comp f) = ratio g * ratio f := by rcases hne with ⟨x, y, hα⟩ have hgf := (edist_eq (g.comp f) x y).symm - simp_rw [coe_comp, Function.comp, edist_eq, ← mul_assoc, ENNReal.mul_eq_mul_right hα.1 hα.2] + simp_rw [coe_comp, Function.comp, edist_eq, ← mul_assoc, ENNReal.mul_left_inj hα.1 hα.2] at hgf rwa [← ENNReal.coe_inj, ENNReal.coe_mul] diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index c1899369609a7..caf9feed00c8f 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -85,7 +85,6 @@ instance : Inhabited GHSpace := ⟨Quot.mk _ ⟨⟨{0}, isCompact_singleton⟩, singleton_nonempty _⟩⟩ /-- A metric space representative of any abstract point in `GHSpace` -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not yet ported; removed @[nolint has_nonempty_instance]; why? def GHSpace.Rep (p : GHSpace) : Type := (Quotient.out p : NonemptyCompacts ℓ_infty_ℝ) @@ -312,7 +311,7 @@ theorem hausdorffDist_optimal {X : Type u} [MetricSpace X] [CompactSpace X] [Non let Fb := candidatesBOfCandidates F Fgood have : hausdorffDist (range (optimalGHInjl X Y)) (range (optimalGHInjr X Y)) ≤ HD Fb := hausdorffDist_optimal_le_HD _ _ (candidatesBOfCandidates_mem F Fgood) - refine le_trans this (le_of_forall_le_of_dense fun r hr => ?_) + refine le_trans this (le_of_forall_gt_imp_ge_of_dense fun r hr => ?_) have I1 : ∀ x : X, (⨅ y, Fb (inl x, inr y)) ≤ r := by intro x have : f (inl x) ∈ (p : Set _) := Φrange ▸ (mem_range_self _) diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index c7494ebd7c280..5eedc8d457f84 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -445,7 +445,6 @@ def premetricOptimalGHDist : PseudoMetricSpace (X ⊕ Y) where attribute [local instance] premetricOptimalGHDist /-- A metric space which realizes the optimal coupling between `X` and `Y` -/ --- @[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): This linter does not exist yet. def OptimalGHCoupling : Type _ := @SeparationQuotient (X ⊕ Y) (premetricOptimalGHDist X Y).toUniformSpace.toTopologicalSpace @@ -481,7 +480,7 @@ the Hausdorff distance in the optimal coupling, although we only prove here the we need. -/ theorem hausdorffDist_optimal_le_HD {f} (h : f ∈ candidatesB X Y) : hausdorffDist (range (optimalGHInjl X Y)) (range (optimalGHInjr X Y)) ≤ HD f := by - refine le_trans (le_of_forall_le_of_dense fun r hr => ?_) (HD_optimalGHDist_le X Y f h) + refine le_trans (le_of_forall_gt_imp_ge_of_dense fun r hr => ?_) (HD_optimalGHDist_le X Y f h) have A : ∀ x ∈ range (optimalGHInjl X Y), ∃ y ∈ range (optimalGHInjr X Y), dist x y ≤ r := by rintro _ ⟨z, rfl⟩ have I1 : (⨆ x, ⨅ y, optimalGHDist X Y (inl x, inr y)) < r := diff --git a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean index ec335783eba45..0235fc0d81322 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean @@ -232,7 +232,7 @@ theorem bsupr_limsup_dimH (s : Set X) : ⨆ x ∈ s, limsup dimH (𝓝[s] x).sma refine le_antisymm (iSup₂_le fun x _ => ?_) ?_ · refine limsup_le_of_le isCobounded_le_of_bot ?_ exact eventually_smallSets.2 ⟨s, self_mem_nhdsWithin, fun t => dimH_mono⟩ - · refine le_of_forall_ge_of_dense fun r hr => ?_ + · refine le_of_forall_lt_imp_le_of_dense fun r hr => ?_ rcases exists_mem_nhdsWithin_lt_dimH_of_lt_dimH hr with ⟨x, hxs, hxr⟩ refine le_iSup₂_of_le x hxs ?_; rw [limsup_eq]; refine le_sInf fun b hb => ?_ rcases eventually_smallSets.1 hb with ⟨t, htx, ht⟩ diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index b4b370cbcce52..8857a7993af8a 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -272,7 +272,6 @@ theorem MetricSpace.isometry_induced (f : α → β) (hf : f.Injective) [m : Met -- such a bijection need not exist /-- `α` and `β` are isometric if there is an isometric bijection between them. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure IsometryEquiv (α : Type u) (β : Type v) [PseudoEMetricSpace α] [PseudoEMetricSpace β] extends α ≃ β where isometry_toFun : Isometry toFun diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 543d4da8c0d76..4762d0f29d1bb 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -235,7 +235,6 @@ variable [MetricSpace α] {s : Opens α} /-- A type synonym for a subset `s` of a metric space, on which we will construct another metric for which it will be complete. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] def CompleteCopy {α : Type*} [MetricSpace α] (s : Opens α) : Type _ := s namespace CompleteCopy diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean index 6b6cd86da53a2..9f284fb00b881 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean @@ -91,7 +91,7 @@ lemma exists_isCompact_closedBall [WeaklyLocallyCompactSpace α] (x : α) : theorem biInter_gt_closedBall (x : α) (r : ℝ) : ⋂ r' > r, closedBall x r' = closedBall x r := by ext - simp [forall_gt_ge_iff] + simp [forall_gt_imp_ge_iff_le_of_dense] theorem biInter_gt_ball (x : α) (r : ℝ) : ⋂ r' > r, ball x r' = closedBall x r := by ext @@ -100,7 +100,7 @@ theorem biInter_gt_ball (x : α) (r : ℝ) : ⋂ r' > r, ball x r' = closedBall theorem biUnion_lt_ball (x : α) (r : ℝ) : ⋃ r' < r, ball x r' = ball x r := by ext rw [← not_iff_not] - simp [forall_lt_le_iff] + simp [forall_lt_imp_le_iff_le_of_dense] theorem biUnion_lt_closedBall (x : α) (r : ℝ) : ⋃ r' < r, closedBall x r' = ball x r := by ext diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 931e8748cbfce..e26ec3b5c4b6d 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -70,7 +70,7 @@ protected theorem _root_.Topology.IsInducing.noetherianSpace [NoetherianSpace α @[deprecated (since := "2024-10-28")] alias _root_.Inducing.noetherianSpace := IsInducing.noetherianSpace -/-- [Stacks: Lemma 0052 (1)](https://stacks.math.columbia.edu/tag/0052)-/ +@[stacks 0052 "(1)"] instance NoetherianSpace.set [NoetherianSpace α] (s : Set α) : NoetherianSpace s := IsInducing.subtypeVal.noetherianSpace @@ -191,7 +191,7 @@ theorem NoetherianSpace.exists_finset_irreducible [NoetherianSpace α] (s : Clos simpa [Set.exists_finite_iff_finset, Finset.sup_id_eq_sSup] using NoetherianSpace.exists_finite_set_closeds_irreducible s -/-- [Stacks: Lemma 0052 (2)](https://stacks.math.columbia.edu/tag/0052) -/ +@[stacks 0052 "(2)"] theorem NoetherianSpace.finite_irreducibleComponents [NoetherianSpace α] : (irreducibleComponents α).Finite := by obtain ⟨S : Set (Set α), hSf, hSc, hSi, hSU⟩ := @@ -201,7 +201,7 @@ theorem NoetherianSpace.finite_irreducibleComponents [NoetherianSpace α] : rcases isIrreducible_iff_sUnion_isClosed.1 hs.1 S hSc (hSU ▸ Set.subset_univ _) with ⟨t, htS, ht⟩ rwa [ht.antisymm (hs.2 (hSi _ htS) ht)] -/-- [Stacks: Lemma 0052 (3)](https://stacks.math.columbia.edu/tag/0052) -/ +@[stacks 0052 "(3)"] theorem NoetherianSpace.exists_open_ne_empty_le_irreducibleComponent [NoetherianSpace α] (Z : Set α) (H : Z ∈ irreducibleComponents α) : ∃ o : Set α, IsOpen o ∧ o ≠ ∅ ∧ o ≤ Z := by diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index a6da5bc6bc39a..a178c181fd1cd 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -49,7 +49,6 @@ variable {X X' : Type*} {Y Y' : Type*} {Z Z' : Type*} [TopologicalSpace Z] [TopologicalSpace Z'] /-- Partial homeomorphisms, defined on open subsets of the space -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. @[nolint has_nonempty_instance] structure PartialHomeomorph (X : Type*) (Y : Type*) [TopologicalSpace X] [TopologicalSpace Y] extends PartialEquiv X Y where open_source : IsOpen source @@ -691,7 +690,8 @@ protected def trans' (h : e.target = e'.source) : PartialHomeomorph X Z where continuousOn_invFun := e.continuousOn_symm.comp e'.continuousOn_symm <| h.symm ▸ e'.symm_mapsTo /-- Composing two partial homeomorphisms, by restricting to the maximal domain where their -composition is well defined. -/ +composition is well defined. +Within the `Manifold` namespace, there is the notation `e ≫ₕ f` for this. -/ @[trans] protected def trans : PartialHomeomorph X Z := PartialHomeomorph.trans' (e.symm.restrOpen e'.source e'.open_source).symm @@ -1375,10 +1375,6 @@ lemma lift_openEmbedding_apply (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding simp_rw [e.lift_openEmbedding_toFun] apply hf.injective.extend_apply -@[simp] -lemma lift_openEmbedding_symm (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : - (e.lift_openEmbedding hf).symm = f ∘ e.symm := rfl - @[simp] lemma lift_openEmbedding_source (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : (e.lift_openEmbedding hf).source = f '' e.source := rfl @@ -1387,4 +1383,34 @@ lemma lift_openEmbedding_source (e : PartialHomeomorph X Z) (hf : IsOpenEmbeddin lemma lift_openEmbedding_target (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : (e.lift_openEmbedding hf).target = e.target := rfl +@[simp] +lemma lift_openEmbedding_symm (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : + (e.lift_openEmbedding hf).symm = f ∘ e.symm := rfl + +@[simp] +lemma lift_openEmbedding_symm_source (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : + (e.lift_openEmbedding hf).symm.source = e.target := rfl + +@[simp] +lemma lift_openEmbedding_symm_target (e : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : + (e.lift_openEmbedding hf).symm.target = f '' e.source := by + rw [PartialHomeomorph.symm_target, e.lift_openEmbedding_source] + +lemma lift_openEmbedding_trans_apply + (e e' : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) (z : Z) : + (e.lift_openEmbedding hf).symm.trans (e'.lift_openEmbedding hf) z = (e.symm.trans e') z := by + simp [hf.injective.extend_apply e'] + +@[simp] +lemma lift_openEmbedding_trans (e e' : PartialHomeomorph X Z) (hf : IsOpenEmbedding f) : + (e.lift_openEmbedding hf).symm.trans (e'.lift_openEmbedding hf) = e.symm.trans e' := by + ext x + · exact e.lift_openEmbedding_trans_apply e' hf x + · simp [hf.injective.extend_apply e] + · simp_rw [PartialHomeomorph.trans_source, e.lift_openEmbedding_symm_source, e.symm_source, + e.lift_openEmbedding_symm, e'.lift_openEmbedding_source] + refine ⟨fun ⟨hx, ⟨y, hy, hxy⟩⟩ ↦ ⟨hx, ?_⟩, fun ⟨hx, hx'⟩ ↦ ⟨hx, mem_image_of_mem f hx'⟩⟩ + rw [mem_preimage]; rw [comp_apply] at hxy + exact (hf.injective hxy) ▸ hy + end PartialHomeomorph diff --git a/Mathlib/Topology/Separation/Regular.lean b/Mathlib/Topology/Separation/Regular.lean index 3be8602edb54b..ae49d2084217e 100644 --- a/Mathlib/Topology/Separation/Regular.lean +++ b/Mathlib/Topology/Separation/Regular.lean @@ -3,8 +3,9 @@ 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.Topology.Separation.Hausdorff +import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.Compactness.Lindelof +import Mathlib.Topology.Separation.Hausdorff /-! # Regular, normal, T₃, T₄ and T₅ spaces @@ -637,8 +638,8 @@ theorem connectedComponent_eq_iInter_isClopen [T2Space X] [CompactSpace X] (x : ⟨s ∩ v, H2, mem_inter H.2.1 h1⟩ /-- `ConnectedComponents X` is Hausdorff when `X` is Hausdorff and compact -/ +@[stacks 0900 "The Stacks entry proves profiniteness."] instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (ConnectedComponents X) := by - -- Proof follows that of: https://stacks.math.columbia.edu/tag/0900 -- Fix 2 distinct connected components, with points a and b refine ⟨ConnectedComponents.surjective_coe.forall₂.2 fun a b ne => ?_⟩ rw [ConnectedComponents.coe_ne_coe] at ne diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index bac92aeaffc4e..47ad1534b482b 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -164,6 +164,9 @@ theorem mk_inf_mk {U V : Set α} {hU : IsOpen U} {hV : IsOpen V} : theorem coe_inf (s t : Opens α) : (↑(s ⊓ t) : Set α) = ↑s ∩ ↑t := rfl +@[simp] +lemma mem_inf {s t : Opens α} {x : α} : x ∈ s ⊓ t ↔ x ∈ s ∧ x ∈ t := Iff.rfl + @[simp, norm_cast] theorem coe_sup (s t : Opens α) : (↑(s ⊔ t) : Set α) = ↑s ∪ ↑t := rfl @@ -172,6 +175,9 @@ theorem coe_sup (s t : Opens α) : (↑(s ⊔ t) : Set α) = ↑s ∪ ↑t := theorem coe_bot : ((⊥ : Opens α) : Set α) = ∅ := rfl +@[simp] +lemma mem_bot {x : α} : x ∈ (⊥ : Opens α) ↔ False := Iff.rfl + @[simp] theorem mk_empty : (⟨∅, isOpen_empty⟩ : Opens α) = ⊥ := rfl @[simp, norm_cast] diff --git a/Mathlib/Topology/Sheaves/Forget.lean b/Mathlib/Topology/Sheaves/Forget.lean index f32fb2bb507d4..458a5633cd652 100644 --- a/Mathlib/Topology/Sheaves/Forget.lean +++ b/Mathlib/Topology/Sheaves/Forget.lean @@ -41,11 +41,8 @@ The important special case is when that preserves limits and reflects isomorphisms. Then to check the sheaf condition it suffices to check it on the underlying sheaf of types. -Another useful example is the forgetful functor `TopCommRingCat ⥤ TopCat`. - -See . -In fact we prove a stronger version with arbitrary target category. --/ +Another useful example is the forgetful functor `TopCommRingCat ⥤ TopCat`. -/ +@[stacks 0073 "In fact we prove a stronger version with arbitrary target category."] theorem isSheaf_iff_isSheaf_comp' {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] (G : C ⥤ D) [G.ReflectsIsomorphisms] [HasLimitsOfSize.{v, v} C] [PreservesLimitsOfSize.{v, v} G] {X : TopCat.{v}} (F : Presheaf C X) : Presheaf.IsSheaf F ↔ Presheaf.IsSheaf (F ⋙ G) := diff --git a/Mathlib/Topology/Sheaves/Presheaf.lean b/Mathlib/Topology/Sheaves/Presheaf.lean index d509f463163f3..f35935ec8b97b 100644 --- a/Mathlib/Topology/Sheaves/Presheaf.lean +++ b/Mathlib/Topology/Sheaves/Presheaf.lean @@ -38,7 +38,6 @@ variable (C : Type u) [Category.{v} C] namespace TopCat /-- The category of `C`-valued presheaves on a (bundled) topological space `X`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] def Presheaf (X : TopCat.{w}) : Type max u v w := (Opens X)ᵒᵖ ⥤ C diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index efcf49d420ddd..40c9acfeeb8ef 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -46,7 +46,6 @@ namespace ShrinkingLemma This type is equipped with the following partial order: `v ≤ v'` if `v.carrier ⊆ v'.carrier` and `v i = v' i` for `i ∈ v.carrier`. We will use Zorn's lemma to prove that this type has a maximal element, then show that the maximal element must have `carrier = univ`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. @[nolint has_nonempty_instance] @[ext] structure PartialRefinement (u : ι → Set X) (s : Set X) (p : Set X → Prop) where /-- A family of sets that form a partial refinement of `u`. -/ toFun : ι → Set X diff --git a/Mathlib/Topology/Ultrafilter.lean b/Mathlib/Topology/Ultrafilter.lean index 6329b619e4334..5927cb6d15ae8 100644 --- a/Mathlib/Topology/Ultrafilter.lean +++ b/Mathlib/Topology/Ultrafilter.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Jeremy Avigad -/ import Mathlib.Order.Filter.Lift import Mathlib.Topology.Basic -import Mathlib.Order.Filter.Ultrafilter +import Mathlib.Order.Filter.Ultrafilter.Basic /-! # Characterization of basic topological properties in terms of ultrafilters -/ diff --git a/Mathlib/Topology/UniformSpace/Equiv.lean b/Mathlib/Topology/UniformSpace/Equiv.lean index 8f0d5e13b5aaf..53230fada3f82 100644 --- a/Mathlib/Topology/UniformSpace/Equiv.lean +++ b/Mathlib/Topology/UniformSpace/Equiv.lean @@ -30,7 +30,6 @@ variable {α : Type u} {β : Type*} {γ : Type*} {δ : Type*} -- not all spaces are homeomorphic to each other /-- Uniform isomorphism between `α` and `β` -/ ---@[nolint has_nonempty_instance] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not yet ported structure UniformEquiv (α : Type*) (β : Type*) [UniformSpace α] [UniformSpace β] extends α ≃ β where /-- Uniform continuity of the function -/ diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index 1d053c0503ab5..580dea46f38a6 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -660,3 +660,36 @@ theorem uniformly_extend_unique {g : α → γ} (hg : ∀ b, g (e b) = f b) (hc IsDenseInducing.extend_unique _ hg hc end UniformExtension + +section DenseExtension + +variable {α β : Type*} [UniformSpace α] [UniformSpace β] + +theorem isUniformInducing_val (s : Set α) : + IsUniformInducing (@Subtype.val α s) := ⟨uniformity_setCoe⟩ + +namespace Dense + +variable {s : Set α} {f : s → β} + +theorem extend_exists [CompleteSpace β] (hs : Dense s) (hf : UniformContinuous f) (a : α) : + ∃ b, Tendsto f (comap (↑) (𝓝 a)) (𝓝 b) := + uniformly_extend_exists (isUniformInducing_val s) hs.denseRange_val hf a + +theorem extend_spec [CompleteSpace β] (hs : Dense s) (hf : UniformContinuous f) (a : α) : + Tendsto f (comap (↑) (𝓝 a)) (𝓝 (hs.extend f a)) := + uniformly_extend_spec (isUniformInducing_val s) hs.denseRange_val hf a + +theorem uniformContinuous_extend [CompleteSpace β] (hs : Dense s) (hf : UniformContinuous f) : + UniformContinuous (hs.extend f) := + uniformContinuous_uniformly_extend (isUniformInducing_val s) hs.denseRange_val hf + +variable [T0Space β] + +theorem extend_of_ind (hs : Dense s) (hf : UniformContinuous f) (x : s) : + hs.extend f x = f x := + IsDenseInducing.extend_eq_at _ hf.continuous.continuousAt + +end Dense + +end DenseExtension diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 411f0e25ccac7..89fe23dc6ebf9 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Algebra.Order.Group.Indicator import Mathlib.Analysis.Normed.Affine.AddTorsor import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.SpecificLimits.Basic diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index 9ba91b204f66f..6ac2a4a88260c 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -534,7 +534,7 @@ theorem coordChange_linear_comp (i j k : ι) : exact Z.coordChange_comp i j k x hx v /-- The index set of a vector bundle core, as a convenience function for dot notation -/ -@[nolint unusedArguments] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was `nolint has_nonempty_instance` +@[nolint unusedArguments] def Index := ι /-- The base space of a vector bundle core, as a convenience function for dot notation -/ @@ -543,7 +543,7 @@ def Base := B /-- The fiber of a vector bundle core, as a convenience function for dot notation and typeclass inference -/ -@[nolint unusedArguments] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was `nolint has_nonempty_instance` +@[nolint unusedArguments] def Fiber : B → Type _ := Z.toFiberBundleCore.Fiber @@ -745,7 +745,6 @@ The field `exists_coordChange` is stated as an existential statement (instead of fields), since it depends on propositional information (namely `e e' ∈ pretrivializationAtlas`). This makes it inconvenient to explicitly define a `coordChange` function when constructing a `VectorPrebundle`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): was @[nolint has_nonempty_instance] structure VectorPrebundle where pretrivializationAtlas : Set (Pretrivialization F (π F E)) pretrivialization_linear' : ∀ e, e ∈ pretrivializationAtlas → e.IsLinear R diff --git a/Mathlib/Util/Notation3.lean b/Mathlib/Util/Notation3.lean index 8021941ea2a4a..17e8a9c0118e2 100644 --- a/Mathlib/Util/Notation3.lean +++ b/Mathlib/Util/Notation3.lean @@ -442,6 +442,24 @@ def getPrettyPrintOpt (opt? : Option (TSyntax ``prettyPrintOpt)) : Bool := else true +/-- +If `pp.tagAppFns` is true and the head of the current expression is a constant, +then delaborates the head and uses it for the ref. +This causes tokens inside the syntax to refer to this constant. +A consequence is that docgen will linkify the tokens. +-/ +def withHeadRefIfTagAppFns (d : Delab) : Delab := do + let tagAppFns ← getPPOption getPPTagAppFns + if tagAppFns && (← getExpr).getAppFn.consumeMData.isConst then + -- Delaborate the head to register term info and get a syntax we can use for the ref. + -- The syntax `f` itself is thrown away. + let f ← withNaryFn <| withOptionAtCurrPos `pp.tagAppFns true delab + let stx ← withRef f d + -- Annotate to ensure that the full syntax still refers to the whole expression. + annotateTermInfo stx + else + d + /-- `notation3` declares notation using Lean-3-style syntax. @@ -584,7 +602,7 @@ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attr let delabName := name ++ `delab let matcher ← ms.foldrM (fun m t => `($(m.2) >=> $t)) (← `(pure)) trace[notation3] "matcher:{indentD matcher}" - let mut result ← `(`($pat)) + let mut result ← `(withHeadRefIfTagAppFns `($pat)) for (name, id) in boundIdents.toArray do match boundType.getD name .normal with | .normal => result ← `(MatchState.delabVar s $(quote name) (some e) >>= fun $id => $result) diff --git a/MathlibTest/MinImports.lean b/MathlibTest/MinImports.lean index 87a0a011d07ad..ad749a05da97f 100644 --- a/MathlibTest/MinImports.lean +++ b/MathlibTest/MinImports.lean @@ -77,7 +77,12 @@ lemma hi (n : ℕ) : n = n := by extract_goal; rfl section Linter.MinImports set_option linter.minImports.increases false -set_option linter.minImports true + +set_option linter.minImports false +/-- info: Counting imports from here. -/ +#guard_msgs in +#import_bumps + /-- warning: Imports increased to [Init.Guard, Mathlib.Data.Int.Notation] diff --git a/MathlibTest/Nat/log.lean b/MathlibTest/Nat/log.lean new file mode 100644 index 0000000000000..950bfa1034cb4 --- /dev/null +++ b/MathlibTest/Nat/log.lean @@ -0,0 +1,10 @@ +import Mathlib.Data.Nat.Log + +/-! +This used to fail (ran out of heartbeats) but with a new faster `Nat.logC` tagged `csimp`, it +succeeds. +-/ + +/-- info: 10000000 -/ +#guard_msgs in +#eval Nat.log 2 (2 ^ 10000000) diff --git a/MathlibTest/Recall.lean b/MathlibTest/Recall.lean index a665edc1a8f34..01092a624dc92 100644 --- a/MathlibTest/Recall.lean +++ b/MathlibTest/Recall.lean @@ -1,7 +1,7 @@ import Mathlib.Tactic.Recall import Mathlib.Analysis.Calculus.Deriv.Basic import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric set_option linter.style.setOption false -- Remark: When the test is run by make/CI, this option is not set, so we set it here. diff --git a/MathlibTest/casesm.lean b/MathlibTest/casesm.lean index 2a765ce9e8e7b..d8d1862d4560f 100644 --- a/MathlibTest/casesm.lean +++ b/MathlibTest/casesm.lean @@ -81,6 +81,36 @@ example : True ∧ True ∧ True := by · guard_target = True; constructorm True · guard_target = True ∧ True; constructorm* True, _∧_ +example (n : Nat) : True := by + fail_if_success casesm! Nat -- two constructors, so `casesm!` doesn't fire + trivial + +example (h : Array Nat) : True := by + casesm! Array _ + trivial + +example (h : Array Nat) : True := by + casesm Array _ + -- user facing name is preserved: + guard_hyp h : List Nat + trivial + +example (n : Nat) (h : n = 0) : True := by + casesm Nat + · trivial + · -- user facing name is preserved: + guard_hyp h : n + 1 = 0 + trivial + +example (h : P ∧ Q) : True := by + casesm _ ∧ _ + -- user facing name is not used here, because there are multiple new hypotheses. + fail_if_success guard_hyp h : P + rename_i p q + guard_hyp p : P + guard_hyp q : Q + trivial + section AuxDecl variable {p q r : Prop} variable (h : p ∧ q ∨ p ∧ r) diff --git a/MathlibTest/fast_instance.lean b/MathlibTest/fast_instance.lean new file mode 100644 index 0000000000000..0e74123b6222a --- /dev/null +++ b/MathlibTest/fast_instance.lean @@ -0,0 +1,139 @@ +import Mathlib.Tactic.FastInstance +import Mathlib.Logic.Function.Defs +import Mathlib.Tactic.Spread + +namespace testing +set_option autoImplicit false + +-- For debugging: +--set_option trace.Elab.fast_instance true + +/-! +Testing a diamond: CommSemigroup +-/ + +class Mul (α : Type*) where + mul : α → α → α + +class Semigroup (α : Type*) extends Mul α where + mul_assoc (x y z : α) : mul x (mul y z) = mul (mul x y) z + +class CommMagma (α : Type*) extends Mul α where + mul_comm (x y : α) : mul x y = mul y x + +class CommSemigroup (α : Type*) extends Semigroup α, CommMagma α + +structure Wrapped (α : Type*) where + val : α + +variable {α : Type*} + +theorem val_injective : Function.Injective (Wrapped.val (α := α)) + | ⟨_⟩, ⟨_⟩, rfl => rfl + +instance [Mul α] : Mul (Wrapped α) where mul m n := ⟨Mul.mul m.1 n.1⟩ + +@[reducible] def Function.Injective.semigroup {α β : Type*} [Mul α] [Semigroup β] + (f : α → β) (hf : Function.Injective f) + (hmul : ∀ x y, f (Mul.mul x y) = Mul.mul (f x) (f y)) : + Semigroup α := + { ‹Mul α› with + mul_assoc := fun x y z => by apply hf; rw [hmul, hmul, hmul, hmul, Semigroup.mul_assoc] } + +@[reducible] def Function.Injective.commMagma {α β : Type*} [Mul α] [CommMagma β] + (f : α → β) (hf : Function.Injective f) + (hmul : ∀ x y, f (Mul.mul x y) = Mul.mul (f x) (f y)) : + CommMagma α where + mul_comm x y := by + apply hf + rw [hmul, hmul, CommMagma.mul_comm] + +@[reducible] def Function.Injective.commSemigroup {α β : Type*} [Mul α] [CommSemigroup β] + (f : α → β) (hf : Function.Injective f) + (hmul : ∀ x y, f (Mul.mul x y) = Mul.mul (f x) (f y)) : + CommSemigroup α where + toSemigroup := Function.Injective.semigroup f hf hmul + __ := Function.Injective.commMagma f hf hmul + +instance instSemigroup [Semigroup α] : Semigroup (Wrapped α) := + fast_instance% Function.Injective.semigroup _ val_injective (fun _ _ => rfl) + +instance instCommSemigroup [CommSemigroup α] : CommSemigroup (Wrapped α) := + fast_instance% Function.Injective.commSemigroup _ val_injective (fun _ _ => rfl) + +/-- +info: def testing.instSemigroup.{u_1} : {α : Type u_1} → [inst : Semigroup α] → Semigroup (Wrapped α) := +fun {α} [inst : Semigroup α] => @Semigroup.mk (Wrapped α) (@instMulWrapped α (@Semigroup.toMul α inst)) ⋯ +-/ +#guard_msgs in +set_option pp.explicit true in +#print instSemigroup +/-- +info: def testing.instCommSemigroup.{u_1} : {α : Type u_1} → [inst : CommSemigroup α] → CommSemigroup (Wrapped α) := +fun {α} [inst : CommSemigroup α] => + @CommSemigroup.mk (Wrapped α) (@instSemigroup α (@CommSemigroup.toSemigroup α inst)) ⋯ +-/ +#guard_msgs in +set_option pp.explicit true in +#print instCommSemigroup + + +/-! +Non-defeq error +-/ +instance : Mul Nat := ⟨(· * · )⟩ + +/-- +error: Provided instance + { mul := fun x y => y * x } +is not defeq to inferred instance + instMulNat + +Use `set_option trace.Elab.fast_instance true` to analyze the error. + +Trace of fields visited: [] +--- +info: { mul := fun x y => y * x } : Mul Nat +-/ +#guard_msgs in +#check fast_instance% { mul := fun x y => y * x : Mul Nat } + + +/-! +Checking handling of non-structure classes. +-/ + +class Dec (p : Prop) where + [dec : Decidable p] + +axiom It : Prop + +/-- warning: declaration uses 'sorry' -/ +#guard_msgs in +abbrev dec1 : Decidable It := isTrue sorry + +/-- warning: declaration uses 'sorry' -/ +#guard_msgs in +def dec2 : Decidable It := isTrue sorry + +/-- info: @Dec.mk It (@isTrue It dec1.proof_1) : Dec It -/ +#guard_msgs in +set_option pp.explicit true in +#check fast_instance% { dec := dec1 : Dec It } + +/-- +error: Provided instance does not reduce to a constructor application + dec2 +Reduces to an application of testing.dec2. + +This instance is not a structure and not canonical. Use a separate 'instance' command to define it. + +Use `set_option trace.Elab.fast_instance true` to analyze the error. + +Trace of fields visited: [testing.Dec.dec] +--- +info: @Dec.mk It dec2 : Dec It +-/ +#guard_msgs in +set_option pp.explicit true in +#check fast_instance% { dec := dec2 : Dec It } diff --git a/MathlibTest/notation3.lean b/MathlibTest/notation3.lean index c4473b94e10e7..9b3fc8d5f1f25 100644 --- a/MathlibTest/notation3.lean +++ b/MathlibTest/notation3.lean @@ -20,6 +20,50 @@ notation3 "∀ᶠ " (...) " in " f ", " r:(scoped p => Filter.eventually p f) => /-- info: ∀ᶠ (x : ℕ) in Filter.atTop, x < 3 : Prop -/ #guard_msgs in #check ∀ᶠ x in Filter.atTop, x < 3 +/-! +Test that `pp.tagAppFns` causes tokens to be tagged with head constant. +-/ + +open Lean in +def findWithTag (tag : Nat) (f : Format) : List Format := + match f with + | .nil => [] + | .line => [] + | .align _ => [] + | .text _ => [] + | .nest _ f' => findWithTag tag f' + | .append f' f'' => findWithTag tag f' ++ findWithTag tag f'' + | .group f' _ => findWithTag tag f' + | .tag t f' => (if t = tag then [f'] else []) ++ findWithTag tag f' + +open Lean Elab Term in +def testTagAppFns (n : Name) : TermElabM Unit := do + let stx ← `(∀ᶠ x in Filter.atTop, x < 3) + let e ← elabTermAndSynthesize stx none + let f ← Meta.ppExprWithInfos e + -- Find tags for the constant `n` + let tags : Array Nat := f.infos.fold (init := #[]) fun tags tag info => + match info with + | .ofTermInfo info | .ofDelabTermInfo info => + if info.expr.isConstOf n then + tags.push tag + else + tags + | _ => tags + let fmts := tags.map (findWithTag · f.fmt) + unless fmts.all (!·.isEmpty) do throwError "missing tag" + let fmts := fmts.toList.flatten + logInfo m!"{repr <| fmts.map (·.pretty.trim)}" + +section +/-- info: [] -/ +#guard_msgs in #eval testTagAppFns ``Filter.eventually +set_option pp.tagAppFns true +/-- info: ["∀ᶠ", "in", ","] -/ +#guard_msgs in #eval testTagAppFns ``Filter.eventually +end + + -- Testing lambda expressions: notation3 "∀ᶠ' " f ", " p => Filter.eventually (fun x => (p : _ → _) x) f /-- info: ∀ᶠ' Filter.atTop, fun x ↦ x < 3 : Prop -/ diff --git a/MathlibTest/polyrith.lean b/MathlibTest/polyrith.lean index 58c192b53bdc6..1e95db80d41cc 100644 --- a/MathlibTest/polyrith.lean +++ b/MathlibTest/polyrith.lean @@ -504,7 +504,7 @@ example (a b c : ℤ) (h1 : a = b) (h2 : b = 1) : polyrith example (x y : ℚ) (h1 : x + y = 3) (h2 : 3*x = 7) : - x*x*y + y*x*y + 6*x = 3*x*y + 14 := + x*x*y + y*x*y + 6*x = 3*x*y + 14 := by polyrith example (x y z w : ℚ) (hzw : z = w) : x*z + 2*y*z = x*w + 2*y*w := by diff --git a/MathlibTest/positivity.lean b/MathlibTest/positivity.lean index 80b80dab4bb9e..a9b0c4816cbd6 100644 --- a/MathlibTest/positivity.lean +++ b/MathlibTest/positivity.lean @@ -1,5 +1,5 @@ import Mathlib.Tactic.Positivity -import Mathlib.Data.Complex.Exponential +import Mathlib.Data.Complex.Trigonometric import Mathlib.Data.Real.Sqrt import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Analysis.SpecialFunctions.Pow.Real diff --git a/Shake/Main.lean b/Shake/Main.lean index 8d5dcb0265572..6652df137ecb7 100644 --- a/Shake/Main.lean +++ b/Shake/Main.lean @@ -458,14 +458,19 @@ def main (args : List String) : IO UInt32 := do else pure none -- Read the config file - let cfg ← if let some file := cfgFile then + -- `isValidCfgFile` is `false` if and only if the config file is present and invalid. + let (cfg, isValidCfgFile) ← if let some file := cfgFile then try - IO.ofExcept (Json.parse (← IO.FS.readFile file) >>= fromJson? (α := ShakeCfg)) + pure (← IO.ofExcept (Json.parse (← IO.FS.readFile file) >>= fromJson? (α := ShakeCfg)), true) catch e => + -- The `cfgFile` is invalid, so we print the error and return `isValidCfgFile = false`. println! "{e.toString}" - pure {} - else pure {} - + pure ({}, false) + else pure ({}, true) + if !isValidCfgFile then + IO.println s!"Invalid config file '{cfgFile.get!}'" + IO.Process.exit 1 + else -- the list of root modules let mods := if args.mods.isEmpty then #[`Mathlib] else args.mods -- Only submodules of `pkg` will be edited or have info reported on them diff --git a/docs/100.yaml b/docs/100.yaml index 381ca2ffe57b6..2e1768b687e6a 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -8,24 +8,24 @@ 1: title : The Irrationality of the Square Root of 2 decl : irrational_sqrt_two - author : mathlib + authors: mathlib 2: title : Fundamental Theorem of Algebra decl : Complex.exists_root - author : Chris Hughes + authors: Chris Hughes 3: title : The Denumerability of the Rational Numbers decl : Rat.instDenumerable - author : Chris Hughes + authors: Chris Hughes 4: title : Pythagorean Theorem decl : EuclideanGeometry.dist_sq_eq_dist_sq_add_dist_sq_iff_angle_eq_pi_div_two - author : Joseph Myers + authors: Joseph Myers 5: title : Prime Number Theorem 6: title : Gödel’s Incompleteness Theorem - author : Shogo Saito + authors: Shogo Saito links : First incompleteness theorem: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/First.lean Second incompleteness theorem: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/Second.lean @@ -35,21 +35,21 @@ decls : - legendreSym.quadratic_reciprocity - jacobiSym.quadratic_reciprocity - author : Chris Hughes (first) and Michael Stoll (second) + authors: Chris Hughes (first) and Michael Stoll (second) 8: title : The Impossibility of Trisecting the Angle and Doubling the Cube 9: title : The Area of a Circle decl : Theorems100.area_disc - author : James Arthur and Benjamin Davidson and Andrew Souther + authors: James Arthur and Benjamin Davidson and Andrew Souther 10: title : Euler’s Generalization of Fermat’s Little Theorem decl : Nat.ModEq.pow_totient - author : Chris Hughes + authors: Chris Hughes 11: title : The Infinitude of Primes decl : Nat.exists_infinite_primes - author : Jeremy Avigad + authors: Jeremy Avigad 12: title : The Independence of the Parallel Postulate 13: @@ -57,46 +57,46 @@ 14: title : Euler’s Summation of 1 + (1/2)^2 + (1/3)^2 + …. decl : hasSum_zeta_two - author : Marc Masdeu and David Loeffler + authors: Marc Masdeu and David Loeffler 15: title : Fundamental Theorem of Integral Calculus decls : - intervalIntegral.integral_hasStrictDerivAt_of_tendsto_ae_right - intervalIntegral.integral_eq_sub_of_hasDeriv_right_of_le - author : Yury G. Kudryashov (first) and Benjamin Davidson (second) + authors: Yury G. Kudryashov (first) and Benjamin Davidson (second) 16: title : Insolvability of General Higher Degree Equations (Abel-Ruffini Theorem) - author : Thomas Browning + authors: Thomas Browning decl : AbelRuffini.exists_not_solvable_by_rad 17: title : De Moivre’s Formula decl : Complex.cos_add_sin_mul_I_pow - author : Abhimanyu Pallavi Sudhir + authors: Abhimanyu Pallavi Sudhir 18: title : Liouville’s Theorem and the Construction of Transcendental Numbers - author : Jujian Zhang + authors: Jujian Zhang decl : Liouville.transcendental 19: title : Four Squares Theorem decl : Nat.sum_four_squares - author : Chris Hughes + authors: Chris Hughes 20: title : All Primes (1 mod 4) Equal the Sum of Two Squares decl : Nat.Prime.sq_add_sq - author : Chris Hughes + authors: Chris Hughes 21: title : Green’s Theorem 22: title : The Non-Denumerability of the Continuum decl : Cardinal.not_countable_real - author : Floris van Doorn + authors: Floris van Doorn 23: title : Formula for Pythagorean Triples decl : PythagoreanTriple.classification - author : Paul van Wamelen + authors: Paul van Wamelen 24: title : The Independence of the Continuum Hypothesis - author : Jesse Michael Han and Floris van Doorn + authors: Jesse Michael Han and Floris van Doorn links : result: https://github.com/flypitch/flypitch/blob/master/src/summary.lean website: https://flypitch.github.io/ @@ -104,26 +104,26 @@ 25: title : Schroeder-Bernstein Theorem decl : Function.Embedding.schroeder_bernstein - author : Mario Carneiro + authors: Mario Carneiro 26: title : Leibniz’s Series for Pi decl : Real.tendsto_sum_pi_div_four - author : Benjamin Davidson + authors: Benjamin Davidson 27: title : Sum of the Angles of a Triangle decl : EuclideanGeometry.angle_add_angle_add_angle_eq_pi - author : Joseph Myers + authors: Joseph Myers 28: title : Pascal’s Hexagon Theorem 29: title : Feuerbach’s Theorem 30: title : The Ballot Problem - author : Bhavik Mehta and Kexing Ying + authors: Bhavik Mehta and Kexing Ying decl : Ballot.ballot_problem 31: title : Ramsey’s Theorem - author : Bhavik Mehta + authors: Bhavik Mehta links : result: https://github.com/b-mehta/combinatorics/blob/extras/src/inf_ramsey.lean 32: @@ -135,156 +135,156 @@ 34: title : Divergence of the Harmonic Series decl : Real.tendsto_sum_range_one_div_nat_succ_atTop - author : Anatole Dedecker and Yury Kudryashov + authors: Anatole Dedecker and Yury Kudryashov 35: title : Taylor’s Theorem decls : - taylor_mean_remainder_lagrange - taylor_mean_remainder_cauchy - author : Moritz Doll + authors: Moritz Doll 36: title : Brouwer Fixed Point Theorem - author : Brendan Seamas Murphy + authors: Brendan Seamas Murphy links : result: https://github.com/Shamrock-Frost/BrouwerFixedPoint/blob/master/src/brouwer_fixed_point.lean 37: title : The Solution of a Cubic - author : Jeoff Lee + authors: Jeoff Lee decl : Theorems100.cubic_eq_zero_iff 38: title : Arithmetic Mean/Geometric Mean decl : Real.geom_mean_le_arith_mean_weighted - author : Yury G. Kudryashov + authors: Yury G. Kudryashov 39: title : Solutions to Pell’s Equation decls : - Pell.eq_pell - Pell.exists_of_not_isSquare - author : Mario Carneiro (first), Michael Stoll (second) + authors: Mario Carneiro (first), Michael Stoll (second) note : "In `pell.eq_pell`, `d` is defined to be `a*a - 1` for an arbitrary `a > 1`." 40: title : Minkowski’s Fundamental Theorem decl : MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure - author : Alex J. Best and Yaël Dillies + authors: Alex J. Best and Yaël Dillies 41: title : Puiseux’s Theorem 42: title : Sum of the Reciprocals of the Triangular Numbers - author : Jalex Stark and Yury Kudryashov + authors: Jalex Stark and Yury Kudryashov decl : Theorems100.inverse_triangle_sum 43: title : The Isoperimetric Theorem 44: title : The Binomial Theorem decl : add_pow - author : Chris Hughes + authors: Chris Hughes 45: title : The Partition Theorem - author : Bhavik Mehta and Aaron Anderson + authors: Bhavik Mehta and Aaron Anderson decl : Theorems100.partition_theorem 46: title : The Solution of the General Quartic Equation decl : Theorems100.quartic_eq_zero_iff - author : Thomas Zhu + authors: Thomas Zhu 47: title : The Central Limit Theorem 48: title : Dirichlet’s Theorem decl : Nat.setOf_prime_and_eq_mod_infinite - author : David Loeffler, Michael Stoll + authors: David Loeffler, Michael Stoll 49: title : The Cayley-Hamilton Theorem decl : Matrix.aeval_self_charpoly - author : Kim Morrison + authors: Kim Morrison 50: title : The Number of Platonic Solids 51: title : Wilson’s Lemma decl : ZMod.wilsons_lemma - author : Chris Hughes + authors: Chris Hughes 52: title : The Number of Subsets of a Set decl : Finset.card_powerset - author : mathlib + authors: mathlib 53: title : Pi is Transcendental 54: title : Königsberg Bridges Problem decl : Konigsberg.not_isEulerian - author : Kyle Miller + authors: Kyle Miller 55: title : Product of Segments of Chords decl : EuclideanGeometry.mul_dist_eq_mul_dist_of_cospherical_of_angle_eq_pi - author : Manuel Candales + authors: Manuel Candales 56: title : The Hermite-Lindemann Transcendence Theorem 57: title : Heron’s Formula decl : Theorems100.heron - author : Matt Kempster + authors: Matt Kempster 58: title : Formula for the Number of Combinations decls : - Finset.card_powersetCard - Finset.mem_powersetCard - author : mathlib + authors: mathlib 59: title : The Laws of Large Numbers decl : ProbabilityTheory.strong_law_ae - author : Sébastien Gouëzel + authors: Sébastien Gouëzel 60: title : Bezout’s Theorem decl : Nat.gcd_eq_gcd_ab - author : mathlib + authors: mathlib 61: title : Theorem of Ceva 62: title : Fair Games Theorem decl : MeasureTheory.submartingale_iff_expected_stoppedValue_mono - author : Kexing Ying + authors: Kexing Ying 63: title : Cantor’s Theorem decl: Function.cantor_surjective - author : Johannes Hölzl and Mario Carneiro + authors: Johannes Hölzl and Mario Carneiro 64: title : L’Hopital’s Rule decl : deriv.lhopital_zero_nhds - author : Anatole Dedecker + authors: Anatole Dedecker 65: title : Isosceles Triangle Theorem decl : EuclideanGeometry.angle_eq_angle_of_dist_eq - author : Joseph Myers + authors: Joseph Myers 66: title : Sum of a Geometric Series decls : - geom_sum_Ico - NNReal.hasSum_geometric - author : Sander R. Dahmen (finite) and Johannes Hölzl (infinite) + authors: Sander R. Dahmen (finite) and Johannes Hölzl (infinite) 67: title : e is Transcendental - author : Jujian Zhang + authors: Jujian Zhang links : result: https://github.com/jjaassoonn/transcendental/blob/master/src/e_transcendental.lean website: https://jjaassoonn.github.io/transcendental/ 68: title : Sum of an arithmetic series decl : Finset.sum_range_id - author : Johannes Hölzl + authors: Johannes Hölzl 69: title : Greatest Common Divisor Algorithm decls : - EuclideanDomain.gcd - EuclideanDomain.gcd_dvd - EuclideanDomain.dvd_gcd - author : mathlib + authors: mathlib 70: title : The Perfect Number Theorem decl : Theorems100.Nat.eq_two_pow_mul_prime_mersenne_of_even_perfect - author : Aaron Anderson + authors: Aaron Anderson 71: title : Order of a Subgroup decl : Subgroup.card_subgroup_dvd_card - author : mathlib + authors: mathlib 72: title : Sylow’s Theorem decls : @@ -292,42 +292,42 @@ - Sylow.isPretransitive_of_finite - Sylow.card_dvd_index - card_sylow_modEq_one - author : Chris Hughes + authors: Chris Hughes 73: title : Ascending or Descending Sequences (Erdős–Szekeres Theorem) decl : Theorems100.erdos_szekeres - author : Bhavik Mehta + authors: Bhavik Mehta 74: title : The Principle of Mathematical Induction decl : Nat note : Automatically generated when defining the natural numbers - author : Leonardo de Moura + authors: Leonardo de Moura 75: title : The Mean Value Theorem decl : exists_deriv_eq_slope - author : Yury G. Kudryashov + authors: Yury G. Kudryashov 76: title : Fourier Series decls : - fourierCoeff - hasSum_fourier_series_L2 - author : Heather Macbeth + authors: Heather Macbeth 77: title : Sum of kth powers decls : - sum_range_pow - sum_Ico_pow - author : Moritz Firsching and Fabian Kruse and Ashvni Narayanan + authors: Moritz Firsching and Fabian Kruse and Ashvni Narayanan 78: title : The Cauchy-Schwarz Inequality decls : - inner_mul_inner_self_le - norm_inner_le_norm - author : Zhouhang Zhou + authors: Zhouhang Zhou 79: title : The Intermediate Value Theorem decl : intermediate_value_Icc - author : Rob Lewis and Chris Hughes + authors: Rob Lewis and Chris Hughes 80: title : The Fundamental Theorem of Arithmetic decls : @@ -336,33 +336,33 @@ - EuclideanDomain.to_principal_ideal_domain - UniqueFactorizationMonoid - UniqueFactorizationMonoid.factors_unique - author : Chris Hughes + authors: Chris Hughes note : "it also has a generalized version, by showing that every Euclidean domain is a unique factorization domain, and showing that the integers form a Euclidean domain." 81: title : Divergence of the Prime Reciprocal Series decls : - Theorems100.Real.tendsto_sum_one_div_prime_atTop - not_summable_one_div_on_primes - author : Manuel Candales (archive), Michael Stoll (mathlib) + authors: Manuel Candales (archive), Michael Stoll (mathlib) 82: title : Dissection of Cubes (J.E. Littlewood’s ‘elegant’ proof) decl : Theorems100.«82».cannot_cube_a_cube - author : Floris van Doorn + authors: Floris van Doorn 83: title : The Friendship Theorem decl : Theorems100.friendship_theorem - author : Aaron Anderson and Jalex Stark and Kyle Miller + authors: Aaron Anderson and Jalex Stark and Kyle Miller 84: title : Morley’s Theorem 85: title : Divisibility by 3 Rule - author : Kim Morrison + authors: Kim Morrison decls : - Nat.three_dvd_iff 86: title : Lebesgue Measure and Integration decl : MeasureTheory.lintegral - author : Johannes Hölzl + authors: Johannes Hölzl 87: title : Desargues’s Theorem 88: @@ -370,22 +370,22 @@ decls : - card_derangements_eq_numDerangements - numDerangements_sum - author : Henry Swanson + authors: Henry Swanson 89: title : The Factor and Remainder Theorems decls : - Polynomial.dvd_iff_isRoot - Polynomial.mod_X_sub_C_eq_C_eval - author : Chris Hughes + authors: Chris Hughes 90: title : Stirling’s Formula decls : - Stirling.tendsto_stirlingSeq_sqrt_pi - author : Moritz Firsching and Fabian Kruse and Nikolas Kuhn and Heather Macbeth + authors: Moritz Firsching and Fabian Kruse and Nikolas Kuhn and Heather Macbeth 91: title : The Triangle Inequality decl : norm_add_le - author : Zhouhang Zhou + authors: Zhouhang Zhou 92: title : Pick’s Theorem 93: @@ -393,15 +393,15 @@ decls : - Theorems100.birthday - Theorems100.birthday_measure - author : Eric Rodriguez + authors: Eric Rodriguez 94: title : The Law of Cosines decl : EuclideanGeometry.dist_sq_eq_dist_sq_add_dist_sq_sub_two_mul_dist_mul_dist_mul_cos_angle - author : Joseph Myers + authors: Joseph Myers 95: title : Ptolemy’s Theorem decl : EuclideanGeometry.mul_dist_add_mul_dist_eq_mul_dist_of_cospherical - author : Manuel Candales + authors: Manuel Candales 96: title : Principle of Inclusion/Exclusion decls : @@ -411,24 +411,24 @@ - Finset.inclusion_exclusion_card_inf_compl links : github : https://github.com/NeilStrickland/lean_lib/blob/f88d162da2f990b87c4d34f5f46bbca2bbc5948e/src/combinatorics/matching.lean#L304 - author : Neil Strickland (outside mathlib), Yaël Dillies (in mathlib) + authors: Neil Strickland (outside mathlib), Yaël Dillies (in mathlib) 97: title : Cramer’s Rule decl : Matrix.mulVec_cramer - author : Anne Baanen + authors: Anne Baanen 98: title : Bertrand’s Postulate decl : Nat.bertrand - author : Bolton Bailey and Patrick Stevens + authors: Bolton Bailey and Patrick Stevens 99: title : Buffon Needle Problem links : decls : - BuffonsNeedle.buffon_short - BuffonsNeedle.buffon_long - author : Enrico Z. Borba + authors: Enrico Z. Borba 100: title : Descartes Rule of Signs links : github : https://github.com/Timeroot/lean-descartes-signs/ - author : Alex Meiburg + authors: Alex Meiburg diff --git a/docs/1000.yaml b/docs/1000.yaml index a1aecd443a62c..fc2fe3e686f5e 100644 --- a/docs/1000.yaml +++ b/docs/1000.yaml @@ -17,7 +17,7 @@ Q11518: title: Pythagorean theorem decl: EuclideanGeometry.dist_sq_eq_dist_sq_add_dist_sq_iff_angle_eq_pi_div_two - author: Joseph Myers + authors: Joseph Myers Q12524: title: Schwenk's theorem @@ -25,14 +25,14 @@ Q12524: Q26708: title: Binomial theorem decl: add_pow - author: Chris Hughes + authors: Chris Hughes Q32182: title: Balinski's theorem Q33481: title: Arrow's impossibility theorem - author: Benjamin Davidson, Andrew Souther + authors: Benjamin Davidson, Andrew Souther date: 2021 url: https://github.com/asouther4/lean-social-choice/blob/master/src/arrows_theorem.lean @@ -56,7 +56,7 @@ Q132469: Q137164: title: Besicovitch covering theorem decl: Besicovitch.exists_disjoint_closedBall_covering_ae - author: Sébastien Gouëzel + authors: Sébastien Gouëzel date: 2021 Q154210: @@ -75,7 +75,7 @@ Q174955: Q179208: title: Cayley's theorem decl: Equiv.Perm.subgroupOfMulAction - author: Eric Wieser + authors: Eric Wieser date: 2021 Q179467: @@ -97,7 +97,7 @@ Q180345: Q182505: title: Bayes' theorem decl: ProbabilityTheory.cond_eq_inv_mul_cond_mul - author: Rishikesh Vaishnav + authors: Rishikesh Vaishnav date: 2022 Q184410: @@ -106,7 +106,7 @@ Q184410: Q188295: title: Fermat's little theorem decl: ZMod.pow_card_sub_one_eq_one - author: Aaron Anderson + authors: Aaron Anderson date: 2020 Q188745: @@ -118,7 +118,7 @@ Q189136: decls: - exists_deriv_eq_slope - exists_ratio_deriv_eq_ratio_slope - author: Yury G. Kudryashov + authors: Yury G. Kudryashov date: 2019 Q190026: @@ -134,25 +134,25 @@ Q190391: Q190556: title: De Moivre's theorem decl: Complex.cos_add_sin_mul_I_pow - author: Abhimanyu Pallavi Sudhir + authors: Abhimanyu Pallavi Sudhir date: 2019 Q191693: title: Lebesgue's decomposition theorem decl: MeasureTheory.Measure.haveLebesgueDecomposition_of_sigmaFinite - author: Kexing Ying + authors: Kexing Ying date: 2021 Q192760: title: Fundamental theorem of algebra decl: Complex.isAlgClosed - author: Chris Hughes + authors: Chris Hughes date: 2019 Q193286: title: Rolle's theorem decl: exists_deriv_eq_zero - author: Yury G. Kudryashov + authors: Yury G. Kudryashov date: 2019 Q193878: @@ -174,7 +174,7 @@ Q195133: Q200787: title: Gödel's incompleteness theorem - author: Shogo Saito + authors: Shogo Saito # TODO: how to display this data nicely? want two URLs, right? # results: # - First: https://github.com/FormalizedFormalLogic/Incompleteness/blob/master/Incompleteness/Arith/First.lean @@ -184,7 +184,7 @@ Q200787: Q203565: title: Solutions of a general cubic equation decl: Theorems100.cubic_eq_zero_iff - author: Jeoff Lee + authors: Jeoff Lee date: 2022 Q204884: @@ -202,7 +202,7 @@ Q207244: Q208416: title: Independence of the continuum hypothesis url: https://github.com/flypitch/flypitch/blob/d72904c17fbb874f01ffe168667ba12663a7b853/src/summary.lean#L90 - author: Floris van Doorn and Jesse Michael Han + authors: Floris van Doorn and Jesse Michael Han date: 2019-09-17 Q208756: @@ -240,7 +240,7 @@ Q242045: Q245098: title: Intermediate value theorem decl: intermediate_value_Icc - author: Rob Lewis and Chris Hughes + authors: Rob Lewis and Chris Hughes Q245098X: title: Bolzano's theorem @@ -267,7 +267,7 @@ Q256303: Q257387: title: Vitali theorem # i.e. existence of the Vitali set - author: Ching-Tsun Chou + authors: Ching-Tsun Chou url: https://github.com/ctchou/my_lean/blob/main/MyLean/NonMeasurable.lean comment: "mathlib4 pull request at https://github.com/leanprover-community/mathlib4/pull/20722" @@ -295,7 +295,7 @@ Q273037: Q276082: title: Wilson's theorem decl: ZMod.wilsons_lemma - author: Chris Hughes + authors: Chris Hughes Q280116: title: Odd number theorem @@ -331,7 +331,7 @@ Q318767: decls: - Complex.tendsto_tsum_powerSeries_nhdsWithin_stolzCone - Real.tendsto_tsum_powerSeries_nhdsWithin_lt - author: Jeremy Tan + authors: Jeremy Tan Q321237: title: Green's theorem @@ -401,9 +401,9 @@ Q420714: Q422187: title: Myhill–Nerode theorem - # two in-progress formalisations, none merged as of end of 2024 - # https://github.com/atarnoam/lean-automata/blob/main/src/regular_languages.lean - # https://github.com/leanprover-community/mathlib4/pull/11311 + decl: Language.isRegular_iff_finite_range_leftQuotient + authors: Chris Wong + date: 2024-03-24 Q425432: title: Laurent expansion theorem @@ -414,7 +414,7 @@ Q428134: Q459547: title: Ptolemy's theorem decl: EuclideanGeometry.mul_dist_add_mul_dist_eq_mul_dist_of_cospherical - author: Manuel Candales + authors: Manuel Candales Q467205: title: Fundamental theorems of welfare economics @@ -431,25 +431,25 @@ Q468391: Q470877: title: Stirling's theorem decl: Stirling.tendsto_stirlingSeq_sqrt_pi - author: Moritz Firsching and Fabian Kruse and Nikolas Kuhn and Heather Macbeth + authors: Moritz Firsching and Fabian Kruse and Nikolas Kuhn and Heather Macbeth Q472883: title: Quadratic reciprocity theorem decls: - legendreSym.quadratic_reciprocity - jacobiSym.quadratic_reciprocity - author: Chris Hughes (first) and Michael Stoll (second) + authors: Chris Hughes (first) and Michael Stoll (second) Q474881: title: Cantor's theorem # a.k.a. Cantor's diagonal argument decl: Function.cantor_surjective - author: Johannes Hölzl and Mario Carneiro + authors: Johannes Hölzl and Mario Carneiro Q476776: title: Solutions of a general quartic equation decl: Theorems100.quartic_eq_zero_iff - author: Thomas Zhu + authors: Thomas Zhu Q487132: title: Infinite monkey theorem @@ -510,7 +510,7 @@ Q544369: Q550402: title: Dirichlet's theorem on arithmetic progressions decl: Nat.setOf_prime_and_eq_mod_infinite - author: David Loeffler, Michael Stoll + authors: David Loeffler, Michael Stoll Q552367: title: Berge's theorem @@ -564,7 +564,7 @@ Q608294: Q609612: title: Knaster–Tarski theorem decl: fixedPoints.completeLattice - author: Kenny Lau + authors: Kenny Lau date: 2018 Q612021: @@ -591,7 +591,7 @@ Q632546X: Q632546: title: Bertrand's postulate decl: Nat.bertrand - author: Bolton Bailey and Patrick Stevens + authors: Bolton Bailey and Patrick Stevens Q637418: title: Napoleon's theorem @@ -599,7 +599,7 @@ Q637418: Q642620: title: Krein–Milman theorem decl: closure_convexHull_extremePoints - author: Yaël Dillies + authors: Yaël Dillies date: 2022 Q643513: @@ -639,7 +639,7 @@ Q656645: Q656772: title: Cayley–Hamilton theorem decl: Matrix.aeval_self_charpoly - author: Kim Morrison + authors: Kim Morrison Q657469: title: Lefschetz fixed-point theorem @@ -649,7 +649,7 @@ Q657469X: Q657482: title: Abel–Ruffini theorem - author: Thomas Browning + authors: Thomas Browning decl: AbelRuffini.exists_not_solvable_by_rad Q657903: @@ -668,7 +668,7 @@ Q670235: - EuclideanDomain.to_principal_ideal_domain - UniqueFactorizationMonoid - UniqueFactorizationMonoid.factors_unique - author: Chris Hughes + authors: Chris Hughes comment: "it also has a generalized version, by showing that every Euclidean domain is a unique factorization domain, and showing that the integers form a Euclidean domain." Q671663: @@ -692,13 +692,13 @@ Q693083: Q716171: title: Erdős–Ginzburg–Ziv theorem decl: ZMod.erdos_ginzburg_ziv - author: Yaël Dillies + authors: Yaël Dillies date: 2023 Q718875: title: Erdős–Ko–Rado theorem decl: Finset.erdos_ko_rado - author: Bhavik Mehta, Yaël Dillies + authors: Bhavik Mehta, Yaël Dillies date: 2020 Q719966: @@ -707,7 +707,7 @@ Q719966: Q720469: title: Chevalley–Warning theorem decl: char_dvd_card_solutions_of_sum_lt - author: Johan Commelin + authors: Johan Commelin date: 2019 Q721695: @@ -747,7 +747,7 @@ Q744440: Q748233: title: Sylvester–Gallai theorem - author: Bhavik Mehta + authors: Bhavik Mehta url: https://github.com/YaelDillies/LeanCamCombi/blob/6a6a670f324b2af82ae17f21f9d51ac0bc859f6f/LeanCamCombi/SylvesterChvatal.lean#L610 date: 2022 @@ -757,7 +757,7 @@ Q751120: Q752375: title: Extreme value theorem decl: IsCompact.exists_isMinOn - author: Sébastien Gouëzel + authors: Sébastien Gouëzel date: 2018 Q755986: @@ -769,13 +769,13 @@ Q755991: Q756946: title: Lagrange's four-square theorem decl: Nat.sum_four_squares - author: Chris Hughes + authors: Chris Hughes date: 2019 Q764287: title: Van der Waerden's theorem decl: Combinatorics.exists_mono_homothetic_copy - author: David Wärn + authors: David Wärn # TODO: this is only some version; should this entry be modified accordingly? Q765987: @@ -831,7 +831,7 @@ Q830513: Q834025: title: Cauchy integral theorem decl: Complex.circleIntegral_div_sub_of_differentiable_on_off_countable - author: Yury Kudryashov + authors: Yury Kudryashov date: 2021 Q834211: @@ -839,7 +839,7 @@ Q834211: Q837506: title: Theorem on friends and strangers - author: Bhavik Mehta + authors: Bhavik Mehta url: https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/Small.20Ramsey.20numbers/near/490303695 comment: "will enter mathlib in 2024" @@ -853,7 +853,7 @@ Q841893: Q842953: title: Lebesgue's density theorem decl: Besicovitch.ae_tendsto_measure_inter_div - author: Oliver Nash + authors: Oliver Nash date: 2022 Q844612: @@ -892,7 +892,7 @@ Q853067: decls: - Pell.eq_pell - Pell.exists_of_not_isSquare - author: Mario Carneiro (first), Michael Stoll (second) + authors: Mario Carneiro (first), Michael Stoll (second) comment: "In `pell.eq_pell`, `d` is defined to be `a*a - 1` for an arbitrary `a > 1`." Q856032: @@ -969,14 +969,14 @@ Q913849: Q914517: title: Fermat's theorem on sums of two squares decl: Nat.Prime.sq_add_sq - author: Chris Hughes + authors: Chris Hughes Q915474: title: Robin's theorem Q918099: title: Ramsey's theorem - author: Bhavik Mehta + authors: Bhavik Mehta url: https://github.com/b-mehta/combinatorics/blob/extras/src/inf_ramsey.lean Q922012: @@ -1013,7 +1013,7 @@ Q931404: Q939927: title: Stone–Weierstrass theorem decl: ContinuousMap.subalgebra_topologicalClosure_eq_top_of_separatesPoints - author: Scott Morrison and Heather Macbeth + authors: Scott Morrison and Heather Macbeth date: 2021-05-01 Q942046: @@ -1030,7 +1030,7 @@ Q944297: Q948664: title: Kneser's addition theorem - author: Mantas Bakšys, Yaël Dillies + authors: Mantas Bakšys, Yaël Dillies url: https://github.com/YaelDillies/LeanCamCombi/blob/master/LeanCamCombi/Kneser/Kneser.lean date: 2022 @@ -1065,7 +1065,7 @@ Q976033: Q976607: title: Erdős–Szekeres theorem decl: Theorems100.erdos_szekeres - author: Bhavik Mehta + authors: Bhavik Mehta Q977912: title: Stolper–Samuelson theorem @@ -1091,12 +1091,12 @@ Q1008566: Q1032886: title: Hales–Jewett theorem decl: Combinatorics.Line.exists_mono_in_high_dimension - author: David Wärn + authors: David Wärn Q1033910: title: Cantor–Bernstein–Schroeder theorem decl: Function.Embedding.schroeder_bernstein - author: Mario Carneiro + authors: Mario Carneiro Q1037559: title: Pólya enumeration theorem @@ -1111,7 +1111,7 @@ Q1046232: Q1047749: title: Turán's theorem decl: SimpleGraph.isTuranMaximal_iff_nonempty_iso_turanGraph - author: Jeremy Tan + authors: Jeremy Tan Q1048589: title: Hohenberg–Kohn theorems @@ -1130,7 +1130,7 @@ Q1050203: Q1050932: title: Hartogs's theorem url: https://github.com/girving/ray/blob/main/Ray/Hartogs/Hartogs.lean - author: Geoffrey Irving + authors: Geoffrey Irving Q1051404: title: Cesàro's theorem @@ -1220,7 +1220,7 @@ Q1077741: Q1082910: title: Euler's partition theorem - author: Bhavik Mehta and Aaron Anderson + authors: Bhavik Mehta and Aaron Anderson decl: Theorems100.partition_theorem Q1095330: @@ -1231,7 +1231,7 @@ Q1095330: Q1097021: title: Minkowski's theorem decl: MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure - author: Alex J. Best and Yaël Dillies + authors: Alex J. Best and Yaël Dillies date: 2021 Q1103054: @@ -1282,7 +1282,7 @@ Q1137206: decls: - taylor_mean_remainder_lagrange - taylor_mean_remainder_cauchy - author: Moritz Doll + authors: Moritz Doll Q1139041: title: Cauchy's theorem @@ -1309,7 +1309,7 @@ Q1143540: Q1144897: title: Brouwer fixed-point theorem - author: Brendan Seamas Murphy + authors: Brendan Seamas Murphy url: https://github.com/Shamrock-Frost/BrouwerFixedPoint/blob/master/src/brouwer_fixed_point.lean comment: "in Lean 3" @@ -1391,7 +1391,7 @@ Q1191319: Q1191709: title: Egorov's theorem decl: MeasureTheory.tendstoUniformlyOn_of_ae_tendsto - author: Kexing Ying + authors: Kexing Ying Q1191862: title: Rouché's theorem @@ -1418,7 +1418,7 @@ Q1217677: decls: - intervalIntegral.integral_hasStrictDerivAt_of_tendsto_ae_right - intervalIntegral.integral_eq_sub_of_hasDeriv_right_of_le - author: Yury G. Kudryashov (first) and Benjamin Davidson (second) + authors: Yury G. Kudryashov (first) and Benjamin Davidson (second) Q1227061: title: Khinchin's theorem @@ -1435,7 +1435,7 @@ Q1242398: Q1243340: title: Birkhoff–Von Neumann theorem decl: doublyStochastic_eq_convexHull_permMatrix - author: Bhavik Mehta + authors: Bhavik Mehta date: 2024 Q1249069: @@ -1544,7 +1544,7 @@ Q1425529: Q1426292: title: Banach–Steinhaus theorem decl: banach_steinhaus - author: Jireh Loreaux + authors: Jireh Loreaux date: 2021 Q1434158: @@ -1581,7 +1581,7 @@ Q1505529: Q1506253: title: Euclid's theorem decl: Nat.exists_infinite_primes - author: Jeremy Avigad + authors: Jeremy Avigad Q1529941: title: Peter–Weyl theorem @@ -1595,7 +1595,7 @@ Q1535225: Q1542114: title: Bézout's theorem decl: Nat.gcd_eq_gcd_ab - author: mathlib + authors: mathlib Q1543149: title: Sokhatsky–Weierstrass theorem @@ -1615,7 +1615,7 @@ Q1564541: Q1566341: title: Hindman's theorem decl: Hindman.FP_partition_regular - author: David Wärn + authors: David Wärn Q1566372: title: Haag's theorem @@ -1626,7 +1626,7 @@ Q1568612: Q1568811: title: Hahn decomposition theorem decl: MeasureTheory.SignedMeasure.exists_isCompl_positive_negative - author: Kexing Ying + authors: Kexing Ying date: 2021 Q1572474: @@ -1659,7 +1659,7 @@ Q1632301: Q1632433: title: Helly's theorem - author: Vasily Nesterov + authors: Vasily Nesterov decl: Convex.helly_theorem date: 2023 @@ -1678,7 +1678,7 @@ Q1683356: Q1687147: title: Sprague–Grundy theorem decl: SetTheory.PGame.equiv_nim_grundyValue - author: Fox Thomson + authors: Fox Thomson date: 2020 Q1694565: @@ -1822,7 +1822,7 @@ Q2022775: Q2027347: title: Optional stopping theorem decl: MeasureTheory.submartingale_iff_expected_stoppedValue_mono - author: Kexing Ying, Rémy Degenne + authors: Kexing Ying, Rémy Degenne date: 2022 Q2028341: @@ -1918,13 +1918,13 @@ Q2226774: Q2226786: title: Sperner's theorem decl: IsAntichain.sperner - author: Bhavik Mehta, Alena Gusakov, Yaël Dillies + authors: Bhavik Mehta, Alena Gusakov, Yaël Dillies date: 2022 Q2226800: title: Schur–Zassenhaus theorem decl: Subgroup.exists_left_complement'_of_coprime - author: Thomas Browning + authors: Thomas Browning date: 2021 Q2226803: @@ -1936,7 +1936,7 @@ Q2226807: Q2226855: title: Sharkovskii's theorem # TODO: Find the code - author: Bhavik Mehta + authors: Bhavik Mehta date: 2022 Q2226868: @@ -1947,7 +1947,7 @@ Q2226962: Q2253746: title: Bertrand's ballot theorem - author: Bhavik Mehta and Kexing Ying + authors: Bhavik Mehta and Kexing Ying decl: Ballot.ballot_problem Q2266397: @@ -1965,7 +1965,7 @@ Q2270905: Q2275191: title: Lebesgue differentiation theorem decl: VitaliFamily.ae_tendsto_lintegral_div - author: Sébastien Gouëzel + authors: Sébastien Gouëzel date: 2021 Q2277040: @@ -2120,7 +2120,7 @@ Q2916568: Q2919401: title: Ostrowski's theorem decl: Rat.AbsoluteValue.equiv_real_or_padic - author: David Kurniadi Angdinata, Fabrizio Barroero, Laura Capuano, Nirvana Coppola, María Inés de Frutos-Fernández, Sam van Gool, Silvain Rideau-Kikuchi, Amos Turchet, Francesco Veneziano + authors: David Kurniadi Angdinata, Fabrizio Barroero, Laura Capuano, Nirvana Coppola, María Inés de Frutos-Fernández, Sam van Gool, Silvain Rideau-Kikuchi, Amos Turchet, Francesco Veneziano date: 2024 Q2981012: @@ -2168,7 +2168,7 @@ Q3229335: Q3229352: title: Vitali covering theorem decl: Vitali.exists_disjoint_covering_ae - author: Sébastien Gouëzel + authors: Sébastien Gouëzel date: 2021 Q3345678: @@ -2206,7 +2206,7 @@ Q3526993: Q3526996: title: Kolmogorov extension theorem - author: Rémy Degenne, Peter Pfaffelhuber + authors: Rémy Degenne, Peter Pfaffelhuber date: 2023 url: https://github.com/RemyDegenne/kolmogorov_extension4 @@ -2279,7 +2279,7 @@ Q3527100: Q3527102: title: Kruskal–Katona theorem decl: Finset.kruskal_katona - author: Bhavik Mehta, Yaël Dillies + authors: Bhavik Mehta, Yaël Dillies date: 2020 Q3527110: @@ -2339,7 +2339,7 @@ Q3527214: Q3527215: title: Hilbert projection theorem decl: exists_norm_eq_iInf_of_complete_convex - author: Zhouhang Zhou + authors: Zhouhang Zhou date: 2019 Q3527217: @@ -2372,13 +2372,13 @@ Q3527247: Q3527250: title: Hadamard three-lines theorem decl: Complex.HadamardThreeLines.norm_le_interpStrip_of_mem_verticalClosedStrip - author: Xavier Généreux + authors: Xavier Généreux date: 2023 Q3527263: title: Kleene fixed-point theorem decl: fixedPoints.lfp_eq_sSup_iterate - author: Ira Fesefeldt + authors: Ira Fesefeldt date: 2024 Q3527279: @@ -2436,7 +2436,7 @@ Q3984052: Q3984053: title: Fourier inversion theorem decl: Continuous.fourier_inversion - author: Sébastien Gouëzel + authors: Sébastien Gouëzel date: 2024 Q3984056: @@ -2469,7 +2469,7 @@ Q4272645: Q4378868: title: Phragmén–Lindelöf theorem decl: PhragmenLindelof.horizontal_strip - author: Yury G. Kudryashov + authors: Yury G. Kudryashov date: 2022 Q4378889: @@ -2560,7 +2560,7 @@ Q4677985: Q4695203: title: Four functions theorem decl: four_functions_theorem - author: Yaël Dillies + authors: Yaël Dillies date: 2023 Q4700718: @@ -2632,7 +2632,7 @@ Q4827308: Q4830725: title: Ax–Grothendieck theorem decl: ax_grothendieck_zeroLocus - author: Chris Hughes + authors: Chris Hughes date: 2023 Q4832965: @@ -2848,7 +2848,7 @@ Q5166389: Q5171652: title: Corners theorem decl: corners_theorem_nat - author: Yaël Dillies, Bhavik Mehta + authors: Yaël Dillies, Bhavik Mehta date: 2022 Q5172143: @@ -3000,7 +3000,7 @@ Q5503689: Q5504427: title: Friendship theorem decl: Theorems100.friendship_theorem - author: Aaron Anderson and Jalex Stark and Kyle Miller + authors: Aaron Anderson and Jalex Stark and Kyle Miller Q5505098: title: Frobenius determinant theorem @@ -3276,7 +3276,7 @@ Q6749789: Q6757284: title: Marcinkiewicz theorem - author: Jim Portegies + authors: Jim Portegies date: 2024-09-06 url: https://github.com/fpvandoorn/carleson/blob/master/Carleson/RealInterpolation.lean # will become part of mathlib; update URL in this case! @@ -3532,7 +3532,7 @@ Q7433034: Q7433182: title: Schwartz–Zippel theorem decl: MvPolynomial.schwartz_zippel_sup_sum - author: Bolton Bailey, Yaël Dillies + authors: Bolton Bailey, Yaël Dillies date: 2023 Q7433295: @@ -3693,7 +3693,7 @@ Q7857350: Q7888360: title: Structure theorem for finitely generated modules over a principal ideal domain decl: Module.equiv_free_prod_directSum - author: Pierre-Alexandre Bazin + authors: Pierre-Alexandre Bazin date: 2022 Q7894110: @@ -3759,7 +3759,7 @@ Q8066795: Q8074796: title: Zsigmondy's theorem # TODO: Find the code - author: Mantas Bakšys + authors: Mantas Bakšys date: 2022 Q8081891: @@ -3781,7 +3781,7 @@ Q10942247: Q11352023: title: Vitali convergence theorem decl: MeasureTheory.tendstoInMeasure_iff_tendsto_Lp - author: Igor Khavkine + authors: Igor Khavkine date: 2024 Q11573495: @@ -3820,7 +3820,7 @@ Q15895894: Q16251580: title: Matiyasevich's theorem decl: Pell.matiyasevic - author: Mario Carneiro + authors: Mario Carneiro date: 2017 Q16680059: @@ -3841,7 +3841,7 @@ Q17003552: Q17005116: title: Birkhoff's representation theorem decl: LatticeHom.birkhoffSet - author: Yaël Dillies + authors: Yaël Dillies date: 2022 Q17008559: @@ -3911,7 +3911,7 @@ Q18206266: title: Euclid–Euler theorem # a.k.a. perfect number theorem decl: Theorems100.Nat.eq_two_pow_mul_prime_mersenne_of_even_perfect - author: Aaron Anderson + authors: Aaron Anderson Q18630480: title: Euler's quadrilateral theorem @@ -3931,7 +3931,7 @@ Q20971632: Q22952648: title: Uncountability of the continuum decl: Cardinal.not_countable_real - author: Floris van Doorn + authors: Floris van Doorn Q25099402: title: Kolmogorov–Arnold representation theorem diff --git a/docs/overview.yaml b/docs/overview.yaml index 2d11b73912178..8a37e088447d8 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -377,7 +377,7 @@ Probability Theory: independent events: 'ProbabilityTheory.iIndepSet' independent sigma-algebras: 'ProbabilityTheory.iIndep' conditional probability: 'ProbabilityTheory.cond' - conditional expectation: 'MeasureTheory.condexp' + conditional expectation: 'MeasureTheory.condExp' Random variables and their laws: discrete law: 'PMF' probability density function: 'MeasureTheory.HasPDF' @@ -427,17 +427,63 @@ Geometry: Nullstellensatz : 'MvPolynomial.vanishingIdeal_zeroLocus_eq_radical' Combinatorics: - Graph theory: - simple graph: 'SimpleGraph' - degree-sum formula: 'SimpleGraph.sum_degrees_eq_twice_card_edges' - matching: 'SimpleGraph.Subgraph.IsMatching' - adjacency matrix: 'SimpleGraph.adjMatrix' Pigeonhole principles: finite: 'Fintype.exists_ne_map_eq_of_card_lt' infinite: 'Finite.exists_infinite_fiber' strong pigeonhole principle: 'Fintype.exists_lt_card_fiber_of_mul_lt_card' Transversals: Hall's marriage theorem: 'Finset.all_card_le_biUnion_card_iff_exists_injective' + Enumerative combinatorics: + Catalan numbers: 'catalan' + Bell numbers: 'Multiset.bell' + Dyck words: 'DyckWord' + Set families: + # Correlation inequalities: + Four functions theorem: 'four_functions_theorem' + Harris-Kleitman inequality: 'IsLowerSet.le_card_inter_finset' + # Antichains: + Ahlswede-Zhang identity: 'AhlswedeZhang.infSum_eq_one' + LYM inequality: 'Finset.sum_card_slice_div_choose_le_one' + Sperner's theorem: 'IsAntichain.sperner' + # Intersecting families: + Maximum size of an intersecting family: 'Set.Intersecting.is_max_iff_card_eq' + Kleitman's theorem: 'Finset.card_biUnion_le_of_intersecting' + # VC dimension: + VC dimension: 'Finset.vcDim' + Sauer-Shelah lemma: 'Finset.card_shatterer_le_sum_vcDim' + Kruskal-Katona theorem: 'Finset.kruskal_katona' + Graph theory: + simple graph: 'SimpleGraph' + degree-sum formula: 'SimpleGraph.sum_degrees_eq_twice_card_edges' + matching: 'SimpleGraph.Subgraph.IsMatching' + adjacency matrix: 'SimpleGraph.adjMatrix' + Strongly regular graphs: 'SimpleGraph.IsSRGWith' + # Extremal graph theory: + Turán's theorem: 'SimpleGraph.isTuranMaximal_iff_nonempty_iso_turanGraph' + Regularity lemma: 'szemeredi_regularity' + Triangle counting lemma: 'SimpleGraph.triangle_counting' + Triangle removal lemma: 'SimpleGraph.triangle_removal' + Ramsey theory: + van der Waerden theorem: 'Combinatorics.exists_mono_homothetic_copy' + Hales-Jewett theorem: 'Combinatorics.Line.exists_mono_in_high_dimension' + Hindman's theorem: 'Hindman.FS_partition_regular' + Additive combinatorics: + # Avoiding patterns: + sets without arithmetic progressions: 'ThreeAPFree' + Roth's theorem: 'roth_3ap_theorem_nat' + Corners theorem: corners_theorem_nat + Behrend's construction: 'Behrend.roth_lower_bound' + # Growth: + doubling constant: 'Finset.addConst' + Ruzsa covering lemma: 'Finset.exists_subset_add_sub' + Ruzsa triangle inequality: 'Finset.ruzsa_triangle_inequality_sub_sub_sub' + Plünnecke-Petridis inequality: 'Finset.pluennecke_petridis_inequality_add' + Small tripling implies small powers: 'Finset.small_pow_of_small_tripling' + Dyson transform: 'Finset.addDysonETransform' + Cauchy-Davenport theorem: 'ZMod.cauchy_davenport' + Erdős–Ginzburg–Ziv theorem: 'ZMod.erdos_ginzburg_ziv' + Additive energy: 'Finset.addEnergy' + Freiman homomorphisms: 'IsAddFreimanHom' Dynamics: Circle dynamics: diff --git a/docs/references.bib b/docs/references.bib index 9136aacea3cb3..817246b9e23f7 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -21,6 +21,21 @@ @Book{ abramsky_gabbay_maibaum_1994 zbl = {0829.68111} } +@Book{ Adamek_Rosicky_1994, + author = {Adámek, Jiří and Rosický, Jiří}, + title = {Locally presentable and accessible categories}, + series = {London Mathematical Society Lecture Note Series}, + volume = {189}, + publisher = {Cambridge University Press, Cambridge}, + year = {1994}, + pages = {xiv+316}, + isbn = {0-521-42261-2}, + mrclass = {18Axx (18-02)}, + mrnumber = {1294136}, + doi = {10.1017/CBO9780511600579}, + url = {https://doi.org/10.1017/CBO9780511600579} +} + @Book{ Adamek_Rosicky_Vitale_2010, place = {Cambridge}, series = {Cambridge Tracts in Mathematics}, @@ -276,6 +291,17 @@ @Book{ behrends1979 zbl = {0436.46013} } +@Book{ bensonfarb1993, + author = {Benson Farb, R. Keith Dennis}, + series = {Graduate Texts in Mathematics}, + title = {Noncommutative Algebra}, + year = {1993}, + publisher = {Springer}, + language = {English}, + volume = {144}, + url = {https://link.springer.com/book/10.1007/978-1-4612-0889-1} +} + @Article{ bergelson1985, author = {Bergelson, Vitaly}, title = {Sets of Recurrence of Zm-Actions and Properties of Sets of diff --git a/lake-manifest.json b/lake-manifest.json index 7cc60c94958f9..7373536114b5d 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "5b23a1297aba9683f231c4b1a7ab4076af4ad53d", + "rev": "383137545783429af1317226778fd44869a711e2", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/scripts/add_deprecations.sh b/scripts/add_deprecations.sh index ae89ebfeaf5d5..f3601dda6daed 100755 --- a/scripts/add_deprecations.sh +++ b/scripts/add_deprecations.sh @@ -1,4 +1,4 @@ -#! /bin/bash +#! /usr/bin/env bash : <<'BASH_MODULE_DOC' diff --git a/scripts/import_trans_difference.sh b/scripts/import_trans_difference.sh index 30902e69f971e..8c72ba69ef2ed 100755 --- a/scripts/import_trans_difference.sh +++ b/scripts/import_trans_difference.sh @@ -67,7 +67,7 @@ git checkout "${currCommit}" printf '\n\n
Import changes for all files\n\n%s\n\n
\n' "$( printf "|Files|Import difference|\n|-|-|\n" - (awk -F, -v all="${all}" -v ghLimit='261752' -v newFiles="$( + (gawk -F, -v all="${all}" -v ghLimit='261752' -v newFiles="$( # we pass the "A"dded files with respect to master, converting them to module names git diff --name-only --diff-filter=A master | tr '\n' , | sed 's=\.lean,=,=g; s=/=.=g' )" ' diff --git a/scripts/lint-style.lean b/scripts/lint-style.lean index b4b7d93b9f283..7f1b2cdd34467 100644 --- a/scripts/lint-style.lean +++ b/scripts/lint-style.lean @@ -29,6 +29,12 @@ def findImports (path : System.FilePath) : IO (Array Lean.Name) := do return (← Lean.parseImports' (← IO.FS.readFile path) path.toString) |>.map (fun imp ↦ imp.module) |>.erase `Init +/-- Additional imports generated by `mk_all`. -/ +def explicitImports : Array Lean.Name := #[`Batteries, `Std] + +/-- Remove the additional imports generated by `mk_all` so that only mathlib modules remain. -/ +def eraseExplicitImports (names : Array Lean.Name) : Array Lean.Name := + explicitImports.foldl Array.erase names /-- Check that `Mathlib.Init` is transitively imported in all of Mathlib. @@ -39,7 +45,7 @@ Return `true` iff there was an error. def checkInitImports : IO Bool := do -- Find any file in the Mathlib directory which does not contain any Mathlib import. -- We simply parse `Mathlib.lean`, as CI ensures this file is up to date. - let allModuleNames := (← findImports "Mathlib.lean").erase `Batteries + let allModuleNames := eraseExplicitImports (← findImports "Mathlib.lean") let mut modulesWithoutMathlibImports := #[] let mut importsHeaderLinter := #[] for module in allModuleNames do @@ -107,8 +113,9 @@ def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do let mut allModuleNames := #[] for s in ["Archive.lean", "Counterexamples.lean", "Mathlib.lean"] do allModuleNames := allModuleNames.append (← findImports s) - -- Note: since "Batteries" is added explicitly to "Mathlib.lean", we remove it here manually. - allModuleNames := allModuleNames.erase `Batteries + -- Note: since "Batteries" and "Std" are added explicitly to "Mathlib.lean", we remove them here + -- manually. + allModuleNames := eraseExplicitImports allModuleNames -- Read the `nolints` file, with manual exceptions for the linter. -- NB. We pass these lints to `lintModules` explicitly to prevent cache invalidation bugs: diff --git a/scripts/mk_all.lean b/scripts/mk_all.lean index 00bacbc8eaefb..ae79948610542 100644 --- a/scripts/mk_all.lean +++ b/scripts/mk_all.lean @@ -55,9 +55,9 @@ def mkAllCLI (args : Parsed) : IO UInt32 := do for d in libs.reverse do -- reverse to create `Mathlib/Tactic.lean` before `Mathlib.lean` let fileName := addExtension d "lean" let mut allFiles ← getAllModulesSorted git d - -- mathlib exception: manually import Batteries in `Mathlib.lean` + -- mathlib exception: manually import Std and Batteries in `Mathlib.lean` if d == "Mathlib" then - allFiles := #["Batteries"] ++ allFiles + allFiles := #["Std", "Batteries"] ++ allFiles let fileContent := ("\n".intercalate (allFiles.map ("import " ++ ·)).toList).push '\n' if !(← pathExists fileName) then if check then diff --git a/scripts/nolints.json b/scripts/nolints.json index d2824c5a15099..d52df230c4996 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -35,8 +35,6 @@ ["docBlame", "measurability!?"], ["docBlame", "tacticCancel_denoms_"], ["docBlame", "tacticUse_finite_instance"], - ["docBlame", "«term∮_InC(_,_),_»"], - ["docBlame", "«term⨍_In_.._,_»"], ["docBlame", "when"], ["docBlame", "whenM"], ["docBlame", "Action.V"], @@ -154,30 +152,19 @@ ["docBlame", "ImplicitFunctionData.pt"], ["docBlame", "ImplicitFunctionData.rightDeriv"], ["docBlame", "ImplicitFunctionData.rightFun"], - ["docBlame", "InnerProduct.«term_†»"], ["docBlame", "Int.decNonneg"], ["docBlame", "IntermediateField.delabAdjoinNotation"], ["docBlame", "IsAdjoinRoot.map"], ["docBlame", "JordanHolderLattice.IsMaximal"], ["docBlame", "JordanHolderLattice.Iso"], - ["docBlame", "Kronecker.«term_⊗ₖ_»"], - ["docBlame", "Kronecker.«term_⊗ₖₜ[_]_»"], - ["docBlame", "Kronecker.«term_⊗ₖₜ_»"], ["docBlame", "Lean.ExportM"], ["docBlame", "LinearPMap.domain"], ["docBlame", "LinearPMap.sSup"], - ["docBlame", "LinearPMap.«term_†»"], ["docBlame", "LinearPMap.toFun"], ["docBlame", "LinearPMap.toFun'"], ["docBlame", "LinearRecurrence.coeffs"], ["docBlame", "LinearRecurrence.order"], ["docBlame", "MagmaCat.forget_obj_eq_coe"], - ["docBlame", "Manifold.«term_≫_»"], - ["docBlame", "Manifold.«term_≫ₕ_»"], - ["docBlame", "Manifold.«term𝒅»"], - ["docBlame", "Manifold.«term𝒅ₕ»"], - ["docBlame", "Matrix.«term_⊕ᵥ_»"], - ["docBlame", "Matrix.«term_⊙_»"], ["docBlame", "MaximalSpectrum.asIdeal"], ["docBlame", "MeasureTheory.«term_[_|_]»"], ["docBlame", "MeasureTheory.«term_→₁[_]_»"], @@ -221,7 +208,6 @@ ["docBlame", "NFA.start"], ["docBlame", "NFA.step"], ["docBlame", "NNDist.nndist"], - ["docBlame", "NNReal.«termℝ≥0»"], ["docBlame", "Nat.discriminate"], ["docBlame", "Nat.findX"], ["docBlame", "Nat.ltByCases"], @@ -249,8 +235,6 @@ ["docBlame", "PowerBasis.gen"], ["docBlame", "Pretrivialization.baseSet"], ["docBlame", "PrimeSpectrum.asIdeal"], - ["docBlame", "ProbabilityTheory.«termEVar[_]»"], - ["docBlame", "ProbabilityTheory.«termVar[_]»"], ["docBlame", "ProbabilityTheory.«term_=ₐₛ_»"], ["docBlame", "ProbabilityTheory.«term_≤ₐₛ_»"], ["docBlame", "ProbabilityTheory.«term_⟦_|_⟧»"], @@ -293,8 +277,6 @@ ["docBlame", "Stream'.unfolds"], ["docBlame", "StrictWeakOrder.Equiv"], ["docBlame", "Submodule.«term_∙_»"], - ["docBlame", "Topology.«termI^_»"], - ["docBlame", "Topology.termπ_"], ["docBlame", "Traversable.foldMap"], ["docBlame", "Traversable.foldl"], ["docBlame", "Traversable.foldlm"], @@ -504,7 +486,6 @@ ["docBlame", "Tactic.Elementwise.tacticElementwise___"], ["docBlame", "Tactic.NormNum.evalNatGCD"], ["docBlame", "ToAdditive.additiveTestUnsafe.visit"], - ["docBlame", "Topology.Homotopy.termΩ"], ["docBlame", "WittVector.Isocrystal.frob"], ["docBlame", "AlgebraicGeometry.PresheafedSpace.Hom.base"], ["docBlame", "AlgebraicGeometry.PresheafedSpace.Hom.c"], diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index 6fdfdc2ba7ecd..be856e67cdf24 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -2695,14 +2695,14 @@ MeasureTheory.AEStronglyMeasurable'.of_subsingleton' MeasureTheory.ae_withDensity_iff' MeasureTheory.ae_withDensity_iff_ae_restrict' MeasureTheory.average_eq' -MeasureTheory.condexp_bot' -MeasureTheory.condexpIndL1Fin_smul' -MeasureTheory.condexpIndL1_smul' -MeasureTheory.condexpInd_smul' -MeasureTheory.condexpIndSMul_smul' -MeasureTheory.condexpL1CLM_of_aestronglyMeasurable' -MeasureTheory.condexpL1_of_aestronglyMeasurable' -MeasureTheory.condexp_of_aestronglyMeasurable' +MeasureTheory.condExp_bot' +MeasureTheory.condExpIndL1Fin_smul' +MeasureTheory.condExpIndL1_smul' +MeasureTheory.condExpInd_smul' +MeasureTheory.condExpIndSMul_smul' +MeasureTheory.condExpL1CLM_of_aestronglyMeasurable' +MeasureTheory.condExpL1_of_aestronglyMeasurable' +MeasureTheory.condExp_of_aestronglyMeasurable' MeasureTheory.Content.innerContent_mono' MeasureTheory.diracProba_toMeasure_apply' MeasureTheory.eLpNorm_add_le' @@ -2833,7 +2833,7 @@ MeasureTheory.lowerCrossingTime_stabilize' MeasureTheory.Lp.ae_tendsto_of_cauchy_eLpNorm' MeasureTheory.Lp.eLpNorm'_lim_le_liminf_eLpNorm' MeasureTheory.Lp.eLpNorm'_sum_norm_sub_le_tsum_of_cauchy_eLpNorm' -MeasureTheory.lpMeas.aeStronglyMeasurable' +MeasureTheory.lpMeas.aestronglyMeasurable MeasureTheory.Lp.norm_const' MeasureTheory.Lp.simpleFunc.eq' MeasureTheory.Lp.tendsto_Lp_iff_tendsto_ℒp' @@ -2895,7 +2895,6 @@ MeasureTheory.memℒp_finset_sum' MeasureTheory.Memℒp.integrable_norm_rpow' MeasureTheory.Memℒp.meas_ge_lt_top' MeasureTheory.mem_lpMeas_iff_aeStronglyMeasurable' -MeasureTheory.mem_lpMeasSubgroup_iff_aeStronglyMeasurable' MeasureTheory.Memℒp.mono' MeasureTheory.norm_indicatorConstLp' MeasureTheory.norm_setIntegral_le_of_norm_le_const' @@ -2962,7 +2961,7 @@ MeasureTheory.stoppedProcess_eq'' MeasureTheory.stoppedValue_eq' MeasureTheory.stoppedValue_piecewise_const' MeasureTheory.stoppedValue_sub_eq_sum' -MeasureTheory.StronglyMeasurable.aeStronglyMeasurable' +MeasureTheory.StronglyMeasurable.aestronglyMeasurable MeasureTheory.StronglyMeasurable.const_smul' MeasureTheory.StronglyMeasurable.integral_kernel_prod_left' MeasureTheory.StronglyMeasurable.integral_kernel_prod_left'' @@ -2993,7 +2992,7 @@ MeasureTheory.withDensityᵥ_neg' MeasureTheory.withDensityᵥ_smul' MeasureTheory.withDensityᵥ_smul_eq_withDensityᵥ_withDensity' MeasureTheory.withDensityᵥ_sub' -MeasureTheory.zero_mem_ℒp' +MeasureTheory.Memℒp.zero' mem_ball_iff_norm'' mem_ball_iff_norm''' mem_closedBall_iff_norm'' @@ -3935,8 +3934,8 @@ ProbabilityTheory.cgf_zero' ProbabilityTheory.cond_apply' ProbabilityTheory.cond_cond_eq_cond_inter' ProbabilityTheory.uniformOn_inter' -ProbabilityTheory.condexp_ae_eq_integral_condexpKernel' -ProbabilityTheory.condexpKernel_ae_eq_condexp' +ProbabilityTheory.condExp_ae_eq_integral_condExpKernel' +ProbabilityTheory.condExpKernel_ae_eq_condExp' ProbabilityTheory.CondIndepSets.condIndep' ProbabilityTheory.cond_mul_eq_inter' ProbabilityTheory.evariance_def' diff --git a/scripts/noshake.json b/scripts/noshake.json index ef5a9e86f0813..5a6590b022253 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -96,6 +96,7 @@ "Mathlib.Tactic.ExtractGoal", "Mathlib.Tactic.FBinop", "Mathlib.Tactic.FailIfNoProgress", + "Mathlib.Tactic.FastInstance", "Mathlib.Tactic.FieldSimp", "Mathlib.Tactic.FinCases", "Mathlib.Tactic.Find", @@ -430,7 +431,7 @@ ["Mathlib.Tactic.CategoryTheory.ToApp"], "Mathlib.CategoryTheory.Bicategory.Functor.Lax": ["Mathlib.Tactic.CategoryTheory.ToApp"], - "Mathlib.CategoryTheory.Bicategory.Adjunction": + "Mathlib.CategoryTheory.Bicategory.Adjunction.Basic": ["Mathlib.Tactic.CategoryTheory.Bicategory.Basic"], "Mathlib.Analysis.Normed.Operator.LinearIsometry": ["Mathlib.Algebra.Star.Basic"], @@ -447,11 +448,17 @@ "Mathlib.Algebra.ZeroOne.Lemmas": ["Batteries.Tactic.Init"], "Mathlib.Algebra.Star.Module": ["Mathlib.Algebra.Module.LinearMap.Star"], "Mathlib.Algebra.Ring.CentroidHom": ["Mathlib.Algebra.Algebra.Defs"], + "Mathlib.Algebra.Order.Ring.Nat": ["Mathlib.Algebra.Order.Group.Nat"], "Mathlib.Algebra.Order.Quantale": ["Mathlib.Tactic.Variable"], "Mathlib.Algebra.Order.Pi": ["Mathlib.Algebra.ZeroOne.Lemmas"], + "Mathlib.Algebra.Order.Nonneg.Field": ["Mathlib.Algebra.Order.Field.InjSurj"], + "Mathlib.Algebra.Order.Field.Subfield": + ["Mathlib.Algebra.Order.Field.InjSurj"], "Mathlib.Algebra.Order.CauSeq.Basic": ["Mathlib.Data.Setoid.Basic"], "Mathlib.Algebra.MonoidAlgebra.Basic": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], + "Mathlib.Algebra.Module.Submodule.Order": + ["Mathlib.Algebra.Order.Group.InjSurj"], "Mathlib.Algebra.Module.LinearMap.Star": ["Mathlib.Algebra.Module.Equiv", "Mathlib.Algebra.Module.Equiv.Defs", @@ -459,10 +466,10 @@ "Mathlib.Algebra.Module.LinearMap.Basic": ["Mathlib.Algebra.Star.Basic"], "Mathlib.Algebra.Module.Equiv": ["Mathlib.Algebra.Star.Basic"], "Mathlib.Algebra.Homology.ModuleCat": ["Mathlib.Algebra.Homology.Homotopy"], - "Mathlib.Algebra.Group.Idempotent": ["Mathlib.Data.Subtype"], "Mathlib.Algebra.Group.Units.Basic": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Units": ["Mathlib.Tactic.Subsingleton"], "Mathlib.Algebra.Group.Pi.Basic": ["Batteries.Tactic.Classical"], + "Mathlib.Algebra.Group.Idempotent": ["Mathlib.Data.Subtype"], "Mathlib.Algebra.Group.Defs": ["Mathlib.Tactic.OfNat"], "Mathlib.Algebra.GeomSum": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset"], "Mathlib.Algebra.Category.Ring.Basic": @@ -474,8 +481,6 @@ "Mathlib.Algebra.BigOperators.Group.List": ["Mathlib.Data.List.ProdSigma"], "Mathlib.Algebra.Algebra.Subalgebra.Order": ["Mathlib.Algebra.Module.Submodule.Order"], - "Mathlib.Algebra.Order.Ring.Nat": - ["Mathlib.Algebra.Order.Group.Nat"], "Batteries.Tactic.OpenPrivate": ["Lean.Parser.Module"], "Batteries.Tactic.Omega.OmegaM": ["Batteries.Tactic.Omega.Int"], "Batteries.Tactic.Omega.Frontend": ["Batteries.Tactic.Omega.Logic"], diff --git a/scripts/polyrith_sage.py b/scripts/polyrith_sage.py deleted file mode 100644 index 9833e5923864e..0000000000000 --- a/scripts/polyrith_sage.py +++ /dev/null @@ -1,115 +0,0 @@ -# This file is part of the `polyrith` tactic in `src/tactic/polyrith.lean`. -# It interfaces between Lean and the Sage web interface. - -import json -from os.path import join, dirname -import sys -from typing import Dict, Any -import urllib.error -import urllib.parse -import urllib.request - -# These functions are used to format the output of Sage for parsing in Lean. -# They are stored here as a string since they are passed to Sage via the web API. -with open(join(dirname(__file__), "polyrith_sage_helper.py"), encoding='utf8') as f: - polynomial_formatting_functions = f.read() - -# future extensions may change behavior depending on the base type -def type_str(type): - return "QQ" - -def create_query(type: str, n_vars: int, eq_list, target): - """ Create a query to invoke Sage's `MPolynomial_libsingular.lift`. See - https://github.com/sagemath/sage/blob/f8df80820dc7321dc9b18c9644c3b8315999670b/src/sage/rings/polynomial/multi_polynomial_libsingular.pyx#L4472-L4518 - for a description of this method. """ - var_list = [f"var{i}" for i in range(n_vars)] + ['aux'] - query = f''' -if {n_vars!r} != 0: - P = PolynomialRing({type_str(type)}, {var_list}) - [{", ".join(var_list)}] = P.gens() - p = P({target}) - if p==0: - # The "radicalization trick" implemented below does not work if the target polynomial p is 0 - # since it requires substituting 1/p. - print('1;'+serialize_polynomials(len({eq_list})*[P(0)])) - else: - # Implements the trick described in 2.2 of arxiv.org/pdf/1007.3615.pdf - # for testing membership in the radical. - gens = {eq_list} + [1 - p*aux] - I = P.ideal(gens) - coeffs = P(1).lift(I) - power = max(cf.degree(aux) for cf in coeffs) - coeffs = [P(cf.subs(aux = 1/p)*p^power) for cf in coeffs[:int(-1)]] - print(str(power)+';'+serialize_polynomials(coeffs)) -else: - # workaround for a Sage shortcoming with `n_vars = 0`, - # `TypeError: no conversion of this ring to a Singular ring defined` - # In this case, there is no need to look for membership in the *radical*; - # we just check for membership in the ideal, and return exponent 1 - # if coefficients are found. - P = PolynomialRing({type_str(type)}, 'var', 1) - p = P({target}) - I = P.ideal({eq_list}) - coeffs = p.lift(I) - print('1;'+serialize_polynomials(coeffs)) -''' - return query - -class EvaluationError(Exception): - def __init__(self, ename, evalue, message='Error in Sage communication'): - self.ename = ename - self.evalue = evalue - self.message = message - super().__init__(self.message) - -def parse_response(resp: str) -> Dict[str, Any]: - exp, data = resp.split(';', 1) - return dict(power=int(exp), coeffs=json.loads(data)) - - -def evaluate_in_sage(query: str) -> Dict[str, Any]: - data = urllib.parse.urlencode({'code': query}).encode('utf-8') - headers = {'Content-Type': 'application/x-www-form-urlencoded', - 'User-Agent': 'LeanProver (https://leanprover-community.github.io/)'} - req = urllib.request.Request('https://sagecell.sagemath.org/service', data=data, headers=headers) - with urllib.request.urlopen(req) as response: - response_data = response.read().decode() - response = json.loads(response_data) - if response['success']: - return parse_response(response.get('stdout')) - elif 'execute_reply' in response and 'ename' in response['execute_reply'] and 'evalue' in response['execute_reply']: - raise EvaluationError(response['execute_reply']['ename'], response['execute_reply']['evalue']) - else: - raise Exception(response) - -def main(): - '''The system args contain the following: - 0 - the path to this python file - 1 - a string containing "true" or "false" depending on whether polyrith was called with trace enabled - 2 - a string representing the base type of the target - 3 - the number of variables used - 4 - a list of the polynomial hypotheses/proof terms in terms of the variables - 5 - a single polynomial representing the target - - This returns a json object with format: - ``` - { success: bool, - data: Optional[list[str]], - trace: Optional[str], - name: Optional[str], - value: Optional[str] } - ``` - ''' - command = create_query(sys.argv[2], int(sys.argv[3]), sys.argv[4], sys.argv[5]) - final_query = polynomial_formatting_functions + "\n" + command - if sys.argv[1] == 'true': # trace dry run enabled - output = dict(success=True, trace=command) - else: - try: - output = dict(success=True, data=evaluate_in_sage(final_query)) - except EvaluationError as e: - output = dict(success=False, name=e.ename, value=e.evalue) - print(json.dumps(output)) - -if __name__ == "__main__": - main() diff --git a/scripts/polyrith_sage_helper.py b/scripts/polyrith_sage_helper.py index 44f05122530fb..78355df4754dc 100644 --- a/scripts/polyrith_sage_helper.py +++ b/scripts/polyrith_sage_helper.py @@ -1,13 +1,44 @@ # this file will be run by the remote sage server, so should not import local files. -from typing import Iterable +import json -def q_arr(coeff: QQ) -> str: - return "[" + str(coeff.numerator()) + "," + str(coeff.denominator()) + "]" +def q_arr(coeff: QQ) -> list[int]: + return [int(coeff.numerator()), int(coeff.denominator())] -def arr(args: Iterable[str]) -> str: - return "[" + ",".join(args) + "]" +def serialize_polynomials(power, coeffs) -> str: + return json.dumps({'power': int(power), 'coeffs': [ + [[[[int(t[0]), int(t[1])] for t in etuple.sparse_iter()], q_arr(coeff)] + for etuple, coeff in c.dict().items()] for c in coeffs + ]}) -def serialize_polynomials(coeffs) -> str: - return arr( - arr(arr([arr(arr([str(t[0]), str(t[1])]) for t in etuple.sparse_iter()), q_arr(coeff)]) - for etuple, coeff in c.dict().items()) for c in coeffs) +def main(base_ring, atoms: int, make_hyps, make_target): + if atoms != 0: + vars_str = ['var' + str(i) for i in range(atoms)] + ['aux'] + P = PolynomialRing(base_ring, vars_str) + *vars, aux = P.gens() + hyps = make_hyps(vars) + p = P(make_target(vars)) + if p == 0: + # The 'radicalization trick' implemented below does not work if + # the target polynomial p is 0 since it requires substituting 1/p. + print(serialize_polynomials(1, len(hyps)*[P(0)])) + else: + # Implements the trick described in 2.2 of arxiv.org/pdf/1007.3615.pdf + # for testing membership in the radical. + gens = hyps + [1 - p*aux] + I = P.ideal(gens) + coeffs = P(1).lift(I) + power = max(cf.degree(aux) for cf in coeffs) + coeffs = [P(cf.subs(aux = 1/p)*p^power) for cf in coeffs[:int(-1)]] + print(serialize_polynomials(power, coeffs)) + else: + # workaround for a Sage shortcoming with `atoms = 0`, + # `TypeError: no conversion of this ring to a Singular ring defined` + # In this case, there is no need to look for membership in the *radical*; + # we just check for membership in the ideal, and return exponent 1 + # if coefficients are found. + P = PolynomialRing(base_ring, 'var', 1) + hyps = make_hyps([]) + p = P(make_target([])) + I = P.ideal(hyps) + coeffs = p.lift(I) + print(serialize_polynomials(1, coeffs)) diff --git a/scripts/yaml_check.py b/scripts/yaml_check.py index d9e9a51cfc842..cc375456c8c6f 100644 --- a/scripts/yaml_check.py +++ b/scripts/yaml_check.py @@ -58,7 +58,7 @@ class HundredTheorem: # like |decl|, but a list of declarations (if one theorem is split into multiple declarations) (optional) decls: Optional[List[str]] = None # name(s) of the author(s) of this formalization (optional) - author: Optional[str] = None + authors: Optional[str] = None # Date of the formalization, in the form `YYYY`, `YYYY-MM` or `YYYY-MM-DD` (optional) date: Optional[str] = None links: Optional[Mapping[str, str]] = None @@ -87,7 +87,7 @@ class ThousandPlusTheorem: # like |decl|, but a list of declarations (if one theorem is split into multiple declarations) (optional) decls: Optional[List[str]] = None # name(s) of the author(s) of this formalization (optional) - author: Optional[str] = None + authors: Optional[str] = None # Date of the formalization, in the form `YYYY`, `YYYY-MM` or `YYYY-MM-DD` (optional) date: Optional[str] = None # for external projects, an URL referring to the result diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index de3e5d3ca70c3..fbc811d1671a9 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -66,11 +66,13 @@ has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) has_awaiting_author = any(reaction['emoji_name'] == 'writing' for reaction in reactions) + has_closed = any(reaction['emoji_name'] == 'closed-pr' for reaction in reactions) match = pr_pattern.search(content) if match: print(f"matched: '{message}'") # removing previous emoji reactions + # if the emoji is a custom emoji, add the fields `emoji_code` and `reaction_type` as well print("Removing previous reactions, if present.") if has_peace_sign: print('Removing peace_sign') @@ -102,6 +104,15 @@ "emoji_name": "writing" }) print(f"result: '{result}'") + if has_closed: + print('Removing closed-pr') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "closed-pr", + "emoji_code": "61293", # 61282 was the earlier version of the emoji + "reaction_type": "realm_emoji", + }) + print(f"result: '{result}'") # applying appropriate emoji reaction @@ -124,6 +135,12 @@ "message_id": message['id'], "emoji_name": "writing" }) + elif LABEL == 'closed': + print('adding closed-pr') + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "closed-pr" + }) elif LABEL == 'unlabeled': print('awaiting-author removed') # the reaction was already removed.