From 43881b58cc09b474a672296c7252e001840ac542 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 7 May 2024 08:14:36 +0000 Subject: [PATCH 01/52] chore: Move sign of power lemmas (#11986) * Move the sign of power lemmas from `Algebra.Parity` to `Algebra.GroupPower.Order`. * For this to work, I must swap the order of import between `Algebra.GroupPower.Order` and `Algebra.Parity`. This means that I need to weaken one `assert_not_exists` to allow importing `Data.Set.Defs`. This is inconsequential. * Use them to golf and deprecate the `bit0`/`bit1` lemmas * Delete the deprecated `pow_bit0_abs`, `pow_bit0_pos_of_neg`, `pow_bit1_neg` Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Algebra/GroupPower/Order.lean | 170 +++++++++++------- Mathlib/Algebra/Order/Ring/Abs.lean | 4 + Mathlib/Algebra/Parity.lean | 105 +---------- .../Analysis/Calculus/FDeriv/Symmetric.lean | 2 +- Mathlib/Data/Rat/Sqrt.lean | 1 - .../Order/SuccPred/LinearLocallyFinite.lean | 6 +- 6 files changed, 115 insertions(+), 173 deletions(-) diff --git a/Mathlib/Algebra/GroupPower/Order.lean b/Mathlib/Algebra/GroupPower/Order.lean index a9390bb31067a..94097bb8eef7b 100644 --- a/Mathlib/Algebra/GroupPower/Order.lean +++ b/Mathlib/Algebra/GroupPower/Order.lean @@ -6,17 +6,16 @@ Authors: Jeremy Avigad, Robert Y. Lewis import Mathlib.Algebra.GroupPower.CovariantClass import Mathlib.Algebra.GroupPower.Ring import Mathlib.Algebra.Order.Ring.Canonical +import Mathlib.Algebra.Parity #align_import algebra.group_power.order from "leanprover-community/mathlib"@"00f91228655eecdcd3ac97a7fd8dbcb139fe990a" /-! # Lemmas about the interaction of power operations with order - -Note that some lemmas are in `Algebra/GroupPower/Lemmas.lean` as they import files which -depend on this file. -/ -assert_not_exists Set.range +-- We should need only a minimal development of sets in order to get here. +assert_not_exists Set.Subsingleton open Function Int @@ -298,22 +297,9 @@ theorem sq_pos_of_pos (ha : 0 < a) : 0 < a ^ 2 := pow_pos ha _ end StrictOrderedSemiring section StrictOrderedRing -set_option linter.deprecated false - variable [StrictOrderedRing R] {a : R} -theorem pow_bit0_pos_of_neg (ha : a < 0) (n : ℕ) : 0 < a ^ bit0 n := by - rw [pow_bit0'] - exact pow_pos (mul_pos_of_neg_of_neg ha ha) _ -#align pow_bit0_pos_of_neg pow_bit0_pos_of_neg - -theorem pow_bit1_neg (ha : a < 0) (n : ℕ) : a ^ bit1 n < 0 := by - rw [bit1, pow_succ'] - exact mul_neg_of_neg_of_pos ha (pow_bit0_pos_of_neg ha n) -#align pow_bit1_neg pow_bit1_neg - -theorem sq_pos_of_neg (ha : a < 0) : 0 < a ^ 2 := - pow_bit0_pos_of_neg ha 1 +lemma sq_pos_of_neg (ha : a < 0) : 0 < a ^ 2 := by rw [sq]; exact mul_pos_of_neg_of_neg ha ha #align sq_pos_of_neg sq_pos_of_neg end StrictOrderedRing @@ -394,6 +380,16 @@ theorem lt_of_mul_self_lt_mul_self (hb : 0 ≤ b) : a * a < b * b → a < b := b exact lt_of_pow_lt_pow_left _ hb #align lt_of_mul_self_lt_mul_self lt_of_mul_self_lt_mul_self +/-! +### Lemmas for canonically linear ordered semirings or linear ordered rings + +The slightly unusual typeclass assumptions `[LinearOrderedSemiring R] [ExistsAddOfLE R]` cover two +more familiar settings: +* `[LinearOrderedRing R]`, eg `ℤ`, `ℚ` or `ℝ` +* `[CanonicallyLinearOrderedSemiring R]` (although we don't actually have this typeclass), eg `ℕ`, + `ℚ≥0` or `ℝ≥0` +-/ + variable [ExistsAddOfLE R] lemma add_sq_le : (a + b) ^ 2 ≤ 2 * (a ^ 2 + b ^ 2) := by @@ -425,11 +421,10 @@ lemma add_pow_le (ha : 0 ≤ a) (hb : 0 ≤ b) : ∀ n, (a + b) ^ n ≤ 2 ^ (n - · exact mul_add_mul_le_mul_add_mul (pow_le_pow_left ha hab _) hab · exact mul_add_mul_le_mul_add_mul' (pow_le_pow_left hb hba _) hba --- TODO: State using `Even` -protected lemma Even.add_pow_le (hn : ∃ k, 2 * k = n) : +protected lemma Even.add_pow_le (hn : Even n) : (a + b) ^ n ≤ 2 ^ (n - 1) * (a ^ n + b ^ n) := by obtain ⟨n, rfl⟩ := hn - rw [pow_mul] + rw [← two_mul, pow_mul] calc _ ≤ (2 * (a ^ 2 + b ^ 2)) ^ n := pow_le_pow_left (sq_nonneg _) add_sq_le _ _ = 2 ^ n * (a ^ 2 + b ^ 2) ^ n := by -- TODO: Should be `Nat.cast_commute` @@ -442,76 +437,113 @@ protected lemma Even.add_pow_le (hn : ∃ k, 2 * k = n) : · rfl · simp [Nat.two_mul] -end LinearOrderedSemiring +lemma Even.pow_nonneg (hn : Even n) (a : R) : 0 ≤ a ^ n := by + obtain ⟨k, rfl⟩ := hn; rw [pow_add]; exact mul_self_nonneg _ +#align even.pow_nonneg Even.pow_nonneg + +lemma Even.pow_pos (hn : Even n) (ha : a ≠ 0) : 0 < a ^ n := + (hn.pow_nonneg _).lt_of_ne' (pow_ne_zero _ ha) +#align even.pow_pos Even.pow_pos + +lemma Even.pow_pos_iff (hn : Even n) (h₀ : n ≠ 0) : 0 < a ^ n ↔ a ≠ 0 := by + obtain ⟨k, rfl⟩ := hn; rw [pow_add, mul_self_pos (α := R), pow_ne_zero_iff (by simpa using h₀)] +#align even.pow_pos_iff Even.pow_pos_iff + +lemma Odd.pow_neg_iff (hn : Odd n) : a ^ n < 0 ↔ a < 0 := by + refine ⟨lt_imp_lt_of_le_imp_le (pow_nonneg · _), fun ha ↦ ?_⟩ + obtain ⟨k, rfl⟩ := hn + rw [pow_succ] + exact mul_neg_of_pos_of_neg ((even_two_mul _).pow_pos ha.ne) ha +#align odd.pow_neg_iff Odd.pow_neg_iff -section LinearOrderedRing -variable [LinearOrderedRing R] {a b : R} {n : ℕ} +lemma Odd.pow_nonneg_iff (hn : Odd n) : 0 ≤ a ^ n ↔ 0 ≤ a := + le_iff_le_iff_lt_iff_lt.2 hn.pow_neg_iff +#align odd.pow_nonneg_iff Odd.pow_nonneg_iff + +lemma Odd.pow_nonpos_iff (hn : Odd n) : a ^ n ≤ 0 ↔ a ≤ 0 := by + rw [le_iff_lt_or_eq, le_iff_lt_or_eq, hn.pow_neg_iff, pow_eq_zero_iff] + rintro rfl; simp [Odd, eq_comm (a := 0)] at hn +#align odd.pow_nonpos_iff Odd.pow_nonpos_iff + +lemma Odd.pow_pos_iff (hn : Odd n) : 0 < a ^ n ↔ 0 < a := lt_iff_lt_of_le_iff_le hn.pow_nonpos_iff +#align odd.pow_pos_iff Odd.pow_pos_iff + +alias ⟨_, Odd.pow_nonpos⟩ := Odd.pow_nonpos_iff +alias ⟨_, Odd.pow_neg⟩ := Odd.pow_neg_iff +#align odd.pow_nonpos Odd.pow_nonpos +#align odd.pow_neg Odd.pow_neg + +lemma Odd.strictMono_pow (hn : Odd n) : StrictMono fun a : R => a ^ n := by + have hn₀ : n ≠ 0 := by rintro rfl; simp [Odd, eq_comm (a := 0)] at hn + intro a b hab + obtain ha | ha := le_total 0 a + · exact pow_lt_pow_left hab ha hn₀ + obtain hb | hb := lt_or_le 0 b + · exact (hn.pow_nonpos ha).trans_lt (pow_pos hb _) + obtain ⟨c, hac⟩ := exists_add_of_le ha + obtain ⟨d, hbd⟩ := exists_add_of_le hb + have hd := nonneg_of_le_add_right (hb.trans_eq hbd) + refine lt_of_add_lt_add_right (a := c ^ n + d ^ n) ?_ + dsimp + calc + a ^ n + (c ^ n + d ^ n) = d ^ n := by + rw [← add_assoc, hn.pow_add_pow_eq_zero hac.symm, zero_add] + _ < c ^ n := pow_lt_pow_left ?_ hd hn₀ + _ = b ^ n + (c ^ n + d ^ n) := by rw [add_left_comm, hn.pow_add_pow_eq_zero hbd.symm, add_zero] + refine lt_of_add_lt_add_right (a := a + b) ?_ + rwa [add_rotate', ← hbd, add_zero, add_left_comm, ← add_assoc, ← hac, zero_add] +#align odd.strict_mono_pow Odd.strictMono_pow + +lemma sq_pos_iff {a : R} : 0 < a ^ 2 ↔ a ≠ 0 := even_two.pow_pos_iff two_ne_zero +#align sq_pos_iff sq_pos_iff + +alias ⟨_, sq_pos_of_ne_zero⟩ := sq_pos_iff +alias pow_two_pos_of_ne_zero := sq_pos_of_ne_zero +#align sq_pos_of_ne_zero sq_pos_of_ne_zero +#align pow_two_pos_of_ne_zero pow_two_pos_of_ne_zero + +lemma pow_four_le_pow_two_of_pow_two_le (h : a ^ 2 ≤ b) : a ^ 4 ≤ b ^ 2 := + (pow_mul a 2 2).symm ▸ pow_le_pow_left (sq_nonneg a) h 2 +#align pow_four_le_pow_two_of_pow_two_le pow_four_le_pow_two_of_pow_two_le section deprecated set_option linter.deprecated false -theorem pow_bit0_nonneg (a : R) (n : ℕ) : 0 ≤ a ^ bit0 n := by - rw [pow_bit0] - exact mul_self_nonneg _ +@[deprecated Even.pow_nonneg] -- 2024-04-06 +lemma pow_bit0_nonneg (a : R) (n : ℕ) : 0 ≤ a ^ bit0 n := (even_bit0 _).pow_nonneg _ #align pow_bit0_nonneg pow_bit0_nonneg -theorem pow_bit0_pos {a : R} (h : a ≠ 0) (n : ℕ) : 0 < a ^ bit0 n := - (pow_bit0_nonneg a n).lt_of_ne (pow_ne_zero _ h).symm +@[deprecated Even.pow_pos] -- 2024-04-06 +lemma pow_bit0_pos {a : R} (h : a ≠ 0) (n : ℕ) : 0 < a ^ bit0 n := (even_bit0 _).pow_pos h #align pow_bit0_pos pow_bit0_pos -theorem pow_bit0_pos_iff (a : R) {n : ℕ} (hn : n ≠ 0) : 0 < a ^ bit0 n ↔ a ≠ 0 := by - refine' ⟨fun h => _, fun h => pow_bit0_pos h n⟩ - rintro rfl - rw [zero_pow (Nat.bit0_ne_zero hn)] at h - exact lt_irrefl _ h +@[deprecated Even.pow_pos_iff] -- 2024-04-06 +lemma pow_bit0_pos_iff (a : R) {n : ℕ} (hn : n ≠ 0) : 0 < a ^ bit0 n ↔ a ≠ 0 := + (even_bit0 _).pow_pos_iff (by simpa [bit0]) #align pow_bit0_pos_iff pow_bit0_pos_iff -@[simp] -lemma pow_bit1_neg_iff : a ^ bit1 n < 0 ↔ a < 0 := - ⟨fun h ↦ not_le.1 fun h' => not_le.2 h <| pow_nonneg h' _, fun ha ↦ pow_bit1_neg ha n⟩ +@[simp, deprecated Odd.pow_neg_iff] -- 2024-04-06 +lemma pow_bit1_neg_iff : a ^ bit1 n < 0 ↔ a < 0 := (odd_bit1 _).pow_neg_iff #align pow_bit1_neg_iff pow_bit1_neg_iff -@[simp] -lemma pow_bit1_nonneg_iff : 0 ≤ a ^ bit1 n ↔ 0 ≤ a := le_iff_le_iff_lt_iff_lt.2 pow_bit1_neg_iff +@[simp, deprecated Odd.pow_nonneg_iff] -- 2024-04-06 +lemma pow_bit1_nonneg_iff : 0 ≤ a ^ bit1 n ↔ 0 ≤ a := (odd_bit1 _).pow_nonneg_iff #align pow_bit1_nonneg_iff pow_bit1_nonneg_iff -@[simp] -lemma pow_bit1_nonpos_iff : a ^ bit1 n ≤ 0 ↔ a ≤ 0 := by - simp only [le_iff_lt_or_eq, pow_bit1_neg_iff, pow_eq_zero_iff']; simp [bit1] +@[simp, deprecated Odd.pow_nonpos_iff] -- 2024-04-06 +lemma pow_bit1_nonpos_iff : a ^ bit1 n ≤ 0 ↔ a ≤ 0 := (odd_bit1 _).pow_nonpos_iff #align pow_bit1_nonpos_iff pow_bit1_nonpos_iff -@[simp] -lemma pow_bit1_pos_iff : 0 < a ^ bit1 n ↔ 0 < a := lt_iff_lt_of_le_iff_le pow_bit1_nonpos_iff +@[simp, deprecated Odd.pow_pos_iff] -- 2024-04-06 +lemma pow_bit1_pos_iff : 0 < a ^ bit1 n ↔ 0 < a := (odd_bit1 _).pow_pos_iff #align pow_bit1_pos_iff pow_bit1_pos_iff -lemma strictMono_pow_bit1 (n : ℕ) : StrictMono (· ^ bit1 n : R → R) := by - intro a b hab - rcases le_total a 0 with ha | ha - · rcases le_or_lt b 0 with hb | hb - · rw [← neg_lt_neg_iff, ← neg_pow_bit1, ← neg_pow_bit1] - exact pow_lt_pow_left (neg_lt_neg hab) (neg_nonneg.2 hb) n.bit1_ne_zero - · exact (pow_bit1_nonpos_iff.2 ha).trans_lt (pow_bit1_pos_iff.2 hb) - · exact pow_lt_pow_left hab ha n.bit1_ne_zero +@[deprecated Odd.strictMono_pow] -- 2024-04-06 +lemma strictMono_pow_bit1 (n : ℕ) : StrictMono (· ^ bit1 n : R → R) := (odd_bit1 _).strictMono_pow #align strict_mono_pow_bit1 strictMono_pow_bit1 end deprecated - -lemma sq_pos_iff {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {a : R} : - 0 < a ^ 2 ↔ a ≠ 0 := by - rw [← pow_ne_zero_iff two_ne_zero, (sq_nonneg a).lt_iff_ne, ne_comm] -#align sq_pos_iff sq_pos_iff - -alias ⟨_, sq_pos_of_ne_zero⟩ := sq_pos_iff -#align sq_pos_of_ne_zero sq_pos_of_ne_zero - -alias pow_two_pos_of_ne_zero := sq_pos_of_ne_zero -#align pow_two_pos_of_ne_zero pow_two_pos_of_ne_zero - -lemma pow_four_le_pow_two_of_pow_two_le (h : a ^ 2 ≤ b) : a ^ 4 ≤ b ^ 2 := - (pow_mul a 2 2).symm ▸ pow_le_pow_left (sq_nonneg a) h 2 -#align pow_four_le_pow_two_of_pow_two_le pow_four_le_pow_two_of_pow_two_le - -end LinearOrderedRing +end LinearOrderedSemiring namespace MonoidHom diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index b503986c67085..e3efa2cd8cafa 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -63,6 +63,10 @@ lemma abs_pow (a : α) (n : ℕ) : |a ^ n| = |a| ^ n := (absHom.toMonoidHom : α lemma pow_abs (a : α) (n : ℕ) : |a| ^ n = |a ^ n| := (abs_pow a n).symm #align pow_abs pow_abs +lemma Even.pow_abs (hn : Even n) (a : α) : |a| ^ n = a ^ n := by + rw [← abs_pow, abs_eq_self]; exact hn.pow_nonneg _ +#align even.pow_abs Even.pow_abs + lemma abs_neg_one_pow (n : ℕ) : |(-1 : α) ^ n| = 1 := by rw [← pow_abs, abs_neg, abs_one, one_pow] #align abs_neg_one_pow abs_neg_one_pow diff --git a/Mathlib/Algebra/Parity.lean b/Mathlib/Algebra/Parity.lean index 4286fdf35f7e8..64ec77999e215 100644 --- a/Mathlib/Algebra/Parity.lean +++ b/Mathlib/Algebra/Parity.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import Mathlib.Algebra.Group.Opposite -import Mathlib.Algebra.Order.Ring.Abs +import Mathlib.Algebra.GroupPower.Ring +import Mathlib.Algebra.Order.Group.Abs +import Mathlib.Algebra.Order.Ring.Canonical +import Mathlib.Data.Nat.Cast.Basic import Mathlib.Data.Nat.Cast.Commute import Mathlib.Data.Set.Defs @@ -279,7 +282,7 @@ theorem even_iff_exists_two_mul (a : α) : Even a ↔ ∃ b, a = 2 * b := by simp [even_iff_exists_two_nsmul] #align even_iff_exists_two_mul even_iff_exists_two_mul -theorem even_iff_two_dvd {a : α} : Even a ↔ 2 ∣ a := by simp [Even, Dvd.dvd, two_mul] +theorem even_iff_two_dvd : Even a ↔ 2 ∣ a := by simp [Even, Dvd.dvd, two_mul] #align even_iff_two_dvd even_iff_two_dvd alias ⟨Even.two_dvd, _⟩ := even_iff_two_dvd @@ -338,7 +341,7 @@ def Odd (a : α) : Prop := #align odd Odd set_option linter.deprecated false in -theorem odd_iff_exists_bit1 {a : α} : Odd a ↔ ∃ b, a = bit1 b := +theorem odd_iff_exists_bit1 : Odd a ↔ ∃ b, a = bit1 b := exists_congr fun b => by rw [two_mul] rfl @@ -522,99 +525,3 @@ theorem odd_abs [LinearOrder α] : Odd (abs a) ↔ Odd a := by #align odd_abs odd_abs end Ring - -section Powers - -/-! -### Lemmas for canonically linear ordered semirings or linear ordered rings - -The slightly unusual typeclass assumptions `[LinearOrderedSemiring R] [ExistsAddOfLE R]` cover two -more familiar settings: -* `[LinearOrderedRing R]`, eg `ℤ`, `ℚ` or `ℝ` -* `[CanonicallyLinearOrderedSemiring R]` (although we don't actually have this typeclass), eg `ℕ`, - `ℚ≥0` or `ℝ≥0` --/ - -section LinearOrderedSemiring -variable [LinearOrderedSemiring R] [ExistsAddOfLE R] {a b : R} {n : ℕ} - -theorem Even.pow_nonneg (hn : Even n) (a : R) : 0 ≤ a ^ n := by - obtain ⟨k, rfl⟩ := hn; rw [pow_add]; exact mul_self_nonneg _ -#align even.pow_nonneg Even.pow_nonneg - -theorem Even.pow_pos (hn : Even n) (ha : a ≠ 0) : 0 < a ^ n := - (hn.pow_nonneg _).lt_of_ne' (pow_ne_zero _ ha) -#align even.pow_pos Even.pow_pos - -theorem Odd.pow_neg_iff (hn : Odd n) : a ^ n < 0 ↔ a < 0 := by - refine ⟨lt_imp_lt_of_le_imp_le (pow_nonneg · _), fun ha ↦ ?_⟩ - obtain ⟨k, rfl⟩ := hn - rw [pow_succ] - exact mul_neg_of_pos_of_neg ((even_two_mul _).pow_pos ha.ne) ha -#align odd.pow_neg_iff Odd.pow_neg_iff - -theorem Odd.pow_nonneg_iff (hn : Odd n) : 0 ≤ a ^ n ↔ 0 ≤ a := - le_iff_le_iff_lt_iff_lt.2 hn.pow_neg_iff -#align odd.pow_nonneg_iff Odd.pow_nonneg_iff - -theorem Odd.pow_nonpos_iff (hn : Odd n) : a ^ n ≤ 0 ↔ a ≤ 0 := by - rw [le_iff_lt_or_eq, le_iff_lt_or_eq, hn.pow_neg_iff, pow_eq_zero_iff] - rintro rfl; simp [Odd, eq_comm (a := 0)] at hn -#align odd.pow_nonpos_iff Odd.pow_nonpos_iff - -theorem Odd.pow_pos_iff (hn : Odd n) : 0 < a ^ n ↔ 0 < a := - lt_iff_lt_of_le_iff_le hn.pow_nonpos_iff -#align odd.pow_pos_iff Odd.pow_pos_iff - -alias ⟨_, Odd.pow_nonpos⟩ := Odd.pow_nonpos_iff -#align odd.pow_nonpos Odd.pow_nonpos - -alias ⟨_, Odd.pow_neg⟩ := Odd.pow_neg_iff -#align odd.pow_neg Odd.pow_neg - -theorem Even.pow_pos_iff (hn : Even n) (h₀ : n ≠ 0) : 0 < a ^ n ↔ a ≠ 0 := - ⟨fun h ha => by - rw [ha, zero_pow h₀] at h - exact lt_irrefl 0 h, hn.pow_pos⟩ -#align even.pow_pos_iff Even.pow_pos_iff - -lemma Odd.strictMono_pow (hn : Odd n) : StrictMono fun a : R => a ^ n := by - have hn₀ : n ≠ 0 := by rintro rfl; simp [Odd, eq_comm (a := 0)] at hn - intro a b hab - obtain ha | ha := le_total 0 a - · exact pow_lt_pow_left hab ha hn₀ - obtain hb | hb := lt_or_le 0 b - · exact (hn.pow_nonpos ha).trans_lt (pow_pos hb _) - obtain ⟨c, hac⟩ := exists_add_of_le ha - obtain ⟨d, hbd⟩ := exists_add_of_le hb - have hd := nonneg_of_le_add_right (hb.trans_eq hbd) - refine lt_of_add_lt_add_right (a := c ^ n + d ^ n) ?_ - dsimp - calc - a ^ n + (c ^ n + d ^ n) = d ^ n := by - rw [← add_assoc, hn.pow_add_pow_eq_zero hac.symm, zero_add] - _ < c ^ n := pow_lt_pow_left ?_ hd hn₀ - _ = b ^ n + (c ^ n + d ^ n) := by rw [add_left_comm, hn.pow_add_pow_eq_zero hbd.symm, add_zero] - refine lt_of_add_lt_add_right (a := a + b) ?_ - rwa [add_rotate', ← hbd, add_zero, add_left_comm, ← add_assoc, ← hac, zero_add] -#align odd.strict_mono_pow Odd.strictMono_pow - -end LinearOrderedSemiring - -section LinearOrderedRing -variable [LinearOrderedRing R] {a : R} {n : ℕ} - -theorem Even.pow_abs {p : ℕ} (hp : Even p) (a : R) : |a| ^ p = a ^ p := by - rw [← abs_pow, abs_eq_self] - exact hp.pow_nonneg _ -#align even.pow_abs Even.pow_abs - -set_option linter.deprecated false in -@[simp] -theorem pow_bit0_abs (a : R) (p : ℕ) : |a| ^ bit0 p = a ^ bit0 p := - (even_bit0 _).pow_abs _ -#align pow_bit0_abs pow_bit0_abs - -end LinearOrderedRing - -end Powers diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index 8d8d12aa212b0..0911897131359 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -169,7 +169,7 @@ theorem Convex.taylor_approx_two_segment {v w : E} (hv : x + v ∈ interior s) zero_smul, Ne, not_false_iff, bit0_eq_zero, zero_pow] abel · simp only [Real.norm_eq_abs, abs_mul, add_nonneg (norm_nonneg v) (norm_nonneg w), abs_of_nonneg, - hpos.le, mul_assoc, pow_bit0_abs, norm_nonneg, abs_pow] + hpos.le, mul_assoc, norm_nonneg, abs_pow] #align convex.taylor_approx_two_segment Convex.taylor_approx_two_segment /-- One can get `f'' v w` as the limit of `h ^ (-2)` times the alternate sum of the values of `f` diff --git a/Mathlib/Data/Rat/Sqrt.lean b/Mathlib/Data/Rat/Sqrt.lean index c04ef753fc65f..9b2dcdc120027 100644 --- a/Mathlib/Data/Rat/Sqrt.lean +++ b/Mathlib/Data/Rat/Sqrt.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Order.Ring.Abs -import Mathlib.Algebra.Parity import Mathlib.Data.Rat.Order import Mathlib.Data.Rat.Lemmas import Mathlib.Data.Int.Sqrt diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index b132f01ab1c76..08a6a9c123051 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -3,12 +3,12 @@ Copyright (c) 2022 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Order.Interval.Finset.Defs -import Mathlib.Order.SuccPred.Basic -import Mathlib.Order.Hom.Basic import Mathlib.Data.Countable.Basic import Mathlib.Data.Nat.Cast.Order import Mathlib.Logic.Encodable.Basic +import Mathlib.Order.Hom.Basic +import Mathlib.Order.Interval.Finset.Defs +import Mathlib.Order.SuccPred.Basic #align_import order.succ_pred.linear_locally_finite from "leanprover-community/mathlib"@"2705404e701abc6b3127da906f40bae062a169c9" From ad2ccae9e38b7209f29c6cb817f0a017ab2f287a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 7 May 2024 08:14:37 +0000 Subject: [PATCH 02/52] chore: remove pointless @[reducible] attributes (#12691) This is a backport of changes made by @Ruben-VandeVelde. These attributes had no effect, and will become illegal on v4.9.0. Co-authored-by: Scott Morrison --- Mathlib/Algebra/Group/Aut.lean | 6 +++++- Mathlib/Algebra/Polynomial/Module/Basic.lean | 1 - Mathlib/Algebra/Ring/CompTypeclasses.lean | 4 ---- Mathlib/Logic/Equiv/TransferInstance.lean | 2 +- Mathlib/Order/Heyting/Hom.lean | 2 +- 5 files changed, 7 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index 17cd8d4e0f8e2..b3a86a3f08ff0 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -31,12 +31,16 @@ MulAut, AddAut variable {A : Type*} {M : Type*} {G : Type*} /-- The group of multiplicative automorphisms. -/ -@[to_additive (attr := reducible) "The group of additive automorphisms."] +@[reducible, to_additive "The group of additive automorphisms."] def MulAut (M : Type*) [Mul M] := M ≃* M #align mul_aut MulAut #align add_aut AddAut +-- Note that `(attr := reducible)` in `to_additive` currently doesn't work, +-- so we add the reducible attribute manually. +attribute [reducible] AddAut + namespace MulAut variable (M) [Mul M] diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index 936483cbf7e75..3ee894a51bf2a 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -282,7 +282,6 @@ theorem induction_linear {P : PolynomialModule R M → Prop} (f : PolynomialModu Finsupp.induction_linear f h0 hadd hsingle #align polynomial_module.induction_linear PolynomialModule.induction_linear -@[semireducible] noncomputable instance polynomialModule : Module R[X] (PolynomialModule R M) := inferInstanceAs (Module R[X] (Module.AEval' (Finsupp.lmapDomain M R Nat.succ))) #align polynomial_module.polynomial_module PolynomialModule.polynomialModule diff --git a/Mathlib/Algebra/Ring/CompTypeclasses.lean b/Mathlib/Algebra/Ring/CompTypeclasses.lean index 07b7003664234..452a41db70080 100644 --- a/Mathlib/Algebra/Ring/CompTypeclasses.lean +++ b/Mathlib/Algebra/Ring/CompTypeclasses.lean @@ -126,8 +126,6 @@ instance triples₂ {σ₂₁ : R₂ →+* R₁} [RingHomInvPair σ₁₂ σ₂ This is not an instance, as for equivalences that are involutions, a better instance would be `RingHomInvPair e e`. Indeed, this declaration is not currently used in mathlib. - -See note [reducible non-instances]. -/ theorem of_ringEquiv (e : R₁ ≃+* R₂) : RingHomInvPair (↑e : R₁ →+* R₂) ↑e.symm := ⟨e.symm_toRingHom_comp_toRingHom, e.symm.symm_toRingHom_comp_toRingHom⟩ @@ -137,8 +135,6 @@ theorem of_ringEquiv (e : R₁ ≃+* R₂) : RingHomInvPair (↑e : R₁ →+* R Swap the direction of a `RingHomInvPair`. This is not an instance as it would loop, and better instances are often available and may often be preferrable to using this one. Indeed, this declaration is not currently used in mathlib. - -See note [reducible non-instances]. -/ theorem symm (σ₁₂ : R₁ →+* R₂) (σ₂₁ : R₂ →+* R₁) [RingHomInvPair σ₁₂ σ₂₁] : RingHomInvPair σ₂₁ σ₁₂ := diff --git a/Mathlib/Logic/Equiv/TransferInstance.lean b/Mathlib/Logic/Equiv/TransferInstance.lean index ae28ef520a7c0..991766095f8da 100644 --- a/Mathlib/Logic/Equiv/TransferInstance.lean +++ b/Mathlib/Logic/Equiv/TransferInstance.lean @@ -134,7 +134,7 @@ noncomputable instance [Small.{v} α] (R : Type*) [SMul R α] : SMul R (Shrink.{ (equivShrink α).symm.smul R /-- Transfer `Pow` across an `Equiv` -/ -@[to_additive (attr := reducible) existing smul] +@[reducible, to_additive existing smul] protected def pow (N : Type*) [Pow β N] : Pow α N := ⟨fun x n => e.symm (e x ^ n)⟩ #align equiv.has_pow Equiv.pow diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index 0c9df0599467e..10c8ed91bfc08 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -185,7 +185,7 @@ end Equiv variable [FunLike F α β] --- Porting note: Revisit this issue to see if it works in Lean 4. -/ +-- Porting note: Revisit this issue to see if it works in Lean 4. /-- This can't be an instance because of typeclass loops. -/ lemma BoundedLatticeHomClass.toBiheytingHomClass [BooleanAlgebra α] [BooleanAlgebra β] [BoundedLatticeHomClass F α β] : BiheytingHomClass F α β := From 4b6177e2a98f2afb04fcad75bbe1bd397a1fa9ba Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Tue, 7 May 2024 08:14:38 +0000 Subject: [PATCH 03/52] fix(Cache): reference to Std (#12725) Follow-up to #12720, the find/replace missed a spot. Co-authored-by: Mario Carneiro --- Cache/IO.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Cache/IO.lean b/Cache/IO.lean index cfab57deba792..1dcf6db24b7dd 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -133,7 +133,7 @@ private def CacheM.getContext : IO CacheM.Context := do ("Archive", root), ("Counterexamples", root), ("Aesop", LAKEPACKAGESDIR / "aesop"), - ("Std", LAKEPACKAGESDIR / "std"), + ("Batteries", LAKEPACKAGESDIR / "batteries"), ("Cli", LAKEPACKAGESDIR / "Cli"), ("ProofWidgets", LAKEPACKAGESDIR / "proofwidgets"), ("Qq", LAKEPACKAGESDIR / "Qq"), From 859845f1833f51ea06dbb1e0dae75e39693006e8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 7 May 2024 09:25:39 +0000 Subject: [PATCH 04/52] refactor: Redefine 3AP-free sets (#12701) The current definition allows `{0, 1}` as a 3AP-free set in `ZMod 2`, which is really inconvenient from the point of view of interfacing with corners. This PR redefines it so that `0, 1, 0` in `ZMod 2` is a 3AP and `{0, 1}` is not 3AP-free. Also take the opportunity to reorder the binders to a more convenient fashion, to move the heavy result `threeAPFree_sphere` to `Combinatorics.Additive.Behrend` where it is used (it won't be used anywhere else) and to drop `rothNumberNat_isBigO_id` since it's a trivial result that pulls in a bunch of analysis. --- .../Additive/AP/Three/Behrend.lean | 31 ++- .../Combinatorics/Additive/AP/Three/Defs.lean | 217 ++++++------------ 2 files changed, 99 insertions(+), 149 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index e54b9e0f3778a..0062374e02b5d 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -44,13 +44,36 @@ integer points on that sphere and map them onto `ℕ` in a way that preserves ar 3AP-free, Salem-Spencer, Behrend construction, arithmetic progression, sphere, strictly convex -/ - open Nat hiding log - -open Finset Real - +open Finset Metric Real open scoped BigOperators Pointwise +/-- The frontier of a closed strictly convex set only contains trivial arithmetic progressions. +The idea is that an arithmetic progression is contained on a line and the frontier of a strictly +convex set does not contain lines. -/ +lemma threeAPFree_frontier {𝕜 E : Type*} [LinearOrderedField 𝕜] [TopologicalSpace E] + [AddCommMonoid E] [Module 𝕜 E] {s : Set E} (hs₀ : IsClosed s) (hs₁ : StrictConvex 𝕜 s) : + ThreeAPFree (frontier s) := by + intro a ha b hb c hc habc + obtain rfl : (1 / 2 : 𝕜) • a + (1 / 2 : 𝕜) • c = b := by + rwa [← smul_add, one_div, inv_smul_eq_iff₀ (show (2 : 𝕜) ≠ 0 by norm_num), two_smul] + have := + hs₁.eq (hs₀.frontier_subset ha) (hs₀.frontier_subset hc) one_half_pos one_half_pos + (add_halves _) hb.2 + simp [this, ← add_smul] + ring_nf + simp +#align add_salem_spencer_frontier threeAPFree_frontier + +lemma threeAPFree_sphere {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + [StrictConvexSpace ℝ E] (x : E) (r : ℝ) : ThreeAPFree (sphere x r) := by + obtain rfl | hr := eq_or_ne r 0 + · rw [sphere_zero] + exact threeAPFree_singleton _ + · convert threeAPFree_frontier isClosed_ball (strictConvex_closedBall ℝ x r) + exact (frontier_closedBall _ hr).symm +#align add_salem_spencer_sphere threeAPFree_sphere + namespace Behrend variable {α β : Type*} {n d k N : ℕ} {x : Fin n → ℕ} diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean index dd03366ebb78c..e76b1e38316ce 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Group.Freiman -import Mathlib.Analysis.Asymptotics.Asymptotics -import Mathlib.Analysis.Convex.StrictConvexSpace +import Mathlib.Data.Nat.Interval +import Mathlib.Data.Set.Pointwise.SMul #align_import combinatorics.additive.salem_spencer from "leanprover-community/mathlib"@"acf5258c81d0bc7cb254ed026c1352e685df306c" @@ -44,10 +44,8 @@ the size of the biggest 3AP-free subset of `{0, ..., n - 1}`. 3AP-free, Salem-Spencer, Roth, arithmetic progression, average, three-free -/ - -open Finset Function Metric Nat - -open Pointwise +open Finset Function Nat +open scoped Pointwise variable {F α β 𝕜 E : Type*} @@ -65,7 +63,7 @@ three. -/ of length three. This is also sometimes called a **non averaging set** or **Salem-Spencer set**."] -def ThreeGPFree : Prop := ∀ ⦃a b c⦄, a ∈ s → b ∈ s → c ∈ s → a * b = c * c → a = b +def ThreeGPFree : Prop := ∀ ⦃a⦄, a ∈ s → ∀ ⦃b⦄, b ∈ s → ∀ ⦃c⦄, c ∈ s → a * c = b * b → a = b #align mul_salem_spencer ThreeGPFree #align add_salem_spencer ThreeAPFree @@ -73,14 +71,13 @@ def ThreeGPFree : Prop := ∀ ⦃a b c⦄, a ∈ s → b ∈ s → c ∈ s → a @[to_additive "Whether a given finset is 3AP-free is decidable."] instance ThreeGPFree.instDecidable [DecidableEq α] {s : Finset α} : Decidable (ThreeGPFree (s : Set α)) := - decidable_of_iff (∀ a ∈ s, ∀ b ∈ s, ∀ c ∈ s, a * b = c * c → a = b) - ⟨fun h a b c ha hb hc => h a ha b hb c hc, fun h _ ha _ hb _ hc => h ha hb hc⟩ + decidable_of_iff (∀ a ∈ s, ∀ b ∈ s, ∀ c ∈ s, a * c = b * b → a = b) Iff.rfl variable {s t} @[to_additive] theorem ThreeGPFree.mono (h : t ⊆ s) (hs : ThreeGPFree s) : ThreeGPFree t := - fun _ _ _ ha hb hc => hs (h ha) (h hb) (h hc) + fun _ ha _ hb _ hc ↦ hs (h ha) (h hb) (h hc) #align mul_salem_spencer.mono ThreeGPFree.mono #align add_salem_spencer.mono ThreeAPFree.mono @@ -91,7 +88,7 @@ theorem threeGPFree_empty : ThreeGPFree (∅ : Set α) := fun _ _ _ ha => ha.eli @[to_additive] theorem Set.Subsingleton.threeGPFree (hs : s.Subsingleton) : ThreeGPFree s := - fun _ _ _ ha hb _ _ => hs ha hb + fun _ ha _ hb _ _ _ ↦ hs ha hb #align set.subsingleton.mul_salem_spencer Set.Subsingleton.threeGPFree #align set.subsingleton.add_salem_spencer Set.Subsingleton.threeAPFree @@ -103,7 +100,7 @@ theorem threeGPFree_singleton (a : α) : ThreeGPFree ({a} : Set α) := @[to_additive ThreeAPFree.prod] theorem ThreeGPFree.prod {t : Set β} (hs : ThreeGPFree s) (ht : ThreeGPFree t) : - ThreeGPFree (s ×ˢ t) := fun _ _ _ ha hb hc h => + ThreeGPFree (s ×ˢ t) := fun _ ha _ hb _ hc h ↦ Prod.ext (hs ha.1 hb.1 hc.1 (Prod.ext_iff.1 h).1) (ht ha.2 hb.2 hc.2 (Prod.ext_iff.1 h).2) #align mul_salem_spencer.prod ThreeGPFree.prod #align add_salem_spencer.prod ThreeAPFree.prod @@ -111,7 +108,7 @@ theorem ThreeGPFree.prod {t : Set β} (hs : ThreeGPFree s) (ht : ThreeGPFree t) @[to_additive] theorem threeGPFree_pi {ι : Type*} {α : ι → Type*} [∀ i, Monoid (α i)] {s : ∀ i, Set (α i)} (hs : ∀ i, ThreeGPFree (s i)) : ThreeGPFree ((univ : Set ι).pi s) := - fun _ _ _ ha hb hc h => + fun _ ha _ hb _ hc h ↦ funext fun i => hs i (ha i trivial) (hb i trivial) (hc i trivial) <| congr_fun h i #align mul_salem_spencer_pi threeGPFree_pi #align add_salem_spencer_pi threeAPFree_pi @@ -119,15 +116,16 @@ theorem threeGPFree_pi {ι : Type*} {α : ι → Type*} [∀ i, Monoid (α i)] { end Monoid section CommMonoid +variable [CommMonoid α] [CommMonoid β] {s A : Set α} {t B : Set β} {f : α → β} {a : α} variable [CommMonoid α] [CommMonoid β] {s : Set α} {a : α} @[to_additive] theorem ThreeGPFree.of_image [FunLike F α β] [FreimanHomClass F s β 2] (f : F) (hf : s.InjOn f) (h : ThreeGPFree (f '' s)) : ThreeGPFree s := - fun _ _ _ ha hb hc habc => hf ha hb <| + fun _ ha _ hb _ hc habc => hf ha hb <| h (mem_image_of_mem _ ha) (mem_image_of_mem _ hb) (mem_image_of_mem _ hc) <| - map_mul_map_eq_map_mul_map f ha hb hc hc habc + map_mul_map_eq_map_mul_map f ha hc hb hb habc #align mul_salem_spencer.of_image ThreeGPFree.of_image #align add_salem_spencer.of_image ThreeAPFree.of_image @@ -135,8 +133,8 @@ theorem ThreeGPFree.of_image [FunLike F α β] [FreimanHomClass F s β 2] (f : F @[to_additive] theorem ThreeGPFree.image [FunLike F α β] [MulHomClass F α β] (f : F) (hf : (s * s).InjOn f) (h : ThreeGPFree s) : ThreeGPFree (f '' s) := by - rintro _ _ _ ⟨a, ha, rfl⟩ ⟨b, hb, rfl⟩ ⟨c, hc, rfl⟩ habc - rw [h ha hb hc (hf (mul_mem_mul ha hb) (mul_mem_mul hc hc) <| by rwa [map_mul, map_mul])] + rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ habc + rw [h ha hb hc (hf (mul_mem_mul ha hc) (mul_mem_mul hb hb) <| by rwa [map_mul, map_mul])] #align mul_salem_spencer.image ThreeGPFree.image #align add_salem_spencer.image ThreeAPFree.image @@ -146,73 +144,56 @@ section CancelCommMonoid variable [CancelCommMonoid α] {s : Set α} {a : α} -@[to_additive] -theorem threeGPFree_insert : ThreeGPFree (insert a s) ↔ ThreeGPFree s ∧ - (∀ ⦃b c⦄, b ∈ s → c ∈ s → a * b = c * c → a = b) ∧ - ∀ ⦃b c⦄, b ∈ s → c ∈ s → b * c = a * a → b = c := by - refine' ⟨fun hs => ⟨hs.mono (subset_insert _ _), - fun b c hb hc => hs (Or.inl rfl) (Or.inr hb) (Or.inr hc), - fun b c hb hc => hs (Or.inr hb) (Or.inr hc) (Or.inl rfl)⟩, _⟩ - rintro ⟨hs, ha, ha'⟩ b c d hb hc hd h +lemma ThreeGPFree.eq_right (hs : ThreeGPFree s) : + ∀ ⦃a⦄, a ∈ s → ∀ ⦃b⦄, b ∈ s → ∀ ⦃c⦄, c ∈ s → a * c = b * b → b = c := by + rintro a ha b hb c hc habc + obtain rfl := hs ha hb hc habc + simpa using habc.symm + +@[to_additive] lemma threeGPFree_insert : + ThreeGPFree (insert a s) ↔ ThreeGPFree s ∧ + (∀ ⦃b⦄, b ∈ s → ∀ ⦃c⦄, c ∈ s → a * c = b * b → a = b) ∧ + ∀ ⦃b⦄, b ∈ s → ∀ ⦃c⦄, c ∈ s → b * c = a * a → b = a := by + refine ⟨fun hs ↦ ⟨hs.mono (subset_insert _ _), + fun b hb c hc ↦ hs (Or.inl rfl) (Or.inr hb) (Or.inr hc), + fun b hb c hc ↦ hs (Or.inr hb) (Or.inl rfl) (Or.inr hc)⟩, ?_⟩ + rintro ⟨hs, ha, ha'⟩ b hb c hc d hd h rw [mem_insert_iff] at hb hc hd obtain rfl | hb := hb <;> obtain rfl | hc := hc · rfl all_goals obtain rfl | hd := hd - · exact (mul_left_cancel h).symm + · exact (ha' hc hc h.symm).symm · exact ha hc hd h · exact mul_right_cancel h - · exact (ha hb hd <| (mul_comm _ _).trans h).symm - · exact ha' hb hc h + · exact ha' hb hd h + · obtain rfl := ha hc hb ((mul_comm _ _).trans h) + exact ha' hb hc h · exact hs hb hc hd h #align mul_salem_spencer_insert threeGPFree_insert #align add_salem_spencer_insert threeAPFree_insert -@[to_additive (attr := simp)] -theorem threeGPFree_pair (a b : α) : ThreeGPFree ({a, b} : Set α) := by - rw [threeGPFree_insert] - refine' ⟨threeGPFree_singleton _, _, _⟩ - · rintro c d (rfl : c = b) (rfl : d = c) - exact mul_right_cancel - · rintro c d (rfl : c = b) (rfl : d = c) _ - rfl -#align mul_salem_spencer_pair threeGPFree_pair -#align add_salem_spencer_pair threeAPFree_pair - @[to_additive] -theorem ThreeGPFree.mul_left (hs : ThreeGPFree s) : ThreeGPFree ((a * ·) '' s) := by - rintro _ _ _ ⟨b, hb, rfl⟩ ⟨c, hc, rfl⟩ ⟨d, hd, rfl⟩ h - rw [mul_mul_mul_comm, mul_mul_mul_comm a d] at h - rw [hs hb hc hd (mul_left_cancel h)] -#align mul_salem_spencer.mul_left ThreeGPFree.mul_left -#align add_salem_spencer.add_left ThreeAPFree.add_left +theorem ThreeGPFree.smul_set (hs : ThreeGPFree s) : ThreeGPFree (a • s) := by + rintro _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ _ ⟨d, hd, rfl⟩ h + exact congr_arg (a • ·) $ hs hb hc hd $ by simpa [mul_mul_mul_comm _ _ a] using h +#align mul_salem_spencer.mul_left ThreeGPFree.smul_set +#align add_salem_spencer.add_left ThreeAPFree.vadd_set -@[to_additive] -theorem ThreeGPFree.mul_right (hs : ThreeGPFree s) : ThreeGPFree ((· * a) '' s) := by - rintro _ _ _ ⟨b, hb, rfl⟩ ⟨c, hc, rfl⟩ ⟨d, hd, rfl⟩ h - rw [mul_mul_mul_comm, mul_mul_mul_comm d] at h - rw [hs hb hc hd (mul_right_cancel h)] -#align mul_salem_spencer.mul_right ThreeGPFree.mul_right -#align add_salem_spencer.add_right ThreeAPFree.add_right +#noalign mul_salem_spencer.mul_right +#noalign add_salem_spencer.add_right @[to_additive] -theorem threeGPFree_mul_left_iff : ThreeGPFree ((a * ·) '' s) ↔ ThreeGPFree s := - ⟨fun hs b c d hb hc hd h => +theorem threeGPFree_smul_set : ThreeGPFree ((a * ·) '' s) ↔ ThreeGPFree s := + ⟨fun hs b hb c hc d hd h ↦ mul_left_cancel (hs (mem_image_of_mem _ hb) (mem_image_of_mem _ hc) (mem_image_of_mem _ hd) <| by rw [mul_mul_mul_comm, h, mul_mul_mul_comm]), - ThreeGPFree.mul_left⟩ -#align mul_salem_spencer_mul_left_iff threeGPFree_mul_left_iff -#align add_salem_spencer_add_left_iff threeAPFree_add_left_iff + ThreeGPFree.smul_set⟩ +#align mul_salem_spencer_mul_left_iff threeGPFree_smul_set +#align add_salem_spencer_add_left_iff threeAPFree_vadd_set -@[to_additive] -theorem threeGPFree_mul_right_iff : ThreeGPFree ((· * a) '' s) ↔ ThreeGPFree s := - ⟨fun hs b c d hb hc hd h => - mul_right_cancel - (hs (Set.mem_image_of_mem _ hb) (Set.mem_image_of_mem _ hc) (Set.mem_image_of_mem _ hd) <| by - rw [mul_mul_mul_comm, h, mul_mul_mul_comm]), - ThreeGPFree.mul_right⟩ -#align mul_salem_spencer_mul_right_iff threeGPFree_mul_right_iff -#align add_salem_spencer_add_right_iff threeAPFree_add_right_iff +#noalign mul_salem_spencer_mul_right_iff +#noalign add_salem_spencer_add_right_iff end CancelCommMonoid @@ -223,10 +204,10 @@ variable [OrderedCancelCommMonoid α] {s : Set α} {a : α} @[to_additive] theorem threeGPFree_insert_of_lt (hs : ∀ i ∈ s, i < a) : ThreeGPFree (insert a s) ↔ - ThreeGPFree s ∧ ∀ ⦃b c⦄, b ∈ s → c ∈ s → a * b = c * c → a = b := by + ThreeGPFree s ∧ ∀ ⦃b⦄, b ∈ s → ∀ ⦃c⦄, c ∈ s → a * c = b * b → a = b := by refine' threeGPFree_insert.trans _ rw [← and_assoc] - exact and_iff_left fun b c hb hc h => ((mul_lt_mul_of_lt_of_lt (hs _ hb) (hs _ hc)).ne h).elim + exact and_iff_left fun b hb c hc h => ((mul_lt_mul_of_lt_of_lt (hs _ hb) (hs _ hc)).ne h).elim #align mul_salem_spencer_insert_of_lt threeGPFree_insert_of_lt #align add_salem_spencer_insert_of_lt threeAPFree_insert_of_lt @@ -236,77 +217,39 @@ section CancelCommMonoidWithZero variable [CancelCommMonoidWithZero α] [NoZeroDivisors α] {s : Set α} {a : α} -theorem ThreeGPFree.mul_left₀ (hs : ThreeGPFree s) (ha : a ≠ 0) : +theorem ThreeGPFree.smul_set₀ (hs : ThreeGPFree s) (ha : a ≠ 0) : ThreeGPFree ((a * ·) '' s) := by - rintro _ _ _ ⟨b, hb, rfl⟩ ⟨c, hc, rfl⟩ ⟨d, hd, rfl⟩ h - rw [mul_mul_mul_comm, mul_mul_mul_comm a d] at h - rw [hs hb hc hd (mul_left_cancel₀ (mul_ne_zero ha ha) h)] -#align mul_salem_spencer.mul_left₀ ThreeGPFree.mul_left₀ - -theorem ThreeGPFree.mul_right₀ (hs : ThreeGPFree s) (ha : a ≠ 0) : - ThreeGPFree ((· * a) '' s) := by - rintro _ _ _ ⟨b, hb, rfl⟩ ⟨c, hc, rfl⟩ ⟨d, hd, rfl⟩ h - rw [mul_mul_mul_comm, mul_mul_mul_comm d] at h - rw [hs hb hc hd (mul_right_cancel₀ (mul_ne_zero ha ha) h)] -#align mul_salem_spencer.mul_right₀ ThreeGPFree.mul_right₀ - -theorem threeGPFree_mul_left_iff₀ (ha : a ≠ 0) : - ThreeGPFree ((a * ·) '' s) ↔ ThreeGPFree s := - ⟨fun hs b c d hb hc hd h => + rintro _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ _ ⟨d, hd, rfl⟩ h + exact congr_arg (a • ·) $ hs hb hc hd $ by simpa [mul_mul_mul_comm _ _ a, ha] using h +#align mul_salem_spencer.mul_left₀ ThreeGPFree.smul_set₀ + +#noalign mul_salem_spencer.mul_right₀.mul_right₀ + +theorem threeGPFree_smul_set₀ (ha : a ≠ 0) : ThreeGPFree (a • s) ↔ ThreeGPFree s := + ⟨fun hs b hb c hc d hd h ↦ mul_left_cancel₀ ha (hs (Set.mem_image_of_mem _ hb) (Set.mem_image_of_mem _ hc) (Set.mem_image_of_mem _ hd) <| by - rw [mul_mul_mul_comm, h, mul_mul_mul_comm]), - fun hs => hs.mul_left₀ ha⟩ -#align mul_salem_spencer_mul_left_iff₀ threeGPFree_mul_left_iff₀ + rw [smul_eq_mul, smul_eq_mul, mul_mul_mul_comm, h, mul_mul_mul_comm]), + fun hs => hs.smul_set₀ ha⟩ +#align mul_salem_spencer_mul_left_iff₀ threeGPFree_smul_set₀ -theorem threeGPFree_mul_right_iff₀ (ha : a ≠ 0) : - ThreeGPFree ((· * a) '' s) ↔ ThreeGPFree s := - ⟨fun hs b c d hb hc hd h => - mul_right_cancel₀ ha - (hs (Set.mem_image_of_mem _ hb) (Set.mem_image_of_mem _ hc) (Set.mem_image_of_mem _ hd) <| by - rw [mul_mul_mul_comm, h, mul_mul_mul_comm]), - fun hs => hs.mul_right₀ ha⟩ -#align mul_salem_spencer_mul_right_iff₀ threeGPFree_mul_right_iff₀ +#noalign mul_salem_spencer_mul_right_iff₀ end CancelCommMonoidWithZero section Nat theorem threeAPFree_iff_eq_right {s : Set ℕ} : - ThreeAPFree s ↔ ∀ ⦃a b c⦄, a ∈ s → b ∈ s → c ∈ s → a + b = c + c → a = c := by - refine' forall₄_congr fun a b c _ => forall₃_congr fun _ _ habc => ⟨_, _⟩ + ThreeAPFree s ↔ ∀ ⦃a⦄, a ∈ s → ∀ ⦃b⦄, b ∈ s → ∀ ⦃c⦄, c ∈ s → a + c = b + b → a = c := by + refine' forall₄_congr fun a _ha b hb => forall₃_congr fun c hc habc => ⟨_, _⟩ + · rintro rfl + exact (add_left_cancel habc).symm · rintro rfl simp_rw [← two_mul] at habc exact mul_left_cancel₀ two_ne_zero habc - · rintro rfl - exact (add_left_cancel habc).symm #align add_salem_spencer_iff_eq_right threeAPFree_iff_eq_right end Nat - -/-- The frontier of a closed strictly convex set only contains trivial arithmetic progressions. -The idea is that an arithmetic progression is contained on a line and the frontier of a strictly -convex set does not contain lines. -/ -theorem threeAPFree_frontier [LinearOrderedField 𝕜] [TopologicalSpace E] [AddCommMonoid E] - [Module 𝕜 E] {s : Set E} (hs₀ : IsClosed s) (hs₁ : StrictConvex 𝕜 s) : - ThreeAPFree (frontier s) := by - intro a b c ha hb hc habc - obtain rfl : (1 / 2 : 𝕜) • a + (1 / 2 : 𝕜) • b = c := by - rwa [← smul_add, one_div, inv_smul_eq_iff₀ (show (2 : 𝕜) ≠ 0 by norm_num), two_smul] - exact - hs₁.eq (hs₀.frontier_subset ha) (hs₀.frontier_subset hb) one_half_pos one_half_pos - (add_halves _) hc.2 -#align add_salem_spencer_frontier threeAPFree_frontier - -theorem threeAPFree_sphere [NormedAddCommGroup E] [NormedSpace ℝ E] [StrictConvexSpace ℝ E] - (x : E) (r : ℝ) : ThreeAPFree (sphere x r) := by - obtain rfl | hr := eq_or_ne r 0 - · rw [sphere_zero] - exact threeAPFree_singleton _ - · convert threeAPFree_frontier isClosed_ball (strictConvex_closedBall ℝ x r) - exact (frontier_closedBall _ hr).symm -#align add_salem_spencer_sphere threeAPFree_sphere - end ThreeAPFree open Finset @@ -356,11 +299,11 @@ theorem ThreeGPFree.le_mulRothNumber (hs : ThreeGPFree (s : Set α)) (h : s ⊆ #align add_salem_spencer.le_add_roth_number ThreeAPFree.le_addRothNumber @[to_additive] -theorem ThreeGPFree.roth_number_eq (hs : ThreeGPFree (s : Set α)) : +theorem ThreeGPFree.mulRothNumber_eq (hs : ThreeGPFree (s : Set α)) : mulRothNumber s = s.card := (mulRothNumber_le _).antisymm <| hs.le_mulRothNumber <| Subset.refl _ -#align mul_salem_spencer.roth_number_eq ThreeGPFree.roth_number_eq -#align add_salem_spencer.roth_number_eq ThreeAPFree.roth_number_eq +#align mul_salem_spencer.roth_number_eq ThreeGPFree.mulRothNumber_eq +#align add_salem_spencer.roth_number_eq ThreeAPFree.addRothNumber_eq @[to_additive (attr := simp)] theorem mulRothNumber_empty : mulRothNumber (∅ : Finset α) = 0 := @@ -370,7 +313,7 @@ theorem mulRothNumber_empty : mulRothNumber (∅ : Finset α) = 0 := @[to_additive (attr := simp)] theorem mulRothNumber_singleton (a : α) : mulRothNumber ({a} : Finset α) = 1 := by - refine' ThreeGPFree.roth_number_eq _ + refine' ThreeGPFree.mulRothNumber_eq _ rw [coe_singleton] exact threeGPFree_singleton a #align mul_roth_number_singleton mulRothNumber_singleton @@ -429,11 +372,9 @@ theorem mulRothNumber_map_mul_left : obtain ⟨u, hus, rfl⟩ := hus rw [coe_map] at hu rw [← hcard, card_map] - exact (threeGPFree_mul_left_iff.1 hu).le_mulRothNumber hus + exact (threeGPFree_smul_set.1 hu).le_mulRothNumber hus · obtain ⟨u, hus, hcard, hu⟩ := mulRothNumber_spec s - have h : ThreeGPFree (u.map <| mulLeftEmbedding a : Set α) := by - rw [coe_map] - exact hu.mul_left + have h : ThreeGPFree (u.map <| mulLeftEmbedding a : Set α) := by rw [coe_map]; exact hu.smul_set convert h.le_mulRothNumber (map_subset_map.2 hus) using 1 rw [card_map, hcard] #align mul_roth_number_map_mul_left mulRothNumber_map_mul_left @@ -509,18 +450,4 @@ theorem addRothNumber_Ico (a b : ℕ) : addRothNumber (Ico a b) = rothNumberNat exact (add_tsub_cancel_of_le h).symm #align add_roth_number_Ico addRothNumber_Ico -open Asymptotics Filter - -theorem rothNumberNat_isBigOWith_id : - IsBigOWith 1 atTop (fun N => (rothNumberNat N : ℝ)) fun N => (N : ℝ) := - isBigOWith_of_le _ <| by simpa only [Real.norm_natCast, Nat.cast_le] using rothNumberNat_le -set_option linter.uppercaseLean3 false in -#align roth_number_nat_is_O_with_id rothNumberNat_isBigOWith_id - -/-- The Roth number has the trivial bound `rothNumberNat N = O(N)`. -/ -theorem rothNumberNat_isBigO_id : (fun N => (rothNumberNat N : ℝ)) =O[atTop] fun N => (N : ℝ) := - rothNumberNat_isBigOWith_id.isBigO -set_option linter.uppercaseLean3 false in -#align roth_number_nat_is_O_id rothNumberNat_isBigO_id - end rothNumberNat From 0477177214eb78b0cb5ba949dfa7dc34b85c08a9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier=20G=C3=A9n=C3=A9reux?= Date: Tue, 7 May 2024 09:50:32 +0000 Subject: [PATCH 05/52] feat: Hadamard three-lines theorem (#7919) This PR introduces a version of the Hadamard three-lines theorem (see [here](https://en.wikipedia.org/wiki/Hadamard_three-lines_theorem)), which is a prerequisite for norm interpolation. This project was suggested by @dupuisf earlier this year. This is my first decently sized PR so thank you for your patience if I have made any mistakes! --- Mathlib.lean | 1 + Mathlib/Analysis/Complex/Hadamard.lean | 419 +++++++++++++++++++++++++ 2 files changed, 420 insertions(+) create mode 100644 Mathlib/Analysis/Complex/Hadamard.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3cd40686c967f..e20d01bb67ce2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -841,6 +841,7 @@ import Mathlib.Analysis.Complex.CauchyIntegral import Mathlib.Analysis.Complex.Circle import Mathlib.Analysis.Complex.Conformal import Mathlib.Analysis.Complex.Convex +import Mathlib.Analysis.Complex.Hadamard import Mathlib.Analysis.Complex.HalfPlane import Mathlib.Analysis.Complex.Isometry import Mathlib.Analysis.Complex.Liouville diff --git a/Mathlib/Analysis/Complex/Hadamard.lean b/Mathlib/Analysis/Complex/Hadamard.lean new file mode 100644 index 0000000000000..7dce70fa892c8 --- /dev/null +++ b/Mathlib/Analysis/Complex/Hadamard.lean @@ -0,0 +1,419 @@ +/- +Copyright (c) 2023 Xavier Généreux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Généreux +-/ + +import Mathlib.Analysis.SpecialFunctions.Pow.Deriv +import Mathlib.Analysis.Complex.PhragmenLindelof + +/-! +# Hadamard three-lines Theorem + +In this file we present a proof of a special case of Hadamard's three-lines Theorem. + +## Main result + +- `norm_le_interp_of_mem_verticalClosedStrip` : +Hadamard three-line theorem on `re ⁻¹' [0,1]`: If `f` is a bounded function, continuous on +`re ⁻¹' [0,1]` and differentiable on `re ⁻¹' (0,1)`, then for +`M(x) := sup ((norm ∘ f) '' (re ⁻¹' {x}))`, that is `M(x)` is the supremum of the absolute value of +`f` along the vertical lines `re z = x`, we have that `∀ z ∈ re ⁻¹' [0,1]` the inequality +`‖f(z)‖ ≤ M(0) ^ (1 - z.re) * M(1) ^ z.re` holds. This can be seen to be equivalent to the statement +that `log M(re z)` is a convex function on `[0,1]`. + +- `norm_le_interp_of_mem_verticalClosedStrip'` : +Variant of the above lemma in simpler terms. In particular, it makes no mention of the helper +functions defined in this file. + +## Main definitions + +- `Complex.HadamardThreeLines.verticalStrip` : + The vertical strip defined by : `re ⁻¹' Ioo a b` + +- `Complex.HadamardThreeLines.verticalClosedStrip` : + The vertical strip defined by : `re ⁻¹' Icc a b` + +- `Complex.HadamardThreeLines.sSupNormIm` : + The supremum function on vertical lines defined by : `sSup {|f(z)| : z.re = x}` + +- `Complex.HadamardThreeLines.interpStrip` : + The interpolation between the `sSupNormIm` on the edges of the vertical strip. + +- `Complex.HadamardThreeLines.invInterpStrip` : + Inverse of the interpolation between the `sSupNormIm` on the edges of the vertical strip. + +- `Complex.HadamardThreeLines.F` : + Function defined by `f` times `invInterpStrip`. Convenient form for proofs. + +## Note + +The proof follows from Phragmén-Lindelöf when both frontiers are not everywhere zero. +We then use a limit argument to cover the case when either of the sides are `0`. +-/ + + +open Set Filter Function Complex Topology + +namespace Complex +namespace HadamardThreeLines + +/-- The vertical strip in the complex plane containing all `z ∈ ℂ` such that `z.re ∈ Ioo a b`. -/ +def verticalStrip (a : ℝ) (b : ℝ) : Set ℂ := re ⁻¹' Ioo a b + +/-- The vertical strip in the complex plane containing all `z ∈ ℂ` such that `z.re ∈ Icc a b`. -/ +def verticalClosedStrip (a : ℝ) (b : ℝ) : Set ℂ := re ⁻¹' Icc a b + +/-- The supremum of the norm of `f` on imaginary lines. (Fixed real part) +This is also known as the function `M` -/ +noncomputable def sSupNormIm {E : Type*} [NormedAddCommGroup E] + (f : ℂ → E) (x : ℝ) : ℝ := + sSup ((norm ∘ f) '' (re ⁻¹' {x})) + +section invInterpStrip + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] (f : ℂ → E) (z : ℂ) + +/-- +The inverse of the interpolation of `sSupNormIm` on the two boundaries. +In other words, this is the inverse of the right side of the target inequality: +`|f(z)| ≤ |M(0) ^ (1-z)| * |M(1) ^ z|`. + +Shifting this by a positive epsilon allows us to prove the case when either of the boundaries +is zero.-/ +noncomputable def invInterpStrip (ε : ℝ) : ℂ := + (ε + sSupNormIm f 0) ^ (z - 1) * (ε + sSupNormIm f 1) ^ (-z) + +/-- A function useful for the proofs steps. We will aim to show that it is bounded by 1. -/ +noncomputable def F (ε : ℝ) := fun z ↦ invInterpStrip f z ε • f z + +/-- `sSup` of `norm` is nonneg applied to the image of `f` on the vertical line `re z = x` -/ +lemma sSupNormIm_nonneg (x : ℝ) : 0 ≤ sSupNormIm f x := by + apply Real.sSup_nonneg + rintro y ⟨z1, _, hz2⟩ + simp only [← hz2, comp, norm_nonneg] + +/-- `sSup` of `norm` translated by `ε > 0` is positive applied to the image of `f` on the +vertical line `re z = x` -/ +lemma sSupNormIm_eps_pos {ε : ℝ} (hε : ε > 0) (x : ℝ) : 0 < ε + sSupNormIm f x := by + linarith [sSupNormIm_nonneg f x] + +/-- Useful rewrite for the absolute value of `invInterpStrip`-/ +lemma abs_invInterpStrip {ε : ℝ} (hε : ε > 0) : + abs (invInterpStrip f z ε) = + (ε + sSupNormIm f 0) ^ (z.re - 1) * (ε + sSupNormIm f 1) ^ (-z.re) := by + simp only [invInterpStrip, map_mul] + repeat rw [← ofReal_add] + repeat rw [abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε _) _] + simp only [sub_re, one_re, neg_re] + +/-- The function `invInterpStrip` is `diffContOnCl`. -/ +lemma diffContOnCl_invInterpStrip {ε : ℝ} (hε : ε > 0) : + DiffContOnCl ℂ (fun z ↦ invInterpStrip f z ε) (verticalStrip 0 1) := by + apply Differentiable.diffContOnCl + apply Differentiable.mul + · apply Differentiable.const_cpow (Differentiable.sub_const (differentiable_id') 1) _ + left + rw [← ofReal_add, ofReal_ne_zero] + simp only [ne_eq, ne_of_gt (sSupNormIm_eps_pos f hε 0), not_false_eq_true] + · apply Differentiable.const_cpow (Differentiable.neg differentiable_id') + apply Or.inl + rw [← ofReal_add, ofReal_ne_zero] + exact (ne_of_gt (sSupNormIm_eps_pos f hε 1)) + +/-- If `f` is bounded on the unit vertical strip, then `f` is bounded by `sSupNormIm` there. -/ +lemma norm_le_sSupNormIm (f : ℂ → E) (z : ℂ) (hD : z ∈ verticalClosedStrip 0 1) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) : + ‖f z‖ ≤ sSupNormIm f (z.re) := by + refine le_csSup ?_ ?_ + · apply BddAbove.mono (image_subset (norm ∘ f) _) hB + exact preimage_mono (singleton_subset_iff.mpr hD) + · apply mem_image_of_mem (norm ∘ f) + simp only [mem_preimage, mem_singleton] + +/-- Alternative version of `norm_le_sSupNormIm` with a strict inequality and a positive `ε`. -/ +lemma norm_lt_sSupNormIm_eps (f : ℂ → E) (ε : ℝ) (hε : ε > 0) (z : ℂ) + (hD : z ∈ verticalClosedStrip 0 1) (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) : + ‖f z‖ < ε + sSupNormIm f (z.re) := + lt_add_of_pos_of_le hε (norm_le_sSupNormIm f z hD hB) + +/-- When the function `f` is bounded above on a vertical strip, then so is `F`. -/ +lemma F_BddAbove (f : ℂ → E) (ε : ℝ) (hε : ε > 0) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) : + BddAbove ((norm ∘ (F f ε)) '' (verticalClosedStrip 0 1)) := by + -- Rewriting goal + simp only [F, image_congr, comp_apply, map_mul, invInterpStrip] + rw [bddAbove_def] at * + rcases hB with ⟨B, hB⟩ + -- Using bound + use ((max 1 ((ε + sSupNormIm f 0) ^ (-(1 : ℝ)))) * max 1 ((ε + sSupNormIm f 1) ^ (-(1 : ℝ)))) * B + simp only [mem_image, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] + intros z hset + specialize hB (‖f z‖) (by simpa [image_congr, mem_image, comp_apply] using ⟨z, hset, rfl⟩) + -- Proof that the bound is correct + simp only [norm_smul, norm_mul, ← ofReal_add] + gcongr + -- Bounding individual terms + · by_cases hM0_one : 1 ≤ ε + sSupNormIm f 0 + -- `1 ≤ (sSupNormIm f 0)` + · apply le_trans _ (le_max_left _ _) + simp only [norm_eq_abs, abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε 0), sub_re, + one_re, Real.rpow_le_one_of_one_le_of_nonpos hM0_one (sub_nonpos.mpr hset.2)] + -- `0 < (sSupNormIm f 0) < 1` + · rw [not_le] at hM0_one; apply le_trans _ (le_max_right _ _) + simp only [norm_eq_abs, abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε 0), sub_re, + one_re] + apply Real.rpow_le_rpow_of_exponent_ge (sSupNormIm_eps_pos f hε 0) (le_of_lt hM0_one) _ + simp only [neg_le_sub_iff_le_add, le_add_iff_nonneg_left, hset.1] + · by_cases hM1_one : 1 ≤ ε + sSupNormIm f 1 + -- `1 ≤ sSupNormIm f 1` + · apply le_trans _ (le_max_left _ _) + simp only [norm_eq_abs, abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε 1), sub_re, + one_re, neg_re, Real.rpow_le_one_of_one_le_of_nonpos + hM1_one (Right.neg_nonpos_iff.mpr hset.1)] + -- `0 < sSupNormIm f 1 < 1` + · rw [not_le] at hM1_one; apply le_trans _ (le_max_right _ _) + simp only [norm_eq_abs, abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε 1), sub_re, + one_re, neg_re, Real.rpow_le_rpow_of_exponent_ge (sSupNormIm_eps_pos f hε 1) + (le_of_lt hM1_one) (neg_le_neg_iff.mpr hset.2)] + +/-- Proof that `F` is bounded by one one the edges. -/ +lemma F_edge_le_one (f : ℂ → E) (ε : ℝ) (hε : ε > 0) (z : ℂ) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) (hz : z ∈ re ⁻¹' {0, 1}) : + ‖F f ε z‖ ≤ 1 := by + simp only [F, norm_smul, norm_eq_abs, map_mul, abs_cpow_eq_rpow_re_of_pos, + abs_invInterpStrip f z hε, sSupNormIm_eps_pos f hε 1, + sub_re, one_re, neg_re] + rcases hz with hz0 | hz1 + -- `z.re = 0` + · simp only [hz0, zero_sub, Real.rpow_neg_one, neg_zero, Real.rpow_zero, mul_one, + inv_mul_le_iff (sSupNormIm_eps_pos f hε 0)] + rw [← hz0] + apply le_of_lt (norm_lt_sSupNormIm_eps f ε hε _ _ hB) + simp only [verticalClosedStrip, mem_preimage, zero_le_one, left_mem_Icc, hz0] + -- `z.re = 1` + · rw [mem_singleton_iff] at hz1 + simp only [hz1, one_mul, Real.rpow_zero, sub_self, Real.rpow_neg_one, + inv_mul_le_iff (sSupNormIm_eps_pos f hε 1), mul_one] + rw [← hz1] + apply le_of_lt (norm_lt_sSupNormIm_eps f ε hε _ _ hB) + simp only [verticalClosedStrip, mem_preimage, zero_le_one, hz1, right_mem_Icc] + +theorem norm_mul_invInterpStrip_le_one_of_mem_verticalClosedStrip (f : ℂ → E) (ε : ℝ) (hε : 0 < ε) + (z : ℂ) (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) (hz : z ∈ verticalClosedStrip 0 1) : + ‖F f ε z‖ ≤ 1 := by + apply PhragmenLindelof.vertical_strip + (DiffContOnCl.smul (diffContOnCl_invInterpStrip f hε) hd) _ + (fun x hx ↦ F_edge_le_one f ε hε x hB (Or.inl hx)) + (fun x hx ↦ F_edge_le_one f ε hε x hB (Or.inr hx)) hz.1 hz.2 + use 0 + rw [sub_zero, div_one] + refine ⟨ Real.pi_pos, ?_⟩ + obtain ⟨BF, hBF⟩ := F_BddAbove f ε hε hB + simp only [comp_apply, mem_upperBounds, mem_image, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂] at hBF + use BF + rw [Asymptotics.isBigO_iff] + use 1 + rw [eventually_inf_principal] + apply eventually_of_forall + intro x hx + norm_num + exact (hBF x ((preimage_mono Ioo_subset_Icc_self) hx)).trans + ((le_of_lt (lt_add_one BF)).trans (Real.add_one_le_exp BF)) + +end invInterpStrip + +----- + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] (f : ℂ → E) + +/-- +The interpolation of `sSupNormIm` on the two boundaries. +In other words, this is the right side of the target inequality: +`|f(z)| ≤ |M(0) ^ (1-z)| * |M(1) ^ z|`. + +Note that if `(sSupNormIm f 0) = 0 ∨ (sSupNormIm f 1) = 0` then the power is not continuous +since `0 ^ 0 = 1`. Hence the use of `ite`. -/ +noncomputable def interpStrip (z : ℂ) : ℂ := + if (sSupNormIm f 0) = 0 ∨ (sSupNormIm f 1) = 0 + then 0 + else (sSupNormIm f 0) ^ (1-z) * (sSupNormIm f 1) ^ z + +/-- Rewrite for `InterpStrip` when `0 < sSupNormIm f 0` and `0 < sSupNormIm f 1`. -/ +lemma interpStrip_eq_of_pos (z : ℂ) (h0 : 0 < sSupNormIm f 0) (h1 : 0 < sSupNormIm f 1) : + interpStrip f z = (sSupNormIm f 0) ^ (1 - z) * (sSupNormIm f 1) ^ z := by + simp only [ne_of_gt h0, ne_of_gt h1, interpStrip, if_false, or_false] + +/-- Rewrite for `InterpStrip` when `0 = sSupNormIm f 0` or `0 = sSupNormIm f 1`. -/ +lemma interpStrip_eq_of_zero (z : ℂ) (h : sSupNormIm f 0 = 0 ∨ sSupNormIm f 1 = 0) : + interpStrip f z = 0 := + if_pos h + +/-- Rewrite for `InterpStrip` on the open vertical strip. -/ +lemma interpStrip_eq_of_mem_verticalStrip (z : ℂ) (hz : z ∈ verticalStrip 0 1): + interpStrip f z = (sSupNormIm f 0) ^ (1 - z) * (sSupNormIm f 1) ^ z := by + by_cases h : sSupNormIm f 0 = 0 ∨ sSupNormIm f 1 = 0 + · rw [interpStrip_eq_of_zero _ z h] + rcases h with h0 | h1 + · simp only [h0, ofReal_zero, zero_eq_mul, cpow_eq_zero_iff, ne_eq, true_and, ofReal_eq_zero] + left + rw [sub_eq_zero, eq_comm] + simp only [ne_eq, ext_iff, one_re, ne_of_lt hz.2, or_iff_left, false_and, not_false_eq_true] + · simp only [h1, ofReal_zero, zero_eq_mul, cpow_eq_zero_iff, ofReal_eq_zero, ne_eq, true_and] + right + rw [eq_comm] + simp only [ne_eq, ext_iff, zero_re, ne_of_lt hz.1, or_iff_left, false_and, not_false_eq_true] + · push_neg at h + replace h : (0 < sSupNormIm f 0) ∧ (0 < sSupNormIm f 1) := + ⟨(lt_of_le_of_ne (sSupNormIm_nonneg f 0) (ne_comm.mp h.1)), + (lt_of_le_of_ne (sSupNormIm_nonneg f 1) (ne_comm.mp h.2))⟩ + exact interpStrip_eq_of_pos f z h.1 h.2 + +lemma diffContOnCl_interpStrip : + DiffContOnCl ℂ (interpStrip f) (verticalStrip 0 1) := by + by_cases h : sSupNormIm f 0 = 0 ∨ sSupNormIm f 1 = 0 + -- Case everywhere 0 + · eta_expand; simp_rw [interpStrip_eq_of_zero f _ h]; exact diffContOnCl_const + -- Case nowhere 0 + · push_neg at h + rcases h with ⟨h0, h1⟩ + rw [ne_comm] at h0 h1 + apply Differentiable.diffContOnCl + intro z + eta_expand + simp_rw [interpStrip_eq_of_pos f _ (lt_of_le_of_ne (sSupNormIm_nonneg f 0) h0) + (lt_of_le_of_ne (sSupNormIm_nonneg f 1) h1)] + refine DifferentiableAt.mul ?_ ?_ + · apply DifferentiableAt.const_cpow (DifferentiableAt.const_sub (differentiableAt_id') 1) _ + left; simp only [Ne, ofReal_eq_zero]; rwa [eq_comm] + · refine DifferentiableAt.const_cpow ?_ ?_ + · apply differentiableAt_id' + · left; simp only [Ne, ofReal_eq_zero]; rwa [eq_comm] + +lemma norm_le_interpStrip_of_mem_verticalClosedStrip_eps (ε : ℝ) (hε : ε > 0) (z : ℂ) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) + (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) (hz : z ∈ verticalClosedStrip 0 1) : + ‖f z‖ ≤ ‖((ε + sSupNormIm f 0) ^ (1-z) * (ε + sSupNormIm f 1) ^ z : ℂ)‖ := by + simp only [F, abs_invInterpStrip _ _ hε, norm_smul, norm_mul, norm_eq_abs, + ← ofReal_add, abs_cpow_eq_rpow_re_of_pos (sSupNormIm_eps_pos f hε _) _, sub_re, one_re] + rw [← mul_inv_le_iff, ← one_mul (((ε + sSupNormIm f 1) ^ z.re)), ← mul_inv_le_iff', + ← Real.rpow_neg_one, ← Real.rpow_neg_one] + · simp only [← Real.rpow_mul (le_of_lt (sSupNormIm_eps_pos f hε _)), + mul_neg, mul_one, neg_sub, mul_assoc] + simpa [F, abs_invInterpStrip _ _ hε, norm_smul, mul_comm] using + norm_mul_invInterpStrip_le_one_of_mem_verticalClosedStrip f ε hε z hd hB hz + · simp only [Real.rpow_pos_of_pos (sSupNormIm_eps_pos f hε _) z.re] + · simp only [Real.rpow_pos_of_pos (sSupNormIm_eps_pos f hε _) (1-z.re)] + +lemma eventuallyle (z : ℂ) (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) + (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) (hz : z ∈ verticalStrip 0 1) : + (fun _ : ℝ ↦ ‖f z‖) ≤ᶠ[𝓝[>] 0] + (fun ε ↦ ‖((ε + sSupNormIm f 0) ^ (1 - z) * (ε + sSupNormIm f 1) ^ z : ℂ)‖) := by + filter_upwards [self_mem_nhdsWithin] with ε (hε : 0 < ε) using + norm_le_interpStrip_of_mem_verticalClosedStrip_eps f ε hε z hB hd + (mem_of_mem_of_subset hz (preimage_mono Ioo_subset_Icc_self)) + +lemma norm_le_interpStrip_of_mem_verticalStrip_zero (z : ℂ) + (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) (hz : z ∈ verticalStrip 0 1) : + ‖f z‖ ≤ ‖interpStrip f z‖ := by + apply tendsto_le_of_eventuallyLE _ _ (eventuallyle f z hB hd hz) + · apply tendsto_inf_left + simp only [tendsto_const_nhds_iff] + -- Proof that we can let epsilon tend to zero. + · rw [interpStrip_eq_of_mem_verticalStrip _ _ hz] + convert ContinuousWithinAt.tendsto _ using 2 + · simp only [ofReal_zero, zero_add] + · simp_rw [← ofReal_add, norm_eq_abs] + have : ∀ x ∈ Ioi 0, (x + sSupNormIm f 0) ^ (1 - z.re) * (x + (sSupNormIm f 1)) ^ z.re + = abs (↑(x + sSupNormIm f 0) ^ (1 - z) * ↑(x + sSupNormIm f 1) ^ z) := by + intro x hx + simp only [map_mul] + repeat rw [abs_cpow_eq_rpow_re_of_nonneg (le_of_lt (sSupNormIm_eps_pos f hx _)) _] + · simp only [sub_re, one_re] + · simpa using (ne_comm.mpr (ne_of_lt hz.1)) + · simpa [sub_eq_zero] using (ne_comm.mpr (ne_of_lt hz.2)) + apply tendsto_nhdsWithin_congr this _ + simp only [zero_add] + rw [map_mul, abs_cpow_eq_rpow_re_of_nonneg (sSupNormIm_nonneg _ _) _, + abs_cpow_eq_rpow_re_of_nonneg (sSupNormIm_nonneg _ _) _] + · apply Tendsto.mul + · apply Tendsto.rpow_const + · nth_rw 2 [← zero_add (sSupNormIm f 0)] + exact Tendsto.add_const (sSupNormIm f 0) (tendsto_nhdsWithin_of_tendsto_nhds + (Continuous.tendsto continuous_id' _)) + · right; simp only [sub_nonneg, le_of_lt hz.2] + · apply Tendsto.rpow_const + · nth_rw 2 [← zero_add (sSupNormIm f 1)] + exact Tendsto.add_const (sSupNormIm f 1) (tendsto_nhdsWithin_of_tendsto_nhds + (Continuous.tendsto continuous_id' _)) + · right; simp only [sub_nonneg, le_of_lt hz.1] + · simpa using (ne_comm.mpr (ne_of_lt hz.1)) + · simpa [sub_eq_zero] using (ne_comm.mpr (ne_of_lt hz.2)) + +/-- +**Hadamard three-line theorem** on `re ⁻¹' [0,1]`: If `f` is a bounded function, continuous on the +closed strip `re ⁻¹' [0,1]` and differentiable on open strip `re ⁻¹' (0,1)`, then for +`M(x) := sup ((norm ∘ f) '' (re ⁻¹' {x}))` we have that for all `z` in the closed strip +`re ⁻¹' [0,1]` the inequality `‖f(z)‖ ≤ M(0) ^ (1 - z.re) * M(1) ^ z.re` holds. -/ +lemma norm_le_interpStrip_of_mem_verticalClosedStrip (f : ℂ → E) {z : ℂ} + (hz : z ∈ verticalClosedStrip 0 1) (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) : + ‖f z‖ ≤ ‖interpStrip f z‖ := by + apply le_on_closure (fun w hw ↦ norm_le_interpStrip_of_mem_verticalStrip_zero f w hd hB hw) + (Continuous.comp_continuousOn' continuous_norm hd.2) + (Continuous.comp_continuousOn' continuous_norm (diffContOnCl_interpStrip f).2) + rwa [verticalClosedStrip, ← closure_Ioo zero_ne_one, ← closure_preimage_re] at hz + +/-- **Hadamard three-line theorem** on `re ⁻¹' [0,1]` (Variant in simpler terms): Let `f` be a +bounded function, continuous on the closed strip `re ⁻¹' [0,1]` and differentiable on open strip +`re ⁻¹' (0,1)`. If, for all `z.re = 0`, `‖f z‖ ≤ a` for some `a ∈ ℝ` and, similarly, for all +`z.re = 1`, `‖f z‖ ≤ b` for some `b ∈ ℝ` then for all `z` in the closed strip +`re ⁻¹' [0,1]` the inequality `‖f(z)‖ ≤ a ^ (1 - z.re) * b ^ z.re` holds. -/ +lemma norm_le_interp_of_mem_verticalClosedStrip' (f : ℂ → E) {z : ℂ} {a b : ℝ} + (hz : z ∈ verticalClosedStrip 0 1) (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) + (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) + (ha : ∀ z ∈ re ⁻¹' {0}, ‖f z‖ ≤ a) (hb : ∀ z ∈ re ⁻¹' {1}, ‖f z‖ ≤ b) : + ‖f z‖ ≤ a ^ (1 - z.re) * b ^ z.re := by + have : ‖interpStrip f z‖ ≤ (sSupNormIm f 0) ^ (1 - z.re) * (sSupNormIm f 1) ^ z.re := by + by_cases h : sSupNormIm f 0 = 0 ∨ sSupNormIm f 1 = 0 + · rw [interpStrip_eq_of_zero f z h, norm_zero, mul_nonneg_iff] + left + exact ⟨Real.rpow_nonneg (sSupNormIm_nonneg f _) _, + Real.rpow_nonneg (sSupNormIm_nonneg f _) _ ⟩ + · push_neg at h + rcases h with ⟨h0, h1⟩ + rw [ne_comm] at h0 h1 + simp_rw [interpStrip_eq_of_pos f _ (lt_of_le_of_ne (sSupNormIm_nonneg f 0) h0) + (lt_of_le_of_ne (sSupNormIm_nonneg f 1) h1)] + simp only [norm_eq_abs, map_mul] + rw [abs_cpow_eq_rpow_re_of_pos ((Ne.le_iff_lt h0).mp (sSupNormIm_nonneg f _)) _] + 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 (Real.rpow_nonneg _ _) + specialize hb 1 + simp only [mem_preimage, one_re, mem_singleton_iff, forall_true_left] at hb + exact (norm_nonneg _).trans hb + · apply Real.rpow_le_rpow (sSupNormIm_nonneg f _) _ (sub_nonneg.mpr hz.2) + · rw [sSupNormIm] + apply csSup_le _ + · simpa [comp_apply, mem_image, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂] using ha + · use ‖(f 0)‖, 0 + simp only [mem_preimage, zero_re, mem_singleton_iff, comp_apply, + and_self] + · apply Real.rpow_le_rpow (sSupNormIm_nonneg f _) _ hz.1 + · rw [sSupNormIm] + apply csSup_le _ + · simpa [comp_apply, mem_image, forall_exists_index, + and_imp, forall_apply_eq_imp_iff₂] using hb + · use ‖(f 1)‖, 1 + simp only [mem_preimage, one_re, mem_singleton_iff, comp_apply, + and_self] + +end HadamardThreeLines +end Complex From d9cc9fafd24512795e3b25143436a81a7ecf1537 Mon Sep 17 00:00:00 2001 From: Etienne Date: Tue, 7 May 2024 09:50:33 +0000 Subject: [PATCH 06/52] feat: operations on indicatorConstLp (#12464) Add lemmas to rewrite the sum/difference of two `indicatorConstLp` with same support. The distance between to Lp functions is the same as the norm of their difference (even if $p < 1$ and it's not a norm). Co-authored-by: Etienne <66847262+EtienneC30@users.noreply.github.com> --- Mathlib/MeasureTheory/Function/LpSpace.lean | 22 ++++++++++++++++++--- 1 file changed, 19 insertions(+), 3 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index f5d39e12b381e..7b6b02d7f5b28 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -306,6 +306,8 @@ protected theorem edist_dist (f g : Lp E p μ) : edist f g = .ofReal (dist f g) protected theorem dist_edist (f g : Lp E p μ) : dist f g = (edist f g).toReal := MeasureTheory.Lp.dist_def .. +theorem dist_eq_norm (f g : Lp E p μ) : dist f g = ‖f - g‖ := rfl + @[simp] theorem edist_toLp_toLp (f g : α → E) (hf : Memℒp f p μ) (hg : Memℒp g p μ) : edist (hf.toLp f) (hg.toLp g) = snorm (f - g) p μ := by @@ -762,6 +764,20 @@ def indicatorConstLp (p : ℝ≥0∞) (hs : MeasurableSet s) (hμs : μ s ≠ Memℒp.toLp (s.indicator fun _ => c) (memℒp_indicator_const p hs c (Or.inr hμs)) #align measure_theory.indicator_const_Lp MeasureTheory.indicatorConstLp +/-- A version of `Set.indicator_add` for `MeasureTheory.indicatorConstLp`.-/ +theorem indicatorConstLp_add {c' : E} : + indicatorConstLp p hs hμs c + indicatorConstLp p hs hμs c' = + indicatorConstLp p hs hμs (c + c') := by + simp_rw [indicatorConstLp, ← Memℒp.toLp_add, indicator_add] + rfl + +/-- A version of `Set.indicator_sub` for `MeasureTheory.indicatorConstLp`.-/ +theorem indicatorConstLp_sub {c' : E} : + indicatorConstLp p hs hμs c - indicatorConstLp p hs hμs c' = + indicatorConstLp p hs hμs (c - c') := by + simp_rw [indicatorConstLp, ← Memℒp.toLp_sub, indicator_sub] + rfl + theorem indicatorConstLp_coeFn : ⇑(indicatorConstLp p hs hμs c) =ᵐ[μ] s.indicator fun _ => c := Memℒp.coeFn_toLp (memℒp_indicator_const p hs c (Or.inr hμs)) #align measure_theory.indicator_const_Lp_coe_fn MeasureTheory.indicatorConstLp_coeFn @@ -805,13 +821,13 @@ theorem norm_indicatorConstLp_le : ENNReal.toReal_rpow, ENNReal.ofReal_toReal] exact ENNReal.rpow_ne_top_of_nonneg (by positivity) hμs -theorem edist_indicatorConstLp_eq_nnnorm {t : Set α} (ht : MeasurableSet t) (hμt : μ t ≠ ∞) : +theorem edist_indicatorConstLp_eq_nnnorm {t : Set α} {ht : MeasurableSet t} {hμt : μ t ≠ ∞} : edist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = ‖indicatorConstLp p (hs.symmDiff ht) (measure_symmDiff_ne_top hμs hμt) c‖₊ := by unfold indicatorConstLp rw [Lp.edist_toLp_toLp, snorm_indicator_sub_indicator, Lp.coe_nnnorm_toLp] -theorem dist_indicatorConstLp_eq_norm {t : Set α} (ht : MeasurableSet t) (hμt : μ t ≠ ∞) : +theorem dist_indicatorConstLp_eq_norm {t : Set α} {ht : MeasurableSet t} {hμt : μ t ≠ ∞} : dist (indicatorConstLp p hs hμs c) (indicatorConstLp p ht hμt c) = ‖indicatorConstLp p (hs.symmDiff ht) (measure_symmDiff_ne_top hμs hμt) c‖ := by rw [Lp.dist_edist, edist_indicatorConstLp_eq_nnnorm, ENNReal.coe_toReal, Lp.coe_nnnorm] @@ -1616,7 +1632,7 @@ theorem ae_tendsto_of_cauchy_snorm [CompleteSpace E] {f : ℕ → α → E} refine' cauchySeq_of_le_tendsto_0 (fun n => (B n).toReal) _ _ · intro n m N hnN hmN specialize hx N n m hnN hmN - rw [dist_eq_norm, ← ENNReal.toReal_ofReal (norm_nonneg _), + rw [_root_.dist_eq_norm, ← ENNReal.toReal_ofReal (norm_nonneg _), ENNReal.toReal_le_toReal ENNReal.ofReal_ne_top (ENNReal.ne_top_of_tsum_ne_top hB N)] rw [← ofReal_norm_eq_coe_nnnorm] at hx exact hx.le From aa79770644c6a6cdc0a83319df79793eb2c0029d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 7 May 2024 10:24:03 +0000 Subject: [PATCH 07/52] feat: `withDensity` of an s-finite measure is s-finite (#12591) --- .../MeasureTheory/Measure/WithDensity.lean | 45 +++++++++++++++++++ 1 file changed, 45 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index 7668015af6ed2..5fb7b8dbefb5b 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -585,6 +585,51 @@ lemma SigmaFinite.withDensity_ofReal [SigmaFinite μ] {f : α → ℝ} (hf : AEM SigmaFinite (μ.withDensity (fun x ↦ ENNReal.ofReal (f x))) := by exact SigmaFinite.withDensity_of_ne_top hf.ennreal_ofReal (ae_of_all _ (by simp)) +section SFinite + +/-- Auxiliary lemma for `sFinite_withDensity_of_measurable`. -/ +lemma sFinite_withDensity_of_sigmaFinite_of_measurable (μ : Measure α) [SigmaFinite μ] + {f : α → ℝ≥0∞} (hf : Measurable f) : + SFinite (μ.withDensity f) := by + let s := {x | f x = ∞} + have hs : MeasurableSet s := hf (measurableSet_singleton _) + rw [← restrict_add_restrict_compl (μ := μ.withDensity f) hs, restrict_withDensity hs, + restrict_withDensity hs.compl, ← withDensity_indicator hs, ← withDensity_indicator hs.compl] + have h1 : SFinite (μ.withDensity (s.indicator f)) := by + have h_eq_sum : s.indicator f = ∑' n : ℕ, s.indicator 1 := by + ext x + rw [tsum_apply] + swap; · rw [Pi.summable]; exact fun _ ↦ ENNReal.summable + simp_rw [Set.indicator_apply] + split_ifs with hx + · simp only [Set.mem_setOf_eq, s] at hx + simp [hx, ENNReal.tsum_const_eq_top_of_ne_zero] + · simp + rw [h_eq_sum, withDensity_tsum (fun _ ↦ measurable_one.indicator hs)] + have : SigmaFinite (μ.withDensity (s.indicator 1)) := by + refine SigmaFinite.withDensity_of_ne_top' (measurable_one.indicator hs).aemeasurable + (fun x ↦ ?_) + simp only [Set.indicator_apply, Pi.one_apply, ne_eq] + split_ifs with h <;> simp [h] + infer_instance + have h2 : SigmaFinite (μ.withDensity (sᶜ.indicator f)) := by + refine SigmaFinite.withDensity_of_ne_top' (hf.indicator hs.compl).aemeasurable (fun x ↦ ?_) + simp only [Set.indicator_apply, Set.mem_compl_iff, Set.mem_setOf_eq, ite_not, ne_eq, s] + split_ifs with h <;> simp [h] + infer_instance + +/-- If `μ` is s-finite and `f` is measurable, then `μ.withDensity f` is s-finite. +TODO: extend this to all functions and make it an instance. -/ +lemma sFinite_withDensity_of_measurable (μ : Measure α) [SFinite μ] + {f : α → ℝ≥0∞} (hf : Measurable f) : + SFinite (μ.withDensity f) := by + rw [← sum_sFiniteSeq μ, withDensity_sum] + have : ∀ n, SFinite ((sFiniteSeq μ n).withDensity f) := + fun n ↦ sFinite_withDensity_of_sigmaFinite_of_measurable _ hf + infer_instance + +end SFinite + variable [TopologicalSpace α] [OpensMeasurableSpace α] [IsLocallyFiniteMeasure μ] lemma IsLocallyFiniteMeasure.withDensity_coe {f : α → ℝ≥0} (hf : Continuous f) : From 740f52cf67acc2daedb94b53329bbb03870fad94 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 7 May 2024 10:24:04 +0000 Subject: [PATCH 08/52] feat: add `*.of_comp_iff` lemmas for `Inducing`, `Embedding`, etc. (#12639) --- .../Topology/FiberBundle/Constructions.lean | 2 +- Mathlib/Topology/LocalAtTarget.lean | 4 ++-- Mathlib/Topology/Maps.lean | 21 ++++++++++++------- .../UniformSpace/UniformEmbedding.lean | 10 +++++++++ 4 files changed, 27 insertions(+), 10 deletions(-) diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index 7974a3471d076..26722ef4a374c 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -249,7 +249,7 @@ variable [∀ x, Zero (E₁ x)] [∀ x, Zero (E₂ x)] [∀ x : B, TopologicalSp /-- The product of two fiber bundles is a fiber bundle. -/ noncomputable instance FiberBundle.prod : FiberBundle (F₁ × F₂) (E₁ ×ᵇ E₂) where totalSpaceMk_inducing' b := by - rw [(Prod.inducing_diag F₁ E₁ F₂ E₂).inducing_iff] + rw [← (Prod.inducing_diag F₁ E₁ F₂ E₂).of_comp_iff] exact (totalSpaceMk_inducing F₁ E₁ b).prod_map (totalSpaceMk_inducing F₂ E₂ b) trivializationAtlas' := { e | ∃ (e₁ : Trivialization F₁ (π F₁ E₁)) (e₂ : Trivialization F₂ (π F₂ E₂)) diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index cefbc8f7bea97..5fbec387d429a 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -28,7 +28,7 @@ variable {s : Set β} {ι : Type*} {U : ι → Opens β} (hU : iSup U = ⊤) theorem Set.restrictPreimage_inducing (s : Set β) (h : Inducing f) : Inducing (s.restrictPreimage f) := by - simp_rw [inducing_subtype_val.inducing_iff, inducing_iff_nhds, restrictPreimage, + simp_rw [← inducing_subtype_val.of_comp_iff, inducing_iff_nhds, restrictPreimage, MapsTo.coe_restrict, restrict_eq, ← @Filter.comap_comap _ _ _ _ _ f, Function.comp_apply] at h ⊢ intro a rw [← h, ← inducing_subtype_val.nhds_eq_comap] @@ -128,7 +128,7 @@ theorem isClosedMap_iff_isClosedMap_of_iSup_eq_top : theorem inducing_iff_inducing_of_iSup_eq_top (h : Continuous f) : Inducing f ↔ ∀ i, Inducing ((U i).1.restrictPreimage f) := by - simp_rw [inducing_subtype_val.inducing_iff, inducing_iff_nhds, restrictPreimage, + simp_rw [← inducing_subtype_val.of_comp_iff, inducing_iff_nhds, restrictPreimage, MapsTo.coe_restrict, restrict_eq, ← @Filter.comap_comap _ _ _ _ _ f] constructor · intro H i x diff --git a/Mathlib/Topology/Maps.lean b/Mathlib/Topology/Maps.lean index 507c0f2ad7c14..8f086c07f39ba 100644 --- a/Mathlib/Topology/Maps.lean +++ b/Mathlib/Topology/Maps.lean @@ -66,6 +66,12 @@ protected theorem Inducing.comp (hg : Inducing g) (hf : Inducing f) : ⟨by rw [hf.induced, hg.induced, induced_compose]⟩ #align inducing.comp Inducing.comp +theorem Inducing.of_comp_iff (hg : Inducing g) : + Inducing (g ∘ f) ↔ Inducing f := by + refine ⟨fun h ↦ ?_, hg.comp⟩ + rw [inducing_iff, hg.induced, induced_compose, h.induced] +#align inducing.inducing_iff Inducing.of_comp_iff + theorem inducing_of_inducing_compose (hf : Continuous f) (hg : Continuous g) (hgf : Inducing (g ∘ f)) : Inducing f := ⟨le_antisymm (by rwa [← continuous_iff_le_induced]) @@ -137,13 +143,6 @@ protected theorem continuous (hf : Inducing f) : Continuous f := hf.continuous_iff.mp continuous_id #align inducing.continuous Inducing.continuous -protected theorem inducing_iff (hg : Inducing g) : - Inducing f ↔ Inducing (g ∘ f) := by - refine' ⟨fun h => hg.comp h, fun hgf => inducing_of_inducing_compose _ hg.continuous hgf⟩ - rw [hg.continuous_iff] - exact hgf.continuous -#align inducing.inducing_iff Inducing.inducing_iff - theorem closure_eq_preimage_closure_image (hf : Inducing f) (s : Set X) : closure s = f ⁻¹' closure (f '' s) := by ext x @@ -206,6 +205,9 @@ protected theorem Embedding.comp (hg : Embedding g) (hf : Embedding f) : { hg.toInducing.comp hf.toInducing with inj := fun _ _ h => hf.inj <| hg.inj h } #align embedding.comp Embedding.comp +theorem Embedding.of_comp_iff (hg : Embedding g) : Embedding (g ∘ f) ↔ Embedding f := by + simp_rw [embedding_iff, hg.toInducing.of_comp_iff, hg.inj.of_comp_iff f] + theorem embedding_of_embedding_compose (hf : Continuous f) (hg : Continuous g) (hgf : Embedding (g ∘ f)) : Embedding f := { induced := (inducing_of_inducing_compose hf hg hgf.toInducing).induced @@ -698,6 +700,11 @@ theorem comp (hg : ClosedEmbedding g) (hf : ClosedEmbedding f) : ⟨hg.toEmbedding.comp hf.toEmbedding, (hg.isClosedMap.comp hf.isClosedMap).isClosed_range⟩ #align closed_embedding.comp ClosedEmbedding.comp +theorem of_comp_iff (hg : ClosedEmbedding g) : + ClosedEmbedding (g ∘ f) ↔ ClosedEmbedding f := by + simp_rw [closedEmbedding_iff, hg.toEmbedding.of_comp_iff, Set.range_comp, + ← hg.closed_iff_image_closed] + theorem closure_image_eq (hf : ClosedEmbedding f) (s : Set X) : closure (f '' s) = f '' closure s := hf.isClosedMap.closure_image_eq_of_continuous hf.continuous s diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index e27a9806326cb..c62d495523247 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -73,6 +73,12 @@ theorem UniformInducing.comp {g : β → γ} (hg : UniformInducing g) {f : α ⟨by rw [← hf.1, ← hg.1, comap_comap]; rfl⟩ #align uniform_inducing.comp UniformInducing.comp +theorem UniformInducing.of_comp_iff {g : β → γ} (hg : UniformInducing g) {f : α → β} : + UniformInducing (g ∘ f) ↔ UniformInducing f := by + refine ⟨fun h ↦ ?_, hg.comp⟩ + rw [uniformInducing_iff, ← hg.comap_uniformity, comap_comap, ← h.comap_uniformity, + Function.comp, Function.comp] + theorem UniformInducing.basis_uniformity {f : α → β} (hf : UniformInducing f) {ι : Sort*} {p : ι → Prop} {s : ι → Set (β × β)} (H : (𝓤 β).HasBasis p s) : (𝓤 α).HasBasis p fun i => Prod.map f f ⁻¹' s i := @@ -182,6 +188,10 @@ theorem UniformEmbedding.comp {g : β → γ} (hg : UniformEmbedding g) {f : α { hg.toUniformInducing.comp hf.toUniformInducing with inj := hg.inj.comp hf.inj } #align uniform_embedding.comp UniformEmbedding.comp +theorem UniformEmbedding.of_comp_iff {g : β → γ} (hg : UniformEmbedding g) {f : α → β} : + UniformEmbedding (g ∘ f) ↔ UniformEmbedding f := by + simp_rw [uniformEmbedding_iff, hg.toUniformInducing.of_comp_iff, hg.inj.of_comp_iff f] + theorem Equiv.uniformEmbedding {α β : Type*} [UniformSpace α] [UniformSpace β] (f : α ≃ β) (h₁ : UniformContinuous f) (h₂ : UniformContinuous f.symm) : UniformEmbedding f := uniformEmbedding_iff'.2 ⟨f.injective, h₁, by rwa [← Equiv.prodCongr_apply, ← map_equiv_symm]⟩ From 466ce46b56681a73bf81a3593b08336366378597 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Pol=27tta=20/=20Miyahara=20K=C5=8D?= Date: Tue, 7 May 2024 10:24:05 +0000 Subject: [PATCH 09/52] =?UTF-8?q?refactor(Mathlib/Init/Data/List/Instances?= =?UTF-8?q?):=20deduplicate=20`Decidable=20(=E2=88=80=20x,=20x=20=E2=88=88?= =?UTF-8?q?=20l=20=E2=86=92=20p=20x)`=20&=20`Decidable=20(=E2=88=83=20x,?= =?UTF-8?q?=20x=20=E2=88=88=20l=20=E2=86=92=20p=20x)`=20instances=20(#1272?= =?UTF-8?q?2)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Data/Int/Range.lean | 1 - Mathlib/Init/Data/List/Instances.lean | 30 +++------------------------ 2 files changed, 3 insertions(+), 28 deletions(-) diff --git a/Mathlib/Data/Int/Range.lean b/Mathlib/Data/Int/Range.lean index 9a26dcea59bd0..1e0d5673a77c1 100644 --- a/Mathlib/Data/Int/Range.lean +++ b/Mathlib/Data/Int/Range.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau -/ import Mathlib.Algebra.Order.Ring.Int -import Mathlib.Init.Data.List.Instances #align_import data.int.range from "leanprover-community/mathlib"@"7b78d1776212a91ecc94cf601f83bdcc46b04213" diff --git a/Mathlib/Init/Data/List/Instances.lean b/Mathlib/Init/Data/List/Instances.lean index 510ab6be754a1..becd08b9a83c5 100644 --- a/Mathlib/Init/Data/List/Instances.lean +++ b/Mathlib/Init/Data/List/Instances.lean @@ -59,32 +59,8 @@ instance instAlternative : Alternative List.{u} where #noalign list.bin_tree_to_list -variable {α : Type u} {p : α → Prop} [DecidablePred p] - --- To work around lean4#2552, we call specific `Decidable` instances and use `match` on them, --- as opposed to using `if`. -instance decidableExistsMem : ∀ (l : List α), Decidable (∃ x ∈ l, p x) - | [] => isFalse (by simp) - | x::xs => - match ‹DecidablePred p› x with - | isTrue h₁ => isTrue ⟨x, mem_cons_self _ _, h₁⟩ - | isFalse h₁ => match decidableExistsMem xs with - | isTrue h₂ => isTrue <| by - rcases h₂ with ⟨y, hm, hp⟩ - exact ⟨y, mem_cons_of_mem _ hm, hp⟩ - | isFalse h₂ => isFalse <| by - rintro ⟨y, hm, hp⟩ - cases mem_cons.1 hm with - | inl h => rw [h] at hp; contradiction - | inr h => exact absurd ⟨y, h, hp⟩ h₂ -#align list.decidable_bex List.decidableExistsMem - -instance decidableForallMem (l : List α) : Decidable (∀ x ∈ l, p x) := - match (inferInstance : Decidable <| ∃ x ∈ l, ¬ p x) with - | isFalse h => isTrue fun x hx => match ‹DecidablePred p› x with - | isTrue h' => h' - | isFalse h' => False.elim <| h ⟨x, hx, h'⟩ - | isTrue h => isFalse <| let ⟨x, h, np⟩ := h; fun al => np (al x h) -#align list.decidable_ball List.decidableForallMem +#align list.decidable_bex List.decidableBEx + +#align list.decidable_ball List.decidableBAll end List From 248b41f5309cfef76df808284e925ae83ef49aa8 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Tue, 7 May 2024 10:24:07 +0000 Subject: [PATCH 10/52] refactor(Topology/Connected/TotallyDisconnected): use Set.Pairwise (#12724) Co-authored-by: Moritz Firsching --- Mathlib/Topology/Connected/Separation.lean | 2 +- Mathlib/Topology/Connected/TotallyDisconnected.lean | 9 ++++----- Mathlib/Topology/Separation.lean | 4 ++-- 3 files changed, 7 insertions(+), 8 deletions(-) diff --git a/Mathlib/Topology/Connected/Separation.lean b/Mathlib/Topology/Connected/Separation.lean index 17dcd75b7ff1c..57db6948dc729 100644 --- a/Mathlib/Topology/Connected/Separation.lean +++ b/Mathlib/Topology/Connected/Separation.lean @@ -22,7 +22,7 @@ section TotallySeparated /-- A totally separated space is T2. -/ instance TotallySeparatedSpace.t2Space [TotallySeparatedSpace X] : T2Space X where t2 x y h := by - obtain ⟨u, v, h₁, h₂, h₃, h₄, _, h₅⟩ := isTotallySeparated_univ x trivial y trivial h + obtain ⟨u, v, h₁, h₂, h₃, h₄, _, h₅⟩ := isTotallySeparated_univ trivial trivial h exact ⟨u, v, h₁, h₂, h₃, h₄, h₅⟩ end TotallySeparated diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index 86e84f906566e..3f143cba75f26 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -185,12 +185,11 @@ end TotallyDisconnected section TotallySeparated --- todo: reformulate using `Set.Pairwise` /-- A set `s` is called totally separated if any two points of this set can be separated by two disjoint open sets covering `s`. -/ def IsTotallySeparated (s : Set α) : Prop := - ∀ x ∈ s, ∀ y ∈ s, x ≠ y → - ∃ u v : Set α, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ s ⊆ u ∪ v ∧ Disjoint u v + Set.Pairwise s fun x y => + ∃ u v : Set α, IsOpen u ∧ IsOpen v ∧ x ∈ u ∧ y ∈ v ∧ s ⊆ u ∪ v ∧ Disjoint u v #align is_totally_separated IsTotallySeparated theorem isTotallySeparated_empty : IsTotallySeparated (∅ : Set α) := fun _ => False.elim @@ -207,7 +206,7 @@ theorem isTotallyDisconnected_of_isTotallySeparated {s : Set α} (H : IsTotallyS obtain ⟨u : Set α, v : Set α, hu : IsOpen u, hv : IsOpen v, hxu : x ∈ u, hyv : y ∈ v, hs : s ⊆ u ∪ v, huv⟩ := - H x (hts x_in) y (hts y_in) h + H (hts x_in) (hts y_in) h refine' (ht _ _ hu hv (hts.trans hs) ⟨x, x_in, hxu⟩ ⟨y, y_in, hyv⟩).ne_empty _ rw [huv.inter_eq, inter_empty] #align is_totally_disconnected_of_is_totally_separated isTotallyDisconnected_of_isTotallySeparated @@ -239,7 +238,7 @@ theorem exists_isClopen_of_totally_separated {α : Type*} [TopologicalSpace α] [TotallySeparatedSpace α] {x y : α} (hxy : x ≠ y) : ∃ U : Set α, IsClopen U ∧ x ∈ U ∧ y ∈ Uᶜ := by obtain ⟨U, V, hU, hV, Ux, Vy, f, disj⟩ := - TotallySeparatedSpace.isTotallySeparated_univ x (Set.mem_univ x) y (Set.mem_univ y) hxy + TotallySeparatedSpace.isTotallySeparated_univ (Set.mem_univ x) (Set.mem_univ y) hxy have hU := isClopen_inter_of_disjoint_cover_clopen isClopen_univ f hU hV disj rw [univ_inter _] at hU rw [← Set.subset_compl_iff_disjoint_right, subset_compl_comm] at disj diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 51154f997be83..8549e9923892e 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -2404,8 +2404,8 @@ theorem compact_t2_tot_disc_iff_tot_sep : TotallyDisconnectedSpace X ↔ Totally rw [connectedComponent_eq_iInter_isClopen, mem_iInter] rintro ⟨w : Set X, hw : IsClopen w, hy : y ∈ w⟩ by_contra hx - exact hyp wᶜ w hw.1.isOpen_compl hw.2 hx hy (@isCompl_compl _ w _).symm.codisjoint.top_le - disjoint_compl_left + exact hyp ⟨wᶜ, w, hw.1.isOpen_compl, hw.2, hx, hy, (@isCompl_compl _ w _).symm.codisjoint.top_le, + disjoint_compl_left⟩ #align compact_t2_tot_disc_iff_tot_sep compact_t2_tot_disc_iff_tot_sep variable [TotallyDisconnectedSpace X] From 787f9595dd18b3430e11ff28ef9e6284d63c1981 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Tue, 7 May 2024 11:27:30 +0000 Subject: [PATCH 11/52] feat: easy results about ring of integers. (#12704) Those have been forgotten in #12386. --- Mathlib/NumberTheory/NumberField/Basic.lean | 8 ++++++++ Mathlib/NumberTheory/NumberField/Units/Basic.lean | 2 ++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index 674f9c0cdf83e..47cc68d615b02 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -94,6 +94,10 @@ instance : NoZeroSMulDivisors (𝓞 K) K := inferInstanceAs (NoZeroSMulDivisors (integralClosure _ _) _) instance : Nontrivial (𝓞 K) := inferInstanceAs (Nontrivial (integralClosure _ _)) +instance {L : Type*} [Ring L] [Algebra K L] : Algebra (𝓞 K) L := + inferInstanceAs (Algebra (integralClosure _ _) L) +instance {L : Type*} [Ring L] [Algebra K L] : IsScalarTower (𝓞 K) K L := + inferInstanceAs (IsScalarTower (integralClosure _ _) K L) variable {K} @@ -112,6 +116,10 @@ lemma coe_eq_algebraMap (x : 𝓞 K) : (x : K) = algebraMap _ _ x := rfl theorem ext_iff {x y : 𝓞 K} : x = y ↔ (x : K) = (y : K) := Subtype.ext_iff +@[norm_cast] +theorem eq_iff {x y : 𝓞 K} : (x : K) = (y : K) ↔ x = y := + NumberField.RingOfIntegers.ext_iff.symm + @[simp] lemma map_mk (x : K) (hx) : algebraMap (𝓞 K) K ⟨x, hx⟩ = x := rfl lemma coe_mk {x : K} (hx) : ((⟨x, hx⟩ : 𝓞 K) : K) = x := rfl diff --git a/Mathlib/NumberTheory/NumberField/Units/Basic.lean b/Mathlib/NumberTheory/NumberField/Units/Basic.lean index a25341e9ae1e8..462ed4681d670 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Basic.lean @@ -71,6 +71,8 @@ theorem coe_injective : Function.Injective ((↑) : (𝓞 K)ˣ → K) := variable {K} +theorem coe_coe (u : (𝓞 K)ˣ) : ((u : 𝓞 K) : K) = (u : K) := rfl + theorem coe_mul (x y : (𝓞 K)ˣ) : ((x * y : (𝓞 K)ˣ) : K) = (x : K) * (y : K) := rfl theorem coe_pow (x : (𝓞 K)ˣ) (n : ℕ) : ((x ^ n : (𝓞 K)ˣ) : K) = (x : K) ^ n := by From 77f39e122a43ab41bf67a3115fd6bcfaa23e63f0 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 7 May 2024 12:53:12 +0000 Subject: [PATCH 12/52] feat: FloorRing.exists_prime_mul_pow_div_factorial_lt_one (#12191) A few technical lemmas in support of Lindemann-Weierstrass. Co-authored-by: negiizhao --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Floor/Prime.lean | 42 ++++++++++++++++++++++++++ Mathlib/Data/Nat/Prime.lean | 31 +++++++++++++++++++ 3 files changed, 74 insertions(+) create mode 100644 Mathlib/Algebra/Order/Floor/Prime.lean diff --git a/Mathlib.lean b/Mathlib.lean index e20d01bb67ce2..a7d11b94c3637 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -466,6 +466,7 @@ import Mathlib.Algebra.Order.Field.Power import Mathlib.Algebra.Order.Field.Subfield import Mathlib.Algebra.Order.Floor import Mathlib.Algebra.Order.Floor.Div +import Mathlib.Algebra.Order.Floor.Prime import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Group.Action import Mathlib.Algebra.Order.Group.Bounds diff --git a/Mathlib/Algebra/Order/Floor/Prime.lean b/Mathlib/Algebra/Order/Floor/Prime.lean new file mode 100644 index 0000000000000..7b12c8042f3ea --- /dev/null +++ b/Mathlib/Algebra/Order/Floor/Prime.lean @@ -0,0 +1,42 @@ +/- +Copyright (c) 2022 Yuyang Zhao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuyang Zhao +-/ + +import Mathlib.Algebra.Order.Floor +import Mathlib.Data.Nat.Prime + +/-! +# Existence of a sufficiently large prime for which `a * c ^ p / (p - 1)! < 1` + +This is a technical result used in the proof of the Lindemann-Weierstrass theorem. +-/ + +namespace FloorRing + +open scoped Nat + +variable {K : Type*} + +theorem exists_prime_mul_pow_lt_factorial [LinearOrderedRing K] [FloorRing K] (n : ℕ) (a c : K) : + ∃ p > n, p.Prime ∧ a * c ^ p < (p - 1)! := by + obtain ⟨p, pn, pp, h⟩ := n.exists_prime_mul_pow_lt_factorial ⌈|a|⌉.natAbs ⌈|c|⌉.natAbs + use p, pn, pp + calc a * c ^ p + _ ≤ |a * c ^ p| := le_abs_self _ + _ ≤ ⌈|a|⌉ * (⌈|c|⌉ : K) ^ p := ?_ + _ = ↑(Int.natAbs ⌈|a|⌉ * Int.natAbs ⌈|c|⌉ ^ p) := ?_ + _ < ↑(p - 1)! := Nat.cast_lt.mpr h + · rw [abs_mul, abs_pow] + gcongr <;> try first | positivity | apply Int.le_ceil + · simp_rw [Nat.cast_mul, Nat.cast_pow, Int.cast_natAbs, + abs_eq_self.mpr (Int.ceil_nonneg (abs_nonneg (_ : K)))] + +theorem exists_prime_mul_pow_div_factorial_lt_one [LinearOrderedField K] [FloorRing K] + (n : ℕ) (a c : K) : + ∃ p > n, p.Prime ∧ a * c ^ p / (p - 1)! < 1 := by + simp_rw [div_lt_one (α := K) (Nat.cast_pos.mpr (Nat.factorial_pos _))] + exact exists_prime_mul_pow_lt_factorial .. + +end FloorRing diff --git a/Mathlib/Data/Nat/Prime.lean b/Mathlib/Data/Nat/Prime.lean index 94687c8a48f56..fc618c2d50d3a 100644 --- a/Mathlib/Data/Nat/Prime.lean +++ b/Mathlib/Data/Nat/Prime.lean @@ -769,6 +769,37 @@ lemma Prime.pow_inj {p q m n : ℕ} (hp : p.Prime) (hq : q.Prime) (Prime.dvd_of_dvd_pow hq <| h.symm ▸ dvd_pow_self q (succ_ne_zero n)) exact ⟨H, succ_inj'.mp <| Nat.pow_right_injective hq.two_le (H ▸ h)⟩ +theorem exists_pow_lt_factorial (c : ℕ) : ∃ n0 > 1, ∀ n ≥ n0, c ^ n < (n - 1)! := by + refine ⟨2 * (c ^ 2 + 1), ?_, ?_⟩ + · omega + intro n hn + 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 ?_ + convert_to (c ^ 2) ^ (c ^ 2 + d + 1) < (c ^ 2 + (c ^ 2 + d + 1))! + · rw [← pow_mul, ← pow_add] + congr 1 + omega + · congr + omega + refine lt_of_lt_of_le ?_ Nat.factorial_mul_pow_le_factorial + rw [← one_mul (_ ^ _ : ℕ)] + exact Nat.mul_lt_mul_of_le_of_lt (Nat.one_le_of_lt (Nat.factorial_pos _)) + (Nat.pow_lt_pow_left (Nat.lt_succ_self _) (Nat.succ_ne_zero _)) (Nat.factorial_pos _) + +theorem exists_mul_pow_lt_factorial (a : ℕ) (c : ℕ) : ∃ n0, ∀ n ≥ n0, a * c ^ n < (n - 1)! := by + obtain ⟨n0, hn, h⟩ := Nat.exists_pow_lt_factorial (a * c) + refine ⟨n0, fun n hn => lt_of_le_of_lt ?_ (h n hn)⟩ + rw [mul_pow] + refine Nat.mul_le_mul_right _ (Nat.le_self_pow ?_ _) + omega + +theorem exists_prime_mul_pow_lt_factorial (n a c : ℕ) : ∃ p > n, p.Prime ∧ a * c ^ p < (p - 1)! := + have ⟨n0, h⟩ := Nat.exists_mul_pow_lt_factorial a c + have ⟨p, hp, prime_p⟩ := (max (n + 1) n0).exists_infinite_primes + ⟨p, (le_max_left _ _).trans hp, prime_p, h _ <| le_of_max_le_right hp⟩ + /-- The type of prime numbers -/ def Primes := { p : ℕ // p.Prime } From 661ebd0d094b230f0de56038fb74ee8cb9bc5d36 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 7 May 2024 12:53:14 +0000 Subject: [PATCH 13/52] feat: add gcd_neg, neg_gcd (#12593) --- Mathlib/Algebra/Associated.lean | 18 ++++++++++++ Mathlib/Algebra/GCDMonoid/Basic.lean | 43 +++++++++++++++++++++------- 2 files changed, 51 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Associated.lean b/Mathlib/Algebra/Associated.lean index 085e7a3e21d10..1f54d31fc730a 100644 --- a/Mathlib/Algebra/Associated.lean +++ b/Mathlib/Algebra/Associated.lean @@ -403,6 +403,9 @@ protected theorem refl [Monoid α] (x : α) : x ~ᵤ x := ⟨1, by simp⟩ #align associated.refl Associated.refl +protected theorem rfl [Monoid α] {x : α} : x ~ᵤ x := + .refl x + instance [Monoid α] : IsRefl α Associated := ⟨Associated.refl⟩ @@ -561,6 +564,9 @@ protected theorem Associated.dvd [Monoid α] {a b : α} : a ~ᵤ b → a ∣ b : ⟨u, hu.symm⟩ #align associated.dvd Associated.dvd +protected theorem Associated.dvd' [Monoid α] {a b : α} (h : a ~ᵤ b) : b ∣ a := + h.symm.dvd + protected theorem Associated.dvd_dvd [Monoid α] {a b : α} (h : a ~ᵤ b) : a ∣ b ∧ b ∣ a := ⟨h.dvd, h.symm.dvd⟩ #align associated.dvd_dvd Associated.dvd_dvd @@ -608,6 +614,18 @@ theorem Associated.ne_zero_iff [MonoidWithZero α] {a b : α} (h : a ~ᵤ b) : a not_congr h.eq_zero_iff #align associated.ne_zero_iff Associated.ne_zero_iff +theorem Associated.neg_left [Monoid α] [HasDistribNeg α] {a b : α} (h : Associated a b) : + Associated (-a) b := + let ⟨u, hu⟩ := h; ⟨-u, by simp [hu]⟩ + +theorem Associated.neg_right [Monoid α] [HasDistribNeg α] {a b : α} (h : Associated a b) : + Associated a (-b) := + h.symm.neg_left.symm + +theorem Associated.neg_neg [Monoid α] [HasDistribNeg α] {a b : α} (h : Associated a b) : + Associated (-a) (-b) := + h.neg_left.neg_right + protected theorem Associated.prime [CommMonoidWithZero α] {p q : α} (h : p ~ᵤ q) (hp : Prime p) : Prime q := ⟨h.ne_zero_iff.1 hp.ne_zero, diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index b4373f0a11464..b587b2f3d8c26 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -191,6 +191,11 @@ theorem dvd_antisymm_of_normalize_eq {a b : α} (ha : normalize a = a) (hb : nor ha ▸ hb ▸ normalize_eq_normalize hab hba #align dvd_antisymm_of_normalize_eq dvd_antisymm_of_normalize_eq +theorem Associated.eq_of_normalized + {a b : α} (h : Associated a b) (ha : normalize a = a) (hb : normalize b = b) : + a = b := + dvd_antisymm_of_normalize_eq ha hb h.dvd h.dvd' + --can be proven by simp theorem dvd_normalize_iff {a b : α} : a ∣ normalize b ↔ a ∣ b := Units.dvd_mul_right @@ -435,6 +440,11 @@ theorem gcd_dvd_gcd [GCDMonoid α] {a b c d : α} (hab : a ∣ b) (hcd : c ∣ d dvd_gcd ((gcd_dvd_left _ _).trans hab) ((gcd_dvd_right _ _).trans hcd) #align gcd_dvd_gcd gcd_dvd_gcd +protected theorem Associated.gcd [GCDMonoid α] + {a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) : + Associated (gcd a₁ b₁) (gcd a₂ b₂) := + associated_of_dvd_dvd (gcd_dvd_gcd ha.dvd hb.dvd) (gcd_dvd_gcd ha.dvd' hb.dvd') + @[simp] theorem gcd_same [NormalizedGCDMonoid α] (a : α) : gcd a a = normalize a := gcd_eq_normalize (gcd_dvd_left _ _) (dvd_gcd (dvd_refl a) (dvd_refl a)) @@ -714,6 +724,24 @@ theorem Irreducible.gcd_eq_one_iff [NormalizedGCDMonoid α] {x y : α} (hx : Irr gcd x y = 1 ↔ ¬(x ∣ y) := by rw [← hx.isUnit_gcd_iff, ← normalize_eq_one, NormalizedGCDMonoid.normalize_gcd] +section Neg + +variable [HasDistribNeg α] + +lemma gcd_neg' [GCDMonoid α] {a b : α} : Associated (gcd a (-b)) (gcd a b) := + Associated.gcd .rfl (.neg_left .rfl) + +lemma gcd_neg [NormalizedGCDMonoid α] {a b : α} : gcd a (-b) = gcd a b := + gcd_neg'.eq_of_normalized (normalize_gcd _ _) (normalize_gcd _ _) + +lemma neg_gcd' [GCDMonoid α] {a b : α} : Associated (gcd (-a) b) (gcd a b) := + Associated.gcd (.neg_left .rfl) .rfl + +lemma neg_gcd [NormalizedGCDMonoid α] {a b : α} : gcd (-a) b = gcd a b := + neg_gcd'.eq_of_normalized (normalize_gcd _ _) (normalize_gcd _ _) + +end Neg + end GCD section LCM @@ -799,6 +827,11 @@ theorem lcm_dvd_lcm [GCDMonoid α] {a b c d : α} (hab : a ∣ b) (hcd : c ∣ d lcm_dvd (hab.trans (dvd_lcm_left _ _)) (hcd.trans (dvd_lcm_right _ _)) #align lcm_dvd_lcm lcm_dvd_lcm +protected theorem Associated.lcm [GCDMonoid α] + {a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) : + Associated (lcm a₁ b₁) (lcm a₂ b₂) := + associated_of_dvd_dvd (lcm_dvd_lcm ha.dvd hb.dvd) (lcm_dvd_lcm ha.dvd' hb.dvd') + @[simp] theorem lcm_units_coe_left [NormalizedGCDMonoid α] (u : αˣ) (a : α) : lcm (↑u) a = normalize a := lcm_eq_normalize (lcm_dvd Units.coe_dvd dvd_rfl) (dvd_lcm_right _ _) @@ -1424,16 +1457,6 @@ theorem normalize_eq_one {a : G₀} (h0 : a ≠ 0) : normalize a = 1 := by simp end CommGroupWithZero -theorem Associated.gcd [CancelCommMonoidWithZero α] [GCDMonoid α] - {a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) : - Associated (gcd a₁ b₁) (gcd a₂ b₂) := - associated_of_dvd_dvd (gcd_dvd_gcd ha.dvd hb.dvd) (gcd_dvd_gcd ha.symm.dvd hb.symm.dvd) - -theorem Associated.lcm [CancelCommMonoidWithZero α] [GCDMonoid α] - {a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) : - Associated (lcm a₁ b₁) (lcm a₂ b₂) := - associated_of_dvd_dvd (lcm_dvd_lcm ha.dvd hb.dvd) (lcm_dvd_lcm ha.symm.dvd hb.symm.dvd) - namespace Associates variable [CancelCommMonoidWithZero α] [GCDMonoid α] From 9ef305b3458090295906a0bb870b46b624cbd9b2 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 7 May 2024 13:42:22 +0000 Subject: [PATCH 14/52] fix: generalize std_basis_eq_basis_mul_basis (#12716) The 1 has been silently interpreted as a natural since it was added in https://github.com/leanprover-community/mathlib/pull/3244. --- Mathlib/Data/Matrix/Basis.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 82f78b1673d93..0001cc692a0f0 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -85,7 +85,8 @@ theorem matrix_eq_sum_std_basis [Fintype m] [Fintype n] (x : Matrix m n α) : -- this is not completely trivial because we are indexing by two types, instead of one -- TODO: add `std_basis_vec` theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : - stdBasisMatrix i j 1 = vecMulVec (fun i' => ite (i = i') 1 0) fun j' => ite (j = j') 1 0 := by + stdBasisMatrix i j (1 : α) = + vecMulVec (fun i' => ite (i = i') 1 0) fun j' => ite (j = j') 1 0 := by ext i' j' -- Porting note: was `norm_num [std_basis_matrix, vec_mul_vec]` though there are no numerals -- involved. From 838141a23f640ed3d382e8eab6b3cbfedca44d5a Mon Sep 17 00:00:00 2001 From: samvang <59202064+samvang@users.noreply.github.com> Date: Tue, 7 May 2024 14:07:59 +0000 Subject: [PATCH 15/52] feat(Mathlib/Order/Ideal): ideals of sets are stable under directed unions and chains (#12651) We prove that the directed union of a non-empty directed union of ideals of subsets of a preorder is an ideal, and in particular ideals of subsets of a preorder are stable under unions. In order to prove it, we also add `directed_on_sUnion`, a variant of `directed_on_iUnion`. --- Mathlib/Data/Set/Lattice.lean | 30 +++++++++++++++++++++--------- Mathlib/Order/Ideal.lean | 19 +++++++++++++++++++ 2 files changed, 40 insertions(+), 9 deletions(-) diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index f3fb255050451..5df4f7a65313d 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -534,15 +534,6 @@ theorem diff_iInter (s : Set β) (t : ι → Set β) : (s \ ⋂ i, t i) = ⋃ i, rw [diff_eq, compl_iInter, inter_iUnion]; rfl #align set.diff_Inter Set.diff_iInter -theorem directed_on_iUnion {r} {f : ι → Set α} (hd : Directed (· ⊆ ·) f) - (h : ∀ x, DirectedOn r (f x)) : DirectedOn r (⋃ x, f x) := by - simp only [DirectedOn, exists_prop, mem_iUnion, exists_imp] - exact fun a₁ b₁ fb₁ a₂ b₂ fb₂ => - let ⟨z, zb₁, zb₂⟩ := hd b₁ b₂ - let ⟨x, xf, xa₁, xa₂⟩ := h z a₁ (zb₁ fb₁) a₂ (zb₂ fb₂) - ⟨x, ⟨z, xf⟩, xa₁, xa₂⟩ -#align set.directed_on_Union Set.directed_on_iUnion - theorem iUnion_inter_subset {ι α} {s t : ι → Set α} : ⋃ i, s i ∩ t i ⊆ (⋃ i, s i) ∩ ⋃ i, t i := le_iSup_inf_iSup s t #align set.Union_inter_subset Set.iUnion_inter_subset @@ -2041,6 +2032,27 @@ theorem iUnion_univ_pi {ι : α → Type*} (t : (a : α) → ι a → Set (π a) end Pi +section Directed + +theorem directedOn_iUnion {r} {f : ι → Set α} (hd : Directed (· ⊆ ·) f) + (h : ∀ x, DirectedOn r (f x)) : DirectedOn r (⋃ x, f x) := by + simp only [DirectedOn, exists_prop, mem_iUnion, exists_imp] + exact fun a₁ b₁ fb₁ a₂ b₂ fb₂ => + let ⟨z, zb₁, zb₂⟩ := hd b₁ b₂ + let ⟨x, xf, xa₁, xa₂⟩ := h z a₁ (zb₁ fb₁) a₂ (zb₂ fb₂) + ⟨x, ⟨z, xf⟩, xa₁, xa₂⟩ +#align set.directed_on_Union Set.directedOn_iUnion + +@[deprecated (since := "2024-05-05")] +alias directed_on_iUnion := directedOn_iUnion + +theorem directedOn_sUnion {r} {S : Set (Set α)} (hd : DirectedOn (· ⊆ ·) S) + (h : ∀ x ∈ S, DirectedOn r x) : DirectedOn r (⋃₀ S) := by + rw [sUnion_eq_iUnion] + exact directedOn_iUnion (directedOn_iff_directed.mp hd) (fun i ↦ h i.1 i.2) + +end Directed + end Set namespace Function diff --git a/Mathlib/Order/Ideal.lean b/Mathlib/Order/Ideal.lean index 47528762ffed1..b367d4e9516b5 100644 --- a/Mathlib/Order/Ideal.lean +++ b/Mathlib/Order/Ideal.lean @@ -5,6 +5,7 @@ Authors: David Wärn -/ import Mathlib.Logic.Encodable.Basic import Mathlib.Order.Atoms +import Mathlib.Order.Chain import Mathlib.Order.UpperLower.Basic import Mathlib.Data.Set.Subsingleton @@ -600,4 +601,22 @@ theorem cofinal_meets_idealOfCofinals (i : ι) : ∃ x : P, x ∈ 𝒟 i ∧ x end IdealOfCofinals +section sUnion + +variable [Preorder P] + +/-- A non-empty directed union of ideals of sets in a preorder is an ideal. -/ +lemma isIdeal_sUnion_of_directedOn {C : Set (Set P)} (hidl : ∀ I ∈ C, IsIdeal I) + (hD : DirectedOn (· ⊆ ·) C) (hNe : C.Nonempty) : IsIdeal C.sUnion := by + refine ⟨isLowerSet_sUnion (fun I hI ↦ (hidl I hI).1), Set.nonempty_sUnion.2 ?_, + directedOn_sUnion hD (fun J hJ => (hidl J hJ).3)⟩ + let ⟨I, hI⟩ := hNe + exact ⟨I, ⟨hI, (hidl I hI).2⟩⟩ + +/-- A union of a nonempty chain of ideals of sets is an ideal. -/ +lemma isIdeal_sUnion_of_isChain {C : Set (Set P)} (hidl : ∀ I ∈ C, IsIdeal I) + (hC : IsChain (· ⊆ ·) C) (hNe : C.Nonempty) : IsIdeal C.sUnion := + isIdeal_sUnion_of_directedOn hidl hC.directedOn hNe + +end sUnion end Order From f1308d7a5ade62fcc7aaf100d63e7b79fbebf752 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 7 May 2024 15:21:13 +0000 Subject: [PATCH 16/52] refactor: Turn Freiman homs into predicates (#12546) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The bundled hom approach to Freiman homomorphisms is a fail. It is very hard to use and doesn't bring anything. That is because Freiman homs don't have much structure and we do not study them under this angle. This PR replaces the bundled homs by predicates and moves the file from `Algebra.Group.Freiman` to `Combinatorics.Additive.FreimanHom` so that: * It is clear these are combinatorial objects, not algebraic ones. * The diff isn't completely mangled. Nothing was kept so it's not worth comparing via a naïve git diff. Also fix a few oddities accidentally introduced in #12701. --- Mathlib.lean | 2 +- Mathlib/Algebra/Group/Freiman.lean | 560 ------------------ .../Additive/AP/Three/Behrend.lean | 6 +- .../Combinatorics/Additive/AP/Three/Defs.lean | 103 +++- .../Combinatorics/Additive/FreimanHom.lean | 369 ++++++++++++ 5 files changed, 455 insertions(+), 585 deletions(-) delete mode 100644 Mathlib/Algebra/Group/Freiman.lean create mode 100644 Mathlib/Combinatorics/Additive/FreimanHom.lean diff --git a/Mathlib.lean b/Mathlib.lean index a7d11b94c3637..90e3ae954b91b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -200,7 +200,6 @@ import Mathlib.Algebra.Group.Embedding import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Algebra.Group.Ext -import Mathlib.Algebra.Group.Freiman import Mathlib.Algebra.Group.Hom.Basic import Mathlib.Algebra.Group.Hom.CompTypeclasses import Mathlib.Algebra.Group.Hom.Defs @@ -1602,6 +1601,7 @@ import Mathlib.Combinatorics.Additive.AP.Three.Behrend import Mathlib.Combinatorics.Additive.AP.Three.Defs import Mathlib.Combinatorics.Additive.ETransform import Mathlib.Combinatorics.Additive.Energy +import Mathlib.Combinatorics.Additive.FreimanHom import Mathlib.Combinatorics.Additive.PluenneckeRuzsa import Mathlib.Combinatorics.Additive.RuzsaCovering import Mathlib.Combinatorics.Colex diff --git a/Mathlib/Algebra/Group/Freiman.lean b/Mathlib/Algebra/Group/Freiman.lean deleted file mode 100644 index cbefca49ff85a..0000000000000 --- a/Mathlib/Algebra/Group/Freiman.lean +++ /dev/null @@ -1,560 +0,0 @@ -/- -Copyright (c) 2022 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import Mathlib.Algebra.BigOperators.Multiset.Basic -import Mathlib.Data.FunLike.Basic - -#align_import algebra.hom.freiman from "leanprover-community/mathlib"@"f694c7dead66f5d4c80f446c796a5aad14707f0e" - -/-! -# Freiman homomorphisms - -In this file, we define Freiman homomorphisms. An `n`-Freiman homomorphism on `A` is a function -`f : α → β` such that `f (x₁) * ... * f (xₙ) = f (y₁) * ... * f (yₙ)` for all -`x₁, ..., xₙ, y₁, ..., yₙ ∈ A` such that `x₁ * ... * xₙ = y₁ * ... * yₙ`. In particular, any -`MulHom` is a Freiman homomorphism. - -They are of interest in additive combinatorics. - -## Main declaration - -* `FreimanHom`: Freiman homomorphism. -* `AddFreimanHom`: Additive Freiman homomorphism. - -## Notation - -* `A →*[n] β`: Multiplicative `n`-Freiman homomorphism on `A` -* `A →+[n] β`: Additive `n`-Freiman homomorphism on `A` - -## Implementation notes - -In the context of combinatorics, we are interested in Freiman homomorphisms over sets which are not -necessarily closed under addition/multiplication. This means we must parametrize them with a set in -an `AddMonoid`/`Monoid` instead of the `AddMonoid`/`Monoid` itself. - -## References - -[Yufei Zhao, *18.225: Graph Theory and Additive Combinatorics*](https://yufeizhao.com/gtac/) - -## TODO - -`MonoidHom.toFreimanHom` could be relaxed to `MulHom.toFreimanHom` by proving -`(s.map f).prod = (t.map f).prod` directly by induction instead of going through `f s.prod`. - -Define `n`-Freiman isomorphisms. - -Affine maps induce Freiman homs. Concretely, provide the `AddFreimanHomClass (α →ₐ[𝕜] β) A β n` -instance. --/ - - -open Multiset - -variable {F α β γ δ G : Type*} - -/-- An additive `n`-Freiman homomorphism is a map which preserves sums of `n` elements. -/ -structure AddFreimanHom (A : Set α) (β : Type*) [AddCommMonoid α] [AddCommMonoid β] (n : ℕ) where - /-- The underlying function. -/ - toFun : α → β - /-- An additive `n`-Freiman homomorphism preserves sums of `n` elements. -/ - map_sum_eq_map_sum' {s t : Multiset α} (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) - (hs : Multiset.card s = n) (ht : Multiset.card t = n) (h : s.sum = t.sum) : - (s.map toFun).sum = (t.map toFun).sum -#align add_freiman_hom AddFreimanHom - -/-- An `n`-Freiman homomorphism on a set `A` is a map which preserves products of `n` elements. -/ -@[to_additive AddFreimanHom] -structure FreimanHom (A : Set α) (β : Type*) [CommMonoid α] [CommMonoid β] (n : ℕ) where - /-- The underlying function. -/ - toFun : α → β - /-- An `n`-Freiman homomorphism preserves products of `n` elements. -/ - map_prod_eq_map_prod' {s t : Multiset α} (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) - (hs : Multiset.card s = n) (ht : Multiset.card t = n) (h : s.prod = t.prod) : - (s.map toFun).prod = (t.map toFun).prod -#align freiman_hom FreimanHom - -/- Porting note: local notation given a name, conflict with Algebra.Hom.GroupAction - see https://github.com/leanprover/lean4/issues/2000 -/ -@[inherit_doc] -notation:25 (name := «AddFreimanHomLocal≺») A " →+[" n:25 "] " β:0 => AddFreimanHom A β n - -/- Porting note: local notation given a name, conflict with Algebra.Hom.GroupAction - see https://github.com/leanprover/lean4/issues/2000 -/ -@[inherit_doc] -notation:25 (name := «FreimanHomLocal≺») A " →*[" n:25 "] " β:0 => FreimanHom A β n - -/-- `AddFreimanHomClass F A β n` states that `F` is a type of `n`-ary sums-preserving morphisms. -You should extend this class when you extend `AddFreimanHom`. -/ -class AddFreimanHomClass (F : Type*) (A : outParam <| Set α) (β : outParam <| Type*) - [AddCommMonoid α] [AddCommMonoid β] (n : ℕ) [FunLike F α β] : Prop where - /-- An additive `n`-Freiman homomorphism preserves sums of `n` elements. -/ - map_sum_eq_map_sum' (f : F) {s t : Multiset α} (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) - (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) (hs : Multiset.card s = n) (ht : Multiset.card t = n) - (h : s.sum = t.sum) : - (s.map f).sum = (t.map f).sum -#align add_freiman_hom_class AddFreimanHomClass - -/-- `FreimanHomClass F A β n` states that `F` is a type of `n`-ary products-preserving morphisms. -You should extend this class when you extend `FreimanHom`. -/ -@[to_additive AddFreimanHomClass - "`AddFreimanHomClass F A β n` states that `F` is a type of `n`-ary - sums-preserving morphisms. You should extend this class when you extend `AddFreimanHom`."] -class FreimanHomClass (F : Type*) (A : outParam <| Set α) (β : outParam <| Type*) [CommMonoid α] - [CommMonoid β] (n : ℕ) [FunLike F α β] : Prop where - /-- An `n`-Freiman homomorphism preserves products of `n` elements. -/ - map_prod_eq_map_prod' (f : F) {s t : Multiset α} (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) - (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) (hs : Multiset.card s = n) (ht : Multiset.card t = n) - (h : s.prod = t.prod) : - (s.map f).prod = (t.map f).prod -#align freiman_hom_class FreimanHomClass - -variable [FunLike F α β] - -section CommMonoid - -variable [CommMonoid α] [CommMonoid β] [CommMonoid γ] [CommMonoid δ] [CommGroup G] {A : Set α} - {B : Set β} {C : Set γ} {n : ℕ} {a b c d : α} - -/- porting note: inserted following def & instance for consistent coercion behaviour, -see also Algebra.Hom.Group for similar -/ -/-- Turn an element of a type `F` satisfying `FreimanHomClass F A β n` into an actual -`FreimanHom`. This is declared as the default coercion from `F` to `FreimanHom A β n`. -/ -@[to_additive (attr := coe) - " Turn an element of a type `F` satisfying `AddFreimanHomClass F A β n` into an actual - `AddFreimanHom`. This is declared as the default coercion from `F` to `AddFreimanHom A β n`."] -def _root_.FreimanHomClass.toFreimanHom [FreimanHomClass F A β n] (f : F) : A →*[n] β where - toFun := DFunLike.coe f - map_prod_eq_map_prod' := FreimanHomClass.map_prod_eq_map_prod' f - -/-- Any type satisfying `SMulHomClass` can be cast into `MulActionHom` via - `SMulHomClass.toMulActionHom`. -/ -instance [FreimanHomClass F A β n] : CoeTC F (A →*[n] β) := - ⟨FreimanHomClass.toFreimanHom⟩ - - -@[to_additive] -theorem map_prod_eq_map_prod [FreimanHomClass F A β n] (f : F) {s t : Multiset α} - (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) - (hs : Multiset.card s = n) (ht : Multiset.card t = n) - (h : s.prod = t.prod) : (s.map f).prod = (t.map f).prod := - FreimanHomClass.map_prod_eq_map_prod' f hsA htA hs ht h -#align map_prod_eq_map_prod map_prod_eq_map_prod -#align map_sum_eq_map_sum map_sum_eq_map_sum - -@[to_additive] -theorem map_mul_map_eq_map_mul_map [FreimanHomClass F A β 2] (f : F) (ha : a ∈ A) (hb : b ∈ A) - (hc : c ∈ A) (hd : d ∈ A) (h : a * b = c * d) : f a * f b = f c * f d := by - simp_rw [← prod_pair] at h ⊢ - refine' map_prod_eq_map_prod f _ _ (card_pair _ _) (card_pair _ _) h <;> simp [ha, hb, hc, hd] -#align map_mul_map_eq_map_mul_map map_mul_map_eq_map_mul_map -#align map_add_map_eq_map_add_map map_add_map_eq_map_add_map - -namespace FreimanHom - -@[to_additive] -instance instFunLike : FunLike (A →*[n] β) α β where - coe := toFun - coe_injective' f g h := by cases f; cases g; congr -#align freiman_hom.fun_like FreimanHom.instFunLike -#align add_freiman_hom.fun_like AddFreimanHom.instFunLike - -@[to_additive addFreimanHomClass] -instance freimanHomClass : FreimanHomClass (A →*[n] β) A β n where - map_prod_eq_map_prod' := map_prod_eq_map_prod' -#align freiman_hom.freiman_hom_class FreimanHom.freimanHomClass -#align add_freiman_hom.freiman_hom_class AddFreimanHom.addFreimanHomClass - --- Porting note: not helpful in lean4 --- /-- Helper instance for when there's too many metavariables to apply `DFunLike.hasCoeToFun` --- directly. -/ --- @[to_additive --- "Helper instance for when there's too many metavariables to apply --- `fun_like.has_coe_to_fun` directly."] --- instance : CoeFun (A →*[n] β) fun _ => α → β := --- ⟨toFun⟩ - -initialize_simps_projections FreimanHom (toFun → apply) -initialize_simps_projections AddFreimanHom (toFun → apply) - -@[to_additive (attr := simp)] -theorem toFun_eq_coe (f : A →*[n] β) : f.toFun = f := - rfl -#align freiman_hom.to_fun_eq_coe FreimanHom.toFun_eq_coe -#align add_freiman_hom.to_fun_eq_coe AddFreimanHom.toFun_eq_coe - -@[to_additive (attr := ext)] -theorem ext ⦃f g : A →*[n] β⦄ (h : ∀ x, f x = g x) : f = g := - DFunLike.ext f g h -#align freiman_hom.ext FreimanHom.ext -#align add_freiman_hom.ext AddFreimanHom.ext - -@[to_additive (attr := simp)] -theorem coe_mk (f : α → β) - (h : - ∀ s t : Multiset α, - (∀ ⦃x⦄, x ∈ s → x ∈ A) → - (∀ ⦃x⦄, x ∈ t → x ∈ A) → - Multiset.card s = n → Multiset.card t = n → - s.prod = t.prod → (s.map f).prod = (t.map f).prod) : - ⇑(mk f (h _ _)) = f := - rfl -#align freiman_hom.coe_mk FreimanHom.coe_mk -#align add_freiman_hom.coe_mk AddFreimanHom.coe_mk - -@[to_additive (attr := simp)] -theorem mk_coe (f : A →*[n] β) (h) : mk f h = f := - ext fun _ => rfl -#align freiman_hom.mk_coe FreimanHom.mk_coe -#align add_freiman_hom.mk_coe AddFreimanHom.mk_coe - -/-- The identity map from a commutative monoid to itself. -/ -@[to_additive (attr := simps) "The identity map from an additive commutative monoid to itself."] -protected def id (A : Set α) (n : ℕ) : A →*[n] α where - toFun x := x - map_prod_eq_map_prod' _ _ _ _ h := by rw [map_id', map_id', h] -#align freiman_hom.id FreimanHom.id -#align add_freiman_hom.id AddFreimanHom.id -#align freiman_hom.id_apply FreimanHom.id_apply - -/-- Composition of Freiman homomorphisms as a Freiman homomorphism. -/ -@[to_additive "Composition of additive Freiman homomorphisms as an additive Freiman homomorphism."] -protected def comp (f : B →*[n] γ) (g : A →*[n] β) (hAB : A.MapsTo g B) : A →*[n] γ where - toFun := f ∘ g - map_prod_eq_map_prod' hsA htA hs ht h := by - rw [← map_map, ← map_map] - apply map_prod_eq_map_prod f _ _ ((card_map _ _).trans hs) - · rwa [card_map] - · apply (map_prod_eq_map_prod g hsA htA hs ht h) - · simpa using fun a h => hAB (hsA h) - · simpa using fun a h => hAB (htA h) -#align freiman_hom.comp FreimanHom.comp -#align add_freiman_hom.comp AddFreimanHom.comp - -@[to_additive (attr := simp)] -theorem coe_comp (f : B →*[n] γ) (g : A →*[n] β) {hfg} : ⇑(f.comp g hfg) = f ∘ g := - rfl -#align freiman_hom.coe_comp FreimanHom.coe_comp -#align add_freiman_hom.coe_comp AddFreimanHom.coe_comp - -@[to_additive] -theorem comp_apply (f : B →*[n] γ) (g : A →*[n] β) {hfg} (x : α) : f.comp g hfg x = f (g x) := - rfl -#align freiman_hom.comp_apply FreimanHom.comp_apply -#align add_freiman_hom.comp_apply AddFreimanHom.comp_apply - -@[to_additive] -theorem comp_assoc (f : A →*[n] β) (g : B →*[n] γ) (h : C →*[n] δ) {hf hhg hgf} - {hh : A.MapsTo (g.comp f hgf) C} : (h.comp g hhg).comp f hf = h.comp (g.comp f hgf) hh := - rfl -#align freiman_hom.comp_assoc FreimanHom.comp_assoc -#align add_freiman_hom.comp_assoc AddFreimanHom.comp_assoc - -@[to_additive (attr := simp)] -theorem cancel_right {g₁ g₂ : B →*[n] γ} {f : A →*[n] β} (hf : Function.Surjective f) {hg₁ hg₂} : - g₁.comp f hg₁ = g₂.comp f hg₂ ↔ g₁ = g₂ := - ⟨fun h => ext <| hf.forall.2 <| DFunLike.ext_iff.1 h, fun h => h ▸ rfl⟩ -#align freiman_hom.cancel_right FreimanHom.cancel_right -#align add_freiman_hom.cancel_right AddFreimanHom.cancel_right - -@[to_additive] -theorem cancel_right_on {g₁ g₂ : B →*[n] γ} {f : A →*[n] β} (hf : A.SurjOn f B) {hf'} : - A.EqOn (g₁.comp f hf') (g₂.comp f hf') ↔ B.EqOn g₁ g₂ := - by simp [hf.cancel_right hf'] -#align freiman_hom.cancel_right_on FreimanHom.cancel_right_on -#align add_freiman_hom.cancel_right_on AddFreimanHom.cancel_right_on - -@[to_additive] -theorem cancel_left_on {g : B →*[n] γ} {f₁ f₂ : A →*[n] β} (hg : B.InjOn g) {hf₁ hf₂} : - A.EqOn (g.comp f₁ hf₁) (g.comp f₂ hf₂) ↔ A.EqOn f₁ f₂ := - by simp [hg.cancel_left hf₁ hf₂] -#align freiman_hom.cancel_left_on FreimanHom.cancel_left_on -#align add_freiman_hom.cancel_left_on AddFreimanHom.cancel_left_on - -@[to_additive (attr := simp)] -theorem comp_id (f : A →*[n] β) {hf} : f.comp (FreimanHom.id A n) hf = f := - ext fun _ => rfl -#align freiman_hom.comp_id FreimanHom.comp_id -#align add_freiman_hom.comp_id AddFreimanHom.comp_id - -@[to_additive (attr := simp)] -theorem id_comp (f : A →*[n] β) {hf} : (FreimanHom.id B n).comp f hf = f := - ext fun _ => rfl -#align freiman_hom.id_comp FreimanHom.id_comp -#align add_freiman_hom.id_comp AddFreimanHom.id_comp - -/-- `FreimanHom.const A n b` is the Freiman homomorphism sending everything to `b`. -/ -@[to_additive "`AddFreimanHom.const An b` is the Freiman homomorphism sending everything to `b`."] -def const (A : Set α) (n : ℕ) (b : β) : A →*[n] β where - toFun _ := b - map_prod_eq_map_prod' _ _ hs ht _ := by - simp only [map_const', hs, prod_replicate, ht] -#align freiman_hom.const FreimanHom.const -#align add_freiman_hom.const AddFreimanHom.const - -@[to_additive (attr := simp)] -theorem const_apply (n : ℕ) (b : β) (x : α) : const A n b x = b := - rfl -#align freiman_hom.const_apply FreimanHom.const_apply -#align add_freiman_hom.const_apply AddFreimanHom.const_apply - -@[to_additive (attr := simp)] -theorem const_comp (n : ℕ) (c : γ) (f : A →*[n] β) {hf} : (const B n c).comp f hf = const A n c := - rfl -#align freiman_hom.const_comp FreimanHom.const_comp -#align add_freiman_hom.const_comp AddFreimanHom.const_comp - -/-- `1` is the Freiman homomorphism sending everything to `1`. -/ -@[to_additive "`0` is the Freiman homomorphism sending everything to `0`."] -instance : One (A →*[n] β) := - ⟨const A n 1⟩ - -@[to_additive (attr := simp)] -theorem one_apply (x : α) : (1 : A →*[n] β) x = 1 := - rfl -#align freiman_hom.one_apply FreimanHom.one_apply -#align add_freiman_hom.zero_apply AddFreimanHom.zero_apply - -@[to_additive (attr := simp)] -theorem one_comp (f : A →*[n] β) {hf} : (1 : B →*[n] γ).comp f hf = 1 := - rfl -#align freiman_hom.one_comp FreimanHom.one_comp -#align add_freiman_hom.zero_comp AddFreimanHom.zero_comp - -@[to_additive] -instance : Inhabited (A →*[n] β) := - ⟨1⟩ - -/-- `f * g` is the Freiman homomorphism sends `x` to `f x * g x`. -/ -@[to_additive "`f + g` is the Freiman homomorphism sending `x` to `f x + g x`."] -instance : Mul (A →*[n] β) := - ⟨fun f g => - { toFun := fun x => f x * g x - map_prod_eq_map_prod' := fun hsA htA hs ht h => by - rw [prod_map_mul, prod_map_mul] - rw [map_prod_eq_map_prod f hsA htA hs ht h] - rw [map_prod_eq_map_prod g hsA htA hs ht h]}⟩ - -@[to_additive (attr := simp)] -theorem mul_apply (f g : A →*[n] β) (x : α) : (f * g) x = f x * g x := - rfl -#align freiman_hom.mul_apply FreimanHom.mul_apply -#align add_freiman_hom.add_apply AddFreimanHom.add_apply - -@[to_additive] -theorem mul_comp (g₁ g₂ : B →*[n] γ) (f : A →*[n] β) {hg hg₁ hg₂} : - (g₁ * g₂).comp f hg = g₁.comp f hg₁ * g₂.comp f hg₂ := - rfl -#align freiman_hom.mul_comp FreimanHom.mul_comp -#align add_freiman_hom.add_comp AddFreimanHom.add_comp - -/-- If `f` is a Freiman homomorphism to a commutative group, then `f⁻¹` is the Freiman homomorphism -sending `x` to `(f x)⁻¹`. -/ -@[to_additive - "If `f` is a Freiman homomorphism to an additive commutative group, then `-f` is the - Freiman homomorphism sending `x` to `-f x`."] -instance : Inv (A →*[n] G) := - ⟨fun f => - { toFun := fun x => (f x)⁻¹ - map_prod_eq_map_prod' := fun hsA htA hs ht h => by - rw [prod_map_inv, prod_map_inv, map_prod_eq_map_prod f hsA htA hs ht h] }⟩ - -@[to_additive (attr := simp)] -theorem inv_apply (f : A →*[n] G) (x : α) : f⁻¹ x = (f x)⁻¹ := - rfl -#align freiman_hom.inv_apply FreimanHom.inv_apply -#align add_freiman_hom.neg_apply AddFreimanHom.neg_apply - -@[to_additive (attr := simp)] -theorem inv_comp (f : B →*[n] G) (g : A →*[n] β) {hf hf'} : f⁻¹.comp g hf = (f.comp g hf')⁻¹ := - ext fun _ => rfl -#align freiman_hom.inv_comp FreimanHom.inv_comp -#align add_freiman_hom.neg_comp AddFreimanHom.neg_comp - -/-- If `f` and `g` are Freiman homomorphisms to a commutative group, then `f / g` is the Freiman -homomorphism sending `x` to `f x / g x`. -/ -@[to_additive - "If `f` and `g` are additive Freiman homomorphisms to an additive commutative group, - then `f - g` is the additive Freiman homomorphism sending `x` to `f x - g x`"] -instance : Div (A →*[n] G) := - ⟨fun f g => - { toFun := fun x => f x / g x - map_prod_eq_map_prod' := fun hsA htA hs ht h => by - rw [prod_map_div, prod_map_div, map_prod_eq_map_prod f hsA htA hs ht h, - map_prod_eq_map_prod g hsA htA hs ht h] }⟩ - -@[to_additive (attr := simp)] -theorem div_apply (f g : A →*[n] G) (x : α) : (f / g) x = f x / g x := - rfl -#align freiman_hom.div_apply FreimanHom.div_apply -#align add_freiman_hom.sub_apply AddFreimanHom.sub_apply - -@[to_additive (attr := simp)] -theorem div_comp (f₁ f₂ : B →*[n] G) (g : A →*[n] β) {hf hf₁ hf₂} : - (f₁ / f₂).comp g hf = f₁.comp g hf₁ / f₂.comp g hf₂ := - ext fun _ => rfl -#align freiman_hom.div_comp FreimanHom.div_comp -#align add_freiman_hom.sub_comp AddFreimanHom.sub_comp - -/-! ### Instances -/ - - -/-- `A →*[n] β` is a `CommMonoid`. -/ -@[to_additive "`α →+[n] β` is an `AddCommMonoid`."] -instance commMonoid : CommMonoid (A →*[n] β) where - mul_assoc a b c := by - ext - apply mul_assoc - one_mul a := by - ext - apply one_mul - mul_one a := by - ext - apply mul_one - mul_comm a b := by - ext - apply mul_comm -#align freiman_hom.comm_monoid FreimanHom.commMonoid -#align add_freiman_hom.add_comm_monoid AddFreimanHom.addCommMonoid - -/-- If `β` is a commutative group, then `A →*[n] β` is a commutative group too. -/ -@[to_additive - "If `β` is an additive commutative group, then `A →*[n] β` is an additive commutative - group too."] -instance commGroup {β} [CommGroup β] : CommGroup (A →*[n] β) := - { FreimanHom.commMonoid with - div_eq_mul_inv := by - intros - ext - apply div_eq_mul_inv - mul_left_inv := by - intros - ext - apply mul_left_inv} -#align freiman_hom.comm_group FreimanHom.commGroup -#align add_freiman_hom.add_comm_group AddFreimanHom.addCommGroup - -end FreimanHom - -/-! ### Hom hierarchy -/ - - ---TODO: change to `MonoidHomClass F A β → FreimanHomClass F A β n` once `map_multiset_prod` is --- generalized -/-- A monoid homomorphism is naturally a `FreimanHom` on its entire domain. - -We can't leave the domain `A : Set α` of the `FreimanHom` a free variable, since it wouldn't be -inferrable. -/ -@[to_additive AddMonoidHom.addFreimanHomClass - " An additive monoid homomorphism is naturally an `AddFreimanHom` on its entire - domain. - - We can't leave the domain `A : Set α` of the `AddFreimanHom` a free variable, since it - wouldn't be inferrable."] -instance MonoidHom.freimanHomClass : FreimanHomClass (α →* β) Set.univ β n where - map_prod_eq_map_prod' f s t _ _ _ _ h := by - rw [← f.map_multiset_prod, h, f.map_multiset_prod] -#align monoid_hom.freiman_hom_class MonoidHom.freimanHomClass -#align add_monoid_hom.freiman_hom_class AddMonoidHom.addFreimanHomClass - -/-- A `MonoidHom` is naturally a `FreimanHom`. -/ -@[to_additive AddMonoidHom.toAddFreimanHom - "An `AddMonoidHom` is naturally an `AddFreimanHom`"] -def MonoidHom.toFreimanHom (A : Set α) (n : ℕ) (f : α →* β) : A →*[n] β where - toFun := f - map_prod_eq_map_prod' _ _ := - map_prod_eq_map_prod f (fun _ _ => Set.mem_univ _) fun _ _ => Set.mem_univ _ -#align monoid_hom.to_freiman_hom MonoidHom.toFreimanHom -#align add_monoid_hom.to_add_freiman_hom AddMonoidHom.toAddFreimanHom - -@[to_additive (attr := simp) toAddFreimanHom_coe] -theorem MonoidHom.toFreimanHom_coe (f : α →* β) : (f.toFreimanHom A n : α → β) = f := - rfl -#align monoid_hom.to_freiman_hom_coe MonoidHom.toFreimanHom_coe -#align add_monoid_hom.to_freiman_hom_coe AddMonoidHom.toAddFreimanHom_coe - -@[to_additive AddMonoidHom.toAddFreimanHom_injective] -theorem MonoidHom.toFreimanHom_injective : - Function.Injective (MonoidHom.toFreimanHom A n : (α →* β) → A →*[n] β) := fun f g h => - by rwa [toFreimanHom, toFreimanHom, FreimanHom.mk.injEq, DFunLike.coe_fn_eq] at h -#align monoid_hom.to_freiman_hom_injective MonoidHom.toFreimanHom_injective -#align add_monoid_hom.to_freiman_hom_injective AddMonoidHom.toAddFreimanHom_injective - -end CommMonoid - -section CancelCommMonoid - -variable [CommMonoid α] [CancelCommMonoid β] {A : Set α} {m n : ℕ} - -@[to_additive] -theorem map_prod_eq_map_prod_of_le [FreimanHomClass F A β n] (f : F) {s t : Multiset α} - (hsA : ∀ x ∈ s, x ∈ A) (htA : ∀ x ∈ t, x ∈ A) - (hs : Multiset.card s = m) (ht : Multiset.card t = m) - (hst : s.prod = t.prod) (h : m ≤ n) : (s.map f).prod = (t.map f).prod := by - obtain rfl | hm := m.eq_zero_or_pos - · rw [card_eq_zero] at hs ht - rw [hs, ht] - simp? [← hs, card_pos_iff_exists_mem] at hm says - simp only [← hs, gt_iff_lt, card_pos_iff_exists_mem] at hm - obtain ⟨a, ha⟩ := hm - suffices - ((s + Multiset.replicate (n - m) a).map f).prod = - ((t + Multiset.replicate (n - m) a).map f).prod by - simp_rw [Multiset.map_add, prod_add] at this - exact mul_right_cancel this - replace ha := hsA _ ha - apply map_prod_eq_map_prod f (A := A) (β := β) (n := n) (fun x hx => _) (fun x hx => _) _ _ _ - -- Porting note: below could be golfed when wlog is available - · intro x hx - rw [mem_add] at hx - cases' hx with hx hx - · exact hsA x hx - · rwa [eq_of_mem_replicate hx] - · intro x hx - rw [mem_add] at hx - cases' hx with hx hx - · exact htA x hx - · rwa [eq_of_mem_replicate hx] - · rw [_root_.map_add, card_replicate, hs]; simp [h] - · rw [_root_.map_add, card_replicate, ht]; simp [h] - · rw [prod_add, prod_add, hst] -#align map_prod_eq_map_prod_of_le map_prod_eq_map_prod_of_le -#align map_sum_eq_map_sum_of_le map_sum_eq_map_sum_of_le - -/-- `α →*[n] β` is naturally included in `A →*[m] β` for any `m ≤ n`. -/ -@[to_additive AddFreimanHom.toAddFreimanHom - "`α →+[n] β` is naturally included in `α →+[m] β` - for any `m ≤ n`"] -def FreimanHom.toFreimanHom (h : m ≤ n) (f : A →*[n] β) : A →*[m] β where - toFun := f - map_prod_eq_map_prod' hsA htA hs ht hst := map_prod_eq_map_prod_of_le f hsA htA hs ht hst h -#align freiman_hom.to_freiman_hom FreimanHom.toFreimanHom -#align add_freiman_hom.to_add_freiman_hom AddFreimanHom.toAddFreimanHom - -/-- An `n`-Freiman homomorphism is also a `m`-Freiman homomorphism for any `m ≤ n`. -/ -@[to_additive AddFreimanHom.addFreimanHomClass_of_le - "An additive `n`-Freiman homomorphism is - also an additive `m`-Freiman homomorphism for any `m ≤ n`."] -theorem FreimanHom.FreimanHomClass_of_le [FreimanHomClass F A β n] (h : m ≤ n) : - FreimanHomClass F A β m where - map_prod_eq_map_prod' f _ _ hsA htA hs ht hst := - map_prod_eq_map_prod_of_le f hsA htA hs ht hst h -#align freiman_hom.freiman_hom_class_of_le FreimanHom.FreimanHomClass_of_le -#align add_freiman_hom.add_freiman_hom_class_of_le AddFreimanHom.addFreimanHomClass_of_le - -@[to_additive (attr := simp) AddFreimanHom.toAddFreimanHom_coe] -theorem FreimanHom.toFreimanHom_coe (h : m ≤ n) (f : A →*[n] β) : - (f.toFreimanHom h : α → β) = f := - rfl -#align freiman_hom.to_freiman_hom_coe FreimanHom.toFreimanHom_coe -#align add_freiman_hom.to_add_freiman_hom_coe AddFreimanHom.toAddFreimanHom_coe - -@[to_additive AddFreimanHom.toAddFreimanHom_injective] -theorem FreimanHom.toFreimanHom_injective (h : m ≤ n) : - Function.Injective (FreimanHom.toFreimanHom h : (A →*[n] β) → A →*[m] β) := fun f g hfg => - FreimanHom.ext <| by convert DFunLike.ext_iff.1 hfg using 0 -#align freiman_hom.to_freiman_hom_injective FreimanHom.toFreimanHom_injective -#align add_freiman_hom.to_freiman_hom_injective AddFreimanHom.toAddFreimanHom_injective - -end CancelCommMonoid diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 0062374e02b5d..e21df625a4443 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -195,8 +195,8 @@ nonrec theorem threeAPFree_sphere : ThreeAPFree (sphere n d k : Set (Fin n → { toFun := fun f => ((↑) : ℕ → ℝ) ∘ f map_zero' := funext fun _ => cast_zero map_add' := fun _ _ => funext fun _ => cast_add _ _ } - refine' ThreeAPFree.of_image (f.toAddFreimanHom (sphere n d k : Set (Fin n → ℕ)) 2) _ _ - · exact cast_injective.comp_left.injOn _ + refine ThreeAPFree.of_image (AddMonoidHomClass.isAddFreimanHom f (Set.mapsTo_image _ _)) + (cast_injective.comp_left.injOn _) (Set.subset_univ _) ?_ refine' (threeAPFree_sphere 0 (√↑k)).mono (Set.image_subset_iff.2 fun x => _) rw [Set.mem_preimage, mem_sphere_zero_iff_norm] exact norm_of_mem_sphere @@ -205,7 +205,7 @@ nonrec theorem threeAPFree_sphere : ThreeAPFree (sphere n d k : Set (Fin n → theorem threeAPFree_image_sphere : ThreeAPFree ((sphere n d k).image (map (2 * d - 1)) : Set ℕ) := by rw [coe_image] - refine' ThreeAPFree.image (α := Fin n → ℕ) (β := ℕ) (s := sphere n d k) (map (2 * d - 1)) + refine' ThreeAPFree.image' (α := Fin n → ℕ) (β := ℕ) (s := sphere n d k) (map (2 * d - 1)) (map_injOn.mono _) threeAPFree_sphere · exact x rw [Set.add_subset_iff] diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean index e76b1e38316ce..29b4371d86115 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import Mathlib.Algebra.Group.Freiman +import Mathlib.Combinatorics.Additive.FreimanHom import Mathlib.Data.Nat.Interval import Mathlib.Data.Set.Pointwise.SMul @@ -118,25 +118,55 @@ end Monoid section CommMonoid variable [CommMonoid α] [CommMonoid β] {s A : Set α} {t B : Set β} {f : α → β} {a : α} -variable [CommMonoid α] [CommMonoid β] {s : Set α} {a : α} - -@[to_additive] -theorem ThreeGPFree.of_image [FunLike F α β] [FreimanHomClass F s β 2] (f : F) - (hf : s.InjOn f) (h : ThreeGPFree (f '' s)) : ThreeGPFree s := - fun _ ha _ hb _ hc habc => hf ha hb <| - h (mem_image_of_mem _ ha) (mem_image_of_mem _ hb) (mem_image_of_mem _ hc) <| - map_mul_map_eq_map_mul_map f ha hc hb hb habc +/-- Arithmetic progressions of length three are preserved under `2`-Freiman homomorphisms. -/ +@[to_additive +"Arithmetic progressions of length three are preserved under `2`-Freiman homomorphisms."] +lemma ThreeGPFree.of_image (hf : IsMulFreimanHom 2 s t f) (hf' : s.InjOn f) (hAs : A ⊆ s) + (hA : ThreeGPFree (f '' A)) : ThreeGPFree A := + fun _ ha _ hb _ hc habc ↦ hf' (hAs ha) (hAs hb) <| hA (mem_image_of_mem _ ha) + (mem_image_of_mem _ hb) (mem_image_of_mem _ hc) <| + hf.mul_eq_mul (hAs ha) (hAs hc) (hAs hb) (hAs hb) habc #align mul_salem_spencer.of_image ThreeGPFree.of_image #align add_salem_spencer.of_image ThreeAPFree.of_image --- TODO: Generalize to Freiman homs +/-- Arithmetic progressions of length three are preserved under `2`-Freiman isomorphisms. -/ +@[to_additive +"Arithmetic progressions of length three are preserved under `2`-Freiman isomorphisms."] +lemma threeGPFree_image (hf : IsMulFreimanIso 2 s t f) (hAs : A ⊆ s) : + ThreeGPFree (f '' A) ↔ ThreeGPFree A := by + rw [ThreeGPFree, ThreeGPFree] + have := (hf.bijOn.injOn.mono hAs).bijOn_image (f := f) + simp (config := { contextual := true }) only + [((hf.bijOn.injOn.mono hAs).bijOn_image (f := f)).forall, + hf.mul_eq_mul (hAs _) (hAs _) (hAs _) (hAs _), this.injOn.eq_iff] + +@[to_additive] alias ⟨_, ThreeGPFree.image⟩ := threeGPFree_image +#align mul_salem_spencer.image ThreeGPFree.image +#align add_salem_spencer.image ThreeAPFree.image + +/-- Arithmetic progressions of length three are preserved under `2`-Freiman homomorphisms. -/ +@[to_additive] +lemma IsMulFreimanHom.threeGPFree (hf : IsMulFreimanHom 2 s t f) (hf' : s.InjOn f) + (ht : ThreeGPFree t) : ThreeGPFree s := + fun _ ha _ hb _ hc habc ↦ hf' ha hb <| ht (hf.mapsTo ha) (hf.mapsTo hb) (hf.mapsTo hc) <| + hf.mul_eq_mul ha hc hb hb habc + +/-- Arithmetic progressions of length three are preserved under `2`-Freiman isomorphisms. -/ +@[to_additive] +lemma IsMulFreimanIso.threeGPFree_congr (hf : IsMulFreimanIso 2 s t f) : + ThreeGPFree s ↔ ThreeGPFree t where + mpr := hf.isMulFreimanHom.threeGPFree hf.bijOn.injOn + mp hs a hfa b hfb c hfc habc := by + obtain ⟨a, ha, rfl⟩ := hf.bijOn.surjOn hfa + obtain ⟨b, hb, rfl⟩ := hf.bijOn.surjOn hfb + obtain ⟨c, hc, rfl⟩ := hf.bijOn.surjOn hfc + exact congr_arg f $ hs ha hb hc $ (hf.mul_eq_mul ha hc hb hb).1 habc + @[to_additive] -theorem ThreeGPFree.image [FunLike F α β] [MulHomClass F α β] (f : F) (hf : (s * s).InjOn f) +theorem ThreeGPFree.image' [FunLike F α β] [MulHomClass F α β] (f : F) (hf : (s * s).InjOn f) (h : ThreeGPFree s) : ThreeGPFree (f '' s) := by rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ habc rw [h ha hb hc (hf (mul_mem_mul ha hc) (mul_mem_mul hb hb) <| by rwa [map_mul, map_mul])] -#align mul_salem_spencer.image ThreeGPFree.image -#align add_salem_spencer.image ThreeAPFree.image end CommMonoid @@ -182,13 +212,11 @@ theorem ThreeGPFree.smul_set (hs : ThreeGPFree s) : ThreeGPFree (a • s) := by #noalign mul_salem_spencer.mul_right #noalign add_salem_spencer.add_right -@[to_additive] -theorem threeGPFree_smul_set : ThreeGPFree ((a * ·) '' s) ↔ ThreeGPFree s := - ⟨fun hs b hb c hc d hd h ↦ - mul_left_cancel +@[to_additive] lemma threeGPFree_smul_set : ThreeGPFree (a • s) ↔ ThreeGPFree s where + mp hs b hb c hc d hd h := mul_left_cancel (hs (mem_image_of_mem _ hb) (mem_image_of_mem _ hc) (mem_image_of_mem _ hd) <| by - rw [mul_mul_mul_comm, h, mul_mul_mul_comm]), - ThreeGPFree.smul_set⟩ + rw [mul_mul_mul_comm, smul_eq_mul, smul_eq_mul, mul_mul_mul_comm, h]) + mpr := ThreeGPFree.smul_set #align mul_salem_spencer_mul_left_iff threeGPFree_smul_set #align add_salem_spencer_add_left_iff threeAPFree_vadd_set @@ -217,8 +245,7 @@ section CancelCommMonoidWithZero variable [CancelCommMonoidWithZero α] [NoZeroDivisors α] {s : Set α} {a : α} -theorem ThreeGPFree.smul_set₀ (hs : ThreeGPFree s) (ha : a ≠ 0) : - ThreeGPFree ((a * ·) '' s) := by +lemma ThreeGPFree.smul_set₀ (hs : ThreeGPFree s) (ha : a ≠ 0) : ThreeGPFree (a • s) := by rintro _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ _ ⟨d, hd, rfl⟩ h exact congr_arg (a • ·) $ hs hb hc hd $ by simpa [mul_mul_mul_comm _ _ a, ha] using h #align mul_salem_spencer.mul_left₀ ThreeGPFree.smul_set₀ @@ -359,6 +386,40 @@ theorem mulRothNumber_lt_of_forall_not_threeGPFree end Monoid +section CommMonoid +variable [CommMonoid α] [CommMonoid β] [DecidableEq β] {A : Finset α} {B : Finset β} {f : α → β} + +/-- Arithmetic progressions can be pushed forward along bijective 2-Freiman homs. -/ +@[to_additive "Arithmetic progressions can be pushed forward along bijective 2-Freiman homs."] +lemma IsMulFreimanHom.mulRothNumber_mono (hf : IsMulFreimanHom 2 A B f) (hf' : Set.BijOn f A B) : + mulRothNumber B ≤ mulRothNumber A := by + obtain ⟨s, hsB, hcard, hs⟩ := mulRothNumber_spec B + have hsA : invFunOn f A '' s ⊆ A := + (hf'.surjOn.mapsTo_invFunOn.mono (coe_subset.2 hsB) Subset.rfl).image_subset + have hfsA : Set.SurjOn f A s := hf'.surjOn.mono Subset.rfl (coe_subset.2 hsB) + rw [← hcard, ← s.card_image_of_injOn ((invFunOn_injOn_image f _).mono hfsA)] + refine ThreeGPFree.le_mulRothNumber ?_ (mod_cast hsA) + rw [coe_image] + + simpa using (hf.subset hsA hfsA.bijOn_subset.mapsTo).threeGPFree (hf'.injOn.mono hsA) hs + +/-- Arithmetic progressions are preserved under 2-Freiman isos. -/ +@[to_additive "Arithmetic progressions are preserved under 2-Freiman isos."] +lemma IsMulFreimanIso.mulRothNumber_congr (hf : IsMulFreimanIso 2 A B f) : + mulRothNumber A = mulRothNumber B := by + refine le_antisymm ?_ (hf.isMulFreimanHom.mulRothNumber_mono hf.bijOn) + obtain ⟨s, hsA, hcard, hs⟩ := mulRothNumber_spec A + rw [← coe_subset] at hsA + have hfs : Set.InjOn f s := hf.bijOn.injOn.mono hsA + have := (hf.subset hsA hfs.bijOn_image).threeGPFree_congr.1 hs + rw [← coe_image] at this + rw [← hcard, ← Finset.card_image_of_injOn hfs] + refine this.le_mulRothNumber ?_ + rw [← coe_subset, coe_image] + exact (hf.bijOn.mapsTo.mono hsA Subset.rfl).image_subset + +end CommMonoid + section CancelCommMonoid variable [CancelCommMonoid α] (s : Finset α) (a : α) diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean new file mode 100644 index 0000000000000..3db46ddddaced --- /dev/null +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -0,0 +1,369 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.BigOperators.Multiset.Basic +import Mathlib.Data.FunLike.Basic +import Mathlib.Data.Set.Pointwise.Basic + +#align_import algebra.hom.freiman from "leanprover-community/mathlib"@"f694c7dead66f5d4c80f446c796a5aad14707f0e" + +/-! +# Freiman homomorphisms + +In this file, we define Freiman homomorphisms and isomorphism. + +An `n`-Freiman homomorphism from `A` to `B` is a function `f : α → β` such that `f '' A ⊆ B` and +`f x₁ * ... * f xₙ = f y₁ * ... * f yₙ` for all `x₁, ..., xₙ, y₁, ..., yₙ ∈ A` such that +`x₁ * ... * xₙ = y₁ * ... * yₙ`. In particular, any `MulHom` is a Freiman homomorphism. + +An `n`-Freiman isomorphism from `A` to `B` is a function `f : α → β` bijective between `A` and `B` +such that `f x₁ * ... * f xₙ = f y₁ * ... * f yₙ ↔ x₁ * ... * xₙ = y₁ * ... * yₙ` for all +`x₁, ..., xₙ, y₁, ..., yₙ ∈ A`. In particular, any `MulEquiv` is a Freiman isomorphism. + +They are of interest in additive combinatorics. + +## Main declaration + +* `IsMulFreimanHom`: Predicate for a function to be a multiplicative Freiman homomorphism. +* `IsAddFreimanHom`: Predicate for a function to be an additive Freiman homomorphism. +* `IsMulFreimanIso`: Predicate for a function to be a multiplicative Freiman isomorphism. +* `IsAddFreimanIso`: Predicate for a function to be an additive Freiman isomorphism. + +## Implementation notes + +In the context of combinatorics, we are interested in Freiman homomorphisms over sets which are not +necessarily closed under addition/multiplication. This means we must parametrize them with a set in +an `AddMonoid`/`Monoid` instead of the `AddMonoid`/`Monoid` itself. + +## References + +[Yufei Zhao, *18.225: Graph Theory and Additive Combinatorics*](https://yufeizhao.com/gtac/) + +## TODO + +* `MonoidHomClass.isMulFreimanHom` could be relaxed to `MulHom.toFreimanHom` by proving + `(s.map f).prod = (t.map f).prod` directly by induction instead of going through `f s.prod`. +* Affine maps are Freiman homs. +-/ + +open Multiset Set +open scoped Pointwise + +variable {F α β γ : Type*} + +section CommMonoid +variable [CommMonoid α] [CommMonoid β] [CommMonoid γ] {A A₁ A₂ : Set α} + {B B₁ B₂ : Set β} {C : Set γ} {f f₁ f₂ : α → β} {g : β → γ} {m n : ℕ} + +/-- An additive `n`-Freiman homomorphism from a set `A` to a set `B` is a map which preserves sums +of `n` elements. -/ +structure IsAddFreimanHom [AddCommMonoid α] [AddCommMonoid β] (n : ℕ) (A : Set α) (B : Set β) + (f : α → β) : Prop where + mapsTo : MapsTo f A B + /-- An additive `n`-Freiman homomorphism preserves sums of `n` elements. -/ + map_sum_eq_map_sum ⦃s t : Multiset α⦄ (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) + (hs : Multiset.card s = n) (ht : Multiset.card t = n) (h : s.sum = t.sum) : + (s.map f).sum = (t.map f).sum + +/-- An `n`-Freiman homomorphism from a set `A` to a set `B` is a map which preserves products of `n` +elements. -/ +@[to_additive] +structure IsMulFreimanHom (n : ℕ) (A : Set α) (B : Set β) (f : α → β) : Prop where + mapsTo : MapsTo f A B + /-- An `n`-Freiman homomorphism preserves products of `n` elements. -/ + map_prod_eq_map_prod ⦃s t : Multiset α⦄ (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) + (hs : Multiset.card s = n) (ht : Multiset.card t = n) (h : s.prod = t.prod) : + (s.map f).prod = (t.map f).prod +#align map_prod_eq_map_prod IsMulFreimanHom.map_prod_eq_map_prod +#align map_sum_eq_map_sum IsAddFreimanHom.map_sum_eq_map_sum + +/-- An additive `n`-Freiman homomorphism from a set `A` to a set `B` is a bijective map which +preserves sums of `n` elements. -/ +structure IsAddFreimanIso [AddCommMonoid α] [AddCommMonoid β] (n : ℕ) (A : Set α) (B : Set β) + (f : α → β) : Prop where + bijOn : BijOn f A B + /-- An additive `n`-Freiman homomorphism preserves sums of `n` elements. -/ + map_sum_eq_map_sum ⦃s t : Multiset α⦄ (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) + (hs : Multiset.card s = n) (ht : Multiset.card t = n) : + (s.map f).sum = (t.map f).sum ↔ s.sum = t.sum + +/-- An `n`-Freiman homomorphism from a set `A` to a set `B` is a map which preserves products of `n` +elements. -/ +@[to_additive] +structure IsMulFreimanIso (n : ℕ) (A : Set α) (B : Set β) (f : α → β) : Prop where + bijOn : BijOn f A B + /-- An `n`-Freiman homomorphism preserves products of `n` elements. -/ + map_prod_eq_map_prod ⦃s t : Multiset α⦄ (hsA : ∀ ⦃x⦄, x ∈ s → x ∈ A) (htA : ∀ ⦃x⦄, x ∈ t → x ∈ A) + (hs : Multiset.card s = n) (ht : Multiset.card t = n) : + (s.map f).prod = (t.map f).prod ↔ s.prod = t.prod + +@[to_additive] +lemma IsMulFreimanIso.isMulFreimanHom (hf : IsMulFreimanIso n A B f) : IsMulFreimanHom n A B f where + mapsTo := hf.bijOn.mapsTo + map_prod_eq_map_prod _s _t hsA htA hs ht := (hf.map_prod_eq_map_prod hsA htA hs ht).2 + +@[to_additive] +lemma IsMulFreimanHom.mul_eq_mul (hf : IsMulFreimanHom 2 A B f) {a b c d : α} + (ha : a ∈ A) (hb : b ∈ A) (hc : c ∈ A) (hd : d ∈ A) (h : a * b = c * d) : + f a * f b = f c * f d := by + simp_rw [← prod_pair] at h ⊢ + refine' hf.map_prod_eq_map_prod _ _ (card_pair _ _) (card_pair _ _) h <;> simp [ha, hb, hc, hd] +#align map_mul_map_eq_map_mul_map IsMulFreimanHom.mul_eq_mul +#align map_add_map_eq_map_add_map IsAddFreimanHom.add_eq_add + +@[to_additive] +lemma IsMulFreimanIso.mul_eq_mul (hf : IsMulFreimanIso 2 A B f) {a b c d : α} + (ha : a ∈ A) (hb : b ∈ A) (hc : c ∈ A) (hd : d ∈ A) : + f a * f b = f c * f d ↔ a * b = c * d := by + simp_rw [← prod_pair] + refine' hf.map_prod_eq_map_prod _ _ (card_pair _ _) (card_pair _ _) <;> simp [ha, hb, hc, hd] + +/-- Characterisation of `2`-Freiman homs. -/ +@[to_additive "Characterisation of `2`-Freiman homs."] +lemma isMulFreimanHom_two : + IsMulFreimanHom 2 A B f ↔ MapsTo f A B ∧ ∀ a ∈ A, ∀ b ∈ A, ∀ c ∈ A, ∀ d ∈ A, + a * b = c * d → f a * f b = f c * f d where + mp hf := ⟨hf.mapsTo, fun a ha b hb c hc d hd ↦ hf.mul_eq_mul ha hb hc hd⟩ + mpr hf := ⟨hf.1, by aesop (add simp [Multiset.card_eq_two])⟩ + +@[to_additive] lemma isMulFreimanHom_id (hA : A₁ ⊆ A₂) : IsMulFreimanHom n A₁ A₂ id where + mapsTo := hA + map_prod_eq_map_prod s t _ _ _ _ h := by simpa using h + +@[to_additive] lemma isMulFreimanIso_id : IsMulFreimanIso n A A id where + bijOn := bijOn_id _ + map_prod_eq_map_prod s t _ _ _ _ := by simp + +@[to_additive] lemma IsMulFreimanHom.comp (hg : IsMulFreimanHom n B C g) + (hf : IsMulFreimanHom n A B f) : IsMulFreimanHom n A C (g ∘ f) where + mapsTo := hg.mapsTo.comp hf.mapsTo + map_prod_eq_map_prod s t hsA htA hs ht h := by + rw [← map_map, ← map_map] + refine hg.map_prod_eq_map_prod ?_ ?_ (by rwa [card_map]) (by rwa [card_map]) + (hf.map_prod_eq_map_prod hsA htA hs ht h) + · simpa using fun a h ↦ hf.mapsTo (hsA h) + · simpa using fun a h ↦ hf.mapsTo (htA h) + +@[to_additive] lemma IsMulFreimanIso.comp (hg : IsMulFreimanIso n B C g) + (hf : IsMulFreimanIso n A B f) : IsMulFreimanIso n A C (g ∘ f) where + bijOn := hg.bijOn.comp hf.bijOn + map_prod_eq_map_prod s t hsA htA hs ht := by + rw [← map_map, ← map_map] + rw[ hg.map_prod_eq_map_prod _ _ (by rwa [card_map]) (by rwa [card_map]), + hf.map_prod_eq_map_prod hsA htA hs ht] + · simpa using fun a h ↦ hf.bijOn.mapsTo (hsA h) + · simpa using fun a h ↦ hf.bijOn.mapsTo (htA h) + +@[to_additive] lemma IsMulFreimanHom.subset (hA : A₁ ⊆ A₂) (hf : IsMulFreimanHom n A₂ B₂ f) + (hf' : MapsTo f A₁ B₁) : IsMulFreimanHom n A₁ B₁ f where + mapsTo := hf' + __ := hf.comp (isMulFreimanHom_id hA) + +@[to_additive] lemma IsMulFreimanHom.superset (hB : B₁ ⊆ B₂) (hf : IsMulFreimanHom n A B₁ f) : + IsMulFreimanHom n A B₂ f := (isMulFreimanHom_id hB).comp hf + +@[to_additive] lemma IsMulFreimanIso.subset (hA : A₁ ⊆ A₂) (hf : IsMulFreimanIso n A₂ B₂ f) + (hf' : BijOn f A₁ B₁) : IsMulFreimanIso n A₁ B₁ f where + bijOn := hf' + map_prod_eq_map_prod s t hsA htA hs ht := by + refine hf.map_prod_eq_map_prod (fun a ha ↦ hA (hsA ha)) (fun a ha ↦ hA (htA ha)) hs ht + +@[to_additive] +lemma isMulFreimanHom_const {b : β} (hb : b ∈ B) : IsMulFreimanHom n A B fun _ ↦ b where + mapsTo _ _ := hb + map_prod_eq_map_prod s t _ _ hs ht _ := by simp only [map_const', hs, prod_replicate, ht] + +@[to_additive (attr := simp)] +lemma isMulFreimanIso_empty : IsMulFreimanIso n (∅ : Set α) (∅ : Set β) f where + bijOn := bijOn_empty _ + map_prod_eq_map_prod s t hs ht := by + simp [eq_zero_of_forall_not_mem hs, eq_zero_of_forall_not_mem ht] + +@[to_additive] lemma IsMulFreimanHom.mul (h₁ : IsMulFreimanHom n A B₁ f₁) + (h₂ : IsMulFreimanHom n A B₂ f₂) : IsMulFreimanHom n A (B₁ * B₂) (f₁ * f₂) where + -- TODO: Extract `Set.MapsTo.mul` from this proof + mapsTo a ha := mul_mem_mul (h₁.mapsTo ha) (h₂.mapsTo ha) + map_prod_eq_map_prod s t hsA htA hs ht h := by + rw [Pi.mul_def, prod_map_mul, prod_map_mul, h₁.map_prod_eq_map_prod hsA htA hs ht h, + h₂.map_prod_eq_map_prod hsA htA hs ht h] + +@[to_additive] lemma MonoidHomClass.isMulFreimanHom [FunLike F α β] [MonoidHomClass F α β] (f : F) + (hfAB : MapsTo f A B) : IsMulFreimanHom n A B f where + mapsTo := hfAB + map_prod_eq_map_prod s t _ _ _ _ h := by rw [← map_multiset_prod, h, map_multiset_prod] + +@[to_additive] lemma MulEquivClass.isMulFreimanIso [EquivLike F α β] [MulEquivClass F α β] (f : F) + (hfAB : BijOn f A B) : IsMulFreimanIso n A B f where + bijOn := hfAB + map_prod_eq_map_prod s t _ _ _ _ := by + rw [← map_multiset_prod, ← map_multiset_prod, EquivLike.apply_eq_iff_eq] + +end CommMonoid + +section CancelCommMonoid +variable [CommMonoid α] [CancelCommMonoid β] {A : Set α} {B : Set β} {f : α → β} {m n : ℕ} + +@[to_additive] +lemma IsMulFreimanHom.mono (hmn : m ≤ n) (hf : IsMulFreimanHom n A B f) : + IsMulFreimanHom m A B f where + mapsTo := hf.mapsTo + map_prod_eq_map_prod s t hsA htA hs ht h := by + obtain rfl | hm := m.eq_zero_or_pos + · rw [card_eq_zero] at hs ht + rw [hs, ht] + simp only [← hs, card_pos_iff_exists_mem] at hm + obtain ⟨a, ha⟩ := hm + suffices ((s + replicate (n - m) a).map f).prod = ((t + replicate (n - m) a).map f).prod by + simp_rw [Multiset.map_add, prod_add] at this + exact mul_right_cancel this + replace ha := hsA ha + refine hf.map_prod_eq_map_prod (fun a ha ↦ ?_) (fun a ha ↦ ?_) ?_ ?_ ?_ + · rw [Multiset.mem_add] at ha + obtain ha | ha := ha + · exact hsA ha + · rwa [eq_of_mem_replicate ha] + · rw [Multiset.mem_add] at ha + obtain ha | ha := ha + · exact htA ha + · rwa [eq_of_mem_replicate ha] + · rw [_root_.map_add, card_replicate, hs, Nat.add_sub_cancel' hmn] + · rw [_root_.map_add, card_replicate, ht, Nat.add_sub_cancel' hmn] + · rw [prod_add, prod_add, h] +#align map_prod_eq_map_prod_of_le IsMulFreimanHom.mono +#align map_sum_eq_map_sum_of_le IsAddFreimanHom.mono + +end CancelCommMonoid + +section CancelCommMonoid +variable [CancelCommMonoid α] [CancelCommMonoid β] {A : Set α} {B : Set β} {f : α → β} {m n : ℕ} + +@[to_additive] +lemma IsMulFreimanIso.mono {hmn : m ≤ n} (hf : IsMulFreimanIso n A B f) : + IsMulFreimanIso m A B f where + bijOn := hf.bijOn + map_prod_eq_map_prod s t hsA htA hs ht := by + obtain rfl | hm := m.eq_zero_or_pos + · rw [card_eq_zero] at hs ht + simp [hs, ht] + simp only [← hs, card_pos_iff_exists_mem] at hm + obtain ⟨a, ha⟩ := hm + suffices + ((s + replicate (n - m) a).map f).prod = ((t + replicate (n - m) a).map f).prod ↔ + (s + replicate (n - m) a).prod = (t + replicate (n - m) a).prod by + simpa only [Multiset.map_add, prod_add, mul_right_cancel_iff] using this + replace ha := hsA ha + refine hf.map_prod_eq_map_prod (fun a ha ↦ ?_) (fun a ha ↦ ?_) ?_ ?_ + · rw [Multiset.mem_add] at ha + obtain ha | ha := ha + · exact hsA ha + · rwa [eq_of_mem_replicate ha] + · rw [Multiset.mem_add] at ha + obtain ha | ha := ha + · exact htA ha + · rwa [eq_of_mem_replicate ha] + · rw [_root_.map_add, card_replicate, hs, Nat.add_sub_cancel' hmn] + · rw [_root_.map_add, card_replicate, ht, Nat.add_sub_cancel' hmn] + +end CancelCommMonoid + +section DivisionCommMonoid +variable [CommMonoid α] [DivisionCommMonoid β] {A : Set α} {B : Set β} {f : α → β} {m n : ℕ} + +@[to_additive] +lemma IsMulFreimanHom.inv (hf : IsMulFreimanHom n A B f) : IsMulFreimanHom n A B⁻¹ f⁻¹ where + -- TODO: Extract `Set.MapsTo.inv` from this proof + mapsTo a ha := inv_mem_inv.2 (hf.mapsTo ha) + map_prod_eq_map_prod s t hsA htA hs ht h := by + rw [Pi.inv_def, prod_map_inv, prod_map_inv, hf.map_prod_eq_map_prod hsA htA hs ht h] + +@[to_additive] lemma IsMulFreimanHom.div {β : Type*} [DivisionCommMonoid β] {B₁ B₂ : Set β} + {f₁ f₂ : α → β} (h₁ : IsMulFreimanHom n A B₁ f₁) (h₂ : IsMulFreimanHom n A B₂ f₂) : + IsMulFreimanHom n A (B₁ / B₂) (f₁ / f₂) where + -- TODO: Extract `Set.MapsTo.div` from this proof + mapsTo a ha := div_mem_div (h₁.mapsTo ha) (h₂.mapsTo ha) + map_prod_eq_map_prod s t hsA htA hs ht h := by + rw [Pi.div_def, prod_map_div, prod_map_div, h₁.map_prod_eq_map_prod hsA htA hs ht h, + h₂.map_prod_eq_map_prod hsA htA hs ht h] + +end DivisionCommMonoid + +#noalign add_freiman_hom +#noalign freiman_hom +#noalign add_freiman_hom_class +#noalign freiman_hom_class +#noalign freiman_hom.fun_like +#noalign add_freiman_hom.fun_like +#noalign freiman_hom.freiman_hom_class +#noalign add_freiman_hom.freiman_hom_class +#noalign freiman_hom.to_fun_eq_coe +#noalign add_freiman_hom.to_fun_eq_coe +#noalign freiman_hom.ext +#noalign add_freiman_hom.ext +#noalign freiman_hom.coe_mk +#noalign add_freiman_hom.coe_nat_eq_mk +#noalign freiman_hom.mk_coe +#noalign add_freiman_hom.mk_coe +#noalign add_freiman_hom.id +#noalign freiman_hom.id_apply +#noalign freiman_hom.comp +#noalign add_freiman_hom.comp +#noalign freiman_hom.coe_comp +#noalign add_freiman_hom.coe_comp +#noalign freiman_hom.comp_apply +#noalign add_freiman_hom.comp_apply +#noalign freiman_hom.comp_assoc +#noalign add_freiman_hom.comp_assoc +#noalign freiman_hom.cancel_right +#noalign add_freiman_hom.cancel_right +#noalign freiman_hom.cancel_right_on +#noalign add_freiman_hom.cancel_right_on +#noalign freiman_hom.cancel_left_on +#noalign add_freiman_hom.cancel_left_on +#noalign freiman_hom.comp_id +#noalign add_freiman_hom.comp_id +#noalign freiman_hom.id_comp +#noalign add_freiman_hom.id_comp +#noalign freiman_hom.const +#noalign add_freiman_hom.const +#noalign freiman_hom.const_apply +#noalign add_freiman_hom.const_apply +#noalign freiman_hom.const_comp +#noalign add_freiman_hom.const_comp +#noalign freiman_hom.one_apply +#noalign add_freiman_hom.zero_apply +#noalign freiman_hom.one_comp +#noalign add_freiman_hom.zero_comp +#noalign freiman_hom.mul_apply +#noalign add_freiman_hom.add_apply +#noalign freiman_hom.mul_comp +#noalign add_freiman_hom.add_comp +#noalign freiman_hom.inv_apply +#noalign add_freiman_hom.neg_apply +#noalign freiman_hom.inv_comp +#noalign add_freiman_hom.neg_comp +#noalign freiman_hom.div_apply +#noalign add_freiman_hom.sub_apply +#noalign freiman_hom.div_comp +#noalign add_freiman_hom.sub_comp +#noalign freiman_hom.comm_monoid +#noalign add_freiman_hom.add_comm_monoid +#noalign freiman_hom.comm_group +#noalign add_freiman_hom.add_comm_group +#noalign monoid_hom.freiman_hom_class +#noalign add_monoid_hom.freiman_hom_class +#noalign monoid_hom.to_freiman_hom +#noalign add_monoid_hom.to_add_freiman_hom +#noalign monoid_hom.to_freiman_hom_coe +#noalign add_monoid_hom.to_freiman_hom_coe +#noalign monoid_hom.to_freiman_hom_injective +#noalign add_monoid_hom.to_freiman_hom_injective +#noalign freiman_hom.to_freiman_hom +#noalign add_freiman_hom.to_add_freiman_hom +#noalign freiman_hom.freiman_hom_class_of_le +#noalign add_freiman_hom.add_freiman_hom_class_of_le +#noalign freiman_hom.to_freiman_hom_coe +#noalign add_freiman_hom.to_add_freiman_hom_coe +#noalign freiman_hom.to_freiman_hom_injective +#noalign add_freiman_hom.to_freiman_hom_injective From 5cafc71ab1ecc86aec7c5aa03739bc8f1a93d6f4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 7 May 2024 16:59:52 +0000 Subject: [PATCH 17/52] further renaming Std->Batteries (#12727) Co-authored-by: Scott Morrison --- Mathlib.lean | 2 +- .../AlgebraicTopology/TopologicalSimplex.lean | 5 +--- Mathlib/Analysis/SpecialFunctions/Exp.lean | 5 +--- Mathlib/Data/Array/Lemmas.lean | 2 +- Mathlib/Data/Bool/AllAny.lean | 4 +-- Mathlib/Data/DList/Basic.lean | 2 +- Mathlib/Data/DList/Defs.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 20 +++++++-------- Mathlib/Data/HashMap.lean | 4 +-- Mathlib/Data/Int/Defs.lean | 4 +-- Mathlib/Data/List/Basic.lean | 25 ++++++++++--------- Mathlib/Data/List/Count.lean | 2 +- Mathlib/Data/List/Defs.lean | 6 ++--- Mathlib/Data/List/Indexes.lean | 2 +- Mathlib/Data/Nat/Bits.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 22 ++++++++-------- Mathlib/Data/Rat/Defs.lean | 16 ++++++------ Mathlib/Data/Rat/Field.lean | 2 +- Mathlib/Data/Rat/Init.lean | 2 +- Mathlib/Data/String/Defs.lean | 2 +- Mathlib/Data/Sum/Basic.lean | 2 +- Mathlib/Init/Classes/Order.lean | 2 +- Mathlib/Init/Control/Lawful.lean | 3 ++- Mathlib/Init/Data/Bool/Basic.lean | 3 ++- Mathlib/Init/Data/Int/DivMod.lean | 2 +- Mathlib/Init/Data/Int/Lemmas.lean | 2 +- Mathlib/Init/Data/List/Basic.lean | 2 +- Mathlib/Init/Data/List/Instances.lean | 2 +- Mathlib/Init/Data/List/Lemmas.lean | 4 +-- Mathlib/Init/Data/Option/Basic.lean | 2 +- Mathlib/Init/Data/Option/Init/Lemmas.lean | 2 +- Mathlib/Init/Data/Option/Lemmas.lean | 2 +- Mathlib/Init/Data/Rat/Basic.lean | 2 +- Mathlib/Lean/Name.lean | 2 +- Mathlib/Logic/Basic.lean | 4 +-- Mathlib/Logic/Unique.lean | 2 +- Mathlib/Tactic.lean | 2 +- Mathlib/Tactic/Attr/Core.lean | 2 +- Mathlib/Tactic/CC/Datatypes.lean | 2 +- Mathlib/Tactic/Congr!.lean | 2 +- Mathlib/Tactic/ExtractGoal.lean | 2 +- Mathlib/Tactic/FunProp.lean | 2 +- Mathlib/Tactic/FunProp/Core.lean | 2 +- Mathlib/Tactic/FunProp/Mor.lean | 2 +- .../FunProp/{ToStd.lean => ToBatteries.lean} | 0 Mathlib/Tactic/Linarith/Frontend.lean | 2 +- Mathlib/Tactic/Linarith/Verification.lean | 2 +- Mathlib/Tactic/Lint.lean | 2 +- Mathlib/Tactic/Monotonicity/Basic.lean | 2 +- Mathlib/Tactic/Nontriviality/Core.lean | 2 +- Mathlib/Tactic/Replace.lean | 2 +- Mathlib/Util/Delaborators.lean | 2 +- docs/Conv/Guide.lean | 13 +++++----- docs/overview.yaml | 8 +++--- scripts/nolints.json | 5 ---- scripts/noshake.json | 2 +- 56 files changed, 110 insertions(+), 117 deletions(-) rename Mathlib/Tactic/FunProp/{ToStd.lean => ToBatteries.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 90e3ae954b91b..df4ae97c027c1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3684,7 +3684,7 @@ import Mathlib.Tactic.FunProp.Mor import Mathlib.Tactic.FunProp.RefinedDiscrTree import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems -import Mathlib.Tactic.FunProp.ToStd +import Mathlib.Tactic.FunProp.ToBatteries import Mathlib.Tactic.FunProp.Types import Mathlib.Tactic.GCongr import Mathlib.Tactic.GCongr.Core diff --git a/Mathlib/AlgebraicTopology/TopologicalSimplex.lean b/Mathlib/AlgebraicTopology/TopologicalSimplex.lean index 6e53cfb6f5d5c..9fb3e82a3025e 100644 --- a/Mathlib/AlgebraicTopology/TopologicalSimplex.lean +++ b/Mathlib/AlgebraicTopology/TopologicalSimplex.lean @@ -55,10 +55,7 @@ def toTopMap {x y : SimplexCategory} (f : x ⟶ y) (g : x.toTopObj) : y.toTopObj · apply Set.pairwiseDisjoint_filter⟩ #align simplex_category.to_Top_map SimplexCategory.toTopMap --- Adaptation note: nightly-2024-04-23 --- The simpNF linter now times out on this lemma. --- See https://github.com/leanprover-community/mathlib4/issues/12227 -@[simp, nolint simpNF] +@[simp] theorem coe_toTopMap {x y : SimplexCategory} (f : x ⟶ y) (g : x.toTopObj) (i : y) : toTopMap f g i = ∑ j in Finset.univ.filter (f · = i), g j := rfl diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index b0f7d0b9cb8c7..84a260cb9101e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -450,10 +450,7 @@ lemma HasSum.rexp {ι} {f : ι → ℝ} {a : ℝ} (h : HasSum f a) : HasProd (re namespace Complex --- Adaptation note: nightly-2024-04-01 --- The simpNF linter now times out on this lemma. --- See https://github.com/leanprover-community/mathlib4/issues/12226 -@[simp, nolint simpNF] +@[simp] theorem comap_exp_cobounded : comap exp (cobounded ℂ) = comap re atTop := calc comap exp (cobounded ℂ) = comap re (comap Real.exp atTop) := by diff --git a/Mathlib/Data/Array/Lemmas.lean b/Mathlib/Data/Array/Lemmas.lean index 521e025342b09..923f9c0990c57 100644 --- a/Mathlib/Data/Array/Lemmas.lean +++ b/Mathlib/Data/Array/Lemmas.lean @@ -12,7 +12,7 @@ import Mathlib.Data.List.Basic Porting note: Following the discussion on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929- mathlib4/topic/porting.20.60.2Earray.60.20files.20whose.20PR's.20were.20closed.3F), these -will wait until Std4 has finalized `DArray` and `Array'` types so we can translates +will wait until Batteries has finalized `DArray` and `Array'` types so we can translates apples to apples. `align` for lemmas about lean3/mathlib3 versions of d_array and array with mathport output diff --git a/Mathlib/Data/Bool/AllAny.lean b/Mathlib/Data/Bool/AllAny.lean index e024629f42f0f..2e8b962b101de 100644 --- a/Mathlib/Data/Bool/AllAny.lean +++ b/Mathlib/Data/Bool/AllAny.lean @@ -21,7 +21,7 @@ variable {α : Type*} {p : α → Prop} [DecidablePred p] {l : List α} {a : α} namespace List --- Porting note: in Std +-- Porting note: in Batteries #align list.all_nil List.all_nil #align list.all_cons List.all_consₓ @@ -36,7 +36,7 @@ theorem all_iff_forall_prop : (all l fun a => p a) ↔ ∀ a ∈ l, p a := by simp only [all_iff_forall, Bool.of_decide_iff] #align list.all_iff_forall_prop List.all_iff_forall_prop --- Porting note: in Std +-- Porting note: in Batteries #align list.any_nil List.any_nil #align list.any_cons List.any_consₓ diff --git a/Mathlib/Data/DList/Basic.lean b/Mathlib/Data/DList/Basic.lean index 0c3103443f75a..2df33cb464e8d 100644 --- a/Mathlib/Data/DList/Basic.lean +++ b/Mathlib/Data/DList/Basic.lean @@ -12,7 +12,7 @@ import Mathlib.Tactic.TypeStar /-! # Difference list -This file provides a few results about `DList`, which is defined in `Std`. +This file provides a few results about `DList`, which is defined in `Batteries`. A difference list is a function that, given a list, returns the original content of the difference list prepended to the given list. It is useful to represent elements of a given type diff --git a/Mathlib/Data/DList/Defs.lean b/Mathlib/Data/DList/Defs.lean index be59b95876500..10f72d34e8842 100644 --- a/Mathlib/Data/DList/Defs.lean +++ b/Mathlib/Data/DList/Defs.lean @@ -12,7 +12,7 @@ import Mathlib.Tactic.Cases /-! # Difference list -This file provides a few results about `DList`, which is defined in `Std`. +This file provides a few results about `DList`, which is defined in `Batteries`. A difference list is a function that, given a list, returns the original content of the difference list prepended to the given list. It is useful to represent elements of a given type diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index a4a8f5ed72e2a..2e7412b4274c6 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -304,7 +304,7 @@ instance {n : ℕ} [NeZero n] : One (Fin n) := ⟨ofNat'' 1⟩ #align fin.coe_zero Fin.val_zero /-- -The `Fin.val_zero` in `Std` only applies in `Fin (n+1)`. +The `Fin.val_zero` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ @[simp] @@ -315,7 +315,7 @@ theorem val_zero' (n : ℕ) [NeZero n] : ((0 : Fin n) : ℕ) = 0 := #align fin.mk_zero Fin.mk_zero /-- -The `Fin.zero_le` in `Std` only applies in `Fin (n+1)`. +The `Fin.zero_le` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ @[simp] @@ -327,7 +327,7 @@ protected theorem zero_le' [NeZero n] (a : Fin n) : 0 ≤ a := #align fin.not_lt_zero Fin.not_lt_zero /-- -The `Fin.pos_iff_ne_zero` in `Std` only applies in `Fin (n+1)`. +The `Fin.pos_iff_ne_zero` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ theorem pos_iff_ne_zero' [NeZero n] (a : Fin n) : 0 < a ↔ a ≠ 0 := by @@ -780,7 +780,7 @@ theorem zero_ne_one' [NeZero n] : (0 : Fin (n + 1)) ≠ 1 := Fin.ne_of_lt one_po #align fin.succ_zero_eq_one' Fin.succ_zero_eq_one /-- -The `Fin.succ_one_eq_two` in `Std` only applies in `Fin (n+2)`. +The `Fin.succ_one_eq_two` in `Lean` only applies in `Fin (n+2)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ @[simp] @@ -803,7 +803,7 @@ theorem succ_one_eq_two' [NeZero n] : Fin.succ (1 : Fin (n + 1)) = 2 := by #align fin.lt_add_one_iff Fin.lt_add_one_iff /-- -The `Fin.le_zero_iff` in `Std` only applies in `Fin (n+1)`. +The `Fin.le_zero_iff` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ @[simp] @@ -816,7 +816,7 @@ theorem le_zero_iff' {n : ℕ} [NeZero n] {k : Fin n} : k ≤ 0 ↔ k = 0 := #align fin.coe_cast_lt Fin.coe_castLT #align fin.cast_lt_mk Fin.castLT_mk --- Move to Std? +-- Move to Batteries? @[simp] theorem cast_refl {n : Nat} (h : n = n) : Fin.cast h = id := rfl @@ -1045,7 +1045,7 @@ theorem exists_fin_succ' {P : Fin (n + 1) → Prop} : fun h => h.elim (fun ⟨i, hi⟩ => ⟨i.castSucc, hi⟩) (fun h => ⟨.last _, h⟩)⟩ /-- -The `Fin.castSucc_zero` in `Std` only applies in `Fin (n+1)`. +The `Fin.castSucc_zero` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ @[simp] @@ -1056,14 +1056,14 @@ theorem castSucc_zero' [NeZero n] : castSucc (0 : Fin n) = 0 := /-- `castSucc i` is positive when `i` is positive. -The `Fin.castSucc_pos` in `Std` only applies in `Fin (n+1)`. +The `Fin.castSucc_pos` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ theorem castSucc_pos' [NeZero n] {i : Fin n} (h : 0 < i) : 0 < castSucc i := by simpa [lt_iff_val_lt_val] using h #align fin.cast_succ_pos Fin.castSucc_pos' /-- -The `Fin.castSucc_eq_zero_iff` in `Std` only applies in `Fin (n+1)`. +The `Fin.castSucc_eq_zero_iff` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ @[simp] @@ -1072,7 +1072,7 @@ theorem castSucc_eq_zero_iff' [NeZero n] (a : Fin n) : castSucc a = 0 ↔ a = 0 #align fin.cast_succ_eq_zero_iff Fin.castSucc_eq_zero_iff' /-- -The `Fin.castSucc_ne_zero_iff` in `Std` only applies in `Fin (n+1)`. +The `Fin.castSucc_ne_zero_iff` in `Lean` only applies in `Fin (n+1)`. This one instead uses a `NeZero n` typeclass hypothesis. -/ theorem castSucc_ne_zero_iff' [NeZero n] (a : Fin n) : castSucc a ≠ 0 ↔ a ≠ 0 := diff --git a/Mathlib/Data/HashMap.lean b/Mathlib/Data/HashMap.lean index f8e1e56ac5314..7929e522e2f81 100644 --- a/Mathlib/Data/HashMap.lean +++ b/Mathlib/Data/HashMap.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 -As `HashMap` has been completely reimplemented in `Std`, +As `HashMap` has been completely reimplemented in `Batteries`, nothing from the mathlib3 file `data.hash_map` is reflected here. The porting header is just here to mark that no further work on `data.hash_map` is desired. -/ @@ -16,7 +16,7 @@ import Mathlib.Mathport.Rename /-! # Additional API for `HashMap` and `RBSet`. -These should be replaced by proper implementations in Std. +These should be replaced by proper implementations in Batteries. -/ set_option autoImplicit true diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 037e07998d208..5c59bb8df703e 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -20,7 +20,7 @@ See note [foundational algebra order theory]. ## TODO Split this file into: -* `Data.Int.Init` (or maybe `Data.Int.Std`?) for lemmas that could go to Std +* `Data.Int.Init` (or maybe `Data.Int.Batteries`?) for lemmas that could go to Batteries * `Data.Int.Basic` for the lemmas that require mathlib definitions -/ @@ -47,7 +47,7 @@ variable {m n : ℕ} #align int.of_nat_nat_abs_eq_of_nonneg Int.ofNat_natAbs_eq_of_nonnegₓ #align int.nat_abs_of_neg_succ_of_nat Int.natAbs_negSucc --- TODO: Tag in Std +-- TODO: Tag in Lean attribute [simp] natAbs_pos protected lemma one_pos : 0 < (1 : Int) := Int.zero_lt_one diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 254bfac5d96e5..db21afe70a8eb 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -319,7 +319,7 @@ theorem append_subset_of_subset_of_subset {l₁ l₂ l : List α} (l₁subl : l fun _ h ↦ (mem_append.1 h).elim (@l₁subl _) (@l₂subl _) #align list.append_subset_of_subset_of_subset List.append_subset_of_subset_of_subset --- Porting note: in Std +-- Porting note: in Batteries #align list.append_subset_iff List.append_subset alias ⟨eq_nil_of_subset_nil, _⟩ := subset_nil @@ -350,7 +350,7 @@ theorem append_eq_has_append {L₁ L₂ : List α} : List.append L₁ L₂ = L #align list.append_eq_nil List.append_eq_nil --- Porting note: in Std +-- Porting note: in Batteries #align list.nil_eq_append_iff List.nil_eq_append @[deprecated] alias append_eq_cons_iff := append_eq_cons -- 2024-03-24 @@ -579,9 +579,9 @@ theorem concat_eq_reverse_cons (a : α) (l : List α) : concat l a = reverse (a #align list.length_reverse List.length_reverse -- Porting note: This one was @[simp] in mathlib 3, --- but Std contains a competing simp lemma reverse_map. +-- but Lean contains a competing simp lemma reverse_map. -- For now we remove @[simp] to avoid simplification loops. --- TODO: Change Std lemma to match mathlib 3? +-- TODO: Change Lean lemma to match mathlib 3? theorem map_reverse (f : α → β) (l : List α) : map f (reverse l) = reverse (map f l) := (reverse_map f l).symm #align list.map_reverse List.map_reverse @@ -1083,7 +1083,7 @@ theorem eq_nil_of_sublist_nil {l : List α} (s : l <+ []) : l = [] := eq_nil_of_subset_nil <| s.subset #align list.eq_nil_of_sublist_nil List.eq_nil_of_sublist_nil --- Porting note: this lemma seems to have been renamed on the occasion of its move to Std4 +-- Porting note: this lemma seems to have been renamed on the occasion of its move to Batteries alias sublist_nil_iff_eq_nil := sublist_nil #align list.sublist_nil_iff_eq_nil List.sublist_nil_iff_eq_nil @@ -2668,12 +2668,14 @@ variable {p : α → Bool} {l : List α} {a : α} #align list.find_nil List.find?_nil -- @[simp] --- Later porting note (at time of this lemma moving to Std): removing attribute `nolint simpNF` +-- Later porting note (at time of this lemma moving to Batteries): +-- removing attribute `nolint simpNF` attribute [simp 1100] find?_cons_of_pos #align list.find_cons_of_pos List.find?_cons_of_pos -- @[simp] --- Later porting note (at time of this lemma moving to Std): removing attribute `nolint simpNF` +-- Later porting note (at time of this lemma moving to Batteries): +-- removing attribute `nolint simpNF` attribute [simp 1100] find?_cons_of_neg #align list.find_cons_of_neg List.find?_cons_of_neg @@ -2771,14 +2773,13 @@ end Lookmap #align list.filter_map_nil List.filterMap_nil --- Porting note: List.filterMap is given @[simp] in Std.Data.List.Init.Lemmas --- @[simp] --- Later porting note (at time of this lemma moving to Std): removing attribute `nolint simpNF` +-- Later porting note (at time of this lemma moving to Batteries): +-- removing attribute `nolint simpNF` attribute [simp 1100] filterMap_cons_none #align list.filter_map_cons_none List.filterMap_cons_none --- @[simp] --- Later porting note (at time of this lemma moving to Std): removing attribute `nolint simpNF` +-- Later porting note (at time of this lemma moving to Batteries): +-- removing attribute `nolint simpNF` attribute [simp 1100] filterMap_cons_some #align list.filter_map_cons_some List.filterMap_cons_some diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 6b256ce865e31..b10a901d6c283 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -12,7 +12,7 @@ import Mathlib.Data.List.Basic This file proves basic properties of `List.countP` and `List.count`, which count the number of elements of a list satisfying a predicate and equal to a given element respectively. Their -definitions can be found in `Std.Data.List.Basic`. +definitions can be found in `Batteries.Data.List.Basic`. -/ assert_not_exists Set.range diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 80739459f459f..fb15d21f11073 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -20,7 +20,7 @@ proofs about these definitions, those are contained in other files in `Data.List -/ -- Porting note --- Many of the definitions in `Data.List.Defs` were already defined upstream in `Std4` +-- Many of the definitions in `Data.List.Defs` were already defined upstream in `Batteries` -- These have been annotated with `#align`s -- To make this easier for review, the `#align`s have been placed in order of occurrence -- in `mathlib` @@ -150,7 +150,7 @@ end foldIdxM section mapIdxM -- Porting note: This was defined in `mathlib` with an `Applicative` --- constraint on `m` and have been `#align`ed to the `Std` versions defined +-- constraint on `m` and have been `#align`ed to the `Batteries` versions defined -- with a `Monad` typeclass constraint. -- Since all `Monad`s are `Applicative` this won't cause issues -- downstream & `Monad`ic code is more performant per Mario C @@ -358,7 +358,7 @@ def destutter (R : α → α → Prop) [DecidableRel R] : List α → List α #align list.reduce_option List.reduceOption -- Porting note: replace ilast' by getLastD #align list.ilast' List.ilast' --- Porting note: remove last' from Std +-- Porting note: remove last' from Batteries #align list.last' List.getLast? #align list.rotate List.rotate #align list.rotate' List.rotate' diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 6f3e429fabe02..c88ed84cd6e44 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -267,7 +267,7 @@ theorem findIdxs_eq_map_indexesValues (p : α → Prop) [DecidablePred p] (as : Bool.cond_decide] #align list.find_indexes_eq_map_indexes_values List.findIdxs_eq_map_indexesValues -section FindIdx -- TODO: upstream to Std +section FindIdx -- TODO: upstream to Batteries theorem findIdx_eq_length {p : α → Bool} {xs : List α} : xs.findIdx p = xs.length ↔ ∀ x ∈ xs, ¬p x := by diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 3cf09804e1958..5f366a1f0d118 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -184,7 +184,7 @@ lemma shiftLeft'_false : ∀ n, shiftLeft' false m n = m <<< n rw [Nat.mul_comm, Nat.mul_assoc, ← Nat.pow_succ]; simp simp [shiftLeft_eq, shiftLeft', bit_val, shiftLeft'_false, this] -/-- Std4 takes the unprimed name for `Nat.shiftLeft_eq m n : m <<< n = m * 2 ^ n`. -/ +/-- Lean takes the unprimed name for `Nat.shiftLeft_eq m n : m <<< n = m * 2 ^ n`. -/ @[simp] lemma shiftLeft_eq' (m n : Nat) : shiftLeft m n = m <<< n := rfl @[simp] lemma shiftRight_eq (m n : Nat) : shiftRight m n = m >>> n := rfl diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 5006adff84169..4acfababd4f88 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -31,13 +31,13 @@ See note [foundational algebra order theory]. ## TODO Split this file into: -* `Data.Nat.Init` (or maybe `Data.Nat.Std`?) for lemmas that could go to Std +* `Data.Nat.Init` (or maybe `Data.Nat.Batteries`?) for lemmas that could go to Batteries * `Data.Nat.Basic` for the lemmas that require mathlib definitions -/ library_note "foundational algebra order theory"/-- -Std has a home-baked development of the algebraic and order theoretic theory of `ℕ` and `ℤ` which, -in particular, is not typeclass-mediated. This is useful to set up the algebra and finiteness +Batteries has a home-baked development of the algebraic and order theoretic theory of `ℕ` and `ℤ +which, in particular, is not typeclass-mediated. This is useful to set up the algebra and finiteness libraries in mathlib (naturals and integers show up as indices/offsets in lists, cardinality in finsets, powers in groups, ...). @@ -95,7 +95,7 @@ lemma succ_succ_ne_one (n : ℕ) : n.succ.succ ≠ 1 := by simp lemma one_lt_succ_succ (n : ℕ) : 1 < n.succ.succ := succ_lt_succ <| succ_pos n #align nat.one_lt_succ_succ Nat.one_lt_succ_succ --- Moved to Std +-- Moved to Batteries #align nat.succ_le_succ_iff Nat.succ_le_succ_iff #align nat.succ_lt_succ_iff Nat.succ_lt_succ_iff #align nat.le_pred_of_lt Nat.le_pred_of_lt @@ -174,7 +174,7 @@ lemma one_le_of_lt (h : a < b) : 1 ≤ b := Nat.lt_of_le_of_lt (Nat.zero_le _) h #align nat.min_eq_zero_iff Nat.min_eq_zero_iff #align nat.max_eq_zero_iff Nat.max_eq_zero_iff --- Moved to Std +-- Moved to Batteries #align nat.succ_eq_one_add Nat.succ_eq_one_add #align nat.one_add Nat.one_add #align nat.zero_max Nat.zero_max @@ -1237,7 +1237,7 @@ protected lemma mul_dvd_mul_iff_right (hc : 0 < c) : a * c ∣ b * c ↔ a ∣ b #align nat.dvd_one Nat.dvd_one #align nat.mod_mod_of_dvd Nat.mod_mod_of_dvd --- Moved to Std +-- Moved to Batteries #align nat.mod_mod Nat.mod_mod #align nat.mod_add_mod Nat.mod_add_mod #align nat.add_mod_mod Nat.add_mod_mod @@ -1251,7 +1251,7 @@ lemma add_mod_eq_add_mod_left (c : ℕ) (H : a % d = b % d) : (c + a) % d = (c + rw [Nat.add_comm, add_mod_eq_add_mod_right _ H, Nat.add_comm] #align nat.add_mod_eq_add_mod_left Nat.add_mod_eq_add_mod_left --- Moved to Std +-- Moved to Batteries #align nat.mul_mod Nat.mul_mod lemma mul_dvd_of_dvd_div (hcb : c ∣ b) (h : a ∣ b / c) : c * a ∣ b := @@ -1272,7 +1272,7 @@ protected theorem div_le_div {a b c d : ℕ} (h1 : a ≤ b) (h2 : d ≤ c) (h3 : calc a / c ≤ b / c := Nat.div_le_div_right h1 _ ≤ b / d := Nat.div_le_div_left h2 (Nat.pos_of_ne_zero h3) --- Moved to Std +-- Moved to Batteries #align nat.mul_div_le Nat.mul_div_le lemma lt_mul_div_succ (a : ℕ) (hb : 0 < b) : a < b * (a / b + 1) := by @@ -1280,7 +1280,7 @@ lemma lt_mul_div_succ (a : ℕ) (hb : 0 < b) : a < b * (a / b + 1) := by exact lt_succ_self _ #align nat.lt_mul_div_succ Nat.lt_mul_div_succ --- TODO: Std4 claimed this name but flipped the order of multiplication +-- TODO: Batteries claimed this name but flipped the order of multiplication lemma mul_add_mod' (a b c : ℕ) : (a * b + c) % b = c % b := by rw [Nat.mul_comm, Nat.mul_add_mod] #align nat.mul_add_mod Nat.mul_add_mod' @@ -1473,7 +1473,7 @@ private lemma AM_GM : {a b : ℕ} → (4 * a * b ≤ (a + b) * (a + b)) Nat.mul_one, Nat.add_assoc, Nat.add_left_comm, Nat.add_le_add_iff_left] using Nat.add_le_add_right (@AM_GM a b) 4 --- These two lemmas seem like they belong to `Std.Data.Nat.Basic`. +-- These two lemmas seem like they belong to `Batteries.Data.Nat.Basic`. lemma sqrt.iter_sq_le (n guess : ℕ) : sqrt.iter n guess * sqrt.iter n guess ≤ n := by unfold sqrt.iter @@ -1511,7 +1511,7 @@ lemma sqrt.lt_iter_succ_sq (n guess : ℕ) (hn : n < (guess + 1) * (guess + 1)) · simpa only [dif_neg h] using hn #align nat.sqrt Nat.sqrt --- Porting note: the implementation òf `Nat.sqrt` in `Std` no longer needs `sqrt_aux`. +-- Porting note: the implementation of `Nat.sqrt` in `Batteries` no longer needs `sqrt_aux`. #noalign nat.sqrt_aux_dec #noalign nat.sqrt_aux #noalign nat.sqrt_aux_0 diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index d8856e2ae97ff..4327a6a199946 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -51,7 +51,7 @@ theorem ofInt_eq_cast (n : ℤ) : ofInt n = Int.cast n := rfl #align rat.of_int_eq_cast Rat.ofInt_eq_cast --- TODO: Replace `Rat.ofNat_num`/`Rat.ofNat_den` in Std +-- TODO: Replace `Rat.ofNat_num`/`Rat.ofNat_den` in Batteries -- See note [no_index around OfNat.ofNat] @[simp] lemma num_ofNat (n : ℕ) : num (no_index (OfNat.ofNat n)) = OfNat.ofNat n := rfl @[simp] lemma den_ofNat (n : ℕ) : den (no_index (OfNat.ofNat n)) = 1 := rfl @@ -62,7 +62,7 @@ theorem ofInt_eq_cast (n : ℤ) : ofInt n = Int.cast n := @[simp, norm_cast] lemma den_natCast (n : ℕ) : den n = 1 := rfl #align rat.coe_nat_denom Rat.den_natCast --- TODO: Replace `intCast_num`/`intCast_den` the names in Std +-- TODO: Replace `intCast_num`/`intCast_den` the names in Batteries @[simp, norm_cast] lemma num_intCast (n : ℤ) : (n : ℚ).num = n := rfl #align rat.coe_int_num Rat.num_intCast @@ -114,14 +114,14 @@ theorem divInt_ne_zero {a b : ℤ} (b0 : b ≠ 0) : a /. b ≠ 0 ↔ a ≠ 0 := #align rat.mk_eq Rat.divInt_eq_iff #align rat.div_mk_div_cancel_left Rat.divInt_mul_right --- Porting note: this can move to Std4 +-- Porting note: this can move to Batteries theorem normalize_eq_mk' (n : Int) (d : Nat) (h : d ≠ 0) (c : Nat.gcd (Int.natAbs n) d = 1) : normalize n d h = mk' n d h c := (mk_eq_normalize ..).symm --- TODO: Rename `mkRat_num_den` in Std +-- TODO: Rename `mkRat_num_den` in Batteries @[simp] alias mkRat_num_den' := mkRat_self --- TODO: Rename `Rat.divInt_self` to `Rat.num_divInt_den` in Std +-- TODO: Rename `Rat.divInt_self` to `Rat.num_divInt_den` in Batteries lemma num_divInt_den (q : ℚ) : q.num /. q.den = q := divInt_self _ #align rat.num_denom Rat.num_divInt_den @@ -131,7 +131,7 @@ lemma mk'_eq_divInt {n d h c} : (⟨n, d, h, c⟩ : ℚ) = n /. d := (num_divInt theorem intCast_eq_divInt (z : ℤ) : (z : ℚ) = z /. 1 := mk'_eq_divInt #align rat.coe_int_eq_mk Rat.intCast_eq_divInt --- TODO: Rename `divInt_self` in Std to `num_divInt_den` +-- TODO: Rename `divInt_self` in Batteries to `num_divInt_den` @[simp] lemma divInt_self' {n : ℤ} (hn : n ≠ 0) : n /. n = 1 := by simpa using divInt_mul_right (n := 1) (d := 1) hn @@ -160,7 +160,7 @@ def numDenCasesOn''.{u} {C : ℚ → Sort u} (a : ℚ) #align rat.add Rat.add --- Porting note: there's already an instance for `Add ℚ` is in Std. +-- Porting note: there's already an instance for `Add ℚ` is in Batteries. theorem lift_binop_eq (f : ℚ → ℚ → ℚ) (f₁ : ℤ → ℤ → ℤ → ℤ → ℤ) (f₂ : ℤ → ℤ → ℤ → ℤ → ℤ) (fv : @@ -229,7 +229,7 @@ lemma mk'_mul_mk' (n₁ n₂ : ℤ) (d₁ d₂ : ℕ) (hd₁ hd₂ hnd₁ hnd₂ lemma mul_eq_mkRat (q r : ℚ) : q * r = mkRat (q.num * r.num) (q.den * r.den) := by rw [mul_def, normalize_eq_mkRat] --- TODO: Rename `divInt_eq_iff` in Std to `divInt_eq_divInt` +-- TODO: Rename `divInt_eq_iff` in Batteries to `divInt_eq_divInt` alias divInt_eq_divInt := divInt_eq_iff @[deprecated] alias mul_num_den := mul_eq_mkRat diff --git a/Mathlib/Data/Rat/Field.lean b/Mathlib/Data/Rat/Field.lean index fa9af42709cb3..7aed995b8a3d8 100644 --- a/Mathlib/Data/Rat/Field.lean +++ b/Mathlib/Data/Rat/Field.lean @@ -26,7 +26,7 @@ was defined in `Mathlib.Data.Rat.Defs`. We have to define the field structure in a separate file to avoid cyclic imports: the `Field` class contains a map from `ℚ` (see `Field`'s docstring for the rationale), so we have a dependency `Rat.Field → Field → Rat` that is reflected in the import -hierarchy `Mathlib.Data.Rat.Basic → Mathlib.Algebra.Field.Defs → Std.Data.Rat`. +hierarchy `Mathlib.Data.Rat.Basic → Mathlib.Algebra.Field.Defs → Batteries.Data.Rat`. ## Tags diff --git a/Mathlib/Data/Rat/Init.lean b/Mathlib/Data/Rat/Init.lean index a2a16ec36e4e9..3a7cd10bb8ea3 100644 --- a/Mathlib/Data/Rat/Init.lean +++ b/Mathlib/Data/Rat/Init.lean @@ -17,7 +17,7 @@ import Batteries.Classes.RatCast This file declares `ℚ` notation for the rationals and defines the nonnegative rationals `ℚ≥0`. -This file is eligible to upstreaming to Std. +This file is eligible to upstreaming to Batteries. -/ @[inherit_doc] notation "ℚ" => Rat diff --git a/Mathlib/Data/String/Defs.lean b/Mathlib/Data/String/Defs.lean index 86e7fb82cb50e..d8b8c55547b62 100644 --- a/Mathlib/Data/String/Defs.lean +++ b/Mathlib/Data/String/Defs.lean @@ -32,7 +32,7 @@ def replicate (n : Nat) (c : Char) : String := ⟨List.replicate n c⟩ -- TODO bring this definition in line with the above, either by: --- adding `List.rightpad` to Std and changing the definition of `rightpad` here to match +-- adding `List.rightpad` to Batteries and changing the definition of `rightpad` here to match -- or by changing the definition of `leftpad` above to match this /-- Pad `s : String` with repeated occurrences of `c : Char` on the right until it's of length `n`. If `s` is initially larger than `n`, just return `s`. -/ diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index 77b16d5e073b8..1b781290b7a60 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -11,7 +11,7 @@ import Mathlib.Tactic.MkIffOfInductiveProp /-! # Additional lemmas about sum types -Most of the former contents of this file have been moved to Std. +Most of the former contents of this file have been moved to Batteries. -/ diff --git a/Mathlib/Init/Classes/Order.lean b/Mathlib/Init/Classes/Order.lean index cfb0d03f3f564..dc864617893d0 100644 --- a/Mathlib/Init/Classes/Order.lean +++ b/Mathlib/Init/Classes/Order.lean @@ -8,7 +8,7 @@ import Batteries.Classes.Order import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align ordering.swap_inj Ordering.swap_inj diff --git a/Mathlib/Init/Control/Lawful.lean b/Mathlib/Init/Control/Lawful.lean index d38489800317c..a1cb5eb4d00c0 100644 --- a/Mathlib/Init/Control/Lawful.lean +++ b/Mathlib/Init/Control/Lawful.lean @@ -252,7 +252,8 @@ As discussed in https://github.com/leanprover/std4/pull/416, it should be possible for core to expose the lawfulness of `IO` as part of the opaque interface, which would remove the need for these proofs anyway. -These are not in Std because Std does not want to deal with the churn from such a core refactor. +These are not in Batteries because Batteries does not want to deal with the churn from such a core +refactor. -/ instance : LawfulMonad (EIO ε) := inferInstanceAs <| LawfulMonad (EStateM _ _) diff --git a/Mathlib/Init/Data/Bool/Basic.lean b/Mathlib/Init/Data/Bool/Basic.lean index d7773bc48ac9c..f036fe50831c3 100644 --- a/Mathlib/Init/Data/Bool/Basic.lean +++ b/Mathlib/Init/Data/Bool/Basic.lean @@ -10,7 +10,8 @@ import Batteries.Data.Bool # Boolean operations In Lean 3 this file contained the definitions of `cond`, `bor`, `band` and `bnot`, -the boolean functions. These are now in Lean 4 core or Std (as `cond`, `or`, `and`, `not`, `xor`). +the boolean functions. +These are now in Lean 4 core or Batteries (as `cond`, `or`, `and`, `not`, `xor`). -/ #align bor or diff --git a/Mathlib/Init/Data/Int/DivMod.lean b/Mathlib/Init/Data/Int/DivMod.lean index 28f5327dfe170..bc97ffbddeaeb 100644 --- a/Mathlib/Init/Data/Int/DivMod.lean +++ b/Mathlib/Init/Data/Int/DivMod.lean @@ -8,7 +8,7 @@ import Batteries.Data.Int.DivMod import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align int.zero_div Int.zero_div diff --git a/Mathlib/Init/Data/Int/Lemmas.lean b/Mathlib/Init/Data/Int/Lemmas.lean index e0a2a36a339b0..88d49d6ca1ad0 100644 --- a/Mathlib/Init/Data/Int/Lemmas.lean +++ b/Mathlib/Init/Data/Int/Lemmas.lean @@ -8,7 +8,7 @@ import Batteries.Data.Int.Order import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align int.default_eq_zero Int.default_eq_zero diff --git a/Mathlib/Init/Data/List/Basic.lean b/Mathlib/Init/Data/List/Basic.lean index f0da206a51538..de24518a148be 100644 --- a/Mathlib/Init/Data/List/Basic.lean +++ b/Mathlib/Init/Data/List/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Mathport.Rename import Mathlib.Init.Data.Nat.Notation import Batteries.Data.List.Basic /-! -Definitions for `List` not (yet) in `Std` +Definitions for `List` not (yet) in `Batteries` -/ set_option autoImplicit true diff --git a/Mathlib/Init/Data/List/Instances.lean b/Mathlib/Init/Data/List/Instances.lean index becd08b9a83c5..cdb78d8b6fcf8 100644 --- a/Mathlib/Init/Data/List/Instances.lean +++ b/Mathlib/Init/Data/List/Instances.lean @@ -9,7 +9,7 @@ import Mathlib.Mathport.Rename #align_import init.data.list.instances from "leanprover-community/lean"@"9af482290ef68e8aaa5ead01aa7b09b7be7019fd" /-! -# Decidable and Monad instances for `List` not (yet) in `Std` +# Decidable and Monad instances for `List` not (yet) in `Batteries` -/ universe u v w diff --git a/Mathlib/Init/Data/List/Lemmas.lean b/Mathlib/Init/Data/List/Lemmas.lean index c4dd242a97d14..4411459613028 100644 --- a/Mathlib/Init/Data/List/Lemmas.lean +++ b/Mathlib/Init/Data/List/Lemmas.lean @@ -8,7 +8,7 @@ import Mathlib.Mathport.Rename import Mathlib.Tactic.Cases /-! -Lemmas for `List` not (yet) in `Std` +Lemmas for `List` not (yet) in `Batteries` -/ #align_import init.data.list.lemmas from "leanprover-community/lean"@"4a03bdeb31b3688c31d02d7ff8e0ff2e5d6174db" @@ -89,7 +89,7 @@ theorem not_exists_mem_nil (p : α → Prop) : ¬∃ x ∈ @nil α, p x := /-! list subset -/ #align list.subset List.Subset --- This is relying on an automatically generated instance name from Std. +-- This is relying on an automatically generated instance name from Batteries. #align list.has_subset List.instHasSubset_batteries #align list.nil_subset List.nil_subset #align list.subset.refl List.Subset.refl diff --git a/Mathlib/Init/Data/Option/Basic.lean b/Mathlib/Init/Data/Option/Basic.lean index e949c38c90b26..12f2a8cd1d0f4 100644 --- a/Mathlib/Init/Data/Option/Basic.lean +++ b/Mathlib/Init/Data/Option/Basic.lean @@ -7,7 +7,7 @@ Authors: Johan Commelin import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align option.mem_def Option.mem_def diff --git a/Mathlib/Init/Data/Option/Init/Lemmas.lean b/Mathlib/Init/Data/Option/Init/Lemmas.lean index 1a221a37deb62..91d39b84cfb70 100644 --- a/Mathlib/Init/Data/Option/Init/Lemmas.lean +++ b/Mathlib/Init/Data/Option/Init/Lemmas.lean @@ -7,7 +7,7 @@ Authors: Johan Commelin import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align option.map_none' Option.map_none' diff --git a/Mathlib/Init/Data/Option/Lemmas.lean b/Mathlib/Init/Data/Option/Lemmas.lean index 86f1af2eebe69..17a2f72f0b307 100644 --- a/Mathlib/Init/Data/Option/Lemmas.lean +++ b/Mathlib/Init/Data/Option/Lemmas.lean @@ -8,7 +8,7 @@ import Batteries.Data.Option.Lemmas import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align option.mem_iff Option.mem_iff diff --git a/Mathlib/Init/Data/Rat/Basic.lean b/Mathlib/Init/Data/Rat/Basic.lean index 68bf33568c438..658ae90d7d812 100644 --- a/Mathlib/Init/Data/Rat/Basic.lean +++ b/Mathlib/Init/Data/Rat/Basic.lean @@ -8,7 +8,7 @@ import Batteries.Data.Rat.Lemmas import Mathlib.Mathport.Rename /-! -# Align statements for declarations from Std +# Align statements for declarations from Batteries -/ #align rat.ext Rat.ext diff --git a/Mathlib/Lean/Name.lean b/Mathlib/Lean/Name.lean index 8a021f050e82d..49bbebc08c5d6 100644 --- a/Mathlib/Lean/Name.lean +++ b/Mathlib/Lean/Name.lean @@ -42,7 +42,7 @@ def allNamesByModule (p : Name → Bool) : CoreM (Batteries.HashMap Name (Array (← getEnv).constants.foldM (init := Batteries.HashMap.empty) fun names n _ => do if p n && !(← isBlackListed n) then let some m ← findModuleOf? n | return names - -- TODO use `Std.HashMap.modify` when we bump Std4 (or `alter` if that is written). + -- TODO use `Batteries.HashMap.modify` when we bump Batteries (or `alter` if that is written). match names.find? m with | some others => return names.insert m (others.push n) | none => return names.insert m #[n] diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 8b22edb2b92a7..4d5e8f9761912 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -553,7 +553,7 @@ theorem eq_equivalence {α : Sort*} : Equivalence (@Eq α) := ⟨Eq.refl, @Eq.symm _, @Eq.trans _⟩ #align eq_equivalence eq_equivalence --- These were migrated to Std but the `@[simp]` attributes were (mysteriously?) removed. +-- These were migrated to Batteries but the `@[simp]` attributes were (mysteriously?) removed. attribute [simp] eq_mp_eq_cast eq_mpr_eq_cast #align eq_mp_eq_cast eq_mp_eq_cast @@ -637,7 +637,7 @@ theorem pi_congr {β' : α → Sort _} (h : ∀ a, β a = β' a) : (∀ a, β a) #align pi_congr pi_congr -- Porting note: some higher order lemmas such as `forall₂_congr` and `exists₂_congr` --- were moved to `Std4` +-- were moved to `Batteries` theorem forall₂_imp {p q : ∀ a, β a → Prop} (h : ∀ a b, p a b → q a b) : (∀ a b, p a b) → ∀ a b, q a b := diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index d8d0c092aac33..117877544a465 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -254,7 +254,7 @@ theorem uniqueElim_const {β : Sort*} {_ : Unique ι} (x : β) (i : ι) : end Pi --- TODO: Mario turned this off as a simp lemma in Std, wanting to profile it. +-- TODO: Mario turned this off as a simp lemma in Batteries, wanting to profile it. attribute [local simp] eq_iff_true_of_subsingleton in theorem Unique.bijective {A B} [Unique A] [Unique B] {f : A → B} : Function.Bijective f := by rw [Function.bijective_iff_has_inverse] diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index f6a2e4ffeb3f5..eedcd3e9c8fbc 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -73,7 +73,7 @@ import Mathlib.Tactic.FunProp.Mor import Mathlib.Tactic.FunProp.RefinedDiscrTree import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems -import Mathlib.Tactic.FunProp.ToStd +import Mathlib.Tactic.FunProp.ToBatteries import Mathlib.Tactic.FunProp.Types import Mathlib.Tactic.GCongr import Mathlib.Tactic.GCongr.Core diff --git a/Mathlib/Tactic/Attr/Core.lean b/Mathlib/Tactic/Attr/Core.lean index 7c31a2e729112..55ef5e39e7b18 100644 --- a/Mathlib/Tactic/Attr/Core.lean +++ b/Mathlib/Tactic/Attr/Core.lean @@ -11,7 +11,7 @@ import Batteries.Logic In Lean 4, an attribute declared with `register_simp_attr` cannot be used in the same file. So, we declare all `simp` attributes used in `Mathlib` in `Mathlib/Tactic/Attr/Register` and tag lemmas -from the core library and the `Std` library with these attributes in this file. +from the core library and the `Batteries` library with these attributes in this file. -/ attribute [simp] id_map' diff --git a/Mathlib/Tactic/CC/Datatypes.lean b/Mathlib/Tactic/CC/Datatypes.lean index 0ae0eb83b3eed..8da4d00ab8f9d 100644 --- a/Mathlib/Tactic/CC/Datatypes.lean +++ b/Mathlib/Tactic/CC/Datatypes.lean @@ -23,7 +23,7 @@ This file is ported from C++ code, so many declarations lack documents. universe u -open Lean Meta Elab Tactic Std +open Lean Meta Elab Tactic namespace Mathlib.Tactic.CC diff --git a/Mathlib/Tactic/Congr!.lean b/Mathlib/Tactic/Congr!.lean index 72d21f827df6a..1fa47aa9f98be 100644 --- a/Mathlib/Tactic/Congr!.lean +++ b/Mathlib/Tactic/Congr!.lean @@ -512,7 +512,7 @@ private theorem eq_imp_of_iff_imp {p : x = y → Prop} (h : (he : x ↔ y) → p /-- Does `Lean.MVarId.intros` but then cleans up the introduced hypotheses, removing anything that is trivial. If there are any patterns in the current `CongrMetaM` state then instead -of `Lean.MVarId.intros` it does `Std.Tactic.RCases.rintro`. +of `Lean.MVarId.intros` it does `Lean.Elab..Tactic.RCases.rintro`. Cleaning up includes: - deleting hypotheses of the form `HEq x x`, `x = x`, and `x ↔ x`. diff --git a/Mathlib/Tactic/ExtractGoal.lean b/Mathlib/Tactic/ExtractGoal.lean index de0104e7f5406..56ac437da9864 100644 --- a/Mathlib/Tactic/ExtractGoal.lean +++ b/Mathlib/Tactic/ExtractGoal.lean @@ -47,7 +47,7 @@ example {z : Int} : ∃ n : Nat, ↑n = z := by apply int_eq_nat -- works ``` -However, importing `Std.Classes.Cast`, makes `extract_goal` produce a different theorem +However, importing `Batteries.Classes.Cast`, makes `extract_goal` produce a different theorem ```lean import Batteries.Classes.Cast diff --git a/Mathlib/Tactic/FunProp.lean b/Mathlib/Tactic/FunProp.lean index 995dacd9754e2..24b8c3de402c1 100644 --- a/Mathlib/Tactic/FunProp.lean +++ b/Mathlib/Tactic/FunProp.lean @@ -7,5 +7,5 @@ import Mathlib.Tactic.FunProp.Mor import Mathlib.Tactic.FunProp.RefinedDiscrTree import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems -import Mathlib.Tactic.FunProp.ToStd +import Mathlib.Tactic.FunProp.ToBatteries import Mathlib.Tactic.FunProp.Types diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index c844beec3004f..50257bfacc642 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -7,7 +7,7 @@ import Lean import Batteries.Tactic.Exact import Mathlib.Tactic.FunProp.Theorems -import Mathlib.Tactic.FunProp.ToStd +import Mathlib.Tactic.FunProp.ToBatteries import Mathlib.Tactic.FunProp.Types import Mathlib.Lean.Expr.Basic diff --git a/Mathlib/Tactic/FunProp/Mor.lean b/Mathlib/Tactic/FunProp/Mor.lean index a03f75b8ea753..d90b08481574c 100644 --- a/Mathlib/Tactic/FunProp/Mor.lean +++ b/Mathlib/Tactic/FunProp/Mor.lean @@ -5,7 +5,7 @@ Authors: Tomas Skrivan -/ import Lean import Mathlib.Data.FunLike.Basic -import Mathlib.Tactic.FunProp.ToStd +import Mathlib.Tactic.FunProp.ToBatteries /-! diff --git a/Mathlib/Tactic/FunProp/ToStd.lean b/Mathlib/Tactic/FunProp/ToBatteries.lean similarity index 100% rename from Mathlib/Tactic/FunProp/ToStd.lean rename to Mathlib/Tactic/FunProp/ToBatteries.lean diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index 05125efbda29d..650babe34a113 100644 --- a/Mathlib/Tactic/Linarith/Frontend.lean +++ b/Mathlib/Tactic/Linarith/Frontend.lean @@ -80,7 +80,7 @@ disequality hypotheses, since this would lead to a number of runs exponential in disequalities in the context. The Fourier-Motzkin oracle is very modular. It can easily be replaced with another function of type -`certificateOracle := List Comp → ℕ → TacticM ((Std.HashMap ℕ ℕ))`, +`certificateOracle := List Comp → ℕ → TacticM ((Batteries.HashMap ℕ ℕ))`, which takes a list of comparisons and the largest variable index appearing in those comparisons, and returns a map from comparison indices to coefficients. An alternate oracle can be specified in the `LinarithConfig` object. diff --git a/Mathlib/Tactic/Linarith/Verification.lean b/Mathlib/Tactic/Linarith/Verification.lean index d85aebe47273a..9f4091b9c9f41 100644 --- a/Mathlib/Tactic/Linarith/Verification.lean +++ b/Mathlib/Tactic/Linarith/Verification.lean @@ -204,7 +204,7 @@ def proveFalseByLinarith (cfg : LinarithConfig) : MVarId → List Expr → MetaM trace[linarith.detail] "{comps}" let oracle := cfg.oracle.getD FourierMotzkin.produceCertificate -- perform the elimination and fail if no contradiction is found. - let certificate : Std.HashMap Nat Nat ← try + let certificate : Batteries.HashMap Nat Nat ← try oracle comps max_var catch e => trace[linarith] e.toMessageData diff --git a/Mathlib/Tactic/Lint.lean b/Mathlib/Tactic/Lint.lean index 602bbeb9188b8..e4729098f8630 100644 --- a/Mathlib/Tactic/Lint.lean +++ b/Mathlib/Tactic/Lint.lean @@ -12,7 +12,7 @@ import Batteries.Tactic.Lint In this file we define additional linters for mathlib. -Perhaps these should be moved to Std in the future. +Perhaps these should be moved to Batteries in the future. -/ namespace Std.Tactic.Lint diff --git a/Mathlib/Tactic/Monotonicity/Basic.lean b/Mathlib/Tactic/Monotonicity/Basic.lean index d826233293671..0b0ba3b09f108 100644 --- a/Mathlib/Tactic/Monotonicity/Basic.lean +++ b/Mathlib/Tactic/Monotonicity/Basic.lean @@ -24,7 +24,7 @@ for this in Lean 3 was `mono*`. Both `mono` and `mono*` implement this behavior -/ open Lean Elab Tactic Parser Tactic -open Std Tactic SolveByElim +open Tactic SolveByElim namespace Mathlib.Tactic.Monotonicity diff --git a/Mathlib/Tactic/Nontriviality/Core.lean b/Mathlib/Tactic/Nontriviality/Core.lean index 6343a1b9c3aba..91e22c1f077f6 100644 --- a/Mathlib/Tactic/Nontriviality/Core.lean +++ b/Mathlib/Tactic/Nontriviality/Core.lean @@ -12,7 +12,7 @@ import Mathlib.Tactic.Attr.Core set_option autoImplicit true namespace Mathlib.Tactic.Nontriviality -open Lean Elab Meta Tactic Qq Std.Tactic +open Lean Elab Meta Tactic Qq theorem subsingleton_or_nontrivial_elim {p : Prop} {α : Type u} (h₁ : Subsingleton α → p) (h₂ : Nontrivial α → p) : p := diff --git a/Mathlib/Tactic/Replace.lean b/Mathlib/Tactic/Replace.lean index 3c840a1603498..7a4da71738e92 100644 --- a/Mathlib/Tactic/Replace.lean +++ b/Mathlib/Tactic/Replace.lean @@ -8,7 +8,7 @@ import Mathlib.Tactic.Have /-! # Extending `replace` -This file extends the `replace` tactic from `Std` to allow the addition of hypotheses to +This file extends the `replace` tactic from `Batteries` to allow the addition of hypotheses to the context without requiring their proofs to be provided immediately. As a style choice, this should not be used in mathlib; but is provided for downstream users who diff --git a/Mathlib/Util/Delaborators.lean b/Mathlib/Util/Delaborators.lean index de400557aca97..a37232c230318 100644 --- a/Mathlib/Util/Delaborators.lean +++ b/Mathlib/Util/Delaborators.lean @@ -29,7 +29,7 @@ def piNotation := leading_parser:leadPrec /-- Dependent function type (a "pi type"). The notation `Π x ∈ s, β x` is short for `Π x, x ∈ s → β x`. -/ --- A copy of forall notation from `Std.Util.ExtendedBinder` for pi notation +-- A copy of forall notation from `Batteries.Util.ExtendedBinder` for pi notation syntax "Π " binderIdent binderPred ", " term : term macro_rules diff --git a/docs/Conv/Guide.lean b/docs/Conv/Guide.lean index b419d8ff53302..a44f6d7745ab3 100644 --- a/docs/Conv/Guide.lean +++ b/docs/Conv/Guide.lean @@ -58,9 +58,9 @@ any remaining goals (note that in `conv` mode, every goal can be solved for by ` then it uses the resulting `lhs = rhs` proof to rewrite the goal in the surrounding normal tactic mode. -## Conv tactics from Lean 4, Std4, and Mathlib4 +## Conv tactics from Lean 4, Batteries, and Mathlib -Unless they're annotated with "Std4" or "Mathlib", the following tactics are defined +Unless they're annotated with "Batteries" or "Mathlib", the following tactics are defined in Lean 4 core. ### Control @@ -157,7 +157,7 @@ in Lean 4 core. * `change t` changes the expression to `t` if the expression and `t` are definitionally equal. * `equals t => tacticSeq` changes the current expression, say `e`, to `t`, and asks you to prove - the equality `e = t`. (Std4) + the equality `e = t`. (Batteries) * `rw [thms...]` rewrites the expression using the given theorems. The syntax is similar to `rw`. @@ -175,7 +175,7 @@ in Lean 4 core. * `refine e` applies `e` to the goal (which remember is `⊢ lhs = ?rhs`) using the `refine` tactic. Strange results may occur if the placeholders in `e` are not equalities. -* `exact e` closes the goal, where `e : lhs = ?rhs`. (Std4) +* `exact e` closes the goal, where `e : lhs = ?rhs`. (Batteries) * Mathlib provides a number of tactics as `conv` tactics: * `abel` and `abel_nf` @@ -214,9 +214,10 @@ and expands all local variables). * `fail_if_success convSeq` fails if the `conv` sequence succeeds. -* `guard_expr` and `guard_target` for asserting that certain expressions are equal to others. (Std4) +* `guard_expr` and `guard_target` for asserting that certain expressions are equal to others. + (Batteries) -* `unreachable!`, which is the same as the `unreachable!` tactic. (Std4) +* `unreachable!`, which is the same as the `unreachable!` tactic. (Batteriess) * `run_tac doSeq` evaluates a monadic value and runs it as a tactic using `tactic'`. (Mathlib) diff --git a/docs/overview.yaml b/docs/overview.yaml index 5e56b885b2459..995a355786065 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -453,7 +453,7 @@ Data structures: list: 'List' array: 'Array' # buffer: 'buffer' - difference list: 'Std.DList' + difference list: 'Batteries.DList' lazy list: 'LazyList' stream: 'Stream''' sequence: 'Stream''.Seq' @@ -467,14 +467,14 @@ Data structures: Maps: key-value map: 'AList' - red-black map: 'Std.RBMap' - hash map: 'Std.HashMap' + red-black map: 'Batteries.RBMap' + hash map: 'Batteries.HashMap' finitely supported function: 'Finsupp' finite map: 'Finmap' Trees: tree: 'Tree' - red-black tree: 'Std.RBSet' + red-black tree: 'Batteries.RBSet' size-balanced binary search tree: 'Ordnode' W type: 'WType' diff --git a/scripts/nolints.json b/scripts/nolints.json index fef5fd0e976eb..0c0e267bdbc8f 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -789,11 +789,6 @@ ["docBlame", "SlimCheck.TestResult.toString"], ["docBlame", "SlimCheck.Testable.run"], ["docBlame", "SlimCheck.Testable.runProp"], - ["docBlame", "Std.Prec.arrow"], - ["docBlame", "Std.Prec.max"], - ["docBlame", "Std.Prec.maxPlus"], - ["docBlame", "Std.Priority.default"], - ["docBlame", "Std.Priority.max"], ["docBlame", "Stream'.WSeq.«term_~ʷ_»"], ["docBlame", "String.toAsciiByteArray.loop"], ["docBlame", "Submodule.quotientPi_aux.invFun"], diff --git a/scripts/noshake.json b/scripts/noshake.json index 455cf27ca9276..08fa3decea166 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -239,7 +239,7 @@ "Mathlib.Tactic.GCongr": ["Mathlib.Tactic.Positivity.Core"], "Mathlib.Tactic.FunProp.Types": ["Mathlib.Logic.Function.Basic"], "Mathlib.Tactic.FunProp.RefinedDiscrTree": ["Mathlib.Algebra.Group.Pi.Basic"], - "Mathlib.Tactic.FunProp.MorExt": ["Mathlib.Tactic.FunProp.ToStd"], + "Mathlib.Tactic.FunProp.MorExt": ["Mathlib.Tactic.FunProp.ToBatteries"], "Mathlib.Tactic.FunProp.Mor": ["Mathlib.Data.FunLike.Basic"], "Mathlib.Tactic.FunProp.Measurable": ["Mathlib.Analysis.Calculus.ContDiff.Basic", From 3d7266526f9c3398ed9322475a680d3ec836b6b8 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Tue, 7 May 2024 18:06:04 +0000 Subject: [PATCH 18/52] chore(Order.Copy): clean up structure instance constructions (#12729) After the port, this used a `refine' {foo := foo, bar := bar, ..}` pattern to build the instances. We clean this up. --- Mathlib/Order/Copy.lean | 129 +++++++++++++++++----------------------- 1 file changed, 53 insertions(+), 76 deletions(-) diff --git a/Mathlib/Order/Copy.lean b/Mathlib/Order/Copy.lean index 1abe3e7ecb3ef..23c7da50b6035 100644 --- a/Mathlib/Order/Copy.lean +++ b/Mathlib/Order/Copy.lean @@ -22,8 +22,6 @@ universe u variable {α : Type u} --- Porting note: mathlib3 proof uses `refine { top := top, bot := bot, .. }` but this does not work --- anymore /-- A function to create a provable equal copy of a bounded order with possibly different definitional equalities. -/ def BoundedOrder.copy {h : LE α} {h' : LE α} (c : @BoundedOrder α h') @@ -34,51 +32,37 @@ def BoundedOrder.copy {h : LE α} {h' : LE α} (c : @BoundedOrder α h') (@OrderBot.mk α h { bot := bot } (fun _ ↦ by simp [eq_bot, le_eq])) #align bounded_order.copy BoundedOrder.copy --- Porting note: original proof uses --- `all_goals { abstract { subst_vars, casesI c, simp_rw le_eq, assumption } }` /-- A function to create a provable equal copy of a lattice with possibly different definitional equalities. -/ def Lattice.copy (c : Lattice α) (le : α → α → Prop) (eq_le : le = (by infer_instance : LE α).le) (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) - (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) : Lattice α := by - refine' { le := le, sup := sup, inf := inf, lt := fun a b ↦ le a b ∧ ¬ le b a.. } - · intros; simp [eq_le] - · intro _ _ _ hab hbc; rw [eq_le] at hab hbc ⊢; exact le_trans hab hbc - · intros; simp [eq_le] - · intro _ _ hab hba; simp_rw [eq_le] at hab hba; exact le_antisymm hab hba - · intros; simp [eq_le, eq_sup] - · intros; simp [eq_le, eq_sup] - · intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_sup, hac, hbc] - · intros; simp [eq_le, eq_inf] - · intros; simp [eq_le, eq_inf] - · intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_inf, hac, hbc] + (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) : Lattice α where + le := le + sup := sup + inf := inf + lt := fun a b ↦ le a b ∧ ¬ le b a + le_refl := by intros; simp [eq_le] + le_trans := by intro _ _ _ hab hbc; rw [eq_le] at hab hbc ⊢; exact le_trans hab hbc + le_antisymm := by intro _ _ hab hba; simp_rw [eq_le] at hab hba; exact le_antisymm hab hba + le_sup_left := by intros; simp [eq_le, eq_sup] + le_sup_right := by intros; simp [eq_le, eq_sup] + sup_le := by intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_sup, hac, hbc] + inf_le_left := by intros; simp [eq_le, eq_inf] + inf_le_right := by intros; simp [eq_le, eq_inf] + le_inf := by intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_inf, hac, hbc] #align lattice.copy Lattice.copy --- Porting note: original proof uses --- `all_goals { abstract { subst_vars, casesI c, simp_rw le_eq, assumption } }` /-- A function to create a provable equal copy of a distributive lattice with possibly different definitional equalities. -/ def DistribLattice.copy (c : DistribLattice α) (le : α → α → Prop) (eq_le : le = (by infer_instance : LE α).le) (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) - (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) : DistribLattice α := by - refine' { le := le, sup := sup, inf := inf, lt := fun a b ↦ le a b ∧ ¬ le b a.. } - · intros; simp [eq_le] - · intro _ _ _ hab hbc; rw [eq_le] at hab hbc ⊢; exact le_trans hab hbc - · intros; simp [eq_le] - · intro _ _ hab hba; simp_rw [eq_le] at hab hba; exact le_antisymm hab hba - · intros; simp [eq_le, eq_sup] - · intros; simp [eq_le, eq_sup] - · intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_sup, hac, hbc] - · intros; simp [eq_le, eq_inf] - · intros; simp [eq_le, eq_inf] - · intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_inf, hac, hbc] - · intros; simp [eq_le, eq_inf, eq_sup, le_sup_inf] + (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) : DistribLattice α where + toLattice := Lattice.copy (@DistribLattice.toLattice α c) le eq_le sup eq_sup inf eq_inf + le_sup_inf := by intros; simp [eq_le, eq_sup, eq_inf, le_sup_inf] #align distrib_lattice.copy DistribLattice.copy --- Porting note: original proof uses --- `all_goals { abstract { subst_vars, casesI c, simp_rw le_eq, assumption } }` /-- A function to create a provable equal copy of a complete lattice with possibly different definitional equalities. -/ def CompleteLattice.copy (c : CompleteLattice α) @@ -89,19 +73,20 @@ def CompleteLattice.copy (c : CompleteLattice α) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : - CompleteLattice α := by - refine' { Lattice.copy (@CompleteLattice.toLattice α c) le eq_le sup eq_sup inf eq_inf with - le := le, top := top, bot := bot, sup := sup, inf := inf, sSup := sSup, sInf := sInf.. } - · intro _ _ h; simp [eq_le, eq_sSup, le_sSup _ _ h] - · intro _ _ h; simpa [eq_le, eq_sSup] using h - · intro _ _ h; simp [eq_le, eq_sInf, sInf_le _ _ h] - · intro _ _ h; simpa [eq_le, eq_sInf] using h - · intros; simp [eq_le, eq_top] - · intros; simp [eq_le, eq_bot] + CompleteLattice α where + toLattice := Lattice.copy (@CompleteLattice.toLattice α c) le eq_le sup eq_sup inf eq_inf + top := top + bot := bot + sSup := sSup + sInf := sInf + le_sSup := by intro _ _ h; simp [eq_le, eq_sSup, le_sSup _ _ h] + sSup_le := by intro _ _ h; simpa [eq_le, eq_sSup] using h + sInf_le := by intro _ _ h; simp [eq_le, eq_sInf, sInf_le _ _ h] + le_sInf := by intro _ _ h; simpa [eq_le, eq_sInf] using h + le_top := by intros; simp [eq_le, eq_top] + bot_le := by intros; simp [eq_le, eq_bot] #align complete_lattice.copy CompleteLattice.copy --- Porting note: original proof uses --- `all_goals { abstract { subst_vars, casesI c, simp_rw le_eq, assumption } }` /-- A function to create a provable equal copy of a frame with possibly different definitional equalities. -/ def Frame.copy (c : Frame α) (le : α → α → Prop) (eq_le : le = (by infer_instance : LE α).le) @@ -110,11 +95,11 @@ def Frame.copy (c : Frame α) (le : α → α → Prop) (eq_le : le = (by infer_ (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) - (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : Frame α := - { 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 with - 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] } + (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] #align frame.copy Frame.copy -- Porting note: original proof uses @@ -127,11 +112,11 @@ def Coframe.copy (c : Coframe α) (le : α → α → Prop) (eq_le : le = (by in (sup : α → α → α) (eq_sup : sup = (by infer_instance : Sup α).sup) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) - (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : Coframe α := - { 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 with - 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] } + (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] #align coframe.copy Coframe.copy /-- A function to create a provable equal copy of a complete distributive lattice @@ -144,11 +129,11 @@ def CompleteDistribLattice.copy (c : CompleteDistribLattice α) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : - CompleteDistribLattice α := - { 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 with } + 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 #align complete_distrib_lattice.copy CompleteDistribLattice.copy -- Porting note: original proof uses @@ -161,21 +146,13 @@ def ConditionallyCompleteLattice.copy (c : ConditionallyCompleteLattice α) (inf : α → α → α) (eq_inf : inf = (by infer_instance : Inf α).inf) (sSup : Set α → α) (eq_sSup : sSup = (by infer_instance : SupSet α).sSup) (sInf : Set α → α) (eq_sInf : sInf = (by infer_instance : InfSet α).sInf) : - ConditionallyCompleteLattice α := by - refine' { le := le, sup := sup, inf := inf, sSup := sSup, sInf := sInf.. } - · intro a b; exact le a b ∧ ¬ le b a - · intros; simp [eq_le] - · intro _ _ _ hab hbc; rw [eq_le] at hab hbc ⊢; exact le_trans hab hbc - · intros; simp [eq_le] - · intro _ _ hab hba; simp_rw [eq_le] at hab hba; exact le_antisymm hab hba - · intros; simp [eq_le, eq_sup] - · intros; simp [eq_le, eq_sup] - · intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_sup, hac, hbc] - · intros; simp [eq_le, eq_inf] - · intros; simp [eq_le, eq_inf] - · intro _ _ _ hac hbc; simp_rw [eq_le] at hac hbc ⊢; simp [eq_inf, hac, hbc] - · intro _ _ hb h; subst_vars; exact le_csSup _ _ hb h - · intro _ _ hb h; subst_vars; exact csSup_le _ _ hb h - · intro _ _ hb h; subst_vars; exact csInf_le _ _ hb h - · intro _ _ hb h; subst_vars; exact le_csInf _ _ hb h + ConditionallyCompleteLattice α where + toLattice := Lattice.copy (@ConditionallyCompleteLattice.toLattice α c) + le eq_le sup eq_sup inf eq_inf + sSup := sSup + sInf := sInf + le_csSup := by intro _ _ hb h; subst_vars; exact le_csSup _ _ hb h + csSup_le := by intro _ _ hb h; subst_vars; exact csSup_le _ _ hb h + csInf_le := by intro _ _ hb h; subst_vars; exact csInf_le _ _ hb h + le_csInf := by intro _ _ hb h; subst_vars; exact le_csInf _ _ hb h #align conditionally_complete_lattice.copy ConditionallyCompleteLattice.copy From e0219f6c814ea3790a743b8f316bce697d79dcf0 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Tue, 7 May 2024 18:06:05 +0000 Subject: [PATCH 19/52] chore: split Mathlib.Algebra.Lie.Killing (#12735) --- Mathlib.lean | 2 + Mathlib/Algebra/Lie/Killing.lean | 700 +---------------------- Mathlib/Algebra/Lie/TraceForm.lean | 428 ++++++++++++++ Mathlib/Algebra/Lie/Weights/Killing.lean | 284 +++++++++ 4 files changed, 732 insertions(+), 682 deletions(-) create mode 100644 Mathlib/Algebra/Lie/TraceForm.lean create mode 100644 Mathlib/Algebra/Lie/Weights/Killing.lean diff --git a/Mathlib.lean b/Mathlib.lean index df4ae97c027c1..2fd01119c36fe 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -367,10 +367,12 @@ import Mathlib.Algebra.Lie.Solvable import Mathlib.Algebra.Lie.Subalgebra import Mathlib.Algebra.Lie.Submodule import Mathlib.Algebra.Lie.TensorProduct +import Mathlib.Algebra.Lie.TraceForm import Mathlib.Algebra.Lie.UniversalEnveloping import Mathlib.Algebra.Lie.Weights.Basic import Mathlib.Algebra.Lie.Weights.Cartan import Mathlib.Algebra.Lie.Weights.Chain +import Mathlib.Algebra.Lie.Weights.Killing import Mathlib.Algebra.Lie.Weights.Linear import Mathlib.Algebra.LinearRecurrence import Mathlib.Algebra.ModEq diff --git a/Mathlib/Algebra/Lie/Killing.lean b/Mathlib/Algebra/Lie/Killing.lean index 3fbe54a6a81dc..dba9091d91333 100644 --- a/Mathlib/Algebra/Lie/Killing.lean +++ b/Mathlib/Algebra/Lie/Killing.lean @@ -3,414 +3,38 @@ Copyright (c) 2023 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ -import Mathlib.Algebra.DirectSum.LinearMap -import Mathlib.Algebra.GroupWithZero.NonZeroDivisors -import Mathlib.Algebra.Lie.Nilpotent import Mathlib.Algebra.Lie.Semisimple -import Mathlib.Algebra.Lie.Weights.Cartan -import Mathlib.Algebra.Lie.Weights.Chain -import Mathlib.Algebra.Lie.Weights.Linear -import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure -import Mathlib.LinearAlgebra.BilinearForm.Orthogonal -import Mathlib.LinearAlgebra.PID -import Mathlib.LinearAlgebra.Trace +import Mathlib.Algebra.Lie.TraceForm /-! -# The trace and Killing forms of a Lie algebra. +# Lie algebras with non-degenerate Killing forms. -Let `L` be a Lie algebra with coefficients in a commutative ring `R`. Suppose `M` is a finite, free -`R`-module and we have a representation `φ : L → End M`. This data induces a natural bilinear form -`B` on `L`, called the trace form associated to `M`; it is defined as `B(x, y) = Tr (φ x) (φ y)`. +In characteristic zero, the following three conditions are equivalent: + 1. The solvable radical of a Lie algebra is trivial + 2. A Lie algebra is a direct sum of its simple ideals + 3. A Lie algebra has non-degenerate Killing form -In the special case that `M` is `L` itself and `φ` is the adjoint representation, the trace form -is known as the Killing form. +In positive characteristic, it is still true that 3 implies 2, and that 2 implies 1, but there are +counterexamples to the remaining implications. Thus condition 3 is the strongest assumption. +Furthermore, much of the Cartan-Killing classification of semisimple Lie algebras in characteristic +zero, continues to hold in positive characteristic (over a perfect field) if the Lie algebra has a +non-degenerate Killing form. -We define the trace / Killing form in this file and prove some basic properties. +This file contains basic definitions and results for such Lie algebras. ## Main definitions - - * `LieModule.traceForm`: a finite, free representation of a Lie algebra `L` induces a bilinear form - on `L` called the trace Form. - * `LieModule.traceForm_eq_zero_of_isNilpotent`: the trace form induced by a nilpotent - representation of a Lie algebra vanishes. - * `killingForm`: the adjoint representation of a (finite, free) Lie algebra `L` induces a bilinear - form on `L` via the trace form construction. * `LieAlgebra.IsKilling`: a typeclass encoding the fact that a Lie algebra has a non-singular Killing form. - * `LieAlgebra.IsKilling.ker_restrict_eq_bot_of_isCartanSubalgebra`: if the Killing form of - a Lie algebra is non-singular, it remains non-singular when restricted to a Cartan subalgebra. - * `LieAlgebra.IsKilling.isSemisimple`: if a Lie algebra has non-singular Killing form then it is - semisimple. - * `LieAlgebra.IsKilling.instIsLieAbelianOfIsCartanSubalgebra`: if the Killing form of a Lie - algebra is non-singular, then its Cartan subalgebras are Abelian. - * `LieAlgebra.IsKilling.span_weight_eq_top`: given a splitting Cartan subalgebra `H` of a - finite-dimensional Lie algebra with non-singular Killing form, the corresponding roots span the - dual space of `H`. - * `LieAlgebra.IsKilling.coroot`: the coroot corresponding to a root. - * `LieAlgebra.IsKilling.isCompl_ker_weight_span_coroot`: given a root `α` with respect to a Cartan - subalgebra `H`, we have a natural decomposition of `H` as the kernel of `α` and the span of the - coroot corresponding to `α`. + * `LieAlgebra.IsKilling.instIsSemisimple`: if a Lie algebra has non-singular Killing form then it + is semisimple. ## TODO * Prove that in characteristic zero, a semisimple Lie algebra has non-singular Killing form. --/ - -variable (R K L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] - [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] - [Module.Free R M] [Module.Finite R M] - [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [FiniteDimensional K M] - -local notation "φ" => LieModule.toEndomorphism R L M - -open LinearMap (trace) -open Set BigOperators FiniteDimensional - -namespace LieModule - -/-- A finite, free representation of a Lie algebra `L` induces a bilinear form on `L` called -the trace Form. See also `killingForm`. -/ -noncomputable def traceForm : LinearMap.BilinForm R L := - ((LinearMap.mul _ _).compl₁₂ (φ).toLinearMap (φ).toLinearMap).compr₂ (trace R M) - -lemma traceForm_apply_apply (x y : L) : - traceForm R L M x y = trace R _ (φ x ∘ₗ φ y) := - rfl - -lemma traceForm_comm (x y : L) : traceForm R L M x y = traceForm R L M y x := - LinearMap.trace_mul_comm R (φ x) (φ y) - -lemma traceForm_isSymm : LinearMap.IsSymm (traceForm R L M) := LieModule.traceForm_comm R L M - -@[simp] lemma traceForm_flip : LinearMap.flip (traceForm R L M) = traceForm R L M := - Eq.symm <| LinearMap.ext₂ <| traceForm_comm R L M - -/-- The trace form of a Lie module is compatible with the action of the Lie algebra. - -See also `LieModule.traceForm_apply_lie_apply'`. -/ -lemma traceForm_apply_lie_apply (x y z : L) : - traceForm R L M ⁅x, y⁆ z = traceForm R L M x ⁅y, z⁆ := by - calc traceForm R L M ⁅x, y⁆ z - = trace R _ (φ ⁅x, y⁆ ∘ₗ φ z) := by simp only [traceForm_apply_apply] - _ = trace R _ ((φ x * φ y - φ y * φ x) * φ z) := ?_ - _ = trace R _ (φ x * (φ y * φ z)) - trace R _ (φ y * (φ x * φ z)) := ?_ - _ = trace R _ (φ x * (φ y * φ z)) - trace R _ (φ x * (φ z * φ y)) := ?_ - _ = traceForm R L M x ⁅y, z⁆ := ?_ - · simp only [LieHom.map_lie, Ring.lie_def, ← LinearMap.mul_eq_comp] - · simp only [sub_mul, mul_sub, map_sub, mul_assoc] - · simp only [LinearMap.trace_mul_cycle' R (φ x) (φ z) (φ y)] - · simp only [traceForm_apply_apply, LieHom.map_lie, Ring.lie_def, mul_sub, map_sub, - ← LinearMap.mul_eq_comp] - -/-- Given a representation `M` of a Lie algebra `L`, the action of any `x : L` is skew-adjoint wrt -the trace form. -/ -lemma traceForm_apply_lie_apply' (x y z : L) : - traceForm R L M ⁅x, y⁆ z = - traceForm R L M y ⁅x, z⁆ := - calc traceForm R L M ⁅x, y⁆ z - = - traceForm R L M ⁅y, x⁆ z := by rw [← lie_skew x y, map_neg, LinearMap.neg_apply] - _ = - traceForm R L M y ⁅x, z⁆ := by rw [traceForm_apply_lie_apply] - -/-- This lemma justifies the terminology "invariant" for trace forms. -/ -@[simp] lemma lie_traceForm_eq_zero (x : L) : ⁅x, traceForm R L M⁆ = 0 := by - ext y z - rw [LieHom.lie_apply, LinearMap.sub_apply, Module.Dual.lie_apply, LinearMap.zero_apply, - LinearMap.zero_apply, traceForm_apply_lie_apply', sub_self] - -@[simp] lemma traceForm_eq_zero_of_isNilpotent [IsReduced R] [IsNilpotent R L M] : - traceForm R L M = 0 := by - ext x y - simp only [traceForm_apply_apply, LinearMap.zero_apply, ← isNilpotent_iff_eq_zero] - apply LinearMap.isNilpotent_trace_of_isNilpotent - exact isNilpotent_toEndomorphism_of_isNilpotent₂ R L M x y - -@[simp] -lemma traceForm_weightSpace_eq [IsDomain R] [IsPrincipalIdealRing R] - [LieAlgebra.IsNilpotent R L] [IsNoetherian R M] [LinearWeights R L M] (χ : L → R) (x y : L) : - traceForm R L (weightSpace M χ) x y = finrank R (weightSpace M χ) • (χ x * χ y) := by - set d := finrank R (weightSpace M χ) - have h₁ : χ y • d • χ x - χ y • χ x • (d : R) = 0 := by simp [mul_comm (χ x)] - have h₂ : χ x • d • χ y = d • (χ x * χ y) := by - simpa [nsmul_eq_mul, smul_eq_mul] using mul_left_comm (χ x) d (χ y) - have := traceForm_eq_zero_of_isNilpotent R L (shiftedWeightSpace R L M χ) - replace this := LinearMap.congr_fun (LinearMap.congr_fun this x) y - rwa [LinearMap.zero_apply, LinearMap.zero_apply, traceForm_apply_apply, - shiftedWeightSpace.toEndomorphism_eq, shiftedWeightSpace.toEndomorphism_eq, - ← LinearEquiv.conj_comp, LinearMap.trace_conj', LinearMap.comp_sub, LinearMap.sub_comp, - LinearMap.sub_comp, map_sub, map_sub, map_sub, LinearMap.comp_smul, LinearMap.smul_comp, - LinearMap.comp_id, LinearMap.id_comp, LinearMap.map_smul, LinearMap.map_smul, - trace_toEndomorphism_weightSpace, trace_toEndomorphism_weightSpace, - LinearMap.comp_smul, LinearMap.smul_comp, LinearMap.id_comp, map_smul, map_smul, - LinearMap.trace_id, ← traceForm_apply_apply, h₁, h₂, sub_zero, sub_eq_zero] at this - -/-- The upper and lower central series of `L` are orthogonal wrt the trace form of any Lie module -`M`. -/ -lemma traceForm_eq_zero_if_mem_lcs_of_mem_ucs {x y : L} (k : ℕ) - (hx : x ∈ (⊤ : LieIdeal R L).lcs L k) (hy : y ∈ (⊥ : LieIdeal R L).ucs k) : - traceForm R L M x y = 0 := by - induction' k with k ih generalizing x y - · replace hy : y = 0 := by simpa using hy - simp [hy] - · rw [LieSubmodule.ucs_succ, LieSubmodule.mem_normalizer] at hy - simp_rw [LieIdeal.lcs_succ, ← LieSubmodule.mem_coeSubmodule, - LieSubmodule.lieIdeal_oper_eq_linear_span', LieSubmodule.mem_top, true_and] at hx - refine Submodule.span_induction hx ?_ ?_ (fun z w hz hw ↦ ?_) (fun t z hz ↦ ?_) - · rintro - ⟨z, w, hw, rfl⟩ - rw [← lie_skew, map_neg, LinearMap.neg_apply, neg_eq_zero, traceForm_apply_lie_apply] - exact ih hw (hy _) - · simp - · simp [hz, hw] - · simp [hz] - -lemma traceForm_apply_eq_zero_of_mem_lcs_of_mem_center {x y : L} - (hx : x ∈ lowerCentralSeries R L L 1) (hy : y ∈ LieAlgebra.center R L) : - traceForm R L M x y = 0 := by - apply traceForm_eq_zero_if_mem_lcs_of_mem_ucs R L M 1 - · simpa using hx - · simpa using hy - --- This is barely worth having: it usually follows from `LieModule.traceForm_eq_zero_of_isNilpotent` -@[simp] lemma traceForm_eq_zero_of_isTrivial [IsTrivial L M] : - traceForm R L M = 0 := by - ext x y - suffices φ x ∘ₗ φ y = 0 by simp [traceForm_apply_apply, this] - ext m - simp - -/-- Given a bilinear form `B` on a representation `M` of a nilpotent Lie algebra `L`, if `B` is -invariant (in the sense that the action of `L` is skew-adjoint wrt `B`) then components of the -Fitting decomposition of `M` are orthogonal wrt `B`. -/ -lemma eq_zero_of_mem_weightSpace_mem_posFitting [LieAlgebra.IsNilpotent R L] - {B : LinearMap.BilinForm R M} (hB : ∀ (x : L) (m n : M), B ⁅x, m⁆ n = - B m ⁅x, n⁆) - {m₀ m₁ : M} (hm₀ : m₀ ∈ weightSpace M (0 : L → R)) (hm₁ : m₁ ∈ posFittingComp R L M) : - B m₀ m₁ = 0 := by - replace hB : ∀ x (k : ℕ) m n, B m ((φ x ^ k) n) = (- 1 : R) ^ k • B ((φ x ^ k) m) n := by - intro x k - induction k with - | zero => simp - | succ k ih => - intro m n - replace hB : ∀ m, B m (φ x n) = (- 1 : R) • B (φ x m) n := by simp [hB] - have : (-1 : R) ^ k • (-1 : R) = (-1 : R) ^ (k + 1) := by rw [pow_succ (-1 : R), smul_eq_mul] - conv_lhs => rw [pow_succ, LinearMap.mul_eq_comp, LinearMap.comp_apply, ih, hB, - ← (φ x).comp_apply, ← LinearMap.mul_eq_comp, ← pow_succ', ← smul_assoc, this] - suffices ∀ (x : L) m, m ∈ posFittingCompOf R M x → B m₀ m = 0 by - apply LieSubmodule.iSup_induction _ hm₁ this (map_zero _) - aesop - clear hm₁ m₁; intro x m₁ hm₁ - simp only [mem_weightSpace, Pi.zero_apply, zero_smul, sub_zero] at hm₀ - obtain ⟨k, hk⟩ := hm₀ x - obtain ⟨m, rfl⟩ := (mem_posFittingCompOf R x m₁).mp hm₁ k - simp [hB, hk] - -lemma trace_toEndomorphism_eq_zero_of_mem_lcs - {k : ℕ} {x : L} (hk : 1 ≤ k) (hx : x ∈ lowerCentralSeries R L L k) : - trace R _ (toEndomorphism R L M x) = 0 := by - replace hx : x ∈ lowerCentralSeries R L L 1 := antitone_lowerCentralSeries _ _ _ hk hx - replace hx : x ∈ Submodule.span R {m | ∃ u v : L, ⁅u, v⁆ = m} := by - rw [lowerCentralSeries_succ, ← LieSubmodule.mem_coeSubmodule, - LieSubmodule.lieIdeal_oper_eq_linear_span'] at hx - simpa using hx - refine Submodule.span_induction (p := fun x ↦ trace R _ (toEndomorphism R L M x) = 0) hx - (fun y ⟨u, v, huv⟩ ↦ ?_) ?_ (fun u v hu hv ↦ ?_) (fun t u hu ↦ ?_) - · simp [← huv] - · simp - · simp [hu, hv] - · simp [hu] - -@[simp] -lemma traceForm_lieSubalgebra_mk_left (L' : LieSubalgebra R L) {x : L} (hx : x ∈ L') (y : L') : - traceForm R L' M ⟨x, hx⟩ y = traceForm R L M x y := - rfl - -@[simp] -lemma traceForm_lieSubalgebra_mk_right (L' : LieSubalgebra R L) {x : L'} {y : L} (hy : y ∈ L') : - traceForm R L' M x ⟨y, hy⟩ = traceForm R L M x y := - rfl - -open TensorProduct - -variable [LieAlgebra.IsNilpotent R L] [IsDomain R] [IsPrincipalIdealRing R] - -lemma traceForm_eq_sum_weightSpaceOf [IsTriangularizable R L M] (z : L) : - traceForm R L M = - ∑ χ in (finite_weightSpaceOf_ne_bot R L M z).toFinset, traceForm R L (weightSpaceOf M χ z) := by - ext x y - have hxy : ∀ χ : R, MapsTo ((toEndomorphism R L M x).comp (toEndomorphism R L M y)) - (weightSpaceOf M χ z) (weightSpaceOf M χ z) := - fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm - have hfin : {χ : R | (weightSpaceOf M χ z : Submodule R M) ≠ ⊥}.Finite := by - convert finite_weightSpaceOf_ne_bot R L M z - exact LieSubmodule.coeSubmodule_eq_bot_iff (weightSpaceOf M _ _) - classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpaceOf R L M z) - (IsTriangularizable.iSup_eq_top z) - simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, - LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] - exact Finset.sum_congr (by simp) (fun χ _ ↦ rfl) - --- In characteristic zero (or even just `LinearWeights R L M`) a stronger result holds (no --- `⊓ LieAlgebra.center R L`) TODO prove this using `LieModule.traceForm_eq_sum_finrank_nsmul_mul`. -lemma lowerCentralSeries_one_inf_center_le_ker_traceForm : - lowerCentralSeries R L L 1 ⊓ LieAlgebra.center R L ≤ LinearMap.ker (traceForm R L M) := by - /- Sketch of proof (due to Zassenhaus): - - Let `z ∈ lowerCentralSeries R L L 1 ⊓ LieAlgebra.center R L` and `x : L`. We must show that - `trace (φ x ∘ φ z) = 0` where `φ z : End R M` indicates the action of `z` on `M` (and likewise - for `φ x`). - - Because `z` belongs to the indicated intersection, it has two key properties: - (a) the trace of the action of `z` vanishes on any Lie module of `L` - (see `LieModule.trace_toEndomorphism_eq_zero_of_mem_lcs`), - (b) `z` commutes with all elements of `L`. - - If `φ x` were triangularizable, we could write `M` as a direct sum of generalized eigenspaces of - `φ x`. Because `L` is nilpotent these are all Lie submodules, thus Lie modules in their own right, - and thus by (a) above we learn that `trace (φ z) = 0` restricted to each generalized eigenspace. - Because `z` commutes with `x`, this forces `trace (φ x ∘ φ z) = 0` on each generalized eigenspace, - and so by summing the traces on each generalized eigenspace we learn the total trace is zero, as - required (see `LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero`). - - To cater for the fact that `φ x` may not be triangularizable, we first extend the scalars from `R` - to `AlgebraicClosure (FractionRing R)` and argue using the action of `A ⊗ L` on `A ⊗ M`. -/ - rintro z ⟨hz : z ∈ lowerCentralSeries R L L 1, hzc : z ∈ LieAlgebra.center R L⟩ - ext x - rw [traceForm_apply_apply, LinearMap.zero_apply] - let A := AlgebraicClosure (FractionRing R) - suffices algebraMap R A (trace R _ ((φ z).comp (φ x))) = 0 by - have _i : NoZeroSMulDivisors R A := NoZeroSMulDivisors.trans R (FractionRing R) A - rw [← map_zero (algebraMap R A)] at this - exact NoZeroSMulDivisors.algebraMap_injective R A this - rw [← LinearMap.trace_baseChange, LinearMap.baseChange_comp, ← toEndomorphism_baseChange, - ← toEndomorphism_baseChange] - replace hz : 1 ⊗ₜ z ∈ lowerCentralSeries A (A ⊗[R] L) (A ⊗[R] L) 1 := by - simp only [lowerCentralSeries_succ, lowerCentralSeries_zero] at hz ⊢ - rw [← LieSubmodule.baseChange_top, ← LieSubmodule.lie_baseChange] - exact Submodule.tmul_mem_baseChange_of_mem 1 hz - replace hzc : 1 ⊗ₜ[R] z ∈ LieAlgebra.center A (A ⊗[R] L) := by - simp only [mem_maxTrivSubmodule] at hzc ⊢ - intro y - exact y.induction_on rfl (fun a u ↦ by simp [hzc u]) (fun u v hu hv ↦ by simp [hu, hv]) - apply LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero - · exact IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) - · exact fun μ ↦ trace_toEndomorphism_eq_zero_of_mem_lcs A (A ⊗[R] L) - (weightSpaceOf (A ⊗[R] M) μ (1 ⊗ₜ x)) (le_refl 1) hz - · exact commute_toEndomorphism_of_mem_center_right (A ⊗[R] M) hzc (1 ⊗ₜ x) - -/-- A nilpotent Lie algebra with a representation whose trace form is non-singular is Abelian. -/ -lemma isLieAbelian_of_ker_traceForm_eq_bot (h : LinearMap.ker (traceForm R L M) = ⊥) : - IsLieAbelian L := by - simpa only [← disjoint_lowerCentralSeries_maxTrivSubmodule_iff R L L, disjoint_iff_inf_le, - LieIdeal.coe_to_lieSubalgebra_to_submodule, LieSubmodule.coeSubmodule_eq_bot_iff, h] - using lowerCentralSeries_one_inf_center_le_ker_traceForm R L M - -end LieModule - -namespace LieSubmodule - -open LieModule (traceForm) - -variable {R L M} -variable [IsDomain R] [IsPrincipalIdealRing R] - (N : LieSubmodule R L M) (I : LieIdeal R L) (h : I ≤ N.idealizer) (x : L) {y : L} (hy : y ∈ I) - -lemma trace_eq_trace_restrict_of_le_idealizer - (hy' : ∀ m ∈ N, (φ x ∘ₗ φ y) m ∈ N := fun m _ ↦ N.lie_mem (N.mem_idealizer.mp (h hy) m)) : - trace R M (φ x ∘ₗ φ y) = trace R N ((φ x ∘ₗ φ y).restrict hy') := by - suffices ∀ m, ⁅x, ⁅y, m⁆⁆ ∈ N by simp [(φ x ∘ₗ φ y).trace_restrict_eq_of_forall_mem _ this] - exact fun m ↦ N.lie_mem (h hy m) - -lemma traceForm_eq_of_le_idealizer : - traceForm R I N = (traceForm R L M).restrict I := by - ext ⟨x, hx⟩ ⟨y, hy⟩ - change _ = trace R M (φ x ∘ₗ φ y) - rw [N.trace_eq_trace_restrict_of_le_idealizer I h x hy] - rfl - -/-- Note that this result is slightly stronger than it might look at first glance: we only assume -that `N` is trivial over `I` rather than all of `L`. This means that it applies in the important -case of an Abelian ideal (which has `M = L` and `N = I`). -/ -lemma traceForm_eq_zero_of_isTrivial [LieModule.IsTrivial I N] : - trace R M (φ x ∘ₗ φ y) = 0 := by - let hy' : ∀ m ∈ N, (φ x ∘ₗ φ y) m ∈ N := fun m _ ↦ N.lie_mem (N.mem_idealizer.mp (h hy) m) - suffices (φ x ∘ₗ φ y).restrict hy' = 0 by - simp [this, N.trace_eq_trace_restrict_of_le_idealizer I h x hy] - ext n - suffices ⁅y, (n : M)⁆ = 0 by simp [this] - exact Submodule.coe_eq_zero.mpr (LieModule.IsTrivial.trivial (⟨y, hy⟩ : I) n) - -end LieSubmodule - -section LieAlgebra - -variable [Module.Free R L] [Module.Finite R L] - -/-- A finite, free (as an `R`-module) Lie algebra `L` carries a bilinear form on `L`. - -This is a specialisation of `LieModule.traceForm` to the adjoint representation of `L`. -/ -noncomputable abbrev killingForm : LinearMap.BilinForm R L := LieModule.traceForm R L L - -open LieAlgebra in -lemma killingForm_apply_apply (x y : L) : killingForm R L x y = trace R L (ad R L x ∘ₗ ad R L y) := - LieModule.traceForm_apply_apply R L L x y - -lemma killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting - (H : LieSubalgebra R L) [LieAlgebra.IsNilpotent R H] - {x₀ x₁ : L} - (hx₀ : x₀ ∈ LieAlgebra.zeroRootSubalgebra R L H) - (hx₁ : x₁ ∈ LieModule.posFittingComp R H L) : - killingForm R L x₀ x₁ = 0 := - LieModule.eq_zero_of_mem_weightSpace_mem_posFitting R H L - (fun x y z ↦ LieModule.traceForm_apply_lie_apply' R L L x y z) hx₀ hx₁ - -namespace LieIdeal - -variable (I : LieIdeal R L) - -/-- The orthogonal complement of an ideal with respect to the killing form is an ideal. -/ -noncomputable def killingCompl : LieIdeal R L := - { __ := (killingForm R L).orthogonal I.toSubmodule - lie_mem := by - intro x y hy - simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, - Submodule.mem_toAddSubmonoid, LinearMap.BilinForm.mem_orthogonal_iff, - LieSubmodule.mem_coeSubmodule, LinearMap.BilinForm.IsOrtho] - intro z hz - rw [← LieModule.traceForm_apply_lie_apply] - exact hy _ <| lie_mem_left _ _ _ _ _ hz } - -@[simp] lemma toSubmodule_killingCompl : - LieSubmodule.toSubmodule I.killingCompl = (killingForm R L).orthogonal I.toSubmodule := - rfl - -@[simp] lemma mem_killingCompl {x : L} : - x ∈ I.killingCompl ↔ ∀ y ∈ I, killingForm R L y x = 0 := by - rfl - -lemma coe_killingCompl_top : - killingCompl R L ⊤ = LinearMap.ker (killingForm R L) := by - ext x - simp [LinearMap.ext_iff, LinearMap.BilinForm.IsOrtho, LieModule.traceForm_comm R L L x] -variable [IsDomain R] [IsPrincipalIdealRing R] - -lemma killingForm_eq : - killingForm R I = (killingForm R L).restrict I := - LieSubmodule.traceForm_eq_of_le_idealizer I I <| by simp - -lemma restrict_killingForm : - (killingForm R L).restrict I = LieModule.traceForm R I L := - rfl - -@[simp] lemma le_killingCompl_top_of_isLieAbelian [IsLieAbelian I] : - I ≤ LieIdeal.killingCompl R L ⊤ := by - intro x (hx : x ∈ I) - simp only [mem_killingCompl, LieSubmodule.mem_top, forall_true_left] - intro y - rw [LieModule.traceForm_apply_apply] - exact LieSubmodule.traceForm_eq_zero_of_isTrivial I I (by simp) _ hx +-/ -end LieIdeal +variable (R L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] namespace LieAlgebra @@ -426,7 +50,7 @@ attribute [simp] IsKilling.killingCompl_top_eq_bot namespace IsKilling -variable [IsKilling R L] +variable [Module.Free R L] [Module.Finite R L] [IsKilling R L] @[simp] lemma ker_killingForm_eq_bot : LinearMap.ker (killingForm R L) = ⊥ := by @@ -436,299 +60,13 @@ lemma killingForm_nondegenerate : (killingForm R L).Nondegenerate := by simp [LinearMap.BilinForm.nondegenerate_iff_ker_eq_bot] -/-- If the Killing form of a Lie algebra is non-singular, it remains non-singular when restricted -to a Cartan subalgebra. -/ -lemma ker_restrict_eq_bot_of_isCartanSubalgebra - [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : - LinearMap.ker ((killingForm R L).restrict H) = ⊥ := by - have h : Codisjoint (rootSpace H 0) (LieModule.posFittingComp R H L) := - (LieModule.isCompl_weightSpace_zero_posFittingComp R H L).codisjoint - replace h : Codisjoint (H : Submodule R L) (LieModule.posFittingComp R H L : Submodule R L) := by - rwa [codisjoint_iff, ← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.sup_coe_toSubmodule, - LieSubmodule.top_coeSubmodule, rootSpace_zero_eq R L H, LieSubalgebra.coe_toLieSubmodule, - ← codisjoint_iff] at h - suffices this : ∀ m₀ ∈ H, ∀ m₁ ∈ LieModule.posFittingComp R H L, killingForm R L m₀ m₁ = 0 by - simp [LinearMap.BilinForm.ker_restrict_eq_of_codisjoint h this] - intro m₀ h₀ m₁ h₁ - exact killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting R L H (le_zeroRootSubalgebra R L H h₀) h₁ - -lemma restrict_killingForm (H : LieSubalgebra R L) : - (killingForm R L).restrict H = LieModule.traceForm R H L := - rfl - -@[simp] lemma ker_traceForm_eq_bot_of_isCartanSubalgebra - [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : - LinearMap.ker (LieModule.traceForm R H L) = ⊥ := - ker_restrict_eq_bot_of_isCartanSubalgebra R L H - -lemma traceForm_cartan_nondegenerate - [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : - (LieModule.traceForm R H L).Nondegenerate := by - simp [LinearMap.BilinForm.nondegenerate_iff_ker_eq_bot] - /-- The converse of this is true over a field of characteristic zero. There are counterexamples over fields with positive characteristic. -/ -instance isSemisimple [IsDomain R] [IsPrincipalIdealRing R] : IsSemisimple R L := by +instance instIsSemisimple [IsDomain R] [IsPrincipalIdealRing R] : IsSemisimple R L := by refine' (isSemisimple_iff_no_abelian_ideals R L).mpr fun I hI ↦ _ rw [eq_bot_iff, ← killingCompl_top_eq_bot] exact I.le_killingCompl_top_of_isLieAbelian --- TODO: formalize a positive-characteristic counterexample to the above instance - -instance instIsLieAbelianOfIsCartanSubalgebra - [IsDomain R] [IsPrincipalIdealRing R] [IsArtinian R L] - (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : - IsLieAbelian H := - LieModule.isLieAbelian_of_ker_traceForm_eq_bot R H L <| - ker_restrict_eq_bot_of_isCartanSubalgebra R L H - -end IsKilling - -end LieAlgebra - -end LieAlgebra - -section Field - -open LieModule FiniteDimensional -open Submodule (span subset_span) - -namespace LieModule - -variable [LieAlgebra.IsNilpotent K L] [LinearWeights K L M] [IsTriangularizable K L M] - -lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : - traceForm K L M x y = ∑ χ : Weight K L M, finrank K (weightSpace M χ) • (χ x * χ y) := by - have hxy : ∀ χ : Weight K L M, MapsTo (toEndomorphism K L M x ∘ₗ toEndomorphism K L M y) - (weightSpace M χ) (weightSpace M χ) := - fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm - classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpace' K L M) - (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_weightSpace_eq_top' K L M) - simp_rw [traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict hds hxy, - ← traceForm_weightSpace_eq K L M _ x y] - rfl - -lemma traceForm_eq_sum_finrank_nsmul : - traceForm K L M = ∑ χ : Weight K L M, finrank K (weightSpace M χ) • - (χ : L →ₗ[K] K).smulRight (χ : L →ₗ[K] K) := by - ext - rw [traceForm_eq_sum_finrank_nsmul_mul, ← Finset.sum_attach] - simp - --- The reverse inclusion should also hold: TODO prove this! -lemma range_traceForm_le_span_weight : - LinearMap.range (traceForm K L M) ≤ span K (range (Weight.toLinear K L M)) := by - rintro - ⟨x, rfl⟩ - 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)] - exact Submodule.smul_mem _ _ <| Submodule.smul_mem _ _ <| subset_span <| mem_range_self χ - -end LieModule - -namespace LieAlgebra - -variable [FiniteDimensional K L] - (H : LieSubalgebra K L) [H.IsCartanSubalgebra] [IsTriangularizable K H L] - -/-- For any `α` and `β`, the corresponding root spaces are orthogonal with respect to the Killing -form, provided `α + β ≠ 0`. -/ -lemma killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero {α β : H → K} {x y : L} - (hx : x ∈ rootSpace H α) (hy : y ∈ rootSpace H β) (hαβ : α + β ≠ 0) : - killingForm K L x y = 0 := by - /- If `ad R L z` is semisimple for all `z ∈ H` then writing `⟪x, y⟫ = killingForm K L x y`, there - is a slick proof of this lemma that requires only invariance of the Killing form as follows. - For any `z ∈ H`, we have: - `α z • ⟪x, y⟫ = ⟪α z • x, y⟫ = ⟪⁅z, x⁆, y⟫ = - ⟪x, ⁅z, y⁆⟫ = - ⟪x, β z • y⟫ = - β z • ⟪x, y⟫`. - Since this is true for any `z`, we thus have: `(α + β) • ⟪x, y⟫ = 0`, and hence the result. - However the semisimplicity of `ad R L z` is (a) non-trivial and (b) requires the assumption - that `K` is characteristic 0 (possibly perfect field would suffice) and `L` is semisimple. -/ - let σ : (H → K) → (H → K) := fun γ ↦ α + (β + γ) - have hσ : ∀ γ, σ γ ≠ γ := fun γ ↦ by simpa only [σ, ← add_assoc] using add_left_ne_self.mpr hαβ - let f : Module.End K L := (ad K L x) ∘ₗ (ad K L y) - have hf : ∀ γ, MapsTo f (rootSpace H γ) (rootSpace H (σ γ)) := fun γ ↦ - (mapsTo_toEndomorphism_weightSpace_add_of_mem_rootSpace K L H L α (β + γ) hx).comp <| - mapsTo_toEndomorphism_weightSpace_add_of_mem_rootSpace K L H L β γ hy - classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpace K H L) - (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_weightSpace_eq_top K H L) - exact LinearMap.trace_eq_zero_of_mapsTo_ne hds σ hσ hf - -/-- Elements of the `α` root space which are Killing-orthogonal to the `-α` root space are -Killing-orthogonal to all of `L`. -/ -lemma mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg - {α : H → K} {x : L} (hx : x ∈ rootSpace H α) - (hx' : ∀ y ∈ rootSpace H (-α), killingForm K L x y = 0) : - x ∈ LinearMap.ker (killingForm K L) := by - rw [LinearMap.mem_ker] - ext y - have hy : y ∈ ⨆ β, rootSpace H β := by simp [iSup_weightSpace_eq_top K H L] - induction hy using LieSubmodule.iSup_induction' - · next β y hy => - by_cases hαβ : α + β = 0 - · exact hx' _ (add_eq_zero_iff_neg_eq.mp hαβ ▸ hy) - · exact killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero K L H hx hy hαβ - · simp - · simp_all - -namespace IsKilling - -variable [IsKilling K L] - -variable {K L} in -/-- The restriction of the Killing form to a Cartan subalgebra, as a linear equivalence to the -dual. -/ -@[simps!] -noncomputable def cartanEquivDual : - H ≃ₗ[K] Module.Dual K H := - (traceForm K H L).toDual <| traceForm_cartan_nondegenerate K L H - -/-- This is Proposition 4.18 from [carter2005] except that we use -`LieModule.exists_forall_lie_eq_smul` instead of Lie's theorem (and so avoid -assuming `K` has characteristic zero). -/ -lemma cartanEquivDual_symm_apply_mem_corootSpace (α : Weight K H L) : - (cartanEquivDual H).symm α ∈ corootSpace α := by - obtain ⟨e : L, he₀ : e ≠ 0, he : ∀ x, ⁅x, e⁆ = α x • e⟩ := exists_forall_lie_eq_smul K H L α - have heα : e ∈ rootSpace H α := (mem_weightSpace L α e).mpr fun x ↦ ⟨1, by simp [← he x]⟩ - obtain ⟨f, hfα, hf⟩ : ∃ f ∈ rootSpace H (-α), killingForm K L e f ≠ 0 := by - contrapose! he₀ - simpa using mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg K L H heα he₀ - suffices ⁅e, f⁆ = killingForm K L e f • ((cartanEquivDual H).symm α : L) from - (mem_corootSpace α).mpr <| Submodule.subset_span ⟨(killingForm K L e f)⁻¹ • e, - Submodule.smul_mem _ _ heα, f, hfα, by simpa [inv_smul_eq_iff₀ hf]⟩ - set α' := (cartanEquivDual H).symm α - rw [← sub_eq_zero, ← Submodule.mem_bot (R := K), ← ker_killingForm_eq_bot] - apply mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg (α := (0 : H → K)) - · simp only [rootSpace_zero_eq, LieSubalgebra.mem_toLieSubmodule] - refine sub_mem ?_ (H.smul_mem _ α'.property) - simpa using mapsTo_toEndomorphism_weightSpace_add_of_mem_rootSpace K L H L α (-α) heα hfα - · intro z hz - replace hz : z ∈ H := by simpa using hz - replace he : ⁅z, e⁆ = α ⟨z, hz⟩ • e := by simpa using he ⟨z, hz⟩ - have hαz : killingForm K L α' (⟨z, hz⟩ : H) = α ⟨z, hz⟩ := - LinearMap.BilinForm.apply_toDual_symm_apply (hB := traceForm_cartan_nondegenerate K L H) _ _ - simp [traceForm_comm K L L ⁅e, f⁆, ← traceForm_apply_lie_apply, he, mul_comm _ (α ⟨z, hz⟩), hαz] - -/-- Given a splitting Cartan subalgebra `H` of a finite-dimensional Lie algebra with non-singular -Killing form, the corresponding roots span the dual space of `H`. -/ -@[simp] -lemma span_weight_eq_top : - span K (range (Weight.toLinear K H L)) = ⊤ := by - refine eq_top_iff.mpr (le_trans ?_ (LieModule.range_traceForm_le_span_weight K H L)) - rw [← traceForm_flip K H L, ← LinearMap.dualAnnihilator_ker_eq_range_flip, - ker_traceForm_eq_bot_of_isCartanSubalgebra, Submodule.dualAnnihilator_bot] - -@[simp] -lemma iInf_ker_weight_eq_bot : - ⨅ α : Weight K H L, α.ker = ⊥ := by - rw [← Subspace.dualAnnihilator_inj, Subspace.dualAnnihilator_iInf_eq, - Submodule.dualAnnihilator_bot] - simp [← LinearMap.range_dualMap_eq_dualAnnihilator_ker, ← Submodule.span_range_eq_iSup] - -@[simp] -lemma corootSpace_zero_eq_bot : - corootSpace (0 : H → K) = ⊥ := by - refine eq_bot_iff.mpr fun x hx ↦ ?_ - suffices {x | ∃ y ∈ H, ∃ z ∈ H, ⁅y, z⁆ = x} = {0} by simpa [mem_corootSpace, this] using hx - refine eq_singleton_iff_unique_mem.mpr ⟨⟨0, H.zero_mem, 0, H.zero_mem, zero_lie 0⟩, ?_⟩ - rintro - ⟨y, hy, z, hz, rfl⟩ - suffices ⁅(⟨y, hy⟩ : H), (⟨z, hz⟩ : H)⁆ = 0 by - simpa only [Subtype.ext_iff, LieSubalgebra.coe_bracket, ZeroMemClass.coe_zero] using this - simp - -section CharZero - -variable {K H L} -variable [CharZero K] - -/-- The contrapositive of this result is very useful, taking `x` to be the element of `H` -corresponding to a root `α` under the identification between `H` and `H^*` provided by the Killing -form. -/ -lemma eq_zero_of_apply_eq_zero_of_mem_corootSpace - (x : H) (α : H → K) (hαx : α x = 0) (hx : x ∈ corootSpace α) : - x = 0 := by - rcases eq_or_ne α 0 with rfl | hα; · simpa using hx - replace hx : x ∈ ⨅ β : Weight K H L, β.ker := by - refine (Submodule.mem_iInf _).mpr fun β ↦ ?_ - obtain ⟨a, b, hb, hab⟩ := - exists_forall_mem_corootSpace_smul_add_eq_zero L α β hα β.weightSpace_ne_bot - simpa [hαx, hb.ne'] using hab _ hx - simpa using hx - -lemma disjoint_ker_weight_corootSpace (α : Weight K H L) : - Disjoint α.ker (corootSpace α) := by - rw [disjoint_iff] - refine (Submodule.eq_bot_iff _).mpr fun x ⟨hαx, hx⟩ ↦ ?_ - replace hαx : α x = 0 := by simpa using hαx - exact eq_zero_of_apply_eq_zero_of_mem_corootSpace x α hαx hx - -/-- The coroot corresponding to a root. -/ -noncomputable def coroot (α : Weight K H L) : H := - 2 • (α <| (cartanEquivDual H).symm α)⁻¹ • (cartanEquivDual H).symm α - -lemma root_apply_cartanEquivDual_symm_ne_zero {α : Weight K H L} (hα : α.IsNonZero) : - α ((cartanEquivDual H).symm α) ≠ 0 := by - contrapose! hα - suffices (cartanEquivDual H).symm α ∈ α.ker ⊓ corootSpace α by - rw [(disjoint_ker_weight_corootSpace α).eq_bot] at this - simpa using this - exact Submodule.mem_inf.mp ⟨hα, cartanEquivDual_symm_apply_mem_corootSpace K L H α⟩ - -lemma root_apply_coroot {α : Weight K H L} (hα : α.IsNonZero) : - α (coroot α) = 2 := by - rw [← Weight.coe_coe] - simpa [coroot] using inv_mul_cancel (root_apply_cartanEquivDual_symm_ne_zero hα) - -@[simp] lemma coroot_eq_zero_iff {α : Weight K H L} : - coroot α = 0 ↔ α.IsZero := by - refine ⟨fun hα ↦ ?_, fun hα ↦ ?_⟩ - · by_contra contra - simpa [hα, ← α.coe_coe, map_zero] using root_apply_coroot contra - · simp [coroot, Weight.coe_toLinear_eq_zero_iff.mpr hα] - -lemma coe_corootSpace_eq_span_singleton (α : Weight K H L) : - LieSubmodule.toSubmodule (corootSpace α) = K ∙ coroot α := by - if hα : α.IsZero then - simp [hα.eq, coroot_eq_zero_iff.mpr hα] - else - set α' := (cartanEquivDual H).symm α - have hα' : (cartanEquivDual H).symm α ≠ 0 := by simpa using hα - have h₁ : (K ∙ coroot α) = K ∙ α' := by - 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 _ - have h₂ : (K ∙ (cartanEquivDual H).symm α : Submodule K H) ≤ corootSpace α := by - simpa using cartanEquivDual_symm_apply_mem_corootSpace K L H α - suffices finrank K (LieSubmodule.toSubmodule (corootSpace α)) ≤ 1 by - rw [← finrank_span_singleton (K := K) hα'] at this - exact h₁ ▸ (eq_of_le_of_finrank_le h₂ this).symm - have h : finrank K H = finrank K α.ker + 1 := - (Module.Dual.finrank_ker_add_one_of_ne_zero <| Weight.coe_toLinear_ne_zero_iff.mpr hα).symm - simpa [h] using Submodule.finrank_add_finrank_le_of_disjoint (disjoint_ker_weight_corootSpace α) - -@[simp] -lemma corootSpace_eq_bot_iff {α : Weight K H L} : - corootSpace α = ⊥ ↔ α.IsZero := by - simp [← LieSubmodule.coeSubmodule_eq_bot_iff, coe_corootSpace_eq_span_singleton α] - -lemma isCompl_ker_weight_span_coroot (α : Weight K H L) : - IsCompl α.ker (K ∙ coroot α) := by - if hα : α.IsZero then - simpa [Weight.coe_toLinear_eq_zero_iff.mpr hα, coroot_eq_zero_iff.mpr hα, Weight.ker] - using isCompl_top_bot - else - rw [← coe_corootSpace_eq_span_singleton] - apply Module.Dual.isCompl_ker_of_disjoint_of_ne_bot (by aesop) - (disjoint_ker_weight_corootSpace α) - replace hα : corootSpace α ≠ ⊥ := by simpa using hα - rwa [ne_eq, ← LieSubmodule.coe_toSubmodule_eq_iff] at hα - -end CharZero - end IsKilling section LieEquiv @@ -762,5 +100,3 @@ alias _root_.LieEquiv.isKilling := LieAlgebra.isKilling_of_equiv end LieEquiv end LieAlgebra - -end Field diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean new file mode 100644 index 0000000000000..955d2858eba0a --- /dev/null +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -0,0 +1,428 @@ +/- +Copyright (c) 2023 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.LinearAlgebra.PID +import Mathlib.Algebra.DirectSum.LinearMap +import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure +import Mathlib.LinearAlgebra.BilinearForm.Orthogonal +import Mathlib.Algebra.Lie.Weights.Cartan +import Mathlib.Algebra.Lie.Weights.Linear + +/-! +# The trace and Killing forms of a Lie algebra. + +Let `L` be a Lie algebra with coefficients in a commutative ring `R`. Suppose `M` is a finite, free +`R`-module and we have a representation `φ : L → End M`. This data induces a natural bilinear form +`B` on `L`, called the trace form associated to `M`; it is defined as `B(x, y) = Tr (φ x) (φ y)`. + +In the special case that `M` is `L` itself and `φ` is the adjoint representation, the trace form +is known as the Killing form. + +We define the trace / Killing form in this file and prove some basic properties. + +## Main definitions + + * `LieModule.traceForm`: a finite, free representation of a Lie algebra `L` induces a bilinear form + on `L` called the trace Form. + * `LieModule.traceForm_eq_zero_of_isNilpotent`: the trace form induced by a nilpotent + representation of a Lie algebra vanishes. + * `killingForm`: the adjoint representation of a (finite, free) Lie algebra `L` induces a bilinear + form on `L` via the trace form construction. +-/ + +variable (R K L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] + [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] + [Module.Free R M] [Module.Finite R M] + +local notation "φ" => LieModule.toEndomorphism R L M + +open LinearMap (trace) +open Set BigOperators FiniteDimensional + +namespace LieModule + +/-- A finite, free representation of a Lie algebra `L` induces a bilinear form on `L` called +the trace Form. See also `killingForm`. -/ +noncomputable def traceForm : LinearMap.BilinForm R L := + ((LinearMap.mul _ _).compl₁₂ (φ).toLinearMap (φ).toLinearMap).compr₂ (trace R M) + +lemma traceForm_apply_apply (x y : L) : + traceForm R L M x y = trace R _ (φ x ∘ₗ φ y) := + rfl + +lemma traceForm_comm (x y : L) : traceForm R L M x y = traceForm R L M y x := + LinearMap.trace_mul_comm R (φ x) (φ y) + +lemma traceForm_isSymm : LinearMap.IsSymm (traceForm R L M) := LieModule.traceForm_comm R L M + +@[simp] lemma traceForm_flip : LinearMap.flip (traceForm R L M) = traceForm R L M := + Eq.symm <| LinearMap.ext₂ <| traceForm_comm R L M + +/-- The trace form of a Lie module is compatible with the action of the Lie algebra. + +See also `LieModule.traceForm_apply_lie_apply'`. -/ +lemma traceForm_apply_lie_apply (x y z : L) : + traceForm R L M ⁅x, y⁆ z = traceForm R L M x ⁅y, z⁆ := by + calc traceForm R L M ⁅x, y⁆ z + = trace R _ (φ ⁅x, y⁆ ∘ₗ φ z) := by simp only [traceForm_apply_apply] + _ = trace R _ ((φ x * φ y - φ y * φ x) * φ z) := ?_ + _ = trace R _ (φ x * (φ y * φ z)) - trace R _ (φ y * (φ x * φ z)) := ?_ + _ = trace R _ (φ x * (φ y * φ z)) - trace R _ (φ x * (φ z * φ y)) := ?_ + _ = traceForm R L M x ⁅y, z⁆ := ?_ + · simp only [LieHom.map_lie, Ring.lie_def, ← LinearMap.mul_eq_comp] + · simp only [sub_mul, mul_sub, map_sub, mul_assoc] + · simp only [LinearMap.trace_mul_cycle' R (φ x) (φ z) (φ y)] + · simp only [traceForm_apply_apply, LieHom.map_lie, Ring.lie_def, mul_sub, map_sub, + ← LinearMap.mul_eq_comp] + +/-- Given a representation `M` of a Lie algebra `L`, the action of any `x : L` is skew-adjoint wrt +the trace form. -/ +lemma traceForm_apply_lie_apply' (x y z : L) : + traceForm R L M ⁅x, y⁆ z = - traceForm R L M y ⁅x, z⁆ := + calc traceForm R L M ⁅x, y⁆ z + = - traceForm R L M ⁅y, x⁆ z := by rw [← lie_skew x y, map_neg, LinearMap.neg_apply] + _ = - traceForm R L M y ⁅x, z⁆ := by rw [traceForm_apply_lie_apply] + +/-- This lemma justifies the terminology "invariant" for trace forms. -/ +@[simp] lemma lie_traceForm_eq_zero (x : L) : ⁅x, traceForm R L M⁆ = 0 := by + ext y z + rw [LieHom.lie_apply, LinearMap.sub_apply, Module.Dual.lie_apply, LinearMap.zero_apply, + LinearMap.zero_apply, traceForm_apply_lie_apply', sub_self] + +@[simp] lemma traceForm_eq_zero_of_isNilpotent [IsReduced R] [IsNilpotent R L M] : + traceForm R L M = 0 := by + ext x y + simp only [traceForm_apply_apply, LinearMap.zero_apply, ← isNilpotent_iff_eq_zero] + apply LinearMap.isNilpotent_trace_of_isNilpotent + exact isNilpotent_toEndomorphism_of_isNilpotent₂ R L M x y + +@[simp] +lemma traceForm_weightSpace_eq [IsDomain R] [IsPrincipalIdealRing R] + [LieAlgebra.IsNilpotent R L] [IsNoetherian R M] [LinearWeights R L M] (χ : L → R) (x y : L) : + traceForm R L (weightSpace M χ) x y = finrank R (weightSpace M χ) • (χ x * χ y) := by + set d := finrank R (weightSpace M χ) + have h₁ : χ y • d • χ x - χ y • χ x • (d : R) = 0 := by simp [mul_comm (χ x)] + have h₂ : χ x • d • χ y = d • (χ x * χ y) := by + simpa [nsmul_eq_mul, smul_eq_mul] using mul_left_comm (χ x) d (χ y) + have := traceForm_eq_zero_of_isNilpotent R L (shiftedWeightSpace R L M χ) + replace this := LinearMap.congr_fun (LinearMap.congr_fun this x) y + rwa [LinearMap.zero_apply, LinearMap.zero_apply, traceForm_apply_apply, + shiftedWeightSpace.toEndomorphism_eq, shiftedWeightSpace.toEndomorphism_eq, + ← LinearEquiv.conj_comp, LinearMap.trace_conj', LinearMap.comp_sub, LinearMap.sub_comp, + LinearMap.sub_comp, map_sub, map_sub, map_sub, LinearMap.comp_smul, LinearMap.smul_comp, + LinearMap.comp_id, LinearMap.id_comp, LinearMap.map_smul, LinearMap.map_smul, + trace_toEndomorphism_weightSpace, trace_toEndomorphism_weightSpace, + LinearMap.comp_smul, LinearMap.smul_comp, LinearMap.id_comp, map_smul, map_smul, + LinearMap.trace_id, ← traceForm_apply_apply, h₁, h₂, sub_zero, sub_eq_zero] at this + +/-- The upper and lower central series of `L` are orthogonal wrt the trace form of any Lie module +`M`. -/ +lemma traceForm_eq_zero_if_mem_lcs_of_mem_ucs {x y : L} (k : ℕ) + (hx : x ∈ (⊤ : LieIdeal R L).lcs L k) (hy : y ∈ (⊥ : LieIdeal R L).ucs k) : + traceForm R L M x y = 0 := by + induction' k with k ih generalizing x y + · replace hy : y = 0 := by simpa using hy + simp [hy] + · rw [LieSubmodule.ucs_succ, LieSubmodule.mem_normalizer] at hy + simp_rw [LieIdeal.lcs_succ, ← LieSubmodule.mem_coeSubmodule, + LieSubmodule.lieIdeal_oper_eq_linear_span', LieSubmodule.mem_top, true_and] at hx + refine Submodule.span_induction hx ?_ ?_ (fun z w hz hw ↦ ?_) (fun t z hz ↦ ?_) + · rintro - ⟨z, w, hw, rfl⟩ + rw [← lie_skew, map_neg, LinearMap.neg_apply, neg_eq_zero, traceForm_apply_lie_apply] + exact ih hw (hy _) + · simp + · simp [hz, hw] + · simp [hz] + +lemma traceForm_apply_eq_zero_of_mem_lcs_of_mem_center {x y : L} + (hx : x ∈ lowerCentralSeries R L L 1) (hy : y ∈ LieAlgebra.center R L) : + traceForm R L M x y = 0 := by + apply traceForm_eq_zero_if_mem_lcs_of_mem_ucs R L M 1 + · simpa using hx + · simpa using hy + +-- This is barely worth having: it usually follows from `LieModule.traceForm_eq_zero_of_isNilpotent` +@[simp] lemma traceForm_eq_zero_of_isTrivial [IsTrivial L M] : + traceForm R L M = 0 := by + ext x y + suffices φ x ∘ₗ φ y = 0 by simp [traceForm_apply_apply, this] + ext m + simp + +/-- Given a bilinear form `B` on a representation `M` of a nilpotent Lie algebra `L`, if `B` is +invariant (in the sense that the action of `L` is skew-adjoint wrt `B`) then components of the +Fitting decomposition of `M` are orthogonal wrt `B`. -/ +lemma eq_zero_of_mem_weightSpace_mem_posFitting [LieAlgebra.IsNilpotent R L] + {B : LinearMap.BilinForm R M} (hB : ∀ (x : L) (m n : M), B ⁅x, m⁆ n = - B m ⁅x, n⁆) + {m₀ m₁ : M} (hm₀ : m₀ ∈ weightSpace M (0 : L → R)) (hm₁ : m₁ ∈ posFittingComp R L M) : + B m₀ m₁ = 0 := by + replace hB : ∀ x (k : ℕ) m n, B m ((φ x ^ k) n) = (- 1 : R) ^ k • B ((φ x ^ k) m) n := by + intro x k + induction k with + | zero => simp + | succ k ih => + intro m n + replace hB : ∀ m, B m (φ x n) = (- 1 : R) • B (φ x m) n := by simp [hB] + have : (-1 : R) ^ k • (-1 : R) = (-1 : R) ^ (k + 1) := by rw [pow_succ (-1 : R), smul_eq_mul] + conv_lhs => rw [pow_succ, LinearMap.mul_eq_comp, LinearMap.comp_apply, ih, hB, + ← (φ x).comp_apply, ← LinearMap.mul_eq_comp, ← pow_succ', ← smul_assoc, this] + suffices ∀ (x : L) m, m ∈ posFittingCompOf R M x → B m₀ m = 0 by + apply LieSubmodule.iSup_induction _ hm₁ this (map_zero _) + aesop + clear hm₁ m₁; intro x m₁ hm₁ + simp only [mem_weightSpace, Pi.zero_apply, zero_smul, sub_zero] at hm₀ + obtain ⟨k, hk⟩ := hm₀ x + obtain ⟨m, rfl⟩ := (mem_posFittingCompOf R x m₁).mp hm₁ k + simp [hB, hk] + +lemma trace_toEndomorphism_eq_zero_of_mem_lcs + {k : ℕ} {x : L} (hk : 1 ≤ k) (hx : x ∈ lowerCentralSeries R L L k) : + trace R _ (toEndomorphism R L M x) = 0 := by + replace hx : x ∈ lowerCentralSeries R L L 1 := antitone_lowerCentralSeries _ _ _ hk hx + replace hx : x ∈ Submodule.span R {m | ∃ u v : L, ⁅u, v⁆ = m} := by + rw [lowerCentralSeries_succ, ← LieSubmodule.mem_coeSubmodule, + LieSubmodule.lieIdeal_oper_eq_linear_span'] at hx + simpa using hx + refine Submodule.span_induction (p := fun x ↦ trace R _ (toEndomorphism R L M x) = 0) hx + (fun y ⟨u, v, huv⟩ ↦ ?_) ?_ (fun u v hu hv ↦ ?_) (fun t u hu ↦ ?_) + · simp [← huv] + · simp + · simp [hu, hv] + · simp [hu] + +@[simp] +lemma traceForm_lieSubalgebra_mk_left (L' : LieSubalgebra R L) {x : L} (hx : x ∈ L') (y : L') : + traceForm R L' M ⟨x, hx⟩ y = traceForm R L M x y := + rfl + +@[simp] +lemma traceForm_lieSubalgebra_mk_right (L' : LieSubalgebra R L) {x : L'} {y : L} (hy : y ∈ L') : + traceForm R L' M x ⟨y, hy⟩ = traceForm R L M x y := + rfl + +open TensorProduct + +variable [LieAlgebra.IsNilpotent R L] [IsDomain R] [IsPrincipalIdealRing R] + +lemma traceForm_eq_sum_weightSpaceOf [IsTriangularizable R L M] (z : L) : + traceForm R L M = + ∑ χ in (finite_weightSpaceOf_ne_bot R L M z).toFinset, traceForm R L (weightSpaceOf M χ z) := by + ext x y + have hxy : ∀ χ : R, MapsTo ((toEndomorphism R L M x).comp (toEndomorphism R L M y)) + (weightSpaceOf M χ z) (weightSpaceOf M χ z) := + fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm + have hfin : {χ : R | (weightSpaceOf M χ z : Submodule R M) ≠ ⊥}.Finite := by + convert finite_weightSpaceOf_ne_bot R L M z + exact LieSubmodule.coeSubmodule_eq_bot_iff (weightSpaceOf M _ _) + classical + have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top + (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpaceOf R L M z) + (IsTriangularizable.iSup_eq_top z) + simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, + LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] + exact Finset.sum_congr (by simp) (fun χ _ ↦ rfl) + +-- In characteristic zero (or even just `LinearWeights R L M`) a stronger result holds (no +-- `⊓ LieAlgebra.center R L`) TODO prove this using `LieModule.traceForm_eq_sum_finrank_nsmul_mul`. +lemma lowerCentralSeries_one_inf_center_le_ker_traceForm : + lowerCentralSeries R L L 1 ⊓ LieAlgebra.center R L ≤ LinearMap.ker (traceForm R L M) := by + /- Sketch of proof (due to Zassenhaus): + + Let `z ∈ lowerCentralSeries R L L 1 ⊓ LieAlgebra.center R L` and `x : L`. We must show that + `trace (φ x ∘ φ z) = 0` where `φ z : End R M` indicates the action of `z` on `M` (and likewise + for `φ x`). + + Because `z` belongs to the indicated intersection, it has two key properties: + (a) the trace of the action of `z` vanishes on any Lie module of `L` + (see `LieModule.trace_toEndomorphism_eq_zero_of_mem_lcs`), + (b) `z` commutes with all elements of `L`. + + If `φ x` were triangularizable, we could write `M` as a direct sum of generalized eigenspaces of + `φ x`. Because `L` is nilpotent these are all Lie submodules, thus Lie modules in their own right, + and thus by (a) above we learn that `trace (φ z) = 0` restricted to each generalized eigenspace. + Because `z` commutes with `x`, this forces `trace (φ x ∘ φ z) = 0` on each generalized eigenspace, + and so by summing the traces on each generalized eigenspace we learn the total trace is zero, as + required (see `LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero`). + + To cater for the fact that `φ x` may not be triangularizable, we first extend the scalars from `R` + to `AlgebraicClosure (FractionRing R)` and argue using the action of `A ⊗ L` on `A ⊗ M`. -/ + rintro z ⟨hz : z ∈ lowerCentralSeries R L L 1, hzc : z ∈ LieAlgebra.center R L⟩ + ext x + rw [traceForm_apply_apply, LinearMap.zero_apply] + let A := AlgebraicClosure (FractionRing R) + suffices algebraMap R A (trace R _ ((φ z).comp (φ x))) = 0 by + have _i : NoZeroSMulDivisors R A := NoZeroSMulDivisors.trans R (FractionRing R) A + rw [← map_zero (algebraMap R A)] at this + exact NoZeroSMulDivisors.algebraMap_injective R A this + rw [← LinearMap.trace_baseChange, LinearMap.baseChange_comp, ← toEndomorphism_baseChange, + ← toEndomorphism_baseChange] + replace hz : 1 ⊗ₜ z ∈ lowerCentralSeries A (A ⊗[R] L) (A ⊗[R] L) 1 := by + simp only [lowerCentralSeries_succ, lowerCentralSeries_zero] at hz ⊢ + rw [← LieSubmodule.baseChange_top, ← LieSubmodule.lie_baseChange] + exact Submodule.tmul_mem_baseChange_of_mem 1 hz + replace hzc : 1 ⊗ₜ[R] z ∈ LieAlgebra.center A (A ⊗[R] L) := by + simp only [mem_maxTrivSubmodule] at hzc ⊢ + intro y + exact y.induction_on rfl (fun a u ↦ by simp [hzc u]) (fun u v hu hv ↦ by simp [hu, hv]) + apply LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero + · exact IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) + · exact fun μ ↦ trace_toEndomorphism_eq_zero_of_mem_lcs A (A ⊗[R] L) + (weightSpaceOf (A ⊗[R] M) μ (1 ⊗ₜ x)) (le_refl 1) hz + · exact commute_toEndomorphism_of_mem_center_right (A ⊗[R] M) hzc (1 ⊗ₜ x) + +/-- A nilpotent Lie algebra with a representation whose trace form is non-singular is Abelian. -/ +lemma isLieAbelian_of_ker_traceForm_eq_bot (h : LinearMap.ker (traceForm R L M) = ⊥) : + IsLieAbelian L := by + simpa only [← disjoint_lowerCentralSeries_maxTrivSubmodule_iff R L L, disjoint_iff_inf_le, + LieIdeal.coe_to_lieSubalgebra_to_submodule, LieSubmodule.coeSubmodule_eq_bot_iff, h] + using lowerCentralSeries_one_inf_center_le_ker_traceForm R L M + +end LieModule + +namespace LieSubmodule + +open LieModule (traceForm) + +variable {R L M} +variable [IsDomain R] [IsPrincipalIdealRing R] + (N : LieSubmodule R L M) (I : LieIdeal R L) (h : I ≤ N.idealizer) (x : L) {y : L} (hy : y ∈ I) + +lemma trace_eq_trace_restrict_of_le_idealizer + (hy' : ∀ m ∈ N, (φ x ∘ₗ φ y) m ∈ N := fun m _ ↦ N.lie_mem (N.mem_idealizer.mp (h hy) m)) : + trace R M (φ x ∘ₗ φ y) = trace R N ((φ x ∘ₗ φ y).restrict hy') := by + suffices ∀ m, ⁅x, ⁅y, m⁆⁆ ∈ N by simp [(φ x ∘ₗ φ y).trace_restrict_eq_of_forall_mem _ this] + exact fun m ↦ N.lie_mem (h hy m) + +lemma traceForm_eq_of_le_idealizer : + traceForm R I N = (traceForm R L M).restrict I := by + ext ⟨x, hx⟩ ⟨y, hy⟩ + change _ = trace R M (φ x ∘ₗ φ y) + rw [N.trace_eq_trace_restrict_of_le_idealizer I h x hy] + rfl + +/-- Note that this result is slightly stronger than it might look at first glance: we only assume +that `N` is trivial over `I` rather than all of `L`. This means that it applies in the important +case of an Abelian ideal (which has `M = L` and `N = I`). -/ +lemma traceForm_eq_zero_of_isTrivial [LieModule.IsTrivial I N] : + trace R M (φ x ∘ₗ φ y) = 0 := by + let hy' : ∀ m ∈ N, (φ x ∘ₗ φ y) m ∈ N := fun m _ ↦ N.lie_mem (N.mem_idealizer.mp (h hy) m) + suffices (φ x ∘ₗ φ y).restrict hy' = 0 by + simp [this, N.trace_eq_trace_restrict_of_le_idealizer I h x hy] + ext n + suffices ⁅y, (n : M)⁆ = 0 by simp [this] + exact Submodule.coe_eq_zero.mpr (LieModule.IsTrivial.trivial (⟨y, hy⟩ : I) n) + +end LieSubmodule + +section LieAlgebra + +variable [Module.Free R L] [Module.Finite R L] + +/-- A finite, free (as an `R`-module) Lie algebra `L` carries a bilinear form on `L`. + +This is a specialisation of `LieModule.traceForm` to the adjoint representation of `L`. -/ +noncomputable abbrev killingForm : LinearMap.BilinForm R L := LieModule.traceForm R L L + +open LieAlgebra in +lemma killingForm_apply_apply (x y : L) : killingForm R L x y = trace R L (ad R L x ∘ₗ ad R L y) := + LieModule.traceForm_apply_apply R L L x y + +lemma killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting + (H : LieSubalgebra R L) [LieAlgebra.IsNilpotent R H] + {x₀ x₁ : L} + (hx₀ : x₀ ∈ LieAlgebra.zeroRootSubalgebra R L H) + (hx₁ : x₁ ∈ LieModule.posFittingComp R H L) : + killingForm R L x₀ x₁ = 0 := + LieModule.eq_zero_of_mem_weightSpace_mem_posFitting R H L + (fun x y z ↦ LieModule.traceForm_apply_lie_apply' R L L x y z) hx₀ hx₁ + +namespace LieIdeal + +variable (I : LieIdeal R L) + +/-- The orthogonal complement of an ideal with respect to the killing form is an ideal. -/ +noncomputable def killingCompl : LieIdeal R L := + { __ := (killingForm R L).orthogonal I.toSubmodule + lie_mem := by + intro x y hy + simp only [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, + Submodule.mem_toAddSubmonoid, LinearMap.BilinForm.mem_orthogonal_iff, + LieSubmodule.mem_coeSubmodule, LinearMap.BilinForm.IsOrtho] + intro z hz + rw [← LieModule.traceForm_apply_lie_apply] + exact hy _ <| lie_mem_left _ _ _ _ _ hz } + +@[simp] lemma toSubmodule_killingCompl : + LieSubmodule.toSubmodule I.killingCompl = (killingForm R L).orthogonal I.toSubmodule := + rfl + +@[simp] lemma mem_killingCompl {x : L} : + x ∈ I.killingCompl ↔ ∀ y ∈ I, killingForm R L y x = 0 := by + rfl + +lemma coe_killingCompl_top : + killingCompl R L ⊤ = LinearMap.ker (killingForm R L) := by + ext x + simp [LinearMap.ext_iff, LinearMap.BilinForm.IsOrtho, LieModule.traceForm_comm R L L x] + +variable [IsDomain R] [IsPrincipalIdealRing R] + +lemma killingForm_eq : + killingForm R I = (killingForm R L).restrict I := + LieSubmodule.traceForm_eq_of_le_idealizer I I <| by simp + +lemma restrict_killingForm : + (killingForm R L).restrict I = LieModule.traceForm R I L := + rfl + +@[simp] lemma le_killingCompl_top_of_isLieAbelian [IsLieAbelian I] : + I ≤ LieIdeal.killingCompl R L ⊤ := by + intro x (hx : x ∈ I) + simp only [mem_killingCompl, LieSubmodule.mem_top, forall_true_left] + intro y + rw [LieModule.traceForm_apply_apply] + exact LieSubmodule.traceForm_eq_zero_of_isTrivial I I (by simp) _ hx + +end LieIdeal + +open LieModule FiniteDimensional +open Submodule (span subset_span) + +namespace LieModule + +variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [FiniteDimensional K M] +variable [LieAlgebra.IsNilpotent K L] [LinearWeights K L M] [IsTriangularizable K L M] + +lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : + traceForm K L M x y = ∑ χ : Weight K L M, finrank K (weightSpace M χ) • (χ x * χ y) := by + have hxy : ∀ χ : Weight K L M, MapsTo (toEndomorphism K L M x ∘ₗ toEndomorphism K L M y) + (weightSpace M χ) (weightSpace M χ) := + fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm + classical + have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top + (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpace' K L M) + (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_weightSpace_eq_top' K L M) + simp_rw [traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict hds hxy, + ← traceForm_weightSpace_eq K L M _ x y] + rfl + +lemma traceForm_eq_sum_finrank_nsmul : + traceForm K L M = ∑ χ : Weight K L M, finrank K (weightSpace M χ) • + (χ : L →ₗ[K] K).smulRight (χ : L →ₗ[K] K) := by + ext + rw [traceForm_eq_sum_finrank_nsmul_mul, ← Finset.sum_attach] + simp + +-- The reverse inclusion should also hold: TODO prove this! +lemma range_traceForm_le_span_weight : + LinearMap.range (traceForm K L M) ≤ span K (range (Weight.toLinear K L M)) := by + rintro - ⟨x, rfl⟩ + 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)] + exact Submodule.smul_mem _ _ <| Submodule.smul_mem _ _ <| subset_span <| mem_range_self χ + +end LieModule diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean new file mode 100644 index 0000000000000..0e4c93e2238fa --- /dev/null +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -0,0 +1,284 @@ +/- +Copyright (c) 2024 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import Mathlib.Algebra.Lie.Killing +import Mathlib.Algebra.Lie.Weights.Chain + +/-! +# Roots of Lie algebras with non-degenerate Killing forms + +The file contains definitions and results about roots of Lie algebras with non-degenerate Killing +forms. + +## Main definitions + * `LieAlgebra.IsKilling.ker_restrict_eq_bot_of_isCartanSubalgebra`: if the Killing form of + a Lie algebra is non-singular, it remains non-singular when restricted to a Cartan subalgebra. + * `LieAlgebra.IsKilling.instIsLieAbelianOfIsCartanSubalgebra`: if the Killing form of a Lie + algebra is non-singular, then its Cartan subalgebras are Abelian. + * `LieAlgebra.IsKilling.span_weight_eq_top`: given a splitting Cartan subalgebra `H` of a + finite-dimensional Lie algebra with non-singular Killing form, the corresponding roots span the + dual space of `H`. + * `LieAlgebra.IsKilling.coroot`: the coroot corresponding to a root. + * `LieAlgebra.IsKilling.isCompl_ker_weight_span_coroot`: given a root `α` with respect to a Cartan + subalgebra `H`, we have a natural decomposition of `H` as the kernel of `α` and the span of the + coroot corresponding to `α`. + +-/ + +variable (R K L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] [Field K] [LieAlgebra K L] + +namespace LieAlgebra + +namespace IsKilling + +variable [IsKilling R L] [Module.Finite R L] [Module.Free R L] + +/-- If the Killing form of a Lie algebra is non-singular, it remains non-singular when restricted +to a Cartan subalgebra. -/ +lemma ker_restrict_eq_bot_of_isCartanSubalgebra + [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : + LinearMap.ker ((killingForm R L).restrict H) = ⊥ := by + have h : Codisjoint (rootSpace H 0) (LieModule.posFittingComp R H L) := + (LieModule.isCompl_weightSpace_zero_posFittingComp R H L).codisjoint + replace h : Codisjoint (H : Submodule R L) (LieModule.posFittingComp R H L : Submodule R L) := by + rwa [codisjoint_iff, ← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.sup_coe_toSubmodule, + LieSubmodule.top_coeSubmodule, rootSpace_zero_eq R L H, LieSubalgebra.coe_toLieSubmodule, + ← codisjoint_iff] at h + suffices this : ∀ m₀ ∈ H, ∀ m₁ ∈ LieModule.posFittingComp R H L, killingForm R L m₀ m₁ = 0 by + simp [LinearMap.BilinForm.ker_restrict_eq_of_codisjoint h this] + intro m₀ h₀ m₁ h₁ + exact killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting R L H (le_zeroRootSubalgebra R L H h₀) h₁ + +lemma restrict_killingForm (H : LieSubalgebra R L) : + (killingForm R L).restrict H = LieModule.traceForm R H L := + rfl + +@[simp] lemma ker_traceForm_eq_bot_of_isCartanSubalgebra + [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : + LinearMap.ker (LieModule.traceForm R H L) = ⊥ := + ker_restrict_eq_bot_of_isCartanSubalgebra R L H + +lemma traceForm_cartan_nondegenerate + [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : + (LieModule.traceForm R H L).Nondegenerate := by + simp [LinearMap.BilinForm.nondegenerate_iff_ker_eq_bot] + +instance instIsLieAbelianOfIsCartanSubalgebra + [IsDomain R] [IsPrincipalIdealRing R] [IsArtinian R L] + (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : + IsLieAbelian H := + LieModule.isLieAbelian_of_ker_traceForm_eq_bot R H L <| + ker_restrict_eq_bot_of_isCartanSubalgebra R L H + +end IsKilling + +section Field + +open FiniteDimensional LieModule Set +open Submodule (span subset_span) + +variable [FiniteDimensional K L] + (H : LieSubalgebra K L) [H.IsCartanSubalgebra] [IsTriangularizable K H L] + +/-- For any `α` and `β`, the corresponding root spaces are orthogonal with respect to the Killing +form, provided `α + β ≠ 0`. -/ +lemma killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero {α β : H → K} {x y : L} + (hx : x ∈ rootSpace H α) (hy : y ∈ rootSpace H β) (hαβ : α + β ≠ 0) : + killingForm K L x y = 0 := by + /- If `ad R L z` is semisimple for all `z ∈ H` then writing `⟪x, y⟫ = killingForm K L x y`, there + is a slick proof of this lemma that requires only invariance of the Killing form as follows. + For any `z ∈ H`, we have: + `α z • ⟪x, y⟫ = ⟪α z • x, y⟫ = ⟪⁅z, x⁆, y⟫ = - ⟪x, ⁅z, y⁆⟫ = - ⟪x, β z • y⟫ = - β z • ⟪x, y⟫`. + Since this is true for any `z`, we thus have: `(α + β) • ⟪x, y⟫ = 0`, and hence the result. + However the semisimplicity of `ad R L z` is (a) non-trivial and (b) requires the assumption + that `K` is a perfect field and `L` has non-degenerate Killing form. -/ + let σ : (H → K) → (H → K) := fun γ ↦ α + (β + γ) + have hσ : ∀ γ, σ γ ≠ γ := fun γ ↦ by simpa only [σ, ← add_assoc] using add_left_ne_self.mpr hαβ + let f : Module.End K L := (ad K L x) ∘ₗ (ad K L y) + have hf : ∀ γ, MapsTo f (rootSpace H γ) (rootSpace H (σ γ)) := fun γ ↦ + (mapsTo_toEndomorphism_weightSpace_add_of_mem_rootSpace K L H L α (β + γ) hx).comp <| + mapsTo_toEndomorphism_weightSpace_add_of_mem_rootSpace K L H L β γ hy + classical + have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top + (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpace K H L) + (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_weightSpace_eq_top K H L) + exact LinearMap.trace_eq_zero_of_mapsTo_ne hds σ hσ hf + +/-- Elements of the `α` root space which are Killing-orthogonal to the `-α` root space are +Killing-orthogonal to all of `L`. -/ +lemma mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg + {α : H → K} {x : L} (hx : x ∈ rootSpace H α) + (hx' : ∀ y ∈ rootSpace H (-α), killingForm K L x y = 0) : + x ∈ LinearMap.ker (killingForm K L) := by + rw [LinearMap.mem_ker] + ext y + have hy : y ∈ ⨆ β, rootSpace H β := by simp [iSup_weightSpace_eq_top K H L] + induction hy using LieSubmodule.iSup_induction' + · next β y hy => + by_cases hαβ : α + β = 0 + · exact hx' _ (add_eq_zero_iff_neg_eq.mp hαβ ▸ hy) + · exact killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero K L H hx hy hαβ + · simp + · simp_all + +namespace IsKilling + +variable [IsKilling K L] + +variable {K L} in +/-- The restriction of the Killing form to a Cartan subalgebra, as a linear equivalence to the +dual. -/ +@[simps!] +noncomputable def cartanEquivDual : + H ≃ₗ[K] Module.Dual K H := + (traceForm K H L).toDual <| traceForm_cartan_nondegenerate K L H + +/-- This is Proposition 4.18 from [carter2005] except that we use +`LieModule.exists_forall_lie_eq_smul` instead of Lie's theorem (and so avoid +assuming `K` has characteristic zero). -/ +lemma cartanEquivDual_symm_apply_mem_corootSpace (α : Weight K H L) : + (cartanEquivDual H).symm α ∈ corootSpace α := by + obtain ⟨e : L, he₀ : e ≠ 0, he : ∀ x, ⁅x, e⁆ = α x • e⟩ := exists_forall_lie_eq_smul K H L α + have heα : e ∈ rootSpace H α := (mem_weightSpace L α e).mpr fun x ↦ ⟨1, by simp [← he x]⟩ + obtain ⟨f, hfα, hf⟩ : ∃ f ∈ rootSpace H (-α), killingForm K L e f ≠ 0 := by + contrapose! he₀ + simpa using mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg K L H heα he₀ + suffices ⁅e, f⁆ = killingForm K L e f • ((cartanEquivDual H).symm α : L) from + (mem_corootSpace α).mpr <| Submodule.subset_span ⟨(killingForm K L e f)⁻¹ • e, + Submodule.smul_mem _ _ heα, f, hfα, by simpa [inv_smul_eq_iff₀ hf]⟩ + set α' := (cartanEquivDual H).symm α + rw [← sub_eq_zero, ← Submodule.mem_bot (R := K), ← ker_killingForm_eq_bot] + apply mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg (α := (0 : H → K)) + · simp only [rootSpace_zero_eq, LieSubalgebra.mem_toLieSubmodule] + refine sub_mem ?_ (H.smul_mem _ α'.property) + simpa using mapsTo_toEndomorphism_weightSpace_add_of_mem_rootSpace K L H L α (-α) heα hfα + · intro z hz + replace hz : z ∈ H := by simpa using hz + replace he : ⁅z, e⁆ = α ⟨z, hz⟩ • e := by simpa using he ⟨z, hz⟩ + have hαz : killingForm K L α' (⟨z, hz⟩ : H) = α ⟨z, hz⟩ := + LinearMap.BilinForm.apply_toDual_symm_apply (hB := traceForm_cartan_nondegenerate K L H) _ _ + simp [traceForm_comm K L L ⁅e, f⁆, ← traceForm_apply_lie_apply, he, mul_comm _ (α ⟨z, hz⟩), hαz] + +/-- Given a splitting Cartan subalgebra `H` of a finite-dimensional Lie algebra with non-singular +Killing form, the corresponding roots span the dual space of `H`. -/ +@[simp] +lemma span_weight_eq_top : + span K (range (Weight.toLinear K H L)) = ⊤ := by + refine eq_top_iff.mpr (le_trans ?_ (LieModule.range_traceForm_le_span_weight K H L)) + rw [← traceForm_flip K H L, ← LinearMap.dualAnnihilator_ker_eq_range_flip, + ker_traceForm_eq_bot_of_isCartanSubalgebra, Submodule.dualAnnihilator_bot] + +@[simp] +lemma iInf_ker_weight_eq_bot : + ⨅ α : Weight K H L, α.ker = ⊥ := by + rw [← Subspace.dualAnnihilator_inj, Subspace.dualAnnihilator_iInf_eq, + Submodule.dualAnnihilator_bot] + simp [← LinearMap.range_dualMap_eq_dualAnnihilator_ker, ← Submodule.span_range_eq_iSup] + +@[simp] +lemma corootSpace_zero_eq_bot : + corootSpace (0 : H → K) = ⊥ := by + refine eq_bot_iff.mpr fun x hx ↦ ?_ + suffices {x | ∃ y ∈ H, ∃ z ∈ H, ⁅y, z⁆ = x} = {0} by simpa [mem_corootSpace, this] using hx + refine eq_singleton_iff_unique_mem.mpr ⟨⟨0, H.zero_mem, 0, H.zero_mem, zero_lie 0⟩, ?_⟩ + rintro - ⟨y, hy, z, hz, rfl⟩ + suffices ⁅(⟨y, hy⟩ : H), (⟨z, hz⟩ : H)⁆ = 0 by + simpa only [Subtype.ext_iff, LieSubalgebra.coe_bracket, ZeroMemClass.coe_zero] using this + simp + +section CharZero + +variable {K H L} +variable [CharZero K] + +/-- The contrapositive of this result is very useful, taking `x` to be the element of `H` +corresponding to a root `α` under the identification between `H` and `H^*` provided by the Killing +form. -/ +lemma eq_zero_of_apply_eq_zero_of_mem_corootSpace + (x : H) (α : H → K) (hαx : α x = 0) (hx : x ∈ corootSpace α) : + x = 0 := by + rcases eq_or_ne α 0 with rfl | hα; · simpa using hx + replace hx : x ∈ ⨅ β : Weight K H L, β.ker := by + refine (Submodule.mem_iInf _).mpr fun β ↦ ?_ + obtain ⟨a, b, hb, hab⟩ := + exists_forall_mem_corootSpace_smul_add_eq_zero L α β hα β.weightSpace_ne_bot + simpa [hαx, hb.ne'] using hab _ hx + simpa using hx + +lemma disjoint_ker_weight_corootSpace (α : Weight K H L) : + Disjoint α.ker (corootSpace α) := by + rw [disjoint_iff] + refine (Submodule.eq_bot_iff _).mpr fun x ⟨hαx, hx⟩ ↦ ?_ + replace hαx : α x = 0 := by simpa using hαx + exact eq_zero_of_apply_eq_zero_of_mem_corootSpace x α hαx hx + +/-- The coroot corresponding to a root. -/ +noncomputable def coroot (α : Weight K H L) : H := + 2 • (α <| (cartanEquivDual H).symm α)⁻¹ • (cartanEquivDual H).symm α + +lemma root_apply_cartanEquivDual_symm_ne_zero {α : Weight K H L} (hα : α.IsNonZero) : + α ((cartanEquivDual H).symm α) ≠ 0 := by + contrapose! hα + suffices (cartanEquivDual H).symm α ∈ α.ker ⊓ corootSpace α by + rw [(disjoint_ker_weight_corootSpace α).eq_bot] at this + simpa using this + exact Submodule.mem_inf.mp ⟨hα, cartanEquivDual_symm_apply_mem_corootSpace K L H α⟩ + +lemma root_apply_coroot {α : Weight K H L} (hα : α.IsNonZero) : + α (coroot α) = 2 := by + rw [← Weight.coe_coe] + simpa [coroot] using inv_mul_cancel (root_apply_cartanEquivDual_symm_ne_zero hα) + +@[simp] lemma coroot_eq_zero_iff {α : Weight K H L} : + coroot α = 0 ↔ α.IsZero := by + refine ⟨fun hα ↦ ?_, fun hα ↦ ?_⟩ + · by_contra contra + simpa [hα, ← α.coe_coe, map_zero] using root_apply_coroot contra + · simp [coroot, Weight.coe_toLinear_eq_zero_iff.mpr hα] + +lemma coe_corootSpace_eq_span_singleton (α : Weight K H L) : + LieSubmodule.toSubmodule (corootSpace α) = K ∙ coroot α := by + if hα : α.IsZero then + simp [hα.eq, coroot_eq_zero_iff.mpr hα] + else + set α' := (cartanEquivDual H).symm α + have hα' : (cartanEquivDual H).symm α ≠ 0 := by simpa using hα + have h₁ : (K ∙ coroot α) = K ∙ α' := by + 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 _ + have h₂ : (K ∙ (cartanEquivDual H).symm α : Submodule K H) ≤ corootSpace α := by + simpa using cartanEquivDual_symm_apply_mem_corootSpace K L H α + suffices finrank K (LieSubmodule.toSubmodule (corootSpace α)) ≤ 1 by + rw [← finrank_span_singleton (K := K) hα'] at this + exact h₁ ▸ (eq_of_le_of_finrank_le h₂ this).symm + have h : finrank K H = finrank K α.ker + 1 := + (Module.Dual.finrank_ker_add_one_of_ne_zero <| Weight.coe_toLinear_ne_zero_iff.mpr hα).symm + simpa [h] using Submodule.finrank_add_finrank_le_of_disjoint (disjoint_ker_weight_corootSpace α) + +@[simp] +lemma corootSpace_eq_bot_iff {α : Weight K H L} : + corootSpace α = ⊥ ↔ α.IsZero := by + simp [← LieSubmodule.coeSubmodule_eq_bot_iff, coe_corootSpace_eq_span_singleton α] + +lemma isCompl_ker_weight_span_coroot (α : Weight K H L) : + IsCompl α.ker (K ∙ coroot α) := by + if hα : α.IsZero then + simpa [Weight.coe_toLinear_eq_zero_iff.mpr hα, coroot_eq_zero_iff.mpr hα, Weight.ker] + using isCompl_top_bot + else + rw [← coe_corootSpace_eq_span_singleton] + apply Module.Dual.isCompl_ker_of_disjoint_of_ne_bot (by aesop) + (disjoint_ker_weight_corootSpace α) + replace hα : corootSpace α ≠ ⊥ := by simpa using hα + rwa [ne_eq, ← LieSubmodule.coe_toSubmodule_eq_iff] at hα + +end CharZero + +end IsKilling + +end Field + +end LieAlgebra From 0b80a0a765326c5309cc57f9a0ffaa8748dcc5ab Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 7 May 2024 18:47:04 +0000 Subject: [PATCH 20/52] refactor: Make `CompleteDistribLattice` extend `Coframe` (#12730) This spares us to write one instance by hand and also reduces the diff of #10560 a bit (after which we really need that `extends` clause) --- Mathlib/Order/CompleteBooleanAlgebra.lean | 9 +-------- 1 file changed, 1 insertion(+), 8 deletions(-) diff --git a/Mathlib/Order/CompleteBooleanAlgebra.lean b/Mathlib/Order/CompleteBooleanAlgebra.lean index d1ef857fb1589..c1fc95e8290e6 100644 --- a/Mathlib/Order/CompleteBooleanAlgebra.lean +++ b/Mathlib/Order/CompleteBooleanAlgebra.lean @@ -76,19 +76,12 @@ open Order /-- A complete distributive lattice is a complete lattice whose `⊔` and `⊓` respectively distribute over `⨅` and `⨆`. -/ -class CompleteDistribLattice (α : Type*) extends Frame α where - iInf_sup_le_sup_sInf : ∀ a s, ⨅ b ∈ s, a ⊔ b ≤ a ⊔ sInf s +class CompleteDistribLattice (α : Type*) extends Frame α, Coframe α #align complete_distrib_lattice CompleteDistribLattice /-- In a complete distributive lattice, `⊔` distributes over `⨅`. -/ add_decl_doc CompleteDistribLattice.iInf_sup_le_sup_sInf --- See note [lower instance priority] -instance (priority := 100) CompleteDistribLattice.toCoframe [CompleteDistribLattice α] : - Coframe α := - { ‹CompleteDistribLattice α› with } -#align complete_distrib_lattice.to_coframe CompleteDistribLattice.toCoframe - /-- A completely distributive lattice is a complete lattice whose `⨅` and `⨆` distribute over each other. -/ class CompletelyDistribLattice (α : Type u) extends CompleteLattice α where From b3037511489de57034366e564bdacc547da7b815 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 7 May 2024 20:39:28 +0000 Subject: [PATCH 21/52] feat: `mapsTo`, `surjOn` and `injOn` are preserved by `Filter.map` (#12385) And use it to simplify the proof of [completeSpace_iff_isComplete_range](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/UniformSpace/UniformEmbedding.html#completeSpace_iff_isComplete_range). I also add the consequence of the previous result that any uniform embedding from a complete space to a hausdorff space is a closed embedding (I agree this may be a bit unrelated, but I was in the area). Notes about the filter results: - the left/right inverse results are not used, but I think they can be helpful. I'm okay with removing them. - the new API doesn't match very well the existing one, but after way more time that I'm willing to admit I came to the conclusion that it doesn't make sense to make it better without reorganizing and splitting the whole file. --- Mathlib/Data/Set/Function.lean | 4 ++ Mathlib/Order/Filter/Basic.lean | 72 +++++++++++++++++++ .../UniformSpace/CompleteSeparated.lean | 11 +-- .../UniformSpace/UniformEmbedding.lean | 37 +++------- 4 files changed, 93 insertions(+), 31 deletions(-) diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 3a7041ef39b3d..3ec7c591aaba8 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -970,6 +970,10 @@ theorem eqOn_comp_right_iff : s.EqOn (g₁ ∘ f) (g₂ ∘ f) ↔ (f '' s).EqOn (s.surjOn_image f).cancel_right <| s.mapsTo_image f #align set.eq_on_comp_right_iff Set.eqOn_comp_right_iff +theorem SurjOn.forall {p : β → Prop} (hf : s.SurjOn f t) (hf' : s.MapsTo f t) : + (∀ y ∈ t, p y) ↔ (∀ x ∈ s, p (f x)) := + ⟨fun H x hx ↦ H (f x) (hf' hx), fun H _y hy ↦ let ⟨x, hx, hxy⟩ := hf hy; hxy ▸ H x hx⟩ + end surjOn /-! ### Bijectivity -/ diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index a983e7c584730..5967e3e80a1a9 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -2163,6 +2163,9 @@ theorem map_pure (f : α → β) (a : α) : map f (pure a) = pure (f a) := rfl #align filter.map_pure Filter.map_pure +theorem pure_le_principal (a : α) : pure a ≤ 𝓟 s ↔ a ∈ s := by + simp + @[simp] theorem join_pure (f : Filter α) : join (pure f) = f := rfl #align filter.join_pure Filter.join_pure @@ -2304,6 +2307,37 @@ theorem _root_.Function.Commute.filter_comap {f g : α → α} (h : Function.Com h.semiconj.filter_comap #align function.commute.filter_comap Function.Commute.filter_comap +section + +open Filter + +theorem _root_.Function.LeftInverse.filter_map {f : α → β} {g : β → α} (hfg : LeftInverse g f) : + LeftInverse (map g) (map f) := fun F ↦ by + rw [map_map, hfg.comp_eq_id, map_id] + +theorem _root_.Function.LeftInverse.filter_comap {f : α → β} {g : β → α} (hfg : LeftInverse g f) : + RightInverse (comap g) (comap f) := fun F ↦ by + rw [comap_comap, hfg.comp_eq_id, comap_id] + +nonrec theorem _root_.Function.RightInverse.filter_map {f : α → β} {g : β → α} + (hfg : RightInverse g f) : RightInverse (map g) (map f) := + hfg.filter_map + +nonrec theorem _root_.Function.RightInverse.filter_comap {f : α → β} {g : β → α} + (hfg : RightInverse g f) : LeftInverse (comap g) (comap f) := + hfg.filter_comap + +theorem _root_.Set.LeftInvOn.filter_map_Iic {f : α → β} {g : β → α} (hfg : LeftInvOn g f s) : + LeftInvOn (map g) (map f) (Iic <| 𝓟 s) := fun F (hF : F ≤ 𝓟 s) ↦ by + have : (g ∘ f) =ᶠ[𝓟 s] id := by simpa only [eventuallyEq_principal] using hfg + rw [map_map, map_congr (this.filter_mono hF), map_id] + +nonrec theorem _root_.Set.RightInvOn.filter_map_Iic {f : α → β} {g : β → α} + (hfg : RightInvOn g f t) : RightInvOn (map g) (map f) (Iic <| 𝓟 t) := + hfg.filter_map_Iic + +end + @[simp] theorem comap_principal {t : Set β} : comap m (𝓟 t) = 𝓟 (m ⁻¹' t) := Filter.ext fun _ => ⟨fun ⟨_u, hu, b⟩ => (preimage_mono hu).trans b, @@ -3302,6 +3336,44 @@ theorem Filter.EventuallyEq.comp_tendsto {f' : α → β} (H : f =ᶠ[l] f') {g hg.eventually H #align filter.eventually_eq.comp_tendsto Filter.EventuallyEq.comp_tendsto +theorem Filter.map_mapsTo_Iic_iff_tendsto {m : α → β} : + MapsTo (map m) (Iic F) (Iic G) ↔ Tendsto m F G := + ⟨fun hm ↦ hm right_mem_Iic, fun hm _ ↦ hm.mono_left⟩ + +alias ⟨_, Filter.Tendsto.map_mapsTo_Iic⟩ := Filter.map_mapsTo_Iic_iff_tendsto + +theorem Filter.map_mapsTo_Iic_iff_mapsTo {m : α → β} : + MapsTo (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 t) ↔ MapsTo m s t := + by rw [map_mapsTo_Iic_iff_tendsto, tendsto_principal_principal, MapsTo] + +alias ⟨_, Set.MapsTo.filter_map_Iic⟩ := Filter.map_mapsTo_Iic_iff_mapsTo + +-- TODO(Anatole): unify with the global case + +theorem Filter.map_surjOn_Iic_iff_le_map {m : α → β} : + SurjOn (map m) (Iic F) (Iic G) ↔ G ≤ map m F := by + refine ⟨fun hm ↦ ?_, fun hm ↦ ?_⟩ + · rcases hm right_mem_Iic with ⟨H, (hHF : H ≤ F), rfl⟩ + exact map_mono hHF + · have : RightInvOn (F ⊓ comap m ·) (map m) (Iic G) := + fun H (hHG : H ≤ G) ↦ by simpa [Filter.push_pull] using hHG.trans hm + exact this.surjOn fun H _ ↦ mem_Iic.mpr inf_le_left + +theorem Filter.map_surjOn_Iic_iff_surjOn {m : α → β} : + SurjOn (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 t) ↔ SurjOn m s t := by + rw [map_surjOn_Iic_iff_le_map, map_principal, principal_mono, SurjOn] + +alias ⟨_, Set.SurjOn.filter_map_Iic⟩ := Filter.map_surjOn_Iic_iff_surjOn + +theorem Filter.filter_injOn_Iic_iff_injOn {m : α → β} : + InjOn (map m) (Iic <| 𝓟 s) ↔ InjOn m s := by + refine ⟨fun hm x hx y hy hxy ↦ ?_, fun hm F hF G hG ↦ ?_⟩ + · rwa [← pure_injective.eq_iff, ← map_pure, ← map_pure, hm.eq_iff, pure_injective.eq_iff] + at hxy <;> rwa [mem_Iic, pure_le_principal] + · simp [map_eq_map_iff_of_injOn (le_principal_iff.mp hF) (le_principal_iff.mp hG) hm] + +alias ⟨_, Set.InjOn.filter_map_Iic⟩ := Filter.filter_injOn_Iic_iff_injOn + namespace Filter /-- Construct a filter from a property that is stable under finite unions. diff --git a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean index 7d5d8da26814e..17b539df7f84e 100644 --- a/Mathlib/Topology/UniformSpace/CompleteSeparated.lean +++ b/Mathlib/Topology/UniformSpace/CompleteSeparated.lean @@ -3,9 +3,7 @@ Copyright (c) 2019 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.UniformSpace.Cauchy -import Mathlib.Topology.UniformSpace.Separation -import Mathlib.Topology.DenseEmbedding +import Mathlib.Topology.UniformSpace.UniformEmbedding #align_import topology.uniform_space.complete_separated from "leanprover-community/mathlib"@"b363547b3113d350d053abdf2884e9850a56b205" @@ -20,7 +18,7 @@ open Filter open Topology Filter -variable {α : Type*} +variable {α β : Type*} /-- In a separated space, a complete set is closed. -/ theorem IsComplete.isClosed [UniformSpace α] [T0Space α] {s : Set α} (h : IsComplete s) : @@ -32,6 +30,11 @@ theorem IsComplete.isClosed [UniformSpace α] [T0Space α] {s : Set α} (h : IsC rwa [(tendsto_nhds_unique' ha inf_le_left fy : a = y)] #align is_complete.is_closed IsComplete.isClosed +theorem UniformEmbedding.toClosedEmbedding [UniformSpace α] [UniformSpace β] [CompleteSpace α] + [T0Space β] {f : α → β} (hf : UniformEmbedding f) : + ClosedEmbedding f := + ⟨hf.embedding, hf.toUniformInducing.isComplete_range.isClosed⟩ + namespace DenseInducing open Filter diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index c62d495523247..761310269770c 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -291,38 +291,18 @@ theorem UniformEmbedding.prod {α' : Type*} {β' : Type*} [UniformSpace α'] [Un { h₁.toUniformInducing.prod h₂.toUniformInducing with inj := h₁.inj.Prod_map h₂.inj } #align uniform_embedding.prod UniformEmbedding.prod -theorem isComplete_of_complete_image {m : α → β} {s : Set α} (hm : UniformInducing m) - (hs : IsComplete (m '' s)) : IsComplete s := by - intro f hf hfs - rw [le_principal_iff] at hfs - obtain ⟨_, ⟨x, hx, rfl⟩, hyf⟩ : ∃ y ∈ m '' s, map m f ≤ 𝓝 y - · exact hs (f.map m) (hf.map hm.uniformContinuous) (le_principal_iff.2 (image_mem_map hfs)) - rw [map_le_iff_le_comap, ← nhds_induced, ← hm.inducing.induced] at hyf - exact ⟨x, hx, hyf⟩ -#align is_complete_of_complete_image isComplete_of_complete_image - -theorem IsComplete.completeSpace_coe {s : Set α} (hs : IsComplete s) : CompleteSpace s := - completeSpace_iff_isComplete_univ.2 <| - isComplete_of_complete_image uniformEmbedding_subtype_val.toUniformInducing <| by simp [hs] -#align is_complete.complete_space_coe IsComplete.completeSpace_coe - /-- A set is complete iff its image under a uniform inducing map is complete. -/ theorem isComplete_image_iff {m : α → β} {s : Set α} (hm : UniformInducing m) : IsComplete (m '' s) ↔ IsComplete s := by - refine' ⟨isComplete_of_complete_image hm, fun c => _⟩ - haveI : CompleteSpace s := c.completeSpace_coe - set m' : s → β := m ∘ (↑) - suffices IsComplete (range m') by rwa [range_comp, Subtype.range_coe] at this - have hm' : UniformInducing m' := hm.comp uniformEmbedding_subtype_val.toUniformInducing - intro f hf hfm - rw [Filter.le_principal_iff] at hfm - have cf' : Cauchy (comap m' f) := - hf.comap' hm'.comap_uniformity.le (NeBot.comap_of_range_mem hf.1 hfm) - rcases CompleteSpace.complete cf' with ⟨x, hx⟩ - rw [hm'.inducing.nhds_eq_comap, comap_le_comap_iff hfm] at hx - exact ⟨m' x, mem_range_self _, hx⟩ + have fact1 : SurjOn (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 <| m '' s) := surjOn_image .. |>.filter_map_Iic + have fact2 : MapsTo (map m) (Iic <| 𝓟 s) (Iic <| 𝓟 <| m '' s) := mapsTo_image .. |>.filter_map_Iic + simp_rw [IsComplete, imp.swap (a := Cauchy _), ← mem_Iic (b := 𝓟 _), fact1.forall fact2, + hm.cauchy_map_iff, exists_mem_image, map_le_iff_le_comap, hm.inducing.nhds_eq_comap] #align is_complete_image_iff isComplete_image_iff +alias ⟨isComplete_of_complete_image, _⟩ := isComplete_image_iff +#align is_complete_of_complete_image isComplete_of_complete_image + theorem completeSpace_iff_isComplete_range {f : α → β} (hf : UniformInducing f) : CompleteSpace α ↔ IsComplete (range f) := by rw [completeSpace_iff_isComplete_univ, ← isComplete_image_iff hf, image_univ] @@ -354,6 +334,9 @@ theorem completeSpace_coe_iff_isComplete {s : Set α} : CompleteSpace s ↔ IsCo rw [Subtype.range_coe] #align complete_space_coe_iff_is_complete completeSpace_coe_iff_isComplete +alias ⟨_, IsComplete.completeSpace_coe⟩ := completeSpace_coe_iff_isComplete +#align is_complete.complete_space_coe IsComplete.completeSpace_coe + theorem IsClosed.completeSpace_coe [CompleteSpace α] {s : Set α} (hs : IsClosed s) : CompleteSpace s := hs.isComplete.completeSpace_coe From 4c7c71d37bb6151e2a6464c031ce5f4ba4950ba8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Dagur=20T=C3=B3mas=20=C3=81sgeirsson?= Date: Tue, 7 May 2024 21:46:46 +0000 Subject: [PATCH 22/52] feat(LightProfinite): (co)limits (#9513) This PR constructs pullbacks and finite coproducts in `LightProfinite`. --- Mathlib.lean | 1 + .../Topology/Category/CompHaus/Limits.lean | 39 +-- .../Category/LightProfinite/Basic.lean | 28 +- .../Category/LightProfinite/Limits.lean | 243 ++++++++++++++++++ .../Topology/Category/Profinite/Limits.lean | 40 ++- Mathlib/Topology/Category/Stonean/Limits.lean | 61 ++--- 6 files changed, 311 insertions(+), 101 deletions(-) create mode 100644 Mathlib/Topology/Category/LightProfinite/Limits.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2fd01119c36fe..ee72500dd3ea7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3912,6 +3912,7 @@ import Mathlib.Topology.Category.CompHaus.Projective import Mathlib.Topology.Category.Compactum import Mathlib.Topology.Category.LightProfinite.Basic import Mathlib.Topology.Category.LightProfinite.IsLight +import Mathlib.Topology.Category.LightProfinite.Limits import Mathlib.Topology.Category.Locale import Mathlib.Topology.Category.Profinite.AsLimit import Mathlib.Topology.Category.Profinite.Basic diff --git a/Mathlib/Topology/Category/CompHaus/Limits.lean b/Mathlib/Topology/Category/CompHaus/Limits.lean index 60856ea566928..b7809f86de225 100644 --- a/Mathlib/Topology/Category/CompHaus/Limits.lean +++ b/Mathlib/Topology/Category/CompHaus/Limits.lean @@ -176,26 +176,17 @@ lemma finiteCoproduct.hom_ext {B : CompHaus} (f g : finiteCoproduct X ⟶ B) apply_fun (fun q => q x) at h exact h -/-- -The coproduct cocone associated to the explicit finite coproduct. --/ -@[simps] -def finiteCoproduct.cocone : Limits.Cocone (Discrete.functor X) where - pt := finiteCoproduct X - ι := Discrete.natTrans fun ⟨a⟩ => finiteCoproduct.ι X a - -/-- -The explicit finite coproduct cocone is a colimit cocone. --/ -@[simps] -def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cocone X) where - desc := fun s => finiteCoproduct.desc _ fun a => s.ι.app ⟨a⟩ - fac := fun s ⟨a⟩ => finiteCoproduct.ι_desc _ _ _ - uniq := fun s m hm => finiteCoproduct.hom_ext _ _ _ fun a => by - specialize hm ⟨a⟩ - ext t - apply_fun (fun q => q t) at hm - exact hm +/-- The coproduct cocone associated to the explicit finite coproduct. -/ +abbrev finiteCoproduct.cofan : Limits.Cofan X := + Cofan.mk (finiteCoproduct X) (finiteCoproduct.ι X) + +/-- The explicit finite coproduct cocone is a colimit cocone. -/ +def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cofan X) := + mkCofanColimit _ + (fun s ↦ desc _ fun a ↦ s.inj a) + (fun _ _ ↦ ι_desc _ _ _) + fun _ _ hm ↦ finiteCoproduct.hom_ext _ _ _ fun a ↦ + (DFunLike.ext _ _ fun t ↦ congrFun (congrArg DFunLike.coe (hm a)) t) section Iso @@ -207,9 +198,7 @@ def coproductIsoCoproduct : finiteCoproduct X ≅ ∐ X := theorem Sigma.ι_comp_toFiniteCoproduct (a : α) : (Limits.Sigma.ι X a) ≫ (coproductIsoCoproduct X).inv = finiteCoproduct.ι X a := by - dsimp [coproductIsoCoproduct] - simp only [Limits.colimit.comp_coconePointUniqueUpToIso_inv, finiteCoproduct.cocone_pt, - finiteCoproduct.cocone_ι, Discrete.natTrans_app] + simp [coproductIsoCoproduct] /-- The homeomorphism from the explicit finite coproducts to the abstract coproduct. -/ noncomputable @@ -248,8 +237,8 @@ section Terminal /-- A one-element space is terminal in `CompHaus` -/ def isTerminalPUnit : IsTerminal (CompHaus.of PUnit.{u + 1}) := - haveI : ∀ X, Unique (X ⟶ CompHaus.of PUnit.{u + 1}) := fun X => - ⟨⟨⟨fun _ => PUnit.unit, continuous_const⟩⟩, fun f => by ext; aesop⟩ + haveI : ∀ X, Unique (X ⟶ CompHaus.of PUnit.{u + 1}) := fun _ ↦ + ⟨⟨⟨fun _ => PUnit.unit, continuous_const⟩⟩, fun _ ↦ rfl⟩ Limits.IsTerminal.ofUnique _ /-- The isomorphism from an arbitrary terminal object of `CompHaus` to a one-element space. -/ diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 33a2794c472af..fb60ee99da909 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -36,6 +36,7 @@ structure LightProfinite : Type (u+1) where isLimit : IsLimit cone /-- A finite set can be regarded as a light profinite set as the limit of the constant diagram. -/ +@[simps] def FintypeCat.toLightProfinite (X : FintypeCat) : LightProfinite where diagram := (Functor.const _).obj X cone := { @@ -58,6 +59,7 @@ theorem ext {Y : LightProfinite} {a b : Y.cone.pt} Given a functor from `ℕᵒᵖ` to finite sets we can take its limit in `Profinite` and obtain a light profinite set.  -/ +@[simps] noncomputable def of (F : ℕᵒᵖ ⥤ FintypeCat) : LightProfinite.{u} where diagram := F isLimit := limit.isLimit (F ⋙ FintypeCat.toProfinite) @@ -79,8 +81,6 @@ instance : lightToProfinite.Faithful := show (inducedFunctor _).Faithful from in instance : lightToProfinite.Full := show (inducedFunctor _).Full from inferInstance -instance : lightToProfinite.ReflectsEpimorphisms := inferInstance - instance {X : LightProfinite} : TopologicalSpace ((forget LightProfinite).obj X) := (inferInstance : TopologicalSpace X.cone.pt) @@ -94,6 +94,7 @@ instance {X : LightProfinite} : T2Space ((forget LightProfinite).obj X) := (inferInstance : T2Space X.cone.pt ) /-- The explicit functor `FintypeCat ⥤ LightProfinite`. -/ +@[simps] def fintypeCatToLightProfinite : FintypeCat ⥤ LightProfinite.{u} where obj X := X.toLightProfinite map f := FintypeCat.toProfinite.map f @@ -104,10 +105,21 @@ instance : fintypeCatToLightProfinite.Faithful where instance : fintypeCatToLightProfinite.Full where map_surjective f := ⟨fun x ↦ f x, rfl⟩ +/-- The fully faithful embedding of `LightProfinite` in `TopCat`. -/ +@[simps!] +def toTopCat : LightProfinite ⥤ TopCat := + lightToProfinite ⋙ Profinite.toTopCat + +instance : toTopCat.Faithful := Functor.Faithful.comp _ _ + +instance : toTopCat.Full := Functor.Full.comp _ _ + end LightProfinite noncomputable section EssentiallySmall +open LightProfinite + /-- This is an auxiliary definition used to show that `LightProfinite` is essentially small. -/ structure LightProfinite' : Type u where /-- The diagram takes values in a small category equivalent to `FintypeCat`. -/ @@ -130,13 +142,11 @@ instance : LightProfinite'.toLightFunctor.{u}.Full where map_surjective f := ⟨f, rfl⟩ instance : LightProfinite'.toLightFunctor.{u}.EssSurj where - mem_essImage Y := by - let i : limit (((Y.diagram ⋙ Skeleton.equivalence.inverse) ⋙ Skeleton.equivalence.functor) ⋙ - toProfinite) ≅ Y.cone.pt := (Limits.lim.mapIso (isoWhiskerRight ((Functor.associator _ _ _) ≪≫ - (isoWhiskerLeft Y.diagram Skeleton.equivalence.counitIso)) toProfinite)) ≪≫ - IsLimit.conePointUniqueUpToIso (limit.isLimit _) Y.isLimit - exact ⟨⟨Y.diagram ⋙ Skeleton.equivalence.inverse⟩, ⟨⟨i.hom, i.inv, i.hom_inv_id, i.inv_hom_id⟩⟩⟩ - -- why can't I just write `i` instead of `⟨i.hom, i.inv, i.hom_inv_id, i.inv_hom_id⟩`? + mem_essImage Y := + ⟨⟨Y.diagram ⋙ Skeleton.equivalence.inverse⟩, ⟨lightToProfinite.preimageIso ( + (Limits.lim.mapIso (isoWhiskerRight ((isoWhiskerLeft Y.diagram + Skeleton.equivalence.counitIso)) toProfinite)) ≪≫ + (limit.isLimit _).conePointUniqueUpToIso Y.isLimit)⟩⟩ instance : LightProfinite'.toLightFunctor.IsEquivalence := Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ diff --git a/Mathlib/Topology/Category/LightProfinite/Limits.lean b/Mathlib/Topology/Category/LightProfinite/Limits.lean new file mode 100644 index 0000000000000..a3fa134a70f77 --- /dev/null +++ b/Mathlib/Topology/Category/LightProfinite/Limits.lean @@ -0,0 +1,243 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.Topology.Category.LightProfinite.IsLight +import Mathlib.Topology.Category.Profinite.Limits +/-! + +# Explicit limits and colimits + +This file collects some constructions of explicit limits and colimits in `LightProfinite`, +which may be useful due to their definitional properties. + +## Main definitions + +* `LightProfinite.pullback`: Explicit pullback, defined in the "usual" way as a subset of the + product. + +* `LightProfinite.finiteCoproduct`: Explicit finite coproducts, defined as a disjoint union. + +-/ + +universe u w + +open CategoryTheory Profinite TopologicalSpace Limits + +namespace LightProfinite + +section Pullback + +instance {X Y B : Profinite.{u}} (f : X ⟶ B) (g : Y ⟶ B) [X.IsLight] [Y.IsLight] : + (Profinite.pullback f g).IsLight := by + let i : Profinite.pullback f g ⟶ Profinite.of (X × Y) := ⟨fun x ↦ x.val, continuous_induced_dom⟩ + have : Mono i := by + rw [Profinite.mono_iff_injective] + exact Subtype.val_injective + exact isLight_of_mono i + +variable {X Y B : LightProfinite.{u}} (f : X ⟶ B) (g : Y ⟶ B) + +/-- +The "explicit" pullback of two morphisms `f, g` in `LightProfinite`, whose underlying profinite set +is the set of pairs `(x, y)` such that `f x = g y`, with the topology induced by the product. +-/ +noncomputable def pullback : LightProfinite.{u} := + ofIsLight (Profinite.pullback (lightToProfinite.map f) (lightToProfinite.map g)) + +/-- The projection from the pullback to the first component. -/ +def pullback.fst : pullback f g ⟶ X where + toFun := fun ⟨⟨x, _⟩, _⟩ ↦ x + continuous_toFun := Continuous.comp continuous_fst continuous_subtype_val + +/-- The projection from the pullback to the second component. -/ +def pullback.snd : pullback f g ⟶ Y where + toFun := fun ⟨⟨_, y⟩, _⟩ ↦ y + continuous_toFun := Continuous.comp continuous_snd continuous_subtype_val + +@[reassoc] +lemma pullback.condition : pullback.fst f g ≫ f = pullback.snd f g ≫ g := by + ext ⟨_, h⟩ + exact h + +/-- +Construct a morphism to the explicit pullback given morphisms to the factors +which are compatible with the maps to the base. +This is essentially the universal property of the pullback. +-/ +def pullback.lift {Z : LightProfinite.{u}} (a : Z ⟶ X) (b : Z ⟶ Y) (w : a ≫ f = b ≫ g) : + Z ⟶ pullback f g where + toFun := fun z ↦ ⟨⟨a z, b z⟩, by apply_fun (· z) at w; exact w⟩ + continuous_toFun := by + apply Continuous.subtype_mk + rw [continuous_prod_mk] + exact ⟨a.continuous, b.continuous⟩ + +@[reassoc (attr := simp)] +lemma pullback.lift_fst {Z : LightProfinite.{u}} (a : Z ⟶ X) (b : Z ⟶ Y) (w : a ≫ f = b ≫ g) : + pullback.lift f g a b w ≫ pullback.fst f g = a := rfl + +@[reassoc (attr := simp)] +lemma pullback.lift_snd {Z : LightProfinite.{u}} (a : Z ⟶ X) (b : Z ⟶ Y) (w : a ≫ f = b ≫ g) : + pullback.lift f g a b w ≫ pullback.snd f g = b := rfl + +lemma pullback.hom_ext {Z : LightProfinite.{u}} (a b : Z ⟶ pullback f g) + (hfst : a ≫ pullback.fst f g = b ≫ pullback.fst f g) + (hsnd : a ≫ pullback.snd f g = b ≫ pullback.snd f g) : a = b := by + ext z + apply_fun (· z) at hfst hsnd + apply Subtype.ext + apply Prod.ext + · exact hfst + · exact hsnd + +/-- The pullback cone whose cone point is the explicit pullback. -/ +@[simps! pt π] +noncomputable def pullback.cone : Limits.PullbackCone f g := + Limits.PullbackCone.mk (pullback.fst f g) (pullback.snd f g) (pullback.condition f g) + +/-- The explicit pullback cone is a limit cone. -/ +@[simps! lift] +def pullback.isLimit : Limits.IsLimit (pullback.cone f g) := + Limits.PullbackCone.isLimitAux _ + (fun s ↦ pullback.lift f g s.fst s.snd s.condition) + (fun _ ↦ pullback.lift_fst _ _ _ _ _) + (fun _ ↦ pullback.lift_snd _ _ _ _ _) + (fun _ _ hm ↦ pullback.hom_ext _ _ _ _ (hm .left) (hm .right)) + +end Pullback + +section FiniteCoproduct + +instance {α : Type w} [Finite α] (X : α → Profinite.{max u w}) [∀ a, (X a).IsLight] : + (Profinite.finiteCoproduct X).IsLight where + countable_clopens := by + refine @Function.Surjective.countable ((a : α) → Clopens (X a)) _ inferInstance + (fun f ↦ ⟨⋃ (a : α), Sigma.mk a '' (f a).1, ?_⟩) ?_ + · apply isClopen_iUnion_of_finite + intro i + exact ⟨isClosedMap_sigmaMk _ (f i).2.1, isOpenMap_sigmaMk _ (f i).2.2⟩ + · intro ⟨s, ⟨hsc, hso⟩⟩ + rw [isOpen_sigma_iff] at hso + rw [isClosed_sigma_iff] at hsc + refine ⟨fun i ↦ ⟨_, ⟨hsc i, hso i⟩⟩, ?_⟩ + simp only [Subtype.mk.injEq] + ext ⟨i, xi⟩ + refine ⟨fun hx ↦ ?_, fun hx ↦ ?_⟩ + · simp only [Clopens.coe_mk, Set.mem_iUnion] at hx + obtain ⟨_, _, hj, hxj⟩ := hx + simpa [hxj] using hj + · simp only [Clopens.coe_mk, Set.mem_iUnion] + refine ⟨i, xi, (by simpa using hx), rfl⟩ + +variable {α : Type w} [Finite α] (X : α → LightProfinite.{max u w}) + +/-- +The "explicit" coproduct of a finite family of objects in `LightProfinite`, whose underlying +profinite set is the disjoint union with its usual topology. +-/ +noncomputable +def finiteCoproduct : LightProfinite := + ofIsLight (Profinite.finiteCoproduct.{u, w} fun a ↦ (X a).toProfinite) + +/-- The inclusion of one of the factors into the explicit finite coproduct. -/ +def finiteCoproduct.ι (a : α) : X a ⟶ finiteCoproduct X where + toFun := (⟨a, ·⟩) + continuous_toFun := continuous_sigmaMk (σ := fun a ↦ (X a).toProfinite) + +/-- +To construct a morphism from the explicit finite coproduct, it suffices to +specify a morphism from each of its factors. This is the universal property of the coproduct. +-/ +def finiteCoproduct.desc {B : LightProfinite.{max u w}} (e : (a : α) → (X a ⟶ B)) : + finiteCoproduct X ⟶ B where + toFun := fun ⟨a, x⟩ ↦ e a x + continuous_toFun := by + apply continuous_sigma + intro a + exact (e a).continuous + +@[reassoc (attr := simp)] +lemma finiteCoproduct.ι_desc {B : LightProfinite.{max u w}} (e : (a : α) → (X a ⟶ B)) (a : α) : + finiteCoproduct.ι X a ≫ finiteCoproduct.desc X e = e a := rfl + +lemma finiteCoproduct.hom_ext {B : LightProfinite.{max u w}} (f g : finiteCoproduct X ⟶ B) + (h : ∀ a : α, finiteCoproduct.ι X a ≫ f = finiteCoproduct.ι X a ≫ g) : f = g := by + ext ⟨a, x⟩ + specialize h a + apply_fun (· x) at h + exact h + +/-- The coproduct cocone associated to the explicit finite coproduct. -/ +noncomputable abbrev finiteCoproduct.cofan : Limits.Cofan X := + Cofan.mk (finiteCoproduct X) (finiteCoproduct.ι X) + +/-- The explicit finite coproduct cocone is a colimit cocone. -/ +def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cofan X) := + mkCofanColimit _ + (fun s ↦ desc _ fun a ↦ s.inj a) + (fun s a ↦ ι_desc _ _ _) + fun s m hm ↦ finiteCoproduct.hom_ext _ _ _ fun a ↦ + (by ext t; exact congrFun (congrArg DFunLike.coe (hm a)) t) + +end FiniteCoproduct + +section HasPreserves + +instance (n : ℕ) (F : Discrete (Fin n) ⥤ LightProfinite) : + HasColimit (Discrete.functor (F.obj ∘ Discrete.mk) : Discrete (Fin n) ⥤ LightProfinite) where + exists_colimit := ⟨⟨finiteCoproduct.cofan _, finiteCoproduct.isColimit _⟩⟩ + +instance : HasFiniteCoproducts LightProfinite where + out _ := { has_colimit := fun _ ↦ hasColimitOfIso Discrete.natIsoFunctor } + +instance {X Y B : LightProfinite} (f : X ⟶ B) (g : Y ⟶ B) : HasLimit (cospan f g) where + exists_limit := ⟨⟨pullback.cone f g, pullback.isLimit f g⟩⟩ + +instance : HasPullbacks LightProfinite where + has_limit F := hasLimitOfIso (diagramIsoCospan F).symm + +noncomputable +instance : PreservesFiniteCoproducts lightToProfinite := by + refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ + suffices PreservesColimit (Discrete.functor (F.obj ∘ Discrete.mk)) lightToProfinite by + exact preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + apply preservesColimitOfPreservesColimitCocone (finiteCoproduct.isColimit _) + exact Profinite.finiteCoproduct.isColimit _ + +noncomputable +instance : PreservesLimitsOfShape WalkingCospan lightToProfinite := by + suffices ∀ {X Y B} (f : X ⟶ B) (g : Y ⟶ B), PreservesLimit (cospan f g) lightToProfinite from + ⟨fun {F} ↦ preservesLimitOfIsoDiagram _ (diagramIsoCospan F).symm⟩ + intro _ _ _ f g + apply preservesLimitOfPreservesLimitCone (pullback.isLimit f g) + exact (isLimitMapConePullbackConeEquiv lightToProfinite (pullback.condition f g)).symm + (Profinite.pullback.isLimit _ _) + +instance (X : LightProfinite) : Unique (X ⟶ (FintypeCat.of PUnit.{u+1}).toLightProfinite) := + ⟨⟨⟨fun _ ↦ PUnit.unit, continuous_const⟩⟩, fun _ ↦ rfl⟩ + +/-- A one-element space is terminal in `LightProfinite` -/ +def isTerminalPUnit : IsTerminal ((FintypeCat.of PUnit.{u+1}).toLightProfinite) := + Limits.IsTerminal.ofUnique _ + +instance : HasTerminal LightProfinite.{u} := + Limits.hasTerminal_of_unique (FintypeCat.of PUnit.{u+1}).toLightProfinite + +/-- The isomorphism from an arbitrary terminal object of `LightProfinite` to a one-element space. -/ +noncomputable def terminalIsoPUnit : + ⊤_ LightProfinite.{u} ≅ (FintypeCat.of PUnit.{u+1}).toLightProfinite := + terminalIsTerminal.uniqueUpToIso LightProfinite.isTerminalPUnit + +noncomputable instance : PreservesFiniteCoproducts LightProfinite.toTopCat.{u} where + preserves _ _ := (inferInstance : + PreservesColimitsOfShape _ (LightProfinite.lightToProfinite.{u} ⋙ Profinite.toTopCat.{u})) + +noncomputable instance : PreservesLimitsOfShape WalkingCospan LightProfinite.toTopCat.{u} := + (inferInstance : PreservesLimitsOfShape WalkingCospan + (LightProfinite.lightToProfinite.{u} ⋙ Profinite.toTopCat.{u})) + +end HasPreserves + +end LightProfinite diff --git a/Mathlib/Topology/Category/Profinite/Limits.lean b/Mathlib/Topology/Category/Profinite/Limits.lean index 4c4ec29e657c0..8fefb8fd7b56e 100644 --- a/Mathlib/Topology/Category/Profinite/Limits.lean +++ b/Mathlib/Topology/Category/Profinite/Limits.lean @@ -23,7 +23,7 @@ which may be useful due to their definitional properties. namespace Profinite -universe u +universe u w open CategoryTheory Limits @@ -130,7 +130,7 @@ end Pullbacks section FiniteCoproducts -variable {α : Type} [Finite α] (X : α → Profinite.{u}) +variable {α : Type w} [Finite α] (X : α → Profinite.{max u w}) /-- The coproduct of a finite family of objects in `Profinite`, constructed as the disjoint @@ -148,7 +148,7 @@ To construct a morphism from the explicit finite coproduct, it suffices to specify a morphism from each of its factors. This is essentially the universal property of the coproduct. -/ -def finiteCoproduct.desc {B : Profinite.{u}} (e : (a : α) → (X a ⟶ B)) : +def finiteCoproduct.desc {B : Profinite.{max u w}} (e : (a : α) → (X a ⟶ B)) : finiteCoproduct X ⟶ B where toFun := fun ⟨a, x⟩ => e a x continuous_toFun := by @@ -157,10 +157,10 @@ def finiteCoproduct.desc {B : Profinite.{u}} (e : (a : α) → (X a ⟶ B)) : exact (e a).continuous @[reassoc (attr := simp)] -lemma finiteCoproduct.ι_desc {B : Profinite.{u}} (e : (a : α) → (X a ⟶ B)) (a : α) : +lemma finiteCoproduct.ι_desc {B : Profinite.{max u w}} (e : (a : α) → (X a ⟶ B)) (a : α) : finiteCoproduct.ι X a ≫ finiteCoproduct.desc X e = e a := rfl -lemma finiteCoproduct.hom_ext {B : Profinite.{u}} (f g : finiteCoproduct X ⟶ B) +lemma finiteCoproduct.hom_ext {B : Profinite.{max u w}} (f g : finiteCoproduct X ⟶ B) (h : ∀ a : α, finiteCoproduct.ι X a ≫ f = finiteCoproduct.ι X a ≫ g) : f = g := by ext ⟨a, x⟩ specialize h a @@ -168,21 +168,16 @@ lemma finiteCoproduct.hom_ext {B : Profinite.{u}} (f g : finiteCoproduct X ⟶ B exact h /-- The coproduct cocone associated to the explicit finite coproduct. -/ -@[simps] -def finiteCoproduct.cocone : Limits.Cocone (Discrete.functor X) where - pt := finiteCoproduct X - ι := Discrete.natTrans fun ⟨a⟩ => finiteCoproduct.ι X a +abbrev finiteCoproduct.cofan : Limits.Cofan X := + Cofan.mk (finiteCoproduct X) (finiteCoproduct.ι X) /-- The explicit finite coproduct cocone is a colimit cocone. -/ -@[simps] -def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cocone X) where - desc := fun s => finiteCoproduct.desc _ fun a => s.ι.app ⟨a⟩ - fac := fun s ⟨a⟩ => finiteCoproduct.ι_desc _ _ _ - uniq := fun s m hm => finiteCoproduct.hom_ext _ _ _ fun a => by - specialize hm ⟨a⟩ - ext t - apply_fun (· t) at hm - exact hm +def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cofan X) := + mkCofanColimit _ + (fun s ↦ desc _ fun a ↦ s.inj a) + (fun s a ↦ ι_desc _ _ _) + fun s m hm ↦ finiteCoproduct.hom_ext _ _ _ fun a ↦ + (by ext t; exact congrFun (congrArg DFunLike.coe (hm a)) t) section Iso @@ -193,13 +188,12 @@ Limits.IsColimit.coconePointUniqueUpToIso (finiteCoproduct.isColimit X) (Limits. theorem Sigma.ι_comp_toFiniteCoproduct (a : α) : (Limits.Sigma.ι X a) ≫ (coproductIsoCoproduct X).inv = finiteCoproduct.ι X a := by - simp only [coproductIsoCoproduct, Limits.colimit.comp_coconePointUniqueUpToIso_inv, - finiteCoproduct.cocone_pt, finiteCoproduct.cocone_ι, Discrete.natTrans_app] + simp [coproductIsoCoproduct] /-- The homeomorphism from the explicit finite coproducts to the abstract coproduct. -/ noncomputable def coproductHomeoCoproduct : finiteCoproduct X ≃ₜ (∐ X : _) := -Profinite.homeoOfIso (coproductIsoCoproduct X) + Profinite.homeoOfIso (coproductIsoCoproduct X) end Iso @@ -223,6 +217,10 @@ instance : PreservesFiniteCoproducts profiniteToCompHaus := by apply preservesColimitOfPreservesColimitCocone (Profinite.finiteCoproduct.isColimit _) exact CompHaus.finiteCoproduct.isColimit _ +noncomputable instance : PreservesFiniteCoproducts Profinite.toTopCat.{u} where + preserves _ _:= (inferInstance : + PreservesColimitsOfShape _ (profiniteToCompHaus.{u} ⋙ compHausToTop.{u})) + instance : FinitaryExtensive Profinite := finitaryExtensive_of_preserves_and_reflects profiniteToCompHaus diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index ac85d30e6c17e..b877fbe3dd023 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -73,59 +73,29 @@ lemma finiteCoproduct.hom_ext {B : Stonean.{u}} (f g : finiteCoproduct X ⟶ B) exact h /-- The coproduct cocone associated to the explicit finite coproduct. -/ -@[simps] -def finiteCoproduct.cocone (F : Discrete α ⥤ Stonean) : - Cocone F where - pt := finiteCoproduct (F.obj ∘ Discrete.mk) - ι := Discrete.natTrans fun a => finiteCoproduct.ι (F.obj ∘ Discrete.mk) a.as +abbrev finiteCoproduct.cofan : Limits.Cofan X := + Cofan.mk (finiteCoproduct X) (finiteCoproduct.ι X) /-- The explicit finite coproduct cocone is a colimit cocone. -/ -@[simps] -def finiteCoproduct.isColimit (F : Discrete α ⥤ Stonean) : - IsColimit (finiteCoproduct.cocone F) where - desc := fun s => finiteCoproduct.desc _ fun a => s.ι.app ⟨a⟩ - fac := fun s ⟨a⟩ => finiteCoproduct.ι_desc _ _ _ - uniq := fun s m hm => finiteCoproduct.hom_ext _ _ _ fun a => by - specialize hm ⟨a⟩ - ext t - apply_fun (fun q => q t) at hm - exact hm - -/-- The category of extremally disconnected spaces has finite coproducts. --/ -instance hasFiniteCoproducts : HasFiniteCoproducts Stonean.{u} where - out _ := { - has_colimit := fun F => { - exists_colimit := ⟨{ - cocone := finiteCoproduct.cocone F - isColimit := finiteCoproduct.isColimit F }⟩ } } +def finiteCoproduct.isColimit : Limits.IsColimit (finiteCoproduct.cofan X) := + mkCofanColimit _ + (fun s ↦ desc _ fun a ↦ s.inj a) + (fun s a ↦ ι_desc _ _ _) + fun s m hm ↦ finiteCoproduct.hom_ext _ _ _ fun a ↦ + (by ext t; exact congrFun (congrArg DFunLike.coe (hm a)) t) -/-- -A coproduct cocone associated to the explicit finite coproduct with cone point `finiteCoproduct X`. --/ -@[simps] -def finiteCoproduct.explicitCocone : Limits.Cocone (Discrete.functor X) where - pt := finiteCoproduct X - ι := Discrete.natTrans fun ⟨a⟩ => finiteCoproduct.ι X a +instance (n : ℕ) (F : Discrete (Fin n) ⥤ Stonean) : + HasColimit (Discrete.functor (F.obj ∘ Discrete.mk) : Discrete (Fin n) ⥤ Stonean) where + exists_colimit := ⟨⟨finiteCoproduct.cofan _, finiteCoproduct.isColimit _⟩⟩ -/-- -The more explicit finite coproduct cocone is a colimit cocone. --/ -@[simps] -def finiteCoproduct.isColimit' : Limits.IsColimit (finiteCoproduct.explicitCocone X) where - desc := fun s => finiteCoproduct.desc _ fun a => s.ι.app ⟨a⟩ - fac := fun s ⟨a⟩ => finiteCoproduct.ι_desc _ _ _ - uniq := fun s m hm => finiteCoproduct.hom_ext _ _ _ fun a => by - specialize hm ⟨a⟩ - ext t - apply_fun (fun q => q t) at hm - exact hm +instance : HasFiniteCoproducts Stonean where + out _ := { has_colimit := fun _ ↦ hasColimitOfIso Discrete.natIsoFunctor } /-- The isomorphism from the explicit finite coproducts to the abstract coproduct. -/ noncomputable def coproductIsoCoproduct : finiteCoproduct X ≅ ∐ X := Limits.IsColimit.coconePointUniqueUpToIso - (finiteCoproduct.isColimit' X) (Limits.colimit.isColimit _) + (finiteCoproduct.isColimit X) (Limits.colimit.isColimit _) /-- The inclusion maps into the explicit finite coproduct are open embeddings. -/ lemma finiteCoproduct.openEmbedding_ι {α : Type} [Finite α] (Z : α → Stonean.{u}) (a : α) : @@ -139,8 +109,7 @@ lemma Sigma.openEmbedding_ι {α : Type} [Finite α] (Z : α → Stonean.{u}) (a convert finiteCoproduct.openEmbedding_ι Z a ext x change ((Sigma.ι Z a) ≫ (coproductIsoCoproduct Z).inv) x = _ - simp only [coproductIsoCoproduct, colimit.comp_coconePointUniqueUpToIso_inv, - finiteCoproduct.explicitCocone_pt, finiteCoproduct.explicitCocone_ι, Discrete.natTrans_app] + simp [coproductIsoCoproduct] instance : PreservesFiniteCoproducts Stonean.toCompHaus := by refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ From 294ff6a495e19ae5a32188033d6e487ca42aded5 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Tue, 7 May 2024 22:42:43 +0000 Subject: [PATCH 23/52] perf(BundledCats): more explicit universe annotations (#12741) Testing for leanprover/lean4#4085 exposed possible regressions typeclass synthesis which can be ameliorated using explicit universes. We do this and we add references to the issue leanprover-community/mathlib4#12737 tracking these changes. --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 2 + Mathlib/AlgebraicGeometry/FunctionField.lean | 2 + .../GammaSpecAdjunction.lean | 2 + .../Morphisms/RingHomProperties.lean | 6 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 86 ++++++++++--------- Mathlib/AlgebraicGeometry/Properties.lean | 4 +- Mathlib/AlgebraicGeometry/Restrict.lean | 75 ++++++++-------- Mathlib/AlgebraicGeometry/Scheme.lean | 7 +- Mathlib/AlgebraicGeometry/Spec.lean | 3 + .../RingedSpace/LocallyRingedSpace.lean | 72 ++++++++-------- 10 files changed, 143 insertions(+), 116 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 41a0a48d36758..07c578db76020 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -32,6 +32,8 @@ 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 linter.uppercaseLean3 false noncomputable section diff --git a/Mathlib/AlgebraicGeometry/FunctionField.lean b/Mathlib/AlgebraicGeometry/FunctionField.lean index fe71ea653ff30..6a2b7d31d81d7 100644 --- a/Mathlib/AlgebraicGeometry/FunctionField.lean +++ b/Mathlib/AlgebraicGeometry/FunctionField.lean @@ -19,6 +19,8 @@ This is a field when the scheme is integral. field. This map is injective. -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + set_option linter.uppercaseLean3 false universe u v diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index ce9ac08db3925..23558e2471aa4 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -32,6 +32,8 @@ case the unit and the counit would switch to each other. -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + set_option linter.uppercaseLean3 false noncomputable section diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 18e434329b42b..6b5eb8f6e83fd 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -34,6 +34,8 @@ Further more, these properties are stable under compositions (resp. base change) -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + universe u open CategoryTheory Opposite TopologicalSpace CategoryTheory.Limits AlgebraicGeometry @@ -68,7 +70,7 @@ theorem RespectsIso.basicOpen_iff (hP : RespectsIso @P) {X Y : Scheme.{u}} [IsAf congr 1 #align ring_hom.respects_iso.basic_open_iff RingHom.RespectsIso.basicOpen_iff -theorem RespectsIso.basicOpen_iff_localization (hP : RespectsIso @P) {X Y : Scheme} [IsAffine X] +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 ?_ @@ -134,7 +136,7 @@ def sourceAffineLocally : AffineTargetMorphismProperty := fun X _ f _ => /-- 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 := +abbrev affineLocally : MorphismProperty Scheme.{u} := targetAffineLocally (sourceAffineLocally @P) #align algebraic_geometry.affine_locally AlgebraicGeometry.affineLocally diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index b8ad8ecb9f967..3b3b601890457 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -14,6 +14,8 @@ import Mathlib.CategoryTheory.Limits.Shapes.CommSq -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + set_option linter.uppercaseLean3 false noncomputable section @@ -31,11 +33,11 @@ variable {C : Type u} [Category.{v} C] /-- A morphism of Schemes is an open immersion if it is an open immersion as a morphism of LocallyRingedSpaces -/ -abbrev IsOpenImmersion {X Y : Scheme} (f : X ⟶ Y) : Prop := +abbrev IsOpenImmersion {X Y : Scheme.{u}} (f : X ⟶ Y) : Prop := LocallyRingedSpace.IsOpenImmersion f #align algebraic_geometry.IsOpenImmersion AlgebraicGeometry.IsOpenImmersion -instance IsOpenImmersion.comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) +instance IsOpenImmersion.comp {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) [IsOpenImmersion f] [IsOpenImmersion g] : IsOpenImmersion (f ≫ g) := LocallyRingedSpace.IsOpenImmersion.comp f g @@ -43,7 +45,7 @@ namespace LocallyRingedSpace.IsOpenImmersion /-- To show that a locally ringed space is a scheme, it suffices to show that it has a jointly surjective family of open immersions from affine schemes. -/ -protected def scheme (X : LocallyRingedSpace) +protected def scheme (X : LocallyRingedSpace.{u}) (h : ∀ x : X, ∃ (R : CommRingCat) (f : Spec.toLocallyRingedSpace.obj (op R) ⟶ X), @@ -63,7 +65,7 @@ protected def scheme (X : LocallyRingedSpace) end LocallyRingedSpace.IsOpenImmersion -theorem IsOpenImmersion.isOpen_range {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImmersion f] : +theorem IsOpenImmersion.isOpen_range {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] : IsOpen (Set.range f.1.base) := H.base_open.isOpen_range #align algebraic_geometry.IsOpenImmersion.open_range AlgebraicGeometry.IsOpenImmersion.isOpen_range @@ -105,7 +107,7 @@ variable {X Y Z : Scheme.{u}} (𝒰 : OpenCover X) (f : X ⟶ Z) (g : Y ⟶ Z) variable [∀ x, HasPullback (𝒰.map x ≫ f) g] /-- The affine cover of a scheme. -/ -def affineCover (X : Scheme) : OpenCover X where +def affineCover (X : Scheme.{u}) : OpenCover X where J := X.carrier obj x := Spec.obj <| Opposite.op (X.local_affine x).choose_spec.choose map x := @@ -166,7 +168,7 @@ def openCoverOfIsIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] : OpenCover Y wh /-- We construct an open cover from another, by providing the needed fields and showing that the provided fields are isomorphic with the original open cover. -/ @[simps J obj map] -def OpenCover.copy {X : Scheme} (𝒰 : OpenCover X) (J : Type*) (obj : J → Scheme) +def OpenCover.copy {X : Scheme.{u}} (𝒰 : OpenCover X) (J : Type*) (obj : J → Scheme) (map : ∀ i, obj i ⟶ X) (e₁ : J ≃ 𝒰.J) (e₂ : ∀ i, obj i ≅ 𝒰.obj (e₁ i)) (e₂ : ∀ i, map i = (e₂ i).hom ≫ 𝒰.map (e₁ i)) : OpenCover X := { J, obj, map @@ -193,7 +195,7 @@ def OpenCover.pushforwardIso {X Y : Scheme.{u}} (𝒰 : OpenCover.{v} X) (f : X /-- Adding an open immersion into an open cover gives another open cover. -/ @[simps] -def OpenCover.add {X : Scheme} (𝒰 : X.OpenCover) {Y : Scheme} (f : Y ⟶ X) [IsOpenImmersion f] : +def OpenCover.add {X Y : Scheme.{u}} (𝒰 : X.OpenCover) (f : Y ⟶ X) [IsOpenImmersion f] : X.OpenCover where J := Option 𝒰.J obj i := Option.rec Y 𝒰.obj i @@ -208,11 +210,11 @@ def OpenCover.add {X : Scheme} (𝒰 : X.OpenCover) {Y : Scheme} (f : Y ⟶ X) [ -- about open immersion (which in turn needs the open cover API). -- attribute [local reducible] CommRingCat.of CommRingCat.ofHom -instance val_base_isIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : IsIso f.1.base := +instance val_base_isIso {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.1.base := Scheme.forgetToTop.map_isIso f #align algebraic_geometry.Scheme.val_base_is_iso AlgebraicGeometry.Scheme.val_base_isIso -instance basic_open_isOpenImmersion {R : CommRingCat} (f : R) : +instance basic_open_isOpenImmersion {R : CommRingCat.{u}} (f : R) : AlgebraicGeometry.IsOpenImmersion (Scheme.Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away f))).op) := by apply SheafedSpace.IsOpenImmersion.of_stalk_iso (H := ?_) @@ -222,7 +224,7 @@ instance basic_open_isOpenImmersion {R : CommRingCat} (f : R) : #align algebraic_geometry.Scheme.basic_open_IsOpenImmersion AlgebraicGeometry.Scheme.basic_open_isOpenImmersion /-- The basic open sets form an affine open cover of `Spec R`. -/ -def affineBasisCoverOfAffine (R : CommRingCat) : OpenCover (Spec.obj (Opposite.op R)) where +def affineBasisCoverOfAffine (R : CommRingCat.{u}) : OpenCover (Spec.obj (Opposite.op R)) where J := R obj r := Spec.obj (Opposite.op <| CommRingCat.of <| Localization.Away r) map r := Spec.map (Quiver.Hom.op (algebraMap R (Localization.Away r) : _)) @@ -239,21 +241,21 @@ def affineBasisCoverOfAffine (R : CommRingCat) : OpenCover (Spec.obj (Opposite.o /-- We may bind the basic open sets of an open affine cover to form an affine cover that is also a basis. -/ -def affineBasisCover (X : Scheme) : OpenCover X := +def affineBasisCover (X : Scheme.{u}) : OpenCover X := X.affineCover.bind fun _ => affineBasisCoverOfAffine _ #align algebraic_geometry.Scheme.affine_basis_cover AlgebraicGeometry.Scheme.affineBasisCover /-- The coordinate ring of a component in the `affine_basis_cover`. -/ -def affineBasisCoverRing (X : Scheme) (i : X.affineBasisCover.J) : CommRingCat := +def affineBasisCoverRing (X : Scheme.{u}) (i : X.affineBasisCover.J) : CommRingCat := CommRingCat.of <| @Localization.Away (X.local_affine i.1).choose_spec.choose _ i.2 #align algebraic_geometry.Scheme.affine_basis_cover_ring AlgebraicGeometry.Scheme.affineBasisCoverRing -theorem affineBasisCover_obj (X : Scheme) (i : X.affineBasisCover.J) : +theorem affineBasisCover_obj (X : Scheme.{u}) (i : X.affineBasisCover.J) : X.affineBasisCover.obj i = Spec.obj (op <| X.affineBasisCoverRing i) := rfl #align algebraic_geometry.Scheme.affine_basis_cover_obj AlgebraicGeometry.Scheme.affineBasisCover_obj -theorem affineBasisCover_map_range (X : Scheme) (x : X) +theorem affineBasisCover_map_range (X : Scheme.{u}) (x : X) (r : (X.local_affine x).choose_spec.choose) : Set.range (X.affineBasisCover.map ⟨x, r⟩).1.base = (X.affineCover.map x).1.base '' (PrimeSpectrum.basicOpen r).1 := by @@ -263,7 +265,7 @@ theorem affineBasisCover_map_range (X : Scheme) (x : X) exact (PrimeSpectrum.localization_away_comap_range (Localization.Away r) r : _) #align algebraic_geometry.Scheme.affine_basis_cover_map_range AlgebraicGeometry.Scheme.affineBasisCover_map_range -theorem affineBasisCover_is_basis (X : Scheme) : +theorem affineBasisCover_is_basis (X : Scheme.{u}) : TopologicalSpace.IsTopologicalBasis {x : Set X | ∃ a : X.affineBasisCover.J, x = Set.range (X.affineBasisCover.map a).1.base} := by @@ -286,7 +288,7 @@ theorem affineBasisCover_is_basis (X : Scheme) : /-- Every open cover of a quasi-compact scheme can be refined into a finite subcover. -/ @[simps! obj map] -def OpenCover.finiteSubcover {X : Scheme} (𝒰 : OpenCover X) [H : CompactSpace X] : +def OpenCover.finiteSubcover {X : Scheme.{u}} (𝒰 : OpenCover X) [H : CompactSpace X] : OpenCover X := by have := @CompactSpace.elim_nhds_subcover _ _ H (fun x : X => Set.range (𝒰.map (𝒰.f x)).1.base) @@ -356,12 +358,12 @@ instance toSchemeHom_isOpenImmersion : AlgebraicGeometry.IsOpenImmersion (toSche H #align algebraic_geometry.PresheafedSpace.IsOpenImmersion.to_Scheme_hom_IsOpenImmersion AlgebraicGeometry.PresheafedSpace.IsOpenImmersionₓ.toSchemeHom_isOpenImmersionₓ -theorem scheme_eq_of_locallyRingedSpace_eq {X Y : Scheme} +theorem scheme_eq_of_locallyRingedSpace_eq {X Y : Scheme.{u}} (H : X.toLocallyRingedSpace = Y.toLocallyRingedSpace) : X = Y := by cases X; cases Y; congr #align algebraic_geometry.PresheafedSpace.IsOpenImmersion.Scheme_eq_of_LocallyRingedSpace_eq AlgebraicGeometry.PresheafedSpace.IsOpenImmersionₓ.scheme_eq_of_locallyRingedSpace_eq -theorem scheme_toScheme {X Y : Scheme} (f : X ⟶ Y) [AlgebraicGeometry.IsOpenImmersion f] : +theorem scheme_toScheme {X Y : Scheme.{u}} (f : X ⟶ Y) [AlgebraicGeometry.IsOpenImmersion f] : toScheme Y f.1 = X := by apply scheme_eq_of_locallyRingedSpace_eq exact locallyRingedSpace_toLocallyRingedSpace f @@ -373,24 +375,24 @@ end PresheafedSpace.IsOpenImmersion /-- The restriction of a Scheme along an open embedding. -/ @[simps! (config := .lemmasOnly) carrier, simps! presheaf_map presheaf_obj] -def Scheme.restrict {U : TopCat} (X : Scheme) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) : +def Scheme.restrict {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) : Scheme := { PresheafedSpace.IsOpenImmersion.toScheme X (X.toPresheafedSpace.ofRestrict h) with toPresheafedSpace := X.toPresheafedSpace.restrict h } #align algebraic_geometry.Scheme.restrict AlgebraicGeometry.Scheme.restrict lemma Scheme.restrict_toPresheafedSpace - {U : TopCat} (X : Scheme) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) : + {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) : (X.restrict h).toPresheafedSpace = X.toPresheafedSpace.restrict h := rfl /-- The canonical map from the restriction to the subspace. -/ @[simps!] -def Scheme.ofRestrict {U : TopCat} (X : Scheme) {f : U ⟶ TopCat.of X} +def Scheme.ofRestrict {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) : X.restrict h ⟶ X := X.toLocallyRingedSpace.ofRestrict h #align algebraic_geometry.Scheme.ofRestrict AlgebraicGeometry.Scheme.ofRestrict -instance IsOpenImmersion.ofRestrict {U : TopCat} (X : Scheme) {f : U ⟶ TopCat.of X} +instance IsOpenImmersion.ofRestrict {U : TopCat.{u}} (X : Scheme.{u}) {f : U ⟶ TopCat.of X} (h : OpenEmbedding f) : IsOpenImmersion (X.ofRestrict h) := show PresheafedSpace.IsOpenImmersion (X.toPresheafedSpace.ofRestrict h) by infer_instance #align algebraic_geometry.IsOpenImmersion.ofRestrict AlgebraicGeometry.IsOpenImmersion.ofRestrict @@ -405,29 +407,29 @@ instance (priority := 100) of_isIso [IsIso g] : IsOpenImmersion g := (show IsIso ((inducedFunctor _).map g) by infer_instance) #align algebraic_geometry.IsOpenImmersion.of_is_iso AlgebraicGeometry.IsOpenImmersion.of_isIso -theorem to_iso {X Y : Scheme} (f : X ⟶ Y) [h : IsOpenImmersion f] [Epi f.1.base] : IsIso f := +theorem to_iso {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsOpenImmersion f] [Epi f.1.base] : IsIso f := @isIso_of_reflects_iso _ _ _ _ _ _ f (Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) (@PresheafedSpace.IsOpenImmersion.to_iso _ _ _ _ f.1 h _) _ #align algebraic_geometry.IsOpenImmersion.to_iso AlgebraicGeometry.IsOpenImmersion.to_iso -theorem of_stalk_iso {X Y : Scheme} (f : X ⟶ Y) (hf : OpenEmbedding f.1.base) +theorem of_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) (hf : OpenEmbedding f.1.base) [∀ x, IsIso (PresheafedSpace.stalkMap f.1 x)] : IsOpenImmersion f := SheafedSpace.IsOpenImmersion.of_stalk_iso f.1 hf #align algebraic_geometry.IsOpenImmersion.of_stalk_iso AlgebraicGeometry.IsOpenImmersion.of_stalk_iso -theorem iff_stalk_iso {X Y : Scheme} (f : X ⟶ Y) : +theorem iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : IsOpenImmersion f ↔ OpenEmbedding f.1.base ∧ ∀ x, IsIso (PresheafedSpace.stalkMap f.1 x) := ⟨fun H => ⟨H.1, inferInstance⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.of_stalk_iso _ _ f h₁ h₂⟩ #align algebraic_geometry.IsOpenImmersion.iff_stalk_iso AlgebraicGeometry.IsOpenImmersion.iff_stalk_iso -theorem _root_.AlgebraicGeometry.isIso_iff_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) : +theorem _root_.AlgebraicGeometry.isIso_iff_isOpenImmersion {X Y : Scheme.{u}} (f : X ⟶ Y) : IsIso f ↔ IsOpenImmersion f ∧ Epi f.1.base := ⟨fun _ => ⟨inferInstance, inferInstance⟩, fun ⟨h₁, h₂⟩ => @IsOpenImmersion.to_iso _ _ f h₁ h₂⟩ #align algebraic_geometry.is_iso_iff_IsOpenImmersion AlgebraicGeometry.isIso_iff_isOpenImmersion -theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme} (f : X ⟶ Y) : +theorem _root_.AlgebraicGeometry.isIso_iff_stalk_iso {X Y : Scheme.{u}} (f : X ⟶ Y) : IsIso f ↔ IsIso f.1.base ∧ ∀ x, IsIso (PresheafedSpace.stalkMap f.1 x) := by rw [isIso_iff_isOpenImmersion, IsOpenImmersion.iff_stalk_iso, and_comm, ← and_assoc] refine' and_congr ⟨_, _⟩ Iff.rfl @@ -628,31 +630,31 @@ def isoOfRangeEq [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.bas #align algebraic_geometry.IsOpenImmersion.iso_of_range_eq AlgebraicGeometry.IsOpenImmersion.isoOfRangeEq @[simp, reassoc] -lemma isoOfRangeEq_hom_fac {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) +lemma isoOfRangeEq_hom_fac {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) [IsOpenImmersion f] [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.base) : (isoOfRangeEq f g e).hom ≫ g = f := lift_fac _ _ (le_of_eq e) @[simp, reassoc] -lemma isoOfRangeEq_inv_fac {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) +lemma isoOfRangeEq_inv_fac {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) [IsOpenImmersion f] [IsOpenImmersion g] (e : Set.range f.1.base = Set.range g.1.base) : (isoOfRangeEq f g e).inv ≫ f = g := lift_fac _ _ (le_of_eq e.symm) /-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ -abbrev _root_.AlgebraicGeometry.Scheme.Hom.opensFunctor {X Y : Scheme} (f : X ⟶ Y) +abbrev _root_.AlgebraicGeometry.Scheme.Hom.opensFunctor {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] : Opens X ⥤ Opens Y := H.openFunctor #align algebraic_geometry.Scheme.hom.opens_functor AlgebraicGeometry.Scheme.Hom.opensFunctor /-- The isomorphism `Γ(X, U) ⟶ Γ(Y, f(U))` induced by an open immersion `f : X ⟶ Y`. -/ -def _root_.AlgebraicGeometry.Scheme.Hom.invApp {X Y : Scheme} (f : X ⟶ Y) +def _root_.AlgebraicGeometry.Scheme.Hom.invApp {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] (U) : X.presheaf.obj (op U) ⟶ Y.presheaf.obj (op (f.opensFunctor.obj U)) := H.invApp U #align algebraic_geometry.Scheme.hom.inv_app AlgebraicGeometry.Scheme.Hom.invApp -theorem app_eq_inv_app_app_of_comp_eq_aux {X Y U : Scheme} (f : Y ⟶ U) (g : U ⟶ X) (fg : Y ⟶ X) +theorem app_eq_inv_app_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) : (Opens.map f.1.base).obj V = (Opens.map fg.1.base).obj (g.opensFunctor.obj V) := by subst H @@ -663,7 +665,7 @@ theorem app_eq_inv_app_app_of_comp_eq_aux {X Y U : Scheme} (f : Y ⟶ U) (g : U #align algebraic_geometry.IsOpenImmersion.app_eq_inv_app_app_of_comp_eq_aux AlgebraicGeometry.IsOpenImmersion.app_eq_inv_app_app_of_comp_eq_aux /-- The `fg` argument is to avoid nasty stuff about dependent types. -/ -theorem app_eq_invApp_app_of_comp_eq {X Y U : Scheme} (f : Y ⟶ U) (g : U ⟶ X) (fg : Y ⟶ X) +theorem app_eq_invApp_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) : f.1.c.app (op V) = Scheme.Hom.invApp g _ ≫ @@ -678,7 +680,7 @@ theorem app_eq_invApp_app_of_comp_eq {X Y U : Scheme} (f : Y ⟶ U) (g : U ⟶ X convert Y.presheaf.map_id _ #align algebraic_geometry.IsOpenImmersion.app_eq_inv_app_app_of_comp_eq AlgebraicGeometry.IsOpenImmersion.app_eq_invApp_app_of_comp_eq -theorem lift_app {X Y U : Scheme} (f : U ⟶ Y) (g : X ⟶ Y) [IsOpenImmersion f] (H) +theorem lift_app {X Y U : Scheme.{u}} (f : U ⟶ Y) (g : X ⟶ Y) [IsOpenImmersion f] (H) (V : Opens U) : (IsOpenImmersion.lift f g H).1.c.app (op V) = Scheme.Hom.invApp f _ ≫ @@ -695,7 +697,7 @@ end IsOpenImmersion namespace Scheme -theorem image_basicOpen {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImmersion f] {U : Opens X} +theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] {U : Opens X} (r : X.presheaf.obj (op U)) : f.opensFunctor.obj (X.basicOpen r) = Y.basicOpen (Scheme.Hom.invApp f U r) := by have e := Scheme.preimage_basicOpen f (Scheme.Hom.invApp f U r) @@ -717,7 +719,7 @@ theorem image_basicOpen {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImmersion f] {U /-- The image of an open immersion as an open set. -/ @[simps] -def Hom.opensRange {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImmersion f] : Opens Y := +def Hom.opensRange {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] : Opens Y := ⟨_, H.base_open.isOpen_range⟩ #align algebraic_geometry.Scheme.hom.opens_range AlgebraicGeometry.Scheme.Hom.opensRange @@ -726,7 +728,7 @@ end Scheme /-- Given an open cover on `X`, we may pull them back along a morphism `W ⟶ X` to obtain an open cover of `W`. -/ @[simps] -def Scheme.OpenCover.pullbackCover {X : Scheme} (𝒰 : X.OpenCover) {W : Scheme} (f : W ⟶ X) : +def Scheme.OpenCover.pullbackCover {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ X) : W.OpenCover where J := 𝒰.J obj x := pullback f (𝒰.map x) @@ -749,7 +751,7 @@ def Scheme.OpenCover.pullbackCover {X : Scheme} (𝒰 : X.OpenCover) {W : Scheme an open cover of `W`. This is similar to `Scheme.OpenCover.pullbackCover`, but here we take `pullback (𝒰.map x) f` instead of `pullback f (𝒰.map x)`. -/ @[simps] -def Scheme.OpenCover.pullbackCover' {X : Scheme} (𝒰 : X.OpenCover) {W : Scheme} (f : W ⟶ X) : +def Scheme.OpenCover.pullbackCover' {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ X) : W.OpenCover where J := 𝒰.J obj x := pullback (𝒰.map x) f @@ -767,7 +769,7 @@ def Scheme.OpenCover.pullbackCover' {X : Scheme} (𝒰 : X.OpenCover) {W : Schem exact ⟨y, h⟩ · rw [← TopCat.epi_iff_surjective]; infer_instance -theorem Scheme.OpenCover.iUnion_range {X : Scheme} (𝒰 : X.OpenCover) : +theorem Scheme.OpenCover.iUnion_range {X : Scheme.{u}} (𝒰 : X.OpenCover) : ⋃ i, Set.range (𝒰.map i).1.base = Set.univ := by rw [Set.eq_univ_iff_forall] intro x @@ -775,12 +777,12 @@ theorem Scheme.OpenCover.iUnion_range {X : Scheme} (𝒰 : X.OpenCover) : exact ⟨𝒰.f x, 𝒰.Covers x⟩ #align algebraic_geometry.Scheme.open_cover.Union_range AlgebraicGeometry.Scheme.OpenCover.iUnion_range -theorem Scheme.OpenCover.iSup_opensRange {X : Scheme} (𝒰 : X.OpenCover) : +theorem Scheme.OpenCover.iSup_opensRange {X : Scheme.{u}} (𝒰 : X.OpenCover) : ⨆ i, Scheme.Hom.opensRange (𝒰.map i) = ⊤ := Opens.ext <| by rw [Opens.coe_iSup]; exact 𝒰.iUnion_range #align algebraic_geometry.Scheme.open_cover.supr_opens_range AlgebraicGeometry.Scheme.OpenCover.iSup_opensRange -theorem Scheme.OpenCover.compactSpace {X : Scheme} (𝒰 : X.OpenCover) [Finite 𝒰.J] +theorem Scheme.OpenCover.compactSpace {X : Scheme.{u}} (𝒰 : X.OpenCover) [Finite 𝒰.J] [H : ∀ i, CompactSpace (𝒰.obj i)] : CompactSpace X := by cases nonempty_fintype 𝒰.J rw [← isCompact_univ_iff, ← 𝒰.iUnion_range] @@ -812,7 +814,7 @@ def Scheme.OpenCover.inter {X : Scheme.{u}} (𝒰₁ : Scheme.OpenCover.{v₁} X /-- 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 : s → Opens X) +def Scheme.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 diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index 2f9fa60070005..d40b69c1a2e3f 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -23,8 +23,10 @@ We provide some basic properties of schemes are reduced. -/ -universe u +-- Explicit universe annotations were used in this file to improve perfomance #12737 + +universe u open TopologicalSpace Opposite CategoryTheory CategoryTheory.Limits TopCat diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 11aecab1ec778..f16df212bbb70 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -17,6 +17,8 @@ import Mathlib.AlgebraicGeometry.OpenImmersion -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + set_option linter.uppercaseLean3 false noncomputable section @@ -27,13 +29,13 @@ open CategoryTheory.Limits namespace AlgebraicGeometry -universe v v₁ v₂ u +universe v v₁ v₂ u u₁ -variable {C : Type u} [Category.{v} C] +variable {C : Type u₁} [Category.{v} C] section -variable (X : Scheme) +variable (X : Scheme.{u}) /-- `f ⁻¹ᵁ U` is notation for `(Opens.map f.1.base).obj U`, the preimage of an open set `U` under `f`. -/ @@ -44,22 +46,22 @@ notation3:90 f:91 "⁻¹ᵁ " U:90 => (Opens.map (f : LocallyRingedSpace.Hom _ _ notation3:60 X:60 " ∣_ᵤ " U:61 => Scheme.restrict X (U : Opens X).openEmbedding /-- The restriction of a scheme to an open subset. -/ -abbrev Scheme.ιOpens {X : Scheme} (U : Opens X.carrier) : X ∣_ᵤ U ⟶ X := X.ofRestrict _ +abbrev Scheme.ιOpens {X : Scheme.{u}} (U : Opens X.carrier) : X ∣_ᵤ U ⟶ X := X.ofRestrict _ -lemma Scheme.ofRestrict_val_c_app_self {X : Scheme} (U : Opens X) : +lemma Scheme.ofRestrict_val_c_app_self {X : Scheme.{u}} (U : Opens X) : (X.ofRestrict U.openEmbedding).1.c.app (op U) = X.presheaf.map (eqToHom (by simp)).op := rfl -lemma Scheme.eq_restrict_presheaf_map_eqToHom {X : Scheme} (U : Opens X) {V W : Opens U} +lemma Scheme.eq_restrict_presheaf_map_eqToHom {X : Scheme.{u}} (U : Opens X) {V W : Opens U} (e : U.openEmbedding.isOpenMap.functor.obj V = U.openEmbedding.isOpenMap.functor.obj W) : X.presheaf.map (eqToHom e).op = (X ∣_ᵤ U).presheaf.map (eqToHom <| U.openEmbedding.functor_obj_injective e).op := rfl -instance ΓRestrictAlgebra {X : Scheme} {Y : TopCat} {f : Y ⟶ X} (hf : OpenEmbedding f) : +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 #align algebraic_geometry.Γ_restrict_algebra AlgebraicGeometry.ΓRestrictAlgebra -lemma Scheme.map_basicOpen' (X : Scheme) (U : Opens X) (r : Scheme.Γ.obj (op <| X ∣_ᵤ U)) : +lemma Scheme.map_basicOpen' (X : Scheme.{u}) (U : Opens X) (r : Scheme.Γ.obj (op <| X ∣_ᵤ U)) : U.openEmbedding.isOpenMap.functor.obj ((X ∣_ᵤ U).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 _ @@ -69,11 +71,11 @@ lemma Scheme.map_basicOpen' (X : Scheme) (U : Opens X) (r : Scheme.Γ.obj (op <| congr exact PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ _ _ -lemma Scheme.map_basicOpen (X : Scheme) (U : Opens X) (r : Scheme.Γ.obj (op <| X ∣_ᵤ U)) : +lemma Scheme.map_basicOpen (X : Scheme.{u}) (U : Opens X) (r : Scheme.Γ.obj (op <| X ∣_ᵤ U)) : U.openEmbedding.isOpenMap.functor.obj ((X ∣_ᵤ U).basicOpen r) = X.basicOpen r := by rw [Scheme.map_basicOpen', Scheme.basicOpen_res_eq] -lemma Scheme.map_basicOpen_map (X : Scheme) (U : Opens X) (r : X.presheaf.obj (op U)) : +lemma Scheme.map_basicOpen_map (X : Scheme.{u}) (U : Opens X) (r : X.presheaf.obj (op U)) : U.openEmbedding.isOpenMap.functor.obj ((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] @@ -176,7 +178,7 @@ 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 V : Opens X.carrier) : +def Scheme.restrictRestrictComm (X : Scheme.{u}) (U V : Opens X.carrier) : 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, @@ -188,7 +190,7 @@ def Scheme.restrictRestrictComm (X : Scheme) (U V : Opens X.carrier) : /-- If `V` is an open subset of `U`, then `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V`. -/ noncomputable -def Scheme.restrictRestrict (X : Scheme) (U : Opens X.carrier) (V : Opens (X ∣_ᵤ U).carrier) : +def Scheme.restrictRestrict (X : Scheme.{u}) (U : Opens X.carrier) (V : Opens (X ∣_ᵤ U).carrier) : X ∣_ᵤ U ∣_ᵤ V ≅ X ∣_ᵤ U.openEmbedding.isOpenMap.functor.obj V := by refine IsOpenImmersion.isoOfRangeEq (ιOpens _ ≫ ιOpens U) (ιOpens _) ?_ simp only [Scheme.restrict_carrier, Scheme.ofRestrict_val_base, Scheme.comp_coeBase, @@ -197,27 +199,27 @@ def Scheme.restrictRestrict (X : Scheme) (U : Opens X.carrier) (V : Opens (X ∣ rfl @[simp, reassoc] -lemma Scheme.restrictRestrict_hom_restrict (X : Scheme) (U : Opens X.carrier) +lemma Scheme.restrictRestrict_hom_restrict (X : Scheme.{u}) (U : Opens X.carrier) (V : Opens (X ∣_ᵤ U).carrier) : (X.restrictRestrict U V).hom ≫ ιOpens _ = ιOpens V ≫ ιOpens U := IsOpenImmersion.isoOfRangeEq_hom_fac _ _ _ @[simp, reassoc] -lemma Scheme.restrictRestrict_inv_restrict_restrict (X : Scheme) (U : Opens X.carrier) +lemma Scheme.restrictRestrict_inv_restrict_restrict (X : Scheme.{u}) (U : Opens X.carrier) (V : Opens (X ∣_ᵤ U).carrier) : (X.restrictRestrict U V).inv ≫ ιOpens V ≫ ιOpens U = ιOpens _ := IsOpenImmersion.isoOfRangeEq_inv_fac _ _ _ /-- If `U = V`, then `X ∣_ U` is isomorphic to `X ∣_ V`. -/ noncomputable -def Scheme.restrictIsoOfEq (X : Scheme) {U V : Opens X.carrier} (e : U = V) : +def Scheme.restrictIsoOfEq (X : Scheme.{u}) {U V : Opens X.carrier} (e : U = V) : X ∣_ᵤ U ≅ X ∣_ᵤ V := by exact IsOpenImmersion.isoOfRangeEq (ιOpens U) (ιOpens V) (by rw [e]) end /-- The restriction of an isomorphism onto an open set. -/ -noncomputable abbrev Scheme.restrictMapIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] +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) (H := PresheafedSpace.IsOpenImmersion.comp (hf := inferInstance) (hg := inferInstance)) @@ -232,7 +234,7 @@ noncomputable abbrev Scheme.restrictMapIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f 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} (f : X ⟶ Y) (U : Opens Y) : +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 (X.ofRestrict _) _ rw [IsOpenImmersion.range_pullback_fst_of_right] @@ -242,19 +244,19 @@ def pullbackRestrictIsoRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : #align algebraic_geometry.pullback_restrict_iso_restrict AlgebraicGeometry.pullbackRestrictIsoRestrict @[simp, reassoc] -theorem pullbackRestrictIsoRestrict_inv_fst {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : +theorem pullbackRestrictIsoRestrict_inv_fst {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : (pullbackRestrictIsoRestrict f U).inv ≫ pullback.fst = X.ofRestrict _ := by delta pullbackRestrictIsoRestrict; simp #align algebraic_geometry.pullback_restrict_iso_restrict_inv_fst AlgebraicGeometry.pullbackRestrictIsoRestrict_inv_fst @[simp, reassoc] -theorem pullbackRestrictIsoRestrict_hom_restrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : +theorem pullbackRestrictIsoRestrict_hom_restrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : (pullbackRestrictIsoRestrict f U).hom ≫ Scheme.ιOpens (f ⁻¹ᵁ U) = pullback.fst := by delta pullbackRestrictIsoRestrict; simp #align algebraic_geometry.pullback_restrict_iso_restrict_hom_restrict AlgebraicGeometry.pullbackRestrictIsoRestrict_hom_restrict /-- The restriction of a morphism `X ⟶ Y` onto `X |_{f ⁻¹ U} ⟶ Y |_ U`. -/ -def morphismRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : X ∣_ᵤ f ⁻¹ᵁ U ⟶ Y ∣_ᵤ U := +def morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : X ∣_ᵤ f ⁻¹ᵁ U ⟶ Y ∣_ᵤ U := (pullbackRestrictIsoRestrict f U).inv ≫ pullback.snd #align algebraic_geometry.morphism_restrict AlgebraicGeometry.morphismRestrict @@ -262,19 +264,19 @@ def morphismRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : X ∣_ᵤ f infixl:85 " ∣_ " => morphismRestrict @[simp, reassoc] -theorem pullbackRestrictIsoRestrict_hom_morphismRestrict {X Y : Scheme} (f : X ⟶ Y) +theorem pullbackRestrictIsoRestrict_hom_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : (pullbackRestrictIsoRestrict f U).hom ≫ f ∣_ U = pullback.snd := Iso.hom_inv_id_assoc _ _ #align algebraic_geometry.pullback_restrict_iso_restrict_hom_morphism_restrict AlgebraicGeometry.pullbackRestrictIsoRestrict_hom_morphismRestrict @[simp, reassoc] -theorem morphismRestrict_ι {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : +theorem morphismRestrict_ι {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : (f ∣_ U) ≫ Scheme.ιOpens U = Scheme.ιOpens (f ⁻¹ᵁ U) ≫ f := by delta morphismRestrict rw [Category.assoc, pullback.condition.symm, pullbackRestrictIsoRestrict_inv_fst_assoc] #align algebraic_geometry.morphism_restrict_ι AlgebraicGeometry.morphismRestrict_ι -theorem isPullback_morphismRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : +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 delta morphismRestrict rw [← Category.id_comp f] @@ -285,7 +287,7 @@ theorem isPullback_morphismRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : erw [pullbackRestrictIsoRestrict_inv_fst]; rw [Category.comp_id] #align algebraic_geometry.is_pullback_morphism_restrict AlgebraicGeometry.isPullback_morphismRestrict -theorem morphismRestrict_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U : Opens Z) : +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 rw [← pullbackRightPullbackFstIso_inv_snd_snd] @@ -298,21 +300,21 @@ theorem morphismRestrict_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U : pullbackRestrictIsoRestrict_inv_fst_assoc] #align algebraic_geometry.morphism_restrict_comp AlgebraicGeometry.morphismRestrict_comp -instance {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U : Opens Y) : IsIso (f ∣_ U) := by +instance {X Y : Scheme.{u}} (f : X ⟶ Y) [IsIso f] (U : Opens Y) : IsIso (f ∣_ U) := by delta morphismRestrict; infer_instance -theorem morphismRestrict_base_coe {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (x) : +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 := congr_arg (fun f => PresheafedSpace.Hom.base (LocallyRingedSpace.Hom.val f) x) (morphismRestrict_ι f U) #align algebraic_geometry.morphism_restrict_base_coe AlgebraicGeometry.morphismRestrict_base_coe -theorem morphismRestrict_val_base {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : +theorem morphismRestrict_val_base {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : ⇑(f ∣_ U).1.base = U.1.restrictPreimage f.1.base := funext fun x => Subtype.ext (morphismRestrict_base_coe f U x) #align algebraic_geometry.morphism_restrict_val_base AlgebraicGeometry.morphismRestrict_val_base -theorem image_morphismRestrict_preimage {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : +theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : (f ⁻¹ᵁ U).openEmbedding.isOpenMap.functor.obj ((f ∣_ U) ⁻¹ᵁ V) = f ⁻¹ᵁ (U.openEmbedding.isOpenMap.functor.obj V) := by ext1 @@ -334,7 +336,7 @@ theorem image_morphismRestrict_preimage {X Y : Scheme} (f : X ⟶ Y) (U : Opens exact morphismRestrict_base_coe f U ⟨x, hx⟩ #align algebraic_geometry.image_morphism_restrict_preimage AlgebraicGeometry.image_morphismRestrict_preimage -theorem morphismRestrict_c_app {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : +theorem morphismRestrict_c_app {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens U) : (f ∣_ U).1.c.app (op V) = f.1.c.app (op (U.openEmbedding.isOpenMap.functor.obj V)) ≫ X.presheaf.map (eqToHom (image_morphismRestrict_preimage f U V)).op := by @@ -354,7 +356,7 @@ theorem morphismRestrict_c_app {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : O congr 1 #align algebraic_geometry.morphism_restrict_c_app AlgebraicGeometry.morphismRestrict_c_app -theorem Γ_map_morphismRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : +theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) : Scheme.Γ.map (f ∣_ U).op = Y.presheaf.map (eqToHom <| U.openEmbedding_obj_top.symm).op ≫ f.1.c.app (op U) ≫ X.presheaf.map (eqToHom <| (f ⁻¹ᵁ U).openEmbedding_obj_top).op := by @@ -365,7 +367,8 @@ theorem Γ_map_morphismRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) : /-- Restricting a morphism onto the image of an open immersion is isomorphic to the base change along the immersion. -/ -def morphismRestrictOpensRange {X Y U : Scheme} (f : X ⟶ Y) (g : U ⟶ Y) [hg : IsOpenImmersion g] : +def morphismRestrictOpensRange + {X Y U : Scheme.{u}} (f : X ⟶ Y) (g : U ⟶ Y) [hg : IsOpenImmersion g] : Arrow.mk (f ∣_ Scheme.Hom.opensRange g) ≅ Arrow.mk (pullback.snd : pullback f g ⟶ _) := by let V : Opens Y := Scheme.Hom.opensRange g let e := @@ -383,13 +386,13 @@ def morphismRestrictOpensRange {X Y U : Scheme} (f : X ⟶ Y) (g : U ⟶ Y) [hg /-- 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} (f : X ⟶ Y) {U V : Opens Y} (e : U = V) : +def morphismRestrictEq {X Y : Scheme.{u}} (f : X ⟶ Y) {U V : Opens Y} (e : U = V) : Arrow.mk (f ∣_ U) ≅ Arrow.mk (f ∣_ V) := eqToIso (by subst e; rfl) #align algebraic_geometry.morphism_restrict_eq AlgebraicGeometry.morphismRestrictEq /-- Restricting a morphism twice is isomorphic to one restriction. -/ -def morphismRestrictRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : Opens (Y ∣_ᵤ U)) : +def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (V : Opens (Y ∣_ᵤ U)) : Arrow.mk (f ∣_ U ∣_ V) ≅ Arrow.mk (f ∣_ U.openEmbedding.isOpenMap.functor.obj V) := by refine Arrow.isoMk' _ _ (Scheme.restrictRestrict _ _ _ ≪≫ Scheme.restrictIsoOfEq _ ?_) (Scheme.restrictRestrict _ _ _) ?_ @@ -410,7 +413,7 @@ def morphismRestrictRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : Ope #align algebraic_geometry.morphism_restrict_restrict AlgebraicGeometry.morphismRestrictRestrict /-- Restricting a morphism twice onto a basic open set is isomorphic to one restriction. -/ -def morphismRestrictRestrictBasicOpen {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) +def morphismRestrictRestrictBasicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (r : Y.presheaf.obj (op U)) : Arrow.mk (f ∣_ U ∣_ @@ -432,7 +435,7 @@ def morphismRestrictRestrictBasicOpen {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) /-- The stalk map of a restriction of a morphism is isomorphic to the stalk map of the original map. -/ -def morphismRestrictStalkMap {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (x) : +def morphismRestrictStalkMap {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) (x) : Arrow.mk (PresheafedSpace.stalkMap (f ∣_ U).1 x) ≅ Arrow.mk (PresheafedSpace.stalkMap f.1 x.1) := by fapply Arrow.isoMk' @@ -453,7 +456,7 @@ def morphismRestrictStalkMap {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (x) : #align algebraic_geometry.morphism_restrict_stalk_map AlgebraicGeometry.morphismRestrictStalkMap -instance {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) [IsOpenImmersion f] : +instance {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Opens Y) [IsOpenImmersion f] : IsOpenImmersion (f ∣_ U) := by delta morphismRestrict exact PresheafedSpace.IsOpenImmersion.comp _ _ diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 74322e4e06714..472756f673465 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -19,7 +19,12 @@ A morphism of schemes is just a morphism of the underlying locally ringed spaces -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + set_option linter.uppercaseLean3 false + +universe u + noncomputable section open TopologicalSpace @@ -245,7 +250,7 @@ def Spec : CommRingCatᵒᵖ ⥤ Scheme where /-- The empty scheme. -/ @[simps] -def empty.{u} : Scheme.{u} where +def empty : Scheme where carrier := TopCat.of PEmpty presheaf := (CategoryTheory.Functor.const _).obj (CommRingCat.of PUnit) IsSheaf := Presheaf.isSheaf_of_isTerminal _ CommRingCat.punitIsTerminal diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 500cf63dc52aa..70a5a65c987ef 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -34,6 +34,9 @@ The adjunction `Γ ⊣ Spec` is constructed in `Mathlib/AlgebraicGeometry/GammaS -/ + +-- Explicit universe annotations were used in this file to improve perfomance #12737 + noncomputable section universe u v diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index e3bccbe776aa3..dfac00c93d8ff 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -16,6 +16,10 @@ stalks are local rings), and morphisms between these (morphisms in `SheafedSpace `is_local_ring_hom` on the stalk maps). -/ +-- Explicit universe annotations were used in this file to improve perfomance #12737 + +universe u + open CategoryTheory open TopCat @@ -33,7 +37,7 @@ such that all the stalks are local rings. A morphism of locally ringed spaces is a morphism of ringed spaces such that the morphisms induced on stalks are local ring homomorphisms. -/ -structure LocallyRingedSpace extends SheafedSpace CommRingCat where +structure LocallyRingedSpace extends SheafedSpace CommRingCat.{u} where /-- Stalks of a locally ringed space are local rings. -/ localRing : ∀ x, LocalRing (presheaf.stalk x) set_option linter.uppercaseLean3 false in @@ -43,7 +47,7 @@ attribute [instance] LocallyRingedSpace.localRing namespace LocallyRingedSpace -variable (X : LocallyRingedSpace) +variable (X : LocallyRingedSpace.{u}) /-- An alias for `to_SheafedSpace`, where the result type is a `RingedSpace`. This allows us to use dot-notation for the `RingedSpace` namespace. @@ -59,7 +63,7 @@ def toTopCat : TopCat := set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.to_Top AlgebraicGeometry.LocallyRingedSpace.toTopCat -instance : CoeSort LocallyRingedSpace (Type*) := +instance : CoeSort LocallyRingedSpace (Type u) := ⟨fun X : LocallyRingedSpace => (X.toTopCat : Type _)⟩ instance (x : X) : LocalRing (X.stalk x) := @@ -76,7 +80,7 @@ set_option linter.uppercaseLean3 false in /-- A morphism of locally ringed spaces is a morphism of ringed spaces such that the morphisms induced on stalks are local ring homomorphisms. -/ @[ext] -structure Hom (X Y : LocallyRingedSpace) : Type _ where +structure Hom (X Y : LocallyRingedSpace.{u}) : Type _ where /-- the underlying morphism between ringed spaces -/ val : X.toSheafedSpace ⟶ Y.toSheafedSpace /-- the underlying morphism induces a local ring homomorphism on stalks -/ @@ -87,14 +91,14 @@ set_option linter.uppercaseLean3 false in instance : Quiver LocallyRingedSpace := ⟨Hom⟩ -@[ext] lemma Hom.ext' (X Y : LocallyRingedSpace) {f g : X ⟶ Y} (h : f.val = g.val) : f = g := +@[ext] lemma Hom.ext' (X Y : LocallyRingedSpace.{u}) {f g : X ⟶ Y} (h : f.val = g.val) : f = g := Hom.ext _ _ h -- TODO perhaps we should make a bundled `LocalRing` and return one here? -- TODO define `sheaf.stalk` so we can write `X.𝒪.stalk` here? /-- The stalk of a locally ringed space, just as a `CommRing`. -/ -noncomputable def stalk (X : LocallyRingedSpace) (x : X) : CommRingCat := +noncomputable def stalk (X : LocallyRingedSpace.{u}) (x : X) : CommRingCat := X.presheaf.stalk x set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.stalk AlgebraicGeometry.LocallyRingedSpace.stalk @@ -106,31 +110,31 @@ instance stalkLocal (x : X) : LocalRing <| X.stalk x := X.localRing x /-- A morphism of locally ringed spaces `f : X ⟶ Y` induces a local ring homomorphism from `Y.stalk (f x)` to `X.stalk x` for any `x : X`. -/ -noncomputable def stalkMap {X Y : LocallyRingedSpace} (f : X ⟶ Y) (x : X) : +noncomputable def stalkMap {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : Y.stalk (f.1.1 x) ⟶ X.stalk x := PresheafedSpace.stalkMap f.1 x set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.stalk_map AlgebraicGeometry.LocallyRingedSpace.stalkMap -instance {X Y : LocallyRingedSpace} (f : X ⟶ Y) (x : X) : IsLocalRingHom (stalkMap f x) := +instance {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : IsLocalRingHom (stalkMap f x) := f.2 x -instance {X Y : LocallyRingedSpace} (f : X ⟶ Y) (x : X) : +instance {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) (x : X) : IsLocalRingHom (PresheafedSpace.stalkMap f.1 x) := f.2 x /-- The identity morphism on a locally ringed space. -/ @[simps] -def id (X : LocallyRingedSpace) : Hom X X := +def id (X : LocallyRingedSpace.{u}) : Hom X X := ⟨𝟙 _, fun x => by erw [PresheafedSpace.stalkMap.id]; apply isLocalRingHom_id⟩ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.id AlgebraicGeometry.LocallyRingedSpace.id -instance (X : LocallyRingedSpace) : Inhabited (Hom X X) := +instance (X : LocallyRingedSpace.{u}) : Inhabited (Hom X X) := ⟨id X⟩ /-- Composition of morphisms of locally ringed spaces. -/ -def comp {X Y Z : LocallyRingedSpace} (f : Hom X Y) (g : Hom Y Z) : Hom X Z := +def comp {X Y Z : LocallyRingedSpace.{u}} (f : Hom X Y) (g : Hom Y Z) : Hom X Z := ⟨f.val ≫ g.val, fun x => by erw [PresheafedSpace.stalkMap.comp] exact @isLocalRingHom_comp _ _ _ _ _ _ _ _ (f.2 _) (g.2 _)⟩ @@ -138,7 +142,7 @@ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.comp AlgebraicGeometry.LocallyRingedSpace.comp /-- The category of locally ringed spaces. -/ -instance : Category LocallyRingedSpace where +instance : Category LocallyRingedSpace.{u} where Hom := Hom id := id comp {X Y Z} f g := comp f g @@ -148,7 +152,7 @@ instance : Category LocallyRingedSpace where /-- The forgetful functor from `LocallyRingedSpace` to `SheafedSpace CommRing`. -/ @[simps] -def forgetToSheafedSpace : LocallyRingedSpace ⥤ SheafedSpace CommRingCat where +def forgetToSheafedSpace : LocallyRingedSpace.{u} ⥤ SheafedSpace CommRingCat.{u} where obj X := X.toSheafedSpace map {X Y} f := f.1 set_option linter.uppercaseLean3 false in @@ -159,13 +163,13 @@ instance : forgetToSheafedSpace.Faithful where /-- The forgetful functor from `LocallyRingedSpace` to `Top`. -/ @[simps!] -def forgetToTop : LocallyRingedSpace ⥤ TopCat := +def forgetToTop : LocallyRingedSpace.{u} ⥤ TopCat.{u} := forgetToSheafedSpace ⋙ SheafedSpace.forget _ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.forget_to_Top AlgebraicGeometry.LocallyRingedSpace.forgetToTop @[simp] -theorem comp_val {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) : +theorem comp_val {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).val = f.val ≫ g.val := rfl set_option linter.uppercaseLean3 false in @@ -174,13 +178,13 @@ set_option linter.uppercaseLean3 false in -- Porting note: complains that `(f ≫ g).val.c` can be further simplified -- so changed to its simp normal form `(f.val ≫ g.val).c` @[simp] -theorem comp_val_c {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) : +theorem comp_val_c {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : (f.1 ≫ g.1).c = g.val.c ≫ (Presheaf.pushforward _ g.val.base).map f.val.c := rfl set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.comp_val_c AlgebraicGeometry.LocallyRingedSpace.comp_val_c -theorem comp_val_c_app {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (Opens Z)ᵒᵖ) : +theorem comp_val_c_app {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (Opens Z)ᵒᵖ) : (f ≫ g).val.c.app U = g.val.c.app U ≫ f.val.c.app (op <| (Opens.map g.val.base).obj U.unop) := rfl set_option linter.uppercaseLean3 false in @@ -192,8 +196,8 @@ spaces can be lifted to a morphism `X ⟶ Y` as locally ringed spaces. See also `iso_of_SheafedSpace_iso`. -/ @[simps] -def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace} (f : X.toSheafedSpace ⟶ Y.toSheafedSpace) - [IsIso f] : X ⟶ Y := +def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}} + (f : X.toSheafedSpace ⟶ Y.toSheafedSpace) [IsIso f] : X ⟶ Y := Hom.mk f fun x => -- Here we need to see that the stalk maps are really local ring homomorphisms. -- This can be solved by type class inference, because stalk maps of isomorphisms @@ -210,7 +214,7 @@ This is related to the property that the functor `forget_to_SheafedSpace` reflec In fact, it is slightly stronger as we do not require `f` to come from a morphism between _locally_ ringed spaces. -/ -def isoOfSheafedSpaceIso {X Y : LocallyRingedSpace} (f : X.toSheafedSpace ≅ Y.toSheafedSpace) : +def isoOfSheafedSpaceIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace ≅ Y.toSheafedSpace) : X ≅ Y where hom := homOfSheafedSpaceHomOfIsIso f.hom inv := homOfSheafedSpaceHomOfIsIso f.inv @@ -224,7 +228,7 @@ instance : forgetToSheafedSpace.ReflectsIsomorphisms where reflects {_ _} f i := ⟨homOfSheafedSpaceHomOfIsIso (CategoryTheory.inv (forgetToSheafedSpace.map f)), Hom.ext _ _ (IsIso.hom_inv_id (I := i)), Hom.ext _ _ (IsIso.inv_hom_id (I := i))⟩ } -instance is_sheafedSpace_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) [IsIso f] : IsIso f.1 := +instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.1 := LocallyRingedSpace.forgetToSheafedSpace.map_isIso f set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.is_SheafedSpace_iso AlgebraicGeometry.LocallyRingedSpace.is_sheafedSpace_iso @@ -232,7 +236,7 @@ set_option linter.uppercaseLean3 false in /-- The restriction of a locally ringed space along an open embedding. -/ @[simps!] -def restrict {U : TopCat} (X : LocallyRingedSpace) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : +def restrict {U : TopCat} (X : LocallyRingedSpace.{u}) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : LocallyRingedSpace where localRing := by intro x @@ -244,15 +248,15 @@ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.restrict AlgebraicGeometry.LocallyRingedSpace.restrict /-- The canonical map from the restriction to the subspace. -/ -def ofRestrict {U : TopCat} (X : LocallyRingedSpace) {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : - X.restrict h ⟶ X := +def ofRestrict {U : TopCat} (X : LocallyRingedSpace.{u}) + {f : U ⟶ X.toTopCat} (h : OpenEmbedding f) : X.restrict h ⟶ X := ⟨X.toPresheafedSpace.ofRestrict h, fun _ => inferInstance⟩ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.of_restrict AlgebraicGeometry.LocallyRingedSpace.ofRestrict /-- The restriction of a locally ringed space `X` to the top subspace is isomorphic to `X` itself. -/ -def restrictTopIso (X : LocallyRingedSpace) : +def restrictTopIso (X : LocallyRingedSpace.{u}) : X.restrict (Opens.openEmbedding ⊤) ≅ X := isoOfSheafedSpaceIso X.toSheafedSpace.restrictTopIso set_option linter.uppercaseLean3 false in @@ -260,7 +264,7 @@ set_option linter.uppercaseLean3 false in /-- The global sections, notated Gamma. -/ -def Γ : LocallyRingedSpaceᵒᵖ ⥤ CommRingCat := +def Γ : LocallyRingedSpace.{u}ᵒᵖ ⥤ CommRingCat.{u} := forgetToSheafedSpace.op ⋙ SheafedSpace.Γ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.Γ AlgebraicGeometry.LocallyRingedSpace.Γ @@ -271,28 +275,28 @@ set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.Γ_def AlgebraicGeometry.LocallyRingedSpace.Γ_def @[simp] -theorem Γ_obj (X : LocallyRingedSpaceᵒᵖ) : Γ.obj X = X.unop.presheaf.obj (op ⊤) := +theorem Γ_obj (X : LocallyRingedSpace.{u}ᵒᵖ) : Γ.obj X = X.unop.presheaf.obj (op ⊤) := rfl set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.Γ_obj AlgebraicGeometry.LocallyRingedSpace.Γ_obj -theorem Γ_obj_op (X : LocallyRingedSpace) : Γ.obj (op X) = X.presheaf.obj (op ⊤) := +theorem Γ_obj_op (X : LocallyRingedSpace.{u}) : Γ.obj (op X) = X.presheaf.obj (op ⊤) := rfl set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.Γ_obj_op AlgebraicGeometry.LocallyRingedSpace.Γ_obj_op @[simp] -theorem Γ_map {X Y : LocallyRingedSpaceᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.1.c.app (op ⊤) := +theorem Γ_map {X Y : LocallyRingedSpace.{u}ᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.1.c.app (op ⊤) := rfl set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.Γ_map AlgebraicGeometry.LocallyRingedSpace.Γ_map -theorem Γ_map_op {X Y : LocallyRingedSpace} (f : X ⟶ Y) : Γ.map f.op = f.1.c.app (op ⊤) := +theorem Γ_map_op {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) : Γ.map f.op = f.1.c.app (op ⊤) := rfl set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.Γ_map_op AlgebraicGeometry.LocallyRingedSpace.Γ_map_op -theorem preimage_basicOpen {X Y : LocallyRingedSpace} (f : X ⟶ Y) {U : Opens Y} +theorem preimage_basicOpen {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) {U : Opens Y} (s : Y.presheaf.obj (op U)) : (Opens.map f.1.base).obj (Y.toRingedSpace.basicOpen s) = @RingedSpace.basicOpen X.toRingedSpace ((Opens.map f.1.base).obj U) (f.1.c.app _ s) := by @@ -311,7 +315,7 @@ set_option linter.uppercaseLean3 false in -- This actually holds for all ringed spaces with nontrivial stalks. @[simp] -theorem basicOpen_zero (X : LocallyRingedSpace) (U : Opens X.carrier) : +theorem basicOpen_zero (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) : X.toRingedSpace.basicOpen (0 : X.presheaf.obj <| op U) = ⊥ := by ext x simp only [RingedSpace.basicOpen, Opens.coe_mk, Set.mem_image, Set.mem_setOf_eq, Subtype.exists, @@ -324,7 +328,7 @@ theorem basicOpen_zero (X : LocallyRingedSpace) (U : Opens X.carrier) : set_option linter.uppercaseLean3 false in #align algebraic_geometry.LocallyRingedSpace.basic_open_zero AlgebraicGeometry.LocallyRingedSpace.basicOpen_zero -instance component_nontrivial (X : LocallyRingedSpace) (U : Opens X.carrier) [hU : Nonempty U] : +instance component_nontrivial (X : LocallyRingedSpace.{u}) (U : Opens X.carrier) [hU : Nonempty U] : Nontrivial (X.presheaf.obj <| op U) := (X.presheaf.germ hU.some).domain_nontrivial set_option linter.uppercaseLean3 false in From ba9d411b0518c2de8057c3488cbfbf68ce5e5439 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 7 May 2024 23:13:53 +0000 Subject: [PATCH 24/52] feat: Positivity extension for `Finset.prod` (#9365) Followup to #9365 From LeanAPAP Co-authored-by: Eric Wieser Co-authored-by: Alex J Best --- .../Order/BigOperators/Ring/Finset.lean | 58 +++++++++++++++++++ Mathlib/Analysis/Analytic/Composition.lean | 4 +- Mathlib/Analysis/Analytic/Inverse.lean | 3 +- .../NormedSpace/BoundedLinearMaps.lean | 30 ++++------ .../NormedSpace/Multilinear/Basic.lean | 38 +++++------- .../NormedSpace/Multilinear/Curry.lean | 6 +- Mathlib/Data/Nat/Choose/Multinomial.lean | 2 +- Mathlib/Data/Nat/Factorial/BigOperators.lean | 3 +- .../MeasureTheory/Integral/TorusIntegral.lean | 3 +- test/positivity.lean | 21 +++++++ 10 files changed, 114 insertions(+), 54 deletions(-) diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 1c8f13e1aa0a1..5ba5ee09a25db 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -206,3 +206,61 @@ lemma IsAbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCom #align is_absolute_value.map_prod IsAbsoluteValue.map_prod end AbsoluteValue + +namespace Mathlib.Meta.Positivity +open Qq Lean Meta Finset + +private alias ⟨_, prod_ne_zero⟩ := prod_ne_zero_iff + +/-- The `positivity` extension which proves that `∏ i in s, f i` is nonnegative if `f` is, and +positive if each `f i` is. + +TODO: The following example does not work +``` +example (s : Finset ℕ) (f : ℕ → ℤ) (hf : ∀ n, 0 ≤ f n) : 0 ≤ s.prod f := by positivity +``` +because `compareHyp` can't look for assumptions behind binders. +-/ +@[positivity Finset.prod _ _] +def evalFinsetProd : PositivityExt where eval {u α} zα pα e := do + match e with + | ~q(@Finset.prod $ι _ $instα $s $f) => + let i : Q($ι) ← mkFreshExprMVarQ q($ι) .syntheticOpaque + have body : Q($α) := Expr.betaRev f #[i] + let rbody ← core zα pα body + let _instαmon ← synthInstanceQ q(CommMonoidWithZero $α) + + -- Try to show that the product is positive + let p_pos : Option Q(0 < $e) := ← do + let .positive pbody := rbody | pure none -- Fail if the body is not provably positive + -- TODO(quote4#38): We must name the following, else `assertInstancesCommute` loops. + let .some _instαzeroone ← trySynthInstanceQ q(ZeroLEOneClass $α) | pure none + let .some _instαposmul ← trySynthInstanceQ q(PosMulStrictMono $α) | pure none + let .some _instαnontriv ← trySynthInstanceQ q(Nontrivial $α) | pure none + assertInstancesCommute + let pr : Q(∀ i, 0 < $f i) ← mkLambdaFVars #[i] pbody (binderInfoForMVars := .default) + return some q(prod_pos fun i _ ↦ $pr i) + if let some p_pos := p_pos then return .positive p_pos + + -- Try to show that the product is nonnegative + let p_nonneg : Option Q(0 ≤ $e) := ← do + let .some pbody := rbody.toNonneg + | return none -- Fail if the body is not provably nonnegative + let pr : Q(∀ i, 0 ≤ $f i) ← mkLambdaFVars #[i] pbody (binderInfoForMVars := .default) + -- TODO(quote4#38): We must name the following, else `assertInstancesCommute` loops. + let .some _instαzeroone ← trySynthInstanceQ q(ZeroLEOneClass $α) | pure none + let .some _instαposmul ← trySynthInstanceQ q(PosMulMono $α) | pure none + assertInstancesCommute + return some q(prod_nonneg fun i _ ↦ $pr i) + if let some p_nonneg := p_nonneg then return .nonnegative p_nonneg + + -- Fall back to showing that the product is nonzero + let pbody ← rbody.toNonzero + let pr : Q(∀ i, $f i ≠ 0) ← mkLambdaFVars #[i] pbody (binderInfoForMVars := .default) + -- TODO(quote4#38): We must name the following, else `assertInstancesCommute` loops. + let _instαnontriv ← synthInstanceQ q(Nontrivial $α) + let _instαnozerodiv ← synthInstanceQ q(NoZeroDivisors $α) + assertInstancesCommute + return .nonzero q(prod_ne_zero fun i _ ↦ $pr i) + +end Mathlib.Meta.Positivity diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 2d23fb07258b2..5da081023ea5f 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -331,9 +331,7 @@ the norms of the relevant bits of `q` and `p`. -/ theorem compAlongComposition_norm {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) : ‖q.compAlongComposition p c‖ ≤ ‖q c.length‖ * ∏ i, ‖p (c.blocksFun i)‖ := - ContinuousMultilinearMap.opNorm_le_bound _ - (mul_nonneg (norm_nonneg _) (Finset.prod_nonneg fun _i _hi => norm_nonneg _)) - (compAlongComposition_bound _ _ _) + ContinuousMultilinearMap.opNorm_le_bound _ (by positivity) (compAlongComposition_bound _ _ _) #align formal_multilinear_series.comp_along_composition_norm FormalMultilinearSeries.compAlongComposition_norm theorem compAlongComposition_nnnorm {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index b51b9dac274f1..f0d914db19b87 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -476,8 +476,7 @@ theorem radius_rightInv_pos_of_radius_pos_aux2 {n : ℕ} (hn : 2 ≤ n + 1) gcongr apply (compAlongComposition_norm _ _ _).trans gcongr - · exact prod_nonneg fun j _ => norm_nonneg _ - · apply hp + apply hp _ = I * a + I * C * diff --git a/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean b/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean index 455de6a2ef748..c2be04a15b2ba 100644 --- a/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/NormedSpace/BoundedLinearMaps.lean @@ -217,24 +217,18 @@ operation. -/ theorem isBoundedLinearMap_prod_multilinear {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] : IsBoundedLinearMap 𝕜 fun p : ContinuousMultilinearMap 𝕜 E F × ContinuousMultilinearMap 𝕜 E G => - p.1.prod p.2 := - { map_add := fun p₁ p₂ => by - ext1 m - rfl - map_smul := fun c p => by - ext1 m - rfl - bound := - ⟨1, zero_lt_one, fun p => by - rw [one_mul] - apply ContinuousMultilinearMap.opNorm_le_bound _ (norm_nonneg _) _ - intro m - rw [ContinuousMultilinearMap.prod_apply, norm_prod_le_iff] - constructor - · exact (p.1.le_opNorm m).trans (mul_le_mul_of_nonneg_right (norm_fst_le p) - (Finset.prod_nonneg fun i _ => norm_nonneg _)) - · exact (p.2.le_opNorm m).trans (mul_le_mul_of_nonneg_right (norm_snd_le p) - (Finset.prod_nonneg fun i _ => norm_nonneg _))⟩ } + p.1.prod p.2 where + map_add p₁ p₂ := by ext : 1; rfl + map_smul c p := by ext : 1; rfl + bound := by + refine ⟨1, zero_lt_one, fun p ↦ ?_⟩ + rw [one_mul] + apply ContinuousMultilinearMap.opNorm_le_bound _ (norm_nonneg _) _ + intro m + rw [ContinuousMultilinearMap.prod_apply, norm_prod_le_iff] + constructor + · exact (p.1.le_opNorm m).trans (mul_le_mul_of_nonneg_right (norm_fst_le p) <| by positivity) + · exact (p.2.le_opNorm m).trans (mul_le_mul_of_nonneg_right (norm_snd_le p) <| by positivity) #align is_bounded_linear_map_prod_multilinear isBoundedLinearMap_prod_multilinear /-- Given a fixed continuous linear map `g`, associating to a continuous multilinear map `f` the diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 74b5d188840b4..0dc7f02aed35f 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -241,10 +241,8 @@ continuous. -/ theorem continuous_of_bound (C : ℝ) (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : Continuous f := by let D := max C 1 have D_pos : 0 ≤ D := le_trans zero_le_one (le_max_right _ _) - replace H : ∀ m, ‖f m‖ ≤ D * ∏ i, ‖m i‖ := by - intro m - apply le_trans (H m) (mul_le_mul_of_nonneg_right (le_max_left _ _) _) - exact prod_nonneg fun (i : ι) _ => norm_nonneg (m i) + replace H (m) : ‖f m‖ ≤ D * ∏ i, ‖m i‖ := + (H m).trans (mul_le_mul_of_nonneg_right (le_max_left _ _) <| by positivity) refine' continuous_iff_continuousAt.2 fun m => _ refine' continuousAt_of_locally_lipschitz zero_lt_one @@ -362,7 +360,7 @@ variable {f m} theorem le_mul_prod_of_le_opNorm_of_le {C : ℝ} {b : ι → ℝ} (hC : ‖f‖ ≤ C) (hm : ∀ i, ‖m i‖ ≤ b i) : ‖f m‖ ≤ C * ∏ i, b i := (f.le_opNorm m).trans <| mul_le_mul hC (prod_le_prod (fun _ _ ↦ norm_nonneg _) fun _ _ ↦ hm _) - (prod_nonneg fun _ _ ↦ norm_nonneg _) ((opNorm_nonneg _).trans hC) + (by positivity) ((opNorm_nonneg _).trans hC) @[deprecated] alias le_mul_prod_of_le_op_norm_of_le := le_mul_prod_of_le_opNorm_of_le -- 2024-02-02 @@ -400,8 +398,7 @@ theorem le_of_opNorm_le {C : ℝ} (h : ‖f‖ ≤ C) : ‖f m‖ ≤ C * ∏ i, variable (f) theorem ratio_le_opNorm : (‖f m‖ / ∏ i, ‖m i‖) ≤ ‖f‖ := - div_le_of_nonneg_of_le_mul (prod_nonneg fun _ _ => norm_nonneg _) (opNorm_nonneg _) - (f.le_opNorm m) + div_le_of_nonneg_of_le_mul (by positivity) (opNorm_nonneg _) (f.le_opNorm m) #align continuous_multilinear_map.ratio_le_op_norm ContinuousMultilinearMap.ratio_le_opNorm @[deprecated] alias ratio_le_op_norm := ratio_le_opNorm -- deprecated on 2024-02-02 @@ -768,9 +765,8 @@ theorem MultilinearMap.mkContinuous_norm_le (f : MultilinearMap 𝕜 E G) {C : nonnegative. -/ theorem MultilinearMap.mkContinuous_norm_le' (f : MultilinearMap 𝕜 E G) {C : ℝ} (H : ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖) : ‖f.mkContinuous C H‖ ≤ max C 0 := - ContinuousMultilinearMap.opNorm_le_bound _ (le_max_right _ _) fun m => - (H m).trans <| - mul_le_mul_of_nonneg_right (le_max_left _ _) (prod_nonneg fun _ _ => norm_nonneg _) + ContinuousMultilinearMap.opNorm_le_bound _ (le_max_right _ _) fun m ↦ (H m).trans <| + mul_le_mul_of_nonneg_right (le_max_left _ _) <| by positivity #align multilinear_map.mk_continuous_norm_le' MultilinearMap.mkContinuous_norm_le' namespace ContinuousMultilinearMap @@ -1029,8 +1025,7 @@ def flipMultilinear (f : G →L[𝕜] ContinuousMultilinearMap 𝕜 E G') : LinearMap.mkContinuous_apply, Pi.smul_apply, AddHom.coe_mk] } ‖f‖ fun m => by dsimp only [MultilinearMap.coe_mk] - refine LinearMap.mkContinuous_norm_le _ - (mul_nonneg (norm_nonneg f) (prod_nonneg fun i _ => norm_nonneg (m i))) _ + exact LinearMap.mkContinuous_norm_le _ (by positivity) _ #align continuous_linear_map.flip_multilinear ContinuousLinearMap.flipMultilinear #align continuous_linear_map.flip_multilinear_apply_apply ContinuousLinearMap.flipMultilinear_apply_apply @@ -1103,7 +1098,7 @@ def mkContinuousMultilinear (f : MultilinearMap 𝕜 E (MultilinearMap 𝕜 E' G simp only [coe_mk] refine ((f m).mkContinuous_norm_le' _).trans_eq ?_ rw [max_mul_of_nonneg, zero_mul] - exact prod_nonneg fun _ _ => norm_nonneg _ + positivity #align multilinear_map.mk_continuous_multilinear MultilinearMap.mkContinuousMultilinear @[simp] @@ -1134,7 +1129,7 @@ set_option linter.uppercaseLean3 false theorem norm_compContinuousLinearMap_le (g : ContinuousMultilinearMap 𝕜 E₁ G) (f : ∀ i, E i →L[𝕜] E₁ i) : ‖g.compContinuousLinearMap f‖ ≤ ‖g‖ * ∏ i, ‖f i‖ := - opNorm_le_bound _ (mul_nonneg (norm_nonneg _) <| prod_nonneg fun i _ => norm_nonneg _) fun m => + opNorm_le_bound _ (by positivity) fun m => calc ‖g fun i => f i (m i)‖ ≤ ‖g‖ * ∏ i, ‖f i (m i)‖ := g.le_opNorm _ _ ≤ ‖g‖ * ∏ i, ‖f i‖ * ‖m i‖ := @@ -1194,7 +1189,7 @@ theorem compContinuousLinearMapL_apply (g : ContinuousMultilinearMap 𝕜 E₁ G variable (G) in theorem norm_compContinuousLinearMapL_le (f : ∀ i, E i →L[𝕜] E₁ i) : ‖compContinuousLinearMapL (G := G) f‖ ≤ ∏ i, ‖f i‖ := - LinearMap.mkContinuous_norm_le _ (prod_nonneg fun _ _ => norm_nonneg _) _ + LinearMap.mkContinuous_norm_le _ (by positivity) _ #align continuous_multilinear_map.norm_comp_continuous_linear_mapL_le ContinuousMultilinearMap.norm_compContinuousLinearMapL_le /-- `ContinuousMultilinearMap.compContinuousLinearMap` as a bundled continuous linear map. @@ -1369,8 +1364,6 @@ addition of `Finset.prod` where needed). The duplication could be avoided by ded case from the multilinear case via a currying isomorphism. However, this would mess up imports, and it is more satisfactory to have the simplest case as a standalone proof. -/ instance completeSpace [CompleteSpace G] : CompleteSpace (ContinuousMultilinearMap 𝕜 E G) := by - have nonneg : ∀ v : ∀ i, E i, 0 ≤ ∏ i, ‖v i‖ := fun v => - Finset.prod_nonneg fun i _ => norm_nonneg _ -- We show that every Cauchy sequence converges. refine' Metric.complete_of_cauchySeq_tendsto fun f hf => _ -- We now expand out the definition of a Cauchy sequence, @@ -1379,12 +1372,13 @@ instance completeSpace [CompleteSpace G] : CompleteSpace (ContinuousMultilinearM have cau : ∀ v, CauchySeq fun n => f n v := by intro v apply cauchySeq_iff_le_tendsto_0.2 ⟨fun n => b n * ∏ i, ‖v i‖, _, _, _⟩ - · intro - exact mul_nonneg (b0 _) (nonneg v) + · intro n + have := b0 n + positivity · intro n m N hn hm rw [dist_eq_norm] apply le_trans ((f n - f m).le_opNorm v) _ - exact mul_le_mul_of_nonneg_right (b_bound n m N hn hm) (nonneg v) + exact mul_le_mul_of_nonneg_right (b_bound n m N hn hm) <| by positivity · simpa using b_lim.mul tendsto_const_nhds -- We assemble the limits points of those Cauchy sequences -- (which exist as `G` is complete) @@ -1409,7 +1403,7 @@ instance completeSpace [CompleteSpace G] : CompleteSpace (ContinuousMultilinearM have A : ∀ n, ‖f n v‖ ≤ (b 0 + ‖f 0‖) * ∏ i, ‖v i‖ := by intro n apply le_trans ((f n).le_opNorm _) _ - apply mul_le_mul_of_nonneg_right _ (nonneg v) + apply mul_le_mul_of_nonneg_right _ <| by positivity calc ‖f n‖ = ‖f n - f 0 + f 0‖ := by congr 1 @@ -1429,7 +1423,7 @@ instance completeSpace [CompleteSpace G] : CompleteSpace (ContinuousMultilinearM have A : ∀ᶠ m in atTop, ‖(f n - f m) v‖ ≤ b n * ∏ i, ‖v i‖ := by refine' eventually_atTop.2 ⟨n, fun m hm => _⟩ apply le_trans ((f n - f m).le_opNorm _) _ - exact mul_le_mul_of_nonneg_right (b_bound n m n le_rfl hm) (nonneg v) + exact mul_le_mul_of_nonneg_right (b_bound n m n le_rfl hm) <| by positivity have B : Tendsto (fun m => ‖(f n - f m) v‖) atTop (𝓝 ‖(f n - Fcont) v‖) := Tendsto.norm (tendsto_const_nhds.sub (hF v)) exact le_of_tendsto B A diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean index e085b0f9e8c0b..b5896ae869221 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Curry.lean @@ -63,8 +63,7 @@ theorem ContinuousLinearMap.norm_map_tail_le ‖f (m 0) (tail m)‖ ≤ ‖f‖ * ∏ i, ‖m i‖ := calc ‖f (m 0) (tail m)‖ ≤ ‖f (m 0)‖ * ∏ i, ‖(tail m) i‖ := (f (m 0)).le_opNorm _ - _ ≤ ‖f‖ * ‖m 0‖ * ∏ i, ‖(tail m) i‖ := - (mul_le_mul_of_nonneg_right (f.le_opNorm _) (prod_nonneg fun _ _ => norm_nonneg _)) + _ ≤ ‖f‖ * ‖m 0‖ * ∏ i, ‖tail m i‖ := mul_le_mul_of_nonneg_right (f.le_opNorm _) <| by positivity _ = ‖f‖ * (‖m 0‖ * ∏ i, ‖(tail m) i‖) := by ring _ = ‖f‖ * ∏ i, ‖m i‖ := by rw [prod_univ_succ] @@ -269,8 +268,7 @@ def ContinuousMultilinearMap.curryRight (f : ContinuousMultilinearMap 𝕜 Ei G) simp } f'.mkContinuous ‖f‖ fun m => by simp only [f', MultilinearMap.coe_mk] - exact LinearMap.mkContinuous_norm_le _ - (mul_nonneg (norm_nonneg _) (prod_nonneg fun _ _ => norm_nonneg _)) _ + exact LinearMap.mkContinuous_norm_le _ (by positivity) _ #align continuous_multilinear_map.curry_right ContinuousMultilinearMap.curryRight @[simp] diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index a9e0378d8566d..27a01d251fb26 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -65,7 +65,7 @@ lemma multinomial_cons (ha : a ∉ s) (f : α → ℕ) : multinomial, mul_assoc, mul_left_comm _ (f a)!, Nat.div_mul_cancel (prod_factorial_dvd_factorial_sum _ _), ← mul_assoc, Nat.choose_symm_add, Nat.add_choose_mul_factorial_mul_factorial, Finset.sum_cons] - exact prod_pos fun i _ ↦ by positivity + positivity lemma multinomial_insert [DecidableEq α] (ha : a ∉ s) (f : α → ℕ) : multinomial (insert a s) f = (f a + ∑ i in s, f i).choose (f a) * multinomial s f := by diff --git a/Mathlib/Data/Nat/Factorial/BigOperators.lean b/Mathlib/Data/Nat/Factorial/BigOperators.lean index 74c65728722a5..3d7cf6af12894 100644 --- a/Mathlib/Data/Nat/Factorial/BigOperators.lean +++ b/Mathlib/Data/Nat/Factorial/BigOperators.lean @@ -28,8 +28,7 @@ lemma monotone_factorial : Monotone factorial := fun _ _ => factorial_le variable {α : Type*} (s : Finset α) (f : α → ℕ) -theorem prod_factorial_pos : 0 < ∏ i in s, (f i)! := - Finset.prod_pos fun i _ => factorial_pos (f i) +theorem prod_factorial_pos : 0 < ∏ i in s, (f i)! := by positivity #align nat.prod_factorial_pos Nat.prod_factorial_pos theorem prod_factorial_dvd_factorial_sum : (∏ i in s, (f i)!) ∣ (∑ i in s, f i)! := by diff --git a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean index b56d0f5fe3046..d3e3ce6ef54e3 100644 --- a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean @@ -199,8 +199,7 @@ theorem norm_torusIntegral_le_of_norm_le_const {C : ℝ} (hf : ∀ θ, ‖f (tor ‖(∏ i : Fin n, R i * exp (θ i * I) * I : ℂ) • f (torusMap c R θ)‖ = (∏ i : Fin n, |R i|) * ‖f (torusMap c R θ)‖ := by simp [norm_smul] - _ ≤ (∏ i : Fin n, |R i|) * C := - mul_le_mul_of_nonneg_left (hf _) (Finset.prod_nonneg fun _ _ => abs_nonneg _) + _ ≤ (∏ i : Fin n, |R i|) * C := mul_le_mul_of_nonneg_left (hf _) <| by positivity _ = ((2 * π) ^ (n : ℕ) * ∏ i, |R i|) * C := by simp only [Pi.zero_def, Real.volume_Icc_pi_toReal fun _ => Real.two_pi_pos.le, sub_zero, Fin.prod_const, mul_assoc, mul_comm ((2 * π) ^ (n : ℕ))] diff --git a/test/positivity.lean b/test/positivity.lean index 3150b776d1a80..5e58a2e5776d3 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -382,6 +382,27 @@ example (s : Finset ℕ) (f : ℕ → ℕ) (a : ℕ) : 0 ≤ s.sum (f a) := by p set_option linter.unusedVariables false in example (f : ℕ → ℕ) (hf : 0 ≤ f 0) : 0 ≤ ∑ n in Finset.range 10, f n := by positivity +example (n : ℕ) : ∏ j in range n, (-1) ≠ 0 := by positivity +example (n : ℕ) (a : ℕ → ℤ) : 0 ≤ ∏ j in range n, a j^2 := by positivity +example (a : ULift.{2} ℕ → ℤ) (s : Finset (ULift.{2} ℕ)) : 0 ≤ ∏ j in s, a j^2 := by positivity +example (n : ℕ) (a : ℕ → ℤ) : 0 ≤ ∏ j : Fin 8, ∏ i in range n, (a j^2 + i ^ 2) := by positivity +example (n : ℕ) (a : ℕ → ℤ) : 0 < ∏ j : Fin (n + 1), (a j^2 + 1) := by positivity +example (a : ℕ → ℤ) : 0 < ∏ j in ({1} : Finset ℕ), (a j^2 + 1) := by + have : Finset.Nonempty {1} := singleton_nonempty 1 + positivity +example (s : Finset ℕ) : 0 ≤ ∏ j in s, j := by positivity +example (s : Finset ℕ) : 0 ≤ s.sum id := by positivity +example (s : Finset ℕ) (f : ℕ → ℕ) (a : ℕ) : 0 ≤ s.sum (f a) := by positivity + +-- Make sure that the extension doesn't produce an invalid term by accidentally unifying `?n` with +-- `0` because of the `hf` assumption +set_option linter.unusedVariables false in +example (f : ℕ → ℕ) (hf : 0 ≤ f 0) : 0 ≤ ∏ n in Finset.range 10, f n := by positivity + +-- Make sure that `positivity` isn't too greedy by trying to prove that a product is positive +-- because its body is even if multiplication isn't strictly monotone +example [OrderedCommSemiring α] {a : α} (ha : 0 < a) : 0 ≤ ∏ _i in {(0 : α)}, a := by positivity + /- ## Other extensions -/ example [Zero β] [PartialOrder β] [FunLike F α β] [NonnegHomClass F α β] From 424683e862f4b80437c11a749dba0b98cffbf1ec Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 7 May 2024 23:13:54 +0000 Subject: [PATCH 25/52] chore: Rename `smul_one_eq_coe` to `smul_one_eq_cast` (#12621) --- Archive/Wiedijk100Theorems/FriendshipGraphs.lean | 2 +- Mathlib/Algebra/CharP/Basic.lean | 2 +- Mathlib/Algebra/Field/Defs.lean | 10 +++++++--- Mathlib/Algebra/Module/BigOperators.lean | 2 +- Mathlib/Algebra/Module/Defs.lean | 11 +++++++---- Mathlib/Algebra/Polynomial/Basic.lean | 4 ++-- Mathlib/Algebra/Polynomial/UnitTrinomial.lean | 2 +- Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean | 2 +- Mathlib/Analysis/MeanInequalities.lean | 4 ++-- Mathlib/Analysis/Normed/Group/AddCircle.lean | 2 +- Mathlib/Analysis/NormedSpace/Basic.lean | 4 ++-- Mathlib/Analysis/NormedSpace/Star/Multiplier.lean | 4 ++-- .../Combinatorics/SimpleGraph/Regularity/Energy.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 4 ++-- Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean | 4 ++-- Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean | 2 +- Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean | 4 ++-- .../RingTheory/Polynomial/Eisenstein/IsIntegral.lean | 2 +- .../RingTheory/WittVector/FrobeniusFractionField.lean | 2 +- 19 files changed, 38 insertions(+), 31 deletions(-) diff --git a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean index 1b30d5b7718b2..796e551654f30 100644 --- a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean +++ b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean @@ -206,7 +206,7 @@ end Nonempty theorem adjMatrix_sq_mul_const_one_of_regular (hd : G.IsRegularOfDegree d) : G.adjMatrix R * of (fun _ _ => 1) = of (fun _ _ => (d : R)) := by ext x - simp only [← hd x, degree, adjMatrix_mul_apply, sum_const, Nat.smul_one_eq_coe, + simp only [← hd x, degree, adjMatrix_mul_apply, sum_const, Nat.smul_one_eq_cast, of_apply] #align theorems_100.friendship.adj_matrix_sq_mul_const_one_of_regular Theorems100.Friendship.adjMatrix_sq_mul_const_one_of_regular diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index a74bdba2f3672..c0f71ddc226cf 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -137,7 +137,7 @@ theorem CharP.cast_card_eq_zero [AddGroupWithOne R] [Fintype R] : (Fintype.card #align char_p.cast_card_eq_zero CharP.cast_card_eq_zero theorem CharP.addOrderOf_one (R) [Semiring R] : CharP R (addOrderOf (1 : R)) := - ⟨fun n => by rw [← Nat.smul_one_eq_coe, addOrderOf_dvd_iff_nsmul_eq_zero]⟩ + ⟨fun n => by rw [← Nat.smul_one_eq_cast, addOrderOf_dvd_iff_nsmul_eq_zero]⟩ #align char_p.add_order_of_one CharP.addOrderOf_one theorem CharP.intCast_eq_zero_iff [AddGroupWithOne R] (p : ℕ) [CharP R p] (a : ℤ) : diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index b06596ac3edf0..993c9b3d1f879 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -200,7 +200,9 @@ lemma smul_def (q : ℚ≥0) (a : α) : q • a = q * a := DivisionSemiring.nnqs variable (α) -@[simp] lemma smul_one_eq_coe (q : ℚ≥0) : q • (1 : α) = q := by rw [NNRat.smul_def, mul_one] +@[simp] lemma smul_one_eq_cast (q : ℚ≥0) : q • (1 : α) = q := by rw [NNRat.smul_def, mul_one] + +@[deprecated (since := "2024-05-03")] alias smul_one_eq_coe := smul_one_eq_cast end NNRat @@ -221,9 +223,11 @@ theorem smul_def (a : ℚ) (x : K) : a • x = ↑a * x := DivisionRing.qsmul_de #align rat.smul_def Rat.smul_def @[simp] -theorem smul_one_eq_coe (A : Type*) [DivisionRing A] (m : ℚ) : m • (1 : A) = ↑m := by +theorem smul_one_eq_cast (A : Type*) [DivisionRing A] (m : ℚ) : m • (1 : A) = ↑m := by rw [Rat.smul_def, mul_one] -#align rat.smul_one_eq_coe Rat.smul_one_eq_coe +#align rat.smul_one_eq_coe Rat.smul_one_eq_cast + +@[deprecated (since := "2024-05-03")] alias smul_one_eq_coe := smul_one_eq_cast end Rat diff --git a/Mathlib/Algebra/Module/BigOperators.lean b/Mathlib/Algebra/Module/BigOperators.lean index 977c708bca13f..cb7feff3cc5d5 100644 --- a/Mathlib/Algebra/Module/BigOperators.lean +++ b/Mathlib/Algebra/Module/BigOperators.lean @@ -50,7 +50,7 @@ theorem Finset.sum_smul_sum {f : α → R} {g : β → M} {s : Finset α} {t : F end AddCommMonoid theorem Finset.cast_card [CommSemiring R] (s : Finset α) : (s.card : R) = ∑ a in s, 1 := by - rw [Finset.sum_const, Nat.smul_one_eq_coe] + rw [Finset.sum_const, Nat.smul_one_eq_cast] #align finset.cast_card Finset.cast_card open Finset diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 6fc278913501c..6e8ccbb8f7b7a 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -628,15 +628,18 @@ end NoZeroSMulDivisors -- Porting note (#10618): simp can prove this --@[simp] -theorem Nat.smul_one_eq_coe {R : Type*} [Semiring R] (m : ℕ) : m • (1 : R) = ↑m := by +theorem Nat.smul_one_eq_cast {R : Type*} [Semiring R] (m : ℕ) : m • (1 : R) = ↑m := by rw [nsmul_eq_mul, mul_one] -#align nat.smul_one_eq_coe Nat.smul_one_eq_coe +#align nat.smul_one_eq_coe Nat.smul_one_eq_cast -- Porting note (#10618): simp can prove this --@[simp] -theorem Int.smul_one_eq_coe {R : Type*} [Ring R] (m : ℤ) : m • (1 : R) = ↑m := by +theorem Int.smul_one_eq_cast {R : Type*} [Ring R] (m : ℤ) : m • (1 : R) = ↑m := by rw [zsmul_eq_mul, mul_one] -#align int.smul_one_eq_coe Int.smul_one_eq_coe +#align int.smul_one_eq_coe Int.smul_one_eq_cast + +@[deprecated (since := "2024-05-03")] alias Nat.smul_one_eq_coe := Nat.smul_one_eq_cast +@[deprecated (since := "2024-05-03")] alias Int.smul_one_eq_coe := Int.smul_one_eq_cast assert_not_exists Multiset assert_not_exists Set.indicator diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 34d4eff5e090a..908e7f91af162 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -1262,7 +1262,7 @@ section DivisionSemiring variable [DivisionSemiring R] lemma nnqsmul_eq_C_mul (q : ℚ≥0) (f : R[X]) : q • f = Polynomial.C (q : R) * f := by - rw [← NNRat.smul_one_eq_coe, ← Polynomial.smul_C, C_1, smul_one_mul] + rw [← NNRat.smul_one_eq_cast, ← Polynomial.smul_C, C_1, smul_one_mul] end DivisionSemiring @@ -1271,7 +1271,7 @@ section DivisionRing variable [DivisionRing R] theorem qsmul_eq_C_mul (a : ℚ) (f : R[X]) : a • f = Polynomial.C (a : R) * f := by - rw [← Rat.smul_one_eq_coe, ← Polynomial.smul_C, C_1, smul_one_mul] + rw [← Rat.smul_one_eq_cast, ← Polynomial.smul_C, C_1, smul_one_mul] #align polynomial.rat_smul_eq_C_mul Polynomial.qsmul_eq_C_mul end DivisionRing diff --git a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean index 91ce67a0ebb84..eb70273d858cc 100644 --- a/Mathlib/Algebra/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Algebra/Polynomial/UnitTrinomial.lean @@ -208,7 +208,7 @@ theorem isUnitTrinomial_iff' : Int.sq_eq_one_of_sq_le_three ((single_le_sum (fun k _ => sq_nonneg (p.coeff k)) hk).trans hp.le) (mem_support_iff.mp hk) refine' isUnitTrinomial_iff.mpr ⟨_, fun k hk => isUnit_ofPowEqOne (key k hk) two_ne_zero⟩ - rw [sum_def, sum_congr rfl key, sum_const, Nat.smul_one_eq_coe] at hp + rw [sum_def, sum_congr rfl key, sum_const, Nat.smul_one_eq_cast] at hp exact Nat.cast_injective hp #align polynomial.is_unit_trinomial_iff' Polynomial.isUnitTrinomial_iff' diff --git a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean index aa767b2cc6b95..23495c800e262 100644 --- a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean @@ -134,7 +134,7 @@ theorem Asymptotics.isLittleO_sum_range_of_tendsto_zero {α : Type*} [NormedAddC {f : ℕ → α} (h : Tendsto f atTop (𝓝 0)) : (fun n => ∑ i in range n, f i) =o[atTop] fun n => (n : ℝ) := by have := ((isLittleO_one_iff ℝ).2 h).sum_range fun i => zero_le_one - simp only [sum_const, card_range, Nat.smul_one_eq_coe] at this + simp only [sum_const, card_range, Nat.smul_one_eq_cast] at this exact this tendsto_natCast_atTop_atTop #align asymptotics.is_o_sum_range_of_tendsto_zero Asymptotics.isLittleO_sum_range_of_tendsto_zero diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index d1b716f4c9a79..27ecafd98508e 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -468,7 +468,7 @@ theorem rpow_sum_le_const_mul_sum_rpow (f : ι → ℝ≥0) {p : ℝ} (hp : 1 rw [← hpq.div_conj_eq_sub_one] ring simpa only [NNReal.mul_rpow, ← NNReal.rpow_mul, hp₁, hq, one_mul, one_rpow, rpow_one, - Pi.one_apply, sum_const, Nat.smul_one_eq_coe] using + Pi.one_apply, sum_const, Nat.smul_one_eq_cast] using NNReal.rpow_le_rpow (inner_le_Lp_mul_Lq s 1 f hpq.symm) hpq.nonneg #align nnreal.rpow_sum_le_const_mul_sum_rpow NNReal.rpow_sum_le_const_mul_sum_rpow @@ -834,7 +834,7 @@ theorem rpow_sum_le_const_mul_sum_rpow (hp : 1 ≤ p) : rw [← hpq.div_conj_eq_sub_one] ring simpa only [ENNReal.mul_rpow_of_nonneg _ _ hpq.nonneg, ← ENNReal.rpow_mul, hp₁, hq, coe_one, - one_mul, one_rpow, rpow_one, Pi.one_apply, sum_const, Nat.smul_one_eq_coe] using + one_mul, one_rpow, rpow_one, Pi.one_apply, sum_const, Nat.smul_one_eq_cast] using ENNReal.rpow_le_rpow (inner_le_Lp_mul_Lq s 1 f hpq.symm) hpq.nonneg #align ennreal.rpow_sum_le_const_mul_sum_rpow ENNReal.rpow_sum_le_const_mul_sum_rpow diff --git a/Mathlib/Analysis/Normed/Group/AddCircle.lean b/Mathlib/Analysis/Normed/Group/AddCircle.lean index d3a979c490527..284ed2ddd941f 100644 --- a/Mathlib/Analysis/Normed/Group/AddCircle.lean +++ b/Mathlib/Analysis/Normed/Group/AddCircle.lean @@ -110,7 +110,7 @@ theorem norm_eq {x : ℝ} : ‖(x : AddCircle p)‖ = |x - round (p⁻¹ * x) * · simp only [QuotientAddGroup.mk'_apply, Real.norm_eq_abs, le_csInf_iff h₁ h₂] rintro b' ⟨b, hb, rfl⟩ simp only [mem_setOf, QuotientAddGroup.eq_iff_sub_mem, mem_zmultiples_iff, - smul_one_eq_coe] at hb + smul_one_eq_cast] at hb obtain ⟨z, hz⟩ := hb rw [(by rw [hz]; abel : x = b - z), fract_sub_int, ← abs_sub_round_eq_min] convert round_le b 0 diff --git a/Mathlib/Analysis/NormedSpace/Basic.lean b/Mathlib/Analysis/NormedSpace/Basic.lean index d1d1d7e3f8ccd..3408fc161d25e 100644 --- a/Mathlib/Analysis/NormedSpace/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Basic.lean @@ -67,7 +67,7 @@ instance NormedField.to_boundedSMul : BoundedSMul 𝕜 𝕜 := variable (𝕜) in theorem norm_zsmul [NormedSpace 𝕜 E] (n : ℤ) (x : E) : ‖n • x‖ = ‖(n : 𝕜)‖ * ‖x‖ := by - rw [← norm_smul, ← Int.smul_one_eq_coe, smul_assoc, one_smul] + rw [← norm_smul, ← Int.smul_one_eq_cast, smul_assoc, one_smul] #align norm_zsmul norm_zsmul theorem eventually_nhds_norm_smul_sub_lt (c : 𝕜) (x : E) {ε : ℝ} (h : 0 < ε) : @@ -347,7 +347,7 @@ norm. -/ instance normedAlgebraRat {𝕜} [NormedDivisionRing 𝕜] [CharZero 𝕜] [NormedAlgebra ℝ 𝕜] : NormedAlgebra ℚ 𝕜 where norm_smul_le q x := by - rw [← smul_one_smul ℝ q x, Rat.smul_one_eq_coe, norm_smul, Rat.norm_cast_real] + rw [← smul_one_smul ℝ q x, Rat.smul_one_eq_cast, norm_smul, Rat.norm_cast_real] #align normed_algebra_rat normedAlgebraRat instance PUnit.normedAlgebra : NormedAlgebra 𝕜 PUnit where diff --git a/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean b/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean index 63a0d393bc6c2..2b61f71325f16 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean +++ b/Mathlib/Analysis/NormedSpace/Star/Multiplier.lean @@ -188,13 +188,13 @@ instance instNatCast : NatCast 𝓜(𝕜, A) where natCast n := ⟨n, fun x y => by rw [Prod.snd_natCast, Prod.fst_natCast] - simp only [← Nat.smul_one_eq_coe, smul_apply, one_apply, mul_smul_comm, smul_mul_assoc]⟩ + simp only [← Nat.smul_one_eq_cast, smul_apply, one_apply, mul_smul_comm, smul_mul_assoc]⟩ instance instIntCast : IntCast 𝓜(𝕜, A) where intCast n := ⟨n, fun x y => by rw [Prod.snd_intCast, Prod.fst_intCast] - simp only [← Int.smul_one_eq_coe, smul_apply, one_apply, mul_smul_comm, smul_mul_assoc]⟩ + simp only [← Int.smul_one_eq_cast, smul_apply, one_apply, mul_smul_comm, smul_mul_assoc]⟩ instance instPow : Pow 𝓜(𝕜, A) ℕ where pow a n := diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index 828077c86c0b0..b899283f22b5a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -51,7 +51,7 @@ theorem energy_le_one : P.energy G ≤ 1 := ∑ uv in P.parts.offDiag, G.edgeDensity uv.1 uv.2 ^ 2 ≤ P.parts.offDiag.card • (1 : ℚ) := sum_le_card_nsmul _ _ 1 fun uv _ => (sq_le_one_iff <| G.edgeDensity_nonneg _ _).2 <| G.edgeDensity_le_one _ _ - _ = P.parts.offDiag.card := Nat.smul_one_eq_coe _ + _ = P.parts.offDiag.card := Nat.smul_one_eq_cast _ _ ≤ _ := by rw [offDiag_card, one_mul] norm_cast diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 35a047e500ae5..db8dcbd5e5ed6 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -117,14 +117,14 @@ theorem addOrderOf_coe (a : ℕ) {n : ℕ} (n0 : n ≠ 0) : addOrderOf (a : ZMod cases' a with a · simp only [Nat.zero_eq, Nat.cast_zero, addOrderOf_zero, Nat.gcd_zero_right, Nat.pos_of_ne_zero n0, Nat.div_self] - rw [← Nat.smul_one_eq_coe, addOrderOf_nsmul' _ a.succ_ne_zero, ZMod.addOrderOf_one] + rw [← Nat.smul_one_eq_cast, addOrderOf_nsmul' _ a.succ_ne_zero, ZMod.addOrderOf_one] #align zmod.add_order_of_coe ZMod.addOrderOf_coe /-- This lemma works in the case in which `a ≠ 0`. The version where `ZMod n` is not infinite, i.e. `n ≠ 0`, is `addOrderOf_coe`. -/ @[simp] theorem addOrderOf_coe' {a : ℕ} (n : ℕ) (a0 : a ≠ 0) : addOrderOf (a : ZMod n) = n / n.gcd a := by - rw [← Nat.smul_one_eq_coe, addOrderOf_nsmul' _ a0, ZMod.addOrderOf_one] + rw [← Nat.smul_one_eq_cast, addOrderOf_nsmul' _ a0, ZMod.addOrderOf_one] #align zmod.add_order_of_coe' ZMod.addOrderOf_coe' /-- We have that `ringChar (ZMod n) = n`. -/ diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index d05879b426908..41df0d3c916b3 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -276,7 +276,7 @@ theorem sum_eq_card_of_is_trivial {ψ : AddChar R R'} (hψ : ¬IsNontrivial ψ) ∑ a, ψ a = Fintype.card R := by simp only [IsNontrivial] at hψ push_neg at hψ - simp only [hψ, Finset.sum_const, Nat.smul_one_eq_coe] + simp only [hψ, Finset.sum_const, Nat.smul_one_eq_cast] rfl #align add_char.sum_eq_card_of_is_trivial AddChar.sum_eq_card_of_is_trivial @@ -290,7 +290,7 @@ theorem sum_mulShift {R : Type*} [CommRing R] [Fintype R] [DecidableEq R] (hψ : IsPrimitive ψ) : ∑ x : R, ψ (x * b) = if b = 0 then Fintype.card R else 0 := by split_ifs with h · -- case `b = 0` - simp only [h, mul_zero, map_zero_one, Finset.sum_const, Nat.smul_one_eq_coe] + simp only [h, mul_zero, map_zero_one, Finset.sum_const, Nat.smul_one_eq_cast] rfl · -- case `b ≠ 0` simp_rw [mul_comm] diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean index a43106c1407cd..55d29ba3e19db 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussSum.lean @@ -317,7 +317,7 @@ theorem FiniteField.two_pow_card {F : Type*} [Fintype F] [Field F] (hF : ringCha -- Porting note: original proof -- ext; congr; apply pow_one ext (x : Fin 8); rw [← map_nsmul_pow ψ₈char]; congr 2; - rw [Nat.smul_one_eq_coe, Fin.cast_val_eq_self x] + rw [Nat.smul_one_eq_cast, Fin.cast_val_eq_self x] have h₂ : (0 + 1 * τ ^ 1 + 0 + -1 * τ ^ 3 + 0 + -1 * τ ^ 5 + 0 + 1 * τ ^ 7) ^ 2 = 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) := by ring have h₃ : 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) = ↑8 := by diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index 0b7473129a32c..6b797a47c3f18 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -31,7 +31,7 @@ open scoped BigOperators theorem eval_one_cyclotomic_prime {R : Type*} [CommRing R] {p : ℕ} [hn : Fact p.Prime] : eval 1 (cyclotomic p R) = p := by simp only [cyclotomic_prime, eval_X, one_pow, Finset.sum_const, eval_pow, eval_finset_sum, - Finset.card_range, smul_one_eq_coe] + Finset.card_range, smul_one_eq_cast] #align polynomial.eval_one_cyclotomic_prime Polynomial.eval_one_cyclotomic_prime -- @[simp] -- Porting note (#10618): simp already proves this @@ -43,7 +43,7 @@ theorem eval₂_one_cyclotomic_prime {R S : Type*} [CommRing R] [Semiring S] (f theorem eval_one_cyclotomic_prime_pow {R : Type*} [CommRing R] {p : ℕ} (k : ℕ) [hn : Fact p.Prime] : eval 1 (cyclotomic (p ^ (k + 1)) R) = p := by simp only [cyclotomic_prime_pow_eq_geom_sum hn.out, eval_X, one_pow, Finset.sum_const, eval_pow, - eval_finset_sum, Finset.card_range, smul_one_eq_coe] + eval_finset_sum, Finset.card_range, smul_one_eq_cast] #align polynomial.eval_one_cyclotomic_prime_pow Polynomial.eval_one_cyclotomic_prime_pow -- @[simp] -- Porting note (#10618): simp already proves this diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index 7aa9a24555ac8..e31143d9b5ce7 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -106,7 +106,7 @@ theorem cyclotomic_prime_pow_comp_X_add_one_isEisensteinAt [hp : Fact p.Prime] ( · exact ⟨p ^ n, by rw [pow_succ']⟩ · rw [coeff_zero_eq_eval_zero, eval_comp, cyclotomic_prime_pow_eq_geom_sum hp.out, eval_add, eval_X, eval_one, zero_add, eval_finset_sum] - simp only [eval_pow, eval_X, one_pow, sum_const, card_range, Nat.smul_one_eq_coe, + simp only [eval_pow, eval_X, one_pow, sum_const, card_range, Nat.smul_one_eq_cast, submodule_span_eq, Ideal.submodule_span_eq, Ideal.span_singleton_pow, Ideal.mem_span_singleton] intro h diff --git a/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean b/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean index 02ddc8eab4ea3..bbaecc250f071 100644 --- a/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean +++ b/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean @@ -81,7 +81,7 @@ theorem succNthDefiningPoly_degree [IsDomain k] (n : ℕ) (a₁ a₂ : 𝕎 k) ( (succNthDefiningPoly p n a₁ a₂ bs).degree = p := by have : (X ^ p * C (a₁.coeff 0 ^ p ^ (n + 1))).degree = (p : WithBot ℕ) := by rw [degree_mul, degree_C] - · simp only [Nat.cast_withBot, add_zero, degree_X, degree_pow, Nat.smul_one_eq_coe] + · simp only [Nat.cast_withBot, add_zero, degree_X, degree_pow, Nat.smul_one_eq_cast] · exact pow_ne_zero _ ha₁ have : (X ^ p * C (a₁.coeff 0 ^ p ^ (n + 1)) - X * C (a₂.coeff 0 ^ p ^ (n + 1))).degree = (p : WithBot ℕ) := by From 893e43e193c6af2aa6dd94b62a5c4575a2a8381b Mon Sep 17 00:00:00 2001 From: Brendan Seamas Murphy Date: Tue, 7 May 2024 23:13:55 +0000 Subject: [PATCH 26/52] doc(Mathlib/LinearAlgebra/TensorProduct/RightExactness): Fix typo in copyright string (#12738) Two letters in the author's name were transposed. --- Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean index 70359b8de9588..99838be39f8ee 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Antoine Chambert-Loir. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Antoine Chambetr-Loir +Authors: Antoine Chambert-Loir -/ import Mathlib.Algebra.Exact From b1cd4c0ea79dca2d30ef8a2b014892eaf11b2869 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 8 May 2024 00:17:59 +0000 Subject: [PATCH 27/52] chore: backported changes for nightly-2024-05-07 (#12747) This changes become necessary on `nightly-2024-05-07`, but work already on `master`. Co-authored-by: Scott Morrison --- Mathlib/Algebra/Order/BigOperators/Group/List.lean | 3 ++- Mathlib/Data/BinaryHeap.lean | 4 ++-- Mathlib/Data/UnionFind.lean | 2 +- 3 files changed, 5 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 348ef2d73b143..abb6bcfaf812f 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -76,7 +76,8 @@ lemma prod_lt_prod' [Preorder M] [CovariantClass M M (· * ·) (· < ·)] (h₁ : ∀ i ∈ l, f i ≤ g i) (h₂ : ∃ i ∈ l, f i < g i) : (l.map f).prod < (l.map g).prod := by induction' l with i l ihl · rcases h₂ with ⟨_, ⟨⟩, _⟩ - simp only [forall_mem_cons, exists_mem_cons, map_cons, prod_cons] at h₁ h₂ ⊢ + simp only [forall_mem_cons, map_cons, prod_cons] at h₁ ⊢ + simp only [mem_cons, exists_eq_or_imp] at h₂ cases h₂ · exact mul_lt_mul_of_lt_of_le ‹_› (prod_le_prod' h₁.2) · exact mul_lt_mul_of_le_of_lt h₁.1 <| ihl h₁.2 ‹_› diff --git a/Mathlib/Data/BinaryHeap.lean b/Mathlib/Data/BinaryHeap.lean index 1422924223bf3..0c2b758bdab98 100644 --- a/Mathlib/Data/BinaryHeap.lean +++ b/Mathlib/Data/BinaryHeap.lean @@ -134,7 +134,7 @@ def insertExtractMax {lt} (self : BinaryHeap α lt) (x : α) : α × BinaryHeap | some m => if lt x m then let a := self.1.set ⟨0, size_pos_of_max e⟩ x - (m, ⟨heapifyDown lt a ⟨0, by simp [a]; exact size_pos_of_max e⟩⟩) + (m, ⟨heapifyDown lt a ⟨0, by simp only [Array.size_set, a]; exact size_pos_of_max e⟩⟩) else (x, self) /-- `O(log n)`. Equivalent to `(self.max, self.popMax.insert x)`. -/ @@ -143,7 +143,7 @@ def replaceMax {lt} (self : BinaryHeap α lt) (x : α) : Option α × BinaryHeap | none => (none, ⟨self.1.push x⟩) | some m => let a := self.1.set ⟨0, size_pos_of_max e⟩ x - (some m, ⟨heapifyDown lt a ⟨0, by simp [a]; exact size_pos_of_max e⟩⟩) + (some m, ⟨heapifyDown lt a ⟨0, by simp only [Array.size_set, a]; exact size_pos_of_max e⟩⟩) /-- `O(log n)`. Replace the value at index `i` by `x`. Assumes that `x ≤ self.get i`. -/ def decreaseKey {lt} (self : BinaryHeap α lt) (i : Fin self.size) (x : α) : BinaryHeap α lt where diff --git a/Mathlib/Data/UnionFind.lean b/Mathlib/Data/UnionFind.lean index 1f505ba204ba7..50cf56f1d940d 100644 --- a/Mathlib/Data/UnionFind.lean +++ b/Mathlib/Data/UnionFind.lean @@ -277,7 +277,7 @@ def link (self : UnionFind α) (x y : Fin self.size) have : UFModel.Agrees arr₁ (·.parent) parent := hm.1.set (fun i h ↦ by simp [parent]; rw [if_neg h.symm]) (fun _ ↦ by simp [parent]) have H1 : UFModel.Agrees arr₂ (·.parent) parent := by - simp only [arr₂, getElem_fin]; split + simp only [arr₂, Fin.getElem_fin]; split · exact this.set (fun i h ↦ by simp [h.symm]) fun _ ↦ by simp [ne, hm.parent_eq', ny, parent] · exact this have : UFModel.Agrees arr₁ (·.rank) (fun i : Fin n ↦ m.rank i) := From 628b26666181cf42e401d9babbcc650423c90dc4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 8 May 2024 01:12:54 +0000 Subject: [PATCH 28/52] feat(UniformConvergenceTopology): prove `CompleteSpace` (#10844) --- Mathlib/Topology/UniformSpace/Equiv.lean | 7 ++- .../UniformConvergenceTopology.lean | 55 ++++++++++++++++++- 2 files changed, 57 insertions(+), 5 deletions(-) diff --git a/Mathlib/Topology/UniformSpace/Equiv.lean b/Mathlib/Topology/UniformSpace/Equiv.lean index c2f386c67a5d0..ab63ab2af7e91 100644 --- a/Mathlib/Topology/UniformSpace/Equiv.lean +++ b/Mathlib/Topology/UniformSpace/Equiv.lean @@ -240,14 +240,17 @@ protected theorem uniformInducing (h : α ≃ᵤ β) : UniformInducing h := simp only [symm_comp_self, uniformInducing_id] #align uniform_equiv.uniform_inducing UniformEquiv.uniformInducing -theorem comap_eq (h : α ≃ᵤ β) : UniformSpace.comap h ‹_› = ‹_› := by - ext : 1; exact h.uniformInducing.comap_uniformity +theorem comap_eq (h : α ≃ᵤ β) : UniformSpace.comap h ‹_› = ‹_› := + h.uniformInducing.comap_uniformSpace #align uniform_equiv.comap_eq UniformEquiv.comap_eq protected theorem uniformEmbedding (h : α ≃ᵤ β) : UniformEmbedding h := ⟨h.uniformInducing, h.injective⟩ #align uniform_equiv.uniform_embedding UniformEquiv.uniformEmbedding +theorem completeSpace_iff (h : α ≃ᵤ β) : CompleteSpace α ↔ CompleteSpace β := + completeSpace_congr h.uniformEmbedding + /-- Uniform equiv given a uniform embedding. -/ noncomputable def ofUniformEmbedding (f : α → β) (hf : UniformEmbedding f) : α ≃ᵤ Set.range f where diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index 2389e25ab80ba..71329eeb5d49f 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -135,9 +135,7 @@ uniform convergence noncomputable section -open scoped Classical -open Topology Uniformity Filter - +open scoped Classical Topology Uniformity open Set Filter section TypeAlias @@ -777,6 +775,27 @@ protected theorem gen_mem_nhds (f : α →ᵤ[𝔖] β) (hs : s ∈ 𝔖) {V : S rw [UniformOnFun.nhds_eq] apply_rules [mem_iInf_of_mem, mem_principal_self] +theorem uniformContinuous_ofUniformFun : + UniformContinuous fun f : α →ᵤ β ↦ ofFun 𝔖 (UniformFun.toFun f) := by + simp only [UniformContinuous, UniformOnFun.uniformity_eq, tendsto_iInf, tendsto_principal, + (UniformFun.hasBasis_uniformity _ _).eventually_iff] + exact fun _ _ U hU ↦ ⟨U, hU, fun f hf x _ ↦ hf x⟩ + +/-- The uniformity on `α →ᵤ[𝔖] β` is the same as the uniformity on `α →ᵤ β`, +provided that `Set.univ ∈ 𝔖`. + +Here we formulate it as a `UniformEquiv`. -/ +def uniformEquivUniformFun (h : univ ∈ 𝔖) : (α →ᵤ[𝔖] β) ≃ᵤ (α →ᵤ β) where + toFun f := UniformFun.ofFun <| toFun _ f + invFun f := ofFun _ <| UniformFun.toFun f + left_inv _ := rfl + right_inv _ := rfl + uniformContinuous_toFun := by + simp only [UniformContinuous, (UniformFun.hasBasis_uniformity _ _).tendsto_right_iff] + intro U hU + filter_upwards [UniformOnFun.gen_mem_uniformity _ _ h hU] with f hf x using hf x (mem_univ _) + uniformContinuous_invFun := uniformContinuous_ofUniformFun _ _ + /-- Let `u₁`, `u₂` be two uniform structures on `γ` and `𝔖₁ 𝔖₂ : Set (Set α)`. If `u₁ ≤ u₂` and `𝔖₂ ⊆ 𝔖₁` then `𝒱(α, γ, 𝔖₁, u₁) ≤ 𝒱(α, γ, 𝔖₂, u₂)`. -/ protected theorem mono ⦃u₁ u₂ : UniformSpace γ⦄ (hu : u₁ ≤ u₂) ⦃𝔖₁ 𝔖₂ : Set (Set α)⦄ @@ -794,6 +813,11 @@ theorem uniformContinuous_eval_of_mem {x : α} (hxs : x ∈ s) (hs : s ∈ 𝔖) (UniformOnFun.uniformContinuous_restrict α β 𝔖 hs) #align uniform_on_fun.uniform_continuous_eval_of_mem UniformOnFun.uniformContinuous_eval_of_mem +theorem uniformContinuous_eval_of_mem_sUnion {x : α} (hx : x ∈ ⋃₀ 𝔖) : + UniformContinuous ((Function.eval x : (α → β) → β) ∘ toFun 𝔖) := + let ⟨_s, hs, hxs⟩ := hx + uniformContinuous_eval_of_mem _ _ hxs hs + variable {β} {𝔖} /-- If `u` is a family of uniform structures on `γ`, then @@ -936,6 +960,24 @@ protected lemma continuous_rng_iff {X : Type*} [TopologicalSpace X] {f : X → ( tendstoUniformlyOn_iff_tendstoUniformly_comp_coe, @forall_swap X] rfl +instance [CompleteSpace β] : CompleteSpace (α →ᵤ[𝔖] β) := by + rcases isEmpty_or_nonempty β + · infer_instance + · refine ⟨fun {F} hF ↦ ?_⟩ + have := hF.1 + have : ∀ x ∈ ⋃₀ 𝔖, ∃ y : β, Tendsto (toFun 𝔖 · x) F (𝓝 y) := fun x hx ↦ + CompleteSpace.complete (hF.map (uniformContinuous_eval_of_mem_sUnion _ _ hx)) + choose! g hg using this + use ofFun 𝔖 g + simp_rw [UniformOnFun.nhds_eq_of_basis _ _ uniformity_hasBasis_closed, le_iInf₂_iff, + le_principal_iff] + intro s hs U ⟨hU, hUc⟩ + rcases cauchy_iff.mp hF |>.2 _ <| UniformOnFun.gen_mem_uniformity _ _ hs hU + with ⟨V, hV, hVU⟩ + filter_upwards [hV] with f hf x hx + refine hUc.mem_of_tendsto ((hg x ⟨s, hs, hx⟩).prod_mk_nhds tendsto_const_nhds) ?_ + filter_upwards [hV] with g' hg' using hVU (mk_mem_prod hg' hf) _ hx + /-- The natural bijection between `α → β × γ` and `(α → β) × (α → γ)`, upgraded to a uniform isomorphism between `α →ᵤ[𝔖] β × γ` and `(α →ᵤ[𝔖] β) × (α →ᵤ[𝔖] γ)`. -/ protected def uniformEquivProdArrow [UniformSpace γ] : @@ -1005,3 +1047,10 @@ theorem isClosed_setOf_continuous_of_le [t : TopologicalSpace α] simpa only [continuous_iSup_dom, continuous_coinduced_dom] using fun s hs ↦ (hcont s hs).restrict end UniformOnFun + +namespace UniformFun + +instance {α β : Type*} [UniformSpace β] [CompleteSpace β] : CompleteSpace (α →ᵤ β) := + (UniformOnFun.uniformEquivUniformFun β {univ} (mem_singleton _)).completeSpace_iff.1 inferInstance + +end UniformFun From 9edd2350d9396700cf5a446372b94f082479154c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 8 May 2024 07:05:35 +0000 Subject: [PATCH 29/52] chore: Remove positivity extension for `Finset.card` stub (#12740) The extension was already implemented in #10610. --- Mathlib/Data/Fintype/Card.lean | 28 ---------------------------- 1 file changed, 28 deletions(-) diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index e3d0a1dae4218..8d18900862ab7 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -1293,31 +1293,3 @@ theorem Fintype.induction_subsingleton_or_nontrivial {P : ∀ (α) [Fintype α], rw [hn] at hlt exact ih (Fintype.card β) hlt _ rfl #align fintype.induction_subsingleton_or_nontrivial Fintype.induction_subsingleton_or_nontrivial - -namespace Tactic - ---open Positivity - -private theorem card_univ_pos (α : Type*) [Fintype α] [Nonempty α] : - 0 < (Finset.univ : Finset α).card := - Finset.univ_nonempty.card_pos - ---Porting note(https://github.com/leanprover-community/mathlib4/issues/6038): restore --- /-- Extension for the `positivity` tactic: `Finset.card s` is positive if `s` is nonempty. -/ --- @[positivity] --- unsafe def positivity_finset_card : expr → tactic strictness --- | q(Finset.card $(s)) => do --- let p --- ← -- TODO: Partial decision procedure for `Finset.nonempty` --- to_expr --- ``(Finset.Nonempty $(s)) >>= --- find_assumption --- positive <$> mk_app `` Finset.Nonempty.card_pos [p] --- | q(@Fintype.card $(α) $(i)) => positive <$> mk_mapp `` Fintype.card_pos [α, i, none] --- | e => --- pp e >>= --- fail ∘ --- format.bracket "The expression `" "` isn't of the form `Finset.card s` or `Fintype.card α`" --- #align tactic.positivity_finset_card tactic.positivity_finset_card - -end Tactic From 2b89926ebc4e019b35b0327289ee422461419ef1 Mon Sep 17 00:00:00 2001 From: Etienne Date: Wed, 8 May 2024 08:07:34 +0000 Subject: [PATCH 30/52] feat: continuity of measure without monotonicity assumptions (#12447) Add four theorems about continuity from below/above of measures. Contrary to those already existing, no assumptions about the sequence of sets. Show that the measures of the partial unions/intersections converge to the measure of the union/intersection. --- .../MeasureTheory/Measure/MeasureSpace.lean | 61 ++++++++++++++++++- 1 file changed, 59 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index e9435c60190b6..cde400cdee133 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -513,6 +513,19 @@ theorem measure_iUnion_eq_iSup [Countable ι] {s : ι → Set α} (hd : Directed _ ≤ ⨆ n, μ (t n) := le_iSup (μ ∘ t) N #align measure_theory.measure_Union_eq_supr MeasureTheory.measure_iUnion_eq_iSup +/-- Continuity from below: the measure of the union of a sequence of +(not necessarily measurable) sets is the supremum of the measures of the partial unions. -/ +theorem measure_iUnion_eq_iSup' {α ι : Type*} [MeasurableSpace α] {μ : Measure α} + [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] + {f : ι → Set α} : μ (⋃ i, f i) = ⨆ i, μ (Accumulate f i) := by + have hd : Directed (· ⊆ ·) (Accumulate f) := by + intro i j + rcases directed_of (· ≤ ·) i j with ⟨k, rik, rjk⟩ + exact ⟨k, biUnion_subset_biUnion_left fun l rli ↦ le_trans rli rik, + biUnion_subset_biUnion_left fun l rlj ↦ le_trans rlj rjk⟩ + rw [← iUnion_accumulate] + exact measure_iUnion_eq_iSup hd + theorem measure_biUnion_eq_iSup {s : ι → Set α} {t : Set ι} (ht : t.Countable) (hd : DirectedOn ((· ⊆ ·) on s) t) : μ (⋃ i ∈ t, s i) = ⨆ i ∈ t, μ (s i) := by haveI := ht.toEncodable @@ -542,14 +555,48 @@ theorem measure_iInter_eq_iInf [Countable ι] {s : ι → Set α} (h : ∀ i, Me · exact hd.mono_comp _ fun _ _ => diff_subset_diff_right #align measure_theory.measure_Inter_eq_infi MeasureTheory.measure_iInter_eq_iInf -/-- Continuity from below: the measure of the union of an increasing sequence of measurable sets -is the limit of the measures. -/ +/-- Continuity from above: the measure of the intersection of a sequence of +measurable sets is the infimum of the measures of the partial intersections. -/ +theorem measure_iInter_eq_iInf' {α ι : Type*} [MeasurableSpace α] {μ : Measure α} + [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] + {f : ι → Set α} (h : ∀ i, MeasurableSet (f i)) (hfin : ∃ i, μ (f i) ≠ ∞) : + μ (⋂ i, f i) = ⨅ i, μ (⋂ j ≤ i, f j) := by + let s := fun i ↦ ⋂ j ≤ i, f j + have iInter_eq : ⋂ i, f i = ⋂ i, s i := by + ext x; simp [s]; constructor + · exact fun h _ j _ ↦ h j + · intro h i + rcases directed_of (· ≤ ·) i i with ⟨j, rij, -⟩ + exact h j i rij + have ms : ∀ i, MeasurableSet (s i) := + fun i ↦ MeasurableSet.biInter (countable_univ.mono <| subset_univ _) fun i _ ↦ h i + have hd : Directed (· ⊇ ·) s := by + intro i j + rcases directed_of (· ≤ ·) i j with ⟨k, rik, rjk⟩ + exact ⟨k, biInter_subset_biInter_left fun j rji ↦ le_trans rji rik, + biInter_subset_biInter_left fun i rij ↦ le_trans rij rjk⟩ + have hfin' : ∃ i, μ (s i) ≠ ∞ := by + rcases hfin with ⟨i, hi⟩ + rcases directed_of (· ≤ ·) i i with ⟨j, rij, -⟩ + exact ⟨j, ne_top_of_le_ne_top hi <| measure_mono <| biInter_subset_of_mem rij⟩ + exact iInter_eq ▸ measure_iInter_eq_iInf ms hd hfin' + +/-- Continuity from below: the measure of the union of an increasing sequence of (not necessarily +measurable) sets is the limit of the measures. -/ theorem tendsto_measure_iUnion [Preorder ι] [IsDirected ι (· ≤ ·)] [Countable ι] {s : ι → Set α} (hm : Monotone s) : Tendsto (μ ∘ s) atTop (𝓝 (μ (⋃ n, s n))) := by rw [measure_iUnion_eq_iSup hm.directed_le] exact tendsto_atTop_iSup fun n m hnm => measure_mono <| hm hnm #align measure_theory.tendsto_measure_Union MeasureTheory.tendsto_measure_iUnion +/-- Continuity from below: the measure of the union of a sequence of (not necessarily measurable) +sets is the limit of the measures of the partial unions. -/ +theorem tendsto_measure_iUnion' {α ι : Type*} [MeasurableSpace α] {μ : Measure α} [Countable ι] + [Preorder ι] [IsDirected ι (· ≤ ·)] {f : ι → Set α} : + Tendsto (fun i ↦ μ (Accumulate f i)) atTop (𝓝 (μ (⋃ i, f i))) := by + rw [measure_iUnion_eq_iSup'] + exact tendsto_atTop_iSup fun i j hij ↦ measure_mono <| monotone_accumulate hij + /-- Continuity from above: the measure of the intersection of a decreasing sequence of measurable sets is the limit of the measures. -/ theorem tendsto_measure_iInter [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] {s : ι → Set α} @@ -559,6 +606,16 @@ theorem tendsto_measure_iInter [Countable ι] [Preorder ι] [IsDirected ι (· exact tendsto_atTop_iInf fun n m hnm => measure_mono <| hm hnm #align measure_theory.tendsto_measure_Inter MeasureTheory.tendsto_measure_iInter +/-- Continuity from above: the measure of the intersection of a sequence of measurable +sets such that one has finite measure is the limit of the measures of the partial intersections. -/ +theorem tendsto_measure_iInter' {α ι : Type*} [MeasurableSpace α] {μ : Measure α} [Countable ι] + [Preorder ι] [IsDirected ι (· ≤ ·)] {f : ι → Set α} (hm : ∀ i, MeasurableSet (f i)) + (hf : ∃ i, μ (f i) ≠ ∞) : + Tendsto (fun i ↦ μ (⋂ j ∈ {j | j ≤ i}, f j)) atTop (𝓝 (μ (⋂ i, f i))) := by + rw [measure_iInter_eq_iInf' hm hf] + exact tendsto_atTop_iInf + fun i j hij ↦ measure_mono <| biInter_subset_biInter_left fun k hki ↦ le_trans hki hij + /-- The measure of the intersection of a decreasing sequence of measurable sets indexed by a linear order with first countable topology is the limit of the measures. -/ theorem tendsto_measure_biInter_gt {ι : Type*} [LinearOrder ι] [TopologicalSpace ι] From 994f298c6c02a897c98574d2e11f3ba7b78d0578 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 8 May 2024 08:43:22 +0000 Subject: [PATCH 31/52] chore: Deduplicate `zpow_ne_zero_of_ne_zero` and `zpow_ne_zero` (#12734) Both lemmas are the same up to a binder swap. Also rename `zpow_eq_zero` to `eq_zero_of_zpow_ne_zero`. --- Mathlib/Algebra/GroupWithZero/Power.lean | 31 +++++++++---------- .../Asymptotics/SuperpolynomialDecay.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/ZPow.lean | 2 +- Mathlib/NumberTheory/Padics/PadicNorm.lean | 4 +-- 4 files changed, 18 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Power.lean b/Mathlib/Algebra/GroupWithZero/Power.lean index f81a5389fadd0..554a7d20afebf 100644 --- a/Mathlib/Algebra/GroupWithZero/Power.lean +++ b/Mathlib/Algebra/GroupWithZero/Power.lean @@ -56,7 +56,7 @@ section ZPow open Int -variable {G₀ : Type*} [GroupWithZero G₀] +variable {G₀ : Type*} [GroupWithZero G₀] {a : G₀} {n : ℤ} theorem zero_zpow : ∀ z : ℤ, z ≠ 0 → (0 : G₀) ^ z = 0 | (n : ℕ), h => by @@ -141,30 +141,27 @@ theorem Commute.zpow_zpow_self₀ (a : G₀) (m n : ℤ) : Commute (a ^ m) (a ^ (Commute.refl a).zpow_zpow₀ m n #align commute.zpow_zpow_self₀ Commute.zpow_zpow_self₀ -theorem zpow_ne_zero_of_ne_zero {a : G₀} (ha : a ≠ 0) : ∀ z : ℤ, a ^ z ≠ 0 - | (_ : ℕ) => by - rw [zpow_natCast] - exact pow_ne_zero _ ha - | -[_+1] => by - rw [zpow_negSucc] - exact inv_ne_zero (pow_ne_zero _ ha) -#align zpow_ne_zero_of_ne_zero zpow_ne_zero_of_ne_zero +theorem zpow_ne_zero {a : G₀} : ∀ n : ℤ, a ≠ 0 → a ^ n ≠ 0 + | (_ : ℕ) => by rw [zpow_natCast]; exact pow_ne_zero _ + | -[_+1] => fun ha ↦ by rw [zpow_negSucc]; exact inv_ne_zero (pow_ne_zero _ ha) +#align zpow_ne_zero_of_ne_zero zpow_ne_zero +#align zpow_ne_zero zpow_ne_zero theorem zpow_sub₀ {a : G₀} (ha : a ≠ 0) (z1 z2 : ℤ) : a ^ (z1 - z2) = a ^ z1 / a ^ z2 := by rw [sub_eq_add_neg, zpow_add₀ ha, zpow_neg, div_eq_mul_inv] #align zpow_sub₀ zpow_sub₀ -theorem zpow_eq_zero {x : G₀} {n : ℤ} (h : x ^ n = 0) : x = 0 := - by_contradiction fun hx => zpow_ne_zero_of_ne_zero hx n h -#align zpow_eq_zero zpow_eq_zero +lemma eq_zero_of_zpow_eq_zero : a ^ n = 0 → a = 0 := not_imp_not.1 (zpow_ne_zero _) +#align zpow_eq_zero eq_zero_of_zpow_eq_zero + +@[deprecated (since := "2024-05-07")] alias zpow_ne_zero_of_ne_zero := zpow_ne_zero +@[deprecated (since := "2024-05-07")] alias zpow_eq_zero := eq_zero_of_zpow_eq_zero -theorem zpow_eq_zero_iff {a : G₀} {n : ℤ} (hn : n ≠ 0) : a ^ n = 0 ↔ a = 0 := - ⟨zpow_eq_zero, fun ha => ha.symm ▸ zero_zpow _ hn⟩ +lemma zpow_eq_zero_iff (hn : n ≠ 0) : a ^ n = 0 ↔ a = 0 := + ⟨eq_zero_of_zpow_eq_zero, fun ha => ha.symm ▸ zero_zpow _ hn⟩ #align zpow_eq_zero_iff zpow_eq_zero_iff -theorem zpow_ne_zero {x : G₀} (n : ℤ) : x ≠ 0 → x ^ n ≠ 0 := - mt zpow_eq_zero -#align zpow_ne_zero zpow_ne_zero +lemma zpow_ne_zero_iff (hn : n ≠ 0) : a ^ n ≠ 0 ↔ a ≠ 0 := (zpow_eq_zero_iff hn).ne theorem zpow_neg_mul_zpow_self (n : ℤ) {x : G₀} (h : x ≠ 0) : x ^ (-n) * x ^ n = 1 := by rw [zpow_neg] diff --git a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean index 7f72383ed0cb8..a82de6b54cb05 100644 --- a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean +++ b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean @@ -330,7 +330,7 @@ theorem superpolynomialDecay_iff_isBigO (hk : Tendsto k l atTop) : refine' (superpolynomialDecay_iff_zpow_tendsto_zero f hk).trans _ have hk0 : ∀ᶠ x in l, k x ≠ 0 := hk.eventually_ne_atTop 0 refine' ⟨fun h z => _, fun h z => _⟩ - · refine' isBigO_of_div_tendsto_nhds (hk0.mono fun x hx hxz => absurd (zpow_eq_zero hxz) hx) 0 _ + · refine isBigO_of_div_tendsto_nhds (hk0.mono fun x hx hxz ↦ absurd hxz (zpow_ne_zero _ hx)) 0 ?_ have : (fun a : α => k a ^ z)⁻¹ = fun a : α => k a ^ (-z) := funext fun x => by simp rw [div_eq_mul_inv, mul_comm f, this] exact h (-z) diff --git a/Mathlib/Analysis/Calculus/Deriv/ZPow.lean b/Mathlib/Analysis/Calculus/Deriv/ZPow.lean index cacaf98bf2cd9..38a6e5249671c 100644 --- a/Mathlib/Analysis/Calculus/Deriv/ZPow.lean +++ b/Mathlib/Analysis/Calculus/Deriv/ZPow.lean @@ -47,7 +47,7 @@ theorem hasStrictDerivAt_zpow (m : ℤ) (x : 𝕜) (h : x ≠ 0 ∨ 0 ≤ m) : rcases lt_trichotomy m 0 with (hm | hm | hm) · have hx : x ≠ 0 := h.resolve_right hm.not_le have := (hasStrictDerivAt_inv ?_).scomp _ (this (-m) (neg_pos.2 hm)) <;> - [skip; exact zpow_ne_zero_of_ne_zero hx _] + [skip; exact zpow_ne_zero _ hx] simp only [(· ∘ ·), zpow_neg, one_div, inv_inv, smul_eq_mul] at this convert this using 1 rw [sq, mul_inv, inv_inv, Int.cast_neg, neg_mul, neg_mul_neg, ← zpow_add₀ hx, mul_assoc, ← diff --git a/Mathlib/NumberTheory/Padics/PadicNorm.lean b/Mathlib/NumberTheory/Padics/PadicNorm.lean index 1cc955ce7c3d1..e117e0982b841 100644 --- a/Mathlib/NumberTheory/Padics/PadicNorm.lean +++ b/Mathlib/NumberTheory/Padics/PadicNorm.lean @@ -129,7 +129,7 @@ variable [hp : Fact p.Prime] /-- If `q ≠ 0`, then `padicNorm p q ≠ 0`. -/ protected theorem nonzero {q : ℚ} (hq : q ≠ 0) : padicNorm p q ≠ 0 := by rw [padicNorm.eq_zpow_of_nonzero hq] - apply zpow_ne_zero_of_ne_zero + apply zpow_ne_zero exact mod_cast ne_of_gt hp.1.pos #align padic_norm.nonzero padicNorm.nonzero @@ -138,7 +138,7 @@ theorem zero_of_padicNorm_eq_zero {q : ℚ} (h : padicNorm p q = 0) : q = 0 := b apply by_contradiction; intro hq unfold padicNorm at h; rw [if_neg hq] at h apply absurd h - apply zpow_ne_zero_of_ne_zero + apply zpow_ne_zero exact mod_cast hp.1.ne_zero #align padic_norm.zero_of_padic_norm_eq_zero padicNorm.zero_of_padicNorm_eq_zero From 89b4b590b4938e14211b6296f8ecf0c1940cc999 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 8 May 2024 09:34:49 +0000 Subject: [PATCH 32/52] feat: In characteristic p, `Nat.cast` is injective on `Iio p` (#12542) Add `CharP.natCast_injOn_Iio`. --- Mathlib/Algebra/CharP/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index c0f71ddc226cf..ba9b833ec6219 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -172,6 +172,10 @@ theorem CharP.natCast_eq_natCast [AddMonoidWithOne R] [IsRightCancelAdd R] (p : Nat.sub_add_cancel hle, eq_comm] #align char_p.nat_cast_eq_nat_cast CharP.natCast_eq_natCast +lemma CharP.natCast_injOn_Iio (R) [AddMonoidWithOne R] (p : ℕ) [CharP R p] [IsRightCancelAdd R] : + (Set.Iio p).InjOn ((↑) : ℕ → R) := + fun _a ha _b hb hab ↦ ((natCast_eq_natCast _ _).1 hab).eq_of_lt_of_lt ha hb + theorem CharP.intCast_eq_intCast_mod [AddGroupWithOne R] (p : ℕ) [CharP R p] {a : ℤ} : (a : R) = a % (p : ℤ) := (CharP.intCast_eq_intCast R p).mpr (Int.mod_modEq a p).symm From 2113c23b8292c6f1b71de1b7aa6fca33a69c5148 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 8 May 2024 11:06:09 +0000 Subject: [PATCH 33/52] refactor(CategoryTheory): make Functor.IsEquivalence a Prop (#12462) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `Functor.IsEquivalence` is now the combination of `F.Full`, `F.Faithful` and `F.EssSurj`. That such a functor is an equivalence (i.e. there is a quasi-inverse, etc.) is now contained in the definition `F.asEquivalence`. This refactor is to avoid any diamond in the `IsEquivalence` instances (and also for `IsLeftAdjoint` and `IsRightAdjoint` which are also made `Props`). If data is important, one should use `Equivalence` or `Adjunction` instead. This refactor also avoids some duplication of code between the structures `Equivalence` and the former `IsEquivalence`. `IsLeftAdjoint` and `IsRightAdjoint` are moved to the `Functor` namespace. The namespaces have also been fixed in the file `CategoryTheory.Adjunction.FullyFaithful`. Proofs have been tidied in `CategoryTheory.Monoidal.NaturalTransformation`. Definitions which relied on `IsLeft/RightAdjoint` have been modified in order to contain the data of an adjunction (e.g. `Reflective`, `MonadicRightAdjoint`, etc.). Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../Algebra/Category/AlgebraCat/Basic.lean | 3 +- .../Category/GroupCat/Adjunctions.lean | 16 +- .../Category/GroupCat/ZModuleEquivalence.lean | 3 +- .../Category/ModuleCat/Adjunctions.lean | 4 +- .../Category/ModuleCat/ChangeOfRings.lean | 26 +- .../Category/ModuleCat/Monoidal/Closed.lean | 7 +- .../Algebra/Category/MonCat/Adjunctions.lean | 4 +- .../Algebra/Category/Ring/Adjunctions.lean | 4 +- Mathlib/Algebra/Homology/LocalCohomology.lean | 13 +- Mathlib/AlgebraicGeometry/AffineScheme.lean | 19 +- .../GammaSpecAdjunction.lean | 24 +- .../DoldKan/Compatibility.lean | 49 +- .../DoldKan/Equivalence.lean | 8 +- .../AlgebraicTopology/SimplexCategory.lean | 5 +- .../Adjunction/AdjointFunctorTheorems.lean | 30 +- Mathlib/CategoryTheory/Adjunction/Basic.lean | 201 ++++---- Mathlib/CategoryTheory/Adjunction/Comma.lean | 37 +- .../CategoryTheory/Adjunction/Evaluation.lean | 8 +- .../Adjunction/FullyFaithful.lean | 54 ++- .../CategoryTheory/Adjunction/Lifting.lean | 54 +-- Mathlib/CategoryTheory/Adjunction/Limits.lean | 51 +- Mathlib/CategoryTheory/Adjunction/Over.lean | 4 +- .../CategoryTheory/Adjunction/Reflective.lean | 84 ++-- .../CategoryTheory/Bicategory/SingleObj.lean | 15 +- Mathlib/CategoryTheory/CatCommSq.lean | 4 - Mathlib/CategoryTheory/Closed/Cartesian.lean | 51 +- .../Closed/FunctorCategory.lean | 15 +- Mathlib/CategoryTheory/Closed/Ideal.lean | 78 +-- Mathlib/CategoryTheory/Closed/Monoidal.lean | 88 ++-- Mathlib/CategoryTheory/Closed/Types.lean | 29 +- Mathlib/CategoryTheory/Comma/Arrow.lean | 4 +- Mathlib/CategoryTheory/Comma/Basic.lean | 12 +- Mathlib/CategoryTheory/Comma/Over.lean | 8 +- .../CategoryTheory/Comma/StructuredArrow.lean | 18 +- .../CategoryTheory/ConnectedComponents.lean | 3 +- .../EffectiveEpi/Preserves.lean | 3 +- Mathlib/CategoryTheory/Equivalence.lean | 452 ++++++------------ Mathlib/CategoryTheory/EssentialImage.lean | 9 +- Mathlib/CategoryTheory/Filtered/Basic.lean | 16 +- Mathlib/CategoryTheory/FintypeCat.lean | 5 +- Mathlib/CategoryTheory/Functor/EpiMono.lean | 4 +- Mathlib/CategoryTheory/Functor/Flat.lean | 4 +- .../CategoryTheory/Functor/FullyFaithful.lean | 4 + .../Functor/KanExtension/Basic.lean | 4 +- .../Idempotents/FunctorExtension.lean | 30 +- .../CategoryTheory/Idempotents/Karoubi.lean | 8 +- .../CategoryTheory/Limits/ConeCategory.lean | 12 +- Mathlib/CategoryTheory/Limits/Cones.lean | 14 +- .../Limits/Constructions/Over/Products.lean | 2 +- Mathlib/CategoryTheory/Limits/Final.lean | 36 +- Mathlib/CategoryTheory/Limits/HasLimits.lean | 4 +- Mathlib/CategoryTheory/Limits/IsLimit.lean | 32 +- Mathlib/CategoryTheory/Limits/Over.lean | 4 +- .../Limits/Preserves/Shapes/Zero.lean | 4 +- Mathlib/CategoryTheory/Limits/Presheaf.lean | 23 +- .../CategoryTheory/Limits/Shapes/Images.lean | 5 +- .../Limits/Shapes/NormalMono/Basic.lean | 4 - Mathlib/CategoryTheory/Limits/VanKampen.lean | 3 +- .../Localization/Composition.lean | 1 + .../Localization/Equivalence.lean | 10 +- .../Localization/FiniteProducts.lean | 2 +- .../Localization/LocalizerMorphism.lean | 61 ++- .../Localization/Predicate.lean | 60 +-- Mathlib/CategoryTheory/Monad/Adjunction.lean | 83 +++- Mathlib/CategoryTheory/Monad/Algebra.lean | 28 +- Mathlib/CategoryTheory/Monad/Limits.lean | 58 +-- Mathlib/CategoryTheory/Monad/Monadicity.lean | 328 ++++++------- Mathlib/CategoryTheory/Monoidal/Functor.lean | 40 +- .../Monoidal/NaturalTransformation.lean | 100 ++-- .../Monoidal/OfHasFiniteProducts.lean | 55 ++- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 2 +- .../Monoidal/Rigid/OfEquivalence.lean | 29 +- Mathlib/CategoryTheory/Monoidal/Skeleton.lean | 4 +- .../CategoryTheory/Monoidal/Subcategory.lean | 7 +- .../CategoryTheory/Monoidal/Transport.lean | 8 +- Mathlib/CategoryTheory/Pi/Basic.lean | 2 +- Mathlib/CategoryTheory/Preadditive/Mat.lean | 3 +- Mathlib/CategoryTheory/Shift/Basic.lean | 14 - Mathlib/CategoryTheory/Sites/Adjunction.lean | 8 +- .../Sites/ConcreteSheafification.lean | 6 +- Mathlib/CategoryTheory/Sites/LeftExact.lean | 2 - .../CategoryTheory/Sites/Sheafification.lean | 32 +- Mathlib/CategoryTheory/Skeletal.lean | 33 +- .../CategoryTheory/Subobject/MonoOver.lean | 6 +- Mathlib/CategoryTheory/UnivLE.lean | 3 +- .../RepresentationTheory/Action/Monoidal.lean | 53 +- Mathlib/RepresentationTheory/Rep.lean | 7 +- Mathlib/Topology/Category/CompHaus/Basic.lean | 3 +- Mathlib/Topology/Category/Compactum.lean | 3 +- .../Category/LightProfinite/Basic.lean | 3 +- .../Topology/Category/Profinite/Basic.lean | 2 +- .../Topology/Category/TopCat/Adjunctions.lean | 8 +- Mathlib/Topology/Category/UniformSpace.lean | 4 +- Mathlib/Topology/Sheaves/Functors.lean | 5 +- Mathlib/Topology/Sheaves/Skyscraper.lean | 12 +- 95 files changed, 1306 insertions(+), 1494 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 752e4a9df022f..b5a3d0b496918 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -190,8 +190,7 @@ def adj : free.{u} R ⊣ forget (AlgebraCat.{u} R) := rfl } #align Algebra.adj AlgebraCat.adj -instance : IsRightAdjoint (forget (AlgebraCat.{u} R)) := - ⟨_, adj R⟩ +instance : (forget (AlgebraCat.{u} R)).IsRightAdjoint := (adj R).isRightAdjoint end AlgebraCat diff --git a/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean b/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean index aa5965565a24a..db4c46c6ed250 100644 --- a/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean @@ -81,8 +81,8 @@ def adj : free ⊣ forget AddCommGroupCat.{u} := apply FreeAbelianGroup.lift_comp } #align AddCommGroup.adj AddCommGroupCat.adj -instance : IsRightAdjoint (forget AddCommGroupCat.{u}) := - ⟨_, adj⟩ +instance : (forget AddCommGroupCat.{u}).IsRightAdjoint := + ⟨_, ⟨adj⟩⟩ /-- As an example, we now give a high-powered proof that the monomorphisms in `AddCommGroup` are just the injective functions. @@ -127,8 +127,8 @@ def adj : free ⊣ forget GroupCat.{u} := apply FreeGroup.lift.of } #align Group.adj GroupCat.adj -instance : IsRightAdjoint (forget GroupCat.{u}) := - ⟨_, adj⟩ +instance : (forget GroupCat.{u}).IsRightAdjoint := + ⟨_, ⟨adj⟩⟩ end GroupCat @@ -196,8 +196,8 @@ def GroupCat.forget₂MonAdj : forget₂ GroupCat MonCat ⊣ MonCat.units.{u} wh homEquiv_counit := MonoidHom.ext fun _ => rfl #align Group.forget₂_Mon_adj GroupCat.forget₂MonAdj -instance : IsRightAdjoint MonCat.units.{u} := - ⟨_, GroupCat.forget₂MonAdj⟩ +instance : MonCat.units.{u}.IsRightAdjoint := + ⟨_, ⟨GroupCat.forget₂MonAdj⟩⟩ /-- The functor taking a monoid to its subgroup of units. -/ @[simps] @@ -225,5 +225,5 @@ def CommGroupCat.forget₂CommMonAdj : forget₂ CommGroupCat CommMonCat ⊣ Com homEquiv_counit := MonoidHom.ext fun _ => rfl #align CommGroup.forget₂_CommMon_adj CommGroupCat.forget₂CommMonAdj -instance : IsRightAdjoint CommMonCat.units.{u} := - ⟨_, CommGroupCat.forget₂CommMonAdj⟩ +instance : CommMonCat.units.{u}.IsRightAdjoint := + ⟨_, ⟨CommGroupCat.forget₂CommMonAdj⟩⟩ diff --git a/Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean b/Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean index 720ab949cc2ff..0f87a12105345 100644 --- a/Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean +++ b/Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean @@ -50,8 +50,7 @@ set_option linter.uppercaseLean3 false in #align Module.forget₂_AddCommGroup_ess_surj ModuleCat.forget₂_addCommGroupCat_essSurj noncomputable instance forget₂AddCommGroupIsEquivalence : - (forget₂ (ModuleCat ℤ) AddCommGroupCat.{u}).IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj (forget₂ (ModuleCat ℤ) AddCommGroupCat) + (forget₂ (ModuleCat ℤ) AddCommGroupCat.{u}).IsEquivalence where set_option linter.uppercaseLean3 false in #align Module.forget₂_AddCommGroup_is_equivalence ModuleCat.forget₂AddCommGroupIsEquivalence diff --git a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean index 10ffabe4f6d78..0db6fc36d0e3d 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Adjunctions.lean @@ -57,8 +57,8 @@ def adj : free R ⊣ forget (ModuleCat.{u} R) := (Finsupp.sum_mapDomain_index_addMonoidHom fun y => (smulAddHom R M).flip (g y)).symm } #align Module.adj ModuleCat.adj -instance : IsRightAdjoint (forget (ModuleCat.{u} R)) := - ⟨_, adj R⟩ +instance : (forget (ModuleCat.{u} R)).IsRightAdjoint := + (adj R).isRightAdjoint end diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index e19381b603660..9daa574cffe25 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -242,8 +242,10 @@ abbrev restrictScalarsComp := restrictScalarsComp'.{v} f g _ rfl end -instance restrictScalarsIsEquivalenceOfRingEquiv {R S} [Ring R] [Ring S] (e : R ≃+* S) : - (ModuleCat.restrictScalars e.toRingHom).IsEquivalence where +/-- The equivalence of categories `ModuleCat S ≌ ModuleCat R` induced by `e : R ≃+* S`. -/ +def restrictScalarsEquivalenceOfRingEquiv {R S} [Ring R] [Ring S] (e : R ≃+* S) : + ModuleCat S ≌ ModuleCat R where + functor := ModuleCat.restrictScalars e.toRingHom inverse := ModuleCat.restrictScalars e.symm unitIso := NatIso.ofComponents (fun M ↦ LinearEquiv.toModuleIso' { __ := AddEquiv.refl M @@ -253,6 +255,10 @@ instance restrictScalarsIsEquivalenceOfRingEquiv {R S} [Ring R] [Ring S] (e : R map_smul' := fun r m ↦ congr_arg (· • (_ : M)) (e.left_inv r)}) (by intros; rfl) functor_unitIso_comp := by intros; rfl +instance restrictScalars_isEquivalence_of_ringEquiv {R S} [Ring R] [Ring S] (e : R ≃+* S) : + (ModuleCat.restrictScalars e.toRingHom).IsEquivalence := + (restrictScalarsEquivalenceOfRingEquiv e).isEquivalence_functor + open TensorProduct variable {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) @@ -612,12 +618,12 @@ def restrictCoextendScalarsAdj {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] #align category_theory.Module.restrict_coextend_scalars_adj ModuleCat.restrictCoextendScalarsAdj instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : - CategoryTheory.IsLeftAdjoint (restrictScalars f) := - ⟨_, restrictCoextendScalarsAdj f⟩ + (restrictScalars f).IsLeftAdjoint := + (restrictCoextendScalarsAdj f).isLeftAdjoint instance {R : Type u₁} {S : Type u₂} [Ring R] [Ring S] (f : R →+* S) : - CategoryTheory.IsRightAdjoint (coextendScalars f) := - ⟨_, restrictCoextendScalarsAdj f⟩ + (coextendScalars f).IsRightAdjoint := + (restrictCoextendScalarsAdj f).isRightAdjoint namespace ExtendRestrictScalarsAdj @@ -851,12 +857,12 @@ def extendRestrictScalarsAdj {R : Type u₁} {S : Type u₂} [CommRing R] [CommR #align category_theory.Module.extend_restrict_scalars_adj ModuleCat.extendRestrictScalarsAdj instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) : - CategoryTheory.IsLeftAdjoint (extendScalars f) := - ⟨_, extendRestrictScalarsAdj f⟩ + (extendScalars f).IsLeftAdjoint := + (extendRestrictScalarsAdj f).isLeftAdjoint instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S) : - CategoryTheory.IsRightAdjoint (restrictScalars f) := - ⟨_, extendRestrictScalarsAdj f⟩ + (restrictScalars f).IsRightAdjoint := + (extendRestrictScalarsAdj f).isRightAdjoint noncomputable instance preservesLimitRestrictScalars {R : Type*} {S : Type*} [Ring R] [Ring S] (f : R →+* S) {J : Type*} [Category J] diff --git a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean index 5be7db233669b..810d35d5aaf83 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean @@ -46,16 +46,15 @@ set_option linter.uppercaseLean3 false in instance : MonoidalClosed (ModuleCat.{u} R) where closed M := - { isAdj := - { right := (linearCoyoneda R (ModuleCat.{u} R)).obj (op M) - adj := Adjunction.mkOfHomEquiv + { rightAdj := (linearCoyoneda R (ModuleCat.{u} R)).obj (op M) + adj := Adjunction.mkOfHomEquiv { homEquiv := fun N P => monoidalClosedHomEquiv M N P -- Porting note: this proof was automatic in mathlib3 homEquiv_naturality_left_symm := by intros apply TensorProduct.ext' intro m n - rfl } } } + rfl } } theorem ihom_map_apply {M N P : ModuleCat.{u} R} (f : N ⟶ P) (g : ModuleCat.of R (M ⟶ N)) : (ihom M).map f g = g ≫ f := diff --git a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean index 2502496c84979..8daccf2e56539 100644 --- a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean @@ -80,5 +80,5 @@ def adj : free ⊣ forget MonCat.{u} := homEquiv_naturality_left_symm := fun _ _ => FreeMonoid.hom_eq (fun _ => rfl) } #align adj adj -instance : IsRightAdjoint (forget MonCat.{u}) := - ⟨_, adj⟩ +instance : (forget MonCat.{u}).IsRightAdjoint := + ⟨_, ⟨adj⟩⟩ diff --git a/Mathlib/Algebra/Category/Ring/Adjunctions.lean b/Mathlib/Algebra/Category/Ring/Adjunctions.lean index 817d862c2ce3f..5bc168771ec83 100644 --- a/Mathlib/Algebra/Category/Ring/Adjunctions.lean +++ b/Mathlib/Algebra/Category/Ring/Adjunctions.lean @@ -60,7 +60,7 @@ def adj : free ⊣ forget CommRingCat.{u} := RingHom.ext fun x => eval₂_cast_comp f (Int.castRingHom Y) g x } #align CommRing.adj CommRingCat.adj -instance : IsRightAdjoint (forget CommRingCat.{u}) := - ⟨_, adj⟩ +instance : (forget CommRingCat.{u}).IsRightAdjoint := + ⟨_, ⟨adj⟩⟩ end CommRingCat diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index af4a95306848d..a9373ea015d77 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -273,12 +273,19 @@ def SelfLERadical.cast (hJK : J.radical = K.radical) : SelfLERadical J ⥤ SelfL #align local_cohomology.self_le_radical.cast localCohomology.SelfLERadical.cast -- TODO generalize this to the equivalence of full categories for any `iff`. -instance SelfLERadical.castIsEquivalence (hJK : J.radical = K.radical) : - (SelfLERadical.cast hJK).IsEquivalence where +/-- The equivalence of categories `SelfLERadical J ≌ SelfLERadical K` +when `J.radical = K.radical`. -/ +def SelfLERadical.castEquivalence (hJK : J.radical = K.radical) : + SelfLERadical J ≌ SelfLERadical K where + functor := SelfLERadical.cast hJK inverse := SelfLERadical.cast hJK.symm unitIso := Iso.refl _ counitIso := Iso.refl _ -#align local_cohomology.self_le_radical.cast_is_equivalence localCohomology.SelfLERadical.castIsEquivalence + +instance SelfLERadical.cast_isEquivalence (hJK : J.radical = K.radical) : + (SelfLERadical.cast hJK).IsEquivalence := + (castEquivalence hJK).isEquivalence_functor +#align local_cohomology.self_le_radical.cast_is_equivalence localCohomology.SelfLERadical.cast_isEquivalence /-- The natural isomorphism between local cohomology defined using the `of_self_le_radical` diagram, assuming `J.radical = K.radical`. -/ diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 07c578db76020..f465c6e04db6a 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -69,8 +69,8 @@ def Scheme.isoSpec (X : Scheme) [IsAffine X] : X ≅ Scheme.Spec.obj (op <| Sche /-- Construct an affine scheme from a scheme and the information that it is affine. Also see `AffineScheme.of` for a typeclass version. -/ @[simps] -def AffineScheme.mk (X : Scheme) (h : IsAffine X) : AffineScheme := - ⟨X, @mem_essImage_of_unit_isIso _ _ _ _ _ _ _ h.1⟩ +def AffineScheme.mk (X : Scheme) (_ : IsAffine X) : AffineScheme := + ⟨X, mem_essImage_of_unit_isIso (adj := ΓSpec.adjunction) _⟩ #align algebraic_geometry.AffineScheme.mk AlgebraicGeometry.AffineScheme.mk /-- Construct an affine scheme from a scheme. Also see `AffineScheme.mk` for a non-typeclass @@ -86,7 +86,8 @@ def AffineScheme.ofHom {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : #align algebraic_geometry.AffineScheme.of_hom AlgebraicGeometry.AffineScheme.ofHom theorem mem_Spec_essImage (X : Scheme) : X ∈ Scheme.Spec.essImage ↔ IsAffine X := - ⟨fun h => ⟨Functor.essImage.unit_isIso h⟩, fun h => @mem_essImage_of_unit_isIso _ _ _ _ _ _ X h.1⟩ + ⟨fun h => ⟨Functor.essImage.unit_isIso h⟩, + fun _ => mem_essImage_of_unit_isIso (adj := ΓSpec.adjunction) _⟩ #align algebraic_geometry.mem_Spec_ess_image AlgebraicGeometry.mem_Spec_essImage instance isAffineAffineScheme (X : AffineScheme.{u}) : IsAffine X.obj := @@ -141,10 +142,12 @@ def equivCommRingCat : AffineScheme ≌ CommRingCatᵒᵖ := equivEssImageOfReflective.symm #align algebraic_geometry.AffineScheme.equiv_CommRing AlgebraicGeometry.AffineScheme.equivCommRingCat +instance : Γ.{u}.rightOp.IsEquivalence := equivCommRingCat.isEquivalence_functor + +instance : Γ.{u}.rightOp.op.IsEquivalence := equivCommRingCat.op.isEquivalence_functor + instance ΓIsEquiv : Γ.{u}.IsEquivalence := - haveI : Γ.{u}.rightOp.op.IsEquivalence := - Functor.IsEquivalence.ofEquivalence equivCommRingCat.op - Functor.isEquivalenceTrans Γ.{u}.rightOp.op (opOpEquivalence _).functor + inferInstanceAs (Γ.{u}.rightOp.op ⋙ (opOpEquivalence _).functor).IsEquivalence #align algebraic_geometry.AffineScheme.Γ_is_equiv AlgebraicGeometry.AffineScheme.ΓIsEquiv instance hasColimits : HasColimits AffineScheme.{u} := @@ -156,9 +159,7 @@ instance hasLimits : HasLimits AffineScheme.{u} := by haveI : HasLimits AffineScheme.{u}ᵒᵖᵒᵖ := Limits.hasLimits_op_of_hasColimits exact Adjunction.has_limits_of_equivalence (opOpEquivalence AffineScheme.{u}).inverse -noncomputable instance Γ_preservesLimits : PreservesLimits Γ.{u}.rightOp := - @Adjunction.isEquivalencePreservesLimits _ _ _ _ Γ.rightOp - (Functor.IsEquivalence.ofEquivalence equivCommRingCat) +noncomputable instance Γ_preservesLimits : PreservesLimits Γ.{u}.rightOp := inferInstance noncomputable instance forgetToScheme_preservesLimits : PreservesLimits forgetToScheme := by apply (config := { allowSynthFailures := true }) diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 23558e2471aa4..b8da3f1b38621 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -480,31 +480,31 @@ instance Spec.preservesLimits : Limits.PreservesLimits Scheme.Spec := /-- Spec is a full functor. -/ instance : Spec.toLocallyRingedSpace.Full := - R_full_of_counit_isIso ΓSpec.locallyRingedSpaceAdjunction + ΓSpec.locallyRingedSpaceAdjunction.R_full_of_counit_isIso instance Spec.full : Scheme.Spec.Full := - R_full_of_counit_isIso ΓSpec.adjunction + ΓSpec.adjunction.R_full_of_counit_isIso #align algebraic_geometry.Spec.full AlgebraicGeometry.Spec.full /-- Spec is a faithful functor. -/ instance : Spec.toLocallyRingedSpace.Faithful := - R_faithful_of_counit_isIso ΓSpec.locallyRingedSpaceAdjunction + ΓSpec.locallyRingedSpaceAdjunction.R_faithful_of_counit_isIso instance Spec.faithful : Scheme.Spec.Faithful := - R_faithful_of_counit_isIso ΓSpec.adjunction + ΓSpec.adjunction.R_faithful_of_counit_isIso #align algebraic_geometry.Spec.faithful AlgebraicGeometry.Spec.faithful -instance : IsRightAdjoint Spec.toLocallyRingedSpace := - ⟨_, ΓSpec.locallyRingedSpaceAdjunction⟩ +instance : Spec.toLocallyRingedSpace.IsRightAdjoint := + (ΓSpec.locallyRingedSpaceAdjunction).isRightAdjoint -instance : IsRightAdjoint Scheme.Spec := - ⟨_, ΓSpec.adjunction⟩ +instance : Scheme.Spec.IsRightAdjoint := + (ΓSpec.adjunction).isRightAdjoint -instance : Reflective Spec.toLocallyRingedSpace := - ⟨⟩ +instance : Reflective Spec.toLocallyRingedSpace where + adj := ΓSpec.locallyRingedSpaceAdjunction -instance Spec.reflective : Reflective Scheme.Spec := - ⟨⟩ +instance Spec.reflective : Reflective Scheme.Spec where + adj := ΓSpec.adjunction #align algebraic_geometry.Spec.reflective AlgebraicGeometry.Spec.reflective end AlgebraicGeometry diff --git a/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean b/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean index 950a9be62bd5a..d381654f5453c 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Compatibility.lean @@ -64,10 +64,7 @@ variable {eA} {e'} /-- An intermediate equivalence `A ≅ B'` whose functor is `F` and whose inverse is `e'.inverse ⋙ eA.inverse`. -/ @[simps! functor] -def equivalence₁ : A ≌ B' := - letI : F.IsEquivalence := - Functor.IsEquivalence.ofIso hF (Functor.IsEquivalence.ofEquivalence (equivalence₀ eA e')) - F.asEquivalence +def equivalence₁ : A ≌ B' := (equivalence₀ eA e').changeFunctor hF #align algebraic_topology.dold_kan.compatibility.equivalence₁ AlgebraicTopology.DoldKan.Compatibility.equivalence₁ theorem equivalence₁_inverse : (equivalence₁ hF).inverse = e'.inverse ⋙ eA.inverse := @@ -88,9 +85,7 @@ def equivalence₁CounitIso : (e'.inverse ⋙ eA.inverse) ⋙ F ≅ 𝟭 B' := theorem equivalence₁CounitIso_eq : (equivalence₁ hF).counitIso = equivalence₁CounitIso hF := by ext Y - dsimp [equivalence₁, equivalence₀, Functor.IsEquivalence.inverse, - Functor.IsEquivalence.ofEquivalence] - simp + simp [equivalence₁, equivalence₀] #align algebraic_topology.dold_kan.compatibility.equivalence₁_counit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalence₁CounitIso_eq /-- The unit isomorphism of the equivalence `equivalence₁` between `A` and `B'`. -/ @@ -107,8 +102,7 @@ def equivalence₁UnitIso : 𝟭 A ≅ F ⋙ e'.inverse ⋙ eA.inverse := theorem equivalence₁UnitIso_eq : (equivalence₁ hF).unitIso = equivalence₁UnitIso hF := by ext X - dsimp [equivalence₁, NatIso.hcomp, Functor.IsEquivalence.ofEquivalence] - simp + simp [equivalence₁] #align algebraic_topology.dold_kan.compatibility.equivalence₁_unit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalence₁UnitIso_eq /-- An intermediate equivalence `A ≅ B` obtained as the composition of `equivalence₁` and @@ -169,15 +163,12 @@ variable {eB} whose inverse is `G : B ≅ A`. -/ @[simps! inverse] def equivalence : A ≌ B := - letI : G.IsEquivalence := by - refine' Functor.IsEquivalence.ofIso _ - (Functor.IsEquivalence.ofEquivalence (equivalence₂ eB hF).symm) - calc - eB.functor ⋙ e'.inverse ⋙ eA.inverse ≅ (eB.functor ⋙ e'.inverse) ⋙ eA.inverse := Iso.refl _ - _ ≅ (G ⋙ eA.functor) ⋙ eA.inverse := isoWhiskerRight hG _ - _ ≅ G ⋙ 𝟭 A := isoWhiskerLeft _ eA.unitIso.symm - _ ≅ G := Functor.rightUnitor G - G.asEquivalence.symm + ((equivalence₂ eB hF).changeInverse + (calc eB.functor ⋙ e'.inverse ⋙ eA.inverse ≅ + (eB.functor ⋙ e'.inverse) ⋙ eA.inverse := (Functor.associator _ _ _).symm + _ ≅ (G ⋙ eA.functor) ⋙ eA.inverse := isoWhiskerRight hG _ + _ ≅ G ⋙ 𝟭 A := isoWhiskerLeft _ eA.unitIso.symm + _ ≅ G := G.rightUnitor)) #align algebraic_topology.dold_kan.compatibility.equivalence AlgebraicTopology.DoldKan.Compatibility.equivalence theorem equivalence_functor : (equivalence hF hG).functor = F ⋙ eB.inverse := @@ -223,16 +214,14 @@ variable {η hF hG} theorem equivalenceCounitIso_eq : (equivalence hF hG).counitIso = equivalenceCounitIso η := by ext1; apply NatTrans.ext; ext Y - dsimp [equivalence, Functor.asEquivalence, Functor.IsEquivalence.ofEquivalence] - rw [equivalenceCounitIso_hom_app, Functor.IsEquivalence.ofIso_unitIso_inv_app] - dsimp - simp only [comp_id, id_comp, F.map_comp, assoc, - equivalence₂CounitIso_eq, equivalence₂CounitIso_hom_app, - ← eB.inverse.map_comp_assoc, ← τ₀_hom_app, hη, τ₁_hom_app] + dsimp [equivalence] + simp only [comp_id, id_comp, Functor.map_comp, equivalence₂CounitIso_eq, + equivalence₂CounitIso_hom_app, assoc, equivalenceCounitIso_hom_app] + simp only [← eB.inverse.map_comp_assoc, ← τ₀_hom_app, hη, τ₁_hom_app] erw [hF.inv.naturality_assoc, hF.inv.naturality_assoc] dsimp congr 2 - simp only [assoc, ← e'.functor.map_comp_assoc, Equivalence.fun_inv_map, + simp only [← e'.functor.map_comp_assoc, Equivalence.fun_inv_map, assoc, Iso.inv_hom_id_app_assoc, hG.inv_hom_id_app] dsimp rw [comp_id, eA.functor_unitIso_comp, e'.functor.map_id, id_comp, hF.inv_hom_id_app_assoc] @@ -274,13 +263,11 @@ variable {ε hF hG} theorem equivalenceUnitIso_eq : (equivalence hF hG).unitIso = equivalenceUnitIso hG ε := by ext1; apply NatTrans.ext; ext X - dsimp [equivalence, Functor.asEquivalence, Functor.IsEquivalence.ofEquivalence, - Functor.IsEquivalence.inverse] - rw [Functor.IsEquivalence.ofIso_counitIso_inv_app] - dsimp - erw [id_comp, comp_id] + dsimp [equivalence] + simp only [assoc, comp_id, equivalenceUnitIso_hom_app] + erw [id_comp] simp only [equivalence₂UnitIso_eq eB hF, equivalence₂UnitIso_hom_app, - assoc, equivalenceUnitIso_hom_app, ← eA.inverse.map_comp_assoc, ← hε, υ_hom_app] + ← eA.inverse.map_comp_assoc, assoc, ← hε, υ_hom_app] #align algebraic_topology.dold_kan.compatibility.equivalence_unit_iso_eq AlgebraicTopology.DoldKan.Compatibility.equivalenceUnitIso_eq end Compatibility diff --git a/Mathlib/AlgebraicTopology/DoldKan/Equivalence.lean b/Mathlib/AlgebraicTopology/DoldKan/Equivalence.lean index 4a21fa1269dde..2f90ef1515b8c 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Equivalence.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Equivalence.lean @@ -165,12 +165,8 @@ set_option linter.uppercaseLean3 false in /-- The Dold-Kan equivalence for abelian categories -/ @[simps! functor] -def equivalence : SimplicialObject A ≌ ChainComplex A ℕ := by - let F : SimplicialObject A ⥤ _ := Idempotents.DoldKan.N - let hF : F.IsEquivalence := Functor.IsEquivalence.ofEquivalence Idempotents.DoldKan.equivalence - letI : (N : SimplicialObject A ⥤ _).IsEquivalence := - Functor.IsEquivalence.ofIso comparisonN.symm hF - exact N.asEquivalence +def equivalence : SimplicialObject A ≌ ChainComplex A ℕ := + (Idempotents.DoldKan.equivalence (C := A)).changeFunctor comparisonN.symm #align category_theory.abelian.dold_kan.equivalence CategoryTheory.Abelian.DoldKan.equivalence theorem equivalence_inverse : (equivalence : SimplicialObject A ≌ _).inverse = Γ := diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index b2264830ae014..d3b88ac4859c3 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -490,8 +490,7 @@ instance : skeletalFunctor.EssSurj where show f (f.symm i) ≤ f (f.symm j) simpa only [OrderIso.apply_symm_apply]⟩⟩ -noncomputable instance isEquivalence : skeletalFunctor.IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj skeletalFunctor +noncomputable instance isEquivalence : skeletalFunctor.IsEquivalence where #align simplex_category.skeletal_functor.is_equivalence SimplexCategory.SkeletalFunctor.isEquivalence end SkeletalFunctor @@ -506,7 +505,7 @@ end Skeleton /-- `SimplexCategory` is a skeleton of `NonemptyFinLinOrd`. -/ -noncomputable def isSkeletonOf : +lemma isSkeletonOf : IsSkeletonOf NonemptyFinLinOrd SimplexCategory skeletalFunctor where skel := skeletal eqv := SkeletalFunctor.isEquivalence diff --git a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean index ffafad2a03199..c3538f767cdc7 100644 --- a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean +++ b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean @@ -66,20 +66,20 @@ variable {D : Type u} [Category.{v} D] variable (G : D ⥤ C) /-- If `G : D ⥤ C` is a right adjoint it satisfies the solution set condition. -/ -theorem solutionSetCondition_of_isRightAdjoint [IsRightAdjoint G] : SolutionSetCondition G := by +theorem solutionSetCondition_of_isRightAdjoint [G.IsRightAdjoint] : SolutionSetCondition G := by intro A refine' - ⟨PUnit, fun _ => (leftAdjoint G).obj A, fun _ => (Adjunction.ofRightAdjoint G).unit.app A, _⟩ + ⟨PUnit, fun _ => G.leftAdjoint.obj A, fun _ => (Adjunction.ofIsRightAdjoint G).unit.app A, _⟩ intro B h - refine' ⟨PUnit.unit, ((Adjunction.ofRightAdjoint G).homEquiv _ _).symm h, _⟩ + refine' ⟨PUnit.unit, ((Adjunction.ofIsRightAdjoint G).homEquiv _ _).symm h, _⟩ rw [← Adjunction.homEquiv_unit, Equiv.apply_symm_apply] #align category_theory.solution_set_condition_of_is_right_adjoint CategoryTheory.solutionSetCondition_of_isRightAdjoint /-- The general adjoint functor theorem says that if `G : D ⥤ C` preserves limits and `D` has them, if `G` satisfies the solution set condition then `G` is a right adjoint. -/ -noncomputable def isRightAdjointOfPreservesLimitsOfSolutionSetCondition [HasLimits D] - [PreservesLimits G] (hG : SolutionSetCondition G) : IsRightAdjoint G := by +lemma isRightAdjoint_of_preservesLimits_of_solutionSetCondition [HasLimits D] + [PreservesLimits G] (hG : SolutionSetCondition G) : G.IsRightAdjoint := by refine @isRightAdjointOfStructuredArrowInitials _ _ _ _ G ?_ intro A specialize hG A @@ -91,7 +91,7 @@ noncomputable def isRightAdjointOfPreservesLimitsOfSolutionSetCondition [HasLimi exact ⟨i, ⟨StructuredArrow.homMk _ t⟩⟩ obtain ⟨T, hT⟩ := has_weakly_initial_of_weakly_initial_set_and_hasProducts hB' apply hasInitial_of_weakly_initial_and_hasWideEqualizers hT -#align category_theory.is_right_adjoint_of_preserves_limits_of_solution_set_condition CategoryTheory.isRightAdjointOfPreservesLimitsOfSolutionSetCondition +#align category_theory.is_right_adjoint_of_preserves_limits_of_solution_set_condition CategoryTheory.isRightAdjoint_of_preservesLimits_of_solutionSetCondition end GeneralAdjointFunctorTheorem @@ -102,24 +102,24 @@ variable {D : Type u'} [Category.{v} D] /-- The special adjoint functor theorem: if `G : D ⥤ C` preserves limits and `D` is complete, well-powered and has a small coseparating set, then `G` has a left adjoint. -/ -noncomputable def isRightAdjointOfPreservesLimitsOfIsCoseparating [HasLimits D] [WellPowered D] +lemma isRightAdjoint_of_preservesLimits_of_isCoseparating [HasLimits D] [WellPowered D] {𝒢 : Set D} [Small.{v} 𝒢] (h𝒢 : IsCoseparating 𝒢) (G : D ⥤ C) [PreservesLimits G] : - IsRightAdjoint G := + G.IsRightAdjoint := have : ∀ A, HasInitial (StructuredArrow A G) := fun A => hasInitial_of_isCoseparating (StructuredArrow.isCoseparating_proj_preimage A G h𝒢) isRightAdjointOfStructuredArrowInitials _ -#align category_theory.is_right_adjoint_of_preserves_limits_of_is_coseparating CategoryTheory.isRightAdjointOfPreservesLimitsOfIsCoseparating +#align category_theory.is_right_adjoint_of_preserves_limits_of_is_coseparating CategoryTheory.isRightAdjoint_of_preservesLimits_of_isCoseparating /-- The special adjoint functor theorem: if `F : C ⥤ D` preserves colimits and `C` is cocomplete, well-copowered and has a small separating set, then `F` has a right adjoint. -/ -noncomputable def isLeftAdjointOfPreservesColimitsOfIsSeparatig [HasColimits C] [WellPowered Cᵒᵖ] +lemma isLeftAdjoint_of_preservesColimits_of_isSeparating [HasColimits C] [WellPowered Cᵒᵖ] {𝒢 : Set C} [Small.{v} 𝒢] (h𝒢 : IsSeparating 𝒢) (F : C ⥤ D) [PreservesColimits F] : - IsLeftAdjoint F := + F.IsLeftAdjoint := have : ∀ A, HasTerminal (CostructuredArrow F A) := fun A => hasTerminal_of_isSeparating (CostructuredArrow.isSeparating_proj_preimage F A h𝒢) - isLeftAdjointOfCostructuredArrowTerminals _ -#align category_theory.is_left_adjoint_of_preserves_colimits_of_is_separatig CategoryTheory.isLeftAdjointOfPreservesColimitsOfIsSeparatig + isLeftAdjoint_of_costructuredArrowTerminals _ +#align category_theory.is_left_adjoint_of_preserves_colimits_of_is_separatig CategoryTheory.isLeftAdjoint_of_preservesColimits_of_isSeparating end SpecialAdjointFunctorTheorem @@ -131,7 +131,7 @@ theorem hasColimits_of_hasLimits_of_isCoseparating [HasLimits C] [WellPowered C] [Small.{v} 𝒢] (h𝒢 : IsCoseparating 𝒢) : HasColimits C := { has_colimits_of_shape := fun _ _ => hasColimitsOfShape_iff_isRightAdjoint_const.2 - ⟨isRightAdjointOfPreservesLimitsOfIsCoseparating h𝒢 _⟩ } + (isRightAdjoint_of_preservesLimits_of_isCoseparating h𝒢 _) } #align category_theory.limits.has_colimits_of_has_limits_of_is_coseparating CategoryTheory.Limits.hasColimits_of_hasLimits_of_isCoseparating /-- A consequence of the special adjoint functor theorem: if `C` is cocomplete, well-copowered and @@ -140,7 +140,7 @@ theorem hasLimits_of_hasColimits_of_isSeparating [HasColimits C] [WellPowered C [Small.{v} 𝒢] (h𝒢 : IsSeparating 𝒢) : HasLimits C := { has_limits_of_shape := fun _ _ => hasLimitsOfShape_iff_isLeftAdjoint_const.2 - ⟨isLeftAdjointOfPreservesColimitsOfIsSeparatig h𝒢 _⟩ } + (isLeftAdjoint_of_preservesColimits_of_isSeparating h𝒢 _) } #align category_theory.limits.has_limits_of_has_colimits_of_is_separating CategoryTheory.Limits.hasLimits_of_hasColimits_of_isSeparating end Limits diff --git a/Mathlib/CategoryTheory/Adjunction/Basic.lean b/Mathlib/CategoryTheory/Adjunction/Basic.lean index 20ce9ae8a12d5..e0e8fc72512ae 100644 --- a/Mathlib/CategoryTheory/Adjunction/Basic.lean +++ b/Mathlib/CategoryTheory/Adjunction/Basic.lean @@ -22,9 +22,9 @@ We provide various useful constructors: * `adjunctionOfEquivLeft` / `adjunctionOfEquivRight` witness that these constructions give adjunctions. -There are also typeclasses `IsLeftAdjoint` / `IsRightAdjoint`, carrying data witnessing -that a given functor is a left or right adjoint. -Given `[IsLeftAdjoint F]`, a right adjoint of `F` can be constructed as `rightAdjoint F`. +There are also typeclasses `IsLeftAdjoint` / `IsRightAdjoint`, which asserts the +existence of a adjoint functor. Given `[F.IsLeftAdjoint]`, a chosen right +adjoint can be obtained as `F.rightAdjoint`. `Adjunction.comp` composes adjunctions. @@ -81,43 +81,41 @@ structure Adjunction (F : C ⥤ D) (G : D ⥤ C) where /-- The notation `F ⊣ G` stands for `Adjunction F G` representing that `F` is left adjoint to `G` -/ infixl:15 " ⊣ " => Adjunction -/-- A class giving a chosen right adjoint to the functor `left`. -/ -class IsLeftAdjoint (left : C ⥤ D) where - /-- The right adjoint to `left` -/ - right : D ⥤ C - /-- The adjunction between `left` and `right` -/ - adj : left ⊣ right -#align category_theory.is_left_adjoint CategoryTheory.IsLeftAdjoint - -/-- A class giving a chosen left adjoint to the functor `right`. -/ -class IsRightAdjoint (right : D ⥤ C) where - /-- The left adjoint to `right` -/ - left : C ⥤ D - /-- The adjunction between `left` and `right` -/ - adj : left ⊣ right -#align category_theory.is_right_adjoint CategoryTheory.IsRightAdjoint - -/-- Extract the left adjoint from the instance giving the chosen adjoint. -/ -def leftAdjoint (R : D ⥤ C) [IsRightAdjoint R] : C ⥤ D := - IsRightAdjoint.left R -#align category_theory.left_adjoint CategoryTheory.leftAdjoint - -/-- Extract the right adjoint from the instance giving the chosen adjoint. -/ -def rightAdjoint (L : C ⥤ D) [IsLeftAdjoint L] : D ⥤ C := - IsLeftAdjoint.right L -#align category_theory.right_adjoint CategoryTheory.rightAdjoint +namespace Functor + +/-- A class asserting the existence of a right adjoint. -/ +class IsLeftAdjoint (left : C ⥤ D) : Prop where + exists_rightAdjoint : ∃ (right : D ⥤ C), Nonempty (left ⊣ right) +#align category_theory.is_left_adjoint CategoryTheory.Functor.IsLeftAdjoint + +/-- A class asserting the existence of a left adjoint. -/ +class IsRightAdjoint (right : D ⥤ C) : Prop where + exists_leftAdjoint : ∃ (left : C ⥤ D), Nonempty (left ⊣ right) +#align category_theory.is_right_adjoint CategoryTheory.Functor.IsRightAdjoint + +/-- A chosen left adjoint to a functor that is a right adjoint. -/ +noncomputable def leftAdjoint (R : D ⥤ C) [IsRightAdjoint R] : C ⥤ D := + (IsRightAdjoint.exists_leftAdjoint (right := R)).choose +#align category_theory.left_adjoint CategoryTheory.Functor.leftAdjoint + +/-- A chosen right adjoint to a functor that is a left adjoint. -/ +noncomputable def rightAdjoint (L : C ⥤ D) [IsLeftAdjoint L] : D ⥤ C := + (IsLeftAdjoint.exists_rightAdjoint (left := L)).choose +#align category_theory.right_adjoint CategoryTheory.Functor.rightAdjoint + +end Functor /-- The adjunction associated to a functor known to be a left adjoint. -/ -def Adjunction.ofLeftAdjoint (left : C ⥤ D) [IsLeftAdjoint left] : - Adjunction left (rightAdjoint left) := - IsLeftAdjoint.adj -#align category_theory.adjunction.of_left_adjoint CategoryTheory.Adjunction.ofLeftAdjoint +noncomputable def Adjunction.ofIsLeftAdjoint (left : C ⥤ D) [left.IsLeftAdjoint] : + left ⊣ left.rightAdjoint := + Functor.IsLeftAdjoint.exists_rightAdjoint.choose_spec.some +#align category_theory.adjunction.of_left_adjoint CategoryTheory.Adjunction.ofIsLeftAdjoint /-- The adjunction associated to a functor known to be a right adjoint. -/ -def Adjunction.ofRightAdjoint (right : C ⥤ D) [IsRightAdjoint right] : - Adjunction (leftAdjoint right) right := - IsRightAdjoint.adj -#align category_theory.adjunction.of_right_adjoint CategoryTheory.Adjunction.ofRightAdjoint +noncomputable def Adjunction.ofIsRightAdjoint (right : C ⥤ D) [right.IsRightAdjoint] : + right.leftAdjoint ⊣ right := + Functor.IsRightAdjoint.exists_leftAdjoint.choose_spec.some +#align category_theory.adjunction.of_right_adjoint CategoryTheory.Adjunction.ofIsRightAdjoint namespace Adjunction @@ -130,7 +128,13 @@ attribute [simp] homEquiv_unit homEquiv_counit section -variable {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) {X' X : C} {Y Y' : D} +variable {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) + +lemma isLeftAdjoint : F.IsLeftAdjoint := ⟨_, ⟨adj⟩⟩ + +lemma isRightAdjoint : G.IsRightAdjoint := ⟨_, ⟨adj⟩⟩ + +variable {X' X : C} {Y Y' : D} theorem homEquiv_id (X : C) : adj.homEquiv X _ (𝟙 _) = adj.unit.app X := by simp #align category_theory.adjunction.hom_equiv_id CategoryTheory.Adjunction.homEquiv_id @@ -437,20 +441,6 @@ def ofNatIsoRight {F : C ⥤ D} {G H : D ⥤ C} (adj : F ⊣ G) (iso : G ≅ H) { homEquiv := fun X Y => (adj.homEquiv X Y).trans (equivHomsetRightOfNatIso iso) } #align category_theory.adjunction.of_nat_iso_right CategoryTheory.Adjunction.ofNatIsoRight -/-- Transport being a right adjoint along a natural isomorphism. -/ -def rightAdjointOfNatIso {F G : C ⥤ D} (h : F ≅ G) [r : IsRightAdjoint F] : IsRightAdjoint G - where - left := r.left - adj := ofNatIsoRight r.adj h -#align category_theory.adjunction.right_adjoint_of_nat_iso CategoryTheory.Adjunction.rightAdjointOfNatIso - -/-- Transport being a left adjoint along a natural isomorphism. -/ -def leftAdjointOfNatIso {F G : C ⥤ D} (h : F ≅ G) [r : IsLeftAdjoint F] : IsLeftAdjoint G - where - right := r.right - adj := ofNatIsoLeft r.adj h -#align category_theory.adjunction.left_adjoint_of_nat_iso CategoryTheory.Adjunction.leftAdjointOfNatIso - section variable {E : Type u₃} [ℰ : Category.{v₃} E] {H : D ⥤ E} {I : E ⥤ D} @@ -467,22 +457,6 @@ def comp (adj₁ : F ⊣ G) (adj₂ : H ⊣ I) : F ⋙ H ⊣ I ⋙ G (Functor.associator _ _ _).hom ≫ (whiskerLeft I <| whiskerRight adj₁.counit H) ≫ adj₂.counit #align category_theory.adjunction.comp CategoryTheory.Adjunction.comp -/-- If `F` and `G` are left adjoints then `F ⋙ G` is a left adjoint too. -/ -instance leftAdjointOfComp {E : Type u₃} [Category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E) - [Fl : IsLeftAdjoint F] [Gl : IsLeftAdjoint G] : IsLeftAdjoint (F ⋙ G) - where - right := Gl.right ⋙ Fl.right - adj := Fl.adj.comp Gl.adj -#align category_theory.adjunction.left_adjoint_of_comp CategoryTheory.Adjunction.leftAdjointOfComp - -/-- If `F` and `G` are right adjoints then `F ⋙ G` is a right adjoint too. -/ -instance rightAdjointOfComp {E : Type u₃} [Category.{v₃} E] {F : C ⥤ D} {G : D ⥤ E} - [Fr : IsRightAdjoint F] [Gr : IsRightAdjoint G] : IsRightAdjoint (F ⋙ G) - where - left := Gr.left ⋙ Fr.left - adj := Gr.adj.comp Fr.adj -#align category_theory.adjunction.right_adjoint_of_comp CategoryTheory.Adjunction.rightAdjointOfComp - end section ConstructLeft @@ -590,27 +564,28 @@ noncomputable def toEquivalence (adj : F ⊣ G) [∀ X, IsIso (adj.unit.app X)] counitIso := NatIso.ofComponents fun Y => asIso (adj.counit.app Y) #align category_theory.adjunction.to_equivalence CategoryTheory.Adjunction.toEquivalence +end Adjunction + +open Adjunction + /-- If the unit and counit for the adjunction corresponding to a right adjoint functor are (pointwise) isomorphisms, then the functor is an equivalence of categories. -/ -@[simps!] -noncomputable def isRightAdjointToIsEquivalence [IsRightAdjoint G] - [∀ X, IsIso ((Adjunction.ofRightAdjoint G).unit.app X)] - [∀ Y, IsIso ((Adjunction.ofRightAdjoint G).counit.app Y)] : G.IsEquivalence := - Functor.IsEquivalence.ofEquivalenceInverse (Adjunction.ofRightAdjoint G).toEquivalence -#align category_theory.adjunction.is_right_adjoint_to_is_equivalence CategoryTheory.Adjunction.isRightAdjointToIsEquivalence - -end Adjunction - -open Adjunction +lemma Functor.isEquivalence_of_isRightAdjoint (G : C ⥤ D) [IsRightAdjoint G] + [∀ X, IsIso ((Adjunction.ofIsRightAdjoint G).unit.app X)] + [∀ Y, IsIso ((Adjunction.ofIsRightAdjoint G).counit.app Y)] : G.IsEquivalence := + (Adjunction.ofIsRightAdjoint G).toEquivalence.isEquivalence_inverse +#align category_theory.adjunction.is_right_adjoint_to_is_equivalence CategoryTheory.Functor.isEquivalence_of_isRightAdjoint namespace Equivalence +variable (e : C ≌ D) + /-- The adjunction given by an equivalence of categories. (To obtain the opposite adjunction, simply use `e.symm.toAdjunction`. -/ @[pp_dot, simps! unit counit] -def toAdjunction (e : C ≌ D) : e.functor ⊣ e.inverse := +def toAdjunction : e.functor ⊣ e.inverse := mkOfUnitCounit ⟨e.unit, e.counit, by ext @@ -622,42 +597,68 @@ def toAdjunction (e : C ≌ D) : e.functor ⊣ e.inverse := simp only [id_comp] exact e.unit_inverse_comp _⟩ #align category_theory.equivalence.to_adjunction CategoryTheory.Equivalence.toAdjunction - #align category_theory.equivalence.as_equivalence_to_adjunction_unit CategoryTheory.Equivalence.toAdjunction_unitₓ #align category_theory.equivalence.as_equivalence_to_adjunction_counit CategoryTheory.Equivalence.toAdjunction_counitₓ +lemma isLeftAdjoint_functor : e.functor.IsLeftAdjoint where + exists_rightAdjoint := ⟨_, ⟨e.toAdjunction⟩⟩ + +lemma isRightAdjoint_inverse : e.inverse.IsRightAdjoint where + exists_leftAdjoint := ⟨_, ⟨e.toAdjunction⟩⟩ + +lemma isLeftAdjoint_inverse : e.inverse.IsLeftAdjoint := + e.symm.isLeftAdjoint_functor + +lemma isRightAdjoint_functor : e.functor.IsRightAdjoint := + e.symm.isRightAdjoint_inverse + end Equivalence namespace Functor +/-- If `F` and `G` are left adjoints then `F ⋙ G` is a left adjoint too. -/ +instance isLeftAdjoint_comp {E : Type u₃} [Category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E) + [F.IsLeftAdjoint] [G.IsLeftAdjoint] : (F ⋙ G).IsLeftAdjoint where + exists_rightAdjoint := + ⟨_, ⟨(Adjunction.ofIsLeftAdjoint F).comp (Adjunction.ofIsLeftAdjoint G)⟩⟩ +#align category_theory.adjunction.left_adjoint_of_comp CategoryTheory.Functor.isLeftAdjoint_comp + +/-- If `F` and `G` are right adjoints then `F ⋙ G` is a right adjoint too. -/ +instance isRightAdjoint_comp {E : Type u₃} [Category.{v₃} E] {F : C ⥤ D} {G : D ⥤ E} + [IsRightAdjoint F] [IsRightAdjoint G] : IsRightAdjoint (F ⋙ G) where + exists_leftAdjoint := + ⟨_, ⟨(Adjunction.ofIsRightAdjoint G).comp (Adjunction.ofIsRightAdjoint F)⟩⟩ +#align category_theory.adjunction.right_adjoint_of_comp CategoryTheory.Functor.isRightAdjoint_comp + +/-- Transport being a right adjoint along a natural isomorphism. -/ +lemma isRightAdjoint_of_iso {F G : C ⥤ D} (h : F ≅ G) [F.IsRightAdjoint] : + IsRightAdjoint G where + exists_leftAdjoint := ⟨_, ⟨(Adjunction.ofIsRightAdjoint F).ofNatIsoRight h⟩⟩ +#align category_theory.adjunction.right_adjoint_of_nat_iso CategoryTheory.Functor.isRightAdjoint_of_iso + +/-- Transport being a left adjoint along a natural isomorphism. -/ +lemma isLeftAdjoint_of_iso {F G : C ⥤ D} (h : F ≅ G) [IsLeftAdjoint F] : + IsLeftAdjoint G where + exists_rightAdjoint := ⟨_, ⟨(Adjunction.ofIsLeftAdjoint F).ofNatIsoLeft h⟩⟩ +#align category_theory.adjunction.left_adjoint_of_nat_iso CategoryTheory.Functor.isLeftAdjoint_of_iso + + /-- An equivalence `E` is left adjoint to its inverse. -/ -def adjunction (E : C ⥤ D) [IsEquivalence E] : E ⊣ E.inv := +noncomputable def adjunction (E : C ⥤ D) [IsEquivalence E] : E ⊣ E.inv := E.asEquivalence.toAdjunction #align category_theory.functor.adjunction CategoryTheory.Functor.adjunction /-- If `F` is an equivalence, it's a left adjoint. -/ -instance (priority := 10) leftAdjointOfEquivalence {F : C ⥤ D} [IsEquivalence F] : IsLeftAdjoint F - where - right := _ - adj := Functor.adjunction F -#align category_theory.functor.left_adjoint_of_equivalence CategoryTheory.Functor.leftAdjointOfEquivalence - -@[simp] -theorem rightAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : rightAdjoint F = inv F := - rfl -#align category_theory.functor.right_adjoint_of_is_equivalence CategoryTheory.Functor.rightAdjoint_of_isEquivalence +instance (priority := 10) isLeftAdjoint_of_isEquivalence {F : C ⥤ D} [F.IsEquivalence] : + IsLeftAdjoint F := + F.asEquivalence.isLeftAdjoint_functor +#align category_theory.functor.left_adjoint_of_equivalence CategoryTheory.Functor.isLeftAdjoint_of_isEquivalence /-- If `F` is an equivalence, it's a right adjoint. -/ -instance (priority := 10) rightAdjointOfEquivalence {F : C ⥤ D} [IsEquivalence F] : IsRightAdjoint F - where - left := _ - adj := Functor.adjunction F.inv -#align category_theory.functor.right_adjoint_of_equivalence CategoryTheory.Functor.rightAdjointOfEquivalence - -@[simp] -theorem leftAdjoint_of_isEquivalence {F : C ⥤ D} [IsEquivalence F] : leftAdjoint F = inv F := - rfl -#align category_theory.functor.left_adjoint_of_is_equivalence CategoryTheory.Functor.leftAdjoint_of_isEquivalence +instance (priority := 10) isRightAdjoint_of_isEquivalence {F : C ⥤ D} [F.IsEquivalence] : + IsRightAdjoint F := + F.asEquivalence.isRightAdjoint_functor +#align category_theory.functor.right_adjoint_of_equivalence CategoryTheory.Functor.isRightAdjoint_of_isEquivalence end Functor diff --git a/Mathlib/CategoryTheory/Adjunction/Comma.lean b/Mathlib/CategoryTheory/Adjunction/Comma.lean index c26283b0bebff..c8a3f07bbf56d 100644 --- a/Mathlib/CategoryTheory/Adjunction/Comma.lean +++ b/Mathlib/CategoryTheory/Adjunction/Comma.lean @@ -76,10 +76,8 @@ def adjunctionOfStructuredArrowInitials : leftAdjointOfStructuredArrowInitials G #align category_theory.adjunction_of_structured_arrow_initials CategoryTheory.adjunctionOfStructuredArrowInitials /-- If each structured arrow category on `G` has an initial object, `G` is a right adjoint. -/ -def isRightAdjointOfStructuredArrowInitials : IsRightAdjoint G - where - left := _ - adj := adjunctionOfStructuredArrowInitials G +lemma isRightAdjointOfStructuredArrowInitials : G.IsRightAdjoint where + exists_leftAdjoint := ⟨_, ⟨adjunctionOfStructuredArrowInitials G⟩⟩ #align category_theory.is_right_adjoint_of_structured_arrow_initials CategoryTheory.isRightAdjointOfStructuredArrowInitials end OfInitials @@ -128,11 +126,10 @@ def adjunctionOfCostructuredArrowTerminals : G ⊣ rightAdjointOfCostructuredArr #align category_theory.adjunction_of_costructured_arrow_terminals CategoryTheory.adjunctionOfCostructuredArrowTerminals /-- If each costructured arrow category on `G` has a terminal object, `G` is a left adjoint. -/ -def isLeftAdjointOfCostructuredArrowTerminals : IsLeftAdjoint G - where - right := rightAdjointOfCostructuredArrowTerminals G - adj := Adjunction.adjunctionOfEquivRight _ _ -#align category_theory.is_left_adjoint_of_costructured_arrow_terminals CategoryTheory.isLeftAdjointOfCostructuredArrowTerminals +lemma isLeftAdjoint_of_costructuredArrowTerminals : G.IsLeftAdjoint where + exists_rightAdjoint := + ⟨rightAdjointOfCostructuredArrowTerminals G, ⟨Adjunction.adjunctionOfEquivRight _ _⟩⟩ +#align category_theory.is_left_adjoint_of_costructured_arrow_terminals CategoryTheory.isLeftAdjoint_of_costructuredArrowTerminals end OfTerminals @@ -169,16 +166,16 @@ def mkTerminalOfRightAdjoint (h : F ⊣ G) (A : D) : end -theorem nonempty_isRightAdjoint_iff_hasInitial_structuredArrow {G : D ⥤ C} : - Nonempty (IsRightAdjoint G) ↔ ∀ A, HasInitial (StructuredArrow A G) := - ⟨fun ⟨h⟩ A => (mkInitialOfLeftAdjoint _ h.adj A).hasInitial, fun _ => - ⟨isRightAdjointOfStructuredArrowInitials _⟩⟩ -#align category_theory.nonempty_is_right_adjoint_iff_has_initial_structured_arrow CategoryTheory.nonempty_isRightAdjoint_iff_hasInitial_structuredArrow - -theorem nonempty_isLeftAdjoint_iff_hasTerminal_costructuredArrow {F : C ⥤ D} : - Nonempty (IsLeftAdjoint F) ↔ ∀ A, HasTerminal (CostructuredArrow F A) := - ⟨fun ⟨h⟩ A => (mkTerminalOfRightAdjoint _ h.adj A).hasTerminal, fun _ => - ⟨isLeftAdjointOfCostructuredArrowTerminals _⟩⟩ -#align category_theory.nonempty_is_left_adjoint_iff_has_terminal_costructured_arrow CategoryTheory.nonempty_isLeftAdjoint_iff_hasTerminal_costructuredArrow +theorem isRightAdjoint_iff_hasInitial_structuredArrow {G : D ⥤ C} : + G.IsRightAdjoint ↔ ∀ A, HasInitial (StructuredArrow A G) := + ⟨fun _ A => (mkInitialOfLeftAdjoint _ (Adjunction.ofIsRightAdjoint G) A).hasInitial, + fun _ => isRightAdjointOfStructuredArrowInitials _⟩ +#align category_theory.nonempty_is_right_adjoint_iff_has_initial_structured_arrow CategoryTheory.isRightAdjoint_iff_hasInitial_structuredArrow + +theorem isLeftAdjoint_iff_hasTerminal_costructuredArrow {F : C ⥤ D} : + F.IsLeftAdjoint ↔ ∀ A, HasTerminal (CostructuredArrow F A) := + ⟨fun _ A => (mkTerminalOfRightAdjoint _ (Adjunction.ofIsLeftAdjoint F) A).hasTerminal, + fun _ => isLeftAdjoint_of_costructuredArrowTerminals _⟩ +#align category_theory.nonempty_is_left_adjoint_iff_has_terminal_costructured_arrow CategoryTheory.isLeftAdjoint_iff_hasTerminal_costructuredArrow end CategoryTheory diff --git a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean index 84fade269826f..0d4b845781676 100644 --- a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean +++ b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean @@ -74,8 +74,8 @@ def evaluationAdjunctionRight (c : C) : evaluationLeftAdjoint D c ⊣ (evaluatio homEquiv_naturality_right := by intros; dsimp; simp } #align category_theory.evaluation_adjunction_right CategoryTheory.evaluationAdjunctionRight -instance evaluationIsRightAdjoint (c : C) : IsRightAdjoint ((evaluation _ D).obj c) := - ⟨_, evaluationAdjunctionRight _ _⟩ +instance evaluationIsRightAdjoint (c : C) : ((evaluation _ D).obj c).IsRightAdjoint := + ⟨_, ⟨evaluationAdjunctionRight _ _⟩⟩ #align category_theory.evaluation_is_right_adjoint CategoryTheory.evaluationIsRightAdjoint theorem NatTrans.mono_iff_mono_app {F G : C ⥤ D} (η : F ⟶ G) : Mono η ↔ ∀ c, Mono (η.app c) := by @@ -133,8 +133,8 @@ def evaluationAdjunctionLeft (c : C) : (evaluation _ _).obj c ⊣ evaluationRigh Fan.mk_pt, Fan.mk_π_app, Discrete.natTrans_app, Category.comp_id] } } #align category_theory.evaluation_adjunction_left CategoryTheory.evaluationAdjunctionLeft -instance evaluationIsLeftAdjoint (c : C) : IsLeftAdjoint ((evaluation _ D).obj c) := - ⟨_, evaluationAdjunctionLeft _ _⟩ +instance evaluationIsLeftAdjoint (c : C) : ((evaluation _ D).obj c).IsLeftAdjoint := + ⟨_, ⟨evaluationAdjunctionLeft _ _⟩⟩ #align category_theory.evaluation_is_left_adjoint CategoryTheory.evaluationIsLeftAdjoint theorem NatTrans.epi_iff_epi_app {F G : C ⥤ D} (η : F ⟶ G) : Epi η ↔ ∀ c, Epi (η.app c) := by diff --git a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean index 5596f26d468d6..c95fa51adb725 100644 --- a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean @@ -23,7 +23,7 @@ The statements from Riehl 4.5.13 for adjoints which are either full, or faithful open CategoryTheory -namespace CategoryTheory +namespace CategoryTheory.Adjunction universe v₁ v₂ u₁ u₂ @@ -58,7 +58,7 @@ instance unit_isIso_of_L_fully_faithful [L.Full] [L.Faithful] : IsIso (Adjunctio rw [← h.unit_naturality] simp⟩⟩⟩ set_option linter.uppercaseLean3 false in -#align category_theory.unit_is_iso_of_L_fully_faithful CategoryTheory.unit_isIso_of_L_fully_faithful +#align category_theory.unit_is_iso_of_L_fully_faithful CategoryTheory.Adjunction.unit_isIso_of_L_fully_faithful /-- If the right adjoint is fully faithful, then the counit is an isomorphism. @@ -80,7 +80,7 @@ instance counit_isIso_of_R_fully_faithful [R.Full] [R.Faithful] : IsIso (Adjunct rw [← h.counit_naturality] simp⟩⟩⟩ set_option linter.uppercaseLean3 false in -#align category_theory.counit_is_iso_of_R_fully_faithful CategoryTheory.counit_isIso_of_R_fully_faithful +#align category_theory.counit_is_iso_of_R_fully_faithful CategoryTheory.Adjunction.counit_isIso_of_R_fully_faithful /-- If the unit of an adjunction is an isomorphism, then its inverse on the image of L is given by L whiskered with the counit. -/ @@ -88,14 +88,14 @@ by L whiskered with the counit. -/ theorem inv_map_unit {X : C} [IsIso (h.unit.app X)] : inv (L.map (h.unit.app X)) = h.counit.app (L.obj X) := IsIso.inv_eq_of_hom_inv_id (h.left_triangle_components X) -#align category_theory.inv_map_unit CategoryTheory.inv_map_unit +#align category_theory.inv_map_unit CategoryTheory.Adjunction.inv_map_unit /-- If the unit is an isomorphism, bundle one has an isomorphism `L ⋙ R ⋙ L ≅ L`. -/ @[simps!] noncomputable def whiskerLeftLCounitIsoOfIsIsoUnit [IsIso h.unit] : L ⋙ R ⋙ L ≅ L := (L.associator R L).symm ≪≫ isoWhiskerRight (asIso h.unit).symm L ≪≫ Functor.leftUnitor _ set_option linter.uppercaseLean3 false in -#align category_theory.whisker_left_L_counit_iso_of_is_iso_unit CategoryTheory.whiskerLeftLCounitIsoOfIsIsoUnit +#align category_theory.whisker_left_L_counit_iso_of_is_iso_unit CategoryTheory.Adjunction.whiskerLeftLCounitIsoOfIsIsoUnit /-- If the counit of an adjunction is an isomorphism, then its inverse on the image of R is given by R whiskered with the unit. -/ @@ -103,20 +103,20 @@ by R whiskered with the unit. -/ theorem inv_counit_map {X : D} [IsIso (h.counit.app X)] : inv (R.map (h.counit.app X)) = h.unit.app (R.obj X) := IsIso.inv_eq_of_inv_hom_id (h.right_triangle_components X) -#align category_theory.inv_counit_map CategoryTheory.inv_counit_map +#align category_theory.inv_counit_map CategoryTheory.Adjunction.inv_counit_map /-- If the counit of an is an isomorphism, one has an isomorphism `(R ⋙ L ⋙ R) ≅ R`. -/ @[simps!] noncomputable def whiskerLeftRUnitIsoOfIsIsoCounit [IsIso h.counit] : R ⋙ L ⋙ R ≅ R := (R.associator L R).symm ≪≫ isoWhiskerRight (asIso h.counit) R ≪≫ Functor.leftUnitor _ set_option linter.uppercaseLean3 false in -#align category_theory.whisker_left_R_unit_iso_of_is_iso_counit CategoryTheory.whiskerLeftRUnitIsoOfIsIsoCounit +#align category_theory.whisker_left_R_unit_iso_of_is_iso_counit CategoryTheory.Adjunction.whiskerLeftRUnitIsoOfIsIsoCounit /-- If the unit is an isomorphism, then the left adjoint is full-/ lemma L_full_of_unit_isIso [IsIso h.unit] : L.Full where map_surjective {X Y} f := ⟨h.homEquiv _ (L.obj Y) f ≫ inv (h.unit.app Y), by simp⟩ set_option linter.uppercaseLean3 false in -#align category_theory.L_full_of_unit_is_iso CategoryTheory.L_full_of_unit_isIso +#align category_theory.L_full_of_unit_is_iso CategoryTheory.Adjunction.L_full_of_unit_isIso /-- If the unit is an isomorphism, then the left adjoint is faithful-/ theorem L_faithful_of_unit_isIso [IsIso h.unit] : L.Faithful := @@ -124,13 +124,13 @@ theorem L_faithful_of_unit_isIso [IsIso h.unit] : L.Faithful := rw [← (h.homEquiv X (L.obj Y)).apply_eq_iff_eq] at H simpa using H =≫ inv (h.unit.app Y)⟩ set_option linter.uppercaseLean3 false in -#align category_theory.L_faithful_of_unit_is_iso CategoryTheory.L_faithful_of_unit_isIso +#align category_theory.L_faithful_of_unit_is_iso CategoryTheory.Adjunction.L_faithful_of_unit_isIso /-- If the counit is an isomorphism, then the right adjoint is full-/ lemma R_full_of_counit_isIso [IsIso h.counit] : R.Full where map_surjective {X Y} f := ⟨inv (h.counit.app X) ≫ (h.homEquiv (R.obj X) Y).symm f, by simp⟩ set_option linter.uppercaseLean3 false in -#align category_theory.R_full_of_counit_is_iso CategoryTheory.R_full_of_counit_isIso +#align category_theory.R_full_of_counit_is_iso CategoryTheory.Adjunction.R_full_of_counit_isIso /-- If the counit is an isomorphism, then the right adjoint is faithful-/ theorem R_faithful_of_counit_isIso [IsIso h.counit] : R.Faithful := @@ -138,7 +138,7 @@ theorem R_faithful_of_counit_isIso [IsIso h.counit] : R.Faithful := rw [← (h.homEquiv (R.obj X) Y).symm.apply_eq_iff_eq] at H simpa using inv (h.counit.app X) ≫= H⟩ set_option linter.uppercaseLean3 false in -#align category_theory.R_faithful_of_counit_is_iso CategoryTheory.R_faithful_of_counit_isIso +#align category_theory.R_faithful_of_counit_is_iso CategoryTheory.Adjunction.R_faithful_of_counit_isIso instance whiskerLeft_counit_iso_of_L_fully_faithful [L.Full] [L.Faithful] : IsIso (whiskerLeft L h.counit) := by @@ -147,7 +147,7 @@ instance whiskerLeft_counit_iso_of_L_fully_faithful [L.Full] [L.Faithful] : rw [this] infer_instance set_option linter.uppercaseLean3 false in -#align category_theory.whisker_left_counit_iso_of_L_fully_faithful CategoryTheory.whiskerLeft_counit_iso_of_L_fully_faithful +#align category_theory.whisker_left_counit_iso_of_L_fully_faithful CategoryTheory.Adjunction.whiskerLeft_counit_iso_of_L_fully_faithful instance whiskerRight_counit_iso_of_L_fully_faithful [L.Full] [L.Faithful] : IsIso (whiskerRight h.counit R) := by @@ -156,7 +156,7 @@ instance whiskerRight_counit_iso_of_L_fully_faithful [L.Full] [L.Faithful] : rw [this] infer_instance set_option linter.uppercaseLean3 false in -#align category_theory.whisker_right_counit_iso_of_L_fully_faithful CategoryTheory.whiskerRight_counit_iso_of_L_fully_faithful +#align category_theory.whisker_right_counit_iso_of_L_fully_faithful CategoryTheory.Adjunction.whiskerRight_counit_iso_of_L_fully_faithful instance whiskerLeft_unit_iso_of_R_fully_faithful [R.Full] [R.Faithful] : IsIso (whiskerLeft R h.unit) := by @@ -165,7 +165,7 @@ instance whiskerLeft_unit_iso_of_R_fully_faithful [R.Full] [R.Faithful] : rw [this] infer_instance set_option linter.uppercaseLean3 false in -#align category_theory.whisker_left_unit_iso_of_R_fully_faithful CategoryTheory.whiskerLeft_unit_iso_of_R_fully_faithful +#align category_theory.whisker_left_unit_iso_of_R_fully_faithful CategoryTheory.Adjunction.whiskerLeft_unit_iso_of_R_fully_faithful instance whiskerRight_unit_iso_of_R_fully_faithful [R.Full] [R.Faithful] : IsIso (whiskerRight h.unit L) := by @@ -174,7 +174,7 @@ instance whiskerRight_unit_iso_of_R_fully_faithful [R.Full] [R.Faithful] : rw [this] infer_instance set_option linter.uppercaseLean3 false in -#align category_theory.whisker_right_unit_iso_of_R_fully_faithful CategoryTheory.whiskerRight_unit_iso_of_R_fully_faithful +#align category_theory.whisker_right_unit_iso_of_R_fully_faithful CategoryTheory.Adjunction.whiskerRight_unit_iso_of_R_fully_faithful instance [L.Faithful] [L.Full] {Y : C} : IsIso (h.counit.app (L.obj Y)) := isIso_of_hom_comp_eq_id _ (h.left_triangle_components Y) @@ -208,4 +208,26 @@ lemma isIso_unit_app_of_iso [R.Faithful] [R.Full] {X : D} {Y : C} (e : Y ≅ R.o IsIso (h.unit.app Y) := (isIso_unit_app_iff_mem_essImage h).mpr ⟨X, ⟨e.symm⟩⟩ -end CategoryTheory +instance [R.IsEquivalence] : IsIso h.unit := by + have := fun Y => isIso_unit_app_of_iso h (R.objObjPreimageIso Y).symm + apply NatIso.isIso_of_isIso_app + +instance [L.IsEquivalence] : IsIso h.counit := by + have := fun X => isIso_counit_app_of_iso h (L.objObjPreimageIso X).symm + apply NatIso.isIso_of_isIso_app + +lemma isEquivalence_left_of_isEquivalence_right [R.IsEquivalence] : L.IsEquivalence := + h.toEquivalence.isEquivalence_functor + +lemma isEquivalence_right_of_isEquivalence_left [L.IsEquivalence] : R.IsEquivalence := + h.toEquivalence.isEquivalence_inverse + +instance [L.IsEquivalence] : IsIso h.unit := by + have := h.isEquivalence_right_of_isEquivalence_left + infer_instance + +instance [R.IsEquivalence] : IsIso h.counit := by + have := h.isEquivalence_left_of_isEquivalence_right + infer_instance + +end CategoryTheory.Adjunction diff --git a/Mathlib/CategoryTheory/Adjunction/Lifting.lean b/Mathlib/CategoryTheory/Adjunction/Lifting.lean index 8987a4fc612e0..74659ef6c0637 100644 --- a/Mathlib/CategoryTheory/Adjunction/Lifting.lean +++ b/Mathlib/CategoryTheory/Adjunction/Lifting.lean @@ -175,30 +175,30 @@ pairs, then a functor `R : A ⥤ B` has a left adjoint if the composite `R ⋙ U Note the converse is true (with weaker assumptions), by `Adjunction.comp`. See https://ncatlab.org/nlab/show/adjoint+triangle+theorem -/ -noncomputable def adjointTriangleLift {U : B ⥤ C} {F : C ⥤ B} (R : A ⥤ B) (adj₁ : F ⊣ U) +lemma adjointTriangleLift {U : B ⥤ C} {F : C ⥤ B} (R : A ⥤ B) (adj₁ : F ⊣ U) [∀ X : B, RegularEpi (adj₁.counit.app X)] [HasReflexiveCoequalizers A] - [IsRightAdjoint (R ⋙ U)] : IsRightAdjoint R where - left := LiftAdjoint.constructLeftAdjoint R _ adj₁ (Adjunction.ofRightAdjoint _) - adj := Adjunction.adjunctionOfEquivLeft _ _ -#align category_theory.adjoint_triangle_lift CategoryTheory.adjointTriangleLift + [(R ⋙ U).IsRightAdjoint ] : R.IsRightAdjoint where + exists_leftAdjoint := + ⟨LiftAdjoint.constructLeftAdjoint R _ adj₁ (Adjunction.ofIsRightAdjoint _), + ⟨Adjunction.adjunctionOfEquivLeft _ _⟩⟩ /-- If `R ⋙ U` has a left adjoint, the domain of `R` has reflexive coequalizers and `U` is a monadic functor, then `R` has a left adjoint. This is a special case of `adjointTriangleLift` which is often more useful in practice. -/ -noncomputable def monadicAdjointTriangleLift (U : B ⥤ C) [MonadicRightAdjoint U] {R : A ⥤ B} - [HasReflexiveCoequalizers A] [IsRightAdjoint (R ⋙ U)] : IsRightAdjoint R := by - let R' : A ⥤ _ := R ⋙ Monad.comparison (Adjunction.ofRightAdjoint U) - rsuffices : IsRightAdjoint R' - · let this : IsRightAdjoint (R' ⋙ (Monad.comparison (Adjunction.ofRightAdjoint U)).inv) := by +lemma monadicAdjointTriangleLift (U : B ⥤ C) [MonadicRightAdjoint U] {R : A ⥤ B} + [HasReflexiveCoequalizers A] [(R ⋙ U).IsRightAdjoint] : R.IsRightAdjoint := by + let R' : A ⥤ _ := R ⋙ Monad.comparison (monadicAdjunction U) + rsuffices : R'.IsRightAdjoint + · let this : (R' ⋙ (Monad.comparison (monadicAdjunction U)).inv).IsRightAdjoint := by infer_instance - · let this : R' ⋙ (Monad.comparison (Adjunction.ofRightAdjoint U)).inv ≅ R := - (isoWhiskerLeft R (Monad.comparison _).asEquivalence.unitIso.symm : _) ≪≫ R.rightUnitor - exact Adjunction.rightAdjointOfNatIso this - let this : IsRightAdjoint (R' ⋙ Monad.forget (Adjunction.ofRightAdjoint U).toMonad) := - Adjunction.rightAdjointOfNatIso - (isoWhiskerLeft R (Monad.comparisonForget (Adjunction.ofRightAdjoint U)).symm : _) - let this : ∀ X, RegularEpi ((Monad.adj (Adjunction.ofRightAdjoint U).toMonad).counit.app X) := by + · refine' ((Adjunction.ofIsRightAdjoint + (R' ⋙ (Monad.comparison (monadicAdjunction U)).inv)).ofNatIsoRight _).isRightAdjoint + exact isoWhiskerLeft R (Monad.comparison _).asEquivalence.unitIso.symm ≪≫ R.rightUnitor + let this : (R' ⋙ Monad.forget (monadicAdjunction U).toMonad).IsRightAdjoint := by + refine' ((Adjunction.ofIsRightAdjoint (R ⋙ U)).ofNatIsoRight _).isRightAdjoint + exact isoWhiskerLeft R (Monad.comparisonForget (monadicAdjunction U)).symm + let this : ∀ X, RegularEpi ((Monad.adj (monadicAdjunction U).toMonad).counit.app X) := by intro X simp only [Monad.adj_counit] exact ⟨_, _, _, _, Monad.beckAlgebraCoequalizer X⟩ @@ -222,12 +222,12 @@ Then `Q` has a left adjoint if `R` has a left adjoint. See https://ncatlab.org/nlab/show/adjoint+lifting+theorem -/ -noncomputable def adjointSquareLift (Q : A ⥤ B) (V : B ⥤ D) (U : A ⥤ C) (R : C ⥤ D) - (comm : U ⋙ R ≅ Q ⋙ V) [IsRightAdjoint U] [IsRightAdjoint V] [IsRightAdjoint R] - [∀ X, RegularEpi ((Adjunction.ofRightAdjoint V).counit.app X)] [HasReflexiveCoequalizers A] : - IsRightAdjoint Q := - letI := Adjunction.rightAdjointOfNatIso comm - adjointTriangleLift Q (Adjunction.ofRightAdjoint V) +lemma adjointSquareLift (Q : A ⥤ B) (V : B ⥤ D) (U : A ⥤ C) (R : C ⥤ D) + (comm : U ⋙ R ≅ Q ⋙ V) [U.IsRightAdjoint] [V.IsRightAdjoint] [R.IsRightAdjoint] + [∀ X, RegularEpi ((Adjunction.ofIsRightAdjoint V).counit.app X)] [HasReflexiveCoequalizers A] : + Q.IsRightAdjoint := + have := ((Adjunction.ofIsRightAdjoint (U ⋙ R)).ofNatIsoRight comm).isRightAdjoint + adjointTriangleLift Q (Adjunction.ofIsRightAdjoint V) #align category_theory.adjoint_square_lift CategoryTheory.adjointSquareLift /-- Suppose we have a commutative square of functors @@ -243,10 +243,10 @@ Then `Q` has a left adjoint if `R` has a left adjoint. See https://ncatlab.org/nlab/show/adjoint+lifting+theorem -/ -noncomputable def monadicAdjointSquareLift (Q : A ⥤ B) (V : B ⥤ D) (U : A ⥤ C) (R : C ⥤ D) - (comm : U ⋙ R ≅ Q ⋙ V) [IsRightAdjoint U] [MonadicRightAdjoint V] [IsRightAdjoint R] - [HasReflexiveCoequalizers A] : IsRightAdjoint Q := - letI := Adjunction.rightAdjointOfNatIso comm +lemma monadicAdjointSquareLift (Q : A ⥤ B) (V : B ⥤ D) (U : A ⥤ C) (R : C ⥤ D) + (comm : U ⋙ R ≅ Q ⋙ V) [U.IsRightAdjoint] [MonadicRightAdjoint V] [R.IsRightAdjoint] + [HasReflexiveCoequalizers A] : Q.IsRightAdjoint := + have := ((Adjunction.ofIsRightAdjoint (U ⋙ R)).ofNatIsoRight comm).isRightAdjoint monadicAdjointTriangleLift V #align category_theory.monadic_adjoint_square_lift CategoryTheory.monadicAdjointSquareLift diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index 57205d3107479..46ab14d5fbfc6 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -79,12 +79,11 @@ def functorialityCounit : #align category_theory.adjunction.functoriality_counit CategoryTheory.Adjunction.functorialityCounit /-- The functor `Cocones.functoriality K F : Cocone K ⥤ Cocone (K ⋙ F)` is a left adjoint. -/ -def functorialityIsLeftAdjoint : IsLeftAdjoint (Cocones.functoriality K F) where - right := functorialityRightAdjoint adj K - adj := mkOfUnitCounit +def functorialityAdjunction : Cocones.functoriality K F ⊣ functorialityRightAdjoint adj K := + mkOfUnitCounit { unit := functorialityUnit adj K - counit := functorialityCounit adj K } -#align category_theory.adjunction.functoriality_is_left_adjoint CategoryTheory.Adjunction.functorialityIsLeftAdjoint + counit := functorialityCounit adj K} +#align category_theory.adjunction.functoriality_is_left_adjoint CategoryTheory.Adjunction.functorialityAdjunction /-- A left adjoint preserves colimits. @@ -96,17 +95,19 @@ def leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F where { preserves := fun hc => IsColimit.isoUniqueCoconeMorphism.inv fun _ => @Equiv.unique _ _ (IsColimit.isoUniqueCoconeMorphism.hom hc _) - ((adj.functorialityIsLeftAdjoint _).adj.homEquiv _ _) } } + ((adj.functorialityAdjunction _).homEquiv _ _) } } #align category_theory.adjunction.left_adjoint_preserves_colimits CategoryTheory.Adjunction.leftAdjointPreservesColimits -- see Note [lower instance priority] -instance (priority := 100) isEquivalencePreservesColimits (E : C ⥤ D) [E.IsEquivalence] : +noncomputable instance (priority := 100) isEquivalencePreservesColimits + (E : C ⥤ D) [E.IsEquivalence] : PreservesColimitsOfSize.{v, u} E := leftAdjointPreservesColimits E.adjunction #align category_theory.adjunction.is_equivalence_preserves_colimits CategoryTheory.Adjunction.isEquivalencePreservesColimits -- see Note [lower instance priority] -instance (priority := 100) isEquivalenceReflectsColimits (E : D ⥤ C) [E.IsEquivalence] : +noncomputable instance (priority := 100) isEquivalenceReflectsColimits + (E : D ⥤ C) [E.IsEquivalence] : ReflectsColimitsOfSize.{v, u} E where reflectsColimitsOfShape := { reflectsColimit := @@ -126,7 +127,7 @@ noncomputable instance (priority := 100) isEquivalenceCreatesColimits (H : D ⥤ #align category_theory.adjunction.is_equivalence_creates_colimits CategoryTheory.Adjunction.isEquivalenceCreatesColimits -- verify the preserve_colimits instance works as expected: -example (E : C ⥤ D) [E.IsEquivalence] (c : Cocone K) (h : IsColimit c) : +noncomputable example (E : C ⥤ D) [E.IsEquivalence] (c : Cocone K) (h : IsColimit c) : IsColimit (E.mapCocone c) := PreservesColimit.preserves h @@ -194,12 +195,11 @@ def functorialityCounit' : #align category_theory.adjunction.functoriality_counit' CategoryTheory.Adjunction.functorialityCounit' /-- The functor `Cones.functoriality K G : Cone K ⥤ Cone (K ⋙ G)` is a right adjoint. -/ -def functorialityIsRightAdjoint : IsRightAdjoint (Cones.functoriality K G) where - left := functorialityLeftAdjoint adj K - adj := mkOfUnitCounit +def functorialityAdjunction' : functorialityLeftAdjoint adj K ⊣ Cones.functoriality K G := + mkOfUnitCounit { unit := functorialityUnit' adj K counit := functorialityCounit' adj K } -#align category_theory.adjunction.functoriality_is_right_adjoint CategoryTheory.Adjunction.functorialityIsRightAdjoint +#align category_theory.adjunction.functoriality_is_right_adjoint CategoryTheory.Adjunction.functorialityAdjunction' /-- A right adjoint preserves limits. @@ -211,17 +211,19 @@ def rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G where { preserves := fun hc => IsLimit.isoUniqueConeMorphism.inv fun _ => @Equiv.unique _ _ (IsLimit.isoUniqueConeMorphism.hom hc _) - ((adj.functorialityIsRightAdjoint _).adj.homEquiv _ _).symm } } + ((adj.functorialityAdjunction' _).homEquiv _ _).symm } } #align category_theory.adjunction.right_adjoint_preserves_limits CategoryTheory.Adjunction.rightAdjointPreservesLimits -- see Note [lower instance priority] -instance (priority := 100) isEquivalencePreservesLimits (E : D ⥤ C) [E.IsEquivalence] : +noncomputable instance (priority := 100) isEquivalencePreservesLimits + (E : D ⥤ C) [E.IsEquivalence] : PreservesLimitsOfSize.{v, u} E := - rightAdjointPreservesLimits E.inv.adjunction + rightAdjointPreservesLimits E.asEquivalence.symm.toAdjunction #align category_theory.adjunction.is_equivalence_preserves_limits CategoryTheory.Adjunction.isEquivalencePreservesLimits -- see Note [lower instance priority] -instance (priority := 100) isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : +noncomputable instance (priority := 100) isEquivalenceReflectsLimits + (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E where reflectsLimitsOfShape := { reflectsLimit := @@ -240,7 +242,8 @@ noncomputable instance (priority := 100) isEquivalenceCreatesLimits (H : D ⥤ C #align category_theory.adjunction.is_equivalence_creates_limits CategoryTheory.Adjunction.isEquivalenceCreatesLimits -- verify the preserve_limits instance works as expected: -example (E : D ⥤ C) [E.IsEquivalence] (c : Cone K) (h : IsLimit c) : IsLimit (E.mapCone c) := +noncomputable example (E : D ⥤ C) [E.IsEquivalence] (c : Cone K) (h : IsLimit c) : + IsLimit (E.mapCone c) := PreservesLimit.preserves h theorem hasLimit_comp_equivalence (E : D ⥤ C) [E.IsEquivalence] [HasLimit K] : HasLimit (K ⋙ E) := @@ -348,15 +351,15 @@ namespace Functor variable {J C D : Type*} [Category J] [Category C] [Category D] (F : C ⥤ D) -instance [IsLeftAdjoint F] : PreservesColimitsOfShape J F := - (Adjunction.ofLeftAdjoint F).leftAdjointPreservesColimits.preservesColimitsOfShape +noncomputable instance [IsLeftAdjoint F] : PreservesColimitsOfShape J F := + (Adjunction.ofIsLeftAdjoint F).leftAdjointPreservesColimits.preservesColimitsOfShape -instance [IsLeftAdjoint F] : PreservesColimits F where +noncomputable instance [IsLeftAdjoint F] : PreservesColimits F where -instance [IsRightAdjoint F] : PreservesLimitsOfShape J F := - (Adjunction.ofRightAdjoint F).rightAdjointPreservesLimits.preservesLimitsOfShape +noncomputable instance [IsRightAdjoint F] : PreservesLimitsOfShape J F := + (Adjunction.ofIsRightAdjoint F).rightAdjointPreservesLimits.preservesLimitsOfShape -instance [IsRightAdjoint F] : PreservesLimits F where +noncomputable instance [IsRightAdjoint F] : PreservesLimits F where end Functor diff --git a/Mathlib/CategoryTheory/Adjunction/Over.lean b/Mathlib/CategoryTheory/Adjunction/Over.lean index 863a0afb20dd6..63b00b7beca65 100644 --- a/Mathlib/CategoryTheory/Adjunction/Over.lean +++ b/Mathlib/CategoryTheory/Adjunction/Over.lean @@ -49,7 +49,7 @@ def forgetAdjStar [HasBinaryProducts C] : Over.forget X ⊣ star X := /-- Note that the binary products assumption is necessary: the existence of a right adjoint to `Over.forget X` is equivalent to the existence of each binary product `X ⨯ -`. -/ -instance [HasBinaryProducts C] : IsLeftAdjoint (Over.forget X) := - ⟨_, forgetAdjStar X⟩ +instance [HasBinaryProducts C] : (Over.forget X).IsLeftAdjoint := + ⟨_, ⟨forgetAdjStar X⟩⟩ end CategoryTheory diff --git a/Mathlib/CategoryTheory/Adjunction/Reflective.lean b/Mathlib/CategoryTheory/Adjunction/Reflective.lean index 6bcf9f71d2b0d..ee29007babd8c 100644 --- a/Mathlib/CategoryTheory/Adjunction/Reflective.lean +++ b/Mathlib/CategoryTheory/Adjunction/Reflective.lean @@ -33,19 +33,33 @@ variable [Category.{v₁} C] [Category.{v₂} D] [Category.{v₃} E] /-- A functor is *reflective*, or *a reflective inclusion*, if it is fully faithful and right adjoint. -/ -class Reflective (R : D ⥤ C) extends IsRightAdjoint R, R.Full, R.Faithful +class Reflective (R : D ⥤ C) extends R.Full, R.Faithful where + /-- a choice of a left adjoint to `R` -/ + L : C ⥤ D + /-- `R` is a right adjoint -/ + adj : L ⊣ R #align category_theory.reflective CategoryTheory.Reflective -variable {i : D ⥤ C} +variable (i : D ⥤ C) + +/-- The reflector `C ⥤ D` when `R : D ⥤ C` is reflective. -/ +def reflector [Reflective i] : C ⥤ D := Reflective.L (R := i) + +/-- The adjunction `reflector i ⊣ i` when `i` is reflective. -/ +def reflectorAdjunction [Reflective i] : reflector i ⊣ i := Reflective.adj + +instance [Reflective i] : i.IsRightAdjoint := ⟨_, ⟨reflectorAdjunction i⟩⟩ + +instance [Reflective i] : (reflector i).IsLeftAdjoint := ⟨_, ⟨reflectorAdjunction i⟩⟩ -- TODO: This holds more generally for idempotent adjunctions, not just reflective adjunctions. /-- For a reflective functor `i` (with left adjoint `L`), with unit `η`, we have `η_iL = iL η`. -/ theorem unit_obj_eq_map_unit [Reflective i] (X : C) : - (ofRightAdjoint i).unit.app (i.obj ((leftAdjoint i).obj X)) = - i.map ((leftAdjoint i).map ((ofRightAdjoint i).unit.app X)) := by - rw [← cancel_mono (i.map ((ofRightAdjoint i).counit.app ((leftAdjoint i).obj X))), ← - i.map_comp] + (reflectorAdjunction i).unit.app (i.obj ((reflector i).obj X)) = + i.map ((reflector i).map ((reflectorAdjunction i).unit.app X)) := by + rw [← cancel_mono (i.map ((reflectorAdjunction i).counit.app ((reflector i).obj X))), + ← i.map_comp] simp #align category_theory.unit_obj_eq_map_unit CategoryTheory.unit_obj_eq_map_unit @@ -55,9 +69,11 @@ When restricted to objects in `D` given by `i : D ⥤ C`, the unit is an isomorp More generally this applies to objects essentially in the reflective subcategory, see `Functor.essImage.unit_isIso`. -/ -example [Reflective i] {B : D} : IsIso ((ofRightAdjoint i).unit.app (i.obj B)) := +example [Reflective i] {B : D} : IsIso ((reflectorAdjunction i).unit.app (i.obj B)) := inferInstance +variable {i} + /-- If `A` is essentially in the image of a reflective functor `i`, then `η_A` is an isomorphism. This gives that the "witness" for `A` being in the essential image can instead be given as the reflection of `A`, with the isomorphism as `η_A`. @@ -65,45 +81,47 @@ reflection of `A`, with the isomorphism as `η_A`. (For any `B` in the reflective subcategory, we automatically have that `ε_B` is an iso.) -/ theorem Functor.essImage.unit_isIso [Reflective i] {A : C} (h : A ∈ i.essImage) : - IsIso ((ofRightAdjoint i).unit.app A) := by + IsIso ((reflectorAdjunction i).unit.app A) := by rwa [isIso_unit_app_iff_mem_essImage] #align category_theory.functor.ess_image.unit_is_iso CategoryTheory.Functor.essImage.unit_isIso /-- If `η_A` is an isomorphism, then `A` is in the essential image of `i`. -/ -theorem mem_essImage_of_unit_isIso [IsRightAdjoint i] (A : C) - [IsIso ((ofRightAdjoint i).unit.app A)] : A ∈ i.essImage := - ⟨(leftAdjoint i).obj A, ⟨(asIso ((ofRightAdjoint i).unit.app A)).symm⟩⟩ +theorem mem_essImage_of_unit_isIso {L : C ⥤ D} (adj : L ⊣ i) (A : C) + [IsIso (adj.unit.app A)] : A ∈ i.essImage := + ⟨L.obj A, ⟨(asIso (adj.unit.app A)).symm⟩⟩ #align category_theory.mem_ess_image_of_unit_is_iso CategoryTheory.mem_essImage_of_unit_isIso /-- If `η_A` is a split monomorphism, then `A` is in the reflective subcategory. -/ theorem mem_essImage_of_unit_isSplitMono [Reflective i] {A : C} - [IsSplitMono ((ofRightAdjoint i).unit.app A)] : A ∈ i.essImage := by - let η : 𝟭 C ⟶ leftAdjoint i ⋙ i := (ofRightAdjoint i).unit - haveI : IsIso (η.app (i.obj ((leftAdjoint i).obj A))) := + [IsSplitMono ((reflectorAdjunction i).unit.app A)] : A ∈ i.essImage := by + let η : 𝟭 C ⟶ reflector i ⋙ i := (reflectorAdjunction i).unit + haveI : IsIso (η.app (i.obj ((reflector i).obj A))) := Functor.essImage.unit_isIso ((i.obj_mem_essImage _)) have : Epi (η.app A) := by refine @epi_of_epi _ _ _ _ _ (retraction (η.app A)) (η.app A) ?_ rw [show retraction _ ≫ η.app A = _ from η.naturality (retraction (η.app A))] - apply epi_comp (η.app (i.obj ((leftAdjoint i).obj A))) + apply epi_comp (η.app (i.obj ((reflector i).obj A))) haveI := isIso_of_epi_of_isSplitMono (η.app A) - exact mem_essImage_of_unit_isIso A + exact mem_essImage_of_unit_isIso (reflectorAdjunction i) A #align category_theory.mem_ess_image_of_unit_is_split_mono CategoryTheory.mem_essImage_of_unit_isSplitMono /-- Composition of reflective functors. -/ instance Reflective.comp (F : C ⥤ D) (G : D ⥤ E) [Reflective F] [Reflective G] : Reflective (F ⋙ G) where + L := reflector G ⋙ reflector F + adj := (reflectorAdjunction G).comp (reflectorAdjunction F) #align category_theory.reflective.comp CategoryTheory.Reflective.comp /-- (Implementation) Auxiliary definition for `unitCompPartialBijective`. -/ def unitCompPartialBijectiveAux [Reflective i] (A : C) (B : D) : - (A ⟶ i.obj B) ≃ (i.obj ((leftAdjoint i).obj A) ⟶ i.obj B) := - ((Adjunction.ofRightAdjoint i).homEquiv _ _).symm.trans (equivOfFullyFaithful i) + (A ⟶ i.obj B) ≃ (i.obj ((reflector i).obj A) ⟶ i.obj B) := + ((reflectorAdjunction i).homEquiv _ _).symm.trans (equivOfFullyFaithful i) #align category_theory.unit_comp_partial_bijective_aux CategoryTheory.unitCompPartialBijectiveAux /-- The description of the inverse of the bijection `unitCompPartialBijectiveAux`. -/ theorem unitCompPartialBijectiveAux_symm_apply [Reflective i] {A : C} {B : D} - (f : i.obj ((leftAdjoint i).obj A) ⟶ i.obj B) : - (unitCompPartialBijectiveAux _ _).symm f = (ofRightAdjoint i).unit.app A ≫ f := by + (f : i.obj ((reflector i).obj A) ⟶ i.obj B) : + (unitCompPartialBijectiveAux _ _).symm f = (reflectorAdjunction i).unit.app A ≫ f := by simp [unitCompPartialBijectiveAux] #align category_theory.unit_comp_partial_bijective_aux_symm_apply CategoryTheory.unitCompPartialBijectiveAux_symm_apply @@ -119,22 +137,22 @@ from the point of view of objects in `D`, `A` and `i.obj (L.obj A)` look the sam that `η.app A` is an isomorphism. -/ def unitCompPartialBijective [Reflective i] (A : C) {B : C} (hB : B ∈ i.essImage) : - (A ⟶ B) ≃ (i.obj ((leftAdjoint i).obj A) ⟶ B) := + (A ⟶ B) ≃ (i.obj ((reflector i).obj A) ⟶ B) := calc (A ⟶ B) ≃ (A ⟶ i.obj (Functor.essImage.witness hB)) := Iso.homCongr (Iso.refl _) hB.getIso.symm _ ≃ (i.obj _ ⟶ i.obj (Functor.essImage.witness hB)) := unitCompPartialBijectiveAux _ _ - _ ≃ (i.obj ((leftAdjoint i).obj A) ⟶ B) := + _ ≃ (i.obj ((reflector i).obj A) ⟶ B) := Iso.homCongr (Iso.refl _) (Functor.essImage.getIso hB) #align category_theory.unit_comp_partial_bijective CategoryTheory.unitCompPartialBijective @[simp] theorem unitCompPartialBijective_symm_apply [Reflective i] (A : C) {B : C} (hB : B ∈ i.essImage) - (f) : (unitCompPartialBijective A hB).symm f = (ofRightAdjoint i).unit.app A ≫ f := by + (f) : (unitCompPartialBijective A hB).symm f = (reflectorAdjunction i).unit.app A ≫ f := by simp [unitCompPartialBijective, unitCompPartialBijectiveAux_symm_apply] #align category_theory.unit_comp_partial_bijective_symm_apply CategoryTheory.unitCompPartialBijective_symm_apply theorem unitCompPartialBijective_symm_natural [Reflective i] (A : C) {B B' : C} (h : B ⟶ B') - (hB : B ∈ i.essImage) (hB' : B' ∈ i.essImage) (f : i.obj ((leftAdjoint i).obj A) ⟶ B) : + (hB : B ∈ i.essImage) (hB' : B' ∈ i.essImage) (f : i.obj ((reflector i).obj A) ⟶ B) : (unitCompPartialBijective A hB').symm (f ≫ h) = (unitCompPartialBijective A hB).symm f ≫ h := by simp #align category_theory.unit_comp_partial_bijective_symm_natural CategoryTheory.unitCompPartialBijective_symm_natural @@ -146,7 +164,7 @@ theorem unitCompPartialBijective_natural [Reflective i] (A : C) {B B' : C} (h : #align category_theory.unit_comp_partial_bijective_natural CategoryTheory.unitCompPartialBijective_natural instance [Reflective i] (X : Functor.EssImageSubcategory i) : - IsIso (NatTrans.app (ofRightAdjoint i).unit X.obj) := + IsIso (NatTrans.app (reflectorAdjunction i).unit X.obj) := Functor.essImage.unit_isIso X.property -- Porting note: the following auxiliary definition and the next two lemmas were @@ -154,8 +172,8 @@ Functor.essImage.unit_isIso X.property /-- The counit isomorphism of the equivalence `D ≌ i.EssImageSubcategory` given by `equivEssImageOfReflective` when the functor `i` is reflective. -/ def equivEssImageOfReflective_counitIso_app [Reflective i] (X : Functor.EssImageSubcategory i) : - ((Functor.essImageInclusion i ⋙ leftAdjoint i) ⋙ Functor.toEssImage i).obj X ≅ X := by - refine Iso.symm (@asIso _ _ X _ ((ofRightAdjoint i).unit.app X.obj) ?_) + ((Functor.essImageInclusion i ⋙ reflector i) ⋙ Functor.toEssImage i).obj X ≅ X := by + refine Iso.symm (@asIso _ _ X _ ((reflectorAdjunction i).unit.app X.obj) ?_) refine @isIso_of_reflects_iso _ _ _ _ _ _ _ i.essImageInclusion ?_ _ dsimp exact inferInstance @@ -163,7 +181,7 @@ def equivEssImageOfReflective_counitIso_app [Reflective i] (X : Functor.EssImage lemma equivEssImageOfReflective_map_counitIso_app_hom [Reflective i] (X : Functor.EssImageSubcategory i) : (Functor.essImageInclusion i).map (equivEssImageOfReflective_counitIso_app X).hom = - inv (NatTrans.app (ofRightAdjoint i).unit X.obj) := by + inv (NatTrans.app (reflectorAdjunction i).unit X.obj) := by simp only [Functor.comp_obj, Functor.essImageInclusion_obj, Functor.toEssImage_obj_obj, equivEssImageOfReflective_counitIso_app, asIso, Iso.symm_mk, Functor.essImageInclusion_map, Functor.id_obj] @@ -172,7 +190,7 @@ lemma equivEssImageOfReflective_map_counitIso_app_hom [Reflective i] lemma equivEssImageOfReflective_map_counitIso_app_inv [Reflective i] (X : Functor.EssImageSubcategory i) : (Functor.essImageInclusion i).map (equivEssImageOfReflective_counitIso_app X).inv = - (NatTrans.app (ofRightAdjoint i).unit X.obj) := rfl + (NatTrans.app (reflectorAdjunction i).unit X.obj) := rfl /-- If `i : D ⥤ C` is reflective, the inverse functor of `i ≌ F.essImage` can be explicitly defined by the reflector. -/ @@ -180,20 +198,20 @@ defined by the reflector. -/ def equivEssImageOfReflective [Reflective i] : D ≌ i.EssImageSubcategory where functor := i.toEssImage - inverse := i.essImageInclusion ⋙ (leftAdjoint i : _) + inverse := i.essImageInclusion ⋙ reflector i unitIso := - NatIso.ofComponents (fun X => (asIso <| (ofRightAdjoint i).counit.app X).symm) + NatIso.ofComponents (fun X => (asIso <| (reflectorAdjunction i).counit.app X).symm) (by intro X Y f dsimp rw [IsIso.comp_inv_eq, Category.assoc, IsIso.eq_inv_comp] - exact ((ofRightAdjoint i).counit.naturality f).symm) + exact ((reflectorAdjunction i).counit.naturality f).symm) counitIso := NatIso.ofComponents equivEssImageOfReflective_counitIso_app (by intro X Y f apply (Functor.essImageInclusion i).map_injective - have h := ((ofRightAdjoint i).unit.naturality f).symm + have h := ((reflectorAdjunction i).unit.naturality f).symm rw [Functor.id_map] at h erw [Functor.map_comp, Functor.map_comp, equivEssImageOfReflective_map_counitIso_app_hom, diff --git a/Mathlib/CategoryTheory/Bicategory/SingleObj.lean b/Mathlib/CategoryTheory/Bicategory/SingleObj.lean index 545f7a122422c..4778799c9f070 100644 --- a/Mathlib/CategoryTheory/Bicategory/SingleObj.lean +++ b/Mathlib/CategoryTheory/Bicategory/SingleObj.lean @@ -86,14 +86,19 @@ def endMonoidalStarFunctor : MonoidalFunctor (EndMonoidal (MonoidalSingleObj.sta when we promote a monoidal category to a single object bicategory, and the original monoidal category. -/ -noncomputable def endMonoidalStarFunctorIsEquivalence : - (endMonoidalStarFunctor C).IsEquivalence where +@[simps functor inverse_obj inverse_map unitIso counitIso] +noncomputable def endMonoidalStarFunctorEquivalence : + EndMonoidal (MonoidalSingleObj.star C) ≌ C where + functor := (endMonoidalStarFunctor C).toFunctor inverse := { obj := fun X => X map := fun f => f } - unitIso := NatIso.ofComponents fun X => asIso (𝟙 _) - counitIso := NatIso.ofComponents fun X => asIso (𝟙 _) -#align category_theory.monoidal_single_obj.End_monoidal_star_functor_is_equivalence CategoryTheory.MonoidalSingleObj.endMonoidalStarFunctorIsEquivalence + unitIso := Iso.refl _ + counitIso := Iso.refl _ + +instance endMonoidalStarFunctor_isEquivalence : (endMonoidalStarFunctor C).IsEquivalence := + (endMonoidalStarFunctorEquivalence C).isEquivalence_functor +#align category_theory.monoidal_single_obj.End_monoidal_star_functor_is_equivalence CategoryTheory.MonoidalSingleObj.endMonoidalStarFunctor_isEquivalence end MonoidalSingleObj diff --git a/Mathlib/CategoryTheory/CatCommSq.lean b/Mathlib/CategoryTheory/CatCommSq.lean index a08c2c5adb304..1f791f064aa43 100644 --- a/Mathlib/CategoryTheory/CatCommSq.lean +++ b/Mathlib/CategoryTheory/CatCommSq.lean @@ -95,10 +95,6 @@ def hInvEquiv : CatCommSq T.functor L R B.functor ≃ CatCommSq T.inverse R L B. end -instance hInv' [h : CatCommSq T L R B] [T.IsEquivalence] [B.IsEquivalence] : - CatCommSq T.inv R L B.inv := - hInv T.asEquivalence L R B.asEquivalence h - end CatCommSq end CategoryTheory diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index fd59ded043436..e25e074b8aff1 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -50,6 +50,13 @@ abbrev Exponentiable {C : Type u} [Category.{v} C] [HasFiniteProducts C] (X : C) Closed X #align category_theory.exponentiable CategoryTheory.Exponentiable +/-- Constructor for `Exponentiable X` which takes as an input an adjunction +`MonoidalCategory.tensorLeft X ⊣ exp` for some functor `exp : C ⥤ C`. -/ +abbrev Exponentiable.mk {C : Type u} [Category.{v} C] [HasFiniteProducts C] (X : C) + (exp : C ⥤ C) (adj : MonoidalCategory.tensorLeft X ⊣ exp) : + Exponentiable X where + adj := adj + /-- If `X` and `Y` are exponentiable then `X ⨯ Y` is. This isn't an instance because it's not usually how we want to construct exponentials, we'll usually prove all objects are exponential uniformly. @@ -78,10 +85,9 @@ abbrev CartesianClosed (C : Type u) [Category.{v} C] [HasFiniteProducts C] := -- Porting note: added to ease the port of `CategoryTheory.Closed.Types` /-- Constructor for `CartesianClosed C`. -/ def CartesianClosed.mk (C : Type u) [Category.{v} C] [HasFiniteProducts C] - (h : ∀ X, IsLeftAdjoint (@MonoidalCategory.tensorLeft _ _ - (monoidalOfHasFiniteProducts C) X)) : - CartesianClosed C := - ⟨fun X => ⟨h X⟩⟩ + (exp : ∀ (X : C), Exponentiable X) : + CartesianClosed C where + closed X := exp X variable {C : Type u} [Category.{v} C] (A B : C) {X X' Y Y' Z : C} variable [HasFiniteProducts C] [Exponentiable A] @@ -203,12 +209,12 @@ theorem uncurry_natural_left (f : X ⟶ X') (g : X' ⟶ A ⟹ Y) : @[simp] theorem uncurry_curry (f : A ⨯ X ⟶ Y) : uncurry (curry f) = f := - (Closed.isAdj.adj.homEquiv _ _).left_inv f + (Closed.adj.homEquiv _ _).left_inv f #align category_theory.cartesian_closed.uncurry_curry CategoryTheory.CartesianClosed.uncurry_curry @[simp] theorem curry_uncurry (f : X ⟶ A ⟹ Y) : curry (uncurry f) = f := - (Closed.isAdj.adj.homEquiv _ _).right_inv f + (Closed.adj.homEquiv _ _).right_inv f #align category_theory.cartesian_closed.curry_uncurry CategoryTheory.CartesianClosed.curry_uncurry -- Porting note: extra `(exp.adjunction A)` argument was needed for elaboration to succeed. @@ -239,11 +245,11 @@ theorem curry_id_eq_coev (A X : C) [Exponentiable A] : curry (𝟙 _) = (exp.coe #align category_theory.cartesian_closed.curry_id_eq_coev CategoryTheory.CartesianClosed.curry_id_eq_coev theorem curry_injective : Function.Injective (curry : (A ⨯ Y ⟶ X) → (Y ⟶ A ⟹ X)) := - (Closed.isAdj.adj.homEquiv _ _).injective + (Closed.adj.homEquiv _ _).injective #align category_theory.cartesian_closed.curry_injective CategoryTheory.CartesianClosed.curry_injective theorem uncurry_injective : Function.Injective (uncurry : (Y ⟶ A ⟹ X) → (A ⨯ Y ⟶ X)) := - (Closed.isAdj.adj.homEquiv _ _).symm.injective + (Closed.adj.homEquiv _ _).symm.injective #align category_theory.cartesian_closed.uncurry_injective CategoryTheory.CartesianClosed.uncurry_injective end CartesianClosed @@ -405,33 +411,8 @@ variable [HasFiniteProducts D] Note we didn't require any coherence between the choice of finite products here, since we transport along the `prodComparison` isomorphism. -/ -def cartesianClosedOfEquiv (e : C ≌ D) [h : CartesianClosed C] : CartesianClosed D where - closed X := - { isAdj := by - haveI q : Exponentiable (e.inverse.obj X) := inferInstance - have : IsLeftAdjoint (prod.functor.obj (e.inverse.obj X)) := q.isAdj - have : e.functor ⋙ prod.functor.obj X ⋙ e.inverse ≅ - prod.functor.obj (e.inverse.obj X) := by - apply NatIso.ofComponents _ _ - · intro Y - apply asIso (prodComparison e.inverse X (e.functor.obj Y)) ≪≫ _ - apply prod.mapIso (Iso.refl _) (e.unitIso.app Y).symm - · intro Y Z g - dsimp - simp [prodComparison, prod.comp_lift, ← e.inverse.map_comp, ← e.inverse.map_comp_assoc] - -- I wonder if it would be a good idea to - -- make `map_comp` a simp lemma the other way round - · have : IsLeftAdjoint (e.functor ⋙ prod.functor.obj X ⋙ e.inverse) := - Adjunction.leftAdjointOfNatIso this.symm - have : IsLeftAdjoint (e.inverse ⋙ e.functor ⋙ prod.functor.obj X ⋙ e.inverse) := - Adjunction.leftAdjointOfComp e.inverse _ - have : - (e.inverse ⋙ e.functor ⋙ prod.functor.obj X ⋙ e.inverse) ⋙ e.functor ≅ - prod.functor.obj X := by - apply isoWhiskerRight e.counitIso (prod.functor.obj X ⋙ e.inverse ⋙ e.functor) ≪≫ _ - change prod.functor.obj X ⋙ e.inverse ⋙ e.functor ≅ prod.functor.obj X - apply isoWhiskerLeft (prod.functor.obj X) e.counitIso - apply Adjunction.leftAdjointOfNatIso this } +def cartesianClosedOfEquiv (e : C ≌ D) [CartesianClosed C] : CartesianClosed D := + MonoidalClosed.ofEquiv (e.inverse.toMonoidalFunctorOfHasFiniteProducts) e.symm.toAdjunction #align category_theory.cartesian_closed_of_equiv CategoryTheory.cartesianClosedOfEquiv end Functor diff --git a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean index b028622bd0573..96b5df2a9020c 100644 --- a/Mathlib/CategoryTheory/Closed/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Closed/FunctorCategory.lean @@ -65,24 +65,19 @@ closed in the functor category `F : D ⥤ C` with the pointwise monoidal structu -- Porting note: removed `@[simps]`, as some of the generated lemmas were failing the simpNF linter, -- and none of the generated lemmmas was actually used in mathlib3. instance closed (F : D ⥤ C) : Closed F where - isAdj := - { right := closedIhom F - adj := Adjunction.mkOfUnitCounit + rightAdj := closedIhom F + adj := + Adjunction.mkOfUnitCounit { unit := closedUnit F - counit := closedCounit F } } + counit := closedCounit F } #align category_theory.functor.closed CategoryTheory.Functor.closed /-- If `C` is a monoidal closed category and `D` is groupoid, then the functor category `D ⥤ C`, with the pointwise monoidal structure, is monoidal closed. -/ -@[simps!] +@[simps! closed_adj] instance monoidalClosed : MonoidalClosed (D ⥤ C) where - closed := by infer_instance #align category_theory.functor.monoidal_closed CategoryTheory.Functor.monoidalClosed --- These lemmas have always been bad (#7657), but leanprover/lean4#2644 made `simp` start noticing -attribute [nolint simpNF] Functor.monoidalClosed_closed_isAdj_adj_homEquiv_apply_app - Functor.monoidalClosed_closed_isAdj_adj_homEquiv_symm_apply_app - theorem ihom_map (F : D ⥤ C) {G H : D ⥤ C} (f : G ⟶ H) : (ihom F).map f = (closedIhom F).map f := rfl #align category_theory.functor.ihom_map CategoryTheory.Functor.ihom_map diff --git a/Mathlib/CategoryTheory/Closed/Ideal.lean b/Mathlib/CategoryTheory/Closed/Ideal.lean index 2c030b8e01687..b2a6dfeb030f8 100644 --- a/Mathlib/CategoryTheory/Closed/Ideal.lean +++ b/Mathlib/CategoryTheory/Closed/Ideal.lean @@ -7,6 +7,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts import Mathlib.CategoryTheory.Monad.Limits import Mathlib.CategoryTheory.Adjunction.FullyFaithful +import Mathlib.CategoryTheory.Adjunction.Limits import Mathlib.CategoryTheory.Adjunction.Reflective import Mathlib.CategoryTheory.Closed.Cartesian import Mathlib.CategoryTheory.Subterminal @@ -78,12 +79,12 @@ the presence of a natural isomorphism `i ⋙ exp A ⋙ leftAdjoint i ⋙ i ≅ i The converse is given in `ExponentialIdeal.mk_of_iso`. -/ def exponentialIdealReflective (A : C) [Reflective i] [ExponentialIdeal i] : - i ⋙ exp A ⋙ leftAdjoint i ⋙ i ≅ i ⋙ exp A := by + i ⋙ exp A ⋙ reflector i ⋙ i ≅ i ⋙ exp A := by symm apply NatIso.ofComponents _ _ · intro X haveI := Functor.essImage.unit_isIso (ExponentialIdeal.exp_closed (i.obj_mem_essImage X) A) - apply asIso ((Adjunction.ofRightAdjoint i).unit.app (A ⟹ i.obj X)) + apply asIso ((reflectorAdjunction i).unit.app (A ⟹ i.obj X)) · simp [asIso] #align category_theory.exponential_ideal_reflective CategoryTheory.exponentialIdealReflective @@ -91,7 +92,7 @@ def exponentialIdealReflective (A : C) [Reflective i] [ExponentialIdeal i] : is an exponential ideal. -/ theorem ExponentialIdeal.mk_of_iso [Reflective i] - (h : ∀ A : C, i ⋙ exp A ⋙ leftAdjoint i ⋙ i ≅ i ⋙ exp A) : ExponentialIdeal i := by + (h : ∀ A : C, i ⋙ exp A ⋙ reflector i ⋙ i ≅ i ⋙ exp A) : ExponentialIdeal i := by apply ExponentialIdeal.mk' intro B A exact ⟨_, ⟨(h A).app B⟩⟩ @@ -120,9 +121,9 @@ variable [HasFiniteProducts C] [Reflective i] [CartesianClosed C] [HasFiniteProd This is the converse of `preservesBinaryProductsOfExponentialIdeal`. -/ instance (priority := 10) exponentialIdeal_of_preservesBinaryProducts - [PreservesLimitsOfShape (Discrete WalkingPair) (leftAdjoint i)] : ExponentialIdeal i := by - let ir := Adjunction.ofRightAdjoint i - let L : C ⥤ D := leftAdjoint i + [PreservesLimitsOfShape (Discrete WalkingPair) (reflector i)] : ExponentialIdeal i := by + let ir := reflectorAdjunction i + let L : C ⥤ D := reflector i let η : 𝟭 C ⟶ L ⋙ i := ir.unit let ε : i ⋙ L ⟶ 𝟭 D := ir.counit apply ExponentialIdeal.mk' @@ -149,18 +150,17 @@ itself cartesian closed. def cartesianClosedOfReflective : CartesianClosed D := { __ := monoidalOfHasFiniteProducts D -- Porting note (#10754): added this instance closed := fun B => - { isAdj := - { right := i ⋙ exp (i.obj B) ⋙ leftAdjoint i - adj := by - apply Adjunction.restrictFullyFaithful i i (exp.adjunction (i.obj B)) - · symm - refine' NatIso.ofComponents (fun X => _) (fun f => _) - · haveI := - Adjunction.rightAdjointPreservesLimits.{0, 0} (Adjunction.ofRightAdjoint i) - apply asIso (prodComparison i B X) - · dsimp [asIso] - rw [prodComparison_natural, Functor.map_id] - · apply (exponentialIdealReflective i _).symm } } } + { rightAdj :=i ⋙ exp (i.obj B) ⋙ reflector i + adj := by + apply Adjunction.restrictFullyFaithful i i (exp.adjunction (i.obj B)) + · symm + refine' NatIso.ofComponents (fun X => _) (fun f => _) + · haveI := + Adjunction.rightAdjointPreservesLimits.{0, 0} (reflectorAdjunction i) + apply asIso (prodComparison i B X) + · dsimp [asIso] + rw [prodComparison_natural, Functor.map_id] + · apply (exponentialIdealReflective i _).symm } } #align category_theory.cartesian_closed_of_reflective CategoryTheory.cartesianClosedOfReflective -- It's annoying that I need to do this. @@ -177,26 +177,26 @@ Together these help show that `L` preserves binary products. This should be cons *internal implementation* towards `preservesBinaryProductsOfExponentialIdeal`. -/ noncomputable def bijection (A B : C) (X : D) : - ((leftAdjoint i).obj (A ⨯ B) ⟶ X) ≃ ((leftAdjoint i).obj A ⨯ (leftAdjoint i).obj B ⟶ X) := + ((reflector i).obj (A ⨯ B) ⟶ X) ≃ ((reflector i).obj A ⨯ (reflector i).obj B ⟶ X) := calc - _ ≃ (A ⨯ B ⟶ i.obj X) := (Adjunction.ofRightAdjoint i).homEquiv _ _ + _ ≃ (A ⨯ B ⟶ i.obj X) := (reflectorAdjunction i).homEquiv _ _ _ ≃ (B ⨯ A ⟶ i.obj X) := (Limits.prod.braiding _ _).homCongr (Iso.refl _) _ ≃ (A ⟶ B ⟹ i.obj X) := (exp.adjunction _).homEquiv _ _ - _ ≃ (i.obj ((leftAdjoint i).obj A) ⟶ B ⟹ i.obj X) := + _ ≃ (i.obj ((reflector i).obj A) ⟶ B ⟹ i.obj X) := (unitCompPartialBijective _ (ExponentialIdeal.exp_closed (i.obj_mem_essImage _) _)) - _ ≃ (B ⨯ i.obj ((leftAdjoint i).obj A) ⟶ i.obj X) := ((exp.adjunction _).homEquiv _ _).symm - _ ≃ (i.obj ((leftAdjoint i).obj A) ⨯ B ⟶ i.obj X) := + _ ≃ (B ⨯ i.obj ((reflector i).obj A) ⟶ i.obj X) := ((exp.adjunction _).homEquiv _ _).symm + _ ≃ (i.obj ((reflector i).obj A) ⨯ B ⟶ i.obj X) := ((Limits.prod.braiding _ _).homCongr (Iso.refl _)) - _ ≃ (B ⟶ i.obj ((leftAdjoint i).obj A) ⟹ i.obj X) := (exp.adjunction _).homEquiv _ _ - _ ≃ (i.obj ((leftAdjoint i).obj B) ⟶ i.obj ((leftAdjoint i).obj A) ⟹ i.obj X) := + _ ≃ (B ⟶ i.obj ((reflector i).obj A) ⟹ i.obj X) := (exp.adjunction _).homEquiv _ _ + _ ≃ (i.obj ((reflector i).obj B) ⟶ i.obj ((reflector i).obj A) ⟹ i.obj X) := (unitCompPartialBijective _ (ExponentialIdeal.exp_closed (i.obj_mem_essImage _) _)) - _ ≃ (i.obj ((leftAdjoint i).obj A) ⨯ i.obj ((leftAdjoint i).obj B) ⟶ i.obj X) := + _ ≃ (i.obj ((reflector i).obj A) ⨯ i.obj ((reflector i).obj B) ⟶ i.obj X) := ((exp.adjunction _).homEquiv _ _).symm - _ ≃ (i.obj ((leftAdjoint i).obj A ⨯ (leftAdjoint i).obj B) ⟶ i.obj X) := - haveI : PreservesLimits i := (Adjunction.ofRightAdjoint i).rightAdjointPreservesLimits + _ ≃ (i.obj ((reflector i).obj A ⨯ (reflector i).obj B) ⟶ i.obj X) := + haveI : PreservesLimits i := (reflectorAdjunction i).rightAdjointPreservesLimits haveI := preservesSmallestLimitsOfPreservesLimits i Iso.homCongr (PreservesLimitPair.iso _ _ _).symm (Iso.refl (i.obj X)) - _ ≃ ((leftAdjoint i).obj A ⨯ (leftAdjoint i).obj B ⟶ X) := (equivOfFullyFaithful _).symm + _ ≃ ((reflector i).obj A ⨯ (reflector i).obj B ⟶ X) := (equivOfFullyFaithful _).symm #align category_theory.bijection CategoryTheory.bijection theorem bijection_symm_apply_id (A B : C) : @@ -211,15 +211,15 @@ theorem bijection_symm_apply_id (A B : C) : dsimp only [Functor.comp_obj] rw [prod.comp_lift_assoc, prod.lift_snd, prod.lift_fst_assoc, prod.lift_fst_comp_snd_comp, ← Adjunction.eq_homEquiv_apply, Adjunction.homEquiv_unit, Iso.comp_inv_eq, assoc] - rw [PreservesLimitPair.iso_hom i ((leftAdjoint i).obj A) ((leftAdjoint i).obj B)] + rw [PreservesLimitPair.iso_hom i ((reflector i).obj A) ((reflector i).obj B)] apply prod.hom_ext · rw [Limits.prod.map_fst, assoc, assoc, prodComparison_fst, ← i.map_comp, prodComparison_fst] - apply (Adjunction.ofRightAdjoint i).unit.naturality + apply (reflectorAdjunction i).unit.naturality · rw [Limits.prod.map_snd, assoc, assoc, prodComparison_snd, ← i.map_comp, prodComparison_snd] - apply (Adjunction.ofRightAdjoint i).unit.naturality + apply (reflectorAdjunction i).unit.naturality #align category_theory.bijection_symm_apply_id CategoryTheory.bijection_symm_apply_id -theorem bijection_natural (A B : C) (X X' : D) (f : (leftAdjoint i).obj (A ⨯ B) ⟶ X) (g : X ⟶ X') : +theorem bijection_natural (A B : C) (X X' : D) (f : (reflector i).obj (A ⨯ B) ⟶ X) (g : X ⟶ X') : bijection i _ _ _ (f ≫ g) = bijection i _ _ _ f ≫ g := by dsimp [bijection] -- Porting note: added @@ -236,7 +236,7 @@ theorem bijection_natural (A B : C) (X X' : D) (f : (leftAdjoint i).obj (A ⨯ B The bijection allows us to show that `prodComparison L A B` is an isomorphism, where the inverse is the forward map of the identity morphism. -/ -theorem prodComparison_iso (A B : C) : IsIso (prodComparison (leftAdjoint i) A B) := +theorem prodComparison_iso (A B : C) : IsIso (prodComparison (reflector i) A B) := ⟨⟨bijection i _ _ _ (𝟙 _), by rw [← (bijection i _ _ _).injective.eq_iff, bijection_natural, ← bijection_symm_apply_id, Equiv.apply_symm_apply, id_comp], @@ -250,10 +250,10 @@ If a reflective subcategory is an exponential ideal, then the reflector preserve This is the converse of `exponentialIdeal_of_preserves_binary_products`. -/ noncomputable def preservesBinaryProductsOfExponentialIdeal : - PreservesLimitsOfShape (Discrete WalkingPair) (leftAdjoint i) where + PreservesLimitsOfShape (Discrete WalkingPair) (reflector i) where preservesLimit {K} := letI := PreservesLimitPair.ofIsoProdComparison - (leftAdjoint i) (K.obj ⟨WalkingPair.left⟩) (K.obj ⟨WalkingPair.right⟩) + (reflector i) (K.obj ⟨WalkingPair.left⟩) (K.obj ⟨WalkingPair.right⟩) Limits.preservesLimitOfIsoDiagram _ (diagramIsoPair K).symm #align category_theory.preserves_binary_products_of_exponential_ideal CategoryTheory.preservesBinaryProductsOfExponentialIdeal @@ -261,10 +261,10 @@ noncomputable def preservesBinaryProductsOfExponentialIdeal : If a reflective subcategory is an exponential ideal, then the reflector preserves finite products. -/ noncomputable def preservesFiniteProductsOfExponentialIdeal (J : Type) [Fintype J] : - PreservesLimitsOfShape (Discrete J) (leftAdjoint i) := by + PreservesLimitsOfShape (Discrete J) (reflector i) := by letI := preservesBinaryProductsOfExponentialIdeal i - letI := leftAdjointPreservesTerminalOfReflective.{0} i - apply preservesFiniteProductsOfPreservesBinaryAndTerminal (leftAdjoint i) J + letI : PreservesLimitsOfShape _ (reflector i) := leftAdjointPreservesTerminalOfReflective.{0} i + apply preservesFiniteProductsOfPreservesBinaryAndTerminal (reflector i) J #align category_theory.preserves_finite_products_of_exponential_ideal CategoryTheory.preservesFiniteProductsOfExponentialIdeal end diff --git a/Mathlib/CategoryTheory/Closed/Monoidal.lean b/Mathlib/CategoryTheory/Closed/Monoidal.lean index 7e957c3c97732..ce61aca35df54 100644 --- a/Mathlib/CategoryTheory/Closed/Monoidal.lean +++ b/Mathlib/CategoryTheory/Closed/Monoidal.lean @@ -32,13 +32,15 @@ open Category MonoidalCategory -- so definitional properties of instances may be important. /-- An object `X` is (right) closed if `(X ⊗ -)` is a left adjoint. -/ class Closed {C : Type u} [Category.{v} C] [MonoidalCategory.{v} C] (X : C) where - /-- a choice of a left adjoint for `tensorLeft X` -/ - isAdj : IsLeftAdjoint (tensorLeft X) + /-- a choice of a right adjoint for `tensorLeft X` -/ + rightAdj : C ⥤ C + /-- `tensorLeft X` is a left adjoint -/ + adj : tensorLeft X ⊣ rightAdj #align category_theory.closed CategoryTheory.Closed /-- A monoidal category `C` is (right) monoidal closed if every object is (right) closed. -/ class MonoidalClosed (C : Type u) [Category.{v} C] [MonoidalCategory.{v} C] where - closed : ∀ X : C, Closed X + closed (X : C) : Closed X := by infer_instance #align category_theory.monoidal_closed CategoryTheory.MonoidalClosed attribute [instance 100] MonoidalClosed.closed @@ -50,10 +52,7 @@ This isn't an instance because it's not usually how we want to construct interna we'll usually prove all objects are closed uniformly. -/ def tensorClosed {X Y : C} (hX : Closed X) (hY : Closed Y) : Closed (X ⊗ Y) where - isAdj := by - haveI := hX.isAdj - haveI := hY.isAdj - exact Adjunction.leftAdjointOfNatIso (MonoidalCategory.tensorLeftTensor _ _).symm + adj := (hY.adj.comp hX.adj).ofNatIsoLeft (MonoidalCategory.tensorLeftTensor X Y).symm #align category_theory.tensor_closed CategoryTheory.tensorClosed /-- The unit object is always closed. @@ -61,22 +60,8 @@ This isn't an instance because most of the time we'll prove closedness for all o rather than just for this one. -/ def unitClosed : Closed (𝟙_ C) where - isAdj := - { right := 𝟭 C - adj := - Adjunction.mkOfHomEquiv - { homEquiv := fun X _ => - { toFun := fun a => (leftUnitor X).inv ≫ a - invFun := fun a => (leftUnitor X).hom ≫ a - left_inv := by aesop_cat - right_inv := by aesop_cat } - homEquiv_naturality_left_symm := fun f g => by - dsimp - rw [leftUnitor_naturality_assoc] - -- This used to be automatic before leanprover/lean4#2644 - homEquiv_naturality_right := by -- aesop failure - dsimp - simp }} + rightAdj := 𝟭 C + adj := Adjunction.id.ofNatIsoLeft (MonoidalCategory.leftUnitorNatIso C).symm #align category_theory.unit_closed CategoryTheory.unitClosed variable (A B : C) {X X' Y Y' Z : C} @@ -85,14 +70,14 @@ variable [Closed A] /-- This is the internal hom `A ⟶[C] -`. -/ def ihom : C ⥤ C := - (@Closed.isAdj _ _ _ A _).right + Closed.rightAdj (X := A) #align category_theory.ihom CategoryTheory.ihom namespace ihom /-- The adjunction between `A ⊗ -` and `A ⟹ -`. -/ def adjunction : tensorLeft A ⊣ ihom A := - Closed.isAdj.adj + Closed.adj #align category_theory.ihom.adjunction CategoryTheory.ihom.adjunction /-- The evaluation natural transformation. -/ @@ -201,12 +186,12 @@ theorem uncurry_natural_left (f : X ⟶ X') (g : X' ⟶ A ⟶[C] Y) : @[simp] theorem uncurry_curry (f : A ⊗ X ⟶ Y) : uncurry (curry f) = f := - (Closed.isAdj.adj.homEquiv _ _).left_inv f + (Closed.adj.homEquiv _ _).left_inv f #align category_theory.monoidal_closed.uncurry_curry CategoryTheory.MonoidalClosed.uncurry_curry @[simp] theorem curry_uncurry (f : X ⟶ A ⟶[C] Y) : curry (uncurry f) = f := - (Closed.isAdj.adj.homEquiv _ _).right_inv f + (Closed.adj.homEquiv _ _).right_inv f #align category_theory.monoidal_closed.curry_uncurry CategoryTheory.MonoidalClosed.curry_uncurry theorem curry_eq_iff (f : A ⊗ Y ⟶ X) (g : Y ⟶ A ⟶[C] X) : curry f = g ↔ f = uncurry g := @@ -227,11 +212,11 @@ theorem curry_eq (g : A ⊗ Y ⟶ X) : curry g = (ihom.coev A).app Y ≫ (ihom A #align category_theory.monoidal_closed.curry_eq CategoryTheory.MonoidalClosed.curry_eq theorem curry_injective : Function.Injective (curry : (A ⊗ Y ⟶ X) → (Y ⟶ A ⟶[C] X)) := - (Closed.isAdj.adj.homEquiv _ _).injective + (Closed.adj.homEquiv _ _).injective #align category_theory.monoidal_closed.curry_injective CategoryTheory.MonoidalClosed.curry_injective theorem uncurry_injective : Function.Injective (uncurry : (Y ⟶ A ⟶[C] X) → (A ⊗ Y ⟶ X)) := - (Closed.isAdj.adj.homEquiv _ _).symm.injective + (Closed.adj.homEquiv _ _).symm.injective #align category_theory.monoidal_closed.uncurry_injective CategoryTheory.MonoidalClosed.uncurry_injective variable (A X) @@ -303,15 +288,15 @@ section OfEquiv variable {D : Type u₂} [Category.{v₂} D] [MonoidalCategory.{v₂} D] +variable (F : MonoidalFunctor C D) {G : D ⥤ C} (adj : F.toFunctor ⊣ G) + [F.IsEquivalence] [MonoidalClosed D] /-- Transport the property of being monoidal closed across a monoidal equivalence of categories -/ -noncomputable def ofEquiv (F : MonoidalFunctor C D) [F.IsEquivalence] - [h : MonoidalClosed D] : MonoidalClosed C where +noncomputable def ofEquiv : MonoidalClosed C where closed X := - { isAdj := by - haveI q : Closed (F.obj X) := inferInstance - haveI : IsLeftAdjoint (tensorLeft (F.obj X)) := q.isAdj - have i := (MonoidalFunctor.commTensorLeft F X).compInvIso - exact Adjunction.leftAdjointOfNatIso i } + { rightAdj := F.toFunctor ⋙ ihom (F.obj X) ⋙ G + adj := (adj.comp ((ihom.adjunction (F.obj X)).comp + adj.toEquivalence.symm.toAdjunction)).ofNatIsoLeft + (Iso.compInverseIso (H := adj.toEquivalence) (MonoidalFunctor.commTensorLeft F X)) } #align category_theory.monoidal_closed.of_equiv CategoryTheory.MonoidalClosed.ofEquiv /-- Suppose we have a monoidal equivalence `F : C ≌ D`, with `D` monoidal closed. We can pull the @@ -319,13 +304,13 @@ monoidal closed instance back along the equivalence. For `X, Y, Z : C`, this lem resulting currying map `Hom(X ⊗ Y, Z) → Hom(Y, (X ⟶[C] Z))`. (`X ⟶[C] Z` is defined to be `F⁻¹(F(X) ⟶[D] F(Z))`, so currying in `C` is given by essentially conjugating currying in `D` by `F.`) -/ -theorem ofEquiv_curry_def (F : MonoidalFunctor C D) [F.IsEquivalence] - [MonoidalClosed D] {X Y Z : C} (f : X ⊗ Y ⟶ Z) : - @MonoidalClosed.curry _ _ _ _ _ _ ((MonoidalClosed.ofEquiv F).1 _) f = - (F.1.1.adjunction.homEquiv Y ((ihom _).obj _)) - (MonoidalClosed.curry - (F.1.1.inv.adjunction.homEquiv (F.1.1.obj X ⊗ F.1.1.obj Y) Z - (((F.commTensorLeft X).compInvIso).hom.app Y ≫ f))) := +theorem ofEquiv_curry_def {X Y Z : C} (f : X ⊗ Y ⟶ Z) : + letI := ofEquiv F adj + MonoidalClosed.curry f = + adj.homEquiv Y ((ihom (F.obj X)).obj (F.obj Z)) + (MonoidalClosed.curry (adj.toEquivalence.symm.toAdjunction.homEquiv (F.obj X ⊗ F.obj Y) Z + ((Iso.compInverseIso (H := adj.toEquivalence) + (MonoidalFunctor.commTensorLeft F X)).hom.app Y ≫ f))) := rfl #align category_theory.monoidal_closed.of_equiv_curry_def CategoryTheory.MonoidalClosed.ofEquiv_curry_def @@ -334,15 +319,14 @@ monoidal closed instance back along the equivalence. For `X, Y, Z : C`, this lem resulting uncurrying map `Hom(Y, (X ⟶[C] Z)) → Hom(X ⊗ Y ⟶ Z)`. (`X ⟶[C] Z` is defined to be `F⁻¹(F(X) ⟶[D] F(Z))`, so uncurrying in `C` is given by essentially conjugating uncurrying in `D` by `F.`) -/ -theorem ofEquiv_uncurry_def (F : MonoidalFunctor C D) [F.IsEquivalence] - [MonoidalClosed D] {X Y Z : C} - (f : Y ⟶ (@ihom _ _ _ X <| (MonoidalClosed.ofEquiv F).1 X).obj Z) : - @MonoidalClosed.uncurry _ _ _ _ _ _ ((MonoidalClosed.ofEquiv F).1 _) f = - ((F.commTensorLeft X).compInvIso).inv.app Y ≫ - (F.1.1.inv.adjunction.homEquiv (F.1.1.obj X ⊗ F.1.1.obj Y) Z).symm - (MonoidalClosed.uncurry - ((F.1.1.adjunction.homEquiv Y ((ihom (F.1.1.obj X)).obj (F.1.1.obj Z))).symm f)) := - rfl +theorem ofEquiv_uncurry_def {X Y Z : C} : + letI := ofEquiv F adj + ∀ (f : Y ⟶ (ihom X).obj Z), MonoidalClosed.uncurry f = + ((Iso.compInverseIso (H := adj.toEquivalence) + (MonoidalFunctor.commTensorLeft F X)).inv.app Y) ≫ + (adj.toEquivalence.symm.toAdjunction.homEquiv _ _).symm + (MonoidalClosed.uncurry ((adj.homEquiv _ _).symm f)) := + fun _ => rfl #align category_theory.monoidal_closed.of_equiv_uncurry_def CategoryTheory.MonoidalClosed.ofEquiv_uncurry_def end OfEquiv diff --git a/Mathlib/CategoryTheory/Closed/Types.lean b/Mathlib/CategoryTheory/Closed/Types.lean index 8606594d1591a..10444c5948599 100644 --- a/Mathlib/CategoryTheory/Closed/Types.lean +++ b/Mathlib/CategoryTheory/Closed/Types.lean @@ -31,22 +31,24 @@ variable {C : Type v₂} [Category.{v₁} C] section CartesianClosed -instance (X : Type v₁) : IsLeftAdjoint (Types.binaryProductFunctor.obj X) where - right := - { obj := fun Y => X ⟶ Y - map := fun f g => g ≫ f } - adj := - Adjunction.mkOfUnitCounit - { unit := { app := fun Z (z : Z) x => ⟨x, z⟩ } - counit := { app := fun Z xf => xf.2 xf.1 } } +/-- The adjunction `Limits.Types.binaryProductFunctor.obj X ⊣ coyoneda.obj (Opposite.op X)` +for any `X : Type v₁`. -/ +def Types.binaryProductAdjunction (X : Type v₁) : + Limits.Types.binaryProductFunctor.obj X ⊣ coyoneda.obj (Opposite.op X) := + Adjunction.mkOfUnitCounit + { unit := { app := fun Z (z : Z) x => ⟨x, z⟩ } + counit := { app := fun Z xf => xf.2 xf.1 } } + +instance (X : Type v₁) : (Types.binaryProductFunctor.obj X).IsLeftAdjoint := + ⟨_, ⟨Types.binaryProductAdjunction X⟩⟩ -- Porting note: this instance should be moved to a higher file. instance : HasFiniteProducts (Type v₁) := hasFiniteProducts_of_hasProducts.{v₁} _ -instance : CartesianClosed (Type v₁) := - CartesianClosed.mk _ - (fun X => Adjunction.leftAdjointOfNatIso (Types.binaryProductIsoProd.app X)) +instance : CartesianClosed (Type v₁) := CartesianClosed.mk _ + (fun X => Exponentiable.mk _ _ + ((Types.binaryProductAdjunction X).ofNatIsoLeft (Types.binaryProductIsoProd.app X))) -- Porting note: in mathlib3, the assertion was for `(C ⥤ Type u₁)`, but then Lean4 was -- confused with universes. It makes no harm to relax the universe assumptions here. @@ -55,9 +57,10 @@ instance {C : Type u₁} [Category.{v₁} C] : HasFiniteProducts (C ⥤ Type u instance {C : Type v₁} [SmallCategory C] : CartesianClosed (C ⥤ Type v₁) := CartesianClosed.mk _ - (fun F => + (fun F => by letI := FunctorCategory.prodPreservesColimits F - isLeftAdjointOfPreservesColimits (prod.functor.obj F)) + have := isLeftAdjointOfPreservesColimits (prod.functor.obj F) + exact Exponentiable.mk _ _ (Adjunction.ofIsLeftAdjoint (prod.functor.obj F))) end CartesianClosed diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index 8bb688ea64892..22c77aea59012 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -360,9 +360,9 @@ def mapArrowEquivalence (e : C ≌ D) : Arrow C ≌ Arrow D where unitIso := Functor.mapIso (mapArrowFunctor C C) e.unitIso counitIso := Functor.mapIso (mapArrowFunctor D D) e.counitIso -instance isEquivalenceMapArrow (F : C ⥤ D) [IsEquivalence F] : +instance isEquivalence_mapArrow (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.mapArrow := - IsEquivalence.ofEquivalence (mapArrowEquivalence (asEquivalence F)) + (mapArrowEquivalence (asEquivalence F)).isEquivalence_functor end Functor diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index e304e9490ded4..bf05bc59e8639 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -341,10 +341,8 @@ instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.EssSurj] : (preLeft F L R) ⟨isoMk (F.objObjPreimageIso _) (Iso.refl _) (by simp)⟩⟩ /-- If `F` is an equivalence, then so is `preLeft F L R`. -/ -noncomputable def isEquivalencePreLeft (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.IsEquivalence] : - (preLeft F L R).IsEquivalence := - have := Equivalence.essSurj_of_equivalence F - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance isEquivalence_preLeft (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.IsEquivalence] : + (preLeft F L R).IsEquivalence where /-- The functor `(F ⋙ L, R) ⥤ (L, R)` -/ @[simps] @@ -372,10 +370,8 @@ instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.EssSurj] : (preRight L F R ⟨isoMk (Iso.refl _) (F.objObjPreimageIso _) (by simp [← R.map_comp])⟩⟩ /-- If `F` is an equivalence, then so is `preRight L F R`. -/ -noncomputable def isEquivalencePreRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.IsEquivalence] : - (preRight L F R).IsEquivalence := - have := Equivalence.essSurj_of_equivalence F - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance isEquivalence_preRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.IsEquivalence] : + (preRight L F R).IsEquivalence where /-- The functor `(L, R) ⥤ (L ⋙ F, R ⋙ F)` -/ @[simps] diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 2fba56ef0f6c0..4af6bcb9d97c5 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -329,9 +329,9 @@ instance (F : D ⥤ T) (X : T) [F.EssSurj] : (toOver F X).EssSurj := show (CostructuredArrow.pre _ _ _).EssSurj from inferInstance /-- An equivalence `F` induces an equivalence `CostructuredArrow F X ≌ Over X`. -/ -noncomputable def isEquivalenceToOver (F : D ⥤ T) (X : T) [F.IsEquivalence] : +instance isEquivalence_toOver (F : D ⥤ T) (X : T) [F.IsEquivalence] : (toOver F X).IsEquivalence := - CostructuredArrow.isEquivalencePre _ _ _ + CostructuredArrow.isEquivalence_pre _ _ _ end CostructuredArrow @@ -567,9 +567,9 @@ instance (X : T) (F : D ⥤ T) [F.EssSurj] : (toUnder X F).EssSurj := show (StructuredArrow.pre _ _ _).EssSurj from inferInstance /-- An equivalence `F` induces an equivalence `StructuredArrow X F ≌ Under X`. -/ -noncomputable def isEquivalenceToUnder (X : T) (F : D ⥤ T) [F.IsEquivalence] : +instance isEquivalence_toUnder (X : T) (F : D ⥤ T) [F.IsEquivalence] : (toUnder X F).IsEquivalence := - StructuredArrow.isEquivalencePre _ _ _ + StructuredArrow.isEquivalence_pre _ _ _ end StructuredArrow diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean index e325ccaa28b98..73b62fed6932e 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean @@ -302,9 +302,9 @@ instance (S : D) (F : B ⥤ C) (G : C ⥤ D) [F.EssSurj] : (pre S F G).EssSurj : show (Comma.preRight _ _ _).EssSurj from inferInstance /-- If `F` is an equivalence, then so is the functor `(S, F ⋙ G) ⥤ (S, G)`. -/ -noncomputable def isEquivalencePre (S : D) (F : B ⥤ C) (G : C ⥤ D) [F.IsEquivalence] : +instance isEquivalence_pre (S : D) (F : B ⥤ C) (G : C ⥤ D) [F.IsEquivalence] : (pre S F G).IsEquivalence := - Comma.isEquivalencePreRight _ _ _ + Comma.isEquivalence_preRight _ _ _ /-- The functor `(S, F) ⥤ (G(S), F ⋙ G)`. -/ @[simps] @@ -324,9 +324,8 @@ instance (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] : (post S F G).EssSurj whe mem_essImage h := ⟨mk (G.preimage h.hom), ⟨isoMk (Iso.refl _) (by simp)⟩⟩ /-- If `G` is fully faithful, then `post S F G : (S, F) ⥤ (G(S), F ⋙ G)` is an equivalence. -/ -noncomputable def isEquivalencePost (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] [G.Faithful] : - (post S F G).IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance isEquivalence_post (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] [G.Faithful] : + (post S F G).IsEquivalence where instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by @@ -640,9 +639,9 @@ instance (F : B ⥤ C) (G : C ⥤ D) (S : D) [F.EssSurj] : (pre F G S).EssSurj : show (Comma.preLeft _ _ _).EssSurj from inferInstance /-- If `F` is an equivalence, then so is the functor `(F ⋙ G, S) ⥤ (G, S)`. -/ -noncomputable def isEquivalencePre (F : B ⥤ C) (G : C ⥤ D) (S : D) [F.IsEquivalence] : +instance isEquivalence_pre (F : B ⥤ C) (G : C ⥤ D) (S : D) [F.IsEquivalence] : (pre F G S).IsEquivalence := - Comma.isEquivalencePreLeft _ _ _ + Comma.isEquivalence_preLeft _ _ _ /-- The functor `(F, S) ⥤ (F ⋙ G, G(S))`. -/ @[simps] @@ -662,9 +661,8 @@ instance (F : B ⥤ C) (G : C ⥤ D) (S : C) [G.Full] : (post F G S).EssSurj whe mem_essImage h := ⟨mk (G.preimage h.hom), ⟨isoMk (Iso.refl _) (by simp)⟩⟩ /-- If `G` is fully faithful, then `post F G S : (F, S) ⥤ (F ⋙ G, G(S))` is an equivalence. -/ -noncomputable def isEquivalencePost (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] [G.Faithful] : - (post F G S).IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance isEquivalence_post (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] [G.Faithful] : + (post F G S).IsEquivalence where instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by diff --git a/Mathlib/CategoryTheory/ConnectedComponents.lean b/Mathlib/CategoryTheory/ConnectedComponents.lean index 861c1f0f73a86..5fc65af9fc3ce 100644 --- a/Mathlib/CategoryTheory/ConnectedComponents.lean +++ b/Mathlib/CategoryTheory/ConnectedComponents.lean @@ -144,8 +144,7 @@ instance : (decomposedTo J).Faithful where instance : (decomposedTo J).EssSurj where mem_essImage j := ⟨⟨_, j, rfl⟩, ⟨Iso.refl _⟩⟩ -instance : (decomposedTo J).IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance : (decomposedTo J).IsEquivalence where -- Porting note: it was originally @[simps (config := { rhsMd := semireducible }) Functor] /-- This gives that any category is equivalent to a disjoint union of connected categories. -/ diff --git a/Mathlib/CategoryTheory/EffectiveEpi/Preserves.lean b/Mathlib/CategoryTheory/EffectiveEpi/Preserves.lean index 1406ae4a8ec28..8b2c520772926 100644 --- a/Mathlib/CategoryTheory/EffectiveEpi/Preserves.lean +++ b/Mathlib/CategoryTheory/EffectiveEpi/Preserves.lean @@ -209,8 +209,7 @@ instance (F : C ⥤ D) [IsEquivalence F] : F.ReflectsEffectiveEpiFamilies where reflects {α B} X π _ := by let i : (a : α) → X a ⟶ (inv F).obj (F.obj (X a)) := fun a ↦ (asEquivalence F).unit.app _ have : EffectiveEpiFamily X (fun a ↦ (i a) ≫ (inv F).map (F.map (π a))) := inferInstance - simp only [asEquivalence_functor, asEquivalence_inverse, IsEquivalence.inv_fun_map, comp_obj, - id_obj, Iso.hom_inv_id_app_assoc, i] at this + simp only [inv_fun_map, Iso.hom_inv_id_app_assoc, i] at this have : EffectiveEpiFamily X (fun a ↦ (π a ≫ (asEquivalence F).unit.app B) ≫ (asEquivalence F).unitInv.app _) := inferInstance simpa diff --git a/Mathlib/CategoryTheory/Equivalence.lean b/Mathlib/CategoryTheory/Equivalence.lean index 610be20a7a1d3..fd0b2d014ba90 100644 --- a/Mathlib/CategoryTheory/Equivalence.lean +++ b/Mathlib/CategoryTheory/Equivalence.lean @@ -37,18 +37,19 @@ if it is full, faithful and essentially surjective. ## Main definitions * `Equivalence`: bundled (half-)adjoint equivalences of categories -* `IsEquivalence`: type class on a functor `F` containing the data of the inverse `G` as well as - the natural isomorphisms `η` and `ε`. * `Functor.EssSurj`: type class on a functor `F` containing the data of the preimages and the isomorphisms `F.obj (preimage d) ≅ d`. +* `Functor.IsEquivalence`: type class on a functor `F` which is full, faithful and +essentially surjective. ## Main results * `Equivalence.mk`: upgrade an equivalence to a (half-)adjoint equivalence -* `Functor.IsEquivalence.equivOfIso`: when `F` and `G` are isomorphic functors, +* `isEquivalence_iff_of_iso`: when `F` and `G` are isomorphic functors, `F` is an equivalence iff `G` is. -* `Functor.IsEquivalence.ofFullyFaithfullyEssSurj`: a fully faithful essentially surjective -functor is an equivalence. +* `Functor.asEquivalenceFunctor`: construction of an equivalence of categories from +a functor `F` which satisfies the property `F.IsEquivalence` (i.e. `F` is full, faithful +and essentially surjective). ## Notations @@ -76,6 +77,7 @@ universe v₁ v₂ v₃ u₁ u₂ u₃ See -/ +@[ext] structure Equivalence (C : Type u₁) (D : Type u₂) [Category.{v₁} C] [Category.{v₂} D] where mk' :: /-- A functor in one direction -/ functor : C ⥤ D @@ -466,105 +468,140 @@ theorem pow_neg_one (e : C ≌ C) : e ^ (-1 : ℤ) = e.symm := -- TODO as necessary, add the natural isomorphisms `(e^a).trans e^b ≅ e^(a+b)`. -- At this point, we haven't even defined the category of equivalences. +-- Note: the better formulation of this would involve `HasShift`. end -end Equivalence +/-- The functor of an equivalence of categories is essentially surjective. -namespace Functor +See . +-/ +instance essSurj_functor (e : C ≌ E) : e.functor.EssSurj := + ⟨fun Y => ⟨e.inverse.obj Y, ⟨e.counitIso.app Y⟩⟩⟩ -/-- A functor that is part of a (half) adjoint equivalence -/ -class IsEquivalence (F : C ⥤ D) where mk' :: - /-- The inverse functor to `F` -/ - inverse : D ⥤ C - /-- Composition `F ⋙ inverse` is isomorphic to the identity. -/ - unitIso : 𝟭 C ≅ F ⋙ inverse - /-- Composition `inverse ⋙ F` is isomorphic to the identity. -/ - counitIso : inverse ⋙ F ≅ 𝟭 D - /-- The natural isomorphisms are inverse. -/ - functor_unitIso_comp : - ∀ X : C, - F.map ((unitIso.hom : 𝟭 C ⟶ F ⋙ inverse).app X) ≫ counitIso.hom.app (F.obj X) = - 𝟙 (F.obj X) := by - aesop_cat +instance essSurj_inverse (e : C ≌ E) : e.inverse.EssSurj := + e.symm.essSurj_functor + +/-- The functor of an equivalence of categories is faithful. + +See . +-/ +instance faithful_functor (e : C ≌ E) : e.functor.Faithful where + map_injective {X Y f g} h := by + rw [← cancel_mono (e.unit.app Y), ← cancel_epi (e.unitInv.app X), + ← e.inv_fun_map _ _ f, ← e.inv_fun_map _ _ g, h] + +instance faithful_inverse (e : C ≌ E) : e.inverse.Faithful := + e.symm.faithful_functor + +/-- The functor of an equivalence of categories is full. + +See . +-/ +instance full_functor (e : C ≌ E) : e.functor.Full where + map_surjective {X Y} f := + ⟨e.unitIso.hom.app X ≫ e.inverse.map f ≫ e.unitIso.inv.app Y, + e.inverse.map_injective (by simp)⟩ + +instance full_inverse (e : C ≌ E) : e.inverse.Full := + e.symm.full_functor + +/-- If `e : C ≌ D` is an equivalence of categories, and `iso : e.functor ≅ G` is +an isomorphism, then there is an equivalence of categories whose functor is `G`. -/ +@[simps!] +def changeFunctor (e : C ≌ D) {G : C ⥤ D} (iso : e.functor ≅ G) : C ≌ D where + functor := G + inverse := e.inverse + unitIso := e.unitIso ≪≫ isoWhiskerRight iso _ + counitIso := isoWhiskerLeft _ iso.symm ≪≫ e.counitIso + +/-- Compatibility of `changeFunctor` with identity isomorphisms of functors -/ +theorem changeFunctor_refl (e : C ≌ D) : e.changeFunctor (Iso.refl _) = e := by aesop_cat + +/-- Compatibility of `changeFunctor` with the composition of isomorphisms of functors -/ +theorem changeFunctor_trans (e : C ≌ D) {G G' : C ⥤ D} (iso₁ : e.functor ≅ G) (iso₂ : G ≅ G') : + (e.changeFunctor iso₁).changeFunctor iso₂ = e.changeFunctor (iso₁ ≪≫ iso₂) := by aesop_cat + +/-- If `e : C ≌ D` is an equivalence of categories, and `iso : e.functor ≅ G` is +an isomorphism, then there is an equivalence of categories whose inverse is `G`. -/ +@[simps!] +def changeInverse (e : C ≌ D) {G : D ⥤ C} (iso : e.inverse ≅ G) : C ≌ D where + functor := e.functor + inverse := G + unitIso := e.unitIso ≪≫ isoWhiskerLeft _ iso + counitIso := isoWhiskerRight iso.symm _ ≪≫ e.counitIso + functor_unitIso_comp X := by + dsimp + rw [← map_comp_assoc, assoc, iso.hom_inv_id_app, comp_id, functor_unit_comp] + +end Equivalence + +/-- A functor is an equivalence of categories if it is faithful, full and +essentially surjective. -/ +class Functor.IsEquivalence (F : C ⥤ D) : Prop where + faithful : F.Faithful := by infer_instance + full : F.Full := by infer_instance + essSurj : F.EssSurj := by infer_instance #align category_theory.is_equivalence CategoryTheory.Functor.IsEquivalence -#align category_theory.is_equivalence.unit_iso CategoryTheory.Functor.IsEquivalence.unitIso -#align category_theory.is_equivalence.counit_iso CategoryTheory.Functor.IsEquivalence.counitIso -#align category_theory.is_equivalence.functor_unit_iso_comp CategoryTheory.Functor.IsEquivalence.functor_unitIso_comp -attribute [reassoc (attr := simp)] IsEquivalence.functor_unitIso_comp +instance Equivalence.isEquivalence_functor (F : C ≌ D) : IsEquivalence F.functor where +#align category_theory.is_equivalence.of_equivalence CategoryTheory.Equivalence.isEquivalence_functor -namespace IsEquivalence +instance Equivalence.isEquivalence_inverse (F : C ≌ D) : IsEquivalence F.inverse := + F.symm.isEquivalence_functor +#align category_theory.is_equivalence.of_equivalence_inverse CategoryTheory.Equivalence.isEquivalence_inverse -instance ofEquivalence (F : C ≌ D) : IsEquivalence F.functor := - { F with } -#align category_theory.is_equivalence.of_equivalence CategoryTheory.Functor.IsEquivalence.ofEquivalence +namespace Functor -instance ofEquivalenceInverse (F : C ≌ D) : IsEquivalence F.inverse := - IsEquivalence.ofEquivalence F.symm -#align category_theory.is_equivalence.of_equivalence_inverse CategoryTheory.Functor.IsEquivalence.ofEquivalenceInverse +namespace IsEquivalence -open Equivalence +attribute [instance] faithful full essSurj /-- To see that a functor is an equivalence, it suffices to provide an inverse functor `G` such that `F ⋙ G` and `G ⋙ F` are naturally isomorphic to identity functors. -/ -protected def mk {F : C ⥤ D} (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : IsEquivalence F := - ⟨G, adjointifyη η ε, ε, adjointify_η_ε η ε⟩ +protected lemma mk' {F : C ⥤ D} (G : D ⥤ C) (η : 𝟭 C ≅ F ⋙ G) (ε : G ⋙ F ≅ 𝟭 D) : + IsEquivalence F := + inferInstanceAs (IsEquivalence (Equivalence.mk F G η ε).functor) #align category_theory.is_equivalence.mk CategoryTheory.Functor.IsEquivalence.mk end IsEquivalence -/-- Interpret a functor that is an equivalence as an equivalence. -/ -def asEquivalence (F : C ⥤ D) [IsEquivalence F] : C ≌ D := - ⟨F, IsEquivalence.inverse F, IsEquivalence.unitIso, IsEquivalence.counitIso, - IsEquivalence.functor_unitIso_comp⟩ -#align category_theory.functor.as_equivalence CategoryTheory.Functor.asEquivalence - -instance isEquivalenceRefl : IsEquivalence (𝟭 C) := - IsEquivalence.ofEquivalence Equivalence.refl -#align category_theory.functor.is_equivalence_refl CategoryTheory.Functor.isEquivalenceRefl - -/-- The inverse functor of a functor that is an equivalence. -/ -def inv (F : C ⥤ D) [IsEquivalence F] : D ⥤ C := - IsEquivalence.inverse F +/-- A quasi-inverse `D ⥤ C` to a functor that `F : C ⥤ D` that is an equivalence, +i.e. faithful, full, and essentially surjective. -/ +noncomputable def inv (F : C ⥤ D) [F.IsEquivalence] : D ⥤ C + where + obj X := F.objPreimage X + map {X Y} f := F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv) + map_id X := by apply F.map_injective; aesop_cat + map_comp {X Y Z} f g := by apply F.map_injective; simp #align category_theory.functor.inv CategoryTheory.Functor.inv -instance isEquivalenceInv (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.inv := - IsEquivalence.ofEquivalence F.asEquivalence.symm -#align category_theory.functor.is_equivalence_inv CategoryTheory.Functor.isEquivalenceInv - -@[simp] -theorem asEquivalence_functor (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.functor = F := - rfl +/-- Interpret a functor that is an equivalence as an equivalence. + +See . -/ +@[simps functor] +noncomputable def asEquivalence (F : C ⥤ D) [F.IsEquivalence] : C ≌ D where + functor := F + inverse := F.inv + unitIso := NatIso.ofComponents + (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) + (fun f => F.map_injective (by simp [inv])) + counitIso := NatIso.ofComponents F.objObjPreimageIso (by simp [inv]) +#align category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Functor.asEquivalence #align category_theory.functor.as_equivalence_functor CategoryTheory.Functor.asEquivalence_functor -@[simp] -theorem asEquivalence_inverse (F : C ⥤ D) [IsEquivalence F] : F.asEquivalence.inverse = inv F := - rfl -#align category_theory.functor.as_equivalence_inverse CategoryTheory.Functor.asEquivalence_inverse - -@[simp] -theorem asEquivalence_unit {F : C ⥤ D} [IsEquivalence F] : - F.asEquivalence.unitIso = IsEquivalence.unitIso := - rfl -#align category_theory.functor.as_equivalence_unit CategoryTheory.Functor.asEquivalence_unit - -@[simp] -theorem asEquivalence_counit {F : C ⥤ D} [IsEquivalence F] : - F.asEquivalence.counitIso = IsEquivalence.counitIso := - rfl -#align category_theory.functor.as_equivalence_counit CategoryTheory.Functor.asEquivalence_counit +instance isEquivalence_refl : IsEquivalence (𝟭 C) := + Equivalence.refl.isEquivalence_functor +#align category_theory.functor.is_equivalence_refl CategoryTheory.Functor.isEquivalence_refl -@[simp] -theorem inv_inv (F : C ⥤ D) [IsEquivalence F] : inv (inv F) = F := - rfl -#align category_theory.functor.inv_inv CategoryTheory.Functor.inv_inv +instance isEquivalence_inv (F : C ⥤ D) [IsEquivalence F] : IsEquivalence F.inv := + F.asEquivalence.symm.isEquivalence_functor +#align category_theory.functor.is_equivalence_inv CategoryTheory.Functor.isEquivalence_inv variable {E : Type u₃} [Category.{v₃} E] -instance isEquivalenceTrans (F : C ⥤ D) (G : D ⥤ E) [IsEquivalence F] [IsEquivalence G] : - IsEquivalence (F ⋙ G) := - IsEquivalence.ofEquivalence (Equivalence.trans (asEquivalence F) (asEquivalence G)) -#align category_theory.functor.is_equivalence_trans CategoryTheory.Functor.isEquivalenceTrans +instance isEquivalence_trans (F : C ⥤ D) (G : D ⥤ E) [IsEquivalence F] [IsEquivalence G] : + IsEquivalence (F ⋙ G) where +#align category_theory.functor.is_equivalence_trans CategoryTheory.Functor.isEquivalence_trans instance (F : C ⥤ D) [IsEquivalence F] : IsEquivalence ((whiskeringLeft C D E).obj F) := (inferInstance : IsEquivalence (Equivalence.congrLeft F.asEquivalence).inverse) @@ -574,212 +611,59 @@ instance (F : C ⥤ D) [IsEquivalence F] : IsEquivalence ((whiskeringRight E C D end Functor -namespace Equivalence - -@[simp] -theorem functor_inv (E : C ≌ D) : E.functor.inv = E.inverse := - rfl -#align category_theory.equivalence.functor_inv CategoryTheory.Equivalence.functor_inv - -@[simp] -theorem inverse_inv (E : C ≌ D) : E.inverse.inv = E.functor := - rfl -#align category_theory.equivalence.inverse_inv CategoryTheory.Equivalence.inverse_inv - -@[simp] -theorem isEquivalence_unitIso (E : C ≌ D) : IsEquivalence.unitIso = E.unitIso := - rfl - -@[simp] -theorem isEquivalence_counitIso (E : C ≌ D) : IsEquivalence.counitIso = E.counitIso := - rfl - -@[simp] -theorem functor_asEquivalence (E : C ≌ D) : E.functor.asEquivalence = E := by - cases E - congr -#align category_theory.equivalence.functor_as_equivalence CategoryTheory.Equivalence.functor_asEquivalence - -@[simp] -theorem inverse_asEquivalence (E : C ≌ D) : E.inverse.asEquivalence = E.symm := by - cases E - congr -#align category_theory.equivalence.inverse_as_equivalence CategoryTheory.Equivalence.inverse_asEquivalence - -end Equivalence +namespace Functor -namespace Functor.IsEquivalence @[simp] theorem fun_inv_map (F : C ⥤ D) [IsEquivalence F] (X Y : D) (f : X ⟶ Y) : F.map (F.inv.map f) = F.asEquivalence.counit.app X ≫ f ≫ F.asEquivalence.counitInv.app Y := by erw [NatIso.naturality_2] rfl -#align category_theory.is_equivalence.fun_inv_map CategoryTheory.Functor.IsEquivalence.fun_inv_map +#align category_theory.is_equivalence.fun_inv_map CategoryTheory.Functor.fun_inv_map @[simp] theorem inv_fun_map (F : C ⥤ D) [IsEquivalence F] (X Y : C) (f : X ⟶ Y) : F.inv.map (F.map f) = F.asEquivalence.unitInv.app X ≫ f ≫ F.asEquivalence.unit.app Y := by erw [NatIso.naturality_1] rfl -#align category_theory.is_equivalence.inv_fun_map CategoryTheory.Functor.IsEquivalence.inv_fun_map +#align category_theory.is_equivalence.inv_fun_map CategoryTheory.Functor.inv_fun_map -/-- When a functor `F` is an equivalence of categories, and `G` is isomorphic to `F`, then -`G` is also an equivalence of categories. -/ -@[simps!] -def ofIso {F G : C ⥤ D} (e : F ≅ G) (hF : IsEquivalence F) : IsEquivalence G - where - inverse := hF.inverse - unitIso := hF.unitIso ≪≫ NatIso.hcomp e (Iso.refl hF.inverse) - counitIso := NatIso.hcomp (Iso.refl hF.inverse) e.symm ≪≫ hF.counitIso - functor_unitIso_comp X := by - dsimp [NatIso.hcomp] - erw [id_comp, F.map_id, comp_id] - apply (cancel_epi (e.hom.app X)).mp - slice_lhs 1 2 => rw [← e.hom.naturality] - slice_lhs 2 3 => rw [← NatTrans.vcomp_app', e.hom_inv_id] - simp only [NatTrans.id_app, id_comp, comp_id, F.map_comp, assoc] - erw [hF.counitIso.hom.naturality] - slice_lhs 1 2 => rw [functor_unitIso_comp] - simp only [Functor.id_map, id_comp] -#align category_theory.is_equivalence.of_iso CategoryTheory.Functor.IsEquivalence.ofIso - -/-- Compatibility of `ofIso` with the composition of isomorphisms of functors -/ -theorem ofIso_trans {F G H : C ⥤ D} (e : F ≅ G) (e' : G ≅ H) (hF : IsEquivalence F) : - ofIso e' (ofIso e hF) = ofIso (e ≪≫ e') hF := by - dsimp [ofIso] - congr 1 <;> ext X <;> dsimp [NatIso.hcomp] - · simp only [id_comp, assoc, Functor.map_comp] - · simp only [Functor.map_id, comp_id, id_comp, assoc] -#align category_theory.is_equivalence.of_iso_trans CategoryTheory.Functor.IsEquivalence.ofIso_trans - -/-- Compatibility of `ofIso` with identity isomorphisms of functors -/ -theorem ofIso_refl (F : C ⥤ D) (hF : IsEquivalence F) : ofIso (Iso.refl F) hF = hF := by - rcases hF with ⟨Finv, Funit, Fcounit, Fcomp⟩ - dsimp [ofIso] - congr 1 <;> ext X <;> dsimp [NatIso.hcomp] - · simp only [comp_id, map_id] - · simp only [id_comp, map_id] -#align category_theory.is_equivalence.of_iso_refl CategoryTheory.Functor.IsEquivalence.ofIso_refl - -/-- When `F` and `G` are two isomorphic functors, then `F` is an equivalence iff `G` is. -/ -@[simps] -def equivOfIso {F G : C ⥤ D} (e : F ≅ G) : IsEquivalence F ≃ IsEquivalence G - where - toFun := ofIso e - invFun := ofIso e.symm - left_inv hF := by rw [ofIso_trans, Iso.self_symm_id, ofIso_refl] - right_inv hF := by rw [ofIso_trans, Iso.symm_self_id, ofIso_refl] -#align category_theory.is_equivalence.equiv_of_iso CategoryTheory.Functor.IsEquivalence.equivOfIso +lemma isEquivalence_of_iso {F G : C ⥤ D} (e : F ≅ G) [F.IsEquivalence] : G.IsEquivalence := + ((asEquivalence F).changeFunctor e).isEquivalence_functor + +lemma isEquivalence_iff_of_iso {F G : C ⥤ D} (e : F ≅ G) : + F.IsEquivalence ↔ G.IsEquivalence := + ⟨fun _ => isEquivalence_of_iso e, fun _ => isEquivalence_of_iso e.symm⟩ /-- If `G` and `F ⋙ G` are equivalence of categories, then `F` is also an equivalence. -/ -@[simp] -def cancelCompRight {E : Type*} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hG : IsEquivalence G) - (_ : IsEquivalence (F ⋙ G)) : IsEquivalence F := - ofIso (Functor.associator F G G.inv ≪≫ NatIso.hcomp (Iso.refl F) hG.unitIso.symm ≪≫ rightUnitor F) - (Functor.isEquivalenceTrans (F ⋙ G) G.inv) -#align category_theory.is_equivalence.cancel_comp_right CategoryTheory.Functor.IsEquivalence.cancelCompRight +lemma isEquivalence_of_comp_right {E : Type*} [Category E] (F : C ⥤ D) (G : D ⥤ E) + [IsEquivalence G] [IsEquivalence (F ⋙ G)] : IsEquivalence F := by + rw [isEquivalence_iff_of_iso (F.rightUnitor.symm ≪≫ isoWhiskerLeft F (G.asEquivalence.unitIso))] + exact ((F ⋙ G).asEquivalence.trans G.asEquivalence.symm).isEquivalence_functor +#align category_theory.is_equivalence.cancel_comp_right CategoryTheory.Functor.isEquivalence_of_comp_right /-- If `F` and `F ⋙ G` are equivalence of categories, then `G` is also an equivalence. -/ -@[simp] -def cancelCompLeft {E : Type*} [Category E] (F : C ⥤ D) (G : D ⥤ E) (hF : IsEquivalence F) - (_ : IsEquivalence (F ⋙ G)) : IsEquivalence G := - ofIso - ((Functor.associator F.inv F G).symm ≪≫ NatIso.hcomp hF.counitIso (Iso.refl G) ≪≫ leftUnitor G) - (Functor.isEquivalenceTrans F.inv (F ⋙ G)) -#align category_theory.is_equivalence.cancel_comp_left CategoryTheory.Functor.IsEquivalence.cancelCompLeft - -end IsEquivalence +lemma isEquivalence_of_comp_left {E : Type*} [Category E] (F : C ⥤ D) (G : D ⥤ E) + [IsEquivalence F] [IsEquivalence (F ⋙ G)] : IsEquivalence G := by + rw [isEquivalence_iff_of_iso (G.leftUnitor.symm ≪≫ + isoWhiskerRight F.asEquivalence.counitIso.symm G)] + exact (F.asEquivalence.symm.trans (F ⋙ G).asEquivalence).isEquivalence_functor +#align category_theory.is_equivalence.cancel_comp_left CategoryTheory.Functor.isEquivalence_of_comp_left end Functor namespace Equivalence -/-- An equivalence is essentially surjective. - -See . --/ -theorem essSurj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : F.EssSurj := - ⟨fun Y => ⟨F.inv.obj Y, ⟨F.asEquivalence.counitIso.app Y⟩⟩⟩ -#align category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.essSurj_of_equivalence - --- see Note [lower instance priority] -/-- An equivalence is faithful. - -See . --/ -instance (priority := 100) faithfulOfEquivalence (F : C ⥤ D) [IsEquivalence F] : F.Faithful where - map_injective := @fun X Y f g h => by - have p : F.inv.map (F.map f) = F.inv.map (F.map g) := congrArg F.inv.map h - simpa only [cancel_epi, cancel_mono, IsEquivalence.inv_fun_map] using p -#align category_theory.equivalence.faithful_of_equivalence CategoryTheory.Equivalence.faithfulOfEquivalence - --- see Note [lower instance priority] -/-- An equivalence is full. - -See . --/ -instance (priority := 100) full_of_equivalence (F : C ⥤ D) [IsEquivalence F] : F.Full where - map_surjective {X Y} f := - ⟨F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y, - F.inv.map_injective <| by - simpa only [IsEquivalence.inv_fun_map, assoc, Iso.inv_hom_id_app_assoc, - Iso.inv_hom_id_app] using comp_id _⟩ -#align category_theory.equivalence.full_of_equivalence CategoryTheory.Equivalence.full_of_equivalence - -@[simps] -private noncomputable def equivalenceInverse (F : C ⥤ D) [F.Full] [F.Faithful] [F.EssSurj] : D ⥤ C - where - obj X := F.objPreimage X - map {X Y} f := F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv) - map_id X := by apply F.map_injective; aesop_cat - map_comp {X Y Z} f g := by apply F.map_injective; simp --- #align category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse -/- Porting note: this is a private def in mathlib -/ - --- -/-- A functor which is full, faithful, and essentially surjective is an equivalence. - -See . --/ -noncomputable def _root_.CategoryTheory.Functor.IsEquivalence.ofFullyFaithfullyEssSurj - (F : C ⥤ D) [F.Full] [F.Faithful] [F.EssSurj] : - IsEquivalence F := - IsEquivalence.mk (equivalenceInverse F) - (NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm) - fun f => by - apply F.map_injective - aesop_cat) - (NatIso.ofComponents F.objObjPreimageIso) -#align category_theory.equivalence.of_fully_faithfully_ess_surj CategoryTheory.Functor.IsEquivalence.ofFullyFaithfullyEssSurj - --- deprecated on 2024-04-06 -@[deprecated] alias ofFullyFaithfullyEssSurj := Functor.IsEquivalence.ofFullyFaithfullyEssSurj - -@[simp] -theorem functor_map_inj_iff (e : C ≌ D) {X Y : C} (f g : X ⟶ Y) : - e.functor.map f = e.functor.map g ↔ f = g := - ⟨fun h => e.functor.map_injective h, fun h => h ▸ rfl⟩ -#align category_theory.equivalence.functor_map_inj_iff CategoryTheory.Equivalence.functor_map_inj_iff - -@[simp] -theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) : - e.inverse.map f = e.inverse.map g ↔ f = g := - functor_map_inj_iff e.symm f g -#align category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff - instance essSurjInducedFunctor {C' : Type*} (e : C' ≃ D) : (inducedFunctor e).EssSurj where mem_essImage Y := ⟨e.symm Y, by simpa using ⟨default⟩⟩ #align category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.essSurjInducedFunctor noncomputable instance inducedFunctorOfEquiv {C' : Type*} (e : C' ≃ D) : - IsEquivalence (inducedFunctor e) := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ + IsEquivalence (inducedFunctor e) where #align category_theory.equivalence.induced_functor_of_equiv CategoryTheory.Equivalence.inducedFunctorOfEquiv noncomputable instance fullyFaithfulToEssImage (F : C ⥤ D) [F.Full] [F.Faithful] : - IsEquivalence F.toEssImage := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj F.toEssImage + IsEquivalence F.toEssImage where #align category_theory.equivalence.fully_faithful_to_ess_image CategoryTheory.Equivalence.fullyFaithfulToEssImage end Equivalence @@ -788,69 +672,41 @@ namespace Iso variable {E : Type u₃} [Category.{v₃} E] {F : C ⥤ E} {G : C ⥤ D} {H : D ⥤ E} -/-- Construct an isomorphism `F ⋙ H.inv ≅ G` from an isomorphism `F ≅ G ⋙ H`. -/ -@[simps!] -def compInvIso [h : IsEquivalence H] (i : F ≅ G ⋙ H) : F ⋙ H.inv ≅ G := - isoWhiskerRight i H.inv ≪≫ - associator G H H.inv ≪≫ isoWhiskerLeft G h.unitIso.symm ≪≫ G.rightUnitor -#align category_theory.comp_inv_iso CategoryTheory.Iso.compInvIso - /-- Construct an isomorphism `F ⋙ H.inverse ≅ G` from an isomorphism `F ≅ G ⋙ H.functor`. -/ @[simps!] def compInverseIso {H : D ≌ E} (i : F ≅ G ⋙ H.functor) : F ⋙ H.inverse ≅ G := - i.compInvIso - -/-- Construct an isomorphism `G ≅ F ⋙ H.inv` from an isomorphism `G ⋙ H ≅ F`. -/ -@[simps!] -def isoCompInv [IsEquivalence H] (i : G ⋙ H ≅ F) : G ≅ F ⋙ H.inv := - (compInvIso i.symm).symm -#align category_theory.iso_comp_inv CategoryTheory.Iso.isoCompInv + isoWhiskerRight i H.inverse ≪≫ + associator G _ H.inverse ≪≫ isoWhiskerLeft G H.unitIso.symm ≪≫ G.rightUnitor /-- Construct an isomorphism `G ≅ F ⋙ H.inverse` from an isomorphism `G ⋙ H.functor ≅ F`. -/ @[simps!] def isoCompInverse {H : D ≌ E} (i : G ⋙ H.functor ≅ F) : G ≅ F ⋙ H.inverse := - i.isoCompInv - -/-- Construct an isomorphism `G.inv ⋙ F ≅ H` from an isomorphism `F ≅ G ⋙ H`. -/ -@[simps!] -def invCompIso [h : IsEquivalence G] (i : F ≅ G ⋙ H) : G.inv ⋙ F ≅ H := - isoWhiskerLeft G.inv i ≪≫ - (associator G.inv G H).symm ≪≫ isoWhiskerRight h.counitIso H ≪≫ H.leftUnitor -#align category_theory.inv_comp_iso CategoryTheory.Iso.invCompIso + G.rightUnitor.symm ≪≫ isoWhiskerLeft G H.unitIso ≪≫ (associator _ _ _).symm ≪≫ + isoWhiskerRight i H.inverse /-- Construct an isomorphism `G.inverse ⋙ F ≅ H` from an isomorphism `F ≅ G.functor ⋙ H`. -/ @[simps!] def inverseCompIso {G : C ≌ D} (i : F ≅ G.functor ⋙ H) : G.inverse ⋙ F ≅ H := - i.invCompIso - -/-- Construct an isomorphism `H ≅ G.inv ⋙ F` from an isomorphism `G ⋙ H ≅ F`. -/ -@[simps!] -def isoInvComp [IsEquivalence G] (i : G ⋙ H ≅ F) : H ≅ G.inv ⋙ F := - (invCompIso i.symm).symm -#align category_theory.iso_inv_comp CategoryTheory.Iso.isoInvComp + isoWhiskerLeft G.inverse i ≪≫ (associator _ _ _).symm ≪≫ + isoWhiskerRight G.counitIso H ≪≫ H.leftUnitor /-- Construct an isomorphism `H ≅ G.inverse ⋙ F` from an isomorphism `G.functor ⋙ H ≅ F`. -/ @[simps!] def isoInverseComp {G : C ≌ D} (i : G.functor ⋙ H ≅ F) : H ≅ G.inverse ⋙ F := - i.isoInvComp + H.leftUnitor.symm ≪≫ isoWhiskerRight G.counitIso.symm H ≪≫ associator _ _ _ + ≪≫ isoWhiskerLeft G.inverse i end Iso --- deprecated on 2024-02-07 -@[deprecated] alias compInvIso := Iso.compInvIso -@[deprecated] alias isoCompInv := Iso.isoCompInv -@[deprecated] alias invCompIso := Iso.invCompIso -@[deprecated] alias isoInvComp := Iso.isoInvComp - -- deprecated on 2024-04-06 @[deprecated] alias IsEquivalence := Functor.IsEquivalence -@[deprecated] alias IsEquivalence.fun_inv_map := Functor.IsEquivalence.fun_inv_map -@[deprecated] alias IsEquivalence.inv_fun_map := Functor.IsEquivalence.inv_fun_map -@[deprecated] alias IsEquivalence.ofIso := Functor.IsEquivalence.ofIso -@[deprecated] alias IsEquivalence.ofIso_trans := Functor.IsEquivalence.ofIso_trans -@[deprecated] alias IsEquivalence.ofIso_refl := Functor.IsEquivalence.ofIso_refl -@[deprecated] alias IsEquivalence.equivOfIso := Functor.IsEquivalence.equivOfIso -@[deprecated] alias IsEquivalence.cancelCompRight := Functor.IsEquivalence.cancelCompRight -@[deprecated] alias IsEquivalence.cancelCompLeft := Functor.IsEquivalence.cancelCompLeft +@[deprecated] alias IsEquivalence.fun_inv_map := Functor.fun_inv_map +@[deprecated] alias IsEquivalence.inv_fun_map := Functor.inv_fun_map +@[deprecated] alias IsEquivalence.ofIso := Equivalence.changeFunctor +@[deprecated] alias IsEquivalence.ofIso_trans := Equivalence.changeFunctor_trans +@[deprecated] alias IsEquivalence.ofIso_refl := Equivalence.changeFunctor_refl +@[deprecated] alias IsEquivalence.equivOfIso := Functor.isEquivalence_iff_of_iso +@[deprecated] alias IsEquivalence.cancelCompRight := Functor.isEquivalence_of_comp_right +@[deprecated] alias IsEquivalence.cancelCompLeft := Functor.isEquivalence_of_comp_left end CategoryTheory diff --git a/Mathlib/CategoryTheory/EssentialImage.lean b/Mathlib/CategoryTheory/EssentialImage.lean index 7a34db16b6f58..e04cd1e362667 100644 --- a/Mathlib/CategoryTheory/EssentialImage.lean +++ b/Mathlib/CategoryTheory/EssentialImage.lean @@ -23,13 +23,14 @@ a functor decomposes into an essentially surjective functor and a fully faithful -/ -universe v₁ v₂ u₁ u₂ +universe v₁ v₂ v₃ u₁ u₂ u₃ noncomputable section namespace CategoryTheory -variable {C : Type u₁} {D : Type u₂} [Category.{v₁} C] [Category.{v₂} D] {F : C ⥤ D} +variable {C : Type u₁} {D : Type u₂} {E : Type u₃} + [Category.{v₁} C] [Category.{v₂} D] [Category.{v₃} E] {F : C ⥤ D} namespace Functor @@ -173,6 +174,10 @@ theorem essSurj_of_surj (h : Function.Surjective F.obj) : EssSurj F where lemma essSurj_of_iso {F G : C ⥤ D} [EssSurj F] (α : F ≅ G) : EssSurj G where mem_essImage Y := Functor.essImage.ofNatIso α (EssSurj.mem_essImage Y) +instance essSurj_comp (F : C ⥤ D) (G : D ⥤ E) [F.EssSurj] [G.EssSurj] : + (F ⋙ G).EssSurj where + mem_essImage Z := ⟨_, ⟨G.mapIso (F.objObjPreimageIso _) ≪≫ G.objObjPreimageIso Z⟩⟩ + end Functor -- deprecated on 2024-04-06 diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index b40274054d846..3aefe9338d1d3 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -221,8 +221,8 @@ theorem of_right_adjoint {L : D ⥤ C} {R : C ⥤ D} (h : L ⊣ R) : IsFilteredO /-- If `C` is filtered or empty, and we have a right adjoint functor `R : C ⥤ D`, then `D` is filtered or empty. -/ -theorem of_isRightAdjoint (R : C ⥤ D) [IsRightAdjoint R] : IsFilteredOrEmpty D := - of_right_adjoint (Adjunction.ofRightAdjoint R) +theorem of_isRightAdjoint (R : C ⥤ D) [R.IsRightAdjoint] : IsFilteredOrEmpty D := + of_right_adjoint (Adjunction.ofIsRightAdjoint R) /-- Being filtered or empty is preserved by equivalence of categories. -/ theorem of_equivalence (h : C ≌ D) : IsFilteredOrEmpty D := @@ -353,8 +353,8 @@ theorem of_right_adjoint {L : D ⥤ C} {R : C ⥤ D} (h : L ⊣ R) : IsFiltered #align category_theory.is_filtered.of_right_adjoint CategoryTheory.IsFiltered.of_right_adjoint /-- If `C` is filtered, and we have a right adjoint functor `R : C ⥤ D`, then `D` is filtered. -/ -theorem of_isRightAdjoint (R : C ⥤ D) [IsRightAdjoint R] : IsFiltered D := - of_right_adjoint (Adjunction.ofRightAdjoint R) +theorem of_isRightAdjoint (R : C ⥤ D) [R.IsRightAdjoint] : IsFiltered D := + of_right_adjoint (Adjunction.ofIsRightAdjoint R) #align category_theory.is_filtered.of_is_right_adjoint CategoryTheory.IsFiltered.of_isRightAdjoint /-- Being filtered is preserved by equivalence of categories. -/ @@ -705,8 +705,8 @@ theorem of_left_adjoint {L : C ⥤ D} {R : D ⥤ C} (h : L ⊣ R) : IsCofiltered /-- If `C` is cofiltered or empty, and we have a left adjoint functor `L : C ⥤ D`, then `D` is cofiltered or empty. -/ -theorem of_isLeftAdjoint (L : C ⥤ D) [IsLeftAdjoint L] : IsCofilteredOrEmpty D := - of_left_adjoint (Adjunction.ofLeftAdjoint L) +theorem of_isLeftAdjoint (L : C ⥤ D) [L.IsLeftAdjoint] : IsCofilteredOrEmpty D := + of_left_adjoint (Adjunction.ofIsLeftAdjoint L) /-- Being cofiltered or empty is preserved by equivalence of categories. -/ theorem of_equivalence (h : C ≌ D) : IsCofilteredOrEmpty D := @@ -839,8 +839,8 @@ theorem of_left_adjoint {L : C ⥤ D} {R : D ⥤ C} (h : L ⊣ R) : IsCofiltered #align category_theory.is_cofiltered.of_left_adjoint CategoryTheory.IsCofiltered.of_left_adjoint /-- If `C` is cofiltered, and we have a left adjoint functor `L : C ⥤ D`, then `D` is cofiltered. -/ -theorem of_isLeftAdjoint (L : C ⥤ D) [IsLeftAdjoint L] : IsCofiltered D := - of_left_adjoint (Adjunction.ofLeftAdjoint L) +theorem of_isLeftAdjoint (L : C ⥤ D) [L.IsLeftAdjoint] : IsCofiltered D := + of_left_adjoint (Adjunction.ofIsLeftAdjoint L) #align category_theory.is_cofiltered.of_is_left_adjoint CategoryTheory.IsCofiltered.of_isLeftAdjoint /-- Being cofiltered is preserved by equivalence of categories. -/ diff --git a/Mathlib/CategoryTheory/FintypeCat.lean b/Mathlib/CategoryTheory/FintypeCat.lean index 3ef1f9f2b31e5..43ef87ee7bb72 100644 --- a/Mathlib/CategoryTheory/FintypeCat.lean +++ b/Mathlib/CategoryTheory/FintypeCat.lean @@ -199,8 +199,7 @@ instance : incl.EssSurj := { hom := F.symm ∘ ULift.down inv := ULift.up ∘ F }⟩ -noncomputable instance : incl.IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +noncomputable instance : incl.IsEquivalence where /-- The equivalence between `Fintype.Skeleton` and `Fintype`. -/ noncomputable def equivalence : Skeleton ≌ FintypeCat := @@ -218,7 +217,7 @@ set_option linter.uppercaseLean3 false in end Skeleton /-- `Fintype.Skeleton` is a skeleton of `Fintype`. -/ -noncomputable def isSkeleton : IsSkeletonOf FintypeCat Skeleton Skeleton.incl where +lemma isSkeleton : IsSkeletonOf FintypeCat Skeleton Skeleton.incl where skel := Skeleton.is_skeletal eqv := by infer_instance set_option linter.uppercaseLean3 false in diff --git a/Mathlib/CategoryTheory/Functor/EpiMono.lean b/Mathlib/CategoryTheory/Functor/EpiMono.lean index b5d8a1f641179..1921c81c0de8b 100644 --- a/Mathlib/CategoryTheory/Functor/EpiMono.lean +++ b/Mathlib/CategoryTheory/Functor/EpiMono.lean @@ -184,7 +184,7 @@ theorem preservesEpimorphsisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : instance (priority := 100) preservesEpimorphisms_of_isLeftAdjoint (F : C ⥤ D) [IsLeftAdjoint F] : PreservesEpimorphisms F := - preservesEpimorphsisms_of_adjunction (Adjunction.ofLeftAdjoint F) + preservesEpimorphsisms_of_adjunction (Adjunction.ofIsLeftAdjoint F) #align category_theory.functor.preserves_epimorphisms_of_is_left_adjoint CategoryTheory.Functor.preservesEpimorphisms_of_isLeftAdjoint theorem preservesMonomorphisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : F ⊣ G) : @@ -199,7 +199,7 @@ theorem preservesMonomorphisms_of_adjunction {F : C ⥤ D} {G : D ⥤ C} (adj : instance (priority := 100) preservesMonomorphisms_of_isRightAdjoint (F : C ⥤ D) [IsRightAdjoint F] : PreservesMonomorphisms F := - preservesMonomorphisms_of_adjunction (Adjunction.ofRightAdjoint F) + preservesMonomorphisms_of_adjunction (Adjunction.ofIsRightAdjoint F) #align category_theory.functor.preserves_monomorphisms_of_is_right_adjoint CategoryTheory.Functor.preservesMonomorphisms_of_isRightAdjoint instance (priority := 100) reflectsMonomorphisms_of_faithful (F : C ⥤ D) [Faithful F] : diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index aa9a8a111b043..04f66b1b913bb 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -67,9 +67,9 @@ class RepresentablyFlat (F : C ⥤ D) : Prop where attribute [instance] RepresentablyFlat.cofiltered -instance RepresentablyFlat.of_isRightAdjoint (F : C ⥤ D) [IsRightAdjoint F] : +instance RepresentablyFlat.of_isRightAdjoint (F : C ⥤ D) [F.IsRightAdjoint] : RepresentablyFlat F where - cofiltered _ := IsCofiltered.of_isInitial _ (mkInitialOfLeftAdjoint _ (.ofRightAdjoint F) _) + cofiltered _ := IsCofiltered.of_isInitial _ (mkInitialOfLeftAdjoint _ (.ofIsRightAdjoint F) _) theorem RepresentablyFlat.id : RepresentablyFlat (𝟭 C) := inferInstance #align category_theory.representably_flat.id CategoryTheory.RepresentablyFlat.id diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index eb02c5952db9f..b6ea9e25b7a57 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -64,6 +64,10 @@ theorem map_injective (F : C ⥤ D) [Faithful F] : Faithful.map_injective #align category_theory.functor.map_injective CategoryTheory.Functor.map_injective +lemma map_injective_iff (F : C ⥤ D) [Faithful F] {X Y : C} (f g : X ⟶ Y) : + F.map f = F.map g ↔ f = g := + ⟨fun h => F.map_injective h, fun h => by rw [h]⟩ + theorem mapIso_injective (F : C ⥤ D) [Faithful F] : Function.Injective <| (F.mapIso : (X ≅ Y) → (F.obj X ≅ F.obj Y)) := fun _ _ h => Iso.ext (map_injective F (congr_arg Iso.hom h : _)) diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 77e2c2c2a0adb..573a21abf8297 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -272,7 +272,7 @@ variable (L L' : C ⥤ D) (iso₁ : L ≅ L') (F F' : C ⥤ H) (e : D ≌ D') (i when `e` is an equivalence. -/ noncomputable def rightExtensionEquivalenceOfPostcomp₁ : RightExtension (L ⋙ e.functor) F ≌ RightExtension L F := by - have := CostructuredArrow.isEquivalencePre ((whiskeringLeft D D' H).obj e.functor) + have := CostructuredArrow.isEquivalence_pre ((whiskeringLeft D D' H).obj e.functor) ((whiskeringLeft C D H).obj L) F exact Functor.asEquivalence (CostructuredArrow.pre ((whiskeringLeft D D' H).obj e.functor) ((whiskeringLeft C D H).obj L) F) @@ -285,7 +285,7 @@ lemma hasRightExtension_iff_postcomp₁ : when `e` is an equivalence. -/ noncomputable def leftExtensionEquivalenceOfPostcomp₁ : LeftExtension (L ⋙ e.functor) F ≌ LeftExtension L F := by - have := StructuredArrow.isEquivalencePre F ((whiskeringLeft D D' H).obj e.functor) + have := StructuredArrow.isEquivalence_pre F ((whiskeringLeft D D' H).obj e.functor) ((whiskeringLeft C D H).obj L) exact Functor.asEquivalence (StructuredArrow.pre F ((whiskeringLeft D D' H).obj e.functor) ((whiskeringLeft C D H).obj L)) diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean index 49331b20f52b9..2f34a506bcc57 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean @@ -197,9 +197,6 @@ section IsIdempotentComplete variable [IsIdempotentComplete D] -noncomputable instance : (toKaroubi D).IsEquivalence := - toKaroubiIsEquivalence D - /-- The equivalence of categories `(C ⥤ D) ≌ (Karoubi C ⥤ Karoubi D)` when `D` is idempotent complete. -/ @[simp] @@ -220,7 +217,7 @@ when `D` is idempotent complete. -/ @[simps!] noncomputable def functorExtension : (C ⥤ D) ⥤ Karoubi C ⥤ D := functorExtension₂ C D ⋙ - (whiskeringRight (Karoubi C) (Karoubi D) D).obj (toKaroubiIsEquivalence D).inverse + (whiskeringRight (Karoubi C) (Karoubi D) D).obj (toKaroubiEquivalence D).inverse #align category_theory.idempotents.functor_extension CategoryTheory.Idempotents.functorExtension /-- The equivalence `(C ⥤ D) ≌ (Karoubi C ⥤ D)` when `D` is idempotent complete. -/ @@ -237,22 +234,15 @@ noncomputable instance : (functorExtension C D).IsEquivalence := by rw [← karoubiUniversal_functor_eq] infer_instance --- Porting note: added to avoid a timeout in the following definition -lemma isEquivalence_whiskeringLeft_obj_toKaroubi_aux : - ((whiskeringLeft C (Karoubi C) D).obj (toKaroubi C) ⋙ - (whiskeringRight C D (Karoubi D)).obj (toKaroubi D) ⋙ - (whiskeringRight C (Karoubi D) D).obj (Functor.inv (toKaroubi D))) = - (karoubiUniversal C D).inverse := by - rfl - -noncomputable instance : ((whiskeringLeft C (Karoubi C) D).obj (toKaroubi C)).IsEquivalence := - Functor.IsEquivalence.cancelCompRight _ - ((whiskeringRight C _ _).obj (toKaroubi D) ⋙ (whiskeringRight C _ _).obj (toKaroubi D).inv) - (Functor.IsEquivalence.ofEquivalence <| (toKaroubi D).asEquivalence.congrRight.trans - (toKaroubi D).asEquivalence.symm.congrRight) - (by - rw [isEquivalence_whiskeringLeft_obj_toKaroubi_aux] - infer_instance) +instance : ((whiskeringLeft C (Karoubi C) D).obj (toKaroubi C)).IsEquivalence := by + have : ((whiskeringLeft C (Karoubi C) D).obj (toKaroubi C) ⋙ + (whiskeringRight C D (Karoubi D)).obj (toKaroubi D) ⋙ + (whiskeringRight C (Karoubi D) D).obj (Functor.inv (toKaroubi D))).IsEquivalence := by + change (karoubiUniversal C D).inverse.IsEquivalence + infer_instance + exact Functor.isEquivalence_of_comp_right _ + ((whiskeringRight C _ _).obj (toKaroubi D) ⋙ + (whiskeringRight C (Karoubi D) D).obj (Functor.inv (toKaroubi D))) variable {C D} diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index 0c39d62329850..c80c1c1fef942 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -246,14 +246,12 @@ instance [IsIdempotentComplete C] : (toKaroubi C).EssSurj := inv := ⟨e, by erw [comp_id, ← h₂, assoc, h₁, comp_id]⟩ }⟩ /-- If `C` is idempotent complete, the functor `toKaroubi : C ⥤ Karoubi C` is an equivalence. -/ -def toKaroubiIsEquivalence [IsIdempotentComplete C] : (toKaroubi C).IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj (toKaroubi C) -#align category_theory.idempotents.to_karoubi_is_equivalence CategoryTheory.Idempotents.toKaroubiIsEquivalence +instance toKaroubi_isEquivalence [IsIdempotentComplete C] : (toKaroubi C).IsEquivalence where +#align category_theory.idempotents.to_karoubi_is_equivalence CategoryTheory.Idempotents.toKaroubi_isEquivalence /-- The equivalence `C ≅ Karoubi C` when `C` is idempotent complete. -/ def toKaroubiEquivalence [IsIdempotentComplete C] : C ≌ Karoubi C := - haveI := toKaroubiIsEquivalence C - Functor.asEquivalence (toKaroubi C) + (toKaroubi C).asEquivalence #align category_theory.idempotents.to_karoubi_equivalence CategoryTheory.Idempotents.toKaroubiEquivalence instance toKaroubiEquivalence_functor_additive [Preadditive C] [IsIdempotentComplete C] : diff --git a/Mathlib/CategoryTheory/Limits/ConeCategory.lean b/Mathlib/CategoryTheory/Limits/ConeCategory.lean index a6628d9d71225..017b1e0cd245f 100644 --- a/Mathlib/CategoryTheory/Limits/ConeCategory.lean +++ b/Mathlib/CategoryTheory/Limits/ConeCategory.lean @@ -178,15 +178,15 @@ theorem hasLimit_iff_hasTerminal_cone (F : J ⥤ C) : HasLimit F ↔ HasTerminal #align category_theory.limits.has_limit_iff_has_terminal_cone CategoryTheory.Limits.hasLimit_iff_hasTerminal_cone theorem hasLimitsOfShape_iff_isLeftAdjoint_const : - HasLimitsOfShape J C ↔ Nonempty (IsLeftAdjoint (const J : C ⥤ _)) := + HasLimitsOfShape J C ↔ IsLeftAdjoint (const J : C ⥤ _) := calc HasLimitsOfShape J C ↔ ∀ F : J ⥤ C, HasLimit F := ⟨fun h => h.has_limit, fun h => HasLimitsOfShape.mk⟩ _ ↔ ∀ F : J ⥤ C, HasTerminal (Cone F) := forall_congr' hasLimit_iff_hasTerminal_cone _ ↔ ∀ F : J ⥤ C, HasTerminal (CostructuredArrow (const J) F) := (forall_congr' fun F => (Cone.equivCostructuredArrow F).hasTerminal_iff) - _ ↔ Nonempty (IsLeftAdjoint (const J : C ⥤ _)) := - nonempty_isLeftAdjoint_iff_hasTerminal_costructuredArrow.symm + _ ↔ (IsLeftAdjoint (const J : C ⥤ _)) := + isLeftAdjoint_iff_hasTerminal_costructuredArrow.symm #align category_theory.limits.has_limits_of_shape_iff_is_left_adjoint_const CategoryTheory.Limits.hasLimitsOfShape_iff_isLeftAdjoint_const theorem IsLimit.liftConeMorphism_eq_isTerminal_from {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) @@ -348,15 +348,15 @@ theorem hasColimit_iff_hasInitial_cocone (F : J ⥤ C) : HasColimit F ↔ HasIni #align category_theory.limits.has_colimit_iff_has_initial_cocone CategoryTheory.Limits.hasColimit_iff_hasInitial_cocone theorem hasColimitsOfShape_iff_isRightAdjoint_const : - HasColimitsOfShape J C ↔ Nonempty (IsRightAdjoint (const J : C ⥤ _)) := + HasColimitsOfShape J C ↔ IsRightAdjoint (const J : C ⥤ _) := calc HasColimitsOfShape J C ↔ ∀ F : J ⥤ C, HasColimit F := ⟨fun h => h.has_colimit, fun h => HasColimitsOfShape.mk⟩ _ ↔ ∀ F : J ⥤ C, HasInitial (Cocone F) := forall_congr' hasColimit_iff_hasInitial_cocone _ ↔ ∀ F : J ⥤ C, HasInitial (StructuredArrow F (const J)) := (forall_congr' fun F => (Cocone.equivStructuredArrow F).hasInitial_iff) - _ ↔ Nonempty (IsRightAdjoint (const J : C ⥤ _)) := - nonempty_isRightAdjoint_iff_hasInitial_structuredArrow.symm + _ ↔ (IsRightAdjoint (const J : C ⥤ _)) := + isRightAdjoint_iff_hasInitial_structuredArrow.symm #align category_theory.limits.has_colimits_of_shape_iff_is_right_adjoint_const CategoryTheory.Limits.hasColimitsOfShape_iff_isRightAdjoint_const theorem IsColimit.descCoconeMorphism_eq_isInitial_to {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) diff --git a/Mathlib/CategoryTheory/Limits/Cones.lean b/Mathlib/CategoryTheory/Limits/Cones.lean index bfdb50d831f69..755cbecee5563 100644 --- a/Mathlib/CategoryTheory/Limits/Cones.lean +++ b/Mathlib/CategoryTheory/Limits/Cones.lean @@ -747,36 +747,38 @@ def mapCoconeMorphism {c c' : Cocone F} (f : c ⟶ c') : H.mapCocone c ⟶ H.map /-- If `H` is an equivalence, we invert `H.mapCone` and get a cone for `F` from a cone for `F ⋙ H`. -/ -def mapConeInv [IsEquivalence H] (c : Cone (F ⋙ H)) : Cone F := +noncomputable def mapConeInv [IsEquivalence H] (c : Cone (F ⋙ H)) : Cone F := (Limits.Cones.functorialityEquivalence F (asEquivalence H)).inverse.obj c #align category_theory.functor.map_cone_inv CategoryTheory.Functor.mapConeInv /-- `mapCone` is the left inverse to `mapConeInv`. -/ -def mapConeMapConeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone (F ⋙ H)) : +noncomputable def mapConeMapConeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] + (c : Cone (F ⋙ H)) : mapCone H (mapConeInv H c) ≅ c := (Limits.Cones.functorialityEquivalence F (asEquivalence H)).counitIso.app c #align category_theory.functor.map_cone_map_cone_inv CategoryTheory.Functor.mapConeMapConeInv /-- `MapCone` is the right inverse to `mapConeInv`. -/ -def mapConeInvMapCone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone F) : +noncomputable def mapConeInvMapCone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cone F) : mapConeInv H (mapCone H c) ≅ c := (Limits.Cones.functorialityEquivalence F (asEquivalence H)).unitIso.symm.app c #align category_theory.functor.map_cone_inv_map_cone CategoryTheory.Functor.mapConeInvMapCone /-- If `H` is an equivalence, we invert `H.mapCone` and get a cone for `F` from a cone for `F ⋙ H`. -/ -def mapCoconeInv [IsEquivalence H] (c : Cocone (F ⋙ H)) : Cocone F := +noncomputable def mapCoconeInv [IsEquivalence H] (c : Cocone (F ⋙ H)) : Cocone F := (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).inverse.obj c #align category_theory.functor.map_cocone_inv CategoryTheory.Functor.mapCoconeInv /-- `mapCocone` is the left inverse to `mapCoconeInv`. -/ -def mapCoconeMapCoconeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone (F ⋙ H)) : +noncomputable def mapCoconeMapCoconeInv {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] + (c : Cocone (F ⋙ H)) : mapCocone H (mapCoconeInv H c) ≅ c := (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).counitIso.app c #align category_theory.functor.map_cocone_map_cocone_inv CategoryTheory.Functor.mapCoconeMapCoconeInv /-- `mapCocone` is the right inverse to `mapCoconeInv`. -/ -def mapCoconeInvMapCocone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone F) : +noncomputable def mapCoconeInvMapCocone {F : J ⥤ D} (H : D ⥤ C) [IsEquivalence H] (c : Cocone F) : mapCoconeInv H (mapCocone H c) ≅ c := (Limits.Cocones.functorialityEquivalence F (asEquivalence H)).unitIso.symm.app c #align category_theory.functor.map_cocone_inv_map_cocone CategoryTheory.Functor.mapCoconeInvMapCocone diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean b/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean index 0d8ccdae3d806..009ac9f2bc172 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean @@ -130,7 +130,7 @@ theorem has_over_limit_discrete_of_widePullback_limit {B : C} (F : Discrete J [HasLimit (widePullbackDiagramOfDiagramOver B F)] : HasLimit F := HasLimit.mk { cone := _ - isLimit := IsLimit.ofRightAdjoint (conesEquiv B F).functor + isLimit := IsLimit.ofRightAdjoint (conesEquiv B F).symm.toAdjunction (limit.isLimit (widePullbackDiagramOfDiagramOver B F)) } #align category_theory.over.construct_products.has_over_limit_discrete_of_wide_pullback_limit CategoryTheory.Over.ConstructProducts.has_over_limit_discrete_of_widePullback_limit diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index c160f2d85c965..3d2b764d8bb65 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -152,12 +152,12 @@ theorem initial_of_adjunction {L : C ⥤ D} {R : D ⥤ C} (adj : L ⊣ R) : Init Or.inr ⟨CostructuredArrow.homMk (adj.homEquiv g.left d g.hom) (by simp [u])⟩)) } #align category_theory.functor.initial_of_adjunction CategoryTheory.Functor.initial_of_adjunction -instance (priority := 100) final_of_isRightAdjoint (F : C ⥤ D) [h : IsRightAdjoint F] : Final F := - final_of_adjunction h.adj +instance (priority := 100) final_of_isRightAdjoint (F : C ⥤ D) [IsRightAdjoint F] : Final F := + final_of_adjunction (Adjunction.ofIsRightAdjoint F) #align category_theory.functor.final_of_is_right_adjoint CategoryTheory.Functor.final_of_isRightAdjoint -instance (priority := 100) initial_of_isLeftAdjoint (F : C ⥤ D) [h : IsLeftAdjoint F] : Initial F := - initial_of_adjunction h.adj +instance (priority := 100) initial_of_isLeftAdjoint (F : C ⥤ D) [IsLeftAdjoint F] : Initial F := + initial_of_adjunction (Adjunction.ofIsLeftAdjoint F) #align category_theory.functor.initial_of_is_left_adjoint CategoryTheory.Functor.initial_of_isLeftAdjoint theorem final_of_natIso {F F' : C ⥤ D} [Final F] (i : F ≅ F') : Final F' where @@ -679,51 +679,39 @@ variable {E : Type u₃} [Category.{v₃} E] (F : C ⥤ D) (G : D ⥤ E) /-- The hypotheses also imply that `G` is final, see `final_of_comp_full_faithful'`. -/ theorem final_of_comp_full_faithful [Full G] [Faithful G] [Final (F ⋙ G)] : Final F where - out d := - have := StructuredArrow.isEquivalencePost d F G - isConnected_of_equivalent (StructuredArrow.post d F G).asEquivalence.symm + out d := isConnected_of_equivalent (StructuredArrow.post d F G).asEquivalence.symm /-- The hypotheses also imply that `G` is initial, see `initial_of_comp_full_faithful'`. -/ theorem initial_of_comp_full_faithful [Full G] [Faithful G] [Initial (F ⋙ G)] : Initial F where - out d := - have := CostructuredArrow.isEquivalencePost d F G - isConnected_of_equivalent (CostructuredArrow.post F G d).asEquivalence.symm + out d := isConnected_of_equivalent (CostructuredArrow.post F G d).asEquivalence.symm /-- See also the strictly more general `final_comp` below. -/ theorem final_comp_equivalence [Final F] [IsEquivalence G] : Final (F ⋙ G) := - let i : F ≅ (F ⋙ G) ⋙ G.inv := isoWhiskerLeft F IsEquivalence.unitIso + let i : F ≅ (F ⋙ G) ⋙ G.inv := isoWhiskerLeft F G.asEquivalence.unitIso have : Final ((F ⋙ G) ⋙ G.inv) := final_of_natIso i final_of_comp_full_faithful (F ⋙ G) G.inv /-- See also the strictly more general `initial_comp` below. -/ theorem initial_comp_equivalence [Initial F] [IsEquivalence G] : Initial (F ⋙ G) := - let i : F ≅ (F ⋙ G) ⋙ G.inv := isoWhiskerLeft F IsEquivalence.unitIso + let i : F ≅ (F ⋙ G) ⋙ G.inv := isoWhiskerLeft F G.asEquivalence.unitIso have : Initial ((F ⋙ G) ⋙ G.inv) := initial_of_natIso i initial_of_comp_full_faithful (F ⋙ G) G.inv /-- See also the strictly more general `final_comp` below. -/ theorem final_equivalence_comp [IsEquivalence F] [Final G] : Final (F ⋙ G) where - out d := - have := StructuredArrow.isEquivalencePre d F G - isConnected_of_equivalent (StructuredArrow.pre d F G).asEquivalence.symm + out d := isConnected_of_equivalent (StructuredArrow.pre d F G).asEquivalence.symm /-- See also the strictly more general `inital_comp` below. -/ theorem initial_equivalence_comp [IsEquivalence F] [Initial G] : Initial (F ⋙ G) where - out d := - have := CostructuredArrow.isEquivalencePre F G d - isConnected_of_equivalent (CostructuredArrow.pre F G d).asEquivalence.symm + out d := isConnected_of_equivalent (CostructuredArrow.pre F G d).asEquivalence.symm /-- See also the strictly more general `final_of_final_comp` below. -/ theorem final_of_equivalence_comp [IsEquivalence F] [Final (F ⋙ G)] : Final G where - out d := - have := StructuredArrow.isEquivalencePre d F G - isConnected_of_equivalent (StructuredArrow.pre d F G).asEquivalence + out d := isConnected_of_equivalent (StructuredArrow.pre d F G).asEquivalence /-- See also the strictly more general `initial_of_initial_comp` below. -/ theorem initial_of_equivalence_comp [IsEquivalence F] [Initial (F ⋙ G)] : Initial G where - out d := - have := CostructuredArrow.isEquivalencePre F G d - isConnected_of_equivalent (CostructuredArrow.pre F G d).asEquivalence + out d := isConnected_of_equivalent (CostructuredArrow.pre F G d).asEquivalence /-- See also the strictly more general `final_iff_comp_final_full_faithful` below. -/ theorem final_iff_comp_equivalence [IsEquivalence G] : Final F ↔ Final (F ⋙ G) := diff --git a/Mathlib/CategoryTheory/Limits/HasLimits.lean b/Mathlib/CategoryTheory/Limits/HasLimits.lean index d7fd397cb1334..5345bb89793b4 100644 --- a/Mathlib/CategoryTheory/Limits/HasLimits.lean +++ b/Mathlib/CategoryTheory/Limits/HasLimits.lean @@ -580,7 +580,7 @@ def constLimAdj : (const J : C ⥤ J ⥤ C) ⊣ lim where #align category_theory.limits.const_lim_adj CategoryTheory.Limits.constLimAdj instance : IsRightAdjoint (lim : (J ⥤ C) ⥤ C) := - ⟨_, constLimAdj⟩ + ⟨_, ⟨constLimAdj⟩⟩ end LimFunctor @@ -1184,7 +1184,7 @@ def colimConstAdj : (colim : (J ⥤ C) ⥤ C) ⊣ const J where #align category_theory.limits.colim_const_adj CategoryTheory.Limits.colimConstAdj instance : IsLeftAdjoint (colim : (J ⥤ C) ⥤ C) := - ⟨_, colimConstAdj⟩ + ⟨_, ⟨colimConstAdj⟩⟩ end ColimFunctor diff --git a/Mathlib/CategoryTheory/Limits/IsLimit.lean b/Mathlib/CategoryTheory/Limits/IsLimit.lean index 94e8a60c203fc..7093bfce63e1a 100644 --- a/Mathlib/CategoryTheory/Limits/IsLimit.lean +++ b/Mathlib/CategoryTheory/Limits/IsLimit.lean @@ -236,9 +236,10 @@ theorem hom_ext (h : IsLimit t) {W : C} {f f' : W ⟶ t.pt} /-- Given a right adjoint functor between categories of cones, the image of a limit cone is a limit cone. -/ -def ofRightAdjoint {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} (h : Cone G ⥤ Cone F) - [IsRightAdjoint h] {c : Cone G} (t : IsLimit c) : IsLimit (h.obj c) := - mkConeMorphism (fun s => (Adjunction.ofRightAdjoint h).homEquiv s c (t.liftConeMorphism _)) +def ofRightAdjoint {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} {left : Cone F ⥤ Cone G} + {right : Cone G ⥤ Cone F} + (adj : left ⊣ right) {c : Cone G} (t : IsLimit c) : IsLimit (right.obj c) := + mkConeMorphism (fun s => adj.homEquiv s c (t.liftConeMorphism _)) fun _ _ => (Adjunction.eq_homEquiv_apply _ _ _).2 t.uniq_cone_morphism #align category_theory.limits.is_limit.of_right_adjoint CategoryTheory.Limits.IsLimit.ofRightAdjoint @@ -248,8 +249,8 @@ across the equivalence. def ofConeEquiv {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} (h : Cone G ≌ Cone F) {c : Cone G} : IsLimit (h.functor.obj c) ≃ IsLimit c where - toFun P := ofIsoLimit (ofRightAdjoint h.inverse P) (h.unitIso.symm.app c) - invFun := ofRightAdjoint h.functor + toFun P := ofIsoLimit (ofRightAdjoint h.toAdjunction P) (h.unitIso.symm.app c) + invFun := ofRightAdjoint h.symm.toAdjunction left_inv := by aesop_cat right_inv := by aesop_cat #align category_theory.limits.is_limit.of_cone_equiv CategoryTheory.Limits.IsLimit.ofConeEquiv @@ -258,7 +259,7 @@ def ofConeEquiv {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} (h : Cone G theorem ofConeEquiv_apply_desc {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} (h : Cone G ≌ Cone F) {c : Cone G} (P : IsLimit (h.functor.obj c)) (s) : (ofConeEquiv h P).lift s = - ((h.unitIso.hom.app s).hom ≫ (h.functor.inv.map (P.liftConeMorphism (h.functor.obj s))).hom) ≫ + ((h.unitIso.hom.app s).hom ≫ (h.inverse.map (P.liftConeMorphism (h.functor.obj s))).hom) ≫ (h.unitIso.inv.app c).hom := rfl #align category_theory.limits.is_limit.of_cone_equiv_apply_desc CategoryTheory.Limits.IsLimit.ofConeEquiv_apply_desc @@ -340,14 +341,14 @@ open CategoryTheory.Equivalence /-- If `s : Cone F` is a limit cone, so is `s` whiskered by an equivalence `e`. -/ def whiskerEquivalence {s : Cone F} (P : IsLimit s) (e : K ≌ J) : IsLimit (s.whisker e.functor) := - ofRightAdjoint (Cones.whiskeringEquivalence e).functor P + ofRightAdjoint (Cones.whiskeringEquivalence e).symm.toAdjunction P #align category_theory.limits.is_limit.whisker_equivalence CategoryTheory.Limits.IsLimit.whiskerEquivalence /-- If `s : Cone F` whiskered by an equivalence `e` is a limit cone, so is `s`. -/ def ofWhiskerEquivalence {s : Cone F} (e : K ≌ J) (P : IsLimit (s.whisker e.functor)) : IsLimit s := equivIsoLimit ((Cones.whiskeringEquivalence e).unitIso.app s).symm - (ofRightAdjoint (Cones.whiskeringEquivalence e).inverse P : _) + (ofRightAdjoint (Cones.whiskeringEquivalence e).toAdjunction P) #align category_theory.limits.is_limit.of_whisker_equivalence CategoryTheory.Limits.IsLimit.ofWhiskerEquivalence /-- Given an equivalence of diagrams `e`, `s` is a limit cone iff `s.whisker e.functor` is. @@ -761,10 +762,11 @@ theorem hom_ext (h : IsColimit t) {W : C} {f f' : t.pt ⟶ W} /-- Given a left adjoint functor between categories of cocones, the image of a colimit cocone is a colimit cocone. -/ -def ofLeftAdjoint {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} (h : Cocone G ⥤ Cocone F) - [IsLeftAdjoint h] {c : Cocone G} (t : IsColimit c) : IsColimit (h.obj c) := +def ofLeftAdjoint {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} {left : Cocone G ⥤ Cocone F} + {right : Cocone F ⥤ Cocone G} (adj : left ⊣ right) {c : Cocone G} (t : IsColimit c) : + IsColimit (left.obj c) := mkCoconeMorphism - (fun s => ((Adjunction.ofLeftAdjoint h).homEquiv c s).symm (t.descCoconeMorphism _)) fun _ _ => + (fun s => (adj.homEquiv c s).symm (t.descCoconeMorphism _)) fun _ _ => (Adjunction.homEquiv_apply_eq _ _ _).1 t.uniq_cocone_morphism #align category_theory.limits.is_colimit.of_left_adjoint CategoryTheory.Limits.IsColimit.ofLeftAdjoint @@ -774,8 +776,8 @@ we can transport a colimiting cocone across the equivalence. def ofCoconeEquiv {D : Type u₄} [Category.{v₄} D] {G : K ⥤ D} (h : Cocone G ≌ Cocone F) {c : Cocone G} : IsColimit (h.functor.obj c) ≃ IsColimit c where - toFun P := ofIsoColimit (ofLeftAdjoint h.inverse P) (h.unitIso.symm.app c) - invFun := ofLeftAdjoint h.functor + toFun P := ofIsoColimit (ofLeftAdjoint h.symm.toAdjunction P) (h.unitIso.symm.app c) + invFun := ofLeftAdjoint h.toAdjunction left_inv := by aesop_cat right_inv := by aesop_cat #align category_theory.limits.is_colimit.of_cocone_equiv CategoryTheory.Limits.IsColimit.ofCoconeEquiv @@ -868,7 +870,7 @@ open CategoryTheory.Equivalence -/ def whiskerEquivalence {s : Cocone F} (P : IsColimit s) (e : K ≌ J) : IsColimit (s.whisker e.functor) := - ofLeftAdjoint (Cocones.whiskeringEquivalence e).functor P + ofLeftAdjoint (Cocones.whiskeringEquivalence e).toAdjunction P #align category_theory.limits.is_colimit.whisker_equivalence CategoryTheory.Limits.IsColimit.whiskerEquivalence /-- If `s : Cocone F` whiskered by an equivalence `e` is a colimit cocone, so is `s`. @@ -876,7 +878,7 @@ def whiskerEquivalence {s : Cocone F} (P : IsColimit s) (e : K ≌ J) : def ofWhiskerEquivalence {s : Cocone F} (e : K ≌ J) (P : IsColimit (s.whisker e.functor)) : IsColimit s := equivIsoColimit ((Cocones.whiskeringEquivalence e).unitIso.app s).symm - (ofLeftAdjoint (Cocones.whiskeringEquivalence e).inverse P : _) + (ofLeftAdjoint (Cocones.whiskeringEquivalence e).symm.toAdjunction P) #align category_theory.limits.is_colimit.of_whisker_equivalence CategoryTheory.Limits.IsColimit.ofWhiskerEquivalence /-- Given an equivalence of diagrams `e`, `s` is a colimit cocone iff `s.whisker e.functor` is. diff --git a/Mathlib/CategoryTheory/Limits/Over.lean b/Mathlib/CategoryTheory/Limits/Over.lean index aa98179e055cb..7711bad009ad4 100644 --- a/Mathlib/CategoryTheory/Limits/Over.lean +++ b/Mathlib/CategoryTheory/Limits/Over.lean @@ -142,8 +142,8 @@ def pullbackComp {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) : pullback (f ≫ g) (((mapPullbackAdj _).comp (mapPullbackAdj _)).ofNatIsoLeft (Over.mapComp _ _).symm) #align category_theory.over.pullback_comp CategoryTheory.Over.pullbackComp -instance pullbackIsRightAdjoint {A B : C} (f : A ⟶ B) : IsRightAdjoint (pullback f) := - ⟨_, mapPullbackAdj f⟩ +instance pullbackIsRightAdjoint {A B : C} (f : A ⟶ B) : (pullback f).IsRightAdjoint := + ⟨_, ⟨mapPullbackAdj f⟩⟩ #align category_theory.over.pullback_is_right_adjoint CategoryTheory.Over.pullbackIsRightAdjoint end diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean index 3e94415964244..631d1a633da5e 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean @@ -74,7 +74,7 @@ theorem map_eq_zero_iff (F : C ⥤ D) [PreservesZeroMorphisms F] [Faithful F] {X instance (priority := 100) preservesZeroMorphisms_of_isLeftAdjoint (F : C ⥤ D) [IsLeftAdjoint F] : PreservesZeroMorphisms F where map_zero X Y := by - let adj := Adjunction.ofLeftAdjoint F + let adj := Adjunction.ofIsLeftAdjoint F calc F.map (0 : X ⟶ Y) = F.map 0 ≫ F.map (adj.unit.app Y) ≫ adj.counit.app (F.obj Y) := ?_ _ = F.map 0 ≫ F.map ((rightAdjoint F).map (0 : F.obj X ⟶ _)) ≫ adj.counit.app (F.obj Y) := ?_ @@ -88,7 +88,7 @@ instance (priority := 100) preservesZeroMorphisms_of_isLeftAdjoint (F : C ⥤ D) instance (priority := 100) preservesZeroMorphisms_of_isRightAdjoint (G : C ⥤ D) [IsRightAdjoint G] : PreservesZeroMorphisms G where map_zero X Y := by - let adj := Adjunction.ofRightAdjoint G + let adj := Adjunction.ofIsRightAdjoint G calc G.map (0 : X ⟶ Y) = adj.unit.app (G.obj X) ≫ G.map (adj.counit.app X) ≫ G.map 0 := ?_ _ = adj.unit.app (G.obj X) ≫ G.map ((leftAdjoint G).map (0 : _ ⟶ G.obj X)) ≫ G.map 0 := ?_ diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index 4b5e8e550d97e..e058a3caa2e58 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -411,23 +411,20 @@ noncomputable def uniqueExtensionAlongYoneda (L : (Cᵒᵖ ⥤ Type u₁) ⥤ natIsoOfNatIsoOnRepresentables _ _ (hL ≪≫ (isExtensionAlongYoneda _).symm) #align category_theory.unique_extension_along_yoneda CategoryTheory.uniqueExtensionAlongYoneda -/-- If `L` preserves colimits and `ℰ` has them, then it is a left adjoint. This is a special case of -`isLeftAdjointOfPreservesColimits` used to prove that. --/ -noncomputable def isLeftAdjointOfPreservesColimitsAux (L : (Cᵒᵖ ⥤ Type u₁) ⥤ ℰ) - [PreservesColimits L] : IsLeftAdjoint L where - right := restrictedYoneda (yoneda ⋙ L) - adj := (yonedaAdjunction _).ofNatIsoLeft (uniqueExtensionAlongYoneda _ L (Iso.refl _)).symm -#align category_theory.is_left_adjoint_of_preserves_colimits_aux CategoryTheory.isLeftAdjointOfPreservesColimitsAux +/-- Auxiliary definition for `isLeftAdjointOfPreservesColimits`. -/ +noncomputable def adjunctionOfPreservesColimitsAux (L : (Cᵒᵖ ⥤ Type u₁) ⥤ ℰ) + [PreservesColimits L] : L ⊣ restrictedYoneda (yoneda ⋙ L) := + (yonedaAdjunction _).ofNatIsoLeft (uniqueExtensionAlongYoneda _ L (Iso.refl _)).symm +#align category_theory.is_left_adjoint_of_preserves_colimits_aux CategoryTheory.adjunctionOfPreservesColimitsAux /-- If `L` preserves colimits and `ℰ` has them, then it is a left adjoint. Note this is a (partial) converse to `leftAdjointPreservesColimits`. -/ -noncomputable def isLeftAdjointOfPreservesColimits (L : (C ⥤ Type u₁) ⥤ ℰ) [PreservesColimits L] : - IsLeftAdjoint L := - let e : _ ⥤ Type u₁ ≌ _ ⥤ Type u₁ := (opOpEquivalence C).congrLeft - let _ := isLeftAdjointOfPreservesColimitsAux (e.functor ⋙ L : _) - Adjunction.leftAdjointOfNatIso (e.invFunIdAssoc _) +lemma isLeftAdjointOfPreservesColimits (L : (C ⥤ Type u₁) ⥤ ℰ) [PreservesColimits L] : + L.IsLeftAdjoint := + ⟨_, ⟨((opOpEquivalence C).congrLeft.symm.toAdjunction.comp + (adjunctionOfPreservesColimitsAux ((opOpEquivalence C).congrLeft.functor ⋙ L))).ofNatIsoLeft + ((opOpEquivalence C).congrLeft.invFunIdAssoc L)⟩⟩ #align category_theory.is_left_adjoint_of_preserves_colimits CategoryTheory.isLeftAdjointOfPreservesColimits end SmallCategory diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index b758b784c505b..0d5fd45342b71 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -1051,8 +1051,9 @@ theorem hasStrongEpiMonoFactorisations_imp_of_isEquivalence (F : C ⥤ D) [IsEqu e := F.asEquivalence.counitIso.inv.app X ≫ F.map em.e m := F.map em.m ≫ F.asEquivalence.counitIso.hom.app Y fac := by - simpa only [Category.assoc, ← F.map_comp_assoc, em.fac, IsEquivalence.fun_inv_map, - Iso.inv_hom_id_app, Iso.inv_hom_id_app_assoc] using Category.comp_id _ }⟩ + simp only [asEquivalence_functor, Category.assoc, ← F.map_comp_assoc, + MonoFactorisation.fac, fun_inv_map, id_obj, Iso.inv_hom_id_app, Category.comp_id, + Iso.inv_hom_id_app_assoc] }⟩ #align category_theory.functor.has_strong_epi_mono_factorisations_imp_of_is_equivalence CategoryTheory.Functor.hasStrongEpiMonoFactorisations_imp_of_isEquivalence end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean index 8eb03b1e819cd..538316152a664 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean @@ -54,8 +54,6 @@ attribute [inherit_doc NormalMono] NormalMono.Z NormalMono.g NormalMono.w Normal section -attribute [local instance] Equivalence.essSurj_of_equivalence - /-- If `F` is an equivalence and `F.map f` is a normal mono, then `f` is a normal mono. -/ def equivalenceReflectsNormalMono {D : Type u₂} [Category.{v₁} D] [HasZeroMorphisms D] (F : C ⥤ D) [F.IsEquivalence] {X Y : C} {f : X ⟶ Y} (hf : NormalMono (F.map f)) : NormalMono f where @@ -171,8 +169,6 @@ attribute [inherit_doc NormalEpi] NormalEpi.W NormalEpi.g NormalEpi.w NormalEpi. section -attribute [local instance] Equivalence.essSurj_of_equivalence - /-- If `F` is an equivalence and `F.map f` is a normal epi, then `f` is a normal epi. -/ def equivalenceReflectsNormalEpi {D : Type u₂} [Category.{v₁} D] [HasZeroMorphisms D] (F : C ⥤ D) [F.IsEquivalence] {X Y : C} {f : X ⟶ Y} (hf : NormalEpi (F.map f)) : NormalEpi f where diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index 4457ae959f92a..143536dd6b7e6 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -226,8 +226,7 @@ theorem IsVanKampenColimit.mapCocone_iff (G : C ⥤ D) {F : J ⥤ C} {c : Cocone let e : F ⋙ G ⋙ Functor.inv G ≅ F := NatIso.hcomp (Iso.refl F) G.asEquivalence.unitIso.symm apply IsVanKampenColimit.of_mapCocone G.inv apply (IsVanKampenColimit.precompose_isIso_iff e.inv).mp - refine hc.of_iso (Cocones.ext (G.asEquivalence.unitIso.app c.pt) ?_) - simp [e, Functor.asEquivalence]⟩ + exact hc.of_iso (Cocones.ext (G.asEquivalence.unitIso.app c.pt) (fun j => (by simp [e])))⟩ theorem IsUniversalColimit.whiskerEquivalence {K : Type*} [Category K] (e : J ≌ K) {F : K ⥤ C} {c : Cocone F} (hc : IsUniversalColimit c) : diff --git a/Mathlib/CategoryTheory/Localization/Composition.lean b/Mathlib/CategoryTheory/Localization/Composition.lean index ad904eda0105d..2df736f2914e3 100644 --- a/Mathlib/CategoryTheory/Localization/Composition.lean +++ b/Mathlib/CategoryTheory/Localization/Composition.lean @@ -26,6 +26,7 @@ namespace CategoryTheory variable {C₁ : Type u₁} {C₂ : Type u₂} {C₃ : Type u₃} {E : Type u₄} [Category.{v₁} C₁] [Category.{v₂} C₂] [Category.{v₃} C₃] [Category.{v₄} E] {L₁ : C₁ ⥤ C₂} {L₂ : C₂ ⥤ C₃} {W₁ : MorphismProperty C₁} {W₂ : MorphismProperty C₂} + namespace Localization /-- Under some conditions on the `MorphismProperty`, functors satisfying the strict diff --git a/Mathlib/CategoryTheory/Localization/Equivalence.lean b/Mathlib/CategoryTheory/Localization/Equivalence.lean index 3f08956dc107d..986f7a61c1bd4 100644 --- a/Mathlib/CategoryTheory/Localization/Equivalence.lean +++ b/Mathlib/CategoryTheory/Localization/Equivalence.lean @@ -52,8 +52,8 @@ lemma equivalence_counitIso_app (X : C₂) : rw [comp_id] /-- Basic constructor of an equivalence between localized categories -/ -noncomputable def isEquivalence : G'.IsEquivalence := - Functor.IsEquivalence.ofEquivalence (equivalence L₁ W₁ L₂ W₂ G G' F F' α β) +lemma isEquivalence : G'.IsEquivalence := + (equivalence L₁ W₁ L₂ W₂ G G' F F' α β).isEquivalence_functor end Localization @@ -74,8 +74,8 @@ lemma of_equivalence_source (L₁ : C₁ ⥤ D) (W₁ : MorphismProperty C₁) (W₂.Q.mapArrow.mapIso e)).1 (Localization.inverts W₂.Q W₂ _ hf') exact { inverts := hW₂ - nonempty_isEquivalence := - ⟨Localization.isEquivalence W₂.Q W₂ L₁ W₁ L₂ (Construction.lift L₂ hW₂) + isEquivalence := + Localization.isEquivalence W₂.Q W₂ L₁ W₁ L₂ (Construction.lift L₂ hW₂) (E.functor ⋙ W₂.Q) (Localization.lift (E.functor ⋙ W₂.Q) h L₁) (by calc L₂ ⋙ lift (E.functor ⋙ W₂.Q) h L₁ ≅ _ := (leftUnitor _).symm @@ -90,7 +90,7 @@ lemma of_equivalence_source (L₁ : C₁ ⥤ D) (W₁ : MorphismProperty C₁) _ ≅ (E.inverse ⋙ E.functor) ⋙ W₂.Q := (Functor.associator _ _ _).symm _ ≅ 𝟭 C₂ ⋙ W₂.Q := isoWhiskerRight E.counitIso _ _ ≅ W₂.Q := leftUnitor _) - (Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ (Lifting.iso W₂.Q W₂ _ _) ≪≫ iso)⟩ } + (Functor.associator _ _ _ ≪≫ isoWhiskerLeft _ (Lifting.iso W₂.Q W₂ _ _) ≪≫ iso) } /-- If `L₁ : C₁ ⥤ D₁` is a localization functor for `W₁ : MorphismProperty C₁`, then if we transport this functor `L₁` via equivalences `C₁ ≌ C₂` and `D₁ ≌ D₂` to get a functor diff --git a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean index 045f79314e84d..e69a4d3c58363 100644 --- a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean +++ b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean @@ -89,7 +89,7 @@ lemma hasProductsOfShape (J : Type) [Finite J] [HasProductsOfShape J C] (hW : W.IsStableUnderProductsOfShape J) : HasProductsOfShape J D := hasLimitsOfShape_iff_isLeftAdjoint_const.2 - ⟨⟨_, HasProductsOfShapeAux.adj L hW⟩⟩ + (HasProductsOfShapeAux.adj L hW).isLeftAdjoint /-- When `C` has finite products indexed by `J`, `W : MorphismProperty C` contains identities and is stable by products indexed by `J`, diff --git a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean index df1206087632c..d352137b2ba86 100644 --- a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean +++ b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean @@ -94,53 +94,50 @@ 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. -/ -noncomputable def isEquivalence_imp [G.IsEquivalence] : G'.IsEquivalence := by - let E₁ := Localization.uniq L₁ L₁' W₁ - let E₂ := Localization.uniq L₂ L₂' W₂ - let e : L₁ ⋙ G ⋙ E₂.functor ≅ L₁ ⋙ E₁.functor ⋙ G' := - calc - L₁ ⋙ G ⋙ E₂.functor ≅ Φ.functor ⋙ L₂ ⋙ E₂.functor := - (Functor.associator _ _ _).symm ≪≫ - isoWhiskerRight (CatCommSq.iso Φ.functor L₁ L₂ G).symm E₂.functor ≪≫ - Functor.associator _ _ _ - _ ≅ Φ.functor ⋙ L₂' := isoWhiskerLeft Φ.functor (compUniqFunctor L₂ L₂' W₂) - _ ≅ L₁' ⋙ G' := CatCommSq.iso Φ.functor L₁' L₂' G' - _ ≅ L₁ ⋙ E₁.functor ⋙ G' := - isoWhiskerRight (compUniqFunctor L₁ L₁' W₁).symm G' ≪≫ Functor.associator _ _ _ - exact Functor.IsEquivalence.cancelCompLeft E₁.functor G' inferInstance - (Functor.IsEquivalence.ofIso - (liftNatIso L₁ W₁ _ _ (G ⋙ E₂.functor) (E₁.functor ⋙ G') e) inferInstance) - -lemma nonempty_isEquivalence_iff : Nonempty (G.IsEquivalence) ↔ Nonempty (G'.IsEquivalence) := by - constructor - · rintro ⟨e⟩ - exact ⟨Φ.isEquivalence_imp L₁ L₂ G L₁' L₂' G'⟩ - · rintro ⟨e'⟩ - exact ⟨Φ.isEquivalence_imp L₁' L₂' G' L₁ L₂ G⟩ +lemma isEquivalence_imp [G.IsEquivalence] : G'.IsEquivalence := by + let E₁ := Localization.uniq L₁ L₁' W₁ + let E₂ := Localization.uniq L₂ L₂' W₂ + let e : L₁ ⋙ G ⋙ E₂.functor ≅ L₁ ⋙ E₁.functor ⋙ G' := + calc + L₁ ⋙ G ⋙ E₂.functor ≅ Φ.functor ⋙ L₂ ⋙ E₂.functor := + (Functor.associator _ _ _).symm ≪≫ + isoWhiskerRight (CatCommSq.iso Φ.functor L₁ L₂ G).symm E₂.functor ≪≫ + Functor.associator _ _ _ + _ ≅ Φ.functor ⋙ L₂' := isoWhiskerLeft Φ.functor (compUniqFunctor L₂ L₂' W₂) + _ ≅ L₁' ⋙ G' := CatCommSq.iso Φ.functor L₁' L₂' G' + _ ≅ L₁ ⋙ E₁.functor ⋙ G' := + 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' + +lemma isEquivalence_iff : G.IsEquivalence ↔ G'.IsEquivalence := + ⟨fun _ => Φ.isEquivalence_imp L₁ L₂ G L₁' L₂' G', + fun _ => Φ.isEquivalence_imp L₁' L₂' G' L₁ L₂ G⟩ end /-- Condition that a `LocalizerMorphism` induces an equivalence on the localized categories -/ class IsLocalizedEquivalence : Prop := /-- the induced functor on the constructed localized categories is an equivalence -/ - nonempty_isEquivalence : Nonempty ((Φ.localizedFunctor W₁.Q W₂.Q).IsEquivalence) + isEquivalence : (Φ.localizedFunctor W₁.Q W₂.Q).IsEquivalence lemma IsLocalizedEquivalence.mk' [CatCommSq Φ.functor L₁ L₂ G] [G.IsEquivalence] : Φ.IsLocalizedEquivalence where - nonempty_isEquivalence := by - rw [Φ.nonempty_isEquivalence_iff W₁.Q W₂.Q (Φ.localizedFunctor W₁.Q W₂.Q) L₁ L₂ G] - exact ⟨inferInstance⟩ + isEquivalence := by + rw [Φ.isEquivalence_iff W₁.Q W₂.Q (Φ.localizedFunctor W₁.Q W₂.Q) L₁ L₂ G] + exact inferInstance /-- If a `LocalizerMorphism` is a localized equivalence, then any compatible functor between the localized categories is an equivalence. -/ -noncomputable def isEquivalence [h : Φ.IsLocalizedEquivalence] [CatCommSq Φ.functor L₁ L₂ G] : - G.IsEquivalence := Nonempty.some (by - rw [Φ.nonempty_isEquivalence_iff L₁ L₂ G W₁.Q W₂.Q (Φ.localizedFunctor W₁.Q W₂.Q)] - exact h.nonempty_isEquivalence) +lemma isEquivalence [h : Φ.IsLocalizedEquivalence] [CatCommSq Φ.functor L₁ L₂ G] : + G.IsEquivalence := (by + rw [Φ.isEquivalence_iff L₁ L₂ G W₁.Q W₂.Q (Φ.localizedFunctor W₁.Q W₂.Q)] + exact h.isEquivalence) /-- If a `LocalizerMorphism` is a localized equivalence, then the induced functor on the localized categories is an equivalence -/ -noncomputable instance localizedFunctor_isEquivalence [Φ.IsLocalizedEquivalence] : +instance localizedFunctor_isEquivalence [Φ.IsLocalizedEquivalence] : (Φ.localizedFunctor L₁ L₂).IsEquivalence := Φ.isEquivalence L₁ L₂ _ diff --git a/Mathlib/CategoryTheory/Localization/Predicate.lean b/Mathlib/CategoryTheory/Localization/Predicate.lean index 60119682a70a3..403bc4c27ff50 100644 --- a/Mathlib/CategoryTheory/Localization/Predicate.lean +++ b/Mathlib/CategoryTheory/Localization/Predicate.lean @@ -51,15 +51,14 @@ class IsLocalization : Prop where /-- the functor inverts the given `MorphismProperty` -/ inverts : W.IsInvertedBy L /-- the induced functor from the constructed localized category is an equivalence -/ - nonempty_isEquivalence : Nonempty (IsEquivalence (Localization.Construction.lift L inverts)) + isEquivalence : IsEquivalence (Localization.Construction.lift L inverts) #align category_theory.functor.is_localization CategoryTheory.Functor.IsLocalization instance q_isLocalization : W.Q.IsLocalization W where inverts := W.Q_inverts - nonempty_isEquivalence := by + isEquivalence := by suffices Localization.Construction.lift W.Q W.Q_inverts = 𝟭 _ by - apply Nonempty.intro rw [this] infer_instance apply Localization.Construction.uniq @@ -126,25 +125,11 @@ theorem IsLocalization.mk' (h₁ : Localization.StrictUniversalPropertyFixedTarg (h₂ : Localization.StrictUniversalPropertyFixedTarget L W W.Localization) : IsLocalization L W := { inverts := h₁.inverts - nonempty_isEquivalence := - Nonempty.intro - { inverse := h₂.lift W.Q W.Q_inverts - unitIso := - eqToIso - (Localization.Construction.uniq _ _ - (by - simp only [← Functor.assoc, Localization.Construction.fac, h₂.fac, - Functor.comp_id])) - counitIso := - eqToIso - (h₁.uniq _ _ - (by - simp only [← Functor.assoc, h₂.fac, Localization.Construction.fac, - Functor.comp_id])) - functor_unitIso_comp := fun X => by - simp only [eqToIso.hom, eqToHom_app, eqToHom_map, eqToHom_trans, - eqToHom_refl] - rfl } } + isEquivalence := IsEquivalence.mk' (h₂.lift W.Q W.Q_inverts) + (eqToIso (Localization.Construction.uniq _ _ (by + simp only [← Functor.assoc, Localization.Construction.fac, h₂.fac, Functor.comp_id]))) + (eqToIso (h₁.uniq _ _ (by + simp only [← Functor.assoc, h₂.fac, Localization.Construction.fac, Functor.comp_id]))) } #align category_theory.functor.is_localization.mk' CategoryTheory.Functor.IsLocalization.mk' theorem IsLocalization.for_id (hW : W ≤ MorphismProperty.isomorphisms C) : (𝟭 C).IsLocalization W := @@ -171,7 +156,7 @@ def isoOfHom {X Y : C} (f : X ⟶ Y) (hf : W f) : L.obj X ≅ L.obj Y := #align category_theory.localization.iso_of_hom CategoryTheory.Localization.isoOfHom instance : (Localization.Construction.lift L (inverts L W)).IsEquivalence := - (inferInstance : L.IsLocalization W).nonempty_isEquivalence.some + (inferInstance : L.IsLocalization W).isEquivalence /-- A chosen equivalence of categories `W.Localization ≅ D` for a functor `L : C ⥤ D` which satisfies `L.IsLocalization W`. This shall be used in @@ -215,20 +200,13 @@ def whiskeringLeftFunctor : (D ⥤ E) ⥤ W.FunctorsInverting E := #align category_theory.localization.whiskering_left_functor CategoryTheory.Localization.whiskeringLeftFunctor instance : (whiskeringLeftFunctor L W E).IsEquivalence := by - refine' - Functor.IsEquivalence.ofIso _ - (Functor.IsEquivalence.ofEquivalence - ((Equivalence.congrLeft (equivalenceFromModel L W).symm).trans - (Construction.whiskeringLeftEquivalence W E))) - exact - NatIso.ofComponents - (fun F => - eqToIso - (by - ext - change (W.Q ⋙ Localization.Construction.lift L (inverts L W)) ⋙ F = L ⋙ F - rw [Construction.fac])) - fun τ => by + let iso : (whiskeringLeft (MorphismProperty.Localization W) D E).obj + (equivalenceFromModel L W).functor ⋙ + (Construction.whiskeringLeftEquivalence W E).functor ≅ whiskeringLeftFunctor L W E := + NatIso.ofComponents (fun F => eqToIso (by + ext + change (W.Q ⋙ Localization.Construction.lift L (inverts L W)) ⋙ F = L ⋙ F + rw [Construction.fac])) (fun τ => by ext dsimp [Construction.whiskeringLeftEquivalence, equivalenceFromModel, whiskerLeft] erw [NatTrans.comp_app, NatTrans.comp_app, eqToHom_app, eqToHom_app, eqToHom_refl, @@ -236,7 +214,8 @@ instance : (whiskeringLeftFunctor L W E).IsEquivalence := by · rfl all_goals change (W.Q ⋙ Localization.Construction.lift L (inverts L W)) ⋙ _ = L ⋙ _ - rw [Construction.fac] + rw [Construction.fac]) + exact Functor.isEquivalence_of_iso iso /-- The equivalence of categories `(D ⥤ E) ≌ (W.FunctorsInverting E)` induced by the composition with a localization functor `L : C ⥤ D` with respect to @@ -422,8 +401,7 @@ theorem of_iso {L₁ L₂ : C ⥤ D} (e : L₁ ≅ L₂) [L₁.IsLocalization W] let F₂ := Localization.Construction.lift L₂ h exact { inverts := h - nonempty_isEquivalence := - Nonempty.intro (IsEquivalence.ofIso (liftNatIso W.Q W L₁ L₂ F₁ F₂ e) inferInstance) } + isEquivalence := Functor.isEquivalence_of_iso (liftNatIso W.Q W L₁ L₂ F₁ F₂ e) } #align category_theory.functor.is_localization.of_iso CategoryTheory.Functor.IsLocalization.of_iso /-- If `L : C ⥤ D` is a localization for `W : MorphismProperty C`, then it is also @@ -438,7 +416,7 @@ theorem of_equivalence_target {E : Type*} [Category E] (L' : C ⥤ E) (eq : D let e' : F₁ ⋙ eq.functor ≅ F₂ := liftNatIso W.Q W (L ⋙ eq.functor) L' _ _ e exact { inverts := h - nonempty_isEquivalence := Nonempty.intro (IsEquivalence.ofIso e' inferInstance) } + isEquivalence := Functor.isEquivalence_of_iso e' } #align category_theory.functor.is_localization.of_equivalence_target CategoryTheory.Functor.IsLocalization.of_equivalence_target lemma of_isEquivalence (L : C ⥤ D) (W : MorphismProperty C) diff --git a/Mathlib/CategoryTheory/Monad/Adjunction.lean b/Mathlib/CategoryTheory/Monad/Adjunction.lean index 05a99b1a84b82..229774e8981a3 100644 --- a/Mathlib/CategoryTheory/Monad/Adjunction.lean +++ b/Mathlib/CategoryTheory/Monad/Adjunction.lean @@ -190,26 +190,68 @@ instance (G : Comonad C) : (Comonad.comparison G.adj).EssSurj where /-- A right adjoint functor `R : D ⥤ C` is *monadic* if the comparison functor `Monad.comparison R` from `D` to the category of Eilenberg-Moore algebras for the adjunction is an equivalence. -/ -class MonadicRightAdjoint (R : D ⥤ C) extends IsRightAdjoint R where - eqv : (Monad.comparison (Adjunction.ofRightAdjoint R)).IsEquivalence +class MonadicRightAdjoint (R : D ⥤ C) where + /-- a choice of left adjoint for `R` -/ + L : C ⥤ D + /-- `R` is a right adjoint -/ + adj : L ⊣ R + eqv : (Monad.comparison adj).IsEquivalence #align category_theory.monadic_right_adjoint CategoryTheory.MonadicRightAdjoint +/-- The left adjoint functor to `R` given by `[MonadicRightAdjoint R]`. -/ +def monadicLeftAdjoint (R : D ⥤ C) [MonadicRightAdjoint R] : C ⥤ D := + MonadicRightAdjoint.L (R := R) + +/-- The adjunction `monadicLeftAdjoint R ⊣ R` given by `[MonadicRightAdjoint R]`. -/ +def monadicAdjunction (R : D ⥤ C) [MonadicRightAdjoint R] : + monadicLeftAdjoint R ⊣ R := + MonadicRightAdjoint.adj + +instance (R : D ⥤ C) [MonadicRightAdjoint R] : + (Monad.comparison (monadicAdjunction R)).IsEquivalence := + MonadicRightAdjoint.eqv + +instance (R : D ⥤ C) [MonadicRightAdjoint R] : R.IsRightAdjoint := + (monadicAdjunction R).isRightAdjoint + +noncomputable instance (T : Monad C) : MonadicRightAdjoint T.forget where + adj := T.adj + eqv := { } + /-- A left adjoint functor `L : C ⥤ D` is *comonadic* if the comparison functor `Comonad.comparison L` from `C` to the category of Eilenberg-Moore algebras for the adjunction is an equivalence. -/ -class ComonadicLeftAdjoint (L : C ⥤ D) extends IsLeftAdjoint L where - eqv : (Comonad.comparison (Adjunction.ofLeftAdjoint L)).IsEquivalence +class ComonadicLeftAdjoint (L : C ⥤ D) where + /-- a choice of right adjoint for `L` -/ + R : D ⥤ C + /-- `L` is a right adjoint -/ + adj : L ⊣ R + eqv : (Comonad.comparison adj).IsEquivalence #align category_theory.comonadic_left_adjoint CategoryTheory.ComonadicLeftAdjoint -noncomputable instance (T : Monad C) : MonadicRightAdjoint T.forget := - ⟨Functor.IsEquivalence.ofFullyFaithfullyEssSurj (Monad.comparison T.adj)⟩ +/-- The right adjoint functor to `L` given by `[ComonadicLeftAdjoint L]`. -/ +def comonadicRightAdjoint (L : C ⥤ D) [ComonadicLeftAdjoint L] : D ⥤ C := + ComonadicLeftAdjoint.R (L := L) + +/-- The adjunction `L ⊣ comonadicRightAdjoint L` given by `[ComonadicLeftAdjoint L]`. -/ +def comonadicAdjunction (L : C ⥤ D) [ComonadicLeftAdjoint L] : + L ⊣ comonadicRightAdjoint L := + ComonadicLeftAdjoint.adj + +instance (L : C ⥤ D) [ComonadicLeftAdjoint L] : + (Comonad.comparison (comonadicAdjunction L)).IsEquivalence := + ComonadicLeftAdjoint.eqv + +instance (L : C ⥤ D) [ComonadicLeftAdjoint L] : L.IsLeftAdjoint := + (comonadicAdjunction L).isLeftAdjoint -noncomputable instance (G : Comonad C) : ComonadicLeftAdjoint G.forget := - ⟨Functor.IsEquivalence.ofFullyFaithfullyEssSurj (Comonad.comparison G.adj)⟩ +noncomputable instance (G : Comonad C) : ComonadicLeftAdjoint G.forget where + adj := G.adj + eqv := { } -- TODO: This holds more generally for idempotent adjunctions, not just reflective adjunctions. -instance μ_iso_of_reflective [Reflective R] : IsIso (Adjunction.ofRightAdjoint R).toMonad.μ := by +instance μ_iso_of_reflective [Reflective R] : IsIso (reflectorAdjunction R).toMonad.μ := by dsimp infer_instance #align category_theory.μ_iso_of_reflective CategoryTheory.μ_iso_of_reflective @@ -219,34 +261,34 @@ attribute [instance] ComonadicLeftAdjoint.eqv namespace Reflective -instance [Reflective R] (X : (Adjunction.ofRightAdjoint R).toMonad.Algebra) : - IsIso ((Adjunction.ofRightAdjoint R).unit.app X.A) := +instance [Reflective R] (X : (reflectorAdjunction R).toMonad.Algebra) : + IsIso ((reflectorAdjunction R).unit.app X.A) := ⟨⟨X.a, ⟨X.unit, by dsimp only [Functor.id_obj] - rw [← (Adjunction.ofRightAdjoint R).unit_naturality] + rw [← (reflectorAdjunction R).unit_naturality] dsimp only [Functor.comp_obj, Adjunction.toMonad_coe] rw [unit_obj_eq_map_unit, ← Functor.map_comp, ← Functor.map_comp] erw [X.unit] simp⟩⟩⟩ instance comparison_essSurj [Reflective R] : - (Monad.comparison (Adjunction.ofRightAdjoint R)).EssSurj := by - refine' ⟨fun X => ⟨(leftAdjoint R).obj X.A, ⟨_⟩⟩⟩ + (Monad.comparison (reflectorAdjunction R)).EssSurj := by + refine' ⟨fun X => ⟨(reflector R).obj X.A, ⟨_⟩⟩⟩ symm refine' Monad.Algebra.isoMk _ _ - · exact asIso ((Adjunction.ofRightAdjoint R).unit.app X.A) + · exact asIso ((reflectorAdjunction R).unit.app X.A) dsimp only [Functor.comp_map, Monad.comparison_obj_a, asIso_hom, Functor.comp_obj, Monad.comparison_obj_A, Adjunction.toMonad_coe] - rw [← cancel_epi ((Adjunction.ofRightAdjoint R).unit.app X.A)] + rw [← cancel_epi ((reflectorAdjunction R).unit.app X.A)] dsimp only [Functor.id_obj, Functor.comp_obj] rw [Adjunction.unit_naturality_assoc, Adjunction.right_triangle_components, comp_id] apply (X.unit_assoc _).symm #align category_theory.reflective.comparison_ess_surj CategoryTheory.Reflective.comparison_essSurj -instance comparison_full [R.Full] [IsRightAdjoint R] : - (Monad.comparison (Adjunction.ofRightAdjoint R)).Full where +lemma comparison_full [R.Full] {L : C ⥤ D} (adj : L ⊣ R): + (Monad.comparison adj).Full where map_surjective f := ⟨R.preimage f.f, by aesop_cat⟩ #align category_theory.reflective.comparison_full CategoryTheory.Reflective.comparison_full @@ -257,9 +299,10 @@ end Reflective -- see Note [lower instance priority] /-- Any reflective inclusion has a monadic right adjoint. cf Prop 5.3.3 of [Riehl][riehl2017] -/ -noncomputable instance (priority := 100) monadicOfReflective [Reflective R] : +instance (priority := 100) monadicOfReflective [Reflective R] : MonadicRightAdjoint R where - eqv := Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ + adj := reflectorAdjunction R + eqv := { full := Reflective.comparison_full _ } #align category_theory.monadic_of_reflective CategoryTheory.monadicOfReflective end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monad/Algebra.lean b/Mathlib/CategoryTheory/Monad/Algebra.lean index 874cf32eddab9..f2de843bfe406 100644 --- a/Mathlib/CategoryTheory/Monad/Algebra.lean +++ b/Mathlib/CategoryTheory/Monad/Algebra.lean @@ -239,18 +239,8 @@ theorem algebra_mono_of_mono {X Y : Algebra T} (f : X ⟶ Y) [h : Mono f.f] : Mo (forget T).mono_of_mono_map h #align category_theory.monad.algebra_mono_of_mono CategoryTheory.Monad.algebra_mono_of_mono -instance : IsRightAdjoint T.forget := - ⟨T.free, T.adj⟩ - -@[simp] -theorem leftAdjoint_forget : leftAdjoint T.forget = T.free := - rfl -#align category_theory.monad.left_adjoint_forget CategoryTheory.Monad.leftAdjoint_forget - -@[simp] -theorem ofRightAdjoint_forget : Adjunction.ofRightAdjoint T.forget = T.adj := - rfl -#align category_theory.monad.of_right_adjoint_forget CategoryTheory.Monad.ofRightAdjoint_forget +instance : T.forget.IsRightAdjoint := + ⟨T.free, ⟨T.adj⟩⟩ /-- Given a monad morphism from `T₂` to `T₁`, we get a functor from the algebras of `T₁` to algebras of @@ -520,18 +510,8 @@ theorem algebra_mono_of_mono {X Y : Coalgebra G} (f : X ⟶ Y) [h : Mono f.f] : (forget G).mono_of_mono_map h #align category_theory.comonad.algebra_mono_of_mono CategoryTheory.Comonad.algebra_mono_of_mono -instance : IsLeftAdjoint G.forget := - ⟨_, G.adj⟩ - -@[simp] -theorem rightAdjoint_forget : rightAdjoint G.forget = G.cofree := - rfl -#align category_theory.comonad.right_adjoint_forget CategoryTheory.Comonad.rightAdjoint_forget - -@[simp] -theorem ofLeftAdjoint_forget : Adjunction.ofLeftAdjoint G.forget = G.adj := - rfl -#align category_theory.comonad.of_left_adjoint_forget CategoryTheory.Comonad.ofLeftAdjoint_forget +instance : G.forget.IsLeftAdjoint := + ⟨_, ⟨G.adj⟩⟩ end Comonad diff --git a/Mathlib/CategoryTheory/Monad/Limits.lean b/Mathlib/CategoryTheory/Monad/Limits.lean index c2c18aec963af..8f4ff1e35b0c8 100644 --- a/Mathlib/CategoryTheory/Monad/Limits.lean +++ b/Mathlib/CategoryTheory/Monad/Limits.lean @@ -280,59 +280,59 @@ variable {J : Type u} [Category.{v} J] instance comp_comparison_forget_hasLimit (F : J ⥤ D) (R : D ⥤ C) [MonadicRightAdjoint R] [HasLimit (F ⋙ R)] : - HasLimit ((F ⋙ Monad.comparison (Adjunction.ofRightAdjoint R)) ⋙ Monad.forget _) := + HasLimit ((F ⋙ Monad.comparison (monadicAdjunction R)) ⋙ Monad.forget _) := @hasLimitOfIso _ _ _ _ (F ⋙ R) _ _ - (isoWhiskerLeft F (Monad.comparisonForget (Adjunction.ofRightAdjoint R)).symm) + (isoWhiskerLeft F (Monad.comparisonForget (monadicAdjunction R)).symm) #align category_theory.comp_comparison_forget_has_limit CategoryTheory.comp_comparison_forget_hasLimit instance comp_comparison_hasLimit (F : J ⥤ D) (R : D ⥤ C) [MonadicRightAdjoint R] - [HasLimit (F ⋙ R)] : HasLimit (F ⋙ Monad.comparison (Adjunction.ofRightAdjoint R)) := - Monad.hasLimit_of_comp_forget_hasLimit (F ⋙ Monad.comparison (Adjunction.ofRightAdjoint R)) + [HasLimit (F ⋙ R)] : HasLimit (F ⋙ Monad.comparison (monadicAdjunction R)) := + Monad.hasLimit_of_comp_forget_hasLimit (F ⋙ Monad.comparison (monadicAdjunction R)) #align category_theory.comp_comparison_has_limit CategoryTheory.comp_comparison_hasLimit /-- Any monadic functor creates limits. -/ noncomputable def monadicCreatesLimits (R : D ⥤ C) [MonadicRightAdjoint R] : CreatesLimitsOfSize.{v, u} R := - createsLimitsOfNatIso (Monad.comparisonForget (Adjunction.ofRightAdjoint R)) + createsLimitsOfNatIso (Monad.comparisonForget (monadicAdjunction R)) #align category_theory.monadic_creates_limits CategoryTheory.monadicCreatesLimits /-- The forgetful functor from the Eilenberg-Moore category for a monad creates any colimit which the monad itself preserves. -/ noncomputable def monadicCreatesColimitOfPreservesColimit (R : D ⥤ C) (K : J ⥤ D) - [MonadicRightAdjoint R] [PreservesColimit (K ⋙ R) (leftAdjoint R ⋙ R)] - [PreservesColimit ((K ⋙ R) ⋙ leftAdjoint R ⋙ R) (leftAdjoint R ⋙ R)] : + [MonadicRightAdjoint R] [PreservesColimit (K ⋙ R) (monadicLeftAdjoint R ⋙ R)] + [PreservesColimit ((K ⋙ R) ⋙ monadicLeftAdjoint R ⋙ R) (monadicLeftAdjoint R ⋙ R)] : CreatesColimit K R := by -- Porting note: It would be nice to have a variant of apply which introduces goals for missing -- instances. - letI A := Monad.comparison (Adjunction.ofRightAdjoint R) - letI B := Monad.forget (Adjunction.toMonad (Adjunction.ofRightAdjoint R)) - let i : (K ⋙ Monad.comparison (Adjunction.ofRightAdjoint R)) ⋙ Monad.forget _ ≅ K ⋙ R := + letI A := Monad.comparison (monadicAdjunction R) + letI B := Monad.forget (Adjunction.toMonad (monadicAdjunction R)) + let i : (K ⋙ Monad.comparison (monadicAdjunction R)) ⋙ Monad.forget _ ≅ K ⋙ R := Functor.associator _ _ _ ≪≫ - isoWhiskerLeft K (Monad.comparisonForget (Adjunction.ofRightAdjoint R)) + isoWhiskerLeft K (Monad.comparisonForget (monadicAdjunction R)) letI : PreservesColimit ((K ⋙ A) ⋙ Monad.forget - (Adjunction.toMonad (Adjunction.ofRightAdjoint R))) - (Adjunction.toMonad (Adjunction.ofRightAdjoint R)).toFunctor := by + (Adjunction.toMonad (monadicAdjunction R))) + (Adjunction.toMonad (monadicAdjunction R)).toFunctor := by dsimp exact preservesColimitOfIsoDiagram _ i.symm letI : PreservesColimit - (((K ⋙ A) ⋙ Monad.forget (Adjunction.toMonad (Adjunction.ofRightAdjoint R))) ⋙ - (Adjunction.toMonad (Adjunction.ofRightAdjoint R)).toFunctor) - (Adjunction.toMonad (Adjunction.ofRightAdjoint R)).toFunctor := by + (((K ⋙ A) ⋙ Monad.forget (Adjunction.toMonad (monadicAdjunction R))) ⋙ + (Adjunction.toMonad (monadicAdjunction R)).toFunctor) + (Adjunction.toMonad (monadicAdjunction R)).toFunctor := by dsimp - exact preservesColimitOfIsoDiagram _ (isoWhiskerRight i (leftAdjoint R ⋙ R)).symm + exact preservesColimitOfIsoDiagram _ (isoWhiskerRight i (monadicLeftAdjoint R ⋙ R)).symm letI : CreatesColimit (K ⋙ A) B := CategoryTheory.Monad.forgetCreatesColimit _ letI : CreatesColimit K (A ⋙ B) := CategoryTheory.compCreatesColimit _ _ - let e := Monad.comparisonForget (Adjunction.ofRightAdjoint R) + let e := Monad.comparisonForget (monadicAdjunction R) apply createsColimitOfNatIso e #align category_theory.monadic_creates_colimit_of_preserves_colimit CategoryTheory.monadicCreatesColimitOfPreservesColimit /-- A monadic functor creates any colimits of shapes it preserves. -/ noncomputable def monadicCreatesColimitsOfShapeOfPreservesColimitsOfShape (R : D ⥤ C) [MonadicRightAdjoint R] [PreservesColimitsOfShape J R] : CreatesColimitsOfShape J R := - letI : PreservesColimitsOfShape J (leftAdjoint R) := by - apply (Adjunction.leftAdjointPreservesColimits (Adjunction.ofRightAdjoint R)).1 - letI : PreservesColimitsOfShape J (leftAdjoint R ⋙ R) := by + letI : PreservesColimitsOfShape J (monadicLeftAdjoint R) := by + apply (Adjunction.leftAdjointPreservesColimits (monadicAdjunction R)).1 + letI : PreservesColimitsOfShape J (monadicLeftAdjoint R ⋙ R) := by apply CategoryTheory.Limits.compPreservesColimitsOfShape _ _ ⟨monadicCreatesColimitOfPreservesColimit _ _⟩ #align category_theory.monadic_creates_colimits_of_shape_of_preserves_colimits_of_shape CategoryTheory.monadicCreatesColimitsOfShapeOfPreservesColimitsOfShape @@ -368,13 +368,13 @@ theorem hasLimits_of_reflective (R : D ⥤ C) [HasLimitsOfSize.{v, u} C] [Reflec theorem hasColimitsOfShape_of_reflective (R : D ⥤ C) [Reflective R] [HasColimitsOfShape J C] : HasColimitsOfShape J D where has_colimit := fun F => by - let c := (leftAdjoint R).mapCocone (colimit.cocone (F ⋙ R)) + let c := (monadicLeftAdjoint R).mapCocone (colimit.cocone (F ⋙ R)) letI : PreservesColimitsOfShape J _ := - (Adjunction.ofRightAdjoint R).leftAdjointPreservesColimits.1 - let t : IsColimit c := isColimitOfPreserves (leftAdjoint R) (colimit.isColimit _) + (monadicAdjunction R).leftAdjointPreservesColimits.1 + let t : IsColimit c := isColimitOfPreserves (monadicLeftAdjoint R) (colimit.isColimit _) apply HasColimit.mk ⟨_, (IsColimit.precomposeInvEquiv _ _).symm t⟩ apply - (isoWhiskerLeft F (asIso (Adjunction.ofRightAdjoint R).counit) : _) ≪≫ F.rightUnitor + (isoWhiskerLeft F (asIso (monadicAdjunction R).counit) : _) ≪≫ F.rightUnitor #align category_theory.has_colimits_of_shape_of_reflective CategoryTheory.hasColimitsOfShape_of_reflective /-- If `C` has colimits then any reflective subcategory has colimits. -/ @@ -387,17 +387,17 @@ theorem hasColimits_of_reflective (R : D ⥤ C) [Reflective R] [HasColimitsOfSiz limit. -/ noncomputable def leftAdjointPreservesTerminalOfReflective (R : D ⥤ C) [Reflective R] : - PreservesLimitsOfShape (Discrete.{v} PEmpty) (leftAdjoint R) where + PreservesLimitsOfShape (Discrete.{v} PEmpty) (monadicLeftAdjoint R) where preservesLimit {K} := by let F := Functor.empty.{v} D - letI : PreservesLimit (F ⋙ R) (leftAdjoint R) := by + letI : PreservesLimit (F ⋙ R) (monadicLeftAdjoint R) := by constructor intro c h haveI : HasLimit (F ⋙ R) := ⟨⟨⟨c, h⟩⟩⟩ haveI : HasLimit F := hasLimit_of_reflective F R apply isLimitChangeEmptyCone D (limit.isLimit F) - apply (asIso ((Adjunction.ofRightAdjoint R).counit.app _)).symm.trans - apply (leftAdjoint R).mapIso + apply (asIso ((monadicAdjunction R).counit.app _)).symm.trans + apply (monadicLeftAdjoint R).mapIso letI := monadicCreatesLimits.{v, v} R let A := CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit F R apply (A.preserves (limit.isLimit F)).conePointUniqueUpToIso h diff --git a/Mathlib/CategoryTheory/Monad/Monadicity.lean b/Mathlib/CategoryTheory/Monad/Monadicity.lean index bae5504c6500d..33d86a33ee42e 100644 --- a/Mathlib/CategoryTheory/Monad/Monadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Monadicity.lean @@ -53,26 +53,18 @@ namespace MonadicityInternal variable {C : Type u₁} {D : Type u₂} variable [Category.{v₁} C] [Category.{v₁} D] -variable {G : D ⥤ C} [IsRightAdjoint G] - --- An unfortunate consequence of the local notation is that it is only recognised if there --- is some syntax to separate it from dot notation. --- Porting note: in Lean 3 you could write `F .map A.a` for example, but in Lean 4 that would --- mean passing two arguments to `F` (`.map` and `A.a`). We write `(F).map A.a` instead. --- As a warning, this pretty prints as `F.map A.a`. -local notation3 "F" => leftAdjoint G -local notation3 "adj" => Adjunction.ofRightAdjoint G +variable {G : D ⥤ C} {F : C ⥤ D} (adj : F ⊣ G) /-- The "main pair" for an algebra `(A, α)` is the pair of morphisms `(F α, ε_FA)`. It is always a reflexive pair, and will be used to construct the left adjoint to the comparison functor and show it is an equivalence. -/ -instance main_pair_reflexive (A : (adj).toMonad.Algebra) : - IsReflexivePair ((F).map A.a) ((adj).counit.app ((F).obj A.A)) := by - apply IsReflexivePair.mk' ((F).map ((adj).unit.app _)) _ _ - · rw [← (F).map_comp, ← (F).map_id] - exact congr_arg (F).map A.unit - · rw [(adj).left_triangle_components] +instance main_pair_reflexive (A : adj.toMonad.Algebra) : + IsReflexivePair (F.map A.a) (adj.counit.app (F.obj A.A)) := by + apply IsReflexivePair.mk' (F.map (adj.unit.app _)) _ _ + · rw [← F.map_comp, ← F.map_id] + exact congr_arg F.map A.unit + · rw [adj.left_triangle_components] rfl #align category_theory.monad.monadicity_internal.main_pair_reflexive CategoryTheory.Monad.MonadicityInternal.main_pair_reflexive @@ -80,17 +72,17 @@ instance main_pair_reflexive (A : (adj).toMonad.Algebra) : `G`-split pair, and will be used to construct the left adjoint to the comparison functor and show it is an equivalence. -/ -instance main_pair_G_split (A : (adj).toMonad.Algebra) : - G.IsSplitPair ((F).map A.a) - ((adj).counit.app ((F).obj A.A)) where +instance main_pair_G_split (A : adj.toMonad.Algebra) : + G.IsSplitPair (F.map A.a) + (adj.counit.app (F.obj A.A)) where splittable := ⟨_, _, ⟨beckSplitCoequalizer A⟩⟩ set_option linter.uppercaseLean3 false in #align category_theory.monad.monadicity_internal.main_pair_G_split CategoryTheory.Monad.MonadicityInternal.main_pair_G_split /-- The object function for the left adjoint to the comparison functor. -/ -def comparisonLeftAdjointObj (A : (adj).toMonad.Algebra) - [HasCoequalizer ((F).map A.a) ((adj).counit.app _)] : D := - coequalizer ((F).map A.a) ((adj).counit.app _) +def comparisonLeftAdjointObj (A : adj.toMonad.Algebra) + [HasCoequalizer (F.map A.a) (adj.counit.app _)] : D := + coequalizer (F.map A.a) (adj.counit.app _) #align category_theory.monad.monadicity_internal.comparison_left_adjoint_obj CategoryTheory.Monad.MonadicityInternal.comparisonLeftAdjointObj /-- @@ -98,20 +90,20 @@ We have a bijection of homsets which will be used to construct the left adjoint functor. -/ @[simps!] -def comparisonLeftAdjointHomEquiv (A : (adj).toMonad.Algebra) (B : D) - [HasCoequalizer ((F).map A.a) ((adj).counit.app ((F).obj A.A))] : - (comparisonLeftAdjointObj A ⟶ B) ≃ (A ⟶ (comparison adj).obj B) := +def comparisonLeftAdjointHomEquiv (A : adj.toMonad.Algebra) (B : D) + [HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] : + (comparisonLeftAdjointObj adj A ⟶ B) ≃ (A ⟶ (comparison adj).obj B) := calc - (comparisonLeftAdjointObj A ⟶ B) ≃ { f : (F).obj A.A ⟶ B // _ } := + (comparisonLeftAdjointObj adj A ⟶ B) ≃ { f : F.obj A.A ⟶ B // _ } := Cofork.IsColimit.homIso (colimit.isColimit _) B - _ ≃ { g : A.A ⟶ G.obj B // G.map ((F).map g) ≫ G.map ((adj).counit.app B) = A.a ≫ g } := by - refine' ((adj).homEquiv _ _).subtypeEquiv _ + _ ≃ { g : A.A ⟶ G.obj B // G.map (F.map g) ≫ G.map (adj.counit.app B) = A.a ≫ g } := by + refine' (adj.homEquiv _ _).subtypeEquiv _ intro f - rw [← ((adj).homEquiv _ _).injective.eq_iff, Adjunction.homEquiv_naturality_left, - (adj).homEquiv_unit, (adj).homEquiv_unit, G.map_comp] + rw [← (adj.homEquiv _ _).injective.eq_iff, Adjunction.homEquiv_naturality_left, + adj.homEquiv_unit, adj.homEquiv_unit, G.map_comp] dsimp - rw [(adj).right_triangle_components_assoc, ← G.map_comp, (F).map_comp, Category.assoc, - (adj).counit_naturality, (adj).left_triangle_components_assoc] + rw [adj.right_triangle_components_assoc, ← G.map_comp, F.map_comp, Category.assoc, + adj.counit_naturality, adj.left_triangle_components_assoc] apply eq_comm _ ≃ (A ⟶ (comparison adj).obj B) := { toFun := fun g => @@ -125,19 +117,19 @@ def comparisonLeftAdjointHomEquiv (A : (adj).toMonad.Algebra) (B : D) /-- Construct the adjunction to the comparison functor. -/ def leftAdjointComparison - [∀ A : (adj).toMonad.Algebra, HasCoequalizer ((F).map A.a) - ((adj).counit.app ((F).obj A.A))] : - (adj).toMonad.Algebra ⥤ D := by + [∀ A : adj.toMonad.Algebra, HasCoequalizer (F.map A.a) + (adj.counit.app (F.obj A.A))] : + adj.toMonad.Algebra ⥤ D := by refine' Adjunction.leftAdjointOfEquiv (G := comparison adj) - (F_obj := fun A => comparisonLeftAdjointObj A) (fun A B => _) _ + (F_obj := fun A => comparisonLeftAdjointObj adj A) (fun A B => _) _ · apply comparisonLeftAdjointHomEquiv · intro A B B' g h ext1 -- Porting note: the goal was previously closed by the following, which succeeds until -- `Category.assoc`. -- dsimp [comparisonLeftAdjointHomEquiv] - -- rw [← (adj).homEquiv_naturality_right, Category.assoc] + -- rw [← adj.homEquiv_naturality_right, Category.assoc] simp [Cofork.IsColimit.homIso] #align category_theory.monad.monadicity_internal.left_adjoint_comparison CategoryTheory.Monad.MonadicityInternal.leftAdjointComparison @@ -145,97 +137,102 @@ def leftAdjointComparison -/ @[simps! counit] def comparisonAdjunction - [∀ A : (adj).toMonad.Algebra, HasCoequalizer ((F).map A.a) - ((adj).counit.app ((F).obj A.A))] : - leftAdjointComparison ⊣ comparison (adj) := + [∀ A : adj.toMonad.Algebra, HasCoequalizer (F.map A.a) + (adj.counit.app (F.obj A.A))] : + leftAdjointComparison adj ⊣ comparison adj := Adjunction.adjunctionOfEquivLeft _ _ #align category_theory.monad.monadicity_internal.comparison_adjunction CategoryTheory.Monad.MonadicityInternal.comparisonAdjunction +variable {adj} + theorem comparisonAdjunction_unit_f_aux - [∀ A : (adj).toMonad.Algebra, HasCoequalizer ((F).map A.a) - ((adj).counit.app ((F).obj A.A))] - (A : (adj).toMonad.Algebra) : - (comparisonAdjunction.unit.app A).f = - (adj).homEquiv A.A _ - (coequalizer.π ((F).map A.a) ((adj).counit.app ((F).obj A.A))) := - congr_arg ((adj).homEquiv _ _) (Category.comp_id _) + [∀ A : adj.toMonad.Algebra, HasCoequalizer (F.map A.a) + (adj.counit.app (F.obj A.A))] + (A : adj.toMonad.Algebra) : + ((comparisonAdjunction adj).unit.app A).f = + adj.homEquiv A.A _ + (coequalizer.π (F.map A.a) (adj.counit.app (F.obj A.A))) := + congr_arg (adj.homEquiv _ _) (Category.comp_id _) #align category_theory.monad.monadicity_internal.comparison_adjunction_unit_f_aux CategoryTheory.Monad.MonadicityInternal.comparisonAdjunction_unit_f_aux /-- This is a cofork which is helpful for establishing monadicity: the morphism from the Beck coequalizer to this cofork is the unit for the adjunction on the comparison functor. -/ @[simps! pt] -def unitCofork (A : (adj).toMonad.Algebra) - [HasCoequalizer ((F).map A.a) ((adj).counit.app ((F).obj A.A))] : - Cofork (G.map ((F).map A.a)) (G.map ((adj).counit.app ((F).obj A.A))) := - Cofork.ofπ (G.map (coequalizer.π ((F).map A.a) ((adj).counit.app ((F).obj A.A)))) +def unitCofork (A : adj.toMonad.Algebra) + [HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] : + Cofork (G.map (F.map A.a)) (G.map (adj.counit.app (F.obj A.A))) := + Cofork.ofπ (G.map (coequalizer.π (F.map A.a) (adj.counit.app (F.obj A.A)))) (by change _ = G.map _ ≫ _ rw [← G.map_comp, coequalizer.condition, G.map_comp]) #align category_theory.monad.monadicity_internal.unit_cofork CategoryTheory.Monad.MonadicityInternal.unitCofork @[simp] -theorem unitCofork_π (A : (adj).toMonad.Algebra) - [HasCoequalizer ((F).map A.a) ((adj).counit.app ((F).obj A.A))] : - (unitCofork A).π = G.map (coequalizer.π ((F).map A.a) ((adj).counit.app ((F).obj A.A))) := +theorem unitCofork_π (A : adj.toMonad.Algebra) + [HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] : + (unitCofork A).π = G.map (coequalizer.π (F.map A.a) (adj.counit.app (F.obj A.A))) := rfl #align category_theory.monad.monadicity_internal.unit_cofork_π CategoryTheory.Monad.MonadicityInternal.unitCofork_π theorem comparisonAdjunction_unit_f - [∀ A : (adj).toMonad.Algebra, HasCoequalizer ((F).map A.a) - ((adj).counit.app ((F).obj A.A))] - (A : (adj).toMonad.Algebra) : - (comparisonAdjunction.unit.app A).f = (beckCoequalizer A).desc (unitCofork A) := by + [∀ A : adj.toMonad.Algebra, HasCoequalizer (F.map A.a) + (adj.counit.app (F.obj A.A))] + (A : adj.toMonad.Algebra) : + ((comparisonAdjunction adj).unit.app A).f = (beckCoequalizer A).desc (unitCofork A) := by apply Limits.Cofork.IsColimit.hom_ext (beckCoequalizer A) rw [Cofork.IsColimit.π_desc] dsimp only [beckCofork_π, unitCofork_π] - rw [comparisonAdjunction_unit_f_aux, ← (adj).homEquiv_naturality_left A.a, coequalizer.condition, - (adj).homEquiv_naturality_right, (adj).homEquiv_unit, Category.assoc] - apply (adj).right_triangle_components_assoc + rw [comparisonAdjunction_unit_f_aux, ← adj.homEquiv_naturality_left A.a, coequalizer.condition, + adj.homEquiv_naturality_right, adj.homEquiv_unit, Category.assoc] + apply adj.right_triangle_components_assoc #align category_theory.monad.monadicity_internal.comparison_adjunction_unit_f CategoryTheory.Monad.MonadicityInternal.comparisonAdjunction_unit_f +variable (adj) + /-- The cofork which describes the counit of the adjunction: the morphism from the coequalizer of this pair to this morphism is the counit. -/ @[simps!] def counitCofork (B : D) : - Cofork ((F).map (G.map ((adj).counit.app B))) - ((adj).counit.app ((F).obj (G.obj B))) := - Cofork.ofπ ((adj).counit.app B) ((adj).counit_naturality _) + Cofork (F.map (G.map (adj.counit.app B))) + (adj.counit.app (F.obj (G.obj B))) := + Cofork.ofπ (adj.counit.app B) (adj.counit_naturality _) #align category_theory.monad.monadicity_internal.counit_cofork CategoryTheory.Monad.MonadicityInternal.counitCofork +variable {adj} in /-- The unit cofork is a colimit provided `G` preserves it. -/ -def unitColimitOfPreservesCoequalizer (A : (adj).toMonad.Algebra) - [HasCoequalizer ((F).map A.a) ((adj).counit.app ((F).obj A.A))] - [PreservesColimit (parallelPair ((F).map A.a) ((adj).counit.app ((F).obj A.A))) G] : +def unitColimitOfPreservesCoequalizer (A : adj.toMonad.Algebra) + [HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] + [PreservesColimit (parallelPair (F.map A.a) (adj.counit.app (F.obj A.A))) G] : IsColimit (unitCofork (G := G) A) := isColimitOfHasCoequalizerOfPreservesColimit G _ _ #align category_theory.monad.monadicity_internal.unit_colimit_of_preserves_coequalizer CategoryTheory.Monad.MonadicityInternal.unitColimitOfPreservesCoequalizer /-- The counit cofork is a colimit provided `G` reflects it. -/ def counitCoequalizerOfReflectsCoequalizer (B : D) - [ReflectsColimit (parallelPair ((F).map (G.map ((adj).counit.app B))) - ((adj).counit.app ((F).obj (G.obj B)))) G] : - IsColimit (counitCofork (G := G) B) := + [ReflectsColimit (parallelPair (F.map (G.map (adj.counit.app B))) + (adj.counit.app (F.obj (G.obj B)))) G] : + IsColimit (counitCofork (adj := adj) B) := isColimitOfIsColimitCoforkMap G _ (beckCoequalizer ((comparison adj).obj B)) #align category_theory.monad.monadicity_internal.counit_coequalizer_of_reflects_coequalizer CategoryTheory.Monad.MonadicityInternal.counitCoequalizerOfReflectsCoequalizer -- Porting note: Lean 3 didn't seem to need this instance - [∀ A : (adj).toMonad.Algebra, HasCoequalizer ((F).map A.a) ((adj).counit.app ((F).obj A.A))] + [∀ A : adj.toMonad.Algebra, HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] (B : D) : HasColimit (parallelPair - ((F).map (G.map (NatTrans.app (adj).counit B))) - (NatTrans.app (adj).counit ((F).obj (G.obj B)))) := + (F.map (G.map (NatTrans.app adj.counit B))) + (NatTrans.app adj.counit (F.obj (G.obj B)))) := inferInstanceAs <| HasCoequalizer - ((F).map ((comparison (adj)).obj B).a) - ((adj).counit.app ((F).obj ((comparison (adj)).obj B).A)) + (F.map ((comparison adj).obj B).a) + (adj.counit.app (F.obj ((comparison adj).obj B).A)) theorem comparisonAdjunction_counit_app - [∀ A : (adj).toMonad.Algebra, HasCoequalizer ((F).map A.a) ((adj).counit.app ((F).obj A.A))] - (B : D) : (comparisonAdjunction (G := G)).counit.app B = colimit.desc _ (counitCofork B) := by + [∀ A : adj.toMonad.Algebra, HasCoequalizer (F.map A.a) (adj.counit.app (F.obj A.A))] (B : D) : + (comparisonAdjunction adj).counit.app B = colimit.desc _ (counitCofork adj B) := by apply coequalizer.hom_ext change - coequalizer.π _ _ ≫ coequalizer.desc (((adj).homEquiv _ B).symm (𝟙 _)) _ = + coequalizer.π _ _ ≫ coequalizer.desc ((adj.homEquiv _ B).symm (𝟙 _)) _ = coequalizer.π _ _ ≫ coequalizer.desc _ _ simp #align category_theory.monad.monadicity_internal.comparison_adjunction_counit_app CategoryTheory.Monad.MonadicityInternal.comparisonAdjunction_counit_app @@ -246,8 +243,9 @@ open CategoryTheory Adjunction Monad MonadicityInternal variable {C : Type u₁} {D : Type u₂} variable [Category.{v₁} C] [Category.{v₁} D] -variable {G : D ⥤ C} +variable {G : D ⥤ C} {F : C ⥤ D} (adj : F ⊣ G) +variable (G) in /-- If `G` is monadic, it creates colimits of `G`-split pairs. This is the "boring" direction of Beck's monadicity theorem, the converse is given in `monadicOfCreatesGSplitCoequalizers`. @@ -265,14 +263,8 @@ def createsGSplitCoequalizersOfMonadic [MonadicRightAdjoint G] ⦃A B⦄ (f g : set_option linter.uppercaseLean3 false in #align category_theory.monad.creates_G_split_coequalizers_of_monadic CategoryTheory.Monad.createsGSplitCoequalizersOfMonadic -variable [IsRightAdjoint G] - section BeckMonadicity --- Porting note: these local notations were not used in this section before -local notation3 "F" => leftAdjoint G -local notation3 "adj" => Adjunction.ofRightAdjoint G - -- Porting note: added these to replace parametric instances lean4#2311 -- When this is fixed the proofs below that struggle with instances should be reviewed. -- [∀ ⦃A B⦄ (f g : A ⟶ B) [G.IsSplitPair f g], HasCoequalizer f g] @@ -283,9 +275,9 @@ class HasCoequalizerOfIsSplitPair (G : D ⥤ C) : Prop where -- instance {A B} (f g : A ⟶ B) [G.IsSplitPair f g] [HasCoequalizerOfIsSplitPair G] : -- HasCoequalizer f g := HasCoequalizerOfIsSplitPair.out f g -instance [HasCoequalizerOfIsSplitPair G] : ∀ (A : Algebra (toMonad (ofRightAdjoint G))), - HasCoequalizer ((leftAdjoint G).map A.a) - ((ofRightAdjoint G).counit.app ((leftAdjoint G).obj A.A)) := +instance [HasCoequalizerOfIsSplitPair G] : ∀ (A : Algebra adj.toMonad), + HasCoequalizer (F.map A.a) + (adj.counit.app (F.obj A.A)) := fun _ => HasCoequalizerOfIsSplitPair.out G _ _ -- Porting note: added these to replace parametric instances lean4#2311 @@ -296,9 +288,9 @@ class PreservesColimitOfIsSplitPair (G : D ⥤ C) where instance {A B} (f g : A ⟶ B) [G.IsSplitPair f g] [PreservesColimitOfIsSplitPair G] : PreservesColimit (parallelPair f g) G := PreservesColimitOfIsSplitPair.out f g -instance [PreservesColimitOfIsSplitPair G] : ∀ (A : Algebra (toMonad (ofRightAdjoint G))), - PreservesColimit (parallelPair ((leftAdjoint G).map A.a) - (NatTrans.app (Adjunction.ofRightAdjoint G).counit ((leftAdjoint G).obj A.A))) G := +instance [PreservesColimitOfIsSplitPair G] : ∀ (A : Algebra adj.toMonad), + PreservesColimit (parallelPair (F.map A.a) + (NatTrans.app adj.counit (F.obj A.A))) G := fun _ => PreservesColimitOfIsSplitPair.out _ _ -- Porting note: added these to replace parametric instances lean4#2311 @@ -309,45 +301,43 @@ class ReflectsColimitOfIsSplitPair (G : D ⥤ C) where instance {A B} (f g : A ⟶ B) [G.IsSplitPair f g] [ReflectsColimitOfIsSplitPair G] : ReflectsColimit (parallelPair f g) G := ReflectsColimitOfIsSplitPair.out f g -instance [ReflectsColimitOfIsSplitPair G] : ∀ (A : Algebra (toMonad (ofRightAdjoint G))), - ReflectsColimit (parallelPair ((leftAdjoint G).map A.a) - (NatTrans.app (Adjunction.ofRightAdjoint G).counit ((leftAdjoint G).obj A.A))) G := +instance [ReflectsColimitOfIsSplitPair G] : ∀ (A : Algebra adj.toMonad), + ReflectsColimit (parallelPair (F.map A.a) + (NatTrans.app adj.counit (F.obj A.A))) G := fun _ => ReflectsColimitOfIsSplitPair.out _ _ /-- To show `G` is a monadic right adjoint, we can show it preserves and reflects `G`-split coequalizers, and `C` has them. -/ def monadicOfHasPreservesReflectsGSplitCoequalizers [HasCoequalizerOfIsSplitPair G] - [PreservesColimitOfIsSplitPair G] [ReflectsColimitOfIsSplitPair G] : MonadicRightAdjoint G := by - let i : IsRightAdjoint (comparison (ofRightAdjoint G)) := ⟨_, comparisonAdjunction⟩ - constructor - let _ : ∀ X : (ofRightAdjoint G).toMonad.Algebra, - IsIso ((ofRightAdjoint (comparison (ofRightAdjoint G))).unit.app X) := by - intro X - -- Porting note: passing instances through - apply @isIso_of_reflects_iso _ _ _ _ _ _ _ (Monad.forget (ofRightAdjoint G).toMonad) ?_ _ - · change IsIso (comparisonAdjunction.unit.app X).f - rw [comparisonAdjunction_unit_f] - change - IsIso - (IsColimit.coconePointUniqueUpToIso (beckCoequalizer X) - (unitColimitOfPreservesCoequalizer X)).hom - exact IsIso.of_iso (IsColimit.coconePointUniqueUpToIso _ _) - let _ : ∀ Y : D, IsIso ((ofRightAdjoint (comparison (ofRightAdjoint G))).counit.app Y) := by - intro Y - change IsIso (comparisonAdjunction.counit.app Y) - rw [comparisonAdjunction_counit_app] - -- Porting note: passing instances through - change IsIso (IsColimit.coconePointUniqueUpToIso _ ?_).hom - infer_instance - -- Porting note: passing instances through - apply @counitCoequalizerOfReflectsCoequalizer _ _ _ _ _ _ _ ?_ - letI _ : - G.IsSplitPair ((leftAdjoint G).map (G.map ((Adjunction.ofRightAdjoint G).counit.app Y))) - ((Adjunction.ofRightAdjoint G).counit.app ((leftAdjoint G).obj (G.obj Y))) := - MonadicityInternal.main_pair_G_split ((comparison (Adjunction.ofRightAdjoint G)).obj Y) - infer_instance - exact Adjunction.isRightAdjointToIsEquivalence + [PreservesColimitOfIsSplitPair G] [ReflectsColimitOfIsSplitPair G] : + MonadicRightAdjoint G where + adj := adj + eqv := by + have : ∀ (X : Algebra adj.toMonad), IsIso ((comparisonAdjunction adj).unit.app X) := by + intro X + apply @isIso_of_reflects_iso _ _ _ _ _ _ _ (Monad.forget adj.toMonad) ?_ _ + · change IsIso ((comparisonAdjunction adj).unit.app X).f + rw [comparisonAdjunction_unit_f] + change + IsIso + (IsColimit.coconePointUniqueUpToIso (beckCoequalizer X) + (unitColimitOfPreservesCoequalizer X)).hom + exact IsIso.of_iso (IsColimit.coconePointUniqueUpToIso _ _) + have : ∀ (Y : D), IsIso ((comparisonAdjunction adj).counit.app Y) := by + intro Y + rw [comparisonAdjunction_counit_app] + -- Porting note: passing instances through + change IsIso (IsColimit.coconePointUniqueUpToIso _ ?_).hom + infer_instance + -- Porting note: passing instances through + apply @counitCoequalizerOfReflectsCoequalizer _ _ _ _ _ _ _ _ ?_ + letI _ : + G.IsSplitPair (F.map (G.map (adj.counit.app Y))) + (adj.counit.app (F.obj (G.obj Y))) := + MonadicityInternal.main_pair_G_split _ ((comparison adj).obj Y) + infer_instance + exact (comparisonAdjunction adj).toEquivalence.isEquivalence_inverse set_option linter.uppercaseLean3 false in #align category_theory.monad.monadic_of_has_preserves_reflects_G_split_coequalizers CategoryTheory.Monad.monadicOfHasPreservesReflectsGSplitCoequalizers @@ -359,9 +349,9 @@ class CreatesColimitOfIsSplitPair (G : D ⥤ C) where instance {A B} (f g : A ⟶ B) [G.IsSplitPair f g] [CreatesColimitOfIsSplitPair G] : CreatesColimit (parallelPair f g) G := CreatesColimitOfIsSplitPair.out f g -instance [CreatesColimitOfIsSplitPair G] : ∀ (A : Algebra (toMonad (ofRightAdjoint G))), - CreatesColimit (parallelPair ((leftAdjoint G).map A.a) - (NatTrans.app (Adjunction.ofRightAdjoint G).counit ((leftAdjoint G).obj A.A))) G := +instance [CreatesColimitOfIsSplitPair G] : ∀ (A : Algebra adj.toMonad), + CreatesColimit (parallelPair (F.map A.a) + (NatTrans.app adj.counit (F.obj A.A))) G := fun _ => CreatesColimitOfIsSplitPair.out _ _ /-- @@ -374,16 +364,10 @@ def monadicOfCreatesGSplitCoequalizers [CreatesColimitOfIsSplitPair G] : let I {A B} (f g : A ⟶ B) [G.IsSplitPair f g] : HasColimit (parallelPair f g ⋙ G) := by apply @hasColimitOfIso _ _ _ _ _ _ ?_ (diagramIsoParallelPair.{v₁} _) exact inferInstanceAs <| HasCoequalizer (G.map f) (G.map g) - apply @monadicOfHasPreservesReflectsGSplitCoequalizers _ _ _ _ _ _ ?_ ?_ ?_ - · constructor - intros - apply hasColimit_of_created (parallelPair _ _) G - · constructor - intros - infer_instance - · constructor - intros - infer_instance + have : HasCoequalizerOfIsSplitPair G := ⟨fun _ _ => hasColimit_of_created (parallelPair _ _) G⟩ + have : PreservesColimitOfIsSplitPair G := ⟨by intros; infer_instance⟩ + have : ReflectsColimitOfIsSplitPair G := ⟨by intros; infer_instance⟩ + exact monadicOfHasPreservesReflectsGSplitCoequalizers adj set_option linter.uppercaseLean3 false in #align category_theory.monad.monadic_of_creates_G_split_coequalizers CategoryTheory.Monad.monadicOfCreatesGSplitCoequalizers @@ -393,11 +377,10 @@ coequalizers of `G`-split pairs and `C` has coequalizers of `G`-split pairs, the def monadicOfHasPreservesGSplitCoequalizersOfReflectsIsomorphisms [G.ReflectsIsomorphisms] [HasCoequalizerOfIsSplitPair G] [PreservesColimitOfIsSplitPair G] : MonadicRightAdjoint G := by - apply (config := {allowSynthFailures := true}) @monadicOfHasPreservesReflectsGSplitCoequalizers - constructor - intro _ _ f g _ - have := HasCoequalizerOfIsSplitPair.out G f g - apply reflectsColimitOfReflectsIsomorphisms + have : ReflectsColimitOfIsSplitPair G := ⟨fun f g _ => by + have := HasCoequalizerOfIsSplitPair.out G f g + apply reflectsColimitOfReflectsIsomorphisms⟩ + apply monadicOfHasPreservesReflectsGSplitCoequalizers adj set_option linter.uppercaseLean3 false in #align category_theory.monad.monadic_of_has_preserves_G_split_coequalizers_of_reflects_isomorphisms CategoryTheory.Monad.monadicOfHasPreservesGSplitCoequalizersOfReflectsIsomorphisms @@ -407,10 +390,6 @@ section ReflexiveMonadicity variable [HasReflexiveCoequalizers D] [G.ReflectsIsomorphisms] --- Porting note: these local notations were not used in this section before -local notation3 "F" => leftAdjoint G -local notation3 "adj" => Adjunction.ofRightAdjoint G - -- Porting note: added these to replace parametric instances lean4#2311 -- [∀ ⦃A B⦄ (f g : A ⟶ B) [G.IsReflexivePair f g], PreservesColimit (parallelPair f g) G] : class PreservesColimitOfIsReflexivePair (G : C ⥤ D) where @@ -419,9 +398,9 @@ class PreservesColimitOfIsReflexivePair (G : C ⥤ D) where instance {A B} (f g : A ⟶ B) [IsReflexivePair f g] [PreservesColimitOfIsReflexivePair G] : PreservesColimit (parallelPair f g) G := PreservesColimitOfIsReflexivePair.out f g -instance [PreservesColimitOfIsReflexivePair G] : ∀ X : Algebra (toMonad (ofRightAdjoint G)), - PreservesColimit (parallelPair ((leftAdjoint G).map X.a) - (NatTrans.app (Adjunction.ofRightAdjoint G).counit ((leftAdjoint G).obj X.A))) G := +instance [PreservesColimitOfIsReflexivePair G] : ∀ X : Algebra adj.toMonad, + PreservesColimit (parallelPair (F.map X.a) + (NatTrans.app adj.counit (F.obj X.A))) G := fun _ => PreservesColimitOfIsReflexivePair.out _ _ variable [PreservesColimitOfIsReflexivePair G] @@ -429,35 +408,30 @@ variable [PreservesColimitOfIsReflexivePair G] /-- Reflexive (crude) monadicity theorem. If `G` has a right adjoint, `D` has and `G` preserves reflexive coequalizers and `G` reflects isomorphisms, then `G` is monadic. -/ -def monadicOfHasPreservesReflexiveCoequalizersOfReflectsIsomorphisms : MonadicRightAdjoint G := by - letI i : IsRightAdjoint (comparison (Adjunction.ofRightAdjoint G)) := - ⟨_, comparisonAdjunction⟩ - constructor - let _ : ∀ X : (Adjunction.ofRightAdjoint G).toMonad.Algebra, - IsIso ((Adjunction.ofRightAdjoint - (comparison (Adjunction.ofRightAdjoint G))).unit.app X) := by - intro X - apply - @isIso_of_reflects_iso _ _ _ _ _ _ _ (Monad.forget (Adjunction.ofRightAdjoint G).toMonad) ?_ _ - · change IsIso (comparisonAdjunction.unit.app X).f - rw [comparisonAdjunction_unit_f] - change - IsIso - (IsColimit.coconePointUniqueUpToIso (beckCoequalizer X) - (unitColimitOfPreservesCoequalizer X)).hom - apply IsIso.of_iso (IsColimit.coconePointUniqueUpToIso _ _) - let _ : ∀ Y : D, - IsIso ((ofRightAdjoint (comparison (Adjunction.ofRightAdjoint G))).counit.app Y) := by - intro Y - change IsIso (comparisonAdjunction.counit.app Y) - rw [comparisonAdjunction_counit_app] - -- Porting note: passing instances through - change IsIso (IsColimit.coconePointUniqueUpToIso _ ?_).hom - infer_instance - -- Porting note: passing instances through - apply @counitCoequalizerOfReflectsCoequalizer _ _ _ _ _ _ _ ?_ - apply reflectsColimitOfReflectsIsomorphisms - exact Adjunction.isRightAdjointToIsEquivalence +def monadicOfHasPreservesReflexiveCoequalizersOfReflectsIsomorphisms : MonadicRightAdjoint G where + adj := adj + eqv := by + have : ∀ (X : Algebra adj.toMonad), IsIso ((comparisonAdjunction adj).unit.app X) := by + intro X + apply + @isIso_of_reflects_iso _ _ _ _ _ _ _ (Monad.forget adj.toMonad) ?_ _ + · change IsIso ((comparisonAdjunction adj).unit.app X).f + rw [comparisonAdjunction_unit_f] + change + IsIso + (IsColimit.coconePointUniqueUpToIso (beckCoequalizer X) + (unitColimitOfPreservesCoequalizer X)).hom + apply IsIso.of_iso (IsColimit.coconePointUniqueUpToIso _ _) + have : ∀ (Y : D), IsIso ((comparisonAdjunction adj).counit.app Y) := by + intro Y + rw [comparisonAdjunction_counit_app] + -- Porting note: passing instances through + change IsIso (IsColimit.coconePointUniqueUpToIso _ ?_).hom + infer_instance + -- Porting note: passing instances through + apply @counitCoequalizerOfReflectsCoequalizer _ _ _ _ _ _ _ _ ?_ + apply reflectsColimitOfReflectsIsomorphisms + exact (comparisonAdjunction adj).toEquivalence.isEquivalence_inverse #align category_theory.monad.monadic_of_has_preserves_reflexive_coequalizers_of_reflects_isomorphisms CategoryTheory.Monad.monadicOfHasPreservesReflexiveCoequalizersOfReflectsIsomorphisms end ReflexiveMonadicity diff --git a/Mathlib/CategoryTheory/Monoidal/Functor.lean b/Mathlib/CategoryTheory/Monoidal/Functor.lean index a636408026073..2e88bd48841e2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Functor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Functor.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Jendrusch, Scott Morrison, Bhavik Mehta -/ import Mathlib.CategoryTheory.Monoidal.Category -import Mathlib.CategoryTheory.Adjunction.Basic +import Mathlib.CategoryTheory.Adjunction.FullyFaithful import Mathlib.CategoryTheory.Products.Basic #align_import category_theory.monoidal.functor from "leanprover-community/mathlib"@"3d7987cda72abc473c7cdbbb075170e9ac620042" @@ -522,11 +522,22 @@ theorem prod'_toLaxMonoidalFunctor : end MonoidalFunctor +section + +-- TODO: The definitions below would be slightly better phrased if, in addition to +-- `MonoidalFunctor` (which extends `Functor`), we had a data valued type class +-- `Functor.Monoidal` (resp. `Functor.LaxMonoidal`) so that the definitions below +-- could be phrased in terms of `F : C ⥤ D`, `G : D ⥤ D`, `h : F ⊣ G` and `[F.Monoidal]`. +-- Then, in the case of an equivalence (see `monoidalInverse`), we could just take as +-- input an equivalence of categories `e : C ≌ D` and the data `[e.functor.Monoidal]`. + +variable (F : MonoidalFunctor C D) {G : D ⥤ C} (h : F.toFunctor ⊣ G) + /-- If we have a right adjoint functor `G` to a monoidal functor `F`, then `G` has a lax monoidal structure as well. -/ @[simp] -noncomputable def monoidalAdjoint (F : MonoidalFunctor C D) {G : D ⥤ C} (h : F.toFunctor ⊣ G) : +noncomputable def monoidalAdjoint : LaxMonoidalFunctor D C where toFunctor := G ε := h.homEquiv _ _ (inv F.ε) @@ -580,18 +591,23 @@ noncomputable def monoidalAdjoint (F : MonoidalFunctor C D) {G : D ⥤ C} (h : F simp #align category_theory.monoidal_adjoint CategoryTheory.monoidalAdjoint +instance [F.IsEquivalence] : IsIso (monoidalAdjoint F h).ε := by + dsimp + rw [Adjunction.homEquiv_unit] + infer_instance + +instance (X Y : D) [F.IsEquivalence] : IsIso ((monoidalAdjoint F h).μ X Y) := by + dsimp + rw [Adjunction.homEquiv_unit] + infer_instance + /-- If a monoidal functor `F` is an equivalence of categories then its inverse is also monoidal. -/ @[simps] -noncomputable def monoidalInverse (F : MonoidalFunctor C D) [F.IsEquivalence] : - MonoidalFunctor D C - where - toLaxMonoidalFunctor := monoidalAdjoint F (asEquivalence _).toAdjunction - ε_isIso := by - dsimp [Equivalence.toAdjunction] - infer_instance - μ_isIso X Y := by - dsimp [Equivalence.toAdjunction] - infer_instance +noncomputable def monoidalInverse [F.IsEquivalence] : + MonoidalFunctor D C where + toLaxMonoidalFunctor := monoidalAdjoint F h #align category_theory.monoidal_inverse CategoryTheory.monoidalInverse +end + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 772ceb6a7a33a..9d35e330a0163 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -3,6 +3,7 @@ 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.Adjunction.FullyFaithful import Mathlib.CategoryTheory.Monoidal.Functor import Mathlib.CategoryTheory.FullSubcategory @@ -177,71 +178,48 @@ instance isIso_of_isIso_app (α : F ⟶ G) [∀ X : C, IsIso (α.app X)] : IsIso end MonoidalNatIso -noncomputable section - -/-- The unit of a monoidal equivalence can be upgraded to a monoidal natural transformation. -/ -@[simps! toNatTrans] -- Porting note: have to manually specify the toNatTrans projection -def monoidalUnit (F : MonoidalFunctor C D) [F.IsEquivalence] : - LaxMonoidalFunctor.id C ⟶ F.toLaxMonoidalFunctor ⊗⋙ (monoidalInverse F).toLaxMonoidalFunctor := - let e := F.toFunctor.asEquivalence - { toNatTrans := e.unit - tensor := fun X Y => by - -- This proof is not pretty; golfing welcome! - dsimp - simp only [Adjunction.homEquiv_unit, Adjunction.homEquiv_naturality_right, - id_comp, assoc] - simp only [← Functor.map_comp, assoc] - erw [e.counit_app_functor, e.counit_app_functor, - F.toLaxMonoidalFunctor.μ_natural, IsIso.inv_hom_id_assoc] - simp only [Functor.IsEquivalence.inv_fun_map] - slice_rhs 2 3 => erw [Iso.hom_inv_id_app] - dsimp - simp only [CategoryTheory.Category.id_comp] - slice_rhs 1 2 => - rw [← tensor_comp, Iso.hom_inv_id_app, Iso.hom_inv_id_app] - dsimp - rw [tensor_id] - simp } +variable (F : MonoidalFunctor C D) {G : D ⥤ C} (h : F.toFunctor ⊣ G) + +/-- The unit of a adjunction can be upgraded to a monoidal natural transformation. -/ +def monoidalUnit : + LaxMonoidalFunctor.id C ⟶ F.toLaxMonoidalFunctor ⊗⋙ (monoidalAdjoint F h) where + toNatTrans := h.unit + tensor X Y := by + dsimp + simp only [id_comp, comp_id, assoc, Adjunction.homEquiv_unit, + ← h.unit_naturality_assoc, ← Functor.map_comp, + F.map_tensor, IsIso.hom_inv_id_assoc, ← tensor_comp_assoc, + Adjunction.left_triangle_components, tensorHom_id, id_whiskerRight, + IsIso.inv_hom_id, map_id] #align category_theory.monoidal_unit CategoryTheory.monoidalUnit -instance (F : MonoidalFunctor C D) [F.IsEquivalence] : IsIso (monoidalUnit F) := - haveI : ∀ X : C, IsIso ((monoidalUnit F).toNatTrans.app X) := by - dsimp; infer_instance - MonoidalNatIso.isIso_of_isIso_app _ - -/-- The counit of a monoidal equivalence can be upgraded to a monoidal natural transformation. -/ -@[simps! toNatTrans] -- Porting note: have to manually specify the toNatTrans projection -def monoidalCounit (F : MonoidalFunctor C D) [F.IsEquivalence] : - (monoidalInverse F).toLaxMonoidalFunctor ⊗⋙ F.toLaxMonoidalFunctor ⟶ LaxMonoidalFunctor.id D := - let e := F.toFunctor.asEquivalence - { toNatTrans := e.counit - unit := by - dsimp - simp only [comp_id, assoc, Functor.map_inv, Functor.map_comp, - NatIso.inv_inv_app, IsIso.inv_comp, Functor.IsEquivalence.fun_inv_map, - Adjunction.homEquiv_unit] - erw [e.counit_app_functor, ← e.functor.map_comp_assoc, Iso.hom_inv_id_app] - dsimp; simp - tensor := fun X Y => by - dsimp - simp only [Adjunction.homEquiv_unit, Adjunction.homEquiv_naturality_right, assoc, - comp_id, Functor.map_comp] - simp only [Functor.IsEquivalence.fun_inv_map] - erw [e.counit_app_functor] - simp only [assoc] - erw [← e.functor.map_comp_assoc] - simp only [CategoryTheory.Iso.inv_hom_id_app, CategoryTheory.Iso.inv_hom_id_app_assoc] - erw [Iso.hom_inv_id_app, CategoryTheory.Functor.map_id] - simp only [id_comp, CategoryTheory.Iso.inv_hom_id_app, - CategoryTheory.IsIso.hom_inv_id_assoc] - erw [comp_id] } +/-- The unit of a adjunction can be upgraded to a monoidal natural transformation. -/ +@[simps] +def monoidalCounit : + (monoidalAdjoint F h) ⊗⋙ F.toLaxMonoidalFunctor ⟶ LaxMonoidalFunctor.id D where + toNatTrans := h.counit + tensor X Y := by + have eq := h.counit_naturality (F.μ (G.obj X) (G.obj Y)) =≫ inv (F.μ _ _) + simp only [assoc, IsIso.hom_inv_id, comp_id] at eq + dsimp + simp only [Adjunction.homEquiv_unit, comp_id, assoc, + map_comp, map_inv, h.counit_naturality, ← eq, + h.left_triangle_components_assoc, + IsIso.inv_hom_id_assoc, IsIso.hom_inv_id_assoc] + unit := by + have eq := h.counit.naturality F.ε + dsimp at eq ⊢ + rw [Adjunction.homEquiv_unit, map_inv, map_comp, assoc, assoc, map_inv, + ← cancel_mono F.ε, assoc, assoc, assoc, ← eq, IsIso.inv_hom_id_assoc, + Adjunction.left_triangle_components, comp_id, id_comp] #align category_theory.monoidal_counit CategoryTheory.monoidalCounit -instance (F : MonoidalFunctor C D) [F.IsEquivalence] : IsIso (monoidalCounit F) := - haveI : ∀ X : D, IsIso ((monoidalCounit F).toNatTrans.app X) := - by dsimp; infer_instance - MonoidalNatIso.isIso_of_isIso_app _ +instance [F.IsEquivalence] : IsIso (monoidalUnit F h) := by + dsimp [monoidalUnit] + infer_instance -end +instance [F.IsEquivalence] : IsIso (monoidalCounit F h) := by + dsimp [monoidalCounit] + infer_instance end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean index 9c8a704d23cf8..e26e17d18f1b1 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Simon Hudon -/ import Mathlib.CategoryTheory.Monoidal.Braided.Basic -import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal #align_import category_theory.monoidal.of_has_finite_products from "leanprover-community/mathlib"@"f153a85a8dc0a96ce9133fed69e34df72f7f191f" @@ -252,4 +252,55 @@ def symmetricOfHasFiniteCoproducts [HasInitial C] [HasBinaryCoproducts C] : end +section + +attribute [local instance] monoidalOfHasFiniteProducts + +variable {C} +variable {D : Type*} [Category D] (F : C ⥤ D) + [HasTerminal C] [HasBinaryProducts C] + [HasTerminal D] [HasBinaryProducts D] + [PreservesLimit (Functor.empty.{0} C) F] + [PreservesLimitsOfShape (Discrete WalkingPair) F] + +/-- Promote a finite products preserving functor to a monoidal functor between +categories equipped with the monoidal category structure given by finite products. -/ +@[simps] +def Functor.toMonoidalFunctorOfHasFiniteProducts : MonoidalFunctor C D where + toFunctor := F + ε := (asIso (terminalComparison F)).inv + μ X Y := (asIso (prodComparison F X Y)).inv + μ_natural_left {X Y} f X' := by simpa using (prodComparison_inv_natural F f (𝟙 X')).symm + μ_natural_right {X Y} X' g := by simpa using (prodComparison_inv_natural F (𝟙 X') g).symm + associativity X Y Z := by + dsimp only [monoidalOfHasFiniteProducts.associator_hom] + rw [← cancel_epi (prod.map (prodComparison F X Y) (𝟙 (F.obj Z)))] + dsimp + simp only [prod.map_map_assoc, IsIso.hom_inv_id, Category.comp_id, prod.map_id_id, + Category.id_comp, prod.lift_map_assoc, IsIso.inv_comp_eq] + rw [← cancel_mono (prodComparison F X (Y ⨯ Z))] + simp only [Category.assoc, IsIso.inv_hom_id, Category.comp_id, prod.comp_lift, + prod.map_fst_assoc, prodComparison_fst, prodComparison_fst_assoc] + ext + · simp [-Functor.map_comp, ← F.map_comp] + · rw [← cancel_mono (prodComparison F Y Z)] + ext + all_goals simp [-Functor.map_comp, ← F.map_comp] + left_unitality Y := by + rw [← cancel_epi (prod.map (terminalComparison F) (𝟙 (F.obj Y)))] + dsimp + simp only [prod.map_map_assoc, IsIso.hom_inv_id, Category.comp_id, prod.map_id_id, + Category.id_comp, IsIso.eq_inv_comp] + erw [prod.map_snd, Category.comp_id, prodComparison_snd] + right_unitality X := by + rw [← cancel_epi (prod.map (𝟙 (F.obj X)) (terminalComparison F))] + dsimp + simp only [prod.map_map_assoc, Category.comp_id, IsIso.hom_inv_id, prod.map_id_id, + Category.id_comp, IsIso.eq_inv_comp] + erw [prod.map_fst, Category.comp_id, prodComparison_fst] + +instance [F.IsEquivalence] : F.toMonoidalFunctorOfHasFiniteProducts.IsEquivalence := by assumption + +end + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index 99ce115a5cb47..cc4402b6d46af 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -388,7 +388,7 @@ convenient to define the internal hom as `Y →ₗ[k] X` rather than `ᘁY ⊗ X naturally isomorphic). -/ def closedOfHasLeftDual (Y : C) [HasLeftDual Y] : Closed Y where - isAdj := ⟨_, tensorLeftAdjunction (ᘁY) Y⟩ + adj := tensorLeftAdjunction (ᘁY) Y #align category_theory.closed_of_has_left_dual CategoryTheory.closedOfHasLeftDual /-- `tensorLeftHomEquiv` commutes with tensoring on the right -/ diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean index b9ff0069151b8..2dff897753da9 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/OfEquivalence.lean @@ -49,44 +49,47 @@ def exactPairingOfFullyFaithful [F.Full] [F.Faithful] (X Y : C) (F.toFunctor.preimage (inv F.ε ≫ η_ _ _ ≫ F.μ _ _)) (by simp) (by simp) #align category_theory.exact_pairing_of_fully_faithful CategoryTheory.exactPairingOfFullyFaithful +variable {F} +variable {G : D ⥤ C} (adj : F.toFunctor ⊣ G) [F.IsEquivalence] + /-- Pull back a left dual along an equivalence. -/ -def hasLeftDualOfEquivalence [F.IsEquivalence] (X : C) [HasLeftDual (F.obj X)] : +def hasLeftDualOfEquivalence (X : C) [HasLeftDual (F.obj X)] : HasLeftDual X where - leftDual := F.toFunctor.inv.obj (ᘁ(F.obj X)) + leftDual := G.obj (ᘁ(F.obj X)) exact := by -- Porting note: in Lean3, `apply exactPairingOfFullyFaithful F _ _` automatically -- created the goals for `ExactPairing` type class refine @exactPairingOfFullyFaithful _ _ _ _ _ _ F _ _ _ _ ?_ - refine @exactPairingCongrLeft _ _ _ _ _ _ ?_ (F.toFunctor.asEquivalence.counitIso.app _) + refine @exactPairingCongrLeft _ _ _ _ _ _ ?_ (adj.toEquivalence.counitIso.app _) dsimp infer_instance #align category_theory.has_left_dual_of_equivalence CategoryTheory.hasLeftDualOfEquivalence /-- Pull back a right dual along an equivalence. -/ -def hasRightDualOfEquivalence [F.IsEquivalence] (X : C) [HasRightDual (F.obj X)] : +def hasRightDualOfEquivalence (X : C) [HasRightDual (F.obj X)] : HasRightDual X where - rightDual := F.toFunctor.inv.obj ((F.obj X)ᘁ) + rightDual := G.obj ((F.obj X)ᘁ) exact := by refine @exactPairingOfFullyFaithful _ _ _ _ _ _ F _ _ _ _ ?_ - refine @exactPairingCongrRight _ _ _ _ _ _ ?_ (F.toFunctor.asEquivalence.counitIso.app _) + refine @exactPairingCongrRight _ _ _ _ _ _ ?_ (adj.toEquivalence.counitIso.app _) dsimp infer_instance #align category_theory.has_right_dual_of_equivalence CategoryTheory.hasRightDualOfEquivalence /-- Pull back a left rigid structure along an equivalence. -/ -def leftRigidCategoryOfEquivalence [F.IsEquivalence] [LeftRigidCategory D] : - LeftRigidCategory C where leftDual X := hasLeftDualOfEquivalence F X +def leftRigidCategoryOfEquivalence [LeftRigidCategory D] : + LeftRigidCategory C where leftDual X := hasLeftDualOfEquivalence adj X #align category_theory.left_rigid_category_of_equivalence CategoryTheory.leftRigidCategoryOfEquivalence /-- Pull back a right rigid structure along an equivalence. -/ -def rightRigidCategoryOfEquivalence [F.IsEquivalence] [RightRigidCategory D] : - RightRigidCategory C where rightDual X := hasRightDualOfEquivalence F X +def rightRigidCategoryOfEquivalence [RightRigidCategory D] : + RightRigidCategory C where rightDual X := hasRightDualOfEquivalence adj X #align category_theory.right_rigid_category_of_equivalence CategoryTheory.rightRigidCategoryOfEquivalence /-- Pull back a rigid structure along an equivalence. -/ -def rigidCategoryOfEquivalence [F.IsEquivalence] [RigidCategory D] : RigidCategory C where - leftDual X := hasLeftDualOfEquivalence F X - rightDual X := hasRightDualOfEquivalence F X +def rigidCategoryOfEquivalence [RigidCategory D] : RigidCategory C where + leftDual X := hasLeftDualOfEquivalence adj X + rightDual X := hasRightDualOfEquivalence adj X #align category_theory.rigid_category_of_equivalence CategoryTheory.rigidCategoryOfEquivalence end CategoryTheory diff --git a/Mathlib/CategoryTheory/Monoidal/Skeleton.lean b/Mathlib/CategoryTheory/Monoidal/Skeleton.lean index b497063356a3a..aef288593c11e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Skeleton.lean +++ b/Mathlib/CategoryTheory/Monoidal/Skeleton.lean @@ -57,7 +57,7 @@ The skeleton of a monoidal category can be viewed as a monoid, where the multipl the tensor product, and satisfies the monoid axioms since it is a skeleton. -/ noncomputable instance instMonoid : Monoid (Skeleton C) := - monoidOfSkeletalMonoidal (skeletonIsSkeleton _).skel + monoidOfSkeletalMonoidal (skeleton_isSkeleton _).skel /-- The skeleton of a braided monoidal category has a braided monoidal structure itself, induced by the equivalence. -/ @@ -71,7 +71,7 @@ multiplication is given by the tensor product, and satisfies the monoid axioms s skeleton. -/ noncomputable instance instCommMonoid [BraidedCategory C] : CommMonoid (Skeleton C) := - commMonoidOfSkeletalBraided (skeletonIsSkeleton _).skel + commMonoidOfSkeletalBraided (skeleton_isSkeleton _).skel end Skeleton diff --git a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean index 339558e52cbd3..23144ec791e4d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Subcategory.lean +++ b/Mathlib/CategoryTheory/Monoidal/Subcategory.lean @@ -217,9 +217,8 @@ variable [ClosedPredicate P] instance fullMonoidalClosedSubcategory : MonoidalClosed (FullSubcategory P) where closed X := - { isAdj := - { right := - FullSubcategory.lift P (fullSubcategoryInclusion P ⋙ ihom X.1) fun Y => prop_ihom X.2 Y.2 + { rightAdj := FullSubcategory.lift P (fullSubcategoryInclusion P ⋙ ihom X.1) + fun Y => prop_ihom X.2 Y.2 adj := Adjunction.mkOfUnitCounit { unit := @@ -229,7 +228,7 @@ instance fullMonoidalClosedSubcategory : MonoidalClosed (FullSubcategory P) wher { app := fun Y => (ihom.ev X.1).app Y.1 naturality := fun Y Z f => ihom.ev_naturality X.1 f } left_triangle := by ext Y; simp [FullSubcategory.comp_def, FullSubcategory.id_def] - right_triangle := by ext Y; simp [FullSubcategory.comp_def, FullSubcategory.id_def] } } } + right_triangle := by ext Y; simp [FullSubcategory.comp_def, FullSubcategory.id_def] } } #align category_theory.monoidal_category.full_monoidal_closed_subcategory CategoryTheory.MonoidalCategory.fullMonoidalClosedSubcategory @[simp] diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index e9d47e988537a..5d47f1ea0c5d7 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -201,18 +201,18 @@ instance instIsEquivalence_fromTransported (e : C ≌ D) : -/ @[simps!] def toTransported (e : C ≌ D) : MonoidalFunctor C (Transported e) := - monoidalInverse (fromTransported e) + monoidalInverse (fromTransported e) e.symm.toAdjunction #align category_theory.monoidal.to_transported CategoryTheory.Monoidal.toTransported instance (e : C ≌ D) : (toTransported e).IsEquivalence := - inferInstanceAs (e.functor.IsEquivalence) + e.isEquivalence_functor /-- The unit isomorphism upgrades to a monoidal isomorphism. -/ @[simps! hom inv] def transportedMonoidalUnitIso (e : C ≌ D) : LaxMonoidalFunctor.id C ≅ (toTransported e).toLaxMonoidalFunctor ⊗⋙ (fromTransported e).toLaxMonoidalFunctor := - asIso (monoidalCounit (fromTransported e)) |>.symm + asIso (monoidalCounit (fromTransported e) e.symm.toAdjunction) |>.symm #align category_theory.monoidal.transported_monoidal_unit_iso CategoryTheory.Monoidal.transportedMonoidalUnitIso /-- The counit isomorphism upgrades to a monoidal isomorphism. -/ @@ -220,7 +220,7 @@ def transportedMonoidalUnitIso (e : C ≌ D) : def transportedMonoidalCounitIso (e : C ≌ D) : (fromTransported e).toLaxMonoidalFunctor ⊗⋙ (toTransported e).toLaxMonoidalFunctor ≅ LaxMonoidalFunctor.id (Transported e) := - asIso (monoidalUnit (fromTransported e)) |>.symm + asIso (monoidalUnit (fromTransported e) e.symm.toAdjunction) |>.symm #align category_theory.monoidal.transported_monoidal_counit_iso CategoryTheory.Monoidal.transportedMonoidalCounitIso end CategoryTheory.Monoidal diff --git a/Mathlib/CategoryTheory/Pi/Basic.lean b/Mathlib/CategoryTheory/Pi/Basic.lean index e68b1572a60dd..1f92ed000fbae 100644 --- a/Mathlib/CategoryTheory/Pi/Basic.lean +++ b/Mathlib/CategoryTheory/Pi/Basic.lean @@ -389,7 +389,7 @@ def pi (E : ∀ i, C i ≌ D i) : (∀ i, C i) ≌ (∀ i, D i) where instance (F : ∀ i, C i ⥤ D i) [∀ i, (F i).IsEquivalence] : (Functor.pi F).IsEquivalence := - Functor.IsEquivalence.ofEquivalence (pi (fun i => (F i).asEquivalence)) + (pi (fun i => (F i).asEquivalence)).isEquivalence_functor end Equivalence diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 1a5ae8552dd52..c94bd9316d3df 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -651,10 +651,11 @@ instance : (equivalenceSingleObjInverse R).EssSurj where ⟨{ ι := X X := fun _ => PUnit.unit }, ⟨eqToIso (by dsimp; cases X; congr)⟩⟩ +instance : (equivalenceSingleObjInverse R).IsEquivalence where + /-- The categorical equivalence between the category of matrices over a ring, and the category of matrices over that ring considered as a single-object category. -/ def equivalenceSingleObj : Mat R ≌ Mat_ (SingleObj Rᵐᵒᵖ) := - haveI := Functor.IsEquivalence.ofFullyFaithfullyEssSurj (equivalenceSingleObjInverse R) (equivalenceSingleObjInverse R).asEquivalence.symm set_option linter.uppercaseLean3 false in #align category_theory.Mat.equivalence_single_obj CategoryTheory.Mat.equivalenceSingleObj diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index 121c50211dd51..edb2b8f73491f 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -452,20 +452,6 @@ instance (i : A) : (shiftFunctor C i).IsEquivalence := by change (shiftEquiv C i).functor.IsEquivalence infer_instance -@[simp] -theorem shiftFunctor_inv (i : A) : (shiftFunctor C i).inv = shiftFunctor C (-i) := - rfl -#align category_theory.shift_functor_inv CategoryTheory.shiftFunctor_inv - -section - -/-- Shifting by `n` is an essentially surjective functor. -/ -instance shiftFunctor_essSurj (i : A) : (shiftFunctor C i).EssSurj := - Equivalence.essSurj_of_equivalence _ -#align category_theory.shift_functor_ess_surj CategoryTheory.shiftFunctor_essSurj - -end - variable {C} /-- Shifting by `i` and then shifting by `-i` is the identity. -/ diff --git a/Mathlib/CategoryTheory/Sites/Adjunction.lean b/Mathlib/CategoryTheory/Sites/Adjunction.lean index 19196e0d987e5..bde19b949c88e 100644 --- a/Mathlib/CategoryTheory/Sites/Adjunction.lean +++ b/Mathlib/CategoryTheory/Sites/Adjunction.lean @@ -98,8 +98,8 @@ def adjunction (adj : G ⊣ F) : composeAndSheafify J G ⊣ sheafCompose J F := set_option linter.uppercaseLean3 false in #align category_theory.Sheaf.adjunction CategoryTheory.Sheaf.adjunction -instance [IsRightAdjoint F] : IsRightAdjoint (sheafCompose J F) := - ⟨_, adjunction J (Adjunction.ofRightAdjoint F)⟩ +instance [F.IsRightAdjoint] : (sheafCompose J F).IsRightAdjoint := + (adjunction J (Adjunction.ofIsRightAdjoint F)).isRightAdjoint section ForgetToType @@ -149,8 +149,8 @@ theorem adjunctionToTypes_counit_app_val {G : Type max v u ⥤ D} (adj : G ⊣ f set_option linter.uppercaseLean3 false in #align category_theory.Sheaf.adjunction_to_types_counit_app_val CategoryTheory.Sheaf.adjunctionToTypes_counit_app_val -instance [IsRightAdjoint (forget D)] : IsRightAdjoint (sheafForget J : Sheaf J D ⥤ _) := - ⟨_, adjunctionToTypes J (Adjunction.ofRightAdjoint (forget D))⟩ +instance [(forget D).IsRightAdjoint ] : (sheafForget J : Sheaf J D ⥤ _).IsRightAdjoint := + (adjunctionToTypes J (Adjunction.ofIsRightAdjoint (forget D))).isRightAdjoint end ForgetToType diff --git a/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean b/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean index 04bc44102fa04..10fb901d30b33 100644 --- a/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean +++ b/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean @@ -640,10 +640,10 @@ noncomputable def plusPlusAdjunction : plusPlusSheaf J D ⊣ sheafToPresheaf J D rw [Category.assoc] } #align category_theory.sheafification_adjunction CategoryTheory.plusPlusAdjunction -noncomputable instance sheafToPresheafIsRightAdjoint : IsRightAdjoint (sheafToPresheaf J D) := - ⟨_, plusPlusAdjunction J D⟩ +instance sheafToPresheaf_isRightAdjoint : (sheafToPresheaf J D).IsRightAdjoint := + (plusPlusAdjunction J D).isRightAdjoint set_option linter.uppercaseLean3 false in -#align category_theory.Sheaf_to_presheaf_is_right_adjoint CategoryTheory.sheafToPresheafIsRightAdjoint +#align category_theory.Sheaf_to_presheaf_is_right_adjoint CategoryTheory.sheafToPresheaf_isRightAdjoint instance presheaf_mono_of_mono {F G : Sheaf J D} (f : F ⟶ G) [Mono f] : Mono f.1 := (sheafToPresheaf J D).map_mono _ diff --git a/Mathlib/CategoryTheory/Sites/LeftExact.lean b/Mathlib/CategoryTheory/Sites/LeftExact.lean index 30c91630569a0..ef34767d78f42 100644 --- a/Mathlib/CategoryTheory/Sites/LeftExact.lean +++ b/Mathlib/CategoryTheory/Sites/LeftExact.lean @@ -258,8 +258,6 @@ instance preservesfiniteLimits_presheafToSheaf [HasFiniteLimits D] : intros infer_instance -instance : HasWeakSheafify J D := ⟨sheafToPresheafIsRightAdjoint J D⟩ - variable (J D) /-- `plusPlusSheaf` is isomorphic to an arbitrary choice of left adjoint. -/ diff --git a/Mathlib/CategoryTheory/Sites/Sheafification.lean b/Mathlib/CategoryTheory/Sites/Sheafification.lean index 75dff87a70b28..2547e2053b259 100644 --- a/Mathlib/CategoryTheory/Sites/Sheafification.lean +++ b/Mathlib/CategoryTheory/Sites/Sheafification.lean @@ -30,7 +30,7 @@ variable (A : Type u₂) [Category.{v₂} A] /-- A proposition saying that the inclusion functor from sheaves to presheaves admits a left adjoint. -/ -abbrev HasWeakSheafify := Nonempty (IsRightAdjoint (sheafToPresheaf J A)) +abbrev HasWeakSheafify : Prop := (sheafToPresheaf J A).IsRightAdjoint /-- `HasSheafify` means that the inclusion functor from sheaves to presheaves admits a left exact @@ -41,42 +41,36 @@ Given a finite limit preserving functor `F : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A` and -/ class HasSheafify : Prop where isRightAdjoint : HasWeakSheafify J A - isLeftExact : letI := isRightAdjoint.some - Nonempty (PreservesFiniteLimits (leftAdjoint (sheafToPresheaf J A))) + isLeftExact : Nonempty (PreservesFiniteLimits ((sheafToPresheaf J A).leftAdjoint)) instance [HasSheafify J A] : HasWeakSheafify J A := HasSheafify.isRightAdjoint -instance [IsRightAdjoint <| sheafToPresheaf J A] : HasWeakSheafify J A := ⟨inferInstance⟩ - noncomputable section -instance [h : HasWeakSheafify J A] : IsRightAdjoint (sheafToPresheaf J A) := h.some - -instance [HasSheafify J A] : PreservesFiniteLimits (leftAdjoint (sheafToPresheaf J A)) := +instance [HasSheafify J A] : PreservesFiniteLimits ((sheafToPresheaf J A).leftAdjoint) := HasSheafify.isLeftExact.some theorem HasSheafify.mk' {F : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A} (adj : F ⊣ sheafToPresheaf J A) [PreservesFiniteLimits F] : HasSheafify J A where - isRightAdjoint := ⟨⟨F, adj⟩⟩ - isLeftExact := - let i : (h : IsRightAdjoint (sheafToPresheaf J A) := ⟨F, adj⟩) → - F ≅ leftAdjoint (sheafToPresheaf J A) := fun _ ↦ - adj.leftAdjointUniq (Adjunction.ofRightAdjoint (sheafToPresheaf J A)) - ⟨⟨fun _ ↦ preservesLimitsOfShapeOfNatIso (i _)⟩⟩ + isRightAdjoint := ⟨F, ⟨adj⟩⟩ + isLeftExact := ⟨by + have : (sheafToPresheaf J A).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ + exact ⟨fun _ _ _ ↦ preservesLimitsOfShapeOfNatIso + (adj.leftAdjointUniq (Adjunction.ofIsRightAdjoint (sheafToPresheaf J A)))⟩⟩ /-- The sheafification functor, left adjoint to the inclusion. -/ def presheafToSheaf [HasWeakSheafify J A] : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A := - leftAdjoint (sheafToPresheaf J A) + (sheafToPresheaf J A).leftAdjoint instance [HasSheafify J A] : PreservesFiniteLimits (presheafToSheaf J A) := HasSheafify.isLeftExact.some /-- The sheafification-inclusion adjunction. -/ def sheafificationAdjunction [HasWeakSheafify J A] : - presheafToSheaf J A ⊣ sheafToPresheaf J A := IsRightAdjoint.adj + presheafToSheaf J A ⊣ sheafToPresheaf J A := Adjunction.ofIsRightAdjoint _ -instance [HasWeakSheafify J A] : IsLeftAdjoint <| presheafToSheaf J A where - adj := sheafificationAdjunction J A +instance [HasWeakSheafify J A] : (presheafToSheaf J A).IsLeftAdjoint := + ⟨_, ⟨sheafificationAdjunction J A⟩⟩ end @@ -141,7 +135,7 @@ theorem isIso_toSheafify {P : Cᵒᵖ ⥤ D} (hP : Presheaf.IsSheaf J P) : IsIso rfl · change (sheafToPresheaf _ _).map _ ≫ _ = _ change _ ≫ (sheafificationAdjunction J D).unit.app ((sheafToPresheaf J D).obj ⟨P, hP⟩) = _ - erw [← inv_counit_map (sheafificationAdjunction J D) (X := ⟨P, hP⟩), comp_inv_eq_id] + erw [← (sheafificationAdjunction J D).inv_counit_map (X := ⟨P, hP⟩), comp_inv_eq_id] /-- If `P` is a sheaf, then `P` is isomorphic to `sheafify J P`. -/ noncomputable def isoSheafify {P : Cᵒᵖ ⥤ D} (hP : Presheaf.IsSheaf J P) : P ≅ sheafify J P := diff --git a/Mathlib/CategoryTheory/Skeletal.lean b/Mathlib/CategoryTheory/Skeletal.lean index bbe372b716fcd..ebad5166289c5 100644 --- a/Mathlib/CategoryTheory/Skeletal.lean +++ b/Mathlib/CategoryTheory/Skeletal.lean @@ -44,11 +44,11 @@ def Skeletal : Prop := /-- `IsSkeletonOf C D F` says that `F : D ⥤ C` exhibits `D` as a skeletal full subcategory of `C`, in particular `F` is a (strong) equivalence and `D` is skeletal. -/ -structure IsSkeletonOf (F : D ⥤ C) where +structure IsSkeletonOf (F : D ⥤ C) : Prop where /-- The category `D` has isomorphic objects equal -/ skel : Skeletal D /-- The functor `F` is an equivalence -/ - eqv : F.IsEquivalence + eqv : F.IsEquivalence := by infer_instance #align category_theory.is_skeleton_of CategoryTheory.IsSkeletonOf attribute [local instance] isIsomorphicSetoid @@ -98,8 +98,7 @@ noncomputable instance : (fromSkeleton C).Faithful := by instance : (fromSkeleton C).EssSurj where mem_essImage X := ⟨Quotient.mk' X, Quotient.mk_out X⟩ -- Porting note: named this instance -noncomputable instance fromSkeleton.isEquivalence : (fromSkeleton C).IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj (fromSkeleton C) +noncomputable instance fromSkeleton.isEquivalence : (fromSkeleton C).IsEquivalence where /-- The equivalence between the skeleton and the category itself. -/ noncomputable def skeletonEquivalence : Skeleton C ≌ C := @@ -113,10 +112,10 @@ theorem skeleton_skeletal : Skeletal (Skeleton C) := by #align category_theory.skeleton_skeletal CategoryTheory.skeleton_skeletal /-- The `skeleton` of `C` given by choice is a skeleton of `C`. -/ -noncomputable def skeletonIsSkeleton : IsSkeletonOf C (Skeleton C) (fromSkeleton C) where +lemma skeleton_isSkeleton : IsSkeletonOf C (Skeleton C) (fromSkeleton C) where skel := skeleton_skeletal C eqv := fromSkeleton.isEquivalence C -#align category_theory.skeleton_is_skeleton CategoryTheory.skeletonIsSkeleton +#align category_theory.skeleton_is_skeleton CategoryTheory.skeleton_isSkeleton section @@ -266,18 +265,19 @@ noncomputable def fromThinSkeleton : ThinSkeleton C ⥤ C where (Nonempty.some (Quotient.mk_out X)).hom ≫ f.le.some ≫ (Nonempty.some (Quotient.mk_out Y)).inv #align category_theory.thin_skeleton.from_thin_skeleton CategoryTheory.ThinSkeleton.fromThinSkeleton -noncomputable instance fromThinSkeletonEquivalence : (fromThinSkeleton C).IsEquivalence where +/-- The equivalence between the thin skeleton and the category itself. -/ +noncomputable def equivalence : ThinSkeleton C ≌ C where + functor := fromThinSkeleton C inverse := toThinSkeleton C counitIso := NatIso.ofComponents fun X => Nonempty.some (Quotient.mk_out X) unitIso := NatIso.ofComponents fun x => Quotient.recOnSubsingleton x fun X => eqToIso (Quotient.sound ⟨(Nonempty.some (Quotient.mk_out X)).symm⟩) -#align category_theory.thin_skeleton.from_thin_skeleton_equivalence CategoryTheory.ThinSkeleton.fromThinSkeletonEquivalence - -/-- The equivalence between the thin skeleton and the category itself. -/ -noncomputable def equivalence : ThinSkeleton C ≌ C := - (fromThinSkeleton C).asEquivalence #align category_theory.thin_skeleton.equivalence CategoryTheory.ThinSkeleton.equivalence +noncomputable instance fromThinSkeleton_isEquivalence : (fromThinSkeleton C).IsEquivalence := + (equivalence C).isEquivalence_functor +#align category_theory.thin_skeleton.from_thin_skeleton_equivalence CategoryTheory.ThinSkeleton.fromThinSkeleton_isEquivalence + variable {C} theorem equiv_of_both_ways {X Y : C} (f : X ⟶ Y) (g : Y ⟶ X) : X ≈ Y := @@ -314,15 +314,14 @@ theorem map_iso_eq {F₁ F₂ : D ⥤ C} (h : F₁ ≅ F₂) : map F₁ = map F #align category_theory.thin_skeleton.map_iso_eq CategoryTheory.ThinSkeleton.map_iso_eq /-- `fromThinSkeleton C` exhibits the thin skeleton as a skeleton. -/ -noncomputable def thinSkeletonIsSkeleton : IsSkeletonOf C (ThinSkeleton C) (fromThinSkeleton C) +lemma thinSkeleton_isSkeleton : IsSkeletonOf C (ThinSkeleton C) (fromThinSkeleton C) where skel := skeletal - eqv := ThinSkeleton.fromThinSkeletonEquivalence C -#align category_theory.thin_skeleton.thin_skeleton_is_skeleton CategoryTheory.ThinSkeleton.thinSkeletonIsSkeleton +#align category_theory.thin_skeleton.thin_skeleton_is_skeleton CategoryTheory.ThinSkeleton.thinSkeleton_isSkeleton -noncomputable instance isSkeletonOfInhabited : +instance isSkeletonOfInhabited : Inhabited (IsSkeletonOf C (ThinSkeleton C) (fromThinSkeleton C)) := - ⟨thinSkeletonIsSkeleton⟩ + ⟨thinSkeleton_isSkeleton⟩ #align category_theory.thin_skeleton.is_skeleton_of_inhabited CategoryTheory.ThinSkeleton.isSkeletonOfInhabited end diff --git a/Mathlib/CategoryTheory/Subobject/MonoOver.lean b/Mathlib/CategoryTheory/Subobject/MonoOver.lean index eebc0b14c2ca7..44b31637e0b13 100644 --- a/Mathlib/CategoryTheory/Subobject/MonoOver.lean +++ b/Mathlib/CategoryTheory/Subobject/MonoOver.lean @@ -403,11 +403,11 @@ def imageForgetAdj : image ⊣ forget X := exact (Over.w k).symm } } #align category_theory.mono_over.image_forget_adj CategoryTheory.MonoOver.imageForgetAdj -instance : IsRightAdjoint (forget X) where - left := image - adj := imageForgetAdj +instance : (forget X).IsRightAdjoint := + ⟨_, ⟨imageForgetAdj⟩⟩ instance reflective : Reflective (forget X) where + adj := imageForgetAdj #align category_theory.mono_over.reflective CategoryTheory.MonoOver.reflective /-- Forgetting that a monomorphism over `X` is a monomorphism, then taking its image, diff --git a/Mathlib/CategoryTheory/UnivLE.lean b/Mathlib/CategoryTheory/UnivLE.lean index f18934ecbe155..e6b8c95764011 100644 --- a/Mathlib/CategoryTheory/UnivLE.lean +++ b/Mathlib/CategoryTheory/UnivLE.lean @@ -34,8 +34,7 @@ theorem UnivLE_iff_essSurj : UnivLE.{max u v, v} ↔ (uliftFunctor.{u, v} : Type v ⥤ Type max u v).EssSurj := ⟨fun _ => inferInstance, fun w => UnivLE.ofEssSurj w⟩ -instance [UnivLE.{max u v, v}] : uliftFunctor.{u, v}.IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj uliftFunctor +instance [UnivLE.{max u v, v}] : uliftFunctor.{u, v}.IsEquivalence where def UnivLE.witness [UnivLE.{max u v, v}] : Type u ⥤ Type v := uliftFunctor.{v, u} ⋙ (uliftFunctor.{u, v}).inv diff --git a/Mathlib/RepresentationTheory/Action/Monoidal.lean b/Mathlib/RepresentationTheory/Action/Monoidal.lean index 1adabe8513107..7fed696495b9e 100644 --- a/Mathlib/RepresentationTheory/Action/Monoidal.lean +++ b/Mathlib/RepresentationTheory/Action/Monoidal.lean @@ -194,17 +194,23 @@ def functorCategoryMonoidalEquivalence : MonoidalFunctor (Action V G) (SingleObj set_option linter.uppercaseLean3 false in #align Action.functor_category_monoidal_equivalence Action.functorCategoryMonoidalEquivalence +/-- Upgrading the functor `(SingleObj G ⥤ V) ⥤ Action V G` to a monoidal functor. -/ +def functorCategoryMonoidalEquivalenceInverse : MonoidalFunctor (SingleObj G ⥤ V) (Action V G) := + Monoidal.toTransported (Action.functorCategoryEquivalence _ _).symm + +/-- The adjunction (which is an equivalence) between the categories +`Action V G` and `(SingleObj G ⥤ V)`. -/ +def functorCategoryMonoidalAdjunction : + (functorCategoryMonoidalEquivalence V G).toFunctor ⊣ + (functorCategoryMonoidalEquivalenceInverse V G).toFunctor := + (Action.functorCategoryEquivalence _ _).toAdjunction + instance : (functorCategoryMonoidalEquivalence V G).IsEquivalence := by change (Action.functorCategoryEquivalence _ _).functor.IsEquivalence; infer_instance @[simp] theorem functorCategoryMonoidalEquivalence.μ_app (A B : Action V G) : - ((functorCategoryMonoidalEquivalence V G).μ A B).app PUnit.unit = 𝟙 _ := by - dsimp only [functorCategoryMonoidalEquivalence] - simp only [Monoidal.fromTransported_toLaxMonoidalFunctor_μ, NatTrans.comp_app] - -- Porting note: Lean3 was able to see through some defeq, as the mathlib3 proof was - -- show (𝟙 A.V ⊗ 𝟙 B.V) ≫ 𝟙 (A.V ⊗ B.V) ≫ (𝟙 A.V ⊗ 𝟙 B.V) = 𝟙 (A.V ⊗ B.V) - -- simp only [monoidal_category.tensor_id, category.comp_id] + ((functorCategoryMonoidalEquivalence V G).μ A B).app PUnit.unit = 𝟙 _ := rfl set_option linter.uppercaseLean3 false in #align Action.functor_category_monoidal_equivalence.μ_app Action.functorCategoryMonoidalEquivalence.μ_app @@ -229,32 +235,18 @@ set_option linter.uppercaseLean3 false in #align Action.functor_category_monoidal_equivalence.ε_app Action.functorCategoryMonoidalEquivalence.ε_app @[simp] -theorem functorCategoryMonoidalEquivalence.inv_counit_app_hom (A : Action V G) : - ((functorCategoryMonoidalEquivalence _ _).inv.adjunction.counit.app A).hom = 𝟙 _ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.inv_counit_app_hom Action.functorCategoryMonoidalEquivalence.inv_counit_app_hom - -@[simp] -theorem functorCategoryMonoidalEquivalence.counit_app (A : SingleObj G ⥤ V) : - ((functorCategoryMonoidalEquivalence _ _).adjunction.counit.app A).app PUnit.unit = 𝟙 _ := - rfl -set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.counit_app Action.functorCategoryMonoidalEquivalence.counit_app - -@[simp] -theorem functorCategoryMonoidalEquivalence.inv_unit_app_app (A : SingleObj G ⥤ V) : - ((functorCategoryMonoidalEquivalence _ _).inv.adjunction.unit.app A).app PUnit.unit = 𝟙 _ := +theorem functorCategoryMonoidalAdjunction.unit_app_hom (A : Action V G) : + ((functorCategoryMonoidalAdjunction _ _).unit.app A).hom = 𝟙 _ := rfl set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.inv_unit_app_app Action.functorCategoryMonoidalEquivalence.inv_unit_app_app +#align Action.functor_category_monoidal_equivalence.inv_counit_app_hom Action.functorCategoryMonoidalAdjunction.unit_app_hom @[simp] -theorem functorCategoryMonoidalEquivalence.unit_app_hom (A : Action V G) : - ((functorCategoryMonoidalEquivalence _ _).adjunction.unit.app A).hom = 𝟙 _ := +theorem functorCategoryMonoidalAdjunction.counit_app_app (A : SingleObj G ⥤ V) : + ((functorCategoryMonoidalAdjunction _ _).counit.app A).app PUnit.unit = 𝟙 _ := rfl set_option linter.uppercaseLean3 false in -#align Action.functor_category_monoidal_equivalence.unit_app_hom Action.functorCategoryMonoidalEquivalence.unit_app_hom +#align Action.functor_category_monoidal_equivalence.inv_unit_app_app Action.functorCategoryMonoidalAdjunction.counit_app_app @[simp] theorem functorCategoryMonoidalEquivalence.functor_map {A B : Action V G} (f : A ⟶ B) : @@ -265,7 +257,8 @@ set_option linter.uppercaseLean3 false in @[simp] theorem functorCategoryMonoidalEquivalence.inverse_map {A B : SingleObj G ⥤ V} (f : A ⟶ B) : - (functorCategoryMonoidalEquivalence _ _).inv.map f = FunctorCategoryEquivalence.inverse.map f := + (functorCategoryMonoidalEquivalenceInverse _ _).map f = + FunctorCategoryEquivalence.inverse.map f := rfl set_option linter.uppercaseLean3 false in #align Action.functor_category_monoidal_equivalence.inverse_map Action.functorCategoryMonoidalEquivalence.inverse_map @@ -277,21 +270,21 @@ instance [RightRigidCategory V] : RightRigidCategory (SingleObj (H : MonCat.{u}) /-- If `V` is right rigid, so is `Action V G`. -/ instance [RightRigidCategory V] : RightRigidCategory (Action V H) := - rightRigidCategoryOfEquivalence (functorCategoryMonoidalEquivalence V _) + rightRigidCategoryOfEquivalence (functorCategoryMonoidalAdjunction V _) instance [LeftRigidCategory V] : LeftRigidCategory (SingleObj (H : MonCat.{u}) ⥤ V) := by change LeftRigidCategory (SingleObj H ⥤ V); infer_instance /-- If `V` is left rigid, so is `Action V G`. -/ instance [LeftRigidCategory V] : LeftRigidCategory (Action V H) := - leftRigidCategoryOfEquivalence (functorCategoryMonoidalEquivalence V _) + leftRigidCategoryOfEquivalence (functorCategoryMonoidalAdjunction V _) instance [RigidCategory V] : RigidCategory (SingleObj (H : MonCat.{u}) ⥤ V) := by change RigidCategory (SingleObj H ⥤ V); infer_instance /-- If `V` is rigid, so is `Action V G`. -/ instance [RigidCategory V] : RigidCategory (Action V H) := - rigidCategoryOfEquivalence (functorCategoryMonoidalEquivalence V _) + rigidCategoryOfEquivalence (functorCategoryMonoidalAdjunction V _) variable {V H} variable (X : Action V H) diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 21580ea3399b9..be968c1e2094f 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -474,15 +474,14 @@ set_option linter.uppercaseLean3 false in #align Rep.hom_equiv_symm_apply_hom Rep.homEquiv_symm_apply_hom instance : MonoidalClosed (Rep k G) where - closed := fun A => - { isAdj := - { right := Rep.ihom A + closed A := + { rightAdj := Rep.ihom A adj := Adjunction.mkOfHomEquiv ( { homEquiv := Rep.homEquiv A homEquiv_naturality_left_symm := fun _ _ => Action.Hom.ext _ _ (TensorProduct.ext' fun _ _ => rfl) homEquiv_naturality_right := fun _ _ => Action.Hom.ext _ _ (LinearMap.ext - fun _ => LinearMap.ext fun _ => rfl) })}} + fun _ => LinearMap.ext fun _ => rfl) })} @[simp] theorem ihom_obj_ρ_def (A B : Rep k G) : ((ihom A).obj B).ρ = ((Rep.ihom A).obj B).ρ := diff --git a/Mathlib/Topology/Category/CompHaus/Basic.lean b/Mathlib/Topology/Category/CompHaus/Basic.lean index 17092e04e5de2..f173db8d8829f 100644 --- a/Mathlib/Topology/Category/CompHaus/Basic.lean +++ b/Mathlib/Topology/Category/CompHaus/Basic.lean @@ -258,7 +258,8 @@ set_option linter.uppercaseLean3 false in /-- The category of compact Hausdorff spaces is reflective in the category of topological spaces. -/ noncomputable instance compHausToTop.reflective : Reflective compHausToTop where - toIsRightAdjoint := ⟨topToCompHaus, Adjunction.adjunctionOfEquivLeft _ _⟩ + L := topToCompHaus + adj := Adjunction.adjunctionOfEquivLeft _ _ set_option linter.uppercaseLean3 false in #align CompHaus_to_Top.reflective compHausToTop.reflective diff --git a/Mathlib/Topology/Category/Compactum.lean b/Mathlib/Topology/Category/Compactum.lean index 9732c060ba053..3dc440d5b043a 100644 --- a/Mathlib/Topology/Category/Compactum.lean +++ b/Mathlib/Topology/Category/Compactum.lean @@ -479,8 +479,7 @@ instance essSurj : compactumToCompHaus.EssSurj := #align Compactum_to_CompHaus.ess_surj compactumToCompHaus.essSurj /-- The functor `compactumToCompHaus` is an equivalence of categories. -/ -noncomputable instance isEquivalence : compactumToCompHaus.IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance isEquivalence : compactumToCompHaus.IsEquivalence where #align Compactum_to_CompHaus.is_equivalence compactumToCompHaus.isEquivalence end compactumToCompHaus diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index fb60ee99da909..cb1cdc22d7274 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -148,8 +148,7 @@ instance : LightProfinite'.toLightFunctor.{u}.EssSurj where Skeleton.equivalence.counitIso)) toProfinite)) ≪≫ (limit.isLimit _).conePointUniqueUpToIso Y.isLimit)⟩⟩ -instance : LightProfinite'.toLightFunctor.IsEquivalence := - Functor.IsEquivalence.ofFullyFaithfullyEssSurj _ +instance : LightProfinite'.toLightFunctor.IsEquivalence where /-- The equivalence beween `LightProfinite` and a small category. -/ def LightProfinite.equivSmall : LightProfinite.{u} ≌ LightProfinite'.{u} := diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index 2f067c84493c2..5ac806cd9b8e5 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -281,7 +281,7 @@ def toProfiniteAdjToCompHaus : CompHaus.toProfinite ⊣ profiniteToCompHaus := /-- The category of profinite sets is reflective in the category of compact Hausdorff spaces -/ instance toCompHaus.reflective : Reflective profiniteToCompHaus where - toIsRightAdjoint := ⟨CompHaus.toProfinite, Profinite.toProfiniteAdjToCompHaus⟩ + adj := Profinite.toProfiniteAdjToCompHaus #align Profinite.to_CompHaus.reflective Profinite.toCompHaus.reflective noncomputable instance toCompHaus.createsLimits : CreatesLimits profiniteToCompHaus := diff --git a/Mathlib/Topology/Category/TopCat/Adjunctions.lean b/Mathlib/Topology/Category/TopCat/Adjunctions.lean index 2028388ec6586..af901101f6b54 100644 --- a/Mathlib/Topology/Category/TopCat/Adjunctions.lean +++ b/Mathlib/Topology/Category/TopCat/Adjunctions.lean @@ -45,10 +45,10 @@ def adj₂ : forget TopCat.{u} ⊣ trivial := set_option linter.uppercaseLean3 false in #align Top.adj₂ TopCat.adj₂ -instance : IsRightAdjoint (forget TopCat.{u}) := - ⟨_, adj₁⟩ +instance : (forget TopCat.{u}).IsRightAdjoint := + ⟨_, ⟨adj₁⟩⟩ -instance : IsLeftAdjoint (forget TopCat.{u}) := - ⟨_, adj₂⟩ +instance : (forget TopCat.{u}).IsLeftAdjoint := + ⟨_, ⟨adj₂⟩⟩ end TopCat diff --git a/Mathlib/Topology/Category/UniformSpace.lean b/Mathlib/Topology/Category/UniformSpace.lean index 0d15631a01ba3..64478175c7338 100644 --- a/Mathlib/Topology/Category/UniformSpace.lean +++ b/Mathlib/Topology/Category/UniformSpace.lean @@ -234,10 +234,8 @@ noncomputable def adj : completionFunctor ⊣ forget₂ CpltSepUniformSpace Unif rfl } #align UniformSpace.adj UniformSpaceCat.adj -noncomputable instance : IsRightAdjoint (forget₂ CpltSepUniformSpace UniformSpaceCat) := - ⟨completionFunctor, adj⟩ - noncomputable instance : Reflective (forget₂ CpltSepUniformSpace UniformSpaceCat) where + adj := adj map_surjective f := ⟨f, rfl⟩ open CategoryTheory.Limits diff --git a/Mathlib/Topology/Sheaves/Functors.lean b/Mathlib/Topology/Sheaves/Functors.lean index 758cba09f034a..05446960d24e5 100644 --- a/Mathlib/Topology/Sheaves/Functors.lean +++ b/Mathlib/Topology/Sheaves/Functors.lean @@ -105,9 +105,8 @@ def pullbackPushforwardAdjunction (f : X ⟶ Y) : pullback A f ⊣ pushforward A f := (Opens.map f).sheafAdjunctionContinuous _ _ _ -instance : IsLeftAdjoint (pullback A f) := ⟨_, pullbackPushforwardAdjunction A f⟩ -instance : IsRightAdjoint (pushforward A f) := ⟨_, pullbackPushforwardAdjunction A f⟩ - +instance : (pullback A f).IsLeftAdjoint := (pullbackPushforwardAdjunction A f).isLeftAdjoint +instance : (pushforward A f).IsRightAdjoint := (pullbackPushforwardAdjunction A f).isRightAdjoint end Sheaf diff --git a/Mathlib/Topology/Sheaves/Skyscraper.lean b/Mathlib/Topology/Sheaves/Skyscraper.lean index 7101f31e12970..f7e5199691e64 100644 --- a/Mathlib/Topology/Sheaves/Skyscraper.lean +++ b/Mathlib/Topology/Sheaves/Skyscraper.lean @@ -387,11 +387,11 @@ def skyscraperPresheafStalkAdjunction [HasColimits C] : rfl #align skyscraper_presheaf_stalk_adjunction skyscraperPresheafStalkAdjunction -instance [HasColimits C] : IsRightAdjoint (skyscraperPresheafFunctor p₀ : C ⥤ Presheaf C X) := - ⟨_, skyscraperPresheafStalkAdjunction _⟩ +instance [HasColimits C] : (skyscraperPresheafFunctor p₀ : C ⥤ Presheaf C X).IsRightAdjoint := + (skyscraperPresheafStalkAdjunction _).isRightAdjoint -instance [HasColimits C] : IsLeftAdjoint (Presheaf.stalkFunctor C p₀) := - ⟨_, skyscraperPresheafStalkAdjunction _⟩ +instance [HasColimits C] : (Presheaf.stalkFunctor C p₀).IsLeftAdjoint := + (skyscraperPresheafStalkAdjunction _).isLeftAdjoint /-- Taking stalks of a sheaf is the left adjoint functor to `skyscraperSheafFunctor` -/ @@ -411,8 +411,8 @@ def stalkSkyscraperSheafAdjunction [HasColimits C] : homEquiv_counit {𝓐} c f := (skyscraperPresheafStalkAdjunction p₀).homEquiv_counit #align stalk_skyscraper_sheaf_adjunction stalkSkyscraperSheafAdjunction -instance [HasColimits C] : IsRightAdjoint (skyscraperSheafFunctor p₀ : C ⥤ Sheaf C X) := - ⟨_, stalkSkyscraperSheafAdjunction _⟩ +instance [HasColimits C] : (skyscraperSheafFunctor p₀ : C ⥤ Sheaf C X).IsRightAdjoint := + (stalkSkyscraperSheafAdjunction _).isRightAdjoint end From ded3fb6262f4fb8850ea4d5fffe6bad7d3d82ef5 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 8 May 2024 11:39:14 +0000 Subject: [PATCH 34/52] =?UTF-8?q?Feat:=20add=20various=20result=20about=20?= =?UTF-8?q?the=20norm,=20relative=20to=20`=E2=84=A4`,=20of=20`=CE=B6=20-?= =?UTF-8?q?=201`.=20(#11792)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We add various result about the norm, relative to `ℤ`, of `ζ - 1`. These results already exist for the norm relative to `ℚ`. From the flt3 project at LFTCM2024. - [x] depends on: #12386 Co-authored-by: Anne Baanen Co-authored-by: Vierkantor Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- .../Cyclotomic/PrimitiveRoots.lean | 106 +++++++++++------- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 96 ++++++++++++++-- .../RingTheory/Localization/NormTrace.lean | 8 ++ 3 files changed, 163 insertions(+), 47 deletions(-) diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 4794f1ce1b74a..953e390f23390 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -38,11 +38,11 @@ details section. * `IsPrimitiveRoot.sub_one_norm_eq_eval_cyclotomic`: if `Irreducible (cyclotomic n K)` (in particular for `K = ℚ`), then the norm of `ζ - 1` is `eval 1 (cyclotomic n ℤ)`, for a primitive root `ζ`. We also prove the analogous of this result for `zeta`. -* `IsPrimitiveRoot.pow_sub_one_norm_prime_pow_ne_two` : if +* `IsPrimitiveRoot.norm_pow_sub_one_of_prime_pow_ne_two` : if `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` `p ^ (k - s + 1) ≠ 2`. See the following lemmas for similar results. We also prove the analogous of this result for `zeta`. -* `IsPrimitiveRoot.sub_one_norm_prime_ne_two` : if `Irreducible (cyclotomic (p ^ (k + 1)) K)` +* `IsPrimitiveRoot.norm_sub_one_of_prime_ne_two` : if `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `ζ - 1` is `p`. We also prove the analogous of this result for `zeta`. * `IsPrimitiveRoot.embeddingsEquivPrimitiveRoots`: the equivalence between `L →ₐ[K] A` @@ -399,7 +399,7 @@ open scoped Cyclotomic /-- If `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `p ^ (k - s + 1) ≠ 2`. See the next lemmas for similar results. -/ -theorem pow_sub_one_norm_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) +theorem norm_pow_sub_one_of_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) [hpri : Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : norm K (ζ ^ (p : ℕ) ^ s - 1) = (p : K) ^ (p : ℕ) ^ s := by @@ -461,15 +461,15 @@ theorem pow_sub_one_norm_prime_pow_ne_two {k s : ℕ} (hζ : IsPrimitiveRoot ζ exact (Nat.sub_add_cancel hs).symm rw [Hex, pow_add] at this exact mul_left_cancel₀ (pow_ne_zero _ hpri.out.ne_zero) this -#align is_primitive_root.pow_sub_one_norm_prime_pow_ne_two IsPrimitiveRoot.pow_sub_one_norm_prime_pow_ne_two +#align is_primitive_root.pow_sub_one_norm_prime_pow_ne_two IsPrimitiveRoot.norm_pow_sub_one_of_prime_pow_ne_two /-- If `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `p ≠ 2`. -/ -theorem pow_sub_one_norm_prime_ne_two {k : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) +theorem norm_pow_sub_one_of_prime_ne_two {k : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) [hpri : Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) {s : ℕ} (hs : s ≤ k) (hodd : p ≠ 2) : norm K (ζ ^ (p : ℕ) ^ s - 1) = (p : K) ^ (p : ℕ) ^ s := by - refine' hζ.pow_sub_one_norm_prime_pow_ne_two hirr hs fun h => _ + refine' hζ.norm_pow_sub_one_of_prime_pow_ne_two hirr hs fun h => _ have coe_two : ((2 : ℕ+) : ℕ) = 2 := by norm_cast rw [← PNat.coe_inj, coe_two, PNat.pow_coe, ← pow_one 2] at h -- Porting note: the proof is slightly different because of coercions. @@ -477,31 +477,31 @@ theorem pow_sub_one_norm_prime_ne_two {k : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p eq_of_prime_pow_eq (prime_iff.1 hpri.out) (prime_iff.1 Nat.prime_two) (k - s).succ_pos h exact hodd (PNat.coe_injective h) -#align is_primitive_root.pow_sub_one_norm_prime_ne_two IsPrimitiveRoot.pow_sub_one_norm_prime_ne_two +#align is_primitive_root.pow_sub_one_norm_prime_ne_two IsPrimitiveRoot.norm_pow_sub_one_of_prime_ne_two /-- If `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `ζ - 1` is `p`. -/ -theorem sub_one_norm_prime_ne_two {k : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) +theorem norm_sub_one_of_prime_ne_two {k : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) [hpri : Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (h : p ≠ 2) : norm K (ζ - 1) = p := by - simpa using hζ.pow_sub_one_norm_prime_ne_two hirr k.zero_le h -#align is_primitive_root.sub_one_norm_prime_ne_two IsPrimitiveRoot.sub_one_norm_prime_ne_two + simpa using hζ.norm_pow_sub_one_of_prime_ne_two hirr k.zero_le h +#align is_primitive_root.sub_one_norm_prime_ne_two IsPrimitiveRoot.norm_sub_one_of_prime_ne_two /-- If `Irreducible (cyclotomic p K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `ζ - 1` is `p`. -/ -theorem sub_one_norm_prime [hpri : Fact (p : ℕ).Prime] [hcyc : IsCyclotomicExtension {p} K L] - (hζ : IsPrimitiveRoot ζ p) (hirr : Irreducible (cyclotomic p K)) (h : p ≠ 2) : - norm K (ζ - 1) = p := by +theorem norm_sub_one_of_prime_ne_two' [hpri : Fact (p : ℕ).Prime] + [hcyc : IsCyclotomicExtension {p} K L] (hζ : IsPrimitiveRoot ζ p) + (hirr : Irreducible (cyclotomic p K)) (h : p ≠ 2) : norm K (ζ - 1) = p := by replace hirr : Irreducible (cyclotomic (p ^ (0 + 1) : ℕ) K) := by simp [hirr] replace hζ : IsPrimitiveRoot ζ (p ^ (0 + 1) : ℕ) := by simp [hζ] haveI : IsCyclotomicExtension {p ^ (0 + 1)} K L := by simp [hcyc] - simpa using sub_one_norm_prime_ne_two hζ hirr h -#align is_primitive_root.sub_one_norm_prime IsPrimitiveRoot.sub_one_norm_prime + simpa using norm_sub_one_of_prime_ne_two hζ hirr h +#align is_primitive_root.sub_one_norm_prime IsPrimitiveRoot.norm_sub_one_of_prime_ne_two' /-- If `Irreducible (cyclotomic (2 ^ (k + 1)) K)` (in particular for `K = ℚ`), then the norm of `ζ ^ (2 ^ k) - 1` is `(-2) ^ (2 ^ k)`. -/ -- Porting note: writing `(2 : ℕ+)` was not needed (similarly everywhere). -theorem pow_sub_one_norm_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ (k + 1))) +theorem norm_pow_sub_one_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ (k + 1))) [IsCyclotomicExtension {(2 : ℕ+) ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (2 ^ (k + 1)) K)) : norm K (ζ ^ 2 ^ k - 1) = (-2 : K) ^ 2 ^ k := by @@ -517,11 +517,11 @@ theorem pow_sub_one_norm_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ (k + 1))) IsCyclotomicExtension.finrank L hirr, pow_coe, show ((2 : ℕ+) : ℕ) = 2 from rfl, totient_prime_pow Nat.prime_two (zero_lt_succ k), succ_sub_succ_eq_sub, tsub_zero] simp -#align is_primitive_root.pow_sub_one_norm_two IsPrimitiveRoot.pow_sub_one_norm_two +#align is_primitive_root.pow_sub_one_norm_two IsPrimitiveRoot.norm_pow_sub_one_two /-- If `Irreducible (cyclotomic (2 ^ k) K)` (in particular for `K = ℚ`) and `k` is at least `2`, then the norm of `ζ - 1` is `2`. -/ -theorem sub_one_norm_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ k)) (hk : 2 ≤ k) +theorem norm_sub_one_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ k)) (hk : 2 ≤ k) [H : IsCyclotomicExtension {(2 : ℕ+) ^ k} K L] (hirr : Irreducible (cyclotomic (2 ^ k) K)) : norm K (ζ - 1) = 2 := by have : 2 < (2 : ℕ+) ^ k := by @@ -535,11 +535,11 @@ theorem sub_one_norm_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ k)) (hk : 2 obtain ⟨k₁, hk₁⟩ := exists_eq_succ_of_ne_zero (lt_of_lt_of_le zero_lt_two hk).ne.symm -- Porting note: the proof is slightly different because of coercions. simpa [hk₁, show ((2 : ℕ+) : ℕ) = 2 from rfl] using sub_one_norm_eq_eval_cyclotomic hζ this hirr -#align is_primitive_root.sub_one_norm_two IsPrimitiveRoot.sub_one_norm_two +#align is_primitive_root.sub_one_norm_two IsPrimitiveRoot.norm_sub_one_two /-- If `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `k ≠ 0` and `s ≤ k`. -/ -theorem pow_sub_one_norm_prime_pow_of_ne_zero {k s : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) +theorem norm_pow_sub_one_eq_prime_pow_of_ne_zero {k s : ℕ} (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) [hpri : Fact (p : ℕ).Prime] [hcycl : IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hs : s ≤ k) (hk : k ≠ 0) : norm K (ζ ^ (p : ℕ) ^ s - 1) = (p : K) ^ (p : ℕ) ^ s := by @@ -559,11 +559,11 @@ theorem pow_sub_one_norm_prime_pow_of_ne_zero {k s : ℕ} (hζ : IsPrimitiveRoot at hζ hirr hcycl ⊢ obtain ⟨k₁, hk₁⟩ := Nat.exists_eq_succ_of_ne_zero hk -- Porting note: the proof is slightly different because of coercions. - rw [hζ.pow_sub_one_norm_two hirr, hk₁, _root_.pow_succ', pow_mul, neg_eq_neg_one_mul, mul_pow, - neg_one_sq, one_mul, ← pow_mul, ← _root_.pow_succ'] + rw [hζ.norm_pow_sub_one_two hirr, hk₁, _root_.pow_succ', pow_mul, neg_eq_neg_one_mul, + mul_pow, neg_one_sq, one_mul, ← pow_mul, ← _root_.pow_succ'] simp - · exact hζ.pow_sub_one_norm_prime_pow_ne_two hirr hs htwo -#align is_primitive_root.pow_sub_one_norm_prime_pow_of_ne_zero IsPrimitiveRoot.pow_sub_one_norm_prime_pow_of_ne_zero + · exact hζ.norm_pow_sub_one_of_prime_pow_ne_two hirr hs htwo +#align is_primitive_root.pow_sub_one_norm_prime_pow_of_ne_zero IsPrimitiveRoot.norm_pow_sub_one_eq_prime_pow_of_ne_zero end Field @@ -584,47 +584,75 @@ theorem norm_zeta_eq_one [IsCyclotomicExtension {n} K L] (hn : n ≠ 2) /-- If `IsPrimePow (n : ℕ)`, `n ≠ 2` and `Irreducible (cyclotomic n K)` (in particular for `K = ℚ`), then the norm of `zeta n K L - 1` is `(n : ℕ).minFac`. -/ -theorem isPrimePow_norm_zeta_sub_one (hn : IsPrimePow (n : ℕ)) [IsCyclotomicExtension {n} K L] +theorem norm_zeta_sub_one_of_isPrimePow (hn : IsPrimePow (n : ℕ)) [IsCyclotomicExtension {n} K L] (hirr : Irreducible (cyclotomic (n : ℕ) K)) (h : n ≠ 2) : norm K (zeta n K L - 1) = (n : ℕ).minFac := (zeta_spec n K L).sub_one_norm_isPrimePow hn hirr h -#align is_cyclotomic_extension.is_prime_pow_norm_zeta_sub_one IsCyclotomicExtension.isPrimePow_norm_zeta_sub_one +#align is_cyclotomic_extension.is_prime_pow_norm_zeta_sub_one IsCyclotomicExtension.norm_zeta_sub_one_of_isPrimePow /-- If `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `(zeta (p ^ (k + 1)) K L) ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `p ^ (k - s + 1) ≠ 2`. -/ -theorem prime_ne_two_pow_norm_zeta_pow_sub_one {k : ℕ} [Fact (p : ℕ).Prime] +theorem norm_zeta_pow_sub_one_of_prime_pow_ne_two {k : ℕ} [Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) {s : ℕ} (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : norm K (zeta (p ^ (k + 1)) K L ^ (p : ℕ) ^ s - 1) = (p : K) ^ (p : ℕ) ^ s := - (zeta_spec _ K L).pow_sub_one_norm_prime_pow_ne_two hirr hs htwo -#align is_cyclotomic_extension.prime_ne_two_pow_norm_zeta_pow_sub_one IsCyclotomicExtension.prime_ne_two_pow_norm_zeta_pow_sub_one + (zeta_spec _ K L).norm_pow_sub_one_of_prime_pow_ne_two hirr hs htwo +#align is_cyclotomic_extension.prime_ne_two_pow_norm_zeta_pow_sub_one IsCyclotomicExtension.norm_zeta_pow_sub_one_of_prime_pow_ne_two /-- If `Irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `zeta (p ^ (k + 1)) K L - 1` is `p`. -/ -theorem prime_ne_two_pow_norm_zeta_sub_one {k : ℕ} [Fact (p : ℕ).Prime] +theorem norm_zeta_pow_sub_one_of_prime_ne_two {k : ℕ} [Fact (p : ℕ).Prime] [IsCyclotomicExtension {p ^ (k + 1)} K L] (hirr : Irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (h : p ≠ 2) : norm K (zeta (p ^ (k + 1)) K L - 1) = p := - (zeta_spec _ K L).sub_one_norm_prime_ne_two hirr h -#align is_cyclotomic_extension.prime_ne_two_pow_norm_zeta_sub_one IsCyclotomicExtension.prime_ne_two_pow_norm_zeta_sub_one + (zeta_spec _ K L).norm_sub_one_of_prime_ne_two hirr h +#align is_cyclotomic_extension.prime_ne_two_pow_norm_zeta_sub_one IsCyclotomicExtension.norm_zeta_pow_sub_one_of_prime_ne_two /-- If `Irreducible (cyclotomic p K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `zeta p K L - 1` is `p`. -/ -theorem prime_ne_two_norm_zeta_sub_one [Fact (p : ℕ).Prime] +theorem norm_zeta_sub_one_of_prime_ne_two [Fact (p : ℕ).Prime] [IsCyclotomicExtension {p} K L] (hirr : Irreducible (cyclotomic p K)) (h : p ≠ 2) : norm K (zeta p K L - 1) = p := - (zeta_spec _ K L).sub_one_norm_prime hirr h -#align is_cyclotomic_extension.prime_ne_two_norm_zeta_sub_one IsCyclotomicExtension.prime_ne_two_norm_zeta_sub_one + (zeta_spec _ K L).norm_sub_one_of_prime_ne_two' hirr h +#align is_cyclotomic_extension.prime_ne_two_norm_zeta_sub_one IsCyclotomicExtension.norm_zeta_sub_one_of_prime_ne_two /-- If `Irreducible (cyclotomic (2 ^ k) K)` (in particular for `K = ℚ`) and `k` is at least `2`, then the norm of `zeta (2 ^ k) K L - 1` is `2`. -/ -theorem two_pow_norm_zeta_sub_one {k : ℕ} (hk : 2 ≤ k) [IsCyclotomicExtension {(2 : ℕ+) ^ k} K L] - (hirr : Irreducible (cyclotomic (2 ^ k) K)) : norm K (zeta ((2 : ℕ+) ^ k) K L - 1) = 2 := - sub_one_norm_two (zeta_spec ((2 : ℕ+) ^ k) K L) hk hirr -#align is_cyclotomic_extension.two_pow_norm_zeta_sub_one IsCyclotomicExtension.two_pow_norm_zeta_sub_one +theorem norm_zeta_pow_sub_one_two {k : ℕ} (hk : 2 ≤ k) + [IsCyclotomicExtension {(2 : ℕ+) ^ k} K L] (hirr : Irreducible (cyclotomic (2 ^ k) K)) : + norm K (zeta ((2 : ℕ+) ^ k) K L - 1) = 2 := + norm_sub_one_two (zeta_spec ((2 : ℕ+) ^ k) K L) hk hirr +#align is_cyclotomic_extension.two_pow_norm_zeta_sub_one IsCyclotomicExtension.norm_zeta_pow_sub_one_two end IsCyclotomicExtension +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_prime_pow_ne_two := + IsPrimitiveRoot.norm_pow_sub_one_of_prime_pow_ne_two +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_prime_ne_two := + IsPrimitiveRoot.norm_pow_sub_one_of_prime_ne_two +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.sub_one_norm_prime_ne_two := + IsPrimitiveRoot.norm_sub_one_of_prime_ne_two +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.sub_one_norm_prime := + IsPrimitiveRoot.norm_sub_one_of_prime_ne_two' +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_two := + IsPrimitiveRoot.norm_pow_sub_one_two +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.sub_one_norm_two := + IsPrimitiveRoot.norm_sub_one_two +@[deprecated (since := "2024-04-02")] alias IsPrimitiveRoot.pow_sub_one_norm_prime_pow_of_ne_zero := + IsPrimitiveRoot.norm_pow_sub_one_eq_prime_pow_of_ne_zero +@[deprecated (since := "2024-04-02")] alias IsCyclotomicExtension.isPrimePow_norm_zeta_sub_one := + IsCyclotomicExtension.norm_zeta_sub_one_of_isPrimePow +@[deprecated (since := "2024-04-02")] + alias IsCyclotomicExtension.prime_ne_two_pow_norm_zeta_pow_sub_one := + IsCyclotomicExtension.norm_zeta_pow_sub_one_of_prime_pow_ne_two +@[deprecated (since := "2024-04-02")] + alias IsCyclotomicExtension.prime_ne_two_pow_norm_zeta_sub_one := + IsCyclotomicExtension.norm_zeta_pow_sub_one_of_prime_ne_two +@[deprecated (since := "2024-04-02")] alias IsCyclotomicExtension.prime_ne_two_norm_zeta_sub_one := + IsCyclotomicExtension.norm_zeta_sub_one_of_prime_ne_two +@[deprecated (since := "2024-04-02")] alias IsCyclotomicExtension.two_pow_norm_zeta_sub_one := + IsCyclotomicExtension.norm_zeta_pow_sub_one_two + end Norm diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 5d08f15987681..d717e92331d86 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -24,7 +24,6 @@ integers of a `p ^ n`-th cyclotomic extension of `ℚ`. of cyclotomic fields. -/ - universe u open Algebra IsCyclotomicExtension Polynomial NumberField @@ -292,7 +291,7 @@ theorem zeta_sub_one_prime_of_ne_two [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] rw [← Algebra.norm_localization (Sₘ := K) ℤ (nonZeroDivisors ℤ)] simp only [PNat.pow_coe, id.map_eq_id, RingHomCompTriple.comp_eq, RingHom.coe_coe, Subalgebra.coe_val, algebraMap_int_eq, map_natCast] - exact hζ.sub_one_norm_prime_ne_two (Polynomial.cyclotomic.irreducible_rat (PNat.pos _)) hodd + exact hζ.norm_sub_one_of_prime_ne_two (Polynomial.cyclotomic.irreducible_rat (PNat.pos _)) hodd /-- `ζ - 1` is prime if `ζ` is a primitive `2 ^ (k + 1)`-th root of unity. See `zeta_sub_one_prime` for a general statement. -/ @@ -312,13 +311,16 @@ theorem zeta_sub_one_prime_of_two_pow [IsCyclotomicExtension {(2 : ℕ+) ^ (k + rw [← Algebra.norm_localization (Sₘ := K) ℤ (nonZeroDivisors ℤ)] simp only [Nat.zero_eq, PNat.pow_coe, id.map_eq_id, RingHomCompTriple.comp_eq, RingHom.coe_coe, Subalgebra.coe_val, algebraMap_int_eq, map_neg, map_ofNat] - simpa using hζ.pow_sub_one_norm_two (cyclotomic.irreducible_rat (by simp)) + simpa only [zero_add, pow_one, AddSubgroupClass.coe_sub, OneMemClass.coe_one, Nat.zero_eq, + pow_zero] + using hζ.norm_pow_sub_one_two (cyclotomic.irreducible_rat + (by simp only [Nat.zero_eq, zero_add, pow_one, Nat.ofNat_pos])) convert Int.prime_two apply RingHom.injective_int (algebraMap ℤ ℚ) rw [← Algebra.norm_localization (Sₘ := K) ℤ (nonZeroDivisors ℤ)] simp only [PNat.pow_coe, id.map_eq_id, RingHomCompTriple.comp_eq, RingHom.coe_coe, Subalgebra.coe_val, algebraMap_int_eq, map_natCast] - exact hζ.sub_one_norm_two Nat.AtLeastTwo.prop (cyclotomic.irreducible_rat (by simp)) + exact hζ.norm_sub_one_two Nat.AtLeastTwo.prop (cyclotomic.irreducible_rat (by simp)) /-- `ζ - 1` is prime if `ζ` is a primitive `p ^ (k + 1)`-th root of unity. -/ theorem zeta_sub_one_prime [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] @@ -331,16 +333,94 @@ theorem zeta_sub_one_prime [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] /-- `ζ - 1` is prime if `ζ` is a primitive `p`-th root of unity. -/ theorem zeta_sub_one_prime' [h : IsCyclotomicExtension {p} ℚ K] (hζ : IsPrimitiveRoot ζ p) : Prime ((hζ.toInteger - 1)) := by - convert zeta_sub_one_prime (k := 0) (by simpa) - simpa + convert zeta_sub_one_prime (k := 0) (by simpa only [zero_add, pow_one]) + simpa only [zero_add, pow_one] theorem subOneIntegralPowerBasis_gen_prime [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) : - Prime hζ.subOneIntegralPowerBasis.gen := by simpa using hζ.zeta_sub_one_prime + Prime hζ.subOneIntegralPowerBasis.gen := by + simpa only [subOneIntegralPowerBasis_gen] using hζ.zeta_sub_one_prime theorem subOneIntegralPowerBasis'_gen_prime [IsCyclotomicExtension {p} ℚ K] (hζ : IsPrimitiveRoot ζ ↑p) : - Prime hζ.subOneIntegralPowerBasis'.gen := by simpa using hζ.zeta_sub_one_prime' + Prime hζ.subOneIntegralPowerBasis'.gen := by + simpa only [subOneIntegralPowerBasis'_gen] using hζ.zeta_sub_one_prime' + +/-- The norm, relative to `ℤ`, of `ζ ^ p ^ s - 1` in a `p ^ (k + 1)`-th cyclotomic extension of `ℚ` +is p ^ p ^ s` if `s ≤ k` and `p ^ (k - s + 1) ≠ 2`. -/ +lemma norm_toInteger_pow_sub_one_of_prime_pow_ne_two [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) {s : ℕ} (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : + Algebra.norm ℤ (hζ.toInteger ^ (p : ℕ) ^ s - 1) = p ^ (p : ℕ) ^ s := by + have : NumberField K := IsCyclotomicExtension.numberField {p ^ (k + 1)} ℚ K + rw [Algebra.norm_eq_iff ℤ (Sₘ := K) (Rₘ := ℚ) rfl.le] + simp [hζ.norm_pow_sub_one_of_prime_pow_ne_two + (cyclotomic.irreducible_rat (by simp only [PNat.pow_coe, gt_iff_lt, PNat.pos, pow_pos])) + hs htwo] + +/-- The norm, relative to `ℤ`, of `ζ ^ 2 ^ k - 1` in a `2 ^ (k + 1)`-th cyclotomic extension of `ℚ` +is `(-2) ^ 2 ^ k`. -/ +lemma norm_toInteger_pow_sub_one_of_two [IsCyclotomicExtension {2 ^ (k + 1)} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑((2 : ℕ+) ^ (k + 1))) : + Algebra.norm ℤ (hζ.toInteger ^ 2 ^ k - 1) = (-2) ^ (2 : ℕ) ^ k := by + have : NumberField K := IsCyclotomicExtension.numberField {2 ^ (k + 1)} ℚ K + rw [Algebra.norm_eq_iff ℤ (Sₘ := K) (Rₘ := ℚ) rfl.le] + simp [hζ.norm_pow_sub_one_two (cyclotomic.irreducible_rat (pow_pos (by decide) _))] + +/-- The norm, relative to `ℤ`, of `ζ ^ p ^ s - 1` in a `p ^ (k + 1)`-th cyclotomic extension of `ℚ` +is `p ^ p ^ s` if `s ≤ k` and `p ≠ 2`. -/ +lemma norm_toInteger_pow_sub_one_of_prime_ne_two [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) {s : ℕ} (hs : s ≤ k) (hodd : p ≠ 2) : + Algebra.norm ℤ (hζ.toInteger ^ (p : ℕ) ^ s - 1) = p ^ (p : ℕ) ^ s := by + refine hζ.norm_toInteger_pow_sub_one_of_prime_pow_ne_two hs (fun h ↦ hodd ?_) + suffices h : (p : ℕ) = 2 from PNat.coe_injective h + apply eq_of_prime_pow_eq hp.out.prime Nat.prime_two.prime (k - s).succ_pos + rw [pow_one] + exact congr_arg Subtype.val h + +/-- The norm, relative to `ℤ`, of `ζ - 1` in a `p ^ (k + 1)`-th cyclotomic extension of `ℚ` is +`p` if `p ≠ 2`. -/ +lemma norm_toInteger_sub_one_of_prime_ne_two [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) (hodd : p ≠ 2) : + Algebra.norm ℤ (hζ.toInteger - 1) = p := by + simpa only [pow_zero, pow_one] using + hζ.norm_toInteger_pow_sub_one_of_prime_ne_two (Nat.zero_le _) hodd + +/-- The norm, relative to `ℤ`, of `ζ - 1` in a `p`-th cyclotomic extension of `ℚ` is `p` if +`p ≠ 2`. -/ +lemma norm_toInteger_sub_one_of_prime_ne_two' [hcycl : IsCyclotomicExtension {p} ℚ K] + (hζ : IsPrimitiveRoot ζ p) (h : p ≠ 2) : Algebra.norm ℤ (hζ.toInteger - 1) = p := by + have : IsCyclotomicExtension {p ^ (0 + 1)} ℚ K := by simpa using hcycl + replace hζ : IsPrimitiveRoot ζ (p ^ (0 + 1)) := by simpa using hζ + exact hζ.norm_toInteger_sub_one_of_prime_ne_two h + +/-- The norm, relative to `ℤ`, of `ζ - 1` in a `p ^ (k + 1)`-th cyclotomic extension of `ℚ` is +a prime if `p ^ (k + 1) ≠ 2`. -/ +lemma prime_norm_toInteger_sub_one_of_prime_pow_ne_two [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) (htwo : p ^ (k + 1) ≠ 2) : + Prime (Algebra.norm ℤ (hζ.toInteger - 1)) := by + have := hζ.norm_toInteger_pow_sub_one_of_prime_pow_ne_two (zero_le _) htwo + simp only [pow_zero, pow_one] at this + rw [this] + exact Nat.prime_iff_prime_int.1 hp.out + +/-- The norm, relative to `ℤ`, of `ζ - 1` in a `p ^ (k + 1)`-th cyclotomic extension of `ℚ` is +a prime if `p ≠ 2`. -/ +lemma prime_norm_toInteger_sub_one_of_prime_ne_two [hcycl : IsCyclotomicExtension {p ^ (k + 1)} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) (hodd : p ≠ 2) : + Prime (Algebra.norm ℤ (hζ.toInteger - 1)) := by + have := hζ.norm_toInteger_sub_one_of_prime_ne_two hodd + simp only [pow_zero, pow_one] at this + rw [this] + exact Nat.prime_iff_prime_int.1 hp.out + +/-- The norm, relative to `ℤ`, of `ζ - 1` in a `p`-th cyclotomic extension of `ℚ` is a prime if +`p ≠ 2`. -/ +lemma prime_norm_toInteger_sub_one_of_prime_ne_two' [hcycl : IsCyclotomicExtension {p} ℚ K] + (hζ : IsPrimitiveRoot ζ ↑p) (hodd : p ≠ 2) : + Prime (Algebra.norm ℤ (hζ.toInteger - 1)) := by + have : IsCyclotomicExtension {p ^ (0 + 1)} ℚ K := by simpa using hcycl + replace hζ : IsPrimitiveRoot ζ (p ^ (0 + 1)) := by simpa using hζ + exact hζ.prime_norm_toInteger_sub_one_of_prime_ne_two hodd end IsPrimitiveRoot diff --git a/Mathlib/RingTheory/Localization/NormTrace.lean b/Mathlib/RingTheory/Localization/NormTrace.lean index 546d51ebff230..3bb4c30768a86 100644 --- a/Mathlib/RingTheory/Localization/NormTrace.lean +++ b/Mathlib/RingTheory/Localization/NormTrace.lean @@ -69,6 +69,14 @@ theorem Algebra.norm_localization [Module.Free R S] [Module.Finite R S] (a : S) Algebra.norm_eq_matrix_det b, RingHom.map_det, ← Algebra.map_leftMulMatrix_localization] #align algebra.norm_localization Algebra.norm_localization +variable {M} in +/-- The norm of `a : S` in `R` can be computed in `Sₘ`. -/ +lemma Algebra.norm_eq_iff [Module.Free R S] [Module.Finite R S] {a : S} {b : R} + (hM : M ≤ nonZeroDivisors R) : Algebra.norm R a = b ↔ + (Algebra.norm Rₘ) ((algebraMap S Sₘ) a) = algebraMap R Rₘ b := + ⟨fun h ↦ h.symm ▸ Algebra.norm_localization _ M _, fun h ↦ + IsLocalization.injective Rₘ hM <| h.symm ▸ (Algebra.norm_localization R M a).symm⟩ + /-- Let `S` be an extension of `R` and `Rₘ Sₘ` be localizations at `M` of `R S` respectively. Then the trace of `a : Sₘ` over `Rₘ` is the trace of `a : S` over `R` if `S` is free as `R`-module. -/ From feb8ef420aa099763cb2a9dd72bf028a175d6878 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 8 May 2024 12:42:11 +0000 Subject: [PATCH 35/52] feat(CategoryTheory): more functoriality for Comma categories (#12206) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces the functor `map : Comma L R ⥤ Comma L' R'` induced by three functors `F₁`, `F₂`, `F` and two natural transformations `F₁ ⋙ L' ⟶ L ⋙ F` and `R ⋙ F ⟶ F₂ ⋙ R'`. Variants for `StructuredArrow` and `CostructuredArrow` are also introduced. All the basic functors between `Comma` categories are special cases of this one, but their definition was not changed: it would break too many defeq. --- Mathlib/CategoryTheory/Comma/Basic.lean | 165 +++++++++++------- .../CategoryTheory/Comma/StructuredArrow.lean | 54 ++++++ Mathlib/CategoryTheory/DiscreteCategory.lean | 5 + 3 files changed, 162 insertions(+), 62 deletions(-) diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index bf05bc59e8639..f4eba4336d152 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -48,12 +48,15 @@ comma, slice, coslice, over, under, arrow namespace CategoryTheory +open Category + -- declare the `v`'s first; see `CategoryTheory.Category` for an explanation universe v₁ v₂ v₃ v₄ v₅ u₁ u₂ u₃ u₄ u₅ variable {A : Type u₁} [Category.{v₁} A] variable {B : Type u₂} [Category.{v₂} B] variable {T : Type u₃} [Category.{v₃} T] +variable {A' B' T' : Type*} [Category A'] [Category B'] [Category T'] /-- The objects of the comma category are triples of an object `left : A`, an object `right : B` and a morphism `hom : L.obj left ⟶ R.obj right`. -/ @@ -196,6 +199,68 @@ def isoMk {X Y : Comma L₁ R₁} (l : X.left ≅ Y.left) (r : X.right ≅ Y.rig simp } #align category_theory.comma.iso_mk CategoryTheory.Comma.isoMk +section + +variable {L R} +variable {L' : A' ⥤ T'} {R' : B' ⥤ T'} + {F₁ : A ⥤ A'} {F₂ : B ⥤ B'} {F : T ⥤ T'} + (α : F₁ ⋙ L' ⟶ L ⋙ F) (β : R ⋙ F ⟶ F₂ ⋙ R') + +/-- The functor `Comma L R ⥤ Comma L' R'` induced by three functors `F₁`, `F₂`, `F` +and two natural transformations `F₁ ⋙ L' ⟶ L ⋙ F` and `R ⋙ F ⟶ F₂ ⋙ R'`. -/ +@[simps] +def map : Comma L R ⥤ Comma L' R' where + obj X := + { left := F₁.obj X.left + right := F₂.obj X.right + hom := α.app X.left ≫ F.map X.hom ≫ β.app X.right } + map {X Y} φ := + { left := F₁.map φ.left + right := F₂.map φ.right + w := by + dsimp + rw [assoc, assoc] + erw [α.naturality_assoc, ← β.naturality] + dsimp + rw [← F.map_comp_assoc, ← F.map_comp_assoc, φ.w] } + +instance faithful_map [F₁.Faithful] [F₂.Faithful] : (map α β).Faithful where + map_injective {X Y} f g h := by + ext + · exact F₁.map_injective (congr_arg CommaMorphism.left h) + · exact F₂.map_injective (congr_arg CommaMorphism.right h) + +instance full_map [F.Faithful] [F₁.Full] [F₂.Full] [IsIso α] [IsIso β] : (map α β).Full where + map_surjective {X Y} φ := + ⟨{ left := F₁.preimage φ.left + right := F₂.preimage φ.right + w := F.map_injective (by + rw [← cancel_mono (β.app _), ← cancel_epi (α.app _), F.map_comp, F.map_comp, + assoc, assoc] + erw [← α.naturality_assoc, β.naturality] + dsimp + rw [F₁.image_preimage, F₂.image_preimage] + simpa using φ.w) }, by aesop_cat⟩ + +instance essSurj_map [F₁.EssSurj] [F₂.EssSurj] [F.Full] [IsIso α] [IsIso β] : + (map α β).EssSurj where + mem_essImage X := + ⟨{ left := F₁.objPreimage X.left + right := F₂.objPreimage X.right + hom := F.preimage ((inv α).app _ ≫ L'.map (F₁.objObjPreimageIso X.left).hom ≫ + X.hom ≫ R'.map (F₂.objObjPreimageIso X.right).inv ≫ (inv β).app _) }, + ⟨isoMk (F₁.objObjPreimageIso X.left) (F₂.objObjPreimageIso X.right) (by + dsimp + simp only [NatIso.isIso_inv_app, Functor.comp_obj, Functor.map_preimage, assoc, + IsIso.inv_hom_id, comp_id, IsIso.hom_inv_id_assoc] + rw [← R'.map_comp, Iso.inv_hom_id, R'.map_id, comp_id])⟩⟩ + +noncomputable instance isEquivalenceMap + [F₁.IsEquivalence] [F₂.IsEquivalence] [F.Faithful] [F.Full] [IsIso α] [IsIso β] : + (map α β).IsEquivalence where + +end + /-- A natural transformation `L₁ ⟶ L₂` induces a functor `Comma L₂ R ⥤ Comma L₁ R`. -/ @[simps] def mapLeft (l : L₁ ⟶ L₂) : Comma L₂ R ⥤ Comma L₁ R where @@ -210,39 +275,25 @@ def mapLeft (l : L₁ ⟶ L₂) : Comma L₂ R ⥤ Comma L₁ R where /-- The functor `Comma L R ⥤ Comma L R` induced by the identity natural transformation on `L` is naturally isomorphic to the identity functor. -/ -@[simps] -def mapLeftId : mapLeft R (𝟙 L) ≅ 𝟭 _ where - hom := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } - inv := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } +@[simps!] +def mapLeftId : mapLeft R (𝟙 L) ≅ 𝟭 _ := + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) #align category_theory.comma.map_left_id CategoryTheory.Comma.mapLeftId /-- The functor `Comma L₁ R ⥤ Comma L₃ R` induced by the composition of two natural transformations `l : L₁ ⟶ L₂` and `l' : L₂ ⟶ L₃` is naturally isomorphic to the composition of the two functors induced by these natural transformations. -/ -@[simps] +@[simps!] def mapLeftComp (l : L₁ ⟶ L₂) (l' : L₂ ⟶ L₃) : - mapLeft R (l ≫ l') ≅ mapLeft R l' ⋙ mapLeft R l where - hom := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } - inv := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } + mapLeft R (l ≫ l') ≅ mapLeft R l' ⋙ mapLeft R l := + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) #align category_theory.comma.map_left_comp CategoryTheory.Comma.mapLeftComp /-- Two equal natural transformations `L₁ ⟶ L₂` yield naturally isomorphic functors `Comma L₁ R ⥤ Comma L₂ R`. -/ @[simps!] def mapLeftEq (l l' : L₁ ⟶ L₂) (h : l = l') : mapLeft R l ≅ mapLeft R l' := - NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _) (by aesop_cat)) (by aesop_cat) + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) /-- A natural isomorphism `L₁ ≅ L₂` induces an equivalence of categories `Comma L₁ R ≌ Comma L₂ R`. -/ @@ -266,39 +317,25 @@ def mapRight (r : R₁ ⟶ R₂) : Comma L R₁ ⥤ Comma L R₂ where /-- The functor `Comma L R ⥤ Comma L R` induced by the identity natural transformation on `R` is naturally isomorphic to the identity functor. -/ -@[simps] -def mapRightId : mapRight L (𝟙 R) ≅ 𝟭 _ where - hom := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } - inv := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } +@[simps!] +def mapRightId : mapRight L (𝟙 R) ≅ 𝟭 _ := + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) #align category_theory.comma.map_right_id CategoryTheory.Comma.mapRightId /-- The functor `Comma L R₁ ⥤ Comma L R₃` induced by the composition of the natural transformations `r : R₁ ⟶ R₂` and `r' : R₂ ⟶ R₃` is naturally isomorphic to the composition of the functors induced by these natural transformations. -/ -@[simps] +@[simps!] def mapRightComp (r : R₁ ⟶ R₂) (r' : R₂ ⟶ R₃) : - mapRight L (r ≫ r') ≅ mapRight L r ⋙ mapRight L r' where - hom := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } - inv := - { app := fun X => - { left := 𝟙 _ - right := 𝟙 _ } } + mapRight L (r ≫ r') ≅ mapRight L r ⋙ mapRight L r' := + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) #align category_theory.comma.map_right_comp CategoryTheory.Comma.mapRightComp /-- Two equal natural transformations `R₁ ⟶ R₂` yield naturally isomorphic functors `Comma L R₁ ⥤ Comma L R₂`. -/ @[simps!] def mapRightEq (r r' : R₁ ⟶ R₂) (h : r = r') : mapRight L r ≅ mapRight L r' := - NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _) (by aesop_cat)) (by aesop_cat) + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) /-- A natural isomorphism `R₁ ≅ R₂` induces an equivalence of categories `Comma L R₁ ≌ Comma L R₂`. -/ @@ -327,18 +364,20 @@ def preLeft (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) : Comma (F ⋙ L) R ⥤ Co w := by simpa using f.w } #align category_theory.comma.pre_left CategoryTheory.Comma.preLeft -instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.Faithful] : (preLeft F L R).Faithful where - map_injective {X Y} f g h := hom_ext _ _ (F.map_injective (congrArg CommaMorphism.left h)) - (by apply congrArg CommaMorphism.right h) +/-- `Comma.preLeft` is a particular case of `Comma.map`, +but with better definitional properties. -/ +def preLeftIso (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) : + preLeft F L R ≅ map (F ⋙ L).rightUnitor.inv (R.rightUnitor.hom ≫ R.leftUnitor.inv) := + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) + +instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.Faithful] : (preLeft F L R).Faithful := + Functor.Faithful.of_iso (preLeftIso F L R).symm -instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.Full] : (preLeft F L R).Full where - map_surjective {X Y} f := - ⟨CommaMorphism.mk (F.preimage f.left) f.right (by simpa using f.w), by aesop_cat⟩ +instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.Full] : (preLeft F L R).Full := + Functor.Full.of_iso (preLeftIso F L R).symm -instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.EssSurj] : (preLeft F L R).EssSurj where - mem_essImage Y := - ⟨Comma.mk (F.objPreimage Y.left) Y.right ((L.mapIso (F.objObjPreimageIso _)).hom ≫ Y.hom), - ⟨isoMk (F.objObjPreimageIso _) (Iso.refl _) (by simp)⟩⟩ +instance (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.EssSurj] : (preLeft F L R).EssSurj := + Functor.essSurj_of_iso (preLeftIso F L R).symm /-- If `F` is an equivalence, then so is `preLeft F L R`. -/ instance isEquivalence_preLeft (F : C ⥤ A) (L : A ⥤ T) (R : B ⥤ T) [F.IsEquivalence] : @@ -356,18 +395,20 @@ def preRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) : Comma L (F ⋙ R) ⥤ C right := F.map f.right } #align category_theory.comma.pre_right CategoryTheory.Comma.preRight -instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.Faithful] : (preRight L F R).Faithful where - map_injective {X Y } f g h := hom_ext _ _ (by apply congrArg CommaMorphism.left h) - (F.map_injective (congrArg CommaMorphism.right h)) +/-- `Comma.preRight` is a particular case of `Comma.map`, +but with better definitional properties. -/ +def preRightIso (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) : + preRight L F R ≅ map (L.leftUnitor.hom ≫ L.rightUnitor.inv) (F ⋙ R).rightUnitor.hom := + NatIso.ofComponents (fun X => isoMk (Iso.refl _) (Iso.refl _)) + +instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.Faithful] : (preRight L F R).Faithful := + Functor.Faithful.of_iso (preRightIso L F R).symm -instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.Full] : (preRight L F R).Full where - map_surjective {X Y} f := - ⟨CommaMorphism.mk f.left (F.preimage f.right) (by simpa using f.w), by aesop_cat⟩ +instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.Full] : (preRight L F R).Full := + Functor.Full.of_iso (preRightIso L F R).symm -instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.EssSurj] : (preRight L F R).EssSurj where - mem_essImage Y := - ⟨Comma.mk Y.left (F.objPreimage Y.right) (Y.hom ≫ (R.mapIso (F.objObjPreimageIso _)).inv), - ⟨isoMk (Iso.refl _) (F.objObjPreimageIso _) (by simp [← R.map_comp])⟩⟩ +instance (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.EssSurj] : (preRight L F R).EssSurj := + Functor.essSurj_of_iso (preRightIso L F R).symm /-- If `F` is an equivalence, then so is `preRight L F R`. -/ instance isEquivalence_preRight (L : A ⥤ T) (F : C ⥤ B) (R : B ⥤ T) [F.IsEquivalence] : diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean index 73b62fed6932e..8fd6a313f1491 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean @@ -327,6 +327,33 @@ instance (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] : (post S F G).EssSurj whe instance isEquivalence_post (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] [G.Faithful] : (post S F G).IsEquivalence where +section + +variable {L : D} {R : C ⥤ D} {L' : B} {R' : A ⥤ B} {F : C ⥤ A} {G : D ⥤ B} + (α : L' ⟶ G.obj L) (β : R ⋙ G ⟶ F ⋙ R') + +/-- The functor `StructuredArrow L R ⥤ StructuredArrow L' R'` that is deduced from +a natural transformation `R ⋙ G ⟶ F ⋙ R'` and a morphism `L' ⟶ G.obj L.` -/ +@[simps!] +def map₂ : StructuredArrow L R ⥤ StructuredArrow L' R' := + Comma.map (F₁ := 𝟭 (Discrete PUnit)) (Discrete.natTrans (fun _ => α)) β + +instance faithful_map₂ [F.Faithful] : (map₂ α β).Faithful := by + apply Comma.faithful_map + +instance full_map₂ [G.Faithful] [F.Full] [IsIso α] [IsIso β] : (map₂ α β).Full := by + apply Comma.full_map + +instance essSurj_map₂ [F.EssSurj] [G.Full] [IsIso α] [IsIso β] : (map₂ α β).EssSurj := by + apply Comma.essSurj_map + +noncomputable instance isEquivalenceMap₂ + [F.IsEquivalence] [G.Faithful] [G.Full] [IsIso α] [IsIso β] : + (map₂ α β).IsEquivalence := by + apply Comma.isEquivalenceMap + +end + instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by suffices (proj S T).obj ⁻¹' 𝒢 = Set.range fun f : ΣG : 𝒢, S ⟶ T.obj G => mk f.2 by @@ -664,6 +691,33 @@ instance (F : B ⥤ C) (G : C ⥤ D) (S : C) [G.Full] : (post F G S).EssSurj whe instance isEquivalence_post (S : C) (F : B ⥤ C) (G : C ⥤ D) [G.Full] [G.Faithful] : (post F G S).IsEquivalence where +section + +variable {U : A ⥤ B} {V : B} {F : C ⥤ A} {G : D ⥤ B} + (α : F ⋙ U ⟶ S ⋙ G) (β : G.obj T ⟶ V) + +/-- The functor `CostructuredArrow S T ⥤ CostructuredArrow U V` that is deduced from +a natural transformation `F ⋙ U ⟶ S ⋙ G` and a morphism `G.obj T ⟶ V` -/ +@[simps!] +def map₂ : CostructuredArrow S T ⥤ CostructuredArrow U V := + Comma.map (F₂ := 𝟭 (Discrete PUnit)) α (Discrete.natTrans (fun _ => β)) + +instance faithful_map₂ [F.Faithful] : (map₂ α β).Faithful := by + apply Comma.faithful_map + +instance full_map₂ [G.Faithful] [F.Full] [IsIso α] [IsIso β] : (map₂ α β).Full := by + apply Comma.full_map + +instance essSurj_map₂ [F.EssSurj] [G.Full] [IsIso α] [IsIso β] : (map₂ α β).EssSurj := by + apply Comma.essSurj_map + +noncomputable instance isEquivalenceMap₂ + [F.IsEquivalence] [G.Faithful] [G.Full] [IsIso α] [IsIso β] : + (map₂ α β).IsEquivalence := by + apply Comma.isEquivalenceMap + +end + instance small_proj_preimage_of_locallySmall {𝒢 : Set C} [Small.{v₁} 𝒢] [LocallySmall.{v₁} D] : Small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := by suffices (proj S T).obj ⁻¹' 𝒢 = Set.range fun f : ΣG : 𝒢, S.obj G ⟶ T => mk f.2 by diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index 7f6838e38fd5b..f73e631a5a83b 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -225,6 +225,11 @@ def natIso {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.o simp #align category_theory.discrete.nat_iso CategoryTheory.Discrete.natIso +instance {I : Type*} {F G : Discrete I ⥤ C} (f : ∀ i, F.obj i ⟶ G.obj i) [∀ i, IsIso (f i)] : + IsIso (Discrete.natTrans f) := by + change IsIso (Discrete.natIso (fun i => asIso (f i))).hom + infer_instance + @[simp] theorem natIso_app {I : Type u₁} {F G : Discrete I ⥤ C} (f : ∀ i : Discrete I, F.obj i ≅ G.obj i) (i : Discrete I) : (Discrete.natIso f).app i = f i := by aesop_cat From 6d1b4e774edf08348d000edaf2facef71c183c90 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 8 May 2024 13:10:27 +0000 Subject: [PATCH 36/52] feat: iterated derivative of a continuous multilinear map (#12207) Consider a multilinear map `f : E^n -> F`. Then its `k`-th derivative can be written as follows, given a basepoint `x : E^n` and `k` vectors `v_1, ..., v_k : E^n`: ``` D^k f (x; v_1, ..., v_k) = \sum f (x_1, (v_{e.symm 2})_2, x_3, ..., (v_{e.symm n))_n) ``` where in the sum at each index `i` one uses either the `i`-th coordinate of `x`, or the `i`-th coordinate of one of the vectors `v_j`, and each `v_j` has to appear exactly once in this expression. The sum is parameterized by the embeddings `e` of `Fin k` in `Fin n`. This expression is a little bit cumbersome to write down and to use, but it gives nice explicit bounds on `||D^k f||` that are useful to study Fourier transforms. The PR introduces the formal definition of each term in the sum (as an object which is multilinear in the coordinates of `x` outside the range of `e`, and in the `v_j`), and then proves that the sum is indeed the iterated derivative of `f`. --- Mathlib/Analysis/Analytic/Basic.lean | 2 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 86 ++++++++++++++++++- .../NormedSpace/Multilinear/Basic.lean | 82 +++++++++++++++++- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 71 +++++++++++++++ 4 files changed, 237 insertions(+), 4 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 2072ddabe699e..44e5ae4cb73be 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -1084,7 +1084,7 @@ $$ The corresponding power series has thus a `k`-th coefficient equal to $\sum_{n} \binom{n}{k} p_n y^{n-k}$. In the general case where `pₙ` is a multilinear map, this has to be interpreted suitably: instead of having a binomial coefficient, one should sum over all -possible subsets `s` of `Fin n` of cardinal `k`, and attribute `z` to the indices in `s` and +possible subsets `s` of `Fin n` of cardinality `k`, and attribute `z` to the indices in `s` and `y` to the indices outside of `s`. In this paragraph, we implement this. The new power series is called `p.changeOrigin y`. Then, we diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 39f368a34c54e..600f189be33d6 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -7,6 +7,7 @@ import Mathlib.Analysis.Analytic.Basic import Mathlib.Analysis.Analytic.CPolynomial import Mathlib.Analysis.Calculus.Deriv.Basic import Mathlib.Analysis.Calculus.ContDiff.Defs +import Mathlib.Analysis.Calculus.FDeriv.Add #align_import analysis.calculus.fderiv_analytic from "leanprover-community/mathlib"@"3bce8d800a6f2b8f63fe1e588fd76a9ff4adcebe" @@ -15,12 +16,14 @@ import Mathlib.Analysis.Calculus.ContDiff.Defs A function expressible as a power series at a point has a Frechet derivative there. Also the special case in terms of `deriv` when the domain is 1-dimensional. --/ +As an application, we show that continuous multilinear maps are smooth. We also compute their +iterated derivatives, in `ContinuousMultilinearMap.iteratedFDeriv_eq`. +-/ open Filter Asymptotics -open scoped ENNReal +open scoped ENNReal BigOperators universe u v @@ -347,6 +350,85 @@ protected theorem hasFDerivAt [DecidableEq ι] : HasFDerivAt f (f.linearDeriv x) convert f.hasFiniteFPowerSeriesOnBall.hasFDerivAt (y := x) ENNReal.coe_lt_top rw [zero_add] +/-- Technical lemma used in the proof of `hasFTaylorSeriesUpTo_iteratedFDeriv`, to compare sums +over embedding of `Fin k` and `Fin (k + 1)`. -/ +private lemma _root_.Equiv.succ_embeddingFinSucc_fst_symm_apply {ι : Type*} [DecidableEq ι] + {n : ℕ} (e : Fin (n+1) ↪ ι) {k : ι} + (h'k : k ∈ Set.range (Equiv.embeddingFinSucc n ι e).1) (hk : k ∈ Set.range e) : + Fin.succ ((Equiv.embeddingFinSucc n ι e).1.toEquivRange.symm ⟨k, h'k⟩) + = e.toEquivRange.symm ⟨k, hk⟩ := by + rcases hk with ⟨j, rfl⟩ + have hj : j ≠ 0 := by + rintro rfl + simp at h'k + simp only [Function.Embedding.toEquivRange_symm_apply_self] + have : e j = (Equiv.embeddingFinSucc n ι e).1 (Fin.pred j hj) := by simp + simp_rw [this] + simp [-Equiv.embeddingFinSucc_fst] + +/-- A continuous multilinear function `f` admits a Taylor series, whose successive terms are given +by `f.iteratedFDeriv n`. This is the point of the definition of `f.iteratedFDeriv`. -/ +theorem hasFTaylorSeriesUpTo_iteratedFDeriv : + HasFTaylorSeriesUpTo ⊤ f (fun v n ↦ f.iteratedFDeriv n v) := by + classical + constructor + · simp [ContinuousMultilinearMap.iteratedFDeriv] + · rintro n - x + suffices H : curryLeft (f.iteratedFDeriv (Nat.succ n) x) = (∑ e : Fin n ↪ ι, + ((iteratedFDerivComponent f e.toEquivRange).linearDeriv + (Pi.compRightL 𝕜 _ Subtype.val x)) ∘L (Pi.compRightL 𝕜 _ Subtype.val)) by + have A : HasFDerivAt (f.iteratedFDeriv n) (∑ e : Fin n ↪ ι, + ((iteratedFDerivComponent f e.toEquivRange).linearDeriv (Pi.compRightL 𝕜 _ Subtype.val x)) + ∘L (Pi.compRightL 𝕜 _ Subtype.val)) x := by + apply HasFDerivAt.sum (fun s _hs ↦ ?_) + exact (ContinuousMultilinearMap.hasFDerivAt _ _).comp x (ContinuousLinearMap.hasFDerivAt _) + rwa [← H] at A + ext v m + simp only [ContinuousMultilinearMap.iteratedFDeriv, curryLeft_apply, sum_apply, + iteratedFDerivComponent_apply, Finset.univ_sigma_univ, + Pi.compRightL_apply, ContinuousLinearMap.coe_sum', ContinuousLinearMap.coe_comp', + Finset.sum_apply, Function.comp_apply, linearDeriv_apply, Finset.sum_sigma'] + rw [← (Equiv.embeddingFinSucc n ι).sum_comp] + congr with e + congr with k + by_cases hke : k ∈ Set.range e + · simp only [hke, ↓reduceDite] + split_ifs with hkf + · simp only [← Equiv.succ_embeddingFinSucc_fst_symm_apply e hkf hke, Fin.cons_succ] + · obtain rfl : k = e 0 := by + rcases hke with ⟨j, rfl⟩ + simpa using hkf + simp only [Function.Embedding.toEquivRange_symm_apply_self, Fin.cons_zero, Function.update, + Pi.compRightL_apply] + split_ifs with h + · congr! + · exfalso + apply h + simp_rw [← Equiv.embeddingFinSucc_snd e] + · have hkf : k ∉ Set.range (Equiv.embeddingFinSucc n ι e).1 := by + contrapose! hke + rw [Equiv.embeddingFinSucc_fst] at hke + exact Set.range_comp_subset_range _ _ hke + simp only [hke, hkf, ↓reduceDite, Pi.compRightL, + ContinuousLinearMap.coe_mk', LinearMap.coe_mk, AddHom.coe_mk] + rw [Function.update_noteq] + contrapose! hke + rw [show k = _ from Subtype.ext_iff_val.1 hke, Equiv.embeddingFinSucc_snd e] + exact Set.mem_range_self _ + · rintro n - + apply continuous_finset_sum _ (fun e _ ↦ ?_) + exact (ContinuousMultilinearMap.coe_continuous _).comp (ContinuousLinearMap.continuous _) + +theorem iteratedFDeriv_eq (n : ℕ) : + iteratedFDeriv 𝕜 n f = f.iteratedFDeriv n := + funext fun x ↦ (f.hasFTaylorSeriesUpTo_iteratedFDeriv.eq_iteratedFDeriv (m := n) le_top x).symm + +theorem norm_iteratedFDeriv_le (n : ℕ) (x : (i : ι) → E i) : + ‖iteratedFDeriv 𝕜 n f x‖ + ≤ Nat.descFactorial (Fintype.card ι) n * ‖f‖ * ‖x‖ ^ (Fintype.card ι - n) := by + rw [f.iteratedFDeriv_eq] + exact f.norm_iteratedFDeriv_le' n x + lemma cPolynomialAt : CPolynomialAt 𝕜 f x := f.hasFiniteFPowerSeriesOnBall.cPolynomialAt_of_mem (by simp only [Metric.emetric_ball_top, Set.mem_univ]) diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index 0dc7f02aed35f..6d55de83b5bf4 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Sophie Morel, Yury Kudryashov -/ import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace +import Mathlib.Logic.Embedding.Basic +import Mathlib.Data.Fintype.CardEmbedding import Mathlib.Topology.Algebra.Module.Multilinear.Topology #align_import analysis.normed_space.multilinear from "leanprover-community/mathlib"@"f40476639bac089693a489c9e354ebd75dc0f886" @@ -72,7 +74,6 @@ We use the following type variables in this file: * `G`, `G'` : normed vector spaces over `𝕜`. -/ - universe u v v' wE wE₁ wE' wEi wG wG' section Seminorm @@ -1306,6 +1307,85 @@ theorem compContinuousLinearMapEquivL_apply (g : ContinuousMultilinearMap 𝕜 E rfl #align continuous_multilinear_map.comp_continuous_linear_map_equivL_apply ContinuousMultilinearMap.compContinuousLinearMapEquivL_apply +/-- One of the components of the iterated derivative of a continuous multilinear map. Given a +bijection `e` between a type `α` (typically `Fin k`) and a subset `s` of `ι`, this component is a +continuous multilinear map of `k` vectors `v₁, ..., vₖ`, mapping them +to `f (x₁, (v_{e.symm 2})₂, x₃, ...)`, where at indices `i` in `s` one uses the `i`-th coordinate of +the vector `v_{e.symm i}` and otherwise one uses the `i`-th coordinate of a reference vector `x`. +This is continuous multilinear in the components of `x` outside of `s`, and in the `v_j`. -/ +noncomputable def iteratedFDerivComponent {α : Type*} [Fintype α] [DecidableEq ι] + (f : ContinuousMultilinearMap 𝕜 E₁ G) {s : Set ι} (e : α ≃ s) [DecidablePred (· ∈ s)] : + ContinuousMultilinearMap 𝕜 (fun (i : {a : ι // a ∉ s}) ↦ E₁ i) + (ContinuousMultilinearMap 𝕜 (fun (_ : α) ↦ (∀ i, E₁ i)) G) := + (f.toMultilinearMap.iteratedFDerivComponent e).mkContinuousMultilinear ‖f‖ <| by + intro x m + simp only [MultilinearMap.iteratedFDerivComponent, MultilinearMap.domDomRestrictₗ, + MultilinearMap.coe_mk, MultilinearMap.domDomRestrict_apply, coe_coe] + apply (f.le_opNorm _).trans _ + rw [← prod_compl_mul_prod s.toFinset, mul_assoc] + gcongr + · apply le_of_eq + have : ∀ x, x ∈ s.toFinsetᶜ ↔ (fun x ↦ x ∉ s) x := by simp + rw [prod_subtype _ this] + congr with i + simp [i.2] + · rw [prod_subtype _ (fun _ ↦ s.mem_toFinset), ← Equiv.prod_comp e.symm] + apply Finset.prod_le_prod (fun i _ ↦ norm_nonneg _) (fun i _ ↦ ?_) + simpa only [i.2, ↓reduceDite, Subtype.coe_eta] using norm_le_pi_norm (m (e.symm i)) ↑i + +@[simp] lemma iteratedFDerivComponent_apply {α : Type*} [Fintype α] [DecidableEq ι] + (f : ContinuousMultilinearMap 𝕜 E₁ G) {s : Set ι} (e : α ≃ s) [DecidablePred (· ∈ s)] + (v : ∀ i : {a : ι // a ∉ s}, E₁ i) (w : α → (∀ i, E₁ i)) : + f.iteratedFDerivComponent e v w = + f (fun j ↦ if h : j ∈ s then w (e.symm ⟨j, h⟩) j else v ⟨j, h⟩) := by + simp [iteratedFDerivComponent, MultilinearMap.iteratedFDerivComponent, + MultilinearMap.domDomRestrictₗ] + +lemma norm_iteratedFDerivComponent_le {α : Type*} [Fintype α] [DecidableEq ι] + (f : ContinuousMultilinearMap 𝕜 E₁ G) {s : Set ι} (e : α ≃ s) [DecidablePred (· ∈ s)] + (x : (i : ι) → E₁ i) : + ‖f.iteratedFDerivComponent e (x ·)‖ ≤ ‖f‖ * ‖x‖ ^ (Fintype.card ι - Fintype.card α) := calc + ‖f.iteratedFDerivComponent e (fun i ↦ x i)‖ + ≤ ‖f.iteratedFDerivComponent e‖ * ∏ i : {a : ι // a ∉ s}, ‖x i‖ := + ContinuousMultilinearMap.le_opNorm _ _ + _ ≤ ‖f‖ * ∏ _i : {a : ι // a ∉ s}, ‖x‖ := by + gcongr + · exact MultilinearMap.mkContinuousMultilinear_norm_le _ (norm_nonneg _) _ + · exact fun _ _ ↦ norm_nonneg _ + · exact norm_le_pi_norm _ _ + _ = ‖f‖ * ‖x‖ ^ (Fintype.card {a : ι // a ∉ s}) := by rw [prod_const, card_univ] + _ = ‖f‖ * ‖x‖ ^ (Fintype.card ι - Fintype.card α) := by simp [Fintype.card_congr e] + +open Classical in +/-- The `k`-th iterated derivative of a continuous multilinear map `f` at the point `x`. It is a +continuous multilinear map of `k` vectors `v₁, ..., vₖ` (with the same type as `x`), mapping them +to `∑ f (x₁, (v_{i₁})₂, x₃, ...)`, where at each index `j` one uses either `xⱼ` or one +of the `(vᵢ)ⱼ`, and each `vᵢ` has to be used exactly once. +The sum is parameterized by the embeddings of `Fin k` in the index type `ι` (or, equivalently, +by the subsets `s` of `ι` of cardinality `k` and then the bijections between `Fin k` and `s`). + +The fact that this is indeed the iterated Fréchet derivative is proved in +`ContinuousMultilinearMap.iteratedFDeriv_eq`. +-/ +protected def iteratedFDeriv (f : ContinuousMultilinearMap 𝕜 E₁ G) (k : ℕ) (x : (i : ι) → E₁ i) : + ContinuousMultilinearMap 𝕜 (fun (_ : Fin k) ↦ (∀ i, E₁ i)) G := + ∑ e : Fin k ↪ ι, iteratedFDerivComponent f e.toEquivRange (Pi.compRightL 𝕜 _ Subtype.val x) + +/-- Controlling the norm of `f.iteratedFDeriv` when `f` is continuous multilinear. For the same +bound on the iterated derivative of `f` in the calculus sense, +see `ContinuousMultilinearMap.norm_iteratedFDeriv_le`. -/ +lemma norm_iteratedFDeriv_le' (f : ContinuousMultilinearMap 𝕜 E₁ G) (k : ℕ) (x : (i : ι) → E₁ i) : + ‖f.iteratedFDeriv k x‖ + ≤ Nat.descFactorial (Fintype.card ι) k * ‖f‖ * ‖x‖ ^ (Fintype.card ι - k) := by + classical + calc ‖f.iteratedFDeriv k x‖ + _ ≤ ∑ e : Fin k ↪ ι, ‖iteratedFDerivComponent f e.toEquivRange (fun i ↦ x i)‖ := norm_sum_le _ _ + _ ≤ ∑ _ : Fin k ↪ ι, ‖f‖ * ‖x‖ ^ (Fintype.card ι - k) := by + gcongr with e _ + simpa using norm_iteratedFDerivComponent_le f e.toEquivRange x + _ = Nat.descFactorial (Fintype.card ι) k * ‖f‖ * ‖x‖ ^ (Fintype.card ι - k) := by + simp [card_univ, mul_assoc] + end ContinuousMultilinearMap end Seminorm diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 1a0f917d83536..e081a0b137ac5 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -9,6 +9,7 @@ import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Fintype.Sort import Mathlib.Data.List.FinRange import Mathlib.LinearAlgebra.Pi +import Mathlib.Logic.Equiv.Fintype #align_import linear_algebra.multilinear.basic from "leanprover-community/mathlib"@"78fdf68dcd2fdb3fe64c0dd6f88926a49418a6ea" @@ -780,6 +781,13 @@ lemma domDomRestrict_aux [DecidableEq ι] (P : ι → Prop) [DecidablePred P] rw [Function.update_noteq h''] · simp only [h', ne_eq, Subtype.mk.injEq, dite_false] +lemma domDomRestrict_aux_right [DecidableEq ι] (P : ι → Prop) [DecidablePred P] + [DecidableEq {a // ¬ P a}] + (x : (i : {a // P a}) → M₁ i) (z : (i : {a // ¬ P a}) → M₁ i) (i : {a : ι // ¬ P a}) + (c : M₁ i) : (fun j ↦ if h : P j then x ⟨j, h⟩ else Function.update z i c ⟨j, h⟩) = + Function.update (fun j => if h : P j then x ⟨j, h⟩ else z ⟨j, h⟩) i c := by + simpa only [dite_not] using domDomRestrict_aux _ z (fun j ↦ x ⟨j.1, not_not.mp j.2⟩) i c + /-- Given a multilinear map `f` on `(i : ι) → M i`, a (decidable) predicate `P` on `ι` and an element `z` of `(i : {a // ¬ P a}) → M₁ i`, construct a multilinear map on `(i : {a // P a}) → M₁ i)` whose value at `x` is `f` evaluated at the vector with `i`th coordinate @@ -787,6 +795,8 @@ an element `z` of `(i : {a // ¬ P a}) → M₁ i`, construct a multilinear map The naming is similar to `MultilinearMap.domDomCongr`: here we are applying the restriction to the domain of the domain. + +For a linear map version, see `MultilinearMap.domDomRestrictₗ`. -/ def domDomRestrict (f : MultilinearMap R M₁ M₂) (P : ι → Prop) [DecidablePred P] (z : (i : {a : ι // ¬ P a}) → M₁ i) : @@ -1004,6 +1014,67 @@ variable [CommSemiring R] [∀ i, AddCommMonoid (M₁ i)] [∀ i, AddCommMonoid section variable {M₁' : ι → Type*} [Π i, AddCommMonoid (M₁' i)] [Π i, Module R (M₁' i)] +/-- Given a predicate `P`, one may associate to a multilinear map `f` a multilinear map +from the elements satisfying `P` to the multilinear maps on elements not satisfying `P`. +In other words, splitting the variables into two subsets one gets a multilinear map into +multilinear maps. +This is a linear map version of the function `MultilinearMap.domDomRestrict`. -/ +def domDomRestrictₗ (f : MultilinearMap R M₁ M₂) (P : ι → Prop) [DecidablePred P] : + MultilinearMap R (fun (i : {a : ι // ¬ P a}) => M₁ i) + (MultilinearMap R (fun (i : {a : ι // P a}) => M₁ i) M₂) where + toFun := fun z ↦ domDomRestrict f P z + map_add' := by + intro h m i x y + classical + ext v + simp [domDomRestrict_aux_right] + map_smul' := by + intro h m i c x + classical + ext v + simp [domDomRestrict_aux_right] + +lemma iteratedFDeriv_aux {α : Type*} [DecidableEq α] + (s : Set ι) [DecidableEq { x // x ∈ s }] (e : α ≃ s) + (m : α → ((i : ι) → M₁ i)) (a : α) (z : (i : ι) → M₁ i) : + (fun i ↦ update m a z (e.symm i) i) = + (fun i ↦ update (fun j ↦ m (e.symm j) j) (e a) (z (e a)) i) := by + ext i + rcases eq_or_ne a (e.symm i) with rfl | hne + · rw [Equiv.apply_symm_apply e i, update_same, update_same] + · rw [update_noteq hne.symm, update_noteq fun h ↦ (Equiv.symm_apply_apply .. ▸ h ▸ hne) rfl] + +/-- One of the components of the iterated derivative of a multilinear map. Given a bijection `e` +between a type `α` (typically `Fin k`) and a subset `s` of `ι`, this component is a multilinear map +of `k` vectors `v₁, ..., vₖ`, mapping them to `f (x₁, (v_{e.symm 2})₂, x₃, ...)`, where at +indices `i` in `s` one uses the `i`-th coordinate of the vector `v_{e.symm i}` and otherwise one +uses the `i`-th coordinate of a reference vector `x`. +This is multilinear in the components of `x` outside of `s`, and in the `v_j`. -/ +noncomputable def iteratedFDerivComponent {α : Type*} + (f : MultilinearMap R M₁ M₂) {s : Set ι} (e : α ≃ s) [DecidablePred (· ∈ s)] : + MultilinearMap R (fun (i : {a : ι // a ∉ s}) ↦ M₁ i) + (MultilinearMap R (fun (_ : α) ↦ (∀ i, M₁ i)) M₂) where + toFun := fun z ↦ + { toFun := fun v ↦ domDomRestrictₗ f (fun i ↦ i ∈ s) z (fun i ↦ v (e.symm i) i) + map_add' := by classical simp [iteratedFDeriv_aux] + map_smul' := by classical simp [iteratedFDeriv_aux] } + map_add' := by intros; ext; simp + map_smul' := by intros; ext; simp + +open Classical in +/-- The `k`-th iterated derivative of a multilinear map `f` at the point `x`. It is a multilinear +map of `k` vectors `v₁, ..., vₖ` (with the same type as `x`), mapping them +to `∑ f (x₁, (v_{i₁})₂, x₃, ...)`, where at each index `j` one uses either `xⱼ` or one +of the `(vᵢ)ⱼ`, and each `vᵢ` has to be used exactly once. +The sum is parameterized by the embeddings of `Fin k` in the index type `ι` (or, equivalently, +by the subsets `s` of `ι` of cardinality `k` and then the bijections between `Fin k` and `s`). + +For the continuous version, see `ContinuousMultilinearMap.iteratedFDeriv`. -/ +protected noncomputable def iteratedFDeriv [Fintype ι] + (f : MultilinearMap R M₁ M₂) (k : ℕ) (x : (i : ι) → M₁ i) : + MultilinearMap R (fun (_ : Fin k) ↦ (∀ i, M₁ i)) M₂ := + ∑ e : Fin k ↪ ι, iteratedFDerivComponent f e.toEquivRange (fun i ↦ x i) + /-- If `f` is a collection of linear maps, then the construction `MultilinearMap.compLinearMap` sending a multilinear map `g` to `g (f₁ ⬝ , ..., fₙ ⬝ )` is linear in `g`. -/ @[simps] def compLinearMapₗ (f : Π (i : ι), M₁ i →ₗ[R] M₁' i) : From 945d7cf3cfc5c0ce3527cbe3e52ed8c9ec4d1247 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Wed, 8 May 2024 13:10:28 +0000 Subject: [PATCH 37/52] chore(ConcreteCategory): remove refine' pattern (#12755) Similar to #12729, we should avoid constructing data using the `refine' { foo := foo, bar := bar .. }` pattern. Building data with tactics is always a dicey proposition and should be a last resort. --- .../ConcreteCategory/BundledHom.lean | 15 +++++++-------- 1 file changed, 7 insertions(+), 8 deletions(-) diff --git a/Mathlib/CategoryTheory/ConcreteCategory/BundledHom.lean b/Mathlib/CategoryTheory/ConcreteCategory/BundledHom.lean index eb67db5aa525d..6058bcddf4e1f 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/BundledHom.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/BundledHom.lean @@ -60,14 +60,13 @@ set_option synthInstance.checkSynthOrder false in This instance generates the type-class problem `BundledHom ?m`. Currently that is not a problem, as there are almost no instances of `BundledHom`. -/ -instance category : Category (Bundled c) := by - refine' { Hom := fun X Y => @hom X Y X.str Y.str - id := fun X => @BundledHom.id c hom 𝒞 X X.str - comp := @fun X Y Z f g => @BundledHom.comp c hom 𝒞 X Y Z X.str Y.str Z.str g f - comp_id := _ - id_comp := _ - assoc := _ } <;> intros <;> apply 𝒞.hom_ext <;> - aesop_cat +instance category : Category (Bundled c) where + Hom := fun X Y => hom X.str Y.str + id := fun X => BundledHom.id 𝒞 (α := X) X.str + comp := fun {X Y Z} f g => BundledHom.comp 𝒞 (α := X) (β := Y) (γ := Z) X.str Y.str Z.str g f + comp_id _ := by apply 𝒞.hom_ext; simp + assoc _ _ _ := by apply 𝒞.hom_ext; aesop_cat + id_comp _ := by apply 𝒞.hom_ext; simp #align category_theory.bundled_hom.category CategoryTheory.BundledHom.category /-- A category given by `BundledHom` is a concrete category. -/ From 49c10909cf40bf43c6257ed74933dd0fd0c0435c Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Wed, 8 May 2024 13:44:41 +0000 Subject: [PATCH 38/52] chore(Order/Interval/Finset/Defs): golf (#12752) Co-authored-by: Moritz Firsching --- Mathlib/Order/Interval/Finset/Defs.lean | 98 +++++++------------------ 1 file changed, 26 insertions(+), 72 deletions(-) diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index 6a459d37b15d9..810717c7a9eea 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -710,60 +710,32 @@ instance OrderDual.instLocallyFiniteOrder : LocallyFiniteOrder αᵒᵈ where finset_mem_Ioc _ _ _ := (mem_Ico (α := α)).trans and_comm finset_mem_Ioo _ _ _ := (mem_Ioo (α := α)).trans and_comm -theorem Icc_toDual : Icc (toDual a) (toDual b) = (Icc b a).map toDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Icc, mem_Icc (α := α)] - exact and_comm +theorem Icc_toDual : Icc (toDual a) (toDual b) = (Icc b a).map toDual.toEmbedding := map_refl.symm #align Icc_to_dual Icc_toDual -theorem Ico_toDual : Ico (toDual a) (toDual b) = (Ioc b a).map toDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Ico, mem_Ioc (α := α)] - exact and_comm +theorem Ico_toDual : Ico (toDual a) (toDual b) = (Ioc b a).map toDual.toEmbedding := map_refl.symm #align Ico_to_dual Ico_toDual -theorem Ioc_toDual : Ioc (toDual a) (toDual b) = (Ico b a).map toDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Ioc, mem_Ico (α := α)] - exact and_comm +theorem Ioc_toDual : Ioc (toDual a) (toDual b) = (Ico b a).map toDual.toEmbedding := map_refl.symm #align Ioc_to_dual Ioc_toDual -theorem Ioo_toDual : Ioo (toDual a) (toDual b) = (Ioo b a).map toDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Ioo, mem_Ioo (α := α)] - exact and_comm +theorem Ioo_toDual : Ioo (toDual a) (toDual b) = (Ioo b a).map toDual.toEmbedding := map_refl.symm #align Ioo_to_dual Ioo_toDual -theorem Icc_ofDual (a b : αᵒᵈ) : Icc (ofDual a) (ofDual b) = (Icc b a).map ofDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Icc, mem_Icc (α := αᵒᵈ)] - exact and_comm +theorem Icc_ofDual (a b : αᵒᵈ) : Icc (ofDual a) (ofDual b) = (Icc b a).map ofDual.toEmbedding := + map_refl.symm #align Icc_of_dual Icc_ofDual -theorem Ico_ofDual (a b : αᵒᵈ) : Ico (ofDual a) (ofDual b) = (Ioc b a).map ofDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Ico, mem_Ioc (α := αᵒᵈ)] - exact and_comm +theorem Ico_ofDual (a b : αᵒᵈ) : Ico (ofDual a) (ofDual b) = (Ioc b a).map ofDual.toEmbedding := + map_refl.symm #align Ico_of_dual Ico_ofDual -theorem Ioc_ofDual (a b : αᵒᵈ) : Ioc (ofDual a) (ofDual b) = (Ico b a).map ofDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Ioc, mem_Ico (α := αᵒᵈ)] - exact and_comm +theorem Ioc_ofDual (a b : αᵒᵈ) : Ioc (ofDual a) (ofDual b) = (Ico b a).map ofDual.toEmbedding := + map_refl.symm #align Ioc_of_dual Ioc_ofDual -theorem Ioo_ofDual (a b : αᵒᵈ) : Ioo (ofDual a) (ofDual b) = (Ioo b a).map ofDual.toEmbedding := by - refine' Eq.trans _ map_refl.symm - ext c - rw [mem_Ioo, mem_Ioo (α := αᵒᵈ)] - exact and_comm +theorem Ioo_ofDual (a b : αᵒᵈ) : Ioo (ofDual a) (ofDual b) = (Ioo b a).map ofDual.toEmbedding := + map_refl.symm #align Ioo_of_dual Ioo_ofDual end LocallyFiniteOrder @@ -841,20 +813,17 @@ namespace Prod instance [LocallyFiniteOrder α] [LocallyFiniteOrder β] [DecidableRel ((· ≤ ·) : α × β → α × β → Prop)] : LocallyFiniteOrder (α × β) := LocallyFiniteOrder.ofIcc' (α × β) (fun a b => Icc a.fst b.fst ×ˢ Icc a.snd b.snd) fun a b x => by - rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm] - rfl + rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm, le_def, le_def] instance [LocallyFiniteOrderTop α] [LocallyFiniteOrderTop β] [DecidableRel ((· ≤ ·) : α × β → α × β → Prop)] : LocallyFiniteOrderTop (α × β) := LocallyFiniteOrderTop.ofIci' (α × β) (fun a => Ici a.fst ×ˢ Ici a.snd) fun a x => by - rw [mem_product, mem_Ici, mem_Ici] - rfl + rw [mem_product, mem_Ici, mem_Ici, le_def] instance [LocallyFiniteOrderBot α] [LocallyFiniteOrderBot β] [DecidableRel ((· ≤ ·) : α × β → α × β → Prop)] : LocallyFiniteOrderBot (α × β) := LocallyFiniteOrderBot.ofIic' (α × β) (fun a => Iic a.fst ×ˢ Iic a.snd) fun a x => by - rw [mem_product, mem_Iic, mem_Iic] - rfl + rw [mem_product, mem_Iic, mem_Iic, le_def] theorem Icc_eq [LocallyFiniteOrder α] [LocallyFiniteOrder β] [DecidableRel ((· ≤ ·) : α × β → α × β → Prop)] (p q : α × β) : @@ -960,8 +929,7 @@ instance locallyFiniteOrder : LocallyFiniteOrder (WithTop α) where | (a : α), ⊤, ⊤ => by simp [WithTop.some, WithTop.top, insertNone] | (a : α), ⊤, (x : α) => by simp only [some, le_eq_subset, some_le_some, le_top, and_true] - rw [some_mem_insertNone] - simp + rw [some_mem_insertNone, mem_Ici] | (a : α), (b : α), ⊤ => by simp only [Embedding.some, mem_map, mem_Icc, and_false, exists_const, some, le_top, top_le_iff] @@ -1285,8 +1253,8 @@ theorem BddBelow.finite_of_bddAbove [Preorder α] [LocallyFiniteOrder α] theorem Set.finite_iff_bddAbove [SemilatticeSup α] [LocallyFiniteOrder α] [OrderBot α] : s.Finite ↔ BddAbove s := - ⟨fun h ↦ ⟨h.toFinset.sup id, fun x hx ↦ Finset.le_sup (f := id) (by simpa)⟩, - fun ⟨m, hm⟩ ↦ (Set.finite_Icc ⊥ m).subset (fun x hx ↦ ⟨bot_le, hm hx⟩)⟩ + ⟨fun h ↦ ⟨h.toFinset.sup id, fun _ hx ↦ Finset.le_sup (f := id) ((Finite.mem_toFinset h).mpr hx)⟩, + fun ⟨m, hm⟩ ↦ (Set.finite_Icc ⊥ m).subset (fun _ hx ↦ ⟨bot_le, hm hx⟩)⟩ theorem Set.finite_iff_bddBelow [SemilatticeInf α] [LocallyFiniteOrder α] [OrderTop α] : s.Finite ↔ BddBelow s := @@ -1296,8 +1264,10 @@ theorem Set.finite_iff_bddBelow_bddAbove [Nonempty α] [Lattice α] [LocallyFini s.Finite ↔ BddBelow s ∧ BddAbove s := by obtain (rfl | hs) := s.eq_empty_or_nonempty · simp only [Set.finite_empty, bddBelow_empty, bddAbove_empty, and_self] - exact ⟨fun h ↦ ⟨⟨h.toFinset.inf' (by simpa) id, fun x hx ↦ Finset.inf'_le id (by simpa)⟩, - ⟨h.toFinset.sup' (by simpa) id, fun x hx ↦ Finset.le_sup' id (by simpa)⟩⟩, + exact ⟨fun h ↦ ⟨⟨h.toFinset.inf' ((Finite.toFinset_nonempty h).mpr hs) id, + fun x hx ↦ Finset.inf'_le id ((Finite.mem_toFinset h).mpr hx)⟩, + ⟨h.toFinset.sup' ((Finite.toFinset_nonempty h).mpr hs) id, fun x hx ↦ Finset.le_sup' id + ((Finite.mem_toFinset h).mpr hx)⟩⟩, fun ⟨h₀, h₁⟩ ↦ BddBelow.finite_of_bddAbove h₀ h₁⟩ end Finite @@ -1350,32 +1320,16 @@ instance (priority := low) [Preorder α] [DecidableRel ((· : α) < ·)] [Locall exact fun _ => b.property instance [Preorder α] [LocallyFiniteOrderBot α] : Finite { x : α // x ≤ y } := by - apply Set.Finite.to_subtype - convert (Finset.Iic y).finite_toSet using 1 - ext - simp - rfl + simpa only [coe_Iic] using (Finset.Iic y).finite_toSet instance [Preorder α] [LocallyFiniteOrderBot α] : Finite { x : α // x < y } := by - apply Set.Finite.to_subtype - convert (Finset.Iio y).finite_toSet using 1 - ext - simp - rfl + simpa only [coe_Iio] using (Finset.Iio y).finite_toSet instance [Preorder α] [LocallyFiniteOrderTop α] : Finite { x : α // y ≤ x } := by - apply Set.Finite.to_subtype - convert (Finset.Ici y).finite_toSet using 1 - ext - simp - rfl + simpa only [coe_Ici] using (Finset.Ici y).finite_toSet instance [Preorder α] [LocallyFiniteOrderTop α] : Finite { x : α // y < x } := by - apply Set.Finite.to_subtype - convert (Finset.Ioi y).finite_toSet using 1 - ext - simp - rfl + simpa only [coe_Ioi] using (Finset.Ioi y).finite_toSet namespace Set variable {α : Type*} [Preorder α] From 063b1255bfc75dd729a3f168dbf1838a55213f49 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 8 May 2024 15:10:41 +0000 Subject: [PATCH 39/52] fix: typo in polyrith (#12754) --- Mathlib/Tactic/Polyrith.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 7acbd079047ff..92ab11e3d03ad 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -422,7 +422,7 @@ elab_rules : tactic | `(tactic| polyrith%$tk $[only%$onlyTk]? $[[$hyps,*]]?) => do let hyps ← hyps.map (·.getElems) |>.getD #[] |>.mapM (elabTerm · none) let traceMe ← Lean.isTracingEnabledFor `Meta.Tactic.polyrith - match ← polyrith (← getMainGoal) onlyTk.isNone hyps traceMe with + match ← polyrith (← getMainGoal) onlyTk.isSome hyps traceMe with | .ok stx => replaceMainGoal [] if !traceMe then Lean.Meta.Tactic.TryThis.addSuggestion tk stx From 9d8f6783f64176fcf8a328020945b9d79f3a8711 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 8 May 2024 15:35:53 +0000 Subject: [PATCH 40/52] feat(Algebra/Category/ModuleCat): the category of presheaves of modules has colimits (#12713) --- Mathlib.lean | 1 + .../Category/ModuleCat/ChangeOfRings.lean | 13 ++ .../Algebra/Category/ModuleCat/Colimits.lean | 7 + .../Category/ModuleCat/Presheaf/Colimits.lean | 159 ++++++++++++++++++ .../Category/ModuleCat/Presheaf/Limits.lean | 3 - 5 files changed, 180 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean diff --git a/Mathlib.lean b/Mathlib.lean index ee72500dd3ea7..25ce23f473fdb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -88,6 +88,7 @@ import Mathlib.Algebra.Category.ModuleCat.Monoidal.Basic import Mathlib.Algebra.Category.ModuleCat.Monoidal.Closed import Mathlib.Algebra.Category.ModuleCat.Monoidal.Symmetric import Mathlib.Algebra.Category.ModuleCat.Presheaf +import Mathlib.Algebra.Category.ModuleCat.Presheaf.Colimits import Mathlib.Algebra.Category.ModuleCat.Presheaf.Limits import Mathlib.Algebra.Category.ModuleCat.Products import Mathlib.Algebra.Category.ModuleCat.Projective diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 9daa574cffe25..2221518135dcb 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.Algebra.Category.ModuleCat.EpiMono +import Mathlib.Algebra.Category.ModuleCat.Colimits import Mathlib.Algebra.Category.ModuleCat.Limits import Mathlib.RingTheory.TensorProduct.Basic @@ -873,4 +874,16 @@ noncomputable instance preservesLimitRestrictScalars have hc' := isLimitOfPreserves (forget₂ _ AddCommGroupCat) hc exact isLimitOfReflects (forget₂ _ AddCommGroupCat) hc'⟩ +instance preservesColimitRestrictScalars {R S : Type*} [Ring R] [Ring S] + (f : R →+* S) {J : Type*} [Category J] (F : J ⥤ ModuleCat.{v} S) + [HasColimit (F ⋙ forget₂ _ AddCommGroupCat)] : + PreservesColimit F (ModuleCat.restrictScalars.{v} f) := by + have : HasColimit ((F ⋙ restrictScalars f) ⋙ forget₂ (ModuleCat R) AddCommGroupCat) := + inferInstanceAs (HasColimit (F ⋙ forget₂ _ AddCommGroupCat)) + apply preservesColimitOfPreservesColimitCocone (HasColimit.isColimitColimitCocone F) + apply isColimitOfReflects (forget₂ _ AddCommGroupCat) + apply isColimitOfPreserves (forget₂ (ModuleCat.{v} S) AddCommGroupCat.{v}) + exact HasColimit.isColimitColimitCocone F + + end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean index 3ca028fe8d29e..1c452e36296f8 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean @@ -101,6 +101,10 @@ instance : HasColimit F := ⟨_, isColimitColimitCocone F⟩ noncomputable instance : PreservesColimit F (forget₂ _ AddCommGroupCat) := preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) +noncomputable instance reflectsColimit : + ReflectsColimit F (forget₂ (ModuleCat.{w'} R) AddCommGroupCat) := + reflectsColimitOfReflectsIsomorphisms _ _ + end HasColimit variable (J R) @@ -108,6 +112,9 @@ variable (J R) instance hasColimitsOfShape [HasColimitsOfShape J AddCommGroupCat.{w'}] : HasColimitsOfShape J (ModuleCat.{w'} R) where +noncomputable instance reflectsColimitsOfShape [HasColimitsOfShape J AddCommGroupCat.{w'}] : + ReflectsColimitsOfShape J (forget₂ (ModuleCat.{w'} R) AddCommGroupCat) where + instance hasColimitsOfSize [HasColimitsOfSize.{v, u} AddCommGroupCat.{w'}] : HasColimitsOfSize.{v, u} (ModuleCat.{w'} R) where diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean new file mode 100644 index 0000000000000..fefcc51f4fd18 --- /dev/null +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean @@ -0,0 +1,159 @@ +/- +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.ModuleCat.Presheaf +import Mathlib.Algebra.Category.ModuleCat.Colimits + +/-! # Colimits in categories of presheaves of modules + +In this file, it is shown that under suitable assumptions, +colimits exist in the category `PresheafOfModules R`. + +-/ + +universe v v₁ v₂ u₁ u₂ u u' + +open CategoryTheory Category Limits + +namespace PresheafOfModules + +variable {C : Type u₁} [Category.{v₁} C] {R : Cᵒᵖ ⥤ RingCat.{u}} + {J : Type u₂} [Category.{v₂} J] + (F : J ⥤ PresheafOfModules.{v} R) + +section Colimits + +variable [∀ {X Y : Cᵒᵖ} (f : X ⟶ Y), PreservesColimit (F ⋙ evaluation R Y) + (ModuleCat.restrictScalars (R.map f))] + +/-- A cocone in the category `PresheafOfModules R` is colimit if it is so after the application +of the functors `evaluation R X` for all `X`. -/ +def evaluationJointlyReflectsColimits (c : Cocone F) + (hc : ∀ (X : Cᵒᵖ), IsColimit ((evaluation R X).mapCocone c)) : IsColimit c where + desc s := Hom.mk'' (fun X => (hc X).desc ((evaluation R X).mapCocone s)) (fun X Y f => by + apply (hc X).hom_ext + intro j + erw [(hc X).fac_assoc ((evaluation R X).mapCocone s) j, ← restrictionApp_naturality_assoc] + rw [← Functor.map_comp] + erw [(hc Y).fac ((evaluation R Y).mapCocone s), restrictionApp_naturality] + rfl) + fac s j := by + ext1 X + exact (hc X).fac ((evaluation R X).mapCocone s) j + uniq s m hm := by + ext1 X + apply (hc X).uniq ((evaluation R X).mapCocone s) + intro j + dsimp + rw [← hm] + rfl + +variable [∀ X, HasColimit (F ⋙ evaluation R X)] + +instance {X Y : Cᵒᵖ} (f : X ⟶ Y) : + HasColimit (F ⋙ evaluation R Y ⋙ (ModuleCat.restrictScalars (R.map f))) := + ⟨_, isColimitOfPreserves (ModuleCat.restrictScalars (R.map f)) + (colimit.isColimit (F ⋙ evaluation R Y))⟩ + +/-- Given `F : J ⥤ PresheafOfModules.{v} R`, this is the `BundledCorePresheafOfModules R` which +corresponds to the presheaf of modules which sends `X` to the colimit of `F ⋙ evaluation R X`. -/ +@[simps] +noncomputable def colimitBundledCore : BundledCorePresheafOfModules R where + obj X := colimit (F ⋙ evaluation R X) + map {X Y} f := colimMap (whiskerLeft F (restriction R f)) ≫ + (preservesColimitIso (ModuleCat.restrictScalars (R.map f)) (F ⋙ evaluation R Y)).inv + map_id X := colimit.hom_ext (fun j => by + dsimp + rw [ι_colimMap_assoc, whiskerLeft_app, restriction_app] + erw [ι_preservesColimitsIso_inv (G := ModuleCat.restrictScalars (R.map (𝟙 X))), + ModuleCat.restrictScalarsId'App_inv_naturality] + rw [restrictionApp_id] + rfl) + map_comp {X Y Z} f g := colimit.hom_ext (fun j => by + dsimp + rw [ι_colimMap_assoc, whiskerLeft_app, restriction_app, assoc, ι_colimMap_assoc] + erw [ι_preservesColimitsIso_inv (G := ModuleCat.restrictScalars (R.map (f ≫ g))), + ι_preservesColimitsIso_inv_assoc (G := ModuleCat.restrictScalars (R.map f))] + rw [← Functor.map_comp_assoc, ι_colimMap_assoc] + erw [ι_preservesColimitsIso_inv (G := ModuleCat.restrictScalars (R.map g))] + rw [restrictionApp_comp, ModuleCat.restrictScalarsComp'_inv_app, assoc, assoc, + whiskerLeft_app, whiskerLeft_app, restriction_app, restriction_app] + simp only [Functor.map_comp, assoc] + rfl) + +/-- Given `F : J ⥤ PresheafOfModules.{v} R`, this is the canonical map +`F.obj j ⟶ (colimitBundledCore F).toPresheafOfModules` for all `j : J`. -/ +noncomputable def colimitCoconeιApp (j : J) : + F.obj j ⟶ (colimitBundledCore F).toPresheafOfModules := + PresheafOfModules.Hom.mk'' (fun X => colimit.ι (F ⋙ evaluation R X) j) (fun X Y f => by + dsimp + erw [colimit.ι_desc_assoc, assoc, ← ι_preservesColimitsIso_inv] + rfl) + +@[reassoc (attr := simp)] +lemma colimitCoconeιApp_naturality {i j : J} (f : i ⟶ j) : + F.map f ≫ colimitCoconeιApp F j = colimitCoconeιApp F i := by + ext1 X + exact colimit.w (F ⋙ evaluation R X) f + +/-- The (colimit) cocone for `F : J ⥤ PresheafOfModules.{v} R` that is constructed from +the colimit of `F ⋙ evaluation R X` for all `X`. -/ +@[simps] +noncomputable def colimitCocone : Cocone F where + pt := (colimitBundledCore F).toPresheafOfModules + ι := { app := colimitCoconeιApp F } + +/-- The cocone `colimitCocone F` is colimit for any `F : J ⥤ PresheafOfModules.{v} R`. -/ +noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) := + evaluationJointlyReflectsColimits _ _ (fun _ => colimit.isColimit _) + +instance hasColimit : HasColimit F := ⟨_, isColimitColimitCocone F⟩ + +noncomputable instance evaluationPreservesColimit (X : Cᵒᵖ) : + PreservesColimit F (evaluation R X) := + preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) + +variable [∀ X, PreservesColimit F + (evaluation R X ⋙ forget₂ (ModuleCat (R.obj X)) AddCommGroupCat)] + +noncomputable instance toPresheafPreservesColimit : + PreservesColimit F (toPresheaf R) := + preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + (Limits.evaluationJointlyReflectsColimits _ + (fun X => isColimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGroupCat) + (isColimitColimitCocone F))) + +end Colimits + +variable (R J) + +section HasColimitsOfShape + +variable [HasColimitsOfShape J AddCommGroupCat.{v}] + +instance hasColimitsOfShape : HasColimitsOfShape J (PresheafOfModules.{v} R) where + +noncomputable instance evaluationPreservesColimitsOfShape (X : Cᵒᵖ) : + PreservesColimitsOfShape J (evaluation R X : PresheafOfModules.{v} R ⥤ _) where + +noncomputable instance toPresheafPreservesColimitsOfShape : + PreservesColimitsOfShape J (toPresheaf.{v} R) where + +end HasColimitsOfShape + +namespace Finite + +instance hasFiniteColimits : HasFiniteColimits (PresheafOfModules.{v} R) := + ⟨fun _ => inferInstance⟩ + +noncomputable instance evaluationPreservesFiniteColimits (X : Cᵒᵖ) : + PreservesFiniteColimits (evaluation.{v} R X) where + +noncomputable instance toPresheafPreservesFiniteColimits : + PreservesFiniteColimits (toPresheaf R) where + +end Finite + +end PresheafOfModules diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean index c5b87e8127cbf..b5cb79ba5c6b5 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean @@ -13,9 +13,6 @@ import Mathlib.CategoryTheory.Limits.FunctorCategory In this file, it is shown that under suitable assumptions, limits exist in the category `PresheafOfModules R`. -## TODO -* do the same for colimits - -/ universe v v₁ v₂ u₁ u₂ u u' From 54a54ded180ca32fdc608779ed3447c213dc3a76 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 8 May 2024 15:35:54 +0000 Subject: [PATCH 41/52] fix: remove MathlibExtras from mk_all.sh (#12763) This directory was removed in #12548. --- scripts/mk_all.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/mk_all.sh b/scripts/mk_all.sh index cd69ce8ef2641..d5712ee52b116 100755 --- a/scripts/mk_all.sh +++ b/scripts/mk_all.sh @@ -1,6 +1,6 @@ #! /usr/bin/env bash cd "$(git rev-parse --show-toplevel)" || exit 1 -for gp in Mathlib MathlibExtras Mathlib/Tactic Counterexamples Archive +for gp in Mathlib Mathlib/Tactic Counterexamples Archive do git ls-files "$gp/*.lean" | LC_ALL=C sort | sed 's/\.lean//;s,/,.,g;s/^/import /' > "$gp.lean" done From e6ea7987e2bef631d2be529252a5ff581e12362e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 8 May 2024 16:11:19 +0000 Subject: [PATCH 42/52] feat: `positivity` extension for `posPart`, `negPart` (#10681) --- Mathlib/Tactic/Positivity/Basic.lean | 23 +++++++++++++++++++++++ scripts/noshake.json | 1 + test/positivity.lean | 4 ++++ 3 files changed, 28 insertions(+) diff --git a/Mathlib/Tactic/Positivity/Basic.lean b/Mathlib/Tactic/Positivity/Basic.lean index d10e7f4edc057..ef67fdf4e2bbc 100644 --- a/Mathlib/Tactic/Positivity/Basic.lean +++ b/Mathlib/Tactic/Positivity/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Heather Macbeth, Yaël Dillies -/ import Mathlib.Algebra.GroupPower.Order +import Mathlib.Algebra.Order.Group.PosPart import Mathlib.Data.Int.CharZero import Mathlib.Algebra.Order.Ring.Int import Mathlib.Data.Nat.Factorial.Basic @@ -450,3 +451,25 @@ def evalRatDen : PositivityExt where eval {u α} _ _ e := do assumeInstancesCommute pure $ .positive q(den_pos $a) | _, _ => throwError "not Rat.num" + +/-- Extension for `posPart`. `a⁺` is always nonegative, and positive if `a` is. -/ +@[positivity _⁺] +def evalPosPart : PositivityExt where eval zα pα e := do + match e with + | ~q(@posPart _ $instαlat $instαgrp $a) => + assertInstancesCommute + -- FIXME: There seems to be a bug in `Positivity.core` that makes it fail (instead of returning + -- `.none`) here sometimes. See eg the first test for `posPart`. This is why we need `catchNone` + match ← catchNone (core zα pα a) with + | .positive pf => return .positive q(posPart_pos $pf) + | _ => return .nonnegative q(posPart_nonneg $a) + | _ => throwError "not `posPart`" + +/-- Extension for `negPart`. `a⁻` is always nonegative. -/ +@[positivity _⁻] +def evalNegPart : PositivityExt where eval _ _ e := do + match e with + | ~q(@negPart _ $instαlat $instαgrp $a) => + assertInstancesCommute + return .nonnegative q(negPart_nonneg $a) + | _ => throwError "not `negPart`" diff --git a/scripts/noshake.json b/scripts/noshake.json index 08fa3decea166..69ec9572e9508 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -222,6 +222,7 @@ "Mathlib.Tactic.Positivity.Finset": ["Mathlib.Data.Fintype.Card"], "Mathlib.Tactic.Positivity.Basic": ["Mathlib.Algebra.GroupPower.Order", + "Mathlib.Algebra.Order.Group.PosPart", "Mathlib.Data.Int.CharZero", "Mathlib.Data.Nat.Factorial.Basic"], "Mathlib.Tactic.NormNum.Ineq": diff --git a/test/positivity.lean b/test/positivity.lean index 5e58a2e5776d3..decca1cf3385c 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -196,6 +196,10 @@ example (hq : 0 ≤ q) : 0 ≤ q.num := by positivity end +example (a : ℤ) : 0 ≤ a⁺ := by positivity +example (a : ℤ) (ha : 0 < a) : 0 < a⁺ := by positivity +example (a : ℤ) : 0 ≤ a⁻ := by positivity + /-! ### Exponentiation -/ example [OrderedSemiring α] [Nontrivial α] (a : α) : 0 < a ^ 0 := by positivity From 59b7bc4871691da68166f821185b40bf5fb9240a Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Wed, 8 May 2024 18:28:53 +0000 Subject: [PATCH 43/52] refactor(Opposite): `Opposite.mk` -> `Opposite.op` (#12636) It was observed on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/making.20.60op.60.20irreducible) that we should really have `op` as the custom name for the `Opposite` constructor as long as it remains a structure. --- .../Morphisms/RingHomProperties.lean | 3 ++- Mathlib/CategoryTheory/Preadditive/Injective.lean | 4 ++-- Mathlib/CategoryTheory/Sites/CompatiblePlus.lean | 9 +++------ Mathlib/Data/Opposite.lean | 13 +++++-------- Mathlib/Geometry/RingedSpace/OpenImmersion.lean | 8 +++++--- Mathlib/Topology/Sheaves/Skyscraper.lean | 1 - 6 files changed, 17 insertions(+), 21 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 6b5eb8f6e83fd..cf5af48cd1f0c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -308,7 +308,8 @@ theorem sourceAffineLocally_of_source_openCover {X Y : Scheme.{u}} (f : X ⟶ Y) swap · refine' X.presheaf.map (@homOfLE _ _ ((IsOpenMap.functor _).obj _) ((IsOpenMap.functor _).obj _) _).op - rw [unop_op, unop_op, Opens.openEmbedding_obj_top, Opens.openEmbedding_obj_top] + 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 : _) diff --git a/Mathlib/CategoryTheory/Preadditive/Injective.lean b/Mathlib/CategoryTheory/Preadditive/Injective.lean index da4f68035a0ea..cccd3796845bd 100644 --- a/Mathlib/CategoryTheory/Preadditive/Injective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Injective.lean @@ -161,11 +161,11 @@ instance {β : Type v} (c : β → C) [HasZeroMorphisms C] [HasBiproduct c] [∀ ext simp only [Category.assoc, biproduct.lift_π, comp_factorThru] -instance {P : Cᵒᵖ} [Projective P] : Injective (unop P) where +instance {P : Cᵒᵖ} [Projective P] : Injective no_index (unop P) where factors g f mono := ⟨(@Projective.factorThru Cᵒᵖ _ P _ _ _ g.op f.op _).unop, Quiver.Hom.op_inj (by simp)⟩ -instance {J : Cᵒᵖ} [Injective J] : Projective (unop J) where +instance {J : Cᵒᵖ} [Injective J] : Projective no_index (unop J) where factors f e he := ⟨(@factorThru Cᵒᵖ _ J _ _ _ f.op e.op _).unop, Quiver.Hom.op_inj (by simp)⟩ diff --git a/Mathlib/CategoryTheory/Sites/CompatiblePlus.lean b/Mathlib/CategoryTheory/Sites/CompatiblePlus.lean index 86291a996dfe0..fa9080af35d54 100644 --- a/Mathlib/CategoryTheory/Sites/CompatiblePlus.lean +++ b/Mathlib/CategoryTheory/Sites/CompatiblePlus.lean @@ -207,13 +207,10 @@ theorem whiskerRight_toPlus_comp_plusCompIso_hom : simp only [ι_plusCompIso_hom, Functor.map_comp, Category.assoc] simp only [← Category.assoc] congr 1 - -- Porting note: this used to work with `ext` -- See https://github.com/leanprover-community/mathlib4/issues/5229 - apply Multiequalizer.hom_ext - delta Cover.toMultiequalizer - simp only [diagramCompIso_hom_ι, Category.assoc, ← F.map_comp] - simp only [unop_op, limit.lift_π, Multifork.ofι_π_app, Functor.comp_obj, Functor.comp_map, - implies_true] + apply Multiequalizer.hom_ext; intro a + rw [Category.assoc, diagramCompIso_hom_ι, ← F.map_comp] + simp only [unop_op, limit.lift_π, Multifork.ofι_π_app, Functor.comp_obj, Functor.comp_map] #align category_theory.grothendieck_topology.whisker_right_to_plus_comp_plus_comp_iso_hom CategoryTheory.GrothendieckTopology.whiskerRight_toPlus_comp_plusCompIso_hom @[simp] diff --git a/Mathlib/Data/Opposite.lean b/Mathlib/Data/Opposite.lean index 27fb4a76304df..ecf0018e7b25e 100644 --- a/Mathlib/Data/Opposite.lean +++ b/Mathlib/Data/Opposite.lean @@ -33,10 +33,15 @@ variable (α : Sort u) -/ structure Opposite := + /-- The canonical map `α → αᵒᵖ`. -/ + op :: /-- The canonical map `αᵒᵖ → α`. -/ unop : α #align opposite Opposite #align opposite.unop Opposite.unop +#align opposite.op Opposite.op + +-- Porting note: pp_nodot has not been implemented for Opposite.op @[inherit_doc] notation:max -- Use a high right binding power (like that of postfix ⁻¹) so that, for example, @@ -47,12 +52,6 @@ namespace Opposite variable {α} -/-- The canonical map `α → αᵒᵖ`. -/ --- Porting note: pp_nodot has not been implemented. ---@[pp_nodot] -def op (x : α) : αᵒᵖ := ⟨x⟩ -#align opposite.op Opposite.op - theorem op_injective : Function.Injective (op : α → αᵒᵖ) := fun _ _ => congr_arg Opposite.unop #align opposite.op_injective Opposite.op_injective @@ -64,14 +63,12 @@ theorem op_unop (x : αᵒᵖ) : op (unop x) = x := rfl #align opposite.op_unop Opposite.op_unop -@[simp] theorem unop_op (x : α) : unop (op x) = x := rfl #align opposite.unop_op Opposite.unop_op -- We could prove these by `Iff.rfl`, but that would make these eligible for `dsimp`. That would be -- a bad idea because `Opposite` is irreducible. -@[simp] theorem op_inj_iff (x y : α) : op x = op y ↔ x = y := op_injective.eq_iff #align opposite.op_inj_iff Opposite.op_inj_iff diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index feb2c4d771bf2..469c8cef45550 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -64,7 +64,7 @@ namespace AlgebraicGeometry universe v v₁ v₂ u -variable {C : Type*} [Category C] +variable {C : Type u} [Category.{v} C] /-- An open immersion of PresheafedSpaces is an open embedding `f : X ⟶ U ⊆ Y` of the underlying spaces, such that the sheaf map `Y(V) ⟶ f _* X(V)` is an iso for each `V ⊆ U`. @@ -295,7 +295,6 @@ instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} congr apply Subsingleton.helim rw [this] - rfl · infer_instance #align algebraic_geometry.PresheafedSpace.is_open_immersion.of_restrict AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.ofRestrict @@ -373,7 +372,10 @@ def pullbackConeOfLeftFst : intro U V i induction U using Opposite.rec' induction V using Opposite.rec' - simp only [Quiver.Hom.unop_op, Category.assoc, Functor.op_map, inv_naturality_assoc] + simp only [Quiver.Hom.unop_op, Category.assoc, Functor.op_map] + -- Note: this doesn't fire in `simp` because of reduction of the term via structure eta + -- before discrimination tree key generation + rw [inv_naturality_assoc] -- Porting note: the following lemmas are not picked up by `simp` -- See https://github.com/leanprover-community/mathlib4/issues/5026 erw [g.c.naturality_assoc, TopCat.Presheaf.pushforwardObj_map, ← Y.presheaf.map_comp, diff --git a/Mathlib/Topology/Sheaves/Skyscraper.lean b/Mathlib/Topology/Sheaves/Skyscraper.lean index f7e5199691e64..a861e784e9f3e 100644 --- a/Mathlib/Topology/Sheaves/Skyscraper.lean +++ b/Mathlib/Topology/Sheaves/Skyscraper.lean @@ -307,7 +307,6 @@ theorem to_skyscraper_fromStalk {𝓕 : Presheaf C X} {c : C} (f : 𝓕 ⟶ skys split_ifs with h · erw [← Category.assoc, colimit.ι_desc, Category.assoc, eqToHom_trans, eqToHom_refl, Category.comp_id] - simp only [unop_op, op_unop] · exact ((if_neg h).symm.ndrec terminalIsTerminal).hom_ext .. #align stalk_skyscraper_presheaf_adjunction_auxs.to_skyscraper_from_stalk StalkSkyscraperPresheafAdjunctionAuxs.to_skyscraper_fromStalk From ba7e114658fe6add6aad39c24a6b499101de60e1 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Thu, 9 May 2024 00:52:22 +0000 Subject: [PATCH 44/52] feat(Data/Nat/Digits): `digits_head` and `ofDigits_mod_eq_head` (#11129) Co-authored-by: Moritz Firsching --- Mathlib/Algebra/BigOperators/List/Basic.lean | 2 +- Mathlib/Computability/TMToPartrec.lean | 2 -- Mathlib/Data/List/Basic.lean | 3 +++ Mathlib/Data/Nat/Defs.lean | 2 ++ Mathlib/Data/Nat/Digits.lean | 16 ++++++++++++++++ 5 files changed, 22 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/List/Basic.lean b/Mathlib/Algebra/BigOperators/List/Basic.lean index 68e5c7958861d..4ff97ad3c3ff3 100644 --- a/Mathlib/Algebra/BigOperators/List/Basic.lean +++ b/Mathlib/Algebra/BigOperators/List/Basic.lean @@ -574,7 +574,7 @@ If desired, we could add a class stating that `default = 0`. /-- This relies on `default ℕ = 0`. -/ theorem headI_add_tail_sum (L : List ℕ) : L.headI + L.tail.sum = L.sum := by - cases L <;> simp; rfl + cases L <;> simp #align list.head_add_tail_sum List.headI_add_tail_sum /-- This relies on `default ℕ = 0`. -/ diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 481b0c322c45c..75c3db48727e2 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -1182,7 +1182,6 @@ def trNat (n : ℕ) : List Γ' := theorem trNat_zero : trNat 0 = [] := by rw [trNat, Nat.cast_zero]; rfl #align turing.partrec_to_TM2.tr_nat_zero Turing.PartrecToTM2.trNat_zero -@[simp] theorem trNat_default : trNat default = [] := trNat_zero #align turing.partrec_to_TM2.tr_nat_default Turing.PartrecToTM2.trNat_default @@ -1550,7 +1549,6 @@ theorem pred_ok (q₁ q₂ s v) (c d : List Γ') : ∃ s', rcases v with (_ | ⟨_ | n, v⟩) · refine' ⟨none, TransGen.single _⟩ simp - rfl · refine' ⟨some Γ'.cons, TransGen.single _⟩ simp refine' ⟨none, _⟩ diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index db21afe70a8eb..22eb8d4054b1d 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -764,6 +764,9 @@ theorem getLast?_append {l₁ l₂ : List α} {x : α} (h : x ∈ l₂.getLast?) /-! ### head(!?) and tail -/ +@[simp] +theorem head!_nil [Inhabited α] : ([] : List α).head! = default := rfl + @[simp] theorem head_cons_tail (x : List α) (h : x ≠ []) : x.head h :: x.tail = x := by cases x <;> simp at h ⊢ diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 4acfababd4f88..f499725a23f49 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -66,6 +66,8 @@ variable {a b c d m n k : ℕ} {p q : ℕ → Prop} instance instNontrivial : Nontrivial ℕ := ⟨⟨0, 1, Nat.zero_ne_one⟩⟩ +@[simp] theorem default_eq_zero : default = 0 := rfl + attribute [gcongr] Nat.succ_le_succ attribute [simp] Nat.not_lt_zero Nat.succ_ne_zero Nat.succ_ne_self Nat.zero_ne_one Nat.one_ne_zero Nat.min_eq_left Nat.min_eq_right Nat.max_eq_left Nat.max_eq_right diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 06347ebbab00a..c6619ce27948b 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -192,6 +192,9 @@ theorem ofDigits_eq_sum_mapIdx (b : ℕ) (L : List ℕ) : Or.inl hl #align nat.of_digits_eq_sum_map_with_index Nat.ofDigits_eq_sum_mapIdx +@[simp] +theorem ofDigits_nil {b : ℕ} : ofDigits b [] = 0 := rfl + @[simp] theorem ofDigits_singleton {b n : ℕ} : ofDigits b [n] = n := by simp [ofDigits] #align nat.of_digits_singleton Nat.ofDigits_singleton @@ -660,6 +663,19 @@ theorem ofDigits_mod (b k : ℕ) (L : List ℕ) : ofDigits b L % k = ofDigits (b ofDigits_modEq b k L #align nat.of_digits_mod Nat.ofDigits_mod +theorem ofDigits_mod_eq_head! (b : ℕ) (l : List ℕ) : ofDigits b l % b = l.head! % b := by + induction l <;> simp [Nat.ofDigits, Int.ModEq] + +theorem head!_digits {b n : ℕ} (h : b ≠ 1) : (Nat.digits b n).head! = n % b := by + by_cases hb : 1 < b + · rcases n with _ | n + · simp + · nth_rw 2 [← Nat.ofDigits_digits b (n + 1)] + rw [Nat.ofDigits_mod_eq_head! _ _] + exact (Nat.mod_eq_of_lt (Nat.digits_lt_base hb <| List.head!_mem_self <| + Nat.digits_ne_nil_iff_ne_zero.mpr <| Nat.succ_ne_zero n)).symm + · rcases n with _ | _ <;> simp_all [show b = 0 by omega] + theorem ofDigits_zmodeq' (b b' : ℤ) (k : ℕ) (h : b ≡ b' [ZMOD k]) (L : List ℕ) : ofDigits b L ≡ ofDigits b' L [ZMOD k] := by induction' L with d L ih From 47d4fa91aedcbd429282046dd1e8c80404a13afc Mon Sep 17 00:00:00 2001 From: Raghuram <102684766+raghuram-13@users.noreply.github.com> Date: Thu, 9 May 2024 04:14:54 +0000 Subject: [PATCH 45/52] feat(Algebra/Ring/Ext): prove extensionality lemmas for Ring and similar typeclasses (#9511) Add a file `Algebra/Ring/Ext`, proving `ext` lemmas for all of the ring-like classes (for example, anything from `NonUnitalNonAssocSemiring` to `CommRing`), stating that two instances of such a class are equal if they define the same addition and multiplication operations. Also prove `ext_iff` variants for each class, and injectivity lemmas for projections from classes to parent classes. --- Mathlib.lean | 1 + Mathlib/Algebra/Ring/Ext.lean | 534 ++++++++++++++++++++++++++++++++++ 2 files changed, 535 insertions(+) create mode 100644 Mathlib/Algebra/Ring/Ext.lean diff --git a/Mathlib.lean b/Mathlib.lean index 25ce23f473fdb..d10438cf9a07b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -612,6 +612,7 @@ import Mathlib.Algebra.Ring.Defs import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Algebra.Ring.Divisibility.Lemmas import Mathlib.Algebra.Ring.Equiv +import Mathlib.Algebra.Ring.Ext import Mathlib.Algebra.Ring.Fin import Mathlib.Algebra.Ring.Hom.Basic import Mathlib.Algebra.Ring.Hom.Defs diff --git a/Mathlib/Algebra/Ring/Ext.lean b/Mathlib/Algebra/Ring/Ext.lean new file mode 100644 index 0000000000000..9d38fa233e300 --- /dev/null +++ b/Mathlib/Algebra/Ring/Ext.lean @@ -0,0 +1,534 @@ +/- +Copyright (c) 2024 Raghuram Sundararajan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Raghuram Sundararajan +-/ +import Mathlib.Algebra.Ring.Defs +import Mathlib.Algebra.Group.Ext + +/-! +# Extensionality lemmas for rings and similar structures + +In this file we prove extensionality lemmas for the ring-like structures defined in +`Mathlib/Algebra/Ring/Defs.lean`, ranging from `NonUnitalNonAssocSemiring` to `CommRing`. These +extensionality lemmas take the form of asserting that two algebraic structures on a type are equal +whenever the addition and multiplication defined by them are both the same. + +## Implementation details + +We follow `Mathlib/Algebra/Group/Ext.lean` in using the term `(letI := i; HMul.hMul : R → R → R)` to +refer to the multiplication specified by a typeclass instance `i` on a type `R` (and similarly for +addition). We abbreviate these using some local notations. + +Since `Mathlib/Algebra/Group/Ext.lean` proved several injectivity lemmas, we do so as well — even if +sometimes we don't need them to prove extensionality. + +## Tags +semiring, ring, extensionality +-/ + +local macro:max "local_hAdd[" type:term ", " inst:term "]" : term => + `(term| (letI := $inst; HAdd.hAdd : $type → $type → $type)) +local macro:max "local_hMul[" type:term ", " inst:term "]" : term => + `(term| (letI := $inst; HMul.hMul : $type → $type → $type)) + +universe u + +variable {R : Type u} + +/-! ### Distrib -/ +namespace Distrib + +@[ext] theorem ext ⦃inst₁ inst₂ : Distrib R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + -- Split into `add` and `mul` functions and properties. + rcases inst₁ with @⟨⟨⟩, ⟨⟩⟩ + rcases inst₂ with @⟨⟨⟩, ⟨⟩⟩ + -- Prove equality of parts using function extensionality. + congr + +theorem ext_iff {inst₁ inst₂ : Distrib R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end Distrib + +/-! ### NonUnitalNonAssocSemiring -/ +namespace NonUnitalNonAssocSemiring + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalNonAssocSemiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + -- Split into `AddMonoid` instance, `mul` function and properties. + rcases inst₁ with @⟨_, ⟨⟩⟩ + rcases inst₂ with @⟨_, ⟨⟩⟩ + -- Prove equality of parts using already-proved extensionality lemmas. + congr; ext : 1; assumption + +theorem toDistrib_injective : Function.Injective (@toDistrib R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocSemiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalNonAssocSemiring + +/-! ### NonUnitalSemiring -/ +namespace NonUnitalSemiring + +theorem toNonUnitalNonAssocSemiring_injective : + Function.Injective (@toNonUnitalNonAssocSemiring R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalSemiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toNonUnitalNonAssocSemiring_injective <| + NonUnitalNonAssocSemiring.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : NonUnitalSemiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalSemiring + +/-! ### NonAssocSemiring and its ancestors + +This section also includes results for `AddMonoidWithOne`, `AddCommMonoidWithOne`, etc. +as these are considered implementation detail of the ring classes. +TODO consider relocating these lemmas. +-/ +/- TODO consider relocating these lemmas. -/ +@[ext] theorem AddMonoidWithOne.ext ⦃inst₁ inst₂ : AddMonoidWithOne R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_one : (letI := inst₁; One.one : R) = (letI := inst₂; One.one : R)) : + inst₁ = inst₂ := by + have h_monoid : inst₁.toAddMonoid = inst₂.toAddMonoid := by ext : 1; exact h_add + have h_zero' : inst₁.toZero = inst₂.toZero := congrArg (·.toZero) h_monoid + have h_one' : inst₁.toOne = inst₂.toOne := + congrArg One.mk h_one + have h_natCast : inst₁.toNatCast.natCast = inst₂.toNatCast.natCast := by + funext n; induction n with + | zero => rewrite [inst₁.natCast_zero, inst₂.natCast_zero] + exact congrArg (@Zero.zero R) h_zero' + | succ n h => rw [inst₁.natCast_succ, inst₂.natCast_succ, h_add] + exact congrArg₂ _ h h_one + rcases inst₁ with @⟨⟨⟩⟩; rcases inst₂ with @⟨⟨⟩⟩ + congr + +theorem AddCommMonoidWithOne.toAddMonoidWithOne_injective : + Function.Injective (@AddCommMonoidWithOne.toAddMonoidWithOne R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem AddCommMonoidWithOne.ext ⦃inst₁ inst₂ : AddCommMonoidWithOne R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_one : (letI := inst₁; One.one : R) = (letI := inst₂; One.one : R)) : + inst₁ = inst₂ := + AddCommMonoidWithOne.toAddMonoidWithOne_injective <| + AddMonoidWithOne.ext h_add h_one + +namespace NonAssocSemiring + +/- The best place to prove that the `NatCast` is determined by the other operations is probably in +an extensionality lemma for `AddMonoidWithOne`, in which case we may as well do the typeclasses +defined in `Mathlib/Algebra/GroupWithZero/Defs.lean` as well. -/ +@[ext] theorem ext ⦃inst₁ inst₂ : NonAssocSemiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + have h : inst₁.toNonUnitalNonAssocSemiring = inst₂.toNonUnitalNonAssocSemiring := by + ext : 1 <;> assumption + have h_zero : (inst₁.toMulZeroClass).toZero.zero = (inst₂.toMulZeroClass).toZero.zero := + congrArg (fun inst => (inst.toMulZeroClass).toZero.zero) h + have h_one' : (inst₁.toMulZeroOneClass).toMulOneClass.toOne + = (inst₂.toMulZeroOneClass).toMulOneClass.toOne := + congrArg (@MulOneClass.toOne R) <| by ext : 1; exact h_mul + have h_one : (inst₁.toMulZeroOneClass).toMulOneClass.toOne.one + = (inst₂.toMulZeroOneClass).toMulOneClass.toOne.one := + congrArg (@One.one R) h_one' + have : inst₁.toAddCommMonoidWithOne = inst₂.toAddCommMonoidWithOne := + by ext : 1 <;> assumption + have : inst₁.toNatCast = inst₂.toNatCast := + congrArg (·.toNatCast) this + -- Split into `NonUnitalNonAssocSemiring`, `One` and `natCast` instances. + cases inst₁; cases inst₂ + congr + +theorem toNonUnitalNonAssocSemiring_injective : + Function.Injective (@toNonUnitalNonAssocSemiring R) := by + intro _ _ _ + ext <;> congr + +theorem ext_iff {inst₁ inst₂ : NonAssocSemiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonAssocSemiring + +/-! ### NonUnitalNonAssocRing -/ +namespace NonUnitalNonAssocRing + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalNonAssocRing R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + -- Split into `AddCommGroup` instance, `mul` function and properties. + rcases inst₁ with @⟨_, ⟨⟩⟩; rcases inst₂ with @⟨_, ⟨⟩⟩ + congr; (ext : 1; assumption) + +theorem toNonUnitalNonAssocSemiring_injective : + Function.Injective (@toNonUnitalNonAssocSemiring R) := by + intro _ _ h + -- Use above extensionality lemma to prove injectivity by showing that `h_add` and `h_mul` hold. + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocRing R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalNonAssocRing + +/-! ### NonUnitalRing -/ +namespace NonUnitalRing + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalRing R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + have : inst₁.toNonUnitalNonAssocRing = inst₂.toNonUnitalNonAssocRing := by + ext : 1 <;> assumption + -- Split into fields and prove they are equal using the above. + cases inst₁; cases inst₂ + congr + +theorem toNonUnitalSemiring_injective : + Function.Injective (@toNonUnitalSemiring R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem toNonUnitalNonAssocring_injective : + Function.Injective (@toNonUnitalNonAssocRing R) := by + intro _ _ _ + ext <;> congr + +theorem ext_iff {inst₁ inst₂ : NonUnitalRing R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalRing + +/-! ### NonAssocRing and its ancestors + +This section also includes results for `AddGroupWithOne`, `AddCommGroupWithOne`, etc. +as these are considered implementation detail of the ring classes. +TODO consider relocating these lemmas. -/ +@[ext] theorem AddGroupWithOne.ext ⦃inst₁ inst₂ : AddGroupWithOne R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_one : (letI := inst₁; One.one : R) = (letI := inst₂; One.one)) : + inst₁ = inst₂ := by + have : inst₁.toAddMonoidWithOne = inst₂.toAddMonoidWithOne := + AddMonoidWithOne.ext h_add h_one + have : inst₁.toNatCast = inst₂.toNatCast := congrArg (·.toNatCast) this + have h_group : inst₁.toAddGroup = inst₂.toAddGroup := by ext : 1; exact h_add + -- Extract equality of necessary substructures from h_group + injection h_group with h_group; injection h_group + have : inst₁.toIntCast.intCast = inst₂.toIntCast.intCast := by + funext n; cases n with + | ofNat n => rewrite [Int.ofNat_eq_coe, inst₁.intCast_ofNat, inst₂.intCast_ofNat]; congr + | negSucc n => rewrite [inst₁.intCast_negSucc, inst₂.intCast_negSucc]; congr + rcases inst₁ with @⟨⟨⟩⟩; rcases inst₂ with @⟨⟨⟩⟩ + congr + +@[ext] theorem AddCommGroupWithOne.ext ⦃inst₁ inst₂ : AddCommGroupWithOne R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_one : (letI := inst₁; One.one : R) = (letI := inst₂; One.one)) : + inst₁ = inst₂ := by + have : inst₁.toAddCommGroup = inst₂.toAddCommGroup := + AddCommGroup.ext h_add + have : inst₁.toAddGroupWithOne = inst₂.toAddGroupWithOne := + AddGroupWithOne.ext h_add h_one + injection this with _ h_addMonoidWithOne; injection h_addMonoidWithOne + cases inst₁; cases inst₂ + congr + +namespace NonAssocRing + +@[ext] theorem ext ⦃inst₁ inst₂ : NonAssocRing R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + have h₁ : inst₁.toNonUnitalNonAssocRing = inst₂.toNonUnitalNonAssocRing := by + ext : 1 <;> assumption + have h₂ : inst₁.toNonAssocSemiring = inst₂.toNonAssocSemiring := by + ext : 1 <;> assumption + -- Mathematically non-trivial fact: `intCast` is determined by the rest. + have h₃ : inst₁.toAddCommGroupWithOne = inst₂.toAddCommGroupWithOne := + AddCommGroupWithOne.ext h_add (congrArg (·.toOne.one) h₂) + cases inst₁; cases inst₂ + congr <;> solve| injection h₁ | injection h₂ | injection h₃ + +theorem toNonAssocSemiring_injective : + Function.Injective (@toNonAssocSemiring R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem toNonUnitalNonAssocring_injective : + Function.Injective (@toNonUnitalNonAssocRing R) := by + intro _ _ _ + ext <;> congr + +theorem ext_iff {inst₁ inst₂ : NonAssocRing R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonAssocRing + +/-! ### Semiring -/ +namespace Semiring + +@[ext] theorem ext ⦃inst₁ inst₂ : Semiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + -- Show that enough substructures are equal. + have h₁ : inst₁.toNonUnitalSemiring = inst₂.toNonUnitalSemiring := by + ext : 1 <;> assumption + have h₂ : inst₁.toNonAssocSemiring = inst₂.toNonAssocSemiring := by + ext : 1 <;> assumption + have h₃ : (inst₁.toMonoidWithZero).toMonoid = (inst₂.toMonoidWithZero).toMonoid := by + ext : 1; exact h_mul + -- Split into fields and prove they are equal using the above. + cases inst₁; cases inst₂ + congr <;> solve| injection h₁ | injection h₂ | injection h₃ + +theorem toNonUnitalSemiring_injective : + Function.Injective (@toNonUnitalSemiring R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem toNonAssocSemiring_injective : + Function.Injective (@toNonAssocSemiring R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem ext_iff {inst₁ inst₂ : Semiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end Semiring + +/-! ### Ring -/ +namespace Ring + +@[ext] theorem ext ⦃inst₁ inst₂ : Ring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := by + -- Show that enough substructures are equal. + have h₁ : inst₁.toSemiring = inst₂.toSemiring := by + ext : 1 <;> assumption + have h₂ : inst₁.toNonAssocRing = inst₂.toNonAssocRing := by + ext : 1 <;> assumption + /- We prove that the `SubNegMonoid`s are equal because they are one + field away from `Sub` and `Neg`, enabling use of `injection`. -/ + have h₃ : (inst₁.toAddCommGroup).toAddGroup.toSubNegMonoid + = (inst₂.toAddCommGroup).toAddGroup.toSubNegMonoid := + congrArg (@AddGroup.toSubNegMonoid R) <| by ext : 1; exact h_add + -- Split into fields and prove they are equal using the above. + cases inst₁; cases inst₂ + congr <;> solve | injection h₂ | injection h₃ + +theorem toNonUnitalRing_injective : + Function.Injective (@toNonUnitalRing R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem toNonAssocRing_injective : + Function.Injective (@toNonAssocRing R) := by + intro _ _ _ + ext <;> congr + +theorem toSemiring_injective : + Function.Injective (@toSemiring R) := by + intro _ _ h + ext x y + · exact congrArg (·.toAdd.add x y) h + · exact congrArg (·.toMul.mul x y) h + +theorem ext_iff {inst₁ inst₂ : Ring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext ·)⟩ + +end Ring + +/-! ### NonUnitalNonAssocCommSemiring -/ +namespace NonUnitalNonAssocCommSemiring + +theorem toNonUnitalNonAssocSemiring_injective : + Function.Injective (@toNonUnitalNonAssocSemiring R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalNonAssocCommSemiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toNonUnitalNonAssocSemiring_injective <| + NonUnitalNonAssocSemiring.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocCommSemiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalNonAssocCommSemiring + +/-! ### NonUnitalCommSemiring -/ +namespace NonUnitalCommSemiring + +theorem toNonUnitalSemiring_injective : + Function.Injective (@toNonUnitalSemiring R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalCommSemiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toNonUnitalSemiring_injective <| + NonUnitalSemiring.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : NonUnitalCommSemiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalCommSemiring + +-- At present, there is no `NonAssocCommSemiring` in Mathlib. + +/-! ### NonUnitalNonAssocCommRing -/ +namespace NonUnitalNonAssocCommRing + +theorem toNonUnitalNonAssocRing_injective : + Function.Injective (@toNonUnitalNonAssocRing R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalNonAssocCommRing R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toNonUnitalNonAssocRing_injective <| + NonUnitalNonAssocRing.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocCommRing R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalNonAssocCommRing + +/-! ### NonUnitalCommRing -/ +namespace NonUnitalCommRing + +theorem toNonUnitalRing_injective : + Function.Injective (@toNonUnitalRing R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : NonUnitalCommRing R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toNonUnitalRing_injective <| + NonUnitalRing.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : NonUnitalCommRing R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end NonUnitalCommRing + +-- At present, there is no `NonAssocCommRing` in Mathlib. + +/-! ### CommSemiring -/ +namespace CommSemiring + +theorem toSemiring_injective : + Function.Injective (@toSemiring R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : CommSemiring R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toSemiring_injective <| + Semiring.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : CommSemiring R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ + +end CommSemiring + +/-! ### CommRing -/ +namespace CommRing + +theorem toRing_injective : Function.Injective (@toRing R) := by + rintro ⟨⟩ ⟨⟩ _; congr + +@[ext] theorem ext ⦃inst₁ inst₂ : CommRing R⦄ + (h_add : local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) + (h_mul : local_hMul[R, inst₁] = local_hMul[R, inst₂]) : + inst₁ = inst₂ := + toRing_injective <| Ring.ext h_add h_mul + +theorem ext_iff {inst₁ inst₂ : CommRing R} : + inst₁ = inst₂ ↔ + (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ + (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := + ⟨by rintro rfl; constructor <;> rfl, And.elim (ext ·)⟩ + +end CommRing From b72ddd8f67eba47cffd3f3422d6c01f1e666c23f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 May 2024 05:29:02 +0000 Subject: [PATCH 46/52] feat: backport changes from lean-pr-testing-4096 (#12772) These changes will be required after https://github.com/leanprover/lean4/pull/4096, but they all look like positive readability changes to me anyway (the things being filled in are sort of hard to work out with time travelling to later lines in the proofs!), so I think we can backport them all to `master`. Co-authored-by: Scott Morrison --- Mathlib/Algebra/Order/CauSeq/Completion.lean | 2 +- Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean | 4 ++-- Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean | 2 +- Mathlib/Computability/AkraBazzi/AkraBazzi.lean | 2 +- Mathlib/GroupTheory/Coset.lean | 2 +- Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean | 2 +- Mathlib/MeasureTheory/Measure/Haar/Unique.lean | 2 +- Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean | 2 +- Mathlib/Order/PrimeIdeal.lean | 2 +- Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean | 2 +- Mathlib/Topology/ContinuousFunction/Bounded.lean | 2 +- 11 files changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/Algebra/Order/CauSeq/Completion.lean b/Mathlib/Algebra/Order/CauSeq/Completion.lean index e835ea0fc6a27..43f63c26c1587 100644 --- a/Mathlib/Algebra/Order/CauSeq/Completion.lean +++ b/Mathlib/Algebra/Order/CauSeq/Completion.lean @@ -236,7 +236,7 @@ theorem inv_mk {f} (hf) : (@mk α _ β _ abv _ f)⁻¹ = mk (inv f hf) := #align cau_seq.completion.inv_mk CauSeq.Completion.inv_mk theorem cau_seq_zero_ne_one : ¬(0 : CauSeq _ abv) ≈ 1 := fun h => - have : LimZero (1 - 0) := Setoid.symm h + have : LimZero (1 - 0 : CauSeq _ abv) := Setoid.symm h have : LimZero 1 := by simpa by apply one_ne_zero <| const_limZero.1 this #align cau_seq.completion.cau_seq_zero_ne_one CauSeq.Completion.cau_seq_zero_ne_one diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean index 06304fe75c151..e63fdf1c2ad26 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean @@ -160,7 +160,7 @@ variable (𝕜 E F) ‖fst 𝕜 E F‖ = 1 := by refine le_antisymm (norm_fst_le ..) ?_ let ⟨e, he⟩ := exists_ne (0 : E) - have : ‖e‖ ≤ _ * max ‖e‖ ‖0‖ := (fst 𝕜 E F).le_opNorm (e, 0) + have : ‖e‖ ≤ _ * max ‖e‖ ‖(0 : F)‖ := (fst 𝕜 E F).le_opNorm (e, 0) rw [norm_zero, max_eq_left (norm_nonneg e)] at this rwa [← mul_le_mul_iff_of_pos_right (norm_pos_iff.mpr he), one_mul] @@ -170,7 +170,7 @@ variable (𝕜 E F) ‖snd 𝕜 E F‖ = 1 := by refine le_antisymm (norm_snd_le ..) ?_ let ⟨f, hf⟩ := exists_ne (0 : F) - have : ‖f‖ ≤ _ * max ‖0‖ ‖f‖ := (snd 𝕜 E F).le_opNorm (0, f) + have : ‖f‖ ≤ _ * max ‖(0 : E)‖ ‖f‖ := (snd 𝕜 E F).le_opNorm (0, f) rw [norm_zero, max_eq_right (norm_nonneg f)] at this rwa [← mul_le_mul_iff_of_pos_right (norm_pos_iff.mpr hf), one_mul] diff --git a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean index 35b5932ff26a3..60ffd6fc7e1e9 100644 --- a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean @@ -96,7 +96,7 @@ theorem not_integrableOn_Ioi_rpow (s : ℝ) : ¬ IntegrableOn (fun x ↦ x ^ s) · have : IntegrableOn (fun x ↦ x ^ s) (Ioo (0 : ℝ) 1) := h.mono Ioo_subset_Ioi_self le_rfl rw [integrableOn_Ioo_rpow_iff zero_lt_one] at this exact hs.not_lt this - · have : IntegrableOn (fun x ↦ x ^ s) (Ioi 1) := h.mono (Ioi_subset_Ioi zero_le_one) le_rfl + · have : IntegrableOn (fun x ↦ x ^ s) (Ioi (1 : ℝ)) := h.mono (Ioi_subset_Ioi zero_le_one) le_rfl rw [integrableOn_Ioi_rpow_iff zero_lt_one] at this exact hs.not_lt this diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index ab91ba6657b6b..914f92776b9ff 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -520,7 +520,7 @@ lemma tendsto_zero_sumCoeffsExp : Tendsto (fun (p : ℝ) => ∑ i, a i * (b i) ^ linarith lemma tendsto_atTop_sumCoeffsExp : Tendsto (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) atBot atTop := by - have h₁ : Tendsto (fun p => (a (max_bi b) : ℝ) * b (max_bi b) ^ p) atBot atTop := + have h₁ : Tendsto (fun p : ℝ => (a (max_bi b) : ℝ) * b (max_bi b) ^ p) atBot atTop := Tendsto.mul_atTop (R.a_pos (max_bi b)) (by simp) <| tendsto_rpow_atBot_of_base_lt_one _ (by have := R.b_pos (max_bi b); linarith) (R.b_lt_one _) diff --git a/Mathlib/GroupTheory/Coset.lean b/Mathlib/GroupTheory/Coset.lean index 3e22b20e912cd..0f4b46c1664b1 100644 --- a/Mathlib/GroupTheory/Coset.lean +++ b/Mathlib/GroupTheory/Coset.lean @@ -148,7 +148,7 @@ variable [Monoid α] (s : Submonoid α) @[to_additive mem_own_leftAddCoset] theorem mem_own_leftCoset (a : α) : a ∈ a • (s : Set α) := - suffices a * 1 ∈ a • ↑s by simpa + suffices a * 1 ∈ a • (s : Set α) by simpa mem_leftCoset a (one_mem s : 1 ∈ s) #align mem_own_left_coset mem_own_leftCoset #align mem_own_left_add_coset mem_own_leftAddCoset diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 6a5351e8839ba..b7885fae4ed40 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -118,7 +118,7 @@ def invertibleEquivDetInvertible : Invertible A ≃ Invertible A.det where variable {A B} theorem mul_eq_one_comm : A * B = 1 ↔ B * A = 1 := - suffices ∀ A B, A * B = 1 → B * A = 1 from ⟨this A B, this B A⟩ + suffices ∀ A B : Matrix n n α, A * B = 1 → B * A = 1 from ⟨this A B, this B A⟩ fun A B h => by letI : Invertible B.det := detInvertibleOfLeftInverse _ _ h letI : Invertible B := invertibleOfDetInvertible B diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index 6a74cbf3e0ea6..39ef8d37ab32f 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -492,7 +492,7 @@ lemma measure_preimage_isMulLeftInvariant_eq_smul_of_hasCompactSupport · simpa using (u_mem n).2.le have I1 := I μ' (by infer_instance) simp_rw [M] at I1 - have J1 : ∫ (x : G), indicator {1} (fun _ ↦ 1) (f x) ∂μ' + have J1 : ∫ (x : G), indicator {1} (fun _ ↦ (1 : ℝ)) (f x) ∂μ' = ∫ (x : G), indicator {1} (fun _ ↦ 1) (f x) ∂(haarScalarFactor μ' μ • μ) := tendsto_nhds_unique I1 (I (haarScalarFactor μ' μ • μ) (by infer_instance)) have J2 : ENNReal.toReal (μ' (f ⁻¹' {1})) diff --git a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean index 9a94e2cc82300..1e5a747035295 100644 --- a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean @@ -88,7 +88,7 @@ theorem prime_of_nat_prime_of_mod_four_eq_three (p : ℕ) [hp : Fact p.Prime] (h irreducible_iff_prime.1 <| by_contradiction fun hpi => let ⟨a, b, hab⟩ := sq_add_sq_of_nat_prime_of_not_irreducible p hpi - have : ∀ a b : ZMod 4, a ^ 2 + b ^ 2 ≠ ↑p := by + have : ∀ a b : ZMod 4, a ^ 2 + b ^ 2 ≠ (p : ZMod 4) := by erw [← ZMod.natCast_mod p 4, hp3]; decide this a b (hab ▸ by simp) #align gaussian_int.prime_of_nat_prime_of_mod_four_eq_three GaussianInt.prime_of_nat_prime_of_mod_four_eq_three diff --git a/Mathlib/Order/PrimeIdeal.lean b/Mathlib/Order/PrimeIdeal.lean index 00ca6fc656f29..137da1310f5c3 100644 --- a/Mathlib/Order/PrimeIdeal.lean +++ b/Mathlib/Order/PrimeIdeal.lean @@ -158,7 +158,7 @@ instance (priority := 100) IsMaximal.isPrime [IsMaximal I] : IsPrime I := by let J := I ⊔ principal x have hJuniv : (J : Set P) = Set.univ := IsMaximal.maximal_proper (lt_sup_principal_of_not_mem ‹_›) - have hyJ : y ∈ ↑J := Set.eq_univ_iff_forall.mp hJuniv y + have hyJ : y ∈ (J : Set P) := Set.eq_univ_iff_forall.mp hJuniv y rw [coe_sup_eq] at hyJ rcases hyJ with ⟨a, ha, b, hb, hy⟩ rw [hy] diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean index 21b12711a50b1..86bc20cf4ce1f 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean @@ -175,7 +175,7 @@ theorem dvd_pow_natDegree_of_eval₂_eq_zero {f : R →+* A} (hf : Function.Inje (Ideal.mem_span_singleton.mpr <| dvd_refl x)).pow_natDegree_le_of_root_of_monic_mem _ ((monic_scaleRoots_iff x).mpr hp) _ le_rfl rw [injective_iff_map_eq_zero'] at hf - have : eval₂ _ _ (p.scaleRoots x) = 0 := scaleRoots_eval₂_eq_zero f h + have : eval₂ f _ (p.scaleRoots x) = 0 := scaleRoots_eval₂_eq_zero f h rwa [hz, Polynomial.eval₂_at_apply, hf] at this #align polynomial.dvd_pow_nat_degree_of_eval₂_eq_zero Polynomial.dvd_pow_natDegree_of_eval₂_eq_zero diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index c37c0582b37d3..39ec9997fa04d 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -581,7 +581,7 @@ theorem arzela_ascoli₂ (s : Set β) (hs : IsCompact s) (A : Set (α →ᵇ β) IsCompact A := by /- This version is deduced from the previous one by restricting to the compact type in the target, using compactness there and then lifting everything to the original space. -/ - have M : LipschitzWith 1 (↑) := LipschitzWith.subtype_val s + have M : LipschitzWith 1 Subtype.val := LipschitzWith.subtype_val s let F : (α →ᵇ s) → α →ᵇ β := comp (↑) M refine' IsCompact.of_isClosed_subset ((_ : IsCompact (F ⁻¹' A)).image (continuous_comp M)) closed fun f hf => _ From ad9935b346590a398c264d9b4bb33f3b6f1928ae Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 May 2024 08:07:51 +0000 Subject: [PATCH 47/52] chore: fix deprecation warnings (#12775) In preparation for https://github.com/leanprover/lean4/pull/3969 landing, we fix some deprecation warnings. --- Mathlib/Algebra/Algebra/Equiv.lean | 2 +- Mathlib/Algebra/DirectSum/Algebra.lean | 2 +- Mathlib/Algebra/Group/Submonoid/Units.lean | 4 ++-- .../Algebra/GroupWithZero/Units/Basic.lean | 2 ++ Mathlib/AlgebraicGeometry/Pullbacks.lean | 2 +- .../AlgebraicTopology/DoldKan/GammaCompN.lean | 2 +- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 10 +++++----- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 6 +++--- Mathlib/Analysis/Calculus/Darboux.lean | 2 +- Mathlib/Analysis/Complex/AbsMax.lean | 2 +- .../Analysis/Complex/LocallyUniformLimit.lean | 2 +- Mathlib/Analysis/Complex/OpenMapping.lean | 4 ++-- .../Analysis/Complex/PhragmenLindelof.lean | 2 +- Mathlib/Analysis/Convex/Gauge.lean | 4 ++-- Mathlib/Analysis/Convex/KreinMilman.lean | 4 ++-- Mathlib/Analysis/Convolution.lean | 13 ++++++------ .../Analysis/Distribution/SchwartzSpace.lean | 2 +- .../Analysis/InnerProductSpace/Rayleigh.lean | 4 ++-- .../Analysis/NormedSpace/FiniteDimension.lean | 2 +- .../NormedSpace/HahnBanach/Separation.lean | 8 +++++--- Mathlib/Analysis/NormedSpace/Spectrum.lean | 2 +- Mathlib/Analysis/Seminorm.lean | 6 +++--- .../CategoryTheory/Comma/StructuredArrow.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Discrete.lean | 2 +- Mathlib/Data/Finset/Sort.lean | 20 +++++++++---------- Mathlib/Data/Finsupp/Basic.lean | 2 +- Mathlib/Data/List/Basic.lean | 4 ++++ Mathlib/Data/List/DropRight.lean | 10 +++++----- Mathlib/Data/List/Forall2.lean | 3 +++ Mathlib/Data/List/Indexes.lean | 1 + Mathlib/Data/List/Infix.lean | 2 +- Mathlib/Data/List/InsertNth.lean | 2 ++ Mathlib/Data/List/NodupEquivFin.lean | 3 ++- Mathlib/Data/List/Perm.lean | 1 + Mathlib/Data/List/Sort.lean | 6 +++++- Mathlib/Data/List/Zip.lean | 6 ++++-- Mathlib/Data/MLList/BestFirst.lean | 4 ++-- Mathlib/Data/Nat/Nth.lean | 2 +- .../RotationNumber/TranslationNumber.lean | 4 ++-- Mathlib/GroupTheory/Perm/Cycle/Concrete.lean | 10 ++++++++++ Mathlib/GroupTheory/Perm/List.lean | 1 + .../Function/LpSeminorm/Basic.lean | 2 +- .../Function/StronglyMeasurable/Basic.lean | 2 +- Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 2 +- .../NumberTheory/LSeries/HurwitzZetaEven.lean | 6 +++--- .../ModularForms/JacobiTheta/OneVariable.lean | 2 +- Mathlib/NumberTheory/NumberField/Basic.lean | 2 +- Mathlib/NumberTheory/ZetaValues.lean | 2 +- Mathlib/Probability/Density.lean | 5 +++-- Mathlib/RingTheory/FreeCommRing.lean | 2 +- Mathlib/RingTheory/Trace.lean | 2 +- Mathlib/Topology/Algebra/Order/Compact.lean | 10 +++++----- Mathlib/Topology/Algebra/Order/Rolle.lean | 4 ++-- .../Topology/ContinuousFunction/Ideals.lean | 2 +- .../MetricSpace/GromovHausdorffRealized.lean | 2 +- .../MetricSpace/HausdorffDistance.lean | 8 +++----- Mathlib/Topology/MetricSpace/ProperSpace.lean | 2 +- Mathlib/Topology/ProperMap.lean | 6 ++++-- Mathlib/Topology/UniformSpace/Basic.lean | 2 ++ 59 files changed, 137 insertions(+), 100 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 228041aff0e15..45e7651e606fb 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -232,7 +232,7 @@ nonrec theorem map_sum {ι : Type*} (f : ι → A₁) (s : Finset ι) : theorem map_finsupp_sum {α : Type*} [Zero α] {ι : Type*} (f : ι →₀ α) (g : ι → α → A₁) : e (f.sum g) = f.sum fun i b => e (g i b) := - e.map_sum _ _ + _root_.map_sum e _ _ #align alg_equiv.map_finsupp_sum AlgEquiv.map_finsupp_sum -- Porting note: Added [coe] attribute diff --git a/Mathlib/Algebra/DirectSum/Algebra.lean b/Mathlib/Algebra/DirectSum/Algebra.lean index fdb8c89aaca47..85b2c03bf1b09 100644 --- a/Mathlib/Algebra/DirectSum/Algebra.lean +++ b/Mathlib/Algebra/DirectSum/Algebra.lean @@ -72,7 +72,7 @@ instance : Algebra R (⨁ i, A i) where toFun := (DirectSum.of A 0).comp GAlgebra.toFun map_zero' := AddMonoidHom.map_zero _ map_add' := AddMonoidHom.map_add _ - map_one' := (DirectSum.of A 0).congr_arg GAlgebra.map_one + map_one' := DFunLike.congr_arg (DirectSum.of A 0) GAlgebra.map_one map_mul' a b := by simp only [AddMonoidHom.comp_apply] rw [of_mul_of] diff --git a/Mathlib/Algebra/Group/Submonoid/Units.lean b/Mathlib/Algebra/Group/Submonoid/Units.lean index 89e48fc8afa75..ef959d7f891e4 100644 --- a/Mathlib/Algebra/Group/Submonoid/Units.lean +++ b/Mathlib/Algebra/Group/Submonoid/Units.lean @@ -105,11 +105,11 @@ lemma inv_val_mem_of_mem_units (S : Submonoid M) {x : Mˣ} (h : x ∈ S.units) : @[to_additive] lemma coe_inv_val_mul_coe_val (S : Submonoid M) {x : Sˣ} : - ((x⁻¹ : Sˣ) : M) * ((x : Sˣ) : M) = 1 := S.subtype.congr_arg x.inv_mul + ((x⁻¹ : Sˣ) : M) * ((x : Sˣ) : M) = 1 := DFunLike.congr_arg S.subtype x.inv_mul @[to_additive] lemma coe_val_mul_coe_inv_val (S : Submonoid M) {x : Sˣ} : - ((x : Sˣ) : M) * ((x⁻¹ : Sˣ) : M) = 1 := S.subtype.congr_arg x.mul_inv + ((x : Sˣ) : M) * ((x⁻¹ : Sˣ) : M) = 1 := DFunLike.congr_arg S.subtype x.mul_inv @[to_additive] lemma mk_inv_mul_mk_eq_one (S : Submonoid M) {x : Mˣ} (h : x ∈ S.units) : diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 0cc69ff3a8732..5125c1cbbbac5 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -375,6 +375,7 @@ lemma div_eq_one_iff_eq (hb : b ≠ 0) : a / b = 1 ↔ a = b := hb.isUnit.div_eq lemma div_mul_cancel_right₀ (hb : b ≠ 0) (a : G₀) : b / (a * b) = a⁻¹ := hb.isUnit.div_mul_cancel_right _ +set_option linter.deprecated false in @[deprecated div_mul_cancel_right₀] -- 2024-03-20 lemma div_mul_left (hb : b ≠ 0) : b / (a * b) = 1 / a := hb.isUnit.div_mul_left #align div_mul_left div_mul_left @@ -441,6 +442,7 @@ instance (priority := 100) CommGroupWithZero.toDivisionCommMonoid : lemma div_mul_cancel_left₀ (ha : a ≠ 0) (b : G₀) : a / (a * b) = b⁻¹ := ha.isUnit.div_mul_cancel_left _ +set_option linter.deprecated false in @[deprecated div_mul_cancel_left₀] -- 2024-03-22 lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := ha.isUnit.div_mul_right _ #align div_mul_right div_mul_right diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index 60dd064391dfb..938258158822b 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -550,7 +550,7 @@ theorem hasPullback_of_cover : HasPullback f g := instance affine_hasPullback {A B C : CommRingCat} (f : Spec.obj (Opposite.op A) ⟶ Spec.obj (Opposite.op C)) (g : Spec.obj (Opposite.op B) ⟶ Spec.obj (Opposite.op C)) : HasPullback f g := by - rw [← Spec.image_preimage f, ← Spec.image_preimage g] + rw [← Spec.map_preimage f, ← Spec.map_preimage g] exact ⟨⟨⟨_, isLimitOfHasPullbackOfPreservesLimit Spec (Spec.preimage f) (Spec.preimage g)⟩⟩⟩ #align algebraic_geometry.Scheme.pullback.affine_has_pullback AlgebraicGeometry.Scheme.Pullback.affine_hasPullback diff --git a/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean b/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean index 2199904012cb0..1b129df80bd1a 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/GammaCompN.lean @@ -191,7 +191,7 @@ lemma whiskerLeft_toKaroubi_N₂Γ₂_hom : whiskerLeft (toKaroubi (ChainComplex C ℕ)) N₂Γ₂.hom = N₂Γ₂ToKaroubiIso.hom ≫ N₁Γ₀.hom := by let e : _ ≅ toKaroubi (ChainComplex C ℕ) ⋙ 𝟭 _ := N₂Γ₂ToKaroubiIso ≪≫ N₁Γ₀ have h := ((whiskeringLeft _ _ (Karoubi (ChainComplex C ℕ))).obj - (toKaroubi (ChainComplex C ℕ))).image_preimage e.hom + (toKaroubi (ChainComplex C ℕ))).map_preimage e.hom dsimp only [whiskeringLeft, N₂Γ₂, Functor.preimageIso] at h ⊢ exact h diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 08c46af2a81c0..a3fa496d55be6 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -250,8 +250,8 @@ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_left {f : E → F} (g : F (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = g.compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := - (((hf.ftaylorSeriesWithin hs).continuousLinearMap_comp g).eq_ftaylor_series_of_uniqueDiffOn hi hs - hx).symm + (((hf.ftaylorSeriesWithin hs).continuousLinearMap_comp g).eq_iteratedFDerivWithin_of_uniqueDiffOn + hi hs hx).symm #align continuous_linear_map.iterated_fderiv_within_comp_left ContinuousLinearMap.iteratedFDerivWithin_comp_left /-- The iterated derivative of the composition with a linear map on the left is @@ -416,8 +416,8 @@ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_right {f : E → F} (g : G (hx : g x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : iteratedFDerivWithin 𝕜 i (f ∘ g) (g ⁻¹' s) x = (iteratedFDerivWithin 𝕜 i f s (g x)).compContinuousLinearMap fun _ => g := - (((hf.ftaylorSeriesWithin hs).compContinuousLinearMap g).eq_ftaylor_series_of_uniqueDiffOn hi h's - hx).symm + (((hf.ftaylorSeriesWithin hs).compContinuousLinearMap g).eq_iteratedFDerivWithin_of_uniqueDiffOn + hi h's hx).symm #align continuous_linear_map.iterated_fderiv_within_comp_right ContinuousLinearMap.iteratedFDerivWithin_comp_right /-- The iterated derivative within a set of the composition with a linear equiv on the right is @@ -1304,7 +1304,7 @@ theorem iteratedFDerivWithin_add_apply {f g : E → F} (hf : ContDiffOn 𝕜 i f iteratedFDerivWithin 𝕜 i (f + g) s x = iteratedFDerivWithin 𝕜 i f s x + iteratedFDerivWithin 𝕜 i g s x := Eq.symm <| ((hf.ftaylorSeriesWithin hu).add - (hg.ftaylorSeriesWithin hu)).eq_ftaylor_series_of_uniqueDiffOn le_rfl hu hx + (hg.ftaylorSeriesWithin hu)).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl hu hx #align iterated_fderiv_within_add_apply iteratedFDerivWithin_add_apply /-- The iterated derivative of the sum of two functions is the sum of the iterated derivatives. diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index c5cd200530eab..73133a99cd043 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -1036,14 +1036,14 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs have : p x m.succ = ftaylorSeriesWithin 𝕜 f s x m.succ := by change p x m.succ = iteratedFDerivWithin 𝕜 m.succ f s x rw [← iteratedFDerivWithin_inter_open o_open xo] - exact (Hp.mono ho).eq_ftaylor_series_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hx, xo⟩ + exact (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hx, xo⟩ rw [← this, ← hasFDerivWithinAt_inter (IsOpen.mem_nhds o_open xo)] have A : ∀ y ∈ s ∩ o, p y m = ftaylorSeriesWithin 𝕜 f s y m := by rintro y ⟨hy, yo⟩ change p y m = iteratedFDerivWithin 𝕜 m f s y rw [← iteratedFDerivWithin_inter_open o_open yo] exact - (Hp.mono ho).eq_ftaylor_series_of_uniqueDiffOn (WithTop.coe_le_coe.2 (Nat.le_succ m)) + (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn (WithTop.coe_le_coe.2 (Nat.le_succ m)) (hs.inter o_open) ⟨hy, yo⟩ exact ((Hp.mono ho).fderivWithin m (WithTop.coe_lt_coe.2 (lt_add_one m)) x ⟨hx, xo⟩).congr @@ -1060,7 +1060,7 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs rintro y ⟨hy, yo⟩ change p y m = iteratedFDerivWithin 𝕜 m f s y rw [← iteratedFDerivWithin_inter_open o_open yo] - exact (Hp.mono ho).eq_ftaylor_series_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hy, yo⟩ + exact (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hy, yo⟩ exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm #align cont_diff_on.ftaylor_series_within ContDiffOn.ftaylorSeriesWithin diff --git a/Mathlib/Analysis/Calculus/Darboux.lean b/Mathlib/Analysis/Calculus/Darboux.lean index 9dbf2e7479411..728683cf11792 100644 --- a/Mathlib/Analysis/Calculus/Darboux.lean +++ b/Mathlib/Analysis/Calculus/Darboux.lean @@ -35,7 +35,7 @@ theorem exists_hasDerivWithinAt_eq_of_gt_of_lt (hab : a ≤ b) intro x hx simpa using (hf x hx).sub ((hasDerivWithinAt_id x _).const_mul m) obtain ⟨c, cmem, hc⟩ : ∃ c ∈ Icc a b, IsMinOn g (Icc a b) c := - isCompact_Icc.exists_forall_le (nonempty_Icc.2 <| hab) fun x hx => (hg x hx).continuousWithinAt + isCompact_Icc.exists_isMinOn (nonempty_Icc.2 <| hab) fun x hx => (hg x hx).continuousWithinAt have cmem' : c ∈ Ioo a b := by rcases cmem.1.eq_or_lt with (rfl | hac) -- Show that `c` can't be equal to `a` diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index f72aca855c3a3..faaee5d29fc46 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -371,7 +371,7 @@ theorem exists_mem_frontier_isMaxOn_norm [FiniteDimensional ℂ E] {f : E → F} ∃ z ∈ frontier U, IsMaxOn (norm ∘ f) (closure U) z := by have hc : IsCompact (closure U) := hb.isCompact_closure obtain ⟨w, hwU, hle⟩ : ∃ w ∈ closure U, IsMaxOn (norm ∘ f) (closure U) w := - hc.exists_forall_ge hne.closure hd.continuousOn.norm + hc.exists_isMaxOn hne.closure hd.continuousOn.norm rw [closure_eq_interior_union_frontier, mem_union] at hwU cases' hwU with hwU hwU; rotate_left; · exact ⟨w, hwU, hle⟩ have : interior U ≠ univ := ne_top_of_le_ne_top hc.ne_univ interior_subset_closure diff --git a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean index 4b16c0b5ae4dc..0979b0dd21cd6 100644 --- a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean +++ b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean @@ -81,7 +81,7 @@ theorem norm_cderiv_lt (hr : 0 < r) (hfM : ∀ w ∈ sphere z r, ‖f w‖ < M) obtain ⟨L, hL1, hL2⟩ : ∃ L < M, ∀ w ∈ sphere z r, ‖f w‖ ≤ L := by have e1 : (sphere z r).Nonempty := NormedSpace.sphere_nonempty.mpr hr.le have e2 : ContinuousOn (fun w => ‖f w‖) (sphere z r) := continuous_norm.comp_continuousOn hf - obtain ⟨x, hx, hx'⟩ := (isCompact_sphere z r).exists_forall_ge e1 e2 + obtain ⟨x, hx, hx'⟩ := (isCompact_sphere z r).exists_isMaxOn e1 e2 exact ⟨‖f x‖, hfM x hx, hx'⟩ exact (norm_cderiv_le hr hL2).trans_lt ((div_lt_div_right hr).mpr hL1) #align complex.norm_cderiv_lt Complex.norm_cderiv_lt diff --git a/Mathlib/Analysis/Complex/OpenMapping.lean b/Mathlib/Analysis/Complex/OpenMapping.lean index 74c43bfa7ebcd..6e7a5cc3cb48f 100644 --- a/Mathlib/Analysis/Complex/OpenMapping.lean +++ b/Mathlib/Analysis/Complex/OpenMapping.lean @@ -100,9 +100,9 @@ theorem AnalyticAt.eventually_constant_or_nhds_le_map_nhds_aux (hf : AnalyticAt have h8 : (sphere z₀ r).Nonempty := NormedSpace.sphere_nonempty.mpr hr.le have h9 : ContinuousOn (fun x => ‖f x - f z₀‖) (sphere z₀ r) := continuous_norm.comp_continuousOn ((h6.sub_const (f z₀)).continuousOn_ball.mono sphere_subset_closedBall) - obtain ⟨x, hx, hfx⟩ := (isCompact_sphere z₀ r).exists_forall_le h8 h9 + obtain ⟨x, hx, hfx⟩ := (isCompact_sphere z₀ r).exists_isMinOn h8 h9 refine ⟨‖f x - f z₀‖ / 2, half_pos (norm_sub_pos_iff.mpr (h7 x hx)), ?_⟩ - exact (h6.ball_subset_image_closedBall hr (fun z hz => hfx z hz) (not_eventually.mp h)).trans + exact (h6.ball_subset_image_closedBall hr (fun z hz => hfx hz) (not_eventually.mp h)).trans (image_subset f (closedBall_subset_closedBall inf_le_right)) #align analytic_at.eventually_constant_or_nhds_le_map_nhds_aux AnalyticAt.eventually_constant_or_nhds_le_map_nhds_aux diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index 6c136ca454d19..e0941b88b8f74 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -709,7 +709,7 @@ theorem right_half_plane_of_tendsto_zero_on_real (hd : DiffContOnCl ℂ f {z | 0 rcases h₀ with ⟨x₀, hx₀, hne⟩ have hlt : ‖(0 : E)‖ < ‖f x₀‖ := by rwa [norm_zero, norm_pos_iff] suffices ∀ᶠ x : ℝ in cocompact ℝ ⊓ 𝓟 (Ici 0), ‖f x‖ ≤ ‖f x₀‖ by - simpa only [exists_prop] using hfc.norm.exists_forall_ge' isClosed_Ici hx₀ this + simpa only [exists_prop] using hfc.norm.exists_isMaxOn' isClosed_Ici hx₀ this rw [cocompact_eq_atBot_atTop, inf_sup_right, (disjoint_atBot_principal_Ici (0 : ℝ)).eq_bot, bot_sup_eq] exact (hre.norm.eventually <| ge_mem_nhds hlt).filter_mono inf_le_left diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index 2987e2c9d1500..09cc83ae21ae7 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -609,7 +609,7 @@ theorem gauge_closedBall (hr : 0 ≤ r) (x : E) : gauge (closedBall (0 : E) r) x filter_upwards [self_mem_nhdsWithin] with R hR rw [← gauge_ball (hr.trans hR.out.le)] refine gauge_mono ?_ (closedBall_subset_ball hR) _ - exact (absorbent_ball_zero hr').subset ball_subset_closedBall + exact (absorbent_ball_zero hr').mono ball_subset_closedBall theorem mul_gauge_le_norm (hs : Metric.ball (0 : E) r ⊆ s) : r * gauge s x ≤ ‖x‖ := by obtain hr | hr := le_or_lt r 0 @@ -624,7 +624,7 @@ theorem Convex.lipschitzWith_gauge {r : ℝ≥0} (hc : Convex ℝ s) (hr : 0 < r LipschitzWith.of_le_add_mul _ fun x y => calc gauge s x = gauge s (y + (x - y)) := by simp - _ ≤ gauge s y + gauge s (x - y) := gauge_add_le hc (this.subset hs) _ _ + _ ≤ gauge s y + gauge s (x - y) := gauge_add_le hc (this.mono hs) _ _ _ ≤ gauge s y + ‖x - y‖ / r := add_le_add_left ((gauge_mono this hs (x - y)).trans_eq (gauge_ball hr.le _)) _ _ = gauge s y + r⁻¹ * dist x y := by rw [dist_eq_norm, div_eq_inv_mul, NNReal.coe_inv] diff --git a/Mathlib/Analysis/Convex/KreinMilman.lean b/Mathlib/Analysis/Convex/KreinMilman.lean index 3c8d0a4cab7a4..551032d8bda83 100644 --- a/Mathlib/Analysis/Convex/KreinMilman.lean +++ b/Mathlib/Analysis/Convex/KreinMilman.lean @@ -69,7 +69,7 @@ theorem IsCompact.extremePoints_nonempty (hscomp : IsCompact s) (hsnemp : s.None by_contra hyx obtain ⟨l, hl⟩ := geometric_hahn_banach_point_point hyx obtain ⟨z, hzt, hz⟩ := - (hscomp.of_isClosed_subset htclos hst.1).exists_forall_ge ⟨x, hxt⟩ + (hscomp.of_isClosed_subset htclos hst.1).exists_isMaxOn ⟨x, hxt⟩ l.continuous.continuousOn have h : IsExposed ℝ t ({ z ∈ t | ∀ w ∈ t, l w ≤ l z }) := fun _ => ⟨l, rfl⟩ rw [← hBmin ({ z ∈ t | ∀ w ∈ t, l w ≤ l z }) @@ -100,7 +100,7 @@ theorem closure_convexHull_extremePoints (hscomp : IsCompact s) (hAconv : Convex obtain ⟨l, r, hlr, hrx⟩ := geometric_hahn_banach_closed_point (convex_convexHull _ _).closure isClosed_closure hxt have h : IsExposed ℝ s ({ y ∈ s | ∀ z ∈ s, l z ≤ l y }) := fun _ => ⟨l, rfl⟩ - obtain ⟨z, hzA, hz⟩ := hscomp.exists_forall_ge ⟨x, hxA⟩ l.continuous.continuousOn + obtain ⟨z, hzA, hz⟩ := hscomp.exists_isMaxOn ⟨x, hxA⟩ l.continuous.continuousOn obtain ⟨y, hy⟩ := (h.isCompact hscomp).extremePoints_nonempty ⟨z, hzA, hz⟩ linarith [hlr _ (subset_closure <| subset_convexHull _ _ <| h.isExtreme.extremePoints_subset_extremePoints hy), hy.1.2 x hxA] diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 25c6a6d22b838..430c3b34d436f 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -256,13 +256,13 @@ theorem AEStronglyMeasurable.convolution_integrand_snd (hf : AEStronglyMeasurabl (hg : AEStronglyMeasurable g μ) (x : G) : AEStronglyMeasurable (fun t => L (f t) (g (x - t))) μ := hf.convolution_integrand_snd' L <| - hg.mono' <| (quasiMeasurePreserving_sub_left_of_right_invariant μ x).absolutelyContinuous + hg.mono_ac <| (quasiMeasurePreserving_sub_left_of_right_invariant μ x).absolutelyContinuous #align measure_theory.ae_strongly_measurable.convolution_integrand_snd MeasureTheory.AEStronglyMeasurable.convolution_integrand_snd theorem AEStronglyMeasurable.convolution_integrand_swap_snd (hf : AEStronglyMeasurable f μ) (hg : AEStronglyMeasurable g μ) (x : G) : AEStronglyMeasurable (fun t => L (f (x - t)) (g t)) μ := - (hf.mono' + (hf.mono_ac (quasiMeasurePreserving_sub_left_of_right_invariant μ x).absolutelyContinuous).convolution_integrand_swap_snd' L hg @@ -274,7 +274,7 @@ theorem ConvolutionExistsAt.ofNorm {x₀ : G} (hmf : AEStronglyMeasurable f μ) (hmg : AEStronglyMeasurable g μ) : ConvolutionExistsAt f g x₀ L μ := h.ofNorm' L hmf <| - hmg.mono' (quasiMeasurePreserving_sub_left_of_right_invariant μ x₀).absolutelyContinuous + hmg.mono_ac (quasiMeasurePreserving_sub_left_of_right_invariant μ x₀).absolutelyContinuous #align convolution_exists_at.of_norm MeasureTheory.ConvolutionExistsAt.ofNorm end Left @@ -288,7 +288,7 @@ theorem AEStronglyMeasurable.convolution_integrand (hf : AEStronglyMeasurable f (hg : AEStronglyMeasurable g μ) : AEStronglyMeasurable (fun p : G × G => L (f p.2) (g (p.1 - p.2))) (μ.prod ν) := hf.convolution_integrand' L <| - hg.mono' (quasiMeasurePreserving_sub_of_right_invariant μ ν).absolutelyContinuous + hg.mono_ac (quasiMeasurePreserving_sub_of_right_invariant μ ν).absolutelyContinuous #align measure_theory.ae_strongly_measurable.convolution_integrand MeasureTheory.AEStronglyMeasurable.convolution_integrand theorem Integrable.convolution_integrand (hf : Integrable f ν) (hg : Integrable g μ) : @@ -383,7 +383,7 @@ theorem _root_.BddAbove.convolutionExistsAt [MeasurableAdd₂ G] [SigmaFinite μ refine' BddAbove.convolutionExistsAt' L _ hs h2s hf _ · simp_rw [← sub_eq_neg_add, hbg] · have : AEStronglyMeasurable g (map (fun t : G => x₀ - t) μ) := - hmg.mono' (quasiMeasurePreserving_sub_left_of_right_invariant μ x₀).absolutelyContinuous + hmg.mono_ac (quasiMeasurePreserving_sub_left_of_right_invariant μ x₀).absolutelyContinuous apply this.mono_measure exact map_mono restrict_le_self (measurable_const.sub measurable_id') #align bdd_above.convolution_exists_at BddAbove.convolutionExistsAt @@ -968,7 +968,8 @@ theorem convolution_assoc (hL : ∀ (x : E) (y : E') (z : E''), L₂ (L x y) z = (μ.prod ν) := by refine' L₃.aestronglyMeasurable_comp₂ hf.snd _ refine' L₄.aestronglyMeasurable_comp₂ hg.fst _ - refine' (hk.mono' _).comp_measurable ((measurable_const.sub measurable_snd).sub measurable_fst) + refine' (hk.mono_ac _).comp_measurable + ((measurable_const.sub measurable_snd).sub measurable_fst) refine' QuasiMeasurePreserving.absolutelyContinuous _ refine' QuasiMeasurePreserving.prod_of_left ((measurable_const.sub measurable_snd).sub measurable_fst) (eventually_of_forall fun y => _) diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 03266bd2403e6..5f251ff7d496f 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -655,7 +655,7 @@ lemma _root_.ContinuousLinearMap.hasTemperateGrowth (f : E →L[ℝ] F) : apply Function.HasTemperateGrowth.of_fderiv ?_ f.differentiable (k := 1) (C := ‖f‖) (fun x ↦ ?_) · have : fderiv ℝ f = fun _ ↦ f := by ext1 v; simp only [ContinuousLinearMap.fderiv] simpa [this] using .const _ - · exact (f.le_op_norm x).trans (by simp [mul_add]) + · exact (f.le_opNorm x).trans (by simp [mul_add]) variable [NormedAddCommGroup D] [NormedSpace ℝ D] variable [MeasurableSpace D] [BorelSpace D] [SecondCountableTopology D] [FiniteDimensional ℝ D] diff --git a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean index 73535314f7688..bf25b941f24f8 100644 --- a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean +++ b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean @@ -248,7 +248,7 @@ theorem hasEigenvalue_iSup_of_finiteDimensional (hT : T.IsSymmetric) : have H₂ : (sphere (0 : E) ‖x‖).Nonempty := ⟨x, by simp⟩ -- key point: in finite dimension, a continuous function on the sphere has a max obtain ⟨x₀, hx₀', hTx₀⟩ := - H₁.exists_forall_ge H₂ T'.val.reApplyInnerSelf_continuous.continuousOn + H₁.exists_isMaxOn H₂ T'.val.reApplyInnerSelf_continuous.continuousOn have hx₀ : ‖x₀‖ = ‖x‖ := by simpa using hx₀' have : IsMaxOn T'.val.reApplyInnerSelf (sphere 0 ‖x₀‖) x₀ := by simpa only [← hx₀] using hTx₀ have hx₀_ne : x₀ ≠ 0 := by @@ -268,7 +268,7 @@ theorem hasEigenvalue_iInf_of_finiteDimensional (hT : T.IsSymmetric) : have H₂ : (sphere (0 : E) ‖x‖).Nonempty := ⟨x, by simp⟩ -- key point: in finite dimension, a continuous function on the sphere has a min obtain ⟨x₀, hx₀', hTx₀⟩ := - H₁.exists_forall_le H₂ T'.val.reApplyInnerSelf_continuous.continuousOn + H₁.exists_isMinOn H₂ T'.val.reApplyInnerSelf_continuous.continuousOn have hx₀ : ‖x₀‖ = ‖x‖ := by simpa using hx₀' have : IsMinOn T'.val.reApplyInnerSelf (sphere 0 ‖x₀‖) x₀ := by simpa only [← hx₀] using hTx₀ have hx₀_ne : x₀ ≠ 0 := by diff --git a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean b/Mathlib/Analysis/NormedSpace/FiniteDimension.lean index ebee3f70a8cf9..38fd38c9c5392 100644 --- a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean +++ b/Mathlib/Analysis/NormedSpace/FiniteDimension.lean @@ -299,7 +299,7 @@ theorem Basis.opNNNorm_le {ι : Type*} [Fintype ι] (v : Basis ι 𝕜 E) {u : E set φ := v.equivFunL.toContinuousLinearMap calc ‖u e‖₊ = ‖u (∑ i, v.equivFun e i • v i)‖₊ := by rw [v.sum_equivFun] - _ = ‖∑ i, v.equivFun e i • (u <| v i)‖₊ := by simp [u.map_sum, LinearMap.map_smul] + _ = ‖∑ i, v.equivFun e i • (u <| v i)‖₊ := by simp [map_sum, LinearMap.map_smul] _ ≤ ∑ i, ‖v.equivFun e i • (u <| v i)‖₊ := nnnorm_sum_le _ _ _ = ∑ i, ‖v.equivFun e i‖₊ * ‖u (v i)‖₊ := by simp only [nnnorm_smul] _ ≤ ∑ i, ‖v.equivFun e i‖₊ * M := by gcongr; apply hu diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean index 6bd45c3b2ace0..4da5ec28f5b95 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean @@ -159,11 +159,13 @@ theorem geometric_hahn_banach_compact_closed (hs₁ : Convex ℝ s) (hs₂ : IsC · exact ⟨0, 1, 2, fun a _ha => by norm_num, by norm_num, by simp⟩ obtain ⟨U, V, hU, hV, hU₁, hV₁, sU, tV, disj'⟩ := disj.exists_open_convexes hs₁ hs₂ ht₁ ht₂ obtain ⟨f, u, hf₁, hf₂⟩ := geometric_hahn_banach_open_open hU₁ hU hV₁ hV disj' - obtain ⟨x, hx₁, hx₂⟩ := hs₂.exists_forall_ge hs f.continuous.continuousOn + obtain ⟨x, hx₁, hx₂⟩ := hs₂.exists_isMaxOn hs f.continuous.continuousOn have : f x < u := hf₁ x (sU hx₁) exact - ⟨f, (f x + u) / 2, u, fun a ha => by linarith [hx₂ a ha], by linarith, fun b hb => - hf₂ b (tV hb)⟩ + ⟨f, (f x + u) / 2, u, + fun a ha => by have := hx₂ ha; dsimp at this; linarith, + by linarith, + fun b hb => hf₂ b (tV hb)⟩ #align geometric_hahn_banach_compact_closed geometric_hahn_banach_compact_closed /-- A version of the **Hahn-Banach theorem**: given disjoint convex sets `s`, `t` where `s` is diff --git a/Mathlib/Analysis/NormedSpace/Spectrum.lean b/Mathlib/Analysis/NormedSpace/Spectrum.lean index 8b495d9258443..0b1c6b19c56f1 100644 --- a/Mathlib/Analysis/NormedSpace/Spectrum.lean +++ b/Mathlib/Analysis/NormedSpace/Spectrum.lean @@ -156,7 +156,7 @@ theorem spectralRadius_le_nnnorm [NormOneClass A] (a : A) : spectralRadius 𝕜 theorem exists_nnnorm_eq_spectralRadius_of_nonempty [ProperSpace 𝕜] {a : A} (ha : (σ a).Nonempty) : ∃ k ∈ σ a, (‖k‖₊ : ℝ≥0∞) = spectralRadius 𝕜 a := by - obtain ⟨k, hk, h⟩ := (spectrum.isCompact a).exists_forall_ge ha continuous_nnnorm.continuousOn + obtain ⟨k, hk, h⟩ := (spectrum.isCompact a).exists_isMaxOn ha continuous_nnnorm.continuousOn exact ⟨k, hk, le_antisymm (le_iSup₂ (α := ℝ≥0∞) k hk) (iSup₂_le <| mod_cast h)⟩ #align spectrum.exists_nnnorm_eq_spectral_radius_of_nonempty spectrum.exists_nnnorm_eq_spectralRadius_of_nonempty diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index d4a529b5d3f83..04f163f4ddd60 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -1039,19 +1039,19 @@ protected theorem absorbent_ball_zero (hr : 0 < r) : Absorbent 𝕜 (ball p (0 : /-- Closed seminorm-balls at the origin are absorbent. -/ protected theorem absorbent_closedBall_zero (hr : 0 < r) : Absorbent 𝕜 (closedBall p (0 : E) r) := - (p.absorbent_ball_zero hr).subset (p.ball_subset_closedBall _ _) + (p.absorbent_ball_zero hr).mono (p.ball_subset_closedBall _ _) #align seminorm.absorbent_closed_ball_zero Seminorm.absorbent_closedBall_zero /-- Seminorm-balls containing the origin are absorbent. -/ protected theorem absorbent_ball (hpr : p x < r) : Absorbent 𝕜 (ball p x r) := by - refine' (p.absorbent_ball_zero <| sub_pos.2 hpr).subset fun y hy => _ + refine' (p.absorbent_ball_zero <| sub_pos.2 hpr).mono fun y hy => _ rw [p.mem_ball_zero] at hy exact p.mem_ball.2 ((map_sub_le_add p _ _).trans_lt <| add_lt_of_lt_sub_right hy) #align seminorm.absorbent_ball Seminorm.absorbent_ball /-- Seminorm-balls containing the origin are absorbent. -/ protected theorem absorbent_closedBall (hpr : p x < r) : Absorbent 𝕜 (closedBall p x r) := by - refine' (p.absorbent_closedBall_zero <| sub_pos.2 hpr).subset fun y hy => _ + refine' (p.absorbent_closedBall_zero <| sub_pos.2 hpr).mono fun y hy => _ rw [p.mem_closedBall_zero] at hy exact p.mem_closedBall.2 ((map_sub_le_add p _ _).trans <| add_le_of_le_sub_right hy) #align seminorm.absorbent_closed_ball Seminorm.absorbent_closedBall diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean index 8fd6a313f1491..e3b3defc53be6 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean @@ -281,7 +281,7 @@ noncomputable def mkIdInitial [T.Full] [T.Faithful] : IsInitial (mk (𝟙 (T.obj apply CommaMorphism.ext · aesop_cat · apply T.map_injective - simpa only [homMk_right, T.image_preimage, ← w m] using (Category.id_comp _).symm + simpa only [homMk_right, T.map_preimage, ← w m] using (Category.id_comp _).symm #align category_theory.structured_arrow.mk_id_initial CategoryTheory.StructuredArrow.mkIdInitial variable {A : Type u₃} [Category.{v₃} A] {B : Type u₄} [Category.{v₄} B] @@ -645,7 +645,7 @@ noncomputable def mkIdTerminal [S.Full] [S.Faithful] : IsTerminal (mk (𝟙 (S.o rintro c m - ext apply S.map_injective - simpa only [homMk_left, S.image_preimage, ← w m] using (Category.comp_id _).symm + simpa only [homMk_left, S.map_preimage, ← w m] using (Category.comp_id _).symm #align category_theory.costructured_arrow.mk_id_terminal CategoryTheory.CostructuredArrow.mkIdTerminal variable {A : Type u₃} [Category.{v₃} A] {B : Type u₄} [Category.{v₄} B] diff --git a/Mathlib/CategoryTheory/Monoidal/Discrete.lean b/Mathlib/CategoryTheory/Monoidal/Discrete.lean index 641d322d3126d..b3920af26f985 100644 --- a/Mathlib/CategoryTheory/Monoidal/Discrete.lean +++ b/Mathlib/CategoryTheory/Monoidal/Discrete.lean @@ -51,7 +51,7 @@ discrete monoidal categories. def Discrete.monoidalFunctor (F : M →* N) : MonoidalFunctor (Discrete M) (Discrete N) where obj X := Discrete.mk (F X.as) - map f := Discrete.eqToHom (F.congr_arg (eq_of_hom f)) + map f := Discrete.eqToHom (DFunLike.congr_arg F (eq_of_hom f)) ε := Discrete.eqToHom F.map_one.symm μ X Y := Discrete.eqToHom (F.map_mul X.as Y.as).symm #align category_theory.discrete.monoidal_functor CategoryTheory.Discrete.monoidalFunctor diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index 5630aedf19407..192127da3295e 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -95,30 +95,30 @@ theorem sort_sorted_gt (s : Finset α) : List.Sorted (· > ·) (sort (· ≥ ·) (sort_sorted _ _).gt_of_ge (sort_nodup _ _) theorem sorted_zero_eq_min'_aux (s : Finset α) (h : 0 < (s.sort (· ≤ ·)).length) (H : s.Nonempty) : - (s.sort (· ≤ ·)).nthLe 0 h = s.min' H := by + (s.sort (· ≤ ·)).get ⟨0, h⟩ = s.min' H := by let l := s.sort (· ≤ ·) apply le_antisymm · have : s.min' H ∈ l := (Finset.mem_sort (α := α) (· ≤ ·)).mpr (s.min'_mem H) obtain ⟨i, hi⟩ : ∃ i, l.get i = s.min' H := List.mem_iff_get.1 this rw [← hi] - exact (s.sort_sorted (· ≤ ·)).rel_nthLe_of_le _ _ (Nat.zero_le i) + exact (s.sort_sorted (· ≤ ·)).rel_get_of_le (Nat.zero_le i) · have : l.get ⟨0, h⟩ ∈ s := (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l 0 h) exact s.min'_le _ this #align finset.sorted_zero_eq_min'_aux Finset.sorted_zero_eq_min'_aux theorem sorted_zero_eq_min' {s : Finset α} {h : 0 < (s.sort (· ≤ ·)).length} : - (s.sort (· ≤ ·)).nthLe 0 h = s.min' (card_pos.1 <| by rwa [length_sort] at h) := + (s.sort (· ≤ ·)).get ⟨0, h⟩ = s.min' (card_pos.1 <| by rwa [length_sort] at h) := sorted_zero_eq_min'_aux _ _ _ #align finset.sorted_zero_eq_min' Finset.sorted_zero_eq_min' theorem min'_eq_sorted_zero {s : Finset α} {h : s.Nonempty} : - s.min' h = (s.sort (· ≤ ·)).nthLe 0 (by rw [length_sort]; exact card_pos.2 h) := + s.min' h = (s.sort (· ≤ ·)).get ⟨0, (by rw [length_sort]; exact card_pos.2 h)⟩ := (sorted_zero_eq_min'_aux _ _ _).symm #align finset.min'_eq_sorted_zero Finset.min'_eq_sorted_zero theorem sorted_last_eq_max'_aux (s : Finset α) (h : (s.sort (· ≤ ·)).length - 1 < (s.sort (· ≤ ·)).length) (H : s.Nonempty) : - (s.sort (· ≤ ·)).nthLe ((s.sort (· ≤ ·)).length - 1) h = s.max' H := by + (s.sort (· ≤ ·)).get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ = s.max' H := by let l := s.sort (· ≤ ·) apply le_antisymm · have : l.get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ ∈ s := @@ -127,20 +127,20 @@ theorem sorted_last_eq_max'_aux (s : Finset α) · have : s.max' H ∈ l := (Finset.mem_sort (α := α) (· ≤ ·)).mpr (s.max'_mem H) obtain ⟨i, hi⟩ : ∃ i, l.get i = s.max' H := List.mem_iff_get.1 this rw [← hi] - exact (s.sort_sorted (· ≤ ·)).rel_nthLe_of_le _ _ (Nat.le_sub_one_of_lt i.prop) + exact (s.sort_sorted (· ≤ ·)).rel_get_of_le (Nat.le_sub_one_of_lt i.prop) #align finset.sorted_last_eq_max'_aux Finset.sorted_last_eq_max'_aux theorem sorted_last_eq_max' {s : Finset α} {h : (s.sort (· ≤ ·)).length - 1 < (s.sort (· ≤ ·)).length} : - (s.sort (· ≤ ·)).nthLe ((s.sort (· ≤ ·)).length - 1) h = + (s.sort (· ≤ ·)).get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ = s.max' (by rw [length_sort] at h; exact card_pos.1 (lt_of_le_of_lt bot_le h)) := sorted_last_eq_max'_aux _ _ _ #align finset.sorted_last_eq_max' Finset.sorted_last_eq_max' theorem max'_eq_sorted_last {s : Finset α} {h : s.Nonempty} : s.max' h = - (s.sort (· ≤ ·)).nthLe ((s.sort (· ≤ ·)).length - 1) - (by simpa using Nat.sub_lt (card_pos.mpr h) zero_lt_one) := + (s.sort (· ≤ ·)).get ⟨(s.sort (· ≤ ·)).length - 1, + by simpa using Nat.sub_lt (card_pos.mpr h) zero_lt_one⟩ := (sorted_last_eq_max'_aux _ _ _).symm #align finset.max'_eq_sorted_last Finset.max'_eq_sorted_last @@ -174,7 +174,7 @@ theorem orderIsoOfFin_symm_apply (s : Finset α) {k : ℕ} (h : s.card = k) (x : theorem orderEmbOfFin_apply (s : Finset α) {k : ℕ} (h : s.card = k) (i : Fin k) : s.orderEmbOfFin h i = - (s.sort (· ≤ ·)).nthLe i (by rw [length_sort, h]; exact i.2) := + (s.sort (· ≤ ·)).get ⟨i, by rw [length_sort, h]; exact i.2⟩ := rfl #align finset.order_emb_of_fin_apply Finset.orderEmbOfFin_apply diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 832b4e69a179c..d394b5dcb2069 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -538,7 +538,7 @@ theorem mapDomain_finset_sum {f : α → β} {s : Finset ι} {v : ι → α → theorem mapDomain_sum [Zero N] {f : α → β} {s : α →₀ N} {v : α → N → α →₀ M} : mapDomain f (s.sum v) = s.sum fun a b => mapDomain f (v a b) := - (mapDomain.addMonoidHom f : (α →₀ M) →+ β →₀ M).map_finsupp_sum _ _ + map_finsupp_sum (mapDomain.addMonoidHom f : (α →₀ M) →+ β →₀ M) _ _ #align finsupp.map_domain_sum Finsupp.mapDomain_sum theorem mapDomain_support [DecidableEq β] {f : α → β} {s : α →₀ M} : diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 22eb8d4054b1d..ad7740ee54045 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -2949,6 +2949,8 @@ theorem span_eq_take_drop (l : List α) : span p l = (takeWhile p l, dropWhile p #align list.take_while_append_drop List.takeWhile_append_dropWhile +-- TODO update to use `get` instead of `nthLe` +set_option linter.deprecated false in theorem dropWhile_nthLe_zero_not (l : List α) (hl : 0 < (l.dropWhile p).length) : ¬p ((l.dropWhile p).nthLe 0 hl) := by induction' l with hd tl IH @@ -2994,6 +2996,8 @@ theorem takeWhile_eq_self_iff : takeWhile p l = l ↔ ∀ x ∈ l, p x := by · by_cases hp : p x <;> simp [hp, takeWhile_cons, IH] #align list.take_while_eq_self_iff List.takeWhile_eq_self_iff +-- TODO update to use `get` instead of `nthLe` +set_option linter.deprecated false in @[simp] theorem takeWhile_eq_nil_iff : takeWhile p l = [] ↔ ∀ hl : 0 < l.length, ¬p (l.nthLe 0 hl) := by induction' l with x xs IH diff --git a/Mathlib/Data/List/DropRight.lean b/Mathlib/Data/List/DropRight.lean index af4e4af4df329..d6ff682d0a3c2 100644 --- a/Mathlib/Data/List/DropRight.lean +++ b/Mathlib/Data/List/DropRight.lean @@ -141,7 +141,7 @@ theorem rdropWhile_eq_nil_iff : rdropWhile p l = [] ↔ ∀ x ∈ l, p x := by s -- it is in this file because it requires `List.Infix` @[simp] -theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p (l.nthLe 0 hl) := by +theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p (l.get ⟨0, hl⟩) := by cases' l with hd tl · simp only [dropWhile, true_iff] intro h @@ -150,13 +150,13 @@ theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p · rw [dropWhile] refine' ⟨fun h => _, fun h => _⟩ · intro _ H - rw [nthLe, get] at H + rw [get] at H refine' (cons_ne_self hd tl) (Sublist.antisymm _ (sublist_cons _ _)) rw [← h] simp only [H] exact List.IsSuffix.sublist (dropWhile_suffix p) · have := h (by simp only [length, Nat.succ_pos]) - rw [nthLe, get] at this + rw [get] at this simp_rw [this] #align list.drop_while_eq_self_iff List.dropWhile_eq_self_iff @@ -168,10 +168,10 @@ theorem rdropWhile_eq_self_iff : rdropWhile p l = l ↔ ∀ hl : l ≠ [], ¬p ( refine' ⟨fun h hl => _, fun h hl => _⟩ · rw [← length_pos, ← length_reverse] at hl have := h hl - rwa [nthLe, get_reverse'] at this + rwa [get_reverse'] at this · rw [length_reverse, length_pos] at hl have := h hl - rwa [nthLe, get_reverse'] + rwa [get_reverse'] #align list.rdrop_while_eq_self_iff List.rdropWhile_eq_self_iff variable (p) (l) diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index 9b6a5ddd36f3d..492b6d05737e8 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -162,6 +162,7 @@ theorem Forall₂.get : | _, _, Forall₂.cons ha _, 0, _, _ => ha | _, _, Forall₂.cons _ hl, succ _, _, _ => hl.get _ _ +set_option linter.deprecated false in @[deprecated (since := "2024-05-05")] theorem Forall₂.nthLe {x y} (h : Forall₂ R x y) ⦃i : ℕ⦄ (hx : i < x.length) (hy : i < y.length) : R (x.nthLe i hx) (y.nthLe i hy) := h.get hx hy #align list.forall₂.nth_le List.Forall₂.nthLe @@ -175,6 +176,7 @@ theorem forall₂_of_length_eq_of_get : (forall₂_of_length_eq_of_get (succ.inj hl) fun i h₁ h₂ => h i.succ (succ_lt_succ h₁) (succ_lt_succ h₂)) +set_option linter.deprecated false in @[deprecated (since := "2024-05-05")] theorem forall₂_of_length_eq_of_nthLe {x y} (H : x.length = y.length) (H' : ∀ i h₁ h₂, R (x.nthLe i h₁) (y.nthLe i h₂)) : Forall₂ R x y := forall₂_of_length_eq_of_get H H' @@ -184,6 +186,7 @@ theorem forall₂_iff_get {l₁ : List α} {l₂ : List β} : Forall₂ R l₁ l₂ ↔ l₁.length = l₂.length ∧ ∀ i h₁ h₂, R (l₁.get ⟨i, h₁⟩) (l₂.get ⟨i, h₂⟩) := ⟨fun h => ⟨h.length_eq, h.get⟩, fun h => forall₂_of_length_eq_of_get h.1 h.2⟩ +set_option linter.deprecated false in @[deprecated (since := "2024-05-05")] theorem forall₂_iff_nthLe {l₁ : List α} {l₂ : List β} : Forall₂ R l₁ l₂ ↔ l₁.length = l₂.length ∧ ∀ i h₁ h₂, R (l₁.nthLe i h₁) (l₂.nthLe i h₂) := forall₂_iff_get diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index c88ed84cd6e44..91f3b8c66cf09 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -211,6 +211,7 @@ theorem length_mapIdx {α β} (l : List α) (f : ℕ → α → β) : (l.mapIdx theorem mapIdx_eq_nil {α β} {f : ℕ → α → β} {l : List α} : List.mapIdx f l = [] ↔ l = [] := by rw [List.mapIdx_eq_enum_map, List.map_eq_nil, List.enum_eq_nil] +set_option linter.deprecated false in @[simp, deprecated] -- 2023-02-11 theorem nthLe_mapIdx {α β} (l : List α) (f : ℕ → α → β) (i : ℕ) (h : i < l.length) (h' : i < (l.mapIdx f).length := h.trans_le (l.length_mapIdx f).ge) : diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 286fce2e7ccd5..35cf3193ba862 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -335,7 +335,7 @@ protected theorem IsPrefix.filterMap (h : l₁ <+: l₂) (f : α → Option β) protected theorem IsPrefix.reduceOption {l₁ l₂ : List (Option α)} (h : l₁ <+: l₂) : l₁.reduceOption <+: l₂.reduceOption := - h.filter_map id + h.filterMap id #align list.is_prefix.reduce_option List.IsPrefix.reduceOption #align list.is_prefix.filter List.IsPrefix.filter diff --git a/Mathlib/Data/List/InsertNth.lean b/Mathlib/Data/List/InsertNth.lean index 5b5fd11213e8b..561811b6a0b3e 100644 --- a/Mathlib/Data/List/InsertNth.lean +++ b/Mathlib/Data/List/InsertNth.lean @@ -147,6 +147,7 @@ theorem get_insertNth_of_lt (l : List α) (x : α) (n k : ℕ) (hn : k < n) (hk · rw [Nat.succ_lt_succ_iff] at hn simpa using IH _ _ hn _ +set_option linter.deprecated false in @[deprecated get_insertNth_of_lt] -- 2023-01-05 theorem nthLe_insertNth_of_lt : ∀ (l : List α) (x : α) (n k : ℕ), k < n → ∀ (hk : k < l.length) (hk' : k < (insertNth n x l).length := hk.trans_le (length_le_length_insertNth _ _ _)), @@ -166,6 +167,7 @@ theorem get_insertNth_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length · simp only [Nat.succ_le_succ_iff, length] at hn simpa using IH _ hn +set_option linter.deprecated false in @[simp, deprecated get_insertNth_self] theorem nthLe_insertNth_self (l : List α) (x : α) (n : ℕ) (hn : n ≤ l.length) (hn' : n < (insertNth n x l).length := (by rwa [length_insertNth _ _ hn, Nat.lt_succ_iff])) : diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index a39b9497f5244..dd276cc229bdb 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -40,7 +40,7 @@ for a version giving an equivalence when there is decidable equality. -/ @[simps] def getBijectionOfForallMemList (l : List α) (nd : l.Nodup) (h : ∀ x : α, x ∈ l) : { f : Fin l.length → α // Function.Bijective f } := - ⟨fun i => l.get i, fun _ _ h => Fin.ext <| (nd.nthLe_inj_iff _ _).1 h, + ⟨fun i => l.get i, fun _ _ h => nd.get_inj_iff.1 h, fun x => let ⟨i, hl⟩ := List.mem_iff_get.1 (h x) ⟨i, hl⟩⟩ @@ -233,6 +233,7 @@ theorem duplicate_iff_exists_distinct_get {l : List α} {x : α} : · simpa using h · simpa using h' +set_option linter.deprecated false in /-- An element `x : α` of `l : List α` is a duplicate iff it can be found at two distinct indices `n m : ℕ` inside the list `l`. -/ diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 956bfc3a26b09..e167aebee84d8 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -778,6 +778,7 @@ theorem get_permutations'Aux (s : List α) (x : α) (n : ℕ) · simpa [get] using IH _ _ #align list.nth_le_permutations'_aux List.get_permutations'Aux +set_option linter.deprecated false in @[deprecated get_permutations'Aux] -- 2024-04-23 theorem nthLe_permutations'Aux (s : List α) (x : α) (n : ℕ) (hn : n < length (permutations'Aux x s)) : diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 4986a37274d06..7440e39aa3309 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -135,6 +135,8 @@ theorem Sorted.rel_get_of_lt {l : List α} (h : l.Sorted r) {a b : Fin l.length} r (l.get a) (l.get b) := List.pairwise_iff_get.1 h _ _ hab +set_option linter.deprecated false in +@[deprecated Sorted.rel_get_of_lt (since := "2024-05-08")] theorem Sorted.rel_nthLe_of_lt {l : List α} (h : l.Sorted r) {a b : ℕ} (ha : a < l.length) (hb : b < l.length) (hab : a < b) : r (l.nthLe a ha) (l.nthLe b hb) := List.pairwise_iff_get.1 h ⟨a, ha⟩ ⟨b, hb⟩ hab @@ -145,6 +147,8 @@ theorem Sorted.rel_get_of_le [IsRefl α r] {l : List α} (h : l.Sorted r) {a b : rcases hab.eq_or_lt with (rfl | hlt) exacts [refl _, h.rel_get_of_lt hlt] +set_option linter.deprecated false in +@[deprecated Sorted.rel_get_of_le (since := "2024-05-08")] theorem Sorted.rel_nthLe_of_le [IsRefl α r] {l : List α} (h : l.Sorted r) {a b : ℕ} (ha : a < l.length) (hb : b < l.length) (hab : a ≤ b) : r (l.nthLe a ha) (l.nthLe b hb) := h.rel_get_of_le hab @@ -156,7 +160,7 @@ theorem Sorted.rel_of_mem_take_of_mem_drop {l : List α} (h : List.Sorted r l) { obtain ⟨⟨ix, hix⟩, rfl⟩ := get_of_mem hx rw [get_take', get_drop'] rw [length_take] at hix - exact h.rel_nthLe_of_lt _ _ (Nat.lt_add_right _ (lt_min_iff.mp hix).left) + exact h.rel_get_of_lt (Nat.lt_add_right _ (lt_min_iff.mp hix).left) #align list.sorted.rel_of_mem_take_of_mem_drop List.Sorted.rel_of_mem_take_of_mem_drop end Sorted diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index 630085df1d021..f994967d1f71b 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -298,7 +298,8 @@ theorem get_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : Fin ⟨l.get ⟨i, lt_length_left_of_zipWith i.isLt⟩, l'.get ⟨i, lt_length_right_of_zipWith i.isLt⟩, by rw [get?_eq_get], by rw [get?_eq_get]; exact ⟨rfl, rfl⟩⟩ -@[simp] +set_option linter.deprecated false in +@[simp, deprecated get_zipWith (since := "2024-05-09")] theorem nthLe_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : ℕ} {h : i < (zipWith f l l').length} : (zipWith f l l').nthLe i h = @@ -312,7 +313,8 @@ theorem get_zip {l : List α} {l' : List β} {i : Fin (zip l l').length} : (l.get ⟨i, lt_length_left_of_zip i.isLt⟩, l'.get ⟨i, lt_length_right_of_zip i.isLt⟩) := get_zipWith -@[simp] +set_option linter.deprecated false in +@[simp, deprecated get_zip (since := "2024-05-09")] theorem nthLe_zip {l : List α} {l' : List β} {i : ℕ} {h : i < (zip l l').length} : (zip l l').nthLe i h = (l.nthLe i (lt_length_left_of_zip h), l'.nthLe i (lt_length_right_of_zip h)) := diff --git a/Mathlib/Data/MLList/BestFirst.lean b/Mathlib/Data/MLList/BestFirst.lean index 20030696241d3..8ec003793df80 100644 --- a/Mathlib/Data/MLList/BestFirst.lean +++ b/Mathlib/Data/MLList/BestFirst.lean @@ -128,7 +128,7 @@ def insertAndEject if q.size < max then q.insert n l else - match q.max with + match q.max? with | none => RBMap.empty | some m => q.insert n l |>.erase m.1 @@ -172,7 +172,7 @@ This may require improving estimates of priorities and shuffling the queue. partial def popWithBound (q : BestFirstQueue prio ε m β maxSize) : m (Option (((a : α) × (ε a) × β) × BestFirstQueue prio ε m β maxSize)) := do let q' ← ensureFirstIsBest q - match q'.min with + match q'.min? with | none => -- The queue is empty, nothing to return. return none diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index b12759c1b8da6..88020efde8bce 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -70,7 +70,7 @@ theorem nth_eq_getD_sort (h : (setOf p).Finite) (n : ℕ) : theorem nth_eq_orderEmbOfFin (hf : (setOf p).Finite) {n : ℕ} (hn : n < hf.toFinset.card) : nth p n = hf.toFinset.orderEmbOfFin rfl ⟨n, hn⟩ := by - rw [nth_eq_getD_sort hf, Finset.orderEmbOfFin_apply, List.getD_eq_get]; rfl + rw [nth_eq_getD_sort hf, Finset.orderEmbOfFin_apply, List.getD_eq_get] #align nat.nth_eq_order_emb_of_fin Nat.nth_eq_orderEmbOfFin theorem nth_strictMonoOn (hf : (setOf p).Finite) : diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 87ff6338a46af..b055590be4486 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -906,7 +906,7 @@ theorem forall_map_sub_of_Icc (P : ℝ → Prop) (h : ∀ x ∈ Icc (0 : ℝ) 1, theorem translationNumber_lt_of_forall_lt_add (hf : Continuous f) {z : ℝ} (hz : ∀ x, f x < x + z) : τ f < z := by obtain ⟨x, -, hx⟩ : ∃ x ∈ Icc (0 : ℝ) 1, ∀ y ∈ Icc (0 : ℝ) 1, f y - y ≤ f x - x := - isCompact_Icc.exists_forall_ge (nonempty_Icc.2 zero_le_one) + isCompact_Icc.exists_isMaxOn (nonempty_Icc.2 zero_le_one) (hf.sub continuous_id).continuousOn refine' lt_of_le_of_lt _ (sub_lt_iff_lt_add'.2 <| hz x) apply translationNumber_le_of_le_add @@ -917,7 +917,7 @@ theorem translationNumber_lt_of_forall_lt_add (hf : Continuous f) {z : ℝ} (hz theorem lt_translationNumber_of_forall_add_lt (hf : Continuous f) {z : ℝ} (hz : ∀ x, x + z < f x) : z < τ f := by obtain ⟨x, -, hx⟩ : ∃ x ∈ Icc (0 : ℝ) 1, ∀ y ∈ Icc (0 : ℝ) 1, f x - x ≤ f y - y - exact isCompact_Icc.exists_forall_le (nonempty_Icc.2 zero_le_one) + exact isCompact_Icc.exists_isMinOn (nonempty_Icc.2 zero_le_one) (hf.sub continuous_id).continuousOn refine' lt_of_lt_of_le (lt_sub_iff_add_lt'.2 <| hz x) _ apply le_translationNumber_of_add_le diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 78ef46bb81629..71239953bbf7e 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -242,10 +242,20 @@ theorem length_toList_pos_of_mem_support (h : x ∈ p.support) : 0 < length (toL zero_lt_two.trans_le (two_le_length_toList_iff_mem_support.mpr h) #align equiv.perm.length_to_list_pos_of_mem_support Equiv.Perm.length_toList_pos_of_mem_support +theorem get_toList (n : ℕ) (hn : n < length (toList p x)) : (toList p x).get ⟨n, hn⟩ = (p ^ n) x := + by simp [toList] + +theorem toList_get_zero (h : x ∈ p.support) : + (toList p x).get ⟨0, (length_toList_pos_of_mem_support _ _ h)⟩ = x := by simp [toList] + +set_option linter.deprecated false in +@[deprecated get_toList (since := "2024-05-08")] theorem nthLe_toList (n : ℕ) (hn : n < length (toList p x)) : (toList p x).nthLe n hn = (p ^ n) x := by simp [toList] #align equiv.perm.nth_le_to_list Equiv.Perm.nthLe_toList +set_option linter.deprecated false in +@[deprecated toList_get_zero (since := "2024-05-08")] theorem toList_nthLe_zero (h : x ∈ p.support) : (toList p x).nthLe 0 (length_toList_pos_of_mem_support _ _ h) = x := by simp [toList] #align equiv.perm.to_list_nth_le_zero Equiv.Perm.toList_nthLe_zero diff --git a/Mathlib/GroupTheory/Perm/List.lean b/Mathlib/GroupTheory/Perm/List.lean index 0c8aa822834c5..4559c0eb5442d 100644 --- a/Mathlib/GroupTheory/Perm/List.lean +++ b/Mathlib/GroupTheory/Perm/List.lean @@ -294,6 +294,7 @@ theorem formPerm_pow_apply_get (l : List α) (h : Nodup l) (n k : ℕ) (hk : k < · simp [Nat.mod_eq_of_lt hk] · simp [pow_succ', mul_apply, hn, formPerm_apply_get _ h, Nat.succ_eq_add_one, ← Nat.add_assoc] +set_option linter.deprecated false in @[deprecated formPerm_pow_apply_get] -- 2024-04-23 theorem formPerm_pow_apply_nthLe (l : List α) (h : Nodup l) (n k : ℕ) (hk : k < l.length) : (formPerm l ^ n) (l.nthLe k hk) = diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index b23e303e70054..c236aa8a6a895 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -667,7 +667,7 @@ theorem snorm_one_smul_measure {f : α → F} (c : ℝ≥0∞) : snorm f 1 (c theorem Memℒp.of_measure_le_smul {μ' : Measure α} (c : ℝ≥0∞) (hc : c ≠ ∞) (hμ'_le : μ' ≤ c • μ) {f : α → E} (hf : Memℒp f p μ) : Memℒp f p μ' := by - refine' ⟨hf.1.mono' (Measure.absolutelyContinuous_of_le_smul hμ'_le), _⟩ + refine' ⟨hf.1.mono_ac (Measure.absolutelyContinuous_of_le_smul hμ'_le), _⟩ refine' (snorm_mono_measure f hμ'_le).trans_lt _ by_cases hc0 : c = 0 · simp [hc0] diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 26b2dce43935f..305a88d8e5327 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -1617,7 +1617,7 @@ theorem comp_measurable {γ : Type*} {_ : MeasurableSpace γ} {_ : MeasurableSpa theorem comp_quasiMeasurePreserving {γ : Type*} {_ : MeasurableSpace γ} {_ : MeasurableSpace α} {f : γ → α} {μ : Measure γ} {ν : Measure α} (hg : AEStronglyMeasurable g ν) (hf : QuasiMeasurePreserving f μ ν) : AEStronglyMeasurable (g ∘ f) μ := - (hg.mono' hf.absolutelyContinuous).comp_measurable hf.measurable + (hg.mono_ac hf.absolutelyContinuous).comp_measurable hf.measurable #align measure_theory.ae_strongly_measurable.comp_quasi_measure_preserving MeasureTheory.AEStronglyMeasurable.comp_quasiMeasurePreserving theorem comp_measurePreserving {γ : Type*} {_ : MeasurableSpace γ} {_ : MeasurableSpace α} diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index e05f00c0f8532..cdf98ee54ced5 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -718,7 +718,7 @@ theorem div_mem_nhds_one_of_haar_pos (μ : Measure G) [IsHaarMeasure μ] [Locall rcases MeasurableSet.exists_lt_isCompact hE hEpos with ⟨K, KE, K_comp, K_meas⟩ refine ⟨closure K, ?_, K_comp.closure, isClosed_closure, ?_⟩ · exact K_comp.closure_subset_measurableSet hE KE - · rwa [K_comp.measure_closure_eq_of_group] + · rwa [K_comp.measure_closure] filter_upwards [eventually_nhds_one_measure_smul_diff_lt hK K_closed hKpos.ne' (μ := μ)] with g hg have : ¬Disjoint (g • K) K := fun hd ↦ by rw [hd.symm.sdiff_eq_right, measure_smul] at hg diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index a4d15512078e7..794b169139f81 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -218,7 +218,7 @@ lemma hasSum_nat_cosKernel₀ (a : ℝ) {t : ℝ} (ht : 0 < t) : HasSum (fun n : ℕ ↦ 2 * Real.cos (2 * π * a * (n + 1)) * rexp (-π * (n + 1) ^ 2 * t)) (cosKernel a t - 1) := by rw [← hasSum_ofReal, ofReal_sub, ofReal_one] - have := (hasSum_int_cosKernel a ht).sum_nat_of_sum_int + have := (hasSum_int_cosKernel a ht).nat_add_neg rw [← hasSum_nat_add_iff' 1] at this simp_rw [Finset.sum_range_one, Nat.cast_zero, neg_zero, Int.cast_zero, zero_pow two_ne_zero, mul_zero, zero_mul, Complex.exp_zero, Real.exp_zero, ofReal_one, mul_one, Int.cast_neg, @@ -531,7 +531,7 @@ lemma hasSum_nat_completedCosZeta (a : ℝ) {s : ℂ} (hs : 1 < re s) : (completedCosZeta a s) := by have aux : ((|0| : ℤ) : ℂ) ^ s = 0 := by rw [abs_zero, Int.cast_zero, zero_cpow (ne_zero_of_one_lt_re hs)] - have hint := (hasSum_int_completedCosZeta a hs).sum_nat_of_sum_int + have hint := (hasSum_int_completedCosZeta a hs).nat_add_neg rw [aux, div_zero, zero_div, add_zero] at hint refine hint.congr_fun fun n ↦ ?_ split_ifs with h @@ -757,7 +757,7 @@ lemma hasSum_int_cosZeta (a : ℝ) {s : ℂ} (hs : 1 < re s) : /-- Formula for `cosZeta` as a Dirichlet series in the convergence range, with sum over `ℕ`. -/ lemma hasSum_nat_cosZeta (a : ℝ) {s : ℂ} (hs : 1 < re s) : HasSum (fun n : ℕ ↦ Real.cos (2 * π * a * n) / (n : ℂ) ^ s) (cosZeta a s) := by - have := (hasSum_int_cosZeta a hs).sum_nat_of_sum_int + have := (hasSum_int_cosZeta a hs).nat_add_neg simp_rw [abs_neg, Int.cast_neg, Nat.abs_cast, Int.cast_natCast, mul_neg, abs_zero, Int.cast_zero, zero_cpow (ne_zero_of_one_lt_re hs), div_zero, zero_div, add_zero, ← add_div, div_right_comm _ _ (2 : ℂ)] at this diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/OneVariable.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/OneVariable.lean index 9d2d5aa9cab3b..bbcba8b9c8509 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/OneVariable.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/OneVariable.lean @@ -76,7 +76,7 @@ theorem hasSum_nat_jacobiTheta {τ : ℂ} (hτ : 0 < im τ) : HasSum (fun n : ℕ => cexp (π * I * ((n : ℂ) + 1) ^ 2 * τ)) ((jacobiTheta τ - 1) / 2) := by have := hasSum_jacobiTheta₂_term 0 hτ simp_rw [jacobiTheta₂_term, mul_zero, zero_add, ← jacobiTheta_eq_jacobiTheta₂] at this - have := this.sum_nat_of_sum_int + have := this.nat_add_neg rw [← hasSum_nat_add_iff' 1] at this simp_rw [Finset.sum_range_one, Int.cast_neg, Int.cast_natCast, Nat.cast_zero, neg_zero, Int.cast_zero, sq (0 : ℂ), mul_zero, zero_mul, neg_sq, ← mul_two, diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index 47cc68d615b02..550d74faf9e78 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -329,7 +329,7 @@ attribute [-instance] algebraRat is a number field. -/ instance {f : Polynomial ℚ} [hf : Fact (Irreducible f)] : NumberField (AdjoinRoot f) where to_charZero := charZero_of_injective_algebraMap (algebraMap ℚ _).injective - to_finiteDimensional := by convert (AdjoinRoot.powerBasis hf.out.ne_zero).finiteDimensional + to_finiteDimensional := by convert (AdjoinRoot.powerBasis hf.out.ne_zero).finite end diff --git a/Mathlib/NumberTheory/ZetaValues.lean b/Mathlib/NumberTheory/ZetaValues.lean index 677e7d3ab4029..a2bd05342743a 100644 --- a/Mathlib/NumberTheory/ZetaValues.lean +++ b/Mathlib/NumberTheory/ZetaValues.lean @@ -236,7 +236,7 @@ theorem hasSum_one_div_nat_pow_mul_fourier {k : ℕ} (hk : 2 ≤ k) {x : ℝ} (h (fun n : ℕ => (1 : ℂ) / (n : ℂ) ^ k * (fourier n (x : 𝕌) + (-1 : ℂ) ^ k * fourier (-n) (x : 𝕌))) (-(2 * π * I) ^ k / k ! * bernoulliFun k x) := by - convert (hasSum_one_div_pow_mul_fourier_mul_bernoulliFun hk hx).sum_nat_of_sum_int using 1 + convert (hasSum_one_div_pow_mul_fourier_mul_bernoulliFun hk hx).nat_add_neg using 1 · ext1 n rw [Int.cast_neg, mul_add, ← mul_assoc] conv_rhs => rw [neg_eq_neg_one_mul, mul_pow, ← div_div] diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index d7bbc3a0eec1b..83854775aec31 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -243,7 +243,8 @@ theorem integrable_pdf_smul_iff [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X Integrable (fun x => (pdf X ℙ μ x).toReal • f x) μ ↔ Integrable (fun x => f (X x)) ℙ := by -- Porting note: using `erw` because `rw` doesn't recognize `(f <| X ·)` as `f ∘ X` -- https://github.com/leanprover-community/mathlib4/issues/5164 - erw [← integrable_map_measure (hf.mono' HasPDF.absolutelyContinuous) (HasPDF.aemeasurable X ℙ μ), + erw [← integrable_map_measure (hf.mono_ac HasPDF.absolutelyContinuous) + (HasPDF.aemeasurable X ℙ μ), map_eq_withDensity_pdf X ℙ μ, pdf_def, integrable_rnDeriv_smul_iff HasPDF.absolutelyContinuous] eta_reduce rw [withDensity_rnDeriv_eq _ _ HasPDF.absolutelyContinuous] @@ -254,7 +255,7 @@ function `f`, `f ∘ X` is a random variable with expectation `∫ x, pdf X x where `μ` is a measure on the codomain of `X`. -/ theorem integral_pdf_smul [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X ℙ μ] {f : E → F} (hf : AEStronglyMeasurable f μ) : ∫ x, (pdf X ℙ μ x).toReal • f x ∂μ = ∫ x, f (X x) ∂ℙ := by - rw [← integral_map (HasPDF.aemeasurable X ℙ μ) (hf.mono' HasPDF.absolutelyContinuous), + rw [← integral_map (HasPDF.aemeasurable X ℙ μ) (hf.mono_ac HasPDF.absolutelyContinuous), map_eq_withDensity_pdf X ℙ μ, pdf_def, integral_rnDeriv_smul HasPDF.absolutelyContinuous, withDensity_rnDeriv_eq _ _ HasPDF.absolutelyContinuous] #align measure_theory.pdf.integral_fun_mul_eq_integral MeasureTheory.pdf.integral_pdf_smul diff --git a/Mathlib/RingTheory/FreeCommRing.lean b/Mathlib/RingTheory/FreeCommRing.lean index 9763da662034e..86dad072d2c13 100644 --- a/Mathlib/RingTheory/FreeCommRing.lean +++ b/Mathlib/RingTheory/FreeCommRing.lean @@ -131,7 +131,7 @@ private def liftToMultiset : (α → R) ≃ (Multiplicative (Multiset α) →* R let x' := Multiplicative.toAdd x show (Multiset.map (fun a => F' {a}) x').sum = F' x' by erw [← Multiset.map_map (fun x => F' x) (fun x => {x}), ← AddMonoidHom.map_multiset_sum] - exact F.congr_arg (Multiset.sum_map_singleton x') + exact DFunLike.congr_arg F (Multiset.sum_map_singleton x') /-- Lift a map `α → R` to an additive group homomorphism `FreeCommRing α → R`. -/ def lift : (α → R) ≃ (FreeCommRing α →+* R) := diff --git a/Mathlib/RingTheory/Trace.lean b/Mathlib/RingTheory/Trace.lean index 6152d9aa25897..259a62dedfc8b 100644 --- a/Mathlib/RingTheory/Trace.lean +++ b/Mathlib/RingTheory/Trace.lean @@ -589,7 +589,7 @@ theorem det_traceMatrix_ne_zero' [IsSeparable K L] : det (traceMatrix K pb.basis suffices algebraMap K (AlgebraicClosure L) (det (traceMatrix K pb.basis)) ≠ 0 by refine' mt (fun ht => _) this rw [ht, RingHom.map_zero] - haveI : FiniteDimensional K L := pb.finiteDimensional + haveI : FiniteDimensional K L := pb.finite let e : Fin pb.dim ≃ (L →ₐ[K] AlgebraicClosure L) := (Fintype.equivFinOfCardEq ?_).symm · rw [RingHom.map_det, RingHom.mapMatrix_apply, traceMatrix_eq_embeddingsMatrixReindex_mul_trans K _ _ e, diff --git a/Mathlib/Topology/Algebra/Order/Compact.lean b/Mathlib/Topology/Algebra/Order/Compact.lean index 6420b2aff8c57..d0e57479b2550 100644 --- a/Mathlib/Topology/Algebra/Order/Compact.lean +++ b/Mathlib/Topology/Algebra/Order/Compact.lean @@ -306,7 +306,7 @@ theorem ContinuousOn.exists_isMinOn' [ClosedIicTopology α] {s : Set β} {f : β rcases (hasBasis_cocompact.inf_principal _).eventually_iff.1 hc with ⟨K, hK, hKf⟩ have hsub : insert x₀ (K ∩ s) ⊆ s := insert_subset_iff.2 ⟨h₀, inter_subset_right _ _⟩ obtain ⟨x, hx, hxf⟩ : ∃ x ∈ insert x₀ (K ∩ s), ∀ y ∈ insert x₀ (K ∩ s), f x ≤ f y := - ((hK.inter_right hsc).insert x₀).exists_forall_le (insert_nonempty _ _) (hf.mono hsub) + ((hK.inter_right hsc).insert x₀).exists_isMinOn (insert_nonempty _ _) (hf.mono hsub) refine' ⟨x, hsub hx, fun y hy => _⟩ by_cases hyK : y ∈ K exacts [hxf _ (Or.inr ⟨hyK, hy⟩), (hxf _ (Or.inl rfl)).trans (hKf ⟨hyK, hy⟩)] @@ -340,9 +340,9 @@ theorem ContinuousOn.exists_forall_ge' [ClosedIciTopology α] {s : Set β} {f : away from compact sets, then it has a global minimum. -/ theorem Continuous.exists_forall_le' [ClosedIicTopology α] {f : β → α} (hf : Continuous f) (x₀ : β) (h : ∀ᶠ x in cocompact β, f x₀ ≤ f x) : ∃ x : β, ∀ y : β, f x ≤ f y := - let ⟨x, _, hx⟩ := hf.continuousOn.exists_forall_le' isClosed_univ (mem_univ x₀) + let ⟨x, _, hx⟩ := hf.continuousOn.exists_isMinOn' isClosed_univ (mem_univ x₀) (by rwa [principal_univ, inf_top_eq]) - ⟨x, fun y => hx y (mem_univ y)⟩ + ⟨x, fun y => hx (mem_univ y)⟩ #align continuous.exists_forall_le' Continuous.exists_forall_le' /-- The **extreme value theorem**: if a continuous function `f` is smaller than a value in its range @@ -442,9 +442,9 @@ theorem IsCompact.sSup_lt_iff_of_continuous [ClosedIciTopology α] {f : β → sSup (f '' K) < y ↔ ∀ x ∈ K, f x < y := by refine' ⟨fun h x hx => (le_csSup (hK.bddAbove_image hf) <| mem_image_of_mem f hx).trans_lt h, fun h => _⟩ - obtain ⟨x, hx, h2x⟩ := hK.exists_forall_ge h0K hf + obtain ⟨x, hx, h2x⟩ := hK.exists_isMaxOn h0K hf refine' (csSup_le (h0K.image f) _).trans_lt (h x hx) - rintro _ ⟨x', hx', rfl⟩; exact h2x x' hx' + rintro _ ⟨x', hx', rfl⟩; exact h2x hx' #align is_compact.Sup_lt_iff_of_continuous IsCompact.sSup_lt_iff_of_continuous theorem IsCompact.lt_sInf_iff_of_continuous [ClosedIicTopology α] {f : β → α} {K : Set β} diff --git a/Mathlib/Topology/Algebra/Order/Rolle.lean b/Mathlib/Topology/Algebra/Order/Rolle.lean index 4e304d734ef42..5223e414bf609 100644 --- a/Mathlib/Topology/Algebra/Order/Rolle.lean +++ b/Mathlib/Topology/Algebra/Order/Rolle.lean @@ -39,9 +39,9 @@ theorem exists_Ioo_extr_on_Icc (hab : a < b) (hfc : ContinuousOn f (Icc a b)) (h have ne : (Icc a b).Nonempty := nonempty_Icc.2 (le_of_lt hab) -- Consider absolute min and max points obtain ⟨c, cmem, cle⟩ : ∃ c ∈ Icc a b, ∀ x ∈ Icc a b, f c ≤ f x := - isCompact_Icc.exists_forall_le ne hfc + isCompact_Icc.exists_isMinOn ne hfc obtain ⟨C, Cmem, Cge⟩ : ∃ C ∈ Icc a b, ∀ x ∈ Icc a b, f x ≤ f C := - isCompact_Icc.exists_forall_ge ne hfc + isCompact_Icc.exists_isMaxOn ne hfc by_cases hc : f c = f a · by_cases hC : f C = f a · have : ∀ x ∈ Icc a b, f x = f a := fun x hx => le_antisymm (hC ▸ Cge x hx) (hc ▸ cle x hx) diff --git a/Mathlib/Topology/ContinuousFunction/Ideals.lean b/Mathlib/Topology/ContinuousFunction/Ideals.lean index 9998ad746b4b0..61c6589779785 100644 --- a/Mathlib/Topology/ContinuousFunction/Ideals.lean +++ b/Mathlib/Topology/ContinuousFunction/Ideals.lean @@ -292,7 +292,7 @@ theorem idealOfSet_ofIdeal_eq_closure (I : Ideal C(X, 𝕜)) : obtain ⟨c, hc, hgc'⟩ : ∃ c > 0, ∀ y : X, y ∈ t → c ≤ g' y := t.eq_empty_or_nonempty.elim (fun ht' => ⟨1, zero_lt_one, fun y hy => False.elim (by rwa [ht'] at hy)⟩) fun ht' => - let ⟨x, hx, hx'⟩ := ht.isCompact.exists_forall_le ht' (map_continuous g').continuousOn + let ⟨x, hx, hx'⟩ := ht.isCompact.exists_isMinOn ht' (map_continuous g').continuousOn ⟨g' x, hgt' x hx, hx'⟩ obtain ⟨g, hg, hgc⟩ := exists_mul_le_one_eqOn_ge g' hc refine' ⟨g * g', _, hg, hgc.mono hgc'⟩ diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index ba5a30f1c4272..fc68f42a36698 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -420,7 +420,7 @@ variable (X : Type u) (Y : Type v) [MetricSpace X] [CompactSpace X] [Nonempty X] we can finally select a candidate minimizing `HD`. This will be the candidate realizing the optimal coupling. -/ private theorem exists_minimizer : ∃ f ∈ candidatesB X Y, ∀ g ∈ candidatesB X Y, HD f ≤ HD g := - isCompact_candidatesB.exists_forall_le candidatesB_nonempty HD_continuous.continuousOn + isCompact_candidatesB.exists_isMinOn candidatesB_nonempty HD_continuous.continuousOn private def optimalGHDist : Cb X Y := Classical.choose (exists_minimizer X Y) diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index 3158193259e03..9051564b58808 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -248,8 +248,7 @@ theorem _root_.IsOpen.exists_iUnion_isClosed {U : Set α} (hU : IsOpen U) : theorem _root_.IsCompact.exists_infEdist_eq_edist (hs : IsCompact s) (hne : s.Nonempty) (x : α) : ∃ y ∈ s, infEdist x s = edist x y := by have A : Continuous fun y => edist x y := continuous_const.edist continuous_id - obtain ⟨y, ys, hy⟩ : ∃ y ∈ s, ∀ z, z ∈ s → edist x y ≤ edist x z := - hs.exists_forall_le hne A.continuousOn + obtain ⟨y, ys, hy⟩ := hs.exists_isMinOn hne A.continuousOn exact ⟨y, ys, le_antisymm (infEdist_le_edist_of_mem ys) (by rwa [le_infEdist])⟩ #align is_compact.exists_inf_edist_eq_edist IsCompact.exists_infEdist_eq_edist @@ -258,12 +257,11 @@ theorem exists_pos_forall_lt_edist (hs : IsCompact s) (ht : IsClosed t) (hst : D rcases s.eq_empty_or_nonempty with (rfl | hne) · use 1 simp - obtain ⟨x, hx, h⟩ : ∃ x ∈ s, ∀ y ∈ s, infEdist x t ≤ infEdist y t := - hs.exists_forall_le hne continuous_infEdist.continuousOn + obtain ⟨x, hx, h⟩ := hs.exists_isMinOn hne continuous_infEdist.continuousOn have : 0 < infEdist x t := pos_iff_ne_zero.2 fun H => hst.le_bot ⟨hx, (mem_iff_infEdist_zero_of_closed ht).mpr H⟩ rcases ENNReal.lt_iff_exists_nnreal_btwn.1 this with ⟨r, h₀, hr⟩ - exact ⟨r, ENNReal.coe_pos.mp h₀, fun y hy z hz => hr.trans_le <| le_infEdist.1 (h y hy) z hz⟩ + exact ⟨r, ENNReal.coe_pos.mp h₀, fun y hy z hz => hr.trans_le <| le_infEdist.1 (h hy) z hz⟩ #align emetric.exists_pos_forall_lt_edist EMetric.exists_pos_forall_lt_edist end InfEdist diff --git a/Mathlib/Topology/MetricSpace/ProperSpace.lean b/Mathlib/Topology/MetricSpace/ProperSpace.lean index 70d387f4ba3d8..6b0a54f53afaa 100644 --- a/Mathlib/Topology/MetricSpace/ProperSpace.lean +++ b/Mathlib/Topology/MetricSpace/ProperSpace.lean @@ -137,7 +137,7 @@ theorem exists_pos_lt_subset_ball (hr : 0 < r) (hs : IsClosed s) (h : s ⊆ ball have : IsCompact s := (isCompact_closedBall x r).of_isClosed_subset hs (h.trans ball_subset_closedBall) obtain ⟨y, hys, hy⟩ : ∃ y ∈ s, s ⊆ closedBall x (dist y x) := - this.exists_forall_ge hne (continuous_id.dist continuous_const).continuousOn + this.exists_isMaxOn hne (continuous_id.dist continuous_const).continuousOn have hyr : dist y x < r := h hys rcases exists_between hyr with ⟨r', hyr', hrr'⟩ exact ⟨r', ⟨dist_nonneg.trans_lt hyr', hrr'⟩, hy.trans <| closedBall_subset_ball hyr'⟩ diff --git a/Mathlib/Topology/ProperMap.lean b/Mathlib/Topology/ProperMap.lean index 6b56d7288c21e..0d0a9397445f2 100644 --- a/Mathlib/Topology/ProperMap.lean +++ b/Mathlib/Topology/ProperMap.lean @@ -287,8 +287,10 @@ lemma isProperMap_restr_of_proper_of_closed {U : Set X} (hf : IsProperMap f) (hU IsProperMap.comp (isProperMap_subtype_val_of_closed hU) hf /-- The range of a proper map is closed. -/ -lemma IsProperMap.closed_range (hf : IsProperMap f) : IsClosed (range f) := - hf.isClosedMap.closed_range +lemma IsProperMap.isClosed_range (hf : IsProperMap f) : IsClosed (range f) := + hf.isClosedMap.isClosed_range + +@[deprecated (since := "2024-05-08")] alias IsProperMap.closed_range := IsProperMap.isClosed_range /-- Version of `isProperMap_iff_isClosedMap_and_compact_fibers` in terms of `cofinite` and `cocompact`. Only works when the codomain is `T1`. -/ diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 6185fbc96c308..40356cdf45df6 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -1157,6 +1157,8 @@ protected theorem UniformSpace.le_sInf {tt : Set (UniformSpace α)} {t : Uniform (h : ∀ t' ∈ tt, t ≤ t') : t ≤ sInf tt := show 𝓤[t] ≤ ⨅ u ∈ tt, 𝓤[u] from le_iInf₂ h +set_option linter.deprecated false in +-- TODO update this code to avoid the deprecation instance : Top (UniformSpace α) := ⟨.ofNhdsEqComap ⟨⊤, le_top, le_top, le_top⟩ ⊤ fun x ↦ by simp only [nhds_top, comap_top]⟩ From e50d4102559e198e35fb447e995021e680656385 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 9 May 2024 09:01:36 +0000 Subject: [PATCH 48/52] feat(Rat, NNRat): `q * q.den = q.num` (#12739) and a few other lemmas. From LeanAPAP --- Archive/Imo/Imo2013Q5.lean | 2 +- Mathlib/Data/NNRat/Defs.lean | 22 +++++++++++++++++++ Mathlib/Data/Rat/Defs.lean | 10 +++++++++ Mathlib/Data/Rat/Lemmas.lean | 9 -------- .../RingTheory/Localization/FractionRing.lean | 2 +- 5 files changed, 34 insertions(+), 11 deletions(-) diff --git a/Archive/Imo/Imo2013Q5.lean b/Archive/Imo/Imo2013Q5.lean index 4b439868c75db..9d549f5d85d2e 100644 --- a/Archive/Imo/Imo2013Q5.lean +++ b/Archive/Imo/Imo2013Q5.lean @@ -241,7 +241,7 @@ theorem imo2013_q5 (f : ℚ → ℝ) (H1 : ∀ x y, 0 < x → 0 < y → f (x * y -- we need the top of the fraction to be strictly greater than 1 in order -- to apply `fixed_point_of_gt_1`. intro x hx - have H₀ : x * x.den = x.num := Rat.mul_den_eq_num + have H₀ : x * x.den = x.num := x.mul_den_eq_num have H : x * (↑(2 * x.den) : ℚ) = (↑(2 * x.num) : ℚ) := by push_cast; linear_combination 2 * H₀ set x2denom := 2 * x.den set x2num := 2 * x.num diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 7a4718a258894..79eb0c1ba042b 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -403,9 +403,31 @@ lemma mk_divInt (n d : ℕ) : lemma divNat_inj (h₁ : d₁ ≠ 0) (h₂ : d₂ ≠ 0) : divNat n₁ d₁ = divNat n₂ d₂ ↔ n₁ * d₂ = n₂ * d₁ := by rw [← coe_inj]; simp [Rat.mkRat_eq_iff, h₁, h₂]; norm_cast +@[simp] lemma divNat_zero (n : ℕ) : divNat n 0 = 0 := by simp [divNat]; rfl + @[simp] lemma num_divNat_den (q : ℚ≥0) : divNat q.num q.den = q := ext $ by rw [← (q : ℚ).mkRat_num_den']; simp [num_coe, den_coe] +lemma natCast_eq_divNat (n : ℕ) : (n : ℚ≥0) = divNat n 1 := (num_divNat_den _).symm + +lemma divNat_mul_divNat (n₁ n₂ : ℕ) {d₁ d₂} (hd₁ : d₁ ≠ 0) (hd₂ : d₂ ≠ 0) : + divNat n₁ d₁ * divNat n₂ d₂ = divNat (n₁ * n₂) (d₁ * d₂) := by + ext; push_cast; exact Rat.divInt_mul_divInt _ _ (mod_cast hd₁) (mod_cast hd₂) + +lemma divNat_mul_left {a : ℕ} (ha : a ≠ 0) (n d : ℕ) : divNat (a * n) (a * d) = divNat n d := by + ext; push_cast; exact Rat.divInt_mul_left (mod_cast ha) + +lemma divNat_mul_right {a : ℕ} (ha : a ≠ 0) (n d : ℕ) : divNat (n * a) (d * a) = divNat n d := by + ext; push_cast; exact Rat.divInt_mul_right (mod_cast ha) + +@[simp] lemma mul_den_eq_num (q : ℚ≥0) : q * q.den = q.num := by + ext + push_cast + rw [← Int.cast_natCast, ← den_coe, ← Int.cast_natCast q.num, ← num_coe] + exact Rat.mul_den_eq_num _ + +@[simp] lemma den_mul_eq_num (q : ℚ≥0) : q.den * q = q.num := by rw [mul_comm, mul_den_eq_num] + /-- Define a (dependent) function or prove `∀ r : ℚ, p r` by dealing with nonnegative rational numbers of the form `n / d` with `d ≠ 0` and `n`, `d` coprime. -/ @[elab_as_elim] diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 4327a6a199946..d9c55a22faa14 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -574,6 +574,16 @@ theorem natCast_eq_divInt (n : ℕ) : ↑n = n /. 1 := by rw [← Int.cast_natCast, intCast_eq_divInt] #align rat.coe_nat_eq_mk Rat.natCast_eq_divInt +@[simp] lemma mul_den_eq_num (q : ℚ) : q * q.den = q.num := by + suffices (q.num /. ↑q.den) * (↑q.den /. 1) = q.num /. 1 by + conv => pattern (occs := 1) q; (rw [← num_divInt_den q]) + simp only [intCast_eq_divInt, natCast_eq_divInt, num_divInt_den] at this ⊢; assumption + have : (q.den : ℤ) ≠ 0 := mod_cast q.den_ne_zero + rw [divInt_mul_divInt _ _ this one_ne_zero, mul_comm (q.den : ℤ) 1, divInt_mul_right this] +#align rat.mul_denom_eq_num Rat.mul_den_eq_num + +@[simp] lemma den_mul_eq_num (q : ℚ) : q.den * q = q.num := by rw [mul_comm, mul_den_eq_num] + -- 2024-04-05 @[deprecated] alias coe_int_eq_divInt := intCast_eq_divInt @[deprecated] alias coe_int_div_eq_divInt := intCast_div_eq_divInt diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index b9472c869999c..ab92ae5f0d070 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -177,15 +177,6 @@ protected theorem inv_neg (q : ℚ) : (-q)⁻¹ = -q⁻¹ := by simp only [Rat.neg_divInt, Rat.inv_divInt', eq_self_iff_true, Rat.divInt_neg] #align rat.inv_neg Rat.inv_neg -@[simp] -theorem mul_den_eq_num {q : ℚ} : q * q.den = q.num := by - suffices (q.num /. ↑q.den) * (↑q.den /. 1) = q.num /. 1 by - conv => pattern (occs := 1) q; (rw [← num_divInt_den q]) - simp only [intCast_eq_divInt, natCast_eq_divInt, num_divInt_den] at this ⊢; assumption - have : (q.den : ℤ) ≠ 0 := ne_of_gt (mod_cast q.pos) - rw [Rat.divInt_mul_divInt _ _ this one_ne_zero, mul_comm (q.den : ℤ) 1, divInt_mul_right this] -#align rat.mul_denom_eq_num Rat.mul_den_eq_num - theorem den_div_cast_eq_one_iff (m n : ℤ) (hn : n ≠ 0) : ((m : ℚ) / n).den = 1 ↔ n ∣ m := by replace hn : (n : ℚ) ≠ 0 := by rwa [Ne, ← Int.cast_zero, coe_int_inj] constructor diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index b2ccdc82dccca..83b8e46fa66b3 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -54,7 +54,7 @@ instance Rat.isFractionRing : IsFractionRing ℤ ℚ where simpa only [eq_intCast, isUnit_iff_ne_zero, Int.cast_eq_zero, Ne, Subtype.coe_mk] using hx surj' := by rintro ⟨n, d, hd, h⟩ - refine' ⟨⟨n, ⟨d, _⟩⟩, Rat.mul_den_eq_num⟩ + refine' ⟨⟨n, ⟨d, _⟩⟩, Rat.mul_den_eq_num _⟩ rw [mem_nonZeroDivisors_iff_ne_zero, Int.natCast_ne_zero_iff_pos] exact Nat.zero_lt_of_ne_zero hd exists_of_eq {x y} := by From 51625e908ad8a7478f13ad7a109d3bb32c2d610d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 9 May 2024 10:48:39 +0000 Subject: [PATCH 49/52] refactor(Measure): use FunLike (#12684) Co-authored-by: Eric Wieser --- Mathlib/MeasureTheory/Constructions/Pi.lean | 16 ++++--- .../MeasureTheory/Decomposition/Jordan.lean | 4 +- .../MeasureTheory/Decomposition/Lebesgue.lean | 2 +- .../Group/FundamentalDomain.lean | 2 +- Mathlib/MeasureTheory/Measure/Dirac.lean | 4 +- .../MeasureTheory/Measure/EverywherePos.lean | 4 +- .../MeasureTheory/Measure/FiniteMeasure.lean | 8 ++-- .../Measure/FiniteMeasureProd.lean | 5 +-- Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 3 +- Mathlib/MeasureTheory/Measure/Hausdorff.lean | 10 ++--- .../Measure/Lebesgue/EqHaar.lean | 3 +- .../MeasureTheory/Measure/MeasureSpace.lean | 35 +++++---------- .../Measure/MeasureSpaceDef.lean | 45 +++++++++---------- .../Measure/ProbabilityMeasure.lean | 13 ++---- Mathlib/MeasureTheory/Measure/Restrict.lean | 14 +++--- Mathlib/MeasureTheory/Measure/Tilted.lean | 9 ++-- Mathlib/MeasureTheory/Measure/Trim.lean | 2 +- .../MeasureTheory/Measure/Typeclasses.lean | 7 ++- .../Probability/ConditionalProbability.lean | 13 +++--- Mathlib/Probability/Kernel/Composition.lean | 3 +- .../Kernel/Disintegration/Density.lean | 5 +-- Mathlib/Probability/Kernel/RadonNikodym.lean | 10 ++--- Mathlib/Probability/Variance.lean | 2 +- 23 files changed, 91 insertions(+), 128 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 5668bf5e83b4b..6e9521f349db8 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -249,10 +249,12 @@ instance sigmaFinite_tprod (l : List δ) (μ : ∀ i, Measure (π i)) [∀ i, Si theorem tprod_tprod (l : List δ) (μ : ∀ i, Measure (π i)) [∀ i, SigmaFinite (μ i)] (s : ∀ i, Set (π i)) : Measure.tprod l μ (Set.tprod l s) = (l.map fun i => (μ i) (s i)).prod := by - induction' l with i l ih; · simp - rw [tprod_cons, Set.tprod] - erw [prod_prod] -- TODO: why `rw` fails? - rw [map_cons, prod_cons, ih] + induction l with + | nil => simp + | cons a l ih => + rw [tprod_cons, Set.tprod] + erw [prod_prod] -- TODO: why `rw` fails? + rw [map_cons, prod_cons, ih] #align measure_theory.measure.tprod_tprod MeasureTheory.Measure.tprod_tprod end Tprod @@ -843,10 +845,10 @@ theorem measurePreserving_piUnique {π : ι → Type*} [Unique ι] {m : ∀ i, M set e := MeasurableEquiv.piUnique π have : (piPremeasure fun i => (μ i).toOuterMeasure) = Measure.map e.symm (μ default) := by ext1 s - rw [piPremeasure, Fintype.prod_unique, e.symm.map_apply] + rw [piPremeasure, Fintype.prod_unique, e.symm.map_apply, coe_toOuterMeasure] congr 1; exact e.toEquiv.image_eq_preimage s - simp_rw [Measure.pi, OuterMeasure.pi, this, boundedBy_eq_self, toOuterMeasure_toMeasure, - MeasurableEquiv.map_map_symm] + simp_rw [Measure.pi, OuterMeasure.pi, this, ← coe_toOuterMeasure, boundedBy_eq_self, + toOuterMeasure_toMeasure, MeasurableEquiv.map_map_symm] theorem volume_preserving_piUnique (π : ι → Type*) [Unique ι] [∀ i, MeasureSpace (π i)] : MeasurePreserving (MeasurableEquiv.piUnique π) volume volume := diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index 6ecaf800b86fe..d7582de3f324b 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -189,8 +189,8 @@ theorem toSignedMeasure_smul (r : ℝ≥0) : (r • j).toSignedMeasure = r • j -- Porting note: removed `rfl` after the `rw` by adding further steps. rw [VectorMeasure.smul_apply, toSignedMeasure, toSignedMeasure, toSignedMeasure_sub_apply hi, toSignedMeasure_sub_apply hi, smul_sub, smul_posPart, - smul_negPart, ← ENNReal.toReal_smul, ← ENNReal.toReal_smul, smul_toOuterMeasure, - OuterMeasure.coe_smul, Pi.smul_apply, smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply] + smul_negPart, ← ENNReal.toReal_smul, ← ENNReal.toReal_smul, Measure.smul_apply, + Measure.smul_apply] #align measure_theory.jordan_decomposition.to_signed_measure_smul MeasureTheory.JordanDecomposition.toSignedMeasure_smul /-- A Jordan decomposition provides a Hahn decomposition. -/ diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index 10f4536cb90d5..873bdf8e97303 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -221,7 +221,7 @@ lemma absolutelyContinuous_withDensity_rnDeriv [HaveLebesgueDecomposition ν μ] simp only [nonpos_iff_eq_zero, add_eq_zero] constructor · refine hμν ?_ - simp only [add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply, add_eq_zero] + simp only [coe_add, Pi.add_apply, add_eq_zero] constructor · exact measure_mono_null (Set.inter_subset_right _ _) ht1 · exact measure_mono_null (Set.inter_subset_left _ _) hνs diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index 8c1567590ea69..8870735d06e9a 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -942,7 +942,7 @@ theorem IsFundamentalDomain.quotientMeasureEqMeasurePreimage_of_zero QuotientMeasureEqMeasurePreimage ν (0 : Measure (Quotient α_mod_G)) := by apply fund_dom_s.quotientMeasureEqMeasurePreimage ext U meas_U - simp only [zero_toOuterMeasure, OuterMeasure.coe_zero, Pi.zero_apply] + simp only [Measure.coe_zero, Pi.zero_apply] convert (measure_inter_null_of_null_right (h := vol_s) (Quotient.mk α_mod_G ⁻¹' U)).symm rw [measure_map_restrict_apply (meas_U := meas_U)] diff --git a/Mathlib/MeasureTheory/Measure/Dirac.lean b/Mathlib/MeasureTheory/Measure/Dirac.lean index ffead1aa256a9..9f8549aacfff7 100644 --- a/Mathlib/MeasureTheory/Measure/Dirac.lean +++ b/Mathlib/MeasureTheory/Measure/Dirac.lean @@ -65,8 +65,8 @@ theorem map_dirac {f : α → β} (hf : Measurable f) (a : α) : (dirac a).map f lemma map_const (μ : Measure α) (c : β) : μ.map (fun _ ↦ c) = (μ Set.univ) • dirac c := by ext s hs - simp only [aemeasurable_const, measurable_const, smul_toOuterMeasure, OuterMeasure.coe_smul, - Pi.smul_apply, dirac_apply' _ hs, smul_eq_mul] + simp only [aemeasurable_const, measurable_const, Measure.coe_smul, Pi.smul_apply, + dirac_apply' _ hs, smul_eq_mul] classical rw [Measure.map_apply measurable_const hs, Set.preimage_const] by_cases hsc : c ∈ s diff --git a/Mathlib/MeasureTheory/Measure/EverywherePos.lean b/Mathlib/MeasureTheory/Measure/EverywherePos.lean index a13922b5c1831..ce97a307136dd 100644 --- a/Mathlib/MeasureTheory/Measure/EverywherePos.lean +++ b/Mathlib/MeasureTheory/Measure/EverywherePos.lean @@ -109,7 +109,7 @@ lemma measure_eq_zero_of_subset_diff_everywherePosSubset its everywhere positive subset. -/ lemma everywherePosSubset_ae_eq [OpensMeasurableSpace α] [InnerRegular μ] (hs : MeasurableSet s) : μ.everywherePosSubset s =ᵐ[μ] s := by - simp only [ae_eq_set, diff_eq_empty.mpr (everywherePosSubset_subset μ s), OuterMeasure.empty', + simp only [ae_eq_set, diff_eq_empty.mpr (everywherePosSubset_subset μ s), measure_empty, true_and, (hs.diff hs.everywherePosSubset).measure_eq_iSup_isCompact, ENNReal.iSup_eq_zero] intro k hk h'k exact measure_eq_zero_of_subset_diff_everywherePosSubset h'k hk @@ -121,7 +121,7 @@ lemma everywherePosSubset_ae_eq_of_measure_ne_top μ.everywherePosSubset s =ᵐ[μ] s := by have A : μ (s \ μ.everywherePosSubset s) ≠ ∞ := ((measure_mono (diff_subset _ _ )).trans_lt h's.lt_top).ne - simp only [ae_eq_set, diff_eq_empty.mpr (everywherePosSubset_subset μ s), OuterMeasure.empty', + simp only [ae_eq_set, diff_eq_empty.mpr (everywherePosSubset_subset μ s), measure_empty, true_and, (hs.diff hs.everywherePosSubset).measure_eq_iSup_isCompact_of_ne_top A, ENNReal.iSup_eq_zero] intro k hk h'k diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 6f83a54a485eb..657fa9af35145 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -222,8 +222,7 @@ variable {R : Type*} [SMul R ℝ≥0] [SMul R ℝ≥0∞] [IsScalarTower R ℝ instance instSMul : SMul R (FiniteMeasure Ω) where smul (c : R) μ := ⟨c • (μ : Measure Ω), MeasureTheory.isFiniteMeasureSMulOfNNRealTower⟩ --- Porting note: with `simp` here the `coeFn` lemmas below fall prey to `simpNF`: the LHS simplifies -@[norm_cast] +@[simp, norm_cast] theorem toMeasure_zero : ((↑) : FiniteMeasure Ω → Measure Ω) 0 = 0 := rfl #align measure_theory.finite_measure.coe_zero MeasureTheory.FiniteMeasure.toMeasure_zero @@ -240,9 +239,8 @@ theorem toMeasure_smul (c : R) (μ : FiniteMeasure Ω) : ↑(c • μ) = c • ( rfl #align measure_theory.finite_measure.coe_smul MeasureTheory.FiniteMeasure.toMeasure_smul -@[simp, norm_cast] -theorem coeFn_zero : (⇑(0 : FiniteMeasure Ω) : Set Ω → ℝ≥0) = (0 : Set Ω → ℝ≥0) := by - funext; rfl +@[norm_cast] +theorem coeFn_zero : (⇑(0 : FiniteMeasure Ω) : Set Ω → ℝ≥0) = (0 : Set Ω → ℝ≥0) := rfl #align measure_theory.finite_measure.coe_fn_zero MeasureTheory.FiniteMeasure.coeFn_zero @[simp, norm_cast] diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean index b32e05b93ac83..bd80ac705eea2 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasureProd.lean @@ -71,7 +71,7 @@ lemma prod_prod (s : Set α) (t : Set β) : μ.prod ν (s ×ˢ t) = μ s * ν t apply Subtype.ext simp only [val_eq_toMeasure, toMeasure_map, toMeasure_prod, Measure.map_fst_prod] ext s _ - simp only [Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, smul_eq_mul] + simp only [Measure.coe_smul, Pi.smul_apply, smul_eq_mul] have aux := coeFn_smul_apply (ν univ) μ s simpa using congr_arg ENNReal.ofNNReal aux.symm @@ -79,8 +79,7 @@ lemma prod_prod (s : Set α) (t : Set β) : μ.prod ν (s ×ˢ t) = μ s * ν t apply Subtype.ext simp only [val_eq_toMeasure, toMeasure_map, toMeasure_prod, Measure.map_fst_prod] ext s _ - simp only [Measure.map_snd_prod, Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, - Pi.smul_apply, smul_eq_mul] + simp only [Measure.map_snd_prod, Measure.coe_smul, Pi.smul_apply, smul_eq_mul] have aux := coeFn_smul_apply (μ univ) ν s simpa using congr_arg ENNReal.ofNNReal aux.symm diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index cdf98ee54ced5..8688faf459444 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -636,8 +636,7 @@ instance isMulLeftInvariant_haarMeasure (K₀ : PositiveCompacts G) : @[to_additive] theorem haarMeasure_self {K₀ : PositiveCompacts G} : haarMeasure K₀ K₀ = 1 := by haveI : LocallyCompactSpace G := K₀.locallyCompactSpace_of_group - simp only [haarMeasure, smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, - smul_eq_mul] + simp only [haarMeasure, coe_smul, Pi.smul_apply, smul_eq_mul] rw [← OuterRegular.measure_closure_eq_of_isCompact K₀.isCompact, Content.measure_apply _ isClosed_closure.measurableSet, ENNReal.inv_mul_cancel] · exact (haarContent_outerMeasure_closure_pos K₀).ne' diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 7159602d76a95..9c9f1e13c2463 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -468,7 +468,7 @@ end Measure theorem OuterMeasure.coe_mkMetric [MeasurableSpace X] [BorelSpace X] (m : ℝ≥0∞ → ℝ≥0∞) : ⇑(OuterMeasure.mkMetric m : OuterMeasure X) = Measure.mkMetric m := by - rw [← Measure.mkMetric_toOuterMeasure] + rw [← Measure.mkMetric_toOuterMeasure, Measure.coe_toOuterMeasure] #align measure_theory.outer_measure.coe_mk_metric MeasureTheory.OuterMeasure.coe_mkMetric namespace Measure @@ -870,12 +870,8 @@ variable {f : X → Y} {d : ℝ} theorem hausdorffMeasure_image (hf : Isometry f) (hd : 0 ≤ d ∨ Surjective f) (s : Set X) : μH[d] (f '' s) = μH[d] s := by simp only [hausdorffMeasure, ← OuterMeasure.coe_mkMetric, ← OuterMeasure.comap_apply] - -- Porting note: this proof was slightly nicer before the port - simp only [mkMetric_toOuterMeasure] - have : 0 ≤ d → Monotone fun r : ℝ≥0∞ ↦ r ^ d := by - exact fun hd x y hxy => ENNReal.rpow_le_rpow hxy hd - have := OuterMeasure.isometry_comap_mkMetric (fun (r : ℝ≥0∞) => r ^ d) hf (hd.imp_left this) - congr + rw [OuterMeasure.isometry_comap_mkMetric _ hf (hd.imp_left _)] + exact ENNReal.monotone_rpow_of_nonneg #align isometry.hausdorff_measure_image Isometry.hausdorffMeasure_image theorem hausdorffMeasure_preimage (hf : Isometry f) (hd : 0 ≤ d ∨ Surjective f) (s : Set Y) : diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 4e09341327e38..08209cec60ce9 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -377,8 +377,7 @@ theorem addHaar_preimage_smul {r : ℝ} (hr : r ≠ 0) (s : Set E) : μ ((r • ·) ⁻¹' s) = Measure.map (r • ·) μ s := ((Homeomorph.smul (isUnit_iff_ne_zero.2 hr).unit).toMeasurableEquiv.map_apply s).symm _ = ENNReal.ofReal (abs (r ^ finrank ℝ E)⁻¹) * μ s := by - rw [map_addHaar_smul μ hr, smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, - smul_eq_mul] + rw [map_addHaar_smul μ hr, coe_smul, Pi.smul_apply, smul_eq_mul] #align measure_theory.measure.add_haar_preimage_smul MeasureTheory.Measure.addHaar_preimage_smul /-- Rescaling a set by a factor `r` multiplies its measure by `abs (r ^ dim)`. -/ diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index cde400cdee133..660870a1fd67e 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -745,10 +745,7 @@ theorem toMeasure_toOuterMeasure (m : OuterMeasure α) (h : ms ≤ m.caratheodor rfl #align measure_theory.to_measure_to_outer_measure MeasureTheory.toMeasure_toOuterMeasure --- Porting note: A coercion is directly elaborated in Lean4, so the LHS is simplified by --- `toMeasure_toOuterMeasure` even if this theorem has high priority. --- Instead of this theorem, we give `simp` attr to `OuterMeasure.trim_eq`. --- @[simp] +@[simp] theorem toMeasure_apply (m : OuterMeasure α) (h : ms ≤ m.caratheodory) {s : Set α} (hs : MeasurableSet s) : m.toMeasure h s = m s := m.trim_eq hs @@ -890,7 +887,7 @@ instance instSMul [MeasurableSpace α] : SMul R (Measure α) := ⟨fun c μ => { toOuterMeasure := c • μ.toOuterMeasure m_iUnion := fun s hs hd => by - simp_rw [OuterMeasure.smul_apply, measure_iUnion hd hs] + simp_rw [OuterMeasure.smul_apply, coe_toOuterMeasure, measure_iUnion hd hs] simpa using (ENNReal.tsum_mul_left (a := c • 1)).symm trimmed := by rw [OuterMeasure.trim_smul, μ.trimmed] }⟩ #align measure_theory.measure.has_smul MeasureTheory.Measure.instSMul @@ -971,10 +968,7 @@ instance instModule [Semiring R] [Module R ℝ≥0∞] [IsScalarTower R ℝ≥0 toOuterMeasure_injective smul_toOuterMeasure #align measure_theory.measure.module MeasureTheory.Measure.instModule --- Porting note: A coercion is directly elaborated in Lean4, so the LHS is simplified by --- `smul_toOuterMeasure` even if this theorem has high priority. --- Instead of this theorem, we give `simp` attr to `nnreal_smul_coe_apply`. --- @[simp] +@[simp] theorem coe_nnreal_smul_apply {_m : MeasurableSpace α} (c : ℝ≥0) (μ : Measure α) (s : Set α) : (c • μ) s = c * μ s := rfl @@ -1041,7 +1035,7 @@ theorem toOuterMeasure_le : μ₁.toOuterMeasure ≤ μ₂.toOuterMeasure ↔ μ #align measure_theory.measure.to_outer_measure_le MeasureTheory.Measure.toOuterMeasure_le theorem le_iff : μ₁ ≤ μ₂ ↔ ∀ s, MeasurableSet s → μ₁ s ≤ μ₂ s := by - rw [← toOuterMeasure_le, ← OuterMeasure.le_trim_iff, μ₂.trimmed] + erw [← toOuterMeasure_le, ← OuterMeasure.le_trim_iff, μ₂.trimmed] #align measure_theory.measure.le_iff MeasureTheory.Measure.le_iff theorem le_intro (h : ∀ s, MeasurableSet s → s.Nonempty → μ₁ s ≤ μ₂ s) : μ₁ ≤ μ₂ := @@ -1078,8 +1072,8 @@ theorem sInf_caratheodory (s : Set α) (hs : MeasurableSet s) : MeasurableSet[(sInf (toOuterMeasure '' m)).caratheodory] s := by rw [OuterMeasure.sInf_eq_boundedBy_sInfGen] refine' OuterMeasure.boundedBy_caratheodory fun t => _ - simp only [OuterMeasure.sInfGen, le_iInf_iff, forall_mem_image, - measure_eq_iInf t] + simp only [OuterMeasure.sInfGen, le_iInf_iff, forall_mem_image, measure_eq_iInf t, + coe_toOuterMeasure] intro μ hμ u htu _hu have hm : ∀ {s t}, s ⊆ t → OuterMeasure.sInfGen (toOuterMeasure '' m) s ≤ μ t := by intro s t hst @@ -1320,7 +1314,7 @@ theorem map_toOuterMeasure (hf : AEMeasurable f μ) : (μ.map f).toOuterMeasure = (OuterMeasure.map f μ.toOuterMeasure).trim := by rw [← trimmed, OuterMeasure.trim_eq_trim_iff] intro s hs - rw [map_apply_of_aemeasurable hf hs, OuterMeasure.map_apply] + simp [hf, hs] #align measure_theory.measure.map_to_outer_measure MeasureTheory.Measure.map_toOuterMeasure @[simp] lemma map_eq_zero_iff (hf : AEMeasurable f μ) : μ.map f = 0 ↔ μ = 0 := by @@ -1395,7 +1389,7 @@ def comapₗ [MeasurableSpace α] (f : α → β) : Measure β →ₗ[ℝ≥0∞ theorem comapₗ_apply {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) (hfi : Injective f) (hf : ∀ s, MeasurableSet s → MeasurableSet (f '' s)) (μ : Measure β) (hs : MeasurableSet s) : comapₗ f μ s = μ (f '' s) := by - rw [comapₗ, dif_pos, liftLinear_apply _ hs, OuterMeasure.comap_apply] + rw [comapₗ, dif_pos, liftLinear_apply _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] exact ⟨hfi, hf⟩ #align measure_theory.measure.comapₗ_apply MeasureTheory.Measure.comapₗ_apply @@ -1413,7 +1407,7 @@ theorem comap_apply₀ [MeasurableSpace α] (f : α → β) (μ : Measure β) (h (hf : ∀ s, MeasurableSet s → NullMeasurableSet (f '' s) μ) (hs : NullMeasurableSet s (comap f μ)) : comap f μ s = μ (f '' s) := by rw [comap, dif_pos (And.intro hfi hf)] at hs ⊢ - rw [toMeasure_apply₀ _ _ hs, OuterMeasure.comap_apply] + rw [toMeasure_apply₀ _ _ hs, OuterMeasure.comap_apply, coe_toOuterMeasure] #align measure_theory.measure.comap_apply₀ MeasureTheory.Measure.comap_apply₀ theorem le_comap_apply {β} [MeasurableSpace α] {mβ : MeasurableSpace β} (f : α → β) (μ : Measure β) @@ -1572,9 +1566,6 @@ theorem sum_fintype [Fintype ι] (μ : ι → Measure α) : sum μ = ∑ i, μ i simp only [sum_apply, finset_sum_apply, hs, tsum_fintype] #align measure_theory.measure.sum_fintype MeasureTheory.Measure.sum_fintype --- Porting note: The LHS is simplified by --- `sum_fintype` even if this theorem has high priority. -@[simp, nolint simpNF] theorem sum_coe_finset (s : Finset ι) (μ : ι → Measure α) : (sum fun i : s => μ i) = ∑ i in s, μ i := by rw [sum_fintype, Finset.sum_coe_sort s μ] #align measure_theory.measure.sum_coe_finset MeasureTheory.Measure.sum_coe_finset @@ -1584,17 +1575,15 @@ theorem ae_sum_eq [Countable ι] (μ : ι → Measure α) : (sum μ).ae = ⨆ i, Filter.ext fun _ => ae_sum_iff.trans mem_iSup.symm #align measure_theory.measure.ae_sum_eq MeasureTheory.Measure.ae_sum_eq --- @[simp] -- Porting note (#10618): simp can prove this theorem sum_bool (f : Bool → Measure α) : sum f = f true + f false := by rw [sum_fintype, Fintype.sum_bool] #align measure_theory.measure.sum_bool MeasureTheory.Measure.sum_bool --- @[simp] -- Porting note (#10618): simp can prove this theorem sum_cond (μ ν : Measure α) : (sum fun b => cond b μ ν) = μ + ν := sum_bool _ #align measure_theory.measure.sum_cond MeasureTheory.Measure.sum_cond --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem sum_of_empty [IsEmpty ι] (μ : ι → Measure α) : sum μ = 0 := by rw [← measure_univ_eq_zero, sum_apply _ MeasurableSet.univ, tsum_empty] #align measure_theory.measure.sum_of_empty MeasureTheory.Measure.sum_of_empty @@ -1691,7 +1680,7 @@ protected theorem smul [Monoid R] [DistribMulAction R ℝ≥0∞] [IsScalarTower protected lemma add (h1 : μ₁ ≪ ν) (h2 : μ₂ ≪ ν') : μ₁ + μ₂ ≪ ν + ν' := by intro s hs - simp only [add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ + simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ exact ⟨h1 hs.1, h2 hs.2⟩ lemma add_left_iff {μ₁ μ₂ ν : Measure α} : @@ -1705,7 +1694,7 @@ lemma add_left_iff {μ₁ μ₂ ν : Measure α} : lemma add_right (h1 : μ ≪ ν) (ν' : Measure α) : μ ≪ ν + ν' := by intro s hs - simp only [add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ + simp only [coe_add, Pi.add_apply, add_eq_zero] at hs ⊢ exact h1 hs.1 end AbsolutelyContinuous diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 8bd7408a411c1..20ac9032356b6 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -78,23 +78,21 @@ namespace MeasureTheory measurable sets, with the additional assumption that the outer measure is the canonical extension of the restricted measure. -/ structure Measure (α : Type*) [MeasurableSpace α] extends OuterMeasure α where - m_iUnion ⦃f : ℕ → Set α⦄ : - (∀ i, MeasurableSet (f i)) → - Pairwise (Disjoint on f) → toOuterMeasure (⋃ i, f i) = ∑' i, toOuterMeasure (f i) + m_iUnion ⦃f : ℕ → Set α⦄ : (∀ i, MeasurableSet (f i)) → Pairwise (Disjoint on f) → + toOuterMeasure (⋃ i, f i) = ∑' i, toOuterMeasure (f i) trimmed : toOuterMeasure.trim = toOuterMeasure #align measure_theory.measure MeasureTheory.Measure -/-- Measure projections for a measure space. +theorem Measure.toOuterMeasure_injective [MeasurableSpace α] : + Injective (toOuterMeasure : Measure α → OuterMeasure α) + | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl => rfl +#align measure_theory.measure.to_outer_measure_injective MeasureTheory.Measure.toOuterMeasure_injective -For measurable sets this returns the measure assigned by the `measureOf` field in `Measure`. -But we can extend this to _all_ sets, but using the outer measure. This gives us monotonicity and -subadditivity for all sets. --/ -instance Measure.instCoeFun [MeasurableSpace α] : CoeFun (Measure α) fun _ => Set α → ℝ≥0∞ := - ⟨fun m => m.toOuterMeasure⟩ -#align measure_theory.measure.has_coe_to_fun MeasureTheory.Measure.instCoeFun +instance Measure.instFunLike [MeasurableSpace α] : FunLike (Measure α) (Set α) ℝ≥0∞ where + coe μ := μ.toOuterMeasure + coe_injective' | ⟨_, _, _⟩, ⟨_, _, _⟩, h => toOuterMeasure_injective <| DFunLike.coe_injective h -attribute [coe] Measure.toOuterMeasure +#noalign measure_theory.measure.has_coe_to_fun section @@ -104,7 +102,6 @@ namespace Measure /-! ### General facts about measures -/ - /-- Obtain a measure by giving a countably additive function that sends `∅` to `0`. -/ def ofMeasurable (m : ∀ s : Set α, MeasurableSet s → ℝ≥0∞) (m0 : m ∅ MeasurableSet.empty = 0) (mU : @@ -132,11 +129,6 @@ theorem ofMeasurable_apply {m : ∀ s : Set α, MeasurableSet s → ℝ≥0∞} inducedOuterMeasure_eq m0 mU hs #align measure_theory.measure.of_measurable_apply MeasureTheory.Measure.ofMeasurable_apply -theorem toOuterMeasure_injective : Injective (toOuterMeasure : Measure α → OuterMeasure α) := - fun ⟨m₁, u₁, h₁⟩ ⟨m₂, _u₂, _h₂⟩ _h => by - congr -#align measure_theory.measure.to_outer_measure_injective MeasureTheory.Measure.toOuterMeasure_injective - @[ext] theorem ext (h : ∀ s, MeasurableSet s → μ₁ s = μ₂ s) : μ₁ = μ₂ := toOuterMeasure_injective <| by @@ -152,15 +144,20 @@ theorem ext_iff' : μ₁ = μ₂ ↔ ∀ s, μ₁ s = μ₂ s := end Measure -#noalign measure_theory.coe_to_outer_measure +@[simp] theorem Measure.coe_toOuterMeasure (μ : Measure α) : ⇑μ.toOuterMeasure = μ := rfl +#align measure_theory.coe_to_outer_measure MeasureTheory.Measure.coe_toOuterMeasure -#noalign measure_theory.to_outer_measure_apply +theorem Measure.toOuterMeasure_apply (μ : Measure α) (s : Set α) : + μ.toOuterMeasure s = μ s := + rfl +#align measure_theory.to_outer_measure_apply MeasureTheory.Measure.toOuterMeasure_apply -theorem measure_eq_trim (s : Set α) : μ s = μ.toOuterMeasure.trim s := by rw [μ.trimmed] +theorem measure_eq_trim (s : Set α) : μ s = μ.toOuterMeasure.trim s := by + rw [μ.trimmed, μ.coe_toOuterMeasure] #align measure_theory.measure_eq_trim MeasureTheory.measure_eq_trim theorem measure_eq_iInf (s : Set α) : μ s = ⨅ (t) (_ : s ⊆ t) (_ : MeasurableSet t), μ t := by - rw [measure_eq_trim, OuterMeasure.trim_eq_iInf] + rw [measure_eq_trim, OuterMeasure.trim_eq_iInf, μ.coe_toOuterMeasure] #align measure_theory.measure_eq_infi MeasureTheory.measure_eq_iInf /-- A variant of `measure_eq_iInf` which has a single `iInf`. This is useful when applying a @@ -187,7 +184,7 @@ theorem measure_eq_extend (hs : MeasurableSet s) : exact hs #align measure_theory.measure_eq_extend MeasureTheory.measure_eq_extend --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem measure_empty : μ ∅ = 0 := μ.empty #align measure_theory.measure_empty MeasureTheory.measure_empty @@ -277,7 +274,7 @@ theorem measure_iUnion_null [Countable ι] {s : ι → Set α} : (∀ i, μ (s i μ.toOuterMeasure.iUnion_null #align measure_theory.measure_Union_null MeasureTheory.measure_iUnion_null --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem measure_iUnion_null_iff [Countable ι] {s : ι → Set α} : μ (⋃ i, s i) = 0 ↔ ∀ i, μ (s i) = 0 := μ.toOuterMeasure.iUnion_null_iff diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index 4dfaaf65579a3..5600282f7b004 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -353,8 +353,8 @@ def normalize : ProbabilityMeasure Ω := refine' ⟨_⟩ -- Porting note: paying the price that this isn't `simp` lemma now. rw [FiniteMeasure.toMeasure_smul] - simp only [Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, - Measure.nnreal_smul_coe_apply, ne_eq, mass_zero_iff, ENNReal.coe_inv zero, ennreal_mass] + simp only [Measure.coe_smul, Pi.smul_apply, Measure.nnreal_smul_coe_apply, ne_eq, + mass_zero_iff, ENNReal.coe_inv zero, ennreal_mass] rw [← Ne, ← ENNReal.coe_ne_zero, ennreal_mass] at zero exact ENNReal.inv_mul_cancel zero μ.prop.measure_univ_lt_top.ne } #align measure_theory.finite_measure.normalize MeasureTheory.FiniteMeasure.normalize @@ -362,15 +362,10 @@ def normalize : ProbabilityMeasure Ω := @[simp] theorem self_eq_mass_mul_normalize (s : Set Ω) : μ s = μ.mass * μ.normalize s := by obtain rfl | h := eq_or_ne μ 0 - · simp only [zero_mass, coeFn_zero, Pi.zero_apply, zero_mul] - rfl + · simp have mass_nonzero : μ.mass ≠ 0 := by rwa [μ.mass_nonzero_iff] simp only [normalize, dif_neg mass_nonzero] - change μ s = mass μ * ((mass μ)⁻¹ • μ) s - -- Porting note: this `change` is a hack, but I had trouble coming up with something better - simp only [toMeasure_smul, Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, - Measure.nnreal_smul_coe_apply, ne_eq, mass_zero_iff, ENNReal.toNNReal_mul, ENNReal.toNNReal_coe, - mul_inv_cancel_left₀ mass_nonzero] + simp [ProbabilityMeasure.coe_mk, toMeasure_smul, mul_inv_cancel_left₀ mass_nonzero] #align measure_theory.finite_measure.self_eq_mass_mul_normalize MeasureTheory.FiniteMeasure.self_eq_mass_mul_normalize theorem self_eq_mass_smul_normalize : μ = μ.mass • μ.normalize.toFiniteMeasure := by diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 0f3b8268c8317..6580dba0637d0 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -60,12 +60,9 @@ theorem restrict_toOuterMeasure_eq_toOuterMeasure_restrict (h : MeasurableSet s) toMeasure_toOuterMeasure, OuterMeasure.restrict_trim h, μ.trimmed] #align measure_theory.measure.restrict_to_outer_measure_eq_to_outer_measure_restrict MeasureTheory.Measure.restrict_toOuterMeasure_eq_toOuterMeasure_restrict -theorem restrict_apply₀ (ht : NullMeasurableSet t (μ.restrict s)) : μ.restrict s t = μ (t ∩ s) := - (toMeasure_apply₀ _ (fun s' hs' t => by - suffices μ (s ∩ t) = μ (s ∩ t ∩ s') + μ ((s ∩ t) \ s') by - simpa [← Set.inter_assoc, Set.inter_comm _ s, ← inter_diff_assoc] - exact le_toOuterMeasure_caratheodory _ _ hs' _) ht).trans <| by - simp only [OuterMeasure.restrict_apply] +theorem restrict_apply₀ (ht : NullMeasurableSet t (μ.restrict s)) : μ.restrict s t = μ (t ∩ s) := by + rw [← restrictₗ_apply, restrictₗ, liftLinear_apply₀ _ ht, OuterMeasure.restrict_apply, + coe_toOuterMeasure] #align measure_theory.measure.restrict_apply₀ MeasureTheory.Measure.restrict_apply₀ /-- If `t` is a measurable set, then the measure of `t` with respect to the restriction of @@ -106,8 +103,9 @@ the measure to `s` equals the outer measure of `t ∩ s`. This is an alternate v `Measure.restrict_apply`, requiring that `s` is measurable instead of `t`. -/ @[simp] theorem restrict_apply' (hs : MeasurableSet s) : μ.restrict s t = μ (t ∩ s) := by - rw [Measure.restrict_toOuterMeasure_eq_toOuterMeasure_restrict hs, - OuterMeasure.restrict_apply s t _] + rw [← toOuterMeasure_apply, + Measure.restrict_toOuterMeasure_eq_toOuterMeasure_restrict hs, + OuterMeasure.restrict_apply s t _, toOuterMeasure_apply] #align measure_theory.measure.restrict_apply' MeasureTheory.Measure.restrict_apply' theorem restrict_apply₀' (hs : NullMeasurableSet s μ) : μ.restrict s t = μ (t ∩ s) := by diff --git a/Mathlib/MeasureTheory/Measure/Tilted.lean b/Mathlib/MeasureTheory/Measure/Tilted.lean index e35bdb7e197e2..a69aec40e9d30 100644 --- a/Mathlib/MeasureTheory/Measure/Tilted.lean +++ b/Mathlib/MeasureTheory/Measure/Tilted.lean @@ -109,9 +109,8 @@ lemma tilted_apply_eq_ofReal_integral' {s : Set α} (f : α → ℝ) (hs : Measu · rw [tilted_apply' _ _ hs, ← ofReal_integral_eq_lintegral_ofReal] · exact hf.integrableOn.div_const _ · exact ae_of_all _ (fun _ ↦ by positivity) - · simp only [hf, not_false_eq_true, tilted_of_not_integrable, Measure.zero_toOuterMeasure, - OuterMeasure.coe_zero, Pi.zero_apply, integral_undef hf, div_zero, integral_zero, - ENNReal.ofReal_zero] + · simp only [hf, not_false_eq_true, tilted_of_not_integrable, Measure.coe_zero, + Pi.zero_apply, integral_undef hf, div_zero, integral_zero, ENNReal.ofReal_zero] lemma tilted_apply_eq_ofReal_integral [SFinite μ] (f : α → ℝ) (s : Set α) : μ.tilted f s = ENNReal.ofReal (∫ a in s, exp (f a) / ∫ x, exp (f x) ∂μ ∂μ) := by @@ -119,9 +118,7 @@ lemma tilted_apply_eq_ofReal_integral [SFinite μ] (f : α → ℝ) (s : Set α) · rw [tilted_apply _ _, ← ofReal_integral_eq_lintegral_ofReal] · exact hf.integrableOn.div_const _ · exact ae_of_all _ (fun _ ↦ by positivity) - · simp only [hf, not_false_eq_true, tilted_of_not_integrable, Measure.zero_toOuterMeasure, - OuterMeasure.coe_zero, Pi.zero_apply, integral_undef hf, div_zero, integral_zero, - ENNReal.ofReal_zero] + · simp [tilted_of_not_integrable hf, integral_undef hf] instance isFiniteMeasure_tilted : IsFiniteMeasure (μ.tilted f) := by by_cases hf : Integrable (fun x ↦ exp (f x)) μ diff --git a/Mathlib/MeasureTheory/Measure/Trim.lean b/Mathlib/MeasureTheory/Measure/Trim.lean index f836cf45d904b..b7782d9dbb9f0 100644 --- a/Mathlib/MeasureTheory/Measure/Trim.lean +++ b/Mathlib/MeasureTheory/Measure/Trim.lean @@ -52,7 +52,7 @@ theorem zero_trim (hm : m ≤ m0) : (0 : Measure α).trim hm = (0 : @Measure α #align measure_theory.zero_trim MeasureTheory.zero_trim theorem trim_measurableSet_eq (hm : m ≤ m0) (hs : @MeasurableSet α m s) : μ.trim hm s = μ s := by - rw [Measure.trim, toMeasure_apply (ms := m) _ _ hs] + rw [Measure.trim, toMeasure_apply (ms := m) _ _ hs, Measure.coe_toOuterMeasure] #align measure_theory.trim_measurable_set_eq MeasureTheory.trim_measurableSet_eq theorem le_trim (hm : m ≤ m0) : μ s ≤ μ.trim hm s := by diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 1e585b62cb9c7..eec9c96cea0b6 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -498,7 +498,7 @@ theorem FiniteAtFilter.exists_mem_basis {f : Filter α} (hμ : FiniteAtFilter μ #align measure_theory.measure.finite_at_filter.exists_mem_basis MeasureTheory.Measure.FiniteAtFilter.exists_mem_basis theorem finiteAtBot {m0 : MeasurableSpace α} (μ : Measure α) : μ.FiniteAtFilter ⊥ := - ⟨∅, mem_bot, by simp only [OuterMeasure.empty', zero_lt_top]⟩ + ⟨∅, mem_bot, by simp only [measure_empty, zero_lt_top]⟩ #align measure_theory.measure.finite_at_bot MeasureTheory.Measure.finiteAtBot /-- `μ` has finite spanning sets in `C` if there is a countable sequence of sets in `C` that have @@ -1026,7 +1026,7 @@ theorem sigmaFinite_of_le (μ : Measure α) [hs : SigmaFinite μ] (h : ν ≤ μ ext s hs rw [← ENNReal.add_right_inj (measure_mono (inter_subset_right s _) |>.trans_lt <| measure_spanningSets_lt_top μ i).ne] - simp only [ext_iff', add_toOuterMeasure, OuterMeasure.coe_add, Pi.add_apply] at h + simp only [ext_iff', coe_add, Pi.add_apply] at h simp [hs, h] @[simp] lemma add_left_inj (μ ν₁ ν₂ : Measure α) [SigmaFinite μ] : @@ -1096,8 +1096,7 @@ instance SMul.sigmaFinite {μ : Measure α} [SigmaFinite μ] (c : ℝ≥0) : set_mem := fun _ ↦ trivial finite := by intro i - simp only [smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, - nnreal_smul_coe_apply] + simp only [Measure.coe_smul, Pi.smul_apply, nnreal_smul_coe_apply] exact ENNReal.mul_lt_top ENNReal.coe_ne_top (measure_spanningSets_lt_top μ i).ne spanning := iUnion_spanningSets μ }⟩ diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index 24612c177be55..3bb4930797efb 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -92,8 +92,8 @@ theorem cond_isProbabilityMeasure_of_finite (hcs : μ s ≠ 0) (hs : μ s ≠ IsProbabilityMeasure μ[|s] := ⟨by unfold ProbabilityTheory.cond - simp only [Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, - MeasurableSet.univ, Measure.restrict_apply, Set.univ_inter, smul_eq_mul] + simp only [Measure.coe_smul, Pi.smul_apply, MeasurableSet.univ, Measure.restrict_apply, + Set.univ_inter, smul_eq_mul] exact ENNReal.inv_mul_cancel hcs hs⟩ /-- The conditional probability measure of any finite measure on any set of positive measure @@ -104,9 +104,8 @@ theorem cond_isProbabilityMeasure [IsFiniteMeasure μ] (hcs : μ s ≠ 0) : instance cond_isFiniteMeasure : IsFiniteMeasure μ[|s] := by constructor - simp only [Measure.smul_toOuterMeasure, OuterMeasure.coe_smul, Pi.smul_apply, MeasurableSet.univ, - Measure.restrict_apply, Set.univ_inter, smul_eq_mul, ProbabilityTheory.cond, - ← ENNReal.div_eq_inv_mul] + simp only [Measure.coe_smul, Pi.smul_apply, MeasurableSet.univ, Measure.restrict_apply, + Set.univ_inter, smul_eq_mul, ProbabilityTheory.cond, ← ENNReal.div_eq_inv_mul] exact ENNReal.div_self_le_one.trans_lt ENNReal.one_lt_top theorem cond_toMeasurable_eq : @@ -234,8 +233,8 @@ lemma sum_meas_smul_cond_fiber {X : Ω → α} (hX : Measurable X) (μ : Measure ext E hE calc _ = ∑ x, μ (X ⁻¹' {x} ∩ E) := by - simp only [Measure.coe_finset_sum, smul_toOuterMeasure, OuterMeasure.coe_smul, - Finset.sum_apply, Pi.smul_apply, smul_eq_mul] + simp only [Measure.coe_finset_sum, Measure.coe_smul, Finset.sum_apply, + 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 diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 43cb16faff4c0..0afae13eb3457 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -580,8 +580,7 @@ lemma compProd_add_right (μ : kernel α β) (κ η : kernel (α × β) γ) [IsSFiniteKernel μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] : μ ⊗ₖ (κ + η) = μ ⊗ₖ κ + μ ⊗ₖ η := by ext a s hs - simp only [compProd_apply _ _ _ hs, coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, - OuterMeasure.coe_add] + simp only [compProd_apply _ _ _ hs, coeFn_add, Pi.add_apply, Measure.coe_add] rw [lintegral_add_left] exact measurable_kernel_prod_mk_left' hs a diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index b7a7918012a99..81c302631d139 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -391,8 +391,7 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : kernel α (γ × β)) simp only [le_eq_subset, prod_subset_prod_iff, subset_rfl, true_and] exact Or.inl <| hseq hmm' · exact ⟨0, measure_ne_top _ _⟩ - · simp only [prod_empty, OuterMeasure.empty', ne_eq, not_true_eq_false, false_or, h0, - not_false_iff] + · exact .inr h0 lemma tendsto_densityProcess_atTop_of_antitone (κ : kernel α (γ × β)) (ν : kernel α γ) [IsFiniteKernel κ] (n : ℕ) (a : α) (x : γ) @@ -674,7 +673,7 @@ lemma tendsto_integral_density_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKern · refine ⟨0, measure_ne_top _ _⟩ convert h rw [← prod_iInter, hseq_iInter] - simp only [ne_eq, prod_empty, OuterMeasure.empty', forall_exists_index] + simp lemma tendsto_density_atTop_ae_of_antitone (hκν : fst κ ≤ ν) [IsFiniteKernel ν] (a : α) (seq : ℕ → Set β) (hseq : Antitone seq) (hseq_iInter : ⋂ i, seq i = ∅) diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index e89500c10fb88..d4b6b25a2c4a4 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -170,8 +170,7 @@ lemma withDensity_one_sub_rnDerivAux (κ η : kernel α γ) [IsFiniteKernel κ] = κ a s + η a s := by rw [this] simp - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add] - at h + simp only [coeFn_add, Pi.add_apply, Measure.coe_add] at h rwa [withDensity_rnDerivAux, add_comm, ENNReal.add_right_inj (measure_ne_top _ _)] at h have : ∀ b, (Real.toNNReal b : ℝ≥0∞) = ENNReal.ofReal b := fun _ ↦ rfl simp_rw [this, ENNReal.ofReal_sub _ (rnDerivAux_nonneg h_le), ENNReal.ofReal_one] @@ -401,8 +400,7 @@ lemma rnDeriv_add_singularPart (κ η : kernel α γ) [IsFiniteKernel κ] [IsFin withDensity η (rnDeriv κ η) + singularPart κ η = κ := by ext a s hs rw [← inter_union_diff s (mutuallySingularSetSlice κ η a)] - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, - OuterMeasure.coe_add] + simp only [coeFn_add, Pi.add_apply, Measure.coe_add] have hm := measurableSet_mutuallySingularSetSlice κ η a simp only [measure_union (Disjoint.mono (inter_subset_right _ _) subset_rfl disjoint_sdiff_right) (hs.diff hm)] @@ -459,7 +457,7 @@ lemma singularPart_eq_zero_iff_measure_eq_zero (κ η : kernel α γ) simp_rw [ext_iff, Measure.ext_iff] at h_eq_add specialize h_eq_add a (mutuallySingularSetSlice κ η a) (measurableSet_mutuallySingularSetSlice κ η a) - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add, + simp only [coeFn_add, Pi.add_apply, Measure.coe_add, withDensity_rnDeriv_mutuallySingularSetSlice κ η, zero_add] at h_eq_add rw [← h_eq_add] exact singularPart_eq_zero_iff_apply_eq_zero κ η a @@ -471,7 +469,7 @@ lemma withDensity_rnDeriv_eq_zero_iff_measure_eq_zero (κ η : kernel α γ) simp_rw [ext_iff, Measure.ext_iff] at h_eq_add specialize h_eq_add a (mutuallySingularSetSlice κ η a)ᶜ (measurableSet_mutuallySingularSetSlice κ η a).compl - simp only [coeFn_add, Pi.add_apply, Measure.add_toOuterMeasure, OuterMeasure.coe_add, + simp only [coeFn_add, Pi.add_apply, Measure.coe_add, singularPart_compl_mutuallySingularSetSlice κ η, add_zero] at h_eq_add rw [← h_eq_add] exact withDensity_rnDeriv_eq_zero_iff_apply_eq_zero κ η a diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index 2109e0867211b..1ad72c6f4cecd 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -151,7 +151,7 @@ theorem evariance_eq_zero_iff (hX : AEMeasurable X μ) : using hω · rw [hω] simp - · measurability + · exact (hX.sub_const _).ennnorm.pow_const _ -- TODO `measurability` and `fun_prop` fail #align probability_theory.evariance_eq_zero_iff ProbabilityTheory.evariance_eq_zero_iff theorem evariance_mul (c : ℝ) (X : Ω → ℝ) (μ : Measure Ω) : From fcceeb1dd138f3969ddff60db830c196ae1b761f Mon Sep 17 00:00:00 2001 From: kkytola Date: Thu, 9 May 2024 11:24:52 +0000 Subject: [PATCH 50/52] feat: BoundedSub class to generalize subtraction of bounded continuous functions. (#12559) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Introduce type class `BoundedSub`, which allows to generalize subtraction of `BoundedContinuousFunction` to the case where the values are not necessarily in `SeminormedAddGroup`. The most important case for measure theory applications that is allowed by the generalization is bounded continuous functions with values in `ℝ≥0`. Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- Mathlib.lean | 1 + .../Topology/Bornology/BoundedOperation.lean | 110 ++++++++++++++++++ .../Topology/ContinuousFunction/Bounded.lean | 37 +++--- .../Topology/MetricSpace/PseudoMetric.lean | 18 +++ 4 files changed, 151 insertions(+), 15 deletions(-) create mode 100644 Mathlib/Topology/Bornology/BoundedOperation.lean diff --git a/Mathlib.lean b/Mathlib.lean index d10438cf9a07b..298c2347f5bd7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3904,6 +3904,7 @@ import Mathlib.Topology.Bases import Mathlib.Topology.Basic import Mathlib.Topology.Bornology.Absorbs import Mathlib.Topology.Bornology.Basic +import Mathlib.Topology.Bornology.BoundedOperation import Mathlib.Topology.Bornology.Constructions import Mathlib.Topology.Bornology.Hom import Mathlib.Topology.Category.Born diff --git a/Mathlib/Topology/Bornology/BoundedOperation.lean b/Mathlib/Topology/Bornology/BoundedOperation.lean new file mode 100644 index 0000000000000..e84b8fd1366b6 --- /dev/null +++ b/Mathlib/Topology/Bornology/BoundedOperation.lean @@ -0,0 +1,110 @@ +/- +Copyright (c) 2024 Kalle Kytölä +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kalle Kytölä +-/ +import Mathlib.Analysis.Normed.Group.Basic + +/-! +# Bounded operations + +This file introduces type classes for bornologically bounded operations. + +In particular, when combined with type classes which guarantee continuity of the same operations, +we can equip bounded continuous functions with the corresponding operations. + +## Main definitions + +* `BoundedSub R`: a class guaranteeing boundedness of subtraction. + +TODO: +* Add bounded multiplication. (So that, e.g., multiplication works in `X →ᵇ ℝ≥0`.) + +-/ + +section bounded_sub +/-! +### Bounded subtraction +-/ + +open Pointwise + +/-- A typeclass saying that `(p : R × R) ↦ p.1 - p.2` maps any pair of bounded sets to a bounded +set. This property automatically holds for seminormed additive groups, but it also holds, e.g., +for `ℝ≥0`. -/ +class BoundedSub (R : Type*) [Bornology R] [Sub R] : Prop where + isBounded_sub : ∀ {s t : Set R}, + Bornology.IsBounded s → Bornology.IsBounded t → Bornology.IsBounded (s - t) + +variable {R : Type*} + +lemma isBounded_sub [Bornology R] [Sub R] [BoundedSub R] {s t : Set R} + (hs : Bornology.IsBounded s) (ht : Bornology.IsBounded t) : + Bornology.IsBounded (s - t) := BoundedSub.isBounded_sub hs ht + +lemma sub_bounded_of_bounded_of_bounded {X : Type*} [PseudoMetricSpace R] [Sub R] [BoundedSub R] + {f g : X → R} (f_bdd : ∃ C, ∀ x y, dist (f x) (f y) ≤ C) + (g_bdd : ∃ C, ∀ x y, dist (g x) (g y) ≤ C) : + ∃ C, ∀ x y, dist ((f - g) x) ((f - g) y) ≤ C := by + obtain ⟨C, hC⟩ := Metric.isBounded_iff.mp <| + isBounded_sub (Metric.isBounded_range_iff.mpr f_bdd) (Metric.isBounded_range_iff.mpr g_bdd) + use C + intro x y + exact hC (Set.sub_mem_sub (Set.mem_range_self (f := f) x) (Set.mem_range_self (f := g) x)) + (Set.sub_mem_sub (Set.mem_range_self (f := f) y) (Set.mem_range_self (f := g) y)) + +end bounded_sub + +section SeminormedAddCommGroup +/-! +### Bounded operations in seminormed additive commutative groups +-/ + +variable {R : Type*} [SeminormedAddCommGroup R] + +lemma SeminormedAddCommGroup.lipschitzWith_sub : + LipschitzWith 2 (fun (p : R × R) ↦ p.1 - p.2) := by + convert LipschitzWith.prod_fst.sub LipschitzWith.prod_snd + norm_num + +instance : BoundedSub R where + isBounded_sub := by + intro s t hs ht + rw [Metric.isBounded_iff] at hs ht ⊢ + obtain ⟨Cf, hf⟩ := hs + obtain ⟨Cg, hg⟩ := ht + use Cf + Cg + intro z hz w hw + obtain ⟨x₁, hx₁, y₁, hy₁, z_eq⟩ := Set.mem_sub.mp hz + obtain ⟨x₂, hx₂, y₂, hy₂, w_eq⟩ := Set.mem_sub.mp hw + rw [← w_eq, ← z_eq, dist_eq_norm] + calc ‖x₁ - y₁ - (x₂ - y₂)‖ + _ = ‖x₁ - x₂ + (y₂ - y₁)‖ := by + rw [sub_sub_sub_comm, sub_eq_add_neg, neg_sub] + _ ≤ ‖x₁ - x₂‖ + ‖y₂ - y₁‖ := norm_add_le _ _ + _ ≤ Cf + Cg := + add_le_add (mem_closedBall_iff_norm.mp (hf hx₁ hx₂)) + (mem_closedBall_iff_norm.mp (hg hy₂ hy₁)) + +end SeminormedAddCommGroup + +section NNReal +/-! +### Bounded operations in ℝ≥0 +-/ + +open scoped NNReal + +noncomputable instance : BoundedSub ℝ≥0 where + isBounded_sub := by + intro s t hs _ + obtain ⟨c, hc⟩ := (Metric.isBounded_iff_subset_ball 0).mp hs + apply (Metric.isBounded_iff_subset_ball 0).mpr + use c + intro z hz + obtain ⟨a, a_in_s, b, _, z_eq⟩ := Set.mem_sub.mp hz + have key := hc a_in_s + simp only [NNReal.ball_zero_eq_Ico, ← z_eq, Set.mem_Ico, zero_le, true_and, gt_iff_lt] at * + exact tsub_lt_of_lt key + +end NNReal diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index 39ec9997fa04d..1367c04083dcc 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -9,6 +9,7 @@ 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.Topology.Bornology.BoundedOperation #align_import topology.continuous_function.bounded from "leanprover-community/mathlib"@"5dc275ec639221ca4d5f56938eb966f6ad9bc89f" @@ -767,6 +768,27 @@ theorem sum_apply {ι : Type*} (s : Finset ι) (f : ι → α →ᵇ β) (a : α end CommHasLipschitzAdd +section sub + +variable [TopologicalSpace α] +variable {R : Type*} [PseudoMetricSpace R] [Sub R] [BoundedSub R] [ContinuousSub R] +variable (f g : α →ᵇ R) + +/-- The pointwise difference of two bounded continuous functions is again bounded continuous. -/ +instance instSub : Sub (α →ᵇ R) where + sub f g := + { toFun := fun x ↦ (f x - g x), + map_bounded' := sub_bounded_of_bounded_of_bounded f.map_bounded' g.map_bounded' } + +theorem sub_apply {x : α} : (f - g) x = f x - g x := rfl +#align bounded_continuous_function.sub_apply BoundedContinuousFunction.sub_apply + +@[simp] +theorem coe_sub : ⇑(f - g) = f - g := rfl +#align bounded_continuous_function.coe_sub BoundedContinuousFunction.coe_sub + +end sub + section NormedAddCommGroup /- In this section, if `β` is a normed group, then we show that the space of bounded @@ -928,14 +950,6 @@ instance : Neg (α →ᵇ β) := ofNormedAddCommGroup (-f) f.continuous.neg ‖f‖ fun x => norm_neg ((⇑f) x) ▸ f.norm_coe_le_norm x⟩ -/-- The pointwise difference of two bounded continuous functions is again bounded continuous. -/ -instance instSub : Sub (α →ᵇ β) := - ⟨fun f g => - ofNormedAddCommGroup (f - g) (f.continuous.sub g.continuous) (‖f‖ + ‖g‖) fun x => by - simp only [sub_eq_add_neg] - exact le_trans (norm_add_le _ _) - (add_le_add (f.norm_coe_le_norm x) <| norm_neg ((⇑g) x) ▸ g.norm_coe_le_norm x)⟩ - @[simp] theorem coe_neg : ⇑(-f) = -f := rfl #align bounded_continuous_function.coe_neg BoundedContinuousFunction.coe_neg @@ -943,13 +957,6 @@ theorem coe_neg : ⇑(-f) = -f := rfl theorem neg_apply : (-f) x = -f x := rfl #align bounded_continuous_function.neg_apply BoundedContinuousFunction.neg_apply -@[simp] -theorem coe_sub : ⇑(f - g) = f - g := rfl -#align bounded_continuous_function.coe_sub BoundedContinuousFunction.coe_sub - -theorem sub_apply : (f - g) x = f x - g x := rfl -#align bounded_continuous_function.sub_apply BoundedContinuousFunction.sub_apply - @[simp] theorem mkOfCompact_neg [CompactSpace α] (f : C(α, β)) : mkOfCompact (-f) = -mkOfCompact f := rfl #align bounded_continuous_function.mk_of_compact_neg BoundedContinuousFunction.mkOfCompact_neg diff --git a/Mathlib/Topology/MetricSpace/PseudoMetric.lean b/Mathlib/Topology/MetricSpace/PseudoMetric.lean index d68b679507d24..3df6331ced896 100644 --- a/Mathlib/Topology/MetricSpace/PseudoMetric.lean +++ b/Mathlib/Topology/MetricSpace/PseudoMetric.lean @@ -1592,6 +1592,24 @@ theorem NNReal.le_add_nndist (a b : ℝ≥0) : a ≤ b + nndist a b := by exact le_of_abs_le (dist_eq a b).ge #align nnreal.le_add_nndist NNReal.le_add_nndist +lemma NNReal.ball_zero_eq_Ico' (c : ℝ≥0) : + Metric.ball (0 : ℝ≥0) c.toReal = Set.Ico 0 c := by ext x; simp + +lemma NNReal.ball_zero_eq_Ico (c : ℝ) : + Metric.ball (0 : ℝ≥0) c = Set.Ico 0 c.toNNReal := by + by_cases c_pos : 0 < c + · convert NNReal.ball_zero_eq_Ico' ⟨c, c_pos.le⟩ + simp [Real.toNNReal, c_pos.le] + simp [not_lt.mp c_pos] + +lemma NNReal.closedBall_zero_eq_Icc' (c : ℝ≥0) : + Metric.closedBall (0 : ℝ≥0) c.toReal = Set.Icc 0 c := by ext x; simp + +lemma NNReal.closedBall_zero_eq_Icc {c : ℝ} (c_nn : 0 ≤ c) : + Metric.closedBall (0 : ℝ≥0) c = Set.Icc 0 c.toNNReal := by + convert NNReal.closedBall_zero_eq_Icc' ⟨c, c_nn⟩ + simp [Real.toNNReal, c_nn] + end NNReal section ULift From 690d741a2041d5670fd10338b772280cca9b47cd Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 May 2024 12:14:53 +0000 Subject: [PATCH 51/52] chore: more deprecations (#12783) --- Mathlib/CategoryTheory/Comma/Basic.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/Discriminant.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index f4eba4336d152..d8940c2f3c405 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -239,7 +239,7 @@ instance full_map [F.Faithful] [F₁.Full] [F₂.Full] [IsIso α] [IsIso β] : ( assoc, assoc] erw [← α.naturality_assoc, β.naturality] dsimp - rw [F₁.image_preimage, F₂.image_preimage] + rw [F₁.map_preimage, F₂.map_preimage] simpa using φ.w) }, by aesop_cat⟩ instance essSurj_map [F₁.EssSurj] [F₂.EssSurj] [F.Full] [IsIso α] [IsIso β] : diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index fe2b8906b5426..9103e79d5a238 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -102,8 +102,8 @@ theorem discr_prime_pow_ne_two [IsCyclotomicExtension {p ^ (k + 1)} K L] [hp : F replace H := congr_arg (Algebra.norm K) H have hnorm : (norm K) (ζ ^ (p : ℕ) ^ k - 1) = (p : K) ^ (p : ℕ) ^ k := by by_cases hp : p = 2 - · exact mod_cast hζ.pow_sub_one_norm_prime_pow_of_ne_zero hirr le_rfl (hp2 hp) - · exact mod_cast hζ.pow_sub_one_norm_prime_ne_two hirr le_rfl hp + · exact mod_cast hζ.norm_pow_sub_one_eq_prime_pow_of_ne_zero hirr le_rfl (hp2 hp) + · exact mod_cast hζ.norm_pow_sub_one_of_prime_ne_two hirr le_rfl hp rw [MonoidHom.map_mul, hnorm, MonoidHom.map_mul, ← map_natCast (algebraMap K L), Algebra.norm_algebraMap, finrank L hirr] at H conv_rhs at H => -- Porting note: need to drill down to successfully rewrite the totient From fa33d89fc6f3890f5bf7188abae01b6baacab8d5 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 May 2024 16:03:52 +0000 Subject: [PATCH 52/52] feat: basic definition of comonoid objects (#10091) Co-authored-by: Scott Morrison Co-authored-by: Scott Morrison --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Monoidal/Comon_.lean | 171 ++++++++++++++++++++ 2 files changed, 172 insertions(+) create mode 100644 Mathlib/CategoryTheory/Monoidal/Comon_.lean diff --git a/Mathlib.lean b/Mathlib.lean index 298c2347f5bd7..b4f5569dc57c5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1442,6 +1442,7 @@ import Mathlib.CategoryTheory.Monoidal.Category import Mathlib.CategoryTheory.Monoidal.Center import Mathlib.CategoryTheory.Monoidal.CoherenceLemmas import Mathlib.CategoryTheory.Monoidal.CommMon_ +import Mathlib.CategoryTheory.Monoidal.Comon_ import Mathlib.CategoryTheory.Monoidal.Discrete import Mathlib.CategoryTheory.Monoidal.End import Mathlib.CategoryTheory.Monoidal.Free.Basic diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean new file mode 100644 index 0000000000000..85fd2f2198266 --- /dev/null +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2024 Lean FRO LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Mathlib.CategoryTheory.Monoidal.CoherenceLemmas +import Mathlib.CategoryTheory.Limits.Shapes.Terminal + +/-! +# The category of comonoids in a monoidal category. + +We define comonoids in a monoidal category `C`. + +## TODO +* `Comon_ C ≌ Mon_ (Cᵒᵖ)` +* An oplax monoidal functor takes comonoid objects to comonoid objects. + That is, a oplax monoidal functor `F : C ⥤ D` induces a functor `Comon_ C ⥤ Comon_ D`. +* Comonoid objects in `C` are "just" + oplax monoidal functors from the trivial monoidal category to `C`. +* The category of comonoids in a braided monoidal category is monoidal. + (It may suffice to transfer this across the equivalent to monoids in the opposite category.) + +-/ + +universe v₁ v₂ u₁ u₂ u + +open CategoryTheory MonoidalCategory + +variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] + +/-- A comonoid object internal to a monoidal category. + +When the monoidal category is preadditive, this is also sometimes called a "coalgebra object". +-/ +structure Comon_ where + /-- The underlying object of a comonoid object. -/ + X : C + /-- The counit of a comonoid object. -/ + counit : X ⟶ 𝟙_ C + /-- The comultiplication morphism of a comonoid object. -/ + comul : X ⟶ X ⊗ X + counit_comul : comul ≫ (counit ▷ X) = (λ_ X).inv := by aesop_cat + comul_counit : comul ≫ (X ◁ counit) = (ρ_ X).inv := by aesop_cat + comul_assoc : comul ≫ (X ◁ comul) ≫ (α_ X X X).inv = comul ≫ (comul ▷ X) := by aesop_cat + +attribute [reassoc] Comon_.counit_comul Comon_.comul_counit + +attribute [reassoc (attr := simp)] Comon_.comul_assoc + +namespace Comon_ + +/-- The trivial comonoid object. We later show this is terminal in `Comon_ C`. +-/ +@[simps] +def trivial : Comon_ C where + X := 𝟙_ C + counit := 𝟙 _ + comul := (λ_ _).inv + comul_assoc := by coherence + counit_comul := by coherence + comul_counit := by coherence + +instance : Inhabited (Comon_ C) := + ⟨trivial C⟩ + +variable {C} +variable {M : Comon_ C} + +@[reassoc (attr := simp)] +theorem counit_comul_hom {Z : C} (f : M.X ⟶ Z) : M.comul ≫ (M.counit ⊗ f) = f ≫ (λ_ Z).inv := by + rw [leftUnitor_inv_naturality, tensorHom_def, counit_comul_assoc] + +@[reassoc (attr := simp)] +theorem comul_counit_hom {Z : C} (f : M.X ⟶ Z) : M.comul ≫ (f ⊗ M.counit) = f ≫ (ρ_ Z).inv := by + rw [rightUnitor_inv_naturality, tensorHom_def', comul_counit_assoc] + +@[reassoc (attr := simp)] theorem comul_assoc_flip : + M.comul ≫ (M.comul ▷ M.X) ≫ (α_ M.X M.X M.X).hom = M.comul ≫ (M.X ◁ M.comul) := by + simp [← comul_assoc_assoc] + +/-- A morphism of comonoid objects. -/ +@[ext] +structure Hom (M N : Comon_ C) where + /-- The underlying morphism of a morphism of comonoid objects. -/ + hom : M.X ⟶ N.X + hom_counit : hom ≫ N.counit = M.counit := by aesop_cat + hom_comul : hom ≫ N.comul = M.comul ≫ (hom ⊗ hom) := by aesop_cat + +attribute [reassoc (attr := simp)] Hom.hom_counit Hom.hom_comul + +/-- The identity morphism on a comonoid object. -/ +@[simps] +def id (M : Comon_ C) : Hom M M where + hom := 𝟙 M.X + +instance homInhabited (M : Comon_ C) : Inhabited (Hom M M) := + ⟨id M⟩ + +/-- Composition of morphisms of monoid objects. -/ +@[simps] +def comp {M N O : Comon_ C} (f : Hom M N) (g : Hom N O) : Hom M O where + hom := f.hom ≫ g.hom + +instance : Category (Comon_ C) where + Hom M N := Hom M N + id := id + comp f g := comp f g + +@[ext] lemma ext {X Y : Comon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext _ _ w + +@[simp] theorem id_hom' (M : Comon_ C) : (𝟙 M : Hom M M).hom = 𝟙 M.X := rfl + +@[simp] +theorem comp_hom' {M N K : Comon_ C} (f : M ⟶ N) (g : N ⟶ K) : (f ≫ g).hom = f.hom ≫ g.hom := + rfl + +section + +variable (C) + +/-- The forgetful functor from comonoid objects to the ambient category. -/ +@[simps] +def forget : Comon_ C ⥤ C where + obj A := A.X + map f := f.hom + +end + +instance forget_faithful : (@forget C _ _).Faithful where + +instance {A B : Comon_ C} (f : A ⟶ B) [e : IsIso ((forget C).map f)] : IsIso f.hom := e + +/-- The forgetful functor from comonoid objects to the ambient category reflects isomorphisms. -/ +instance : (forget C).ReflectsIsomorphisms where + reflects f e := + ⟨⟨{ hom := inv f.hom }, by aesop_cat⟩⟩ + +/-- Construct an isomorphism of monoids by giving an isomorphism between the underlying objects +and checking compatibility with unit and multiplication only in the forward direction. +-/ +@[simps] +def mkIso {M N : Comon_ C} (f : M.X ≅ N.X) (f_counit : f.hom ≫ N.counit = M.counit) + (f_comul : f.hom ≫ N.comul = M.comul ≫ (f.hom ⊗ f.hom)) : M ≅ N where + hom := + { hom := f.hom + hom_counit := f_counit + hom_comul := f_comul } + inv := + { hom := f.inv + hom_counit := by rw [← f_counit]; simp + hom_comul := by + rw [← cancel_epi f.hom] + slice_rhs 1 2 => rw [f_comul] + simp } + +instance uniqueHomToTrivial (A : Comon_ C) : Unique (A ⟶ trivial C) where + default := + { hom := A.counit + hom_counit := by dsimp; simp + hom_comul := by dsimp; simp [A.comul_counit, unitors_inv_equal] } + uniq f := by + ext; simp + rw [← Category.comp_id f.hom] + erw [f.hom_counit] + +open CategoryTheory.Limits + +instance : HasTerminal (Comon_ C) := + hasTerminal_of_unique (trivial C) + +end Comon_