diff --git a/.github/workflows/lint_and_suggest_pr.yml b/.github/workflows/lint_and_suggest_pr.yml index ed118d591859f..64f06540402cd 100644 --- a/.github/workflows/lint_and_suggest_pr.yml +++ b/.github/workflows/lint_and_suggest_pr.yml @@ -21,8 +21,15 @@ jobs: with: python-version: 3.8 + - name: install elan + run: | + set -o pipefail + curl -sSfL https://github.com/leanprover/elan/releases/download/v3.1.1/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz + ./elan-init -y --default-toolchain none + echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" + - name: lint - continue-on-error: true + continue-on-error: true # allows the following `reviewdog` step to add GitHub suggestions run: | ./scripts/lint-style.sh --fix @@ -37,7 +44,7 @@ jobs: sudo apt-get install -y bibtool - name: lint references.bib - continue-on-error: true + continue-on-error: true # allows the following `reviewdog` step to add GitHub suggestions run: | ./scripts/lint-bib.sh @@ -65,6 +72,7 @@ jobs: echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - name: update {Mathlib, Tactic, Counterexamples, Archive}.lean + continue-on-error: true # allows the following `reviewdog` step to add GitHub suggestions run: lake exe mk_all - name: suggester / import list diff --git a/.github/workflows/nightly_bump_toolchain.yml b/.github/workflows/nightly_bump_toolchain.yml index 0a8ce5492ff5e..ed8cd271a71b1 100644 --- a/.github/workflows/nightly_bump_toolchain.yml +++ b/.github/workflows/nightly_bump_toolchain.yml @@ -1,8 +1,9 @@ name: Bump lean-toolchain on nightly-testing on: + workflow_dispatch: schedule: - - cron: '0 10 * * *' # 11AM CET/2AM PT, 3 hours after lean4 starts building the nightly. + - cron: '0 10/3 * * *' # Run every three hours, starting at 11AM CET/2AM PT. jobs: update-toolchain: @@ -30,5 +31,6 @@ jobs: git config user.name "leanprover-community-mathlib4-bot" git config user.email "leanprover-community-mathlib4-bot@users.noreply.github.com" git add lean-toolchain - git commit -m "chore: bump to ${RELEASE_TAG}" + # Don't fail if there's nothing to commit + git commit -m "chore: bump to ${RELEASE_TAG}" || true git push origin nightly-testing diff --git a/Archive.lean b/Archive.lean index cc2f9d415c61a..df1f30db3e121 100644 --- a/Archive.lean +++ b/Archive.lean @@ -40,6 +40,8 @@ import Archive.Imo.Imo2019Q2 import Archive.Imo.Imo2019Q4 import Archive.Imo.Imo2020Q2 import Archive.Imo.Imo2021Q1 +import Archive.Imo.Imo2024Q1 +import Archive.Imo.Imo2024Q2 import Archive.Imo.Imo2024Q6 import Archive.MiuLanguage.Basic import Archive.MiuLanguage.DecisionNec diff --git a/Archive/Imo/Imo1972Q5.lean b/Archive/Imo/Imo1972Q5.lean index 81dd6603874e6..288c9e77dabdf 100644 --- a/Archive/Imo/Imo1972Q5.lean +++ b/Archive/Imo/Imo1972Q5.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Ruben Van de Velde, Stanislas Polu -/ import Mathlib.Data.Real.Basic -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # IMO 1972 Q5 diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index 7af3a289d5446..ac4e0dfb0f196 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -116,7 +116,7 @@ theorem A_fibre_over_contestant_card (c : C) : apply Finset.card_image_of_injOn unfold Set.InjOn rintro ⟨a, p⟩ h ⟨a', p'⟩ h' rfl - aesop + aesop (add simp AgreedTriple.contestant) 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 diff --git a/Archive/Imo/Imo2024Q1.lean b/Archive/Imo/Imo2024Q1.lean new file mode 100644 index 0000000000000..74975bfd430bd --- /dev/null +++ b/Archive/Imo/Imo2024Q1.lean @@ -0,0 +1,184 @@ +/- +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.Algebra.BigOperators.Intervals +import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Algebra.Order.ToIntervalMod +import Mathlib.Data.Real.Archimedean +import Mathlib.Tactic.Peel + +/-! +# IMO 2024 Q1 + +Determine all real numbers $\alpha$ such that, for every positive integer $n$, the integer +\[ +\lfloor \alpha \rfloor + \lfloor 2\alpha \rfloor + \cdots + \lfloor n\alpha \rfloor +\] +is a multiple of~$n$. + +We follow Solution 3 from the +[official solutions](https://www.imo2024.uk/s/IMO-2024-Paper-1-Solutions.pdf). First reducing +modulo 2, any answer that is not a multiple of 2 is inductively shown to be contained in a +decreasing sequence of intervals, with empty intersection. +-/ + + +namespace Imo2024Q1 + +/-- The condition of the problem. -/ +def Condition (α : ℝ) : Prop := (∀ n : ℕ, 0 < n → (n : ℤ) ∣ ∑ i ∈ Finset.Icc 1 n, ⌊i * α⌋) + +lemma condition_two_mul_int (m : ℤ) : Condition (2 * m) := by + rintro n - + suffices (n : ℤ) ∣ ∑ i ∈ Finset.Icc 0 n, ⌊((i * (2 * m) : ℤ) : ℝ)⌋ by + rw [← Nat.Icc_insert_succ_left n.zero_le, Finset.sum_insert_zero (by norm_num)] at this + exact_mod_cast this + simp_rw [Int.floor_intCast, ← Finset.sum_mul, ← Nat.Ico_succ_right, ← Finset.range_eq_Ico, + ← mul_assoc] + refine dvd_mul_of_dvd_left ?_ _ + rw [← Nat.cast_sum, ← Nat.cast_ofNat (n := 2), ← Nat.cast_mul, Finset.sum_range_id_mul_two] + simp + +lemma condition_sub_two_mul_int_iff {α : ℝ} (m : ℤ) : Condition (α - 2 * m) ↔ Condition α := by + unfold Condition + peel with n hn + refine dvd_iff_dvd_of_dvd_sub ?_ + simp_rw [← Finset.sum_sub_distrib, mul_sub] + norm_cast + simp_rw [Int.floor_sub_int, sub_sub_cancel_left] + convert condition_two_mul_int (-m) n hn + norm_cast + rw [Int.floor_intCast] + simp + +lemma condition_toIcoMod_iff {α : ℝ} : + Condition (toIcoMod (by norm_num : (0 : ℝ) < 2) 0 α) ↔ Condition α := by + rw [toIcoMod, zsmul_eq_mul, mul_comm, condition_sub_two_mul_int_iff] + +namespace Condition + +variable {α : ℝ} (hc : Condition α) + +lemma mem_Ico_one_of_mem_Ioo (h : α ∈ Set.Ioo 0 2) : α ∈ Set.Ico 1 2 := by + rcases h with ⟨h0, h2⟩ + refine ⟨?_, h2⟩ + by_contra! hn + have hr : 1 < ⌈α⁻¹⌉₊ := by + rw [Nat.lt_ceil] + exact_mod_cast one_lt_inv h0 hn + replace hc := hc ⌈α⁻¹⌉₊ (zero_lt_one.trans hr) + refine hr.ne' ?_ + suffices ⌈α⁻¹⌉₊ = (1 : ℤ) from mod_cast this + refine Int.eq_one_of_dvd_one (Int.zero_le_ofNat _) ?_ + convert hc + rw [← Finset.add_sum_Ico_eq_sum_Icc hr.le] + convert (add_zero _).symm + · rw [Int.floor_eq_iff] + refine ⟨?_, ?_⟩ + · rw [Int.cast_one] + calc 1 ≤ α⁻¹ * α := by simp [h0.ne'] + _ ≤ ⌈α⁻¹⌉₊ * α := by gcongr; exact Nat.le_ceil _ + · calc ⌈α⁻¹⌉₊ * α < (α⁻¹ + 1) * α := by gcongr; exact Nat.ceil_lt_add_one (inv_nonneg.2 h0.le) + _ = 1 + α := by field_simp [h0.ne'] + _ ≤ (1 : ℕ) + 1 := by gcongr; norm_cast + · refine Finset.sum_eq_zero ?_ + intro x hx + rw [Int.floor_eq_zero_iff] + refine ⟨by positivity, ?_⟩ + rw [Finset.mem_Ico, Nat.lt_ceil] at hx + calc x * α < α⁻¹ * α := by gcongr; exact hx.2 + _ ≤ 1 := by simp [h0.ne'] + +lemma mem_Ico_n_of_mem_Ioo (h : α ∈ Set.Ioo 0 2) + {n : ℕ} (hn : 0 < n) : α ∈ Set.Ico ((2 * n - 1) / n : ℝ) 2 := by + suffices ∑ i ∈ Finset.Icc 1 n, ⌊i * α⌋ = n ^ 2 ∧ α ∈ Set.Ico ((2 * n - 1) / n : ℝ) 2 from this.2 + induction' n, hn using Nat.le_induction with k kpos hk + · obtain ⟨h1, h2⟩ := hc.mem_Ico_one_of_mem_Ioo h + simp only [zero_add, Finset.Icc_self, Finset.sum_singleton, Nat.cast_one, one_mul, one_pow, + Int.floor_eq_iff, Int.cast_one, mul_one, div_one, Set.mem_Ico, tsub_le_iff_right] + exact ⟨⟨h1, by linarith⟩, by linarith, h2⟩ + · rcases hk with ⟨hks, hkl, hk2⟩ + have hs : (∑ i ∈ Finset.Icc 1 (k + 1), ⌊i * α⌋) = + ⌊(k + 1 : ℕ) * α⌋ + ((k : ℕ) : ℤ) ^ 2 := by + have hn11 : k + 1 ∉ Finset.Icc 1 k := by + rw [Finset.mem_Icc] + omega + rw [← Nat.Icc_insert_succ_right (Nat.le_add_left 1 k), Finset.sum_insert hn11, hks] + replace hc := hc (k + 1) k.succ_pos + rw [hs] at hc ⊢ + have hkl' : 2 * k ≤ ⌊(k + 1 : ℕ) * α⌋ := by + rw [Int.le_floor] + calc ((2 * k : ℤ) : ℝ) = ((2 * k : ℤ) : ℝ) + 0 := (add_zero _).symm + _ ≤ ((2 * k : ℤ) : ℝ) + (k - 1) / k := by gcongr; norm_cast; positivity + _ = (k + 1 : ℕ) * ((2 * (k : ℕ) - 1) / ((k : ℕ) : ℝ) : ℝ) := by + field_simp + ring + _ ≤ (k + 1 : ℕ) * α := by gcongr + have hk2' : ⌊(k + 1 : ℕ) * α⌋ < (k + 1 : ℕ) * 2 := by + rw [Int.floor_lt] + push_cast + gcongr + have hk : ⌊(k + 1 : ℕ) * α⌋ = 2 * k ∨ ⌊(k + 1 : ℕ) * α⌋ = 2 * k + 1 := by omega + have hk' : ⌊(k + 1 : ℕ) * α⌋ = 2 * k + 1 := by + rcases hk with hk | hk + · rw [hk] at hc + have hc' : ((k + 1 : ℕ) : ℤ) ∣ ((k + 1 : ℕ) : ℤ) * ((k + 1 : ℕ) : ℤ) - 1 := by + convert hc using 1 + push_cast + ring + rw [dvd_sub_right (dvd_mul_right _ _), ← isUnit_iff_dvd_one, Int.isUnit_iff] at hc' + omega + · exact hk + rw [hk'] + refine ⟨?_, ?_, h.2⟩ + · push_cast + ring + · rw [Int.floor_eq_iff] at hk' + rw [div_le_iff (by norm_cast; omega), mul_comm α] + convert hk'.1 + push_cast + ring + +end Condition + +lemma not_condition_of_mem_Ioo {α : ℝ} (h : α ∈ Set.Ioo 0 2) : ¬Condition α := by + intro hc + let n : ℕ := ⌊(2 - α)⁻¹⌋₊ + 1 + have hn : 0 < n := by omega + have hna := (hc.mem_Ico_n_of_mem_Ioo h hn).1 + rcases h with ⟨-, h2⟩ + have hna' : 2 - (n : ℝ)⁻¹ ≤ α := by + convert hna using 1 + field_simp + rw [sub_eq_add_neg, ← le_sub_iff_add_le', neg_le, neg_sub] at hna' + rw [le_inv (by linarith) (mod_cast hn), ← not_lt] at hna' + apply hna' + exact_mod_cast Nat.lt_floor_add_one (_ : ℝ) + +lemma condition_iff_of_mem_Ico {α : ℝ} (h : α ∈ Set.Ico 0 2) : Condition α ↔ α = 0 := by + refine ⟨?_, ?_⟩ + · intro hc + rcases Set.eq_left_or_mem_Ioo_of_mem_Ico h with rfl | ho + · rfl + · exact False.elim (not_condition_of_mem_Ioo ho hc) + · rintro rfl + convert condition_two_mul_int 0 + norm_num + +/-- This is to be determined by the solver of the original problem. -/ +def solutionSet : Set ℝ := {α : ℝ | ∃ m : ℤ, α = 2 * m} + +theorem result (α : ℝ) : Condition α ↔ α ∈ solutionSet := by + refine ⟨fun h ↦ ?_, ?_⟩ + · rw [← condition_toIcoMod_iff, condition_iff_of_mem_Ico (toIcoMod_mem_Ico' _ _), + ← AddCommGroup.modEq_iff_toIcoMod_eq_left, AddCommGroup.ModEq] at h + simp_rw [sub_zero] at h + rcases h with ⟨m, rfl⟩ + rw [zsmul_eq_mul, mul_comm] + simp [solutionSet] + · rintro ⟨m, rfl⟩ + exact condition_two_mul_int m + +end Imo2024Q1 diff --git a/Archive/Imo/Imo2024Q2.lean b/Archive/Imo/Imo2024Q2.lean new file mode 100644 index 0000000000000..0801f9425db8f --- /dev/null +++ b/Archive/Imo/Imo2024Q2.lean @@ -0,0 +1,190 @@ +/- +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.FieldTheory.Finite.Basic + +/-! +# IMO 2024 Q2 + +Determine all pairs $(a,b)$ of positive integers for which there exist positive integers +$g$ and $N$ such that +\[ \gcd(a^n + b, b^n + a) = g \] +holds for all integers $n \ge N$. + +We consider the sequence modulo `ab+1`; if the exponent is `-1` modulo `φ(ab+1)`, the terms +are zero modulo `ab+1`, so `ab+1` divides `g`, and all sufficiently large terms, so all terms, +from which we conclude that `a=b=1`. +-/ + + +namespace Imo2024Q2 + +open scoped Nat + +/-- The condition of the problem. -/ +def Condition (a b : ℕ) : Prop := + 0 < a ∧ 0 < b ∧ ∃ g N : ℕ, 0 < g ∧ 0 < N ∧ ∀ n : ℕ, N ≤ n → Nat.gcd (a ^ n + b) (b ^ n + a) = g + +lemma dvd_pow_iff_of_dvd_sub {a b d n : ℕ} {z : ℤ} (ha : a.Coprime d) + (hd : (φ d : ℤ) ∣ (n : ℤ) - z) : + d ∣ a ^ n + b ↔ (((ZMod.unitOfCoprime _ ha) ^ z : (ZMod d)ˣ) : ZMod d) + b = 0 := by + rcases hd with ⟨k, hk⟩ + rw [← ZMod.natCast_zmod_eq_zero_iff_dvd] + convert Iff.rfl + push_cast + congr + suffices (((ZMod.unitOfCoprime _ ha) ^ z : (ZMod d)ˣ) : ZMod d) = + (((ZMod.unitOfCoprime _ ha) ^ (n : ℤ) : (ZMod d)ˣ) : ZMod d) by + convert this + rw [sub_eq_iff_eq_add] at hk + rw [hk, zpow_add, zpow_mul] + norm_cast + rw [ZMod.pow_totient, one_zpow, one_mul] + +namespace Condition + +variable {a b : ℕ} (h : Condition a b) + +lemma a_pos : 0 < a := h.1 + +lemma b_pos : 0 < b := h.2.1 + +/-- The value of `g` in the problem (determined by `a` and `b`). -/ +noncomputable def g : ℕ := h.2.2.choose + +lemma g_spec : ∃ N : ℕ, 0 < h.g ∧ 0 < N ∧ ∀ n : ℕ, N ≤ n → Nat.gcd (a ^ n + b) (b ^ n + a) = h.g := + h.2.2.choose_spec + +/-- The value of `N` in the problem (any sufficiently large value). -/ +noncomputable def N : ℕ := h.g_spec.choose + +lemma N_spec : 0 < h.g ∧ 0 < h.N ∧ ∀ n : ℕ, h.N ≤ n → Nat.gcd (a ^ n + b) (b ^ n + a) = h.g := + h.g_spec.choose_spec + +lemma g_pos : 0 < h.g := h.N_spec.1 + +lemma N_pos : 0 < h.N := h.N_spec.2.1 + +lemma gcd_eq_g {n : ℕ} (hn : h.N ≤ n) : Nat.gcd (a ^ n + b) (b ^ n + a) = h.g := h.N_spec.2.2 n hn + +protected lemma symm : Condition b a := by + refine ⟨h.b_pos, h.a_pos, h.g, h.N, h.g_pos, h.N_pos, fun n hn ↦ ?_⟩ + rw [Nat.gcd_comm] + exact h.gcd_eq_g hn + +lemma dvd_g_of_le_N_of_dvd {n : ℕ} (hn : h.N ≤ n) {d : ℕ} (hab : d ∣ a ^ n + b) + (hba : d ∣ b ^ n + a) : d ∣ h.g := by + rw [← h.gcd_eq_g hn, Nat.dvd_gcd_iff] + exact ⟨hab, hba⟩ + +lemma a_coprime_ab_add_one : a.Coprime (a * b + 1) := by + simp + +/-- A sufficiently large value of n, congruent to `-1` mod `φ (a * b + 1)`. -/ +noncomputable def large_n : ℕ := (max h.N h.symm.N + 1) * φ (a * b + 1) - 1 + +lemma symm_large_n : h.symm.large_n = h.large_n := by + simp_rw [large_n] + congr 2 + · rw [max_comm] + · rw [mul_comm] + +lemma N_le_large_n : h.N ≤ h.large_n := by + have hp : 0 < φ (a * b + 1) := Nat.totient_pos.2 (Nat.add_pos_right _ zero_lt_one) + rw [large_n, add_mul, one_mul, Nat.add_sub_assoc (Nat.one_le_of_lt hp)] + suffices h.N ≤ h.N * φ (a * b + 1) + (φ (a * b + 1) - 1) by + refine this.trans ?_ + gcongr + simp + exact Nat.le_add_right_of_le (Nat.le_mul_of_pos_right _ hp) + +lemma dvd_large_n_sub_neg_one : (φ (a * b + 1) : ℤ) ∣ (h.large_n : ℤ) - (-1 : ℤ) := by + simp [large_n] + +/-- A sufficiently large value of n, congruent to `0` mod `φ (a * b + 1)`. -/ +noncomputable def large_n_0 : ℕ := (max h.N h.symm.N) * φ (a * b + 1) + +lemma symm_large_n_0 : h.symm.large_n_0 = h.large_n_0 := by + simp_rw [large_n_0] + congr 1 + · rw [max_comm] + · rw [mul_comm] + +lemma N_le_large_n_0 : h.N ≤ h.large_n_0 := by + have hp : 0 < φ (a * b + 1) := Nat.totient_pos.2 (Nat.add_pos_right _ zero_lt_one) + rw [large_n_0] + suffices h.N ≤ h.N * φ (a * b + 1) by + refine this.trans ?_ + gcongr + simp + exact Nat.le_mul_of_pos_right _ hp + +lemma dvd_large_n_0_sub_zero : (φ (a * b + 1) : ℤ) ∣ (h.large_n_0 : ℤ) - (0 : ℤ) := by + simp [large_n_0] + +lemma ab_add_one_dvd_a_pow_large_n_add_b : a * b + 1 ∣ a ^ h.large_n + b := by + rw [dvd_pow_iff_of_dvd_sub a_coprime_ab_add_one h.dvd_large_n_sub_neg_one, zpow_neg, zpow_one] + suffices ((ZMod.unitOfCoprime _ a_coprime_ab_add_one : (ZMod (a * b + 1))ˣ) : ZMod (a * b + 1)) * + ((((ZMod.unitOfCoprime _ a_coprime_ab_add_one)⁻¹ : (ZMod (a * b + 1))ˣ) : + ZMod (a * b + 1)) + (b : ZMod (a * b + 1))) = 0 from + (IsUnit.mul_right_eq_zero (ZMod.unitOfCoprime _ a_coprime_ab_add_one).isUnit).1 this + rw [mul_add] + norm_cast + simp only [mul_right_inv, Units.val_one, ZMod.coe_unitOfCoprime] + norm_cast + convert ZMod.natCast_self (a * b + 1) using 2 + exact add_comm _ _ + +lemma ab_add_one_dvd_b_pow_large_n_add_a : a * b + 1 ∣ b ^ h.large_n + a := by + convert h.symm.ab_add_one_dvd_a_pow_large_n_add_b using 1 + · rw [mul_comm] + · rw [h.symm_large_n] + +lemma ab_add_one_dvd_g : a * b + 1 ∣ h.g := + h.dvd_g_of_le_N_of_dvd h.N_le_large_n h.ab_add_one_dvd_a_pow_large_n_add_b + h.ab_add_one_dvd_b_pow_large_n_add_a + +lemma ab_add_one_dvd_a_pow_large_n_0_add_b : a * b + 1 ∣ a ^ h.large_n_0 + b := by + refine h.ab_add_one_dvd_g.trans ?_ + rw [← h.gcd_eq_g h.N_le_large_n_0] + exact Nat.gcd_dvd_left _ _ + +lemma ab_add_one_dvd_b_add_one : a * b + 1 ∣ b + 1 := by + rw [add_comm b] + suffices a * b + 1 ∣ a ^ h.large_n_0 + b by + rw [dvd_pow_iff_of_dvd_sub a_coprime_ab_add_one h.dvd_large_n_0_sub_zero, zpow_zero] at this + rw [← ZMod.natCast_zmod_eq_zero_iff_dvd] + push_cast + norm_cast at this + exact h.ab_add_one_dvd_a_pow_large_n_0_add_b + +lemma a_eq_one : a = 1 := by + have hle : a * b + 1 ≤ b + 1 := Nat.le_of_dvd (by omega) h.ab_add_one_dvd_b_add_one + rw [add_le_add_iff_right] at hle + suffices a ≤ 1 by + have hp := h.a_pos + omega + have hle' : a * b ≤ 1 * b := by + simpa using hle + exact Nat.le_of_mul_le_mul_right hle' h.b_pos + +lemma b_eq_one : b = 1 := h.symm.a_eq_one + +end Condition + +/-- This is to be determined by the solver of the original problem. -/ +def solutionSet : Set (ℕ × ℕ) := {(1, 1)} + +theorem result (a b : ℕ) : Condition a b ↔ (a, b) ∈ solutionSet := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · simp only [solutionSet, Set.mem_singleton_iff, Prod.mk.injEq] + exact ⟨h.a_eq_one, h.b_eq_one⟩ + · simp only [solutionSet, Set.mem_singleton_iff, Prod.mk.injEq] at h + rcases h with ⟨rfl, rfl⟩ + rw [Condition] + refine ⟨by decide, by decide, 2, 1, ?_⟩ + simp + +end Imo2024Q2 diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index f2005282278b2..ba3cc0bc34669 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gihan Marasingha -/ import Archive.MiuLanguage.Basic +import Mathlib.Data.List.Basic import Mathlib.Data.List.Count import Mathlib.Data.Nat.ModEq import Mathlib.Tactic.Ring @@ -71,9 +72,10 @@ theorem count_equiv_one_or_two_mod3_of_derivable (en : Miustr) : any_goals apply mod3_eq_1_or_mod3_eq_2 h_ih -- Porting note: `simp_rw [count_append]` usually doesn't work · left; rw [count_append, count_append]; rfl - · right; simp_rw [count_append, count_cons, if_false, two_mul]; simp + · right; simp_rw [count_append, count_cons, beq_iff_eq, ite_false, add_zero, two_mul] · left; rw [count_append, count_append, count_append] - simp_rw [count_cons_self, count_nil, count_cons, ite_false, add_right_comm, add_mod_right] + simp_rw [count_cons_self, count_nil, count_cons, beq_iff_eq, ite_false, add_right_comm, + add_mod_right] simp · left; rw [count_append, count_append, count_append] simp only [ne_eq, not_false_eq_true, count_cons_of_ne, count_nil, add_zero] @@ -127,7 +129,7 @@ theorem goodm_of_rule1 (xs : Miustr) (h₁ : Derivable (xs ++ ↑[I])) (h₂ : G · change ¬M ∈ tail (xs ++ ↑([I] ++ [U])) rw [← append_assoc, tail_append_singleton_of_ne_nil] · simp_rw [mem_append, mem_singleton, or_false]; exact nmtail - · exact append_ne_nil_of_ne_nil_left this _ + · exact append_ne_nil_of_left_ne_nil this _ theorem goodm_of_rule2 (xs : Miustr) (_ : Derivable (M :: xs)) (h₂ : Goodm (M :: xs)) : Goodm (↑(M :: xs) ++ xs) := by diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 4e460e9939ba2..10eaa8e26942b 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -251,7 +251,7 @@ theorem count_I_eq_length_of_count_U_zero_and_neg_mem {ys : Miustr} (hu : count · -- case `x = M` gives a contradiction. exfalso; exact hm (mem_cons_self M xs) · -- case `x = I` - rw [count_cons, if_pos rfl, length, succ_inj'] + rw [count_cons, beq_self_eq_true, if_pos rfl, length, succ_inj'] apply hxs · simpa only [count] · rw [mem_cons, not_or] at hm; exact hm.2 @@ -307,7 +307,9 @@ theorem ind_hyp_suf (k : ℕ) (ys : Miustr) (hu : count U ys = succ k) (hdec : D use as, bs refine ⟨rfl, ?_, ?_, ?_⟩ · -- Porting note: `simp_rw [count_append]` didn't work - rw [count_append] at hu; simp_rw [count_cons, if_true, add_succ, succ_inj'] at hu + rw [count_append] at hu + simp_rw [count_cons, beq_self_eq_true, if_true, add_succ, beq_iff_eq, reduceIte, add_zero, + succ_inj'] at hu rwa [count_append, count_append] · apply And.intro rfl rw [cons_append, cons_append] diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index ea112d84d5fcd..08871a933bd8c 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -8,7 +8,7 @@ import Mathlib.Tactic.FinCases import Mathlib.Tactic.ApplyFun import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.LinearAlgebra.Dual -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Data.Real.Sqrt /-! diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index ee61e72761e35..0d0bde8da3d54 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.Analysis.Calculus.LocalExtr.Polynomial -import Mathlib.Analysis.Complex.Polynomial +import Mathlib.Analysis.Complex.Polynomial.Basic import Mathlib.FieldTheory.AbelRuffini import Mathlib.RingTheory.RootsOfUnity.Minpoly import Mathlib.RingTheory.EisensteinCriterion diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index b037dd8fab914..8b45732a49ddb 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -115,18 +115,19 @@ theorem counted_succ_succ (p q : ℕ) : obtain ⟨hl₀, hl₁, hl₂⟩ := hl obtain hlast | hlast := hl₂ (l.head hlnil) (List.head_mem hlnil) · refine Or.inl ⟨l.tail, ⟨?_, ?_, ?_⟩, ?_⟩ - · rw [List.count_tail l 1 hlnil, hl₀, hlast, if_pos rfl, Nat.add_sub_cancel] + · rw [List.count_tail l 1 hlnil, hl₀, hlast, beq_self_eq_true, if_pos rfl, Nat.add_sub_cancel] · rw [List.count_tail l (-1) hlnil, hl₁, hlast, if_neg (by decide), Nat.sub_zero] · exact fun x hx => hl₂ x (List.mem_of_mem_tail hx) · rw [← hlast, List.head_cons_tail] · refine Or.inr ⟨l.tail, ⟨?_, ?_, ?_⟩, ?_⟩ · rw [List.count_tail l 1 hlnil, hl₀, hlast, if_neg (by decide), Nat.sub_zero] - · rw [List.count_tail l (-1) hlnil, hl₁, hlast, if_pos rfl, Nat.add_sub_cancel] + · rw [List.count_tail l (-1) hlnil, hl₁, hlast, beq_self_eq_true, if_pos rfl, + Nat.add_sub_cancel] · exact fun x hx => hl₂ x (List.mem_of_mem_tail hx) · rw [← hlast, List.head_cons_tail] · rintro (⟨t, ⟨ht₀, ht₁, ht₂⟩, rfl⟩ | ⟨t, ⟨ht₀, ht₁, ht₂⟩, rfl⟩) · refine ⟨?_, ?_, ?_⟩ - · rw [List.count_cons, if_pos rfl, ht₀] + · rw [List.count_cons, beq_self_eq_true, if_pos rfl, ht₀] · rw [List.count_cons, if_neg, ht₁] norm_num · rintro x (_ | _) @@ -134,7 +135,7 @@ theorem counted_succ_succ (p q : ℕ) : · refine ⟨?_, ?_, ?_⟩ · rw [List.count_cons, if_neg, ht₀] norm_num - · rw [List.count_cons, if_pos rfl, ht₁] + · rw [List.count_cons, beq_self_eq_true, if_pos rfl, ht₁] · rintro x (_ | _) exacts [Or.inr rfl, ht₂ x (by tauto)] diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 0179465f55758..38232baf5bdab 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -235,7 +235,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) rw [← hw₂] exact dvd_mul_left _ _ · intro i - simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq] + simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq'] split_ifs with h · rcases hf₄ i h with ⟨w, hw₁, hw₂⟩ rwa [← hw₂, Nat.mul_div_cancel _ (hs i h)] @@ -245,7 +245,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) rcases hi with ⟨j, hj₁, hj₂⟩ rwa [Multiset.eq_of_mem_replicate hj₂] · ext i - simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq] + simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq'] simp only [ne_eq, Multiset.mem_toFinset, not_not, smul_eq_mul, ite_mul, zero_mul, Finsupp.coe_mk] split_ifs with h diff --git a/Mathlib.lean b/Mathlib.lean index 2cffde07c87f8..e4e32c10ec3af 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -49,6 +49,7 @@ import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.BigOperators.Ring.List import Mathlib.Algebra.BigOperators.Ring.Multiset +import Mathlib.Algebra.BigOperators.Ring.Nat import Mathlib.Algebra.BigOperators.RingEquiv import Mathlib.Algebra.BigOperators.WithTop import Mathlib.Algebra.Bounds @@ -291,6 +292,9 @@ import Mathlib.Algebra.Group.ZeroOne import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.Algebra.GroupWithZero.Action.Opposite +import Mathlib.Algebra.GroupWithZero.Action.Pi +import Mathlib.Algebra.GroupWithZero.Action.Prod +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.GroupWithZero.Basic import Mathlib.Algebra.GroupWithZero.Center import Mathlib.Algebra.GroupWithZero.Commute @@ -334,6 +338,7 @@ import Mathlib.Algebra.Homology.Embedding.Boundary import Mathlib.Algebra.Homology.Embedding.Extend import Mathlib.Algebra.Homology.Embedding.IsSupported import Mathlib.Algebra.Homology.Embedding.Restriction +import Mathlib.Algebra.Homology.Embedding.TruncGE import Mathlib.Algebra.Homology.ExactSequence import Mathlib.Algebra.Homology.Factorizations.Basic import Mathlib.Algebra.Homology.Functor @@ -564,6 +569,7 @@ import Mathlib.Algebra.Order.Group.PosPart import Mathlib.Algebra.Order.Group.Prod import Mathlib.Algebra.Order.Group.Synonym import Mathlib.Algebra.Order.Group.TypeTags +import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Algebra.Order.Group.Unbundled.Basic import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.GroupWithZero.Canonical @@ -733,6 +739,7 @@ import Mathlib.Algebra.Ring.Prod import Mathlib.Algebra.Ring.Rat import Mathlib.Algebra.Ring.Regular import Mathlib.Algebra.Ring.Semiconj +import Mathlib.Algebra.Ring.Semireal.Defs import Mathlib.Algebra.Ring.Subring.Basic import Mathlib.Algebra.Ring.Subring.MulOpposite import Mathlib.Algebra.Ring.Subring.Order @@ -742,6 +749,7 @@ import Mathlib.Algebra.Ring.Subsemiring.Basic import Mathlib.Algebra.Ring.Subsemiring.MulOpposite import Mathlib.Algebra.Ring.Subsemiring.Order import Mathlib.Algebra.Ring.Subsemiring.Pointwise +import Mathlib.Algebra.Ring.SumsOfSquares import Mathlib.Algebra.Ring.ULift import Mathlib.Algebra.Ring.Units import Mathlib.Algebra.Ring.WithZero @@ -791,6 +799,7 @@ import Mathlib.AlgebraicGeometry.Limits import Mathlib.AlgebraicGeometry.Modules.Presheaf import Mathlib.AlgebraicGeometry.Modules.Sheaf import Mathlib.AlgebraicGeometry.Modules.Tilde +import Mathlib.AlgebraicGeometry.Morphisms.Affine import Mathlib.AlgebraicGeometry.Morphisms.Basic import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion import Mathlib.AlgebraicGeometry.Morphisms.Constructors @@ -810,6 +819,7 @@ import Mathlib.AlgebraicGeometry.PrimeSpectrum.IsOpenComapC import Mathlib.AlgebraicGeometry.PrimeSpectrum.Maximal import Mathlib.AlgebraicGeometry.PrimeSpectrum.Module import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian +import Mathlib.AlgebraicGeometry.PrimeSpectrum.TensorProduct import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Scheme import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Topology @@ -872,6 +882,7 @@ import Mathlib.Analysis.Analytic.Meromorphic import Mathlib.Analysis.Analytic.Polynomial import Mathlib.Analysis.Analytic.RadiusLiminf import Mathlib.Analysis.Analytic.Uniqueness +import Mathlib.Analysis.Analytic.Within import Mathlib.Analysis.Asymptotics.AsymptoticEquivalent import Mathlib.Analysis.Asymptotics.Asymptotics import Mathlib.Analysis.Asymptotics.SpecificAsymptotics @@ -986,7 +997,8 @@ import Mathlib.Analysis.Complex.LocallyUniformLimit import Mathlib.Analysis.Complex.OpenMapping import Mathlib.Analysis.Complex.OperatorNorm import Mathlib.Analysis.Complex.PhragmenLindelof -import Mathlib.Analysis.Complex.Polynomial +import Mathlib.Analysis.Complex.Polynomial.Basic +import Mathlib.Analysis.Complex.Polynomial.UnitTrinomial import Mathlib.Analysis.Complex.ReImTopology import Mathlib.Analysis.Complex.RealDeriv import Mathlib.Analysis.Complex.RemovableSingularity @@ -1052,6 +1064,21 @@ import Mathlib.Analysis.Convex.Strong import Mathlib.Analysis.Convex.Topology import Mathlib.Analysis.Convex.Uniform import Mathlib.Analysis.Convolution +import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Basic +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Instances +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Order +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Restrict +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unitary +import Mathlib.Analysis.CstarAlgebra.Exponential +import Mathlib.Analysis.CstarAlgebra.GelfandDuality +import Mathlib.Analysis.CstarAlgebra.Matrix +import Mathlib.Analysis.CstarAlgebra.Multiplier +import Mathlib.Analysis.CstarAlgebra.Spectrum +import Mathlib.Analysis.CstarAlgebra.Unitization import Mathlib.Analysis.Distribution.AEEqOfIntegralContDiff import Mathlib.Analysis.Distribution.FourierSchwartz import Mathlib.Analysis.Distribution.SchwartzSpace @@ -1103,6 +1130,11 @@ import Mathlib.Analysis.MeanInequalities import Mathlib.Analysis.MeanInequalitiesPow import Mathlib.Analysis.MellinInversion import Mathlib.Analysis.MellinTransform +import Mathlib.Analysis.Normed.Algebra.Basic +import Mathlib.Analysis.Normed.Algebra.Spectrum +import Mathlib.Analysis.Normed.Algebra.TrivSqZeroExt +import Mathlib.Analysis.Normed.Algebra.Unitization +import Mathlib.Analysis.Normed.Algebra.UnitizationL1 import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Field.InfiniteSum import Mathlib.Analysis.Normed.Field.UnitBall @@ -1132,35 +1164,45 @@ import Mathlib.Analysis.Normed.Group.Submodule import Mathlib.Analysis.Normed.Group.Tannery import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Analysis.Normed.Group.ZeroAtInfty +import Mathlib.Analysis.Normed.Lp.LpEquiv +import Mathlib.Analysis.Normed.Lp.PiLp +import Mathlib.Analysis.Normed.Lp.ProdLp +import Mathlib.Analysis.Normed.Lp.WithLp +import Mathlib.Analysis.Normed.Lp.lpSpace +import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.Analysis.Normed.Module.Complemented +import Mathlib.Analysis.Normed.Module.Completion +import Mathlib.Analysis.Normed.Module.Dual +import Mathlib.Analysis.Normed.Module.FiniteDimension +import Mathlib.Analysis.Normed.Module.Ray +import Mathlib.Analysis.Normed.Module.Span +import Mathlib.Analysis.Normed.Module.WeakDual import Mathlib.Analysis.Normed.MulAction +import Mathlib.Analysis.Normed.Operator.Banach +import Mathlib.Analysis.Normed.Operator.BanachSteinhaus +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps +import Mathlib.Analysis.Normed.Operator.Compact +import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap +import Mathlib.Analysis.Normed.Operator.LinearIsometry +import Mathlib.Analysis.Normed.Operator.WeakOperatorTopology import Mathlib.Analysis.Normed.Order.Basic import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.Normed.Order.UpperLower import Mathlib.Analysis.Normed.Ring.Seminorm import Mathlib.Analysis.Normed.Ring.SeminormFromBounded +import Mathlib.Analysis.Normed.Ring.Units import Mathlib.Analysis.NormedSpace.AddTorsor import Mathlib.Analysis.NormedSpace.AddTorsorBases import Mathlib.Analysis.NormedSpace.AffineIsometry -import Mathlib.Analysis.NormedSpace.Algebra import Mathlib.Analysis.NormedSpace.BallAction -import Mathlib.Analysis.NormedSpace.Banach -import Mathlib.Analysis.NormedSpace.BanachSteinhaus -import Mathlib.Analysis.NormedSpace.Basic -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps -import Mathlib.Analysis.NormedSpace.CompactOperator -import Mathlib.Analysis.NormedSpace.Complemented -import Mathlib.Analysis.NormedSpace.Completion import Mathlib.Analysis.NormedSpace.ConformalLinearMap import Mathlib.Analysis.NormedSpace.Connected import Mathlib.Analysis.NormedSpace.ContinuousAffineMap -import Mathlib.Analysis.NormedSpace.ContinuousLinearMap -import Mathlib.Analysis.NormedSpace.Dual import Mathlib.Analysis.NormedSpace.DualNumber import Mathlib.Analysis.NormedSpace.ENorm import Mathlib.Analysis.NormedSpace.Exponential import Mathlib.Analysis.NormedSpace.Extend import Mathlib.Analysis.NormedSpace.Extr -import Mathlib.Analysis.NormedSpace.FiniteDimension import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.NormedSpace.HahnBanach.Extension import Mathlib.Analysis.NormedSpace.HahnBanach.SeparatingDual @@ -1168,8 +1210,6 @@ import Mathlib.Analysis.NormedSpace.HahnBanach.Separation import Mathlib.Analysis.NormedSpace.HomeomorphBall import Mathlib.Analysis.NormedSpace.IndicatorFunction import Mathlib.Analysis.NormedSpace.Int -import Mathlib.Analysis.NormedSpace.LinearIsometry -import Mathlib.Analysis.NormedSpace.LpEquiv import Mathlib.Analysis.NormedSpace.MStructure import Mathlib.Analysis.NormedSpace.MatrixExponential import Mathlib.Analysis.NormedSpace.MazurUlam @@ -1183,39 +1223,14 @@ import Mathlib.Analysis.NormedSpace.OperatorNorm.Mul import Mathlib.Analysis.NormedSpace.OperatorNorm.NNNorm import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace import Mathlib.Analysis.NormedSpace.OperatorNorm.Prod -import Mathlib.Analysis.NormedSpace.PiLp import Mathlib.Analysis.NormedSpace.PiTensorProduct.InjectiveSeminorm import Mathlib.Analysis.NormedSpace.PiTensorProduct.ProjectiveSeminorm import Mathlib.Analysis.NormedSpace.Pointwise -import Mathlib.Analysis.NormedSpace.ProdLp import Mathlib.Analysis.NormedSpace.QuaternionExponential import Mathlib.Analysis.NormedSpace.RCLike -import Mathlib.Analysis.NormedSpace.Ray import Mathlib.Analysis.NormedSpace.Real import Mathlib.Analysis.NormedSpace.RieszLemma -import Mathlib.Analysis.NormedSpace.Span -import Mathlib.Analysis.NormedSpace.Spectrum import Mathlib.Analysis.NormedSpace.SphereNormEquiv -import Mathlib.Analysis.NormedSpace.Star.Basic -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus.Instances -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus.Order -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus.Restrict -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus.Unitary -import Mathlib.Analysis.NormedSpace.Star.Exponential -import Mathlib.Analysis.NormedSpace.Star.GelfandDuality -import Mathlib.Analysis.NormedSpace.Star.Matrix -import Mathlib.Analysis.NormedSpace.Star.Multiplier -import Mathlib.Analysis.NormedSpace.Star.Spectrum -import Mathlib.Analysis.NormedSpace.Star.Unitization -import Mathlib.Analysis.NormedSpace.TrivSqZeroExt -import Mathlib.Analysis.NormedSpace.Unitization -import Mathlib.Analysis.NormedSpace.UnitizationL1 -import Mathlib.Analysis.NormedSpace.Units -import Mathlib.Analysis.NormedSpace.WeakDual -import Mathlib.Analysis.NormedSpace.WeakOperatorTopology -import Mathlib.Analysis.NormedSpace.WithLp -import Mathlib.Analysis.NormedSpace.lpSpace import Mathlib.Analysis.ODE.Gronwall import Mathlib.Analysis.ODE.PicardLindelof import Mathlib.Analysis.Oscillation @@ -1768,8 +1783,10 @@ import Mathlib.CategoryTheory.Sites.LeftExact import Mathlib.CategoryTheory.Sites.Limits import Mathlib.CategoryTheory.Sites.Localization import Mathlib.CategoryTheory.Sites.LocallyBijective +import Mathlib.CategoryTheory.Sites.LocallyFullyFaithful import Mathlib.CategoryTheory.Sites.LocallyInjective import Mathlib.CategoryTheory.Sites.LocallySurjective +import Mathlib.CategoryTheory.Sites.MayerVietorisSquare import Mathlib.CategoryTheory.Sites.NonabelianCohomology.H1 import Mathlib.CategoryTheory.Sites.OneHypercover import Mathlib.CategoryTheory.Sites.Over @@ -1780,6 +1797,7 @@ import Mathlib.CategoryTheory.Sites.PreservesSheafification import Mathlib.CategoryTheory.Sites.Pretopology import Mathlib.CategoryTheory.Sites.Pullback import Mathlib.CategoryTheory.Sites.Sheaf +import Mathlib.CategoryTheory.Sites.SheafCohomology.Basic import Mathlib.CategoryTheory.Sites.SheafHom import Mathlib.CategoryTheory.Sites.SheafOfTypes import Mathlib.CategoryTheory.Sites.Sheafification @@ -1827,6 +1845,7 @@ import Mathlib.Combinatorics.Additive.Corner.Roth import Mathlib.Combinatorics.Additive.Dissociation import Mathlib.Combinatorics.Additive.ETransform import Mathlib.Combinatorics.Additive.Energy +import Mathlib.Combinatorics.Additive.ErdosGinzburgZiv import Mathlib.Combinatorics.Additive.FreimanHom import Mathlib.Combinatorics.Additive.PluenneckeRuzsa import Mathlib.Combinatorics.Additive.RuzsaCovering @@ -2188,6 +2207,7 @@ import Mathlib.Data.List.Pairwise import Mathlib.Data.List.Palindrome import Mathlib.Data.List.Perm import Mathlib.Data.List.Permutation +import Mathlib.Data.List.Pi import Mathlib.Data.List.Prime import Mathlib.Data.List.ProdSigma import Mathlib.Data.List.Range @@ -2277,6 +2297,7 @@ import Mathlib.Data.Nat.Choose.Cast import Mathlib.Data.Nat.Choose.Central import Mathlib.Data.Nat.Choose.Dvd import Mathlib.Data.Nat.Choose.Factorization +import Mathlib.Data.Nat.Choose.Lucas import Mathlib.Data.Nat.Choose.Multinomial import Mathlib.Data.Nat.Choose.Sum import Mathlib.Data.Nat.Choose.Vandermonde @@ -2317,7 +2338,6 @@ import Mathlib.Data.Nat.Periodic import Mathlib.Data.Nat.Prime.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.Nat.PrimeFin -import Mathlib.Data.Nat.PrimeNormNum import Mathlib.Data.Nat.Set import Mathlib.Data.Nat.Size import Mathlib.Data.Nat.Squarefree @@ -2654,6 +2674,7 @@ import Mathlib.GroupTheory.Congruence.Opposite import Mathlib.GroupTheory.Coprod.Basic import Mathlib.GroupTheory.CoprodI import Mathlib.GroupTheory.Coset +import Mathlib.GroupTheory.CosetCover import Mathlib.GroupTheory.Coxeter.Basic import Mathlib.GroupTheory.Coxeter.Inversion import Mathlib.GroupTheory.Coxeter.Length @@ -2682,16 +2703,14 @@ import Mathlib.GroupTheory.GroupAction.FixedPoints import Mathlib.GroupTheory.GroupAction.FixingSubgroup import Mathlib.GroupTheory.GroupAction.Group import Mathlib.GroupTheory.GroupAction.Hom +import Mathlib.GroupTheory.GroupAction.IterateAct import Mathlib.GroupTheory.GroupAction.Period -import Mathlib.GroupTheory.GroupAction.Pi import Mathlib.GroupTheory.GroupAction.Pointwise -import Mathlib.GroupTheory.GroupAction.Prod import Mathlib.GroupTheory.GroupAction.Quotient import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.GroupAction.SubMulAction.Pointwise import Mathlib.GroupTheory.GroupAction.Support -import Mathlib.GroupTheory.GroupAction.Units import Mathlib.GroupTheory.HNNExtension import Mathlib.GroupTheory.Index import Mathlib.GroupTheory.MonoidLocalization.Basic @@ -2753,7 +2772,6 @@ import Mathlib.Init.Data.Int.Order import Mathlib.Init.Data.List.Basic import Mathlib.Init.Data.List.Instances import Mathlib.Init.Data.List.Lemmas -import Mathlib.Init.Data.Nat.GCD import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Init.Data.Quot import Mathlib.Init.Data.Sigma.Basic @@ -3053,7 +3071,8 @@ import Mathlib.MeasureTheory.Constructions.Cylinders import Mathlib.MeasureTheory.Constructions.EventuallyMeasurable import Mathlib.MeasureTheory.Constructions.HaarToSphere import Mathlib.MeasureTheory.Constructions.Pi -import Mathlib.MeasureTheory.Constructions.Polish +import Mathlib.MeasureTheory.Constructions.Polish.Basic +import Mathlib.MeasureTheory.Constructions.Polish.EmbeddingReal import Mathlib.MeasureTheory.Constructions.Prod.Basic import Mathlib.MeasureTheory.Constructions.Prod.Integral import Mathlib.MeasureTheory.Constructions.Projective @@ -3225,7 +3244,6 @@ import Mathlib.MeasureTheory.OuterMeasure.Operations import Mathlib.MeasureTheory.PiSystem import Mathlib.MeasureTheory.SetAlgebra import Mathlib.MeasureTheory.SetSemiring -import Mathlib.MeasureTheory.Tactic import Mathlib.ModelTheory.Algebra.Field.Basic import Mathlib.ModelTheory.Algebra.Field.CharP import Mathlib.ModelTheory.Algebra.Ring.Basic @@ -3346,7 +3364,9 @@ import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody import Mathlib.NumberTheory.NumberField.ClassNumber import Mathlib.NumberTheory.NumberField.Discriminant import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.NumberTheory.NumberField.EquivReindex import Mathlib.NumberTheory.NumberField.FractionalIdeal +import Mathlib.NumberTheory.NumberField.House import Mathlib.NumberTheory.NumberField.Norm import Mathlib.NumberTheory.NumberField.Units.Basic import Mathlib.NumberTheory.NumberField.Units.DirichletTheorem @@ -3539,6 +3559,7 @@ import Mathlib.Order.PrimeIdeal import Mathlib.Order.PrimeSeparator import Mathlib.Order.PropInstances import Mathlib.Order.Radical +import Mathlib.Order.Rel.GaloisConnection import Mathlib.Order.RelClasses import Mathlib.Order.RelIso.Basic import Mathlib.Order.RelIso.Group @@ -3594,6 +3615,7 @@ import Mathlib.Probability.Kernel.Disintegration.CondCdf import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes +import Mathlib.Probability.Kernel.Disintegration.StandardBorel import Mathlib.Probability.Kernel.Disintegration.Unique import Mathlib.Probability.Kernel.IntegralCompProd import Mathlib.Probability.Kernel.Invariance @@ -3737,6 +3759,7 @@ import Mathlib.RingTheory.Ideal.Pointwise import Mathlib.RingTheory.Ideal.Prod import Mathlib.RingTheory.Ideal.Quotient import Mathlib.RingTheory.Ideal.QuotientOperations +import Mathlib.RingTheory.Idempotents import Mathlib.RingTheory.Int.Basic import Mathlib.RingTheory.IntegralClosure.Algebra.Basic import Mathlib.RingTheory.IntegralClosure.Algebra.Defs @@ -3862,6 +3885,7 @@ import Mathlib.RingTheory.ReesAlgebra import Mathlib.RingTheory.Regular.IsSMulRegular import Mathlib.RingTheory.Regular.RegularSequence import Mathlib.RingTheory.RingHom.Finite +import Mathlib.RingTheory.RingHom.FinitePresentation import Mathlib.RingTheory.RingHom.FiniteType import Mathlib.RingTheory.RingHom.Integral import Mathlib.RingTheory.RingHom.Surjective @@ -3875,6 +3899,7 @@ import Mathlib.RingTheory.SimpleModule import Mathlib.RingTheory.Smooth.Basic import Mathlib.RingTheory.Smooth.StandardSmooth import Mathlib.RingTheory.Support +import Mathlib.RingTheory.SurjectiveOnStalks import Mathlib.RingTheory.TensorProduct.Basic import Mathlib.RingTheory.TensorProduct.MvPolynomial import Mathlib.RingTheory.Trace.Basic @@ -4252,6 +4277,7 @@ import Mathlib.Topology.Algebra.Order.Rolle import Mathlib.Topology.Algebra.Order.UpperLower import Mathlib.Topology.Algebra.Polynomial import Mathlib.Topology.Algebra.PontryaginDual +import Mathlib.Topology.Algebra.ProperAction import Mathlib.Topology.Algebra.ProperConstSMul import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Algebra.Ring.Ideal @@ -4334,6 +4360,7 @@ import Mathlib.Topology.Compactness.PseudometrizableLindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.CompletelyRegular import Mathlib.Topology.Connected.Basic +import Mathlib.Topology.Connected.Clopen import Mathlib.Topology.Connected.LocallyConnected import Mathlib.Topology.Connected.PathConnected import Mathlib.Topology.Connected.Separation @@ -4346,17 +4373,14 @@ import Mathlib.Topology.ContinuousFunction.CocompactMap import Mathlib.Topology.ContinuousFunction.Compact import Mathlib.Topology.ContinuousFunction.CompactlySupported import Mathlib.Topology.ContinuousFunction.ContinuousMapZero -import Mathlib.Topology.ContinuousFunction.FunctionalCalculus import Mathlib.Topology.ContinuousFunction.Ideals import Mathlib.Topology.ContinuousFunction.LocallyConstant -import Mathlib.Topology.ContinuousFunction.NonUnitalFunctionalCalculus import Mathlib.Topology.ContinuousFunction.Ordered import Mathlib.Topology.ContinuousFunction.Polynomial import Mathlib.Topology.ContinuousFunction.Sigma import Mathlib.Topology.ContinuousFunction.StarOrdered import Mathlib.Topology.ContinuousFunction.StoneWeierstrass import Mathlib.Topology.ContinuousFunction.T0Sierpinski -import Mathlib.Topology.ContinuousFunction.UniqueCFC import Mathlib.Topology.ContinuousFunction.Units import Mathlib.Topology.ContinuousFunction.Weierstrass import Mathlib.Topology.ContinuousFunction.ZeroAtInfty @@ -4368,6 +4392,7 @@ import Mathlib.Topology.Defs.Filter import Mathlib.Topology.Defs.Induced import Mathlib.Topology.Defs.Sequences import Mathlib.Topology.DenseEmbedding +import Mathlib.Topology.DerivedSet import Mathlib.Topology.DiscreteQuotient import Mathlib.Topology.DiscreteSubset import Mathlib.Topology.EMetricSpace.Basic @@ -4544,7 +4569,6 @@ import Mathlib.Topology.Specialization import Mathlib.Topology.Spectral.Hom import Mathlib.Topology.StoneCech import Mathlib.Topology.Support -import Mathlib.Topology.Tactic import Mathlib.Topology.TietzeExtension import Mathlib.Topology.UniformSpace.AbsoluteValue import Mathlib.Topology.UniformSpace.AbstractCompletion diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 74e3e0c3b1502..7cba7ba8d66c3 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -330,10 +330,6 @@ theorem algebra_compatible_smul (r : R) (m : M) : r • m = (algebraMap R A) r theorem algebraMap_smul (r : R) (m : M) : (algebraMap R A) r • m = r • m := (algebra_compatible_smul A r m).symm -theorem intCast_smul {k V : Type*} [CommRing k] [AddCommGroup V] [Module k V] (r : ℤ) (x : V) : - (r : k) • x = r • x := - algebraMap_smul k r x - theorem NoZeroSMulDivisors.trans (R A M : Type*) [CommRing R] [Ring A] [IsDomain A] [Algebra R A] [AddCommGroup M] [Module R M] [Module A M] [IsScalarTower R A M] [NoZeroSMulDivisors R A] [NoZeroSMulDivisors A M] : NoZeroSMulDivisors R M := by @@ -476,12 +472,12 @@ end algebraMap section surjective variable {R S} [CommSemiring R] [Semiring S] [Algebra R S] -variable (h : Function.Surjective (algebraMap R S)) variable {M N} [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module S M] [IsScalarTower R S M] variable [Module R N] [Module S N] [IsScalarTower R S N] /-- If `R →+* S` is surjective, then `S`-linear maps between modules are exactly `R`-linear maps. -/ -def LinearMap.extendScalarsOfSurjectiveEquiv : (M →ₗ[R] N) ≃ₗ[R] (M →ₗ[S] N) where +def LinearMap.extendScalarsOfSurjectiveEquiv (h : Function.Surjective (algebraMap R S)) : + (M →ₗ[R] N) ≃ₗ[R] (M →ₗ[S] N) where toFun f := { __ := f, map_smul' := fun r x ↦ by obtain ⟨r, rfl⟩ := h r; simp } map_add' _ _ := rfl map_smul' _ _ := rfl @@ -490,18 +486,22 @@ def LinearMap.extendScalarsOfSurjectiveEquiv : (M →ₗ[R] N) ≃ₗ[R] (M → right_inv f := rfl /-- If `R →+* S` is surjective, then `R`-linear maps are also `S`-linear. -/ -abbrev LinearMap.extendScalarsOfSurjective (l : M →ₗ[R] N) : M →ₗ[S] N := +abbrev LinearMap.extendScalarsOfSurjective (h : Function.Surjective (algebraMap R S)) + (l : M →ₗ[R] N) : M →ₗ[S] N := extendScalarsOfSurjectiveEquiv h l -@[simp] -lemma LinearMap.extendScalarsOfSurjective_apply (l : M →ₗ[R] N) (x) : - l.extendScalarsOfSurjective h x = l x := rfl - /-- If `R →+* S` is surjective, then `R`-linear isomorphisms are also `S`-linear. -/ -def LinearEquiv.extendScalarsOfSurjective (f : M ≃ₗ[R] N) : M ≃ₗ[S] N where +def LinearEquiv.extendScalarsOfSurjective (h : Function.Surjective (algebraMap R S)) + (f : M ≃ₗ[R] N) : M ≃ₗ[S] N where __ := f map_smul' r x := by obtain ⟨r, rfl⟩ := h r; simp +variable (h : Function.Surjective (algebraMap R S)) + +@[simp] +lemma LinearMap.extendScalarsOfSurjective_apply (l : M →ₗ[R] N) (x) : + l.extendScalarsOfSurjective h x = l x := rfl + @[simp] lemma LinearEquiv.extendScalarsOfSurjective_apply (f : M ≃ₗ[R] N) (x) : f.extendScalarsOfSurjective h x = f x := rfl diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index d24d25b2125df..5df753173fa61 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Algebra.Algebra.Hom -import Mathlib.GroupTheory.GroupAction.Prod +import Mathlib.Algebra.GroupWithZero.Action.Prod /-! # Morphisms of non-unital algebras diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index 5f84e6a89b749..4359122d1476f 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -949,18 +949,15 @@ noncomputable def iSupLift [Nonempty ι] (K : ι → NonUnitalSubalgebra R A) (d map_mul' := by dsimp apply Set.iUnionLift_binary (coe_iSup_of_directed dir) dir _ (fun _ => (· * ·)) - on_goal 3 => rw [coe_iSup_of_directed dir] all_goals simp map_add' := by dsimp apply Set.iUnionLift_binary (coe_iSup_of_directed dir) dir _ (fun _ => (· + ·)) - on_goal 3 => rw [coe_iSup_of_directed dir] all_goals simp map_smul' := fun r => by dsimp apply Set.iUnionLift_unary (coe_iSup_of_directed dir) _ (fun _ x => r • x) (fun _ _ => rfl) - on_goal 2 => rw [coe_iSup_of_directed dir] all_goals simp } variable [Nonempty ι] {K : ι → NonUnitalSubalgebra R A} {dir : Directed (· ≤ ·) K} diff --git a/Mathlib/Algebra/Algebra/Prod.lean b/Mathlib/Algebra/Algebra/Prod.lean index bfc06f65a0c2f..312790886ed37 100644 --- a/Mathlib/Algebra/Algebra/Prod.lean +++ b/Mathlib/Algebra/Algebra/Prod.lean @@ -94,4 +94,10 @@ def prodEquiv : (A →ₐ[R] B) × (A →ₐ[R] C) ≃ (A →ₐ[R] B × C) wher left_inv f := by ext <;> rfl right_inv f := by ext <;> rfl +/-- `Prod.map` of two algebra homomorphisms. -/ +def prodMap {D : Type*} [Semiring D] [Algebra R D] (f : A →ₐ[R] B) (g : C →ₐ[R] D) : + A × C →ₐ[R] B × D := + { toRingHom := f.toRingHom.prodMap g.toRingHom + commutes' := fun r => by simp [commutes] } + end AlgHom diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 0c24ee13d0d32..23a330818b47f 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -548,21 +548,6 @@ theorem rangeRestrict_surjective (f : A →ₐ[R] B) : Function.Surjective (f.ra let ⟨x, hx⟩ := hy ⟨x, SetCoe.ext hx⟩ -/-- The equalizer of two R-algebra homomorphisms -/ -def equalizer (ϕ ψ : A →ₐ[R] B) : Subalgebra R A where - carrier := { a | ϕ a = ψ a } - zero_mem' := by simp only [Set.mem_setOf_eq, map_zero] - one_mem' := by simp only [Set.mem_setOf_eq, map_one] - add_mem' {x y} (hx : ϕ x = ψ x) (hy : ϕ y = ψ y) := by - rw [Set.mem_setOf_eq, map_add, map_add, hx, hy] - mul_mem' {x y} (hx : ϕ x = ψ x) (hy : ϕ y = ψ y) := by - rw [Set.mem_setOf_eq, map_mul, map_mul, hx, hy] - algebraMap_mem' x := by rw [Set.mem_setOf_eq, AlgHom.commutes, AlgHom.commutes] - -@[simp] -theorem mem_equalizer (ϕ ψ : A →ₐ[R] B) (x : A) : x ∈ ϕ.equalizer ψ ↔ ϕ x = ψ x := - Iff.rfl - /-- The range of a morphism of algebras is a fintype, if the domain is a fintype. Note that this instance can cause a diamond with `Subtype.fintype` if `B` is also a fintype. -/ @@ -1137,3 +1122,51 @@ theorem mem_subalgebraOfSubring {x : R} {S : Subring R} : x ∈ subalgebraOfSubr Iff.rfl end Int + +section Equalizer + +namespace AlgHom + +variable {R A B : Type*} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] +variable {F : Type*} [FunLike F A B] [AlgHomClass F R A B] + +/-- The equalizer of two R-algebra homomorphisms -/ +def equalizer (ϕ ψ : F) : Subalgebra R A where + carrier := { a | ϕ a = ψ a } + zero_mem' := by simp only [Set.mem_setOf_eq, map_zero] + one_mem' := by simp only [Set.mem_setOf_eq, map_one] + add_mem' {x y} (hx : ϕ x = ψ x) (hy : ϕ y = ψ y) := by + rw [Set.mem_setOf_eq, map_add, map_add, hx, hy] + mul_mem' {x y} (hx : ϕ x = ψ x) (hy : ϕ y = ψ y) := by + rw [Set.mem_setOf_eq, map_mul, map_mul, hx, hy] + algebraMap_mem' x := by + simp only [Set.mem_setOf_eq, AlgHomClass.commutes] + +@[simp] +theorem mem_equalizer (φ ψ : F) (x : A) : x ∈ equalizer φ ψ ↔ φ x = ψ x := + Iff.rfl + +theorem equalizer_toSubmodule {φ ψ : F} : + Subalgebra.toSubmodule (equalizer φ ψ) = LinearMap.eqLocus φ ψ := rfl + +@[simp] +theorem equalizer_eq_top {φ ψ : F} : equalizer φ ψ = ⊤ ↔ φ = ψ := by + simp [SetLike.ext_iff, DFunLike.ext_iff] + +@[simp] +theorem equalizer_same (φ : F) : equalizer φ φ = ⊤ := equalizer_eq_top.2 rfl + +theorem le_equalizer {φ ψ : F} {S : Subalgebra R A} : S ≤ equalizer φ ψ ↔ Set.EqOn φ ψ S := Iff.rfl + +theorem eqOn_sup {φ ψ : F} {S T : Subalgebra R A} (hS : Set.EqOn φ ψ S) (hT : Set.EqOn φ ψ T) : + Set.EqOn φ ψ ↑(S ⊔ T) := by + rw [← le_equalizer] at hS hT ⊢ + exact sup_le hS hT + +theorem ext_on_codisjoint {φ ψ : F} {S T : Subalgebra R A} (hST : Codisjoint S T) + (hS : Set.EqOn φ ψ S) (hT : Set.EqOn φ ψ T) : φ = ψ := + DFunLike.ext _ _ fun _ ↦ eqOn_sup hS hT <| hST.eq_top.symm ▸ trivial + +end AlgHom + +end Equalizer diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Directed.lean b/Mathlib/Algebra/Algebra/Subalgebra/Directed.lean index b32762eca0dda..1615f57fc9f98 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Directed.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Directed.lean @@ -24,9 +24,9 @@ open Algebra variable {R A B : Type*} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] variable (S : Subalgebra R A) -variable {ι : Type*} [Nonempty ι] {K : ι → Subalgebra R A} (dir : Directed (· ≤ ·) K) +variable {ι : Type*} [Nonempty ι] {K : ι → Subalgebra R A} -theorem coe_iSup_of_directed : ↑(iSup K) = ⋃ i, (K i : Set A) := +theorem coe_iSup_of_directed (dir : Directed (· ≤ ·) K) : ↑(iSup K) = ⋃ i, (K i : Set A) := let s : Subalgebra R A := { __ := Subsemiring.copy _ _ (Subsemiring.coe_iSup_of_directed dir).symm algebraMap_mem' := fun _ ↦ Set.mem_iUnion.2 @@ -36,15 +36,15 @@ theorem coe_iSup_of_directed : ↑(iSup K) = ⋃ i, (K i : Set A) := this.symm ▸ rfl variable (K) -variable (f : ∀ i, K i →ₐ[R] B) (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) - (T : Subalgebra R A) (hT : T = iSup K) -- Porting note (#11215): TODO: turn `hT` into an assumption `T ≤ iSup K`. -- That's what `Set.iUnionLift` needs -- Porting note: the proofs of `map_{zero,one,add,mul}` got a bit uglier, probably unification trbls /-- Define an algebra homomorphism on a directed supremum of subalgebras by defining it on each subalgebra, and proving that it agrees on the intersection of subalgebras. -/ -noncomputable def iSupLift : ↥T →ₐ[R] B := +noncomputable def iSupLift (dir : Directed (· ≤ ·) K) (f : ∀ i, K i →ₐ[R] B) + (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) + (T : Subalgebra R A) (hT : T = iSup K): ↥T →ₐ[R] B := { toFun := Set.iUnionLift (fun i => ↑(K i)) (fun i x => f i x) (fun i j x hxi hxj => by let ⟨k, hik, hjk⟩ := dir i j @@ -57,36 +57,41 @@ noncomputable def iSupLift : ↥T →ₐ[R] B := map_mul' := by subst hT; dsimp apply Set.iUnionLift_binary (coe_iSup_of_directed dir) dir _ (fun _ => (· * ·)) - on_goal 3 => rw [coe_iSup_of_directed dir] all_goals simp map_add' := by subst hT; dsimp apply Set.iUnionLift_binary (coe_iSup_of_directed dir) dir _ (fun _ => (· + ·)) - on_goal 3 => rw [coe_iSup_of_directed dir] all_goals simp commutes' := fun r => by dsimp apply Set.iUnionLift_const _ (fun _ => algebraMap R _ r) <;> simp } -variable {K dir f hf T hT} @[simp] -theorem iSupLift_inclusion {i : ι} (x : K i) (h : K i ≤ T) : +theorem iSupLift_inclusion {dir : Directed (· ≤ ·) K} {f : ∀ i, K i →ₐ[R] B} + {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} + {T : Subalgebra R A} {hT : T = iSup K} {i : ι} (x : K i) (h : K i ≤ T) : iSupLift K dir f hf T hT (inclusion h x) = f i x := by dsimp [iSupLift, inclusion] rw [Set.iUnionLift_inclusion] @[simp] -theorem iSupLift_comp_inclusion {i : ι} (h : K i ≤ T) : +theorem iSupLift_comp_inclusion {dir : Directed (· ≤ ·) K} {f : ∀ i, K i →ₐ[R] B} + {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} + {T : Subalgebra R A} {hT : T = iSup K} {i : ι} (h : K i ≤ T) : (iSupLift K dir f hf T hT).comp (inclusion h) = f i := by ext; simp @[simp] -theorem iSupLift_mk {i : ι} (x : K i) (hx : (x : A) ∈ T) : +theorem iSupLift_mk {dir : Directed (· ≤ ·) K} {f : ∀ i, K i →ₐ[R] B} + {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} + {T : Subalgebra R A} {hT : T = iSup K} {i : ι} (x : K i) (hx : (x : A) ∈ T) : iSupLift K dir f hf T hT ⟨x, hx⟩ = f i x := by dsimp [iSupLift, inclusion] rw [Set.iUnionLift_mk] -theorem iSupLift_of_mem {i : ι} (x : T) (hx : (x : A) ∈ K i) : +theorem iSupLift_of_mem {dir : Directed (· ≤ ·) K} {f : ∀ i, K i →ₐ[R] B} + {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} + {T : Subalgebra R A} {hT : T = iSup K} {i : ι} (x : T) (hx : (x : A) ∈ K i) : iSupLift K dir f hf T hT x = f i ⟨x, hx⟩ := by dsimp [iSupLift, inclusion] rw [Set.iUnionLift_of_mem] diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index e8b60e6443565..37fd87f4d697b 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -845,7 +845,7 @@ theorem mul_eq_one_iff {x y : Associates α} : x * y = 1 ↔ x = 1 ∧ y = 1 := (Quotient.inductionOn₂ x y fun a b h => have : a * b ~ᵤ 1 := Quotient.exact h ⟨Quotient.sound <| associated_one_of_associated_mul_one this, - Quotient.sound <| associated_one_of_associated_mul_one <| by rwa [mul_comm] at this⟩) + Quotient.sound <| associated_one_of_associated_mul_one (b := a) (by rwa [mul_comm])⟩) (by simp (config := { contextual := true })) theorem units_eq_one (u : (Associates α)ˣ) : u = 1 := diff --git a/Mathlib/Algebra/BigOperators/Associated.lean b/Mathlib/Algebra/BigOperators/Associated.lean index 1ad0220d73544..c0dc9dbb0050f 100644 --- a/Mathlib/Algebra/BigOperators/Associated.lean +++ b/Mathlib/Algebra/BigOperators/Associated.lean @@ -22,9 +22,9 @@ local infixl:50 " ~ᵤ " => Associated namespace Prime -variable [CommMonoidWithZero α] {p : α} (hp : Prime p) +variable [CommMonoidWithZero α] {p : α} -theorem exists_mem_multiset_dvd {s : Multiset α} : p ∣ s.prod → ∃ a ∈ s, p ∣ a := +theorem exists_mem_multiset_dvd (hp : Prime p) {s : Multiset α} : p ∣ s.prod → ∃ a ∈ s, p ∣ a := Multiset.induction_on s (fun h => (hp.not_dvd_one h).elim) fun a s ih h => have : p ∣ a * s.prod := by simpa using h match hp.dvd_or_dvd this with @@ -33,12 +33,13 @@ theorem exists_mem_multiset_dvd {s : Multiset α} : p ∣ s.prod → ∃ a ∈ s let ⟨a, has, h⟩ := ih h ⟨a, Multiset.mem_cons_of_mem has, h⟩ -theorem exists_mem_multiset_map_dvd {s : Multiset β} {f : β → α} : +theorem exists_mem_multiset_map_dvd (hp : Prime p) {s : Multiset β} {f : β → α} : p ∣ (s.map f).prod → ∃ a ∈ s, p ∣ f a := fun h => by simpa only [exists_prop, Multiset.mem_map, exists_exists_and_eq_and] using hp.exists_mem_multiset_dvd h -theorem exists_mem_finset_dvd {s : Finset β} {f : β → α} : p ∣ s.prod f → ∃ i ∈ s, p ∣ f i := +theorem exists_mem_finset_dvd (hp : Prime p) {s : Finset β} {f : β → α} : + p ∣ s.prod f → ∃ i ∈ s, p ∣ f i := hp.exists_mem_multiset_map_dvd end Prime diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index a8d0c959a9e3c..e8acaf6380069 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -406,7 +406,7 @@ theorem prod_take_ofFn {n : ℕ} (f : Fin n → α) (i : ℕ) : · have A : (ofFn f).take i = (ofFn f).take i.succ := by rw [← length_ofFn f] at h have : length (ofFn f) ≤ i := not_lt.mp h - rw [take_all_of_le this, take_all_of_le (le_trans this (Nat.le_succ _))] + rw [take_of_length_le this, take_of_length_le (le_trans this (Nat.le_succ _))] have B : ∀ j : Fin n, ((j : ℕ) < i.succ) = ((j : ℕ) < i) := by intro j have : (j : ℕ) < i := lt_of_lt_of_le j.2 (not_lt.mp h) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 10e1ea0d701af..e5a2b3e91b3f4 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -46,11 +46,15 @@ open Fin Function namespace Finset -/-- `∏ x ∈ s, f x` is the product of `f x` -as `x` ranges over the elements of the finite set `s`. --/ +/-- `∏ x ∈ s, f x` is the product of `f x` as `x` ranges over the elements of the finite set `s`. + +When the index type is a `Fintype`, the notation `∏ x, f x`, is a shorthand for +`∏ x ∈ Finset.univ, f x`. -/ @[to_additive "`∑ x ∈ s, f x` is the sum of `f x` as `x` ranges over the elements -of the finite set `s`."] +of the finite set `s`. + +When the index type is a `Fintype`, the notation `∑ x, f x`, is a shorthand for +`∑ x ∈ Finset.univ, f x`."] protected def prod [CommMonoid β] (s : Finset α) (f : α → β) : β := (s.1.map f).prod @@ -2089,18 +2093,23 @@ theorem disjoint_finset_sum_right {β : Type*} {i : Finset β} {f : β → Multi {a : Multiset α} : Multiset.Disjoint a (i.sum f) ↔ ∀ b ∈ i, Multiset.Disjoint a (f b) := by simpa only [disjoint_comm] using disjoint_finset_sum_left +@[simp] +lemma mem_sum {s : Finset ι} {m : ι → Multiset α} : a ∈ ∑ i ∈ s, m i ↔ ∃ i ∈ s, a ∈ m i := by + induction' s using Finset.cons_induction <;> simp [*] + variable [DecidableEq α] -@[simp] theorem toFinset_sum_count_eq (s : Multiset α) : ∑ a in s.toFinset, s.count a = card s := by simpa using (Finset.sum_multiset_map_count s (fun _ => (1 : ℕ))).symm -@[simp] -theorem sum_count_eq [Fintype α] (s : Multiset α) : ∑ a, s.count a = Multiset.card s := by +@[simp] lemma sum_count_eq_card {s : Finset α} {m : Multiset α} (hms : ∀ a ∈ m, a ∈ s) : + ∑ a ∈ s, m.count a = card m := by rw [← toFinset_sum_count_eq, ← Finset.sum_filter_ne_zero] - congr - ext - simp + congr with a + simpa using hms a + +@[deprecated sum_count_eq_card (since := "2024-07-21")] +theorem sum_count_eq [Fintype α] (s : Multiset α) : ∑ a, s.count a = Multiset.card s := by simp theorem count_sum' {s : Finset β} {a : α} {f : β → Multiset α} : count a (∑ x ∈ s, f x) = ∑ x ∈ s, count a (f x) := by diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 51a8921c84de9..63d796abc7c03 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -297,9 +297,10 @@ lemma prod_map_eq_pow_single [DecidableEq α] {l : List α} (a : α) (f : α → · rw [map_nil, prod_nil, count_nil, _root_.pow_zero] · specialize h a fun a' ha' hfa' => hf a' ha' (mem_cons_of_mem _ hfa') rw [List.map_cons, List.prod_cons, count_cons, h] + simp only [beq_iff_eq] split_ifs with ha' · rw [ha', _root_.pow_succ'] - · rw [hf a' (Ne.symm ha') (List.mem_cons_self a' as), one_mul, add_zero] + · rw [hf a' ha' (List.mem_cons_self a' as), one_mul, add_zero] @[to_additive] lemma prod_eq_pow_single [DecidableEq M] (a : M) (h : ∀ a', a' ≠ a → a' ∈ l → a' = 1) : @@ -463,7 +464,7 @@ theorem prod_set' (L : List G) (n : ℕ) (a : G) : split_ifs with hn · rw [mul_comm _ a, mul_assoc a, prod_drop_succ L n hn, mul_comm _ (drop n L).prod, ← mul_assoc (take n L).prod, prod_take_mul_prod_drop, mul_comm a, mul_assoc] - · simp only [take_all_of_le (le_of_not_lt hn), prod_nil, mul_one, + · simp only [take_of_length_le (le_of_not_lt hn), prod_nil, mul_one, drop_eq_nil_of_le ((le_of_not_lt hn).trans n.le_succ)] @[to_additive] @@ -477,9 +478,9 @@ lemma prod_map_ite_eq {A : Type*} [DecidableEq A] (l : List A) (f g : A → G) ( rw [ih] clear ih by_cases hx : x = a - · simp only [hx, ite_true, div_pow, pow_add, pow_one, div_eq_mul_inv, mul_assoc, mul_comm, - mul_left_comm, mul_inv_cancel_left] - · simp only [hx, ite_false, ne_comm.mp hx, add_zero, mul_assoc, mul_comm (g x) _] + · simp only [hx, ↓reduceIte, div_eq_mul_inv, beq_self_eq_true, pow_add, pow_one, mul_comm, + mul_assoc, mul_left_comm, mul_inv_cancel_left] + · simp only [hx, ↓reduceIte, div_pow, mul_comm (g x) _, mul_assoc, beq_iff_eq, add_zero] end CommGroup @@ -578,13 +579,14 @@ theorem sum_map_count_dedup_filter_eq_countP (p : α → Bool) (l : List α) : match p a with | true => simp only [List.map_cons, List.sum_cons, List.count_eq_zero.2 ha, zero_add] | false => simp only - · by_cases hp : p a - · refine _root_.trans (sum_map_eq_nsmul_single a _ fun _ h _ => by simp [h]) ?_ + · simp only [beq_iff_eq] + by_cases hp : p a + · refine _root_.trans (sum_map_eq_nsmul_single a _ fun _ h _ => by simp [h.symm]) ?_ simp [hp, count_dedup] · refine _root_.trans (List.sum_eq_zero fun n hn => ?_) (by simp [hp]) obtain ⟨a', ha'⟩ := List.mem_map.1 hn split_ifs at ha' with ha - · simp only [ha, mem_filter, mem_dedup, find?, mem_cons, true_or, hp, + · simp only [ha.symm, mem_filter, mem_dedup, find?, mem_cons, true_or, hp, and_false, false_and] at ha' · exact ha'.2.symm @@ -630,10 +632,6 @@ lemma ranges_join (l : List ℕ) : l.ranges.join = range l.sum := by simp [range lemma mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by simp [mem_mem_ranges_iff_lt_natSum] -@[simp] -theorem length_join (L : List (List α)) : length (join L) = sum (map length L) := by - induction L <;> [rfl; simp only [*, join, map, sum_cons, length_append]] - lemma countP_join (p : α → Bool) : ∀ L : List (List α), countP p L.join = (L.map (countP p)).sum | [] => rfl | a :: l => by rw [join, countP_append, map_cons, sum_cons, countP_join _ l] @@ -643,7 +641,8 @@ lemma count_join [BEq α] (L : List (List α)) (a : α) : L.join.count a = (L.ma @[simp] theorem length_bind (l : List α) (f : α → List β) : - length (List.bind l f) = sum (map (length ∘ f) l) := by rw [List.bind, length_join, map_map] + length (List.bind l f) = sum (map (length ∘ f) l) := by + rw [List.bind, length_join, map_map, Nat.sum_eq_listSum] lemma countP_bind (p : β → Bool) (l : List α) (f : α → List β) : countP p (l.bind f) = sum (map (countP p ∘ f) l) := by rw [List.bind, countP_join, map_map] diff --git a/Mathlib/Algebra/BigOperators/Ring/Nat.lean b/Mathlib/Algebra/BigOperators/Ring/Nat.lean new file mode 100644 index 0000000000000..1b3759dd9bbe0 --- /dev/null +++ b/Mathlib/Algebra/BigOperators/Ring/Nat.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2024 Pim Otte. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Pim Otte +-/ +import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.Ring.Parity + +/-! +# Big operators on a finset in the natural numbers + +This file contains the results concerning the interaction of finset big operators with natural +numbers. +-/ + +variable {ι : Type*} + +namespace Finset + +lemma even_sum_iff_even_card_odd {s : Finset ι} (f : ι → ℕ) : + Even (∑ i ∈ s, f i) ↔ Even (s.filter fun x ↦ Odd (f x)).card := by + rw [← Finset.sum_filter_add_sum_filter_not _ (fun x ↦ Even (f x)), Nat.even_add] + simp only [Finset.mem_filter, and_imp, imp_self, implies_true, Finset.even_sum, true_iff] + rw [Nat.even_iff, Finset.sum_nat_mod, Finset.sum_filter] + simp (config := { contextual := true }) only [← Nat.odd_iff_not_even, Nat.odd_iff.mp] + simp_rw [← Finset.sum_filter, ← Nat.even_iff, Finset.card_eq_sum_ones] + +lemma odd_sum_iff_odd_card_odd {s : Finset ι} (f : ι → ℕ) : + Odd (∑ i ∈ s, f i) ↔ Odd (s.filter fun x ↦ Odd (f x)).card := by + simp only [Nat.odd_iff_not_even, even_sum_iff_even_card_odd] + +end Finset diff --git a/Mathlib/Algebra/Category/Grp/Adjunctions.lean b/Mathlib/Algebra/Category/Grp/Adjunctions.lean index aa8c406045283..92b143f959c69 100644 --- a/Mathlib/Algebra/Category/Grp/Adjunctions.lean +++ b/Mathlib/Algebra/Category/Grp/Adjunctions.lean @@ -3,8 +3,9 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Johannes Hölzl -/ -import Mathlib.Algebra.Category.Grp.Basic +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.GroupTheory.FreeAbelianGroup +import Mathlib.CategoryTheory.Limits.Shapes.Types /-! # Adjunctions regarding the category of (abelian) groups @@ -35,7 +36,7 @@ noncomputable section universe u -open CategoryTheory +open CategoryTheory Limits namespace AddCommGrp @@ -73,6 +74,9 @@ def adj : free ⊣ forget AddCommGrp.{u} := simp only [Equiv.symm_symm] apply FreeAbelianGroup.lift_comp } +instance : free.{u}.IsLeftAdjoint := + ⟨_, ⟨adj⟩⟩ + instance : (forget AddCommGrp.{u}).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ @@ -84,6 +88,18 @@ the monomorphisms in `AddCommGroup` are just the injective functions. example {G H : AddCommGrp.{u}} (f : G ⟶ H) [Mono f] : Function.Injective f := (mono_iff_injective (f : G → H)).mp (Functor.map_mono (forget AddCommGrp) f) +instance : (free.{u}).PreservesMonomorphisms where + preserves {X Y} f _ := by + by_cases hX : IsEmpty X + · constructor + intros + apply (IsInitial.isInitialObj free _ + ((Types.initial_iff_empty X).2 hX).some).isZero.eq_of_tgt + · simp only [not_isEmpty_iff] at hX + have hf : Function.Injective f := by rwa [← mono_iff_injective] + obtain ⟨g, hg⟩ := hf.hasLeftInverse + have : IsSplitMono f := IsSplitMono.mk' { retraction := g } + infer_instance end AddCommGrp diff --git a/Mathlib/Algebra/Category/Grp/EpiMono.lean b/Mathlib/Algebra/Category/Grp/EpiMono.lean index 325ef9f2d901f..4ead5efd13225 100644 --- a/Mathlib/Algebra/Category/Grp/EpiMono.lean +++ b/Mathlib/Algebra/Category/Grp/EpiMono.lean @@ -33,7 +33,7 @@ variable [Group A] [Group B] @[to_additive] theorem ker_eq_bot_of_cancel {f : A →* B} (h : ∀ u v : f.ker →* A, f.comp u = f.comp v → u = v) : - f.ker = ⊥ := by simpa using _root_.congr_arg range (h f.ker.subtype 1 (by aesop_cat)) + f.ker = ⊥ := by simpa using congr_arg range (h f.ker.subtype 1 (by aesop_cat)) end diff --git a/Mathlib/Algebra/Category/Grp/Kernels.lean b/Mathlib/Algebra/Category/Grp/Kernels.lean index 818c11a8823e6..8f367b2c6f460 100644 --- a/Mathlib/Algebra/Category/Grp/Kernels.lean +++ b/Mathlib/Algebra/Category/Grp/Kernels.lean @@ -42,6 +42,6 @@ def cokernelIsColimit : IsColimit <| cokernelCocone f := (fun s => lift _ _ <| (range_le_ker_iff _ _).mpr <| CokernelCofork.condition s) (fun _ => rfl) (fun _ _ h => have : Epi (cokernelCocone f).π := (epi_iff_surjective _).mpr <| mk'_surjective _ - (cancel_epi _).mp <| by simpa only [parallelPair_obj_one] using h) + (cancel_epi (cokernelCocone f).π).mp <| by simpa only [parallelPair_obj_one] using h) end AddCommGrp diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index 829c9afc4cff2..1f731dd378510 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -111,7 +111,7 @@ section Terminal def punitIsTerminal : IsTerminal (CommRingCat.of.{u} PUnit) := by refine IsTerminal.ofUnique (h := fun X => ⟨⟨⟨⟨1, rfl⟩, fun _ _ => rfl⟩, ?_, ?_⟩, ?_⟩) · rfl - · intros; simp; ext + · intros; simp only [coe_of, Pi.one_apply, self_eq_add_right]; ext · intros f; ext; rfl instance commRingCat_hasStrictTerminalObjects : HasStrictTerminalObjects CommRingCat.{u} := by diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index 4af5f4ab4dc76..ce2ecc6aef10e 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -110,6 +110,11 @@ theorem add_pow_char_pow [CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] { (x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n := add_pow_char_pow_of_commute _ _ _ (Commute.all _ _) +theorem add_pow_eq_add_pow_mod_mul_pow_add_pow_div + [CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] {n : ℕ} (x y : R) : + (x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := by + rw [← add_pow_char, ← pow_mul, ← pow_add, Nat.mod_add_div] + theorem sub_pow_char [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] (x y : R) : (x - y) ^ p = x ^ p - y ^ p := sub_pow_char_of_commute _ _ _ (Commute.all _ _) @@ -118,6 +123,11 @@ theorem sub_pow_char_pow [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] {n : (x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n := sub_pow_char_pow_of_commute _ _ _ (Commute.all _ _) +theorem sub_pow_eq_sub_pow_mod_mul_pow_sub_pow_div + [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] {n : ℕ} (x y : R) : + (x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := by + rw [← sub_pow_char, ← pow_mul, ← pow_add, Nat.mod_add_div] + theorem CharP.neg_one_pow_char [Ring R] (p : ℕ) [CharP R p] [Fact p.Prime] : (-1 : R) ^ p = -1 := by rw [eq_neg_iff_add_eq_zero] diff --git a/Mathlib/Algebra/CharZero/Lemmas.lean b/Mathlib/Algebra/CharZero/Lemmas.lean index dd6c53ec469bc..0c637d4a69ff7 100644 --- a/Mathlib/Algebra/CharZero/Lemmas.lean +++ b/Mathlib/Algebra/CharZero/Lemmas.lean @@ -149,7 +149,7 @@ section RingHom variable {R S : Type*} [NonAssocSemiring R] [NonAssocSemiring S] theorem RingHom.charZero (ϕ : R →+* S) [CharZero S] : CharZero R := - ⟨fun a b h => CharZero.cast_injective (by rw [← map_natCast ϕ, ← map_natCast ϕ, h])⟩ + ⟨fun a b h => CharZero.cast_injective (R := S) (by rw [← map_natCast ϕ, ← map_natCast ϕ, h])⟩ theorem RingHom.charZero_iff {ϕ : R →+* S} (hϕ : Function.Injective ϕ) : CharZero R ↔ CharZero S := ⟨fun hR => diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean b/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean index 8f97c5671ad18..af131395506b3 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/TerminatesIffRat.lean @@ -49,7 +49,8 @@ We want to show that the computation of a continued fraction `GenContFract.of v` terminates if and only if `v ∈ ℚ`. In this section, we show the implication from left to right. We first show that every finite convergent corresponds to a rational number `q` and then use the -finite correctness proof (`of_correctness_of_terminates`) of `GenContFract.of` to show that `v = ↑q`. +finite correctness proof (`of_correctness_of_terminates`) of `GenContFract.of` to show that +`v = ↑q`. -/ diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean index 052a03011554d..c14beeefe520e 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean @@ -270,8 +270,8 @@ theorem of_s_of_int (a : ℤ) : (of (a : K)).s = Stream'.Seq.nil := variable {K} (v) -/-- Recurrence for the `GenContFract.of` an element `v` of `K` in terms of that of the inverse of the -fractional part of `v`. +/-- Recurrence for the `GenContFract.of` an element `v` of `K` in terms of that of the inverse of +the fractional part of `v`. -/ theorem of_s_succ (n : ℕ) : (of v).s.get? (n + 1) = (of (fract v)⁻¹).s.get? n := by rcases eq_or_ne (fract v) 0 with h | h diff --git a/Mathlib/Algebra/DirectSum/Basic.lean b/Mathlib/Algebra/DirectSum/Basic.lean index 9ccaa4b053ac5..cb22d873d8661 100644 --- a/Mathlib/Algebra/DirectSum/Basic.lean +++ b/Mathlib/Algebra/DirectSum/Basic.lean @@ -111,6 +111,8 @@ def mk (s : Finset ι) : (∀ i : (↑s : Set ι), β i.1) →+ ⨁ i, β i wher def of (i : ι) : β i →+ ⨁ i, β i := DFinsupp.singleAddHom β i +variable {β} + @[simp] theorem of_eq_same (i : ι) (x : β i) : (of _ i x) i = x := DFinsupp.single_eq_same @@ -121,6 +123,16 @@ theorem of_eq_of_ne (i j : ι) (x : β i) (h : i ≠ j) : (of _ i x) j = 0 := lemma of_apply {i : ι} (j : ι) (x : β i) : of β i x j = if h : i = j then Eq.recOn h x else 0 := DFinsupp.single_apply +theorem mk_apply_of_mem {s : Finset ι} {f : ∀ i : (↑s : Set ι), β i.val} {n : ι} (hn : n ∈ s) : + mk β s f n = f ⟨n, hn⟩ := by + dsimp only [Finset.coe_sort_coe, mk, AddMonoidHom.coe_mk, ZeroHom.coe_mk, DFinsupp.mk_apply] + rw [dif_pos hn] + +theorem mk_apply_of_not_mem {s : Finset ι} {f : ∀ i : (↑s : Set ι), β i.val} {n : ι} (hn : n ∉ s) : + mk β s f n = 0 := by + dsimp only [Finset.coe_sort_coe, mk, AddMonoidHom.coe_mk, ZeroHom.coe_mk, DFinsupp.mk_apply] + rw [dif_neg hn] + @[simp] theorem support_zero [∀ (i : ι) (x : β i), Decidable (x ≠ 0)] : (0 : ⨁ i, β i).support = ∅ := DFinsupp.support_zero @@ -144,8 +156,6 @@ theorem sum_univ_of [Fintype ι] (x : ⨁ i, β i) : rw [DFinsupp.finset_sum_apply] simp [of_apply] -variable {β} - theorem mk_injective (s : Finset ι) : Function.Injective (mk β s) := DFinsupp.mk_injective s @@ -326,6 +336,13 @@ protected def coeAddMonoidHom {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [ [AddSubmonoidClass S M] (A : ι → S) : (⨁ i, A i) →+ M := toAddMonoid fun i => AddSubmonoidClass.subtype (A i) +theorem coeAddMonoidHom_eq_dfinsupp_sum {M S : Type*} [DecidableEq M] [AddCommMonoid M] + [SetLike S M] [AddSubmonoidClass S M] (A : ι → S) (x : DirectSum ι fun i => A i) : + DirectSum.coeAddMonoidHom A x = DFinsupp.sum x fun i => (fun x : A i => ↑x) := by + simp only [DirectSum.coeAddMonoidHom, toAddMonoid, DFinsupp.liftAddHom, AddEquiv.coe_mk, + Equiv.coe_fn_mk] + exact DFinsupp.sumAddHom_apply _ x + @[simp] theorem coeAddMonoidHom_of {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S M] [AddSubmonoidClass S M] (A : ι → S) (i : ι) (x : A i) : @@ -337,7 +354,7 @@ theorem coe_of_apply {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S (of (fun i ↦ {x // x ∈ A i}) i x j : M) = if i = j then x else 0 := by obtain rfl | h := Decidable.eq_or_ne i j · rw [DirectSum.of_eq_same, if_pos rfl] - · rw [DirectSum.of_eq_of_ne _ _ _ _ h, if_neg h, ZeroMemClass.coe_zero, ZeroMemClass.coe_zero] + · rw [DirectSum.of_eq_of_ne _ _ _ h, if_neg h, ZeroMemClass.coe_zero, ZeroMemClass.coe_zero] /-- The `DirectSum` formed by a collection of additive submonoids (or subgroups, or submodules) of `M` is said to be internal if the canonical map `(⨁ i, A i) →+ M` is bijective. @@ -354,4 +371,17 @@ theorem IsInternal.addSubmonoid_iSup_eq_top {M : Type*} [DecidableEq ι] [AddCom rw [AddSubmonoid.iSup_eq_mrange_dfinsupp_sumAddHom, AddMonoidHom.mrange_top_iff_surjective] exact Function.Bijective.surjective h +variable {M S : Type*} [DecidableEq ι] [DecidableEq M] [AddCommMonoid M] [SetLike S M] + [AddSubmonoidClass S M] (A : ι → S) + +theorem support_subset (x : DirectSum ι fun i => A i) : + (Function.support fun i => (x i : M)) ⊆ ↑(DFinsupp.support x) := by + intro m + simp only [Function.mem_support, Finset.mem_coe, DFinsupp.mem_support_toFun, not_imp_not, + ZeroMemClass.coe_eq_zero, imp_self] + +theorem finite_support (x : DirectSum ι fun i => A i) : + (Function.support fun i => (x i : M)).Finite := + Set.Finite.subset (DFinsupp.support x : Set ι).toFinite (DirectSum.support_subset _ x) + end DirectSum diff --git a/Mathlib/Algebra/DirectSum/Decomposition.lean b/Mathlib/Algebra/DirectSum/Decomposition.lean index 55aee463f8f67..45f344b63d1ae 100644 --- a/Mathlib/Algebra/DirectSum/Decomposition.lean +++ b/Mathlib/Algebra/DirectSum/Decomposition.lean @@ -128,7 +128,7 @@ theorem decompose_of_mem_same {x : M} {i : ι} (hx : x ∈ ℳ i) : (decompose theorem decompose_of_mem_ne {x : M} {i j : ι} (hx : x ∈ ℳ i) (hij : i ≠ j) : (decompose ℳ x j : M) = 0 := by - rw [decompose_of_mem _ hx, DirectSum.of_eq_of_ne _ _ _ _ hij, ZeroMemClass.coe_zero] + rw [decompose_of_mem _ hx, DirectSum.of_eq_of_ne _ _ _ hij, ZeroMemClass.coe_zero] theorem degree_eq_of_mem_mem {x : M} {i j : ι} (hxi : x ∈ ℳ i) (hxj : x ∈ ℳ j) (hx : x ≠ 0) : i = j := by @@ -178,7 +178,7 @@ theorem decompose_symm_sum {ι'} (s : Finset ι') (f : ι' → ⨁ i, ℳ i) : theorem sum_support_decompose [∀ (i) (x : ℳ i), Decidable (x ≠ 0)] (r : M) : (∑ i ∈ (decompose ℳ r).support, (decompose ℳ r i : M)) = r := by conv_rhs => - rw [← (decompose ℳ).symm_apply_apply r, ← sum_support_of (fun i ↦ ℳ i) (decompose ℳ r)] + rw [← (decompose ℳ).symm_apply_apply r, ← sum_support_of (decompose ℳ r)] rw [decompose_symm_sum] simp_rw [decompose_symm_of] diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 669f9931abb13..8eb72e07f3608 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -173,7 +173,7 @@ theorem coe_mul_apply_eq_dfinsupp_sum [AddMonoid ι] [SetLike.GradedMonoid A] · subst h rw [of_eq_same] rfl - · rw [of_eq_of_ne _ _ _ _ h] + · rw [of_eq_of_ne _ _ _ h] rfl theorem coe_of_mul_apply_aux [AddMonoid ι] [SetLike.GradedMonoid A] {i : ι} (r : A i) diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index b8cc3fcbd8159..828f0ad54d9e2 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -269,6 +269,13 @@ indexed by `ι`. This is `DirectSum.coeAddMonoidHom` as a `LinearMap`. -/ def coeLinearMap : (⨁ i, A i) →ₗ[R] M := toModule R ι M fun i ↦ (A i).subtype +theorem coeLinearMap_eq_dfinsupp_sum [DecidableEq M] (x : DirectSum ι fun i => A i) : + coeLinearMap A x = DFinsupp.sum x fun i => (fun x : A i => ↑x) := by + simp only [coeLinearMap, toModule, DFinsupp.lsum, LinearEquiv.coe_mk, LinearMap.coe_mk, + AddHom.coe_mk] + rw [DFinsupp.sumAddHom_apply] + simp only [LinearMap.toAddMonoidHom_coe, Submodule.coeSubtype] + @[simp] theorem coeLinearMap_of (i : ι) (x : A i) : DirectSum.coeLinearMap A (of (fun i ↦ A i) i x) = x := -- Porting note: spelled out arguments. (I don't know how this works.) @@ -286,7 +293,7 @@ theorem IsInternal.ofBijective_coeLinearMap_same (h : IsInternal A) theorem IsInternal.ofBijective_coeLinearMap_of_ne (h : IsInternal A) {i j : ι} (hij : i ≠ j) (x : A i) : (LinearEquiv.ofBijective (coeLinearMap A) h).symm x j = 0 := by - rw [← coeLinearMap_of, LinearEquiv.ofBijective_symm_apply_apply, of_eq_of_ne _ i j _ hij] + rw [← coeLinearMap_of, LinearEquiv.ofBijective_symm_apply_apply, of_eq_of_ne i j _ hij] theorem IsInternal.ofBijective_coeLinearMap_of_mem (h : IsInternal A) {i : ι} {x : M} (hx : x ∈ A i) : diff --git a/Mathlib/Algebra/Exact.lean b/Mathlib/Algebra/Exact.lean index f18e91cf12c76..87008b6a28f9b 100644 --- a/Mathlib/Algebra/Exact.lean +++ b/Mathlib/Algebra/Exact.lean @@ -109,10 +109,9 @@ variable {X₁ X₂ X₃ Y₁ Y₂ Y₃ : Type*} [AddCommMonoid X₁] [AddCommMo [AddCommMonoid Y₁] [AddCommMonoid Y₂] [AddCommMonoid Y₃] (e₁ : X₁ ≃+ Y₁) (e₂ : X₂ ≃+ Y₂) (e₃ : X₃ ≃+ Y₃) {f₁₂ : X₁ →+ X₂} {f₂₃ : X₂ →+ X₃} {g₁₂ : Y₁ →+ Y₂} {g₂₃ : Y₂ →+ Y₃} - (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) -lemma of_ladder_addEquiv_of_exact (H : Exact f₁₂ f₂₃) : Exact g₁₂ g₂₃ := by +lemma of_ladder_addEquiv_of_exact (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact f₁₂ f₂₃) : Exact g₁₂ g₂₃ := by have h₁₂ := DFunLike.congr_fun comm₁₂ have h₂₃ := DFunLike.congr_fun comm₂₃ dsimp at h₁₂ h₂₃ @@ -126,7 +125,8 @@ lemma of_ladder_addEquiv_of_exact (H : Exact f₁₂ f₂₃) : Exact g₁₂ g obtain ⟨x₁, rfl⟩ := (H x₂).1 (e₃.injective (by rw [← h₂₃, hx₂, map_zero])) exact ⟨e₁ x₁, by rw [h₁₂]⟩ -lemma of_ladder_addEquiv_of_exact' (H : Exact g₁₂ g₂₃) : Exact f₁₂ f₂₃ := by +lemma of_ladder_addEquiv_of_exact' (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact g₁₂ g₂₃) : Exact f₁₂ f₂₃ := by refine of_ladder_addEquiv_of_exact e₁.symm e₂.symm e₃.symm ?_ ?_ H · ext y₁ obtain ⟨x₁, rfl⟩ := e₁.surjective y₁ @@ -137,7 +137,8 @@ lemma of_ladder_addEquiv_of_exact' (H : Exact g₁₂ g₂₃) : Exact f₁₂ f apply e₃.injective simpa using DFunLike.congr_fun comm₂₃.symm x₂ -lemma iff_of_ladder_addEquiv : Exact g₁₂ g₂₃ ↔ Exact f₁₂ f₂₃ := by +lemma iff_of_ladder_addEquiv (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) : Exact g₁₂ g₂₃ ↔ Exact f₁₂ f₂₃ := by constructor · exact of_ladder_addEquiv_of_exact' e₁ e₂ e₃ comm₁₂ comm₂₃ · exact of_ladder_addEquiv_of_exact e₁ e₂ e₃ comm₁₂ comm₂₃ diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index 1d0578a2e4614..30ccec87b7083 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -228,7 +228,9 @@ noncomputable abbrev DivisionRing.ofIsUnitOrEqZero [Ring R] (h : ∀ a : R, IsUn toRing := ‹Ring R› __ := groupWithZeroOfIsUnitOrEqZero h nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q a => rfl /-- Constructs a `Field` structure on a `CommRing` consisting only of units and 0. -/ -- See note [reducible non-instances] diff --git a/Mathlib/Algebra/Field/Opposite.lean b/Mathlib/Algebra/Field/Opposite.lean index de970262cb925..cde950bef7a7c 100644 --- a/Mathlib/Algebra/Field/Opposite.lean +++ b/Mathlib/Algebra/Field/Opposite.lean @@ -34,6 +34,7 @@ instance instDivisionSemiring [DivisionSemiring α] : DivisionSemiring αᵐᵒ __ := instSemiring __ := instGroupWithZero nnqsmul := _ + nnqsmul_def := fun q a => rfl nnratCast_def q := unop_injective $ by rw [unop_nnratCast, unop_div, unop_natCast, unop_natCast, NNRat.cast_def, div_eq_mul_inv, Nat.cast_comm] @@ -41,6 +42,7 @@ instance instDivisionRing [DivisionRing α] : DivisionRing αᵐᵒᵖ where __ := instRing __ := instDivisionSemiring qsmul := _ + qsmul_def := fun q a => rfl ratCast_def q := unop_injective <| by rw [unop_ratCast, Rat.cast_def, unop_div, unop_natCast, unop_intCast, Int.commute_cast, div_eq_mul_inv] @@ -60,6 +62,7 @@ instance instDivisionSemiring [DivisionSemiring α] : DivisionSemiring αᵃᵒ __ := instSemiring __ := instGroupWithZero nnqsmul := _ + nnqsmul_def := fun q a => rfl nnratCast_def q := unop_injective $ by rw [unop_nnratCast, unop_div, unop_natCast, unop_natCast, NNRat.cast_def, div_eq_mul_inv] @@ -67,6 +70,7 @@ instance instDivisionRing [DivisionRing α] : DivisionRing αᵃᵒᵖ where __ := instRing __ := instDivisionSemiring qsmul := _ + qsmul_def := fun q a => rfl ratCast_def q := unop_injective <| by rw [unop_ratCast, Rat.cast_def, unop_div, unop_natCast, unop_intCast, div_eq_mul_inv] diff --git a/Mathlib/Algebra/Field/Rat.lean b/Mathlib/Algebra/Field/Rat.lean index 4631de2fcb645..fa3f4d1154ec5 100644 --- a/Mathlib/Algebra/Field/Rat.lean +++ b/Mathlib/Algebra/Field/Rat.lean @@ -29,7 +29,9 @@ instance instField : Field ℚ where __ := commRing __ := commGroupWithZero nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q a => rfl nnratCast_def q := by rw [← NNRat.den_coe, ← Int.cast_natCast q.num, ← NNRat.num_coe]; exact(num_div_den _).symm ratCast_def q := (num_div_den _).symm diff --git a/Mathlib/Algebra/FreeAlgebra.lean b/Mathlib/Algebra/FreeAlgebra.lean index bb42056e5700f..1a8ca31f56a0f 100644 --- a/Mathlib/Algebra/FreeAlgebra.lean +++ b/Mathlib/Algebra/FreeAlgebra.lean @@ -234,7 +234,7 @@ instance : Semiring (FreeAlgebra R X) where __ := instDistrib R X natCast n := Quot.mk _ (n : R) natCast_zero := by simp; rfl - natCast_succ n := by simp; exact Quot.sound Rel.add_scalar + natCast_succ n := by simpa using Quot.sound Rel.add_scalar instance : Inhabited (FreeAlgebra R X) := ⟨0⟩ diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 17734817cb862..c551106241d8a 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1498,23 +1498,21 @@ instance (priority := 100) normal_of_comm {G : Type*} [CommGroup G] (H : Subgrou namespace Normal -variable (nH : H.Normal) - @[to_additive] -theorem conj_mem' (n : G) (hn : n ∈ H) (g : G) : +theorem conj_mem' (nH : H.Normal) (n : G) (hn : n ∈ H) (g : G) : g⁻¹ * n * g ∈ H := by convert nH.conj_mem n hn g⁻¹ rw [inv_inv] @[to_additive] -theorem mem_comm {a b : G} (h : a * b ∈ H) : b * a ∈ H := by +theorem mem_comm (nH : H.Normal) {a b : G} (h : a * b ∈ H) : b * a ∈ H := by have : a⁻¹ * (a * b) * a⁻¹⁻¹ ∈ H := nH.conj_mem (a * b) h a⁻¹ -- Porting note: Previous code was: -- simpa simp_all only [inv_mul_cancel_left, inv_inv] @[to_additive] -theorem mem_comm_iff {a b : G} : a * b ∈ H ↔ b * a ∈ H := +theorem mem_comm_iff (nH : H.Normal) {a b : G} : a * b ∈ H ↔ b * a ∈ H := ⟨nH.mem_comm, nH.mem_comm⟩ end Normal diff --git a/Mathlib/Algebra/Group/Subgroup/Order.lean b/Mathlib/Algebra/Group/Subgroup/Order.lean index c6924130e1d6e..36a107b58fb58 100644 --- a/Mathlib/Algebra/Group/Subgroup/Order.lean +++ b/Mathlib/Algebra/Group/Subgroup/Order.lean @@ -3,11 +3,10 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Ruben Van de Velde -/ - +import Mathlib.Order.Atoms import Mathlib.Algebra.Group.Subgroup.Basic -import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Group.InjSurj -import Mathlib.Order.Atoms +import Mathlib.Algebra.Order.Group.Unbundled.Abs /-! # Facts about ordered structures and ordered instances on subgroups diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index c24d31780fa4b..d3a7554aeb851 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -297,13 +297,15 @@ theorem map_id (S : Submonoid M) : S.map (MonoidHom.id M) = S := section GaloisCoinsertion -variable {ι : Type*} {f : F} (hf : Function.Injective f) +variable {ι : Type*} {f : F} /-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ @[to_additive " `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. "] -def gciMapComap : GaloisCoinsertion (map f) (comap f) := +def gciMapComap (hf : Function.Injective f) : GaloisCoinsertion (map f) (comap f) := (gc_map_comap f).toGaloisCoinsertion fun S x => by simp [mem_comap, mem_map, hf.eq_iff] +variable (hf : Function.Injective f) + @[to_additive] theorem comap_map_eq_of_injective (S : Submonoid M) : (S.map f).comap f = S := (gciMapComap hf).u_l_eq _ @@ -999,7 +1001,7 @@ theorem eq_bot_iff_forall : S = ⊥ ↔ ∀ x ∈ S, x = (1 : M) := theorem eq_bot_of_subsingleton [Subsingleton S] : S = ⊥ := by rw [eq_bot_iff_forall] intro y hy - simpa using _root_.congr_arg ((↑) : S → M) <| Subsingleton.elim (⟨y, hy⟩ : S) 1 + simpa using congr_arg ((↑) : S → M) <| Subsingleton.elim (⟨y, hy⟩ : S) 1 @[to_additive] theorem nontrivial_iff_exists_ne_one (S : Submonoid M) : Nontrivial S ↔ ∃ x ∈ S, x ≠ (1 : M) := diff --git a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean index be07bf397537c..e55595fdcf520 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean @@ -301,13 +301,15 @@ theorem map_id (S : Subsemigroup M) : S.map (MulHom.id M) = S := section GaloisCoinsertion -variable {ι : Type*} {f : M →ₙ* N} (hf : Function.Injective f) +variable {ι : Type*} {f : M →ₙ* N} /-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ @[to_additive " `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. "] -def gciMapComap : GaloisCoinsertion (map f) (comap f) := +def gciMapComap (hf : Function.Injective f) : GaloisCoinsertion (map f) (comap f) := (gc_map_comap f).toGaloisCoinsertion fun S x => by simp [mem_comap, mem_map, hf.eq_iff] +variable (hf : Function.Injective f) + @[to_additive] theorem comap_map_eq_of_injective (S : Subsemigroup M) : (S.map f).comap f = S := (gciMapComap hf).u_l_eq _ diff --git a/Mathlib/GroupTheory/GroupAction/Pi.lean b/Mathlib/Algebra/GroupWithZero/Action/Pi.lean similarity index 92% rename from Mathlib/GroupTheory/GroupAction/Pi.lean rename to Mathlib/Algebra/GroupWithZero/Action/Pi.lean index fc4e3e81db122..a68e1a97b1d8a 100644 --- a/Mathlib/GroupTheory/GroupAction/Pi.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Pi.lean @@ -6,19 +6,18 @@ Authors: Simon Hudon, Patrick Massot import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.Algebra.GroupWithZero.Defs -import Mathlib.Data.Set.Function +import Mathlib.Tactic.Common /-! -# Pi instances for multiplicative actions +# Pi instances for multiplicative actions with zero -This file defines instances for `MulAction` and related structures on `Pi` types. +This file defines instances for `MulActionWithZero` and related structures on `Pi` types. ## See also -* `GroupTheory.GroupAction.option` -* `GroupTheory.GroupAction.prod` -* `GroupTheory.GroupAction.sigma` -* `GroupTheory.GroupAction.sum` +* `Algebra.GroupWithZero.Action.Opposite` +* `Algebra.GroupWithZero.Action.Prod` +* `Algebra.GroupWithZero.Action.Units` -/ diff --git a/Mathlib/GroupTheory/GroupAction/Prod.lean b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean similarity index 79% rename from Mathlib/GroupTheory/GroupAction/Prod.lean rename to Mathlib/Algebra/GroupWithZero/Action/Prod.lean index de48351427da4..5d43252e1d1a7 100644 --- a/Mathlib/GroupTheory/GroupAction/Prod.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean @@ -7,29 +7,20 @@ import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.GroupWithZero.Action.Defs /-! -# Prod instances for additive and multiplicative actions -This file defines instances for binary product of additive and multiplicative actions and provides -scalar multiplication as a homomorphism from `α × β` to `β`. -## Main declarations -* `smulMulHom`/`smulMonoidHom`: Scalar multiplication bundled as a multiplicative/monoid - homomorphism. -## See also -* `Mathlib.GroupTheory.GroupAction.Option` -* `Mathlib.GroupTheory.GroupAction.Pi` -* `Mathlib.GroupTheory.GroupAction.Sigma` -* `Mathlib.GroupTheory.GroupAction.Sum` +# Prod instances for multiplicative actions with zero + +This file defines instances for `MulActionWithZero` and related structures on `α × β` -# Porting notes -The `to_additive` attribute can be used to generate both the `smul` and `vadd` lemmas -from the corresponding `pow` lemmas, as explained on zulip here: -https://leanprover.zulipchat.com/#narrow/near/316087838 +## See also -This was not done as part of the port in order to stay as close as possible to the mathlib3 code. +* `Algebra.GroupWithZero.Action.Opposite` +* `Algebra.GroupWithZero.Action.Pi` +* `Algebra.GroupWithZero.Action.Units` -/ assert_not_exists MonoidWithZero -variable {M N P E α β : Type*} +variable {M N α β : Type*} namespace Prod diff --git a/Mathlib/GroupTheory/GroupAction/Units.lean b/Mathlib/Algebra/GroupWithZero/Action/Units.lean similarity index 70% rename from Mathlib/GroupTheory/GroupAction/Units.lean rename to Mathlib/Algebra/GroupWithZero/Action/Units.lean index 067430b131a68..e56fe65a0de40 100644 --- a/Mathlib/GroupTheory/GroupAction/Units.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Units.lean @@ -5,29 +5,30 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Action.Units import Mathlib.Algebra.GroupWithZero.Action.Defs -import Mathlib.Tactic.Common -/-! # Group actions on and by `Mˣ` +/-! +# Multiplicative actions with zero on and by `Mˣ` -This file provides the action of a unit on a type `α`, `SMul Mˣ α`, in the presence of -`SMul M α`, with the obvious definition stated in `Units.smul_def`. This definition preserves -`MulAction` and `DistribMulAction` structures too. +This file provides the multiplicative actions with zero of a unit on a type `α`, `SMul Mˣ α`, in the +presence of `SMulWithZero M α`, with the obvious definition stated in `Units.smul_def`. -Additionally, a `MulAction G M` for some group `G` satisfying some additional properties admits a -`MulAction G Mˣ` structure, again with the obvious definition stated in `Units.coe_smul`. -These instances use a primed name. +Additionally, a `MulDistribMulAction G M` for some group `G` satisfying some additional properties +admits a `MulDistribMulAction G Mˣ` structure, again with the obvious definition stated in +`Units.coe_smul`. This instance uses a primed name. -The results are repeated for `AddUnits` and `VAdd` where relevant. --/ +## See also +* `Algebra.GroupWithZero.Action.Opposite` +* `Algebra.GroupWithZero.Action.Pi` +* `Algebra.GroupWithZero.Action.Prod` +-/ -variable {G H M N : Type*} {α : Type*} +variable {G M α : Type*} namespace Units /-! ### Action of the units of `M` on a type `α` -/ - @[to_additive] instance [Monoid M] [SMul M α] : SMul Mˣ α where smul m a := (m : M) • a diff --git a/Mathlib/Algebra/GroupWithZero/Divisibility.lean b/Mathlib/Algebra/GroupWithZero/Divisibility.lean index 9b138c99c2e9a..9b07a41369654 100644 --- a/Mathlib/Algebra/GroupWithZero/Divisibility.lean +++ b/Mathlib/Algebra/GroupWithZero/Divisibility.lean @@ -129,7 +129,10 @@ end MonoidWithZero section CancelCommMonoidWithZero -variable [CancelCommMonoidWithZero α] [Subsingleton αˣ] {a b : α} {m n : ℕ} +variable [CancelCommMonoidWithZero α] {a b : α} {m n : ℕ} + +section Subsingleton +variable [Subsingleton αˣ] theorem dvd_antisymm : a ∣ b → b ∣ a → a = b := by rintro ⟨c, rfl⟩ ⟨d, hcd⟩ @@ -152,6 +155,8 @@ theorem eq_of_forall_dvd (h : ∀ c, a ∣ c ↔ b ∣ c) : a = b := theorem eq_of_forall_dvd' (h : ∀ c, c ∣ a ↔ c ∣ b) : a = b := ((h _).1 dvd_rfl).antisymm <| (h _).2 dvd_rfl +end Subsingleton + lemma pow_dvd_pow_iff (ha₀ : a ≠ 0) (ha : ¬IsUnit a) : a ^ n ∣ a ^ m ↔ n ≤ m := by constructor · intro h diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 9aa93f3327f8b..70110ada7d9fc 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Nontriviality import Mathlib.Tactic.Spread import Mathlib.Util.AssertExists +import Mathlib.Data.Subtype /-! # Lemmas about units in a `MonoidWithZero` or a `GroupWithZero`. diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index 9febf19604313..af2fe768fa1ac 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Algebra.Group.Units.Hom +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.GroupWithZero.Commute import Mathlib.Algebra.GroupWithZero.Hom -import Mathlib.GroupTheory.GroupAction.Units /-! # Further lemmas about units in a `MonoidWithZero` or a `GroupWithZero`. @@ -30,7 +30,7 @@ end Commute section MonoidWithZero variable [GroupWithZero G₀] [Nontrivial M₀] [MonoidWithZero M₀'] [FunLike F G₀ M₀] - [MonoidWithZeroHomClass F G₀ M₀] [FunLike F' G₀ M₀'] [MonoidWithZeroHomClass F' G₀ M₀'] + [MonoidWithZeroHomClass F G₀ M₀] [FunLike F' G₀ M₀'] (f : F) {a : G₀} @@ -41,8 +41,8 @@ theorem map_ne_zero : f a ≠ 0 ↔ a ≠ 0 := theorem map_eq_zero : f a = 0 ↔ a = 0 := not_iff_not.1 (map_ne_zero f) - -theorem eq_on_inv₀ (f g : F') (h : f a = g a) : f a⁻¹ = g a⁻¹ := by +theorem eq_on_inv₀ [MonoidWithZeroHomClass F' G₀ M₀'] (f g : F') (h : f a = g a) : + f a⁻¹ = g a⁻¹ := by rcases eq_or_ne a 0 with (rfl | ha) · rw [inv_zero, map_zero, map_zero] · exact (IsUnit.mk0 a ha).eq_on_inv f g h diff --git a/Mathlib/Algebra/Homology/CommSq.lean b/Mathlib/Algebra/Homology/CommSq.lean index 974a57e2f64cf..90dad419b073b 100644 --- a/Mathlib/Algebra/Homology/CommSq.lean +++ b/Mathlib/Algebra/Homology/CommSq.lean @@ -38,16 +38,14 @@ variable {C : Type*} [Category C] [Preadditive C] section Pushout variable {f : X₁ ⟶ X₂} {g : X₁ ⟶ X₃} {inl : X₂ ⟶ X₄} {inr : X₃ ⟶ X₄} - (sq : CommSq f g inl inr) - /-- The cokernel cofork attached to a commutative square in a preadditive category. -/ -noncomputable abbrev CommSq.cokernelCofork : +noncomputable abbrev CommSq.cokernelCofork (sq : CommSq f g inl inr) : CokernelCofork (biprod.lift f (-g)) := CokernelCofork.ofπ (biprod.desc inl inr) (by simp [sq.w]) /-- A commutative square in a preadditive category is a pushout square iff the corresponding diagram `X₁ ⟶ X₂ ⊞ X₃ ⟶ X₄ ⟶ 0` makes `X₄` a cokernel. -/ -noncomputable def CommSq.isColimitEquivIsColimitCokernelCofork : +noncomputable def CommSq.isColimitEquivIsColimitCokernelCofork (sq : CommSq f g inl inr) : IsColimit (PushoutCocone.mk _ _ sq.w) ≃ IsColimit sq.cokernelCofork where toFun h := Cofork.IsColimit.mk _ @@ -104,16 +102,15 @@ end Pushout section Pullback variable {fst : X₁ ⟶ X₂} {snd : X₁ ⟶ X₃} {f : X₂ ⟶ X₄} {g : X₃ ⟶ X₄} - (sq : CommSq fst snd f g) /-- The kernel fork attached to a commutative square in a preadditive category. -/ -noncomputable abbrev CommSq.kernelFork : +noncomputable abbrev CommSq.kernelFork (sq : CommSq fst snd f g) : KernelFork (biprod.desc f (-g)) := KernelFork.ofι (biprod.lift fst snd) (by simp [sq.w]) /-- A commutative square in a preadditive category is a pullback square iff the corresponding diagram `0 ⟶ X₁ ⟶ X₂ ⊞ X₃ ⟶ X₄ ⟶ 0` makes `X₁` a kernel. -/ -noncomputable def CommSq.isLimitEquivIsLimitKernelFork : +noncomputable def CommSq.isLimitEquivIsLimitKernelFork (sq : CommSq fst snd f g) : IsLimit (PullbackCone.mk _ _ sq.w) ≃ IsLimit sq.kernelFork where toFun h := Fork.IsLimit.mk _ diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean index 5cab8d20ffe60..6245ca808b6f0 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean @@ -300,8 +300,59 @@ lemma homAddEquiv_apply (α : Ext X Y n) : homAddEquiv α = α.hom := rfl end +variable (X Y Z) in +/-- The composition of `Ext`, as a bilinear map. -/ +@[simps!] +noncomputable def bilinearComp (a b c : ℕ) (h : a + b = c) : + Ext X Y a →+ Ext Y Z b →+ Ext X Z c := + AddMonoidHom.mk' (fun α ↦ AddMonoidHom.mk' (fun β ↦ α.comp β h) (by simp)) (by aesop) + +/-- The postcomposition `Ext X Y a →+ Ext X Z b` with `β : Ext Y Z n` when `a + n = b`. -/ +noncomputable abbrev postcomp (β : Ext Y Z n) (X : C) {a b : ℕ} (h : a + n = b) : + Ext X Y a →+ Ext X Z b := + (bilinearComp X Y Z a n b h).flip β + +/-- The precomposition `Ext Y Z a →+ Ext X Z b` with `α : Ext X Y n` when `n + a = b`. -/ +noncomputable abbrev precomp (α : Ext X Y n) (Z : C) {a b : ℕ} (h : n + a = b) : + Ext Y Z a →+ Ext X Z b := + bilinearComp X Y Z n a b h α + end Ext +/-- Auxiliary definition for `extFunctor`. -/ +@[simps] +noncomputable def extFunctorObj (X : C) (n : ℕ) : C ⥤ AddCommGrp.{w} where + obj Y := AddCommGrp.of (Ext X Y n) + map f := AddCommGrp.ofHom ((Ext.mk₀ f).postcomp _ (add_zero n)) + map_comp f f' := by + ext α + dsimp [AddCommGrp.ofHom] + rw [← Ext.mk₀_comp_mk₀] + symm + apply Ext.comp_assoc + omega + +/-- The functor `Cᵒᵖ ⥤ C ⥤ AddCommGrp` which sends `X : C` and `Y : C` +to `Ext X Y n`. -/ +@[simps] +noncomputable def extFunctor (n : ℕ) : Cᵒᵖ ⥤ C ⥤ AddCommGrp.{w} where + obj X := extFunctorObj X.unop n + map {X₁ X₂} f := + { app := fun Y ↦ AddCommGrp.ofHom (AddMonoidHom.mk' + (fun α ↦ (Ext.mk₀ f.unop).comp α (zero_add _)) (by simp)) + naturality := fun {Y₁ Y₂} g ↦ by + ext α + dsimp + symm + apply Ext.comp_assoc + all_goals omega } + map_comp {X₁ X₂ X₃} f f' := by + ext Y α + dsimp + rw [← Ext.mk₀_comp_mk₀] + apply Ext.comp_assoc + all_goals omega + end Abelian end CategoryTheory diff --git a/Mathlib/Algebra/Homology/Embedding/Boundary.lean b/Mathlib/Algebra/Homology/Embedding/Boundary.lean index f0f5f70e7f186..046809631f6a2 100644 --- a/Mathlib/Algebra/Homology/Embedding/Boundary.lean +++ b/Mathlib/Algebra/Homology/Embedding/Boundary.lean @@ -36,7 +36,7 @@ namespace ComplexShape namespace Embedding -variable {ι ι' : Type*} (c : ComplexShape ι) (c' : ComplexShape ι') (e : Embedding c c') +variable {ι ι' : Type*} {c : ComplexShape ι} {c' : ComplexShape ι'} (e : Embedding c c') /-- The lower boundary of an embedding `e : Embedding c c'`, as a predicate on `ι`. It is satisfied by `j : ι` when there exists `i' : ι'` not in the image of `e.f` @@ -87,7 +87,7 @@ lemma prev_f_of_not_boundaryGE [e.IsRelIff] {i j : ι} (hij : c.prev j = i) exact hij' (by simpa only [hij] using hi) variable {e} in -lemma BoundaryGE.false {j : ι} (hj : e.BoundaryGE j) [e.IsTruncLE] : False := by +lemma BoundaryGE.false_of_isTruncLE {j : ι} (hj : e.BoundaryGE j) [e.IsTruncLE] : False := by obtain ⟨i, hi⟩ := e.mem_prev hj.1 exact hj.2 i (by simpa only [hi] using hj.1) @@ -140,7 +140,7 @@ lemma next_f_of_not_boundaryLE [e.IsRelIff] {j k : ι} (hjk : c.next j = k) exact hjk' (by simpa only [hjk] using hk) variable {e} in -lemma BoundaryLE.false {j : ι} (hj : e.BoundaryLE j) [e.IsTruncGE] : False := by +lemma BoundaryLE.false_of_isTruncGE {j : ι} (hj : e.BoundaryLE j) [e.IsTruncGE] : False := by obtain ⟨k, hk⟩ := e.mem_next hj.1 exact hj.2 k (by simpa only [hk] using hj.1) diff --git a/Mathlib/Algebra/Homology/Embedding/TruncGE.lean b/Mathlib/Algebra/Homology/Embedding/TruncGE.lean new file mode 100644 index 0000000000000..1cb72b2a91514 --- /dev/null +++ b/Mathlib/Algebra/Homology/Embedding/TruncGE.lean @@ -0,0 +1,147 @@ +/- +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.Boundary +import Mathlib.Algebra.Homology.Embedding.Extend +import Mathlib.Algebra.Homology.ShortComplex.HomologicalComplex + +/-! +# The canonical truncation + +Given an embedding `e : Embedding c c'` of complex shapes which +satisfies `e.IsTruncGE` and `K : HomologicalComplex C c'`, +we define `K.truncGE' e : HomologicalComplex C c` +and `K.truncGE e : HomologicalComplex C c'` which are the canonical +truncations of `K` relative to `e`. + +For example, if `e` is the embedding `embeddingUpIntGE p` of `ComplexShape.up ℕ` +in `ComplexShape.up ℤ` which sends `n : ℕ` to `p + n` and `K : CochainComplex C ℤ`, +then `K.truncGE' e : CochainComplex C ℕ` is the following complex: + +`Q ⟶ K.X (p + 1) ⟶ K.X (p + 2) ⟶ K.X (p + 3) ⟶ ...` + +where in degree `0`, the object `Q` identifies to the cokernel +of `K.X (p - 1) ⟶ K.X p` (this is `K.opcycles p`). Then, the +cochain complex `K.truncGE e` is indexed by `ℤ`, and has the +following shape: + +`... ⟶ 0 ⟶ 0 ⟶ 0 ⟶ Q ⟶ K.X (p + 1) ⟶ K.X (p + 2) ⟶ K.X (p + 3) ⟶ ...` + +where `Q` is in degree `p`. + +## TODO +* construct a morphism `K.πTruncGE e : K ⟶ K.truncGE e` and show that +it induces an isomorphism in homology in degrees in the image of `e.f`. + +-/ + +open CategoryTheory Limits ZeroObject Category + +variable {ι ι' : Type*} {c : ComplexShape ι} {c' : ComplexShape ι'} + {C : Type*} [Category C] [HasZeroMorphisms C] [HasZeroObject C] + +namespace HomologicalComplex + +variable (K L M : HomologicalComplex C c') (φ : K ⟶ L) (φ' : L ⟶ M) + (e : c.Embedding c') [e.IsTruncGE] + [∀ i', K.HasHomology i'] [∀ i', L.HasHomology i'] [∀ i', M.HasHomology i'] + +namespace truncGE' + +open Classical in +/-- The `X` field of `truncGE'`. -/ +noncomputable def X (i : ι) : C := + if e.BoundaryGE i + then K.opcycles (e.f i) + else K.X (e.f i) + +/-- The isomorphism `truncGE'.X K e i ≅ K.opcycles (e.f i)` when `e.BoundaryGE i` holds.-/ +noncomputable def XIsoOpcycles {i : ι} (hi : e.BoundaryGE i) : + X K e i ≅ K.opcycles (e.f i) := + eqToIso (if_pos hi) + +/-- The isomorphism `truncGE'.X K e i ≅ K.X (e.f i)` when `e.BoundaryGE i` does not hold.-/ +noncomputable def XIso {i : ι} (hi : ¬ e.BoundaryGE i) : + X K e i ≅ K.X (e.f i) := + eqToIso (if_neg hi) + +open Classical in +/-- The `d` field of `truncGE'`. -/ +noncomputable def d (i j : ι) : X K e i ⟶ X K e j := + if hij : c.Rel i j + then + if hi : e.BoundaryGE i + then (truncGE'.XIsoOpcycles K e hi).hom ≫ K.fromOpcycles (e.f i) (e.f j) ≫ + (XIso K e (e.not_boundaryGE_next hij)).inv + else (XIso K e hi).hom ≫ K.d (e.f i) (e.f j) ≫ + (XIso K e (e.not_boundaryGE_next hij)).inv + else 0 + +@[reassoc (attr := simp)] +lemma d_comp_d (i j k : ι) : d K e i j ≫ d K e j k = 0 := by + dsimp [d] + by_cases hij : c.Rel i j + · by_cases hjk : c.Rel j k + · rw [dif_pos hij, dif_pos hjk, dif_neg (e.not_boundaryGE_next hij)] + split_ifs <;> simp + · rw [dif_neg hjk, comp_zero] + · rw [dif_neg hij, zero_comp] + +end truncGE' + +/-- The canonical truncation of a homological complex relative to an embedding +of complex shapes `e` which satisfies `e.IsTruncGE`. -/ +noncomputable def truncGE' : HomologicalComplex C c where + X := truncGE'.X K e + d := truncGE'.d K e + shape _ _ h := dif_neg h + +/-- The isomorphism `(K.truncGE' e).X i ≅ K.X i'` when `e.f i = i'` +and `e.BoundaryGE i` does not hold. -/ +noncomputable def truncGE'XIso {i : ι} {i' : ι'} (hi' : e.f i = i') (hi : ¬ e.BoundaryGE i) : + (K.truncGE' e).X i ≅ K.X i' := + (truncGE'.XIso K e hi) ≪≫ eqToIso (by subst hi'; rfl) + +/-- The isomorphism `(K.truncGE' e).X i ≅ K.opcycles i'` when `e.f i = i'` +and `e.BoundaryGE i` holds. -/ +noncomputable def truncGE'XIsoOpcycles {i : ι} {i' : ι'} (hi' : e.f i = i') (hi : e.BoundaryGE i) : + (K.truncGE' e).X i ≅ K.opcycles i' := + (truncGE'.XIsoOpcycles K e hi) ≪≫ eqToIso (by subst hi'; rfl) + +lemma truncGE'_d_eq {i j : ι} (hij : c.Rel i j) {i' j' : ι'} + (hi' : e.f i = i') (hj' : e.f j = j') (hi : ¬ e.BoundaryGE i) : + (K.truncGE' e).d i j = (K.truncGE'XIso e hi' hi).hom ≫ K.d i' j' ≫ + (K.truncGE'XIso e hj' (e.not_boundaryGE_next hij)).inv := by + dsimp [truncGE', truncGE'.d] + rw [dif_pos hij, dif_neg hi] + subst hi' hj' + simp [truncGE'XIso] + +lemma truncGE'_d_eq_fromOpcycles {i j : ι} (hij : c.Rel i j) {i' j' : ι'} + (hi' : e.f i = i') (hj' : e.f j = j') (hi : e.BoundaryGE i) : + (K.truncGE' e).d i j = (K.truncGE'XIsoOpcycles e hi' hi).hom ≫ K.fromOpcycles i' j' ≫ + (K.truncGE'XIso e hj' (e.not_boundaryGE_next hij)).inv := by + dsimp [truncGE', truncGE'.d] + rw [dif_pos hij, dif_pos hi] + subst hi' hj' + simp [truncGE'XIso, truncGE'XIsoOpcycles] + +/-- The canonical truncation of a homological complex relative to an embedding +of complex shapes `e` which satisfies `e.IsTruncGE`. -/ +noncomputable def truncGE : HomologicalComplex C c' := (K.truncGE' e).extend e + +/-- The isomorphism `(K.truncGE e).X i' ≅ K.X i'` when `e.f i = i'` +and `e.BoundaryGE i` does not hold. -/ +noncomputable def truncGEXIso {i : ι} {i' : ι'} (hi' : e.f i = i') (hi : ¬ e.BoundaryGE i) : + (K.truncGE e).X i' ≅ K.X i' := + (K.truncGE' e).extendXIso e hi' ≪≫ K.truncGE'XIso e hi' hi + +/-- The isomorphism `(K.truncGE e).X i' ≅ K.opcycles i'` when `e.f i = i'` +and `e.BoundaryGE i` holds. -/ +noncomputable def truncGEXIsoOpcycles {i : ι} {i' : ι'} (hi' : e.f i = i') (hi : e.BoundaryGE i) : + (K.truncGE e).X i' ≅ K.opcycles i' := + (K.truncGE' e).extendXIso e hi' ≪≫ K.truncGE'XIsoOpcycles e hi' hi + +end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean index 2b9e64ece11b7..6c3f809ed38d6 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Exact.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Exact.lean @@ -703,37 +703,35 @@ variable [Balanced C] namespace Exact -variable (hS : S.Exact) - -lemma isIso_f' (h : S.LeftHomologyData) [Mono S.f] : +lemma isIso_f' (hS : S.Exact) (h : S.LeftHomologyData) [Mono S.f] : IsIso h.f' := by have := hS.epi_f' h have := mono_of_mono_fac h.f'_i exact isIso_of_mono_of_epi h.f' -lemma isIso_toCycles [Mono S.f] [S.HasLeftHomology] : +lemma isIso_toCycles (hS : S.Exact) [Mono S.f] [S.HasLeftHomology]: IsIso S.toCycles := hS.isIso_f' _ -lemma isIso_g' (h : S.RightHomologyData) [Epi S.g] : +lemma isIso_g' (hS : S.Exact) (h : S.RightHomologyData) [Epi S.g] : IsIso h.g' := by have := hS.mono_g' h have := epi_of_epi_fac h.p_g' exact isIso_of_mono_of_epi h.g' -lemma isIso_fromOpcycles [Epi S.g] [S.HasRightHomology] : +lemma isIso_fromOpcycles (hS : S.Exact) [Epi S.g] [S.HasRightHomology] : IsIso S.fromOpcycles := hS.isIso_g' _ /-- In a balanced category, if a short complex `S` is exact and `S.f` is a mono, then `S.X₁` is the kernel of `S.g`. -/ -noncomputable def fIsKernel [Mono S.f] : IsLimit (KernelFork.ofι S.f S.zero) := by +noncomputable def fIsKernel (hS : S.Exact) [Mono S.f] : IsLimit (KernelFork.ofι S.f S.zero) := by have := hS.hasHomology have := hS.isIso_toCycles exact IsLimit.ofIsoLimit S.cyclesIsKernel (Fork.ext (asIso S.toCycles).symm (by simp)) -lemma map_of_mono_of_preservesKernel (F : C ⥤ D) +lemma map_of_mono_of_preservesKernel (hS : S.Exact) (F : C ⥤ D) [F.PreservesZeroMorphisms] [(S.map F).HasHomology] (_ : Mono S.f) (_ : PreservesLimit (parallelPair S.g 0) F) : (S.map F).Exact := @@ -741,13 +739,14 @@ lemma map_of_mono_of_preservesKernel (F : C ⥤ D) /-- In a balanced category, if a short complex `S` is exact and `S.g` is an epi, then `S.X₃` is the cokernel of `S.g`. -/ -noncomputable def gIsCokernel [Epi S.g] : IsColimit (CokernelCofork.ofπ S.g S.zero) := by +noncomputable def gIsCokernel (hS : S.Exact) [Epi S.g] : + IsColimit (CokernelCofork.ofπ S.g S.zero) := by have := hS.hasHomology have := hS.isIso_fromOpcycles exact IsColimit.ofIsoColimit S.opcyclesIsCokernel (Cofork.ext (asIso S.fromOpcycles) (by simp)) -lemma map_of_epi_of_preservesCokernel (F : C ⥤ D) +lemma map_of_epi_of_preservesCokernel (hS : S.Exact) (F : C ⥤ D) [F.PreservesZeroMorphisms] [(S.map F).HasHomology] (_ : Epi S.g) (_ : PreservesColimit (parallelPair S.f 0) F) : (S.map F).Exact := @@ -755,29 +754,29 @@ lemma map_of_epi_of_preservesCokernel (F : C ⥤ D) /-- If a short complex `S` in a balanced category is exact and such that `S.f` is a mono, then a morphism `k : A ⟶ S.X₂` such that `k ≫ S.g = 0` lifts to a morphism `A ⟶ S.X₁`. -/ -noncomputable def lift {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : +noncomputable def lift (hS : S.Exact) {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : A ⟶ S.X₁ := hS.fIsKernel.lift (KernelFork.ofι k hk) @[reassoc (attr := simp)] -lemma lift_f {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : +lemma lift_f (hS : S.Exact) {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : hS.lift k hk ≫ S.f = k := Fork.IsLimit.lift_ι _ -lemma lift' {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : +lemma lift' (hS : S.Exact) {A : C} (k : A ⟶ S.X₂) (hk : k ≫ S.g = 0) [Mono S.f] : ∃ (l : A ⟶ S.X₁), l ≫ S.f = k := ⟨hS.lift k hk, by simp⟩ /-- If a short complex `S` in a balanced category is exact and such that `S.g` is an epi, then a morphism `k : S.X₂ ⟶ A` such that `S.f ≫ k = 0` descends to a morphism `S.X₃ ⟶ A`. -/ -noncomputable def desc {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : +noncomputable def desc (hS : S.Exact) {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : S.X₃ ⟶ A := hS.gIsCokernel.desc (CokernelCofork.ofπ k hk) @[reassoc (attr := simp)] -lemma g_desc {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : +lemma g_desc (hS : S.Exact) {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : S.g ≫ hS.desc k hk = k := Cofork.IsColimit.π_desc (hS.gIsCokernel) -lemma desc' {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : +lemma desc' (hS : S.Exact) {A : C} (k : S.X₂ ⟶ A) (hk : S.f ≫ k = 0) [Epi S.g] : ∃ (l : S.X₃ ⟶ A), S.g ≫ l = k := ⟨hS.desc k hk, by simp⟩ diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index fc0d579ff6eee..1aa00c6f31818 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -27,7 +27,7 @@ theorem isPrimePow_def : IsPrimePow n ↔ ∃ (p : R) (k : ℕ), Prime p ∧ 0 < natural `k` such that `n` can be written as `p^(k+1)`. -/ theorem isPrimePow_iff_pow_succ : IsPrimePow n ↔ ∃ (p : R) (k : ℕ), Prime p ∧ p ^ (k + 1) = n := (isPrimePow_def _).trans - ⟨fun ⟨p, k, hp, hk, hn⟩ => ⟨_, _, hp, by rwa [Nat.sub_add_cancel hk]⟩, fun ⟨p, k, hp, hn⟩ => + ⟨fun ⟨p, k, hp, hk, hn⟩ => ⟨p, k - 1, hp, by rwa [Nat.sub_add_cancel hk]⟩, fun ⟨p, k, hp, hn⟩ => ⟨_, _, hp, Nat.succ_pos', hn⟩⟩ theorem not_isPrimePow_zero [NoZeroDivisors R] : ¬IsPrimePow (0 : R) := by diff --git a/Mathlib/Algebra/Lie/DirectSum.lean b/Mathlib/Algebra/Lie/DirectSum.lean index 03d1186e39bd6..7d4a6efc56417 100644 --- a/Mathlib/Algebra/Lie/DirectSum.lean +++ b/Mathlib/Algebra/Lie/DirectSum.lean @@ -124,8 +124,8 @@ theorem lie_of_of_ne [DecidableEq ι] {i j : ι} (hij : i ≠ j) (x : L i) (y : refine DFinsupp.ext fun k => ?_ rw [bracket_apply] obtain rfl | hik := Decidable.eq_or_ne i k - · rw [of_eq_of_ne _ _ _ _ hij.symm, lie_zero, zero_apply] - · rw [of_eq_of_ne _ _ _ _ hik, zero_lie, zero_apply] + · rw [of_eq_of_ne _ _ _ hij.symm, lie_zero, zero_apply] + · rw [of_eq_of_ne _ _ _ hik, zero_lie, zero_apply] @[simp] theorem lie_of [DecidableEq ι] {i j : ι} (x : L i) (y : L j) : diff --git a/Mathlib/Algebra/Lie/Sl2.lean b/Mathlib/Algebra/Lie/Sl2.lean index f6d7bf223160c..799f311607911 100644 --- a/Mathlib/Algebra/Lie/Sl2.lean +++ b/Mathlib/Algebra/Lie/Sl2.lean @@ -154,8 +154,8 @@ lemma pow_toEnd_f_ne_zero_of_eq_nat · next i IH => have : ((i + 1) * (n - i) : ℤ) • (toEnd R L M f ^ i) m = 0 := by have := congr_arg (⁅e, ·⁆) H - simpa [zsmul_eq_smul_cast R, P.lie_e_pow_succ_toEnd_f, hn] using this - rw [zsmul_eq_smul_cast R, smul_eq_zero, Int.cast_eq_zero, mul_eq_zero, sub_eq_zero, + simpa [← Int.cast_smul_eq_nsmul R, P.lie_e_pow_succ_toEnd_f, hn] using this + rw [← Int.cast_smul_eq_nsmul R, smul_eq_zero, Int.cast_eq_zero, mul_eq_zero, sub_eq_zero, Nat.cast_inj, ← @Nat.cast_one ℤ, ← Nat.cast_add, Nat.cast_eq_zero] at this simp only [add_eq_zero, one_ne_zero, and_false, false_or] at this exact (hi.trans_eq (this.resolve_right (IH (i.le_succ.trans hi)))).not_lt i.lt_succ_self diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index b99b673c5b65f..6392851189478 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -418,7 +418,7 @@ lemma range_traceForm_le_span_weight : rw [LieModule.traceForm_eq_sum_finrank_nsmul, LinearMap.coeFn_sum, Finset.sum_apply] refine Submodule.sum_mem _ fun χ _ ↦ ?_ simp_rw [LinearMap.smul_apply, LinearMap.coe_smulRight, Weight.toLinear_apply, - nsmul_eq_smul_cast (R := K)] + ← Nat.cast_smul_eq_nsmul K] exact Submodule.smul_mem _ _ <| Submodule.smul_mem _ _ <| subset_span <| mem_range_self χ end LieModule diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index 85bba329103f4..e17fe80cd32a6 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -265,7 +265,7 @@ lemma weightSpace_chainTopCoeff_add_one_nsmul_add : lemma weightSpace_chainTopCoeff_add_one_zsmul_add : weightSpace M ((chainTopCoeff α β + 1 : ℤ) • α + β : L → R) = ⊥ := by - rw [← weightSpace_chainTopCoeff_add_one_nsmul_add α β hα, nsmul_eq_smul_cast ℤ, + rw [← weightSpace_chainTopCoeff_add_one_nsmul_add α β hα, ← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_add, Nat.cast_one] lemma weightSpace_nsmul_add_ne_bot_of_le {n} (hn : n ≤ chainTopCoeff α β) : @@ -285,10 +285,10 @@ lemma weightSpace_zsmul_add_ne_bot {n : ℤ} (hn : -chainBotCoeff α β ≤ n) (hn' : n ≤ chainTopCoeff α β) : weightSpace M (n • α + β : L → R) ≠ ⊥ := by rcases n with (n | n) - · simp only [Int.ofNat_eq_coe, Nat.cast_le, ← nsmul_eq_smul_cast] at hn' ⊢ + · simp only [Int.ofNat_eq_coe, Nat.cast_le, Nat.cast_smul_eq_nsmul] at hn' ⊢ exact weightSpace_nsmul_add_ne_bot_of_le α β hn' · simp only [Int.negSucc_eq, ← Nat.cast_succ, neg_le_neg_iff, Nat.cast_le] at hn ⊢ - rw [neg_smul, ← smul_neg, ← nsmul_eq_smul_cast] + rw [neg_smul, ← smul_neg, Nat.cast_smul_eq_nsmul] exact weightSpace_nsmul_add_ne_bot_of_le (-α) β hn lemma weightSpace_neg_zsmul_add_ne_bot {n : ℕ} (hn : n ≤ chainBotCoeff α β) : @@ -308,7 +308,7 @@ def chainBot (α : L → R) (β : Weight R L M) : Weight R L M := lemma coe_chainTop' : (chainTop α β : L → R) = chainTopCoeff α β • α + β := rfl @[simp] lemma coe_chainTop : (chainTop α β : L → R) = (chainTopCoeff α β : ℤ) • α + β := by - rw [← nsmul_eq_smul_cast ℤ]; rfl + rw [Nat.cast_smul_eq_nsmul ℤ]; rfl @[simp] lemma coe_chainBot : (chainBot α β : L → R) = (-chainBotCoeff α β : ℤ) • α + β := rfl @[simp] lemma chainTop_neg : chainTop (-α) β = chainBot α β := by ext; simp diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index 190ec7f235eb0..f9fb2a4d49c1e 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -397,7 +397,7 @@ lemma coe_corootSpace_eq_span_singleton (α : Weight K H L) : suffices (K ∙ coroot α) = K ∙ α' by rw [coe_corootSpace_eq_span_singleton']; exact this.symm have : IsUnit (2 * (α α')⁻¹) := by simpa using root_apply_cartanEquivDual_symm_ne_zero hα change (K ∙ (2 • (α α')⁻¹ • α')) = _ - simpa [nsmul_eq_smul_cast (R := K), smul_smul] using Submodule.span_singleton_smul_eq this _ + simpa [← Nat.cast_smul_eq_nsmul K, smul_smul] using Submodule.span_singleton_smul_eq this _ @[simp] lemma corootSpace_eq_bot_iff {α : Weight K H L} : diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index 0e2306a73fb76..184ab5adcd0f0 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -54,7 +54,7 @@ private lemma chainLength_aux {x} (hx : x ∈ rootSpace H (chainTop α β)) : ⟨hx', by rw [← lie_eq_smul_of_mem_rootSpace hx]; rfl, by rwa [weightSpace_add_chainTop α β hα] at this⟩ obtain ⟨μ, hμ⟩ := this.exists_nat - exact ⟨μ, by rw [nsmul_eq_smul_cast K, ← hμ, lie_eq_smul_of_mem_rootSpace hx]⟩ + exact ⟨μ, by rw [← Nat.cast_smul_eq_nsmul K, ← hμ, lie_eq_smul_of_mem_rootSpace hx]⟩ /-- The length of the `α`-chain through `β`. See `chainBotCoeff_add_chainTopCoeff`. -/ def chainLength (α β : Weight K H L) : ℕ := @@ -78,7 +78,7 @@ lemma chainLength_nsmul {x} (hx : x ∈ rootSpace H (chainTop α β)) : lemma chainLength_smul {x} (hx : x ∈ rootSpace H (chainTop α β)) : (chainLength α β : K) • x = ⁅coroot α, x⁆ := by - rw [← nsmul_eq_smul_cast, chainLength_nsmul _ _ hx] + rw [Nat.cast_smul_eq_nsmul, chainLength_nsmul _ _ hx] lemma apply_coroot_eq_cast' : β (coroot α) = ↑(chainLength α β - 2 * chainTopCoeff α β : ℤ) := by @@ -121,7 +121,7 @@ lemma rootSpace_neg_nsmul_add_chainTop_of_lt {n : ℕ} (hn : chainLength α β < linarith [this, hn] have H₂ : ((1 + n + chainTopCoeff (-α) W) • α + chainTop (-α) W : H → K) = (chainTopCoeff α β + 1) • α + β := by - simp only [Weight.coe_neg, nsmul_eq_smul_cast ℤ, Nat.cast_add, Nat.cast_one, coe_chainTop, + simp only [Weight.coe_neg, ← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_add, Nat.cast_one, coe_chainTop, smul_neg, ← neg_smul, hW, ← add_assoc, ← add_smul, ← sub_eq_add_neg] congr 2 ring @@ -136,8 +136,8 @@ lemma chainTopCoeff_le_chainLength : chainTopCoeff α β ≤ chainLength α β : intro e apply weightSpace_nsmul_add_ne_bot_of_le α β (Nat.sub_le (chainTopCoeff α β) (chainLength α β).succ) - rw [nsmul_eq_smul_cast ℤ, Nat.cast_sub e, sub_smul, sub_eq_neg_add, - add_assoc, ← coe_chainTop, ← nsmul_eq_smul_cast] + rw [← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_sub e, sub_smul, sub_eq_neg_add, + add_assoc, ← coe_chainTop, Nat.cast_smul_eq_nsmul] exact rootSpace_neg_nsmul_add_chainTop_of_lt α β hα (Nat.lt_succ_self _) lemma chainBotCoeff_add_chainTopCoeff : @@ -149,17 +149,17 @@ lemma chainBotCoeff_add_chainTopCoeff : ← not_lt, ← Nat.succ_le, chainBotCoeff, ← Weight.coe_neg] intro e apply weightSpace_nsmul_add_ne_bot_of_le _ _ e - rw [nsmul_eq_smul_cast ℤ, Nat.cast_succ, Nat.cast_sub (chainTopCoeff_le_chainLength α β), + rw [← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_succ, Nat.cast_sub (chainTopCoeff_le_chainLength α β), LieModule.Weight.coe_neg, smul_neg, ← neg_smul, neg_add_rev, neg_sub, sub_eq_neg_add, ← add_assoc, ← neg_add_rev, add_smul, add_assoc, ← coe_chainTop, neg_smul, - ← @Nat.cast_one ℤ, ← Nat.cast_add, ← nsmul_eq_smul_cast] + ← @Nat.cast_one ℤ, ← Nat.cast_add, Nat.cast_smul_eq_nsmul] exact rootSpace_neg_nsmul_add_chainTop_of_lt α β hα (Nat.lt_succ_self _) · rw [← not_lt] intro e apply rootSpace_neg_nsmul_add_chainTop_of_le α β e - rw [← Nat.succ_add, nsmul_eq_smul_cast ℤ, ← neg_smul, coe_chainTop, ← add_assoc, ← add_smul, - Nat.cast_add, neg_add, add_assoc, neg_add_self, add_zero, neg_smul, ← smul_neg, - ← nsmul_eq_smul_cast] + rw [← Nat.succ_add, ← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, coe_chainTop, ← add_assoc, + ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_self, add_zero, neg_smul, ← smul_neg, + Nat.cast_smul_eq_nsmul] exact weightSpace_chainTopCoeff_add_one_nsmul_add (-α) β (Weight.IsNonZero.neg hα) lemma chainTopCoeff_add_chainBotCoeff : @@ -192,7 +192,7 @@ lemma le_chainBotCoeff_of_rootSpace_ne_top (n : ℤ) (hn : rootSpace H (-n • rw [Nat.cast_lt, ← @Nat.add_lt_add_iff_right (chainTopCoeff α β), chainBotCoeff_add_chainTopCoeff] at hn have := rootSpace_neg_nsmul_add_chainTop_of_lt α β hα hn - rwa [nsmul_eq_smul_cast ℤ, ← neg_smul, coe_chainTop, ← add_assoc, + rwa [← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, coe_chainTop, ← add_assoc, ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_self, add_zero] at this /-- Members of the `α`-chain through `β` are the only roots of the form `β - kα`. -/ @@ -211,7 +211,7 @@ lemma rootSpace_zsmul_add_ne_bot_iff (n : ℤ) : simp only [neg_sub, tsub_le_iff_right, ← Nat.cast_add, Nat.cast_le, chainBotCoeff_add_chainTopCoeff] at h₂ have := rootSpace_neg_nsmul_add_chainTop_of_le α β h₂ - rwa [coe_chainTop, nsmul_eq_smul_cast ℤ, ← neg_smul, + rwa [coe_chainTop, ← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, ← add_assoc, ← add_smul, ← sub_eq_neg_add] at this lemma rootSpace_zsmul_add_ne_bot_iff_mem (n : ℤ) : @@ -293,7 +293,7 @@ lemma rootSpace_one_div_two_smul : rootSpace H ((2⁻¹ : K) • α) = ⊥ := by let W : Weight K H L := ⟨_, h⟩ have hW : 2 • (W : H → K) = α := by show 2 • (2⁻¹ : K) • (α : H → K) = α - rw [nsmul_eq_smul_cast K, smul_smul]; simp + rw [← Nat.cast_smul_eq_nsmul K, smul_smul]; simp apply α.weightSpace_ne_bot have := rootSpace_two_smul W (fun (e : (W : H → K) = 0) ↦ hα <| by apply_fun (2 • ·) at e; simpa [hW] using e) @@ -313,7 +313,7 @@ lemma eq_neg_one_or_eq_zero_or_eq_one_of_eq_smul (k : K) (h : (β : H → K) = k mul_eq_mul_left_iff, OfNat.ofNat_ne_zero, or_false] at H rw [← Int.cast_natCast, ← Int.cast_natCast (chainTopCoeff α β), ← Int.cast_sub] at H have := (rootSpace_zsmul_add_ne_bot_iff_mem α 0 hα (n - chainTopCoeff α β)).mp - (by rw [zsmul_eq_smul_cast K, ← H, ← h, Weight.coe_zero, add_zero]; exact β.2) + (by rw [← Int.cast_smul_eq_nsmul K, ← H, ← h, Weight.coe_zero, add_zero]; exact β.2) rw [chainTopCoeff_zero_right α hα, chainBotCoeff_zero_right α hα, Nat.cast_one] at this set k' : ℤ := n - chainTopCoeff α β subst H @@ -330,7 +330,7 @@ lemma eq_neg_one_or_eq_zero_or_eq_one_of_eq_smul (k : K) (h : (β : H → K) = k · simp only [tsub_le_iff_right, le_add_iff_nonneg_right, Nat.cast_nonneg, neg_sub, true_and] rw [← Nat.cast_add, chainBotCoeff_add_chainTopCoeff, hn] omega - rw [h, hk, zsmul_eq_smul_cast K, ← add_smul] at this + rw [h, hk, ← Int.cast_smul_eq_nsmul K, ← add_smul] at this simp only [Int.cast_sub, Int.cast_natCast, sub_add_sub_cancel', add_sub_cancel_left, ne_eq] at this cases this (rootSpace_one_div_two_smul α hα) @@ -351,8 +351,8 @@ def reflectRoot (α β : Weight K H L) : Weight K H L where weightSpace_ne_bot' := by by_cases hα : α.IsZero · simpa [hα.eq] using β.weightSpace_ne_bot - rw [sub_eq_neg_add, apply_coroot_eq_cast α β, ← neg_smul, ← Int.cast_neg, ← zsmul_eq_smul_cast, - rootSpace_zsmul_add_ne_bot_iff α β hα] + rw [sub_eq_neg_add, apply_coroot_eq_cast α β, ← neg_smul, ← Int.cast_neg, + Int.cast_smul_eq_nsmul, rootSpace_zsmul_add_ne_bot_iff α β hα] omega lemma reflectRoot_isNonZero (α β : Weight K H L) (hβ : β.IsNonZero) : diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 23e8a7b1b35c4..1edc9c1b8fdf9 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -104,7 +104,7 @@ section Module variable [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] instance [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := - ⟨fun {c x} hcx ↦ by rwa [nsmul_eq_smul_cast ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ + ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ end Module diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index bd4b6c5df495e..25c5c9fd0303f 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Group.Hom.End +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Ring.Invertible import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.SMulWithZero import Mathlib.Data.Int.Cast.Lemmas -import Mathlib.GroupTheory.GroupAction.Units /-! # Modules over a ring @@ -326,18 +326,27 @@ section variable (R) /-- `nsmul` is equal to any other module structure via a cast. -/ -theorem nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := by +lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by induction' n with n ih · rw [Nat.cast_zero, zero_smul, zero_smul] · rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] +/-- `nsmul` is equal to any other module structure via a cast. -/ +-- See note [no_index around OfNat.ofNat] +lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : + (no_index OfNat.ofNat n : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. + +/-- `nsmul` is equal to any other module structure via a cast. -/ +@[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] +lemma nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := (Nat.cast_smul_eq_nsmul ..).symm + end /-- Convert back any exotic `ℕ`-smul to the canonical instance. This should not be needed since in mathlib all `AddCommMonoid`s should normally have exactly one `ℕ`-module structure by design. -/ -theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : - @SMul.smul ℕ M h.toSMul n x = n • x := by rw [nsmul_eq_smul_cast ℕ n x, Nat.cast_id]; rfl +theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : @SMul.smul ℕ M h.toSMul n x = n • x := + Nat.cast_smul_eq_nsmul .. /-- All `ℕ`-module structures are equal. Not an instance since in mathlib all `AddCommMonoid` should normally have exactly one `ℕ`-module structure by design. -/ @@ -361,18 +370,24 @@ section variable (R) /-- `zsmul` is equal to any other module structure via a cast. -/ -theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := - have : (smulAddHom ℤ M).flip b = ((smulAddHom R M).flip b).comp (Int.castAddHom R) := by +lemma Int.cast_smul_eq_nsmul (n : ℤ) (b : M) : (n : R) • b = n • b := + have : ((smulAddHom R M).flip b).comp (Int.castAddHom R) = (smulAddHom ℤ M).flip b := by apply AddMonoidHom.ext_int simp DFunLike.congr_fun this n +@[deprecated (since := "2024-07-23")] alias intCast_smul := Int.cast_smul_eq_nsmul + +/-- `zsmul` is equal to any other module structure via a cast. -/ +@[deprecated Int.cast_smul_eq_nsmul (since := "2024-07-23")] +theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := (Int.cast_smul_eq_nsmul ..).symm + end /-- Convert back any exotic `ℤ`-smul to the canonical instance. This should not be needed since in mathlib all `AddCommGroup`s should normally have exactly one `ℤ`-module structure by design. -/ -theorem int_smul_eq_zsmul (h : Module ℤ M) (n : ℤ) (x : M) : - @SMul.smul ℤ M h.toSMul n x = n • x := by rw [zsmul_eq_smul_cast ℤ n x, Int.cast_id]; rfl +theorem int_smul_eq_zsmul (h : Module ℤ M) (n : ℤ) (x : M) : @SMul.smul ℤ M h.toSMul n x = n • x := + Int.cast_smul_eq_nsmul .. /-- All `ℤ`-module structures are equal. Not an instance since in mathlib all `AddCommGroup` should normally have exactly one `ℤ`-module structure by design. -/ @@ -385,12 +400,12 @@ end AddCommGroup theorem map_intCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Ring R] [Ring S] [Module R M] [Module S M₂] (x : ℤ) (a : M) : - f ((x : R) • a) = (x : S) • f a := by simp only [← zsmul_eq_smul_cast, map_zsmul] + f ((x : R) • a) = (x : S) • f a := by simp only [Int.cast_smul_eq_nsmul, map_zsmul] theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by - simp only [← nsmul_eq_smul_cast, AddMonoidHom.map_nsmul, map_nsmul] + simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] instance AddCommGroup.intIsScalarTower {R : Type u} {M : Type v} [Ring R] [AddCommGroup M] [Module R M] : IsScalarTower ℤ R M where @@ -463,11 +478,8 @@ section Nat variable [NoZeroSMulDivisors R M] [CharZero R] variable (R) (M) -theorem Nat.noZeroSMulDivisors : NoZeroSMulDivisors ℕ M := - ⟨by - intro c x - rw [nsmul_eq_smul_cast R, smul_eq_zero] - simp⟩ +theorem Nat.noZeroSMulDivisors : NoZeroSMulDivisors ℕ M where + eq_zero_or_eq_zero_of_smul_eq_zero {c x} := by rw [← Nat.cast_smul_eq_nsmul R, smul_eq_zero]; simp -- Porting note: left-hand side never simplifies when using simp on itself --@[simp] @@ -483,7 +495,7 @@ variable (R M) zero as well. Usually `M` is an `R`-algebra. -/ theorem CharZero.of_module (M) [AddCommMonoidWithOne M] [CharZero M] [Module R M] : CharZero R := by refine ⟨fun m n h => @Nat.cast_injective M _ _ _ _ ?_⟩ - rw [← nsmul_one, ← nsmul_one, nsmul_eq_smul_cast R m (1 : M), nsmul_eq_smul_cast R n (1 : M), h] + rw [← nsmul_one, ← nsmul_one, ← Nat.cast_smul_eq_nsmul R, ← Nat.cast_smul_eq_nsmul R, h] end Module @@ -549,7 +561,7 @@ theorem smul_left_injective {x : M} (hx : x ≠ 0) : Function.Injective fun c : end SMulInjective instance [NoZeroSMulDivisors ℤ M] : NoZeroSMulDivisors ℕ M := - ⟨fun {c x} hcx ↦ by rwa [nsmul_eq_smul_cast ℤ c x, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ + ⟨fun {c x} hcx ↦ by rwa [← Nat.cast_smul_eq_nsmul ℤ, smul_eq_zero, Nat.cast_eq_zero] at hcx⟩ variable (R M) @@ -561,7 +573,7 @@ scalar torsion. -/ theorem CharZero.of_noZeroSMulDivisors [Nontrivial M] [NoZeroSMulDivisors ℤ M] : CharZero R := by refine ⟨fun {n m h} ↦ ?_⟩ obtain ⟨x, hx⟩ := exists_ne (0 : M) - replace h : (n : ℤ) • x = (m : ℤ) • x := by simp [zsmul_eq_smul_cast R, h] + replace h : (n : ℤ) • x = (m : ℤ) • x := by simp [← Nat.cast_smul_eq_nsmul R, h] simpa using smul_left_injective ℤ hx h end Module diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 43a6008a35f44..89315cfde15fe 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -49,7 +49,7 @@ def restrictScalars (f : M ≃ₗ[S] M₂) : M ≃ₗ[R] M₂ := right_inv := f.right_inv } theorem restrictScalars_injective : - Function.Injective (restrictScalars R : (M ≃ₗ[S] M₂) → M ≃ₗ[R] M₂) := fun _ _ h => + Function.Injective (restrictScalars R : (M ≃ₗ[S] M₂) → M ≃ₗ[R] M₂) := fun _ _ h ↦ ext (LinearEquiv.congr_fun h : _) @[simp] @@ -61,12 +61,12 @@ end RestrictScalars theorem _root_.Module.End_isUnit_iff [Module R M] (f : Module.End R M) : IsUnit f ↔ Function.Bijective f := - ⟨fun h => + ⟨fun h ↦ Function.bijective_iff_has_inverse.mpr <| ⟨h.unit.inv, ⟨Module.End_isUnit_inv_apply_apply_of_isUnit h, Module.End_isUnit_apply_inv_apply_of_isUnit h⟩⟩, - fun H => + fun H ↦ let e : M ≃ₗ[R] M := { f, Equiv.ofBijective f H with } ⟨⟨_, e.symm, LinearMap.ext e.right_inv, LinearMap.ext e.left_inv⟩, rfl⟩⟩ @@ -79,8 +79,8 @@ instance automorphismGroup : Group (M ≃ₗ[R] M) where one := LinearEquiv.refl R M inv f := f.symm mul_assoc f g h := rfl - mul_one f := ext fun x => rfl - one_mul f := ext fun x => rfl + mul_one f := ext fun x ↦ rfl + one_mul f := ext fun x ↦ rfl mul_left_inv f := ext <| f.left_inv @[simp] @@ -91,7 +91,7 @@ lemma coe_toLinearMap_one : (↑(1 : M ≃ₗ[R] M) : M →ₗ[R] M) = LinearMap @[simp] lemma coe_toLinearMap_mul {e₁ e₂ : M ≃ₗ[R] M} : - (↑(e₁ * e₂) : M →ₗ[R] M) = (e₁ : M →ₗ[R] M) * (e₂ : M →ₗ[R] M) := by + (↑(e₁ * e₂) : M →ₗ[R] M) = (e₁ : M →ₗ[R] M) * (e₂ : M →ₗ[R] M) := rfl theorem coe_pow (e : M ≃ₗ[R] M) (n : ℕ) : ⇑(e ^ n) = e^[n] := hom_coe_pow _ rfl (fun _ _ ↦ rfl) _ _ @@ -122,7 +122,7 @@ protected theorem smul_def (f : M ≃ₗ[R] M) (a : M) : f • a = f a := /-- `LinearEquiv.applyDistribMulAction` is faithful. -/ instance apply_faithfulSMul : FaithfulSMul (M ≃ₗ[R] M) M := - ⟨@fun _ _ => LinearEquiv.ext⟩ + ⟨@fun _ _ ↦ LinearEquiv.ext⟩ instance apply_smulCommClass : SMulCommClass R (M ≃ₗ[R] M) M where smul_comm r e m := (e.map_smul r m).symm @@ -141,10 +141,10 @@ variable [Module R M] [Module R M₂] [Subsingleton M] [Subsingleton M₂] @[simps] def ofSubsingleton : M ≃ₗ[R] M₂ := { (0 : M →ₗ[R] M₂) with - toFun := fun _ => 0 - invFun := fun _ => 0 - left_inv := fun _ => Subsingleton.elim _ _ - right_inv := fun _ => Subsingleton.elim _ _ } + toFun := fun _ ↦ 0 + invFun := fun _ ↦ 0 + left_inv := fun _ ↦ Subsingleton.elim _ _ + right_inv := fun _ ↦ Subsingleton.elim _ _ } @[simp] theorem ofSubsingleton_self : ofSubsingleton M M = refl R M := by @@ -217,15 +217,14 @@ theorem coe_toLinearEquiv_symm (h : ∀ (c : R) (x), e (c • x) = c • e x) : /-- An additive equivalence between commutative additive monoids is a linear equivalence between ℕ-modules -/ def toNatLinearEquiv : M ≃ₗ[ℕ] M₂ := - e.toLinearEquiv fun c a => by rw [map_nsmul] + e.toLinearEquiv fun c a ↦ by rw [map_nsmul] @[simp] theorem coe_toNatLinearEquiv : ⇑e.toNatLinearEquiv = e := rfl @[simp] -theorem toNatLinearEquiv_toAddEquiv : ↑e.toNatLinearEquiv = e := by - ext +theorem toNatLinearEquiv_toAddEquiv : ↑e.toNatLinearEquiv = e := rfl @[simp] @@ -256,7 +255,7 @@ variable (e : M ≃+ M₂) /-- An additive equivalence between commutative additive groups is a linear equivalence between ℤ-modules -/ def toIntLinearEquiv : M ≃ₗ[ℤ] M₂ := - e.toLinearEquiv fun c a => e.toAddMonoidHom.map_zsmul a c + e.toLinearEquiv fun c a ↦ e.toAddMonoidHom.map_zsmul a c @[simp] theorem coe_toIntLinearEquiv : ⇑e.toIntLinearEquiv = e := @@ -305,12 +304,12 @@ See note [bundled maps over different rings]. @[simps] def ringLmapEquivSelf [Module S M] [SMulCommClass R S M] : (R →ₗ[R] M) ≃ₗ[S] M := { applyₗ' S (1 : R) with - toFun := fun f => f 1 + toFun := fun f ↦ f 1 invFun := smulRight (1 : R →ₗ[R] R) - left_inv := fun f => by + left_inv := fun f ↦ by ext simp only [coe_smulRight, one_apply, smul_eq_mul, ← map_smul f, mul_one] - right_inv := fun x => by simp } + right_inv := fun x ↦ by simp } end LinearMap @@ -323,10 +322,10 @@ def addMonoidHomLequivNat {A B : Type*} (R : Type*) [Semiring R] [AddCommMonoid where toFun := AddMonoidHom.toNatLinearMap invFun := LinearMap.toAddMonoidHom - map_add' := by intros; ext; rfl - map_smul' := by intros; ext; rfl - left_inv := by intro f; ext; rfl - right_inv := by intro f; ext; rfl + map_add' _ _ := rfl + map_smul' _ _ := rfl + left_inv _ := rfl + right_inv _ := rfl /-- The `R`-linear equivalence between additive morphisms `A →+ B` and `ℤ`-linear morphisms `A →ₗ[ℤ] B`. @@ -337,17 +336,17 @@ def addMonoidHomLequivInt {A B : Type*} (R : Type*) [Semiring R] [AddCommGroup A where toFun := AddMonoidHom.toIntLinearMap invFun := LinearMap.toAddMonoidHom - map_add' := by intros; ext; rfl - map_smul' := by intros; ext; rfl - left_inv := by intro f; ext; rfl - right_inv := by intro f; ext; rfl + map_add' _ _ := rfl + map_smul' _ _ := rfl + left_inv _ := rfl + right_inv _ := rfl /-- Ring equivalence between additive group endomorphisms of an `AddCommGroup` `A` and `ℤ`-module endomorphisms of `A.` -/ @[simps] def addMonoidEndRingEquivInt (A : Type*) [AddCommGroup A] : AddMonoid.End A ≃+* Module.End ℤ A := { addMonoidHomLequivInt (B := A) ℤ with - map_mul' := fun _ _ => rfl } + map_mul' := fun _ _ ↦ rfl } namespace LinearEquiv @@ -391,7 +390,6 @@ instance : Unique (M ≃ₛₗ[σ₁₂] M₂) where uniq _ := toLinearMap_injective (Subsingleton.elim _ _) default := 0 - end Module instance uniqueOfSubsingleton [Subsingleton R] [Subsingleton R₂] : Unique (M ≃ₛₗ[σ₁₂] M₂) := by @@ -411,12 +409,8 @@ variable (V V₂ R) Differs from `TensorProduct.curry`. -/ protected def curry : (V × V₂ → R) ≃ₗ[R] V → V₂ → R := { Equiv.curry _ _ _ with - map_add' := fun _ _ => by - ext - rfl - map_smul' := fun _ _ => by - ext - rfl } + map_add' := fun _ _ ↦ rfl + map_smul' := fun _ _ ↦ rfl } @[simp] theorem coe_curry : ⇑(LinearEquiv.curry R V V₂) = curry := @@ -501,7 +495,7 @@ linear isomorphism between the two function spaces. -/ 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₁) + 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 @@ -571,8 +565,7 @@ theorem conj_comp (e : M ≃ₗ[R] M₂) (f g : Module.End R M) : arrowCongr_comp e e e f g theorem conj_trans (e₁ : M ≃ₗ[R] M₂) (e₂ : M₂ ≃ₗ[R] M₃) : - e₁.conj.trans e₂.conj = (e₁.trans e₂).conj := by - ext f x + e₁.conj.trans e₂.conj = (e₁.trans e₂).conj := rfl @[simp] @@ -652,7 +645,7 @@ theorem funLeft_surjective_of_injective (f : m → n) (hf : Injective f) : Surjective (funLeft R M f) := by classical intro g - refine ⟨fun x => if h : ∃ y, f y = x then g h.choose else 0, ?_⟩ + refine ⟨fun x ↦ if h : ∃ y, f y = x then g h.choose else 0, ?_⟩ ext dsimp only [funLeft_apply] split_ifs with w @@ -677,10 +670,10 @@ open LinearMap construct a linear equivalence `(n → M) ≃ₗ[R] (m → M)` -/ def funCongrLeft (e : m ≃ n) : (n → M) ≃ₗ[R] m → M := LinearEquiv.ofLinear (funLeft R M e) (funLeft R M e.symm) - (LinearMap.ext fun x => - funext fun i => by rw [id_apply, ← funLeft_comp, Equiv.symm_comp_self, LinearMap.funLeft_id]) - (LinearMap.ext fun x => - funext fun i => by rw [id_apply, ← funLeft_comp, Equiv.self_comp_symm, LinearMap.funLeft_id]) + (LinearMap.ext fun x ↦ + funext fun i ↦ by rw [id_apply, ← funLeft_comp, Equiv.symm_comp_self, LinearMap.funLeft_id]) + (LinearMap.ext fun x ↦ + funext fun i ↦ by rw [id_apply, ← funLeft_comp, Equiv.self_comp_symm, LinearMap.funLeft_id]) @[simp] theorem funCongrLeft_apply (e : m ≃ n) (x : n → M) : funCongrLeft R M e x = funLeft R M e x := diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 06b3913b4ce10..1513cc839f617 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -147,10 +147,10 @@ instance : Coe (M ≃ₛₗ[σ] M₂) (M →ₛₗ[σ] M₂) := -- This exists for compatibility, previously `≃ₗ[R]` extended `≃` instead of `≃+`. /-- The equivalence of types underlying a linear equivalence. -/ -def toEquiv : (M ≃ₛₗ[σ] M₂) → M ≃ M₂ := fun f => f.toAddEquiv.toEquiv +def toEquiv : (M ≃ₛₗ[σ] M₂) → M ≃ M₂ := fun f ↦ f.toAddEquiv.toEquiv theorem toEquiv_injective : Function.Injective (toEquiv : (M ≃ₛₗ[σ] M₂) → M ≃ M₂) := - fun ⟨⟨⟨_, _⟩, _⟩, _, _, _⟩ ⟨⟨⟨_, _⟩, _⟩, _, _, _⟩ h => + fun ⟨⟨⟨_, _⟩, _⟩, _, _, _⟩ ⟨⟨⟨_, _⟩, _⟩, _, _, _⟩ h ↦ (LinearEquiv.mk.injEq _ _ _ _ _ _ _ _).mpr ⟨LinearMap.ext (congr_fun (Equiv.mk.inj h).1), (Equiv.mk.inj h).2⟩ @@ -159,7 +159,7 @@ theorem toEquiv_inj {e₁ e₂ : M ≃ₛₗ[σ] M₂} : e₁.toEquiv = e₂.toE toEquiv_injective.eq_iff theorem toLinearMap_injective : Injective (toLinearMap : (M ≃ₛₗ[σ] M₂) → M →ₛₗ[σ] M₂) := - fun _ _ H => toEquiv_injective <| Equiv.ext <| LinearMap.congr_fun H + fun _ _ H ↦ toEquiv_injective <| Equiv.ext <| LinearMap.congr_fun H @[simp, norm_cast] theorem toLinearMap_inj {e₁ e₂ : M ≃ₛₗ[σ] M₂} : (↑e₁ : M →ₛₗ[σ] M₂) = e₂ ↔ e₁ = e₂ := @@ -257,7 +257,7 @@ def symm (e : M ≃ₛₗ[σ] M₂) : M₂ ≃ₛₗ[σ'] M := e.toEquiv.symm with toFun := e.toLinearMap.inverse e.invFun e.left_inv e.right_inv invFun := e.toEquiv.symm.invFun - map_smul' := fun r x => by dsimp only; rw [map_smulₛₗ] } + map_smul' := fun r x ↦ by dsimp only; rw [map_smulₛₗ] } -- Porting note: this is new /-- See Note [custom simps projection] -/ @@ -428,7 +428,7 @@ theorem comp_coe [Module R M] [Module R M₂] [Module R M₃] (f : M ≃ₗ[R] M @[simp] theorem mk_coe (f h₁ h₂) : (LinearEquiv.mk e f h₁ h₂ : M ≃ₛₗ[σ] M₂) = e := - ext fun _ => rfl + ext fun _ ↦ rfl protected theorem map_add (a b : M) : e (a + b) = e a + e b := map_add e a b @@ -460,7 +460,7 @@ theorem symm_bijective [Module R M] [Module S M₂] [RingHomInvPair σ' σ] [Rin @[simp] theorem mk_coe' (f h₁ h₂ h₃ h₄) : (LinearEquiv.mk ⟨⟨f, h₁⟩, h₂⟩ (⇑e) h₃ h₄ : M₂ ≃ₛₗ[σ'] M) = e.symm := - symm_bijective.injective <| ext fun _ => rfl + symm_bijective.injective <| ext fun _ ↦ rfl /-- Auxilliary definition to avoid looping in `dsimp` with `LinearEquiv.symm_mk`. -/ protected def symm_mk.aux (f h₁ h₂ h₃ h₄) := (⟨⟨⟨e, h₁⟩, h₂⟩, f, h₃, h₄⟩ : M ≃ₛₗ[σ] M₂).symm diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index b019967a63be4..8857617afb083 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -478,21 +478,22 @@ variable [Semiring R₁] [Semiring R₂] [Semiring R₃] variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] variable {module_M₁ : Module R₁ M₁} {module_M₂ : Module R₂ M₂} {module_M₃ : Module R₃ M₃} variable {σ₁₂ : R₁ →+* R₂} {σ₂₃ : R₂ →+* R₃} {σ₁₃ : R₁ →+* R₃} -variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] -variable (f : M₂ →ₛₗ[σ₂₃] M₃) (g : M₁ →ₛₗ[σ₁₂] M₂) /-- Composition of two linear maps is a linear map -/ -def comp : M₁ →ₛₗ[σ₁₃] M₃ where +def comp [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] (f : M₂ →ₛₗ[σ₂₃] M₃) (g : M₁ →ₛₗ[σ₁₂] M₂) : + M₁ →ₛₗ[σ₁₃] M₃ where toFun := f ∘ g map_add' := by simp only [map_add, forall_const, Function.comp_apply] -- Note that #8386 changed `map_smulₛₗ` to `map_smulₛₗ _` map_smul' r x := by simp only [Function.comp_apply, map_smulₛₗ _, RingHomCompTriple.comp_apply] +variable [RingHomCompTriple σ₁₂ σ₂₃ σ₁₃] +variable (f : M₂ →ₛₗ[σ₂₃] M₃) (g : M₁ →ₛₗ[σ₁₂] M₂) + /-- `∘ₗ` is notation for composition of two linear (not semilinear!) maps into a linear map. This is useful when Lean is struggling to infer the `RingHomCompTriple` instance. -/ notation3:80 (name := compNotation) f:81 " ∘ₗ " g:80 => - @LinearMap.comp _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ (RingHom.id _) (RingHom.id _) (RingHom.id _) - RingHomCompTriple.ids f g + LinearMap.comp (σ₁₂ := RingHom.id _) (σ₂₃ := RingHom.id _) (σ₁₃ := RingHom.id _) f g theorem comp_apply (x : M₁) : f.comp g x = f (g x) := rfl diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index 884ccffe367a1..a6d7f5857710f 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -44,7 +44,7 @@ variable [Semiring R] [AddCommMonoid M] [AddCommGroup N₁] [Module R M] [Module instance : One (Module.End R M) := ⟨LinearMap.id⟩ -instance : Mul (Module.End R M) := ⟨LinearMap.comp⟩ +instance : Mul (Module.End R M) := ⟨fun f g => LinearMap.comp f g⟩ theorem one_eq_id : (1 : Module.End R M) = id := rfl @@ -364,7 +364,7 @@ variable (f g : M →ₗ[R] M₂) /-- Composition by `f : M₂ → M₃` is a linear map from the space of linear maps `M → M₂` to the space of linear maps `M → M₃`. -/ def compRight (f : M₂ →ₗ[R] M₃) : (M →ₗ[R] M₂) →ₗ[R] M →ₗ[R] M₃ where - toFun := f.comp + toFun g := f.comp g map_add' _ _ := LinearMap.ext fun _ => map_add f _ _ map_smul' _ _ := LinearMap.ext fun _ => map_smul f _ _ diff --git a/Mathlib/Algebra/Module/Pi.lean b/Mathlib/Algebra/Module/Pi.lean index bb6b0ef86a064..002ec529cf852 100644 --- a/Mathlib/Algebra/Module/Pi.lean +++ b/Mathlib/Algebra/Module/Pi.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ +import Mathlib.Algebra.GroupWithZero.Action.Pi import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Regular.SMul import Mathlib.Algebra.Ring.Pi -import Mathlib.GroupTheory.GroupAction.Pi /-! # Pi instances for modules diff --git a/Mathlib/Algebra/Module/Prod.lean b/Mathlib/Algebra/Module/Prod.lean index 6401606639813..11b934777a06a 100644 --- a/Mathlib/Algebra/Module/Prod.lean +++ b/Mathlib/Algebra/Module/Prod.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Eric Wieser -/ +import Mathlib.Algebra.GroupWithZero.Action.Prod import Mathlib.Algebra.Module.Defs -import Mathlib.GroupTheory.GroupAction.Prod /-! # Prod instances for module and multiplicative actions diff --git a/Mathlib/Algebra/Module/Rat.lean b/Mathlib/Algebra/Module/Rat.lean index 4babcfa3b11cc..8c134bd1eb1a9 100644 --- a/Mathlib/Algebra/Module/Rat.lean +++ b/Mathlib/Algebra/Module/Rat.lean @@ -60,6 +60,6 @@ instance SMulCommClass.rat' {R : Type u} {M : Type v} [Semiring R] [AddCommGroup instance (priority := 100) RatModule.noZeroSMulDivisors [AddCommGroup M] [Module ℚ M] : NoZeroSMulDivisors ℤ M := ⟨fun {k} {x : M} h => by - simpa only [zsmul_eq_smul_cast ℚ k x, smul_eq_zero, Rat.zero_iff_num_zero] using h⟩ + simpa only [← Int.cast_smul_eq_nsmul ℚ k x, smul_eq_zero, Rat.zero_iff_num_zero] using h⟩ -- Porting note: old proof was: --⟨fun {k x} h => by simpa [zsmul_eq_smul_cast ℚ k x] using h⟩ diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 6be8336dd218c..d701d5f5cc447 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -158,7 +158,7 @@ theorem map_equivMapOfInjective_symm_apply (f : F) (i : Injective f) (p : Submod i.eq_iff, LinearEquiv.apply_symm_apply] /-- The pullback of a submodule `p ⊆ M₂` along `f : M → M₂` -/ -def comap (f : F) (p : Submodule R₂ M₂) : Submodule R M := +def comap [SemilinearMapClass F σ₁₂ M M₂] (f : F) (p : Submodule R₂ M₂) : Submodule R M := { p.toAddSubmonoid.comap f with carrier := f ⁻¹' p -- Note: #8386 added `map_smulₛₗ _` @@ -247,16 +247,17 @@ theorem le_comap_map [RingHomSurjective σ₁₂] (f : F) (p : Submodule R M) : section GaloisInsertion -variable {f : F} (hf : Surjective f) -variable [RingHomSurjective σ₁₂] +variable [RingHomSurjective σ₁₂] {f : F} /-- `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. -/ -def giMapComap : GaloisInsertion (map f) (comap f) := +def giMapComap (hf : Surjective f) : GaloisInsertion (map f) (comap f) := (gc_map_comap f).toGaloisInsertion fun S x hx => by rcases hf x with ⟨y, rfl⟩ simp only [mem_map, mem_comap] exact ⟨y, hx, rfl⟩ +variable (hf : Surjective f) + theorem map_comap_eq_of_surjective (p : Submodule R₂ M₂) : (p.comap f).map f = p := (giMapComap hf).l_u_eq _ @@ -292,16 +293,18 @@ end GaloisInsertion section GaloisCoinsertion -variable [RingHomSurjective σ₁₂] {f : F} (hf : Injective f) +variable [RingHomSurjective σ₁₂] {f : F} /-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ -def gciMapComap : GaloisCoinsertion (map f) (comap f) := +def gciMapComap (hf : Injective f) : GaloisCoinsertion (map f) (comap f) := (gc_map_comap f).toGaloisCoinsertion fun S x => by simp [mem_comap, mem_map, forall_exists_index, and_imp] intro y hy hxy rw [hf.eq_iff] at hxy rwa [← hxy] +variable (hf : Injective f) + theorem comap_map_eq_of_injective (p : Submodule R M) : (p.map f).comap f = p := (gciMapComap hf).u_l_eq _ diff --git a/Mathlib/Algebra/Module/Zlattice/Basic.lean b/Mathlib/Algebra/Module/Zlattice/Basic.lean index 59d356c57d1e0..4025c1293c57a 100644 --- a/Mathlib/Algebra/Module/Zlattice/Basic.lean +++ b/Mathlib/Algebra/Module/Zlattice/Basic.lean @@ -94,13 +94,13 @@ def ceil (m : E) : span ℤ (Set.range b) := ∑ i, ⌈b.repr m i⌉ • b.restr @[simp] theorem repr_floor_apply (m : E) (i : ι) : b.repr (floor b m) i = ⌊b.repr m i⌋ := by - classical simp only [floor, zsmul_eq_smul_cast K, b.repr.map_smul, Finsupp.single_apply, + classical simp only [floor, ← Int.cast_smul_eq_nsmul K, b.repr.map_smul, Finsupp.single_apply, Finset.sum_apply', Basis.repr_self, Finsupp.smul_single', mul_one, Finset.sum_ite_eq', coe_sum, Finset.mem_univ, if_true, coe_smul_of_tower, Basis.restrictScalars_apply, map_sum] @[simp] theorem repr_ceil_apply (m : E) (i : ι) : b.repr (ceil b m) i = ⌈b.repr m i⌉ := by - classical simp only [ceil, zsmul_eq_smul_cast K, b.repr.map_smul, Finsupp.single_apply, + classical simp only [ceil, ← Int.cast_smul_eq_nsmul K, b.repr.map_smul, Finsupp.single_apply, Finset.sum_apply', Basis.repr_self, Finsupp.smul_single', mul_one, Finset.sum_ite_eq', coe_sum, Finset.mem_univ, if_true, coe_smul_of_tower, Basis.restrictScalars_apply, map_sum] @@ -542,8 +542,8 @@ theorem Zlattice.rank [hs : IsZlattice K L] : finrank ℤ L = finrank K E := by rwa [Ne, add_eq_zero_iff_eq_neg.not, neg_inj, Rat.coe_int_inj, ← Ne] apply (smul_mem_iff _ h_nz).mp refine span_subset_span ℤ ℚ _ ?_ - rwa [add_smul, neg_smul, SetLike.mem_coe, ← Zspan.fract_eq_fract, ← zsmul_eq_smul_cast ℚ, - ← zsmul_eq_smul_cast ℚ] + rwa [add_smul, neg_smul, SetLike.mem_coe, ← Zspan.fract_eq_fract, Int.cast_smul_eq_nsmul ℚ, + Int.cast_smul_eq_nsmul ℚ] · -- To prove that `finrank K E ≤ finrank ℤ L`, we use the fact `b` generates `E` over `K` -- and thus `finrank K E ≤ card b = finrank ℤ L` rw [← topEquiv.finrank_eq, ← h_spanE] diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index a9f2125a76945..3176c280a11c6 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -948,10 +948,11 @@ variable [Monoid G] [CommSemiring k] {V : Type u₃} {W : Type u₄} [AddCommMon [Module (MonoidAlgebra k G) V] [IsScalarTower k (MonoidAlgebra k G) V] [AddCommMonoid W] [Module k W] [Module (MonoidAlgebra k G) W] [IsScalarTower k (MonoidAlgebra k G) W] (f : V →ₗ[k] W) - (h : ∀ (g : G) (v : V), f (single g (1 : k) • v) = single g (1 : k) • f v) /-- Build a `k[G]`-linear map from a `k`-linear map and evidence that it is `G`-equivariant. -/ -def equivariantOfLinearOfComm : V →ₗ[MonoidAlgebra k G] W where +def equivariantOfLinearOfComm + (h : ∀ (g : G) (v : V), f (single g (1 : k) • v) = single g (1 : k) • f v) : + V →ₗ[MonoidAlgebra k G] W where toFun := f map_add' v v' := by simp map_smul' c v := by @@ -964,6 +965,8 @@ def equivariantOfLinearOfComm : V →ₗ[MonoidAlgebra k G] W where erw [algebraMap_smul (MonoidAlgebra k G) r, algebraMap_smul (MonoidAlgebra k G) r, f.map_smul, h g v, of_apply] +variable (h : ∀ (g : G) (v : V), f (single g (1 : k) • v) = single g (1 : k) • f v) + @[simp] theorem equivariantOfLinearOfComm_apply (v : V) : (equivariantOfLinearOfComm f h) v = f v := rfl diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 5cc5f0da71706..a47dbbb4d9829 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -435,6 +435,10 @@ theorem totalDegree_finset_sum {ι : Type*} (s : Finset ι) (f : ι → MvPolyno · rw [Finset.sum_cons, Finset.sup_cons, sup_eq_max] exact (MvPolynomial.totalDegree_add _ _).trans (max_le_max le_rfl hind) +lemma totalDegree_finsetSum_le {ι : Type*} {s : Finset ι} {f : ι → MvPolynomial σ R} {d : ℕ} + (hf : ∀ i ∈ s, (f i).totalDegree ≤ d) : (s.sum f).totalDegree ≤ d := + (totalDegree_finset_sum ..).trans $ Finset.sup_le hf + lemma degreeOf_le_totalDegree (f : MvPolynomial σ R) (i : σ) : f.degreeOf i ≤ f.totalDegree := degreeOf_le_iff.mpr fun d hd ↦ (eq_or_ne (d i) 0).elim (·.trans_le zero_le') fun h ↦ (Finset.single_le_sum (fun _ _ ↦ zero_le') <| Finsupp.mem_support_iff.mpr h).trans diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index 51c9e9a3b1d46..8cc4c25cbbcbf 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -244,4 +244,19 @@ lemma map_nsmul_piAntidiag_univ [Fintype ι] (m : ℕ) {n : ℕ} (hn : n ≠ 0) simpa using map_nsmul_piAntidiag (univ : Finset ι) m hn end Nat + +lemma map_sym_eq_piAntidiag [DecidableEq ι] (s : Finset ι) (n : ℕ) : + (s.sym n).map ⟨fun m a ↦ m.1.count a, Multiset.count_injective.comp Sym.coe_injective⟩ = + piAntidiag s n := by + ext f + simp only [Sym.val_eq_coe, mem_map, mem_sym_iff, Embedding.coeFn_mk, funext_iff, Sym.exists, + Sym.mem_mk, Sym.coe_mk, exists_and_left, exists_prop, mem_piAntidiag, ne_eq] + constructor + · rintro ⟨m, hm, rfl, hf⟩ + simpa [← hf, Multiset.sum_count_eq_card hm] + · rintro ⟨rfl, hf⟩ + refine ⟨∑ a ∈ s, f a • {a}, ?_, ?_⟩ + · simp (config := { contextual := true }) + · simpa [Multiset.count_sum', Multiset.count_singleton, not_imp_comm, eq_comm (a := 0)] using hf + end Finset diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 920c30bdab385..7872ed429ca63 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -172,7 +172,7 @@ variable [CanonicallyOrderedCommMonoid M] {l : List M} cases' lt_or_le n L.length with h h · rw [prod_take_succ _ _ h] exact le_self_mul - · simp [take_all_of_le h, take_all_of_le (le_trans h (Nat.le_succ _))] + · simp [take_of_length_le h, take_of_length_le (le_trans h (Nat.le_succ _))] end CanonicallyOrderedCommMonoid end List diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 51add01a7229f..f65506259576e 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -75,11 +75,11 @@ lemma prod_lt_prod (hf : ∀ i ∈ s, 0 < f i) (hfg : ∀ i ∈ s, f i ≤ g i) obtain ⟨i, hi, hilt⟩ := hlt rw [← insert_erase hi, prod_insert (not_mem_erase _ _), prod_insert (not_mem_erase _ _)] have := posMulStrictMono_iff_mulPosStrictMono.1 ‹PosMulStrictMono R› - refine mul_lt_mul_of_le_of_lt' hilt ?_ ?_ ?_ + refine mul_lt_mul_of_pos_of_nonneg' hilt ?_ ?_ ?_ · exact prod_le_prod (fun j hj => le_of_lt (hf j (mem_of_mem_erase hj))) (fun _ hj ↦ hfg _ <| mem_of_mem_erase hj) - · exact (hf i hi).le.trans hilt.le · exact prod_pos fun j hj => hf j (mem_of_mem_erase hj) + · exact (hf i hi).le.trans hilt.le lemma prod_lt_prod_of_nonempty (hf : ∀ i ∈ s, 0 < f i) (hfg : ∀ i ∈ s, f i < g i) (h_ne : s.Nonempty) : diff --git a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean index 7a0c6d5df04ee..1e545d67b0704 100644 --- a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean +++ b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean @@ -172,7 +172,7 @@ lemma of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ lemma of_mono_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) (hnm : ∀ n ≥ m, f n ≤ f n.succ) : IsCauSeq abs f := - (of_decreasing_bounded _ (by simpa using ham) $ by simpa using hnm).of_neg + (of_decreasing_bounded (-f) (a := a) (m := m) (by simpa using ham) $ by simpa using hnm).of_neg lemma geo_series [Nontrivial β] (x : β) (hx1 : abv x < 1) : IsCauSeq abv fun n ↦ ∑ m ∈ range n, x ^ m := by diff --git a/Mathlib/Algebra/Order/CauSeq/Completion.lean b/Mathlib/Algebra/Order/CauSeq/Completion.lean index 921fcf3294a3d..6a8ebe09f7d4b 100644 --- a/Mathlib/Algebra/Order/CauSeq/Completion.lean +++ b/Mathlib/Algebra/Order/CauSeq/Completion.lean @@ -133,10 +133,10 @@ theorem ofRat_mul (x y : β) : ofRat (x * y) = (ofRat x * ofRat y : Cauchy abv) := congr_arg mk (const_mul _ _) -private theorem zero_def : 0 = @mk _ _ _ _ abv _ 0 := +private theorem zero_def : 0 = mk (abv := abv) 0 := rfl -private theorem one_def : 1 = @mk _ _ _ _ abv _ 1 := +private theorem one_def : 1 = mk (abv := abv) 1 := rfl instance Cauchy.ring : Ring (Cauchy abv) := @@ -205,12 +205,12 @@ theorem inv_zero : (0 : (Cauchy abv))⁻¹ = 0 := congr_arg mk <| by rw [dif_pos] <;> [rfl; exact zero_limZero] @[simp] -theorem inv_mk {f} (hf) : (@mk α _ β _ abv _ f)⁻¹ = mk (inv f hf) := +theorem inv_mk {f} (hf) : (mk (abv := abv) f)⁻¹ = mk (inv f hf) := congr_arg mk <| by rw [dif_neg] theorem cau_seq_zero_ne_one : ¬(0 : CauSeq _ abv) ≈ 1 := fun h => have : LimZero (1 - 0 : CauSeq _ abv) := Setoid.symm h - have : LimZero 1 := by simpa + have : LimZero (1 : CauSeq _ abv) := by simpa by apply one_ne_zero <| const_limZero.1 this theorem zero_ne_one : (0 : (Cauchy abv)) ≠ 1 := fun h => cau_seq_zero_ne_one <| mk_eq.1 h diff --git a/Mathlib/Algebra/Order/Floor/Div.lean b/Mathlib/Algebra/Order/Floor/Div.lean index 72640c3c63593..33bc2b55c54a5 100644 --- a/Mathlib/Algebra/Order/Floor/Div.lean +++ b/Mathlib/Algebra/Order/Floor/Div.lean @@ -3,10 +3,10 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.GroupWithZero.Action.Pi import Mathlib.Algebra.Order.Module.Defs import Mathlib.Algebra.Order.Pi import Mathlib.Data.Finsupp.Order -import Mathlib.GroupTheory.GroupAction.Pi import Mathlib.Order.GaloisConnection /-! diff --git a/Mathlib/Algebra/Order/Group/Abs.lean b/Mathlib/Algebra/Order/Group/Abs.lean index 3f5109e906231..3d6763c037cfa 100644 --- a/Mathlib/Algebra/Order/Group/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Abs.lean @@ -3,9 +3,8 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ -import Mathlib.Algebra.Group.Even -import Mathlib.Algebra.Order.Group.Lattice import Mathlib.Algebra.Order.Group.Defs +import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Algebra.Order.Monoid.Unbundled.Pow /-! @@ -24,242 +23,6 @@ open Function variable {α : Type*} -section Lattice -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⁻¹ - -@[inherit_doc mabs] -macro:max atomic("|" noWs) a:term noWs "|ₘ" : term => `(mabs $a) - -@[inherit_doc abs] -macro:max atomic("|" noWs) a:term noWs "|" : term => `(abs $a) - -/-- Unexpander for the notation `|a|ₘ` for `mabs a`. -Tries to add discretionary parentheses in unparseable cases. -/ -@[app_unexpander abs] -def mabs.unexpander : Lean.PrettyPrinter.Unexpander - | `($_ $a) => - match a with - | `(|$_|ₘ) | `(-$_) => `(|($a)|ₘ) - | _ => `(|$a|ₘ) - | _ => throw () - -/-- Unexpander for the notation `|a|` for `abs a`. -Tries to add discretionary parentheses in unparseable cases. -/ -@[app_unexpander abs] -def abs.unexpander : Lean.PrettyPrinter.Unexpander - | `($_ $a) => - match a with - | `(|$_|) | `(-$_) => `(|($a)|) - | _ => `(|$a|) - | _ => throw () - -@[to_additive] lemma mabs_le' : |a|ₘ ≤ b ↔ a ≤ b ∧ a⁻¹ ≤ b := sup_le_iff - -@[to_additive] lemma le_mabs_self (a : α) : a ≤ |a|ₘ := le_sup_left - -@[to_additive] lemma inv_le_mabs (a : α) : a⁻¹ ≤ |a|ₘ := le_sup_right - -@[to_additive] lemma mabs_le_mabs (h₀ : a ≤ b) (h₁ : a⁻¹ ≤ b) : |a|ₘ ≤ |b|ₘ := - (mabs_le'.2 ⟨h₀, h₁⟩).trans (le_mabs_self b) - -@[to_additive (attr := simp)] lemma mabs_inv (a : α) : |a⁻¹|ₘ = |a|ₘ := by simp [mabs, sup_comm] - -@[to_additive] lemma mabs_div_comm (a b : α) : |a / b|ₘ = |b / a|ₘ := by rw [← mabs_inv, inv_div] - -variable [CovariantClass α α (· * ·) (· ≤ ·)] - -@[to_additive] lemma mabs_of_one_le (h : 1 ≤ a) : |a|ₘ = a := - sup_eq_left.2 <| (inv_le_one'.2 h).trans h - -@[to_additive] lemma mabs_of_one_lt (h : 1 < a) : |a|ₘ = a := mabs_of_one_le h.le - -@[to_additive] lemma mabs_of_le_one (h : a ≤ 1) : |a|ₘ = a⁻¹ := - sup_eq_right.2 <| h.trans (one_le_inv'.2 h) - -@[to_additive] lemma mabs_of_lt_one (h : a < 1) : |a|ₘ = a⁻¹ := mabs_of_le_one h.le - -@[to_additive] lemma mabs_le_mabs_of_one_le (ha : 1 ≤ a) (hab : a ≤ b) : |a|ₘ ≤ |b|ₘ := by - rwa [mabs_of_one_le ha, mabs_of_one_le (ha.trans hab)] - -attribute [gcongr] abs_le_abs_of_nonneg - -@[to_additive (attr := simp)] lemma mabs_one : |(1 : α)|ₘ = 1 := mabs_of_one_le le_rfl - -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - -@[to_additive (attr := simp) abs_nonneg] lemma one_le_mabs (a : α) : 1 ≤ |a|ₘ := by - apply pow_two_semiclosed _ - rw [mabs, pow_two, mul_sup, sup_mul, ← pow_two, mul_left_inv, sup_comm, ← sup_assoc] - apply le_sup_right - -@[to_additive (attr := simp)] lemma mabs_mabs (a : α) : |(|a|ₘ)|ₘ = |a|ₘ := - mabs_of_one_le <| one_le_mabs a - -end Group - -section CommGroup -variable [CommGroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} - --- Banasiak Proposition 2.12, Zaanen 2nd lecture -/-- The absolute value satisfies the triangle inequality. -/ -@[to_additive "The absolute value satisfies the triangle inequality."] -lemma mabs_mul_le (a b : α) : |a * b|ₘ ≤ |a|ₘ * |b|ₘ := by - apply sup_le - · exact mul_le_mul' (le_mabs_self a) (le_mabs_self b) - · rw [mul_inv] - exact mul_le_mul' (inv_le_mabs _) (inv_le_mabs _) - -@[to_additive] -lemma mabs_mabs_div_mabs_le (a b : α) : |(|a|ₘ / |b|ₘ)|ₘ ≤ |a / b|ₘ := by - rw [mabs, sup_le_iff] - constructor - · apply div_le_iff_le_mul.2 - convert mabs_mul_le (a / b) b - rw [div_mul_cancel] - · rw [div_eq_mul_inv, mul_inv_rev, inv_inv, mul_inv_le_iff_le_mul, mabs_div_comm] - convert mabs_mul_le (b / a) a - · rw [div_mul_cancel] - -@[to_additive] lemma sup_div_inf_eq_mabs_div (a b : α) : (a ⊔ b) / (a ⊓ b) = |b / a|ₘ := by - simp_rw [sup_div, div_inf, div_self', sup_comm, sup_sup_sup_comm, sup_idem] - rw [← inv_div, sup_comm (b := _ / _), ← mabs, sup_eq_left] - exact one_le_mabs _ - -@[to_additive two_nsmul_sup_eq_add_add_abs_sub] -lemma sup_sq_eq_mul_mul_mabs_div (a b : α) : (a ⊔ b) ^ 2 = a * b * |b / a|ₘ := by - rw [← inf_mul_sup a b, ← sup_div_inf_eq_mabs_div, div_eq_mul_inv, ← mul_assoc, mul_comm, - mul_assoc, ← pow_two, inv_mul_cancel_left] - -@[to_additive two_nsmul_inf_eq_add_sub_abs_sub] -lemma inf_sq_eq_mul_div_mabs_div (a b : α) : (a ⊓ b) ^ 2 = a * b / |b / a|ₘ := by - rw [← inf_mul_sup a b, ← sup_div_inf_eq_mabs_div, div_eq_mul_inv, div_eq_mul_inv, mul_inv_rev, - inv_inv, mul_assoc, mul_inv_cancel_comm_assoc, ← pow_two] - --- See, e.g. Zaanen, Lectures on Riesz Spaces --- 3rd lecture -@[to_additive] -lemma mabs_div_sup_mul_mabs_div_inf (a b c : α) : - |(a ⊔ c) / (b ⊔ c)|ₘ * |(a ⊓ c) / (b ⊓ c)|ₘ = |a / b|ₘ := by - letI : DistribLattice α := CommGroup.toDistribLattice α - calc - |(a ⊔ c) / (b ⊔ c)|ₘ * |(a ⊓ c) / (b ⊓ c)|ₘ = - (b ⊔ c ⊔ (a ⊔ c)) / ((b ⊔ c) ⊓ (a ⊔ c)) * |(a ⊓ c) / (b ⊓ c)|ₘ := by - rw [sup_div_inf_eq_mabs_div] - _ = (b ⊔ c ⊔ (a ⊔ c)) / ((b ⊔ c) ⊓ (a ⊔ c)) * ((b ⊓ c ⊔ a ⊓ c) / (b ⊓ c ⊓ (a ⊓ c))) := by - rw [sup_div_inf_eq_mabs_div (b ⊓ c) (a ⊓ c)] - _ = (b ⊔ a ⊔ c) / (b ⊓ a ⊔ c) * (((b ⊔ a) ⊓ c) / (b ⊓ a ⊓ c)) := by - rw [← sup_inf_right, ← inf_sup_right, sup_assoc, sup_comm c (a ⊔ c), sup_right_idem, - sup_assoc, inf_assoc, inf_comm c (a ⊓ c), inf_right_idem, inf_assoc] - _ = (b ⊔ a ⊔ c) * ((b ⊔ a) ⊓ c) / ((b ⊓ a ⊔ c) * (b ⊓ a ⊓ c)) := by rw [div_mul_div_comm] - _ = (b ⊔ a) * c / ((b ⊓ a) * c) := by - rw [mul_comm, inf_mul_sup, mul_comm (b ⊓ a ⊔ c), inf_mul_sup] - _ = (b ⊔ a) / (b ⊓ a) := by - rw [div_eq_mul_inv, mul_inv_rev, mul_assoc, mul_inv_cancel_left, ← div_eq_mul_inv] - _ = |a / b|ₘ := by rw [sup_div_inf_eq_mabs_div] - -@[to_additive] lemma mabs_sup_div_sup_le_mabs (a b c : α) : |(a ⊔ c) / (b ⊔ c)|ₘ ≤ |a / b|ₘ := by - apply le_of_mul_le_of_one_le_left _ (one_le_mabs _); rw [mabs_div_sup_mul_mabs_div_inf] - -@[to_additive] lemma mabs_inf_div_inf_le_mabs (a b c : α) : |(a ⊓ c) / (b ⊓ c)|ₘ ≤ |a / b|ₘ := by - apply le_of_mul_le_of_one_le_right _ (one_le_mabs _); rw [mabs_div_sup_mul_mabs_div_inf] - --- Commutative case, Zaanen, 3rd lecture --- For the non-commutative case, see Birkhoff Theorem 19 (27) -@[to_additive Birkhoff_inequalities] -lemma m_Birkhoff_inequalities (a b c : α) : - |(a ⊔ c) / (b ⊔ c)|ₘ ⊔ |(a ⊓ c) / (b ⊓ c)|ₘ ≤ |a / b|ₘ := - sup_le (mabs_sup_div_sup_le_mabs a b c) (mabs_inf_div_inf_le_mabs a b c) - -end CommGroup -end Lattice - -section LinearOrder -variable [Group α] [LinearOrder α] {a b : α} - -@[to_additive] lemma mabs_choice (x : α) : |x|ₘ = x ∨ |x|ₘ = x⁻¹ := max_choice _ _ - -@[to_additive] lemma le_mabs : a ≤ |b|ₘ ↔ a ≤ b ∨ a ≤ b⁻¹ := le_max_iff - -@[to_additive] lemma mabs_eq_max_inv : |a|ₘ = max a a⁻¹ := rfl - -@[to_additive] lemma lt_mabs : a < |b|ₘ ↔ a < b ∨ a < b⁻¹ := lt_max_iff - -@[to_additive] lemma mabs_by_cases (P : α → Prop) (h1 : P a) (h2 : P a⁻¹) : P |a|ₘ := - sup_ind _ _ h1 h2 - -@[to_additive] lemma eq_or_eq_inv_of_mabs_eq (h : |a|ₘ = b) : a = b ∨ a = b⁻¹ := by - simpa only [← h, eq_comm (a := |a|ₘ), inv_eq_iff_eq_inv] using mabs_choice a - -@[to_additive] lemma mabs_eq_mabs : |a|ₘ = |b|ₘ ↔ a = b ∨ a = b⁻¹ := by - refine ⟨fun h ↦ ?_, by rintro (h | h) <;> simp [h, abs_neg]⟩ - obtain rfl | rfl := eq_or_eq_inv_of_mabs_eq h <;> - simpa only [inv_eq_iff_eq_inv (a := |b|ₘ), inv_inv, inv_inj, or_comm] using mabs_choice b - -@[to_additive] lemma isSquare_mabs : IsSquare |a|ₘ ↔ IsSquare a := - mabs_by_cases (IsSquare · ↔ _) Iff.rfl isSquare_inv - -variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} - -@[to_additive (attr := simp) abs_pos] lemma one_lt_mabs : 1 < |a|ₘ ↔ a ≠ 1 := by - obtain ha | rfl | ha := lt_trichotomy a 1 - · simp [mabs_of_lt_one ha, neg_pos, ha.ne, ha] - · simp - · simp [mabs_of_one_lt ha, ha, ha.ne'] - -@[to_additive abs_pos_of_pos] lemma one_lt_mabs_pos_of_one_lt (h : 1 < a) : 1 < |a|ₘ := - one_lt_mabs.2 h.ne' - -@[to_additive abs_pos_of_neg] lemma one_lt_mabs_of_lt_one (h : a < 1) : 1 < |a|ₘ := - one_lt_mabs.2 h.ne - -@[to_additive] lemma inv_mabs_le (a : α) : |a|ₘ⁻¹ ≤ a := by - obtain h | h := le_total 1 a - · simpa [mabs_of_one_le h] using (inv_le_one'.2 h).trans h - · simp [mabs_of_le_one h] - -@[to_additive add_abs_nonneg] lemma one_le_mul_mabs (a : α) : 1 ≤ a * |a|ₘ := by - rw [← mul_right_inv a]; exact mul_le_mul_left' (inv_le_mabs a) _ - -@[to_additive] lemma inv_mabs_le_inv (a : α) : |a|ₘ⁻¹ ≤ a⁻¹ := by simpa using inv_mabs_le a⁻¹ - -variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] - -@[to_additive] lemma mabs_ne_one : |a|ₘ ≠ 1 ↔ a ≠ 1 := - (one_le_mabs a).gt_iff_ne.symm.trans one_lt_mabs - -@[to_additive (attr := simp)] lemma mabs_eq_one : |a|ₘ = 1 ↔ a = 1 := not_iff_not.1 mabs_ne_one - -@[to_additive (attr := simp) abs_nonpos_iff] lemma mabs_le_one : |a|ₘ ≤ 1 ↔ a = 1 := - (one_le_mabs a).le_iff_eq.trans mabs_eq_one - -@[to_additive] lemma mabs_le_mabs_of_le_one (ha : a ≤ 1) (hab : b ≤ a) : |a|ₘ ≤ |b|ₘ := by - rw [mabs_of_le_one ha, mabs_of_le_one (hab.trans ha)]; exact inv_le_inv_iff.mpr hab - -@[to_additive] lemma mabs_lt : |a|ₘ < b ↔ b⁻¹ < a ∧ a < b := - max_lt_iff.trans <| and_comm.trans <| by rw [inv_lt'] - -@[to_additive] lemma inv_lt_of_mabs_lt (h : |a|ₘ < b) : b⁻¹ < a := (mabs_lt.mp h).1 - -@[to_additive] lemma lt_of_mabs_lt : |a|ₘ < b → a < b := (le_mabs_self _).trans_lt - -@[to_additive] lemma max_div_min_eq_mabs' (a b : α) : max a b / min a b = |a / b|ₘ := by - rcases le_total a b with ab | ba - · rw [max_eq_right ab, min_eq_left ab, mabs_of_le_one, inv_div] - rwa [div_le_one'] - · rw [max_eq_left ba, min_eq_right ba, mabs_of_one_le] - rwa [one_le_div'] - -@[to_additive] lemma max_div_min_eq_mabs (a b : α) : max a b / min a b = |b / a|ₘ := by - rw [mabs_div_comm, max_div_min_eq_mabs'] - -end LinearOrder - section LinearOrderedCommGroup variable [LinearOrderedCommGroup α] {a b : α} @@ -454,33 +217,3 @@ theorem max_zero_add_max_neg_zero_eq_abs_self (a : α) : max a 0 + max (-a) 0 = rcases le_total 0 a with (ha | ha) <;> simp [ha] end LinearOrderedAddCommGroup - -namespace LatticeOrderedAddCommGroup -variable [Lattice α] [AddCommGroup α] {s t : Set α} - -/-- A set `s` in a lattice ordered group is *solid* if for all `x ∈ s` and all `y ∈ α` such that -`|y| ≤ |x|`, then `y ∈ s`. -/ -def IsSolid (s : Set α) : Prop := ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, |y| ≤ |x| → y ∈ s - -/-- The solid closure of a subset `s` is the smallest superset of `s` that is solid. -/ -def solidClosure (s : Set α) : Set α := {y | ∃ x ∈ s, |y| ≤ |x|} - -lemma isSolid_solidClosure (s : Set α) : IsSolid (solidClosure s) := - fun _ ⟨y, hy, hxy⟩ _ hzx ↦ ⟨y, hy, hzx.trans hxy⟩ - -lemma solidClosure_min (hst : s ⊆ t) (ht : IsSolid t) : solidClosure s ⊆ t := - fun _ ⟨_, hy, hxy⟩ ↦ ht (hst hy) hxy - -end LatticeOrderedAddCommGroup - -namespace Pi -variable {ι : Type*} {α : ι → Type*} [∀ i, AddGroup (α i)] [∀ i, Lattice (α i)] - -@[simp] lemma abs_apply (f : ∀ i, α i) (i : ι) : |f| i = |f i| := rfl - -lemma abs_def (f : ∀ i, α i) : |f| = fun i ↦ |f i| := rfl - -end Pi - -@[deprecated (since := "2024-01-13")] alias neg_le_abs_self := neg_le_abs -@[deprecated (since := "2024-01-13")] alias neg_abs_le_self := neg_abs_le diff --git a/Mathlib/Algebra/Order/Group/Indicator.lean b/Mathlib/Algebra/Order/Group/Indicator.lean index c2841b39aa391..85070045eaeab 100644 --- a/Mathlib/Algebra/Order/Group/Indicator.lean +++ b/Mathlib/Algebra/Order/Group/Indicator.lean @@ -4,9 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Indicator -import Mathlib.Algebra.Order.Group.Abs -import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Algebra.Order.Group.Defs +import Mathlib.Algebra.Order.Group.Synonym +import Mathlib.Algebra.Order.Group.Unbundled.Abs +import Mathlib.Algebra.Order.Monoid.Canonical.Defs /-! # Support of a function in an order diff --git a/Mathlib/Algebra/Order/Group/OrderIso.lean b/Mathlib/Algebra/Order/Group/OrderIso.lean index e1bbf732ac8c9..f4cd51d207bb2 100644 --- a/Mathlib/Algebra/Order/Group/OrderIso.lean +++ b/Mathlib/Algebra/Order/Group/OrderIso.lean @@ -35,7 +35,7 @@ variable (α) @[to_additive (attr := simps!) "`x ↦ -x` as an order-reversing equivalence."] def OrderIso.inv : α ≃o αᵒᵈ where toEquiv := (Equiv.inv α).trans OrderDual.toDual - map_rel_iff' {_ _} := @inv_le_inv_iff α _ _ _ _ _ _ + map_rel_iff' {_ _} := inv_le_inv_iff (α := α) end @@ -55,7 +55,7 @@ theorem le_inv' : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := @[to_additive (attr := simps!) "`x ↦ a - x` as an order-reversing equivalence."] def OrderIso.divLeft (a : α) : α ≃o αᵒᵈ where toEquiv := (Equiv.divLeft a).trans OrderDual.toDual - map_rel_iff' {_ _} := @div_le_div_iff_left α _ _ _ _ _ _ _ + map_rel_iff' {_ _} := div_le_div_iff_left (α := α) _ end TypeclassesLeftRightLE diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index 29c63ff9e8a06..f820ae23444e5 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -3,8 +3,7 @@ Copyright (c) 2021 Christopher Hoskin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin, Yaël Dillies -/ -import Mathlib.Algebra.Order.Group.Abs -import Mathlib.Algebra.Order.Group.Lattice +import Mathlib.Algebra.Order.Group.Unbundled.Abs /-! # Positive & negative parts diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean new file mode 100644 index 0000000000000..d39f570ebf6a2 --- /dev/null +++ b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean @@ -0,0 +1,290 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Order.Group.Lattice +-- import Mathlib.Algebra.Order.Group.Defs + +/-! +# Absolute values in ordered groups + +The absolute value of an element in a group which is also a lattice is its supremum with its +negation. This generalizes the usual absolute value on real numbers (`|x| = max x (-x)`). + +## Notations + +- `|a|`: The *absolute value* of an element `a` of an additive lattice ordered group +- `|a|ₘ`: The *absolute value* of an element `a` of a multiplicative lattice ordered group +-/ + +open Function + +variable {α : Type*} + +section Lattice +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⁻¹ + +@[inherit_doc mabs] +macro:max atomic("|" noWs) a:term noWs "|ₘ" : term => `(mabs $a) + +@[inherit_doc abs] +macro:max atomic("|" noWs) a:term noWs "|" : term => `(abs $a) + +/-- Unexpander for the notation `|a|ₘ` for `mabs a`. +Tries to add discretionary parentheses in unparseable cases. -/ +@[app_unexpander abs] +def mabs.unexpander : Lean.PrettyPrinter.Unexpander + | `($_ $a) => + match a with + | `(|$_|ₘ) | `(-$_) => `(|($a)|ₘ) + | _ => `(|$a|ₘ) + | _ => throw () + +/-- Unexpander for the notation `|a|` for `abs a`. +Tries to add discretionary parentheses in unparseable cases. -/ +@[app_unexpander abs] +def abs.unexpander : Lean.PrettyPrinter.Unexpander + | `($_ $a) => + match a with + | `(|$_|) | `(-$_) => `(|($a)|) + | _ => `(|$a|) + | _ => throw () + +@[to_additive] lemma mabs_le' : |a|ₘ ≤ b ↔ a ≤ b ∧ a⁻¹ ≤ b := sup_le_iff + +@[to_additive] lemma le_mabs_self (a : α) : a ≤ |a|ₘ := le_sup_left + +@[to_additive] lemma inv_le_mabs (a : α) : a⁻¹ ≤ |a|ₘ := le_sup_right + +@[to_additive] lemma mabs_le_mabs (h₀ : a ≤ b) (h₁ : a⁻¹ ≤ b) : |a|ₘ ≤ |b|ₘ := + (mabs_le'.2 ⟨h₀, h₁⟩).trans (le_mabs_self b) + +@[to_additive (attr := simp)] lemma mabs_inv (a : α) : |a⁻¹|ₘ = |a|ₘ := by simp [mabs, sup_comm] + +@[to_additive] lemma mabs_div_comm (a b : α) : |a / b|ₘ = |b / a|ₘ := by rw [← mabs_inv, inv_div] + +variable [CovariantClass α α (· * ·) (· ≤ ·)] + +@[to_additive] lemma mabs_of_one_le (h : 1 ≤ a) : |a|ₘ = a := + sup_eq_left.2 <| (inv_le_one'.2 h).trans h + +@[to_additive] lemma mabs_of_one_lt (h : 1 < a) : |a|ₘ = a := mabs_of_one_le h.le + +@[to_additive] lemma mabs_of_le_one (h : a ≤ 1) : |a|ₘ = a⁻¹ := + sup_eq_right.2 <| h.trans (one_le_inv'.2 h) + +@[to_additive] lemma mabs_of_lt_one (h : a < 1) : |a|ₘ = a⁻¹ := mabs_of_le_one h.le + +@[to_additive] lemma mabs_le_mabs_of_one_le (ha : 1 ≤ a) (hab : a ≤ b) : |a|ₘ ≤ |b|ₘ := by + rwa [mabs_of_one_le ha, mabs_of_one_le (ha.trans hab)] + +attribute [gcongr] abs_le_abs_of_nonneg + +@[to_additive (attr := simp)] lemma mabs_one : |(1 : α)|ₘ = 1 := mabs_of_one_le le_rfl + +variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] + +@[to_additive (attr := simp) abs_nonneg] lemma one_le_mabs (a : α) : 1 ≤ |a|ₘ := by + apply pow_two_semiclosed _ + rw [mabs, pow_two, mul_sup, sup_mul, ← pow_two, mul_left_inv, sup_comm, ← sup_assoc] + apply le_sup_right + +@[to_additive (attr := simp)] lemma mabs_mabs (a : α) : |(|a|ₘ)|ₘ = |a|ₘ := + mabs_of_one_le <| one_le_mabs a + +end Group + +section CommGroup +variable [CommGroup α] [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} + +-- Banasiak Proposition 2.12, Zaanen 2nd lecture +/-- The absolute value satisfies the triangle inequality. -/ +@[to_additive "The absolute value satisfies the triangle inequality."] +lemma mabs_mul_le (a b : α) : |a * b|ₘ ≤ |a|ₘ * |b|ₘ := by + apply sup_le + · exact mul_le_mul' (le_mabs_self a) (le_mabs_self b) + · rw [mul_inv] + exact mul_le_mul' (inv_le_mabs _) (inv_le_mabs _) + +@[to_additive] +lemma mabs_mabs_div_mabs_le (a b : α) : |(|a|ₘ / |b|ₘ)|ₘ ≤ |a / b|ₘ := by + rw [mabs, sup_le_iff] + constructor + · apply div_le_iff_le_mul.2 + convert mabs_mul_le (a / b) b + rw [div_mul_cancel] + · rw [div_eq_mul_inv, mul_inv_rev, inv_inv, mul_inv_le_iff_le_mul, mabs_div_comm] + convert mabs_mul_le (b / a) a + · rw [div_mul_cancel] + +@[to_additive] lemma sup_div_inf_eq_mabs_div (a b : α) : (a ⊔ b) / (a ⊓ b) = |b / a|ₘ := by + simp_rw [sup_div, div_inf, div_self', sup_comm, sup_sup_sup_comm, sup_idem] + rw [← inv_div, sup_comm (b := _ / _), ← mabs, sup_eq_left] + exact one_le_mabs _ + +@[to_additive two_nsmul_sup_eq_add_add_abs_sub] +lemma sup_sq_eq_mul_mul_mabs_div (a b : α) : (a ⊔ b) ^ 2 = a * b * |b / a|ₘ := by + rw [← inf_mul_sup a b, ← sup_div_inf_eq_mabs_div, div_eq_mul_inv, ← mul_assoc, mul_comm, + mul_assoc, ← pow_two, inv_mul_cancel_left] + +@[to_additive two_nsmul_inf_eq_add_sub_abs_sub] +lemma inf_sq_eq_mul_div_mabs_div (a b : α) : (a ⊓ b) ^ 2 = a * b / |b / a|ₘ := by + rw [← inf_mul_sup a b, ← sup_div_inf_eq_mabs_div, div_eq_mul_inv, div_eq_mul_inv, mul_inv_rev, + inv_inv, mul_assoc, mul_inv_cancel_comm_assoc, ← pow_two] + +-- See, e.g. Zaanen, Lectures on Riesz Spaces +-- 3rd lecture +@[to_additive] +lemma mabs_div_sup_mul_mabs_div_inf (a b c : α) : + |(a ⊔ c) / (b ⊔ c)|ₘ * |(a ⊓ c) / (b ⊓ c)|ₘ = |a / b|ₘ := by + letI : DistribLattice α := CommGroup.toDistribLattice α + calc + |(a ⊔ c) / (b ⊔ c)|ₘ * |(a ⊓ c) / (b ⊓ c)|ₘ = + (b ⊔ c ⊔ (a ⊔ c)) / ((b ⊔ c) ⊓ (a ⊔ c)) * |(a ⊓ c) / (b ⊓ c)|ₘ := by + rw [sup_div_inf_eq_mabs_div] + _ = (b ⊔ c ⊔ (a ⊔ c)) / ((b ⊔ c) ⊓ (a ⊔ c)) * ((b ⊓ c ⊔ a ⊓ c) / (b ⊓ c ⊓ (a ⊓ c))) := by + rw [sup_div_inf_eq_mabs_div (b ⊓ c) (a ⊓ c)] + _ = (b ⊔ a ⊔ c) / (b ⊓ a ⊔ c) * (((b ⊔ a) ⊓ c) / (b ⊓ a ⊓ c)) := by + rw [← sup_inf_right, ← inf_sup_right, sup_assoc, sup_comm c (a ⊔ c), sup_right_idem, + sup_assoc, inf_assoc, inf_comm c (a ⊓ c), inf_right_idem, inf_assoc] + _ = (b ⊔ a ⊔ c) * ((b ⊔ a) ⊓ c) / ((b ⊓ a ⊔ c) * (b ⊓ a ⊓ c)) := by rw [div_mul_div_comm] + _ = (b ⊔ a) * c / ((b ⊓ a) * c) := by + rw [mul_comm, inf_mul_sup, mul_comm (b ⊓ a ⊔ c), inf_mul_sup] + _ = (b ⊔ a) / (b ⊓ a) := by + rw [div_eq_mul_inv, mul_inv_rev, mul_assoc, mul_inv_cancel_left, ← div_eq_mul_inv] + _ = |a / b|ₘ := by rw [sup_div_inf_eq_mabs_div] + +@[to_additive] lemma mabs_sup_div_sup_le_mabs (a b c : α) : |(a ⊔ c) / (b ⊔ c)|ₘ ≤ |a / b|ₘ := by + apply le_of_mul_le_of_one_le_left _ (one_le_mabs _); rw [mabs_div_sup_mul_mabs_div_inf] + +@[to_additive] lemma mabs_inf_div_inf_le_mabs (a b c : α) : |(a ⊓ c) / (b ⊓ c)|ₘ ≤ |a / b|ₘ := by + apply le_of_mul_le_of_one_le_right _ (one_le_mabs _); rw [mabs_div_sup_mul_mabs_div_inf] + +-- Commutative case, Zaanen, 3rd lecture +-- For the non-commutative case, see Birkhoff Theorem 19 (27) +@[to_additive Birkhoff_inequalities] +lemma m_Birkhoff_inequalities (a b c : α) : + |(a ⊔ c) / (b ⊔ c)|ₘ ⊔ |(a ⊓ c) / (b ⊓ c)|ₘ ≤ |a / b|ₘ := + sup_le (mabs_sup_div_sup_le_mabs a b c) (mabs_inf_div_inf_le_mabs a b c) + +end CommGroup +end Lattice + +section LinearOrder +variable [Group α] [LinearOrder α] {a b : α} + +@[to_additive] lemma mabs_choice (x : α) : |x|ₘ = x ∨ |x|ₘ = x⁻¹ := max_choice _ _ + +@[to_additive] lemma le_mabs : a ≤ |b|ₘ ↔ a ≤ b ∨ a ≤ b⁻¹ := le_max_iff + +@[to_additive] lemma mabs_eq_max_inv : |a|ₘ = max a a⁻¹ := rfl + +@[to_additive] lemma lt_mabs : a < |b|ₘ ↔ a < b ∨ a < b⁻¹ := lt_max_iff + +@[to_additive] lemma mabs_by_cases (P : α → Prop) (h1 : P a) (h2 : P a⁻¹) : P |a|ₘ := + sup_ind _ _ h1 h2 + +@[to_additive] lemma eq_or_eq_inv_of_mabs_eq (h : |a|ₘ = b) : a = b ∨ a = b⁻¹ := by + simpa only [← h, eq_comm (a := |a|ₘ), inv_eq_iff_eq_inv] using mabs_choice a + +@[to_additive] lemma mabs_eq_mabs : |a|ₘ = |b|ₘ ↔ a = b ∨ a = b⁻¹ := by + refine ⟨fun h ↦ ?_, by rintro (h | h) <;> simp [h, abs_neg]⟩ + obtain rfl | rfl := eq_or_eq_inv_of_mabs_eq h <;> + simpa only [inv_eq_iff_eq_inv (a := |b|ₘ), inv_inv, inv_inj, or_comm] using mabs_choice b + +@[to_additive] lemma isSquare_mabs : IsSquare |a|ₘ ↔ IsSquare a := + mabs_by_cases (IsSquare · ↔ _) Iff.rfl isSquare_inv + +variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} + +@[to_additive (attr := simp) abs_pos] lemma one_lt_mabs : 1 < |a|ₘ ↔ a ≠ 1 := by + obtain ha | rfl | ha := lt_trichotomy a 1 + · simp [mabs_of_lt_one ha, neg_pos, ha.ne, ha] + · simp + · simp [mabs_of_one_lt ha, ha, ha.ne'] + +@[to_additive abs_pos_of_pos] lemma one_lt_mabs_pos_of_one_lt (h : 1 < a) : 1 < |a|ₘ := + one_lt_mabs.2 h.ne' + +@[to_additive abs_pos_of_neg] lemma one_lt_mabs_of_lt_one (h : a < 1) : 1 < |a|ₘ := + one_lt_mabs.2 h.ne + +@[to_additive] lemma inv_mabs_le (a : α) : |a|ₘ⁻¹ ≤ a := by + obtain h | h := le_total 1 a + · simpa [mabs_of_one_le h] using (inv_le_one'.2 h).trans h + · simp [mabs_of_le_one h] + +@[to_additive add_abs_nonneg] lemma one_le_mul_mabs (a : α) : 1 ≤ a * |a|ₘ := by + rw [← mul_right_inv a]; exact mul_le_mul_left' (inv_le_mabs a) _ + +@[to_additive] lemma inv_mabs_le_inv (a : α) : |a|ₘ⁻¹ ≤ a⁻¹ := by simpa using inv_mabs_le a⁻¹ + +variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] + +@[to_additive] lemma mabs_ne_one : |a|ₘ ≠ 1 ↔ a ≠ 1 := + (one_le_mabs a).gt_iff_ne.symm.trans one_lt_mabs + +@[to_additive (attr := simp)] lemma mabs_eq_one : |a|ₘ = 1 ↔ a = 1 := not_iff_not.1 mabs_ne_one + +@[to_additive (attr := simp) abs_nonpos_iff] lemma mabs_le_one : |a|ₘ ≤ 1 ↔ a = 1 := + (one_le_mabs a).le_iff_eq.trans mabs_eq_one + +@[to_additive] lemma mabs_le_mabs_of_le_one (ha : a ≤ 1) (hab : b ≤ a) : |a|ₘ ≤ |b|ₘ := by + rw [mabs_of_le_one ha, mabs_of_le_one (hab.trans ha)]; exact inv_le_inv_iff.mpr hab + +@[to_additive] lemma mabs_lt : |a|ₘ < b ↔ b⁻¹ < a ∧ a < b := + max_lt_iff.trans <| and_comm.trans <| by rw [inv_lt'] + +@[to_additive] lemma inv_lt_of_mabs_lt (h : |a|ₘ < b) : b⁻¹ < a := (mabs_lt.mp h).1 + +@[to_additive] lemma lt_of_mabs_lt : |a|ₘ < b → a < b := (le_mabs_self _).trans_lt + +@[to_additive] lemma max_div_min_eq_mabs' (a b : α) : max a b / min a b = |a / b|ₘ := by + rcases le_total a b with ab | ba + · rw [max_eq_right ab, min_eq_left ab, mabs_of_le_one, inv_div] + rwa [div_le_one'] + · rw [max_eq_left ba, min_eq_right ba, mabs_of_one_le] + rwa [one_le_div'] + +@[to_additive] lemma max_div_min_eq_mabs (a b : α) : max a b / min a b = |b / a|ₘ := by + rw [mabs_div_comm, max_div_min_eq_mabs'] + +end LinearOrder + +namespace LatticeOrderedAddCommGroup +variable [Lattice α] [AddCommGroup α] {s t : Set α} + +/-- A set `s` in a lattice ordered group is *solid* if for all `x ∈ s` and all `y ∈ α` such that +`|y| ≤ |x|`, then `y ∈ s`. -/ +def IsSolid (s : Set α) : Prop := ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, |y| ≤ |x| → y ∈ s + +/-- The solid closure of a subset `s` is the smallest superset of `s` that is solid. -/ +def solidClosure (s : Set α) : Set α := {y | ∃ x ∈ s, |y| ≤ |x|} + +lemma isSolid_solidClosure (s : Set α) : IsSolid (solidClosure s) := + fun _ ⟨y, hy, hxy⟩ _ hzx ↦ ⟨y, hy, hzx.trans hxy⟩ + +lemma solidClosure_min (hst : s ⊆ t) (ht : IsSolid t) : solidClosure s ⊆ t := + fun _ ⟨_, hy, hxy⟩ ↦ ht (hst hy) hxy + +end LatticeOrderedAddCommGroup + +namespace Pi +variable {ι : Type*} {α : ι → Type*} [∀ i, AddGroup (α i)] [∀ i, Lattice (α i)] + +@[simp] lemma abs_apply (f : ∀ i, α i) (i : ι) : |f| i = |f i| := rfl + +lemma abs_def (f : ∀ i, α i) : |f| = fun i ↦ |f i| := rfl + +end Pi + +@[deprecated (since := "2024-01-13")] alias neg_le_abs_self := neg_le_abs +@[deprecated (since := "2024-01-13")] alias neg_abs_le_self := neg_abs_le diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 7221d786025e6..bbd29978dc36c 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -120,10 +120,10 @@ theorem mul_inv_le_of_le_mul (hab : a ≤ b * c) : a * c⁻¹ ≤ b := by · exact le_of_le_mul_right h (by simpa [h] using hab) theorem inv_le_one₀ (ha : a ≠ 0) : a⁻¹ ≤ 1 ↔ 1 ≤ a := - @inv_le_one' _ _ _ _ <| Units.mk0 a ha + inv_le_one' (a := Units.mk0 a ha) theorem one_le_inv₀ (ha : a ≠ 0) : 1 ≤ a⁻¹ ↔ a ≤ 1 := - @one_le_inv' _ _ _ _ <| Units.mk0 a ha + one_le_inv' (a := Units.mk0 a ha) theorem le_mul_inv_iff₀ (hc : c ≠ 0) : a ≤ b * c⁻¹ ↔ a * c ≤ b := ⟨fun h ↦ inv_inv c ▸ mul_inv_le_of_le_mul h, le_mul_inv_of_mul_le hc⟩ @@ -168,10 +168,10 @@ theorem mul_lt_right₀ (c : α) (h : a < b) (hc : c ≠ 0) : a * c < b * c := b exact le_of_le_mul_right hc h theorem inv_lt_one₀ (ha : a ≠ 0) : a⁻¹ < 1 ↔ 1 < a := - @inv_lt_one' _ _ _ _ <| Units.mk0 a ha + inv_lt_one' (a := Units.mk0 a ha) theorem one_lt_inv₀ (ha : a ≠ 0) : 1 < a⁻¹ ↔ a < 1 := - @one_lt_inv' _ _ _ _ <| Units.mk0 a ha + one_lt_inv' (a := Units.mk0 a ha) theorem inv_lt_inv₀ (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ < b⁻¹ ↔ b < a := show (Units.mk0 a ha)⁻¹ < (Units.mk0 b hb)⁻¹ ↔ Units.mk0 b hb < Units.mk0 a ha from diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 3100a9ae94237..3fdf5f2d9d810 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -215,11 +215,8 @@ theorem le_of_mul_le_mul_right [MulPosReflectLE α] (bc : b * a ≤ c * a) (a0 : @ContravariantClass.elim α>0 α (fun x y => y * x) (· ≤ ·) _ ⟨a, a0⟩ _ _ bc alias lt_of_mul_lt_mul_of_nonneg_left := lt_of_mul_lt_mul_left - alias lt_of_mul_lt_mul_of_nonneg_right := lt_of_mul_lt_mul_right - alias le_of_mul_le_mul_of_pos_left := le_of_mul_le_mul_left - alias le_of_mul_le_mul_of_pos_right := le_of_mul_le_mul_right @[simp] @@ -245,45 +242,99 @@ alias mul_le_mul_iff_of_pos_right := mul_le_mul_right alias mul_lt_mul_iff_of_pos_left := mul_lt_mul_left alias mul_lt_mul_iff_of_pos_right := mul_lt_mul_right -theorem mul_lt_mul_of_pos_of_nonneg [PosMulStrictMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c < d) - (a0 : 0 < a) (d0 : 0 ≤ d) : a * c < b * d := +theorem mul_le_mul_of_nonneg [PosMulMono α] [MulPosMono α] + (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 ≤ d) : a * c ≤ b * d := + (mul_le_mul_of_nonneg_left h₂ a0).trans (mul_le_mul_of_nonneg_right h₁ d0) + +@[deprecated (since := "2024-07-13")] +alias mul_le_mul_of_le_of_le := mul_le_mul_of_nonneg + +theorem mul_le_mul_of_nonneg' [PosMulMono α] [MulPosMono α] + (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) (b0 : 0 ≤ b) : a * c ≤ b * d := + (mul_le_mul_of_nonneg_right h₁ c0).trans (mul_le_mul_of_nonneg_left h₂ b0) + +theorem mul_lt_mul_of_le_of_lt_of_pos_of_nonneg [PosMulStrictMono α] [MulPosMono α] + (h₁ : a ≤ b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 ≤ d) : a * c < b * d := (mul_lt_mul_of_pos_left h₂ a0).trans_le (mul_le_mul_of_nonneg_right h₁ d0) -theorem mul_lt_mul_of_le_of_le' [PosMulStrictMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c < d) - (b0 : 0 < b) (c0 : 0 ≤ c) : a * c < b * d := +alias mul_lt_mul_of_pos_of_nonneg := mul_lt_mul_of_le_of_lt_of_pos_of_nonneg + +theorem mul_lt_mul_of_le_of_lt_of_nonneg_of_pos [PosMulStrictMono α] [MulPosMono α] + (h₁ : a ≤ b) (h₂ : c < d) (c0 : 0 ≤ c) (b0 : 0 < b) : a * c < b * d := (mul_le_mul_of_nonneg_right h₁ c0).trans_lt (mul_lt_mul_of_pos_left h₂ b0) -theorem mul_lt_mul_of_nonneg_of_pos [PosMulMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c ≤ d) - (a0 : 0 ≤ a) (d0 : 0 < d) : a * c < b * d := +alias mul_lt_mul_of_nonneg_of_pos' := mul_lt_mul_of_le_of_lt_of_nonneg_of_pos + +@[deprecated (since := "2024-07-13")] +alias mul_lt_mul_of_le_of_le' := mul_lt_mul_of_le_of_lt_of_nonneg_of_pos + +theorem mul_lt_mul_of_lt_of_le_of_nonneg_of_pos [PosMulMono α] [MulPosStrictMono α] + (h₁ : a < b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 < d) : a * c < b * d := (mul_le_mul_of_nonneg_left h₂ a0).trans_lt (mul_lt_mul_of_pos_right h₁ d0) -theorem mul_lt_mul_of_le_of_lt' [PosMulMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c ≤ d) - (b0 : 0 ≤ b) (c0 : 0 < c) : a * c < b * d := +alias mul_lt_mul_of_nonneg_of_pos := mul_lt_mul_of_lt_of_le_of_nonneg_of_pos + +theorem mul_lt_mul_of_lt_of_le_of_pos_of_nonneg [PosMulMono α] [MulPosStrictMono α] + (h₁ : a < b) (h₂ : c ≤ d) (c0 : 0 < c) (b0 : 0 ≤ b) : a * c < b * d := (mul_lt_mul_of_pos_right h₁ c0).trans_le (mul_le_mul_of_nonneg_left h₂ b0) -theorem mul_lt_mul_of_pos_of_pos [PosMulStrictMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c < d) - (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := +alias mul_lt_mul_of_pos_of_nonneg' := mul_lt_mul_of_lt_of_le_of_pos_of_nonneg + +@[deprecated (since := "2024-07-13")] +alias mul_lt_mul_of_le_of_lt' := mul_lt_mul_of_lt_of_le_of_pos_of_nonneg + +theorem mul_lt_mul_of_pos [PosMulStrictMono α] [MulPosStrictMono α] + (h₁ : a < b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := (mul_lt_mul_of_pos_left h₂ a0).trans (mul_lt_mul_of_pos_right h₁ d0) -theorem mul_lt_mul_of_lt_of_lt' [PosMulStrictMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c < d) - (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := +@[deprecated (since := "2024-07-13")] +alias mul_lt_mul_of_pos_of_pos := mul_lt_mul_of_pos + +theorem mul_lt_mul_of_pos' [PosMulStrictMono α] [MulPosStrictMono α] + (h₁ : a < b) (h₂ : c < d) (c0 : 0 < c) (b0 : 0 < b) : a * c < b * d := (mul_lt_mul_of_pos_right h₁ c0).trans (mul_lt_mul_of_pos_left h₂ b0) -theorem mul_lt_of_mul_lt_of_nonneg_left [PosMulMono α] (h : a * b < c) (hdb : d ≤ b) (ha : 0 ≤ a) : +@[deprecated (since := "2024-07-13")] +alias mul_lt_mul_of_lt_of_lt' := mul_lt_mul_of_pos' + +alias mul_le_mul := mul_le_mul_of_nonneg' +attribute [gcongr] mul_le_mul + +alias mul_lt_mul := mul_lt_mul_of_pos_of_nonneg' + +alias mul_lt_mul' := mul_lt_mul_of_nonneg_of_pos' + +theorem mul_le_of_mul_le_of_nonneg_left [PosMulMono α] (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 ≤ a) : + a * d ≤ c := + (mul_le_mul_of_nonneg_left hle a0).trans h + +theorem mul_lt_of_mul_lt_of_nonneg_left [PosMulMono α] (h : a * b < c) (hle : d ≤ b) (a0 : 0 ≤ a) : a * d < c := - (mul_le_mul_of_nonneg_left hdb ha).trans_lt h + (mul_le_mul_of_nonneg_left hle a0).trans_lt h -theorem lt_mul_of_lt_mul_of_nonneg_left [PosMulMono α] (h : a < b * c) (hcd : c ≤ d) (hb : 0 ≤ b) : +theorem le_mul_of_le_mul_of_nonneg_left [PosMulMono α] (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 ≤ b) : + a ≤ b * d := + h.trans (mul_le_mul_of_nonneg_left hle b0) + +theorem lt_mul_of_lt_mul_of_nonneg_left [PosMulMono α] (h : a < b * c) (hle : c ≤ d) (b0 : 0 ≤ b) : a < b * d := - h.trans_le <| mul_le_mul_of_nonneg_left hcd hb + h.trans_le (mul_le_mul_of_nonneg_left hle b0) + +theorem mul_le_of_mul_le_of_nonneg_right [MulPosMono α] (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 ≤ b) : + d * b ≤ c := + (mul_le_mul_of_nonneg_right hle b0).trans h -theorem mul_lt_of_mul_lt_of_nonneg_right [MulPosMono α] (h : a * b < c) (hda : d ≤ a) (hb : 0 ≤ b) : +theorem mul_lt_of_mul_lt_of_nonneg_right [MulPosMono α] (h : a * b < c) (hle : d ≤ a) (b0 : 0 ≤ b) : d * b < c := - (mul_le_mul_of_nonneg_right hda hb).trans_lt h + (mul_le_mul_of_nonneg_right hle b0).trans_lt h + +theorem le_mul_of_le_mul_of_nonneg_right [MulPosMono α] (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 ≤ c) : + a ≤ d * c := + h.trans (mul_le_mul_of_nonneg_right hle c0) -theorem lt_mul_of_lt_mul_of_nonneg_right [MulPosMono α] (h : a < b * c) (hbd : b ≤ d) (hc : 0 ≤ c) : +theorem lt_mul_of_lt_mul_of_nonneg_right [MulPosMono α] (h : a < b * c) (hle : b ≤ d) (c0 : 0 ≤ c) : a < d * c := - h.trans_le <| mul_le_mul_of_nonneg_right hbd hc + h.trans_le (mul_le_mul_of_nonneg_right hle c0) end Preorder @@ -397,35 +448,25 @@ theorem pos_iff_pos_of_mul_pos [PosMulReflectLT α] [MulPosReflectLT α] (hab : 0 < a ↔ 0 < b := ⟨pos_of_mul_pos_right hab ∘ le_of_lt, pos_of_mul_pos_left hab ∘ le_of_lt⟩ -theorem mul_le_mul_of_le_of_le [PosMulMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 ≤ a) - (d0 : 0 ≤ d) : a * c ≤ b * d := - (mul_le_mul_of_nonneg_left h₂ a0).trans <| mul_le_mul_of_nonneg_right h₁ d0 +/-- Assumes left strict covariance. -/ +theorem Left.mul_lt_mul_of_nonneg [PosMulStrictMono α] [MulPosMono α] + (h₁ : a < b) (h₂ : c < d) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : a * c < b * d := + mul_lt_mul_of_le_of_lt_of_nonneg_of_pos h₁.le h₂ c0 (a0.trans_lt h₁) -@[gcongr] -theorem mul_le_mul [PosMulMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) - (b0 : 0 ≤ b) : a * c ≤ b * d := - (mul_le_mul_of_nonneg_right h₁ c0).trans <| mul_le_mul_of_nonneg_left h₂ b0 +/-- Assumes right strict covariance. -/ +theorem Right.mul_lt_mul_of_nonneg [PosMulMono α] [MulPosStrictMono α] + (h₁ : a < b) (h₂ : c < d) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : a * c < b * d := + mul_lt_mul_of_lt_of_le_of_nonneg_of_pos h₁ h₂.le a0 (c0.trans_lt h₂) + +alias mul_lt_mul_of_nonneg := Left.mul_lt_mul_of_nonneg + +alias mul_lt_mul'' := Left.mul_lt_mul_of_nonneg +attribute [gcongr] mul_lt_mul'' theorem mul_self_le_mul_self [PosMulMono α] [MulPosMono α] (ha : 0 ≤ a) (hab : a ≤ b) : a * a ≤ b * b := mul_le_mul hab hab ha <| ha.trans hab -theorem mul_le_of_mul_le_of_nonneg_left [PosMulMono α] (h : a * b ≤ c) (hle : d ≤ b) - (a0 : 0 ≤ a) : a * d ≤ c := - (mul_le_mul_of_nonneg_left hle a0).trans h - -theorem le_mul_of_le_mul_of_nonneg_left [PosMulMono α] (h : a ≤ b * c) (hle : c ≤ d) - (b0 : 0 ≤ b) : a ≤ b * d := - h.trans (mul_le_mul_of_nonneg_left hle b0) - -theorem mul_le_of_mul_le_of_nonneg_right [MulPosMono α] (h : a * b ≤ c) (hle : d ≤ a) - (b0 : 0 ≤ b) : d * b ≤ c := - (mul_le_mul_of_nonneg_right hle b0).trans h - -theorem le_mul_of_le_mul_of_nonneg_right [MulPosMono α] (h : a ≤ b * c) (hle : b ≤ d) - (c0 : 0 ≤ c) : a ≤ d * c := - h.trans (mul_le_mul_of_nonneg_right hle c0) - end Preorder section PartialOrder @@ -536,7 +577,6 @@ theorem pos_and_pos_or_neg_and_neg_of_mul_pos [PosMulMono α] [MulPosMono α] (h · refine Or.inl ⟨ha, lt_imp_lt_of_le_imp_le (fun hb => ?_) hab⟩ exact mul_nonpos_of_nonneg_of_nonpos ha.le hb - theorem neg_of_mul_pos_right [PosMulMono α] [MulPosMono α] (h : 0 < a * b) (ha : a ≤ 0) : b < 0 := ((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_left fun h => h.1.not_le ha).2 @@ -546,17 +586,21 @@ theorem neg_of_mul_pos_left [PosMulMono α] [MulPosMono α] (h : 0 < a * b) (ha theorem neg_iff_neg_of_mul_pos [PosMulMono α] [MulPosMono α] (hab : 0 < a * b) : a < 0 ↔ b < 0 := ⟨neg_of_mul_pos_right hab ∘ le_of_lt, neg_of_mul_pos_left hab ∘ le_of_lt⟩ -theorem Left.neg_of_mul_neg_left [PosMulMono α] (h : a * b < 0) (h1 : 0 ≤ a) : b < 0 := - lt_of_not_ge fun h2 : b ≥ 0 => (Left.mul_nonneg h1 h2).not_lt h +theorem Left.neg_of_mul_neg_right [PosMulMono α] (h : a * b < 0) (a0 : 0 ≤ a) : b < 0 := + lt_of_not_ge fun b0 : b ≥ 0 => (Left.mul_nonneg a0 b0).not_lt h + +alias neg_of_mul_neg_right := Left.neg_of_mul_neg_right + +theorem Right.neg_of_mul_neg_right [MulPosMono α] (h : a * b < 0) (a0 : 0 ≤ a) : b < 0 := + lt_of_not_ge fun b0 : b ≥ 0 => (Right.mul_nonneg a0 b0).not_lt h -theorem Right.neg_of_mul_neg_left [MulPosMono α] (h : a * b < 0) (h1 : 0 ≤ a) : b < 0 := - lt_of_not_ge fun h2 : b ≥ 0 => (Right.mul_nonneg h1 h2).not_lt h +theorem Left.neg_of_mul_neg_left [PosMulMono α] (h : a * b < 0) (b0 : 0 ≤ b) : a < 0 := + lt_of_not_ge fun a0 : a ≥ 0 => (Left.mul_nonneg a0 b0).not_lt h -theorem Left.neg_of_mul_neg_right [PosMulMono α] (h : a * b < 0) (h1 : 0 ≤ b) : a < 0 := - lt_of_not_ge fun h2 : a ≥ 0 => (Left.mul_nonneg h2 h1).not_lt h +alias neg_of_mul_neg_left := Left.neg_of_mul_neg_left -theorem Right.neg_of_mul_neg_right [MulPosMono α] (h : a * b < 0) (h1 : 0 ≤ b) : a < 0 := - lt_of_not_ge fun h2 : a ≥ 0 => (Right.mul_nonneg h2 h1).not_lt h +theorem Right.neg_of_mul_neg_left [MulPosMono α] (h : a * b < 0) (b0 : 0 ≤ b) : a < 0 := + lt_of_not_ge fun a0 : a ≥ 0 => (Right.mul_nonneg a0 b0).not_lt h end LinearOrder diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index 1fb431acf73b2..f8a5ec5424139 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -182,15 +182,18 @@ end OrderedZero section OrderedAddCommGroup variable [OrderedAddCommGroup α] [OrderedAddCommMonoid β] [i : FunLike F α β] -variable [iamhc : AddMonoidHomClass F α β] (f : F) +variable (f : F) -theorem monotone_iff_map_nonneg : Monotone (f : α → β) ↔ ∀ a, 0 ≤ a → 0 ≤ f a := +theorem monotone_iff_map_nonneg [iamhc : AddMonoidHomClass F α β] : + Monotone (f : α → β) ↔ ∀ a, 0 ≤ a → 0 ≤ f a := ⟨fun h a => by rw [← map_zero f] apply h, fun h a b hl => by rw [← sub_add_cancel b a, map_add f] exact le_add_of_nonneg_left (h _ <| sub_nonneg.2 hl)⟩ +variable [iamhc : AddMonoidHomClass F α β] + theorem antitone_iff_map_nonpos : Antitone (f : α → β) ↔ ∀ a, 0 ≤ a → f a ≤ 0 := monotone_toDual_comp_iff.symm.trans <| monotone_iff_map_nonneg (β := βᵒᵈ) (iamhc := iamhc) _ @@ -202,7 +205,8 @@ theorem antitone_iff_map_nonneg : Antitone (f : α → β) ↔ ∀ a ≤ 0, 0 variable [CovariantClass β β (· + ·) (· < ·)] -theorem strictMono_iff_map_pos : StrictMono (f : α → β) ↔ ∀ a, 0 < a → 0 < f a := by +theorem strictMono_iff_map_pos [iamhc : AddMonoidHomClass F α β] : + StrictMono (f : α → β) ↔ ∀ a, 0 < a → 0 < f a := by refine ⟨fun h a => ?_, fun h a b hl => ?_⟩ · rw [← map_zero f] apply h diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index 09ae52da277f9..2103932e456c4 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -183,15 +183,17 @@ instance NonemptyInterval.hasNSMul [AddMonoid α] [Preorder α] [CovariantClass section Pow -variable [Monoid α] [Preorder α] [CovariantClass α α (· * ·) (· ≤ ·)] - [CovariantClass α α (swap (· * ·)) (· ≤ ·)] +variable [Monoid α] [Preorder α] @[to_additive existing] -instance NonemptyInterval.hasPow : Pow (NonemptyInterval α) ℕ := +instance NonemptyInterval.hasPow + [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : + Pow (NonemptyInterval α) ℕ := ⟨fun s n => ⟨s.toProd ^ n, pow_le_pow_left' s.fst_le_snd _⟩⟩ namespace NonemptyInterval +variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] variable (s : NonemptyInterval α) (a : α) (n : ℕ) @[to_additive (attr := simp) toProd_nsmul] diff --git a/Mathlib/Algebra/Order/Module/Defs.lean b/Mathlib/Algebra/Order/Module/Defs.lean index cac83f26595dc..fec86336439c9 100644 --- a/Mathlib/Algebra/Order/Module/Defs.lean +++ b/Mathlib/Algebra/Order/Module/Defs.lean @@ -1052,38 +1052,57 @@ end Pi section Lift variable {γ : Type*} [Zero α] [Preorder α] [Zero β] [Preorder β] [Zero γ] [Preorder γ] - [SMul α β] [SMul α γ] (f : β → γ) (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) - (smul : ∀ (a : α) b, f (a • b) = a • f b) (zero : f 0 = 0) + [SMul α β] [SMul α γ] (f : β → γ) -lemma PosSMulMono.lift [PosSMulMono α γ] : PosSMulMono α β where +lemma PosSMulMono.lift [PosSMulMono α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) : PosSMulMono α β where elim a ha b₁ b₂ hb := by simp only [← hf, smul] at *; exact smul_le_smul_of_nonneg_left hb ha -lemma PosSMulStrictMono.lift [PosSMulStrictMono α γ] : PosSMulStrictMono α β where +lemma PosSMulStrictMono.lift [PosSMulStrictMono α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) : PosSMulStrictMono α β where elim a ha b₁ b₂ hb := by simp only [← lt_iff_lt_of_le_iff_le' hf hf, smul] at *; exact smul_lt_smul_of_pos_left hb ha -lemma PosSMulReflectLE.lift [PosSMulReflectLE α γ] : PosSMulReflectLE α β where +lemma PosSMulReflectLE.lift [PosSMulReflectLE α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) : PosSMulReflectLE α β where elim a ha b₁ b₂ h := hf.1 <| le_of_smul_le_smul_left (by simpa only [smul] using hf.2 h) ha -lemma PosSMulReflectLT.lift [PosSMulReflectLT α γ] : PosSMulReflectLT α β where +lemma PosSMulReflectLT.lift [PosSMulReflectLT α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) : PosSMulReflectLT α β where elim a ha b₁ b₂ h := by simp only [← lt_iff_lt_of_le_iff_le' hf hf, smul] at *; exact lt_of_smul_lt_smul_left h ha -lemma SMulPosMono.lift [SMulPosMono α γ] : SMulPosMono α β where +lemma SMulPosMono.lift [SMulPosMono α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) + (zero : f 0 = 0) : SMulPosMono α β where elim b hb a₁ a₂ ha := by simp only [← hf, zero, smul] at *; exact smul_le_smul_of_nonneg_right ha hb -lemma SMulPosStrictMono.lift [SMulPosStrictMono α γ] : SMulPosStrictMono α β where +lemma SMulPosStrictMono.lift [SMulPosStrictMono α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) + (zero : f 0 = 0) : SMulPosStrictMono α β where elim b hb a₁ a₂ ha := by simp only [← lt_iff_lt_of_le_iff_le' hf hf, zero, smul] at * exact smul_lt_smul_of_pos_right ha hb -lemma SMulPosReflectLE.lift [SMulPosReflectLE α γ] : SMulPosReflectLE α β where +lemma SMulPosReflectLE.lift [SMulPosReflectLE α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) + (zero : f 0 = 0) : SMulPosReflectLE α β where elim b hb a₁ a₂ h := by simp only [← hf, ← lt_iff_lt_of_le_iff_le' hf hf, zero, smul] at * exact le_of_smul_le_smul_right h hb -lemma SMulPosReflectLT.lift [SMulPosReflectLT α γ] : SMulPosReflectLT α β where +lemma SMulPosReflectLT.lift [SMulPosReflectLT α γ] + (hf : ∀ {b₁ b₂}, f b₁ ≤ f b₂ ↔ b₁ ≤ b₂) + (smul : ∀ (a : α) b, f (a • b) = a • f b) + (zero : f 0 = 0) : SMulPosReflectLT α β where elim b hb a₁ a₂ h := by simp only [← hf, ← lt_iff_lt_of_le_iff_le' hf hf, zero, smul] at * exact lt_of_smul_lt_smul_right h hb diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index f5e39354ca3b6..08d3cc925d652 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -45,7 +45,7 @@ theorem one_le_pow_of_one_le' {a : M} (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n @[to_additive nsmul_nonpos] theorem pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := - @one_le_pow_of_one_le' Mᵒᵈ _ _ _ _ H n + one_le_pow_of_one_le' (M := Mᵒᵈ) H n @[to_additive (attr := gcongr) nsmul_le_nsmul_left] theorem pow_le_pow_right' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := @@ -69,7 +69,7 @@ theorem one_lt_pow' {a : M} (ha : 1 < a) {k : ℕ} (hk : k ≠ 0) : 1 < a ^ k := @[to_additive nsmul_neg] theorem pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := - @one_lt_pow' Mᵒᵈ _ _ _ _ ha k hk + one_lt_pow' (M := Mᵒᵈ) ha hk @[to_additive (attr := gcongr) nsmul_lt_nsmul_left] theorem pow_lt_pow_right' [CovariantClass M M (· * ·) (· < ·)] {a : M} {n m : ℕ} (ha : 1 < a) @@ -192,7 +192,7 @@ theorem one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ @[to_additive] theorem pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 := - @one_le_pow_iff Mᵒᵈ _ _ _ _ _ hn + one_le_pow_iff (M := Mᵒᵈ) hn @[to_additive nsmul_pos_iff] theorem one_lt_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 < x ^ n ↔ 1 < x := diff --git a/Mathlib/Algebra/Order/Pi.lean b/Mathlib/Algebra/Order/Pi.lean index 5c262d911813e..b89f80e5f895b 100644 --- a/Mathlib/Algebra/Order/Pi.lean +++ b/Mathlib/Algebra/Order/Pi.lean @@ -102,19 +102,19 @@ variable {β} [Nonempty β] @[to_additive (attr := simp) const_nonneg] theorem one_le_const : 1 ≤ const β a ↔ 1 ≤ a := - @const_le_const _ _ _ _ 1 _ + const_le_const @[to_additive (attr := simp) const_pos] theorem one_lt_const : 1 < const β a ↔ 1 < a := - @const_lt_const _ _ _ _ 1 a + const_lt_const @[to_additive (attr := simp)] theorem const_le_one : const β a ≤ 1 ↔ a ≤ 1 := - @const_le_const _ _ _ _ _ 1 + const_le_const @[to_additive (attr := simp) const_neg'] theorem const_lt_one : const β a < 1 ↔ a < 1 := - @const_lt_const _ _ _ _ _ 1 + const_lt_const end const diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 234159d4be10b..278ac4ae53594 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -113,8 +113,8 @@ variable [StrictOrderedSemiring R] {a x y : R} {n m : ℕ} theorem pow_lt_pow_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, n ≠ 0 → x ^ n < y ^ n | 0, hn => by contradiction | n + 1, _ => by - simpa only [pow_succ] using - mul_lt_mul_of_le_of_le' (pow_le_pow_left hx h.le _) h (pow_pos (hx.trans_lt h) _) hx + simpa only [pow_succ] using mul_lt_mul_of_le_of_lt_of_nonneg_of_pos + (pow_le_pow_left hx h.le _) h hx (pow_pos (hx.trans_lt h) _) /-- See also `pow_left_strictMono` and `Nat.pow_left_strictMono`. -/ lemma pow_left_strictMonoOn (hn : n ≠ 0) : StrictMonoOn (· ^ n : R → R) {a | 0 ≤ a} := diff --git a/Mathlib/Algebra/Order/Ring/Canonical.lean b/Mathlib/Algebra/Order/Ring/Canonical.lean index f64f1200e3475..750d5c7e2a2ff 100644 --- a/Mathlib/Algebra/Order/Ring/Canonical.lean +++ b/Mathlib/Algebra/Order/Ring/Canonical.lean @@ -80,7 +80,7 @@ protected lemma mul_lt_mul_of_lt_of_lt [PosMulStrictMono α] (hab : a < b) (hcd obtain rfl | hc := eq_zero_or_pos c · rw [mul_zero] exact mul_pos ((zero_le _).trans_lt hab) hcd - · exact mul_lt_mul_of_lt_of_lt' hab hcd ((zero_le _).trans_lt hab) hc + · exact mul_lt_mul_of_pos' hab hcd hc ((zero_le _).trans_lt hab) end CanonicallyOrderedCommSemiring end CanonicallyOrderedCommSemiring diff --git a/Mathlib/Algebra/Order/Ring/Cast.lean b/Mathlib/Algebra/Order/Ring/Cast.lean index 2f8d74e4b82bf..06e3f2297e000 100644 --- a/Mathlib/Algebra/Order/Ring/Cast.lean +++ b/Mathlib/Algebra/Order/Ring/Cast.lean @@ -24,15 +24,17 @@ open Function Nat variable {R : Type*} namespace Int -section OrderedRing -variable [OrderedRing R] +section OrderedAddCommGroupWithOne + +variable [AddCommGroupWithOne R] [PartialOrder R] [CovariantClass R R (· + ·) (· ≤ ·)] +variable [ZeroLEOneClass R] [NeZero (1 : R)] lemma cast_mono : Monotone (Int.cast : ℤ → R) := by intro m n h rw [← sub_nonneg] at h lift n - m to ℕ using h with k hk rw [← sub_nonneg, ← cast_sub, ← hk, cast_natCast] - exact k.cast_nonneg + exact k.cast_nonneg' variable [Nontrivial R] {m n : ℤ} @@ -56,7 +58,7 @@ lemma cast_strictMono : StrictMono (fun x : ℤ => (x : R)) := @[simp] lemma cast_lt_zero : (n : R) < 0 ↔ n < 0 := by rw [← cast_zero, cast_lt] -end OrderedRing +end OrderedAddCommGroupWithOne section LinearOrderedRing variable [LinearOrderedRing R] {a b n : ℤ} {x : R} diff --git a/Mathlib/Algebra/Order/Ring/Star.lean b/Mathlib/Algebra/Order/Ring/Star.lean index 46cde55c71e5a..da189507528bc 100644 --- a/Mathlib/Algebra/Order/Ring/Star.lean +++ b/Mathlib/Algebra/Order/Ring/Star.lean @@ -54,7 +54,7 @@ See note [reducible non-instances]. -/ abbrev toOrderedCommSemiring (R : Type*) [CommSemiring R] [PartialOrder R] [StarRing R] [StarOrderedRing R] : OrderedCommSemiring R where add_le_add_left _ _ := add_le_add_left - zero_le_one := by simpa using star_mul_self_nonneg (1 : R) + zero_le_one := zero_le_one mul_comm := mul_comm mul_le_mul_of_nonneg_left _ _ _ := mul_le_mul_of_nonneg_left mul_le_mul_of_nonneg_right a b c := by simpa only [mul_comm _ c] using mul_le_mul_of_nonneg_left @@ -68,7 +68,7 @@ See note [reducible non-instances]. -/ abbrev toOrderedCommRing (R : Type*) [CommRing R] [PartialOrder R] [StarRing R] [StarOrderedRing R] : OrderedCommRing R where add_le_add_left _ _ := add_le_add_left - zero_le_one := by simpa using star_mul_self_nonneg (1 : R) + zero_le_one := zero_le_one mul_comm := mul_comm mul_nonneg _ _ := let _ := toOrderedCommSemiring R; mul_nonneg diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index a64baa31f0510..d35cda83bc542 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -327,12 +327,6 @@ section StrictOrderedSemiring variable [Semiring α] [PartialOrder α] [ZeroLEOneClass α] [PosMulStrictMono α] [MulPosStrictMono α] [CovariantClass α α (· + ·) (· < ·)] {a b c d : α} -theorem mul_lt_mul (hac : a < c) (hbd : b ≤ d) (hb : 0 < b) (hc : 0 ≤ c) : a * b < c * d := - (mul_lt_mul_of_pos_right hac hb).trans_le <| mul_le_mul_of_nonneg_left hbd hc - -theorem mul_lt_mul' (hac : a ≤ c) (hbd : b < d) (hb : 0 ≤ b) (hc : 0 < c) : a * b < c * d := - (mul_le_mul_of_nonneg_right hac hb).trans_lt <| mul_lt_mul_of_pos_left hbd hc - @[simp] theorem pow_pos (H : 0 < a) : ∀ n : ℕ, 0 < a ^ n | 0 => by @@ -359,10 +353,6 @@ protected theorem Decidable.mul_lt_mul'' [@DecidableRel α (· ≤ ·)] (h1 : a h4.lt_or_eq_dec.elim (fun b0 => mul_lt_mul h1 h2.le b0 <| h3.trans h1.le) fun b0 => by rw [← b0, mul_zero]; exact mul_pos (h3.trans_lt h1) (h4.trans_lt h2) -@[gcongr] -theorem mul_lt_mul'' : a < c → b < d → 0 ≤ a → 0 ≤ b → a * b < c * d := by classical - exact Decidable.mul_lt_mul'' - theorem lt_mul_left (hn : 0 < a) (hm : 1 < b) : a < b * a := by convert mul_lt_mul_of_pos_right hm hn rw [one_mul] @@ -526,12 +516,6 @@ theorem nonneg_of_mul_nonneg_left (h : 0 ≤ a * b) (hb : 0 < b) : 0 ≤ a := theorem nonneg_of_mul_nonneg_right (h : 0 ≤ a * b) (ha : 0 < a) : 0 ≤ b := le_of_not_gt fun hb => (mul_neg_of_pos_of_neg ha hb).not_le h -theorem neg_of_mul_neg_left (h : a * b < 0) (hb : 0 ≤ b) : a < 0 := - lt_of_not_ge fun ha => (mul_nonneg ha hb).not_lt h - -theorem neg_of_mul_neg_right (h : a * b < 0) (ha : 0 ≤ a) : b < 0 := - lt_of_not_ge fun hb => (mul_nonneg ha hb).not_lt h - theorem nonpos_of_mul_nonpos_left (h : a * b ≤ 0) (hb : 0 < b) : a ≤ 0 := le_of_not_gt fun ha : a > 0 => (mul_pos ha hb).not_le h diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 19fed8773f848..fbfa554def41c 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -1050,16 +1050,16 @@ instance commRing [CommRing R] : CommRing R[X] := section NonzeroSemiring -variable [Semiring R] [Nontrivial R] +variable [Semiring R] -instance nontrivial : Nontrivial R[X] := by +instance nontrivial [Nontrivial R] : Nontrivial R[X] := by have h : Nontrivial R[ℕ] := by infer_instance rcases h.exists_pair_ne with ⟨x, y, hxy⟩ refine ⟨⟨⟨x⟩, ⟨y⟩, ?_⟩⟩ simp [hxy] @[simp] -theorem X_ne_zero : (X : R[X]) ≠ 0 := +theorem X_ne_zero [Nontrivial R] : (X : R[X]) ≠ 0 := mt (congr_arg fun p => coeff p 1) (by simp) end NonzeroSemiring diff --git a/Mathlib/Algebra/Polynomial/DenomsClearable.lean b/Mathlib/Algebra/Polynomial/DenomsClearable.lean index 22e12f9048050..a9004b5762139 100644 --- a/Mathlib/Algebra/Polynomial/DenomsClearable.lean +++ b/Mathlib/Algebra/Polynomial/DenomsClearable.lean @@ -87,7 +87,7 @@ theorem one_le_pow_mul_abs_eval_div {K : Type*} [LinearOrderedField K] {f : ℤ[ rw [eq_intCast, one_div_mul_cancel] rw [Int.cast_ne_zero] exact b0.ne.symm) - obtain Fa := _root_.congr_arg abs hF + obtain Fa := congr_arg abs hF rw [eq_one_div_of_mul_eq_one_left bu, eq_intCast, eq_intCast, abs_mul] at Fa rw [abs_of_pos (pow_pos (Int.cast_pos.mpr b0) _ : 0 < (b : K) ^ _), one_div, eq_intCast] at Fa rw [div_eq_mul_inv, ← Fa, ← Int.cast_abs, ← Int.cast_one, Int.cast_le] diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index dae29c648387c..dd585a94cd62a 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -86,7 +86,7 @@ theorem div_wf_lemma (h : degree q ≤ degree p ∧ p ≠ 0) (hq : Monic q) : have hp : leadingCoeff p ≠ 0 := mt leadingCoeff_eq_zero.1 h.2 have hq0 : q ≠ 0 := hq.ne_zero_of_polynomial_ne h.2 have hlt : natDegree q ≤ natDegree p := - Nat.cast_le.1 + (Nat.cast_le (α := WithBot ℕ)).1 (by rw [← degree_eq_natDegree h.2, ← degree_eq_natDegree hq0]; exact h.1) degree_sub_lt (by diff --git a/Mathlib/Algebra/Polynomial/Inductions.lean b/Mathlib/Algebra/Polynomial/Inductions.lean index 877abf4376217..f5a98d3c5ebe0 100644 --- a/Mathlib/Algebra/Polynomial/Inductions.lean +++ b/Mathlib/Algebra/Polynomial/Inductions.lean @@ -163,11 +163,12 @@ theorem degree_pos_induction_on {P : R[X] → Prop} (p : R[X]) (h0 : 0 < degree (hC : ∀ {a}, a ≠ 0 → P (C a * X)) (hX : ∀ {p}, 0 < degree p → P p → P (p * X)) (hadd : ∀ {p} {a}, 0 < degree p → P p → P (p + C a)) : P p := recOnHorner p (fun h => by rw [degree_zero] at h; exact absurd h (by decide)) - (fun p a _ _ ih h0 => - have : 0 < degree p := - lt_of_not_ge fun h => - not_lt_of_ge degree_C_le <| by rwa [eq_C_of_degree_le_zero h, ← C_add] at h0 - hadd this (ih this)) + (fun p a heq0 _ ih h0 => + (have : 0 < degree p := + (lt_of_not_ge fun h => + not_lt_of_ge (degree_C_le (a := a)) <| + by rwa [eq_C_of_degree_le_zero h, ← C_add,heq0,zero_add] at h0) + hadd this (ih this))) (fun p _ ih h0' => if h0 : 0 < degree p then hX h0 (ih h0) else by diff --git a/Mathlib/Algebra/Polynomial/Monic.lean b/Mathlib/Algebra/Polynomial/Monic.lean index 4315e4c1ae052..07cb29792f10d 100644 --- a/Mathlib/Algebra/Polynomial/Monic.lean +++ b/Mathlib/Algebra/Polynomial/Monic.lean @@ -293,36 +293,40 @@ section Injective open Function -variable [Semiring S] {f : R →+* S} (hf : Injective f) +variable [Semiring S] {f : R →+* S} -theorem degree_map_eq_of_injective (p : R[X]) : degree (p.map f) = degree p := +theorem degree_map_eq_of_injective (hf : Injective f) (p : R[X]) : degree (p.map f) = degree p := letI := Classical.decEq R if h : p = 0 then by simp [h] else degree_map_eq_of_leadingCoeff_ne_zero _ (by rw [← f.map_zero]; exact mt hf.eq_iff.1 (mt leadingCoeff_eq_zero.1 h)) -theorem natDegree_map_eq_of_injective (p : R[X]) : natDegree (p.map f) = natDegree p := +theorem natDegree_map_eq_of_injective (hf : Injective f) (p : R[X]) : + natDegree (p.map f) = natDegree p := natDegree_eq_of_degree_eq (degree_map_eq_of_injective hf p) -theorem leadingCoeff_map' (p : R[X]) : leadingCoeff (p.map f) = f (leadingCoeff p) := by +theorem leadingCoeff_map' (hf : Injective f) (p : R[X]) : + leadingCoeff (p.map f) = f (leadingCoeff p) := by unfold leadingCoeff rw [coeff_map, natDegree_map_eq_of_injective hf p] -theorem nextCoeff_map (p : R[X]) : (p.map f).nextCoeff = f p.nextCoeff := by +theorem nextCoeff_map (hf : Injective f) (p : R[X]) : (p.map f).nextCoeff = f p.nextCoeff := by unfold nextCoeff rw [natDegree_map_eq_of_injective hf] split_ifs <;> simp [*] -theorem leadingCoeff_of_injective (p : R[X]) : leadingCoeff (p.map f) = f (leadingCoeff p) := by +theorem leadingCoeff_of_injective (hf : Injective f) (p : R[X]) : + leadingCoeff (p.map f) = f (leadingCoeff p) := by delta leadingCoeff rw [coeff_map f, natDegree_map_eq_of_injective hf p] -theorem monic_of_injective {p : R[X]} (hp : (p.map f).Monic) : p.Monic := by +theorem monic_of_injective (hf : Injective f) {p : R[X]} (hp : (p.map f).Monic) : p.Monic := by apply hf rw [← leadingCoeff_of_injective hf, hp.leadingCoeff, f.map_one] -theorem _root_.Function.Injective.monic_map_iff {p : R[X]} : p.Monic ↔ (p.map f).Monic := +theorem _root_.Function.Injective.monic_map_iff (hf : Injective f) {p : R[X]} : + p.Monic ↔ (p.map f).Monic := ⟨Monic.map _, Polynomial.monic_of_injective hf⟩ end Injective diff --git a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean index 1db55429e94bd..f22387feb2113 100644 --- a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.Algebra.Polynomial.Mirror -import Mathlib.Analysis.Complex.Polynomial +import Mathlib.Data.Int.Order.Units /-! # Unit Trinomials @@ -22,6 +22,7 @@ This file defines irreducible trinomials and proves an irreducibility criterion. -/ +assert_not_exists TopologicalSpace namespace Polynomial @@ -320,30 +321,6 @@ theorem irreducible_of_isCoprime (hp : p.IsUnitTrinomial) (h : IsCoprime p p.mir Irreducible p := irreducible_of_coprime hp fun _ => h.isUnit_of_dvd' -/-- A unit trinomial is irreducible if it has no complex roots in common with its mirror -/ -theorem irreducible_of_coprime' (hp : IsUnitTrinomial p) - (h : ∀ z : ℂ, ¬(aeval z p = 0 ∧ aeval z (mirror p) = 0)) : Irreducible p := by - refine hp.irreducible_of_coprime fun q hq hq' => ?_ - suffices ¬0 < q.natDegree by - rcases hq with ⟨p, rfl⟩ - replace hp := hp.leadingCoeff_isUnit - rw [leadingCoeff_mul] at hp - replace hp := isUnit_of_mul_isUnit_left hp - rw [not_lt, Nat.le_zero] at this - rwa [eq_C_of_natDegree_eq_zero this, isUnit_C, ← this] - intro hq'' - rw [natDegree_pos_iff_degree_pos] at hq'' - rw [← degree_map_eq_of_injective (algebraMap ℤ ℂ).injective_int] at hq'' - cases' Complex.exists_root hq'' with z hz - rw [IsRoot, eval_map, ← aeval_def] at hz - refine h z ⟨?_, ?_⟩ - · cases' hq with g' hg' - rw [hg', aeval_mul, hz, zero_mul] - · cases' hq' with g' hg' - rw [hg', aeval_mul, hz, zero_mul] - --- TODO: Develop more theory (e.g., it suffices to check that `aeval z p ≠ 0` for `z = 0` --- and `z` a root of unity) end IsUnitTrinomial end Polynomial diff --git a/Mathlib/Algebra/Regular/Basic.lean b/Mathlib/Algebra/Regular/Basic.lean index fac1ea50bc528..0a534b873b9c3 100644 --- a/Mathlib/Algebra/Regular/Basic.lean +++ b/Mathlib/Algebra/Regular/Basic.lean @@ -122,7 +122,7 @@ element, then `b` is left-regular. -/ @[to_additive "If an element `b` becomes add-left-regular after adding to it on the left an add-left-regular element, then `b` is add-left-regular."] theorem IsLeftRegular.of_mul (ab : IsLeftRegular (a * b)) : IsLeftRegular b := - Function.Injective.of_comp (by rwa [comp_mul_left a b]) + Function.Injective.of_comp (f := (a * ·)) (by rwa [comp_mul_left a b]) /-- An element is left-regular if and only if multiplying it on the left by a left-regular element is left-regular. -/ @@ -284,12 +284,12 @@ variable [Monoid R] {a b : R} /-- An element admitting a left inverse is left-regular. -/ @[to_additive "An element admitting a left additive opposite is add-left-regular."] theorem isLeftRegular_of_mul_eq_one (h : b * a = 1) : IsLeftRegular a := - @IsLeftRegular.of_mul R _ _ _ (by rw [h]; exact isRegular_one.left) + IsLeftRegular.of_mul (a := b) (by rw [h]; exact isRegular_one.left) /-- An element admitting a right inverse is right-regular. -/ @[to_additive "An element admitting a right additive opposite is add-right-regular."] theorem isRightRegular_of_mul_eq_one (h : a * b = 1) : IsRightRegular a := - IsRightRegular.of_mul (by rw [h]; exact isRegular_one.right) + IsRightRegular.of_mul (a := b) (by rw [h]; exact isRegular_one.right) /-- If `R` is a monoid, an element in `Rˣ` is regular. -/ @[to_additive "If `R` is an additive monoid, an element in `add_units R` is add-regular."] diff --git a/Mathlib/Algebra/Regular/SMul.lean b/Mathlib/Algebra/Regular/SMul.lean index e599b977e2e02..fca8434b15e2c 100644 --- a/Mathlib/Algebra/Regular/SMul.lean +++ b/Mathlib/Algebra/Regular/SMul.lean @@ -129,10 +129,7 @@ variable {M} /-- An element of `R` admitting a left inverse is `M`-regular. -/ theorem of_mul_eq_one (h : a * b = 1) : IsSMulRegular M b := - of_mul - (by - rw [h] - exact one M) + of_mul (a := a) (by rw [h]; exact one M) /-- Any power of an `M`-regular element is `M`-regular. -/ theorem pow (n : ℕ) (ra : IsSMulRegular M a) : IsSMulRegular M (a ^ n) := by diff --git a/Mathlib/Algebra/Ring/Divisibility/Lemmas.lean b/Mathlib/Algebra/Ring/Divisibility/Lemmas.lean index a2574fef5777e..0db93ca03b73f 100644 --- a/Mathlib/Algebra/Ring/Divisibility/Lemmas.lean +++ b/Mathlib/Algebra/Ring/Divisibility/Lemmas.lean @@ -32,14 +32,14 @@ lemma dvd_zsmul_of_dvd [NonUnitalRing R] {x y : R} (z : ℤ) (h : x ∣ y) : x namespace Commute -variable {x y : R} {n m p : ℕ} (hp : n + m ≤ p + 1) +variable {x y : R} {n m p : ℕ} section Semiring -variable [Semiring R] (h_comm : Commute x y) +variable [Semiring R] -lemma pow_dvd_add_pow_of_pow_eq_zero_right (hy : y ^ n = 0) : - x ^ m ∣ (x + y) ^ p := by +lemma pow_dvd_add_pow_of_pow_eq_zero_right (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (hy : y ^ n = 0) : x ^ m ∣ (x + y) ^ p := by rw [h_comm.add_pow'] refine Finset.dvd_sum fun ⟨i, j⟩ hij ↦ ?_ replace hij : i + j = p := by simpa using hij @@ -48,45 +48,45 @@ lemma pow_dvd_add_pow_of_pow_eq_zero_right (hy : y ^ n = 0) : · exact dvd_mul_of_dvd_left (pow_dvd_pow x hi) _ · simp [pow_eq_zero_of_le (by omega : n ≤ j) hy] -lemma pow_dvd_add_pow_of_pow_eq_zero_left (hx : x ^ n = 0) : - y ^ m ∣ (x + y) ^ p := +lemma pow_dvd_add_pow_of_pow_eq_zero_left (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (hx : x ^ n = 0) : y ^ m ∣ (x + y) ^ p := add_comm x y ▸ h_comm.symm.pow_dvd_add_pow_of_pow_eq_zero_right hp hx end Semiring section Ring -variable [Ring R] (h_comm : Commute x y) +variable [Ring R] -lemma pow_dvd_pow_of_sub_pow_eq_zero (h : (x - y) ^ n = 0) : - x ^ m ∣ y ^ p := by +lemma pow_dvd_pow_of_sub_pow_eq_zero (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (h : (x - y) ^ n = 0) : x ^ m ∣ y ^ p := by rw [← sub_add_cancel y x] apply (h_comm.symm.sub_left rfl).pow_dvd_add_pow_of_pow_eq_zero_left hp _ rw [← neg_sub x y, neg_pow, h, mul_zero] -lemma pow_dvd_pow_of_add_pow_eq_zero (h : (x + y) ^ n = 0) : - x ^ m ∣ y ^ p := by +lemma pow_dvd_pow_of_add_pow_eq_zero (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (h : (x + y) ^ n = 0) : x ^ m ∣ y ^ p := by rw [← neg_neg y, neg_pow'] apply dvd_mul_of_dvd_left apply h_comm.neg_right.pow_dvd_pow_of_sub_pow_eq_zero hp simpa -lemma pow_dvd_sub_pow_of_pow_eq_zero_right (hy : y ^ n = 0) : - x ^ m ∣ (x - y) ^ p := +lemma pow_dvd_sub_pow_of_pow_eq_zero_right (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (hy : y ^ n = 0) : x ^ m ∣ (x - y) ^ p := (sub_right rfl h_comm).pow_dvd_pow_of_sub_pow_eq_zero hp (by simpa) -lemma pow_dvd_sub_pow_of_pow_eq_zero_left (hx : x ^ n = 0) : - y ^ m ∣ (x - y) ^ p := by +lemma pow_dvd_sub_pow_of_pow_eq_zero_left (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (hx : x ^ n = 0) : y ^ m ∣ (x - y) ^ p := by rw [← neg_sub y x, neg_pow'] apply dvd_mul_of_dvd_left exact h_comm.symm.pow_dvd_sub_pow_of_pow_eq_zero_right hp hx -lemma add_pow_dvd_pow_of_pow_eq_zero_right (hx : x ^ n = 0) : - (x + y) ^ m ∣ y ^ p := +lemma add_pow_dvd_pow_of_pow_eq_zero_right (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (hx : x ^ n = 0) : (x + y) ^ m ∣ y ^ p := (h_comm.add_left rfl).pow_dvd_pow_of_sub_pow_eq_zero hp (by simpa) -lemma add_pow_dvd_pow_of_pow_eq_zero_left (hy : y ^ n = 0) : - (x + y) ^ m ∣ x ^ p := +lemma add_pow_dvd_pow_of_pow_eq_zero_left (hp : n + m ≤ p + 1) (h_comm : Commute x y) + (hy : y ^ n = 0) : (x + y) ^ m ∣ x ^ p := add_comm x y ▸ h_comm.symm.add_pow_dvd_pow_of_pow_eq_zero_right hp hy end Ring diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 6bf12f1ee448a..a440563a5296e 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -439,10 +439,10 @@ section variable {_ : NonAssocSemiring α} {_ : NonAssocSemiring β} (f : α →+* β) {x y : α} -theorem congr_fun {f g : α →+* β} (h : f = g) (x : α) : f x = g x := +protected theorem congr_fun {f g : α →+* β} (h : f = g) (x : α) : f x = g x := DFunLike.congr_fun h x -theorem congr_arg (f : α →+* β) {x y : α} (h : x = y) : f x = f y := +protected theorem congr_arg (f : α →+* β) {x y : α} (h : x = y) : f x = f y := DFunLike.congr_arg f h theorem coe_inj ⦃f g : α →+* β⦄ (h : (f : α → β) = g) : f = g := diff --git a/Mathlib/Algebra/Ring/Semireal/Defs.lean b/Mathlib/Algebra/Ring/Semireal/Defs.lean new file mode 100644 index 0000000000000..1943e7609eb01 --- /dev/null +++ b/Mathlib/Algebra/Ring/Semireal/Defs.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2024 Florent Schaffhauser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Florent Schaffhauser +-/ + +import Mathlib.Algebra.Ring.SumsOfSquares +import Mathlib.Algebra.Order.Field.Defs + +/-! +# Semireal rings + +A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a sum of +squares. Note that `-1` does not make sense in a semiring. + +For instance, linearly ordered fields are semireal, because sums of squares are positive and `-1` is +not. More generally, linearly ordered semirings in which `a ≤ b → ∃ c, a + c = b` holds, are +semireal. + +## Main declaration + +- `isSemireal`: the predicate asserting that a commutative ring `R` is semireal. + +## References + +- *An introduction to real algebra*, by T.Y. Lam. Rocky Mountain J. Math. 14(4): 767-814 (1984). +[lam_1984](https://doi.org/10.1216/RMJ-1984-14-4-767) +-/ + +variable (R : Type*) + +/-- +A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a sum of +squares. Note that `-1` does not make sense in a semiring. Below we define the class `isSemiReal R` +for all additive monoid `R` equipped with a multiplication, a multiplicative unit and a negation. +-/ +@[mk_iff] +class isSemireal [AddMonoid R] [Mul R] [One R] [Neg R] : Prop where + non_trivial : (0 : R) ≠ 1 + neg_one_not_SumSq : ¬isSumSq (-1 : R) + +instance [LinearOrderedField R] : isSemireal R where + non_trivial := zero_ne_one + neg_one_not_SumSq := fun h ↦ (not_le (α := R)).2 neg_one_lt_zero h.nonneg diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index cdb15249b2126..ef70be292c08d 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -665,7 +665,9 @@ instance instField : Field (center K) where inv_zero := Subtype.ext inv_zero -- TODO: use a nicer defeq nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q x => rfl @[simp] theorem center.coe_inv (a : center K) : ((a⁻¹ : center K) : K) = (a : K)⁻¹ := @@ -848,7 +850,7 @@ theorem exists_list_of_mem_closure {s : Set R} {x : R} (h : x ∈ closure s) : ⟨l ++ m, fun t ht => (List.mem_append.1 ht).elim (hl1 t) (hm1 t), by simp [hl2, hm2]⟩) fun x ⟨L, hL⟩ => ⟨L.map (List.cons (-1)), - List.forall_mem_map_iff.2 fun j hj => List.forall_mem_cons.2 ⟨Or.inr rfl, hL.1 j hj⟩, + List.forall_mem_map.2 fun j hj => List.forall_mem_cons.2 ⟨Or.inr rfl, hL.1 j hj⟩, hL.2 ▸ List.recOn L (by simp) (by simp (config := { contextual := true }) [List.map_cons, add_comm])⟩ diff --git a/Mathlib/Algebra/Ring/SumsOfSquares.lean b/Mathlib/Algebra/Ring/SumsOfSquares.lean new file mode 100644 index 0000000000000..0d555aa10890b --- /dev/null +++ b/Mathlib/Algebra/Ring/SumsOfSquares.lean @@ -0,0 +1,101 @@ +/- +Copyright (c) 2024 Florent Schaffhauser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Florent Schaffhauser +-/ + +import Mathlib.Algebra.Ring.Defs +import Mathlib.Algebra.Group.Submonoid.Basic +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Order.Ring.Defs + +/-! +# Sums of squares + +We introduce sums of squares in a type `R` endowed with an `[Add R]`, `[Zero R]` and `[Mul R]` +instances. Sums of squares in `R` are defined by an inductive predicate `isSumSq : R → Prop`: +`0 : R` is a sum of squares and if `S` is a sum of squares, then for all `a : R`, `a * a + S` is a +sum of squares in `R`. + +## Main declaration + +- The predicate `isSumSq : R → Prop`, defining the property of being a sum of squares in `R`. + +## Auxiliary declarations + +- The type `SumSqIn R` : in an additive monoid with multiplication `R`, we introduce the type +`SumSqIn R` as the submonoid of `R` whose carrier is the subset `{S : R | isSumSq S}`. + +## Theorems + +- `isSumSq.add`: if `S1` and `S2` are sums of squares in `R`, then so is `S1 + S2`. +- `SumSq.nonneg`: in a linearly ordered semiring `R` with an `[ExistsAddOfLE R]` instance, sums of +squares are non-negative. +- `SquaresInSumSq`: a square is a sum of squares. +- `SquaresAddClosure`: the submonoid of `R` generated by the squares in `R` is the set of sums of +squares in `R`. + +-/ + +variable {R : Type*} [Mul R] + +/-- +In a type `R` with an addition, a zero element and a multiplication, the property of being a sum of +squares is defined by an inductive predicate: `0 : R` is a sum of squares and if `S` is a sum of +squares, then for all `a : R`, `a * a + S` is a sum of squares in `R`. +-/ +@[mk_iff] +inductive isSumSq [Add R] [Zero R] : R → Prop + | zero : isSumSq 0 + | sq_add (a S : R) (pS : isSumSq S) : isSumSq (a * a + S) + +/-- +If `S1` and `S2` are sums of squares in a semiring `R`, then `S1 + S2` is a sum of squares in `R`. +-/ +theorem isSumSq.add [AddMonoid R] {S1 S2 : R} (p1 : isSumSq S1) + (p2 : isSumSq S2) : isSumSq (S1 + S2) := by + induction p1 with + | zero => rw [zero_add]; exact p2 + | sq_add a S pS ih => rw [add_assoc]; exact isSumSq.sq_add a (S + S2) ih + +variable (R) in +/-- +In an additive monoid with multiplication `R`, the type `SumSqIn R` is the submonoid of sums of +squares in `R`. +-/ +def SumSqIn [AddMonoid R] : AddSubmonoid R where + carrier := {S : R | isSumSq S} + zero_mem' := isSumSq.zero + add_mem' := isSumSq.add + +/-- +In an additive monoid with multiplication, every square is a sum of squares. By definition, a square +in `R` is a term `x : R` such that `x = y * y` for some `y : R` and in Mathlib this is known as +`IsSquare R` (see Mathlib.Algebra.Group.Even). +-/ +theorem SquaresInSumSq [AddMonoid R] {x : R} (px : IsSquare x) : x ∈ SumSqIn R := by + rcases px with ⟨y, py⟩ + rw [py, ← AddMonoid.add_zero (y * y)] + exact isSumSq.sq_add _ _ isSumSq.zero + +open AddSubmonoid in +/-- +In an additive monoid with multiplication `R`, the submonoid generated by the squares is the set of +sums of squares. +-/ +theorem SquaresAddClosure [AddMonoid R] : closure IsSquare = SumSqIn R := by + refine le_antisymm (closure_le.2 (fun x hx ↦ SquaresInSumSq hx)) (fun x hx ↦ ?_) + induction hx with + | zero => exact zero_mem _ + | sq_add a S _ ih => exact AddSubmonoid.add_mem _ (subset_closure ⟨a, rfl⟩) ih + +/-- +Let `R` be a linearly ordered semiring in which the property `a ≤ b → ∃ c, a + c = b` holds +(e.g. `R = ℕ`). If `S : R` is a sum of squares in `R`, then `0 ≤ S`. This is used in +`Mathlib.Algebra.Ring.Semireal.Defs` to show that linearly ordered fields are semireal. +-/ +theorem isSumSq.nonneg {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {S : R} + (pS : isSumSq S) : 0 ≤ S := by + induction pS with + | zero => simp only [le_refl] + | sq_add x S _ ih => apply add_nonneg ?_ ih; simp only [← pow_two x, sq_nonneg] diff --git a/Mathlib/Algebra/Star/CHSH.lean b/Mathlib/Algebra/Star/CHSH.lean index f29d176ce3e65..6adb8f4bf21a8 100644 --- a/Mathlib/Algebra/Star/CHSH.lean +++ b/Mathlib/Algebra/Star/CHSH.lean @@ -176,7 +176,7 @@ theorem tsirelson_inequality [OrderedRing R] [StarRing R] [StarOrderedRing R] [A A₀ * B₀ + A₀ * B₁ + A₁ * B₀ - A₁ * B₁ ≤ √2 ^ 3 • (1 : R) := by -- abel will create `ℤ` multiplication. We will `simp` them away to `ℝ` multiplication. have M : ∀ (m : ℤ) (a : ℝ) (x : R), m • a • x = ((m : ℝ) * a) • x := fun m a x => by - rw [zsmul_eq_smul_cast ℝ, ← mul_smul] + rw [← Int.cast_smul_eq_nsmul ℝ, ← mul_smul] let P := (√2)⁻¹ • (A₁ + A₀) - B₀ let Q := (√2)⁻¹ • (A₁ - A₀) + B₁ have w : √2 ^ 3 • (1 : R) - A₀ * B₀ - A₀ * B₁ - A₁ * B₀ + A₁ * B₁ = (√2)⁻¹ • (P ^ 2 + Q ^ 2) := by diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index fdc9184e81bcf..56a02b6d6dfb7 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -967,20 +967,17 @@ noncomputable def iSupLift [Nonempty ι] (K : ι → NonUnitalStarSubalgebra R A inclusion_mk, Eq.ndrec, id_eq, eq_mpr_eq_cast, ZeroMemClass.coe_zero, AddSubmonoid.mk_add_mk, Set.inclusion_mk] apply Set.iUnionLift_binary (coe_iSup_of_directed dir) dir _ (fun _ => (· * ·)) - on_goal 3 => rw [coe_iSup_of_directed dir] all_goals simp map_add' := by dsimp only [SetLike.coe_sort_coe, NonUnitalAlgHom.coe_comp, Function.comp_apply, inclusion_mk, Eq.ndrec, id_eq, eq_mpr_eq_cast] apply Set.iUnionLift_binary (coe_iSup_of_directed dir) dir _ (fun _ => (· + ·)) - on_goal 3 => rw [coe_iSup_of_directed dir] all_goals simp map_smul' := fun r => by dsimp only [SetLike.coe_sort_coe, NonUnitalAlgHom.coe_comp, Function.comp_apply, inclusion_mk, Eq.ndrec, id_eq, eq_mpr_eq_cast] apply Set.iUnionLift_unary (coe_iSup_of_directed dir) _ (fun _ x => r • x) (fun _ _ => rfl) - on_goal 2 => rw [coe_iSup_of_directed dir] all_goals simp map_star' := by dsimp only [SetLike.coe_sort_coe, NonUnitalStarAlgHom.comp_apply, inclusion_mk, Eq.ndrec, @@ -990,7 +987,6 @@ noncomputable def iSupLift [Nonempty ι] (K : ι → NonUnitalStarSubalgebra R A NonUnitalAlgHom.coe_mk] apply Set.iUnionLift_unary (coe_iSup_of_directed dir) _ (fun _ x => star x) (fun _ _ => rfl) - on_goal 2 => rw [coe_iSup_of_directed dir] all_goals simp [map_star] } variable [Nonempty ι] {K : ι → NonUnitalStarSubalgebra R A} {dir : Directed (· ≤ ·) K} diff --git a/Mathlib/Algebra/Star/Order.lean b/Mathlib/Algebra/Star/Order.lean index 2d79cdea834ed..3d1200121de45 100644 --- a/Mathlib/Algebra/Star/Order.lean +++ b/Mathlib/Algebra/Star/Order.lean @@ -247,6 +247,9 @@ end NonUnitalSemiring section Semiring variable [Semiring R] [PartialOrder R] [StarRing R] [StarOrderedRing R] +instance : ZeroLEOneClass R where + zero_le_one := by simpa using star_mul_self_nonneg (1 : R) + @[simp] lemma one_le_star_iff {x : R} : 1 ≤ star x ↔ 1 ≤ x := by simpa using star_le_star_iff (x := 1) (y := x) diff --git a/Mathlib/Algebra/Star/RingQuot.lean b/Mathlib/Algebra/Star/RingQuot.lean index 1a992bd1528a6..c8633430a1659 100644 --- a/Mathlib/Algebra/Star/RingQuot.lean +++ b/Mathlib/Algebra/Star/RingQuot.lean @@ -18,9 +18,10 @@ variable {R : Type u} [Semiring R] (r : R → R → Prop) section StarRing -variable [StarRing R] (hr : ∀ a b, r a b → r (star a) (star b)) +variable [StarRing R] -theorem Rel.star ⦃a b : R⦄ (h : Rel r a b) : Rel r (star a) (star b) := by +theorem Rel.star (hr : ∀ a b, r a b → r (star a) (star b)) + ⦃a b : R⦄ (h : Rel r a b) : Rel r (star a) (star b) := by induction h with | of h => exact Rel.of (hr _ _ h) | add_left _ h => rw [star_add, star_add] @@ -30,7 +31,7 @@ theorem Rel.star ⦃a b : R⦄ (h : Rel r a b) : Rel r (star a) (star b) := by | mul_right _ h => rw [star_mul, star_mul] exact Rel.mul_left h -private irreducible_def star' : RingQuot r → RingQuot r +private irreducible_def star' (hr : ∀ a b, r a b → r (star a) (star b)) : RingQuot r → RingQuot r | ⟨a⟩ => ⟨Quot.map (star : R → R) (Rel.star r hr) a⟩ theorem star'_quot (hr : ∀ a b, r a b → r (star a) (star b)) {a} : diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 69be7add4a50c..bb80b87af6d6a 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -33,10 +33,6 @@ We also define predicates about affine schemes and affine open sets. -- Explicit universe annotations were used in this file to improve perfomance #12737 - -set_option allowUnsafeReducibility true in -attribute [local reducible] Quiver.Hom - noncomputable section open CategoryTheory CategoryTheory.Limits Opposite TopologicalSpace @@ -179,15 +175,14 @@ noncomputable instance forgetToScheme_preservesLimits : PreservesLimits forgetTo end AffineScheme /-- An open subset of a scheme is affine if the open subscheme is affine. -/ -def IsAffineOpen {X : Scheme} (U : Opens X) : Prop := - IsAffine (X ∣_ᵤ U) +def IsAffineOpen {X : Scheme} (U : X.Opens) : Prop := + IsAffine U /-- The set of affine opens as a subset of `opens X`. -/ -def Scheme.affineOpens (X : Scheme) : Set (Opens X) := - {U : Opens X | IsAffineOpen U} +def Scheme.affineOpens (X : Scheme) : Set X.Opens := + {U : X.Opens | IsAffineOpen U} -instance {Y : Scheme.{u}} (U : Y.affineOpens) : - IsAffine (Scheme.restrict Y <| Opens.openEmbedding U.val) := +instance {Y : Scheme.{u}} (U : Y.affineOpens) : IsAffine U := U.property theorem isAffineOpen_opensRange {X Y : Scheme} [IsAffine X] (f : X ⟶ Y) @@ -195,7 +190,7 @@ theorem isAffineOpen_opensRange {X Y : Scheme} [IsAffine X] (f : X ⟶ Y) refine isAffine_of_isIso (IsOpenImmersion.isoOfRangeEq f (Y.ofRestrict _) ?_).inv exact Subtype.range_val.symm -theorem isAffineOpen_top (X : Scheme) [IsAffine X] : IsAffineOpen (⊤ : Opens X) := by +theorem isAffineOpen_top (X : Scheme) [IsAffine X] : IsAffineOpen (⊤ : X.Opens) := by convert isAffineOpen_opensRange (𝟙 X) ext1 exact Set.range_id.symm @@ -223,7 +218,7 @@ theorem isBasis_affine_open (X : Scheme) : Opens.IsBasis X.affineOpens := by rcases hS with ⟨i, rfl⟩ exact isAffineOpen_opensRange _ -theorem iSup_affineOpens_eq_top (X : Scheme) : ⨆ i : X.affineOpens, (i : Opens X) = ⊤ := by +theorem iSup_affineOpens_eq_top (X : Scheme) : ⨆ i : X.affineOpens, (i : X.Opens) = ⊤ := by apply Opens.ext rw [Opens.coe_iSup] apply IsTopologicalBasis.sUnion_eq @@ -236,7 +231,7 @@ theorem Scheme.map_PrimeSpectrum_basicOpen_of_affine ΓSpec.adjunction_unit_map_basicOpen _ _ theorem isBasis_basicOpen (X : Scheme) [IsAffine X] : - Opens.IsBasis (Set.range (X.basicOpen : Γ(X, ⊤) → Opens X)) := by + Opens.IsBasis (Set.range (X.basicOpen : Γ(X, ⊤) → X.Opens)) := by delta Opens.IsBasis convert PrimeSpectrum.isBasis_basic_opens.inducing (TopCat.homeoOfIso (Scheme.forgetToTop.mapIso X.isoSpec)).inducing using 1 @@ -252,14 +247,14 @@ theorem isBasis_basicOpen (X : Scheme) [IsAffine X] : namespace IsAffineOpen -variable {X Y : Scheme.{u}} {U : Opens X} (hU : IsAffineOpen U) (f : Γ(X, U)) +variable {X Y : Scheme.{u}} {U : X.Opens} (hU : IsAffineOpen U) (f : Γ(X, U)) /-- The open immersion `Spec Γ(X, U) ⟶ X` for an affine `U`. -/ def fromSpec : - (Spec Γ(X, U)) ⟶ X := - haveI : IsAffine (X ∣_ᵤ U) := hU + Spec Γ(X, U) ⟶ X := + haveI : IsAffine U := hU Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op) ≫ - (X ∣_ᵤ U).isoSpec.inv ≫ Scheme.ιOpens U + U.toScheme.isoSpec.inv ≫ U.ι instance isOpenImmersion_fromSpec : IsOpenImmersion hU.fromSpec := by @@ -279,7 +274,7 @@ theorem range_fromSpec : theorem opensRange_fromSpec : Scheme.Hom.opensRange hU.fromSpec = U := Opens.ext (range_fromSpec hU) @[reassoc (attr := simp)] -theorem map_fromSpec {V : Opens X} (hV : IsAffineOpen V) (f : op U ⟶ op V) : +theorem map_fromSpec {V : X.Opens} (hV : IsAffineOpen V) (f : op U ⟶ op V) : Spec.map (X.presheaf.map f) ≫ hU.fromSpec = hV.fromSpec := by have : IsAffine (X.restrictFunctor.obj U).left := hU haveI : IsAffine _ := hV @@ -299,17 +294,17 @@ protected theorem isCompact : theorem image_of_isOpenImmersion (f : X ⟶ Y) [H : IsOpenImmersion f] : IsAffineOpen (f ''ᵁ U) := by have : IsAffine _ := hU - convert isAffineOpen_opensRange (X.ofRestrict U.openEmbedding ≫ f) + convert isAffineOpen_opensRange (U.ι ≫ f) ext1 exact Set.image_eq_range _ _ -theorem preimage_of_isIso {U : Opens Y.carrier} (hU : IsAffineOpen U) (f : X ⟶ Y) [IsIso f] : +theorem preimage_of_isIso {U : Y.Opens} (hU : IsAffineOpen U) (f : X ⟶ Y) [IsIso f] : IsAffineOpen (f ⁻¹ᵁ U) := haveI : IsAffine _ := hU isAffine_of_isIso (f ∣_ U) theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion - (f : AlgebraicGeometry.Scheme.Hom X Y) [H : IsOpenImmersion f] {U : Opens X} : + (f : AlgebraicGeometry.Scheme.Hom X Y) [H : IsOpenImmersion f] {U : X.Opens} : IsAffineOpen (f ''ᵁ U) ↔ IsAffineOpen U := by refine ⟨fun hU => @isAffine_of_isIso _ _ (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.openEmbedding ≫ f) (Y.ofRestrict _) ?_).hom ?_ hU, @@ -334,14 +329,14 @@ def _root_.AlgebraicGeometry.IsOpenImmersion.affineOpensEquiv (f : X ⟶ Y) [H : /-- The affine open sets of an open subscheme corresponds to the affine open sets containing in the subset. -/ @[simps! apply_coe_coe] -def _root_.AlgebraicGeometry.affineOpensRestrict {X : Scheme.{u}} (U : Opens X) : - (X ∣_ᵤ U).affineOpens ≃ { V : X.affineOpens // V ≤ U } := - (IsOpenImmersion.affineOpensEquiv (Scheme.ιOpens U)).trans (Equiv.subtypeEquivProp (by simp)) +def _root_.AlgebraicGeometry.affineOpensRestrict {X : Scheme.{u}} (U : X.Opens) : + U.toScheme.affineOpens ≃ { V : X.affineOpens // V ≤ U } := + (IsOpenImmersion.affineOpensEquiv U.ι).trans (Equiv.subtypeEquivProp (by simp)) @[simp] lemma _root_.AlgebraicGeometry.affineOpensRestrict_symm_apply_coe - {X : Scheme.{u}} (U : Opens X) (V) : - ((affineOpensRestrict U).symm V).1 = Scheme.ιOpens U ⁻¹ᵁ V := rfl + {X : Scheme.{u}} (U : X.Opens) (V) : + ((affineOpensRestrict U).symm V).1 = U.ι ⁻¹ᵁ V := rfl instance (priority := 100) _root_.AlgebraicGeometry.Scheme.compactSpace_of_isAffine (X : Scheme) [IsAffine X] : @@ -355,22 +350,15 @@ theorem fromSpec_preimage_self : rw [Opens.map_coe, Opens.coe_top, ← hU.range_fromSpec, ← Set.image_univ] exact Set.preimage_image_eq _ PresheafedSpace.IsOpenImmersion.base_open.inj -#adaptation_note /-- 2024-04-23 -The backwards compatibility flags don't help here. -/ -set_option maxHeartbeats 400000 in --- Doesn't build without the `IsAffine` instance but the linter complains -@[nolint unusedHavesSuffices] theorem SpecΓIdentity_hom_app_fromSpec : (Scheme.ΓSpecIso Γ(X, U)).hom ≫ hU.fromSpec.app U = (Spec Γ(X, U)).presheaf.map (eqToHom hU.fromSpec_preimage_self).op := by - have : IsAffine _ := hU - delta IsAffineOpen.fromSpec Scheme.isoSpec - rw [Scheme.comp_app, Scheme.comp_app, ΓSpecIso_obj_hom, - Scheme.ofRestrict_app_self] - simp only [Category.assoc] - dsimp only [asIso_inv, Functor.op_obj, unop_op] - rw [← Functor.map_comp_assoc, ← op_comp, eqToHom_trans, Scheme.eq_restrict_presheaf_map_eqToHom, - Scheme.Hom.naturality_assoc, Scheme.inv_app_top, IsIso.hom_inv_id_assoc] + simp only [fromSpec, Scheme.isoSpec, asIso_inv, Scheme.comp_coeBase, Opens.map_comp_obj, + ΓSpecIso_obj_hom, Scheme.Opens.topIso_inv, Opens.map_top, Functor.id_obj, Functor.comp_obj, + Functor.rightOp_obj, Scheme.Γ_obj, unop_op, Scheme.Spec_obj, Scheme.Opens.topIso_hom, + Scheme.comp_app, Scheme.Opens.ι_app_self, Category.assoc, ← Functor.map_comp_assoc, ← op_comp, + eqToHom_trans, Scheme.Opens.eq_presheaf_map_eqToHom, Scheme.Hom.naturality_assoc, + Scheme.inv_app_top, IsIso.hom_inv_id_assoc] simp only [eqToHom_op, eqToHom_map, Spec.map_eqToHom, eqToHom_unop, Scheme.Spec_map_presheaf_map_eqToHom, eqToHom_trans] @@ -409,29 +397,29 @@ theorem basicOpen : (Spec.map (CommRingCat.ofHom <| algebraMap Γ(X, U) (Localization.Away f))) exact Opens.ext (PrimeSpectrum.localization_away_comap_range (Localization.Away f) f).symm -instance [IsAffine X] (r : Γ(X, ⊤)) : IsAffine (X ∣_ᵤ X.basicOpen r) := +instance [IsAffine X] (r : Γ(X, ⊤)) : IsAffine (X.basicOpen r) := (isAffineOpen_top X).basicOpen _ -theorem ιOpens_basicOpen_preimage (r : Γ(X, ⊤)) : - IsAffineOpen (Scheme.ιOpens (X.basicOpen r) ⁻¹ᵁ U) := by - apply (Scheme.ιOpens (X.basicOpen r)).isAffineOpen_iff_of_isOpenImmersion.mp +theorem ι_basicOpen_preimage (r : Γ(X, ⊤)) : + IsAffineOpen ((X.basicOpen r).ι ⁻¹ᵁ U) := by + apply (X.basicOpen r).ι.isAffineOpen_iff_of_isOpenImmersion.mp dsimp [Scheme.Hom.opensFunctor, LocallyRingedSpace.IsOpenImmersion.opensFunctor] rw [Opens.functor_obj_map_obj, Opens.openEmbedding_obj_top, inf_comm, ← Scheme.basicOpen_res _ _ (homOfLE le_top).op] exact hU.basicOpen _ -theorem exists_basicOpen_le {V : Opens X} (x : V) (h : ↑x ∈ U) : +theorem exists_basicOpen_le {V : X.Opens} (x : V) (h : ↑x ∈ U) : ∃ f : Γ(X, U), X.basicOpen f ≤ V ∧ ↑x ∈ X.basicOpen f := by have : IsAffine _ := hU obtain ⟨_, ⟨_, ⟨r, rfl⟩, rfl⟩, h₁, h₂⟩ := - (isBasis_basicOpen (X ∣_ᵤ U)).exists_subset_of_mem_open (x.2 : ⟨x, h⟩ ∈ _) + (isBasis_basicOpen U).exists_subset_of_mem_open (x.2 : ⟨x, h⟩ ∈ _) ((Opens.map U.inclusion).obj V).isOpen have : - Scheme.ιOpens U ''ᵁ ((X ∣_ᵤ U).basicOpen r) = + U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r) := by - refine (Scheme.image_basicOpen (Scheme.ιOpens U) r).trans ?_ + refine (Scheme.image_basicOpen U.ι r).trans ?_ rw [Scheme.basicOpen_res_eq] - simp only [Scheme.restrict_presheaf_obj, Scheme.ofRestrict_appIso, Iso.refl_inv, + simp only [Scheme.Opens.toScheme_presheaf_obj, Scheme.Opens.ι_appIso, Iso.refl_inv, CommRingCat.id_apply] use X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r rw [← this] @@ -475,8 +463,8 @@ instance _root_.AlgebraicGeometry.isLocalization_away_of_isAffine IsLocalization.Away r Γ(X, X.basicOpen r) := isLocalization_basicOpen (isAffineOpen_top X) r -lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Opens Y} (hU : IsAffineOpen U) - {V : Opens X} (hV : IsAffineOpen V) (e) (r : Γ(Y, U)) : +lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (hU : IsAffineOpen U) + {V : X.Opens} (hV : IsAffineOpen V) (e) (r : Γ(Y, U)) : letI := hU.isLocalization_basicOpen r letI := hV.isLocalization_basicOpen (f.appLE U V e r) f.appLE (Y.basicOpen r) (X.basicOpen (f.appLE U V e r)) @@ -489,13 +477,13 @@ lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Opens Y} (hU : IsA RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] -theorem isLocalization_of_eq_basicOpen {V : Opens X} (i : V ⟶ U) (e : V = X.basicOpen f) : +theorem isLocalization_of_eq_basicOpen {V : X.Opens} (i : V ⟶ U) (e : V = X.basicOpen f) : @IsLocalization.Away _ _ f Γ(X, V) _ (X.presheaf.map i.op).toAlgebra := by subst e; convert isLocalization_basicOpen hU f using 3 instance _root_.AlgebraicGeometry.Γ_restrict_isLocalization - (X : Scheme.{u}) [IsAffine X] (r : Scheme.Γ.obj (op X)) : - IsLocalization.Away r (Scheme.Γ.obj (op (X ∣_ᵤ X.basicOpen r))) := + (X : Scheme.{u}) [IsAffine X] (r : Γ(X, ⊤)) : + IsLocalization.Away r Γ(X.basicOpen r, ⊤) := (isAffineOpen_top X).isLocalization_of_eq_basicOpen r _ (Opens.openEmbedding_obj_top _) theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : @@ -507,7 +495,7 @@ theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : rw [Scheme.basicOpen_res] refine (inf_eq_left.mpr ?_).symm -- Porting note: a little help is needed here - convert inf_le_left (α := Opens X) using 1 + convert inf_le_left (α := X.Opens) using 1 apply Scheme.basicOpen_of_isUnit apply Submonoid.leftInv_le_isUnit _ @@ -515,7 +503,7 @@ theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : _).prop theorem _root_.AlgebraicGeometry.exists_basicOpen_le_affine_inter - {V : Opens X} (hV : IsAffineOpen V) (x : X) (hx : x ∈ U ⊓ V) : + {V : X.Opens} (hV : IsAffineOpen V) (x : X) (hx : x ∈ U ⊓ V) : ∃ (f : Γ(X, U)) (g : Γ(X, V)), X.basicOpen f = X.basicOpen g ∧ x ∈ X.basicOpen f := by obtain ⟨f, hf₁, hf₂⟩ := hU.exists_basicOpen_le ⟨x, hx.2⟩ hx.1 obtain ⟨g, hg₁, hg₂⟩ := hV.exists_basicOpen_le ⟨x, hf₂⟩ hx.2 @@ -527,7 +515,7 @@ theorem _root_.AlgebraicGeometry.exists_basicOpen_le_affine_inter /-- The prime ideal of `𝒪ₓ(U)` corresponding to a point `x : U`. -/ noncomputable def primeIdealOf (x : U) : PrimeSpectrum Γ(X, U) := - ((@Scheme.isoSpec (X ∣_ᵤ U) hU).hom ≫ + ((@Scheme.isoSpec U hU).hom ≫ Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top).op)).1.base x theorem fromSpec_primeIdealOf (x : U) : @@ -537,10 +525,9 @@ theorem fromSpec_primeIdealOf (x : U) : -- unnecessary, indeed, the linter did not like it, so I just use `elementwise_of%` instead of -- adding the corresponding lemma in `Scheme.lean` file erw [← elementwise_of% Scheme.comp_val_base] -- now `erw` after #13170 - simp only [Scheme.restrict_carrier, Scheme.restrict_presheaf_obj, unop_op, Category.assoc, ← - Spec.map_comp_assoc, ← Functor.map_comp, ← op_comp, eqToHom_trans, eqToHom_refl, op_id, - CategoryTheory.Functor.map_id, Spec.map_id, Category.id_comp, Iso.hom_inv_id_assoc, - Scheme.ofRestrict_val_base] + simp only [Scheme.Opens.toScheme_presheaf_obj, Category.assoc, ← Spec.map_comp_assoc, + ← Functor.map_comp, ← op_comp, eqToHom_trans, eqToHom_refl, op_id, + CategoryTheory.Functor.map_id, Spec.map_id, Category.id_comp, Iso.hom_inv_id_assoc] rfl -- `rfl` was not needed before #13170 theorem isLocalization_stalk' (y : PrimeSpectrum Γ(X, U)) (hy : hU.fromSpec.1.base y ∈ U) : @@ -580,6 +567,10 @@ def _root_.AlgebraicGeometry.Scheme.affineBasicOpen (X : Scheme) {U : X.affineOpens} (f : Γ(X, U)) : X.affineOpens := ⟨X.basicOpen f, U.prop.basicOpen f⟩ +/-- +In an affine open set `U`, a family of basic open covers `U` iff the sections span `Γ(X, U)`. +See `iSup_basicOpen_of_span_eq_top` for the inverse direction without the affine-ness assuption. +-/ theorem basicOpen_union_eq_self_iff (s : Set Γ(X, U)) : ⨆ f : s, X.basicOpen (f : Γ(X, U)) = U ↔ Ideal.span s = ⊤ := by trans ⋃ i : s, (PrimeSpectrum.basicOpen i.1).1 = Set.univ @@ -618,6 +609,26 @@ theorem self_le_basicOpen_union_iff (s : Set Γ(X, U)) : end IsAffineOpen +/-- +Given a spanning set of `Γ(X, U)`, the corresponding basic open sets cover `U`. +See `IsAffineOpen.basicOpen_union_eq_self_iff` for the inverse direction for affine open sets. +-/ +lemma iSup_basicOpen_of_span_eq_top {X : Scheme} (U) (s : Set Γ(X, U)) + (hs : Ideal.span s = ⊤) : (⨆ i ∈ s, X.basicOpen i) = U := by + apply le_antisymm + · rw [iSup₂_le_iff] + exact fun i _ ↦ X.basicOpen_le i + · intro x hx + obtain ⟨_, ⟨V, hV, rfl⟩, hxV, hVU⟩ := (isBasis_affine_open X).exists_subset_of_mem_open hx U.2 + refine SetLike.mem_of_subset ?_ hxV + rw [← (hV.basicOpen_union_eq_self_iff (X.presheaf.map (homOfLE hVU).op '' s)).mpr + (by rw [← Ideal.map_span, hs, Ideal.map_top])] + simp only [Opens.iSup_mk, Opens.carrier_eq_coe, Set.iUnion_coe_set, Set.mem_image, + Set.iUnion_exists, Set.biUnion_and', Set.iUnion_iUnion_eq_right, Scheme.basicOpen_res, + Opens.coe_inf, Opens.coe_mk, Set.iUnion_subset_iff] + exact fun i hi ↦ (Set.inter_subset_right.trans + (Set.subset_iUnion₂ (s := fun x _ ↦ (X.basicOpen x : Set X)) i hi)) + /-- Let `P` be a predicate on the affine open sets of `X` satisfying 1. If `P` holds on `U`, then `P` holds on the basic open set of every section on `U`. 2. If `P` holds for a family of basic open sets covering `U`, then `P` holds for `U`. @@ -628,7 +639,7 @@ Then `P` holds for every affine open of `X`. This is also known as the **Affine communication lemma** in [*The rising sea*][RisingSea]. -/ @[elab_as_elim] theorem of_affine_open_cover {X : Scheme} {P : X.affineOpens → Prop} - {ι} (U : ι → X.affineOpens) (iSup_U : (⨆ i, U i : Opens X) = ⊤) + {ι} (U : ι → X.affineOpens) (iSup_U : (⨆ i, U i : X.Opens) = ⊤) (V : X.affineOpens) (basicOpen : ∀ (U : X.affineOpens) (f : Γ(X, U)), P U → P (X.affineBasicOpen f)) (openCover : @@ -639,7 +650,7 @@ theorem of_affine_open_cover {X : Scheme} {P : X.affineOpens → Prop} classical have : ∀ (x : V.1), ∃ f : Γ(X, V), ↑x ∈ X.basicOpen f ∧ P (X.affineBasicOpen f) := by intro x - obtain ⟨i, hi⟩ := Opens.mem_iSup.mp (show x.1 ∈ (⨆ i, U i : Opens X) from iSup_U ▸ trivial) + obtain ⟨i, hi⟩ := Opens.mem_iSup.mp (show x.1 ∈ (⨆ i, U i : X.Opens) from iSup_U ▸ trivial) obtain ⟨f, g, e, hf⟩ := exists_basicOpen_le_affine_inter V.prop (U i).prop x ⟨x.prop, hi⟩ refine ⟨f, hf, ?_⟩ convert basicOpen _ g (hU i) using 1 @@ -669,7 +680,8 @@ open ConcreteCategory /-- If `X` is affine, the image of the zero locus of global sections of `X` under `toΓSpecFun` 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] (s : Set Γ(X, ⊤)) : +lemma Scheme.toΓSpec_image_zeroLocus_eq_of_isAffine {X : Scheme.{u}} [IsAffine X] + (s : Set Γ(X, ⊤)) : X.isoSpec.hom.val.base '' X.zeroLocus s = PrimeSpectrum.zeroLocus s := by erw [← X.toΓSpec_preimage_zeroLocus_eq, Set.image_preimage_eq] exact (bijective_of_isIso X.isoSpec.hom.val.base).surjective @@ -721,6 +733,6 @@ alias IsAffineOpen.opensFunctor_map_basicOpen := IsAffineOpen.fromSpec_image_bas @[deprecated (since := "2024-06-21")] alias IsAffineOpen.basicOpenIsAffine := IsAffineOpen.basicOpen @[deprecated (since := "2024-06-21")] -alias IsAffineOpen.mapRestrictBasicOpen := IsAffineOpen.ιOpens_basicOpen_preimage +alias IsAffineOpen.mapRestrictBasicOpen := IsAffineOpen.ι_basicOpen_preimage end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Cover/Open.lean b/Mathlib/AlgebraicGeometry/Cover/Open.lean index 693680f4a4b19..20acd2442916a 100644 --- a/Mathlib/AlgebraicGeometry/Cover/Open.lean +++ b/Mathlib/AlgebraicGeometry/Cover/Open.lean @@ -263,21 +263,6 @@ def OpenCover.inter {X : Scheme.{u}} (𝒰₁ : Scheme.OpenCover.{v₁} X) exact ⟨𝒰₁.covers x, 𝒰₂.covers x⟩ IsOpen x := inferInstance -/-- If `U` is a family of open sets that covers `X`, then `X.restrict U` forms an `X.open_cover`. -/ -@[simps! J obj map] -def openCoverOfSuprEqTop {s : Type*} (X : Scheme.{u}) (U : s → Opens X) - (hU : ⨆ i, U i = ⊤) : X.OpenCover where - J := s - obj i := X.restrict (U i).openEmbedding - map i := X.ofRestrict (U i).openEmbedding - f x := - haveI : x ∈ ⨆ i, U i := hU.symm ▸ show x ∈ (⊤ : Opens X) by trivial - (Opens.mem_iSup.mp this).choose - covers x := by - erw [Subtype.range_coe] - have : x ∈ ⨆ i, U i := hU.symm ▸ show x ∈ (⊤ : Opens X) by trivial - exact (Opens.mem_iSup.mp this).choose_spec - /-- An affine open cover of `X` consists of a family of open immersions into `X` from spectra of rings. @@ -428,7 +413,7 @@ def OpenCover.fromAffineRefinement {X : Scheme.{u}} (𝓤 : X.OpenCover) : /-- If two global sections agree after restriction to each member of an open cover, then they agree globally. -/ -lemma OpenCover.ext_elem {X : Scheme.{u}} {U : Opens X} (f g : Γ(X, U)) (𝒰 : X.OpenCover) +lemma OpenCover.ext_elem {X : Scheme.{u}} {U : X.Opens} (f g : Γ(X, U)) (𝒰 : X.OpenCover) (h : ∀ i : 𝒰.J, (𝒰.map i).app U f = (𝒰.map i).app U g) : f = g := by fapply TopCat.Sheaf.eq_of_locally_eq' X.sheaf (fun i ↦ (𝒰.map (𝒰.f i)).opensRange ⊓ U) _ (fun _ ↦ homOfLE inf_le_right) @@ -444,13 +429,13 @@ lemma OpenCover.ext_elem {X : Scheme.{u}} {U : Opens X} (f g : Γ(X, U)) (𝒰 : /-- If the restriction of a global section to each member of an open cover is zero, then it is globally zero. -/ -lemma zero_of_zero_cover {X : Scheme.{u}} {U : Opens X} (s : Γ(X, U)) (𝒰 : X.OpenCover) +lemma zero_of_zero_cover {X : Scheme.{u}} {U : X.Opens} (s : Γ(X, U)) (𝒰 : X.OpenCover) (h : ∀ i : 𝒰.J, (𝒰.map i).app U s = 0) : s = 0 := 𝒰.ext_elem s 0 (fun i ↦ by rw [map_zero]; exact h i) /-- If a global section is nilpotent on each member of a finite open cover, then `f` is nilpotent. -/ -lemma isNilpotent_of_isNilpotent_cover {X : Scheme.{u}} {U : Opens X} (s : Γ(X, U)) +lemma isNilpotent_of_isNilpotent_cover {X : Scheme.{u}} {U : X.Opens} (s : Γ(X, U)) (𝒰 : X.OpenCover) [Finite 𝒰.J] (h : ∀ i : 𝒰.J, IsNilpotent ((𝒰.map i).app U s)) : IsNilpotent s := by choose fn hfn using h diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index ec029dad172a5..91aaaac6fed72 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -34,13 +34,13 @@ noncomputable abbrev Scheme.functionField [IrreducibleSpace X] : CommRingCat := X.presheaf.stalk (genericPoint X) /-- The restriction map from a component to the function field. -/ -noncomputable abbrev Scheme.germToFunctionField [IrreducibleSpace X] (U : Opens X) +noncomputable abbrev Scheme.germToFunctionField [IrreducibleSpace X] (U : X.Opens) [h : Nonempty U] : Γ(X, U) ⟶ X.functionField := X.presheaf.germ ⟨genericPoint X, ((genericPoint_spec X).mem_open_set_iff U.isOpen).mpr (by simpa using h)⟩ -noncomputable instance [IrreducibleSpace X] (U : Opens X) [Nonempty U] : +noncomputable instance [IrreducibleSpace X] (U : X.Opens) [Nonempty U] : Algebra Γ(X, U) X.functionField := (X.germToFunctionField U).toAlgebra @@ -59,7 +59,7 @@ noncomputable instance [IsIntegral X] : Field X.functionField := by have := (X.presheaf.germ ⟨_, hs⟩).isUnit_map (RingedSpace.isUnit_res_basicOpen _ s) rwa [TopCat.Presheaf.germ_res_apply] at this -theorem germ_injective_of_isIntegral [IsIntegral X] {U : Opens X} (x : U) : +theorem germ_injective_of_isIntegral [IsIntegral X] {U : X.Opens} (x : U) : Function.Injective (X.presheaf.germ x) := by rw [injective_iff_map_eq_zero] intro y hy @@ -69,7 +69,7 @@ theorem germ_injective_of_isIntegral [IsIntegral X] {U : Opens X} (x : U) : haveI : Nonempty W := ⟨⟨_, hW⟩⟩ exact map_injective_of_isIntegral X iU e -theorem Scheme.germToFunctionField_injective [IsIntegral X] (U : Opens X) [Nonempty U] : +theorem Scheme.germToFunctionField_injective [IsIntegral X] (U : X.Opens) [Nonempty U] : Function.Injective (X.germToFunctionField U) := germ_injective_of_isIntegral _ _ @@ -90,7 +90,7 @@ noncomputable instance stalkFunctionFieldAlgebra [IrreducibleSpace X] (x : X) : apply RingHom.toAlgebra exact X.presheaf.stalkSpecializes ((genericPoint_spec X).specializes trivial) -instance functionField_isScalarTower [IrreducibleSpace X] (U : Opens X) (x : U) +instance functionField_isScalarTower [IrreducibleSpace X] (U : X.Opens) (x : U) [Nonempty U] : IsScalarTower Γ(X, U) (X.presheaf.stalk x) X.functionField := by apply IsScalarTower.of_algebraMap_eq' simp_rw [RingHom.algebraMap_toAlgebra] @@ -121,12 +121,11 @@ instance functionField_isFractionRing_of_affine (R : CommRingCat.{u}) [IsDomain ext exact mem_nonZeroDivisors_iff_ne_zero -instance {X : Scheme} [IsIntegral X] {U : Opens X} [hU : Nonempty U] : - IsIntegral (X ∣_ᵤ U) := - haveI : Nonempty (X ∣_ᵤ U) := hU - isIntegral_of_isOpenImmersion (Scheme.ιOpens U) +instance {X : Scheme} [IsIntegral X] {U : X.Opens} [Nonempty U] : + IsIntegral U := + isIntegral_of_isOpenImmersion U.ι -theorem IsAffineOpen.primeIdealOf_genericPoint {X : Scheme} [IsIntegral X] {U : Opens X} +theorem IsAffineOpen.primeIdealOf_genericPoint {X : Scheme} [IsIntegral X] {U : X.Opens} (hU : IsAffineOpen U) [h : Nonempty U] : hU.primeIdealOf ⟨genericPoint X, @@ -136,19 +135,18 @@ theorem IsAffineOpen.primeIdealOf_genericPoint {X : Scheme} [IsIntegral X] {U : delta IsAffineOpen.primeIdealOf convert genericPoint_eq_of_isOpenImmersion - ((X ∣_ᵤ U).isoSpec.hom ≫ Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top).op)) + (U.toScheme.isoSpec.hom ≫ Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top).op)) -- Porting note: this was `ext1` apply Subtype.ext - exact (genericPoint_eq_of_isOpenImmersion (Scheme.ιOpens U)).symm + exact (genericPoint_eq_of_isOpenImmersion U.ι).symm -theorem functionField_isFractionRing_of_isAffineOpen [IsIntegral X] (U : Opens X) - (hU : IsAffineOpen U) [hU' : Nonempty U] : +theorem functionField_isFractionRing_of_isAffineOpen [IsIntegral X] (U : X.Opens) + (hU : IsAffineOpen U) [Nonempty U] : IsFractionRing Γ(X, U) X.functionField := by haveI : IsAffine _ := hU - haveI : Nonempty (X ∣_ᵤ U) := hU' - haveI : IsIntegral (X ∣_ᵤ U) := + haveI : IsIntegral U := @isIntegral_of_isAffine_of_isDomain _ _ _ - (by dsimp; rw [Opens.openEmbedding_obj_top]; infer_instance) + (by rw [Scheme.Opens.toScheme_presheaf_obj, Opens.openEmbedding_obj_top]; infer_instance) delta IsFractionRing Scheme.functionField convert hU.isLocalization_stalk ⟨genericPoint X, _⟩ using 1 rw [hU.primeIdealOf_genericPoint, genericPoint_eq_bot_of_affine] @@ -159,7 +157,7 @@ instance (x : X) : IsAffine (X.affineCover.obj x) := instance [IsIntegral X] (x : X) : IsFractionRing (X.presheaf.stalk x) X.functionField := - let U : Opens X := (X.affineCover.map x).opensRange + let U : X.Opens := (X.affineCover.map x).opensRange have hU : IsAffineOpen U := isAffineOpen_opensRange (X.affineCover.map x) let x : U := ⟨x, X.affineCover.covers x⟩ have : Nonempty U := ⟨x⟩ diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 595194051f03c..7180e550a04f1 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -484,12 +484,9 @@ theorem toOpen_unit_app_val_c_app' {X : Scheme.{u}} (U : Opens (PrimeSpectrum Γ end ΓSpec -theorem ΓSpecIso_obj_hom {X : Scheme.{u}} (U : Opens X) : - (Scheme.ΓSpecIso Γ(X, U)).hom = - Scheme.Γ.map (Spec.map (X.presheaf.map (eqToHom U.openEmbedding_obj_top).op)).op ≫ - (ΓSpec.adjunction.unit.app (X ∣_ᵤ U)).app ⊤ ≫ - X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op := by - dsimp [-Scheme.SpecΓIdentity_hom_app] +theorem ΓSpecIso_obj_hom {X : Scheme.{u}} (U : X.Opens) : + (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).app ⊤ ≫ + (ΓSpec.adjunction.unit.app U).app ⊤ ≫ U.topIso.hom := by rw [ΓSpec.adjunction_unit_app_app_top] -- why can't simp find this simp diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index ebcc74b197b44..a43b1b4879ffe 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -112,7 +112,7 @@ instance : HasInitial Scheme.{u} := instance initial_isEmpty : IsEmpty (⊥_ Scheme) := ⟨fun x => ((initial.to Scheme.empty : _).1.base x).elim⟩ -theorem isAffineOpen_bot (X : Scheme) : IsAffineOpen (⊥ : Opens X) := +theorem isAffineOpen_bot (X : Scheme) : IsAffineOpen (⊥ : X.Opens) := @isAffine_of_isEmpty _ (inferInstanceAs (IsEmpty (∅ : Set X))) instance : HasStrictInitialObjects Scheme := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean new file mode 100644 index 0000000000000..e4979dd055a15 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -0,0 +1,187 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated +import Mathlib.AlgebraicGeometry.Morphisms.IsIso + +/-! + +# Affine morphisms of schemes + +A morphism of schemes `f : X ⟶ Y` is affine if the preimage +of an arbitrary affine open subset of `Y` is affine. + +It is equivalent to ask only that `Y` is covered by affine opens whose preimage is affine. + +## Main results + +- `AlgebraicGeometry.IsAffineHom`: The class of affine morphisms. +- `AlgebraicGeometry.isAffineOpen_of_isAffineOpen_basicOpen`: + If `s` is a spanning set of `Γ(X, U)`, such that each `X.basicOpen i` is affine, + then `U` is also affine. +- `AlgebraicGeometry.isAffineHom_stableUnderBaseChange`: + Affine morphisms are stable under base change. + +We also provide the instance `HasAffineProperty @IsAffineHom fun X _ _ _ ↦ IsAffine X`. + +## TODO +- Affine morphisms are separated. + +-/ + +universe v u + +open CategoryTheory TopologicalSpace Opposite + +namespace AlgebraicGeometry + +variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) + +/-- A morphism of schemes `X ⟶ Y` is affine if +the preimage of any affine open subset of `Y` is affine. -/ +@[mk_iff] +class IsAffineHom {X Y : Scheme} (f : X ⟶ Y) : Prop where + isAffine_preimage : ∀ U : Y.Opens, IsAffineOpen U → IsAffineOpen (f ⁻¹ᵁ U) + +lemma isAffineOpen.preimage {X Y : Scheme} {U : Y.Opens} (hU : IsAffineOpen U) + (f : X ⟶ Y) [IsAffineHom f] : + IsAffineOpen (f ⁻¹ᵁ U) := + IsAffineHom.isAffine_preimage _ hU + +/-- The preimage of an affine open as an `Scheme.affine_opens`. -/ +@[simps] +def affinePreimage {X Y : Scheme} (f : X ⟶ Y) [IsAffineHom f] (U : Y.affineOpens) : + X.affineOpens := + ⟨f ⁻¹ᵁ U.1, IsAffineHom.isAffine_preimage _ U.prop⟩ + +instance (priority := 900) [IsIso f] : IsAffineHom f := + ⟨fun _ hU ↦ hU.preimage_of_isIso f⟩ + +instance (priority := 900) [IsAffineHom f] : QuasiCompact f := + (quasiCompact_iff_forall_affine f).mpr + (fun U hU ↦ (IsAffineHom.isAffine_preimage U hU).isCompact) + +instance [IsAffineHom f] [IsAffineHom g] : IsAffineHom (f ≫ g) := by + constructor + intros U hU + rw [Scheme.comp_val_base, Opens.map_comp_obj] + apply IsAffineHom.isAffine_preimage + apply IsAffineHom.isAffine_preimage + exact hU + +instance : MorphismProperty.IsMultiplicative @IsAffineHom where + id_mem := inferInstance + comp_mem _ _ _ _ := inferInstance + +instance {X : Scheme} (r : Γ(X, ⊤)) : + IsAffineHom (X.basicOpen r).ι := by + constructor + intros U hU + fapply (Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion (X.basicOpen r).ι).mp + convert hU.basicOpen (X.presheaf.map (homOfLE le_top).op r) + rw [X.basicOpen_res] + ext1 + refine Set.image_preimage_eq_inter_range.trans ?_ + erw [Subtype.range_coe] + rfl + +lemma isAffineOpen_of_isAffineOpen_basicOpen_aux (s : Set Γ(X, ⊤)) + (hs : Ideal.span s = ⊤) (hs₂ : ∀ i ∈ s, IsAffineOpen (X.basicOpen i)) : + QuasiSeparatedSpace X := by + rw [quasiSeparatedSpace_iff_affine] + intros U V + obtain ⟨s', hs', e⟩ := (Ideal.span_eq_top_iff_finite _).mp hs + rw [← Set.inter_univ (_ ∩ _), ← Opens.coe_top, ← iSup_basicOpen_of_span_eq_top _ _ e, + ← iSup_subtype'', Opens.coe_iSup, Set.inter_iUnion] + apply isCompact_iUnion + intro i + rw [Set.inter_inter_distrib_right] + refine (hs₂ i (hs' i.2)).isQuasiSeparated _ _ Set.inter_subset_right + (U.1.2.inter (X.basicOpen _).2) ?_ Set.inter_subset_right (V.1.2.inter (X.basicOpen _).2) ?_ + · rw [← Opens.coe_inf, ← X.basicOpen_res _ (homOfLE le_top).op] + exact (U.2.basicOpen _).isCompact + · rw [← Opens.coe_inf, ← X.basicOpen_res _ (homOfLE le_top).op] + exact (V.2.basicOpen _).isCompact + +lemma isAffine_of_isAffineOpen_basicOpen (s : Set Γ(X, ⊤)) + (hs : Ideal.span s = ⊤) (hs₂ : ∀ i ∈ s, IsAffineOpen (X.basicOpen i)) : + IsAffine X := by + have : QuasiSeparatedSpace X := isAffineOpen_of_isAffineOpen_basicOpen_aux s hs hs₂ + have : CompactSpace X := by + obtain ⟨s', hs', e⟩ := (Ideal.span_eq_top_iff_finite _).mp hs + rw [← isCompact_univ_iff, ← Opens.coe_top, ← iSup_basicOpen_of_span_eq_top _ _ e] + simp only [Finset.mem_coe, Opens.iSup_mk, Opens.carrier_eq_coe, Opens.coe_mk] + apply s'.isCompact_biUnion + exact fun i hi ↦ (hs₂ _ (hs' hi)).isCompact + constructor + refine HasAffineProperty.of_iSup_eq_top (P := MorphismProperty.isomorphisms Scheme) + (fun i : s ↦ ⟨PrimeSpectrum.basicOpen i.1, ?_⟩) ?_ (fun i ↦ ⟨?_, ?_⟩) + · show IsAffineOpen _ + simp only [← basicOpen_eq_of_affine] + exact (isAffineOpen_top (Scheme.Spec.obj (op _))).basicOpen _ + · rw [PrimeSpectrum.iSup_basicOpen_eq_top_iff, Subtype.range_coe_subtype, Set.setOf_mem_eq, hs] + · show IsAffineOpen (ΓSpec.adjunction.unit.app X ⁻¹ᵁ PrimeSpectrum.basicOpen i.1) + rw [ΓSpec.adjunction_unit_map_basicOpen] + exact hs₂ _ i.2 + · simp only [Functor.comp_obj, Functor.rightOp_obj, Scheme.Γ_obj, Scheme.Spec_obj, id_eq, + eq_mpr_eq_cast, Functor.id_obj, Opens.map_top, morphismRestrict_app] + apply (config := { allowSynthFailures := true }) IsIso.comp_isIso + convert isIso_ΓSpec_adjunction_unit_app_basicOpen i.1 using 0 + refine congr(IsIso ((ΓSpec.adjunction.unit.app X).app $(?_))) + rw [Opens.openEmbedding_obj_top] + +/-- +If `s` is a spanning set of `Γ(X, U)`, such that each `X.basicOpen i` is affine, then `U` is also +affine. +-/ +lemma isAffineOpen_of_isAffineOpen_basicOpen (U) (s : Set Γ(X, U)) + (hs : Ideal.span s = ⊤) (hs₂ : ∀ i ∈ s, IsAffineOpen (X.basicOpen i)) : + IsAffineOpen U := by + apply isAffine_of_isAffineOpen_basicOpen (U.topIso.inv '' s) + · rw [← Ideal.map_span U.topIso.inv, hs, Ideal.map_top] + · rintro _ ⟨j, hj, rfl⟩ + rw [← (Scheme.Opens.ι _).isAffineOpen_iff_of_isOpenImmersion, Scheme.image_basicOpen] + simpa [Scheme.Opens.toScheme_presheaf_obj] using hs₂ j hj + +instance : HasAffineProperty @IsAffineHom fun X _ _ _ ↦ IsAffine X where + isLocal_affineProperty := by + constructor + · apply AffineTargetMorphismProperty.respectsIso_mk + · rintro X Y Z e _ _ H + have : IsAffine _ := H + exact isAffine_of_isIso e.hom + · exact fun _ _ _ ↦ id + · intro X Y _ f r H + have : IsAffine X := H + show IsAffineOpen _ + rw [Scheme.preimage_basicOpen] + exact (isAffineOpen_top X).basicOpen _ + · intro X Y _ f S hS hS' + apply_fun Ideal.map (f.1.c.app (op ⊤)) at hS + rw [Ideal.map_span, Ideal.map_top] at hS + apply isAffine_of_isAffineOpen_basicOpen _ hS + have : ∀ i : S, IsAffineOpen (f⁻¹ᵁ Y.basicOpen i.1) := hS' + simpa [Scheme.preimage_basicOpen] using this + eq_targetAffineLocally' := by + ext X Y f + simp only [targetAffineLocally, Scheme.affineOpens, Set.coe_setOf, Set.mem_setOf_eq, + Subtype.forall, isAffineHom_iff] + rfl + +lemma isAffineHom_stableUnderBaseChange : + MorphismProperty.StableUnderBaseChange @IsAffineHom := by + apply HasAffineProperty.stableUnderBaseChange + letI := HasAffineProperty.isLocal_affineProperty + apply AffineTargetMorphismProperty.StableUnderBaseChange.mk + introv X hX H + infer_instance + +instance (priority := 100) isAffineHom_of_isAffine [IsAffine X] [IsAffine Y] : IsAffineHom f := + (HasAffineProperty.iff_of_isAffine (P := @IsAffineHom)).mpr inferInstance + +lemma isAffine_of_isAffineHom [IsAffineHom f] [IsAffine Y] : IsAffine X := + (HasAffineProperty.iff_of_isAffine (P := @IsAffineHom) (f := f)).mp inferInstance + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index 91356a3a42e0b..c003c3c7331ed 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -51,7 +51,7 @@ For a morphism property `P` local at the source and `f : X ⟶ Y`, we provide th - `AlgebraicGeometry.IsLocalAtTarget.comp`: `P` is preserved under composition with open immersions at the source. - `AlgebraicGeometry.IsLocalAtTarget.iff_of_iSup_eq_top`: - `P f ↔ ∀ i, P (ιOpens U ≫ f)` for a family `U i` of open sets covering `X`. + `P f ↔ ∀ i, P (U.ι ≫ f)` for a family `U i` of open sets covering `X`. - `AlgebraicGeometry.IsLocalAtTarget.iff_of_openCover`: `P f ↔ ∀ i, P (𝒰.map i ≫ f)` for `𝒰 : X.openCover`. - `AlgebraicGeometry.IsLocalAtTarget.of_isOpenImmersion`: If `P` contains identities then `P` holds @@ -124,9 +124,9 @@ attribute [instance] respectsIso 3. If `P` holds for `f ∣_ U` for an open cover `U` of `Y`, then `P` holds for `f`. -/ protected lemma mk' {P : MorphismProperty Scheme} [P.RespectsIso] - (restrict : ∀ {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y), P f → P (f ∣_ U)) + (restrict : ∀ {X Y : Scheme} (f : X ⟶ Y) (U : Y.Opens), P f → P (f ∣_ U)) (of_sSup_eq_top : - ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) {ι : Type u} (U : ι → Opens Y), iSup U = ⊤ → + ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) {ι : Type u} (U : ι → Y.Opens), iSup U = ⊤ → (∀ i, P (f ∣_ U i)) → P f) : IsLocalAtTarget P := by refine ⟨inferInstance, fun {X Y} f 𝒰 ↦ ⟨?_, fun H ↦ of_sSup_eq_top f _ 𝒰.iSup_opensRange ?_⟩⟩ @@ -150,20 +150,20 @@ lemma of_isPullback {UX UY : Scheme.{u}} {iY : UY ⟶ Y} [IsOpenImmersion iY] rw [← P.cancel_left_of_respectsIso h.isoPullback.inv, h.isoPullback_inv_snd] exact (iff_of_openCover' f (Y.affineCover.add iY)).mp H .none -theorem restrict (hf : P f) (U : Opens Y) : P (f ∣_ U) := +theorem restrict (hf : P f) (U : Y.Opens) : P (f ∣_ U) := of_isPullback (isPullback_morphismRestrict f U).flip hf -lemma of_iSup_eq_top {ι} (U : ι → Opens Y) (hU : iSup U = ⊤) +lemma of_iSup_eq_top {ι} (U : ι → Y.Opens) (hU : iSup U = ⊤) (H : ∀ i, P (f ∣_ U i)) : P f := by refine (IsLocalAtTarget.iff_of_openCover' f (Y.openCoverOfSuprEqTop (s := Set.range U) Subtype.val (by ext; simp [← hU]))).mpr fun i ↦ ?_ obtain ⟨_, i, rfl⟩ := i refine (P.arrow_mk_iso_iff (morphismRestrictOpensRange f _)).mp ?_ - show P (f ∣_ (Scheme.ιOpens (U i)).opensRange) - rw [opensRange_ιOpens] + show P (f ∣_ (U i).ι.opensRange) + rw [Scheme.Opens.opensRange_ι] exact H i -theorem iff_of_iSup_eq_top {ι} (U : ι → Opens Y) (hU : iSup U = ⊤) : +theorem iff_of_iSup_eq_top {ι} (U : ι → Y.Opens) (hU : iSup U = ⊤) : P f ↔ ∀ i, P (f ∣_ U i) := ⟨fun H _ ↦ restrict H _, of_iSup_eq_top U hU⟩ @@ -202,19 +202,19 @@ attribute [instance] respectsIso 3. If `P` holds for `f ∣_ U` for an open cover `U` of `Y`, then `P` holds for `f`. -/ protected lemma mk' {P : MorphismProperty Scheme} [P.RespectsIso] - (restrict : ∀ {X Y : Scheme} (f : X ⟶ Y) (U : Opens X), P f → P (Scheme.ιOpens U ≫ f)) + (restrict : ∀ {X Y : Scheme} (f : X ⟶ Y) (U : X.Opens), P f → P (U.ι ≫ f)) (of_sSup_eq_top : - ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) {ι : Type u} (U : ι → Opens X), iSup U = ⊤ → - (∀ i, P (Scheme.ιOpens (U i) ≫ f)) → P f) : + ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) {ι : Type u} (U : ι → X.Opens), iSup U = ⊤ → + (∀ i, P ((U i).ι ≫ f)) → P f) : IsLocalAtSource P := by refine ⟨inferInstance, fun {X Y} f 𝒰 ↦ ⟨fun H i ↦ ?_, fun H ↦ of_sSup_eq_top f _ 𝒰.iSup_opensRange fun i ↦ ?_⟩⟩ - · rw [← IsOpenImmersion.isoOfRangeEq_hom_fac (𝒰.map i) (Scheme.ιOpens _) - (congr_arg Opens.carrier (opensRange_ιOpens (𝒰.map i).opensRange).symm), Category.assoc, + · rw [← IsOpenImmersion.isoOfRangeEq_hom_fac (𝒰.map i) (Scheme.Opens.ι _) + (congr_arg Opens.carrier (𝒰.map i).opensRange.opensRange_ι.symm), Category.assoc, P.cancel_left_of_respectsIso] exact restrict _ _ H - · rw [← IsOpenImmersion.isoOfRangeEq_inv_fac (𝒰.map i) (Scheme.ιOpens _) - (congr_arg Opens.carrier (opensRange_ιOpens (𝒰.map i).opensRange).symm), Category.assoc, + · rw [← IsOpenImmersion.isoOfRangeEq_inv_fac (𝒰.map i) (Scheme.Opens.ι _) + (congr_arg Opens.carrier (𝒰.map i).opensRange.opensRange_ι.symm), Category.assoc, P.cancel_left_of_respectsIso] exact H _ @@ -234,21 +234,21 @@ lemma comp {UX : Scheme.{u}} (H : P f) (i : UX ⟶ X) [IsOpenImmersion i] : P (i ≫ f) := (iff_of_openCover' f (X.affineCover.add i)).mp H .none -lemma of_iSup_eq_top {ι} (U : ι → Opens X) (hU : iSup U = ⊤) - (H : ∀ i, P (Scheme.ιOpens (U i) ≫ f)) : P f := by +lemma of_iSup_eq_top {ι} (U : ι → X.Opens) (hU : iSup U = ⊤) + (H : ∀ i, P ((U i).ι ≫ f)) : P f := by refine (iff_of_openCover' f (X.openCoverOfSuprEqTop (s := Set.range U) Subtype.val (by ext; simp [← hU]))).mpr fun i ↦ ?_ obtain ⟨_, i, rfl⟩ := i exact H i -theorem iff_of_iSup_eq_top {ι} (U : ι → Opens X) (hU : iSup U = ⊤) : - P f ↔ ∀ i, P (Scheme.ιOpens (U i) ≫ f) := +theorem iff_of_iSup_eq_top {ι} (U : ι → X.Opens) (hU : iSup U = ⊤) : + P f ↔ ∀ i, P ((U i).ι ≫ f) := ⟨fun H _ ↦ comp H _, of_iSup_eq_top U hU⟩ lemma of_openCover (H : ∀ i, P (𝒰.map i ≫ f)) : P f := by refine of_iSup_eq_top (fun i ↦ (𝒰.map i).opensRange) 𝒰.iSup_opensRange fun i ↦ ?_ - rw [← IsOpenImmersion.isoOfRangeEq_inv_fac (𝒰.map i) (Scheme.ιOpens _) - (congr_arg Opens.carrier (opensRange_ιOpens (𝒰.map i).opensRange).symm), Category.assoc, + rw [← IsOpenImmersion.isoOfRangeEq_inv_fac (𝒰.map i) (Scheme.Opens.ι _) + (congr_arg Opens.carrier (𝒰.map i).opensRange.opensRange_ι.symm), Category.assoc, P.cancel_left_of_respectsIso] exact H i @@ -469,7 +469,7 @@ instance (priority := 900) : P.RespectsIso := by infer_instance theorem of_iSup_eq_top - {ι} (U : ι → Y.affineOpens) (hU : ⨆ i, (U i : Opens Y) = ⊤) + {ι} (U : ι → Y.affineOpens) (hU : ⨆ i, (U i : Y.Opens) = ⊤) (hU' : ∀ i, Q (f ∣_ U i)) : P f := by letI := isLocal_affineProperty P @@ -479,12 +479,12 @@ theorem of_iSup_eq_top induction V using of_affine_open_cover U hU with | basicOpen U r h => haveI : IsAffine _ := U.2 - have := AffineTargetMorphismProperty.IsLocal.to_basicOpen (f ∣_ U.1) ((Y.resTop U).inv r) h + have := AffineTargetMorphismProperty.IsLocal.to_basicOpen (f ∣_ U.1) (U.1.topIso.inv r) h exact (Q.arrow_mk_iso_iff (morphismRestrictRestrictBasicOpen f _ r)).mp this | openCover U s hs H => apply AffineTargetMorphismProperty.IsLocal.of_basicOpenCover _ - (s.image (Y.resTop _).inv) (by simp [← Ideal.map_span, hs, Ideal.map_top]) + (s.image (Scheme.Opens.topIso _).inv) (by simp [← Ideal.map_span, hs, Ideal.map_top]) intro ⟨r, hr⟩ obtain ⟨r, hr', rfl⟩ := Finset.mem_image.mp hr exact (Q.arrow_mk_iso_iff @@ -492,7 +492,7 @@ theorem of_iSup_eq_top | hU i => exact hU' i theorem iff_of_iSup_eq_top - {ι} (U : ι → Y.affineOpens) (hU : ⨆ i, (U i : Opens Y) = ⊤) : + {ι} (U : ι → Y.affineOpens) (hU : ⨆ i, (U i : Y.Opens) = ⊤) : P f ↔ ∀ i, Q (f ∣_ U i) := ⟨fun H _ ↦ restrict H _, fun H ↦ HasAffineProperty.of_iSup_eq_top U hU H⟩ @@ -536,7 +536,7 @@ instance (priority := 900) : IsLocalAtTarget P := by have : P (𝒰.pullbackHom f i) := by refine (P.arrow_mk_iso_iff (morphismRestrictEq _ ?_ ≪≫ morphismRestrictOpensRange f (𝒰.map i))).mp (H i) - exact (opensRange_ιOpens _).symm + exact (Scheme.Opens.opensRange_ι _).symm rw [← Q.cancel_left_of_respectsIso (𝒰.pullbackCoverAffineRefinementObjIso f _).inv, 𝒰.pullbackCoverAffineRefinementObjIso_inv_pullbackHom] exact of_isPullback (.of_hasPullback _ _) this @@ -582,6 +582,16 @@ theorem stableUnderBaseChange (hP' : Q.StableUnderBaseChange) : rw [← pullbackSymmetry_hom_comp_snd, Q.cancel_left_of_respectsIso] apply of_isPullback (.of_hasPullback _ _) H) +lemma isLocalAtSource + (H : ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) [IsAffine Y] (𝒰 : Scheme.OpenCover.{u} X), + Q f ↔ ∀ i, Q (𝒰.map i ≫ f)) : IsLocalAtSource P where + iff_of_openCover' {X Y} f 𝒰 := by + simp_rw [IsLocalAtTarget.iff_of_iSup_eq_top _ (iSup_affineOpens_eq_top Y)] + rw [forall_comm] + refine forall_congr' fun U ↦ ?_ + simp_rw [HasAffineProperty.iff_of_isAffine, morphismRestrict_comp] + exact @H _ _ (f ∣_ U.1) U.2 (𝒰.restrict (f ⁻¹ᵁ U.1)) + end HasAffineProperty end targetAffineLocally diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean index 59f2d5c319800..b89116345a9be 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Constructors.lean @@ -77,11 +77,11 @@ theorem HasAffineProperty.diagonal_of_openCover (P) {Q} [HasAffineProperty P Q] ((pullbackDiagonalMapIso _ _ ((𝒰' i).map j) ((𝒰' i).map k)).inv ≫ pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (𝟙 _) _ _) (pullback.snd _ _)).mp _ using 1 · simp - · ext <;> simp + · ext1 <;> simp · simp only [Category.assoc, limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Functor.const_obj_obj, cospan_one, cospan_left, cospan_right, Category.comp_id] convert h𝒰' i j k - ext <;> simp [Scheme.OpenCover.pullbackHom] + ext1 <;> simp [Scheme.OpenCover.pullbackHom] theorem HasAffineProperty.diagonal_of_openCover_diagonal (P) {Q} [HasAffineProperty P Q] @@ -156,7 +156,7 @@ end Diagonal section Universally theorem universally_isLocalAtTarget (P : MorphismProperty Scheme) - (hP₂ : ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) {ι : Type u} (U : ι → Opens Y.carrier) + (hP₂ : ∀ {X Y : Scheme.{u}} (f : X ⟶ Y) {ι : Type u} (U : ι → Y.Opens) (_ : iSup U = ⊤), (∀ i, P (f ∣_ U i)) → P f) : IsLocalAtTarget P.universally := by apply IsLocalAtTarget.mk' · exact fun {X Y} f U => P.universally_stableUnderBaseChange @@ -173,7 +173,7 @@ theorem universally_isLocalAtTarget (P : MorphismProperty Scheme) · simp only [Scheme.restrictIsoOfEq, Category.assoc, morphismRestrict_ι, IsOpenImmersion.isoOfRangeEq_hom_fac_assoc] exact (isPullback_morphismRestrict f' (i₂ ⁻¹ᵁ U i)).paste_vert h - · rw [← cancel_mono (Scheme.ιOpens _)] + · rw [← cancel_mono (Scheme.Opens.ι _)] simp [IsOpenImmersion.isoOfRangeEq_hom_fac_assoc, Scheme.restrictIsoOfEq, morphismRestrict_ι_assoc, h.1.1] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean index 62366add65526..ffc351acd412a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean @@ -14,7 +14,7 @@ A morphism of schemes `f : X ⟶ Y` is locally of finite type if for each affine A morphism of schemes is of finite type if it is both locally of finite type and quasi-compact. -We show that these properties are local, and are stable under compositions. +We show that these properties are local, and are stable under compositions and base change. -/ @@ -37,45 +37,27 @@ class LocallyOfFiniteType (f : X ⟶ Y) : Prop where finiteType_of_affine_subset : ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), (f.appLE U V e).FiniteType -theorem locallyOfFiniteType_eq : @LocallyOfFiniteType = affineLocally @RingHom.FiniteType := by - ext X Y f - rw [locallyOfFiniteType_iff, affineLocally_iff_affineOpens_le] - exact RingHom.finiteType_respectsIso +instance : HasRingHomProperty @LocallyOfFiniteType RingHom.FiniteType where + isLocal_ringHomProperty := RingHom.finiteType_is_local + eq_affineLocally' := by + ext X Y f + rw [locallyOfFiniteType_iff, affineLocally_iff_affineOpens_le] instance (priority := 900) locallyOfFiniteType_of_isOpenImmersion [IsOpenImmersion f] : LocallyOfFiniteType f := - locallyOfFiniteType_eq.symm ▸ RingHom.finiteType_is_local.affineLocally_of_isOpenImmersion f - -instance locallyOfFiniteType_isStableUnderComposition : - MorphismProperty.IsStableUnderComposition @LocallyOfFiniteType := - locallyOfFiniteType_eq.symm ▸ RingHom.finiteType_is_local.affineLocally_isStableUnderComposition + HasRingHomProperty.of_isOpenImmersion instance locallyOfFiniteType_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [hf : LocallyOfFiniteType f] [hg : LocallyOfFiniteType g] : LocallyOfFiniteType (f ≫ g) := MorphismProperty.comp_mem _ f g hf hg theorem locallyOfFiniteType_of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) - [hf : LocallyOfFiniteType (f ≫ g)] : LocallyOfFiniteType f := by - revert hf - rw [locallyOfFiniteType_eq] - apply RingHom.finiteType_is_local.affineLocally_of_comp - introv H - exact RingHom.FiniteType.of_comp_finiteType H - -theorem LocallyOfFiniteType.affine_openCover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) - (𝒰 : Scheme.OpenCover.{u} Y) [∀ i, IsAffine (𝒰.obj i)] - (𝒰' : ∀ i, Scheme.OpenCover.{u} ((𝒰.pullbackCover f).obj i)) [∀ i j, IsAffine ((𝒰' i).obj j)] : - LocallyOfFiniteType f ↔ - ∀ i j, (Scheme.Γ.map ((𝒰' i).map j ≫ pullback.snd _ _).op).FiniteType := - locallyOfFiniteType_eq.symm ▸ RingHom.finiteType_is_local.affine_openCover_iff f 𝒰 𝒰' - -theorem LocallyOfFiniteType.source_openCover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) - (𝒰 : Scheme.OpenCover.{u} X) : LocallyOfFiniteType f ↔ ∀ i, LocallyOfFiniteType (𝒰.map i ≫ f) := - locallyOfFiniteType_eq.symm ▸ RingHom.finiteType_is_local.source_openCover_iff f 𝒰 - -instance locallyOfFiniteType_isLocalAtTarget : IsLocalAtTarget @LocallyOfFiniteType := by - have := RingHom.finiteType_is_local.hasAffinePropertyAffineLocally - rw [← locallyOfFiniteType_eq] at this - infer_instance + [LocallyOfFiniteType (f ≫ g)] : LocallyOfFiniteType f := + HasRingHomProperty.of_comp (fun f g ↦ RingHom.FiniteType.of_comp_finiteType) ‹_› + +open scoped TensorProduct in +lemma locallyOfFiniteType_stableUnderBaseChange : + MorphismProperty.StableUnderBaseChange @LocallyOfFiniteType := + HasRingHomProperty.stableUnderBaseChange RingHom.finiteType_stableUnderBaseChange end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index 5b67ced6ba235..211a07cc5edf0 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -65,7 +65,7 @@ instance quasiCompact_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [QuasiCo theorem isCompactOpen_iff_eq_finset_affine_union {X : Scheme} (U : Set X) : IsCompact U ∧ IsOpen U ↔ ∃ s : Set X.affineOpens, s.Finite ∧ U = ⋃ i ∈ s, i := by apply Opens.IsBasis.isCompact_open_iff_eq_finite_iUnion - (fun (U : X.affineOpens) => (U : Opens X)) + (fun (U : X.affineOpens) => (U : X.Opens)) · rw [Subtype.range_coe]; exact isBasis_affine_open X · exact fun i => i.2.isCompact @@ -77,7 +77,7 @@ theorem isCompactOpen_iff_eq_basicOpen_union {X : Scheme} [IsAffine X] (U : Set theorem quasiCompact_iff_forall_affine : QuasiCompact f ↔ - ∀ U : Opens Y, IsAffineOpen U → IsCompact (f ⁻¹ᵁ U : Set X) := by + ∀ U : Y.Opens, IsAffineOpen U → IsCompact (f ⁻¹ᵁ U : Set X) := by rw [quasiCompact_iff] refine ⟨fun H U hU => H U U.isOpen hU.isCompact, ?_⟩ intro H U hU hU' @@ -85,7 +85,7 @@ theorem quasiCompact_iff_forall_affine : simp only [Set.preimage_iUnion] exact Set.Finite.isCompact_biUnion hS (fun i _ => H i i.prop) -theorem isCompact_basicOpen (X : Scheme) {U : Opens X} (hU : IsCompact (U : Set X)) +theorem isCompact_basicOpen (X : Scheme) {U : X.Opens} (hU : IsCompact (U : Set X)) (f : Γ(X, U)) : IsCompact (X.basicOpen f : Set X) := by classical refine ((isCompactOpen_iff_eq_finset_affine_union _).mpr ?_).1 @@ -109,7 +109,7 @@ theorem isCompact_basicOpen (X : Scheme) {U : Opens X} (hU : IsCompact (U : Set -- Porting note: had to make explicit the first given parameter to `Set.subset_iUnion₂` exact Set.Subset.trans (Set.Subset.rfl : _ ≤ g ⟨i, hi⟩) (@Set.subset_iUnion₂ _ _ _ - (fun (i : Scheme.affineOpens X) (_ : i ∈ Set.range g) => (i : Set X.toPresheafedSpace)) _ + (fun (i : X.affineOpens) (_ : i ∈ Set.range g) => (i : Set X.toPresheafedSpace)) _ (Set.mem_range_self ⟨i, hi⟩)) · rintro ⟨i, hi⟩ ⟨⟨j, hj⟩, hj'⟩ rw [← hj'] @@ -176,18 +176,18 @@ instance (f : X ⟶ Z) (g : Y ⟶ Z) [QuasiCompact f] : QuasiCompact (pullback.s quasiCompact_stableUnderBaseChange.snd f g inferInstance @[elab_as_elim] -theorem compact_open_induction_on {P : Opens X → Prop} (S : Opens X) +theorem compact_open_induction_on {P : X.Opens → Prop} (S : X.Opens) (hS : IsCompact S.1) (h₁ : P ⊥) - (h₂ : ∀ (S : Opens X) (_ : IsCompact S.1) (U : X.affineOpens), P S → P (S ⊔ U)) : + (h₂ : ∀ (S : X.Opens) (_ : IsCompact S.1) (U : X.affineOpens), P S → P (S ⊔ U)) : P S := by classical obtain ⟨s, hs, hs'⟩ := (isCompactOpen_iff_eq_finset_affine_union S.1).mp ⟨hS, S.2⟩ - replace hs' : S = iSup fun i : s => (i : Opens X) := by ext1; simpa using hs' + replace hs' : S = iSup fun i : s => (i : X.Opens) := by ext1; simpa using hs' subst hs' apply @Set.Finite.induction_on _ _ _ hs · convert h₁; rw [iSup_eq_bot]; rintro ⟨_, h⟩; exact h.elim · intro x s _ hs h₄ - have : IsCompact (⨆ i : s, (i : Opens X)).1 := by + have : IsCompact (⨆ i : s, (i : X.Opens)).1 := by refine ((isCompactOpen_iff_eq_finset_affine_union _).mpr ?_).1; exact ⟨s, hs, by simp⟩ convert h₂ _ this x h₄ rw [iSup_subtype, sup_comm] @@ -195,7 +195,7 @@ theorem compact_open_induction_on {P : Opens X → Prop} (S : Opens X) exact iSup_insert theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Scheme) - {U : Opens X} (hU : IsAffineOpen U) (x f : Γ(X, U)) + {U : X.Opens} (hU : IsAffineOpen U) (x f : Γ(X, U)) (H : x |_ X.basicOpen f = 0) : ∃ n : ℕ, f ^ n * x = 0 := by rw [← map_zero (X.presheaf.map (homOfLE <| X.basicOpen_le f : X.basicOpen f ⟶ U).op)] at H obtain ⟨⟨_, n, rfl⟩, e⟩ := (hU.isLocalization_basicOpen f).exists_of_eq H @@ -204,14 +204,14 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Sch /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then `f ^ n * x = 0` for some `n`. -/ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme.{u}) - {U : Opens X} (hU : IsCompact U.1) (x f : Γ(X, U)) + {U : X.Opens} (hU : IsCompact U.1) (x f : Γ(X, U)) (H : x |_ X.basicOpen f = 0) : ∃ n : ℕ, f ^ n * x = 0 := by obtain ⟨s, hs, e⟩ := (isCompactOpen_iff_eq_finset_affine_union U.1).mp ⟨hU, U.2⟩ - replace e : U = iSup fun i : s => (i : Opens X) := by + replace e : U = iSup fun i : s => (i : X.Opens) := by ext1; simpa using e have h₁ : ∀ i : s, i.1.1 ≤ U := by intro i - change (i : Opens X) ≤ U + change (i : X.Opens) ≤ U rw [e] -- Porting note: `exact le_iSup _ _` no longer works exact le_iSup (fun (i : s) => (i : Opens (X.toPresheafedSpace))) _ @@ -233,7 +233,7 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme use Finset.univ.sup n suffices ∀ i : s, X.presheaf.map (homOfLE (h₁ i)).op (f ^ Finset.univ.sup n * x) = 0 by subst e - apply TopCat.Sheaf.eq_of_locally_eq X.sheaf fun i : s => (i : Opens X) + apply TopCat.Sheaf.eq_of_locally_eq X.sheaf fun i : s => (i : X.Opens) intro i rw [map_zero] apply this @@ -249,7 +249,7 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme /-- A section over a compact open of a scheme is nilpotent if and only if its associated basic open is empty. -/ lemma Scheme.isNilpotent_iff_basicOpen_eq_bot_of_isCompact {X : Scheme.{u}} - {U : Opens X} (hU : IsCompact U.carrier) (f : Γ(X, U)) : + {U : X.Opens} (hU : IsCompact (U : Set X)) (f : Γ(X, U)) : IsNilpotent f ↔ X.basicOpen f = ⊥ := by refine ⟨X.basicOpen_eq_bot_of_isNilpotent U f, fun hf ↦ ?_⟩ have h : (1 : Γ(X, U)) |_ X.basicOpen f = (0 : Γ(X, X.basicOpen f)) := by @@ -266,9 +266,9 @@ lemma Scheme.isNilpotent_iff_basicOpen_eq_bot_of_isCompact {X : Scheme.{u}} /-- The zero locus of a set of sections over a compact open of a scheme is `X` if and only if `s` is contained in the nilradical of `Γ(X, U)`. -/ -lemma Scheme.zeroLocus_eq_top_iff_subset_nilradical_of_isCompact {X : Scheme.{u}} {U : Opens X} - (hU : IsCompact U.carrier) (s : Set Γ(X, U)) : - X.zeroLocus s = ⊤ ↔ s ⊆ (nilradical Γ(X, U)).carrier := by +lemma Scheme.zeroLocus_eq_top_iff_subset_nilradical_of_isCompact {X : Scheme.{u}} {U : X.Opens} + (hU : IsCompact (U : Set X)) (s : Set Γ(X, U)) : + X.zeroLocus s = ⊤ ↔ s ⊆ nilradical Γ(X, U) := by simp [Scheme.zeroLocus_def, ← Scheme.isNilpotent_iff_basicOpen_eq_bot_of_isCompact hU, ← mem_nilradical, Set.subset_def] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean index 33d1eacce69d1..83373737cbce2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiSeparated.lean @@ -53,20 +53,18 @@ theorem quasiSeparatedSpace_iff_affine (X : Scheme) : · intro H U V; exact H U V U.1.2 U.2.isCompact V.1.2 V.2.isCompact · intro H suffices - ∀ (U : Opens X) (_ : IsCompact U.1) (V : Opens X) (_ : IsCompact V.1), + ∀ (U : X.Opens) (_ : IsCompact U.1) (V : X.Opens) (_ : IsCompact V.1), IsCompact (U ⊓ V).1 by intro U V hU hU' hV hV'; exact this ⟨U, hU⟩ hU' ⟨V, hV⟩ hV' intro U hU V hV - -- Porting note: it complains "unable to find motive", but telling Lean that motive is - -- underscore is actually sufficient, weird - apply compact_open_induction_on (P := _) V hV + refine compact_open_induction_on V hV ?_ ?_ · simp · intro S _ V hV change IsCompact (U.1 ∩ (S.1 ∪ V.1)) rw [Set.inter_union_distrib_left] apply hV.union clear hV - apply compact_open_induction_on (P := _) U hU + refine compact_open_induction_on U hU ?_ ?_ · simp · intro S _ W hW change IsCompact ((S.1 ∪ W.1) ∩ V.1) @@ -84,8 +82,7 @@ theorem quasiCompact_affineProperty_iff_quasiSeparatedSpace {X Y : Scheme} [IsAf · intro H U V haveI : IsAffine _ := U.2 haveI : IsAffine _ := V.2 - let g : pullback (X.ofRestrict U.1.openEmbedding) (X.ofRestrict V.1.openEmbedding) ⟶ X := - pullback.fst _ _ ≫ X.ofRestrict _ + let g : pullback U.1.ι V.1.ι ⟶ X := pullback.fst _ _ ≫ U.1.ι -- Porting note: `inferInstance` does not work here have : IsOpenImmersion g := PresheafedSpace.IsOpenImmersion.comp _ _ have e := Homeomorph.ofEmbedding _ this.base_open.toEmbedding @@ -169,7 +166,7 @@ instance quasiSeparatedSpace_of_isAffine (X : Scheme) [IsAffine X] : rw [← Scheme.basicOpen_mul] exact ((isAffineOpen_top _).basicOpen _).isCompact -theorem IsAffineOpen.isQuasiSeparated {X : Scheme} {U : Opens X} (hU : IsAffineOpen U) : +theorem IsAffineOpen.isQuasiSeparated {X : Scheme} {U : X.Opens} (hU : IsAffineOpen U) : IsQuasiSeparated (U : Set X) := by rw [isQuasiSeparated_iff_quasiSeparatedSpace] exacts [@AlgebraicGeometry.quasiSeparatedSpace_of_isAffine _ hU, U.isOpen] @@ -187,7 +184,7 @@ theorem QuasiSeparated.of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [Qua (pullbackRightPullbackFstIso g (Z.affineCover.map i) f).hom · exact inferInstance -theorem exists_eq_pow_mul_of_isAffineOpen (X : Scheme) (U : Opens X) (hU : IsAffineOpen U) +theorem exists_eq_pow_mul_of_isAffineOpen (X : Scheme) (U : X.Opens) (hU : IsAffineOpen U) (f : Γ(X, U)) (x : Γ(X, X.basicOpen f)) : ∃ (n : ℕ) (y : Γ(X, U)), y |_ X.basicOpen f = (f |_ X.basicOpen f) ^ n * x := by have := (hU.isLocalization_basicOpen f).2 @@ -211,7 +208,7 @@ theorem exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux_aux {X : To rw [e₁, e₂, mul_left_comm] theorem exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux (X : Scheme) - (S : X.affineOpens) (U₁ U₂ : Opens X) {n₁ n₂ : ℕ} {y₁ : Γ(X, U₁)} + (S : X.affineOpens) (U₁ U₂ : X.Opens) {n₁ n₂ : ℕ} {y₁ : Γ(X, U₁)} {y₂ : Γ(X, U₂)} {f : Γ(X, U₁ ⊔ U₂)} {x : Γ(X, X.basicOpen f)} (h₁ : S.1 ≤ U₁) (h₂ : S.1 ≤ U₂) (e₁ : y₁ |_ X.basicOpen (f |_ U₁) = ((f |_ U₁ |_ X.basicOpen _) ^ n₁) * x |_ X.basicOpen _) @@ -240,14 +237,12 @@ theorem exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux (X : Scheme Subtype.coe_mk] at e ⊢ rw [e] -theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : Opens X) +theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : X.Opens) (hU : IsCompact U.1) (hU' : IsQuasiSeparated U.1) (f : Γ(X, U)) (x : Γ(X, X.basicOpen f)) : ∃ (n : ℕ) (y : Γ(X, U)), y |_ X.basicOpen f = (f |_ X.basicOpen f) ^ n * x := by delta TopCat.Presheaf.restrictOpen TopCat.Presheaf.restrict revert hU' f x - -- Porting note: complains `expected type is not available`, but tell Lean that it is underscore - -- is sufficient - apply compact_open_induction_on (P := _) U hU + refine compact_open_induction_on U hU ?_ ?_ · intro _ f x use 0, f refine @Subsingleton.elim _ @@ -276,19 +271,19 @@ theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : (S ⊓ U.1).2⟩ haveI := hs'.to_subtype cases nonempty_fintype s - replace hs : S ⊓ U.1 = iSup fun i : s => (i : Opens X) := by ext1; simpa using hs + replace hs : S ⊓ U.1 = iSup fun i : s => (i : X.Opens) := by ext1; simpa using hs have hs₁ : ∀ i : s, i.1.1 ≤ S := by - intro i; change (i : Opens X) ≤ S + intro i; change (i : X.Opens) ≤ S refine le_trans ?_ (inf_le_left (b := U.1)) erw [hs] -- Porting note: have to add argument explicitly - exact @le_iSup (Opens X) s _ (fun (i : s) => (i : Opens X)) i + exact @le_iSup X.Opens s _ (fun (i : s) => (i : X.Opens)) i have hs₂ : ∀ i : s, i.1.1 ≤ U.1 := by - intro i; change (i : Opens X) ≤ U + intro i; change (i : X.Opens) ≤ U refine le_trans ?_ (inf_le_right (a := S)) erw [hs] -- Porting note: have to add argument explicitly - exact @le_iSup (Opens X) s _ (fun (i : s) => (i : Opens X)) i + exact @le_iSup X.Opens s _ (fun (i : s) => (i : X.Opens)) i -- On each affine open in the intersection, we have `f ^ (n + n₂) * y₁ = f ^ (n + n₁) * y₂` -- for some `n` since `f ^ n₂ * y₁ = f ^ (n₁ + n₂) * x = f ^ n₁ * y₂` on `X_f`. have := fun i ↦ exists_eq_pow_mul_of_is_compact_of_quasi_separated_space_aux @@ -304,7 +299,7 @@ theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : fapply X.sheaf.eq_of_locally_eq' fun i : s => i.1.1 · refine fun i => homOfLE ?_; erw [hs] -- Porting note: have to add argument explicitly - exact @le_iSup (Opens X) s _ (fun (i : s) => (i : Opens X)) i + exact @le_iSup X.Opens s _ (fun (i : s) => (i : X.Opens)) i · exact le_of_eq hs · intro i delta Scheme.sheaf SheafedSpace.sheaf @@ -332,7 +327,7 @@ theorem exists_eq_pow_mul_of_isCompact_of_isQuasiSeparated (X : Scheme.{u}) (U : /-- If `U` is qcqs, then `Γ(X, D(f)) ≃ Γ(X, U)_f` for every `f : Γ(X, U)`. This is known as the **Qcqs lemma** in [R. Vakil, *The rising sea*][RisingSea]. -/ -theorem is_localization_basicOpen_of_qcqs {X : Scheme} {U : Opens X} (hU : IsCompact U.1) +theorem is_localization_basicOpen_of_qcqs {X : Scheme} {U : X.Opens} (hU : IsCompact U.1) (hU' : IsQuasiSeparated U.1) (f : Γ(X, U)) : IsLocalization.Away f (Γ(X, X.basicOpen f)) := by constructor @@ -352,4 +347,22 @@ theorem is_localization_basicOpen_of_qcqs {X : Scheme} {U : Opens X} (hU : IsCom refine ⟨⟨_, n, rfl⟩, ?_⟩ simpa [mul_comm z] using e +/-- If `U` is qcqs, then `Γ(X, D(f)) ≃ Γ(X, U)_f` for every `f : Γ(X, U)`. +This is known as the **Qcqs lemma** in [R. Vakil, *The rising sea*][RisingSea]. -/ +theorem isIso_ΓSpec_adjunction_unit_app_basicOpen {X : Scheme} [CompactSpace X] + [QuasiSeparatedSpace X] (f : X.presheaf.obj (op ⊤)) : + IsIso ((ΓSpec.adjunction.unit.app X).val.c.app (op (PrimeSpectrum.basicOpen f))) := by + refine @IsIso.of_isIso_comp_right _ _ _ _ _ _ (X.presheaf.map + (eqToHom (ΓSpec.adjunction_unit_map_basicOpen _ _).symm).op) _ ?_ + rw [ConcreteCategory.isIso_iff_bijective, CommRingCat.forget_map] + apply (config := { allowSynthFailures := true }) IsLocalization.bijective + · exact StructureSheaf.IsLocalization.to_basicOpen _ _ + · refine is_localization_basicOpen_of_qcqs ?_ ?_ _ + · exact isCompact_univ + · exact isQuasiSeparated_univ + · rw [← CommRingCat.comp_eq_ring_hom_comp] + simp [RingHom.algebraMap_toAlgebra] + rw [ΓSpec.toOpen_unit_app_val_c_app'_assoc, ← Functor.map_comp] + rfl + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index a03d3e4341cad..ed9f201a49cea 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -27,8 +27,31 @@ the target. (TODO) For the latter `P` should be local on the target (`RingHom.PropertyIsLocal P`), and `affineLocally P` will be local on both the source and the target. - -Further more, these properties are stable under compositions (resp. base change) if `P` is. (TODO) +We also provide the following interface: + +## `HasRingHomProperty` + +`HasRingHomProperty P Q` is a type class asserting that `P` is local at the target and the source, +and for `f : Spec B ⟶ Spec A`, it is equivalent to the ring hom property `Q` on `Γ(f)`. + +For `HasRingHomProperty P Q` and `f : X ⟶ Y`, we provide these API lemmas: +- `AlgebraicGeometry.HasRingHomProperty.iff_appLE`: + `P f` if and only if `Q (f.appLE U V _)` for all affine `U : Opens Y` and `V : Opens X`. +- `AlgebraicGeometry.HasRingHomProperty.iff_of_source_openCover`: + If `Y` is affine, `P f ↔ ∀ i, Q ((𝒰.map i ≫ f).app ⊤)` for an affine open cover `𝒰` of `X`. +- `AlgebraicGeometry.HasRingHomProperty.iff_of_isAffine`: + If `X` and `Y` are affine, then `P f ↔ Q (f.app ⊤)`. +- `AlgebraicGeometry.HasRingHomProperty.Spec_iff`: + `P (Spec.map φ) ↔ Q φ` +- `AlgebraicGeometry.HasRingHomProperty.iff_of_iSup_eq_top`: + If `Y` is affine, `P f ↔ ∀ i, Q (f.appLE ⊤ (U i) _)` for a family `U` of affine opens of `X`. +- `AlgebraicGeometry.HasRingHomProperty.of_isOpenImmersion`: + If `f` is an open immersion then `P f`. +- `AlgebraicGeometry.HasRingHomProperty.stableUnderBaseChange`: + If `Q` is stable under base change, then so is `P`. + +We also provide the instances `P.IsMultiplicative`, `P.IsStableUnderComposition`, +`IsLocalAtTarget P`, `IsLocalAtSource P`. -/ @@ -38,81 +61,24 @@ universe u open CategoryTheory Opposite TopologicalSpace CategoryTheory.Limits AlgebraicGeometry -variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop) - namespace RingHom -variable {P} -theorem RespectsIso.basicOpen_iff (hP : RespectsIso @P) {X Y : Scheme.{u}} [IsAffine X] [IsAffine Y] - (f : X ⟶ Y) (r : Y.presheaf.obj (Opposite.op ⊤)) : - P (Scheme.Γ.map (f ∣_ Y.basicOpen r).op) ↔ - P (@IsLocalization.Away.map (Y.presheaf.obj (Opposite.op ⊤)) _ - (Y.presheaf.obj (Opposite.op <| Y.basicOpen r)) _ _ (X.presheaf.obj (Opposite.op ⊤)) _ - (X.presheaf.obj (Opposite.op <| X.basicOpen (Scheme.Γ.map f.op r))) _ _ - (Scheme.Γ.map f.op) r _ <| @isLocalization_away_of_isAffine X _ (Scheme.Γ.map f.op r)) := by - rw [Γ_map_morphismRestrict, hP.cancel_left_isIso, hP.cancel_right_isIso, - ← hP.cancel_right_isIso (f.val.c.app (Opposite.op (Y.basicOpen r))) - (X.presheaf.map (eqToHom (Scheme.preimage_basicOpen f r).symm).op), ← eq_iff_iff] - congr - delta IsLocalization.Away.map - refine IsLocalization.ringHom_ext (Submonoid.powers r) ?_ - generalize_proofs - haveI i1 := @isLocalization_away_of_isAffine X _ (Scheme.Γ.map f.op r) - -- Porting note: needs to be very explicit here - convert - (@IsLocalization.map_comp (hy := ‹_ ≤ _›) (Y.presheaf.obj <| Opposite.op (Scheme.basicOpen Y r)) - _ _ (isLocalization_away_of_isAffine _) _ _ _ i1).symm using 1 - change Y.presheaf.map _ ≫ _ = _ ≫ X.presheaf.map _ - rw [f.val.c.naturality_assoc] - dsimp - simp only [← X.presheaf.map_comp] - congr 1 - -theorem RespectsIso.basicOpen_iff_localization (hP : RespectsIso @P) {X Y : Scheme.{u}} [IsAffine X] - [IsAffine Y] (f : X ⟶ Y) (r : Y.presheaf.obj (Opposite.op ⊤)) : - P (Scheme.Γ.map (f ∣_ Y.basicOpen r).op) ↔ P (Localization.awayMap (Scheme.Γ.map f.op) r) := by - refine (hP.basicOpen_iff _ _).trans ?_ - -- Porting note: was a one line term mode proof, but this `dsimp` is vital so the term mode - -- one liner is not possible - dsimp - rw [← hP.is_localization_away_iff] - -@[deprecated (since := "2024-03-02")] alias -RespectsIso.ofRestrict_morphismRestrict_iff_of_isAffine := RespectsIso.basicOpen_iff_localization - -theorem RespectsIso.ofRestrict_morphismRestrict_iff (hP : RingHom.RespectsIso @P) {X Y : Scheme.{u}} - [IsAffine Y] (f : X ⟶ Y) (r : Y.presheaf.obj (Opposite.op ⊤)) (U : Opens X.carrier) - (hU : IsAffineOpen U) {V : Opens _} - (e : V = (Scheme.ιOpens <| f ⁻¹ᵁ Y.basicOpen r) ⁻¹ᵁ U) : - P (Scheme.Γ.map (Scheme.ιOpens V ≫ f ∣_ Y.basicOpen r).op) ↔ - P (Localization.awayMap (Scheme.Γ.map (Scheme.ιOpens U ≫ f).op) r) := by - subst e - refine (hP.cancel_right_isIso _ - (Scheme.Γ.mapIso (Scheme.restrictRestrictComm _ _ _).op).inv).symm.trans ?_ - haveI : IsAffine _ := hU - rw [← hP.basicOpen_iff_localization, iff_iff_eq] - congr 1 - simp only [Functor.mapIso_inv, Iso.op_inv, ← Functor.map_comp, ← op_comp, morphismRestrict_comp] - rw [← Category.assoc] - congr 3 - rw [← cancel_mono (Scheme.ιOpens _), Category.assoc, Scheme.restrictRestrictComm, - IsOpenImmersion.isoOfRangeEq_inv_fac, morphismRestrict_ι] - -theorem StableUnderBaseChange.Γ_pullback_fst (hP : StableUnderBaseChange @P) (hP' : RespectsIso @P) +variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop) + +theorem StableUnderBaseChange.pullback_fst_app_top + (hP : StableUnderBaseChange P) (hP' : RespectsIso P) {X Y S : Scheme} [IsAffine X] [IsAffine Y] [IsAffine S] (f : X ⟶ S) (g : Y ⟶ S) - (H : P (Scheme.Γ.map g.op)) : P (Scheme.Γ.map (pullback.fst f g).op) := by + (H : P (g.app ⊤)) : P ((pullback.fst f g).app ⊤) := by -- Porting note (#11224): change `rw` to `erw` erw [← PreservesPullback.iso_inv_fst AffineScheme.forgetToScheme (AffineScheme.ofHom f) (AffineScheme.ofHom g)] - rw [op_comp, Functor.map_comp, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] - have := - _root_.congr_arg Quiver.Hom.unop + rw [Scheme.comp_app, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] + have := congr_arg Quiver.Hom.unop (PreservesPullback.iso_hom_fst AffineScheme.Γ.rightOp (AffineScheme.ofHom f) (AffineScheme.ofHom g)) - simp only [Quiver.Hom.unop_op, Functor.rightOp_map, unop_comp] at this - delta AffineScheme.Γ at this - simp only [Quiver.Hom.unop_op, Functor.comp_map, AffineScheme.forgetToScheme_map, - Functor.op_map] at this + simp only [AffineScheme.Γ, Functor.rightOp_obj, Functor.comp_obj, Functor.op_obj, unop_comp, + AffineScheme.forgetToScheme_obj, Scheme.Γ_obj, Functor.rightOp_map, Functor.comp_map, + Functor.op_map, Quiver.Hom.unop_op, AffineScheme.forgetToScheme_map, Scheme.Γ_map] at this rw [← this, hP'.cancel_right_isIso, ← pushoutIsoUnopPullback_inl_hom, hP'.cancel_right_isIso] exact hP.pushout_inl _ hP' _ _ H @@ -120,422 +86,278 @@ end RingHom namespace AlgebraicGeometry +section affineLocally + +variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop) + /-- For `P` a property of ring homomorphisms, `sourceAffineLocally P` holds for `f : X ⟶ Y` whenever `P` holds for the restriction of `f` on every affine open subset of `X`. -/ def sourceAffineLocally : AffineTargetMorphismProperty := fun X _ f _ => - ∀ U : X.affineOpens, P (Scheme.Γ.map (X.ofRestrict U.1.openEmbedding ≫ f).op) + ∀ U : X.affineOpens, P (f.appLE ⊤ U le_top) /-- For `P` a property of ring homomorphisms, `affineLocally P` holds for `f : X ⟶ Y` if for each affine open `U = Spec A ⊆ Y` and `V = Spec B ⊆ f ⁻¹' U`, the ring hom `A ⟶ B` satisfies `P`. Also see `affineLocally_iff_affineOpens_le`. -/ abbrev affineLocally : MorphismProperty Scheme.{u} := - targetAffineLocally (sourceAffineLocally @P) - -variable {P} + targetAffineLocally (sourceAffineLocally P) -theorem sourceAffineLocally_respectsIso (h₁ : RingHom.RespectsIso @P) : - (sourceAffineLocally @P).toProperty.RespectsIso := by +theorem sourceAffineLocally_respectsIso (h₁ : RingHom.RespectsIso P) : + (sourceAffineLocally P).toProperty.RespectsIso := by apply AffineTargetMorphismProperty.respectsIso_mk · introv H U - rw [← h₁.cancel_right_isIso _ (Scheme.Γ.map (Scheme.restrictMapIso e.inv U.1).hom.op), ← - Functor.map_comp, ← op_comp] - convert H ⟨_, U.prop.preimage_of_isIso e.inv⟩ using 3 - rw [IsOpenImmersion.isoOfRangeEq_hom_fac_assoc, Category.assoc, - e.inv_hom_id_assoc] + have : IsIso (e.hom.appLE (e.hom ''ᵁ U) U.1 (e.hom.preimage_image_eq _).ge) := + inferInstanceAs (IsIso (e.hom.app _ ≫ + X.presheaf.map (eqToHom (e.hom.preimage_image_eq _).symm).op)) + rw [← Scheme.appLE_comp_appLE _ _ ⊤ (e.hom ''ᵁ U) U.1 le_top (e.hom.preimage_image_eq _).ge, + h₁.cancel_right_isIso] + exact H ⟨_, U.prop.image_of_isOpenImmersion e.hom⟩ · introv H U - rw [← Category.assoc, op_comp, Functor.map_comp, h₁.cancel_left_isIso] + rw [Scheme.comp_appLE, h₁.cancel_left_isIso] exact H U -theorem affineLocally_respectsIso (h : RingHom.RespectsIso @P) : (affineLocally @P).RespectsIso := by - have := sourceAffineLocally_respectsIso h - delta affineLocally - infer_instance - -theorem affineLocally_iff_affineOpens_le - (hP : RingHom.RespectsIso @P) {X Y : Scheme.{u}} (f : X ⟶ Y) : - affineLocally.{u} (@P) f ↔ - ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ (Opens.map f.1.base).obj U.1), - P (f.appLE _ _ e) := by - apply forall_congr' - intro U - delta sourceAffineLocally - simp_rw [op_comp, Scheme.Γ.map_comp, Γ_map_morphismRestrict, Category.assoc, Scheme.Γ_map_op, - hP.cancel_left_isIso (Y.presheaf.map (eqToHom _).op)] - constructor - · intro H V e - let U' := (Opens.map f.val.base).obj U.1 - have e'' : (Scheme.Hom.opensFunctor (X.ofRestrict U'.openEmbedding)).obj - (X.ofRestrict U'.openEmbedding⁻¹ᵁ V) = V := by - ext1; refine Set.image_preimage_eq_inter_range.trans (Set.inter_eq_left.mpr ?_) - erw [Subtype.range_val] - exact e - have h : X.ofRestrict U'.openEmbedding ⁻¹ᵁ ↑V ∈ Scheme.affineOpens (X.restrict _) := by - apply (X.ofRestrict U'.openEmbedding).isAffineOpen_iff_of_isOpenImmersion.mp - -- Porting note: was convert V.2 - rw [e''] - convert V.2 - have := H ⟨(Opens.map (X.ofRestrict U'.openEmbedding).1.base).obj V.1, h⟩ - rw [← hP.cancel_right_isIso _ (X.presheaf.map (eqToHom _)), Scheme.Hom.appLE, - Category.assoc, ← X.presheaf.map_comp] - · convert this using 1 - congr 1 - rw [X.presheaf.map_comp] - · simp only [Scheme.ofRestrict_val_c_app, Scheme.restrict_presheaf_map, ← X.presheaf.map_comp] - congr 1 - · dsimp only [Functor.op, unop_op] - rw [Opens.openEmbedding_obj_top] - congr 1 - exact e''.symm - · intro H V - specialize H ⟨_, V.2.image_of_isOpenImmersion (X.ofRestrict _)⟩ (Subtype.coe_image_subset _ _) - rw [← hP.cancel_right_isIso _ (X.presheaf.map (eqToHom _)), Category.assoc] - · convert H - simp only [Scheme.ofRestrict_val_c_app, Scheme.restrict_presheaf_map, ← X.presheaf.map_comp] - congr 1 - · dsimp only [Functor.op, unop_op]; rw [Opens.openEmbedding_obj_top] - --- The `IsLocalization` is not necessary, but Lean errors if it is omitted. -@[nolint unusedHavesSuffices] -theorem scheme_restrict_basicOpen_of_localizationPreserves (h₁ : RingHom.RespectsIso @P) - (h₂ : RingHom.LocalizationPreserves @P) {X Y : Scheme.{u}} [IsAffine Y] (f : X ⟶ Y) - (r : Y.presheaf.obj (op ⊤)) (H : sourceAffineLocally (@P) f) - (U : (X.restrict ((Opens.map f.1.base).obj <| Y.basicOpen r).openEmbedding).affineOpens) : - P (Scheme.Γ.map ((X.restrict ((Opens.map f.1.base).obj <| - Y.basicOpen r).openEmbedding).ofRestrict U.1.openEmbedding ≫ f ∣_ Y.basicOpen r).op) := by - specialize H ⟨_, U.2.image_of_isOpenImmersion (X.ofRestrict _)⟩ - letI i1 : Algebra (Y.presheaf.obj <| Opposite.op ⊤) (Localization.Away r) := - OreLocalization.instAlgebra - haveI : IsLocalization.Away r (Localization.Away r) := inferInstance - exact (h₁.ofRestrict_morphismRestrict_iff f r - ((Scheme.Hom.opensFunctor - (X.ofRestrict ((Opens.map f.1.base).obj <| Y.basicOpen r).openEmbedding)).obj U.1) - (IsAffineOpen.image_of_isOpenImmersion U.2 - (X.ofRestrict ((Opens.map f.1.base).obj <| Y.basicOpen r).openEmbedding)) - (Opens.ext (Set.preimage_image_eq _ Subtype.coe_injective).symm)).mpr (h₂.away r H) - -theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso @P) - (h₂ : RingHom.LocalizationPreserves @P) (h₃ : RingHom.OfLocalizationSpan @P) : - (sourceAffineLocally @P).IsLocal := by +theorem affineLocally_respectsIso (h : RingHom.RespectsIso P) : (affineLocally P).RespectsIso := + letI := sourceAffineLocally_respectsIso P h + inferInstance + +open Scheme in +theorem sourceAffineLocally_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) + (U : Y.Opens) (hU : IsAffineOpen U) : + @sourceAffineLocally P _ _ (f ∣_ U) hU ↔ + ∀ (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U), P (f.appLE U V e) := by + dsimp only [sourceAffineLocally] + simp only [morphismRestrict_appLE] + rw [(affineOpensRestrict (f ⁻¹ᵁ U)).forall_congr_left, Subtype.forall] + refine forall₂_congr fun V h ↦ ?_ + have := (affineOpensRestrict (f ⁻¹ᵁ U)).apply_symm_apply ⟨V, h⟩ + exact f.appLE_congr _ (Opens.ι_image_top _) congr($(this).1.1) _ + +theorem affineLocally_iff_affineOpens_le {X Y : Scheme.{u}} (f : X ⟶ Y) : + affineLocally.{u} P f ↔ + ∀ (U : Y.affineOpens) (V : X.affineOpens) (e : V.1 ≤ f ⁻¹ᵁ U.1), P (f.appLE U V e) := + forall_congr' fun U ↦ sourceAffineLocally_morphismRestrict P f U U.2 + +theorem sourceAffineLocally_isLocal (h₁ : RingHom.RespectsIso P) + (h₂ : RingHom.LocalizationPreserves P) (h₃ : RingHom.OfLocalizationSpan P) : + (sourceAffineLocally P).IsLocal := by constructor - · exact sourceAffineLocally_respectsIso h₁ - · introv H U - apply scheme_restrict_basicOpen_of_localizationPreserves h₁ h₂; assumption + · exact sourceAffineLocally_respectsIso P h₁ + · intro X Y _ f r H + rw [sourceAffineLocally_morphismRestrict] + intro U hU + have : X.basicOpen (f.appLE ⊤ U (by simp) r) = U := by + simp only [Scheme.Hom.appLE, Opens.map_top, CommRingCat.coe_comp_of, RingHom.coe_comp, + Function.comp_apply] + rw [Scheme.basicOpen_res] + simpa using hU + rw [← f.appLE_congr _ rfl this P, + IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2 _ r] + apply (config := { allowSynthFailures := true }) h₂.away + exact H U · introv hs hs' U apply h₃ _ _ hs intro r - have := hs' r ⟨(Opens.map (X.ofRestrict _).1.base).obj U.1, ?_⟩ - · rwa [h₁.ofRestrict_morphismRestrict_iff] at this - · exact U.2 - · rfl - · suffices ∀ (V) (_ : V = (Opens.map f.val.base).obj (Y.basicOpen r.val)), - IsAffineOpen ((Opens.map (X.ofRestrict V.openEmbedding).1.base).obj U.1) by - exact this _ rfl - intro V hV - rw [Scheme.preimage_basicOpen] at hV - subst hV - exact U.2.ιOpens_basicOpen_preimage (Scheme.Γ.map f.op r.1) - -variable (hP : RingHom.PropertyIsLocal @P) - -theorem sourceAffineLocally_of_source_open_cover_aux (h₁ : RingHom.RespectsIso @P) - (h₃ : RingHom.OfLocalizationSpanTarget @P) {X Y : Scheme.{u}} (f : X ⟶ Y) (U : X.affineOpens) - (s : Set (X.presheaf.obj (op U.1))) (hs : Ideal.span s = ⊤) - (hs' : ∀ r : s, P (Scheme.Γ.map (Scheme.ιOpens (X.basicOpen r.1) ≫ f).op)) : - P (Scheme.Γ.map (Scheme.ιOpens U ≫ f).op) := by - apply_fun Ideal.map (X.presheaf.map (eqToHom U.1.openEmbedding_obj_top).op) at hs - rw [Ideal.map_span, Ideal.map_top] at hs - apply h₃.ofIsLocalization h₁ _ _ hs - rintro ⟨s, r, hr, hs⟩ - refine ⟨_, _, _, @AlgebraicGeometry.Γ_restrict_isLocalization (X ∣_ᵤ U.1) U.2 s, ?_⟩ - rw [RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, ← Functor.map_comp, - ← op_comp, ← h₁.cancel_right_isIso _ (Scheme.Γ.mapIso (Scheme.restrictRestrict _ _ _).op).inv] - subst hs - rw [← h₁.cancel_right_isIso _ - (Scheme.Γ.mapIso (Scheme.restrictIsoOfEq _ (Scheme.map_basicOpen_map _ _ _)).op).inv] - simp only [Functor.mapIso_inv, Iso.op_inv, ← Functor.map_comp, ← op_comp, - Scheme.restrictRestrict_inv_restrict_restrict_assoc, Scheme.restrictIsoOfEq, - IsOpenImmersion.isoOfRangeEq_inv_fac_assoc] - exact hs' ⟨r, hr⟩ - -theorem isOpenImmersionCat_comp_of_sourceAffineLocally (h₁ : RingHom.RespectsIso @P) - {X Y Z : Scheme.{u}} [IsAffine X] [IsAffine Z] (f : X ⟶ Y) [IsOpenImmersion f] (g : Y ⟶ Z) - (h₂ : sourceAffineLocally (@P) g) : P (Scheme.Γ.map (f ≫ g).op) := by - rw [← h₁.cancel_right_isIso _ - (Scheme.Γ.map (IsOpenImmersion.isoOfRangeEq (Y.ofRestrict _) f _).hom.op), - ← Functor.map_comp, ← op_comp] - · convert h₂ ⟨_, isAffineOpen_opensRange f⟩ using 3 - · rw [IsOpenImmersion.isoOfRangeEq_hom_fac_assoc] - exact Subtype.range_coe - -end AlgebraicGeometry + 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 simpa [Scheme.Hom.appLE] using Scheme.basicOpen_restrict _ _ _) + rwa [IsAffineOpen.appLE_eq_away_map f (isAffineOpen_top Y) U.2, + ← h₁.is_localization_away_iff] at this + +end affineLocally + +/-- +`HasRingHomProperty P Q` is a type class asserting that `P` is local at the target and the source, +and for `f : Spec B ⟶ Spec A`, it is equivalent to the ring hom property `Q`. +To make the proofs easier, we state it instead as +1. `Q` is local (See `RingHom.PropertyIsLocal`) +2. `P f` if and only if `Q` holds for every `Γ(Y, U) ⟶ Γ(X, V)` for all affine `U`, `V`. +See `HasRingHomProperty.iff_appLE`. +-/ +class HasRingHomProperty (P : MorphismProperty Scheme.{u}) + (Q : outParam (∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop)) : Prop where + isLocal_ringHomProperty : RingHom.PropertyIsLocal Q + eq_affineLocally' : P = affineLocally Q + +namespace HasRingHomProperty + +variable (P : MorphismProperty Scheme.{u}) {Q} [HasRingHomProperty P Q] +variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) + +lemma eq_affineLocally : P = affineLocally Q := eq_affineLocally' + +@[local instance] +lemma HasAffineProperty : HasAffineProperty P (sourceAffineLocally Q) where + isLocal_affineProperty := sourceAffineLocally_isLocal _ + (isLocal_ringHomProperty P).respectsIso + (isLocal_ringHomProperty P).LocalizationPreserves + (isLocal_ringHomProperty P).ofLocalizationSpan + eq_targetAffineLocally' := eq_affineLocally P + +theorem appLE (H : P f) (U : Y.affineOpens) (V : X.affineOpens) (e) : Q (f.appLE U V e) := by + rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H + exact H _ _ _ + +theorem app_top (H : P f) [IsAffine X] [IsAffine Y] : Q (f.app ⊤) := by + rw [Scheme.Hom.app_eq_appLE] + exact appLE P f H ⟨_, isAffineOpen_top _⟩ ⟨_, isAffineOpen_top _⟩ _ + +theorem comp_of_isOpenImmersion [IsOpenImmersion f] (H : P g) : + P (f ≫ g) := by + rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H ⊢ + intro U V e + have : IsIso (f.appLE (f ''ᵁ V) V.1 (f.preimage_image_eq _).ge) := + inferInstanceAs (IsIso (f.app _ ≫ + X.presheaf.map (eqToHom (f.preimage_image_eq _).symm).op)) + rw [← Scheme.appLE_comp_appLE _ _ _ (f ''ᵁ V) V.1 + (Set.image_subset_iff.mpr e) (f.preimage_image_eq _).ge, + (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso] + exact H _ ⟨_, V.2.image_of_isOpenImmersion _⟩ _ + +variable {P f} + +lemma iff_appLE : P f ↔ ∀ (U : Y.affineOpens) (V : X.affineOpens) (e), Q (f.appLE U V e) := by + rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] + +theorem of_source_openCover [IsAffine Y] + (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).app ⊤)) : + P f := by + rw [HasAffineProperty.iff_of_isAffine (P := P)] + intro U + let S i : X.affineOpens := ⟨_, isAffineOpen_opensRange (𝒰.map i)⟩ + induction U using of_affine_open_cover S 𝒰.iSup_opensRange with + | basicOpen U r H => + simp_rw [Scheme.affineBasicOpen_coe, + ← f.appLE_map (U := ⊤) le_top (homOfLE (X.basicOpen_le r)).op] + apply (isLocal_ringHomProperty P).StableUnderComposition _ _ H + have := U.2.isLocalization_basicOpen r + apply (isLocal_ringHomProperty P).HoldsForLocalizationAway _ r + | openCover U s hs H => + apply (isLocal_ringHomProperty P).OfLocalizationSpanTarget.ofIsLocalization + (isLocal_ringHomProperty P).respectsIso _ _ hs + rintro r + refine ⟨_, _, _, IsAffineOpen.isLocalization_basicOpen U.2 r, ?_⟩ + rw [RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, Scheme.Hom.appLE_map] + exact H r + | hU i => + specialize H i + rw [← (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso _ + ((IsOpenImmersion.isoOfRangeEq (𝒰.map i) (S i).1.ι + Subtype.range_coe.symm).inv.app _), ← Scheme.comp_app, + IsOpenImmersion.isoOfRangeEq_inv_fac_assoc, Scheme.comp_app, + Scheme.Opens.ι_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map] at H + exact (f.appLE_congr _ rfl (by simp) Q).mp H + +theorem iff_of_source_openCover [IsAffine Y] (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] : + P f ↔ ∀ i, Q ((𝒰.map i ≫ f).app ⊤) := + ⟨fun H i ↦ app_top P _ (comp_of_isOpenImmersion P (𝒰.map i) f H), of_source_openCover 𝒰⟩ + +theorem iff_of_isAffine [IsAffine X] [IsAffine Y] : + P f ↔ Q (f.app ⊤) := by + rw [iff_of_source_openCover (P := P) (Scheme.openCoverOfIsIso.{u} (𝟙 _))] + simp + +theorem Spec_iff {R S : CommRingCat.{u}} {φ : R ⟶ S} : + P (Spec.map φ) ↔ Q φ := by + have H := (isLocal_ringHomProperty P).respectsIso + rw [iff_of_isAffine (P := P), ← H.cancel_right_isIso _ (Scheme.ΓSpecIso _).hom, + Scheme.ΓSpecIso_naturality, H.cancel_left_isIso] + +theorem of_iSup_eq_top [IsAffine Y] {ι : Type*} + (U : ι → X.affineOpens) (hU : ⨆ i, (U i : Opens X) = ⊤) + (H : ∀ i, Q (f.appLE ⊤ (U i).1 le_top)) : + P f := by + have (i) : IsAffine ((X.openCoverOfSuprEqTop _ hU).obj i) := (U i).2 + refine of_source_openCover (X.openCoverOfSuprEqTop _ hU) fun i ↦ ?_ + simpa [Scheme.Hom.app_eq_appLE] using (f.appLE_congr _ rfl (by simp) Q).mp (H i) + +theorem iff_of_iSup_eq_top [IsAffine Y] {ι : Type*} + (U : ι → X.affineOpens) (hU : ⨆ i, (U i : Opens X) = ⊤) : + P f ↔ ∀ i, Q (f.appLE ⊤ (U i).1 le_top) := + ⟨fun H _ ↦ appLE P f H ⟨_, isAffineOpen_top _⟩ _ le_top, of_iSup_eq_top U hU⟩ + +instance : IsLocalAtSource P := by + apply HasAffineProperty.isLocalAtSource + intros X Y f _ 𝒰 + simp_rw [← HasAffineProperty.iff_of_isAffine (P := P), + iff_of_source_openCover 𝒰.affineRefinement.openCover, + fun i ↦ iff_of_source_openCover (P := P) (f := 𝒰.map i ≫ f) (𝒰.obj i).affineCover] + simp [Scheme.OpenCover.affineRefinement] + +instance : P.ContainsIdentities where + id_mem X := by + rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := P) _ (iSup_affineOpens_eq_top _)] + intro U + have : IsAffine (𝟙 X ⁻¹ᵁ U.1) := U.2 + rw [morphismRestrict_id, iff_of_isAffine (P := P), Scheme.id_app] + exact (isLocal_ringHomProperty P).HoldsForLocalizationAway.of_bijective _ _ + Function.bijective_id + +instance : P.IsStableUnderComposition where + comp_mem {X Y Z} f g hf hg := by + wlog hZ : IsAffine Z generalizing X Y Z + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := P) _ (iSup_affineOpens_eq_top _)] + intro U + rw [morphismRestrict_comp] + exact this _ _ (IsLocalAtTarget.restrict hf _) (IsLocalAtTarget.restrict hg _) U.2 + wlog hY : IsAffine Y generalizing X Y + · rw [IsLocalAtSource.iff_of_openCover (P := P) (Y.affineCover.pullbackCover f)] + intro i + rw [← Scheme.OpenCover.pullbackHom_map_assoc] + exact this _ _ (IsLocalAtTarget.of_isPullback (.of_hasPullback _ _) hf) + (comp_of_isOpenImmersion _ _ _ hg) inferInstance + wlog hX : IsAffine X generalizing X + · rw [IsLocalAtSource.iff_of_openCover (P := P) X.affineCover] + intro i + rw [← Category.assoc] + exact this _ (comp_of_isOpenImmersion _ _ _ hf) inferInstance + rw [iff_of_isAffine (P := P)] at hf hg ⊢ + exact (isLocal_ringHomProperty P).StableUnderComposition _ _ hg hf -open AlgebraicGeometry +theorem of_comp + (H : ∀ {R S T : Type u} [CommRing R] [CommRing S] [CommRing T], + ∀ (f : R →+* S) (g : S →+* T), Q (g.comp f) → Q g) + {X Y Z : Scheme.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (h : P (f ≫ g)) : P f := by + wlog hZ : IsAffine Z generalizing X Y Z + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := P) _ + (g.preimage_iSup_eq_top (iSup_affineOpens_eq_top Z))] + intro U + have H := IsLocalAtTarget.restrict h U.1 + rw [morphismRestrict_comp] at H + exact this H inferInstance + wlog hY : IsAffine Y generalizing X Y + · rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := P) _ (iSup_affineOpens_eq_top Y)] + intro U + have H := comp_of_isOpenImmersion P (f ⁻¹ᵁ U.1).ι (f ≫ g) h + rw [← morphismRestrict_ι_assoc] at H + exact this H inferInstance + wlog hY : IsAffine X generalizing X + · rw [IsLocalAtSource.iff_of_iSup_eq_top (P := P) _ (iSup_affineOpens_eq_top X)] + intro U + have H := comp_of_isOpenImmersion P U.1.ι (f ≫ g) h + rw [← Category.assoc] at H + exact this H inferInstance + rw [iff_of_isAffine (P := P)] at h ⊢ + exact H _ _ h -namespace RingHom.PropertyIsLocal +instance : P.IsMultiplicative where -variable {P} (hP : RingHom.PropertyIsLocal @P) +lemma of_isOpenImmersion [IsOpenImmersion f] : P f := IsLocalAtSource.of_isOpenImmersion f -theorem sourceAffineLocally_of_source_openCover {X Y : Scheme.{u}} (f : X ⟶ Y) [IsAffine Y] - (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, P (Scheme.Γ.map (𝒰.map i ≫ f).op)) : - sourceAffineLocally (@P) f := by - let S i := (⟨⟨Set.range (𝒰.map i).1.base, (𝒰.IsOpen i).base_open.isOpen_range⟩, - isAffineOpen_opensRange (𝒰.map i)⟩ : X.affineOpens) - intro U - -- Porting note: here is what we are eliminating into Lean - apply of_affine_open_cover - (P := fun V => P (Scheme.Γ.map (X.ofRestrict (Opens.openEmbedding V.val) ≫ f).op)) S - 𝒰.iSup_opensRange - · intro U r H - -- Porting note: failing on instance synthesis for an (unspecified) meta variable - -- made φ explicit and forced to use dsimp in the proof - convert hP.StableUnderComposition - (S := Scheme.Γ.obj (Opposite.op (X.restrict <| Opens.openEmbedding U.val))) - (T := Scheme.Γ.obj (Opposite.op (X.restrict <| Opens.openEmbedding (X.basicOpen r)))) - ?_ ?_ H ?_ using 1 - swap - · refine X.presheaf.map - (@homOfLE _ _ ((IsOpenMap.functor _).obj _) ((IsOpenMap.functor _).obj _) ?_).op - dsimp - rw [Opens.openEmbedding_obj_top, Opens.openEmbedding_obj_top] - exact X.basicOpen_le _ - · rw [op_comp, op_comp, Functor.map_comp, Functor.map_comp] - refine (Eq.trans ?_ (Category.assoc (obj := CommRingCat) _ _ _).symm : _) - congr 1 - dsimp - refine Eq.trans ?_ (X.presheaf.map_comp _ _) - change X.presheaf.map _ = _ - congr! - -- Porting note: need to pass Algebra through explicitly - convert @HoldsForLocalizationAway _ hP _ - (Scheme.Γ.obj (Opposite.op (X.restrict (X.basicOpen r).openEmbedding))) _ _ ?_ - (X.presheaf.map (eqToHom U.1.openEmbedding_obj_top).op r) ?_ - · exact RingHom.algebraMap_toAlgebra - (R := Scheme.Γ.obj <| Opposite.op <| X.restrict (U.1.openEmbedding)) - (S := - Scheme.Γ.obj (Opposite.op <| X.restrict (X.affineBasicOpen r).1.openEmbedding)) _|>.symm - · dsimp [Scheme.Γ] - have := U.2 - rw [← U.1.openEmbedding_obj_top] at this - -- Porting note: the second argument of `IsLocalization.Away` is a type, and we want - -- to generate an equality, so using `typeEqs := true` to force allowing type equalities. - convert (config := {typeEqs := true, transparency := .default}) - this.isLocalization_basicOpen _ using 5 - all_goals rw [Opens.openEmbedding_obj_top]; exact (Scheme.basicOpen_res_eq _ _ _).symm - · introv hs hs' - exact sourceAffineLocally_of_source_open_cover_aux hP.respectsIso hP.2 _ _ _ hs hs' - · rintro i - specialize H i - rw [← hP.respectsIso.cancel_right_isIso _ - (Scheme.Γ.map - (IsOpenImmersion.isoOfRangeEq (𝒰.map i) (X.ofRestrict (S i).1.openEmbedding) - Subtype.range_coe.symm).inv.op)] at H - rwa [← Scheme.Γ.map_comp, ← op_comp, IsOpenImmersion.isoOfRangeEq_inv_fac_assoc] at H - -theorem affine_openCover_TFAE {X Y : Scheme.{u}} [IsAffine Y] (f : X ⟶ Y) : - List.TFAE - [sourceAffineLocally (@P) f, - ∃ (𝒰 : Scheme.OpenCover.{u} X) (_ : ∀ i, IsAffine (𝒰.obj i)), - ∀ i : 𝒰.J, P (Scheme.Γ.map (𝒰.map i ≫ f).op), - ∀ (𝒰 : Scheme.OpenCover.{u} X) [∀ i, IsAffine (𝒰.obj i)] (i : 𝒰.J), - P (Scheme.Γ.map (𝒰.map i ≫ f).op), - ∀ {U : Scheme} (g : U ⟶ X) [IsAffine U] [IsOpenImmersion g], - P (Scheme.Γ.map (g ≫ f).op)] := by - tfae_have 1 → 4 - · intro H U g _ hg - specialize H ⟨⟨_, hg.base_open.isOpen_range⟩, isAffineOpen_opensRange g⟩ - rw [← hP.respectsIso.cancel_right_isIso _ (Scheme.Γ.map (IsOpenImmersion.isoOfRangeEq g - (X.ofRestrict (Opens.openEmbedding ⟨_, hg.base_open.isOpen_range⟩)) - Subtype.range_coe.symm).hom.op), - ← Scheme.Γ.map_comp, ← op_comp, IsOpenImmersion.isoOfRangeEq_hom_fac_assoc] at H - exact H - tfae_have 4 → 3 - · intro H 𝒰 _ i; apply H - tfae_have 3 → 2 - · intro H; exact ⟨X.affineCover, inferInstance, H _⟩ - tfae_have 2 → 1 - · rintro ⟨𝒰, _, h𝒰⟩ - exact sourceAffineLocally_of_source_openCover hP f 𝒰 h𝒰 - tfae_finish - -theorem openCover_TFAE {X Y : Scheme.{u}} [IsAffine Y] (f : X ⟶ Y) : - List.TFAE - [sourceAffineLocally (@P) f, - ∃ 𝒰 : Scheme.OpenCover.{u} X, ∀ i : 𝒰.J, sourceAffineLocally (@P) (𝒰.map i ≫ f), - ∀ (𝒰 : Scheme.OpenCover.{u} X) (i : 𝒰.J), sourceAffineLocally (@P) (𝒰.map i ≫ f), - ∀ {U : Scheme} (g : U ⟶ X) [IsOpenImmersion g], sourceAffineLocally (@P) (g ≫ f)] := by - tfae_have 1 → 4 - · intro H U g hg V - -- Porting note: this has metavariable if I put it directly into rw - have := (hP.affine_openCover_TFAE f).out 0 3 - rw [this] at H - haveI : IsAffine _ := V.2 - rw [← Category.assoc] - -- Porting note: Lean could find this previously - have : IsOpenImmersion <| (Scheme.ofRestrict U (Opens.openEmbedding V.val)) ≫ g := - LocallyRingedSpace.IsOpenImmersion.comp _ _ - apply H - tfae_have 4 → 3 - · intro H 𝒰 _ i; apply H - tfae_have 3 → 2 - · intro H; exact ⟨X.affineCover, H _⟩ - tfae_have 2 → 1 - · rintro ⟨𝒰, h𝒰⟩ - -- Porting note: this has metavariable if I put it directly into rw - have := (hP.affine_openCover_TFAE f).out 0 1 - rw [this] - refine ⟨𝒰.bind fun _ => Scheme.affineCover _, ?_, ?_⟩ - · intro i; dsimp; infer_instance - · intro i - specialize h𝒰 i.1 - -- Porting note: this has metavariable if I put it directly into rw - have := (hP.affine_openCover_TFAE (𝒰.map i.fst ≫ f)).out 0 3 - rw [this] at h𝒰 - -- Porting note: this was discharged after the apply previously - have : IsAffine (Scheme.OpenCover.obj - (Scheme.OpenCover.bind 𝒰 fun x ↦ Scheme.affineCover (Scheme.OpenCover.obj 𝒰 x)) i) := by - dsimp; infer_instance - apply @h𝒰 _ (show _ from _) - tfae_finish - -theorem sourceAffineLocally_comp_of_isOpenImmersion {X Y Z : Scheme.{u}} [IsAffine Z] (f : X ⟶ Y) - (g : Y ⟶ Z) [IsOpenImmersion f] (H : sourceAffineLocally (@P) g) : - sourceAffineLocally (@P) (f ≫ g) := by - -- Porting note: more tfae mis-behavior - have := (hP.openCover_TFAE g).out 0 3 - apply this.mp H - -theorem source_affine_openCover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) [IsAffine Y] - (𝒰 : Scheme.OpenCover.{u} X) [∀ i, IsAffine (𝒰.obj i)] : - sourceAffineLocally (@P) f ↔ ∀ i, P (Scheme.Γ.map (𝒰.map i ≫ f).op) := by - -- Porting note: seems like TFAE is misbehaving; this used to be pure term proof but - -- had strange failures where the output of TFAE turned into a metavariable when used despite - -- being correctly displayed in the infoview - refine ⟨fun H => ?_, fun H => ?_⟩ - · have h := (hP.affine_openCover_TFAE f).out 0 2 - apply h.mp - exact H - · have h := (hP.affine_openCover_TFAE f).out 1 0 - apply h.mp - use 𝒰 - -theorem isLocal_sourceAffineLocally : (sourceAffineLocally @P).IsLocal := - sourceAffineLocally_isLocal hP.respectsIso hP.LocalizationPreserves - (@RingHom.PropertyIsLocal.ofLocalizationSpan _ hP) - -theorem hasAffinePropertyAffineLocally : - HasAffineProperty (affineLocally @P) (sourceAffineLocally @P) where - isLocal_affineProperty := hP.isLocal_sourceAffineLocally - eq_targetAffineLocally' := rfl - -theorem affine_openCover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) (𝒰 : Scheme.OpenCover.{u} Y) - [∀ i, IsAffine (𝒰.obj i)] (𝒰' : ∀ i, Scheme.OpenCover.{u} ((𝒰.pullbackCover f).obj i)) - [∀ i j, IsAffine ((𝒰' i).obj j)] : - affineLocally (@P) f ↔ ∀ i j, P (Scheme.Γ.map ((𝒰' i).map j ≫ pullback.snd _ _).op) := - letI := hP.hasAffinePropertyAffineLocally - (HasAffineProperty.iff_of_openCover 𝒰).trans - (forall_congr' fun i => hP.source_affine_openCover_iff _ (𝒰' i)) - -theorem source_openCover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) (𝒰 : Scheme.OpenCover.{u} X) : - affineLocally (@P) f ↔ ∀ i, affineLocally (@P) (𝒰.map i ≫ f) := by - constructor - · intro H i U - rw [morphismRestrict_comp] - apply hP.sourceAffineLocally_comp_of_isOpenImmersion - apply H - · intro H U - haveI : IsAffine _ := U.2 - apply ((hP.openCover_TFAE (f ∣_ U.1)).out 1 0).mp - use 𝒰.pullbackCover (X.ofRestrict _) +lemma stableUnderBaseChange (hP : RingHom.StableUnderBaseChange Q) : P.StableUnderBaseChange := by + apply HasAffineProperty.stableUnderBaseChange + letI := HasAffineProperty.isLocal_affineProperty P + apply AffineTargetMorphismProperty.StableUnderBaseChange.mk + intros X Y S _ _ f g H + rw [← HasAffineProperty.iff_of_isAffine (P := P)] at H ⊢ + wlog hX : IsAffine Y generalizing Y + · rw [IsLocalAtSource.iff_of_openCover (P := P) + (Scheme.Pullback.openCoverOfRight Y.affineCover f g)] intro i - specialize H i U - rw [morphismRestrict_comp] at H - delta morphismRestrict at H - have := sourceAffineLocally_respectsIso hP.respectsIso - rw [Category.assoc, (sourceAffineLocally P).cancel_left_of_respectsIso, - ← (sourceAffineLocally P).cancel_left_of_respectsIso (pullbackSymmetry _ _).hom, - pullbackSymmetry_hom_comp_snd_assoc] at H - exact H - -theorem affineLocally_of_isOpenImmersion (hP : RingHom.PropertyIsLocal @P) {X Y : Scheme.{u}} - (f : X ⟶ Y) [hf : IsOpenImmersion f] : affineLocally (@P) f := by - intro U - haveI H : IsAffine _ := U.2 - rw [← Category.comp_id (f ∣_ U)] - apply hP.sourceAffineLocally_comp_of_isOpenImmersion - -- Porting note: need to excuse Lean from synthesizing an instance - rw [@source_affine_openCover_iff _ hP _ _ _ _ (Scheme.openCoverOfIsIso (𝟙 _)) (_)] - · intro i - let esto := Scheme.Γ.obj (Opposite.op (Y.restrict <| Opens.openEmbedding U.val)) - let eso := Scheme.Γ.obj (Opposite.op ((Scheme.openCoverOfIsIso - (𝟙 (Y.restrict <| Opens.openEmbedding U.val))).obj i)) - have := hP.HoldsForLocalizationAway - convert @this esto eso _ _ ?_ 1 ?_ - · exact (RingHom.algebraMap_toAlgebra (Scheme.Γ.map _)).symm - · exact - @IsLocalization.away_of_isUnit_of_bijective _ _ _ _ (_) _ isUnit_one Function.bijective_id - · intro; exact H - -theorem affineLocally_of_comp - (H : ∀ {R S T : Type u} [CommRing R] [CommRing S] [CommRing T], - ∀ (f : R →+* S) (g : S →+* T), P (g.comp f) → P g) - {X Y Z : Scheme.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (h : affineLocally (@P) (f ≫ g)) : - affineLocally (@P) f := by - let 𝒰 : ∀ i, ((Z.affineCover.pullbackCover (f ≫ g)).obj i).OpenCover := by - intro i - refine Scheme.OpenCover.bind ?_ fun i => Scheme.affineCover _ - apply Scheme.OpenCover.pushforwardIso _ - (pullbackRightPullbackFstIso g (Z.affineCover.map i) f).hom - apply Scheme.Pullback.openCoverOfRight - exact (pullback g (Z.affineCover.map i)).affineCover - have h𝒰 : ∀ i j, IsAffine ((𝒰 i).obj j) := by dsimp [𝒰]; infer_instance - let 𝒰' := (Z.affineCover.pullbackCover g).bind fun i => Scheme.affineCover _ - have h𝒰' : ∀ i, IsAffine (𝒰'.obj i) := by dsimp [𝒰']; infer_instance - rw [hP.affine_openCover_iff f 𝒰' fun i => Scheme.affineCover _] - rw [hP.affine_openCover_iff (f ≫ g) Z.affineCover 𝒰] at h - rintro ⟨i, j⟩ k - dsimp at i j k - specialize h i ⟨j, k⟩ - dsimp only [𝒰, 𝒰', Scheme.OpenCover.bind_map, Scheme.OpenCover.pushforwardIso_obj, - Scheme.Pullback.openCoverOfRight_obj, Scheme.OpenCover.pushforwardIso_map, - Scheme.Pullback.openCoverOfRight_map, Scheme.OpenCover.bind_obj, - Scheme.OpenCover.pullbackCover_obj, Scheme.OpenCover.pullbackCover_map] at h ⊢ - rw [Category.assoc, Category.assoc, pullbackRightPullbackFstIso_hom_snd, - pullback.lift_snd_assoc, Category.assoc, ← Category.assoc, op_comp, Functor.map_comp] at h - exact H _ _ h + simp only [Scheme.Pullback.openCoverOfRight_obj, Scheme.Pullback.openCoverOfRight_map, + limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.comp_id] + apply this _ (comp_of_isOpenImmersion _ _ _ H) inferInstance + rw [iff_of_isAffine (P := P)] at H ⊢ + exact hP.pullback_fst_app_top _ (isLocal_ringHomProperty P).respectsIso _ _ H -theorem affineLocally_isStableUnderComposition : (affineLocally @P).IsStableUnderComposition where - comp_mem {X Y S} f g hf hg := by - let 𝒰 : ∀ i, ((S.affineCover.pullbackCover (f ≫ g)).obj i).OpenCover := by - intro i - refine Scheme.OpenCover.bind ?_ fun i => Scheme.affineCover _ - apply Scheme.OpenCover.pushforwardIso _ - (pullbackRightPullbackFstIso g (S.affineCover.map i) f).hom - apply Scheme.Pullback.openCoverOfRight - exact (pullback g (S.affineCover.map i)).affineCover - -- Porting note: used to be - rw [hP.affine_openCover_iff (f ≫ g) S.affineCover _] - but - -- metavariables cause problems in the instance search - apply (@affine_openCover_iff _ hP _ _ (f ≫ g) S.affineCover _ ?_ ?_).mpr - rotate_left - · exact 𝒰 - · intro i j; dsimp [𝒰] at *; infer_instance - · rintro i ⟨j, k⟩ - dsimp at i j k - dsimp only [𝒰, Scheme.OpenCover.bind_map, Scheme.OpenCover.pushforwardIso_obj, - Scheme.Pullback.openCoverOfRight_obj, Scheme.OpenCover.pushforwardIso_map, - Scheme.Pullback.openCoverOfRight_map, Scheme.OpenCover.bind_obj] - rw [Category.assoc, Category.assoc, pullbackRightPullbackFstIso_hom_snd, - pullback.lift_snd_assoc, Category.assoc, ← Category.assoc, op_comp, Functor.map_comp] - apply hP.StableUnderComposition - · -- Porting note: used to be exact _|>. hg i j but that can't find an instance - apply hP.affine_openCover_iff _ _ _|>.mp - exact hg - · letI := hP.hasAffinePropertyAffineLocally - replace hf := HasAffineProperty.of_isPullback (.of_hasPullback _ - ((pullback g (S.affineCover.map i)).affineCover.map j ≫ pullback.fst _ _)) hf - -- Porting note: again strange behavior of TFAE - have := (hP.affine_openCover_TFAE - (pullback.snd f ((pullback g (S.affineCover.map i)).affineCover.map j ≫ - pullback.fst _ _))).out 0 3 - apply this.mp hf - -end RingHom.PropertyIsLocal +end HasRingHomProperty + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 6c6aa92d204a2..ad7dc34669b3f 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -95,7 +95,7 @@ variable {X : Scheme} /-- If a scheme `X` has a cover by affine opens whose sections are Noetherian rings, then `X` is locally Noetherian. -/ theorem isLocallyNoetherian_of_affine_cover {ι} {S : ι → X.affineOpens} - (hS : (⨆ i, S i : Opens X) = ⊤) + (hS : (⨆ i, S i : X.Opens) = ⊤) (hS' : ∀ i, IsNoetherianRing Γ(X, S i)) : IsLocallyNoetherian X := by refine ⟨fun U => ?_⟩ induction U using of_affine_open_cover S hS with @@ -116,7 +116,7 @@ are noetherian rings. See [Har77], Proposition II.3.2. -/ theorem isLocallyNoetherian_iff_of_iSup_eq_top {ι} {S : ι → X.affineOpens} - (hS : (⨆ i, S i : Opens X) = ⊤) : + (hS : (⨆ i, S i : X.Opens) = ⊤) : IsLocallyNoetherian X ↔ ∀ i, IsNoetherianRing Γ(X, S i) := ⟨fun _ i => IsLocallyNoetherian.component_noetherian (S i), isLocallyNoetherian_of_affine_cover hS⟩ @@ -183,11 +183,11 @@ lemma noetherianSpace_of_isAffine [IsAffine X] [IsNoetherianRing Γ(X, ⊤)] : NoetherianSpace X := (noetherianSpace_iff_of_homeomorph X.isoSpec.inv.homeomorph).mp inferInstance -lemma noetherianSpace_of_isAffineOpen (U : Opens X) (hU : IsAffineOpen U) +lemma noetherianSpace_of_isAffineOpen (U : X.Opens) (hU : IsAffineOpen U) [IsNoetherianRing Γ(X, U)] : NoetherianSpace U := by - have : IsNoetherianRing Γ(X ∣_ᵤ U, ⊤) := isNoetherianRing_of_ringEquiv _ - (X.restrictFunctorΓ.app (op U)).symm.commRingCatIsoToRingEquiv + have : IsNoetherianRing Γ(U, ⊤) := isNoetherianRing_of_ringEquiv _ + (Scheme.restrictFunctorΓ.app (op U)).symm.commRingCatIsoToRingEquiv exact @noetherianSpace_of_isAffine _ hU _ /-- Any open immersion `Z ⟶ X` with `X` locally Noetherian is quasi-compact. @@ -232,7 +232,7 @@ class IsNoetherian (X : Scheme) extends IsLocallyNoetherian X, CompactSpace X : /-- A scheme is Noetherian if and only if it is covered by finitely many affine opens whose sections are noetherian rings. -/ theorem isNoetherian_iff_of_finite_iSup_eq_top {ι} [Finite ι] {S : ι → X.affineOpens} - (hS : (⨆ i, S i : Opens X) = ⊤) : + (hS : (⨆ i, S i : X.Opens) = ⊤) : IsNoetherian X ↔ ∀ i, IsNoetherianRing Γ(X, S i) := by constructor · intro h i @@ -249,7 +249,7 @@ theorem isNoetherian_iff_of_finite_iSup_eq_top {ι} [Finite ι] {S : ι → X.af convert CompactSpace.isCompact_univ have : NoetherianSpace (S i) := by apply noetherianSpace_of_isAffineOpen (S i).1 (S i).2 - apply NoetherianSpace.compactSpace + apply NoetherianSpace.compactSpace (S i) /-- A version of `isNoetherian_iff_of_finite_iSup_eq_top` using `Scheme.OpenCover`. -/ theorem isNoetherian_iff_of_finite_affine_openCover {𝒰 : Scheme.OpenCover.{v, u} X} diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 38528f69f5d47..a9d7afd848943 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -76,22 +76,22 @@ theorem openEmbedding : OpenEmbedding f.1.base := /-- The image of an open immersion as an open set. -/ @[simps] -def opensRange : Opens Y := +def opensRange : Y.Opens := ⟨_, f.openEmbedding.isOpen_range⟩ /-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ -abbrev opensFunctor : Opens X ⥤ Opens Y := +abbrev opensFunctor : X.Opens ⥤ Y.Opens := LocallyRingedSpace.IsOpenImmersion.opensFunctor f /-- `f ''ᵁ U` is notation for the image (as an open set) of `U` under an open immersion `f`. -/ scoped[AlgebraicGeometry] notation3:90 f:91 " ''ᵁ " U:90 => (Scheme.Hom.opensFunctor f).obj U -lemma image_le_image_of_le {U V : Opens X} (e : U ≤ V) : f ''ᵁ U ≤ f ''ᵁ V := by +lemma image_le_image_of_le {U V : X.Opens} (e : U ≤ V) : f ''ᵁ U ≤ f ''ᵁ V := by rintro a ⟨u, hu, rfl⟩ exact Set.mem_image_of_mem (⇑f.val.base) (e hu) @[simp] -lemma opensFunctor_map_homOfLE {U V : Opens X} (e : U ≤ V) : +lemma opensFunctor_map_homOfLE {U V : X.Opens} (e : U ≤ V) : (Scheme.Hom.opensFunctor f).map (homOfLE e) = homOfLE (f.image_le_image_of_le e) := rfl @@ -101,11 +101,11 @@ lemma image_top_eq_opensRange : f ''ᵁ ⊤ = f.opensRange := by simp @[simp] -lemma preimage_image_eq (U : Opens X) : f ⁻¹ᵁ f ''ᵁ U = U := by +lemma preimage_image_eq (U : X.Opens) : f ⁻¹ᵁ f ''ᵁ U = U := by apply Opens.ext simp [Set.preimage_image_eq _ f.openEmbedding.inj] -lemma image_preimage_eq_opensRange_inter (U : Opens Y) : f ''ᵁ f ⁻¹ᵁ U = f.opensRange ⊓ U := by +lemma image_preimage_eq_opensRange_inter (U : Y.Opens) : f ''ᵁ f ⁻¹ᵁ U = f.opensRange ⊓ U := by apply Opens.ext simp [Set.image_preimage_eq_range_inter] @@ -114,7 +114,7 @@ def appIso (U) : Γ(Y, f ''ᵁ U) ≅ Γ(X, U) := (asIso <| LocallyRingedSpace.IsOpenImmersion.invApp f U).symm @[reassoc (attr := simp)] -theorem appIso_inv_naturality {U V : Opens X} (i : op U ⟶ op V) : +theorem appIso_inv_naturality {U V : X.Opens} (i : op U ⟶ op V) : X.presheaf.map i ≫ (f.appIso V).inv = (f.appIso U).inv ≫ Y.presheaf.map (f.opensFunctor.op.map i) := PresheafedSpace.IsOpenImmersion.inv_naturality _ _ @@ -147,8 +147,8 @@ theorem appIso_inv_app (U) : (PresheafedSpace.IsOpenImmersion.invApp_app _ _).trans (by rw [eqToHom_op]) @[reassoc (attr := simp), elementwise] -lemma appLE_appIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U : Opens Y} - {V : Opens X} (e : V ≤ f ⁻¹ᵁ U) : +lemma appLE_appIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U : Y.Opens} + {V : X.Opens} (e : V ≤ f ⁻¹ᵁ U) : f.appLE U V e ≫ (f.appIso V).inv = Y.presheaf.map (homOfLE <| (f.image_le_image_of_le e).trans (f.image_preimage_eq_opensRange_inter U ▸ inf_le_right)).op := by @@ -158,7 +158,7 @@ lemma appLE_appIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U : rfl @[reassoc (attr := simp)] -lemma appIso_inv_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U V : Opens X} +lemma appIso_inv_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U V : X.Opens} (e : V ≤ f ⁻¹ᵁ f ''ᵁ U) : (f.appIso U).inv ≫ f.appLE (f ''ᵁ U) V e = X.presheaf.map (homOfLE (by rwa [preimage_image_eq] at e)).op := by @@ -171,7 +171,7 @@ end Scheme.Hom /-- The open sets of an open subscheme corresponds to the open sets containing in the image. -/ @[simps] def IsOpenImmersion.opensEquiv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] : - Opens X ≃ { U : Opens Y // U ≤ f.opensRange } where + X.Opens ≃ { U : Y.Opens // U ≤ f.opensRange } where toFun U := ⟨f ''ᵁ U, Set.image_subset_range _ _⟩ invFun U := f ⁻¹ᵁ U left_inv _ := Opens.ext (Set.preimage_image_eq _ f.openEmbedding.inj) @@ -205,7 +205,7 @@ lemma IsOpenImmersion.of_isLocalization {R S} [CommRing R] [CommRing S] infer_instance theorem exists_affine_mem_range_and_range_subset - {X : Scheme.{u}} {x : X} {U : Opens X} (hxU : x ∈ U) : + {X : Scheme.{u}} {x : X} {U : X.Opens} (hxU : x ∈ U) : ∃ (R : CommRingCat) (f : Spec R ⟶ X), IsOpenImmersion f ∧ x ∈ Set.range f.1.base ∧ Set.range f.1.base ⊆ U := by obtain ⟨⟨V, hxV⟩, R, ⟨e⟩⟩ := X.2 x @@ -530,7 +530,7 @@ lemma isoOfRangeEq_inv_fac {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) lift_fac _ _ (le_of_eq e.symm) theorem app_eq_invApp_app_of_comp_eq_aux {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : U ⟶ X) (fg : Y ⟶ X) - (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : Opens U) : + (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : U.Opens) : f ⁻¹ᵁ V = fg ⁻¹ᵁ (g ''ᵁ V) := by subst H rw [Scheme.comp_val_base, Opens.map_comp_obj] @@ -540,7 +540,7 @@ theorem app_eq_invApp_app_of_comp_eq_aux {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : /-- The `fg` argument is to avoid nasty stuff about dependent types. -/ theorem app_eq_appIso_inv_app_of_comp_eq {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : U ⟶ X) (fg : Y ⟶ X) - (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : Opens U) : + (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : U.Opens) : f.app V = (g.appIso V).inv ≫ fg.app (g ''ᵁ V) ≫ Y.presheaf.map (eqToHom <| IsOpenImmersion.app_eq_invApp_app_of_comp_eq_aux f g fg H V).op := by subst H @@ -549,7 +549,7 @@ theorem app_eq_appIso_inv_app_of_comp_eq {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : eqToHom_op, eqToHom_refl, CategoryTheory.Functor.map_id, Category.comp_id] theorem lift_app {X Y U : Scheme.{u}} (f : U ⟶ Y) (g : X ⟶ Y) [IsOpenImmersion f] (H) - (V : Opens U) : + (V : U.Opens) : (IsOpenImmersion.lift f g H).app V = (f.appIso V).inv ≫ g.app (f ''ᵁ V) ≫ X.presheaf.map (eqToHom <| IsOpenImmersion.app_eq_invApp_app_of_comp_eq_aux _ _ _ (IsOpenImmersion.lift_fac f g H).symm V).op := @@ -558,25 +558,25 @@ theorem lift_app {X Y U : Scheme.{u}} (f : U ⟶ Y) (g : X ⟶ Y) [IsOpenImmersi /-- If `f` is an open immersion `X ⟶ Y`, the global sections of `X` are naturally isomorphic to the sections of `Y` over the image of `f`. -/ noncomputable -def ΓIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Opens Y) : +def ΓIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Y.Opens) : Γ(X, f⁻¹ᵁ U) ≅ Γ(Y, f.opensRange ⊓ U) := (f.appIso (f⁻¹ᵁ U)).symm ≪≫ Y.presheaf.mapIso (eqToIso <| (f.image_preimage_eq_opensRange_inter U).symm).op @[simp] -lemma ΓIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Opens Y) : +lemma ΓIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Y.Opens) : (ΓIso f U).inv = f.appLE (f.opensRange ⊓ U) (f⁻¹ᵁ U) (by rw [← f.image_preimage_eq_opensRange_inter, f.preimage_image_eq]) := by simp only [ΓIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, eqToIso.inv, eqToHom_op, asIso_inv, IsIso.comp_inv_eq, Iso.symm_inv, Scheme.Hom.appIso_hom', Scheme.Hom.map_appLE] @[reassoc, elementwise] -lemma map_ΓIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Opens Y) : +lemma map_ΓIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Y.Opens) : Y.presheaf.map (homOfLE inf_le_right).op ≫ (ΓIso f U).inv = f.app U := by simp [Scheme.Hom.appLE_eq_app] @[reassoc, elementwise] -lemma ΓIso_hom_map {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Opens Y) : +lemma ΓIso_hom_map {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : Y.Opens) : f.app U ≫ (ΓIso f U).hom = Y.presheaf.map (homOfLE inf_le_right).op := by rw [← map_ΓIso_inv] simp [-ΓIso_inv] @@ -592,7 +592,7 @@ end IsOpenImmersion namespace Scheme -theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] {U : Opens X} +theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] {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) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 338ab3eb4dd8d..dd26335893a81 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -250,9 +250,8 @@ theorem preimage_comap_zeroLocus (s : Set R) : comap f ⁻¹' zeroLocus s = zero theorem comap_injective_of_surjective (f : R →+* S) (hf : Function.Surjective f) : Function.Injective (comap f) := fun x y h => - PrimeSpectrum.ext - (Ideal.comap_injective_of_surjective f hf - (congr_arg PrimeSpectrum.asIdeal h : (comap f x).asIdeal = (comap f y).asIdeal)) + PrimeSpectrum.ext (Ideal.comap_injective_of_surjective f hf + (congr_arg PrimeSpectrum.asIdeal h : (comap f x).asIdeal = (comap f y).asIdeal)) variable (S) @@ -513,6 +512,20 @@ lemma comap_basicOpen (f : R →+* S) (x : R) : TopologicalSpace.Opens.comap (comap f) (basicOpen x) = basicOpen (f x) := rfl +open TopologicalSpace in +lemma iSup_basicOpen_eq_top_iff {ι : Type*} {f : ι → R} : + (⨆ i : ι, PrimeSpectrum.basicOpen (f i)) = ⊤ ↔ Ideal.span (Set.range f) = ⊤ := by + rw [SetLike.ext'_iff, Opens.coe_iSup] + simp only [PrimeSpectrum.basicOpen_eq_zeroLocus_compl, Opens.coe_top, ← Set.compl_iInter, + ← PrimeSpectrum.zeroLocus_iUnion] + rw [← PrimeSpectrum.zeroLocus_empty_iff_eq_top, compl_involutive.eq_iff] + simp only [Set.iUnion_singleton_eq_range, Set.compl_univ, PrimeSpectrum.zeroLocus_span] + +lemma iSup_basicOpen_eq_top_iff' {s : Set R} : + (⨆ i ∈ s, PrimeSpectrum.basicOpen i) = ⊤ ↔ Ideal.span s = ⊤ := by + conv_rhs => rw [← Subtype.range_val (s := s), ← iSup_basicOpen_eq_top_iff] + simp + end BasicOpen section Order @@ -521,20 +534,9 @@ section Order ## The specialization order We endow `PrimeSpectrum R` with a partial order, where `x ≤ y` if and only if `y ∈ closure {x}`. +This instance was defined in `RingTheory/PrimeSpectrum/Basic`. -/ - -instance : PartialOrder (PrimeSpectrum R) := - PartialOrder.lift asIdeal (fun _ _ => PrimeSpectrum.ext) - -@[simp] -theorem asIdeal_le_asIdeal (x y : PrimeSpectrum R) : x.asIdeal ≤ y.asIdeal ↔ x ≤ y := - Iff.rfl - -@[simp] -theorem asIdeal_lt_asIdeal (x y : PrimeSpectrum R) : x.asIdeal < y.asIdeal ↔ x < y := - Iff.rfl - theorem le_iff_mem_closure (x y : PrimeSpectrum R) : x ≤ y ↔ y ∈ closure ({x} : Set (PrimeSpectrum R)) := by rw [← asIdeal_le_asIdeal, ← zeroLocus_vanishingIdeal_eq_closure, mem_zeroLocus, @@ -551,14 +553,6 @@ def nhdsOrderEmbedding : PrimeSpectrum R ↪o Filter (PrimeSpectrum R) := instance : T0Space (PrimeSpectrum R) := ⟨nhdsOrderEmbedding.inj'⟩ -instance [IsDomain R] : OrderBot (PrimeSpectrum R) where - bot := ⟨⊥, Ideal.bot_prime⟩ - bot_le I := @bot_le _ _ _ I.asIdeal - -instance {R : Type*} [Field R] : Unique (PrimeSpectrum R) where - default := ⊥ - uniq x := PrimeSpectrum.ext ((IsSimpleOrder.eq_bot_or_eq_top _).resolve_right x.2.ne_top) - end Order /-- If `x` specializes to `y`, then there is a natural map from the localization of `y` to the diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean new file mode 100644 index 0000000000000..05ca19ffa9555 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.SurjectiveOnStalks +import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic + +/-! + +# Lemmas regarding the prime spectrum of tensor products + +## Main result +- `PrimeSpectrum.embedding_tensorProductTo_of_surjectiveOnStalks`: + If `R →+* T` is surjective on stalks (see Mathlib/RingTheory/SurjectiveOnStalks.lean), + then `Spec(S ⊗[R] T) → Spec S × Spec T` is a topological embedding + (where `Spec S × Spec T` is the cartesian product with the product topology). +-/ + +variable (R S T : Type*) [CommRing R] [CommRing S] [Algebra R S] +variable [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] + +open TensorProduct + +/-- The canonical map from `Spec(S ⊗[R] T)` to the cartesian product `Spec S × Spec T`. -/ +noncomputable +def PrimeSpectrum.tensorProductTo (x : PrimeSpectrum (S ⊗[R] T)) : + PrimeSpectrum S × PrimeSpectrum T := + ⟨comap (algebraMap _ _) x, comap Algebra.TensorProduct.includeRight.toRingHom x⟩ + +lemma PrimeSpectrum.continuous_tensorProductTo : Continuous (tensorProductTo R S T) := + (comap _).2.prod_mk (comap _).2 + +variable (hRT : (algebraMap R T).SurjectiveOnStalks) + +lemma PrimeSpectrum.embedding_tensorProductTo_of_surjectiveOnStalks_aux + (p₁ p₂ : PrimeSpectrum (S ⊗[R] T)) + (h : tensorProductTo R S T p₁ = tensorProductTo R S T p₂) : + p₁ ≤ p₂ := by + let g : T →+* S ⊗[R] T := Algebra.TensorProduct.includeRight.toRingHom + intros x hxp₁ + by_contra hxp₂ + obtain ⟨t, r, a, ht, e⟩ := hRT.exists_mul_eq_tmul x + (p₂.asIdeal.comap g) inferInstance + have h₁ : a ⊗ₜ[R] t ∈ p₁.asIdeal := e ▸ p₁.asIdeal.mul_mem_left (1 ⊗ₜ[R] (r • t)) hxp₁ + have h₂ : a ⊗ₜ[R] t ∉ p₂.asIdeal := e ▸ p₂.asIdeal.primeCompl.mul_mem ht hxp₂ + rw [← mul_one a, ← one_mul t, ← Algebra.TensorProduct.tmul_mul_tmul] at h₁ h₂ + have h₃ : t ∉ p₂.asIdeal.comap g := fun h ↦ h₂ (Ideal.mul_mem_left _ _ h) + have h₄ : a ∉ p₂.asIdeal.comap (algebraMap S (S ⊗[R] T)) := + fun h ↦ h₂ (Ideal.mul_mem_right _ _ h) + replace h₃ : t ∉ p₁.asIdeal.comap g := by + rwa [show p₁.asIdeal.comap g = p₂.asIdeal.comap g from congr($h.2.1)] + replace h₄ : a ∉ p₁.asIdeal.comap (algebraMap S (S ⊗[R] T)) := by + rwa [show p₁.asIdeal.comap (algebraMap S (S ⊗[R] T)) = p₂.asIdeal.comap _ from congr($h.1.1)] + exact p₁.asIdeal.primeCompl.mul_mem h₄ h₃ h₁ + +lemma PrimeSpectrum.embedding_tensorProductTo_of_surjectiveOnStalks : + Embedding (tensorProductTo R S T) := by + refine ⟨?_, fun p₁ p₂ e ↦ + (embedding_tensorProductTo_of_surjectiveOnStalks_aux R S T hRT p₁ p₂ e).antisymm + (embedding_tensorProductTo_of_surjectiveOnStalks_aux R S T hRT p₂ p₁ e.symm)⟩ + let g : T →+* S ⊗[R] T := Algebra.TensorProduct.includeRight.toRingHom + refine ⟨(continuous_tensorProductTo ..).le_induced.antisymm (isBasis_basic_opens.le_iff.mpr ?_)⟩ + rintro _ ⟨f, rfl⟩ + rw [@isOpen_iff_forall_mem_open] + rintro J (hJ : f ∉ J.asIdeal) + obtain ⟨t, r, a, ht, e⟩ := hRT.exists_mul_eq_tmul f + (J.asIdeal.comap g) inferInstance + refine ⟨_, ?_, ⟨_, (basicOpen a).2.prod (basicOpen t).2, rfl⟩, ?_⟩ + · rintro x ⟨hx₁ : a ⊗ₜ[R] (1 : T) ∉ x.asIdeal, hx₂ : (1 : S) ⊗ₜ[R] t ∉ x.asIdeal⟩ + (hx₃ : f ∈ x.asIdeal) + apply x.asIdeal.primeCompl.mul_mem hx₁ hx₂ + rw [Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul, ← e] + exact x.asIdeal.mul_mem_left _ hx₃ + · have : a ⊗ₜ[R] (1 : T) * (1 : S) ⊗ₜ[R] t ∉ J.asIdeal := by + rw [Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul, ← e] + exact J.asIdeal.primeCompl.mul_mem ht hJ + rwa [J.isPrime.mul_mem_iff_mem_or_mem.not, not_or] at this diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 93d4135c835c7..04695224a1f20 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -665,7 +665,8 @@ lemma mk_mem_toSpec_base_apply {f} (x : Proj| pbo f) lemma toSpec_preimage_basicOpen {f} (t : NumDenSameDeg 𝒜 (.powers f)) : - toSpec 𝒜 f ⁻¹ᵁ (sbo (.mk t)) = Opens.comap ⟨_, continuous_subtype_val⟩ (pbo t.num.1) := + (Opens.map (toSpec 𝒜 f).1.base).obj (sbo (.mk t)) = + Opens.comap ⟨_, continuous_subtype_val⟩ (pbo t.num.1) := Opens.ext <| Opens.map_coe _ _ ▸ by convert (ProjIsoSpecTopComponent.ToSpec.preimage_basicOpen f t) exact funext fun _ => toSpec_base_apply_eq _ _ diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index 68bb2d2b3949d..3a6c56a18fe33 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -112,9 +112,9 @@ theorem isReduced_of_isAffine_isReduced [IsAffine X] [_root_.IsReduced Γ(X, ⊤ 3. `P` holds for the entire space of an affine scheme. -/ @[elab_as_elim] -theorem reduce_to_affine_global (P : ∀ {X : Scheme} (_ : Opens X), Prop) - {X : Scheme} (U : Opens X) - (h₁ : ∀ (X : Scheme) (U : Opens X), +theorem reduce_to_affine_global (P : ∀ {X : Scheme} (_ : X.Opens), Prop) + {X : Scheme} (U : X.Opens) + (h₁ : ∀ (X : Scheme) (U : X.Opens), (∀ x : U, ∃ (V : _) (_ : x.1 ∈ V) (_ : V ⟶ U), P V) → P U) (h₂ : ∀ (X Y) (f : X ⟶ Y) [hf : IsOpenImmersion f], ∃ (U : Set X) (V : Set Y) (hU : U = ⊤) (hV : V = Set.range f.1.base), @@ -141,7 +141,7 @@ theorem reduce_to_affine_nbhd (P : ∀ (X : Scheme) (_ : X), Prop) · rw [e] apply h₁ -theorem eq_zero_of_basicOpen_eq_bot {X : Scheme} [hX : IsReduced X] {U : Opens X} +theorem eq_zero_of_basicOpen_eq_bot {X : Scheme} [hX : IsReduced X] {U : X.Opens} (s : Γ(X, U)) (hs : X.basicOpen s = ⊥) : s = 0 := by apply TopCat.Presheaf.section_ext X.sheaf U intro x @@ -174,7 +174,7 @@ theorem eq_zero_of_basicOpen_eq_bot {X : Scheme} [hX : IsReduced X] {U : Opens X rw [hs, map_zero] @[simp] -theorem basicOpen_eq_bot_iff {X : Scheme} [IsReduced X] {U : Opens X} +theorem basicOpen_eq_bot_iff {X : Scheme} [IsReduced X] {U : X.Opens} (s : Γ(X, U)) : X.basicOpen s = ⊥ ↔ s = 0 := by refine ⟨eq_zero_of_basicOpen_eq_bot s, ?_⟩ rintro rfl @@ -184,13 +184,12 @@ theorem basicOpen_eq_bot_iff {X : Scheme} [IsReduced X] {U : Opens X} and `𝒪ₓ(U)` is an integral domain for each `U ≠ ∅`. -/ class IsIntegral : Prop where nonempty : Nonempty X := by infer_instance - component_integral : ∀ (U : Opens X) [Nonempty U], IsDomain Γ(X, U) := by infer_instance + component_integral : ∀ (U : X.Opens) [Nonempty U], IsDomain Γ(X, U) := by infer_instance attribute [instance] IsIntegral.component_integral IsIntegral.nonempty -instance [h : IsIntegral X] : IsDomain Γ(X, ⊤) := - @IsIntegral.component_integral _ _ _ (by - simp only [Set.univ_nonempty, Opens.nonempty_coeSort, Opens.coe_top]) +instance [IsIntegral X] : IsDomain Γ(X, ⊤) := + @IsIntegral.component_integral _ _ _ ⟨Nonempty.some inferInstance, trivial⟩ instance (priority := 900) isReduced_of_isIntegral [IsIntegral X] : IsReduced X := by constructor @@ -203,6 +202,10 @@ instance (priority := 900) isReduced_of_isIntegral [IsIntegral X] : IsReduced X · haveI : Nonempty U := by simpa infer_instance +instance Scheme.component_nontrivial (X : Scheme.{u}) (U : X.Opens) [Nonempty U] : + Nontrivial Γ(X, U) := + LocallyRingedSpace.component_nontrivial (hU := ‹_›) + instance irreducibleSpace_of_isIntegral [IsIntegral X] : IrreducibleSpace X := by by_contra H replace H : ¬IsPreirreducible (⊤ : Set X) := fun h => @@ -212,17 +215,16 @@ instance irreducibleSpace_of_isIntegral [IsIntegral X] : IrreducibleSpace X := b rcases H with ⟨S, T, hS, hT, h₁, h₂, h₃⟩ erw [not_forall] at h₂ h₃ simp_rw [not_forall] at h₂ h₃ - haveI : Nonempty (⟨Sᶜ, hS.1⟩ : Opens X) := ⟨⟨_, h₂.choose_spec.choose_spec⟩⟩ - haveI : Nonempty (⟨Tᶜ, hT.1⟩ : Opens X) := ⟨⟨_, h₃.choose_spec.choose_spec⟩⟩ - haveI : Nonempty (⟨Sᶜ, hS.1⟩ ⊔ ⟨Tᶜ, hT.1⟩ : Opens X) := + haveI : Nonempty (⟨Sᶜ, hS.1⟩ : X.Opens) := ⟨⟨_, h₂.choose_spec.choose_spec⟩⟩ + haveI : Nonempty (⟨Tᶜ, hT.1⟩ : X.Opens) := ⟨⟨_, h₃.choose_spec.choose_spec⟩⟩ + haveI : Nonempty (⟨Sᶜ, hS.1⟩ ⊔ ⟨Tᶜ, hT.1⟩ : X.Opens) := ⟨⟨_, Or.inl h₂.choose_spec.choose_spec⟩⟩ let e : Γ(X, _) ≅ CommRingCat.of _ := (X.sheaf.isProductOfDisjoint ⟨_, hS.1⟩ ⟨_, hT.1⟩ ?_).conePointUniqueUpToIso (CommRingCat.prodFanIsLimit _ _) - · apply (config := { allowSynthFailures := true }) false_of_nontrivial_of_product_domain - · exact e.symm.commRingCatIsoToRingEquiv.toMulEquiv.isDomain _ - · apply X.toLocallyRingedSpace.component_nontrivial - · apply X.toLocallyRingedSpace.component_nontrivial + · have : IsDomain (Γ(X, ⟨Sᶜ, hS.1⟩) × Γ(X, ⟨Tᶜ, hT.1⟩)) := + e.symm.commRingCatIsoToRingEquiv.toMulEquiv.isDomain _ + exact false_of_nontrivial_of_product_domain Γ(X, ⟨Sᶜ, hS.1⟩) Γ(X, ⟨Tᶜ, hT.1⟩) · ext x constructor · rintro ⟨hS, hT⟩ @@ -283,7 +285,7 @@ theorem isIntegral_of_isAffine_of_isDomain [IsAffine X] [Nonempty X] [IsDomain IsIntegral X := isIntegral_of_isOpenImmersion X.isoSpec.hom -theorem map_injective_of_isIntegral [IsIntegral X] {U V : Opens X} (i : U ⟶ V) +theorem map_injective_of_isIntegral [IsIntegral X] {U V : X.Opens} (i : U ⟶ V) [H : Nonempty U] : Function.Injective (X.presheaf.map i.op) := by rw [injective_iff_map_eq_zero] intro x hx diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index b812d46705370..da12840b42d36 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -3,7 +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.AlgebraicGeometry.OpenImmersion +import Mathlib.AlgebraicGeometry.Cover.Open /-! # Restriction of Schemes and Morphisms @@ -11,8 +11,8 @@ import Mathlib.AlgebraicGeometry.OpenImmersion ## Main definition - `AlgebraicGeometry.Scheme.restrict`: The restriction of a scheme along an open embedding. The map `X.restrict f ⟶ X` is `AlgebraicGeometry.Scheme.ofRestrict`. - `X ∣_ᵤ U` is the notation for restricting onto an open set, and `ιOpens U` is a shorthand - for `X.restrict U.open_embedding : X ∣_ᵤ U ⟶ X`. + `U : X.Opens` has a coercion to `Scheme` and `U.ι` is a shorthand + for `X.restrict U.open_embedding : U ⟶ X`. - `AlgebraicGeometry.morphism_restrict`: The restriction of `X ⟶ Y` to `X ∣_ᵤ f ⁻¹ᵁ U ⟶ Y ∣_ᵤ U`. -/ @@ -34,55 +34,127 @@ variable {C : Type u₁} [Category.{v} C] section -variable (X : Scheme.{u}) +variable {X : Scheme.{u}} (U : X.Opens) -/-- `X ∣_ᵤ U` is notation for `X.restrict U.openEmbedding`, the restriction of `X` to an open set - `U` of `X`. -/ -notation3:60 X:60 " ∣_ᵤ " U:61 => Scheme.restrict X (U : Opens X).openEmbedding +namespace Scheme.Opens + +/-- Open subset of a scheme as a scheme. -/ +@[coe] +def toScheme {X : Scheme.{u}} (U : X.Opens) : Scheme.{u} := + X.restrict U.openEmbedding + +instance : CoeOut X.Opens Scheme := ⟨toScheme⟩ /-- The restriction of a scheme to an open subset. -/ -abbrev Scheme.ιOpens {X : Scheme.{u}} (U : Opens X) : X ∣_ᵤ U ⟶ X := X.ofRestrict _ +@[simps! val_base_apply] +def ι : ↑U ⟶ X := X.ofRestrict _ + +instance : IsOpenImmersion U.ι := inferInstanceAs (IsOpenImmersion (X.ofRestrict _)) + +lemma toScheme_carrier : (U : Type u) = (U : Set X) := rfl + +lemma toScheme_presheaf_obj (V) : Γ(U, V) = Γ(X, U.ι ''ᵁ V) := rfl + +@[simp] +lemma toScheme_presheaf_map {V W} (i : V ⟶ W) : + U.toScheme.presheaf.map i = X.presheaf.map (U.ι.opensFunctor.map i.unop).op := rfl + +@[simp] +lemma ι_app (V) : U.ι.app V = X.presheaf.map + (homOfLE (x := U.ι ''ᵁ U.ι ⁻¹ᵁ V) (Set.image_preimage_subset _ _)).op := + rfl -lemma Scheme.ιOpens_image_top {X : Scheme.{u}} (U : Opens X) : (Scheme.ιOpens U) ''ᵁ ⊤ = U := +@[simp] +lemma ι_appLE (V W e) : + U.ι.appLE V W e = + X.presheaf.map (homOfLE (x := U.ι ''ᵁ W) (Set.image_subset_iff.mpr ‹_›)).op := by + simp only [Hom.appLE, ι_app, Functor.op_obj, Opens.carrier_eq_coe, toScheme_presheaf_map, + Quiver.Hom.unop_op, Hom.opensFunctor_map_homOfLE, Opens.coe_inclusion, ← Functor.map_comp] + rfl + +@[simp] +lemma ι_appIso (V) : U.ι.appIso V = Iso.refl _ := + X.ofRestrict_appIso _ _ + +@[simp] +lemma opensRange_ι : U.ι.opensRange = U := + Opens.ext Subtype.range_val + +@[simp] +lemma range_ι : Set.range U.ι.val.base = U := + Subtype.range_val + +lemma ι_image_top : U.ι ''ᵁ ⊤ = U := U.openEmbedding_obj_top -instance Scheme.ιOpens_appLE_isIso {X : Scheme.{u}} (U : Opens X) : - IsIso ((Scheme.ιOpens U).appLE U ⊤ (by simp)) := by - simp only [restrict_presheaf_obj, ofRestrict_appLE] - have : IsIso (homOfLE <| le_of_eq (ιOpens_image_top U)) := - homOfLE_isIso_of_eq _ <| ιOpens_image_top U +@[simp] +lemma ι_preimage_self : U.ι ⁻¹ᵁ U = ⊤ := + Opens.inclusion_map_eq_top _ + +instance ι_appLE_isIso : + IsIso (U.ι.appLE U ⊤ U.ι_preimage_self.ge) := by + simp only [ι, ofRestrict_appLE] + show IsIso (X.presheaf.map (eqToIso U.ι_image_top).hom.op) infer_instance -lemma Scheme.ofRestrict_app_self {X : Scheme.{u}} (U : Opens X) : - (ιOpens U).app U = X.presheaf.map (eqToHom (by simp)).op := rfl +lemma ι_app_self : U.ι.app U = X.presheaf.map (eqToHom (X := U.ι ''ᵁ _) (by simp)).op := rfl -lemma Scheme.eq_restrict_presheaf_map_eqToHom {X : Scheme.{u}} (U : Opens X) {V W : Opens U} - (e : Scheme.ιOpens U ''ᵁ V = ιOpens U ''ᵁ W) : - X.presheaf.map (eqToHom e).op = - (X ∣_ᵤ U).presheaf.map (eqToHom <| U.openEmbedding.functor_obj_injective e).op := rfl +lemma eq_presheaf_map_eqToHom {V W : Opens U} (e : U.ι ''ᵁ V = U.ι ''ᵁ W) : + X.presheaf.map (eqToHom e).op = + U.toScheme.presheaf.map (eqToHom <| U.openEmbedding.functor_obj_injective e).op := rfl @[simp] -lemma opensRange_ιOpens {X : Scheme.{u}} (U : Opens X) : (Scheme.ιOpens U).opensRange = U := - Opens.ext Subtype.range_val +lemma nonempty_iff : Nonempty U.toScheme ↔ (U : Set X).Nonempty := by + simp only [toScheme_carrier, SetLike.coe_sort_coe, nonempty_subtype] + rfl attribute [-simp] eqToHom_op in /-- The global sections of the restriction is isomorphic to the sections on the open set. -/ @[simps!] -def Scheme.resTop (X : Scheme.{u}) (U : Opens X) : Γ(X ∣_ᵤ U, ⊤) ≅ Γ(X, U) := - X.presheaf.mapIso (eqToIso U.openEmbedding_obj_top.symm).op +def topIso : Γ(U, ⊤) ≅ Γ(X, U) := + X.presheaf.mapIso (eqToIso U.ι_image_top.symm).op + +/-- The stalks of an open subscheme are isomorphic to the stalks of the original scheme. -/ +def stalkIso {X : Scheme.{u}} (U : X.Opens) (x : U) : + U.toScheme.presheaf.stalk x ≅ X.presheaf.stalk x.1 := + X.restrictStalkIso (Opens.openEmbedding _) _ + +@[reassoc (attr := simp)] +lemma germ_stalkIso_hom {X : Scheme.{u}} (U : X.Opens) + {V : TopologicalSpace.Opens U} (x : V) : + U.toScheme.presheaf.germ x ≫ (U.stalkIso x.1).hom = + X.presheaf.germ ⟨x.1.1, show x.1.1 ∈ U.ι ''ᵁ V from ⟨x.1, x.2, rfl⟩⟩ := + PresheafedSpace.restrictStalkIso_hom_eq_germ _ _ _ _ _ + +end Scheme.Opens + +/-- If `U` is a family of open sets that covers `X`, then `X.restrict U` forms an `X.open_cover`. -/ +@[simps! J obj map] +def Scheme.openCoverOfSuprEqTop {s : Type*} (X : Scheme.{u}) (U : s → X.Opens) + (hU : ⨆ i, U i = ⊤) : X.OpenCover where + J := s + obj i := U i + map i := (U i).ι + f x := + haveI : x ∈ ⨆ i, U i := hU.symm ▸ show x ∈ (⊤ : X.Opens) by trivial + (Opens.mem_iSup.mp this).choose + covers x := by + erw [Subtype.range_coe] + have : x ∈ ⨆ i, U i := hU.symm ▸ show x ∈ (⊤ : X.Opens) by trivial + exact (Opens.mem_iSup.mp this).choose_spec /-- The open sets of an open subscheme corresponds to the open sets containing in the subset. -/ @[simps!] -def opensRestrict {X : Scheme.{u}} (U : Opens X) : - Opens (X ∣_ᵤ U) ≃ { V : Opens X // V ≤ U } := - (IsOpenImmersion.opensEquiv (Scheme.ιOpens U)).trans (Equiv.subtypeEquivProp (by simp)) +def opensRestrict : + Scheme.Opens U ≃ { V : X.Opens // V ≤ U } := + (IsOpenImmersion.opensEquiv (U.ι)).trans (Equiv.subtypeEquivProp (by simp)) -instance ΓRestrictAlgebra {X : Scheme.{u}} {Y : TopCat.{u}} {f : Y ⟶ X} (hf : OpenEmbedding f) : - Algebra (Scheme.Γ.obj (op X)) (Scheme.Γ.obj (op <| X.restrict hf)) := - (Scheme.Γ.map (X.ofRestrict hf).op).toAlgebra +instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : + Algebra (Γ(X, ⊤)) Γ(U, ⊤) := + (U.ι.app ⊤).toAlgebra -lemma Scheme.map_basicOpen' (X : Scheme.{u}) (U : Opens X) (r : Γ(X ∣_ᵤ U, ⊤)) : - Scheme.ιOpens U ''ᵁ ((X ∣_ᵤ U).basicOpen r) = X.basicOpen +lemma Scheme.map_basicOpen' (r : Γ(U, ⊤)) : + U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen (X.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen (X.ofRestrict U.openEmbedding) r).trans ?_ erw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.openEmbedding_obj_top).op] @@ -91,88 +163,75 @@ lemma Scheme.map_basicOpen' (X : Scheme.{u}) (U : Opens X) (r : Γ(X ∣_ᵤ U, congr exact PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ _ _ -lemma Scheme.map_basicOpen (X : Scheme.{u}) (U : Opens X) (r : Γ(X ∣_ᵤ U, ⊤)) : - ιOpens U ''ᵁ ((X ∣_ᵤ U).basicOpen r) = X.basicOpen r := by +lemma Scheme.Opens.ι_image_basicOpen (r : Γ(U, ⊤)) : + U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen r := by rw [Scheme.map_basicOpen', Scheme.basicOpen_res_eq] -lemma Scheme.map_basicOpen_map (X : Scheme.{u}) (U : Opens X) (r : Γ(X, U)) : - ιOpens U ''ᵁ ((X ∣_ᵤ U).basicOpen <| - X.presheaf.map (eqToHom U.openEmbedding_obj_top).op r) = X.basicOpen r := by - rw [Scheme.map_basicOpen', Scheme.basicOpen_res_eq, Scheme.basicOpen_res_eq] +lemma Scheme.map_basicOpen_map (r : Γ(X, U)) : + U.ι ''ᵁ (U.toScheme.basicOpen <| U.topIso.inv r) = X.basicOpen r := by + simp only [Scheme.Opens.toScheme_presheaf_obj] + rw [Scheme.map_basicOpen', Scheme.basicOpen_res_eq, Scheme.Opens.topIso_inv, + Scheme.basicOpen_res_eq X] -- Porting note: `simps` can't synthesize `obj_left, obj_hom, mapLeft` +variable (X) in /-- The functor taking open subsets of `X` to open subschemes of `X`. -/ -- @[simps obj_left obj_hom mapLeft] -def Scheme.restrictFunctor : Opens X ⥤ Over X where - obj U := Over.mk (ιOpens U) +def Scheme.restrictFunctor : X.Opens ⥤ Over X where + obj U := Over.mk U.ι map {U V} i := Over.homMk - (IsOpenImmersion.lift (ιOpens V) (ιOpens U) <| by - dsimp [restrict, ofRestrict, LocallyRingedSpace.ofRestrict, Opens.coe_inclusion] - rw [Subtype.range_val, Subtype.range_val] - exact i.le) + (IsOpenImmersion.lift V.ι U.ι <| by simpa using i.le) (IsOpenImmersion.lift_fac _ _ _) map_id U := by ext1 dsimp only [Over.homMk_left, Over.id_left] - rw [← cancel_mono (ιOpens U), Category.id_comp, + rw [← cancel_mono U.ι, Category.id_comp, IsOpenImmersion.lift_fac] map_comp {U V W} i j := by ext1 dsimp only [Over.homMk_left, Over.comp_left] - rw [← cancel_mono (ιOpens W), Category.assoc] + rw [← cancel_mono W.ι, Category.assoc] iterate 3 rw [IsOpenImmersion.lift_fac] -@[simp] lemma Scheme.restrictFunctor_obj_left (U : Opens X) : - (X.restrictFunctor.obj U).left = X ∣_ᵤ U := rfl +@[simp] lemma Scheme.restrictFunctor_obj_left (U : X.Opens) : + (X.restrictFunctor.obj U).left = U := rfl -@[simp] lemma Scheme.restrictFunctor_obj_hom (U : Opens X) : - (X.restrictFunctor.obj U).hom = Scheme.ιOpens U := rfl +@[simp] lemma Scheme.restrictFunctor_obj_hom (U : X.Opens) : + (X.restrictFunctor.obj U).hom = U.ι := rfl -@[simp] lemma Scheme.restrictFunctor_map_left {U V : Opens X} (i : U ⟶ V) : - (X.restrictFunctor.map i).left = IsOpenImmersion.lift (ιOpens V) (ιOpens U) (by - dsimp [ofRestrict, LocallyRingedSpace.ofRestrict, Opens.inclusion] - -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [ContinuousMap.coe_mk, ContinuousMap.coe_mk]; rw [Subtype.range_val, Subtype.range_val] - exact i.le) := rfl +@[simp] lemma Scheme.restrictFunctor_map_left {U V : X.Opens} (i : U ⟶ V) : + (X.restrictFunctor.map i).left = IsOpenImmersion.lift (V.ι) U.ι (by simpa using i.le) := rfl -- Porting note: the `by ...` used to be automatically done by unification magic @[reassoc] -theorem Scheme.restrictFunctor_map_ofRestrict {U V : Opens X} (i : U ⟶ V) : - (X.restrictFunctor.map i).1 ≫ ιOpens V = ιOpens U := - IsOpenImmersion.lift_fac _ _ (by - dsimp [restrict, ofRestrict, LocallyRingedSpace.ofRestrict] - rw [Subtype.range_val, Subtype.range_val] - exact i.le) - -theorem Scheme.restrictFunctor_map_base {U V : Opens X} (i : U ⟶ V) : - (X.restrictFunctor.map i).1.1.base = (Opens.toTopCat _).map i := by +theorem Scheme.restrictFunctor_map_ofRestrict {U V : X.Opens} (i : U ⟶ V) : + (X.restrictFunctor.map i).1 ≫ V.ι = U.ι := + IsOpenImmersion.lift_fac _ _ (by simpa using i.le) + +theorem Scheme.restrictFunctor_map_base {U V : X.Opens} (i : U ⟶ V) : + (X.restrictFunctor.map i).1.val.base = (Opens.toTopCat _).map i := by ext a; refine Subtype.ext ?_ -- Porting note: `ext` did not pick up `Subtype.ext` - exact (congr_arg (fun f : X.restrict U.openEmbedding ⟶ X => f.1.base a) + exact (congr_arg (fun f : X.restrict U.openEmbedding ⟶ X => f.val.base a) (X.restrictFunctor_map_ofRestrict i)) -theorem Scheme.restrictFunctor_map_app_aux {U V : Opens X} (i : U ⟶ V) (W : Opens V) : - ιOpens U ''ᵁ ((X.restrictFunctor.map i).1 ⁻¹ᵁ W) ≤ ιOpens V ''ᵁ W := by +theorem Scheme.restrictFunctor_map_app_aux {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : + U.ι ''ᵁ ((X.restrictFunctor.map i).1 ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by simp only [← SetLike.coe_subset_coe, IsOpenMap.functor_obj_coe, Set.image_subset_iff, Scheme.restrictFunctor_map_base, Opens.map_coe, Opens.inclusion_apply] rintro _ h exact ⟨_, h, rfl⟩ -theorem Scheme.restrictFunctor_map_app {U V : Opens X} (i : U ⟶ V) (W : Opens V) : +theorem Scheme.restrictFunctor_map_app {U V : X.Opens} (i : U ⟶ V) (W : Opens V) : (X.restrictFunctor.map i).1.app W = X.presheaf.map (homOfLE <| X.restrictFunctor_map_app_aux i W).op := by - have e₁ := - Scheme.congr_app (X.restrictFunctor_map_ofRestrict i) (ιOpens V ''ᵁ W) - rw [Scheme.comp_app] at e₁ - -- Porting note: `Opens.map_functor_eq` need more help - have e₂ := (X.restrictFunctor.map i).1.naturality (eqToHom <| W.map_functor_eq (U := V)).op - rw [← IsIso.eq_inv_comp] at e₂ - dsimp [restrict] at e₁ e₂ ⊢ - rw [e₂, W.adjunction_counit_map_functor (U := V), ← IsIso.eq_inv_comp, IsIso.inv_comp_eq, - ← IsIso.eq_comp_inv] at e₁ - simp_rw [eqToHom_map (Opens.map _), eqToHom_map (IsOpenMap.functor _), ← Functor.map_inv, - ← Functor.map_comp] at e₁ - rw [e₁] + have e₁ := Scheme.congr_app (X.restrictFunctor_map_ofRestrict i) (V.ι ''ᵁ W) + have : V.ι ⁻¹ᵁ V.ι ''ᵁ W = W := W.map_functor_eq (U := V) + have e₂ := (X.restrictFunctor.map i).1.naturality (eqToIso this).hom.op + have e₃ := e₂.symm.trans e₁ + dsimp at e₃ ⊢ + rw [← IsIso.eq_comp_inv, ← Functor.map_inv, ← Functor.map_comp] at e₃ + rw [e₃, ← Functor.map_comp] congr 1 /-- The functor that restricts to open subschemes and then takes global section is @@ -189,79 +248,65 @@ def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ S /-- `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V ∣_ U` -/ noncomputable -def Scheme.restrictRestrictComm (X : Scheme.{u}) (U V : Opens X) : - X ∣_ᵤ U ∣_ᵤ ιOpens U ⁻¹ᵁ V ≅ X ∣_ᵤ V ∣_ᵤ ιOpens V ⁻¹ᵁ U := by - refine IsOpenImmersion.isoOfRangeEq (ιOpens _ ≫ ιOpens U) (ιOpens _ ≫ ιOpens V) ?_ - simp only [Scheme.restrict_carrier, Scheme.ofRestrict_val_base, Scheme.comp_coeBase, - TopCat.coe_comp, Opens.coe_inclusion, Set.range_comp, Opens.map] - rw [Subtype.range_val, Subtype.range_val] - dsimp - rw [Set.image_preimage_eq_inter_range, Set.image_preimage_eq_inter_range, - Subtype.range_val, Subtype.range_val, Set.inter_comm] +def Scheme.restrictRestrictComm (X : Scheme.{u}) (U V : X.Opens) : + (U.ι ⁻¹ᵁ V).toScheme ≅ V.ι ⁻¹ᵁ U := + IsOpenImmersion.isoOfRangeEq (Opens.ι _ ≫ U.ι) (Opens.ι _ ≫ V.ι) <| by + simp [Set.image_preimage_eq_inter_range, Set.inter_comm (U : Set X), Set.range_comp] /-- If `V` is an open subset of `U`, then `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V`. -/ noncomputable -def Scheme.restrictRestrict (X : Scheme.{u}) (U : Opens X) (V : Opens (X ∣_ᵤ U)) : - X ∣_ᵤ U ∣_ᵤ V ≅ X ∣_ᵤ ιOpens U ''ᵁ V := by - refine IsOpenImmersion.isoOfRangeEq (ιOpens _ ≫ ιOpens U) (ιOpens _) ?_ - simp only [Scheme.restrict_carrier, Scheme.ofRestrict_val_base, Scheme.comp_coeBase, - TopCat.coe_comp, Opens.coe_inclusion, Set.range_comp, Opens.map] - rw [Subtype.range_val, Subtype.range_val] - rfl +def Scheme.restrictRestrict (X : Scheme.{u}) (U : X.Opens) (V : U.toScheme.Opens) : + V.toScheme ≅ U.ι ''ᵁ V := + IsOpenImmersion.isoOfRangeEq (Opens.ι _ ≫ U.ι) (Opens.ι _) (by simp [Set.range_comp]) @[simp, reassoc] -lemma Scheme.restrictRestrict_hom_restrict (X : Scheme.{u}) (U : Opens X) - (V : Opens (X ∣_ᵤ U)) : - (X.restrictRestrict U V).hom ≫ ιOpens _ = ιOpens V ≫ ιOpens U := +lemma Scheme.restrictRestrict_hom_restrict (X : Scheme.{u}) (U : X.Opens) + (V : U.toScheme.Opens) : + (X.restrictRestrict U V).hom ≫ Opens.ι _ = V.ι ≫ U.ι := IsOpenImmersion.isoOfRangeEq_hom_fac _ _ _ @[simp, reassoc] -lemma Scheme.restrictRestrict_inv_restrict_restrict (X : Scheme.{u}) (U : Opens X) - (V : Opens (X ∣_ᵤ U)) : - (X.restrictRestrict U V).inv ≫ ιOpens V ≫ ιOpens U = ιOpens _ := +lemma Scheme.restrictRestrict_inv_restrict_restrict (X : Scheme.{u}) (U : X.Opens) + (V : U.toScheme.Opens) : + (X.restrictRestrict U V).inv ≫ V.ι ≫ U.ι = Opens.ι _ := IsOpenImmersion.isoOfRangeEq_inv_fac _ _ _ /-- If `U = V`, then `X ∣_ U` is isomorphic to `X ∣_ V`. -/ noncomputable -def Scheme.restrictIsoOfEq (X : Scheme.{u}) {U V : Opens X} (e : U = V) : - X ∣_ᵤ U ≅ X ∣_ᵤ V := - IsOpenImmersion.isoOfRangeEq (ιOpens U) (ιOpens V) (by rw [e]) +def Scheme.restrictIsoOfEq (X : Scheme.{u}) {U V : X.Opens} (e : U = V) : + (U : Scheme.{u}) ≅ V := + IsOpenImmersion.isoOfRangeEq U.ι (V.ι) (by rw [e]) end /-- The restriction of an isomorphism onto an open set. -/ noncomputable abbrev Scheme.restrictMapIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] - (U : Opens Y) : X ∣_ᵤ f ⁻¹ᵁ U ≅ Y ∣_ᵤ U := by - apply IsOpenImmersion.isoOfRangeEq (f := X.ofRestrict _ ≫ f) (Y.ofRestrict _) _ - dsimp [restrict] - rw [Set.range_comp, Subtype.range_val, Subtype.range_coe] - refine @Set.image_preimage_eq _ _ f.1.base U.1 ?_ - rw [← TopCat.epi_iff_surjective] - infer_instance + (U : Y.Opens) : (f ⁻¹ᵁ U).toScheme ≅ U := by + apply IsOpenImmersion.isoOfRangeEq (f := (f ⁻¹ᵁ U).ι ≫ f) U.ι _ + dsimp + rw [Set.range_comp, Opens.range_ι, Opens.range_ι] + refine @Set.image_preimage_eq _ _ f.val.base U.1 f.homeomorph.surjective section MorphismRestrict /-- Given a morphism `f : X ⟶ Y` and an open set `U ⊆ Y`, we have `X ×[Y] U ≅ X |_{f ⁻¹ U}` -/ -def pullbackRestrictIsoRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : - pullback f (Scheme.ιOpens U) ≅ X ∣_ᵤ f ⁻¹ᵁ U := by - refine IsOpenImmersion.isoOfRangeEq (pullback.fst f _) (X.ofRestrict _) ?_ - rw [IsOpenImmersion.range_pullback_fst_of_right] - dsimp [Opens.coe_inclusion, Scheme.restrict] - rw [Subtype.range_val, Subtype.range_coe] - rfl +def pullbackRestrictIsoRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + pullback f (U.ι) ≅ f ⁻¹ᵁ U := by + refine IsOpenImmersion.isoOfRangeEq (pullback.fst f _) (Scheme.Opens.ι _) ?_ + simp [IsOpenImmersion.range_pullback_fst_of_right] @[simp, reassoc] -theorem pullbackRestrictIsoRestrict_inv_fst {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : - (pullbackRestrictIsoRestrict f U).inv ≫ pullback.fst f _ = X.ofRestrict _ := by +theorem pullbackRestrictIsoRestrict_inv_fst {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + (pullbackRestrictIsoRestrict f U).inv ≫ pullback.fst f _ = (f ⁻¹ᵁ U).ι := by delta pullbackRestrictIsoRestrict; simp @[simp, reassoc] -theorem pullbackRestrictIsoRestrict_hom_restrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : - (pullbackRestrictIsoRestrict f U).hom ≫ Scheme.ιOpens (f ⁻¹ᵁ U) = pullback.fst f _ := by +theorem pullbackRestrictIsoRestrict_hom_restrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + (pullbackRestrictIsoRestrict f U).hom ≫ (f ⁻¹ᵁ U).ι = pullback.fst f _ := by delta pullbackRestrictIsoRestrict; simp /-- The restriction of a morphism `X ⟶ Y` onto `X |_{f ⁻¹ U} ⟶ Y |_ U`. -/ -def morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : X ∣_ᵤ f ⁻¹ᵁ U ⟶ Y ∣_ᵤ U := +def morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : (f ⁻¹ᵁ U).toScheme ⟶ U := (pullbackRestrictIsoRestrict f U).inv ≫ pullback.snd _ _ /-- the notation for restricting a morphism of scheme to an open subset of the target scheme -/ @@ -269,17 +314,17 @@ infixl:85 " ∣_ " => morphismRestrict @[simp, reassoc] theorem pullbackRestrictIsoRestrict_hom_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) - (U : Opens Y) : (pullbackRestrictIsoRestrict f U).hom ≫ f ∣_ U = pullback.snd _ _ := + (U : Y.Opens) : (pullbackRestrictIsoRestrict f U).hom ≫ f ∣_ U = pullback.snd _ _ := Iso.hom_inv_id_assoc _ _ @[simp, reassoc] -theorem morphismRestrict_ι {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : - (f ∣_ U) ≫ Scheme.ιOpens U = Scheme.ιOpens (f ⁻¹ᵁ U) ≫ f := by +theorem morphismRestrict_ι {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + (f ∣_ U) ≫ U.ι = (f ⁻¹ᵁ U).ι ≫ f := by delta morphismRestrict rw [Category.assoc, pullback.condition.symm, pullbackRestrictIsoRestrict_inv_fst_assoc] -theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : - IsPullback (f ∣_ U) (Scheme.ιOpens (f ⁻¹ᵁ U)) (Scheme.ιOpens U) f := by +theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + IsPullback (f ∣_ U) (f ⁻¹ᵁ U).ι U.ι f := by delta morphismRestrict rw [← Category.id_comp f] refine @@ -288,6 +333,11 @@ theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens -- Porting note: changed `rw` to `erw` erw [pullbackRestrictIsoRestrict_inv_fst]; rw [Category.comp_id] +@[simp] +lemma morphismRestrict_id {X : Scheme.{u}} (U : X.Opens) : 𝟙 X ∣_ U = 𝟙 _ := by + rw [← cancel_mono U.ι, morphismRestrict_ι, Category.comp_id, Category.id_comp] + rfl + theorem morphismRestrict_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U : Opens Z) : (f ≫ g) ∣_ U = f ∣_ g ⁻¹ᵁ U ≫ g ∣_ U := by delta morphismRestrict @@ -299,25 +349,26 @@ theorem morphismRestrict_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) ( rw [pullbackRestrictIsoRestrict_inv_fst, pullbackRightPullbackFstIso_inv_snd_fst, ← pullback.condition, pullbackRestrictIsoRestrict_inv_fst_assoc, pullbackRestrictIsoRestrict_inv_fst_assoc] + rfl -instance {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] (U : Opens Y) : IsIso (f ∣_ U) := by +instance {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] (U : Y.Opens) : IsIso (f ∣_ U) := by delta morphismRestrict; infer_instance -theorem morphismRestrict_base_coe {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (x) : - @Coe.coe U Y (⟨fun x => x.1⟩) ((f ∣_ U).1.base x) = f.1.base x.1 := +theorem morphismRestrict_base_coe {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (x) : + @Coe.coe U Y (⟨fun x => x.1⟩) ((f ∣_ U).val.base x) = f.val.base x.1 := congr_arg (fun f => PresheafedSpace.Hom.base (LocallyRingedSpace.Hom.val f) x) (morphismRestrict_ι f U) -theorem morphismRestrict_val_base {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : - ⇑(f ∣_ U).1.base = U.1.restrictPreimage f.1.base := +theorem morphismRestrict_val_base {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + ⇑(f ∣_ U).val.base = U.1.restrictPreimage f.val.base := funext fun x => Subtype.ext (morphismRestrict_base_coe f U x) -theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : - Scheme.ιOpens (f ⁻¹ᵁ U) ''ᵁ ((f ∣_ U) ⁻¹ᵁ V) = f ⁻¹ᵁ (Scheme.ιOpens U ''ᵁ V) := by +theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Opens U) : + (f ⁻¹ᵁ U).ι ''ᵁ ((f ∣_ U) ⁻¹ᵁ V) = f ⁻¹ᵁ (U.ι ''ᵁ V) := by ext1 ext x constructor - · rintro ⟨⟨x, hx⟩, hx' : (f ∣_ U).1.base _ ∈ V, rfl⟩ + · rintro ⟨⟨x, hx⟩, hx' : (f ∣_ U).val.base _ ∈ V, rfl⟩ refine ⟨⟨_, hx⟩, ?_, rfl⟩ -- Porting note: this rewrite was not necessary rw [SetLike.mem_coe] @@ -326,39 +377,43 @@ theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Op refine Subtype.ext ?_ exact (morphismRestrict_base_coe f U ⟨x, hx⟩).symm · rintro ⟨⟨x, hx⟩, hx' : _ ∈ V.1, rfl : x = _⟩ - refine ⟨⟨_, hx⟩, (?_ : (f ∣_ U).1.base ⟨x, hx⟩ ∈ V.1), rfl⟩ + refine ⟨⟨_, hx⟩, (?_ : (f ∣_ U).val.base ⟨x, hx⟩ ∈ V.1), rfl⟩ convert hx' -- Porting note: `ext1` is compiling refine Subtype.ext ?_ exact morphismRestrict_base_coe f U ⟨x, hx⟩ +lemma eqToHom_eq_homOfLE {C} [Preorder C] {X Y : C} (e : X = Y) : eqToHom e = homOfLE e.le := rfl + open Scheme in -theorem morphismRestrict_app {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : - (f ∣_ U).app V = f.app (ιOpens U ''ᵁ V) ≫ +theorem morphismRestrict_app {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : U.toScheme.Opens) : + (f ∣_ U).app V = f.app (U.ι ''ᵁ V) ≫ X.presheaf.map (eqToHom (image_morphismRestrict_preimage f U V)).op := by - have := Scheme.congr_app (morphismRestrict_ι f U) (Scheme.ιOpens U ''ᵁ V) - simp only [restrict_presheaf_obj, Hom.app_eq_appLE, eqToHom_op, Hom.appLE_map] - simp only [comp_coeBase, Opens.map_comp_obj, restrict_presheaf_obj, - Hom.app_eq_appLE, comp_appLE, ofRestrict_appLE, Hom.appLE_map, eqToHom_op, - restrict_presheaf_map, eqToHom_unop] at this - have e : ιOpens U ⁻¹ᵁ (ιOpens U ''ᵁ V) = V := + have := Scheme.congr_app (morphismRestrict_ι f U) (U.ι ''ᵁ V) + simp only [Scheme.preimage_comp, Opens.toScheme_presheaf_obj, Hom.app_eq_appLE, comp_appLE, + Opens.ι_appLE, eqToHom_op, Opens.toScheme_presheaf_map, eqToHom_unop] at this + have e : U.ι ⁻¹ᵁ (U.ι ''ᵁ V) = V := Opens.ext (Set.preimage_image_eq _ Subtype.coe_injective) - have e' : (f ∣_ U) ⁻¹ᵁ V = (f ∣_ U) ⁻¹ᵁ ιOpens U ⁻¹ᵁ ιOpens U ''ᵁ V := by rw [e] - rw [← (f ∣_ U).appLE_map' _ e', ← (f ∣_ U).map_appLE' _ e, Scheme.restrict_presheaf_map, - Scheme.restrict_presheaf_map, this, Hom.appLE_map] + have e' : (f ∣_ U) ⁻¹ᵁ V = (f ∣_ U) ⁻¹ᵁ U.ι ⁻¹ᵁ U.ι ''ᵁ V := by rw [e] + simp only [Opens.toScheme_presheaf_obj, Hom.app_eq_appLE, eqToHom_op, Hom.appLE_map] + rw [← (f ∣_ U).appLE_map' _ e', ← (f ∣_ U).map_appLE' _ e] + simp only [Opens.toScheme_presheaf_obj, eqToHom_eq_homOfLE, Opens.toScheme_presheaf_map, + Quiver.Hom.unop_op, Hom.opensFunctor_map_homOfLE] + rw [this, Hom.appLE_map, Hom.appLE_map, Hom.appLE_map] @[simp] -theorem morphismRestrict_app' {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : +theorem morphismRestrict_app' {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Opens U) : (f ∣_ U).app V = f.appLE _ _ (image_morphismRestrict_preimage f U V).le := morphismRestrict_app f U V @[simp] -theorem morphismRestrict_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V W e) : - (f ∣_ U).appLE V W e = f.appLE (Scheme.ιOpens U ''ᵁ V) (Scheme.ιOpens (f ⁻¹ᵁ U) ''ᵁ W) +theorem morphismRestrict_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V W e) : + (f ∣_ U).appLE V W e = f.appLE (U.ι ''ᵁ V) ((f ⁻¹ᵁ U).ι ''ᵁ W) ((Set.image_subset _ e).trans (image_morphismRestrict_preimage f U V).le) := by - rw [Scheme.Hom.appLE, morphismRestrict_app', Scheme.restrict_presheaf_map, Scheme.Hom.appLE_map] + rw [Scheme.Hom.appLE, morphismRestrict_app', Scheme.Opens.toScheme_presheaf_map, + Scheme.Hom.appLE_map] -theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : +theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : Scheme.Γ.map (f ∣_ U).op = Y.presheaf.map (eqToHom U.openEmbedding_obj_top.symm).op ≫ f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).openEmbedding_obj_top).op := by @@ -370,10 +425,10 @@ along the immersion. -/ def morphismRestrictOpensRange {X Y U : Scheme.{u}} (f : X ⟶ Y) (g : U ⟶ Y) [hg : IsOpenImmersion g] : Arrow.mk (f ∣_ g.opensRange) ≅ Arrow.mk (pullback.snd f g) := by - let V : Opens Y := g.opensRange + let V : Y.Opens := g.opensRange let e := - IsOpenImmersion.isoOfRangeEq g (Scheme.ιOpens V) Subtype.range_coe.symm - let t : pullback f g ⟶ pullback f (Scheme.ιOpens V) := + IsOpenImmersion.isoOfRangeEq g V.ι Subtype.range_coe.symm + let t : pullback f g ⟶ pullback f V.ι := pullback.map _ _ _ _ (𝟙 _) e.hom (𝟙 _) (by rw [Category.comp_id, Category.id_comp]) (by rw [Category.comp_id, IsOpenImmersion.isoOfRangeEq_hom_fac]) symm @@ -385,13 +440,13 @@ def morphismRestrictOpensRange /-- The restrictions onto two equal open sets are isomorphic. This currently has bad defeqs when unfolded, but it should not matter for now. Replace this definition if better defeqs are needed. -/ -def morphismRestrictEq {X Y : Scheme.{u}} (f : X ⟶ Y) {U V : Opens Y} (e : U = V) : +def morphismRestrictEq {X Y : Scheme.{u}} (f : X ⟶ Y) {U V : Y.Opens} (e : U = V) : Arrow.mk (f ∣_ U) ≅ Arrow.mk (f ∣_ V) := eqToIso (by subst e; rfl) /-- Restricting a morphism twice is isomorphic to one restriction. -/ -def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens (Y ∣_ᵤ U)) : - Arrow.mk (f ∣_ U ∣_ V) ≅ Arrow.mk (f ∣_ Scheme.ιOpens U ''ᵁ V) := by +def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : U.toScheme.Opens) : + Arrow.mk (f ∣_ U ∣_ V) ≅ Arrow.mk (f ∣_ U.ι ''ᵁ V) := by refine Arrow.isoMk' _ _ (Scheme.restrictRestrict _ _ _ ≪≫ Scheme.restrictIsoOfEq _ ?_) (Scheme.restrictRestrict _ _ _) ?_ · ext x @@ -402,7 +457,7 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : exact ⟨_, h₂, rfl⟩ · rintro ⟨⟨a, h₁⟩, h₂, rfl : a = _⟩ exact ⟨⟨x, h₁⟩, h₂, rfl⟩ - · rw [← cancel_mono (Scheme.ιOpens _), Iso.trans_hom, Category.assoc, Category.assoc, + · rw [← cancel_mono (Scheme.Opens.ι _), Iso.trans_hom, Category.assoc, Category.assoc, Category.assoc, morphismRestrict_ι, Scheme.restrictIsoOfEq, IsOpenImmersion.isoOfRangeEq_hom_fac_assoc, Scheme.restrictRestrict_hom_restrict_assoc, @@ -410,50 +465,65 @@ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : morphismRestrict_ι_assoc, morphismRestrict_ι] /-- Restricting a morphism twice onto a basic open set is isomorphic to one restriction. -/ -def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (r : Γ(Y, U)) : +def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (r : Γ(Y, U)) : Arrow.mk (f ∣_ U ∣_ - (Y ∣_ᵤ U).basicOpen (Y.presheaf.map (eqToHom U.openEmbedding_obj_top).op r)) ≅ + U.toScheme.basicOpen (Y.presheaf.map (eqToHom U.openEmbedding_obj_top).op r)) ≅ Arrow.mk (f ∣_ Y.basicOpen r) := by refine morphismRestrictRestrict _ _ _ ≪≫ morphismRestrictEq _ ?_ - have e := Scheme.preimage_basicOpen (Y.ofRestrict U.openEmbedding) r - rw [Scheme.ofRestrict_app, Opens.adjunction_counit_app_self, eqToHom_op] at e - rw [← (Y.restrict U.openEmbedding).basicOpen_res_eq _ (eqToHom U.inclusion_map_eq_top).op] + have e := Scheme.preimage_basicOpen U.ι r + rw [Scheme.Opens.ι_app] at e + rw [← U.toScheme.basicOpen_res_eq _ (eqToHom U.inclusion_map_eq_top).op] erw [← comp_apply] erw [← Y.presheaf.map_comp] rw [eqToHom_op, eqToHom_op, eqToHom_map, eqToHom_trans] erw [← e] ext1 - dsimp [Opens.map_coe, Scheme.restrict] - rw [Set.image_preimage_eq_inter_range, Set.inter_eq_left, Subtype.range_val] + dsimp [Opens.map_coe] + rw [Set.image_preimage_eq_inter_range, Set.inter_eq_left, Scheme.Opens.range_ι] exact Y.basicOpen_le r /-- The stalk map of a restriction of a morphism is isomorphic to the stalk map of the original map. -/ -def morphismRestrictStalkMap {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (x) : +def morphismRestrictStalkMap {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (x) : Arrow.mk (PresheafedSpace.stalkMap (f ∣_ U).1 x) ≅ Arrow.mk (PresheafedSpace.stalkMap f.1 x.1) := by fapply Arrow.isoMk' - · refine Y.restrictStalkIso U.openEmbedding ((f ∣_ U).1.1 x) ≪≫ TopCat.Presheaf.stalkCongr _ ?_ + · refine U.stalkIso ((f ∣_ U).1.1 x) ≪≫ TopCat.Presheaf.stalkCongr _ ?_ apply Inseparable.of_eq exact morphismRestrict_base_coe f U x - · exact X.restrictStalkIso (Opens.openEmbedding _) _ + · exact Scheme.Opens.stalkIso _ _ · apply TopCat.Presheaf.stalk_hom_ext intro V hxV change ↑(f ⁻¹ᵁ U) at x simp only [Scheme.restrict_presheaf_obj, unop_op, Opens.coe_inclusion, Iso.trans_hom, TopCat.Presheaf.stalkCongr_hom, Category.assoc, Scheme.restrict_toPresheafedSpace] - rw [PresheafedSpace.restrictStalkIso_hom_eq_germ_assoc, + rw [Scheme.Opens.germ_stalkIso_hom_assoc, TopCat.Presheaf.germ_stalkSpecializes'_assoc, PresheafedSpace.stalkMap_germ'_assoc, PresheafedSpace.stalkMap_germ', ← Scheme.Hom.app, ← Scheme.Hom.app, morphismRestrict_app, - PresheafedSpace.restrictStalkIso_hom_eq_germ, Category.assoc, TopCat.Presheaf.germ_res] - rfl + Scheme.Opens.germ_stalkIso_hom, Category.assoc, TopCat.Presheaf.germ_res] -instance {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) [IsOpenImmersion f] : +instance {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) [IsOpenImmersion f] : IsOpenImmersion (f ∣_ U) := by delta morphismRestrict exact PresheafedSpace.IsOpenImmersion.comp _ _ end MorphismRestrict +/-- The restriction of an open cover to an open subset. -/ +@[simps! J obj map] +noncomputable +def Scheme.OpenCover.restrict {X : Scheme.{u}} (𝒰 : X.OpenCover) (U : Opens X) : + U.toScheme.OpenCover := by + refine copy (𝒰.pullbackCover U.ι) 𝒰.J _ (𝒰.map · ∣_ U) (Equiv.refl _) + (fun i ↦ IsOpenImmersion.isoOfRangeEq (Opens.ι _) (pullback.snd _ _) ?_) ?_ + · erw [IsOpenImmersion.range_pullback_snd_of_left U.ι (𝒰.map i)] + rw [Opens.opensRange_ι] + exact Subtype.range_val + · intro i + rw [← cancel_mono U.ι] + simp only [morphismRestrict_ι, pullbackCover_J, Equiv.refl_apply, pullbackCover_obj, + pullbackCover_map, Category.assoc, pullback.condition] + erw [IsOpenImmersion.isoOfRangeEq_hom_fac_assoc] + end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index c6457fb631a8c..25382d56711b3 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -50,6 +50,12 @@ structure Scheme extends LocallyRingedSpace where namespace Scheme +instance : CoeSort Scheme Type* where + coe X := X.carrier + +/-- The type of open sets of a scheme. -/ +abbrev Opens (X : Scheme) : Type* := TopologicalSpace.Opens X + /-- A morphism between schemes is a morphism between the underlying locally ringed spaces. -/ -- @[nolint has_nonempty_instance] -- Porting note(#5171): linter not ported yet def Hom (X Y : Scheme) : Type* := @@ -63,12 +69,14 @@ instance : Category Scheme := /-- `f ⁻¹ᵁ U` is notation for `(Opens.map f.1.base).obj U`, the preimage of an open set `U` under `f`. -/ scoped[AlgebraicGeometry] notation3:90 f:91 " ⁻¹ᵁ " U:90 => - (Opens.map (f : LocallyRingedSpace.Hom _ _).val.base).obj U + @Prefunctor.obj (Scheme.Opens _) _ (Scheme.Opens _) _ + (Opens.map (f : LocallyRingedSpace.Hom _ _).val.base).toPrefunctor U /-- `Γ(X, U)` is notation for `X.presheaf.obj (op U)`. -/ scoped[AlgebraicGeometry] notation3 "Γ(" X ", " U ")" => (PresheafedSpace.presheaf (SheafedSpace.toPresheafedSpace - (LocallyRingedSpace.toSheafedSpace (Scheme.toLocallyRingedSpace X)))).obj (op U) + (LocallyRingedSpace.toSheafedSpace (Scheme.toLocallyRingedSpace X)))).obj + (op (α := Scheme.Opens _) U) instance {X : Scheme.{u}} : Subsingleton Γ(X, ⊥) := CommRingCat.subsingleton_of_isTerminal X.sheaf.isTerminalOfEmpty @@ -80,16 +88,13 @@ lemma Hom.continuous {X Y : Scheme} (f : X ⟶ Y) : Continuous f.1.base := f.1.b protected abbrev sheaf (X : Scheme) := X.toSheafedSpace.sheaf -instance : CoeSort Scheme Type* where - coe X := X.carrier - namespace Hom -variable {X Y : Scheme.{u}} (f : Hom X Y) {U U' : Opens Y} {V V' : Opens X} +variable {X Y : Scheme.{u}} (f : Hom X Y) {U U' : Y.Opens} {V V' : X.Opens} /-- Given a morphism of schemes `f : X ⟶ Y`, and open `U ⊆ Y`, this is the induced map `Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U)`. -/ -abbrev app (U : Opens Y) : Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U) := +abbrev app (U : Y.Opens) : Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U) := f.1.c.app (op U) @[reassoc] @@ -99,7 +104,7 @@ lemma naturality (i : op U' ⟶ op U) : /-- Given a morphism of schemes `f : X ⟶ Y`, and open sets `U ⊆ Y`, `V ⊆ f ⁻¹' U`, this is the induced map `Γ(Y, U) ⟶ Γ(X, V)`. -/ -def appLE (U : Opens Y) (V : Opens X) (e : V ≤ f ⁻¹ᵁ U) : Γ(Y, U) ⟶ Γ(X, V) := +def appLE (U : Y.Opens) (V : X.Opens) (e : V ≤ f ⁻¹ᵁ U) : Γ(Y, U) ⟶ Γ(X, V) := f.app U ≫ X.presheaf.map (homOfLE e).op @[reassoc (attr := simp)] @@ -125,11 +130,11 @@ lemma map_appLE' (e : V ≤ f ⁻¹ᵁ U) (i : U' = U) : Y.presheaf.map (eqToHom i).op ≫ f.appLE U' V (i ▸ e) = f.appLE U V e := map_appLE _ _ _ -lemma app_eq_appLE {U : Opens Y} : +lemma app_eq_appLE {U : Y.Opens} : f.app U = f.appLE U _ le_rfl := by simp [Hom.appLE] -lemma appLE_eq_app {U : Opens Y} : +lemma appLE_eq_app {U : Y.Opens} : f.appLE U (f ⁻¹ᵁ U) le_rfl = f.app U := (app_eq_appLE f).symm @@ -142,8 +147,25 @@ lemma appLE_congr (e : V ≤ f ⁻¹ᵁ U) (e₁ : U = U') (e₂ : V = V') noncomputable def homeomorph [IsIso f] : X ≃ₜ Y := TopCat.homeoOfIso (asIso <| f.val.base) +@[ext (iff := false)] +protected lemma ext {f g : X ⟶ Y} (h_base : f.1.base = g.1.base) + (h_app : ∀ U, f.app U ≫ X.presheaf.map + (eqToHom congr((Opens.map $h_base.symm).obj U)).op = g.app U) : f = g := + LocallyRingedSpace.Hom.ext <| SheafedSpace.ext _ _ h_base + (TopCat.Presheaf.ext fun U ↦ by simpa using h_app U) + +lemma preimage_iSup {ι} (U : ι → Opens Y) : f ⁻¹ᵁ iSup U = ⨆ i, f ⁻¹ᵁ U i := + Opens.ext (by simp) + +lemma preimage_iSup_eq_top {ι} {U : ι → Opens Y} (hU : iSup U = ⊤) : + ⨆ i, f ⁻¹ᵁ U i = ⊤ := f.preimage_iSup U ▸ hU ▸ rfl + end Hom +@[simp] +lemma preimage_comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U) : + (f ≫ g) ⁻¹ᵁ U = f ⁻¹ᵁ g ⁻¹ᵁ U := rfl + /-- The forgetful functor from `Scheme` to `LocallyRingedSpace`. -/ @[simps!] def forgetToLocallyRingedSpace : Scheme ⥤ LocallyRingedSpace := @@ -181,7 +203,7 @@ theorem id_val_base (X : Scheme) : (𝟙 X : _).1.base = 𝟙 _ := rfl @[simp] -theorem id_app {X : Scheme} (U : Opens X) : +theorem id_app {X : Scheme} (U : X.Opens) : (𝟙 X : _).app U = 𝟙 _ := rfl @[reassoc] @@ -227,7 +249,7 @@ theorem congr_app {X Y : Scheme} {f g : X ⟶ Y} (e : f = g) (U) : f.app U = g.app U ≫ X.presheaf.map (eqToHom (by subst e; rfl)).op := by subst e; dsimp; simp -theorem app_eq {X Y : Scheme} (f : X ⟶ Y) {U V : Opens Y.carrier} (e : U = V) : +theorem app_eq {X Y : Scheme} (f : X ⟶ Y) {U V : Y.Opens} (e : U = V) : f.app U = Y.presheaf.map (eqToHom e.symm).op ≫ f.app V ≫ @@ -241,7 +263,7 @@ theorem eqToHom_c_app {X Y : Scheme} (e : X = Y) (U) : -- Porting note: in `AffineScheme.lean` file, `eqToHom_op` can't be used in `(e)rw` or `simp(_rw)` -- when terms get very complicated. See `AlgebraicGeometry.IsAffineOpen.isLocalization_stalk_aux`. -lemma presheaf_map_eqToHom_op (X : Scheme) (U V : Opens X) (i : U = V) : +lemma presheaf_map_eqToHom_op (X : Scheme) (U V : X.Opens) (i : U = V) : X.presheaf.map (eqToHom i).op = eqToHom (i ▸ rfl) := by rw [eqToHom_op, eqToHom_map] @@ -262,7 +284,7 @@ instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U) : IsIso (f.app U) := NatIso.isIso_app_of_isIso _ _ @[simp] -theorem inv_app {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U : Opens X) : +theorem inv_app {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U : X.Opens) : (inv f).app U = X.presheaf.map (eqToHom (show (f ≫ inv f) ⁻¹ᵁ U = U by rw [IsIso.hom_inv_id]; rfl)).op ≫ inv (f.app ((inv f) ⁻¹ᵁ U)) := by @@ -321,7 +343,7 @@ lemma Spec.map_inv {R S : CommRingCat} (f : R ⟶ S) [IsIso f] : section -variable {R S : CommRingCat} (f : R ⟶ S) +variable {R S : CommRingCat.{u}} (f : R ⟶ S) -- The lemmas below are not tagged simp to respect the abstraction. lemma Spec_carrier (R : CommRingCat.{u}) : (Spec R).carrier = PrimeSpectrum R := rfl @@ -329,7 +351,6 @@ lemma Spec_sheaf (R : CommRingCat.{u}) : (Spec R).sheaf = Spec.structureSheaf R lemma Spec_presheaf (R : CommRingCat.{u}) : (Spec R).presheaf = (Spec.structureSheaf R).1 := rfl lemma Spec.map_base : (Spec.map f).1.base = PrimeSpectrum.comap f := rfl -set_option maxHeartbeats 800000 in lemma Spec.map_app (U) : (Spec.map f).app U = StructureSheaf.comap f U (Spec.map f ⁻¹ᵁ U) le_rfl := rfl @@ -414,17 +435,17 @@ lemma toOpen_eq (U) : section BasicOpen -variable (X : Scheme) {V U : Opens X} (f g : Γ(X, U)) +variable (X : Scheme) {V U : X.Opens} (f g : Γ(X, U)) /-- The subset of the underlying space where the given section does not vanish. -/ -def basicOpen : Opens X := +def basicOpen : X.Opens := X.toLocallyRingedSpace.toRingedSpace.basicOpen f @[simp] theorem mem_basicOpen (x : U) : ↑x ∈ X.basicOpen f ↔ IsUnit (X.presheaf.germ x f) := RingedSpace.mem_basicOpen _ _ _ -theorem mem_basicOpen_top' {U : Opens X} (f : Γ(X, U)) (x : X) : +theorem mem_basicOpen_top' {U : X.Opens} (f : Γ(X, U)) (x : X) : x ∈ X.basicOpen f ↔ ∃ (m : x ∈ U), IsUnit (X.presheaf.germ (⟨x, m⟩ : U) f) := by fconstructor · rintro ⟨y, hy1, rfl⟩ @@ -457,18 +478,18 @@ lemma basicOpen_restrict (i : V ⟶ U) (f : Γ(X, U)) : (Scheme.basicOpen_res _ _ _).trans_le inf_le_right @[simp] -theorem preimage_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Opens Y} (r : Γ(Y, U)) : +theorem preimage_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (r : Γ(Y, U)) : f ⁻¹ᵁ (Y.basicOpen r) = X.basicOpen (f.app U r) := LocallyRingedSpace.preimage_basicOpen f r -lemma basicOpen_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens X) (V : Opens Y) (e : U ≤ f ⁻¹ᵁ V) +lemma basicOpen_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : X.Opens) (V : Y.Opens) (e : U ≤ f ⁻¹ᵁ V) (s : Γ(Y, V)) : X.basicOpen (f.appLE V U e s) = U ⊓ f ⁻¹ᵁ (Y.basicOpen s) := by simp only [preimage_basicOpen, Hom.appLE, CommRingCat.coe_comp_of, RingHom.coe_comp, Function.comp_apply] rw [basicOpen_res] @[simp] -theorem basicOpen_zero (U : Opens X) : X.basicOpen (0 : Γ(X, U)) = ⊥ := +theorem basicOpen_zero (U : X.Opens) : X.basicOpen (0 : Γ(X, U)) = ⊥ := LocallyRingedSpace.basicOpen_zero _ U @[simp] @@ -481,7 +502,7 @@ lemma basicOpen_pow {n : ℕ} (h : 0 < n) : X.basicOpen (f ^ n) = X.basicOpen f theorem basicOpen_of_isUnit {f : Γ(X, U)} (hf : IsUnit f) : X.basicOpen f = U := RingedSpace.basicOpen_of_isUnit _ hf -instance algebra_section_section_basicOpen {X : Scheme} {U : Opens X} (f : Γ(X, U)) : +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).toAlgebra @@ -495,27 +516,27 @@ variable (X : Scheme.{u}) The zero locus of a set of sections `s` over an open set `U` is the closed set consisting of the complement of `U` and of all points of `U`, where all elements of `f` vanish. -/ -def zeroLocus {U : Opens X} (s : Set Γ(X, U)) : Set X := X.toRingedSpace.zeroLocus s +def zeroLocus {U : X.Opens} (s : Set Γ(X, U)) : Set X := X.toRingedSpace.zeroLocus s -lemma zeroLocus_def {U : Opens X} (s : Set Γ(X, U)) : +lemma zeroLocus_def {U : X.Opens} (s : Set Γ(X, U)) : X.zeroLocus s = ⋂ f ∈ s, (X.basicOpen f).carrierᶜ := rfl -lemma zeroLocus_isClosed {U : Opens X} (s : Set Γ(X, U)) : +lemma zeroLocus_isClosed {U : X.Opens} (s : Set Γ(X, U)) : IsClosed (X.zeroLocus s) := X.toRingedSpace.zeroLocus_isClosed s -lemma zeroLocus_singleton {U : Opens X} (f : Γ(X, U)) : +lemma zeroLocus_singleton {U : X.Opens} (f : Γ(X, U)) : X.zeroLocus {f} = (X.basicOpen f).carrierᶜ := X.toRingedSpace.zeroLocus_singleton f @[simp] -lemma zeroLocus_empty_eq_univ {U : Opens X} : +lemma zeroLocus_empty_eq_univ {U : X.Opens} : X.zeroLocus (∅ : Set Γ(X, U)) = Set.univ := X.toRingedSpace.zeroLocus_empty_eq_univ @[simp] -lemma mem_zeroLocus_iff {U : Opens X} (s : Set Γ(X, U)) (x : X) : +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 @@ -541,7 +562,7 @@ theorem basicOpen_eq_of_affine' {R : CommRingCat} (f : Γ(Spec R, ⊤)) : convert basicOpen_eq_of_affine ((Scheme.ΓSpecIso R).hom f) exact (Iso.hom_inv_id_apply (Scheme.ΓSpecIso R) f).symm -theorem Scheme.Spec_map_presheaf_map_eqToHom {X : Scheme} {U V : Opens X} (h : U = V) (W) : +theorem Scheme.Spec_map_presheaf_map_eqToHom {X : Scheme} {U V : X.Opens} (h : U = V) (W) : (Spec.map (X.presheaf.map (eqToHom h).op)).app W = eqToHom (by cases h; dsimp; simp) := by have : Scheme.Spec.map (X.presheaf.map (𝟙 (op U))).op = 𝟙 _ := by rw [X.presheaf.map_id, op_id, Scheme.Spec.map_id] diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 59c7611d72ee1..0da8e64004430 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -22,7 +22,7 @@ A morphism from `Spec(O_x)` to `X`, which is defined with the help of an affine neighborhood `U` of `x`. -/ noncomputable def IsAffineOpen.fromSpecStalk - {X : Scheme} {U : Opens X} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : + {X : Scheme} {U : X.Opens} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : Spec (X.stalk x) ⟶ X := Spec.map (X.presheaf.germ ⟨x, hxU⟩) ≫ hU.fromSpec @@ -30,7 +30,7 @@ noncomputable def IsAffineOpen.fromSpecStalk The morphism from `Spec(O_x)` to `X` given by `IsAffineOpen.fromSpec` does not depend on the affine open neighborhood of `x` we choose. -/ -theorem IsAffineOpen.fromSpecStalk_eq {X : Scheme} (x : X) {U V : TopologicalSpace.Opens X} +theorem IsAffineOpen.fromSpecStalk_eq {X : Scheme} (x : X) {U V : X.Opens} (hU : IsAffineOpen U) (hV : IsAffineOpen V) (hxU : x ∈ U) (hxV : x ∈ V) : hU.fromSpecStalk hxU = hV.fromSpecStalk hxV := by obtain ⟨U', h₁, h₂, h₃ : U' ≤ U ⊓ V⟩ := @@ -56,7 +56,7 @@ noncomputable def Scheme.fromSpecStalk (X : Scheme) (x : X) : @[simp] theorem IsAffineOpen.fromSpecStalk_eq_fromSpecStalk - {X : Scheme} {U : Opens X} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : + {X : Scheme} {U : X.Opens} (hU : IsAffineOpen U) {x : X} (hxU : x ∈ U) : hU.fromSpecStalk hxU = X.fromSpecStalk x := fromSpecStalk_eq .. end AlgebraicGeometry diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index e081745fcd87b..753771bcacc87 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -50,6 +50,13 @@ Additionally, let `f` be a function from `E` to `F`. * `AnalyticAt 𝕜 f x`: there exists a power series `p` such that holds `HasFPowerSeriesAt f p x`. * `AnalyticOn 𝕜 f s`: the function `f` is analytic at every point of `s`. +We also define versions of `HasFPowerSeriesOnBall`, `AnalyticAt`, and `AnalyticOn` restricted to a +set, similar to `ContinuousWithinAt`. See `Mathlib.Analysis.Analytic.Within` for basic properties. + +* `AnalyticWithinAt 𝕜 f s x` means a power series at `x` converges to `f` on `𝓝[s] x`, and + `f` is continuous within `s` at `x`. +* `AnalyticWithinOn 𝕜 f s t` means `∀ x ∈ t, AnalyticWithinAt 𝕜 f s x`. + We develop the basic properties of these notions, notably: * If a function admits a power series, it is continuous (see `HasFPowerSeriesOnBall.continuousOn` and `HasFPowerSeriesAt.continuousAt` and @@ -347,11 +354,28 @@ structure HasFPowerSeriesOnBall (f : E → F) (p : FormalMultilinearSeries 𝕜 hasSum : ∀ {y}, y ∈ EMetric.ball (0 : E) r → HasSum (fun n : ℕ => p n fun _ : Fin n => y) (f (x + y)) +/-- Analogue of `HasFPowerSeriesOnBall` where convergence is required only on a set `s`. -/ +structure HasFPowerSeriesWithinOnBall (f : E → F) (p : FormalMultilinearSeries 𝕜 E F) (s : Set E) + (x : E) (r : ℝ≥0∞) : Prop where + /-- `p` converges on `ball 0 r` -/ + r_le : r ≤ p.radius + /-- The radius of convergence is positive -/ + r_pos : 0 < r + /-- `p converges to f` within `s` -/ + hasSum : ∀ {y}, x + y ∈ s → y ∈ EMetric.ball (0 : E) r → + HasSum (fun n : ℕ => p n fun _ : Fin n => y) (f (x + y)) + /-- We require `ContinuousWithinAt f s x` to ensure `f x` is nice -/ + continuousWithinAt : ContinuousWithinAt f s x + /-- Given a function `f : E → F` and a formal multilinear series `p`, we say that `f` has `p` as a power series around `x` if `f (x + y) = ∑' pₙ yⁿ` for all `y` in a neighborhood of `0`. -/ def HasFPowerSeriesAt (f : E → F) (p : FormalMultilinearSeries 𝕜 E F) (x : E) := ∃ r, HasFPowerSeriesOnBall f p x r +/-- Analogue of `HasFPowerSeriesAt` where convergence is required only on a set `s`. -/ +def HasFPowerSeriesWithinAt (f : E → F) (p : FormalMultilinearSeries 𝕜 E F) (s : Set E) (x : E) := + ∃ r, HasFPowerSeriesWithinOnBall f p s x r + variable (𝕜) /-- Given a function `f : E → F`, we say that `f` is analytic at `x` if it admits a convergent power @@ -359,11 +383,20 @@ series expansion around `x`. -/ def AnalyticAt (f : E → F) (x : E) := ∃ p : FormalMultilinearSeries 𝕜 E F, HasFPowerSeriesAt f p x +/-- `f` is analytic within `s` at `x` if it has a power series at `x` that converges on `𝓝[s] x` -/ +def AnalyticWithinAt (f : E → F) (s : Set E) (x : E) : Prop := + ∃ p : FormalMultilinearSeries 𝕜 E F, HasFPowerSeriesWithinAt f p s x + /-- Given a function `f : E → F`, we say that `f` is analytic on a set `s` if it is analytic around every point of `s`. -/ def AnalyticOn (f : E → F) (s : Set E) := ∀ x, x ∈ s → AnalyticAt 𝕜 f x +/-- `f` is analytic within `s` if it is analytic within `s` at each point of `t`. Note that +this is weaker than `AnalyticOn 𝕜 f s`, as `f` is allowed to be arbitrary outside `s`. -/ +def AnalyticWithinOn (f : E → F) (s : Set E) : Prop := + ∀ x ∈ s, AnalyticWithinAt 𝕜 f s x + variable {𝕜} theorem HasFPowerSeriesOnBall.hasFPowerSeriesAt (hf : HasFPowerSeriesOnBall f p x r) : diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 743f4565a50a5..a2e1dfd5a7861 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -934,7 +934,7 @@ the first two blocks of `a` and its last three blocks, giving `a.gather b = [11, def gather (a : Composition n) (b : Composition a.length) : Composition n where blocks := (a.blocks.splitWrtComposition b).map sum blocks_pos := by - rw [forall_mem_map_iff] + rw [forall_mem_map] intro j hj suffices H : ∀ i ∈ j, 1 ≤ i by calc 0 < j.length := length_pos_of_mem_splitWrtComposition hj diff --git a/Mathlib/Analysis/Analytic/Uniqueness.lean b/Mathlib/Analysis/Analytic/Uniqueness.lean index b8d90ff2f5450..b18db8f4e34b7 100644 --- a/Mathlib/Analysis/Analytic/Uniqueness.lean +++ b/Mathlib/Analysis/Analytic/Uniqueness.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Analytic.Linear import Mathlib.Analysis.Analytic.Composition -import Mathlib.Analysis.NormedSpace.Completion +import Mathlib.Analysis.Normed.Module.Completion /-! # Uniqueness principle for analytic functions diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean new file mode 100644 index 0000000000000..630e2ae62bc1d --- /dev/null +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -0,0 +1,367 @@ +/- +Copyright (c) 2024 Geoffrey Irving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Geoffrey Irving +-/ +import Mathlib.Analysis.Analytic.Constructions +import Mathlib.Analysis.Calculus.FDeriv.Analytic + +/-! +# Properties of analyticity restricted to a set + +From `Mathlib.Analysis.Analytic.Basic`, we have the definitons + +1. `AnalyticWithinAt 𝕜 f s x` means a power series at `x` converges to `f` on `𝓝[s] x`, and + `f` is continuous within `s` at `x`. +2. `AnalyticWithinOn 𝕜 f s t` means `∀ x ∈ t, AnalyticWithinAt 𝕜 f s x`. + +This means there exists an extension of `f` which is analytic and agrees with `f` on `s ∪ {x}`, but +`f` is allowed to be arbitrary elsewhere. Requiring `ContinuousWithinAt` is essential if `x ∉ s`: +it is required for composition and smoothness to follow without extra hypotheses (we could +alternately require convergence at `x` even if `x ∉ s`). + +Here we prove basic properties of these definitions. Where convenient we assume completeness of the +ambient space, which allows us to related `AnalyticWithinAt` to analyticity of a local extension. +-/ + +noncomputable section + +open Topology Filter ENNReal + +open Set Filter + +variable {α : Type*} + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + +variable {E F G H : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] + [NormedSpace 𝕜 F] [NormedAddCommGroup G] [NormedSpace 𝕜 G] [NormedAddCommGroup H] + [NormedSpace 𝕜 H] + +/-! +### Basic properties +-/ + +@[simp] lemma hasFPowerSeriesWithinOnBall_univ {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} + {x : E} {r : ℝ≥0∞} : + HasFPowerSeriesWithinOnBall f p univ x r ↔ HasFPowerSeriesOnBall f p x r := by + constructor + · intro h + exact ⟨h.r_le, h.r_pos, fun {y} m ↦ h.hasSum (mem_univ _) m⟩ + · intro h + refine ⟨h.r_le, h.r_pos, fun {y} _ m => h.hasSum m, ?_⟩ + exact (h.continuousOn.continuousAt (EMetric.ball_mem_nhds x h.r_pos)).continuousWithinAt + +@[simp] lemma hasFPowerSeriesWithinAt_univ {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} {x : E} : + HasFPowerSeriesWithinAt f p univ x ↔ HasFPowerSeriesAt f p x := by + simp only [HasFPowerSeriesWithinAt, hasFPowerSeriesWithinOnBall_univ, HasFPowerSeriesAt] + +@[simp] lemma analyticWithinAt_univ {f : E → F} {x : E} : + AnalyticWithinAt 𝕜 f univ x ↔ AnalyticAt 𝕜 f x := by + simp only [AnalyticWithinAt, hasFPowerSeriesWithinAt_univ, AnalyticAt] + +lemma analyticWithinOn_univ {f : E → F} : + AnalyticWithinOn 𝕜 f univ ↔ AnalyticOn 𝕜 f univ := by + simp only [AnalyticWithinOn, analyticWithinAt_univ, AnalyticOn] + +lemma HasFPowerSeriesWithinAt.continuousWithinAt {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} + {s : Set E} {x : E} (h : HasFPowerSeriesWithinAt f p s x) : ContinuousWithinAt f s x := by + rcases h with ⟨r, h⟩ + exact h.continuousWithinAt + +lemma AnalyticWithinAt.continuousWithinAt {f : E → F} {s : Set E} {x : E} + (h : AnalyticWithinAt 𝕜 f s x) : ContinuousWithinAt f s x := by + rcases h with ⟨p, h⟩ + exact h.continuousWithinAt + +/-- `AnalyticWithinAt` is trivial if `{x} ∈ 𝓝[s] x` -/ +lemma analyticWithinAt_of_singleton_mem {f : E → F} {s : Set E} {x : E} (h : {x} ∈ 𝓝[s] x) : + AnalyticWithinAt 𝕜 f s x := by + have fc : ContinuousWithinAt f s x := + Filter.Tendsto.mono_left (tendsto_pure_nhds _ _) (Filter.le_pure_iff.mpr h) + rcases mem_nhdsWithin.mp h with ⟨t, ot, xt, st⟩ + rcases Metric.mem_nhds_iff.mp (ot.mem_nhds xt) with ⟨r, r0, rt⟩ + exact ⟨constFormalMultilinearSeries 𝕜 E (f x), .ofReal r, { + r_le := by simp only [FormalMultilinearSeries.constFormalMultilinearSeries_radius, le_top] + r_pos := by positivity + hasSum := by + intro y ys yr + simp only [subset_singleton_iff, mem_inter_iff, and_imp] at st + specialize st (x + y) (rt (by simpa using yr)) ys + simp only [st] + apply (hasFPowerSeriesOnBall_const (e := 0)).hasSum + simp only [Metric.emetric_ball_top, mem_univ] + continuousWithinAt := fc + }⟩ + +/-- Analyticity implies analyticity within any `s` -/ +lemma AnalyticAt.analyticWithinAt {f : E → F} {s : Set E} {x : E} (h : AnalyticAt 𝕜 f x) : + AnalyticWithinAt 𝕜 f s x := by + rcases h with ⟨p, r, hp⟩ + exact ⟨p, r, { + r_le := hp.r_le + r_pos := hp.r_pos + hasSum := fun {y} _ yr ↦ hp.hasSum yr + continuousWithinAt := + (hp.continuousOn.continuousAt (EMetric.ball_mem_nhds x hp.r_pos)).continuousWithinAt + }⟩ + +/-- Analyticity on `s` implies analyticity within `s` -/ +lemma AnalyticOn.analyticWithinOn {f : E → F} {s : Set E} (h : AnalyticOn 𝕜 f s) : + AnalyticWithinOn 𝕜 f s := + fun x m ↦ (h x m).analyticWithinAt + +lemma AnalyticWithinOn.continuousOn {f : E → F} {s : Set E} (h : AnalyticWithinOn 𝕜 f s) : + ContinuousOn f s := + fun x m ↦ (h x m).continuousWithinAt + +/-- If `f` is `AnalyticWithinOn` near each point in a set, it is `AnalyticWithinOn` the set -/ +lemma analyticWithinOn_of_locally_analyticWithinOn {f : E → F} {s : Set E} + (h : ∀ x ∈ s, ∃ u, IsOpen u ∧ x ∈ u ∧ AnalyticWithinOn 𝕜 f (s ∩ u)) : + AnalyticWithinOn 𝕜 f s := by + intro x m + rcases h x m with ⟨u, ou, xu, fu⟩ + rcases Metric.mem_nhds_iff.mp (ou.mem_nhds xu) with ⟨r, r0, ru⟩ + rcases fu x ⟨m, xu⟩ with ⟨p, t, fp⟩ + exact ⟨p, min (.ofReal r) t, { + r_pos := lt_min (by positivity) fp.r_pos + r_le := min_le_of_right_le fp.r_le + hasSum := by + intro y ys yr + simp only [EMetric.mem_ball, lt_min_iff, edist_lt_ofReal, dist_zero_right] at yr + apply fp.hasSum ⟨ys, ru ?_⟩ + · simp only [EMetric.mem_ball, yr] + · simp only [Metric.mem_ball, dist_self_add_left, yr] + continuousWithinAt := by + refine (fu.continuousOn x ⟨m, xu⟩).mono_left (le_of_eq ?_) + exact nhdsWithin_eq_nhdsWithin xu ou (by simp only [inter_assoc, inter_self]) + }⟩ + +/-- On open sets, `AnalyticOn` and `AnalyticWithinOn` coincide -/ +@[simp] lemma IsOpen.analyticWithinOn_iff_analyticOn {f : E → F} {s : Set E} (hs : IsOpen s) : + AnalyticWithinOn 𝕜 f s ↔ AnalyticOn 𝕜 f s := by + refine ⟨?_, AnalyticOn.analyticWithinOn⟩ + intro hf x m + rcases Metric.mem_nhds_iff.mp (hs.mem_nhds m) with ⟨r, r0, rs⟩ + rcases hf x m with ⟨p, t, fp⟩ + exact ⟨p, min (.ofReal r) t, { + r_pos := lt_min (by positivity) fp.r_pos + r_le := min_le_of_right_le fp.r_le + hasSum := by + intro y ym + simp only [EMetric.mem_ball, lt_min_iff, edist_lt_ofReal, dist_zero_right] at ym + refine fp.hasSum (rs ?_) ym.2 + simp only [Metric.mem_ball, dist_self_add_left, ym.1] + }⟩ + +/-! +### Equivalence to analyticity of a local extension + +We show that `HasFPowerSeriesWithinOnBall`, `HasFPowerSeriesWithinAt`, and `AnalyticWithinAt` are +equivalent to the existence of a local extension with full analyticity. We do not yet show a +result for `AnalyticWithinOn`, as this requires a bit more work to show that local extensions can +be stitched together. +-/ + +/-- `f` has power series `p` at `x` iff some local extension of `f` has that series -/ +lemma hasFPowerSeriesWithinOnBall_iff_exists_hasFPowerSeriesOnBall [CompleteSpace F] {f : E → F} + {p : FormalMultilinearSeries 𝕜 E F} {s : Set E} {x : E} {r : ℝ≥0∞} : + HasFPowerSeriesWithinOnBall f p s x r ↔ + ContinuousWithinAt f s x ∧ ∃ g, EqOn f g (s ∩ EMetric.ball x r) ∧ + HasFPowerSeriesOnBall g p x r := by + constructor + · intro h + refine ⟨h.continuousWithinAt, fun y ↦ p.sum (y - x), ?_, ?_⟩ + · intro y ⟨ys,yb⟩ + simp only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub] at yb + have e0 := p.hasSum (x := y - x) ?_ + have e1 := (h.hasSum (y := y - x) ?_ ?_) + · simp only [add_sub_cancel] at e1 + exact e1.unique e0 + · simpa only [add_sub_cancel] + · simpa only [EMetric.mem_ball, edist_eq_coe_nnnorm] + · simp only [EMetric.mem_ball, edist_eq_coe_nnnorm] + exact lt_of_lt_of_le yb h.r_le + · refine ⟨h.r_le, h.r_pos, ?_⟩ + intro y lt + simp only [add_sub_cancel_left] + apply p.hasSum + simp only [EMetric.mem_ball] at lt ⊢ + exact lt_of_lt_of_le lt h.r_le + · intro ⟨mem, g, hfg, hg⟩ + refine ⟨hg.r_le, hg.r_pos, ?_, mem⟩ + intro y ys lt + rw [hfg] + · exact hg.hasSum lt + · refine ⟨ys, ?_⟩ + simpa only [EMetric.mem_ball, edist_eq_coe_nnnorm_sub, add_sub_cancel_left, sub_zero] using lt + +/-- `f` has power series `p` at `x` iff some local extension of `f` has that series -/ +lemma hasFPowerSeriesWithinAt_iff_exists_hasFPowerSeriesAt [CompleteSpace F] {f : E → F} + {p : FormalMultilinearSeries 𝕜 E F} {s : Set E} {x : E} : + HasFPowerSeriesWithinAt f p s x ↔ + ContinuousWithinAt f s x ∧ ∃ g, f =ᶠ[𝓝[s] x] g ∧ HasFPowerSeriesAt g p x := by + constructor + · intro ⟨r, h⟩ + rcases hasFPowerSeriesWithinOnBall_iff_exists_hasFPowerSeriesOnBall.mp h with ⟨fc, g, e, h⟩ + refine ⟨fc, g, ?_, ⟨r, h⟩⟩ + refine Filter.eventuallyEq_iff_exists_mem.mpr ⟨_, ?_, e⟩ + exact inter_mem_nhdsWithin _ (EMetric.ball_mem_nhds _ h.r_pos) + · intro ⟨mem, g, hfg, ⟨r, hg⟩⟩ + simp only [eventuallyEq_nhdsWithin_iff, Metric.eventually_nhds_iff] at hfg + rcases hfg with ⟨e, e0, hfg⟩ + refine ⟨min r (.ofReal e), ?_⟩ + refine hasFPowerSeriesWithinOnBall_iff_exists_hasFPowerSeriesOnBall.mpr ⟨mem, g, ?_, ?_⟩ + · intro y ⟨ys, xy⟩ + refine hfg ?_ ys + simp only [EMetric.mem_ball, lt_min_iff, edist_lt_ofReal] at xy + exact xy.2 + · exact hg.mono (lt_min hg.r_pos (by positivity)) (min_le_left _ _) + +/-- `f` is analytic within `s` at `x` iff some local extension of `f` is analytic at `x` -/ +lemma analyticWithinAt_iff_exists_analyticAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} : + AnalyticWithinAt 𝕜 f s x ↔ + ContinuousWithinAt f s x ∧ ∃ g, f =ᶠ[𝓝[s] x] g ∧ AnalyticAt 𝕜 g x := by + simp only [AnalyticWithinAt, AnalyticAt, hasFPowerSeriesWithinAt_iff_exists_hasFPowerSeriesAt] + tauto + +/-- If `f` is analytic within `s` at `x`, some local extension of `f` is analytic at `x` -/ +lemma AnalyticWithinAt.exists_analyticAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} + (h : AnalyticWithinAt 𝕜 f s x) : ∃ g, f x = g x ∧ f =ᶠ[𝓝[s] x] g ∧ AnalyticAt 𝕜 g x := by + by_cases s0 : 𝓝[s] x = ⊥ + · refine ⟨fun _ ↦ f x, rfl, ?_, analyticAt_const⟩ + simp only [EventuallyEq, s0, eventually_bot] + · rcases analyticWithinAt_iff_exists_analyticAt.mp h with ⟨_, g, fg, hg⟩ + refine ⟨g, ?_, fg, hg⟩ + exact tendsto_nhds_unique' ⟨s0⟩ h.continuousWithinAt + (hg.continuousAt.continuousWithinAt.congr' fg.symm) + +/-! +### Congruence + +We require completeness to use equivalence to locally extensions, but this is nonessential. +-/ + +lemma AnalyticWithinAt.congr_of_eventuallyEq [CompleteSpace F] {f g : E → F} {s : Set E} {x : E} + (hf : AnalyticWithinAt 𝕜 f s x) (hs : f =ᶠ[𝓝[s] x] g) (hx : f x = g x) : + AnalyticWithinAt 𝕜 g s x := by + rcases hf.exists_analyticAt with ⟨f', fx, ef, hf'⟩ + rw [analyticWithinAt_iff_exists_analyticAt] + have eg := hs.symm.trans ef + refine ⟨?_, f', eg, hf'⟩ + exact hf'.continuousAt.continuousWithinAt.congr_of_eventuallyEq eg (hx.symm.trans fx) + +lemma AnalyticWithinAt.congr [CompleteSpace F] {f g : E → F} {s : Set E} {x : E} + (hf : AnalyticWithinAt 𝕜 f s x) (hs : EqOn f g s) (hx : f x = g x) : + AnalyticWithinAt 𝕜 g s x := + hf.congr_of_eventuallyEq hs.eventuallyEq_nhdsWithin hx + +lemma AnalyticWithinOn.congr [CompleteSpace F] {f g : E → F} {s : Set E} + (hf : AnalyticWithinOn 𝕜 f s) (hs : EqOn f g s) : + AnalyticWithinOn 𝕜 g s := + fun x m ↦ (hf x m).congr hs (hs m) + +/-! +### Monotonicity w.r.t. the set we're analytic within +-/ + +lemma HasFPowerSeriesWithinOnBall.mono {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} + {s t : Set E} {x : E} {r : ℝ≥0∞} (h : HasFPowerSeriesWithinOnBall f p t x r) + (hs : s ⊆ t) : HasFPowerSeriesWithinOnBall f p s x r where + r_le := h.r_le + r_pos := h.r_pos + hasSum {_} ys yb := h.hasSum (hs ys) yb + continuousWithinAt := h.continuousWithinAt.mono hs + +lemma HasFPowerSeriesWithinAt.mono {f : E → F} {p : FormalMultilinearSeries 𝕜 E F} + {s t : Set E} {x : E} (h : HasFPowerSeriesWithinAt f p t x) + (hs : s ⊆ t) : HasFPowerSeriesWithinAt f p s x := by + rcases h with ⟨r, hr⟩ + exact ⟨r, hr.mono hs⟩ + +lemma AnalyticWithinAt.mono {f : E → F} {s t : Set E} {x : E} (h : AnalyticWithinAt 𝕜 f t x) + (hs : s ⊆ t) : AnalyticWithinAt 𝕜 f s x := by + rcases h with ⟨p, hp⟩ + exact ⟨p, hp.mono hs⟩ + +lemma AnalyticWithinOn.mono {f : E → F} {s t : Set E} (h : AnalyticWithinOn 𝕜 f t) + (hs : s ⊆ t) : AnalyticWithinOn 𝕜 f s := + fun _ m ↦ (h _ (hs m)).mono hs + +/-! +### Analyticity within respects composition + +Currently we require `CompleteSpace`s to use equivalence to local extensions, but this is not +essential. +-/ + +lemma AnalyticWithinAt.comp [CompleteSpace F] [CompleteSpace G] {f : F → G} {g : E → F} {s : Set F} + {t : Set E} {x : E} (hf : AnalyticWithinAt 𝕜 f s (g x)) (hg : AnalyticWithinAt 𝕜 g t x) + (h : MapsTo g t s) : AnalyticWithinAt 𝕜 (f ∘ g) t x := by + rcases hf.exists_analyticAt with ⟨f', _, ef, hf'⟩ + rcases hg.exists_analyticAt with ⟨g', gx, eg, hg'⟩ + refine analyticWithinAt_iff_exists_analyticAt.mpr ⟨?_, f' ∘ g', ?_, ?_⟩ + · exact hf.continuousWithinAt.comp hg.continuousWithinAt h + · have gt := hg.continuousWithinAt.tendsto_nhdsWithin h + filter_upwards [eg, gt.eventually ef] + intro y gy fgy + simp only [Function.comp_apply, fgy, ← gy] + · exact hf'.comp_of_eq hg' gx.symm + +lemma AnalyticWithinOn.comp [CompleteSpace F] [CompleteSpace G] {f : F → G} {g : E → F} {s : Set F} + {t : Set E} (hf : AnalyticWithinOn 𝕜 f s) (hg : AnalyticWithinOn 𝕜 g t) (h : MapsTo g t s) : + AnalyticWithinOn 𝕜 (f ∘ g) t := + fun x m ↦ (hf _ (h m)).comp (hg x m) h + +/-! +### Analyticity within implies smoothness +-/ + +lemma AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} + (h : AnalyticWithinAt 𝕜 f s x) {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x := by + rcases h.exists_analyticAt with ⟨g, fx, fg, hg⟩ + exact hg.contDiffAt.contDiffWithinAt.congr_of_eventuallyEq fg fx + +lemma AnalyticWithinOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E} + (h : AnalyticWithinOn 𝕜 f s) {n : ℕ∞} : ContDiffOn 𝕜 n f s := + fun x m ↦ (h x m).contDiffWithinAt + +/-! +### Analyticity within respects products +-/ + +lemma HasFPowerSeriesWithinOnBall.prod {e : E} {f : E → F} {g : E → G} {s : Set E} {r t : ℝ≥0∞} + {p : FormalMultilinearSeries 𝕜 E F} {q : FormalMultilinearSeries 𝕜 E G} + (hf : HasFPowerSeriesWithinOnBall f p s e r) (hg : HasFPowerSeriesWithinOnBall g q s e t) : + HasFPowerSeriesWithinOnBall (fun x ↦ (f x, g x)) (p.prod q) s e (min r t) where + r_le := by + rw [p.radius_prod_eq_min] + exact min_le_min hf.r_le hg.r_le + r_pos := lt_min hf.r_pos hg.r_pos + hasSum := by + intro y m hy + simp_rw [FormalMultilinearSeries.prod, ContinuousMultilinearMap.prod_apply] + refine (hf.hasSum m ?_).prod_mk (hg.hasSum m ?_) + · exact EMetric.mem_ball.mpr (lt_of_lt_of_le hy (min_le_left _ _)) + · exact EMetric.mem_ball.mpr (lt_of_lt_of_le hy (min_le_right _ _)) + continuousWithinAt := hf.continuousWithinAt.prod hg.continuousWithinAt + +lemma HasFPowerSeriesWithinAt.prod {e : E} {f : E → F} {g : E → G} {s : Set E} + {p : FormalMultilinearSeries 𝕜 E F} {q : FormalMultilinearSeries 𝕜 E G} + (hf : HasFPowerSeriesWithinAt f p s e) (hg : HasFPowerSeriesWithinAt g q s e) : + HasFPowerSeriesWithinAt (fun x ↦ (f x, g x)) (p.prod q) s e := by + rcases hf with ⟨_, hf⟩ + rcases hg with ⟨_, hg⟩ + exact ⟨_, hf.prod hg⟩ + +lemma AnalyticWithinAt.prod {e : E} {f : E → F} {g : E → G} {s : Set E} + (hf : AnalyticWithinAt 𝕜 f s e) (hg : AnalyticWithinAt 𝕜 g s e) : + AnalyticWithinAt 𝕜 (fun x ↦ (f x, g x)) s e := by + rcases hf with ⟨_, hf⟩ + rcases hg with ⟨_, hg⟩ + exact ⟨_, hf.prod hg⟩ + +lemma AnalyticWithinOn.prod {f : E → F} {g : E → G} {s : Set E} + (hf : AnalyticWithinOn 𝕜 f s) (hg : AnalyticWithinOn 𝕜 g s) : + AnalyticWithinOn 𝕜 (fun x ↦ (f x, g x)) s := + fun x hx ↦ (hf x hx).prod (hg x hx) diff --git a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean index 4432211cb0d39..e3f7304c5b371 100644 --- a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean @@ -5,7 +5,7 @@ Authors: Anatole Dedecker -/ import Mathlib.Analysis.Normed.Order.Basic import Mathlib.Analysis.Asymptotics.Asymptotics -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # A collection of specific asymptotic results @@ -134,7 +134,7 @@ theorem Filter.Tendsto.cesaro_smul {E : Type*} [NormedAddCommGroup E] [NormedSpa · filter_upwards [Ici_mem_atTop 1] with n npos have nposℝ : (0 : ℝ) < n := Nat.cast_pos.2 npos simp only [smul_sub, sum_sub_distrib, sum_const, card_range, sub_right_inj] - rw [nsmul_eq_smul_cast ℝ, smul_smul, inv_mul_cancel nposℝ.ne', one_smul] + rw [← Nat.cast_smul_eq_nsmul ℝ, smul_smul, inv_mul_cancel nposℝ.ne', one_smul] · filter_upwards [Ici_mem_atTop 1] with n npos have nposℝ : (0 : ℝ) < n := Nat.cast_pos.2 npos rw [Algebra.id.smul_eq_mul, inv_mul_cancel nposℝ.ne'] diff --git a/Mathlib/Analysis/Asymptotics/Theta.lean b/Mathlib/Analysis/Asymptotics/Theta.lean index 7b1e59915af11..e352e6b93cbfc 100644 --- a/Mathlib/Analysis/Asymptotics/Theta.lean +++ b/Mathlib/Analysis/Asymptotics/Theta.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import Mathlib.Analysis.Asymptotics.Asymptotics -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # Asymptotic equivalence up to a constant diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean index 87a847be0c6bd..1f335549d5cb1 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Yury Kudryashov -/ import Mathlib.Analysis.Calculus.ContDiff.Basic -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension /-! # Infinitely smooth "bump" functions diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 2bae32a5090ab..3d4f7a6ddaa30 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1572,8 +1572,7 @@ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : induction' n using ENat.nat_induction with n IH Itop · intro m hm refine ⟨{ y : R | IsUnit y }, ?_, ?_⟩ - · simp [nhdsWithin_univ] - exact x.nhds + · simpa [nhdsWithin_univ] using x.nhds · use ftaylorSeriesWithin 𝕜 inverse univ rw [le_antisymm hm bot_le, hasFTaylorSeriesUpToOn_zero_iff] constructor diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index 8603960471769..7b0d3e391e4fd 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn -/ import Mathlib.Analysis.Calculus.ContDiff.Basic -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension /-! # Higher differentiability in finite dimensions. diff --git a/Mathlib/Analysis/Calculus/Deriv/Slope.lean b/Mathlib/Analysis/Calculus/Deriv/Slope.lean index 2c4d6c983487b..3b9f51dde30cf 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Slope.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Slope.lean @@ -95,8 +95,7 @@ theorem range_derivWithin_subset_closure_span_image range (derivWithin f s) ⊆ closure (Submodule.span 𝕜 (f '' t)) := by rintro - ⟨x, rfl⟩ rcases eq_or_neBot (𝓝[s \ {x}] x) with H|H - · simp [derivWithin, fderivWithin, H] - exact subset_closure (zero_mem _) + · simpa [derivWithin, fderivWithin, H] using subset_closure (zero_mem _) by_cases H' : DifferentiableWithinAt 𝕜 f s x; swap · rw [derivWithin_zero_of_not_differentiableWithinAt H'] exact subset_closure (zero_mem _) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Linear.lean b/Mathlib/Analysis/Calculus/FDeriv/Linear.lean index 8ce423fb50829..eba65ff473261 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Linear.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Linear.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Sébastien Gouëzel, Yury Kudryashov -/ import Mathlib.Analysis.Calculus.FDeriv.Basic -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps /-! # The derivative of bounded linear maps diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index b8f92c7933150..b02054cbd951d 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -5,8 +5,8 @@ Authors: Sébastien Gouëzel, Yury Kudryashov -/ import Mathlib.Analysis.Calculus.Deriv.Basic import Mathlib.Analysis.Calculus.Deriv.Slope -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.MeasureTheory.Constructions.BorelSpace.ContinuousLinearMap import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic diff --git a/Mathlib/Analysis/Calculus/Implicit.lean b/Mathlib/Analysis/Calculus/Implicit.lean index 27254f8a83e91..72620c2d28d46 100644 --- a/Mathlib/Analysis/Calculus/Implicit.lean +++ b/Mathlib/Analysis/Calculus/Implicit.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Analysis.Calculus.InverseFunctionTheorem.FDeriv import Mathlib.Analysis.Calculus.FDeriv.Add import Mathlib.Analysis.Calculus.FDeriv.Prod -import Mathlib.Analysis.NormedSpace.Complemented +import Mathlib.Analysis.Normed.Module.Complemented /-! # Implicit function theorem diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 21185cf2aaf84..8517d2665d2bd 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Sébastien Gouëzel -/ -import Mathlib.Analysis.NormedSpace.Banach +import Mathlib.Analysis.Normed.Operator.Banach import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace import Mathlib.Topology.PartialHomeomorph /-! diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FiniteDimensional.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FiniteDimensional.lean index ae54ae40b6836..09f63bcf9f6d0 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FiniteDimensional.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FiniteDimensional.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Calculus.InverseFunctionTheorem.ApproximatesLinearOn -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension /-! # A lemma about `ApproximatesLinearOn` that needs `FiniteDimensional` diff --git a/Mathlib/Analysis/Calculus/ParametricIntegral.lean b/Mathlib/Analysis/Calculus/ParametricIntegral.lean index 890a66527d02d..e6088a80f8437 100644 --- a/Mathlib/Analysis/Calculus/ParametricIntegral.lean +++ b/Mathlib/Analysis/Calculus/ParametricIntegral.lean @@ -273,8 +273,7 @@ theorem hasDerivAt_integral_of_dominated_loc_of_lip {F' : α → E} (ε_pos : 0 hF'_int refine ⟨hF'_int, ?_⟩ by_cases hE : CompleteSpace E; swap - · simp [integral, hE] - exact hasDerivAt_const x₀ 0 + · simpa [integral, hE] using hasDerivAt_const x₀ 0 simp_rw [hasDerivAt_iff_hasFDerivAt] at h_diff ⊢ simpa only [(· ∘ ·), ContinuousLinearMap.integral_comp_comm _ hF'_int] using key diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 4447eb2f032ed..46a730bff90c4 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Calculus.LineDeriv.Measurable -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar import Mathlib.Analysis.BoundedVariation import Mathlib.MeasureTheory.Group.Integral diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index bbf1a8eb5de3c..af0dc8bc44f58 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Convex.Topology -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Analysis.SpecificLimits.Basic /-! diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index d08edc10fc519..931e271b760d3 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import Mathlib.Analysis.Complex.CauchyIntegral -import Mathlib.Analysis.NormedSpace.Completion +import Mathlib.Analysis.Normed.Module.Completion import Mathlib.Analysis.NormedSpace.Extr import Mathlib.Topology.Order.ExtrClosure diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 9dfed31969079..be27eb26ad65a 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -185,6 +185,10 @@ theorem nnnorm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : theorem norm_eq_one_of_pow_eq_one {ζ : ℂ} {n : ℕ} (h : ζ ^ n = 1) (hn : n ≠ 0) : ‖ζ‖ = 1 := congr_arg Subtype.val (nnnorm_eq_one_of_pow_eq_one h hn) +lemma le_of_eq_sum_of_eq_sum_norm {ι : Type*} {a b : ℝ} (f : ι → ℂ) (s : Finset ι) (ha₀ : 0 ≤ a) + (ha : a = ∑ i ∈ s, f i) (hb : b = ∑ i ∈ s, (‖f i‖ : ℂ)) : a ≤ b := by + norm_cast at hb; rw [← Complex.abs_of_nonneg ha₀, ha, hb]; exact norm_sum_le s f + theorem equivRealProd_apply_le (z : ℂ) : ‖equivRealProd z‖ ≤ abs z := by simp [Prod.norm_def, abs_re_le_abs, abs_im_le_abs] diff --git a/Mathlib/Analysis/Complex/Conformal.lean b/Mathlib/Analysis/Complex/Conformal.lean index e614087c3457f..9c3c4ee064bd1 100644 --- a/Mathlib/Analysis/Complex/Conformal.lean +++ b/Mathlib/Analysis/Complex/Conformal.lean @@ -5,7 +5,7 @@ Authors: Yourong Zang -/ import Mathlib.Analysis.Complex.Isometry import Mathlib.Analysis.NormedSpace.ConformalLinearMap -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.Data.Complex.FiniteDimensional /-! diff --git a/Mathlib/Analysis/Complex/Hadamard.lean b/Mathlib/Analysis/Complex/Hadamard.lean index e5a142469d29d..ad0f8231ce845 100644 --- a/Mathlib/Analysis/Complex/Hadamard.lean +++ b/Mathlib/Analysis/Complex/Hadamard.lean @@ -395,7 +395,7 @@ lemma norm_le_interp_of_mem_verticalClosedStrip' (f : ℂ → E) {z : ℂ} {a b rw [abs_cpow_eq_rpow_re_of_pos ((Ne.le_iff_lt h1).mp (sSupNormIm_nonneg f _)) _] simp only [sub_re, one_re, le_refl] apply (norm_le_interpStrip_of_mem_verticalClosedStrip f hz hd hB).trans (this.trans _) - apply mul_le_mul_of_le_of_le _ _ (Real.rpow_nonneg (sSupNormIm_nonneg f _) _) + apply mul_le_mul_of_nonneg _ _ (Real.rpow_nonneg (sSupNormIm_nonneg f _) _) · apply (Real.rpow_nonneg _ _) specialize hb 1 simp only [mem_preimage, one_re, mem_singleton_iff, forall_true_left] at hb diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index de4412f55a026..13ec934118ccc 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -5,7 +5,7 @@ Authors: Yury G. Kudryashov -/ import Mathlib.Analysis.Complex.CauchyIntegral import Mathlib.Analysis.Calculus.FDeriv.Analytic -import Mathlib.Analysis.NormedSpace.Completion +import Mathlib.Analysis.Normed.Module.Completion /-! # Liouville's theorem diff --git a/Mathlib/Analysis/Complex/Polynomial.lean b/Mathlib/Analysis/Complex/Polynomial/Basic.lean similarity index 100% rename from Mathlib/Analysis/Complex/Polynomial.lean rename to Mathlib/Analysis/Complex/Polynomial/Basic.lean diff --git a/Mathlib/Analysis/Complex/Polynomial/UnitTrinomial.lean b/Mathlib/Analysis/Complex/Polynomial/UnitTrinomial.lean new file mode 100644 index 0000000000000..1e925a903dc63 --- /dev/null +++ b/Mathlib/Analysis/Complex/Polynomial/UnitTrinomial.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2022 Thomas Browning. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Browning +-/ +import Mathlib.Algebra.Polynomial.UnitTrinomial +import Mathlib.Analysis.Complex.Polynomial.Basic + +/-! +# Irreducibility of unit trinomials + +## TODO + +Develop more theory (e.g., it suffices to check that `aeval z p ≠ 0` for `z = 0` and `z` a root of +unity). +-/ + +namespace Polynomial.IsUnitTrinomial +variable {p : ℤ[X]} + +/-- A unit trinomial is irreducible if it has no complex roots in common with its mirror. -/ +theorem irreducible_of_coprime' (hp : IsUnitTrinomial p) + (h : ∀ z : ℂ, ¬(aeval z p = 0 ∧ aeval z (mirror p) = 0)) : Irreducible p := by + refine hp.irreducible_of_coprime fun q hq hq' => ?_ + suffices ¬0 < q.natDegree by + rcases hq with ⟨p, rfl⟩ + replace hp := hp.leadingCoeff_isUnit + rw [leadingCoeff_mul] at hp + replace hp := isUnit_of_mul_isUnit_left hp + rw [not_lt, Nat.le_zero] at this + rwa [eq_C_of_natDegree_eq_zero this, isUnit_C, ← this] + intro hq'' + rw [natDegree_pos_iff_degree_pos] at hq'' + rw [← degree_map_eq_of_injective (algebraMap ℤ ℂ).injective_int] at hq'' + cases' Complex.exists_root hq'' with z hz + rw [IsRoot, eval_map, ← aeval_def] at hz + refine h z ⟨?_, ?_⟩ + · cases' hq with g' hg' + rw [hg', aeval_mul, hz, zero_mul] + · cases' hq' with g' hg' + rw [hg', aeval_mul, hz, zero_mul] + +end Polynomial.IsUnitTrinomial diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index 03344352df02b..bd775685ea4b8 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -273,7 +273,7 @@ theorem convex_Iio (r : β) : Convex 𝕜 (Iio r) := by _ = r := Convex.combo_self hab _ theorem convex_Ioi (r : β) : Convex 𝕜 (Ioi r) := - @convex_Iio 𝕜 βᵒᵈ _ _ _ _ r + convex_Iio (β := βᵒᵈ) r theorem convex_Ioo (r s : β) : Convex 𝕜 (Ioo r s) := Ioi_inter_Iio.subst ((convex_Ioi r).inter <| convex_Iio s) @@ -325,27 +325,27 @@ theorem MonotoneOn.convex_lt (hf : MonotoneOn f s) (hs : Convex 𝕜 s) (r : β) theorem MonotoneOn.convex_ge (hf : MonotoneOn f s) (hs : Convex 𝕜 s) (r : β) : Convex 𝕜 ({ x ∈ s | r ≤ f x }) := - @MonotoneOn.convex_le 𝕜 Eᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual hs r + MonotoneOn.convex_le (E := Eᵒᵈ) (β := βᵒᵈ) hf.dual hs r theorem MonotoneOn.convex_gt (hf : MonotoneOn f s) (hs : Convex 𝕜 s) (r : β) : Convex 𝕜 ({ x ∈ s | r < f x }) := - @MonotoneOn.convex_lt 𝕜 Eᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual hs r + MonotoneOn.convex_lt (E := Eᵒᵈ) (β := βᵒᵈ) hf.dual hs r theorem AntitoneOn.convex_le (hf : AntitoneOn f s) (hs : Convex 𝕜 s) (r : β) : Convex 𝕜 ({ x ∈ s | f x ≤ r }) := - @MonotoneOn.convex_ge 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r + MonotoneOn.convex_ge (β := βᵒᵈ) hf hs r theorem AntitoneOn.convex_lt (hf : AntitoneOn f s) (hs : Convex 𝕜 s) (r : β) : Convex 𝕜 ({ x ∈ s | f x < r }) := - @MonotoneOn.convex_gt 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r + MonotoneOn.convex_gt (β := βᵒᵈ) hf hs r theorem AntitoneOn.convex_ge (hf : AntitoneOn f s) (hs : Convex 𝕜 s) (r : β) : Convex 𝕜 ({ x ∈ s | r ≤ f x }) := - @MonotoneOn.convex_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r + MonotoneOn.convex_le (β := βᵒᵈ) hf hs r theorem AntitoneOn.convex_gt (hf : AntitoneOn f s) (hs : Convex 𝕜 s) (r : β) : Convex 𝕜 ({ x ∈ s | r < f x }) := - @MonotoneOn.convex_lt 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r + MonotoneOn.convex_lt (β := βᵒᵈ) hf hs r theorem Monotone.convex_le (hf : Monotone f) (r : β) : Convex 𝕜 { x | f x ≤ r } := Set.sep_univ.subst ((hf.monotoneOn univ).convex_le convex_univ r) diff --git a/Mathlib/Analysis/Convex/Body.lean b/Mathlib/Analysis/Convex/Body.lean index b288fe49b26a3..e8edd08e74e9f 100644 --- a/Mathlib/Analysis/Convex/Body.lean +++ b/Mathlib/Analysis/Convex/Body.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul A. Reichert -/ import Mathlib.Analysis.Convex.Basic -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Topology.MetricSpace.HausdorffDistance /-! diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 255e081170c57..a14b272ac8de7 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -424,8 +424,7 @@ theorem convexHull_prod (s : Set E) (t : Set F) : prod_subset_iff.2 fun _ hx _ => mk_mem_convexHull_prod hx theorem convexHull_add (s t : Set E) : convexHull R (s + t) = convexHull R s + convexHull R t := by - simp_rw [← image2_add, ← image_prod, ← IsLinearMap.isLinearMap_add.image_convexHull, - convexHull_prod] + simp_rw [← add_image_prod, ← IsLinearMap.isLinearMap_add.image_convexHull, convexHull_prod] variable (R E) @@ -549,13 +548,14 @@ lemma AffineIndependent.convexHull_inter' (hs : AffineIndependent R ((↑) : ↑ end section pi -variable {𝕜 ι : Type*} {E : ι → Type*} [Fintype ι] [LinearOrderedField 𝕜] +variable {𝕜 ι : Type*} {E : ι → Type*} [Finite ι] [LinearOrderedField 𝕜] [Π i, AddCommGroup (E i)] [Π i, Module 𝕜 (E i)] {s : Set ι} {t : Π i, Set (E i)} {x : Π i, E i} open Finset Fintype lemma mem_convexHull_pi (h : ∀ i ∈ s, x i ∈ convexHull 𝕜 (t i)) : x ∈ convexHull 𝕜 (s.pi t) := by - wlog hs : s = Set.univ + cases nonempty_fintype ι + wlog hs : s = Set.univ generalizing s t · rw [← pi_univ_ite] refine this (fun i _ ↦ ?_) rfl split_ifs with hi diff --git a/Mathlib/Analysis/Convex/Function.lean b/Mathlib/Analysis/Convex/Function.lean index c3782b07a850e..261e230a5eb97 100644 --- a/Mathlib/Analysis/Convex/Function.lean +++ b/Mathlib/Analysis/Convex/Function.lean @@ -90,6 +90,32 @@ theorem concaveOn_id {s : Set β} (hs : Convex 𝕜 s) : ConcaveOn 𝕜 s _root_ intros rfl⟩ +section congr + +variable {g : E → β} + +theorem ConvexOn.congr (hf : ConvexOn 𝕜 s f) (hfg : EqOn f g s) : ConvexOn 𝕜 s g := + ⟨hf.1, fun x hx y hy a b ha hb hab => by + simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha hb hab)] using hf.2 hx hy ha hb hab⟩ + +theorem ConcaveOn.congr (hf : ConcaveOn 𝕜 s f) (hfg : EqOn f g s) : ConcaveOn 𝕜 s g := + ⟨hf.1, fun x hx y hy a b ha hb hab => by + simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha hb hab)] using hf.2 hx hy ha hb hab⟩ + +theorem StrictConvexOn.congr (hf : StrictConvexOn 𝕜 s f) (hfg : EqOn f g s) : + StrictConvexOn 𝕜 s g := + ⟨hf.1, fun x hx y hy hxy a b ha hb hab => by + simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha.le hb.le hab)] using + hf.2 hx hy hxy ha hb hab⟩ + +theorem StrictConcaveOn.congr (hf : StrictConcaveOn 𝕜 s f) (hfg : EqOn f g s) : + StrictConcaveOn 𝕜 s g := + ⟨hf.1, fun x hx y hy hxy a b ha hb hab => by + simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha.le hb.le hab)] using + hf.2 hx hy hxy ha hb hab⟩ + +end congr + theorem ConvexOn.subset {t : Set E} (hf : ConvexOn 𝕜 t f) (hst : s ⊆ t) (hs : Convex 𝕜 s) : ConvexOn 𝕜 s f := ⟨hs, fun _ hx _ hy => hf.2 (hst hx) (hst hy)⟩ @@ -185,6 +211,14 @@ theorem convexOn_const (c : β) (hs : Convex 𝕜 s) : ConvexOn 𝕜 s fun _ : E theorem concaveOn_const (c : β) (hs : Convex 𝕜 s) : ConcaveOn 𝕜 s fun _ => c := convexOn_const (β := βᵒᵈ) _ hs +theorem ConvexOn.add_const (hf : ConvexOn 𝕜 s f) (b : β) : + ConvexOn 𝕜 s (f + fun _ => b) := + hf.add (convexOn_const _ hf.1) + +theorem ConcaveOn.add_const (hf : ConcaveOn 𝕜 s f) (b : β) : + ConcaveOn 𝕜 s (f + fun _ => b) := + hf.add (concaveOn_const _ hf.1) + theorem convexOn_of_convex_epigraph (h : Convex 𝕜 { p : E × β | p.1 ∈ s ∧ f p.1 ≤ p.2 }) : ConvexOn 𝕜 s f := ⟨fun x hx y hy a b ha hb hab => (@h (x, f x) ⟨hx, le_rfl⟩ (y, f y) ⟨hy, le_rfl⟩ a b ha hb hab).1, @@ -479,6 +513,14 @@ theorem StrictConcaveOn.add (hf : StrictConcaveOn 𝕜 s f) (hg : StrictConcaveO StrictConcaveOn 𝕜 s (f + g) := hf.dual.add hg +theorem StrictConvexOn.add_const {γ : Type*} {f : E → γ} [OrderedCancelAddCommMonoid γ] + [Module 𝕜 γ] (hf : StrictConvexOn 𝕜 s f) (b : γ) : StrictConvexOn 𝕜 s (f + fun _ => b) := + hf.add_convexOn (convexOn_const _ hf.1) + +theorem StrictConcaveOn.add_const {γ : Type*} {f : E → γ} [OrderedCancelAddCommMonoid γ] + [Module 𝕜 γ] (hf : StrictConcaveOn 𝕜 s f) (b : γ) : StrictConcaveOn 𝕜 s (f + fun _ => b) := + hf.add_concaveOn (concaveOn_const _ hf.1) + end DistribMulAction section Module diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index 4f571cf775603..e11bb4b933256 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -30,7 +30,7 @@ open Fintype Finset Set namespace Convex -variable {ι 𝕜 E : Type*} [DecidableEq ι] [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] +variable {ι 𝕜 E : Type*} [LinearOrderedField 𝕜] [AddCommGroup E] [Module 𝕜 E] /-- **Radon's theorem on convex sets**. @@ -67,7 +67,7 @@ open FiniteDimensional variable [FiniteDimensional 𝕜 E] /-- Corner case for `helly_theorem'`. -/ -private lemma helly_theorem_corner {F : ι → Set E} {s : Finset ι} [DecidableEq ι] +private lemma helly_theorem_corner {F : ι → Set E} {s : Finset ι} (h_card_small : s.card ≤ finrank 𝕜 E + 1) (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := h_inter s (by simp) h_card_small @@ -80,6 +80,7 @@ theorem helly_theorem' {F : ι → Set E} {s : Finset ι} (h_convex : ∀ i ∈ s, Convex 𝕜 (F i)) (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := by + classical obtain h_card | h_card := lt_or_le s.card (finrank 𝕜 E + 1) · exact helly_theorem_corner (le_of_lt h_card) h_inter generalize hn : s.card = n @@ -88,7 +89,7 @@ theorem helly_theorem' {F : ι → Set E} {s : Finset ι} · exact helly_theorem_corner (le_of_eq hn) h_inter /- Construct a family of vectors indexed by `ι` such that the vector corresponding to `i : ι` is an arbitrary element of the intersection of all `F j` except `F i`. -/ - let a (i : s) : E := Set.Nonempty.some (s := ⋂ j ∈ (s.erase i), F j) <| by + let a (i : s) : E := Set.Nonempty.some (s := ⋂ j ∈ s.erase i, F j) <| by apply hk (s := s.erase i) · exact fun i hi ↦ h_convex i (mem_of_mem_erase hi) · intro J hJ_ss hJ_card @@ -186,6 +187,7 @@ theorem helly_theorem_compact' [TopologicalSpace E] [T2Space E] {F : ι → Set (h_convex : ∀ i, Convex 𝕜 (F i)) (h_compact : ∀ i, IsCompact (F i)) (h_inter : ∀ I : Finset ι, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i, F i).Nonempty := by + classical /- If `ι` is empty the statement is trivial. -/ cases' isEmpty_or_nonempty ι with _ h_nonempty · simp only [iInter_of_empty, Set.univ_nonempty] diff --git a/Mathlib/Analysis/Convex/StrictConvexSpace.lean b/Mathlib/Analysis/Convex/StrictConvexSpace.lean index 3fbfd8251e952..ea048d3b233ce 100644 --- a/Mathlib/Analysis/Convex/StrictConvexSpace.lean +++ b/Mathlib/Analysis/Convex/StrictConvexSpace.lean @@ -8,7 +8,7 @@ import Mathlib.Analysis.Convex.Strict import Mathlib.Analysis.Normed.Order.Basic import Mathlib.Analysis.NormedSpace.AddTorsor import Mathlib.Analysis.NormedSpace.Pointwise -import Mathlib.Analysis.NormedSpace.Ray +import Mathlib.Analysis.Normed.Module.Ray /-! # Strictly convex spaces diff --git a/Mathlib/Analysis/NormedSpace/Star/Basic.lean b/Mathlib/Analysis/CstarAlgebra/Basic.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Star/Basic.lean rename to Mathlib/Analysis/CstarAlgebra/Basic.lean index f323691c26004..deaff8d8703e5 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Basic.lean +++ b/Mathlib/Analysis/CstarAlgebra/Basic.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Analysis.Normed.Group.Hom -import Mathlib.Analysis.NormedSpace.Basic -import Mathlib.Analysis.NormedSpace.LinearIsometry +import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.Analysis.Normed.Operator.LinearIsometry import Mathlib.Algebra.Star.SelfAdjoint import Mathlib.Algebra.Star.Subalgebra import Mathlib.Algebra.Star.Unitary diff --git a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Basic.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Basic.lean index 5d79a9ac114a5..fce7f222738c0 100644 --- a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Star.GelfandDuality +import Mathlib.Analysis.CstarAlgebra.GelfandDuality import Mathlib.Topology.Algebra.StarSubalgebra /-! # Continuous functional calculus diff --git a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Instances.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 5a00fa924c60f..eeed22259eeb0 100644 --- a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -3,11 +3,11 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus.Restrict -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus -import Mathlib.Analysis.NormedSpace.Star.Spectrum -import Mathlib.Analysis.NormedSpace.Star.Unitization -import Mathlib.Topology.ContinuousFunction.UniqueCFC +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Restrict +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Basic +import Mathlib.Analysis.CstarAlgebra.Spectrum +import Mathlib.Analysis.CstarAlgebra.Unitization +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique /-! # Instances of the continuous functional calculus @@ -467,7 +467,7 @@ lemma spectrum_star_mul_self_nonneg {b : A} : ∀ x ∈ spectrum ℝ (star b * b refine SpectrumRestricts.nnreal_add ?_ ?_ ?_ h_c_spec₀ · exact IsSelfAdjoint.smul (by rfl) <| ((ℜ c).prop.pow 2).add ((ℑ c).prop.pow 2) · exact (IsSelfAdjoint.star_mul_self c).neg - · rw [nsmul_eq_smul_cast ℝ] + · rw [← Nat.cast_smul_eq_nsmul ℝ] refine (ℜ c).2.sq_spectrumRestricts.nnreal_add ((ℜ c).2.pow 2) ((ℑ c).2.pow 2) (ℑ c).2.sq_spectrumRestricts |>.smul_of_nonneg <| by norm_num have h_c_spec₂ : SpectrumRestricts (star c * c) ContinuousMap.realToNNReal := by diff --git a/Mathlib/Topology/ContinuousFunction/NonUnitalFunctionalCalculus.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean similarity index 98% rename from Mathlib/Topology/ContinuousFunction/NonUnitalFunctionalCalculus.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 33da281b44754..e47f681cce41e 100644 --- a/Mathlib/Topology/ContinuousFunction/NonUnitalFunctionalCalculus.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -6,7 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Algebra.Quasispectrum import Mathlib.Topology.ContinuousFunction.Compact import Mathlib.Topology.ContinuousFunction.ContinuousMapZero -import Mathlib.Topology.ContinuousFunction.FunctionalCalculus +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital import Mathlib.Topology.UniformSpace.CompactConvergence /-! @@ -14,10 +14,11 @@ import Mathlib.Topology.UniformSpace.CompactConvergence This file defines a generic API for the *continuous functional calculus* in *non-unital* algebras which is suitable in a wide range of settings. The design is intended to match as closely as -possible that for unital algebras in `Mathlib.Topology.ContinuousFunction.FunctionalCalculus`. -Changes to either file should be mirrored in its counterpart whenever possible. The underlying -reasons for the design decisions in the unital case apply equally in the non-unital case. See the -module documentation in that file for more information. +possible that for unital algebras in +`Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital`. Changes to either file should +be mirrored in its counterpart whenever possible. The underlying reasons for the design decisions in +the unital case apply equally in the non-unital case. See the module documentation in that file for +more information. A continuous functional calculus for an element `a : A` in a non-unital topological `R`-algebra is a continuous extension of the polynomial functional calculus (i.e., `Polynomial.aeval`) for @@ -181,7 +182,7 @@ junk value `0`. This is the primary declaration intended for widespread use of the continuous functional calculus for non-unital algebras, and all the API applies to this declaration. For more information, see the -module documentation for `Topology.ContinuousFunction.FunctionalCalculus`. -/ +module documentation for `Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital`. -/ noncomputable irreducible_def cfcₙ (f : R → R) (a : A) : A := if h : p a ∧ ContinuousOn f (σₙ R a) ∧ f 0 = 0 then cfcₙHom h.1 ⟨⟨_, h.2.1.restrict⟩, h.2.2⟩ diff --git a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Order.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Order.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Order.lean index aa2694ee6a41e..c407b6c5369aa 100644 --- a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ -import Mathlib.Analysis.NormedSpace.Star.ContinuousFunctionalCalculus.Instances +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Instances import Mathlib.Topology.ContinuousFunction.StarOrdered -import Mathlib.Analysis.NormedSpace.Star.Unitization +import Mathlib.Analysis.CstarAlgebra.Unitization /-! # Facts about star-ordered rings that depend on the continuous functional calculus diff --git a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Restrict.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Restrict.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index 630a9f4bb66f8..304e1476b9402 100644 --- a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Spectrum -import Mathlib.Topology.ContinuousFunction.NonUnitalFunctionalCalculus +import Mathlib.Analysis.Normed.Algebra.Spectrum +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital /-! # Restriction of the continuous functional calculus to a scalar subring diff --git a/Mathlib/Topology/ContinuousFunction/UniqueCFC.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unique.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/UniqueCFC.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 0e0056ab86909..4b6726cf226cb 100644 --- a/Mathlib/Topology/ContinuousFunction/UniqueCFC.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Spectrum -import Mathlib.Topology.ContinuousFunction.NonUnitalFunctionalCalculus +import Mathlib.Analysis.Normed.Algebra.Spectrum +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital import Mathlib.Topology.ContinuousFunction.StoneWeierstrass /-! diff --git a/Mathlib/Topology/ContinuousFunction/FunctionalCalculus.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unital.lean similarity index 99% rename from Mathlib/Topology/ContinuousFunction/FunctionalCalculus.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unital.lean index fd69722f360b3..aa701997722b1 100644 --- a/Mathlib/Topology/ContinuousFunction/FunctionalCalculus.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -171,8 +171,9 @@ properties that it is a continuous star algebra homomorphism mapping the (restri identity to `a`. This is the necessary tool used to establish `cfcHom_comp` and the more common variant `cfc_comp`. -This class has instances, which can be found in `Mathlib.Topology.ContinuousFunction.UniqueCFC`, in -each of the common cases `ℂ`, `ℝ` and `ℝ≥0` as a consequence of the Stone-Weierstrass theorem. +This class has instances, which can be found in +`Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique`, in each of the common cases +`ℂ`, `ℝ` and `ℝ≥0` as a consequence of the Stone-Weierstrass theorem. This class is separate from `ContinuousFunctionalCalculus` primarily because we will later use `SpectrumRestricts` to derive an instance of `ContinuousFunctionalCalculus` on a scalar subring @@ -269,7 +270,7 @@ not continuous on the spectrum of `a`, then `cfc f a` returns the junk value `0` This is the primary declaration intended for widespread use of the continuous functional calculus, and all the API applies to this declaration. For more information, see the module documentation -for `Topology.ContinuousFunction.FunctionalCalculus`. -/ +for `Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital`. -/ noncomputable irreducible_def cfc (f : R → R) (a : A) : A := if h : p a ∧ ContinuousOn f (spectrum R a) then cfcHom h.1 ⟨_, h.2.restrict⟩ @@ -836,7 +837,7 @@ lemma algebraMap_le_of_le_spectrum {r : R} {a : A} (h : ∀ x ∈ spectrum R a, exact algebraMap_le_cfc id r a h lemma cfc_le_one (f : R → R) (a : A) (h : ∀ x ∈ spectrum R a, f x ≤ 1) : cfc f a ≤ 1 := by - apply cfc_cases (· ≤ 1) _ _ (by simpa using star_mul_self_nonneg (1 : A)) fun hf ha ↦ ?_ + apply cfc_cases (· ≤ 1) _ _ (by simp) fun hf ha ↦ ?_ rw [← map_one (cfcHom ha (R := R))] apply cfcHom_mono ha simpa [ContinuousMap.le_def] using h diff --git a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Unitary.lean b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unitary.lean similarity index 96% rename from Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Unitary.lean rename to Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unitary.lean index af231ee75eec9..b9be85d92cdb3 100644 --- a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Unitary.lean +++ b/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unitary.lean @@ -5,7 +5,7 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.Complex.Circle import Mathlib.Tactic.Peel -import Mathlib.Topology.ContinuousFunction.FunctionalCalculus +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital /-! # Conditions on unitary elements imposed by the continuous functional calculus diff --git a/Mathlib/Analysis/NormedSpace/Star/Exponential.lean b/Mathlib/Analysis/CstarAlgebra/Exponential.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Star/Exponential.lean rename to Mathlib/Analysis/CstarAlgebra/Exponential.lean diff --git a/Mathlib/Analysis/NormedSpace/Star/GelfandDuality.lean b/Mathlib/Analysis/CstarAlgebra/GelfandDuality.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Star/GelfandDuality.lean rename to Mathlib/Analysis/CstarAlgebra/GelfandDuality.lean index 181f350647bd8..0db306dbcac02 100644 --- a/Mathlib/Analysis/NormedSpace/Star/GelfandDuality.lean +++ b/Mathlib/Analysis/CstarAlgebra/GelfandDuality.lean @@ -3,9 +3,9 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Star.Spectrum +import Mathlib.Analysis.CstarAlgebra.Spectrum import Mathlib.Analysis.Normed.Group.Quotient -import Mathlib.Analysis.NormedSpace.Algebra +import Mathlib.Analysis.Normed.Algebra.Basic import Mathlib.Topology.ContinuousFunction.Units import Mathlib.Topology.ContinuousFunction.Compact import Mathlib.Topology.Algebra.Algebra diff --git a/Mathlib/Analysis/NormedSpace/Star/Matrix.lean b/Mathlib/Analysis/CstarAlgebra/Matrix.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Star/Matrix.lean rename to Mathlib/Analysis/CstarAlgebra/Matrix.lean diff --git a/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean b/Mathlib/Analysis/CstarAlgebra/Multiplier.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Star/Multiplier.lean rename to Mathlib/Analysis/CstarAlgebra/Multiplier.lean index 0d11731c83106..bc423401f3149 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean +++ b/Mathlib/Analysis/CstarAlgebra/Multiplier.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux, Jon Bannon -/ import Mathlib.Analysis.NormedSpace.OperatorNorm.Completeness -import Mathlib.Analysis.NormedSpace.Star.Unitization +import Mathlib.Analysis.CstarAlgebra.Unitization import Mathlib.Analysis.SpecialFunctions.Pow.NNReal /-! @@ -447,7 +447,7 @@ maps `Lₐ Rₐ : A →L[𝕜] A` given by left- and right-multiplication by `a` Warning: if `A = 𝕜`, then this is a coercion which is not definitionally equal to the `algebraMap 𝕜 𝓜(𝕜, 𝕜)` coercion, but these are propositionally equal. See `DoubleCentralizer.coe_eq_algebraMap` below. -/ --- Porting note: added `noncomputable` because an IR check failed? +-- Porting note: added `noncomputable`; IR check does not recognise `ContinuousLinearMap.mul` @[coe] protected noncomputable def coe (a : A) : 𝓜(𝕜, A) := { fst := ContinuousLinearMap.mul 𝕜 A a diff --git a/Mathlib/Analysis/NormedSpace/Star/Spectrum.lean b/Mathlib/Analysis/CstarAlgebra/Spectrum.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/Star/Spectrum.lean rename to Mathlib/Analysis/CstarAlgebra/Spectrum.lean index e304de3349988..0b1a907133349 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Spectrum.lean +++ b/Mathlib/Analysis/CstarAlgebra/Spectrum.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Star.Basic -import Mathlib.Analysis.NormedSpace.Spectrum +import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.SpecialFunctions.Exponential import Mathlib.Algebra.Star.StarAlgHom diff --git a/Mathlib/Analysis/NormedSpace/Star/Unitization.lean b/Mathlib/Analysis/CstarAlgebra/Unitization.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/Star/Unitization.lean rename to Mathlib/Analysis/CstarAlgebra/Unitization.lean index e92da57355c79..2b9a2688c3699 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Unitization.lean +++ b/Mathlib/Analysis/CstarAlgebra/Unitization.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Star.Basic -import Mathlib.Analysis.NormedSpace.Unitization +import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.Normed.Algebra.Unitization /-! # The minimal unitization of a C⋆-algebra This file shows that when `E` is a C⋆-algebra (over a densely normed field `𝕜`), that the minimal diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 1409bbe01e2f9..85133c9920df7 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -257,25 +257,13 @@ instance instNSMul : SMul ℕ 𝓢(E, F) := ⟨fun c f => { toFun := c • (f : E → F) smooth' := (f.smooth _).const_smul c - decay' := by - have : c • (f : E → F) = (c : ℝ) • f := by - ext x - simp only [Pi.smul_apply, smul_apply] - exact nsmul_eq_smul_cast _ _ _ - simp only [this] - exact ((c : ℝ) • f).decay' }⟩ + decay' := by simpa [← Nat.cast_smul_eq_nsmul ℝ] using ((c : ℝ) • f).decay' }⟩ instance instZSMul : SMul ℤ 𝓢(E, F) := ⟨fun c f => { toFun := c • (f : E → F) smooth' := (f.smooth _).const_smul c - decay' := by - have : c • (f : E → F) = (c : ℝ) • f := by - ext x - simp only [Pi.smul_apply, smul_apply] - exact zsmul_eq_smul_cast _ _ _ - simp only [this] - exact ((c : ℝ) • f).decay' }⟩ + decay' := by simpa [← Int.cast_smul_eq_nsmul ℝ] using ((c : ℝ) • f).decay' }⟩ end SMul @@ -604,8 +592,7 @@ def _root_.MeasureTheory.Measure.integrablePower (μ : Measure D) : ℕ := lemma integrable_pow_neg_integrablePower (μ : Measure D) [h : μ.HasTemperateGrowth] : Integrable (fun x ↦ (1 + ‖x‖) ^ (- (μ.integrablePower : ℝ))) μ := by - simp [Measure.integrablePower, h] - exact h.exists_integrable.choose_spec + simpa [Measure.integrablePower, h] using h.exists_integrable.choose_spec instance _root_.MeasureTheory.Measure.IsFiniteMeasure.instHasTemperateGrowth {μ : Measure D} [h : IsFiniteMeasure μ] : μ.HasTemperateGrowth := ⟨⟨0, by simp⟩⟩ diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 2c7c0ae99634b..ebe6b648e9ec1 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -54,7 +54,7 @@ open Set Function Finset MeasureTheory Measure Filter noncomputable section -variable {ι : Type*} [Fintype ι] [DecidableEq ι] +variable {ι : Type*} [Fintype ι] local prefix:max "#" => Fintype.card @@ -65,6 +65,10 @@ variable {A : ι → Type*} [∀ i, MeasurableSpace (A i)] namespace MeasureTheory +section DecidableEq + +variable [DecidableEq ι] + namespace GridLines /-- The "grid-lines operation" (not a standard name) which is central in the inductive proof of the @@ -278,6 +282,8 @@ theorem lintegral_prod_lintegral_pow_le convert lintegral_mul_prod_lintegral_pow_le μ h2 h3 hf using 2 field_simp +end DecidableEq + /-! ## The Gagliardo-Nirenberg-Sobolev inequality -/ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] @@ -294,6 +300,7 @@ theorem lintegral_pow_le_pow_lintegral_fderiv_aux {u : (ι → ℝ) → F} (hu : ContDiff ℝ 1 u) (h2u : HasCompactSupport u) : ∫⁻ x, (‖u x‖₊ : ℝ≥0∞) ^ p ≤ (∫⁻ x, ‖fderiv ℝ u x‖₊) ^ p := by + classical /- For a function `f` in one variable and `t ∈ ℝ` we have `|f(t)| = `|∫_{-∞}^t Df(s)∂s| ≤ ∫_ℝ |Df(s)| ∂s` where we use the fundamental theorem of calculus. For each `x ∈ ℝⁿ` we let `u` vary in one of the `n` coordinates and apply the inequality above. @@ -477,8 +484,8 @@ theorem snorm_le_snorm_fderiv_of_eq_inner {u : E → F'} have h2p : (p : ℝ) < n := by have : 0 < p⁻¹ - (n : ℝ)⁻¹ := NNReal.coe_lt_coe.mpr (pos_iff_ne_zero.mpr (inv_ne_zero hp'0)) |>.trans_eq hp' - simp [sub_pos] at this - rwa [inv_lt_inv _ (zero_lt_one.trans_le (NNReal.coe_le_coe.mpr hp))] at this + rwa [NNReal.coe_inv, sub_pos, + inv_lt_inv _ (zero_lt_one.trans_le (NNReal.coe_le_coe.mpr hp))] at this exact_mod_cast hn have h0n : 2 ≤ n := Nat.succ_le_of_lt <| Nat.one_lt_cast.mp <| hp.trans_lt h2p have hn : NNReal.IsConjExponent n n' := .conjExponent (by norm_cast) diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index bf585884576da..752bd81941924 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -7,8 +7,8 @@ import Mathlib.Algebra.DirectSum.Module import Mathlib.Algebra.Module.LinearMap.Basic import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.Convex.Uniform -import Mathlib.Analysis.NormedSpace.Completion -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps +import Mathlib.Analysis.Normed.Module.Completion +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps /-! # Inner product space diff --git a/Mathlib/Analysis/InnerProductSpace/Dual.lean b/Mathlib/Analysis/InnerProductSpace/Dual.lean index 7a1712104d21f..e6af204a8c575 100644 --- a/Mathlib/Analysis/InnerProductSpace/Dual.lean +++ b/Mathlib/Analysis/InnerProductSpace/Dual.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Analysis.InnerProductSpace.Projection -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual /-! # The Fréchet-Riesz representation theorem diff --git a/Mathlib/Analysis/InnerProductSpace/NormPow.lean b/Mathlib/Analysis/InnerProductSpace/NormPow.lean index bf01e70b29384..1d53811b6834c 100644 --- a/Mathlib/Analysis/InnerProductSpace/NormPow.lean +++ b/Mathlib/Analysis/InnerProductSpace/NormPow.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Heather Macbeth -/ import Mathlib.Analysis.InnerProductSpace.Calculus -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual import Mathlib.Analysis.SpecialFunctions.Pow.Deriv /-! @@ -29,7 +29,7 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] theorem hasFDerivAt_norm_rpow (x : E) {p : ℝ} (hp : 1 < p) : HasFDerivAt (fun x : E ↦ ‖x‖ ^ p) ((p * ‖x‖ ^ (p - 2)) • innerSL ℝ x) x := by by_cases hx : x = 0 - · simp [hx] + · simp only [hx, norm_zero, map_zero, smul_zero] have h2p : 0 < p - 1 := sub_pos.mpr hp rw [HasFDerivAt, hasFDerivAtFilter_iff_isLittleO] calc (fun x : E ↦ ‖x‖ ^ p - ‖(0 : E)‖ ^ p - 0) @@ -46,7 +46,7 @@ theorem hasFDerivAt_norm_rpow (x : E) {p : ℝ} (hp : 1 < p) : simp_rw [mul_one, isBigO_norm_left (f' := fun x ↦ x), sub_zero, isBigO_refl] · apply HasStrictFDerivAt.hasFDerivAt convert (hasStrictFDerivAt_norm_sq x).rpow_const (p := p / 2) (by simp [hx]) using 0 - simp_rw [← Real.rpow_natCast_mul (norm_nonneg _), nsmul_eq_smul_cast ℝ, smul_smul] + simp_rw [← Real.rpow_natCast_mul (norm_nonneg _), ← Nat.cast_smul_eq_nsmul ℝ, smul_smul] ring_nf -- doesn't close the goal? congr! 2 ring diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index 62c5a45c134b1..4629fbeeaa6c4 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Sébastien Gouëzel, Heather Macbeth -/ import Mathlib.Analysis.InnerProductSpace.Projection -import Mathlib.Analysis.NormedSpace.PiLp +import Mathlib.Analysis.Normed.Lp.PiLp import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.LinearAlgebra.UnitaryGroup @@ -649,7 +649,8 @@ def Complex.isometryOfOrthonormal (v : OrthonormalBasis (Fin 2) ℝ F) : ℂ ≃ @[simp] theorem Complex.map_isometryOfOrthonormal (v : OrthonormalBasis (Fin 2) ℝ F) (f : F ≃ₗᵢ[ℝ] F') : Complex.isometryOfOrthonormal (v.map f) = (Complex.isometryOfOrthonormal v).trans f := by - simp [Complex.isometryOfOrthonormal, LinearIsometryEquiv.trans_assoc, OrthonormalBasis.map] + simp only [isometryOfOrthonormal, OrthonormalBasis.map, LinearIsometryEquiv.symm_trans, + LinearIsometryEquiv.symm_symm] -- Porting note: `LinearIsometryEquiv.trans_assoc` doesn't trigger in the `simp` above rw [LinearIsometryEquiv.trans_assoc] @@ -941,7 +942,7 @@ variable {m n : Type*} namespace Matrix -variable [Fintype m] [Fintype n] [DecidableEq n] +variable [Fintype n] [DecidableEq n] /-- `Matrix.toLin'` adapted for `EuclideanSpace 𝕜 _`. -/ def toEuclideanLin : Matrix m n 𝕜 ≃ₗ[𝕜] EuclideanSpace 𝕜 n →ₗ[𝕜] EuclideanSpace 𝕜 m := @@ -975,13 +976,13 @@ theorem toEuclideanLin_apply_piLp_equiv_symm (M : Matrix m n 𝕜) (v : n → rfl -- `Matrix.toEuclideanLin` is the same as `Matrix.toLin` applied to `PiLp.basisFun`, -theorem toEuclideanLin_eq_toLin : +theorem toEuclideanLin_eq_toLin [Finite m] : (toEuclideanLin : Matrix m n 𝕜 ≃ₗ[𝕜] _) = Matrix.toLin (PiLp.basisFun _ _ _) (PiLp.basisFun _ _ _) := rfl open EuclideanSpace in -lemma toEuclideanLin_eq_toLin_orthonormal : +lemma toEuclideanLin_eq_toLin_orthonormal [Fintype m] : toEuclideanLin = toLin (basisFun n 𝕜).toBasis (basisFun m 𝕜).toBasis := rfl diff --git a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean index 5ec5a19767405..c38eea54c37bc 100644 --- a/Mathlib/Analysis/InnerProductSpace/ProdL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/ProdL2.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ import Mathlib.Analysis.InnerProductSpace.PiL2 -import Mathlib.Analysis.NormedSpace.ProdLp +import Mathlib.Analysis.Normed.Lp.ProdLp /-! # `L²` inner product space structure on products of inner product spaces diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index 68224a1d9faaa..a9ba99321c951 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll, Frédéric Dupuis, Heather Macbeth -/ import Mathlib.Analysis.InnerProductSpace.Basic -import Mathlib.Analysis.NormedSpace.Banach +import Mathlib.Analysis.Normed.Operator.Banach import Mathlib.LinearAlgebra.SesquilinearForm /-! diff --git a/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean b/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean index 3c224a67e5e93..34f154ad37514 100644 --- a/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean +++ b/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean @@ -5,7 +5,7 @@ Authors: Frédéric Dupuis -/ import Mathlib.Analysis.InnerProductSpace.Dual -import Mathlib.Analysis.NormedSpace.WeakOperatorTopology +import Mathlib.Analysis.Normed.Operator.WeakOperatorTopology /-! # The weak operator topology in Hilbert spaces diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index 9917be34248f1..5980076f937fc 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ import Mathlib.Analysis.InnerProductSpace.Projection -import Mathlib.Analysis.NormedSpace.lpSpace +import Mathlib.Analysis.Normed.Lp.lpSpace import Mathlib.Analysis.InnerProductSpace.PiL2 /-! @@ -55,7 +55,7 @@ We also define a *predicate* `IsHilbertSum 𝕜 G V`, where `V : Π i, G i → ## Main results * `lp.instInnerProductSpace`: Construction of the inner product space instance on the Hilbert sum - `lp G 2`. Note that from the file `Analysis.NormedSpace.lpSpace`, the space `lp G 2` already + `lp G 2`. Note that from the file `Analysis.Normed.Lp.lpSpace`, the space `lp G 2` already held a normed space instance (`lp.normedSpace`), and if each `G i` is a Hilbert space (i.e., complete), then `lp G 2` was already known to be complete (`lp.completeSpace`). So the work here is to define the inner product and show it is compatible. diff --git a/Mathlib/Analysis/LocallyConvex/Barrelled.lean b/Mathlib/Analysis/LocallyConvex/Barrelled.lean index c40472dc09c0e..c92b0399ae1c7 100644 --- a/Mathlib/Analysis/LocallyConvex/Barrelled.lean +++ b/Mathlib/Analysis/LocallyConvex/Barrelled.lean @@ -15,7 +15,7 @@ Banach-Steinhaus theorem for maps from a barrelled space to a space equipped wit of seminorms generating the topology (i.e `WithSeminorms q` for some family of seminorms `q`). The more standard Banach-Steinhaus theorem for normed spaces is then deduced from that in -`Analysis.NormedSpace.BanachSteinhaus`. +`Analysis.Normed.Operator.BanachSteinhaus`. ## Main definitions diff --git a/Mathlib/Analysis/LocallyConvex/Basic.lean b/Mathlib/Analysis/LocallyConvex/Basic.lean index 7a41fb06bc4b4..c5bf0dfa6a4e3 100644 --- a/Mathlib/Analysis/LocallyConvex/Basic.lean +++ b/Mathlib/Analysis/LocallyConvex/Basic.lean @@ -5,7 +5,7 @@ Authors: Jean Lo, Bhavik Mehta, Yaël Dillies -/ import Mathlib.Analysis.Convex.Basic import Mathlib.Analysis.Convex.Hull -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Topology.Bornology.Absorbs /-! diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index ac667ab6c7305..56d0adc4ea7c3 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth, Eric Wieser -/ -import Mathlib.Analysis.NormedSpace.PiLp +import Mathlib.Analysis.Normed.Lp.PiLp import Mathlib.Analysis.InnerProductSpace.PiL2 /-! @@ -44,7 +44,7 @@ of a matrix. The norm induced by the identification of `Matrix m n 𝕜` with `EuclideanSpace n 𝕜 →L[𝕜] EuclideanSpace m 𝕜` (i.e., the ℓ² operator norm) can be found in -`Analysis.NormedSpace.Star.Matrix`. It is separated to avoid extraneous imports in this file. +`Analysis.CstarAlgebra.Matrix`. It is separated to avoid extraneous imports in this file. -/ noncomputable section @@ -403,7 +403,7 @@ private theorem norm_unitOf (a : α) : ‖unitOf a‖₊ = 1 := by set_option tactic.skipAssignedInstances false in private theorem mul_unitOf (a : α) : a * unitOf a = algebraMap _ _ (‖a‖₊ : ℝ) := by - simp [unitOf] + simp only [unitOf, coe_nnnorm] split_ifs with h · simp [h] · rw [mul_smul_comm, mul_inv_cancel h, Algebra.algebraMap_eq_smul_one] diff --git a/Mathlib/Analysis/NormedSpace/Algebra.lean b/Mathlib/Analysis/Normed/Algebra/Basic.lean similarity index 94% rename from Mathlib/Analysis/NormedSpace/Algebra.lean rename to Mathlib/Analysis/Normed/Algebra/Basic.lean index 27f24cb09fcb2..d44b75699bb05 100644 --- a/Mathlib/Analysis/NormedSpace/Algebra.lean +++ b/Mathlib/Analysis/Normed/Algebra/Basic.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Topology.Algebra.Module.CharacterSpace -import Mathlib.Analysis.NormedSpace.WeakDual -import Mathlib.Analysis.NormedSpace.Spectrum +import Mathlib.Analysis.Normed.Module.WeakDual +import Mathlib.Analysis.Normed.Algebra.Spectrum /-! # Normed algebras diff --git a/Mathlib/Analysis/NormedSpace/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Spectrum.lean rename to Mathlib/Analysis/Normed/Algebra/Spectrum.lean index f07d7a10b0fb6..1d83a00d0260b 100644 --- a/Mathlib/Analysis/NormedSpace/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -6,11 +6,11 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Algebra.Quasispectrum import Mathlib.FieldTheory.IsAlgClosed.Spectrum import Mathlib.Analysis.Complex.Liouville -import Mathlib.Analysis.Complex.Polynomial +import Mathlib.Analysis.Complex.Polynomial.Basic import Mathlib.Analysis.Analytic.RadiusLiminf import Mathlib.Topology.Algebra.Module.CharacterSpace import Mathlib.Analysis.NormedSpace.Exponential -import Mathlib.Analysis.NormedSpace.UnitizationL1 +import Mathlib.Analysis.Normed.Algebra.UnitizationL1 import Mathlib.Tactic.ContinuousFunctionalCalculus /-! diff --git a/Mathlib/Analysis/NormedSpace/TrivSqZeroExt.lean b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/TrivSqZeroExt.lean rename to Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean index a01763371a27d..eaa57d0c34175 100644 --- a/Mathlib/Analysis/NormedSpace/TrivSqZeroExt.lean +++ b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Analysis.NormedSpace.Exponential -import Mathlib.Analysis.NormedSpace.ProdLp +import Mathlib.Analysis.Normed.Lp.ProdLp import Mathlib.Topology.Instances.TrivSqZeroExt /-! @@ -81,8 +81,9 @@ variable [Field 𝕜] [CharZero 𝕜] [Ring R] [AddCommGroup M] theorem snd_expSeries_of_smul_comm (x : tsze R M) (hx : MulOpposite.op x.fst • x.snd = x.fst • x.snd) (n : ℕ) : snd (expSeries 𝕜 (tsze R M) (n + 1) fun _ => x) = (expSeries 𝕜 R n fun _ => x.fst) • x.snd := by - simp_rw [expSeries_apply_eq, snd_smul, snd_pow_of_smul_comm _ _ hx, nsmul_eq_smul_cast 𝕜 (n + 1), - smul_smul, smul_assoc, Nat.factorial_succ, Nat.pred_succ, Nat.cast_mul, mul_inv_rev, + simp_rw [expSeries_apply_eq, snd_smul, snd_pow_of_smul_comm _ _ hx, + ← Nat.cast_smul_eq_nsmul 𝕜 (n + 1), smul_smul, smul_assoc, Nat.factorial_succ, Nat.pred_succ, + Nat.cast_mul, mul_inv_rev, inv_mul_cancel_right₀ ((Nat.cast_ne_zero (R := 𝕜)).mpr <| Nat.succ_ne_zero n)] /-- If `exp R x.fst` converges to `e` then `(exp R x).snd` converges to `e • x.snd`. -/ diff --git a/Mathlib/Analysis/NormedSpace/Unitization.lean b/Mathlib/Analysis/Normed/Algebra/Unitization.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Unitization.lean rename to Mathlib/Analysis/Normed/Algebra/Unitization.lean diff --git a/Mathlib/Analysis/NormedSpace/UnitizationL1.lean b/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/UnitizationL1.lean rename to Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean index 1b074ef996d59..f92a28cdee905 100644 --- a/Mathlib/Analysis/NormedSpace/UnitizationL1.lean +++ b/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Algebra.Algebra.Unitization -import Mathlib.Analysis.NormedSpace.ProdLp +import Mathlib.Analysis.Normed.Lp.ProdLp /-! # Unitization equipped with the $L^1$ norm diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index 0e0dd4c98725b..d2126e4827d15 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -196,7 +196,7 @@ instance Prod.normOneClass [SeminormedAddCommGroup α] [One α] [NormOneClass α instance Pi.normOneClass {ι : Type*} {α : ι → Type*} [Nonempty ι] [Fintype ι] [∀ i, SeminormedAddCommGroup (α i)] [∀ i, One (α i)] [∀ i, NormOneClass (α i)] : NormOneClass (∀ i, α i) := - ⟨by simp [Pi.norm_def]; exact Finset.sup_const Finset.univ_nonempty 1⟩ + ⟨by simpa [Pi.norm_def] using Finset.sup_const Finset.univ_nonempty 1⟩ instance MulOpposite.normOneClass [SeminormedAddCommGroup α] [One α] [NormOneClass α] : NormOneClass αᵐᵒᵖ := diff --git a/Mathlib/Analysis/Normed/Group/Quotient.lean b/Mathlib/Analysis/Normed/Group/Quotient.lean index ff3dcfdaed0ba..7406b8e2cf248 100644 --- a/Mathlib/Analysis/Normed/Group/Quotient.lean +++ b/Mathlib/Analysis/Normed/Group/Quotient.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Riccardo Brasca -/ -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Analysis.Normed.Group.Hom import Mathlib.RingTheory.Ideal.QuotientOperations import Mathlib.Topology.MetricSpace.HausdorffDistance diff --git a/Mathlib/Analysis/NormedSpace/LpEquiv.lean b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/LpEquiv.lean rename to Mathlib/Analysis/Normed/Lp/LpEquiv.lean index 8a1b21ffd89f6..d43a285b86665 100644 --- a/Mathlib/Analysis/NormedSpace/LpEquiv.lean +++ b/Mathlib/Analysis/Normed/Lp/LpEquiv.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.lpSpace -import Mathlib.Analysis.NormedSpace.PiLp +import Mathlib.Analysis.Normed.Lp.lpSpace +import Mathlib.Analysis.Normed.Lp.PiLp import Mathlib.Topology.ContinuousFunction.Bounded /-! diff --git a/Mathlib/Analysis/NormedSpace/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/PiLp.lean rename to Mathlib/Analysis/Normed/Lp/PiLp.lean index b0ea7df8fc7e7..5ceff479c7c2d 100644 --- a/Mathlib/Analysis/NormedSpace/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel, Jireh Loreaux import Mathlib.Analysis.MeanInequalities import Mathlib.Data.Fintype.Order import Mathlib.LinearAlgebra.Matrix.Basis -import Mathlib.Analysis.NormedSpace.WithLp +import Mathlib.Analysis.Normed.Lp.WithLp /-! # `L^p` distance on finite products of metric spaces @@ -615,8 +615,8 @@ instance instBoundedSMul [SeminormedRing 𝕜] [∀ i, SeminormedAddCommGroup ( exact nnnorm_smul_le c (WithLp.equiv ∞ (∀ i, β i) f) · have hp0 : 0 < p.toReal := zero_lt_one.trans_le hp have hpt : p ≠ ⊤ := p.toReal_pos_iff_ne_top.mp hp0 - rw [nnnorm_eq_sum hpt, nnnorm_eq_sum hpt, one_div, NNReal.rpow_inv_le_iff hp0, NNReal.mul_rpow, - ← NNReal.rpow_mul, inv_mul_cancel hp0.ne', NNReal.rpow_one, Finset.mul_sum] + rw [nnnorm_eq_sum hpt, nnnorm_eq_sum hpt, one_div, NNReal.rpow_inv_le_iff hp0, + NNReal.mul_rpow, ← NNReal.rpow_mul, inv_mul_cancel hp0.ne', NNReal.rpow_one, Finset.mul_sum] simp_rw [← NNReal.mul_rpow, smul_apply] exact Finset.sum_le_sum fun i _ => NNReal.rpow_le_rpow (nnnorm_smul_le _ _) hp0.le diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/Normed/Lp/ProdLp.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/ProdLp.lean rename to Mathlib/Analysis/Normed/Lp/ProdLp.lean index 043e005fe84be..76b7315100d34 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/Normed/Lp/ProdLp.lean @@ -5,7 +5,7 @@ Authors: Moritz Doll, Sébastien Gouëzel, Jireh Loreaux -/ import Mathlib.Analysis.MeanInequalities -import Mathlib.Analysis.NormedSpace.WithLp +import Mathlib.Analysis.Normed.Lp.WithLp /-! # `L^p` distance on products of two metric spaces @@ -32,7 +32,7 @@ statements for the coordinate functions, for instance. # Implementation notes -This file is a straight-forward adaptation of `Mathlib.Analysis.NormedSpace.PiLp`. +This file is a straight-forward adaptation of `Mathlib.Analysis.Normed.Lp.PiLp`. -/ diff --git a/Mathlib/Analysis/NormedSpace/WithLp.lean b/Mathlib/Analysis/Normed/Lp/WithLp.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/WithLp.lean rename to Mathlib/Analysis/Normed/Lp/WithLp.lean diff --git a/Mathlib/Analysis/NormedSpace/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/lpSpace.lean rename to Mathlib/Analysis/Normed/Lp/lpSpace.lean diff --git a/Mathlib/Analysis/NormedSpace/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean similarity index 67% rename from Mathlib/Analysis/NormedSpace/Basic.lean rename to Mathlib/Analysis/Normed/Module/Basic.lean index df276bbecf987..fe58c989cdb6a 100644 --- a/Mathlib/Analysis/NormedSpace/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -493,3 +493,178 @@ def NormedAlgebra.restrictScalars : NormedAlgebra 𝕜 E := end NormedAlgebra end RestrictScalars + +section Core +/-! +### Structures for constructing new normed spaces + +This section contains tools meant for constructing new normed spaces. These allow one to easily +construct all the relevant instances (distances measures, etc) while proving only a minimal +set of axioms. Furthermore, tools are provided to add a norm structure to a type that already +has a preexisting uniformity or bornology: in such cases, it is necessary to keep the preexisting +instances, while ensuring that the norm induces the same uniformity/bornology. +-/ + +open scoped Uniformity Bornology + +/-- A structure encapsulating minimal axioms needed to defined a seminormed vector space, as found +in textbooks. This is meant to be used to easily define `SeminormedAddCommGroup E` instances from +scratch on a type with no preexisting distance or topology. -/ +structure SeminormedAddCommGroup.Core (𝕜 : Type*) (E : Type*) [NormedField 𝕜] [AddCommGroup E] + [Norm E] [Module 𝕜 E] : Prop where + norm_nonneg (x : E) : 0 ≤ ‖x‖ + norm_smul (c : 𝕜) (x : E) : ‖c • x‖ = ‖c‖ * ‖x‖ + norm_triangle (x y : E) : ‖x + y‖ ≤ ‖x‖ + ‖y‖ + +/-- Produces a `PseudoMetricSpace E` instance from a `SeminormedAddCommGroup.Core`. Note that +if this is used to define an instance on a type, it also provides a new uniformity and +topology on the type. See note [reducible non-instances]. -/ +abbrev PseudoMetricSpace.ofSeminormedAddCommGroupCore {𝕜 E : Type*} [NormedField 𝕜] [AddCommGroup E] + [Norm E] [Module 𝕜 E] (core : SeminormedAddCommGroup.Core 𝕜 E) : + PseudoMetricSpace E where + dist x y := ‖x - y‖ + dist_self x := by + show ‖x - x‖ = 0 + simp only [sub_self] + have : (0 : E) = (0 : 𝕜) • (0 : E) := by simp + rw [this, core.norm_smul] + simp + dist_comm x y := by + show ‖x - y‖ = ‖y - x‖ + have : y - x = (-1 : 𝕜) • (x - y) := by simp + rw [this, core.norm_smul] + simp + dist_triangle x y z := by + show ‖x - z‖ ≤ ‖x - y‖ + ‖y - z‖ + have : x - z = (x - y) + (y - z) := by abel + rw [this] + exact core.norm_triangle _ _ + edist_dist x y := by exact (ENNReal.ofReal_eq_coe_nnreal _).symm + +/-- Produces a `PseudoEMetricSpace E` instance from a `SeminormedAddCommGroup.Core`. Note that +if this is used to define an instance on a type, it also provides a new uniformity and +topology on the type. See note [reducible non-instances]. -/ +abbrev PseudoEMetricSpace.ofSeminormedAddCommGroupCore {𝕜 E : Type*} [NormedField 𝕜] + [AddCommGroup E] [Norm E] [Module 𝕜 E] + (core : SeminormedAddCommGroup.Core 𝕜 E) : PseudoEMetricSpace E := + (PseudoMetricSpace.ofSeminormedAddCommGroupCore core).toPseudoEMetricSpace + +/-- Produces a `PseudoEMetricSpace E` instance from a `SeminormedAddCommGroup.Core` on a type that +already has an existing uniform space structure. This requires a proof that the uniformity induced +by the norm is equal to the preexisting uniformity. See note [reducible non-instances]. -/ +abbrev PseudoMetricSpace.ofSeminormedAddCommGroupCoreReplaceUniformity {𝕜 E : Type*} [NormedField 𝕜] + [AddCommGroup E] [Norm E] [Module 𝕜 E] [U : UniformSpace E] + (core : SeminormedAddCommGroup.Core 𝕜 E) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace + (self := PseudoEMetricSpace.ofSeminormedAddCommGroupCore core)]) : + PseudoMetricSpace E := + .replaceUniformity (.ofSeminormedAddCommGroupCore core) H + +open Bornology in +/-- Produces a `PseudoEMetricSpace E` instance from a `SeminormedAddCommGroup.Core` on a type that +already has a preexisting uniform space structure and a preexisting bornology. This requires proofs +that the uniformity induced by the norm is equal to the preexisting uniformity, and likewise for +the bornology. See note [reducible non-instances]. -/ +abbrev PseudoMetricSpace.ofSeminormedAddCommGroupCoreReplaceAll {𝕜 E : Type*} [NormedField 𝕜] + [AddCommGroup E] [Norm E] [Module 𝕜 E] [U : UniformSpace E] [B : Bornology E] + (core : SeminormedAddCommGroup.Core 𝕜 E) + (HU : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace + (self := PseudoEMetricSpace.ofSeminormedAddCommGroupCore core)]) + (HB : ∀ s : Set E, @IsBounded _ B s + ↔ @IsBounded _ (PseudoMetricSpace.ofSeminormedAddCommGroupCore core).toBornology s) : + PseudoMetricSpace E := + .replaceBornology (.replaceUniformity (.ofSeminormedAddCommGroupCore core) HU) HB + +/-- Produces a `SeminormedAddCommGroup E` instance from a `SeminormedAddCommGroup.Core`. Note that +if this is used to define an instance on a type, it also provides a new distance measure from the +norm. it must therefore not be used on a type with a preexisting distance measure or topology. +See note [reducible non-instances]. -/ +abbrev SeminormedAddCommGroup.ofCore {𝕜 : Type*} {E : Type*} [NormedField 𝕜] [AddCommGroup E] + [Norm E] [Module 𝕜 E] (core : SeminormedAddCommGroup.Core 𝕜 E) : SeminormedAddCommGroup E := + { PseudoMetricSpace.ofSeminormedAddCommGroupCore core with } + +/-- Produces a `SeminormedAddCommGroup E` instance from a `SeminormedAddCommGroup.Core` on a type +that already has an existing uniform space structure. This requires a proof that the uniformity +induced by the norm is equal to the preexisting uniformity. See note [reducible non-instances]. -/ +abbrev SeminormedAddCommGroup.ofCoreReplaceUniformity {𝕜 : Type*} {E : Type*} [NormedField 𝕜] + [AddCommGroup E] [Norm E] [Module 𝕜 E] [U : UniformSpace E] + (core : SeminormedAddCommGroup.Core 𝕜 E) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace + (self := PseudoEMetricSpace.ofSeminormedAddCommGroupCore core)]) : + SeminormedAddCommGroup E := + { PseudoMetricSpace.ofSeminormedAddCommGroupCoreReplaceUniformity core H with } + +open Bornology in +/-- Produces a `SeminormedAddCommGroup E` instance from a `SeminormedAddCommGroup.Core` on a type +that already has a preexisting uniform space structure and a preexisting bornology. This requires +proofs that the uniformity induced by the norm is equal to the preexisting uniformity, and likewise +for the bornology. See note [reducible non-instances]. -/ +abbrev SeminormedAddCommGroup.ofCoreReplaceAll {𝕜 : Type*} {E : Type*} [NormedField 𝕜] + [AddCommGroup E] [Norm E] [Module 𝕜 E] [U : UniformSpace E] [B : Bornology E] + (core : SeminormedAddCommGroup.Core 𝕜 E) + (HU : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace + (self := PseudoEMetricSpace.ofSeminormedAddCommGroupCore core)]) + (HB : ∀ s : Set E, @IsBounded _ B s + ↔ @IsBounded _ (PseudoMetricSpace.ofSeminormedAddCommGroupCore core).toBornology s) : + SeminormedAddCommGroup E := + { PseudoMetricSpace.ofSeminormedAddCommGroupCoreReplaceAll core HU HB with } + +/-- A structure encapsulating minimal axioms needed to defined a normed vector space, as found +in textbooks. This is meant to be used to easily define `NormedAddCommGroup E` and `NormedSpace E` +instances from scratch on a type with no preexisting distance or topology. -/ +structure NormedSpace.Core (𝕜 : Type*) (E : Type*) [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] + [Norm E] extends SeminormedAddCommGroup.Core 𝕜 E : Prop where + norm_eq_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 + +variable {𝕜 : Type*} {E : Type*} [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] [Norm E] + +/-- Produces a `NormedAddCommGroup E` instance from a `NormedSpace.Core`. Note that if this is +used to define an instance on a type, it also provides a new distance measure from the norm. +it must therefore not be used on a type with a preexisting distance measure. +See note [reducible non-instances]. -/ +abbrev NormedAddCommGroup.ofCore (core : NormedSpace.Core 𝕜 E) : NormedAddCommGroup E := + { SeminormedAddCommGroup.ofCore core.toCore with + eq_of_dist_eq_zero := by + intro x y h + rw [← sub_eq_zero, ← core.norm_eq_zero_iff] + exact h } + +/-- Produces a `NormedAddCommGroup E` instance from a `NormedAddCommGroup.Core` on a type +that already has an existing uniform space structure. This requires a proof that the uniformity +induced by the norm is equal to the preexisting uniformity. See note [reducible non-instances]. -/ +abbrev NormedAddCommGroup.ofCoreReplaceUniformity [U : UniformSpace E] (core : NormedSpace.Core 𝕜 E) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace + (self := PseudoEMetricSpace.ofSeminormedAddCommGroupCore core.toCore)]) : + NormedAddCommGroup E := + { SeminormedAddCommGroup.ofCoreReplaceUniformity core.toCore H with + eq_of_dist_eq_zero := by + intro x y h + rw [← sub_eq_zero, ← core.norm_eq_zero_iff] + exact h } + +open Bornology in +/-- Produces a `NormedAddCommGroup E` instance from a `NormedAddCommGroup.Core` on a type +that already has a preexisting uniform space structure and a preexisting bornology. This requires +proofs that the uniformity induced by the norm is equal to the preexisting uniformity, and likewise +for the bornology. See note [reducible non-instances]. -/ +abbrev NormedAddCommGroup.ofCoreReplaceAll [U : UniformSpace E] [B : Bornology E] + (core : NormedSpace.Core 𝕜 E) + (HU : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace + (self := PseudoEMetricSpace.ofSeminormedAddCommGroupCore core.toCore)]) + (HB : ∀ s : Set E, @IsBounded _ B s + ↔ @IsBounded _ (PseudoMetricSpace.ofSeminormedAddCommGroupCore core.toCore).toBornology s) : + NormedAddCommGroup E := + { SeminormedAddCommGroup.ofCoreReplaceAll core.toCore HU HB with + eq_of_dist_eq_zero := by + intro x y h + rw [← sub_eq_zero, ← core.norm_eq_zero_iff] + exact h } + +/-- Produces a `NormedSpace 𝕜 E` instance from a `NormedSpace.Core`. This is meant to be used +on types where the `NormedAddCommGroup E` instance has also been defined using `core`. +See note [reducible non-instances]. -/ +abbrev NormedSpace.ofCore {𝕜 : Type*} {E : Type*} [NormedField 𝕜] [SeminormedAddCommGroup E] + [Module 𝕜 E] (core : NormedSpace.Core 𝕜 E) : NormedSpace 𝕜 E where + norm_smul_le r x := by rw [core.norm_smul r x] + +end Core diff --git a/Mathlib/Analysis/NormedSpace/Complemented.lean b/Mathlib/Analysis/Normed/Module/Complemented.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Complemented.lean rename to Mathlib/Analysis/Normed/Module/Complemented.lean index 019b3c21606c1..a909039a4c387 100644 --- a/Mathlib/Analysis/NormedSpace/Complemented.lean +++ b/Mathlib/Analysis/Normed/Module/Complemented.lean @@ -3,7 +3,7 @@ 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.Analysis.NormedSpace.Banach +import Mathlib.Analysis.Normed.Operator.Banach import Mathlib.Topology.Algebra.Module.FiniteDimension /-! diff --git a/Mathlib/Analysis/NormedSpace/Completion.lean b/Mathlib/Analysis/Normed/Module/Completion.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Completion.lean rename to Mathlib/Analysis/Normed/Module/Completion.lean diff --git a/Mathlib/Analysis/NormedSpace/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Dual.lean rename to Mathlib/Analysis/Normed/Module/Dual.lean diff --git a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/FiniteDimension.lean rename to Mathlib/Analysis/Normed/Module/FiniteDimension.lean diff --git a/Mathlib/Analysis/NormedSpace/Ray.lean b/Mathlib/Analysis/Normed/Module/Ray.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Ray.lean rename to Mathlib/Analysis/Normed/Module/Ray.lean diff --git a/Mathlib/Analysis/NormedSpace/Span.lean b/Mathlib/Analysis/Normed/Module/Span.lean similarity index 95% rename from Mathlib/Analysis/NormedSpace/Span.lean rename to Mathlib/Analysis/Normed/Module/Span.lean index b4919b263a67f..2996a80f34809 100644 --- a/Mathlib/Analysis/NormedSpace/Span.lean +++ b/Mathlib/Analysis/Normed/Module/Span.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ -import Mathlib.Analysis.NormedSpace.LinearIsometry -import Mathlib.Analysis.NormedSpace.ContinuousLinearMap -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Operator.LinearIsometry +import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap +import Mathlib.Analysis.Normed.Module.Basic /-! # The span of a single vector diff --git a/Mathlib/Analysis/NormedSpace/WeakDual.lean b/Mathlib/Analysis/Normed/Module/WeakDual.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/WeakDual.lean rename to Mathlib/Analysis/Normed/Module/WeakDual.lean index 8bbedf55c95a6..71dc2da624481 100644 --- a/Mathlib/Analysis/NormedSpace/WeakDual.lean +++ b/Mathlib/Analysis/Normed/Module/WeakDual.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä, Yury Kudryashov -/ -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual import Mathlib.Analysis.NormedSpace.OperatorNorm.Completeness /-! diff --git a/Mathlib/Analysis/NormedSpace/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Banach.lean rename to Mathlib/Analysis/Normed/Operator/Banach.lean diff --git a/Mathlib/Analysis/NormedSpace/BanachSteinhaus.lean b/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/BanachSteinhaus.lean rename to Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean diff --git a/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean rename to Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean index c4124aa813d1a..b43fd4cefc393 100644 --- a/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ import Mathlib.Analysis.NormedSpace.Multilinear.Basic -import Mathlib.Analysis.NormedSpace.Units +import Mathlib.Analysis.Normed.Ring.Units import Mathlib.Analysis.NormedSpace.OperatorNorm.Completeness import Mathlib.Analysis.NormedSpace.OperatorNorm.Mul diff --git a/Mathlib/Analysis/NormedSpace/CompactOperator.lean b/Mathlib/Analysis/Normed/Operator/Compact.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/CompactOperator.lean rename to Mathlib/Analysis/Normed/Operator/Compact.lean diff --git a/Mathlib/Analysis/NormedSpace/ContinuousLinearMap.lean b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/ContinuousLinearMap.lean rename to Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean diff --git a/Mathlib/Analysis/NormedSpace/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/LinearIsometry.lean rename to Mathlib/Analysis/Normed/Operator/LinearIsometry.lean diff --git a/Mathlib/Analysis/NormedSpace/WeakOperatorTopology.lean b/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean similarity index 98% rename from Mathlib/Analysis/NormedSpace/WeakOperatorTopology.lean rename to Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean index e5e51b23f4d5d..b7a5ea7136815 100644 --- a/Mathlib/Analysis/NormedSpace/WeakOperatorTopology.lean +++ b/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean @@ -5,7 +5,7 @@ Authors: Frédéric Dupuis -/ import Mathlib.Analysis.LocallyConvex.WithSeminorms -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual /-! # The weak operator topology @@ -53,7 +53,8 @@ open scoped Topology `E →WOT[𝕜] F`. -/ @[irreducible] def ContinuousLinearMapWOT (𝕜 : Type*) (E : Type*) (F : Type*) [Semiring 𝕜] [AddCommGroup E] - [TopologicalSpace E] [Module 𝕜 E] [AddCommGroup F] [TopologicalSpace F] [Module 𝕜 F] := E →L[𝕜] F + [TopologicalSpace E] [Module 𝕜 E] [AddCommGroup F] [TopologicalSpace F] [Module 𝕜 F] := + E →L[𝕜] F @[inherit_doc] notation:25 E " →WOT[" 𝕜 "]" F => ContinuousLinearMapWOT 𝕜 E F @@ -214,7 +215,7 @@ all `x` and `y`. -/ def seminorm (x : E) (y : F⋆) : Seminorm 𝕜 (E →WOT[𝕜] F) where toFun A := ‖y (A x)‖ map_zero' := by simp - add_le' A B := by simp; exact norm_add_le _ _ + add_le' A B := by simpa using norm_add_le _ _ neg' A := by simp smul' r A := by simp @@ -227,7 +228,7 @@ def seminormFamily : SeminormFamily 𝕜 (E →WOT[𝕜] F) (E × F⋆) := lemma hasBasis_seminorms : (𝓝 (0 : E →WOT[𝕜] F)).HasBasis (seminormFamily 𝕜 E F).basisSets id := by let p := seminormFamily 𝕜 E F rw [nhds_induced, nhds_pi] - simp [map_zero, zero_apply] + simp only [map_zero, Pi.zero_apply] have h := Filter.hasBasis_pi (fun _ : (E × F⋆) ↦ Metric.nhds_basis_ball (x := 0)) |>.comap (inducingFn 𝕜 E F) refine h.to_hasBasis' ?_ ?_ diff --git a/Mathlib/Analysis/NormedSpace/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/Units.lean rename to Mathlib/Analysis/Normed/Ring/Units.lean diff --git a/Mathlib/Analysis/NormedSpace/AddTorsor.lean b/Mathlib/Analysis/NormedSpace/AddTorsor.lean index 81f14c9cd40d2..070c79f4057a1 100644 --- a/Mathlib/Analysis/NormedSpace/AddTorsor.lean +++ b/Mathlib/Analysis/NormedSpace/AddTorsor.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Yury Kudryashov -/ import Mathlib.Algebra.CharP.Invertible -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Analysis.Normed.Group.AddTorsor import Mathlib.LinearAlgebra.AffineSpace.AffineSubspace import Mathlib.Topology.Instances.RealVectorSpace @@ -194,8 +194,8 @@ end invertibleTwo variable (𝕜) in theorem dist_pointReflection_right (p q : P) : dist (Equiv.pointReflection p q) q = ‖(2 : 𝕜)‖ * dist p q := by - simp [dist_eq_norm_vsub V, Equiv.pointReflection_vsub_right (G := V), - nsmul_eq_smul_cast 𝕜, norm_smul] + simp [dist_eq_norm_vsub V, Equiv.pointReflection_vsub_right (G := V), ← Nat.cast_smul_eq_nsmul 𝕜, + norm_smul] variable (𝕜) in theorem dist_right_pointReflection (p q : P) : diff --git a/Mathlib/Analysis/NormedSpace/AddTorsorBases.lean b/Mathlib/Analysis/NormedSpace/AddTorsorBases.lean index 6fd629c82244c..1dc65e797cba1 100644 --- a/Mathlib/Analysis/NormedSpace/AddTorsorBases.lean +++ b/Mathlib/Analysis/NormedSpace/AddTorsorBases.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional /-! diff --git a/Mathlib/Analysis/NormedSpace/AffineIsometry.lean b/Mathlib/Analysis/NormedSpace/AffineIsometry.lean index b285fa1ef7bba..7b4aca414185d 100644 --- a/Mathlib/Analysis/NormedSpace/AffineIsometry.lean +++ b/Mathlib/Analysis/NormedSpace/AffineIsometry.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ import Mathlib.Algebra.CharP.Invertible -import Mathlib.Analysis.NormedSpace.LinearIsometry +import Mathlib.Analysis.Normed.Operator.LinearIsometry import Mathlib.Analysis.Normed.Group.AddTorsor -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.LinearAlgebra.AffineSpace.Restrict import Mathlib.Tactic.FailIfNoProgress diff --git a/Mathlib/Analysis/NormedSpace/BallAction.lean b/Mathlib/Analysis/NormedSpace/BallAction.lean index 016b1d5949548..83b77c0ba1c7f 100644 --- a/Mathlib/Analysis/NormedSpace/BallAction.lean +++ b/Mathlib/Analysis/NormedSpace/BallAction.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Heather Macbeth -/ import Mathlib.Analysis.Normed.Field.UnitBall -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # Multiplicative actions of/on balls and spheres diff --git a/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean b/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean index 886b5e8f448a9..d73ac04e060f8 100644 --- a/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean +++ b/Mathlib/Analysis/NormedSpace/ConformalLinearMap.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Yourong Zang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang -/ -import Mathlib.Analysis.NormedSpace.Basic -import Mathlib.Analysis.NormedSpace.LinearIsometry +import Mathlib.Analysis.Normed.Module.Basic +import Mathlib.Analysis.Normed.Operator.LinearIsometry /-! # Conformal Linear Maps diff --git a/Mathlib/Analysis/NormedSpace/Connected.lean b/Mathlib/Analysis/NormedSpace/Connected.lean index 5fd608bfa75d6..4137feae466c0 100644 --- a/Mathlib/Analysis/NormedSpace/Connected.lean +++ b/Mathlib/Analysis/NormedSpace/Connected.lean @@ -54,11 +54,11 @@ theorem Set.Countable.isPathConnected_compl_of_one_lt_rank have Ia : c - x = a := by simp only [c, x, smul_add, smul_sub] abel_nf - simp [zsmul_eq_smul_cast ℝ 2] + simp [← Int.cast_smul_eq_nsmul ℝ 2] have Ib : c + x = b := by simp only [c, x, smul_add, smul_sub] abel_nf - simp [zsmul_eq_smul_cast ℝ 2] + simp [← Int.cast_smul_eq_nsmul ℝ 2] have x_ne_zero : x ≠ 0 := by simpa [x] using sub_ne_zero.2 hab.symm obtain ⟨y, hy⟩ : ∃ y, LinearIndependent ℝ ![x, y] := exists_linearIndependent_pair_of_one_lt_rank h x_ne_zero diff --git a/Mathlib/Analysis/NormedSpace/DualNumber.lean b/Mathlib/Analysis/NormedSpace/DualNumber.lean index 13247d8e2d873..4058d986db870 100644 --- a/Mathlib/Analysis/NormedSpace/DualNumber.lean +++ b/Mathlib/Analysis/NormedSpace/DualNumber.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.DualNumber -import Mathlib.Analysis.NormedSpace.TrivSqZeroExt +import Mathlib.Analysis.Normed.Algebra.TrivSqZeroExt /-! # Results on `DualNumber R` related to the norm diff --git a/Mathlib/Analysis/NormedSpace/ENorm.lean b/Mathlib/Analysis/NormedSpace/ENorm.lean index ae70ed5e5c429..fccadf993c631 100644 --- a/Mathlib/Analysis/NormedSpace/ENorm.lean +++ b/Mathlib/Analysis/NormedSpace/ENorm.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # Extended norm diff --git a/Mathlib/Analysis/NormedSpace/Exponential.lean b/Mathlib/Analysis/NormedSpace/Exponential.lean index 2c6be89ca452f..afb211322d7cf 100644 --- a/Mathlib/Analysis/NormedSpace/Exponential.lean +++ b/Mathlib/Analysis/NormedSpace/Exponential.lean @@ -265,7 +265,7 @@ theorem exp_add_of_commute_of_mem_ball [CharZero 𝕂] {x y : 𝔸} (hxy : Commu ext rw [hxy.add_pow' _, Finset.smul_sum] refine tsum_congr fun n => Finset.sum_congr rfl fun kl hkl => ?_ - rw [nsmul_eq_smul_cast 𝕂, smul_smul, smul_mul_smul, ← Finset.mem_antidiagonal.mp hkl, + rw [← Nat.cast_smul_eq_nsmul 𝕂, smul_smul, smul_mul_smul, ← Finset.mem_antidiagonal.mp hkl, Nat.cast_add_choose, Finset.mem_antidiagonal.mp hkl] congr 1 have : (n ! : 𝕂) ≠ 0 := Nat.cast_ne_zero.mpr n.factorial_ne_zero diff --git a/Mathlib/Analysis/NormedSpace/Extr.lean b/Mathlib/Analysis/NormedSpace/Extr.lean index 0f63d9574a0a1..913f2d63cb789 100644 --- a/Mathlib/Analysis/NormedSpace/Extr.lean +++ b/Mathlib/Analysis/NormedSpace/Extr.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Analysis.NormedSpace.Ray +import Mathlib.Analysis.Normed.Module.Ray import Mathlib.Topology.Order.LocalExtr /-! diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean index ee1e8f754ebba..9159e42a85143 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel import Mathlib.Analysis.NormedSpace.HahnBanach.Extension import Mathlib.Analysis.NormedSpace.HahnBanach.Separation import Mathlib.LinearAlgebra.Dual -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps /-! # Spaces with separating dual diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index db2f883ce50c0..1993c366e6d1b 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -6,8 +6,8 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo import Mathlib.Algebra.Algebra.Tower import Mathlib.Analysis.LocallyConvex.WithSeminorms import Mathlib.Topology.Algebra.Module.StrongTopology -import Mathlib.Analysis.NormedSpace.LinearIsometry -import Mathlib.Analysis.NormedSpace.ContinuousLinearMap +import Mathlib.Analysis.Normed.Operator.LinearIsometry +import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Tactic.SuppressCompilation /-! diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 50f49aa06004e..d76f44dd9ab98 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo -/ import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic -import Mathlib.Analysis.NormedSpace.LinearIsometry -import Mathlib.Analysis.NormedSpace.ContinuousLinearMap +import Mathlib.Analysis.Normed.Operator.LinearIsometry +import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap /-! # Operator norm: bilinear maps diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 87d682406a3a7..ba20e4a471025 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -167,7 +167,7 @@ theorem isClosed_image_coe_closedBall (f₀ : E →SL[σ₁₂] F) (r : ℝ) : /-- **Banach-Alaoglu** theorem. The set of functions `f : E → F` that represent continuous linear maps `f : E →SL[σ₁₂] F` at distance `≤ r` from `f₀ : E →SL[σ₁₂] F` is compact in the topology of pointwise convergence. Other versions of this theorem can be found in -`Analysis.NormedSpace.WeakDual`. -/ +`Analysis.Normed.Module.WeakDual`. -/ theorem isCompact_image_coe_closedBall [ProperSpace F] (f₀ : E →SL[σ₁₂] F) (r : ℝ) : IsCompact (((↑) : (E →SL[σ₁₂] F) → E → F) '' closedBall f₀ r) := isCompact_image_coe_of_bounded_of_weak_closed isBounded_closedBall <| diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean index 5a9e29a450b49..664473cd62040 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean @@ -5,7 +5,7 @@ Authors: Jan-David Salchow, Sébastien Gouëzel, Jean Lo -/ import Mathlib.Analysis.NormedSpace.OperatorNorm.Bilinear import Mathlib.Analysis.NormedSpace.OperatorNorm.NNNorm -import Mathlib.Analysis.NormedSpace.Span +import Mathlib.Analysis.Normed.Module.Span /-! # Operator norm for maps on normed spaces diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean index a28383e303833..c0f66c299a7b6 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean @@ -146,8 +146,8 @@ theorem injectiveSeminorm_apply (x : ⨂[𝕜] i, E i) : (_ : SeminormedAddCommGroup G) (_ : NormedSpace 𝕜 G), p = Seminorm.comp (normSeminorm 𝕜 (ContinuousMultilinearMap 𝕜 E G →L[𝕜] G)) (toDualContinuousMultilinearMap G (𝕜 := 𝕜) (E := E))}, p.1 x := by - simp [injectiveSeminorm] - exact Seminorm.sSup_apply dualSeminorms_bounded + simpa only [injectiveSeminorm, Set.coe_setOf, Set.mem_setOf_eq] + using Seminorm.sSup_apply dualSeminorms_bounded theorem norm_eval_le_injectiveSeminorm (f : ContinuousMultilinearMap 𝕜 E F) (x : ⨂[𝕜] i, E i) : ‖lift f.toMultilinearMap x‖ ≤ ‖f‖ * injectiveSeminorm x := by diff --git a/Mathlib/Analysis/NormedSpace/Real.lean b/Mathlib/Analysis/NormedSpace/Real.lean index 49fd28f29ff1d..a71f4ae0f993f 100644 --- a/Mathlib/Analysis/NormedSpace/Real.lean +++ b/Mathlib/Analysis/NormedSpace/Real.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Patrick Massot, Eric Wieser, Yaël Dillies -/ -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Topology.Algebra.Module.Basic /-! diff --git a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean index 68cfc9d622f5b..d40f285042ae8 100644 --- a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean +++ b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean @@ -3,7 +3,7 @@ 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.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # Homeomorphism between a normed space and sphere times `(0, +∞)` diff --git a/Mathlib/Analysis/PSeriesComplex.lean b/Mathlib/Analysis/PSeriesComplex.lean index 1a19ce1828726..ce01efdf979f6 100644 --- a/Mathlib/Analysis/PSeriesComplex.lean +++ b/Mathlib/Analysis/PSeriesComplex.lean @@ -5,7 +5,7 @@ Authors: David Loeffler -/ import Mathlib.Analysis.PSeries -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.Data.Complex.FiniteDimensional /-! diff --git a/Mathlib/Analysis/Quaternion.lean b/Mathlib/Analysis/Quaternion.lean index 14a079dfa5a35..e41480ad1e545 100644 --- a/Mathlib/Analysis/Quaternion.lean +++ b/Mathlib/Analysis/Quaternion.lean @@ -86,7 +86,6 @@ noncomputable instance : NormedDivisionRing ℍ where simp only [norm_eq_sqrt_real_inner, inner_self, normSq.map_mul] exact Real.sqrt_mul normSq_nonneg _ --- Porting note: added `noncomputable` noncomputable instance : NormedAlgebra ℝ ℍ where norm_smul_le := norm_smul_le toAlgebra := Quaternion.algebra diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 590110043a1c7..d7d6e7c20ec0f 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Star.Order -import Mathlib.Analysis.NormedSpace.Star.Basic -import Mathlib.Analysis.NormedSpace.ContinuousLinearMap -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Data.Real.Sqrt import Mathlib.Algebra.Algebra.Field @@ -581,7 +581,7 @@ theorem norm_ofNat (n : ℕ) [n.AtLeastTwo] : ‖(no_index (OfNat.ofNat n) : K) variable (K) in lemma norm_nsmul [NormedAddCommGroup E] [NormedSpace K E] (n : ℕ) (x : E) : ‖n • x‖ = n • ‖x‖ := by - rw [nsmul_eq_smul_cast K, norm_smul, RCLike.norm_natCast, nsmul_eq_mul] + rw [← Nat.cast_smul_eq_nsmul K, norm_smul, RCLike.norm_natCast, nsmul_eq_mul] theorem mul_self_norm (z : K) : ‖z‖ * ‖z‖ = normSq z := by rw [normSq_eq_def', sq] diff --git a/Mathlib/Analysis/RCLike/Lemmas.lean b/Mathlib/Analysis/RCLike/Lemmas.lean index bfe5a76d18a35..9c307a9be695f 100644 --- a/Mathlib/Analysis/RCLike/Lemmas.lean +++ b/Mathlib/Analysis/RCLike/Lemmas.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 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.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.Analysis.RCLike.Basic /-! # Further lemmas about `RCLike` -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 8bf20ddd8c8ab..567ae939e14d0 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -561,16 +561,12 @@ theorem tendsto_arg_nhdsWithin_im_neg_of_re_neg_of_im_zero {z : ℂ} (hre : z.re (𝓝[{ z : ℂ | z.im < 0 }] z) (𝓝 (-π)) by refine H.congr' ?_ have : ∀ᶠ x : ℂ in 𝓝 z, x.re < 0 := continuous_re.tendsto z (gt_mem_nhds hre) - -- Porting note: need to specify the `nhdsWithin` set - filter_upwards [self_mem_nhdsWithin (s := { z : ℂ | z.im < 0 }), - mem_nhdsWithin_of_mem_nhds this] with _ him hre + filter_upwards [self_mem_nhdsWithin, mem_nhdsWithin_of_mem_nhds this] with _ him hre rw [arg, if_neg hre.not_le, if_neg him.not_le] convert (Real.continuousAt_arcsin.comp_continuousWithinAt - ((continuous_im.continuousAt.comp_continuousWithinAt continuousWithinAt_neg).div - -- Porting note: added type hint to assist in goal state below - continuous_abs.continuousWithinAt (s := { z : ℂ | z.im < 0 }) (_ : abs z ≠ 0)) - -- Porting note: specify constant precisely to assist in goal below - ).sub_const π using 1 + ((continuous_im.continuousAt.comp_continuousWithinAt continuousWithinAt_neg).div + continuous_abs.continuousWithinAt _) + ).sub_const π using 1 · simp [him] · lift z to ℝ using him simpa using hre.ne diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index b461a5b57d887..46d637bae1dbd 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ -import Mathlib.Analysis.NormedSpace.Spectrum +import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.SpecialFunctions.Exponential -import Mathlib.Topology.ContinuousFunction.FunctionalCalculus +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital /-! # The exponential and logarithm based on the continuous functional calculus @@ -157,7 +157,7 @@ lemma log_pow (n : ℕ) (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) have ha₂' : ContinuousOn Real.log (spectrum ℝ a) := by fun_prop (disch := assumption) have ha₂'' : ContinuousOn Real.log ((· ^ n) '' spectrum ℝ a) := by fun_prop (disch := aesop) rw [log, ← cfc_pow_id (R := ℝ) a n ha₁, ← cfc_comp' Real.log (· ^ n) a ha₂'', log] - simp_rw [Real.log_pow, nsmul_eq_smul_cast ℝ n, cfc_const_mul (n : ℝ) Real.log a ha₂'] + simp_rw [Real.log_pow, ← Nat.cast_smul_eq_nsmul ℝ n, cfc_const_mul (n : ℝ) Real.log a ha₂'] end real_log end CFC diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 2d18fb2622b49..d5dc0f84e7bc8 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -342,11 +342,11 @@ theorem tendsto_exp_comp_nhds_zero {f : α → ℝ} : theorem openEmbedding_exp : OpenEmbedding exp := isOpen_Ioi.openEmbedding_subtype_val.comp expOrderIso.toHomeomorph.openEmbedding --- Porting note (#11215): TODO: backport & make `@[simp]` +@[simp] theorem map_exp_nhds (x : ℝ) : map exp (𝓝 x) = 𝓝 (exp x) := openEmbedding_exp.map_nhds_eq x --- Porting note (#11215): TODO: backport & make `@[simp]` +@[simp] theorem comap_exp_nhds_exp (x : ℝ) : comap exp (𝓝 (exp x)) = 𝓝 x := (openEmbedding_exp.nhds_eq_comap x).symm diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean index 9d67ea306d5c3..69dbed7be166e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean @@ -42,50 +42,6 @@ open Filter Set MeasureTheory open scoped Nat ENNReal Topology Real -section Convexity - --- Porting note: move the following lemmas to `Analysis.Convex.Function` -variable {𝕜 E β : Type*} {s : Set E} {f g : E → β} [OrderedSemiring 𝕜] [SMul 𝕜 E] [AddCommMonoid E] - [OrderedAddCommMonoid β] - -theorem ConvexOn.congr [SMul 𝕜 β] (hf : ConvexOn 𝕜 s f) (hfg : EqOn f g s) : ConvexOn 𝕜 s g := - ⟨hf.1, fun x hx y hy a b ha hb hab => by - simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha hb hab)] using hf.2 hx hy ha hb hab⟩ - -theorem ConcaveOn.congr [SMul 𝕜 β] (hf : ConcaveOn 𝕜 s f) (hfg : EqOn f g s) : ConcaveOn 𝕜 s g := - ⟨hf.1, fun x hx y hy a b ha hb hab => by - simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha hb hab)] using hf.2 hx hy ha hb hab⟩ - -theorem StrictConvexOn.congr [SMul 𝕜 β] (hf : StrictConvexOn 𝕜 s f) (hfg : EqOn f g s) : - StrictConvexOn 𝕜 s g := - ⟨hf.1, fun x hx y hy hxy a b ha hb hab => by - simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha.le hb.le hab)] using - hf.2 hx hy hxy ha hb hab⟩ - -theorem StrictConcaveOn.congr [SMul 𝕜 β] (hf : StrictConcaveOn 𝕜 s f) (hfg : EqOn f g s) : - StrictConcaveOn 𝕜 s g := - ⟨hf.1, fun x hx y hy hxy a b ha hb hab => by - simpa only [← hfg hx, ← hfg hy, ← hfg (hf.1 hx hy ha.le hb.le hab)] using - hf.2 hx hy hxy ha hb hab⟩ - -theorem ConvexOn.add_const [Module 𝕜 β] (hf : ConvexOn 𝕜 s f) (b : β) : - ConvexOn 𝕜 s (f + fun _ => b) := - hf.add (convexOn_const _ hf.1) - -theorem ConcaveOn.add_const [Module 𝕜 β] (hf : ConcaveOn 𝕜 s f) (b : β) : - ConcaveOn 𝕜 s (f + fun _ => b) := - hf.add (concaveOn_const _ hf.1) - -theorem StrictConvexOn.add_const {γ : Type*} {f : E → γ} [OrderedCancelAddCommMonoid γ] - [Module 𝕜 γ] (hf : StrictConvexOn 𝕜 s f) (b : γ) : StrictConvexOn 𝕜 s (f + fun _ => b) := - hf.add_convexOn (convexOn_const _ hf.1) - -theorem StrictConcaveOn.add_const {γ : Type*} {f : E → γ} [OrderedCancelAddCommMonoid γ] - [Module 𝕜 γ] (hf : StrictConcaveOn 𝕜 s f) (b : γ) : StrictConcaveOn 𝕜 s (f + fun _ => b) := - hf.add_concaveOn (concaveOn_const _ hf.1) - -end Convexity - namespace Real section Convexity diff --git a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean index 6b7e37b0d6c12..6bc1053f456c6 100644 --- a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean +++ b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean @@ -61,9 +61,6 @@ theorem nonneg (x : ℝ) : 0 ≤ expNegInvGlue x := by /-! ### Smoothness of `expNegInvGlue` -Porting note: Yury Kudryashov rewrote the proof while porting, generalizing auxiliary lemmas and -removing some auxiliary definitions. - In this section we prove that the function `f = expNegInvGlue` is infinitely smooth. To do this, we show that $g_p(x)=p(x^{-1})f(x)$ is infinitely smooth for any polynomial `p` with real coefficients. First we show that $g_p(x)$ tends to zero at zero, then we show that it is diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 153d3e4b0ac3d..1d5205947fa4d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -63,7 +63,6 @@ theorem log_stirlingSeq_formula (n : ℕ) : · simp · rw [stirlingSeq, log_div, log_mul, sqrt_eq_rpow, log_rpow, Real.log_pow, tsub_tsub] <;> positivity --- Porting note: generalized from `n.succ` to `n` /-- The sequence `log (stirlingSeq (m + 1)) - log (stirlingSeq (m + 2))` has the series expansion `∑ 1 / (2 * (k + 1) + 1) * (1 / 2 * (m + 1) + 1)^(2 * (k + 1))` diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index d0bd57f707bb2..6dd45e6fe5610 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -156,7 +156,6 @@ theorem nsmul_eq_iff {ψ θ : Angle} {n : ℕ} (hz : n ≠ 0) : QuotientAddGroup.zmultiples_nsmul_eq_nsmul_iff hz theorem two_zsmul_eq_iff {ψ θ : Angle} : (2 : ℤ) • ψ = (2 : ℤ) • θ ↔ ψ = θ ∨ ψ = θ + ↑π := by - -- Porting note: no `Int.natAbs_bit0` anymore have : Int.natAbs 2 = 2 := rfl rw [zsmul_eq_iff two_ne_zero, this, Fin.exists_fin_two, Fin.val_zero, Fin.val_one, zero_smul, add_zero, one_smul, Int.cast_two, diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean index 6ec1f49816bfc..7b8d099c1e98a 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean @@ -177,8 +177,7 @@ theorem integral_cos_pow_eq (n : ℕ) : have L : IntervalIntegrable _ volume 0 (π / 2) := (continuous_sin.pow n).intervalIntegrable _ _ have R : IntervalIntegrable _ volume (π / 2) π := (continuous_sin.pow n).intervalIntegrable _ _ rw [← integral_add_adjacent_intervals L R] - -- Porting note: was `congr 1` but it timeouts - refine congr_arg₂ _ ?_ ?_ + congr 1 · nth_rw 1 [(by ring : 0 = π / 2 - π / 2)] nth_rw 3 [(by ring : π / 2 = π / 2 - 0)] rw [← integral_comp_sub_left] diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index fd9da001e1006..22e2390134e42 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -9,7 +9,7 @@ import Mathlib.Order.Filter.ModEq import Mathlib.Analysis.Asymptotics.Asymptotics import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Data.List.TFAE -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # A collection of specific limit computations diff --git a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean index 1c74c810a035f..28d385c769837 100644 --- a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean +++ b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.InnerProductSpace.Adjoint diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index d9406ee36551b..b5bcc32771f73 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -460,17 +460,13 @@ section ConstructLeft -- constructed from this data. variable {F_obj : C → D} variable (e : ∀ X Y, (F_obj X ⟶ Y) ≃ (X ⟶ G.obj Y)) -variable (he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g) - -private theorem he' {X Y Y'} (f g) : (e X Y').symm (f ≫ G.map g) = (e X Y).symm f ≫ g := by - rw [Equiv.symm_apply_eq, he]; simp /-- Construct a left adjoint functor to `G`, given the functor's value on objects `F_obj` and a bijection `e` between `F_obj X ⟶ Y` and `X ⟶ G.obj Y` satisfying a naturality law `he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g`. Dual to `rightAdjointOfEquiv`. -/ @[simps!] -def leftAdjointOfEquiv : C ⥤ D where +def leftAdjointOfEquiv (he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g) : C ⥤ D where obj := F_obj map {X} {X'} f := (e X (F_obj X')).symm (f ≫ e X' (F_obj X') (𝟙 _)) map_comp := fun f f' => by @@ -480,6 +476,8 @@ def leftAdjointOfEquiv : C ⥤ D where rw [assoc, ← he, id_comp, Equiv.apply_symm_apply] simp +variable (he : ∀ X Y Y' g h, e X Y' (h ≫ g) = e X Y h ≫ G.map g) + /-- Show that the functor given by `leftAdjointOfEquiv` is indeed left adjoint to `G`. Dual to `adjunctionOfRightEquiv`. -/ @[simps!] @@ -487,7 +485,9 @@ def adjunctionOfEquivLeft : leftAdjointOfEquiv e he ⊣ G := mkOfHomEquiv { homEquiv := e homEquiv_naturality_left_symm := fun {X'} {X} {Y} f g => by - have := @he' C _ D _ G F_obj e he + have {X : C} {Y Y' : D} (f : X ⟶ G.obj Y) (g : Y ⟶ Y') : + (e X Y').symm (f ≫ G.map g) = (e X Y).symm f ≫ g := by + rw [Equiv.symm_apply_eq, he]; simp erw [← this, ← Equiv.apply_eq_iff_eq (e X' Y)] simp only [leftAdjointOfEquiv_obj, Equiv.apply_symm_apply, assoc] congr @@ -502,9 +502,9 @@ section ConstructRight -- Construction of a right adjoint, analogous to the above. variable {G_obj : D → C} variable (e : ∀ X Y, (F.obj X ⟶ Y) ≃ (X ⟶ G_obj Y)) -variable (he : ∀ X' X Y f g, e X' Y (F.map f ≫ g) = f ≫ e X Y g) -private theorem he'' {X' X Y} (f g) : F.map f ≫ (e X Y).symm g = (e X' Y).symm (f ≫ g) := by +private theorem he'' (he : ∀ X' X Y f g, e X' Y (F.map f ≫ g) = f ≫ e X Y g) + {X' X Y} (f g) : F.map f ≫ (e X Y).symm g = (e X' Y).symm (f ≫ g) := by rw [Equiv.eq_symm_apply, he]; simp /-- Construct a right adjoint functor to `F`, given the functor's value on objects `G_obj` and @@ -512,7 +512,7 @@ a bijection `e` between `F.obj X ⟶ Y` and `X ⟶ G_obj Y` satisfying a natural `he : ∀ X Y Y' g h, e X' Y (F.map f ≫ g) = f ≫ e X Y g`. Dual to `leftAdjointOfEquiv`. -/ @[simps!] -def rightAdjointOfEquiv : D ⥤ C where +def rightAdjointOfEquiv (he : ∀ X' X Y f g, e X' Y (F.map f ≫ g) = f ≫ e X Y g) : D ⥤ C where obj := G_obj map {Y} {Y'} g := (e (G_obj Y) Y') ((e (G_obj Y) Y).symm (𝟙 _) ≫ g) map_comp := fun {Y} {Y'} {Y''} g g' => by @@ -525,7 +525,8 @@ def rightAdjointOfEquiv : D ⥤ C where /-- Show that the functor given by `rightAdjointOfEquiv` is indeed right adjoint to `F`. Dual to `adjunctionOfEquivRight`. -/ @[simps!] -def adjunctionOfEquivRight : F ⊣ (rightAdjointOfEquiv e he) := +def adjunctionOfEquivRight (he : ∀ X' X Y f g, e X' Y (F.map f ≫ g) = f ≫ e X Y g) : + F ⊣ (rightAdjointOfEquiv e he) := mkOfHomEquiv { homEquiv := e homEquiv_naturality_left_symm := by diff --git a/Mathlib/CategoryTheory/Bicategory/Adjunction.lean b/Mathlib/CategoryTheory/Bicategory/Adjunction.lean index c17f034c110d2..ae25dba05c39c 100644 --- a/Mathlib/CategoryTheory/Bicategory/Adjunction.lean +++ b/Mathlib/CategoryTheory/Bicategory/Adjunction.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ import Mathlib.Tactic.CategoryTheory.Coherence -import Mathlib.CategoryTheory.Bicategory.Coherence /-! # Adjunctions in bicategories diff --git a/Mathlib/CategoryTheory/Bicategory/Basic.lean b/Mathlib/CategoryTheory/Bicategory/Basic.lean index a63052560c68c..7a8cf12061d0f 100644 --- a/Mathlib/CategoryTheory/Bicategory/Basic.lean +++ b/Mathlib/CategoryTheory/Bicategory/Basic.lean @@ -430,7 +430,7 @@ category of functors `(b ⟶ c) ⥤ (a ⟶ c)`. -/ @[simps] def precomposing (a b c : B) : (a ⟶ b) ⥤ (b ⟶ c) ⥤ (a ⟶ c) where obj f := precomp c f - map η := ⟨(η ▷ ·), _⟩ + map η := { app := (η ▷ ·) } /-- Postcomposition of a 1-morphism as a functor. -/ @[simps] @@ -443,7 +443,7 @@ category of functors `(a ⟶ b) ⥤ (a ⟶ c)`. -/ @[simps] def postcomposing (a b c : B) : (b ⟶ c) ⥤ (a ⟶ b) ⥤ (a ⟶ c) where obj f := postcomp a f - map η := ⟨(· ◁ η), _⟩ + map η := { app := (· ◁ η) } /-- Left component of the associator as a natural isomorphism. -/ @[simps!] diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean b/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean index 4930763c58e63..bc1c66a21dae8 100644 --- a/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean +++ b/Mathlib/CategoryTheory/Bicategory/Functor/Oplax.lean @@ -110,6 +110,21 @@ attribute [nolint docBlame] CategoryTheory.OplaxFunctor.mapId variable (F : OplaxFunctor B C) +@[reassoc] +lemma mapComp_assoc_right {a b c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + F.mapComp f (g ≫ h) ≫ F.map f ◁ F.mapComp g h = F.map₂ (α_ f g h).inv ≫ + F.mapComp (f ≫ g) h ≫ F.mapComp f g ▷ F.map h ≫ + (α_ (F.map f) (F.map g) (F.map h)).hom := by + rw [← @map₂_associator, ← F.map₂_comp_assoc] + simp + +@[reassoc] +lemma mapComp_assoc_left {a b c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + F.mapComp (f ≫ g) h ≫ F.mapComp f g ▷ F.map h = + F.map₂ (α_ f g h).hom ≫ F.mapComp f (g ≫ h) ≫ F.map f ◁ F.mapComp g h + ≫ (α_ (F.map f) (F.map g) (F.map h)).inv := by + simp + -- Porting note: `to_prelax_eq_coe` and `to_prelaxFunctor_obj` are -- syntactic tautologies in lean 4 diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/Pseudofunctor.lean b/Mathlib/CategoryTheory/Bicategory/Functor/Pseudofunctor.lean index f075171101cb0..20e57f3061b67 100644 --- a/Mathlib/CategoryTheory/Bicategory/Functor/Pseudofunctor.lean +++ b/Mathlib/CategoryTheory/Bicategory/Functor/Pseudofunctor.lean @@ -115,15 +115,6 @@ attribute [nolint docBlame] CategoryTheory.Pseudofunctor.mapId variable (F : Pseudofunctor B C) --- Porting note: `toPrelaxFunctor_eq_coe` and `to_prelaxFunctor_obj` --- are syntactic tautologies in lean 4 - --- Porting note: removed lemma `to_prelaxFunctor_map` relating the now --- nonexistent `PrelaxFunctor.map` and the now nonexistent `Pseudofunctor.map` - --- Porting note: removed lemma `to_prelaxFunctor_map₂` relating --- `PrelaxFunctor.map₂` to nonexistent `Pseudofunctor.map₂` - /-- The oplax functor associated with a pseudofunctor. -/ @[simps] def toOplax : OplaxFunctor B C where @@ -134,8 +125,6 @@ def toOplax : OplaxFunctor B C where instance hasCoeToOplax : Coe (Pseudofunctor B C) (OplaxFunctor B C) := ⟨toOplax⟩ --- Porting note: `toOplax_eq_coe` is a syntactic tautology in lean 4 - /-- The Lax functor associated with a pseudofunctor. -/ @[simps] def toLax : LaxFunctor B C where @@ -152,7 +141,6 @@ def toLax : LaxFunctor B C where instance hasCoeToLax : Coe (Pseudofunctor B C) (LaxFunctor B C) := ⟨toLax⟩ - /-- The identity pseudofunctor. -/ @[simps] def id (B : Type u₁) [Bicategory.{w₁, v₁} B] : Pseudofunctor B B where @@ -176,6 +164,96 @@ def comp (F : Pseudofunctor B C) (G : Pseudofunctor C D) : Pseudofunctor B D whe map₂_left_unitor f := by dsimp; simp map₂_right_unitor f := by dsimp; simp +section + +variable (F : Pseudofunctor B C) {a b : B} + +@[reassoc] +lemma mapComp_assoc_right_hom {c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + (F.mapComp f (g ≫ h)).hom ≫ F.map f ◁ (F.mapComp g h).hom = F.map₂ (α_ f g h).inv ≫ + (F.mapComp (f ≫ g) h).hom ≫ (F.mapComp f g).hom ▷ F.map h ≫ + (α_ (F.map f) (F.map g) (F.map h)).hom := + F.toOplax.mapComp_assoc_right _ _ _ + +@[reassoc] +lemma mapComp_assoc_left_hom {c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + (F.mapComp (f ≫ g) h).hom ≫ (F.mapComp f g).hom ▷ F.map h = + F.map₂ (α_ f g h).hom ≫ (F.mapComp f (g ≫ h)).hom ≫ F.map f ◁ (F.mapComp g h).hom + ≫ (α_ (F.map f) (F.map g) (F.map h)).inv := + F.toOplax.mapComp_assoc_left _ _ _ + +@[reassoc] +lemma mapComp_assoc_right_inv {c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + F.map f ◁ (F.mapComp g h).inv ≫ (F.mapComp f (g ≫ h)).inv = + (α_ (F.map f) (F.map g) (F.map h)).inv ≫ (F.mapComp f g).inv ▷ F.map h ≫ + (F.mapComp (f ≫ g) h).inv ≫ F.map₂ (α_ f g h).hom := + F.toLax.mapComp_assoc_right _ _ _ + +@[reassoc] +lemma mapComp_assoc_left_inv {c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + (F.mapComp f g).inv ▷ F.map h ≫ (F.mapComp (f ≫ g) h).inv = + (α_ (F.map f) (F.map g) (F.map h)).hom ≫ F.map f ◁ (F.mapComp g h).inv ≫ + (F.mapComp f (g ≫ h)).inv ≫ F.map₂ (α_ f g h).inv := + F.toLax.mapComp_assoc_left _ _ _ + +@[reassoc] +lemma mapComp_id_left_hom (f : a ⟶ b) : (F.mapComp (𝟙 a) f).hom = + F.map₂ (λ_ f).hom ≫ (λ_ (F.map f)).inv ≫ (F.mapId a).inv ▷ F.map f := by + simp + +lemma mapComp_id_left (f : a ⟶ b) : (F.mapComp (𝟙 a) f) = F.map₂Iso (λ_ f) ≪≫ + (λ_ (F.map f)).symm ≪≫ (whiskerRightIso (F.mapId a) (F.map f)).symm := + Iso.ext <| F.mapComp_id_left_hom f + +@[reassoc] +lemma mapComp_id_left_inv (f : a ⟶ b) : (F.mapComp (𝟙 a) f).inv = + (F.mapId a).hom ▷ F.map f ≫ (λ_ (F.map f)).hom ≫ F.map₂ (λ_ f).inv := by + simp [mapComp_id_left] + +lemma whiskerRightIso_mapId (f : a ⟶ b) : whiskerRightIso (F.mapId a) (F.map f) = + (F.mapComp (𝟙 a) f).symm ≪≫ F.map₂Iso (λ_ f) ≪≫ (λ_ (F.map f)).symm := by + simp [mapComp_id_left] + +@[reassoc] +lemma whiskerRight_mapId_hom (f : a ⟶ b) : (F.mapId a).hom ▷ F.map f = + (F.mapComp (𝟙 a) f).inv ≫ F.map₂ (λ_ f).hom ≫ (λ_ (F.map f)).inv := by + simp [whiskerRightIso_mapId] + +@[reassoc] +lemma whiskerRight_mapId_inv (f : a ⟶ b) : (F.mapId a).inv ▷ F.map f = + (λ_ (F.map f)).hom ≫ F.map₂ (λ_ f).inv ≫ (F.mapComp (𝟙 a) f).hom := by + simpa using congrArg (·.inv) (F.whiskerRightIso_mapId f) + +@[reassoc] +lemma mapComp_id_right_hom (f : a ⟶ b) : (F.mapComp f (𝟙 b)).hom = + F.map₂ (ρ_ f).hom ≫ (ρ_ (F.map f)).inv ≫ F.map f ◁ (F.mapId b).inv := by + simp + +lemma mapComp_id_right (f : a ⟶ b) : (F.mapComp f (𝟙 b)) = F.map₂Iso (ρ_ f) ≪≫ + (ρ_ (F.map f)).symm ≪≫ (whiskerLeftIso (F.map f) (F.mapId b)).symm := + Iso.ext <| F.mapComp_id_right_hom f + +@[reassoc] +lemma mapComp_id_right_inv (f : a ⟶ b) : (F.mapComp f (𝟙 b)).inv = + F.map f ◁ (F.mapId b).hom ≫ (ρ_ (F.map f)).hom ≫ F.map₂ (ρ_ f).inv := by + simp [mapComp_id_right] + +lemma whiskerLeftIso_mapId (f : a ⟶ b) : whiskerLeftIso (F.map f) (F.mapId b) = + (F.mapComp f (𝟙 b)).symm ≪≫ F.map₂Iso (ρ_ f) ≪≫ (ρ_ (F.map f)).symm := by + simp [mapComp_id_right] + +@[reassoc] +lemma whiskerLeft_mapId_hom (f : a ⟶ b) : F.map f ◁ (F.mapId b).hom = + (F.mapComp f (𝟙 b)).inv ≫ F.map₂ (ρ_ f).hom ≫ (ρ_ (F.map f)).inv := by + simp [whiskerLeftIso_mapId] + +@[reassoc] +lemma whiskerLeft_mapId_inv (f : a ⟶ b) : F.map f ◁ (F.mapId b).inv = + (ρ_ (F.map f)).hom ≫ F.map₂ (ρ_ f).inv ≫ (F.mapComp f (𝟙 b)).hom := by + simpa using congrArg (·.inv) (F.whiskerLeftIso_mapId f) + +end + /-- Construct a pseudofunctor from an oplax functor whose `mapId` and `mapComp` are isomorphisms. -/ @[simps] def mkOfOplax (F : OplaxFunctor B C) (F' : F.PseudoCore) : Pseudofunctor B C where diff --git a/Mathlib/CategoryTheory/Category/GaloisConnection.lean b/Mathlib/CategoryTheory/Category/GaloisConnection.lean index ab8ddf81cdbca..41362d5292353 100644 --- a/Mathlib/CategoryTheory/Category/GaloisConnection.lean +++ b/Mathlib/CategoryTheory/Category/GaloisConnection.lean @@ -28,8 +28,10 @@ def GaloisConnection.adjunction {l : X → Y} {u : Y → X} (gc : GaloisConnecti gc.monotone_l.functor ⊣ gc.monotone_u.functor := CategoryTheory.Adjunction.mkOfHomEquiv { homEquiv := fun X Y => - ⟨fun f => CategoryTheory.homOfLE (gc.le_u f.le), - fun f => CategoryTheory.homOfLE (gc.l_le f.le), _, _⟩ } + { toFun := fun f => CategoryTheory.homOfLE (gc.le_u f.le) + invFun := fun f => CategoryTheory.homOfLE (gc.l_le f.le) + left_inv := by aesop_cat + right_inv := by aesop_cat } } end diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index 90ddb8166baae..8acbee9382883 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -481,14 +481,13 @@ abbrev δlast (F : ComposableArrows C (n + 1)) := δlastFunctor.obj F section variable {F G : ComposableArrows C (n + 1)} - (α : F.obj' 0 ⟶ G.obj' 0) - (β : F.δ₀ ⟶ G.δ₀) - (w : F.map' 0 1 ≫ app' β 0 = α ≫ G.map' 0 1) + /-- Inductive construction of morphisms in `ComposableArrows C (n + 1)`: in order to construct a morphism `F ⟶ G`, it suffices to provide `α : F.obj' 0 ⟶ G.obj' 0` and `β : F.δ₀ ⟶ G.δ₀` such that `F.map' 0 1 ≫ app' β 0 = α ≫ G.map' 0 1`. -/ -def homMkSucc : F ⟶ G := +def homMkSucc (α : F.obj' 0 ⟶ G.obj' 0) (β : F.δ₀ ⟶ G.δ₀) + (w : F.map' 0 1 ≫ app' β 0 = α ≫ G.map' 0 1) : F ⟶ G := homMk (fun i => match i with | ⟨0, _⟩ => α @@ -498,6 +497,9 @@ def homMkSucc : F ⟶ G := · exact w · exact naturality' β i (i + 1)) +variable (α : F.obj' 0 ⟶ G.obj' 0) (β : F.δ₀ ⟶ G.δ₀) + (w : F.map' 0 1 ≫ app' β 0 = α ≫ G.map' 0 1) + @[simp] lemma homMkSucc_app_zero : (homMkSucc α β w).app 0 = α := rfl diff --git a/Mathlib/CategoryTheory/Core.lean b/Mathlib/CategoryTheory/Core.lean index b5e483ee4183b..e124cbff9a5bc 100644 --- a/Mathlib/CategoryTheory/Core.lean +++ b/Mathlib/CategoryTheory/Core.lean @@ -72,7 +72,7 @@ variable {C} {G : Type u₂} [Groupoid.{v₂} G] /-- A functor from a groupoid to a category C factors through the core of C. -/ def functorToCore (F : G ⥤ C) : G ⥤ Core C where obj X := F.obj X - map f := ⟨F.map f, F.map (Groupoid.inv f), _, _⟩ + map f := { hom := F.map f, inv := F.map (Groupoid.inv f) } /-- We can functorially associate to any functor from a groupoid to the core of a category `C`, a functor from the groupoid to `C`, simply by composing with the embedding `Core C ⥤ C`. diff --git a/Mathlib/CategoryTheory/Endomorphism.lean b/Mathlib/CategoryTheory/Endomorphism.lean index 3c7310719d892..f1a262d510296 100644 --- a/Mathlib/CategoryTheory/Endomorphism.lean +++ b/Mathlib/CategoryTheory/Endomorphism.lean @@ -150,8 +150,8 @@ def toEnd (X : C) : Aut X →* End X := (Units.coeHom (End X)).comp (Aut.unitsEn /-- Isomorphisms induce isomorphisms of the automorphism group -/ def autMulEquivOfIso {X Y : C} (h : X ≅ Y) : Aut X ≃* Aut Y where - toFun x := ⟨h.inv ≫ x.hom ≫ h.hom, h.inv ≫ x.inv ≫ h.hom, _, _⟩ - invFun y := ⟨h.hom ≫ y.hom ≫ h.inv, h.hom ≫ y.inv ≫ h.inv, _, _⟩ + toFun x := { hom := h.inv ≫ x.hom ≫ h.hom, inv := h.inv ≫ x.inv ≫ h.hom } + invFun y := { hom := h.hom ≫ y.hom ≫ h.inv, inv := h.hom ≫ y.inv ≫ h.inv } left_inv _ := by aesop_cat right_inv _ := by aesop_cat map_mul' := by simp [Aut_mul_def] diff --git a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean index f3fab7017415d..d8ac3d794f3a7 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/HomLift.lean @@ -71,14 +71,16 @@ protected lemma id {p : 𝒳 ⥤ 𝒮} {R : 𝒮} {a : 𝒳} (ha : p.obj a = R) section -variable {R S : 𝒮} {a b : 𝒳} (f : R ⟶ S) (φ : a ⟶ b) [p.IsHomLift f φ] +variable {R S : 𝒮} {a b : 𝒳} -lemma domain_eq : p.obj a = R := by +lemma domain_eq (f : R ⟶ S) (φ : a ⟶ b) [p.IsHomLift f φ] : p.obj a = R := by subst_hom_lift p f φ; rfl -lemma codomain_eq : p.obj b = S := by +lemma codomain_eq (f : R ⟶ S) (φ : a ⟶ b) [p.IsHomLift f φ] : p.obj b = S := by subst_hom_lift p f φ; rfl +variable (f : R ⟶ S) (φ : a ⟶ b) [p.IsHomLift f φ] + lemma fac : f = eqToHom (domain_eq p f φ).symm ≫ p.map φ ≫ eqToHom (codomain_eq p f φ) := by subst_hom_lift p f φ; simp diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index 4b20f6d867cf7..f4da1d923d0e6 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -186,14 +186,15 @@ variable {C} variable [IsFilteredOrEmpty C] variable {D : Type u₁} [Category.{v₁} D] -/-- If `C` is filtered or emtpy, and we have a functor `R : C ⥤ D` with a left adjoint, then `D` is +/-- If `C` is filtered or empty, and we have a functor `R : C ⥤ D` with a left adjoint, then `D` is filtered or empty. -/ theorem of_right_adjoint {L : D ⥤ C} {R : C ⥤ D} (h : L ⊣ R) : IsFilteredOrEmpty D := { cocone_objs := fun X Y => - ⟨_, h.homEquiv _ _ (leftToMax _ _), h.homEquiv _ _ (rightToMax _ _), ⟨⟩⟩ + ⟨R.obj (max (L.obj X) (L.obj Y)), + h.homEquiv _ _ (leftToMax _ _), h.homEquiv _ _ (rightToMax _ _), ⟨⟩⟩ cocone_maps := fun X Y f g => - ⟨_, h.homEquiv _ _ (coeqHom _ _), by + ⟨R.obj (coeq (L.map f) (L.map g)), h.homEquiv _ _ (coeqHom _ _), by rw [← h.homEquiv_naturality_left, ← h.homEquiv_naturality_left, coeq_condition]⟩ } /-- If `C` is filtered or empty, and we have a right adjoint functor `R : C ⥤ D`, then `D` is diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index 7e26775dbde61..8845826dfdb7a 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -32,7 +32,7 @@ universe v₁ v₂ v₃ u₁ u₂ u₃ namespace CategoryTheory -variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] {E : Type*} [Category E] namespace Functor @@ -79,7 +79,10 @@ theorem map_preimage (F : C ⥤ D) [Full F] {X Y : C} (f : F.obj X ⟶ F.obj Y) F.map (preimage F f) = f := (F.map_surjective f).choose_spec -variable {F : C ⥤ D} [Full F] [F.Faithful] {X Y Z : C} +variable {F : C ⥤ D} {X Y Z : C} + +section +variable [Full F] [F.Faithful] @[simp] theorem preimage_id : F.preimage (𝟙 (F.obj X)) = 𝟙 X := @@ -110,6 +113,9 @@ theorem preimageIso_mapIso (f : X ≅ Y) : F.preimageIso (F.mapIso f) = f := by ext simp +end + +variable (F) in /-- Structure containing the data of inverse map `(F.obj X ⟶ F.obj Y) ⟶ (X ⟶ Y)` of `F.map` in order to express that `F` is a fully faithful functor. -/ structure FullyFaithful where @@ -122,13 +128,20 @@ namespace FullyFaithful attribute [simp] map_preimage preimage_map +variable (F) in /-- A `FullyFaithful` structure can be obtained from the assumption the `F` is both full and faithful. -/ noncomputable def ofFullyFaithful [F.Full] [F.Faithful] : F.FullyFaithful where preimage := F.preimage -variable {F} +variable (C) in +/-- The identity functor is fully faithful. -/ +@[simps] +def id : (𝟭 C).FullyFaithful where + preimage f := f + +section variable (hF : F.FullyFaithful) /-- The equivalence `(X ⟶ Y) ≃ (F.obj X ⟶ F.obj Y)` given by `h : F.FullyFaithful`. -/ @@ -165,7 +178,8 @@ def preimageIso {X Y : C} (e : F.obj X ≅ F.obj Y) : X ≅ Y where hom_inv_id := hF.map_injective (by simp) inv_hom_id := hF.map_injective (by simp) -lemma isIso_of_isIso_map {X Y : C} (f : X ⟶ Y) [IsIso (F.map f)] : IsIso f := by +lemma isIso_of_isIso_map (hF : F.FullyFaithful) {X Y : C} (f : X ⟶ Y) [IsIso (F.map f)] : + IsIso f := by simpa using (hF.preimageIso (asIso (F.map f))).isIso_hom /-- The equivalence `(X ≅ Y) ≃ (F.obj X ≅ F.obj Y)` given by `h : F.FullyFaithful`. -/ @@ -176,19 +190,13 @@ def isoEquiv {X Y : C} : (X ≅ Y) ≃ (F.obj X ≅ F.obj Y) where left_inv := by aesop_cat right_inv := by aesop_cat -variable (C) in -/-- The identity functor is fully faithful. -/ -@[simps] -def id : (𝟭 C).FullyFaithful where - preimage f := f - -variable {E : Type*} [Category E] - /-- Fully faithful functors are stable by composition. -/ @[simps] def comp {G : D ⥤ E} (hG : G.FullyFaithful) : (F ⋙ G).FullyFaithful where preimage f := hF.preimage (hG.preimage f) +end + /-- If `F ⋙ G` is fully faithful and `G` is faithful, then `F` is fully faithful. -/ def ofCompFaithful {G : D ⥤ E} [G.Faithful] (hFG : (F ⋙ G).FullyFaithful) : F.FullyFaithful where diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean index 18df915fff2cb..c3bd4a60730ab 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean @@ -136,7 +136,8 @@ def IsPointwiseLeftKanExtension.homFrom (G : LeftExtension L F) : E ⟶ G := ext X simpa using (h (L.obj X)).fac (LeftExtension.coconeAt G _) (CostructuredArrow.mk (𝟙 _))) -lemma IsPointwiseLeftKanExtension.hom_ext {G : LeftExtension L F} {f₁ f₂ : E ⟶ G} : f₁ = f₂ := by +lemma IsPointwiseLeftKanExtension.hom_ext (h : E.IsPointwiseLeftKanExtension) + {G : LeftExtension L F} {f₁ f₂ : E ⟶ G} : f₁ = f₂ := by ext Y apply (h Y).hom_ext intro X @@ -160,9 +161,9 @@ lemma IsPointwiseLeftKanExtension.hasLeftKanExtension : HasLeftKanExtension.mk E.right E.hom lemma IsPointwiseLeftKanExtension.isIso_hom [L.Full] [L.Faithful] : - IsIso (E.hom) := by + IsIso (E.hom) := have := fun X => (h (L.obj X)).isIso_hom_app - apply NatIso.isIso_of_isIso_app + NatIso.isIso_of_isIso_app .. end LeftExtension @@ -247,7 +248,8 @@ def IsPointwiseRightKanExtension.homTo (G : RightExtension L F) : G ⟶ E := ext X simpa using (h (L.obj X)).fac (RightExtension.coneAt G _) (StructuredArrow.mk (𝟙 _)) ) -lemma IsPointwiseRightKanExtension.hom_ext {G : RightExtension L F} {f₁ f₂ : G ⟶ E} : f₁ = f₂ := by +lemma IsPointwiseRightKanExtension.hom_ext (h : E.IsPointwiseRightKanExtension) + {G : RightExtension L F} {f₁ f₂ : G ⟶ E} : f₁ = f₂ := by ext Y apply (h Y).hom_ext intro X @@ -270,9 +272,9 @@ lemma IsPointwiseRightKanExtension.hasRightKanExtension : HasRightKanExtension.mk E.left E.hom lemma IsPointwiseRightKanExtension.isIso_hom [L.Full] [L.Faithful] : - IsIso (E.hom) := by + IsIso (E.hom) := have := fun X => (h (L.obj X)).isIso_hom_app - apply NatIso.isIso_of_isIso_app + NatIso.isIso_of_isIso_app .. end RightExtension diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index 82356706b9525..cafeecc3c84e4 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -310,7 +310,7 @@ for all `j : J`, the coproduct of all `X i` such `p i = j` exists. -/ abbrev HasMap : Prop := ∀ (j : J), HasCoproduct (X.mapObjFun p j) variable {X Y} in -lemma hasMap_of_iso [HasMap X p] : HasMap Y p := fun j => by +lemma hasMap_of_iso (e : X ≅ Y) (p: I → J) [HasMap X p] : HasMap Y p := fun j => by have α : Discrete.functor (X.mapObjFun p j) ≅ Discrete.functor (Y.mapObjFun p j) := Discrete.natIso (fun ⟨i, _⟩ => (GradedObject.eval i).mapIso e) exact hasColimitOfIso α.symm diff --git a/Mathlib/CategoryTheory/GradedObject/Unitor.lean b/Mathlib/CategoryTheory/GradedObject/Unitor.lean index c47e96d33d87f..eece8ef2deb7c 100644 --- a/Mathlib/CategoryTheory/GradedObject/Unitor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Unitor.lean @@ -53,7 +53,7 @@ noncomputable def mapBifunctorObjSingle₀ObjIsInitial (a : I × J) (ha : a.1 `p : I × J → J` such that `p ⟨0, j⟩ = j` for all `j`, this is the (colimit) cofan which shall be used to construct the isomorphism `mapBifunctorMapObj F p ((single₀ I).obj X) Y ≅ Y`, see `mapBifunctorLeftUnitor`. -/ -noncomputable def mapBifunctorLeftUnitorCofan (j : J) : +noncomputable def mapBifunctorLeftUnitorCofan (hp : ∀ (j : J), p ⟨0, j⟩ = j) (Y) (j : J) : (((mapBifunctor F I J).obj ((single₀ I).obj X)).obj Y).CofanMapObjFun p j := CofanMapObjFun.mk _ _ _ (Y j) (fun a ha => if ha : a.1 = 0 then @@ -164,7 +164,7 @@ noncomputable def mapBifunctorObjObjSingle₀IsInitial (a : J × I) (ha : a.2 `p : J × I → J` such that `p ⟨j, 0⟩ = j` for all `j`, this is the (colimit) cofan which shall be used to construct the isomorphism `mapBifunctorMapObj F p X ((single₀ I).obj Y) ≅ X`, see `mapBifunctorRightUnitor`. -/ -noncomputable def mapBifunctorRightUnitorCofan (j : J) : +noncomputable def mapBifunctorRightUnitorCofan (hp : ∀ (j : J), p ⟨j, 0⟩ = j) (X) (j : J) : (((mapBifunctor F J I).obj X).obj ((single₀ I).obj Y)).CofanMapObjFun p j := CofanMapObjFun.mk _ _ _ (X j) (fun a ha => if ha : a.2 = 0 then diff --git a/Mathlib/CategoryTheory/GuitartExact/Basic.lean b/Mathlib/CategoryTheory/GuitartExact/Basic.lean index b8c74acb5596a..c87d10938f0ce 100644 --- a/Mathlib/CategoryTheory/GuitartExact/Basic.lean +++ b/Mathlib/CategoryTheory/GuitartExact/Basic.lean @@ -116,17 +116,19 @@ abbrev CostructuredArrowDownwards := section variable (X₁ : C₁) (a : X₂ ⟶ T.obj X₁) (b : L.obj X₁ ⟶ X₃) - (comm : R.map a ≫ w.app X₁ ≫ B.map b = g) /-- Constructor for objects in `w.StructuredArrowRightwards g`. -/ -abbrev StructuredArrowRightwards.mk : w.StructuredArrowRightwards g := +abbrev StructuredArrowRightwards.mk (comm : R.map a ≫ w.app X₁ ≫ B.map b = g) : + w.StructuredArrowRightwards g := StructuredArrow.mk (Y := CostructuredArrow.mk b) (CostructuredArrow.homMk a comm) /-- Constructor for objects in `w.CostructuredArrowDownwards g`. -/ -abbrev CostructuredArrowDownwards.mk : w.CostructuredArrowDownwards g := +abbrev CostructuredArrowDownwards.mk (comm : R.map a ≫ w.app X₁ ≫ B.map b = g) : + w.CostructuredArrowDownwards g := CostructuredArrow.mk (Y := StructuredArrow.mk a) (StructuredArrow.homMk b (by simpa using comm)) +variable (comm : R.map a ≫ w.app X₁ ≫ B.map b = g) variable {w g} lemma StructuredArrowRightwards.mk_surjective diff --git a/Mathlib/CategoryTheory/IsomorphismClasses.lean b/Mathlib/CategoryTheory/IsomorphismClasses.lean index ffb19388879b2..847b0684305c1 100644 --- a/Mathlib/CategoryTheory/IsomorphismClasses.lean +++ b/Mathlib/CategoryTheory/IsomorphismClasses.lean @@ -42,14 +42,14 @@ def isomorphismClasses : Cat.{v, u} ⥤ Type u where map {C D} F := Quot.map F.obj fun X Y ⟨f⟩ => ⟨F.mapIso f⟩ map_id {C} := by -- Porting note: this used to be `tidy` dsimp; apply funext; intro x - apply x.recOn -- Porting note: `induction x` not working yet + apply @Quot.recOn _ _ _ x · intro _ _ p simp only [types_id_apply] · intro _ rfl map_comp {C D E} f g := by -- Porting note(s): idem dsimp; apply funext; intro x - apply x.recOn + apply @Quot.recOn _ _ _ x · intro _ _ _ simp only [types_id_apply] · intro _ diff --git a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean index 4308ebc2e6046..e033aac35d916 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Adjunction.lean @@ -30,16 +30,18 @@ namespace CommSq section variable {A B : C} {X Y : D} {i : A ⟶ B} {p : X ⟶ Y} {u : G.obj A ⟶ X} {v : G.obj B ⟶ Y} - (sq : CommSq u (G.map i) p v) (adj : G ⊣ F) /-- When we have an adjunction `G ⊣ F`, any commutative square where the left map is of the form `G.map i` and the right map is `p` has an "adjoint" commutative square whose left map is `i` and whose right map is `F.map p`. -/ -theorem right_adjoint : CommSq (adj.homEquiv _ _ u) i (F.map p) (adj.homEquiv _ _ v) := +theorem right_adjoint (sq : CommSq u (G.map i) p v) (adj : G ⊣ F) : + CommSq (adj.homEquiv _ _ u) i (F.map p) (adj.homEquiv _ _ v) := ⟨by simp only [Adjunction.homEquiv_unit, assoc, ← F.map_comp, sq.w] rw [F.map_comp, Adjunction.unit_naturality_assoc]⟩ +variable (sq : CommSq u (G.map i) p v) (adj : G ⊣ F) + /-- The liftings of a commutative are in bijection with the liftings of its (right) adjoint square. -/ def rightAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.right_adjoint adj).LiftStruct where @@ -72,19 +74,22 @@ end section variable {A B : C} {X Y : D} {i : A ⟶ B} {p : X ⟶ Y} {u : A ⟶ F.obj X} {v : B ⟶ F.obj Y} - (sq : CommSq u i (F.map p) v) (adj : G ⊣ F) /-- When we have an adjunction `G ⊣ F`, any commutative square where the left map is of the form `i` and the right map is `F.map p` has an "adjoint" commutative square whose left map is `G.map i` and whose right map is `p`. -/ -theorem left_adjoint : CommSq ((adj.homEquiv _ _).symm u) (G.map i) p ((adj.homEquiv _ _).symm v) := +theorem left_adjoint (sq : CommSq u i (F.map p) v) (adj : G ⊣ F) : + CommSq ((adj.homEquiv _ _).symm u) (G.map i) p ((adj.homEquiv _ _).symm v) := ⟨by simp only [Adjunction.homEquiv_counit, assoc, ← G.map_comp_assoc, ← sq.w] rw [G.map_comp, assoc, Adjunction.counit_naturality]⟩ +variable (sq : CommSq u i (F.map p) v) (adj : G ⊣ F) + /-- The liftings of a commutative are in bijection with the liftings of its (left) adjoint square. -/ -def leftAdjointLiftStructEquiv : sq.LiftStruct ≃ (sq.left_adjoint adj).LiftStruct where +def leftAdjointLiftStructEquiv : + sq.LiftStruct ≃ (sq.left_adjoint adj).LiftStruct where toFun l := { l := (adj.homEquiv _ _).symm l.l fac_left := by rw [← adj.homEquiv_naturality_left_symm, l.fac_left] diff --git a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean index 4634384c9f8e7..a194496163cb3 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean @@ -90,20 +90,21 @@ This is a helper lemma for `hasFiniteProductsOfHasBinaryAndTerminal`, which is m than this. -/ private theorem hasProduct_fin : ∀ (n : ℕ) (f : Fin n → C), HasProduct f - | 0 => fun f => by + | 0 => fun _ => letI : HasLimitsOfShape (Discrete (Fin 0)) C := hasLimitsOfShape_of_equivalence (Discrete.equivalence.{0} finZeroEquiv'.symm) - infer_instance - | n + 1 => fun f => by + inferInstance + | n + 1 => fun f => haveI := hasProduct_fin n - apply HasLimit.mk ⟨_, extendFanIsLimit f (limit.isLimit _) (limit.isLimit _)⟩ + HasLimit.mk ⟨_, extendFanIsLimit f (limit.isLimit _) (limit.isLimit _)⟩ /-- If `C` has a terminal object and binary products, then it has finite products. -/ -theorem hasFiniteProducts_of_has_binary_and_terminal : HasFiniteProducts C := by - refine ⟨fun n => ⟨fun K => ?_⟩⟩ - letI := hasProduct_fin n fun n => K.obj ⟨n⟩ - let that : (Discrete.functor fun n => K.obj ⟨n⟩) ≅ K := Discrete.natIso fun ⟨i⟩ => Iso.refl _ - apply @hasLimitOfIso _ _ _ _ _ _ this that +theorem hasFiniteProducts_of_has_binary_and_terminal : HasFiniteProducts C := + ⟨fun n => ⟨fun K => + let this := hasProduct_fin n fun n => K.obj ⟨n⟩ + let that : (Discrete.functor fun n => K.obj ⟨n⟩) ≅ K := Discrete.natIso fun ⟨_⟩ => Iso.refl _ + @hasLimitOfIso _ _ _ _ _ _ this that⟩⟩ + end @@ -222,21 +223,20 @@ This is a helper lemma for `hasCofiniteProductsOfHasBinaryAndTerminal`, which is than this. -/ private theorem hasCoproduct_fin : ∀ (n : ℕ) (f : Fin n → C), HasCoproduct f - | 0 => fun f => by + | 0 => fun _ => letI : HasColimitsOfShape (Discrete (Fin 0)) C := hasColimitsOfShape_of_equivalence (Discrete.equivalence.{0} finZeroEquiv'.symm) - infer_instance - | n + 1 => fun f => by + inferInstance + | n + 1 => fun f => haveI := hasCoproduct_fin n - apply - HasColimit.mk ⟨_, extendCofanIsColimit f (colimit.isColimit _) (colimit.isColimit _)⟩ + HasColimit.mk ⟨_, extendCofanIsColimit f (colimit.isColimit _) (colimit.isColimit _)⟩ /-- If `C` has an initial object and binary coproducts, then it has finite coproducts. -/ -theorem hasFiniteCoproducts_of_has_binary_and_initial : HasFiniteCoproducts C := by - refine ⟨fun n => ⟨fun K => ?_⟩⟩ - letI := hasCoproduct_fin n fun n => K.obj ⟨n⟩ - let that : K ≅ Discrete.functor fun n => K.obj ⟨n⟩ := Discrete.natIso fun ⟨i⟩ => Iso.refl _ - apply @hasColimitOfIso _ _ _ _ _ _ this that +theorem hasFiniteCoproducts_of_has_binary_and_initial : HasFiniteCoproducts C := + ⟨fun n => ⟨fun K => + letI := hasCoproduct_fin n fun n => K.obj ⟨n⟩ + let that : K ≅ Discrete.functor fun n => K.obj ⟨n⟩ := Discrete.natIso fun ⟨_⟩ => Iso.refl _ + @hasColimitOfIso _ _ _ _ _ _ this that⟩⟩ end diff --git a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean index 4e1ecc2a64cc3..824a394d211e9 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean @@ -46,15 +46,16 @@ namespace HasLimitOfHasProductsOfHasEqualizers variable {F : J ⥤ C} {c₁ : Fan F.obj} {c₂ : Fan fun f : Σp : J × J, p.1 ⟶ p.2 => F.obj f.1.2} (s t : c₁.pt ⟶ c₂.pt) - (hs : ∀ f : Σp : J × J, p.1 ⟶ p.2, s ≫ c₂.π.app ⟨f⟩ = c₁.π.app ⟨f.1.1⟩ ≫ F.map f.2) - (ht : ∀ f : Σp : J × J, p.1 ⟶ p.2, t ≫ c₂.π.app ⟨f⟩ = c₁.π.app ⟨f.1.2⟩) (i : Fork s t) /-- (Implementation) Given the appropriate product and equalizer cones, build the cone for `F` which is limiting if the given cones are also. -/ @[simps] -def buildLimit : Cone F where +def buildLimit + (hs : ∀ f : Σp : J × J, p.1 ⟶ p.2, s ≫ c₂.π.app ⟨f⟩ = c₁.π.app ⟨f.1.1⟩ ≫ F.map f.2) + (ht : ∀ f : Σp : J × J, p.1 ⟶ p.2, t ≫ c₂.π.app ⟨f⟩ = c₁.π.app ⟨f.1.2⟩) + (i : Fork s t) : Cone F where pt := i.pt π := { app := fun j => i.ι ≫ c₁.π.app ⟨_⟩ @@ -62,7 +63,10 @@ def buildLimit : Cone F where dsimp rw [Category.id_comp, Category.assoc, ← hs ⟨⟨_, _⟩, f⟩, i.condition_assoc, ht] } -variable {i} +variable + (hs : ∀ f : Σp : J × J, p.1 ⟶ p.2, s ≫ c₂.π.app ⟨f⟩ = c₁.π.app ⟨f.1.1⟩ ≫ F.map f.2) + (ht : ∀ f : Σp : J × J, p.1 ⟶ p.2, t ≫ c₂.π.app ⟨f⟩ = c₁.π.app ⟨f.1.2⟩) + {i : Fork s t} /-- (Implementation) Show the cone constructed in `buildLimit` is limiting, provided the cones used in @@ -252,14 +256,15 @@ namespace HasColimitOfHasCoproductsOfHasCoequalizers variable {F : J ⥤ C} {c₁ : Cofan fun f : Σp : J × J, p.1 ⟶ p.2 => F.obj f.1.1} {c₂ : Cofan F.obj} (s t : c₁.pt ⟶ c₂.pt) - (hs : ∀ f : Σp : J × J, p.1 ⟶ p.2, c₁.ι.app ⟨f⟩ ≫ s = F.map f.2 ≫ c₂.ι.app ⟨f.1.2⟩) - (ht : ∀ f : Σp : J × J, p.1 ⟶ p.2, c₁.ι.app ⟨f⟩ ≫ t = c₂.ι.app ⟨f.1.1⟩) (i : Cofork s t) /-- (Implementation) Given the appropriate coproduct and coequalizer cocones, build the cocone for `F` which is colimiting if the given cocones are also. -/ @[simps] -def buildColimit : Cocone F where +def buildColimit + (hs : ∀ f : Σp : J × J, p.1 ⟶ p.2, c₁.ι.app ⟨f⟩ ≫ s = F.map f.2 ≫ c₂.ι.app ⟨f.1.2⟩) + (ht : ∀ f : Σp : J × J, p.1 ⟶ p.2, c₁.ι.app ⟨f⟩ ≫ t = c₂.ι.app ⟨f.1.1⟩) + (i : Cofork s t) : Cocone F where pt := i.pt ι := { app := fun j => c₂.ι.app ⟨_⟩ ≫ i.π @@ -270,7 +275,10 @@ def buildColimit : Cocone F where simp only [← Category.assoc, eq_whisker (hs f)] rw [Category.comp_id, ← reassoced ⟨⟨_, _⟩, f⟩, i.condition, ← Category.assoc, ht] } -variable {i} +variable + (hs : ∀ f : Σp : J × J, p.1 ⟶ p.2, c₁.ι.app ⟨f⟩ ≫ s = F.map f.2 ≫ c₂.ι.app ⟨f.1.2⟩) + (ht : ∀ f : Σp : J × J, p.1 ⟶ p.2, c₁.ι.app ⟨f⟩ ≫ t = c₂.ι.app ⟨f.1.1⟩) + {i : Cofork s t} /-- (Implementation) Show the cocone constructed in `buildColimit` is colimiting, provided the cocones used in its construction are. diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean b/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean index 86d0270b6d944..873dd0fc1eda0 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean @@ -156,7 +156,7 @@ theorem over_hasTerminal (B : C) : HasTerminal (Over B) where π := { app := fun p => p.as.elim } } isLimit := - { lift := fun s => Over.homMk _ + { lift := fun s => Over.homMk s.pt.hom fac := fun _ j => j.as.elim uniq := fun s m _ => by simp only diff --git a/Mathlib/CategoryTheory/Limits/EpiMono.lean b/Mathlib/CategoryTheory/Limits/EpiMono.lean index c6f3e9eb4fc02..3f66f04420bdd 100644 --- a/Mathlib/CategoryTheory/Limits/EpiMono.lean +++ b/Mathlib/CategoryTheory/Limits/EpiMono.lean @@ -33,9 +33,9 @@ variable {C : Type*} [Category C] {X Y : C} {f : X ⟶ Y} section Mono -variable {c : PullbackCone f f} (hc : IsLimit c) +variable {c : PullbackCone f f} -lemma mono_iff_fst_eq_snd : Mono f ↔ c.fst = c.snd := by +lemma mono_iff_fst_eq_snd (hc : IsLimit c) : Mono f ↔ c.fst = c.snd := by constructor · intro hf simpa only [← cancel_mono f] using c.condition @@ -45,7 +45,7 @@ lemma mono_iff_fst_eq_snd : Mono f ↔ c.fst = c.snd := by obtain ⟨φ, rfl, rfl⟩ := PullbackCone.IsLimit.lift' hc g g' h rw [hf] -lemma mono_iff_isIso_fst : Mono f ↔ IsIso c.fst := by +lemma mono_iff_isIso_fst (hc : IsLimit c) : Mono f ↔ IsIso c.fst := by rw [mono_iff_fst_eq_snd hc] constructor · intro h @@ -61,7 +61,7 @@ lemma mono_iff_isIso_fst : Mono f ↔ IsIso c.fst := by rw [← cancel_mono c.fst, assoc, id_comp, hφ₁, comp_id])⟩ rw [← cancel_epi φ, hφ₁, hφ₂] -lemma mono_iff_isIso_snd : Mono f ↔ IsIso c.snd := +lemma mono_iff_isIso_snd (hc : IsLimit c) : Mono f ↔ IsIso c.snd := mono_iff_isIso_fst (PullbackCone.flipIsLimit hc) variable (f) @@ -77,9 +77,9 @@ end Mono section Epi -variable {c : PushoutCocone f f} (hc : IsColimit c) +variable {c : PushoutCocone f f} -lemma epi_iff_inl_eq_inr : Epi f ↔ c.inl = c.inr := by +lemma epi_iff_inl_eq_inr (hc : IsColimit c) : Epi f ↔ c.inl = c.inr := by constructor · intro hf simpa only [← cancel_epi f] using c.condition @@ -89,7 +89,7 @@ lemma epi_iff_inl_eq_inr : Epi f ↔ c.inl = c.inr := by obtain ⟨φ, rfl, rfl⟩ := PushoutCocone.IsColimit.desc' hc g g' h rw [hf] -lemma epi_iff_isIso_inl : Epi f ↔ IsIso c.inl := by +lemma epi_iff_isIso_inl (hc : IsColimit c) : Epi f ↔ IsIso c.inl := by rw [epi_iff_inl_eq_inr hc] constructor · intro h @@ -105,7 +105,7 @@ lemma epi_iff_isIso_inl : Epi f ↔ IsIso c.inl := by rw [← cancel_epi c.inl, reassoc_of% hφ₁, comp_id])⟩ rw [← cancel_mono φ, hφ₁, hφ₂] -lemma epi_iff_isIso_inr : Epi f ↔ IsIso c.inr := +lemma epi_iff_isIso_inr (hc : IsColimit c) : Epi f ↔ IsIso c.inr := epi_iff_isIso_inl (PushoutCocone.flipIsColimit hc) variable (f) diff --git a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean index 34842fb47fcbb..4f00dd6062184 100644 --- a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean +++ b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean @@ -157,13 +157,11 @@ end section -variable [MonoCoprod C] {I J : Type*} (X : I → C) (ι : J → I) (hι : Function.Injective ι) +variable [MonoCoprod C] {I J : Type*} (X : I → C) (ι : J → I) -section - -variable (c : Cofan X) (c₁ : Cofan (X ∘ ι)) (hc : IsColimit c) (hc₁ : IsColimit c₁) - -lemma mono_of_injective_aux (c₂ : Cofan (fun (k : ((Set.range ι)ᶜ : Set I)) => X k.1)) +lemma mono_of_injective_aux (hι : Function.Injective ι) (c : Cofan X) (c₁ : Cofan (X ∘ ι)) + (hc : IsColimit c) (hc₁ : IsColimit c₁) + (c₂ : Cofan (fun (k : ((Set.range ι)ᶜ : Set I)) => X k.1)) (hc₂ : IsColimit c₂) : Mono (Cofan.IsColimit.desc hc₁ (fun i => c.inj (ι i))) := by classical let e := ((Equiv.ofInjective ι hι).sumCongr (Equiv.refl _)).trans (Equiv.Set.sumCompl _) @@ -172,12 +170,13 @@ lemma mono_of_injective_aux (c₂ : Cofan (fun (k : ((Set.range ι)ᶜ : Set I)) exact IsColimit.ofIsoColimit ((IsColimit.ofCoconeEquiv (Cocones.equivalenceOfReindexing (Discrete.equivalence e) (Iso.refl _))).symm hc) (Cocones.ext (Iso.refl _)) +variable (hι : Function.Injective ι) (c : Cofan X) (c₁ : Cofan (X ∘ ι)) + (hc : IsColimit c) (hc₁ : IsColimit c₁) + lemma mono_of_injective [HasCoproduct (fun (k : ((Set.range ι)ᶜ : Set I)) => X k.1)] : Mono (Cofan.IsColimit.desc hc₁ (fun i => c.inj (ι i))) := mono_of_injective_aux X ι hι c c₁ hc hc₁ _ (colimit.isColimit _) -end - lemma mono_of_injective' [HasCoproduct (X ∘ ι)] [HasCoproduct X] [HasCoproduct (fun (k : ((Set.range ι)ᶜ : Set I)) => X k.1)] : Mono (Sigma.desc (f := X ∘ ι) (fun j => Sigma.ι X (ι j))) := diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean index 1bb67b2f7b98f..c17fc18f42fda 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean @@ -5,6 +5,8 @@ Authors: Bhavik Mehta, Andrew Yang -/ import Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback import Mathlib.CategoryTheory.Limits.Preserves.Basic +import Mathlib.CategoryTheory.Limits.Opposites +import Mathlib.CategoryTheory.Limits.Yoneda /-! # Preserving pullbacks @@ -37,6 +39,30 @@ section Pullback variable {C : Type u₁} [Category.{v₁} C] variable {D : Type u₂} [Category.{v₂} D] + +namespace PullbackCone + +variable {X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} (c : PullbackCone f g) (G : C ⥤ D) + +/-- The image of a pullback cone by a functor. -/ +abbrev map : PullbackCone (G.map f) (G.map g) := + PullbackCone.mk (G.map c.fst) (G.map c.snd) + (by simpa using G.congr_map c.condition) + +/-- The map (as a cone) of a pullback cone is limit iff +the map (as a pullback cone) is limit. -/ +def isLimitMapConeEquiv : + IsLimit (mapCone G c) ≃ IsLimit (c.map G) := + (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₂} _) _).symm.trans <| + IsLimit.equivIsoLimit <| by + refine PullbackCone.ext (Iso.refl _) ?_ ?_ + · dsimp only [fst] + simp + · dsimp only [snd] + simp + +end PullbackCone + variable (G : C ⥤ D) variable {W X Y Z : C} {f : X ⟶ Z} {g : Y ⟶ Z} {h : W ⟶ X} {k : W ⟶ Y} (comm : h ≫ f = k ≫ g) @@ -47,23 +73,21 @@ def isLimitMapConePullbackConeEquiv : IsLimit (PullbackCone.mk (G.map h) (G.map k) (by simp only [← G.map_comp, comm]) : PullbackCone (G.map f) (G.map g)) := - (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₂} _) _).symm.trans <| - IsLimit.equivIsoLimit <| - Cones.ext (Iso.refl _) <| by - rintro (_ | _ | _) <;> dsimp <;> simp only [comp_id, id_comp, G.map_comp] + (PullbackCone.mk _ _ comm).isLimitMapConeEquiv G /-- The property of preserving pullbacks expressed in terms of binary fans. -/ def isLimitPullbackConeMapOfIsLimit [PreservesLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk h k comm)) : have : G.map h ≫ G.map f = G.map k ≫ G.map g := by rw [← G.map_comp, ← G.map_comp,comm] IsLimit (PullbackCone.mk (G.map h) (G.map k) this) := - isLimitMapConePullbackConeEquiv G comm (PreservesLimit.preserves l) + (PullbackCone.isLimitMapConeEquiv _ G).1 (PreservesLimit.preserves l) /-- The property of reflecting pullbacks expressed in terms of binary fans. -/ def isLimitOfIsLimitPullbackConeMap [ReflectsLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk (G.map h) (G.map k) (show G.map h ≫ G.map f = G.map k ≫ G.map g from by simp only [← G.map_comp,comm]))) : IsLimit (PullbackCone.mk h k comm) := - ReflectsLimit.reflects ((isLimitMapConePullbackConeEquiv G comm).symm l) + ReflectsLimit.reflects + ((PullbackCone.isLimitMapConeEquiv (PullbackCone.mk _ _ comm) G).2 l) variable (f g) [PreservesLimit (cospan f g) G] @@ -126,12 +150,42 @@ theorem PreservesPullback.iso_inv_snd : (PreservesPullback.iso G f g).inv ≫ G.map (pullback.snd f g) = pullback.snd _ _ := by simp [PreservesPullback.iso, Iso.inv_comp_eq] +/-- A pullback cone in `C` is limit iff if it is so after the application +of `coyoneda.obj X` for all `X : Cᵒᵖ`. -/ +def PullbackCone.isLimitCoyonedaEquiv (c : PullbackCone f g) : + IsLimit c ≃ ∀ (X : Cᵒᵖ), IsLimit (c.map (coyoneda.obj X)) := + (Cone.isLimitCoyonedaEquiv c).trans + (Equiv.piCongrRight (fun X ↦ c.isLimitMapConeEquiv (coyoneda.obj X))) + end Pullback section Pushout variable {C : Type u₁} [Category.{v₁} C] variable {D : Type u₂} [Category.{v₂} D] + +namespace PushoutCocone + +variable {W X Y : C} {f : W ⟶ X} {g : W ⟶ Y} (c : PushoutCocone f g) (G : C ⥤ D) + +/-- The image of a pullback cone by a functor. -/ +abbrev map : PushoutCocone (G.map f) (G.map g) := + PushoutCocone.mk (G.map c.inl) (G.map c.inr) (by simpa using G.congr_map c.condition) + +/-- The map (as a cocone) of a pushout cocone is colimit iff +the map (as a pushout cocone) is limit. -/ +def isColimitMapCoconeEquiv : + IsColimit (mapCocone G c) ≃ IsColimit (c.map G) := + (IsColimit.precomposeHomEquiv (diagramIsoSpan.{v₂} _).symm _).symm.trans <| + IsColimit.equivIsoColimit <| by + refine PushoutCocone.ext (Iso.refl _) ?_ ?_ + · dsimp only [inl] + simp + · dsimp only [inr] + simp + +end PushoutCocone + variable (G : C ⥤ D) variable {W X Y Z : C} {h : X ⟶ Z} {k : Y ⟶ Z} {f : W ⟶ X} {g : W ⟶ Y} (comm : f ≫ h = g ≫ k) @@ -272,6 +326,19 @@ instance : IsIso (pushoutComparison G f g) := by rw [← PreservesPushout.iso_hom] infer_instance +/-- A pushout cocone in `C` is colimit iff it becomes limit +after the application of `yoneda.obj X` for all `X : C`. -/ +def PushoutCocone.isColimitYonedaEquiv (c : PushoutCocone f g) : + IsColimit c ≃ ∀ (X : C), IsLimit (c.op.map (yoneda.obj X)) := + (Limits.Cocone.isColimitYonedaEquiv c).trans + (Equiv.piCongrRight (fun X ↦ + (IsLimit.whiskerEquivalenceEquiv walkingSpanOpEquiv.symm).trans + ((IsLimit.postcomposeHomEquiv + (isoWhiskerRight (cospanOp f g).symm (yoneda.obj X)) _).symm.trans + (Equiv.trans (IsLimit.equivIsoLimit + (by exact Cones.ext (Iso.refl _) (by rintro (_|_|_) <;> simp))) + (c.op.isLimitMapConeEquiv (yoneda.obj X)))))) + end Pushout end diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index 107d67f5b3840..616a79d752ba6 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -359,8 +359,7 @@ noncomputable def compYonedaIsoYonedaCompLan : (yonedaMap F X) (F.op.lan.obj (yoneda.obj X)) (F.op.lanUnit.app (yoneda.obj X))) _) (𝟙 _) have eq₃ := congr_fun (congr_app (F.op.lanUnit.naturality (yoneda.map f)) _) (𝟙 _) dsimp at eq₁ eq₂ eq₃ - rw [yonedaMap_app_apply] at eq₁ - simp only [yonedaMap_app_apply, Functor.map_id] at eq₂ + simp only [Functor.map_id] at eq₂ simp only [id_comp] at eq₃ simp [yonedaEquiv, eq₁, eq₂, eq₃]) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index ac842f5627917..ecb6dfd87d021 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -961,10 +961,10 @@ variable {C : Type u} [Category.{v} C] [HasZeroMorphisms C] variable {D : Type uD} [Category.{uD'} D] [HasZeroMorphisms D] instance biproduct.ι_mono (f : J → C) [HasBiproduct f] (b : J) : IsSplitMono (biproduct.ι f b) := - IsSplitMono.mk' { retraction := biproduct.desc <| Pi.single b _ } + IsSplitMono.mk' { retraction := biproduct.desc <| Pi.single b (𝟙 (f b)) } instance biproduct.π_epi (f : J → C) [HasBiproduct f] (b : J) : IsSplitEpi (biproduct.π f b) := - IsSplitEpi.mk' { section_ := biproduct.lift <| Pi.single b _ } + IsSplitEpi.mk' { section_ := biproduct.lift <| Pi.single b (𝟙 (f b)) } /-- Auxiliary lemma for `biproduct.uniqueUpToIso`. -/ theorem biproduct.conePointUniqueUpToIso_hom (f : J → C) [HasBiproduct f] {b : Bicone f} diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean index 9107ff51d52a7..4c7ff0cbdedd3 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean @@ -214,8 +214,8 @@ def pullbackDiagonalMapIdIso : pullbackDiagonalMapIso i (𝟙 _) (f ≫ inv (pullback.fst _ _)) (g ≫ inv (pullback.fst _ _)) ≪≫ ?_ · refine @asIso _ _ _ _ (pullback.map _ _ _ _ (𝟙 T) ((pullback.congrHom ?_ ?_).hom) (𝟙 _) ?_ ?_) ?_ - · rw [← Category.comp_id (pullback.snd _ _), ← condition, Category.assoc, IsIso.inv_hom_id_assoc] - · rw [← Category.comp_id (pullback.snd _ _), ← condition, Category.assoc, IsIso.inv_hom_id_assoc] + · rw [← Category.comp_id (pullback.snd ..), ← condition, Category.assoc, IsIso.inv_hom_id_assoc] + · rw [← Category.comp_id (pullback.snd ..), ← condition, Category.assoc, IsIso.inv_hom_id_assoc] · rw [Category.comp_id, Category.id_comp] · ext <;> simp · infer_instance @@ -380,8 +380,8 @@ def pullbackFstFstIso {X Y S X' Y' S' : C} (f : X ⟶ S) (g : Y ⟶ S) (f' : X' pullback.condition_assoc]) inv := pullback.lift - (pullback.lift (pullback.map _ _ _ _ _ _ _ e₁ e₂) (pullback.fst _ _) (pullback.lift_fst _ _ _)) - (pullback.lift (pullback.map _ _ _ _ _ _ _ e₁ e₂) (pullback.snd _ _) (pullback.lift_snd _ _ _)) + (pullback.lift (pullback.map _ _ _ _ _ _ _ e₁ e₂) (pullback.fst _ _) (pullback.lift_fst ..)) + (pullback.lift (pullback.map _ _ _ _ _ _ _ e₁ e₂) (pullback.snd _ _) (pullback.lift_snd ..)) (by rw [pullback.lift_fst, pullback.lift_fst]) hom_inv_id := by -- We could use `ext` here to immediately descend to the leaf goals, @@ -391,8 +391,7 @@ def pullbackFstFstIso {X Y S X' Y' S' : C} (f : X ⟶ S) (g : Y ⟶ S) (f' : X' · apply pullback.hom_ext · simp only [Category.assoc, lift_fst, lift_fst_assoc, Category.id_comp] rw [condition] - · simp [Category.assoc, lift_snd] - rw [condition_assoc, condition] + · simp [Category.assoc, lift_snd, condition_assoc, condition] · simp only [Category.assoc, lift_fst_assoc, lift_snd, lift_fst, Category.id_comp] · apply pullback.hom_ext · apply pullback.hom_ext diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean index 6cfd112238b78..3bc2f8b2e9d02 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Joël Riou +Authors: Scott Morrison, Joël Riou, Calle Sönne -/ import Mathlib.CategoryTheory.CommSq import Mathlib.CategoryTheory.Limits.Opposites @@ -201,6 +201,25 @@ theorem cone_snd (h : IsPullback fst snd f g) : h.cone.snd = snd := noncomputable def isLimit (h : IsPullback fst snd f g) : IsLimit h.cone := h.isLimit'.some +/-- API for PullbackCone.IsLimit.lift for `IsPullback` -/ +noncomputable def lift (hP : IsPullback fst snd f g) {W : C} (h : W ⟶ X) (k : W ⟶ Y) + (w : h ≫ f = k ≫ g) : W ⟶ P := + PullbackCone.IsLimit.lift hP.isLimit h k w + +@[reassoc (attr := simp)] +lemma lift_fst (hP : IsPullback fst snd f g) {W : C} (h : W ⟶ X) (k : W ⟶ Y) + (w : h ≫ f = k ≫ g) : hP.lift h k w ≫ fst = h := + PullbackCone.IsLimit.lift_fst hP.isLimit h k w + +@[reassoc (attr := simp)] +lemma lift_snd (hP : IsPullback fst snd f g) {W : C} (h : W ⟶ X) (k : W ⟶ Y) + (w : h ≫ f = k ≫ g) : hP.lift h k w ≫ snd = k := + PullbackCone.IsLimit.lift_snd hP.isLimit h k w + +lemma hom_ext (hP : IsPullback fst snd f g) {W : C} {k l : W ⟶ P} + (h₀ : k ≫ fst = l ≫ fst) (h₁ : k ≫ snd = l ≫ snd) : k = l := + PullbackCone.IsLimit.hom_ext hP.isLimit h₀ h₁ + /-- If `c` is a limiting pullback cone, then we have an `IsPullback c.fst c.snd f g`. -/ theorem of_isLimit {c : PullbackCone f g} (h : Limits.IsLimit c) : IsPullback c.fst c.snd f g := { w := c.condition @@ -249,6 +268,38 @@ theorem of_hasBinaryProduct [HasBinaryProduct X Y] [HasZeroObject C] [HasZeroMor convert @of_is_product _ _ X Y 0 _ (limit.isLimit _) HasZeroObject.zeroIsTerminal <;> subsingleton +section + +variable {P': C} {fst' : P' ⟶ X} {snd' : P' ⟶ Y} + +/-- Any object at the top left of a pullback square is isomorphic to the object at the top left +of any other pullback square with the same cospan. -/ +noncomputable def isoIsPullback (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : + P ≅ P' := + IsLimit.conePointUniqueUpToIso h.isLimit h'.isLimit + +@[reassoc (attr := simp)] +theorem isoIsPullback_hom_fst (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : + (h.isoIsPullback h').hom ≫ fst' = fst := + IsLimit.conePointUniqueUpToIso_hom_comp h.isLimit h'.isLimit WalkingCospan.left + +@[reassoc (attr := simp)] +theorem isoIsPullback_hom_snd (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : + (h.isoIsPullback h').hom ≫ snd' = snd := + IsLimit.conePointUniqueUpToIso_hom_comp h.isLimit h'.isLimit WalkingCospan.right + +@[reassoc (attr := simp)] +theorem isoIsPullback_inv_fst (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : + (h.isoIsPullback h').inv ≫ fst = fst' := by + simp only [Iso.inv_comp_eq, isoIsPullback_hom_fst] + +@[reassoc (attr := simp)] +theorem isoIsPullback_inv_snd (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : + (h.isoIsPullback h').inv ≫ snd = snd' := by + simp only [Iso.inv_comp_eq, isoIsPullback_hom_snd] + +end + variable {X Y} /-- Any object at the top left of a pullback square is @@ -256,23 +307,23 @@ isomorphic to the pullback provided by the `HasLimit` API. -/ noncomputable def isoPullback (h : IsPullback fst snd f g) [HasPullback f g] : P ≅ pullback f g := (limit.isoLimitCone ⟨_, h.isLimit⟩).symm -@[simp] +@[reassoc (attr := simp)] theorem isoPullback_hom_fst (h : IsPullback fst snd f g) [HasPullback f g] : h.isoPullback.hom ≫ pullback.fst _ _ = fst := by dsimp [isoPullback, cone, CommSq.cone] simp -@[simp] +@[reassoc (attr := simp)] theorem isoPullback_hom_snd (h : IsPullback fst snd f g) [HasPullback f g] : h.isoPullback.hom ≫ pullback.snd _ _ = snd := by dsimp [isoPullback, cone, CommSq.cone] simp -@[simp] +@[reassoc (attr := simp)] theorem isoPullback_inv_fst (h : IsPullback fst snd f g) [HasPullback f g] : h.isoPullback.inv ≫ fst = pullback.fst _ _ := by simp [Iso.inv_comp_eq] -@[simp] +@[reassoc (attr := simp)] theorem isoPullback_inv_snd (h : IsPullback fst snd f g) [HasPullback f g] : h.isoPullback.inv ≫ snd = pullback.snd _ _ := by simp [Iso.inv_comp_eq] @@ -337,6 +388,25 @@ theorem cocone_inr (h : IsPushout f g inl inr) : h.cocone.inr = inr := noncomputable def isColimit (h : IsPushout f g inl inr) : IsColimit h.cocone := h.isColimit'.some +/-- API for PushoutCocone.IsColimit.lift for `IsPushout` -/ +noncomputable def desc (hP : IsPushout f g inl inr) {W : C} (h : X ⟶ W) (k : Y ⟶ W) + (w : f ≫ h = g ≫ k) : P ⟶ W := + PushoutCocone.IsColimit.desc hP.isColimit h k w + +@[reassoc (attr := simp)] +lemma inl_desc (hP : IsPushout f g inl inr) {W : C} (h : X ⟶ W) (k : Y ⟶ W) + (w : f ≫ h = g ≫ k) : inl ≫ hP.desc h k w = h := + PushoutCocone.IsColimit.inl_desc hP.isColimit h k w + +@[reassoc (attr := simp)] +lemma inr_desc (hP : IsPushout f g inl inr) {W : C} (h : X ⟶ W) (k : Y ⟶ W) + (w : f ≫ h = g ≫ k) : inr ≫ hP.desc h k w = k := + PushoutCocone.IsColimit.inr_desc hP.isColimit h k w + +lemma hom_ext (hP : IsPushout f g inl inr) {W : C} {k l : P ⟶ W} + (h₀ : inl ≫ k = inl ≫ l) (h₁ : inr ≫ k = inr ≫ l) : k = l := + PushoutCocone.IsColimit.hom_ext hP.isColimit h₀ h₁ + /-- If `c` is a colimiting pushout cocone, then we have an `IsPushout f g c.inl c.inr`. -/ theorem of_isColimit {c : PushoutCocone f g} (h : Limits.IsColimit c) : IsPushout f g c.inl c.inr := { w := c.condition @@ -386,6 +456,39 @@ theorem of_hasBinaryCoproduct [HasBinaryCoproduct X Y] [HasZeroObject C] [HasZer convert @of_is_coproduct _ _ 0 X Y _ (colimit.isColimit _) HasZeroObject.zeroIsInitial <;> subsingleton +section + +variable {P': C} {inl' : X ⟶ P'} {inr' : Y ⟶ P'} + +/-- Any object at the bottom right of a pushout square is isomorphic to the object at the bottom +right of any other pushout square with the same span. -/ +noncomputable def isoIsPushout (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : + P ≅ P' := + IsColimit.coconePointUniqueUpToIso h.isColimit h'.isColimit + +@[reassoc (attr := simp)] +theorem inl_isoIsPushout_hom (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : + inl ≫ (h.isoIsPushout h').hom = inl' := + IsColimit.comp_coconePointUniqueUpToIso_hom h.isColimit h'.isColimit WalkingSpan.left + +@[reassoc (attr := simp)] +theorem inr_isoIsPushout_hom (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : + inr ≫ (h.isoIsPushout h').hom = inr' := + IsColimit.comp_coconePointUniqueUpToIso_hom h.isColimit h'.isColimit WalkingSpan.right + +@[reassoc (attr := simp)] +theorem inl_isoIsPushout_inv (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : + inl' ≫ (h.isoIsPushout h').inv = inl := by + simp only [Iso.comp_inv_eq, inl_isoIsPushout_hom] + +@[reassoc (attr := simp)] +theorem inr_isoIsPushout_inv (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : + inr' ≫ (h.isoIsPushout h').inv = inr := by + simp only [Iso.comp_inv_eq, inr_isoIsPushout_hom] + +end + + variable {X Y} /-- Any object at the top left of a pullback square is @@ -393,23 +496,23 @@ isomorphic to the pullback provided by the `HasLimit` API. -/ noncomputable def isoPushout (h : IsPushout f g inl inr) [HasPushout f g] : P ≅ pushout f g := (colimit.isoColimitCocone ⟨_, h.isColimit⟩).symm -@[simp] +@[reassoc (attr := simp)] theorem inl_isoPushout_inv (h : IsPushout f g inl inr) [HasPushout f g] : pushout.inl _ _ ≫ h.isoPushout.inv = inl := by dsimp [isoPushout, cocone, CommSq.cocone] simp -@[simp] +@[reassoc (attr := simp)] theorem inr_isoPushout_inv (h : IsPushout f g inl inr) [HasPushout f g] : pushout.inr _ _ ≫ h.isoPushout.inv = inr := by dsimp [isoPushout, cocone, CommSq.cocone] simp -@[simp] +@[reassoc (attr := simp)] theorem inl_isoPushout_hom (h : IsPushout f g inl inr) [HasPushout f g] : inl ≫ h.isoPushout.hom = pushout.inl _ _ := by simp [← Iso.eq_comp_inv] -@[simp] +@[reassoc (attr := simp)] theorem inr_isoPushout_hom (h : IsPushout f g inl inr) [HasPushout f g] : inr ≫ h.isoPushout.hom = pushout.inr _ _ := by simp [← Iso.eq_comp_inv] @@ -487,14 +590,38 @@ end -- Objects here are arranged in a 3x2 grid, and indexed by their xy coordinates. -- Morphisms are named `hᵢⱼ` for a horizontal morphism starting at `(i,j)`, -- and `vᵢⱼ` for a vertical morphism starting at `(i,j)`. -/-- Paste two pullback squares "vertically" to obtain another pullback square. -/ +/-- Paste two pullback squares "vertically" to obtain another pullback square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ +| | +v₁₁ v₁₂ +↓ ↓ +X₂₁ - h₂₁ -> X₂₂ +| | +v₂₁ v₂₂ +↓ ↓ +X₃₁ - h₃₁ -> X₃₂ +``` +-/ theorem paste_vert {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPullback h₁₁ v₁₁ v₁₂ h₂₁) (t : IsPullback h₂₁ v₂₁ v₂₂ h₃₁) : IsPullback h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁ := of_isLimit (bigSquareIsPullback _ _ _ _ _ _ _ s.w t.w t.isLimit s.isLimit) -/-- Paste two pullback squares "horizontally" to obtain another pullback square. -/ +/-- Paste two pullback squares "horizontally" to obtain another pullback square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ - h₁₂ -> X₁₃ +| | | +v₁₁ v₁₂ v₁₃ +↓ ↓ ↓ +X₂₁ - h₂₁ -> X₂₂ - h₂₂ -> X₂₃ +``` +-/ theorem paste_horiz {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} {h₂₂ : X₂₂ ⟶ X₂₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} (s : IsPullback h₁₁ v₁₁ v₁₂ h₂₁) (t : IsPullback h₁₂ v₁₂ v₁₃ h₂₂) : @@ -502,7 +629,21 @@ theorem paste_horiz {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁ (paste_vert s.flip t.flip).flip /-- Given a pullback square assembled from a commuting square on the top and -a pullback square on the bottom, the top square is a pullback square. -/ +a pullback square on the bottom, the top square is a pullback square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ +| | +v₁₁ v₁₂ +↓ ↓ +X₂₁ - h₂₁ -> X₂₂ +| | +v₂₁ v₂₂ +↓ ↓ +X₃₁ - h₃₁ -> X₃₂ +``` +-/ theorem of_bot {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPullback h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁) (p : h₁₁ ≫ v₁₂ = v₁₁ ≫ h₂₁) @@ -510,7 +651,17 @@ theorem of_bot {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : of_isLimit (leftSquareIsPullback _ _ _ _ _ _ _ p t.w t.isLimit s.isLimit) /-- Given a pullback square assembled from a commuting square on the left and -a pullback square on the right, the left square is a pullback square. -/ +a pullback square on the right, the left square is a pullback square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ - h₁₂ -> X₁₃ +| | | +v₁₁ v₁₂ v₁₃ +↓ ↓ ↓ +X₂₁ - h₂₁ -> X₂₂ - h₂₂ -> X₂₃ +``` +-/ theorem of_right {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} {h₂₂ : X₂₂ ⟶ X₂₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} (s : IsPullback (h₁₁ ≫ h₁₂) v₁₁ v₁₃ (h₂₁ ≫ h₂₂)) (p : h₁₁ ≫ v₁₂ = v₁₁ ≫ h₂₁) @@ -529,6 +680,46 @@ theorem paste_horiz_iff {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h IsPullback (h₁₁ ≫ h₁₂) v₁₁ v₁₃ (h₂₁ ≫ h₂₂) ↔ IsPullback h₁₁ v₁₁ v₁₂ h₂₁ := ⟨fun h => h.of_right e s, fun h => h.paste_horiz s⟩ +/-- Variant of `IsPullback.of_right` where `h₁₁` is induced from a morphism `h₁₃ : X₁₁ ⟶ X₁₃`, and +the universal property of the right square. + +The objects fit in the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ - h₁₂ -> X₁₃ +| | | +v₁₁ v₁₂ v₁₃ +↓ ↓ ↓ +X₂₁ - h₂₁ -> X₂₂ - h₂₂ -> X₂₃ +``` +-/ +theorem of_right' {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} + {h₂₂ : X₂₂ ⟶ X₂₃} {h₁₃ : X₁₁ ⟶ X₁₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} + (s : IsPullback h₁₃ v₁₁ v₁₃ (h₂₁ ≫ h₂₂)) (t : IsPullback h₁₂ v₁₂ v₁₃ h₂₂) : + IsPullback (t.lift h₁₃ (v₁₁ ≫ h₂₁) (by rw [s.w, Category.assoc])) v₁₁ v₁₂ h₂₁ := + of_right ((t.lift_fst _ _ _) ▸ s) (t.lift_snd _ _ _) t + +/-- Variant of `IsPullback.of_bot`, where `v₁₁` is induced from a morphism `v₃₁ : X₁₁ ⟶ X₃₁`, and +the universal property of the bottom square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ +| | +v₁₁ v₁₂ +↓ ↓ +X₂₁ - h₂₁ -> X₂₂ +| | +v₂₁ v₂₂ +↓ ↓ +X₃₁ - h₃₁ -> X₃₂ +``` +-/ +theorem of_bot' {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} + {h₃₁ : X₃₁ ⟶ X₃₂} {v₃₁ : X₁₁ ⟶ X₃₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} + (s : IsPullback h₁₁ v₃₁ (v₁₂ ≫ v₂₂) h₃₁) (t : IsPullback h₂₁ v₂₁ v₂₂ h₃₁) : + IsPullback h₁₁ (t.lift (h₁₁ ≫ v₁₂) v₃₁ (by rw [Category.assoc, s.w])) v₁₂ h₂₁ := + of_bot ((t.lift_snd _ _ _) ▸ s) (by simp only [lift_fst]) t + section variable [HasZeroObject C] [HasZeroMorphisms C] @@ -627,6 +818,30 @@ lemma of_id_fst : IsPullback (𝟙 _) f f (𝟙 _) := IsPullback.of_horiz_isIso lemma of_id_snd : IsPullback f (𝟙 _) (𝟙 _) f := IsPullback.of_vert_isIso ⟨by simp⟩ +/-- The following diagram is a pullback +``` +X --f--> Z +| | +id id +v v +X --f--> Z +``` +-/ +lemma id_vert (f : X ⟶ Z) : IsPullback f (𝟙 X) (𝟙 Z) f := + of_vert_isIso ⟨by simp only [Category.id_comp, Category.comp_id]⟩ + +/-- The following diagram is a pullback +``` +X --id--> X +| | +f f +v v +Z --id--> Z +``` +-/ +lemma id_horiz (f : X ⟶ Z) : IsPullback (𝟙 X) f f (𝟙 Z) := + of_horiz_isIso ⟨by simp only [Category.id_comp, Category.comp_id]⟩ + end IsPullback namespace IsPushout @@ -679,14 +894,38 @@ end -- Objects here are arranged in a 3x2 grid, and indexed by their xy coordinates. -- Morphisms are named `hᵢⱼ` for a horizontal morphism starting at `(i,j)`, -- and `vᵢⱼ` for a vertical morphism starting at `(i,j)`. -/-- Paste two pushout squares "vertically" to obtain another pushout square. -/ +/-- Paste two pushout squares "vertically" to obtain another pushout square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ +| | +v₁₁ v₁₂ +↓ ↓ +X₂₁ - h₂₁ -> X₂₂ +| | +v₂₁ v₂₂ +↓ ↓ +X₃₁ - h₃₁ -> X₃₂ +``` +-/ theorem paste_vert {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) (t : IsPushout h₂₁ v₂₁ v₂₂ h₃₁) : IsPushout h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁ := of_isColimit (bigSquareIsPushout _ _ _ _ _ _ _ s.w t.w t.isColimit s.isColimit) -/-- Paste two pushout squares "horizontally" to obtain another pushout square. -/ +/-- Paste two pushout squares "horizontally" to obtain another pushout square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ - h₁₂ -> X₁₃ +| | | +v₁₁ v₁₂ v₁₃ +↓ ↓ ↓ +X₂₁ - h₂₁ -> X₂₂ - h₂₂ -> X₂₃ +``` +-/ theorem paste_horiz {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} {h₂₂ : X₂₂ ⟶ X₂₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} (s : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) (t : IsPushout h₁₂ v₁₂ v₁₃ h₂₂) : @@ -694,32 +933,96 @@ theorem paste_horiz {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁ (paste_vert s.flip t.flip).flip /-- Given a pushout square assembled from a pushout square on the top and -a commuting square on the bottom, the bottom square is a pushout square. -/ -theorem of_bot {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} +a commuting square on the bottom, the bottom square is a pushout square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ +| | +v₁₁ v₁₂ +↓ ↓ +X₂₁ - h₂₁ -> X₂₂ +| | +v₂₁ v₂₂ +↓ ↓ +X₃₁ - h₃₁ -> X₃₂ +``` +-/ +theorem of_top {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPushout h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁) (p : h₂₁ ≫ v₂₂ = v₂₁ ≫ h₃₁) (t : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) : IsPushout h₂₁ v₂₁ v₂₂ h₃₁ := of_isColimit (rightSquareIsPushout _ _ _ _ _ _ _ t.w p t.isColimit s.isColimit) /-- Given a pushout square assembled from a pushout square on the left and -a commuting square on the right, the right square is a pushout square. -/ -theorem of_right {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} +a commuting square on the right, the right square is a pushout square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ - h₁₂ -> X₁₃ +| | | +v₁₁ v₁₂ v₁₃ +↓ ↓ ↓ +X₂₁ - h₂₁ -> X₂₂ - h₂₂ -> X₂₃ +``` +-/ +theorem of_left {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} {h₂₂ : X₂₂ ⟶ X₂₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} (s : IsPushout (h₁₁ ≫ h₁₂) v₁₁ v₁₃ (h₂₁ ≫ h₂₂)) (p : h₁₂ ≫ v₁₃ = v₁₂ ≫ h₂₂) (t : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) : IsPushout h₁₂ v₁₂ v₁₃ h₂₂ := - (of_bot s.flip p.symm t.flip).flip + (of_top s.flip p.symm t.flip).flip theorem paste_vert_iff {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) (e : h₂₁ ≫ v₂₂ = v₂₁ ≫ h₃₁) : IsPushout h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁ ↔ IsPushout h₂₁ v₂₁ v₂₂ h₃₁ := - ⟨fun h => h.of_bot e s, s.paste_vert⟩ + ⟨fun h => h.of_top e s, s.paste_vert⟩ theorem paste_horiz_iff {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} {h₂₂ : X₂₂ ⟶ X₂₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} (s : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) (e : h₁₂ ≫ v₁₃ = v₁₂ ≫ h₂₂) : IsPushout (h₁₁ ≫ h₁₂) v₁₁ v₁₃ (h₂₁ ≫ h₂₂) ↔ IsPushout h₁₂ v₁₂ v₁₃ h₂₂ := - ⟨fun h => h.of_right e s, s.paste_horiz⟩ + ⟨fun h => h.of_left e s, s.paste_horiz⟩ + +/-- Variant of `IsPushout.of_top` where `v₂₂` is induced from a morphism `v₁₃ : X₁₂ ⟶ X₃₂`, and +the universal property of the top square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ +| | +v₁₁ v₁₂ +↓ ↓ +X₂₁ - h₂₁ -> X₂₂ +| | +v₂₁ v₂₂ +↓ ↓ +X₃₁ - h₃₁ -> X₃₂ +``` +-/ +theorem of_top' {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₂₁ : X₂₁ ⟶ X₂₂} {h₃₁ : X₃₁ ⟶ X₃₂} + {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₂ ⟶ X₃₂} {v₂₁ : X₂₁ ⟶ X₃₁} + (s : IsPushout h₁₁ (v₁₁ ≫ v₂₁) v₁₃ h₃₁) (t : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) : + IsPushout h₂₁ v₂₁ (t.desc v₁₃ (v₂₁ ≫ h₃₁) (by rw [s.w, Category.assoc])) h₃₁ := + of_top ((t.inl_desc _ _ _).symm ▸ s) (t.inr_desc _ _ _) t + +/-- Variant of `IsPushout.of_right` where `h₂₂` is induced from a morphism `h₂₃ : X₂₁ ⟶ X₂₃`, and +the universal property of the left square. + +The objects in the statement fit into the following diagram: +``` +X₁₁ - h₁₁ -> X₁₂ - h₁₂ -> X₁₃ +| | | +v₁₁ v₁₂ v₁₃ +↓ ↓ ↓ +X₂₁ - h₂₁ -> X₂₂ - h₂₂ -> X₂₃ +``` +-/ +theorem of_left' {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} {h₁₁ : X₁₁ ⟶ X₁₂} {h₁₂ : X₁₂ ⟶ X₁₃} {h₂₁ : X₂₁ ⟶ X₂₂} + {h₂₃ : X₂₁ ⟶ X₂₃} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₁₃ : X₁₃ ⟶ X₂₃} + (s : IsPushout (h₁₁ ≫ h₁₂) v₁₁ v₁₃ h₂₃) (t : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) : + IsPushout h₁₂ v₁₂ v₁₃ (t.desc (h₁₂ ≫ v₁₃) h₂₃ (by rw [← Category.assoc, s.w])) := + of_left ((t.inr_desc _ _ _).symm ▸ s) (by simp only [inl_desc]) t section @@ -740,7 +1043,7 @@ theorem of_has_biproduct (X Y : C) [HasBinaryBiproduct X Y] : theorem inl_snd' {b : BinaryBicone X Y} (h : b.IsBilimit) : IsPushout b.inl (0 : X ⟶ 0) b.snd (0 : 0 ⟶ Y) := by apply flip - refine of_right ?_ (by simp) (of_isBilimit h) + refine of_left ?_ (by simp) (of_isBilimit h) simp /-- The square @@ -760,7 +1063,7 @@ theorem inl_snd (X Y : C) [HasBinaryBiproduct X Y] : theorem inr_fst' {b : BinaryBicone X Y} (h : b.IsBilimit) : IsPushout b.inr (0 : Y ⟶ 0) b.fst (0 : 0 ⟶ X) := by - refine of_bot ?_ (by simp) (of_isBilimit h) + refine of_top ?_ (by simp) (of_isBilimit h) simp /-- The square @@ -780,7 +1083,7 @@ theorem inr_fst (X Y : C) [HasBinaryBiproduct X Y] : theorem of_is_bilimit' {b : BinaryBicone X Y} (h : b.IsBilimit) : IsPushout b.fst b.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := by - refine IsPushout.of_right ?_ (by simp) (IsPushout.inl_snd' h) + refine IsPushout.of_left ?_ (by simp) (IsPushout.inl_snd' h) simp theorem of_hasBinaryBiproduct (X Y : C) [HasBinaryBiproduct X Y] : @@ -826,6 +1129,30 @@ lemma of_id_fst : IsPushout (𝟙 _) f f (𝟙 _) := IsPushout.of_horiz_isIso lemma of_id_snd : IsPushout f (𝟙 _) (𝟙 _) f := IsPushout.of_vert_isIso ⟨by simp⟩ +/-- The following diagram is a pullback +``` +X --f--> Z +| | +id id +v v +X --f--> Z +``` +-/ +lemma id_vert (f : X ⟶ Z) : IsPushout f (𝟙 X) (𝟙 Z) f := + of_vert_isIso ⟨by simp only [Category.id_comp, Category.comp_id]⟩ + +/-- The following diagram is a pullback +``` +X --id--> X +| | +f f +v v +Z --id--> Z +``` +-/ +lemma id_horiz (f : X ⟶ Z) : IsPushout (𝟙 X) f f (𝟙 Z) := + of_horiz_isIso ⟨by simp only [Category.id_comp, Category.comp_id]⟩ + end IsPushout section Equalizer diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean index 5cfcb54dcfc39..6e410d8ba12e2 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean @@ -137,12 +137,11 @@ instance mono_pullback_to_prod {C : Type*} [Category C] {X Y Z : C} (f : X ⟶ Z · simpa using congrArg (fun f => f ≫ prod.fst) h · simpa using congrArg (fun f => f ≫ prod.snd) h⟩ - /-- The pullback of `f, g` is also the pullback of `f ≫ i, g ≫ i` for any mono `i`. -/ noncomputable def pullbackIsPullbackOfCompMono (f : X ⟶ W) (g : Y ⟶ W) (i : W ⟶ Z) [Mono i] - [HasPullback f g] : IsLimit (PullbackCone.mk (pullback.fst _ _) (pullback.snd _ _) + [HasPullback f g] : IsLimit (PullbackCone.mk (pullback.fst f g) (pullback.snd f g) -- Porting note: following used to be _ - (show (pullback.fst _ _) ≫ f ≫ i = (pullback.snd _ _) ≫ g ≫ i from by + (show (pullback.fst f g) ≫ f ≫ i = (pullback.snd f g) ≫ g ≫ i from by simp only [← Category.assoc]; rw [cancel_mono]; apply pullback.condition)) := PullbackCone.isLimitOfCompMono f g i _ (limit.isLimit (cospan f g)) @@ -314,8 +313,8 @@ instance epi_coprod_to_pushout {C : Type*} [Category C] {X Y Z : C} (f : X ⟶ Y /-- The pushout of `f, g` is also the pullback of `h ≫ f, h ≫ g` for any epi `h`. -/ noncomputable def pushoutIsPushoutOfEpiComp (f : X ⟶ Y) (g : X ⟶ Z) (h : W ⟶ X) [Epi h] - [HasPushout f g] : IsColimit (PushoutCocone.mk (pushout.inl _ _) (pushout.inr _ _) - (show (h ≫ f) ≫ pushout.inl _ _ = (h ≫ g) ≫ pushout.inr _ _ from by + [HasPushout f g] : IsColimit (PushoutCocone.mk (pushout.inl f g) (pushout.inr f g) + (show (h ≫ f) ≫ pushout.inl f g = (h ≫ g) ≫ pushout.inr f g from by simp only [Category.assoc]; rw [cancel_epi]; exact pushout.condition)) := PushoutCocone.isColimitOfEpiComp f g h _ (colimit.isColimit (span f g)) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean index 3c1f84a2a87d9..f56d21c18f129 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/PullbackCone.lean @@ -51,8 +51,8 @@ Various ways of constructing pullback cones: Interaction with `IsLimit`: * `PullbackCone.isLimitAux` and `PullbackCone.isLimitAux'` provide two convenient ways to show that a given `PullbackCone` is a limit cone. -* `PullbackCone.isLimit.mk` provides a convenient way to show that a `PullbackCone` constructed using - `PullbackCone.mk` is a limit cone. +* `PullbackCone.isLimit.mk` provides a convenient way to show that a `PullbackCone` constructed + using `PullbackCone.mk` is a limit cone. * `PullbackCone.IsLimit.lift` and `PullbackCone.IsLimit.lift'` provides convenient ways for constructing the morphisms to the point of a limit `PullbackCone` from the universal property. * `PullbackCone.IsLimit.hom_ext` provides a convenient way to show that two morphisms to the point diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean index 4abb9dfacd9f6..6aa46f0c5a128 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean @@ -610,4 +610,18 @@ lemma IsColimit.isZero_pt {c : Cocone F} (hc : IsColimit c) (hF : IsZero F) : Is end +section + +variable [HasZeroMorphisms C] + +lemma IsTerminal.isZero {X : C} (hX : IsTerminal X) : IsZero X := by + rw [IsZero.iff_id_eq_zero] + apply hX.hom_ext + +lemma IsInitial.isZero {X : C} (hX : IsInitial X) : IsZero X := by + rw [IsZero.iff_id_eq_zero] + apply hX.hom_ext + +end + end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index cb51a7bb1fe5d..d1497d0eaa18d 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -663,7 +663,10 @@ lemma surjective_π_app_zero_of_surjective_map_aux : /-- Given surjections `⋯ ⟶ Xₙ₊₁ ⟶ Xₙ ⟶ ⋯ ⟶ X₀`, the projection map `lim Xₙ ⟶ X₀` is surjective. -/ -lemma surjective_π_app_zero_of_surjective_map : Function.Surjective (c.π.app ⟨0⟩) := by +lemma surjective_π_app_zero_of_surjective_map + (hc : IsLimit c) + (hF : ∀ n, Function.Surjective (F.map (homOfLE (Nat.le_succ n)).op)) : + Function.Surjective (c.π.app ⟨0⟩) := by let i := hc.conePointUniqueUpToIso (limitConeIsLimit F) have : c.π.app ⟨0⟩ = i.hom ≫ (limitCone F).π.app ⟨0⟩ := by simp [i] rw [this] diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index b5d731d568e9a..783314187e22d 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -361,7 +361,8 @@ theorem IsUniversalColimit.map_reflective · rw [Gl.map_comp, hα'', Category.assoc, hc''] dsimp [β] rw [Category.comp_id, Category.assoc] - have : cf.hom ≫ (PreservesPullback.iso _ _ _).hom ≫ pullback.fst _ _ ≫ adj.counit.app _ = 𝟙 _ := by + have : + cf.hom ≫ (PreservesPullback.iso _ _ _).hom ≫ pullback.fst _ _ ≫ adj.counit.app _ = 𝟙 _ := by simp only [IsIso.inv_hom_id, Iso.inv_hom_id_assoc, Category.assoc, pullback.lift_fst_assoc] have : IsIso cf := by apply @Cocones.cocone_iso_of_hom_iso (i := ?_) diff --git a/Mathlib/CategoryTheory/Limits/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Yoneda.lean index fa3b998c61800..07b2b79a18c53 100644 --- a/Mathlib/CategoryTheory/Limits/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Yoneda.lean @@ -107,6 +107,15 @@ def yonedaJointlyReflectsLimits (F : J ⥤ Cᵒᵖ) (c : Cone F) dsimp [Types.isLimitEquivSections, Types.sectionOfCone] rw [eq, Category.comp_id, ← hm, unop_comp]) +/-- A cocone is colimit iff it becomes limit after the +application of `yoneda.obj X` for all `X : C`. -/ +noncomputable def Limits.Cocone.isColimitYonedaEquiv {F : J ⥤ C} (c : Cocone F) : + IsColimit c ≃ ∀ (X : C), IsLimit ((yoneda.obj X).mapCone c.op) where + toFun h X := isLimitOfPreserves _ h.op + invFun h := IsLimit.unop (yonedaJointlyReflectsLimits _ _ h) + left_inv _ := Subsingleton.elim _ _ + right_inv _ := by ext; apply Subsingleton.elim + /-- The cone of `F` corresponding to an element in `(F ⋙ coyoneda.obj X).sections`. -/ @[simps] def Limits.coneOfSectionCompCoyoneda (F : J ⥤ C) (X : Cᵒᵖ) @@ -140,6 +149,14 @@ def coyonedaJointlyReflectsLimits (F : J ⥤ C) (c : Cone F) dsimp at eq rw [eq, Category.id_comp, ← hm] +/-- A cone is limit iff it is so after the application of `coyoneda.obj X` for all `X : Cᵒᵖ`. -/ +noncomputable def Limits.Cone.isLimitCoyonedaEquiv {F : J ⥤ C} (c : Cone F) : + IsLimit c ≃ ∀ (X : Cᵒᵖ), IsLimit ((coyoneda.obj X).mapCone c) where + toFun h X := isLimitOfPreserves _ h + invFun h := coyonedaJointlyReflectsLimits _ _ h + left_inv _ := Subsingleton.elim _ _ + right_inv _ := by ext; apply Subsingleton.elim + end /-- The yoneda embedding `yoneda.obj X : Cᵒᵖ ⥤ Type v` for `X : C` preserves limits. -/ diff --git a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean index 96a8ee1ea2847..ebb9d0bb8ba30 100644 --- a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean +++ b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean @@ -103,7 +103,7 @@ variable [CatCommSq Φ.functor L₁ L₂ G] /-- If a localizer morphism induces an equivalence on some choice of localized categories, it will be so for any choice of localized categoriees. -/ -lemma isEquivalence_imp [G.IsEquivalence] : G'.IsEquivalence := by +lemma isEquivalence_imp [G.IsEquivalence] : G'.IsEquivalence := let E₁ := Localization.uniq L₁ L₁' W₁ let E₂ := Localization.uniq L₂ L₂' W₂ let e : L₁ ⋙ G ⋙ E₂.functor ≅ L₁ ⋙ E₁.functor ⋙ G' := @@ -118,7 +118,7 @@ lemma isEquivalence_imp [G.IsEquivalence] : G'.IsEquivalence := by isoWhiskerRight (compUniqFunctor L₁ L₁' W₁).symm G' ≪≫ Functor.associator _ _ _ have := Functor.isEquivalence_of_iso (liftNatIso L₁ W₁ _ _ (G ⋙ E₂.functor) (E₁.functor ⋙ G') e) - exact Functor.isEquivalence_of_comp_left E₁.functor G' + Functor.isEquivalence_of_comp_left E₁.functor G' lemma isEquivalence_iff : G.IsEquivalence ↔ G'.IsEquivalence := ⟨fun _ => Φ.isEquivalence_imp L₁ L₂ G L₁' L₂' G', diff --git a/Mathlib/CategoryTheory/Localization/Prod.lean b/Mathlib/CategoryTheory/Localization/Prod.lean index 8129a05d899a6..68e1839c40576 100644 --- a/Mathlib/CategoryTheory/Localization/Prod.lean +++ b/Mathlib/CategoryTheory/Localization/Prod.lean @@ -38,17 +38,18 @@ namespace Localization namespace StrictUniversalPropertyFixedTarget -variable {E : Type u₅} [Category.{v₅} E] - (F : C₁ × C₂ ⥤ E) (hF : (W₁.prod W₂).IsInvertedBy F) +variable {E : Type u₅} [Category.{v₅} E] (F : C₁ × C₂ ⥤ E) /-- Auxiliary definition for `prodLift`. -/ -noncomputable def prodLift₁ : +noncomputable def prodLift₁ (hF : (W₁.prod W₂).IsInvertedBy F) : W₁.Localization ⥤ C₂ ⥤ E := Construction.lift (curry.obj F) (fun _ _ f₁ hf₁ => by haveI : ∀ (X₂ : C₂), IsIso (((curry.obj F).map f₁).app X₂) := fun X₂ => hF _ ⟨hf₁, MorphismProperty.id_mem _ _⟩ apply NatIso.isIso_of_isIso_app) +variable (hF : (W₁.prod W₂).IsInvertedBy F) + lemma prod_fac₁ : W₁.Q ⋙ prodLift₁ F hF = curry.obj F := Construction.fac _ _ diff --git a/Mathlib/CategoryTheory/Monad/Equalizer.lean b/Mathlib/CategoryTheory/Monad/Equalizer.lean index ebe9532b85a62..a898227f6d7c1 100644 --- a/Mathlib/CategoryTheory/Monad/Equalizer.lean +++ b/Mathlib/CategoryTheory/Monad/Equalizer.lean @@ -47,7 +47,8 @@ def CofreeEqualizer.topMap : (Comonad.cofree T).obj X.A ⟶ (Comonad.cofree T). /-- The bottom map in the equalizer diagram we will construct. -/ @[simps] -def CofreeEqualizer.bottomMap : (Comonad.cofree T).obj X.A ⟶ (Comonad.cofree T).obj (T.obj X.A) where +def CofreeEqualizer.bottomMap : + (Comonad.cofree T).obj X.A ⟶ (Comonad.cofree T).obj (T.obj X.A) where f := T.δ.app X.A h := T.coassoc X.A diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index 235ef39fd8215..f0ba170eb47a1 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.CategoryTheory.Monoidal.Free.Coherence import Mathlib.CategoryTheory.Monoidal.Discrete import Mathlib.CategoryTheory.Monoidal.NaturalTransformation import Mathlib.CategoryTheory.Monoidal.Opposite diff --git a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean index e900909325f36..430680b7fe1cf 100644 --- a/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean +++ b/Mathlib/CategoryTheory/Monoidal/CoherenceLemmas.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Jendrusch, Scott Morrison, Bhavik Mehta, Jakob von Raumer -/ import Mathlib.Tactic.CategoryTheory.Coherence -import Mathlib.CategoryTheory.Monoidal.Free.Coherence /-! # Lemmas which are consequences of monoidal coherence diff --git a/Mathlib/CategoryTheory/Monoidal/Opposite.lean b/Mathlib/CategoryTheory/Monoidal/Opposite.lean index 25d404670f914..cecef27b955d5 100644 --- a/Mathlib/CategoryTheory/Monoidal/Opposite.lean +++ b/Mathlib/CategoryTheory/Monoidal/Opposite.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.CategoryTheory.Monoidal.Free.Coherence import Mathlib.Tactic.CategoryTheory.Coherence /-! diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index c9d09acb7b3a9..9a2f2fa079c42 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 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.Monoidal.Free.Coherence import Mathlib.Tactic.CategoryTheory.Coherence import Mathlib.CategoryTheory.Closed.Monoidal import Mathlib.Tactic.ApplyFun diff --git a/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean b/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean index ee138b0e38a78..633d6872c7ad0 100644 --- a/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean +++ b/Mathlib/CategoryTheory/Monoidal/Types/Coyoneda.lean @@ -25,14 +25,7 @@ open Opposite open MonoidalCategory --- Porting note: made it noncomputable. --- `failed to compile definition, consider marking it as 'noncomputable' because it` --- `depends on 'CategoryTheory.typesMonoidal', and it does not have executable code` --- I don't know if that is a problem, might need to change it back in the future, but --- if so it might be better to fix then instead of at the moment of porting. - /-- `(𝟙_ C ⟶ -)` is a lax monoidal functor to `Type`. -/ -noncomputable def coyonedaTensorUnit (C : Type u) [Category.{v} C] [MonoidalCategory C] : LaxMonoidalFunctor C (Type v) := .ofTensorHom (F := coyoneda.obj (op (𝟙_ C))) diff --git a/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean b/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean index df95bec3470bc..61a44cda69610 100644 --- a/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean +++ b/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean @@ -84,7 +84,7 @@ instance Monad.algebraPreadditive : Preadditive (Monad.Algebra T) where zsmul_neg' := by intros ext - simp only [negSucc_zsmul, neg_inj, nsmul_eq_smul_cast ℤ] + simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] add_left_neg := by intros ext @@ -165,7 +165,7 @@ instance Comonad.coalgebraPreadditive : Preadditive (Comonad.Coalgebra U) where zsmul_neg' := by intros ext - simp only [negSucc_zsmul, neg_inj, nsmul_eq_smul_cast ℤ] + simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] add_left_neg := by intros ext diff --git a/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean b/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean index f86c3b692d9bf..c9856b25bf323 100644 --- a/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean @@ -84,7 +84,7 @@ instance Endofunctor.algebraPreadditive : Preadditive (Endofunctor.Algebra F) wh zsmul_neg' := by intros apply Algebra.Hom.ext - simp only [negSucc_zsmul, neg_inj, nsmul_eq_smul_cast ℤ] + simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] add_left_neg := by intros apply Algebra.Hom.ext @@ -162,7 +162,7 @@ instance Endofunctor.coalgebraPreadditive : Preadditive (Endofunctor.Coalgebra F zsmul_neg' := by intros apply Coalgebra.Hom.ext - simp only [negSucc_zsmul, neg_inj, nsmul_eq_smul_cast ℤ] + simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] add_left_neg := by intros apply Coalgebra.Hom.ext diff --git a/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean b/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean index f4978f59a06ff..4ff63ae23affb 100644 --- a/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean +++ b/Mathlib/CategoryTheory/Preadditive/HomOrthogonal.lean @@ -130,7 +130,7 @@ theorem matrixDecomposition_id (o : HomOrthogonal s) {α : Type} [Finite α] {f split_ifs with h · cases h simp - · simp at h + · simp only [Subtype.mk.injEq] at h -- Porting note: used to be `convert comp_zero`, but that does not work anymore have : biproduct.ι (fun a ↦ s (f a)) a ≫ biproduct.π (fun b ↦ s (f b)) b = 0 := by simpa using biproduct.ι_π_ne _ (Ne.symm h) diff --git a/Mathlib/CategoryTheory/Quotient.lean b/Mathlib/CategoryTheory/Quotient.lean index 780ceeefbc696..5a1f9831b2055 100644 --- a/Mathlib/CategoryTheory/Quotient.lean +++ b/Mathlib/CategoryTheory/Quotient.lean @@ -141,10 +141,9 @@ theorem functor_map_eq_iff [h : Congruence r] {X Y : C} (f f' : X ⟶ Y) : simpa only [compClosure_eq_self r] using h.equivalence variable {D : Type _} [Category D] (F : C ⥤ D) - (H : ∀ (x y : C) (f₁ f₂ : x ⟶ y), r f₁ f₂ → F.map f₁ = F.map f₂) /-- The induced functor on the quotient category. -/ -def lift : Quotient r ⥤ D where +def lift (H : ∀ (x y : C) (f₁ f₂ : x ⟶ y), r f₁ f₂ → F.map f₁ = F.map f₂) : Quotient r ⥤ D where obj a := F.obj a.as map := @fun a b hf ↦ Quot.liftOn hf (fun f ↦ F.map f) @@ -156,6 +155,8 @@ def lift : Quotient r ⥤ D where rintro a b c ⟨f⟩ ⟨g⟩ exact F.map_comp f g +variable (H : ∀ (x y : C) (f₁ f₂ : x ⟶ y), r f₁ f₂ → F.map f₁ = F.map f₂) + theorem lift_spec : functor r ⋙ lift r F H = F := by apply Functor.ext; rotate_left · rintro X diff --git a/Mathlib/CategoryTheory/Quotient/Linear.lean b/Mathlib/CategoryTheory/Quotient/Linear.lean index 95b4414887b42..bdb9c5b50e4a7 100644 --- a/Mathlib/CategoryTheory/Quotient/Linear.lean +++ b/Mathlib/CategoryTheory/Quotient/Linear.lean @@ -25,12 +25,12 @@ namespace Quotient variable {R C : Type*} [Semiring R] [Category C] [Preadditive C] [Linear R C] (r : HomRel C) [Congruence r] - (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) namespace Linear /-- The scalar multiplications on morphisms in `Quotient R`. -/ -def smul (X Y : Quotient r) : SMul R (X ⟶ Y) where +def smul (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) + (X Y : Quotient r) : SMul R (X ⟶ Y) where smul a := Quot.lift (fun g => Quot.mk _ (a • g)) (fun f₁ f₂ h₁₂ => by dsimp simp only [compClosure_eq_self] at h₁₂ @@ -39,14 +39,16 @@ def smul (X Y : Quotient r) : SMul R (X ⟶ Y) where exact hr _ _ _ h₁₂) @[simp] -lemma smul_eq (a : R) {X Y : C} (f : X ⟶ Y) : +lemma smul_eq (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) + (a : R) {X Y : C} (f : X ⟶ Y) : letI := smul r hr a • (functor r).map f = (functor r).map (a • f) := rfl -variable [Preadditive (Quotient r)] [(functor r).Additive] /-- Auxiliary definition for `Quotient.Linear.module`. -/ -def module' (X Y : C) : Module R ((functor r).obj X ⟶ (functor r).obj Y) := +def module' (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) + [Preadditive (Quotient r)] [(functor r).Additive] (X Y : C) : + Module R ((functor r).obj X ⟶ (functor r).obj Y) := letI := smul r hr ((functor r).obj X) ((functor r).obj Y) { smul_zero := fun a => by dsimp @@ -74,18 +76,20 @@ def module' (X Y : C) : Module R ((functor r).obj X ⟶ (functor r).obj Y) := rw [add_smul, Functor.map_add] } /-- Auxiliary definition for `Quotient.linear`. -/ -def module (X Y : Quotient r) : Module R (X ⟶ Y) := module' r hr X.as Y.as +def module (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) + [Preadditive (Quotient r)] [(functor r).Additive] (X Y : Quotient r) : + Module R (X ⟶ Y) := module' r hr X.as Y.as end Linear variable (R) -variable [Preadditive (Quotient r)] [(functor r).Additive] /-- Assuming `Quotient r` has already been endowed with a preadditive category structure such that `functor r : C ⥤ Quotient r` is additive, and that `C` has a `R`-linear category structure compatible with `r`, this is the induced `R`-linear category structure on `Quotient r`. -/ -def linear : Linear R (Quotient r) := by +def linear (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) + [Preadditive (Quotient r)] [(functor r).Additive] : Linear R (Quotient r) := by letI := Linear.module r hr exact { smul_comp := by @@ -101,7 +105,10 @@ def linear : Linear R (Quotient r) := by rw [Linear.smul_eq, ← Functor.map_comp, ← Functor.map_comp, Linear.smul_eq, Linear.comp_smul] } -instance linear_functor : letI := linear R r hr; Functor.Linear R (functor r) := by +instance linear_functor + (hr : ∀ (a : R) ⦃X Y : C⦄ (f₁ f₂ : X ⟶ Y) (_ : r f₁ f₂), r (a • f₁) (a • f₂)) + [Preadditive (Quotient r)] [(functor r).Additive] : + letI := linear R r hr; Functor.Linear R (functor r) := by letI := linear R r hr; exact { } end Quotient diff --git a/Mathlib/CategoryTheory/Quotient/Preadditive.lean b/Mathlib/CategoryTheory/Quotient/Preadditive.lean index c7b2d6abfc81a..80e1204fe4404 100644 --- a/Mathlib/CategoryTheory/Quotient/Preadditive.lean +++ b/Mathlib/CategoryTheory/Quotient/Preadditive.lean @@ -20,13 +20,13 @@ namespace CategoryTheory namespace Quotient variable {C : Type _} [Category C] [Preadditive C] (r : HomRel C) [Congruence r] - (hr : ∀ ⦃X Y : C⦄ (f₁ f₂ g₁ g₂ : X ⟶ Y) (_ : r f₁ f₂) (_ : r g₁ g₂), r (f₁ + g₁) (f₂ + g₂)) namespace Preadditive /-- The addition on the morphisms in the category `Quotient r` when `r` is compatible with the addition. -/ -def add {X Y : Quotient r} (f g : X ⟶ Y) : X ⟶ Y := +def add (hr : ∀ ⦃X Y : C⦄ (f₁ f₂ g₁ g₂ : X ⟶ Y) (_ : r f₁ f₂) (_ : r g₁ g₂), r (f₁ + g₁) (f₂ + g₂)) + {X Y : Quotient r} (f g : X ⟶ Y) : X ⟶ Y := Quot.liftOn₂ f g (fun a b => Quot.mk _ (a + b)) (fun f g₁ g₂ h₁₂ => by simp only [compClosure_iff_self] at h₁₂ @@ -39,7 +39,8 @@ def add {X Y : Quotient r} (f g : X ⟶ Y) : X ⟶ Y := /-- The negation on the morphisms in the category `Quotient r` when `r` is compatible with the addition. -/ -def neg {X Y : Quotient r} (f : X ⟶ Y) : X ⟶ Y := +def neg (hr : ∀ ⦃X Y : C⦄ (f₁ f₂ g₁ g₂ : X ⟶ Y) (_ : r f₁ f₂) (_ : r g₁ g₂), r (f₁ + g₁) (f₂ + g₂)) + {X Y : Quotient r} (f : X ⟶ Y) : X ⟶ Y := Quot.liftOn f (fun a => Quot.mk _ (-a)) (fun f g => by intro hfg @@ -52,7 +53,9 @@ end Preadditive /-- The preadditive structure on the category `Quotient r` when `r` is compatible with the addition. -/ -def preadditive : Preadditive (Quotient r) where +def preadditive + (hr : ∀ ⦃X Y : C⦄ (f₁ f₂ g₁ g₂ : X ⟶ Y) (_ : r f₁ f₂) (_ : r g₁ g₂), r (f₁ + g₁) (f₂ + g₂)) : + Preadditive (Quotient r) where homGroup P Q := let iZ : Zero (P ⟶ Q) := { zero := Quot.mk _ 0 } @@ -75,7 +78,8 @@ def preadditive : Preadditive (Quotient r) where rintro _ _ _ ⟨_⟩ ⟨_⟩ ⟨_⟩ exact congr_arg (functor r).map (by apply Preadditive.comp_add) -lemma functor_additive : +lemma functor_additive + (hr : ∀ ⦃X Y : C⦄ (f₁ f₂ g₁ g₂ : X ⟶ Y) (_ : r f₁ f₂) (_ : r g₁ g₂), r (f₁ + g₁) (f₂ + g₂)) : letI := preadditive r hr (functor r).Additive := letI := preadditive r hr diff --git a/Mathlib/CategoryTheory/Sites/Canonical.lean b/Mathlib/CategoryTheory/Sites/Canonical.lean index 76e3700b4ff2d..756cf0d75719f 100644 --- a/Mathlib/CategoryTheory/Sites/Canonical.lean +++ b/Mathlib/CategoryTheory/Sites/Canonical.lean @@ -233,19 +233,20 @@ theorem isSheaf_of_representable {J : GrothendieckTopology C} (hJ : Subcanonical Presieve.isSheaf_of_le _ hJ (Sheaf.isSheaf_of_representable P) variable {J} -variable (hJ : Subcanonical J) /-- If `J` is subcanonical, we obtain a "Yoneda" functor from the defining site into the sheaf category. -/ @[simps] -def yoneda : C ⥤ Sheaf J (Type v) where +def yoneda (hJ : Subcanonical J) : C ⥤ Sheaf J (Type v) where obj X := ⟨CategoryTheory.yoneda.obj X, by rw [isSheaf_iff_isSheaf_of_type] apply hJ.isSheaf_of_representable⟩ map f := ⟨CategoryTheory.yoneda.map f⟩ +variable (hJ : Subcanonical J) + /-- The yoneda embedding into the presheaf category factors through the one to the sheaf category. diff --git a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean index 00595c51688e7..ef1204a732163 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean @@ -96,8 +96,7 @@ lemma extensiveTopology.presheafIsLocallySurjective_iff [FinitaryPreExtensive C] obtain ⟨x, hx⟩ := h X y convert (extensiveTopology C).top_mem' X rw [← Sieve.id_mem_iff_eq_top] - simp [Presheaf.imageSieve] - exact ⟨x, hx⟩ + simpa [Presheaf.imageSieve] using ⟨x, hx⟩ lemma extensiveTopology.isLocallySurjective_iff [FinitaryExtensive C] {F G : Sheaf (extensiveTopology C) D} (f : F ⟶ G) diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index bccb645c7cdd2..330fcdfb5f49f 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -55,7 +55,7 @@ instance : F.IsCoverDense (coherentTopology _) := by theorem exists_effectiveEpiFamily_iff_mem_induced (X : C) (S : Sieve X) : (∃ (α : Type) (_ : Finite α) (Y : α → C) (π : (a : α) → (Y a ⟶ X)), EffectiveEpiFamily Y π ∧ (∀ a : α, (S.arrows) (π a)) ) ↔ - (S ∈ F.inducedTopologyOfIsCoverDense (coherentTopology _) X) := by + (S ∈ F.inducedTopology (coherentTopology _) X) := by refine ⟨fun ⟨α, _, Y, π, ⟨H₁, H₂⟩⟩ ↦ ?_, fun hS ↦ ?_⟩ · apply (mem_sieves_iff_hasEffectiveEpiFamily (Sieve.functorPushforward _ S)).mpr refine ⟨α, inferInstance, fun i => F.obj (Y i), @@ -77,25 +77,19 @@ theorem exists_effectiveEpiFamily_iff_mem_induced (X : C) (S : Sieve X) : lemma eq_induced : haveI := F.reflects_precoherent coherentTopology C = - F.inducedTopologyOfIsCoverDense (coherentTopology _) := by + F.inducedTopology (coherentTopology _) := by ext X S have := F.reflects_precoherent rw [← exists_effectiveEpiFamily_iff_mem_induced F X] rw [← coherentTopology.mem_sieves_iff_hasEffectiveEpiFamily S] -lemma coverPreserving : haveI := F.reflects_precoherent - CoverPreserving (coherentTopology _) (coherentTopology _) F := by - rw [eq_induced F] - apply LocallyCoverDense.inducedTopology_coverPreserving - -instance coverLifting : haveI := F.reflects_precoherent - F.IsCocontinuous (coherentTopology _) (coherentTopology _) := by - rw [eq_induced F] - apply LocallyCoverDense.inducedTopology_isCocontinuous +instance : haveI := F.reflects_precoherent; + F.IsDenseSubsite (coherentTopology C) (coherentTopology D) where + functorPushforward_mem_iff := by simp_rw [eq_induced F]; rfl -instance isContinuous : haveI := F.reflects_precoherent - F.IsContinuous (coherentTopology _) (coherentTopology _) := - Functor.IsCoverDense.isContinuous _ _ _ (coverPreserving F) +lemma coverPreserving : haveI := F.reflects_precoherent + CoverPreserving (coherentTopology _) (coherentTopology _) F := + IsDenseSubsite.coverPreserving _ _ _ section SheafEquiv @@ -114,7 +108,7 @@ noncomputable def equivalence (A : Type u₃) [Category.{v₃} A] [∀ X, HasLimitsOfShape (StructuredArrow X F.op) A] : haveI := F.reflects_precoherent Sheaf (coherentTopology C) A ≌ Sheaf (coherentTopology D) A := - Functor.IsCoverDense.sheafEquivOfCoverPreservingCoverLifting F _ _ _ + Functor.IsDenseSubsite.sheafEquiv F _ _ _ end SheafEquiv @@ -138,7 +132,7 @@ def equivalence' (A : Type u₃) [Category.{v₃} A] [∀ X, HasLimitsOfShape (StructuredArrow X F.op) A] : haveI := F.reflects_precoherent Sheaf (coherentTopology C) A ≌ Sheaf (coherentTopology D) A := - Functor.IsCoverDense.sheafEquivOfCoverPreservingCoverLifting F _ _ _ + Functor.IsDenseSubsite.sheafEquiv F _ _ _ end RegularExtensive @@ -161,7 +155,7 @@ instance : F.IsCoverDense (regularTopology _) := by theorem exists_effectiveEpi_iff_mem_induced (X : C) (S : Sieve X) : (∃ (Y : C) (π : Y ⟶ X), EffectiveEpi π ∧ S.arrows π) ↔ - (S ∈ F.inducedTopologyOfIsCoverDense (regularTopology _) X) := by + (S ∈ F.inducedTopology (regularTopology _) X) := by refine ⟨fun ⟨Y, π, ⟨H₁, H₂⟩⟩ ↦ ?_, fun hS ↦ ?_⟩ · apply (mem_sieves_iff_hasEffectiveEpi (Sieve.functorPushforward _ S)).mpr refine ⟨F.obj Y, F.map π, ⟨?_, Sieve.image_mem_functorPushforward F S H₂⟩⟩ @@ -179,25 +173,19 @@ theorem exists_effectiveEpi_iff_mem_induced (X : C) (S : Sieve X) : lemma eq_induced : haveI := F.reflects_preregular regularTopology C = - F.inducedTopologyOfIsCoverDense (regularTopology _) := by + F.inducedTopology (regularTopology _) := by ext X S have := F.reflects_preregular rw [← exists_effectiveEpi_iff_mem_induced F X] rw [← mem_sieves_iff_hasEffectiveEpi S] -lemma coverPreserving : haveI := F.reflects_preregular - CoverPreserving (regularTopology _) (regularTopology _) F := by - rw [eq_induced F] - apply LocallyCoverDense.inducedTopology_coverPreserving - -instance coverLifting : haveI := F.reflects_preregular - F.IsCocontinuous (regularTopology _) (regularTopology _) := by - rw [eq_induced F] - apply LocallyCoverDense.inducedTopology_isCocontinuous +instance : haveI := F.reflects_preregular; + F.IsDenseSubsite (regularTopology C) (regularTopology D) where + functorPushforward_mem_iff := by simp_rw [eq_induced F]; rfl -instance isContinuous : haveI := F.reflects_preregular - F.IsContinuous (regularTopology _) (regularTopology _) := - Functor.IsCoverDense.isContinuous _ _ _ (coverPreserving F) +lemma coverPreserving : haveI := F.reflects_preregular + CoverPreserving (regularTopology _) (regularTopology _) F := + IsDenseSubsite.coverPreserving _ _ _ section SheafEquiv @@ -216,7 +204,7 @@ noncomputable def equivalence (A : Type u₃) [Category.{v₃} A] [∀ X, HasLimitsOfShape (StructuredArrow X F.op) A] : haveI := F.reflects_preregular Sheaf (regularTopology C) A ≌ Sheaf (regularTopology D) A := - Functor.IsCoverDense.sheafEquivOfCoverPreservingCoverLifting F _ _ _ + Functor.IsDenseSubsite.sheafEquiv F _ _ _ end SheafEquiv diff --git a/Mathlib/CategoryTheory/Sites/Continuous.lean b/Mathlib/CategoryTheory/Sites/Continuous.lean index ab505278e5043..91469d869e53b 100644 --- a/Mathlib/CategoryTheory/Sites/Continuous.lean +++ b/Mathlib/CategoryTheory/Sites/Continuous.lean @@ -151,10 +151,8 @@ lemma isContinuous_comp' {F₁ : C ⥤ D} {F₂ : D ⥤ E} {F₁₂ : C ⥤ E} section -variable [PreservesOneHypercovers.{w} F J K] - [GrothendieckTopology.IsGeneratedByOneHypercovers.{w} J] - lemma op_comp_isSheaf_of_preservesOneHypercovers + [PreservesOneHypercovers.{w} F J K] [GrothendieckTopology.IsGeneratedByOneHypercovers.{w} J] (P : Dᵒᵖ ⥤ A) (hP : Presheaf.IsSheaf K P) : Presheaf.IsSheaf J (F.op ⋙ P) := by rw [Presheaf.isSheaf_iff_of_isGeneratedByOneHypercovers.{w}] @@ -162,7 +160,8 @@ lemma op_comp_isSheaf_of_preservesOneHypercovers exact ⟨(E.toPreOneHypercover.isLimitMapMultiforkEquiv F P) ((E.map F K).isLimitMultifork ⟨P, hP⟩)⟩ -lemma isContinuous_of_preservesOneHypercovers : +lemma isContinuous_of_preservesOneHypercovers + [PreservesOneHypercovers.{w} F J K] [GrothendieckTopology.IsGeneratedByOneHypercovers.{w} J] : IsContinuous.{t} F J K where op_comp_isSheafOfTypes := by rintro ⟨P, hP⟩ diff --git a/Mathlib/CategoryTheory/Sites/CoversTop.lean b/Mathlib/CategoryTheory/Sites/CoversTop.lean index ba3231bcbb209..d36281d946589 100644 --- a/Mathlib/CategoryTheory/Sites/CoversTop.lean +++ b/Mathlib/CategoryTheory/Sites/CoversTop.lean @@ -107,22 +107,20 @@ noncomputable def familyOfElements (X : C) : namespace IsCompatible variable {x} -variable (hx : x.IsCompatible) -lemma familyOfElements_apply {X Z : C} (f : Z ⟶ X) (i : I) (φ : Z ⟶ Y i) : +lemma familyOfElements_apply (hx : x.IsCompatible) {X Z : C} (f : Z ⟶ X) (i : I) (φ : Z ⟶ Y i) : familyOfElements x X f ⟨i, ⟨φ⟩⟩ = F.map φ.op (x i) := by apply hx -lemma familyOfElements_isCompatible (X : C) : +lemma familyOfElements_isCompatible (hx : x.IsCompatible) (X : C) : (familyOfElements x X).Compatible := by intro Y₁ Y₂ Z g₁ g₂ f₁ f₂ ⟨i₁, ⟨φ₁⟩⟩ ⟨i₂, ⟨φ₂⟩⟩ _ simpa [hx.familyOfElements_apply f₁ i₁ φ₁, hx.familyOfElements_apply f₂ i₂ φ₂] using hx Z i₁ i₂ (g₁ ≫ φ₁) (g₂ ≫ φ₂) variable {J} -variable (hY : J.CoversTop Y) (hF : IsSheaf J F) -lemma exists_unique_section : +lemma exists_unique_section (hx : x.IsCompatible) (hY : J.CoversTop Y) (hF : IsSheaf J F) : ∃! (s : F.sections), ∀ (i : I), s.1 (Opposite.op (Y i)) = x i := by have H := (isSheaf_iff_isSheaf_of_type _ _).1 hF apply exists_unique_of_exists_of_unique @@ -148,6 +146,8 @@ lemma exists_unique_section : · intro y₁ y₂ hy₁ hy₂ exact hY.sections_ext ⟨F, hF⟩ (fun i => by rw [hy₁, hy₂]) +variable (hx : x.IsCompatible) (hY : J.CoversTop Y) (hF : IsSheaf J F) + /-- The section of a sheaf of types which lifts a compatible family of elements indexed by objects which cover the terminal object. -/ noncomputable def section_ : F.sections := (hx.exists_unique_section hY hF).choose diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean index cc8d05d373b88..5e35b47125044 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean @@ -7,6 +7,7 @@ import Mathlib.CategoryTheory.Sites.Sheaf import Mathlib.CategoryTheory.Sites.CoverLifting import Mathlib.CategoryTheory.Sites.CoverPreserving import Mathlib.CategoryTheory.Adjunction.FullyFaithful +import Mathlib.CategoryTheory.Sites.LocallyFullyFaithful /-! # Dense subsites @@ -14,25 +15,23 @@ import Mathlib.CategoryTheory.Adjunction.FullyFaithful We define `IsCoverDense` functors into sites as functors such that there exists a covering sieve that factors through images of the functor for each object in `D`. -We will primarily consider cover-dense functors that are also full, since this notion is in general -not well-behaved otherwise. Note that https://ncatlab.org/nlab/show/dense+sub-site indeed has a -weaker notion of cover-dense that loosens this requirement, but it would not have all the properties -we would need, and some sheafification would be needed for here and there. - ## Main results -- `CategoryTheory.Functor.IsCoverDense.Types.presheafHom`: If `G : C ⥤ (D, K)` is full +- `CategoryTheory.Functor.IsCoverDense.Types.presheafHom`: If `G : C ⥤ (D, K)` is locally-full and cover-dense, then given any presheaf `ℱ` and sheaf `ℱ'` on `D`, and a morphism `α : G ⋙ ℱ ⟶ G ⋙ ℱ'`, we may glue them together to obtain a morphism of presheaves `ℱ ⟶ ℱ'`. - `CategoryTheory.Functor.IsCoverDense.sheafIso`: If `ℱ` above is a sheaf and `α` is an iso, then the result is also an iso. -- `CategoryTheory.Functor.IsCoverDense.iso_of_restrict_iso`: If `G : C ⥤ (D, K)` is full +- `CategoryTheory.Functor.IsCoverDense.iso_of_restrict_iso`: If `G : C ⥤ (D, K)` is locally-full and cover-dense, then given any sheaves `ℱ, ℱ'` on `D`, and a morphism `α : ℱ ⟶ ℱ'`, then `α` is an iso if `G ⋙ ℱ ⟶ G ⋙ ℱ'` is iso. -- `CategoryTheory.Functor.IsCoverDense.sheafEquivOfCoverPreservingCoverLifting`: - If `G : (C, J) ⥤ (D, K)` is fully-faithful, cover-lifting, cover-preserving, and cover-dense, - then it will induce an equivalence of categories of sheaves valued in a complete category. +- `CategoryTheory.Functor.IsDenseSubsite`: + The functor `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)` if `G` is cover-dense, + locally fully-faithful, and `S` is a cover of `C` iff the image of `S` in `D` is a cover. +- `CategoryTheory.Functor.IsDenseSubsite.sheafEquiv`: + If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, + it induces an equivalence of category of sheaves valued in a category with suitable limits. ## References @@ -124,15 +123,15 @@ theorem ext (ℱ : SheafOfTypes K) (X : D) {s t : ℱ.val.obj (op X)} variable {G} -theorem functorPullback_pushforward_covering [Full G] {X : C} +theorem functorPullback_pushforward_covering [G.IsLocallyFull K] {X : C} (T : K (G.obj X)) : (T.val.functorPullback G).functorPushforward G ∈ K (G.obj X) := by - refine K.superset_covering ?_ (K.bind_covering T.property - fun Y f _ => G.is_cover_of_isCoverDense K Y) - rintro Y _ ⟨Z, _, f, hf, ⟨W, g, f', ⟨rfl⟩⟩, rfl⟩ - use W; use G.preimage (f' ≫ f); use g - constructor - · simpa using T.val.downward_closed hf f' - · simp + refine K.transitive T.2 _ fun Y iYX hiYX ↦ ?_ + apply K.transitive (G.is_cover_of_isCoverDense _ _) _ + rintro W _ ⟨Z, iWZ, iZY, rfl⟩ + rw [Sieve.pullback_comp]; apply K.pullback_stable; clear W iWZ + apply K.superset_covering ?_ (G.functorPushforward_imageSieve_mem _ (iZY ≫ iYX)) + rintro W _ ⟨V, iVZ, iWV, ⟨iVX, e⟩, rfl⟩ + exact ⟨_, iVX, iWV, by simpa [e] using T.1.downward_closed hiYX (G.map iVZ ≫ iZY), by simp [e]⟩ /-- (Implementation). Given a hom between the pullbacks of two sheaves, we can whisker it with `coyoneda` to obtain a hom between the pullbacks of the sheaves of maps from `X`. @@ -155,12 +154,24 @@ theorem sheaf_eq_amalgamation (ℱ : Sheaf K A) {X : A} {U : D} {T : Sieve U} (h t = (ℱ.cond X T hT).amalgamate x hx := (ℱ.cond X T hT).isSeparatedFor x t _ h ((ℱ.cond X T hT).isAmalgamation hx) -variable [Full G] +variable [G.IsLocallyFull K] namespace Types variable {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : SheafOfTypes.{v} K} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) +theorem naturality_apply {X Y : C} (i : G.obj X ⟶ G.obj Y) (x) : + ℱ'.1.map i.op (α.app _ x) = α.app _ (ℱ.map i.op x) := by + have {X Y} (i : X ⟶ Y) (x) : + ℱ'.1.map (G.map i).op (α.app _ x) = α.app _ (ℱ.map (G.map i).op x) := by + exact congr_fun (α.naturality i.op).symm x + refine IsLocallyFull.ext G _ i fun V iVX iVY e ↦ ?_ + simp only [comp_obj, types_comp_apply, ← FunctorToTypes.map_comp_apply, ← op_comp, ← e, this] + +@[reassoc] +theorem naturality {X Y : C} (i : G.obj X ⟶ G.obj Y) : + α.app _ ≫ ℱ'.1.map i.op = ℱ.map i.op ≫ α.app _ := types_ext _ _ (naturality_apply α i) + /-- (Implementation). Given a section of `ℱ` on `X`, we can obtain a family of elements valued in `ℱ'` that is defined on a cover generated by the images of `G`. -/ @@ -181,20 +192,20 @@ noncomputable def pushforwardFamily {X} (x : ℱ.obj (op X)) : /-- (Implementation). The `pushforwardFamily` defined is compatible. -/ theorem pushforwardFamily_compatible {X} (x : ℱ.obj (op X)) : (pushforwardFamily α x).Compatible := by - intro Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ e - apply IsCoverDense.ext G - intro Y f - simp only [pushforwardFamily, ← FunctorToTypes.map_comp_apply, ← op_comp] - change (ℱ.map _ ≫ α.app (op _) ≫ ℱ'.val.map _) _ = (ℱ.map _ ≫ α.app (op _) ≫ ℱ'.val.map _) _ - rw [← G.map_preimage (f ≫ g₁ ≫ _)] - rw [← G.map_preimage (f ≫ g₂ ≫ _)] - erw [← α.naturality (G.preimage _).op] - erw [← α.naturality (G.preimage _).op] - refine congr_fun ?_ x - simp only [Functor.comp_map, ← Category.assoc, Functor.op_map, Quiver.Hom.unop_op, - ← ℱ.map_comp, ← op_comp, G.map_preimage] - congr 3 - simp [e] + suffices ∀ {Z W₁ W₂} (iWX₁ : G.obj W₁ ⟶ X) (iWX₂ : G.obj W₂ ⟶ X) (iZW₁ : Z ⟶ G.obj W₁) + (iZW₂ : Z ⟶ G.obj W₂), iZW₁ ≫ iWX₁ = iZW₂ ≫ iWX₂ → + ℱ'.1.map iZW₁.op (α.app _ (ℱ.map iWX₁.op x)) = ℱ'.1.map iZW₂.op (α.app _ (ℱ.map iWX₂.op x)) by + rintro Y₁ Y₂ Z iZY₁ iZY₂ f₁ f₂ h₁ h₂ e + simp only [pushforwardFamily, ← FunctorToTypes.map_comp_apply, ← op_comp] + generalize Nonempty.some h₁ = l₁ + generalize Nonempty.some h₂ = l₂ + obtain ⟨W₁, iYW₁, iWX₁, rfl⟩ := l₁ + obtain ⟨W₂, iYW₂, iWX₂, rfl⟩ := l₂ + exact this _ _ _ _ (by simpa only [Category.assoc] using e) + introv e + refine ext G _ _ fun V iVZ ↦ ?_ + simp only [← op_comp, ← FunctorToTypes.map_comp_apply, ← Functor.map_comp, naturality_apply, + Category.assoc, e] /-- (Implementation). The morphism `ℱ(X) ⟶ ℱ'(X)` given by gluing the `pushforwardFamily`. -/ noncomputable def appHom (X : D) : ℱ.obj (op X) ⟶ ℱ'.val.obj (op X) := fun x => @@ -204,20 +215,10 @@ noncomputable def appHom (X : D) : ℱ.obj (op X) ⟶ ℱ'.val.obj (op X) := fun @[simp] theorem pushforwardFamily_apply {X} (x : ℱ.obj (op X)) {Y : C} (f : G.obj Y ⟶ X) : pushforwardFamily α x f (Presieve.in_coverByImage G f) = α.app (op Y) (ℱ.map f.op x) := by - unfold pushforwardFamily - -- Porting note: congr_fun was more powerful in Lean 3; I had to explicitly supply - -- the type of the first input here even though it's obvious (there is a unique occurrence - -- of x on each side of the equality) - refine congr_fun (?_ : - (fun t => ℱ'.val.map ((Nonempty.some (_ : coverByImage G X f)).lift.op) - (α.app (op (Nonempty.some (_ : coverByImage G X f)).1) - (ℱ.map ((Nonempty.some (_ : coverByImage G X f)).map.op) t))) = - (fun t => α.app (op Y) (ℱ.map (f.op) t))) x - rw [← G.map_preimage (Nonempty.some _ : Presieve.CoverByImageStructure _ _).lift] - change ℱ.map _ ≫ α.app (op _) ≫ ℱ'.val.map _ = ℱ.map f.op ≫ α.app (op Y) - erw [← α.naturality (G.preimage _).op] - simp only [← Functor.map_comp, ← Category.assoc, Functor.comp_map, G.map_preimage, G.op_map, - Quiver.Hom.unop_op, ← op_comp, Presieve.CoverByImageStructure.fac] + simp only [pushforwardFamily_def, op_obj] + generalize Nonempty.some (Presieve.in_coverByImage G f) = l + obtain ⟨W, iYW, iWX, rfl⟩ := l + simp only [← op_comp, ← FunctorToTypes.map_comp_apply, naturality_apply] @[simp] theorem appHom_restrict {X : D} {Y : C} (f : op X ⟶ op (G.obj Y)) (x) : @@ -251,8 +252,10 @@ noncomputable def appIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val intro Y f simp -/-- Given a natural transformation `G ⋙ ℱ ⟶ G ⋙ ℱ'` between presheaves of types, where `G` is -full and cover-dense, and `ℱ'` is a sheaf, we may obtain a natural transformation between sheaves. +/-- +Given a natural transformation `G ⋙ ℱ ⟶ G ⋙ ℱ'` between presheaves of types, +where `G` is locally-full and cover-dense, and `ℱ'` is a sheaf, +we may obtain a natural transformation between sheaves. -/ @[simps] noncomputable def presheafHom (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : ℱ ⟶ ℱ'.val where @@ -264,16 +267,20 @@ noncomputable def presheafHom (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : ℱ simp only [appHom_restrict, types_comp_apply, ← FunctorToTypes.map_comp_apply] -- Porting note: Lean 3 proof continued with a rewrite but we're done here -/-- Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of types, where `G` is full -and cover-dense, and `ℱ, ℱ'` are sheaves, we may obtain a natural isomorphism between presheaves. +/-- +Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of types, +where `G` is locally-full and cover-dense, and `ℱ, ℱ'` are sheaves, +we may obtain a natural isomorphism between presheaves. -/ @[simps!] noncomputable def presheafIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : ℱ.val ≅ ℱ'.val := NatIso.ofComponents (fun X => appIso i (unop X)) @(presheafHom i.hom).naturality -/-- Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of types, where `G` is full -and cover-dense, and `ℱ, ℱ'` are sheaves, we may obtain a natural isomorphism between sheaves. +/-- +Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of types, +where `G` is locally-full and cover-dense, and `ℱ, ℱ'` are sheaves, +we may obtain a natural isomorphism between sheaves. -/ @[simps] noncomputable def sheafIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : @@ -332,17 +339,19 @@ noncomputable def sheafYonedaHom (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : ext X x exact congr_fun (((sheafCoyonedaHom α).app X).naturality i) x -/-- Given a natural transformation `G ⋙ ℱ ⟶ G ⋙ ℱ'` between presheaves of arbitrary category, -where `G` is full and cover-dense, and `ℱ'` is a sheaf, we may obtain a natural transformation -between presheaves. +/-- +Given a natural transformation `G ⋙ ℱ ⟶ G ⋙ ℱ'` between presheaves of arbitrary category, +where `G` is locally-full and cover-dense, and `ℱ'` is a sheaf, we may obtain a natural +transformation between presheaves. -/ noncomputable def sheafHom (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : ℱ ⟶ ℱ'.val := let α' := sheafYonedaHom α { app := fun X => yoneda.preimage (α'.app X) naturality := fun X Y f => yoneda.map_injective (by simpa using α'.naturality f) } -/-- Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of arbitrary category, -where `G` is full and cover-dense, and `ℱ', ℱ` are sheaves, +/-- +Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of arbitrary category, +where `G` is locally-full and cover-dense, and `ℱ', ℱ` are sheaves, we may obtain a natural isomorphism between presheaves. -/ @[simps!] @@ -364,8 +373,9 @@ noncomputable def presheafIso {ℱ ℱ' : Sheaf K A} (i : G.op ⋙ ℱ.val ≅ G haveI : IsIso (sheafHom i.hom) := by apply NatIso.isIso_of_isIso_app apply asIso (sheafHom i.hom) -/-- Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of arbitrary category, -where `G` is full and cover-dense, and `ℱ', ℱ` are sheaves, +/-- +Given a natural isomorphism `G ⋙ ℱ ≅ G ⋙ ℱ'` between presheaves of arbitrary category, +where `G` is locally-full and cover-dense, and `ℱ', ℱ` are sheaves, we may obtain a natural isomorphism between presheaves. -/ @[simps] @@ -379,8 +389,7 @@ noncomputable def sheafIso {ℱ ℱ' : Sheaf K A} (i : G.op ⋙ ℱ.val ≅ G.op ext1 apply (presheafIso i).inv_hom_id -/-- The constructed `sheafHom α` is equal to `α` when restricted onto `C`. --/ +/-- The constructed `sheafHom α` is equal to `α` when restricted onto `C`. -/ theorem sheafHom_restrict_eq (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : whiskerLeft G.op (sheafHom α) = α := by ext X @@ -400,14 +409,14 @@ theorem sheafHom_restrict_eq (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : congr 1 simp only [Category.assoc] congr 1 - rw [← G.map_preimage hf.some.map] - symm - apply α.naturality (G.preimage hf.some.map).op - -- porting note; Lean 3 needed a random `inferInstance` for cleanup here; not necessary in lean 4 + have := naturality_apply (G := G) (ℱ := ℱ ⋙ coyoneda.obj (op <| (G.op ⋙ ℱ).obj X)) + (ℱ' := ⟨_, ℱ'.2 ((G.op ⋙ ℱ).obj X)⟩) (whiskerRight α (coyoneda.obj _)) hf.some.map (𝟙 _) + simpa using this variable (G) -/-- If the pullback map is obtained via whiskering, +/-- +If the pullback map is obtained via whiskering, then the result `sheaf_hom (whisker_left G.op α)` is equal to `α`. -/ theorem sheafHom_eq (α : ℱ ⟶ ℱ'.val) : sheafHom (whiskerLeft G.op α) = α := by @@ -429,7 +438,8 @@ theorem sheafHom_eq (α : ℱ ⟶ ℱ'.val) : sheafHom (whiskerLeft G.op α) = variable {G} -/-- A full and cover-dense functor `G` induces an equivalence between morphisms into a sheaf and +/-- +A locally-full and cover-dense functor `G` induces an equivalence between morphisms into a sheaf and morphisms over the restrictions via `G`. -/ noncomputable def restrictHomEquivHom : (G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) ≃ (ℱ ⟶ ℱ'.val) where @@ -438,8 +448,8 @@ noncomputable def restrictHomEquivHom : (G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) ≃ left_inv := sheafHom_restrict_eq right_inv := sheafHom_eq _ -/-- Given a full and cover-dense functor `G` and a natural transformation of sheaves `α : ℱ ⟶ ℱ'`, -if the pullback of `α` along `G` is iso, then `α` is also iso. +/-- Given a locally-full and cover-dense functor `G` and a natural transformation of sheaves +`α : ℱ ⟶ ℱ'`, if the pullback of `α` along `G` is iso, then `α` is also iso. -/ theorem iso_of_restrict_iso {ℱ ℱ' : Sheaf K A} (α : ℱ ⟶ ℱ') (i : IsIso (whiskerLeft G.op α.val)) : IsIso α := by @@ -449,20 +459,21 @@ theorem iso_of_restrict_iso {ℱ ℱ' : Sheaf K A} (α : ℱ ⟶ ℱ') (i : IsIs variable (G K) -/-- A fully faithful cover-dense functor preserves compatible families. -/ -lemma compatiblePreserving [Faithful G] : CompatiblePreserving K G := by +/-- A locally-fully-faithful and cover-dense functor preserves compatible families. -/ +lemma compatiblePreserving [G.IsLocallyFaithful K] : CompatiblePreserving K G := by constructor intro ℱ Z T x hx Y₁ Y₂ X f₁ f₂ g₁ g₂ hg₁ hg₂ eq apply Functor.IsCoverDense.ext G intro W i - simp only [← FunctorToTypes.map_comp_apply, ← op_comp] - rw [← G.map_preimage (i ≫ f₁)] - rw [← G.map_preimage (i ≫ f₂)] - apply hx (G.preimage (i ≫ f₁)) ((G.preimage (i ≫ f₂))) hg₁ hg₂ - apply G.map_injective - simp [eq] - -lemma isContinuous [Faithful G] (Hp : CoverPreserving J K G) : G.IsContinuous J K := + refine IsLocallyFull.ext G _ (i ≫ f₁) fun V₁ iVW iV₁Y₁ e₁ ↦ ?_ + refine IsLocallyFull.ext G _ (G.map iVW ≫ i ≫ f₂) fun V₂ iV₂V₁ iV₂Y₂ e₂ ↦ ?_ + refine IsLocallyFaithful.ext G _ (iV₂V₁ ≫ iV₁Y₁ ≫ g₁) (iV₂Y₂ ≫ g₂) (by simp [e₁, e₂, eq]) ?_ + intro V₃ iV₃ e₄ + simp only [← op_comp, ← FunctorToTypes.map_comp_apply, ← e₁, ← e₂, ← Functor.map_comp] + apply hx + simpa using e₄ + +lemma isContinuous [G.IsLocallyFaithful K] (Hp : CoverPreserving J K G) : G.IsContinuous J K := isContinuous_of_coverPreserving (compatiblePreserving K G) Hp instance full_sheafPushforwardContinuous [G.IsContinuous J K] : @@ -483,51 +494,183 @@ end IsCoverDense /-- If `G : C ⥤ D` is cover dense and full, then the map `(P ⟶ Q) → (G.op ⋙ P ⟶ G.op ⋙ Q)` is bijective when `Q` is a sheaf`. -/ lemma whiskerLeft_obj_map_bijective_of_isCoverDense (G : C ⥤ D) - [G.IsCoverDense K] [G.Full] {A : Type*} [Category A] + [G.IsCoverDense K] [G.IsLocallyFull K] {A : Type*} [Category A] (P Q : Dᵒᵖ ⥤ A) (hQ : Presheaf.IsSheaf K Q) : Function.Bijective (((whiskeringLeft Cᵒᵖ Dᵒᵖ A).obj G.op).map : (P ⟶ Q) → _) := (IsCoverDense.restrictHomEquivHom (ℱ' := ⟨Q, hQ⟩)).symm.bijective +variable {A : Type*} [Category A] +variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) (G : C ⥤ D) + +/-- The functor `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)` +if `G` is cover-dense, locally fully-faithful, +and `S` is a cover of `C` if and only if the image of `S` in `D` is a cover. -/ +class IsDenseSubsite : Prop where + isCoverDense' : G.IsCoverDense K := by infer_instance + isLocallyFull' : G.IsLocallyFull K := by infer_instance + isLocallyFaithful' : G.IsLocallyFaithful K := by infer_instance + functorPushforward_mem_iff : ∀ {X : C} {S : Sieve X}, S.functorPushforward G ∈ K _ ↔ S ∈ J _ + +namespace IsDenseSubsite + +variable [G.IsDenseSubsite J K] + +lemma isCoverDense : G.IsCoverDense K := isCoverDense' J +lemma isLocallyFull : G.IsLocallyFull K := isLocallyFull' J +lemma isLocallyFaithful : G.IsLocallyFaithful K := isLocallyFaithful' J + +lemma coverPreserving [G.IsDenseSubsite J K] : CoverPreserving J K G := + ⟨functorPushforward_mem_iff.mpr⟩ + +instance (priority := 900) [G.IsDenseSubsite J K] : G.IsContinuous J K := + letI := IsDenseSubsite.isCoverDense J K G + letI := IsDenseSubsite.isLocallyFull J K G + letI := IsDenseSubsite.isLocallyFaithful J K G + IsCoverDense.isContinuous J K G (IsDenseSubsite.coverPreserving J K G) + +instance (priority := 900) [G.IsDenseSubsite J K] : G.IsCocontinuous J K where + cover_lift hS := + letI := IsDenseSubsite.isCoverDense J K G + letI := IsDenseSubsite.isLocallyFull J K G + IsDenseSubsite.functorPushforward_mem_iff.mp + (IsCoverDense.functorPullback_pushforward_covering ⟨_, hS⟩) + +instance full_sheafPushforwardContinuous [G.IsDenseSubsite J K] : + Full (G.sheafPushforwardContinuous A J K) := + letI := IsDenseSubsite.isCoverDense J K G + letI := IsDenseSubsite.isLocallyFull J K G + inferInstance + +instance faithful_sheafPushforwardContinuous [G.IsDenseSubsite J K] : + Faithful (G.sheafPushforwardContinuous A J K) := + letI := IsDenseSubsite.isCoverDense J K G + letI := IsDenseSubsite.isLocallyFull J K G + inferInstance + +lemma imageSieve_mem [G.IsDenseSubsite J K] {U V} (f : G.obj U ⟶ G.obj V) : + G.imageSieve f ∈ J _ := + letI := IsDenseSubsite.isLocallyFull J K G + IsDenseSubsite.functorPushforward_mem_iff.mp (G.functorPushforward_imageSieve_mem K f) + +lemma equalizer_mem [G.IsDenseSubsite J K] {U V} (f₁ f₂ : U ⟶ V) (e : G.map f₁ = G.map f₂) : + Sieve.equalizer f₁ f₂ ∈ J _ := + letI := IsDenseSubsite.isLocallyFaithful J K G + IsDenseSubsite.functorPushforward_mem_iff.mp (G.functorPushforward_equalizer_mem K f₁ f₂ e) + +end IsDenseSubsite + end Functor end CategoryTheory -namespace CategoryTheory.Functor.IsCoverDense +namespace CategoryTheory.Functor.IsDenseSubsite -open CategoryTheory +open CategoryTheory Opposite universe w' variable {C D : Type*} [Category C] [Category D] -variable (G : C ⥤ D) [Full G] [Faithful G] +variable (G : C ⥤ D) variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) variable {A : Type w} [Category.{w'} A] [∀ X, Limits.HasLimitsOfShape (StructuredArrow X G.op) A] -variable [G.IsCoverDense K] [G.IsContinuous J K] [G.IsCocontinuous J K] +variable [G.IsDenseSubsite J K] + +lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : + IsIso ((yoneda.map ((G.op.ranCounit.app Y.val).app (op U))).app (op X)) := by + rw [isIso_iff_bijective] + constructor + · intro f₁ f₂ e + apply (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).hom_ext + rintro ⟨⟨⟨⟩⟩, ⟨W⟩, g⟩ + obtain ⟨g, rfl⟩ : ∃ g' : G.obj W ⟶ G.obj U, g = g'.op := ⟨g.unop, rfl⟩ + simp only [id_obj, comp_obj, StructuredArrow.proj_obj, RightExtension.coneAt_pt, + RightExtension.mk_left, RightExtension.coneAt_π_app, const_obj_obj, op_obj, + whiskeringLeft_obj_obj, RightExtension.mk_hom] + apply (Y.2 X _ (IsDenseSubsite.imageSieve_mem J K G g)).isSeparatedFor.ext + rintro V iVW ⟨iVU, e'⟩ + have := congr($e ≫ Y.1.map iVU.op) + simp only [comp_obj, yoneda_map_app, Category.assoc, coyoneda_obj_obj, comp_map, + coyoneda_obj_map, ← NatTrans.naturality, op_obj, op_map, Quiver.Hom.unop_op, ← map_comp_assoc, + ← op_comp, ← e'] at this ⊢ + erw [← NatTrans.naturality] at this + exact this + · intro f + have (X Y Z) (f : X ⟶ Y) (g : G.obj Y ⟶ G.obj Z) (hf : G.imageSieve g f) : Exists _ := hf + choose l hl using this + let c : Limits.Cone (StructuredArrow.proj (op (G.obj U)) G.op ⋙ Y.val) := by + refine ⟨X, ⟨fun g ↦ ?_, ?_⟩⟩ + · refine Y.2.amalgamate ⟨_, IsDenseSubsite.imageSieve_mem J K G g.hom.unop⟩ + (fun I ↦ f ≫ Y.1.map (l _ _ _ _ _ I.hf).op) fun I₁ I₂ r ↦ ?_ + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (r.g₁ ≫ l _ _ _ _ _ I₁.hf) + (r.g₂ ≫ l _ _ _ _ _ I₂.hf) ?_)).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + · simp only [const_obj_obj, op_obj, map_comp, hl] + simp only [← map_comp_assoc, r.w] + · simp [← map_comp, ← op_comp, hiUV] + · rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ + dsimp at g₁ g₂ i hi + obtain rfl : g₂ = g₁ ≫ (G.map i.unop).op := by simpa only [Category.id_comp] using hi + obtain ⟨g, rfl⟩ : ∃ g' : G.obj W₁ ⟶ G.obj U, g₁ = g'.op := ⟨g₁.unop, rfl⟩ + obtain ⟨i, rfl⟩ : ∃ i' : W₂ ⟶ W₁, i = i'.op := ⟨i.unop, rfl⟩ + simp only [const_obj_obj, id_obj, comp_obj, StructuredArrow.proj_obj, const_obj_map, op_obj, + unop_comp, Quiver.Hom.unop_op, Category.id_comp, comp_map, StructuredArrow.proj_map] + apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (G.map i ≫ g)⟩ + intro I + simp only [Presheaf.IsSheaf.amalgamate_map, Category.assoc, ← Functor.map_comp, ← op_comp] + let I' : GrothendieckTopology.Cover.Arrow ⟨_, IsDenseSubsite.imageSieve_mem J K G g⟩ := + ⟨_, I.f ≫ i, ⟨l _ _ _ _ _ I.hf, by simp [hl]⟩⟩ + refine Eq.trans ?_ (Y.2.amalgamate_map _ _ _ I').symm + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) + (l _ _ _ _ _ I'.hf) (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + simp [← Functor.map_comp, ← op_comp, hiUV] + refine ⟨(isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).lift c, ?_⟩ + · have := (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).fac c (.mk (𝟙 _)) + simp only [id_obj, comp_obj, StructuredArrow.proj_obj, StructuredArrow.mk_right, + RightExtension.coneAt_pt, RightExtension.mk_left, RightExtension.coneAt_π_app, + const_obj_obj, op_obj, StructuredArrow.mk_hom_eq_self, map_id, whiskeringLeft_obj_obj, + RightExtension.mk_hom, Category.id_comp, StructuredArrow.mk_left, unop_id] at this + simp only [id_obj, yoneda_map_app, this] + apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (𝟙 (G.obj U))⟩ _ _ fun I ↦ ?_ + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) + I.f (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + simp [← Functor.map_comp, ← op_comp, hiUV] instance (Y : Sheaf J A) : IsIso ((G.sheafAdjunctionCocontinuous A J K).counit.app Y) := by - haveI : IsIso ((sheafToPresheaf J A).map - ((G.sheafAdjunctionCocontinuous A J K).counit.app Y)) := by - dsimp - rw [sheafAdjunctionCocontinuous_counit_app_val] - infer_instance - apply ReflectsIsomorphisms.reflects (sheafToPresheaf J A) + apply (config := { allowSynthFailures := true }) + ReflectsIsomorphisms.reflects (sheafToPresheaf J A) + rw [NatTrans.isIso_iff_isIso_app] + intro ⟨U⟩ + apply (config := { allowSynthFailures := true }) ReflectsIsomorphisms.reflects yoneda + rw [NatTrans.isIso_iff_isIso_app] + intro ⟨X⟩ + simp only [comp_obj, sheafToPresheaf_obj, sheafPushforwardContinuous_obj_val_obj, yoneda_obj_obj, + id_obj, sheafToPresheaf_map, sheafAdjunctionCocontinuous_counit_app_val, ranAdjunction_counit] + exact isIso_ranCounit_app_of_isDenseSubsite G J K Y U X variable (A) -/-- Given a functor between small sites that is cover-dense, cover-preserving, and cover-lifting, -it induces an equivalence of category of sheaves valued in a complete category. +/-- +If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, +it induces an equivalence of category of sheaves valued in a category with suitable limits. -/ @[simps! functor inverse] -noncomputable def sheafEquivOfCoverPreservingCoverLifting : Sheaf J A ≌ Sheaf K A := +noncomputable def sheafEquiv : Sheaf J A ≌ Sheaf K A := (G.sheafAdjunctionCocontinuous A J K).toEquivalence.symm variable [HasWeakSheafify J A] [HasWeakSheafify K A] /-- The natural isomorphism exhibiting the compatibility of -`sheafEquivOfCoverPreservingCoverLifting` with sheafification. -/ +`IsDenseSubsite.sheafEquiv` with sheafification. -/ noncomputable -abbrev sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility : +abbrev sheafEquivSheafificationCompatibility : (whiskeringLeft _ _ A).obj G.op ⋙ presheafToSheaf _ _ ≅ - presheafToSheaf _ _ ⋙ (sheafEquivOfCoverPreservingCoverLifting G J K A).inverse := by + presheafToSheaf _ _ ⋙ (sheafEquiv G J K A).inverse := by apply Functor.pushforwardContinuousSheafificationCompatibility -end CategoryTheory.Functor.IsCoverDense +end IsDenseSubsite + +@[deprecated (since := "2024-07-23")] +alias IsCoverDense.sheafEquivOfCoverPreservingCoverLifting := IsDenseSubsite.sheafEquiv +@[deprecated (since := "2024-07-23")] +alias IsCoverDense.sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility := + IsDenseSubsite.sheafEquivSheafificationCompatibility + +end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/Discrete.lean b/Mathlib/CategoryTheory/Sites/Discrete.lean index 978e0410f3691..2b59807d4ad46 100644 --- a/Mathlib/CategoryTheory/Sites/Discrete.lean +++ b/Mathlib/CategoryTheory/Sites/Discrete.lean @@ -76,17 +76,17 @@ lemma isDiscrete_iff_isIso_counit_app {L : A ⥤ Sheaf J A} (adj : L ⊣ (sheafS section Equivalence variable {D : Type*} [Category D] (K : GrothendieckTopology D) [HasWeakSheafify K A] -variable (G : C ⥤ D) [G.Full] [G.Faithful] +variable (G : C ⥤ D) [∀ (X : Dᵒᵖ), HasLimitsOfShape (StructuredArrow X G.op) A] - [G.IsCoverDense K] [G.IsContinuous J K] [G.IsCocontinuous J K] (ht' : IsTerminal (G.obj t)) + [G.IsDenseSubsite J K] (ht' : IsTerminal (G.obj t)) variable [(constantSheaf J A).Faithful] [(constantSheaf J A).Full] -open Functor.IsCoverDense +open Functor.IsDenseSubsite noncomputable example : let e : Sheaf J A ≌ Sheaf K A := - sheafEquivOfCoverPreservingCoverLifting G J K A + sheafEquiv G J K A e.inverse ⋙ (sheafSections J A).obj (op t) ≅ (sheafSections K A).obj (op (G.obj t)) := Iso.refl _ @@ -99,10 +99,10 @@ property of a sheaf of being a discrete object is invariant under equivalence of -/ noncomputable def equivCommuteConstant : let e : Sheaf J A ≌ Sheaf K A := - sheafEquivOfCoverPreservingCoverLifting G J K A + sheafEquiv G J K A constantSheaf J A ⋙ e.functor ≅ constantSheaf K A := let e : Sheaf J A ≌ Sheaf K A := - sheafEquivOfCoverPreservingCoverLifting G J K A + sheafEquiv G J K A (Adjunction.leftAdjointUniq ((constantSheafAdj J A ht).comp e.toAdjunction) (constantSheafAdj K A ht')) @@ -115,10 +115,10 @@ property of a sheaf of being a discrete object is invariant under equivalence of -/ noncomputable def equivCommuteConstant' : let e : Sheaf J A ≌ Sheaf K A := - sheafEquivOfCoverPreservingCoverLifting G J K A + sheafEquiv G J K A constantSheaf J A ≅ constantSheaf K A ⋙ e.inverse := let e : Sheaf J A ≌ Sheaf K A := - sheafEquivOfCoverPreservingCoverLifting G J K A + sheafEquiv G J K A isoWhiskerLeft (constantSheaf J A) e.unitIso ≪≫ isoWhiskerRight (equivCommuteConstant J A ht K G ht') e.inverse @@ -128,7 +128,7 @@ categories. -/ lemma isDiscrete_iff_of_equivalence (F : Sheaf K A) : let e : Sheaf J A ≌ Sheaf K A := - sheafEquivOfCoverPreservingCoverLifting G J K A + sheafEquiv G J K A haveI : (constantSheaf K A).Faithful := Functor.Faithful.of_iso (equivCommuteConstant J A ht K G ht') haveI : (constantSheaf K A).Full := diff --git a/Mathlib/CategoryTheory/Sites/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Equivalence.lean index 5808937e197f4..853a30fb44a3a 100644 --- a/Mathlib/CategoryTheory/Sites/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Equivalence.lean @@ -48,51 +48,9 @@ variable (A : Type u₃) [Category.{v₃} A] namespace Equivalence -theorem locallyCoverDense : LocallyCoverDense J e.inverse := by - intro X T - convert T.prop - ext Z f - constructor - · rintro ⟨_, _, g', hg, rfl⟩ - exact T.val.downward_closed hg g' - · intro hf - refine ⟨e.functor.obj Z, (Adjunction.homEquiv e.toAdjunction _ _).symm f, e.unit.app Z, ?_, ?_⟩ - · simp only [Adjunction.homEquiv_counit, Functor.id_obj, Equivalence.toAdjunction_counit, - Sieve.functorPullback_apply, Presieve.functorPullback_mem, Functor.map_comp, - Equivalence.inv_fun_map, Functor.comp_obj, Category.assoc, Equivalence.unit_inverse_comp, - Category.comp_id] - exact T.val.downward_closed hf _ - · simp - -theorem coverPreserving : CoverPreserving J (e.locallyCoverDense J).inducedTopology e.functor where - cover_preserve {U S} h := by - change _ ∈ J.sieves (e.inverse.obj (e.functor.obj U)) - convert J.pullback_stable (e.unitInv.app U) h - ext Z f - rw [← Sieve.functorPushforward_comp] - simp only [Sieve.functorPushforward_apply, Presieve.functorPushforward, exists_and_left, id_obj, - comp_obj, Sieve.pullback_apply] - constructor - · rintro ⟨W, g, hg, x, rfl⟩ - rw [Category.assoc] - apply S.downward_closed - simpa using S.downward_closed hg _ - · intro hf - exact ⟨_, e.unitInv.app Z ≫ f ≫ e.unitInv.app U, S.downward_closed hf _, - e.unit.app Z ≫ e.unit.app _, by simp⟩ - -instance : IsCoverDense e.functor (e.locallyCoverDense J).inducedTopology where - is_cover U := by - change _ ∈ J.sieves _ - convert J.top_mem (e.inverse.obj U) - ext Y f - simp only [Sieve.functorPushforward_apply, Presieve.functorPushforward, exists_and_left, - Sieve.top_apply, iff_true] - exact ⟨e.functor.obj Y, (Adjunction.homEquiv e.toAdjunction _ _).symm f, - Presieve.in_coverByImage _ _, e.unit.app _, by simp⟩ - -instance : IsCoverDense e.inverse J where +instance (priority := 900) [G.IsEquivalence] : IsCoverDense G J where is_cover U := by + let e := (asEquivalence G).symm convert J.top_mem U ext Y f simp only [Sieve.functorPushforward_apply, Presieve.functorPushforward, exists_and_left, @@ -102,16 +60,28 @@ instance : IsCoverDense e.inverse J where replace := Sieve.downward_closed _ this (e.unit.app Y) simpa [g] using this +instance : e.functor.IsDenseSubsite J (e.inverse.inducedTopology J) := by + have : J = e.functor.inducedTopology (e.inverse.inducedTopology J) := by + ext X S + show _ ↔ _ ∈ J.sieves _ + erw [← GrothendieckTopology.pullback_mem_iff_of_isIso (i := e.unit.app X)] + congr!; ext Y f; simp + nth_rw 1 [this] + infer_instance + +theorem coverPreserving : CoverPreserving J (e.inverse.inducedTopology J) e.functor := + IsDenseSubsite.coverPreserving _ _ _ + /-- A class saying that the equivalence `e` transports the Grothendieck topology `J` to `K`. -/ class TransportsGrothendieckTopology : Prop where /-- `K` is equal to the induced topology. -/ - eq_inducedTopology : K = (e.locallyCoverDense J).inducedTopology + eq_inducedTopology : K = e.inverse.inducedTopology J -instance : e.TransportsGrothendieckTopology J (e.locallyCoverDense J).inducedTopology := ⟨rfl⟩ +instance : e.TransportsGrothendieckTopology J (e.inverse.inducedTopology J) := ⟨rfl⟩ variable [e.TransportsGrothendieckTopology J K] -theorem eq_inducedTopology_of_transports : K = (e.locallyCoverDense J).inducedTopology := +theorem eq_inducedTopology_of_transports : K = e.inverse.inducedTopology J := TransportsGrothendieckTopology.eq_inducedTopology instance : IsContinuous e.functor J K := by @@ -120,7 +90,7 @@ instance : IsContinuous e.functor J K := by instance : IsContinuous e.inverse K J := by rw [eq_inducedTopology_of_transports J K e] - exact IsCoverDense.isContinuous _ _ _ (e.locallyCoverDense J).inducedTopology_coverPreserving + infer_instance /-- The functor in the equivalence of sheaf categories. -/ @[simps!] @@ -215,13 +185,13 @@ end Equivalence variable [EssentiallySmall.{w} C] variable (B : Type*) [Category B] (F : A ⥤ B) -variable [HasSheafify ((equivSmallModel C).locallyCoverDense J).inducedTopology A] -variable [((equivSmallModel C).locallyCoverDense J).inducedTopology.HasSheafCompose F] +variable [HasSheafify ((equivSmallModel C).inverse.inducedTopology J) A] +variable [((equivSmallModel C).inverse.inducedTopology J).HasSheafCompose F] /-- Transport to a small model and sheafify there. -/ noncomputable def smallSheafify : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A := (equivSmallModel C).transportAndSheafify J - ((equivSmallModel C).locallyCoverDense J).inducedTopology A + ((equivSmallModel C).inverse.inducedTopology J) A /-- Transporting to a small model and sheafifying there is left adjoint to the underlying presheaf @@ -232,22 +202,22 @@ def smallSheafificationAdjunction : smallSheafify J A ⊣ sheafToPresheaf J A := (equivSmallModel C).transportSheafificationAdjunction J _ A noncomputable instance hasSheafifyEssentiallySmallSite : HasSheafify J A := - (equivSmallModel C).hasSheafify J ((equivSmallModel C).locallyCoverDense J).inducedTopology A + (equivSmallModel C).hasSheafify J ((equivSmallModel C).inverse.inducedTopology J) A instance hasSheafComposeEssentiallySmallSite : HasSheafCompose J F := - (equivSmallModel C).hasSheafCompose J ((equivSmallModel C).locallyCoverDense J).inducedTopology F + (equivSmallModel C).hasSheafCompose J ((equivSmallModel C).inverse.inducedTopology J) F instance hasLimitsEssentiallySmallSite - [HasLimits <| Sheaf ((equivSmallModel C).locallyCoverDense J).inducedTopology A] : + [HasLimits <| Sheaf ((equivSmallModel C).inverse.inducedTopology J) A] : HasLimitsOfSize.{max v₃ w, max v₃ w} <| Sheaf J A := Adjunction.has_limits_of_equivalence ((equivSmallModel C).sheafCongr J - ((equivSmallModel C).locallyCoverDense J).inducedTopology A).functor + ((equivSmallModel C).inverse.inducedTopology J) A).functor instance hasColimitsEssentiallySmallSite - [HasColimits <| Sheaf ((equivSmallModel C).locallyCoverDense J).inducedTopology A] : + [HasColimits <| Sheaf ((equivSmallModel C).inverse.inducedTopology J) A] : HasColimitsOfSize.{max v₃ w, max v₃ w} <| Sheaf J A := Adjunction.has_colimits_of_equivalence ((equivSmallModel C).sheafCongr J - ((equivSmallModel C).locallyCoverDense J).inducedTopology A).functor + ((equivSmallModel C).inverse.inducedTopology J) A).functor namespace GrothendieckTopology diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index c0a9e3738620d..f7773b4fc0fb2 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -121,6 +121,14 @@ theorem top_mem (X : C) : ⊤ ∈ J X := theorem pullback_stable (f : Y ⟶ X) (hS : S ∈ J X) : S.pullback f ∈ J Y := J.pullback_stable' f hS +variable {J} in +@[simp] +lemma pullback_mem_iff_of_isIso {i : X ⟶ Y} [IsIso i] {S : Sieve Y} : + S.pullback i ∈ J _ ↔ S ∈ J _ := by + refine ⟨fun H ↦ ?_, J.pullback_stable i⟩ + convert J.pullback_stable (inv i) H + rw [← Sieve.pullback_comp, IsIso.inv_hom_id, Sieve.pullback_id] + theorem transitive (hS : S ∈ J X) (R : Sieve X) (h : ∀ ⦃Y⦄ ⦃f : Y ⟶ X⦄, S f → R.pullback f ∈ J Y) : R ∈ J X := J.transitive' hS R h diff --git a/Mathlib/CategoryTheory/Sites/InducedTopology.lean b/Mathlib/CategoryTheory/Sites/InducedTopology.lean index 3d71abeb6d788..48981facbaeb7 100644 --- a/Mathlib/CategoryTheory/Sites/InducedTopology.lean +++ b/Mathlib/CategoryTheory/Sites/InducedTopology.lean @@ -29,37 +29,35 @@ Given a fully faithful cover-dense functor `G : C ⥤ (D, K)` between small site -/ - namespace CategoryTheory universe v u -open Limits Opposite Presieve - -section +open Limits Opposite Presieve CategoryTheory -variable {C : Type*} [Category C] {D : Type*} [Category D] {G : C ⥤ D} -variable {J : GrothendieckTopology C} {K : GrothendieckTopology D} +variable {C : Type*} [Category C] {D : Type*} [Category D] (G : C ⥤ D) +variable {J : GrothendieckTopology C} (K : GrothendieckTopology D) variable (A : Type v) [Category.{u} A] +namespace Functor + -- variables (A) [full G] [faithful G] /-- We say that a functor `C ⥤ D` into a site is "locally dense" if for each covering sieve `T` in `D`, `T ∩ mor(C)` generates a covering sieve in `D`. -/ -def LocallyCoverDense (K : GrothendieckTopology D) (G : C ⥤ D) : Prop := - ∀ ⦃X : C⦄ (T : K (G.obj X)), (T.val.functorPullback G).functorPushforward G ∈ K (G.obj X) - -namespace LocallyCoverDense +class LocallyCoverDense : Prop where + functorPushforward_functorPullback_mem : + ∀ ⦃X : C⦄ (T : K (G.obj X)), (T.val.functorPullback G).functorPushforward G ∈ K (G.obj X) -variable [G.Full] [G.Faithful] (Hld : LocallyCoverDense K G) +variable [G.LocallyCoverDense K] [G.IsLocallyFull K] [G.IsLocallyFaithful K] -theorem pushforward_cover_iff_cover_pullback {X : C} (S : Sieve X) : +theorem pushforward_cover_iff_cover_pullback [G.Full] [G.Faithful] {X : C} (S : Sieve X) : K _ (S.functorPushforward G) ↔ ∃ T : K (G.obj X), T.val.functorPullback G = S := by constructor · intro hS exact ⟨⟨_, hS⟩, (Sieve.fullyFaithfulFunctorGaloisCoinsertion G X).u_l_eq S⟩ · rintro ⟨T, rfl⟩ - exact Hld T + exact LocallyCoverDense.functorPushforward_functorPullback_mem T /-- If a functor `G : C ⥤ (D, K)` is fully faithful and locally dense, then the set `{ T ∩ mor(C) | T ∈ K }` is a grothendieck topology of `C`. @@ -71,15 +69,23 @@ def inducedTopology : GrothendieckTopology C where change K _ _ rw [Sieve.functorPushforward_top] exact K.top_mem _ - pullback_stable' X Y S f hS := by - have : S.pullback f = ((S.functorPushforward G).pullback (G.map f)).functorPullback G := by - conv_lhs => rw [← (Sieve.fullyFaithfulFunctorGaloisCoinsertion G X).u_l_eq S] - ext - change (S.functorPushforward G) _ ↔ (S.functorPushforward G) _ - rw [G.map_comp] - rw [this] - change K _ _ - apply Hld ⟨_, K.pullback_stable (G.map f) hS⟩ + pullback_stable' X Y S iYX hS := by + apply K.transitive (LocallyCoverDense.functorPushforward_functorPullback_mem + ⟨_, K.pullback_stable (G.map iYX) hS⟩) + rintro Z _ ⟨U, iUY, iZU, ⟨W, iWX, iUW, hiWX, e₁⟩, rfl⟩ + rw [Sieve.pullback_comp] + apply K.pullback_stable + clear iZU Z + apply K.transitive (G.functorPushforward_imageSieve_mem _ iUW) + rintro Z _ ⟨U₁, iU₁U, iZU₁, ⟨iU₁W, e₂⟩, rfl⟩ + rw [Sieve.pullback_comp] + apply K.pullback_stable + clear iZU₁ Z + apply K.superset_covering ?_ (G.functorPushforward_equalizer_mem _ + (iU₁U ≫ iUY ≫ iYX) (iU₁W ≫ iWX) (by simp [e₁, e₂])) + rintro Z _ ⟨U₂, iU₂U₁, iZU₂, e₃ : _ = _, rfl⟩ + refine ⟨_, iU₂U₁ ≫ iU₁U ≫ iUY, iZU₂, ?_, by simp⟩ + simpa [e₃] using S.downward_closed hiWX (iU₂U₁ ≫ iU₁W) transitive' X S hS S' H' := by apply K.transitive hS rintro Y _ ⟨Z, g, i, hg, rfl⟩ @@ -91,74 +97,47 @@ def inducedTopology : GrothendieckTopology C where simp /-- `G` is cover-lifting wrt the induced topology. -/ -theorem inducedTopology_isCocontinuous : G.IsCocontinuous Hld.inducedTopology K := - ⟨@fun _ S hS => Hld ⟨S, hS⟩⟩ +instance inducedTopology_isCocontinuous : G.IsCocontinuous (G.inducedTopology K) K := + ⟨@fun _ S hS => LocallyCoverDense.functorPushforward_functorPullback_mem ⟨S, hS⟩⟩ /-- `G` is cover-preserving wrt the induced topology. -/ -theorem inducedTopology_coverPreserving : CoverPreserving Hld.inducedTopology K G := +theorem inducedTopology_coverPreserving : CoverPreserving (G.inducedTopology K) K G := ⟨@fun _ _ hS => hS⟩ -end LocallyCoverDense - -variable (G K) +instance (priority := 900) locallyCoverDense_of_isCoverDense [G.IsCoverDense K] : + G.LocallyCoverDense K where + functorPushforward_functorPullback_mem _ _ := + IsCoverDense.functorPullback_pushforward_covering _ -theorem Functor.locallyCoverDense_of_isCoverDense [Full G] [G.IsCoverDense K] : - LocallyCoverDense K G := by - intro X T - refine K.superset_covering ?_ (K.bind_covering T.property - fun Y f _ => G.is_cover_of_isCoverDense _ Y) - rintro Y _ ⟨Z, _, f, hf, ⟨W, g, f', rfl : _ = _⟩, rfl⟩ - use W; use G.preimage (f' ≫ f); use g - constructor - · simpa using T.val.downward_closed hf f' - · simp +instance (priority := 900) [G.IsCoverDense K] : G.IsDenseSubsite (G.inducedTopology K) K where + functorPushforward_mem_iff := Iff.rfl -/-- Given a fully faithful cover-dense functor `G : C ⥤ (D, K)`, we may induce a topology on `C`. --/ -abbrev Functor.inducedTopologyOfIsCoverDense [Full G] [Faithful G] [G.IsCoverDense K] : - GrothendieckTopology C := - (G.locallyCoverDense_of_isCoverDense K).inducedTopology +@[deprecated (since := "2024-07-23")] +alias inducedTopologyOfIsCoverDense := inducedTopology variable (J) -theorem over_forget_locallyCoverDense (X : C) : LocallyCoverDense J (Over.forget X) := by - intro Y T - convert T.property - ext Z f - constructor - · rintro ⟨_, _, g', hg, rfl⟩ - exact T.val.downward_closed hg g' - · intro hf - exact ⟨Over.mk (f ≫ Y.hom), Over.homMk f, 𝟙 _, hf, (Category.id_comp _).symm⟩ - -end - -section SmallSite - -variable {C : Type v} [SmallCategory C] {D : Type v} [SmallCategory D] {G : C ⥤ D} -variable {J : GrothendieckTopology C} {K : GrothendieckTopology D} -variable (A : Type u) [Category.{v} A] - -instance [G.Full] [G.Faithful] [G.IsCoverDense K] : - Functor.IsContinuous G (G.inducedTopologyOfIsCoverDense K) K := by - apply Functor.IsCoverDense.isContinuous - exact (G.locallyCoverDense_of_isCoverDense K).inducedTopology_coverPreserving - -instance [G.Full] [G.Faithful] [G.IsCoverDense K] : - Functor.IsCocontinuous G (G.inducedTopologyOfIsCoverDense K) K := - (G.locallyCoverDense_of_isCoverDense K).inducedTopology_isCocontinuous +instance over_forget_locallyCoverDense (X : C) : (Over.forget X).LocallyCoverDense J where + functorPushforward_functorPullback_mem Y T := by + convert T.property + ext Z f + constructor + · rintro ⟨_, _, g', hg, rfl⟩ + exact T.val.downward_closed hg g' + · intro hf + exact ⟨Over.mk (f ≫ Y.hom), Over.homMk f, 𝟙 _, hf, (Category.id_comp _).symm⟩ /-- Cover-dense functors induces an equivalence of categories of sheaves. This is known as the comparison lemma. It requires that the sites are small and the value category is complete. -/ -noncomputable def Functor.sheafInducedTopologyEquivOfIsCoverDense [Full G] [Faithful G] - [G.IsCoverDense K] [HasLimits A] : - Sheaf (G.inducedTopologyOfIsCoverDense K) A ≌ Sheaf K A := - Functor.IsCoverDense.sheafEquivOfCoverPreservingCoverLifting G - (G.inducedTopologyOfIsCoverDense K) K A +noncomputable def sheafInducedTopologyEquivOfIsCoverDense + [G.IsCoverDense K] [∀ (X : Dᵒᵖ), HasLimitsOfShape (StructuredArrow X G.op) A] : + Sheaf (G.inducedTopology K) A ≌ Sheaf K A := + Functor.IsDenseSubsite.sheafEquiv G + (G.inducedTopology K) K A -end SmallSite +end Functor end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean index 80d9d38f5a63b..ccd585a8efb58 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafFor.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafFor.lean @@ -676,15 +676,17 @@ lemma FamilyOfElements.isAmalgamation_iff_ofArrows (x : FamilyOfElements P (ofAr namespace Arrows.Compatible -variable {x : (i : I) → P.obj (op (X i))} (hx : Compatible P π x) +variable {x : (i : I) → P.obj (op (X i))} variable {P π} -theorem exists_familyOfElements : +theorem exists_familyOfElements (hx : Compatible P π x) : ∃ (x' : FamilyOfElements P (ofArrows X π)), ∀ (i : I), x' _ (ofArrows.mk i) = x i := by choose i h h' using @ofArrows_surj _ _ _ _ _ π exact ⟨fun Y f hf ↦ P.map (eqToHom (h f hf).symm).op (x _), fun j ↦ (hx _ j (X j) _ (𝟙 _) <| by rw [← h', id_comp]).trans <| by simp⟩ +variable (hx : Compatible P π x) + /-- A `FamilyOfElements` associated to an explicit family of elements. -/ diff --git a/Mathlib/CategoryTheory/Sites/Localization.lean b/Mathlib/CategoryTheory/Sites/Localization.lean index e419625fe9f13..d25e8fd6b17af 100644 --- a/Mathlib/CategoryTheory/Sites/Localization.lean +++ b/Mathlib/CategoryTheory/Sites/Localization.lean @@ -45,18 +45,19 @@ lemma W_sheafToPreheaf_map_iff_isIso {F₁ F₂ : Sheaf J A} (φ : F₁ ⟶ F₂ section Adjunction -variable {G : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A} (adj : G ⊣ sheafToPresheaf J A) +variable {G : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A} -lemma W_adj_unit_app (P : Cᵒᵖ ⥤ A) : J.W (adj.unit.app P) := by +lemma W_adj_unit_app (adj : G ⊣ sheafToPresheaf J A) (P : Cᵒᵖ ⥤ A) : J.W (adj.unit.app P) := by rw [W_eq_W_range_sheafToPresheaf_obj] exact LeftBousfield.W_adj_unit_app adj P -lemma W_iff_isIso_map_of_adjunction {P₁ P₂ : Cᵒᵖ ⥤ A} (f : P₁ ⟶ P₂) : +lemma W_iff_isIso_map_of_adjunction (adj : G ⊣ sheafToPresheaf J A) + {P₁ P₂ : Cᵒᵖ ⥤ A} (f : P₁ ⟶ P₂) : J.W f ↔ IsIso (G.map f) := by rw [W_eq_W_range_sheafToPresheaf_obj] exact LeftBousfield.W_iff_isIso_map adj f -lemma W_eq_inverseImage_isomorphisms_of_adjunction : +lemma W_eq_inverseImage_isomorphisms_of_adjunction (adj : G ⊣ sheafToPresheaf J A) : J.W = (MorphismProperty.isomorphisms _).inverseImage G := by rw [W_eq_W_range_sheafToPresheaf_obj, LeftBousfield.W_eq_inverseImage_isomorphisms adj] diff --git a/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean b/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean new file mode 100644 index 0000000000000..780045646d719 --- /dev/null +++ b/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.CategoryTheory.Sites.LocallySurjective + +/-! +# Locally fully faithful functors into sites + +## Main results + +- `CategoryTheory.Functor.IsLocallyFull`: + A functor `G : C ⥤ D` is locally full wrt a topology on `D` if for every `f : G.obj U ⟶ G.obj V`, + the set of `G.map fᵢ : G.obj Wᵢ ⟶ G.obj U` such that `G.map fᵢ ≫ f` is + in the image of `G` is a coverage of the topology on `D`. +- `CategoryTheory.Functor.IsLocallyFaithful`: + A functor `G : C ⥤ D` is locally faithful wrt a topology on `D` if for every `f₁ f₂ : U ⟶ V` whose + image in `D` are equal, the set of `G.map gᵢ : G.obj Wᵢ ⟶ G.obj U` such that `gᵢ ≫ f₁ = gᵢ ≫ f₂` + is a coverage of the topology on `D`. + +## References + +* [caramello2020]: Olivia Caramello, *Denseness conditions, morphisms and equivalences of toposes* + +-/ + +universe w vC vD uC uD + +namespace CategoryTheory + +variable {C : Type uC} [Category.{vC} C] {D : Type uD} [Category.{vD} D] (G : C ⥤ D) +variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) + +/-- +For a functor `G : C ⥤ D`, and a morphism `f : G.obj U ⟶ G.obj V`, +`Functor.imageSieve G f` is the sieve of `U` +consisting of those arrows whose composition with `f` has a lift in `G`. + +This is the image sieve of `f` under `yonedaMap G V` and hence the name. +See `Functor.imageSieve_eq_imageSieve`. +-/ +def Functor.imageSieve {U V : C} (f : G.obj U ⟶ G.obj V) : Sieve U where + arrows Y i := ∃ l, G.map l = G.map i ≫ f + downward_closed := by + rintro Y₁ Y₂ i₁ ⟨l, hl⟩ i₂ + exact ⟨i₂ ≫ l, by simp [hl]⟩ + +@[simp] +lemma Functor.imageSieve_map {U V : C} (f : U ⟶ V) : G.imageSieve (G.map f) = ⊤ := by + ext W g; simpa using ⟨g ≫ f, by simp⟩ + +/-- +For two arrows `f₁ f₂ : U ⟶ V`, the arrows `i` such that `i ≫ f₁ = i ≫ f₂` forms a sieve. +-/ +@[simps] +def Sieve.equalizer {U V : C} (f₁ f₂ : U ⟶ V) : Sieve U where + arrows Y i := i ≫ f₁ = i ≫ f₂ + downward_closed := by aesop + +@[simp] +lemma Sieve.equalizer_self {U V : C} (f : U ⟶ V) : equalizer f f = ⊤ := by ext; simp + +lemma Sieve.equalizer_eq_equalizerSieve {U V : C} (f₁ f₂ : U ⟶ V) : + Sieve.equalizer f₁ f₂ = Presheaf.equalizerSieve (F := yoneda.obj _) f₁ f₂ := rfl + +lemma Functor.imageSieve_eq_imageSieve {D : Type uD} [Category.{vC} D] (G : C ⥤ D) {U V : C} + (f : G.obj U ⟶ G.obj V) : + G.imageSieve f = Presheaf.imageSieve (yonedaMap G V) f := rfl + +open Presieve Opposite + +namespace Functor + +/-- +A functor `G : C ⥤ D` is locally full wrt a topology on `D` if for every `f : G.obj U ⟶ G.obj V`, +the set of `G.map fᵢ : G.obj Wᵢ ⟶ G.obj U` such that `G.map fᵢ ≫ f` is +in the image of `G` is a coverage of the topology on `D`. +-/ +class IsLocallyFull : Prop where + functorPushforward_imageSieve_mem : ∀ {U V} (f : G.obj U ⟶ G.obj V), + (G.imageSieve f).functorPushforward G ∈ K _ + +/-- +A functor `G : C ⥤ D` is locally faithful wrt a topology on `D` if for every `f₁ f₂ : U ⟶ V` whose +image in `D` are equal, the set of `G.map gᵢ : G.obj Wᵢ ⟶ G.obj U` such that `gᵢ ≫ f₁ = gᵢ ≫ f₂` +is a coverage of the topology on `D`. +-/ +class IsLocallyFaithful : Prop where + functorPushforward_equalizer_mem : ∀ {U V : C} (f₁ f₂ : U ⟶ V), G.map f₁ = G.map f₂ → + (Sieve.equalizer f₁ f₂).functorPushforward G ∈ K _ + +lemma functorPushforward_imageSieve_mem [G.IsLocallyFull K] {U V} (f : G.obj U ⟶ G.obj V) : + (G.imageSieve f).functorPushforward G ∈ K _ := + Functor.IsLocallyFull.functorPushforward_imageSieve_mem _ + +lemma functorPushforward_equalizer_mem + [G.IsLocallyFaithful K] {U V} (f₁ f₂ : U ⟶ V) (e : G.map f₁ = G.map f₂) : + (Sieve.equalizer f₁ f₂).functorPushforward G ∈ K _ := + Functor.IsLocallyFaithful.functorPushforward_equalizer_mem _ _ e + +variable {K} +variable {A : Type*} [Category A] (G : C ⥤ D) + +theorem IsLocallyFull.ext [G.IsLocallyFull K] + (ℱ : SheafOfTypes K) {X Y : C} (i : G.obj X ⟶ G.obj Y) + {s t : ℱ.val.obj (op (G.obj X))} + (h : ∀ ⦃Z : C⦄ (j : Z ⟶ X) (f : Z ⟶ Y), G.map f = G.map j ≫ i → + ℱ.1.map (G.map j).op s = ℱ.1.map (G.map j).op t) : s = t := by + apply (ℱ.cond _ (G.functorPushforward_imageSieve_mem K i)).isSeparatedFor.ext + rintro Z _ ⟨W, iWX, iZW, ⟨iWY, e⟩, rfl⟩ + simp [h iWX iWY e] + +theorem IsLocallyFaithful.ext [G.IsLocallyFaithful K] (ℱ : SheafOfTypes K) + {X Y : C} (i₁ i₂ : X ⟶ Y) (e : G.map i₁ = G.map i₂) + {s t : ℱ.val.obj (op (G.obj X))} + (h : ∀ ⦃Z : C⦄ (j : Z ⟶ X), j ≫ i₁ = j ≫ i₂ → + ℱ.1.map (G.map j).op s = ℱ.1.map (G.map j).op t) : s = t := by + apply (ℱ.cond _ (G.functorPushforward_equalizer_mem K i₁ i₂ e)).isSeparatedFor.ext + rintro Z _ ⟨W, iWX, iZW, hiWX, rfl⟩ + simp [h iWX hiWX] + +instance (priority := 900) IsLocallyFull.of_full [G.Full] : G.IsLocallyFull K where + functorPushforward_imageSieve_mem f := by + rw [← G.map_preimage f] + simp only [Functor.imageSieve_map, Sieve.functorPushforward_top, GrothendieckTopology.top_mem] + +instance (priority := 900) IsLocallyFaithful.of_faithful [G.Faithful] : G.IsLocallyFaithful K where + functorPushforward_equalizer_mem f₁ f₂ e := by obtain rfl := G.map_injective e; simp + +end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean new file mode 100644 index 0000000000000..b89884b96132b --- /dev/null +++ b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean @@ -0,0 +1,67 @@ +/- +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.Pullback.Square +import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.CategoryTheory.Sites.Sheafification + +/-! +# Mayer-Vietoris squares + +The purpose of this file is to allow the formalization of long exact +Mayer-Vietoris sequences in sheaf cohomology. If `X₄` is an open subset +of a topological space that is covered by two open subsets `X₂` and `X₃`, +it is known that there is a long exact sequence +`... ⟶ H^q(X₄) ⟶ H^q(X₂) ⊞ H^q(X₃) ⟶ H^q(X₁) ⟶ H^{q+1}(X₄) ⟶ ...` +where `X₁` is the intersection of `X₂` and `X₃`, and `H^q` are the +cohomology groups with values in an abelian sheaf. + +In this file, we introduce a structure +`GrothendieckTopology.MayerVietorisSquare` which extends `Squace C`, +and asserts properties which shall imply the existence of long +exact Mayer-Vietoris sequences in sheaf cohomology (TODO). +We require that the map `X₁ ⟶ X₃` is a monomorphism and +that the square in `C` becomes a pushout square in +the category of sheaves after the application of the +functor `yoneda ⋙ presheafToSheaf J _`. Note that in the +standard case of a covering by two open subsets, all +the morphisms in the square would be monomorphisms, +but this dissymetry allows the example of Nisnevich distinguished +squares in the case of the Nisnevich topology on schemes (in which case +`f₂₄ : X₂ ⟶ X₄` shall be an open immersion and +`f₃₄ : X₃ ⟶ X₄` an étale map that is an isomorphism over +the closed (reduced) subscheme `X₄ - X₂`, +and `X₁` shall be the pullback of `f₂₄` and `f₃₄`.). + +## TODO +* show that when "evaluating" a sheaf on a Mayer-Vietoris +square, one obtains a pullback square +* provide constructors for `MayerVietorisSquare` + +## References +* https://stacks.math.columbia.edu/tag/08GL + +-/ +universe v u + +namespace CategoryTheory + +namespace GrothendieckTopology + +variable {C : Type u} [Category.{v} C] + (J : GrothendieckTopology C) [HasWeakSheafify J (Type v)] + +/-- A Mayer-Vietoris square in a category `C` equipped with a Grothendieck +topology consists of a commutative square `f₁₂ ≫ f₂₄ = f₁₃ ≫ f₃₄` in `C` +such that `f₁₃` is a monomorphism and that the square becomes a +pushout square in the category of sheaves of sets. -/ +structure MayerVietorisSquare extends Square C where + mono_f₁₃ : Mono toSquare.f₁₃ := by infer_instance + /-- the square becomes a pushout square in the category of sheaves of types -/ + isPushout : (toSquare.map (yoneda ⋙ presheafToSheaf J _)).IsPushout + +end GrothendieckTopology + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/SheafCohomology/Basic.lean b/Mathlib/CategoryTheory/Sites/SheafCohomology/Basic.lean new file mode 100644 index 0000000000000..ba6a5261a40ff --- /dev/null +++ b/Mathlib/CategoryTheory/Sites/SheafCohomology/Basic.lean @@ -0,0 +1,84 @@ +/- +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.Category.Grp.Abelian +import Mathlib.Algebra.Category.Grp.Adjunctions +import Mathlib.Algebra.Homology.DerivedCategory.Ext +import Mathlib.CategoryTheory.Sites.Abelian +import Mathlib.CategoryTheory.Sites.ConstantSheaf + +/-! +# Sheaf cohomology + +Let `C` be a category equipped with a Grothendieck topology `J`. +We define the cohomology types `Sheaf.H F n` of an abelian +sheaf `F` on the site `(C, J)` for all `n : ℕ`. These abelian +groups are defined as the `Ext`-groups from the constant abelian +sheaf with values `ℤ` (actually `ULift ℤ`) to `F`. + +We also define `Sheaf.cohomologyPresheaf F n : Cᵒᵖ ⥤ AddCommGrp` +which is the presheaf which sends `U` to the `n`th `Ext`-group +from the free abelian sheaf generated by the presheaf +of sets `yoneda.obj U` to `F`. + +## TODO +* if `U` is a terminal object of `C`, define an isomorphism +`(F.cohomologyPresheaf n).obj (Opposite.op U) ≃+ Sheaf.H F n`. +* if `U : C`, define an isomorphism +`(F.cohomologyPresheaf n).obj (Opposite.op U) ≃+ Sheaf.H (F.over U) n`. + +-/ + +universe w' w v u + +namespace CategoryTheory + +open Abelian + +variable {C : Type u} [Category.{v} C] {J : GrothendieckTopology C} + +namespace Sheaf + +section + +variable (F : Sheaf J AddCommGrp.{w}) + [HasSheafify J AddCommGrp.{w}] [HasExt.{w'} (Sheaf J AddCommGrp.{w})] + +/-- The cohomology of an abelian sheaf in degree `n`. -/ +def H (n : ℕ) : Type w' := + Ext ((constantSheaf J AddCommGrp.{w}).obj (AddCommGrp.of (ULift ℤ))) F n + +noncomputable instance (n : ℕ) : AddCommGroup (F.H n) := by + dsimp only [H] + infer_instance + +end + +section + +variable [HasSheafify J AddCommGrp.{v}] [HasExt.{w'} (Sheaf J AddCommGrp.{v})] + +variable (J) in +/-- The bifunctor which sends an abelian sheaf `F` and an object `U` to the +`n`th Ext-group from the free abelian sheaf generated by the +presheaf of sets `yoneda.obj U` to `F`. -/ +noncomputable def cohomologyPresheafFunctor (n : ℕ) : + Sheaf J AddCommGrp.{v} ⥤ Cᵒᵖ ⥤ AddCommGrp.{w'} := + Functor.flip + (Functor.op (yoneda ⋙ (whiskeringRight _ _ _).obj + AddCommGrp.free ⋙ presheafToSheaf _ _) ⋙ extFunctor n) + +/-- Given an abelian sheaf `F`, this is the presheaf which sends `U` +to the `n`th Ext-group from the free abelian sheaf generated by the +presheaf of sets `yoneda.obj U` to `F`. -/ +noncomputable abbrev cohomologyPresheaf (F : Sheaf J AddCommGrp.{v}) (n : ℕ) : + Cᵒᵖ ⥤ AddCommGrp.{w'} := + (cohomologyPresheafFunctor J n).obj F + +end + +end Sheaf + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index 4d301502130fd..aa647e488a110 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -414,6 +414,13 @@ theorem generate_of_singleton_isSplitEpi (f : Y ⟶ X) [IsSplitEpi f] : theorem generate_top : generate (⊤ : Presieve X) = ⊤ := generate_of_contains_isSplitEpi (𝟙 _) ⟨⟩ +@[simp] +lemma comp_mem_iff (i : X ⟶ Y) (f : Y ⟶ Z) [IsIso i] (S : Sieve Z) : + S (i ≫ f) ↔ S f := by + refine ⟨fun H ↦ ?_, fun H ↦ S.downward_closed H _⟩ + convert S.downward_closed H (inv i) + simp + /-- The sieve of `X` generated by family of morphisms `Y i ⟶ X`. -/ abbrev ofArrows {I : Type*} {X : C} (Y : I → C) (f : ∀ i, Y i ⟶ X) : Sieve X := @@ -705,7 +712,7 @@ def essSurjFullFunctorGaloisInsertion [F.EssSurj] [F.Full] (X : C) : apply (functor_galoisConnection F X).toGaloisInsertion intro S Y f hf refine ⟨_, F.preimage ((F.objObjPreimageIso Y).hom ≫ f), (F.objObjPreimageIso Y).inv, ?_⟩ - simpa using S.downward_closed hf _ + simpa using hf /-- When `F` is fully faithful, the galois connection is a galois coinsertion. -/ def fullyFaithfulFunctorGaloisCoinsertion [F.Full] [F.Faithful] (X : C) : @@ -717,6 +724,29 @@ def fullyFaithfulFunctorGaloisCoinsertion [F.Full] [F.Faithful] (X : C) : rw [F.map_injective h₂] exact S.downward_closed h₁ _ +lemma functorPushforward_functor (S : Sieve X) (e : C ≌ D) : + S.functorPushforward e.functor = (S.pullback (e.unitInv.app X)).functorPullback e.inverse := by + ext Y iYX + constructor + · rintro ⟨Z, iZX, iYZ, hiZX, rfl⟩ + simpa using S.downward_closed hiZX (e.inverse.map iYZ ≫ e.unitInv.app Z) + · intro H + exact ⟨_, e.inverse.map iYX ≫ e.unitInv.app X, e.counitInv.app Y, by simpa using H, by simp⟩ + +@[simp] +lemma mem_functorPushforward_functor {Y : D} {S : Sieve X} {e : C ≌ D} {f : Y ⟶ e.functor.obj X} : + S.functorPushforward e.functor f ↔ S (e.inverse.map f ≫ e.unitInv.app X) := + congr($(S.functorPushforward_functor e).arrows f) + +lemma functorPushforward_inverse {X : D} (S : Sieve X) (e : C ≌ D) : + S.functorPushforward e.inverse = (S.pullback (e.counit.app X)).functorPullback e.functor := + Sieve.functorPushforward_functor S e.symm + +@[simp] +lemma mem_functorPushforward_inverse {X : D} {S : Sieve X} {e : C ≌ D} {f : Y ⟶ e.inverse.obj X} : + S.functorPushforward e.inverse f ↔ S (e.functor.map f ≫ e.counit.app X) := + congr($(S.functorPushforward_inverse e).arrows f) + end Functor /-- A sieve induces a presheaf. -/ diff --git a/Mathlib/CategoryTheory/SmallObject/Construction.lean b/Mathlib/CategoryTheory/SmallObject/Construction.lean index 80a69bf63bdba..0ec6400d63876 100644 --- a/Mathlib/CategoryTheory/SmallObject/Construction.lean +++ b/Mathlib/CategoryTheory/SmallObject/Construction.lean @@ -62,7 +62,7 @@ variable {C : Type u} [Category.{v} C] {I : Type w} {A B : I → C} (f : ∀ i, section -variable {S : C} {X Y Z : C} (πX : X ⟶ S) (πY : Y ⟶ S) (φ : X ⟶ Y) (hφ : φ ≫ πY = πX) +variable {S : C} {X Y Z : C} (πX : X ⟶ S) (πY : Y ⟶ S) (φ : X ⟶ Y) /-- Given a family of morphisms `f i : A i ⟶ B i` and a morphism `πX : X ⟶ S`, this type parametrizes the commutative squares with a morphism `f i` on the left @@ -143,10 +143,12 @@ lemma ιFunctorObj_πFunctorObj : ιFunctorObj f πX ≫ πFunctorObj f πX = π /-- The canonical morphism `∐ (functorObjSrcFamily f πX) ⟶ ∐ (functorObjSrcFamily f πY)` induced by a morphism in `φ : X ⟶ Y` such that `φ ≫ πX = πY`. -/ -noncomputable def functorMapSrc : +noncomputable def functorMapSrc (hφ : φ ≫ πY = πX) : ∐ (functorObjSrcFamily f πX) ⟶ ∐ functorObjSrcFamily f πY := Sigma.map' (fun x => FunctorObjIndex.mk x.i (x.t ≫ φ) x.b (by simp [hφ])) (fun _ => 𝟙 _) +variable (hφ : φ ≫ πY = πX) + @[reassoc] lemma ι_functorMapSrc (i : I) (t : A i ⟶ X) (b : B i ⟶ S) (w : t ≫ πX = f i ≫ b) (t' : A i ⟶ Y) (fac : t ≫ φ = t') : @@ -164,7 +166,7 @@ lemma functorMapSrc_functorObjTop : /-- The canonical morphism `∐ functorObjTgtFamily f πX ⟶ ∐ functorObjTgtFamily f πY` induced by a morphism in `φ : X ⟶ Y` such that `φ ≫ πX = πY`. -/ -noncomputable def functorMapTgt : +noncomputable def functorMapTgt (hφ : φ ≫ πY = πX) : ∐ functorObjTgtFamily f πX ⟶ ∐ functorObjTgtFamily f πY := Sigma.map' (fun x => FunctorObjIndex.mk x.i (x.t ≫ φ) x.b (by simp [hφ])) (fun _ => 𝟙 _) diff --git a/Mathlib/CategoryTheory/Square.lean b/Mathlib/CategoryTheory/Square.lean index 0e7335ceb4b16..53622ff002ab1 100644 --- a/Mathlib/CategoryTheory/Square.lean +++ b/Mathlib/CategoryTheory/Square.lean @@ -35,13 +35,13 @@ maps (`arrowArrowEquivalence'`). -/ -universe v u +universe v v' u u' namespace CategoryTheory open Category -variable (C : Type u) [Category.{v} C] +variable (C : Type u) [Category.{v} C] {D : Type u'} [Category.{v'} D] /-- The category of commutative squares in a category. -/ structure Square where @@ -265,6 +265,33 @@ def evaluation₄ : Square C ⥤ C where obj sq := sq.X₄ map φ := φ.τ₄ +/-- The image of a commutative square by a functor. -/ +@[simps] +def map (sq : Square C) (F : C ⥤ D) : Square D where + f₁₂ := F.map sq.f₁₂ + f₁₃ := F.map sq.f₁₃ + f₂₄ := F.map sq.f₂₄ + f₃₄ := F.map sq.f₃₄ + fac := by simpa using F.congr_map sq.fac + end Square +namespace Functor + +/-- The functor `Square C ⥤ Square D` induced by a functor `C ⥤ D`. -/ +@[simps] +def mapSquare (F : C ⥤ D) : Square C ⥤ Square D where + obj sq := sq.map F + map φ := + { τ₁ := F.map φ.τ₁ + τ₂ := F.map φ.τ₂ + τ₃ := F.map φ.τ₃ + τ₄ := F.map φ.τ₄ + comm₁₂ := by simpa only [Functor.map_comp] using F.congr_map φ.comm₁₂ + comm₁₃ := by simpa only [Functor.map_comp] using F.congr_map φ.comm₁₃ + comm₂₄ := by simpa only [Functor.map_comp] using F.congr_map φ.comm₂₄ + comm₃₄ := by simpa only [Functor.map_comp] using F.congr_map φ.comm₃₄ } + +end Functor + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index af564adfb94e9..f3cbf94ce4ae3 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -611,13 +611,12 @@ variable {D : Type*} [Category.{v₁} D] (F : C ⥤ D) /-- The natural transformation `yoneda.obj X ⟶ F.op ⋙ yoneda.obj (F.obj X)` when `F : C ⥤ D` and `X : C`. -/ -def yonedaMap (X : C) : yoneda.obj X ⟶ F.op ⋙ yoneda.obj (F.obj X) := - yonedaEquiv.symm (𝟙 _) +def yonedaMap (X : C) : yoneda.obj X ⟶ F.op ⋙ yoneda.obj (F.obj X) where + app X f := F.map f @[simp] lemma yonedaMap_app_apply {Y : C} {X : Cᵒᵖ} (f : X.unop ⟶ Y) : - (yonedaMap F Y).app X f = F.map f := by - simp [yonedaMap, yonedaEquiv] + (yonedaMap F Y).app X f = F.map f := rfl end diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean new file mode 100644 index 0000000000000..09fd0c79b7e07 --- /dev/null +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -0,0 +1,203 @@ +/- +Copyright (c) 2023 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Data.Multiset.Fintype +import Mathlib.FieldTheory.ChevalleyWarning +import Mathlib.RingTheory.UniqueFactorizationDomain + +/-! +# The Erdős–Ginzburg–Ziv theorem + +This file proves the Erdős–Ginzburg–Ziv theorem as a corollary of Chevalley-Warning. This theorem +states that among any (not necessarily distinct) `2 * n - 1` elements of `ZMod n`, we can find `n` +elements of sum zero. + +## Main declarations + +* `Int.erdos_ginzburg_ziv`: The Erdős–Ginzburg–Ziv theorem stated using sequences in `ℤ` +* `ZMod.erdos_ginzburg_ziv`: The Erdős–Ginzburg–Ziv theorem stated using sequences in `ZMod n` +-/ + +open Finset MvPolynomial +open scoped BigOperators + +variable {ι : Type*} + +section prime +variable {p : ℕ} [Fact p.Prime] {s : Finset ι} + +set_option linter.unusedVariables false in +/-- The first multivariate polynomial used in the proof of Erdős–Ginzburg–Ziv. -/ +private noncomputable def f₁ (s : Finset ι) (a : ι → ZMod p) : MvPolynomial s (ZMod p) := + ∑ i, X i ^ (p - 1) + +/-- The second multivariate polynomial used in the proof of Erdős–Ginzburg–Ziv. -/ +private noncomputable def f₂ (s : Finset ι) (a : ι → ZMod p) : MvPolynomial s (ZMod p) := + ∑ i : s, a i • X i ^ (p - 1) + +private lemma totalDegree_f₁_add_totalDegree_f₂ {a : ι → ZMod p} : + (f₁ s a).totalDegree + (f₂ s a).totalDegree < 2 * p - 1 := by + calc + _ ≤ (p - 1) + (p - 1) := by + gcongr <;> apply totalDegree_finsetSum_le <;> rintro i _ + · exact (totalDegree_X_pow ..).le + · exact (totalDegree_smul_le ..).trans (totalDegree_X_pow ..).le + _ < 2 * p - 1 := by have := (Fact.out : p.Prime).two_le; omega + +/-- The prime case of the **Erdős–Ginzburg–Ziv theorem** for `ℤ/pℤ`. + +Any sequence of `2 * p - 1` elements of `ZMod p` contains a subsequence of `p` elements whose sum is +zero. -/ +private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : s.card = 2 * p - 1) : + ∃ t ⊆ s, t.card = p ∧ ∑ i ∈ t, a i = 0 := by + haveI : NeZero p := inferInstance + classical + -- Let `N` be the number of common roots of our polynomials `f₁` and `f₂` (`f s ff` and `f s tt`). + set N := Fintype.card {x // eval x (f₁ s a) = 0 ∧ eval x (f₂ s a) = 0} + -- Zero is a common root to `f₁` and `f₂`, so `N` is nonzero + let zero_sol : {x // eval x (f₁ s a) = 0 ∧ eval x (f₂ s a) = 0} := + ⟨0, by simp [f₁, f₂, map_sum, (Fact.out : p.Prime).one_lt, tsub_eq_zero_iff_le]⟩ + have hN₀ : 0 < N := @Fintype.card_pos _ _ ⟨zero_sol⟩ + have hs' : 2 * p - 1 = Fintype.card s := by simp [hs] + -- Chevalley-Warning gives us that `p ∣ n` because the total degrees of `f₁` and `f₂` are at most + -- `p - 1`, and we have `2 * p - 1 > 2 * (p - 1)` variables. + have hpN : p ∣ N := char_dvd_card_solutions_of_add_lt p + (totalDegree_f₁_add_totalDegree_f₂.trans_eq hs') + -- Hence, `2 ≤ p ≤ N` and we can make a common root `x ≠ 0`. + obtain ⟨x, hx⟩ := Fintype.exists_ne_of_one_lt_card ((Fact.out : p.Prime).one_lt.trans_le $ + Nat.le_of_dvd hN₀ hpN) zero_sol + -- This common root gives us the required subsequence, namely the `i ∈ s` such that `x i ≠ 0`. + refine ⟨(s.attach.filter fun a ↦ x.1 a ≠ 0).map ⟨(↑), Subtype.val_injective⟩, ?_, ?_, ?_⟩ + · simp (config := { contextual := true }) [subset_iff] + -- From `f₁ x = 0`, we get that `p` divides the number of `a` such that `x a ≠ 0`. + · rw [card_map] + refine Nat.eq_of_dvd_of_lt_two_mul (Finset.card_pos.2 ?_).ne' ?_ $ + (Finset.card_filter_le _ _).trans_lt ?_ + -- This number is nonzero because `x ≠ 0`. + · rw [← Subtype.coe_ne_coe, Function.ne_iff] at hx + exact hx.imp (fun a ha ↦ mem_filter.2 ⟨Finset.mem_attach _ _, ha⟩) + · rw [← CharP.cast_eq_zero_iff (ZMod p), ← Finset.sum_boole] + simpa only [f₁, map_sum, ZMod.pow_card_sub_one, map_pow, eval_X] using x.2.1 + -- And it is at most `2 * p - 1`, so it must be `p`. + · rw [Finset.card_attach, hs] + exact tsub_lt_self (mul_pos zero_lt_two (Fact.out : p.Prime).pos) zero_lt_one + -- From `f₂ x = 0`, we get that `p` divides the sum of the `a ∈ s` such that `x a ≠ 0`. + · simpa [f₂, ZMod.pow_card_sub_one, Finset.sum_filter] using x.2.2 + +/-- The prime case of the **Erdős–Ginzburg–Ziv theorem** for `ℤ`. + +Any sequence of `2 * p - 1` elements of `ℤ` contains a subsequence of `p` elements whose sum is +divisible by `p`. -/ +private theorem Int.erdos_ginzburg_ziv_prime (a : ι → ℤ) (hs : s.card = 2 * p - 1) : + ∃ t ⊆ s, t.card = p ∧ ↑p ∣ ∑ i ∈ t, a i := by + simpa [← Int.cast_sum, ZMod.intCast_zmod_eq_zero_iff_dvd] + using ZMod.erdos_ginzburg_ziv_prime (Int.cast ∘ a) hs + +end prime + +section composite +variable {n : ℕ} {s : Finset ι} + +/-- The **Erdős–Ginzburg–Ziv theorem** for `ℤ`. + +Any sequence of at least `2 * n - 1` elements of `ℤ` contains a subsequence of `n` elements whose +sum is divisible by `n`. -/ +theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : + ∃ t ⊆ s, t.card = n ∧ ↑n ∣ ∑ i ∈ t, a i := by + classical + -- Do induction on the prime factorisation of `n`. Note that we will apply the induction + -- hypothesis with `ι := Finset ι`, so we need to generalise. + induction n using Nat.prime_composite_induction generalizing ι + -- When `n := 0`, we can set `t := ∅`. + case zero => exact ⟨∅, by simp⟩ + -- When `n := 1`, we can take `t` to be any subset of `s` of size `2 * n - 1`. + case one => simpa using exists_subset_card_eq hs + -- When `n := p` is prime, we use the prime case `Int.erdos_ginzburg_ziv_prime`. + case prime p hp => + haveI := Fact.mk hp + obtain ⟨t, hts, ht⟩ := exists_subset_card_eq hs + obtain ⟨u, hut, hu⟩ := Int.erdos_ginzburg_ziv_prime a ht + exact ⟨u, hut.trans hts, hu⟩ + -- When `n := m * n` is composite, we pick (by induction hypothesis on `n`) `2 * m - 1` sets of + -- size `n` and sums divisible by `n`. Then by induction hypothesis (on `m`) we can pick `m` of + -- these sets whose sum is divisible by `m * n`. + case composite m hm ihm n hn ihn => + -- First, show that it is enough to have those `2 * m - 1` sets. + suffices ∀ k ≤ 2 * m - 1, ∃ 𝒜 : Finset (Finset ι), 𝒜.card = k ∧ + (𝒜 : Set (Finset ι)).Pairwise _root_.Disjoint ∧ + ∀ ⦃t⦄, t ∈ 𝒜 → t ⊆ s ∧ t.card = n ∧ ↑n ∣ ∑ i ∈ t, a i by + -- Assume `𝒜` is a family of `2 * m - 1` sets, each of size `n` and sum divisible by `n`. + obtain ⟨𝒜, h𝒜card, h𝒜disj, h𝒜⟩ := this _ le_rfl + -- By induction hypothesis on `m`, find a subfamily `ℬ` of size `m` such that the sum over + -- `t ∈ ℬ` of `(∑ i ∈ t, a i) / n` is divisible by `m`. + obtain ⟨ℬ, hℬ𝒜, hℬcard, hℬ⟩ := ihm (fun t ↦ (∑ i ∈ t, a i) / n) h𝒜card.ge + -- We are done. + refine ⟨ℬ.biUnion fun x ↦ x, biUnion_subset.2 fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).1, ?_, ?_⟩ + · rw [card_biUnion (h𝒜disj.mono hℬ𝒜), sum_const_nat fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).2.1, hℬcard] + rwa [sum_biUnion, natCast_mul, mul_comm, ← Int.dvd_div_iff_mul_dvd, Int.sum_div] + · exact fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).2.2 + · exact dvd_sum fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).2.2 + · exact h𝒜disj.mono hℬ𝒜 + -- Now, let's find those `2 * m - 1` sets. + rintro k hk + -- We induct on the size `k ≤ 2 * m - 1` of the family we are constructing. + induction' k with k ih + -- For `k = 0`, the empty family trivially works. + · exact ⟨∅, by simp⟩ + -- At `k + 1`, call `𝒜` the existing family of size `k ≤ 2 * m - 2`. + obtain ⟨𝒜, h𝒜card, h𝒜disj, h𝒜⟩ := ih (Nat.le_of_succ_le hk) + -- There are at least `2 * (m * n) - 1 - k * n ≥ 2 * m - 1` elements in `s` that have not been + -- taken in any element of `𝒜`. + have : 2 * n - 1 ≤ (s \ 𝒜.biUnion id).card := by + calc + _ ≤ (2 * m - k) * n - 1 := by gcongr; omega + _ = (2 * (m * n) - 1) - ∑ t ∈ 𝒜, t.card := by + rw [tsub_mul, mul_assoc, tsub_right_comm, sum_const_nat fun t ht ↦ (h𝒜 ht).2.1, h𝒜card] + _ ≤ s.card - (𝒜.biUnion id).card := by gcongr; exact card_biUnion_le + _ ≤ (s \ 𝒜.biUnion id).card := le_card_sdiff .. + -- So by the induction hypothesis on `n` we can find a new set `t` of size `n` and sum divisible + -- by `n`. + obtain ⟨t₀, ht₀, ht₀card, ht₀sum⟩ := ihn a this + -- This set is distinct and disjoint from the previous ones, so we are done. + have : t₀ ∉ 𝒜 := by + rintro h + obtain rfl : n = 0 := by + simpa [← card_eq_zero, ht₀card] using sdiff_disjoint.mono ht₀ $ subset_biUnion_of_mem id h + omega + refine ⟨𝒜.cons t₀ this, by rw [card_cons, h𝒜card], ?_, ?_⟩ + · simp only [cons_eq_insert, coe_insert, Set.pairwise_insert_of_symmetric symmetric_disjoint, + mem_coe, ne_eq] + exact ⟨h𝒜disj, fun t ht _ ↦ sdiff_disjoint.mono ht₀ $ subset_biUnion_of_mem id ht⟩ + · simp only [cons_eq_insert, mem_insert, forall_eq_or_imp, and_assoc] + exact ⟨ht₀.trans sdiff_subset, ht₀card, ht₀sum, h𝒜⟩ + +/-- The **Erdős–Ginzburg–Ziv theorem** for `ℤ/nℤ`. + +Any sequence of at least `2 * n - 1` elements of `ZMod n` contains a subsequence of `n` elements +whose sum is zero. -/ +theorem ZMod.erdos_ginzburg_ziv (a : ι → ZMod n) (hs : 2 * n - 1 ≤ s.card) : + ∃ t ⊆ s, t.card = n ∧ ∑ i ∈ t, a i = 0 := by + simpa [← ZMod.intCast_zmod_eq_zero_iff_dvd] using Int.erdos_ginzburg_ziv (ZMod.cast ∘ a) hs + +/-- The **Erdős–Ginzburg–Ziv theorem** for `ℤ` for multiset. + +Any multiset of at least `2 * n - 1` elements of `ℤ` contains a submultiset of `n` elements whose +sum is divisible by `n`. -/ +theorem Int.erdos_ginzburg_ziv_multiset (s : Multiset ℤ) (hs : 2 * n - 1 ≤ Multiset.card s) : + ∃ t ≤ s, Multiset.card t = n ∧ ↑n ∣ t.sum := by + obtain ⟨t, hts, ht⟩ := Int.erdos_ginzburg_ziv (s := s.toEnumFinset) Prod.fst (by simpa using hs) + exact ⟨t.1.map Prod.fst, Multiset.map_fst_le_of_subset_toEnumFinset hts, by simpa using ht⟩ + +/-- The **Erdős–Ginzburg–Ziv theorem** for `ℤ/nℤ` for multiset. + +Any multiset of at least `2 * n - 1` elements of `ℤ` contains a submultiset of `n` elements whose +sum is divisible by `n`. -/ +theorem ZMod.erdos_ginzburg_ziv_multiset (s : Multiset (ZMod n)) + (hs : 2 * n - 1 ≤ Multiset.card s) : ∃ t ≤ s, Multiset.card t = n ∧ t.sum = 0 := by + obtain ⟨t, hts, ht⟩ := ZMod.erdos_ginzburg_ziv (s := s.toEnumFinset) Prod.fst (by simpa using hs) + exact ⟨t.1.map Prod.fst, Multiset.map_fst_le_of_subset_toEnumFinset hts, by simpa using ht⟩ + +end composite diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index f9b5c18959239..401c788269c57 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -87,7 +87,6 @@ Composition, partition -/ - open List variable {n : ℕ} @@ -188,7 +187,7 @@ theorem sizeUpTo_zero : c.sizeUpTo 0 = 0 := by simp [sizeUpTo] theorem sizeUpTo_ofLength_le (i : ℕ) (h : c.length ≤ i) : c.sizeUpTo i = n := by dsimp [sizeUpTo] convert c.blocks_sum - exact take_all_of_le h + exact take_of_length_le h @[simp] theorem sizeUpTo_length : c.sizeUpTo c.length = n := @@ -259,7 +258,7 @@ theorem orderEmbOfFin_boundaries : def embedding (i : Fin c.length) : Fin (c.blocksFun i) ↪o Fin n := (Fin.natAddOrderEmb <| c.sizeUpTo i).trans <| Fin.castLEOrderEmb <| calc - c.sizeUpTo i + c.blocksFun i = c.sizeUpTo (i + 1) := (c.sizeUpTo_succ _).symm + c.sizeUpTo i + c.blocksFun i = c.sizeUpTo (i + 1) := (c.sizeUpTo_succ i.2).symm _ ≤ c.sizeUpTo c.length := monotone_sum_take _ i.2 _ = n := c.sizeUpTo_length @@ -622,7 +621,7 @@ theorem sum_take_map_length_splitWrtComposition (l : List α) (c : Composition l theorem getElem_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi : i < (l.splitWrtCompositionAux ns).length) : - (l.splitWrtCompositionAux ns)[i] = + (l.splitWrtCompositionAux ns)[i] = (l.take (ns.take (i + 1)).sum).drop (ns.take i).sum := by induction' ns with n ns IH generalizing l i · cases hi @@ -639,13 +638,13 @@ block of the composition. -/ theorem getElem_splitWrtComposition' (l : List α) (c : Composition n) {i : ℕ} (hi : i < (l.splitWrtComposition c).length) : (l.splitWrtComposition c)[i] = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := - getElem_splitWrtCompositionAux _ _ _ + getElem_splitWrtCompositionAux _ _ hi -- Porting note: restatement of `get_splitWrtComposition` theorem getElem_splitWrtComposition (l : List α) (c : Composition n) (i : Nat) (h : i < (l.splitWrtComposition c).length) : (l.splitWrtComposition c)[i] = (l.take (c.sizeUpTo (i + 1))).drop (c.sizeUpTo i) := - getElem_splitWrtComposition' _ _ _ + getElem_splitWrtComposition' _ _ h @[deprecated getElem_splitWrtCompositionAux (since := "2024-06-12")] theorem get_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi) : @@ -869,7 +868,7 @@ theorem mem_boundaries_iff_exists_blocks_sum_take_eq {j : Fin (n + 1)} : exact this.symm theorem blocks_sum : c.blocks.sum = n := by - have : c.blocks.take c.length = c.blocks := take_all_of_le (by simp [blocks]) + have : c.blocks.take c.length = c.blocks := take_of_length_le (by simp [blocks]) rw [← this, c.blocks_partial_sum c.length_lt_card_boundaries, c.boundary_length] rfl diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 3abea4837ef24..d88c1fba8060d 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -92,6 +92,34 @@ def ofSums (n : ℕ) (l : Multiset ℕ) (hl : l.sum = n) : Partition n where @[simps!] def ofMultiset (l : Multiset ℕ) : Partition l.sum := ofSums _ l rfl +/-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. -/ +def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition where + parts := s.1.dedup.map s.1.count + parts_pos := by simp [Multiset.count_pos] + parts_sum := by + show ∑ a ∈ s.1.toFinset, count a s.1 = n + rw [toFinset_sum_count_eq] + exact s.2 + +variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] + +@[simp] lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : + ofSym (s.map e) = ofSym s := by + simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] + rw [Multiset.dedup_map_of_injective e.injective] + simp only [map_map, Function.comp_apply] + congr; funext i + rw [← Multiset.count_map_eq_count' e _ e.injective] + +/-- An equivalence between `σ` and `τ` induces an equivalence between the subtypes of `Sym σ n` and +`Sym τ n` corresponding to a given partition. -/ +def ofSymShapeEquiv (μ : Partition n) (e : σ ≃ τ) : + {x : Sym σ n // ofSym x = μ} ≃ {x : Sym τ n // ofSym x = μ} where + toFun := fun x => ⟨Sym.equivCongr e x, by simp [ofSym_map, x.2]⟩ + invFun := fun x => ⟨Sym.equivCongr e.symm x, by simp [ofSym_map, x.2]⟩ + left_inv := by intro x; simp + right_inv := by intro x; simp + /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl @@ -115,6 +143,9 @@ instance UniquePartitionZero : Unique (Partition 0) where instance UniquePartitionOne : Unique (Partition 1) where uniq _ := Partition.ext <| by simp +@[simp] lemma ofSym_one (s : Sym σ 1) : ofSym s = indiscrete 1 := by + ext; simp + /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. (For `i = 0`, `Partition.non_zero` combined with `Multiset.count_eq_zero_of_not_mem` gives that diff --git a/Mathlib/Combinatorics/HalesJewett.lean b/Mathlib/Combinatorics/HalesJewett.lean index 50203d214f81b..1c5d874c88cde 100644 --- a/Mathlib/Combinatorics/HalesJewett.lean +++ b/Mathlib/Combinatorics/HalesJewett.lean @@ -4,14 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn -/ import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Data.Countable.Small import Mathlib.Data.Fintype.Option import Mathlib.Data.Fintype.Pi +import Mathlib.Data.Fintype.Prod +import Mathlib.Data.Fintype.Shrink import Mathlib.Data.Fintype.Sum /-! # The Hales-Jewett theorem -We prove the Hales-Jewett theorem and deduce Van der Waerden's theorem as a corollary. +We prove the Hales-Jewett theorem. We deduce Van der Waerden's theorem and the multidimensional +Hales-Jewett theorem as corollaries. The Hales-Jewett theorem is a result in Ramsey theory dealing with *combinatorial lines*. Given an 'alphabet' `α : Type*` and `a b : α`, an example of a combinatorial line in `α^5` is @@ -22,6 +26,10 @@ huge) finite type `ι` such that whenever `ι → α` is `κ`-colored (i.e. for the idea of *color focusing* and a *product argument*. See the proof of `Combinatorics.Line.exists_mono_in_high_dimension'` for details. +*Combinatorial subspaces* are higher-dimensional analogues of combinatorial lines. See +`Combinatorics.Subspace`. The multidimensional Hales-Jewett theorem generalises the statement above +from combinatorial lines to combinatorial subspaces of a fixed dimension. + The version of Van der Waerden's theorem in this file states that whenever a commutative monoid `M` is finitely colored and `S` is a finite subset, there exists a monochromatic homothetic copy of `S`. This follows from the Hales-Jewett theorem by considering the map `(ι → S) → M` sending `v` @@ -29,8 +37,9 @@ to `∑ i : ι, v i`, which sends a combinatorial line to a homothetic copy of ` ## Main results -- `Combinatorics.Line.exists_mono_in_high_dimension`: the Hales-Jewett theorem. -- `Combinatorics.exists_mono_homothetic_copy`: a generalization of Van der Waerden's theorem. +- `Combinatorics.Line.exists_mono_in_high_dimension`: The Hales-Jewett theorem. +- `Combinatorics.Subspace.exists_mono_in_high_dimension`: The multidimensional Hales-Jewett theorem. +- `Combinatorics.exists_mono_homothetic_copy`: A generalization of Van der Waerden's theorem. ## Implementation details @@ -57,13 +66,98 @@ combinatorial line, Ramsey theory, arithmetic progression -/ - +open Function open scoped Classical universe u v +variable {η α ι κ : Type*} namespace Combinatorics +/-- The type of combinatorial subspaces. A subspace `l : Subspace η α ι` in the hypercube `ι → α` +defines a function `(η → α) → ι → α` from `η → α` to the hypercube, such that for each coordinate +`i : ι` and direction `e : η`, the function `fun x ↦ l x i` is either `fun x ↦ x e` for some +direction `e : η` or constant. We require subspaces to be non-degenerate in the sense that, for +every `e : η`, `fun x ↦ l x i` is `fun x ↦ x e` for at least one `i`. + +Formally, a subspace is represented by a word `l.idxFun : ι → α ⊕ η` which says whether +`fun x ↦ l x i` is `fun x ↦ x e` (corresponding to `l.idxFun i = Sum.inr e`) or constantly `a` +(corresponding to `l.idxFun i = Sum.inl a`). + +When `α` has size `1` there can be many elements of `Subspace η α ι` defining the same function. -/ +@[ext] +structure Subspace (η α ι : Type*) where + /-- The word representing a combinatorial subspace. `l.idxfun i = Sum.inr e` means that + `l x i = x e` for all `x` and `l.idxfun i = some a` means that `l x i = a` for all `x`. -/ + idxFun : ι → α ⊕ η + /-- We require combinatorial subspaces to be nontrivial in the sense that `fun x ↦ l x i` is + `fun x ↦ x e` for at least one coordinate `i`. -/ + proper : ∀ e, ∃ i, idxFun i = Sum.inr e + +namespace Subspace +variable {η α ι κ : Type*} {l : Subspace η α ι} {x : η → α} {i : ι} {a : α} {e : η} + +/-- The combinatorial subspace corresponding to the identity embedding `(ι → α) → (ι → α)`. -/ +instance : Inhabited (Subspace ι α ι) := ⟨⟨Sum.inr, fun i ↦ ⟨i, rfl⟩⟩⟩ + +/-- Consider a subspace `l : Subspace η α ι` as a function `(η → α) → ι → α`. -/ +@[coe] def toFun (l : Subspace η α ι) (x : η → α) (i : ι) : α := (l.idxFun i).elim id x + +instance instCoeFun : CoeFun (Subspace η α ι) (fun _ ↦ (η → α) → ι → α) := ⟨toFun⟩ + +lemma coe_apply (l : Subspace η α ι) (x : η → α) (i : ι) : l x i = (l.idxFun i).elim id x := rfl + +-- Note: This is not made a `FunLike` instance to avoid having two syntactically different coercions +lemma coe_injective [Nontrivial α] : Injective ((⇑) : Subspace η α ι → (η → α) → ι → α) := by + rintro l m hlm + ext i + simp only [funext_iff] at hlm + cases hl : idxFun l i with + | inl a => + obtain ⟨b, hba⟩ := exists_ne a + cases hm : idxFun m i <;> simpa [hl, hm, hba.symm, coe_apply] using hlm (const _ b) i + | inr e => + cases hm : idxFun m i with + | inl a => + obtain ⟨b, hba⟩ := exists_ne a + simpa [hl, hm, hba, coe_apply] using hlm (const _ b) i + | inr f => + obtain ⟨a, b, hab⟩ := exists_pair_ne α + simp only [Sum.inr.injEq] + by_contra! hef + simpa [hl, hm, hef, hab, coe_apply] using hlm (Function.update (const _ a) f b) i + +lemma apply_def (l : Subspace η α ι) (x : η → α) (i : ι) : l x i = (l.idxFun i).elim id x := rfl +lemma apply_inl (h : l.idxFun i = Sum.inl a) : l x i = a := by simp [apply_def, h] +lemma apply_inr (h : l.idxFun i = Sum.inr e) : l x i = x e := by simp [apply_def, h] + +/-- Given a coloring `C` of `ι → α` and a combinatorial subspace `l` of `ι → α`, `l.IsMono C` +means that `l` is monochromatic with regard to `C`. -/ +def IsMono (C : (ι → α) → κ) (l : Subspace η α ι) : Prop := ∃ c, ∀ x, C (l x) = c + +variable {η' α' ι' : Type*} + +/-- Change the index types of a subspace. -/ +def reindex (l : Subspace η α ι) (eη : η ≃ η') (eα : α ≃ α') (eι : ι ≃ ι') : Subspace η' α' ι' where + idxFun i := (l.idxFun <| eι.symm i).map eα eη + proper e := (eι.exists_congr fun i ↦ by cases h : idxFun l i <;> + simp [*, Function.funext_iff, Equiv.eq_symm_apply]).1 <| l.proper <| eη.symm e + +@[simp] lemma reindex_apply (l : Subspace η α ι) (eη : η ≃ η') (eα : α ≃ α') (eι : ι ≃ ι') (x i) : + l.reindex eη eα eι x i = eα (l (eα.symm ∘ x ∘ eη) <| eι.symm i) := by + cases h : l.idxFun (eι.symm i) <;> simp [h, reindex, coe_apply] + +@[simp] lemma reindex_isMono {eη : η ≃ η'} {eα : α ≃ α'} {eι : ι ≃ ι'} {C : (ι' → α') → κ} : + (l.reindex eη eα eι).IsMono C ↔ l.IsMono fun x ↦ C <| eα ∘ x ∘ eι.symm := by + simp only [IsMono, funext (reindex_apply _ _ _ _ _), coe_apply] + exact exists_congr fun c ↦ (eη.arrowCongr eα).symm.forall_congr <| by aesop + +protected lemma IsMono.reindex {eη : η ≃ η'} {eα : α ≃ α'} {eι : ι ≃ ι'} {C : (ι → α) → κ} + (hl : l.IsMono C) : (l.reindex eη eα eι).IsMono fun x ↦ C <| eα.symm ∘ x ∘ eι := by + simp [reindex_isMono, Function.comp.assoc]; simpa [← Function.comp.assoc] + +end Subspace + /-- The type of combinatorial lines. A line `l : Line α ι` in the hypercube `ι → α` defines a function `α → ι → α` from `α` to the hypercube, such that for each coordinate `i : ι`, the function `fun x ↦ l x i` is either `id` or constant. We require lines to be nontrivial in the sense that @@ -74,6 +168,7 @@ Formally, a line is represented by a word `l.idxFun : ι → Option α` which sa `l.idxFun i = some y`). When `α` has size `1` there can be many elements of `Line α ι` defining the same function. -/ +@[ext] structure Line (α ι : Type*) where /-- The word representing a combinatorial line. `l.idxfun i = none` means that `l x i = x` for all `x` and `l.idxfun i = some y` means that `l x i = y`. -/ @@ -83,15 +178,60 @@ structure Line (α ι : Type*) where proper : ∃ i, idxFun i = none namespace Line +variable {l : Line α ι} {i : ι} {a x : α} + +/-- Consider a line `l : Line α ι` as a function `α → ι → α`. -/ +@[coe] def toFun (l : Line α ι) (x : α) (i : ι) : α := (l.idxFun i).getD x -- This lets us treat a line `l : Line α ι` as a function `α → ι → α`. -instance (α ι) : CoeFun (Line α ι) fun _ => α → ι → α := +instance instCoeFun : CoeFun (Line α ι) fun _ => α → ι → α := ⟨fun l x i => (l.idxFun i).getD x⟩ +lemma coe_apply (l : Line α ι) (x : α) (i : ι) : l x i = (l.idxFun i).getD x := rfl + +-- Note: This is not made a `FunLike` instance to avoid having two syntactically different coercions +lemma coe_injective [Nontrivial α] : Injective ((⇑) : Line α ι → α → ι → α) := by + rintro l m hlm + ext i a + obtain ⟨b, hba⟩ := exists_ne a + simp only [Option.mem_def, funext_iff] at hlm ⊢ + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · cases hi : idxFun m i <;> simpa [@eq_comm _ a, hi, h, hba] using hlm b i + · cases hi : idxFun l i <;> simpa [@eq_comm _ a, hi, h, hba] using hlm b i + /-- A line is monochromatic if all its points are the same color. -/ def IsMono {α ι κ} (C : (ι → α) → κ) (l : Line α ι) : Prop := ∃ c, ∀ x, C (l x) = c +/-- Consider a line as a one-dimensional subspace. -/ +def toSubspaceUnit (l : Line α ι) : Subspace Unit α ι where + idxFun i := (l.idxFun i).elim (.inr ()) .inl + proper _ := l.proper.imp fun i hi ↦ by simp [hi] + +@[simp] lemma toSubspaceUnit_apply (l : Line α ι) (a) : ⇑l.toSubspaceUnit a = l (a ()) := by + ext i; cases h : l.idxFun i <;> simp [toSubspaceUnit, h, Subspace.coe_apply] + +@[simp] lemma toSubspaceUnit_isMono {C : (ι → α) → κ} : l.toSubspaceUnit.IsMono C ↔ l.IsMono C := by + simp only [Subspace.IsMono, toSubspaceUnit_apply, IsMono] + exact exists_congr fun c ↦ ⟨fun h a ↦ h fun _ ↦ a, fun h a ↦ h _⟩ + +protected alias ⟨_, IsMono.toSubspaceUnit⟩ := toSubspaceUnit_isMono + +/-- Consider a line in `ι → η → α` as a `η`-dimensional subspace in `ι × η → α`. -/ +def toSubspace (l : Line (η → α) ι) : Subspace η α (ι × η) where + idxFun ie := (l.idxFun ie.1).elim (.inr ie.2) (fun f ↦ .inl <| f ie.2) + proper e := let ⟨i, hi⟩ := l.proper; ⟨(i, e), by simp [hi]⟩ + +@[simp] lemma toSubspace_apply (l : Line (η → α) ι) (a ie) : + ⇑l.toSubspace a ie = l a ie.1 ie.2 := by + cases h : l.idxFun ie.1 <;> simp [toSubspace, h, coe_apply, Subspace.coe_apply] + +@[simp] lemma toSubspace_isMono {l : Line (η → α) ι} {C : (ι × η → α) → κ} : + l.toSubspace.IsMono C ↔ l.IsMono fun x : ι → η → α ↦ C fun (i, e) ↦ x i e := by + simp [Subspace.IsMono, IsMono, funext (toSubspace_apply _ _)] + +protected alias ⟨_, IsMono.toSubspace⟩ := toSubspace_isMono + /-- The diagonal line. It is the identity at every coordinate. -/ def diagonal (α ι) [Nonempty ι] : Line α ι where idxFun _ := none @@ -136,7 +276,7 @@ instance {α ι κ} (C : (ι → Option α) → κ) : Inhabited (ColorFocused C) refine ⟨⟨0, fun _ => none, fun h => ?_, Multiset.nodup_zero⟩⟩ simp only [Multiset.not_mem_zero, IsEmpty.forall_iff] -/-- A function `f : α → α'` determines a function `line α ι → line α' ι`. For a coordinate `i`, +/-- A function `f : α → α'` determines a function `line α ι → line α' ι`. For a coordinate `i` `l.map f` is the identity at `i` if `l` is, and constantly `f y` if `l` is constantly `y` at `i`. -/ def map {α α' ι} (f : α → α') (l : Line α ι) : Line α' ι where idxFun i := (l.idxFun i).map f @@ -157,19 +297,16 @@ def prod {α ι ι'} (l : Line α ι) (l' : Line α ι') : Line α (ι ⊕ ι') idxFun := Sum.elim l.idxFun l'.idxFun proper := ⟨Sum.inl l.proper.choose, l.proper.choose_spec⟩ -theorem apply {α ι} (l : Line α ι) (x : α) : l x = fun i => (l.idxFun i).getD x := - rfl +theorem apply_def (l : Line α ι) (x : α) : l x = fun i => (l.idxFun i).getD x := rfl theorem apply_none {α ι} (l : Line α ι) (x : α) (i : ι) (h : l.idxFun i = none) : l x i = x := by - simp only [Option.getD_none, h, l.apply] + simp only [Option.getD_none, h, l.apply_def] -theorem apply_of_ne_none {α ι} (l : Line α ι) (x : α) (i : ι) (h : l.idxFun i ≠ none) : - some (l x i) = l.idxFun i := by rw [l.apply, Option.getD_of_ne_none h] +lemma apply_some (h : l.idxFun i = some a) : l x i = a := by simp [l.apply_def, h] @[simp] theorem map_apply {α α' ι} (f : α → α') (l : Line α ι) (x : α) : l.map f (f x) = f ∘ l x := by - simp only [Line.apply, Line.map, Option.getD_map] - rfl + simp only [Line.apply_def, Line.map, Option.getD_map, comp_def] @[simp] theorem vertical_apply {α ι ι'} (v : ι → α) (l : Line α ι') (x : α) : @@ -302,8 +439,9 @@ private theorem exists_mono_in_high_dimension' : -- Finally, we really do have `r+1` lines! · rw [Multiset.card_cons, Multiset.card_map, sr]) -/-- The Hales-Jewett theorem: for any finite types `α` and `κ`, there exists a finite type `ι` such -that whenever the hypercube `ι → α` is `κ`-colored, there is a monochromatic combinatorial line. -/ +/-- The **Hales-Jewett theorem**: For any finite types `α` and `κ`, there exists a finite type `ι` +such that whenever the hypercube `ι → α` is `κ`-colored, there is a monochromatic combinatorial +line. -/ theorem exists_mono_in_high_dimension (α : Type u) [Finite α] (κ : Type v) [Finite κ] : ∃ (ι : Type) (_ : Fintype ι), ∀ C : (ι → α) → κ, ∃ l : Line α ι, l.IsMono C := let ⟨ι, ιfin, hι⟩ := exists_mono_in_high_dimension'.{u,v} α (ULift.{u,v} κ) @@ -343,4 +481,32 @@ theorem exists_mono_homothetic_copy {M κ : Type*} [AddCommMonoid M] (S : Finset obtain ⟨y, hy⟩ := Option.ne_none_iff_exists.mp hi.right simp_rw [← hy, Option.map_some', Option.getD] +namespace Subspace + +/-- The **multidimensional Hales-Jewett theorem**, aka **extended Hales-Jewett theorem**: For any +finite types `η`, `α` and `κ`, there exists a finite type `ι` such that whenever the hypercube +`ι → α` is `κ`-colored, there is a monochromatic combinatorial subspace of dimension `η`. -/ +theorem exists_mono_in_high_dimension (α κ η) [Finite α] [Finite κ] [Finite η] : + ∃ (ι : Type) (_ : Fintype ι), ∀ C : (ι → α) → κ, ∃ l : Subspace η α ι, l.IsMono C := by + cases nonempty_fintype η + obtain ⟨ι, _, hι⟩ := Line.exists_mono_in_high_dimension (Shrink.{0} η → α) κ + refine ⟨ι × Shrink η, inferInstance, fun C ↦ ?_⟩ + obtain ⟨l, hl⟩ := hι fun x ↦ C fun (i, e) ↦ x i e + refine ⟨l.toSubspace.reindex (equivShrink.{0} η).symm (Equiv.refl _) (Equiv.refl _), ?_⟩ + convert hl.toSubspace.reindex + simp + +/-- A variant of the **extended Hales-Jewett theorem** `exists_mono_in_high_dimension` where the +returned type is some `Fin n` instead of a general fintype. -/ +theorem exists_mono_in_high_dimension_fin (α κ η) [Fintype α] [Fintype κ] [Fintype η] : + ∃ n, ∀ C : (Fin n → α) → κ, ∃ l : Subspace η α (Fin n), l.IsMono C := by + obtain ⟨ι, ιfin, hι⟩ := exists_mono_in_high_dimension α κ η + refine ⟨Fintype.card ι, fun C ↦ ?_⟩ + obtain ⟨l, c, cl⟩ := hι fun v ↦ C (v ∘ (Fintype.equivFin _).symm) + refine ⟨⟨l.idxFun ∘ (Fintype.equivFin _).symm, fun e ↦ ?_⟩, c, cl⟩ + obtain ⟨i, hi⟩ := l.proper e + use Fintype.equivFin _ i + simpa using hi + +end Subspace end Combinatorics diff --git a/Mathlib/Combinatorics/Quiver/Covering.lean b/Mathlib/Combinatorics/Quiver/Covering.lean index 75129c7d420e1..28e85b2f31ed9 100644 --- a/Mathlib/Combinatorics/Quiver/Covering.lean +++ b/Mathlib/Combinatorics/Quiver/Covering.lean @@ -250,7 +250,7 @@ end Prefunctor.IsCovering section HasInvolutiveReverse -variable [HasInvolutiveReverse U] [HasInvolutiveReverse V] [Prefunctor.MapReverse φ] +variable [HasInvolutiveReverse U] [HasInvolutiveReverse V] /-- In a quiver with involutive inverses, the star and costar at every vertex are equivalent. This map is induced by `Quiver.reverse`. -/ @@ -271,13 +271,15 @@ theorem Quiver.starEquivCostar_symm_apply {u v : U} (e : u ⟶ v) : (Quiver.starEquivCostar v).symm (Quiver.Costar.mk e) = Quiver.Star.mk (reverse e) := rfl +variable [Prefunctor.MapReverse φ] + theorem Prefunctor.costar_conj_star (u : U) : φ.costar u = Quiver.starEquivCostar (φ.obj u) ∘ φ.star u ∘ (Quiver.starEquivCostar u).symm := by ext ⟨v, f⟩ <;> simp theorem Prefunctor.bijective_costar_iff_bijective_star (u : U) : Bijective (φ.costar u) ↔ Bijective (φ.star u) := by - rw [Prefunctor.costar_conj_star, EquivLike.comp_bijective, EquivLike.bijective_comp] + rw [Prefunctor.costar_conj_star φ, EquivLike.comp_bijective, EquivLike.bijective_comp] theorem Prefunctor.isCovering_of_bijective_star (h : ∀ u, Bijective (φ.star u)) : φ.IsCovering := ⟨h, fun u => (φ.bijective_costar_iff_bijective_star u).2 (h u)⟩ diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 1ffbe6512bf1d..b4677cec67585 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -99,8 +99,7 @@ theorem card_div_choose_le_card_shadow_div_choose (hr : r ≠ 0) rw [tsub_add_eq_add_tsub hr', add_tsub_add_eq_tsub_right] at h𝒜 apply le_of_mul_le_mul_right _ (pos_iff_ne_zero.2 hr) convert Nat.mul_le_mul_right ((Fintype.card α).choose r) h𝒜 using 1 - · simp [mul_assoc, Nat.choose_succ_right_eq] - exact Or.inl (mul_comm _ _) + · simpa [mul_assoc, Nat.choose_succ_right_eq] using Or.inl (mul_comm _ _) · simp only [mul_assoc, choose_succ_right_eq, mul_eq_mul_left_iff] exact Or.inl (mul_comm _ _) · exact Nat.choose_pos hr' diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index e1551a84f9f72..d53cf1aa710f7 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -63,7 +63,7 @@ Use `aesop_graph?` to pass along a `Try this` suggestion when using `aesop_graph -/ macro (name := aesop_graph?) "aesop_graph?" c:Aesop.tactic_clause* : tactic => `(tactic| - aesop $c* + aesop? $c* (config := { introsTransparency? := some .default, terminal := true }) (rule_sets := [$(Lean.mkIdent `SimpleGraph):ident])) diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index 5f80b48fb9ffd..3cf3baf0178be 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -125,9 +125,8 @@ theorem isClique_map_iff {f : α ↪ β} {t : Set β} : obtain (hs | hs) := s.subsingleton_or_nontrivial · simp [hs, IsClique.of_subsingleton] simp [or_iff_right hs.not_subsingleton, Set.image_eq_image f.injective] -section DecidableEq -variable [DecidableEq β] {f : α ↪ β} {t : Finset β} +variable {f : α ↪ β} {t : Finset β} theorem isClique_map_finset_iff_of_nontrivial (ht : t.Nontrivial) : (G.map f).IsClique t ↔ ∃ (s : Finset α), G.IsClique s ∧ s.map f = t := by @@ -153,8 +152,6 @@ protected theorem IsClique.finsetMap {f : α ↪ β} {s : Finset α} (h : G.IsCl (G.map f).IsClique (s.map f) := by simpa -end DecidableEq - end Clique /-! ### `n`-cliques -/ @@ -191,7 +188,7 @@ protected theorem IsNClique.map (h : G.IsNClique n s) {f : α ↪ β} : (G.map f).IsNClique n (s.map f) := ⟨by rw [coe_map]; exact h.1.map, (card_map _).trans h.2⟩ -theorem isNClique_map_iff [DecidableEq β] (hn : 1 < n) {t : Finset β} {f : α ↪ β} : +theorem isNClique_map_iff (hn : 1 < n) {t : Finset β} {f : α ↪ β} : (G.map f).IsNClique n t ↔ ∃ s : Finset α, G.IsNClique n s ∧ s.map f = t := by rw [isNClique_iff, isClique_map_finset_iff, or_and_right, or_iff_right (by rintro ⟨h', rfl⟩; exact h'.not_lt hn)] @@ -326,7 +323,7 @@ theorem CliqueFree.comap {H : SimpleGraph β} (f : H ↪g G) : G.CliqueFree n intro h; contrapose h exact not_cliqueFree_of_top_embedding <| f.comp (topEmbeddingOfNotCliqueFree h) -@[simp] theorem cliqueFree_map_iff {f : α ↪ β} [DecidableEq β] [Nonempty α] : +@[simp] theorem cliqueFree_map_iff {f : α ↪ β} [Nonempty α] : (G.map f).CliqueFree n ↔ G.CliqueFree n := by obtain (hle | hlt) := le_or_lt n 1 · obtain (rfl | rfl) := Nat.le_one_iff_eq_zero_or_eq_one.1 hle diff --git a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean index c63edb3d36028..455f7edb1ac0e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean @@ -104,7 +104,7 @@ theorem Coloring.colorClasses_finite [Finite α] : C.colorClasses.Finite := theorem Coloring.card_colorClasses_le [Fintype α] [Fintype C.colorClasses] : Fintype.card C.colorClasses ≤ Fintype.card α := by - simp [colorClasses] + simp only [colorClasses] -- Porting note: brute force instance declaration `[Fintype (Setoid.classes (Setoid.ker C))]` haveI : Fintype (Setoid.classes (Setoid.ker C)) := by assumption convert Setoid.card_classes_ker_le C diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean index b474ff749d94b..7849182e88987 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean @@ -323,6 +323,15 @@ theorem reverse_reverse {u v : V} (p : G.Walk u v) : p.reverse.reverse = p := by | nil => rfl | cons _ _ ih => simp [ih] +theorem reverse_surjective {u v : V} : Function.Surjective (reverse : G.Walk u v → _) := + RightInverse.surjective reverse_reverse + +theorem reverse_injective {u v : V} : Function.Injective (reverse : G.Walk u v → _) := + RightInverse.injective reverse_reverse + +theorem reverse_bijective {u v : V} : Function.Bijective (reverse : G.Walk u v → _) := + And.intro reverse_injective reverse_surjective + @[simp] theorem length_nil {u : V} : (nil : G.Walk u u).length = 0 := rfl @@ -1076,7 +1085,7 @@ theorem count_edges_takeUntil_le_one {u v w : V} (p : G.Walk v w) (h : u ∈ p.s simp · rw [edges_cons, List.count_cons] split_ifs with h'' - · rw [Sym2.eq_iff] at h'' + · simp only [beq_iff_eq, Sym2.eq, Sym2.rel_iff'] at h'' obtain ⟨rfl, rfl⟩ | ⟨rfl, rfl⟩ := h'' · exact (h' rfl).elim · cases p' <;> simp! @@ -1118,11 +1127,11 @@ theorem darts_dropUntil_subset {u v w : V} (p : G.Walk v w) (h : u ∈ p.support theorem edges_takeUntil_subset {u v w : V} (p : G.Walk v w) (h : u ∈ p.support) : (p.takeUntil u h).edges ⊆ p.edges := - List.map_subset _ (p.darts_takeUntil_subset h) + List.Subset.map _ (p.darts_takeUntil_subset h) theorem edges_dropUntil_subset {u v w : V} (p : G.Walk v w) (h : u ∈ p.support) : (p.dropUntil u h).edges ⊆ p.edges := - List.map_subset _ (p.darts_dropUntil_subset h) + List.Subset.map _ (p.darts_dropUntil_subset h) theorem length_takeUntil_le {u v w : V} (p : G.Walk v w) (h : u ∈ p.support) : (p.takeUntil u h).length ≤ p.length := by @@ -1138,20 +1147,20 @@ theorem length_dropUntil_le {u v w : V} (p : G.Walk v w) (h : u ∈ p.support) : protected theorem IsTrail.takeUntil {u v w : V} {p : G.Walk v w} (hc : p.IsTrail) (h : u ∈ p.support) : (p.takeUntil u h).IsTrail := - IsTrail.of_append_left (by rwa [← take_spec _ h] at hc) + IsTrail.of_append_left (q := p.dropUntil u h) (by rwa [← take_spec _ h] at hc) protected theorem IsTrail.dropUntil {u v w : V} {p : G.Walk v w} (hc : p.IsTrail) (h : u ∈ p.support) : (p.dropUntil u h).IsTrail := - IsTrail.of_append_right (by rwa [← take_spec _ h] at hc) + IsTrail.of_append_right (p := p.takeUntil u h) (by rwa [← take_spec _ h] at hc) protected theorem IsPath.takeUntil {u v w : V} {p : G.Walk v w} (hc : p.IsPath) (h : u ∈ p.support) : (p.takeUntil u h).IsPath := - IsPath.of_append_left (by rwa [← take_spec _ h] at hc) + IsPath.of_append_left (q := p.dropUntil u h) (by rwa [← take_spec _ h] at hc) -- Porting note: p was previously accidentally an explicit argument protected theorem IsPath.dropUntil {u v w : V} {p : G.Walk v w} (hc : p.IsPath) (h : u ∈ p.support) : (p.dropUntil u h).IsPath := - IsPath.of_append_right (by rwa [← take_spec _ h] at hc) + IsPath.of_append_right (p := p.takeUntil u h) (by rwa [← take_spec _ h] at hc) /-- Rotate a loop walk such that it is centered at the given vertex. -/ def rotate {u v : V} (c : G.Walk v v) (h : u ∈ c.support) : G.Walk u u := @@ -1369,7 +1378,7 @@ theorem darts_bypass_subset {u v : V} (p : G.Walk u v) : p.bypass.darts ⊆ p.da exact List.cons_subset_cons _ ih theorem edges_bypass_subset {u v : V} (p : G.Walk u v) : p.bypass.edges ⊆ p.edges := - List.map_subset _ p.darts_bypass_subset + List.Subset.map _ p.darts_bypass_subset theorem darts_toPath_subset {u v : V} (p : G.Walk u v) : (p.toPath : G.Walk u v).darts ⊆ p.darts := darts_bypass_subset _ @@ -2064,6 +2073,22 @@ lemma mem_coe_supp_of_adj {v w : V} {H : Subgraph G} {c : ConnectedComponent H.c end ConnectedComponent +-- TODO: Extract as lemma about general equivalence relation +lemma pairwise_disjoint_supp_connectedComponent (G : SimpleGraph V) : + Pairwise fun c c' : ConnectedComponent G ↦ Disjoint c.supp c'.supp := by + simp_rw [Set.disjoint_left] + intro _ _ h a hsx hsy + rw [ConnectedComponent.mem_supp_iff] at hsx hsy + rw [hsx] at hsy + exact h hsy + +-- TODO: Extract as lemma about general equivalence relation +lemma iUnion_connectedComponentSupp (G : SimpleGraph V) : + ⋃ c : G.ConnectedComponent, c.supp = Set.univ := by + refine Set.eq_univ_of_forall fun v ↦ ⟨G.connectedComponentMk v, ?_⟩ + simp only [Set.mem_range, SetLike.mem_coe] + exact ⟨by use G.connectedComponentMk v; exact rfl, rfl⟩ + theorem Preconnected.set_univ_walk_nonempty (hconn : G.Preconnected) (u v : V) : (Set.univ : Set (G.Walk u v)).Nonempty := by rw [← Set.nonempty_iff_univ_nonempty] diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index 80e60b12a4d2b..ccdb9857a4362 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -3,8 +3,10 @@ Copyright (c) 2021 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Algebra.BigOperators.Ring.Nat import Mathlib.Combinatorics.SimpleGraph.Connectivity import Mathlib.Combinatorics.SimpleGraph.Subgraph +import Mathlib.SetTheory.Cardinal.Finite /-! # Counting walks of a given length @@ -169,6 +171,19 @@ instance : Decidable G.Connected := by instance instDecidableMemSupp (c : G.ConnectedComponent) (v : V) : Decidable (v ∈ c.supp) := c.recOn (fun w ↦ decidable_of_iff (G.Reachable v w) $ by simp) (fun _ _ _ _ ↦ Subsingleton.elim _ _) + +lemma odd_card_iff_odd_components : Odd (Nat.card V) ↔ + Odd (Nat.card ({(c : ConnectedComponent G) | Odd (Nat.card c.supp)})) := by + rw [Nat.card_eq_fintype_card] + simp only [← (set_fintype_card_eq_univ_iff _).mpr G.iUnion_connectedComponentSupp, + ConnectedComponent.mem_supp_iff, Fintype.card_subtype_compl, + ← Set.toFinset_card, Set.toFinset_iUnion ConnectedComponent.supp] + rw [Finset.card_biUnion + (fun x _ y _ hxy ↦ Set.disjoint_toFinset.mpr (pairwise_disjoint_supp_connectedComponent _ hxy))] + simp_rw [Set.toFinset_card, ← Nat.card_eq_fintype_card] + rw [Nat.card_eq_fintype_card, Fintype.card_ofFinset] + exact (Finset.odd_sum_iff_odd_card_odd (fun x : G.ConnectedComponent ↦ Nat.card x.supp)) + end Finite end WalkCounting diff --git a/Mathlib/Combinatorics/SimpleGraph/Density.lean b/Mathlib/Combinatorics/SimpleGraph/Density.lean index ea7e5fa5d7430..d58cb3ac6b46b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Density.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Density.lean @@ -237,14 +237,16 @@ end Asymmetric section Symmetric -variable (r : α → α → Prop) [DecidableRel r] {s s₁ s₂ t t₁ t₂ : Finset α} {a b : α} -variable {r} (hr : Symmetric r) +variable {r : α → α → Prop} [DecidableRel r] {s s₁ s₂ t t₁ t₂ : Finset α} {a b : α} @[simp] -theorem swap_mem_interedges_iff {x : α × α} : x.swap ∈ interedges r s t ↔ x ∈ interedges r t s := by +theorem swap_mem_interedges_iff (hr : Symmetric r) {x : α × α} : + x.swap ∈ interedges r s t ↔ x ∈ interedges r t s := by rw [mem_interedges_iff, mem_interedges_iff, hr.iff] exact and_left_comm +variable (hr : Symmetric r) + theorem mk_mem_interedges_comm : (a, b) ∈ interedges r s t ↔ (b, a) ∈ interedges r t s := @swap_mem_interedges_iff _ _ _ _ _ hr (b, a) @@ -253,7 +255,8 @@ theorem card_interedges_comm (s t : Finset α) : (interedges r s t).card = (inte (fun _ _ _ _ h ↦ Prod.swap_injective h) fun x h ↦ ⟨x.swap, (swap_mem_interedges_iff hr).2 h, x.swap_swap⟩ -theorem edgeDensity_comm (s t : Finset α) : edgeDensity r s t = edgeDensity r t s := by +theorem edgeDensity_comm (hr : Symmetric r) (s t : Finset α) : + edgeDensity r s t = edgeDensity r t s := by rw [edgeDensity, mul_comm, card_interedges_comm hr, edgeDensity] end Symmetric diff --git a/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean index 919a2d3bfb1ae..b2aac193efab8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Finsubgraph.lean @@ -61,6 +61,9 @@ instance : Sup G.Finsubgraph := instance : Inf G.Finsubgraph := ⟨fun G₁ G₂ => ⟨G₁ ⊓ G₂, G₁.2.subset inter_subset_left⟩⟩ +instance instSDiff : SDiff G.Finsubgraph where + sdiff G₁ G₂ := ⟨G₁ \ G₂, G₁.2.subset (Subgraph.verts_mono sdiff_le)⟩ + @[simp, norm_cast] lemma coe_bot : (⊥ : G.Finsubgraph) = (⊥ : G.Subgraph) := rfl @[simp, norm_cast] @@ -69,17 +72,28 @@ lemma coe_sup (G₁ G₂ : G.Finsubgraph) : ↑(G₁ ⊔ G₂) = (G₁ ⊔ G₂ @[simp, norm_cast] lemma coe_inf (G₁ G₂ : G.Finsubgraph) : ↑(G₁ ⊓ G₂) = (G₁ ⊓ G₂ : G.Subgraph) := rfl -instance : DistribLattice G.Finsubgraph := - Subtype.coe_injective.distribLattice _ (fun _ _ => rfl) fun _ _ => rfl +@[simp, norm_cast] +lemma coe_sdiff (G₁ G₂ : G.Finsubgraph) : ↑(G₁ \ G₂) = (G₁ \ G₂ : G.Subgraph) := rfl + +instance instGeneralizedCoheytingAlgebra : GeneralizedCoheytingAlgebra G.Finsubgraph := + Subtype.coe_injective.generalizedCoheytingAlgebra _ coe_sup coe_inf coe_bot coe_sdiff section Finite variable [Finite V] instance instTop : Top G.Finsubgraph where top := ⟨⊤, finite_univ⟩ +instance instHasCompl : HasCompl G.Finsubgraph where compl G' := ⟨G'ᶜ, Set.toFinite _⟩ +instance instHNot : HNot G.Finsubgraph where hnot G' := ⟨¬G', Set.toFinite _⟩ +instance instHImp : HImp G.Finsubgraph where himp G₁ G₂ := ⟨G₁ ⇨ G₂, Set.toFinite _⟩ instance instSupSet : SupSet G.Finsubgraph where sSup s := ⟨⨆ G ∈ s, ↑G, Set.toFinite _⟩ instance instInfSet : InfSet G.Finsubgraph where sInf s := ⟨⨅ G ∈ s, ↑G, Set.toFinite _⟩ -@[simp, norm_cast] lemma coe_top : ↑(⊤ : G.Finsubgraph) = (⊤ : G.Subgraph) := rfl +@[simp, norm_cast] lemma coe_top : (⊤ : G.Finsubgraph) = (⊤ : G.Subgraph) := rfl +@[simp, norm_cast] lemma coe_compl (G' : G.Finsubgraph) : ↑(G'ᶜ) = (G'ᶜ : G.Subgraph) := rfl +@[simp, norm_cast] lemma coe_hnot (G' : G.Finsubgraph) : ↑(¬G') = (¬G' : G.Subgraph) := rfl + +@[simp, norm_cast] +lemma coe_himp (G₁ G₂ : G.Finsubgraph) : ↑(G₁ ⇨ G₂) = (G₁ ⇨ G₂ : G.Subgraph) := rfl @[simp, norm_cast] lemma coe_sSup (s : Set G.Finsubgraph) : sSup s = (⨆ G ∈ s, G : G.Subgraph) := rfl @@ -97,6 +111,7 @@ lemma coe_iInf {ι : Sort*} (f : ι → G.Finsubgraph) : ⨅ i, f i = (⨅ i, f instance instCompletelyDistribLattice : CompletelyDistribLattice G.Finsubgraph := Subtype.coe_injective.completelyDistribLattice _ coe_sup coe_inf coe_sSup coe_sInf coe_top coe_bot + coe_compl coe_himp coe_hnot coe_sdiff end Finite end Finsubgraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index f43e517aba690..aa00ce2aee1b6 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -79,8 +79,8 @@ lemma IsHamiltonianCycle.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective intro x rcases p with (_ | ⟨y, p⟩) · cases hp.ne_nil rfl - simp only [support_cons, List.count_cons, List.map_cons, List.head_cons, hf.injective.eq_iff, - add_tsub_cancel_right] + simp only [support_cons, List.count_cons, List.map_cons, List.head_cons, beq_iff_eq, + hf.injective.eq_iff, add_tsub_cancel_right] exact hp.isHamiltonian_tail _ lemma isHamiltonianCycle_isCycle_and_isHamiltonian_tail : diff --git a/Mathlib/Combinatorics/SimpleGraph/Metric.lean b/Mathlib/Combinatorics/SimpleGraph/Metric.lean index f0780089fb17e..51dfeb1a6cbc0 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Metric.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Metric.lean @@ -1,19 +1,21 @@ /- Copyright (c) 2022 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kyle Miller, Vincent Beffara +Authors: Kyle Miller, Vincent Beffara, Rida Hamadani -/ import Mathlib.Combinatorics.SimpleGraph.Connectivity -import Mathlib.Data.Nat.Lattice +import Mathlib.Data.ENat.Lattice /-! # Graph metric -This module defines the `SimpleGraph.dist` function, which takes -pairs of vertices to the length of the shortest walk between them. +This module defines the `SimpleGraph.edist` function, which takes pairs of vertices to the length of +the shortest walk between them, or `⊤` if they are disconnected. It also defines `SimpleGraph.dist` +which is the `ℕ`-valued version of `SimpleGraph.edist`. ## Main definitions +- `SimpleGraph.edist` is the graph extended metric. - `SimpleGraph.dist` is the graph metric. ## TODO @@ -21,10 +23,6 @@ pairs of vertices to the length of the shortest walk between them. - Provide an additional computable version of `SimpleGraph.dist` for when `G` is connected. -- Evaluate `Nat` vs `ENat` for the codomain of `dist`, or potentially - having an additional `edist` when the objects under consideration are - disconnected graphs. - - When directed graphs exist, a directed notion of distance, likely `ENat`-valued. @@ -41,91 +39,181 @@ variable {V : Type*} (G : SimpleGraph V) /-! ## Metric -/ +section edist + +/-- +The extended distance between two vertices is the length of the shortest walk between them. +It is `⊤` if no such walk exists. +-/ +noncomputable def edist (u v : V) : ℕ∞ := + ⨅ w : G.Walk u v, w.length + +variable {G} {u v w : V} + +theorem edist_eq_sInf : G.edist u v = sInf (Set.range fun w : G.Walk u v ↦ (w.length : ℕ∞)) := rfl + +protected theorem Reachable.exists_walk_length_eq_edist (hr : G.Reachable u v) : + ∃ p : G.Walk u v, p.length = G.edist u v := + csInf_mem <| Set.range_nonempty_iff_nonempty.mpr hr -/-- The distance between two vertices is the length of the shortest walk between them. -If no such walk exists, this uses the junk value of `0`. -/ +protected theorem Connected.exists_walk_length_eq_edist (hconn : G.Connected) (u v : V) : + ∃ p : G.Walk u v, p.length = G.edist u v := + (hconn u v).exists_walk_length_eq_edist + +theorem edist_le (p : G.Walk u v) : + G.edist u v ≤ p.length := + sInf_le ⟨p, rfl⟩ +protected alias Walk.edist_le := edist_le + +@[simp] +theorem edist_eq_zero_iff : + G.edist u v = 0 ↔ u = v := by + apply Iff.intro <;> simp [edist, ENat.iInf_eq_zero] + +theorem edist_self : edist G v v = 0 := + edist_eq_zero_iff.mpr rfl + +theorem edist_pos_of_ne (hne : u ≠ v) : + 0 < G.edist u v := + pos_iff_ne_zero.mpr <| edist_eq_zero_iff.ne.mpr hne + +lemma edist_eq_top_of_not_reachable (h : ¬G.Reachable u v) : + G.edist u v = ⊤ := by + simp [edist, not_reachable_iff_isEmpty_walk.mp h] + +theorem reachable_of_edist_ne_top (h : G.edist u v ≠ ⊤) : + G.Reachable u v := + not_not.mp <| edist_eq_top_of_not_reachable.mt h + +lemma exists_walk_of_edist_ne_top (h : G.edist u v ≠ ⊤) : + ∃ p : G.Walk u v, p.length = G.edist u v := + (reachable_of_edist_ne_top h).exists_walk_length_eq_edist + +protected theorem edist_triangle : G.edist u w ≤ G.edist u v + G.edist v w := by + rcases eq_or_ne (G.edist u v) ⊤ with huv | huv + case inl => simp [huv] + case inr => + rcases eq_or_ne (G.edist v w) ⊤ with hvw | hvw + case inl => simp [hvw] + case inr => + obtain ⟨p, hp⟩ := exists_walk_of_edist_ne_top huv + obtain ⟨q, hq⟩ := exists_walk_of_edist_ne_top hvw + rw [← hp, ← hq, ← Nat.cast_add, ← Walk.length_append] + exact edist_le _ + +theorem edist_comm : G.edist u v = G.edist v u := by + rw [edist_eq_sInf, ← Set.image_univ, ← Set.image_univ_of_surjective Walk.reverse_surjective, + ← Set.image_comp, Set.image_univ, Function.comp_def] + simp_rw [Walk.length_reverse, ← edist_eq_sInf] + +lemma exists_walk_of_edist_eq_coe {k : ℕ} (h : G.edist u v = k) : + ∃ p : G.Walk u v, p.length = k := + have : G.edist u v ≠ ⊤ := by rw [h]; exact ENat.coe_ne_top _ + have ⟨p, hp⟩ := exists_walk_of_edist_ne_top this + ⟨p, Nat.cast_injective (hp.trans h)⟩ + +lemma edist_ne_top_iff_reachable : G.edist u v ≠ ⊤ ↔ G.Reachable u v := by + refine ⟨reachable_of_edist_ne_top, fun h ↦ ?_⟩ + by_contra hx + simp only [edist, iInf_eq_top, ENat.coe_ne_top] at hx + exact h.elim hx + +/-- +The extended distance between vertices is equal to `1` if and only if these vertices are adjacent. +-/ +@[simp] +theorem edist_eq_one_iff_adj : G.edist u v = 1 ↔ G.Adj u v := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · obtain ⟨w, hw⟩ := exists_walk_of_edist_ne_top <| by rw [h]; simp + exact w.adj_of_length_eq_one <| Nat.cast_eq_one.mp <| h ▸ hw + · exact le_antisymm (edist_le h.toWalk) (ENat.one_le_iff_pos.mpr <| edist_pos_of_ne h.ne) + +end edist + +section dist + +/-- +The distance between two vertices is the length of the shortest walk between them. +If no such walk exists, this uses the junk value of `0`. +-/ noncomputable def dist (u v : V) : ℕ := - sInf (Set.range (Walk.length : G.Walk u v → ℕ)) + (G.edist u v).toNat + +variable {G} {u v w : V} -variable {G} +theorem dist_eq_sInf : G.dist u v = sInf (Set.range (Walk.length : G.Walk u v → ℕ)) := + ENat.iInf_toNat -protected theorem Reachable.exists_walk_of_dist {u v : V} (hr : G.Reachable u v) : +protected theorem Reachable.exists_walk_length_eq_dist (hr : G.Reachable u v) : ∃ p : G.Walk u v, p.length = G.dist u v := - Nat.sInf_mem (Set.range_nonempty_iff_nonempty.mpr hr) + dist_eq_sInf ▸ Nat.sInf_mem (Set.range_nonempty_iff_nonempty.mpr hr) -protected theorem Connected.exists_walk_of_dist (hconn : G.Connected) (u v : V) : +protected theorem Connected.exists_walk_length_eq_dist (hconn : G.Connected) (u v : V) : ∃ p : G.Walk u v, p.length = G.dist u v := - (hconn u v).exists_walk_of_dist + dist_eq_sInf ▸ (hconn u v).exists_walk_length_eq_dist -theorem dist_le {u v : V} (p : G.Walk u v) : G.dist u v ≤ p.length := - Nat.sInf_le ⟨p, rfl⟩ +theorem dist_le (p : G.Walk u v) : G.dist u v ≤ p.length := + dist_eq_sInf ▸ Nat.sInf_le ⟨p, rfl⟩ @[simp] -theorem dist_eq_zero_iff_eq_or_not_reachable {u v : V} : - G.dist u v = 0 ↔ u = v ∨ ¬G.Reachable u v := by simp [dist, Nat.sInf_eq_zero, Reachable] +theorem dist_eq_zero_iff_eq_or_not_reachable : + G.dist u v = 0 ↔ u = v ∨ ¬G.Reachable u v := by simp [dist_eq_sInf, Nat.sInf_eq_zero, Reachable] -theorem dist_self {v : V} : dist G v v = 0 := by simp +theorem dist_self : dist G v v = 0 := by simp -protected theorem Reachable.dist_eq_zero_iff {u v : V} (hr : G.Reachable u v) : +protected theorem Reachable.dist_eq_zero_iff (hr : G.Reachable u v) : G.dist u v = 0 ↔ u = v := by simp [hr] -protected theorem Reachable.pos_dist_of_ne {u v : V} (h : G.Reachable u v) (hne : u ≠ v) : +protected theorem Reachable.pos_dist_of_ne (h : G.Reachable u v) (hne : u ≠ v) : 0 < G.dist u v := Nat.pos_of_ne_zero (by simp [h, hne]) -protected theorem Connected.dist_eq_zero_iff (hconn : G.Connected) {u v : V} : +protected theorem Connected.dist_eq_zero_iff (hconn : G.Connected) : G.dist u v = 0 ↔ u = v := by simp [hconn u v] -protected theorem Connected.pos_dist_of_ne {u v : V} (hconn : G.Connected) (hne : u ≠ v) : +protected theorem Connected.pos_dist_of_ne (hconn : G.Connected) (hne : u ≠ v) : 0 < G.dist u v := - Nat.pos_of_ne_zero (by intro h; exact False.elim (hne (hconn.dist_eq_zero_iff.mp h))) + Nat.pos_of_ne_zero fun h ↦ False.elim <| hne <| (hconn.dist_eq_zero_iff).mp h -theorem dist_eq_zero_of_not_reachable {u v : V} (h : ¬G.Reachable u v) : G.dist u v = 0 := by +theorem dist_eq_zero_of_not_reachable (h : ¬G.Reachable u v) : G.dist u v = 0 := by simp [h] -theorem nonempty_of_pos_dist {u v : V} (h : 0 < G.dist u v) : +theorem nonempty_of_pos_dist (h : 0 < G.dist u v) : (Set.univ : Set (G.Walk u v)).Nonempty := by + rw [dist_eq_sInf] at h simpa [Set.range_nonempty_iff_nonempty, Set.nonempty_iff_univ_nonempty] using Nat.nonempty_of_pos_sInf h -protected theorem Connected.dist_triangle (hconn : G.Connected) {u v w : V} : +protected theorem Connected.dist_triangle (hconn : G.Connected) : G.dist u w ≤ G.dist u v + G.dist v w := by - obtain ⟨p, hp⟩ := hconn.exists_walk_of_dist u v - obtain ⟨q, hq⟩ := hconn.exists_walk_of_dist v w + obtain ⟨p, hp⟩ := hconn.exists_walk_length_eq_dist u v + obtain ⟨q, hq⟩ := hconn.exists_walk_length_eq_dist v w rw [← hp, ← hq, ← Walk.length_append] apply dist_le -private theorem dist_comm_aux {u v : V} (h : G.Reachable u v) : G.dist u v ≤ G.dist v u := by - obtain ⟨p, hp⟩ := h.symm.exists_walk_of_dist - rw [← hp, ← Walk.length_reverse] - apply dist_le +theorem dist_comm : G.dist u v = G.dist v u := by + rw [dist, dist, edist_comm] -theorem dist_comm {u v : V} : G.dist u v = G.dist v u := by - by_cases h : G.Reachable u v - · apply le_antisymm (dist_comm_aux h) (dist_comm_aux h.symm) - · have h' : ¬G.Reachable v u := fun h' => absurd h'.symm h - simp [h, h', dist_eq_zero_of_not_reachable] - -lemma dist_ne_zero_iff_ne_and_reachable {u v : V} : G.dist u v ≠ 0 ↔ u ≠ v ∧ G.Reachable u v := by +lemma dist_ne_zero_iff_ne_and_reachable : G.dist u v ≠ 0 ↔ u ≠ v ∧ G.Reachable u v := by rw [ne_eq, dist_eq_zero_iff_eq_or_not_reachable.not] push_neg; rfl -lemma Reachable.of_dist_ne_zero {u v : V} (h : G.dist u v ≠ 0) : G.Reachable u v := +lemma Reachable.of_dist_ne_zero (h : G.dist u v ≠ 0) : G.Reachable u v := (dist_ne_zero_iff_ne_and_reachable.mp h).2 -lemma exists_walk_of_dist_ne_zero {u v : V} (h : G.dist u v ≠ 0) : +lemma exists_walk_of_dist_ne_zero (h : G.dist u v ≠ 0) : ∃ p : G.Walk u v, p.length = G.dist u v := - (Reachable.of_dist_ne_zero h).exists_walk_of_dist + (Reachable.of_dist_ne_zero h).exists_walk_length_eq_dist -/- The distance between vertices is equal to `1` if and only if these vertices are adjacent. -/ -theorem dist_eq_one_iff_adj {u v : V} : G.dist u v = 1 ↔ G.Adj u v := by - refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ - · let ⟨w, hw⟩ := exists_walk_of_dist_ne_zero <| ne_zero_of_eq_one h - exact w.adj_of_length_eq_one <| h ▸ hw - · have : h.toWalk.length = 1 := Walk.length_cons _ _ - exact ge_antisymm (h.reachable.pos_dist_of_ne h.ne) (this ▸ dist_le _) +/-- +The distance between vertices is equal to `1` if and only if these vertices are adjacent. +-/ +@[simp] +theorem dist_eq_one_iff_adj : G.dist u v = 1 ↔ G.Adj u v := by + rw [dist, ENat.toNat_eq_iff, ENat.coe_one, edist_eq_one_iff_adj] + decide -theorem Walk.isPath_of_length_eq_dist {u v : V} (p : G.Walk u v) (hp : p.length = G.dist u v) : +theorem Walk.isPath_of_length_eq_dist (p : G.Walk u v) (hp : p.length = G.dist u v) : p.IsPath := by classical have : p.bypass = p := by @@ -136,14 +224,16 @@ theorem Walk.isPath_of_length_eq_dist {u v : V} (p : G.Walk u v) (hp : p.length rw [← this] apply Walk.bypass_isPath -lemma Reachable.exists_path_of_dist {u v : V} (hr : G.Reachable u v) : +lemma Reachable.exists_path_of_dist (hr : G.Reachable u v) : ∃ (p : G.Walk u v), p.IsPath ∧ p.length = G.dist u v := by - obtain ⟨p, h⟩ := hr.exists_walk_of_dist + obtain ⟨p, h⟩ := hr.exists_walk_length_eq_dist exact ⟨p, p.isPath_of_length_eq_dist h, h⟩ lemma Connected.exists_path_of_dist (hconn : G.Connected) (u v : V) : ∃ (p : G.Walk u v), p.IsPath ∧ p.length = G.dist u v := by - obtain ⟨p, h⟩ := hconn.exists_walk_of_dist u v + obtain ⟨p, h⟩ := hconn.exists_walk_length_eq_dist u v exact ⟨p, p.isPath_of_length_eq_dist h, h⟩ +end dist + end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index 5e723dbc905f4..68777353d2638 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -397,8 +397,8 @@ instance : BoundedOrder (Subgraph G) where le_top x := ⟨Set.subset_univ _, fun _ _ => x.adj_sub⟩ bot_le _ := ⟨Set.empty_subset _, fun _ _ => False.elim⟩ --- Note that subgraphs do not form a Boolean algebra, because of `verts`. -instance : CompletelyDistribLattice G.Subgraph := +/-- Note that subgraphs do not form a Boolean algebra, because of `verts`. -/ +def completelyDistribLatticeMinimalAxioms : CompletelyDistribLattice.MinimalAxioms G.Subgraph := { Subgraph.distribLattice with le := (· ≤ ·) sup := (· ⊔ ·) @@ -422,6 +422,9 @@ instance : CompletelyDistribLattice G.Subgraph := iInf_iSup_eq := fun f => Subgraph.ext (by simpa using iInf_iSup_eq) (by ext; simp [Classical.skolem]) } +instance : CompletelyDistribLattice G.Subgraph := + .ofMinimalAxioms completelyDistribLatticeMinimalAxioms + @[gcongr] lemma verts_mono {H H' : G.Subgraph} (h : H ≤ H') : H.verts ⊆ H'.verts := h.1 lemma verts_monotone : Monotone (verts : G.Subgraph → Set V) := fun _ _ h ↦ h.1 @@ -844,7 +847,7 @@ theorem neighborSet_subgraphOfAdj_of_ne_of_ne {u v w : V} (hvw : G.Adj v w) (hv theorem neighborSet_subgraphOfAdj [DecidableEq V] {u v w : V} (hvw : G.Adj v w) : (G.subgraphOfAdj hvw).neighborSet u = (if u = v then {w} else ∅) ∪ if u = w then {v} else ∅ := by - split_ifs <;> subst_vars <;> simp [*, Set.singleton_def] + split_ifs <;> subst_vars <;> simp [*] theorem singletonSubgraph_fst_le_subgraphOfAdj {u v : V} {h : G.Adj u v} : G.singletonSubgraph u ≤ G.subgraphOfAdj h := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Trails.lean b/Mathlib/Combinatorics/SimpleGraph/Trails.lean index 9104132477e42..32df7b9bd5787 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Trails.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Trails.lean @@ -90,7 +90,7 @@ theorem IsEulerian.isTrail {u v : V} {p : G.Walk u v} (h : p.IsEulerian) : p.IsT intro e by_cases he : e ∈ p.edges · exact (h e (edges_subset_edgeSet _ he)).le - · simp [he] + · simp [List.count_eq_zero_of_not_mem he] theorem IsEulerian.mem_edges_iff {u v : V} {p : G.Walk u v} (h : p.IsEulerian) {e : Sym2 V} : e ∈ p.edges ↔ e ∈ G.edgeSet := diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index b7574c58cbad5..9eb128cce610b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -83,11 +83,10 @@ lemma LocallyLinear.map (f : α ↪ β) (hG : G.LocallyLinear) : (G.map f).Local · rw [← Equiv.coe_toEmbedding, ← map_symm] exact LocallyLinear.map _ -variable [DecidableEq α] - lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : G.EdgeDisjointTriangles ↔ ∀ ⦃e : Sym2 α⦄, ¬ e.IsDiag → {s ∈ G.cliqueSet 3 | e ∈ (s : Finset α).sym2}.Subsingleton := by + classical have (a b) (hab : a ≠ b) : {s ∈ (G.cliqueSet 3 : Set (Finset α)) | s(a, b) ∈ (s : Finset α).sym2} = {s | G.Adj a b ∧ ∃ c, G.Adj a c ∧ G.Adj b c ∧ s = {a, b, c}} := by ext s @@ -131,7 +130,7 @@ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : alias ⟨EdgeDisjointTriangles.mem_sym2_subsingleton, _⟩ := edgeDisjointTriangles_iff_mem_sym2_subsingleton -variable [Fintype α] [DecidableRel G.Adj] +variable [DecidableEq α] [Fintype α] [DecidableRel G.Adj] instance EdgeDisjointTriangles.instDecidable : Decidable G.EdgeDisjointTriangles := decidable_of_iff ((G.cliqueFinset 3 : Set (Finset α)).Pairwise fun x y ↦ ((x ∩ y).card ≤ 1)) $ by @@ -159,7 +158,8 @@ lemma LocallyLinear.card_edgeFinset (hG : G.LocallyLinear) : refine hG.edgeDisjointTriangles.card_edgeFinset_le.antisymm' ?_ rw [← mul_comm, ← mul_one (Finset.card _)] refine card_mul_le_card_mul (fun e s ↦ e ∈ s.sym2) ?_ ?_ - · simpa [Sym2.forall, Nat.one_le_iff_ne_zero, -card_eq_zero, card_ne_zero, Finset.Nonempty] + · simpa [Sym2.forall, Nat.one_le_iff_ne_zero, -Finset.card_eq_zero, Finset.card_ne_zero, + Finset.Nonempty] using hG.2 simp only [mem_cliqueFinset_iff, is3Clique_iff, forall_exists_index, and_imp] rintro _ a b c hab hac hbc rfl @@ -175,7 +175,7 @@ lemma LocallyLinear.card_edgeFinset (hG : G.LocallyLinear) : end LocallyLinear variable (G ε) -variable [Fintype α] [DecidableEq α] [DecidableRel G.Adj] [DecidableRel H.Adj] +variable [Fintype α] [DecidableRel G.Adj] [DecidableRel H.Adj] /-- A simple graph is *`ε`-far from triangle-free* if one must remove at least `ε * (card α) ^ 2` edges to make it triangle-free. -/ @@ -192,6 +192,10 @@ alias ⟨farFromTriangleFree.le_card_sub_card, _⟩ := farFromTriangleFree_iff nonrec theorem FarFromTriangleFree.mono (hε : G.FarFromTriangleFree ε) (h : δ ≤ ε) : G.FarFromTriangleFree δ := hε.mono <| by gcongr +section DecidableEq + +variable [DecidableEq α] + theorem FarFromTriangleFree.cliqueFinset_nonempty' (hH : H ≤ G) (hG : G.FarFromTriangleFree ε) (hcard : G.edgeFinset.card - H.edgeFinset.card < ε * (card α ^ 2 : ℕ)) : (H.cliqueFinset 3).Nonempty := @@ -238,12 +242,16 @@ lemma farFromTriangleFree_of_disjoint_triangles (tris : Finset (Finset α)) (Nat.cast_le.2 $ farFromTriangleFree_of_disjoint_triangles_aux htris pd hG hH) protected lemma EdgeDisjointTriangles.farFromTriangleFree (hG : G.EdgeDisjointTriangles) - (tris_big : ε * (card α ^ 2 : ℕ) ≤ (G.cliqueFinset 3).card) : G.FarFromTriangleFree ε := + (tris_big : ε * (card α ^ 2 : ℕ) ≤ (G.cliqueFinset 3).card) : + G.FarFromTriangleFree ε := farFromTriangleFree_of_disjoint_triangles _ Subset.rfl (by simpa using hG) tris_big +end DecidableEq + variable [Nonempty α] lemma FarFromTriangleFree.lt_half (hG : G.FarFromTriangleFree ε) : ε < 2⁻¹ := by + classical by_contra! hε refine lt_irrefl (ε * card α ^ 2) ?_ have hε₀ : 0 < ε := hε.trans_lt' (by norm_num) @@ -281,8 +289,8 @@ theorem CliqueFree.not_farFromTriangleFree (hG : G.CliqueFree 3) (hε : 0 < ε) theorem FarFromTriangleFree.not_cliqueFree (hG : G.FarFromTriangleFree ε) (hε : 0 < ε) : ¬G.CliqueFree 3 := fun h => (hG.nonpos h).not_lt hε -theorem FarFromTriangleFree.cliqueFinset_nonempty (hG : G.FarFromTriangleFree ε) (hε : 0 < ε) : - (G.cliqueFinset 3).Nonempty := +theorem FarFromTriangleFree.cliqueFinset_nonempty [DecidableEq α] + (hG : G.FarFromTriangleFree ε) (hε : 0 < ε) : (G.cliqueFinset 3).Nonempty := nonempty_of_ne_empty <| cliqueFinset_eq_empty_iff.not.2 <| hG.not_cliqueFree hε end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Turan.lean b/Mathlib/Combinatorics/SimpleGraph/Turan.lean index ca69ed95ceed8..7309b4fe49ef4 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Turan.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Turan.lean @@ -126,10 +126,11 @@ end Defs namespace IsTuranMaximal -variable {s t u : V} (h : G.IsTuranMaximal r) +variable {s t u : V} /-- In a Turán-maximal graph, non-adjacent vertices have the same degree. -/ -lemma degree_eq_of_not_adj (hn : ¬G.Adj s t) : G.degree s = G.degree t := by +lemma degree_eq_of_not_adj (h : G.IsTuranMaximal r) (hn : ¬G.Adj s t) : + G.degree s = G.degree t := by rw [IsTuranMaximal] at h; contrapose! h; intro cf wlog hd : G.degree t < G.degree s generalizing G t s · replace hd : G.degree s < G.degree t := lt_of_le_of_ne (le_of_not_lt hd) h @@ -139,7 +140,8 @@ lemma degree_eq_of_not_adj (hn : ¬G.Adj s t) : G.degree s = G.degree t := by omega /-- In a Turán-maximal graph, non-adjacency is transitive. -/ -lemma not_adj_trans (hts : ¬G.Adj t s) (hsu : ¬G.Adj s u) : ¬G.Adj t u := by +lemma not_adj_trans (h : G.IsTuranMaximal r) (hts : ¬G.Adj t s) (hsu : ¬G.Adj s u) : + ¬G.Adj t u := by have hst : ¬G.Adj s t := fun a ↦ hts a.symm have dst := h.degree_eq_of_not_adj hst have dsu := h.degree_eq_of_not_adj hsu @@ -165,6 +167,8 @@ lemma not_adj_trans (hts : ¬G.Adj t s) (hsu : ¬G.Adj s u) : ¬G.Adj t u := by have l3 : 0 < G.degree u := by rw [G.degree_pos_iff_exists_adj u]; use t, h.symm omega +variable (h : G.IsTuranMaximal r) + /-- In a Turán-maximal graph, non-adjacency is an equivalence relation. -/ theorem equivalence_not_adj : Equivalence (¬G.Adj · ·) where refl := by simp @@ -255,7 +259,8 @@ theorem card_parts : h.finpartition.parts.card = min (Fintype.card V) r := by /-- **Turán's theorem**, forward direction. Any `r + 1`-cliquefree Turán-maximal graph on `n` vertices is isomorphic to `turanGraph n r`. -/ -theorem nonempty_iso_turanGraph : Nonempty (G ≃g turanGraph (Fintype.card V) r) := by +theorem nonempty_iso_turanGraph (h : G.IsTuranMaximal r) : + Nonempty (G ≃g turanGraph (Fintype.card V) r) := by obtain ⟨zm, zp⟩ := h.isEquipartition.exists_partPreservingEquiv use (Equiv.subtypeUnivEquiv mem_univ).symm.trans zm intro a b diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 1bad52fb250d8..7e6c9dafd5dc9 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -452,8 +452,8 @@ theorem ofRowLens_to_rowLens_eq_self {μ : YoungDiagram} : ofRowLens _ (rowLens_ /-- The right_inv direction of the equivalence -/ theorem rowLens_ofRowLens_eq_self {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (hpos : ∀ x ∈ w, 0 < x) : (ofRowLens w hw).rowLens = w := - List.ext_get (rowLens_length_ofRowLens hpos) fun i _ h₂ => - get_rowLens.trans <| rowLen_ofRowLens ⟨i, h₂⟩ + List.ext_get (rowLens_length_ofRowLens hpos) fun i h₁ h₂ => + (get_rowLens (h := h₁)).trans <| rowLen_ofRowLens ⟨i, h₂⟩ /-- Equivalence between Young diagrams and weakly decreasing lists of positive natural numbers. A Young diagram `μ` is equivalent to a list of row lengths. -/ diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index d67b71d5ea80d..6f0e19fdcf5c8 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -139,13 +139,13 @@ lemma isLittleO_self_div_log_id : (fun (n : ℕ) => n / log n ^ 2) =o[atTop] (fun (n : ℕ) => (n : ℝ)) := by calc (fun (n : ℕ) => (n : ℝ) / log n ^ 2) = fun (n : ℕ) => (n : ℝ) * ((log n) ^ 2)⁻¹ := by simp_rw [div_eq_mul_inv] - _ =o[atTop] fun (n : ℕ) => (n : ℝ) * 1⁻¹ := by + _ =o[atTop] fun (n : ℕ) => (n : ℝ) * 1⁻¹ := by refine IsBigO.mul_isLittleO (isBigO_refl _ _) ?_ refine IsLittleO.inv_rev ?main ?zero case zero => simp case main => calc _ = (fun (_ : ℕ) => ((1 : ℝ) ^ 2)) := by simp - _ =o[atTop] (fun (n : ℕ) => (log n)^2) := + _ =o[atTop] (fun (n : ℕ) => (log n)^2) := IsLittleO.pow (IsLittleO.natCast_atTop <| isLittleO_const_log_atTop) (by norm_num) _ = (fun (n : ℕ) => (n : ℝ)) := by ext; simp @@ -191,8 +191,8 @@ lemma eventually_bi_mul_le_r : ∀ᶠ (n : ℕ) in atTop, ∀ i, (b (min_bi b) / lemma bi_min_div_two_lt_one : b (min_bi b) / 2 < 1 := by have gt_zero : 0 < b (min_bi b) := R.b_pos (min_bi b) - calc b (min_bi b) / 2 < b (min_bi b) := by aesop (add safe apply div_two_lt_of_pos) - _ < 1 := R.b_lt_one _ + calc b (min_bi b) / 2 < b (min_bi b) := by aesop (add safe apply div_two_lt_of_pos) + _ < 1 := R.b_lt_one _ lemma bi_min_div_two_pos : 0 < b (min_bi b) / 2 := div_pos (R.b_pos _) (by norm_num) @@ -459,12 +459,12 @@ lemma strictAntiOn_one_add_smoothingFn : StrictAntiOn (fun (x : ℝ) => (1 : ℝ lemma isEquivalent_smoothingFn_sub_self (i : α) : (fun (n : ℕ) => ε (b i * n) - ε n) ~[atTop] fun n => -log (b i) / (log n)^2 := by calc (fun (n : ℕ) => 1 / log (b i * n) - 1 / log n) - =ᶠ[atTop] fun (n : ℕ) => (log n - log (b i * n)) / (log (b i * n) * log n) := by + =ᶠ[atTop] fun (n : ℕ) => (log n - log (b i * n)) / (log (b i * n) * log n) := by filter_upwards [eventually_gt_atTop 1, R.eventually_log_b_mul_pos] with n hn hn' have h_log_pos : 0 < log n := Real.log_pos <| by aesop simp only [one_div] rw [inv_sub_inv (by have := hn' i; positivity) (by aesop)] - _ =ᶠ[atTop] (fun (n : ℕ) => (log n - log (b i) - log n) / ((log (b i) + log n) * log n)) := by + _ =ᶠ[atTop] (fun (n : ℕ) ↦ (log n - log (b i) - log n) / ((log (b i) + log n) * log n)) := by filter_upwards [eventually_ne_atTop 0] with n hn have : 0 < b i := R.b_pos i rw [log_mul (by positivity) (by aesop), sub_add_eq_sub_sub] @@ -585,8 +585,8 @@ lemma asympBound_def' {n : ℕ} : simp [asympBound_def, sumTransform, mul_add, mul_one, Finset.sum_Ico_eq_sum_range] lemma asympBound_pos (n : ℕ) (hn : 0 < n) : 0 < asympBound g a b n := by - calc 0 < (n : ℝ) ^ p a b * (1 + 0) := by aesop (add safe Real.rpow_pos_of_pos) - _ ≤ asympBound g a b n := by + calc 0 < (n : ℝ) ^ p a b * (1 + 0) := by aesop (add safe Real.rpow_pos_of_pos) + _ ≤ asympBound g a b n := by simp only [asympBound_def'] gcongr n^p a b * (1 + ?_) have := R.g_nonneg @@ -656,8 +656,8 @@ lemma eventually_atTop_sumTransform_le : refine hn₂ u ?_ rw [Set.mem_Icc] refine ⟨?_, by norm_cast; omega⟩ - calc c₁ * n ≤ r i n := by exact hn₁ i - _ ≤ u := by exact_mod_cast hu'.1 + calc c₁ * n ≤ r i n := by exact hn₁ i + _ ≤ u := by exact_mod_cast hu'.1 _ ≤ n ^ (p a b) * (∑ _u ∈ Finset.Ico (r i n) n, c₂ * g n / n ^ ((p a b) + 1)) := by gcongr n ^ (p a b) * (Finset.Ico (r i n) n).sum (fun _ => c₂ * g n / ?_) with u hu rw [Finset.mem_Ico] at hu @@ -695,7 +695,7 @@ lemma eventually_atTop_sumTransform_ge : cases le_or_gt 0 (p a b + 1) with | inl hp => -- 0 ≤ (p a b) + 1 calc sumTransform (p a b) g (r i n) n - = n ^ (p a b) * (∑ u ∈ Finset.Ico (r i n) n, g u / u ^ ((p a b) + 1)) := by rfl + = n ^ (p a b) * (∑ u ∈ Finset.Ico (r i n) n, g u / u ^ ((p a b) + 1)) := rfl _ ≥ n ^ (p a b) * (∑ u ∈ Finset.Ico (r i n) n, c₂ * g n / u^((p a b) + 1)) := by gcongr with u hu rw [Finset.mem_Ico] at hu @@ -703,8 +703,8 @@ lemma eventually_atTop_sumTransform_ge : refine hn₂ u ?_ rw [Set.mem_Icc] refine ⟨?_, by norm_cast; omega⟩ - calc c₁ * n ≤ r i n := by exact hn₁ i - _ ≤ u := by exact_mod_cast hu'.1 + calc c₁ * n ≤ r i n := by exact hn₁ i + _ ≤ u := by exact_mod_cast hu'.1 _ ≥ n ^ (p a b) * (∑ _u ∈ Finset.Ico (r i n) n, c₂ * g n / n ^ ((p a b) + 1)) := by gcongr with u hu · rw [Finset.mem_Ico] at hu @@ -729,7 +729,7 @@ lemma eventually_atTop_sumTransform_ge : gcongr; exact min_le_left _ _ | inr hp => -- (p a b) + 1 < 0 calc sumTransform (p a b) g (r i n) n - = n ^ (p a b) * (∑ u ∈ Finset.Ico (r i n) n, g u / u^((p a b) + 1)) := by rfl + = n ^ (p a b) * (∑ u ∈ Finset.Ico (r i n) n, g u / u^((p a b) + 1)) := by rfl _ ≥ n ^ (p a b) * (∑ u ∈ Finset.Ico (r i n) n, c₂ * g n / u ^ ((p a b) + 1)) := by gcongr with u hu rw [Finset.mem_Ico] at hu @@ -1004,9 +1004,9 @@ lemma rpow_p_mul_one_sub_smoothingFn_le : filter_upwards [eventually_ne_atTop 0] with n hn have hn' : (n : ℝ) ≠ 0 := by positivity simp [← mul_div_assoc, ← Real.rpow_add_one hn'] - _ = fun (n : ℕ) => (n : ℝ) ^ (p a b) * (1 / (log n)^2) := by + _ = fun (n : ℕ) => (n : ℝ) ^ (p a b) * (1 / (log n)^2) := by simp_rw [mul_div, mul_one] - _ =Θ[atTop] fun (n : ℕ) => (b i) ^ (p a b) * n ^ (p a b) * (1 / (log n)^2) := by + _ =Θ[atTop] fun (n : ℕ) => (b i) ^ (p a b) * n ^ (p a b) * (1 / (log n)^2) := by refine IsTheta.symm ?_ simp_rw [mul_assoc] refine IsTheta.const_mul_left ?_ (isTheta_refl _ _) @@ -1017,7 +1017,7 @@ lemma rpow_p_mul_one_sub_smoothingFn_le : have h_main : (fun (n : ℕ) => q (r i n) - q (b i * n)) ≤ᶠ[atTop] fun (n : ℕ) => (b i) ^ (p a b) * n ^ (p a b) * (ε (b i * n) - ε n) := by calc (fun (n : ℕ) => q (r i n) - q (b i * n)) - ≤ᶠ[atTop] fun (n : ℕ) => ‖q (r i n) - q (b i * n)‖ := by + ≤ᶠ[atTop] fun (n : ℕ) => ‖q (r i n) - q (b i * n)‖ := by filter_upwards with _; exact le_norm_self _ _ ≤ᶠ[atTop] fun (n : ℕ) => ‖(b i) ^ (p a b) * n ^ (p a b) * (ε (b i * n) - ε n)‖ := h_main_norm @@ -1099,7 +1099,7 @@ lemma rpow_p_mul_one_add_smoothingFn_ge : filter_upwards [eventually_ne_atTop 0] with n hn have hn' : (n : ℝ) ≠ 0 := by positivity simp [← mul_div_assoc, ← Real.rpow_add_one hn'] - _ = fun (n : ℕ) => (n : ℝ) ^ (p a b) * (1 / (log n) ^ 2) := by simp_rw [mul_div, mul_one] + _ = fun (n : ℕ) => (n : ℝ) ^ (p a b) * (1 / (log n) ^ 2) := by simp_rw [mul_div, mul_one] _ =Θ[atTop] fun (n : ℕ) => (b i) ^ (p a b) * n ^ (p a b) * (1 / (log n) ^ 2) := by refine IsTheta.symm ?_ simp_rw [mul_assoc] @@ -1153,10 +1153,10 @@ lemma base_nonempty {n : ℕ} (hn : 0 < n) : (Finset.Ico (⌊b (min_bi b) / 2 * let b' := b (min_bi b) have hb_pos : 0 < b' := R.b_pos _ simp_rw [Finset.nonempty_Ico] - exact_mod_cast calc ⌊b' / 2 * n⌋₊ ≤ b' / 2 * n := by exact Nat.floor_le (by positivity) - _ < 1 / 2 * n := by gcongr; exact R.b_lt_one (min_bi b) - _ ≤ 1 * n := by gcongr; norm_num - _ = n := by simp + exact_mod_cast calc ⌊b' / 2 * n⌋₊ ≤ b' / 2 * n := by exact Nat.floor_le (by positivity) + _ < 1 / 2 * n := by gcongr; exact R.b_lt_one (min_bi b) + _ ≤ 1 * n := by gcongr; norm_num + _ = n := by simp /-- The main proof of the upper bound part of the Akra-Bazzi theorem. The factor `1 - ε n` does not change the asymptotic order, but is needed for the induction step to go @@ -1214,9 +1214,9 @@ lemma T_isBigO_smoothingFn_mul_asympBound : induction n using Nat.strongInductionOn with | ind n h_ind => have b_mul_n₀_le_ri i : ⌊b' * ↑n₀⌋₊ ≤ r i n := by - exact_mod_cast calc ⌊b' * (n₀ : ℝ)⌋₊ ≤ b' * n₀ := Nat.floor_le <| by positivity - _ ≤ b' * n := by gcongr - _ ≤ r i n := h_bi_le_r n hn i + exact_mod_cast calc ⌊b' * (n₀ : ℝ)⌋₊ ≤ b' * n₀ := Nat.floor_le <| by positivity + _ ≤ b' * n := by gcongr + _ ≤ r i n := h_bi_le_r n hn i have g_pos : 0 ≤ g n := R.g_nonneg n (by positivity) calc T n = (∑ i, a i * T (r i n)) + g n := by exact R.h_rec n <| n₀_ge_Rn₀.trans hn @@ -1452,7 +1452,7 @@ theorem isBigO_asympBound : T =O[atTop] asympBound g a b := by /-- The **Akra-Bazzi theorem**: `T ∈ Ω(n^p (1 + ∑_u^n g(u) / u^{p+1}))` -/ theorem isBigO_symm_asympBound : asympBound g a b =O[atTop] T := by - calc asympBound g a b = (fun n => 1 * asympBound g a b n) := by simp + calc asympBound g a b = (fun n => 1 * asympBound g a b n) := by simp _ ~[atTop] (fun n => (1 + ε n) * asympBound g a b n) := by refine IsEquivalent.mul (IsEquivalent.symm ?_) IsEquivalent.refl rw [Function.const_def, isEquivalent_const_iff_tendsto one_ne_zero, diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index d954c5efef75f..9777467e4dbd1 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -84,7 +84,7 @@ theorem merge' {f g : α →. σ} (hf : Partrec f) (hg : Partrec g) : simp only [k', exists_prop, mem_coe, mem_bind_iff, Option.mem_def] at h' obtain ⟨n, hn, hx⟩ := h' have := (H _).1 _ hn - simp [mem_decode₂, encode_injective.eq_iff] at this + simp only [decode₂_encode, coe_some, bind_some, mem_map_iff] at this obtain ⟨a', ha, rfl⟩ | ⟨a', ha, rfl⟩ := this <;> simp only [encodek, Option.some_inj] at hx <;> rw [hx] at ha · exact Or.inl ha @@ -214,8 +214,8 @@ theorem rice₂ (C : Set Code) (H : ∀ cf cg, eval cf = eval cg → (cf ∈ C (Partrec.nat_iff.1 <| eval_part.comp (const cf) Computable.id) (Partrec.nat_iff.1 <| eval_part.comp (const cg) Computable.id) ((hC _).1 fC), fun h => by { - obtain rfl | rfl := h <;> simp [ComputablePred, Set.mem_empty_iff_false] <;> - exact ⟨by infer_instance, Computable.const _⟩ }⟩ + obtain rfl | rfl := h <;> simpa [ComputablePred, Set.mem_empty_iff_false] using + ⟨by infer_instance, Computable.const _⟩ }⟩ /-- The Halting problem is recursively enumerable -/ theorem halting_problem_re (n) : RePred fun c => (eval c n).Dom := @@ -304,13 +304,13 @@ theorem tail {n f} (hf : @Partrec' n f) : @Partrec' n.succ fun v => f v.tail := protected theorem bind {n f g} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : @Partrec' n fun v => (f v).bind fun a => g (a ::ᵥ v) := (@comp n (n + 1) g (fun i => Fin.cases f (fun i v => some (v.get i)) i) hg fun i => by - refine Fin.cases ?_ (fun i => ?_) i <;> simp [*] - exact prim (Nat.Primrec'.get _)).of_eq + refine Fin.cases ?_ (fun i => ?_) i <;> simp [*] + exact prim (Nat.Primrec'.get _)).of_eq fun v => by simp [mOfFn, Part.bind_assoc, pure] protected theorem map {n f} {g : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : @Partrec' n fun v => (f v).map fun a => g (a ::ᵥ v) := by - simp [(Part.bind_some_eq_map _ _).symm]; exact hf.bind hg + simpa [(Part.bind_some_eq_map _ _).symm] using hf.bind hg /-- Analogous to `Nat.Partrec'` for `ℕ`-valued functions, a predicate for partial recursive vector-valued functions. -/ diff --git a/Mathlib/Computability/Partrec.lean b/Mathlib/Computability/Partrec.lean index 9d5f49c144d9f..2064a39e90bb2 100644 --- a/Mathlib/Computability/Partrec.lean +++ b/Mathlib/Computability/Partrec.lean @@ -34,9 +34,7 @@ variable (p : ℕ →. Bool) private def lbp (m n : ℕ) : Prop := m = n + 1 ∧ ∀ k ≤ n, false ∈ p k -variable (H : ∃ n, true ∈ p n ∧ ∀ k < n, (p k).Dom) - -private def wf_lbp : WellFounded (lbp p) := +private def wf_lbp (H : ∃ n, true ∈ p n ∧ ∀ k < n, (p k).Dom) : WellFounded (lbp p) := ⟨by let ⟨n, pn⟩ := H suffices ∀ m k, n ≤ k + m → Acc (lbp p) k by exact fun a => this _ _ (Nat.le_add_left _ _) @@ -45,6 +43,8 @@ private def wf_lbp : WellFounded (lbp p) := · injection mem_unique pn.1 (a _ kn) · exact IH _ (by rw [Nat.add_right_comm]; exact kn)⟩ +variable (H : ∃ n, true ∈ p n ∧ ∀ k < n, (p k).Dom) + def rfindX : { n // true ∈ p n ∧ ∀ m < n, false ∈ p m } := suffices ∀ k, (∀ n < k, false ∈ p n) → { n // true ∈ p n ∧ ∀ m < n, false ∈ p m } from this 0 fun n => (Nat.not_lt_zero _).elim @@ -492,7 +492,7 @@ theorem rfindOpt {f : α → ℕ → Option σ} (hf : Computable₂ f) : theorem nat_casesOn_right {f : α → ℕ} {g : α → σ} {h : α → ℕ →. σ} (hf : Computable f) (hg : Computable g) (hh : Partrec₂ h) : Partrec fun a => (f a).casesOn (some (g a)) (h a) := (nat_rec hf hg (hh.comp fst (pred.comp <| hf.comp fst)).to₂).of_eq fun a => by - simp; cases' f a with n <;> simp + simp only [PFun.coe_val, Nat.pred_eq_sub_one]; cases' f a with n <;> simp refine ext fun b => ⟨fun H => ?_, fun H => ?_⟩ · rcases mem_bind_iff.1 H with ⟨c, _, h₂⟩ exact h₂ @@ -629,7 +629,7 @@ theorem nat_strong_rec (f : α → ℕ → σ) {g : α → List σ → Option σ option_map (hg.comp (fst.comp <| fst.comp fst) snd) (to₂ <| list_concat.comp (snd.comp fst) snd))).of_eq fun a => by - simp; induction' a.2 with n IH; · rfl + induction' a.2 with n IH; · rfl simp [IH, H, List.range_succ] theorem list_ofFn : @@ -725,10 +725,9 @@ theorem fix_aux {α σ} (f : α →. σ ⊕ α) (a : α) (b : σ) : intro a₂ h₂ IH k hk rcases PFun.mem_fix_iff.1 h₂ with (h₂ | ⟨a₃, am₃, _⟩) · refine ⟨k.succ, ?_, fun m mk km => ⟨a₂, ?_⟩⟩ - · simp [F] - exact Or.inr ⟨_, hk, h₂⟩ + · simpa [F] using Or.inr ⟨_, hk, h₂⟩ · rwa [le_antisymm (Nat.le_of_lt_succ mk) km] - · rcases IH _ am₃ k.succ (by simp [F]; exact ⟨_, hk, am₃⟩) with ⟨n, hn₁, hn₂⟩ + · rcases IH _ am₃ k.succ (by simpa [F] using ⟨_, hk, am₃⟩) with ⟨n, hn₁, hn₂⟩ refine ⟨n, hn₁, fun m mn km => ?_⟩ cases' km.lt_or_eq_dec with km km · exact hn₂ _ mn km @@ -744,6 +743,6 @@ theorem fix {f : α →. σ ⊕ α} (hf : Partrec f) : Partrec (PFun.fix f) := b have hp : Partrec₂ p := hF.map ((sum_casesOn Computable.id (const true).to₂ (const false).to₂).comp snd).to₂ exact (hp.rfind.bind (hF.bind (sum_casesOn_right snd snd.to₂ none.to₂).to₂).to₂).of_eq fun a => - ext fun b => by simp [p]; apply fix_aux f + ext fun b => by simpa [p] using fix_aux f _ _ end Partrec diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index 5bb50ca1e5675..0bb7463dbe554 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -64,7 +64,7 @@ theorem rfind' {f} (hf : Nat.Partrec f) : (Primrec₂.pair.comp (Primrec.fst.comp <| Primrec.unpair.comp Primrec.fst) (Primrec.nat_add.comp Primrec.snd (Primrec.snd.comp <| Primrec.unpair.comp Primrec.fst))).to_comp)) - simp at this; exact this + simpa /-- Code for partial recursive functions from ℕ to ℕ. See `Nat.Partrec.Code.eval` for the interpretation of these constructors. @@ -79,7 +79,6 @@ inductive Code : Type | prec : Code → Code → Code | rfind' : Code → Code --- Porting note: `Nat.Partrec.Code.recOn` is noncomputable in Lean4, so we make it computable. compile_inductive% Code end Nat.Partrec @@ -315,7 +314,6 @@ theorem rec_prim' {α σ} [Primcodable α] [Primcodable σ] {c : α → Code} (h snd.pair <| nat_div2.comp <| nat_div2.comp snd refine (nat_strong_rec (fun a n => F a (ofNat Code n)) this.to₂ fun a n => ?_) |>.comp .id (encode_iff.2 hc) |>.of_eq fun a => by simp - simp iterate 4 cases' n with n; · simp [ofNatCode_eq, ofNatCode]; rfl simp only [G]; rw [List.length_map, List.length_range] let m := n.div2.div2 @@ -426,7 +424,6 @@ theorem rec_computable {α σ} [Primcodable α] [Primcodable σ] {c : α → Cod snd.pair <| nat_div2.comp <| nat_div2.comp snd refine (nat_strong_rec (fun a n => F a (ofNat Code n)) this.to₂ fun a n => ?_) |>.comp .id (encode_iff.2 hc) |>.of_eq fun a => by simp - simp iterate 4 cases' n with n; · simp [ofNatCode_eq, ofNatCode]; rfl simp only [G]; rw [List.length_map, List.length_range] let m := n.div2.div2 diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 8ebca3b6f50d9..28f554c1f9089 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -763,7 +763,7 @@ private theorem list_foldl' {f : α → List β} {g : α → σ} {h : α → σ hG) suffices ∀ a n, F a n = (((f a).take n).foldl (fun s b => h a (s, b)) (g a), (f a).drop n) by refine hF.of_eq fun a => ?_ - rw [this, List.take_all_of_le (length_le_encode _)] + rw [this, List.take_of_length_le (length_le_encode _)] introv dsimp only [F] generalize f a = l @@ -993,7 +993,7 @@ theorem nat_strong_rec (f : α → ℕ → σ) {g : α → List σ → Option σ option_map (hg.comp (fst.comp fst) snd) (to₂ <| list_concat.comp (snd.comp fst) snd))).of_eq fun a n => by - simp; induction' n with n IH; · rfl + induction' n with n IH; · rfl simp [IH, H, List.range_succ] theorem listLookup [DecidableEq α] : Primrec₂ (List.lookup : α → List (α × β) → Option β) := @@ -1191,7 +1191,7 @@ theorem vector_toList_iff {n} {f : α → Vector β n} : (Primrec fun a => (f a) subtype_val_iff theorem vector_cons {n} : Primrec₂ (@Vector.cons α n) := - vector_toList_iff.1 <| by simp; exact list_cons.comp fst (vector_toList_iff.2 snd) + vector_toList_iff.1 <| by simpa using list_cons.comp fst (vector_toList_iff.2 snd) theorem vector_length {n} : Primrec (@Vector.length α n) := const _ @@ -1212,7 +1212,7 @@ theorem list_ofFn : ∀ {n} {f : Fin n → α → σ}, (∀ i, Primrec (f i)) → Primrec fun a => List.ofFn fun i => f i a | 0, _, _ => const [] | n + 1, f, hf => by - simp [List.ofFn_succ]; exact list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) + simpa [List.ofFn_succ] using list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Primrec (f i)) : Primrec fun a => Vector.ofFn fun i => f i a := @@ -1357,7 +1357,7 @@ protected theorem encode : ∀ {n}, @Primrec' n encode theorem sqrt : @Primrec' 1 fun v => v.head.sqrt := by suffices H : ∀ n : ℕ, n.sqrt = n.rec 0 fun x y => if x.succ < y.succ * y.succ then y else y.succ by - simp [H] + simp only [H, succ_eq_add_one] have := @prec' 1 _ _ (fun v => by diff --git a/Mathlib/Computability/Reduce.lean b/Mathlib/Computability/Reduce.lean index f4ca560de6c97..32ee465296600 100644 --- a/Mathlib/Computability/Reduce.lean +++ b/Mathlib/Computability/Reduce.lean @@ -416,11 +416,11 @@ protected theorem add_le {d₁ d₂ d₃ : ManyOneDegree} : d₁ + d₂ ≤ d₃ @[simp] protected theorem le_add_left (d₁ d₂ : ManyOneDegree) : d₁ ≤ d₁ + d₂ := - (ManyOneDegree.add_le.1 (by rfl)).1 + (ManyOneDegree.add_le.1 (le_refl _)).1 @[simp] protected theorem le_add_right (d₁ d₂ : ManyOneDegree) : d₂ ≤ d₁ + d₂ := - (ManyOneDegree.add_le.1 (by rfl)).2 + (ManyOneDegree.add_le.1 (le_refl _)).2 instance instSemilatticeSup : SemilatticeSup ManyOneDegree := { ManyOneDegree.instPartialOrder with diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index 7ee610ea75e33..4400baf666d22 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -238,7 +238,7 @@ theorem mul_rmatch_iff (P Q : RegularExpression α) (x : List α) : subst hu repeat rw [rmatch] at h₂ simp [h₂] - · rw [rmatch]; simp [deriv] + · rw [rmatch]; simp only [deriv] split_ifs with hepsilon · rw [add_rmatch_iff, ih] constructor diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 762a216837a6e..d31e42e629f70 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -867,7 +867,6 @@ inductive Λ' | pred (q₁ q₂ : Λ') | ret (k : Cont') --- Porting note: `Turing.PartrecToTM2.Λ'.rec` is noncomputable in Lean4, so we make it computable. compile_inductive% Code compile_inductive% Cont' compile_inductive% K' @@ -1331,7 +1330,7 @@ theorem trNat_natEnd (n) : ∀ x ∈ trNat n, natEnd x = false := theorem trList_ne_consₗ : ∀ (l), ∀ x ∈ trList l, x ≠ Γ'.consₗ | a :: l, x, h => by - simp [trList] at h + simp only [trList, List.mem_append, List.mem_cons] at h obtain h | rfl | h := h · rintro rfl cases trNat_natEnd _ _ h @@ -1434,7 +1433,7 @@ theorem pred_ok (q₁ q₂ s v) (c d : List Γ') : ∃ s', simp only [TM2.step, trList, trNat.eq_1, trNum, Nat.cast_succ, Num.add_one, Num.succ, List.tail_cons, List.headI_cons] cases' (n : Num) with a - · simp [trPosNum, trNum, show Num.zero.succ' = PosNum.one from rfl] + · simp only [trPosNum, List.singleton_append, List.nil_append] refine TransGen.head rfl ?_ simp only [Option.mem_def, TM2.stepAux, elim_main, List.head?_cons, Option.some.injEq, decide_False, List.tail_cons, elim_update_main, ne_eq, Function.update_noteq, elim_rev, @@ -1572,7 +1571,7 @@ theorem tr_init (c v) : theorem tr_eval (c v) : eval (TM2.step tr) (init c v) = halt <$> Code.eval c v := by obtain ⟨i, h₁, h₂⟩ := tr_init c v refine Part.ext fun x => ?_ - rw [reaches_eval h₂.to_reflTransGen]; simp [-TM2.step] + rw [reaches_eval h₂.to_reflTransGen]; simp only [Part.map_eq_map, Part.mem_map_iff] refine ⟨fun h => ?_, ?_⟩ · obtain ⟨c, hc₁, hc₂⟩ := tr_eval_rev tr_respects h₁ h simp [stepNormal_eval] at hc₂ @@ -1769,7 +1768,7 @@ theorem supports_union {K₁ K₂ S} : Supports (K₁ ∪ K₂) S ↔ Supports K theorem supports_biUnion {K : Option Γ' → Finset Λ'} {S} : Supports (Finset.univ.biUnion K) S ↔ ∀ a, Supports (K a) S := by - simp [Supports]; apply forall_swap + simpa [Supports] using forall_swap theorem head_supports {S k q} (H : (q : Λ').Supports S) : (head k q).Supports S := fun _ => by dsimp only; split_ifs <;> exact H diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 6175146bba7d8..dd43164fc11a9 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -2308,7 +2308,7 @@ theorem tr_respects_aux₂ {k : K} {q : Stmt₂₁} {v : σ} {S : ∀ k, List ( TM1.stepAux (trStAct q o) v ((Tape.move Dir.right)^[(S k).length] (Tape.mk' ∅ (addBottom L))) = TM1.stepAux q v' ((Tape.move Dir.right)^[(S' k).length] (Tape.mk' ∅ (addBottom L'))) := by - dsimp only; simp; cases o with simp only [stWrite, stVar, trStAct, TM1.stepAux] + simp only [Function.update_same]; cases o with simp only [stWrite, stVar, trStAct, TM1.stepAux] | push f => have := Tape.write_move_right_n fun a : Γ' ↦ (a.1, update a.2 k (some (f v))) refine diff --git a/Mathlib/Condensed/TopCatAdjunction.lean b/Mathlib/Condensed/TopCatAdjunction.lean index 027e0369e98bc..c4f72257d9754 100644 --- a/Mathlib/Condensed/TopCatAdjunction.lean +++ b/Mathlib/Condensed/TopCatAdjunction.lean @@ -136,14 +136,14 @@ instance : topCatToCondensedSet.Faithful := topCatAdjunction.faithful_R_of_epi_c open CompactlyGenerated -instance (X : CondensedSet.{u}) : CompactlyGeneratedSpace.{u, u+1} X.toTopCat := by - apply compactlyGeneratedSpace_of_continuous_maps +instance (X : CondensedSet.{u}) : UCompactlyGeneratedSpace.{u, u+1} X.toTopCat := by + apply uCompactlyGeneratedSpace_of_continuous_maps intro Y _ f h rw [continuous_coinduced_dom, continuous_sigma_iff] exact fun ⟨S, s⟩ ↦ h S ⟨_, continuous_coinducingCoprod X _⟩ -instance (X : CondensedSet.{u}) : CompactlyGeneratedSpace.{u, u+1} (condensedSetToTopCat.obj X) := - inferInstanceAs (CompactlyGeneratedSpace.{u, u+1} X.toTopCat) +instance (X : CondensedSet.{u}) : UCompactlyGeneratedSpace.{u, u+1} (condensedSetToTopCat.obj X) := + inferInstanceAs (UCompactlyGeneratedSpace.{u, u+1} X.toTopCat) /-- The functor from condensed sets to topological spaces lands in compactly generated spaces. -/ def condensedSetToCompactlyGenerated : CondensedSet.{u} ⥤ CompactlyGenerated.{u, u+1} where @@ -172,12 +172,12 @@ noncomputable def compactlyGeneratedAdjunction : The counit of the adjunction `condensedSetToCompactlyGenerated ⊣ compactlyGeneratedToCondensedSet` is a homeomorphism. -/ -def compactlyGeneratedAdjunctionCounitHomeo (X : TopCat.{u+1}) [CompactlyGeneratedSpace.{u} X] : +def compactlyGeneratedAdjunctionCounitHomeo (X : TopCat.{u+1}) [UCompactlyGeneratedSpace.{u} X] : X.toCondensedSet.toTopCat ≃ₜ X where toEquiv := topCatAdjunctionCounitEquiv X continuous_toFun := (topCatAdjunctionCounit X).continuous continuous_invFun := by - apply continuous_from_compactlyGeneratedSpace + apply continuous_from_uCompactlyGeneratedSpace exact fun _ _ ↦ continuous_coinducingCoprod X.toCondensedSet _ /-- diff --git a/Mathlib/Control/LawfulFix.lean b/Mathlib/Control/LawfulFix.lean index dc39801dad874..d485806271845 100644 --- a/Mathlib/Control/LawfulFix.lean +++ b/Mathlib/Control/LawfulFix.lean @@ -158,9 +158,8 @@ theorem fix_le {X : (a : _) → Part <| β a} (hX : f X ≤ X) : Part.fix f ≤ · apply hX variable {f} -variable (hc : Continuous f) -theorem fix_eq : Part.fix f = f (Part.fix f) := by +theorem fix_eq (hc : Continuous f) : Part.fix f = f (Part.fix f) := by rw [fix_eq_ωSup f, hc] apply le_antisymm · apply ωSup_le_ωSup_of_le _ diff --git a/Mathlib/Control/Monad/Cont.lean b/Mathlib/Control/Monad/Cont.lean index b202aab6e3077..4a719782d7e74 100644 --- a/Mathlib/Control/Monad/Cont.lean +++ b/Mathlib/Control/Monad/Cont.lean @@ -104,7 +104,10 @@ instance (ε) [MonadExcept ε m] : MonadExcept ε (ContT r m) where end ContT -variable {m : Type u → Type v} [Monad m] +variable {m : Type u → Type v} + +section +variable [Monad m] def ExceptT.mkLabel {α β ε} : Label (Except.{u, u} ε α) m β → Label α (ExceptT ε m) β | ⟨f⟩ => ⟨fun a => monadLift <| f (Except.ok a)⟩ @@ -183,6 +186,8 @@ def WriterT.callCC' [MonadCont m] {α β ω : Type _} [Monoid ω] WriterT.mk <| MonadCont.callCC (WriterT.run ∘ f ∘ WriterT.mkLabel' : Label (α × ω) m β → m (α × ω)) +end + instance (ω) [Monad m] [EmptyCollection ω] [MonadCont m] : MonadCont (WriterT ω m) where callCC := WriterT.callCC @@ -202,7 +207,7 @@ nonrec def StateT.callCC {σ} [MonadCont m] {α β : Type _} instance {σ} [MonadCont m] : MonadCont (StateT σ m) where callCC := StateT.callCC -instance {σ} [MonadCont m] [LawfulMonadCont m] : LawfulMonadCont (StateT σ m) where +instance {σ} [Monad m] [MonadCont m] [LawfulMonadCont m] : LawfulMonadCont (StateT σ m) where callCC_bind_right := by intros simp only [callCC, StateT.callCC, StateT.run_bind, callCC_bind_right]; ext; rfl @@ -228,7 +233,7 @@ nonrec def ReaderT.callCC {ε} [MonadCont m] {α β : Type _} instance {ρ} [MonadCont m] : MonadCont (ReaderT ρ m) where callCC := ReaderT.callCC -instance {ρ} [MonadCont m] [LawfulMonadCont m] : LawfulMonadCont (ReaderT ρ m) where +instance {ρ} [Monad m] [MonadCont m] [LawfulMonadCont m] : LawfulMonadCont (ReaderT ρ m) where callCC_bind_right := by intros; simp only [callCC, ReaderT.callCC, ReaderT.run_bind, callCC_bind_right]; ext; rfl callCC_bind_left := by diff --git a/Mathlib/Control/Monad/Writer.lean b/Mathlib/Control/Monad/Writer.lean index f02096455cd33..46479df03cf3d 100644 --- a/Mathlib/Control/Monad/Writer.lean +++ b/Mathlib/Control/Monad/Writer.lean @@ -33,14 +33,14 @@ class MonadWriter (ω : outParam (Type u)) (M : Type u → Type v) where export MonadWriter (tell listen pass) -variable {M : Type u → Type v} [Monad M] {α ω ρ σ : Type u} +variable {M : Type u → Type v} {α ω ρ σ : Type u} instance [MonadWriter ω M] : MonadWriter ω (ReaderT ρ M) where tell w := (tell w : M _) listen x r := listen <| x r pass x r := pass <| x r -instance [MonadWriter ω M] : MonadWriter ω (StateT σ M) where +instance [Monad M] [MonadWriter ω M] : MonadWriter ω (StateT σ M) where tell w := (tell w : M _) listen x s := (fun ((a,w), s) ↦ ((a,s), w)) <$> listen (x s) pass x s := pass <| (fun ((a, f), s) ↦ ((a, s), f)) <$> (x s) @@ -57,7 +57,7 @@ protected def runThe (ω : Type u) (cmd : WriterT ω M α) : M (α × ω) := cmd @[ext] protected theorem ext {ω : Type u} (x x' : WriterT ω M α) (h : x.run = x'.run) : x = x' := h -variable {ω : Type u} {α β : Type u} +variable {ω : Type u} {α β : Type u} [Monad M] /-- Creates an instance of Monad, with an explicitly given empty and append operation. diff --git a/Mathlib/Data/Analysis/Topology.lean b/Mathlib/Data/Analysis/Topology.lean index 1e7c2003dc521..f8db5f20dd999 100644 --- a/Mathlib/Data/Analysis/Topology.lean +++ b/Mathlib/Data/Analysis/Topology.lean @@ -153,10 +153,8 @@ theorem ext [T : TopologicalSpace α] {σ : Type*} {F : Ctop α σ} (H₁ : ∀ variable [TopologicalSpace α] --- Porting note: add non-computable : because --- > ... it depends on `Inter.inter`, and it does not have executable code. /-- The topological space realizer made of the open sets. -/ -protected noncomputable def id : Realizer α := +protected def id : Realizer α := ⟨{ x : Set α // IsOpen x }, { f := Subtype.val top := fun _ ↦ ⟨univ, isOpen_univ⟩ diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index c90acc6869201..2c12c922b8012 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -5,13 +5,13 @@ Authors: Johannes Hölzl, Kenny Lau -/ import Mathlib.Algebra.BigOperators.GroupWithZero.Finset import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.GroupWithZero.Action.Pi +import Mathlib.Algebra.Module.LinearMap.Defs import Mathlib.Data.Finset.Preimage import Mathlib.Data.Fintype.Quotient import Mathlib.Data.Set.Finite import Mathlib.GroupTheory.GroupAction.BigOperators -import Mathlib.GroupTheory.GroupAction.Pi import Mathlib.Order.ConditionallyCompleteLattice.Basic -import Mathlib.Algebra.Module.LinearMap.Defs /-! # Dependent functions with finite support diff --git a/Mathlib/Data/DFinsupp/WellFounded.lean b/Mathlib/Data/DFinsupp/WellFounded.lean index ca6413ccb83d3..e10cffb7ca8e8 100644 --- a/Mathlib/Data/DFinsupp/WellFounded.lean +++ b/Mathlib/Data/DFinsupp/WellFounded.lean @@ -105,12 +105,12 @@ theorem Lex.acc_of_single_erase [DecidableEq ι] {x : Π₀ i, α i} (i : ι) (InvImage.accessible snd <| hs.prod_gameAdd hu) convert piecewise_single_erase x i -variable (hbot : ∀ ⦃i a⦄, ¬s i a 0) -theorem Lex.acc_zero : Acc (DFinsupp.Lex r s) 0 := +theorem Lex.acc_zero (hbot : ∀ ⦃i a⦄, ¬s i a 0) : Acc (DFinsupp.Lex r s) 0 := Acc.intro 0 fun _ ⟨_, _, h⟩ => (hbot h).elim -theorem Lex.acc_of_single [DecidableEq ι] [∀ (i) (x : α i), Decidable (x ≠ 0)] (x : Π₀ i, α i) : +theorem Lex.acc_of_single (hbot : ∀ ⦃i a⦄, ¬s i a 0) [DecidableEq ι] + [∀ (i) (x : α i), Decidable (x ≠ 0)] (x : Π₀ i, α i) : (∀ i ∈ x.support, Acc (DFinsupp.Lex r s) <| single i (x i)) → Acc (DFinsupp.Lex r s) x := by generalize ht : x.support = t; revert x classical @@ -123,9 +123,8 @@ theorem Lex.acc_of_single [DecidableEq ι] [∀ (i) (x : α i), Decidable (x ≠ rw [erase_ne (ha.ne_of_not_mem hb)] exact h a (Finset.mem_insert_of_mem ha) -variable (hs : ∀ i, WellFounded (s i)) - -theorem Lex.acc_single [DecidableEq ι] {i : ι} (hi : Acc (rᶜ ⊓ (· ≠ ·)) i) : +theorem Lex.acc_single (hbot : ∀ ⦃i a⦄, ¬s i a 0) (hs : ∀ i, WellFounded (s i)) + [DecidableEq ι] {i : ι} (hi : Acc (rᶜ ⊓ (· ≠ ·)) i) : ∀ a, Acc (DFinsupp.Lex r s) (single i a) := by induction' hi with i _ ih refine fun a => WellFounded.induction (hs i) @@ -146,14 +145,17 @@ theorem Lex.acc_single [DecidableEq ι] {i : ι} (hi : Acc (rᶜ ⊓ (· ≠ ·) exact Lex.acc_zero hbot · exact ih _ ⟨h, hij.symm⟩ _ -theorem Lex.acc [DecidableEq ι] [∀ (i) (x : α i), Decidable (x ≠ 0)] (x : Π₀ i, α i) +theorem Lex.acc (hbot : ∀ ⦃i a⦄, ¬s i a 0) (hs : ∀ i, WellFounded (s i)) + [DecidableEq ι] [∀ (i) (x : α i), Decidable (x ≠ 0)] (x : Π₀ i, α i) (h : ∀ i ∈ x.support, Acc (rᶜ ⊓ (· ≠ ·)) i) : Acc (DFinsupp.Lex r s) x := Lex.acc_of_single hbot x fun i hi => Lex.acc_single hbot hs (h i hi) _ -theorem Lex.wellFounded (hr : WellFounded <| rᶜ ⊓ (· ≠ ·)) : WellFounded (DFinsupp.Lex r s) := +theorem Lex.wellFounded (hbot : ∀ ⦃i a⦄, ¬s i a 0) (hs : ∀ i, WellFounded (s i)) + (hr : WellFounded <| rᶜ ⊓ (· ≠ ·)) : WellFounded (DFinsupp.Lex r s) := ⟨fun x => by classical exact Lex.acc hbot hs x fun i _ => hr.apply i⟩ -theorem Lex.wellFounded' [IsTrichotomous ι r] (hr : WellFounded (Function.swap r)) : +theorem Lex.wellFounded' (hbot : ∀ ⦃i a⦄, ¬s i a 0) (hs : ∀ i, WellFounded (s i)) + [IsTrichotomous ι r] (hr : WellFounded (Function.swap r)) : WellFounded (DFinsupp.Lex r s) := Lex.wellFounded hbot hs <| Subrelation.wf (fun {i j} h => ((@IsTrichotomous.trichotomous ι r _ i j).resolve_left h.1).resolve_left h.2) hr diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index 26b8b4fe62200..83efe7eab8e81 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -43,7 +43,23 @@ lemma coe_iSup : BddAbove (range f) → ↑(⨆ i, f i) = ⨆ i, (f i : ℕ∞) @[norm_cast] lemma coe_iInf [Nonempty ι] : ↑(⨅ i, f i) = ⨅ i, (f i : ℕ∞) := WithTop.coe_iInf (OrderBot.bddBelow _) -variable {s : Set ℕ∞} +@[simp] +lemma iInf_eq_top_of_isEmpty [IsEmpty ι] : ⨅ i, (f i : ℕ∞) = ⊤ := + iInf_coe_eq_top.mpr ‹_› + +lemma iInf_toNat : (⨅ i, (f i : ℕ∞)).toNat = ⨅ i, f i := by + cases isEmpty_or_nonempty ι + · simp + · norm_cast + +lemma iInf_eq_zero : ⨅ i, (f i : ℕ∞) = 0 ↔ ∃ i, f i = 0 := by + cases isEmpty_or_nonempty ι + · simp + · norm_cast + rw [iInf, Nat.sInf_eq_zero] + exact ⟨fun h ↦ by simp_all, .inl⟩ + +variable {f : ι → ℕ∞} {s : Set ℕ∞} lemma sSup_eq_zero : sSup s = 0 ↔ ∀ a ∈ s, a = 0 := sSup_eq_bot diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 471b289ab0ed9..a2b9e02f3bc43 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -120,7 +120,7 @@ theorem antidiagonalTuple_zero_right : ∀ k, antidiagonalTuple k 0 = [0] @[simp] theorem antidiagonalTuple_one (n : ℕ) : antidiagonalTuple 1 n = [![n]] := by simp_rw [antidiagonalTuple, antidiagonal, List.range_succ, List.map_append, List.map_singleton, - tsub_self, List.append_bind, List.bind_singleton, List.bind_map] + tsub_self, List.bind_append, List.bind_singleton, List.bind_map] conv_rhs => rw [← List.nil_append [![n]]] congr 1 simp_rw [List.bind_eq_nil, List.mem_range, List.map_eq_nil] diff --git a/Mathlib/Data/FinEnum.lean b/Mathlib/Data/FinEnum.lean index ce48da840ffa3..82688911d453d 100644 --- a/Mathlib/Data/FinEnum.lean +++ b/Mathlib/Data/FinEnum.lean @@ -3,9 +3,9 @@ Copyright (c) 2019 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ -import Mathlib.Control.Monad.Basic import Mathlib.Data.Fintype.Basic import Mathlib.Data.List.ProdSigma +import Mathlib.Data.List.Pi /-! Type class for finitely enumerable types. The property is stronger @@ -68,7 +68,7 @@ theorem nodup_toList [FinEnum α] : List.Nodup (toList α) := by /-- create a `FinEnum` instance using a surjection -/ def ofSurjective {β} (f : β → α) [DecidableEq α] [FinEnum β] (h : Surjective f) : FinEnum α := - ofList ((toList β).map f) (by intro; simp; exact h _) + ofList ((toList β).map f) (by intro; simpa using h _) /-- create a `FinEnum` instance using an injection -/ noncomputable def ofInjective {α β} (f : α → β) [DecidableEq α] [FinEnum β] (h : Injective f) : @@ -176,55 +176,30 @@ instance (priority := 100) [FinEnum α] : Fintype α where elems := univ.map (equiv).symm.toEmbedding complete := by intros; simp -/-- For `Pi.cons x xs y f` create a function where every `i ∈ xs` is mapped to `f i` and -`x` is mapped to `y` -/ -def Pi.cons [DecidableEq α] (x : α) (xs : List α) (y : β x) (f : ∀ a, a ∈ xs → β a) : - ∀ a, a ∈ (x :: xs : List α) → β a - | b, h => if h' : b = x then cast (by rw [h']) y else f b (List.mem_of_ne_of_mem h' h) - -/-- Given `f` a function whose domain is `x :: xs`, produce a function whose domain -is restricted to `xs`. -/ -def Pi.tail {x : α} {xs : List α} (f : ∀ a, a ∈ (x :: xs : List α) → β a) : ∀ a, a ∈ xs → β a - | a, h => f a (List.mem_cons_of_mem _ h) - -/-- `pi xs f` creates the list of functions `g` such that, for `x ∈ xs`, `g x ∈ f x` -/ -def pi {β : α → Type max u v} [DecidableEq α] : - ∀ xs : List α, (∀ a, List (β a)) → List (∀ a, a ∈ xs → β a) - | [], _ => [fun x h => (List.not_mem_nil x h).elim] - | x :: xs, fs => FinEnum.Pi.cons x xs <$> fs x <*> pi xs fs - -theorem mem_pi {β : α → Type _} [FinEnum α] [∀ a, FinEnum (β a)] (xs : List α) - (f : ∀ a, a ∈ xs → β a) : f ∈ pi xs fun x => toList (β x) := by - induction' xs with xs_hd xs_tl xs_ih <;> simp [pi, -List.map_eq_map, monad_norm, functor_norm] - · ext a ⟨⟩ - · exists Pi.cons xs_hd xs_tl (f _ (List.mem_cons_self _ _)) - constructor - · exact ⟨_, rfl⟩ - exists Pi.tail f - constructor - · apply xs_ih - · ext x h - simp only [Pi.cons] - split_ifs - · subst x - rfl - · rfl +end FinEnum + +namespace List +variable {α : Type*} [FinEnum α] {β : α → Type*} [∀ a, FinEnum (β a)] +open FinEnum + +theorem mem_pi_toList (xs : List α) + (f : ∀ a, a ∈ xs → β a) : f ∈ pi xs fun x => toList (β x) := + (mem_pi _ _).mpr fun _ _ ↦ mem_toList _ /-- enumerate all functions whose domain and range are finitely enumerable -/ -def pi.enum (β : α → Type (max u v)) [FinEnum α] [∀ a, FinEnum (β a)] : List (∀ a, β a) := - (pi.{u, v} (toList α) fun x => toList (β x)).map (fun f x => f x (mem_toList _)) +def Pi.enum (β : α → Type*) [∀ a, FinEnum (β a)] : List (∀ a, β a) := + (pi (toList α) fun x => toList (β x)).map (fun f x => f x (mem_toList _)) -theorem pi.mem_enum {β : α → Type (max u v)} [FinEnum α] [∀ a, FinEnum (β a)] (f : ∀ a, β a) : - f ∈ pi.enum.{u, v} β := by simp [pi.enum]; refine ⟨fun a _ => f a, mem_pi _ _, rfl⟩ +theorem Pi.mem_enum (f : ∀ a, β a) : + f ∈ Pi.enum β := by simpa [Pi.enum] using ⟨fun a _ => f a, mem_pi_toList _ _, rfl⟩ -instance pi.finEnum {β : α → Type (max u v)} [FinEnum α] [∀ a, FinEnum (β a)] : - FinEnum (∀ a, β a) := - ofList (pi.enum.{u, v} _) fun _ => pi.mem_enum _ +instance Pi.finEnum : FinEnum (∀ a, β a) := + ofList (Pi.enum _) fun _ => Pi.mem_enum _ instance pfunFinEnum (p : Prop) [Decidable p] (α : p → Type) [∀ hp, FinEnum (α hp)] : FinEnum (∀ hp : p, α hp) := if hp : p then - ofList ((toList (α hp)).map fun x _ => x) (by intro x; simp; exact ⟨x hp, rfl⟩) + ofList ((toList (α hp)).map fun x _ => x) (by intro x; simpa using ⟨x hp, rfl⟩) else ofList [fun hp' => (hp hp').elim] (by intro; simp; ext hp'; cases hp hp') -end FinEnum +end List diff --git a/Mathlib/Data/Finite/Card.lean b/Mathlib/Data/Finite/Card.lean index bd58be6e8a76f..d993c3f18a6bc 100644 --- a/Mathlib/Data/Finite/Card.lean +++ b/Mathlib/Data/Finite/Card.lean @@ -188,14 +188,16 @@ theorem card_union_le (s t : Set α) : Nat.card (↥(s ∪ t)) ≤ Nat.card s + namespace Finite -variable {s t : Set α} (ht : t.Finite) +variable {s t : Set α} -theorem card_lt_card (hsub : s ⊂ t) : Nat.card s < Nat.card t := by +theorem card_lt_card (ht : t.Finite) (hsub : s ⊂ t) : Nat.card s < Nat.card t := by have : Fintype t := Finite.fintype ht have : Fintype s := Finite.fintype (subset ht (subset_of_ssubset hsub)) simp only [Nat.card_eq_fintype_card] exact Set.card_lt_card hsub +variable (ht : t.Finite) + theorem eq_of_subset_of_card_le (hsub : s ⊆ t) (hcard : Nat.card t ≤ Nat.card s) : s = t := (eq_or_ssubset_of_subset hsub).elim id fun h ↦ absurd hcard <| not_le_of_lt <| ht.card_lt_card h diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index c6dcdab6ca7b0..33cbfd541c26d 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -305,7 +305,6 @@ theorem subset_def : s ⊆ t ↔ s.1 ⊆ t.1 := theorem ssubset_def : s ⊂ t ↔ s ⊆ t ∧ ¬t ⊆ s := Iff.rfl -@[simp] theorem Subset.refl (s : Finset α) : s ⊆ s := Multiset.Subset.refl _ @@ -2801,9 +2800,7 @@ theorem mem_toList {a : α} {s : Finset α} : a ∈ s.toList ↔ a ∈ s := theorem toList_eq_nil {s : Finset α} : s.toList = [] ↔ s = ∅ := Multiset.toList_eq_nil.trans val_eq_zero -@[simp] -theorem empty_toList {s : Finset α} : s.toList.isEmpty ↔ s = ∅ := - List.isEmpty_iff_eq_nil.trans toList_eq_nil +theorem empty_toList {s : Finset α} : s.toList.isEmpty ↔ s = ∅ := by simp @[simp] theorem toList_empty : (∅ : Finset α).toList = [] := diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 51fe7ad3257c3..16af175533226 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -162,18 +162,25 @@ lemma noncommProd_induction (s : Multiset α) (comm) variable [FunLike F α β] @[to_additive] -protected theorem noncommProd_map_aux [MonoidHomClass F α β] (s : Multiset α) +protected theorem map_noncommProd_aux [MonoidHomClass F α β] (s : Multiset α) (comm : { x | x ∈ s }.Pairwise Commute) (f : F) : { x | x ∈ s.map f }.Pairwise Commute := by simp only [Multiset.mem_map] rintro _ ⟨x, hx, rfl⟩ _ ⟨y, hy, rfl⟩ _ exact (comm.of_refl hx hy).map f @[to_additive] -theorem noncommProd_map [MonoidHomClass F α β] (s : Multiset α) (comm) (f : F) : - f (s.noncommProd comm) = (s.map f).noncommProd (Multiset.noncommProd_map_aux s comm f) := by +theorem map_noncommProd [MonoidHomClass F α β] (s : Multiset α) (comm) (f : F) : + f (s.noncommProd comm) = (s.map f).noncommProd (Multiset.map_noncommProd_aux s comm f) := by induction s using Quotient.inductionOn simpa using map_list_prod f _ +@[deprecated (since := "2024-07-23")] alias noncommProd_map := map_noncommProd +@[deprecated (since := "2024-07-23")] alias noncommSum_map := map_noncommSum +@[deprecated (since := "2024-07-23")] +protected alias noncommProd_map_aux := Multiset.map_noncommProd_aux +@[deprecated (since := "2024-07-23")] +protected alias noncommSum_map_aux := Multiset.map_noncommSum_aux + @[to_additive noncommSum_eq_card_nsmul] theorem noncommProd_eq_pow_card (s : Multiset α) (comm) (m : α) (h : ∀ x ∈ s, x = m) : s.noncommProd comm = m ^ Multiset.card s := by @@ -308,10 +315,13 @@ theorem noncommProd_singleton (a : α) (f : α → β) : variable [FunLike F β γ] @[to_additive] -theorem noncommProd_map [MonoidHomClass F β γ] (s : Finset α) (f : α → β) (comm) (g : F) : +theorem map_noncommProd [MonoidHomClass F β γ] (s : Finset α) (f : α → β) (comm) (g : F) : g (s.noncommProd f comm) = s.noncommProd (fun i => g (f i)) fun x hx y hy _ => (comm.of_refl hx hy).map g := by - simp [noncommProd, Multiset.noncommProd_map] + simp [noncommProd, Multiset.map_noncommProd] + +@[deprecated (since := "2024-07-23")] alias noncommProd_map := map_noncommProd +@[deprecated (since := "2024-07-23")] alias noncommSum_map := map_noncommSum @[to_additive noncommSum_eq_card_nsmul] theorem noncommProd_eq_pow_card (s : Finset α) (f : α → β) (comm) (m : β) (h : ∀ x ∈ s, f x = m) : @@ -402,7 +412,7 @@ theorem noncommProd_mul_single [Fintype ι] [DecidableEq ι] (x : ∀ i, M i) : (univ.noncommProd (fun i => Pi.mulSingle i (x i)) fun i _ j _ _ => Pi.mulSingle_apply_commute x i j) = x := by ext i - apply (univ.noncommProd_map (fun i ↦ MonoidHom.mulSingle M i (x i)) ?a + apply (univ.map_noncommProd (fun i ↦ MonoidHom.mulSingle M i (x i)) ?a (Pi.evalMonoidHom M i)).trans case a => intro i _ j _ _ @@ -426,7 +436,7 @@ theorem _root_.MonoidHom.pi_ext [Finite ι] [DecidableEq ι] {f g : (∀ i, M i) (h : ∀ i x, f (Pi.mulSingle i x) = g (Pi.mulSingle i x)) : f = g := by cases nonempty_fintype ι ext x - rw [← noncommProd_mul_single x, univ.noncommProd_map, univ.noncommProd_map] + rw [← noncommProd_mul_single x, univ.map_noncommProd, univ.map_noncommProd] congr 1 with i; exact h i (x i) end FinitePi diff --git a/Mathlib/Data/Finset/Pi.lean b/Mathlib/Data/Finset/Pi.lean index 24b46e6ce8874..2a31ffb1eecdd 100644 --- a/Mathlib/Data/Finset/Pi.lean +++ b/Mathlib/Data/Finset/Pi.lean @@ -8,6 +8,10 @@ import Mathlib.Data.Multiset.Pi /-! # The cartesian product of finsets + +## Main definitions + +* `Finset.pi`: Cartesian product of finsets indexed by a finset. -/ open Function diff --git a/Mathlib/Data/Finset/PiInduction.lean b/Mathlib/Data/Finset/PiInduction.lean index 8cab081b0fa03..d040789ab2d0d 100644 --- a/Mathlib/Data/Finset/PiInduction.lean +++ b/Mathlib/Data/Finset/PiInduction.lean @@ -103,6 +103,6 @@ theorem induction_on_pi_min [∀ i, LinearOrder (α i)] {p : (∀ i, Finset (α ∀ (g : ∀ i, Finset (α i)) (i : ι) (x : α i), (∀ y ∈ g i, x < y) → p g → p (update g i (insert x (g i)))) : p f := - @induction_on_pi_max ι (fun i ↦ (α i)ᵒᵈ) _ _ _ _ _ _ h0 step + induction_on_pi_max (α := fun i ↦ (α i)ᵒᵈ) _ h0 step end Finset diff --git a/Mathlib/Data/Finset/SMulAntidiagonal.lean b/Mathlib/Data/Finset/SMulAntidiagonal.lean index c6d6f6b0a1d17..86b0c882b6255 100644 --- a/Mathlib/Data/Finset/SMulAntidiagonal.lean +++ b/Mathlib/Data/Finset/SMulAntidiagonal.lean @@ -27,8 +27,8 @@ open Pointwise namespace Set @[to_additive] -theorem IsPWO.smul [PartialOrder G] [PartialOrder P] [SMul G P] [IsOrderedCancelSMul G P] {s : Set G} - {t : Set P} (hs : s.IsPWO) (ht : t.IsPWO) : IsPWO (s • t) := by +theorem IsPWO.smul [PartialOrder G] [PartialOrder P] [SMul G P] [IsOrderedCancelSMul G P] + {s : Set G} {t : Set P} (hs : s.IsPWO) (ht : t.IsPWO) : IsPWO (s • t) := by rw [← @image_smul_prod] exact (hs.prod ht).image_of_monotone (monotone_fst.smul monotone_snd) diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 9a938afe77d30..74ecc12d9a715 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -181,8 +181,9 @@ alias _root_.MvPolynomial.degree_eq_zero_iff := degree_eq_zero_iff @[simp] theorem degree_zero : degree (0 : σ →₀ ℕ) = 0 := by rw [degree_eq_zero_iff] -theorem degree_eq_weight_one (d : σ →₀ ℕ) : - degree d = weight 1 d := by +theorem degree_eq_weight_one : + degree (σ := σ) = weight 1 := by + ext d simp only [degree, weight_apply, Pi.one_apply, smul_eq_mul, mul_one, Finsupp.sum] @[deprecated degree_eq_weight_one (since := "2024-07-20")] diff --git a/Mathlib/Data/Finsupp/WellFounded.lean b/Mathlib/Data/Finsupp/WellFounded.lean index 86acc681a9b21..01fa1cb138541 100644 --- a/Mathlib/Data/Finsupp/WellFounded.lean +++ b/Mathlib/Data/Finsupp/WellFounded.lean @@ -27,23 +27,24 @@ variable {α N : Type*} namespace Finsupp -variable [Zero N] {r : α → α → Prop} {s : N → N → Prop} (hbot : ∀ ⦃n⦄, ¬s n 0) - (hs : WellFounded s) +variable [Zero N] {r : α → α → Prop} {s : N → N → Prop} /-- Transferred from `DFinsupp.Lex.acc`. See the top of that file for an explanation for the appearance of the relation `rᶜ ⊓ (≠)`. -/ -theorem Lex.acc (x : α →₀ N) (h : ∀ a ∈ x.support, Acc (rᶜ ⊓ (· ≠ ·)) a) : +theorem Lex.acc (hbot : ∀ ⦃n⦄, ¬s n 0) (hs : WellFounded s) (x : α →₀ N) + (h : ∀ a ∈ x.support, Acc (rᶜ ⊓ (· ≠ ·)) a) : Acc (Finsupp.Lex r s) x := by rw [lex_eq_invImage_dfinsupp_lex] classical refine InvImage.accessible toDFinsupp (DFinsupp.Lex.acc (fun _ => hbot) (fun _ => hs) _ ?_) simpa only [toDFinsupp_support] using h -theorem Lex.wellFounded (hr : WellFounded <| rᶜ ⊓ (· ≠ ·)) : WellFounded (Finsupp.Lex r s) := +theorem Lex.wellFounded (hbot : ∀ ⦃n⦄, ¬s n 0) (hs : WellFounded s) + (hr : WellFounded <| rᶜ ⊓ (· ≠ ·)) : WellFounded (Finsupp.Lex r s) := ⟨fun x => Lex.acc hbot hs x fun a _ => hr.apply a⟩ -theorem Lex.wellFounded' [IsTrichotomous α r] (hr : WellFounded (Function.swap r)) : - WellFounded (Finsupp.Lex r s) := +theorem Lex.wellFounded' (hbot : ∀ ⦃n⦄, ¬s n 0) (hs : WellFounded s) + [IsTrichotomous α r] (hr : WellFounded (Function.swap r)) : WellFounded (Finsupp.Lex r s) := (lex_eq_invImage_dfinsupp_lex r s).symm ▸ InvImage.wf _ (DFinsupp.Lex.wellFounded' (fun _ => hbot) (fun _ => hs) hr) diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index f00e3af6445d9..14bdf3e4149ee 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -959,7 +959,7 @@ variable [Fintype α] (p : α → Prop) [DecidablePred p] /-- Given a fintype `α` and a predicate `p`, associate to a proof that there is a unique element of `α` satisfying `p` this unique element, as an element of the corresponding subtype. -/ def chooseX (hp : ∃! a : α, p a) : { a // p a } := - ⟨Finset.choose p univ (by simp; exact hp), Finset.choose_property _ _ _⟩ + ⟨Finset.choose p univ (by simpa), Finset.choose_property _ _ _⟩ /-- Given a fintype `α` and a predicate `p`, associate to a proof that there is a unique element of `α` satisfying `p` this unique element, as an element of `α`. -/ diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index e429515015b5e..63d6af4844767 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -431,7 +431,8 @@ theorem card_lt_of_injective_of_not_mem (f : α → β) (h : Function.Injective calc card α = (univ.map ⟨f, h⟩).card := (card_map _).symm _ < card β := - Finset.card_lt_univ_of_not_mem <| by rwa [← mem_coe, coe_map, coe_univ, Set.image_univ] + Finset.card_lt_univ_of_not_mem (x := b) <| by + rwa [← mem_coe, coe_map, coe_univ, Set.image_univ] theorem card_lt_of_injective_not_surjective (f : α → β) (h : Function.Injective f) (h' : ¬Function.Surjective f) : card α < card β := @@ -755,7 +756,7 @@ theorem Fintype.card_subtype_le [Fintype α] (p : α → Prop) [DecidablePred p] theorem Fintype.card_subtype_lt [Fintype α] {p : α → Prop} [DecidablePred p] {x : α} (hx : ¬p x) : Fintype.card { x // p x } < Fintype.card α := - Fintype.card_lt_of_injective_of_not_mem (↑) Subtype.coe_injective <| by + Fintype.card_lt_of_injective_of_not_mem (b := x) (↑) Subtype.coe_injective <| by rwa [Subtype.range_coe_subtype] theorem Fintype.card_subtype [Fintype α] (p : α → Prop) [DecidablePred p] : diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index e0b09df15d902..bdd198e51a477 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -64,10 +64,10 @@ theorem card_filter_univ_succ (p : Fin (n + 1) → Prop) [DecidablePred p] : (card_filter_univ_succ' p).trans (by split_ifs <;> simp [add_comm 1]) theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Vector α n) : - (univ.filter fun i => a = v.get i).card = v.toList.count a := by + (univ.filter fun i => v.get i = a).card = v.toList.count a := by induction' v with n x xs hxs · simp · simp_rw [card_filter_univ_succ', Vector.get_cons_zero, Vector.toList_cons, Function.comp, - Vector.get_cons_succ, hxs, List.count_cons, add_comm (ite (a = x) 1 0)] + Vector.get_cons_succ, hxs, List.count_cons, add_comm (ite (x = a) 1 0), beq_iff_eq] end Fin diff --git a/Mathlib/Data/Fintype/Order.lean b/Mathlib/Data/Fintype/Order.lean index 4ebbbf169e14a..de38202522c22 100644 --- a/Mathlib/Data/Fintype/Order.lean +++ b/Mathlib/Data/Fintype/Order.lean @@ -91,11 +91,10 @@ noncomputable abbrev toCompleteLattice [Lattice α] [BoundedOrder α] : Complete sInf_le := fun _ _ ha => Finset.inf_le (Set.mem_toFinset.mpr ha) le_sInf := fun s _ ha => Finset.le_inf fun b hb => ha _ <| Set.mem_toFinset.mp hb --- Porting note: `convert` doesn't work as well as it used to. -- See note [reducible non-instances] /-- A finite bounded distributive lattice is completely distributive. -/ -noncomputable abbrev toCompleteDistribLattice [DistribLattice α] [BoundedOrder α] : - CompleteDistribLattice α where +noncomputable abbrev toCompleteDistribLatticeMinimalAxioms [DistribLattice α] [BoundedOrder α] : + CompleteDistribLattice.MinimalAxioms α where __ := toCompleteLattice α iInf_sup_le_sup_sInf := fun a s => by convert (Finset.inf_sup_distrib_left s.toFinset id a).ge using 1 @@ -108,6 +107,11 @@ noncomputable abbrev toCompleteDistribLattice [DistribLattice α] [BoundedOrder simp_rw [Set.mem_toFinset] rfl +-- See note [reducible non-instances] +/-- A finite bounded distributive lattice is completely distributive. -/ +noncomputable abbrev toCompleteDistribLattice [DistribLattice α] [BoundedOrder α] : + CompleteDistribLattice α := .ofMinimalAxioms (toCompleteDistribLatticeMinimalAxioms _) + -- See note [reducible non-instances] /-- A finite bounded linear order is complete. -/ noncomputable abbrev toCompleteLinearOrder @@ -168,12 +172,13 @@ noncomputable instance Bool.completeAtomicBooleanAlgebra : CompleteAtomicBoolean variable {α : Type*} {r : α → α → Prop} [IsTrans α r] {β γ : Type*} [Nonempty γ] {f : γ → α} - [Finite β] (D : Directed r f) + [Finite β] -theorem Directed.finite_set_le {s : Set γ} (hs : s.Finite) : ∃ z, ∀ i ∈ s, r (f i) (f z) := by +theorem Directed.finite_set_le (D : Directed r f) {s : Set γ} (hs : s.Finite) : + ∃ z, ∀ i ∈ s, r (f i) (f z) := by convert D.finset_le hs.toFinset; rw [Set.Finite.mem_toFinset] -theorem Directed.finite_le (g : β → γ) : ∃ z, ∀ i, r (f (g i)) (f z) := by +theorem Directed.finite_le (D : Directed r f) (g : β → γ) : ∃ z, ∀ i, r (f (g i)) (f z) := by classical obtain ⟨z, hz⟩ := D.finite_set_le (Set.finite_range g) exact ⟨z, fun i => hz (g i) ⟨i, rfl⟩⟩ diff --git a/Mathlib/Data/Fintype/Quotient.lean b/Mathlib/Data/Fintype/Quotient.lean index 80a90364a5535..83851ba7b8a46 100644 --- a/Mathlib/Data/Fintype/Quotient.lean +++ b/Mathlib/Data/Fintype/Quotient.lean @@ -32,10 +32,10 @@ def Quotient.finChoiceAux {ι : Type*} [DecidableEq ι] {α : ι → Type*} [S : · exact fun a l => ⟦fun j h => if e : j = i then by rw [e]; exact a else l _ ((List.mem_cons.1 h).resolve_left e)⟧ refine fun a₁ l₁ a₂ l₂ h₁ h₂ => Quotient.sound fun j h => ?_ - by_cases e : j = i <;> simp [e] - · subst j + by_cases e : j = i + · simp [e]; subst j exact h₁ - · exact h₂ _ _ + · simpa [e] using h₂ _ _ theorem Quotient.finChoiceAux_eq {ι : Type*} [DecidableEq ι] {α : ι → Type*} [S : ∀ i, Setoid (α i)] : diff --git a/Mathlib/Data/Fintype/Shrink.lean b/Mathlib/Data/Fintype/Shrink.lean index 87f4df7f842f7..208fa66907a5a 100644 --- a/Mathlib/Data/Fintype/Shrink.lean +++ b/Mathlib/Data/Fintype/Shrink.lean @@ -15,5 +15,8 @@ variable {α : Type u} [Fintype α] noncomputable instance Shrink.instFintype : Fintype (Shrink.{v} α) := .ofEquiv _ (equivShrink _) +instance Shrink.instFinite {α : Type u} [Finite α] : Finite (Shrink.{v} α) := + .of_equiv _ (equivShrink _) + @[simp] lemma Fintype.card_shrink [Fintype (Shrink.{v} α)] : card (Shrink.{v} α) = card α := card_congr (equivShrink _).symm diff --git a/Mathlib/Data/Int/AbsoluteValue.lean b/Mathlib/Data/Int/AbsoluteValue.lean index 817aaf32927a5..51c7fca75dade 100644 --- a/Mathlib/Data/Int/AbsoluteValue.lean +++ b/Mathlib/Data/Int/AbsoluteValue.lean @@ -3,10 +3,10 @@ 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.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Order.AbsoluteValue import Mathlib.Data.Int.Cast.Lemmas -import Mathlib.GroupTheory.GroupAction.Units /-! # Absolute values and the integers diff --git a/Mathlib/Data/Int/GCD.lean b/Mathlib/Data/Int/GCD.lean index e7ad1d7d1f3d8..108bf803cbc6f 100644 --- a/Mathlib/Data/Int/GCD.lean +++ b/Mathlib/Data/Int/GCD.lean @@ -3,11 +3,11 @@ Copyright (c) 2018 Guy Leroy. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sangwoo Jo (aka Jason), Guy Leroy, Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Int import Mathlib.Algebra.GroupWithZero.Semiconj -import Mathlib.Data.Nat.GCD.Basic import Mathlib.Order.Bounds.Basic +import Mathlib.Algebra.Group.Commute.Units +import Mathlib.Data.Nat.GCD.Basic /-! # Extended GCD and divisibility over ℤ @@ -221,7 +221,7 @@ theorem gcd_eq_zero_iff {i j : ℤ} : gcd i j = 0 ↔ i = 0 ∧ j = 0 := by rw [gcd, Nat.gcd_eq_zero_iff, natAbs_eq_zero, natAbs_eq_zero] theorem gcd_pos_iff {i j : ℤ} : 0 < gcd i j ↔ i ≠ 0 ∨ j ≠ 0 := - pos_iff_ne_zero.trans <| gcd_eq_zero_iff.not.trans not_and_or + Nat.pos_iff_ne_zero.trans <| gcd_eq_zero_iff.not.trans not_and_or theorem gcd_div {i j k : ℤ} (H1 : k ∣ i) (H2 : k ∣ j) : gcd (i / k) (j / k) = gcd i j / natAbs k := by diff --git a/Mathlib/Data/LazyList/Basic.lean b/Mathlib/Data/LazyList/Basic.lean index 3c6b2e8783757..ad67f3d3f60aa 100644 --- a/Mathlib/Data/LazyList/Basic.lean +++ b/Mathlib/Data/LazyList/Basic.lean @@ -11,9 +11,7 @@ import Mathlib.Lean.Thunk /-! ## Definitions on lazy lists -This file contains various definitions and proofs on lazy lists. - -TODO: This file will soon be deprecated. +This file is entirely deprecated, and contains various definitions and proofs on lazy lists. -/ -- The whole file is full of deprecations about LazyList @@ -26,6 +24,7 @@ namespace LazyList open Function /-- Isomorphism between strict and lazy lists. -/ +@[deprecated (since := "2024-07-22")] def listEquivLazyList (α : Type*) : List α ≃ LazyList α where toFun := LazyList.ofList invFun := LazyList.toList @@ -40,10 +39,12 @@ def listEquivLazyList (α : Type*) : List α ≃ LazyList α where · simp [toList, ofList] · simpa [ofList, toList] +@[deprecated (since := "2024-07-22")] instance : Traversable LazyList where map := @LazyList.traverse Id _ traverse := @LazyList.traverse +@[deprecated (since := "2024-07-22")] instance : LawfulTraversable LazyList := by apply Equiv.isLawfulTraversable' listEquivLazyList <;> intros <;> ext <;> rename_i f xs · induction' xs using LazyList.rec with _ _ _ _ ih @@ -66,7 +67,8 @@ instance : LawfulTraversable LazyList := by Function.comp, Thunk.pure, ofList] · apply ih -@[simp] theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by +@[deprecated (since := "2024-07-22"), simp] +theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by induction x using LazyList.rec (motive_2 := fun xs => xs.get.bind singleton = xs.get) with | nil => simp [LazyList.bind] | cons h t ih => @@ -75,6 +77,7 @@ instance : LawfulTraversable LazyList := by simp [ih] | mk f ih => simp_all +@[deprecated (since := "2024-07-22")] instance : LawfulMonad LazyList := LawfulMonad.mk' (id_map := by intro α xs @@ -99,8 +102,7 @@ instance : LawfulMonad LazyList := LawfulMonad.mk' | case1 => simp only [Thunk.pure, LazyList.bind, LazyList.traverse, Id.pure_eq] | case2 _ _ ih => - simp [LazyList.bind, LazyList.traverse, Seq.seq, Id.map_eq, append, Thunk.pure] - rw [← ih] + simp only [Thunk.pure, LazyList.bind, append, Thunk.get_mk, comp_apply, ← ih] simp only [Thunk.get, append, singleton, Thunk.pure]) end LazyList diff --git a/Mathlib/Data/List/AList.lean b/Mathlib/Data/List/AList.lean index b53d7963a81e5..ffb84b46157e2 100644 --- a/Mathlib/Data/List/AList.lean +++ b/Mathlib/Data/List/AList.lean @@ -441,7 +441,7 @@ theorem union_comm_of_disjoint {s₁ s₂ : AList β} (h : Disjoint s₁ s₂) : (s₁ ∪ s₂).entries ~ (s₂ ∪ s₁).entries := lookup_ext (AList.nodupKeys _) (AList.nodupKeys _) (by - intros; simp + intros; simp only [union_entries, Option.mem_def, dlookup_kunion_eq_some] constructor <;> intro h' · cases' h' with h' h' · right diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index b2fda9b22c586..d655d005f411e 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -202,7 +202,7 @@ alias ⟨eq_nil_of_subset_nil, _⟩ := subset_nil theorem map_subset_iff {l₁ l₂ : List α} (f : α → β) (h : Injective f) : map f l₁ ⊆ map f l₂ ↔ l₁ ⊆ l₂ := by - refine ⟨?_, map_subset f⟩; intro h2 x hx + refine ⟨?_, Subset.map f⟩; intro h2 x hx rcases mem_map.1 (h2 (mem_map_of_mem f hx)) with ⟨x', hx', hxx'⟩ cases h hxx'; exact hx' @@ -324,9 +324,6 @@ theorem reverse_bijective : Bijective (@reverse α) := theorem reverse_inj {l₁ l₂ : List α} : reverse l₁ = reverse l₂ ↔ l₁ = l₂ := reverse_injective.eq_iff -theorem reverse_eq_iff {l l' : List α} : l.reverse = l' ↔ l = l'.reverse := - reverse_involutive.eq_iff - theorem concat_eq_reverse_cons (a : α) (l : List α) : concat l a = reverse (a :: reverse l) := by simp only [concat_eq_append, reverse_cons, reverse_reverse] @@ -345,20 +342,15 @@ theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <; /-! ### getLast -/ -@[simp] -theorem getLast_cons {a : α} {l : List α} : - ∀ h : l ≠ nil, getLast (a :: l) (cons_ne_nil a l) = getLast l h := by - induction l <;> intros - · contradiction - · rfl +attribute [simp] getLast_cons theorem getLast_append_singleton {a : α} (l : List α) : - getLast (l ++ [a]) (append_ne_nil_of_ne_nil_right l (cons_ne_nil a _)) = a := by + getLast (l ++ [a]) (append_ne_nil_of_right_ne_nil l (cons_ne_nil a _)) = a := by simp [getLast_append] -- Porting note: name should be fixed upstream theorem getLast_append' (l₁ l₂ : List α) (h : l₂ ≠ []) : - getLast (l₁ ++ l₂) (append_ne_nil_of_ne_nil_right l₁ h) = getLast l₂ h := by + getLast (l₁ ++ l₂) (append_ne_nil_of_right_ne_nil l₁ h) = getLast l₂ h := by induction' l₁ with _ _ ih · simp · simp only [cons_append] @@ -563,10 +555,6 @@ theorem tail_append_of_ne_nil (l l' : List α) (h : l ≠ []) : (l ++ l').tail = · contradiction · simp -theorem getElem_eq_getElem? (l : List α) (i : Nat) (h : i < l.length) : - l[i] = l[i]?.get (by simp [getElem?_eq_getElem, h]) := by - simp [getElem_eq_iff] - theorem get_eq_get? (l : List α) (i : Fin l.length) : l.get i = (l.get? i).get (by simp [getElem?_eq_getElem]) := by simp [getElem_eq_iff] @@ -732,7 +720,7 @@ theorem sublist_cons_of_sublist (a : α) (h : l₁ <+ l₂) : l₁ <+ a :: l₂ theorem tail_sublist : ∀ l : List α, tail l <+ l | [] => .slnil - | a::l => sublist_cons a l + | a::l => sublist_cons_self a l @[gcongr] protected theorem Sublist.tail : ∀ {l₁ l₂ : List α}, l₁ <+ l₂ → tail l₁ <+ tail l₂ | _, _, slnil => .slnil @@ -1094,8 +1082,6 @@ theorem get_set_of_ne {l : List α} {i j : ℕ} (h : i ≠ j) (a : α) @[deprecated (since := "2024-06-21")] alias map_congr := map_congr_left -@[deprecated (since := "2024-06-21")] alias map_eq_map_iff := map_inj_left - theorem bind_pure_eq_map (f : α → β) (l : List α) : l.bind (pure ∘ f) = map f l := .symm <| map_eq_bind .. @@ -1322,18 +1308,6 @@ theorem foldl_fixed {a : α} : ∀ l : List β, foldl (fun a _ => a) a l = a := theorem foldr_fixed {b : β} : ∀ l : List α, foldr (fun _ b => b) b l = b := foldr_fixed' fun _ => rfl -@[simp] -theorem foldl_join (f : α → β → α) : - ∀ (a : α) (L : List (List β)), foldl f a (join L) = foldl (foldl f) a L - | a, [] => rfl - | a, l :: L => by simp only [join, foldl_append, foldl_cons, foldl_join f (foldl f a l) L] - -@[simp] -theorem foldr_join (f : α → β → β) : - ∀ (b : β) (L : List (List α)), foldr f b (join L) = foldr (fun l b => foldr f b l) b L - | a, [] => rfl - | a, l :: L => by simp only [join, foldr_append, foldr_join f a L, foldr_cons] - -- Porting note (#10618): simp can prove this -- @[simp] theorem foldr_eta : ∀ l : List α, foldr cons [] l = l := by @@ -1657,7 +1631,7 @@ variable (p : α → Bool) (xs ys : List α) (ls : List (List α)) (f : List α theorem splitAt_eq_take_drop (n : ℕ) (l : List α) : splitAt n l = (take n l, drop n l) := by by_cases h : n < l.length <;> rw [splitAt, go_eq_take_drop] · rw [if_pos h]; rfl - · rw [if_neg h, take_all_of_le <| le_of_not_lt h, drop_eq_nil_of_le <| le_of_not_lt h] + · rw [if_neg h, take_of_length_le <| le_of_not_lt h, drop_eq_nil_of_le <| le_of_not_lt h] where go_eq_take_drop (n : ℕ) (l xs : List α) (acc : Array α) : splitAt.go l xs n acc = if n < xs.length then (acc.toList ++ take n xs, drop n xs) else (l, []) := by @@ -1732,7 +1706,7 @@ theorem splitOnP_spec (as : List α) : · rw [if_pos h, h, map, cons_append, zipWith, nil_append, join, cons_append, cons_inj_right] exact ih · rw [if_neg h, eq_false_of_ne_true h, join_zipWith (splitOnP_ne_nil _ _) - (append_ne_nil_of_ne_nil_right _ (cons_ne_nil [] [])), cons_inj_right] + (append_ne_nil_of_right_ne_nil _ (cons_ne_nil [] [])), cons_inj_right] exact ih where join_zipWith {xs ys : List (List α)} {a : α} (hxs : xs ≠ []) (hys : ys ≠ []) : @@ -1814,7 +1788,7 @@ theorem modifyLast.go_append_one (f : α → α) (a : α) (tl : List α) (r : Ar | cons hd tl => simp only [cons_append] rw [modifyLast.go, modifyLast.go] - case x_3 | x_3 => exact append_ne_nil_of_ne_nil_right tl (cons_ne_nil a []) + case x_3 | x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) rw [modifyLast.go_append_one _ _ tl _, modifyLast.go_append_one _ _ tl (Array.push #[] hd)] simp only [Array.toListAppend_eq, Array.push_data, Array.data_toArray, nil_append, append_assoc] @@ -1826,7 +1800,7 @@ theorem modifyLast_append_one (f : α → α) (a : α) (l : List α) : | cons _ tl => simp only [cons_append, modifyLast] rw [modifyLast.go] - case x_3 => exact append_ne_nil_of_ne_nil_right tl (cons_ne_nil a []) + case x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) rw [modifyLast.go_append_one, Array.toListAppend_eq, Array.push_data, Array.data_toArray, nil_append, cons_append, nil_append, cons_inj_right] exact modifyLast_append_one _ _ tl @@ -1848,8 +1822,6 @@ end ModifyLast /-! ### map for partial functions -/ -@[simp] lemma attach_nil : ([] : List α).attach = [] := rfl - theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l) : SizeOf.sizeOf x < SizeOf.sizeOf l := by induction' l with h t ih <;> cases hx <;> rw [cons.sizeOf_spec] @@ -1857,121 +1829,8 @@ theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l · specialize ih ‹_› omega -@[simp] -theorem pmap_eq_map (p : α → Prop) (f : α → β) (l : List α) (H) : - @pmap _ _ p (fun a _ => f a) l H = map f l := by - induction l <;> [rfl; simp only [*, pmap, map]] - -theorem pmap_congr {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a, q a → β} (l : List α) {H₁ H₂} - (h : ∀ a ∈ l, ∀ (h₁ h₂), f a h₁ = g a h₂) : pmap f l H₁ = pmap g l H₂ := by - induction' l with _ _ ih - · rfl - · rw [pmap, pmap, h _ (mem_cons_self _ _), ih fun a ha => h a (mem_cons_of_mem _ ha)] - -theorem map_pmap {p : α → Prop} (g : β → γ) (f : ∀ a, p a → β) (l H) : - map g (pmap f l H) = pmap (fun a h => g (f a h)) l H := by - induction l <;> [rfl; simp only [*, pmap, map]] - -theorem pmap_map {p : β → Prop} (g : ∀ b, p b → γ) (f : α → β) (l H) : - pmap g (map f l) H = pmap (fun a h => g (f a) h) l fun a h => H _ (mem_map_of_mem _ h) := by - induction l <;> [rfl; simp only [*, pmap, map]] - -theorem pmap_eq_map_attach {p : α → Prop} (f : ∀ a, p a → β) (l H) : - pmap f l H = l.attach.map fun x => f x.1 (H _ x.2) := by - rw [attach, attachWith, map_pmap]; exact pmap_congr l fun _ _ _ _ => rfl - --- @[simp] -- Porting note (#10959): lean 4 simp can't rewrite with this -theorem attach_map_coe' (l : List α) (f : α → β) : - (l.attach.map fun (i : {i // i ∈ l}) => f i) = l.map f := by - rw [attach, attachWith, map_pmap]; exact pmap_eq_map _ _ _ _ - -theorem attach_map_val' (l : List α) (f : α → β) : (l.attach.map fun i => f i.val) = l.map f := - attach_map_coe' _ _ - -@[simp] -theorem attach_map_val (l : List α) : l.attach.map Subtype.val = l := - (attach_map_coe' _ _).trans l.map_id --- Porting note: coe is expanded eagerly, so "attach_map_coe" would have the same syntactic form. - -@[simp] -theorem mem_attach (l : List α) : ∀ x, x ∈ l.attach - | ⟨a, h⟩ => by - have := mem_map.1 (by rw [attach_map_val] <;> exact h) - rcases this with ⟨⟨_, _⟩, m, rfl⟩ - exact m - -@[simp] -theorem mem_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H b} : - b ∈ pmap f l H ↔ ∃ (a : _) (h : a ∈ l), f a (H a h) = b := by - simp only [pmap_eq_map_attach, mem_map, mem_attach, true_and_iff, Subtype.exists, eq_comm] - -@[simp] -theorem length_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H} : length (pmap f l H) = length l := by - induction l <;> [rfl; simp only [*, pmap, length]] - -@[simp] -theorem length_attach (L : List α) : L.attach.length = L.length := - length_pmap - -@[simp] -theorem pmap_eq_nil {p : α → Prop} {f : ∀ a, p a → β} {l H} : pmap f l H = [] ↔ l = [] := by - rw [← length_eq_zero, length_pmap, length_eq_zero] - -@[simp] -theorem attach_eq_nil (l : List α) : l.attach = [] ↔ l = [] := - pmap_eq_nil - -theorem getLast_pmap (p : α → Prop) (f : ∀ a, p a → β) (l : List α) - (hl₁ : ∀ a ∈ l, p a) (hl₂ : l ≠ []) : - (l.pmap f hl₁).getLast (mt List.pmap_eq_nil.1 hl₂) = - f (l.getLast hl₂) (hl₁ _ (List.getLast_mem hl₂)) := by - induction' l with l_hd l_tl l_ih - · apply (hl₂ rfl).elim - · by_cases hl_tl : l_tl = [] - · simp [hl_tl] - · simp only [pmap] - rw [getLast_cons, l_ih _ hl_tl] - simp only [getLast_cons hl_tl] - -theorem getElem?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) (n : ℕ) : - (pmap f l h)[n]? = Option.pmap f l[n]? fun x H => h x (getElem?_mem H) := by - induction' l with hd tl hl generalizing n - · simp - · cases' n with n - · simp only [Option.pmap] - split <;> simp_all - · simp only [hl, pmap, Option.pmap, getElem?_cons_succ] - split <;> rename_i h₁ _ <;> split <;> rename_i h₂ _ - · simp_all - · simp at h₂ - simp_all - · simp_all - · simp_all - -theorem get?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) (n : ℕ) : - get? (pmap f l h) n = Option.pmap f (get? l n) fun x H => h x (get?_mem H) := by - simp only [get?_eq_getElem?] - simp [getElem?_pmap, h] - -theorem getElem_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : ℕ} - (hn : n < (pmap f l h).length) : - (pmap f l h)[n] = - f (l[n]'(@length_pmap _ _ p f l h ▸ hn)) - (h _ (getElem_mem l n (@length_pmap _ _ p f l h ▸ hn))) := by - induction' l with hd tl hl generalizing n - · simp only [length, pmap] at hn - exact absurd hn (not_lt_of_le n.zero_le) - · cases n - · simp - · simp [hl] - -theorem get_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : ℕ} - (hn : n < (pmap f l h).length) : - get (pmap f l h) ⟨n, hn⟩ = - f (get l ⟨n, @length_pmap _ _ p f l h ▸ hn⟩) - (h _ (get_mem l n (@length_pmap _ _ p f l h ▸ hn))) := by - simp only [get_eq_getElem] - simp [getElem_pmap] +@[deprecated attach_map_coe (since := "2024-07-29")] alias attach_map_coe' := attach_map_coe +@[deprecated attach_map_val (since := "2024-07-29")] alias attach_map_val' := attach_map_val set_option linter.deprecated false in @[deprecated get_pmap (since := "2023-01-05")] @@ -1982,22 +1841,6 @@ theorem nthLe_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : (h _ (get_mem l n (@length_pmap _ _ p f l h ▸ hn))) := get_pmap .. -theorem pmap_append {p : ι → Prop} (f : ∀ a : ι, p a → α) (l₁ l₂ : List ι) - (h : ∀ a ∈ l₁ ++ l₂, p a) : - (l₁ ++ l₂).pmap f h = - (l₁.pmap f fun a ha => h a (mem_append_left l₂ ha)) ++ - l₂.pmap f fun a ha => h a (mem_append_right l₁ ha) := by - induction' l₁ with _ _ ih - · rfl - · dsimp only [pmap, cons_append] - rw [ih] - -theorem pmap_append' {p : α → Prop} (f : ∀ a : α, p a → β) (l₁ l₂ : List α) - (h₁ : ∀ a ∈ l₁, p a) (h₂ : ∀ a ∈ l₂, p a) : - ((l₁ ++ l₂).pmap f fun a ha => (List.mem_append.1 ha).elim (h₁ a) (h₂ a)) = - l₁.pmap f h₁ ++ l₂.pmap f h₂ := - pmap_append f l₁ l₂ _ - /-! ### find -/ section find? @@ -2150,14 +1993,19 @@ theorem filter_eq_foldr (p : α → Bool) (l : List α) : filter p l = foldr (fun a out => bif p a then a :: out else out) [] l := by induction l <;> simp [*, filter]; rfl +#adaptation_note +/-- +This has to be temporarily renamed to avoid an unintentional collision. +The prime should be removed at nightly-2024-07-27. +-/ @[simp] -theorem filter_subset (l : List α) : filter p l ⊆ l := +theorem filter_subset' (l : List α) : filter p l ⊆ l := (filter_sublist l).subset theorem of_mem_filter {a : α} {l} (h : a ∈ filter p l) : p a := (mem_filter.1 h).2 theorem mem_of_mem_filter {a : α} {l} (h : a ∈ filter p l) : a ∈ l := - filter_subset l h + filter_subset' l h theorem mem_filter_of_mem {a : α} {l} (h₁ : a ∈ l) (h₂ : p a) : a ∈ filter p l := mem_filter.2 ⟨h₁, h₂⟩ @@ -2203,7 +2051,7 @@ lemma filter_attach (l : List α) (p : α → Bool) : (l.filter p).attach.map (Subtype.map id fun x => mem_of_mem_filter) := map_injective_iff.2 Subtype.coe_injective <| by simp_rw [map_map, (· ∘ ·), Subtype.map, id, ← Function.comp_apply (g := Subtype.val), - ← filter_map, attach_map_val] + ← filter_map, attach_map_subtype_val] lemma filter_comm (q) (l : List α) : filter p (filter q l) = filter q (filter p l) := by simp [and_comm] @@ -2674,6 +2522,24 @@ theorem dropSlice_eq (xs : List α) (n m : ℕ) : dropSlice n m xs = xs.take n + · cases xs <;> simp [dropSlice] · cases xs <;> simp [dropSlice, *, Nat.succ_add] +@[simp] +theorem length_dropSlice (i j : ℕ) (xs : List α) : + (List.dropSlice i j xs).length = xs.length - min j (xs.length - i) := by + induction xs generalizing i j with + | nil => simp + | cons x xs xs_ih => + cases i <;> simp only [List.dropSlice] + · cases j with + | zero => simp + | succ n => simp_all [xs_ih]; omega + · simp [xs_ih]; omega + +theorem length_dropSlice_lt (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : + (List.dropSlice i j xs).length < xs.length := by + simp; omega + +set_option linter.deprecated false in +@[deprecated (since := "2024-07-25")] theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : SizeOf.sizeOf (List.dropSlice i j xs) < SizeOf.sizeOf xs := by induction xs generalizing i j hj with @@ -2720,8 +2586,7 @@ lemma lookup_graph (f : α → β) {a : α} {as : List α} (h : a ∈ as) : · exact (List.not_mem_nil _ h).elim · by_cases ha : a = a' · simp [ha, lookup_cons] - · simp [lookup_cons, beq_false_of_ne ha] - exact ih (List.mem_of_ne_of_mem ha h) + · simpa [lookup_cons, beq_false_of_ne ha] using ih (List.mem_of_ne_of_mem ha h) end lookup diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 9a1aa895d2b25..1e128c3d42e2b 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -3,7 +3,7 @@ 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.Basic +import Mathlib.Data.Nat.Defs /-! # Counting in lists @@ -23,25 +23,15 @@ variable {α : Type*} {l : List α} namespace List -section CountP - -variable (p q : α → Bool) - -theorem length_filter_lt_length_iff_exists (l) : - length (filter p l) < length l ↔ ∃ x ∈ l, ¬p x := by - simpa [length_eq_countP_add_countP p l, countP_eq_length_filter] using - countP_pos (fun x => ¬p x) (l := l) - --- Porting note: `Lean.Internal.coeM` forces us to type-ascript `{x // x ∈ l}` -lemma countP_attach (l : List α) : l.attach.countP (fun a : {x // x ∈ l} ↦ p a) = l.countP p := by - simp_rw [← Function.comp_apply (g := Subtype.val), ← countP_map, attach_map_val] - -end CountP - /-! ### count -/ section Count +@[simp] +theorem count_map_of_injective {β} [DecidableEq α] [DecidableEq β] (l : List α) (f : α → β) + (hf : Function.Injective f) (x : α) : count (f x) (map f l) = count x l := by + simp only [count, countP_map, (· ∘ ·), hf.beq_eq] + variable [DecidableEq α] @[deprecated (since := "2023-08-23")] @@ -50,15 +40,6 @@ theorem count_cons' (a b : α) (l : List α) : simp only [count, beq_iff_eq, countP_cons, Nat.add_right_inj] simp only [eq_comm] -@[simp] -lemma count_attach (a : {x // x ∈ l}) : l.attach.count a = l.count ↑a := - Eq.trans (countP_congr fun _ _ => by simp [Subtype.ext_iff]) <| countP_attach _ _ - -@[simp] -theorem count_map_of_injective {β} [DecidableEq α] [DecidableEq β] (l : List α) (f : α → β) - (hf : Function.Injective f) (x : α) : count (f x) (map f l) = count x l := by - simp only [count, countP_map, (· ∘ ·), hf.beq_eq] - end Count end List diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 6bb5c64dac925..dcf48a84137c4 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -168,7 +168,7 @@ theorem next_getLast_cons (h : x ∈ l) (y : α) (h : x ∈ y :: l) (hy : x ≠ rw [← get?_eq_get, dropLast_eq_take, get?_eq_getElem?, getElem?_take, getElem?_cons_zero, Option.some_inj] at hk' · exact hy (Eq.symm hk') - rw [length_cons, Nat.pred_succ] + rw [length_cons] exact length_pos_of_mem (by assumption) suffices k + 1 = l.length by simp [this] at hk cases' l with hd tl diff --git a/Mathlib/Data/List/Destutter.lean b/Mathlib/Data/List/Destutter.lean index 714f67f548c69..5f80cabc72631 100644 --- a/Mathlib/Data/List/Destutter.lean +++ b/Mathlib/Data/List/Destutter.lean @@ -60,7 +60,7 @@ theorem destutter'_sublist (a) : l.destutter' R a <+ a :: l := by rw [destutter'] split_ifs · exact Sublist.cons₂ a (hl b) - · exact (hl a).trans ((l.sublist_cons b).cons_cons a) + · exact (hl a).trans ((l.sublist_cons_self b).cons_cons a) theorem mem_destutter' (a) : a ∈ l.destutter' R a := by induction' l with b l hl diff --git a/Mathlib/Data/List/DropRight.lean b/Mathlib/Data/List/DropRight.lean index 625075fafd908..56bd597ad9ef2 100644 --- a/Mathlib/Data/List/DropRight.lean +++ b/Mathlib/Data/List/DropRight.lean @@ -131,7 +131,7 @@ theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p refine ⟨fun h => ?_, fun h => ?_⟩ · intro _ H rw [get] at H - refine (cons_ne_self hd tl) (Sublist.antisymm ?_ (sublist_cons _ _)) + refine (cons_ne_self hd tl) (Sublist.antisymm ?_ (sublist_cons_self _ _)) rw [← h] simp only [H] exact List.IsSuffix.sublist (dropWhile_suffix p) diff --git a/Mathlib/Data/List/EditDistance/Defs.lean b/Mathlib/Data/List/EditDistance/Defs.lean index e429103588fc1..61e5a51bccc9c 100644 --- a/Mathlib/Data/List/EditDistance/Defs.lean +++ b/Mathlib/Data/List/EditDistance/Defs.lean @@ -246,7 +246,7 @@ theorem suffixLevenshtein_cons₁_fst (x : α) (xs ys) : theorem suffixLevenshtein_cons_cons_fst_get_zero (x : α) (xs y ys) (w : 0 < (suffixLevenshtein C (x :: xs) (y :: ys)).val.length) : - (suffixLevenshtein C (x :: xs) (y :: ys)).1[0] = + (suffixLevenshtein C (x :: xs) (y :: ys)).1[0]'w = let ⟨dx, _⟩ := suffixLevenshtein C xs (y :: ys) let ⟨dy, _⟩ := suffixLevenshtein C (x :: xs) ys let ⟨dxy, _⟩ := suffixLevenshtein C xs ys @@ -296,4 +296,4 @@ theorem levenshtein_cons_cons min (C.delete x + levenshtein C xs (y :: ys)) (min (C.insert y + levenshtein C (x :: xs) ys) (C.substitute x y + levenshtein C xs ys)) := - suffixLevenshtein_cons_cons_fst_get_zero _ _ _ _ _ + suffixLevenshtein_cons_cons_fst_get_zero .. diff --git a/Mathlib/Data/List/Enum.lean b/Mathlib/Data/List/Enum.lean index 08ad0a86d2413..e7bd524eb313e 100644 --- a/Mathlib/Data/List/Enum.lean +++ b/Mathlib/Data/List/Enum.lean @@ -3,7 +3,9 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Yakov Pechersky, Eric Wieser -/ -import Mathlib.Data.List.Basic +import Batteries.Tactic.Alias +import Mathlib.Tactic.TypeStar +import Mathlib.Data.Nat.Notation /-! # Properties of `List.enum` @@ -13,53 +15,21 @@ namespace List variable {α β : Type*} -@[simp] -theorem getElem?_enumFrom : - ∀ n (l : List α) m, (enumFrom n l)[m]? = l[m]?.map fun a => (n + m, a) - | n, [], m => rfl - | n, a :: l, 0 => by simp - | n, a :: l, m + 1 => by - simp only [enumFrom_cons, getElem?_cons_succ] - exact (getElem?_enumFrom (n + 1) l m).trans <| by rw [Nat.add_right_comm]; rfl - theorem get?_enumFrom (n) (l : List α) (m) : get? (enumFrom n l) m = (get? l m).map fun a => (n + m, a) := by simp @[deprecated (since := "2024-04-06")] alias enumFrom_get? := get?_enumFrom -@[simp] -theorem getElem?_enum (l : List α) (n : Nat) : (enum l)[n]? = l[n]?.map fun a => (n, a) := by - rw [enum, getElem?_enumFrom, Nat.zero_add] - theorem get?_enum (l : List α) (n) : get? (enum l) n = (get? l n).map fun a => (n, a) := by simp @[deprecated (since := "2024-04-06")] alias enum_get? := get?_enum -@[simp] -theorem enumFrom_map_snd : ∀ (n) (l : List α), map Prod.snd (enumFrom n l) = l - | _, [] => rfl - | _, _ :: _ => congr_arg (cons _) (enumFrom_map_snd _ _) - -@[simp] -theorem enum_map_snd (l : List α) : map Prod.snd (enum l) = l := - enumFrom_map_snd _ _ - -@[simp] -theorem getElem_enumFrom (l : List α) (n) (i : Nat) (h : i < (l.enumFrom n).length) : - (l.enumFrom n)[i] = (n + i, l[i]'(by simpa [enumFrom_length] using h)) := by - simp [getElem_eq_getElem?] - theorem get_enumFrom (l : List α) (n) (i : Fin (l.enumFrom n).length) : (l.enumFrom n).get i = (n + i, l.get (i.cast enumFrom_length)) := by simp -@[simp] -theorem getElem_enum (l : List α) (i : Nat) (h : i < l.enum.length) : - l.enum[i] = (i, l[i]'(by simpa [enum_length] using h)) := by - simp [enum] - theorem get_enum (l : List α) (i : Fin l.enum.length) : l.enum.get i = (i.1, l.get (i.cast enum_length)) := by simp @@ -83,77 +53,4 @@ theorem mk_mem_enum_iff_get? {i : ℕ} {x : α} {l : List α} : (i, x) ∈ enum theorem mem_enum_iff_get? {x : ℕ × α} {l : List α} : x ∈ enum l ↔ l.get? x.1 = x.2 := mk_mem_enum_iff_get? -theorem le_fst_of_mem_enumFrom {x : ℕ × α} {n : ℕ} {l : List α} (h : x ∈ enumFrom n l) : - n ≤ x.1 := - (mk_mem_enumFrom_iff_le_and_get?_sub.1 h).1 - -theorem fst_lt_add_of_mem_enumFrom {x : ℕ × α} {n : ℕ} {l : List α} (h : x ∈ enumFrom n l) : - x.1 < n + length l := by - rcases mem_iff_get.1 h with ⟨i, rfl⟩ - simpa using i.is_lt - -theorem fst_lt_of_mem_enum {x : ℕ × α} {l : List α} (h : x ∈ enum l) : x.1 < length l := by - simpa using fst_lt_add_of_mem_enumFrom h - -theorem snd_mem_of_mem_enumFrom {x : ℕ × α} {n : ℕ} {l : List α} (h : x ∈ enumFrom n l) : x.2 ∈ l := - enumFrom_map_snd n l ▸ mem_map_of_mem _ h - -theorem snd_mem_of_mem_enum {x : ℕ × α} {l : List α} (h : x ∈ enum l) : x.2 ∈ l := - snd_mem_of_mem_enumFrom h - -theorem mem_enumFrom {x : α} {i j : ℕ} (xs : List α) (h : (i, x) ∈ xs.enumFrom j) : - j ≤ i ∧ i < j + xs.length ∧ x ∈ xs := - ⟨le_fst_of_mem_enumFrom h, fst_lt_add_of_mem_enumFrom h, snd_mem_of_mem_enumFrom h⟩ - -@[simp] -theorem enumFrom_singleton (x : α) (n : ℕ) : enumFrom n [x] = [(n, x)] := - rfl - -@[simp] -theorem enum_singleton (x : α) : enum [x] = [(0, x)] := - rfl - -theorem enumFrom_append (xs ys : List α) (n : ℕ) : - enumFrom n (xs ++ ys) = enumFrom n xs ++ enumFrom (n + xs.length) ys := by - induction' xs with x xs IH generalizing ys n - · simp - · rw [cons_append, enumFrom_cons, IH, ← cons_append, ← enumFrom_cons, length, Nat.add_right_comm, - Nat.add_assoc] - -theorem enum_append (xs ys : List α) : enum (xs ++ ys) = enum xs ++ enumFrom xs.length ys := by - simp [enum, enumFrom_append] - -theorem map_fst_add_enumFrom_eq_enumFrom (l : List α) (n k : ℕ) : - map (Prod.map (· + n) id) (enumFrom k l) = enumFrom (n + k) l := - ext_get? fun i ↦ by simp [(· ∘ ·), Nat.add_comm, Nat.add_left_comm] - -theorem map_fst_add_enum_eq_enumFrom (l : List α) (n : ℕ) : - map (Prod.map (· + n) id) (enum l) = enumFrom n l := - map_fst_add_enumFrom_eq_enumFrom l _ _ - -theorem enumFrom_cons' (n : ℕ) (x : α) (xs : List α) : - enumFrom n (x :: xs) = (n, x) :: (enumFrom n xs).map (Prod.map Nat.succ id) := by - rw [enumFrom_cons, Nat.add_comm, ← map_fst_add_enumFrom_eq_enumFrom] - -theorem enum_cons' (x : α) (xs : List α) : - enum (x :: xs) = (0, x) :: (enum xs).map (Prod.map Nat.succ id) := - enumFrom_cons' _ _ _ - -theorem enumFrom_map (n : ℕ) (l : List α) (f : α → β) : - enumFrom n (l.map f) = (enumFrom n l).map (Prod.map id f) := by - induction' l with hd tl IH - · rfl - · rw [map_cons, enumFrom_cons', enumFrom_cons', map_cons, map_map, IH, map_map] - rfl - -theorem enum_map (l : List α) (f : α → β) : (l.map f).enum = l.enum.map (Prod.map id f) := - enumFrom_map _ _ _ - -@[simp] -theorem enumFrom_eq_nil {n : ℕ} {l : List α} : List.enumFrom n l = [] ↔ l = [] := by - cases l <;> simp - -@[simp] -theorem enum_eq_nil {l : List α} : List.enum l = [] ↔ l = [] := enumFrom_eq_nil - end List diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index dcd67a6eaf546..bf666aecb4c27 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -5,6 +5,7 @@ Authors: Jannis Limperg -/ import Mathlib.Data.List.OfFn import Mathlib.Data.List.Range +import Mathlib.Data.List.Zip /-! # Lemmas about List.*Idx functions. diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index f809824bf8436..cefbcfd118662 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Data.List.Lattice -import Mathlib.Data.List.Range import Mathlib.Data.Bool.Basic import Mathlib.Init.Data.Nat.Lemmas +import Mathlib.Order.Lattice /-! # Intervals in ℕ diff --git a/Mathlib/Data/List/Iterate.lean b/Mathlib/Data/List/Iterate.lean index 76fffcce7b531..dc981a78f8e38 100644 --- a/Mathlib/Data/List/Iterate.lean +++ b/Mathlib/Data/List/Iterate.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Miyahara Kō -/ -import Mathlib.Data.List.Range import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Data.List.Defs +import Mathlib.Data.Set.Function /-! # iterate diff --git a/Mathlib/Data/List/Join.lean b/Mathlib/Data/List/Join.lean index e6f20b71c8f26..c8ecd55aad8c5 100644 --- a/Mathlib/Data/List/Join.lean +++ b/Mathlib/Data/List/Join.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn, Mario Carneiro, Martin Dvorak -/ import Mathlib.Data.List.Basic -import Batteries.Data.Nat.Lemmas /-! # Join of a list of lists @@ -42,10 +41,8 @@ theorem join_filter_not_isEmpty : @[simp] theorem join_filter_ne_nil [DecidablePred fun l : List α => l ≠ []] {L : List (List α)} : join (L.filter fun l => l ≠ []) = L.join := by - simp [join_filter_not_isEmpty, ← isEmpty_iff_eq_nil] - -theorem join_join (l : List (List (List α))) : l.join.join = (l.map join).join := by - induction l <;> simp [*] + simp only [ne_eq, ← isEmpty_iff_eq_nil, Bool.not_eq_true, Bool.decide_eq_false, + join_filter_not_isEmpty] /-- See `List.length_join` for the corresponding statement using `List.sum`. -/ lemma length_join' (L : List (List α)) : length (join L) = Nat.sum (map length L) := by diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 34f4a787a74af..181f14ce99536 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -122,7 +122,7 @@ theorem mem_inter_of_mem_of_mem (h₁ : a ∈ l₁) (h₂ : a ∈ l₂) : a ∈ mem_filter_of_mem h₁ <| by simpa using h₂ theorem inter_subset_left {l₁ l₂ : List α} : l₁ ∩ l₂ ⊆ l₁ := - filter_subset _ + filter_subset' _ theorem inter_subset_right {l₁ l₂ : List α} : l₁ ∩ l₂ ⊆ l₂ := fun _ => mem_of_mem_inter_right @@ -200,10 +200,12 @@ theorem count_bagInter {a : α} : by_cases hb : b ∈ l₂ · rw [cons_bagInter_of_pos _ hb, count_cons, count_cons, count_bagInter, count_erase, ← Nat.add_min_add_right] - by_cases ab : a = b - · rw [if_pos ab, Nat.sub_add_cancel] - rwa [succ_le_iff, count_pos_iff_mem, ab] - · rw [if_neg ab, Nat.sub_zero, Nat.add_zero, Nat.add_zero] + by_cases ba : b = a + · simp only [beq_iff_eq] + rw [if_pos ba, Nat.sub_add_cancel] + rwa [succ_le_iff, count_pos_iff_mem, ← ba] + · simp only [beq_iff_eq] + rw [if_neg ba, Nat.sub_zero, Nat.add_zero, Nat.add_zero] · rw [cons_bagInter_of_neg _ hb, count_bagInter] by_cases ab : a = b · rw [← ab] at hb @@ -221,8 +223,9 @@ theorem bagInter_sublist_left : ∀ l₁ l₂ : List α, l₁.bagInter l₂ <+ l theorem bagInter_nil_iff_inter_nil : ∀ l₁ l₂ : List α, l₁.bagInter l₂ = [] ↔ l₁ ∩ l₂ = [] | [], l₂ => by simp | b :: l₁, l₂ => by - by_cases h : b ∈ l₂ <;> simp [h] - exact bagInter_nil_iff_inter_nil l₁ l₂ + by_cases h : b ∈ l₂ + · simp [h] + · simpa [h] using bagInter_nil_iff_inter_nil l₁ l₂ end BagInter diff --git a/Mathlib/Data/List/NatAntidiagonal.lean b/Mathlib/Data/List/NatAntidiagonal.lean index 856fdacc2be03..fbbf06a9c47bb 100644 --- a/Mathlib/Data/List/NatAntidiagonal.lean +++ b/Mathlib/Data/List/NatAntidiagonal.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Data.List.Nodup -import Mathlib.Data.List.Range /-! # Antidiagonals in ℕ × ℕ as lists diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 625740e5c4d20..30fd2bbe130b3 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -226,8 +226,8 @@ theorem nodup_map_iff {f : α → β} {l : List α} (hf : Injective f) : Nodup ( @[simp] theorem nodup_attach {l : List α} : Nodup (attach l) ↔ Nodup l := - ⟨fun h => attach_map_val l ▸ h.map fun _ _ => Subtype.eq, fun h => - Nodup.of_map Subtype.val ((attach_map_val l).symm ▸ h)⟩ + ⟨fun h => attach_map_subtype_val l ▸ h.map fun _ _ => Subtype.eq, fun h => + Nodup.of_map Subtype.val ((attach_map_subtype_val l).symm ▸ h)⟩ alias ⟨Nodup.of_attach, Nodup.attach⟩ := nodup_attach diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index ec3e96ab365c3..a4498de9ea470 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -139,7 +139,7 @@ attribute [simp] nil_subperm theorem subperm_nil : List.Subperm l [] ↔ l = [] := ⟨fun h ↦ length_eq_zero.1 <| Nat.le_zero.1 h.length_le, by rintro rfl; rfl⟩ -lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons _ _⟩ +lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ _⟩ lemma count_eq_count_filter_add [DecidableEq α] (P : α → Prop) [DecidablePred P] (l : List α) (a : α) : @@ -156,8 +156,8 @@ theorem Perm.foldr_eq {f : α → β → β} {l₁ l₂ : List α} (lcomm : Left intro b induction p using Perm.recOnSwap' generalizing b with | nil => rfl - | cons _ _ r => simp; rw [r b] - | swap' _ _ _ r => simp; rw [lcomm, r b] + | cons _ _ r => simp [r b] + | swap' _ _ _ r => simp only [foldr_cons]; rw [lcomm, r b] | trans _ _ r₁ r₂ => exact Eq.trans (r₁ b) (r₂ b) section @@ -217,8 +217,6 @@ section variable [DecidableEq α] --- attribute [congr] - theorem Perm.bagInter_right {l₁ l₂ : List α} (t : List α) (h : l₁ ~ l₂) : l₁.bagInter t ~ l₂.bagInter t := by induction' h with x _ _ _ _ x y _ _ _ _ _ _ ih_1 ih_2 generalizing t; · simp @@ -247,7 +245,7 @@ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by - simp (config := { contextual := true }) [count_replicate, h, h.symm, this, count_eq_zero] + simp (config := { contextual := true }) [count_replicate, h, this, count_eq_zero, Ne.symm] trans ∀ c, c ∈ l → c = b ∨ c = a · simp [subset_def, or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] @@ -255,13 +253,10 @@ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h -- @[congr] theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := perm_iff_count.2 fun a => - if h : a ∈ l₁ then by simp [nodup_dedup, h, p.subset h] else by simp [h, mt p.mem_iff.2 h] - --- attribute [congr] - --- @[congr] - --- @[congr] + if h : a ∈ l₁ then by + simp [nodup_dedup, h, p.subset h] + else by + simp [mem_dedup, h, not_false_eq_true, count_eq_zero_of_not_mem, mt p.mem_iff.2 h] theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : l ∩ (t₁ ++ t₂) ~ l ∩ t₁ ++ l ∩ t₂ := by @@ -291,7 +286,9 @@ theorem Perm.bind_left (l : List α) {f g : α → List β} (h : ∀ a ∈ l, f theorem bind_append_perm (l : List α) (f g : α → List β) : l.bind f ++ l.bind g ~ l.bind fun x => f x ++ g x := by - induction' l with a l IH <;> simp + induction' l with a l IH + · simp + simp only [bind_cons, append_assoc] refine (Perm.trans ?_ (IH.append_left _)).append_left _ rw [← append_assoc, ← append_assoc] exact perm_append_comm.append_right _ @@ -318,18 +315,13 @@ theorem perm_lookmap (f : α → Option α) {l₁ l₂ : List α} lookmap f l₁ ~ lookmap f l₂ := by induction' p with a l₁ l₂ p IH a b l l₁ l₂ l₃ p₁ _ IH₁ IH₂; · simp · cases h : f a - · simp [h] - exact IH (pairwise_cons.1 H).2 + · simpa [h] using IH (pairwise_cons.1 H).2 · simp [lookmap_cons_some _ _ h, p] · cases' h₁ : f a with c <;> cases' h₂ : f b with d - · simp [h₁, h₂] - apply swap - · simp [h₁, lookmap_cons_some _ _ h₂] - apply swap - · simp [lookmap_cons_some _ _ h₁, h₂] - apply swap - · simp [lookmap_cons_some _ _ h₁, lookmap_cons_some _ _ h₂] - rcases (pairwise_cons.1 H).1 _ (mem_cons.2 (Or.inl rfl)) _ h₂ _ h₁ with ⟨rfl, rfl⟩ + · simpa [h₁, h₂] using swap _ _ _ + · simpa [h₁, lookmap_cons_some _ _ h₂] using swap _ _ _ + · simpa [lookmap_cons_some _ _ h₁, h₂] using swap _ _ _ + · rcases (pairwise_cons.1 H).1 _ (mem_cons.2 (Or.inl rfl)) _ h₂ _ h₁ with ⟨rfl, rfl⟩ exact Perm.refl _ · refine (IH₁ H).trans (IH₂ ((p₁.pairwise_iff ?_).1 H)) intro x y h c hc d hd @@ -350,7 +342,7 @@ theorem Perm.drop_inter [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ y have h₀ : n = xs.length - n' := by rwa [Nat.sub_sub_self] have h₁ : n' ≤ xs.length := Nat.sub_le .. have h₂ : xs.drop n = (xs.reverse.take n').reverse := by - rw [take_reverse _ h₁, h₀, reverse_reverse] + rw [take_reverse h₁, h₀, reverse_reverse] rw [h₂] apply (reverse_perm _).trans rw [inter_reverse] @@ -537,7 +529,7 @@ theorem count_permutations'Aux_self [DecidableEq α] (l : List α) (x : α) : count (x :: l) (permutations'Aux x l) = length (takeWhile (x = ·) l) + 1 := by induction' l with y l IH generalizing x · simp [takeWhile, count] - · rw [permutations'Aux, DecEq_eq, count_cons_self] + · rw [permutations'Aux, count_cons_self] by_cases hx : x = y · subst hx simpa [takeWhile, Nat.succ_inj', DecEq_eq] using IH _ diff --git a/Mathlib/Data/List/Pi.lean b/Mathlib/Data/List/Pi.lean new file mode 100644 index 0000000000000..25244772bc6f4 --- /dev/null +++ b/Mathlib/Data/List/Pi.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2023 Yuyang Zhao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuyang Zhao +-/ +import Mathlib.Data.Multiset.Pi + +/-! +# The cartesian product of lists + +## Main definitions + +* `List.pi`: Cartesian product of lists indexed by a list. +-/ + +namespace List + +namespace Pi +variable {ι : Type*} [DecidableEq ι] {α : ι → Sort*} + +/-- Given `α : ι → Sort*`, `Pi.nil α` is the trivial dependent function out of the empty list. -/ +def nil (α : ι → Sort*) : (∀ i ∈ ([] : List ι), α i) := + nofun + +variable {i : ι} {l : List ι} + +/-- Given `f` a function whose domain is `i :: l`, get its value at `i`. -/ +def head (f : ∀ j ∈ i :: l, α j) : α i := + f i (mem_cons_self _ _) + +/-- Given `f` a function whose domain is `i :: l`, produce a function whose domain +is restricted to `l`. -/ +def tail (f : ∀ j ∈ i :: l, α j) : ∀ j ∈ l, α j := + fun j hj ↦ f j (mem_cons_of_mem _ hj) + +variable (i l) + +/-- Given `α : ι → Sort*`, a list `l` and a term `i`, as well as a term `a : α i` and a +function `f` such that `f j : α j` for all `j` in `l`, `Pi.cons a f` is a function `g` such +that `g k : α k` for all `k` in `i :: l`. -/ +def cons (a : α i) (f : ∀ j ∈ l, α j) : ∀ j ∈ i :: l, α j := + Multiset.Pi.cons (α := ι) l _ a f + +variable {i l} + +lemma cons_def (a : α i) (f : ∀ j ∈ l, α j) : cons _ _ a f = + fun j hj ↦ if h : j = i then h.symm.rec a else f j <| (mem_cons.1 hj).resolve_left h := + rfl + +@[simp] lemma _root_.Multiset.Pi.cons_coe {l : List ι} (a : α i) (f : ∀ j ∈ l, α j) : + Multiset.Pi.cons l _ a f = cons _ _ a f := + rfl + +@[simp] lemma cons_eta (f : ∀ j ∈ i :: l, α j) : + cons _ _ (head f) (tail f) = f := + Multiset.Pi.cons_eta (α := ι) (m := l) f + +lemma cons_map (a : α i) (f : ∀ j ∈ l, α j) + {α' : ι → Sort*} (φ : ∀ ⦃j⦄, α j → α' j) : + cons _ _ (φ a) (fun j hj ↦ φ (f j hj)) = (fun j hj ↦ φ ((cons _ _ a f) j hj)) := + Multiset.Pi.cons_map _ _ _ + +lemma forall_rel_cons_ext {r : ∀ ⦃i⦄, α i → α i → Prop} {a₁ a₂ : α i} {f₁ f₂ : ∀ j ∈ l, α j} + (ha : r a₁ a₂) (hf : ∀ (i : ι) (hi : i ∈ l), r (f₁ i hi) (f₂ i hi)) : + ∀ j hj, r (cons _ _ a₁ f₁ j hj) (cons _ _ a₂ f₂ j hj) := + Multiset.Pi.forall_rel_cons_ext (α := ι) (m := l) ha hf + +end Pi + +variable {ι : Type*} [DecidableEq ι] {α : ι → Type*} + +/-- `pi xs f` creates the list of functions `g` such that, for `x ∈ xs`, `g x ∈ f x` -/ +def pi : ∀ l : List ι, (∀ i, List (α i)) → List (∀ i, i ∈ l → α i) + | [], _ => [List.Pi.nil α] + | i :: l, fs => (fs i).bind (fun b ↦ (pi l fs).map (List.Pi.cons _ _ b)) + +@[simp] lemma pi_nil (t : ∀ i, List (α i)) : + pi [] t = [Pi.nil α] := + rfl + +@[simp] lemma pi_cons (i : ι) (l : List ι) (t : ∀ j, List (α j)) : + pi (i :: l) t = ((t i).bind fun b ↦ (pi l t).map <| Pi.cons _ _ b) := + rfl + +lemma _root_.Multiset.pi_coe (l : List ι) (fs : ∀ i, List (α i)) : + (l : Multiset ι).pi (fs ·) = (↑(pi l fs) : Multiset (∀ i ∈ l, α i)) := by + induction' l with i l ih + · change Multiset.pi 0 _ = _ + simp only [Multiset.coe_singleton, Multiset.pi_zero, pi_nil, Multiset.singleton_inj] + ext i hi + simp at hi + · change Multiset.pi (i ::ₘ ↑l) _ = _ + simp [ih, Multiset.coe_bind, - Multiset.cons_coe] + +lemma mem_pi {l : List ι} (fs : ∀ i, List (α i)) : + ∀ f : ∀ i ∈ l, α i, (f ∈ pi l fs) ↔ (∀ i (hi : i ∈ l), f i hi ∈ fs i) := by + intros f + convert @Multiset.mem_pi ι _ α ↑l (fs ·) f using 1 + rw [Multiset.pi_coe] + rfl + +end List diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 833ec5889075a..64aec7b5b00f9 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Scott Morrison -/ import Mathlib.Data.List.Chain -import Mathlib.Data.List.Enum import Mathlib.Data.List.Nodup import Mathlib.Data.List.Pairwise -import Mathlib.Data.List.Zip import Batteries.Data.Nat.Lemmas /-! @@ -29,16 +27,6 @@ namespace List variable {α : Type u} -@[simp] theorem range'_one {s step : ℕ} : range' s 1 step = [s] := rfl - -theorem pairwise_lt_range' : ∀ s n (step := 1) (_ : 0 < step := by simp), - Pairwise (· < ·) (range' s n step) - | _, 0, _, _ => Pairwise.nil - | s, n + 1, _, h => chain_iff_pairwise.1 (chain_lt_range' s n h) - -theorem nodup_range' (s n : ℕ) (step := 1) (h : 0 < step := by simp) : Nodup (range' s n step) := - (pairwise_lt_range' s n step h).imp _root_.ne_of_lt - set_option linter.deprecated false in @[simp] theorem nthLe_range' {n m step} (i) (H : i < (range' n m step).length) : @@ -48,20 +36,6 @@ set_option linter.deprecated false in theorem nthLe_range'_1 {n m} (i) (H : i < (range' n m).length) : nthLe (range' n m) i H = n + i := by simp -theorem pairwise_lt_range (n : ℕ) : Pairwise (· < ·) (range n) := by - simp (config := {decide := true}) only [range_eq_range', pairwise_lt_range'] - -theorem pairwise_le_range (n : ℕ) : Pairwise (· ≤ ·) (range n) := - Pairwise.imp (@le_of_lt ℕ _) (pairwise_lt_range _) - -theorem take_range (m n : ℕ) : take m (range n) = range (min m n) := by - apply List.ext_getElem - · simp - · simp (config := { contextual := true }) [← getElem_take, Nat.lt_min] - -theorem nodup_range (n : ℕ) : Nodup (range n) := by - simp (config := {decide := true}) only [range_eq_range', nodup_range'] - theorem chain'_range_succ (r : ℕ → ℕ → Prop) (n : ℕ) : Chain' r (range n.succ) ↔ ∀ m < n, r m m.succ := by rw [range_succ] @@ -77,12 +51,6 @@ theorem chain_range_succ (r : ℕ → ℕ → Prop) (n a : ℕ) : rw [range_succ_eq_map, chain_cons, and_congr_right_iff, ← chain'_range_succ, range_succ_eq_map] exact fun _ => Iff.rfl -theorem pairwise_gt_iota (n : ℕ) : Pairwise (· > ·) (iota n) := by - simpa only [iota_eq_reverse_range', pairwise_reverse] using pairwise_lt_range' 1 n - -theorem nodup_iota (n : ℕ) : Nodup (iota n) := - (pairwise_gt_iota n).imp _root_.ne_of_gt - /-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ def finRange (n : ℕ) : List (Fin n) := (range n).pmap Fin.mk fun _ => List.mem_range.1 @@ -115,21 +83,6 @@ theorem pairwise_lt_finRange (n : ℕ) : Pairwise (· < ·) (finRange n) := theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := (List.pairwise_le_range n).pmap (by simp) (by simp) -theorem enum_eq_zip_range (l : List α) : l.enum = (range l.length).zip l := - zip_of_prod (enum_map_fst _) (enum_map_snd _) - -@[simp] -theorem unzip_enum_eq_prod (l : List α) : l.enum.unzip = (range l.length, l) := by - simp only [enum_eq_zip_range, unzip_zip, length_range] - -theorem enumFrom_eq_zip_range' (l : List α) {n : ℕ} : l.enumFrom n = (range' n l.length).zip l := - zip_of_prod (enumFrom_map_fst _ _) (enumFrom_map_snd _ _) - -@[simp] -theorem unzip_enumFrom_eq_prod (l : List α) {n : ℕ} : - (l.enumFrom n).unzip = (range' n l.length, l) := by - simp only [enumFrom_eq_zip_range', unzip_zip, length_range'] - set_option linter.deprecated false in @[simp] theorem nthLe_range {n} (i) (H : i < (range n).length) : nthLe (range n) i H = i := diff --git a/Mathlib/Data/List/ReduceOption.lean b/Mathlib/Data/List/ReduceOption.lean index ab51077e14104..f86d28a812369 100644 --- a/Mathlib/Data/List/ReduceOption.lean +++ b/Mathlib/Data/List/ReduceOption.lean @@ -63,7 +63,7 @@ theorem reduceOption_length_lt_iff {l : List (Option α)} : l.reduceOption.length < l.length ↔ none ∈ l := by rw [Nat.lt_iff_le_and_ne, and_iff_right (reduceOption_length_le l), Ne, reduceOption_length_eq_iff] - induction l <;> simp [*] + induction l <;> simp rw [@eq_comm _ none, ← Option.not_isSome_iff_eq_none, Decidable.imp_iff_not_or] theorem reduceOption_singleton (x : Option α) : [x].reduceOption = x.toList := by cases x <;> rfl diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index f2baca0bc0d74..854cb1f9533fe 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yakov Pechersky -/ import Mathlib.Data.List.Nodup -import Mathlib.Data.List.Zip import Mathlib.Data.Nat.Defs import Mathlib.Data.List.Infix @@ -181,8 +180,8 @@ theorem zipWith_rotate_distrib {β γ : Type*} (f : α → β → γ) (l : List (h : l.length = l'.length) : (zipWith f l l').rotate n = zipWith f (l.rotate n) (l'.rotate n) := by rw [rotate_eq_drop_append_take_mod, rotate_eq_drop_append_take_mod, - rotate_eq_drop_append_take_mod, h, zipWith_append, ← zipWith_distrib_drop, ← - zipWith_distrib_take, List.length_zipWith, h, min_self] + rotate_eq_drop_append_take_mod, h, zipWith_append, ← drop_zipWith, ← + take_zipWith, List.length_zipWith, h, min_self] rw [length_drop, length_drop, h] attribute [local simp] rotate_cons_succ diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index ab70fae9f0d4b..fd7522f0feb23 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Sean Leather -/ -import Mathlib.Data.List.Range import Mathlib.Data.List.Perm +import Mathlib.Data.List.Pairwise /-! # Utilities for lists of sigmas diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 2ad4282d7db08..46474c94663ad 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -246,8 +246,9 @@ theorem perm_orderedInsert (a) : ∀ l : List α, orderedInsert r a l ~ a :: l · simpa [orderedInsert, h] using ((perm_orderedInsert a l).cons _).trans (Perm.swap _ _ _) theorem orderedInsert_count [DecidableEq α] (L : List α) (a b : α) : - count a (L.orderedInsert r b) = count a L + if a = b then 1 else 0 := by + count a (L.orderedInsert r b) = count a L + if b = a then 1 else 0 := by rw [(L.perm_orderedInsert r b).count_eq, count_cons] + simp theorem perm_insertionSort : ∀ l : List α, insertionSort r l ~ l | [] => Perm.nil diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index f933fca99a9b2..7df5ef9099e10 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -119,7 +119,7 @@ theorem sublistsAux_eq_bind : List.reverseRecOn r (by simp [sublistsAux]) (fun r l ih => by - rw [append_bind, ← ih, bind_singleton, sublistsAux, foldl_append] + rw [bind_append, ← ih, bind_singleton, sublistsAux, foldl_append] simp [sublistsAux]) @[csimp] theorem sublists_eq_sublistsFast : @sublists = @sublistsFast := by @@ -177,7 +177,7 @@ theorem length_sublists (l : List α) : length (sublists l) = 2 ^ length l := by theorem map_pure_sublist_sublists (l : List α) : map pure l <+ sublists l := by induction' l using reverseRecOn with l a ih <;> simp only [map, map_append, sublists_concat] - · simp only [sublists_nil, sublist_cons] + · simp only [sublists_nil, sublist_cons_self] exact ((append_sublist_append_left _).2 <| singleton_sublist.2 <| mem_map.2 ⟨[], mem_sublists.2 (nil_sublist _), by rfl⟩).trans ((append_sublist_append_right _).2 ih) @@ -411,7 +411,7 @@ theorem range_bind_sublistsLen_perm (l : List α) : simp_rw [← List.map_bind, ← cons_append] rw [← List.singleton_append, ← List.sublistsLen_zero tl] refine Perm.append ?_ (l_ih.map _) - rw [List.range_succ, append_bind, bind_singleton, + rw [List.range_succ, bind_append, bind_singleton, sublistsLen_of_length_lt (Nat.lt_succ_self _), append_nil, ← List.bind_map Nat.succ fun n => sublistsLen n tl, ← bind_cons 0 _ fun n => sublistsLen n tl, ← range_succ_eq_map] diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index ff9b6fc3ddfba..eaf32bb1f3904 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -269,7 +269,7 @@ protected theorem Sublist.sym (n : ℕ) {xs ys : List α} (h : xs <+ ys) : xs.sy · exact h.sym (n + 1) theorem sym_sublist_sym_cons {a : α} : xs.sym n <+ (a :: xs).sym n := - (sublist_cons a xs).sym n + (sublist_cons_self a xs).sym n theorem mem_of_mem_of_mem_sym {n : ℕ} {xs : List α} {a : α} {z : Sym α n} (ha : a ∈ z) (hz : z ∈ xs.sym n) : a ∈ xs := diff --git a/Mathlib/Data/List/TFAE.lean b/Mathlib/Data/List/TFAE.lean index dc6adea7587cc..f9bce9204f88b 100644 --- a/Mathlib/Data/List/TFAE.lean +++ b/Mathlib/Data/List/TFAE.lean @@ -79,7 +79,7 @@ example (P₁ P₂ P₃ : ℕ → Prop) (H : ∀ n, [P₁ n, P₂ n, P₃ n].TFA -/ theorem forall_tfae {α : Type*} (l : List (α → Prop)) (H : ∀ a : α, (l.map (fun p ↦ p a)).TFAE) : (l.map (fun p ↦ ∀ a, p a)).TFAE := by - simp only [TFAE, List.forall_mem_map_iff] + simp only [TFAE, List.forall_mem_map] intros p₁ hp₁ p₂ hp₂ exact forall_congr' fun a ↦ H a (p₁ a) (mem_map_of_mem (fun p ↦ p a) hp₁) (p₂ a) (mem_map_of_mem (fun p ↦ p a) hp₂) @@ -98,7 +98,7 @@ example (P₁ P₂ P₃ : ℕ → Prop) (H : ∀ n, [P₁ n, P₂ n, P₃ n].TFA -/ theorem exists_tfae {α : Type*} (l : List (α → Prop)) (H : ∀ a : α, (l.map (fun p ↦ p a)).TFAE) : (l.map (fun p ↦ ∃ a, p a)).TFAE := by - simp only [TFAE, List.forall_mem_map_iff] + simp only [TFAE, List.forall_mem_map] intros p₁ hp₁ p₂ hp₂ exact exists_congr fun a ↦ H a (p₁ a) (mem_map_of_mem (fun p ↦ p a) hp₁) (p₂ a) (mem_map_of_mem (fun p ↦ p a) hp₂) diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index b3e510eb899fa..5de1a6e0841c2 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -44,80 +44,10 @@ theorem forall_zipWith {f : α → β → γ} {p : γ → Prop} : simp only [length_cons, succ_inj'] at h simp [forall_zipWith h] -theorem lt_length_left_of_zipWith {f : α → β → γ} {i : ℕ} {l : List α} {l' : List β} - (h : i < (zipWith f l l').length) : i < l.length := by rw [length_zipWith] at h; omega - -theorem lt_length_right_of_zipWith {f : α → β → γ} {i : ℕ} {l : List α} {l' : List β} - (h : i < (zipWith f l l').length) : i < l'.length := by rw [length_zipWith] at h; omega - -theorem lt_length_left_of_zip {i : ℕ} {l : List α} {l' : List β} (h : i < (zip l l').length) : - i < l.length := - lt_length_left_of_zipWith h - -theorem lt_length_right_of_zip {i : ℕ} {l : List α} {l' : List β} (h : i < (zip l l').length) : - i < l'.length := - lt_length_right_of_zipWith h - -theorem mem_zip {a b} : ∀ {l₁ : List α} {l₂ : List β}, (a, b) ∈ zip l₁ l₂ → a ∈ l₁ ∧ b ∈ l₂ - | _ :: l₁, _ :: l₂, h => by - cases' h with _ _ _ h - · simp - · have := mem_zip h - exact ⟨Mem.tail _ this.1, Mem.tail _ this.2⟩ - -theorem unzip_eq_map : ∀ l : List (α × β), unzip l = (l.map Prod.fst, l.map Prod.snd) - | [] => rfl - | (a, b) :: l => by simp only [unzip_cons, map_cons, unzip_eq_map l] - -theorem unzip_left (l : List (α × β)) : (unzip l).1 = l.map Prod.fst := by simp only [unzip_eq_map] - -theorem unzip_right (l : List (α × β)) : (unzip l).2 = l.map Prod.snd := by simp only [unzip_eq_map] - theorem unzip_swap (l : List (α × β)) : unzip (l.map Prod.swap) = (unzip l).swap := by simp only [unzip_eq_map, map_map] rfl -theorem zip_unzip : ∀ l : List (α × β), zip (unzip l).1 (unzip l).2 = l - | [] => rfl - | (a, b) :: l => by simp only [unzip_cons, zip_cons_cons, zip_unzip l] - -theorem unzip_zip_left : - ∀ {l₁ : List α} {l₂ : List β}, length l₁ ≤ length l₂ → (unzip (zip l₁ l₂)).1 = l₁ - | [], l₂, _ => rfl - | l₁, [], h => by rw [eq_nil_of_length_eq_zero (Nat.eq_zero_of_le_zero h)]; rfl - | a :: l₁, b :: l₂, h => by - simp only [zip_cons_cons, unzip_cons, unzip_zip_left (le_of_succ_le_succ h)] - -theorem unzip_zip_right {l₁ : List α} {l₂ : List β} (h : length l₂ ≤ length l₁) : - (unzip (zip l₁ l₂)).2 = l₂ := by rw [← zip_swap, unzip_swap]; exact unzip_zip_left h - -theorem unzip_zip {l₁ : List α} {l₂ : List β} (h : length l₁ = length l₂) : - unzip (zip l₁ l₂) = (l₁, l₂) := by - rw [← Prod.mk.eta (p := unzip (zip l₁ l₂)), - unzip_zip_left (le_of_eq h), unzip_zip_right (ge_of_eq h)] - -theorem zip_of_prod {l : List α} {l' : List β} {lp : List (α × β)} (hl : lp.map Prod.fst = l) - (hr : lp.map Prod.snd = l') : lp = l.zip l' := by - rw [← hl, ← hr, ← zip_unzip lp, ← unzip_left, ← unzip_right, zip_unzip, zip_unzip] - -theorem map_prod_left_eq_zip {l : List α} (f : α → β) : - (l.map fun x => (x, f x)) = l.zip (l.map f) := by - rw [← zip_map'] - congr - exact map_id _ - -theorem map_prod_right_eq_zip {l : List α} (f : α → β) : - (l.map fun x => (f x, x)) = (l.map f).zip l := by - rw [← zip_map'] - congr - exact map_id _ - -theorem zipWith_comm (f : α → β → γ) : - ∀ (la : List α) (lb : List β), zipWith f la lb = zipWith (fun b a => f a b) lb la - | [], _ => List.zipWith_nil_right.symm - | _ :: _, [] => rfl - | _ :: as, _ :: bs => congr_arg _ (zipWith_comm f as bs) - @[congr] theorem zipWith_congr (f g : α → β → γ) (la : List α) (lb : List β) (h : List.Forall₂ (fun a b => f a b = g a b) la lb) : zipWith f la lb = zipWith g la lb := by @@ -125,16 +55,6 @@ theorem zipWith_congr (f g : α → β → γ) (la : List α) (lb : List β) · rfl · exact congr_arg₂ _ hfg ih -theorem zipWith_comm_of_comm (f : α → α → β) (comm : ∀ x y : α, f x y = f y x) (l l' : List α) : - zipWith f l l' = zipWith f l' l := by - rw [zipWith_comm] - simp only [comm] - -@[simp] -theorem zipWith_same (f : α → α → δ) : ∀ l : List α, zipWith f l l = l.map fun a => f a a - | [] => rfl - | _ :: xs => congr_arg _ (zipWith_same f xs) - theorem zipWith_zipWith_left (f : δ → γ → ε) (g : α → β → δ) : ∀ (la : List α) (lb : List β) (lc : List γ), zipWith f (zipWith g la lb) lc = zipWith3 (fun a b c => f (g a b) c) la lb lc @@ -185,11 +105,11 @@ theorem unzip_revzip (l : List α) : (revzip l).unzip = (l, l.reverse) := @[simp] theorem revzip_map_fst (l : List α) : (revzip l).map Prod.fst = l := by - rw [← unzip_left, unzip_revzip] + rw [← unzip_fst, unzip_revzip] @[simp] theorem revzip_map_snd (l : List α) : (revzip l).map Prod.snd = l.reverse := by - rw [← unzip_right, unzip_revzip] + rw [← unzip_snd, unzip_revzip] theorem reverse_revzip (l : List α) : reverse l.revzip = revzip l.reverse := by rw [← zip_unzip (revzip l).reverse] @@ -221,30 +141,10 @@ theorem get?_zip_with_eq_some (f : α → β → γ) (l₁ : List α) (l₂ : Li ∃ x y, l₁.get? i = some x ∧ l₂.get? i = some y ∧ f x y = z := by simp [getElem?_zip_with_eq_some] -theorem getElem?_zip_eq_some (l₁ : List α) (l₂ : List β) (z : α × β) (i : ℕ) : - (zip l₁ l₂)[i]? = some z ↔ l₁[i]? = some z.1 ∧ l₂[i]? = some z.2 := by - cases z - rw [zip, getElem?_zip_with_eq_some]; constructor - · rintro ⟨x, y, h₀, h₁, h₂⟩ - simpa [h₀, h₁] using h₂ - · rintro ⟨h₀, h₁⟩ - exact ⟨_, _, h₀, h₁, rfl⟩ - theorem get?_zip_eq_some (l₁ : List α) (l₂ : List β) (z : α × β) (i : ℕ) : (zip l₁ l₂).get? i = some z ↔ l₁.get? i = some z.1 ∧ l₂.get? i = some z.2 := by simp [getElem?_zip_eq_some] -@[simp] -theorem getElem_zipWith {f : α → β → γ} {l : List α} {l' : List β} - {i : Nat} {h : i < (zipWith f l l').length} : - (zipWith f l l')[i] = - f (l[i]'(lt_length_left_of_zipWith h)) - (l'[i]'(lt_length_right_of_zipWith h)) := by - rw [← Option.some_inj, ← getElem?_eq_getElem, getElem?_zip_with_eq_some] - exact - ⟨l[i]'(lt_length_left_of_zipWith h), l'[i]'(lt_length_right_of_zipWith h), - by rw [getElem?_eq_getElem], by rw [getElem?_eq_getElem]; exact ⟨rfl, rfl⟩⟩ - @[deprecated getElem_zipWith (since := "2024-06-12")] theorem get_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : Fin (zipWith f l l').length} : (zipWith f l l').get i = @@ -260,12 +160,6 @@ theorem nthLe_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : f (l.nthLe i (lt_length_left_of_zipWith h)) (l'.nthLe i (lt_length_right_of_zipWith h)) := get_zipWith (i := ⟨i, h⟩) -@[simp] -theorem getElem_zip {l : List α} {l' : List β} {i : Nat} {h : i < (zip l l').length} : - (zip l l')[i] = - (l[i]'(lt_length_left_of_zip h), l'[i]'(lt_length_right_of_zip h)) := - getElem_zipWith (h := h) - @[deprecated getElem_zip (since := "2024-06-12")] theorem get_zip {l : List α} {l' : List β} {i : Fin (zip l l').length} : (zip l l').get i = @@ -304,23 +198,4 @@ theorem map_uncurry_zip_eq_zipWith (f : α → β → γ) (l : List α) (l' : Li · simp · simp [hl] -section Distrib - -/-! ### Operations that can be applied before or after a `zip_with` -/ - - -variable (f : α → β → γ) (l : List α) (l' : List β) (n : ℕ) - -theorem zipWith_distrib_reverse (h : l.length = l'.length) : - (zipWith f l l').reverse = zipWith f l.reverse l'.reverse := by - induction' l with hd tl hl generalizing l' - · simp - · cases' l' with hd' tl' - · simp - · simp only [Nat.add_left_inj, length] at h - have : tl.reverse.length = tl'.reverse.length := by simp [h] - simp [hl _ h, zipWith_append _ _ _ _ _ this] - -end Distrib - end List diff --git a/Mathlib/Data/Matrix/RowCol.lean b/Mathlib/Data/Matrix/RowCol.lean index b621c008f32d2..2529d88fd19ee 100644 --- a/Mathlib/Data/Matrix/RowCol.lean +++ b/Mathlib/Data/Matrix/RowCol.lean @@ -54,15 +54,16 @@ variable {ι : Type*} theorem row_apply (v : n → α) (i : ι) (j) : row ι v i j = v j := rfl -theorem col_injective [Inhabited ι] : Function.Injective (col ι : (m → α) → Matrix m ι α) := - fun _x _y h => funext fun i => congr_fun₂ h i default +theorem col_injective [Nonempty ι] : Function.Injective (col ι : (m → α) → Matrix m ι α) := by + inhabit ι + exact fun _x _y h => funext fun i => congr_fun₂ h i default -@[simp] theorem col_inj [Inhabited ι] {v w : m → α} : col ι v = col ι w ↔ v = w := +@[simp] theorem col_inj [Nonempty ι] {v w : m → α} : col ι v = col ι w ↔ v = w := col_injective.eq_iff @[simp] theorem col_zero [Zero α] : col ι (0 : m → α) = 0 := rfl -@[simp] theorem col_eq_zero [Zero α] [Inhabited ι] (v : m → α) : col ι v = 0 ↔ v = 0 := col_inj +@[simp] theorem col_eq_zero [Zero α] [Nonempty ι] (v : m → α) : col ι v = 0 ↔ v = 0 := col_inj @[simp] theorem col_add [Add α] (v w : m → α) : col ι (v + w) = col ι v + col ι w := by @@ -74,15 +75,16 @@ theorem col_smul [SMul R α] (x : R) (v : m → α) : col ι (x • v) = x • c ext rfl -theorem row_injective [Inhabited ι] : Function.Injective (row ι : (n → α) → Matrix ι n α) := - fun _x _y h => funext fun j => congr_fun₂ h default j +theorem row_injective [Nonempty ι] : Function.Injective (row ι : (n → α) → Matrix ι n α) := by + inhabit ι + exact fun _x _y h => funext fun j => congr_fun₂ h default j -@[simp] theorem row_inj [Inhabited ι] {v w : n → α} : row ι v = row ι w ↔ v = w := +@[simp] theorem row_inj [Nonempty ι] {v w : n → α} : row ι v = row ι w ↔ v = w := row_injective.eq_iff @[simp] theorem row_zero [Zero α] : row ι (0 : n → α) = 0 := rfl -@[simp] theorem row_eq_zero [Zero α] [Inhabited ι] (v : n → α) : row ι v = 0 ↔ v = 0 := row_inj +@[simp] theorem row_eq_zero [Zero α] [Nonempty ι] (v : n → α) : row ι v = 0 ↔ v = 0 := row_inj @[simp] theorem row_add [Add α] (v w : m → α) : row ι (v + w) = row ι v + row ι w := by diff --git a/Mathlib/Data/Matroid/Map.lean b/Mathlib/Data/Matroid/Map.lean index 5209658a08d7e..74b9dcb61b27b 100644 --- a/Mathlib/Data/Matroid/Map.lean +++ b/Mathlib/Data/Matroid/Map.lean @@ -147,7 +147,8 @@ def comap (N : Matroid β) (f : α → β) : Matroid α := obtain ⟨J, hJ⟩ := (N ↾ range f).existsMaximalSubsetProperty_indep (f '' X) (by simp) (f '' I) (by simpa) (image_subset _ hIX) - simp only [restrict_indep_iff, image_subset_iff, mem_maximals_iff, mem_setOf_eq, and_imp] at hJ + simp only [restrict_indep_iff, image_subset_iff, mem_maximals_iff, mem_setOf_eq, + and_imp] at hJ obtain ⟨J₀, hIJ₀, hJ₀X, hbj⟩ := hIinj.bijOn_image.exists_extend_of_subset hIX (image_subset f hJ.1.2.1) (image_subset_iff.2 <| preimage_mono hJ.1.2.2) @@ -516,6 +517,7 @@ def mapSetEquiv (M : Matroid α) {E : Set β} (e : M.E ≃ E) : Matroid β := @[simp] lemma mapSetEquiv.ground (M : Matroid α) {E : Set β} (e : M.E ≃ E) : (M.mapSetEquiv e).E = E := rfl +end mapSetEquiv section mapEmbedding /-- Map `M : Matroid α` across an embedding defined on all of `α` -/ @@ -544,6 +546,24 @@ lemma Basis.mapEmbedding {X : Set α} (hIX : M.Basis I X) (f : α ↪ β) : (M.mapEmbedding f).Basis (f '' I) (f '' X) := by apply hIX.map +@[simp] lemma mapEmbedding_base_iff {f : α ↪ β} {B : Set β} : + (M.mapEmbedding f).Base B ↔ M.Base (f ⁻¹' B) ∧ B ⊆ range f := by + rw [mapEmbedding, map_base_iff] + refine ⟨?_, fun ⟨h,h'⟩ ↦ ⟨f ⁻¹' B, h, by rwa [eq_comm, image_preimage_eq_iff]⟩⟩ + rintro ⟨B, hB, rfl⟩ + rw [preimage_image_eq _ f.injective] + exact ⟨hB, image_subset_range _ _⟩ + +@[simp] lemma mapEmbedding_basis_iff {f : α ↪ β} {I X : Set β} : + (M.mapEmbedding f).Basis I X ↔ M.Basis (f ⁻¹' I) (f ⁻¹' X) ∧ I ⊆ X ∧ X ⊆ range f := by + rw [mapEmbedding, map_basis_iff'] + refine ⟨?_, fun ⟨hb, hIX, hX⟩ ↦ ?_⟩ + · rintro ⟨I, X, hIX, rfl, rfl⟩ + simp [preimage_image_eq _ f.injective, image_subset f hIX.subset, hIX] + obtain ⟨X, rfl⟩ := subset_range_iff_exists_image_eq.1 hX + obtain ⟨I, -, rfl⟩ := subset_image_iff.1 hIX + exact ⟨I, X, by simpa [preimage_image_eq _ f.injective] using hb⟩ + instance [M.Nonempty] {f : α ↪ β} : (M.mapEmbedding f).Nonempty := inferInstanceAs (M.map f f.injective.injOn).Nonempty @@ -585,6 +605,13 @@ lemma mapEquiv_eq_map (f : α ≃ β) : M.mapEquiv f = M.map f f.injective.injOn rw [mapEquiv_eq_map, map_base_iff] exact ⟨by rintro ⟨I, hI, rfl⟩; simpa, fun h ↦ ⟨_, h, by simp⟩⟩ +@[simp] lemma mapEquiv_basis_iff {α β : Type*} {M : Matroid α} (f : α ≃ β) {I X : Set β} : + (M.mapEquiv f).Basis I X ↔ M.Basis (f.symm '' I) (f.symm '' X) := by + rw [mapEquiv_eq_map, map_basis_iff'] + refine ⟨fun h ↦ ?_, fun h ↦ ⟨_, _, h, by simp, by simp⟩⟩ + obtain ⟨I, X, hIX, rfl, rfl⟩ := h + simpa + instance [M.Nonempty] {f : α ≃ β} : (M.mapEquiv f).Nonempty := inferInstanceAs (M.map f f.injective.injOn).Nonempty @@ -626,6 +653,25 @@ lemma restrictSubtype_inter_indep_iff : (M.restrictSubtype X).Indep (X ↓∩ I) ↔ M.Indep (X ∩ I) := by simp [restrictSubtype, Subtype.val_injective.injOn] +lemma restrictSubtype_basis_iff {Y : Set α} {I X : Set Y} : + (M.restrictSubtype Y).Basis I X ↔ M.Basis' I X := by + rw [restrictSubtype, comap_basis_iff, and_iff_right Subtype.val_injective.injOn, + and_iff_left_of_imp, basis_restrict_iff', basis'_iff_basis_inter_ground] + · simp + exact fun h ↦ (image_subset_image_iff Subtype.val_injective).1 h.subset + +lemma restrictSubtype_base_iff {B : Set X} : (M.restrictSubtype X).Base B ↔ M.Basis' B X := by + rw [restrictSubtype, comap_base_iff] + simp [Subtype.val_injective.injOn, Subset.rfl, basis_restrict_iff', basis'_iff_basis_inter_ground] + +@[simp] lemma restrictSubtype_ground_base_iff {B : Set M.E} : + (M.restrictSubtype M.E).Base B ↔ M.Base B := by + rw [restrictSubtype_base_iff, basis'_iff_basis, basis_ground_iff] + +@[simp] lemma restrictSubtype_ground_basis_iff {I X : Set M.E} : + (M.restrictSubtype M.E).Basis I X ↔ M.Basis I X := by + rw [restrictSubtype_basis_iff, basis'_iff_basis] + lemma eq_of_restrictSubtype_eq {N : Matroid α} (hM : M.E = E) (hN : N.E = E) (h : M.restrictSubtype E = N.restrictSubtype E) : M = N := by subst hM @@ -640,7 +686,7 @@ lemma eq_of_restrictSubtype_eq {N : Matroid α} (hM : M.E = E) (hN : N.E = E) lemma restrictSubtype_dual' (hM : M.E = E) : (M.restrictSubtype E)✶ = M✶.restrictSubtype E := by rw [← hM, restrictSubtype_dual] -/-- `M.restrictSubtype M.E` is isomorphic to `M ↾ X`. -/ +/-- `M.restrictSubtype X` is isomorphic to `M ↾ X`. -/ @[simp] lemma map_val_restrictSubtype_eq (M : Matroid α) (X : Set α) : (M.restrictSubtype X).map (↑) Subtype.val_injective.injOn = M ↾ X := by simp [restrictSubtype, map_comap, Subset.rfl] @@ -671,6 +717,4 @@ instance [M.RkPos] : (M.restrictSubtype M.E).RkPos := by end restrictSubtype -end mapSetEquiv - end Matroid diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index ee8808b3ccf1d..bfc1ded9ed46a 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -416,9 +416,7 @@ theorem coe_toList (s : Multiset α) : (s.toList : Multiset α) = s := theorem toList_eq_nil {s : Multiset α} : s.toList = [] ↔ s = 0 := by rw [← coe_eq_zero, coe_toList] -@[simp] -theorem empty_toList {s : Multiset α} : s.toList.isEmpty ↔ s = 0 := - isEmpty_iff_eq_nil.trans toList_eq_nil +theorem empty_toList {s : Multiset α} : s.toList.isEmpty ↔ s = 0 := by simp @[simp] theorem toList_zero : (Multiset.toList 0 : List α) = [] := @@ -498,7 +496,8 @@ theorem le_zero : s ≤ 0 ↔ s = 0 := theorem lt_cons_self (s : Multiset α) (a : α) : s < a ::ₘ s := Quot.inductionOn s fun l => suffices l <+~ a :: l ∧ ¬l ~ a :: l by simpa [lt_iff_le_and_ne] - ⟨(sublist_cons _ _).subperm, fun p => _root_.ne_of_lt (lt_succ_self (length l)) p.length_eq⟩ + ⟨(sublist_cons_self _ _).subperm, + fun p => _root_.ne_of_lt (lt_succ_self (length l)) p.length_eq⟩ theorem le_cons_self (s : Multiset α) (a : α) : s ≤ a ::ₘ s := le_of_lt <| lt_cons_self _ _ @@ -636,11 +635,16 @@ theorem mem_of_mem_nsmul {a : α} {s : Multiset α} {n : ℕ} (h : a ∈ n • s exact h.elim ih id @[simp] -theorem mem_nsmul {a : α} {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : a ∈ n • s ↔ a ∈ s := by - refine ⟨mem_of_mem_nsmul, fun h => ?_⟩ - obtain ⟨n, rfl⟩ := exists_eq_succ_of_ne_zero h0 +theorem mem_nsmul {a : α} {s : Multiset α} {n : ℕ} : a ∈ n • s ↔ n ≠ 0 ∧ a ∈ s := by + refine ⟨fun ha ↦ ⟨?_, mem_of_mem_nsmul ha⟩, fun h => ?_⟩ + · rintro rfl + simp [zero_nsmul] at ha + obtain ⟨n, rfl⟩ := exists_eq_succ_of_ne_zero h.1 rw [succ_nsmul, mem_add] - exact Or.inr h + exact Or.inr h.2 + +lemma mem_nsmul_of_ne_zero {a : α} {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : a ∈ n • s ↔ a ∈ s := by + simp [*] theorem nsmul_cons {s : Multiset α} (n : ℕ) (a : α) : n • (a ::ₘ s) = n • ({a} : Multiset α) + n • s := by @@ -1022,7 +1026,7 @@ theorem map_hcongr {β' : Type v} {m : Multiset α} {f : α → β} {f' : α → theorem forall_mem_map_iff {f : α → β} {p : β → Prop} {s : Multiset α} : (∀ y ∈ s.map f, p y) ↔ ∀ x ∈ s, p (f x) := - Quotient.inductionOn' s fun _L => List.forall_mem_map_iff + Quotient.inductionOn' s fun _L => List.forall_mem_map @[simp, norm_cast] lemma map_coe (f : α → β) (l : List α) : map f l = l.map f := rfl @@ -1142,7 +1146,7 @@ theorem map_const (s : Multiset α) (b : β) : map (const α b) s = replicate (c theorem eq_of_mem_map_const {b₁ b₂ : β} {l : List α} (h : b₁ ∈ map (Function.const α b₂) l) : b₁ = b₂ := - eq_of_mem_replicate <| by rwa [map_const] at h + eq_of_mem_replicate (n := card (l : Multiset α)) <| by rwa [map_const] at h @[simp] theorem map_le_map {f : α → β} {s t : Multiset α} (h : s ≤ t) : map f s ≤ map f t := @@ -1344,7 +1348,7 @@ theorem pmap_eq_map_attach {p : α → Prop} (f : ∀ a, p a → β) (s) : -- @[simp] -- Porting note: Left hand does not simplify theorem attach_map_val' (s : Multiset α) (f : α → β) : (s.attach.map fun i => f i.val) = s.map f := - Quot.inductionOn s fun l => congr_arg _ <| List.attach_map_coe' l f + Quot.inductionOn s fun l => congr_arg _ <| List.attach_map_coe l f @[simp] theorem attach_map_val (s : Multiset α) : s.attach.map Subtype.val = s := @@ -1651,7 +1655,8 @@ theorem sub_add_inter (s t : Multiset α) : s - t + s ∩ t = s := by · rw [cons_inter_of_neg _ h, sub_cons, erase_of_not_mem h, IH] theorem sub_inter (s t : Multiset α) : s - s ∩ t = s - t := - add_right_cancel <| by rw [sub_add_inter s t, tsub_add_cancel_of_le (inter_le_left s t)] + add_right_cancel (b := s ∩ t) <| by + rw [sub_add_inter s t, tsub_add_cancel_of_le (inter_le_left s t)] end @@ -1716,7 +1721,7 @@ theorem filter_cons_of_neg {a : α} (s) : ¬p a → filter p (a ::ₘ s) = filte @[simp] theorem mem_filter {a : α} {s} : a ∈ filter p s ↔ a ∈ s ∧ p a := - Quot.inductionOn s fun _l => by simpa using List.mem_filter (p := (p ·)) + Quot.inductionOn s fun _l => by simp theorem of_mem_filter {a : α} {s} (h : a ∈ filter p s) : p a := (mem_filter.1 h).2 @@ -2134,9 +2139,10 @@ theorem count_replicate_self (a : α) (n : ℕ) : count a (replicate n a) = n := convert List.count_replicate_self a n rw [← coe_count, coe_replicate] -theorem count_replicate (a b : α) (n : ℕ) : count a (replicate n b) = if a = b then n else 0 := by +theorem count_replicate (a b : α) (n : ℕ) : count a (replicate n b) = if b = a then n else 0 := by convert List.count_replicate a b n rw [← coe_count, coe_replicate] + simp @[simp] theorem count_erase_self (a : α) (s : Multiset α) : count a (erase s a) = count a s - 1 := @@ -2200,6 +2206,9 @@ theorem ext {s t : Multiset α} : s = t ↔ ∀ a, count a s = count a t := theorem ext' {s t : Multiset α} : (∀ a, count a s = count a t) → s = t := ext.2 +lemma count_injective : Injective fun (s : Multiset α) a ↦ s.count a := + fun _s _t hst ↦ ext' $ congr_fun hst + @[simp] theorem coe_inter (s t : List α) : (s ∩ t : Multiset α) = (s.bagInter t : List α) := by ext; simp @@ -2258,7 +2267,7 @@ theorem replicate_inter (n : ℕ) (x : α) (s : Multiset α) : replicate n x ∩ s = replicate (min n (s.count x)) x := by ext y rw [count_inter, count_replicate, count_replicate] - by_cases h : y = x + by_cases h : x = y · simp only [h, if_true] · simp only [h, if_false, Nat.zero_min] diff --git a/Mathlib/Data/Multiset/FinsetOps.lean b/Mathlib/Data/Multiset/FinsetOps.lean index 54883cdea421c..0e2b4e57ed01b 100644 --- a/Mathlib/Data/Multiset/FinsetOps.lean +++ b/Mathlib/Data/Multiset/FinsetOps.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Multiset.Dedup +import Mathlib.Data.List.Infix /-! # Preparations for defining operations on `Finset`. diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index 973273b1efcc5..9f48dea81db32 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -103,6 +103,37 @@ theorem Multiset.mem_toEnumFinset (m : Multiset α) (p : α × ℕ) : theorem Multiset.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 + +@[simp] lemma map_toEnumFinset_fst (m : Multiset α) : m.toEnumFinset.val.map Prod.fst = m := by + ext a; simp [count_map, ← Finset.filter_val, eq_comm (a := a)] + +@[simp] lemma image_toEnumFinset_fst (m : Multiset α) : + m.toEnumFinset.image Prod.fst = m.toFinset := by + rw [Finset.image, Multiset.map_toEnumFinset_fst] + +@[simp] lemma map_fst_le_of_subset_toEnumFinset {s : Finset (α × ℕ)} (hsm : s ⊆ m.toEnumFinset) : + s.1.map Prod.fst ≤ m := by + simp_rw [le_iff_count, count_map] + rintro a + obtain ha | ha := (s.1.filter fun x ↦ a = x.1).card.eq_zero_or_pos + · rw [ha] + exact Nat.zero_le _ + obtain ⟨n, han, hn⟩ : ∃ n ≥ card (s.1.filter fun x ↦ a = x.1) - 1, (a, n) ∈ s := by + by_contra! h + replace h : s.filter (·.1 = a) ⊆ {a} ×ˢ .range (card (s.1.filter fun x ↦ a = x.1) - 1) := by + simpa (config := { contextual := true }) [forall_swap (β := _ = a), Finset.subset_iff, + imp_not_comm, not_le, Nat.lt_sub_iff_add_lt] using h + have : card (s.1.filter fun x ↦ a = x.1) ≤ card (s.1.filter fun x ↦ a = x.1) - 1 := by + simpa [Finset.card, eq_comm] using Finset.card_mono h + 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₂) : m₁.toEnumFinset ⊆ m₂.toEnumFinset := by @@ -112,17 +143,8 @@ theorem Multiset.toEnumFinset_mono {m₁ m₂ : Multiset α} (h : m₁ ≤ m₂) @[simp] theorem Multiset.toEnumFinset_subset_iff {m₁ m₂ : Multiset α} : - m₁.toEnumFinset ⊆ m₂.toEnumFinset ↔ m₁ ≤ m₂ := by - refine ⟨fun h ↦ ?_, Multiset.toEnumFinset_mono⟩ - rw [Multiset.le_iff_count] - intro x - by_cases hx : x ∈ m₁ - · apply Nat.le_of_pred_lt - have : (x, m₁.count x - 1) ∈ m₁.toEnumFinset := by - rw [Multiset.mem_toEnumFinset] - exact Nat.pred_lt (ne_of_gt (Multiset.count_pos.mpr hx)) - simpa only [Multiset.mem_toEnumFinset] using h this - · simp [hx] + 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 `(↑)`. -/ @@ -169,27 +191,6 @@ theorem Multiset.map_univ_coeEmbedding (m : Multiset α) : exists_prop, exists_eq_right_right, exists_eq_right, Multiset.mem_toEnumFinset, iff_self_iff, true_and_iff] -theorem Multiset.toEnumFinset_filter_eq (m : Multiset α) (x : α) : - (m.toEnumFinset.filter fun p ↦ x = p.1) = - (Finset.range (m.count x)).map ⟨Prod.mk x, Prod.mk.inj_left x⟩ := by - ext ⟨y, i⟩ - simp only [eq_comm, Finset.mem_filter, Multiset.mem_toEnumFinset, Finset.mem_map, - Finset.mem_range, Function.Embedding.coeFn_mk, Prod.mk.inj_iff, exists_prop, - exists_eq_right_right', and_congr_left_iff] - rintro rfl - rfl - -@[simp] -theorem Multiset.map_toEnumFinset_fst (m : Multiset α) : m.toEnumFinset.val.map Prod.fst = m := by - ext x - simp only [Multiset.count_map, ← Finset.filter_val, Multiset.toEnumFinset_filter_eq, - Finset.map_val, Finset.range_val, Multiset.card_map, Multiset.card_range] - -@[simp] -theorem Multiset.image_toEnumFinset_fst (m : Multiset α) : - m.toEnumFinset.image Prod.fst = m.toFinset := by - rw [Finset.image, Multiset.map_toEnumFinset_fst] - @[simp] theorem Multiset.map_univ_coe (m : Multiset α) : (Finset.univ : Finset m).val.map (fun x : m ↦ (x : α)) = m := by diff --git a/Mathlib/Data/Multiset/Nodup.lean b/Mathlib/Data/Multiset/Nodup.lean index 7ee834987689f..7e24031fb1f8c 100644 --- a/Mathlib/Data/Multiset/Nodup.lean +++ b/Mathlib/Data/Multiset/Nodup.lean @@ -3,8 +3,8 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Data.List.Range import Mathlib.Data.Multiset.Range +import Mathlib.Data.List.Pairwise /-! # The `Nodup` predicate for multisets without duplicate elements. diff --git a/Mathlib/Data/Multiset/Pi.lean b/Mathlib/Data/Multiset/Pi.lean index bc3b443f59811..a74f22d88819b 100644 --- a/Mathlib/Data/Multiset/Pi.lean +++ b/Mathlib/Data/Multiset/Pi.lean @@ -7,6 +7,10 @@ import Mathlib.Data.Multiset.Bind /-! # The cartesian product of multisets + +## Main definitions + +* `Multiset.pi`: Cartesian product of multisets indexed by a multiset. -/ @@ -14,33 +18,35 @@ namespace Multiset section Pi -variable {α : Type*} - open Function -/-- Given `δ : α → Type*`, `Pi.empty δ` is the trivial dependent function out of the empty +namespace Pi +variable {α : Type*} [DecidableEq α] {δ : α → Sort*} + +/-- Given `δ : α → Sort*`, `Pi.empty δ` is the trivial dependent function out of the empty multiset. -/ -def Pi.empty (δ : α → Sort*) : ∀ a ∈ (0 : Multiset α), δ a := +def empty (δ : α → Sort*) : ∀ a ∈ (0 : Multiset α), δ a := nofun -universe u v -variable [DecidableEq α] {β : α → Type u} {δ : α → Sort v} +variable (m : Multiset α) (a : α) -/-- Given `δ : α → Type*`, a multiset `m` and a term `a`, as well as a term `b : δ a` and a +/-- Given `δ : α → Sort*`, a multiset `m` and a term `a`, as well as a term `b : δ a` and a function `f` such that `f a' : δ a'` for all `a'` in `m`, `Pi.cons m a b f` is a function `g` such that `g a'' : δ a''` for all `a''` in `a ::ₘ m`. -/ -def Pi.cons (m : Multiset α) (a : α) (b : δ a) (f : ∀ a ∈ m, δ a) : ∀ a' ∈ a ::ₘ m, δ a' := +def cons (b : δ a) (f : ∀ a ∈ m, δ a) : ∀ a' ∈ a ::ₘ m, δ a' := fun a' ha' => if h : a' = a then Eq.ndrec b h.symm else f a' <| (mem_cons.1 ha').resolve_left h -theorem Pi.cons_same {m : Multiset α} {a : α} {b : δ a} {f : ∀ a ∈ m, δ a} (h : a ∈ a ::ₘ m) : - Pi.cons m a b f a h = b := +variable {m a} + +theorem cons_same {b : δ a} {f : ∀ a ∈ m, δ a} (h : a ∈ a ::ₘ m) : + cons m a b f a h = b := dif_pos rfl -theorem Pi.cons_ne {m : Multiset α} {a a' : α} {b : δ a} {f : ∀ a ∈ m, δ a} (h' : a' ∈ a ::ₘ m) +theorem cons_ne {a a' : α} {b : δ a} {f : ∀ a ∈ m, δ a} (h' : a' ∈ a ::ₘ m) (h : a' ≠ a) : Pi.cons m a b f a' h' = f a' ((mem_cons.1 h').resolve_left h) := dif_neg h -theorem Pi.cons_swap {a a' : α} {b : δ a} {b' : δ a'} {m : Multiset α} {f : ∀ a ∈ m, δ a} +theorem cons_swap {a a' : α} {b : δ a} {b' : δ a'} {m : Multiset α} {f : ∀ a ∈ m, δ a} (h : a ≠ a') : HEq (Pi.cons (a' ::ₘ m) a b (Pi.cons m a' b' f)) (Pi.cons (a ::ₘ m) a' b' (Pi.cons m a b f)) := by apply hfunext rfl @@ -52,15 +58,33 @@ theorem Pi.cons_swap {a a' : α} {b : δ a} {b' : δ a'} {m : Multiset α} {f : all_goals simp [*, Pi.cons_same, Pi.cons_ne] @[simp, nolint simpNF] -- Porting note: false positive, this lemma can prove itself -theorem pi.cons_eta {m : Multiset α} {a : α} (f : ∀ a' ∈ a ::ₘ m, δ a') : - (Pi.cons m a (f _ (mem_cons_self _ _)) fun a' ha' => f a' (mem_cons_of_mem ha')) = f := by +theorem cons_eta {m : Multiset α} {a : α} (f : ∀ a' ∈ a ::ₘ m, δ a') : + (cons m a (f _ (mem_cons_self _ _)) fun a' ha' => f a' (mem_cons_of_mem ha')) = f := by ext a' h' by_cases h : a' = a · subst h rw [Pi.cons_same] · rw [Pi.cons_ne _ h] -theorem Pi.cons_injective {a : α} {b : δ a} {s : Multiset α} (hs : a ∉ s) : +theorem cons_map (b : δ a) (f : ∀ a' ∈ m, δ a') + {δ' : α → Sort*} (φ : ∀ ⦃a'⦄, δ a' → δ' a') : + Pi.cons _ _ (φ b) (fun a' ha' ↦ φ (f a' ha')) = (fun a' ha' ↦ φ ((cons _ _ b f) a' ha')) := by + ext a' ha' + refine (congrArg₂ _ ?_ rfl).trans (apply_dite (@φ _) (a' = a) _ _).symm + ext rfl + rfl + +theorem forall_rel_cons_ext {r : ∀ ⦃a⦄, δ a → δ a → Prop} {b₁ b₂ : δ a} {f₁ f₂ : ∀ a' ∈ m, δ a'} + (hb : r b₁ b₂) (hf : ∀ (a : α) (ha : a ∈ m), r (f₁ a ha) (f₂ a ha)) : + ∀ a ha, r (cons _ _ b₁ f₁ a ha) (cons _ _ b₂ f₂ a ha) := by + intro a ha + dsimp [cons] + split_ifs with H + · cases H + exact hb + · exact hf _ _ + +theorem cons_injective {a : α} {b : δ a} {s : Multiset α} (hs : a ∉ s) : Function.Injective (Pi.cons s a b) := fun f₁ f₂ eq => funext fun a' => funext fun h' => @@ -68,8 +92,13 @@ theorem Pi.cons_injective {a : α} {b : δ a} {s : Multiset α} (hs : a ∉ s) : have : a' ∈ a ::ₘ s := mem_cons_of_mem h' calc f₁ a' h' = Pi.cons s a b f₁ a' this := by rw [Pi.cons_ne this ne.symm] - _ = Pi.cons s a b f₂ a' this := by rw [eq] - _ = f₂ a' h' := by rw [Pi.cons_ne this ne.symm] + _ = Pi.cons s a b f₂ a' this := by rw [eq] + _ = f₂ a' h' := by rw [Pi.cons_ne this ne.symm] + +end Pi + +section +variable {α : Type*} [DecidableEq α] {β : α → Type*} /-- `pi m t` constructs the Cartesian product over `t` indexed by `m`. -/ def pi (m : Multiset α) (t : ∀ a, Multiset (β a)) : Multiset (∀ a ∈ m, β a) := @@ -139,7 +168,9 @@ theorem mem_pi (m : Multiset α) (t : ∀ a, Multiset (β a)) : apply hf' · intro hf refine ⟨_, hf a (mem_cons_self _ _), _, fun a ha => hf a (mem_cons_of_mem ha), ?_⟩ - rw [pi.cons_eta] + rw [Pi.cons_eta] + +end end Pi diff --git a/Mathlib/Data/Multiset/Powerset.lean b/Mathlib/Data/Multiset/Powerset.lean index 42c6e7e46e363..e344eddc1b02f 100644 --- a/Mathlib/Data/Multiset/Powerset.lean +++ b/Mathlib/Data/Multiset/Powerset.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.List.Sublists +import Mathlib.Data.List.Zip import Mathlib.Data.Multiset.Bind /-! diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index cd24dda1d818c..5088ab573657f 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -3,12 +3,9 @@ Copyright (c) 2014 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Data.Nat.Cast.Order.Basic -import Mathlib.Data.Nat.Cast.Basic -import Mathlib.Algebra.CharZero.Defs -import Mathlib.Algebra.Order.Group.Abs -import Mathlib.Data.Nat.Cast.NeZero import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Algebra.Order.Group.Unbundled.Abs +import Mathlib.Data.Nat.Cast.Order.Basic /-! # Cast of natural numbers: lemmas about bundled ordered semirings diff --git a/Mathlib/Data/Nat/ChineseRemainder.lean b/Mathlib/Data/Nat/ChineseRemainder.lean index d6116845cee31..17beeef782677 100644 --- a/Mathlib/Data/Nat/ChineseRemainder.lean +++ b/Mathlib/Data/Nat/ChineseRemainder.lean @@ -146,8 +146,8 @@ theorem chineseRemainderOfMultiset_lt_prod {m : Multiset ι} (nod : m.Nodup) (hs : ∀ i ∈ m, s i ≠ 0) (pp : Set.Pairwise {x | x ∈ m} (Coprime on s)) : chineseRemainderOfMultiset a s nod hs pp < (m.map s).prod := by induction' m using Quot.ind with l - unfold chineseRemainderOfMultiset; simp - exact chineseRemainderOfList_lt_prod a s l + unfold chineseRemainderOfMultiset + simpa using chineseRemainderOfList_lt_prod a s l (List.Nodup.pairwise_of_forall_ne nod pp) (by simpa using hs) /-- The natural number less than `∏ i ∈ t, s i` congruent to diff --git a/Mathlib/Data/Nat/Choose/Dvd.lean b/Mathlib/Data/Nat/Choose/Dvd.lean index 62b7d9e198bc5..e5cca9d05b77b 100644 --- a/Mathlib/Data/Nat/Choose/Dvd.lean +++ b/Mathlib/Data/Nat/Choose/Dvd.lean @@ -31,7 +31,7 @@ lemma dvd_choose (hp : Prime p) (ha : a < p) (hab : b - a < p) (h : p ≤ b) : p this ▸ hp.dvd_choose_add ha hab (this.symm ▸ h) lemma dvd_choose_self (hp : Prime p) (hk : k ≠ 0) (hkp : k < p) : p ∣ choose p k := - hp.dvd_choose hkp (sub_lt ((zero_le _).trans_lt hkp) hk.bot_lt) le_rfl + hp.dvd_choose hkp (sub_lt ((zero_le _).trans_lt hkp) <| zero_lt_of_ne_zero hk) le_rfl end Prime diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean new file mode 100644 index 0000000000000..c896d134554d5 --- /dev/null +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -0,0 +1,100 @@ +/- +Copyright (c) 2023 Gareth Ma. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Gareth Ma +-/ +import Mathlib.Data.ZMod.Basic +import Mathlib.RingTheory.Polynomial.Basic + +/-! +# Lucas's theorem + +This file contains a proof of [Lucas's theorem](https://en.wikipedia.org/wiki/Lucas's_theorem) about +binomial coefficients, which says that for primes `p`, `n` choose `k` is congruent to product of +`n_i` choose `k_i` modulo `p`, where `n_i` and `k_i` are the base-`p` digits of `n` and `k`, +respectively. + +## Main statements + +* `lucas_theorem`: the binomial coefficient `n choose k` is congruent to the product of `n_i choose +k_i` modulo `p`, where `n_i` and `k_i` are the base-`p` digits of `n` and `k`, respectively. +-/ + +open Finset hiding choose + +open Nat BigOperators Polynomial + +namespace Choose + +variable {n k p : ℕ} [Fact p.Prime] + +/-- For primes `p`, `choose n k` is congruent to `choose (n % p) (k % p) * choose (n / p) (k / p)` +modulo `p`. Also see `choose_modEq_choose_mod_mul_choose_div_nat` for the version with `MOD`. -/ +theorem choose_modEq_choose_mod_mul_choose_div : + choose n k ≡ choose (n % p) (k % p) * choose (n / p) (k / p) [ZMOD p] := by + have decompose : ((X : (ZMod p)[X]) + 1) ^ n = (X + 1) ^ (n % p) * (X ^ p + 1) ^ (n / p) := by + simpa using add_pow_eq_add_pow_mod_mul_pow_add_pow_div _ (X : (ZMod p)[X]) 1 + simp only [← ZMod.intCast_eq_intCast_iff, Int.cast_mul, Int.cast_ofNat, + ← coeff_X_add_one_pow _ n k, ← eq_intCast (Int.castRingHom (ZMod p)), ← coeff_map, + Polynomial.map_pow, Polynomial.map_add, Polynomial.map_one, map_X, decompose] + simp only [add_pow, one_pow, mul_one, ← pow_mul, sum_mul_sum] + conv_lhs => + enter [1, 2, k, 2, k'] + rw [← mul_assoc, mul_right_comm _ _ (X ^ (p * k')), ← pow_add, mul_assoc, ← cast_mul] + have h_iff : ∀ x ∈ range (n % p + 1) ×ˢ range (n / p + 1), + k = x.1 + p * x.2 ↔ (k % p, k / p) = x := by + intro ⟨x₁, x₂⟩ hx + rw [Prod.mk.injEq] + constructor <;> intro h + · simp only [mem_product, mem_range] at hx + have h' : x₁ < p := lt_of_lt_of_le hx.left $ mod_lt _ Fin.size_pos' + rw [h, add_mul_mod_self_left, add_mul_div_left _ _ Fin.size_pos', eq_comm (b := x₂)] + exact ⟨mod_eq_of_lt h', self_eq_add_left.mpr (div_eq_of_lt h')⟩ + · rw [← h.left, ← h.right, mod_add_div] + simp only [finset_sum_coeff, coeff_mul_natCast, coeff_X_pow, ite_mul, zero_mul, ← cast_mul] + rw [← sum_product', sum_congr rfl (fun a ha ↦ if_congr (h_iff a ha) rfl rfl), sum_ite_eq] + split_ifs with h + · simp + · rw [mem_product, mem_range, mem_range, not_and_or, lt_succ, not_le, not_lt] at h + cases h <;> simp [choose_eq_zero_of_lt (by tauto)] + +/-- For primes `p`, `choose n k` is congruent to `choose (n % p) (k % p) * choose (n / p) (k / p)` +modulo `p`. Also see `choose_modEq_choose_mod_mul_choose_div` for the version with `ZMOD`. -/ +theorem choose_modEq_choose_mod_mul_choose_div_nat : + choose n k ≡ choose (n % p) (k % p) * choose (n / p) (k / p) [MOD p] := by + rw [← Int.natCast_modEq_iff] + exact_mod_cast choose_modEq_choose_mod_mul_choose_div + +/-- For primes `p`, `choose n k` is congruent to the product of `choose (⌊n / p ^ i⌋ % p) +(⌊k / p ^ i⌋ % p)` over i < a, multiplied by `choose (⌊n / p ^ a⌋) (⌊k / p ^ a⌋)`, modulo `p`. -/ +theorem choose_modEq_choose_mul_prod_range_choose (a : ℕ) : + choose n k ≡ choose (n / p ^ a) (k / p ^ a) * + ∏ i in range a, choose (n / p ^ i % p) (k / p ^ i % p) [ZMOD p] := + match a with + | Nat.zero => by simp + | Nat.succ a => (choose_modEq_choose_mul_prod_range_choose a).trans <| by + rw [prod_range_succ, cast_mul, ← mul_assoc, mul_right_comm] + gcongr + apply choose_modEq_choose_mod_mul_choose_div.trans + simp_rw [pow_succ, Nat.div_div_eq_div_mul, mul_comm] + rfl + +/-- **Lucas's Theorem**: For primes `p`, `choose n k` is congruent to the product of +`choose (⌊n / p ^ i⌋ % p) (⌊k / p ^ i⌋ % p)` over `i` modulo `p`. -/ +theorem choose_modEq_prod_range_choose {a : ℕ} (ha₁ : n < p ^ a) (ha₂ : k < p ^ a) : + choose n k ≡ ∏ i in range a, choose (n / p ^ i % p) (k / p ^ i % p) [ZMOD p] := by + apply (choose_modEq_choose_mul_prod_range_choose a).trans + simp_rw [Nat.div_eq_of_lt ha₁, Nat.div_eq_of_lt ha₂, choose, cast_one, one_mul, cast_prod, + Int.ModEq.refl] + +/-- **Lucas's Theorem**: For primes `p`, `choose n k` is congruent to the product of +`choose (⌊n / p ^ i⌋ % p) (⌊k / p ^ i⌋ % p)` over `i` modulo `p`. -/ +theorem choose_modEq_prod_range_choose_nat {a : ℕ} (ha₁ : n < p ^ a) (ha₂ : k < p ^ a) : + choose n k ≡ ∏ i in range a, choose (n / p ^ i % p) (k / p ^ i % p) [MOD p] := by + rw [← Int.natCast_modEq_iff] + exact_mod_cast choose_modEq_prod_range_choose ha₁ ha₂ + +alias lucas_theorem := choose_modEq_prod_range_choose +alias lucas_theorem_nat := choose_modEq_prod_range_choose_nat + +end Choose diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index 452f3e4d893ce..d41da81f1a646 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller, Pim Otte -/ import Mathlib.Algebra.BigOperators.Fin +import Mathlib.Algebra.Order.Antidiag.Pi import Mathlib.Data.Nat.Choose.Sum import Mathlib.Data.Nat.Factorial.BigOperators import Mathlib.Data.Fin.VecNotation @@ -196,16 +197,58 @@ theorem multinomial_zero [DecidableEq α] : multinomial (0 : Multiset α) = 1 := end Multiset namespace Finset +open _root_.Nat /-! ### Multinomial theorem -/ -variable {α : Type*} [DecidableEq α] (s : Finset α) {R : Type*} +variable {α R : Type*} [DecidableEq α] -/-- The multinomial theorem +section Semiring +variable [Semiring R] - Proof is by induction on the number of summands. --/ -theorem sum_pow_of_commute [Semiring R] (x : α → R) +-- TODO: Can we prove one of the following two from the other one? +/-- The **multinomial theorem**. -/ +lemma sum_pow_eq_sum_piAntidiag_of_commute (s : Finset α) (f : α → R) + (hc : (s : Set α).Pairwise fun i j ↦ Commute (f i) (f j)) (n : ℕ) : + (∑ i in s, f i) ^ n = ∑ k in piAntidiag s n, multinomial s k * + s.noncommProd (fun i ↦ f i ^ k i) (hc.mono' fun i j h ↦ h.pow_pow ..) := by + classical + induction' s using Finset.cons_induction with a s has ih generalizing n + · cases n <;> simp + rw [Finset.sum_cons, piAntidiag_cons, sum_disjiUnion] + simp only [sum_map, Function.Embedding.coeFn_mk, Pi.add_apply, multinomial_cons, + Pi.add_apply, eq_self_iff_true, if_true, Nat.cast_mul, noncommProd_cons, eq_self_iff_true, + if_true, sum_add_distrib, sum_ite_eq', has, if_false, add_zero, + addLeftEmbedding_eq_addRightEmbedding, addRightEmbedding_apply] + suffices ∀ p : ℕ × ℕ, p ∈ antidiagonal n → + ∑ g in piAntidiag s p.2, ((g a + p.1 + s.sum g).choose (g a + p.1) : R) * + multinomial s (g + fun i ↦ ite (i = a) p.1 0) * + (f a ^ (g a + p.1) * s.noncommProd (fun i ↦ f i ^ (g i + ite (i = a) p.1 0)) + ((hc.mono (by simp)).mono' fun i j h ↦ h.pow_pow ..)) = + ∑ g in piAntidiag s p.2, n.choose p.1 * multinomial s g * (f a ^ p.1 * + s.noncommProd (fun i ↦ f i ^ g i) ((hc.mono (by simp)).mono' fun i j h ↦ h.pow_pow ..)) by + rw [sum_congr rfl this] + simp only [Nat.antidiagonal_eq_map, sum_map, Function.Embedding.coeFn_mk] + rw [(Commute.sum_right _ _ _ fun i hi ↦ hc (by simp) (by simp [hi]) + (by simpa [eq_comm] using ne_of_mem_of_not_mem hi has)).add_pow] + simp only [ih (hc.mono (by simp)), sum_mul, mul_sum] + refine sum_congr rfl fun i _ ↦ sum_congr rfl fun g _ ↦ ?_ + rw [← Nat.cast_comm, (Nat.commute_cast (f a ^ i) _).left_comm, mul_assoc] + refine fun p hp ↦ sum_congr rfl fun f hf ↦ ?_ + rw [mem_piAntidiag] at hf + rw [not_imp_comm.1 (hf.2 _) has, zero_add, hf.1] + congr 2 + · rw [mem_antidiagonal.1 hp] + · rw [multinomial_congr] + intro t ht + rw [Pi.add_apply, if_neg, add_zero] + exact ne_of_mem_of_not_mem ht has + refine noncommProd_congr rfl (fun t ht ↦ ?_) _ + rw [if_neg, add_zero] + exact ne_of_mem_of_not_mem ht has + +/-- The **multinomial theorem**. -/ +theorem sum_pow_of_commute [Semiring R] (x : α → R) (s : Finset α) (hc : (s : Set α).Pairwise fun i j => Commute (x i) (x j)) : ∀ n, s.sum x ^ n = @@ -243,13 +286,23 @@ theorem sum_pow_of_commute [Semiring R] (x : α → R) rw [Multiset.card_replicate, Nat.cast_mul, mul_assoc, Nat.cast_comm] congr 1; simp_rw [← mul_assoc, Nat.cast_comm]; rfl +end Semiring + +section CommSemiring +variable [CommSemiring R] {f : α → R} {s : Finset α} + +lemma sum_pow_eq_sum_piAntidiag (s : Finset α) (f : α → R) (n : ℕ) : + (∑ i in s, f i) ^ n = ∑ k in piAntidiag s n, multinomial s k * ∏ i in s, f i ^ k i := by + simp_rw [← noncommProd_eq_prod] + rw [← sum_pow_eq_sum_piAntidiag_of_commute _ _ fun _ _ _ _ _ ↦ Commute.all ..] theorem sum_pow [CommSemiring R] (x : α → R) (n : ℕ) : s.sum x ^ n = ∑ k ∈ s.sym n, k.val.multinomial * (k.val.map x).prod := by conv_rhs => rw [← sum_coe_sort] - convert sum_pow_of_commute s x (fun _ _ _ _ _ => mul_comm _ _) n + convert sum_pow_of_commute x s (fun _ _ _ _ _ ↦ Commute.all ..) n rw [Multiset.noncommProd_eq_prod] +end CommSemiring end Finset namespace Sym diff --git a/Mathlib/Data/Nat/Factorization/Induction.lean b/Mathlib/Data/Nat/Factorization/Induction.lean index 17b5af579a075..1add41840d66f 100644 --- a/Mathlib/Data/Nat/Factorization/Induction.lean +++ b/Mathlib/Data/Nat/Factorization/Induction.lean @@ -74,8 +74,25 @@ def recOnMul {P : ℕ → Sort*} (h0 : P 0) (h1 : P 1) (hp : ∀ p, Prime p → | n + 1 => h _ _ (hp'' p n hp') (hp p hp') recOnPrimeCoprime h0 hp'' fun a b _ _ _ => h a b -/-! ## Lemmas on multiplicative functions -/ +lemma _root_.induction_on_primes {P : ℕ → Prop} (h₀ : P 0) (h₁ : P 1) + (h : ∀ p a : ℕ, p.Prime → P a → P (p * a)) : ∀ n, P n := by + refine recOnPrimePow h₀ h₁ ?_ + rintro a p n hp - - ha + induction' n with n ih + · simpa using ha + · rw [pow_succ', mul_assoc] + exact h _ _ hp ih + +lemma prime_composite_induction {P : ℕ → Prop} (zero : P 0) (one : P 1) + (prime : ∀ p : ℕ, p.Prime → P p) (composite : ∀ a, 2 ≤ a → P a → ∀ b, 2 ≤ b → P b → P (a * b)) + (n : ℕ) : P n := by + refine induction_on_primes zero one ?_ _ + rintro p (_ | _ | a) hp ha + · simpa + · simpa using prime _ hp + · exact composite _ hp.two_le (prime _ hp) _ a.one_lt_succ_succ ha +/-! ## Lemmas on multiplicative functions -/ /-- For any multiplicative function `f` with `f 1 = 1` and any `n ≠ 0`, we can evaluate `f n` by evaluating `f` at `p ^ k` over the factorization of `n` -/ diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index c778cd6013a44..10a0826c4731c 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -77,7 +77,7 @@ theorem prod_primeFactorsList : ∀ {n}, n ≠ 0 → List.prod (primeFactorsList Nat.mul_div_cancel' (minFac_dvd _)] theorem primeFactorsList_prime {p : ℕ} (hp : Nat.Prime p) : p.primeFactorsList = [p] := by - have : p = p - 2 + 2 := (tsub_eq_iff_eq_add_of_le hp.two_le).mp rfl + have : p = p - 2 + 2 := (Nat.sub_add_cancel hp.two_le).symm rw [this, primeFactorsList] simp only [Eq.symm this] have : Nat.minFac p = p := (Nat.prime_def_minFac.mp hp).2 diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index 7d6cafe617b7e..ffce815884fb3 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ +import Mathlib.Order.Lattice import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Algebra.Order.Ring.Nat -import Mathlib.Tactic.NthRewrite +import Mathlib.Algebra.Ring.Nat +import Mathlib.Init.Data.Nat.Lemmas /-! # Definitions and properties of `Nat.gcd`, `Nat.lcm`, and `Nat.coprime` @@ -109,7 +110,7 @@ theorem lcm_dvd_iff {m n k : ℕ} : lcm m n ∣ k ↔ m ∣ k ∧ n ∣ k := ⟨fun h => ⟨(dvd_lcm_left _ _).trans h, (dvd_lcm_right _ _).trans h⟩, and_imp.2 lcm_dvd⟩ theorem lcm_pos {m n : ℕ} : 0 < m → 0 < n → 0 < m.lcm n := by - simp_rw [pos_iff_ne_zero] + simp_rw [Nat.pos_iff_ne_zero] exact lcm_ne_zero theorem lcm_mul_left {m n k : ℕ} : (m * n).lcm (m * k) = m * n.lcm k := by @@ -269,7 +270,7 @@ theorem pow_dvd_pow_iff {a b n : ℕ} (n0 : n ≠ 0) : a ^ n ∣ b ^ n ↔ a ∣ · simp [eq_zero_of_gcd_eq_zero_right g0] rcases exists_coprime' g0 with ⟨g, a', b', g0', co, rfl, rfl⟩ rw [mul_pow, mul_pow] at h - replace h := Nat.dvd_of_mul_dvd_mul_right (pow_pos g0' _) h + replace h := Nat.dvd_of_mul_dvd_mul_right (Nat.pos_pow_of_pos _ g0') h have := pow_dvd_pow a' <| Nat.pos_of_ne_zero n0 rw [pow_one, (co.pow n n).eq_one_of_dvd h] at this simp [eq_one_of_dvd_one this] diff --git a/Mathlib/Data/Nat/Pairing.lean b/Mathlib/Data/Nat/Pairing.lean index 948df6ce3bad4..b33b6e5f5723c 100644 --- a/Mathlib/Data/Nat/Pairing.lean +++ b/Mathlib/Data/Nat/Pairing.lean @@ -79,9 +79,9 @@ theorem pair_eq_pair {a b c d : ℕ} : pair a b = pair c d ↔ a = c ∧ b = d : theorem unpair_lt {n : ℕ} (n1 : 1 ≤ n) : (unpair n).1 < n := by let s := sqrt n simp only [unpair, Nat.sub_le_iff_le_add] - by_cases h : n - s * s < s <;> simp [h] + by_cases h : n - s * s < s <;> simp only [h, ↓reduceIte] · exact lt_of_lt_of_le h (sqrt_le_self _) - · simp at h + · simp only [not_lt] at h have s0 : 0 < s := sqrt_pos.2 n1 exact lt_of_le_of_lt h (Nat.sub_lt n1 (Nat.mul_pos s0 s0)) @@ -97,8 +97,9 @@ theorem unpair_left_le : ∀ n : ℕ, (unpair n).1 ≤ n theorem left_le_pair (a b : ℕ) : a ≤ pair a b := by simpa using unpair_left_le (pair a b) theorem right_le_pair (a b : ℕ) : b ≤ pair a b := by - by_cases h : a < b <;> simp [pair, h] - exact le_trans (le_mul_self _) (Nat.le_add_right _ _) + by_cases h : a < b + · simpa [pair, h] using le_trans (le_mul_self _) (Nat.le_add_right _ _) + · simp [pair, h] theorem unpair_right_le (n : ℕ) : (unpair n).2 ≤ n := by simpa using right_le_pair n.unpair.1 n.unpair.2 @@ -117,10 +118,10 @@ theorem pair_lt_pair_left {a₁ a₂} (b) (h : a₁ < a₂) : pair a₁ b < pair · apply Nat.add_lt_add_right; assumption theorem pair_lt_pair_right (a) {b₁ b₂} (h : b₁ < b₂) : pair a b₁ < pair a b₂ := by - by_cases h₁ : a < b₁ <;> simp [pair, h₁, Nat.add_assoc] - · simp [pair, lt_trans h₁ h, h] - exact mul_self_lt_mul_self h - · by_cases h₂ : a < b₂ <;> simp [pair, h₂, h] + by_cases h₁ : a < b₁ + · simpa [pair, h₁, Nat.add_assoc, lt_trans h₁ h, h] using mul_self_lt_mul_self h + · simp only [pair, h₁, ↓reduceIte, Nat.add_assoc] + by_cases h₂ : a < b₂ <;> simp [pair, h₂, h] simp? at h₁ says simp only [not_lt] at h₁ rw [Nat.add_comm, Nat.add_comm _ a, Nat.add_assoc, Nat.add_lt_add_iff_left] rwa [Nat.add_comm, ← sqrt_lt, sqrt_add_eq] diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index ec0b3644d257b..1111f4009e1b7 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -212,7 +212,7 @@ theorem dom_of_le_natCast {x : PartENat} {y : ℕ} (h : x ≤ y) : x.Dom := by instance decidableLe (x y : PartENat) [Decidable x.Dom] [Decidable y.Dom] : Decidable (x ≤ y) := if hx : x.Dom then - decidable_of_decidable_of_iff (by rw [le_def]) + decidable_of_decidable_of_iff (le_def x y).symm else if hy : y.Dom then isFalse fun h => hx <| dom_of_le_of_dom h hy else isTrue ⟨fun h => (hy h).elim, fun h => (hy h).elim⟩ diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index d12c54509d76c..419ca7d32efae 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -3,11 +3,11 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.Ring.Int +import Mathlib.Order.Bounds.Basic import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Order.Bounds.Basic +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow /-! ## Notable Theorems @@ -39,7 +39,7 @@ theorem Prime.five_le_of_ne_two_of_ne_three {p : ℕ} (hp : p.Prime) (h_two : p | 2 => decide | 3 => decide | 4 => decide - | n + 5 => exact (h.not_le le_add_self).elim + | n + 5 => exact (h.not_le <| le_add_left ..).elim end @@ -113,7 +113,7 @@ theorem Prime.mod_two_eq_one_iff_ne_two {p : ℕ} [Fact p.Prime] : p % 2 = 1 ↔ simp at h theorem coprime_of_dvd' {m n : ℕ} (H : ∀ k, Prime k → k ∣ m → k ∣ n → k ∣ 1) : Coprime m n := - coprime_of_dvd fun k kp km kn => not_le_of_gt kp.one_lt <| le_of_dvd zero_lt_one <| H k kp km kn + coprime_of_dvd fun k kp km kn => not_le_of_gt kp.one_lt <| le_of_dvd Nat.one_pos <| H k kp km kn theorem Prime.dvd_iff_not_coprime {p n : ℕ} (pp : Prime p) : p ∣ n ↔ ¬Coprime p n := iff_not_comm.2 pp.coprime_iff_not_dvd @@ -294,7 +294,7 @@ theorem exists_pow_lt_factorial (c : ℕ) : ∃ n0 > 1, ∀ n ≥ n0, c ^ n < (n obtain ⟨d, rfl⟩ := Nat.exists_eq_add_of_le hn obtain (rfl | c0) := c.eq_zero_or_pos · simp [Nat.factorial_pos] - refine (Nat.le_mul_of_pos_right _ (pow_pos c0 d)).trans_lt ?_ + refine (Nat.le_mul_of_pos_right _ (Nat.pow_pos (n := d) c0)).trans_lt ?_ convert_to (c ^ 2) ^ (c ^ 2 + d + 1) < (c ^ 2 + (c ^ 2 + d + 1))! · rw [← pow_mul, ← pow_add] congr 1 diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 5ae0fe067754f..57010bbdfe78c 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.Ring.Parity import Mathlib.Data.Nat.GCD.Basic /-! @@ -22,7 +23,6 @@ This file deals with prime numbers: natural numbers `p ≥ 2` whose only divisor -/ - open Bool Subtype open Nat @@ -54,7 +54,7 @@ theorem Prime.pos {p : ℕ} (pp : Prime p) : 0 < p := theorem Prime.two_le : ∀ {p : ℕ}, Prime p → 2 ≤ p | 0, h => (not_prime_zero h).elim | 1, h => (not_prime_one h).elim - | _ + 2, _ => le_add_self + | _ + 2, _ => le_add_left 2 _ theorem Prime.one_lt {p : ℕ} : Prime p → 1 < p := Prime.two_le @@ -78,13 +78,12 @@ theorem Prime.eq_one_or_self_of_dvd {p : ℕ} (pp : p.Prime) (m : ℕ) (hm : m theorem prime_def_lt'' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p := by refine ⟨fun h => ⟨h.two_le, h.eq_one_or_self_of_dvd⟩, fun h => ?_⟩ - -- Porting note: needed to make ℕ explicit - have h1 := (@one_lt_two ℕ ..).trans_le h.1 + have h1 := Nat.one_lt_two.trans_le h.1 refine ⟨mt Nat.isUnit_iff.mp h1.ne', fun a b hab => ?_⟩ simp only [Nat.isUnit_iff] apply Or.imp_right _ (h.2 a _) · rintro rfl - rw [← mul_right_inj' (pos_of_gt h1).ne', ← hab, mul_one] + rw [← mul_right_inj' (zero_lt_of_lt h1).ne', ← hab, mul_one] · rw [hab] exact dvd_mul_right _ _ @@ -105,7 +104,7 @@ theorem prime_def_lt' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, 2 ≤ m → m < revert p2 decide · rfl - · exact (h le_add_self l).elim d⟩ + · exact (h (le_add_left _ _) l).elim d⟩ theorem prime_def_le_sqrt {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, 2 ≤ m → m ≤ sqrt p → ¬m ∣ p := prime_def_lt'.trans <| @@ -117,7 +116,7 @@ theorem prime_def_le_sqrt {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, 2 ≤ m → rcases le_total m k with mk | km · exact this mk m2 e · rw [mul_comm] at e - refine this km (lt_of_mul_lt_mul_right ?_ (zero_le m)) e + refine this km (Nat.lt_of_mul_lt_mul_right (a := m) ?_) e rwa [one_mul, ← e]⟩ theorem prime_of_coprime (n : ℕ) (h1 : 1 < n) (h : ∀ m < n, m ≠ 0 → n.Coprime m) : Prime n := by @@ -188,7 +187,7 @@ theorem Prime.eq_two_or_odd' {p : ℕ} (hp : Prime p) : p = 2 ∨ Odd p := section MinFac theorem minFac_lemma (n k : ℕ) (h : ¬n < k * k) : sqrt n - k < sqrt n + 2 - k := - (tsub_lt_tsub_iff_right <| le_sqrt.2 <| le_of_not_gt h).2 <| Nat.lt_add_of_pos_right (by decide) + (Nat.sub_lt_sub_right <| le_sqrt.2 <| le_of_not_gt h) <| Nat.lt_add_of_pos_right (by decide) /-- If `n < k * k`, then `minFacAux n k = n`, if `k | n`, then `minFacAux n k = k`. @@ -237,15 +236,16 @@ theorem minFacAux_has_prop {n : ℕ} (n2 : 2 ≤ n) : ∀ k i, k = 2 * i + 3 → (∀ m, 2 ≤ m → m ∣ n → k ≤ m) → minFacProp n (minFacAux n k) | k => fun i e a => by rw [minFacAux] - by_cases h : n < k * k <;> simp [h] + by_cases h : n < k * k · have pp : Prime n := prime_def_le_sqrt.2 ⟨n2, fun m m2 l d => not_lt_of_ge l <| lt_of_lt_of_le (sqrt_lt.2 h) (a m m2 d)⟩ - exact ⟨n2, dvd_rfl, fun m m2 d => le_of_eq ((dvd_prime_two_le pp m2).1 d).symm⟩ + simpa [h] using ⟨n2, dvd_rfl, fun m m2 d => le_of_eq ((dvd_prime_two_le pp m2).1 d).symm⟩ have k2 : 2 ≤ k := by subst e apply Nat.le_add_left - by_cases dk : k ∣ n <;> simp [dk] + simp only [h, ↓reduceIte] + by_cases dk : k ∣ n <;> simp only [dk, ↓reduceIte] · exact ⟨k2, dk, a⟩ · refine have := minFac_lemma n k h @@ -270,7 +270,7 @@ theorem minFac_has_prop {n : ℕ} (n1 : n ≠ 1) : minFacProp n (minFac n) := by revert n0 n1 rcases n with (_ | _ | _) <;> simp [succ_le_succ] simp only [minFac_eq, Nat.isUnit_iff] - by_cases d2 : 2 ∣ n <;> simp [d2] + by_cases d2 : 2 ∣ n <;> simp only [d2, ↓reduceIte] · exact ⟨le_rfl, d2, fun k k2 _ => k2⟩ · refine minFacAux_has_prop n2 3 0 rfl fun m m2 d => (Nat.eq_or_lt_of_le m2).resolve_left (mt ?_ d2) @@ -381,7 +381,7 @@ theorem minFac_eq_two_iff (n : ℕ) : minFac n = 2 ↔ 2 ∣ n := by have := le_antisymm (Nat.succ_le_of_lt lb) (Nat.lt_succ_iff.mp h') rw [eq_comm, Nat.minFac_eq_one_iff] at this subst this - exact not_lt_of_le (le_of_dvd zero_lt_one h) one_lt_two + exact not_lt_of_le (le_of_dvd lb h) h' theorem factors_lemma {k} : (k + 2) / minFac (k + 2) < k + 2 := div_lt_self (Nat.zero_lt_succ _) (minFac_prime (by diff --git a/Mathlib/Data/Nat/PrimeNormNum.lean b/Mathlib/Data/Nat/PrimeNormNum.lean deleted file mode 100644 index 46e521fdecf22..0000000000000 --- a/Mathlib/Data/Nat/PrimeNormNum.lean +++ /dev/null @@ -1,19 +0,0 @@ -/- -Copyright (c) 2015 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro --/ -import Mathlib.Data.Nat.Factors -import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Tactic.NormNum -import Mathlib.Tactic.NormNum.Prime - -/-! -# Primality prover - -This file provides a `norm_num` extension to prove that natural numbers are prime. - -Porting note: the sole purpose of this file is to mark it as "ported". -This file seems to be tripping up the porting dashboard. - --/ diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index d503d0c141506..5a25a60e4b644 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -157,7 +157,7 @@ theorem minSqFacProp_div (n) {k} (pk : Prime k) (dk : k ∣ n) (dkk : ¬k * k theorem minSqFacAux_has_prop {n : ℕ} (k) (n0 : 0 < n) (i) (e : k = 2 * i + 3) (ih : ∀ m, Prime m → m ∣ n → k ≤ m) : MinSqFacProp n (minSqFacAux n k) := by rw [minSqFacAux] - by_cases h : n < k * k <;> simp [h] + by_cases h : n < k * k <;> simp only [h, ↓reduceDIte] · refine squarefree_iff_prime_squarefree.2 fun p pp d => ?_ have := ih p pp (dvd_trans ⟨_, rfl⟩ d) have := Nat.mul_le_mul this this diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index cbb74d30c85ae..07e422e8a1a7f 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -1384,13 +1384,14 @@ theorem divMod_to_nat_aux {n d : PosNum} {q r : Num} (h₁ : (r : ℕ) + d * ((q apply Num.mem_ofZNum'.trans rw [← ZNum.to_int_inj, Num.cast_toZNum, Num.cast_sub', sub_eq_iff_eq_add, ← Int.natCast_inj] simp - cases' e : Num.ofZNum' (Num.sub' r (Num.pos d)) with r₂ <;> simp [divModAux] - · rw [two_mul] + cases' e : Num.ofZNum' (Num.sub' r (Num.pos d)) with r₂ + · rw [Num.cast_bit0, two_mul] refine ⟨h₁, lt_of_not_ge fun h => ?_⟩ cases' Nat.le.dest h with r₂ e' rw [← Num.to_of_nat r₂, add_comm] at e' cases e.symm.trans (this.2 e'.symm) · have := this.1 e + simp only [Num.cast_bit1] constructor · rwa [two_mul, add_comm _ 1, mul_add, mul_one, ← add_assoc, ← this] · rwa [this, two_mul, add_lt_add_iff_right] at h₂ @@ -1412,9 +1413,9 @@ theorem divMod_to_nat (d n : PosNum) : -- Porting note: `cases'` didn't rewrite at `this`, so `revert` & `intro` are required. revert IH; cases' divMod d n with q r; intro IH simp only [divMod] at IH ⊢ - apply divMod_to_nat_aux <;> simp - · rw [← two_mul, ← two_mul, mul_left_comm, ← mul_add, ← IH.1] - · exact IH.2 + apply divMod_to_nat_aux + · simp; rw [← two_mul, ← two_mul, mul_left_comm, ← mul_add, ← IH.1] + · simpa using IH.2 @[simp] theorem div'_to_nat (n d) : (div' n d : ℕ) = n / d := diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index 844a3c59f0561..a51fcf32a77b9 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -1375,7 +1375,7 @@ theorem Valid'.map_aux {β} [Preorder β] {f : α → β} (f_strict_mono : Stric Valid' (Option.map f a₁) (map f t) (Option.map f a₂) ∧ (map f t).size = t.size := by induction t generalizing a₁ a₂ with | nil => - simp [map]; apply valid'_nil + simp only [map, size_nil, and_true]; apply valid'_nil cases a₁; · trivial cases a₂; · trivial simp only [Bounded] @@ -1408,7 +1408,7 @@ theorem Valid'.erase_aux [@DecidableRel α (· ≤ ·)] (x : α) {t a₁ a₂} ( Valid' a₁ (erase x t) a₂ ∧ Raised (erase x t).size t.size := by induction t generalizing a₁ a₂ with | nil => - simp [erase, Raised]; exact h + simpa [erase, Raised] | node _ t_l t_x t_r t_ih_l t_ih_r => simp only [erase, size_node] have t_ih_l' := t_ih_l h.left diff --git a/Mathlib/Data/PFun.lean b/Mathlib/Data/PFun.lean index c3445adadf513..8ce516786fae8 100644 --- a/Mathlib/Data/PFun.lean +++ b/Mathlib/Data/PFun.lean @@ -246,7 +246,7 @@ theorem mem_fix_iff {f : α →. β ⊕ α} {a : α} {b : β} : injection h₂.symm.trans e with h; simp [h] next e => injection h₂.symm.trans e - · simp [fix] at h₃ + · simp only [fix, Part.mem_assert_iff] at h₃ cases' h₃ with h₃ h₄ refine ⟨⟨_, fun y h' => ?_⟩, ?_⟩ · injection Part.mem_unique h h' with e diff --git a/Mathlib/Data/QPF/Multivariate/Basic.lean b/Mathlib/Data/QPF/Multivariate/Basic.lean index 1ede4e58b3cfe..a0d0fd66952f8 100644 --- a/Mathlib/Data/QPF/Multivariate/Basic.lean +++ b/Mathlib/Data/QPF/Multivariate/Basic.lean @@ -67,10 +67,13 @@ matched because they preserve the properties of QPF. The latter example, * Const each proves that some operations on functors preserves the QPF structure +-/ +set_option linter.longLine false in +/-! ## Reference -! * [Jeremy Avigad, Mario M. Carneiro and Simon Hudon, *Data Types as Quotients of Polynomial Functors*][avigad-carneiro-hudon2019] +[Jeremy Avigad, Mario M. Carneiro and Simon Hudon, *Data Types as Quotients of Polynomial Functors*][avigad-carneiro-hudon2019] -/ diff --git a/Mathlib/Data/QPF/Univariate/Basic.lean b/Mathlib/Data/QPF/Univariate/Basic.lean index abb30f1569c66..3f3dbf6911b6b 100644 --- a/Mathlib/Data/QPF/Univariate/Basic.lean +++ b/Mathlib/Data/QPF/Univariate/Basic.lean @@ -242,9 +242,7 @@ def Fix.dest : Fix F → F (Fix F) := theorem Fix.rec_eq {α : Type _} (g : F α → α) (x : F (Fix F)) : Fix.rec g (Fix.mk x) = g (Fix.rec g <$> x) := by have : recF g ∘ fixToW = Fix.rec g := by - apply funext - apply Quotient.ind - intro x + ext ⟨x⟩ apply recF_eq_of_Wequiv rw [fixToW] apply Wrepr_equiv @@ -270,8 +268,7 @@ theorem Fix.ind_aux (a : q.P.A) (f : q.P.B a → q.P.W) : theorem Fix.ind_rec {α : Type u} (g₁ g₂ : Fix F → α) (h : ∀ x : F (Fix F), g₁ <$> x = g₂ <$> x → g₁ (Fix.mk x) = g₂ (Fix.mk x)) : ∀ x, g₁ x = g₂ x := by - apply Quot.ind - intro x + rintro ⟨x⟩ induction' x with a f ih change g₁ ⟦⟨a, f⟩⟧ = g₂ ⟦⟨a, f⟩⟧ rw [← Fix.ind_aux a f]; apply h @@ -303,8 +300,7 @@ theorem Fix.dest_mk (x : F (Fix F)) : Fix.dest (Fix.mk x) = x := by apply Fix.mk_dest theorem Fix.ind (p : Fix F → Prop) (h : ∀ x : F (Fix F), Liftp p x → p (Fix.mk x)) : ∀ x, p x := by - apply Quot.ind - intro x + rintro ⟨x⟩ induction' x with a f ih change p ⟦⟨a, f⟩⟧ rw [← Fix.ind_aux a f] @@ -374,17 +370,10 @@ theorem Cofix.dest_corec {α : Type u} (g : α → F α) (x : α) : dsimp rw [corecF_eq, abs_map, abs_repr, ← comp_map]; rfl --- Porting note: Needed to add `(motive := _)` to get `Quot.inductionOn` to work private theorem Cofix.bisim_aux (r : Cofix F → Cofix F → Prop) (h' : ∀ x, r x x) (h : ∀ x y, r x y → Quot.mk r <$> Cofix.dest x = Quot.mk r <$> Cofix.dest y) : ∀ x y, r x y → x = y := by - intro x - apply Quot.inductionOn (motive := _) x - clear x - intro x y - apply Quot.inductionOn (motive := _) y - clear y - intro y rxy + rintro ⟨x⟩ ⟨y⟩ rxy apply Quot.sound let r' x y := r (Quot.mk _ x) (Quot.mk _ y) have : IsPrecongr r' := by @@ -400,11 +389,9 @@ private theorem Cofix.bisim_aux (r : Cofix F → Cofix F → Prop) (h' : ∀ x, rw [Quot.sound cuv] apply h' let f : Quot r → Quot r' := - Quot.lift (Quot.lift (Quot.mk r') h₁) - (by - intro c; apply Quot.inductionOn (motive := _) c; clear c - intro c d; apply Quot.inductionOn (motive := _) d; clear d - intro d rcd; apply Quot.sound; apply rcd) + Quot.lift (Quot.lift (Quot.mk r') h₁) <| by + rintro ⟨c⟩ ⟨d⟩ rcd + exact Quot.sound rcd have : f ∘ Quot.mk r ∘ Quot.mk Mcongr = Quot.mk r' := rfl rw [← this, ← PFunctor.map_map _ _ f, ← PFunctor.map_map _ _ (Quot.mk r), abs_map, abs_map, abs_map, h₀] diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index c391a77cc73e6..94891a14ca0a8 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -517,7 +517,9 @@ noncomputable instance instLinearOrderedField : LinearOrderedField ℝ where exact CauSeq.Completion.inv_mul_cancel h inv_zero := by simp [← ofCauchy_zero, ← ofCauchy_inv] nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q a => rfl nnratCast_def q := by rw [← ofCauchy_nnratCast, NNRat.cast_def, ofCauchy_div, ofCauchy_natCast, ofCauchy_natCast] ratCast_def q := by diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index 55bf8698c775d..14e0942f966a5 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -66,8 +66,9 @@ theorem cantorFunctionAux_false (h : f n = false) : cantorFunctionAux c f n = 0 simp [cantorFunctionAux, h] theorem cantorFunctionAux_nonneg (h : 0 ≤ c) : 0 ≤ cantorFunctionAux c f n := by - cases h' : f n <;> simp [h'] - apply pow_nonneg h + cases h' : f n + · simp [h'] + · simpa [h'] using pow_nonneg h _ theorem cantorFunctionAux_eq (h : f n = g n) : cantorFunctionAux c f n = cantorFunctionAux c g n := by simp [cantorFunctionAux, h] diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index acfc3f3a4dba0..214fffe15846f 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -70,7 +70,7 @@ theorem one_div_nonneg : 0 ≤ 1 / p := le_of_lt h.one_div_pos theorem one_div_ne_zero : 1 / p ≠ 0 := ne_of_gt h.one_div_pos -theorem conj_eq : q = p / (p - 1) := by +theorem conj_eq (h : p.IsConjExponent q) : q = p / (p - 1) := by have := h.inv_add_inv_conj rw [← eq_sub_iff_add_eq', inv_eq_iff_eq_inv] at this field_simp [this, h.ne_zero] @@ -86,7 +86,7 @@ theorem sub_one_mul_conj : (p - 1) * q = p := theorem mul_eq_add : p * q = p + q := by simpa only [sub_mul, sub_eq_iff_eq_add, one_mul] using h.sub_one_mul_conj -@[symm] protected lemma symm : q.IsConjExponent p where +@[symm] protected lemma symm (h : p.IsConjExponent q) : q.IsConjExponent p where one_lt := by simpa only [h.conj_eq] using (one_lt_div h.sub_one_pos).mpr (sub_one_lt p) inv_add_inv_conj := by simpa [add_comm] using h.inv_add_inv_conj @@ -159,7 +159,7 @@ lemma inv_ne_zero : p⁻¹ ≠ 0 := h.inv_pos.ne' lemma one_sub_inv : 1 - p⁻¹ = q⁻¹ := tsub_eq_of_eq_add_rev h.inv_add_inv_conj.symm -lemma conj_eq : q = p / (p - 1) := by +lemma conj_eq (h : p.IsConjExponent q) : q = p / (p - 1) := by simpa only [← coe_one, ← NNReal.coe_sub h.one_le, ← NNReal.coe_div, coe_inj] using h.coe.conj_eq lemma conjExponent_eq : conjExponent p = q := h.conj_eq.symm @@ -171,7 +171,7 @@ lemma mul_eq_add : p * q = p + q := by simpa only [← NNReal.coe_mul, ← NNReal.coe_add, NNReal.coe_inj] using h.coe.mul_eq_add @[symm] -protected lemma symm : q.IsConjExponent p where +protected lemma symm (h : p.IsConjExponent q) : q.IsConjExponent p where one_lt := by rw [h.conj_eq] exact (one_lt_div h.sub_one_pos).mpr (tsub_lt_self h.pos zero_lt_one) diff --git a/Mathlib/Data/Semiquot.lean b/Mathlib/Data/Semiquot.lean index 176ecb385af6e..5557146ba9262 100644 --- a/Mathlib/Data/Semiquot.lean +++ b/Mathlib/Data/Semiquot.lean @@ -181,7 +181,7 @@ theorem get_mem {q : Semiquot α} (p) : get q p ∈ q := by unfold get; rw [liftOn_ofMem q _ _ a h]; exact h theorem eq_pure {q : Semiquot α} (p) : q = pure (get q p) := - ext.2 fun a => by simp; exact ⟨fun h => p _ h _ (get_mem _), fun e => e.symm ▸ get_mem _⟩ + ext.2 fun a => by simpa using ⟨fun h => p _ h _ (get_mem _), fun e => e.symm ▸ get_mem _⟩ @[simp] theorem pure_isPure (a : α) : IsPure (pure a) @@ -198,7 +198,7 @@ theorem IsPure.mono {s t : Semiquot α} (st : s ≤ t) (h : IsPure t) : IsPure s theorem IsPure.min {s t : Semiquot α} (h : IsPure t) : s ≤ t ↔ s = t := ⟨fun st => le_antisymm st <| by - rw [eq_pure h, eq_pure (h.mono st)]; simp; exact h _ (get_mem _) _ (st <| get_mem _), + rw [eq_pure h, eq_pure (h.mono st)]; simpa using h _ (get_mem _) _ (st <| get_mem _), le_of_eq⟩ theorem isPure_of_subsingleton [Subsingleton α] (q : Semiquot α) : IsPure q diff --git a/Mathlib/Data/Seq/Computation.lean b/Mathlib/Data/Seq/Computation.lean index 81482d9f15cc5..93868db28c7bc 100644 --- a/Mathlib/Data/Seq/Computation.lean +++ b/Mathlib/Data/Seq/Computation.lean @@ -623,8 +623,9 @@ theorem bind_pure (f : α → β) (s) : bind s (pure ∘ f) = map f s := by match c₁, c₂, h with | _, c₂, Or.inl (Eq.refl _) => cases' destruct c₂ with b cb <;> simp | _, _, Or.inr ⟨s, rfl, rfl⟩ => - apply recOn s <;> intro s <;> simp - exact Or.inr ⟨s, rfl, rfl⟩ + apply recOn s <;> intro s + · simp + · simpa using Or.inr ⟨s, rfl, rfl⟩ · exact Or.inr ⟨s, rfl, rfl⟩ -- Porting note: used to use `rw [bind_pure]` @@ -648,11 +649,11 @@ theorem bind_assoc (s : Computation α) (f : α → Computation β) (g : β → match c₁, c₂, h with | _, c₂, Or.inl (Eq.refl _) => cases' destruct c₂ with b cb <;> simp | _, _, Or.inr ⟨s, rfl, rfl⟩ => - apply recOn s <;> intro s <;> simp - · generalize f s = fs + apply recOn s <;> intro s + · simp only [BisimO, ret_bind]; generalize f s = fs apply recOn fs <;> intro t <;> simp · cases' destruct (g t) with b cb <;> simp - · exact Or.inr ⟨s, rfl, rfl⟩ + · simpa [BisimO] using Or.inr ⟨s, rfl, rfl⟩ · exact Or.inr ⟨s, rfl, rfl⟩ theorem results_bind {s : Computation α} {f : α → Computation β} {a b m n} (h1 : Results s a m) @@ -1095,8 +1096,9 @@ theorem LiftRelRec.lem {R : α → β → Prop} (C : Computation α → Computat simp [h] · simp only [liftRel_think_left] revert h - apply cb.recOn (fun b => _) fun cb' => _ <;> intros _ h <;> simp at h <;> simp [h] - exact IH _ h + apply cb.recOn (fun b => _) fun cb' => _ <;> intros _ h + · simpa using h + · simpa [h] using IH _ h theorem liftRel_rec {R : α → β → Prop} (C : Computation α → Computation β → Prop) (H : ∀ {ca cb}, C ca cb → LiftRelAux R C (destruct ca) (destruct cb)) (ca cb) (Hc : C ca cb) : diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index 337d34009054e..974c780ddf181 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -401,30 +401,39 @@ def ofStream (s : Stream' α) : Seq α := instance coeStream : Coe (Stream' α) (Seq α) := ⟨ofStream⟩ -/-- Embed a `LazyList α` as a sequence. Note that even though this +section MLList + +/-- Embed a `MLList α` as a sequence. Note that even though this is non-meta, it will produce infinite sequences if used with - cyclic `LazyList`s created by meta constructions. -/ -def ofLazyList : MLList Id α → Seq α := + cyclic `MLList`s created by meta constructions. -/ +def ofMLList : MLList Id α → Seq α := corec fun l => match l.uncons with | .none => none | .some (a, l') => some (a, l') -instance coeLazyList : Coe (MLList Id α) (Seq α) := - ⟨ofLazyList⟩ +@[deprecated (since := "2024-07-26")] alias ofLazyList := ofMLList + +instance coeMLList : Coe (MLList Id α) (Seq α) := + ⟨ofMLList⟩ + +@[deprecated (since := "2024-07-26")] alias coeLazyList := coeMLList -/-- Translate a sequence into a `LazyList`. Since `LazyList` and `List` - are isomorphic as non-meta types, this function is necessarily meta. -/ -unsafe def toLazyList : Seq α → MLList Id α +/-- Translate a sequence into a `MLList`. -/ +unsafe def toMLList : Seq α → MLList Id α | s => match destruct s with | none => .nil - | some (a, s') => .cons a (toLazyList s') + | some (a, s') => .cons a (toMLList s') + +@[deprecated (since := "2024-07-26")] alias toLazyList := toMLList + +end MLList /-- Translate a sequence to a list. This function will run forever if run on an infinite sequence. -/ unsafe def forceToList (s : Seq α) : List α := - (toLazyList s).force + (toMLList s).force /-- The sequence of natural numbers some 0, some 1, ... -/ def nats : Seq ℕ := diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 860f1f9c0de85..9d86f9ce4e599 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -602,9 +602,9 @@ theorem destruct_flatten (c : Computation (WSeq α)) : destruct (flatten c) = c match c1, c2, h with | c, _, Or.inl rfl => by cases c.destruct <;> simp | _, _, Or.inr ⟨c, rfl, rfl⟩ => by - induction' c using Computation.recOn with a c' <;> simp - · cases (destruct a).destruct <;> simp - · exact Or.inr ⟨c', rfl, rfl⟩ + induction' c using Computation.recOn with a c' + · simp; cases (destruct a).destruct <;> simp + · simpa using Or.inr ⟨c', rfl, rfl⟩ theorem head_terminates_iff (s : WSeq α) : Terminates (head s) ↔ Terminates (destruct s) := terminates_map_iff _ (destruct s) @@ -1174,7 +1174,9 @@ theorem toList_nil : toList (nil : WSeq α) = Computation.pure [] := destruct_eq_pure rfl theorem toList_ofList (l : List α) : l ∈ toList (ofList l) := by - induction' l with a l IH <;> simp [ret_mem]; exact think_mem (Computation.mem_map _ IH) + induction' l with a l IH + · simp [ret_mem] + · simpa [ret_mem] using think_mem (Computation.mem_map _ IH) @[simp] theorem destruct_ofSeq (s : Seq α) : diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 7c12ff8f89a07..740784235067c 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -453,6 +453,7 @@ theorem setOf_false : { _a : α | False } = ∅ := theorem empty_subset (s : Set α) : ∅ ⊆ s := nofun +@[simp] theorem subset_empty_iff {s : Set α} : s ⊆ ∅ ↔ s = ∅ := (Subset.antisymm_iff.trans <| and_iff_left (empty_subset _)).symm diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index 5bb4138b4840f..497172e356e0d 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -340,6 +340,13 @@ instance fintypeBiUnion' [DecidableEq α] {ι : Type*} (s : Set ι) [Fintype s] [∀ i, Fintype (t i)] : Fintype (⋃ x ∈ s, t x) := Fintype.ofFinset (s.toFinset.biUnion fun x => (t x).toFinset) <| by simp +lemma toFinset_iUnion [Fintype β] [DecidableEq α] (f : β → Set α) + [∀ w, Fintype (f w)] : + Set.toFinset (⋃ (x : β), f x) = + Finset.biUnion (Finset.univ : Finset β) (fun x => (f x).toFinset) := by + ext v + simp only [mem_toFinset, mem_iUnion, Finset.mem_biUnion, Finset.mem_univ, true_and] + section monad attribute [local instance] Set.monad @@ -761,6 +768,11 @@ theorem Finite.of_preimage (h : (f ⁻¹' s).Finite) (hf : Surjective f) : s.Fin theorem Finite.preimage (I : Set.InjOn f (f ⁻¹' s)) (h : s.Finite) : (f ⁻¹' s).Finite := (h.subset (image_preimage_subset f s)).of_finite_image I +theorem Finite.preimage' (h : s.Finite) (hf : ∀ b ∈ s, (f ⁻¹' {b}).Finite) : + (f ⁻¹' s).Finite := by + rw [← Set.biUnion_preimage_singleton] + exact Set.Finite.biUnion h hf + protected lemma Infinite.preimage (hs : s.Infinite) (hf : s ⊆ range f) : (f ⁻¹' s).Infinite := fun h ↦ hs <| finite_of_finite_preimage h hf @@ -1172,10 +1184,10 @@ theorem Infinite.exists_subset_card_eq {s : Set α} (hs : s.Infinite) (n : ℕ) ⟨((Finset.range n).map (hs.natEmbedding _)).map (Embedding.subtype _), by simp⟩ theorem infinite_of_finite_compl [Infinite α] {s : Set α} (hs : sᶜ.Finite) : s.Infinite := fun h => - Set.infinite_univ (by simpa using hs.union h) + Set.infinite_univ (α := α) (by simpa using hs.union h) theorem Finite.infinite_compl [Infinite α] {s : Set α} (hs : s.Finite) : sᶜ.Infinite := fun h => - Set.infinite_univ (by simpa using hs.union h) + Set.infinite_univ (α := α) (by simpa using hs.union h) theorem Infinite.diff {s t : Set α} (hs : s.Infinite) (ht : t.Finite) : (s \ t).Infinite := fun h => hs <| h.of_diff ht @@ -1271,7 +1283,7 @@ theorem infinite_of_forall_exists_gt (h : ∀ a, ∃ b ∈ s, a < b) : s.Infinit (strictMono_nat_of_lt_succ fun n => (h _).choose_spec.2).injective hf theorem infinite_of_forall_exists_lt (h : ∀ a, ∃ b ∈ s, b < a) : s.Infinite := - @infinite_of_forall_exists_gt αᵒᵈ _ _ _ h + infinite_of_forall_exists_gt (α := αᵒᵈ) h end Preorder @@ -1461,12 +1473,12 @@ variable [Preorder α] [IsDirected α (· ≥ ·)] [Nonempty α] {s : Set α} /-- A finite set is bounded below. -/ protected theorem Finite.bddBelow (hs : s.Finite) : BddBelow s := - @Finite.bddAbove αᵒᵈ _ _ _ _ hs + Finite.bddAbove (α := αᵒᵈ) hs /-- A finite union of sets which are all bounded below is still bounded below. -/ theorem Finite.bddBelow_biUnion {I : Set β} {S : β → Set α} (H : I.Finite) : BddBelow (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, BddBelow (S i) := - @Finite.bddAbove_biUnion αᵒᵈ _ _ _ _ _ _ H + Finite.bddAbove_biUnion (α := αᵒᵈ) H theorem infinite_of_not_bddBelow : ¬BddBelow s → s.Infinite := mt Finite.bddBelow diff --git a/Mathlib/Data/Set/List.lean b/Mathlib/Data/Set/List.lean index 430a15171d4c3..bbd932e0d0918 100644 --- a/Mathlib/Data/Set/List.lean +++ b/Mathlib/Data/Set/List.lean @@ -20,10 +20,10 @@ variable {α β : Type*} (l : List α) namespace Set theorem range_list_map (f : α → β) : range (map f) = { l | ∀ x ∈ l, x ∈ range f } := by - refine antisymm (range_subset_iff.2 fun l => forall_mem_map_iff.2 fun y _ => mem_range_self _) + refine antisymm (range_subset_iff.2 fun l => forall_mem_map.2 fun y _ => mem_range_self _) fun l hl => ?_ induction' l with a l ihl; · exact ⟨[], rfl⟩ - rcases ihl fun x hx => hl x <| subset_cons _ _ hx with ⟨l, rfl⟩ + rcases ihl fun x hx => hl x <| subset_cons_self _ _ hx with ⟨l, rfl⟩ rcases hl a (mem_cons_self _ _) with ⟨a, rfl⟩ exact ⟨a :: l, map_cons _ _ _⟩ diff --git a/Mathlib/Data/Set/Sigma.lean b/Mathlib/Data/Set/Sigma.lean index 3ccad0f9ba268..a1805649dbc09 100644 --- a/Mathlib/Data/Set/Sigma.lean +++ b/Mathlib/Data/Set/Sigma.lean @@ -78,8 +78,8 @@ theorem univ_sigma_univ : (@univ ι).sigma (fun _ ↦ @univ (α i)) = univ := ex theorem sigma_univ : s.sigma (fun _ ↦ univ : ∀ i, Set (α i)) = Sigma.fst ⁻¹' s := ext fun _ ↦ and_true_iff _ -@[simp] theorem univ_sigma_preimage (s : Set (Σ i, α i)) : - (@univ ι).sigma (fun i ↦ Sigma.mk i ⁻¹' s) = s := +@[simp] theorem univ_sigma_preimage_mk (s : Set (Σ i, α i)) : + (univ : Set ι).sigma (fun i ↦ Sigma.mk i ⁻¹' s) = s := ext <| by simp @[simp] diff --git a/Mathlib/Data/Set/UnionLift.lean b/Mathlib/Data/Set/UnionLift.lean index 9b8184a95693c..f673e2f82b811 100644 --- a/Mathlib/Data/Set/UnionLift.lean +++ b/Mathlib/Data/Set/UnionLift.lean @@ -104,6 +104,7 @@ theorem iUnionLift_unary (u : T → T) (ui : ∀ i, S i → S i) Set.inclusion (show S i ⊆ T from hT'.symm ▸ Set.subset_iUnion S i) (ui i x)) (uβ : β → β) (h : ∀ (i) (x : S i), f i (ui i x) = uβ (f i x)) (x : T) : iUnionLift S f hf T (le_of_eq hT') (u x) = uβ (iUnionLift S f hf T (le_of_eq hT') x) := by + clear hT -- this prevents the argument from getting inserted by accident. subst hT' cases' Set.mem_iUnion.1 x.prop with i hi rw [iUnionLift_of_mem x hi, ← h i] @@ -125,6 +126,7 @@ theorem iUnionLift_binary (dir : Directed (· ≤ ·) S) (op : T → T → T) (o (opβ : β → β → β) (h : ∀ (i) (x y : S i), f i (opi i x y) = opβ (f i x) (f i y)) (x y : T) : iUnionLift S f hf T (le_of_eq hT') (op x y) = opβ (iUnionLift S f hf T (le_of_eq hT') x) (iUnionLift S f hf T (le_of_eq hT') y) := by + clear hT -- this prevents the argument from getting inserted by accident. subst hT' cases' Set.mem_iUnion.1 x.prop with i hi cases' Set.mem_iUnion.1 y.prop with j hj diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index 2eb6db825cb15..0d11a1762c559 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -233,13 +233,12 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : match a₁, a₂, b₁, b₂, h₁, h₂ with | _, _, _, _, Eq.refl _, HEq.refl _ => rfl - @[simp] theorem «forall» {p : (Σ'a, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ @[simp] -theorem «exists» {p : (Σ'a, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := +theorem exists' {p : (Σ'a, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ /-- A specialized ext lemma for equality of `PSigma` types over an indexed subtype. -/ diff --git a/Mathlib/Data/String/Defs.lean b/Mathlib/Data/String/Defs.lean index 9afdc68d3848f..3e121b543fc6f 100644 --- a/Mathlib/Data/String/Defs.lean +++ b/Mathlib/Data/String/Defs.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Keeley Hoek, Floris van Doorn, Chris Bailey -/ -import Batteries.Data.List.Basic import Batteries.Data.String.Basic +import Batteries.Tactic.Alias /-! # Definitions for `String` diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index 6dd91108e0520..48aaa0d08f165 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -37,28 +37,16 @@ initialize_simps_projections Subtype (val → coe) theorem prop (x : Subtype p) : p x := x.2 -@[simp] -protected theorem «forall» {q : { a // p a } → Prop} : (∀ x, q x) ↔ ∀ a b, q ⟨a, b⟩ := - ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ - /-- An alternative version of `Subtype.forall`. This one is useful if Lean cannot figure out `q` when using `Subtype.forall` from right to left. -/ protected theorem forall' {q : ∀ x, p x → Prop} : (∀ x h, q x h) ↔ ∀ x : { a // p a }, q x x.2 := (@Subtype.forall _ _ fun x ↦ q x.1 x.2).symm -@[simp] -protected theorem «exists» {q : { a // p a } → Prop} : (∃ x, q x) ↔ ∃ a b, q ⟨a, b⟩ := - ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ - /-- An alternative version of `Subtype.exists`. This one is useful if Lean cannot figure out `q` when using `Subtype.exists` from right to left. -/ protected theorem exists' {q : ∀ x, p x → Prop} : (∃ x h, q x h) ↔ ∃ x : { a // p a }, q x x.2 := (@Subtype.exists _ _ fun x ↦ q x.1 x.2).symm -@[ext] -protected theorem ext : ∀ {a1 a2 : { x // p x }}, (a1 : α) = (a2 : α) → a1 = a2 - | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl - theorem heq_iff_coe_eq (h : ∀ x, p x ↔ q x) {a1 : { x // p x }} {a2 : { x // q x }} : HEq a1 a2 ↔ (a1 : α) = (a2 : α) := Eq.rec diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index d877d71c9d408..aa85d9183ee69 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -147,10 +147,20 @@ instance : Membership α (Sym α n) := instance decidableMem [DecidableEq α] (a : α) (s : Sym α n) : Decidable (a ∈ s) := s.1.decidableMem _ +@[simp, norm_cast] lemma coe_mk (s : Multiset α) (h : Multiset.card s = n) : mk s h = s := rfl + @[simp] theorem mem_mk (a : α) (s : Multiset α) (h : Multiset.card s = n) : a ∈ mk s h ↔ a ∈ s := Iff.rfl +lemma «forall» {p : Sym α n → Prop} : + (∀ s : Sym α n, p s) ↔ ∀ (s : Multiset α) (hs : Multiset.card s = n), p (Sym.mk s hs) := by + simp [Sym] + +lemma «exists» {p : Sym α n → Prop} : + (∃ s : Sym α n, p s) ↔ ∃ (s : Multiset α) (hs : Multiset.card s = n), p (Sym.mk s hs) := by + simp [Sym] + @[simp] theorem not_mem_nil (a : α) : ¬ a ∈ (nil : Sym α 0) := Multiset.not_mem_zero a @@ -453,6 +463,19 @@ theorem coe_append (s : Sym α n) (s' : Sym α n') : (s.append s' : Multiset α) theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ s' := Multiset.mem_add +/-- `a ↦ {a}` as an equivalence between `α` and `Sym α 1`. -/ +@[simps apply] +def oneEquiv : α ≃ Sym α 1 where + toFun a := ⟨{a}, by simp⟩ + invFun s := (Equiv.subtypeQuotientEquivQuotientSubtype + (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s).liftOn + (fun l ↦ l.1.head <| List.length_pos.mp <| by simp) + fun ⟨_, _⟩ ⟨_, h⟩ ↦ fun perm ↦ by + obtain ⟨a, rfl⟩ := List.length_eq_one.mp h + exact List.eq_of_mem_singleton (perm.mem_iff.mp <| List.head_mem _) + left_inv a := by rfl + right_inv := by rintro ⟨⟨l⟩, h⟩; obtain ⟨a, rfl⟩ := List.length_eq_one.mp h; rfl + /-- Fill a term `m : Sym α (n - i)` with `i` copies of `a` to obtain a term of `Sym α n`. This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using `Sym.cast`. -/ @@ -499,7 +522,7 @@ theorem fill_filterNe [DecidableEq α] (a : α) (m : Sym α n) : rw [count_add, count_filter, Sym.coe_replicate, count_replicate] obtain rfl | h := eq_or_ne a b · rw [if_pos rfl, if_neg (not_not.2 rfl), zero_add] - · rw [if_pos h, if_neg h.symm, add_zero]) + · rw [if_pos h, if_neg h, add_zero]) theorem filter_ne_fill [DecidableEq α] (a : α) (m : Σi : Fin (n + 1), Sym α (n - i)) (h : a ∉ m.2) : (m.2.fill a m.1).filterNe a = m := diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index c9e5e3bf53b00..1a377872c5e33 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -657,7 +657,7 @@ theorem subtypeVal_diagSub {α : TypeVec n} : subtypeVal (repeatEq α) ⊚ diagS ext i x induction' i with _ _ _ i_ih · simp [comp, diagSub, subtypeVal, prod.diag] - · simp [prod.diag] + · simp only [prod.diag] simp? [comp, diagSub, subtypeVal, prod.diag] at * says simp only [comp, subtypeVal, diagSub] at * apply @i_ih (drop _) diff --git a/Mathlib/Data/Vector/Zip.lean b/Mathlib/Data/Vector/Zip.lean index 5074615cb13be..eee02c2029439 100644 --- a/Mathlib/Data/Vector/Zip.lean +++ b/Mathlib/Data/Vector/Zip.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Data.Vector.Basic -import Mathlib.Data.List.Zip /-! # The `zipWith` operation on vectors. diff --git a/Mathlib/Data/Vector3.lean b/Mathlib/Data/Vector3.lean index 6d040f56b11d9..26d0da1c9b371 100644 --- a/Mathlib/Data/Vector3.lean +++ b/Mathlib/Data/Vector3.lean @@ -181,11 +181,10 @@ theorem append_insert (a : α) (t : Vector3 α m) (v : Vector3 α n) (i : Fin2 ( insert a (b :: t +-+ v) (Eq.recOn (congr_arg (· + 1) e' : _ + 1 = _) (fs (add i k))) = Eq.recOn (congr_arg (· + 1) e' : _ + 1 = _) (b :: t +-+ insert a v i) - rw [← - (Eq.recOn e' rfl : + rw [← (Eq.recOn e' rfl : fs (Eq.recOn e' (i.add k) : Fin2 ((n + k) + 1)) = Eq.recOn (congr_arg (· + 1) e' : _ + 1 = _) (fs (i.add k)))] - simp; rw [IH]; exact Eq.recOn e' rfl + simpa [IH] using Eq.recOn e' rfl end Vector3 diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 38b38fc33116a..0ac9d27dbe010 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -508,6 +508,8 @@ end CharEq end UniversalProperty +variable {m n : ℕ} + theorem intCast_eq_intCast_iff (a b : ℤ) (c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a ≡ b [ZMOD c] := CharP.intCast_eq_intCast (ZMod c) c @@ -559,7 +561,7 @@ theorem val_intCast {n : ℕ} (a : ℤ) [NeZero n] : ↑(a : ZMod n).val = a % n @[deprecated (since := "2024-04-17")] alias val_int_cast := val_intCast -theorem coe_intCast {n : ℕ} (a : ℤ) : cast (a : ZMod n) = a % n := by +theorem coe_intCast (a : ℤ) : cast (a : ZMod n) = a % n := by cases n · rw [Int.ofNat_zero, Int.emod_zero, Int.cast_id]; rfl · rw [← val_intCast, val]; rfl @@ -567,6 +569,18 @@ theorem coe_intCast {n : ℕ} (a : ℤ) : cast (a : ZMod n) = a % n := by @[deprecated (since := "2024-04-17")] alias coe_int_cast := coe_intCast +lemma intCast_cast_add (x y : ZMod n) : (cast (x + y) : ℤ) = (cast x + cast y) % n := by + rw [← ZMod.coe_intCast, Int.cast_add, ZMod.intCast_zmod_cast, ZMod.intCast_zmod_cast] + +lemma intCast_cast_mul (x y : ZMod n) : (cast (x * y) : ℤ) = cast x * cast y % n := by + rw [← ZMod.coe_intCast, Int.cast_mul, ZMod.intCast_zmod_cast, ZMod.intCast_zmod_cast] + +lemma intCast_cast_sub (x y : ZMod n) : (cast (x - y) : ℤ) = (cast x - cast y) % n := by + rw [← ZMod.coe_intCast, Int.cast_sub, ZMod.intCast_zmod_cast, ZMod.intCast_zmod_cast] + +lemma intCast_cast_neg (x : ZMod n) : (cast (-x) : ℤ) = -cast x % n := by + rw [← ZMod.coe_intCast, Int.cast_neg, ZMod.intCast_zmod_cast] + @[simp] theorem val_neg_one (n : ℕ) : (-1 : ZMod n.succ).val = n := by dsimp [val, Fin.coe_neg] @@ -676,6 +690,13 @@ 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 +lemma val_one'' : ∀ {n}, n ≠ 1 → (1 : ZMod n).val = 1 + | 0, _ => rfl + | 1, hn => by cases hn rfl + | n + 2, _ => + haveI : Fact (1 < n + 2) := ⟨by simp⟩ + ZMod.val_one _ + theorem val_add {n : ℕ} [NeZero n] (a b : ZMod n) : (a + b).val = (a.val + b.val) % n := by cases n · cases NeZero.ne 0 rfl @@ -764,6 +785,11 @@ theorem mul_inv_eq_gcd {n : ℕ} (a : ZMod n) : a * a⁻¹ = Nat.gcd a.val n := rfl _ = Nat.gcd a.val n.succ := by rw [← Nat.gcd_eq_gcd_ab a.val n.succ]; rfl +@[simp] protected lemma inv_one (n : ℕ) : (1⁻¹ : ZMod n) = 1 := by + obtain rfl | hn := eq_or_ne n 1 + · exact Subsingleton.elim _ _ + · simpa [ZMod.val_one'' hn] using mul_inv_eq_gcd (1 : ZMod n) + @[simp] theorem natCast_mod (a : ℕ) (n : ℕ) : ((a % n : ℕ) : ZMod n) = a := by conv => @@ -785,6 +811,15 @@ theorem coe_mul_inv_eq_one {n : ℕ} (x : ℕ) (h : Nat.Coprime x n) : rw [Nat.Coprime, Nat.gcd_comm, Nat.gcd_rec] at h rw [mul_inv_eq_gcd, val_natCast, h, Nat.cast_one] +lemma mul_val_inv (hmn : m.Coprime n) : (m * (m⁻¹ : ZMod n).val : ZMod n) = 1 := by + obtain rfl | hn := eq_or_ne n 0 + · simp [m.coprime_zero_right.1 hmn] + haveI : NeZero n := ⟨hn⟩ + rw [ZMod.natCast_zmod_val, ZMod.coe_mul_inv_eq_one _ hmn] + +lemma val_inv_mul (hmn : m.Coprime n) : ((m⁻¹ : ZMod n).val * m : ZMod n) = 1 := by + rw [mul_comm, mul_val_inv hmn] + /-- `unitOfCoprime` makes an element of `(ZMod n)ˣ` given a natural number `x` and a proof that `x` is coprime to `n` -/ def unitOfCoprime {n : ℕ} (x : ℕ) (h : Nat.Coprime x n) : (ZMod n)ˣ := @@ -980,7 +1015,7 @@ theorem neg_val' {n : ℕ} [NeZero n] (a : ZMod n) : (-a).val = (n - a.val) % n calc (-a).val = val (-a) % n := by rw [Nat.mod_eq_of_lt (-a).val_lt] _ = (n - val a) % n := - Nat.ModEq.add_right_cancel' _ + Nat.ModEq.add_right_cancel' (val a) (by rw [Nat.ModEq, ← val_add, add_left_neg, tsub_add_cancel_of_le a.val_le, Nat.mod_self, val_zero]) @@ -1231,7 +1266,9 @@ instance instField : Field (ZMod p) where mul_inv_cancel := mul_inv_cancel_aux p inv_zero := inv_zero p nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun a x => rfl /-- `ZMod p` is an integral domain when `p` is prime. -/ instance (p : ℕ) [hp : Fact p.Prime] : IsDomain (ZMod p) := by @@ -1321,10 +1358,47 @@ theorem lift_comp_coe : ZMod.lift n f ∘ ((↑) : ℤ → _) = f := theorem lift_comp_castAddHom : (ZMod.lift n f).comp (Int.castAddHom (ZMod n)) = f := AddMonoidHom.ext <| lift_castAddHom _ _ +lemma lift_injective {f : {f : ℤ →+ A // f n = 0}} : + Injective (lift n f) ↔ ∀ m, f.1 m = 0 → (m : ZMod n) = 0 := by + simp only [← AddMonoidHom.ker_eq_bot_iff, eq_bot_iff, SetLike.le_def, + ZMod.intCast_surjective.forall, ZMod.lift_coe, AddMonoidHom.mem_ker, AddSubgroup.mem_bot] + end lift end ZMod +section AddGroup +variable {α : Type*} [AddGroup α] {n : ℕ} + +@[simp] +lemma nsmul_zmod_val_inv_nsmul (hn : (Nat.card α).Coprime n) (a : α) : + n • (n⁻¹ : ZMod (Nat.card α)).val • a = a := by + rw [← mul_nsmul', ← mod_natCard_nsmul, ← ZMod.val_natCast, Nat.cast_mul, + ZMod.mul_val_inv hn.symm, ZMod.val_one_eq_one_mod, mod_natCard_nsmul, one_nsmul] + +@[simp] +lemma zmod_val_inv_nsmul_nsmul (hn : (Nat.card α).Coprime n) (a : α) : + (n⁻¹ : ZMod (Nat.card α)).val • n • a = a := by + rw [nsmul_left_comm, nsmul_zmod_val_inv_nsmul hn] + +end AddGroup + +section Group +variable {α : Type*} [Group α] {n : ℕ} + +-- TODO: Without the `existing`, `to_additive` chokes on `Inv (ZMod n)`. +@[to_additive (attr := simp) existing nsmul_zmod_val_inv_nsmul] +lemma pow_zmod_val_inv_pow (hn : (Nat.card α).Coprime n) (a : α) : + (a ^ (n⁻¹ : ZMod (Nat.card α)).val) ^ n = a := by + rw [← pow_mul', ← pow_mod_natCard, ← ZMod.val_natCast, Nat.cast_mul, ZMod.mul_val_inv hn.symm, + ZMod.val_one_eq_one_mod, pow_mod_natCard, pow_one] + +@[to_additive (attr := simp) existing zmod_val_inv_nsmul_nsmul] +lemma pow_pow_zmod_val_inv (hn : (Nat.card α).Coprime n) (a : α) : + (a ^ n) ^ (n⁻¹ : ZMod (Nat.card α)).val = a := by rw [pow_right_comm, pow_zmod_val_inv_pow hn] + +end Group + /-- The range of `(m * · + k)` on natural numbers is the set of elements `≥ k` in the residue class of `k` mod `m`. -/ lemma Nat.range_mul_add (m k : ℕ) : diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 4b6d242970a0b..a3e51f8b1b144 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -69,7 +69,7 @@ example : Int.instUnitsPow = DivInvMonoid.Pow := rfl @[norm_cast] lemma uzpow_natCast (u : ℤˣ) (n : ℕ) : u ^ (n : R) = u ^ n := by change Additive.toMul ((n : R) • Additive.ofMul u) = _ - rw [← nsmul_eq_smul_cast, toMul_nsmul, toMul_ofMul] + rw [Nat.cast_smul_eq_nsmul, toMul_nsmul, toMul_ofMul] -- See note [no_index around OfNat.ofNat] lemma uzpow_coe_nat (s : ℤˣ) (n : ℕ) [n.AtLeastTwo] : @@ -107,6 +107,6 @@ lemma uzpow_neg (s : ℤˣ) (x : R) : s ^ (-x) = (s ^ x)⁻¹ := @[norm_cast] lemma uzpow_intCast (u : ℤˣ) (z : ℤ) : u ^ (z : R) = u ^ z := by change Additive.toMul ((z : R) • Additive.ofMul u) = _ - rw [← zsmul_eq_smul_cast, toMul_zsmul, toMul_ofMul] + rw [Int.cast_smul_eq_nsmul, toMul_zsmul, toMul_ofMul] end CommRing diff --git a/Mathlib/Data/ZMod/Module.lean b/Mathlib/Data/ZMod/Module.lean index 29d4e97bb91e6..1697e15c69a12 100644 --- a/Mathlib/Data/ZMod/Module.lean +++ b/Mathlib/Data/ZMod/Module.lean @@ -52,7 +52,7 @@ theorem map_smul (f : F) (c : ZMod n) (x : M) : f (c • x) = c • f x := by exact map_intCast_smul f _ _ (cast c) x theorem smul_mem (hx : x ∈ K) (c : ZMod n) : c • x ∈ K := by - rw [← ZMod.intCast_zmod_cast c, ← zsmul_eq_smul_cast] + rw [← ZMod.intCast_zmod_cast c, Int.cast_smul_eq_nsmul] exact zsmul_mem hx (cast c) end ZMod diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index ed31b209996ec..656630e6d541c 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -161,7 +161,7 @@ end IsMonoidHom homomorphism."] theorem IsMulHom.to_isMonoidHom [MulOneClass α] [Group β] {f : α → β} (hf : IsMulHom f) : IsMonoidHom f := - { map_one := mul_right_eq_self.1 <| by rw [← hf.map_mul, one_mul] + { map_one := (mul_right_eq_self (a := f 1)).1 <| by rw [← hf.map_mul, one_mul] map_mul := hf.map_mul } namespace IsMonoidHom diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index e4793ad711d8b..5d229f6e86fd0 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -183,8 +183,7 @@ variable [Group G] @[to_additive] theorem mem_norm_comm {s : Set G} (hs : IsNormalSubgroup s) {a b : G} (hab : a * b ∈ s) : b * a ∈ s := by - have h : a⁻¹ * (a * b) * a⁻¹⁻¹ ∈ s := hs.normal (a * b) hab a⁻¹ - simp at h; exact h + simpa using (hs.normal (a * b) hab a⁻¹) @[to_additive] theorem mem_norm_comm_iff {s : Set G} (hs : IsNormalSubgroup s) {a b : G} : a * b ∈ s ↔ b * a ∈ s := @@ -379,8 +378,8 @@ theorem injective_iff_trivial_ker {f : G → H} (hf : IsGroupHom f) : @[to_additive] theorem trivial_ker_iff_eq_one {f : G → H} (hf : IsGroupHom f) : ker f = trivial G ↔ ∀ x, f x = 1 → x = 1 := by - rw [Set.ext_iff]; simp [ker] - exact ⟨fun h x hx => (h x).1 hx, fun h x => ⟨h x, fun hx => by rw [hx, hf.map_one]⟩⟩ + rw [Set.ext_iff] + simpa [ker] using ⟨fun h x hx => (h x).1 hx, fun h x => ⟨h x, fun hx => by rw [hx, hf.map_one]⟩⟩ end IsGroupHom diff --git a/Mathlib/Dynamics/BirkhoffSum/Average.lean b/Mathlib/Dynamics/BirkhoffSum/Average.lean index c6b44a2f75d11..5696fdcfd31ac 100644 --- a/Mathlib/Dynamics/BirkhoffSum/Average.lean +++ b/Mathlib/Dynamics/BirkhoffSum/Average.lean @@ -71,7 +71,7 @@ theorem birkhoffAverage_congr_ring' (S : Type*) [DivisionSemiring S] [Module S M theorem Function.IsFixedPt.birkhoffAverage_eq [CharZero R] {f : α → α} {x : α} (h : IsFixedPt f x) (g : α → M) {n : ℕ} (hn : n ≠ 0) : birkhoffAverage R f g n x = g x := by - rw [birkhoffAverage, h.birkhoffSum_eq, nsmul_eq_smul_cast R, inv_smul_smul₀] + rw [birkhoffAverage, h.birkhoffSum_eq, ← Nat.cast_smul_eq_nsmul R, inv_smul_smul₀] rwa [Nat.cast_ne_zero] end birkhoffAverage diff --git a/Mathlib/Dynamics/Ergodic/Conservative.lean b/Mathlib/Dynamics/Ergodic/Conservative.lean index cdc8c2ac8e3bc..37e1c0c020126 100644 --- a/Mathlib/Dynamics/Ergodic/Conservative.lean +++ b/Mathlib/Dynamics/Ergodic/Conservative.lean @@ -56,7 +56,7 @@ returns back to `s` under some iteration of `f`. -/ structure Conservative (f : α → α) (μ : Measure α) extends QuasiMeasurePreserving f μ μ : Prop where /-- If `f` is a conservative self-map and `s` is a measurable set of nonzero measure, then there exists a point `x ∈ s` that returns to `s` under a non-zero iteration of `f`. -/ - exists_mem_iterate_mem : ∀ ⦃s⦄, MeasurableSet s → μ s ≠ 0 → ∃ x ∈ s, ∃ m ≠ 0, f^[m] x ∈ s + exists_mem_iterate_mem' : ∀ ⦃s⦄, MeasurableSet s → μ s ≠ 0 → ∃ x ∈ s, ∃ m ≠ 0, f^[m] x ∈ s /-- A self-map preserving a finite measure is conservative. -/ protected theorem MeasurePreserving.conservative [IsFiniteMeasure μ] (h : MeasurePreserving f μ μ) : @@ -68,42 +68,51 @@ namespace Conservative /-- The identity map is conservative w.r.t. any measure. -/ protected theorem id (μ : Measure α) : Conservative id μ := { toQuasiMeasurePreserving := QuasiMeasurePreserving.id μ - exists_mem_iterate_mem := fun _ _ h0 => - let ⟨x, hx⟩ := nonempty_of_measure_ne_zero h0 - ⟨x, hx, 1, one_ne_zero, hx⟩ } + exists_mem_iterate_mem' := fun _ _ h0 => by + simpa [exists_ne] using nonempty_of_measure_ne_zero h0 } + +/-- If `f` is a conservative self-map and `s` is a null measurable set of nonzero measure, +then there exists a point `x ∈ s` that returns to `s` under a non-zero iteration of `f`. -/ +theorem exists_mem_iterate_mem (hf : Conservative f μ) + (hsm : NullMeasurableSet s μ) (hs₀ : μ s ≠ 0) : + ∃ x ∈ s, ∃ m ≠ 0, f^[m] x ∈ s := by + rcases hsm.exists_measurable_subset_ae_eq with ⟨t, hsub, htm, hts⟩ + rcases hf.exists_mem_iterate_mem' htm (by rwa [measure_congr hts]) with ⟨x, hxt, m, hm₀, hmt⟩ + exact ⟨x, hsub hxt, m, hm₀, hsub hmt⟩ /-- If `f` is a conservative map and `s` is a measurable set of nonzero measure, then for infinitely many values of `m` a positive measure of points `x ∈ s` returns back to `s` after `m` iterations of `f`. -/ -theorem frequently_measure_inter_ne_zero (hf : Conservative f μ) (hs : MeasurableSet s) +theorem frequently_measure_inter_ne_zero (hf : Conservative f μ) (hs : NullMeasurableSet s μ) (h0 : μ s ≠ 0) : ∃ᶠ m in atTop, μ (s ∩ f^[m] ⁻¹' s) ≠ 0 := by + set t : ℕ → Set α := fun n ↦ s ∩ f^[n] ⁻¹' s + -- Assume that `μ (t n) ≠ 0`, where `t n = s ∩ f^[n] ⁻¹' s`, only for finitely many `n`. by_contra H - simp only [not_frequently, eventually_atTop, Ne, Classical.not_not] at H - rcases H with ⟨N, hN⟩ - induction' N with N ihN - · apply h0 - simpa using hN 0 le_rfl - rw [imp_false] at ihN - push_neg at ihN - rcases ihN with ⟨n, hn, hμn⟩ - set T := s ∩ ⋃ n ≥ N + 1, f^[n] ⁻¹' s - have hT : MeasurableSet T := - hs.inter (MeasurableSet.biUnion (to_countable _) fun _ _ => hf.measurable.iterate _ hs) - have hμT : μ T = 0 := by - convert (measure_biUnion_null_iff <| to_countable _).2 hN - rw [← inter_iUnion₂] - rfl - have : μ ((s ∩ f^[n] ⁻¹' s) \ T) ≠ 0 := by rwa [measure_diff_null hμT] - rcases hf.exists_mem_iterate_mem ((hs.inter (hf.measurable.iterate n hs)).diff hT) this with - ⟨x, ⟨⟨hxs, _⟩, hxT⟩, m, hm0, ⟨_, hxm⟩, _⟩ - refine hxT ⟨hxs, mem_iUnion₂.2 ⟨n + m, ?_, ?_⟩⟩ - · exact add_le_add hn (Nat.one_le_of_lt <| pos_iff_ne_zero.2 hm0) - · rwa [Set.mem_preimage, ← iterate_add_apply] at hxm + -- Let `N` be the maximal `n` such that `μ (t n) ≠ 0`. + obtain ⟨N, hN, hmax⟩ : ∃ N, μ (t N) ≠ 0 ∧ ∀ n > N, μ (t n) = 0 := by + rw [Nat.frequently_atTop_iff_infinite, not_infinite] at H + convert exists_max_image _ (·) H ⟨0, by simpa⟩ using 4 + rw [gt_iff_lt, ← not_le, not_imp_comm, mem_setOf] + have htm {n : ℕ} : NullMeasurableSet (t n) μ := + hs.inter <| hs.preimage <| hf.toQuasiMeasurePreserving.iterate n + -- Then all `t n`, `n > N`, are null sets, hence `T = t N \ ⋃ n > N, t n` has positive measure. + set T := t N \ ⋃ n > N, t n with hT + have hμT : μ T ≠ 0 := by + rwa [hT, measure_diff_null] + exact (measure_biUnion_null_iff {n | N < n}.to_countable).2 hmax + have hTm : NullMeasurableSet T μ := htm.diff <| .biUnion {n | N < n}.to_countable fun _ _ ↦ htm + -- Take `x ∈ T` and `m ≠ 0` such that `f^[m] x ∈ T`. + rcases hf.exists_mem_iterate_mem hTm hμT with ⟨x, hxt, m, hm₀, hmt⟩ + -- Then `N + m > N`, `x ∈ s`, and `f^[N + m] x = f^[N] (f^[m] x) ∈ s`. + -- This contradicts `x ∈ T ⊆ (⋃ n > N, t n)ᶜ`. + refine hxt.2 <| mem_iUnion₂.2 ⟨N + m, ?_, hxt.1.1, ?_⟩ + · simpa [pos_iff_ne_zero] + · simpa only [iterate_add] using hmt.1.2 /-- If `f` is a conservative map and `s` is a measurable set of nonzero measure, then for an arbitrarily large `m` a positive measure of points `x ∈ s` returns back to `s` after `m` iterations of `f`. -/ -theorem exists_gt_measure_inter_ne_zero (hf : Conservative f μ) (hs : MeasurableSet s) +theorem exists_gt_measure_inter_ne_zero (hf : Conservative f μ) (hs : NullMeasurableSet s μ) (h0 : μ s ≠ 0) (N : ℕ) : ∃ m > N, μ (s ∩ f^[m] ⁻¹' s) ≠ 0 := let ⟨m, hm, hmN⟩ := ((hf.frequently_measure_inter_ne_zero hs h0).and_eventually (eventually_gt_atTop N)).exists @@ -111,31 +120,32 @@ theorem exists_gt_measure_inter_ne_zero (hf : Conservative f μ) (hs : Measurabl /-- Poincaré recurrence theorem: given a conservative map `f` and a measurable set `s`, the set of points `x ∈ s` such that `x` does not return to `s` after `≥ n` iterations has measure zero. -/ -theorem measure_mem_forall_ge_image_not_mem_eq_zero (hf : Conservative f μ) (hs : MeasurableSet s) - (n : ℕ) : μ ({ x ∈ s | ∀ m ≥ n, f^[m] x ∉ s }) = 0 := by +theorem measure_mem_forall_ge_image_not_mem_eq_zero (hf : Conservative f μ) + (hs : NullMeasurableSet s μ) (n : ℕ) : + μ ({ x ∈ s | ∀ m ≥ n, f^[m] x ∉ s }) = 0 := by by_contra H - have : MeasurableSet (s ∩ { x | ∀ m ≥ n, f^[m] x ∉ s }) := by + have : NullMeasurableSet (s ∩ { x | ∀ m ≥ n, f^[m] x ∉ s }) μ := by simp only [setOf_forall, ← compl_setOf] - exact - hs.inter (MeasurableSet.biInter (to_countable _) fun m _ => hf.measurable.iterate m hs.compl) + exact hs.inter <| .biInter (to_countable _) fun m _ ↦ + (hs.preimage <| hf.toQuasiMeasurePreserving.iterate m).compl rcases (hf.exists_gt_measure_inter_ne_zero this H) n with ⟨m, hmn, hm⟩ rcases nonempty_of_measure_ne_zero hm with ⟨x, ⟨_, hxn⟩, hxm, -⟩ exact hxn m hmn.lt.le hxm /-- Poincaré recurrence theorem: given a conservative map `f` and a measurable set `s`, almost every point `x ∈ s` returns back to `s` infinitely many times. -/ -theorem ae_mem_imp_frequently_image_mem (hf : Conservative f μ) (hs : MeasurableSet s) : +theorem ae_mem_imp_frequently_image_mem (hf : Conservative f μ) (hs : NullMeasurableSet s μ) : ∀ᵐ x ∂μ, x ∈ s → ∃ᶠ n in atTop, f^[n] x ∈ s := by simp only [frequently_atTop, @forall_swap (_ ∈ s), ae_all_iff] intro n filter_upwards [measure_zero_iff_ae_nmem.1 (hf.measure_mem_forall_ge_image_not_mem_eq_zero hs n)] simp -theorem inter_frequently_image_mem_ae_eq (hf : Conservative f μ) (hs : MeasurableSet s) : +theorem inter_frequently_image_mem_ae_eq (hf : Conservative f μ) (hs : NullMeasurableSet s μ) : (s ∩ { x | ∃ᶠ n in atTop, f^[n] x ∈ s } : Set α) =ᵐ[μ] s := inter_eventuallyEq_left.2 <| hf.ae_mem_imp_frequently_image_mem hs -theorem measure_inter_frequently_image_mem_eq (hf : Conservative f μ) (hs : MeasurableSet s) : +theorem measure_inter_frequently_image_mem_eq (hf : Conservative f μ) (hs : NullMeasurableSet s μ) : μ (s ∩ { x | ∃ᶠ n in atTop, f^[n] x ∈ s }) = μ s := measure_congr (hf.inter_frequently_image_mem_ae_eq hs) @@ -143,17 +153,18 @@ theorem measure_inter_frequently_image_mem_eq (hf : Conservative f μ) (hs : Mea set, then for `μ`-a.e. `x`, if the orbit of `x` visits `s` at least once, then it visits `s` infinitely many times. -/ theorem ae_forall_image_mem_imp_frequently_image_mem (hf : Conservative f μ) - (hs : MeasurableSet s) : ∀ᵐ x ∂μ, ∀ k, f^[k] x ∈ s → ∃ᶠ n in atTop, f^[n] x ∈ s := by + (hs : NullMeasurableSet s μ) : ∀ᵐ x ∂μ, ∀ k, f^[k] x ∈ s → ∃ᶠ n in atTop, f^[n] x ∈ s := by refine ae_all_iff.2 fun k => ?_ - refine (hf.ae_mem_imp_frequently_image_mem (hf.measurable.iterate k hs)).mono fun x hx hk => ?_ + refine (hf.ae_mem_imp_frequently_image_mem + (hs.preimage <| hf.toQuasiMeasurePreserving.iterate k)).mono fun x hx hk => ?_ rw [← map_add_atTop_eq_nat k, frequently_map] refine (hx hk).mono fun n hn => ?_ rwa [add_comm, iterate_add_apply] /-- If `f` is a conservative self-map and `s` is a measurable set of positive measure, then `ae μ`-frequently we have `x ∈ s` and `s` returns to `s` under infinitely many iterations of `f`. -/ -theorem frequently_ae_mem_and_frequently_image_mem (hf : Conservative f μ) (hs : MeasurableSet s) - (h0 : μ s ≠ 0) : ∃ᵐ x ∂μ, x ∈ s ∧ ∃ᶠ n in atTop, f^[n] x ∈ s := +theorem frequently_ae_mem_and_frequently_image_mem (hf : Conservative f μ) + (hs : NullMeasurableSet s μ) (h0 : μ s ≠ 0) : ∃ᵐ x ∂μ, x ∈ s ∧ ∃ᶠ n in atTop, f^[n] x ∈ s := ((frequently_ae_mem_iff.2 h0).and_eventually (hf.ae_mem_imp_frequently_image_mem hs)).mono fun _ hx => ⟨hx.1, hx.2 hx.1⟩ @@ -164,18 +175,19 @@ theorem ae_frequently_mem_of_mem_nhds [TopologicalSpace α] [SecondCountableTopo [OpensMeasurableSpace α] {f : α → α} {μ : Measure α} (h : Conservative f μ) : ∀ᵐ x ∂μ, ∀ s ∈ 𝓝 x, ∃ᶠ n in atTop, f^[n] x ∈ s := by have : ∀ s ∈ countableBasis α, ∀ᵐ x ∂μ, x ∈ s → ∃ᶠ n in atTop, f^[n] x ∈ s := fun s hs => - h.ae_mem_imp_frequently_image_mem (isOpen_of_mem_countableBasis hs).measurableSet + h.ae_mem_imp_frequently_image_mem (isOpen_of_mem_countableBasis hs).nullMeasurableSet refine ((ae_ball_iff <| countable_countableBasis α).2 this).mono fun x hx s hs => ?_ rcases (isBasis_countableBasis α).mem_nhds_iff.1 hs with ⟨o, hoS, hxo, hos⟩ exact (hx o hoS hxo).mono fun n hn => hos hn /-- Iteration of a conservative system is a conservative system. -/ protected theorem iterate (hf : Conservative f μ) (n : ℕ) : Conservative f^[n] μ := by + -- Discharge the trivial case `n = 0` cases' n with n · exact Conservative.id μ - -- Discharge the trivial case `n = 0` refine ⟨hf.1.iterate _, fun s hs hs0 => ?_⟩ - rcases (hf.frequently_ae_mem_and_frequently_image_mem hs hs0).exists with ⟨x, _, hx⟩ + rcases (hf.frequently_ae_mem_and_frequently_image_mem hs.nullMeasurableSet hs0).exists + with ⟨x, _, hx⟩ /- We take a point `x ∈ s` such that `f^[k] x ∈ s` for infinitely many values of `k`, then we choose two of these values `k < l` such that `k ≡ l [MOD (n + 1)]`. Then `f^[k] x ∈ s` and `f^[n + 1]^[(l - k) / (n + 1)] (f^[k] x) = f^[l] x ∈ s`. -/ diff --git a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean index c2aa662af63b2..9e8c651529028 100644 --- a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean +++ b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean @@ -120,6 +120,10 @@ theorem measure_preimage_emb {f : α → β} (hf : MeasurePreserving f μa μb) (hfe : MeasurableEmbedding f) (s : Set β) : μa (f ⁻¹' s) = μb s := by rw [← hf.map_eq, hfe.map_apply] +theorem measure_preimage_equiv {f : α ≃ᵐ β} (hf : MeasurePreserving f μa μb) (s : Set β) : + μa (f ⁻¹' s) = μb s := + measure_preimage_emb hf f.measurableEmbedding s + protected theorem iterate {f : α → α} (hf : MeasurePreserving f μa μa) : ∀ n, MeasurePreserving f^[n] μa μa | 0 => MeasurePreserving.id μa @@ -147,7 +151,8 @@ theorem exists_mem_iterate_mem_of_volume_lt_mul_volume (hf : MeasurePreserving f have : μ (univ : Set α) < ∑ m ∈ Finset.range n, μ (f^[m] ⁻¹' s) := by simpa [B] obtain ⟨i, hi, j, hj, hij, x, hxi : f^[i] x ∈ s, hxj : f^[j] x ∈ s⟩ : ∃ i < n, ∃ j < n, i ≠ j ∧ (f^[i] ⁻¹' s ∩ f^[j] ⁻¹' s).Nonempty := by - simpa using exists_nonempty_inter_of_measure_univ_lt_sum_measure μ (fun m _ ↦ A m) this + simpa using exists_nonempty_inter_of_measure_univ_lt_sum_measure μ + (fun m _ ↦ (A m).nullMeasurableSet) this wlog hlt : i < j generalizing i j · exact this j hj i hi hij.symm hxj hxi (hij.lt_or_lt.resolve_left hlt) refine ⟨f^[i] x, hxi, j - i, ⟨tsub_pos_of_lt hlt, lt_of_le_of_lt (j.sub_le i) hj⟩, ?_⟩ diff --git a/Mathlib/FieldTheory/Galois.lean b/Mathlib/FieldTheory/Galois.lean index 050ca43a1e556..97c0878478581 100644 --- a/Mathlib/FieldTheory/Galois.lean +++ b/Mathlib/FieldTheory/Galois.lean @@ -389,9 +389,11 @@ theorem of_separable_splitting_field [sp : p.IsSplittingField F E] (hp : p.Separ rfl /-- Equivalent characterizations of a Galois extension of finite degree-/ -theorem tfae [FiniteDimensional F E] : - List.TFAE [IsGalois F E, IntermediateField.fixedField (⊤ : Subgroup (E ≃ₐ[F] E)) = ⊥, - Fintype.card (E ≃ₐ[F] E) = finrank F E, ∃ p : F[X], p.Separable ∧ p.IsSplittingField F E] := by +theorem tfae [FiniteDimensional F E] : List.TFAE [ + IsGalois F E, + IntermediateField.fixedField (⊤ : Subgroup (E ≃ₐ[F] E)) = ⊥, + Fintype.card (E ≃ₐ[F] E) = finrank F E, + ∃ p : F[X], p.Separable ∧ p.IsSplittingField F E] := by tfae_have 1 → 2 · exact fun h => OrderIso.map_bot (@intermediateFieldEquivSubgroup F _ E _ _ _ h).symm tfae_have 1 → 3 diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index fdf673b2198b1..8329c6f2f2b24 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -226,8 +226,8 @@ theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) ring! convert congr_arg₂ Add.add (congr_arg (fun t => t • (y : E)) h₁) (congr_arg (fun t => t • v) h₂) using 1 - · simp [a, inner_add_right, inner_smul_right, hvy, real_inner_self_eq_norm_mul_norm, hv, mul_smul, - mul_pow, Real.norm_eq_abs, sq_abs, norm_smul] + · simp only [innerSL_apply, norm_smul, norm_div, RCLike.norm_ofNat, Real.norm_eq_abs, + AddSubgroupClass.coe_norm, mul_pow, div_pow, sq_abs, SetLike.val_smul, mul_smul, a] -- Porting note: used to be simp only [split, add_comm] but get maxRec errors rw [split, add_comm] ac_rfl diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 23f3042d1cf62..f9d20aedccf3b 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Geometry.Manifold.ChartedSpace -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.Analysis.Calculus.ContDiff.Basic /-! diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index daea2bc7fd7c1..1c743e1acc6b1 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -460,6 +460,13 @@ namespace MemLeftTransversals noncomputable def toEquiv (hS : S ∈ Subgroup.leftTransversals (H : Set G)) : G ⧸ H ≃ S := (Equiv.ofBijective _ (Subgroup.mem_leftTransversals_iff_bijective.mp hS)).symm +@[to_additive "A left transversal is finite iff the subgroup has finite index"] +theorem finite_iff + (h : S ∈ Subgroup.leftTransversals (H : Set G)) : + Finite S ↔ H.FiniteIndex := by + rw [← (Subgroup.MemLeftTransversals.toEquiv h).finite_iff] + exact ⟨fun _ ↦ finiteIndex_of_finite_quotient H, fun _ ↦ finite_quotient_of_finiteIndex H⟩ + @[to_additive] theorem mk''_toEquiv (hS : S ∈ Subgroup.leftTransversals (H : Set G)) (q : G ⧸ H) : Quotient.mk'' (toEquiv hS q : G) = q := @@ -498,6 +505,14 @@ noncomputable def toEquiv (hS : S ∈ Subgroup.rightTransversals (H : Set G)) : Quotient (QuotientGroup.rightRel H) ≃ S := (Equiv.ofBijective _ (Subgroup.mem_rightTransversals_iff_bijective.mp hS)).symm +@[to_additive "A right transversal is finite iff the subgroup has finite index"] +theorem finite_iff + (h : S ∈ Subgroup.rightTransversals (H : Set G)) : + Finite S ↔ H.FiniteIndex := by + rw [← (Subgroup.MemRightTransversals.toEquiv h).finite_iff, + (QuotientGroup.quotientRightRelEquivQuotientLeftRel H).finite_iff] + exact ⟨fun _ ↦ finiteIndex_of_finite_quotient H, fun _ ↦ finite_quotient_of_finiteIndex H⟩ + @[to_additive] theorem mk''_toEquiv (hS : S ∈ Subgroup.rightTransversals (H : Set G)) (q : Quotient (QuotientGroup.rightRel H)) : Quotient.mk'' (toEquiv hS q : G) = q := diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 9602ff8977067..2c1b0c3b3405a 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -573,7 +573,7 @@ the order-preserving bijection between the set of additive congruence relations the additive congruence relations on the quotient of `M` by `c`."] def correspondence : { d // c ≤ d } ≃o Con c.Quotient where toFun d := - d.1.mapOfSurjective (↑) _ (by rw [mul_ker_mk_eq]; exact d.2) <| + d.1.mapOfSurjective (↑) (fun x y => rfl) (by rw [mul_ker_mk_eq]; exact d.2) <| @Quotient.exists_rep _ c.toSetoid invFun d := ⟨comap ((↑) : M → c.Quotient) (fun x y => rfl) d, fun x y h => diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 40d527909e6c1..719e0b7b18fae 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -526,7 +526,7 @@ theorem mem_smul_iff {i j : ι} {m₁ : M i} {m₂ : M j} {w : Word M} : intro hm1 split_ifs with h · rcases h with ⟨hnil, rfl⟩ - simp only [List.head?_eq_head _ hnil, Option.some.injEq, ne_eq] + simp only [List.head?_eq_head hnil, Option.some.injEq, ne_eq] constructor · rintro rfl exact Or.inl ⟨_, rfl, rfl⟩ @@ -637,7 +637,7 @@ def toList : ∀ {i j} (_w : NeWord M i j), List (Σi, M i) theorem toList_ne_nil {i j} (w : NeWord M i j) : w.toList ≠ List.nil := by induction w · rintro ⟨rfl⟩ - · apply List.append_ne_nil_of_ne_nil_left + · apply List.append_ne_nil_of_left_ne_nil assumption /-- The first letter of a `NeWord` -/ diff --git a/Mathlib/GroupTheory/Coset.lean b/Mathlib/GroupTheory/Coset.lean index 79ede1aa10aa1..fc3913d35b708 100644 --- a/Mathlib/GroupTheory/Coset.lean +++ b/Mathlib/GroupTheory/Coset.lean @@ -178,7 +178,7 @@ theorem orbit_subgroup_one_eq_self : MulAction.orbit s (1 : α) = s := @[to_additive eq_addCosets_of_normal] theorem eq_cosets_of_normal (N : s.Normal) (g : α) : g • (s : Set α) = op g • s := - Set.ext fun a => by simp [mem_leftCoset_iff, mem_rightCoset_iff]; rw [N.mem_comm_iff] + Set.ext fun a => by simp [mem_leftCoset_iff, mem_rightCoset_iff, N.mem_comm_iff] @[to_additive normal_of_eq_addCosets] theorem normal_of_eq_cosets (h : ∀ g : α, g • (s : Set α) = op g • s) : s.Normal := diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean new file mode 100644 index 0000000000000..3e09c58c89a09 --- /dev/null +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -0,0 +1,387 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir, Richard Copley. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, Richard Copley +-/ + +import Mathlib.GroupTheory.Complement +import Mathlib.LinearAlgebra.Basis.VectorSpace + +/-! # Lemma of B. H. Neumann on coverings of a group by cosets. + +Let the group $G$ be the union of finitely many, let us say $n$, left cosets +of subgroups $C₁$, $C₂$, ..., $Cₙ$: $$ G = ⋃_{i = 1}^n C_i g_i. $$ + +* `Subgroup.exists_finiteIndex_of_leftCoset_cover` + at least one subgroup $C_i$ has finite index in $G$. + +* `Subgroup.leftCoset_cover_filter_FiniteIndex` + the cosets of subgroups of infinite index may be omitted from the covering. + +* `Subgroup.exists_index_le_card_of_leftCoset_cover` : + the index of (at least) one of these subgroups does not exceed $n$. + +* `Subgroup.one_le_sum_inv_index_of_leftCoset_cover` : + the sum of the inverses of the indexes of the $C_i$ is greater than or equal to 1. + +* `Subgroup.pairwiseDisjoint_leftCoset_cover_of_sum_inv_index_eq_one` + If the sum of the inverses of the indexes of the subgroups $C_i$ is equal to 1, + then the cosets of the subgroups of finite index are pairwise disjoint. + +A corollary of `Subgroup.exists_finiteIndex_of_leftCoset_cover` is: + +* `Subspace.union_ne_univ_of_lt_top` : + a vector space over an infinite field cannot be a finite union of proper subspaces. + +This can be used to show that an algebraic extension of fields is determined by the +set of all minimal polynomials (not proved here). + +[1] [Neumann-1954], *Groups Covered By Permutable Subsets*, Lemma 4.1 +[2] +[3] + +-/ + +open scoped Pointwise BigOperators + +namespace Subgroup + +variable {G : Type*} [Group G] + +section leftCoset_cover_const + +@[to_additive] +theorem exists_leftTransversal_of_FiniteIndex + {D H : Subgroup G} [D.FiniteIndex] (hD_le_H : D ≤ H) : + ∃ t : Finset H, + (t : Set H) ∈ leftTransversals (D.subgroupOf H) ∧ + ⋃ g ∈ t, (g : G) • (D : Set G) = H := by + have ⟨t, ht⟩ := exists_left_transversal (D.subgroupOf H) 1 + have hf : t.Finite := (MemLeftTransversals.finite_iff ht.1).mpr inferInstance + refine ⟨hf.toFinset, hf.coe_toFinset.symm ▸ ht.1, ?_⟩ + ext x + suffices (∃ y ∈ t, ∃ d ∈ D, y * d = x) ↔ x ∈ H by simpa using this + constructor + · rintro ⟨⟨y, hy⟩, -, d, h, rfl⟩ + exact H.mul_mem hy (hD_le_H h) + · intro hx + exact ⟨_, (MemLeftTransversals.toFun ht.1 ⟨x, hx⟩).2, _, + MemLeftTransversals.inv_toFun_mul_mem ht.1 ⟨x, hx⟩, mul_inv_cancel_left _ _⟩ + +variable {ι : Type*} {s : Finset ι} {H : Subgroup G} {g : ι → G} + +@[to_additive] +theorem leftCoset_cover_const_iff_surjOn : + ⋃ i ∈ s, g i • (H : Set G) = Set.univ ↔ Set.SurjOn (g · : ι → G ⧸ H) s Set.univ := by + simp [Set.eq_univ_iff_forall, mem_leftCoset_iff, Set.SurjOn, + QuotientGroup.forall_mk, QuotientGroup.eq] + +variable (hcovers : ⋃ i ∈ s, g i • (H : Set G) = Set.univ) + +/-- If `H` is a subgroup of `G` and `G` is the union of a finite family of left cosets of `H` +then `H` has finite index. -/ +@[to_additive] +theorem finiteIndex_of_leftCoset_cover_const : H.FiniteIndex := by + simp_rw [leftCoset_cover_const_iff_surjOn] at hcovers + have := Set.finite_univ_iff.mp <| Set.Finite.of_surjOn _ hcovers s.finite_toSet + exact H.finiteIndex_of_finite_quotient + +@[to_additive] +theorem index_le_of_leftCoset_cover_const : H.index ≤ s.card := by + cases H.index.eq_zero_or_pos with + | inl h => exact h ▸ s.card.zero_le + | inr h => + rw [leftCoset_cover_const_iff_surjOn, Set.surjOn_iff_surjective] at hcovers + exact (Nat.card_le_card_of_surjective _ hcovers).trans_eq (Nat.card_eq_finsetCard _) + +@[to_additive] +theorem pairwiseDisjoint_leftCoset_cover_const_of_index_eq (hind : H.index = s.card) : + Set.PairwiseDisjoint s (g · • (H : Set G)) := by + have : Fintype (G ⧸ H) := fintypeOfIndexNeZero fun h => by + rw [hind, Finset.card_eq_zero] at h + rw [h, ← Finset.set_biUnion_coe, Finset.coe_empty, Set.biUnion_empty] at hcovers + exact Set.empty_ne_univ hcovers + suffices Function.Bijective (g · : s → G ⧸ H) by + intro i hi j hj h' c hi' hj' x hx + specialize hi' hx + specialize hj' hx + rw [mem_leftCoset_iff, SetLike.mem_coe, ← QuotientGroup.eq] at hi' hj' + rw [ne_eq, ← Subtype.mk.injEq (p := (· ∈ (s : Set ι))) i hi j hj] at h' + exact h' <| this.injective <| by simp only [hi', hj'] + rw [Fintype.bijective_iff_surjective_and_card] + constructor + · rwa [leftCoset_cover_const_iff_surjOn, Set.surjOn_iff_surjective] at hcovers + · simp only [Fintype.card_coe, ← hind, index_eq_card, Nat.card_eq_fintype_card] + +end leftCoset_cover_const + +section + +variable {ι : Type*} {H : ι → Subgroup G} {g : ι → G} {s : Finset ι} + (hcovers : ⋃ i ∈ s, (g i) • (H i : Set G) = Set.univ) + +-- Inductive inner part of `Subgroup.exists_finiteIndex_of_leftCoset_cover` +@[to_additive] +theorem exists_finiteIndex_of_leftCoset_cover_aux [DecidableEq (Subgroup G)] + (j : ι) (hj : j ∈ s) (hcovers' : ⋃ i ∈ s.filter (H · = H j), g i • (H i : Set G) ≠ Set.univ) : + ∃ i ∈ s, H i ≠ H j ∧ (H i).FiniteIndex := by + classical + have ⟨n, hn⟩ : ∃ n, n = (s.image H).card := exists_eq + induction n using Nat.strongRec generalizing ι with + | ind n ih => + -- Every left coset of `H j` is contained in a finite union of + -- left cosets of the other subgroups `H k ≠ H j` of the covering. + have ⟨x, hx⟩ : ∃ (x : G), ∀ i ∈ s, H i = H j → (g i : G ⧸ H i) ≠ ↑x := by + simpa [Set.eq_univ_iff_forall, mem_leftCoset_iff, ← QuotientGroup.eq] using hcovers' + replace hx : ∀ (y : G), y • (H j : Set G) ⊆ + ⋃ i ∈ s.filter (H · ≠ H j), (y * x⁻¹ * g i) • (H i : Set G) := by + intro y z hz + simp_rw [Finset.mem_filter, Set.mem_iUnion] + have ⟨i, hi, hmem⟩ : ∃ i ∈ s, x * (y⁻¹ * z) ∈ g i • (H i : Set G) := by + simpa using Set.eq_univ_iff_forall.mp hcovers (x * (y⁻¹ * z)) + rw [mem_leftCoset_iff, SetLike.mem_coe, ← QuotientGroup.eq] at hmem + refine ⟨i, ⟨hi, fun hij => hx i hi hij ?_⟩, ?_⟩ + · rwa [hmem, eq_comm, QuotientGroup.eq, hij, inv_mul_cancel_left, + ← SetLike.mem_coe, ← mem_leftCoset_iff] + · simpa [mem_leftCoset_iff, SetLike.mem_coe, QuotientGroup.eq, mul_assoc] using hmem + -- Thus `G` can also be covered by a finite union `U k, f k • K k` of left cosets + -- of the subgroups `H k ≠ H j`. + let κ := ↥(s.filter (H · ≠ H j)) × Option ↥(s.filter (H · = H j)) + let f : κ → G + | ⟨k₁, some k₂⟩ => g k₂ * x⁻¹ * g k₁ + | ⟨k₁, none⟩ => g k₁ + let K (k : κ) : Subgroup G := H k.1.val + have hK' (k : κ) : K k ∈ (s.image H).erase (H j) := by + have := Finset.mem_filter.mp k.1.property + exact Finset.mem_erase.mpr ⟨this.2, Finset.mem_image_of_mem H this.1⟩ + have hK (k : κ) : K k ≠ H j := ((Finset.mem_erase.mp (hK' k)).left ·) + replace hcovers : ⋃ k ∈ Finset.univ, f k • (K k : Set G) = Set.univ := + Set.iUnion₂_eq_univ_iff.mpr fun y => by + rw [← s.filter_union_filter_neg_eq (H · = H j), Finset.set_biUnion_union] at hcovers + cases (Set.mem_union _ _ _).mp (hcovers.superset (Set.mem_univ y)) with + | inl hy => + have ⟨k, hk, hy⟩ := Set.mem_iUnion₂.mp hy + have hk' : H k = H j := And.right <| by simpa using hk + have ⟨i, hi, hy⟩ := Set.mem_iUnion₂.mp (hx (g k) (hk' ▸ hy)) + exact ⟨⟨⟨i, hi⟩, some ⟨k, hk⟩⟩, Finset.mem_univ _, hy⟩ + | inr hy => + have ⟨i, hi, hy⟩ := Set.mem_iUnion₂.mp hy + exact ⟨⟨⟨i, hi⟩, none⟩, Finset.mem_univ _, hy⟩ + -- Let `H k` be one of the subgroups in this covering. + have ⟨k⟩ : Nonempty κ := not_isEmpty_iff.mp fun hempty => by + rw [Set.iUnion_of_empty] at hcovers + exact Set.empty_ne_univ hcovers + -- If `G` is the union of the cosets of `H k` in the new covering, we are done. + by_cases hcovers' : ⋃ i ∈ Finset.filter (K · = K k) Finset.univ, f i • (K i : Set G) = Set.univ + · rw [Set.iUnion₂_congr fun i hi => by rw [(Finset.mem_filter.mp hi).right]] at hcovers' + exact ⟨k.1, Finset.mem_of_mem_filter k.1.1 k.1.2, hK k, + finiteIndex_of_leftCoset_cover_const hcovers'⟩ + -- Otherwise, by the induction hypothesis, one of the subgroups `H k ≠ H j` has finite index. + have hn' : (Finset.univ.image K).card < n := hn ▸ by + refine ((Finset.card_le_card fun x => ?_).trans_lt <| + Finset.card_erase_lt_of_mem (Finset.mem_image_of_mem H hj)) + rw [mem_image_univ_iff_mem_range, Set.mem_range] + exact fun ⟨k, hk⟩ => hk ▸ hK' k + have ⟨k', hk'⟩ := ih _ hn' hcovers k (Finset.mem_univ k) hcovers' rfl + exact ⟨k'.1.1, Finset.mem_of_mem_filter k'.1.1 k'.1.2, hK k', hk'.2.2⟩ + +/-- Let the group `G` be the union of finitely many left cosets `g i • H i`. +Then at least one subgroup `H i` has finite index in `G`. -/ +@[to_additive] +theorem exists_finiteIndex_of_leftCoset_cover : ∃ k ∈ s, (H k).FiniteIndex := by + classical + have ⟨j, hj⟩ : s.Nonempty := Finset.nonempty_iff_ne_empty.mpr fun hempty => by + rw [hempty, ← Finset.set_biUnion_coe, Finset.coe_empty, Set.biUnion_empty] at hcovers + exact Set.empty_ne_univ hcovers + by_cases hcovers' : ⋃ i ∈ s.filter (H · = H j), g i • (H i : Set G) = Set.univ + · rw [Set.iUnion₂_congr fun i hi => by rw [(Finset.mem_filter.mp hi).right]] at hcovers' + exact ⟨j, hj, finiteIndex_of_leftCoset_cover_const hcovers'⟩ + · have ⟨i, hi, _, hfi⟩ := + exists_finiteIndex_of_leftCoset_cover_aux hcovers j hj hcovers' + exact ⟨i, hi, hfi⟩ + +-- Auxiliary to `leftCoset_cover_filter_FiniteIndex` and `one_le_sum_inv_index_of_leftCoset_cover`. +@[to_additive] +theorem leftCoset_cover_filter_FiniteIndex_aux + [DecidablePred (FiniteIndex : Subgroup G → Prop)] : + (⋃ k ∈ s.filter (fun i => (H i).FiniteIndex), g k • (H k : Set G) = Set.univ) ∧ + (1 ≤ ∑ i ∈ s, ((H i).index : ℚ)⁻¹) ∧ + (∑ i ∈ s, ((H i).index : ℚ)⁻¹ = 1 → Set.PairwiseDisjoint + (s.filter (fun i => (H i).FiniteIndex)) (fun i ↦ g i • (H i : Set G))) := by + classical + let D := ⨅ k ∈ s.filter (fun i => (H i).FiniteIndex), H k + -- `D`, as the finite intersection of subgroups of finite index, also has finite index. + have hD : D.FiniteIndex := finiteIndex_iInf' _ <| by simp + have hD_le {i} (hi : i ∈ s) (hfi : (H i).FiniteIndex) : D ≤ H i := + iInf₂_le i (Finset.mem_filter.mpr ⟨hi, hfi⟩) + -- Each subgroup of finite index in the covering is the union of finitely many cosets of `D`. + choose t ht using fun i hi hfi => + exists_leftTransversal_of_FiniteIndex (H := H i) (hD_le hi hfi) + -- We construct a cover of `G` by the cosets of subgroups of infinite index and of `D`. + let κ := (i : s) × { x // x ∈ if h : (H i.1).FiniteIndex then t i.1 i.2 h else {1} } + let f (k : κ) : G := g k.1 * k.2.val + let K (k : κ) : Subgroup G := if (H k.1).FiniteIndex then D else H k.1 + have hcovers' : ⋃ k ∈ Finset.univ, f k • (K k : Set G) = Set.univ := by + rw [← s.filter_union_filter_neg_eq (fun i => (H i).FiniteIndex)] at hcovers + rw [← hcovers, ← Finset.univ.filter_union_filter_neg_eq (fun k => (H k.1).FiniteIndex), + Finset.set_biUnion_union, Finset.set_biUnion_union] + apply congrArg₂ (· ∪ ·) <;> rw [Set.iUnion_sigma, Set.iUnion_subtype] <;> + refine Set.iUnion_congr fun i => ?_ + · by_cases hfi : (H i).FiniteIndex <;> + simp [← Set.smul_set_iUnion₂, Set.iUnion_subtype, ← leftCoset_assoc, f, K, ht, hfi] + · by_cases hfi : (H i).FiniteIndex <;> + simp [Set.iUnion_subtype, f, K, hfi] + -- There is at least one coset of a subgroup of finite index in the original covering. + -- Therefore a coset of `D` occurs in the new covering. + have ⟨k, hkfi, hk⟩ : ∃ k, (H k.1.1).FiniteIndex ∧ K k = D := + have ⟨j, hj, hjfi⟩ := exists_finiteIndex_of_leftCoset_cover hcovers + have ⟨x, hx⟩ : (t j hj hjfi).Nonempty := Finset.nonempty_coe_sort.mp + (MemLeftTransversals.toEquiv (ht j hj hjfi).1).symm.nonempty + ⟨⟨⟨j, hj⟩, ⟨x, dif_pos hjfi ▸ hx⟩⟩, hjfi, if_pos hjfi⟩ + -- Since `D` is the unique subgroup of finite index whose cosets occur in the new covering, + -- the cosets of the other subgroups can be omitted. + replace hcovers' : ⋃ i ∈ Finset.univ.filter (K · = D), f i • (D : Set G) = Set.univ := by + rw [← hk, Set.iUnion₂_congr fun i hi => by rw [← (Finset.mem_filter.mp hi).2]] + by_contra! h + obtain ⟨i, -, hi⟩ := + exists_finiteIndex_of_leftCoset_cover_aux hcovers' k (Finset.mem_univ k) h + by_cases hfi : (H i.1.1).FiniteIndex <;> simp [K, hfi, hkfi] at hi + -- The result follows by restoring the original cosets of subgroups of finite index + -- from the cosets of `D` into which they have been decomposed. + have hHD (i) : ¬(H i).FiniteIndex → H i ≠ D := fun hfi hD' => (hD' ▸ hfi) hD + have hdensity : ∑ i ∈ s, ((H i).index : ℚ)⁻¹ = + (Finset.univ.filter (K · = D)).card * (D.index : ℚ)⁻¹ := by + rw [eq_mul_inv_iff_mul_eq₀ (Nat.cast_ne_zero.mpr hD.finiteIndex), Finset.sum_mul, + ← Finset.sum_attach, eq_comm, Finset.card_filter, Nat.cast_sum, ← Finset.univ_sigma_univ, + Finset.sum_sigma, Finset.sum_coe_sort_eq_attach] + refine Finset.sum_congr rfl fun i _ => ?_ + by_cases hfi : (H i).FiniteIndex + · rw [← relindex_mul_index (hD_le i.2 hfi), Nat.cast_mul, mul_comm, + mul_inv_cancel_right₀ (Nat.cast_ne_zero.mpr hfi.finiteIndex)] + simpa [K, hfi] using card_left_transversal (ht i.1 i.2 hfi).1 + · rw [of_not_not (FiniteIndex.mk.mt hfi), Nat.cast_zero, inv_zero, zero_mul] + simpa [K, hfi] using hHD i hfi + refine ⟨?_, ?_, ?_⟩ + · rw [← hcovers', Set.iUnion_sigma, Set.iUnion_subtype] + refine Set.iUnion_congr fun i => ?_ + rw [Finset.mem_filter, Set.iUnion_and] + refine Set.iUnion_congr fun hi => ?_ + by_cases hfi : (H i).FiniteIndex <;> + simp [Set.smul_set_iUnion, Set.iUnion_subtype, ← leftCoset_assoc, + f, K, hHD, ← (ht i hi _).2, hi, hfi, hkfi] + · rw [hdensity] + refine le_of_mul_le_mul_right ?_ (Nat.cast_pos.mpr (Nat.pos_of_ne_zero hD.finiteIndex)) + rw [one_mul, mul_assoc, inv_mul_cancel (Nat.cast_ne_zero.mpr hD.finiteIndex), mul_one, + Nat.cast_le] + exact index_le_of_leftCoset_cover_const hcovers' + · rw [hdensity, mul_inv_eq_one₀ (Nat.cast_ne_zero.mpr hD.finiteIndex), + Nat.cast_inj, Finset.coe_filter] + intro h i hi j hj hij c hi' hj' x hx + have hdisjoint := pairwiseDisjoint_leftCoset_cover_const_of_index_eq hcovers' h.symm + -- We know the `f k • K k` are pairwise disjoint and need to prove that the `g i • H i` are. + rw [Set.mem_setOf_eq] at hi hj + have hk' (i) (hi : i ∈ s ∧ (H i).FiniteIndex) (hi' : c ≤ g i • (H i : Set G)) : + ∃ (k : κ), k.1.1 = i ∧ K k = D ∧ x ∈ f k • (D : Set G) := by + rw [← (ht i hi.1 hi.2).2] at hi' + suffices ∃ r : H i, r ∈ t i hi.1 hi.2 ∧ x ∈ (g i * r) • (D : Set G) by + have ⟨r, hr, hxr⟩ := this + refine ⟨⟨⟨i, hi.1⟩, ⟨r, dif_pos hi.2 ▸ hr⟩⟩, rfl, ?_⟩ + simpa [K, f, if_pos hi.2] using hxr + simpa [Set.mem_smul_set_iff_inv_smul_mem, smul_eq_mul, mul_assoc] using hi' hx + have ⟨k₁, hik₁, hk₁, hxk₁⟩ := hk' i hi hi' + have ⟨k₂, hjk₂, hk₂, hxk₂⟩ := hk' j hj hj' + rw [← Set.singleton_subset_iff, ← Set.le_iff_subset] at hxk₁ hxk₂ ⊢ + exact hdisjoint + (Finset.mem_filter.mpr ⟨Finset.mem_univ k₁, hk₁⟩) + (Finset.mem_filter.mpr ⟨Finset.mem_univ k₂, hk₂⟩) + (ne_of_apply_ne Sigma.fst (ne_of_apply_ne Subtype.val (hik₁ ▸ hjk₂ ▸ hij))) + hxk₁ hxk₂ + +/-- Let the group `G` be the union of finitely many left cosets `g i • H i`. +Then the cosets of subgroups of infinite index may be omitted from the covering. -/ +@[to_additive] +theorem leftCoset_cover_filter_FiniteIndex + [DecidablePred (FiniteIndex : Subgroup G → Prop)] : + ⋃ k ∈ s.filter (fun i => (H i).FiniteIndex), g k • (H k : Set G) = Set.univ := + (leftCoset_cover_filter_FiniteIndex_aux hcovers).1 + +/-- Let the group `G` be the union of finitely many left cosets `g i • H i`. Then the +sum of the inverses of the indexes of the subgroups `H i` is greater than or equal to 1. -/ +@[to_additive one_le_sum_inv_index_of_leftCoset_cover] +theorem one_le_sum_inv_index_of_leftCoset_cover : + 1 ≤ ∑ i ∈ s, ((H i).index : ℚ)⁻¹ := + have := Classical.decPred (FiniteIndex : Subgroup G → Prop) + (leftCoset_cover_filter_FiniteIndex_aux hcovers).2.1 + +/-- Let the group `G` be the union of finitely many left cosets `g i • H i`. +If the sum of the inverses of the indexes of the subgroups `H i` is equal to 1, +then the cosets of the subgroups of finite index are pairwise disjoint. -/ +@[to_additive] +theorem pairwiseDisjoint_leftCoset_cover_of_sum_inv_index_eq_one + [DecidablePred (FiniteIndex : Subgroup G → Prop)] : + ∑ i ∈ s, ((H i).index : ℚ)⁻¹ = 1 → + Set.PairwiseDisjoint (s.filter (fun i => (H i).FiniteIndex)) + (fun i ↦ g i • (H i : Set G)) := + (leftCoset_cover_filter_FiniteIndex_aux hcovers).2.2 + +/-- B. H. Neumann Lemma : +If a finite family of cosets of subgroups covers the group, then at least one +of these subgroups has index not exceeding the number of cosets. -/ +@[to_additive] +theorem exists_index_le_card_of_leftCoset_cover : + ∃ i ∈ s, (H i).FiniteIndex ∧ (H i).index ≤ s.card := by + by_contra! h + apply (one_le_sum_inv_index_of_leftCoset_cover hcovers).not_lt + by_cases hs : s = ∅ + · simp only [hs, Finset.sum_empty, show (0 : ℚ) < 1 from rfl] + rw [← ne_eq, ← Finset.nonempty_iff_ne_empty] at hs + have hs' : 0 < s.card := Nat.pos_of_ne_zero (Finset.card_ne_zero.mpr hs) + have hlt : ∀ i ∈ s, ((H i).index : ℚ)⁻¹ < (s.card : ℚ)⁻¹ := fun i hi ↦ by + cases eq_or_ne (H i).index 0 with + | inl hindex => + rwa [hindex, Nat.cast_zero, inv_zero, inv_pos, Nat.cast_pos] + | inr hindex => + exact inv_lt_inv_of_lt (by exact_mod_cast hs') (by exact_mod_cast h i hi ⟨hindex⟩) + apply (Finset.sum_lt_sum_of_nonempty hs hlt).trans_eq + rw [Finset.sum_const, nsmul_eq_mul, mul_inv_cancel (Nat.cast_ne_zero.mpr hs'.ne')] + +end + +end Subgroup + +section Submodule + +variable {R M ι : Type*} [Ring R] [AddCommGroup M] [Module R M] + {p : ι → Submodule R M} {s : Finset ι} + (hcovers : ⋃ i ∈ s, (p i : Set M) = Set.univ) + +theorem Submodule.exists_finiteIndex_of_cover : + ∃ k ∈ s, (p k).toAddSubgroup.FiniteIndex := + have hcovers' : ⋃ i ∈ s, (0 : M) +ᵥ ((p i).toAddSubgroup : Set M) = Set.univ := by + simpa only [zero_vadd] using hcovers + AddSubgroup.exists_finiteIndex_of_leftCoset_cover hcovers' + +end Submodule + +section Subspace + +variable {k E ι : Type*} [DivisionRing k] [Infinite k] [AddCommGroup E] [Module k E] + {s : Finset (Subspace k E)} + +/- A vector space over an infinite field cannot be a finite union of proper subspaces. -/ +theorem Subspace.biUnion_ne_univ_of_ne_top (hs : ∀ p ∈ s, p ≠ ⊤) : + ⋃ p ∈ s, (p : Set E) ≠ Set.univ := by + intro hcovers + have ⟨p, hp, hfi⟩ := Submodule.exists_finiteIndex_of_cover hcovers + have : Finite (E ⧸ p) := AddSubgroup.finite_quotient_of_finiteIndex _ + have : Nontrivial (E ⧸ p) := Submodule.Quotient.nontrivial_of_lt_top p (hs p hp).lt_top + have : Infinite (E ⧸ p) := Module.Free.infinite k (E ⧸ p) + exact not_finite (E ⧸ p) + +/- A vector space over an infinite field cannot be a finite union of proper subspaces. -/ +theorem Subspace.exists_eq_top_of_biUnion_eq_univ (hcovers : ⋃ p ∈ s, (p : Set E) = Set.univ) : + ∃ p ∈ s, p = ⊤ := by + contrapose! hcovers + exact Subspace.biUnion_ne_univ_of_ne_top hcovers + +end Subspace diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 74ae9374289d5..90a3048bbb6de 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -299,9 +299,9 @@ theorem leftInvSeq_take (ω : List B) (j : ℕ) : lis (ω.take j) = (lis ω).take j := by obtain le | ge := Nat.le_or_ge j ω.length · simp only [leftInvSeq_eq_reverse_rightInvSeq_reverse] - rw [List.take_reverse j (by simpa)] + rw [List.take_reverse (by simpa)] nth_rw 1 [← List.reverse_reverse ω] - rw [List.take_reverse j (by simpa)] + rw [List.take_reverse (by simpa)] simp [rightInvSeq_drop] · rw [take_of_length_le ge, take_of_length_le (by simpa)] diff --git a/Mathlib/GroupTheory/FixedPointFree.lean b/Mathlib/GroupTheory/FixedPointFree.lean index aa93848d984de..49ae5c66a0f12 100644 --- a/Mathlib/GroupTheory/FixedPointFree.lean +++ b/Mathlib/GroupTheory/FixedPointFree.lean @@ -35,57 +35,61 @@ end Definitions namespace FixedPointFree -- todo: refactor Mathlib/Algebra/GroupPower/IterateHom to generalize φ to MonoidHomClass -variable [Group G] {φ : G →* G} (hφ : FixedPointFree φ) +variable [Group G] {φ : G →* G} -theorem commutatorMap_injective : Function.Injective (commutatorMap φ) := by +theorem commutatorMap_injective (hφ : FixedPointFree φ) : Function.Injective (commutatorMap φ) := by refine fun x y h ↦ inv_mul_eq_one.mp <| hφ _ ?_ rwa [map_mul, map_inv, eq_inv_mul_iff_mul_eq, ← mul_assoc, ← eq_div_iff_mul_eq', ← division_def] variable [Finite G] -theorem commutatorMap_surjective : Function.Surjective (commutatorMap φ) := +theorem commutatorMap_surjective (hφ : FixedPointFree φ) : Function.Surjective (commutatorMap φ) := Finite.surjective_of_injective hφ.commutatorMap_injective -theorem prod_pow_eq_one {n : ℕ} (hn : φ^[n] = _root_.id) (g : G) : +theorem prod_pow_eq_one (hφ : FixedPointFree φ) {n : ℕ} (hn : φ^[n] = _root_.id) (g : G) : ((List.range n).map (fun k ↦ φ^[k] g)).prod = 1 := by obtain ⟨g, rfl⟩ := commutatorMap_surjective hφ g simp only [commutatorMap_apply, iterate_map_div, ← Function.iterate_succ_apply] rw [List.prod_range_div', Function.iterate_zero_apply, hn, Function.id_def, div_self'] -theorem coe_eq_inv_of_sq_eq_one (h2 : φ^[2] = _root_.id) : ⇑φ = (·⁻¹) := by +theorem coe_eq_inv_of_sq_eq_one (hφ : FixedPointFree φ) (h2 : φ^[2] = _root_.id) : ⇑φ = (·⁻¹) := by ext g have key : 1 * g * φ g = 1 := hφ.prod_pow_eq_one h2 g rwa [one_mul, ← inv_eq_iff_mul_eq_one, eq_comm] at key section Involutive -variable (h2 : Function.Involutive φ) - -theorem coe_eq_inv_of_involutive : ⇑φ = (·⁻¹) := +theorem coe_eq_inv_of_involutive (hφ : FixedPointFree φ) (h2 : Function.Involutive φ) : + ⇑φ = (·⁻¹) := coe_eq_inv_of_sq_eq_one hφ (funext h2) -theorem commute_all_of_involutive (g h : G) : Commute g h := by +theorem commute_all_of_involutive (hφ : FixedPointFree φ) (h2 : Function.Involutive φ) (g h : G) : + Commute g h := by have key := map_mul φ g h rwa [hφ.coe_eq_inv_of_involutive h2, inv_eq_iff_eq_inv, mul_inv_rev, inv_inv, inv_inv] at key /-- If a finite group admits a fixed-point-free involution, then it is commutative. -/ -def commGroupOfInvolutive : CommGroup G := .mk (hφ.commute_all_of_involutive h2) +def commGroupOfInvolutive (hφ : FixedPointFree φ) (h2 : Function.Involutive φ): + CommGroup G := .mk (hφ.commute_all_of_involutive h2) -theorem orderOf_ne_two_of_involutive (g : G) : orderOf g ≠ 2 := by +theorem orderOf_ne_two_of_involutive (hφ : FixedPointFree φ) (h2 : Function.Involutive φ) (g : G) : + orderOf g ≠ 2 := by intro hg have key : φ g = g := by rw [hφ.coe_eq_inv_of_involutive h2, inv_eq_iff_mul_eq_one, ← sq, ← hg, pow_orderOf_eq_one] rw [hφ g key, orderOf_one] at hg contradiction -theorem odd_card_of_involutive : Odd (Nat.card G) := by +theorem odd_card_of_involutive (hφ : FixedPointFree φ) (h2 : Function.Involutive φ) : + Odd (Nat.card G) := by have := Fintype.ofFinite G by_contra h rw [← Nat.even_iff_not_odd, even_iff_two_dvd, Nat.card_eq_fintype_card] at h obtain ⟨g, hg⟩ := exists_prime_orderOf_dvd_card 2 h exact hφ.orderOf_ne_two_of_involutive h2 g hg -theorem odd_orderOf_of_involutive (g : G) : Odd (orderOf g) := +theorem odd_orderOf_of_involutive (hφ : FixedPointFree φ) (h2 : Function.Involutive φ) (g : G) : + Odd (orderOf g) := Odd.of_dvd_nat (hφ.odd_card_of_involutive h2) (orderOf_dvd_natCard g) end Involutive diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index 3646f9a6cdb90..eca08fb36875e 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -168,9 +168,9 @@ theorem Step.diamond_aux : | [], _, [(x3, b3)], _, _, _, _, _, H => by injections; subst_vars; simp | [(x3, b3)], _, [], _, _, _, _, _, H => by injections; subst_vars; simp | [], _, (x3, b3) :: (x4, b4) :: tl, _, _, _, _, _, H => by - injections; subst_vars; simp; right; exact ⟨_, Red.Step.not, Red.Step.cons_not⟩ + injections; subst_vars; right; exact ⟨_, Red.Step.not, Red.Step.cons_not⟩ | (x3, b3) :: (x4, b4) :: tl, _, [], _, _, _, _, _, H => by - injections; subst_vars; simp; right; exact ⟨_, Red.Step.cons_not, Red.Step.not⟩ + injections; subst_vars; right; simpa using ⟨_, Red.Step.cons_not, Red.Step.not⟩ | (x3, b3) :: tl, _, (x4, b4) :: tl2, _, _, _, _, _, H => let ⟨H1, H2⟩ := List.cons.inj H match Step.diamond_aux H2 with @@ -620,8 +620,8 @@ theorem lift.of_eq (x : FreeGroup α) : lift FreeGroup.of x = x := theorem lift.range_le {s : Subgroup β} (H : Set.range f ⊆ s) : (lift f).range ≤ s := by rintro _ ⟨⟨L⟩, rfl⟩; exact List.recOn L s.one_mem fun ⟨x, b⟩ tl ih ↦ - Bool.recOn b (by simp at ih ⊢; exact s.mul_mem (s.inv_mem <| H ⟨x, rfl⟩) ih) - (by simp at ih ⊢; exact s.mul_mem (H ⟨x, rfl⟩) ih) + Bool.recOn b (by simpa using s.mul_mem (s.inv_mem <| H ⟨x, rfl⟩) ih) + (by simpa using s.mul_mem (H ⟨x, rfl⟩) ih) @[to_additive] theorem lift.range_eq_closure : (lift f).range = Subgroup.closure (Set.range f) := by diff --git a/Mathlib/GroupTheory/GroupAction/Group.lean b/Mathlib/GroupTheory/GroupAction/Group.lean index 5f65e01021480..600fba33cc68c 100644 --- a/Mathlib/GroupTheory/GroupAction/Group.lean +++ b/Mathlib/GroupTheory/GroupAction/Group.lean @@ -3,11 +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 -/ -import Mathlib.Algebra.Group.Aut import Mathlib.Algebra.Group.Action.Basic +import Mathlib.Algebra.Group.Aut import Mathlib.Algebra.Group.Invertible.Basic +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.GroupWithZero.Units.Basic -import Mathlib.GroupTheory.GroupAction.Units /-! # Group actions applied to various types of group diff --git a/Mathlib/GroupTheory/GroupAction/IterateAct.lean b/Mathlib/GroupTheory/GroupAction/IterateAct.lean new file mode 100644 index 0000000000000..9259861a738b6 --- /dev/null +++ b/Mathlib/GroupTheory/GroupAction/IterateAct.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Algebra.Ring.Nat +import Mathlib.Algebra.Group.Action.Defs + +/-! +# Monoid action by iterates of a map + +In this file we define `IterateMulAct f`, `f : α → α`, as a one field structure wrapper over `ℕ` +that acts on `α` by iterates of `f`, `⟨n⟩ • x = f^[n] x`. + +It is useful to convert between definitions and theorems about maps and monoid actions. +-/ + +/-- A structure with a single field `val : ℕ` +that additively acts on `α` by `⟨n⟩ +ᵥ x = f^[n] x`. -/ +structure IterateAddAct {α : Type*} (f : α → α) where + /-- The value of `n : IterateAddAct f`. -/ + val : ℕ + +/-- A structure with a single field `val : ℕ` that acts on `α` by `⟨n⟩ • x = f^[n] x`. -/ +@[to_additive (attr := ext)] +structure IterateMulAct {α : Type*} (f : α → α) where + /-- The value of `n : IterateMulAct f`. -/ + val : ℕ + +namespace IterateMulAct + +variable {α : Type*} {f : α → α} + +@[to_additive] +instance instCommMonoid : CommMonoid (IterateMulAct f) where + one := ⟨0⟩ + mul m n := ⟨m.1 + n.1⟩ + mul_assoc a b c := by ext; apply add_assoc + one_mul _ := by ext; apply zero_add + mul_one _ := rfl + mul_comm _ _ := by ext; apply add_comm + npow n a := ⟨n * a.val⟩ + npow_zero _ := by ext; apply zero_mul + npow_succ n a := by ext; apply Nat.succ_mul + +@[to_additive] +instance instMulAction : MulAction (IterateMulAct f) α where + smul n x := f^[n.val] x + one_smul _ := rfl + mul_smul _ _ := Function.iterate_add_apply f _ _ + +@[to_additive (attr := simp)] +theorem mk_smul (n : ℕ) (x : α) : mk (f := f) n • x = f^[n] x := rfl + +end IterateMulAct diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index c300d2371ccc6..eaf0e222c3698 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -645,8 +645,8 @@ theorem exists_normalWord_prod_eq have : a.fst = -a.fst := by have hl : l ≠ [] := by rintro rfl; simp_all have : a.fst = (l.head hl).fst := (List.chain'_cons'.1 chain).1 (l.head hl) - (List.head?_eq_head _ _) hS - rwa [List.head?_eq_head _ hl, Option.map_some', ← this, Option.some_inj] at hx' + (List.head?_eq_head _) hS + rwa [List.head?_eq_head hl, Option.map_some', ← this, Option.some_inj] at hx' simp at this erw [List.map_cons, mul_smul, of_smul_eq_smul, NormalWord.group_smul_def, t_pow_smul_eq_unitsSMul, unitsSMul, dif_neg this, ← hw'2] diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 1434a0a477074..4dfb748379a3c 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -419,6 +419,23 @@ instance : FiniteIndex (⊤ : Subgroup G) := instance [FiniteIndex H] [FiniteIndex K] : FiniteIndex (H ⊓ K) := ⟨index_inf_ne_zero FiniteIndex.finiteIndex FiniteIndex.finiteIndex⟩ +@[to_additive] +theorem finiteIndex_iInf {ι : Type*} [Finite ι] {f : ι → Subgroup G} + (hf : ∀ i, (f i).FiniteIndex) : (⨅ i, f i).FiniteIndex := + ⟨index_iInf_ne_zero fun i => (hf i).finiteIndex⟩ + +@[to_additive] +theorem finiteIndex_iInf' {ι : Type*} {s : Finset ι} + (f : ι → Subgroup G) (hs : ∀ i ∈ s, (f i).FiniteIndex) : + (⨅ i ∈ s, f i).FiniteIndex := by + rw [iInf_subtype'] + exact finiteIndex_iInf fun ⟨i, hi⟩ => hs i hi + +@[to_additive] +instance instFiniteIndex_subgroupOf (H K : Subgroup G) [H.FiniteIndex] : + (H.subgroupOf K).FiniteIndex := + ⟨fun h => H.index_ne_zero_of_finite <| H.index_eq_zero_of_relindex_eq_zero h⟩ + variable {H K} @[to_additive] diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 8eaed2c63a426..8089e5d60ce04 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -503,7 +503,7 @@ theorem map_right_cancel {x y} {c : S} (h : f.toMap (c * x) = f.toMap (c * y)) : @[to_additive] theorem map_left_cancel {x y} {c : S} (h : f.toMap (x * c) = f.toMap (y * c)) : f.toMap x = f.toMap y := - f.map_right_cancel <| by rw [mul_comm _ x, mul_comm _ y, h] + f.map_right_cancel (c := c) <| by rw [mul_comm _ x, mul_comm _ y, h] /-- Given a localization map `f : M →* N`, the surjection sending `(x, y) : M × S` to `f x * (f y)⁻¹`. -/ diff --git a/Mathlib/GroupTheory/NoncommPiCoprod.lean b/Mathlib/GroupTheory/NoncommPiCoprod.lean index 198c33f5d3918..46cb605a70f01 100644 --- a/Mathlib/GroupTheory/NoncommPiCoprod.lean +++ b/Mathlib/GroupTheory/NoncommPiCoprod.lean @@ -167,16 +167,16 @@ variable {G : Type*} [Group G] variable {ι : Type*} [hdec : DecidableEq ι] [hfin : Fintype ι] variable {H : ι → Type*} [∀ i, Group (H i)] variable (ϕ : ∀ i : ι, H i →* G) -variable {hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)} -- We use `f` and `g` to denote elements of `Π (i : ι), H i` variable (f g : ∀ i : ι, H i) namespace MonoidHom - -- The subgroup version of `MonoidHom.noncommPiCoprod_mrange` @[to_additive] -theorem noncommPiCoprod_range : (noncommPiCoprod ϕ hcomm).range = ⨆ i : ι, (ϕ i).range := by +theorem noncommPiCoprod_range + {hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)} : + (noncommPiCoprod ϕ hcomm).range = ⨆ i : ι, (ϕ i).range := by letI := Classical.decEq ι apply le_antisymm · rintro x ⟨f, rfl⟩ @@ -191,6 +191,7 @@ theorem noncommPiCoprod_range : (noncommPiCoprod ϕ hcomm).range = ⨆ i : ι, ( @[to_additive] theorem injective_noncommPiCoprod_of_independent + {hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)} (hind : CompleteLattice.Independent fun i => (ϕ i).range) (hinj : ∀ i, Function.Injective (ϕ i)) : Function.Injective (noncommPiCoprod ϕ hcomm) := by classical @@ -204,10 +205,10 @@ theorem injective_noncommPiCoprod_of_independent apply hinj simp [this i (Finset.mem_univ i)] -variable (hcomm) - @[to_additive] -theorem independent_range_of_coprime_order [Finite ι] [∀ i, Fintype (H i)] +theorem independent_range_of_coprime_order + (hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)) + [Finite ι] [∀ i, Fintype (H i)] (hcoprime : Pairwise fun i j => Nat.Coprime (Fintype.card (H i)) (Fintype.card (H j))) : CompleteLattice.Independent fun i => (ϕ i).range := by cases nonempty_fintype ι @@ -256,10 +257,12 @@ variable (f g : ∀ i : ι, H i) section CommutingSubgroups -- We assume that the elements of different subgroups commute -variable (hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y) +-- with `hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y` @[to_additive] -theorem commute_subtype_of_commute (i j : ι) (hne : i ≠ j) : +theorem commute_subtype_of_commute + (hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y) (i j : ι) + (hne : i ≠ j) : ∀ (x : H i) (y : H j), Commute ((H i).subtype x) ((H j).subtype y) := by rintro ⟨x, hx⟩ ⟨y, hy⟩ exact hcomm hne x y hx hy @@ -268,31 +271,35 @@ theorem commute_subtype_of_commute (i j : ι) (hne : i ≠ j) : commute -/ @[to_additive "The canonical homomorphism from a family of additive subgroups where elements from different subgroups commute"] -def noncommPiCoprod : (∀ i : ι, H i) →* G := +def noncommPiCoprod (hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y) : + (∀ i : ι, H i) →* G := MonoidHom.noncommPiCoprod (fun i => (H i).subtype) (commute_subtype_of_commute hcomm) -variable {hcomm} - @[to_additive (attr := simp)] -theorem noncommPiCoprod_mulSingle (i : ι) (y : H i) : +theorem noncommPiCoprod_mulSingle + {hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y}(i : ι) (y : H i) : noncommPiCoprod hcomm (Pi.mulSingle i y) = y := by apply MonoidHom.noncommPiCoprod_mulSingle @[to_additive] -theorem noncommPiCoprod_range : (noncommPiCoprod hcomm).range = ⨆ i : ι, H i := by +theorem noncommPiCoprod_range + {hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y} : + (noncommPiCoprod hcomm).range = ⨆ i : ι, H i := by simp [noncommPiCoprod, MonoidHom.noncommPiCoprod_range] @[to_additive] -theorem injective_noncommPiCoprod_of_independent (hind : CompleteLattice.Independent H) : +theorem injective_noncommPiCoprod_of_independent + {hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y} + (hind : CompleteLattice.Independent H) : Function.Injective (noncommPiCoprod hcomm) := by apply MonoidHom.injective_noncommPiCoprod_of_independent · simpa using hind · intro i exact Subtype.coe_injective -variable (hcomm) - @[to_additive] -theorem independent_of_coprime_order [Finite ι] [∀ i, Fintype (H i)] +theorem independent_of_coprime_order + (hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y) + [Finite ι] [∀ i, Fintype (H i)] (hcoprime : Pairwise fun i j => Nat.Coprime (Fintype.card (H i)) (Fintype.card (H j))) : CompleteLattice.Independent H := by simpa using diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 99d46bf66e186..9e37e57b20ea2 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -369,15 +369,17 @@ lemma IsOfFinOrder.finite_powers (ha : IsOfFinOrder a) : (powers a : Set G).Fini namespace Commute -variable {x} (h : Commute x y) +variable {x} @[to_additive] -theorem orderOf_mul_dvd_lcm : orderOf (x * y) ∣ Nat.lcm (orderOf x) (orderOf y) := by +theorem orderOf_mul_dvd_lcm (h : Commute x y) : + orderOf (x * y) ∣ Nat.lcm (orderOf x) (orderOf y) := by rw [orderOf, ← comp_mul_left] exact Function.Commute.minimalPeriod_of_comp_dvd_lcm h.function_commute_mul_left @[to_additive] -theorem orderOf_dvd_lcm_mul : orderOf y ∣ Nat.lcm (orderOf x) (orderOf (x * y)) := by +theorem orderOf_dvd_lcm_mul (h : Commute x y): + orderOf y ∣ Nat.lcm (orderOf x) (orderOf (x * y)) := by by_cases h0 : orderOf x = 0 · rw [h0, lcm_zero_left] apply dvd_zero @@ -389,18 +391,20 @@ theorem orderOf_dvd_lcm_mul : orderOf y ∣ Nat.lcm (orderOf x) (orderOf (x * y) (lcm_dvd_iff.2 ⟨(orderOf_pow_dvd _).trans (dvd_lcm_left _ _), dvd_lcm_right _ _⟩) @[to_additive addOrderOf_add_dvd_mul_addOrderOf] -theorem orderOf_mul_dvd_mul_orderOf : orderOf (x * y) ∣ orderOf x * orderOf y := +theorem orderOf_mul_dvd_mul_orderOf (h : Commute x y): + orderOf (x * y) ∣ orderOf x * orderOf y := dvd_trans h.orderOf_mul_dvd_lcm (lcm_dvd_mul _ _) @[to_additive addOrderOf_add_eq_mul_addOrderOf_of_coprime] -theorem orderOf_mul_eq_mul_orderOf_of_coprime (hco : (orderOf x).Coprime (orderOf y)) : - orderOf (x * y) = orderOf x * orderOf y := by +theorem orderOf_mul_eq_mul_orderOf_of_coprime (h : Commute x y) + (hco : (orderOf x).Coprime (orderOf y)) : orderOf (x * y) = orderOf x * orderOf y := by rw [orderOf, ← comp_mul_left] exact h.function_commute_mul_left.minimalPeriod_of_comp_eq_mul_of_coprime hco /-- Commuting elements of finite order are closed under multiplication. -/ @[to_additive "Commuting elements of finite additive order are closed under addition."] -theorem isOfFinOrder_mul (hx : IsOfFinOrder x) (hy : IsOfFinOrder y) : IsOfFinOrder (x * y) := +theorem isOfFinOrder_mul (h : Commute x y) (hx : IsOfFinOrder x) (hy : IsOfFinOrder y) : + IsOfFinOrder (x * y) := orderOf_pos_iff.mp <| pos_of_dvd_of_pos h.orderOf_mul_dvd_mul_orderOf <| mul_pos hx.orderOf_pos hy.orderOf_pos @@ -410,7 +414,7 @@ theorem isOfFinOrder_mul (hx : IsOfFinOrder x) (hy : IsOfFinOrder y) : IsOfFinOr "If each prime factor of `addOrderOf x` has higher multiplicity in `addOrderOf y`, and `x` commutes with `y`, then `x + y` has the same order as `y`."] -theorem orderOf_mul_eq_right_of_forall_prime_mul_dvd (hy : IsOfFinOrder y) +theorem orderOf_mul_eq_right_of_forall_prime_mul_dvd (h : Commute x y) (hy : IsOfFinOrder y) (hdvd : ∀ p : ℕ, p.Prime → p ∣ orderOf x → p * orderOf x ∣ orderOf y) : orderOf (x * y) = orderOf y := by have hoy := hy.orderOf_pos diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 26f7d3425006a..509c1209b4f26 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -181,7 +181,7 @@ nonrec theorem formPerm_eq_formPerm_iff {α : Type*} [DecidableEq α] {s s' : Cy apply @Quotient.inductionOn₂' _ _ _ _ _ s s' intro l l' -- Porting note: was `simpa using formPerm_eq_formPerm_iff` - simp_all + simp only [mk''_eq_coe, nodup_coe_iff, formPerm_coe, coe_eq_coe, length_coe] intro hs hs' constructor <;> intro h <;> simp_all only [formPerm_eq_formPerm_iff] diff --git a/Mathlib/GroupTheory/Perm/List.lean b/Mathlib/GroupTheory/Perm/List.lean index ee692803b41a6..76f10c832bac0 100644 --- a/Mathlib/GroupTheory/Perm/List.lean +++ b/Mathlib/GroupTheory/Perm/List.lean @@ -68,7 +68,7 @@ theorem mem_or_mem_of_zipWith_swap_prod_ne : ∀ {l l' : List α} {x : α}, | _, [], _ => by simp | a::l, b::l', x => fun hx ↦ if h : (zipWith swap l l').prod x = x then - (eq_or_eq_of_swap_apply_ne_self (by simpa [h] using hx)).imp + (eq_or_eq_of_swap_apply_ne_self (a := a) (b := b) (x := x) (by simpa [h] using hx)).imp (by rintro rfl; exact .head _) (by rintro rfl; exact .head _) else (mem_or_mem_of_zipWith_swap_prod_ne h).imp (.tail _) (.tail _) diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 5174a101c08dc..851351d1966d0 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -317,12 +317,12 @@ theorem eq_one_of_smul_normalized (w : CoprodI.Word G) {i : ι} (h : H) dsimp split_ifs with hep · rcases hep with ⟨hnil, rfl⟩ - rw [head?_eq_head _ hnil] + rw [head?_eq_head hnil] simp_all · push_neg at hep by_cases hw : w.toList = [] · simp [hw, Word.fstIdx] - · simp [head?_eq_head _ hw, Word.fstIdx, hep hw] + · simp [head?_eq_head hw, Word.fstIdx, hep hw] theorem ext_smul {w₁ w₂ : NormalWord d} (i : ι) (h : CoprodI.of (φ i w₁.head) • w₁.toWord = diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index e1fb516ea4b68..c5c111db89a9f 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -43,7 +43,6 @@ open Function open scoped Pointwise universe u v w x - namespace QuotientGroup variable {G : Type u} [Group G] (N : Subgroup G) [nN : N.Normal] {H : Type v} [Group H] @@ -133,7 +132,7 @@ theorem eq_iff_div_mem {N : Subgroup G} [nN : N.Normal] {x y : G} : @[to_additive] instance Quotient.commGroup {G : Type*} [CommGroup G] (N : Subgroup G) : CommGroup (G ⧸ N) := - { toGroup := @QuotientGroup.Quotient.group _ _ N N.normal_of_comm + { toGroup := have := N.normal_of_comm; QuotientGroup.Quotient.group N mul_comm := fun a b => Quotient.inductionOn₂' a b fun a b => congr_arg mk (mul_comm a b) } local notation " Q " => G ⧸ N diff --git a/Mathlib/GroupTheory/Subgroup/Center.lean b/Mathlib/GroupTheory/Subgroup/Center.lean index f1eab1e134fe1..0f10d4480d50a 100644 --- a/Mathlib/GroupTheory/Subgroup/Center.lean +++ b/Mathlib/GroupTheory/Subgroup/Center.lean @@ -86,7 +86,7 @@ theorem _root_.CommGroup.center_eq_top {G : Type*} [CommGroup G] : center G = /-- A group is commutative if the center is the whole group -/ def _root_.Group.commGroupOfCenterEqTop (h : center G = ⊤) : CommGroup G := - { (_ : Group G) with + { ‹Group G› with mul_comm := by rw [eq_top_iff'] at h intro x y diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index eaa5bbba30f3b..6eca87ed01fd9 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -138,7 +138,7 @@ theorem IsTorsion.module_of_torsion [Semiring R] [Module R M] (tR : IsTorsion R) fun f => isOfFinAddOrder_iff_nsmul_eq_zero.mpr <| by obtain ⟨n, npos, hn⟩ := (tR 1).exists_nsmul_eq_zero - exact ⟨n, npos, by simp only [nsmul_eq_smul_cast R _ f, ← nsmul_one, hn, zero_smul]⟩ + exact ⟨n, npos, by simp only [← Nat.cast_smul_eq_nsmul R _ f, ← nsmul_one, hn, zero_smul]⟩ /-- A module with a finite ring of scalars is additively torsion. -/ theorem IsTorsion.module_of_finite [Ring R] [Finite R] [Module R M] : IsTorsion M := diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index a4f2df9da8705..fcfd842450b0b 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -131,7 +131,7 @@ theorem transfer_eq_pow_aux (g : G) letI := fintypeOfIndexNeZero hH classical replace key : ∀ (k : ℕ) (g₀ : G), g₀⁻¹ * g ^ k * g₀ ∈ H → g ^ k ∈ H := fun k g₀ hk => - (_root_.congr_arg (· ∈ H) (key k g₀ hk)).mp hk + (congr_arg (· ∈ H) (key k g₀ hk)).mp hk replace key : ∀ q : G ⧸ H, g ^ Function.minimalPeriod (g • ·) q ∈ H := fun q => key (Function.minimalPeriod (g • ·) q) q.out' (QuotientGroup.out'_conj_pow_minimalPeriod_mem H g q) diff --git a/Mathlib/Init/Data/Nat/GCD.lean b/Mathlib/Init/Data/Nat/GCD.lean deleted file mode 100644 index 74154704c6ff9..0000000000000 --- a/Mathlib/Init/Data/Nat/GCD.lean +++ /dev/null @@ -1,32 +0,0 @@ -/- -Copyright (c) 2014 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro - --/ - -import Batteries.Data.Nat.Gcd -import Mathlib.Data.Nat.Notation - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Definitions and properties of gcd, lcm, and coprime --/ - - -open WellFounded - -namespace Nat - -/-! gcd -/ - -theorem gcd_def (x y : ℕ) : gcd x y = if x = 0 then y else gcd (y % x) x := by - cases x <;> simp [Nat.gcd_succ] - -end Nat diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 15b0661f11cb5..110da1dca9c07 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -497,7 +497,8 @@ theorem pointReflection_fixed_iff_of_injective_bit0 {x y : P₁} (h : Injective Equiv.pointReflection_fixed_iff_of_injective_bit0 h set_option linter.deprecated false in -theorem injective_pointReflection_left_of_injective_bit0 (h : Injective (2 • · : V₁ → V₁)) (y : P₁) : +theorem injective_pointReflection_left_of_injective_bit0 + (h : Injective (2 • · : V₁ → V₁)) (y : P₁) : Injective fun x : P₁ => pointReflection k x y := Equiv.injective_pointReflection_left_of_injective_bit0 h y diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index fcf0454fdf1b2..d838fd0ac45a3 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -108,11 +108,11 @@ theorem coe_mk (f : MultilinearMap R (fun _ : ι => M) N) (h) : ⇑(⟨f, h⟩ : M [⋀^ι]→ₗ[R] N) = f := rfl -theorem congr_fun {f g : M [⋀^ι]→ₗ[R] N} (h : f = g) (x : ι → M) : f x = g x := +protected theorem congr_fun {f g : M [⋀^ι]→ₗ[R] N} (h : f = g) (x : ι → M) : f x = g x := congr_arg (fun h : M [⋀^ι]→ₗ[R] N => h x) h -theorem congr_arg (f : M [⋀^ι]→ₗ[R] N) {x y : ι → M} (h : x = y) : f x = f y := - _root_.congr_arg (fun x : ι → M => f x) h +protected theorem congr_arg (f : M [⋀^ι]→ₗ[R] N) {x y : ι → M} (h : x = y) : f x = f y := + congr_arg (fun x : ι → M => f x) h theorem coe_injective : Injective ((↑) : M [⋀^ι]→ₗ[R] N → (ι → M) → N) := DFunLike.coe_injective diff --git a/Mathlib/LinearAlgebra/Basis.lean b/Mathlib/LinearAlgebra/Basis.lean index d180330f27d28..6e29097ca05f5 100644 --- a/Mathlib/LinearAlgebra/Basis.lean +++ b/Mathlib/LinearAlgebra/Basis.lean @@ -316,6 +316,42 @@ theorem coe_map : (b.map f : ι → M') = f ∘ b := end Map +section SMul +variable {G G'} +variable [Group G] [Group G'] +variable [DistribMulAction G M] [DistribMulAction G' M] +variable [SMulCommClass G R M] [SMulCommClass G' R M] + +/-- The action on a `Basis` by acting on each element. + +See also `Basis.unitsSMul` and `Basis.groupSMul`, for the cases when a different action is applied +to each basis element. -/ +instance : SMul G (Basis ι R M) where + smul g b := b.map <| DistribMulAction.toLinearEquiv _ _ g + +@[simp] +theorem smul_apply (g : G) (b : Basis ι R M) (i : ι) : (g • b) i = g • b i := rfl + +@[norm_cast] theorem coe_smul (g : G) (b : Basis ι R M) : ⇑(g • b) = g • ⇑b := rfl + +/-- When the group in question is the automorphisms, `•` coincides with `Basis.map`. -/ +@[simp] +theorem smul_eq_map (g : M ≃ₗ[R] M) (b : Basis ι R M) : g • b = b.map g := rfl + +@[simp] theorem repr_smul (g : G) (b : Basis ι R M) : + (g • b).repr = (DistribMulAction.toLinearEquiv _ _ g).symm.trans b.repr := rfl + +instance : MulAction G (Basis ι R M) := + Function.Injective.mulAction _ DFunLike.coe_injective coe_smul + +instance [SMulCommClass G G' M] : SMulCommClass G G' (Basis ι R M) where + smul_comm _g _g' _b := DFunLike.ext _ _ fun _ => smul_comm _ _ _ + +instance [SMul G G'] [IsScalarTower G G' M] : IsScalarTower G G' (Basis ι R M) where + smul_assoc _g _g' _b := DFunLike.ext _ _ fun _ => smul_assoc _ _ _ + +end SMul + section MapCoeffs variable {R' : Type*} [Semiring R'] [Module R' M] (f : R ≃+* R') diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index 5eea34374dbf5..984b62e9b92a4 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -484,6 +484,11 @@ theorem Basis.det_isEmpty [IsEmpty ι] : e.det = AlternatingMap.constOfIsEmpty R /-- `Basis.det` is not the zero map. -/ theorem Basis.det_ne_zero [Nontrivial R] : e.det ≠ 0 := fun h => by simpa [h] using e.det_self +theorem Basis.smul_det {G} [Group G] [DistribMulAction G M] [SMulCommClass G R M] + (g : G) (v : ι → M) : + (g • e).det v = e.det (g⁻¹ • v) := by + simp_rw [det_apply, toMatrix_smul_left] + theorem is_basis_iff_det {v : ι → M} : LinearIndependent R v ∧ span R (Set.range v) = ⊤ ↔ IsUnit (e.det v) := by constructor diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 2dbf983d46ea4..8b398a6d36163 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -330,8 +330,9 @@ end theorem restrictDom_comp_subtype (s : Set α) [DecidablePred (· ∈ s)] : (restrictDom M R s).comp (Submodule.subtype _) = LinearMap.id := by ext l a - by_cases h : a ∈ s <;> simp [h] - exact ((mem_supported' R l.1).1 l.2 a h).symm + by_cases h : a ∈ s + · simp [h] + simpa [h] using ((mem_supported' R l.1).1 l.2 a h).symm theorem range_restrictDom (s : Set α) [DecidablePred (· ∈ s)] : LinearMap.range (restrictDom M R s) = ⊤ := diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 3ae3fad6d7929..9420ed64bedfb 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1196,7 +1196,7 @@ open Submodule (instead of a data containing type class) -/ theorem mem_span_insert_exchange : x ∈ span K (insert y s) → x ∉ span K s → y ∈ span K (insert x s) := by - simp [mem_span_insert] + simp only [mem_span_insert, forall_exists_index, and_imp] rintro a z hz rfl h refine ⟨a⁻¹, -a⁻¹ • z, smul_mem _ _ hz, ?_⟩ have a0 : a ≠ 0 := by @@ -1453,7 +1453,7 @@ theorem exists_of_linearIndependent_of_finite_span {t : Finset V} theorem exists_finite_card_le_of_finite_of_linearIndependent_of_span (ht : t.Finite) (hs : LinearIndependent K (fun x => x : s → V)) (hst : s ⊆ span K t) : ∃ h : s.Finite, h.toFinset.card ≤ ht.toFinset.card := - have : s ⊆ (span K ↑ht.toFinset : Submodule K V) := by simp; assumption + have : s ⊆ (span K ↑ht.toFinset : Submodule K V) := by simpa let ⟨u, _hust, hsu, Eq⟩ := exists_of_linearIndependent_of_finite_span hs this have : s.Finite := u.finite_toSet.subset hsu ⟨this, by rw [← Eq]; exact Finset.card_le_card <| Finset.coe_subset.mp <| by simp [hsu]⟩ diff --git a/Mathlib/LinearAlgebra/Matrix/Basis.lean b/Mathlib/LinearAlgebra/Matrix/Basis.lean index 7d9a711e995e3..6c0ce89d7f42e 100644 --- a/Mathlib/LinearAlgebra/Matrix/Basis.lean +++ b/Mathlib/LinearAlgebra/Matrix/Basis.lean @@ -98,6 +98,9 @@ theorem toMatrix_isUnitSMul [DecidableEq ι] (e : Basis ι R₂ M₂) {w : ι (hw : ∀ i, IsUnit (w i)) : e.toMatrix (e.isUnitSMul hw) = diagonal w := e.toMatrix_unitsSMul _ +theorem toMatrix_smul_left {G} [Group G] [DistribMulAction G M] [SMulCommClass G R M] (g : G) : + (g • e).toMatrix v = e.toMatrix (g⁻¹ • v) := rfl + @[simp] theorem sum_toMatrix_smul_self [Fintype ι] : ∑ i : ι, e.toMatrix v i j • e i = v j := by simp_rw [e.toMatrix_apply, e.sum_repr] diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index 3a33ad4e62648..4001c12154234 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -6,7 +6,7 @@ Authors: Jon Bannon, Jireh Loreaux import Mathlib.LinearAlgebra.Matrix.Spectrum import Mathlib.LinearAlgebra.Eigenspace.Matrix -import Mathlib.Topology.ContinuousFunction.UniqueCFC +import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique import Mathlib.Topology.ContinuousFunction.Units /-! diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index a5fc4be55ebb3..745b0e49846d7 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -298,6 +298,10 @@ theorem mul_inv_eq_iff_eq_mul_of_invertible (A B C : Matrix n n α) [Invertible ⟨fun h => by rw [← h, inv_mul_cancel_right_of_invertible], fun h => by rw [h, mul_inv_cancel_right_of_invertible]⟩ +lemma inv_mulVec_eq_vec {A : Matrix n n α} [Invertible A] + {u v : n → α} (hM : u = A.mulVec v) : A⁻¹.mulVec u = v := by + rw [hM, Matrix.mulVec_mulVec, Matrix.inv_mul_of_invertible, Matrix.one_mulVec] + lemma mul_right_injective_of_invertible [Invertible A] : Function.Injective (fun (x : Matrix n m α) => A * x) := fun _ _ h => by simpa only [inv_mul_cancel_left_of_invertible] using congr_arg (A⁻¹ * ·) h diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index 4c34a4d035d80..e14a9a860b1f3 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -377,7 +377,8 @@ theorem Matrix.toLinearMap₂_symm : (LinearMap.toMatrix₂ b₁ b₂).symm_symm theorem Matrix.toLinearMap₂_basisFun : - Matrix.toLinearMap₂ (Pi.basisFun R n) (Pi.basisFun R m) = Matrix.toLinearMap₂' R (N₂ := N₂) := by + Matrix.toLinearMap₂ (Pi.basisFun R n) (Pi.basisFun R m) = + Matrix.toLinearMap₂' R (N₂ := N₂) := by ext M simp only [coe_comp, coe_single, Function.comp_apply, toLinearMap₂_apply, Pi.basisFun_repr, toLinearMap₂'_apply] diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index a7ebeafb101e9..23e3a7cb7ae7f 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -7,7 +7,7 @@ import Mathlib.Analysis.InnerProductSpace.Spectrum import Mathlib.Data.Matrix.Rank import Mathlib.LinearAlgebra.Matrix.Diagonal import Mathlib.LinearAlgebra.Matrix.Hermitian -import Mathlib.Analysis.NormedSpace.Star.Matrix +import Mathlib.Analysis.CstarAlgebra.Matrix import Mathlib.Topology.Algebra.Module.FiniteDimension /-! # Spectral theory of hermitian matrices diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index e6128306df2a9..a141f343332d0 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -586,6 +586,22 @@ theorem LinearMap.toMatrix_basis_equiv [Fintype l] [DecidableEq l] (b : Basis l ext i j simp [LinearMap.toMatrix_apply, Matrix.one_apply, Finsupp.single_apply, eq_comm] +theorem LinearMap.toMatrix_smulBasis_left {G} [Group G] [DistribMulAction G M₁] + [SMulCommClass G R M₁] (g : G) (f : M₁ →ₗ[R] M₂) : + LinearMap.toMatrix (g • v₁) v₂ f = + LinearMap.toMatrix v₁ v₂ (f ∘ₗ DistribMulAction.toLinearMap _ _ g) := by + ext + rw [LinearMap.toMatrix_apply, LinearMap.toMatrix_apply] + dsimp + +theorem LinearMap.toMatrix_smulBasis_right {G} [Group G] [DistribMulAction G M₂] + [SMulCommClass G R M₂] (g : G) (f : M₁ →ₗ[R] M₂) : + LinearMap.toMatrix v₁ (g • v₂) f = + LinearMap.toMatrix v₁ v₂ (DistribMulAction.toLinearMap _ _ g⁻¹ ∘ₗ f) := by + ext + rw [LinearMap.toMatrix_apply, LinearMap.toMatrix_apply] + dsimp + end Finite variable {R : Type*} [CommSemiring R] @@ -835,6 +851,15 @@ theorem leftMulMatrix_injective : Function.Injective (leftMulMatrix b) := fun x _ = Algebra.lmul R S x' 1 := by rw [(LinearMap.toMatrix b b).injective h] _ = x' := mul_one x' +@[simp] +theorem smul_leftMulMatrix {G} [Group G] [DistribMulAction G S] + [SMulCommClass G R S] [SMulCommClass G S S] (g : G) (x) : + leftMulMatrix (g • b) x = leftMulMatrix b x := by + ext + simp_rw [leftMulMatrix_apply, LinearMap.toMatrix_apply, coe_lmul_eq_mul, LinearMap.mul_apply', + Basis.repr_smul, Basis.smul_apply, LinearEquiv.trans_apply, + DistribMulAction.toLinearEquiv_symm_apply, mul_smul_comm, inv_smul_smul] + end Lmul section LmulTower diff --git a/Mathlib/LinearAlgebra/PiTensorProduct.lean b/Mathlib/LinearAlgebra/PiTensorProduct.lean index 8b99e81563b37..1175a41ee25f1 100644 --- a/Mathlib/LinearAlgebra/PiTensorProduct.lean +++ b/Mathlib/LinearAlgebra/PiTensorProduct.lean @@ -370,7 +370,7 @@ theorem ext {φ₁ φ₂ : (⨂[R] i, s i) →ₗ[R] E} PiTensorProduct.induction_on' z ?_ fun {x y} hx hy ↦ by rw [φ₁.map_add, φ₂.map_add, hx, hy] · intro r f rw [tprodCoeff_eq_smul_tprod, φ₁.map_smul, φ₂.map_smul] - apply _root_.congr_arg + apply congr_arg exact MultilinearMap.congr_fun H f /-- The pure tensors (i.e. the elements of the image of `PiTensorProduct.tprod`) span diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 61742a6060f19..01447cf465604 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -223,7 +223,7 @@ theorem map_add_add_add_map (x y z : M) : abel theorem map_add_self (x : M) : Q (x + x) = 4 • Q x := by - rw [← two_smul R x, map_smul, nsmul_eq_smul_cast R] + rw [← two_smul R x, map_smul, ← Nat.cast_smul_eq_nsmul R] norm_num -- not @[simp] because it is superseded by `ZeroHomClass.map_zero` @@ -766,7 +766,7 @@ theorem _root_.QuadraticMap.polarBilin_injective (h : IsUnit (2 : R)) : intro Q₁ Q₂ h₁₂ apply h.smul_left_cancel.mp rw [show (2 : R) = (2 : ℕ) by rfl] - simp_rw [← nsmul_eq_smul_cast R, ← QuadraticMap.toQuadraticMap_polarBilin] + simp_rw [Nat.cast_smul_eq_nsmul R, ← QuadraticMap.toQuadraticMap_polarBilin] exact congrArg toQuadraticMap h₁₂ section diff --git a/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean b/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean index a83ef40bf2565..ff5c68946f2d8 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/IsometryEquiv.lean @@ -28,7 +28,8 @@ open QuadraticMap namespace QuadraticMap variable [CommSemiring R] -variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] [AddCommMonoid N] +variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] + [AddCommMonoid N] variable [Module R M] [Module R M₁] [Module R M₂] [Module R M₃] [Module R N] /-- An isometric equivalence between two quadratic spaces `M₁, Q₁` and `M₂, Q₂` over a ring `R`, diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean index d7cca8c036177..8bb7be468d23b 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean @@ -120,8 +120,8 @@ theorem gradedComm_of_tmul_of (i j : ι) (a : 𝒜 i) (b : ℬ j) : dsimp only [LinearEquiv.trans_apply, LinearEquiv.ofLinear_apply] rw [TensorProduct.directSum_lof_tmul_lof, gradedCommAux_lof_tmul, Units.smul_def, -- Note: #8386 specialized `map_smul` to `LinearEquiv.map_smul` to avoid timeouts. - zsmul_eq_smul_cast R, LinearEquiv.map_smul, TensorProduct.directSum_symm_lof_tmul, - ← zsmul_eq_smul_cast, ← Units.smul_def] + ← Int.cast_smul_eq_nsmul R, LinearEquiv.map_smul, TensorProduct.directSum_symm_lof_tmul, + Int.cast_smul_eq_nsmul, ← Units.smul_def] theorem gradedComm_tmul_of_zero (a : ⨁ i, 𝒜 i) (b : ℬ 0) : gradedComm R 𝒜 ℬ (a ⊗ₜ lof R _ ℬ 0 b) = lof R _ ℬ _ b ⊗ₜ a := by @@ -190,7 +190,7 @@ theorem tmul_of_gradedMul_of_tmul (j₁ i₂ : ι) LinearMap.lTensor_tmul] rw [mul_comm j₁ i₂, gradedComm_of_tmul_of] -- the tower smul lemmas elaborate too slowly - rw [Units.smul_def, Units.smul_def, zsmul_eq_smul_cast R, zsmul_eq_smul_cast R] + rw [Units.smul_def, Units.smul_def, ← Int.cast_smul_eq_nsmul R, ← Int.cast_smul_eq_nsmul R] -- Note: #8386 had to specialize `map_smul` to avoid timeouts. rw [← smul_tmul', LinearEquiv.map_smul, tmul_smul, LinearEquiv.map_smul, LinearMap.map_smul] dsimp @@ -237,11 +237,11 @@ theorem gradedMul_assoc (x y z : DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ) : exact DFunLike.congr_fun (DFunLike.congr_fun (DFunLike.congr_fun this x) y) z ext ixa xa ixb xb iya ya iyb yb iza za izb zb dsimp [mA] - simp_rw [tmul_of_gradedMul_of_tmul, Units.smul_def, zsmul_eq_smul_cast R, + simp_rw [tmul_of_gradedMul_of_tmul, Units.smul_def, ← Int.cast_smul_eq_nsmul R, LinearMap.map_smul₂, LinearMap.map_smul, DirectSum.lof_eq_of, DirectSum.of_mul_of, ← DirectSum.lof_eq_of R, tmul_of_gradedMul_of_tmul, DirectSum.lof_eq_of, ← DirectSum.of_mul_of, ← DirectSum.lof_eq_of R, mul_assoc] - simp_rw [← zsmul_eq_smul_cast R, ← Units.smul_def, smul_smul, ← uzpow_add, add_mul, mul_add] + simp_rw [Int.cast_smul_eq_nsmul R, ← Units.smul_def, smul_smul, ← uzpow_add, add_mul, mul_add] congr 2 abel @@ -256,9 +256,9 @@ theorem gradedComm_gradedMul (x y : DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ) : dsimp rw [gradedComm_of_tmul_of, gradedComm_of_tmul_of, tmul_of_gradedMul_of_tmul] -- Note: #8386 had to specialize `map_smul` to avoid timeouts. - simp_rw [Units.smul_def, zsmul_eq_smul_cast R, LinearEquiv.map_smul, LinearMap.map_smul, + simp_rw [Units.smul_def, ← Int.cast_smul_eq_nsmul R, LinearEquiv.map_smul, LinearMap.map_smul, LinearMap.smul_apply] - simp_rw [← zsmul_eq_smul_cast R, ← Units.smul_def, DirectSum.lof_eq_of, DirectSum.of_mul_of, + simp_rw [Int.cast_smul_eq_nsmul R, ← Units.smul_def, DirectSum.lof_eq_of, DirectSum.of_mul_of, ← DirectSum.lof_eq_of R, gradedComm_of_tmul_of, tmul_of_gradedMul_of_tmul, smul_smul, DirectSum.lof_eq_of, ← DirectSum.of_mul_of, ← DirectSum.lof_eq_of R] simp_rw [← uzpow_add, mul_add, add_mul, mul_comm i₁ j₂] diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index c58248b158b1b..0f33c9bae0d79 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -189,8 +189,8 @@ theorem tmul_coe_mul_coe_tmul {j₁ i₂ : ι} (a₁ : A) (b₁ : ℬ j₁) (a simp_rw [lof_eq_of R] rw [LinearEquiv.symm_symm] -- Note: #8386 had to specialize `map_smul` to `LinearEquiv.map_smul` - rw [@Units.smul_def _ _ (_) (_), zsmul_eq_smul_cast R, LinearEquiv.map_smul, map_smul, - ← zsmul_eq_smul_cast R, ← @Units.smul_def _ _ (_) (_)] + rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_nsmul R, LinearEquiv.map_smul, map_smul, + Int.cast_smul_eq_nsmul R, ← @Units.smul_def _ _ (_) (_)] rw [congr_symm_tmul] dsimp simp_rw [decompose_symm_mul, decompose_symm_of, Equiv.symm_apply_apply] @@ -315,8 +315,8 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) ext a₂ b₂ : 2 dsimp rw [tmul_coe_mul_coe_tmul] - rw [@Units.smul_def _ _ (_) (_), zsmul_eq_smul_cast R, map_smul, map_smul, map_smul] - rw [← zsmul_eq_smul_cast R, ← @Units.smul_def _ _ (_) (_)] + rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_nsmul R, map_smul, map_smul, map_smul] + rw [Int.cast_smul_eq_nsmul R, ← @Units.smul_def _ _ (_) (_)] rw [of_symm_of, map_tmul, LinearMap.mul'_apply] simp_rw [AlgHom.toLinearMap_apply, _root_.map_mul] simp_rw [mul_assoc (f a₁), ← mul_assoc _ _ (g b₂), h_anti_commutes, mul_smul_comm, @@ -376,7 +376,7 @@ lemma auxEquiv_comm (x : 𝒜 ᵍ⊗[R] ℬ) : comm 𝒜 ℬ (a ᵍ⊗ₜ b) = (-1 : ℤˣ)^(j * i) • (b ᵍ⊗ₜ a : ℬ ᵍ⊗[R] 𝒜) := (auxEquiv R ℬ 𝒜).injective <| by simp_rw [auxEquiv_comm, auxEquiv_tmul, decompose_coe, ← lof_eq_of R, gradedComm_of_tmul_of, - @Units.smul_def _ _ (_) (_), zsmul_eq_smul_cast R] + @Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_nsmul R] -- Qualified `map_smul` to avoid a TC timeout #8386 erw [LinearMap.map_smul, auxEquiv_tmul] simp_rw [decompose_coe, lof_eq_of] diff --git a/Mathlib/Logic/Nontrivial/Basic.lean b/Mathlib/Logic/Nontrivial/Basic.lean index d8d8d64bcd8b8..5c64525d7d4d1 100644 --- a/Mathlib/Logic/Nontrivial/Basic.lean +++ b/Mathlib/Logic/Nontrivial/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Init.Order.Defs import Mathlib.Logic.Nontrivial.Defs import Mathlib.Tactic.Attr.Register import Mathlib.Data.Prod.Basic -import Mathlib.Data.Subtype import Mathlib.Logic.Function.Basic import Mathlib.Logic.Unique diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 2575f5e1efd46..96d125bfa3e2d 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -238,6 +238,9 @@ variable [TopologicalSpace α] [MeasurableSpace α] [OpensMeasurableSpace α] [T theorem IsOpen.measurableSet (h : IsOpen s) : MeasurableSet s := OpensMeasurableSpace.borel_le _ <| GenerateMeasurable.basic _ h +theorem IsOpen.nullMeasurableSet {μ} (h : IsOpen s) : NullMeasurableSet s μ := + h.measurableSet.nullMeasurableSet + instance (priority := 1000) {s : Set α} [h : HasCountableSeparatingOn α IsOpen s] : CountablySeparated s := by rw [CountablySeparated.subtype_iff] @@ -258,6 +261,9 @@ theorem measurableSet_of_continuousAt {β} [EMetricSpace β] (f : α → β) : theorem IsClosed.measurableSet (h : IsClosed s) : MeasurableSet s := h.isOpen_compl.measurableSet.of_compl +theorem IsClosed.nullMeasurableSet {μ} (h : IsClosed s) : NullMeasurableSet s μ := + h.measurableSet.nullMeasurableSet + theorem IsCompact.measurableSet [T2Space α] (h : IsCompact s) : MeasurableSet s := h.isClosed.measurableSet @@ -265,7 +271,7 @@ theorem IsCompact.measurableSet [T2Space α] (h : IsCompact s) : MeasurableSet s then they can't be separated by a Borel measurable set. -/ theorem Inseparable.mem_measurableSet_iff {x y : γ} (h : Inseparable x y) {s : Set γ} (hs : MeasurableSet s) : x ∈ s ↔ y ∈ s := - hs.induction_on_open (C := fun s ↦ (x ∈ s ↔ y ∈ s)) (fun _ ↦ h.mem_open_iff) (fun s _ hs ↦ hs.not) + hs.induction_on_open (C := fun s ↦ (x ∈ s ↔ y ∈ s)) (fun _ ↦ h.mem_open_iff) (fun s _ ↦ Iff.not) fun _ _ _ h ↦ by simp [h] /-- If `K` is a compact set in an R₁ space and `s ⊇ K` is a Borel measurable superset, diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 46c6490048acd..613feccea1950 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -406,9 +406,7 @@ theorem pi_of_empty {α : Type*} [Fintype α] [IsEmpty α] {β : α → Type*} haveI : ∀ a, SigmaFinite (μ a) := isEmptyElim refine pi_eq fun s _ => ?_ rw [Fintype.prod_empty, dirac_apply_of_mem] - #adaptation_note - /-- 2024-07-18 need to disable elab_as_elim with `(p := _)` due to type-incorrect motive -/ - exact isEmptyElim (p := _) (α := α) + exact isEmptyElim (α := α) lemma volume_pi_eq_dirac {ι : Type*} [Fintype ι] [IsEmpty ι] {α : ι → Type*} [∀ i, MeasureSpace (α i)] (x : ∀ a, α a := isEmptyElim) : diff --git a/Mathlib/MeasureTheory/Constructions/Polish.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean similarity index 94% rename from Mathlib/MeasureTheory/Constructions/Polish.lean rename to Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index a38d7de96d27a..75e4a6fbdd0a2 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -3,10 +3,9 @@ Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Felix Weilacher -/ -import Mathlib.Data.Real.Cardinality -import Mathlib.Topology.MetricSpace.Perfect import Mathlib.MeasureTheory.Constructions.BorelSpace.Metric import Mathlib.Topology.CountableSeparatingOn +import Mathlib.Topology.MetricSpace.Perfect /-! # The Borel sigma-algebra on Polish spaces @@ -1005,58 +1004,3 @@ noncomputable def Equiv.measurableEquiv (e : α ≃ β) : α ≃ᵐ β := by rwa [e.countable_iff] at h end PolishSpace - -namespace MeasureTheory - -variable (α) -variable [MeasurableSpace α] [StandardBorelSpace α] - -theorem exists_nat_measurableEquiv_range_coe_fin_of_finite [Finite α] : - ∃ n : ℕ, Nonempty (α ≃ᵐ range ((↑) : Fin n → ℝ)) := by - obtain ⟨n, ⟨n_equiv⟩⟩ := Finite.exists_equiv_fin α - refine ⟨n, ⟨PolishSpace.Equiv.measurableEquiv (n_equiv.trans ?_)⟩⟩ - exact Equiv.ofInjective _ (Nat.cast_injective.comp Fin.val_injective) - -theorem measurableEquiv_range_coe_nat_of_infinite_of_countable [Infinite α] [Countable α] : - Nonempty (α ≃ᵐ range ((↑) : ℕ → ℝ)) := by - have : PolishSpace (range ((↑) : ℕ → ℝ)) := - Nat.closedEmbedding_coe_real.isClosedMap.isClosed_range.polishSpace - refine ⟨PolishSpace.Equiv.measurableEquiv ?_⟩ - refine (nonempty_equiv_of_countable.some : α ≃ ℕ).trans ?_ - exact Equiv.ofInjective ((↑) : ℕ → ℝ) Nat.cast_injective - -/-- Any standard Borel space is measurably equivalent to a subset of the reals. -/ -theorem exists_subset_real_measurableEquiv : ∃ s : Set ℝ, MeasurableSet s ∧ Nonempty (α ≃ᵐ s) := by - by_cases hα : Countable α - · cases finite_or_infinite α - · obtain ⟨n, h_nonempty_equiv⟩ := exists_nat_measurableEquiv_range_coe_fin_of_finite α - refine ⟨_, ?_, h_nonempty_equiv⟩ - letI : MeasurableSpace (Fin n) := borel (Fin n) - haveI : BorelSpace (Fin n) := ⟨rfl⟩ - apply MeasurableEmbedding.measurableSet_range (mα := by infer_instance) - exact continuous_of_discreteTopology.measurableEmbedding - (Nat.cast_injective.comp Fin.val_injective) - · refine ⟨_, ?_, measurableEquiv_range_coe_nat_of_infinite_of_countable α⟩ - apply MeasurableEmbedding.measurableSet_range (mα := by infer_instance) - exact continuous_of_discreteTopology.measurableEmbedding Nat.cast_injective - · refine - ⟨univ, MeasurableSet.univ, - ⟨(PolishSpace.measurableEquivOfNotCountable hα ?_ : α ≃ᵐ (univ : Set ℝ))⟩⟩ - rw [countable_coe_iff] - exact Cardinal.not_countable_real - -/-- Any standard Borel space embeds measurably into the reals. -/ -theorem exists_measurableEmbedding_real : ∃ f : α → ℝ, MeasurableEmbedding f := by - obtain ⟨s, hs, ⟨e⟩⟩ := exists_subset_real_measurableEquiv α - exact ⟨(↑) ∘ e, (MeasurableEmbedding.subtype_coe hs).comp e.measurableEmbedding⟩ - -/-- A measurable embedding of a standard Borel space into `ℝ`. -/ -noncomputable -def embeddingReal (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := - (exists_measurableEmbedding_real Ω).choose - -lemma measurableEmbedding_embeddingReal (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : - MeasurableEmbedding (embeddingReal Ω) := - (exists_measurableEmbedding_real Ω).choose_spec - -end MeasureTheory diff --git a/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean b/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean new file mode 100644 index 0000000000000..f6436e70180fc --- /dev/null +++ b/Mathlib/MeasureTheory/Constructions/Polish/EmbeddingReal.lean @@ -0,0 +1,66 @@ +/- +Copyright (c) 2023 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Data.Real.Cardinality +import Mathlib.MeasureTheory.Constructions.Polish.Basic + +/-! +# A Polish Borel space is measurably equivalent to a set of reals +-/ + +open Set Function PolishSpace PiNat TopologicalSpace Bornology Metric Filter Topology MeasureTheory + +namespace MeasureTheory +variable (α : Type*) [MeasurableSpace α] [StandardBorelSpace α] + +theorem exists_nat_measurableEquiv_range_coe_fin_of_finite [Finite α] : + ∃ n : ℕ, Nonempty (α ≃ᵐ range ((↑) : Fin n → ℝ)) := by + obtain ⟨n, ⟨n_equiv⟩⟩ := Finite.exists_equiv_fin α + refine ⟨n, ⟨PolishSpace.Equiv.measurableEquiv (n_equiv.trans ?_)⟩⟩ + exact Equiv.ofInjective _ (Nat.cast_injective.comp Fin.val_injective) + +theorem measurableEquiv_range_coe_nat_of_infinite_of_countable [Infinite α] [Countable α] : + Nonempty (α ≃ᵐ range ((↑) : ℕ → ℝ)) := by + have : PolishSpace (range ((↑) : ℕ → ℝ)) := + Nat.closedEmbedding_coe_real.isClosedMap.isClosed_range.polishSpace + refine ⟨PolishSpace.Equiv.measurableEquiv ?_⟩ + refine (nonempty_equiv_of_countable.some : α ≃ ℕ).trans ?_ + exact Equiv.ofInjective ((↑) : ℕ → ℝ) Nat.cast_injective + +/-- Any standard Borel space is measurably equivalent to a subset of the reals. -/ +theorem exists_subset_real_measurableEquiv : ∃ s : Set ℝ, MeasurableSet s ∧ Nonempty (α ≃ᵐ s) := by + by_cases hα : Countable α + · cases finite_or_infinite α + · obtain ⟨n, h_nonempty_equiv⟩ := exists_nat_measurableEquiv_range_coe_fin_of_finite α + refine ⟨_, ?_, h_nonempty_equiv⟩ + letI : MeasurableSpace (Fin n) := borel (Fin n) + haveI : BorelSpace (Fin n) := ⟨rfl⟩ + apply MeasurableEmbedding.measurableSet_range (mα := by infer_instance) + exact continuous_of_discreteTopology.measurableEmbedding + (Nat.cast_injective.comp Fin.val_injective) + · refine ⟨_, ?_, measurableEquiv_range_coe_nat_of_infinite_of_countable α⟩ + apply MeasurableEmbedding.measurableSet_range (mα := by infer_instance) + exact continuous_of_discreteTopology.measurableEmbedding Nat.cast_injective + · refine + ⟨univ, MeasurableSet.univ, + ⟨(PolishSpace.measurableEquivOfNotCountable hα ?_ : α ≃ᵐ (univ : Set ℝ))⟩⟩ + rw [countable_coe_iff] + exact Cardinal.not_countable_real + +/-- Any standard Borel space embeds measurably into the reals. -/ +theorem exists_measurableEmbedding_real : ∃ f : α → ℝ, MeasurableEmbedding f := by + obtain ⟨s, hs, ⟨e⟩⟩ := exists_subset_real_measurableEquiv α + exact ⟨(↑) ∘ e, (MeasurableEmbedding.subtype_coe hs).comp e.measurableEmbedding⟩ + +/-- A measurable embedding of a standard Borel space into `ℝ`. -/ +noncomputable +def embeddingReal (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : Ω → ℝ := + (exists_measurableEmbedding_real Ω).choose + +lemma measurableEmbedding_embeddingReal (Ω : Type*) [MeasurableSpace Ω] [StandardBorelSpace Ω] : + MeasurableEmbedding (embeddingReal Ω) := + (exists_measurableEmbedding_real Ω).choose_spec + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 05bb7381dca5f..bfe4f8c693f3b 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -151,13 +151,13 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S exact (hcb (H _ H')).elim /-- Vitali covering theorem, closed balls version: given a family `t` of closed balls, one can -extract a disjoint subfamily `u ⊆ t` so that all balls in `t` are covered by the 5-times -dilations of balls in `u`. -/ +extract a disjoint subfamily `u ⊆ t` so that all balls in `t` are covered by the τ-times +dilations of balls in `u`, for some `τ > 3`. -/ theorem exists_disjoint_subfamily_covering_enlargment_closedBall [MetricSpace α] (t : Set ι) - (x : ι → α) (r : ι → ℝ) (R : ℝ) (hr : ∀ a ∈ t, r a ≤ R) : + (x : ι → α) (r : ι → ℝ) (R : ℝ) (hr : ∀ a ∈ t, r a ≤ R) (τ : ℝ) (hτ : 3 < τ) : ∃ u ⊆ t, (u.PairwiseDisjoint fun a => closedBall (x a) (r a)) ∧ - ∀ a ∈ t, ∃ b ∈ u, closedBall (x a) (r a) ⊆ closedBall (x b) (5 * r b) := by + ∀ a ∈ t, ∃ b ∈ u, closedBall (x a) (r a) ⊆ closedBall (x b) (τ * r b) := by rcases eq_empty_or_nonempty t with (rfl | _) · exact ⟨∅, Subset.refl _, pairwiseDisjoint_empty, by simp⟩ by_cases ht : ∀ a ∈ t, r a < 0 @@ -171,11 +171,11 @@ theorem exists_disjoint_subfamily_covering_enlargment_closedBall [MetricSpace α fun a ha => ⟨a, ha, by simp only [closedBall_eq_empty.2 (ht a ha), empty_subset]⟩⟩ push_neg at ht let t' := { a ∈ t | 0 ≤ r a } - rcases exists_disjoint_subfamily_covering_enlargment (fun a => closedBall (x a) (r a)) t' r 2 - one_lt_two (fun a ha => ha.2) R (fun a ha => hr a ha.1) fun a ha => + rcases exists_disjoint_subfamily_covering_enlargment (fun a => closedBall (x a) (r a)) t' r + ((τ - 1) / 2) (by linarith) (fun a ha => ha.2) R (fun a ha => hr a ha.1) fun a ha => ⟨x a, mem_closedBall_self ha.2⟩ with ⟨u, ut', u_disj, hu⟩ - have A : ∀ a ∈ t', ∃ b ∈ u, closedBall (x a) (r a) ⊆ closedBall (x b) (5 * r b) := by + have A : ∀ a ∈ t', ∃ b ∈ u, closedBall (x a) (r a) ⊆ closedBall (x b) (τ * r b) := by intro a ha rcases hu a ha with ⟨b, bu, hb, rb⟩ refine ⟨b, bu, ?_⟩ diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 921b661485c88..13feeda01ee5a 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -189,7 +189,7 @@ private theorem measure_of_restrictNonposSeq (hi₂ : ¬s ≤[i] 0) (n : ℕ) | succ n => rw [restrictNonposSeq_succ] have h₁ : ¬s ≤[i \ ⋃ (k : ℕ) (_ : k ≤ n), restrictNonposSeq s i k] 0 := by - refine mt (restrict_le_zero_subset _ ?_ (by simp [Nat.lt_succ_iff]; rfl)) hn + refine mt (restrict_le_zero_subset _ ?_ (by simp [Nat.lt_succ_iff])) hn convert measurable_of_not_restrict_le_zero _ hn using 3 exact funext fun x => by rw [Nat.lt_succ_iff] rcases someExistsOneDivLT_spec h₁ with ⟨_, _, h⟩ diff --git a/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean index b36c4094fb6f5..486b726308834 100644 --- a/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean @@ -43,26 +43,21 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : have hν : ∀ s, ν s ≠ ∞ := measure_ne_top ν have to_nnreal_μ : ∀ s, ((μ s).toNNReal : ℝ≥0∞) = μ s := fun s => ENNReal.coe_toNNReal <| hμ _ have to_nnreal_ν : ∀ s, ((ν s).toNNReal : ℝ≥0∞) = ν s := fun s => ENNReal.coe_toNNReal <| hν _ - have d_split : ∀ s t, MeasurableSet s → MeasurableSet t → d s = d (s \ t) + d (s ∩ t) := by - intro s t _hs ht + have d_split s t (ht : MeasurableSet t) : d s = d (s \ t) + d (s ∩ t) := by dsimp only [d] rw [← measure_inter_add_diff s ht, ← measure_inter_add_diff s ht, ENNReal.toNNReal_add (hμ _) (hμ _), ENNReal.toNNReal_add (hν _) (hν _), NNReal.coe_add, NNReal.coe_add] simp only [sub_eq_add_neg, neg_add] abel - have d_Union : - ∀ s : ℕ → Set α, Monotone s → Tendsto (fun n => d (s n)) atTop (𝓝 (d (⋃ n, s n))) := by - intro s hm + have d_Union (s : ℕ → Set α) (hm : Monotone s) : + Tendsto (fun n => d (s n)) atTop (𝓝 (d (⋃ n, s n))) := by refine Tendsto.sub ?_ ?_ <;> refine NNReal.tendsto_coe.2 <| (ENNReal.tendsto_toNNReal ?_).comp <| tendsto_measure_iUnion hm · exact hμ _ · exact hν _ - have d_Inter : - ∀ s : ℕ → Set α, - (∀ n, MeasurableSet (s n)) → - (∀ n m, n ≤ m → s m ⊆ s n) → Tendsto (fun n => d (s n)) atTop (𝓝 (d (⋂ n, s n))) := by - intro s hs hm + have d_Inter (s : ℕ → Set α) (hs : ∀ n, MeasurableSet (s n)) (hm : ∀ n m, n ≤ m → s m ⊆ s n) : + Tendsto (fun n => d (s n)) atTop (𝓝 (d (⋂ n, s n))) := by refine Tendsto.sub ?_ ?_ <;> refine NNReal.tendsto_coe.2 <| @@ -76,8 +71,7 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : exact measure_mono (subset_univ _) have c_nonempty : c.Nonempty := Nonempty.image _ ⟨_, MeasurableSet.empty⟩ have d_le_γ : ∀ s, MeasurableSet s → d s ≤ γ := fun s hs => le_csSup bdd_c ⟨s, hs, rfl⟩ - have : ∀ n : ℕ, ∃ s : Set α, MeasurableSet s ∧ γ - (1 / 2) ^ n < d s := by - intro n + have (n : ℕ) : ∃ s : Set α, MeasurableSet s ∧ γ - (1 / 2) ^ n < d s := by have : γ - (1 / 2) ^ n < γ := sub_lt_self γ (pow_pos (half_pos zero_lt_one) n) rcases exists_lt_of_lt_csSup c_nonempty this with ⟨r, ⟨s, hs, rfl⟩, hlt⟩ exact ⟨s, hs, hlt⟩ @@ -86,21 +80,17 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : have he₁ : ∀ n, MeasurableSet (e n) := fun n => (he n).1 have he₂ : ∀ n, γ - (1 / 2) ^ n < d (e n) := fun n => (he n).2 let f : ℕ → ℕ → Set α := fun n m => (Finset.Ico n (m + 1)).inf e - have hf : ∀ n m, MeasurableSet (f n m) := by - intro n m + have hf n m : MeasurableSet (f n m) := by simp only [f, Finset.inf_eq_iInf] exact MeasurableSet.biInter (to_countable _) fun i _ => he₁ _ - have f_subset_f : ∀ {a b c d}, a ≤ b → c ≤ d → f a d ⊆ f b c := by - intro a b c d hab hcd + have f_subset_f {a b c d} (hab : a ≤ b) (hcd : c ≤ d) : f a d ⊆ f b c := by simp_rw [f, Finset.inf_eq_iInf] exact biInter_subset_biInter_left (Finset.Ico_subset_Ico hab <| Nat.succ_le_succ hcd) - have f_succ : ∀ n m, n ≤ m → f n (m + 1) = f n m ∩ e (m + 1) := by - intro n m hnm + have f_succ n m (hnm : n ≤ m) : f n (m + 1) = f n m ∩ e (m + 1) := by have : n ≤ m + 1 := le_of_lt (Nat.succ_le_succ hnm) simp_rw [f, Nat.Ico_succ_right_eq_insert_Ico this, Finset.inf_insert, Set.inter_comm] rfl - have le_d_f : ∀ n m, m ≤ n → γ - 2 * (1 / 2) ^ m + (1 / 2) ^ n ≤ d (f m n) := by - intro n m h + have le_d_f n m (h : m ≤ n) : γ - 2 * (1 / 2) ^ m + (1 / 2) ^ n ≤ d (f m n) := by refine Nat.le_induction ?_ ?_ n h · have := he₂ m simp_rw [f, Nat.Ico_succ_singleton, Finset.inf_singleton] @@ -115,11 +105,10 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : simp only [sub_eq_add_neg]; abel _ ≤ d (e (n + 1)) + d (f m n) := add_le_add (le_of_lt <| he₂ _) ih _ ≤ d (e (n + 1)) + d (f m n \ e (n + 1)) + d (f m (n + 1)) := by - rw [f_succ _ _ hmn, d_split (f m n) (e (n + 1)) (hf _ _) (he₁ _), add_assoc] + rw [f_succ _ _ hmn, d_split (f m n) (e (n + 1)) (he₁ _), add_assoc] _ = d (e (n + 1) ∪ f m n) + d (f m (n + 1)) := by rw [d_split (e (n + 1) ∪ f m n) (e (n + 1)), union_diff_left, union_inter_cancel_left] · abel - · exact (he₁ _).union (hf _ _) · exact he₁ _ _ ≤ γ + d (f m (n + 1)) := add_le_add_right (d_le_γ _ <| (he₁ _).union (hf _ _)) _ exact (add_le_add_iff_left γ).1 this @@ -155,7 +144,7 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : (add_le_add_iff_left γ).1 <| calc γ + 0 ≤ d s := by rw [add_zero]; exact γ_le_d_s - _ = d (s \ t) + d t := by rw [d_split _ _ hs ht, inter_eq_self_of_subset_right hts] + _ = d (s \ t) + d t := by rw [d_split s _ ht, inter_eq_self_of_subset_right hts] _ ≤ γ + d t := add_le_add (d_le_γ _ (hs.diff ht)) le_rfl rw [← to_nnreal_μ, ← to_nnreal_ν, ENNReal.coe_le_coe, ← NNReal.coe_le_coe] @@ -166,7 +155,7 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : calc γ + d t ≤ d s + d t := by gcongr _ = d (s ∪ t) := by - rw [d_split _ _ (hs.union ht) ht, union_diff_right, union_inter_cancel_right, + rw [d_split (s ∪ t) _ ht, union_diff_right, union_inter_cancel_right, (subset_compl_iff_disjoint_left.1 hts).sdiff_eq_left] _ ≤ γ + 0 := by rw [add_zero]; exact d_le_γ _ (hs.union ht) diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index 67657135a6c1f..6c90b652e93b4 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Analysis.InnerProductSpace.Basic -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lp import Mathlib.MeasureTheory.Integral.SetIntegral import Mathlib.Order.Filter.Ring diff --git a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean index 26d66a1c3b200..c54cbfa054dea 100644 --- a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean +++ b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean @@ -61,7 +61,8 @@ open scoped ENNReal NNReal Topology BoundedContinuousFunction open MeasureTheory TopologicalSpace ContinuousMap Set Bornology -variable {α : Type*} [MeasurableSpace α] [TopologicalSpace α] [T4Space α] [BorelSpace α] +variable {α : Type*} [TopologicalSpace α] [NormalSpace α] [R1Space α] + [MeasurableSpace α] [BorelSpace α] variable {E : Type*} [NormedAddCommGroup E] {μ : Measure α} {p : ℝ≥0∞} namespace MeasureTheory @@ -159,23 +160,23 @@ theorem Memℒp.exists_hasCompactSupport_snorm_sub_le [WeaklyLocallyCompactSpace ∃ η : ℝ≥0, 0 < η ∧ ∀ s : Set α, μ s ≤ η → snorm (s.indicator fun _x => c) p μ ≤ δ := exists_snorm_indicator_le hp c δpos.ne' have hη_pos' : (0 : ℝ≥0∞) < η := ENNReal.coe_pos.2 ηpos - obtain ⟨s, st, s_compact, μs⟩ : ∃ s, s ⊆ t ∧ IsCompact s ∧ μ (t \ s) < η := - ht.exists_isCompact_diff_lt htμ.ne hη_pos'.ne' + obtain ⟨s, st, s_compact, s_closed, μs⟩ : + ∃ s, s ⊆ t ∧ IsCompact s ∧ IsClosed s ∧ μ (t \ s) < η := + ht.exists_isCompact_isClosed_diff_lt htμ.ne hη_pos'.ne' have hsμ : μ s < ∞ := (measure_mono st).trans_lt htμ have I1 : snorm ((s.indicator fun _y => c) - t.indicator fun _y => c) p μ ≤ δ := by rw [← snorm_neg, neg_sub, ← indicator_diff st] exact hη _ μs.le obtain ⟨k, k_compact, sk⟩ : ∃ k : Set α, IsCompact k ∧ s ⊆ interior k := exists_compact_superset s_compact - rcases exists_continuous_snorm_sub_le_of_closed hp s_compact.isClosed isOpen_interior sk hsμ.ne c - δpos.ne' with - ⟨f, f_cont, I2, _f_bound, f_support, f_mem⟩ + rcases exists_continuous_snorm_sub_le_of_closed hp s_closed isOpen_interior sk hsμ.ne c δpos.ne' + with ⟨f, f_cont, I2, _f_bound, f_support, f_mem⟩ have I3 : snorm (f - t.indicator fun _y => c) p μ ≤ ε := by convert (hδ _ _ (f_mem.aestronglyMeasurable.sub - (aestronglyMeasurable_const.indicator s_compact.measurableSet)) - ((aestronglyMeasurable_const.indicator s_compact.measurableSet).sub + (aestronglyMeasurable_const.indicator s_closed.measurableSet)) + ((aestronglyMeasurable_const.indicator s_closed.measurableSet).sub (aestronglyMeasurable_const.indicator ht)) I2 I1).le using 2 simp only [sub_add_sub_cancel] diff --git a/Mathlib/MeasureTheory/Function/Intersectivity.lean b/Mathlib/MeasureTheory/Function/Intersectivity.lean index 75a21c61ec714..b7f04c60b2445 100644 --- a/Mathlib/MeasureTheory/Function/Intersectivity.lean +++ b/Mathlib/MeasureTheory/Function/Intersectivity.lean @@ -94,8 +94,8 @@ lemma bergelson' {s : ℕ → Set α} (hs : ∀ n, MeasurableSet (s n)) (hr₀ : _ ≤ limsup (f · x) atTop := hx -- This exactly means that the `s n` containing `x` have all their finite intersection non-null. · refine ⟨{n | x ∈ s n}, fun hxs ↦ ?_, fun u hux hu ↦ ?_⟩ - -- This next block proves that a set of strictly positive natural density is infinite, mixed with - -- the fact that `{n | x ∈ s n}` has strictly positive natural density. + -- This next block proves that a set of strictly positive natural density is infinite, mixed + -- with the fact that `{n | x ∈ s n}` has strictly positive natural density. -- TODO: Separate it out to a lemma once we have a natural density API. · refine ENNReal.div_ne_zero.2 ⟨hr₀, measure_ne_top _ _⟩ $ eq_bot_mono hx $ Tendsto.limsup_eq $ tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index a7ccdc1d1c3a5..927ded827f58c 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -8,7 +8,7 @@ import Mathlib.MeasureTheory.Constructions.BorelSpace.ContinuousLinearMap import Mathlib.MeasureTheory.Covering.BesicovitchVectorSpace import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar import Mathlib.Analysis.NormedSpace.Pointwise -import Mathlib.MeasureTheory.Constructions.Polish +import Mathlib.MeasureTheory.Constructions.Polish.Basic import Mathlib.Analysis.Calculus.InverseFunctionTheorem.ApproximatesLinearOn /-! diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean index 74b3bc3b9cce2..254a43c65d92d 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean @@ -91,8 +91,7 @@ theorem snorm_add_le' {f g : α → E} (hf : AEStronglyMeasurable f μ) (hg : AE · have : p ∈ Set.Ioo (0 : ℝ≥0∞) 1 := ⟨hp.bot_lt, h'p⟩ simp only [LpAddConst, if_pos this] · simpa using ENNReal.toReal_mono ENNReal.one_ne_top h'p.le - · simp [LpAddConst_of_one_le h'p] - exact snorm_add_le hf hg h'p + · simpa [LpAddConst_of_one_le h'p] using snorm_add_le hf hg h'p variable (μ E) diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 1b10a30913cdc..57a5448ad474d 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -885,7 +885,7 @@ theorem Memℒp.norm_rpow_div {f : α → E} (hf : Memℒp f p μ) (q : ℝ≥0 by_cases q_top : q = ∞ · simp [q_top] by_cases q_zero : q = 0 - · simp [q_zero] + · simp only [q_zero, ENNReal.zero_toReal, Real.rpow_zero] by_cases p_zero : p = 0 · simp [p_zero] rw [ENNReal.div_zero p_zero] diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 54b90fa9139d6..61b8c0277bf38 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -81,7 +81,7 @@ theorem norm_approxOn_zero_le [OpensMeasurableSpace E] {f : β → E} (hf : Meas (h₀ : (0 : E) ∈ s) [SeparableSpace s] (x : β) (n : ℕ) : ‖approxOn f hf s 0 h₀ n x‖ ≤ ‖f x‖ + ‖f x‖ := by have := edist_approxOn_y0_le hf h₀ x n - simp [edist_comm (0 : E), edist_eq_coe_nnnorm] at this + simp only [edist_comm (0 : E), edist_eq_coe_nnnorm] at this exact mod_cast this theorem tendsto_approxOn_Lp_snorm [OpensMeasurableSpace E] {f : β → E} (hf : Measurable f) @@ -832,9 +832,7 @@ theorem Lp.induction [_i : Fact (1 ≤ p)] (hp_ne_top : p ≠ ∞) (P : Lp E p P (hf.toLp f) → P (hg.toLp g) → P (hf.toLp f + hg.toLp g)) (h_closed : IsClosed { f : Lp E p μ | P f }) : ∀ f : Lp E p μ, P f := by refine fun f => (Lp.simpleFunc.denseRange hp_ne_top).induction_on f h_closed ?_ - #adaptation_note - /-- 2024-07-18 need to disable elab_as_elim with `(P := _)` due to type-incorrect motive -/ - refine Lp.simpleFunc.induction (P := _) (α := α) (E := E) (lt_of_lt_of_le zero_lt_one _i.elim).ne' + refine Lp.simpleFunc.induction (α := α) (E := E) (lt_of_lt_of_le zero_lt_one _i.elim).ne' hp_ne_top ?_ ?_ · exact fun c s => h_ind c · exact fun f g hf hg => h_add hf hg diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 1a720c99630b2..fd82c96b14643 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -3,7 +3,7 @@ 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, Sébastien Gouëzel -/ -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic import Mathlib.MeasureTheory.Function.SimpleFuncDense /-! diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean index 0408d14ff6214..85ecc189b07e7 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Sébastien Gouëzel -/ -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps import Mathlib.MeasureTheory.Function.StronglyMeasurable.Basic import Mathlib.MeasureTheory.Measure.WithDensity import Mathlib.Topology.Algebra.Module.FiniteDimension diff --git a/Mathlib/MeasureTheory/Function/UnifTight.lean b/Mathlib/MeasureTheory/Function/UnifTight.lean index f2db25bc86a2a..76cbd5ba0db9d 100644 --- a/Mathlib/MeasureTheory/Function/UnifTight.lean +++ b/Mathlib/MeasureTheory/Function/UnifTight.lean @@ -65,7 +65,7 @@ theorem unifTight_iff_ennreal {_ : MeasurableSpace α} (f : ι → α → β) (p μ s ≠ ∞ ∧ ∀ i, snorm (sᶜ.indicator (f i)) p μ ≤ ε := by simp only [ENNReal.forall_ennreal, ENNReal.coe_pos] refine (and_iff_left ?_).symm - simp [-ne_eq, zero_lt_top, le_top] + simp only [zero_lt_top, le_top, implies_true, and_true, true_implies] use ∅; simpa only [measure_empty] using zero_ne_top theorem unifTight_iff_real {_ : MeasurableSpace α} (f : ι → α → β) (p : ℝ≥0∞) (μ : Measure α) : diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 03638ac32d760..c45cc18acbfdf 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -98,6 +98,16 @@ instance isMulRightInvariant_smul_nnreal [IsMulRightInvariant μ] (c : ℝ≥0) IsMulRightInvariant (c • μ) := MeasureTheory.isMulRightInvariant_smul (c : ℝ≥0∞) +@[to_additive] +instance IsMulLeftInvariant.smulInvariantMeasure [IsMulLeftInvariant μ] : + SMulInvariantMeasure G G μ := + ⟨fun _x _s hs => measure_preimage_of_map_eq_self (map_mul_left_eq_self _ _) hs.nullMeasurableSet⟩ + +@[to_additive] +instance IsMulRightInvariant.toSMulInvariantMeasure_op [μ.IsMulRightInvariant] : + SMulInvariantMeasure Gᵐᵒᵖ G μ := + ⟨fun _x _s hs => measure_preimage_of_map_eq_self (map_mul_right_eq_self _ _) hs.nullMeasurableSet⟩ + section MeasurableMul variable [MeasurableMul G] @@ -124,16 +134,6 @@ theorem MeasurePreserving.mul_right (μ : Measure G) [IsMulRightInvariant μ] (g MeasurePreserving (fun x => f x * g) μ' μ := (measurePreserving_mul_right μ g).comp hf -@[to_additive] -instance IsMulLeftInvariant.smulInvariantMeasure [IsMulLeftInvariant μ] : - SMulInvariantMeasure G G μ := - ⟨fun x _s hs => (measurePreserving_mul_left μ x).measure_preimage hs⟩ - -@[to_additive] -instance IsMulRightInvariant.toSMulInvariantMeasure_op [μ.IsMulRightInvariant] : - SMulInvariantMeasure Gᵐᵒᵖ G μ := - ⟨fun x _s hs => (measurePreserving_mul_right μ (MulOpposite.unop x)).measure_preimage hs⟩ - @[to_additive] instance Subgroup.smulInvariantMeasure {G α : Type*} [Group G] [MulAction G α] [MeasurableSpace α] {μ : Measure α} [SMulInvariantMeasure G α μ] (H : Subgroup G) : SMulInvariantMeasure H α μ := diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 1e5ac1c15a20f..1a765f962c2b4 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -8,7 +8,7 @@ import Mathlib.Analysis.Calculus.Deriv.Comp import Mathlib.Analysis.Calculus.Deriv.Add import Mathlib.Analysis.Calculus.Deriv.Slope import Mathlib.Analysis.Calculus.Deriv.Mul -import Mathlib.Analysis.NormedSpace.Dual +import Mathlib.Analysis.Normed.Module.Dual import Mathlib.MeasureTheory.Integral.DominatedConvergence import Mathlib.MeasureTheory.Integral.VitaliCaratheodory diff --git a/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean b/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean index 18a6699169add..f1c89ef6de1a1 100644 --- a/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean +++ b/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.MeasureTheory.Measure.WithDensity -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # A lemma about measurability with density under scalar multiplication in normed spaces -/ diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 1d1341ff2747b..0d23d017587cf 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -90,10 +90,8 @@ theorem _root_.Measurable.lmarginal (hf : Measurable f) : Measurable (∫⋯∫ refine hf.comp ?_ rw [measurable_pi_iff]; intro i by_cases hi : i ∈ s - · simp [hi, updateFinset] - exact measurable_pi_iff.1 measurable_snd _ - · simp [hi, updateFinset] - exact measurable_pi_iff.1 measurable_fst _ + · simpa [hi, updateFinset] using measurable_pi_iff.1 measurable_snd _ + · simpa [hi, updateFinset] using measurable_pi_iff.1 measurable_fst _ @[simp] theorem lmarginal_empty (f : (∀ i, π i) → ℝ≥0∞) : ∫⋯∫⁻_∅, f ∂μ = f := by ext1 x diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index ae03484dfd177..988b4fd6310cf 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -195,7 +195,7 @@ theorem lintegral_prod_norm_pow_le {α ι : Type*} [MeasurableSpace α] {μ : Me simp at hp | @insert i₀ s hi₀ ih => rcases eq_or_ne (p i₀) 1 with h2i₀|h2i₀ - · simp [hi₀] + · simp only [hi₀, not_false_eq_true, prod_insert] have h2p : ∀ i ∈ s, p i = 0 := by simpa [hi₀, h2i₀, sum_eq_zero_iff_of_nonneg (fun i hi ↦ h2p i <| mem_insert_of_mem hi)] using hp diff --git a/Mathlib/MeasureTheory/Integral/Periodic.lean b/Mathlib/MeasureTheory/Integral/Periodic.lean index c539e6683a401..5b378bc95c27c 100644 --- a/Mathlib/MeasureTheory/Integral/Periodic.lean +++ b/Mathlib/MeasureTheory/Integral/Periodic.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov, Alex Kontorovich, Heather Macbeth -/ import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar import Mathlib.MeasureTheory.Measure.Haar.Quotient -import Mathlib.MeasureTheory.Constructions.Polish import Mathlib.MeasureTheory.Integral.IntervalIntegral import Mathlib.Topology.Algebra.Order.Floor diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index eee4fd60f51b1..95f26a1e03a47 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -1287,6 +1287,7 @@ theorem coe_top : ↑(⊤ : Subtype (MeasurableSet : Set α → Prop)) = (⊤ : noncomputable instance Subtype.instBooleanAlgebra : BooleanAlgebra (Subtype (MeasurableSet : Set α → Prop)) := Subtype.coe_injective.booleanAlgebra _ coe_union coe_inter coe_top coe_bot coe_compl coe_sdiff + coe_himp @[measurability] theorem measurableSet_blimsup {s : ℕ → Set α} {p : ℕ → Prop} (h : ∀ n, p n → MeasurableSet (s n)) : diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index e0216810f736e..4710b819f0406 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -83,10 +83,6 @@ variable {G : Type*} [Group G] namespace haar --- Porting note: Even in `noncomputable section`, a definition with `to_additive` require --- `noncomputable` to generate an additive definition. --- Please refer to leanprover/lean4#2077. - /-- The index or Haar covering number or ratio of `K` w.r.t. `V`, denoted `(K : V)`: it is the smallest number of (left) translates of `V` that is necessary to cover `K`. It is defined to be 0 if no finite number of translates cover `K`. -/ @@ -341,11 +337,6 @@ theorem nonempty_iInter_clPrehaar (K₀ : PositiveCompacts G) : ### Lemmas about `chaar` -/ - --- Porting note: Even in `noncomputable section`, a definition with `to_additive` require --- `noncomputable` to generate an additive definition. --- Please refer to leanprover/lean4#2077. - /-- This is the "limit" of `prehaar K₀ U K` as `U` becomes a smaller and smaller open neighborhood of `(1 : G)`. More precisely, it is defined to be an arbitrary element in the intersection of all the sets `clPrehaar K₀ V` in `haarProduct K₀`. @@ -463,10 +454,6 @@ theorem is_left_invariant_chaar {K₀ : PositiveCompacts G} (g : G) (K : Compact apply is_left_invariant_prehaar; rw [h2U.interior_eq]; exact ⟨1, h3U⟩ · apply continuous_iff_isClosed.mp this; exact isClosed_singleton --- Porting note: Even in `noncomputable section`, a definition with `to_additive` require --- `noncomputable` to generate an additive definition. --- Please refer to leanprover/lean4#2077. - /-- The function `chaar` interpreted in `ℝ≥0`, as a content -/ @[to_additive "additive version of `MeasureTheory.Measure.haar.haarContent`"] noncomputable def haarContent (K₀ : PositiveCompacts G) : Content G where @@ -518,13 +505,8 @@ open haar ### The Haar measure -/ - variable [TopologicalSpace G] [TopologicalGroup G] [MeasurableSpace G] [BorelSpace G] --- Porting note: Even in `noncomputable section`, a definition with `to_additive` require --- `noncomputable` to generate an additive definition. --- Please refer to leanprover/lean4#2077. - /-- The Haar measure on the locally compact group `G`, scaled so that `haarMeasure K₀ K₀ = 1`. -/ @[to_additive "The Haar measure on the locally compact additive group `G`, scaled so that diff --git a/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean b/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean index 2dd76343e5ba3..df47302132fae 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.MeasureTheory.Measure.Haar.Basic -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.MeasureTheory.Measure.Haar.Unique /-! diff --git a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean index 1cc12766e5dd1..e6af597345ae0 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean @@ -114,10 +114,8 @@ theorem PiLp.volume_preserving_equiv_symm : MeasurePreserving (WithLp.equiv 2 ( lemma volume_euclideanSpace_eq_dirac [IsEmpty ι] : (volume : Measure (EuclideanSpace ℝ ι)) = Measure.dirac 0 := by - ext s hs - simp only [← ((EuclideanSpace.volume_preserving_measurableEquiv ι).symm).measure_preimage hs, - volume_pi_eq_dirac 0, MeasurableEquiv.measurableSet_preimage, hs, dirac_apply', indicator, - mem_preimage, Pi.one_apply] + rw [← ((EuclideanSpace.volume_preserving_measurableEquiv ι).symm).map_eq, + volume_pi_eq_dirac 0, map_dirac (MeasurableEquiv.measurable _)] rfl end PiLp diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index c53c0efd30991..50342e831be4d 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -3,11 +3,11 @@ Copyright (c) 2022 Alex Kontorovich and Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, Heather Macbeth -/ +import Mathlib.Algebra.Group.Opposite +import Mathlib.MeasureTheory.Constructions.Polish.Basic +import Mathlib.MeasureTheory.Group.FundamentalDomain import Mathlib.MeasureTheory.Integral.DominatedConvergence import Mathlib.MeasureTheory.Measure.Haar.Basic -import Mathlib.MeasureTheory.Group.FundamentalDomain -import Mathlib.Algebra.Group.Opposite -import Mathlib.MeasureTheory.Constructions.Polish /-! # Haar quotient measure diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index e6d1d6661fb10..c0bc10b459319 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -365,13 +365,13 @@ theorem measure_union_toMeasurable : μ (s ∪ toMeasurable μ t) = μ (s ∪ t) (measure_toMeasurable _).le theorem sum_measure_le_measure_univ {s : Finset ι} {t : ι → Set α} - (h : ∀ i ∈ s, MeasurableSet (t i)) (H : Set.PairwiseDisjoint (↑s) t) : + (h : ∀ i ∈ s, NullMeasurableSet (t i) μ) (H : Set.Pairwise s (AEDisjoint μ on t)) : (∑ i ∈ s, μ (t i)) ≤ μ (univ : Set α) := by - rw [← measure_biUnion_finset H h] + rw [← measure_biUnion_finset₀ H h] exact measure_mono (subset_univ _) -theorem tsum_measure_le_measure_univ {s : ι → Set α} (hs : ∀ i, MeasurableSet (s i)) - (H : Pairwise (Disjoint on s)) : (∑' i, μ (s i)) ≤ μ (univ : Set α) := by +theorem tsum_measure_le_measure_univ {s : ι → Set α} (hs : ∀ i, NullMeasurableSet (s i) μ) + (H : Pairwise (AEDisjoint μ on s)) : ∑' i, μ (s i) ≤ μ (univ : Set α) := by rw [ENNReal.tsum_eq_iSup_sum] exact iSup_le fun s => sum_measure_le_measure_univ (fun i _hi => hs i) fun i _hi j _hj hij => H hij @@ -379,23 +379,23 @@ theorem tsum_measure_le_measure_univ {s : ι → Set α} (hs : ∀ i, Measurable /-- Pigeonhole principle for measure spaces: if `∑' i, μ (s i) > μ univ`, then one of the intersections `s i ∩ s j` is not empty. -/ theorem exists_nonempty_inter_of_measure_univ_lt_tsum_measure {m : MeasurableSpace α} - (μ : Measure α) {s : ι → Set α} (hs : ∀ i, MeasurableSet (s i)) + (μ : Measure α) {s : ι → Set α} (hs : ∀ i, NullMeasurableSet (s i) μ) (H : μ (univ : Set α) < ∑' i, μ (s i)) : ∃ i j, i ≠ j ∧ (s i ∩ s j).Nonempty := by contrapose! H apply tsum_measure_le_measure_univ hs intro i j hij - exact disjoint_iff_inter_eq_empty.mpr (H i j hij) + exact (disjoint_iff_inter_eq_empty.mpr (H i j hij)).aedisjoint /-- Pigeonhole principle for measure spaces: if `s` is a `Finset` and `∑ i ∈ s, μ (t i) > μ univ`, then one of the intersections `t i ∩ t j` is not empty. -/ theorem exists_nonempty_inter_of_measure_univ_lt_sum_measure {m : MeasurableSpace α} (μ : Measure α) - {s : Finset ι} {t : ι → Set α} (h : ∀ i ∈ s, MeasurableSet (t i)) + {s : Finset ι} {t : ι → Set α} (h : ∀ i ∈ s, NullMeasurableSet (t i) μ) (H : μ (univ : Set α) < ∑ i ∈ s, μ (t i)) : ∃ i ∈ s, ∃ j ∈ s, ∃ _h : i ≠ j, (t i ∩ t j).Nonempty := by contrapose! H apply sum_measure_le_measure_univ h intro i hi j hj hij - exact disjoint_iff_inter_eq_empty.mpr (H i hi j hj hij) + exact (disjoint_iff_inter_eq_empty.mpr (H i hi j hj hij)).aedisjoint /-- If two sets `s` and `t` are included in a set `u`, and `μ s + μ t > μ u`, then `s` intersects `t`. Version assuming that `t` is measurable. -/ @@ -1172,6 +1172,18 @@ theorem map_toOuterMeasure (hf : AEMeasurable f μ) : @[simp] lemma mapₗ_eq_zero_iff (hf : Measurable f) : Measure.mapₗ f μ = 0 ↔ μ = 0 := by rw [mapₗ_apply_of_measurable hf, map_eq_zero_iff hf.aemeasurable] +/-- If `map f μ = μ`, then the measure of the preimage of any null measurable set `s` +is equal to the measure of `s`. +Note that this lemma does not assume (a.e.) measurability of `f`. -/ +lemma measure_preimage_of_map_eq_self {f : α → α} (hf : map f μ = μ) + {s : Set α} (hs : NullMeasurableSet s μ) : μ (f ⁻¹' s) = μ s := by + if hfm : AEMeasurable f μ then + rw [← map_apply₀ hfm, hf] + rwa [hf] + else + rw [map_of_not_aemeasurable hfm] at hf + simp [← hf] + lemma map_ne_zero_iff (hf : AEMeasurable f μ) : μ.map f ≠ 0 ↔ μ ≠ 0 := (map_eq_zero_iff hf).not lemma mapₗ_ne_zero_iff (hf : Measurable f) : Measure.mapₗ f μ ≠ 0 ↔ μ ≠ 0 := (mapₗ_eq_zero_iff hf).not @@ -1216,8 +1228,10 @@ theorem tendsto_ae_map {f : α → β} (hf : AEMeasurable f μ) : Tendsto f (ae /-- Pullback of a `Measure` as a linear map. If `f` sends each measurable set to a measurable set, then for each measurable set `s` we have `comapₗ f μ s = μ (f '' s)`. +Note that if `f` is not injective, this definition assigns `Set.univ` measure zero. + If the linearity is not needed, please use `comap` instead, which works for a larger class of -functions. -/ +functions. `comapₗ` is an auxiliary definition and most lemmas deal with comap. -/ def comapₗ [MeasurableSpace α] (f : α → β) : Measure β →ₗ[ℝ≥0∞] Measure α := if hf : Injective f ∧ ∀ s, MeasurableSet s → MeasurableSet (f '' s) then liftLinear (OuterMeasure.comap f) fun μ s hs t => by @@ -1233,7 +1247,9 @@ theorem comapₗ_apply {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : exact ⟨hfi, hf⟩ /-- Pullback of a `Measure`. If `f` sends each measurable set to a null-measurable set, -then for each measurable set `s` we have `comap f μ s = μ (f '' s)`. -/ +then for each measurable set `s` we have `comap f μ s = μ (f '' s)`. + +Note that if `f` is not injective, this definition assigns `Set.univ` measure zero. -/ def comap [MeasurableSpace α] (f : α → β) (μ : Measure β) : Measure α := if hf : Injective f ∧ ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ then (OuterMeasure.comap f μ.toOuterMeasure).toMeasure fun s hs t => by diff --git a/Mathlib/MeasureTheory/Order/UpperLower.lean b/Mathlib/MeasureTheory/Order/UpperLower.lean index 9209c645a3906..4c8b3268546a2 100644 --- a/Mathlib/MeasureTheory/Order/UpperLower.lean +++ b/Mathlib/MeasureTheory/Order/UpperLower.lean @@ -75,8 +75,8 @@ private lemma aux₀ gcongr; exact subset_inter ((hf₁ _ $ hε' n).trans interior_subset_closure) $ hf₀ _ $ hε' n dsimp have := hε' n - rw [Real.volume_pi_closedBall, Real.volume_pi_closedBall, ← ENNReal.ofReal_div_of_pos, ← div_pow, - mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div] + rw [Real.volume_pi_closedBall, Real.volume_pi_closedBall, ← ENNReal.ofReal_div_of_pos, + ← div_pow, mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div] all_goals positivity /-- If we can fit a small ball inside a set `sᶜ` intersected with any neighborhood of `x`, then the diff --git a/Mathlib/MeasureTheory/OuterMeasure/Induced.lean b/Mathlib/MeasureTheory/OuterMeasure/Induced.lean index ac9d2193337a6..a6ee00bdb14c8 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Induced.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Induced.lean @@ -349,7 +349,7 @@ theorem trim_zero : (0 : OuterMeasure α).trim = 0 := theorem trim_sum_ge {ι} (m : ι → OuterMeasure α) : (sum fun i => (m i).trim) ≤ (sum m).trim := fun s => by - simp [trim_eq_iInf] + simp only [sum_apply, trim_eq_iInf, le_iInf_iff] exact fun t st ht => ENNReal.tsum_le_tsum fun i => iInf_le_of_le t <| iInf_le_of_le st <| iInf_le _ ht diff --git a/Mathlib/MeasureTheory/OuterMeasure/Operations.lean b/Mathlib/MeasureTheory/OuterMeasure/Operations.lean index 144855acbd04f..0267845e31d67 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Operations.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Operations.lean @@ -73,7 +73,7 @@ variable {R' : Type*} [SMul R' ℝ≥0∞] [IsScalarTower R' ℝ≥0∞ ℝ≥0 instance instSMul : SMul R (OuterMeasure α) := ⟨fun c m => { measureOf := fun s => c • m s - empty := by simp; rw [← smul_one_mul c]; simp + empty := by simp only [measure_empty]; rw [← smul_one_mul c]; simp mono := fun {s t} h => by simp only rw [← smul_one_mul c, ← smul_one_mul c (m t)] diff --git a/Mathlib/MeasureTheory/SetAlgebra.lean b/Mathlib/MeasureTheory/SetAlgebra.lean index 08a235e7298b6..4c9f2857117fd 100644 --- a/Mathlib/MeasureTheory/SetAlgebra.lean +++ b/Mathlib/MeasureTheory/SetAlgebra.lean @@ -8,7 +8,7 @@ import Mathlib.MeasureTheory.SetSemiring /-! # Algebra of sets -in this file we define the notion of algebra of sets and give its basic properties. An algebra +In this file we define the notion of algebra of sets and give its basic properties. An algebra of sets is a family of sets containing the empty set and closed by complement and binary union. It is therefore similar to a `σ`-algebra, except that it is not necessarily closed by countable unions. @@ -218,8 +218,8 @@ theorem countable_generateSetAlgebra (h : 𝒜.Countable) : have count_ℬ : ℬ.Countable := by apply h.union have : compl '' 𝒜 = {s | sᶜ ∈ 𝒜} := by - ext s; simp - exact ⟨fun ⟨x, x_mem, hx⟩ ↦ by simp [← hx, x_mem], fun hs ↦ ⟨sᶜ, hs, by simp⟩⟩ + ext s + simpa using ⟨fun ⟨x, x_mem, hx⟩ ↦ by simp [← hx, x_mem], fun hs ↦ ⟨sᶜ, hs, by simp⟩⟩ exact this ▸ h.image compl let f : Set (Set (Set α)) → Set α := fun A ↦ ⋃ a ∈ A, ⋂ t ∈ a, t let 𝒞 := {a | a.Finite ∧ a ⊆ ℬ} diff --git a/Mathlib/MeasureTheory/Tactic.lean b/Mathlib/MeasureTheory/Tactic.lean deleted file mode 100644 index ba1406d0a122b..0000000000000 --- a/Mathlib/MeasureTheory/Tactic.lean +++ /dev/null @@ -1,19 +0,0 @@ -/- -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.Measure.MeasureSpaceDef -import Mathlib.Tactic.Measurability - -/-! -# Tactics for measure theory - -Currently we have one domain-specific tactic for topology: `measurability`. -It is implemented in `Mathlib.Tactic.Measurability`. - -Porting note: the sole purpose of this file is to mark it as "ported". -This file seems to be tripping up the porting dashboard. - --/ diff --git a/Mathlib/ModelTheory/Definability.lean b/Mathlib/ModelTheory/Definability.lean index 69265fd7c24e1..0413ebeacdec9 100644 --- a/Mathlib/ModelTheory/Definability.lean +++ b/Mathlib/ModelTheory/Definability.lean @@ -377,9 +377,9 @@ theorem coe_sdiff (s t : L.DefinableSet A α) : @[simp, norm_cast] lemma coe_himp (s t : L.DefinableSet A α) : ↑(s ⇨ t) = (s ⇨ t : Set (α → M)) := rfl -instance instBooleanAlgebra : BooleanAlgebra (L.DefinableSet A α) := +noncomputable instance instBooleanAlgebra : BooleanAlgebra (L.DefinableSet A α) := Function.Injective.booleanAlgebra (α := L.DefinableSet A α) _ Subtype.coe_injective - coe_sup coe_inf coe_top coe_bot coe_compl coe_sdiff + coe_sup coe_inf coe_top coe_bot coe_compl coe_sdiff coe_himp end DefinableSet diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 4d165729ffbb1..9476da7554d39 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -52,46 +52,36 @@ def listEncode : L.Term α → List (α ⊕ (Σi, L.Functions i)) Sum.inr (⟨_, f⟩ : Σi, L.Functions i)::(List.finRange _).bind fun i => (ts i).listEncode /-- Decodes a list of variables and function symbols as a list of terms. -/ -def listDecode : List (α ⊕ (Σi, L.Functions i)) → List (Option (L.Term α)) +def listDecode : List (α ⊕ (Σi, L.Functions i)) → List (L.Term α) | [] => [] - | Sum.inl a::l => some (var a)::listDecode l + | Sum.inl a::l => (var a)::listDecode l | Sum.inr ⟨n, f⟩::l => - if h : ∀ i : Fin n, ((listDecode l).get? i).join.isSome then - (func f fun i => Option.get _ (h i))::(listDecode l).drop n - else [none] + if h : n ≤ (listDecode l).length then + (func f (fun i => (listDecode l)[i])) :: (listDecode l).drop n + else [] theorem listDecode_encode_list (l : List (L.Term α)) : - listDecode (l.bind listEncode) = l.map Option.some := by + listDecode (l.bind listEncode) = l := by suffices h : ∀ (t : L.Term α) (l : List (α ⊕ (Σi, L.Functions i))), - listDecode (t.listEncode ++ l) = some t::listDecode l by + listDecode (t.listEncode ++ l) = t::listDecode l by induction' l with t l lih · rfl - · rw [bind_cons, h t (l.bind listEncode), lih, List.map] + · rw [bind_cons, h t (l.bind listEncode), lih] intro t induction' t with a n f ts ih <;> intro l · rw [listEncode, singleton_append, listDecode] · rw [listEncode, cons_append, listDecode] have h : listDecode (((finRange n).bind fun i : Fin n => (ts i).listEncode) ++ l) = - (finRange n).map (Option.some ∘ ts) ++ listDecode l := by + (finRange n).map ts ++ listDecode l := by induction' finRange n with i l' l'ih · rfl - · rw [bind_cons, List.append_assoc, ih, map_cons, l'ih, cons_append, Function.comp] - have h' : ∀ i : Fin n, - (listDecode (((finRange n).bind fun i : Fin n => (ts i).listEncode) ++ l))[(i : Nat)]? = - some (some (ts i)) := by - intro i - rw [h, getElem?_append, getElem?_map] - · simp only [Option.map_eq_some', Function.comp_apply, getElem?_eq_some] - refine ⟨i, ⟨lt_of_lt_of_le i.2 (ge_of_eq (length_finRange _)), ?_⟩, rfl⟩ - rw [getElem_finRange, Fin.eta] - · refine lt_of_lt_of_le i.2 ?_ - simp - refine (dif_pos fun i => Option.isSome_iff_exists.2 ⟨ts i, ?_⟩).trans ?_ - · rw [get?_eq_getElem?, Option.join_eq_some, h'] - refine congr (congr rfl (congr rfl (congr rfl (funext fun i => Option.get_of_mem _ ?_)))) ?_ - · simp [h'] - · rw [h, drop_left'] - rw [length_map, length_finRange] + · rw [bind_cons, List.append_assoc, ih, map_cons, l'ih, cons_append] + simp only [h, length_append, length_map, length_finRange, le_add_iff_nonneg_right, + _root_.zero_le, ↓reduceDIte, getElem_fin, cons.injEq, func.injEq, heq_eq_eq, true_and] + refine ⟨funext (fun i => ?_), ?_⟩ + · rw [List.getElem_append, List.getElem_map, List.getElem_finRange] + simp only [length_map, length_finRange, i.2] + · simp only [length_map, length_finRange, drop_left'] /-- An encoding of terms as lists. -/ @[simps] @@ -102,7 +92,8 @@ protected def encoding : Encoding (L.Term α) where decode_encode t := by have h := listDecode_encode_list [t] rw [bind_singleton] at h - simp only [h, Option.join, head?, List.map, Option.some_bind, id] + simp only [Option.join, h, head?_cons, Option.pure_def, Option.bind_eq_bind, Option.some_bind, + id_eq] theorem listEncode_injective : Function.Injective (listEncode : L.Term α → List (α ⊕ (Σi, L.Functions i))) := @@ -146,7 +137,8 @@ instance [Encodable α] [Encodable (Σi, L.Functions i)] : Encodable (L.Term α) Encodable.ofLeftInjection listEncode (fun l => (listDecode l).head?.join) fun t => by simp only rw [← bind_singleton listEncode, listDecode_encode_list] - simp only [Option.join, head?, List.map, Option.some_bind, id] + simp only [Option.join, head?_cons, Option.pure_def, Option.bind_eq_bind, Option.some_bind, + id_eq] instance [h1 : Countable α] [h2 : Countable (Σl, L.Functions l)] : Countable (L.Term α) := by refine mk_le_aleph0_iff.1 (card_le.trans (max_le_iff.2 ?_)) @@ -176,61 +168,63 @@ def sigmaAll : (Σn, L.BoundedFormula α n) → Σn, L.BoundedFormula α n | ⟨n + 1, φ⟩ => ⟨n, φ.all⟩ | _ => default + +@[simp] +lemma sigmaAll_apply {n} {φ : L.BoundedFormula α (n + 1)} : + sigmaAll ⟨n + 1, φ⟩ = ⟨n, φ.all⟩ := rfl + /-- Applies `imp` to two elements of `(Σ n, L.BoundedFormula α n)`, or returns `default` if not possible. -/ def sigmaImp : (Σn, L.BoundedFormula α n) → (Σn, L.BoundedFormula α n) → Σn, L.BoundedFormula α n | ⟨m, φ⟩, ⟨n, ψ⟩ => if h : m = n then ⟨m, φ.imp (Eq.mp (by rw [h]) ψ)⟩ else default +#adaptation_note +/-- +`List.drop_sizeOf_le` is deprecated. +See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Why.20is.20.60Mathlib.2EModelTheory.2EEncoding.60.20using.20.60SizeOf.2EsizeOf.60.3F +for discussion about adapting this code. +-/ +set_option linter.deprecated false in /-- Decodes a list of symbols as a list of formulas. -/ @[simp] -def listDecode : ∀ l : List ((Σk, L.Term (α ⊕ Fin k)) ⊕ ((Σn, L.Relations n) ⊕ ℕ)), - (Σn, L.BoundedFormula α n) × - { l' : List ((Σk, L.Term (α ⊕ Fin k)) ⊕ ((Σn, L.Relations n) ⊕ ℕ)) // - SizeOf.sizeOf l' ≤ max 1 (SizeOf.sizeOf l) } - | Sum.inr (Sum.inr (n + 2))::l => ⟨⟨n, falsum⟩, l, le_max_of_le_right le_add_self⟩ +lemma sigmaImp_apply {n} {φ ψ : L.BoundedFormula α n} : + sigmaImp ⟨n, φ⟩ ⟨n, ψ⟩ = ⟨n, φ.imp ψ⟩ := by + simp only [sigmaImp, ↓reduceDIte, eq_mp_eq_cast, cast_eq] + +/-- Decodes a list of symbols as a list of formulas. -/ +def listDecode : + List ((Σk, L.Term (α ⊕ Fin k)) ⊕ ((Σn, L.Relations n) ⊕ ℕ)) → List (Σn, L.BoundedFormula α n) + | Sum.inr (Sum.inr (n + 2))::l => ⟨n, falsum⟩::(listDecode l) | Sum.inl ⟨n₁, t₁⟩::Sum.inl ⟨n₂, t₂⟩::l => - ⟨if h : n₁ = n₂ then ⟨n₁, equal t₁ (Eq.mp (by rw [h]) t₂)⟩ else default, l, by - simp only [SizeOf.sizeOf, List._sizeOf_1, ← add_assoc] - exact le_max_of_le_right le_add_self⟩ - | Sum.inr (Sum.inl ⟨n, R⟩)::Sum.inr (Sum.inr k)::l => - ⟨if h : ∀ i : Fin n, ((l.map Sum.getLeft?).get? i).join.isSome then + (if h : n₁ = n₂ then ⟨n₁, equal t₁ (Eq.mp (by rw [h]) t₂)⟩ else default)::(listDecode l) + | Sum.inr (Sum.inl ⟨n, R⟩)::Sum.inr (Sum.inr k)::l => ( + if h : ∀ i : Fin n, ((l.map Sum.getLeft?).get? i).join.isSome then if h' : ∀ i, (Option.get _ (h i)).1 = k then ⟨k, BoundedFormula.rel R fun i => Eq.mp (by rw [h' i]) (Option.get _ (h i)).2⟩ else default - else default, - l.drop n, le_max_of_le_right (le_add_left (le_add_left (List.drop_sizeOf_le _ _)))⟩ - | Sum.inr (Sum.inr 0)::l => - have : SizeOf.sizeOf - (↑(listDecode l).2 : List ((Σk, L.Term (α ⊕ Fin k)) ⊕ ((Σn, L.Relations n) ⊕ ℕ))) < - 1 + (1 + 1) + SizeOf.sizeOf l := by - refine lt_of_le_of_lt (listDecode l).2.2 (max_lt ?_ (Nat.lt_add_of_pos_left (by decide))) - rw [add_assoc, lt_add_iff_pos_right, add_pos_iff] - exact Or.inl zero_lt_two - ⟨sigmaImp (listDecode l).1 (listDecode (listDecode l).2).1, - (listDecode (listDecode l).2).2, - le_max_of_le_right - (_root_.trans (listDecode _).2.2 - (max_le (le_add_right le_self_add) - (_root_.trans (listDecode _).2.2 (max_le (le_add_right le_self_add) le_add_self))))⟩ - | Sum.inr (Sum.inr 1)::l => - ⟨sigmaAll (listDecode l).1, (listDecode l).2, - (listDecode l).2.2.trans (max_le_max le_rfl le_add_self)⟩ - | _ => ⟨default, [], le_max_left _ _⟩ + else default)::(listDecode (l.drop n)) + | Sum.inr (Sum.inr 0)::l => if h : 2 ≤ (listDecode l).length + then (sigmaImp (listDecode l)[0] (listDecode l)[1])::(drop 2 (listDecode l)) + else [] + | Sum.inr (Sum.inr 1)::l => if h : 1 ≤ (listDecode l).length + then (sigmaAll (listDecode l)[0])::(drop 1 (listDecode l)) + else [] + | _ => [] + termination_by l => l.length -set_option linter.deprecated false in @[simp] theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : - (listDecode (l.bind fun φ => φ.2.listEncode)).1 = l.headI := by - suffices h : ∀ (φ : Σn, L.BoundedFormula α n) (l), - (listDecode (listEncode φ.2 ++ l)).1 = φ ∧ (listDecode (listEncode φ.2 ++ l)).2.1 = l by - induction' l with φ l _ - · rw [List.nil_bind] + listDecode (l.bind (fun φ => φ.2.listEncode)) = l := by + suffices h : ∀ (φ : Σn, L.BoundedFormula α n) + (l' : List ((Σk, L.Term (α ⊕ Fin k)) ⊕ ((Σn, L.Relations n) ⊕ ℕ))), + (listDecode (listEncode φ.2 ++ l')) = φ::(listDecode l') by + induction' l with φ l ih + · rw [List.bind_nil] simp [listDecode] - · rw [cons_bind, (h φ _).1, headI_cons] + · rw [bind_cons, h φ _, ih] rintro ⟨n, φ⟩ induction' φ with _ _ _ _ φ_n φ_l φ_R ts _ _ _ ih1 ih2 _ _ ih <;> intro l · rw [listEncode, singleton_append, listDecode] - simp only [eq_self_iff_true, heq_iff_eq, and_self_iff] · rw [listEncode, cons_append, cons_append, listDecode, dif_pos] · simp only [eq_mp_eq_cast, cast_eq, eq_self_iff_true, heq_iff_eq, and_self_iff, nil_append] · simp only [eq_self_iff_true, heq_iff_eq, and_self_iff] @@ -242,9 +236,8 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : simp only [Option.join, map_append, map_map, Option.bind_eq_some, id, exists_eq_right, get?_eq_some, length_append, length_map, length_finRange] refine ⟨lt_of_lt_of_le i.2 le_self_add, ?_⟩ - rw [get_append, get_map] - · simp only [Sum.getLeft?, get_finRange, Fin.eta, Function.comp_apply, eq_self_iff_true, - heq_iff_eq, and_self_iff] + rw [get_eq_getElem, getElem_append, getElem_map] + · simp only [getElem_finRange, Fin.eta, Function.comp_apply, Sum.getLeft?] · simp only [length_map, length_finRange, is_lt] rw [dif_pos] swap @@ -254,31 +247,30 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : · intro i obtain ⟨h1, h2⟩ := Option.eq_some_iff_get_eq.1 (h i) rw [h2] - simp only [Sigma.mk.inj_iff, heq_eq_eq, rel.injEq, true_and] + simp only [Option.join, eq_mp_eq_cast, cons.injEq, Sigma.mk.inj_iff, heq_eq_eq, rel.injEq, + true_and] refine ⟨funext fun i => ?_, ?_⟩ · obtain ⟨h1, h2⟩ := Option.eq_some_iff_get_eq.1 (h i) - rw [eq_mp_eq_cast, cast_eq_iff_heq] + rw [cast_eq_iff_heq] exact (Sigma.ext_iff.1 ((Sigma.eta (Option.get _ h1)).trans h2)).2 rw [List.drop_append_eq_append_drop, length_map, length_finRange, Nat.sub_self, drop, drop_eq_nil_of_le, nil_append] rw [length_map, length_finRange] - · rw [listEncode, List.append_assoc, cons_append, listDecode] - simp only [] at * - rw [(ih1 _).1, (ih1 _).2, (ih2 _).1, (ih2 _).2, sigmaImp] - simp only [dite_true] - exact ⟨rfl, trivial⟩ - · rw [listEncode, cons_append, listDecode] - simp only - simp only [] at * - rw [(ih _).1, (ih _).2, sigmaAll] - exact ⟨rfl, rfl⟩ + · simp only [] at * + rw [listEncode, List.append_assoc, cons_append, listDecode] + simp only [ih1, ih2, length_cons, le_add_iff_nonneg_left, _root_.zero_le, ↓reduceDIte, + getElem_cons_zero, getElem_cons_succ, sigmaImp_apply, drop_succ_cons, drop_zero] + · simp only [] at * + rw [listEncode, cons_append, listDecode] + simp only [ih, length_cons, le_add_iff_nonneg_left, _root_.zero_le, ↓reduceDIte, + getElem_cons_zero, sigmaAll_apply, drop_succ_cons, drop_zero] /-- An encoding of bounded formulas as lists. -/ @[simps] protected def encoding : Encoding (Σn, L.BoundedFormula α n) where Γ := (Σk, L.Term (α ⊕ Fin k)) ⊕ ((Σn, L.Relations n) ⊕ ℕ) encode φ := φ.2.listEncode - decode l := (listDecode l).1 + decode l := (listDecode l)[0]? decode_encode φ := by have h := listDecode_encode_list [φ] rw [bind_singleton] at h diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index 5f6ddb9406385..7bdaca5ee4308 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -234,7 +234,7 @@ theorem sumsq_eq_zero (x) : ∀ l, sumsq l x = 0 ↔ l.Forall fun a : Poly α => have t := add_le_add_left (sumsq_nonneg x ps) (p x * p x) rwa [add_zero] at t) (mul_self_nonneg _) - ⟨this, by simp [this] at h; exact h⟩, + ⟨this, by simpa [this] using h⟩, fun ⟨h1, h2⟩ => by rw [add_apply, mul_apply, h1, h2]; rfl⟩ end @@ -242,9 +242,9 @@ end /-- Map the index set of variables, replacing `x_i` with `x_(f i)`. -/ def map {α β} (f : α → β) (g : Poly α) : Poly β := ⟨fun v => g <| v ∘ f, Poly.induction (C := fun g => IsPoly (fun v => g (v ∘ f))) - (fun i => by simp; apply IsPoly.proj) (fun n => by simp; apply IsPoly.const) - (fun f g pf pg => by simp; apply IsPoly.sub pf pg) - (fun f g pf pg => by simp; apply IsPoly.mul pf pg) _⟩ + (fun i => by simpa using IsPoly.proj _) (fun n => by simpa using IsPoly.const _) + (fun f g pf pg => by simpa using IsPoly.sub pf pg) + (fun f g pf pg => by simpa using IsPoly.mul pf pg) _⟩ @[simp] theorem map_apply {α β} (f : α → β) (g : Poly α) (v) : map f g v = g (v ∘ f) := rfl diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index d1bdcd7bcdbb3..4087714eb2195 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Michael Stoll -/ import Mathlib.Analysis.PSeries -import Mathlib.Analysis.NormedSpace.FiniteDimension +import Mathlib.Analysis.Normed.Module.FiniteDimension import Mathlib.Data.Complex.FiniteDimensional /-! diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaValues.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaValues.lean index 210f380e20f9b..7f86bf4003157 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaValues.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaValues.lean @@ -50,7 +50,8 @@ theorem cosZeta_two_mul_nat (hk : k ≠ 0) (hx : x ∈ Icc 0 1) : cosZeta x (2 * k) = (-1) ^ (k + 1) * (2 * π) ^ (2 * k) / 2 / (2 * k)! * ((Polynomial.bernoulli (2 * k)).map (algebraMap ℚ ℂ)).eval (x : ℂ) := by rw [← (hasSum_nat_cosZeta x (?_ : 1 < re (2 * k))).tsum_eq] - · refine Eq.trans ?_ <| (congr_arg ofReal' (hasSum_one_div_nat_pow_mul_cos hk hx).tsum_eq).trans ?_ + · refine Eq.trans ?_ <| + (congr_arg ofReal' (hasSum_one_div_nat_pow_mul_cos hk hx).tsum_eq).trans ?_ · rw [ofReal_tsum] refine tsum_congr fun n ↦ ?_ rw [mul_comm (1 / _), mul_one_div, ofReal_div, mul_assoc (2 * π), mul_comm x n, ← mul_assoc, @@ -77,7 +78,8 @@ theorem sinZeta_two_mul_nat_add_one (hk : k ≠ 0) (hx : x ∈ Icc 0 1) : sinZeta x (2 * k + 1) = (-1) ^ (k + 1) * (2 * π) ^ (2 * k + 1) / 2 / (2 * k + 1)! * ((Polynomial.bernoulli (2 * k + 1)).map (algebraMap ℚ ℂ)).eval (x : ℂ) := by rw [← (hasSum_nat_sinZeta x (?_ : 1 < re (2 * k + 1))).tsum_eq] - · refine Eq.trans ?_ <| (congr_arg ofReal' (hasSum_one_div_nat_pow_mul_sin hk hx).tsum_eq).trans ?_ + · refine Eq.trans ?_ <| + (congr_arg ofReal' (hasSum_one_div_nat_pow_mul_sin hk hx).tsum_eq).trans ?_ · rw [ofReal_tsum] refine tsum_congr fun n ↦ ?_ rw [mul_comm (1 / _), mul_one_div, ofReal_div, mul_assoc (2 * π), mul_comm x n, ← mul_assoc] diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index 07c1ef302e400..642151f5382bc 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -7,7 +7,7 @@ import Mathlib.NumberTheory.Cyclotomic.PrimitiveRoots import Mathlib.FieldTheory.Finite.Trace import Mathlib.Algebra.Group.AddChar import Mathlib.Data.ZMod.Units -import Mathlib.Analysis.Complex.Polynomial +import Mathlib.Analysis.Complex.Polynomial.Basic /-! # Additive characters of finite rings and fields diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 8f9c3d87312b6..e41aeaa24d61b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -186,9 +186,8 @@ theorem convexBodyLT'_mem {x : K} : theorem convexBodyLT'_neg_mem (x : E K) (hx : x ∈ convexBodyLT' K f w₀) : -x ∈ convexBodyLT' K f w₀ := by - simp [Set.mem_prod, Prod.fst_neg, Set.mem_pi, Set.mem_univ, Pi.neg_apply, - mem_ball_zero_iff, norm_neg, Real.norm_eq_abs, forall_true_left, Subtype.forall, - Prod.snd_neg, Complex.norm_eq_abs] at hx ⊢ + simp only [Set.mem_prod, Set.mem_pi, Set.mem_univ, mem_ball, dist_zero_right, Real.norm_eq_abs, + true_implies, Subtype.forall, Prod.fst_neg, Pi.neg_apply, norm_neg, Prod.snd_neg] at hx ⊢ convert hx using 3 split_ifs <;> simp diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 2df080fcb3a5e..954e318cf0962 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -5,6 +5,7 @@ Authors: Xavier Roblot -/ import Mathlib.Data.Real.Pi.Bounds import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody +import Mathlib.Tactic.Rify /-! # Number field discriminant @@ -201,22 +202,15 @@ theorem abs_discr_ge (h : 1 < finrank ℚ K) : /-- **Hermite-Minkowski Theorem**. A nontrivial number field has discriminant greater than `2`. -/ theorem abs_discr_gt_two (h : 1 < finrank ℚ K) : 2 < |discr K| := by - have h₁ : 1 ≤ 3 * π / 4 := by - rw [_root_.le_div_iff (by positivity), ← _root_.div_le_iff' (by positivity), one_mul] - linarith [Real.pi_gt_three] - have h₂ : (9 : ℝ) < π ^ 2 := by - rw [← Real.sqrt_lt (by positivity) (by positivity), show Real.sqrt (9 : ℝ) = 3 from - (Real.sqrt_eq_iff_sq_eq (by positivity) (by positivity)).mpr (by norm_num)] - exact Real.pi_gt_three - refine Int.cast_lt.mp <| lt_of_lt_of_le ?_ (abs_discr_ge h) - rw [← _root_.div_lt_iff' (by positivity), Int.cast_ofNat] - refine lt_of_lt_of_le ?_ (pow_le_pow_right (n := 2) h₁ h) - rw [div_pow, _root_.lt_div_iff (by norm_num), mul_pow, - show (2 : ℝ) / (4 / 9) * 4 ^ 2 = 72 by norm_num, - show (3 : ℝ) ^ 2 = 9 by norm_num, - ← _root_.div_lt_iff' (by positivity), - show (72 : ℝ) / 9 = 8 by norm_num] - linarith [h₂] + rw [← Nat.succ_le_iff] at h + rify + calc + (2 : ℝ) < (4 / 9) * (3 * π / 4) ^ 2 := by + nlinarith [Real.pi_gt_three] + _ ≤ (4 / 9 : ℝ) * (3 * π / 4) ^ finrank ℚ K := by + gcongr + linarith [Real.pi_gt_three] + _ ≤ |(discr K : ℝ)| := mod_cast abs_discr_ge h /-! ### Hermite Theorem diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index dae551b8164b1..9c9cc81cf80ee 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Xavier Roblot -/ import Mathlib.Algebra.Algebra.Hom.Rat -import Mathlib.Analysis.Complex.Polynomial +import Mathlib.Analysis.Complex.Polynomial.Basic import Mathlib.NumberTheory.NumberField.Norm import Mathlib.NumberTheory.NumberField.Basic import Mathlib.RingTheory.Norm.Basic diff --git a/Mathlib/NumberTheory/NumberField/EquivReindex.lean b/Mathlib/NumberTheory/NumberField/EquivReindex.lean new file mode 100644 index 0000000000000..49dbcfc3ef6ec --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/EquivReindex.lean @@ -0,0 +1,75 @@ +/- +Copyright (c) 2024 Michail Karatarakis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michail Karatarakis +-/ +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic + +/-! + +# Reindexed basis +This file introduces an equivalence between the set of embeddings of `K` into `ℂ` and the +index set of the chosen basis of the ring of integers of `K`. + +## Tagshouse +number field, algebraic number +-/ + +variable (K : Type*) [Field K] [NumberField K] + +namespace NumberField + +noncomputable section + +open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset + +/-- An equivalence between the set of embeddings of `K` into `ℂ` and the + index set of the chosen basis of the ring of integers of `K`. -/ +abbrev equivReindex : (K →+* ℂ) ≃ (ChooseBasisIndex ℤ (𝓞 K)) := + Fintype.equivOfCardEq <| + by rw [Embeddings.card, ← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank] + +/-- The basis matrix for the embeddings of `K` into `ℂ`. This matrix is formed by + taking the lattice basis vectors of `K` and reindexing them according to the + equivalence `equivReindex`, then transposing the resulting matrix. -/ +abbrev basisMatrix : Matrix (K →+* ℂ) (K →+* ℂ) ℂ := + (Matrix.of fun i ↦ latticeBasis K (equivReindex K i)) + +variable [DecidableEq (K →+* ℂ)] + +theorem det_of_basisMatrix_non_zero : (basisMatrix K).det ≠ 0 := by + let e : (K →+* ℂ) ≃ ChooseBasisIndex ℤ (𝓞 K) := equivReindex K + let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) + RingHom.equivRatAlgHom + rw [show (basisMatrix K) = N by + ext:2; simp only [N, transpose_apply, latticeBasis_apply, integralBasis_apply, + of_apply, apply_at]; rfl, ← pow_ne_zero_iff two_ne_zero] + convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr + (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) + rw [← Algebra.discr_reindex ℚ (integralBasis K) e.symm] + exact (Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two ℚ ℂ + (fun _ => integralBasis K (e _)) RingHom.equivRatAlgHom).symm + +instance : Invertible (basisMatrix K) := invertibleOfIsUnitDet _ + (Ne.isUnit (det_of_basisMatrix_non_zero K)) + +variable {K} + +theorem canonicalEmbedding_eq_basisMatrix_mulVec (α : K) : + canonicalEmbedding K α = (basisMatrix K).transpose.mulVec + (fun i ↦ (((integralBasis K).reindex (equivReindex K).symm).repr α i : ℂ)) := by + ext i + rw [← (latticeBasis K).sum_repr (canonicalEmbedding K α), ← Equiv.sum_comp (equivReindex K)] + simp only [canonicalEmbedding.integralBasis_repr_apply, mulVec, dotProduct, + transpose_apply, of_apply, Fintype.sum_apply, mul_comm, Basis.repr_reindex, + Finsupp.mapDomain_equiv_apply, Equiv.symm_symm, Pi.smul_apply, smul_eq_mul] + +theorem inverse_basisMatrix_mulVec_eq_repr (α : 𝓞 K) : + ∀ i, ((basisMatrix K).transpose)⁻¹.mulVec (fun j => + canonicalEmbedding K (algebraMap (𝓞 K) K α) j) i = + ((integralBasis K).reindex (equivReindex K).symm).repr α i := fun i => by + rw [inv_mulVec_eq_vec (canonicalEmbedding_eq_basisMatrix_mulVec ((algebraMap (𝓞 K) K) α))] + +end + +end NumberField diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean new file mode 100644 index 0000000000000..dcec8b3574d3f --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -0,0 +1,311 @@ +/- +Copyright (c) 2024 Michail Karatarakis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michail Karatarakis +-/ +import Mathlib.NumberTheory.SiegelsLemma +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic +import Mathlib.NumberTheory.NumberField.EquivReindex + +/-! + +# House of an algebraic number +This file defines the house of an algebraic number `α`, which is +the largest of the modulus of its conjugates. + +## References +* [D. Marcus, *Number Fields*][marcus1977number] +* [Keng, H. L, *Introduction to number theory*][keng1982house] + +## Tagshouse +number field, algebraic number, house +-/ + +variable {K : Type*} [Field K] [NumberField K] + +namespace NumberField + +noncomputable section + +open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset + +attribute [local instance] Matrix.seminormedAddCommGroup + +/-- The house of an algebraic number as the norm of its image by the canonical embedding. -/ +def house (α : K) : ℝ := ‖canonicalEmbedding K α‖ + +/-- The house is the largest of the modulus of the conjugates of an algebraic number. -/ +theorem house_eq_sup' (α : K) : + house α = univ.sup' univ_nonempty (fun φ : K →+* ℂ ↦ ‖φ α‖₊) := by + rw [house, ← coe_nnnorm, nnnorm_eq, ← sup'_eq_sup univ_nonempty] + +theorem house_sum_le_sum_house {ι : Type*} (s : Finset ι) (α : ι → K) : + house (∑ i ∈ s, α i) ≤ ∑ i ∈ s, house (α i) := by + simp only [house, map_sum]; apply norm_sum_le_of_le; intros; rfl + +theorem house_nonneg (α : K) : 0 ≤ house α := norm_nonneg _ + +theorem house_mul_le (α β : K) : house (α * β) ≤ house α * house β := by + simp only [house, _root_.map_mul]; apply norm_mul_le + +@[simp] theorem house_intCast (x : ℤ) : house (x : K) = |x| := by + simp only [house, map_intCast, Pi.intCast_def, pi_norm_const, Complex.norm_eq_abs, + Complex.abs_intCast, Int.cast_abs] + +end + +end NumberField + +namespace NumberField.house + +noncomputable section + +variable (K) + +variable [DecidableEq (K →+* ℂ)] + +open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset + +attribute [local instance] Matrix.seminormedAddCommGroup + +/-- `c` is defined as the product of the maximum absolute + value of the entries of the inverse of the matrix `basisMatrix` and `finrank ℚ K`. -/ +private def c := (finrank ℚ K) * ‖((basisMatrix K).transpose)⁻¹‖ + +private theorem c_nonneg : 0 ≤ c K := by + rw [c, mul_nonneg_iff]; left + exact ⟨by simp only [Nat.cast_nonneg], norm_nonneg ((basisMatrix K).transpose)⁻¹⟩ + +theorem basis_repr_abs_le_const_mul_house (α : 𝓞 K) (i : K →+* ℂ) : + Complex.abs (((integralBasis K).reindex (equivReindex K).symm).repr α i) ≤ + (c K) * house (algebraMap (𝓞 K) K α) := by + let σ := canonicalEmbedding K + calc + _ ≤ ∑ j, ‖((basisMatrix K).transpose)⁻¹‖ * Complex.abs (σ (algebraMap (𝓞 K) K α) j) := ?_ + _ ≤ ∑ _ : K →+* ℂ, ‖fun i j => ((basisMatrix K).transpose)⁻¹ i j‖ + * house (algebraMap (𝓞 K) K α) := ?_ + _ = ↑(finrank ℚ K) * ‖((basisMatrix K).transpose)⁻¹‖ * house (algebraMap (𝓞 K) K α) := ?_ + + · rw [← inverse_basisMatrix_mulVec_eq_repr] + apply le_trans + · apply le_trans (AbsoluteValue.sum_le Complex.abs _ _) + · exact sum_le_sum (fun _ _ => (AbsoluteValue.map_mul Complex.abs _ _).le) + · apply sum_le_sum (fun _ _ => mul_le_mul_of_nonneg_right ?_ + (AbsoluteValue.nonneg Complex.abs _)) + · exact norm_entry_le_entrywise_sup_norm ((basisMatrix K).transpose)⁻¹ + · apply sum_le_sum; intros j _ + apply mul_le_mul_of_nonneg_left _ (norm_nonneg fun i j ↦ ((basisMatrix K).transpose)⁻¹ i j) + · exact norm_le_pi_norm (σ ((algebraMap (𝓞 K) K) α)) j + · rw [sum_const, card_univ, nsmul_eq_mul, Embeddings.card, mul_assoc] + +/-- `newBasis K` defines a reindexed basis of the ring of integers of `K`, + adjusted by the inverse of the equivalence `equivReindex`. -/ +private def newBasis := (RingOfIntegers.basis K).reindex (equivReindex K).symm + +/-- `supOfBasis K` calculates the supremum of the absolute values of + the elements in `newBasis K`. -/ +private def supOfBasis : ℝ := univ.sup' univ_nonempty + fun r ↦ house (algebraMap (𝓞 K) K (newBasis K r)) + +private theorem supOfBasis_nonneg : 0 ≤ supOfBasis K := by + simp only [supOfBasis, le_sup'_iff, mem_univ, and_self, + exists_const, house_nonneg] + +variable {α : Type*} {β : Type*} [Fintype α] [Fintype β] [DecidableEq β] [DecidableEq α] + +variable (a : Matrix α β (𝓞 K)) + +/-- `a' K a` returns the integer coefficients of the basis vector in the + expansion of the product of an algebraic integer and a basis vectors. -/ +private def a' : α → β → (K →+* ℂ) → (K →+* ℂ) → ℤ := fun k l r => + (newBasis K).repr (a k l * (newBasis K) r) + +/--`asiegel K a` the integer matrix of the coefficients of the + product of matrix elements and basis vectors -/ +private def asiegel : Matrix (α × (K →+* ℂ)) (β × (K →+* ℂ)) ℤ := fun k l => a' K a k.1 l.1 l.2 k.2 + +variable (ha : a ≠ 0) + +private theorem asiegel_ne_0 : asiegel K a ≠ 0 := by + simp (config := { unfoldPartialApp := true }) only [asiegel, a'] + simp only [ne_eq] + rw [Function.funext_iff]; intros hs + simp only [Prod.forall] at hs; + apply ha + rw [← Matrix.ext_iff]; intros k' l + specialize hs k' + let ⟨b⟩ := Fintype.card_pos_iff.1 (Fintype.card_pos (α := (K →+* ℂ))) + have := ((newBasis K).repr.map_eq_zero_iff (x := (a k' l * (newBasis K) b))).1 <| by + ext b' + specialize hs b' + rw [Function.funext_iff] at hs + simp only [Prod.forall] at hs + apply hs + simp only [mul_eq_zero] at this + exact this.resolve_right (Basis.ne_zero (newBasis K) b) + +variable {p q : ℕ} + (cardα : Fintype.card α = p) (cardβ : Fintype.card β = q) + (h0p : 0 < p) (hpq : p < q) + (x : β × (K →+* ℂ) → ℤ) + (hxl : x ≠ 0) + (hmulvec0 : asiegel K a *ᵥ x = 0) + (hxbound : ‖x‖ ≤ (q * finrank ℚ K * ‖asiegel K a‖) ^ ((p : ℝ) / (q - p))) + +/-- `ξ` is the the product of `x (l, r)` and the `r`-th basis element of the newBasis of `K`. -/ +private def ξ : β → 𝓞 K := fun l => ∑ r : K →+* ℂ, x (l, r) * (newBasis K r) + +private theorem ξ_ne_0 : ξ K x ≠ 0 := by + intro H + apply hxl + ext ⟨l, r⟩ + rw [Function.funext_iff] at H + have hblin := Basis.linearIndependent (newBasis K) + simp only [zsmul_eq_mul, Fintype.linearIndependent_iff] at hblin + exact hblin (fun r ↦ x (l,r)) (H _) r + +private theorem lin_1 (l k r) : a k l * (newBasis K) r = + ∑ u, (a' K a k l r u) * (newBasis K) u := by + simp only [Basis.sum_repr (newBasis K) (a k l * (newBasis K) r), a', ← zsmul_eq_mul] + +private theorem ξ_mulVec_eq_0 : a *ᵥ ξ K x = 0 := by + funext k; simp only [Pi.zero_apply]; rw [eq_comm] + + have lin_0 : ∀ u, ∑ r, ∑ l, (a' K a k l r u * x (l, r) : 𝓞 K) = 0 := by + intros u + have hξ := ξ_ne_0 K x hxl + rw [Ne, Function.funext_iff, not_forall] at hξ + rcases hξ with ⟨l, hξ⟩ + rw [Function.funext_iff] at hmulvec0 + specialize hmulvec0 ⟨k, u⟩ + simp only [Fintype.sum_prod_type, mulVec, dotProduct, asiegel] at hmulvec0 + rw [sum_comm] at hmulvec0 + exact mod_cast hmulvec0 + + have : 0 = ∑ u, (∑ r, ∑ l, a' K a k l r u * x (l, r) : 𝓞 K) * (newBasis K) u := by + simp only [lin_0, zero_mul, sum_const_zero] + + have : 0 = ∑ r, ∑ l, x (l, r) * ∑ u, a' K a k l r u * (newBasis K) u := by + conv at this => enter [2, 2, u]; rw [sum_mul] + rw [sum_comm] at this + rw [this]; congr 1; ext1 r + conv => enter [1, 2, l]; rw [sum_mul] + rw [sum_comm]; congr 1; ext1 r + rw [mul_sum]; congr 1; ext1 r + ring + rw [sum_comm] at this + rw [this]; congr 1; ext1 l + rw [ξ, mul_sum]; congr 1; ext1 l + rw [← lin_1]; ring + +variable {A : ℝ} (habs : ∀ k l, (house ((algebraMap (𝓞 K) K) (a k l))) ≤ A) + +/-- `c₂` is the product of the maximum of `1` and `c`, and `supOfBasis`. -/ +private abbrev c₂ := max 1 (c K) * (supOfBasis K) + +private theorem c₂_nonneg : 0 ≤ c₂ K := + mul_nonneg (le_trans zero_le_one (le_max_left ..)) (supOfBasis_nonneg _) + +variable (Apos : 0 ≤ A) + +private theorem asiegel_remark : ‖asiegel K a‖ ≤ c₂ K * A := by + rw [Matrix.norm_le_iff] + · intro kr lu + calc + ‖asiegel K a kr lu‖ = |asiegel K a kr lu| := ?_ + _ ≤ (c K) * + house ((algebraMap (𝓞 K) K) (a kr.1 lu.1 * ((newBasis K) lu.2))) := ?_ + _ ≤ (c K) * house ((algebraMap (𝓞 K) K) (a kr.1 lu.1)) * + house ((algebraMap (𝓞 K) K) ((newBasis K) lu.2)) := ?_ + _ ≤ (c K) * A * house ((algebraMap (𝓞 K) K) ((newBasis K) lu.2)) := ?_ + _ ≤ (c K) * A * (supOfBasis K) := ?_ + _ ≤ (c₂ K) * A := ?_ + · simp only [Int.cast_abs, ← Real.norm_eq_abs (asiegel K a kr lu)]; rfl + · have remark := basis_repr_abs_le_const_mul_house K + simp only [Basis.repr_reindex, Finsupp.mapDomain_equiv_apply, + integralBasis_repr_apply, eq_intCast, Rat.cast_intCast, + Complex.abs_intCast] at remark + exact mod_cast remark ((a kr.1 lu.1 * ((newBasis K) lu.2))) kr.2 + · simp only [house, _root_.map_mul, mul_assoc] + exact mul_le_mul_of_nonneg_left (norm_mul_le _ _) (c_nonneg K) + · rw [mul_assoc, mul_assoc] + apply mul_le_mul_of_nonneg_left ?_ (c_nonneg K) + · apply mul_le_mul_of_nonneg_right (habs kr.1 lu.1) ?_ + · exact norm_nonneg ((canonicalEmbedding K) ((algebraMap (𝓞 K) K) + ((newBasis K) lu.2))) + · apply mul_le_mul_of_nonneg_left ?_ (mul_nonneg (c_nonneg K) Apos) + · simp only [supOfBasis, le_sup'_iff, mem_univ]; use lu.2 + · rw [mul_right_comm] + exact mul_le_mul_of_nonneg_right + (mul_le_mul_of_nonneg_right (le_max_right ..) (supOfBasis_nonneg K)) Apos + · rw [mul_nonneg_iff]; left; exact ⟨c₂_nonneg K, Apos⟩ + +/-- `c₁ K` is the product of `finrank ℚ K` and `c₂ K` and depends on `K`. -/ +private def c₁ := finrank ℚ K * c₂ K + +private theorem house_le_bound : ∀ l, house (ξ K x l).1 ≤ (c₁ K) * + ((c₁ K * q * A)^((p : ℝ) / (q - p))) := by + let h := finrank ℚ K + intros l + calc _ = house (algebraMap (𝓞 K) K (∑ r, (x (l, r)) * ((newBasis K) r))) := rfl + _ ≤ ∑ r, house (((algebraMap (𝓞 K) K) (x (l, r))) * + ((algebraMap (𝓞 K) K) ((newBasis K) r))) := ?_ + _ ≤ ∑ r, ‖x (l,r)‖ * house ((algebraMap (𝓞 K) K) ((newBasis K) r)) := ?_ + _ ≤ ∑ r, ‖x (l, r)‖ * (supOfBasis K) := ?_ + _ ≤ ∑ _r : K →+* ℂ, ((↑q * h * ‖asiegel K a‖) ^ ((p : ℝ) / (q - p))) * supOfBasis K := ?_ + _ ≤ h * (c₂ K) * ((q * c₁ K * A) ^ ((p : ℝ) / (q - p))) := ?_ + _ ≤ c₁ K * ((c₁ K * ↑q * A) ^ ((p : ℝ) / (q - p))) := ?_ + · simp_rw [← _root_.map_mul, map_sum]; apply house_sum_le_sum_house + · apply sum_le_sum; intros r _; convert house_mul_le .. + simp only [map_intCast, house_intCast, Int.cast_abs, Int.norm_eq_abs] + · apply sum_le_sum; intros r _; unfold supOfBasis + apply mul_le_mul_of_nonneg_left ?_ (norm_nonneg (x (l,r))) + · simp only [le_sup'_iff, mem_univ, true_and]; use r + · apply sum_le_sum; intros r _ + apply mul_le_mul_of_nonneg_right ?_ (supOfBasis_nonneg K) + exact le_trans (norm_le_pi_norm x ⟨l, r⟩) hxbound + · simp only [Nat.cast_mul, sum_const, card_univ, nsmul_eq_mul] + rw [Embeddings.card, mul_comm _ (supOfBasis K), c₂, c₁, ← mul_assoc] + apply mul_le_mul + · apply mul_le_mul_of_nonneg_left ?_ (Nat.cast_nonneg' _) + · exact le_mul_of_one_le_left (supOfBasis_nonneg K) (le_max_left ..) + · apply Real.rpow_le_rpow (mul_nonneg (mul_nonneg (Nat.cast_nonneg' _) (Nat.cast_nonneg' _)) + (norm_nonneg _)) + · rw [← mul_assoc, mul_assoc (_*_)] + apply mul_le_mul_of_nonneg_left (asiegel_remark K a habs Apos) + (mul_nonneg (Nat.cast_nonneg' _) (Nat.cast_nonneg _)) + · exact div_nonneg (Nat.cast_nonneg' _) (sub_nonneg.2 (mod_cast hpq.le)) + · apply Real.rpow_nonneg + exact mul_nonneg (mul_nonneg (Nat.cast_nonneg' _) (Nat.cast_nonneg' _)) + (norm_nonneg _) + · exact mul_nonneg (Nat.cast_nonneg' _) (mul_nonneg (le_trans zero_le_one (le_max_left ..)) + (supOfBasis_nonneg _)) + · rw [mul_comm (q : ℝ) (c₁ K)]; rfl + +/-- There exists a "small" non-zero algebraic integral solution of an + non-trivial underdetermined system of linear equations with algebraic integer coefficients.-/ +theorem exists_ne_zero_int_vec_house_le : + ∃ (ξ : β → 𝓞 K), ξ ≠ 0 ∧ a *ᵥ ξ = 0 ∧ + ∀ l, house (ξ l).1 ≤ c₁ K * ((c₁ K * q * A) ^ ((p : ℝ) / (q - p))) := by + let h := finrank ℚ K + have hphqh : p * h < q * h := mul_lt_mul_of_pos_right hpq finrank_pos + have h0ph : 0 < p * h := by rw [mul_pos_iff]; constructor; exact ⟨h0p, finrank_pos⟩ + have hfinp : Fintype.card (α × (K →+* ℂ)) = p * h := by + rw [Fintype.card_prod, cardα, Embeddings.card] + have hfinq : Fintype.card (β × (K →+* ℂ)) = q * h := by + rw [Fintype.card_prod, cardβ, Embeddings.card] + have ⟨x, hxl, hmulvec0, hxbound⟩ := + Int.Matrix.exists_ne_zero_int_vec_norm_le' (asiegel K a) + (by rwa [hfinp, hfinq]) (by rwa [hfinp]) (asiegel_ne_0 K a ha) + simp only [hfinp, hfinq, Nat.cast_mul] at hmulvec0 hxbound + rw [← sub_mul, mul_div_mul_right _ _ (mod_cast finrank_pos.ne')] at hxbound + have Apos : 0 ≤ A := by + have ⟨k⟩ := Fintype.card_pos_iff.1 (cardα ▸ h0p) + have ⟨l⟩ := Fintype.card_pos_iff.1 (cardβ ▸ h0p.trans hpq) + exact le_trans (house_nonneg _) (habs k l) + use ξ K x, ξ_ne_0 K x hxl, ξ_mulVec_eq_0 K a x hxl hmulvec0, + house_le_bound K a hpq x hxbound habs Apos + +end + +end NumberField.house diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index 301f32a6dae1d..39025853f4b6f 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -280,7 +280,7 @@ theorem ne_zero_iff_nequiv_zero (f : PadicSeq p) : mk f ≠ 0 ↔ ¬f ≈ 0 := theorem norm_const (q : ℚ) : norm (const (padicNorm p) q) = padicNorm p q := if hq : q = 0 then by - have : const (padicNorm p) q ≈ 0 := by simp [hq]; apply Setoid.refl (const (padicNorm p) 0) + have : const (padicNorm p) q ≈ 0 := by simpa [hq] using Setoid.refl (const (padicNorm p) 0) subst hq; simp [norm, this] else by have : ¬const (padicNorm p) q ≈ 0 := not_equiv_zero_const_of_nonzero hq diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index ba3aaf0cfbb2f..90bc427eeab82 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -486,7 +486,7 @@ variable {p a b : ℕ} [hp : Fact p.Prime] /-- A rewrite lemma for `padicValNat p (a * b)` with conditions `a ≠ 0`, `b ≠ 0`. -/ protected theorem mul : a ≠ 0 → b ≠ 0 → padicValNat p (a * b) = padicValNat p a + padicValNat p b := - mod_cast @padicValRat.mul p _ a b + mod_cast padicValRat.mul (p := p) (q := a) (r := b) protected theorem div_of_dvd (h : b ∣ a) : padicValNat p (a / b) = padicValNat p a - padicValNat p b := by diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 8f88bdb0d3587..10ce7abc5d7bd 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -197,7 +197,7 @@ theorem isPell_pellZd : ∀ n : ℕ, IsPell (pellZd a1 n) | 0 => rfl | n + 1 => by let o := isPell_one a1 - simp; exact Pell.isPell_mul (isPell_pellZd n) o + simpa using Pell.isPell_mul (isPell_pellZd n) o @[simp] theorem pell_eqz (n : ℕ) : xz a1 n * xz a1 n - d a1 * yz a1 n * yz a1 n = 1 := @@ -401,7 +401,7 @@ theorem y_dvd_iff (m n) : yn a1 m ∣ yn a1 n ↔ m ∣ n := theorem xy_modEq_yn (n) : ∀ k, xn a1 (n * k) ≡ xn a1 n ^ k [MOD yn a1 n ^ 2] ∧ yn a1 (n * k) ≡ k * xn a1 n ^ (k - 1) * yn a1 n [MOD yn a1 n ^ 3] - | 0 => by constructor <;> simp <;> exact Nat.ModEq.refl _ + | 0 => by constructor <;> simpa using Nat.ModEq.refl _ | k + 1 => by let ⟨hx, hy⟩ := xy_modEq_yn n k have L : xn a1 (n * k) * xn a1 n + d a1 * yn a1 (n * k) * yn a1 n ≡ @@ -745,7 +745,7 @@ end theorem xy_modEq_of_modEq {a b c} (a1 : 1 < a) (b1 : 1 < b) (h : a ≡ b [MOD c]) : ∀ n, xn a1 n ≡ xn b1 n [MOD c] ∧ yn a1 n ≡ yn b1 n [MOD c] | 0 => by constructor <;> rfl - | 1 => by simp; exact ⟨h, ModEq.refl 1⟩ + | 1 => by simpa using ⟨h, ModEq.refl 1⟩ | n + 2 => ⟨(xy_modEq_of_modEq a1 b1 h n).left.add_right_cancel <| by rw [xn_succ_succ a1, xn_succ_succ b1] @@ -787,7 +787,7 @@ theorem matiyasevic {a k x y} : have vp : 0 < v := strictMono_y a1 (lt_trans zero_lt_one m1) have b1 : 1 < b := have : xn a1 1 < u := strictMono_x a1 m1 - have : a < u := by simp at this; exact this + have : a < u := by simpa using this lt_of_lt_of_le a1 <| by delta ModEq at ba; rw [Nat.mod_eq_of_lt this] at ba; rw [← ba] apply Nat.mod_le diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 2a1fea033fd97..2b91d8670f7dc 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -1001,8 +1001,8 @@ theorem isAtomic_of_isCoatomic_of_complementedLattice_of_isModular [IsCoatomic isCoatomic_dual_iff_isAtomic.1 isCoatomic_of_isAtomic_of_complementedLattice_of_isModular theorem isAtomic_iff_isCoatomic : IsAtomic α ↔ IsCoatomic α := - ⟨fun h => @isCoatomic_of_isAtomic_of_complementedLattice_of_isModular _ _ _ _ _ h, fun h => - @isAtomic_of_isCoatomic_of_complementedLattice_of_isModular _ _ _ _ _ h⟩ + ⟨fun _ => isCoatomic_of_isAtomic_of_complementedLattice_of_isModular, + fun _ => isAtomic_of_isCoatomic_of_complementedLattice_of_isModular⟩ /-- A complemented modular atomic lattice is strongly atomic. Not an instance to prevent loops. -/ diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index ea885c1d8bcb7..942a70f8f6bec 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -751,12 +751,14 @@ protected abbrev Function.Injective.generalizedBooleanAlgebra [Sup α] [Inf α] -- See note [reducible non-instances] /-- Pullback a `BooleanAlgebra` along an injection. -/ protected abbrev Function.Injective.booleanAlgebra [Sup α] [Inf α] [Top α] [Bot α] [HasCompl α] - [SDiff α] [BooleanAlgebra β] (f : α → β) (hf : Injective f) + [SDiff α] [HImp α] [BooleanAlgebra β] (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f aᶜ = (f a)ᶜ) - (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : BooleanAlgebra α where + (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) : + BooleanAlgebra α where __ := hf.generalizedBooleanAlgebra f map_sup map_inf map_bot map_sdiff compl := compl + himp := himp top := ⊤ le_top a := (@le_top β _ _ _).trans map_top.ge bot_le a := map_bot.le.trans bot_le @@ -765,6 +767,7 @@ protected abbrev Function.Injective.booleanAlgebra [Sup α] [Inf α] [Top α] [B sdiff_eq a b := by refine hf ((map_sdiff _ _).trans (sdiff_eq.trans ?_)) rw [map_inf, map_compl] + himp_eq a b := hf $ (map_himp _ _).trans $ himp_eq.trans $ by rw [map_sup, map_compl] end lift diff --git a/Mathlib/Order/BooleanGenerators.lean b/Mathlib/Order/BooleanGenerators.lean index 56e04b99df110..d91b966fb1edd 100644 --- a/Mathlib/Order/BooleanGenerators.lean +++ b/Mathlib/Order/BooleanGenerators.lean @@ -62,13 +62,13 @@ structure BooleanGenerators (S : Set α) : Prop where namespace BooleanGenerators -variable {S : Set α} (hS : BooleanGenerators S) +variable {S : Set α} -lemma mono {T : Set α} (hTS : T ⊆ S) : BooleanGenerators T where +lemma mono (hS : BooleanGenerators S) {T : Set α} (hTS : T ⊆ S) : BooleanGenerators T where isAtom I hI := hS.isAtom I (hTS hI) finitelyAtomistic := fun s a hs ↦ hS.finitelyAtomistic s a (le_trans hs hTS) -lemma atomistic (a : α) (ha : a ≤ sSup S) : ∃ T ⊆ S, a = sSup T := by +lemma atomistic (hS : BooleanGenerators S) (a : α) (ha : a ≤ sSup S) : ∃ T ⊆ S, a = sSup T := by obtain ⟨C, hC, rfl⟩ := IsCompactlyGenerated.exists_sSup_eq a have aux : ∀ b : α, IsCompactElement b → b ≤ sSup S → ∃ T ⊆ S, b = sSup T := by intro b hb hbS @@ -97,13 +97,14 @@ lemma atomistic (a : α) (ha : a ≤ sSup S) : ∃ T ⊆ S, a = sSup T := by rw [← hT₂] exact le_sSup hbC -lemma isAtomistic_of_sSup_eq_top (h : sSup S = ⊤) : IsAtomistic α := by +lemma isAtomistic_of_sSup_eq_top (hS : BooleanGenerators S) (h : sSup S = ⊤) : + IsAtomistic α := by refine ⟨fun a ↦ ?_⟩ obtain ⟨s, hs, hs'⟩ := hS.atomistic a (h ▸ le_top) exact ⟨s, hs', fun I hI ↦ hS.isAtom I (hs hI)⟩ -lemma mem_of_isAtom_of_le_sSup_atoms (a : α) (ha : IsAtom a) (haS : a ≤ sSup S) : - a ∈ S := by +lemma mem_of_isAtom_of_le_sSup_atoms (hS : BooleanGenerators S) (a : α) (ha : IsAtom a) + (haS : a ≤ sSup S) : a ∈ S := by obtain ⟨T, hT, rfl⟩ := hS.atomistic a haS obtain rfl | ⟨a, haT⟩ := T.eq_empty_or_nonempty · simp only [sSup_empty] at ha @@ -113,7 +114,7 @@ lemma mem_of_isAtom_of_le_sSup_atoms (a : α) (ha : IsAtom a) (haS : a ≤ sSup rwa [ha.le_iff_eq, eq_comm] at this exact (hS.isAtom a (hT haT)).1 -lemma sSup_inter {T₁ T₂ : Set α} (hT₁ : T₁ ⊆ S) (hT₂ : T₂ ⊆ S) : +lemma sSup_inter (hS : BooleanGenerators S) {T₁ T₂ : Set α} (hT₁ : T₁ ⊆ S) (hT₂ : T₂ ⊆ S) : sSup (T₁ ∩ T₂) = (sSup T₁) ⊓ (sSup T₂) := by apply le_antisymm · apply le_inf @@ -133,7 +134,8 @@ lemma sSup_inter {T₁ T₂ : Set α} (hT₁ : T₁ ⊆ S) (hT₂ : T₂ ⊆ S) · exact (_root_.le_sSup hI).trans (hX'.ge.trans inf_le_right) /-- A lattice generated by boolean generators is a distributive lattice. -/ -def distribLattice_of_sSup_eq_top (h : sSup S = ⊤) : DistribLattice α where +def distribLattice_of_sSup_eq_top (hS : BooleanGenerators S) (h : sSup S = ⊤) : + DistribLattice α where le_sup_inf a b c := by obtain ⟨Ta, hTa, rfl⟩ := hS.atomistic a (h ▸ le_top) obtain ⟨Tb, hTb, rfl⟩ := hS.atomistic b (h ▸ le_top) @@ -145,26 +147,28 @@ def distribLattice_of_sSup_eq_top (h : sSup S = ⊤) : DistribLattice α where simp only [Set.union_subset_iff, Set.mem_inter_iff, Set.mem_union] tauto -lemma complementedLattice_of_sSup_eq_top (h : sSup S = ⊤) : ComplementedLattice α := by +lemma complementedLattice_of_sSup_eq_top (hS : BooleanGenerators S) (h : sSup S = ⊤) : + ComplementedLattice α := by let _i := hS.distribLattice_of_sSup_eq_top h have _i₁ := isAtomistic_of_sSup_eq_top hS h apply complementedLattice_of_isAtomistic /-- A compactly generated complete lattice generated by boolean generators is a boolean algebra. -/ noncomputable -def booleanAlgebra_of_sSup_eq_top (h : sSup S = ⊤) : BooleanAlgebra α := +def booleanAlgebra_of_sSup_eq_top (hS : BooleanGenerators S) (h : sSup S = ⊤) : BooleanAlgebra α := let _i := hS.distribLattice_of_sSup_eq_top h have := hS.complementedLattice_of_sSup_eq_top h DistribLattice.booleanAlgebraOfComplemented α -lemma sSup_le_sSup_iff_of_atoms (X Y : Set α) (hX : X ⊆ S) (hY : Y ⊆ S) : +lemma sSup_le_sSup_iff_of_atoms (hS : BooleanGenerators S) (X Y : Set α) (hX : X ⊆ S) (hY : Y ⊆ S) : sSup X ≤ sSup Y ↔ X ⊆ Y := by refine ⟨?_, sSup_le_sSup⟩ intro h a ha apply (hS.mono hY).mem_of_isAtom_of_le_sSup_atoms _ _ ((le_sSup ha).trans h) exact (hS.mono hX).isAtom a ha -lemma eq_atoms_of_sSup_eq_top (h : sSup S = ⊤) : S = {a : α | IsAtom a} := by +lemma eq_atoms_of_sSup_eq_top (hS : BooleanGenerators S) (h : sSup S = ⊤) : + S = {a : α | IsAtom a} := by apply le_antisymm · exact hS.isAtom intro a ha diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index fafa4b0f7d415..ba6ccf6886cc6 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -1144,12 +1144,14 @@ theorem image2_lowerBounds_lowerBounds_subset : image2_subset_iff.2 fun _ ha _ hb ↦ mem_lowerBounds_image2 h₀ h₁ ha hb /-- See also `Monotone.map_bddAbove`. -/ -protected theorem BddAbove.image2 : BddAbove s → BddAbove t → BddAbove (image2 f s t) := by +protected theorem BddAbove.image2 (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) : + BddAbove s → BddAbove t → BddAbove (image2 f s t) := by rintro ⟨a, ha⟩ ⟨b, hb⟩ exact ⟨f a b, mem_upperBounds_image2 h₀ h₁ ha hb⟩ /-- See also `Monotone.map_bddBelow`. -/ -protected theorem BddBelow.image2 : BddBelow s → BddBelow t → BddBelow (image2 f s t) := by +protected theorem BddBelow.image2 (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) : + BddBelow s → BddBelow t → BddBelow (image2 f s t) := by rintro ⟨a, ha⟩ ⟨b, hb⟩ exact ⟨f a b, mem_lowerBounds_image2 h₀ h₁ ha hb⟩ diff --git a/Mathlib/Order/CompleteBooleanAlgebra.lean b/Mathlib/Order/CompleteBooleanAlgebra.lean index fa69f65971a68..38fc91ba70639 100644 --- a/Mathlib/Order/CompleteBooleanAlgebra.lean +++ b/Mathlib/Order/CompleteBooleanAlgebra.lean @@ -54,29 +54,68 @@ universe u v w w' variable {α : Type u} {β : Type v} {ι : Sort w} {κ : ι → Sort w'} +/-- Structure containing the minimal axioms required to check that an order is a frame. Do NOT use, +except for implementing `Order.Frame` via `Order.Frame.ofMinimalAxioms`. + +This structure omits the `himp`, `compl` fields, which can be recovered using +`Order.Frame.ofMinimalAxioms`. -/ +class Order.Frame.MinimalAxioms (α : Type u) extends CompleteLattice α where + inf_sSup_le_iSup_inf (a : α) (s : Set α) : a ⊓ sSup s ≤ ⨆ b ∈ s, a ⊓ b + +/-- Structure containing the minimal axioms required to check that an order is a coframe. Do NOT +use, except for implementing `Order.Coframe` via `Order.Coframe.ofMinimalAxioms`. + +This structure omits the `sdiff`, `hnot` fields, which can be recovered using +`Order.Coframe.ofMinimalAxioms`. -/ +class Order.Coframe.MinimalAxioms (α : Type u) extends CompleteLattice α where + iInf_sup_le_sup_sInf (a : α) (s : Set α) : ⨅ b ∈ s, a ⊔ b ≤ a ⊔ sInf s + /-- A frame, aka complete Heyting algebra, is a complete lattice whose `⊓` distributes over `⨆`. -/ -class Order.Frame (α : Type*) extends CompleteLattice α where +class Order.Frame (α : Type*) extends CompleteLattice α, HeytingAlgebra α where /-- `⊓` distributes over `⨆`. -/ inf_sSup_le_iSup_inf (a : α) (s : Set α) : a ⊓ sSup s ≤ ⨆ b ∈ s, a ⊓ b /-- A coframe, aka complete Brouwer algebra or complete co-Heyting algebra, is a complete lattice whose `⊔` distributes over `⨅`. -/ -class Order.Coframe (α : Type*) extends CompleteLattice α where +class Order.Coframe (α : Type*) extends CompleteLattice α, CoheytingAlgebra α where /-- `⊔` distributes over `⨅`. -/ iInf_sup_le_sup_sInf (a : α) (s : Set α) : ⨅ b ∈ s, a ⊔ b ≤ a ⊔ sInf s open Order +/-- Structure containing the minimal axioms required to check that an order is a complete +distributive lattice. Do NOT use, except for implementing `CompleteDistribLattice` via +`CompleteDistribLattice.ofMinimalAxioms`. + +This structure omits the `himp`, `compl`, `sdiff`, `hnot` fields, which can be recovered using +`CompleteDistribLattice.ofMinimalAxioms`. -/ +structure CompleteDistribLattice.MinimalAxioms (α : Type u) + extends CompleteLattice α, Frame.MinimalAxioms α, Coframe.MinimalAxioms α where + +-- We give those projections better name further down +attribute [nolint docBlame] CompleteDistribLattice.MinimalAxioms.toMinimalAxioms + CompleteDistribLattice.MinimalAxioms.toMinimalAxioms_1 + /-- A complete distributive lattice is a complete lattice whose `⊔` and `⊓` respectively distribute over `⨅` and `⨆`. -/ -class CompleteDistribLattice (α : Type*) extends Frame α, Coframe α +class CompleteDistribLattice (α : Type*) extends Frame α, Coframe α, BiheytingAlgebra α /-- In a complete distributive lattice, `⊔` distributes over `⨅`. -/ add_decl_doc CompleteDistribLattice.iInf_sup_le_sup_sInf +/-- Structure containing the minimal axioms required to check that an order is a completely +distributive. Do NOT use, except for implementing `CompletelyDistribLattice` via +`CompletelyDistribLattice.ofMinimalAxioms`. + +This structure omits the `himp`, `compl`, `sdiff`, `hnot` fields, which can be recovered using +`CompletelyDistribLattice.ofMinimalAxioms`. -/ +structure CompletelyDistribLattice.MinimalAxioms (α : Type u) extends CompleteLattice α where + protected iInf_iSup_eq {ι : Type u} {κ : ι → Type u} (f : ∀ a, κ a → α) : + (⨅ a, ⨆ b, f a b) = ⨆ g : ∀ a, κ a, ⨅ a, f a (g a) + /-- A completely distributive lattice is a complete lattice whose `⨅` and `⨆` distribute over each other. -/ -class CompletelyDistribLattice (α : Type u) extends CompleteLattice α where +class CompletelyDistribLattice (α : Type u) extends CompleteLattice α, BiheytingAlgebra α where protected iInf_iSup_eq {ι : Type u} {κ : ι → Type u} (f : ∀ a, κ a → α) : (⨅ a, ⨆ b, f a b) = ⨆ g : ∀ a, κ a, ⨅ a, f a (g a) @@ -84,25 +123,131 @@ theorem le_iInf_iSup [CompleteLattice α] {f : ∀ a, κ a → α} : (⨆ g : ∀ a, κ a, ⨅ a, f a (g a)) ≤ ⨅ a, ⨆ b, f a b := iSup_le fun _ => le_iInf fun a => le_trans (iInf_le _ a) (le_iSup _ _) -theorem iInf_iSup_eq [CompletelyDistribLattice α] {f : ∀ a, κ a → α} : - (⨅ a, ⨆ b, f a b) = ⨆ g : ∀ a, κ a, ⨅ a, f a (g a) := - (le_antisymm · le_iInf_iSup) <| calc +lemma iSup_iInf_le [CompleteLattice α] {f : ∀ a, κ a → α} : + ⨆ a, ⨅ b, f a b ≤ ⨅ g : ∀ a, κ a, ⨆ a, f a (g a) := + le_iInf_iSup (α := αᵒᵈ) + +namespace Order.Frame.MinimalAxioms +variable (minAx : MinimalAxioms α) {s : Set α} {a b : α} + +lemma inf_sSup_eq : a ⊓ sSup s = ⨆ b ∈ s, a ⊓ b := + (minAx.inf_sSup_le_iSup_inf _ _).antisymm iSup_inf_le_inf_sSup + +lemma sSup_inf_eq : sSup s ⊓ b = ⨆ a ∈ s, a ⊓ b := by + simpa only [inf_comm] using @inf_sSup_eq α _ s b + +lemma iSup_inf_eq (f : ι → α) (a : α) : (⨆ i, f i) ⊓ a = ⨆ i, f i ⊓ a := by + rw [iSup, sSup_inf_eq, iSup_range] + +lemma inf_iSup_eq (a : α) (f : ι → α) : (a ⊓ ⨆ i, f i) = ⨆ i, a ⊓ f i := by + simpa only [inf_comm] using minAx.iSup_inf_eq f a + +lemma inf_iSup₂_eq {f : ∀ i, κ i → α} (a : α) : (a ⊓ ⨆ i, ⨆ j, f i j) = ⨆ i, ⨆ j, a ⊓ f i j := by + simp only [inf_iSup_eq] + +/-- The `Order.Frame.MinimalAxioms` element corresponding to a frame. -/ +def of [Frame α] : MinimalAxioms α := { ‹Frame α› with } + +end MinimalAxioms + +/-- Construct a frame instance using the minimal amount of work needed. + +This sets `a ⇨ b := sSup {c | c ⊓ a ≤ b}` and `aᶜ := a ⇨ ⊥`. -/ +-- See note [reducible non instances] +abbrev ofMinimalAxioms (minAx : MinimalAxioms α) : Frame α where + __ := minAx + himp a b := sSup {c | c ⊓ a ≤ b} + le_himp_iff a b c := + ⟨fun h ↦ (inf_le_inf_right _ h).trans (by simp [minAx.sSup_inf_eq]), fun h ↦ le_sSup h⟩ + himp_bot _ := rfl + +end Order.Frame + +namespace Order.Coframe.MinimalAxioms +variable (minAx : MinimalAxioms α) {s : Set α} {a b : α} + +lemma sup_sInf_eq : a ⊔ sInf s = ⨅ b ∈ s, a ⊔ b := + sup_sInf_le_iInf_sup.antisymm (minAx.iInf_sup_le_sup_sInf _ _) + +lemma sInf_sup_eq : sInf s ⊔ b = ⨅ a ∈ s, a ⊔ b := by + simpa only [sup_comm] using @sup_sInf_eq α _ s b + +lemma iInf_sup_eq (f : ι → α) (a : α) : (⨅ i, f i) ⊔ a = ⨅ i, f i ⊔ a := by + rw [iInf, sInf_sup_eq, iInf_range] + +lemma sup_iInf_eq (a : α) (f : ι → α) : (a ⊔ ⨅ i, f i) = ⨅ i, a ⊔ f i := by + simpa only [sup_comm] using minAx.iInf_sup_eq f a + +lemma sup_iInf₂_eq {f : ∀ i, κ i → α} (a : α) : (a ⊔ ⨅ i, ⨅ j, f i j) = ⨅ i, ⨅ j, a ⊔ f i j := by + simp only [sup_iInf_eq] + +/-- The `Order.Coframe.MinimalAxioms` element corresponding to a frame. -/ +def of [Coframe α] : MinimalAxioms α := { ‹Coframe α› with } + +end MinimalAxioms + +/-- Construct a coframe instance using the minimal amount of work needed. + +This sets `a \ b := sInf {c | a ≤ b ⊔ c}` and `¬a := ⊤ \ a`. -/ +-- See note [reducible non instances] +abbrev ofMinimalAxioms (minAx : MinimalAxioms α) : Coframe α where + __ := minAx + sdiff a b := sInf {c | a ≤ b ⊔ c} + sdiff_le_iff a b c := + ⟨fun h ↦ (sup_le_sup_left h _).trans' (by simp [minAx.sup_sInf_eq]), fun h ↦ sInf_le h⟩ + top_sdiff _ := rfl + +end Order.Coframe + +namespace CompleteDistribLattice.MinimalAxioms +variable (minAx : MinimalAxioms α) {s : Set α} {a b : α} + +/-- The `CompleteDistribLattice.MinimalAxioms` element corresponding to a complete distrib lattice. +-/ +def of [CompleteDistribLattice α] : MinimalAxioms α := { ‹CompleteDistribLattice α› with } + +/-- Turn minimal axioms for `CompleteDistribLattice` into minimal axioms for `Order.Frame`. -/ +abbrev toFrame : Frame.MinimalAxioms α := minAx.toMinimalAxioms + +/-- Turn minimal axioms for `CompleteDistribLattice` into minimal axioms for `Order.Coframe`. -/ +abbrev toCoframe : Coframe.MinimalAxioms α where __ := minAx + +end MinimalAxioms + +/-- Construct a complete distrib lattice instance using the minimal amount of work needed. + +This sets `a ⇨ b := sSup {c | c ⊓ a ≤ b}`, `aᶜ := a ⇨ ⊥`, `a \ b := sInf {c | a ≤ b ⊔ c}` and +`¬a := ⊤ \ a`. -/ +-- See note [reducible non instances] +abbrev ofMinimalAxioms (minAx : MinimalAxioms α) : CompleteDistribLattice α where + __ := Frame.ofMinimalAxioms minAx.toFrame + __ := Coframe.ofMinimalAxioms minAx.toCoframe + +end CompleteDistribLattice + +namespace CompletelyDistribLattice.MinimalAxioms +variable (minAx : MinimalAxioms α) + +lemma iInf_iSup_eq' (f : ∀ a, κ a → α) : + let _ := minAx.toCompleteLattice + ⨅ i, ⨆ j, f i j = ⨆ g : ∀ i, κ i, ⨅ i, f i (g i) := by + let _ := minAx.toCompleteLattice + refine le_antisymm ?_ le_iInf_iSup + calc _ = ⨅ a : range (range <| f ·), ⨆ b : a.1, b.1 := by simp_rw [iInf_subtype, iInf_range, iSup_subtype, iSup_range] - _ = _ := CompletelyDistribLattice.iInf_iSup_eq _ + _ = _ := minAx.iInf_iSup_eq _ _ ≤ _ := iSup_le fun g => by refine le_trans ?_ <| le_iSup _ fun a => Classical.choose (g ⟨_, a, rfl⟩).2 refine le_iInf fun a => le_trans (iInf_le _ ⟨range (f a), a, rfl⟩) ?_ rw [← Classical.choose_spec (g ⟨_, a, rfl⟩).2] -theorem iSup_iInf_le [CompleteLattice α] {f : ∀ a, κ a → α} : - (⨆ a, ⨅ b, f a b) ≤ ⨅ g : ∀ a, κ a, ⨆ a, f a (g a) := - le_iInf_iSup (α := αᵒᵈ) - -theorem iSup_iInf_eq [CompletelyDistribLattice α] {f : ∀ a, κ a → α} : - (⨆ a, ⨅ b, f a b) = ⨅ g : ∀ a, κ a, ⨆ a, f a (g a) := by +lemma iSup_iInf_eq (f : ∀ i, κ i → α) : + let _ := minAx.toCompleteLattice + ⨆ i, ⨅ j, f i j = ⨅ g : ∀ i, κ i, ⨆ i, f i (g i) := by + let _ := minAx.toCompleteLattice refine le_antisymm iSup_iInf_le ?_ - rw [iInf_iSup_eq] + rw [minAx.iInf_iSup_eq'] refine iSup_le fun g => ?_ have ⟨a, ha⟩ : ∃ a, ∀ b, ∃ f, ∃ h : a = g f, h ▸ b = f (g f) := of_not_not fun h => by push_neg at h @@ -114,31 +259,64 @@ theorem iSup_iInf_eq [CompletelyDistribLattice α] {f : ∀ a, κ a → α} : obtain ⟨h, rfl, rfl⟩ := ha b exact iInf_le _ _ +/-- Turn minimal axioms for `CompletelyDistribLattice` into minimal axioms for +`CompleteDistribLattice`. -/ +abbrev toCompleteDistribLattice : CompleteDistribLattice.MinimalAxioms α where + __ := minAx + inf_sSup_le_iSup_inf a s := by + let _ := minAx.toCompleteLattice + calc + _ = ⨅ i : ULift.{u} Bool, ⨆ j : match i with | .up true => PUnit.{u + 1} | .up false => s, + match i with + | .up true => a + | .up false => j := by simp [le_antisymm_iff, sSup_eq_iSup', iSup_unique, iInf_bool_eq] + _ ≤ _ := by + simp only [minAx.iInf_iSup_eq, iInf_ulift, iInf_bool_eq, iSup_le_iff] + exact fun x ↦ le_biSup _ (x (.up false)).2 + iInf_sup_le_sup_sInf a s := by + let _ := minAx.toCompleteLattice + calc + _ ≤ ⨆ i : ULift.{u} Bool, ⨅ j : match i with | .up true => PUnit.{u + 1} | .up false => s, + match i with + | .up true => a + | .up false => j := by + simp only [minAx.iSup_iInf_eq, iSup_ulift, iSup_bool_eq, le_iInf_iff] + exact fun x ↦ biInf_le _ (x (.up false)).2 + _ = _ := by simp [le_antisymm_iff, sInf_eq_iInf', iInf_unique, iSup_bool_eq] + +/-- The `CompletelyDistribLattice.MinimalAxioms` element corresponding to a frame. -/ +def of [CompletelyDistribLattice α] : MinimalAxioms α := { ‹CompletelyDistribLattice α› with } + +end MinimalAxioms + +/-- Construct a completely distributive lattice instance using the minimal amount of work needed. + +This sets `a ⇨ b := sSup {c | c ⊓ a ≤ b}`, `aᶜ := a ⇨ ⊥`, `a \ b := sInf {c | a ≤ b ⊔ c}` and +`¬a := ⊤ \ a`. -/ +-- See note [reducible non instances] +abbrev ofMinimalAxioms (minAx : MinimalAxioms α) : CompletelyDistribLattice α where + __ := minAx + __ := CompleteDistribLattice.ofMinimalAxioms minAx.toCompleteDistribLattice + +end CompletelyDistribLattice + +theorem iInf_iSup_eq [CompletelyDistribLattice α] {f : ∀ a, κ a → α} : + (⨅ a, ⨆ b, f a b) = ⨆ g : ∀ a, κ a, ⨅ a, f a (g a) := + CompletelyDistribLattice.MinimalAxioms.of.iInf_iSup_eq' _ + +theorem iSup_iInf_eq [CompletelyDistribLattice α] {f : ∀ a, κ a → α} : + (⨆ a, ⨅ b, f a b) = ⨅ g : ∀ a, κ a, ⨆ a, f a (g a) := + CompletelyDistribLattice.MinimalAxioms.of.iSup_iInf_eq _ + instance (priority := 100) CompletelyDistribLattice.toCompleteDistribLattice [CompletelyDistribLattice α] : CompleteDistribLattice α where - iInf_sup_le_sup_sInf a s := calc - _ = ⨅ b : s, ⨆ x : Bool, cond x a b := by simp_rw [iInf_subtype, iSup_bool_eq, cond] - _ = _ := iInf_iSup_eq - _ ≤ _ := iSup_le fun f => by - if h : ∀ i, f i = false then - simp [h, iInf_subtype, ← sInf_eq_iInf] - else - have ⟨i, h⟩ : ∃ i, f i = true := by simpa using h - refine le_trans (iInf_le _ i) ?_ - simp [h] - inf_sSup_le_iSup_inf a s := calc - _ = ⨅ x : Bool, ⨆ y : cond x PUnit s, match x with | true => a | false => y.1 := by - simp_rw [iInf_bool_eq, cond, iSup_const, iSup_subtype, sSup_eq_iSup] - _ = _ := by exact iInf_iSup_eq - _ ≤ _ := by - simp_rw [iInf_bool_eq] - refine iSup_le fun g => le_trans ?_ (le_iSup _ (g false).1) - refine le_trans ?_ (le_iSup _ (g false).2) - rfl + __ := ‹CompletelyDistribLattice α› + __ := CompleteDistribLattice.ofMinimalAxioms MinimalAxioms.of.toCompleteDistribLattice -- See note [lower instance priority] instance (priority := 100) CompleteLinearOrder.toCompletelyDistribLattice [CompleteLinearOrder α] : CompletelyDistribLattice α where + __ := ‹CompleteLinearOrder α› iInf_iSup_eq {α β} g := by let lhs := ⨅ a, ⨆ b, g a b let rhs := ⨆ h : ∀ a, β a, ⨅ a, g a (h a) @@ -170,6 +348,7 @@ variable [Frame α] {s t : Set α} {a b : α} instance OrderDual.instCoframe : Coframe αᵒᵈ where __ := instCompleteLattice + __ := instCoheytingAlgebra iInf_sup_le_sup_sInf := @Frame.inf_sSup_le_iSup_inf α _ theorem inf_sSup_eq : a ⊓ sSup s = ⨆ b ∈ s, a ⊓ b := @@ -243,11 +422,13 @@ instance (priority := 100) Frame.toDistribLattice : DistribLattice α := instance Prod.instFrame [Frame α] [Frame β] : Frame (α × β) where __ := instCompleteLattice + __ := instHeytingAlgebra inf_sSup_le_iSup_inf a s := by simp [Prod.le_def, sSup_eq_iSup, fst_iSup, snd_iSup, fst_iInf, snd_iInf, inf_iSup_eq] instance Pi.instFrame {ι : Type*} {π : ι → Type*} [∀ i, Frame (π i)] : Frame (∀ i, π i) where __ := instCompleteLattice + __ := instHeytingAlgebra inf_sSup_le_iSup_inf a s i := by simp only [sSup_apply, iSup_apply, inf_apply, inf_iSup_eq, ← iSup_subtype'']; rfl @@ -259,6 +440,7 @@ variable [Coframe α] {s t : Set α} {a b : α} instance OrderDual.instFrame : Frame αᵒᵈ where __ := instCompleteLattice + __ := instHeytingAlgebra inf_sSup_le_iSup_inf := @Coframe.iInf_sup_le_sup_sInf α _ theorem sup_sInf_eq : a ⊔ sInf s = ⨅ b ∈ s, a ⊔ b := @@ -306,11 +488,13 @@ instance (priority := 100) Coframe.toDistribLattice : DistribLattice α where instance Prod.instCoframe [Coframe β] : Coframe (α × β) where __ := instCompleteLattice + __ := instCoheytingAlgebra iInf_sup_le_sup_sInf a s := by simp [Prod.le_def, sInf_eq_iInf, fst_iSup, snd_iSup, fst_iInf, snd_iInf, sup_iInf_eq] instance Pi.instCoframe {ι : Type*} {π : ι → Type*} [∀ i, Coframe (π i)] : Coframe (∀ i, π i) where __ := instCompleteLattice + __ := instCoheytingAlgebra iInf_sup_le_sup_sInf a s i := by simp only [sInf_apply, iInf_apply, sup_apply, sup_iInf_eq, ← iInf_subtype'']; rfl @@ -342,16 +526,19 @@ section CompletelyDistribLattice instance OrderDual.instCompletelyDistribLattice [CompletelyDistribLattice α] : CompletelyDistribLattice αᵒᵈ where __ := instFrame + __ := instCoframe iInf_iSup_eq _ := iSup_iInf_eq (α := α) instance Prod.instCompletelyDistribLattice [CompletelyDistribLattice α] [CompletelyDistribLattice β] : CompletelyDistribLattice (α × β) where __ := instFrame + __ := instCoframe iInf_iSup_eq f := by ext <;> simp [fst_iSup, fst_iInf, snd_iSup, snd_iInf, iInf_iSup_eq] instance Pi.instCompletelyDistribLattice {ι : Type*} {π : ι → Type*} [∀ i, CompletelyDistribLattice (π i)] : CompletelyDistribLattice (∀ i, π i) where __ := instFrame + __ := instCoframe iInf_iSup_eq f := by ext i; simp only [iInf_apply, iSup_apply, iInf_iSup_eq] end CompletelyDistribLattice @@ -361,7 +548,18 @@ A complete Boolean algebra is a Boolean algebra that is also a complete distribu It is only completely distributive if it is also atomic. -/ -class CompleteBooleanAlgebra (α) extends BooleanAlgebra α, CompleteDistribLattice α +-- We do not directly extend `CompleteDistribLattice` to avoid having the `hnot` field +class CompleteBooleanAlgebra (α) extends CompleteLattice α, BooleanAlgebra α where + /-- `⊓` distributes over `⨆`. -/ + inf_sSup_le_iSup_inf (a : α) (s : Set α) : a ⊓ sSup s ≤ ⨆ b ∈ s, a ⊓ b + /-- `⊔` distributes over `⨅`. -/ + iInf_sup_le_sup_sInf (a : α) (s : Set α) : ⨅ b ∈ s, a ⊔ b ≤ a ⊔ sInf s + +-- See note [lower instance priority] +instance (priority := 100) CompleteBooleanAlgebra.toCompleteDistribLattice + [CompleteBooleanAlgebra α] : CompleteDistribLattice α where + __ := ‹CompleteBooleanAlgebra α› + __ := BooleanAlgebra.toBiheytingAlgebra instance Prod.instCompleteBooleanAlgebra [CompleteBooleanAlgebra α] [CompleteBooleanAlgebra β] : CompleteBooleanAlgebra (α × β) where @@ -424,10 +622,24 @@ that is also completely distributive. We take iSup_iInf_eq as the definition here, and prove later on that this implies atomicity. -/ -class CompleteAtomicBooleanAlgebra (α : Type u) extends - CompletelyDistribLattice α, CompleteBooleanAlgebra α where - iInf_sup_le_sup_sInf := CompletelyDistribLattice.toCompleteDistribLattice.iInf_sup_le_sup_sInf - inf_sSup_le_iSup_inf := CompletelyDistribLattice.toCompleteDistribLattice.inf_sSup_le_iSup_inf +-- We do not directly extend `CompletelyDistribLattice` to avoid having the `hnot` field +-- We do not directly extend `CompleteBooleanAlgebra` to avoid having the `inf_sSup_le_iSup_inf` and +-- `iInf_sup_le_sup_sInf` fields +class CompleteAtomicBooleanAlgebra (α : Type u) extends CompleteLattice α, BooleanAlgebra α where + protected iInf_iSup_eq {ι : Type u} {κ : ι → Type u} (f : ∀ a, κ a → α) : + (⨅ a, ⨆ b, f a b) = ⨆ g : ∀ a, κ a, ⨅ a, f a (g a) + +-- See note [lower instance priority] +instance (priority := 100) CompleteAtomicBooleanAlgebra.toCompletelyDistribLattice + [CompleteAtomicBooleanAlgebra α] : CompletelyDistribLattice α where + __ := ‹CompleteAtomicBooleanAlgebra α› + __ := BooleanAlgebra.toBiheytingAlgebra + +-- See note [lower instance priority] +instance (priority := 100) CompleteAtomicBooleanAlgebra.toCompleteBooleanAlgebra + [CompleteAtomicBooleanAlgebra α] : CompleteBooleanAlgebra α where + __ := ‹CompleteAtomicBooleanAlgebra α› + __ := CompletelyDistribLattice.toCompleteDistribLattice instance Prod.instCompleteAtomicBooleanAlgebra [CompleteAtomicBooleanAlgebra α] [CompleteAtomicBooleanAlgebra β] : CompleteAtomicBooleanAlgebra (α × β) where @@ -454,51 +666,115 @@ instance Prop.instCompleteBooleanAlgebra : CompleteBooleanAlgebra Prop := inferI section lift -- See note [reducible non-instances] -/-- Pullback an `Order.Frame` along an injection. -/ -protected abbrev Function.Injective.frame [Sup α] [Inf α] [SupSet α] [InfSet α] [Top α] [Bot α] - [Frame β] (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) - (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) - (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) : - Frame α where +/-- Pullback an `Order.Frame.MinimalAxioms` along an injection. -/ +protected abbrev Function.Injective.frameMinimalAxioms [Sup α] [Inf α] [SupSet α] [InfSet α] [Top α] + [Bot α] (minAx : Frame.MinimalAxioms β) (f : α → β) (hf : Injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) : Frame.MinimalAxioms α where __ := hf.completeLattice f map_sup map_inf map_sSup map_sInf map_top map_bot inf_sSup_le_iSup_inf a s := by change f (a ⊓ sSup s) ≤ f _ - rw [← sSup_image, map_inf, map_sSup s, inf_iSup₂_eq] + rw [← sSup_image, map_inf, map_sSup s, minAx.inf_iSup₂_eq] simp_rw [← map_inf] exact ((map_sSup _).trans iSup_image).ge -- See note [reducible non-instances] -/-- Pullback an `Order.Coframe` along an injection. -/ -protected abbrev Function.Injective.coframe [Sup α] [Inf α] [SupSet α] [InfSet α] [Top α] [Bot α] - [Coframe β] (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) - (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) - (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) : - Coframe α where +/-- Pullback an `Order.Coframe.MinimalAxioms` along an injection. -/ +protected abbrev Function.Injective.coframeMinimalAxioms [Sup α] [Inf α] [SupSet α] [InfSet α] + [Top α] [Bot α] (minAx : Coframe.MinimalAxioms β) (f : α → β) (hf : Injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) : Coframe.MinimalAxioms α where __ := hf.completeLattice f map_sup map_inf map_sSup map_sInf map_top map_bot iInf_sup_le_sup_sInf a s := by change f _ ≤ f (a ⊔ sInf s) - rw [← sInf_image, map_sup, map_sInf s, sup_iInf₂_eq] + rw [← sInf_image, map_sup, map_sInf s, minAx.sup_iInf₂_eq] simp_rw [← map_sup] exact ((map_sInf _).trans iInf_image).le +-- See note [reducible non-instances] +/-- Pullback an `Order.Frame` along an injection. -/ +protected abbrev Function.Injective.frame [Sup α] [Inf α] [SupSet α] [InfSet α] [Top α] [Bot α] + [HasCompl α] [HImp α] [Frame β] (f : α → β) (hf : Injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f aᶜ = (f a)ᶜ) + (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) : Frame α where + __ := hf.frameMinimalAxioms .of f map_sup map_inf map_sSup map_sInf map_top map_bot + __ := hf.heytingAlgebra f map_sup map_inf map_top map_bot map_compl map_himp + +-- See note [reducible non-instances] +/-- Pullback an `Order.Coframe` along an injection. -/ +protected abbrev Function.Injective.coframe [Sup α] [Inf α] [SupSet α] [InfSet α] [Top α] [Bot α] + [HNot α] [SDiff α] [Coframe β] (f : α → β) (hf : Injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_hnot : ∀ a, f (¬a) = ¬f a) + (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : Coframe α where + __ := hf.coframeMinimalAxioms .of f map_sup map_inf map_sSup map_sInf map_top map_bot + __ := hf.coheytingAlgebra f map_sup map_inf map_top map_bot map_hnot map_sdiff + +-- See note [reducible non-instances] +/-- Pullback a `CompleteDistribLattice.MinimalAxioms` along an injection. -/ +protected abbrev Function.Injective.completeDistribLatticeMinimalAxioms [Sup α] [Inf α] [SupSet α] + [InfSet α] [Top α] [Bot α] (minAx : CompleteDistribLattice.MinimalAxioms β) (f : α → β) + (hf : Injective f) (map_sup : let _ := minAx.toCompleteLattice + ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : let _ := minAx.toCompleteLattice + ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : let _ := minAx.toCompleteLattice + ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : let _ := minAx.toCompleteLattice + ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : let _ := minAx.toCompleteLattice + f ⊤ = ⊤) (map_bot : let _ := minAx.toCompleteLattice + f ⊥ = ⊥) : + CompleteDistribLattice.MinimalAxioms α where + __ := hf.frameMinimalAxioms minAx.toFrame f map_sup map_inf map_sSup map_sInf map_top map_bot + __ := hf.coframeMinimalAxioms minAx.toCoframe f map_sup map_inf map_sSup map_sInf map_top map_bot + -- See note [reducible non-instances] /-- Pullback a `CompleteDistribLattice` along an injection. -/ protected abbrev Function.Injective.completeDistribLattice [Sup α] [Inf α] [SupSet α] [InfSet α] - [Top α] [Bot α] [CompleteDistribLattice β] (f : α → β) (hf : Function.Injective f) + [Top α] [Bot α] [HasCompl α] [HImp α] [HNot α] [SDiff α] [CompleteDistribLattice β] (f : α → β) + (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) - (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) : CompleteDistribLattice α where - __ := hf.frame f map_sup map_inf map_sSup map_sInf map_top map_bot - __ := hf.coframe f map_sup map_inf map_sSup map_sInf map_top map_bot + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) + (map_hnot : ∀ a, f (¬a) = ¬f a) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + CompleteDistribLattice α where + __ := hf.frame f map_sup map_inf map_sSup map_sInf map_top map_bot map_compl map_himp + __ := hf.coframe f map_sup map_inf map_sSup map_sInf map_top map_bot map_hnot map_sdiff + +-- See note [reducible non-instances] +/-- Pullback a `CompletelyDistribLattice.MinimalAxioms` along an injection. -/ +protected abbrev Function.Injective.completelyDistribLatticeMinimalAxioms [Sup α] [Inf α] [SupSet α] + [InfSet α] [Top α] [Bot α] (minAx : CompletelyDistribLattice.MinimalAxioms β) (f : α → β) + (hf : Injective f) (map_sup : let _ := minAx.toCompleteLattice + ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : let _ := minAx.toCompleteLattice + ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : let _ := minAx.toCompleteLattice + ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : let _ := minAx.toCompleteLattice + ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : let _ := minAx.toCompleteLattice + f ⊤ = ⊤) (map_bot : let _ := minAx.toCompleteLattice + f ⊥ = ⊥) : + CompletelyDistribLattice.MinimalAxioms α where + __ := hf.completeDistribLatticeMinimalAxioms minAx.toCompleteDistribLattice f map_sup map_inf + map_sSup map_sInf map_top map_bot + iInf_iSup_eq g := hf <| by + simp_rw [iInf, map_sInf, iInf_range, iSup, map_sSup, iSup_range, map_sInf, iInf_range, + minAx.iInf_iSup_eq'] -- See note [reducible non-instances] /-- Pullback a `CompletelyDistribLattice` along an injection. -/ protected abbrev Function.Injective.completelyDistribLattice [Sup α] [Inf α] [SupSet α] [InfSet α] - [Top α] [Bot α] [CompletelyDistribLattice β] (f : α → β) (hf : Function.Injective f) + [Top α] [Bot α] [HasCompl α] [HImp α] [HNot α] [SDiff α] [CompletelyDistribLattice β] + (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) - (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) : CompletelyDistribLattice α where + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) + (map_hnot : ∀ a, f (¬a) = ¬f a) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + CompletelyDistribLattice α where __ := hf.completeLattice f map_sup map_inf map_sSup map_sInf map_top map_bot + __ := hf.biheytingAlgebra f map_sup map_inf map_top map_bot map_compl map_hnot map_himp map_sdiff iInf_iSup_eq g := hf <| by simp_rw [iInf, map_sInf, iInf_range, iSup, map_sSup, iSup_range, map_sInf, iInf_range, iInf_iSup_eq] @@ -506,26 +782,40 @@ protected abbrev Function.Injective.completelyDistribLattice [Sup α] [Inf α] [ -- See note [reducible non-instances] /-- Pullback a `CompleteBooleanAlgebra` along an injection. -/ protected abbrev Function.Injective.completeBooleanAlgebra [Sup α] [Inf α] [SupSet α] [InfSet α] - [Top α] [Bot α] [HasCompl α] [SDiff α] [CompleteBooleanAlgebra β] (f : α → β) - (hf : Function.Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) + [Top α] [Bot α] [HasCompl α] [HImp α] [SDiff α] [CompleteBooleanAlgebra β] (f : α → β) + (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) - (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) + (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : CompleteBooleanAlgebra α where - __ := hf.completeDistribLattice f map_sup map_inf map_sSup map_sInf map_top map_bot - __ := hf.booleanAlgebra f map_sup map_inf map_top map_bot map_compl map_sdiff + __ := hf.completeLattice f map_sup map_inf map_sSup map_sInf map_top map_bot + __ := hf.booleanAlgebra f map_sup map_inf map_top map_bot map_compl map_sdiff map_himp + inf_sSup_le_iSup_inf a s := by + change f (a ⊓ sSup s) ≤ f _ + rw [← sSup_image, map_inf, map_sSup s, inf_iSup₂_eq] + simp_rw [← map_inf] + exact ((map_sSup _).trans iSup_image).ge + iInf_sup_le_sup_sInf a s := by + change f _ ≤ f (a ⊔ sInf s) + rw [← sInf_image, map_sup, map_sInf s, sup_iInf₂_eq] + simp_rw [← map_sup] + exact ((map_sInf _).trans iInf_image).le -- See note [reducible non-instances] /-- Pullback a `CompleteAtomicBooleanAlgebra` along an injection. -/ protected abbrev Function.Injective.completeAtomicBooleanAlgebra [Sup α] [Inf α] [SupSet α] - [InfSet α] [Top α] [Bot α] [HasCompl α] [SDiff α] [CompleteAtomicBooleanAlgebra β] (f : α → β) - (hf : Function.Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) - (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) - (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) - (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + [InfSet α] [Top α] [Bot α] [HasCompl α] [HImp α] [HNot α] [SDiff α] + [CompleteAtomicBooleanAlgebra β] (f : α → β) (hf : Injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) + (map_hnot : ∀ a, f (¬a) = ¬f a) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : CompleteAtomicBooleanAlgebra α where - __ := hf.completelyDistribLattice f map_sup map_inf map_sSup map_sInf map_top map_bot - __ := hf.booleanAlgebra f map_sup map_inf map_top map_bot map_compl map_sdiff + __ := hf.completelyDistribLattice f map_sup map_inf map_sSup map_sInf map_top map_bot map_compl + map_himp map_hnot map_sdiff + __ := hf.booleanAlgebra f map_sup map_inf map_top map_bot map_compl map_sdiff map_himp end lift diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 6a6bab4ffb63c..67041b9023667 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -47,6 +47,12 @@ open Function OrderDual Set variable {α β β₂ γ : Type*} {ι ι' : Sort*} {κ : ι → Sort*} {κ' : ι' → Sort*} +@[simp] lemma iSup_ulift {ι : Type*} [SupSet α] (f : ULift ι → α) : + ⨆ i : ULift ι, f i = ⨆ i, f (.up i) := by simp [iSup]; congr with x; simp + +@[simp] lemma iInf_ulift {ι : Type*} [InfSet α] (f : ULift ι → α) : + ⨅ i : ULift ι, f i = ⨅ i, f (.up i) := by simp [iInf]; congr with x; simp + instance OrderDual.supSet (α) [InfSet α] : SupSet αᵒᵈ := ⟨(sInf : Set α → α)⟩ @@ -838,6 +844,12 @@ theorem iSup_const [Nonempty ι] : ⨆ _ : ι, a = a := by rw [iSup, range_const theorem iInf_const [Nonempty ι] : ⨅ _ : ι, a = a := @iSup_const αᵒᵈ _ _ a _ +lemma iSup_unique [Unique ι] (f : ι → α) : ⨆ i, f i = f default := by + simp only [congr_arg f (Unique.eq_default _), iSup_const] + +lemma iInf_unique [Unique ι] (f : ι → α) : ⨅ i, f i = f default := by + simp only [congr_arg f (Unique.eq_default _), iInf_const] + @[simp] theorem iSup_bot : (⨆ _ : ι, ⊥ : α) = ⊥ := bot_unique iSup_const_le diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 0652540d773a3..ce39d0ed5c740 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -995,7 +995,7 @@ theorem cbiSup_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : · simp [hi] · apply sup_le · rcases isEmpty_or_nonempty (Subtype p) with hp|hp - · simp [iSup_of_empty'] + · simp only [iSup_of_empty'] convert le_ciSup B i₀ simp [hi₀] · apply ciSup_le diff --git a/Mathlib/Order/Copy.lean b/Mathlib/Order/Copy.lean index d0aec8e510463..078e00e0941a8 100644 --- a/Mathlib/Order/Copy.lean +++ b/Mathlib/Order/Copy.lean @@ -188,12 +188,16 @@ def Frame.copy (c : Frame α) (le : α → α → Prop) (eq_le : le = (by infer_ (bot : α) (eq_bot : bot = (by infer_instance : Bot α).bot) (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) + (himp : α → α → α) (eq_himp : himp = (by infer_instance : HImp α).himp) + (compl : α → α) (eq_compl : compl = (by infer_instance : HasCompl α).compl) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : Frame α where toCompleteLattice := CompleteLattice.copy (@Frame.toCompleteLattice α c) le eq_le top eq_top bot eq_bot sup eq_sup inf eq_inf sSup eq_sSup sInf eq_sInf inf_sSup_le_iSup_inf := fun a s => by simp [eq_le, eq_sup, eq_inf, eq_sSup, @Order.Frame.inf_sSup_le_iSup_inf α _ a s] + __ := HeytingAlgebra.copy (@Frame.toHeytingAlgebra α c) + le eq_le top eq_top bot eq_bot sup eq_sup inf eq_inf himp eq_himp compl eq_compl /-- A function to create a provable equal copy of a coframe with possibly different definitional equalities. -/ @@ -202,12 +206,16 @@ def Coframe.copy (c : Coframe α) (le : α → α → Prop) (eq_le : le = (by in (bot : α) (eq_bot : bot = (by infer_instance : Bot α).bot) (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) + (sdiff : α → α → α) (eq_sdiff : sdiff = (by infer_instance : SDiff α).sdiff) + (hnot : α → α) (eq_hnot : hnot = (by infer_instance : HNot α).hnot) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : Coframe α where toCompleteLattice := CompleteLattice.copy (@Coframe.toCompleteLattice α c) le eq_le top eq_top bot eq_bot sup eq_sup inf eq_inf sSup eq_sSup sInf eq_sInf iInf_sup_le_sup_sInf := fun a s => by simp [eq_le, eq_sup, eq_inf, eq_sInf, @Order.Coframe.iInf_sup_le_sup_sInf α _ a s] + __ := CoheytingAlgebra.copy (@Coframe.toCoheytingAlgebra α c) + le eq_le top eq_top bot eq_bot sup eq_sup inf eq_inf sdiff eq_sdiff hnot eq_hnot /-- A function to create a provable equal copy of a complete distributive lattice with possibly different definitional equalities. -/ @@ -217,13 +225,17 @@ def CompleteDistribLattice.copy (c : CompleteDistribLattice α) (bot : α) (eq_bot : bot = (by infer_instance : Bot α).bot) (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) + (sdiff : α → α → α) (eq_sdiff : sdiff = (by infer_instance : SDiff α).sdiff) + (hnot : α → α) (eq_hnot : hnot = (by infer_instance : HNot α).hnot) + (himp : α → α → α) (eq_himp : himp = (by infer_instance : HImp α).himp) + (compl : α → α) (eq_compl : compl = (by infer_instance : HasCompl α).compl) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : CompleteDistribLattice α where - toFrame := Frame.copy (@CompleteDistribLattice.toFrame α c) - le eq_le top eq_top bot eq_bot sup eq_sup inf eq_inf sSup eq_sSup sInf eq_sInf - __ := Coframe.copy (@CompleteDistribLattice.toCoframe α c) - le eq_le top eq_top bot eq_bot sup eq_sup inf eq_inf sSup eq_sSup sInf eq_sInf + toFrame := Frame.copy (@CompleteDistribLattice.toFrame α c) le eq_le top eq_top bot eq_bot sup + eq_sup inf eq_inf himp eq_himp compl eq_compl sSup eq_sSup sInf eq_sInf + __ := Coframe.copy (@CompleteDistribLattice.toCoframe α c) le eq_le top eq_top bot eq_bot sup + eq_sup inf eq_inf sdiff eq_sdiff hnot eq_hnot sSup eq_sSup sInf eq_sInf /-- A function to create a provable equal copy of a conditionally complete lattice with possibly different definitional equalities. -/ diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 6e59bcc0db263..6f4bc221c9556 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -779,8 +779,8 @@ instance : DistribLattice (Filter α) := ⟨t₁, x.sets_of_superset hs inter_subset_left, ht₁, t₂, x.sets_of_superset hs inter_subset_right, ht₂, rfl⟩ } --- The dual version does not hold! `Filter α` is not a `CompleteDistribLattice`. -/ -instance : Coframe (Filter α) := +/-- The dual version does not hold! `Filter α` is not a `CompleteDistribLattice`. -/ +def coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := { Filter.instCompleteLatticeFilter with iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by rw [iInf_subtype'] @@ -793,6 +793,8 @@ instance : Coframe (Filter α) := rw [Finset.inf_insert, sup_inf_left] exact le_inf (iInf_le _ _) ih } +instance instCoframe : Coframe (Filter α) := .ofMinimalAxioms coframeMinimalAxioms + theorem mem_iInf_finset {s : Finset α} {f : α → Filter β} {t : Set β} : (t ∈ ⨅ a ∈ s, f a) ↔ ∃ p : α → Set β, (∀ a ∈ s, p a ∈ f a) ∧ t = ⋂ a ∈ s, p a := by simp only [← Finset.set_biInter_coe, biInter_eq_iInter, iInf_subtype'] @@ -1943,12 +1945,14 @@ The variables in the following lemmas are used as in this diagram: -/ -variable {φ : α → β} {θ : α → γ} {ψ : β → δ} {ρ : γ → δ} (H : ψ ∘ φ = ρ ∘ θ) +variable {φ : α → β} {θ : α → γ} {ψ : β → δ} {ρ : γ → δ} -theorem map_comm (F : Filter α) : map ψ (map φ F) = map ρ (map θ F) := by +theorem map_comm (H : ψ ∘ φ = ρ ∘ θ) (F : Filter α) : + map ψ (map φ F) = map ρ (map θ F) := by rw [Filter.map_map, H, ← Filter.map_map] -theorem comap_comm (G : Filter δ) : comap φ (comap ψ G) = comap θ (comap ρ G) := by +theorem comap_comm (H : ψ ∘ φ = ρ ∘ θ) (G : Filter δ) : + comap φ (comap ψ G) = comap θ (comap ρ G) := by rw [Filter.comap_comap, H, ← Filter.comap_comap] end comm diff --git a/Mathlib/Order/Filter/CardinalInter.lean b/Mathlib/Order/Filter/CardinalInter.lean index fb740679e7d8e..7462de92c6dd8 100644 --- a/Mathlib/Order/Filter/CardinalInter.lean +++ b/Mathlib/Order/Filter/CardinalInter.lean @@ -302,7 +302,7 @@ theorem mem_cardinaleGenerate_iff {s : Set α} {hreg : c.IsRegular} : constructor <;> intro h · induction' h with s hs s t _ st ih S Sct _ ih · refine ⟨{s}, singleton_subset_iff.mpr hs, ?_⟩ - norm_num; exact ⟨IsRegular.nat_lt hreg 1, subset_rfl⟩ + simpa [subset_refl] using IsRegular.nat_lt hreg 1 · exact ⟨∅, ⟨empty_subset g, mk_eq_zero (∅ : Set <| Set α) ▸ IsRegular.nat_lt hreg 0, by simp⟩⟩ · exact Exists.imp (by tauto) ih choose T Tg Tct hT using ih diff --git a/Mathlib/Order/Filter/Cocardinal.lean b/Mathlib/Order/Filter/Cocardinal.lean index 75f9f8fd58f46..ede785eb18061 100644 --- a/Mathlib/Order/Filter/Cocardinal.lean +++ b/Mathlib/Order/Filter/Cocardinal.lean @@ -95,8 +95,7 @@ theorem _root_.Finset.eventually_cocardinal_nmem (s : Finset α) : eventually_cocardinal_nmem_of_card_lt <| lt_of_lt_of_le (finset_card_lt_aleph0 s) (hreg.aleph0_le) theorem eventually_cocardinal_ne (x : α) : ∀ᶠ a in cocardinal α hreg, a ≠ x := by - simp [Set.finite_singleton x] - exact hreg.nat_lt 1 + simpa [Set.finite_singleton x] using hreg.nat_lt 1 /-- The filter defined by all sets that have countable complements. -/ abbrev cocountable : Filter α := cocardinal α Cardinal.isRegular_aleph_one diff --git a/Mathlib/Order/Filter/FilterProduct.lean b/Mathlib/Order/Filter/FilterProduct.lean index 277d830e64eab..95fd3fee7f210 100644 --- a/Mathlib/Order/Filter/FilterProduct.lean +++ b/Mathlib/Order/Filter/FilterProduct.lean @@ -46,11 +46,13 @@ instance instDivisionSemiring [DivisionSemiring β] : DivisionSemiring β* where toSemiring := instSemiring __ := instGroupWithZero nnqsmul := _ + nnqsmul_def := fun q a => rfl instance instDivisionRing [DivisionRing β] : DivisionRing β* where __ := instRing __ := instDivisionSemiring qsmul := _ + qsmul_def := fun q a => rfl instance instSemifield [Semifield β] : Semifield β* where __ := instCommSemiring diff --git a/Mathlib/Order/Filter/Partial.lean b/Mathlib/Order/Filter/Partial.lean index ef40efac659d5..05453a60c9e88 100644 --- a/Mathlib/Order/Filter/Partial.lean +++ b/Mathlib/Order/Filter/Partial.lean @@ -104,7 +104,8 @@ theorem rcomap_sets (r : Rel α β) (f : Filter β) : theorem rcomap_rcomap (r : Rel α β) (s : Rel β γ) (l : Filter γ) : rcomap r (rcomap s l) = rcomap (r.comp s) l := filter_eq <| by - ext t; simp [rcomap_sets, Rel.image, Rel.core_comp]; constructor + ext t; simp only [rcomap_sets, Rel.image, Filter.mem_sets, Set.mem_setOf_eq, Rel.core_comp] + constructor · rintro ⟨u, ⟨v, vsets, hv⟩, h⟩ exact ⟨v, vsets, Set.Subset.trans (Rel.core_mono _ hv) h⟩ rintro ⟨t, tsets, ht⟩ @@ -118,9 +119,9 @@ theorem rtendsto_iff_le_rcomap (r : Rel α β) (l₁ : Filter α) (l₂ : Filter RTendsto r l₁ l₂ ↔ l₁ ≤ l₂.rcomap r := by rw [rtendsto_def] simp_rw [← l₂.mem_sets] - simp [Filter.le_def, rcomap, Rel.mem_image]; constructor - · exact fun h s t tl₂ => mem_of_superset (h t tl₂) - · exact fun h t tl₂ => h _ t tl₂ Set.Subset.rfl + constructor + · simpa [Filter.le_def, rcomap, Rel.mem_image] using fun h s t tl₂ => mem_of_superset (h t tl₂) + · simpa [Filter.le_def, rcomap, Rel.mem_image] using fun h t tl₂ => h _ t tl₂ Set.Subset.rfl -- Interestingly, there does not seem to be a way to express this relation using a forward map. -- Given a filter `f` on `α`, we want a filter `f'` on `β` such that `r.preimage s ∈ f` if @@ -167,9 +168,9 @@ def RTendsto' (r : Rel α β) (l₁ : Filter α) (l₂ : Filter β) := theorem rtendsto'_def (r : Rel α β) (l₁ : Filter α) (l₂ : Filter β) : RTendsto' r l₁ l₂ ↔ ∀ s ∈ l₂, r.preimage s ∈ l₁ := by - unfold RTendsto' rcomap'; simp [le_def, Rel.mem_image]; constructor - · exact fun h s hs => h _ _ hs Set.Subset.rfl - · exact fun h s t ht => mem_of_superset (h t ht) + unfold RTendsto' rcomap'; constructor + · simpa [le_def, Rel.mem_image] using fun h s hs => h _ _ hs Set.Subset.rfl + · simpa [le_def, Rel.mem_image] using fun h s t ht => mem_of_superset (h t ht) theorem tendsto_iff_rtendsto (l₁ : Filter α) (l₂ : Filter β) (f : α → β) : Tendsto f l₁ l₂ ↔ RTendsto (Function.graph f) l₁ l₂ := by diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 5dced7896fd0c..951efec678d5d 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -557,11 +557,11 @@ protected theorem strictMono : StrictMono f := fun _ _ => f.lt_iff_lt.2 protected theorem acc (a : α) : Acc (· < ·) (f a) → Acc (· < ·) a := f.ltEmbedding.acc a -protected theorem wellFounded : +protected theorem wellFounded (f : α ↪o β) : WellFounded ((· < ·) : β → β → Prop) → WellFounded ((· < ·) : α → α → Prop) := f.ltEmbedding.wellFounded -protected theorem isWellOrder [IsWellOrder β (· < ·)] : IsWellOrder α (· < ·) := +protected theorem isWellOrder [IsWellOrder β (· < ·)] (f : α ↪o β) : IsWellOrder α (· < ·) := f.ltEmbedding.isWellOrder /-- An order embedding is also an order embedding between dual orders. -/ @@ -569,13 +569,13 @@ protected def dual : αᵒᵈ ↪o βᵒᵈ := ⟨f.toEmbedding, f.map_rel_iff⟩ /-- A preorder which embeds into a well-founded preorder is itself well-founded. -/ -protected theorem wellFoundedLT [WellFoundedLT β] : WellFoundedLT α where +protected theorem wellFoundedLT [WellFoundedLT β] (f : α ↪o β) : WellFoundedLT α where wf := f.wellFounded IsWellFounded.wf /-- A preorder which embeds into a preorder in which `(· > ·)` is well-founded also has `(· > ·)` well-founded. -/ -protected theorem wellFoundedGT [WellFoundedGT β] : WellFoundedGT α := - @OrderEmbedding.wellFoundedLT αᵒᵈ _ _ _ f.dual _ +protected theorem wellFoundedGT [WellFoundedGT β] (f : α ↪o β) : WellFoundedGT α := + @OrderEmbedding.wellFoundedLT αᵒᵈ _ _ _ _ f.dual /-- A version of `WithBot.map` for order embeddings. -/ @[simps (config := .asFn)] @@ -1225,13 +1225,14 @@ theorem OrderIso.isCompl {x y : α} (h : IsCompl x y) : IsCompl (f x) (f y) := theorem OrderIso.isCompl_iff {x y : α} : IsCompl x y ↔ IsCompl (f x) (f y) := ⟨f.isCompl, fun h => f.symm_apply_apply x ▸ f.symm_apply_apply y ▸ f.symm.isCompl h⟩ -theorem OrderIso.complementedLattice [ComplementedLattice α] : ComplementedLattice β := +theorem OrderIso.complementedLattice [ComplementedLattice α] (f : α ≃o β) : ComplementedLattice β := ⟨fun x => by obtain ⟨y, hy⟩ := exists_isCompl (f.symm x) rw [← f.symm_apply_apply y] at hy exact ⟨f y, f.symm.isCompl_iff.2 hy⟩⟩ -theorem OrderIso.complementedLattice_iff : ComplementedLattice α ↔ ComplementedLattice β := +theorem OrderIso.complementedLattice_iff (f : α ≃o β) : + ComplementedLattice α ↔ ComplementedLattice β := ⟨by intro; exact f.complementedLattice, by intro; exact f.symm.complementedLattice⟩ diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index f108ab75c87f8..2d05cf8e5f8df 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -1031,27 +1031,29 @@ theorem subtype_Ioc_eq : Ioc a b = (Ioc (a : α) b).subtype p := theorem subtype_Ioo_eq : Ioo a b = (Ioo (a : α) b).subtype p := rfl -variable (hp : ∀ ⦃a b x⦄, a ≤ x → x ≤ b → p a → p b → p x) - -theorem map_subtype_embedding_Icc : (Icc a b).map (Embedding.subtype p) = (Icc a b : Finset α) := by +theorem map_subtype_embedding_Icc (hp : ∀ ⦃a b x⦄, a ≤ x → x ≤ b → p a → p b → p x): + (Icc a b).map (Embedding.subtype p) = (Icc a b : Finset α) := by rw [subtype_Icc_eq] refine Finset.subtype_map_of_mem fun x hx => ?_ rw [mem_Icc] at hx exact hp hx.1 hx.2 a.prop b.prop -theorem map_subtype_embedding_Ico : (Ico a b).map (Embedding.subtype p) = (Ico a b : Finset α) := by +theorem map_subtype_embedding_Ico (hp : ∀ ⦃a b x⦄, a ≤ x → x ≤ b → p a → p b → p x): + (Ico a b).map (Embedding.subtype p) = (Ico a b : Finset α) := by rw [subtype_Ico_eq] refine Finset.subtype_map_of_mem fun x hx => ?_ rw [mem_Ico] at hx exact hp hx.1 hx.2.le a.prop b.prop -theorem map_subtype_embedding_Ioc : (Ioc a b).map (Embedding.subtype p) = (Ioc a b : Finset α) := by +theorem map_subtype_embedding_Ioc (hp : ∀ ⦃a b x⦄, a ≤ x → x ≤ b → p a → p b → p x): + (Ioc a b).map (Embedding.subtype p) = (Ioc a b : Finset α) := by rw [subtype_Ioc_eq] refine Finset.subtype_map_of_mem fun x hx => ?_ rw [mem_Ioc] at hx exact hp hx.1.le hx.2 a.prop b.prop -theorem map_subtype_embedding_Ioo : (Ioo a b).map (Embedding.subtype p) = (Ioo a b : Finset α) := by +theorem map_subtype_embedding_Ioo (hp : ∀ ⦃a b x⦄, a ≤ x → x ≤ b → p a → p b → p x): + (Ioo a b).map (Embedding.subtype p) = (Ioo a b : Finset α) := by rw [subtype_Ioo_eq] refine Finset.subtype_map_of_mem fun x hx => ?_ rw [mem_Ioo] at hx diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index a080675266e12..c4559a1ed5f74 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -297,8 +297,8 @@ lemma Nat.strong_decreasing_induction (base : ∃ n, ∀ m > n, P m) (step : ∀ theorem Nat.decreasing_induction_of_infinite (hP : { x | P x }.Infinite) (n : ℕ) : P n := Nat.decreasing_induction_of_not_bddAbove h (mt BddAbove.finite hP) n -theorem Nat.cauchy_induction' (seed : ℕ) (hs : P seed) (hi : ∀ x, seed ≤ x → P x → ∃ y, x < y ∧ P y) - (n : ℕ) : P n := by +theorem Nat.cauchy_induction' (seed : ℕ) (h : ∀ n, P (n + 1) → P n) (hs : P seed) + (hi : ∀ x, seed ≤ x → P x → ∃ y, x < y ∧ P y) (n : ℕ) : P n := by apply Nat.decreasing_induction_of_infinite h fun hf => _ intro hf obtain ⟨m, hP, hm⟩ := hf.exists_maximal_wrt id _ ⟨seed, hs⟩ @@ -309,8 +309,8 @@ theorem Nat.cauchy_induction (seed : ℕ) (hs : P seed) (f : ℕ → ℕ) (hf : ∀ x, seed ≤ x → P x → x < f x ∧ P (f x)) (n : ℕ) : P n := seed.cauchy_induction' h hs (fun x hl hx => ⟨f x, hf x hl hx⟩) n -theorem Nat.cauchy_induction_mul (k seed : ℕ) (hk : 1 < k) (hs : P seed.succ) - (hm : ∀ x, seed < x → P x → P (k * x)) (n : ℕ) : P n := by +theorem Nat.cauchy_induction_mul (h : ∀ (n : ℕ), P (n + 1) → P n) (k seed : ℕ) (hk : 1 < k) + (hs : P seed.succ) (hm : ∀ x, seed < x → P x → P (k * x)) (n : ℕ) : P n := by apply Nat.cauchy_induction h _ hs (k * ·) fun x hl hP => ⟨_, hm x hl hP⟩ intro _ hl _ convert (Nat.mul_lt_mul_right <| seed.succ_pos.trans_le hl).2 hk diff --git a/Mathlib/Order/Interval/Set/Pi.lean b/Mathlib/Order/Interval/Set/Pi.lean index 94b39d8a71331..524744496cce6 100644 --- a/Mathlib/Order/Interval/Set/Pi.lean +++ b/Mathlib/Order/Interval/Set/Pi.lean @@ -51,22 +51,20 @@ theorem piecewise_mem_Icc' {s : Set ι} [∀ j, Decidable (j ∈ s)] {f₁ f₂ section Nonempty -variable [Nonempty ι] - -theorem pi_univ_Ioi_subset : (pi univ fun i ↦ Ioi (x i)) ⊆ Ioi x := fun z hz ↦ +theorem pi_univ_Ioi_subset [Nonempty ι]: (pi univ fun i ↦ Ioi (x i)) ⊆ Ioi x := fun z hz ↦ ⟨fun i ↦ le_of_lt <| hz i trivial, fun h ↦ - (Nonempty.elim ‹Nonempty ι›) fun i ↦ not_lt_of_le (h i) (hz i trivial)⟩ + (‹Nonempty ι›.elim) fun i ↦ not_lt_of_le (h i) (hz i trivial)⟩ -theorem pi_univ_Iio_subset : (pi univ fun i ↦ Iio (x i)) ⊆ Iio x := - @pi_univ_Ioi_subset ι (fun i ↦ (α i)ᵒᵈ) _ x _ +theorem pi_univ_Iio_subset [Nonempty ι]: (pi univ fun i ↦ Iio (x i)) ⊆ Iio x := + pi_univ_Ioi_subset (α := fun i ↦ (α i)ᵒᵈ) x -theorem pi_univ_Ioo_subset : (pi univ fun i ↦ Ioo (x i) (y i)) ⊆ Ioo x y := fun _ hx ↦ +theorem pi_univ_Ioo_subset [Nonempty ι]: (pi univ fun i ↦ Ioo (x i) (y i)) ⊆ Ioo x y := fun _ hx ↦ ⟨(pi_univ_Ioi_subset _) fun i hi ↦ (hx i hi).1, (pi_univ_Iio_subset _) fun i hi ↦ (hx i hi).2⟩ -theorem pi_univ_Ioc_subset : (pi univ fun i ↦ Ioc (x i) (y i)) ⊆ Ioc x y := fun _ hx ↦ +theorem pi_univ_Ioc_subset [Nonempty ι]: (pi univ fun i ↦ Ioc (x i) (y i)) ⊆ Ioc x y := fun _ hx ↦ ⟨(pi_univ_Ioi_subset _) fun i hi ↦ (hx i hi).1, fun i ↦ (hx i trivial).2⟩ -theorem pi_univ_Ico_subset : (pi univ fun i ↦ Ico (x i) (y i)) ⊆ Ico x y := fun _ hx ↦ +theorem pi_univ_Ico_subset [Nonempty ι]: (pi univ fun i ↦ Ico (x i) (y i)) ⊆ Ico x y := fun _ hx ↦ ⟨fun i ↦ (hx i trivial).1, (pi_univ_Iio_subset _) fun i hi ↦ (hx i hi).2⟩ end Nonempty @@ -290,10 +288,9 @@ theorem Icc_diff_pi_univ_Ioo_subset (x y x' y' : ∀ i, α i) : (Icc x y \ pi univ fun i ↦ Ioo (x' i) (y' i)) ⊆ (⋃ i : ι, Icc x (update y i (x' i))) ∪ ⋃ i : ι, Icc (update x i (y' i)) y := by rintro a ⟨⟨hxa, hay⟩, ha'⟩ - simp at ha' - simp only [le_update_iff, ne_eq, not_and, not_forall, not_le, exists_prop, gt_iff_lt, - update_le_iff, mem_union, mem_iUnion, mem_Icc, hxa, hay _, implies_true, and_true, true_and, - hxa _, hay, ← exists_or] + simp only [mem_pi, mem_univ, mem_Ioo, true_implies, not_forall] at ha' + simp only [le_update_iff, update_le_iff, mem_union, mem_iUnion, mem_Icc, + hxa, hay _, hxa _, hay, ← exists_or] rcases ha' with ⟨w, hw⟩ apply Exists.intro w cases lt_or_le (x' w) (a w) <;> simp_all diff --git a/Mathlib/Order/Irreducible.lean b/Mathlib/Order/Irreducible.lean index 4886caa4d91dc..ecc3744d96ae4 100644 --- a/Mathlib/Order/Irreducible.lean +++ b/Mathlib/Order/Irreducible.lean @@ -185,7 +185,7 @@ variable [WellFoundedGT α] elements. This is the order-theoretic analogue of prime factorisation. -/ theorem exists_infIrred_decomposition (a : α) : ∃ s : Finset α, s.inf id = a ∧ ∀ ⦃b⦄, b ∈ s → InfIrred b := - @exists_supIrred_decomposition αᵒᵈ _ _ _ _ + exists_supIrred_decomposition (α := αᵒᵈ) _ end SemilatticeInf diff --git a/Mathlib/Order/Minimal.lean b/Mathlib/Order/Minimal.lean index 7998bdcd8df72..ae33ce46f2024 100644 --- a/Mathlib/Order/Minimal.lean +++ b/Mathlib/Order/Minimal.lean @@ -86,9 +86,9 @@ theorem mem_maximals_setOf_iff {P : α → Prop} : x ∈ maximals r (setOf P) ↔ P x ∧ ∀ ⦃y⦄, P y → r x y → x = y := mem_maximals_iff -theorem mem_minimals_iff : x ∈ minimals r s ↔ x ∈ s ∧ ∀ ⦃y⦄, y ∈ s → r y x → x = y := by - haveI := IsAntisymm.swap r - exact mem_maximals_iff +theorem mem_minimals_iff : x ∈ minimals r s ↔ x ∈ s ∧ ∀ ⦃y⦄, y ∈ s → r y x → x = y := + have := IsAntisymm.swap r + mem_maximals_iff theorem mem_minimals_setOf_iff {P : α → Prop} : x ∈ minimals r (setOf P) ↔ P x ∧ ∀ ⦃y⦄, P y → r y x → x = y := @@ -118,7 +118,9 @@ theorem minimals_eq_minimals_of_subset_of_forall [IsTrans α r] (hts : t ⊆ s) theorem maximals_eq_maximals_of_subset_of_forall [IsTrans α r] (hts : t ⊆ s) (h : ∀ x ∈ s, ∃ y ∈ t, r x y) : maximals r s = maximals r t := - @minimals_eq_minimals_of_subset_of_forall _ _ _ _ (IsAntisymm.swap r) (IsTrans.swap r) hts h + have := IsTrans.swap r + have := IsAntisymm.swap r + minimals_eq_minimals_of_subset_of_forall hts h variable (r s) @@ -126,7 +128,7 @@ theorem maximals_antichain : IsAntichain r (maximals r s) := fun _a ha _b hb hab hab <| eq_of_mem_maximals ha hb.1 h theorem minimals_antichain : IsAntichain r (minimals r s) := - haveI := IsAntisymm.swap r + have := IsAntisymm.swap r (maximals_antichain _ _).swap end IsAntisymm @@ -280,6 +282,7 @@ section Image variable {f : α → β} {r : α → α → Prop} {s : β → β → Prop} section + variable {x : Set α} (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) {a : α} theorem map_mem_minimals (ha : a ∈ minimals r x) : f a ∈ minimals s (f '' x) := diff --git a/Mathlib/Order/ModularLattice.lean b/Mathlib/Order/ModularLattice.lean index 5c1a6aa279325..0f3392396dadb 100644 --- a/Mathlib/Order/ModularLattice.lean +++ b/Mathlib/Order/ModularLattice.lean @@ -216,7 +216,7 @@ theorem sup_lt_sup_of_lt_of_inf_le_inf (hxy : x < y) (hinf : y ⊓ z ≤ x ⊓ z ne_of_lt hxy <| eq_of_le_of_inf_le_of_sup_le (le_of_lt hxy) hinf (le_of_eq hsup.symm) theorem inf_lt_inf_of_lt_of_sup_le_sup (hxy : x < y) (hinf : y ⊔ z ≤ x ⊔ z) : x ⊓ z < y ⊓ z := - @sup_lt_sup_of_lt_of_inf_le_inf αᵒᵈ _ _ _ _ _ hxy hinf + sup_lt_sup_of_lt_of_inf_le_inf (α := αᵒᵈ) hxy hinf /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Artinian, then `M` is Artinian. -/ @@ -243,7 +243,8 @@ theorem wellFounded_gt_exact_sequence {β γ : Type*} [Preorder β] [PartialOrde (f₁ : β → α) (f₂ : α → β) (g₁ : γ → α) (g₂ : α → γ) (gci : GaloisCoinsertion f₁ f₂) (gi : GaloisInsertion g₂ g₁) (hf : ∀ a, f₁ (f₂ a) = a ⊓ K) (hg : ∀ a, g₁ (g₂ a) = a ⊔ K) : WellFounded ((· > ·) : α → α → Prop) := - @wellFounded_lt_exact_sequence αᵒᵈ _ _ γᵒᵈ βᵒᵈ _ _ h₂ h₁ K g₁ g₂ f₁ f₂ gi.dual gci.dual hg hf + wellFounded_lt_exact_sequence (α := αᵒᵈ) (β := γᵒᵈ) (γ := βᵒᵈ) + h₂ h₁ K g₁ g₂ f₁ f₂ gi.dual gci.dual hg hf /-- The diamond isomorphism between the intervals `[a ⊓ b, a]` and `[b, a ⊔ b]` -/ @[simps] diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 8456867e1ced9..2652ed1dda781 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -122,7 +122,7 @@ theorem map_comp : (c.map f).map g = c.map (g.comp f) := @[mono] theorem map_le_map {g : α →o β} (h : f ≤ g) : c.map f ≤ c.map g := - fun i => by simp [mem_map_iff]; exists i; apply h + fun i => by simp only [map_coe, Function.comp_apply]; exists i; apply h /-- `OmegaCompletePartialOrder.Chain.zip` pairs up the elements of two chains that have the same index. -/ @@ -571,11 +571,11 @@ def Simps.apply (h : α →𝒄 β) : α → β := initialize_simps_projections ContinuousHom (toFun → apply) -theorem congr_fun {f g : α →𝒄 β} (h : f = g) (x : α) : f x = g x := +protected theorem congr_fun {f g : α →𝒄 β} (h : f = g) (x : α) : f x = g x := DFunLike.congr_fun h x -theorem congr_arg (f : α →𝒄 β) {x y : α} (h : x = y) : f x = f y := - _root_.congr_arg f h +protected theorem congr_arg (f : α →𝒄 β) {x y : α} (h : x = y) : f x = f y := + congr_arg f h protected theorem monotone (f : α →𝒄 β) : Monotone f := f.monotone' diff --git a/Mathlib/Order/Rel/GaloisConnection.lean b/Mathlib/Order/Rel/GaloisConnection.lean new file mode 100644 index 0000000000000..0edf015bf7466 --- /dev/null +++ b/Mathlib/Order/Rel/GaloisConnection.lean @@ -0,0 +1,101 @@ +/- +Copyright (c) 2024 Lagrange Mathematics and Computing Research Center. All rights reserved. +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 + +In this file, we show that an arbitrary relation `R` between a pair of types `α` and `β` defines +a pair `toDual ∘ R.leftDual` and `R.rightDual ∘ ofDual` of adjoint order-preserving maps between the +corresponding posets `Set α` and `(Set β)ᵒᵈ`. +We define `R.leftFixedPoints` (resp. `R.rightFixedPoints`) as the set of fixed points `J` +(resp. `I`) of `Set α` (resp. `Set β`) such that `rightDual (leftDual J) = J` +(resp. `leftDual (rightDual I) = I`). + +## Main Results + +⋆ `Rel.gc_leftDual_rightDual`: we prove that the maps `toDual ∘ R.leftDual` and + `R.rightDual ∘ ofDual` form a Galois connection. +⋆ `Rel.equivFixedPoints`: we prove that the maps `R.leftDual` and `R.rightDual` induce inverse + bijections between the sets of fixed points. + +## References + +⋆ Engendrement de topologies, démontrabilité et opérations sur les sous-topos, Olivia Caramello and + Laurent Lafforgue (in preparation) + +## Tags + +relation, Galois connection, induced bijection, fixed points +-/ + +variable {α β : Type*} (R : Rel α β) + +namespace Rel + +/-! ### Pairs of adjoint maps defined by relations -/ + +open OrderDual + +/-- `leftDual` maps any set `J` of elements of type `α` to the set `{b : β | ∀ a ∈ J, R a b}` of +elements `b` of type `β` such that `R a b` for every element `a` of `J`. -/ +def leftDual (J : Set α) : Set β := {b : β | ∀ ⦃a⦄, a ∈ J → R a b} + +/-- `rightDual` maps any set `I` of elements of type `β` to the set `{a : α | ∀ b ∈ I, R a b}` +of elements `a` of type `α` such that `R a b` for every element `b` of `I`. -/ +def rightDual (I : Set β) : Set α := {a : α | ∀ ⦃b⦄, b ∈ I → R a b} + +/-- The pair of functions `toDual ∘ leftDual` and `rightDual ∘ ofDual` forms a Galois connection. -/ +theorem gc_leftDual_rightDual : GaloisConnection (toDual ∘ R.leftDual) (R.rightDual ∘ ofDual) := + fun J I ↦ ⟨fun h a ha b hb ↦ h (by simpa) ha, fun h b hb a ha ↦ h (by simpa) hb⟩ + +/-! ### Induced equivalences between fixed points -/ + +/-- `leftFixedPoints` is the set of elements `J : Set α` satisfying `rightDual (leftDual J) = J`. -/ +def leftFixedPoints := {J : Set α | R.rightDual (R.leftDual J) = J} + +/-- `rightFixedPoints` is the set of elements `I : Set β` satisfying `leftDual (rightDual I) = I`. +-/ +def rightFixedPoints := {I : Set β | R.leftDual (R.rightDual I) = I} + +open GaloisConnection + +/-- `leftDual` maps every element `J` to `rightFixedPoints`. -/ +theorem leftDual_mem_rightFixedPoint (J : Set α) : R.leftDual J ∈ R.rightFixedPoints := by + apply le_antisymm + · apply R.gc_leftDual_rightDual.monotone_l; exact R.gc_leftDual_rightDual.le_u_l J + · exact R.gc_leftDual_rightDual.l_u_le (R.leftDual J) + +/-- `rightDual` maps every element `I` to `leftFixedPoints`. -/ +theorem rightDual_mem_leftFixedPoint (I : Set β) : R.rightDual I ∈ R.leftFixedPoints := by + apply le_antisymm + · apply R.gc_leftDual_rightDual.monotone_u; exact R.gc_leftDual_rightDual.l_u_le I + · exact R.gc_leftDual_rightDual.le_u_l (R.rightDual I) + +/-- The maps `leftDual` and `rightDual` induce inverse bijections between the sets of fixed points. +-/ +def equivFixedPoints : R.leftFixedPoints ≃ R.rightFixedPoints where + toFun := fun ⟨J, _⟩ => ⟨R.leftDual J, R.leftDual_mem_rightFixedPoint J⟩ + invFun := fun ⟨I, _⟩ => ⟨R.rightDual I, R.rightDual_mem_leftFixedPoint I⟩ + left_inv J := by cases' J with J hJ; rw [Subtype.mk.injEq, hJ] + right_inv I := by cases' I with I hI; rw [Subtype.mk.injEq, hI] + +theorem rightDual_leftDual_le_of_le {J J' : Set α} (h : J' ∈ R.leftFixedPoints) (h₁ : J ≤ J') : + R.rightDual (R.leftDual J) ≤ J' := by + rw [← h] + apply R.gc_leftDual_rightDual.monotone_u + apply R.gc_leftDual_rightDual.monotone_l + exact h₁ + +theorem leftDual_rightDual_le_of_le {I I' : Set β} (h : I' ∈ R.rightFixedPoints) (h₁ : I ≤ I') : + R.leftDual (R.rightDual I) ≤ I' := by + rw [← h] + apply R.gc_leftDual_rightDual.monotone_l + apply R.gc_leftDual_rightDual.monotone_u + exact h₁ + +end Rel diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 87bf45d24578a..b2fd3766945b1 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -529,7 +529,7 @@ lemma subset_of_eq_of_subset (hab : a = b) (hbc : b ⊆ c) : a ⊆ c := by rwa [ lemma subset_of_subset_of_eq (hab : a ⊆ b) (hbc : b = c) : a ⊆ c := by rwa [← hbc] -@[refl] +@[refl, simp] lemma subset_refl [IsRefl α (· ⊆ ·)] (a : α) : a ⊆ a := refl _ lemma subset_rfl [IsRefl α (· ⊆ ·)] : a ⊆ a := refl _ diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index 77d231d0c77a5..bf1c20483de89 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -1298,11 +1298,11 @@ theorem exists_pred_iterate_iff_le : (∃ n, pred^[n] b = a) ↔ a ≤ b := @[elab_as_elim] theorem Pred.rec {P : α → Prop} {m : α} (h0 : P m) (h1 : ∀ n, n ≤ m → P n → P (pred n)) ⦃n : α⦄ (hmn : n ≤ m) : P n := - @Succ.rec αᵒᵈ _ _ _ _ _ h0 h1 _ hmn + Succ.rec (α := αᵒᵈ) (P := P) h0 h1 hmn theorem Pred.rec_iff {p : α → Prop} (hsucc : ∀ a, p a ↔ p (pred a)) {a b : α} (h : a ≤ b) : p a ↔ p b := - (@Succ.rec_iff αᵒᵈ _ _ _ _ hsucc _ _ h).symm + (Succ.rec_iff (α := αᵒᵈ) hsucc h).symm end PredOrder diff --git a/Mathlib/Order/SuccPred/Relation.lean b/Mathlib/Order/SuccPred/Relation.lean index d7cb2358cc03b..3b9c6848d52a5 100644 --- a/Mathlib/Order/SuccPred/Relation.lean +++ b/Mathlib/Order/SuccPred/Relation.lean @@ -12,7 +12,6 @@ This file contains properties about relations on types with a `SuccOrder` and their closure operations (like the transitive closure). -/ - open Function Order Relation Set section PartialSucc @@ -88,25 +87,25 @@ variable {α : Type*} [PartialOrder α] [PredOrder α] [IsPredArchimedean α] for all `i` between `n` and `m`. -/ theorem reflTransGen_of_pred_of_ge (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc m n, r i (pred i)) (hnm : m ≤ n) : ReflTransGen r n m := - @reflTransGen_of_succ_of_le αᵒᵈ _ _ _ r n m (fun x hx => h x ⟨hx.2, hx.1⟩) hnm + reflTransGen_of_succ_of_le (α := αᵒᵈ) r (fun x hx => h x ⟨hx.2, hx.1⟩) hnm /-- For `n ≤ m`, `(n, m)` is in the reflexive-transitive closure of `~` if `pred i ~ i` for all `i` between `n` and `m`. -/ theorem reflTransGen_of_pred_of_le (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc n m, r (pred i) i) (hmn : n ≤ m) : ReflTransGen r n m := - @reflTransGen_of_succ_of_ge αᵒᵈ _ _ _ r n m (fun x hx => h x ⟨hx.2, hx.1⟩) hmn + reflTransGen_of_succ_of_ge (α := αᵒᵈ) r (fun x hx => h x ⟨hx.2, hx.1⟩) hmn /-- For `m < n`, `(n, m)` is in the transitive closure of a relation `~` for `n ≠ m` if `i ~ pred i` for all `i` between `n` and `m`. -/ theorem transGen_of_pred_of_gt (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc m n, r i (pred i)) (hnm : m < n) : TransGen r n m := - @transGen_of_succ_of_lt αᵒᵈ _ _ _ r _ _ (fun x hx => h x ⟨hx.2, hx.1⟩) hnm + transGen_of_succ_of_lt (α := αᵒᵈ) r (fun x hx => h x ⟨hx.2, hx.1⟩) hnm /-- For `n < m`, `(n, m)` is in the transitive closure of a relation `~` for `n ≠ m` if `pred i ~ i` for all `i` between `n` and `m`. -/ theorem transGen_of_pred_of_lt (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc n m, r (pred i) i) (hmn : n < m) : TransGen r n m := - @transGen_of_succ_of_gt αᵒᵈ _ _ _ r _ _ (fun x hx => h x ⟨hx.2, hx.1⟩) hmn + transGen_of_succ_of_gt (α := αᵒᵈ) r (fun x hx => h x ⟨hx.2, hx.1⟩) hmn end PartialPred @@ -118,21 +117,21 @@ variable {α : Type*} [LinearOrder α] [PredOrder α] [IsPredArchimedean α] for all `i` between `n` and `m`. -/ theorem reflTransGen_of_pred (r : α → α → Prop) {n m : α} (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) : ReflTransGen r n m := - @reflTransGen_of_succ αᵒᵈ _ _ _ r n m (fun x hx => h1 x ⟨hx.2, hx.1⟩) fun x hx => + reflTransGen_of_succ (α := αᵒᵈ) r (fun x hx => h1 x ⟨hx.2, hx.1⟩) fun x hx => h2 x ⟨hx.2, hx.1⟩ /-- For `n ≠ m`, `(n, m)` is in the transitive closure of a relation `~` if `i ~ pred i` and `pred i ~ i` for all `i` between `n` and `m`. -/ theorem transGen_of_pred_of_ne (r : α → α → Prop) {n m : α} (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) (hnm : n ≠ m) : TransGen r n m := - @transGen_of_succ_of_ne αᵒᵈ _ _ _ r n m (fun x hx => h1 x ⟨hx.2, hx.1⟩) + transGen_of_succ_of_ne (α := αᵒᵈ) r (fun x hx => h1 x ⟨hx.2, hx.1⟩) (fun x hx => h2 x ⟨hx.2, hx.1⟩) hnm /-- `(n, m)` is in the transitive closure of a reflexive relation `~` if `i ~ pred i` and `pred i ~ i` for all `i` between `n` and `m`. -/ theorem transGen_of_pred_of_reflexive (r : α → α → Prop) {n m : α} (hr : Reflexive r) (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) : TransGen r n m := - @transGen_of_succ_of_reflexive αᵒᵈ _ _ _ r n m hr (fun x hx => h1 x ⟨hx.2, hx.1⟩) fun x hx => + transGen_of_succ_of_reflexive (α := αᵒᵈ) r hr (fun x hx => h1 x ⟨hx.2, hx.1⟩) fun x hx => h2 x ⟨hx.2, hx.1⟩ end LinearPred diff --git a/Mathlib/Order/UpperLower/Basic.lean b/Mathlib/Order/UpperLower/Basic.lean index a8492f93da9a6..3e55eda001a9c 100644 --- a/Mathlib/Order/UpperLower/Basic.lean +++ b/Mathlib/Order/UpperLower/Basic.lean @@ -531,10 +531,15 @@ instance : SupSet (UpperSet α) := instance : InfSet (UpperSet α) := ⟨fun S => ⟨⋃ s ∈ S, ↑s, isUpperSet_iUnion₂ fun s _ => s.upper⟩⟩ -instance completelyDistribLattice : CompletelyDistribLattice (UpperSet α) := - (toDual.injective.comp SetLike.coe_injective).completelyDistribLattice _ (fun _ _ => rfl) +instance completeLattice : CompleteLattice (UpperSet α) := + (toDual.injective.comp SetLike.coe_injective).completeLattice _ (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl +instance completelyDistribLattice : CompletelyDistribLattice (UpperSet α) := + .ofMinimalAxioms $ + (toDual.injective.comp SetLike.coe_injective).completelyDistribLatticeMinimalAxioms .of _ + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl + instance : Inhabited (UpperSet α) := ⟨⊥⟩ @@ -661,10 +666,14 @@ instance : SupSet (LowerSet α) := instance : InfSet (LowerSet α) := ⟨fun S => ⟨⋂ s ∈ S, ↑s, isLowerSet_iInter₂ fun s _ => s.lower⟩⟩ -instance completelyDistribLattice : CompletelyDistribLattice (LowerSet α) := - SetLike.coe_injective.completelyDistribLattice _ (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) +instance completeLattice : CompleteLattice (LowerSet α) := + SetLike.coe_injective.completeLattice _ (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl +instance completelyDistribLattice : CompletelyDistribLattice (LowerSet α) := + .ofMinimalAxioms $ SetLike.coe_injective.completelyDistribLatticeMinimalAxioms .of _ + (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl + instance : Inhabited (LowerSet α) := ⟨⊥⟩ @@ -920,10 +929,10 @@ noncomputable instance LowerSet.instLinearOrder : LinearOrder (LowerSet α) := b classical exact Lattice.toLinearOrder _ noncomputable instance UpperSet.instCompleteLinearOrder : CompleteLinearOrder (UpperSet α) := - { completelyDistribLattice, instLinearOrder, LinearOrder.toBiheytingAlgebra with } + { completelyDistribLattice, instLinearOrder with } noncomputable instance LowerSet.instCompleteLinearOrder : CompleteLinearOrder (LowerSet α) := - { completelyDistribLattice, instLinearOrder, LinearOrder.toBiheytingAlgebra with } + { completelyDistribLattice, instLinearOrder with } end LinearOrder diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 00be72e844c10..93169e4c45d8d 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -122,9 +122,10 @@ protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] { section LinearOrder -variable [LinearOrder β] (h : WellFounded ((· < ·) : β → β → Prop)) [PartialOrder γ] +variable [LinearOrder β] [PartialOrder γ] -theorem min_le {x : β} {s : Set β} (hx : x ∈ s) (hne : s.Nonempty := ⟨x, hx⟩) : h.min s hne ≤ x := +theorem min_le (h : WellFounded ((· < ·) : β → β → Prop)) {x : β} {s : Set β} (hx : x ∈ s) + (hne : s.Nonempty := ⟨x, hx⟩) : h.min s hne ≤ x := not_lt.1 <| h.not_lt_min _ _ hx private theorem eq_strictMono_iff_eq_range_aux {f g : β → γ} (hf : StrictMono f) @@ -140,7 +141,8 @@ private theorem eq_strictMono_iff_eq_range_aux {f g : β → γ} (hf : StrictMon · rw [← hc] exact hf.monotone hbc -theorem eq_strictMono_iff_eq_range {f g : β → γ} (hf : StrictMono f) (hg : StrictMono g) : +theorem eq_strictMono_iff_eq_range (h : WellFounded ((· < ·) : β → β → Prop)) {f g : β → γ} + (hf : StrictMono f) (hg : StrictMono g) : Set.range f = Set.range g ↔ f = g := ⟨fun hfg => by funext a @@ -150,7 +152,8 @@ theorem eq_strictMono_iff_eq_range {f g : β → γ} (hf : StrictMono f) (hg : S (eq_strictMono_iff_eq_range_aux hg hf hfg.symm fun a hab => (H a hab).symm), congr_arg _⟩ -theorem self_le_of_strictMono {f : β → β} (hf : StrictMono f) : ∀ n, n ≤ f n := by +theorem self_le_of_strictMono (h : WellFounded ((· < ·) : β → β → Prop)) {f : β → β} + (hf : StrictMono f) : ∀ n, n ≤ f n := by by_contra! h₁ have h₂ := h.min_mem _ h₁ exact h.not_lt_min _ h₁ (hf h₂) h₂ diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index c345a1655fd16..9720f0923febe 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -174,6 +174,13 @@ theorem unbot_coe (x : α) (h : (x : WithBot α) ≠ ⊥ := coe_ne_bot) : (x : W instance canLift : CanLift (WithBot α) α (↑) fun r => r ≠ ⊥ where prf x h := ⟨x.unbot h, coe_unbot _ _⟩ +instance instTop [Top α] : Top (WithBot α) where + top := (⊤ : α) + +@[simp, norm_cast] lemma coe_top [Top α] : ((⊤ : α) : WithBot α) = ⊤ := rfl +@[simp, norm_cast] lemma coe_eq_top [Top α] {a : α} : (a : WithBot α) = ⊤ ↔ a = ⊤ := coe_eq_coe +@[simp, norm_cast] lemma top_eq_coe [Top α] {a : α} : ⊤ = (a : WithBot α) ↔ ⊤ = a := coe_eq_coe + section LE variable [LE α] @@ -195,13 +202,6 @@ theorem some_le_some : @LE.le (WithBot α) _ (Option.some a) (Option.some b) ↔ @[simp, deprecated bot_le "Don't mix Option and WithBot" (since := "2024-05-27")] theorem none_le {a : WithBot α} : @LE.le (WithBot α) _ none a := bot_le -instance instTop [Top α] : Top (WithBot α) where - top := (⊤ : α) - -@[simp, norm_cast] lemma coe_top [Top α] : ((⊤ : α) : WithBot α) = ⊤ := rfl -@[simp, norm_cast] lemma coe_eq_top [Top α] {a : α} : (a : WithBot α) = ⊤ ↔ a = ⊤ := coe_eq_coe -@[simp, norm_cast] lemma top_eq_coe [Top α] {a : α} : ⊤ = (a : WithBot α) ↔ ⊤ = a := coe_eq_coe - instance orderTop [OrderTop α] : OrderTop (WithBot α) where le_top o a ha := by cases ha; exact ⟨_, rfl, le_top⟩ @@ -752,6 +752,13 @@ theorem untop_coe (x : α) (h : (x : WithTop α) ≠ ⊤ := coe_ne_top) : (x : W instance canLift : CanLift (WithTop α) α (↑) fun r => r ≠ ⊤ where prf x h := ⟨x.untop h, coe_untop _ _⟩ +instance instBot [Bot α] : Bot (WithTop α) where + bot := (⊥ : α) + +@[simp, norm_cast] lemma coe_bot [Bot α] : ((⊥ : α) : WithTop α) = ⊥ := rfl +@[simp, norm_cast] lemma coe_eq_bot [Bot α] {a : α} : (a : WithTop α) = ⊥ ↔ a = ⊥ := coe_eq_coe +@[simp, norm_cast] lemma bot_eq_coe [Bot α] {a : α} : (⊥ : WithTop α) = a ↔ ⊥ = a := coe_eq_coe + section LE variable [LE α] @@ -797,13 +804,6 @@ instance orderTop : OrderTop (WithTop α) where @[simp, deprecated le_top "Don't mix Option and WithTop" (since := "2024-05-27")] theorem le_none {a : WithTop α} : @LE.le (WithTop α) _ a none := le_top -instance instBot [Bot α] : Bot (WithTop α) where - bot := (⊥ : α) - -@[simp, norm_cast] lemma coe_bot [Bot α] : ((⊥ : α) : WithTop α) = ⊥ := rfl -@[simp, norm_cast] lemma coe_eq_bot [Bot α] {a : α} : (a : WithTop α) = ⊥ ↔ a = ⊥ := coe_eq_coe -@[simp, norm_cast] lemma bot_eq_coe [Bot α] {a : α} : (⊥ : WithTop α) = a ↔ ⊥ = a := coe_eq_coe - instance orderBot [OrderBot α] : OrderBot (WithTop α) where bot_le o a ha := by cases ha; exact ⟨_, rfl, bot_le⟩ diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index a594d40e19c29..de801087d5074 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -224,7 +224,8 @@ lemma sum_meas_smul_cond_fiber {X : Ω → α} (hX : Measurable X) (μ : Measure Pi.smul_apply, smul_eq_mul] simp_rw [mul_comm (μ _), cond_mul_eq_inter _ (hX (.singleton _))] _ = _ := by - have : ⋃ x ∈ Finset.univ, X ⁻¹' {x} ∩ E = E := by simp; ext _; simp + have : ⋃ x ∈ Finset.univ, X ⁻¹' {x} ∩ E = E := by + simp only [Finset.mem_univ, iUnion_true]; ext _; simp rw [← measure_biUnion_finset _ fun _ _ ↦ (hX (.singleton _)).inter hE, this] aesop (add simp [PairwiseDisjoint, Set.Pairwise, Function.onFun, disjoint_left]) diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index c5b8cc7a5a278..3786b0d00174d 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -206,6 +206,12 @@ protected theorem measurable_coe (κ : Kernel α β) {s : Set β} (hs : Measurab Measurable fun a => κ a s := (Measure.measurable_coe hs).comp κ.measurable +lemma apply_congr_of_mem_measurableAtom (κ : Kernel α β) {y' y : α} (hy' : y' ∈ measurableAtom y) : + κ y' = κ y := by + ext s hs + exact mem_of_mem_measurableAtom hy' + (κ.measurable_coe hs (measurableSet_singleton (κ y s))) rfl + lemma IsFiniteKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] (κ : Kernel α β) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) : Integrable (fun x => (κ x s).toReal) μ := by @@ -294,7 +300,8 @@ instance isFiniteKernel_seq (κ : Kernel α β) [h : IsSFiniteKernel κ] (n : IsFiniteKernel (Kernel.seq κ n) := h.tsum_finite.choose_spec.1 n -instance IsSFiniteKernel.sFinite [IsSFiniteKernel κ] (a : α) : SFinite (κ a) := +instance _root_.ProbabilityTheory.IsSFiniteKernel.sFinite [IsSFiniteKernel κ] (a : α) : + SFinite (κ a) := ⟨⟨fun n ↦ seq κ n a, inferInstance, (measure_sum_seq κ a).symm⟩⟩ instance IsSFiniteKernel.add (κ η : Kernel α β) [IsSFiniteKernel κ] [IsSFiniteKernel η] : @@ -443,18 +450,22 @@ lemma sum_const [Countable ι] (μ : ι → Measure β) : rw [const_apply, Measure.sum_apply _ hs, Kernel.sum_apply' _ _ hs] simp only [const_apply] -instance isFiniteKernel_const {μβ : Measure β} [IsFiniteMeasure μβ] : +instance const.instIsFiniteKernel {μβ : Measure β} [IsFiniteMeasure μβ] : IsFiniteKernel (const α μβ) := ⟨⟨μβ Set.univ, measure_lt_top _ _, fun _ => le_rfl⟩⟩ -instance isSFiniteKernel_const {μβ : Measure β} [SFinite μβ] : +instance const.instIsSFiniteKernel {μβ : Measure β} [SFinite μβ] : IsSFiniteKernel (const α μβ) := ⟨fun n ↦ const α (sFiniteSeq μβ n), fun n ↦ inferInstance, by rw [sum_const, sum_sFiniteSeq]⟩ -instance isMarkovKernel_const {μβ : Measure β} [hμβ : IsProbabilityMeasure μβ] : +instance const.instIsMarkovKernel {μβ : Measure β} [hμβ : IsProbabilityMeasure μβ] : IsMarkovKernel (const α μβ) := ⟨fun _ => hμβ⟩ +lemma isSFiniteKernel_const [Nonempty α] {μβ : Measure β} : + IsSFiniteKernel (const α μβ) ↔ SFinite μβ := + ⟨fun h ↦ h.sFinite (Classical.arbitrary α), fun _ ↦ inferInstance⟩ + @[simp] theorem lintegral_const {f : β → ℝ≥0∞} {μ : Measure β} {a : α} : ∫⁻ x, f x ∂const α μ a = ∫⁻ x, f x ∂μ := by rw [const_apply] diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 9adb75968781a..a8650f58e53c3 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -392,9 +392,7 @@ theorem lintegral_compProd' (κ : Kernel α β) [IsSFiniteKernel κ] (η : Kerne · exact fun i j hij b => lintegral_mono fun c => h_mono hij _ congr ext1 n - -- Porting note: Added `(P := _)` - refine SimpleFunc.induction (P := fun f => (∫⁻ (a : β × γ), f a ∂(κ ⊗ₖ η) a = - ∫⁻ (a_1 : β), ∫⁻ (c : γ), f (a_1, c) ∂η (a, a_1) ∂κ a)) ?_ ?_ (F n) + refine SimpleFunc.induction ?_ ?_ (F n) · intro c s hs classical -- Porting note: Added `classical` for `Set.piecewise_eq_indicator` simp (config := { unfoldPartialApp := true }) only [SimpleFunc.const_zero, @@ -768,6 +766,18 @@ instance IsSFiniteKernel.prodMkLeft (κ : Kernel α β) [IsSFiniteKernel κ] : instance IsSFiniteKernel.prodMkRight (κ : Kernel α β) [IsSFiniteKernel κ] : IsSFiniteKernel (prodMkRight γ κ) := by rw [Kernel.prodMkRight]; infer_instance +lemma isSFiniteKernel_prodMkLeft_unit {κ : Kernel α β} : + IsSFiniteKernel (prodMkLeft Unit κ) ↔ IsSFiniteKernel κ := by + refine ⟨fun _ ↦ ?_, fun _ ↦ inferInstance⟩ + change IsSFiniteKernel ((prodMkLeft Unit κ).comap (fun a ↦ ((), a)) (by fun_prop)) + infer_instance + +lemma isSFiniteKernel_prodMkRight_unit {κ : Kernel α β} : + IsSFiniteKernel (prodMkRight Unit κ) ↔ IsSFiniteKernel κ := by + refine ⟨fun _ ↦ ?_, fun _ ↦ inferInstance⟩ + change IsSFiniteKernel ((prodMkRight Unit κ).comap (fun a ↦ (a, ())) (by fun_prop)) + infer_instance + lemma map_prodMkLeft (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) {f : β → δ} (hf : Measurable f) : map (prodMkLeft γ κ) f hf = prodMkLeft γ (map κ f hf) := rfl diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index c304e0e2987b1..13f6e4c87bf7b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -1,504 +1,168 @@ /- -Copyright (c) 2024 Rémy Degenne. All rights reserved. +Copyright (c) 2024 Yaël Dillies, Kin Yau James Wong. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Rémy Degenne +Authors: Yaël Dillies, Kin Yau James Wong, Rémy Degenne -/ import Mathlib.Probability.Kernel.MeasureCompProd -import Mathlib.Probability.Kernel.Disintegration.CondCdf -import Mathlib.Probability.Kernel.Disintegration.Density -import Mathlib.Probability.Kernel.Disintegration.CdfToKernel /-! -# Disintegration of kernels and measures - -Let `κ : Kernel α (β × Ω)` be a finite kernel, where `Ω` is a standard Borel space. Then if `α` is -countable or `β` has a countably generated σ-algebra (for example if it is standard Borel), then -there exists a `Kernel (α × β) Ω` called conditional kernel and denoted by `condKernel κ` such that -`κ = fst κ ⊗ₖ condKernel κ`. -We also define a conditional kernel for a measure `ρ : Measure (β × Ω)`, where `Ω` is a standard -Borel space. This is a `Kernel β Ω` denoted by `ρ.condKernel` such that `ρ = ρ.fst ⊗ₘ ρ.condKernel`. - -In order to obtain a disintegration for any standard Borel space `Ω`, we use that these spaces embed -measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. - -For `κ : Kernel α (β × ℝ)`, the construction of the conditional kernel proceeds as follows: -* Build a measurable function `f : (α × β) → ℚ → ℝ` such that for all measurable sets - `s` and all `q : ℚ`, `∫ x in s, f (a, x) q ∂(Kernel.fst κ a) = (κ a (s ×ˢ Iic (q : ℝ))).toReal`. - We restrict to `ℚ` here to be able to prove the measurability. -* Extend that function to `(α × β) → StieltjesFunction`. See the file `MeasurableStieltjes.lean`. -* Finally obtain from the measurable Stieltjes function a measure on `ℝ` for each element of `α × β` - in a measurable way: we have obtained a `Kernel (α × β) ℝ`. - See the file `CdfToKernel.lean` for that step. - -The first step (building the measurable function on `ℚ`) is done differently depending on whether -`α` is countable or not. -* If `α` is countable, we can provide for each `a : α` a function `f : β → ℚ → ℝ` and proceed as - above to obtain a `Kernel β ℝ`. Since `α` is countable, measurability is not an issue and we can - put those together into a `Kernel (α × β) ℝ`. The construction of that `f` is done in - the `CondCdf.lean` file. -* If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function - `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a - countably generated σ-algebra (this is the case in particular for standard Borel spaces). - See the file `Density.lean`. - -The conditional kernel is defined under the typeclass assumption -`CountableOrCountablyGenerated α β`, which encodes the property -`Countable α ∨ CountablyGenerated β`. - -Properties of integrals involving `condKernel` are collated in the file `Integral.lean`. -The conditional kernel is unique (almost everywhere w.r.t. `fst κ`): this is proved in the file -`Unique.lean`. +# Disintegration of measures and kernels -## Main definitions +This file defines predicates for a kernel to "disintegrate" a measure or a kernel. This kernel is +also called the "conditional kernel" of the measure or kernel. -* `ProbabilityTheory.Kernel.condKernel κ : Kernel (α × β) Ω`: conditional kernel described above. -* `MeasureTheory.Measure.condKernel ρ : Kernel β Ω`: conditional kernel of a measure. +A measure `ρ : Measure (α × Ω)` is disintegrated by a kernel `ρCond : Kernel α Ω` if +`ρ.fst ⊗ₘ ρCond = ρ`. -## Main statements +A kernel `ρ : Kernel α (β × Ω)` is disintegrated by a kernel `κCond : Kernel (α × β) Ω` if +`κ.fst ⊗ₖ κCond = κ`. -* `ProbabilityTheory.Kernel.compProd_fst_condKernel`: `fst κ ⊗ₖ condKernel κ = κ` -* `MeasureTheory.Measure.compProd_fst_condKernel`: `ρ.fst ⊗ₘ ρ.condKernel = ρ` --/ +## Main definitions -open MeasureTheory Set Filter MeasurableSpace +* `MeasureTheory.Measure.IsCondKernel ρ ρCond`: Predicate for the kernel `ρCond` to disintegrate the + measure `ρ`. +* `ProbabilityTheory.Kernel.IsCondKernel κ κCond`: Predicate for the kernel `κ Cond` to disintegrate + the kernel `κ`. -open scoped ENNReal MeasureTheory Topology ProbabilityTheory +Further, if `κ` is an s-finite kernel from a countable `α` such that each measure `κ a` is +disintegrated by some kernel, then `κ` itself is disintegrated by a kernel, namely +`ProbabilityTheory.Kernel.condKernelCountable`. -namespace ProbabilityTheory.Kernel +## See also -variable {α β γ Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} - {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] - [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] - -section Real - -/-! ### Disintegration of kernels from `α` to `γ × ℝ` for countably generated `γ` -/ - -lemma isRatCondKernelCDFAux_density_Iic (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) where - measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (fst κ) measurableSet_Iic - mono' a q r hqr := - ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) - nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ - le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ - tendsto_integral_of_antitone a s hs_anti hs_tendsto := by - let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - refine tendsto_integral_density_of_antitone le_rfl a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) - · refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ - exact mod_cast hs_anti hij - · ext x - simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] - rw [tendsto_atTop_atBot] at hs_tendsto - have ⟨q, hq⟩ := exists_rat_lt x - obtain ⟨i, hi⟩ := hs_tendsto q - refine ⟨i, lt_of_le_of_lt ?_ hq⟩ - exact mod_cast hi i le_rfl - tendsto_integral_of_monotone a s hs_mono hs_tendsto := by - rw [fst_apply' _ _ MeasurableSet.univ] - let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) - refine tendsto_integral_density_of_monotone (le_rfl : fst κ ≤ fst κ) - a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) - · exact fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) - · ext x - simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] - rw [tendsto_atTop_atTop] at hs_tendsto - have ⟨q, hq⟩ := exists_rat_gt x - obtain ⟨i, hi⟩ := hs_tendsto q - refine ⟨i, hq.le.trans ?_⟩ - exact mod_cast hi i le_rfl - integrable a q := integrable_density le_rfl a measurableSet_Iic - setIntegral a A hA q := setIntegral_density le_rfl a measurableSet_Iic hA - -/-- Taking the kernel density of intervals `Iic q` for `q : ℚ` gives a function with the property -`isRatCondKernelCDF`. -/ -lemma isRatCondKernelCDF_density_Iic (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsRatCondKernelCDF (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) := - (isRatCondKernelCDFAux_density_Iic κ).isRatCondKernelCDF - -/-- The conditional kernel CDF of a kernel `κ : Kernel α (γ × ℝ)`, where `γ` is countably generated. +`Mathlib.Probability.Kernel.Disintegration.StandardBorel` for a **construction** of disintegrating +kernels. -/ -noncomputable -def condKernelCDF (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := - stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) - (isRatCondKernelCDF_density_Iic κ).measurable -lemma isCondKernelCDF_condKernelCDF (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsCondKernelCDF (condKernelCDF κ) κ (fst κ) := - isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) +open MeasureTheory Set Filter MeasurableSpace ProbabilityTheory +open scoped ENNReal MeasureTheory Topology -/-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. -A conditional kernel for `κ : Kernel α (γ × ℝ)` where `γ` is countably generated. -/ -noncomputable -def condKernelReal (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : Kernel (α × γ) ℝ := - (isCondKernelCDF_condKernelCDF κ).toKernel +variable {α β Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mΩ : MeasurableSpace Ω} -instance instIsMarkovKernelCondKernelReal (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelReal κ) := by - rw [condKernelReal] - infer_instance - -lemma compProd_fst_condKernelReal (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelReal κ = κ := by - rw [condKernelReal, compProd_toKernel] - -/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and -`ProbabilityTheory.Kernel.condKernel`. -A conditional kernel for `κ : Kernel Unit (α × ℝ)`. -/ -noncomputable -def condKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : Kernel (Unit × α) ℝ := - (isCondKernelCDF_condCDF (κ ())).toKernel - -instance instIsMarkovKernelCondKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelUnitReal κ) := by - rw [condKernelUnitReal] - infer_instance +/-! +### Disintegration of measures -lemma compProd_fst_condKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelUnitReal κ = κ := by - rw [condKernelUnitReal, compProd_toKernel] - ext a - simp +This section provides a predicate for a kernel to disintegrate a measure. +-/ -end Real +namespace MeasureTheory.Measure +variable (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] (ρCond : Kernel α Ω) -section BorelSnd +/-- A kernel `ρCond` is a conditional kernel for a measure `ρ` if it disintegrates it in the sense +that `ρ.fst ⊗ₘ ρCond = ρ`. -/ +class IsCondKernel : Prop where + disintegrate : ρ.fst ⊗ₘ ρCond = ρ -/-! ### Disintegration of kernels on standard Borel spaces +variable [ρ.IsCondKernel ρCond] -Since every standard Borel space embeds measurably into `ℝ`, we can generalize a disintegration -property on `ℝ` to all these spaces. -/ +lemma disintegrate : ρ.fst ⊗ₘ ρCond = ρ := IsCondKernel.disintegrate -open Classical in -/-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. -A Borel space `Ω` embeds measurably into `ℝ` (with embedding `e`), hence we can get a `Kernel α Ω` -from a `Kernel α ℝ` by taking the comap by `e`. -Here we take the comap of a modification of `η : Kernel α ℝ`, useful when `η a` is a probability -measure with all its mass on `range e` almost everywhere with respect to some measure and we want to -ensure that the comap is a Markov kernel. -We thus take the comap by `e` of a kernel defined piecewise: `η` when -`η a (range (embeddingReal Ω))ᶜ = 0`, and an arbitrary deterministic kernel otherwise. -/ -noncomputable -def borelMarkovFromReal (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] - (η : Kernel α ℝ) : - Kernel α Ω := - have he := measurableEmbedding_embeddingReal Ω - let x₀ := (range_nonempty (embeddingReal Ω)).choose - comapRight - (piecewise ((Kernel.measurable_coe η he.measurableSet_range.compl) (measurableSet_singleton 0) : - MeasurableSet {a | η a (range (embeddingReal Ω))ᶜ = 0}) - η (deterministic (fun _ ↦ x₀) measurable_const)) he - -lemma borelMarkovFromReal_apply (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] - (η : Kernel α ℝ) (a : α) : - borelMarkovFromReal Ω η a - = if η a (range (embeddingReal Ω))ᶜ = 0 then (η a).comap (embeddingReal Ω) - else (Measure.dirac (range_nonempty (embeddingReal Ω)).choose).comap (embeddingReal Ω) := by - classical - rw [borelMarkovFromReal, comapRight_apply, piecewise_apply, deterministic_apply] - simp only [mem_preimage, mem_singleton_iff] - split_ifs <;> rfl - -lemma borelMarkovFromReal_apply' (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] - (η : Kernel α ℝ) (a : α) {s : Set Ω} (hs : MeasurableSet s) : - borelMarkovFromReal Ω η a s - = if η a (range (embeddingReal Ω))ᶜ = 0 then η a (embeddingReal Ω '' s) - else (embeddingReal Ω '' s).indicator 1 (range_nonempty (embeddingReal Ω)).choose := by - have he := measurableEmbedding_embeddingReal Ω - rw [borelMarkovFromReal_apply] - split_ifs with h - · rw [Measure.comap_apply _ he.injective he.measurableSet_image' _ hs] - · rw [Measure.comap_apply _ he.injective he.measurableSet_image' _ hs, Measure.dirac_apply] - -/-- When `η` is an s-finite kernel, `borelMarkovFromReal Ω η` is an s-finite kernel. -/ -instance instIsSFiniteKernelBorelMarkovFromReal (η : Kernel α ℝ) [IsSFiniteKernel η] : - IsSFiniteKernel (borelMarkovFromReal Ω η) := - IsSFiniteKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) - -/-- When `η` is a finite kernel, `borelMarkovFromReal Ω η` is a finite kernel. -/ -instance instIsFiniteKernelBorelMarkovFromReal (η : Kernel α ℝ) [IsFiniteKernel η] : - IsFiniteKernel (borelMarkovFromReal Ω η) := - IsFiniteKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) - -/-- When `η` is a Markov kernel, `borelMarkovFromReal Ω η` is a Markov kernel. -/ -instance instIsMarkovKernelBorelMarkovFromReal (η : Kernel α ℝ) [IsMarkovKernel η] : - IsMarkovKernel (borelMarkovFromReal Ω η) := by - refine IsMarkovKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) (fun a ↦ ?_) - classical - rw [piecewise_apply] - split_ifs with h - · rwa [← prob_compl_eq_zero_iff (measurableEmbedding_embeddingReal Ω).measurableSet_range] - · rw [deterministic_apply] - simp [(range_nonempty (embeddingReal Ω)).choose_spec] - -/-- For `κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable)`, the -hypothesis `hη` is `fst κ' ⊗ₖ η = κ'`. The conclusion of the lemma is -`fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) _`. -/ -lemma compProd_fst_borelMarkovFromReal_eq_comapRight_compProd - (κ : Kernel α (β × Ω)) [IsSFiniteKernel κ] (η : Kernel (α × β) ℝ) [IsSFiniteKernel η] - (hη : (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable))) ⊗ₖ η - = map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) : - fst κ ⊗ₖ borelMarkovFromReal Ω η - = comapRight (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) ⊗ₖ η) - (MeasurableEmbedding.id.prod_mk (measurableEmbedding_embeddingReal Ω)) := by - let e := embeddingReal Ω - let he := measurableEmbedding_embeddingReal Ω - let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) - have hη' : fst κ' ⊗ₖ η = κ' := hη - have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := - MeasurableEmbedding.id.prod_mk he - change fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) h_prod_embed - rw [comapRight_compProd_id_prod _ _ he] - have h_fst : fst κ' = fst κ := by - ext a u - unfold_let κ' - rw [fst_apply, map_apply, Measure.map_map measurable_fst h_prod_embed.measurable, fst_apply] - congr - rw [h_fst] - ext a t ht : 2 - simp_rw [compProd_apply _ _ _ ht] - refine lintegral_congr_ae ?_ - have h_ae : ∀ᵐ t ∂(fst κ a), (a, t) ∈ {p : α × β | η p (range e)ᶜ = 0} := by - rw [← h_fst] - have h_compProd : κ' a (univ ×ˢ range e)ᶜ = 0 := by - unfold_let κ' - rw [map_apply'] - swap; · exact (MeasurableSet.univ.prod he.measurableSet_range).compl - suffices Prod.map id e ⁻¹' (univ ×ˢ range e)ᶜ = ∅ by rw [this]; simp - ext x - simp - rw [← hη', compProd_null] at h_compProd - swap; · exact (MeasurableSet.univ.prod he.measurableSet_range).compl - simp only [preimage_compl, mem_univ, mk_preimage_prod_right] at h_compProd - exact h_compProd - filter_upwards [h_ae] with a ha - rw [borelMarkovFromReal, comapRight_apply', comapRight_apply'] - rotate_left - · exact measurable_prod_mk_left ht - · exact measurable_prod_mk_left ht - classical - rw [piecewise_apply, if_pos] - exact ha - -/-- For `κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable)`, the -hypothesis `hη` is `fst κ' ⊗ₖ η = κ'`. With that hypothesis, -`fst κ ⊗ₖ borelMarkovFromReal κ η = κ`.-/ -lemma compProd_fst_borelMarkovFromReal (κ : Kernel α (β × Ω)) [IsSFiniteKernel κ] - (η : Kernel (α × β) ℝ) [IsSFiniteKernel η] - (hη : (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable))) ⊗ₖ η - = map κ (Prod.map (id : β → β) (embeddingReal Ω)) - (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) : - fst κ ⊗ₖ borelMarkovFromReal Ω η = κ := by - let e := embeddingReal Ω - let he := measurableEmbedding_embeddingReal Ω - let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) - have hη' : fst κ' ⊗ₖ η = κ' := hη - have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := - MeasurableEmbedding.id.prod_mk he - have : κ = comapRight κ' h_prod_embed := by - ext c t : 2 - unfold_let κ' - rw [comapRight_apply, map_apply, h_prod_embed.comap_map] - conv_rhs => rw [this, ← hη'] - exact compProd_fst_borelMarkovFromReal_eq_comapRight_compProd κ η hη - -end BorelSnd - -section CountablyGenerated - -open ProbabilityTheory.Kernel +lemma IsCondKernel.isSFiniteKernel (hρ : ρ ≠ 0) : IsSFiniteKernel ρCond := by + contrapose! hρ; rwa [← ρ.disintegrate ρCond, Measure.compProd_of_not_isSFiniteKernel] -/-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. -A conditional kernel for `κ : Kernel α (γ × Ω)` where `γ` is countably generated and `Ω` is -standard Borel. -/ -noncomputable -def condKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : Kernel (α × γ) Ω := - let e := embeddingReal Ω - let he := measurableEmbedding_embeddingReal Ω - let κ' := map κ (Prod.map (id : γ → γ) e) (measurable_id.prod_map he.measurable) - borelMarkovFromReal Ω (condKernelReal κ') - -instance instIsMarkovKernelCondKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelBorel κ) := by - rw [condKernelBorel] - infer_instance - -lemma compProd_fst_condKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelBorel κ = κ := by - rw [condKernelBorel, compProd_fst_borelMarkovFromReal _ _ (compProd_fst_condKernelReal _)] - -end CountablyGenerated - -section Unit - -/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and -`ProbabilityTheory.Kernel.condKernel`. -A conditional kernel for `κ : Kernel Unit (α × Ω)` where `Ω` is standard Borel. -/ -noncomputable -def condKernelUnitBorel (κ : Kernel Unit (α × Ω)) [IsFiniteKernel κ] : Kernel (Unit × α) Ω := - let e := embeddingReal Ω - let he := measurableEmbedding_embeddingReal Ω - let κ' := map κ (Prod.map (id : α → α) e) (measurable_id.prod_map he.measurable) - borelMarkovFromReal Ω (condKernelUnitReal κ') - -instance instIsMarkovKernelCondKernelUnitBorel (κ : Kernel Unit (α × Ω)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelUnitBorel κ) := by - rw [condKernelUnitBorel] - infer_instance - -lemma compProd_fst_condKernelUnitBorel (κ : Kernel Unit (α × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelUnitBorel κ = κ := by - rw [condKernelUnitBorel, - compProd_fst_borelMarkovFromReal _ _ (compProd_fst_condKernelUnitReal _)] - -end Unit - -section Measure - -variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] - -/-- Conditional kernel of a measure on a product space: a Markov kernel such that -`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `MeasureTheory.Measure.compProd_fst_condKernel`). -/ -noncomputable -irreducible_def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : - Kernel α Ω := - comap (condKernelUnitBorel (const Unit ρ)) (fun a ↦ ((), a)) measurable_prod_mk_left - -lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] - (a : α) : - ρ.condKernel a = condKernelUnitBorel (const Unit ρ) ((), a) := by - rw [Measure.condKernel]; rfl - -instance _root_.MeasureTheory.Measure.instIsMarkovKernelCondKernel - (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : IsMarkovKernel ρ.condKernel := by - rw [Measure.condKernel] - infer_instance - -/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a -measure can be written as the composition-product of `ρ.fst` (marginal measure over `α`) and -a Markov kernel from `α` to `Ω`. We call that Markov kernel `ρ.condKernel`. -/ -lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel - (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : - ρ.fst ⊗ₘ ρ.condKernel = ρ := by - have h1 : const Unit (Measure.fst ρ) = fst (const Unit ρ) := by - ext - simp only [fst_apply, Measure.fst, const_apply] - have h2 : prodMkLeft Unit (Measure.condKernel ρ) = condKernelUnitBorel (const Unit ρ) := by - ext - simp only [prodMkLeft_apply, Measure.condKernel_apply] - rw [Measure.compProd, h1, h2, compProd_fst_condKernelUnitBorel] - simp - -/-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ -lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero_of_measurableSet - [MeasurableSingletonClass α] {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by - nth_rewrite 3 [← ρ.compProd_fst_condKernel] +/-- Auxiliary lemma for `IsCondKernel.apply_of_ne_zero`. -/ +private lemma IsCondKernel.apply_of_ne_zero_of_measurableSet [MeasurableSingletonClass α] {x : α} + (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : + ρCond x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by + have := isSFiniteKernel ρ ρCond (by rintro rfl; simp at hx) + nth_rewrite 2 [← ρ.disintegrate ρCond] rw [Measure.compProd_apply (measurableSet_prod.mpr (Or.inl ⟨measurableSet_singleton x, hs⟩))] classical - have : ∀ a, ρ.condKernel a (Prod.mk a ⁻¹' {x} ×ˢ s) - = ({x} : Set α).indicator (fun a ↦ ρ.condKernel a s) a := by - intro a - by_cases hax : a = x - · simp only [hax, singleton_prod, mem_singleton_iff, indicator_of_mem] + have (a) : ρCond a (Prod.mk a ⁻¹' {x} ×ˢ s) = ({x} : Set α).indicator (ρCond · s) a := by + obtain rfl | hax := eq_or_ne a x + · simp only [singleton_prod, mem_singleton_iff, indicator_of_mem] congr with y simp · simp only [singleton_prod, mem_singleton_iff, hax, not_false_eq_true, indicator_of_not_mem] - have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by - ext y - simp [Ne.symm hax] + have : Prod.mk a ⁻¹' (Prod.mk x '' s) = ∅ := by ext y; simp [Ne.symm hax] simp only [this, measure_empty] simp_rw [this] rw [MeasureTheory.lintegral_indicator _ (measurableSet_singleton x)] simp only [Measure.restrict_singleton, lintegral_smul_measure, lintegral_dirac] - rw [← mul_assoc, ENNReal.inv_mul_cancel hx (measure_ne_top ρ.fst _), one_mul] + rw [← mul_assoc, ENNReal.inv_mul_cancel hx (measure_ne_top _ _), one_mul] /-- If the singleton `{x}` has non-zero mass for `ρ.fst`, then for all `s : Set Ω`, -`ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ -lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero [MeasurableSingletonClass α] - {x : α} (hx : ρ.fst {x} ≠ 0) (s : Set Ω) : - ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by - have : ρ.condKernel x s = ((ρ.fst {x})⁻¹ • ρ).comap (fun (y : Ω) ↦ (x, y)) s := by +`ρCond x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ +lemma IsCondKernel.apply_of_ne_zero [MeasurableSingletonClass α] {x : α} + (hx : ρ.fst {x} ≠ 0) (s : Set Ω) : ρCond x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := by + have : ρCond x s = ((ρ.fst {x})⁻¹ • ρ).comap (fun (y : Ω) ↦ (x, y)) s := by congr 2 with s hs - simp [Measure.condKernel_apply_of_ne_zero_of_measurableSet hx hs, + simp [IsCondKernel.apply_of_ne_zero_of_measurableSet _ _ hx hs, (measurableEmbedding_prod_mk_left x).comap_apply] simp [this, (measurableEmbedding_prod_mk_left x).comap_apply, hx] -end Measure +lemma IsCondKernel.isProbabilityMeasure [MeasurableSingletonClass α] {a : α} (ha : ρ.fst {a} ≠ 0) : + IsProbabilityMeasure (ρCond a) := by + constructor + rw [IsCondKernel.apply_of_ne_zero _ _ ha, prod_univ, ← Measure.fst_apply + (measurableSet_singleton _), ENNReal.inv_mul_cancel ha (measure_ne_top _ _)] -section Countable +lemma IsCondKernel.isMarkovKernel [MeasurableSingletonClass α] (hρ : ∀ a, ρ.fst {a} ≠ 0) : + IsMarkovKernel ρCond := ⟨fun _ ↦ isProbabilityMeasure _ _ (hρ _)⟩ -variable [Countable α] +end MeasureTheory.Measure + +/-! +### Disintegration of kernels + +This section provides a predicate for a kernel to disintegrate a kernel. It also proves that if `κ` +is an s-finite kernel from a countable `α` such that each measure `κ a` is disintegrated by some +kernel, then `κ` itself is disintegrated by a kernel, namely +`ProbabilityTheory.Kernel.condKernelCountable`.. +-/ + +namespace ProbabilityTheory.Kernel +variable (κ : Kernel α (β × Ω)) (κCond : Kernel (α × β) Ω) + +/-! #### Predicate for a kernel to disintegrate a kernel -/ + +/-- A kernel `κCond` is a conditional kernel for a kernel `κ` if it disintegrates it in the sense +that `κ.fst ⊗ₖ κCond = κ`. -/ +class IsCondKernel : Prop where + protected disintegrate : κ.fst ⊗ₖ κCond = κ + +instance instIsCondKernel_zero (κCond : Kernel (α × β) Ω) : IsCondKernel 0 κCond where + disintegrate := by simp + +variable [κ.IsCondKernel κCond] + +lemma disintegrate : κ.fst ⊗ₖ κCond = κ := IsCondKernel.disintegrate + +/-! #### Existence of a disintegrating kernel in a countable space -/ + +section Countable +variable [Countable α] (κCond : α → Kernel β Ω) /-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. -A conditional kernel for `κ : Kernel α (β × Ω)` where `α` is countable and `Ω` is standard Borel. -/ -noncomputable -def condKernelCountable (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : Kernel (α × β) Ω where - toFun p := (κ p.1).condKernel p.2 + +A conditional kernel for `κ : Kernel α (β × Ω)` where `α` is countable and `Ω` is a measurable +space. -/ +noncomputable def condKernelCountable (h_atom : ∀ x y, x ∈ measurableAtom y → κCond x = κCond y) : + Kernel (α × β) Ω where + toFun p := κCond p.1 p.2 measurable' := by - change Measurable ((fun q : β × α ↦ (κ q.2).condKernel q.1) ∘ Prod.swap) - refine (measurable_from_prod_countable' (fun a ↦ ?_) ?_).comp measurable_swap - · exact (κ a).condKernel.measurable - · intro y y' x hy' - have : κ y' = κ y := by - ext s hs - exact mem_of_mem_measurableAtom hy' - (Kernel.measurable_coe κ hs (measurableSet_singleton (κ y s))) rfl - simp [this] - -lemma condKernelCountable_apply (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] (p : α × β) : - condKernelCountable κ p = (κ p.1).condKernel p.2 := rfl - -instance instIsMarkovKernelCondKernelCountable (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernelCountable κ) := - ⟨fun p ↦ (Measure.instIsMarkovKernelCondKernel (κ p.1)).isProbabilityMeasure p.2⟩ - -lemma compProd_fst_condKernelCountable (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernelCountable κ = κ := by + change Measurable ((fun q : β × α ↦ (κCond q.2) q.1) ∘ Prod.swap) + refine (measurable_from_prod_countable' (fun a ↦ (κCond a).measurable) ?_).comp measurable_swap + · intro x y hx hy + simpa using DFunLike.congr (h_atom _ _ hy) rfl + +lemma condKernelCountable_apply (h_atom : ∀ x y, x ∈ measurableAtom y → κCond x = κCond y) + (p : α × β) : condKernelCountable κCond h_atom p = κCond p.1 p.2 := rfl + +instance condKernelCountable.instIsMarkovKernel [∀ a, IsMarkovKernel (κCond a)] + (h_atom : ∀ x y, x ∈ measurableAtom y → κCond x = κCond y) : + IsMarkovKernel (condKernelCountable κCond h_atom) where + isProbabilityMeasure p := (‹∀ a, IsMarkovKernel (κCond a)› p.1).isProbabilityMeasure p.2 + +instance condKernelCountable.instIsCondKernel [∀ a, IsMarkovKernel (κCond a)] + (h_atom : ∀ x y, x ∈ measurableAtom y → κCond x = κCond y) (κ : Kernel α (β × Ω)) + [IsSFiniteKernel κ] [∀ a, (κ a).IsCondKernel (κCond a)] : + κ.IsCondKernel (condKernelCountable κCond h_atom) := by + constructor ext a s hs - have h := (κ a).compProd_fst_condKernel - conv_rhs => rw [← h] + conv_rhs => rw [← (κ a).disintegrate (κCond a)] simp_rw [compProd_apply _ _ _ hs, condKernelCountable_apply, Measure.compProd_apply hs] congr end Countable - -section CountableOrCountablyGenerated - -open Classical in - -/-- Conditional kernel of a kernel `κ : Kernel α (β × Ω)`: a Markov kernel such that -`fst κ ⊗ₖ condKernel κ = κ` (see `MeasureTheory.Measure.compProd_fst_condKernel`). -It exists whenever `Ω` is standard Borel and either `α` is countable -or `β` is countably generated. -/ -noncomputable -irreducible_def condKernel [h : CountableOrCountablyGenerated α β] - (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : - Kernel (α × β) Ω := - if hα : Countable α then condKernelCountable κ - else letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ - -/-- `condKernel κ` is a Markov kernel. -/ -instance instIsMarkovKernelCondKernel [CountableOrCountablyGenerated α β] - (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : - IsMarkovKernel (condKernel κ) := by - rw [condKernel_def] - split_ifs <;> infer_instance - -/-- **Disintegration** of finite kernels. -The composition-product of `fst κ` and `condKernel κ` is equal to `κ`. -/ -lemma compProd_fst_condKernel [hαβ : CountableOrCountablyGenerated α β] - (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : - fst κ ⊗ₖ condKernel κ = κ := by - rw [condKernel_def] - split_ifs with h - · exact compProd_fst_condKernelCountable κ - · have := hαβ.countableOrCountablyGenerated.resolve_left h - exact compProd_fst_condKernelBorel κ - -end CountableOrCountablyGenerated - end ProbabilityTheory.Kernel diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 6f717fa5a0d18..5f5de66ae8cd4 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.Basic +import Mathlib.Probability.Kernel.Disintegration.StandardBorel /-! # Lebesgue and Bochner integrals of conditional kernels @@ -38,14 +38,14 @@ variable [CountableOrCountablyGenerated α β] {κ : Kernel α (β × Ω)} [IsFi lemma lintegral_condKernel_mem (a : α) {s : Set (β × Ω)} (hs : MeasurableSet s) : ∫⁻ x, Kernel.condKernel κ (a, x) {y | (x, y) ∈ s} ∂(Kernel.fst κ a) = κ a s := by - conv_rhs => rw [← Kernel.compProd_fst_condKernel κ] + conv_rhs => rw [← κ.disintegrate κ.condKernel] simp_rw [Kernel.compProd_apply _ _ _ hs] lemma setLIntegral_condKernel_eq_measure_prod (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b in s, Kernel.condKernel κ (a, b) t ∂(Kernel.fst κ a) = κ a (s ×ˢ t) := by have : κ a (s ×ˢ t) = (Kernel.fst κ ⊗ₖ Kernel.condKernel κ) a (s ×ˢ t) := by - congr; exact (Kernel.compProd_fst_condKernel κ).symm + congr; exact (κ.disintegrate _).symm rw [this, Kernel.compProd_apply _ _ _ (hs.prod ht)] classical have : ∀ b, Kernel.condKernel κ (a, b) {c | (b, c) ∈ s ×ˢ t} @@ -60,14 +60,14 @@ alias set_lintegral_condKernel_eq_measure_prod := setLIntegral_condKernel_eq_mea lemma lintegral_condKernel (hf : Measurable f) (a : α) : ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫⁻ x, f x ∂(κ a) := by - conv_rhs => rw [← Kernel.compProd_fst_condKernel κ] + conv_rhs => rw [← κ.disintegrate κ.condKernel] rw [Kernel.lintegral_compProd _ _ _ hf] lemma setLIntegral_condKernel (hf : Measurable f) (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b in s, ∫⁻ ω in t, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫⁻ x in s ×ˢ t, f x ∂(κ a) := by - conv_rhs => rw [← Kernel.compProd_fst_condKernel κ] + conv_rhs => rw [← κ.disintegrate κ.condKernel] rw [Kernel.setLIntegral_compProd _ _ _ hf hs ht] @[deprecated (since := "2024-06-29")] @@ -102,21 +102,21 @@ lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_kernel_condKernel (a : (hf : AEStronglyMeasurable f (κ a)) : AEStronglyMeasurable (fun x ↦ ∫ y, f (x, y) ∂(Kernel.condKernel κ (a, x))) (Kernel.fst κ a) := by - rw [← Kernel.compProd_fst_condKernel κ] at hf + rw [← κ.disintegrate κ.condKernel] at hf exact AEStronglyMeasurable.integral_kernel_compProd hf lemma integral_condKernel (a : α) (hf : Integrable f (κ a)) : ∫ b, ∫ ω, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫ x, f x ∂(κ a) := by - conv_rhs => rw [← Kernel.compProd_fst_condKernel κ] - rw [← Kernel.compProd_fst_condKernel κ] at hf + conv_rhs => rw [← κ.disintegrate κ.condKernel] + rw [← κ.disintegrate κ.condKernel] at hf rw [integral_compProd hf] lemma setIntegral_condKernel (a : α) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) (κ a)) : ∫ b in s, ∫ ω in t, f (b, ω) ∂(Kernel.condKernel κ (a, b)) ∂(Kernel.fst κ a) = ∫ x in s ×ˢ t, f x ∂(κ a) := by - conv_rhs => rw [← Kernel.compProd_fst_condKernel κ] - rw [← Kernel.compProd_fst_condKernel κ] at hf + conv_rhs => rw [← κ.disintegrate κ.condKernel] + rw [← κ.disintegrate κ.condKernel] at hf rw [setIntegral_compProd hs ht hf] @[deprecated (since := "2024-04-17")] @@ -156,7 +156,7 @@ variable [CountableOrCountablyGenerated α β] {ρ : Measure (β × Ω)} [IsFini lemma lintegral_condKernel_mem {s : Set (β × Ω)} (hs : MeasurableSet s) : ∫⁻ x, ρ.condKernel x {y | (x, y) ∈ s} ∂ρ.fst = ρ s := by - conv_rhs => rw [← compProd_fst_condKernel ρ] + conv_rhs => rw [← ρ.disintegrate ρ.condKernel] simp_rw [compProd_apply hs] rfl @@ -164,7 +164,7 @@ lemma setLIntegral_condKernel_eq_measure_prod {s : Set β} (hs : MeasurableSet s (ht : MeasurableSet t) : ∫⁻ b in s, ρ.condKernel b t ∂ρ.fst = ρ (s ×ˢ t) := by have : ρ (s ×ˢ t) = (ρ.fst ⊗ₘ ρ.condKernel) (s ×ˢ t) := by - congr; exact (compProd_fst_condKernel ρ).symm + congr; exact (ρ.disintegrate _).symm rw [this, compProd_apply (hs.prod ht)] classical have : ∀ b, ρ.condKernel b (Prod.mk b ⁻¹' s ×ˢ t) @@ -179,14 +179,14 @@ alias set_lintegral_condKernel_eq_measure_prod := setLIntegral_condKernel_eq_mea lemma lintegral_condKernel (hf : Measurable f) : ∫⁻ b, ∫⁻ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫⁻ x, f x ∂ρ := by - conv_rhs => rw [← compProd_fst_condKernel ρ] + conv_rhs => rw [← ρ.disintegrate ρ.condKernel] rw [lintegral_compProd hf] lemma setLIntegral_condKernel (hf : Measurable f) {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) : ∫⁻ b in s, ∫⁻ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫⁻ x in s ×ˢ t, f x ∂ρ := by - conv_rhs => rw [← compProd_fst_condKernel ρ] + conv_rhs => rw [← ρ.disintegrate ρ.condKernel] rw [setLIntegral_compProd hf hs ht] @[deprecated (since := "2024-06-29")] @@ -220,20 +220,20 @@ variable {ρ : Measure (β × Ω)} [IsFiniteMeasure ρ] lemma _root_.MeasureTheory.AEStronglyMeasurable.integral_condKernel (hf : AEStronglyMeasurable f ρ) : AEStronglyMeasurable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := by - rw [← ρ.compProd_fst_condKernel] at hf + rw [← ρ.disintegrate ρ.condKernel] at hf exact AEStronglyMeasurable.integral_kernel_compProd hf lemma integral_condKernel (hf : Integrable f ρ) : ∫ b, ∫ ω, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x, f x ∂ρ := by - conv_rhs => rw [← compProd_fst_condKernel ρ] - rw [← compProd_fst_condKernel ρ] at hf + conv_rhs => rw [← ρ.disintegrate ρ.condKernel] + rw [← ρ.disintegrate ρ.condKernel] at hf rw [integral_compProd hf] lemma setIntegral_condKernel {s : Set β} (hs : MeasurableSet s) {t : Set Ω} (ht : MeasurableSet t) (hf : IntegrableOn f (s ×ˢ t) ρ) : ∫ b in s, ∫ ω in t, f (b, ω) ∂(ρ.condKernel b) ∂ρ.fst = ∫ x in s ×ˢ t, f x ∂ρ := by - conv_rhs => rw [← compProd_fst_condKernel ρ] - rw [← compProd_fst_condKernel ρ] at hf + conv_rhs => rw [← ρ.disintegrate ρ.condKernel] + rw [← ρ.disintegrate ρ.condKernel] at hf rw [setIntegral_compProd hs ht hf] @[deprecated (since := "2024-04-17")] @@ -275,8 +275,8 @@ theorem AEStronglyMeasurable.ae_integrable_condKernel_iff {f : α × Ω → F} (hf : AEStronglyMeasurable f ρ) : (∀ᵐ a ∂ρ.fst, Integrable (fun ω ↦ f (a, ω)) (ρ.condKernel a)) ∧ Integrable (fun a ↦ ∫ ω, ‖f (a, ω)‖ ∂ρ.condKernel a) ρ.fst ↔ Integrable f ρ := by - rw [← ρ.compProd_fst_condKernel] at hf - conv_rhs => rw [← ρ.compProd_fst_condKernel] + rw [← ρ.disintegrate ρ.condKernel] at hf + conv_rhs => rw [← ρ.disintegrate ρ.condKernel] rw [Measure.integrable_compProd_iff hf] theorem Integrable.condKernel_ae {f : α × Ω → F} (hf_int : Integrable f ρ) : diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 17a12848da76a..54e2234837003 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.Data.Complex.Abs -import Mathlib.MeasureTheory.Constructions.Polish import Mathlib.MeasureTheory.Measure.GiryMonad import Mathlib.MeasureTheory.Measure.Stieltjes import Mathlib.Analysis.Normed.Order.Lattice diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean new file mode 100644 index 0000000000000..0a3aec0c493b1 --- /dev/null +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -0,0 +1,471 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.Probability.Kernel.MeasureCompProd +import Mathlib.Probability.Kernel.Disintegration.Basic +import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.Density +import Mathlib.Probability.Kernel.Disintegration.CdfToKernel +import Mathlib.MeasureTheory.Constructions.Polish.EmbeddingReal + +/-! +# Existence of disintegration of measures and kernels for standard Borel spaces + +Let `κ : Kernel α (β × Ω)` be a finite kernel, where `Ω` is a standard Borel space. Then if `α` is +countable or `β` has a countably generated σ-algebra (for example if it is standard Borel), then +there exists a `Kernel (α × β) Ω` called conditional kernel and denoted by `condKernel κ` such that +`κ = fst κ ⊗ₖ condKernel κ`. +We also define a conditional kernel for a measure `ρ : Measure (β × Ω)`, where `Ω` is a standard +Borel space. This is a `Kernel β Ω` denoted by `ρ.condKernel` such that `ρ = ρ.fst ⊗ₘ ρ.condKernel`. + +In order to obtain a disintegration for any standard Borel space `Ω`, we use that these spaces embed +measurably into `ℝ`: it then suffices to define a suitable kernel for `Ω = ℝ`. + +For `κ : Kernel α (β × ℝ)`, the construction of the conditional kernel proceeds as follows: +* Build a measurable function `f : (α × β) → ℚ → ℝ` such that for all measurable sets + `s` and all `q : ℚ`, `∫ x in s, f (a, x) q ∂(Kernel.fst κ a) = (κ a (s ×ˢ Iic (q : ℝ))).toReal`. + We restrict to `ℚ` here to be able to prove the measurability. +* Extend that function to `(α × β) → StieltjesFunction`. See the file `MeasurableStieltjes.lean`. +* Finally obtain from the measurable Stieltjes function a measure on `ℝ` for each element of `α × β` + in a measurable way: we have obtained a `Kernel (α × β) ℝ`. + See the file `CdfToKernel.lean` for that step. + +The first step (building the measurable function on `ℚ`) is done differently depending on whether +`α` is countable or not. +* If `α` is countable, we can provide for each `a : α` a function `f : β → ℚ → ℝ` and proceed as + above to obtain a `Kernel β ℝ`. Since `α` is countable, measurability is not an issue and we can + put those together into a `Kernel (α × β) ℝ`. The construction of that `f` is done in + the `CondCdf.lean` file. +* If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function + `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a + countably generated σ-algebra (this is the case in particular for standard Borel spaces). + See the file `Density.lean`. + +The conditional kernel is defined under the typeclass assumption +`CountableOrCountablyGenerated α β`, which encodes the property +`Countable α ∨ CountablyGenerated β`. + +Properties of integrals involving `condKernel` are collated in the file `Integral.lean`. +The conditional kernel is unique (almost everywhere w.r.t. `fst κ`): this is proved in the file +`Unique.lean`. + +## Main definitions + +* `ProbabilityTheory.Kernel.condKernel κ : Kernel (α × β) Ω`: conditional kernel described above. +* `MeasureTheory.Measure.condKernel ρ : Kernel β Ω`: conditional kernel of a measure. + +## Main statements + +* `ProbabilityTheory.Kernel.compProd_fst_condKernel`: `fst κ ⊗ₖ condKernel κ = κ` +* `MeasureTheory.Measure.compProd_fst_condKernel`: `ρ.fst ⊗ₘ ρ.condKernel = ρ` +-/ + +open MeasureTheory Set Filter MeasurableSpace + +open scoped ENNReal MeasureTheory Topology ProbabilityTheory + +namespace ProbabilityTheory.Kernel + +variable {α β γ Ω : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} + {mγ : MeasurableSpace γ} [MeasurableSpace.CountablyGenerated γ] + [MeasurableSpace Ω] [StandardBorelSpace Ω] [Nonempty Ω] + +section Real + +/-! ### Disintegration of kernels from `α` to `γ × ℝ` for countably generated `γ` -/ + +lemma isRatCondKernelCDFAux_density_Iic (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatCondKernelCDFAux (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) where + measurable := measurable_pi_iff.mpr fun _ ↦ measurable_density κ (fst κ) measurableSet_Iic + mono' a q r hqr := + ae_of_all _ fun c ↦ density_mono_set le_rfl a c (Iic_subset_Iic.mpr (by exact_mod_cast hqr)) + nonneg' a q := ae_of_all _ fun c ↦ density_nonneg le_rfl _ _ _ + le_one' a q := ae_of_all _ fun c ↦ density_le_one le_rfl _ _ _ + tendsto_integral_of_antitone a s hs_anti hs_tendsto := by + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + refine tendsto_integral_density_of_antitone le_rfl a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) + · refine fun i j hij ↦ Iic_subset_Iic.mpr ?_ + exact mod_cast hs_anti hij + · ext x + simp only [mem_iInter, mem_Iic, mem_empty_iff_false, iff_false, not_forall, not_le, s'] + rw [tendsto_atTop_atBot] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_lt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, lt_of_le_of_lt ?_ hq⟩ + exact mod_cast hi i le_rfl + tendsto_integral_of_monotone a s hs_mono hs_tendsto := by + rw [fst_apply' _ _ MeasurableSet.univ] + let s' : ℕ → Set ℝ := fun n ↦ Iic (s n) + refine tendsto_integral_density_of_monotone (le_rfl : fst κ ≤ fst κ) + a s' ?_ ?_ (fun _ ↦ measurableSet_Iic) + · exact fun i j hij ↦ Iic_subset_Iic.mpr (by exact mod_cast hs_mono hij) + · ext x + simp only [mem_iUnion, mem_Iic, mem_univ, iff_true] + rw [tendsto_atTop_atTop] at hs_tendsto + have ⟨q, hq⟩ := exists_rat_gt x + obtain ⟨i, hi⟩ := hs_tendsto q + refine ⟨i, hq.le.trans ?_⟩ + exact mod_cast hi i le_rfl + integrable a q := integrable_density le_rfl a measurableSet_Iic + setIntegral a A hA q := setIntegral_density le_rfl a measurableSet_Iic hA + +/-- Taking the kernel density of intervals `Iic q` for `q : ℚ` gives a function with the property +`isRatCondKernelCDF`. -/ +lemma isRatCondKernelCDF_density_Iic (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsRatCondKernelCDF (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) κ (fst κ) := + (isRatCondKernelCDFAux_density_Iic κ).isRatCondKernelCDF + +/-- The conditional kernel CDF of a kernel `κ : Kernel α (γ × ℝ)`, where `γ` is countably generated. +-/ +noncomputable +def condKernelCDF (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : α × γ → StieltjesFunction := + stieltjesOfMeasurableRat (fun (p : α × γ) q ↦ density κ (fst κ) p.1 p.2 (Iic q)) + (isRatCondKernelCDF_density_Iic κ).measurable + +lemma isCondKernelCDF_condKernelCDF (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsCondKernelCDF (condKernelCDF κ) κ (fst κ) := + isCondKernelCDF_stieltjesOfMeasurableRat (isRatCondKernelCDF_density_Iic κ) + +/-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. +A conditional kernel for `κ : Kernel α (γ × ℝ)` where `γ` is countably generated. -/ +noncomputable +def condKernelReal (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : Kernel (α × γ) ℝ := + (isCondKernelCDF_condKernelCDF κ).toKernel + +instance instIsMarkovKernelCondKernelReal (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelReal κ) := by + rw [condKernelReal] + infer_instance + +lemma compProd_fst_condKernelReal (κ : Kernel α (γ × ℝ)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelReal κ = κ := by + rw [condKernelReal, compProd_toKernel] + +/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and +`ProbabilityTheory.Kernel.condKernel`. +A conditional kernel for `κ : Kernel Unit (α × ℝ)`. -/ +noncomputable +def condKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : Kernel (Unit × α) ℝ := + (isCondKernelCDF_condCDF (κ ())).toKernel + +instance instIsMarkovKernelCondKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelUnitReal κ) := by + rw [condKernelUnitReal] + infer_instance + +instance condKernelUnitReal.instIsCondKernel (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + κ.IsCondKernel κ.condKernelUnitReal where + disintegrate := by rw [condKernelUnitReal, compProd_toKernel]; ext; simp + +@[deprecated disintegrate (since := "2024-07-26")] +lemma compProd_fst_condKernelUnitReal (κ : Kernel Unit (α × ℝ)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelUnitReal κ = κ := disintegrate _ _ + +end Real + +section BorelSnd + +/-! ### Disintegration of kernels on standard Borel spaces + +Since every standard Borel space embeds measurably into `ℝ`, we can generalize a disintegration +property on `ℝ` to all these spaces. -/ + +open Classical in +/-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. +A Borel space `Ω` embeds measurably into `ℝ` (with embedding `e`), hence we can get a `Kernel α Ω` +from a `Kernel α ℝ` by taking the comap by `e`. +Here we take the comap of a modification of `η : Kernel α ℝ`, useful when `η a` is a probability +measure with all its mass on `range e` almost everywhere with respect to some measure and we want to +ensure that the comap is a Markov kernel. +We thus take the comap by `e` of a kernel defined piecewise: `η` when +`η a (range (embeddingReal Ω))ᶜ = 0`, and an arbitrary deterministic kernel otherwise. -/ +noncomputable +def borelMarkovFromReal (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] + (η : Kernel α ℝ) : + Kernel α Ω := + have he := measurableEmbedding_embeddingReal Ω + let x₀ := (range_nonempty (embeddingReal Ω)).choose + comapRight + (piecewise ((Kernel.measurable_coe η he.measurableSet_range.compl) (measurableSet_singleton 0) : + MeasurableSet {a | η a (range (embeddingReal Ω))ᶜ = 0}) + η (deterministic (fun _ ↦ x₀) measurable_const)) he + +lemma borelMarkovFromReal_apply (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] + (η : Kernel α ℝ) (a : α) : + borelMarkovFromReal Ω η a + = if η a (range (embeddingReal Ω))ᶜ = 0 then (η a).comap (embeddingReal Ω) + else (Measure.dirac (range_nonempty (embeddingReal Ω)).choose).comap (embeddingReal Ω) := by + classical + rw [borelMarkovFromReal, comapRight_apply, piecewise_apply, deterministic_apply] + simp only [mem_preimage, mem_singleton_iff] + split_ifs <;> rfl + +lemma borelMarkovFromReal_apply' (Ω : Type*) [Nonempty Ω] [MeasurableSpace Ω] [StandardBorelSpace Ω] + (η : Kernel α ℝ) (a : α) {s : Set Ω} (hs : MeasurableSet s) : + borelMarkovFromReal Ω η a s + = if η a (range (embeddingReal Ω))ᶜ = 0 then η a (embeddingReal Ω '' s) + else (embeddingReal Ω '' s).indicator 1 (range_nonempty (embeddingReal Ω)).choose := by + have he := measurableEmbedding_embeddingReal Ω + rw [borelMarkovFromReal_apply] + split_ifs with h + · rw [Measure.comap_apply _ he.injective he.measurableSet_image' _ hs] + · rw [Measure.comap_apply _ he.injective he.measurableSet_image' _ hs, Measure.dirac_apply] + +/-- When `η` is an s-finite kernel, `borelMarkovFromReal Ω η` is an s-finite kernel. -/ +instance instIsSFiniteKernelBorelMarkovFromReal (η : Kernel α ℝ) [IsSFiniteKernel η] : + IsSFiniteKernel (borelMarkovFromReal Ω η) := + IsSFiniteKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) + +/-- When `η` is a finite kernel, `borelMarkovFromReal Ω η` is a finite kernel. -/ +instance instIsFiniteKernelBorelMarkovFromReal (η : Kernel α ℝ) [IsFiniteKernel η] : + IsFiniteKernel (borelMarkovFromReal Ω η) := + IsFiniteKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) + +/-- When `η` is a Markov kernel, `borelMarkovFromReal Ω η` is a Markov kernel. -/ +instance instIsMarkovKernelBorelMarkovFromReal (η : Kernel α ℝ) [IsMarkovKernel η] : + IsMarkovKernel (borelMarkovFromReal Ω η) := by + refine IsMarkovKernel.comapRight _ (measurableEmbedding_embeddingReal Ω) (fun a ↦ ?_) + classical + rw [piecewise_apply] + split_ifs with h + · rwa [← prob_compl_eq_zero_iff (measurableEmbedding_embeddingReal Ω).measurableSet_range] + · rw [deterministic_apply] + simp [(range_nonempty (embeddingReal Ω)).choose_spec] + +/-- For `κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable)`, the +hypothesis `hη` is `fst κ' ⊗ₖ η = κ'`. The conclusion of the lemma is +`fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) _`. -/ +lemma compProd_fst_borelMarkovFromReal_eq_comapRight_compProd + (κ : Kernel α (β × Ω)) [IsSFiniteKernel κ] (η : Kernel (α × β) ℝ) [IsSFiniteKernel η] + (hη : (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable))) ⊗ₖ η + = map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) : + fst κ ⊗ₖ borelMarkovFromReal Ω η + = comapRight (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) ⊗ₖ η) + (MeasurableEmbedding.id.prod_mk (measurableEmbedding_embeddingReal Ω)) := by + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) + have hη' : fst κ' ⊗ₖ η = κ' := hη + have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := + MeasurableEmbedding.id.prod_mk he + change fst κ ⊗ₖ borelMarkovFromReal Ω η = comapRight (fst κ' ⊗ₖ η) h_prod_embed + rw [comapRight_compProd_id_prod _ _ he] + have h_fst : fst κ' = fst κ := by + ext a u + unfold_let κ' + rw [fst_apply, map_apply, Measure.map_map measurable_fst h_prod_embed.measurable, fst_apply] + congr + rw [h_fst] + ext a t ht : 2 + simp_rw [compProd_apply _ _ _ ht] + refine lintegral_congr_ae ?_ + have h_ae : ∀ᵐ t ∂(fst κ a), (a, t) ∈ {p : α × β | η p (range e)ᶜ = 0} := by + rw [← h_fst] + have h_compProd : κ' a (univ ×ˢ range e)ᶜ = 0 := by + unfold_let κ' + rw [map_apply'] + swap; · exact (MeasurableSet.univ.prod he.measurableSet_range).compl + suffices Prod.map id e ⁻¹' (univ ×ˢ range e)ᶜ = ∅ by rw [this]; simp + ext x + simp + rw [← hη', compProd_null] at h_compProd + swap; · exact (MeasurableSet.univ.prod he.measurableSet_range).compl + simp only [preimage_compl, mem_univ, mk_preimage_prod_right] at h_compProd + exact h_compProd + filter_upwards [h_ae] with a ha + rw [borelMarkovFromReal, comapRight_apply', comapRight_apply'] + rotate_left + · exact measurable_prod_mk_left ht + · exact measurable_prod_mk_left ht + classical + rw [piecewise_apply, if_pos] + exact ha + +/-- For `κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable)`, the +hypothesis `hη` is `fst κ' ⊗ₖ η = κ'`. With that hypothesis, +`fst κ ⊗ₖ borelMarkovFromReal κ η = κ`.-/ +lemma compProd_fst_borelMarkovFromReal (κ : Kernel α (β × Ω)) [IsSFiniteKernel κ] + (η : Kernel (α × β) ℝ) [IsSFiniteKernel η] + (hη : (fst (map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable))) ⊗ₖ η + = map κ (Prod.map (id : β → β) (embeddingReal Ω)) + (measurable_id.prod_map (measurableEmbedding_embeddingReal Ω).measurable)) : + fst κ ⊗ₖ borelMarkovFromReal Ω η = κ := by + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : β → β) e) (measurable_id.prod_map he.measurable) + have hη' : fst κ' ⊗ₖ η = κ' := hη + have h_prod_embed : MeasurableEmbedding (Prod.map (id : β → β) e) := + MeasurableEmbedding.id.prod_mk he + have : κ = comapRight κ' h_prod_embed := by + ext c t : 2 + unfold_let κ' + rw [comapRight_apply, map_apply, h_prod_embed.comap_map] + conv_rhs => rw [this, ← hη'] + exact compProd_fst_borelMarkovFromReal_eq_comapRight_compProd κ η hη + +end BorelSnd + +section CountablyGenerated + +open ProbabilityTheory.Kernel + +/-- Auxiliary definition for `ProbabilityTheory.Kernel.condKernel`. +A conditional kernel for `κ : Kernel α (γ × Ω)` where `γ` is countably generated and `Ω` is +standard Borel. -/ +noncomputable +def condKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : Kernel (α × γ) Ω := + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : γ → γ) e) (measurable_id.prod_map he.measurable) + borelMarkovFromReal Ω (condKernelReal κ') + +instance instIsMarkovKernelCondKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : + IsMarkovKernel (condKernelBorel κ) := by + rw [condKernelBorel] + infer_instance + +instance condKernelBorel.instIsCondKernel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : + κ.IsCondKernel κ.condKernelBorel where + disintegrate := by + rw [condKernelBorel, compProd_fst_borelMarkovFromReal _ _ (compProd_fst_condKernelReal _)] + +@[deprecated disintegrate (since := "2024-07-26")] +lemma compProd_fst_condKernelBorel (κ : Kernel α (γ × Ω)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelBorel κ = κ := disintegrate _ _ + +end CountablyGenerated + +section Unit +variable (κ : Kernel Unit (α × Ω)) [IsFiniteKernel κ] + +/-- Auxiliary definition for `MeasureTheory.Measure.condKernel` and +`ProbabilityTheory.Kernel.condKernel`. +A conditional kernel for `κ : Kernel Unit (α × Ω)` where `Ω` is standard Borel. -/ +noncomputable +def condKernelUnitBorel : Kernel (Unit × α) Ω := + let e := embeddingReal Ω + let he := measurableEmbedding_embeddingReal Ω + let κ' := map κ (Prod.map (id : α → α) e) (measurable_id.prod_map he.measurable) + borelMarkovFromReal Ω (condKernelUnitReal κ') + +instance instIsMarkovKernelCondKernelUnitBorel : IsMarkovKernel κ.condKernelUnitBorel := by + rw [condKernelUnitBorel] + infer_instance + +instance condKernelUnitBorel.instIsCondKernel : κ.IsCondKernel κ.condKernelUnitBorel where + disintegrate := by + rw [condKernelUnitBorel, compProd_fst_borelMarkovFromReal _ _ (disintegrate _ _)] + +@[deprecated disintegrate (since := "2024-07-26")] +lemma compProd_fst_condKernelUnitBorel (κ : Kernel Unit (α × Ω)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelUnitBorel κ = κ := disintegrate _ _ + +end Unit + +section Measure + +variable {ρ : Measure (α × Ω)} [IsFiniteMeasure ρ] + +/-- Conditional kernel of a measure on a product space: a Markov kernel such that +`ρ = ρ.fst ⊗ₘ ρ.condKernel` (see `MeasureTheory.Measure.compProd_fst_condKernel`). -/ +noncomputable +irreducible_def _root_.MeasureTheory.Measure.condKernel (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : + Kernel α Ω := + comap (condKernelUnitBorel (const Unit ρ)) (fun a ↦ ((), a)) measurable_prod_mk_left + +lemma _root_.MeasureTheory.Measure.condKernel_apply (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] + (a : α) : + ρ.condKernel a = condKernelUnitBorel (const Unit ρ) ((), a) := by + rw [Measure.condKernel]; rfl + +instance _root_.MeasureTheory.Measure.condKernel.instIsCondKernel (ρ : Measure (α × Ω)) + [IsFiniteMeasure ρ] : ρ.IsCondKernel ρ.condKernel where + disintegrate := by + have h1 : const Unit (Measure.fst ρ) = fst (const Unit ρ) := by + ext + simp only [fst_apply, Measure.fst, const_apply] + have h2 : prodMkLeft Unit (Measure.condKernel ρ) = condKernelUnitBorel (const Unit ρ) := by + ext + simp only [prodMkLeft_apply, Measure.condKernel_apply] + rw [Measure.compProd, h1, h2, disintegrate] + simp + +instance _root_.MeasureTheory.Measure.instIsMarkovKernelCondKernel + (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : IsMarkovKernel ρ.condKernel := by + rw [Measure.condKernel] + infer_instance + +/-- **Disintegration** of finite product measures on `α × Ω`, where `Ω` is standard Borel. Such a +measure can be written as the composition-product of `ρ.fst` (marginal measure over `α`) and +a Markov kernel from `α` to `Ω`. We call that Markov kernel `ρ.condKernel`. -/ +@[deprecated Measure.disintegrate (since := "2024-07-24")] +lemma _root_.MeasureTheory.Measure.compProd_fst_condKernel + (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] : + ρ.fst ⊗ₘ ρ.condKernel = ρ := ρ.disintegrate ρ.condKernel + +set_option linter.unusedVariables false in +/-- Auxiliary lemma for `condKernel_apply_of_ne_zero`. -/ +@[deprecated Measure.IsCondKernel.apply_of_ne_zero (since := "2024-07-24"), nolint unusedArguments] +lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero_of_measurableSet + [MeasurableSingletonClass α] {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : + ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := + Measure.IsCondKernel.apply_of_ne_zero _ _ hx _ + +/-- If the singleton `{x}` has non-zero mass for `ρ.fst`, then for all `s : Set Ω`, +`ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s)` . -/ +lemma _root_.MeasureTheory.Measure.condKernel_apply_of_ne_zero [MeasurableSingletonClass α] + {x : α} (hx : ρ.fst {x} ≠ 0) (s : Set Ω) : + ρ.condKernel x s = (ρ.fst {x})⁻¹ * ρ ({x} ×ˢ s) := + Measure.IsCondKernel.apply_of_ne_zero _ _ hx _ + +end Measure + +section Countable + +variable [Countable α] + +@[deprecated disintegrate (since := "2024-07-24")] +lemma compProd_fst_condKernelCountable (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] : + fst κ ⊗ₖ condKernelCountable (fun a ↦ (κ a).condKernel) + (fun x y h ↦ by simp [apply_congr_of_mem_measurableAtom _ h]) = κ := disintegrate _ _ + +end Countable + +section CountableOrCountablyGenerated +variable [h : CountableOrCountablyGenerated α β] (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] + +open Classical in + +/-- Conditional kernel of a kernel `κ : Kernel α (β × Ω)`: a Markov kernel such that +`fst κ ⊗ₖ condKernel κ = κ` (see `MeasureTheory.Measure.compProd_fst_condKernel`). +It exists whenever `Ω` is standard Borel and either `α` is countable +or `β` is countably generated. -/ +noncomputable +irreducible_def condKernel : Kernel (α × β) Ω := + if hα : Countable α then + condKernelCountable (fun a ↦ (κ a).condKernel) + fun x y h ↦ by simp [apply_congr_of_mem_measurableAtom _ h] + else letI := h.countableOrCountablyGenerated.resolve_left hα; condKernelBorel κ + +/-- `condKernel κ` is a Markov kernel. -/ +instance instIsMarkovKernelCondKernel : IsMarkovKernel (condKernel κ) := by + rw [condKernel_def] + split_ifs <;> infer_instance + +instance condKernel.instIsCondKernel : κ.IsCondKernel κ.condKernel where + disintegrate := by rw [condKernel_def]; split_ifs with hα <;> exact disintegrate _ _ + +/-- **Disintegration** of finite kernels. +The composition-product of `fst κ` and `condKernel κ` is equal to `κ`. -/ +@[deprecated Kernel.disintegrate (since := "2024-07-26")] +lemma compProd_fst_condKernel : fst κ ⊗ₖ condKernel κ = κ := κ.disintegrate κ.condKernel + +end CountableOrCountablyGenerated + +end ProbabilityTheory.Kernel diff --git a/Mathlib/Probability/Kernel/Disintegration/Unique.lean b/Mathlib/Probability/Kernel/Disintegration/Unique.lean index f0e64036ba09c..a55d8224d6a44 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Unique.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Unique.lean @@ -115,7 +115,7 @@ theorem eq_condKernel_of_measure_eq_compProd (κ : Kernel α Ω) [IsFiniteKernel filter_upwards [heq] with x hx s hs rw [← hx, Kernel.map_apply, Measure.map_apply hf.measurable hs] ext s hs - conv_lhs => rw [← ρ.compProd_fst_condKernel] + conv_lhs => rw [← ρ.disintegrate ρ.condKernel] rw [Measure.compProd_apply hs, Measure.map_apply (measurable_id.prod_map hf.measurable) hs, Measure.compProd_apply] · congr with a @@ -146,7 +146,7 @@ is equal to the conditional kernel of the measure `κ a` applied to `b`. -/ lemma Kernel.condKernel_apply_eq_condKernel [CountableOrCountablyGenerated α β] (κ : Kernel α (β × Ω)) [IsFiniteKernel κ] (a : α) : (fun b ↦ Kernel.condKernel κ (a, b)) =ᵐ[Kernel.fst κ a] (κ a).condKernel := - Kernel.apply_eq_measure_condKernel_of_compProd_eq (compProd_fst_condKernel κ) a + Kernel.apply_eq_measure_condKernel_of_compProd_eq (κ.disintegrate _) a lemma condKernel_const [CountableOrCountablyGenerated α β] (ρ : Measure (β × Ω)) [IsFiniteMeasure ρ] (a : α) : diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index 5ab039d2fccc9..99269703362e9 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -40,6 +40,16 @@ def compProd (μ : Measure α) (κ : Kernel α β) : Measure (α × β) := @[inherit_doc] scoped[ProbabilityTheory] infixl:100 " ⊗ₘ " => MeasureTheory.Measure.compProd +lemma compProd_of_not_sfinite (μ : Measure α) (κ : Kernel α β) (h : ¬ SFinite μ) : + μ ⊗ₘ κ = 0 := by + rw [compProd, Kernel.compProd_of_not_isSFiniteKernel_left, Kernel.zero_apply] + rwa [Kernel.isSFiniteKernel_const] + +lemma compProd_of_not_isSFiniteKernel (μ : Measure α) (κ : Kernel α β) (h : ¬ IsSFiniteKernel κ) : + μ ⊗ₘ κ = 0 := by + rw [compProd, Kernel.compProd_of_not_isSFiniteKernel_right, Kernel.zero_apply] + rwa [Kernel.isSFiniteKernel_prodMkLeft_unit] + @[simp] lemma compProd_zero_left (κ : Kernel α β) : (0 : Measure α) ⊗ₘ κ = 0 := by simp [compProd] @[simp] lemma compProd_zero_right (μ : Measure α) : μ ⊗ₘ (0 : Kernel α β) = 0 := by simp [compProd] diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 45893abe190ed..b369b661f36e1 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -3,9 +3,9 @@ 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.Probability.Martingale.Upcrossing +import Mathlib.MeasureTheory.Constructions.Polish.Basic import Mathlib.MeasureTheory.Function.UniformIntegrable -import Mathlib.MeasureTheory.Constructions.Polish +import Mathlib.Probability.Martingale.Upcrossing /-! diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index 474e4258f7a4a..808459c99ce3a 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -520,7 +520,7 @@ theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = d /- Porting note (#11039): broken proof was simpa [← @intCast_smul k, simplicial_object.δ] -/ simp_rw [alternatingFaceMapComplex_obj_d, AlternatingFaceMapComplex.objD, SimplicialObject.δ, - Functor.comp_map, ← intCast_smul (k := k) ((-1) ^ _ : ℤ), Int.cast_pow, Int.cast_neg, + Functor.comp_map, ← Int.cast_smul_eq_nsmul k ((-1) ^ _ : ℤ), Int.cast_pow, Int.cast_neg, Int.cast_one, Action.sum_hom, Action.smul_hom, Rep.linearization_map_hom] rw [LinearMap.coeFn_sum, Fintype.sum_apply] erw [d_of (k := k) x] diff --git a/Mathlib/RepresentationTheory/Invariants.lean b/Mathlib/RepresentationTheory/Invariants.lean index e07011cf6be72..fa62df0a43948 100644 --- a/Mathlib/RepresentationTheory/Invariants.lean +++ b/Mathlib/RepresentationTheory/Invariants.lean @@ -101,7 +101,7 @@ theorem averageMap_invariant (v : V) : averageMap ρ v ∈ invariants ρ := fun -/ theorem averageMap_id (v : V) (hv : v ∈ invariants ρ) : averageMap ρ v = v := by rw [mem_invariants] at hv - simp [average, map_sum, hv, Finset.card_univ, nsmul_eq_smul_cast k _ v, smul_smul] + simp [average, map_sum, hv, Finset.card_univ, ← Nat.cast_smul_eq_nsmul k _ v, smul_smul] theorem isProj_averageMap : LinearMap.IsProj ρ.invariants ρ.averageMap := ⟨ρ.averageMap_invariant, ρ.averageMap_id⟩ diff --git a/Mathlib/RepresentationTheory/Maschke.lean b/Mathlib/RepresentationTheory/Maschke.lean index 9944c416493ba..caeae09e6064a 100644 --- a/Mathlib/RepresentationTheory/Maschke.lean +++ b/Mathlib/RepresentationTheory/Maschke.lean @@ -122,7 +122,7 @@ theorem equivariantProjection_apply (v : W) : theorem equivariantProjection_condition (v : V) : (π.equivariantProjection G) (i v) = v := by rw [equivariantProjection_apply] simp only [conjugate_i π i h] - rw [Finset.sum_const, Finset.card_univ, nsmul_eq_smul_cast k, smul_smul, + rw [Finset.sum_const, Finset.card_univ, ← Nat.cast_smul_eq_nsmul k, smul_smul, Invertible.invOf_mul_self, one_smul] end diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index e862f5746f25e..7d07464cc61ca 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -89,7 +89,11 @@ section PiFintype In this section we show that `ofTensorProduct` is an isomorphism if `M = R^n`. -/ -variable (ι : Type*) [Fintype ι] [DecidableEq ι] +variable (ι : Type*) + +section DecidableEq + +variable [Fintype ι] [DecidableEq ι] private lemma piEquivOfFintype_comp_ofTensorProduct_eq : piEquivOfFintype I (fun _ : ι ↦ R) ∘ₗ ofTensorProduct I (ι → R) = @@ -134,10 +138,14 @@ def ofTensorProductEquivOfPiFintype : (ofTensorProduct_comp_ofTensorProductInvOfPiFintype I ι) (ofTensorProductInvOfPiFintype_comp_ofTensorProduct I ι) +end DecidableEq + /-- If `M = R^ι`, `ofTensorProduct` is bijective. -/ -lemma ofTensorProduct_bijective_of_pi_of_fintype : - Function.Bijective (ofTensorProduct I (ι → R)) := - EquivLike.bijective (ofTensorProductEquivOfPiFintype I ι) +lemma ofTensorProduct_bijective_of_pi_of_fintype [Finite ι] : + Function.Bijective (ofTensorProduct I (ι → R)) := by + classical + cases nonempty_fintype ι + exact EquivLike.bijective (ofTensorProductEquivOfPiFintype I ι) end PiFintype diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index e34934b8e8a0b..cbe6c2575aaa0 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -299,7 +299,7 @@ theorem sum_comp_sumInv : sum I M ∘ₗ sumInv I M = LinearMap.id := by ext f n simp only [LinearMap.coe_comp, Function.comp_apply, LinearMap.id_coe, id_eq, mk_apply_coe, Submodule.mkQ_apply] - rw [← DirectSum.sum_univ_of _ (((sumInv I M) ((AdicCompletion.mk I (⨁ (j : ι), M j)) f)))] + rw [← DirectSum.sum_univ_of (((sumInv I M) ((AdicCompletion.mk I (⨁ (j : ι), M j)) f)))] simp only [sumInv_apply, map_mk, map_sum, sum_of, val_sum, mk_apply_coe, AdicCauchySequence.map_apply_coe, Submodule.mkQ_apply] simp only [← Submodule.mkQ_apply, ← map_sum] diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index d04d103ccf0b4..9a5254f9a12da 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -10,7 +10,7 @@ import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.Prod import Mathlib.LinearAlgebra.Finsupp import Mathlib.LinearAlgebra.Prod - +import Mathlib.Algebra.Module.Submodule.EqLocus /-! # Adjoining elements to form subalgebras @@ -428,13 +428,19 @@ theorem map_adjoin_singleton (e : A →ₐ[R] B) (x : A) : rw [map_adjoin, Set.image_singleton] theorem adjoin_le_equalizer (φ₁ φ₂ : A →ₐ[R] B) {s : Set A} (h : s.EqOn φ₁ φ₂) : - adjoin R s ≤ φ₁.equalizer φ₂ := + adjoin R s ≤ equalizer φ₁ φ₂ := adjoin_le h theorem ext_of_adjoin_eq_top {s : Set A} (h : adjoin R s = ⊤) ⦃φ₁ φ₂ : A →ₐ[R] B⦄ (hs : s.EqOn φ₁ φ₂) : φ₁ = φ₂ := ext fun _x => adjoin_le_equalizer φ₁ φ₂ hs <| h.symm ▸ trivial +/-- Two algebra morphisms are equal on `Algebra.span s`iff they are equal on s -/ +theorem eqOn_adjoin_iff {φ ψ : A →ₐ[R] B} {s : Set A} : + Set.EqOn φ ψ (adjoin R s) ↔ Set.EqOn φ ψ s := by + have (S : Set A) : S ≤ equalizer φ ψ ↔ Set.EqOn φ ψ S := Iff.rfl + simp only [← this, Set.le_eq_subset, SetLike.coe_subset_coe, adjoin_le_iff] + end AlgHom section NatInt diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index d0ae553ff795a..b0337aad5d373 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -280,7 +280,7 @@ theorem aeval_algHom_eq_zero (ϕ : AdjoinRoot f →ₐ[R] S) : aeval (ϕ (root f @[simp] theorem liftHom_eq_algHom (f : R[X]) (ϕ : AdjoinRoot f →ₐ[R] S) : liftHom f (ϕ (root f)) (aeval_algHom_eq_zero f ϕ) = ϕ := by - suffices ϕ.equalizer (liftHom f (ϕ (root f)) (aeval_algHom_eq_zero f ϕ)) = ⊤ by + suffices AlgHom.equalizer ϕ (liftHom f (ϕ (root f)) (aeval_algHom_eq_zero f ϕ)) = ⊤ by exact (AlgHom.ext fun x => (SetLike.ext_iff.mp this x).mpr Algebra.mem_top).symm rw [eq_top_iff, ← adjoinRoot_eq_top, Algebra.adjoin_le_iff, Set.singleton_subset_iff] exact (@lift_root _ _ _ _ _ _ _ (aeval_algHom_eq_zero f ϕ)).symm diff --git a/Mathlib/RingTheory/AlgebraTower.lean b/Mathlib/RingTheory/AlgebraTower.lean index 87fd6e920a9c7..48c238fa1bb95 100644 --- a/Mathlib/RingTheory/AlgebraTower.lean +++ b/Mathlib/RingTheory/AlgebraTower.lean @@ -126,7 +126,8 @@ variable {R} /-- `Basis.smulTower (b : Basis ι R S) (c : Basis ι S A)` is the `R`-basis on `A` where the `(i, j)`th basis vector is `b i • c j`. -/ -noncomputable def Basis.smulTower {ι : Type v₁} {ι' : Type w₁} (b : Basis ι R S) (c : Basis ι' S A) : +noncomputable +def Basis.smulTower {ι : Type v₁} {ι' : Type w₁} (b : Basis ι R S) (c : Basis ι' S A) : Basis (ι × ι') R A := haveI := c.isScalarTower_finsupp R .ofRepr diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index e15a34f946e5b..db3f789449c62 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -169,7 +169,7 @@ theorem ascPochhammer_smeval_cast (R : Type*) [Semiring R] {S : Type*} [NonAssoc induction' n with n hn · simp only [Nat.zero_eq, ascPochhammer_zero, smeval_one, one_smul] · simp only [ascPochhammer_succ_right, mul_add, smeval_add, smeval_mul_X, ← Nat.cast_comm] - simp only [← C_eq_natCast, smeval_C_mul, hn, ← nsmul_eq_smul_cast R n] + simp only [← C_eq_natCast, smeval_C_mul, hn, Nat.cast_smul_eq_nsmul R n] simp only [nsmul_eq_mul, Nat.cast_id] variable {R S : Type*} @@ -253,7 +253,7 @@ instance Int.instBinomialRing : BinomialRing ℤ where noncomputable instance {R : Type*} [AddCommMonoid R] [Module ℚ≥0 R] [Pow R ℕ] : BinomialRing R where nsmul_right_injective n hn r s hrs := by rw [← one_smul ℚ≥0 r, ← one_smul ℚ≥0 s, show 1 = (n : ℚ≥0)⁻¹ • (n : ℚ≥0) by simp_all] - simp_all only [smul_assoc, ← nsmul_eq_smul_cast] + simp_all only [smul_assoc, Nat.cast_smul_eq_nsmul] multichoose r n := (n.factorial : ℚ≥0)⁻¹ • Polynomial.smeval (ascPochhammer ℕ n) r factorial_nsmul_multichoose r n := by simp only [← smul_assoc] diff --git a/Mathlib/RingTheory/ClassGroup.lean b/Mathlib/RingTheory/ClassGroup.lean index 06012aa52bffb..036bd50a67491 100644 --- a/Mathlib/RingTheory/ClassGroup.lean +++ b/Mathlib/RingTheory/ClassGroup.lean @@ -356,9 +356,7 @@ of global fields. noncomputable instance [IsPrincipalIdealRing R] : Fintype (ClassGroup R) where elems := {1} complete := by - #adaptation_note - /-- 2024-07-18 need to disable elab_as_elim with `(P := _)` due to type-incorrect motive -/ - refine ClassGroup.induction (P := _) (R := R) (FractionRing R) (fun I => ?_) + refine ClassGroup.induction (R := R) (FractionRing R) (fun I => ?_) rw [Finset.mem_singleton] exact ClassGroup.mk_eq_one_iff.mpr (I : FractionalIdeal R⁰ (FractionRing R)).isPrincipal @@ -366,9 +364,7 @@ noncomputable instance [IsPrincipalIdealRing R] : Fintype (ClassGroup R) where theorem card_classGroup_eq_one [IsPrincipalIdealRing R] : Fintype.card (ClassGroup R) = 1 := by rw [Fintype.card_eq_one_iff] use 1 - #adaptation_note - /-- 2024-07-18 need to disable elab_as_elim due to incorrect motive -/ - refine ClassGroup.induction (P := _) (R := R) (FractionRing R) (fun I => ?_) + refine ClassGroup.induction (R := R) (FractionRing R) (fun I => ?_) exact ClassGroup.mk_eq_one_iff.mpr (I : FractionalIdeal R⁰ (FractionRing R)).isPrincipal /-- The class number is `1` iff the ring of integers is a principal ideal domain. -/ diff --git a/Mathlib/RingTheory/Coprime/Basic.lean b/Mathlib/RingTheory/Coprime/Basic.lean index fa7dbdb6959ea..24c7f06bb3275 100644 --- a/Mathlib/RingTheory/Coprime/Basic.lean +++ b/Mathlib/RingTheory/Coprime/Basic.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Ken Lee, Chris Hughes -/ +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Algebra.Ring.Hom.Defs -import Mathlib.GroupTheory.GroupAction.Units import Mathlib.Logic.Basic import Mathlib.Tactic.Ring @@ -365,7 +365,7 @@ namespace IsRelPrime variable {R} [CommRing R] {x y : R} (h : IsRelPrime x y) (z : R) -theorem add_mul_left_left : IsRelPrime (x + y * z) y := +theorem add_mul_left_left (h : IsRelPrime x y) : IsRelPrime (x + y * z) y := @of_add_mul_left_left R _ _ _ (-z) <| by simpa only [mul_neg, add_neg_cancel_right] using h theorem add_mul_right_left : IsRelPrime (x + z * y) y := diff --git a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean index 4385401a68139..9659b2099784c 100644 --- a/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean +++ b/Mathlib/RingTheory/DedekindDomain/AdicValuation.lean @@ -83,6 +83,10 @@ def intValuationDef (r : R) : ℤₘ₀ := theorem intValuationDef_if_pos {r : R} (hr : r = 0) : v.intValuationDef r = 0 := if_pos hr +@[simp] +theorem intValuationDef_zero : v.intValuationDef 0 = 0 := + if_pos rfl + theorem intValuationDef_if_neg {r : R} (hr : r ≠ 0) : v.intValuationDef r = Multiplicative.ofAdd @@ -312,6 +316,10 @@ theorem valuation_of_mk' {r : R} {s : nonZeroDivisors R} : theorem valuation_of_algebraMap (r : R) : v.valuation (algebraMap R K r) = v.intValuation r := by rw [valuation_def, Valuation.extendToLocalization_apply_map_apply] +open scoped algebraMap in +lemma valuation_eq_intValuationDef (r : R) : v.valuation (r : K) = v.intValuationDef r := + Valuation.extendToLocalization_apply_map_apply .. + /-- The `v`-adic valuation on `R` is bounded above by 1. -/ theorem valuation_le_one (r : R) : v.valuation (algebraMap R K r) ≤ 1 := by rw [valuation_of_algebraMap]; exact v.intValuation_le_one r @@ -394,6 +402,11 @@ theorem mem_adicCompletionIntegers {x : v.adicCompletion K} : x ∈ v.adicCompletionIntegers K ↔ (Valued.v x : ℤₘ₀) ≤ 1 := Iff.rfl +theorem not_mem_adicCompletionIntegers {x : v.adicCompletion K} : + x ∉ v.adicCompletionIntegers K ↔ 1 < (Valued.v x : ℤₘ₀) := by + rw [not_congr <| mem_adicCompletionIntegers R K v] + exact not_le + section AlgebraInstances instance (priority := 100) adicValued.has_uniform_continuous_const_smul' : @@ -457,6 +470,28 @@ instance : Algebra R (v.adicCompletionIntegers K) where simp only [Subring.coe_mul, Algebra.smul_def] rfl +variable {R K} in +open scoped algebraMap in -- to make the coercions from `R` fire +/-- The valuation on the completion agrees with the global valuation on elements of the +integer ring. -/ +theorem valuedAdicCompletion_eq_valuation (r : R) : + Valued.v (r : v.adicCompletion K) = v.valuation (r : K) := by + convert Valued.valuedCompletion_apply (r : K) + +variable {R K} in +/-- The valuation on the completion agrees with the global valuation on elements of the field. -/ +theorem valuedAdicCompletion_eq_valuation' (k : K) : + Valued.v (k : v.adicCompletion K) = v.valuation k := by + convert Valued.valuedCompletion_apply k + +variable {R K} in +open scoped algebraMap in -- to make the coercion from `R` fire +/-- A global integer is in the local integers. -/ +lemma coe_mem_adicCompletionIntegers (r : R) : + (r : adicCompletion K v) ∈ adicCompletionIntegers K v := by + rw [mem_adicCompletionIntegers, valuedAdicCompletion_eq_valuation, valuation_eq_intValuationDef] + exact intValuation_le_one v r + @[simp] theorem coe_smul_adicCompletionIntegers (r : R) (x : v.adicCompletionIntegers K) : (↑(r • x) : v.adicCompletion K) = r • (x : v.adicCompletion K) := @@ -477,4 +512,32 @@ instance adicCompletion.instIsScalarTower' : end AlgebraInstances +open nonZeroDivisors algebraMap in +variable {R K} in +lemma adicCompletion.mul_nonZeroDivisor_mem_adicCompletionIntegers (v : HeightOneSpectrum R) + (a : v.adicCompletion K) : ∃ b ∈ R⁰, a * b ∈ v.adicCompletionIntegers K := by + by_cases ha : a ∈ v.adicCompletionIntegers K + · use 1 + simp [ha, Submonoid.one_mem] + · rw [not_mem_adicCompletionIntegers] at ha + -- Let the additive valuation of a be -d with d>0 + obtain ⟨d, hd⟩ : ∃ d : ℤ, Valued.v a = ofAdd d := + Option.ne_none_iff_exists'.mp <| (lt_trans zero_lt_one ha).ne' + rw [hd, WithZero.one_lt_coe, ← ofAdd_zero, ofAdd_lt] at ha + -- let ϖ be a uniformiser + obtain ⟨ϖ, hϖ⟩ := intValuation_exists_uniformizer v + have hϖ0 : ϖ ≠ 0 := by rintro rfl; simp at hϖ + -- use ϖ^d + refine ⟨ϖ^d.natAbs, pow_mem (mem_nonZeroDivisors_of_ne_zero hϖ0) _, ?_⟩ + -- now manually translate the goal (an inequality in ℤₘ₀) to an inequality in ℤ + rw [mem_adicCompletionIntegers, algebraMap.coe_pow, map_mul, hd, map_pow, + valuedAdicCompletion_eq_valuation, valuation_eq_intValuationDef, hϖ, ← WithZero.coe_pow, + ← WithZero.coe_mul, WithZero.coe_le_one, ← toAdd_le, toAdd_mul, toAdd_ofAdd, toAdd_pow, + toAdd_ofAdd, toAdd_one, + show d.natAbs • (-1) = (d.natAbs : ℤ) • (-1) by simp only [nsmul_eq_mul, + Int.natCast_natAbs, smul_eq_mul], + ← Int.eq_natAbs_of_zero_le ha.le, smul_eq_mul] + -- and now it's easy + linarith + end IsDedekindDomain.HeightOneSpectrum diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index ec3da1af1cae9..7b7ad16bb8ddd 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -327,6 +327,12 @@ instance : CommRing (FiniteAdeleRing R K) := instance : Algebra K (FiniteAdeleRing R K) := Subalgebra.algebra (subalgebra R K) +instance : Algebra R (FiniteAdeleRing R K) := + ((algebraMap K (FiniteAdeleRing R K)).comp (algebraMap R K)).toAlgebra + +instance : IsScalarTower R K (FiniteAdeleRing R K) := + IsScalarTower.of_algebraMap_eq' rfl + instance : Coe (FiniteAdeleRing R K) (K_hat R K) where coe := fun x ↦ x.1 @@ -348,6 +354,86 @@ instance : Algebra (R_hat R K) (FiniteAdeleRing R K) where commutes' _ _ := mul_comm _ _ smul_def' r x := rfl +instance : CoeFun (FiniteAdeleRing R K) + (fun _ ↦ ∀ (v : HeightOneSpectrum R), adicCompletion K v) where + coe a v := a.1 v + +open scoped algebraMap in +variable {R K} in +lemma exists_finiteIntegralAdele_iff (a : FiniteAdeleRing R K) : (∃ c : R_hat R K, + a = c) ↔ ∀ (v : HeightOneSpectrum R), a v ∈ adicCompletionIntegers K v := + ⟨by rintro ⟨c, rfl⟩ v; exact (c v).2, fun h ↦ ⟨fun v ↦ ⟨a v, h v⟩, rfl⟩⟩ + +section Topology + +open Classical nonZeroDivisors Multiplicative Additive IsDedekindDomain.HeightOneSpectrum + +open scoped algebraMap -- coercion from R to FiniteAdeleRing R K +open scoped DiscreteValuation + +variable {R K} in +lemma mul_nonZeroDivisor_mem_finiteIntegralAdeles (a : FiniteAdeleRing R K) : + ∃ (b : R⁰) (c : R_hat R K), a * ((b : R) : FiniteAdeleRing R K) = c := by + let S := {v | a v ∉ adicCompletionIntegers K v} + choose b hb h using adicCompletion.mul_nonZeroDivisor_mem_adicCompletionIntegers (R := R) (K := K) + let p := ∏ᶠ v ∈ S, b v (a v) + have hp : p ∈ R⁰ := finprod_mem_induction (· ∈ R⁰) (one_mem _) (fun _ _ => mul_mem) <| + fun _ _ ↦ hb _ _ + use ⟨p, hp⟩ + rw [exists_finiteIntegralAdele_iff] + intro v + by_cases hv : a v ∈ adicCompletionIntegers K v + · exact mul_mem hv <| coe_mem_adicCompletionIntegers _ _ + · dsimp only + have pprod : p = b v (a v) * ∏ᶠ w ∈ S \ {v}, b w (a w) := by + rw [← finprod_mem_singleton (a := v) (f := fun v ↦ b v (a v)), + finprod_mem_mul_diff (singleton_subset_iff.2 ‹v ∈ S›) a.2] + rw [pprod] + push_cast + rw [← mul_assoc] + exact mul_mem (h v (a v)) <| coe_mem_adicCompletionIntegers _ _ + +open scoped Pointwise + +theorem submodulesRingBasis : SubmodulesRingBasis + (fun (r : R⁰) ↦ Submodule.span (R_hat R K) {((r : R) : FiniteAdeleRing R K)}) where + inter i j := ⟨i * j, by + push_cast + simp only [le_inf_iff, Submodule.span_singleton_le_iff_mem, Submodule.mem_span_singleton] + exact ⟨⟨((j : R) : R_hat R K), by rw [mul_comm]; rfl⟩, ⟨((i : R) : R_hat R K), rfl⟩⟩⟩ + leftMul a r := by + rcases mul_nonZeroDivisor_mem_finiteIntegralAdeles a with ⟨b, c, h⟩ + use r * b + rintro x ⟨m, hm, rfl⟩ + simp only [Submonoid.coe_mul, SetLike.mem_coe] at hm + rw [Submodule.mem_span_singleton] at hm ⊢ + rcases hm with ⟨n, rfl⟩ + simp only [LinearMapClass.map_smul, DistribMulAction.toLinearMap_apply, smul_eq_mul] + use n * c + push_cast + rw [mul_left_comm, h, mul_comm _ (c : FiniteAdeleRing R K), + Algebra.smul_def', Algebra.smul_def', ← mul_assoc] + rfl + mul r := ⟨r, by + intro x hx + rw [mem_mul] at hx + rcases hx with ⟨a, ha, b, hb, rfl⟩ + simp only [SetLike.mem_coe, Submodule.mem_span_singleton] at ha hb ⊢ + rcases ha with ⟨m, rfl⟩ + rcases hb with ⟨n, rfl⟩ + use m * n * (r : R) + simp only [Algebra.smul_def', map_mul] + rw [mul_mul_mul_comm, mul_assoc] + rfl⟩ + +instance : TopologicalSpace (FiniteAdeleRing R K) := + SubmodulesRingBasis.topology (submodulesRingBasis R K) + +-- the point of the above: this now works +example : TopologicalRing (FiniteAdeleRing R K) := inferInstance + +end Topology + end FiniteAdeleRing end DedekindDomain diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index 9c3f026f9aadc..59fa5928122fd 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -149,7 +149,7 @@ theorem map_aeval (P : R[X]) (x : A) : induction P using Polynomial.induction_on · simp · simp [add_smul, *] - · simp [mul_smul, nsmul_eq_smul_cast A] + · simp [mul_smul, ← Nat.cast_smul_eq_nsmul A] theorem eqOn_adjoin {s : Set A} (h : Set.EqOn D1 D2 s) : Set.EqOn D1 D2 (adjoin R s) := fun x hx => Algebra.adjoin_induction hx h (fun r => (D1.map_algebraMap r).trans (D2.map_algebraMap r).symm) @@ -423,7 +423,8 @@ lemma leibniz_zpow (a : K) (n : ℤ) : D (a ^ n) = n • a ^ (n - 1) • D a := congr omega · rw [h, zpow_neg, zpow_natCast, leibniz_inv, leibniz_pow, inv_pow, ← pow_mul, ← zpow_natCast, - ← zpow_natCast, nsmul_eq_smul_cast K, zsmul_eq_smul_cast K, smul_smul, smul_smul, smul_smul] + ← zpow_natCast, ← Nat.cast_smul_eq_nsmul K, ← Int.cast_smul_eq_nsmul K, smul_smul, smul_smul, + smul_smul] trans (-n.natAbs * (a ^ ((n.natAbs - 1 : ℕ) : ℤ) / (a ^ ((n.natAbs * 2 : ℕ) : ℤ)))) • D a · ring_nf rw [← zpow_sub₀ ha] diff --git a/Mathlib/RingTheory/Flat/EquationalCriterion.lean b/Mathlib/RingTheory/Flat/EquationalCriterion.lean index fa89effd7fecc..deef6f5c4e908 100644 --- a/Mathlib/RingTheory/Flat/EquationalCriterion.lean +++ b/Mathlib/RingTheory/Flat/EquationalCriterion.lean @@ -220,10 +220,11 @@ Let $M$ be a flat module. Let $R^\iota$ be a finite free module, let $f \in R^{\ element, and let $x \colon R^{\iota} \to M$ be a homomorphism such that $x(f) = 0$. Then there exist a finite free module $R^\kappa$ and homomorphisms $a \colon R^{\iota} \to R^{\kappa}$ and $y \colon R^{\kappa} \to M$ such that $x = y \circ a$ and $a(f) = 0$. -/ -theorem exists_factorization_of_apply_eq_zero [Flat R M] {ι : Type u} [Fintype ι] +theorem exists_factorization_of_apply_eq_zero [Flat R M] {ι : Type u} [_root_.Finite ι] {f : ι →₀ R} {x : (ι →₀ R) →ₗ[R] M} (h : x f = 0) : ∃ (κ : Type u) (_ : Fintype κ) (a : (ι →₀ R) →ₗ[R] (κ →₀ R)) (y : (κ →₀ R) →ₗ[R] M), - x = y ∘ₗ a ∧ a f = 0 := iff_forall_exists_factorization.mp ‹Flat R M› h + x = y ∘ₗ a ∧ a f = 0 := + let ⟨_⟩ := nonempty_fintype ι; iff_forall_exists_factorization.mp ‹Flat R M› h /-- **Equational criterion for flatness** [Stacks 00HK](https://stacks.math.columbia.edu/tag/00HK), backward direction, alternate form. @@ -263,7 +264,7 @@ theorem exists_factorization_of_comp_eq_zero_of_free [Flat R M] {K N : Type u} have (K' : Submodule R K) (hK' : K'.FG) : ∃ (κ : Type u) (_ : Fintype κ) (a : N →ₗ[R] (κ →₀ R)) (y : (κ →₀ R) →ₗ[R] M), x = y ∘ₗ a ∧ K' ≤ LinearMap.ker (a ∘ₗ f) := by revert N - apply Submodule.fg_induction (P := _) (N := K') (hN := hK') + apply Submodule.fg_induction (N := K') (hN := hK') · intro k N _ _ _ _ f x hfx have : x (f k) = 0 := by simpa using LinearMap.congr_fun hfx k simpa using exists_factorization_of_apply_eq_zero_of_free this diff --git a/Mathlib/RingTheory/Flat/Stability.lean b/Mathlib/RingTheory/Flat/Stability.lean index 5b246b916037f..c1d4983e9f581 100644 --- a/Mathlib/RingTheory/Flat/Stability.lean +++ b/Mathlib/RingTheory/Flat/Stability.lean @@ -6,6 +6,8 @@ Authors: Christian Merten import Mathlib.RingTheory.Flat.Basic import Mathlib.RingTheory.IsTensorProduct import Mathlib.LinearAlgebra.TensorProduct.Tower +import Mathlib.RingTheory.Localization.BaseChange +import Mathlib.Algebra.Module.LocalizedModule /-! # Flatness is stable under composition and base change @@ -18,7 +20,9 @@ We show that flatness is stable under composition and base change. then `M` is a flat `R`-module * `Module.Flat.baseChange`: if `M` is a flat `R`-module and `S` is any `R`-algebra, then `S ⊗[R] M` is `S`-flat. - +* `Module.Flat.of_isLocalizedModule`: if `M` is a flat `R`-module and `S` is a submonoid of `R` + then the localization of `M` at `S` is flat as a module + for the localzation of `R` at `S`. -/ universe u v w t @@ -145,4 +149,24 @@ theorem isBaseChange [Module.Flat R M] (N : Type t) [AddCommGroup N] [Module R N end BaseChange +section Localization + +variable {R : Type u} {M Mp : Type*} (Rp : Type v) + [CommRing R] [AddCommGroup M] [Module R M] [CommRing Rp] [Algebra R Rp] + [AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] (f : M →ₗ[R] Mp) + +instance localizedModule [Module.Flat R M] (S : Submonoid R) : Module.Flat (Localization S) + (LocalizedModule S M) := by + fapply Module.Flat.isBaseChange (R := R) (M := M) (S := Localization S) + exact LocalizedModule.mkLinearMap S M + rw [← isLocalizedModule_iff_isBaseChange S] + exact localizedModuleIsLocalizedModule S + +theorem of_isLocalizedModule [Module.Flat R M] (S : Submonoid R) [IsLocalization S Rp] + (f : M →ₗ[R] Mp) [h : IsLocalizedModule S f] : Module.Flat Rp Mp := by + fapply Module.Flat.isBaseChange (R := R) (M := M) (S := Rp) (N := Mp) + exact (isLocalizedModule_iff_isBaseChange S Rp f).mp h + +end Localization + end Module.Flat diff --git a/Mathlib/RingTheory/Generators.lean b/Mathlib/RingTheory/Generators.lean index 33424b336b9aa..2df545921eb18 100644 --- a/Mathlib/RingTheory/Generators.lean +++ b/Mathlib/RingTheory/Generators.lean @@ -392,13 +392,11 @@ instance Cotangent.module : Module S P.Cotangent where smul_add := fun r x y ↦ ext (smul_add (P.σ r) x.val y.val) add_smul := fun r s x ↦ by have := smul_eq_zero_of_mem (P.σ (r + s) - (P.σ r + P.σ s) : P.Ring) (by simp ) x - simp only [sub_smul, add_smul, sub_eq_zero] at this - exact this + simpa only [sub_smul, add_smul, sub_eq_zero] zero_smul := fun x ↦ smul_eq_zero_of_mem (P.σ 0 : P.Ring) (by simp) x one_smul := fun x ↦ by have := smul_eq_zero_of_mem (P.σ 1 - 1 : P.Ring) (by simp) x - simp [sub_eq_zero, sub_smul] at this - exact this + simpa [sub_eq_zero, sub_smul] mul_smul := fun r s x ↦ by have := smul_eq_zero_of_mem (P.σ (r * s) - (P.σ r * P.σ s) : P.Ring) (by simp) x simpa only [sub_smul, mul_smul, sub_eq_zero] using this diff --git a/Mathlib/RingTheory/GradedAlgebra/Basic.lean b/Mathlib/RingTheory/GradedAlgebra/Basic.lean index d18ad7d09de82..11e8d628a73b7 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Basic.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Basic.lean @@ -334,3 +334,29 @@ theorem coe_decompose_mul_of_right_mem (n) [Decidable (i ≤ n)] (b_mem : b ∈ end DirectSum end CanonicalOrder + +namespace DirectSum.IsInternal + +variable {R : Type*} [CommSemiring R] {A : Type*} [Semiring A] [Algebra R A] +variable {ι : Type*} [DecidableEq ι] [AddMonoid ι] +variable {M : ι → Submodule R A} [SetLike.GradedMonoid M] + +-- The following lines were given on Zulip by Adam Topaz +/-- The canonical isomorphism of an internal direct sum with the ambient algebra -/ +noncomputable def coeAlgEquiv (hM : DirectSum.IsInternal M) : + (DirectSum ι fun i => ↥(M i)) ≃ₐ[R] A := + { RingEquiv.ofBijective (DirectSum.coeAlgHom M) hM with commutes' := fun r => by simp } + +/-- Given an `R`-algebra `A` and a family `ι → Submodule R A` of submodules +parameterized by an additive monoid `ι` +and statisfying `SetLike.GradedMonoid M` (essentially, is multiplicative) +such that `DirectSum.IsInternal M` (`A` is the direct sum of the `M i`), +we endow `A` with the structure of a graded algebra. +The submodules are the *homogeneous* parts. -/ +noncomputable def gradedAlgebra (hM : DirectSum.IsInternal M) : GradedAlgebra M := + { (inferInstance : SetLike.GradedMonoid M) with + decompose' := hM.coeAlgEquiv.symm + left_inv := hM.coeAlgEquiv.symm.left_inv + right_inv := hM.coeAlgEquiv.left_inv } + +end DirectSum.IsInternal diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index fd5ec17dad502..4dce21fd3cc22 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -410,7 +410,7 @@ def powers (x : HahnSeries Γ R) (hx : 0 < x.orderTop) : SummableFamily Γ R ℕ · obtain ⟨hi, _, rfl⟩ := mem_addAntidiagonal.1 (mem_coe.1 hij) rw [← zero_add ij.snd, ← add_assoc, add_zero] exact - add_lt_add_right (WithTop.coe_lt_coe.1 (lt_of_lt_of_le hx (orderTop_le_of_coeff_ne_zero hi))) + add_lt_add_right (WithTop.coe_lt_coe.1 (hx.trans_le (orderTop_le_of_coeff_ne_zero hi))) _ · rintro (_ | n) hn · exact Set.mem_union_right _ (Set.mem_singleton 0) @@ -508,9 +508,10 @@ instance instField [Field R] : Field (HahnSeries Γ R) where (unit_aux x (inv_mul_cancel (leadingCoeff_ne_iff.mpr x0))) rw [sub_sub_cancel] at h rw [← mul_assoc, mul_comm x, h] - nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q a => rfl end Inversion diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 6b5fda9c8699a..689fa8e0c1472 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -239,7 +239,7 @@ theorem zero_ne_one_of_proper {I : Ideal α} (h : I ≠ ⊤) : (0 : α) ≠ 1 := I.ne_top_iff_one.1 h <| hz ▸ I.zero_mem theorem bot_prime [IsDomain α] : (⊥ : Ideal α).IsPrime := - ⟨fun h => one_ne_zero (by rwa [Ideal.eq_top_iff_one, Submodule.mem_bot] at h), fun h => + ⟨fun h => one_ne_zero (α := α) (by rwa [Ideal.eq_top_iff_one, Submodule.mem_bot] at h), fun h => mul_eq_zero.mp (by simpa only [Submodule.mem_bot] using h)⟩ /-- An ideal is maximal if it is maximal in the collection of proper ideals. -/ diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 458062650257e..071fae5bc69ae 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -296,14 +296,12 @@ end Surjective section Injective -variable (hf : Function.Injective f) - -theorem comap_bot_le_of_injective : comap f ⊥ ≤ I := by +theorem comap_bot_le_of_injective (hf : Function.Injective f) : comap f ⊥ ≤ I := by refine le_trans (fun x hx => ?_) bot_le rw [mem_comap, Submodule.mem_bot, ← map_zero f] at hx exact Eq.symm (hf hx) ▸ Submodule.zero_mem ⊥ -theorem comap_bot_of_injective : Ideal.comap f ⊥ = ⊥ := +theorem comap_bot_of_injective (hf : Function.Injective f) : Ideal.comap f ⊥ = ⊥ := le_bot_iff.mp (Ideal.comap_bot_le_of_injective f hf) end Injective @@ -349,9 +347,8 @@ variable [FunLike F R S] [RingHomClass F R S] (f : F) {I : Ideal R} section Surjective -variable (hf : Function.Surjective f) - -theorem comap_map_of_surjective (I : Ideal R) : comap f (map f I) = I ⊔ comap f ⊥ := +theorem comap_map_of_surjective (hf : Function.Surjective f) (I : Ideal R) : + comap f (map f I) = I ⊔ comap f ⊥ := le_antisymm (fun r h => let ⟨s, hsi, hfsr⟩ := mem_image_of_mem_map_of_surjective f hf h @@ -361,7 +358,8 @@ theorem comap_map_of_surjective (I : Ideal R) : comap f (map f I) = I ⊔ comap (sup_le (map_le_iff_le_comap.1 le_rfl) (comap_mono bot_le)) /-- Correspondence theorem -/ -def relIsoOfSurjective : Ideal S ≃o { p : Ideal R // comap f ⊥ ≤ p } where +def relIsoOfSurjective (hf : Function.Surjective f) : + Ideal S ≃o { p : Ideal R // comap f ⊥ ≤ p } where toFun J := ⟨comap f J, comap_mono bot_le⟩ invFun I := map f I.1 left_inv J := map_comap_of_surjective f hf J @@ -374,11 +372,11 @@ def relIsoOfSurjective : Ideal S ≃o { p : Ideal R // comap f ⊥ ≤ p } where comap_mono⟩ /-- The map on ideals induced by a surjective map preserves inclusion. -/ -def orderEmbeddingOfSurjective : Ideal S ↪o Ideal R := +def orderEmbeddingOfSurjective (hf : Function.Surjective f) : Ideal S ↪o Ideal R := (relIsoOfSurjective f hf).toRelEmbedding.trans (Subtype.relEmbedding (fun x y => x ≤ y) _) -theorem map_eq_top_or_isMaximal_of_surjective {I : Ideal R} (H : IsMaximal I) : - map f I = ⊤ ∨ IsMaximal (map f I) := by +theorem map_eq_top_or_isMaximal_of_surjective (hf : Function.Surjective f) {I : Ideal R} + (H : IsMaximal I) : map f I = ⊤ ∨ IsMaximal (map f I) := by refine or_iff_not_imp_left.2 fun ne_top => ⟨⟨fun h => ne_top h, fun J hJ => ?_⟩⟩ · refine (relIsoOfSurjective f hf).injective @@ -386,7 +384,8 @@ theorem map_eq_top_or_isMaximal_of_surjective {I : Ideal R} (H : IsMaximal I) : · exact map_le_iff_le_comap.1 (le_of_lt hJ) · exact fun h => hJ.right (le_map_of_comap_le_of_surjective f hf (le_of_eq h.symm)) -theorem comap_isMaximal_of_surjective {K : Ideal S} [H : IsMaximal K] : IsMaximal (comap f K) := by +theorem comap_isMaximal_of_surjective (hf : Function.Surjective f) {K : Ideal S} [H : IsMaximal K]: + IsMaximal (comap f K) := by refine ⟨⟨comap_ne_top _ H.1.1, fun J hJ => ?_⟩⟩ suffices map f J = ⊤ by have := congr_arg (comap f) this @@ -400,7 +399,8 @@ theorem comap_isMaximal_of_surjective {K : Ideal S} [H : IsMaximal K] : IsMaxima rw [comap_map_of_surjective _ hf, sup_eq_left] exact le_trans (comap_mono bot_le) (le_of_lt hJ) -theorem comap_le_comap_iff_of_surjective (I J : Ideal S) : comap f I ≤ comap f J ↔ I ≤ J := +theorem comap_le_comap_iff_of_surjective (hf : Function.Surjective f) (I J : Ideal S) : + comap f I ≤ comap f J ↔ I ≤ J := ⟨fun h => (map_comap_of_surjective f hf I).symm.le.trans (map_le_of_le_comap h), fun h => le_comap_of_map_le ((map_comap_of_surjective f hf I).le.trans h)⟩ @@ -409,10 +409,8 @@ end Surjective section Bijective -variable (hf : Function.Bijective f) - /-- Special case of the correspondence theorem for isomorphic rings -/ -def relIsoOfBijective : Ideal S ≃o Ideal R where +def relIsoOfBijective (hf : Function.Bijective f) : Ideal S ≃o Ideal R where toFun := comap f invFun := map f left_inv := (relIsoOfSurjective f hf.right).left_inv @@ -421,11 +419,13 @@ def relIsoOfBijective : Ideal S ≃o Ideal R where ((relIsoOfSurjective f hf.right).right_inv ⟨J, comap_bot_le_of_injective f hf.left⟩) map_rel_iff' {_ _} := (relIsoOfSurjective f hf.right).map_rel_iff' -theorem comap_le_iff_le_map {I : Ideal R} {K : Ideal S} : comap f K ≤ I ↔ K ≤ map f I := +theorem comap_le_iff_le_map (hf : Function.Bijective f) {I : Ideal R} {K : Ideal S} : + comap f K ≤ I ↔ K ≤ map f I := ⟨fun h => le_map_of_comap_le_of_surjective f hf.right h, fun h => (relIsoOfBijective f hf).right_inv I ▸ comap_mono h⟩ -theorem map.isMaximal {I : Ideal R} (H : IsMaximal I) : IsMaximal (map f I) := by +theorem map.isMaximal (hf : Function.Bijective f) {I : Ideal R} (H : IsMaximal I) : + IsMaximal (map f I) := by refine or_iff_not_imp_left.1 (map_eq_top_or_isMaximal_of_surjective f hf.right H) fun h => H.1.1 ?_ calc diff --git a/Mathlib/RingTheory/Ideal/Quotient.lean b/Mathlib/RingTheory/Ideal/Quotient.lean index 64f99d921bfe9..ca7045bb61f1e 100644 --- a/Mathlib/RingTheory/Ideal/Quotient.lean +++ b/Mathlib/RingTheory/Ideal/Quotient.lean @@ -202,7 +202,9 @@ protected noncomputable abbrev field (I : Ideal R) [hI : I.IsMaximal] : Field (R __ := commRing _ __ := Quotient.groupWithZero _ nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q x => rfl /-- If the quotient by an ideal is a field, then the ideal is maximal. -/ theorem maximal_of_isField (I : Ideal R) (hqf : IsField (R ⧸ I)) : I.IsMaximal := by diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean new file mode 100644 index 0000000000000..0278341a7ee7a --- /dev/null +++ b/Mathlib/RingTheory/Idempotents.lean @@ -0,0 +1,422 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.Algebra.GeomSum +import Mathlib.Algebra.Polynomial.AlgebraMap +import Mathlib.RingTheory.Ideal.QuotientOperations +import Mathlib.RingTheory.Nilpotent.Defs + +/-! + +## Idempotents in rings + +The predicate `IsIdempotentElem` is defined for general monoids in `Algebra/Ring/Idempotents.lean`. +In this file we provide various results regarding idempotent elements in rings. + +## Main definitions + +- `OrthogonalIdempotents`: + A family `{ eᵢ }` of idempotent elements is orthogonal if `eᵢ * eⱼ = 0` for all `i ≠ j`. +- `CompleteOrthogonalIdempotents`: + A family `{ eᵢ }` of orthogonal idempotent elements is complete if `∑ eᵢ = 1`. + +## Main results + +- `CompleteOrthogonalIdempotents.lift_of_isNilpotent_ker`: + If the kernel of `f : R →+* S` consists of nilpotent elements, and `{ eᵢ }` is a family of + complete orthogonal idempotents in the range of `f`, then `{ eᵢ }` is the image of some + complete orthogonal idempotents in `R`. +- `existsUnique_isIdempotentElem_eq_of_ker_isNilpotent`: + If `R` is commutative and the kernel of `f : R →+* S` consists of nilpotent elements, + then every idempotent in the range of `f` lifts to a unique idempotent in `R`. +- `CompleteOrthogonalIdempotents.bijective_pi`: + 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) + +variable {I : Type*} (e : I → R) + +/-- A family `{ eᵢ }` of idempotent elements is orthogonal if `eᵢ * eⱼ = 0` for all `i ≠ j`. -/ +@[mk_iff] +structure OrthogonalIdempotents : Prop where + idem : ∀ i, IsIdempotentElem (e i) + ortho : ∀ i j, i ≠ j → e i * e j = 0 + +variable {e} +variable (he : OrthogonalIdempotents e) + +lemma OrthogonalIdempotents.mul_eq [DecidableEq I] (he : OrthogonalIdempotents e) (i j) : + e i * e j = if i = j then e i else 0 := by + split + · simp [*, (he.idem j).eq] + · exact he.ortho _ _ ‹_› + +lemma OrthogonalIdempotents.iff_mul_eq [DecidableEq I] : + OrthogonalIdempotents e ↔ ∀ i j, e i * e j = if i = j then e i else 0 := + ⟨mul_eq, fun H ↦ ⟨fun i ↦ by simpa using H i i, fun i j e ↦ by simpa [e] using H i j⟩⟩ + +lemma OrthogonalIdempotents.isIdempotentElem_sum [Fintype I] : IsIdempotentElem (∑ i, e i) := by + classical + simp [IsIdempotentElem, Finset.sum_mul, Finset.mul_sum, he.mul_eq] + +lemma OrthogonalIdempotents.map : + OrthogonalIdempotents (f ∘ e) := by + classical + simp [iff_mul_eq, he.mul_eq, ← map_mul f, apply_ite f] + +lemma OrthogonalIdempotents.map_injective_iff (hf : Function.Injective f) : + OrthogonalIdempotents (f ∘ e) ↔ OrthogonalIdempotents e := by + classical + simp [iff_mul_eq, ← hf.eq_iff, apply_ite] + +lemma OrthogonalIdempotents.embedding {J} (i : J ↪ I) : + OrthogonalIdempotents (e ∘ i) := by + classical + simp [iff_mul_eq, he.mul_eq] + +lemma OrthogonalIdempotents.equiv {J} (i : J ≃ I) : + OrthogonalIdempotents (e ∘ i) ↔ OrthogonalIdempotents e := by + classical + simp [iff_mul_eq, i.forall_congr_left] + +lemma OrthogonalIdempotents.unique [Unique I] : + OrthogonalIdempotents e ↔ IsIdempotentElem (e default) := by + simp [orthogonalIdempotents_iff, Unique.forall_iff] + +lemma OrthogonalIdempotents.option [Fintype I] (x) + (hx : IsIdempotentElem x) (hx₁ : x * ∑ i, e i = 0) (hx₂ : (∑ i, e i) * x = 0) : + OrthogonalIdempotents (Option.elim · x e) where + idem i := i.rec hx he.idem + ortho i j ne := by + classical + cases' i with i <;> cases' j with j + · cases ne rfl + · simpa only [mul_assoc, Finset.sum_mul, he.mul_eq, Finset.sum_ite_eq', Finset.mem_univ, + ↓reduceIte, zero_mul] using congr_arg (· * e j) hx₁ + · simpa only [Option.elim_some, Option.elim_none, ← mul_assoc, Finset.mul_sum, he.mul_eq, + Finset.sum_ite_eq, Finset.mem_univ, ↓reduceIte, mul_zero] using congr_arg (e i * ·) hx₂ + · exact he.ortho i j (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] + +/-- +A family `{ eᵢ }` of idempotent elements is complete orthogonal if +1. (orthogonal) `eᵢ * eⱼ = 0` for all `i ≠ j`. +2. (complete) `∑ eᵢ = 1` +-/ +@[mk_iff] +structure CompleteOrthogonalIdempotents (e : I → R) extends OrthogonalIdempotents e : Prop where + complete : ∑ i, e i = 1 + +variable (he : CompleteOrthogonalIdempotents e) + +lemma CompleteOrthogonalIdempotents.unique_iff [Unique I] : + CompleteOrthogonalIdempotents e ↔ e default = 1 := by + rw [completeOrthogonalIdempotents_iff, orthogonalIdempotents_iff] + simp only [Unique.forall_iff, ne_eq, not_true_eq_false, false_implies, and_true, + Finset.univ_unique, Finset.sum_singleton, 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] + 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)] : + CompleteOrthogonalIdempotents (Pi.single (f := R) · 1) := by + refine ⟨⟨by simp [IsIdempotentElem, ← Pi.single_mul], ?_⟩, Finset.univ_sum_single 1⟩ + intros i j hij + ext k + by_cases hi : i = k + · subst hi; simp [hij] + · simp [hi] + +lemma CompleteOrthogonalIdempotents.map : + CompleteOrthogonalIdempotents (f ∘ e) where + __ := he.toOrthogonalIdempotents.map f + complete := by simp only [Function.comp_apply, ← map_sum, he.complete, map_one] + +lemma CompleteOrthogonalIdempotents.map_injective_iff (hf : Function.Injective f) : + CompleteOrthogonalIdempotents (f ∘ e) ↔ CompleteOrthogonalIdempotents e := by + simp [completeOrthogonalIdempotents_iff, ← hf.eq_iff, apply_ite, + OrthogonalIdempotents.map_injective_iff f hf] + +lemma CompleteOrthogonalIdempotents.equiv {J} [Fintype J] (i : J ≃ I) : + CompleteOrthogonalIdempotents (e ∘ i) ↔ CompleteOrthogonalIdempotents e := by + simp only [completeOrthogonalIdempotents_iff, OrthogonalIdempotents.equiv, Function.comp_apply, + and_congr_right_iff, Fintype.sum_equiv i _ e (fun _ ↦ rfl)] + +lemma CompleteOrthogonalIdempotents.option (he : OrthogonalIdempotents e) : + CompleteOrthogonalIdempotents (Option.elim · (1 - ∑ i, e i) e) where + __ := he.option _ he.isIdempotentElem_sum.one_sub + (by simp [sub_mul, he.isIdempotentElem_sum.eq]) (by simp [mul_sub, he.isIdempotentElem_sum.eq]) + complete := by + 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) : + ∃ e' : Fin n → R, CompleteOrthogonalIdempotents e' ∧ f ∘ e' = e := by + cases subsingleton_or_nontrivial R + · choose e' he' using he' + exact ⟨e', .of_subsingleton, funext he'⟩ + cases subsingleton_or_nontrivial S + · obtain ⟨n, hn⟩ := h 1 (Subsingleton.elim _ _) + simp at hn + cases' n with n + · simpa using he.complete + obtain ⟨e', h₁, h₂⟩ := OrthogonalIdempotents.lift_of_isNilpotent_ker f h he.1 he' + refine ⟨_, (equiv (finSuccEquiv n)).mpr + (CompleteOrthogonalIdempotents.option (h₁.embedding (Fin.succEmb _))), funext fun i ↦ ?_⟩ + have (i) : f (e' i) = e i := congr_fun h₂ i + obtain ⟨_ | i, rfl⟩ := (finSuccEquiv n).symm.surjective i + · simp only [Fin.val_succEmb, Function.comp_apply, finSuccEquiv_symm_none, finSuccEquiv_zero, + Option.elim_none, map_sub, map_one, map_sum, this, ← he.complete, sub_eq_iff_eq_add, + Fin.sum_univ_succ] + · simp [this] + +/-- A system of complete orthogonal idempotents lift along nil ideals. -/ +lemma CompleteOrthogonalIdempotents.lift_of_isNilpotent_ker + (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + {e : I → S} (he : CompleteOrthogonalIdempotents e) (he' : ∀ i, e i ∈ f.range) : + ∃ e' : I → R, CompleteOrthogonalIdempotents e' ∧ f ∘ e' = e := by + obtain ⟨e', h₁, h₂⟩ := lift_of_isNilpotent_ker_aux f h + ((equiv (Fintype.equivFin I).symm).mpr he) (fun _ ↦ he' _) + refine ⟨_, ((equiv (Fintype.equivFin I)).mpr h₁), + by ext x; simpa using congr_fun h₂ (Fintype.equivFin I x)⟩ + +theorem eq_of_isNilpotent_sub_of_isIdempotentElem_of_commute {e₁ e₂ : R} + (he₁ : IsIdempotentElem e₁) (he₂ : IsIdempotentElem e₂) (H : IsNilpotent (e₁ - e₂)) + (H' : Commute e₁ e₂) : + e₁ = e₂ := by + have : (e₁ - e₂) ^ 3 = (e₁ - e₂) := by + simp only [pow_succ, pow_zero, mul_sub, one_mul, sub_mul, he₁.eq, he₂.eq, + H'.eq, mul_assoc] + simp only [← mul_assoc, he₁.eq, he₂.eq] + abel + obtain ⟨n, hn⟩ := H + have : (e₁ - e₂) ^ (2 * n + 1) = (e₁ - e₂) := by + clear hn; induction n <;> simp [mul_add, add_assoc, pow_add _ (2 * _) 3, this, ← pow_succ, *] + rwa [pow_succ, two_mul, pow_add, hn, zero_mul, zero_mul, eq_comm, sub_eq_zero] at this + +theorem CompleteOrthogonalIdempotents.of_ker_isNilpotent_of_isMulCentral + (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + (he : ∀ i, IsIdempotentElem (e i)) + (he' : ∀ i, IsMulCentral (e i)) + (he'' : CompleteOrthogonalIdempotents (f ∘ e)) : + CompleteOrthogonalIdempotents e := by + obtain ⟨e', h₁, h₂⟩ := lift_of_isNilpotent_ker f h he'' (fun _ ↦ ⟨_, rfl⟩) + obtain rfl : e = e' := by + ext i + refine eq_of_isNilpotent_sub_of_isIdempotentElem_of_commute + (he _) (h₁.idem _) (h _ ?_) ((he' i).comm _) + simpa [RingHom.mem_ker, sub_eq_zero] using congr_fun h₂.symm i + exact h₁ + +end Ring + +section CommRing + +variable {R S : Type*} [CommRing R] [Ring S] (f : R →+* S) + +theorem eq_of_isNilpotent_sub_of_isIdempotentElem {e₁ e₂ : R} + (he₁ : IsIdempotentElem e₁) (he₂ : IsIdempotentElem e₂) (H : IsNilpotent (e₁ - e₂)) : + e₁ = e₂ := + eq_of_isNilpotent_sub_of_isIdempotentElem_of_commute he₁ he₂ H (.all _ _) + +theorem existsUnique_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 + obtain ⟨e', he₂, rfl⟩ := exists_isIdempotentElem_eq_of_ker_isNilpotent f h e he he' + exact ⟨e', ⟨he₂, rfl⟩, fun x ⟨hx, hx'⟩ ↦ + eq_of_isNilpotent_sub_of_isIdempotentElem hx he₂ + (h _ (by rw [RingHom.mem_ker, map_sub, hx', sub_self]))⟩ + +/-- A family of orthogonal idempotents induces an surjection `R ≃+* ∏ R ⧸ ⟨1 - eᵢ⟩` -/ +lemma OrthogonalIdempotents.surjective_pi {I : Type*} [Finite I] {e : I → R} + (he : OrthogonalIdempotents e) : + Function.Surjective (Pi.ringHom fun i ↦ Ideal.Quotient.mk (Ideal.span {1 - e i})) := by + suffices Pairwise fun i j ↦ IsCoprime (Ideal.span {1 - e i}) (Ideal.span {1 - e j}) by + intro x + obtain ⟨x, rfl⟩ := Ideal.quotientInfToPiQuotient_surj this x + obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x + exact ⟨x, by ext i; simp [Ideal.quotientInfToPiQuotient]⟩ + intros i j hij + rw [Ideal.isCoprime_span_singleton_iff] + exact ⟨1, e i, by simp [mul_sub, sub_mul, he.ortho i j hij]⟩ + +variable {I : Type*} [Fintype I] {e : I → R} + +theorem CompleteOrthogonalIdempotents.of_ker_isNilpotent (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) + (he : ∀ i, IsIdempotentElem (e i)) + (he' : CompleteOrthogonalIdempotents (f ∘ e)) : + CompleteOrthogonalIdempotents e := + of_ker_isNilpotent_of_isMulCentral f h he + (fun _ ↦ Semigroup.mem_center_iff.mpr (mul_comm · _)) he' + +lemma OrthogonalIdempotents.prod_one_sub (he : OrthogonalIdempotents e) : + ∏ i, (1 - e i) = 1 - ∑ i, e i := by + classical + induction' (@Finset.univ I _) using Finset.induction_on with a s has ih + · simp + · suffices (1 - e a) * (1 - ∑ i in s, e i) = 1 - (e a + ∑ i in s, e i) by simp [*] + have : e a * ∑ i in s, e i = 0 := by + rw [Finset.mul_sum, ← Finset.sum_const_zero (s := s)] + exact Finset.sum_congr rfl fun j hj ↦ he.ortho a j (fun e ↦ has (e ▸ hj)) + rw [sub_mul, mul_sub, mul_sub, one_mul, mul_one, one_mul, this, sub_zero, sub_sub, add_comm] + +lemma CompleteOrthogonalIdempotents.prod_one_sub (he : CompleteOrthogonalIdempotents e) : + ∏ i, (1 - e i) = 0 := by + rw [he.1.prod_one_sub, he.complete, sub_self] + +lemma CompleteOrthogonalIdempotents.of_prod_one_sub + (he : OrthogonalIdempotents e) (he' : ∏ i, (1 - e i) = 0) : + CompleteOrthogonalIdempotents e where + __ := he + complete := by rwa [he.prod_one_sub, sub_eq_zero, eq_comm] at he' + +/-- A family of complete orthogonal idempotents induces an isomorphism `R ≃+* ∏ R ⧸ ⟨1 - eᵢ⟩` -/ +lemma CompleteOrthogonalIdempotents.bijective_pi (he : CompleteOrthogonalIdempotents e) : + Function.Bijective (Pi.ringHom fun i ↦ Ideal.Quotient.mk (Ideal.span {1 - e i})) := by + classical + refine ⟨?_, he.1.surjective_pi⟩ + rw [injective_iff_map_eq_zero] + intro x hx + simp [Function.funext_iff, Ideal.Quotient.eq_zero_iff_mem, Ideal.mem_span_singleton] at hx + suffices ∀ s : Finset I, (∏ i in s, (1 - e i)) * x = x by + rw [← this Finset.univ, he.prod_one_sub, zero_mul] + refine fun s ↦ Finset.induction_on s (by simp) ?_ + intros a s has e' + suffices (1 - e a) * x = x by simp [has, mul_assoc, e', this] + obtain ⟨c, rfl⟩ := hx a + rw [← mul_assoc, (he.idem a).one_sub.eq] + +lemma CompleteOrthogonalIdempotents.bijective_pi' (he : CompleteOrthogonalIdempotents (1 - e ·)) : + Function.Bijective (Pi.ringHom fun i ↦ Ideal.Quotient.mk (Ideal.span {e i})) := by + obtain ⟨e', rfl, h⟩ : ∃ e' : I → R, (e' = e) ∧ Function.Bijective (Pi.ringHom fun i ↦ + Ideal.Quotient.mk (Ideal.span {e' i})) := ⟨_, funext (by simp), he.bijective_pi⟩ + exact h + +lemma bijective_pi_of_isIdempotentElem (e : I → R) + (he : ∀ i, IsIdempotentElem (e i)) + (he₁ : ∀ i j, i ≠ j → (1 - e i) * (1 - e j) = 0) (he₂ : ∏ i, e i = 0) : + Function.Bijective (Pi.ringHom fun i ↦ Ideal.Quotient.mk (Ideal.span {e i})) := + (CompleteOrthogonalIdempotents.of_prod_one_sub + ⟨fun i ↦ (he i).one_sub, he₁⟩ (by simpa using he₂)).bijective_pi' + +end CommRing diff --git a/Mathlib/RingTheory/LocalProperties.lean b/Mathlib/RingTheory/LocalProperties.lean index dc4a47e3db568..d676783897cb7 100644 --- a/Mathlib/RingTheory/LocalProperties.lean +++ b/Mathlib/RingTheory/LocalProperties.lean @@ -162,6 +162,15 @@ theorem RingHom.ofLocalizationSpanTarget_iff_finite : obtain ⟨s', h₁, h₂⟩ := (Ideal.span_eq_top_iff_finite s).mp hs exact h s' h₂ fun x => hs' ⟨_, h₁ x.prop⟩ +theorem RingHom.HoldsForLocalizationAway.of_bijective + (H : RingHom.HoldsForLocalizationAway P) (hf : Function.Bijective f) : + P f := by + letI := f.toAlgebra + have := IsLocalization.at_units (.powers (1 : R)) (by simp) + have := IsLocalization.isLocalization_of_algEquiv (.powers (1 : R)) + (AlgEquiv.ofBijective (Algebra.ofId R S) hf) + exact H _ 1 + variable {P f R' S'} theorem RingHom.PropertyIsLocal.respectsIso (hP : RingHom.PropertyIsLocal @P) : diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index 7b04a4219ba33..ba35a7d6cb2be 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -674,6 +674,13 @@ theorem algEquiv_mk' (x : R) (y : M) : algEquiv M S Q (mk' S x y) = mk' Q x y := -- Porting note (#10618): removed `simp`, `simp` can prove it theorem algEquiv_symm_mk' (x : R) (y : M) : (algEquiv M S Q).symm (mk' Q x y) = mk' S x y := by simp +variable (M) in +protected lemma bijective (f : S →+* Q) (hf : f.comp (algebraMap R S) = algebraMap R Q) : + Function.Bijective f := + (show f = IsLocalization.algEquiv M S Q by + apply IsLocalization.ringHom_ext M; rw [hf]; ext; simp) ▸ + (IsLocalization.algEquiv M S Q).toEquiv.bijective + end AlgEquiv section at_units @@ -1083,7 +1090,9 @@ section variable [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] variable (S Rₘ Sₘ) -theorem IsLocalization.map_units_map_submonoid (y : M) : IsUnit (algebraMap R Sₘ y) := by +theorem IsLocalization.map_units_map_submonoid + [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] (y : M) : + IsUnit (algebraMap R Sₘ y) := by rw [IsScalarTower.algebraMap_apply _ S] exact IsLocalization.map_units Sₘ ⟨algebraMap R S y, Algebra.mem_algebraMapSubmonoid_of_mem y⟩ diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index e71996380b7f5..64199851be828 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -125,7 +125,9 @@ noncomputable abbrev toField : Field K where mul_inv_cancel := IsFractionRing.mul_inv_cancel A inv_zero := show IsFractionRing.inv A (0 : K) = 0 by rw [IsFractionRing.inv]; exact dif_pos rfl nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun a x => rfl lemma surjective_iff_isField [IsDomain R] : Function.Surjective (algebraMap R K) ↔ IsField R where mp h := (RingEquiv.ofBijective (algebraMap R K) @@ -215,7 +217,7 @@ noncomputable def map {A B K L : Type*} [CommRing A] [CommRing B] [IsDomain B] [ fields of fractions `K ≃+* L`. -/ noncomputable def fieldEquivOfRingEquiv [Algebra B L] [IsFractionRing B L] (h : A ≃+* B) : K ≃+* L := - ringEquivOfRingEquiv K L h + ringEquivOfRingEquiv (M := nonZeroDivisors A) K (T := nonZeroDivisors B) L h (by ext b show b ∈ h.toEquiv '' _ ↔ _ diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 350df0559d7ee..fc8d3b40ea966 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -68,7 +68,8 @@ 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⟩⟩ -theorem map_comap (J : Ideal S) : Ideal.map (algebraMap R S) (Ideal.comap (algebraMap R S) J) = J := +theorem map_comap [IsLocalization M S] (J : Ideal S) : + Ideal.map (algebraMap R S) (Ideal.comap (algebraMap R S) J) = J := le_antisymm (Ideal.map_le_iff_le_comap.2 le_rfl) fun x hJ => by obtain ⟨r, s, hx⟩ := mk'_surjective M x rw [← hx] at hJ ⊢ diff --git a/Mathlib/RingTheory/MvPolynomial/Basic.lean b/Mathlib/RingTheory/MvPolynomial/Basic.lean index d943124d0e6cb..b63e82050c796 100644 --- a/Mathlib/RingTheory/MvPolynomial/Basic.lean +++ b/Mathlib/RingTheory/MvPolynomial/Basic.lean @@ -178,6 +178,9 @@ lemma algebraMap_def : algebraMap (MvPolynomial σ R) (MvPolynomial σ S) = MvPolynomial.map (algebraMap R S) := rfl +instance : IsScalarTower R (MvPolynomial σ R) (MvPolynomial σ S) := + IsScalarTower.of_algebraMap_eq' (by ext; simp) + end Algebra end MvPolynomial diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 3e3d530b550e2..2cd4d9fdd0bf9 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -108,6 +108,11 @@ theorem isHomogeneous_monomial {d : σ →₀ ℕ} (r : R) {n : ℕ} (hn : d.deg variable (σ) +theorem totalDegree_eq_zero_iff (p : MvPolynomial σ R) : + p.totalDegree = 0 ↔ ∀ (m : σ →₀ ℕ) (_ : m ∈ p.support) (x : σ), m x = 0 := by + rw [← weightedTotalDegree_one, weightedTotalDegree_eq_zero_iff _ p] + exact nonTorsionWeight_of (Function.const σ one_ne_zero) + theorem totalDegree_zero_iff_isHomogeneous {p : MvPolynomial σ R} : p.totalDegree = 0 ↔ IsHomogeneous p 0 := by rw [← weightedTotalDegree_one, @@ -427,14 +432,6 @@ instance HomogeneousSubmodule.gcommSemiring : SetLike.GradedMonoid (homogeneousS one_mem := isHomogeneous_one σ R mul_mem _ _ _ _ := IsHomogeneous.mul -open DirectSum - -noncomputable example : CommSemiring (⨁ i, homogeneousSubmodule σ R i) := - inferInstance - -noncomputable example : Algebra R (⨁ i, homogeneousSubmodule σ R i) := - inferInstance - end IsHomogeneous noncomputable section @@ -453,6 +450,10 @@ open Finset Finsupp variable [CommSemiring R] (n : ℕ) (φ ψ : MvPolynomial σ R) +theorem homogeneousComponent_mem : + homogeneousComponent n φ ∈ homogeneousSubmodule σ R n := + weightedHomogeneousComponent_mem _ φ n + theorem coeff_homogeneousComponent (d : σ →₀ ℕ) : coeff d (homogeneousComponent n φ) = if d.degree = n then coeff d φ else 0 := by rw [degree_eq_weight_one] @@ -494,12 +495,48 @@ theorem sum_homogeneousComponent : simpa [coeff_sum, coeff_homogeneousComponent] exact fun h => (coeff_eq_zero_of_totalDegree_lt h).symm -theorem homogeneousComponent_homogeneous_polynomial (m n : ℕ) (p : MvPolynomial σ R) - (h : p ∈ homogeneousSubmodule σ R n) : homogeneousComponent m p = if m = n then p else 0 := by - convert weightedHomogeneousComponent_weighted_homogeneous_polynomial m n p h +theorem homogeneousComponent_of_mem {m n : ℕ} {p : MvPolynomial σ R} + (h : p ∈ homogeneousSubmodule σ R n) : + homogeneousComponent m p = if m = n then p else 0 := + weightedHomogeneousComponent_of_mem h end HomogeneousComponent end +noncomputable section GradedAlgebra + +/-- The homogeneous submodules form a graded ring. +This instance is used by `DirectSum.commSemiring` and `DirectSum.algebra`. -/ +lemma HomogeneousSubmodule.gradedMonoid : + SetLike.GradedMonoid (homogeneousSubmodule σ R) := + WeightedHomogeneousSubmodule.gradedMonoid + +/-- The decomposition of `MvPolynomial σ R` into homogeneous submodules. -/ +abbrev decomposition : + DirectSum.Decomposition (homogeneousSubmodule σ R) := + weightedDecomposition R (1 : σ → ℕ) + +/-- `MvPolynomial σ R` as a graded algebra, graded by the degree. +We do not make this a global instance because one may want to consider a different +graded algebra structure on `MvPolynomial σ R`, induced by another weight function. +To make it a local instance, you may use +`attribute [local instance] MvPolynomial.gradedAlgebra`. +-/ +abbrev gradedAlgebra : GradedAlgebra (homogeneousSubmodule σ R) := + weightedGradedAlgebra R (1 : σ → ℕ) + +theorem decomposition.decompose'_apply (φ : MvPolynomial σ R) (i : ℕ) : + (decomposition.decompose' φ i : MvPolynomial σ R) = homogeneousComponent i φ := + weightedDecomposition.decompose'_apply R _ φ i + +theorem decomposition.decompose'_eq : + decomposition.decompose' = fun φ : MvPolynomial σ R => + DirectSum.mk (fun i : ℕ => ↥(homogeneousSubmodule σ R i)) (φ.support.image Finsupp.degree) + fun m => ⟨homogeneousComponent m φ, homogeneousComponent_mem m φ⟩ := by + rw [degree_eq_weight_one] + rfl + +end GradedAlgebra + end MvPolynomial diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index ec5166222e2ac..380d5c0dc3152 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -3,10 +3,12 @@ Copyright (c) 2022 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.Algebra.DirectSum.Decomposition import Mathlib.Algebra.GradedMonoid -import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.MvPolynomial.Basic +import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Data.Finsupp.Weight +import Mathlib.RingTheory.GradedAlgebra.Basic /-! # Weighted homogeneous polynomials @@ -276,16 +278,16 @@ theorem weighted_total_degree [SemilatticeSup M] {w : σ → M} (hφ : IsWeighte replace hd := Finsupp.mem_support_iff.mpr hd apply Finset.le_sup hd +end IsWeightedHomogeneous + +variable {R} + /-- The weighted homogeneous submodules form a graded monoid. -/ -instance WeightedHomogeneousSubmodule.gcomm_monoid {w : σ → M} : +lemma WeightedHomogeneousSubmodule.gradedMonoid {w : σ → M} : SetLike.GradedMonoid (weightedHomogeneousSubmodule R w) where one_mem := isWeightedHomogeneous_one R w mul_mem _ _ _ _ := IsWeightedHomogeneous.mul -end IsWeightedHomogeneous - -variable {R} - /-- `weightedHomogeneousComponent w n φ` is the part of `φ` that is weighted homogeneous of weighted degree `n`, with respect to the weights `w`. See `sum_weightedHomogeneousComponent` for the statement that `φ` is equal to the sum @@ -415,8 +417,8 @@ theorem IsWeightedHomogeneous.weightedHomogeneousComponent_ne {m : M} (n : M) · rw [hp zero_coeff]; exact Ne.symm hn /-- The weighted homogeneous components of a weighted homogeneous polynomial. -/ -theorem weightedHomogeneousComponent_weighted_homogeneous_polynomial [DecidableEq M] (m n : M) - (p : MvPolynomial σ R) (h : p ∈ weightedHomogeneousSubmodule R w n) : +theorem weightedHomogeneousComponent_of_mem [DecidableEq M] {m n : M} + {p : MvPolynomial σ R} (h : p ∈ weightedHomogeneousSubmodule R w n) : weightedHomogeneousComponent w m p = if m = n then p else 0 := by simp only [mem_weightedHomogeneousSubmodule] at h ext x @@ -430,6 +432,64 @@ theorem weightedHomogeneousComponent_weighted_homogeneous_polynomial [DecidableE · rfl · simp only [coeff_zero] +theorem weightedHomogeneousComponent_of_isWeightedHomogeneous_same + {m : M} {p : MvPolynomial σ R} (hp : IsWeightedHomogeneous w p m) : + weightedHomogeneousComponent w m p = p := by + classical + ext x + rw [coeff_weightedHomogeneousComponent] + by_cases zero_coeff : coeff x p = 0 + · split_ifs + rfl; rw [zero_coeff] + · rw [hp zero_coeff, if_pos rfl] + +theorem weightedHomogeneousComponent_of_isWeightedHomogeneous_ne + {m n : M} {p : MvPolynomial σ R} (hp : IsWeightedHomogeneous w p m) (hn : n ≠ m) : + weightedHomogeneousComponent w n p = 0 := by + classical + ext x + rw [coeff_weightedHomogeneousComponent] + by_cases zero_coeff : coeff x p = 0 + · split_ifs <;> simp only [zero_coeff, coeff_zero] + · rw [if_neg (by simp only [hp zero_coeff, hn.symm, not_false_eq_true]), coeff_zero] + +variable (R w) + +open DirectSum + +theorem DirectSum.coeLinearMap_eq_dfinsupp_sum [DecidableEq σ] [DecidableEq R] [DecidableEq M] + (x : DirectSum M fun i : M => ↥(weightedHomogeneousSubmodule R w i)) : + (coeLinearMap fun i : M => weightedHomogeneousSubmodule R w i) x = + DFinsupp.sum x (fun _ x => ↑x) := by + rw [_root_.DirectSum.coeLinearMap_eq_dfinsupp_sum] + +theorem DirectSum.coeAddMonoidHom_eq_support_sum [DecidableEq σ] [DecidableEq R] [DecidableEq M] + (x : DirectSum M fun i : M => ↥(weightedHomogeneousSubmodule R w i)) : + (DirectSum.coeAddMonoidHom fun i : M => weightedHomogeneousSubmodule R w i) x = + DFinsupp.sum x (fun _ x => ↑x) := + DirectSum.coeLinearMap_eq_dfinsupp_sum R w x + +theorem DirectSum.coeLinearMap_eq_finsum [DecidableEq M] + (x : DirectSum M fun i : M => ↥(weightedHomogeneousSubmodule R w i)) : + (DirectSum.coeLinearMap fun i : M => weightedHomogeneousSubmodule R w i) x = + finsum fun m => x m := by + classical + rw [DirectSum.coeLinearMap_eq_dfinsupp_sum, DFinsupp.sum, finsum_eq_sum_of_support_subset] + apply DirectSum.support_subset + +theorem weightedHomogeneousComponent_directSum [DecidableEq M] + (x : DirectSum M fun i : M => ↥(weightedHomogeneousSubmodule R w i)) (m : M) : + (weightedHomogeneousComponent w m) + ((DirectSum.coeLinearMap fun i : M => weightedHomogeneousSubmodule R w i) x) = x m := by + classical + rw [DirectSum.coeLinearMap_eq_dfinsupp_sum, DFinsupp.sum, map_sum] + convert @Finset.sum_eq_single M (MvPolynomial σ R) _ (DFinsupp.support x) _ m _ _ + · rw [weightedHomogeneousComponent_of_isWeightedHomogeneous_same (x m).prop] + · intro n _ hmn + rw [weightedHomogeneousComponent_of_isWeightedHomogeneous_ne (x n).prop hmn.symm] + · rw [DFinsupp.not_mem_support_iff] + intro hm; rw [hm, Submodule.coe_zero, map_zero] + end WeightedHomogeneousComponent end AddCommMonoid @@ -514,4 +574,78 @@ theorem weightedTotalDegree_eq_zero_iff (hw : NonTorsionWeight w) (p : MvPolynom end CanonicallyLinearOrderedMonoid +section GradedAlgebra + +/- Here, given a weight `w : σ → M`, where `M` is an additive and commutative monoid, we endow the + ring of multivariate polynomials `MvPolynomial σ R` with the structure of a graded algebra -/ +variable (w : σ → M) [AddCommMonoid M] + +theorem weightedHomogeneousComponent_eq_zero_of_not_mem [DecidableEq M] + (φ : MvPolynomial σ R) (i : M) (hi : i ∉ Finset.image (weight w) φ.support) : + weightedHomogeneousComponent w i φ = 0 := by + apply weightedHomogeneousComponent_eq_zero' + simp only [Finset.mem_image, mem_support_iff, ne_eq, exists_prop, not_exists, not_and] at hi + exact fun m hm ↦ hi m (mem_support_iff.mp hm) + +variable (R) + +/-- The `decompose'` argument of `weightedDecomposition`. -/ +def decompose' [DecidableEq M] := fun φ : MvPolynomial σ R => + DirectSum.mk (fun i : M => ↥(weightedHomogeneousSubmodule R w i)) + (Finset.image (weight w) φ.support) fun m => + ⟨weightedHomogeneousComponent w m φ, weightedHomogeneousComponent_mem w φ m⟩ + +theorem decompose'_apply [DecidableEq M] (φ : MvPolynomial σ R) (m : M) : + (decompose' R w φ m : MvPolynomial σ R) = weightedHomogeneousComponent w m φ := by + rw [decompose'] + by_cases hm : m ∈ Finset.image (weight w) φ.support + · simp only [DirectSum.mk_apply_of_mem hm, Subtype.coe_mk] + · rw [DirectSum.mk_apply_of_not_mem hm, Submodule.coe_zero, + weightedHomogeneousComponent_eq_zero_of_not_mem w φ m hm] + +/-- Given a weight `w`, the decomposition of `MvPolynomial σ R` into weighted homogeneous +submodules -/ +def weightedDecomposition [DecidableEq M] : + DirectSum.Decomposition (weightedHomogeneousSubmodule R w) where + decompose' := decompose' R w + left_inv φ := by + classical + conv_rhs => rw [← sum_weightedHomogeneousComponent w φ] + rw [← DirectSum.sum_support_of (decompose' R w φ)] + simp only [DirectSum.coeAddMonoidHom_of, MvPolynomial.coeff_sum, map_sum, + finsum_eq_sum _ (weightedHomogeneousComponent_finsupp φ)] + apply Finset.sum_congr _ (fun m _ ↦ by rw [decompose'_apply]) + ext m + simp only [DFinsupp.mem_support_toFun, ne_eq, Set.Finite.mem_toFinset, Function.mem_support, + not_iff_not] + conv_lhs => rw [← Subtype.coe_inj] + rw [decompose'_apply, Submodule.coe_zero] + right_inv x := by + classical + apply DFinsupp.ext + intro m + rw [← Subtype.coe_inj, decompose'_apply] + exact weightedHomogeneousComponent_directSum R w x m + + +/-- Given a weight, `MvPolynomial` as a graded algebra -/ +def weightedGradedAlgebra [DecidableEq M] : + GradedAlgebra (weightedHomogeneousSubmodule R w) where + toDecomposition := weightedDecomposition R w + toGradedMonoid := WeightedHomogeneousSubmodule.gradedMonoid + +theorem weightedDecomposition.decompose'_eq [DecidableEq M] : + (weightedDecomposition R w).decompose' = fun φ : MvPolynomial σ R => + DirectSum.mk (fun i : M => ↥(weightedHomogeneousSubmodule R w i)) + (Finset.image (weight w) φ.support) fun m => + ⟨weightedHomogeneousComponent w m φ, weightedHomogeneousComponent_mem w φ m⟩ := rfl + +theorem weightedDecomposition.decompose'_apply [DecidableEq M] + (φ : MvPolynomial σ R) (m : M) : + ((weightedDecomposition R w).decompose' φ m : MvPolynomial σ R) = + weightedHomogeneousComponent w m φ := + MvPolynomial.decompose'_apply R w φ m + +end GradedAlgebra + end MvPolynomial diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index ce3e700700473..3480196f77b5d 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -106,9 +106,9 @@ namespace Commute section Semiring -variable [Semiring R] (h_comm : Commute x y) +variable [Semiring R] -theorem add_pow_eq_zero_of_add_le_succ_of_pow_eq_zero {m n k : ℕ} +theorem add_pow_eq_zero_of_add_le_succ_of_pow_eq_zero (h_comm : Commute x y) {m n k : ℕ} (hx : x ^ m = 0) (hy : y ^ n = 0) (h : m + n ≤ k + 1) : (x + y) ^ k = 0 := by rw [h_comm.add_pow'] @@ -120,12 +120,13 @@ theorem add_pow_eq_zero_of_add_le_succ_of_pow_eq_zero {m n k : ℕ} rw [pow_eq_zero_of_le ?_ hy, mul_zero] linarith [Finset.mem_antidiagonal.mp hij] -theorem add_pow_add_eq_zero_of_pow_eq_zero {m n : ℕ} +theorem add_pow_add_eq_zero_of_pow_eq_zero (h_comm : Commute x y) {m n : ℕ} (hx : x ^ m = 0) (hy : y ^ n = 0) : (x + y) ^ (m + n - 1) = 0 := h_comm.add_pow_eq_zero_of_add_le_succ_of_pow_eq_zero hx hy <| by rw [← Nat.sub_le_iff_le_add] -theorem isNilpotent_add (hx : IsNilpotent x) (hy : IsNilpotent y) : IsNilpotent (x + y) := by +theorem isNilpotent_add (h_comm : Commute x y) (hx : IsNilpotent x) (hy : IsNilpotent y) : + IsNilpotent (x + y) := by obtain ⟨n, hn⟩ := hx obtain ⟨m, hm⟩ := hy exact ⟨_, add_pow_add_eq_zero_of_pow_eq_zero h_comm hn hm⟩ @@ -144,14 +145,14 @@ protected lemma isNilpotent_sum {ι : Type*} {s : Finset ι} {f : ι → R} · exact ih (fun i hi ↦ hnp i (by simp [hi])) (fun i j hi hj ↦ h_comm i j (by simp [hi]) (by simp [hj])) -protected lemma isNilpotent_mul_left_iff (hy : y ∈ nonZeroDivisorsLeft R) : +protected lemma isNilpotent_mul_left_iff (h_comm : Commute x y) (hy : y ∈ nonZeroDivisorsLeft R) : IsNilpotent (x * y) ↔ IsNilpotent x := by refine ⟨?_, h_comm.isNilpotent_mul_left⟩ rintro ⟨k, hk⟩ rw [mul_pow h_comm] at hk exact ⟨k, (nonZeroDivisorsLeft R).pow_mem hy k _ hk⟩ -protected lemma isNilpotent_mul_right_iff (hx : x ∈ nonZeroDivisorsRight R) : +protected lemma isNilpotent_mul_right_iff (h_comm : Commute x y) (hx : x ∈ nonZeroDivisorsRight R) : IsNilpotent (x * y) ↔ IsNilpotent y := by refine ⟨?_, h_comm.isNilpotent_mul_right⟩ rintro ⟨k, hk⟩ @@ -162,9 +163,10 @@ end Semiring section Ring -variable [Ring R] (h_comm : Commute x y) +variable [Ring R] -theorem isNilpotent_sub (hx : IsNilpotent x) (hy : IsNilpotent y) : IsNilpotent (x - y) := by +theorem isNilpotent_sub (h_comm : Commute x y) (hx : IsNilpotent x) (hy : IsNilpotent y) : + IsNilpotent (x - y) := by rw [← neg_right_iff] at h_comm rw [← isNilpotent_neg_iff] at hy rw [sub_eq_add_neg] diff --git a/Mathlib/RingTheory/OreLocalization/Basic.lean b/Mathlib/RingTheory/OreLocalization/Basic.lean index c7e7c1e601599..24690095025f6 100644 --- a/Mathlib/RingTheory/OreLocalization/Basic.lean +++ b/Mathlib/RingTheory/OreLocalization/Basic.lean @@ -459,13 +459,12 @@ section UMP variable {T : Type*} [Monoid T] variable (f : R →* T) (fS : S →* Units T) -variable (hf : ∀ s : S, f s = fS s) /-- The universal lift from a morphism `R →* T`, which maps elements of `S` to units of `T`, to a morphism `R[S⁻¹] →* T`. -/ @[to_additive "The universal lift from a morphism `R →+ T`, which maps elements of `S` to additive-units of `T`, to a morphism `AddOreLocalization R S →+ T`."] -def universalMulHom : R[S⁻¹] →* T where +def universalMulHom (hf : ∀ s : S, f s = fS s) : R[S⁻¹] →* T where -- Porting note(#12129): additional beta reduction needed toFun x := x.liftExpand (fun r s => ((fS s)⁻¹ : Units T) * f r) fun r t s ht => by @@ -489,6 +488,8 @@ def universalMulHom : R[S⁻¹] →* T where one_mul, ← hf, ← mul_assoc, ← map_mul _ _ r₁, ha, map_mul, hf s₂, mul_assoc, ← mul_assoc (fS s₂ : T), (fS s₂).mul_inv, one_mul] +variable (hf : ∀ s : S, f s = fS s) + @[to_additive] theorem universalMulHom_apply {r : R} {s : S} : universalMulHom f fS hf (r /ₒ s) = ((fS s)⁻¹ : Units T) * f r := diff --git a/Mathlib/RingTheory/OreLocalization/Ring.lean b/Mathlib/RingTheory/OreLocalization/Ring.lean index 7ba5fa7d79701..db7f26d3101ad 100644 --- a/Mathlib/RingTheory/OreLocalization/Ring.lean +++ b/Mathlib/RingTheory/OreLocalization/Ring.lean @@ -275,7 +275,9 @@ instance : DivisionRing R[R⁰⁻¹] where mul_inv_cancel := OreLocalization.mul_inv_cancel inv_zero := OreLocalization.inv_zero nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun q a => rfl end DivisionRing diff --git a/Mathlib/RingTheory/PiTensorProduct.lean b/Mathlib/RingTheory/PiTensorProduct.lean index 2efd62c034d68..ed8ee78470106 100644 --- a/Mathlib/RingTheory/PiTensorProduct.lean +++ b/Mathlib/RingTheory/PiTensorProduct.lean @@ -208,7 +208,7 @@ def liftAlgHom {S : Type*} [Semiring S] [Algebra R S] @[simp] lemma tprod_noncommProd {κ : Type*} (s : Finset κ) (x : κ → Π i, A i) (hx) : tprod R (s.noncommProd x hx) = s.noncommProd (fun k => tprod R (x k)) (hx.imp fun _ _ => Commute.tprod) := - Finset.noncommProd_map s x _ (tprodMonoidHom R) + Finset.map_noncommProd s x _ (tprodMonoidHom R) /-- To show two algebra morphisms from finite tensor products are equal, it suffices to show that they agree on elements of the form $1 ⊗ ⋯ ⊗ a ⊗ 1 ⊗ ⋯$. -/ diff --git a/Mathlib/RingTheory/Polynomial/Bernstein.lean b/Mathlib/RingTheory/Polynomial/Bernstein.lean index c426b88109798..c55621cd2dcaf 100644 --- a/Mathlib/RingTheory/Polynomial/Bernstein.lean +++ b/Mathlib/RingTheory/Polynomial/Bernstein.lean @@ -247,7 +247,7 @@ theorem linearIndependent_aux (n k : ℕ) (h : k ≤ n + 1) : rw [this] exact (iterate_derivative_at_1_ne_zero ℚ n k h).symm refine span_induction m ?_ ?_ ?_ ?_ - · simp + · simp only [Set.mem_range, forall_exists_index, forall_apply_eq_imp_iff] rintro ⟨a, w⟩; simp only [Fin.val_mk] rw [iterate_derivative_at_1_eq_zero_of_lt ℚ n ((tsub_lt_tsub_iff_left_of_le h).mpr w)] · simp diff --git a/Mathlib/RingTheory/Polynomial/Selmer.lean b/Mathlib/RingTheory/Polynomial/Selmer.lean index 181d00e902bd9..765741980303a 100644 --- a/Mathlib/RingTheory/Polynomial/Selmer.lean +++ b/Mathlib/RingTheory/Polynomial/Selmer.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import Mathlib.Algebra.Polynomial.UnitTrinomial +import Mathlib.Analysis.Complex.Polynomial.UnitTrinomial import Mathlib.RingTheory.Polynomial.GaussLemma import Mathlib.Tactic.LinearCombination diff --git a/Mathlib/RingTheory/Prime.lean b/Mathlib/RingTheory/Prime.lean index bf3b35f11c612..eba40fbe15828 100644 --- a/Mathlib/RingTheory/Prime.lean +++ b/Mathlib/RingTheory/Prime.lean @@ -5,8 +5,8 @@ Authors: Chris Hughes -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Order.Group.Unbundled.Abs /-! # Prime elements in rings diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index 71061a58ad792..e672235438cc2 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -364,6 +364,37 @@ theorem mem_compl_zeroLocus_iff_not_mem {f : R} {I : PrimeSpectrum R} : I ∈ (zeroLocus {f} : Set (PrimeSpectrum R))ᶜ ↔ f ∉ I.asIdeal := by rw [Set.mem_compl_iff, mem_zeroLocus, Set.singleton_subset_iff]; rfl +section Order + +/-! +## The specialization order + +We endow `PrimeSpectrum R` with a partial order induced from the ideal lattice. +This is exactly the specialization order. +See the corresponding section at `AlgebraicGeometry/PrimeSpectrum/Basic`. +-/ + +instance : PartialOrder (PrimeSpectrum R) := + PartialOrder.lift asIdeal (@PrimeSpectrum.ext _ _) + +@[simp] +theorem asIdeal_le_asIdeal (x y : PrimeSpectrum R) : x.asIdeal ≤ y.asIdeal ↔ x ≤ y := + Iff.rfl + +@[simp] +theorem asIdeal_lt_asIdeal (x y : PrimeSpectrum R) : x.asIdeal < y.asIdeal ↔ x < y := + Iff.rfl + +instance [IsDomain R] : OrderBot (PrimeSpectrum R) where + bot := ⟨⊥, Ideal.bot_prime⟩ + bot_le I := @bot_le _ _ _ I.asIdeal + +instance {R : Type*} [Field R] : Unique (PrimeSpectrum R) where + default := ⊥ + uniq x := PrimeSpectrum.ext ((IsSimpleOrder.eq_bot_or_eq_top _).resolve_right x.2.ne_top) + +end Order + section Noetherian open Submodule diff --git a/Mathlib/RingTheory/RingHom/FinitePresentation.lean b/Mathlib/RingTheory/RingHom/FinitePresentation.lean new file mode 100644 index 0000000000000..21b4cbcffb892 --- /dev/null +++ b/Mathlib/RingTheory/RingHom/FinitePresentation.lean @@ -0,0 +1,182 @@ +/- +Copyright (c) 2024 Christian Merten. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christian Merten +-/ +import Mathlib.RingTheory.Localization.Finiteness +import Mathlib.RingTheory.MvPolynomial.Localization +import Mathlib.RingTheory.RingHom.FiniteType +import Mathlib.RingTheory.Localization.Away.AdjoinRoot + +/-! + +# The meta properties of finitely-presented ring homomorphisms. + +The main result is `RingHom.finitePresentation_isLocal`. + +-/ + +open scoped Pointwise TensorProduct + +namespace RingHom + +attribute [local instance] MvPolynomial.algebraMvPolynomial + +/-- Being finitely-presented is preserved by localizations. -/ +theorem finitePresentation_localizationPreserves : LocalizationPreserves @FinitePresentation := by + introv R hf + letI := f.toAlgebra + letI := ((algebraMap S S').comp f).toAlgebra + let f' : R' →+* S' := IsLocalization.map S' f M.le_comap_map + letI := f'.toAlgebra + haveI : IsScalarTower R R' S' := + IsScalarTower.of_algebraMap_eq' (IsLocalization.map_comp M.le_comap_map).symm + obtain ⟨n, g, hgsurj, hgker⟩ := hf + let MX : Submonoid (MvPolynomial (Fin n) R) := + Algebra.algebraMapSubmonoid (MvPolynomial (Fin n) R) M + haveI : IsLocalization MX (MvPolynomial (Fin n) R') := + inferInstanceAs <| IsLocalization (M.map MvPolynomial.C) (MvPolynomial (Fin n) R') + haveI : IsScalarTower R S S' := IsScalarTower.of_algebraMap_eq' rfl + haveI : IsLocalization (Algebra.algebraMapSubmonoid S M) S' := + inferInstanceAs <| IsLocalization (M.map f) S' + let g' : MvPolynomial (Fin n) R' →ₐ[R'] S' := IsLocalization.mapₐ M R' _ S' g + let k : RingHom.ker g →ₗ[MvPolynomial (Fin n) R] RingHom.ker g' := + AlgHom.toKerIsLocalization M R' _ S' g + have : IsLocalizedModule MX k := AlgHom.toKerIsLocalization_isLocalizedModule M _ _ _ g + have : Module.Finite (MvPolynomial (Fin n) R) (ker g) := Module.Finite.iff_fg.mpr hgker + exact ⟨n, g', IsLocalization.mapₐ_surjective_of_surjective M R' _ S' g hgsurj, + Module.Finite.iff_fg.mp (Module.Finite.of_isLocalizedModule MX k)⟩ + +/-- Being finitely-presented is stable under composition. -/ +theorem finitePresentation_stableUnderComposition : StableUnderComposition @FinitePresentation := by + introv R hf hg + exact hg.comp hf + +/-- If `R` is a ring, then `Rᵣ` is `R`-finitely-presented for any `r : R`. -/ +theorem finitePresentation_holdsForLocalizationAway : + HoldsForLocalizationAway @FinitePresentation := by + introv R _ + suffices Algebra.FinitePresentation R S by + rw [RingHom.FinitePresentation] + convert this; ext; + rw [Algebra.smul_def]; rfl + exact IsLocalization.Away.finitePresentation r + +/-- +If `S` is an `R`-algebra with a surjection from a finitely-presented `R`-algebra `A`, such that +localized at a spanning set `{ r }` of elements of `A`, `Sᵣ` is finitely-presented, then +`S` is finitely presented. +This is almost `finitePresentation_ofLocalizationSpanTarget`. The difference is, +that here the set `t` generates the unit ideal of `A`, while in the general version, +it only generates a quotient of `A`. +-/ +lemma finitePresentation_ofLocalizationSpanTarget_aux + {R S A : Type*} [CommRing R] [CommRing S] [CommRing A] [Algebra R S] [Algebra R A] + [Algebra.FinitePresentation R A] (f : A →ₐ[R] S) (hf : Function.Surjective f) + (t : Finset A) (ht : Ideal.span (t : Set A) = ⊤) + (H : ∀ g : t, Algebra.FinitePresentation R (Localization.Away (f g))) : + Algebra.FinitePresentation R S := by + apply Algebra.FinitePresentation.of_surjective hf + apply ker_fg_of_localizationSpan t ht + intro g + let f' : Localization.Away g.val →ₐ[R] Localization.Away (f g) := + Localization.awayMapₐ f g.val + have (g : t) : Algebra.FinitePresentation R (Localization.Away g.val) := + haveI : Algebra.FinitePresentation A (Localization.Away g.val) := + IsLocalization.Away.finitePresentation g.val + Algebra.FinitePresentation.trans R A (Localization.Away g.val) + apply Algebra.FinitePresentation.ker_fG_of_surjective f' + exact IsLocalization.Away.mapₐ_surjective_of_surjective _ hf + +/-- Finite-presentation can be checked on a standard covering of the target. -/ +theorem finitePresentation_ofLocalizationSpanTarget : + OfLocalizationSpanTarget @FinitePresentation := by + rw [ofLocalizationSpanTarget_iff_finite] + introv R hs H + classical + letI := f.toAlgebra + replace H : ∀ r : s, Algebra.FinitePresentation R (Localization.Away (r : S)) := by + intro r; simp_rw [RingHom.FinitePresentation] at H; + convert H r; ext; simp_rw [Algebra.smul_def]; rfl + /- + We already know that `S` is of finite type over `R`, so we have a surjection + `MvPolynomial (Fin n) R →ₐ[R] S`. To reason about the kernel, we want to check it on the stalks + of preimages of `s`. But the preimages do not necessarily span `MvPolynomial (Fin n) R`, so + we quotient out by an ideal and apply `finitePresentation_ofLocalizationSpanTarget_aux`. + -/ + have hfintype : Algebra.FiniteType R S := by + apply finiteType_ofLocalizationSpanTarget f s hs + intro r + convert_to Algebra.FiniteType R (Localization.Away r.val) + · rw [RingHom.FiniteType] + constructor <;> intro h <;> convert h <;> ext <;> simp_rw [Algebra.smul_def] <;> rfl + · infer_instance + rw [RingHom.FinitePresentation] + obtain ⟨n, f, hf⟩ := Algebra.FiniteType.iff_quotient_mvPolynomial''.mp hfintype + obtain ⟨l, hl⟩ := (Finsupp.mem_span_iff_total S (s : Set S) 1).mp + (show (1 : S) ∈ Ideal.span (s : Set S) by rw [hs]; trivial) + choose g' hg' using (fun g : s ↦ hf g) + choose h' hh' using (fun g : s ↦ hf (l g)) + let I : Ideal (MvPolynomial (Fin n) R) := Ideal.span { ∑ g : s, g' g * h' g - 1 } + let A := MvPolynomial (Fin n) R ⧸ I + have hfI : ∀ a ∈ I, f a = 0 := by + intro p hp + simp only [Finset.univ_eq_attach, I, Ideal.mem_span_singleton] at hp + obtain ⟨q, rfl⟩ := hp + simp only [map_mul, map_sub, map_sum, map_one, hg', hh'] + erw [Finsupp.total_apply_of_mem_supported S (s := s.attach)] at hl + · rw [← hl] + simp only [Finset.coe_sort_coe, smul_eq_mul, mul_comm, sub_self, mul_zero, zero_mul] + · rintro a - + simp + let f' : A →ₐ[R] S := Ideal.Quotient.liftₐ I f hfI + have hf' : Function.Surjective f' := + Ideal.Quotient.lift_surjective_of_surjective I hfI hf + let t : Finset A := Finset.image (fun g ↦ g' g) Finset.univ + have ht : Ideal.span (t : Set A) = ⊤ := by + rw [Ideal.eq_top_iff_one] + have : ∑ g : { x // x ∈ s }, g' g * h' g = (1 : A) := by + apply eq_of_sub_eq_zero + rw [← map_one (Ideal.Quotient.mk I), ← map_sub, Ideal.Quotient.eq_zero_iff_mem] + apply Ideal.subset_span + simp + simp_rw [← this, Finset.univ_eq_attach, map_sum, map_mul] + refine Ideal.sum_mem _ (fun g _ ↦ Ideal.mul_mem_right _ _ <| Ideal.subset_span ?_) + simp [t] + have : Algebra.FinitePresentation R A := by + apply Algebra.FinitePresentation.quotient + simp only [Finset.univ_eq_attach, I] + exact ⟨{∑ g ∈ s.attach, g' g * h' g - 1}, by simp⟩ + have Ht (g : t) : Algebra.FinitePresentation R (Localization.Away (f' g)) := by + have : ∃ (a : S) (hb : a ∈ s), (Ideal.Quotient.mk I) (g' ⟨a, hb⟩) = g.val := by + simpa [t] using g.property + obtain ⟨r, hr, hrr⟩ := this + simp only [f'] + rw [← hrr, Ideal.Quotient.liftₐ_apply, Ideal.Quotient.lift_mk] + simp_rw [coe_coe] + rw [hg'] + apply H + exact finitePresentation_ofLocalizationSpanTarget_aux f' hf' t ht Ht + +/-- Being finitely-presented is a local property of rings. -/ +theorem finitePresentation_isLocal : PropertyIsLocal @FinitePresentation := + ⟨finitePresentation_localizationPreserves, + finitePresentation_ofLocalizationSpanTarget, finitePresentation_stableUnderComposition, + finitePresentation_holdsForLocalizationAway⟩ + +/-- Being finitely-presented respects isomorphisms. -/ +theorem finitePresentation_respectsIso : RingHom.RespectsIso @RingHom.FinitePresentation := + RingHom.finitePresentation_isLocal.respectsIso + +/-- Being finitely-presented is stable under base change. -/ +theorem finitePresentation_stableUnderBaseChange : StableUnderBaseChange @FinitePresentation := by + apply StableUnderBaseChange.mk + · exact finitePresentation_respectsIso + · introv h + replace h : Algebra.FinitePresentation R T := by + rw [RingHom.FinitePresentation] at h; convert h; ext; simp_rw [Algebra.smul_def]; rfl + suffices Algebra.FinitePresentation S (S ⊗[R] T) by + rw [RingHom.FinitePresentation]; convert this; ext; simp_rw [Algebra.smul_def]; rfl + infer_instance + +end RingHom diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index 2d5e588f9ceda..33b0d6650afea 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -19,8 +19,8 @@ presentation, if its jacobian is invertible. Finally, a standard smooth algebra is an algebra that admits a submersive presentation. While every standard smooth algebra is smooth, the converse does not hold. But if `S` is `R`-smooth, -then `S` is `R`-standard smooth locally on `S`, i.e. there exists a set `{ t }` of `S` that generates -the unit ideal, such that `Sₜ` is `R`-standard smooth for every `t` (TODO, see below). +then `S` is `R`-standard smooth locally on `S`, i.e. there exists a set `{ t }` of `S` that +generates the unit ideal, such that `Sₜ` is `R`-standard smooth for every `t` (TODO, see below). ## Main definitions diff --git a/Mathlib/RingTheory/SurjectiveOnStalks.lean b/Mathlib/RingTheory/SurjectiveOnStalks.lean new file mode 100644 index 0000000000000..e73656d303cd0 --- /dev/null +++ b/Mathlib/RingTheory/SurjectiveOnStalks.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Localization.AtPrime +import Mathlib.RingTheory.LocalRing.RingHom.Basic +import Mathlib.RingTheory.TensorProduct.Basic + +/-! +# Ring Homomorphisms surjective on stalks + +In this file, we prove some results on ring homomorphisms surjective on stalks, to be used in +the development of immersions in algebraic geometry. + +A ring homomorphism `R →+* S` is surjective on stalks if `R_p →+* S_q` is surjective for all pairs +of primes `p = f⁻¹(q)`. We show that this property is stable under composition and base change, and +that surjections and localizations satisfy this. + +-/ + +variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] [Algebra R S] +variable {T : Type*} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable {g : S →+* T} {f : R →+* S} + +namespace RingHom + +/-- +A ring homomorphism `R →+* S` is surjective on stalks if `R_p →+* S_q` is surjective for all pairs +of primes `p = f⁻¹(q)`. +-/ +def SurjectiveOnStalks (f : R →+* S) : Prop := + ∀ (P : Ideal S) (_ : P.IsPrime), Function.Surjective (Localization.localRingHom _ P f rfl) + +/-- +`R_p →+* S_q` is surjective if and only if +every `x : S` is of the form `f x / f r` for some `f r ∉ q`. +This is useful when proving `SurjectiveOnStalks`. +-/ +lemma surjective_localRingHom_iff (P : Ideal S) [P.IsPrime] : + Function.Surjective (Localization.localRingHom _ P f rfl) ↔ + ∀ s : S, ∃ x r : R, ∃ c ∉ P, f r ∉ P ∧ c * f r * s = c * f x := by + constructor + · intro H y + obtain ⟨a, ha⟩ := H (IsLocalization.mk' _ y (1 : P.primeCompl)) + obtain ⟨a, t, rfl⟩ := IsLocalization.mk'_surjective (P.comap f).primeCompl a + rw [Localization.localRingHom_mk', IsLocalization.mk'_eq_iff_eq, + Submonoid.coe_one, one_mul, IsLocalization.eq_iff_exists P.primeCompl] at ha + obtain ⟨c, hc⟩ := ha + simp only [← mul_assoc] at hc + exact ⟨_, _, _, c.2, t.2, hc.symm⟩ + · refine fun H y ↦ Localization.ind (fun ⟨y, t, h⟩ ↦ ?_) y + simp only + obtain ⟨yx, ys, yc, hyc, hy, ey⟩ := H y + obtain ⟨tx, ts, yt, hyt, ht, et⟩ := H t + refine ⟨Localization.mk (yx * ts) ⟨ys * tx, Submonoid.mul_mem _ hy ?_⟩, ?_⟩ + · exact fun H ↦ mul_mem (P.primeCompl.mul_mem hyt ht) h (et ▸ Ideal.mul_mem_left _ yt H) + · simp only [Localization.mk_eq_mk', Localization.localRingHom_mk', map_mul f, + IsLocalization.mk'_eq_iff_eq, IsLocalization.eq_iff_exists P.primeCompl] + refine ⟨⟨yc, hyc⟩ * ⟨yt, hyt⟩, ?_⟩ + simp only [Submonoid.coe_mul] + convert congr($(ey.symm) * $(et)) using 1 <;> ring + +lemma surjectiveOnStalks_iff_forall_ideal : + f.SurjectiveOnStalks ↔ + ∀ I : Ideal S, I ≠ ⊤ → ∀ s : S, ∃ x r : R, ∃ c ∉ I, f r ∉ I ∧ c * f r * s = c * f x := by + simp_rw [SurjectiveOnStalks, surjective_localRingHom_iff] + refine ⟨fun H I hI s ↦ ?_, fun H I hI ↦ H I hI.ne_top⟩ + obtain ⟨M, hM, hIM⟩ := I.exists_le_maximal hI + obtain ⟨x, r, c, hc, hr, e⟩ := H M hM.isPrime s + exact ⟨x, r, c, fun h ↦ hc (hIM h), fun h ↦ hr (hIM h), e⟩ + +lemma surjectiveOnStalks_iff_forall_maximal : + f.SurjectiveOnStalks ↔ ∀ (I : Ideal S) (_ : I.IsMaximal), + Function.Surjective (Localization.localRingHom _ I f rfl) := by + refine ⟨fun H I hI ↦ H I hI.isPrime, fun H I hI ↦ ?_⟩ + simp_rw [surjective_localRingHom_iff] at H ⊢ + intro s + obtain ⟨M, hM, hIM⟩ := I.exists_le_maximal hI.ne_top + obtain ⟨x, r, c, hc, hr, e⟩ := H M hM s + exact ⟨x, r, c, fun h ↦ hc (hIM h), fun h ↦ hr (hIM h), e⟩ + +lemma surjectiveOnStalks_iff_forall_maximal' : + f.SurjectiveOnStalks ↔ ∀ I : Ideal S, I.IsMaximal → + ∀ s : S, ∃ x r : R, ∃ c ∉ I, f r ∉ I ∧ c * f r * s = c * f x := by + simp only [surjectiveOnStalks_iff_forall_maximal, surjective_localRingHom_iff] + +lemma surjectiveOnStalks_of_exists_div (h : ∀ x : S, ∃ r s : R, IsUnit (f s) ∧ f s * x = f r) : + SurjectiveOnStalks f := + surjectiveOnStalks_iff_forall_ideal.mpr fun I hI x ↦ + let ⟨r, s, hr, hr'⟩ := h x + ⟨r, s, 1, by simpa [← Ideal.eq_top_iff_one], fun h ↦ hI (I.eq_top_of_isUnit_mem h hr), by simpa⟩ + +lemma surjectiveOnStalks_of_surjective (h : Function.Surjective f) : + SurjectiveOnStalks f := + surjectiveOnStalks_iff_forall_ideal.mpr fun _ _ s ↦ + let ⟨r, hr⟩ := h s + ⟨r, 1, 1, by simpa [← Ideal.eq_top_iff_one], by simpa [← Ideal.eq_top_iff_one], by simp [hr]⟩ + +variable (S) in +lemma surjectiveOnStalks_of_isLocalization [IsLocalization M S] : + SurjectiveOnStalks (algebraMap R S) := by + refine surjectiveOnStalks_of_exists_div fun s ↦ ?_ + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective M s + exact ⟨x, s, IsLocalization.map_units S s, IsLocalization.mk'_spec' S x s⟩ + +lemma SurjectiveOnStalks.comp (hg : SurjectiveOnStalks g) (hf : SurjectiveOnStalks f) : + SurjectiveOnStalks (g.comp f) := by + intros I hI + have := (hg I hI).comp (hf _ (hI.comap g)) + rwa [← RingHom.coe_comp, ← Localization.localRingHom_comp] at this + +lemma SurjectiveOnStalks.of_comp (hg : SurjectiveOnStalks (g.comp f)) : + SurjectiveOnStalks g := by + intros I hI + have := hg I hI + rw [Localization.localRingHom_comp (I.comap (g.comp f)) (I.comap g) _ _ rfl _ rfl, + RingHom.coe_comp] at this + exact this.of_comp + +open TensorProduct + +/-- +If `R → T` is surjective on stalks, and `J` is some prime of `T`, +then every element `x` in `S ⊗[R] T` satisfies `(1 ⊗ r • t) * x = a ⊗ t` for some +`r : R`, `a : S`, and `t : T` such that `r • t ∉ J`. +-/ +lemma SurjectiveOnStalks.exists_mul_eq_tmul + (hf₂ : (algebraMap R T).SurjectiveOnStalks) + (x : S ⊗[R] T) (J : Ideal T) (hJ : J.IsPrime) : + ∃ (t : T) (r : R) (a : S), (r • t ∉ J) ∧ + (1 : S) ⊗ₜ[R] (r • t) * x = a ⊗ₜ[R] t := by + induction x with + | zero => + exact ⟨1, 1, 0, by rw [one_smul]; exact J.primeCompl.one_mem, + by rw [mul_zero, TensorProduct.zero_tmul]⟩ + | tmul x₁ x₂ => + obtain ⟨y, s, c, hs, hc, e⟩ := (surjective_localRingHom_iff _).mp (hf₂ J hJ) x₂ + simp_rw [Algebra.smul_def] + refine ⟨c, s, y • x₁, J.primeCompl.mul_mem hc hs, ?_⟩ + rw [Algebra.TensorProduct.tmul_mul_tmul, one_mul, mul_comm _ c, e, + TensorProduct.smul_tmul, Algebra.smul_def, mul_comm] + | add x₁ x₂ hx₁ hx₂ => + obtain ⟨t₁, r₁, a₁, hr₁, e₁⟩ := hx₁ + obtain ⟨t₂, r₂, a₂, hr₂, e₂⟩ := hx₂ + have : (r₁ * r₂) • (t₁ * t₂) = (r₁ • t₁) * (r₂ • t₂) := by + simp_rw [← smul_eq_mul]; rw [smul_smul_smul_comm] + refine ⟨t₁ * t₂, r₁ * r₂, r₂ • a₁ + r₁ • a₂, this.symm ▸ J.primeCompl.mul_mem hr₁ hr₂, ?_⟩ + rw [this, ← one_mul (1 : S), ← Algebra.TensorProduct.tmul_mul_tmul, mul_add, mul_comm (_ ⊗ₜ _), + mul_assoc, e₁, Algebra.TensorProduct.tmul_mul_tmul, one_mul, smul_mul_assoc, + ← TensorProduct.smul_tmul, mul_comm (_ ⊗ₜ _), mul_assoc, e₂, + Algebra.TensorProduct.tmul_mul_tmul, one_mul, smul_mul_assoc, ← TensorProduct.smul_tmul, + TensorProduct.add_tmul, mul_comm t₁ t₂] + +lemma SurjectiveOnStalks.baseChange + (hf : (algebraMap R T).SurjectiveOnStalks) : + (algebraMap S (S ⊗[R] T)).SurjectiveOnStalks := by + let g : T →+* S ⊗[R] T := Algebra.TensorProduct.includeRight.toRingHom + intros J hJ + rw [surjective_localRingHom_iff] + intro x + obtain ⟨t, r, a, ht, e⟩ := hf.exists_mul_eq_tmul x (J.comap g) inferInstance + refine ⟨a, algebraMap _ _ r, 1 ⊗ₜ (r • t), ht, ?_, ?_⟩ + · intro H + simp only [Algebra.algebraMap_eq_smul_one (A := S), Algebra.TensorProduct.algebraMap_apply, + Algebra.id.map_eq_id, id_apply, smul_tmul, ← Algebra.algebraMap_eq_smul_one (A := T)] at H + rw [Ideal.mem_comap, Algebra.smul_def, g.map_mul] at ht + exact ht (J.mul_mem_right _ H) + · simp only [tmul_smul, Algebra.TensorProduct.algebraMap_apply, Algebra.id.map_eq_id, + RingHomCompTriple.comp_apply, Algebra.smul_mul_assoc, Algebra.TensorProduct.tmul_mul_tmul, + one_mul, mul_one, id_apply, ← e] + rw [Algebra.algebraMap_eq_smul_one, ← smul_tmul', smul_mul_assoc] + +lemma surjectiveOnStalks_iff_of_isLocalRingHom [LocalRing S] [IsLocalRingHom f] : + f.SurjectiveOnStalks ↔ Function.Surjective f := by + refine ⟨fun H x ↦ ?_, fun h ↦ surjectiveOnStalks_of_surjective h⟩ + obtain ⟨y, r, c, hc, hr, e⟩ := + (surjective_localRingHom_iff _).mp (H (LocalRing.maximalIdeal _) inferInstance) x + simp only [LocalRing.mem_maximalIdeal, mem_nonunits_iff, not_not] at hc hr + refine ⟨(isUnit_of_map_unit f r hr).unit⁻¹ * y, ?_⟩ + apply hr.mul_right_injective + apply hc.mul_right_injective + simp only [← _root_.map_mul, ← mul_assoc, IsUnit.mul_val_inv, one_mul, e] + +end RingHom diff --git a/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean b/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean index f8cc47a277b27..a06929ffda005 100644 --- a/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean +++ b/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean @@ -203,8 +203,7 @@ noncomputable def algebraTensorAlgEquiv : @[simp] lemma algebraTensorAlgEquiv_tmul (a : A) (p : MvPolynomial σ R) : algebraTensorAlgEquiv R A (a ⊗ₜ p) = a • MvPolynomial.map (algebraMap R A) p := by - simp [algebraTensorAlgEquiv] - rw [Algebra.smul_def] + simp [algebraTensorAlgEquiv, Algebra.smul_def] rfl @[simp] diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index 2d5e0a1af5d64..74c46a749d367 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -125,7 +125,7 @@ theorem trace_eq_trace_adjoin [FiniteDimensional K L] (x : L) : conv => lhs rw [← IntermediateField.AdjoinSimple.algebraMap_gen K x] - rw [← trace_trace (L := K⟮x⟯), trace_algebraMap, LinearMap.map_smul_of_tower] + rw [← trace_trace (S := K⟮x⟯), trace_algebraMap, LinearMap.map_smul_of_tower] variable {K} diff --git a/Mathlib/RingTheory/Trace/Defs.lean b/Mathlib/RingTheory/Trace/Defs.lean index 2352ad9acd1d9..2b04f816b9079 100644 --- a/Mathlib/RingTheory/Trace/Defs.lean +++ b/Mathlib/RingTheory/Trace/Defs.lean @@ -28,8 +28,8 @@ the roots of the minimal polynomial of `s` over `R`. ## Implementation notes Typically, the trace is defined specifically for finite field extensions. -The definition is as general as possible and the assumption that we have -fields or that the extension is finite is added to the lemmas as needed. +The definition is as general as possible and the assumption that the extension is finite +is added to the lemmas as needed. We only define the trace for left multiplication (`Algebra.leftMulMatrix`, i.e. `LinearMap.mulLeft`). @@ -46,7 +46,6 @@ universe u v w z variable {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] variable [Algebra R S] [Algebra R T] -variable {K L : Type*} [Field K] [Field L] [Algebra K L] variable {ι κ : Type w} [Fintype ι] open FiniteDimensional @@ -97,10 +96,11 @@ theorem trace_algebraMap_of_basis (x : R) : trace R S (algebraMap R S x) = Finty (If `L` is not finite-dimensional over `K`, then `trace` and `finrank` return `0`.) -/ @[simp] -theorem trace_algebraMap (x : K) : trace K L (algebraMap K L x) = finrank K L • x := by - by_cases H : ∃ s : Finset L, Nonempty (Basis s K L) +theorem trace_algebraMap [StrongRankCondition R] [Module.Free R S] (x : R) : + trace R S (algebraMap R S x) = finrank R S • x := by + by_cases H : ∃ s : Finset S, Nonempty (Basis s R S) · rw [trace_algebraMap_of_basis H.choose_spec.some, finrank_eq_card_basis H.choose_spec.some] - · simp [trace_eq_zero_of_not_exists_basis K H, finrank_eq_zero_of_not_exists_basis_finset H] + · simp [trace_eq_zero_of_not_exists_basis R H, finrank_eq_zero_of_not_exists_basis_finset H] theorem trace_trace_of_basis [Algebra S T] [IsScalarTower R S T] {ι κ : Type*} [Finite ι] [Finite κ] (b : Basis ι R S) (c : Basis κ S T) (x : T) : @@ -122,14 +122,16 @@ theorem trace_comp_trace_of_basis [Algebra S T] [IsScalarTower R S T] {ι κ : T rw [LinearMap.comp_apply, LinearMap.restrictScalars_apply, trace_trace_of_basis b c] @[simp] -theorem trace_trace [Algebra K T] [Algebra L T] [IsScalarTower K L T] [FiniteDimensional K L] - [FiniteDimensional L T] (x : T) : trace K L (trace L T x) = trace K T x := - trace_trace_of_basis (Basis.ofVectorSpace K L) (Basis.ofVectorSpace L T) x +theorem trace_trace [Algebra S T] [IsScalarTower R S T] + [Module.Free R S] [Module.Finite R S] [Module.Free S T] [Module.Finite S T] (x : T) : + trace R S (trace S T x) = trace R T x := + trace_trace_of_basis (Module.Free.chooseBasis R S) (Module.Free.chooseBasis S T) x @[simp] -theorem trace_comp_trace [Algebra K T] [Algebra L T] [IsScalarTower K L T] [FiniteDimensional K L] - [FiniteDimensional L T] : (trace K L).comp ((trace L T).restrictScalars K) = trace K T := by - ext; rw [LinearMap.comp_apply, LinearMap.restrictScalars_apply, trace_trace] +theorem trace_comp_trace [Algebra S T] [IsScalarTower R S T] + [Module.Free R S] [Module.Finite R S] [Module.Free S T] [Module.Finite S T] : + (trace R S).comp ((trace S T).restrictScalars R) = trace R T := + LinearMap.ext trace_trace @[simp] theorem trace_prod_apply [Module.Free R S] [Module.Free R T] [Module.Finite R S] [Module.Finite R T] diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index 54b45e232cb0b..9ec78d1c04325 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1970,13 +1970,4 @@ lemma factors_multiset_prod_of_irreducible {s : Multiset ℕ} (h : ∀ x : ℕ, rw [Ne, Multiset.prod_eq_zero_iff] exact fun con ↦ not_irreducible_zero (h 0 con) -lemma _root_.induction_on_primes {P : ℕ → Prop} (h₀ : P 0) (h₁ : P 1) - (h : ∀ p a : ℕ, p.Prime → P a → P (p * a)) (n : ℕ) : P n := by - apply UniqueFactorizationMonoid.induction_on_prime - · exact h₀ - · intro n h - rw [Nat.isUnit_iff.1 h] - exact h₁ - · exact fun a p _ hp ↦ h p a hp.nat_prime - end Nat diff --git a/Mathlib/RingTheory/Valuation/Integers.lean b/Mathlib/RingTheory/Valuation/Integers.lean index 232a27d070215..e727743126fed 100644 --- a/Mathlib/RingTheory/Valuation/Integers.lean +++ b/Mathlib/RingTheory/Valuation/Integers.lean @@ -70,17 +70,15 @@ theorem one_of_isUnit' {x : O} (hx : IsUnit x) (H : ∀ x, v (algebraMap O R x) (algebraMap O R).map_mul, v.map_mul] exact mul_le_mul_left' (H (u⁻¹ : Units O)) _ -variable (hv : Integers v O) - -theorem one_of_isUnit {x : O} (hx : IsUnit x) : v (algebraMap O R x) = 1 := +theorem one_of_isUnit (hv : Integers v O) {x : O} (hx : IsUnit x) : v (algebraMap O R x) = 1 := one_of_isUnit' hx hv.map_le_one /-- Let `O` be the integers of the valuation `v` on some commutative ring `R`. For every element `x` in `O`, `x` is a unit in `O` if and only if the image of `x` in `R` is a unit and has valuation 1. -/ -theorem isUnit_of_one {x : O} (hx : IsUnit (algebraMap O R x)) (hvx : v (algebraMap O R x) = 1) : - IsUnit x := +theorem isUnit_of_one (hv : Integers v O) {x : O} (hx : IsUnit (algebraMap O R x)) + (hvx : v (algebraMap O R x) = 1) : IsUnit x := let ⟨u, hu⟩ := hx have h1 : v u ≤ 1 := hu.symm ▸ hv.2 x have h2 : v (u⁻¹ : Rˣ) ≤ 1 := by @@ -91,7 +89,8 @@ theorem isUnit_of_one {x : O} (hx : IsUnit (algebraMap O R x)) (hvx : v (algebra hv.1 <| by rw [RingHom.map_mul, RingHom.map_one, hr1, hr2, Units.inv_mul]⟩, hv.1 <| hr1.trans hu⟩ -theorem le_of_dvd {x y : O} (h : x ∣ y) : v (algebraMap O R y) ≤ v (algebraMap O R x) := by +theorem le_of_dvd (hv : Integers v O) {x y : O} (h : x ∣ y) : + v (algebraMap O R y) ≤ v (algebraMap O R x) := by let ⟨z, hz⟩ := h rw [← mul_one (v (algebraMap O R x)), hz, RingHom.map_mul, v.map_mul] exact mul_le_mul_left' (hv.2 z) _ @@ -103,11 +102,12 @@ end CommRing section Field variable {F : Type u} {Γ₀ : Type v} [Field F] [LinearOrderedCommGroupWithZero Γ₀] -variable {v : Valuation F Γ₀} {O : Type w} [CommRing O] [Algebra O F] (hv : Integers v O) +variable {v : Valuation F Γ₀} {O : Type w} [CommRing O] [Algebra O F] namespace Integers -theorem dvd_of_le {x y : O} (h : v (algebraMap O F x) ≤ v (algebraMap O F y)) : y ∣ x := +theorem dvd_of_le (hv : Integers v O) {x y : O} + (h : v (algebraMap O F x) ≤ v (algebraMap O F y)) : y ∣ x := by_cases (fun hy : algebraMap O F y = 0 => have hx : x = 0 := @@ -121,10 +121,12 @@ theorem dvd_of_le {x y : O} (h : v (algebraMap O F x) ≤ v (algebraMap O F y)) let ⟨z, hz⟩ := hv.3 this ⟨z, hv.1 <| ((algebraMap O F).map_mul y z).symm ▸ hz.symm ▸ (mul_inv_cancel_left₀ hy _).symm⟩ -theorem dvd_iff_le {x y : O} : x ∣ y ↔ v (algebraMap O F y) ≤ v (algebraMap O F x) := +theorem dvd_iff_le (hv : Integers v O) {x y : O} : + x ∣ y ↔ v (algebraMap O F y) ≤ v (algebraMap O F x) := ⟨hv.le_of_dvd, hv.dvd_of_le⟩ -theorem le_iff_dvd {x y : O} : v (algebraMap O F x) ≤ v (algebraMap O F y) ↔ y ∣ x := +theorem le_iff_dvd (hv : Integers v O) {x y : O} : + v (algebraMap O F x) ≤ v (algebraMap O F y) ↔ y ∣ x := ⟨hv.dvd_of_le, hv.le_of_dvd⟩ /-- @@ -132,11 +134,11 @@ This is the special case of `Valuation.Integers.isUnit_of_one` when the valuatio over a field. Let `v` be a valuation on some field `F` and `O` be its integers. For every element `x` in `O`, `x` is a unit in `O` if and only if the image of `x` in `F` has valuation 1. -/ -theorem isUnit_of_one' {x : O} (hvx : v (algebraMap O F x) = 1) : IsUnit x := by +theorem isUnit_of_one' (hv : Integers v O) {x : O} (hvx : v (algebraMap O F x) = 1) : IsUnit x := by refine isUnit_of_one hv (IsUnit.mk0 _ ?_) hvx simp only [← v.ne_zero_iff, hvx, ne_eq, one_ne_zero, not_false_eq_true] -theorem eq_algebraMap_or_inv_eq_algebraMap (x : F) : +theorem eq_algebraMap_or_inv_eq_algebraMap (hv : Integers v O) (x : F) : ∃ a : O, x = algebraMap O F a ∨ x⁻¹ = algebraMap O F a := by rcases val_le_one_or_val_inv_le_one v x with h | h <;> obtain ⟨a, ha⟩ := exists_of_le_one hv h diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 4d1c44e2c8d8a..0dad8f77f0f4f 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -7,7 +7,7 @@ import Mathlib.RingTheory.Valuation.ValuationRing import Mathlib.RingTheory.Localization.AsSubring import Mathlib.Algebra.Ring.Subring.Pointwise import Mathlib.Algebra.Ring.Action.Field -import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic +import Mathlib.RingTheory.PrimeSpectrum import Mathlib.RingTheory.LocalRing.ResidueField.Basic /-! diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index adb171defa729..0492664ca9471 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -173,7 +173,7 @@ theorem map_frobeniusPoly (n : ℕ) : rw [Rat.natCast_div _ _ (map_frobeniusPoly.key₁ p (n - i) j hj)] simp only [Nat.cast_pow, pow_add, pow_one] suffices - (((p ^ (n - i)).choose (j + 1) : ℚ) * (p : ℚ) ^ (j - v p ⟨j + 1, j.succ_pos⟩) * ↑p * (p ^ n : ℚ)) + (((p ^ (n - i)).choose (j + 1) : ℚ) * (p : ℚ) ^ (j - v p ⟨j + 1, j.succ_pos⟩) * p * (p ^ n : ℚ)) = (p : ℚ) ^ j * p * ↑((p ^ (n - i)).choose (j + 1) * p ^ i) * (p : ℚ) ^ (n - i - v p ⟨j + 1, j.succ_pos⟩) by have aux : ∀ k : ℕ, (p : ℚ)^ k ≠ 0 := by diff --git a/Mathlib/RingTheory/WittVector/WittPolynomial.lean b/Mathlib/RingTheory/WittVector/WittPolynomial.lean index 34031f221be88..0ef21ce915858 100644 --- a/Mathlib/RingTheory/WittVector/WittPolynomial.lean +++ b/Mathlib/RingTheory/WittVector/WittPolynomial.lean @@ -194,7 +194,8 @@ noncomputable def xInTermsOfW [Invertible (p : R)] : ℕ → MvPolynomial ℕ R C ((p : R) ^ (i : ℕ)) * xInTermsOfW i ^ p ^ (n - (i : ℕ))) * C ((⅟ p : R) ^ n) theorem xInTermsOfW_eq [Invertible (p : R)] {n : ℕ} : xInTermsOfW p R n = - (X n - ∑ i ∈ range n, C ((p : R) ^ i) * xInTermsOfW p R i ^ p ^ (n - i)) * C ((⅟p : R) ^ n) := by + (X n - ∑ i ∈ range n, C ((p : R) ^ i) * + xInTermsOfW p R i ^ p ^ (n - i)) * C ((⅟p : R) ^ n) := by rw [xInTermsOfW, ← Fin.sum_univ_eq_sum_range] @[simp] diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index c60d9a2ab7be0..080fe52fe5888 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -808,7 +808,7 @@ theorem lift_mk_le_lift_mk_mul_of_lift_mk_preimage_le {α : Type u} {β : Type v /-- The range of an indexed cardinal function, whose outputs live in a higher universe than the inputs, is always bounded above. -/ theorem bddAbove_range {ι : Type u} (f : ι → Cardinal.{max u v}) : BddAbove (Set.range f) := - ⟨_, by + ⟨sum f, by rintro a ⟨i, rfl⟩ exact le_sum f i⟩ diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index dc894d283983b..c8063908425a6 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -632,7 +632,7 @@ theorem aleph0_le_cof {o} : ℵ₀ ≤ cof o ↔ IsLimit o := by rcases zero_or_succ_or_limit o with (rfl | ⟨o, rfl⟩ | l) · simp [not_zero_isLimit, Cardinal.aleph0_ne_zero] · simp [not_succ_isLimit, Cardinal.one_lt_aleph0] - · simp [l] + · simp only [l, iff_true] refine le_of_not_lt fun h => ?_ cases' Cardinal.lt_aleph0.1 h with n e have := cof_cof o @@ -876,7 +876,7 @@ theorem isRegular_succ {c : Cardinal.{u}} (h : ℵ₀ ≤ c) : IsRegular (succ c ⟨h.trans (le_succ c), succ_le_of_lt (by - cases' Quotient.exists_rep (@succ Cardinal _ _ c) with α αe; simp at αe + cases' Quotient.exists_rep (@succ Cardinal _ _ c) with α αe; simp only [mk'_def] at αe rcases ord_eq α with ⟨r, wo, re⟩ have := ord_isLimit (h.trans (le_succ _)) rw [← αe, re] at this ⊢ diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index 3f9998efe53fc..dcbb87fc05ba1 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -991,27 +991,26 @@ theorem mk_perm_eq_self_power : #(Equiv.Perm α) = #α ^ #α := theorem mk_perm_eq_two_power : #(Equiv.Perm α) = 2 ^ #α := by rw [mk_perm_eq_self_power, power_self_eq (aleph0_le_mk α)] -variable (leq : lift.{v} #α = lift.{u} #β') (eq : #α = #β) - -theorem mk_equiv_eq_arrow_of_lift_eq : #(α ≃ β') = #(α → β') := by +theorem mk_equiv_eq_arrow_of_lift_eq (leq : lift.{v} #α = lift.{u} #β') : + #(α ≃ β') = #(α → β') := by obtain ⟨e⟩ := lift_mk_eq'.mp leq have e₁ := lift_mk_eq'.mpr ⟨.equivCongr (.refl α) e⟩ have e₂ := lift_mk_eq'.mpr ⟨.arrowCongr (.refl α) e⟩ rw [lift_id'.{u,v}] at e₁ e₂ rw [← e₁, ← e₂, lift_inj, mk_perm_eq_self_power, power_def] -theorem mk_equiv_eq_arrow_of_eq : #(α ≃ β) = #(α → β) := +theorem mk_equiv_eq_arrow_of_eq (eq : #α = #β) : #(α ≃ β) = #(α → β) := mk_equiv_eq_arrow_of_lift_eq congr(lift $eq) -theorem mk_equiv_of_lift_eq : #(α ≃ β') = 2 ^ lift.{v} #α := by +theorem mk_equiv_of_lift_eq (leq : lift.{v} #α = lift.{u} #β') : #(α ≃ β') = 2 ^ lift.{v} #α := by erw [← (lift_mk_eq'.2 ⟨.equivCongr (.refl α) (lift_mk_eq'.1 leq).some⟩).trans (lift_id'.{u,v} _), lift_umax.{u,v}, mk_perm_eq_two_power, lift_power, lift_natCast]; rfl -theorem mk_equiv_of_eq : #(α ≃ β) = 2 ^ #α := by rw [mk_equiv_of_lift_eq (lift_inj.mpr eq), lift_id] - -variable (lle : lift.{u} #β' ≤ lift.{v} #α) (le : #β ≤ #α) +theorem mk_equiv_of_eq (eq : #α = #β) : #(α ≃ β) = 2 ^ #α := by + rw [mk_equiv_of_lift_eq (lift_inj.mpr eq), lift_id] -theorem mk_embedding_eq_arrow_of_lift_le : #(β' ↪ α) = #(β' → α) := +theorem mk_embedding_eq_arrow_of_lift_le (lle : lift.{u} #β' ≤ lift.{v} #α) : + #(β' ↪ α) = #(β' → α) := (mk_embedding_le_arrow _ _).antisymm <| by conv_rhs => rw [← (Equiv.embeddingCongr (.refl _) (Cardinal.eq.mp <| mul_eq_self <| aleph0_le_mk α).some).cardinal_eq] @@ -1019,10 +1018,11 @@ theorem mk_embedding_eq_arrow_of_lift_le : #(β' ↪ α) = #(β' → α) := exact ⟨⟨fun f ↦ ⟨fun b ↦ ⟨e b, f b⟩, fun _ _ h ↦ e.injective congr(Prod.fst $h)⟩, fun f g h ↦ funext fun b ↦ congr(Prod.snd <| $h b)⟩⟩ -theorem mk_embedding_eq_arrow_of_le : #(β ↪ α) = #(β → α) := +theorem mk_embedding_eq_arrow_of_le (le : #β ≤ #α) : #(β ↪ α) = #(β → α) := mk_embedding_eq_arrow_of_lift_le (lift_le.mpr le) -theorem mk_surjective_eq_arrow_of_lift_le : #{f : α → β' | Surjective f} = #(α → β') := +theorem mk_surjective_eq_arrow_of_lift_le (lle : lift.{u} #β' ≤ lift.{v} #α) : + #{f : α → β' | Surjective f} = #(α → β') := (mk_set_le _).antisymm <| have ⟨e⟩ : Nonempty (α ≃ α ⊕ β') := by simp_rw [← lift_mk_eq', mk_sum, lift_add, lift_lift]; rw [lift_umax.{u,v}, eq_comm] @@ -1031,7 +1031,7 @@ theorem mk_surjective_eq_arrow_of_lift_le : #{f : α → β' | Surjective f} = # fun f g h ↦ funext fun a ↦ by simpa only [e.apply_symm_apply] using congr_fun (Subtype.ext_iff.mp h) (e.symm <| .inl a)⟩⟩ -theorem mk_surjective_eq_arrow_of_le : #{f : α → β | Surjective f} = #(α → β) := +theorem mk_surjective_eq_arrow_of_le (le : #β ≤ #α) : #{f : α → β | Surjective f} = #(α → β) := mk_surjective_eq_arrow_of_lift_le (lift_le.mpr le) end Function @@ -1206,13 +1206,13 @@ theorem mk_compl_eq_mk_compl_finite_lift {α : Type u} {β : Type v} [Finite α] theorem mk_compl_eq_mk_compl_finite {α β : Type u} [Finite α] {s : Set α} {t : Set β} (h1 : #α = #β) (h : #s = #t) : #(sᶜ : Set α) = #(tᶜ : Set β) := by - rw [← lift_inj.{u, max u v}] - apply mk_compl_eq_mk_compl_finite_lift.{u, u, max u v} + rw [← lift_inj.{u, u}] + apply mk_compl_eq_mk_compl_finite_lift.{u, u, u} <;> rwa [lift_inj] theorem mk_compl_eq_mk_compl_finite_same {α : Type u} [Finite α] {s t : Set α} (h : #s = #t) : #(sᶜ : Set α) = #(tᶜ : Set α) := - mk_compl_eq_mk_compl_finite.{u, u} rfl h + mk_compl_eq_mk_compl_finite.{u} rfl h end compl diff --git a/Mathlib/SetTheory/Cardinal/ToNat.lean b/Mathlib/SetTheory/Cardinal/ToNat.lean index a71d16d717fe2..f37fdaa87cbf3 100644 --- a/Mathlib/SetTheory/Cardinal/ToNat.lean +++ b/Mathlib/SetTheory/Cardinal/ToNat.lean @@ -44,9 +44,10 @@ theorem cast_toNat_of_lt_aleph0 {c : Cardinal} (h : c < ℵ₀) : ↑(toNat c) = lift c to ℕ using h rw [toNat_natCast] -theorem toNat_apply_of_lt_aleph0 {c : Cardinal} (h : c < ℵ₀) : +theorem toNat_apply_of_lt_aleph0 {c : Cardinal.{u}} (h : c < ℵ₀) : toNat c = Classical.choose (lt_aleph0.1 h) := - Nat.cast_injective <| by rw [cast_toNat_of_lt_aleph0 h, ← Classical.choose_spec (lt_aleph0.1 h)] + Nat.cast_injective (R := Cardinal.{u}) <| by + rw [cast_toNat_of_lt_aleph0 h, ← Classical.choose_spec (lt_aleph0.1 h)] theorem toNat_apply_of_aleph0_le {c : Cardinal} (h : ℵ₀ ≤ c) : toNat c = 0 := by simp [h] diff --git a/Mathlib/SetTheory/Game/Nim.lean b/Mathlib/SetTheory/Game/Nim.lean index 82ed1d3c25403..f56bc51986556 100644 --- a/Mathlib/SetTheory/Game/Nim.lean +++ b/Mathlib/SetTheory/Game/Nim.lean @@ -165,7 +165,7 @@ def nimOneRelabelling : nim 1 ≡r star := by rw [nim_def] refine ⟨?_, ?_, fun i => ?_, fun j => ?_⟩ any_goals dsimp; apply Equiv.equivOfUnique - all_goals simp; exact nimZeroRelabelling + all_goals simpa using nimZeroRelabelling theorem nim_one_equiv : nim 1 ≈ star := nimOneRelabelling.equiv diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index ff4a0373ed463..23a23ba8334f5 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -200,7 +200,7 @@ theorem pred_le {a b} : pred a ≤ b ↔ a ≤ succ b := theorem lift_is_succ {o : Ordinal.{v}} : (∃ a, lift.{u} o = succ a) ↔ ∃ a, o = succ a := ⟨fun ⟨a, h⟩ => let ⟨b, e⟩ := lift_down <| show a ≤ lift.{u} o from le_of_lt <| h.symm ▸ lt_succ a - ⟨b, lift_inj.1 <| by rw [h, ← e, lift_succ]⟩, + ⟨b, (lift_inj.{u,v}).1 <| by rw [h, ← e, lift_succ]⟩, fun ⟨a, h⟩ => ⟨lift.{u} a, by simp only [h, lift_succ]⟩⟩ @[simp] @@ -246,12 +246,13 @@ theorem lt_limit {o} (h : IsLimit o) {a} : a < o ↔ ∃ x < o, a < x := by simpa only [not_forall₂, not_le, bex_def] using not_congr (@limit_le _ h a) @[simp] -theorem lift_isLimit (o) : IsLimit (lift o) ↔ IsLimit o := +theorem lift_isLimit (o : Ordinal.{v}) : IsLimit (lift.{u,v} o) ↔ IsLimit o := and_congr (not_congr <| by simpa only [lift_zero] using @lift_inj o 0) - ⟨fun H a h => lift_lt.1 <| by simpa only [lift_succ] using H _ (lift_lt.2 h), fun H a h => by - obtain ⟨a', rfl⟩ := lift_down h.le - rw [← lift_succ, lift_lt] - exact H a' (lift_lt.1 h)⟩ + ⟨fun H a h => (lift_lt.{u,v}).1 <| + by simpa only [lift_succ] using H _ (lift_lt.2 h), fun H a h => by + obtain ⟨a', rfl⟩ := lift_down h.le + rw [← lift_succ, lift_lt] + exact H a' (lift_lt.1 h)⟩ theorem IsLimit.pos {o : Ordinal} (h : IsLimit o) : 0 < o := lt_of_le_of_ne (Ordinal.zero_le _) h.1.symm diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index d8d0dcca21f24..6266f0e4b0ba6 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1061,7 +1061,7 @@ def outOrderBotOfPos {o : Ordinal} (ho : 0 < o) : OrderBot o.out.α where theorem enum_zero_eq_bot {o : Ordinal} (ho : 0 < o) : enum (· < ·) 0 (by rwa [type_lt]) = haveI H := outOrderBotOfPos ho - ⊥ := + (⊥ : (Quotient.out o).α) := rfl /-! ### Universal ordinal -/ @@ -1140,7 +1140,7 @@ open Ordinal @[simp] theorem mk_ordinal_out (o : Ordinal) : #o.out.α = o.card := - (Ordinal.card_type _).symm.trans <| by rw [Ordinal.type_lt] + (Ordinal.card_type (· < ·)).symm.trans <| by rw [Ordinal.type_lt] /-- The ordinal corresponding to a cardinal `c` is the least ordinal whose cardinal is `c`. For the order-embedding version, see `ord.order_embedding`. -/ diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 5b391197b4d17..fb15a7d13a20d 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -406,15 +406,15 @@ theorem add_nfBelow {b} : ∀ {o₁ o₂}, NFBelow o₁ b → NFBelow o₂ b → | 0, _, _, h₂ => h₂ | oadd e n a, o, h₁, h₂ => by have h' := add_nfBelow (h₁.snd.mono <| le_of_lt h₁.lt) h₂ - simp [oadd_add]; revert h'; cases' a + o with e' n' a' <;> intro h' + simp only [oadd_add]; revert h'; cases' a + o with e' n' a' <;> intro h' · exact NFBelow.oadd h₁.fst NFBelow.zero h₁.lt have : ((e.cmp e').Compares e e') := @cmp_compares _ _ h₁.fst h'.fst - cases h : cmp e e' <;> dsimp [addAux] <;> simp [h] + cases h : cmp e e' <;> dsimp [addAux] <;> simp only [h] · exact h' - · simp [h] at this + · simp only [h] at this subst e' exact NFBelow.oadd h'.fst h'.snd h'.lt - · simp [h] at this + · simp only [h] at this exact NFBelow.oadd h₁.fst (NF.below_of_lt this ⟨⟨_, h'⟩⟩) h₁.lt instance add_nf (o₁ o₂) : ∀ [NF o₁] [NF o₂], NF (o₁ + o₂) @@ -452,12 +452,13 @@ theorem sub_nfBelow : ∀ {o₁ o₂ b}, NFBelow o₁ b → NF o₂ → NFBelow have h' := sub_nfBelow h₁.snd h₂.snd simp only [HSub.hSub, Sub.sub, sub] at h' ⊢ have := @cmp_compares _ _ h₁.fst h₂.fst - cases h : cmp e₁ e₂ <;> simp [sub] + cases h : cmp e₁ e₂ · apply NFBelow.zero - · simp only [h, Ordering.compares_eq] at this + · rw [Nat.sub_eq] + simp only [h, Ordering.compares_eq] at this subst e₂ - cases (n₁ : ℕ) - n₂ <;> simp [sub] - · by_cases en : n₁ = n₂ <;> simp [en] + cases (n₁ : ℕ) - n₂ + · by_cases en : n₁ = n₂ <;> simp only [en, ↓reduceIte] · exact h'.mono (le_of_lt h₁.lt) · exact NFBelow.zero · exact NFBelow.oadd h₁.fst h₁.snd h₁.lt @@ -528,7 +529,7 @@ theorem oadd_mul_nfBelow {e₁ n₁ a₁ b₁} (h₁ : NFBelow (oadd e₁ n₁ a | 0, b₂, _ => NFBelow.zero | oadd e₂ n₂ a₂, b₂, h₂ => by have IH := oadd_mul_nfBelow h₁ h₂.snd - by_cases e0 : e₂ = 0 <;> simp [e0, oadd_mul] + by_cases e0 : e₂ = 0 <;> simp only [e0, oadd_mul, ↓reduceIte] · apply NFBelow.oadd h₁.fst h₁.snd simpa using (add_lt_add_iff_left (repr e₁)).2 (lt_of_le_of_lt (Ordinal.zero_le _) h₂.lt) · haveI := h₁.fst @@ -555,11 +556,13 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep have ao : repr a₁ + ω ^ repr e₁ * (n₁ : ℕ) = ω ^ repr e₁ * (n₁ : ℕ) := by apply add_absorp h₁.snd'.repr_lt simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega_pos).2 (natCast_le.2 n₁.2) - by_cases e0 : e₂ = 0 <;> simp [e0, mul] - · cases' Nat.exists_eq_succ_of_ne_zero n₂.ne_zero with x xe + by_cases e0 : e₂ = 0 + · simp [e0, mul] + cases' Nat.exists_eq_succ_of_ne_zero n₂.ne_zero with x xe simp only [xe, h₂.zero_of_zero e0, repr, add_zero] rw [natCast_succ x, add_mul_succ _ ao, mul_assoc] - · haveI := h₁.fst + · simp only [repr] + haveI := h₁.fst haveI := h₂.fst simp only [Mul.mul, mul, e0, ite_false, repr.eq_2, repr_add, opow_add, IH, repr, mul_add] rw [← mul_assoc] @@ -727,7 +730,7 @@ theorem split_add_lt {o e n a m} [NF o] (h : split o = (oadd e n a, m)) : @[simp] theorem mulNat_eq_mul (n o) : mulNat o n = o * ofNat n := by cases o <;> cases n <;> rfl -instance nf_mulNat (o) [NF o] (n) : NF (mulNat o n) := by simp; exact ONote.mul_nf o (ofNat n) +instance nf_mulNat (o) [NF o] (n) : NF (mulNat o n) := by simpa using ONote.mul_nf o (ofNat n) instance nf_opowAux (e a0 a) [NF e] [NF a0] [NF a] : ∀ k m, NF (opowAux e a0 a k m) := by intro k m @@ -827,7 +830,7 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω have ω00 : 0 < ω0 ^ (k : Ordinal) := opow_pos _ (opow_pos _ omega_pos) have Rl : R < ω ^ (repr a0 * succ ↑k) := by by_cases k0 : k = 0 - · simp [R, k0] + · simp only [k0, Nat.cast_zero, succ_zero, mul_one, R] refine lt_of_lt_of_le ?_ (opow_le_opow_right omega_pos (one_le_iff_ne_zero.2 e0)) cases' m with m <;> simp [opowAux, omega_pos] rw [← add_one_eq_succ, ← Nat.cast_succ] diff --git a/Mathlib/SetTheory/Surreal/Multiplication.lean b/Mathlib/SetTheory/Surreal/Multiplication.lean index f8ebbd1a40ac3..d89da018a18f7 100644 --- a/Mathlib/SetTheory/Surreal/Multiplication.lean +++ b/Mathlib/SetTheory/Surreal/Multiplication.lean @@ -206,23 +206,25 @@ lemma ih1_neg_right : IH1 x y → IH1 x (-y) := /-! #### Specialize `ih` to obtain specialized induction hypotheses for P1 -/ -variable (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) -lemma numeric_option_mul (h : IsOption x' x) : (x' * y).Numeric := +lemma numeric_option_mul (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) (h : IsOption x' x) : + (x' * y).Numeric := ih (Args.P1 x' y) (TransGen.single <| cutExpand_pair_left h) -lemma numeric_mul_option (h : IsOption y' y) : (x * y').Numeric := +lemma numeric_mul_option (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) (h : IsOption y' y) : + (x * y').Numeric := ih (Args.P1 x y') (TransGen.single <| cutExpand_pair_right h) -lemma numeric_option_mul_option (hx : IsOption x' x) (hy : IsOption y' y) : (x' * y').Numeric := +lemma numeric_option_mul_option (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) (hx : IsOption x' x) + (hy : IsOption y' y) : (x' * y').Numeric := ih (Args.P1 x' y') ((TransGen.single <| cutExpand_pair_right hy).tail <| cutExpand_pair_left hx) -lemma ih1 : IH1 x y := by +lemma ih1 (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) : IH1 x y := by rintro x₁ x₂ y' h₁ h₂ (rfl|hy) <;> apply ih (Args.P24 _ _ _) on_goal 2 => refine TransGen.tail ?_ (cutExpand_pair_right hy) all_goals exact TransGen.single (cutExpand_double_left h₁ h₂) -lemma ih1_swap : IH1 y x := ih1 <| by +lemma ih1_swap (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) : IH1 y x := ih1 <| by simp_rw [ArgsRel, InvImage, Args.toMultiset, Multiset.pair_comm] at ih ⊢ exact ih @@ -234,17 +236,15 @@ lemma P3_of_ih (hy : Numeric y) (ihyx : IH1 y x) (i k l) : lemma P24_of_ih (ihxy : IH1 x y) (i j) : P24 (x.moveLeft i) (x.moveLeft j) y := ihxy (IsOption.moveLeft i) (IsOption.moveLeft j) (Or.inl rfl) -variable (hx : Numeric x) (hy : Numeric y) - section -variable (ihxy : IH1 x y) (ihyx : IH1 y x) - -lemma mulOption_lt_of_lt (i j k l) (h : x.moveLeft i < x.moveLeft j) : +lemma mulOption_lt_of_lt (hy : y.Numeric) (ihxy : IH1 x y) (ihyx : IH1 y x) (i j k l) + (h : x.moveLeft i < x.moveLeft j) : (⟦mulOption x y i k⟧ : Game) < -⟦mulOption x (-y) j l⟧ := mulOption_lt_iff_P1.2 <| P1_of_lt (P3_of_ih hy ihyx j k l) <| ((P24_of_ih ihxy i j).2 h).1 k -lemma mulOption_lt (i j k l) : (⟦mulOption x y i k⟧ : Game) < -⟦mulOption x (-y) j l⟧ := by +lemma mulOption_lt (hx : x.Numeric) (hy : y.Numeric) (ihxy : IH1 x y) (ihyx : IH1 y x) (i j k l) : + (⟦mulOption x y i k⟧ : Game) < -⟦mulOption x (-y) j l⟧ := by obtain (h|h|h) := lt_or_equiv_or_gt (hx.moveLeft i) (hx.moveLeft j) · exact mulOption_lt_of_lt hy ihxy ihyx i j k l h · have ml := @IsOption.moveLeft @@ -256,7 +256,8 @@ lemma mulOption_lt (i j k l) : (⟦mulOption x y i k⟧ : Game) < -⟦mulOption end /-- P1 follows from the induction hypothesis. -/ -theorem P1_of_ih : (x * y).Numeric := by +theorem P1_of_ih (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) (hx : x.Numeric) (hy : y.Numeric) : + (x * y).Numeric := by have ihxy := ih1 ih have ihyx := ih1_swap ih have ihxyn := ih1_neg_left (ih1_neg_right ihxy) @@ -290,9 +291,7 @@ def IH4 (x₁ x₂ y : PGame) : Prop := /-! #### Specialize `ih'` to obtain specialized induction hypotheses for P2 and P4 -/ -variable (ih' : ∀ a, ArgsRel a (Args.P24 x₁ x₂ y) → P124 a) - -lemma ih₁₂ : IH24 x₁ x₂ y := by +lemma ih₁₂ (ih' : ∀ a, ArgsRel a (Args.P24 x₁ x₂ y) → P124 a) : IH24 x₁ x₂ y := by rw [IH24] refine fun z ↦ ⟨?_, ?_, ?_⟩ <;> refine fun h ↦ ih' (Args.P24 _ _ _) (TransGen.single ?_) @@ -300,13 +299,13 @@ lemma ih₁₂ : IH24 x₁ x₂ y := by · exact (cutExpand_add_left {x₁}).2 (cutExpand_pair_left h) · exact (cutExpand_add_left {x₁}).2 (cutExpand_pair_right h) -lemma ih₂₁ : IH24 x₂ x₁ y := ih₁₂ <| by +lemma ih₂₁ (ih' : ∀ a, ArgsRel a (Args.P24 x₁ x₂ y) → P124 a) : IH24 x₂ x₁ y := ih₁₂ <| by simp_rw [ArgsRel, InvImage, Args.toMultiset, Multiset.pair_comm] at ih' ⊢ suffices {x₁, y, x₂} = {x₂, y, x₁} by rwa [← this] dsimp only [Multiset.insert_eq_cons, ← Multiset.singleton_add] at ih' ⊢ abel -lemma ih4 : IH4 x₁ x₂ y := by +lemma ih4 (ih' : ∀ a, ArgsRel a (Args.P24 x₁ x₂ y) → P124 a) : IH4 x₁ x₂ y := by refine fun z w h ↦ ⟨?_, ?_⟩ all_goals intro h' @@ -315,7 +314,8 @@ lemma ih4 : IH4 x₁ x₂ y := by try exact (cutExpand_add_right {w}).2 <| cutExpand_pair_left h' try exact (cutExpand_add_right {w}).2 <| cutExpand_pair_right h' -lemma numeric_of_ih : (x₁ * y).Numeric ∧ (x₂ * y).Numeric := by +lemma numeric_of_ih (ih' : ∀ a, ArgsRel a (Args.P24 x₁ x₂ y) → P124 a) : + (x₁ * y).Numeric ∧ (x₂ * y).Numeric := by constructor <;> refine ih' (Args.P1 _ _) (TransGen.single ?_) · exact (cutExpand_add_right {y}).2 <| (cutExpand_add_left {x₁}).2 cutExpand_zero · exact (cutExpand_add_right {x₂, y}).2 cutExpand_zero diff --git a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean index e3df7300cf04f..cdae67c853745 100644 --- a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean +++ b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yuma Mizuno. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ -import Mathlib.CategoryTheory.Bicategory.Free +import Mathlib.CategoryTheory.Bicategory.Coherence import Mathlib.Tactic.CategoryTheory.BicategoricalComp /-! diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index e54081435e58a..5e5b0a4436c80 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -3,7 +3,7 @@ Copyright (c) 2022. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Yuma Mizuno, Oleksandr Manzyuk -/ -import Mathlib.CategoryTheory.Monoidal.Free.Basic +import Mathlib.CategoryTheory.Monoidal.Free.Coherence import Mathlib.Lean.Meta import Mathlib.Tactic.CategoryTheory.BicategoryCoherence import Mathlib.Tactic.CategoryTheory.MonoidalComp @@ -23,9 +23,6 @@ are equal. -/ --- Porting note: restore when ported --- import Mathlib.CategoryTheory.Bicategory.CoherenceTactic - universe v u open CategoryTheory FreeMonoidalCategory diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 3f4bd6f37c00a..f318dca852025 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -16,6 +16,9 @@ import ImportGraph.Imports -- Hopefully `lake` will be able to handle tests later. import ProofWidgets +-- Import common Batteries tactics and commands +import Batteries.Tactic.Where + -- Import Mathlib-specific linters. import Mathlib.Tactic.Linter.Lint diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 6d79fe5bfd554..0cfab4e9c4cc1 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Lean.Linter.Util +import Batteries.Data.String.Matcher import Batteries.Tactic.Lint /-! @@ -217,4 +218,53 @@ initialize addLinter cdotLinter end CDotLinter +/-- The "longLine" linter emits a warning on lines longer than 100 characters. +We allow lines containing URLs to be longer, though. -/ +register_option linter.longLine : Bool := { + defValue := true + descr := "enable the longLine linter" +} + +namespace LongLine + +/-- Gets the value of the `linter.longLine` option. -/ +def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.longLine o + +@[inherit_doc Mathlib.Linter.linter.longLine] +def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless getLinterHash (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + -- TODO: once per-project settings are available, + -- revert this hack to make it only apply on `Mathlib` + unless #[`Mathlib, `test, `Archive, `Counterexamples].contains (← getMainModule).getRoot do + return + -- The linter ignores the `#guard_msgs` command, in particular its doc-string. + -- The linter still lints the message guarded by `#guard_msgs`. + if stx.isOfKind ``Lean.guardMsgsCmd then + return + -- if the linter reached the end of the file, then we scan the `import` syntax instead + let stx := ← do + if stx.isOfKind ``Lean.Parser.Command.eoi then + let fname ← getFileName + if !(← System.FilePath.pathExists fname) then return default + let contents ← IO.FS.readFile fname + -- `impMods` is the syntax for the modules imported in the current file + let (impMods, _) ← Parser.parseHeader (Parser.mkInputContext contents fname) + return impMods + else return stx + let sstr := stx.getSubstring? + let fm ← getFileMap + let longLines := ((sstr.getD default).splitOn "\n").filter fun line => + (100 < (fm.toPosition line.stopPos).column) + for line in longLines do + if !(line.containsSubstr "http") then + Linter.logLint linter.longLine (.ofRange ⟨line.startPos, line.stopPos⟩) + m!"This line exceeds the 100 character limit, please shorten it!" + +initialize addLinter longLineLinter + +end LongLine + end Mathlib.Linter diff --git a/Mathlib/Tactic/Linter/MinImports.lean b/Mathlib/Tactic/Linter/MinImports.lean index 7b9cd7c823da3..999aca9d8c390 100644 --- a/Mathlib/Tactic/Linter/MinImports.lean +++ b/Mathlib/Tactic/Linter/MinImports.lean @@ -72,26 +72,21 @@ def minImportsLinter : Linter where run := withSetOptionIn fun stx => do let currentlyUnneededImports := explicitImportsInFile.diff importsSoFar -- we read the current file, to do a custom parsing of the imports: -- this is a hack to obtain some `Syntax` information for the `import X` commands - let fil ← IO.FS.readFile (← getFileName) + let fname ← getFileName + let contents ← IO.FS.readFile fname + -- `impMods` is the syntax for the modules imported in the current file + let (impMods, _) ← Parser.parseHeader (Parser.mkInputContext contents fname) for i in currentlyUnneededImports do - -- looking for the position of 'import i', so we split at ' i' and - -- compute the length of the string until '...import| i' - match fil.splitOn (" " ++ i.toString) with - | a::_::_ => - let al := a.length - -- create a syntax covering the range that should be occupied by import `i` - let impPos : Syntax := .ofRange ⟨⟨al + 1⟩, ⟨al + i.toString.length + 1⟩⟩ - logWarningAt impPos m!"unneeded import '{i}'" + match impMods.find? (·.getId == i) with + | some impPos => logWarningAt impPos m!"unneeded import '{i}'" | _ => dbg_trace f!"'{i}' not found" -- this should be unreachable -- if the linter found new imports that should be added (likely to *reduce* the dependencies) if !newImps.isEmpty then -- format the imports prepending `import ` to each module name let withImport := (newImps.toArray.qsort Name.lt).map (s!"import {·}") - -- create a syntax node supported, likely on the first imported module name - let firstImport : Syntax := match fil.splitOn "import " with - | a::_::_ => .ofRange ⟨⟨a.length⟩, ⟨a.length + "import".length⟩⟩ - | _ => .ofRange ⟨⟨0⟩, ⟨19⟩⟩ - logWarningAt firstImport m!"-- missing imports\n{"\n".intercalate withImport.toList}" + -- log a warning at the first `import`, if there is one. + logWarningAt ((impMods.find? (·.isOfKind `import)).getD default) + m!"-- missing imports\n{"\n".intercalate withImport.toList}" let id ← getId stx let newImports := getIrredundantImports (← getEnv) (← getAllImports stx id) let tot := (newImports.append importsSoFar) diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index fb57a6ed99fd2..425cdc0023e0e 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -5,7 +5,6 @@ Authors: Damiano Testa -/ import Lean.Elab.Command import Lean.Linter.Util -import Batteries.Data.List.Basic import Batteries.Tactic.Unreachable /-! diff --git a/Mathlib/Tactic/NormNum/Eq.lean b/Mathlib/Tactic/NormNum/Eq.lean index cdb5fd5473b14..3c0cdec603d27 100644 --- a/Mathlib/Tactic/NormNum/Eq.lean +++ b/Mathlib/Tactic/NormNum/Eq.lean @@ -17,11 +17,11 @@ namespace Mathlib.Meta.NormNum theorem isNat_eq_false [AddMonoidWithOne α] [CharZero α] : {a b : α} → {a' b' : ℕ} → IsNat a a' → IsNat b b' → Nat.beq a' b' = false → ¬a = b - | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, h => by simp; exact Nat.ne_of_beq_eq_false h + | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, h => by simpa using Nat.ne_of_beq_eq_false h theorem isInt_eq_false [Ring α] [CharZero α] : {a b : α} → {a' b' : ℤ} → IsInt a a' → IsInt b b' → decide (a' = b') = false → ¬a = b - | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, h => by simp; exact of_decide_eq_false h + | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, h => by simpa using of_decide_eq_false h theorem Rat.invOf_denom_swap [Ring α] (n₁ n₂ : ℤ) (a₁ a₂ : α) [Invertible a₁] [Invertible a₂] : n₁ * ⅟a₁ = n₂ * ⅟a₂ ↔ n₁ * a₂ = n₂ * a₁ := by diff --git a/Mathlib/Tactic/NormNum/OfScientific.lean b/Mathlib/Tactic/NormNum/OfScientific.lean index a74e8278b3eac..c494615ef34f0 100644 --- a/Mathlib/Tactic/NormNum/OfScientific.lean +++ b/Mathlib/Tactic/NormNum/OfScientific.lean @@ -23,7 +23,7 @@ variable {α : Type*} theorem isRat_ofScientific_of_true [DivisionRing α] : {m e : ℕ} → {n : ℤ} → {d : ℕ} → IsRat (mkRat m (10 ^ e) : α) n d → IsRat (OfScientific.ofScientific m true e : α) n d - | _, _, _, _, ⟨_, eq⟩ => ⟨_, by + | _, _, _, _, ⟨_, eq⟩ => ⟨‹_›, by rwa [← Rat.cast_ofScientific, ← Rat.ofScientific_eq_ofScientific, Rat.ofScientific_true_def]⟩ -- see note [norm_num lemma function equality] diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index dbdc3466dba42..6f1c5c4089821 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -474,6 +474,34 @@ def findProjectionIndices (strName projName : Name) : MetaM (List Nat) := do let allProjs := pathToField ++ [fullProjName] return allProjs.map (env.getProjectionFnInfo? · |>.get!.i) +/-- +A variant of `Substring.dropPrefix?` that does not consider `toFoo` to be a prefix to `toFoo_1`. +This is checked by inspecting whether the first character of the remaining part is a digit. + +We use this variant because the latter is often a different field with an auto-generated name. +-/ +private def dropPrefixIfNotNumber? (s : String) (pre : Substring) : Option Substring := do + let ret ← Substring.dropPrefix? s pre + -- flag is true when the remaning part is nonempty and starts with a digit. + let flag := ret.toString.data.head?.elim false Char.isDigit + if flag then none else some ret + +/-- A variant of `String.isPrefixOf` that does not consider `toFoo` to be a prefix to `toFoo_1`. -/ +private def isPrefixOfAndNotNumber (s p : String) : Bool := (dropPrefixIfNotNumber? p s).isSome + +/-- A variant of `String.splitOn` that does not split `toFoo_1` into `toFoo` and `1`. -/ +private def splitOnNotNumber (s delim : String) : List String := + (process (s.splitOn delim).reverse "").reverse where + process (arr : List String) (tail : String) := match arr with + | [] => [] + | (x :: xs) => + -- flag is true when this segment is nonempty and starts with a digit. + let flag := x.data.head?.elim false Char.isDigit + if flag then + process xs (tail ++ delim ++ x) + else + List.cons (x ++ tail) (process xs "") + /-- Auxiliary function of `getCompositeOfProjections`. -/ partial def getCompositeOfProjectionsAux (proj : String) (e : Expr) (pos : Array Nat) (args : Array Expr) : MetaM (Expr × Array Nat) := do @@ -482,7 +510,7 @@ partial def getCompositeOfProjectionsAux (proj : String) (e : Expr) (pos : Array throwError "{e} doesn't have a structure as type" let projs := getStructureFieldsFlattened env structName let projInfo := projs.toList.map fun p ↦ do - ((← proj.dropPrefix? (p.lastComponentAsString ++ "_")).toString, p) + ((← dropPrefixIfNotNumber? proj (p.lastComponentAsString ++ "_")).toString, p) let some (projRest, projName) := projInfo.reduceOption.getLast? | throwError "Failed to find constructor {proj.dropRight 1} in structure {structName}." let newE ← mkProjection e projName @@ -1023,7 +1051,8 @@ partial def addProjections (nm : Name) (type lhs rhs : Expr) if !todoNext.isEmpty && str ∉ cfg.notRecursive then let firstTodo := todoNext.head!.1 throwError "Invalid simp lemma {nm.appendAfter firstTodo}.\nProjection \ - {(firstTodo.splitOn "_")[1]!} doesn't exist, because target {str} is not a structure." + {(splitOnNotNumber firstTodo "_")[1]!} doesn't exist, \ + because target {str} is not a structure." if cfg.fullyApplied then addProjection stxProj univs nm tgt lhsAp rhsAp newArgs cfg else @@ -1106,9 +1135,9 @@ partial def addProjections (nm : Name) (type lhs rhs : Expr) trace[simps.debug] "Next todo: {todoNext}" -- check whether all elements in `todo` have a projection as prefix if let some (x, _) := todo.find? fun (x, _) ↦ projs.all - fun proj ↦ !(proj.lastComponentAsString ++ "_").isPrefixOf x then + fun proj ↦ !isPrefixOfAndNotNumber (proj.lastComponentAsString ++ "_") x then let simpLemma := nm.appendAfter x - let neededProj := (x.splitOn "_")[0]! + let neededProj := (splitOnNotNumber x "_")[0]! throwError "Invalid simp lemma {simpLemma}. \ Structure {str} does not have projection {neededProj}.\n\ The known projections are:\ @@ -1120,7 +1149,8 @@ partial def addProjections (nm : Name) (type lhs rhs : Expr) let nms ← projInfo.concatMapM fun ⟨newRhs, proj, projExpr, projNrs, isDefault, isPrefix⟩ ↦ do let newType ← inferType newRhs let newTodo := todo.filterMap - fun (x, stx) ↦ (x.dropPrefix? (proj.lastComponentAsString ++ "_")).map (·.toString, stx) + fun (x, stx) ↦ (dropPrefixIfNotNumber? x (proj.lastComponentAsString ++ "_")).map + (·.toString, stx) -- we only continue with this field if it is default or mentioned in todo if !(isDefault && todo.isEmpty) && newTodo.isEmpty then return #[] let newLhs := projExpr.instantiateLambdasOrApps #[lhsAp] diff --git a/Mathlib/Tactic/Widget/Calc.lean b/Mathlib/Tactic/Widget/Calc.lean index e287eb58106d4..34ea5f388d745 100644 --- a/Mathlib/Tactic/Widget/Calc.lean +++ b/Mathlib/Tactic/Widget/Calc.lean @@ -7,6 +7,8 @@ import Lean.Elab.Tactic.Calc import Mathlib.Data.String.Defs import Mathlib.Tactic.Widget.SelectPanelUtils +import Batteries.CodeAction.Attr +import Batteries.Lean.Position /-! # Calc widget diff --git a/Mathlib/Tactic/Widget/Conv.lean b/Mathlib/Tactic/Widget/Conv.lean index d3deed1858847..50cd77a425f8d 100644 --- a/Mathlib/Tactic/Widget/Conv.lean +++ b/Mathlib/Tactic/Widget/Conv.lean @@ -6,6 +6,7 @@ Authors: Robin Böhne, Wojciech Nawrocki, Patrick Massot import Mathlib.Tactic.Widget.SelectPanelUtils import Mathlib.Data.String.Defs import Batteries.Tactic.Lint +import Batteries.Lean.Position /-! # Conv widget diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 315e0c645e65b..20086013e281e 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -334,7 +334,7 @@ theorem applyId_mem_iff [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs erw [← Option.mem_def, List.mem_dlookup_iff] at h₃ · simp only [Prod.toSigma, List.mem_map, heq_iff_eq, Prod.exists] at h₃ rcases h₃ with ⟨a, b, h₃, h₄, h₅⟩ - apply (List.mem_zip h₃).2 + apply (List.of_mem_zip h₃).2 simp only [List.NodupKeys, List.keys, comp, Prod.fst_toSigma, List.map_map] rwa [List.map_fst_zip _ _ (le_of_eq h₆)] @@ -347,7 +347,7 @@ theorem List.applyId_eq_self [DecidableEq α] {xs ys : List α} (x : α) : simp only [List.keys, not_exists, Prod.toSigma, exists_and_right, exists_eq_right, List.mem_map, Function.comp_apply, List.map_map, Prod.exists] intro y hy - exact h (List.mem_zip hy).1 + exact h (List.of_mem_zip hy).1 theorem applyId_injective [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs) (h₁ : xs ~ ys) : Injective.{u + 1, u + 1} (List.applyId (xs.zip ys)) := by @@ -389,7 +389,7 @@ def Perm.slice [DecidableEq α] (n m : ℕ) : have h₀ : xs' ~ ys.inter xs' := List.Perm.dropSlice_inter _ _ h h' ⟨xs', ys.inter xs', h₀, h'.inter _⟩ -/-- A lazy list, in decreasing order, of sizes that should be +/-- A list, in decreasing order, of sizes that should be sliced off a list of length `n` -/ def sliceSizes : ℕ → MLList Id ℕ+ diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index 8d10ed2d0fd66..76676013565a5 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Scott Morrison, Antoine Chambert-Loir, María Inés de Frutos-Fernández -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Topology.Algebra.Module.Basic @@ -16,15 +16,19 @@ topological algebras. ## Results -This is just a minimal stub for now! +The topological closure of a subalgebra is still a subalgebra, which as an algebra is a +topological algebra. -The topological closure of a subalgebra is still a subalgebra, -which as an algebra is a topological algebra. --/ +In this file we define continuous algebra homomorphisms, as algebra homomorphisms between +topological (semi-)rings which are continuous. The set of continuous algebra homomorphisms between +the topological `R`-algebras `A` and `B` is denoted by `A →R[A] B`. + +TODO: add continuous algebra isomorphisms. +-/ open scoped Classical -open Set TopologicalSpace Algebra +open Set TopologicalSpace Algebra BigOperators open scoped Classical @@ -66,28 +70,470 @@ theorem algebraMapCLM_coe : ⇑(algebraMapCLM R A) = algebraMap R A := theorem algebraMapCLM_toLinearMap : (algebraMapCLM R A).toLinearMap = Algebra.linearMap R A := rfl +/-- If `R` is a discrete topological ring, then any topological ring `S` which is an `R`-algebra +is also a topological `R`-algebra. -/ +instance [TopologicalSemiring A] [DiscreteTopology R] : + ContinuousSMul R A := continuousSMul_of_algebraMap _ _ continuous_of_discreteTopology + end TopologicalAlgebra section TopologicalAlgebra -variable {R : Type*} [CommSemiring R] -variable {A : Type u} [TopologicalSpace A] -variable [Semiring A] [Algebra R A] +section -variable [TopologicalSemiring A] +variable (R : Type*) [CommSemiring R] [TopologicalSpace R] [TopologicalSemiring R] + (A : Type*) [Semiring A] [TopologicalSpace A] [TopologicalSemiring A] + +/-- Continuous algebra homomorphisms between algebras. We only put the type classes that are +necessary for the definition, although in applications `M` and `B` will be topological algebras +over the topological ring `R`. -/ +structure ContinuousAlgHom (R : Type*) [CommSemiring R] (A : Type*) [Semiring A] + [TopologicalSpace A] (B : Type*) [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] + extends A →ₐ[R] B where +-- TODO: replace with `fun_prop` when that is stable + cont : Continuous toFun := by continuity + +@[inherit_doc] +notation:25 A " →A[" R "] " B => ContinuousAlgHom R A B + +namespace ContinuousAlgHom + +section Semiring + +variable {R} {A} +variable [TopologicalSpace R] [TopologicalSemiring R] + +variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] + +instance : FunLike (A →A[R] B) A B where + coe f := f.toAlgHom + coe_injective' f g h := by + cases f; cases g + simp only [mk.injEq] + exact AlgHom.ext (congrFun h) + +instance : AlgHomClass (A →A[R] B) R A B where + map_mul f x y := map_mul f.toAlgHom x y + map_one f := map_one f.toAlgHom + map_add f := map_add f.toAlgHom + map_zero f := map_zero f.toAlgHom + commutes f r := f.toAlgHom.commutes r + +@[simp] +theorem toAlgHom_eq_coe (f : A →A[R] B) : f.toAlgHom = f := rfl + +@[simp, norm_cast] +theorem coe_inj {f g : A →A[R] B} : (f : A →ₐ[R] B) = g ↔ f = g := by + cases f; cases g; simp only [mk.injEq]; exact Eq.congr_right rfl + +@[simp] +theorem coe_mk (f : A →ₐ[R] B) (h) : (mk f h : A →ₐ[R] B) = f := rfl + +@[simp] +theorem coe_mk' (f : A →ₐ[R] B) (h) : (mk f h : A → B) = f := rfl + +@[simp, norm_cast] +theorem coe_coe (f : A →A[R] B) : ⇑(f : A →ₐ[R] B) = f := rfl + +instance : ContinuousMapClass (A →A[R] B) A B where + map_continuous f := f.2 + +@[fun_prop] +protected theorem continuous (f : A →A[R] B) : Continuous f := f.2 + +protected theorem uniformContinuous {E₁ E₂ : Type*} [UniformSpace E₁] [UniformSpace E₂] + [Ring E₁] [Ring E₂] [Algebra R E₁] [Algebra R E₂] [UniformAddGroup E₁] + [UniformAddGroup E₂] (f : E₁ →A[R] E₂) : UniformContinuous f := + uniformContinuous_addMonoidHom_of_continuous f.continuous + +/-- See Note [custom simps projection]. We need to specify this projection explicitly in this case, +because it is a composition of multiple projections. -/ +def Simps.apply (h : A →A[R] B) : A → B := h + +/-- See Note [custom simps projection]. -/ +def Simps.coe (h : A →A[R] B) : A →ₐ[R] B := h + +initialize_simps_projections ContinuousAlgHom (toAlgHom_toFun → apply, toAlgHom → coe) + +@[ext] +theorem ext {f g : A →A[R] B} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h + +/-- Copy of a `ContinuousAlgHom` with a new `toFun` equal to the old one. Useful to fix +definitional equalities. -/ +def copy (f : A →A[R] B) (f' : A → B) (h : f' = ⇑f) : A →A[R] B where + toAlgHom := { + toRingHom := (f : A →A[R] B).toRingHom.copy f' h + commutes' := fun r => by + simp only [AlgHom.toRingHom_eq_coe, h, RingHom.toMonoidHom_eq_coe, OneHom.toFun_eq_coe, + MonoidHom.toOneHom_coe, MonoidHom.coe_coe, RingHom.coe_copy, AlgHomClass.commutes f r] } + cont := show Continuous f' from h.symm ▸ f.continuous + +@[simp] +theorem coe_copy (f : A →A[R] B) (f' : A → B) (h : f' = ⇑f) : ⇑(f.copy f' h) = f' := rfl + +theorem copy_eq (f : A →A[R] B) (f' : A → B) (h : f' = ⇑f) : f.copy f' h = f := DFunLike.ext' h + +protected theorem map_zero (f : A →A[R] B) : f (0 : A) = 0 := map_zero f + +protected theorem map_add (f : A →A[R] B) (x y : A) : f (x + y) = f x + f y := map_add f x y + +protected theorem map_smul (f : A →A[R] B) (c : R) (x : A) : + f (c • x) = c • f x := + map_smul .. + +theorem map_smul_of_tower {R S : Type*} [CommSemiring S] [SMul R A] [Algebra S A] [SMul R B] + [Algebra S B] [MulActionHomClass (A →A[S] B) R A B] (f : A →A[S] B) (c : R) (x : A) : + f (c • x) = c • f x := + map_smul f c x + +protected theorem map_sum {ι : Type*} (f : A →A[R] B) (s : Finset ι) (g : ι → A) : + f (∑ i in s, g i) = ∑ i in s, f (g i) := + map_sum .. + +/-- Any two continuous `R`-algebra morphisms from `R` are equal -/ +@[ext (iff := false)] +theorem ext_ring [TopologicalSpace R] {f g : R →A[R] A} : f = g := + coe_inj.mp (ext_id _ _ _) + +theorem ext_ring_iff [TopologicalSpace R] {f g : R →A[R] A} : f = g ↔ f 1 = g 1 := + ⟨fun h => h ▸ rfl, fun _ => ext_ring ⟩ + +/-- If two continuous algebra maps are equal on a set `s`, then they are equal on the closure +of the `Algebra.adjoin` of this set. -/ +theorem eqOn_closure_adjoin [T2Space B] {s : Set A} {f g : A →A[R] B} (h : Set.EqOn f g s) : + Set.EqOn f g (closure (Algebra.adjoin R s : Set A)) := + Set.EqOn.closure (AlgHom.eqOn_adjoin_iff.mpr h) f.continuous g.continuous + +/-- If the subalgebra generated by a set `s` is dense in the ambient module, then two continuous +algebra maps equal on `s` are equal. -/ +theorem ext_on [T2Space B] {s : Set A} (hs : Dense (Algebra.adjoin R s : Set A)) + {f g : A →A[R] B} (h : Set.EqOn f g s) : f = g := + ext fun x => eqOn_closure_adjoin h (hs x) + +/-- The topological closure of a subalgebra -/ +def _root_.Subalgebra.topologicalClosure (s : Subalgebra R A) : Subalgebra R A where + toSubsemiring := s.toSubsemiring.topologicalClosure + algebraMap_mem' r := by + simp only [Subsemiring.coe_carrier_toSubmonoid, Subsemiring.topologicalClosure_coe, + Subalgebra.coe_toSubsemiring] + apply subset_closure + exact algebraMap_mem s r + +/-- Under a continuous algebra map, the image of the `TopologicalClosure` of a subalgebra is +contained in the `TopologicalClosure` of its image. -/ +theorem _root_.Subalgebra.topologicalClosure_map + [TopologicalSemiring B] (f : A →A[R] B) (s : Subalgebra R A) : + s.topologicalClosure.map f ≤ (s.map f.toAlgHom).topologicalClosure := + image_closure_subset_closure_image f.continuous + +@[simp] +theorem _root_.Subalgebra.topologicalClosure_coe + (s : Subalgebra R A) : + (s.topologicalClosure : Set A) = closure ↑s := rfl + +/-- Under a dense continuous algebra map, a subalgebra +whose `TopologicalClosure` is `⊤` is sent to another such submodule. +That is, the image of a dense subalgebra under a map with dense range is dense. +-/ +theorem _root_.DenseRange.topologicalClosure_map_subalgebra + [TopologicalSemiring B] {f : A →A[R] B} (hf' : DenseRange f) {s : Subalgebra R A} + (hs : s.topologicalClosure = ⊤) : (s.map (f : A →ₐ[R] B)).topologicalClosure = ⊤ := by + rw [SetLike.ext'_iff] at hs ⊢ + simp only [Subalgebra.topologicalClosure_coe, coe_top, ← dense_iff_closure_eq, Subalgebra.coe_map, + AlgHom.coe_coe] at hs ⊢ + exact hf'.dense_image f.continuous hs + +end Semiring + +section id + +variable [Algebra R A] + +/-- The identity map as a continuous algebra homomorphism. -/ +protected def id : A →A[R] A := ⟨AlgHom.id R A, continuous_id⟩ + +instance : One (A →A[R] A) := ⟨ContinuousAlgHom.id R A⟩ + +theorem one_def : (1 : A →A[R] A) = ContinuousAlgHom.id R A := rfl + +theorem id_apply (x : A) : ContinuousAlgHom.id R A x = x := rfl + +@[simp, norm_cast] +theorem coe_id : ((ContinuousAlgHom.id R A) : A →ₐ[R] A) = AlgHom.id R A:= rfl + +@[simp, norm_cast] +theorem coe_id' : ⇑(ContinuousAlgHom.id R A ) = _root_.id := rfl + +@[simp, norm_cast] +theorem coe_eq_id {f : A →A[R] A} : + (f : A →ₐ[R] A) = AlgHom.id R A ↔ f = ContinuousAlgHom.id R A:= by + rw [← coe_id, coe_inj] + +@[simp] +theorem one_apply (x : A) : (1 : A →A[R] A) x = x := rfl + +end id + +section comp + +variable {R} {A} +variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] + {C : Type*} [Semiring C] [Algebra R C] [TopologicalSpace C] + +/-- Composition of continuous algebra homomorphisms. -/ +def comp (g : B →A[R] C) (f : A →A[R] B) : A →A[R] C := + ⟨(g : B →ₐ[R] C).comp (f : A →ₐ[R] B), g.2.comp f.2⟩ + +@[simp, norm_cast] +theorem coe_comp (h : B →A[R] C) (f : A →A[R] B) : + (h.comp f : A →ₐ[R] C) = (h : B →ₐ[R] C).comp (f : A →ₐ[R] B) := rfl + +@[simp, norm_cast] +theorem coe_comp' (h : B →A[R] C) (f : A →A[R] B) : ⇑(h.comp f) = h ∘ f := rfl + +theorem comp_apply (g : B →A[R] C) (f : A →A[R] B) (x : A) : (g.comp f) x = g (f x) := rfl + +@[simp] +theorem comp_id (f : A →A[R] B) : f.comp (ContinuousAlgHom.id R A) = f := + ext fun _x => rfl + +@[simp] +theorem id_comp (f : A →A[R] B) : (ContinuousAlgHom.id R B).comp f = f := + ext fun _x => rfl + +theorem comp_assoc {D : Type*} [Semiring D] [Algebra R D] [TopologicalSpace D] (h : C →A[R] D) + (g : B →A[R] C) (f : A →A[R] B) : (h.comp g).comp f = h.comp (g.comp f) := + rfl + +instance : Mul (A →A[R] A) := ⟨comp⟩ + +theorem mul_def (f g : A →A[R] A) : f * g = f.comp g := rfl + +@[simp] +theorem coe_mul (f g : A →A[R] A) : ⇑(f * g) = f ∘ g := rfl + +theorem mul_apply (f g : A →A[R] A) (x : A) : (f * g) x = f (g x) := rfl + +instance : Monoid (A →A[R] A) where + mul_one _ := ext fun _ => rfl + one_mul _ := ext fun _ => rfl + mul_assoc _ _ _ := ext fun _ => rfl + +theorem coe_pow (f : A →A[R] A) (n : ℕ) : ⇑(f ^ n) = f^[n] := + hom_coe_pow _ rfl (fun _ _ ↦ rfl) _ _ + +/-- coercion from `ContinuousAlgHom` to `AlgHom` as a `RingHom`. -/ +@[simps] +def toAlgHomMonoidHom : (A →A[R] A) →* A →ₐ[R] A where + toFun := (↑) + map_one' := rfl + map_mul' _ _ := rfl + +end comp + +section prod + +variable {R} {A} +variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] + {C : Type*} [Semiring C] [Algebra R C] [TopologicalSpace C] + +/-- The cartesian product of two continuous algebra morphisms as a continuous algebra morphism. -/ +protected def prod (f₁ : A →A[R] B) (f₂ : A →A[R] C) : + A →A[R] B × C := + ⟨(f₁ : A →ₐ[R] B).prod f₂, f₁.2.prod_mk f₂.2⟩ + +@[simp, norm_cast] +theorem coe_prod (f₁ : A →A[R] B) (f₂ : A →A[R] C) : + (f₁.prod f₂ : A →ₐ[R] B × C) = AlgHom.prod f₁ f₂ := + rfl -/-- The closure of a subalgebra in a topological algebra as a subalgebra. -/ -def Subalgebra.topologicalClosure (s : Subalgebra R A) : Subalgebra R A := - { s.toSubsemiring.topologicalClosure with - carrier := closure (s : Set A) - algebraMap_mem' := fun r => s.toSubsemiring.le_topologicalClosure (s.algebraMap_mem r) } +@[simp, norm_cast] +theorem prod_apply (f₁ : A →A[R] B) (f₂ : A →A[R] C) (x : A) : + f₁.prod f₂ x = (f₁ x, f₂ x) := + rfl + +variable {F : Type*} + +instance {D : Type*} [UniformSpace D] [CompleteSpace D] + [Semiring D] [Algebra R D] [T2Space B] + [FunLike F D B] [AlgHomClass F R D B] [ContinuousMapClass F D B] + (f g : F) : CompleteSpace (AlgHom.equalizer f g) := + isClosed_eq (map_continuous f) (map_continuous g) |>.completeSpace_coe + +variable (R A B) + +/-- `Prod.fst` as a `ContinuousAlgHom`. -/ +def fst : A × B →A[R] A where + cont := continuous_fst + toAlgHom := AlgHom.fst R A B + +/-- `Prod.snd` as a `ContinuousAlgHom`. -/ +def snd : A × B →A[R] B where + cont := continuous_snd + toAlgHom := AlgHom.snd R A B + +variable {R A B} + +@[simp, norm_cast] +theorem coe_fst : ↑(fst R A B) = AlgHom.fst R A B := + rfl + +@[simp, norm_cast] +theorem coe_fst' : ⇑(fst R A B) = Prod.fst := + rfl + +@[simp, norm_cast] +theorem coe_snd : ↑(snd R A B) = AlgHom.snd R A B := + rfl + +@[simp, norm_cast] +theorem coe_snd' : ⇑(snd R A B) = Prod.snd := + rfl + +@[simp] +theorem fst_prod_snd : (fst R A B).prod (snd R A B) = ContinuousAlgHom.id R (A × B) := + ext fun ⟨_x, _y⟩ => rfl + +@[simp] +theorem fst_comp_prod (f : A →A[R] B) (g : A →A[R] C) : + (fst R B C).comp (f.prod g) = f := + ext fun _x => rfl + +@[simp] +theorem snd_comp_prod (f : A →A[R] B) (g : A →A[R] C) : + (snd R B C).comp (f.prod g) = g := + ext fun _x => rfl + +/-- `Prod.map` of two continuous algebra homomorphisms. -/ +def prodMap {D : Type*} [Semiring D] [TopologicalSpace D] [Algebra R D] (f₁ : A →A[R] B) + (f₂ : C →A[R] D) : A × C →A[R] B × D := + (f₁.comp (fst R A C)).prod (f₂.comp (snd R A C)) + + +@[simp, norm_cast] +theorem coe_prodMap {D : Type*} [Semiring D] [TopologicalSpace D] [Algebra R D] (f₁ : A →A[R] B) + (f₂ : C →A[R] D) : + (f₁.prodMap f₂ : A × C →ₐ[R] B × D) = (f₁ : A →ₐ[R] B).prodMap (f₂ : C →ₐ[R] D) := + rfl + +@[simp, norm_cast] +theorem coe_prodMap' {D : Type*} [Semiring D] [TopologicalSpace D] [Algebra R D] (f₁ : A →A[R] B) + (f₂ : C →A[R] D) : ⇑(f₁.prodMap f₂) = Prod.map f₁ f₂ := + rfl + +/-- `ContinuousAlgHom.prod` as an `Equiv`. -/ +@[simps apply] +def prodEquiv : (A →A[R] B) × (A →A[R] C) ≃ (A →A[R] B × C) where + toFun f := f.1.prod f.2 + invFun f := ⟨(fst _ _ _).comp f, (snd _ _ _).comp f⟩ + left_inv f := by ext <;> rfl + right_inv f := by ext <;> rfl + +end prod + +section subalgebra + +variable {R A} +variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] + +/-- Restrict codomain of a continuous algebra morphism. -/ +def codRestrict (f : A →A[R] B) (p : Subalgebra R B) (h : ∀ x, f x ∈ p) : A →A[R] p where + cont := f.continuous.subtype_mk _ + toAlgHom := (f : A →ₐ[R] B).codRestrict p h + +@[norm_cast] +theorem coe_codRestrict (f : A →A[R] B) (p : Subalgebra R B) (h : ∀ x, f x ∈ p) : + (f.codRestrict p h : A →ₐ[R] p) = (f : A →ₐ[R] B).codRestrict p h := + rfl + +@[simp] +theorem coe_codRestrict_apply (f : A →A[R] B) (p : Subalgebra R B) (h : ∀ x, f x ∈ p) (x) : + (f.codRestrict p h x : B) = f x := + rfl + +/-- Restrict the codomain of a continuous algebra homomorphism `f` to `f.range`. -/ +@[reducible] +def rangeRestrict (f : A →A[R] B) := + f.codRestrict (@AlgHom.range R A B _ _ _ _ _ f) (@AlgHom.mem_range_self R A B _ _ _ _ _ f) + +@[simp] +theorem coe_rangeRestrict (f : A →A[R] B) : + (f.rangeRestrict : A →ₐ[R] (@AlgHom.range R A B _ _ _ _ _ f)) = + (f : A →ₐ[R] B).rangeRestrict := + rfl + +/-- `Subalgebra.val` as a `ContinuousAlgHom`. -/ +def _root_.Subalgebra.valA (p : Subalgebra R A) : p →A[R] A where + cont := continuous_subtype_val + toAlgHom := p.val + +@[simp, norm_cast] +theorem _root_.Subalgebra.coe_valA (p : Subalgebra R A) : + (p.valA : p →ₐ[R] A) = p.subtype := + rfl @[simp] -theorem Subalgebra.topologicalClosure_coe (s : Subalgebra R A) : - (s.topologicalClosure : Set A) = closure (s : Set A) := +theorem _root_.Subalgebra.coe_valA' (p : Subalgebra R A) : ⇑p.valA = p.subtype := rfl -instance Subalgebra.topologicalSemiring (s : Subalgebra R A) : TopologicalSemiring s := +@[simp] +theorem _root_.Subalgebra.valA_apply (p : Subalgebra R A) (x : p) : p.valA x = x := + rfl + +@[simp] +theorem _root_.Submodule.range_valA (p : Subalgebra R A) : + @AlgHom.range R p A _ _ _ _ _ p.valA = p := + Subalgebra.range_val p + +end subalgebra + +section Ring + + +variable {S : Type*} [Ring S] [TopologicalSpace S] [Algebra R S] {B : Type*} [Ring B] + [TopologicalSpace B] [Algebra R B] + +protected theorem map_neg (f : S →A[R] B) (x : S) : f (-x) = -f x := map_neg f x + +protected theorem map_sub (f : S →A[R] B) (x y : S) : f (x - y) = f x - f y := map_sub f x y + +end Ring + + +section RestrictScalars + +variable {S : Type*} [CommSemiring S] [Algebra R S] {B : Type*} [Ring B] [TopologicalSpace B] + [Algebra R B] [Algebra S B] [IsScalarTower R S B] {C : Type*} [Ring C] [TopologicalSpace C] + [Algebra R C] [Algebra S C] [IsScalarTower R S C] + +/-- If `A` is an `R`-algebra, then a continuous `A`-algebra morphism can be interpreted as a +continuous `R`-algebra morphism. -/ +def restrictScalars (f : B →A[S] C) : B →A[R] C := + ⟨(f : B →ₐ[S] C).restrictScalars R, f.continuous⟩ + +variable {R} + +@[simp] +theorem coe_restrictScalars (f : B →A[S] C) : + (f.restrictScalars R : B →ₐ[R] C) = (f : B →ₐ[S] C).restrictScalars R := + rfl + +@[simp] +theorem coe_restrictScalars' (f : B →A[S] C) : ⇑(f.restrictScalars R) = f := + rfl + +end RestrictScalars + +end ContinuousAlgHom + +end + +variable {R : Type*} [CommSemiring R] +variable {A : Type u} [TopologicalSpace A] +variable [Semiring A] [Algebra R A] +variable [TopologicalSemiring A] + +instance (s : Subalgebra R A) : TopologicalSemiring s := s.toSubsemiring.topologicalSemiring theorem Subalgebra.le_topologicalClosure (s : Subalgebra R A) : s ≤ s.topologicalClosure := diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 215374cc40925..f9621b990d249 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -304,7 +304,7 @@ theorem closure_smul₀ {E} [Zero E] [MulActionWithZero G₀ E] [TopologicalSpac /-- `smul` is a closed map in the second argument. The lemma that `smul` is a closed map in the first argument (for a normed space over a complete -normed field) is `isClosedMap_smul_left` in `Analysis.NormedSpace.FiniteDimension`. -/ +normed field) is `isClosedMap_smul_left` in `Analysis.Normed.Module.FiniteDimension`. -/ theorem isClosedMap_smul_of_ne_zero {c : G₀} (hc : c ≠ 0) : IsClosedMap fun x : α => c • x := (Homeomorph.smulOfNeZero c hc).isClosedMap @@ -315,7 +315,7 @@ theorem IsClosed.smul_of_ne_zero {c : G₀} {s : Set α} (hs : IsClosed s) (hc : /-- `smul` is a closed map in the second argument. The lemma that `smul` is a closed map in the first argument (for a normed space over a complete -normed field) is `isClosedMap_smul_left` in `Analysis.NormedSpace.FiniteDimension`. -/ +normed field) is `isClosedMap_smul_left` in `Analysis.Normed.Module.FiniteDimension`. -/ theorem isClosedMap_smul₀ {E : Type*} [Zero E] [MulActionWithZero G₀ E] [TopologicalSpace E] [T1Space E] [ContinuousConstSMul G₀ E] (c : G₀) : IsClosedMap fun x : E => c • x := by rcases eq_or_ne c 0 with (rfl | hne) diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 12e0336b38fcc..75dd9c5b48682 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -1245,7 +1245,7 @@ theorem QuotientGroup.isClosedMap_coe {H : Subgroup G} (hH : IsCompact (H : Set @[to_additive] lemma subset_mul_closure_one (s : Set G) : s ⊆ s * (closure {1} : Set G) := by - have : s ⊆ s * ({1} : Set G) := by simpa using Subset.rfl + have : s ⊆ s * ({1} : Set G) := by simp exact this.trans (smul_subset_smul_left subset_closure) @[to_additive] diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean index 5e29b0c172c36..c98296b0fb61f 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean @@ -115,7 +115,7 @@ end DivisionSemiring In this section, we prove various results about `(∑' x : ι, f x) * (∑' y : κ, g y)`. Note that we always assume that the family `fun x : ι × κ ↦ f x.1 * g x.2` is summable, since there is no way to deduce this from the summabilities of `f` and `g` in general, but if you are working in a normed -space, you may want to use the analogous lemmas in `Analysis/NormedSpace/Basic` +space, you may want to use the analogous lemmas in `Analysis.Normed.Module.Basic` (e.g `tsum_mul_tsum_of_summable_norm`). We first establish results about arbitrary index types, `ι` and `κ`, and then we specialize to diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index d6d41a0849313..1cfd1e13e0032 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -30,7 +30,7 @@ When `E` is a normed space, this gets us the equivalence of norms in finite dime ## TODO -Generalize more of `Mathlib.Analysis.NormedSpace.FiniteDimension` to general TVSs. +Generalize more of `Mathlib.Analysis.Normed.Module.FiniteDimension` to general TVSs. ## Implementation detail diff --git a/Mathlib/Topology/Algebra/Module/Star.lean b/Mathlib/Topology/Algebra/Module/Star.lean index 303eff798e2ab..4fccc6fec59a0 100644 --- a/Mathlib/Topology/Algebra/Module/Star.lean +++ b/Mathlib/Topology/Algebra/Module/Star.lean @@ -49,11 +49,11 @@ variable (R : Type*) (A : Type*) [Semiring R] [StarMul R] [TrivialStar R] [AddCo [Module R A] [StarAddMonoid A] [StarModule R A] [Invertible (2 : R)] [TopologicalSpace A] theorem continuous_selfAdjointPart [ContinuousAdd A] [ContinuousStar A] [ContinuousConstSMul R A] : - Continuous (@selfAdjointPart R A _ _ _ _ _ _ _ _) := + Continuous (selfAdjointPart R (A := A)) := ((continuous_const_smul _).comp <| continuous_id.add continuous_star).subtype_mk _ theorem continuous_skewAdjointPart [ContinuousSub A] [ContinuousStar A] [ContinuousConstSMul R A] : - Continuous (@skewAdjointPart R A _ _ _ _ _ _ _ _) := + Continuous (skewAdjointPart R (A := A)) := ((continuous_const_smul _).comp <| continuous_id.sub continuous_star).subtype_mk _ theorem continuous_decomposeProdAdjoint [TopologicalAddGroup A] [ContinuousStar A] diff --git a/Mathlib/Topology/Algebra/OpenSubgroup.lean b/Mathlib/Topology/Algebra/OpenSubgroup.lean index 807f4a84deed0..025b15e34f7a6 100644 --- a/Mathlib/Topology/Algebra/OpenSubgroup.lean +++ b/Mathlib/Topology/Algebra/OpenSubgroup.lean @@ -238,13 +238,13 @@ theorem comap_comap {P : Type*} [Group P] [TopologicalSpace P] (K : OpenSubgroup rfl end OpenSubgroup - namespace Subgroup -variable {G : Type*} [Group G] [TopologicalSpace G] [ContinuousMul G] (H : Subgroup G) +variable {G : Type*} [Group G] [TopologicalSpace G] @[to_additive] -theorem isOpen_of_mem_nhds {g : G} (hg : (H : Set G) ∈ 𝓝 g) : IsOpen (H : Set G) := by +theorem isOpen_of_mem_nhds [ContinuousMul G] (H : Subgroup G) {g : G} (hg : (H : Set G) ∈ 𝓝 g) : + IsOpen (H : Set G) := by refine isOpen_iff_mem_nhds.2 fun x hx => ?_ have hg' : g ∈ H := SetLike.mem_coe.1 (mem_of_mem_nhds hg) have : Filter.Tendsto (fun y => y * (x⁻¹ * g)) (𝓝 x) (𝓝 g) := @@ -253,19 +253,20 @@ theorem isOpen_of_mem_nhds {g : G} (hg : (H : Set G) ∈ 𝓝 g) : IsOpen (H : S H.mul_mem_cancel_right (H.mul_mem (H.inv_mem hx) hg')] using this hg @[to_additive] -theorem isOpen_mono {H₁ H₂ : Subgroup G} (h : H₁ ≤ H₂) (h₁ : IsOpen (H₁ : Set G)) : - IsOpen (H₂ : Set G) := +theorem isOpen_mono [ContinuousMul G] {H₁ H₂ : Subgroup G} (h : H₁ ≤ H₂) + (h₁ : IsOpen (H₁ : Set G)) : IsOpen (H₂ : Set G) := isOpen_of_mem_nhds _ <| Filter.mem_of_superset (h₁.mem_nhds <| one_mem H₁) h @[to_additive] -theorem isOpen_of_openSubgroup {U : OpenSubgroup G} (h : ↑U ≤ H) : IsOpen (H : Set G) := +theorem isOpen_of_openSubgroup [ContinuousMul G] (H: Subgroup G) {U : OpenSubgroup G} (h : ↑U ≤ H) : + IsOpen (H : Set G) := isOpen_mono h U.isOpen /-- If a subgroup of a topological group has `1` in its interior, then it is open. -/ @[to_additive "If a subgroup of an additive topological group has `0` in its interior, then it is open."] -theorem isOpen_of_one_mem_interior (h_1_int : (1 : G) ∈ interior (H : Set G)) : - IsOpen (H : Set G) := +theorem isOpen_of_one_mem_interior [ContinuousMul G] (H: Subgroup G) + (h_1_int : (1 : G) ∈ interior (H : Set G)) : IsOpen (H : Set G) := isOpen_of_mem_nhds H <| mem_interior_iff_mem_nhds.1 h_1_int end Subgroup diff --git a/Mathlib/Topology/Algebra/ProperAction.lean b/Mathlib/Topology/Algebra/ProperAction.lean new file mode 100644 index 0000000000000..4834f7f44beb1 --- /dev/null +++ b/Mathlib/Topology/Algebra/ProperAction.lean @@ -0,0 +1,278 @@ +/- +Copyright (c) 2024 Anatole Dedeker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedeker, Etienne Marion, Florestan Martin-Baillon, Vincent Guirardel +-/ +import Mathlib.Topology.Algebra.Group.Basic +import Mathlib.Topology.Algebra.MulAction +import Mathlib.Topology.Maps.Proper.Basic +import Mathlib.Topology.Sequences + +/-! +# Proper group action + +In this file we define proper action of a group on a topological space, and we prove that in this +case the quotient space is T2. We also give equivalent definitions of proper action using +ultrafilters and show the transfer of proper action to a closed subgroup. +We give sufficent conditions on the topological space such that the action is properly discontinuous +(see `ProperlyDiscontinuousSMul`) if and only if it is continuous in +the first variable (see `ContinuousConstSMul`) and proper in the sense defined here. + +## Main definitions + +* `ProperSMul` : a group `G` acts properly on a topological space `X` + if the map `(g, x) ↦ (g • x, x)` is proper, in the sense of `IsProperMap`. + +## Main statements + +* `t2Space_quotient_mulAction_of_properSMul`: If a group `G` acts properly + on a topological space `X`, then the quotient space is Hausdorff (T2). +* `t2Space_of_properSMul_of_t2Group`: If a T2 group acts properly on a topological space, + then this topological space is T2. +* `properlyDiscontinuousSMul_iff_properSMul`: If a discrete group acts on a T2 space `X` such that + `X × X` is compactly generated, then the action is properly discontinuous if and only if it is + continuous in the second variable and proper. This in particular true if `X` is locally compact + or first-countable. + +## Implementation notes + +Concerning `properlyDiscontinuousSMul_iff_properSMul`, this result should be the only one needed +to link properly discontinuous and proper actions. + +TODO: Replace the compactly generated hypothesis by a typeclass instance such that +`WeaklyLocallyCompactSpace.isProperMap_iff_isCompact_preimage` and +`SequentialSpace.isProperMap_iff_isCompact_preimage` are inferred by typeclass inference. + +## References + +* [N. Bourbaki, *General Topology*][bourbaki1966] + +## Tags + +Hausdorff, group action, proper action +-/ + +open Filter Topology Set Prod + +/-- Proper group action in the sense of Bourbaki: +the map `G × X → X × X` is a proper map (see `IsProperMap`). -/ +class ProperVAdd (G X : Type*) [TopologicalSpace G] [TopologicalSpace X] [AddGroup G] + [AddAction G X] : Prop where + /-- Proper group action in the sense of Bourbaki: + the map `G × X → X × X` is a proper map (see `IsProperMap`). -/ + isProperMap_vadd_pair : IsProperMap (fun gx ↦ (gx.1 +ᵥ gx.2, gx.2) : G × X → X × X) + +/-- Proper group action in the sense of Bourbaki: +the map `G × X → X × X` is a proper map (see `IsProperMap`). -/ +@[to_additive existing (attr := mk_iff)] +class ProperSMul (G X : Type*) [TopologicalSpace G] [TopologicalSpace X] [Group G] + [MulAction G X] : Prop where + /-- Proper group action in the sense of Bourbaki: + the map `G × X → X × X` is a proper map (see `IsProperMap`). -/ + isProperMap_smul_pair : IsProperMap (fun gx ↦ (gx.1 • gx.2, gx.2) : G × X → X × X) + +attribute [to_additive existing] properSMul_iff + +variable {G X Y Z : Type*} [Group G] [MulAction G X] [MulAction G Y] +variable [TopologicalSpace G] [TopologicalSpace X] [TopologicalSpace Y] +variable [TopologicalSpace Z] + +/-- If a group acts properly then in particular it acts continuously. -/ +@[to_additive "If a group acts properly then in particular it acts continuously."] +-- See note [lower instance property] +instance (priority := 100) ProperSMul.toContinuousSMul [ProperSMul G X] : ContinuousSMul G X where + continuous_smul := isProperMap_smul_pair.continuous.fst + +/-- A group `G` acts properly on a topological space `X` if and only if for all ultrafilters +`𝒰` on `X × G`, if `𝒰` converges to `(x₁, x₂)` along the map `(g, x) ↦ (g • x, x)`, +then there exists `g : G` such that `g • x₂ = x₁` and `𝒰.fst` converges to `g`. -/ +@[to_additive "A group acts `G` properly on a topological space `X` if and only if +for all ultrafilters `𝒰` on `X`, if `𝒰` converges to `(x₁, x₂)` +along the map `(g, x) ↦ (g • x, x)`, then there exists `g : G` such that `g • x₂ = x₁` +and `𝒰.fst` converges to `g`."] +theorem properSMul_iff_continuousSMul_ultrafilter_tendsto : + ProperSMul G X ↔ ContinuousSMul G X ∧ + (∀ 𝒰 : Ultrafilter (G × X), ∀ x₁ x₂ : X, + Tendsto (fun gx : G × X ↦ (gx.1 • gx.2, gx.2)) 𝒰 (𝓝 (x₁, x₂)) → + ∃ g : G, g • x₂ = x₁ ∧ Tendsto (Prod.fst : G × X → G) 𝒰 (𝓝 g)) := by + refine ⟨fun h ↦ ⟨inferInstance, fun 𝒰 x₁ x₂ h' ↦ ?_⟩, fun ⟨cont, h⟩ ↦ ?_⟩ + · rw [properSMul_iff, isProperMap_iff_ultrafilter] at h + rcases h.2 h' with ⟨gx, hgx1, hgx2⟩ + refine ⟨gx.1, ?_, (continuous_fst.tendsto gx).mono_left hgx2⟩ + simp only [Prod.mk.injEq] at hgx1 + rw [← hgx1.2, hgx1.1] + · rw [properSMul_iff, isProperMap_iff_ultrafilter] + refine ⟨by fun_prop, fun 𝒰 (x₁, x₂) hxx ↦ ?_⟩ + rcases h 𝒰 x₁ x₂ hxx with ⟨g, hg1, hg2⟩ + refine ⟨(g, x₂), by simp_rw [hg1], ?_⟩ + rw [nhds_prod_eq, 𝒰.le_prod] + exact ⟨hg2, (continuous_snd.tendsto _).comp hxx⟩ + +/-- A group `G` acts properly on a T2 topological space `X` if and only if for all ultrafilters +`𝒰` on `X × G`, if `𝒰` converges to `(x₁, x₂)` along the map `(g, x) ↦ (g • x, x)`, +then there exists `g : G` such that `𝒰.fst` converges to `g`. -/ +theorem properSMul_iff_continuousSMul_ultrafilter_tendsto_t2 [T2Space X] : + ProperSMul G X ↔ ContinuousSMul G X ∧ + (∀ 𝒰 : Ultrafilter (G × X), ∀ x₁ x₂ : X, + Tendsto (fun gx : G × X ↦ (gx.1 • gx.2, gx.2)) 𝒰 (𝓝 (x₁, x₂)) → + ∃ g : G, Tendsto (Prod.fst : G × X → G) 𝒰 (𝓝 g)) := by + rw [properSMul_iff_continuousSMul_ultrafilter_tendsto] + refine and_congr_right fun hc ↦ ?_ + congrm ∀ 𝒰 x₁ x₂ hxx, ∃ g, ?_ + exact and_iff_right_of_imp fun hg ↦ tendsto_nhds_unique + (hg.smul ((continuous_snd.tendsto _).comp hxx)) ((continuous_fst.tendsto _).comp hxx) + +/-- If `G` acts properly on `X`, then the quotient space is Hausdorff (T2). -/ +@[to_additive "If `G` acts properly on `X`, then the quotient space is Hausdorff (T2)."] +theorem t2Space_quotient_mulAction_of_properSMul [ProperSMul G X] : + T2Space (Quotient (MulAction.orbitRel G X)) := by + rw [t2_iff_isClosed_diagonal] + set R := MulAction.orbitRel G X + let π : X → Quotient R := Quotient.mk' + have : QuotientMap (Prod.map π π) := + (isOpenMap_quotient_mk'_mul.prod isOpenMap_quotient_mk'_mul).to_quotientMap + (continuous_quotient_mk'.prod_map continuous_quotient_mk') + ((surjective_quotient_mk' _).prodMap (surjective_quotient_mk' _)) + rw [← this.isClosed_preimage] + convert ProperSMul.isProperMap_smul_pair.isClosedMap.isClosed_range + · ext ⟨x₁, x₂⟩ + simp only [mem_preimage, map_apply, mem_diagonal_iff, mem_range, Prod.mk.injEq, Prod.exists, + exists_eq_right] + rw [Quotient.eq_rel, MulAction.orbitRel_apply, MulAction.mem_orbit_iff] + all_goals infer_instance + +/-- If a T2 group acts properly on a topological space, then this topological space is T2. -/ +@[to_additive "If a T2 group acts properly on a topological space, +then this topological space is T2."] +theorem t2Space_of_properSMul_of_t2Group [h_proper : ProperSMul G X] [T2Space G] : T2Space X := by + let f := fun x : X ↦ ((1 : G), x) + have proper_f : IsProperMap f := by + apply isProperMap_of_closedEmbedding + rw [closedEmbedding_iff] + constructor + · let g := fun gx : G × X ↦ gx.2 + have : Function.LeftInverse g f := fun x ↦ by simp + exact this.embedding (by fun_prop) (by fun_prop) + · have : range f = ({1} ×ˢ univ) := by simp + rw [this] + exact isClosed_singleton.prod isClosed_univ + rw [t2_iff_isClosed_diagonal] + let g := fun gx : G × X ↦ (gx.1 • gx.2, gx.2) + have proper_g : IsProperMap g := (properSMul_iff G X).1 h_proper + have : g ∘ f = fun x ↦ (x, x) := by ext x <;> simp + have range_gf : range (g ∘ f) = diagonal X := by simp [this] + rw [← range_gf] + exact (proper_f.comp proper_g).isClosed_range + +/-- If two groups `H` and `G` act on a topological space `X` such that `G` acts properly and +there exists a group homomorphims `H → G` which is a closed embedding compatible with the actions, +then `H` also acts properly on `X`. -/ +@[to_additive "If two groups `H` and `G` act on a topological space `X` such that `G` acts properly +and there exists a group homomorphims `H → G` which is a closed embedding compatible with the +actions, then `H` also acts properly on `X`."] +theorem properSMul_of_closedEmbedding {H : Type*} [Group H] [MulAction H X] [TopologicalSpace H] + [ProperSMul G X] (f : H →* G) (f_clemb : ClosedEmbedding f) + (f_compat : ∀ (h : H) (x : X), f h • x = h • x) : ProperSMul H X where + isProperMap_smul_pair := by + have := isProperMap_of_closedEmbedding f_clemb + have h : IsProperMap (Prod.map f (fun x : X ↦ x)) := IsProperMap.prod_map this isProperMap_id + have : (fun hx : H × X ↦ (hx.1 • hx.2, hx.2)) = (fun hx ↦ (f hx.1 • hx.2, hx.2)) := by + simp [f_compat] + rw [this] + exact h.comp <| ProperSMul.isProperMap_smul_pair + +/-- If `H` is a closed subgroup of `G` and `G` acts properly on X then so does `H`. -/ +@[to_additive "If `H` is a closed subgroup of `G` and `G` acts properly on X then so does `H`."] +instance {H : Subgroup G} [ProperSMul G X] [H_closed : IsClosed (H : Set G)] : ProperSMul H X := + properSMul_of_closedEmbedding H.subtype H_closed.closedEmbedding_subtype_val fun _ _ ↦ rfl + +/-- If a discrete group acts on a T2 space `X` such that `X × X` is compactly generated, +then the action is properly discontinuous if and only if it is continuous in the second variable +and proper. -/ +theorem properlyDiscontinuousSMul_iff_properSMul [T2Space X] [DiscreteTopology G] + [ContinuousConstSMul G X] + (compactlyGenerated : ∀ s : Set (X × X), IsClosed s ↔ ∀ ⦃K⦄, IsCompact K → IsClosed (s ∩ K)) : + ProperlyDiscontinuousSMul G X ↔ ProperSMul G X := by + constructor + · intro h + rw [properSMul_iff] + -- We have to show that `f : (g, x) ↦ (g • x, x)` is proper. + -- Continuity follows from continuity of `g • ·` and the fact that `G` has the + -- discrete topology, thanks to `continuous_of_partial_of_discrete`. + -- Because `X × X` is compactly generated, to show that f is proper + -- it is enough to show that the preimage of a compact set `K` is compact. + refine (isProperMap_iff_isCompact_preimage compactlyGenerated).2 + ⟨(continuous_prod_mk.2 + ⟨continuous_prod_of_discrete_left.2 continuous_const_smul, by fun_prop⟩), + fun K hK ↦ ?_⟩ + -- We set `K' := pr₁(K) ∪ pr₂(K)`, which is compact because `K` is compact and `pr₁` and + -- `pr₂` are continuous. We halso have that `K ⊆ K' × K'`, and `K` is closed because `X` is T2. + -- Therefore `f ⁻¹ (K)` is also closed and `f ⁻¹ (K) ⊆ f ⁻¹ (K' × K')`, thus it suffices to + -- show that `f ⁻¹ (K' × K')` is compact. + let K' := fst '' K ∪ snd '' K + have hK' : IsCompact K' := (hK.image continuous_fst).union (hK.image continuous_snd) + let E := {g : G | Set.Nonempty ((g • ·) '' K' ∩ K')} + -- The set `E` is finite because the action is properly discontinuous. + have fin : Set.Finite E := by + simp_rw [E, nonempty_iff_ne_empty] + exact h.finite_disjoint_inter_image hK' hK' + -- Therefore we can rewrite `f ⁻¹ (K' × K')` as a finite union of compact sets. + have : (fun gx : G × X ↦ (gx.1 • gx.2, gx.2)) ⁻¹' (K' ×ˢ K') = + ⋃ g ∈ E, {g} ×ˢ ((g⁻¹ • ·) '' K' ∩ K') := by + ext gx + simp only [mem_preimage, mem_prod, nonempty_def, mem_inter_iff, mem_image, + exists_exists_and_eq_and, mem_setOf_eq, singleton_prod, iUnion_exists, biUnion_and', + mem_iUnion, exists_prop, E] + constructor + · exact fun ⟨gx_mem, x_mem⟩ ↦ ⟨gx.2, x_mem, gx.1, gx_mem, + ⟨gx.2, ⟨⟨gx.1 • gx.2, gx_mem, by simp⟩, x_mem⟩, rfl⟩⟩ + · rintro ⟨x, -, g, -, ⟨-, ⟨⟨x', x'_mem, rfl⟩, ginvx'_mem⟩, rfl⟩⟩ + exact ⟨by simpa, by simpa⟩ + -- Indeed each set in this finite union is the product of a singleton and + -- the intersection of the compact `K'` with its image by some element `g`, and this image is + -- compact because `g • ·` is continuous. + have : IsCompact ((fun gx : G × X ↦ (gx.1 • gx.2, gx.2)) ⁻¹' (K' ×ˢ K')) := + this ▸ fin.isCompact_biUnion fun g hg ↦ + isCompact_singleton.prod <| (hK'.image <| continuous_const_smul _).inter hK' + -- We conclude as explained above. + exact this.of_isClosed_subset (hK.isClosed.preimage <| + continuous_prod_mk.2 + ⟨continuous_prod_of_discrete_left.2 continuous_const_smul, by fun_prop⟩) <| + preimage_mono fun x hx ↦ ⟨Or.inl ⟨x, hx, rfl⟩, Or.inr ⟨x, hx, rfl⟩⟩ + · intro h; constructor + intro K L hK hL + simp_rw [← nonempty_iff_ne_empty] + -- We want to show that a subset of `G` is finite, but as `G` has the discrete topology it + -- is enough to show that this subset is compact. + apply IsCompact.finite_of_discrete + -- Now set `h : (g, x) ↦ (g⁻¹ • x, x)`, because `f` is proper by hypothesis, so is `h`. + have : IsProperMap (fun gx : G × X ↦ (gx.1⁻¹ • gx.2, gx.2)) := + (IsProperMap.prod_map (Homeomorph.isProperMap (Homeomorph.inv G)) isProperMap_id).comp <| + ProperSMul.isProperMap_smul_pair + --But we also have that `{g | Set.Nonempty ((g • ·) '' K ∩ L)} = h ⁻¹ (K × L)`, which + -- concludes the proof. + have eq : {g | Set.Nonempty ((g • ·) '' K ∩ L)} = + fst '' ((fun gx : G × X ↦ (gx.1⁻¹ • gx.2, gx.2)) ⁻¹' (K ×ˢ L)) := by + simp_rw [nonempty_def] + ext g; constructor + · exact fun ⟨_, ⟨x, x_mem, rfl⟩, hx⟩ ↦ ⟨(g, g • x), ⟨by simpa, hx⟩, rfl⟩ + · rintro ⟨gx, hgx, rfl⟩ + exact ⟨gx.2, ⟨gx.1⁻¹ • gx.2, hgx.1, by simp⟩, hgx.2⟩ + exact eq ▸ IsCompact.image (this.isCompact_preimage <| hK.prod hL) continuous_fst + +/-- If a discrete group acts on a T2 and locally compact space `X`, +then the action is properly discontinuous if and only if it is continuous in the second variable +and proper. -/ +theorem WeaklyLocallyCompactSpace.properlyDiscontinuousSMul_iff_properSMul [T2Space X] + [WeaklyLocallyCompactSpace X] [DiscreteTopology G] [ContinuousConstSMul G X] : + ProperlyDiscontinuousSMul G X ↔ ProperSMul G X := + _root_.properlyDiscontinuousSMul_iff_properSMul + (fun _ ↦ compactlyGenerated_of_weaklyLocallyCompactSpace) + +/-- If a discrete group acts on a T2 and first-countable space `X`, +then the action is properly discontinuous if and only if it is continuous in the second variable +and proper. -/ +theorem FirstCountableTopology.properlyDiscontinuousSMul_iff_properSMul [T2Space X] + [FirstCountableTopology X] [DiscreteTopology G] [ContinuousConstSMul G X] : + ProperlyDiscontinuousSMul G X ↔ ProperSMul G X := + _root_.properlyDiscontinuousSMul_iff_properSMul (fun _ ↦ compactlyGenerated_of_sequentialSpace) diff --git a/Mathlib/Topology/Algebra/UniformConvergence.lean b/Mathlib/Topology/Algebra/UniformConvergence.lean index 9401745a13ff7..adf3afef18aff 100644 --- a/Mathlib/Topology/Algebra/UniformConvergence.lean +++ b/Mathlib/Topology/Algebra/UniformConvergence.lean @@ -224,10 +224,7 @@ instance : UniformGroup (α →ᵤ G) := protected theorem UniformFun.hasBasis_nhds_one_of_basis {p : ι → Prop} {b : ι → Set G} (h : (𝓝 1 : Filter G).HasBasis p b) : (𝓝 1 : Filter (α →ᵤ G)).HasBasis p fun i => { f : α →ᵤ G | ∀ x, toFun f x ∈ b i } := by - have := h.comap fun p : G × G => p.2 / p.1 - rw [← uniformity_eq_comap_nhds_one] at this - convert UniformFun.hasBasis_nhds_of_basis α _ (1 : α →ᵤ G) this - -- Porting note: removed `ext i f` here, as it has already been done by `convert`. + convert UniformFun.hasBasis_nhds_of_basis α _ (1 : α →ᵤ G) h.uniformity_of_nhds_one simp @[to_additive] @@ -255,10 +252,8 @@ protected theorem UniformOnFun.hasBasis_nhds_one_of_basis (𝔖 : Set <| Set α) (h : (𝓝 1 : Filter G).HasBasis p b) : (𝓝 1 : Filter (α →ᵤ[𝔖] G)).HasBasis (fun Si : Set α × ι => Si.1 ∈ 𝔖 ∧ p Si.2) fun Si => { f : α →ᵤ[𝔖] G | ∀ x ∈ Si.1, toFun 𝔖 f x ∈ b Si.2 } := by - have := h.comap fun p : G × G => p.1 / p.2 - rw [← uniformity_eq_comap_nhds_one_swapped] at this + have := h.uniformity_of_nhds_one_swapped convert UniformOnFun.hasBasis_nhds_of_basis α _ 𝔖 (1 : α →ᵤ[𝔖] G) h𝔖₁ h𝔖₂ this - -- Porting note: removed `ext i f` here, as it has already been done by `convert`. simp [UniformOnFun.gen] @[to_additive] diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 2c45510a39caa..4bdd67535c6e8 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -151,7 +151,9 @@ instance instField : Field (hat K) where inv_zero := by simp only [Inv.inv, ite_true] -- TODO: use a better defeq nnqsmul := _ + nnqsmul_def := fun q a => rfl qsmul := _ + qsmul_def := fun a x => rfl instance : TopologicalDivisionRing (hat K) := { Completion.topologicalRing with diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 3b572491eb205..cffb8dc1888fe 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -984,7 +984,7 @@ theorem accPt_iff_frequently (x : X) (C : Set X) : AccPt x (𝓟 C) ↔ ∃ᶠ y simp [acc_principal_iff_cluster, clusterPt_principal_iff_frequently, and_comm] /-- If `x` is an accumulation point of `F` and `F ≤ G`, then -`x` is an accumulation point of `D`. -/ +`x` is an accumulation point of `G`. -/ theorem AccPt.mono {F G : Filter X} (h : AccPt x F) (hFG : F ≤ G) : AccPt x G := NeBot.mono h (inf_le_inf_left _ hFG) diff --git a/Mathlib/Topology/Category/CompactlyGenerated.lean b/Mathlib/Topology/Category/CompactlyGenerated.lean index 334a0a40711ee..624ffb1502587 100644 --- a/Mathlib/Topology/Category/CompactlyGenerated.lean +++ b/Mathlib/Topology/Category/CompactlyGenerated.lean @@ -50,34 +50,34 @@ A topological space `X` is compactly generated if its topology is finer than (an the compactly generated topology, i.e. it is coinduced by the continuous maps from compact Hausdorff spaces to `X`. -/ -class CompactlyGeneratedSpace (X : Type w) [t : TopologicalSpace X] : Prop where +class UCompactlyGeneratedSpace (X : Type w) [t : TopologicalSpace X] : Prop where /-- The topology of `X` is finer than the compactly generated topology. -/ le_compactlyGenerated : t ≤ compactlyGenerated.{u} X -lemma eq_compactlyGenerated {X : Type w} [t : TopologicalSpace X] [CompactlyGeneratedSpace.{u} X] : +lemma eq_compactlyGenerated {X : Type w} [t : TopologicalSpace X] [UCompactlyGeneratedSpace.{u} X] : t = compactlyGenerated.{u} X := by apply le_antisymm - · exact CompactlyGeneratedSpace.le_compactlyGenerated + · exact UCompactlyGeneratedSpace.le_compactlyGenerated · simp only [compactlyGenerated, ← continuous_iff_coinduced_le, continuous_sigma_iff, Sigma.forall] exact fun S f ↦ f.2 instance (X : Type w) [t : TopologicalSpace X] [DiscreteTopology X] : - CompactlyGeneratedSpace.{u} X where + UCompactlyGeneratedSpace.{u} X where le_compactlyGenerated := by rw [DiscreteTopology.eq_bot (t := t)] exact bot_le -lemma continuous_from_compactlyGeneratedSpace {X : Type w} [TopologicalSpace X] - [CompactlyGeneratedSpace.{u} X] {Y : Type*} [TopologicalSpace Y] (f : X → Y) +lemma continuous_from_uCompactlyGeneratedSpace {X : Type w} [TopologicalSpace X] + [UCompactlyGeneratedSpace.{u} X] {Y : Type*} [TopologicalSpace Y] (f : X → Y) (h : ∀ (S : CompHaus.{u}) (g : C(S, X)), Continuous (f ∘ g)) : Continuous f := by - apply continuous_le_dom CompactlyGeneratedSpace.le_compactlyGenerated + apply continuous_le_dom UCompactlyGeneratedSpace.le_compactlyGenerated exact continuous_from_compactlyGenerated f h -lemma compactlyGeneratedSpace_of_continuous_maps {X : Type w} [t : TopologicalSpace X] +lemma uCompactlyGeneratedSpace_of_continuous_maps {X : Type w} [t : TopologicalSpace X] (h : ∀ {Y : Type w} [tY : TopologicalSpace Y] (f : X → Y), (∀ (S : CompHaus.{u}) (g : C(S, X)), Continuous (f ∘ g)) → Continuous f) : - CompactlyGeneratedSpace.{u} X where + UCompactlyGeneratedSpace.{u} X where le_compactlyGenerated := by suffices Continuous[t, compactlyGenerated.{u} X] (id : X → X) by rwa [← continuous_id_iff_le] @@ -94,7 +94,7 @@ structure CompactlyGenerated where /-- The underlying topological space of an object of `CompactlyGenerated`. -/ toTop : TopCat.{w} /-- The underlying topological space is compactly generated. -/ - [is_compactly_generated : CompactlyGeneratedSpace.{u} toTop] + [is_compactly_generated : UCompactlyGeneratedSpace.{u} toTop] namespace CompactlyGenerated @@ -112,7 +112,7 @@ instance : Category.{w, w+1} CompactlyGenerated.{u, w} := instance : ConcreteCategory.{w} CompactlyGenerated.{u, w} := InducedCategory.concreteCategory _ -variable (X : Type w) [TopologicalSpace X] [CompactlyGeneratedSpace.{u} X] +variable (X : Type w) [TopologicalSpace X] [UCompactlyGeneratedSpace.{u} X] /-- Constructor for objects of the category `CompactlyGenerated`. -/ def of : CompactlyGenerated.{u, w} where diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index fcfee7f8869f7..c54867416cfdc 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -819,8 +819,8 @@ instance (α : Type*) [TopologicalSpace α] : NoZeroSMulDivisors ℤ (LocallyCon intro hc ext x apply mul_right_injective₀ hc - simp [LocallyConstant.ext_iff] at h ⊢ - exact h x + simp [LocallyConstant.ext_iff] at h + simpa [LocallyConstant.ext_iff] using h x theorem GoodProducts.linearIndependentSingleton : LinearIndependent ℤ (eval ({fun _ ↦ false} : Set (I → Bool))) := by diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index 6c5866c88232c..4ccef7c06b3e6 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -292,6 +292,10 @@ instance (priority := 900) nhdsWithin_compl_neBot [∀ x : X, NeBot (𝓝[≠] x theorem nhds_infty_eq : 𝓝 (∞ : OnePoint X) = map (↑) (coclosedCompact X) ⊔ pure ∞ := by rw [← nhdsWithin_compl_infty_eq, nhdsWithin_compl_singleton_sup_pure] +theorem tendsto_coe_infty : Tendsto (↑) (coclosedCompact X) (𝓝 (∞ : OnePoint X)) := by + rw [nhds_infty_eq] + exact Filter.Tendsto.mono_right tendsto_map le_sup_left + theorem hasBasis_nhds_infty : (𝓝 (∞ : OnePoint X)).HasBasis (fun s : Set X => IsClosed s ∧ IsCompact s) fun s => (↑) '' sᶜ ∪ {∞} := by diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 0c1d2d894bee8..1071071d3bfaa 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -978,6 +978,9 @@ theorem IsCompact.prod {t : Set Y} (hs : IsCompact s) (ht : IsCompact t) : instance (priority := 100) Finite.compactSpace [Finite X] : CompactSpace X where isCompact_univ := finite_univ.isCompact +instance ULift.compactSpace [CompactSpace X] : CompactSpace (ULift.{v} X) := + ULift.closedEmbedding_down.compactSpace + /-- The product of two compact spaces is compact. -/ instance [CompactSpace X] [CompactSpace Y] : CompactSpace (X × Y) := ⟨by rw [← univ_prod_univ]; exact isCompact_univ.prod isCompact_univ⟩ diff --git a/Mathlib/Topology/Connected/Basic.lean b/Mathlib/Topology/Connected/Basic.lean index c7e5e6c9e6ad8..af3cfad6372d4 100644 --- a/Mathlib/Topology/Connected/Basic.lean +++ b/Mathlib/Topology/Connected/Basic.lean @@ -3,9 +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 -/ -import Mathlib.Data.Set.Image +import Mathlib.Data.Set.Basic import Mathlib.Order.SuccPred.Relation -import Mathlib.Topology.Clopen import Mathlib.Topology.Irreducible /-! @@ -395,12 +394,6 @@ theorem IsPreconnected.subset_right_of_subset_union (hu : IsOpen u) (hv : IsOpen s ⊆ v := hs.subset_left_of_subset_union hv hu huv.symm (union_comm u v ▸ hsuv) hsv --- Porting note: moved up -/-- Preconnected sets are either contained in or disjoint to any given clopen set. -/ -theorem IsPreconnected.subset_isClopen {s t : Set α} (hs : IsPreconnected s) (ht : IsClopen t) - (hne : (s ∩ t).Nonempty) : s ⊆ t := - hs.subset_left_of_subset_union ht.isOpen ht.compl.isOpen disjoint_compl_right (by simp) hne - /-- If a preconnected set `s` intersects an open set `u`, and limit points of `u` inside `s` are contained in `u`, then the whole set `s` is contained in `u`. -/ theorem IsPreconnected.subset_of_closure_inter_subset (hs : IsPreconnected s) (hu : IsOpen u) @@ -459,59 +452,6 @@ theorem isConnected_univ_pi [∀ i, TopologicalSpace (π i)] {s : ∀ i, Set (π rw [← eval_image_univ_pi hne] exact hc.image _ (continuous_apply _).continuousOn -theorem Sigma.isConnected_iff [∀ i, TopologicalSpace (π i)] {s : Set (Σi, π i)} : - IsConnected s ↔ ∃ i t, IsConnected t ∧ s = Sigma.mk i '' t := by - refine ⟨fun hs => ?_, ?_⟩ - · obtain ⟨⟨i, x⟩, hx⟩ := hs.nonempty - have : s ⊆ range (Sigma.mk i) := - hs.isPreconnected.subset_isClopen isClopen_range_sigmaMk ⟨⟨i, x⟩, hx, x, rfl⟩ - exact ⟨i, Sigma.mk i ⁻¹' s, hs.preimage_of_isOpenMap sigma_mk_injective isOpenMap_sigmaMk this, - (Set.image_preimage_eq_of_subset this).symm⟩ - · rintro ⟨i, t, ht, rfl⟩ - exact ht.image _ continuous_sigmaMk.continuousOn - -theorem Sigma.isPreconnected_iff [hι : Nonempty ι] [∀ i, TopologicalSpace (π i)] - {s : Set (Σi, π i)} : IsPreconnected s ↔ ∃ i t, IsPreconnected t ∧ s = Sigma.mk i '' t := by - refine ⟨fun hs => ?_, ?_⟩ - · obtain rfl | h := s.eq_empty_or_nonempty - · exact ⟨Classical.choice hι, ∅, isPreconnected_empty, (Set.image_empty _).symm⟩ - · obtain ⟨a, t, ht, rfl⟩ := Sigma.isConnected_iff.1 ⟨h, hs⟩ - exact ⟨a, t, ht.isPreconnected, rfl⟩ - · rintro ⟨a, t, ht, rfl⟩ - exact ht.image _ continuous_sigmaMk.continuousOn - -theorem Sum.isConnected_iff [TopologicalSpace β] {s : Set (α ⊕ β)} : - IsConnected s ↔ - (∃ t, IsConnected t ∧ s = Sum.inl '' t) ∨ ∃ t, IsConnected t ∧ s = Sum.inr '' t := by - refine ⟨fun hs => ?_, ?_⟩ - · obtain ⟨x | x, hx⟩ := hs.nonempty - · have h : s ⊆ range Sum.inl := - hs.isPreconnected.subset_isClopen isClopen_range_inl ⟨.inl x, hx, x, rfl⟩ - refine Or.inl ⟨Sum.inl ⁻¹' s, ?_, ?_⟩ - · exact hs.preimage_of_isOpenMap Sum.inl_injective isOpenMap_inl h - · exact (image_preimage_eq_of_subset h).symm - · have h : s ⊆ range Sum.inr := - hs.isPreconnected.subset_isClopen isClopen_range_inr ⟨.inr x, hx, x, rfl⟩ - refine Or.inr ⟨Sum.inr ⁻¹' s, ?_, ?_⟩ - · exact hs.preimage_of_isOpenMap Sum.inr_injective isOpenMap_inr h - · exact (image_preimage_eq_of_subset h).symm - · rintro (⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩) - · exact ht.image _ continuous_inl.continuousOn - · exact ht.image _ continuous_inr.continuousOn - -theorem Sum.isPreconnected_iff [TopologicalSpace β] {s : Set (α ⊕ β)} : - IsPreconnected s ↔ - (∃ t, IsPreconnected t ∧ s = Sum.inl '' t) ∨ ∃ t, IsPreconnected t ∧ s = Sum.inr '' t := by - refine ⟨fun hs => ?_, ?_⟩ - · obtain rfl | h := s.eq_empty_or_nonempty - · exact Or.inl ⟨∅, isPreconnected_empty, (Set.image_empty _).symm⟩ - obtain ⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩ := Sum.isConnected_iff.1 ⟨h, hs⟩ - · exact Or.inl ⟨t, ht.isPreconnected, rfl⟩ - · exact Or.inr ⟨t, ht.isPreconnected, rfl⟩ - · rintro (⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩) - · exact ht.image _ continuous_inl.continuousOn - · exact ht.image _ continuous_inr.continuousOn - /-- The connected component of a point is the maximal connected set that contains this point. -/ def connectedComponent (x : α) : Set α := @@ -759,94 +699,6 @@ instance (priority := 100) PreirreducibleSpace.preconnectedSpace (α : Type u) [ instance (priority := 100) IrreducibleSpace.connectedSpace (α : Type u) [TopologicalSpace α] [IrreducibleSpace α] : ConnectedSpace α where toNonempty := IrreducibleSpace.toNonempty -/-- A continuous map from a connected space to a disjoint union `Σ i, π i` can be lifted to one of -the components `π i`. See also `ContinuousMap.exists_lift_sigma` for a version with bundled -`ContinuousMap`s. -/ -theorem Continuous.exists_lift_sigma [ConnectedSpace α] [∀ i, TopologicalSpace (π i)] - {f : α → Σ i, π i} (hf : Continuous f) : - ∃ (i : ι) (g : α → π i), Continuous g ∧ f = Sigma.mk i ∘ g := by - obtain ⟨i, hi⟩ : ∃ i, range f ⊆ range (.mk i) := by - rcases Sigma.isConnected_iff.1 (isConnected_range hf) with ⟨i, s, -, hs⟩ - exact ⟨i, hs.trans_subset (image_subset_range _ _)⟩ - rcases range_subset_range_iff_exists_comp.1 hi with ⟨g, rfl⟩ - refine ⟨i, g, ?_, rfl⟩ - rwa [← embedding_sigmaMk.continuous_iff] at hf - -theorem nonempty_inter [PreconnectedSpace α] {s t : Set α} : - IsOpen s → IsOpen t → s ∪ t = univ → s.Nonempty → t.Nonempty → (s ∩ t).Nonempty := by - simpa only [univ_inter, univ_subset_iff] using @PreconnectedSpace.isPreconnected_univ α _ _ s t - -theorem isClopen_iff [PreconnectedSpace α] {s : Set α} : IsClopen s ↔ s = ∅ ∨ s = univ := - ⟨fun hs => - by_contradiction fun h => - have h1 : s ≠ ∅ ∧ sᶜ ≠ ∅ := - ⟨mt Or.inl h, - mt (fun h2 => Or.inr <| (by rw [← compl_compl s, h2, compl_empty] : s = univ)) h⟩ - let ⟨_, h2, h3⟩ := - nonempty_inter hs.2 hs.1.isOpen_compl (union_compl_self s) (nonempty_iff_ne_empty.2 h1.1) - (nonempty_iff_ne_empty.2 h1.2) - h3 h2, - by rintro (rfl | rfl) <;> [exact isClopen_empty; exact isClopen_univ]⟩ - -theorem IsClopen.eq_univ [PreconnectedSpace α] {s : Set α} (h' : IsClopen s) (h : s.Nonempty) : - s = univ := - (isClopen_iff.mp h').resolve_left h.ne_empty - -section disjoint_subsets - -variable [PreconnectedSpace α] - {s : ι → Set α} (h_nonempty : ∀ i, (s i).Nonempty) (h_disj : Pairwise (Disjoint on s)) - -/-- In a preconnected space, any disjoint family of non-empty clopen subsets has at most one -element. -/ -lemma subsingleton_of_disjoint_isClopen - (h_clopen : ∀ i, IsClopen (s i)) : - Subsingleton ι := by - replace h_nonempty : ∀ i, s i ≠ ∅ := by intro i; rw [← nonempty_iff_ne_empty]; exact h_nonempty i - rw [← not_nontrivial_iff_subsingleton] - by_contra contra - obtain ⟨i, j, h_ne⟩ := contra - replace h_ne : s i ∩ s j = ∅ := by - simpa only [← bot_eq_empty, eq_bot_iff, ← inf_eq_inter, ← disjoint_iff_inf_le] using h_disj h_ne - cases' isClopen_iff.mp (h_clopen i) with hi hi - · exact h_nonempty i hi - · rw [hi, univ_inter] at h_ne - exact h_nonempty j h_ne - -/-- In a preconnected space, any disjoint cover by non-empty open subsets has at most one -element. -/ -lemma subsingleton_of_disjoint_isOpen_iUnion_eq_univ - (h_open : ∀ i, IsOpen (s i)) (h_Union : ⋃ i, s i = univ) : - Subsingleton ι := by - refine subsingleton_of_disjoint_isClopen h_nonempty h_disj (fun i ↦ ⟨?_, h_open i⟩) - rw [← isOpen_compl_iff, compl_eq_univ_diff, ← h_Union, iUnion_diff] - refine isOpen_iUnion (fun j ↦ ?_) - rcases eq_or_ne i j with rfl | h_ne - · simp - · simpa only [(h_disj h_ne.symm).sdiff_eq_left] using h_open j - -/-- In a preconnected space, any finite disjoint cover by non-empty closed subsets has at most one -element. -/ -lemma subsingleton_of_disjoint_isClosed_iUnion_eq_univ [Finite ι] - (h_closed : ∀ i, IsClosed (s i)) (h_Union : ⋃ i, s i = univ) : - Subsingleton ι := by - refine subsingleton_of_disjoint_isClopen h_nonempty h_disj (fun i ↦ ⟨h_closed i, ?_⟩) - rw [← isClosed_compl_iff, compl_eq_univ_diff, ← h_Union, iUnion_diff] - refine isClosed_iUnion_of_finite (fun j ↦ ?_) - rcases eq_or_ne i j with rfl | h_ne - · simp - · simpa only [(h_disj h_ne.symm).sdiff_eq_left] using h_closed j - -end disjoint_subsets - -theorem frontier_eq_empty_iff [PreconnectedSpace α] {s : Set α} : - frontier s = ∅ ↔ s = ∅ ∨ s = univ := - isClopen_iff_frontier_eq_empty.symm.trans isClopen_iff - -theorem nonempty_frontier_iff [PreconnectedSpace α] {s : Set α} : - (frontier s).Nonempty ↔ s.Nonempty ∧ s ≠ univ := by - simp only [nonempty_iff_ne_empty, Ne, frontier_eq_empty_iff, not_or] - theorem Subtype.preconnectedSpace {s : Set α} (h : IsPreconnected s) : PreconnectedSpace s where isPreconnected_univ := by rwa [← inducing_subtype_val.isPreconnected_image, image_univ, Subtype.range_val] @@ -863,356 +715,4 @@ theorem isConnected_iff_connectedSpace {s : Set α} : IsConnected s ↔ Connecte ⟨Subtype.connectedSpace, fun h => ⟨nonempty_subtype.mp h.2, isPreconnected_iff_preconnectedSpace.mpr h.1⟩⟩ -/-- In a preconnected space, given a transitive relation `P`, if `P x y` and `P y x` are true -for `y` close enough to `x`, then `P x y` holds for all `x, y`. This is a version of the fact -that, if an equivalence relation has open classes, then it has a single equivalence class. -/ -lemma PreconnectedSpace.induction₂' [PreconnectedSpace α] (P : α → α → Prop) - (h : ∀ x, ∀ᶠ y in 𝓝 x, P x y ∧ P y x) (h' : Transitive P) (x y : α) : - P x y := by - let u := {z | P x z} - have A : IsClosed u := by - apply isClosed_iff_nhds.2 (fun z hz ↦ ?_) - rcases hz _ (h z) with ⟨t, ht, h't⟩ - exact h' h't ht.2 - have B : IsOpen u := by - apply isOpen_iff_mem_nhds.2 (fun z hz ↦ ?_) - filter_upwards [h z] with t ht - exact h' hz ht.1 - have C : u.Nonempty := ⟨x, (mem_of_mem_nhds (h x)).1⟩ - have D : u = Set.univ := IsClopen.eq_univ ⟨A, B⟩ C - show y ∈ u - simp [D] - -/-- In a preconnected space, if a symmetric transitive relation `P x y` is true for `y` close -enough to `x`, then it holds for all `x, y`. This is a version of the fact that, if an equivalence -relation has open classes, then it has a single equivalence class. -/ -lemma PreconnectedSpace.induction₂ [PreconnectedSpace α] (P : α → α → Prop) - (h : ∀ x, ∀ᶠ y in 𝓝 x, P x y) (h' : Transitive P) (h'' : Symmetric P) (x y : α) : - P x y := by - refine PreconnectedSpace.induction₂' P (fun z ↦ ?_) h' x y - filter_upwards [h z] with a ha - exact ⟨ha, h'' ha⟩ - -/-- In a preconnected set, given a transitive relation `P`, if `P x y` and `P y x` are true -for `y` close enough to `x`, then `P x y` holds for all `x, y`. This is a version of the fact -that, if an equivalence relation has open classes, then it has a single equivalence class. -/ -lemma IsPreconnected.induction₂' {s : Set α} (hs : IsPreconnected s) (P : α → α → Prop) - (h : ∀ x ∈ s, ∀ᶠ y in 𝓝[s] x, P x y ∧ P y x) - (h' : ∀ x y z, x ∈ s → y ∈ s → z ∈ s → P x y → P y z → P x z) - {x y : α} (hx : x ∈ s) (hy : y ∈ s) : P x y := by - let Q : s → s → Prop := fun a b ↦ P a b - show Q ⟨x, hx⟩ ⟨y, hy⟩ - have : PreconnectedSpace s := Subtype.preconnectedSpace hs - apply PreconnectedSpace.induction₂' - · rintro ⟨x, hx⟩ - have Z := h x hx - rwa [nhdsWithin_eq_map_subtype_coe] at Z - · rintro ⟨a, ha⟩ ⟨b, hb⟩ ⟨c, hc⟩ hab hbc - exact h' a b c ha hb hc hab hbc - -/-- In a preconnected set, if a symmetric transitive relation `P x y` is true for `y` close -enough to `x`, then it holds for all `x, y`. This is a version of the fact that, if an equivalence -relation has open classes, then it has a single equivalence class. -/ -lemma IsPreconnected.induction₂ {s : Set α} (hs : IsPreconnected s) (P : α → α → Prop) - (h : ∀ x ∈ s, ∀ᶠ y in 𝓝[s] x, P x y) - (h' : ∀ x y z, x ∈ s → y ∈ s → z ∈ s → P x y → P y z → P x z) - (h'' : ∀ x y, x ∈ s → y ∈ s → P x y → P y x) - {x y : α} (hx : x ∈ s) (hy : y ∈ s) : P x y := by - apply hs.induction₂' P (fun z hz ↦ ?_) h' hx hy - filter_upwards [h z hz, self_mem_nhdsWithin] with a ha h'a - exact ⟨ha, h'' z a hz h'a ha⟩ - -/-- A set `s` is preconnected if and only if for every cover by two open sets that are disjoint on -`s`, it is contained in one of the two covering sets. -/ -theorem isPreconnected_iff_subset_of_disjoint {s : Set α} : - IsPreconnected s ↔ - ∀ u v, IsOpen u → IsOpen v → s ⊆ u ∪ v → s ∩ (u ∩ v) = ∅ → s ⊆ u ∨ s ⊆ v := by - constructor <;> intro h - · intro u v hu hv hs huv - specialize h u v hu hv hs - contrapose! huv - simp [not_subset] at huv - rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ - have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu - have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv - exact h ⟨y, hys, hyu⟩ ⟨x, hxs, hxv⟩ - · intro u v hu hv hs hsu hsv - by_contra H - specialize h u v hu hv hs (Set.not_nonempty_iff_eq_empty.mp H) - apply H - cases' h with h h - · rcases hsv with ⟨x, hxs, hxv⟩ - exact ⟨x, hxs, ⟨h hxs, hxv⟩⟩ - · rcases hsu with ⟨x, hxs, hxu⟩ - exact ⟨x, hxs, ⟨hxu, h hxs⟩⟩ - -/-- A set `s` is connected if and only if -for every cover by a finite collection of open sets that are pairwise disjoint on `s`, -it is contained in one of the members of the collection. -/ -theorem isConnected_iff_sUnion_disjoint_open {s : Set α} : - IsConnected s ↔ - ∀ U : Finset (Set α), (∀ u v : Set α, u ∈ U → v ∈ U → (s ∩ (u ∩ v)).Nonempty → u = v) → - (∀ u ∈ U, IsOpen u) → (s ⊆ ⋃₀ ↑U) → ∃ u ∈ U, s ⊆ u := by - rw [IsConnected, isPreconnected_iff_subset_of_disjoint] - refine ⟨fun ⟨hne, h⟩ U hU hUo hsU => ?_, fun h => ⟨?_, fun u v hu hv hs hsuv => ?_⟩⟩ - · induction U using Finset.induction_on with - | empty => exact absurd (by simpa using hsU) hne.not_subset_empty - | @insert u U uU IH => - simp only [← forall_cond_comm, Finset.forall_mem_insert, Finset.exists_mem_insert, - Finset.coe_insert, sUnion_insert, implies_true, true_and] at * - refine (h _ hUo.1 (⋃₀ ↑U) (isOpen_sUnion hUo.2) hsU ?_).imp_right ?_ - · refine subset_empty_iff.1 fun x ⟨hxs, hxu, v, hvU, hxv⟩ => ?_ - exact ne_of_mem_of_not_mem hvU uU (hU.1 v hvU ⟨x, hxs, hxu, hxv⟩).symm - · exact IH (fun u hu => (hU.2 u hu).2) hUo.2 - · simpa [subset_empty_iff, nonempty_iff_ne_empty] using h ∅ - · rw [← not_nonempty_iff_eq_empty] at hsuv - have := hsuv; rw [inter_comm u] at this - simpa [*, or_imp, forall_and] using h {u, v} - --- Porting note: `IsPreconnected.subset_isClopen` moved up from here - -/-- Preconnected sets are either contained in or disjoint to any given clopen set. -/ -theorem disjoint_or_subset_of_isClopen {s t : Set α} (hs : IsPreconnected s) (ht : IsClopen t) : - Disjoint s t ∨ s ⊆ t := - (disjoint_or_nonempty_inter s t).imp_right <| hs.subset_isClopen ht - -/-- A set `s` is preconnected if and only if -for every cover by two closed sets that are disjoint on `s`, -it is contained in one of the two covering sets. -/ -theorem isPreconnected_iff_subset_of_disjoint_closed : - IsPreconnected s ↔ - ∀ u v, IsClosed u → IsClosed v → s ⊆ u ∪ v → s ∩ (u ∩ v) = ∅ → s ⊆ u ∨ s ⊆ v := by - constructor <;> intro h - · intro u v hu hv hs huv - rw [isPreconnected_closed_iff] at h - specialize h u v hu hv hs - contrapose! huv - simp [not_subset] at huv - rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ - have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu - have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv - exact h ⟨y, hys, hyu⟩ ⟨x, hxs, hxv⟩ - · rw [isPreconnected_closed_iff] - intro u v hu hv hs hsu hsv - by_contra H - specialize h u v hu hv hs (Set.not_nonempty_iff_eq_empty.mp H) - apply H - cases' h with h h - · rcases hsv with ⟨x, hxs, hxv⟩ - exact ⟨x, hxs, ⟨h hxs, hxv⟩⟩ - · rcases hsu with ⟨x, hxs, hxu⟩ - exact ⟨x, hxs, ⟨hxu, h hxs⟩⟩ - -/-- A closed set `s` is preconnected if and only if for every cover by two closed sets that are -disjoint, it is contained in one of the two covering sets. -/ -theorem isPreconnected_iff_subset_of_fully_disjoint_closed {s : Set α} (hs : IsClosed s) : - IsPreconnected s ↔ - ∀ u v, IsClosed u → IsClosed v → s ⊆ u ∪ v → Disjoint u v → s ⊆ u ∨ s ⊆ v := by - refine isPreconnected_iff_subset_of_disjoint_closed.trans ⟨?_, ?_⟩ <;> intro H u v hu hv hss huv - · apply H u v hu hv hss - rw [huv.inter_eq, inter_empty] - have H1 := H (u ∩ s) (v ∩ s) - rw [subset_inter_iff, subset_inter_iff] at H1 - simp only [Subset.refl, and_true] at H1 - apply H1 (hu.inter hs) (hv.inter hs) - · rw [← union_inter_distrib_right] - exact subset_inter hss Subset.rfl - · rwa [disjoint_iff_inter_eq_empty, ← inter_inter_distrib_right, inter_comm] - -theorem IsClopen.connectedComponent_subset {x} (hs : IsClopen s) (hx : x ∈ s) : - connectedComponent x ⊆ s := - isPreconnected_connectedComponent.subset_isClopen hs ⟨x, mem_connectedComponent, hx⟩ - -/-- The connected component of a point is always a subset of the intersection of all its clopen -neighbourhoods. -/ -theorem connectedComponent_subset_iInter_isClopen {x : α} : - connectedComponent x ⊆ ⋂ Z : { Z : Set α // IsClopen Z ∧ x ∈ Z }, Z := - subset_iInter fun Z => Z.2.1.connectedComponent_subset Z.2.2 - -/-- A clopen set is the union of its connected components. -/ -theorem IsClopen.biUnion_connectedComponent_eq {Z : Set α} (h : IsClopen Z) : - ⋃ x ∈ Z, connectedComponent x = Z := - Subset.antisymm (iUnion₂_subset fun _ => h.connectedComponent_subset) fun _ h => - mem_iUnion₂_of_mem h mem_connectedComponent - -/-- The preimage of a connected component is preconnected if the function has connected fibers -and a subset is closed iff the preimage is. -/ -theorem preimage_connectedComponent_connected [TopologicalSpace β] {f : α → β} - (connected_fibers : ∀ t : β, IsConnected (f ⁻¹' {t})) - (hcl : ∀ T : Set β, IsClosed T ↔ IsClosed (f ⁻¹' T)) (t : β) : - IsConnected (f ⁻¹' connectedComponent t) := by - -- The following proof is essentially https://stacks.math.columbia.edu/tag/0377 - -- although the statement is slightly different - have hf : Surjective f := Surjective.of_comp fun t : β => (connected_fibers t).1 - refine ⟨Nonempty.preimage connectedComponent_nonempty hf, ?_⟩ - have hT : IsClosed (f ⁻¹' connectedComponent t) := - (hcl (connectedComponent t)).1 isClosed_connectedComponent - -- To show it's preconnected we decompose (f ⁻¹' connectedComponent t) as a subset of two - -- closed disjoint sets in α. We want to show that it's a subset of either. - rw [isPreconnected_iff_subset_of_fully_disjoint_closed hT] - intro u v hu hv huv uv_disj - -- To do this we decompose connectedComponent t into T₁ and T₂ - -- we will show that connectedComponent t is a subset of either and hence - -- (f ⁻¹' connectedComponent t) is a subset of u or v - let T₁ := { t' ∈ connectedComponent t | f ⁻¹' {t'} ⊆ u } - let T₂ := { t' ∈ connectedComponent t | f ⁻¹' {t'} ⊆ v } - have fiber_decomp : ∀ t' ∈ connectedComponent t, f ⁻¹' {t'} ⊆ u ∨ f ⁻¹' {t'} ⊆ v := by - intro t' ht' - apply isPreconnected_iff_subset_of_disjoint_closed.1 (connected_fibers t').2 u v hu hv - · exact Subset.trans (preimage_mono (singleton_subset_iff.2 ht')) huv - rw [uv_disj.inter_eq, inter_empty] - have T₁_u : f ⁻¹' T₁ = f ⁻¹' connectedComponent t ∩ u := by - apply eq_of_subset_of_subset - · rw [← biUnion_preimage_singleton] - refine iUnion₂_subset fun t' ht' => subset_inter ?_ ht'.2 - rw [hf.preimage_subset_preimage_iff, singleton_subset_iff] - exact ht'.1 - rintro a ⟨hat, hau⟩ - constructor - · exact mem_preimage.1 hat - refine (fiber_decomp (f a) (mem_preimage.1 hat)).resolve_right fun h => ?_ - exact uv_disj.subset_compl_right hau (h rfl) - -- This proof is exactly the same as the above (modulo some symmetry) - have T₂_v : f ⁻¹' T₂ = f ⁻¹' connectedComponent t ∩ v := by - apply eq_of_subset_of_subset - · rw [← biUnion_preimage_singleton] - refine iUnion₂_subset fun t' ht' => subset_inter ?_ ht'.2 - rw [hf.preimage_subset_preimage_iff, singleton_subset_iff] - exact ht'.1 - rintro a ⟨hat, hav⟩ - constructor - · exact mem_preimage.1 hat - · refine (fiber_decomp (f a) (mem_preimage.1 hat)).resolve_left fun h => ?_ - exact uv_disj.subset_compl_left hav (h rfl) - -- Now we show T₁, T₂ are closed, cover connectedComponent t and are disjoint. - have hT₁ : IsClosed T₁ := (hcl T₁).2 (T₁_u.symm ▸ IsClosed.inter hT hu) - have hT₂ : IsClosed T₂ := (hcl T₂).2 (T₂_v.symm ▸ IsClosed.inter hT hv) - have T_decomp : connectedComponent t ⊆ T₁ ∪ T₂ := fun t' ht' => by - rw [mem_union t' T₁ T₂] - cases' fiber_decomp t' ht' with htu htv - · left - exact ⟨ht', htu⟩ - right - exact ⟨ht', htv⟩ - have T_disjoint : Disjoint T₁ T₂ := by - refine Disjoint.of_preimage hf ?_ - rw [T₁_u, T₂_v, disjoint_iff_inter_eq_empty, ← inter_inter_distrib_left, uv_disj.inter_eq, - inter_empty] - -- Now we do cases on whether (connectedComponent t) is a subset of T₁ or T₂ to show - -- that the preimage is a subset of u or v. - cases' (isPreconnected_iff_subset_of_fully_disjoint_closed isClosed_connectedComponent).1 - isPreconnected_connectedComponent T₁ T₂ hT₁ hT₂ T_decomp T_disjoint with h h - · left - rw [Subset.antisymm_iff] at T₁_u - suffices f ⁻¹' connectedComponent t ⊆ f ⁻¹' T₁ - from (this.trans T₁_u.1).trans inter_subset_right - exact preimage_mono h - · right - rw [Subset.antisymm_iff] at T₂_v - suffices f ⁻¹' connectedComponent t ⊆ f ⁻¹' T₂ - from (this.trans T₂_v.1).trans inter_subset_right - exact preimage_mono h - -theorem QuotientMap.preimage_connectedComponent [TopologicalSpace β] {f : α → β} - (hf : QuotientMap f) (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : - f ⁻¹' connectedComponent (f a) = connectedComponent a := - ((preimage_connectedComponent_connected h_fibers (fun _ => hf.isClosed_preimage.symm) - _).subset_connectedComponent mem_connectedComponent).antisymm - (hf.continuous.mapsTo_connectedComponent a) - -theorem QuotientMap.image_connectedComponent [TopologicalSpace β] {f : α → β} (hf : QuotientMap f) - (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : - f '' connectedComponent a = connectedComponent (f a) := by - rw [← hf.preimage_connectedComponent h_fibers, image_preimage_eq _ hf.surjective] - end Preconnected - -section connectedComponentSetoid -/-- The setoid of connected components of a topological space -/ -def connectedComponentSetoid (α : Type*) [TopologicalSpace α] : Setoid α := - ⟨fun x y => connectedComponent x = connectedComponent y, - ⟨fun x => by trivial, fun h1 => h1.symm, fun h1 h2 => h1.trans h2⟩⟩ - -/-- The quotient of a space by its connected components -/ -def ConnectedComponents (α : Type u) [TopologicalSpace α] := - Quotient (connectedComponentSetoid α) - -namespace ConnectedComponents - -/-- Coercion from a topological space to the set of connected components of this space. -/ -def mk : α → ConnectedComponents α := Quotient.mk'' - -instance : CoeTC α (ConnectedComponents α) := ⟨mk⟩ - -@[simp] -theorem coe_eq_coe {x y : α} : - (x : ConnectedComponents α) = y ↔ connectedComponent x = connectedComponent y := - Quotient.eq'' - -theorem coe_ne_coe {x y : α} : - (x : ConnectedComponents α) ≠ y ↔ connectedComponent x ≠ connectedComponent y := - coe_eq_coe.not - -theorem coe_eq_coe' {x y : α} : (x : ConnectedComponents α) = y ↔ x ∈ connectedComponent y := - coe_eq_coe.trans connectedComponent_eq_iff_mem - -instance [Inhabited α] : Inhabited (ConnectedComponents α) := - ⟨mk default⟩ - -instance : TopologicalSpace (ConnectedComponents α) := - inferInstanceAs (TopologicalSpace (Quotient _)) - -theorem surjective_coe : Surjective (mk : α → ConnectedComponents α) := - surjective_quot_mk _ - -theorem quotientMap_coe : QuotientMap (mk : α → ConnectedComponents α) := - quotientMap_quot_mk - -@[continuity] -theorem continuous_coe : Continuous (mk : α → ConnectedComponents α) := - quotientMap_coe.continuous - -@[simp] -theorem range_coe : range (mk : α → ConnectedComponents α) = univ := - surjective_coe.range_eq - -end ConnectedComponents - -/-- The preimage of a singleton in `connectedComponents` is the connected component -of an element in the equivalence class. -/ -theorem connectedComponents_preimage_singleton {x : α} : - (↑) ⁻¹' ({↑x} : Set (ConnectedComponents α)) = connectedComponent x := by - ext y - rw [mem_preimage, mem_singleton_iff, ConnectedComponents.coe_eq_coe'] - -/-- The preimage of the image of a set under the quotient map to `connectedComponents α` -is the union of the connected components of the elements in it. -/ -theorem connectedComponents_preimage_image (U : Set α) : - (↑) ⁻¹' ((↑) '' U : Set (ConnectedComponents α)) = ⋃ x ∈ U, connectedComponent x := by - simp only [connectedComponents_preimage_singleton, preimage_iUnion₂, image_eq_iUnion] - - - -end connectedComponentSetoid - -/-- If every map to `Bool` (a discrete two-element space), that is -continuous on a set `s`, is constant on s, then s is preconnected -/ -theorem isPreconnected_of_forall_constant {s : Set α} - (hs : ∀ f : α → Bool, ContinuousOn f s → ∀ x ∈ s, ∀ y ∈ s, f x = f y) : IsPreconnected s := by - unfold IsPreconnected - by_contra! - rcases this with ⟨u, v, u_op, v_op, hsuv, ⟨x, x_in_s, x_in_u⟩, ⟨y, y_in_s, y_in_v⟩, H⟩ - have hy : y ∉ u := fun y_in_u => eq_empty_iff_forall_not_mem.mp H y ⟨y_in_s, ⟨y_in_u, y_in_v⟩⟩ - have : ContinuousOn u.boolIndicator s := by - apply (continuousOn_boolIndicator_iff_isClopen _ _).mpr ⟨_, _⟩ - · rw [preimage_subtype_coe_eq_compl hsuv H] - exact (v_op.preimage continuous_subtype_val).isClosed_compl - · exact u_op.preimage continuous_subtype_val - simpa [(u.mem_iff_boolIndicator _).mp x_in_u, (u.not_mem_iff_boolIndicator _).mp hy] using - hs _ this x x_in_s y y_in_s - -/-- A `PreconnectedSpace` version of `isPreconnected_of_forall_constant` -/ -theorem preconnectedSpace_of_forall_constant - (hs : ∀ f : α → Bool, Continuous f → ∀ x y, f x = f y) : PreconnectedSpace α := - ⟨isPreconnected_of_forall_constant fun f hf x _ y _ => - hs f (continuous_iff_continuousOn_univ.mpr hf) x y⟩ diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean new file mode 100644 index 0000000000000..a9fa5d0abcd16 --- /dev/null +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -0,0 +1,531 @@ +/- +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 +-/ +import Mathlib.Topology.Clopen +import Mathlib.Topology.Connected.Basic + +/-! +# Connected subsets and their relation to clopen sets + +In this file we show how connected subsets of a topological space are intimately connected +to clopen sets. + +## Main declarations + ++ `IsClopen.biUnion_connectedComponent_eq`: a clopen set is the union of its connected components. ++ `PreconnectedSpace.induction₂`: an induction principle for preconnected spaces. ++ `ConnectedComponents`: The connected compoenents of a topological space, as a quotient type. + +-/ + +open Set Function Topology TopologicalSpace Relation +open scoped Classical + +universe u v + +variable {α : Type u} {β : Type v} {ι : Type*} {π : ι → Type*} [TopologicalSpace α] + {s t u v : Set α} + +section Preconnected + +/-- Preconnected sets are either contained in or disjoint to any given clopen set. -/ +theorem IsPreconnected.subset_isClopen {s t : Set α} (hs : IsPreconnected s) (ht : IsClopen t) + (hne : (s ∩ t).Nonempty) : s ⊆ t := + hs.subset_left_of_subset_union ht.isOpen ht.compl.isOpen disjoint_compl_right (by simp) hne + +theorem Sigma.isConnected_iff [∀ i, TopologicalSpace (π i)] {s : Set (Σi, π i)} : + IsConnected s ↔ ∃ i t, IsConnected t ∧ s = Sigma.mk i '' t := by + refine ⟨fun hs => ?_, ?_⟩ + · obtain ⟨⟨i, x⟩, hx⟩ := hs.nonempty + have : s ⊆ range (Sigma.mk i) := + hs.isPreconnected.subset_isClopen isClopen_range_sigmaMk ⟨⟨i, x⟩, hx, x, rfl⟩ + exact ⟨i, Sigma.mk i ⁻¹' s, hs.preimage_of_isOpenMap sigma_mk_injective isOpenMap_sigmaMk this, + (Set.image_preimage_eq_of_subset this).symm⟩ + · rintro ⟨i, t, ht, rfl⟩ + exact ht.image _ continuous_sigmaMk.continuousOn + +theorem Sigma.isPreconnected_iff [hι : Nonempty ι] [∀ i, TopologicalSpace (π i)] + {s : Set (Σi, π i)} : IsPreconnected s ↔ ∃ i t, IsPreconnected t ∧ s = Sigma.mk i '' t := by + refine ⟨fun hs => ?_, ?_⟩ + · obtain rfl | h := s.eq_empty_or_nonempty + · exact ⟨Classical.choice hι, ∅, isPreconnected_empty, (Set.image_empty _).symm⟩ + · obtain ⟨a, t, ht, rfl⟩ := Sigma.isConnected_iff.1 ⟨h, hs⟩ + exact ⟨a, t, ht.isPreconnected, rfl⟩ + · rintro ⟨a, t, ht, rfl⟩ + exact ht.image _ continuous_sigmaMk.continuousOn + +theorem Sum.isConnected_iff [TopologicalSpace β] {s : Set (α ⊕ β)} : + IsConnected s ↔ + (∃ t, IsConnected t ∧ s = Sum.inl '' t) ∨ ∃ t, IsConnected t ∧ s = Sum.inr '' t := by + refine ⟨fun hs => ?_, ?_⟩ + · obtain ⟨x | x, hx⟩ := hs.nonempty + · have h : s ⊆ range Sum.inl := + hs.isPreconnected.subset_isClopen isClopen_range_inl ⟨.inl x, hx, x, rfl⟩ + refine Or.inl ⟨Sum.inl ⁻¹' s, ?_, ?_⟩ + · exact hs.preimage_of_isOpenMap Sum.inl_injective isOpenMap_inl h + · exact (image_preimage_eq_of_subset h).symm + · have h : s ⊆ range Sum.inr := + hs.isPreconnected.subset_isClopen isClopen_range_inr ⟨.inr x, hx, x, rfl⟩ + refine Or.inr ⟨Sum.inr ⁻¹' s, ?_, ?_⟩ + · exact hs.preimage_of_isOpenMap Sum.inr_injective isOpenMap_inr h + · exact (image_preimage_eq_of_subset h).symm + · rintro (⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩) + · exact ht.image _ continuous_inl.continuousOn + · exact ht.image _ continuous_inr.continuousOn + +theorem Sum.isPreconnected_iff [TopologicalSpace β] {s : Set (α ⊕ β)} : + IsPreconnected s ↔ + (∃ t, IsPreconnected t ∧ s = Sum.inl '' t) ∨ ∃ t, IsPreconnected t ∧ s = Sum.inr '' t := by + refine ⟨fun hs => ?_, ?_⟩ + · obtain rfl | h := s.eq_empty_or_nonempty + · exact Or.inl ⟨∅, isPreconnected_empty, (Set.image_empty _).symm⟩ + obtain ⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩ := Sum.isConnected_iff.1 ⟨h, hs⟩ + · exact Or.inl ⟨t, ht.isPreconnected, rfl⟩ + · exact Or.inr ⟨t, ht.isPreconnected, rfl⟩ + · rintro (⟨t, ht, rfl⟩ | ⟨t, ht, rfl⟩) + · exact ht.image _ continuous_inl.continuousOn + · exact ht.image _ continuous_inr.continuousOn + +/-- A continuous map from a connected space to a disjoint union `Σ i, π i` can be lifted to one of +the components `π i`. See also `ContinuousMap.exists_lift_sigma` for a version with bundled +`ContinuousMap`s. -/ +theorem Continuous.exists_lift_sigma [ConnectedSpace α] [∀ i, TopologicalSpace (π i)] + {f : α → Σ i, π i} (hf : Continuous f) : + ∃ (i : ι) (g : α → π i), Continuous g ∧ f = Sigma.mk i ∘ g := by + obtain ⟨i, hi⟩ : ∃ i, range f ⊆ range (.mk i) := by + rcases Sigma.isConnected_iff.1 (isConnected_range hf) with ⟨i, s, -, hs⟩ + exact ⟨i, hs.trans_subset (image_subset_range _ _)⟩ + rcases range_subset_range_iff_exists_comp.1 hi with ⟨g, rfl⟩ + refine ⟨i, g, ?_, rfl⟩ + rwa [← embedding_sigmaMk.continuous_iff] at hf + +theorem nonempty_inter [PreconnectedSpace α] {s t : Set α} : + IsOpen s → IsOpen t → s ∪ t = univ → s.Nonempty → t.Nonempty → (s ∩ t).Nonempty := by + simpa only [univ_inter, univ_subset_iff] using @PreconnectedSpace.isPreconnected_univ α _ _ s t + +theorem isClopen_iff [PreconnectedSpace α] {s : Set α} : IsClopen s ↔ s = ∅ ∨ s = univ := + ⟨fun hs => + by_contradiction fun h => + have h1 : s ≠ ∅ ∧ sᶜ ≠ ∅ := + ⟨mt Or.inl h, + mt (fun h2 => Or.inr <| (by rw [← compl_compl s, h2, compl_empty] : s = univ)) h⟩ + let ⟨_, h2, h3⟩ := + nonempty_inter hs.2 hs.1.isOpen_compl (union_compl_self s) (nonempty_iff_ne_empty.2 h1.1) + (nonempty_iff_ne_empty.2 h1.2) + h3 h2, + by rintro (rfl | rfl) <;> [exact isClopen_empty; exact isClopen_univ]⟩ + +theorem IsClopen.eq_univ [PreconnectedSpace α] {s : Set α} (h' : IsClopen s) (h : s.Nonempty) : + s = univ := + (isClopen_iff.mp h').resolve_left h.ne_empty + +section disjoint_subsets + +variable [PreconnectedSpace α] + {s : ι → Set α} (h_nonempty : ∀ i, (s i).Nonempty) (h_disj : Pairwise (Disjoint on s)) + +/-- In a preconnected space, any disjoint family of non-empty clopen subsets has at most one +element. -/ +lemma subsingleton_of_disjoint_isClopen + (h_clopen : ∀ i, IsClopen (s i)) : + Subsingleton ι := by + replace h_nonempty : ∀ i, s i ≠ ∅ := by intro i; rw [← nonempty_iff_ne_empty]; exact h_nonempty i + rw [← not_nontrivial_iff_subsingleton] + by_contra contra + obtain ⟨i, j, h_ne⟩ := contra + replace h_ne : s i ∩ s j = ∅ := by + simpa only [← bot_eq_empty, eq_bot_iff, ← inf_eq_inter, ← disjoint_iff_inf_le] using h_disj h_ne + cases' isClopen_iff.mp (h_clopen i) with hi hi + · exact h_nonempty i hi + · rw [hi, univ_inter] at h_ne + exact h_nonempty j h_ne + +/-- In a preconnected space, any disjoint cover by non-empty open subsets has at most one +element. -/ +lemma subsingleton_of_disjoint_isOpen_iUnion_eq_univ + (h_open : ∀ i, IsOpen (s i)) (h_Union : ⋃ i, s i = univ) : + Subsingleton ι := by + refine subsingleton_of_disjoint_isClopen h_nonempty h_disj (fun i ↦ ⟨?_, h_open i⟩) + rw [← isOpen_compl_iff, compl_eq_univ_diff, ← h_Union, iUnion_diff] + refine isOpen_iUnion (fun j ↦ ?_) + rcases eq_or_ne i j with rfl | h_ne + · simp + · simpa only [(h_disj h_ne.symm).sdiff_eq_left] using h_open j + +/-- In a preconnected space, any finite disjoint cover by non-empty closed subsets has at most one +element. -/ +lemma subsingleton_of_disjoint_isClosed_iUnion_eq_univ [Finite ι] + (h_closed : ∀ i, IsClosed (s i)) (h_Union : ⋃ i, s i = univ) : + Subsingleton ι := by + refine subsingleton_of_disjoint_isClopen h_nonempty h_disj (fun i ↦ ⟨h_closed i, ?_⟩) + rw [← isClosed_compl_iff, compl_eq_univ_diff, ← h_Union, iUnion_diff] + refine isClosed_iUnion_of_finite (fun j ↦ ?_) + rcases eq_or_ne i j with rfl | h_ne + · simp + · simpa only [(h_disj h_ne.symm).sdiff_eq_left] using h_closed j + +end disjoint_subsets + +theorem frontier_eq_empty_iff [PreconnectedSpace α] {s : Set α} : + frontier s = ∅ ↔ s = ∅ ∨ s = univ := + isClopen_iff_frontier_eq_empty.symm.trans isClopen_iff + +theorem nonempty_frontier_iff [PreconnectedSpace α] {s : Set α} : + (frontier s).Nonempty ↔ s.Nonempty ∧ s ≠ univ := by + simp only [nonempty_iff_ne_empty, Ne, frontier_eq_empty_iff, not_or] + +/-- In a preconnected space, given a transitive relation `P`, if `P x y` and `P y x` are true +for `y` close enough to `x`, then `P x y` holds for all `x, y`. This is a version of the fact +that, if an equivalence relation has open classes, then it has a single equivalence class. -/ +lemma PreconnectedSpace.induction₂' [PreconnectedSpace α] (P : α → α → Prop) + (h : ∀ x, ∀ᶠ y in 𝓝 x, P x y ∧ P y x) (h' : Transitive P) (x y : α) : + P x y := by + let u := {z | P x z} + have A : IsClosed u := by + apply isClosed_iff_nhds.2 (fun z hz ↦ ?_) + rcases hz _ (h z) with ⟨t, ht, h't⟩ + exact h' h't ht.2 + have B : IsOpen u := by + apply isOpen_iff_mem_nhds.2 (fun z hz ↦ ?_) + filter_upwards [h z] with t ht + exact h' hz ht.1 + have C : u.Nonempty := ⟨x, (mem_of_mem_nhds (h x)).1⟩ + have D : u = Set.univ := IsClopen.eq_univ ⟨A, B⟩ C + show y ∈ u + simp [D] + +/-- In a preconnected space, if a symmetric transitive relation `P x y` is true for `y` close +enough to `x`, then it holds for all `x, y`. This is a version of the fact that, if an equivalence +relation has open classes, then it has a single equivalence class. -/ +lemma PreconnectedSpace.induction₂ [PreconnectedSpace α] (P : α → α → Prop) + (h : ∀ x, ∀ᶠ y in 𝓝 x, P x y) (h' : Transitive P) (h'' : Symmetric P) (x y : α) : + P x y := by + refine PreconnectedSpace.induction₂' P (fun z ↦ ?_) h' x y + filter_upwards [h z] with a ha + exact ⟨ha, h'' ha⟩ + +/-- In a preconnected set, given a transitive relation `P`, if `P x y` and `P y x` are true +for `y` close enough to `x`, then `P x y` holds for all `x, y`. This is a version of the fact +that, if an equivalence relation has open classes, then it has a single equivalence class. -/ +lemma IsPreconnected.induction₂' {s : Set α} (hs : IsPreconnected s) (P : α → α → Prop) + (h : ∀ x ∈ s, ∀ᶠ y in 𝓝[s] x, P x y ∧ P y x) + (h' : ∀ x y z, x ∈ s → y ∈ s → z ∈ s → P x y → P y z → P x z) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) : P x y := by + let Q : s → s → Prop := fun a b ↦ P a b + show Q ⟨x, hx⟩ ⟨y, hy⟩ + have : PreconnectedSpace s := Subtype.preconnectedSpace hs + apply PreconnectedSpace.induction₂' + · rintro ⟨x, hx⟩ + have Z := h x hx + rwa [nhdsWithin_eq_map_subtype_coe] at Z + · rintro ⟨a, ha⟩ ⟨b, hb⟩ ⟨c, hc⟩ hab hbc + exact h' a b c ha hb hc hab hbc + +/-- In a preconnected set, if a symmetric transitive relation `P x y` is true for `y` close +enough to `x`, then it holds for all `x, y`. This is a version of the fact that, if an equivalence +relation has open classes, then it has a single equivalence class. -/ +lemma IsPreconnected.induction₂ {s : Set α} (hs : IsPreconnected s) (P : α → α → Prop) + (h : ∀ x ∈ s, ∀ᶠ y in 𝓝[s] x, P x y) + (h' : ∀ x y z, x ∈ s → y ∈ s → z ∈ s → P x y → P y z → P x z) + (h'' : ∀ x y, x ∈ s → y ∈ s → P x y → P y x) + {x y : α} (hx : x ∈ s) (hy : y ∈ s) : P x y := by + apply hs.induction₂' P (fun z hz ↦ ?_) h' hx hy + filter_upwards [h z hz, self_mem_nhdsWithin] with a ha h'a + exact ⟨ha, h'' z a hz h'a ha⟩ + +/-- A set `s` is preconnected if and only if for every cover by two open sets that are disjoint on +`s`, it is contained in one of the two covering sets. -/ +theorem isPreconnected_iff_subset_of_disjoint {s : Set α} : + IsPreconnected s ↔ + ∀ u v, IsOpen u → IsOpen v → s ⊆ u ∪ v → s ∩ (u ∩ v) = ∅ → s ⊆ u ∨ s ⊆ v := by + constructor <;> intro h + · intro u v hu hv hs huv + specialize h u v hu hv hs + contrapose! huv + simp [not_subset] at huv + rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ + have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu + have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv + exact h ⟨y, hys, hyu⟩ ⟨x, hxs, hxv⟩ + · intro u v hu hv hs hsu hsv + by_contra H + specialize h u v hu hv hs (Set.not_nonempty_iff_eq_empty.mp H) + apply H + cases' h with h h + · rcases hsv with ⟨x, hxs, hxv⟩ + exact ⟨x, hxs, ⟨h hxs, hxv⟩⟩ + · rcases hsu with ⟨x, hxs, hxu⟩ + exact ⟨x, hxs, ⟨hxu, h hxs⟩⟩ + +/-- A set `s` is connected if and only if +for every cover by a finite collection of open sets that are pairwise disjoint on `s`, +it is contained in one of the members of the collection. -/ +theorem isConnected_iff_sUnion_disjoint_open {s : Set α} : + IsConnected s ↔ + ∀ U : Finset (Set α), (∀ u v : Set α, u ∈ U → v ∈ U → (s ∩ (u ∩ v)).Nonempty → u = v) → + (∀ u ∈ U, IsOpen u) → (s ⊆ ⋃₀ ↑U) → ∃ u ∈ U, s ⊆ u := by + rw [IsConnected, isPreconnected_iff_subset_of_disjoint] + refine ⟨fun ⟨hne, h⟩ U hU hUo hsU => ?_, fun h => ⟨?_, fun u v hu hv hs hsuv => ?_⟩⟩ + · induction U using Finset.induction_on with + | empty => exact absurd (by simpa using hsU) hne.not_subset_empty + | @insert u U uU IH => + simp only [← forall_cond_comm, Finset.forall_mem_insert, Finset.exists_mem_insert, + Finset.coe_insert, sUnion_insert, implies_true, true_and] at * + refine (h _ hUo.1 (⋃₀ ↑U) (isOpen_sUnion hUo.2) hsU ?_).imp_right ?_ + · refine subset_empty_iff.1 fun x ⟨hxs, hxu, v, hvU, hxv⟩ => ?_ + exact ne_of_mem_of_not_mem hvU uU (hU.1 v hvU ⟨x, hxs, hxu, hxv⟩).symm + · exact IH (fun u hu => (hU.2 u hu).2) hUo.2 + · simpa [subset_empty_iff, nonempty_iff_ne_empty] using h ∅ + · rw [← not_nonempty_iff_eq_empty] at hsuv + have := hsuv; rw [inter_comm u] at this + simpa [*, or_imp, forall_and] using h {u, v} + +-- Porting note: `IsPreconnected.subset_isClopen` moved up from here + +/-- Preconnected sets are either contained in or disjoint to any given clopen set. -/ +theorem disjoint_or_subset_of_isClopen {s t : Set α} (hs : IsPreconnected s) (ht : IsClopen t) : + Disjoint s t ∨ s ⊆ t := + (disjoint_or_nonempty_inter s t).imp_right <| hs.subset_isClopen ht + +/-- A set `s` is preconnected if and only if +for every cover by two closed sets that are disjoint on `s`, +it is contained in one of the two covering sets. -/ +theorem isPreconnected_iff_subset_of_disjoint_closed : + IsPreconnected s ↔ + ∀ u v, IsClosed u → IsClosed v → s ⊆ u ∪ v → s ∩ (u ∩ v) = ∅ → s ⊆ u ∨ s ⊆ v := by + constructor <;> intro h + · intro u v hu hv hs huv + rw [isPreconnected_closed_iff] at h + specialize h u v hu hv hs + contrapose! huv + simp [not_subset] at huv + rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ + have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu + have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv + exact h ⟨y, hys, hyu⟩ ⟨x, hxs, hxv⟩ + · rw [isPreconnected_closed_iff] + intro u v hu hv hs hsu hsv + by_contra H + specialize h u v hu hv hs (Set.not_nonempty_iff_eq_empty.mp H) + apply H + cases' h with h h + · rcases hsv with ⟨x, hxs, hxv⟩ + exact ⟨x, hxs, ⟨h hxs, hxv⟩⟩ + · rcases hsu with ⟨x, hxs, hxu⟩ + exact ⟨x, hxs, ⟨hxu, h hxs⟩⟩ + +/-- A closed set `s` is preconnected if and only if for every cover by two closed sets that are +disjoint, it is contained in one of the two covering sets. -/ +theorem isPreconnected_iff_subset_of_fully_disjoint_closed {s : Set α} (hs : IsClosed s) : + IsPreconnected s ↔ + ∀ u v, IsClosed u → IsClosed v → s ⊆ u ∪ v → Disjoint u v → s ⊆ u ∨ s ⊆ v := by + refine isPreconnected_iff_subset_of_disjoint_closed.trans ⟨?_, ?_⟩ <;> intro H u v hu hv hss huv + · apply H u v hu hv hss + rw [huv.inter_eq, inter_empty] + have H1 := H (u ∩ s) (v ∩ s) + rw [subset_inter_iff, subset_inter_iff] at H1 + simp only [Subset.refl, and_true] at H1 + apply H1 (hu.inter hs) (hv.inter hs) + · rw [← union_inter_distrib_right] + exact subset_inter hss Subset.rfl + · rwa [disjoint_iff_inter_eq_empty, ← inter_inter_distrib_right, inter_comm] + +theorem IsClopen.connectedComponent_subset {x} (hs : IsClopen s) (hx : x ∈ s) : + connectedComponent x ⊆ s := + isPreconnected_connectedComponent.subset_isClopen hs ⟨x, mem_connectedComponent, hx⟩ + +/-- The connected component of a point is always a subset of the intersection of all its clopen +neighbourhoods. -/ +theorem connectedComponent_subset_iInter_isClopen {x : α} : + connectedComponent x ⊆ ⋂ Z : { Z : Set α // IsClopen Z ∧ x ∈ Z }, Z := + subset_iInter fun Z => Z.2.1.connectedComponent_subset Z.2.2 + +/-- A clopen set is the union of its connected components. -/ +theorem IsClopen.biUnion_connectedComponent_eq {Z : Set α} (h : IsClopen Z) : + ⋃ x ∈ Z, connectedComponent x = Z := + Subset.antisymm (iUnion₂_subset fun _ => h.connectedComponent_subset) fun _ h => + mem_iUnion₂_of_mem h mem_connectedComponent + +/-- The preimage of a connected component is preconnected if the function has connected fibers +and a subset is closed iff the preimage is. -/ +theorem preimage_connectedComponent_connected [TopologicalSpace β] {f : α → β} + (connected_fibers : ∀ t : β, IsConnected (f ⁻¹' {t})) + (hcl : ∀ T : Set β, IsClosed T ↔ IsClosed (f ⁻¹' T)) (t : β) : + IsConnected (f ⁻¹' connectedComponent t) := by + -- The following proof is essentially https://stacks.math.columbia.edu/tag/0377 + -- although the statement is slightly different + have hf : Surjective f := Surjective.of_comp fun t : β => (connected_fibers t).1 + refine ⟨Nonempty.preimage connectedComponent_nonempty hf, ?_⟩ + have hT : IsClosed (f ⁻¹' connectedComponent t) := + (hcl (connectedComponent t)).1 isClosed_connectedComponent + -- To show it's preconnected we decompose (f ⁻¹' connectedComponent t) as a subset of two + -- closed disjoint sets in α. We want to show that it's a subset of either. + rw [isPreconnected_iff_subset_of_fully_disjoint_closed hT] + intro u v hu hv huv uv_disj + -- To do this we decompose connectedComponent t into T₁ and T₂ + -- we will show that connectedComponent t is a subset of either and hence + -- (f ⁻¹' connectedComponent t) is a subset of u or v + let T₁ := { t' ∈ connectedComponent t | f ⁻¹' {t'} ⊆ u } + let T₂ := { t' ∈ connectedComponent t | f ⁻¹' {t'} ⊆ v } + have fiber_decomp : ∀ t' ∈ connectedComponent t, f ⁻¹' {t'} ⊆ u ∨ f ⁻¹' {t'} ⊆ v := by + intro t' ht' + apply isPreconnected_iff_subset_of_disjoint_closed.1 (connected_fibers t').2 u v hu hv + · exact Subset.trans (preimage_mono (singleton_subset_iff.2 ht')) huv + rw [uv_disj.inter_eq, inter_empty] + have T₁_u : f ⁻¹' T₁ = f ⁻¹' connectedComponent t ∩ u := by + apply eq_of_subset_of_subset + · rw [← biUnion_preimage_singleton] + refine iUnion₂_subset fun t' ht' => subset_inter ?_ ht'.2 + rw [hf.preimage_subset_preimage_iff, singleton_subset_iff] + exact ht'.1 + rintro a ⟨hat, hau⟩ + constructor + · exact mem_preimage.1 hat + refine (fiber_decomp (f a) (mem_preimage.1 hat)).resolve_right fun h => ?_ + exact uv_disj.subset_compl_right hau (h rfl) + -- This proof is exactly the same as the above (modulo some symmetry) + have T₂_v : f ⁻¹' T₂ = f ⁻¹' connectedComponent t ∩ v := by + apply eq_of_subset_of_subset + · rw [← biUnion_preimage_singleton] + refine iUnion₂_subset fun t' ht' => subset_inter ?_ ht'.2 + rw [hf.preimage_subset_preimage_iff, singleton_subset_iff] + exact ht'.1 + rintro a ⟨hat, hav⟩ + constructor + · exact mem_preimage.1 hat + · refine (fiber_decomp (f a) (mem_preimage.1 hat)).resolve_left fun h => ?_ + exact uv_disj.subset_compl_left hav (h rfl) + -- Now we show T₁, T₂ are closed, cover connectedComponent t and are disjoint. + have hT₁ : IsClosed T₁ := (hcl T₁).2 (T₁_u.symm ▸ IsClosed.inter hT hu) + have hT₂ : IsClosed T₂ := (hcl T₂).2 (T₂_v.symm ▸ IsClosed.inter hT hv) + have T_decomp : connectedComponent t ⊆ T₁ ∪ T₂ := fun t' ht' => by + rw [mem_union t' T₁ T₂] + cases' fiber_decomp t' ht' with htu htv + · left + exact ⟨ht', htu⟩ + right + exact ⟨ht', htv⟩ + have T_disjoint : Disjoint T₁ T₂ := by + refine Disjoint.of_preimage hf ?_ + rw [T₁_u, T₂_v, disjoint_iff_inter_eq_empty, ← inter_inter_distrib_left, uv_disj.inter_eq, + inter_empty] + -- Now we do cases on whether (connectedComponent t) is a subset of T₁ or T₂ to show + -- that the preimage is a subset of u or v. + cases' (isPreconnected_iff_subset_of_fully_disjoint_closed isClosed_connectedComponent).1 + isPreconnected_connectedComponent T₁ T₂ hT₁ hT₂ T_decomp T_disjoint with h h + · left + rw [Subset.antisymm_iff] at T₁_u + suffices f ⁻¹' connectedComponent t ⊆ f ⁻¹' T₁ + from (this.trans T₁_u.1).trans inter_subset_right + exact preimage_mono h + · right + rw [Subset.antisymm_iff] at T₂_v + suffices f ⁻¹' connectedComponent t ⊆ f ⁻¹' T₂ + from (this.trans T₂_v.1).trans inter_subset_right + exact preimage_mono h + +theorem QuotientMap.preimage_connectedComponent [TopologicalSpace β] {f : α → β} + (hf : QuotientMap f) (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : + f ⁻¹' connectedComponent (f a) = connectedComponent a := + ((preimage_connectedComponent_connected h_fibers (fun _ => hf.isClosed_preimage.symm) + _).subset_connectedComponent mem_connectedComponent).antisymm + (hf.continuous.mapsTo_connectedComponent a) + +theorem QuotientMap.image_connectedComponent [TopologicalSpace β] {f : α → β} (hf : QuotientMap f) + (h_fibers : ∀ y : β, IsConnected (f ⁻¹' {y})) (a : α) : + f '' connectedComponent a = connectedComponent (f a) := by + rw [← hf.preimage_connectedComponent h_fibers, image_preimage_eq _ hf.surjective] + +end Preconnected + +section connectedComponentSetoid +/-- The setoid of connected components of a topological space -/ +def connectedComponentSetoid (α : Type*) [TopologicalSpace α] : Setoid α := + ⟨fun x y => connectedComponent x = connectedComponent y, + ⟨fun x => by trivial, fun h1 => h1.symm, fun h1 h2 => h1.trans h2⟩⟩ + +/-- The quotient of a space by its connected components -/ +def ConnectedComponents (α : Type u) [TopologicalSpace α] := + Quotient (connectedComponentSetoid α) + +namespace ConnectedComponents + +/-- Coercion from a topological space to the set of connected components of this space. -/ +def mk : α → ConnectedComponents α := Quotient.mk'' + +instance : CoeTC α (ConnectedComponents α) := ⟨mk⟩ + +@[simp] +theorem coe_eq_coe {x y : α} : + (x : ConnectedComponents α) = y ↔ connectedComponent x = connectedComponent y := + Quotient.eq'' + +theorem coe_ne_coe {x y : α} : + (x : ConnectedComponents α) ≠ y ↔ connectedComponent x ≠ connectedComponent y := + coe_eq_coe.not + +theorem coe_eq_coe' {x y : α} : (x : ConnectedComponents α) = y ↔ x ∈ connectedComponent y := + coe_eq_coe.trans connectedComponent_eq_iff_mem + +instance [Inhabited α] : Inhabited (ConnectedComponents α) := + ⟨mk default⟩ + +instance : TopologicalSpace (ConnectedComponents α) := + inferInstanceAs (TopologicalSpace (Quotient _)) + +theorem surjective_coe : Surjective (mk : α → ConnectedComponents α) := + surjective_quot_mk _ + +theorem quotientMap_coe : QuotientMap (mk : α → ConnectedComponents α) := + quotientMap_quot_mk + +@[continuity] +theorem continuous_coe : Continuous (mk : α → ConnectedComponents α) := + quotientMap_coe.continuous + +@[simp] +theorem range_coe : range (mk : α → ConnectedComponents α) = univ := + surjective_coe.range_eq + +end ConnectedComponents + +/-- The preimage of a singleton in `connectedComponents` is the connected component +of an element in the equivalence class. -/ +theorem connectedComponents_preimage_singleton {x : α} : + (↑) ⁻¹' ({↑x} : Set (ConnectedComponents α)) = connectedComponent x := by + ext y + rw [mem_preimage, mem_singleton_iff, ConnectedComponents.coe_eq_coe'] + +/-- The preimage of the image of a set under the quotient map to `connectedComponents α` +is the union of the connected components of the elements in it. -/ +theorem connectedComponents_preimage_image (U : Set α) : + (↑) ⁻¹' ((↑) '' U : Set (ConnectedComponents α)) = ⋃ x ∈ U, connectedComponent x := by + simp only [connectedComponents_preimage_singleton, preimage_iUnion₂, image_eq_iUnion] + + + +end connectedComponentSetoid + +/-- If every map to `Bool` (a discrete two-element space), that is +continuous on a set `s`, is constant on s, then s is preconnected -/ +theorem isPreconnected_of_forall_constant {s : Set α} + (hs : ∀ f : α → Bool, ContinuousOn f s → ∀ x ∈ s, ∀ y ∈ s, f x = f y) : IsPreconnected s := by + unfold IsPreconnected + by_contra! + rcases this with ⟨u, v, u_op, v_op, hsuv, ⟨x, x_in_s, x_in_u⟩, ⟨y, y_in_s, y_in_v⟩, H⟩ + have hy : y ∉ u := fun y_in_u => eq_empty_iff_forall_not_mem.mp H y ⟨y_in_s, ⟨y_in_u, y_in_v⟩⟩ + have : ContinuousOn u.boolIndicator s := by + apply (continuousOn_boolIndicator_iff_isClopen _ _).mpr ⟨_, _⟩ + · rw [preimage_subtype_coe_eq_compl hsuv H] + exact (v_op.preimage continuous_subtype_val).isClosed_compl + · exact u_op.preimage continuous_subtype_val + simpa [(u.mem_iff_boolIndicator _).mp x_in_u, (u.not_mem_iff_boolIndicator _).mp hy] using + hs _ this x x_in_s y y_in_s + +/-- A `PreconnectedSpace` version of `isPreconnected_of_forall_constant` -/ +theorem preconnectedSpace_of_forall_constant + (hs : ∀ f : α → Bool, Continuous f → ∀ x y, f x = f y) : PreconnectedSpace α := + ⟨isPreconnected_of_forall_constant fun f hf x _ y _ => + hs f (continuous_iff_continuousOn_univ.mpr hf) x y⟩ diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index 3db96b74d3686..b01fb5f23d02b 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Patrick Massot, Yury Kudryashov -/ -import Mathlib.Topology.Connected.Basic +import Mathlib.Topology.Connected.Clopen /-! # Totally disconnected and totally separated topological spaces diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index cf4a1d0321afc..077c62af81d06 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -971,6 +971,10 @@ nonrec theorem IsClosed.closedEmbedding_subtype_val {s : Set X} (hs : IsClosed s ClosedEmbedding ((↑) : s → X) := closedEmbedding_subtype_val hs +theorem IsClosed.isClosedMap_subtype_val {s : Set X} (hs : IsClosed s) : + IsClosedMap ((↑) : s → X) := + hs.closedEmbedding_subtype_val.isClosedMap + @[continuity, fun_prop] theorem Continuous.subtype_mk {f : Y → X} (h : Continuous f) (hp : ∀ x, p (f x)) : Continuous fun x => (⟨f x, hp x⟩ : Subtype p) := diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index 8ad59a0a99a1a..2a4324e641da9 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -7,8 +7,8 @@ import Mathlib.Algebra.Module.MinimalAxioms import Mathlib.Topology.ContinuousFunction.Algebra import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic -import Mathlib.Analysis.NormedSpace.Star.Basic -import Mathlib.Analysis.NormedSpace.ContinuousLinearMap +import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Topology.Bornology.BoundedOperation /-! @@ -1306,7 +1306,7 @@ instance instModule' : Module (α →ᵇ 𝕜) (α →ᵇ β) := (fun _ _ _ => ext fun _ => mul_smul _ _ _) (fun f => ext fun x => one_smul 𝕜 (f x)) -/- TODO: When `NormedModule` has been added to `Analysis.NormedSpace.Basic`, this +/- TODO: When `NormedModule` has been added to `Analysis.Normed.Module.Basic`, this shows that the space of bounded continuous functions from `α` to `β` is naturally a normed module over the algebra of bounded continuous functions from `α` to `𝕜`. -/ instance : BoundedSMul (α →ᵇ 𝕜) (α →ᵇ β) := diff --git a/Mathlib/Topology/ContinuousFunction/Ideals.lean b/Mathlib/Topology/ContinuousFunction/Ideals.lean index 482ca5755a463..93810a823e816 100644 --- a/Mathlib/Topology/ContinuousFunction/Ideals.lean +++ b/Mathlib/Topology/ContinuousFunction/Ideals.lean @@ -7,7 +7,7 @@ import Mathlib.Topology.Algebra.Algebra import Mathlib.Topology.ContinuousFunction.Compact import Mathlib.Topology.UrysohnsLemma import Mathlib.Analysis.RCLike.Basic -import Mathlib.Analysis.NormedSpace.Units +import Mathlib.Analysis.Normed.Ring.Units import Mathlib.Topology.Algebra.Module.CharacterSpace /-! diff --git a/Mathlib/Topology/ContinuousFunction/Polynomial.lean b/Mathlib/Topology/ContinuousFunction/Polynomial.lean index 37438159697c8..3fd434d44adc3 100644 --- a/Mathlib/Topology/ContinuousFunction/Polynomial.lean +++ b/Mathlib/Topology/ContinuousFunction/Polynomial.lean @@ -220,7 +220,7 @@ theorem polynomialFunctions.eq_adjoin_X (s : Set R) : theorem polynomialFunctions.le_equalizer {A : Type*} [Semiring A] [Algebra R A] (s : Set R) (φ ψ : C(s, R) →ₐ[R] A) (h : φ (toContinuousMapOnAlgHom s X) = ψ (toContinuousMapOnAlgHom s X)) : - polynomialFunctions s ≤ φ.equalizer ψ := by + polynomialFunctions s ≤ AlgHom.equalizer φ ψ := by rw [polynomialFunctions.eq_adjoin_X s] exact φ.adjoin_le_equalizer ψ fun x hx => (Set.mem_singleton_iff.1 hx).symm ▸ h diff --git a/Mathlib/Topology/ContinuousFunction/Units.lean b/Mathlib/Topology/ContinuousFunction/Units.lean index 11689830548af..a8c880cf3a2bc 100644 --- a/Mathlib/Topology/ContinuousFunction/Units.lean +++ b/Mathlib/Topology/ContinuousFunction/Units.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Units +import Mathlib.Analysis.Normed.Ring.Units import Mathlib.Algebra.Algebra.Spectrum import Mathlib.Topology.ContinuousFunction.Algebra diff --git a/Mathlib/Topology/Defs/Sequences.lean b/Mathlib/Topology/Defs/Sequences.lean index 1ea8d58459d3b..d553dc937665e 100644 --- a/Mathlib/Topology/Defs/Sequences.lean +++ b/Mathlib/Topology/Defs/Sequences.lean @@ -72,9 +72,11 @@ variable (X) converging subsequence. -/ @[mk_iff] class SeqCompactSpace : Prop where - seq_compact_univ : IsSeqCompact (univ : Set X) + isSeqCompact_univ : IsSeqCompact (univ : Set X) -export SeqCompactSpace (seq_compact_univ) +export SeqCompactSpace (isSeqCompact_univ) + +@[deprecated (since := "2024-07-25")] alias seq_compact_univ := isSeqCompact_univ /-- A topological space is called a *Fréchet-Urysohn space*, if the sequential closure of any set is equal to its closure. Since one of the inclusions is trivial, we require only the non-trivial one diff --git a/Mathlib/Topology/DerivedSet.lean b/Mathlib/Topology/DerivedSet.lean new file mode 100644 index 0000000000000..cf3d3077f668b --- /dev/null +++ b/Mathlib/Topology/DerivedSet.lean @@ -0,0 +1,108 @@ +/- +Copyright (c) 2024 Daniel Weber. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Daniel Weber +-/ +import Mathlib.Topology.Perfect +import Mathlib.Tactic.Peel + +/-! +# Derived set + +This file defines the derived set of a set, the set of all `AccPt`s of its principal filter, +and proves some properties of it. + +-/ + +open Filter Topology + +variable {X : Type*} [TopologicalSpace X] + +theorem AccPt.map {β : Type*} [TopologicalSpace β] {F : Filter X} {x : X} + (h : AccPt x F) {f : X → β} (hf1 : ContinuousAt f x) (hf2 : Function.Injective f) : + AccPt (f x) (map f F) := by + apply map_neBot (m := f) (hf := h) |>.mono + rw [Filter.map_inf hf2] + gcongr + apply tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ hf1.continuousWithinAt + simpa [hf2.eq_iff] using eventually_mem_nhdsWithin + +/-- +The derived set of a set is the set of all accumulation points of it. +-/ +def derivedSet (A : Set X) : Set X := {x | AccPt x (𝓟 A)} + +@[simp] +lemma mem_derivedSet {A : Set X} {x : X} : x ∈ derivedSet A ↔ AccPt x (𝓟 A) := Iff.rfl + +lemma derivedSet_union (A B : Set X) : derivedSet (A ∪ B) = derivedSet A ∪ derivedSet B := by + ext x + simp [derivedSet, ← sup_principal, accPt_sup] + +lemma derivedSet_mono (A B : Set X) (h : A ⊆ B) : derivedSet A ⊆ derivedSet B := + fun _ hx ↦ hx.mono <| le_principal_iff.mpr <| mem_principal.mpr h + +theorem Continuous.image_derivedSet {β : Type*} [TopologicalSpace β] {A : Set X} {f : X → β} + (hf1 : Continuous f) (hf2 : Function.Injective f) : + f '' derivedSet A ⊆ derivedSet (f '' A) := by + intro x hx + simp only [Set.mem_image, mem_derivedSet] at hx + obtain ⟨y, hy1, rfl⟩ := hx + convert hy1.map hf1.continuousAt hf2 + simp + +lemma derivedSet_subset_closure (A : Set X) : derivedSet A ⊆ closure A := + fun _ hx ↦ mem_closure_iff_clusterPt.mpr hx.clusterPt + +lemma isClosed_iff_derivedSet_subset (A : Set X) : IsClosed A ↔ derivedSet A ⊆ A where + mp h := derivedSet_subset_closure A |>.trans h.closure_subset + mpr h := by + rw [isClosed_iff_clusterPt] + intro a ha + by_contra! nh + have : A = A \ {a} := by simp [nh] + rw [this, ← acc_principal_iff_cluster] at ha + exact nh (h ha) + +/-- In a `T1Space`, the `derivedSet` of the closure of a set is equal to the derived set of the +set itself. + +Note: this doesn't hold in a space with the indiscrete topology. For example, if `X` is a type with +two elements, `x` and `y`, and `A := {x}`, then `closure A = Set.univ` and `derivedSet A = {y}`, +but `derivedSet Set.univ = Set.univ`. -/ +lemma derivedSet_closure [T1Space X] (A : Set X) : derivedSet (closure A) = derivedSet A := by + refine le_antisymm (fun x hx => ?_) (derivedSet_mono _ _ subset_closure) + rw [mem_derivedSet, AccPt, (nhdsWithin_basis_open x {x}ᶜ).inf_principal_neBot_iff] at hx ⊢ + peel hx with u hu _ + obtain ⟨-, hu_open⟩ := hu + exact mem_closure_iff.mp this.some_mem.2 (u ∩ {x}ᶜ) (hu_open.inter isOpen_compl_singleton) + this.some_mem.1 + +@[simp] +lemma isClosed_derivedSet [T1Space X] (A : Set X) : IsClosed (derivedSet A) := by + rw [← derivedSet_closure, isClosed_iff_derivedSet_subset] + apply derivedSet_mono + simp [← isClosed_iff_derivedSet_subset] + +lemma preperfect_iff_subset_derivedSet {U : Set X} : Preperfect U ↔ U ⊆ derivedSet U := + Iff.rfl + +lemma perfect_iff_eq_derivedSet {U : Set X} : Perfect U ↔ U = derivedSet U := by + rw [perfect_def, isClosed_iff_derivedSet_subset, preperfect_iff_subset_derivedSet, + ← subset_antisymm_iff, eq_comm] + +lemma IsPreconnected.inter_derivedSet_nonempty [T1Space X] {U : Set X} (hs : IsPreconnected U) + (a b : Set X) (h : U ⊆ a ∪ b) (ha : (U ∩ derivedSet a).Nonempty) + (hb : (U ∩ derivedSet b).Nonempty) : (U ∩ (derivedSet a ∩ derivedSet b)).Nonempty := by + by_cases hu : U.Nontrivial + · apply isPreconnected_closed_iff.mp hs + · simp + · simp + · trans derivedSet U + · apply hs.preperfect_of_nontrivial hu + · rw [← derivedSet_union] + exact derivedSet_mono _ _ h + · exact ha + · exact hb + · obtain ⟨x, hx⟩ := ha.left.exists_eq_singleton_or_nontrivial.resolve_right hu + simp_all diff --git a/Mathlib/Topology/Germ.lean b/Mathlib/Topology/Germ.lean index 0ff126f3704ac..ea5f439b093ec 100644 --- a/Mathlib/Topology/Germ.lean +++ b/Mathlib/Topology/Germ.lean @@ -6,7 +6,7 @@ Authors: Patrick Massot import Mathlib.Order.Filter.Germ.Basic import Mathlib.Topology.NhdsSet import Mathlib.Topology.LocallyConstant.Basic -import Mathlib.Analysis.NormedSpace.Basic +import Mathlib.Analysis.Normed.Module.Basic /-! # Germs of functions between topological spaces diff --git a/Mathlib/Topology/Homotopy/Basic.lean b/Mathlib/Topology/Homotopy/Basic.lean index 790f6a59a70c9..75273bcf06755 100644 --- a/Mathlib/Topology/Homotopy/Basic.lean +++ b/Mathlib/Topology/Homotopy/Basic.lean @@ -169,10 +169,10 @@ theorem extend_apply_of_mem_I (F : Homotopy f₀ f₁) {t : ℝ} (ht : t ∈ I) F.extend t x = F (⟨t, ht⟩, x) := ContinuousMap.congr_fun (Set.IccExtend_of_mem (zero_le_one' ℝ) F.curry ht) x -theorem congr_fun {F G : Homotopy f₀ f₁} (h : F = G) (x : I × X) : F x = G x := +protected theorem congr_fun {F G : Homotopy f₀ f₁} (h : F = G) (x : I × X) : F x = G x := ContinuousMap.congr_fun (congr_arg _ h) x -theorem congr_arg (F : Homotopy f₀ f₁) {x y : I × X} (h : x = y) : F x = F y := +protected theorem congr_arg (F : Homotopy f₀ f₁) {x y : I × X} (h : x = y) : F x = F y := F.toContinuousMap.congr_arg h end diff --git a/Mathlib/Topology/Homotopy/Product.lean b/Mathlib/Topology/Homotopy/Product.lean index eaf98e0c0a6a5..a898bde525861 100644 --- a/Mathlib/Topology/Homotopy/Product.lean +++ b/Mathlib/Topology/Homotopy/Product.lean @@ -130,11 +130,9 @@ theorem pi_lift (γ : ∀ i, Path (as i) (bs i)) : /-- Composition and products commute. This is `Path.trans_pi_eq_pi_trans` descended to path homotopy classes. -/ theorem comp_pi_eq_pi_comp (γ₀ : ∀ i, Path.Homotopic.Quotient (as i) (bs i)) - (γ₁ : ∀ i, Path.Homotopic.Quotient (bs i) (cs i)) : pi γ₀ ⬝ pi γ₁ = pi fun i => γ₀ i ⬝ γ₁ i := by - apply Quotient.induction_on_pi (p := _) γ₁ - intro a - apply Quotient.induction_on_pi (p := _) γ₀ - intros + (γ₁ : ∀ i, Path.Homotopic.Quotient (bs i) (cs i)) : pi γ₀ ⬝ pi γ₁ = pi fun i ↦ γ₀ i ⬝ γ₁ i := by + induction γ₁ using Quotient.induction_on_pi with | _ a => + induction γ₀ using Quotient.induction_on_pi simp only [pi_lift] rw [← Path.Homotopic.comp_lift, Path.trans_pi_eq_pi_trans, ← pi_lift] rfl @@ -147,16 +145,14 @@ abbrev proj (i : ι) (p : Path.Homotopic.Quotient as bs) : Path.Homotopic.Quotie @[simp] theorem proj_pi (i : ι) (paths : ∀ i, Path.Homotopic.Quotient (as i) (bs i)) : proj i (pi paths) = paths i := by - apply Quotient.induction_on_pi (p := _) paths - intro; unfold proj - rw [pi_lift, ← Path.Homotopic.map_lift] + induction paths using Quotient.induction_on_pi + rw [proj, pi_lift, ← Path.Homotopic.map_lift] congr @[simp] theorem pi_proj (p : Path.Homotopic.Quotient as bs) : (pi fun i => proj i p) = p := by - apply Quotient.inductionOn (motive := _) p - intro; unfold proj - simp_rw [← Path.Homotopic.map_lift] + induction p using Quotient.inductionOn + simp_rw [proj, ← Path.Homotopic.map_lift] erw [pi_lift] congr @@ -189,10 +185,8 @@ variable (r₁ : Path.Homotopic.Quotient a₂ a₃) (r₂ : Path.Homotopic.Quoti /-- Products commute with path composition. This is `trans_prod_eq_prod_trans` descended to the quotient. -/ theorem comp_prod_eq_prod_comp : prod q₁ q₂ ⬝ prod r₁ r₂ = prod (q₁ ⬝ r₁) (q₂ ⬝ r₂) := by - apply Quotient.inductionOn₂ (motive := _) q₁ q₂ - intro a b - apply Quotient.inductionOn₂ (motive := _) r₁ r₂ - intros + induction q₁, q₂ using Quotient.inductionOn₂ + induction r₁, r₂ using Quotient.inductionOn₂ simp only [prod_lift, ← Path.Homotopic.comp_lift, Path.trans_prod_eq_prod_trans] variable {c₁ c₂ : α × β} @@ -208,27 +202,21 @@ abbrev projRight (p : Path.Homotopic.Quotient c₁ c₂) : Path.Homotopic.Quotie /-- Lemmas showing projection is the inverse of product. -/ @[simp] theorem projLeft_prod : projLeft (prod q₁ q₂) = q₁ := by - apply Quotient.inductionOn₂ (motive := _) q₁ q₂ - intro p₁ p₂ - unfold projLeft - rw [prod_lift, ← Path.Homotopic.map_lift] + induction q₁, q₂ using Quotient.inductionOn₂ + rw [projLeft, prod_lift, ← Path.Homotopic.map_lift] congr @[simp] theorem projRight_prod : projRight (prod q₁ q₂) = q₂ := by - apply Quotient.inductionOn₂ (motive := _) q₁ q₂ - intro p₁ p₂ - unfold projRight - rw [prod_lift, ← Path.Homotopic.map_lift] + induction q₁, q₂ using Quotient.inductionOn₂ + rw [projRight, prod_lift, ← Path.Homotopic.map_lift] congr @[simp] theorem prod_projLeft_projRight (p : Path.Homotopic.Quotient (a₁, b₁) (a₂, b₂)) : prod (projLeft p) (projRight p) = p := by - apply Quotient.inductionOn (motive := _) p - intro p' - unfold projLeft; unfold projRight - simp only [← Path.Homotopic.map_lift, prod_lift] + induction p using Quotient.inductionOn + simp only [projLeft, projRight, ← Path.Homotopic.map_lift, prod_lift] congr end Prod diff --git a/Mathlib/Topology/LocallyClosed.lean b/Mathlib/Topology/LocallyClosed.lean index 341f1b8fccf4d..f5502f74a4239 100644 --- a/Mathlib/Topology/LocallyClosed.lean +++ b/Mathlib/Topology/LocallyClosed.lean @@ -24,7 +24,7 @@ import Mathlib.Topology.LocalAtTarget -/ -variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {s t : Set α} {f : α → β} +variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {s t : Set X} {f : X → Y} open scoped Topology Set.Notation @@ -39,7 +39,7 @@ This is the largest set such that `s` is closed in, and `s` is locally closed if This is unnamed in the literature, and this name is due to the fact that `coborder s = (border sᶜ)ᶜ` where `border s = s \ interior s` is the border in the sense of Hausdorff. -/ -def coborder (s : Set α) : Set α := +def coborder (s : Set X) : Set X := (closure s \ s)ᶜ lemma subset_coborder : @@ -74,24 +74,24 @@ lemma coborder_eq_compl_frontier_iff : alias ⟨_, IsOpen.coborder_eq⟩ := coborder_eq_compl_frontier_iff -lemma IsOpenMap.coborder_preimage_subset (hf : IsOpenMap f) (s : Set β) : +lemma IsOpenMap.coborder_preimage_subset (hf : IsOpenMap f) (s : Set Y) : coborder (f ⁻¹' s) ⊆ f ⁻¹' (coborder s) := by rw [coborder, coborder, preimage_compl, preimage_diff, compl_subset_compl] apply diff_subset_diff_left exact hf.preimage_closure_subset_closure_preimage -lemma Continuous.preimage_coborder_subset (hf : Continuous f) (s : Set β) : +lemma Continuous.preimage_coborder_subset (hf : Continuous f) (s : Set Y) : f ⁻¹' (coborder s) ⊆ coborder (f ⁻¹' s) := by rw [coborder, coborder, preimage_compl, preimage_diff, compl_subset_compl] apply diff_subset_diff_left exact hf.closure_preimage_subset s -lemma coborder_preimage (hf : IsOpenMap f) (hf' : Continuous f) (s : Set β) : +lemma coborder_preimage (hf : IsOpenMap f) (hf' : Continuous f) (s : Set Y) : coborder (f ⁻¹' s) = f ⁻¹' (coborder s) := (hf.coborder_preimage_subset s).antisymm (hf'.preimage_coborder_subset s) protected -lemma OpenEmbedding.coborder_preimage (hf : OpenEmbedding f) (s : Set β) : +lemma OpenEmbedding.coborder_preimage (hf : OpenEmbedding f) (s : Set Y) : coborder (f ⁻¹' s) = f ⁻¹' (coborder s) := coborder_preimage hf.isOpenMap hf.continuous s @@ -103,7 +103,7 @@ lemma isClosed_preimage_val_coborder : A set is locally closed if it is the intersection of some open set and some closed set. Also see `isLocallyClosed_tfae`. -/ -def IsLocallyClosed (s : Set α) : Prop := ∃ (U Z : Set α), IsOpen U ∧ IsClosed Z ∧ s = U ∩ Z +def IsLocallyClosed (s : Set X) : Prop := ∃ (U Z : Set X), IsOpen U ∧ IsClosed Z ∧ s = U ∩ Z lemma IsOpen.isLocallyClosed (hs : IsOpen s) : IsLocallyClosed s := ⟨_, _, hs, isClosed_univ, (inter_univ _).symm⟩ @@ -117,16 +117,16 @@ lemma IsLocallyClosed.inter (hs : IsLocallyClosed s) (ht : IsLocallyClosed t) : obtain ⟨U₂, Z₂, hU₂, hZ₂, rfl⟩ := ht refine ⟨_, _, hU₁.inter hU₂, hZ₁.inter hZ₂, inter_inter_inter_comm U₁ Z₁ U₂ Z₂⟩ -lemma IsLocallyClosed.preimage {s : Set β} (hs : IsLocallyClosed s) - {f : α → β} (hf : Continuous f) : +lemma IsLocallyClosed.preimage {s : Set Y} (hs : IsLocallyClosed s) + {f : X → Y} (hf : Continuous f) : IsLocallyClosed (f ⁻¹' s) := by obtain ⟨U, Z, hU, hZ, rfl⟩ := hs exact ⟨_, _, hU.preimage hf, hZ.preimage hf, preimage_inter⟩ nonrec -lemma Inducing.isLocallyClosed_iff {s : Set α} - {f : α → β} (hf : Inducing f) : - IsLocallyClosed s ↔ ∃ s' : Set β, IsLocallyClosed s' ∧ f ⁻¹' s' = s := by +lemma Inducing.isLocallyClosed_iff {s : Set X} + {f : X → Y} (hf : Inducing f) : + IsLocallyClosed s ↔ ∃ s' : Set Y, IsLocallyClosed s' ∧ f ⁻¹' s' = s := by simp_rw [IsLocallyClosed, hf.isOpen_iff, hf.isClosed_iff] constructor · rintro ⟨_, _, ⟨U, hU, rfl⟩, ⟨Z, hZ, rfl⟩, rfl⟩ @@ -134,14 +134,14 @@ lemma Inducing.isLocallyClosed_iff {s : Set α} · rintro ⟨_, ⟨U, Z, hU, hZ, rfl⟩, rfl⟩ exact ⟨_, _, ⟨U, hU, rfl⟩, ⟨Z, hZ, rfl⟩, rfl⟩ -lemma Embedding.isLocallyClosed_iff {s : Set α} - {f : α → β} (hf : Embedding f) : - IsLocallyClosed s ↔ ∃ s' : Set β, IsLocallyClosed s' ∧ s' ∩ range f = f '' s := by +lemma Embedding.isLocallyClosed_iff {s : Set X} + {f : X → Y} (hf : Embedding f) : + IsLocallyClosed s ↔ ∃ s' : Set Y, IsLocallyClosed s' ∧ s' ∩ range f = f '' s := by simp_rw [hf.toInducing.isLocallyClosed_iff, ← (image_injective.mpr hf.inj).eq_iff, image_preimage_eq_inter_range] -lemma IsLocallyClosed.image {s : Set α} (hs : IsLocallyClosed s) - {f : α → β} (hf : Inducing f) (hf' : IsLocallyClosed (range f)) : +lemma IsLocallyClosed.image {s : Set X} (hs : IsLocallyClosed s) + {f : X → Y} (hf : Inducing f) (hf' : IsLocallyClosed (range f)) : IsLocallyClosed (f '' s) := by obtain ⟨t, ht, rfl⟩ := hf.isLocallyClosed_iff.mp hs rw [image_preimage_eq_inter_range] @@ -155,7 +155,7 @@ A set `s` is locally closed if one of the equivalent conditions below hold 4. Every `x ∈ s` has some open neighborhood `U` such that `U ∩ closure s ⊆ s`. 5. `s` is open in the closure of `s`. -/ -lemma isLocallyClosed_tfae (s : Set α) : +lemma isLocallyClosed_tfae (s : Set X) : List.TFAE [ IsLocallyClosed s, IsOpen (coborder s), @@ -205,9 +205,9 @@ lemma IsLocallyClosed.isOpen_preimage_val_closure (hs : IsLocallyClosed s) : open TopologicalSpace -variable {ι : Type*} {U : ι → Opens β} (hU : iSup U = ⊤) +variable {ι : Type*} {U : ι → Opens Y} (hU : iSup U = ⊤) -theorem isLocallyClosed_iff_coe_preimage_of_iSup_eq_top (s : Set β) : +theorem isLocallyClosed_iff_coe_preimage_of_iSup_eq_top (s : Set Y) : IsLocallyClosed s ↔ ∀ i, IsLocallyClosed ((↑) ⁻¹' s : Set (U i)) := by simp_rw [isLocallyClosed_iff_isOpen_coborder] rw [isOpen_iff_coe_preimage_of_iSup_eq_top hU] diff --git a/Mathlib/Topology/LocallyConstant/Basic.lean b/Mathlib/Topology/LocallyConstant/Basic.lean index fb974d2b95f99..cd4f5070a4e86 100644 --- a/Mathlib/Topology/LocallyConstant/Basic.lean +++ b/Mathlib/Topology/LocallyConstant/Basic.lean @@ -230,10 +230,10 @@ theorem toFun_eq_coe (f : LocallyConstant X Y) : f.toFun = f := theorem coe_mk (f : X → Y) (h) : ⇑(⟨f, h⟩ : LocallyConstant X Y) = f := rfl -theorem congr_fun {f g : LocallyConstant X Y} (h : f = g) (x : X) : f x = g x := +protected theorem congr_fun {f g : LocallyConstant X Y} (h : f = g) (x : X) : f x = g x := DFunLike.congr_fun h x -theorem congr_arg (f : LocallyConstant X Y) {x y : X} (h : x = y) : f x = f y := +protected theorem congr_arg (f : LocallyConstant X Y) {x y : X} (h : x = y) : f x = f y := DFunLike.congr_arg f h theorem coe_injective : @Function.Injective (LocallyConstant X Y) (X → Y) (↑) := fun _ _ => diff --git a/Mathlib/Topology/MetricSpace/Closeds.lean b/Mathlib/Topology/MetricSpace/Closeds.lean index f2eff1588b276..5f7dca32b13d8 100644 --- a/Mathlib/Topology/MetricSpace/Closeds.lean +++ b/Mathlib/Topology/MetricSpace/Closeds.lean @@ -201,7 +201,7 @@ instance Closeds.compactSpace [CompactSpace α] : CompactSpace (Closeds α) := · intro x hx have : x ∈ ⋃ y ∈ s, ball y δ := hs (by simp) rcases mem_iUnion₂.1 this with ⟨y, ys, dy⟩ - have : edist y x < δ := by simp at dy; rwa [edist_comm] at dy + have : edist y x < δ := by simpa [edist_comm] exact ⟨y, ⟨ys, ⟨x, hx, this⟩⟩, le_of_lt dy⟩ · rintro x ⟨_, ⟨y, yu, hy⟩⟩ exact ⟨y, yu, le_of_lt hy⟩ @@ -342,7 +342,7 @@ instance NonemptyCompacts.secondCountableTopology [SecondCountableTopology α] : have tc : ∀ x ∈ t, ∃ y ∈ c, edist x y ≤ δ := by intro x hx rcases tb x hx with ⟨y, yv, Dxy⟩ - have : y ∈ c := by simp [c, -mem_image]; exact ⟨yv, ⟨x, hx, Dxy⟩⟩ + have : y ∈ c := by simpa [c, -mem_image] using ⟨yv, ⟨x, hx, Dxy⟩⟩ exact ⟨y, this, le_of_lt Dxy⟩ -- points in `c` are well approximated by points in `t` have ct : ∀ y ∈ c, ∃ x ∈ t, edist y x ≤ δ := by diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 89f069caca884..ffa01e46f437f 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -100,10 +100,10 @@ theorem toFun_eq_coe {f : α →ᵈ β} : f.toFun = (f : α → β) := theorem coe_mk (f : α → β) (h) : ⇑(⟨f, h⟩ : α →ᵈ β) = f := rfl -theorem congr_fun {f g : α →ᵈ β} (h : f = g) (x : α) : f x = g x := +protected theorem congr_fun {f g : α →ᵈ β} (h : f = g) (x : α) : f x = g x := DFunLike.congr_fun h x -theorem congr_arg (f : α →ᵈ β) {x y : α} (h : x = y) : f x = f y := +protected theorem congr_arg (f : α →ᵈ β) {x y : α} (h : x = y) : f x = f y := DFunLike.congr_arg f h @[ext] diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 5f0184bc8e286..8063585c24d8e 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 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.NormedSpace.lpSpace +import Mathlib.Analysis.Normed.Lp.lpSpace import Mathlib.Topology.Sets.Compacts /-! diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index bd747e3764638..1c0e6daf3ca63 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -388,8 +388,7 @@ protected def metricSpaceOfDiscreteUniformity {E : ℕ → Type*} [∀ n, Unifor eq_of_dist_eq_zero := PiNat.eq_of_dist_eq_zero _ _ toUniformSpace := Pi.uniformSpace _ uniformity_dist := by - simp [Pi.uniformity, comap_iInf, gt_iff_lt, preimage_setOf_eq, comap_principal, - PseudoMetricSpace.uniformity_dist, h, idRel] + simp only [Pi.uniformity, h, idRel, comap_principal, preimage_setOf_eq] apply le_antisymm · simp only [le_iInf_iff, le_principal_iff] intro ε εpos diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index 78c9aade74fc3..84aa5a4ff45ae 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -61,7 +61,7 @@ noncomputable def ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x, d x dist_comm x y := NNReal.coe_inj.2 <| by refine reverse_surjective.iInf_congr _ fun l ↦ ?_ - rw [← sum_reverse, zipWith_distrib_reverse, reverse_append, reverse_reverse, + rw [← sum_reverse, reverse_zipWith, reverse_append, reverse_reverse, reverse_singleton, singleton_append, reverse_cons, reverse_reverse, zipWith_comm_of_comm _ dist_comm] simp only [length, length_append] @@ -134,7 +134,7 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x intro m hm rw [← not_lt, Nat.lt_iff_add_one_le, ← hL_len] intro hLm - rw [mem_setOf_eq, take_all_of_le hLm, two_mul, add_le_iff_nonpos_left, nonpos_iff_eq_zero, + rw [mem_setOf_eq, take_of_length_le hLm, two_mul, add_le_iff_nonpos_left, nonpos_iff_eq_zero, sum_eq_zero_iff, ← forall_iff_forall_mem, forall_zipWith, ← chain_append_singleton_iff_forall₂] at hm <;> @@ -154,7 +154,7 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x have hMl' : length (take M l) = M := (length_take _ _).trans (min_eq_left hMl.le) refine (ihn _ hMl _ _ _ hMl').trans ?_ convert hMs.1.out - rw [zipWith_distrib_take, take, take_succ, getElem?_append hMl, getElem?_eq_getElem hMl, + rw [take_zipWith, take, take_succ, getElem?_append hMl, getElem?_eq_getElem hMl, ← Option.coe_def, Option.toList_some, take_append_of_le_length hMl.le, getElem_cons_succ] · exact single_le_sum (fun x _ => zero_le x) _ (mem_iff_get.2 ⟨⟨M, hM_lt⟩, getElem_zipWith⟩) · rcases hMl.eq_or_lt with (rfl | hMl) @@ -169,7 +169,7 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x rw [← sum_take_add_sum_drop L (M + 1), two_mul, add_le_add_iff_left, ← add_le_add_iff_right, sum_take_add_sum_drop, ← two_mul] at hMs' convert hMs' - rwa [zipWith_distrib_drop, drop, drop_append_of_le_length] + rwa [drop_zipWith, drop, drop_append_of_le_length] end PseudoMetricSpace diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index 57edd4a7cc6ed..e7570fcfdd2cf 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -77,45 +77,45 @@ section OrderTopology section Preorder -variable [TopologicalSpace α] [Preorder α] [t : OrderTopology α] +variable [TopologicalSpace α] [Preorder α] -instance : OrderTopology αᵒᵈ := +instance [t : OrderTopology α] : OrderTopology αᵒᵈ := ⟨by convert OrderTopology.topology_eq_generate_intervals (α := α) using 6 apply or_comm⟩ -theorem isOpen_iff_generate_intervals {s : Set α} : +theorem isOpen_iff_generate_intervals [t : OrderTopology α] {s : Set α} : IsOpen s ↔ GenerateOpen { s | ∃ a, s = Ioi a ∨ s = Iio a } s := by rw [t.topology_eq_generate_intervals]; rfl -theorem isOpen_lt' (a : α) : IsOpen { b : α | a < b } := +theorem isOpen_lt' [OrderTopology α] (a : α) : IsOpen { b : α | a < b } := isOpen_iff_generate_intervals.2 <| .basic _ ⟨a, .inl rfl⟩ -theorem isOpen_gt' (a : α) : IsOpen { b : α | b < a } := +theorem isOpen_gt' [OrderTopology α] (a : α) : IsOpen { b : α | b < a } := isOpen_iff_generate_intervals.2 <| .basic _ ⟨a, .inr rfl⟩ -theorem lt_mem_nhds {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 b, a < x := +theorem lt_mem_nhds [OrderTopology α] {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 b, a < x := (isOpen_lt' _).mem_nhds h -theorem le_mem_nhds {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 b, a ≤ x := +theorem le_mem_nhds [OrderTopology α] {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 b, a ≤ x := (lt_mem_nhds h).mono fun _ => le_of_lt -theorem gt_mem_nhds {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 a, x < b := +theorem gt_mem_nhds [OrderTopology α] {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 a, x < b := (isOpen_gt' _).mem_nhds h -theorem ge_mem_nhds {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 a, x ≤ b := +theorem ge_mem_nhds [OrderTopology α] {a b : α} (h : a < b) : ∀ᶠ x in 𝓝 a, x ≤ b := (gt_mem_nhds h).mono fun _ => le_of_lt -theorem nhds_eq_order (a : α) : 𝓝 a = (⨅ b ∈ Iio a, 𝓟 (Ioi b)) ⊓ ⨅ b ∈ Ioi a, 𝓟 (Iio b) := by - rw [t.topology_eq_generate_intervals, nhds_generateFrom] - simp_rw [mem_setOf_eq, @and_comm (a ∈ _), exists_or, or_and_right, iInf_or, iInf_and, iInf_exists, - iInf_inf_eq, iInf_comm (ι := Set α), iInf_iInf_eq_left, mem_Ioi, mem_Iio] +theorem nhds_eq_order [OrderTopology α] (a : α) : + 𝓝 a = (⨅ b ∈ Iio a, 𝓟 (Ioi b)) ⊓ ⨅ b ∈ Ioi a, 𝓟 (Iio b) := by + rw [OrderTopology.topology_eq_generate_intervals (α := α), nhds_generateFrom] + simp_rw [mem_setOf_eq, @and_comm (a ∈ _), exists_or, or_and_right, iInf_or, iInf_and, + iInf_exists, iInf_inf_eq, iInf_comm (ι := Set α), iInf_iInf_eq_left, mem_Ioi, mem_Iio] -theorem tendsto_order {f : β → α} {a : α} {x : Filter β} : +theorem tendsto_order [OrderTopology α] {f : β → α} {a : α} {x : Filter β} : Tendsto f x (𝓝 a) ↔ (∀ a' < a, ∀ᶠ b in x, a' < f b) ∧ ∀ a' > a, ∀ᶠ b in x, f b < a' := by simp only [nhds_eq_order a, tendsto_inf, tendsto_iInf, tendsto_principal]; rfl - -instance tendstoIccClassNhds (a : α) : TendstoIxxClass Icc (𝓝 a) (𝓝 a) := by +instance tendstoIccClassNhds [OrderTopology α] (a : α) : TendstoIxxClass Icc (𝓝 a) (𝓝 a) := by simp only [nhds_eq_order, iInf_subtype'] refine ((hasBasis_iInf_principal_finite _).inf (hasBasis_iInf_principal_finite _)).tendstoIxxClass @@ -123,36 +123,36 @@ instance tendstoIccClassNhds (a : α) : TendstoIxxClass Icc (𝓝 a) (𝓝 a) := refine ((ordConnected_biInter ?_).inter (ordConnected_biInter ?_)).out <;> intro _ _ exacts [ordConnected_Ioi, ordConnected_Iio] -instance tendstoIcoClassNhds (a : α) : TendstoIxxClass Ico (𝓝 a) (𝓝 a) := +instance tendstoIcoClassNhds [OrderTopology α] (a : α) : TendstoIxxClass Ico (𝓝 a) (𝓝 a) := tendstoIxxClass_of_subset fun _ _ => Ico_subset_Icc_self -instance tendstoIocClassNhds (a : α) : TendstoIxxClass Ioc (𝓝 a) (𝓝 a) := +instance tendstoIocClassNhds [OrderTopology α] (a : α) : TendstoIxxClass Ioc (𝓝 a) (𝓝 a) := tendstoIxxClass_of_subset fun _ _ => Ioc_subset_Icc_self -instance tendstoIooClassNhds (a : α) : TendstoIxxClass Ioo (𝓝 a) (𝓝 a) := +instance tendstoIooClassNhds [OrderTopology α] (a : α) : TendstoIxxClass Ioo (𝓝 a) (𝓝 a) := tendstoIxxClass_of_subset fun _ _ => Ioo_subset_Icc_self /-- **Squeeze theorem** (also known as **sandwich theorem**). This version assumes that inequalities hold eventually for the filter. -/ -theorem tendsto_of_tendsto_of_tendsto_of_le_of_le' {f g h : β → α} {b : Filter β} {a : α} - (hg : Tendsto g b (𝓝 a)) (hh : Tendsto h b (𝓝 a)) (hgf : ∀ᶠ b in b, g b ≤ f b) +theorem tendsto_of_tendsto_of_tendsto_of_le_of_le' [OrderTopology α] {f g h : β → α} {b : Filter β} + {a : α} (hg : Tendsto g b (𝓝 a)) (hh : Tendsto h b (𝓝 a)) (hgf : ∀ᶠ b in b, g b ≤ f b) (hfh : ∀ᶠ b in b, f b ≤ h b) : Tendsto f b (𝓝 a) := (hg.Icc hh).of_smallSets <| hgf.and hfh /-- **Squeeze theorem** (also known as **sandwich theorem**). This version assumes that inequalities hold everywhere. -/ -theorem tendsto_of_tendsto_of_tendsto_of_le_of_le {f g h : β → α} {b : Filter β} {a : α} - (hg : Tendsto g b (𝓝 a)) (hh : Tendsto h b (𝓝 a)) (hgf : g ≤ f) (hfh : f ≤ h) : +theorem tendsto_of_tendsto_of_tendsto_of_le_of_le [OrderTopology α] {f g h : β → α} {b : Filter β} + {a : α} (hg : Tendsto g b (𝓝 a)) (hh : Tendsto h b (𝓝 a)) (hgf : g ≤ f) (hfh : f ≤ h) : Tendsto f b (𝓝 a) := tendsto_of_tendsto_of_tendsto_of_le_of_le' hg hh (eventually_of_forall hgf) (eventually_of_forall hfh) -theorem nhds_order_unbounded {a : α} (hu : ∃ u, a < u) (hl : ∃ l, l < a) : +theorem nhds_order_unbounded [OrderTopology α] {a : α} (hu : ∃ u, a < u) (hl : ∃ l, l < a) : 𝓝 a = ⨅ (l) (_ : l < a) (u) (_ : a < u), 𝓟 (Ioo l u) := by simp only [nhds_eq_order, ← inf_biInf, ← biInf_inf, *, ← inf_principal, ← Ioi_inter_Iio]; rfl -theorem tendsto_order_unbounded {f : β → α} {a : α} {x : Filter β} (hu : ∃ u, a < u) - (hl : ∃ l, l < a) (h : ∀ l u, l < a → a < u → ∀ᶠ b in x, l < f b ∧ f b < u) : +theorem tendsto_order_unbounded [OrderTopology α] {f : β → α} {a : α} {x : Filter β} + (hu : ∃ u, a < u) (hl : ∃ l, l < a) (h : ∀ l u, l < a → a < u → ∀ᶠ b in x, l < f b ∧ f b < u) : Tendsto f x (𝓝 a) := by simp only [nhds_order_unbounded hu hl, tendsto_iInf, tendsto_principal] exact fun l hl u => h l u hl @@ -332,43 +332,42 @@ variable [TopologicalSpace α] [LinearOrder α] section OrderTopology -variable [OrderTopology α] - -theorem order_separated {a₁ a₂ : α} (h : a₁ < a₂) : +theorem order_separated [OrderTopology α] {a₁ a₂ : α} (h : a₁ < a₂) : ∃ u v : Set α, IsOpen u ∧ IsOpen v ∧ a₁ ∈ u ∧ a₂ ∈ v ∧ ∀ b₁ ∈ u, ∀ b₂ ∈ v, b₁ < b₂ := let ⟨x, hx, y, hy, h⟩ := h.exists_disjoint_Iio_Ioi ⟨Iio x, Ioi y, isOpen_gt' _, isOpen_lt' _, hx, hy, h⟩ -- see Note [lower instance priority] -instance (priority := 100) OrderTopology.to_orderClosedTopology : OrderClosedTopology α where +instance (priority := 100) OrderTopology.to_orderClosedTopology [OrderTopology α] : + OrderClosedTopology α where isClosed_le' := isOpen_compl_iff.1 <| isOpen_prod_iff.mpr fun a₁ a₂ (h : ¬a₁ ≤ a₂) => have h : a₂ < a₁ := lt_of_not_ge h let ⟨u, v, hu, hv, ha₁, ha₂, h⟩ := order_separated h ⟨v, u, hv, hu, ha₂, ha₁, fun ⟨b₁, b₂⟩ ⟨h₁, h₂⟩ => not_le_of_gt <| h b₂ h₂ b₁ h₁⟩ -theorem exists_Ioc_subset_of_mem_nhds {a : α} {s : Set α} (hs : s ∈ 𝓝 a) (h : ∃ l, l < a) : - ∃ l < a, Ioc l a ⊆ s := +theorem exists_Ioc_subset_of_mem_nhds [OrderTopology α] {a : α} {s : Set α} (hs : s ∈ 𝓝 a) + (h : ∃ l, l < a) : ∃ l < a, Ioc l a ⊆ s := (nhdsWithin_Iic_basis' h).mem_iff.mp (nhdsWithin_le_nhds hs) -theorem exists_Ioc_subset_of_mem_nhds' {a : α} {s : Set α} (hs : s ∈ 𝓝 a) {l : α} (hl : l < a) : - ∃ l' ∈ Ico l a, Ioc l' a ⊆ s := +theorem exists_Ioc_subset_of_mem_nhds' [OrderTopology α] {a : α} {s : Set α} (hs : s ∈ 𝓝 a) {l : α} + (hl : l < a) : ∃ l' ∈ Ico l a, Ioc l' a ⊆ s := let ⟨l', hl'a, hl's⟩ := exists_Ioc_subset_of_mem_nhds hs ⟨l, hl⟩ ⟨max l l', ⟨le_max_left _ _, max_lt hl hl'a⟩, (Ioc_subset_Ioc_left <| le_max_right _ _).trans hl's⟩ -theorem exists_Ico_subset_of_mem_nhds' {a : α} {s : Set α} (hs : s ∈ 𝓝 a) {u : α} (hu : a < u) : - ∃ u' ∈ Ioc a u, Ico a u' ⊆ s := by +theorem exists_Ico_subset_of_mem_nhds' [OrderTopology α] {a : α} {s : Set α} (hs : s ∈ 𝓝 a) {u : α} + (hu : a < u) : ∃ u' ∈ Ioc a u, Ico a u' ⊆ s := by simpa only [OrderDual.exists, exists_prop, dual_Ico, dual_Ioc] using exists_Ioc_subset_of_mem_nhds' (show ofDual ⁻¹' s ∈ 𝓝 (toDual a) from hs) hu.dual -theorem exists_Ico_subset_of_mem_nhds {a : α} {s : Set α} (hs : s ∈ 𝓝 a) (h : ∃ u, a < u) : - ∃ u, a < u ∧ Ico a u ⊆ s := +theorem exists_Ico_subset_of_mem_nhds [OrderTopology α] {a : α} {s : Set α} (hs : s ∈ 𝓝 a) + (h : ∃ u, a < u) : ∃ u, a < u ∧ Ico a u ⊆ s := let ⟨_l', hl'⟩ := h let ⟨l, hl⟩ := exists_Ico_subset_of_mem_nhds' hs hl' ⟨l, hl.1.1, hl.2⟩ -theorem exists_Icc_mem_subset_of_mem_nhdsWithin_Ici {a : α} {s : Set α} (hs : s ∈ 𝓝[≥] a) : - ∃ b, a ≤ b ∧ Icc a b ∈ 𝓝[≥] a ∧ Icc a b ⊆ s := by +theorem exists_Icc_mem_subset_of_mem_nhdsWithin_Ici [OrderTopology α] {a : α} {s : Set α} + (hs : s ∈ 𝓝[≥] a) : ∃ b, a ≤ b ∧ Icc a b ∈ 𝓝[≥] a ∧ Icc a b ⊆ s := by rcases (em (IsMax a)).imp_right not_isMax_iff.mp with (ha | ha) · use a simpa [ha.Ici_eq] using hs @@ -379,12 +378,12 @@ theorem exists_Icc_mem_subset_of_mem_nhdsWithin_Ici {a : α} {s : Set α} (hs : · refine ⟨c, hac.le, Icc_mem_nhdsWithin_Ici' hac, ?_⟩ exact (Icc_subset_Ico_right hcb).trans hbs -theorem exists_Icc_mem_subset_of_mem_nhdsWithin_Iic {a : α} {s : Set α} (hs : s ∈ 𝓝[≤] a) : - ∃ b ≤ a, Icc b a ∈ 𝓝[≤] a ∧ Icc b a ⊆ s := by +theorem exists_Icc_mem_subset_of_mem_nhdsWithin_Iic [OrderTopology α] {a : α} {s : Set α} + (hs : s ∈ 𝓝[≤] a) : ∃ b ≤ a, Icc b a ∈ 𝓝[≤] a ∧ Icc b a ⊆ s := by simpa only [dual_Icc, toDual.surjective.exists] using exists_Icc_mem_subset_of_mem_nhdsWithin_Ici (α := αᵒᵈ) (a := toDual a) hs -theorem exists_Icc_mem_subset_of_mem_nhds {a : α} {s : Set α} (hs : s ∈ 𝓝 a) : +theorem exists_Icc_mem_subset_of_mem_nhds [OrderTopology α] {a : α} {s : Set α} (hs : s ∈ 𝓝 a) : ∃ b c, a ∈ Icc b c ∧ Icc b c ∈ 𝓝 a ∧ Icc b c ⊆ s := by rcases exists_Icc_mem_subset_of_mem_nhdsWithin_Iic (nhdsWithin_le_nhds hs) with ⟨b, hba, hb_nhds, hbs⟩ @@ -394,8 +393,8 @@ theorem exists_Icc_mem_subset_of_mem_nhds {a : α} {s : Set α} (hs : s ∈ 𝓝 rw [← Icc_union_Icc_eq_Icc hba hac, ← nhds_left_sup_nhds_right] exact ⟨union_mem_sup hb_nhds hc_nhds, union_subset hbs hcs⟩ -theorem IsOpen.exists_Ioo_subset [Nontrivial α] {s : Set α} (hs : IsOpen s) (h : s.Nonempty) : - ∃ a b, a < b ∧ Ioo a b ⊆ s := by +theorem IsOpen.exists_Ioo_subset [OrderTopology α] [Nontrivial α] {s : Set α} (hs : IsOpen s) + (h : s.Nonempty) : ∃ a b, a < b ∧ Ioo a b ⊆ s := by obtain ⟨x, hx⟩ : ∃ x, x ∈ s := h obtain ⟨y, hy⟩ : ∃ y, y ≠ x := exists_ne x rcases lt_trichotomy x y with (H | rfl | H) @@ -407,7 +406,7 @@ theorem IsOpen.exists_Ioo_subset [Nontrivial α] {s : Set α} (hs : IsOpen s) (h exists_Ioc_subset_of_mem_nhds (hs.mem_nhds hx) ⟨y, H⟩ exact ⟨l, x, lx, Ioo_subset_Ioc_self.trans hl⟩ -theorem dense_of_exists_between [Nontrivial α] {s : Set α} +theorem dense_of_exists_between [OrderTopology α] [Nontrivial α] {s : Set α} (h : ∀ ⦃a b⦄, a < b → ∃ c ∈ s, a < c ∧ c < b) : Dense s := by refine dense_iff_inter_open.2 fun U U_open U_nonempty => ?_ obtain ⟨a, b, hab, H⟩ : ∃ a b : α, a < b ∧ Ioo a b ⊆ U := U_open.exists_Ioo_subset U_nonempty @@ -417,14 +416,14 @@ theorem dense_of_exists_between [Nontrivial α] {s : Set α} /-- A set in a nontrivial densely linear ordered type is dense in the sense of topology if and only if for any `a < b` there exists `c ∈ s`, `a < c < b`. Each implication requires less typeclass assumptions. -/ -theorem dense_iff_exists_between [DenselyOrdered α] [Nontrivial α] {s : Set α} : +theorem dense_iff_exists_between [OrderTopology α] [DenselyOrdered α] [Nontrivial α] {s : Set α} : Dense s ↔ ∀ a b, a < b → ∃ c ∈ s, a < c ∧ c < b := ⟨fun h _ _ hab => h.exists_between hab, dense_of_exists_between⟩ /-- A set is a neighborhood of `a` if and only if it contains an interval `(l, u)` containing `a`, provided `a` is neither a bottom element nor a top element. -/ -theorem mem_nhds_iff_exists_Ioo_subset' {a : α} {s : Set α} (hl : ∃ l, l < a) (hu : ∃ u, a < u) : - s ∈ 𝓝 a ↔ ∃ l u, a ∈ Ioo l u ∧ Ioo l u ⊆ s := by +theorem mem_nhds_iff_exists_Ioo_subset' [OrderTopology α] {a : α} {s : Set α} (hl : ∃ l, l < a) + (hu : ∃ u, a < u) : s ∈ 𝓝 a ↔ ∃ l u, a ∈ Ioo l u ∧ Ioo l u ⊆ s := by constructor · intro h rcases exists_Ico_subset_of_mem_nhds h hu with ⟨u, au, hu⟩ @@ -435,24 +434,24 @@ theorem mem_nhds_iff_exists_Ioo_subset' {a : α} {s : Set α} (hl : ∃ l, l < a /-- A set is a neighborhood of `a` if and only if it contains an interval `(l, u)` containing `a`. -/ -theorem mem_nhds_iff_exists_Ioo_subset [NoMaxOrder α] [NoMinOrder α] {a : α} {s : Set α} : - s ∈ 𝓝 a ↔ ∃ l u, a ∈ Ioo l u ∧ Ioo l u ⊆ s := +theorem mem_nhds_iff_exists_Ioo_subset [OrderTopology α] [NoMaxOrder α] [NoMinOrder α] {a : α} + {s : Set α} : s ∈ 𝓝 a ↔ ∃ l u, a ∈ Ioo l u ∧ Ioo l u ⊆ s := mem_nhds_iff_exists_Ioo_subset' (exists_lt a) (exists_gt a) -theorem nhds_basis_Ioo' {a : α} (hl : ∃ l, l < a) (hu : ∃ u, a < u) : +theorem nhds_basis_Ioo' [OrderTopology α] {a : α} (hl : ∃ l, l < a) (hu : ∃ u, a < u) : (𝓝 a).HasBasis (fun b : α × α => b.1 < a ∧ a < b.2) fun b => Ioo b.1 b.2 := ⟨fun s => (mem_nhds_iff_exists_Ioo_subset' hl hu).trans <| by simp⟩ -theorem nhds_basis_Ioo [NoMaxOrder α] [NoMinOrder α] (a : α) : +theorem nhds_basis_Ioo [OrderTopology α] [NoMaxOrder α] [NoMinOrder α] (a : α) : (𝓝 a).HasBasis (fun b : α × α => b.1 < a ∧ a < b.2) fun b => Ioo b.1 b.2 := nhds_basis_Ioo' (exists_lt a) (exists_gt a) -theorem Filter.Eventually.exists_Ioo_subset [NoMaxOrder α] [NoMinOrder α] {a : α} {p : α → Prop} - (hp : ∀ᶠ x in 𝓝 a, p x) : ∃ l u, a ∈ Ioo l u ∧ Ioo l u ⊆ { x | p x } := +theorem Filter.Eventually.exists_Ioo_subset [OrderTopology α] [NoMaxOrder α] [NoMinOrder α] {a : α} + {p : α → Prop} (hp : ∀ᶠ x in 𝓝 a, p x) : ∃ l u, a ∈ Ioo l u ∧ Ioo l u ⊆ { x | p x } := mem_nhds_iff_exists_Ioo_subset.1 hp -theorem Dense.topology_eq_generateFrom [DenselyOrdered α] {s : Set α} (hs : Dense s) : - ‹TopologicalSpace α› = .generateFrom (Ioi '' s ∪ Iio '' s) := by +theorem Dense.topology_eq_generateFrom [OrderTopology α] [DenselyOrdered α] {s : Set α} + (hs : Dense s) : ‹TopologicalSpace α› = .generateFrom (Ioi '' s ∪ Iio '' s) := by refine (OrderTopology.topology_eq_generate_intervals (α := α)).trans ?_ refine le_antisymm (generateFrom_anti ?_) (le_generateFrom ?_) · simp only [union_subset_iff, image_subset_iff] @@ -479,7 +478,7 @@ variable (α) /-- Let `α` be a densely ordered linear order with order topology. If `α` is a separable space, then it has second countable topology. Note that the "densely ordered" assumption cannot be dropped, see [double arrow space](https://topology.pi-base.org/spaces/S000093) for a counterexample. -/ -theorem SecondCountableTopology.of_separableSpace_orderTopology [DenselyOrdered α] +theorem SecondCountableTopology.of_separableSpace_orderTopology [OrderTopology α] [DenselyOrdered α] [SeparableSpace α] : SecondCountableTopology α := by rcases exists_countable_dense α with ⟨s, hc, hd⟩ refine ⟨⟨_, ?_, hd.topology_eq_generateFrom⟩⟩ @@ -489,7 +488,7 @@ variable {α} /-- The set of points which are isolated on the right is countable when the space is second-countable. -/ -theorem countable_setOf_covBy_right [SecondCountableTopology α] : +theorem countable_setOf_covBy_right [OrderTopology α] [SecondCountableTopology α] : Set.Countable { x : α | ∃ y, x ⋖ y } := by nontriviality α let s := { x : α | ∃ y, x ⋖ y } @@ -531,14 +530,14 @@ theorem countable_setOf_covBy_right [SecondCountableTopology α] : /-- The set of points which are isolated on the left is countable when the space is second-countable. -/ -theorem countable_setOf_covBy_left [SecondCountableTopology α] : +theorem countable_setOf_covBy_left [OrderTopology α] [SecondCountableTopology α] : Set.Countable { x : α | ∃ y, y ⋖ x } := by convert countable_setOf_covBy_right (α := αᵒᵈ) using 5 exact toDual_covBy_toDual_iff.symm /-- The set of points which are isolated on the left is countable when the space is second-countable. -/ -theorem countable_of_isolated_left' [SecondCountableTopology α] : +theorem countable_of_isolated_left' [OrderTopology α] [SecondCountableTopology α] : Set.Countable { x : α | ∃ y, y < x ∧ Ioo y x = ∅ } := by simpa only [← covBy_iff_Ioo_eq] using countable_setOf_covBy_left @@ -546,8 +545,9 @@ theorem countable_of_isolated_left' [SecondCountableTopology α] : Then the family is countable. This is not a straightforward consequence of second-countability as some of these intervals might be empty (but in fact this can happen only for countably many of them). -/ -theorem Set.PairwiseDisjoint.countable_of_Ioo [SecondCountableTopology α] {y : α → α} {s : Set α} - (h : PairwiseDisjoint s fun x => Ioo x (y x)) (h' : ∀ x ∈ s, x < y x) : s.Countable := +theorem Set.PairwiseDisjoint.countable_of_Ioo [OrderTopology α] [SecondCountableTopology α] + {y : α → α} {s : Set α} (h : PairwiseDisjoint s fun x => Ioo x (y x)) + (h' : ∀ x ∈ s, x < y x) : s.Countable := have : (s \ { x | ∃ y, x ⋖ y }).Countable := (h.subset diff_subset).countable_of_isOpen (fun _ _ => isOpen_Ioo) fun x hx => (h' _ hx.1).exists_lt_lt (mt (Exists.intro (y x)) hx.2) @@ -555,8 +555,8 @@ theorem Set.PairwiseDisjoint.countable_of_Ioo [SecondCountableTopology α] {y : /-- For a function taking values in a second countable space, the set of points `x` for which the image under `f` of `(x, ∞)` is separated above from `f x` is countable. -/ -theorem countable_image_lt_image_Ioi [LinearOrder β] (f : β → α) [SecondCountableTopology α] : - Set.Countable {x | ∃ z, f x < z ∧ ∀ y, x < y → z ≤ f y} := by +theorem countable_image_lt_image_Ioi [OrderTopology α] [LinearOrder β] (f : β → α) + [SecondCountableTopology α] : Set.Countable {x | ∃ z, f x < z ∧ ∀ y, x < y → z ≤ f y} := by /- If the values of `f` are separated above on the right of `x`, there is an interval `(f x, z x)` which is not reached by `f`. This gives a family of disjoint open intervals in `α`. Such a family can only be countable as `α` is second-countable. -/ @@ -591,23 +591,23 @@ theorem countable_image_lt_image_Ioi [LinearOrder β] (f : β → α) [SecondCou /-- For a function taking values in a second countable space, the set of points `x` for which the image under `f` of `(x, ∞)` is separated below from `f x` is countable. -/ -theorem countable_image_gt_image_Ioi [LinearOrder β] (f : β → α) [SecondCountableTopology α] : - Set.Countable {x | ∃ z, z < f x ∧ ∀ y, x < y → f y ≤ z} := +theorem countable_image_gt_image_Ioi [OrderTopology α] [LinearOrder β] (f : β → α) + [SecondCountableTopology α] : Set.Countable {x | ∃ z, z < f x ∧ ∀ y, x < y → f y ≤ z} := countable_image_lt_image_Ioi (α := αᵒᵈ) f /-- For a function taking values in a second countable space, the set of points `x` for which the image under `f` of `(-∞, x)` is separated above from `f x` is countable. -/ -theorem countable_image_lt_image_Iio [LinearOrder β] (f : β → α) [SecondCountableTopology α] : - Set.Countable {x | ∃ z, f x < z ∧ ∀ y, y < x → z ≤ f y} := +theorem countable_image_lt_image_Iio [OrderTopology α] [LinearOrder β] (f : β → α) + [SecondCountableTopology α] : Set.Countable {x | ∃ z, f x < z ∧ ∀ y, y < x → z ≤ f y} := countable_image_lt_image_Ioi (β := βᵒᵈ) f /-- For a function taking values in a second countable space, the set of points `x` for which the image under `f` of `(-∞, x)` is separated below from `f x` is countable. -/ -theorem countable_image_gt_image_Iio [LinearOrder β] (f : β → α) [SecondCountableTopology α] : - Set.Countable {x | ∃ z, z < f x ∧ ∀ y, y < x → f y ≤ z} := +theorem countable_image_gt_image_Iio [OrderTopology α] [LinearOrder β] (f : β → α) + [SecondCountableTopology α] : Set.Countable {x | ∃ z, z < f x ∧ ∀ y, y < x → f y ≤ z} := countable_image_lt_image_Ioi (α := αᵒᵈ) (β := βᵒᵈ) f -instance instIsCountablyGenerated_atTop [SecondCountableTopology α] : +instance instIsCountablyGenerated_atTop [OrderTopology α] [SecondCountableTopology α] : IsCountablyGenerated (atTop : Filter α) := by by_cases h : ∃ (x : α), IsTop x · rcases h with ⟨x, hx⟩ @@ -637,7 +637,7 @@ instance instIsCountablyGenerated_atTop [SecondCountableTopology α] : rw [this] exact ⟨_, (countable_range _).image _, rfl⟩ -instance instIsCountablyGenerated_atBot [SecondCountableTopology α] : +instance instIsCountablyGenerated_atBot [OrderTopology α] [SecondCountableTopology α] : IsCountablyGenerated (atBot : Filter α) := @instIsCountablyGenerated_atTop αᵒᵈ _ _ _ _ @@ -651,7 +651,7 @@ sometimes Lean fails to unify different instances while trying to apply the depe e.g., `ι → ℝ`. -/ -variable {ι : Type*} {π : ι → Type*} [Finite ι] [∀ i, LinearOrder (π i)] +variable [OrderTopology α] {ι : Type*} {π : ι → Type*} [Finite ι] [∀ i, LinearOrder (π i)] [∀ i, TopologicalSpace (π i)] [∀ i, OrderTopology (π i)] {a b x : ∀ i, π i} {a' b' x' : ι → α} theorem pi_Iic_mem_nhds (ha : ∀ i, x i < a i) : Iic a ∈ 𝓝 x := @@ -674,15 +674,14 @@ theorem pi_Icc_mem_nhds' (ha : ∀ i, a' i < x' i) (hb : ∀ i, x' i < b' i) : I variable [Nonempty ι] -theorem pi_Iio_mem_nhds (ha : ∀ i, x i < a i) : Iio a ∈ 𝓝 x := by - refine mem_of_superset (set_pi_mem_nhds Set.finite_univ fun i _ => ?_) (pi_univ_Iio_subset a) - exact Iio_mem_nhds (ha i) +theorem pi_Iio_mem_nhds (ha : ∀ i, x i < a i) : Iio a ∈ 𝓝 x := mem_of_superset + (set_pi_mem_nhds finite_univ fun i _ ↦ Iio_mem_nhds (ha i)) (pi_univ_Iio_subset a) theorem pi_Iio_mem_nhds' (ha : ∀ i, x' i < a' i) : Iio a' ∈ 𝓝 x' := pi_Iio_mem_nhds ha theorem pi_Ioi_mem_nhds (ha : ∀ i, a i < x i) : Ioi a ∈ 𝓝 x := - @pi_Iio_mem_nhds ι (fun i => (π i)ᵒᵈ) _ _ _ _ _ _ _ ha + pi_Iio_mem_nhds (π := fun i => (π i)ᵒᵈ) ha theorem pi_Ioi_mem_nhds' (ha : ∀ i, a' i < x' i) : Ioi a' ∈ 𝓝 x' := pi_Ioi_mem_nhds ha diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index 1dafe6657a3fd..a439497b79088 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -97,10 +97,15 @@ that for every open covering `U : Set (Set X)` of `s` there exists a partition o subordinate to `U`. -/ structure PartitionOfUnity (ι X : Type*) [TopologicalSpace X] (s : Set X := univ) where + /-- The collection of continuous functions underlying this partition of unity -/ toFun : ι → C(X, ℝ) + /-- the supports of the underlying functions are a locally finite family of sets -/ locallyFinite' : LocallyFinite fun i => support (toFun i) + /-- the functions are non-negative -/ nonneg' : 0 ≤ toFun + /-- the functions sum up to one on `s` -/ sum_eq_one' : ∀ x ∈ s, ∑ᶠ i, toFun i x = 1 + /-- the functions sum up to at most one, globally -/ sum_le_one' : ∀ x, ∑ᶠ i, toFun i x ≤ 1 /-- A `BumpCovering ι X s` is an indexed family of functions `f i`, `i : ι`, such that @@ -119,10 +124,15 @@ every open covering `U : Set (Set X)` of `s` there exists a `BumpCovering` of `s subordinate to `U`. -/ structure BumpCovering (ι X : Type*) [TopologicalSpace X] (s : Set X := univ) where + /-- The collections of continuous functions underlying this bump covering -/ toFun : ι → C(X, ℝ) + /-- the supports of the underlying functions are a locally finite family of sets -/ locallyFinite' : LocallyFinite fun i => support (toFun i) + /-- the functions are non-negative -/ nonneg' : 0 ≤ toFun + /-- the functions are each at most one -/ le_one' : toFun ≤ 1 + /-- Each point `x ∈ s` belongs to the interior of `{x | f i x = 1}` for some `i`. -/ eventuallyEq_one' : ∀ x ∈ s, ∃ i, toFun i =ᶠ[𝓝 x] 1 variable {ι : Type u} {X : Type v} [TopologicalSpace X] diff --git a/Mathlib/Topology/Perfect.lean b/Mathlib/Topology/Perfect.lean index 1905b577b1aeb..31622ec28846e 100644 --- a/Mathlib/Topology/Perfect.lean +++ b/Mathlib/Topology/Perfect.lean @@ -168,6 +168,25 @@ theorem Perfect.splitting [T25Space α] (hC : Perfect C) (hnonempty : C.Nonempty exact inter_subset_right apply Disjoint.mono _ _ hUV <;> apply closure_mono <;> exact inter_subset_left +lemma IsPreconnected.preperfect_of_nontrivial [T1Space α] {U : Set α} (hu : U.Nontrivial) + (h : IsPreconnected U) : Preperfect U := by + intro x hx + rw [isPreconnected_closed_iff] at h + specialize h {x} (closure (U \ {x})) isClosed_singleton isClosed_closure ?_ ?_ ?_ + · trans {x} ∪ (U \ {x}) + · simp + apply Set.union_subset_union_right + exact subset_closure + · exact Set.inter_singleton_nonempty.mpr hx + · obtain ⟨y, hy⟩ := Set.Nontrivial.exists_ne hu x + use y + simp only [Set.mem_inter_iff, hy, true_and] + apply subset_closure + simp [hy] + · apply Set.Nonempty.right at h + rw [Set.singleton_inter_nonempty, mem_closure_iff_clusterPt, ← acc_principal_iff_cluster] at h + exact h + end Preperfect section Kernel diff --git a/Mathlib/Topology/Semicontinuous.lean b/Mathlib/Topology/Semicontinuous.lean index 06436863c0514..b48b959356b6c 100644 --- a/Mathlib/Topology/Semicontinuous.lean +++ b/Mathlib/Topology/Semicontinuous.lean @@ -437,16 +437,12 @@ theorem LowerSemicontinuousWithinAt.add' {f g : α → γ} (hf : LowerSemicontin filter_upwards [hf z₁ z₁lt, hg z₂ z₂lt] with z h₁z h₂z have A1 : min (f z) (f x) ∈ u := by by_cases H : f z ≤ f x - · simp [H] - exact h₁ ⟨h₁z, H⟩ - · simp [le_of_not_le H] - exact h₁ ⟨z₁lt, le_rfl⟩ + · simpa [H] using h₁ ⟨h₁z, H⟩ + · simpa [le_of_not_le H] have A2 : min (g z) (g x) ∈ v := by by_cases H : g z ≤ g x - · simp [H] - exact h₂ ⟨h₂z, H⟩ - · simp [le_of_not_le H] - exact h₂ ⟨z₂lt, le_rfl⟩ + · simpa [H] using h₂ ⟨h₂z, H⟩ + · simpa [le_of_not_le H] have : (min (f z) (f x), min (g z) (g x)) ∈ u ×ˢ v := ⟨A1, A2⟩ calc y < min (f z) (f x) + min (g z) (g x) := h this @@ -456,10 +452,8 @@ theorem LowerSemicontinuousWithinAt.add' {f g : α → γ} (hf : LowerSemicontin filter_upwards [hf z₁ z₁lt] with z h₁z have A1 : min (f z) (f x) ∈ u := by by_cases H : f z ≤ f x - · simp [H] - exact h₁ ⟨h₁z, H⟩ - · simp [le_of_not_le H] - exact h₁ ⟨z₁lt, le_rfl⟩ + · simpa [H] using h₁ ⟨h₁z, H⟩ + · simpa [le_of_not_le H] have : (min (f z) (f x), g x) ∈ u ×ˢ v := ⟨A1, xv⟩ calc y < min (f z) (f x) + g x := h this @@ -472,10 +466,8 @@ theorem LowerSemicontinuousWithinAt.add' {f g : α → γ} (hf : LowerSemicontin filter_upwards [hg z₂ z₂lt] with z h₂z have A2 : min (g z) (g x) ∈ v := by by_cases H : g z ≤ g x - · simp [H] - exact h₂ ⟨h₂z, H⟩ - · simp [le_of_not_le H] - exact h₂ ⟨z₂lt, le_rfl⟩ + · simpa [H] using h₂ ⟨h₂z, H⟩ + · simpa [le_of_not_le H] using h₂ ⟨z₂lt, le_rfl⟩ have : (f x, min (g z) (g x)) ∈ u ×ˢ v := ⟨xu, A2⟩ calc y < f x + min (g z) (g x) := h this diff --git a/Mathlib/Topology/Sequences.lean b/Mathlib/Topology/Sequences.lean index ec9d10d544e10..701c434c73709 100644 --- a/Mathlib/Topology/Sequences.lean +++ b/Mathlib/Topology/Sequences.lean @@ -243,7 +243,7 @@ theorem IsSeqCompact.subseq_of_frequently_in {s : Set X} (hs : IsSeqCompact s) { theorem SeqCompactSpace.tendsto_subseq [SeqCompactSpace X] (x : ℕ → X) : ∃ (a : X) (φ : ℕ → ℕ), StrictMono φ ∧ Tendsto (x ∘ φ) atTop (𝓝 a) := - let ⟨a, _, φ, mono, h⟩ := seq_compact_univ fun n => mem_univ (x n) + let ⟨a, _, φ, mono, h⟩ := isSeqCompact_univ fun n => mem_univ (x n) ⟨a, φ, mono, h⟩ section FirstCountableTopology @@ -277,6 +277,27 @@ theorem CompactSpace.tendsto_subseq [CompactSpace X] (x : ℕ → X) : end FirstCountableTopology +section Image + +variable [TopologicalSpace Y] {f : X → Y} + +/-- Sequential compactness of sets is preserved under sequentially continuous functions. -/ +theorem IsSeqCompact.image (f_cont : SeqContinuous f) {K : Set X} (K_cpt : IsSeqCompact K) : + IsSeqCompact (f '' K) := by + intro ys ys_in_fK + choose xs xs_in_K fxs_eq_ys using ys_in_fK + obtain ⟨a, a_in_K, phi, phi_mono, xs_phi_lim⟩ := K_cpt xs_in_K + refine ⟨f a, mem_image_of_mem f a_in_K, phi, phi_mono, ?_⟩ + exact (f_cont xs_phi_lim).congr fun x ↦ fxs_eq_ys (phi x) + +/-- The range of sequentially continuous function on a sequentially compact space is sequentially +compact. -/ +theorem IsSeqCompact.range [SeqCompactSpace X] (f_cont : SeqContinuous f) : + IsSeqCompact (Set.range f) := by + simpa using isSeqCompact_univ.image f_cont + +end Image + end SeqCompact section UniformSpaceSeqCompact diff --git a/Mathlib/Topology/Sets/Closeds.lean b/Mathlib/Topology/Sets/Closeds.lean index 91b42ecb3064a..6c7fd6bd1046a 100644 --- a/Mathlib/Topology/Sets/Closeds.lean +++ b/Mathlib/Topology/Sets/Closeds.lean @@ -162,11 +162,12 @@ theorem iInf_mk {ι} (s : ι → Set α) (h : ∀ i, IsClosed (s i)) : (⨅ i, ⟨s i, h i⟩ : Closeds α) = ⟨⋂ i, s i, isClosed_iInter h⟩ := iInf_def _ -instance : Coframe (Closeds α) := - { inferInstanceAs (CompleteLattice (Closeds α)) with - sInf := sInf - iInf_sup_le_sup_sInf := fun a s => - (SetLike.coe_injective <| by simp only [coe_sup, coe_iInf, coe_sInf, Set.union_iInter₂]).le } +/-- Closed sets in a topological space form a coframe. -/ +def coframeMinimalAxioms : Coframe.MinimalAxioms (Closeds α) where + iInf_sup_le_sup_sInf a s := + (SetLike.coe_injective <| by simp only [coe_sup, coe_iInf, coe_sInf, Set.union_iInter₂]).le + +instance instCoframe : Coframe (Closeds α) := .ofMinimalAxioms coframeMinimalAxioms /-- The term of `TopologicalSpace.Closeds α` corresponding to a singleton. -/ @[simps] @@ -305,6 +306,7 @@ instance : HasCompl (Clopens α) := ⟨fun s => ⟨sᶜ, s.isClopen.compl⟩⟩ instance : BooleanAlgebra (Clopens α) := SetLike.coe_injective.booleanAlgebra _ coe_sup coe_inf coe_top coe_bot coe_compl coe_sdiff + coe_himp instance : Inhabited (Clopens α) := ⟨⊥⟩ diff --git a/Mathlib/Topology/Sets/Compacts.lean b/Mathlib/Topology/Sets/Compacts.lean index cbeacf795f0aa..ca665e3599138 100644 --- a/Mathlib/Topology/Sets/Compacts.lean +++ b/Mathlib/Topology/Sets/Compacts.lean @@ -509,6 +509,7 @@ instance instHImp : HImp (CompactOpens α) where instance instBooleanAlgebra : BooleanAlgebra (CompactOpens α) := SetLike.coe_injective.booleanAlgebra _ coe_sup coe_inf coe_top coe_bot coe_compl coe_sdiff + coe_himp end Top.Compl diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 99afe5bc909bc..fad80ae1e1cf9 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -226,11 +226,12 @@ theorem mem_iSup {ι} {x : α} {s : ι → Opens α} : x ∈ iSup s ↔ ∃ i, x theorem mem_sSup {Us : Set (Opens α)} {x : α} : x ∈ sSup Us ↔ ∃ u ∈ Us, x ∈ u := by simp_rw [sSup_eq_iSup, mem_iSup, exists_prop] -instance : Frame (Opens α) := - { inferInstanceAs (CompleteLattice (Opens α)) with - sSup := sSup - inf_sSup_le_iSup_inf := fun a s => - (ext <| by simp only [coe_inf, coe_iSup, coe_sSup, Set.inter_iUnion₂]).le } +/-- Open sets in a topological space form a frame. -/ +def frameMinimalAxioms : Frame.MinimalAxioms (Opens α) where + inf_sSup_le_iSup_inf a s := + (ext <| by simp only [coe_inf, coe_iSup, coe_sSup, Set.inter_iUnion₂]).le + +instance instFrame : Frame (Opens α) := .ofMinimalAxioms frameMinimalAxioms theorem openEmbedding' (U : Opens α) : OpenEmbedding (Subtype.val : U → α) := U.isOpen.openEmbedding_subtype_val @@ -305,6 +306,12 @@ theorem IsBasis.isCompact_open_iff_eq_finite_iUnion {ι : Type*} (b : ι → Ope simp · exact hb' +lemma IsBasis.le_iff {α} {t₁ t₂ : TopologicalSpace α} + {Us : Set (Opens α)} (hUs : @IsBasis α t₂ Us) : + t₁ ≤ t₂ ↔ ∀ U ∈ Us, IsOpen[t₁] U := by + conv_lhs => rw [hUs.eq_generateFrom] + simp [Set.subset_def, le_generateFrom_iff_subset_isOpen] + @[simp] theorem isCompactElement_iff (s : Opens α) : CompleteLattice.IsCompactElement s ↔ IsCompact (s : Set α) := by diff --git a/Mathlib/Topology/Tactic.lean b/Mathlib/Topology/Tactic.lean deleted file mode 100644 index da80eba4f18fd..0000000000000 --- a/Mathlib/Topology/Tactic.lean +++ /dev/null @@ -1,18 +0,0 @@ -/- -Copyright (c) 2020 Reid Barton. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Reid Barton --/ -import Mathlib.Topology.Basic -import Mathlib.Tactic.Continuity - -/-! -# Tactics for topology - -Currently we have one domain-specific tactic for topology: `continuity`. -It is implemented in `Mathlib.Tactic.Continuity`. - -Porting note: the sole purpose of this file is to mark it as "ported". -This file seems to be tripping up the porting dashboard. - --/ diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index ae7448c91bfa5..c85211f14a68a 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -90,7 +90,7 @@ private theorem symm_gen : map Prod.swap ((𝓤 α).lift' gen) ≤ (𝓤 α).lif (monotone_setOf fun p => @Filter.monotone_mem _ (p.2.val ×ˢ p.1.val))) (by have h := fun p : CauchyFilter α × CauchyFilter α => @Filter.prod_comm _ _ p.2.val p.1.val - simp [f, Function.comp, h, mem_map'] + simp only [Function.comp, h, mem_map, f] exact le_rfl) exact h₁.trans_le h₂ @@ -210,7 +210,7 @@ instance : CompleteSpace (CauchyFilter α) := have : t' ⊆ { y : α | (f', pureCauchy y) ∈ gen t } := fun x hx => (f ×ˢ pure x).sets_of_superset (prod_mem_prod ht' hx) h f.sets_of_superset ht' <| Subset.trans this (preimage_mono ht₂) - ⟨f', by simp [nhds_eq_uniformity]; assumption⟩ + ⟨f', by simpa [nhds_eq_uniformity]⟩ end diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index f0216a540c3e2..d5f3015592180 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Nicolò Cavalleri. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri, Sebastien Gouezel, Heather Macbeth, Patrick Massot, Floris van Doorn -/ -import Mathlib.Analysis.NormedSpace.BoundedLinearMaps +import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps import Mathlib.Topology.FiberBundle.Basic /-! diff --git a/Mathlib/Util/AssertNoSorry.lean b/Mathlib/Util/AssertNoSorry.lean index 49987ad8dade9..e29f256c1db8b 100644 --- a/Mathlib/Util/AssertNoSorry.lean +++ b/Mathlib/Util/AssertNoSorry.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Renshaw -/ -import Lean.Elab.Print -import Lean.Environment +import Lean.Util.CollectAxioms +import Lean.Elab.Command /-! # Defines the `assert_no_sorry` command. @@ -13,12 +13,11 @@ import Lean.Environment Throws an error if the given identifier uses sorryAx. -/ -open Lean.Elab.Command +open Lean Meta Elab Command /-- Throws an error if the given identifier uses sorryAx. -/ elab "assert_no_sorry " n:ident : command => do - let env ← Lean.getEnv - let (_, s) := ((Lean.Elab.Command.CollectAxioms.collect - (← liftCoreM <| Lean.Elab.realizeGlobalConstNoOverloadWithInfo n)).run env).run {} - if s.axioms.contains ``sorryAx + let name ← liftCoreM <| Lean.Elab.realizeGlobalConstNoOverloadWithInfo n + let axioms ← Lean.collectAxioms name + if axioms.contains ``sorryAx then throwError "{n} contains sorry" diff --git a/docs/references.bib b/docs/references.bib index bf45067ce36ff..94df40e8bf081 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -666,6 +666,17 @@ @Book{ calugareanu doi = {10.1007/978-94-015-9588-9} } +@Article{ caramello2020, + title = {Denseness conditions, morphisms and equivalences of + toposes}, + author = {Olivia Caramello}, + year = {2020}, + eprint = {1906.08737}, + archiveprefix = {arXiv}, + primaryclass = {math.CT}, + url = {https://arxiv.org/abs/1906.08737} +} + @Article{ CARBONI1993145, author = {Aurelio Carboni and Stephen Lack and R.F.C. Walters}, doi = {https://doi.org/10.1016/0022-4049(93)90035-R}, @@ -2176,6 +2187,14 @@ @Article{ kelleyVaught1953 doi = {10.2307/1990847} } +@Book{ keng1982house, + author = {Keng, Hua Loo}, + publisher = {Springer}, + title = {Introduction to Number Theory}, + year = {1982}, + pages = {489} +} + @Article{ kleiman1979, author = {Kleiman, Steven Lawrence}, title = {Misconceptions about {$K_X$}}, @@ -2237,6 +2256,19 @@ @Article{ kozen1994 over Kleene algebras.} } +@Article{ lam_1984, + author = {T.Y. Lam}, + title = {{An introduction to real algebra}}, + volume = {14}, + journal = {Rocky Mountain Journal of Mathematics}, + number = {4}, + publisher = {Rocky Mountain Mathematics Consortium}, + pages = {767 -- 814}, + year = {1984}, + doi = {10.1216/RMJ-1984-14-4-767}, + url = {https://doi.org/10.1216/RMJ-1984-14-4-767} +} + @Book{ lam_1999, series = {Graduate Texts in Mathematics}, title = {Lectures on Modules and Rings}, @@ -2634,6 +2666,19 @@ @Book{ Neukirch1992 doi = {10.1007/978-3-662-03983-0} } +@Article{ Neumann-1954, + author = {Neumann, B. H.}, + title = {Groups Covered By Permutable Subsets}, + year = {1954}, + month = apr, + journal = {Journal of the London Mathematical Society}, + volume = {s1-29}, + number = {2}, + pages = {236--248}, + issn = {00246107}, + doi = {10.1112/jlms/s1-29.2.236} +} + @Book{ Okninski1991, place = {New York}, title = {Semigroup algebras}, diff --git a/lake-manifest.json b/lake-manifest.json index ce5b2cbc55af7..58738ed3f39ee 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "983852f1637ef7fc35e7da36e68495e1060ff05c", + "rev": "56bf1ca13e72d10f69fd0e62f04033123b871932", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing-2024-07-23", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "622d52c803db99ff4ea4fb442c1db9e91aed944c", + "rev": "79fb157c6a5061190d169535f8e5cb007914a82e", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -35,10 +35,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "d1b33202c3a29a079f292de65ea438648123b635", + "rev": "c87908619cccadda23f71262e6898b9893bffa36", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.39", + "inputRev": "v0.0.40", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index d380cb361f553..3a31a99334ed5 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -21,7 +21,7 @@ package mathlib where require "leanprover-community" / "batteries" @ "git#nightly-testing-2024-07-23" require "leanprover-community" / "Qq" @ "git#master" require "leanprover-community" / "aesop" @ "git#master" -require "leanprover-community" / "proofwidgets" @ "git#v0.0.39" +require "leanprover-community" / "proofwidgets" @ "git#v0.0.40" require "leanprover-community" / "importGraph" @ "git#nightly-testing" /-! diff --git a/scripts/declarations_diff.sh b/scripts/declarations_diff.sh index 2b2c7fbfcb154..bde0906978c26 100755 --- a/scripts/declarations_diff.sh +++ b/scripts/declarations_diff.sh @@ -1,5 +1,46 @@ #!/usr/bin/env bash + : <<'BASH_DOC_MODULE' + +# The `declarations_diff` script + +The `declarations_diff` script is a text-based script that attempts to find which declarations +have been removed and which declarations have been added in the current PR with respect to `master`. + +In essence, it looks at the output of `git diff origin/master...HEAD`, extracts the lines that +contain one of the keywords + +`theorem` `lemma` `inductive` `structure` `def` `class` `instance` `alias` + +and tries to find exact matches between a removed line and an added line +(e.g. when moving a declaration from one file to another or restructuring within the same file). + +Exact matches are removed. +Among the remaining unmatched lines, the script tries to extract the declaration ids and tries to +find exact matches among those. +No effort is made to keep track of namespaces: if the declaration is `theorem abc ...`, then +`abc` is extracted, even if this happens inside `namespace xyz`. +If a declaration id is added exactly once *and* removed exactly once, +then they are again considered paired and are not shown. + +If a declaration id is either only added or removed or it is added or removed more than once, then +the script will return a count such as + +++--+ thmName + +This means that a declaration `thmName` was added 3 times and removed twice: this can happen +with namespacing, e.g. you could see it with `++-- map_zero`. + +The script uses some heuristics to guide this process. +* It assumes that each keyword (as above) appears on the same line as the corresponding + declaration id --- a line break between `theorem` and `riemannHypothesis` fools the script. +* It deals with declaration modifiers (such as `noncomputable`, `nonrec`, `protected`) and + attributes. +* It is "aware" of "nameless" `instance`s and, rather than looking for a declaration id, + in this case records the whole line `instance ...`. + +BASH_DOC_MODULE + ## we narrow the diff to lines beginning with `theorem`, `lemma` and a few other commands begs="(theorem|lemma|inductive|structure|def|class|instance|alias)" @@ -123,7 +164,9 @@ printf $'
## more verbose report: ./scripts/declarations_diff.sh long ``` -
' + +The doc-module for `script/declarations_diff.sh` contains some details about this script.' + : <` to increase limit use `set_option diagnostics true` to get diagnostic information -/ #guard_msgs (error) in -example (_h₁ : Function.Injective (funFamily ((List.range 128).map (fun _ => 0)).sum)) : true = true := by +example (_h₁ : Function.Injective (funFamily ((List.range 128).map (fun _ => 0)).sum)) : + true = true := by apply_fun funFamily 0 diff --git a/test/borelize.lean b/test/borelize.lean index b82917535d407..b1c24d7ed7ce7 100644 --- a/test/borelize.lean +++ b/test/borelize.lean @@ -2,7 +2,8 @@ import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic set_option autoImplicit true -example [TopologicalSpace α] [inst : MeasurableSpace α] [BorelSpace α] : MeasurableSet (∅ : Set α) := by +example [TopologicalSpace α] [inst : MeasurableSpace α] [BorelSpace α] : + MeasurableSet (∅ : Set α) := by guard_target = @MeasurableSet α inst ∅ borelize α guard_target = @MeasurableSet α (borel α) ∅ diff --git a/test/cancel_denoms.lean b/test/cancel_denoms.lean index a7592132b3c3c..b717b67f56ee2 100644 --- a/test/cancel_denoms.lean +++ b/test/cancel_denoms.lean @@ -89,7 +89,8 @@ example (h : a + b = c) : a/5 + d*(b/4) = c - 4*a/5 + b*2*d/8 - b := by rw [← h] ring -example (h : 2 * (4 * a + d * 5 * b) ≠ (40 * c - 32 * a + b * 2 * 5 * d - 40 * b)) : a/5 + d*(b/4) ≠ c - 4*a/5 + b*2*d/8 - b := by +example (h : 2 * (4 * a + d * 5 * b) ≠ (40 * c - 32 * a + b * 2 * 5 * d - 40 * b)) : + a/5 + d*(b/4) ≠ c - 4*a/5 + b*2*d/8 - b := by cancel_denoms assumption diff --git a/test/congr.lean b/test/congr.lean index 9f5f069d12235..7dd402f4cb5dc 100644 --- a/test/congr.lean +++ b/test/congr.lean @@ -339,6 +339,7 @@ x : α ⊢ inst1 = inst2 -/ #guard_msgs in -example {α : Type} (inst1 : BEq α) [LawfulBEq α] (inst2 : BEq α) [LawfulBEq α] (xs : List α) (x : α) : +example + {α : Type} (inst1 : BEq α) [LawfulBEq α] (inst2 : BEq α) [LawfulBEq α] (xs : List α) (x : α) : @List.erase _ inst1 xs x = @List.erase _ inst2 xs x := by congr! (config := { beqEq := false }) diff --git a/test/fun_prop.lean b/test/fun_prop.lean index b05fa6835bf66..5b7a7d504d31b 100644 --- a/test/fun_prop.lean +++ b/test/fun_prop.lean @@ -33,7 +33,8 @@ It is best to start with the basic lambda calculus rules. There are five of thes - pi rule `∀ i, Measurable (f · i) → Measurable fun x i => f x i` You do not have to provide them all. For example `IsLinearMap` does not have the constant rule. -However, to have any hope at using `fun_prop` successfully you need to at least provide identity and composition rule. +However, to have any hope at using `fun_prop` successfully you need to at least provide identity +and composition rule. -/ attribute [fun_prop] @@ -44,7 +45,8 @@ attribute [fun_prop] measurable_pi_lambda /-! -Measurability also behaves nicely with respect to taking products. Let's mark the product constructor +Measurability also behaves nicely with respect to taking products. +Let's mark the product constructor. -/ attribute [fun_prop] @@ -59,7 +61,8 @@ or ``` Measurable.fst : Measurable f → Measurable fun x => Prod.fst (f x) ``` -Tactic `fun_prop` can work with both versions and it should be sufficient to provide just one of them. +The tactic `fun_prop` can work with both versions; +it should be sufficient to provide just one of them. It does not hurt to provide both of them though. -/ @@ -70,7 +73,8 @@ attribute [fun_prop] Measurable.snd /-! -A silly example on which `measurability` fails and `fun_prop` succeeds. Let's turn on tracing to see what is going on +A silly example on which `measurability` fails and `fun_prop` succeeds. Let's turn on tracing +to see what is going on set_option trace.Meta.Tactic.fun_prop true in -/ example {α} [MeasurableSpace α] (f : α → α → α) (hf : Measurable fun (x,y) => f x y) (a : α) : @@ -82,8 +86,9 @@ example {α} [MeasurableSpace α] (f : α → α → α) (hf : Measurable fun (x /-! To give more complicated examples we mark theorems about arithmetic operations with `@[fun_prop]` -Again we mark both versions of theorems. Internally `fun_prop` says that theorems like `measurable_add` -are in "uncurried form" and theorems like `Measurable.add` are in compositional form. +Again we mark both versions of theorems. Internally `fun_prop` says that theorems like +`measurable_add` are in "uncurried form" and theorems like `Measurable.add` are in compositional +form. -/ attribute [fun_prop] @@ -179,9 +184,11 @@ measurability of `DFunLike.coe` in `f` and `x` separately. The theorem `ContinuousLinearMap.measurable` states measurability in `x` in uncurried form. The theorem `ContinuousLinearMap.measurable_comp` states measurability in `x` in compositional form. The theorem `ContinuousLinearMap.measurable_apply` states measurability in `f` in uncurried form. -The theorem `Measurable.apply_continuousLinearMap` states measurability in `f` in compositional form. +The theorem `Measurable.apply_continuousLinearMap` states measurability in `f` in compositional +form. -/ +set_option linter.longLine false in attribute [fun_prop] ContinuousLinearMap.measurable -- Measurable fun (x : E) => DFunLike.coe L x ContinuousLinearMap.measurable_comp -- Measurable φ → Measurable fun (x : E) => DFunLike.coe L (φ x) @@ -196,13 +203,14 @@ A silly example that everything together works as expected example (f : ℝ → ℝ → (ℝ →L[ℝ] ℝ)) (hf : Continuous (fun (x,y) => f x y)) : Measurable fun x => (f (x / x) (x * x) 1 + x) := by fun_prop +set_option linter.longLine false in /-! In the current state of `fun_prop`, morphism theorems **have to** be stated in compositional form. Sometimes they might work in uncurried form but `fun_prop` is not designed that way right now. -In other cases the function property of `DFunLike.coe` can be stated jointly in `f` and `x`. This is the case of `ContDiff n` and -continuous linear maps. The theorem is `ContDiff.clm_apply`. +In other cases the function property of `DFunLike.coe` can be stated jointly in `f` and `x`. +This is the case of `ContDiff n` and continuous linear maps. The theorem is `ContDiff.clm_apply`. #check ContDiff.clm_apply -- {f : E → F →L[K] G} → {g : E → F} → ContDiff K n f → ContDiff K n g → ContDiff K n fun x => DFunLike.coe (f x) (g x) diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index 07e11712895b1..bd1fad85f4e2b 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -3,13 +3,13 @@ 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.GroupWithZero.Action.Prod +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Polynomial.Basic -import Mathlib.GroupTheory.GroupAction.Prod -import Mathlib.GroupTheory.GroupAction.Units import Mathlib.Data.Complex.Module -import Mathlib.RingTheory.Algebraic import Mathlib.Data.ZMod.Basic +import Mathlib.RingTheory.Algebraic import Mathlib.RingTheory.TensorProduct.Basic /-! # Tests that instances do not form diamonds -/ diff --git a/test/jacobiSym.lean b/test/jacobiSym.lean index bfc8b7c9deff4..916deb2aa2d2d 100644 --- a/test/jacobiSym.lean +++ b/test/jacobiSym.lean @@ -43,6 +43,6 @@ warning: declaration uses 'sorry' info: 1 -/ #guard_msgs in -#eval @legendreSym (2 ^ 11213 - 1) sorry 7 +#eval! @legendreSym (2 ^ 11213 - 1) sorry 7 end Csimp diff --git a/test/lift.lean b/test/lift.lean index 01251808b25d4..a8c84bce02c83 100644 --- a/test/lift.lean +++ b/test/lift.lean @@ -110,7 +110,8 @@ example (n : ℤ) : ℕ := by fail_if_success lift n to ℕ exact 0 -instance canLift_subtype (R : Type _) (s : Set R) : CanLift R {x // x ∈ s} ((↑) : {x // x ∈ s} → R) (fun x => x ∈ s) := +instance canLift_subtype (R : Type _) (s : Set R) : + CanLift R {x // x ∈ s} ((↑) : {x // x ∈ s} → R) (fun x => x ∈ s) := { prf := fun x hx => ⟨⟨x, hx⟩, rfl⟩ } example {R : Type _} {P : R → Prop} (x : R) (hx : P x) : P x := by diff --git a/test/linarith.lean b/test/linarith.lean index 003b2a3bc151a..628c2da86291d 100644 --- a/test/linarith.lean +++ b/test/linarith.lean @@ -189,7 +189,8 @@ example (a b c : Rat) (h2 : (2 : Rat) > 3) : a + b - c ≥ 3 := by -- Verify that we split conjunctions in hypotheses. example (x y : Rat) - (h : 6 + ((x + 4) * x + (6 + 3 * y) * y) = 3 ∧ (x + 4) * x ≥ 0 ∧ (6 + 3 * y) * y ≥ 0) : False := by + (h : 6 + ((x + 4) * x + (6 + 3 * y) * y) = 3 ∧ (x + 4) * x ≥ 0 ∧ (6 + 3 * y) * y ≥ 0) : + False := by fail_if_success linarith (config := {splitHypotheses := false}) linarith diff --git a/test/measurability.lean b/test/measurability.lean index 59ad11260bfcf..2ebf1a799ee68 100644 --- a/test/measurability.lean +++ b/test/measurability.lean @@ -3,7 +3,6 @@ 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.Tactic import Mathlib.MeasureTheory.MeasurableSpace.Basic import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic import Mathlib.MeasureTheory.Function.SpecialFunctions.Basic @@ -13,7 +12,7 @@ import Mathlib.MeasureTheory.Function.StronglyMeasurable.Lemmas open MeasureTheory TopologicalSpace -variable {α β : Type _} [MeasurableSpace α] [MeasurableSpace β] +variable {α β : Type*} [MeasurableSpace α] [MeasurableSpace β] {f g : α → β} {s₁ s₂ : Set α} {t₁ t₂ : Set β} {μ ν : MeasureTheory.Measure α} set_option linter.unusedVariables false diff --git a/test/polyrith.lean b/test/polyrith.lean index 3f330426aeb46..a0bd8b137e1f3 100644 --- a/test/polyrith.lean +++ b/test/polyrith.lean @@ -5,6 +5,10 @@ Authors: Dhruv Bhatia, Robert Y. Lewis, Mario Carneiro -/ import Mathlib.Tactic.Polyrith +-- Except for the `import`, the doc-modules and the following `set_option`, this file is just +-- comments and whitespace. Once the file gets revived, the linting can start! +set_option linter.longLine false + /-! Each call to `polyrith` makes a call to the SageCell web API at diff --git a/test/positivity.lean b/test/positivity.lean index 2b7bd14601d3f..d75fda862bd69 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -95,7 +95,8 @@ end -- example [Nonempty ι] [Zero α] {a : α} (ha : a ≠ 0) : const ι a ≠ 0 := by positivity -- example [Zero α] [PartialOrder α] {a : α} (ha : 0 < a) : 0 ≤ const ι a := by positivity -- example [Zero α] [PartialOrder α] {a : α} (ha : 0 ≤ a) : 0 ≤ const ι a := by positivity --- example [Nonempty ι] [Zero α] [PartialOrder α] {a : α} (ha : 0 < a) : 0 < const ι a := by positivity +-- example [Nonempty ι] [Zero α] [PartialOrder α] {a : α} (ha : 0 < a) : 0 < const ι a := by +-- positivity section ite variable {p : Prop} [Decidable p] {a b : ℤ} @@ -260,8 +261,8 @@ example {a : ℝ} (ha : 0 ≤ a) : 0 < Real.sqrt (a + 3) := by positivity example {a b : ℤ} (ha : 3 < a) : 0 ≤ min a (b ^ 2) := by positivity --- -- test that the tactic can ignore arithmetic operations whose associated extension tactic requires --- -- more typeclass assumptions than are available +-- -- test that the tactic can ignore arithmetic operations whose associated extension tactic +-- -- requires more typeclass assumptions than are available -- example {R : Type _} [Zero R] [Div R] [LinearOrder R] {a b c : R} (h1 : 0 < a) (h2 : 0 < b) -- (h3 : 0 < c) : -- 0 < max (a / b) c := diff --git a/test/push_neg.lean b/test/push_neg.lean index 5442170e836c8..1f1a0c4f85f39 100644 --- a/test/push_neg.lean +++ b/test/push_neg.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Alice Laroche. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Author: Alice Laroche, Frédéric Dupuis, Jireh Loreaux +Authors: Alice Laroche, Frédéric Dupuis, Jireh Loreaux -/ import Mathlib.Tactic.PushNeg