From a1578ecc8f1ed2f7bef68f3ff7c321bfc0c49820 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Fri, 25 Oct 2024 00:50:18 +0000 Subject: [PATCH 01/29] chore: use `Batteries` test driver directly in the lake file (#15897) With leanprover/lean4#4261, we can specify a test driver from a dependency. We do this with `Batteries` test driver. Previously, the test driver from `Batteries` was called via `scripts/test.lean` which is now removed. --- lakefile.lean | 14 ++------------ scripts/test.lean | 21 --------------------- 2 files changed, 2 insertions(+), 33 deletions(-) delete mode 100644 scripts/test.lean diff --git a/lakefile.lean b/lakefile.lean index 697f41bab4e5b..25b17ebe85a33 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -52,6 +52,8 @@ package mathlib where leanOptions := mathlibLeanOptions -- Mathlib also enforces these linter options, which are not active by default. moreServerOptions := mathlibOnlyLinters + -- Use Batteries' test driver for `lake test` + testDriver := "batteries/test" -- These are additional settings which do not affect the lake hash, -- so they can be enabled in CI and disabled locally or vice versa. -- Warning: Do not put any options here that actually change the olean files, @@ -132,18 +134,6 @@ lean_exe pole where -- Executables which import `Lake` must set `-lLake`. weakLinkArgs := #["-lLake"] -/-- -`lake exe test` is a thin wrapper around `lake exe batteries/test`, until -https://github.com/leanprover/lean4/issues/4121 is resolved. - -You can also use it as e.g. `lake exe test conv eval_elab` to only run the named tests. --/ -@[test_driver] -lean_exe test where - -- We could add the above `leanOptions` and `moreServerOptions`: currently, these do not take - -- effect as `test` is a `lean_exe`. With a `lean_lib`, it would work... - srcDir := "scripts" - /-! ## Other configuration -/ diff --git a/scripts/test.lean b/scripts/test.lean deleted file mode 100644 index f02a621637377..0000000000000 --- a/scripts/test.lean +++ /dev/null @@ -1,21 +0,0 @@ -/- -Copyright (c) 2024 Kim Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kim Morrison --/ -open IO.Process -open System - -/-- -Run tests, via the Batteries test runner. - -When https://github.com/leanprover/lean4/issues/4121 -"allow using an upstream executable as a lake `@[test_runner]`" -is resolved, this file can be replaced with a line in `lakefile.lean`. --/ -def main (args : List String) : IO Unit := do - -- ProofWidgets and Batteries may not have been completely built by `lake build` yet, - -- but they are needed by some tests. - _ ← (← spawn { cmd := "lake", args := #["build", "ProofWidgets", "Batteries"] }).wait - let exitcode ← (← spawn { cmd := "lake", args := #["exe", "batteries/test"] ++ args }).wait - exit exitcode.toUInt8 From 87ee1f25bf4ee48b5290edf466511e6e76d7286a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 01:40:44 +0000 Subject: [PATCH 02/29] chore: cleanup of `mergeSort'` (#16902) This PR removes the old `mergeSort'` in Mathlib (which was *not* a stable sort algorithm), replacing all uses with the stable `mergeSort` from Lean, mostly preserving the API. --- Mathlib/Data/List/Sort.lean | 239 +++++++------------------------- Mathlib/Data/Multiset/Sort.lean | 30 ++-- Mathlib/Logic/Equiv/List.lean | 4 +- Mathlib/Order/Defs.lean | 32 +++++ 4 files changed, 102 insertions(+), 203 deletions(-) diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 513e81df6f0a3..6733b9df5e8a3 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -3,39 +3,20 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ +import Batteries.Data.List.Pairwise import Mathlib.Data.List.OfFn import Mathlib.Data.List.Nodup import Mathlib.Order.Fin.Basic -import Batteries.Data.List.Perm /-! # Sorting algorithms on lists In this file we define `List.Sorted r l` to be an alias for `List.Pairwise r l`. This alias is preferred in the case that `r` is a `<` or `≤`-like relation. -Then we define two sorting algorithms: -`List.insertionSort` and `List.mergeSort'`, and prove their correctness. +Then we define the sorting algorithm +`List.insertionSort` and prove its correctness. -/ -#adaptation_note -/-- -`List.mergeSort` has now been implemented in Lean4. -It improves on the one here by being a "stable" sort -(in the sense that a sorted sublist of the original list remains a sublist of the result), -and is also marginally faster. - -However we haven't yet replaced `List.mergeSort'` here. -The obstacle is that `mergeSort'` is written using `r : α → α → Prop` with `[DecidableRel r]`, -while `mergeSort` uses `r : α → α → Bool`. This is hardly insurmountable, -but it's a bit of work that hasn't been done yet. - -`List.mergeSort'` is only used in Mathlib to sort multisets for printing, so this is not critical. - -A pull request cleaning up here, and ideally deprecating or deleting `List.mergeSort'`, -would be welcome. --/ - - open List.Perm universe u v @@ -165,6 +146,14 @@ theorem Sorted.rel_of_mem_take_of_mem_drop {l : List α} (h : List.Sorted r l) { rw [length_take] at hix exact h.rel_get_of_lt (Nat.lt_add_right _ (Nat.lt_min.mp hix).left) +/-- +If a list is sorted with respect to a decidable relation, +then it is sorted with respect to the corresponding Bool-valued relation. +-/ +theorem Sorted.decide [DecidableRel r] (l : List α) (h : Sorted r l) : + Sorted (fun a b => decide (r a b) = true) l := by + refine h.imp fun {a b} h => by simpa using h + end Sorted section Monotone @@ -220,6 +209,11 @@ def insertionSort : List α → List α | [] => [] | b :: l => orderedInsert r b (insertionSort l) +-- A quick check that insertionSort is stable: +example : + insertionSort (fun m n => m / 10 ≤ n / 10) [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] = + [5, 7, 2, 17, 12, 27, 23, 43, 95, 98, 221, 567] := rfl + @[simp] theorem orderedInsert_nil (a : α) : [].orderedInsert r a = [a] := rfl @@ -492,192 +486,61 @@ end Correctness end InsertionSort -/-! ### Merge sort -/ +/-! ### Merge sort +We provide some wrapper functions around the theorems for `mergeSort` provided in Lean, +which rather than using explicit hypotheses for transitivity and totality, +use Mathlib order typeclasses instead. +-/ -section MergeSort - --- TODO(Jeremy): observation: if instead we write (a :: (split l).1, b :: (split l).2), the --- equation compiler can't prove the third equation -/-- Split `l` into two lists of approximately equal length. - - split [1, 2, 3, 4, 5] = ([1, 3, 5], [2, 4]) -/ -@[simp] -def split : List α → List α × List α - | [] => ([], []) - | a :: l => - let (l₁, l₂) := split l - (a :: l₂, l₁) - -theorem split_cons_of_eq (a : α) {l l₁ l₂ : List α} (h : split l = (l₁, l₂)) : - split (a :: l) = (a :: l₂, l₁) := by rw [split, h] +unseal merge mergeSort in +example : + mergeSort [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] (fun m n => m / 10 ≤ n / 10) = + [5, 7, 2, 17, 12, 27, 23, 43, 95, 98, 221, 567] := rfl -@[simp] -theorem map_split (f : α → β) : - ∀ l : List α, (map f l).split = (l.split.1.map f, l.split.2.map f) - | [] => rfl - | a :: l => by simp [map_split] - -@[simp] -theorem mem_split_iff {x : α} : ∀ {l : List α}, x ∈ l.split.1 ∨ x ∈ l.split.2 ↔ x ∈ l - | [] => by simp - | a :: l => by simp_rw [split, mem_cons, or_assoc, or_comm, mem_split_iff] - -theorem length_split_le : - ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → length l₁ ≤ length l ∧ length l₂ ≤ length l - | [], _, _, rfl => ⟨Nat.le_refl 0, Nat.le_refl 0⟩ - | a :: l, l₁', l₂', h => by - cases' e : split l with l₁ l₂ - injection (split_cons_of_eq _ e).symm.trans h; substs l₁' l₂' - cases' length_split_le e with h₁ h₂ - exact ⟨Nat.succ_le_succ h₂, Nat.le_succ_of_le h₁⟩ - -theorem length_split_fst_le (l : List α) : length (split l).1 ≤ length l := - (length_split_le rfl).1 - -theorem length_split_snd_le (l : List α) : length (split l).2 ≤ length l := - (length_split_le rfl).2 - -theorem length_split_lt {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - length l₁ < length (a :: b :: l) ∧ length l₂ < length (a :: b :: l) := by - cases' e : split l with l₁' l₂' - injection (split_cons_of_eq _ (split_cons_of_eq _ e)).symm.trans h; substs l₁ l₂ - cases' length_split_le e with h₁ h₂ - exact ⟨Nat.succ_le_succ (Nat.succ_le_succ h₁), Nat.succ_le_succ (Nat.succ_le_succ h₂)⟩ - -theorem perm_split : ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → l ~ l₁ ++ l₂ - | [], _, _, rfl => Perm.refl _ - | a :: l, l₁', l₂', h => by - cases' e : split l with l₁ l₂ - injection (split_cons_of_eq _ e).symm.trans h; substs l₁' l₂' - exact ((perm_split e).trans perm_append_comm).cons a - -/-- Implementation of a merge sort algorithm to sort a list. -/ -def mergeSort' : List α → List α - | [] => [] - | [a] => [a] - | a :: b :: l => by - -- Porting note: rewrote to make `mergeSort_cons_cons` proof easier - let ls := (split (a :: b :: l)) - have := length_split_fst_le l - have := length_split_snd_le l - exact merge (mergeSort' ls.1) (mergeSort' ls.2) (r · ·) - termination_by l => length l - -@[nolint unusedHavesSuffices] -- Porting note: false positive -theorem mergeSort'_cons_cons {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - mergeSort' r (a :: b :: l) = merge (mergeSort' r l₁) (mergeSort' r l₂) (r · ·) := by - simp only [mergeSort', h] +section MergeSort section Correctness -theorem perm_mergeSort' : ∀ l : List α, mergeSort' r l ~ l - | [] => by simp [mergeSort'] - | [a] => by simp [mergeSort'] - | a :: b :: l => by - cases' e : split (a :: b :: l) with l₁ l₂ - cases' length_split_lt e with h₁ h₂ - rw [mergeSort'_cons_cons r e] - apply (perm_merge (r · ·) _ _).trans - exact - ((perm_mergeSort' l₁).append (perm_mergeSort' l₂)).trans (perm_split e).symm - termination_by l => length l - -@[simp] -theorem mem_mergeSort' {l : List α} {x : α} : x ∈ l.mergeSort' r ↔ x ∈ l := - (perm_mergeSort' r l).mem_iff - -@[simp] -theorem length_mergeSort' (l : List α) : (mergeSort' r l).length = l.length := - (perm_mergeSort' r _).length_eq - section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge l l' (r · ·) ) - | [], [], _, _ => by simp - | [], b :: l', _, h₂ => by simpa using h₂ - | a :: l, [], h₁, _ => by simpa using h₁ - | a :: l, b :: l', h₁, h₂ => by - by_cases h : a ≼ b - · suffices ∀ b' ∈ List.merge l (b :: l') (r · ·) , r a b' by - simpa [h, h₁.of_cons.merge h₂] - intro b' bm - rcases show b' = b ∨ b' ∈ l ∨ b' ∈ l' by - simpa [or_left_comm] using (perm_merge _ _ _).subset bm with - (be | bl | bl') - · subst b' - assumption - · exact rel_of_sorted_cons h₁ _ bl - · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge (a :: l) l' (r · ·) , r b b' by - simpa [h, h₁.merge h₂.of_cons] - intro b' bm - have ba : b ≼ a := (total_of r _ _).resolve_left h - have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge _ _ _).subset bm - rcases this with (be | bl | bl') - · subst b' - assumption - · exact _root_.trans ba (rel_of_sorted_cons h₁ _ bl) - · exact rel_of_sorted_cons h₂ _ bl' +theorem Sorted.merge {l l' : List α} (h : Sorted r l) (h' : Sorted r l') : + Sorted r (merge l l' (r · ·)) := by + simpa using sorted_merge (le := (r · ·)) + (fun a b c h₁ h₂ => by simpa using _root_.trans (by simpa using h₁) (by simpa using h₂)) + (fun a b => by simpa using IsTotal.total a b) + l l' (by simpa using h) (by simpa using h') variable (r) -theorem sorted_mergeSort' : ∀ l : List α, Sorted r (mergeSort' r l) - | [] => by simp [mergeSort'] - | [a] => by simp [mergeSort'] - | a :: b :: l => by - cases' e : split (a :: b :: l) with l₁ l₂ - cases' length_split_lt e with h₁ h₂ - rw [mergeSort'_cons_cons r e] - exact (sorted_mergeSort' l₁).merge (sorted_mergeSort' l₂) - termination_by l => length l - -theorem mergeSort'_eq_self [IsAntisymm α r] {l : List α} : Sorted r l → mergeSort' r l = l := - eq_of_perm_of_sorted (perm_mergeSort' _ _) (sorted_mergeSort' _ _) - -theorem mergeSort'_eq_insertionSort [IsAntisymm α r] (l : List α) : - mergeSort' r l = insertionSort r l := - eq_of_perm_of_sorted ((perm_mergeSort' r l).trans (perm_insertionSort r l).symm) - (sorted_mergeSort' r l) (sorted_insertionSort r l) +/-- Variant of `sorted_mergeSort` using order typeclasses. -/ +theorem sorted_mergeSort' [Preorder α] [DecidableRel ((· : α) ≤ ·)] [IsTotal α (· ≤ ·)] + (l : List α) : Sorted (· ≤ ·) (mergeSort l) := by + simpa using sorted_mergeSort (le := fun a b => a ≤ b) + (fun a b c h₁ h₂ => by simpa using le_trans (by simpa using h₁) (by simpa using h₂)) + (fun a b => by simpa using IsTotal.total a b) + l + +theorem mergeSort_eq_self [LinearOrder α] {l : List α} : Sorted (· ≤ ·) l → mergeSort l = l := + eq_of_perm_of_sorted (mergeSort_perm _ _) (sorted_mergeSort' l) + +theorem mergeSort_eq_insertionSort [IsAntisymm α r] (l : List α) : + mergeSort l (r · ·) = insertionSort r l := + eq_of_perm_of_sorted ((mergeSort_perm l _).trans (perm_insertionSort r l).symm) + (sorted_mergeSort (le := (r · ·)) + (fun a b c h₁ h₂ => by simpa using _root_.trans (by simpa using h₁) (by simpa using h₂)) + (fun a b => by simpa using IsTotal.total a b) + l) + (sorted_insertionSort r l).decide end TotalAndTransitive end Correctness -@[simp] -theorem mergeSort'_nil : [].mergeSort' r = [] := by rw [List.mergeSort'] - -@[simp] -theorem mergeSort'_singleton (a : α) : [a].mergeSort' r = [a] := by rw [List.mergeSort'] - -theorem map_mergeSort' (f : α → β) (l : List α) (hl : ∀ a ∈ l, ∀ b ∈ l, a ≼ b ↔ f a ≼ f b) : - (l.mergeSort' r).map f = (l.map f).mergeSort' s := - match l with - | [] => by simp - | [x] => by simp - | a :: b :: l => by - simp_rw [← mem_split_iff (l := a :: b :: l), or_imp, forall_and] at hl - set l₁ := (split (a :: b :: l)).1 - set l₂ := (split (a :: b :: l)).2 - have e : split (a :: b :: l) = (l₁, l₂) := rfl - have fe : split (f a :: f b :: l.map f) = (l₁.map f, l₂.map f) := by - rw [← map, ← map, map_split, e] - have := length_split_fst_le l - have := length_split_snd_le l - simp_rw [List.map] - rw [List.mergeSort'_cons_cons _ e, List.mergeSort'_cons_cons _ fe, - map_merge, map_mergeSort' _ l₁ hl.1.1, map_mergeSort' _ l₂ hl.2.2] - simp_rw [mem_mergeSort', decide_eq_decide] - exact hl.1.2 - termination_by length l - end MergeSort end sort --- try them out! ---#eval insertionSort (fun m n : ℕ => m ≤ n) [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] ---#eval mergeSort' (fun m n : ℕ => m ≤ n) [5, 27, 221, 95, 17, 43, 7, 2, 98, 567, 23, 12] end List diff --git a/Mathlib/Data/Multiset/Sort.lean b/Mathlib/Data/Multiset/Sort.lean index 997af93b548c0..7c329e6f85d53 100644 --- a/Mathlib/Data/Multiset/Sort.lean +++ b/Mathlib/Data/Multiset/Sort.lean @@ -10,13 +10,12 @@ import Mathlib.Data.Multiset.Basic # Construct a sorted list from a multiset. -/ +variable {α β : Type*} namespace Multiset open List -variable {α β : Type*} - section sort variable (r : α → α → Prop) [DecidableRel r] [IsTrans α r] [IsAntisymm α r] [IsTotal α r] @@ -25,47 +24,52 @@ variable (r' : β → β → Prop) [DecidableRel r'] [IsTrans β r'] [IsAntisymm /-- `sort s` constructs a sorted list from the multiset `s`. (Uses merge sort algorithm.) -/ def sort (s : Multiset α) : List α := - Quot.liftOn s (mergeSort' r) fun _ _ h => - eq_of_perm_of_sorted ((perm_mergeSort' _ _).trans <| h.trans (perm_mergeSort' _ _).symm) - (sorted_mergeSort' r _) (sorted_mergeSort' r _) + Quot.liftOn s (mergeSort · (r · ·)) fun _ _ h => + eq_of_perm_of_sorted ((mergeSort_perm _ _).trans <| h.trans (mergeSort_perm _ _).symm) + (sorted_mergeSort IsTrans.trans + (fun a b => by simpa using IsTotal.total a b) _) + (sorted_mergeSort IsTrans.trans + (fun a b => by simpa using IsTotal.total a b) _) @[simp] -theorem coe_sort (l : List α) : sort r l = mergeSort' r l := +theorem coe_sort (l : List α) : sort r l = mergeSort l (r · ·) := rfl @[simp] theorem sort_sorted (s : Multiset α) : Sorted r (sort r s) := - Quot.inductionOn s fun _l => sorted_mergeSort' r _ + Quot.inductionOn s fun l => by + simpa using sorted_mergeSort (le := (r · ·)) IsTrans.trans + (fun a b => by simpa using IsTotal.total a b) l @[simp] theorem sort_eq (s : Multiset α) : ↑(sort r s) = s := - Quot.inductionOn s fun _ => Quot.sound <| perm_mergeSort' _ _ + Quot.inductionOn s fun _ => Quot.sound <| mergeSort_perm _ _ @[simp] theorem mem_sort {s : Multiset α} {a : α} : a ∈ sort r s ↔ a ∈ s := by rw [← mem_coe, sort_eq] @[simp] theorem length_sort {s : Multiset α} : (sort r s).length = card s := - Quot.inductionOn s <| length_mergeSort' _ + Quot.inductionOn s <| length_mergeSort @[simp] theorem sort_zero : sort r 0 = [] := - List.mergeSort'_nil r + List.mergeSort_nil @[simp] theorem sort_singleton (a : α) : sort r {a} = [a] := - List.mergeSort'_singleton r a + List.mergeSort_singleton a theorem map_sort (f : α → β) (s : Multiset α) (hs : ∀ a ∈ s, ∀ b ∈ s, r a b ↔ r' (f a) (f b)) : (s.sort r).map f = (s.map f).sort r' := by revert s - exact Quot.ind fun _ => List.map_mergeSort' _ _ _ _ + exact Quot.ind fun l h => map_mergeSort (l := l) (by simpa using h) theorem sort_cons (a : α) (s : Multiset α) : (∀ b ∈ s, r a b) → sort r (a ::ₘ s) = a :: sort r s := by refine Quot.inductionOn s fun l => ?_ - simpa [mergeSort'_eq_insertionSort] using insertionSort_cons r + simpa [mergeSort_eq_insertionSort] using insertionSort_cons r (a := a) (l := l) end sort diff --git a/Mathlib/Logic/Equiv/List.lean b/Mathlib/Logic/Equiv/List.lean index bf14924bbdaac..c0753898bf0dc 100644 --- a/Mathlib/Logic/Equiv/List.lean +++ b/Mathlib/Logic/Equiv/List.lean @@ -289,7 +289,7 @@ instance multiset : Denumerable (Multiset α) := raise_lower (List.sorted_cons.2 ⟨fun n _ => Nat.zero_le n, (s.map encode).sort_sorted _⟩) simp [-Multiset.map_coe, this], fun n => by - simp [-Multiset.map_coe, List.mergeSort'_eq_self _ (raise_sorted _ _), lower_raise]⟩ + simp [-Multiset.map_coe, List.mergeSort_eq_self (raise_sorted _ _), lower_raise]⟩ end Multiset @@ -344,7 +344,7 @@ instance finset : Denumerable (Finset α) := raise_lower' (fun n _ => Nat.zero_le n) (Finset.sort_sorted_lt _)], fun n => by simp [-Multiset.map_coe, Finset.map, raise'Finset, Finset.sort, - List.mergeSort'_eq_self (· ≤ ·) ((raise'_sorted _ _).imp (@le_of_lt _ _)), lower_raise']⟩ + List.mergeSort_eq_self ((raise'_sorted _ _).imp (@le_of_lt _ _)), lower_raise']⟩ end Finset diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 75d603ee2c080..339a1a55bc927 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -125,6 +125,38 @@ instance (priority := 90) isAsymm_of_isTrans_of_isIrrefl [IsTrans α r] [IsIrref IsAsymm α r := ⟨fun a _b h₁ h₂ => absurd (_root_.trans h₁ h₂) (irrefl a)⟩ +instance IsIrrefl.decide [DecidableRel r] [IsIrrefl α r] : + IsIrrefl α (fun a b => decide (r a b) = true) where + irrefl := fun a => by simpa using irrefl a + +instance IsRefl.decide [DecidableRel r] [IsRefl α r] : + IsRefl α (fun a b => decide (r a b) = true) where + refl := fun a => by simpa using refl a + +instance IsTrans.decide [DecidableRel r] [IsTrans α r] : + IsTrans α (fun a b => decide (r a b) = true) where + trans := fun a b c => by simpa using trans a b c + +instance IsSymm.decide [DecidableRel r] [IsSymm α r] : + IsSymm α (fun a b => decide (r a b) = true) where + symm := fun a b => by simpa using symm a b + +instance IsAntisymm.decide [DecidableRel r] [IsAntisymm α r] : + IsAntisymm α (fun a b => decide (r a b) = true) where + antisymm := fun a b h₁ h₂ => antisymm _ _ (by simpa using h₁) (by simpa using h₂) + +instance IsAsymm.decide [DecidableRel r] [IsAsymm α r] : + IsAsymm α (fun a b => decide (r a b) = true) where + asymm := fun a b => by simpa using asymm a b + +instance IsTotal.decide [DecidableRel r] [IsTotal α r] : + IsTotal α (fun a b => decide (r a b) = true) where + total := fun a b => by simpa using total a b + +instance IsTrichotomous.decide [DecidableRel r] [IsTrichotomous α r] : + IsTrichotomous α (fun a b => decide (r a b) = true) where + trichotomous := fun a b => by simpa using trichotomous a b + variable (r) @[elab_without_expected_type] lemma irrefl_of [IsIrrefl α r] (a : α) : ¬a ≺ a := irrefl a From fc157081b3adaa6877bf68c0415d54c0126cf47f Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 05:50:35 +0000 Subject: [PATCH 03/29] chore(Algebra/GroupWithZero/Action): reduce theory included with `Group.Action.Defs` (#18190) The `GroupWithZero/Action/Defs.lean` file depends on a bit of theory and in turn defines some results not needed for e.g. modules. This PR tries to ensure that the focus is on the definitions of `SMulZeroClass` and `DistribMulAction`. In the process, I made the following moves: * `Action/End.lean` now contains results involving `Function.End` and some results relating homomorphisms with multiplication * `Action/Faithful.lean` now contains results involving `FaithfulAction` --- Mathlib.lean | 2 + .../Group/Submonoid/DistribMulAction.lean | 2 +- .../Algebra/GroupWithZero/Action/Defs.lean | 116 +----------------- Mathlib/Algebra/GroupWithZero/Action/End.lean | 93 ++++++++++++++ .../GroupWithZero/Action/Faithful.lean | 23 ++++ .../Algebra/GroupWithZero/Action/Prod.lean | 2 +- .../Algebra/GroupWithZero/Action/Units.lean | 35 +++++- Mathlib/Algebra/Module/Defs.lean | 2 + Mathlib/Algebra/Ring/Action/Basic.lean | 2 +- Mathlib/Data/NNRat/Lemmas.lean | 1 + .../GroupTheory/GroupAction/DomAct/Basic.lean | 1 + 11 files changed, 164 insertions(+), 115 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Action/End.lean create mode 100644 Mathlib/Algebra/GroupWithZero/Action/Faithful.lean diff --git a/Mathlib.lean b/Mathlib.lean index e839e2377857a..23c1f16aa849d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -319,6 +319,8 @@ import Mathlib.Algebra.Group.ZeroOne import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End +import Mathlib.Algebra.GroupWithZero.Action.Faithful import Mathlib.Algebra.GroupWithZero.Action.Opposite import Mathlib.Algebra.GroupWithZero.Action.Pi import Mathlib.Algebra.GroupWithZero.Action.Prod diff --git a/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean b/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean index d11f9c6414e1f..15e285f0d6b07 100644 --- a/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean +++ b/Mathlib/Algebra/Group/Submonoid/DistribMulAction.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Group.Submonoid.Operations -import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End /-! # Distributive actions by submonoids diff --git a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean index 0d5f09f7f565a..77eac698ee0d5 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean @@ -3,9 +3,8 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ -import Mathlib.Algebra.Group.Action.Units -import Mathlib.Algebra.Group.Equiv.Basic -import Mathlib.Algebra.GroupWithZero.Units.Basic +import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Algebra.Group.Hom.Defs /-! # Definitions of group actions @@ -13,25 +12,17 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic This file defines a hierarchy of group action type-classes on top of the previously defined notation classes `SMul` and its additive version `VAdd`: -* `MulAction M α` and its additive version `AddAction G P` are typeclasses used for - actions of multiplicative and additive monoids and groups; they extend notation classes - `SMul` and `VAdd` that are defined in `Algebra.Group.Defs`; +* `SMulZeroClass` is a typeclass for an action that preserves zero +* `DistribSMul M A` is a typeclass for an action on an additive monoid (`AddZeroClass`) that + preserves addition and zero * `DistribMulAction M A` is a typeclass for an action of a multiplicative monoid on an additive monoid such that `a • (b + c) = a • b + a • c` and `a • 0 = 0`. The hierarchy is extended further by `Module`, defined elsewhere. -Also provided are typeclasses for faithful and transitive actions, and typeclasses regarding the -interaction of different group actions, - -* `SMulCommClass M N α` and its additive version `VAddCommClass M N α`; -* `IsScalarTower M N α` and its additive version `VAddAssocClass M N α`; -* `IsCentralScalar M α` and its additive version `IsCentralVAdd M N α`. - ## Notation - `a • b` is used as notation for `SMul.smul a b`. -- `a +ᵥ b` is used as notation for `VAdd.vadd a b`. ## Implementation details @@ -51,42 +42,6 @@ open Function variable {M N A B α β : Type*} -/-- `Monoid.toMulAction` is faithful on nontrivial cancellative monoids with zero. -/ -instance CancelMonoidWithZero.faithfulSMul [CancelMonoidWithZero α] [Nontrivial α] : - FaithfulSMul α α where eq_of_smul_eq_smul h := mul_left_injective₀ one_ne_zero (h 1) - -section GroupWithZero -variable [GroupWithZero α] [MulAction α β] {a : α} - -@[simp] lemma inv_smul_smul₀ (ha : a ≠ 0) (x : β) : a⁻¹ • a • x = x := - inv_smul_smul (Units.mk0 a ha) x - -@[simp] -lemma smul_inv_smul₀ (ha : a ≠ 0) (x : β) : a • a⁻¹ • x = x := smul_inv_smul (Units.mk0 a ha) x - -lemma inv_smul_eq_iff₀ (ha : a ≠ 0) {x y : β} : a⁻¹ • x = y ↔ x = a • y := - inv_smul_eq_iff (g := Units.mk0 a ha) - -lemma eq_inv_smul_iff₀ (ha : a ≠ 0) {x y : β} : x = a⁻¹ • y ↔ a • x = y := - eq_inv_smul_iff (g := Units.mk0 a ha) - -@[simp] -lemma Commute.smul_right_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} - (ha : a ≠ 0) : Commute x (a • y) ↔ Commute x y := Commute.smul_right_iff (g := Units.mk0 a ha) - -@[simp] -lemma Commute.smul_left_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} - (ha : a ≠ 0) : Commute (a • x) y ↔ Commute x y := Commute.smul_left_iff (g := Units.mk0 a ha) - -/-- Right scalar multiplication as an order isomorphism. -/ -@[simps] def Equiv.smulRight (ha : a ≠ 0) : β ≃ β where - toFun b := a • b - invFun b := a⁻¹ • b - left_inv := inv_smul_smul₀ ha - right_inv := smul_inv_smul₀ ha - -end GroupWithZero - /-- Typeclass for scalar multiplication that preserves `0` on the right. -/ class SMulZeroClass (M A : Type*) [Zero A] extends SMul M A where /-- Multiplying `0` by a scalar gives `0` -/ @@ -257,22 +212,8 @@ protected abbrev Function.Surjective.distribMulAction [AddMonoid B] [SMul M B] ( (hf : Surjective f) (smul : ∀ (c : M) (x), f (c • x) = c • f x) : DistribMulAction M B := { hf.distribSMul f smul, hf.mulAction f smul with } -/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. - -See also `Function.Surjective.mulActionLeft` and `Function.Surjective.moduleLeft`. --/ -abbrev Function.Surjective.distribMulActionLeft {R S M : Type*} [Monoid R] [AddMonoid M] - [DistribMulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Function.Surjective f) - (hsmul : ∀ (c) (x : M), f c • x = c • x) : DistribMulAction S M := - { hf.distribSMulLeft f hsmul, hf.mulActionLeft f hsmul with } - variable (A) -/-- Compose a `DistribMulAction` with a `MonoidHom`, with action `f r' • m`. -See note [reducible non-instances]. -/ -abbrev DistribMulAction.compHom [Monoid N] (f : N →* M) : DistribMulAction N A := - { DistribSMul.compFun A f, MulAction.compHom A f with } - /-- Each element of the monoid defines an additive monoid homomorphism. -/ @[simps!] def DistribMulAction.toAddMonoidHom (x : M) : A →+ A := @@ -360,13 +301,6 @@ protected abbrev Function.Surjective.mulDistribMulAction [Monoid B] [SMul M B] ( variable (A) -/-- Compose a `MulDistribMulAction` with a `MonoidHom`, with action `f r' • m`. -See note [reducible non-instances]. -/ -abbrev MulDistribMulAction.compHom [Monoid N] (f : N →* M) : MulDistribMulAction N A := - { MulAction.compHom A f with - smul_one := fun x => smul_one (f x), - smul_mul := fun x => smul_mul' (f x) } - /-- Scalar multiplication by `r` as a `MonoidHom`. -/ def MulDistribMulAction.toMonoidHom (r : M) : A →* A where @@ -384,16 +318,6 @@ theorem MulDistribMulAction.toMonoidHom_apply (r : M) (x : A) : @[simp] lemma smul_pow' (r : M) (x : A) (n : ℕ) : r • x ^ n = (r • x) ^ n := (MulDistribMulAction.toMonoidHom _ _).map_pow _ _ -variable (M A) - -/-- Each element of the monoid defines a monoid homomorphism. -/ -@[simps] -def MulDistribMulAction.toMonoidEnd : - M →* Monoid.End A where - toFun := MulDistribMulAction.toMonoidHom A - map_one' := MonoidHom.ext <| one_smul M - map_mul' x y := MonoidHom.ext <| mul_smul x y - end section @@ -409,36 +333,6 @@ theorem smul_div' (r : M) (x y : A) : r • (x / y) = r • x / r • y := end -/-- The tautological action by `AddMonoid.End α` on `α`. - -This generalizes `Function.End.applyMulAction`. -/ -instance AddMonoid.End.applyDistribMulAction [AddMonoid α] : - DistribMulAction (AddMonoid.End α) α where - smul := (· <| ·) - smul_zero := AddMonoidHom.map_zero - smul_add := AddMonoidHom.map_add - one_smul _ := rfl - mul_smul _ _ _ := rfl - -@[simp] -theorem AddMonoid.End.smul_def [AddMonoid α] (f : AddMonoid.End α) (a : α) : f • a = f a := - rfl - -/-- `AddMonoid.End.applyDistribMulAction` is faithful. -/ -instance AddMonoid.End.applyFaithfulSMul [AddMonoid α] : - FaithfulSMul (AddMonoid.End α) α := - ⟨fun {_ _ h} => AddMonoidHom.ext h⟩ - -/-- Each non-zero element of a `GroupWithZero` defines an additive monoid isomorphism of an -`AddMonoid` on which it acts distributively. -This is a stronger version of `DistribMulAction.toAddMonoidHom`. -/ -def DistribMulAction.toAddEquiv₀ {α : Type*} (β : Type*) [GroupWithZero α] [AddMonoid β] - [DistribMulAction α β] (x : α) (hx : x ≠ 0) : β ≃+ β := - { DistribMulAction.toAddMonoidHom β x with - invFun := fun b ↦ x⁻¹ • b - left_inv := fun b ↦ inv_smul_smul₀ hx b - right_inv := fun b ↦ smul_inv_smul₀ hx b } - section Group variable [Group α] [AddMonoid β] [DistribMulAction α β] diff --git a/Mathlib/Algebra/GroupWithZero/Action/End.lean b/Mathlib/Algebra/GroupWithZero/Action/End.lean new file mode 100644 index 0000000000000..e38aaaf10e21d --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Action/End.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.End +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units + +/-! +# Group actions and (endo)morphisms +-/ + +assert_not_exists Equiv.Perm.equivUnitsEnd +assert_not_exists Prod.fst_mul +assert_not_exists Ring + +open Function + +variable {M N A B α β : Type*} + +/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. + +See also `Function.Surjective.mulActionLeft` and `Function.Surjective.moduleLeft`. +-/ +abbrev Function.Surjective.distribMulActionLeft {R S M : Type*} [Monoid R] [AddMonoid M] + [DistribMulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Function.Surjective f) + (hsmul : ∀ (c) (x : M), f c • x = c • x) : DistribMulAction S M := + { hf.distribSMulLeft f hsmul, hf.mulActionLeft f hsmul with } + +section AddMonoid + +variable (A) [AddMonoid A] [Monoid M] [DistribMulAction M A] + +/-- Compose a `DistribMulAction` with a `MonoidHom`, with action `f r' • m`. +See note [reducible non-instances]. -/ +abbrev DistribMulAction.compHom [Monoid N] (f : N →* M) : DistribMulAction N A := + { DistribSMul.compFun A f, MulAction.compHom A f with } + +end AddMonoid + +section Monoid + +variable (A) [Monoid A] [Monoid M] [MulDistribMulAction M A] + +/-- Compose a `MulDistribMulAction` with a `MonoidHom`, with action `f r' • m`. +See note [reducible non-instances]. -/ +abbrev MulDistribMulAction.compHom [Monoid N] (f : N →* M) : MulDistribMulAction N A := + { MulAction.compHom A f with + smul_one := fun x => smul_one (f x), + smul_mul := fun x => smul_mul' (f x) } + +end Monoid + +/-- The tautological action by `AddMonoid.End α` on `α`. + +This generalizes `Function.End.applyMulAction`. -/ +instance AddMonoid.End.applyDistribMulAction [AddMonoid α] : + DistribMulAction (AddMonoid.End α) α where + smul := (· <| ·) + smul_zero := AddMonoidHom.map_zero + smul_add := AddMonoidHom.map_add + one_smul _ := rfl + mul_smul _ _ _ := rfl + +@[simp] +theorem AddMonoid.End.smul_def [AddMonoid α] (f : AddMonoid.End α) (a : α) : f • a = f a := + rfl + +/-- `AddMonoid.End.applyDistribMulAction` is faithful. -/ +instance AddMonoid.End.applyFaithfulSMul [AddMonoid α] : + FaithfulSMul (AddMonoid.End α) α := + ⟨fun {_ _ h} => AddMonoidHom.ext h⟩ + +/-- Each non-zero element of a `GroupWithZero` defines an additive monoid isomorphism of an +`AddMonoid` on which it acts distributively. +This is a stronger version of `DistribMulAction.toAddMonoidHom`. -/ +def DistribMulAction.toAddEquiv₀ {α : Type*} (β : Type*) [GroupWithZero α] [AddMonoid β] + [DistribMulAction α β] (x : α) (hx : x ≠ 0) : β ≃+ β := + { DistribMulAction.toAddMonoidHom β x with + invFun := fun b ↦ x⁻¹ • b + left_inv := fun b ↦ inv_smul_smul₀ hx b + right_inv := fun b ↦ smul_inv_smul₀ hx b } + +variable (M A) in +/-- Each element of the monoid defines a monoid homomorphism. -/ +@[simps] +def MulDistribMulAction.toMonoidEnd [Monoid M] [Monoid A] [MulDistribMulAction M A] : + M →* Monoid.End A where + toFun := MulDistribMulAction.toMonoidHom A + map_one' := MonoidHom.ext <| one_smul M + map_mul' x y := MonoidHom.ext <| mul_smul x y diff --git a/Mathlib/Algebra/GroupWithZero/Action/Faithful.lean b/Mathlib/Algebra/GroupWithZero/Action/Faithful.lean new file mode 100644 index 0000000000000..9fb6f8fc8c954 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Action/Faithful.lean @@ -0,0 +1,23 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.Faithful +import Mathlib.Algebra.GroupWithZero.NeZero + +/-! +# Faithful actions involving groups with zero +-/ + +assert_not_exists Equiv.Perm.equivUnitsEnd +assert_not_exists Prod.fst_mul +assert_not_exists Ring + +open Function + +variable {M N A B α β : Type*} + +/-- `Monoid.toMulAction` is faithful on nontrivial cancellative monoids with zero. -/ +instance CancelMonoidWithZero.faithfulSMul [CancelMonoidWithZero α] [Nontrivial α] : + FaithfulSMul α α where eq_of_smul_eq_smul h := mul_left_injective₀ one_ne_zero (h 1) diff --git a/Mathlib/Algebra/GroupWithZero/Action/Prod.lean b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean index b825703247594..ff3a88f7e4b13 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Prod.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Prod.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Eric Wieser -/ import Mathlib.Algebra.Group.Action.Prod -import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End /-! # Prod instances for multiplicative actions with zero diff --git a/Mathlib/Algebra/GroupWithZero/Action/Units.lean b/Mathlib/Algebra/GroupWithZero/Action/Units.lean index c517ff3399f1e..5a5df0d53edce 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Units.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Units.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Action.Units import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Units.Basic /-! # Multiplicative actions with zero on and by `Mˣ` @@ -23,7 +24,39 @@ admits a `MulDistribMulAction G Mˣ` structure, again with the obvious definitio * `Algebra.GroupWithZero.Action.Prod` -/ -variable {G M α : Type*} +variable {G M α β : Type*} + +section GroupWithZero +variable [GroupWithZero α] [MulAction α β] {a : α} + +@[simp] lemma inv_smul_smul₀ (ha : a ≠ 0) (x : β) : a⁻¹ • a • x = x := + inv_smul_smul (Units.mk0 a ha) x + +@[simp] +lemma smul_inv_smul₀ (ha : a ≠ 0) (x : β) : a • a⁻¹ • x = x := smul_inv_smul (Units.mk0 a ha) x + +lemma inv_smul_eq_iff₀ (ha : a ≠ 0) {x y : β} : a⁻¹ • x = y ↔ x = a • y := + inv_smul_eq_iff (g := Units.mk0 a ha) + +lemma eq_inv_smul_iff₀ (ha : a ≠ 0) {x y : β} : x = a⁻¹ • y ↔ a • x = y := + eq_inv_smul_iff (g := Units.mk0 a ha) + +@[simp] +lemma Commute.smul_right_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} + (ha : a ≠ 0) : Commute x (a • y) ↔ Commute x y := Commute.smul_right_iff (g := Units.mk0 a ha) + +@[simp] +lemma Commute.smul_left_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTower α β β] {x y : β} + (ha : a ≠ 0) : Commute (a • x) y ↔ Commute x y := Commute.smul_left_iff (g := Units.mk0 a ha) + +/-- Right scalar multiplication as an order isomorphism. -/ +@[simps] def Equiv.smulRight (ha : a ≠ 0) : β ≃ β where + toFun b := a • b + invFun b := a⁻¹ • b + left_inv := inv_smul_smul₀ ha + right_inv := smul_inv_smul₀ ha + +end GroupWithZero namespace Units diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 9b2be9d66ea8b..47b020f71c946 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -3,6 +3,8 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ +import Mathlib.Algebra.GroupWithZero.Action.End +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.SMulWithZero import Mathlib.Data.Int.Cast.Lemmas diff --git a/Mathlib/Algebra/Ring/Action/Basic.lean b/Mathlib/Algebra/Ring/Action/Basic.lean index 8466b8daf17c7..d29c1984e45d9 100644 --- a/Mathlib/Algebra/Ring/Action/Basic.lean +++ b/Mathlib/Algebra/Ring/Action/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Ring.Hom.Defs /-! diff --git a/Mathlib/Data/NNRat/Lemmas.lean b/Mathlib/Data/NNRat/Lemmas.lean index 91aadd51015c4..88467e704c1a4 100644 --- a/Mathlib/Data/NNRat/Lemmas.lean +++ b/Mathlib/Data/NNRat/Lemmas.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Field.Rat import Mathlib.Algebra.Group.Indicator +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Order.Field.Rat /-! diff --git a/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean b/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean index d780cdee10131..9e0ab1536c2d5 100644 --- a/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/DomAct/Basic.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.Pi.Lemmas +import Mathlib.Algebra.Group.Action.Faithful import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.Algebra.Ring.Defs From f8829a66213bebeab14f58c568eeea9442dd607b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 07:22:39 +0000 Subject: [PATCH 04/29] chore: cleanup/split of Data.List.Perm (#18211) Hopefully reduces imports already, and opens the possibility of better splits in `Multiset`. --- Mathlib.lean | 4 +- Mathlib/Algebra/BigOperators/Group/List.lean | 3 +- Mathlib/Data/List/Dedup.lean | 7 + .../Data/List/{Perm.lean => Perm/Basic.lean} | 155 +----------------- Mathlib/Data/List/Perm/Lattice.lean | 107 ++++++++++++ Mathlib/Data/List/Perm/Subperm.lean | 82 +++++++++ Mathlib/Data/List/Permutation.lean | 4 +- Mathlib/Data/List/Prime.lean | 1 - Mathlib/Data/List/Sigma.lean | 3 +- Mathlib/Data/List/Sublists.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 4 +- Mathlib/Data/Multiset/Dedup.lean | 1 + Mathlib/Data/Nat/Factors.lean | 1 + test/MkIffOfInductive.lean | 2 +- 14 files changed, 216 insertions(+), 160 deletions(-) rename Mathlib/Data/List/{Perm.lean => Perm/Basic.lean} (57%) create mode 100644 Mathlib/Data/List/Perm/Lattice.lean create mode 100644 Mathlib/Data/List/Perm/Subperm.lean diff --git a/Mathlib.lean b/Mathlib.lean index 23c1f16aa849d..d7a779c1cc9bf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2398,7 +2398,9 @@ import Mathlib.Data.List.NodupEquivFin import Mathlib.Data.List.OfFn import Mathlib.Data.List.Pairwise import Mathlib.Data.List.Palindrome -import Mathlib.Data.List.Perm +import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Perm.Lattice +import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.List.Permutation import Mathlib.Data.List.Pi import Mathlib.Data.List.Prime diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index c61a25c867599..c354261cd7b5d 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -8,12 +8,13 @@ import Mathlib.Algebra.Group.Int import Mathlib.Algebra.Group.Nat import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.Units.Basic -import Mathlib.Data.List.Perm import Mathlib.Data.List.ProdSigma import Mathlib.Data.List.Range import Mathlib.Data.List.Rotate import Mathlib.Data.List.Pairwise import Mathlib.Data.List.Join +import Mathlib.Data.List.Dedup +import Mathlib.Data.List.Perm.Basic /-! # Sums and products from lists diff --git a/Mathlib/Data/List/Dedup.lean b/Mathlib/Data/List/Dedup.lean index 55576850f682a..2c282a581e0bd 100644 --- a/Mathlib/Data/List/Dedup.lean +++ b/Mathlib/Data/List/Dedup.lean @@ -160,4 +160,11 @@ theorem replicate_dedup {x : α} : ∀ {k}, k ≠ 0 → (replicate k x).dedup = theorem count_dedup (l : List α) (a : α) : l.dedup.count a = if a ∈ l then 1 else 0 := by simp_rw [count_eq_of_nodup <| nodup_dedup l, mem_dedup] +theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := + perm_iff_count.2 fun a => + if h : a ∈ l₁ then by + simp [h, nodup_dedup, p.subset h] + else by + simp [h, count_eq_zero_of_not_mem, mt p.mem_iff.2] + end List diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm/Basic.lean similarity index 57% rename from Mathlib/Data/List/Perm.lean rename to Mathlib/Data/List/Perm/Basic.lean index efeb8dd594a06..c665732b59f94 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm/Basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Data.List.Dedup -import Mathlib.Data.List.Lattice import Batteries.Data.List.Perm +import Mathlib.Logic.Relation +import Mathlib.Order.RelClasses +import Mathlib.Data.List.Forall2 /-! # List Permutations @@ -97,28 +98,6 @@ theorem rel_perm (hr : BiUnique r) : (Forall₂ r ⇒ Forall₂ r ⇒ (· ↔ · end Rel -section Subperm - -attribute [trans] Subperm.trans - -end Subperm - -lemma subperm_iff : l₁ <+~ l₂ ↔ ∃ l, l ~ l₂ ∧ l₁ <+ l := by - refine ⟨?_, fun ⟨l, h₁, h₂⟩ ↦ h₂.subperm.trans h₁.subperm⟩ - rintro ⟨l, h₁, h₂⟩ - obtain ⟨l', h₂⟩ := h₂.exists_perm_append - exact ⟨l₁ ++ l', (h₂.trans (h₁.append_right _)).symm, (prefix_append _ _).sublist⟩ - -@[simp] lemma subperm_singleton_iff : l <+~ [a] ↔ l = [] ∨ l = [a] := by - constructor - · rw [subperm_iff] - rintro ⟨s, hla, h⟩ - rwa [perm_singleton.mp hla, sublist_singleton] at h - · rintro (rfl | rfl) - exacts [nil_subperm, Subperm.refl _] - -lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ _⟩ - lemma count_eq_count_filter_add [DecidableEq α] (P : α → Prop) [DecidablePred P] (l : List α) (a : α) : count a l = count a (l.filter P) + count a (l.filter (¬ P ·)) := by @@ -165,70 +144,8 @@ theorem perm_option_toList {o₁ o₂ : Option α} : o₁.toList ~ o₂.toList @[deprecated (since := "2024-10-16")] alias perm_option_to_list := perm_option_toList -alias ⟨subperm.of_cons, subperm.cons⟩ := subperm_cons - --- Porting note: commented out ---attribute [protected] subperm.cons - -theorem cons_subperm_of_mem {a : α} {l₁ l₂ : List α} (d₁ : Nodup l₁) (h₁ : a ∉ l₁) (h₂ : a ∈ l₂) - (s : l₁ <+~ l₂) : a :: l₁ <+~ l₂ := by - rcases s with ⟨l, p, s⟩ - induction s generalizing l₁ with - | slnil => cases h₂ - | @cons r₁ r₂ b s' ih => - simp? at h₂ says simp only [mem_cons] at h₂ - cases' h₂ with e m - · subst b - exact ⟨a :: r₁, p.cons a, s'.cons₂ _⟩ - · rcases ih d₁ h₁ m p with ⟨t, p', s'⟩ - exact ⟨t, p', s'.cons _⟩ - | @cons₂ r₁ r₂ b _ ih => - have bm : b ∈ l₁ := p.subset <| mem_cons_self _ _ - have am : a ∈ r₂ := by - simp only [find?, mem_cons] at h₂ - exact h₂.resolve_left fun e => h₁ <| e.symm ▸ bm - rcases append_of_mem bm with ⟨t₁, t₂, rfl⟩ - have st : t₁ ++ t₂ <+ t₁ ++ b :: t₂ := by simp - rcases ih (d₁.sublist st) (mt (fun x => st.subset x) h₁) am - (Perm.cons_inv <| p.trans perm_middle) with - ⟨t, p', s'⟩ - exact - ⟨b :: t, (p'.cons b).trans <| (swap _ _ _).trans (perm_middle.symm.cons a), s'.cons₂ _⟩ - -protected theorem Nodup.subperm (d : Nodup l₁) (H : l₁ ⊆ l₂) : l₁ <+~ l₂ := - subperm_of_subset d H - -section - -variable [DecidableEq α] - -theorem Perm.bagInter_right {l₁ l₂ : List α} (t : List α) (h : l₁ ~ l₂) : - l₁.bagInter t ~ l₂.bagInter t := by - induction h generalizing t with - | nil => simp - | cons x => by_cases x ∈ t <;> simp [*, Perm.cons] - | swap x y => - by_cases h : x = y - · simp [h] - by_cases xt : x ∈ t <;> by_cases yt : y ∈ t - · simp [xt, yt, mem_erase_of_ne h, mem_erase_of_ne (Ne.symm h), erase_comm, swap] - · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] - · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] - · simp [xt, yt] - | trans _ _ ih_1 ih_2 => exact (ih_1 _).trans (ih_2 _) - -theorem Perm.bagInter_left (l : List α) {t₁ t₂ : List α} (p : t₁ ~ t₂) : - l.bagInter t₁ = l.bagInter t₂ := by - induction' l with a l IH generalizing t₁ t₂ p; · simp - by_cases h : a ∈ t₁ - · simp [h, p.subset h, IH (p.erase _)] - · simp [h, mt p.mem_iff.2 h, IH p] - -theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t₁ ~ t₂) : - l₁.bagInter t₁ ~ l₂.bagInter t₂ := - ht.bagInter_left l₂ ▸ hl.bagInter_right _ - -theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : +theorem perm_replicate_append_replicate + [DecidableEq α] {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by @@ -237,34 +154,6 @@ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h · simp [subset_def, or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] -theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := - perm_iff_count.2 fun a => - if h : a ∈ l₁ then by - simp [h, nodup_dedup, p.subset h] - else by - simp [h, count_eq_zero_of_not_mem, mt p.mem_iff.2] - -theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : - l ∩ (t₁ ++ t₂) ~ l ∩ t₁ ++ l ∩ t₂ := by - induction l with - | nil => simp - | cons x xs l_ih => - by_cases h₁ : x ∈ t₁ - · have h₂ : x ∉ t₂ := h h₁ - simp [*] - by_cases h₂ : x ∈ t₂ - · simp only [*, inter_cons_of_not_mem, false_or, mem_append, inter_cons_of_mem, - not_false_iff] - refine Perm.trans (Perm.cons _ l_ih) ?_ - change [x] ++ xs ∩ t₁ ++ xs ∩ t₂ ~ xs ∩ t₁ ++ ([x] ++ xs ∩ t₂) - rw [← List.append_assoc] - solve_by_elim [Perm.append_right, perm_append_comm] - · simp [*] - -end - - - theorem Perm.bind_left (l : List α) {f g : α → List β} (h : ∀ a ∈ l, f a ~ g a) : l.bind f ~ l.bind g := Perm.join_congr <| by @@ -313,38 +202,4 @@ theorem perm_lookmap (f : α → Option α) {l₁ l₂ : List α} rw [@eq_comm _ y, @eq_comm _ c] apply h d hd c hc -theorem Perm.take_inter [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ ys) - (h' : ys.Nodup) : xs.take n ~ ys.inter (xs.take n) := by - simp only [List.inter] - exact Perm.trans (show xs.take n ~ xs.filter (xs.take n).elem by - conv_lhs => rw [Nodup.take_eq_filter_mem ((Perm.nodup_iff h).2 h')]) - (Perm.filter _ h) - -theorem Perm.drop_inter [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ ys) (h' : ys.Nodup) : - xs.drop n ~ ys.inter (xs.drop n) := by - by_cases h'' : n ≤ xs.length - · let n' := xs.length - n - have h₀ : n = xs.length - n' := by rwa [Nat.sub_sub_self] - have h₁ : xs.drop n = (xs.reverse.take n').reverse := by - rw [take_reverse, h₀, reverse_reverse] - rw [h₁] - apply (reverse_perm _).trans - rw [inter_reverse] - apply Perm.take_inter _ _ h' - apply (reverse_perm _).trans; assumption - · have : drop n xs = [] := by - apply eq_nil_of_length_eq_zero - rw [length_drop, Nat.sub_eq_zero_iff_le] - apply le_of_not_ge h'' - simp [this, List.inter] - -theorem Perm.dropSlice_inter [DecidableEq α] {xs ys : List α} (n m : ℕ) (h : xs ~ ys) - (h' : ys.Nodup) : List.dropSlice n m xs ~ ys ∩ List.dropSlice n m xs := by - simp only [dropSlice_eq] - have : n ≤ n + m := Nat.le_add_right _ _ - have h₂ := h.nodup_iff.2 h' - apply Perm.trans _ (Perm.inter_append _).symm - · exact Perm.append (Perm.take_inter _ h h') (Perm.drop_inter _ h h') - · exact disjoint_take_drop h₂ this - end List diff --git a/Mathlib/Data/List/Perm/Lattice.lean b/Mathlib/Data/List/Perm/Lattice.lean new file mode 100644 index 0000000000000..44b3c318a9ff0 --- /dev/null +++ b/Mathlib/Data/List/Perm/Lattice.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Data.List.Forall2 +import Mathlib.Data.Set.Pairwise.Basic +import Mathlib.Data.List.Basic +import Mathlib.Data.List.Lattice +import Mathlib.Data.List.Nodup + +/-! +# List Permutations and list lattice operations. + +This file develops theory about the `List.Perm` relation and the lattice structure on lists. +-/ + +-- Make sure we don't import algebra +assert_not_exists Monoid + +open Nat + +namespace List +variable {α β : Type*} {l l₁ l₂ : List α} {a : α} + +open Perm (swap) + +variable [DecidableEq α] + +theorem Perm.bagInter_right {l₁ l₂ : List α} (t : List α) (h : l₁ ~ l₂) : + l₁.bagInter t ~ l₂.bagInter t := by + induction h generalizing t with + | nil => simp + | cons x => by_cases x ∈ t <;> simp [*, Perm.cons] + | swap x y => + by_cases h : x = y + · simp [h] + by_cases xt : x ∈ t <;> by_cases yt : y ∈ t + · simp [xt, yt, mem_erase_of_ne h, mem_erase_of_ne (Ne.symm h), erase_comm, swap] + · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] + · simp [xt, yt, mt mem_of_mem_erase, Perm.cons] + · simp [xt, yt] + | trans _ _ ih_1 ih_2 => exact (ih_1 _).trans (ih_2 _) + +theorem Perm.bagInter_left (l : List α) {t₁ t₂ : List α} (p : t₁ ~ t₂) : + l.bagInter t₁ = l.bagInter t₂ := by + induction' l with a l IH generalizing t₁ t₂ p; · simp + by_cases h : a ∈ t₁ + · simp [h, p.subset h, IH (p.erase _)] + · simp [h, mt p.mem_iff.2 h, IH p] + +theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t₁ ~ t₂) : + l₁.bagInter t₁ ~ l₂.bagInter t₂ := + ht.bagInter_left l₂ ▸ hl.bagInter_right _ + +theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : + l ∩ (t₁ ++ t₂) ~ l ∩ t₁ ++ l ∩ t₂ := by + induction l with + | nil => simp + | cons x xs l_ih => + by_cases h₁ : x ∈ t₁ + · have h₂ : x ∉ t₂ := h h₁ + simp [*] + by_cases h₂ : x ∈ t₂ + · simp only [*, inter_cons_of_not_mem, false_or, mem_append, inter_cons_of_mem, + not_false_iff] + refine Perm.trans (Perm.cons _ l_ih) ?_ + change [x] ++ xs ∩ t₁ ++ xs ∩ t₂ ~ xs ∩ t₁ ++ ([x] ++ xs ∩ t₂) + rw [← List.append_assoc] + solve_by_elim [Perm.append_right, perm_append_comm] + · simp [*] + +theorem Perm.take_inter {xs ys : List α} (n : ℕ) (h : xs ~ ys) + (h' : ys.Nodup) : xs.take n ~ ys.inter (xs.take n) := by + simp only [List.inter] + exact Perm.trans (show xs.take n ~ xs.filter (xs.take n).elem by + conv_lhs => rw [Nodup.take_eq_filter_mem ((Perm.nodup_iff h).2 h')]) + (Perm.filter _ h) + +theorem Perm.drop_inter {xs ys : List α} (n : ℕ) (h : xs ~ ys) (h' : ys.Nodup) : + xs.drop n ~ ys.inter (xs.drop n) := by + by_cases h'' : n ≤ xs.length + · let n' := xs.length - n + have h₀ : n = xs.length - n' := by rwa [Nat.sub_sub_self] + have h₁ : xs.drop n = (xs.reverse.take n').reverse := by + rw [take_reverse, h₀, reverse_reverse] + rw [h₁] + apply (reverse_perm _).trans + rw [inter_reverse] + apply Perm.take_inter _ _ h' + apply (reverse_perm _).trans; assumption + · have : drop n xs = [] := by + apply eq_nil_of_length_eq_zero + rw [length_drop, Nat.sub_eq_zero_iff_le] + apply le_of_not_ge h'' + simp [this, List.inter] + +theorem Perm.dropSlice_inter {xs ys : List α} (n m : ℕ) (h : xs ~ ys) + (h' : ys.Nodup) : List.dropSlice n m xs ~ ys ∩ List.dropSlice n m xs := by + simp only [dropSlice_eq] + have : n ≤ n + m := Nat.le_add_right _ _ + have h₂ := h.nodup_iff.2 h' + apply Perm.trans _ (Perm.inter_append _).symm + · exact Perm.append (Perm.take_inter _ h h') (Perm.drop_inter _ h h') + · exact disjoint_take_drop h₂ this + +end List diff --git a/Mathlib/Data/List/Perm/Subperm.lean b/Mathlib/Data/List/Perm/Subperm.lean new file mode 100644 index 0000000000000..4093ff882b050 --- /dev/null +++ b/Mathlib/Data/List/Perm/Subperm.lean @@ -0,0 +1,82 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Batteries.Data.List.Pairwise +import Batteries.Data.List.Perm +import Mathlib.Data.List.Basic + +/-! +# List Sub-permutations + +This file develops theory about the `List.Subperm` relation. + +## Notation + +The notation `<+~` is used for sub-permutations. +-/ + +open Nat + +namespace List +variable {α β : Type*} {l l₁ l₂ : List α} {a : α} + +open Perm + +section Subperm + +attribute [trans] Subperm.trans + +end Subperm + +lemma subperm_iff : l₁ <+~ l₂ ↔ ∃ l, l ~ l₂ ∧ l₁ <+ l := by + refine ⟨?_, fun ⟨l, h₁, h₂⟩ ↦ h₂.subperm.trans h₁.subperm⟩ + rintro ⟨l, h₁, h₂⟩ + obtain ⟨l', h₂⟩ := h₂.exists_perm_append + exact ⟨l₁ ++ l', (h₂.trans (h₁.append_right _)).symm, (prefix_append _ _).sublist⟩ + +@[simp] lemma subperm_singleton_iff : l <+~ [a] ↔ l = [] ∨ l = [a] := by + constructor + · rw [subperm_iff] + rintro ⟨s, hla, h⟩ + rwa [perm_singleton.mp hla, sublist_singleton] at h + · rintro (rfl | rfl) + exacts [nil_subperm, Subperm.refl _] + +lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ _⟩ + +alias ⟨subperm.of_cons, subperm.cons⟩ := subperm_cons + +-- Porting note: commented out +--attribute [protected] subperm.cons + +theorem cons_subperm_of_mem {a : α} {l₁ l₂ : List α} (d₁ : Nodup l₁) (h₁ : a ∉ l₁) (h₂ : a ∈ l₂) + (s : l₁ <+~ l₂) : a :: l₁ <+~ l₂ := by + rcases s with ⟨l, p, s⟩ + induction s generalizing l₁ with + | slnil => cases h₂ + | @cons r₁ r₂ b s' ih => + simp? at h₂ says simp only [mem_cons] at h₂ + cases' h₂ with e m + · subst b + exact ⟨a :: r₁, p.cons a, s'.cons₂ _⟩ + · rcases ih d₁ h₁ m p with ⟨t, p', s'⟩ + exact ⟨t, p', s'.cons _⟩ + | @cons₂ r₁ r₂ b _ ih => + have bm : b ∈ l₁ := p.subset <| mem_cons_self _ _ + have am : a ∈ r₂ := by + simp only [find?, mem_cons] at h₂ + exact h₂.resolve_left fun e => h₁ <| e.symm ▸ bm + rcases append_of_mem bm with ⟨t₁, t₂, rfl⟩ + have st : t₁ ++ t₂ <+ t₁ ++ b :: t₂ := by simp + rcases ih (d₁.sublist st) (mt (fun x => st.subset x) h₁) am + (Perm.cons_inv <| p.trans perm_middle) with + ⟨t, p', s'⟩ + exact + ⟨b :: t, (p'.cons b).trans <| (swap _ _ _).trans (perm_middle.symm.cons a), s'.cons₂ _⟩ + +protected theorem Nodup.subperm (d : Nodup l₁) (H : l₁ ⊆ l₂) : l₁ <+~ l₂ := + subperm_of_subset d H + +end List diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 5bcb74097b405..2798655fad89c 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -6,12 +6,10 @@ Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, M import Mathlib.Data.List.Join import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.List.Count -import Mathlib.Data.List.Dedup import Mathlib.Data.List.Duplicate import Mathlib.Data.List.InsertIdx -import Mathlib.Data.List.Lattice -import Mathlib.Data.List.Perm import Batteries.Data.List.Perm +import Mathlib.Data.List.Perm.Basic /-! # Permutations of a list diff --git a/Mathlib/Data/List/Prime.lean b/Mathlib/Data/List/Prime.lean index 45c8f957e4d93..2dce64c46d9df 100644 --- a/Mathlib/Data/List/Prime.lean +++ b/Mathlib/Data/List/Prime.lean @@ -5,7 +5,6 @@ Authors: Johannes Hölzl, Jens Wagemaker, Anne Baanen -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.List -import Mathlib.Data.List.Perm /-! # Products of lists of prime elements. diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index 06fa67f877f9c..2ca3f245e6cdf 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Sean Leather -/ -import Mathlib.Data.List.Perm import Mathlib.Data.List.Pairwise +import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Nodup /-! # Utilities for lists of sigmas diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index bbe9b5bfedc30..48d8a5c0e48e9 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Choose.Basic -import Mathlib.Data.List.Perm import Mathlib.Data.List.Range +import Mathlib.Data.List.Perm.Basic /-! # sublists diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 667c2e2f5f6f1..56c3a9eefb7ec 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -5,9 +5,11 @@ Authors: Mario Carneiro -/ import Mathlib.Algebra.Group.Nat import Mathlib.Algebra.Order.Sub.Unbundled.Basic -import Mathlib.Data.List.Perm +import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.Set.List import Mathlib.Order.Hom.Basic +import Mathlib.Data.List.Perm.Lattice +import Mathlib.Data.List.Perm.Basic /-! # Multisets diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 4dee768fea017..1c4737c0e1395 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Data.List.Dedup import Mathlib.Data.Multiset.Nodup /-! diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index df65aa467ceea..c3d6f9ad26ecb 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -8,6 +8,7 @@ import Mathlib.Data.Nat.GCD.Basic import Mathlib.Data.Nat.Prime.Defs import Mathlib.Data.List.Prime import Mathlib.Data.List.Sort +import Mathlib.Data.List.Perm.Subperm /-! # Prime numbers diff --git a/test/MkIffOfInductive.lean b/test/MkIffOfInductive.lean index e2c50aa66a11b..340beef419ff7 100644 --- a/test/MkIffOfInductive.lean +++ b/test/MkIffOfInductive.lean @@ -1,5 +1,5 @@ import Mathlib.Tactic.MkIffOfInductiveProp -import Mathlib.Data.List.Perm +import Mathlib.Data.List.Perm.Lattice mk_iff_of_inductive_prop List.Chain test.chain_iff example {α : Type _} (R : α → α → Prop) (a : α) (al : List α) : From 9d502188e7df20d93aded2b312b28259ff1da308 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 25 Oct 2024 07:57:03 +0000 Subject: [PATCH 05/29] chore: remove variables (#18193) #17715 --- Mathlib/Algebra/Module/ZLattice/Covolume.lean | 5 +---- Mathlib/Analysis/Calculus/Deriv/Prod.lean | 7 +------ Mathlib/Analysis/Calculus/FDeriv/Pi.lean | 1 - Mathlib/Analysis/Calculus/FDeriv/Star.lean | 2 +- Mathlib/Analysis/Calculus/Gradient/Basic.lean | 2 +- .../InverseFunctionTheorem/ApproximatesLinearOn.lean | 2 -- Mathlib/Analysis/Calculus/LineDeriv/Basic.lean | 2 +- Mathlib/MeasureTheory/Integral/SetIntegral.lean | 2 +- Mathlib/Order/PrimeIdeal.lean | 2 +- Mathlib/Probability/Distributions/Uniform.lean | 4 ++-- Mathlib/Probability/Kernel/Composition.lean | 6 +++--- Mathlib/Probability/Kernel/Disintegration/CondCDF.lean | 4 ++-- Mathlib/Probability/Kernel/Integral.lean | 2 +- Mathlib/Probability/Kernel/Invariance.lean | 2 +- Mathlib/Probability/Kernel/MeasurableIntegral.lean | 2 +- Mathlib/Probability/Martingale/OptionalSampling.lean | 4 ++-- Mathlib/RingTheory/Flat/Stability.lean | 2 +- Mathlib/RingTheory/HahnSeries/Summable.lean | 2 +- Mathlib/RingTheory/LocalRing/Module.lean | 2 +- Mathlib/RingTheory/Regular/IsSMulRegular.lean | 2 +- Mathlib/Topology/Order/UpperLowerSetTopology.lean | 2 +- 21 files changed, 24 insertions(+), 35 deletions(-) diff --git a/Mathlib/Algebra/Module/ZLattice/Covolume.lean b/Mathlib/Algebra/Module/ZLattice/Covolume.lean index 8b3705ee2aa1e..f5740d29e7505 100644 --- a/Mathlib/Algebra/Module/ZLattice/Covolume.lean +++ b/Mathlib/Algebra/Module/ZLattice/Covolume.lean @@ -33,10 +33,7 @@ open Submodule MeasureTheory Module MeasureTheory Module ZSpan section General -variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] -variable [ProperSpace E] [MeasurableSpace E] -variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice K L] +variable {E : Type*} [NormedAddCommGroup E] [ProperSpace E] [MeasurableSpace E] (L : Submodule ℤ E) /-- The covolume of a `ℤ`-lattice is the volume of some fundamental domain; see `ZLattice.covolume_eq_volume` for the proof that the volume does not depend on the choice of diff --git a/Mathlib/Analysis/Calculus/Deriv/Prod.lean b/Mathlib/Analysis/Calculus/Deriv/Prod.lean index b0640b705f958..dfd2e896128ad 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Prod.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Prod.lean @@ -29,12 +29,7 @@ open Filter Asymptotics Set variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] -variable {f f₀ f₁ g : 𝕜 → F} -variable {f' f₀' f₁' g' : F} -variable {x : 𝕜} -variable {s t : Set 𝕜} -variable {L L₁ L₂ : Filter 𝕜} +variable {f₁ : 𝕜 → F} {f₁' : F} {x : 𝕜} {s : Set 𝕜} {L : Filter 𝕜} section CartesianProduct diff --git a/Mathlib/Analysis/Calculus/FDeriv/Pi.lean b/Mathlib/Analysis/Calculus/FDeriv/Pi.lean index f985ae505ea08..fdc2b96749939 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Pi.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Pi.lean @@ -11,7 +11,6 @@ import Mathlib.Analysis.Calculus.FDeriv.Add variable {𝕜 ι : Type*} [DecidableEq ι] [Fintype ι] [NontriviallyNormedField 𝕜] variable {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] @[fun_prop] theorem hasFDerivAt_update (x : ∀ i, E i) {i : ι} (y : E i) : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Star.lean b/Mathlib/Analysis/Calculus/FDeriv/Star.lean index 48f652a8dd13d..933c9c2c4fdc1 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Star.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Star.lean @@ -27,7 +27,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [StarAddMonoid F] [NormedSpace 𝕜 F] [StarModule 𝕜 F] [ContinuousStar F] -variable {f : E → F} {f' : E →L[𝕜] F} (e : E →L[𝕜] F) {x : E} {s : Set E} {L : Filter E} +variable {f : E → F} {f' : E →L[𝕜] F} {x : E} {s : Set E} {L : Filter E} @[fun_prop] theorem HasStrictFDerivAt.star (h : HasStrictFDerivAt f f' x) : diff --git a/Mathlib/Analysis/Calculus/Gradient/Basic.lean b/Mathlib/Analysis/Calculus/Gradient/Basic.lean index 558b091cb21ac..5ebf3d941d6a8 100644 --- a/Mathlib/Analysis/Calculus/Gradient/Basic.lean +++ b/Mathlib/Analysis/Calculus/Gradient/Basic.lean @@ -252,7 +252,7 @@ section congr /-! ### Congruence properties of the Gradient -/ -variable {f₀ f₁ : F → 𝕜} {f₀' f₁' : F} {x₀ x₁ : F} {s₀ s₁ t : Set F} {L₀ L₁ : Filter F} +variable {f₀ f₁ : F → 𝕜} {f₀' f₁' : F} {t : Set F} theorem Filter.EventuallyEq.hasGradientAtFilter_iff (h₀ : f₀ =ᶠ[L] f₁) (hx : f₀ x = f₁ x) (h₁ : f₀' = f₁') : HasGradientAtFilter f₀ f₀' x L ↔ HasGradientAtFilter f₁ f₁' x L := diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 67b3db1e4cd80..b23236aaf0baf 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -52,8 +52,6 @@ noncomputable section variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] -variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G'] variable {ε : ℝ} open Filter Metric Set diff --git a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean index 3de1ac70f4558..b675314ee9710 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean @@ -475,7 +475,7 @@ section CompRight variable {E : Type*} [AddCommGroup E] [Module 𝕜 E] {E' : Type*} [AddCommGroup E'] [Module 𝕜 E'] - {f : E → F} {f' : F} {x v : E'} {L : E' →ₗ[𝕜] E} + {f : E → F} {f' : F} {x : E'} {L : E' →ₗ[𝕜] E} theorem HasLineDerivAt.of_comp {v : E'} (hf : HasLineDerivAt 𝕜 (f ∘ L) f' x v) : HasLineDerivAt 𝕜 f f' (L x) (L v) := by diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 6376b28c74713..49b3978cb0252 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -66,7 +66,7 @@ variable [MeasurableSpace X] section NormedAddCommGroup variable [NormedAddCommGroup E] [NormedSpace ℝ E] - {f g : X → E} {s t : Set X} {μ ν : Measure X} {l l' : Filter X} + {f g : X → E} {s t : Set X} {μ : Measure X} theorem setIntegral_congr_ae₀ (hs : NullMeasurableSet s μ) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := diff --git a/Mathlib/Order/PrimeIdeal.lean b/Mathlib/Order/PrimeIdeal.lean index e8b6364e05f7b..aef2ade7bdfd4 100644 --- a/Mathlib/Order/PrimeIdeal.lean +++ b/Mathlib/Order/PrimeIdeal.lean @@ -101,7 +101,7 @@ end Preorder section SemilatticeInf -variable [SemilatticeInf P] {x y : P} {I : Ideal P} +variable [SemilatticeInf P] {I : Ideal P} theorem IsPrime.mem_or_mem (hI : IsPrime I) {x y : P} : x ⊓ y ∈ I → x ∈ I ∨ y ∈ I := by contrapose! diff --git a/Mathlib/Probability/Distributions/Uniform.lean b/Mathlib/Probability/Distributions/Uniform.lean index 43f54dab8c27d..77f3a13bb4a8c 100644 --- a/Mathlib/Probability/Distributions/Uniform.lean +++ b/Mathlib/Probability/Distributions/Uniform.lean @@ -48,7 +48,7 @@ noncomputable section namespace MeasureTheory -variable {E : Type*} [MeasurableSpace E] {m : Measure E} {μ : Measure E} +variable {E : Type*} [MeasurableSpace E] {μ : Measure E} namespace pdf @@ -206,7 +206,7 @@ end MeasureTheory namespace PMF -variable {α β γ : Type*} +variable {α : Type*} open scoped Classical NNReal ENNReal diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 9a2db5e2be4bc..5dbebd1c9f25c 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -65,7 +65,7 @@ namespace ProbabilityTheory namespace Kernel -variable {α β ι : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} +variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} section CompositionProduct @@ -335,7 +335,7 @@ end Ae section Restrict -variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [IsSFiniteKernel η] {a : α} +variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [IsSFiniteKernel η] theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) : Kernel.restrict κ hs ⊗ₖ Kernel.restrict η ht = Kernel.restrict (κ ⊗ₖ η) (hs.prod ht) := by @@ -759,7 +759,7 @@ def prodMkLeft (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) : Kernel ( def prodMkRight (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) : Kernel (α × γ) β := comap κ Prod.fst measurable_fst -variable {γ : Type*} {mγ : MeasurableSpace γ} {f : β → γ} {g : γ → α} +variable {γ : Type*} {mγ : MeasurableSpace γ} @[simp] theorem prodMkLeft_apply (κ : Kernel α β) (ca : γ × α) : prodMkLeft γ κ ca = κ ca.snd := diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index a5c415461bb5b..c02e122e3cff2 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -42,7 +42,7 @@ open scoped NNReal ENNReal MeasureTheory Topology namespace MeasureTheory.Measure -variable {α β : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) +variable {α : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ)) /-- Measure on `α` such that for a measurable set `s`, `ρ.IicSnd r s = ρ (s ×ˢ Iic r)`. -/ noncomputable def IicSnd (r : ℝ) : Measure α := @@ -112,7 +112,7 @@ open MeasureTheory namespace ProbabilityTheory -variable {α β ι : Type*} {mα : MeasurableSpace α} +variable {α : Type*} {mα : MeasurableSpace α} attribute [local instance] MeasureTheory.Measure.IsFiniteMeasure.IicSnd diff --git a/Mathlib/Probability/Kernel/Integral.lean b/Mathlib/Probability/Kernel/Integral.lean index f259203acbe03..e97ef25407719 100644 --- a/Mathlib/Probability/Kernel/Integral.lean +++ b/Mathlib/Probability/Kernel/Integral.lean @@ -91,7 +91,7 @@ end Const section Restrict -variable {s t : Set β} +variable {s : Set β} @[simp] theorem integral_restrict (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/Invariance.lean b/Mathlib/Probability/Kernel/Invariance.lean index 58d1a2185d16b..e9d40e85a0dd9 100644 --- a/Mathlib/Probability/Kernel/Invariance.lean +++ b/Mathlib/Probability/Kernel/Invariance.lean @@ -30,7 +30,7 @@ open scoped MeasureTheory ENNReal ProbabilityTheory namespace ProbabilityTheory -variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} +variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} namespace Kernel diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 2d5fff10ca258..63c99fb77ebfa 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -231,7 +231,7 @@ alias _root_.Measurable.set_lintegral_kernel := _root_.Measurable.setLIntegral_k end Lintegral -variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] [IsSFiniteKernel η] +variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] theorem measurableSet_kernel_integrable ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) : MeasurableSet {x | Integrable (f x) (κ x)} := by diff --git a/Mathlib/Probability/Martingale/OptionalSampling.lean b/Mathlib/Probability/Martingale/OptionalSampling.lean index 535963b759aba..90b9489aaa39c 100644 --- a/Mathlib/Probability/Martingale/OptionalSampling.lean +++ b/Mathlib/Probability/Martingale/OptionalSampling.lean @@ -143,11 +143,11 @@ and is a measurable space with the Borel σ-algebra. -/ variable {ι : Type*} [LinearOrder ι] [LocallyFiniteOrder ι] [OrderBot ι] [TopologicalSpace ι] [DiscreteTopology ι] [MeasurableSpace ι] [BorelSpace ι] [MeasurableSpace E] [BorelSpace E] - [SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i n : ι} + [SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i : ι} theorem condexp_stoppedValue_stopping_time_ae_eq_restrict_le (h : Martingale f ℱ μ) (hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) [SigmaFinite (μ.trim hσ.measurableSpace_le)] - (hτ_le : ∀ x, τ x ≤ n) : + (hτ_le : ∀ x, τ x ≤ i) : μ[stoppedValue f τ|hσ.measurableSpace] =ᵐ[μ.restrict {x : Ω | τ x ≤ σ x}] stoppedValue f τ := by rw [ae_eq_restrict_iff_indicator_ae_eq (hτ.measurableSpace_le _ (hτ.measurableSet_le_stopping_time hσ))] diff --git a/Mathlib/RingTheory/Flat/Stability.lean b/Mathlib/RingTheory/Flat/Stability.lean index 57de113408f23..8b0e1c14b0fa9 100644 --- a/Mathlib/RingTheory/Flat/Stability.lean +++ b/Mathlib/RingTheory/Flat/Stability.lean @@ -153,7 +153,7 @@ section Localization variable {R : Type u} {M Mp : Type*} (Rp : Type v) [CommRing R] [AddCommGroup M] [Module R M] [CommRing Rp] [Algebra R Rp] - [AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] (f : M →ₗ[R] Mp) + [AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] instance localizedModule [Module.Flat R M] (S : Submonoid R) : Module.Flat (Localization S) (LocalizedModule S M) := by diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index b4e933391ff50..c3b6bf05ac681 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -42,7 +42,7 @@ open Pointwise noncomputable section -variable {Γ Γ' R V α β : Type*} +variable {Γ R α β : Type*} namespace HahnSeries diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 332bccee7c8f5..d446db6698308 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -28,7 +28,7 @@ This file gathers various results about finite modules over a local ring `(R, `l` is a split injection if and only if `k ⊗ l` is a (split) injection. -/ -variable {R S} [CommRing R] [CommRing S] [Algebra R S] +variable {R} [CommRing R] section diff --git a/Mathlib/RingTheory/Regular/IsSMulRegular.lean b/Mathlib/RingTheory/Regular/IsSMulRegular.lean index efe4d52d7cef9..3619263bfda72 100644 --- a/Mathlib/RingTheory/Regular/IsSMulRegular.lean +++ b/Mathlib/RingTheory/Regular/IsSMulRegular.lean @@ -136,7 +136,7 @@ open Submodule Pointwise variable (M) [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup M'] [Module R M'] [AddCommGroup M''] [Module R M''] - (I : Ideal R) (N : Submodule R M) (r : R) + (N : Submodule R M) (r : R) variable (R) in lemma biUnion_associatedPrimes_eq_compl_regular [IsNoetherianRing R] : diff --git a/Mathlib/Topology/Order/UpperLowerSetTopology.lean b/Mathlib/Topology/Order/UpperLowerSetTopology.lean index 02681dbd9f86d..c3e5e3a52ad12 100644 --- a/Mathlib/Topology/Order/UpperLowerSetTopology.lean +++ b/Mathlib/Topology/Order/UpperLowerSetTopology.lean @@ -95,7 +95,7 @@ protected def rec {β : WithUpperSet α → Sort*} (h : ∀ a, β (toUpperSet a) instance [Nonempty α] : Nonempty (WithUpperSet α) := ‹Nonempty α› instance [Inhabited α] : Inhabited (WithUpperSet α) := ‹Inhabited α› -variable [Preorder α] [Preorder β] [Preorder γ] +variable [Preorder α] [Preorder β] instance : Preorder (WithUpperSet α) := ‹Preorder α› instance : TopologicalSpace (WithUpperSet α) := upperSet α From d85e7dbb201ed04d86c1f27e373dc6f5fc637092 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 08:08:23 +0000 Subject: [PATCH 06/29] chore: split Topology.Separation (#18213) Avoids needing GDelta / UniformSpace in many later files. (Identified using #18156) --- Mathlib.lean | 3 +- Mathlib/Analysis/Convex/Radon.lean | 2 +- Mathlib/Dynamics/FixedPoints/Topology.lean | 2 +- .../Topology/Algebra/InfiniteSum/Defs.lean | 2 +- Mathlib/Topology/Algebra/Semigroup.lean | 2 +- .../Topology/Compactification/OnePoint.lean | 1 - Mathlib/Topology/Connected/Separation.lean | 2 +- Mathlib/Topology/CountableSeparatingOn.lean | 2 +- Mathlib/Topology/DenseEmbedding.lean | 2 +- Mathlib/Topology/DiscreteQuotient.lean | 1 - Mathlib/Topology/DiscreteSubset.lean | 2 +- Mathlib/Topology/ExtendFrom.lean | 2 +- Mathlib/Topology/Filter.lean | 2 +- Mathlib/Topology/IndicatorConstPointwise.lean | 2 +- Mathlib/Topology/Instances/Irrational.lean | 1 + Mathlib/Topology/Order/OrderClosed.lean | 2 +- Mathlib/Topology/Order/Priestley.lean | 2 +- Mathlib/Topology/Perfect.lean | 2 +- Mathlib/Topology/QuasiSeparated.lean | 1 - Mathlib/Topology/SeparatedMap.lean | 2 +- .../Basic.lean} | 103 +-------------- Mathlib/Topology/Separation/GDelta.lean | 125 ++++++++++++++++++ Mathlib/Topology/Separation/NotNormal.lean | 1 - Mathlib/Topology/ShrinkingLemma.lean | 2 +- Mathlib/Topology/Sober.lean | 1 - Mathlib/Topology/Specialization.lean | 1 - Mathlib/Topology/Support.lean | 2 +- Mathlib/Topology/UniformSpace/Compact.lean | 1 - Mathlib/Topology/UniformSpace/Separation.lean | 2 +- .../UniformSpace/UniformConvergence.lean | 1 - 30 files changed, 152 insertions(+), 124 deletions(-) rename Mathlib/Topology/{Separation.lean => Separation/Basic.lean} (96%) create mode 100644 Mathlib/Topology/Separation/GDelta.lean diff --git a/Mathlib.lean b/Mathlib.lean index d7a779c1cc9bf..46bd976d0d3e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4927,7 +4927,8 @@ import Mathlib.Topology.QuasiSeparated import Mathlib.Topology.RestrictGen import Mathlib.Topology.Semicontinuous import Mathlib.Topology.SeparatedMap -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.Separation.GDelta import Mathlib.Topology.Separation.NotNormal import Mathlib.Topology.Sequences import Mathlib.Topology.Sets.Closeds diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index 2a65ed2e8986a..21470d3e027c9 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -6,7 +6,7 @@ Authors: Vasily Nesterov import Mathlib.Analysis.Convex.Combination import Mathlib.Data.Set.Card import Mathlib.LinearAlgebra.AffineSpace.FiniteDimensional -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Radon's theorem on convex sets diff --git a/Mathlib/Dynamics/FixedPoints/Topology.lean b/Mathlib/Dynamics/FixedPoints/Topology.lean index 092a3114204f4..c657d45d24997 100644 --- a/Mathlib/Dynamics/FixedPoints/Topology.lean +++ b/Mathlib/Dynamics/FixedPoints/Topology.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Johannes Hölzl -/ import Mathlib.Dynamics.FixedPoints.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Topological properties of fixed points diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index 92ae820ce40a0..2bf685cd201e1 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Order.Filter.AtTopBot.BigOperators -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Infinite sum and product over a topological monoid diff --git a/Mathlib/Topology/Algebra/Semigroup.lean b/Mathlib/Topology/Algebra/Semigroup.lean index 9069ed9f766ed..327643f2aa4c3 100644 --- a/Mathlib/Topology/Algebra/Semigroup.lean +++ b/Mathlib/Topology/Algebra/Semigroup.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 David Wärn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn -/ -import Mathlib.Topology.Separation import Mathlib.Algebra.Group.Defs +import Mathlib.Topology.Separation.Basic /-! # Idempotents in topological semigroups diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index e870661f60d4f..dec6ca2d9e95f 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yourong Zang, Yury Kudryashov -/ import Mathlib.Data.Fintype.Option -import Mathlib.Topology.Separation import Mathlib.Topology.Sets.Opens /-! diff --git a/Mathlib/Topology/Connected/Separation.lean b/Mathlib/Topology/Connected/Separation.lean index 8ce21856914c0..f5e3e067c4839 100644 --- a/Mathlib/Topology/Connected/Separation.lean +++ b/Mathlib/Topology/Connected/Separation.lean @@ -3,7 +3,7 @@ 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.Separation +import Mathlib.Topology.Separation.Basic /-! # Separation and (dis)connectedness properties of topological spaces. diff --git a/Mathlib/Topology/CountableSeparatingOn.lean b/Mathlib/Topology/CountableSeparatingOn.lean index fd76eddbe41b3..4531e4bdaaa85 100644 --- a/Mathlib/Topology/CountableSeparatingOn.lean +++ b/Mathlib/Topology/CountableSeparatingOn.lean @@ -3,8 +3,8 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Topology.Separation import Mathlib.Order.Filter.CountableSeparatingOn +import Mathlib.Topology.Separation.Basic /-! # Countable separating families of sets in topological spaces diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index 210caa3dca529..f42b8f1149a17 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Reid Barton. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ -import Mathlib.Topology.Separation import Mathlib.Topology.Bases +import Mathlib.Topology.Separation.Basic /-! # Dense embeddings diff --git a/Mathlib/Topology/DiscreteQuotient.lean b/Mathlib/Topology/DiscreteQuotient.lean index 4c20019504bc8..4048f4765b17f 100644 --- a/Mathlib/Topology/DiscreteQuotient.lean +++ b/Mathlib/Topology/DiscreteQuotient.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Calle Sönne, Adam Topaz -/ import Mathlib.Data.Setoid.Partition -import Mathlib.Topology.Separation import Mathlib.Topology.LocallyConstant.Basic /-! diff --git a/Mathlib/Topology/DiscreteSubset.lean b/Mathlib/Topology/DiscreteSubset.lean index 825509afa0b34..70e38c4cc3f54 100644 --- a/Mathlib/Topology/DiscreteSubset.lean +++ b/Mathlib/Topology/DiscreteSubset.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash, Bhavik Mehta, Daniel Weber -/ import Mathlib.Topology.Constructions -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Discrete subsets of topological spaces diff --git a/Mathlib/Topology/ExtendFrom.lean b/Mathlib/Topology/ExtendFrom.lean index 1c043eb2756ea..015142b528778 100644 --- a/Mathlib/Topology/ExtendFrom.lean +++ b/Mathlib/Topology/ExtendFrom.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Anatole Dedecker -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Extending a function from a subset diff --git a/Mathlib/Topology/Filter.lean b/Mathlib/Topology/Filter.lean index 78b2714743dbe..e2239722200e6 100644 --- a/Mathlib/Topology/Filter.lean +++ b/Mathlib/Topology/Filter.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Order.Filter.Lift -import Mathlib.Topology.Separation import Mathlib.Order.Interval.Set.Monotone +import Mathlib.Topology.Separation.Basic /-! # Topology on the set of filters on a type diff --git a/Mathlib/Topology/IndicatorConstPointwise.lean b/Mathlib/Topology/IndicatorConstPointwise.lean index c1f6ad354e8ff..deb980b52a7e8 100644 --- a/Mathlib/Topology/IndicatorConstPointwise.lean +++ b/Mathlib/Topology/IndicatorConstPointwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kalle Kytölä -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Pointwise convergence of indicator functions diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index 7d4b39145e504..ad7ac6309e9fa 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Data.Real.Irrational import Mathlib.Data.Rat.Encodable import Mathlib.Topology.GDelta +import Mathlib.Topology.Separation.GDelta /-! # Topology of irrational numbers diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index e2763d592bc24..9cdc99b9b7005 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Order-closed topologies diff --git a/Mathlib/Topology/Order/Priestley.lean b/Mathlib/Topology/Order/Priestley.lean index ca780c30f0cc8..bdad508392e4a 100644 --- a/Mathlib/Topology/Order/Priestley.lean +++ b/Mathlib/Topology/Order/Priestley.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.UpperLower.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Priestley spaces diff --git a/Mathlib/Topology/Perfect.lean b/Mathlib/Topology/Perfect.lean index c390399234943..88e7d579f3dd7 100644 --- a/Mathlib/Topology/Perfect.lean +++ b/Mathlib/Topology/Perfect.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Perfect Sets diff --git a/Mathlib/Topology/QuasiSeparated.lean b/Mathlib/Topology/QuasiSeparated.lean index 73240b9ea2c7e..f550bb0a8528c 100644 --- a/Mathlib/Topology/QuasiSeparated.lean +++ b/Mathlib/Topology/QuasiSeparated.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.Topology.Separation import Mathlib.Topology.NoetherianSpace /-! diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index dae896a7a7e4b..992ff27b3e79a 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu -/ import Mathlib.Topology.Connected.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Separated maps and locally injective maps out of a topological space. diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation/Basic.lean similarity index 96% rename from Mathlib/Topology/Separation.lean rename to Mathlib/Topology/Separation/Basic.lean index 284f7ef7efa3c..cd3e2041e3180 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Inseparable -import Mathlib.Topology.GDelta /-! # Separation properties of topological spaces. @@ -50,19 +49,13 @@ This file defines the predicate `SeparatedNhds`, and common separation axioms us to conclude that this is equivalent to all subspaces being normal. Such a space is not necessarily Hausdorff or regular, even if it is T₀. * `T5Space`: A T₅ space is a completely normal T₁ space. T₅ implies T₄. -* `PerfectlyNormalSpace`: A perfectly normal space is a normal space such that - closed sets are Gδ. -* `T6Space`: A T₆ space is a Perfectly normal T₁ space. T₆ implies T₅. + +See `Mathlib.Topology.Separation.GDelta` for the definitions of `PerfectlyNormalSpace` and +`T6Space`. Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but occasionally the literature swaps definitions for e.g. T₃ and regular. -### TODO - -* Add perfectly normal and T6 spaces. -* Use `hasSeparatingCovers_iff_separatedNhds` to prove that perfectly normal spaces - are completely normal. - ## Main results ### T₀ spaces @@ -124,6 +117,8 @@ If the space is also Lindelöf: -/ +assert_not_exists UniformSpace + open Function Set Filter Topology TopologicalSpace universe u v @@ -900,36 +895,6 @@ instance (priority := 100) ConnectedSpace.neBot_nhdsWithin_compl_of_nontrivial_o contra (compl_union_self _) (Set.nonempty_compl_of_nontrivial _) (singleton_nonempty _) simp [compl_inter_self {x}] at contra -theorem IsGδ.compl_singleton (x : X) [T1Space X] : IsGδ ({x}ᶜ : Set X) := - isOpen_compl_singleton.isGδ - -@[deprecated (since := "2024-02-15")] alias isGδ_compl_singleton := IsGδ.compl_singleton - -theorem Set.Countable.isGδ_compl {s : Set X} [T1Space X] (hs : s.Countable) : IsGδ sᶜ := by - rw [← biUnion_of_singleton s, compl_iUnion₂] - exact .biInter hs fun x _ => .compl_singleton x - -theorem Set.Finite.isGδ_compl {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ sᶜ := - hs.countable.isGδ_compl - -theorem Set.Subsingleton.isGδ_compl {s : Set X} [T1Space X] (hs : s.Subsingleton) : IsGδ sᶜ := - hs.finite.isGδ_compl - -theorem Finset.isGδ_compl [T1Space X] (s : Finset X) : IsGδ (sᶜ : Set X) := - s.finite_toSet.isGδ_compl - -protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) : - IsGδ ({x} : Set X) := by - rcases (nhds_basis_opens x).exists_antitone_subbasis with ⟨U, hU, h_basis⟩ - rw [← biInter_basis_nhds h_basis.toHasBasis] - exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ - -@[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton - -theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : - IsGδ s := - Finite.induction_on hs .empty fun _ _ ↦ .union (.singleton _) - theorem SeparationQuotient.t1Space_iff : T1Space (SeparationQuotient X) ↔ R0Space X := by rw [r0Space_iff, ((t1Space_TFAE (SeparationQuotient X)).out 0 9 :)] constructor @@ -2418,62 +2383,6 @@ instance [CompletelyNormalSpace X] [R0Space X] : T5Space (SeparationQuotient X) end CompletelyNormal -section PerfectlyNormal - -/-- A topological space `X` is a *perfectly normal space* provided it is normal and -closed sets are Gδ. -/ -class PerfectlyNormalSpace (X : Type u) [TopologicalSpace X] extends NormalSpace X : Prop where - closed_gdelta : ∀ ⦃h : Set X⦄, IsClosed h → IsGδ h - -/-- Lemma that allows the easy conclusion that perfectly normal spaces are completely normal. -/ -theorem Disjoint.hasSeparatingCover_closed_gdelta_right {s t : Set X} [NormalSpace X] - (st_dis : Disjoint s t) (t_cl : IsClosed t) (t_gd : IsGδ t) : HasSeparatingCover s t := by - obtain ⟨T, T_open, T_count, T_int⟩ := t_gd - rcases T.eq_empty_or_nonempty with rfl | T_nonempty - · rw [T_int, sInter_empty] at st_dis - rw [(s.disjoint_univ).mp st_dis] - exact t.hasSeparatingCover_empty_left - obtain ⟨g, g_surj⟩ := T_count.exists_surjective T_nonempty - choose g' g'_open clt_sub_g' clg'_sub_g using fun n ↦ by - apply normal_exists_closure_subset t_cl (T_open (g n).1 (g n).2) - rw [T_int] - exact sInter_subset_of_mem (g n).2 - have clg'_int : t = ⋂ i, closure (g' i) := by - apply (subset_iInter fun n ↦ (clt_sub_g' n).trans subset_closure).antisymm - rw [T_int] - refine subset_sInter fun t tinT ↦ ?_ - obtain ⟨n, gn⟩ := g_surj ⟨t, tinT⟩ - refine iInter_subset_of_subset n <| (clg'_sub_g n).trans ?_ - rw [gn] - use fun n ↦ (closure (g' n))ᶜ - constructor - · rw [← compl_iInter, subset_compl_comm, ← clg'_int] - exact st_dis.subset_compl_left - · refine fun n ↦ ⟨isOpen_compl_iff.mpr isClosed_closure, ?_⟩ - simp only [closure_compl, disjoint_compl_left_iff_subset] - rw [← closure_eq_iff_isClosed.mpr t_cl] at clt_sub_g' - exact subset_closure.trans <| (clt_sub_g' n).trans <| (g'_open n).subset_interior_closure - -instance (priority := 100) PerfectlyNormalSpace.toCompletelyNormalSpace - [PerfectlyNormalSpace X] : CompletelyNormalSpace X where - completely_normal _ _ hd₁ hd₂ := separatedNhds_iff_disjoint.mp <| - hasSeparatingCovers_iff_separatedNhds.mp - ⟨(hd₂.hasSeparatingCover_closed_gdelta_right isClosed_closure <| - closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure, - ((Disjoint.symm hd₁).hasSeparatingCover_closed_gdelta_right isClosed_closure <| - closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure⟩ - -/-- A T₆ space is a perfectly normal T₁ space. -/ -class T6Space (X : Type u) [TopologicalSpace X] extends T1Space X, PerfectlyNormalSpace X : Prop - --- see Note [lower instance priority] -/-- A `T₆` space is a `T₅` space. -/ -instance (priority := 100) T6Space.toT5Space [T6Space X] : T5Space X where - -- follows from type-class inference - -end PerfectlyNormal - - /-- In a compact T₂ space, the connected component of a point equals the intersection of all its clopen neighbourhoods. -/ theorem connectedComponent_eq_iInter_isClopen [T2Space X] [CompactSpace X] (x : X) : @@ -2670,4 +2579,4 @@ instance ConnectedComponents.t2 [T2Space X] [CompactSpace X] : T2Space (Connecte refine ⟨Vᶜ, V, hU.compl.isOpen, hU.isOpen, ?_, hb mem_connectedComponent, disjoint_compl_left⟩ exact fun h => flip Set.Nonempty.ne_empty ha ⟨a, mem_connectedComponent, h⟩ -set_option linter.style.longFile 2800 +set_option linter.style.longFile 2700 diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean new file mode 100644 index 0000000000000..3b58de08a18ee --- /dev/null +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Topology.Compactness.Lindelof +import Mathlib.Topology.Compactness.SigmaCompact +import Mathlib.Topology.Connected.TotallyDisconnected +import Mathlib.Topology.Inseparable +import Mathlib.Topology.GDelta +import Mathlib.Topology.Separation.Basic + +/-! +# Separation properties of topological spaces. + +## Main definitions + +* `PerfectlyNormalSpace`: A perfectly normal space is a normal space such that + closed sets are Gδ. +* `T6Space`: A T₆ space is a Perfectly normal T₁ space. T₆ implies T₅. + +Note that `mathlib` adopts the modern convention that `m ≤ n` if and only if `T_m → T_n`, but +occasionally the literature swaps definitions for e.g. T₃ and regular. + +### TODO + +* Use `hasSeparatingCovers_iff_separatedNhds` to prove that perfectly normal spaces + are completely normal. + +-/ + +open Function Set Filter Topology TopologicalSpace + +universe u v + +variable {X : Type*} {Y : Type*} [TopologicalSpace X] + +section Separation + +theorem IsGδ.compl_singleton (x : X) [T1Space X] : IsGδ ({x}ᶜ : Set X) := + isOpen_compl_singleton.isGδ + +@[deprecated (since := "2024-02-15")] alias isGδ_compl_singleton := IsGδ.compl_singleton + +theorem Set.Countable.isGδ_compl {s : Set X} [T1Space X] (hs : s.Countable) : IsGδ sᶜ := by + rw [← biUnion_of_singleton s, compl_iUnion₂] + exact .biInter hs fun x _ => .compl_singleton x + +theorem Set.Finite.isGδ_compl {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ sᶜ := + hs.countable.isGδ_compl + +theorem Set.Subsingleton.isGδ_compl {s : Set X} [T1Space X] (hs : s.Subsingleton) : IsGδ sᶜ := + hs.finite.isGδ_compl + +theorem Finset.isGδ_compl [T1Space X] (s : Finset X) : IsGδ (sᶜ : Set X) := + s.finite_toSet.isGδ_compl + +protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) : + IsGδ ({x} : Set X) := by + rcases (nhds_basis_opens x).exists_antitone_subbasis with ⟨U, hU, h_basis⟩ + rw [← biInter_basis_nhds h_basis.toHasBasis] + exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ + +@[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton + +theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : + IsGδ s := + Finite.induction_on hs .empty fun _ _ ↦ .union (.singleton _) + + +section PerfectlyNormal + +/-- A topological space `X` is a *perfectly normal space* provided it is normal and +closed sets are Gδ. -/ +class PerfectlyNormalSpace (X : Type u) [TopologicalSpace X] extends NormalSpace X : Prop where + closed_gdelta : ∀ ⦃h : Set X⦄, IsClosed h → IsGδ h + +/-- Lemma that allows the easy conclusion that perfectly normal spaces are completely normal. -/ +theorem Disjoint.hasSeparatingCover_closed_gdelta_right {s t : Set X} [NormalSpace X] + (st_dis : Disjoint s t) (t_cl : IsClosed t) (t_gd : IsGδ t) : HasSeparatingCover s t := by + obtain ⟨T, T_open, T_count, T_int⟩ := t_gd + rcases T.eq_empty_or_nonempty with rfl | T_nonempty + · rw [T_int, sInter_empty] at st_dis + rw [(s.disjoint_univ).mp st_dis] + exact t.hasSeparatingCover_empty_left + obtain ⟨g, g_surj⟩ := T_count.exists_surjective T_nonempty + choose g' g'_open clt_sub_g' clg'_sub_g using fun n ↦ by + apply normal_exists_closure_subset t_cl (T_open (g n).1 (g n).2) + rw [T_int] + exact sInter_subset_of_mem (g n).2 + have clg'_int : t = ⋂ i, closure (g' i) := by + apply (subset_iInter fun n ↦ (clt_sub_g' n).trans subset_closure).antisymm + rw [T_int] + refine subset_sInter fun t tinT ↦ ?_ + obtain ⟨n, gn⟩ := g_surj ⟨t, tinT⟩ + refine iInter_subset_of_subset n <| (clg'_sub_g n).trans ?_ + rw [gn] + use fun n ↦ (closure (g' n))ᶜ + constructor + · rw [← compl_iInter, subset_compl_comm, ← clg'_int] + exact st_dis.subset_compl_left + · refine fun n ↦ ⟨isOpen_compl_iff.mpr isClosed_closure, ?_⟩ + simp only [closure_compl, disjoint_compl_left_iff_subset] + rw [← closure_eq_iff_isClosed.mpr t_cl] at clt_sub_g' + exact subset_closure.trans <| (clt_sub_g' n).trans <| (g'_open n).subset_interior_closure + +instance (priority := 100) PerfectlyNormalSpace.toCompletelyNormalSpace + [PerfectlyNormalSpace X] : CompletelyNormalSpace X where + completely_normal _ _ hd₁ hd₂ := separatedNhds_iff_disjoint.mp <| + hasSeparatingCovers_iff_separatedNhds.mp + ⟨(hd₂.hasSeparatingCover_closed_gdelta_right isClosed_closure <| + closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure, + ((Disjoint.symm hd₁).hasSeparatingCover_closed_gdelta_right isClosed_closure <| + closed_gdelta isClosed_closure).mono (fun ⦃_⦄ a ↦ a) subset_closure⟩ + +/-- A T₆ space is a perfectly normal T₁ space. -/ +class T6Space (X : Type u) [TopologicalSpace X] extends T1Space X, PerfectlyNormalSpace X : Prop + +-- see Note [lower instance priority] +/-- A `T₆` space is a `T₅` space. -/ +instance (priority := 100) T6Space.toT5Space [T6Space X] : T5Space X where + -- follows from type-class inference +end PerfectlyNormal + +end Separation diff --git a/Mathlib/Topology/Separation/NotNormal.lean b/Mathlib/Topology/Separation/NotNormal.lean index a022ac86d5359..348f0ea4324cb 100644 --- a/Mathlib/Topology/Separation/NotNormal.lean +++ b/Mathlib/Topology/Separation/NotNormal.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Data.Real.Cardinality -import Mathlib.Topology.Separation import Mathlib.Topology.TietzeExtension /-! # Not normal topological spaces diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index b88d57ecb3233..7cd187f075dfa 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Reid Barton -/ -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # The shrinking lemma diff --git a/Mathlib/Topology/Sober.lean b/Mathlib/Topology/Sober.lean index 6cf96933116ef..29cbe4476433a 100644 --- a/Mathlib/Topology/Sober.lean +++ b/Mathlib/Topology/Sober.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.Topology.Separation import Mathlib.Topology.Sets.Closeds /-! diff --git a/Mathlib/Topology/Specialization.lean b/Mathlib/Topology/Specialization.lean index 1c7ce45b4bb4b..bd57a351804bd 100644 --- a/Mathlib/Topology/Specialization.lean +++ b/Mathlib/Topology/Specialization.lean @@ -6,7 +6,6 @@ Authors: Yaël Dillies import Mathlib.Order.Category.Preord import Mathlib.Topology.Category.TopCat.Basic import Mathlib.Topology.ContinuousMap.Basic -import Mathlib.Topology.Separation import Mathlib.Topology.Order.UpperLowerSetTopology /-! diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 30537f87b604c..c533d9dbd075b 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -6,7 +6,7 @@ Authors: Floris van Doorn, Patrick Massot import Mathlib.Algebra.GroupWithZero.Indicator import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Algebra.Module.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # The topological support of a function diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index b9bc2156e776b..f8308127388e0 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -5,7 +5,6 @@ Authors: Patrick Massot, Yury Kudryashov -/ import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.Equicontinuity -import Mathlib.Topology.Separation import Mathlib.Topology.Support /-! diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index f50de8d5deebd..60e6f6a085671 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov -/ import Mathlib.Tactic.ApplyFun import Mathlib.Topology.UniformSpace.Basic -import Mathlib.Topology.Separation +import Mathlib.Topology.Separation.Basic /-! # Hausdorff properties of uniform spaces. Separation quotient. diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index 6fe642c53f2e4..aca68d49538ba 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergence.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergence.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 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.Separation import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.UniformSpace.Cauchy From 1232423bfa0c678e65d76b5970e25b91f58374e6 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 08:30:34 +0000 Subject: [PATCH 07/29] chore: re-add `@[simp]` attribute lost in the port (#18121) This PR re-adds the `@[simp]` attributes that had a porting note that they could be proved by `@[simp]`, but where that does not appear to be the case currently. Of course, the `simp` set may have changed sufficiently to require un`@[simp]`ing some of these anyway, but we should add a more up to date comment in that case. --- Mathlib/Combinatorics/Young/YoungDiagram.lean | 1 - Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean | 2 +- Mathlib/GroupTheory/Perm/Sign.lean | 6 +++--- Mathlib/LinearAlgebra/Isomorphisms.lean | 2 +- Mathlib/Logic/Equiv/Defs.lean | 4 ++-- Mathlib/MeasureTheory/Integral/IntervalIntegral.lean | 2 +- Mathlib/MeasureTheory/Measure/NullMeasurable.lean | 1 - Mathlib/RingTheory/AlgebraicIndependent.lean | 12 ++++++++---- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 12 ++++-------- Mathlib/RingTheory/MvPowerSeries/Inverse.lean | 3 +-- Mathlib/RingTheory/PowerSeries/Basic.lean | 6 ++---- Mathlib/RingTheory/PowerSeries/Inverse.lean | 3 +-- Mathlib/SetTheory/Cardinal/ToNat.lean | 3 +-- Mathlib/Topology/Algebra/Module/Basic.lean | 4 ++-- Mathlib/Topology/Category/TopCat/Basic.lean | 4 ++-- Mathlib/Topology/UniformSpace/Equiv.lean | 6 +++--- 16 files changed, 32 insertions(+), 39 deletions(-) diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index fad6a3949ea28..7bb1431dc4004 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -153,7 +153,6 @@ theorem cells_bot : (⊥ : YoungDiagram).cells = ∅ := rfl -- Porting note: removed `↑`, added `.cells` and changed proof --- @[simp] -- Porting note (#10618): simp can prove this @[norm_cast] theorem coe_bot : (⊥ : YoungDiagram).cells = (∅ : Set (ℕ × ℕ)) := by refine Set.eq_of_subset_of_subset ?_ ?_ diff --git a/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean b/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean index b0bcd7be1a4a3..40a58337f4b9e 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/FDeriv.lean @@ -35,7 +35,7 @@ theorem uniqueMDiffOn_iff_uniqueDiffOn : UniqueMDiffOn 𝓘(𝕜, E) s ↔ Uniqu alias ⟨UniqueMDiffOn.uniqueDiffOn, UniqueDiffOn.uniqueMDiffOn⟩ := uniqueMDiffOn_iff_uniqueDiffOn --- Porting note (#10618): was `@[simp, mfld_simps]` but `simp` can prove it +@[simp, mfld_simps] theorem writtenInExtChartAt_model_space : writtenInExtChartAt 𝓘(𝕜, E) 𝓘(𝕜, E') x f = f := rfl diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index 98b302cba6c58..41c9a40ec36c5 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -370,7 +370,7 @@ section SignType.sign variable [Fintype α] ---@[simp] Porting note (#10618): simp can prove +@[simp] theorem sign_mul (f g : Perm α) : sign (f * g) = sign f * sign g := MonoidHom.map_mul sign f g @@ -378,7 +378,7 @@ theorem sign_mul (f g : Perm α) : sign (f * g) = sign f * sign g := theorem sign_trans (f g : Perm α) : sign (f.trans g) = sign g * sign f := by rw [← mul_def, sign_mul] ---@[simp] Porting note (#10618): simp can prove +@[simp] theorem sign_one : sign (1 : Perm α) = 1 := MonoidHom.map_one sign @@ -386,7 +386,7 @@ theorem sign_one : sign (1 : Perm α) = 1 := theorem sign_refl : sign (Equiv.refl α) = 1 := MonoidHom.map_one sign ---@[simp] Porting note (#10618): simp can prove +@[simp] theorem sign_inv (f : Perm α) : sign f⁻¹ = sign f := by rw [MonoidHom.map_inv sign f, Int.units_inv_eq_self] diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index 0dc3fa092cd9c..3a94028500411 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -149,7 +149,7 @@ theorem quotientQuotientEquivQuotientAux_mk (x : M ⧸ S) : quotientQuotientEquivQuotientAux S T h (Quotient.mk x) = mapQ S T LinearMap.id h x := liftQ_apply _ _ _ --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem quotientQuotientEquivQuotientAux_mk_mk (x : M) : quotientQuotientEquivQuotientAux S T h (Quotient.mk (Quotient.mk x)) = Quotient.mk x := by simp diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 0cda134871dce..2650567700550 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -599,7 +599,7 @@ theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort*} theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort*} (F : ∀ a, β₁ a ≃ β₂ a) : (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := rfl --- Porting note (#10618): simp can now prove this, so I have removed `@[simp]` +@[simp] theorem psigmaCongrRight_refl {α} {β : α → Sort*} : (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ' a, β a) := rfl @@ -620,7 +620,7 @@ theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type*} theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type*} (F : ∀ a, β₁ a ≃ β₂ a) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := rfl --- Porting note (#10618): simp can now prove this, so I have removed `@[simp]` +@[simp] theorem sigmaCongrRight_refl {α} {β : α → Type*} : (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := rfl diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 6b904ee39dabd..ff624c6f12d1a 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -120,7 +120,7 @@ theorem intervalIntegrable_const_iff {c : E} : IntervalIntegrable (fun _ => c) μ a b ↔ c = 0 ∨ μ (Ι a b) < ∞ := by simp only [intervalIntegrable_iff, integrableOn_const] -@[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem intervalIntegrable_const [IsLocallyFiniteMeasure μ] {c : E} : IntervalIntegrable (fun _ => c) μ a b := intervalIntegrable_const_iff.2 <| Or.inr measure_Ioc_lt_top diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index 45d1deff59246..a3090e8782b25 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -171,7 +171,6 @@ protected theorem disjointed {f : ℕ → Set α} (h : ∀ i, NullMeasurableSet NullMeasurableSet (disjointed f n) μ := MeasurableSet.disjointed h n --- @[simp] -- Porting note (#10618): simp can prove thisrove this protected theorem const (p : Prop) : NullMeasurableSet { _a : α | p } μ := MeasurableSet.const p diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index a161c703a9c61..1aa1a3d12eece 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -363,13 +363,17 @@ theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply (aeval (fun o : Option ι => o.elim Polynomial.X fun s : ι => Polynomial.C (X s)) y) := rfl ---@[simp] Porting note: removing simp because the linter complains about deterministic timeout +/-- `simp`-normal form of `mvPolynomialOptionEquivPolynomialAdjoin_C` -/ +@[simp] +theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C' + (hx : AlgebraicIndependent R x) (r) : + Polynomial.C (hx.aevalEquiv (C r)) = Polynomial.C (algebraMap _ _ r) := by + simp [aevalEquiv] + theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_C (hx : AlgebraicIndependent R x) (r) : hx.mvPolynomialOptionEquivPolynomialAdjoin (C r) = Polynomial.C (algebraMap _ _ r) := by - rw [AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_apply, aeval_C, - IsScalarTower.algebraMap_apply R (MvPolynomial ι R), ← Polynomial.C_eq_algebraMap, - Polynomial.map_C, RingHom.coe_coe, AlgEquiv.commutes] + simp theorem AlgebraicIndependent.mvPolynomialOptionEquivPolynomialAdjoin_X_none (hx : AlgebraicIndependent R x) : diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index ab05350f3f87a..a9a882a935dd4 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -177,8 +177,7 @@ theorem eq_of_coeff_monomial_ne_zero {m n : σ →₀ ℕ} {a : R} (h : coeff R theorem coeff_comp_monomial (n : σ →₀ ℕ) : (coeff R n).comp (monomial R n) = LinearMap.id := LinearMap.ext <| coeff_monomial_same n --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem coeff_zero (n : σ →₀ ℕ) : coeff R n (0 : MvPowerSeries σ R) = 0 := rfl @@ -434,13 +433,11 @@ theorem constantCoeff_C (a : R) : constantCoeff σ R (C σ R a) = a := theorem constantCoeff_comp_C : (constantCoeff σ R).comp (C σ R) = RingHom.id R := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_zero : constantCoeff σ R 0 = 0 := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_one : constantCoeff σ R 1 = 1 := rfl @@ -454,8 +451,7 @@ theorem isUnit_constantCoeff (φ : MvPowerSeries σ R) (h : IsUnit φ) : IsUnit (constantCoeff σ R φ) := h.map _ --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem coeff_smul (f : MvPowerSeries σ R) (n) (a : R) : coeff _ n (a • f) = a * coeff _ n f := rfl diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index ce619c30ce099..db252318811bf 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -228,8 +228,7 @@ theorem inv_eq_zero {φ : MvPowerSeries σ k} : φ⁻¹ = 0 ↔ constantCoeff σ theorem zero_inv : (0 : MvPowerSeries σ k)⁻¹ = 0 := by rw [inv_eq_zero, constantCoeff_zero] --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem invOfUnit_eq (φ : MvPowerSeries σ k) (h : constantCoeff σ k φ ≠ 0) : invOfUnit φ (Units.mk0 _ h) = φ⁻¹ := rfl diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index bae1dcd02e069..65b589c7373de 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -327,13 +327,11 @@ theorem constantCoeff_C (a : R) : constantCoeff R (C R a) = a := theorem constantCoeff_comp_C : (constantCoeff R).comp (C R) = RingHom.id R := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_zero : constantCoeff R 0 = 0 := rfl --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem constantCoeff_one : constantCoeff R 1 = 1 := rfl diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index 4a9202d3df4a7..18c6913c525cd 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -150,8 +150,7 @@ theorem inv_eq_zero {φ : k⟦X⟧} : φ⁻¹ = 0 ↔ constantCoeff k φ = 0 := theorem zero_inv : (0 : k⟦X⟧)⁻¹ = 0 := MvPowerSeries.zero_inv --- Porting note (#10618): simp can prove this. --- @[simp] +@[simp] theorem invOfUnit_eq (φ : k⟦X⟧) (h : constantCoeff k φ ≠ 0) : invOfUnit φ (Units.mk0 _ h) = φ⁻¹ := MvPowerSeries.invOfUnit_eq _ _ diff --git a/Mathlib/SetTheory/Cardinal/ToNat.lean b/Mathlib/SetTheory/Cardinal/ToNat.lean index f37fdaa87cbf3..44953f885edf7 100644 --- a/Mathlib/SetTheory/Cardinal/ToNat.lean +++ b/Mathlib/SetTheory/Cardinal/ToNat.lean @@ -116,8 +116,7 @@ theorem aleph0_toNat : toNat ℵ₀ = 0 := theorem mk_toNat_eq_card [Fintype α] : toNat #α = Fintype.card α := by simp --- porting note (#10618): simp can prove this --- @[simp] +@[simp] theorem zero_toNat : toNat 0 = 0 := map_zero _ theorem one_toNat : toNat 1 = 1 := map_one _ diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 6293e9dea3c44..fbc19334443c1 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -436,7 +436,7 @@ protected theorem map_zero (f : M₁ →SL[σ₁₂] M₂) : f (0 : M₁) = 0 := protected theorem map_add (f : M₁ →SL[σ₁₂] M₂) (x y : M₁) : f (x + y) = f x + f y := map_add f x y --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] protected theorem map_smulₛₗ (f : M₁ →SL[σ₁₂] M₂) (c : R₁) (x : M₁) : f (c • x) = σ₁₂ c • f x := (toLinearMap _).map_smulₛₗ _ _ @@ -1681,7 +1681,7 @@ theorem map_zero (e : M₁ ≃SL[σ₁₂] M₂) : e (0 : M₁) = 0 := theorem map_add (e : M₁ ≃SL[σ₁₂] M₂) (x y : M₁) : e (x + y) = e x + e y := (e : M₁ →SL[σ₁₂] M₂).map_add x y --- @[simp] -- Porting note (#10618): simp can prove this +@[simp] theorem map_smulₛₗ (e : M₁ ≃SL[σ₁₂] M₂) (c : R₁) (x : M₁) : e (c • x) = σ₁₂ c • e x := (e : M₁ →SL[σ₁₂] M₂).map_smulₛₗ c x diff --git a/Mathlib/Topology/Category/TopCat/Basic.lean b/Mathlib/Topology/Category/TopCat/Basic.lean index 57a2905802c66..87760f303fc0e 100644 --- a/Mathlib/Topology/Category/TopCat/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Basic.lean @@ -53,10 +53,10 @@ instance instFunLike (X Y : TopCat) : FunLike (X ⟶ Y) X Y := instance instContinuousMapClass (X Y : TopCat) : ContinuousMapClass (X ⟶ Y) X Y := inferInstanceAs <| ContinuousMapClass C(X, Y) X Y --- Porting note (#10618): simp can prove this; removed simp +@[simp] theorem id_app (X : TopCat.{u}) (x : ↑X) : (𝟙 X : X ⟶ X) x = x := rfl --- Porting note (#10618): simp can prove this; removed simp +@[simp] theorem comp_app {X Y Z : TopCat.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g : X → Z) x = g (f x) := rfl diff --git a/Mathlib/Topology/UniformSpace/Equiv.lean b/Mathlib/Topology/UniformSpace/Equiv.lean index 6fb396f7b0004..e9f383c879b3b 100644 --- a/Mathlib/Topology/UniformSpace/Equiv.lean +++ b/Mathlib/Topology/UniformSpace/Equiv.lean @@ -178,7 +178,7 @@ theorem symm_comp_self (h : α ≃ᵤ β) : (h.symm : β → α) ∘ h = id := theorem self_comp_symm (h : α ≃ᵤ β) : (h : α → β) ∘ h.symm = id := funext h.apply_symm_apply --- @[simp] -- Porting note (#10618): `simp` can prove this `simp only [Equiv.range_eq_univ]` +@[simp] theorem range_coe (h : α ≃ᵤ β) : range h = univ := h.surjective.range_eq @@ -188,11 +188,11 @@ theorem image_symm (h : α ≃ᵤ β) : image h.symm = preimage h := theorem preimage_symm (h : α ≃ᵤ β) : preimage h.symm = image h := (funext h.toEquiv.image_eq_preimage).symm --- @[simp] -- Porting note (#10618): `simp` can prove this `simp only [Equiv.image_preimage]` +@[simp] theorem image_preimage (h : α ≃ᵤ β) (s : Set β) : h '' (h ⁻¹' s) = s := h.toEquiv.image_preimage s ---@[simp] -- Porting note (#10618): `simp` can prove this `simp only [Equiv.preimage_image]` +@[simp] theorem preimage_image (h : α ≃ᵤ β) (s : Set α) : h ⁻¹' (h '' s) = s := h.toEquiv.preimage_image s From a3663456d2f9d04002039801a734ff0d65f4b4c7 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 25 Oct 2024 09:09:02 +0000 Subject: [PATCH 08/29] chore(RingTheory/Kaehler): fix docstring typo (#18196) --- Mathlib/RingTheory/Kaehler/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index 5492f105c9f9b..d53431a854a9c 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -799,7 +799,7 @@ lemma KaehlerDifferential.exact_mapBaseChange_map : end -/-- The map `I → B ⊗[A] B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ +/-- The map `I → B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ @[simps] noncomputable def KaehlerDifferential.kerToTensor : @@ -811,7 +811,7 @@ def KaehlerDifferential.kerToTensor : algebraMap_eq_smul_one, RingHom.mem_ker.mp x.prop, TensorProduct.zero_tmul, add_zero, RingHom.id_apply] -/-- The map `I/I² → B ⊗[A] B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ +/-- The map `I/I² → B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ noncomputable def KaehlerDifferential.kerCotangentToTensor : (RingHom.ker (algebraMap A B)).Cotangent →ₗ[A] B ⊗[A] Ω[A⁄R] := From 0418dea8a825428fdeb82bb8d0db769a8c98fa58 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 09:18:07 +0000 Subject: [PATCH 09/29] chore(Algebra): delete "`simp` can prove this" porting notes (#18056) This PR checks all porting notes of the form "`simp` can prove this" and if this is indeed the case (and the proof appears to make sense), removes the associated porting note. Co-authored-by: Johan Commelin --- Mathlib/Algebra/Algebra/Equiv.lean | 1 - Mathlib/Algebra/Algebra/Hom.lean | 1 - Mathlib/Algebra/Algebra/NonUnitalHom.lean | 3 --- Mathlib/Algebra/Algebra/Operations.lean | 2 -- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 1 - Mathlib/Algebra/BigOperators/Finsupp.lean | 4 ---- Mathlib/Algebra/CharP/Defs.lean | 1 - Mathlib/Algebra/CubicDiscriminant.lean | 3 --- Mathlib/Algebra/DirectSum/Ring.lean | 1 - Mathlib/Algebra/Field/Subfield.lean | 2 -- Mathlib/Algebra/GeomSum.lean | 2 -- Mathlib/Algebra/Group/Equiv/Basic.lean | 1 - Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 3 --- Mathlib/Algebra/Group/Pointwise/Set/Basic.lean | 1 - Mathlib/Algebra/Group/Subgroup/Basic.lean | 1 - Mathlib/Algebra/Group/Submonoid/Basic.lean | 1 - Mathlib/Algebra/GroupWithZero/Invertible.lean | 1 - Mathlib/Algebra/Homology/HomologicalComplex.lean | 4 ++-- Mathlib/Algebra/Module/Defs.lean | 6 ------ Mathlib/Algebra/Module/LocalizedModule.lean | 2 -- Mathlib/Algebra/Module/Submodule/Range.lean | 1 - Mathlib/Algebra/Module/Torsion.lean | 2 -- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 2 -- Mathlib/Algebra/MonoidAlgebra/Division.lean | 1 - Mathlib/Algebra/MvPolynomial/Expand.lean | 1 - Mathlib/Algebra/MvPolynomial/PDeriv.lean | 1 - Mathlib/Algebra/Order/CauSeq/Completion.lean | 2 -- Mathlib/Algebra/Order/CompleteField.lean | 1 - Mathlib/Algebra/Order/Floor.lean | 4 ---- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 2 -- Mathlib/Algebra/Order/GroupWithZero/Canonical.lean | 2 -- Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean | 2 -- Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean | 8 -------- Mathlib/Algebra/Polynomial/Basic.lean | 1 - Mathlib/Algebra/Polynomial/Coeff.lean | 4 ++-- Mathlib/Algebra/Polynomial/Degree/Definitions.lean | 3 --- Mathlib/Algebra/Polynomial/Derivative.lean | 5 ----- Mathlib/Algebra/Polynomial/Laurent.lean | 1 - Mathlib/Algebra/Polynomial/Reverse.lean | 2 -- Mathlib/Algebra/Quaternion.lean | 5 ----- Mathlib/Algebra/Ring/BooleanRing.lean | 4 ---- Mathlib/Algebra/Ring/Commute.lean | 4 ---- Mathlib/Algebra/Ring/CompTypeclasses.lean | 6 ------ Mathlib/Algebra/Ring/Equiv.lean | 4 ---- Mathlib/Algebra/Ring/Parity.lean | 10 ---------- Mathlib/Algebra/Ring/Subring/Basic.lean | 4 ---- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 1 - Mathlib/Algebra/Squarefree/Basic.lean | 1 - Mathlib/Algebra/Symmetrized.lean | 2 -- Mathlib/Algebra/Tropical/Basic.lean | 2 -- 50 files changed, 4 insertions(+), 125 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 6e995f594999b..d29ef71d49492 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -216,7 +216,6 @@ protected theorem map_mul : ∀ x y, e (x * y) = e x * e y := protected theorem map_one : e 1 = 1 := map_one e --- @[simp] -- Porting note (#10618): simp can prove this @[deprecated map_smul (since := "2024-06-20")] protected theorem map_smul (r : R) (x : A₁) : e (r • x) = r • e x := map_smul _ _ _ diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 0f6f35fd279e4..96b78e4f6c6d4 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -218,7 +218,6 @@ protected theorem map_one : φ 1 = 1 := protected theorem map_pow (x : A) (n : ℕ) : φ (x ^ n) = φ x ^ n := map_pow _ _ _ --- @[simp] -- Porting note (#10618): simp can prove this @[deprecated map_smul (since := "2024-06-26")] protected theorem map_smul (r : R) (x : A) : φ (r • x) = r • φ x := map_smul _ _ _ diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 678960a93a6bf..157c90299af2a 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -245,15 +245,12 @@ theorem coe_mulHom_mk (f : A →ₛₙₐ[φ] B) (h₁ h₂ h₃ h₄) : protected theorem map_smul (f : A →ₛₙₐ[φ] B) (c : R) (x : A) : f (c • x) = (φ c) • f x := map_smulₛₗ _ _ _ --- @[simp] -- Porting note (#10618) : simp can prove this protected theorem map_add (f : A →ₛₙₐ[φ] B) (x y : A) : f (x + y) = f x + f y := map_add _ _ _ --- @[simp] -- Porting note (#10618) : simp can prove this protected theorem map_mul (f : A →ₛₙₐ[φ] B) (x y : A) : f (x * y) = f x * f y := map_mul _ _ _ --- @[simp] -- Porting note (#10618) : simp can prove this protected theorem map_zero (f : A →ₛₙₐ[φ] B) : f 0 = 0 := map_zero _ diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index ec2010316a560..ce82654752d9f 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -184,12 +184,10 @@ theorem mul_bot : M * ⊥ = ⊥ := theorem bot_mul : ⊥ * M = ⊥ := map₂_bot_left _ _ --- @[simp] -- Porting note (#10618): simp can prove this once we have a monoid structure protected theorem one_mul : (1 : Submodule R A) * M = M := by conv_lhs => rw [one_eq_span, ← span_eq M] erw [span_mul_span, one_mul, span_eq] --- @[simp] -- Porting note (#10618): simp can prove this once we have a monoid structure protected theorem mul_one : M * 1 = M := by conv_lhs => rw [one_eq_span, ← span_eq M] erw [span_mul_span, mul_one, span_eq] diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 8e6c9f2dc91ce..a34f329d35336 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -46,7 +46,6 @@ instance SubsemiringClass : SubsemiringClass (Subalgebra R A) A where theorem mem_toSubsemiring {S : Subalgebra R A} {x} : x ∈ S.toSubsemiring ↔ x ∈ S := Iff.rfl --- @[simp] -- Porting note (#10618): simp can prove this theorem mem_carrier {s : Subalgebra R A} {x : A} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index 7d830fdf78f5a..9b241247cc785 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -113,8 +113,6 @@ theorem prod_ite_eq' [DecidableEq α] (f : α →₀ M) (a : α) (b : α → M dsimp [Finsupp.prod] rw [f.support.prod_ite_eq'] --- Porting note (#10618): simp can prove this --- @[simp] theorem sum_ite_self_eq' [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (f.sum fun x v => ite (x = a) v 0) = f a := by classical @@ -406,8 +404,6 @@ theorem equivFunOnFinite_symm_eq_sum [Fintype α] [AddCommMonoid M] (f : α → ext simp --- Porting note (#10618): simp can prove this --- @[simp] theorem liftAddHom_apply_single [AddCommMonoid M] [AddCommMonoid N] (f : α → M →+ N) (a : α) (b : M) : (liftAddHom (α := α) (M := M) (N := N)) f (single a b) = f a b := sum_single_index (f a).map_zero diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index a80dd345ea9ac..9c41f5a2b3741 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -184,7 +184,6 @@ lemma dvd {x : ℕ} (hx : (x : R) = 0) : ringChar R ∣ x := lemma eq_zero [CharZero R] : ringChar R = 0 := eq R 0 --- @[simp] -- Porting note (#10618): simp can prove this lemma Nat.cast_ringChar : (ringChar R : R) = 0 := by rw [ringChar.spec] end ringChar diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 7280b1a8b0e49..42e56fed93eff 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -194,7 +194,6 @@ theorem leadingCoeff_of_c_eq_zero (ha : P.a = 0) (hb : P.b = 0) (hc : P.c = 0) : P.toPoly.leadingCoeff = P.d := by rw [of_c_eq_zero ha hb hc, leadingCoeff_C] --- @[simp] -- porting note (#10618): simp can prove this theorem leadingCoeff_of_c_eq_zero' : (toPoly ⟨0, 0, 0, d⟩).leadingCoeff = d := leadingCoeff_of_c_eq_zero rfl rfl rfl @@ -304,7 +303,6 @@ theorem degree_of_d_eq_zero (ha : P.a = 0) (hb : P.b = 0) (hc : P.c = 0) (hd : P P.toPoly.degree = ⊥ := by rw [of_d_eq_zero ha hb hc hd, degree_zero] --- @[simp] -- porting note (#10618): simp can prove this theorem degree_of_d_eq_zero' : (⟨0, 0, 0, 0⟩ : Cubic R).toPoly.degree = ⊥ := degree_of_d_eq_zero rfl rfl rfl rfl @@ -354,7 +352,6 @@ theorem natDegree_of_c_eq_zero (ha : P.a = 0) (hb : P.b = 0) (hc : P.c = 0) : P.toPoly.natDegree = 0 := by rw [of_c_eq_zero ha hb hc, natDegree_C] --- @[simp] -- porting note (#10618): simp can prove this theorem natDegree_of_c_eq_zero' : (toPoly ⟨0, 0, 0, d⟩).natDegree = 0 := natDegree_of_c_eq_zero rfl rfl rfl diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index b68930289ac05..e30b34a16eff9 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -553,7 +553,6 @@ def toSemiring (f : ∀ i, A i →+ R) (hone : f _ GradedMonoid.GOne.one = 1) simp_rw [of_mul_of, toAddMonoid_of] exact hmul _ _ } --- Porting note (#10618): removed @[simp] as simp can prove this theorem toSemiring_of (f : ∀ i, A i →+ R) (hone hmul) (i : ι) (x : A i) : toSemiring f hone hmul (of _ i x) = f _ x := toAddMonoid_of f i x diff --git a/Mathlib/Algebra/Field/Subfield.lean b/Mathlib/Algebra/Field/Subfield.lean index 30b6731d7f52c..694d6cf7e5504 100644 --- a/Mathlib/Algebra/Field/Subfield.lean +++ b/Mathlib/Algebra/Field/Subfield.lean @@ -176,7 +176,6 @@ instance : SubfieldClass (Subfield K) K where one_mem s := s.one_mem' inv_mem {s} := s.inv_mem' _ --- @[simp] -- Porting note (#10618): simp can prove this (with `coe_toSubring`, which comes later) theorem mem_carrier {s : Subfield K} {x : K} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl @@ -373,7 +372,6 @@ theorem toSubring_subtype_eq_subtype (S : Subfield K) : /-! # Partial order -/ ---@[simp] -- Porting note (#10618): simp can prove this theorem mem_toSubmonoid {s : Subfield K} {x : K} : x ∈ s.toSubmonoid ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 0b34f0c5df0f1..02480271a1392 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -69,8 +69,6 @@ theorem zero_geom_sum : ∀ {n}, ∑ i ∈ range n, (0 : α) ^ i = if n = 0 then theorem one_geom_sum (n : ℕ) : ∑ i ∈ range n, (1 : α) ^ i = n := by simp --- porting note (#10618): simp can prove this --- @[simp] theorem op_geom_sum (x : α) (n : ℕ) : op (∑ i ∈ range n, x ^ i) = ∑ i ∈ range n, op x ^ i := by simp diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 8d70c2f054093..b11426351ebef 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -265,7 +265,6 @@ protected theorem injective (e : M ≃* N) : Function.Injective e := protected theorem surjective (e : M ≃* N) : Function.Surjective e := EquivLike.surjective e --- Porting note (#10618): `simp` can prove this @[to_additive] theorem apply_eq_iff_eq (e : M ≃* N) {x y : M} : e x = e y ↔ x = y := e.injective.eq_iff diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 5b6bf266bb5de..b23f7252ab9a3 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -708,8 +708,6 @@ theorem div_singleton (a : α) : s / {a} = s.image (· / a) := theorem singleton_div (a : α) : {a} / s = s.image (a / ·) := image₂_singleton_left --- @[to_additive (attr := simp)] --- Porting note (#10618): simp can prove this & the additive version @[to_additive] theorem singleton_div_singleton (a b : α) : ({a} : Finset α) / {b} = {a / b} := image₂_singleton @@ -1251,7 +1249,6 @@ theorem vsub_singleton (b : β) : s -ᵥ ({b} : Finset β) = s.image (· -ᵥ b) theorem singleton_vsub (a : β) : ({a} : Finset β) -ᵥ t = t.image (a -ᵥ ·) := image₂_singleton_left --- @[simp] -- Porting note (#10618): simp can prove this theorem singleton_vsub_singleton (a b : β) : ({a} : Finset β) -ᵥ {b} = {a -ᵥ b} := image₂_singleton diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 7dc3e11ce73a6..59cf3b949909d 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -329,7 +329,6 @@ theorem mul_singleton : s * {b} = (· * b) '' s := theorem singleton_mul : {a} * t = (a * ·) '' t := image2_singleton_left --- Porting note (#10618): simp can prove this @[to_additive] theorem singleton_mul_singleton : ({a} : Set α) * {b} = {a * b} := image2_singleton diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index e53212f0a17ad..9f971900d0349 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1923,7 +1923,6 @@ theorem normalClosure_eq_iInf : theorem normalClosure_eq_self (H : Subgroup G) [H.Normal] : normalClosure ↑H = H := le_antisymm (normalClosure_le_normal rfl.subset) le_normalClosure --- @[simp] -- Porting note (#10618): simp can prove this theorem normalClosure_idempotent : normalClosure ↑(normalClosure s) = normalClosure s := normalClosure_eq_self _ diff --git a/Mathlib/Algebra/Group/Submonoid/Basic.lean b/Mathlib/Algebra/Group/Submonoid/Basic.lean index 9113ca59629d9..04690c8d30132 100644 --- a/Mathlib/Algebra/Group/Submonoid/Basic.lean +++ b/Mathlib/Algebra/Group/Submonoid/Basic.lean @@ -266,7 +266,6 @@ theorem sup_eq_closure (N N' : Submonoid M) : N ⊔ N' = closure ((N : Set M) theorem closure_iUnion {ι} (s : ι → Set M) : closure (⋃ i, s i) = ⨆ i, closure (s i) := (Submonoid.gi M).gc.l_iSup --- Porting note (#10618): `simp` can now prove this, so we remove the `@[simp]` attribute @[to_additive] theorem closure_singleton_le_iff_mem (m : M) (p : Submonoid M) : closure {m} ≤ p ↔ m ∈ p := by rw [closure_le, singleton_subset_iff, SetLike.mem_coe] diff --git a/Mathlib/Algebra/GroupWithZero/Invertible.lean b/Mathlib/Algebra/GroupWithZero/Invertible.lean index 6895aebeb4888..73b809c2847b2 100644 --- a/Mathlib/Algebra/GroupWithZero/Invertible.lean +++ b/Mathlib/Algebra/GroupWithZero/Invertible.lean @@ -80,7 +80,6 @@ theorem div_self_of_invertible (a : α) [Invertible a] : a / a = 1 := def invertibleDiv (a b : α) [Invertible a] [Invertible b] : Invertible (a / b) := ⟨b / a, by simp [← mul_div_assoc], by simp [← mul_div_assoc]⟩ --- Porting note (#10618): removed `simp` attribute as `simp` can prove it theorem invOf_div (a b : α) [Invertible a] [Invertible b] [Invertible (a / b)] : ⅟ (a / b) = b / a := invOf_eq_right_inv (by simp [← mul_div_assoc]) diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index 441c75acfbc02..03b17997ea652 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -548,14 +548,14 @@ theorem next_eq (f : Hom C₁ C₂) {i j : ι} (w : c.Rel i j) : obtain rfl := c.next_eq' w simp only [xNextIso, eqToIso_refl, Iso.refl_hom, Iso.refl_inv, comp_id, id_comp] -@[reassoc, elementwise] -- @[simp] -- Porting note (#10618): simp can prove this +@[reassoc, elementwise] theorem comm_from (f : Hom C₁ C₂) (i : ι) : f.f i ≫ C₂.dFrom i = C₁.dFrom i ≫ f.next i := f.comm _ _ attribute [simp 1100] comm_from_assoc attribute [simp] comm_from_apply -@[reassoc, elementwise] -- @[simp] -- Porting note (#10618): simp can prove this +@[reassoc, elementwise] theorem comm_to (f : Hom C₁ C₂) (j : ι) : f.prev j ≫ C₂.dTo j = C₁.dTo j ≫ f.f j := f.comm _ _ diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 47b020f71c946..059efa2c28483 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -185,8 +185,6 @@ variable [Ring R] [AddCommGroup M] [Module R M] (r : R) (x : M) theorem neg_smul : -r • x = -(r • x) := eq_neg_of_add_eq_zero_left <| by rw [← add_smul, neg_add_cancel, zero_smul] --- Porting note (#10618): simp can prove this ---@[simp] theorem neg_smul_neg : -r • -x = r • x := by rw [neg_smul, smul_neg, neg_neg] @[simp] @@ -317,13 +315,9 @@ theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [Fun [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] --- Porting note (#10618): simp can prove this ---@[simp] theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by rw [nsmul_eq_mul, mul_one] --- Porting note (#10618): simp can prove this ---@[simp] theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by rw [zsmul_eq_mul, mul_one] diff --git a/Mathlib/Algebra/Module/LocalizedModule.lean b/Mathlib/Algebra/Module/LocalizedModule.lean index 231ab81c8c788..936fdf3e75b29 100644 --- a/Mathlib/Algebra/Module/LocalizedModule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule.lean @@ -1003,8 +1003,6 @@ theorem mk'_mul_mk' {M M' : Type*} [Semiring M] [Semiring M'] [Algebra R M] [Alg variable {f} -/-- Porting note (#10618): simp can prove this -@[simp] -/ theorem mk'_eq_iff {m : M} {s : S} {m' : M'} : mk' f m s = m' ↔ f m = s • m' := by rw [← smul_inj f s, Submonoid.smul_def, ← mk'_smul, ← Submonoid.smul_def, mk'_cancel] diff --git a/Mathlib/Algebra/Module/Submodule/Range.lean b/Mathlib/Algebra/Module/Submodule/Range.lean index 44306a514a638..7f70cb40186fa 100644 --- a/Mathlib/Algebra/Module/Submodule/Range.lean +++ b/Mathlib/Algebra/Module/Submodule/Range.lean @@ -266,7 +266,6 @@ theorem map_subtype_le (p' : Submodule R p) : map p.subtype p' ≤ p := by /-- Under the canonical linear map from a submodule `p` to the ambient space `M`, the image of the maximal submodule of `p` is just `p`. -/ --- @[simp] -- Porting note (#10618): simp can prove this theorem map_subtype_top : map p.subtype (⊤ : Submodule R p) = p := by simp @[simp] diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 5be854af327cb..096deecf0532a 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -631,7 +631,6 @@ variable (S : Type*) [CommMonoid S] [DistribMulAction S M] [SMulCommClass S R M] theorem mem_torsion'_iff (x : M) : x ∈ torsion' R M S ↔ ∃ a : S, a • x = 0 := Iff.rfl --- @[simp] Porting note (#10618): simp can prove this theorem mem_torsion_iff (x : M) : x ∈ torsion R M ↔ ∃ a : R⁰, a • x = 0 := Iff.rfl @@ -666,7 +665,6 @@ theorem torsion'_torsion'_eq_top : torsion' R (torsion' R M S) S = ⊤ := /-- The torsion submodule of the torsion submodule (viewed as a module) is the full torsion module. -/ --- @[simp] Porting note (#10618): simp can prove this theorem torsion_torsion_eq_top : torsion R (torsion R M) = ⊤ := torsion'_torsion'_eq_top R⁰ diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 2b71376c102f0..3766874adfc29 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -761,7 +761,6 @@ protected noncomputable def opRingEquiv [Monoid G] : theorem opRingEquiv_single [Monoid G] (r : k) (x : G) : MonoidAlgebra.opRingEquiv (op (single x r)) = single (op x) (op r) := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem opRingEquiv_symm_single [Monoid G] (r : kᵐᵒᵖ) (x : Gᵐᵒᵖ) : MonoidAlgebra.opRingEquiv.symm (single x r) = op (single x.unop r.unop) := by simp @@ -1421,7 +1420,6 @@ protected noncomputable def opRingEquiv [AddCommMonoid G] : theorem opRingEquiv_single [AddCommMonoid G] (r : k) (x : G) : AddMonoidAlgebra.opRingEquiv (op (single x r)) = single x (op r) := by simp --- @[simp] -- Porting note (#10618): simp can prove this theorem opRingEquiv_symm_single [AddCommMonoid G] (r : kᵐᵒᵖ) (x : Gᵐᵒᵖ) : AddMonoidAlgebra.opRingEquiv.symm (single x r) = op (single x r.unop) := by simp diff --git a/Mathlib/Algebra/MonoidAlgebra/Division.lean b/Mathlib/Algebra/MonoidAlgebra/Division.lean index 42c2c79cdba87..6ed612e48a0a5 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Division.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Division.lean @@ -129,7 +129,6 @@ theorem modOf_apply_of_exists_add (x : k[G]) (g : G) (g' : G) theorem modOf_apply_add_self (x : k[G]) (g : G) (d : G) : (x %ᵒᶠ g) (d + g) = 0 := modOf_apply_of_exists_add _ _ _ ⟨_, add_comm _ _⟩ --- @[simp] -- Porting note (#10618): simp can prove this theorem modOf_apply_self_add (x : k[G]) (g : G) (d : G) : (x %ᵒᶠ g) (g + d) = 0 := modOf_apply_of_exists_add _ _ _ ⟨_, rfl⟩ diff --git a/Mathlib/Algebra/MvPolynomial/Expand.lean b/Mathlib/Algebra/MvPolynomial/Expand.lean index 8b06a208a6ea6..1252dd9b13b66 100644 --- a/Mathlib/Algebra/MvPolynomial/Expand.lean +++ b/Mathlib/Algebra/MvPolynomial/Expand.lean @@ -29,7 +29,6 @@ noncomputable def expand (p : ℕ) : MvPolynomial σ R →ₐ[R] MvPolynomial σ { (eval₂Hom C fun i ↦ X i ^ p : MvPolynomial σ R →+* MvPolynomial σ R) with commutes' := fun _ ↦ eval₂Hom_C _ _ _ } --- @[simp] -- Porting note (#10618): simp can prove this theorem expand_C (p : ℕ) (r : R) : expand p (C r : MvPolynomial σ R) = C r := eval₂Hom_C _ _ _ diff --git a/Mathlib/Algebra/MvPolynomial/PDeriv.lean b/Mathlib/Algebra/MvPolynomial/PDeriv.lean index 2f9c43a1a8601..9c6e2c3c43e7c 100644 --- a/Mathlib/Algebra/MvPolynomial/PDeriv.lean +++ b/Mathlib/Algebra/MvPolynomial/PDeriv.lean @@ -104,7 +104,6 @@ theorem pderiv_pow {i : σ} {f : MvPolynomial σ R} {n : ℕ} : pderiv i (f ^ n) = n * f ^ (n - 1) * pderiv i f := by rw [(pderiv i).leibniz_pow f n, nsmul_eq_mul, smul_eq_mul, mul_assoc] --- @[simp] -- Porting note (#10618): simp can prove this theorem pderiv_C_mul {f : MvPolynomial σ R} {i : σ} : pderiv i (C a * f) = C a * pderiv i f := by rw [C_mul', Derivation.map_smul, C_mul'] diff --git a/Mathlib/Algebra/Order/CauSeq/Completion.lean b/Mathlib/Algebra/Order/CauSeq/Completion.lean index ff8a280c51ac8..501a26686963e 100644 --- a/Mathlib/Algebra/Order/CauSeq/Completion.lean +++ b/Mathlib/Algebra/Order/CauSeq/Completion.lean @@ -198,8 +198,6 @@ noncomputable instance : Inv (Cauchy abv) := rw [mk_eq.2 fg, ← Ig] at If rw [← mul_one (mk (inv f hf)), ← Ig', ← mul_assoc, If, mul_assoc, Ig', mul_one]⟩ --- porting note (#10618): simp can prove this --- @[simp] theorem inv_zero : (0 : (Cauchy abv))⁻¹ = 0 := congr_arg mk <| by rw [dif_pos] <;> [rfl; exact zero_limZero] diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index d4a66915b8a75..31a14a28943a9 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -216,7 +216,6 @@ theorem inducedMap_inducedMap (a : α) : inducedMap β γ (inducedMap α β a) = eq_of_forall_rat_lt_iff_lt fun q => by rw [coe_lt_inducedMap_iff, coe_lt_inducedMap_iff, Iff.comm, coe_lt_inducedMap_iff] ---@[simp] -- Porting note (#10618): simp can prove it theorem inducedMap_inv_self (b : β) : inducedMap γ β (inducedMap β γ b) = b := by rw [inducedMap_inducedMap, inducedMap_self] diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index da812710841b2..a05e2fffee157 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -251,8 +251,6 @@ theorem preimage_floor_of_ne_zero {n : ℕ} (hn : n ≠ 0) : theorem lt_ceil : n < ⌈a⌉₊ ↔ (n : α) < a := lt_iff_lt_of_le_iff_le ceil_le --- porting note (#10618): simp can prove this --- @[simp] theorem add_one_le_ceil_iff : n + 1 ≤ ⌈a⌉₊ ↔ (n : α) < a := by rw [← Nat.lt_ceil, Nat.add_one_le_iff] @@ -924,8 +922,6 @@ theorem fract_ofNat (n : ℕ) [n.AtLeastTwo] : fract ((no_index (OfNat.ofNat n)) : α) = 0 := fract_natCast n --- porting note (#10618): simp can prove this --- @[simp] theorem fract_floor (a : α) : fract (⌊a⌋ : α) = 0 := fract_intCast _ diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 95bd5c74cb92e..03c9d7579e5f1 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -148,7 +148,6 @@ theorem mul_inv_le_iff_le_mul : a * b⁻¹ ≤ c ↔ a ≤ c * b := theorem le_mul_inv_iff_mul_le : c ≤ a * b⁻¹ ↔ c * b ≤ a := (mul_le_mul_iff_right b).symm.trans <| by rw [inv_mul_cancel_right] --- Porting note (#10618): `simp` can prove this @[to_additive] theorem mul_inv_le_one_iff_le : a * b⁻¹ ≤ 1 ↔ a ≤ b := mul_inv_le_iff_le_mul.trans <| by rw [one_mul] @@ -193,7 +192,6 @@ theorem mul_inv_lt_iff_lt_mul : a * b⁻¹ < c ↔ a < c * b := by theorem lt_mul_inv_iff_mul_lt : c < a * b⁻¹ ↔ c * b < a := (mul_lt_mul_iff_right b).symm.trans <| by rw [inv_mul_cancel_right] --- Porting note (#10618): `simp` can prove this @[to_additive] theorem inv_mul_lt_one_iff_lt : a * b⁻¹ < 1 ↔ a < b := by rw [← mul_lt_mul_iff_right b, inv_mul_cancel_right, one_mul] diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index b0a55e98d4250..4001156b0e699 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -352,11 +352,9 @@ variable [LinearOrder α] {a b c : α} instance linearOrder : LinearOrder (WithZero α) := WithBot.linearOrder --- Porting note (#10618): @[simp] can prove this protected lemma le_max_iff : (a : WithZero α) ≤ max (b : WithZero α) c ↔ a ≤ max b c := by simp only [WithZero.coe_le_coe, le_max_iff] --- Porting note (#10618): @[simp] can prove this protected lemma min_le_iff : min (a : WithZero α) b ≤ c ↔ min a b ≤ c := by simp only [WithZero.coe_le_coe, min_le_iff] diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index f3c8c9e9daa85..b18b3c81d456e 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -226,12 +226,10 @@ theorem min_mul_distrib (a b c : α) : min a (b * c) = min a (min a b * min a c) theorem min_mul_distrib' (a b c : α) : min (a * b) c = min (min a c * min b c) c := by simpa [min_comm _ c] using min_mul_distrib c a b --- Porting note (#10618): no longer `@[simp]`, as `simp` can prove this. @[to_additive] theorem one_min (a : α) : min 1 a = 1 := min_eq_left (one_le a) --- Porting note (#10618): no longer `@[simp]`, as `simp` can prove this. @[to_additive] theorem min_one (a : α) : min a 1 = 1 := min_eq_right (one_le a) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean index f23cb20dc1b90..cd7326219233b 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/WithTop.lean @@ -113,12 +113,8 @@ theorem add_eq_coe : | some a, ⊤, c => by simp | some a, some b, c => by norm_cast; simp --- Porting note (#10618): simp can already prove this. --- @[simp] theorem add_coe_eq_top_iff {x : WithTop α} {y : α} : x + y = ⊤ ↔ x = ⊤ := by simp --- Porting note (#10618): simp can already prove this. --- @[simp] theorem coe_add_eq_top_iff {y : WithTop α} : ↑x + y = ⊤ ↔ y = ⊤ := by simp theorem add_right_cancel_iff [IsRightCancelAdd α] (ha : a ≠ ⊤) : b + a = c + a ↔ b = c := by @@ -542,13 +538,9 @@ theorem bot_lt_add [LT α] {a b : WithBot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ theorem add_eq_coe : a + b = x ↔ ∃ a' b' : α, ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := WithTop.add_eq_coe --- Porting note (#10618): simp can already prove this. --- @[simp] theorem add_coe_eq_bot_iff : a + y = ⊥ ↔ a = ⊥ := WithTop.add_coe_eq_top_iff --- Porting note (#10618): simp can already prove this. --- @[simp] theorem coe_add_eq_bot_iff : ↑x + b = ⊥ ↔ b = ⊥ := WithTop.coe_add_eq_top_iff diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index f095ae814ba7d..57f4bba5519be 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -455,7 +455,6 @@ theorem smul_C {S} [SMulZeroClass S R] (s : S) (r : R) : s • C r = C (s • r) theorem C_pow : C (a ^ n) = C a ^ n := C.map_pow a n --- @[simp] -- Porting note (#10618): simp can prove this theorem C_eq_natCast (n : ℕ) : C (n : R) = (n : R[X]) := map_natCast C n diff --git a/Mathlib/Algebra/Polynomial/Coeff.lean b/Mathlib/Algebra/Polynomial/Coeff.lean index 82ae37b47a959..383ce0d93d741 100644 --- a/Mathlib/Algebra/Polynomial/Coeff.lean +++ b/Mathlib/Algebra/Polynomial/Coeff.lean @@ -346,7 +346,7 @@ theorem natCast_coeff_zero {n : ℕ} {R : Type*} [Semiring R] : (n : R[X]).coeff @[deprecated (since := "2024-04-17")] alias nat_cast_coeff_zero := natCast_coeff_zero -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem natCast_inj {m n : ℕ} {R : Type*} [Semiring R] [CharZero R] : (↑m : R[X]) = ↑n ↔ m = n := by constructor @@ -366,7 +366,7 @@ theorem intCast_coeff_zero {i : ℤ} {R : Type*} [Ring R] : (i : R[X]).coeff 0 = @[deprecated (since := "2024-04-17")] alias int_cast_coeff_zero := intCast_coeff_zero -@[norm_cast] -- @[simp] -- Porting note (#10618): simp can prove this +@[norm_cast] theorem intCast_inj {m n : ℤ} {R : Type*} [Ring R] [CharZero R] : (↑m : R[X]) = ↑n ↔ m = n := by constructor · intro h diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 9ec69895a9634..e5cc71815055f 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -758,11 +758,9 @@ theorem leadingCoeff_C_mul_X (a : R) : leadingCoeff (C a * X) = a := by theorem leadingCoeff_C (a : R) : leadingCoeff (C a) = a := leadingCoeff_monomial a 0 --- @[simp] -- Porting note (#10618): simp can prove this theorem leadingCoeff_X_pow (n : ℕ) : leadingCoeff ((X : R[X]) ^ n) = 1 := by simpa only [C_1, one_mul] using leadingCoeff_C_mul_X_pow (1 : R) n --- @[simp] -- Porting note (#10618): simp can prove this theorem leadingCoeff_X : leadingCoeff (X : R[X]) = 1 := by simpa only [pow_one] using @leadingCoeff_X_pow R _ 1 @@ -774,7 +772,6 @@ theorem monic_X_pow (n : ℕ) : Monic (X ^ n : R[X]) := theorem monic_X : Monic (X : R[X]) := leadingCoeff_X --- @[simp] -- Porting note (#10618): simp can prove this theorem leadingCoeff_one : leadingCoeff (1 : R[X]) = 1 := leadingCoeff_C 1 diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index 5e8f7007d1c5b..899b80f020a97 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -98,7 +98,6 @@ theorem derivative_C_mul_X_sq (a : R) : derivative (C a * X ^ 2) = C (a * 2) * X theorem derivative_X_pow (n : ℕ) : derivative (X ^ n : R[X]) = C (n : R) * X ^ (n - 1) := by convert derivative_C_mul_X_pow (1 : R) n <;> simp --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_X_sq : derivative (X ^ 2 : R[X]) = C 2 * X := by rw [derivative_X_pow, Nat.cast_two, pow_one] @@ -120,16 +119,13 @@ theorem derivative_one : derivative (1 : R[X]) = 0 := theorem derivative_add {f g : R[X]} : derivative (f + g) = derivative f + derivative g := derivative.map_add f g --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_X_add_C (c : R) : derivative (X + C c) = 1 := by rw [derivative_add, derivative_X, derivative_C, add_zero] --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_sum {s : Finset ι} {f : ι → R[X]} : derivative (∑ b ∈ s, f b) = ∑ b ∈ s, derivative (f b) := map_sum .. --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_smul {S : Type*} [Monoid S] [DistribMulAction S R] [IsScalarTower S R R] (s : S) (p : R[X]) : derivative (s • p) = s • derivative p := derivative.map_smul_of_tower s p @@ -571,7 +567,6 @@ theorem iterate_derivative_neg {f : R[X]} {k : ℕ} : derivative^[k] (-f) = -der theorem derivative_sub {f g : R[X]} : derivative (f - g) = derivative f - derivative g := LinearMap.map_sub derivative f g --- Porting note (#10618): removed `simp`: `simp` can prove it. theorem derivative_X_sub_C (c : R) : derivative (X - C c) = 1 := by rw [derivative_sub, derivative_X, derivative_C, sub_zero] diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index 7210a0ec6104f..52e9ac7978fee 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -218,7 +218,6 @@ theorem _root_.Polynomial.toLaurent_C_mul_eq (r : R) (f : R[X]) : theorem _root_.Polynomial.toLaurent_X_pow (n : ℕ) : toLaurent (X ^ n : R[X]) = T n := by simp only [map_pow, Polynomial.toLaurent_X, T_pow, mul_one] --- @[simp] -- Porting note (#10618): simp can prove this theorem _root_.Polynomial.toLaurent_C_mul_X_pow (n : ℕ) (r : R) : toLaurent (Polynomial.C r * X ^ n) = C r * T n := by simp only [_root_.map_mul, Polynomial.toLaurent_C, Polynomial.toLaurent_X_pow] diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 614061ef26c6d..2d9684409878b 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -78,7 +78,6 @@ theorem revAt_add {N O n o : ℕ} (hn : n ≤ N) (ho : o ≤ O) : rw [add_assoc, add_left_comm n' o, ← add_assoc, revAt_le (le_add_right rfl.le)] repeat' rw [add_tsub_cancel_left] --- @[simp] -- Porting note (#10618): simp can prove this theorem revAt_zero (N : ℕ) : revAt N 0 = N := by simp /-- `reflect N f` is the polynomial such that `(reflect N f).coeff i = f.coeff (revAt N i)`. @@ -123,7 +122,6 @@ theorem reflect_C_mul (f : R[X]) (r : R) (N : ℕ) : reflect N (C r * f) = C r * ext simp only [coeff_reflect, coeff_C_mul] --- @[simp] -- Porting note (#10618): simp can prove this (once `reflect_monomial` is in simp scope) theorem reflect_C_mul_X_pow (N n : ℕ) {c : R} : reflect N (C c * X ^ n) = C c * X ^ revAt N n := by ext rw [reflect_C_mul, coeff_C_mul, coeff_C_mul, coeff_X_pow, coeff_reflect] diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 05c61e0d4e530..6fe9d273656aa 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1282,15 +1282,12 @@ instance instDivisionRing : DivisionRing ℍ[R] where nnqsmul_def q x := by rw [← coe_nnratCast, coe_mul_eq_smul]; ext <;> exact NNRat.smul_def _ _ qsmul_def q x := by rw [← coe_ratCast, coe_mul_eq_smul]; ext <;> exact Rat.smul_def _ _ ---@[simp] Porting note (#10618): `simp` can prove it theorem normSq_inv : normSq a⁻¹ = (normSq a)⁻¹ := map_inv₀ normSq _ ---@[simp] Porting note (#10618): `simp` can prove it theorem normSq_div : normSq (a / b) = normSq a / normSq b := map_div₀ normSq a b ---@[simp] Porting note (#10618): `simp` can prove it theorem normSq_zpow (z : ℤ) : normSq (a ^ z) = normSq a ^ z := map_zpow₀ normSq a z @@ -1330,7 +1327,6 @@ theorem mk_quaternionAlgebra_of_infinite [Infinite R] : #(ℍ[R,c₁,c₂]) = #R theorem mk_univ_quaternionAlgebra : #(Set.univ : Set ℍ[R,c₁,c₂]) = #R ^ 4 := by rw [mk_univ, mk_quaternionAlgebra] ---@[simp] Porting note (#10618): `simp` can prove it theorem mk_univ_quaternionAlgebra_of_infinite [Infinite R] : #(Set.univ : Set ℍ[R,c₁,c₂]) = #R := by rw [mk_univ_quaternionAlgebra, pow_four] @@ -1359,7 +1355,6 @@ theorem mk_quaternion_of_infinite [Infinite R] : #(ℍ[R]) = #R := mk_quaternionAlgebra_of_infinite _ _ /-- The cardinality of the quaternions, as a set. -/ ---@[simp] Porting note (#10618): `simp` can prove it theorem mk_univ_quaternion : #(Set.univ : Set ℍ[R]) = #R ^ 4 := mk_univ_quaternionAlgebra _ _ diff --git a/Mathlib/Algebra/Ring/BooleanRing.lean b/Mathlib/Algebra/Ring/BooleanRing.lean index 985d2359dfa85..e88f282f73822 100644 --- a/Mathlib/Algebra/Ring/BooleanRing.lean +++ b/Mathlib/Algebra/Ring/BooleanRing.lean @@ -137,11 +137,9 @@ theorem toBoolAlg_ofBoolAlg (a : AsBoolAlg α) : toBoolAlg (ofBoolAlg a) = a := theorem ofBoolAlg_toBoolAlg (a : α) : ofBoolAlg (toBoolAlg a) = a := rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem toBoolAlg_inj {a b : α} : toBoolAlg a = toBoolAlg b ↔ a = b := Iff.rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem ofBoolAlg_inj {a b : AsBoolAlg α} : ofBoolAlg a = ofBoolAlg b ↔ a = b := Iff.rfl @@ -355,11 +353,9 @@ theorem toBoolRing_ofBoolRing (a : AsBoolRing α) : toBoolRing (ofBoolRing a) = theorem ofBoolRing_toBoolRing (a : α) : ofBoolRing (toBoolRing a) = a := rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem toBoolRing_inj {a b : α} : toBoolRing a = toBoolRing b ↔ a = b := Iff.rfl --- Porting note (#10618): simp can prove this -- @[simp] theorem ofBoolRing_inj {a b : AsBoolRing α} : ofBoolRing a = ofBoolRing b ↔ a = b := Iff.rfl diff --git a/Mathlib/Algebra/Ring/Commute.lean b/Mathlib/Algebra/Ring/Commute.lean index 27887a07dc5bb..381b7109ee286 100644 --- a/Mathlib/Algebra/Ring/Commute.lean +++ b/Mathlib/Algebra/Ring/Commute.lean @@ -77,13 +77,9 @@ section variable [MulOneClass R] [HasDistribNeg R] --- Porting note (#10618): no longer needs to be `@[simp]` since `simp` can prove it. --- @[simp] theorem neg_one_right (a : R) : Commute a (-1) := SemiconjBy.neg_one_right a --- Porting note (#10618): no longer needs to be `@[simp]` since `simp` can prove it. --- @[simp] theorem neg_one_left (a : R) : Commute (-1) a := SemiconjBy.neg_one_left a diff --git a/Mathlib/Algebra/Ring/CompTypeclasses.lean b/Mathlib/Algebra/Ring/CompTypeclasses.lean index 7ab747f493635..6aff4dd2a3908 100644 --- a/Mathlib/Algebra/Ring/CompTypeclasses.lean +++ b/Mathlib/Algebra/Ring/CompTypeclasses.lean @@ -81,22 +81,16 @@ class RingHomInvPair (σ : R₁ →+* R₂) (σ' : outParam (R₂ →+* R₁)) : /-- `σ'` is a left inverse of `σ'` -/ comp_eq₂ : σ.comp σ' = RingHom.id R₂ --- attribute [simp] RingHomInvPair.comp_eq Porting note (#10618): `simp` can prove it - --- attribute [simp] RingHomInvPair.comp_eq₂ Porting note (#10618): `simp` can prove it - variable {σ : R₁ →+* R₂} {σ' : R₂ →+* R₁} namespace RingHomInvPair variable [RingHomInvPair σ σ'] --- @[simp] Porting note (#10618): `simp` can prove it theorem comp_apply_eq {x : R₁} : σ' (σ x) = x := by rw [← RingHom.comp_apply, comp_eq] simp --- @[simp] Porting note (#10618): `simp` can prove it theorem comp_apply_eq₂ {x : R₂} : σ (σ' x) = x := by rw [← RingHom.comp_apply, comp_eq₂] simp diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 8518ce1badcb4..ff0a1107f04c7 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -733,12 +733,10 @@ theorem toMonoidHom_refl : (RingEquiv.refl R).toMonoidHom = MonoidHom.id R := theorem toAddMonoidHom_refl : (RingEquiv.refl R).toAddMonoidHom = AddMonoidHom.id R := rfl --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem toRingHom_apply_symm_toRingHom_apply (e : R ≃+* S) : ∀ y : S, e.toRingHom (e.symm.toRingHom y) = y := e.toEquiv.apply_symm_apply --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem symm_toRingHom_apply_toRingHom_apply (e : R ≃+* S) : ∀ x : R, e.symm.toRingHom (e.toRingHom x) = x := Equiv.symm_apply_apply e.toEquiv @@ -748,13 +746,11 @@ theorem toRingHom_trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : (e₁.trans e₂).toRingHom = e₂.toRingHom.comp e₁.toRingHom := rfl --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem toRingHom_comp_symm_toRingHom (e : R ≃+* S) : e.toRingHom.comp e.symm.toRingHom = RingHom.id _ := by ext simp --- Porting note (#10618): Now other `simp` can do this, so removed `simp` attribute theorem symm_toRingHom_comp_toRingHom (e : R ≃+* S) : e.symm.toRingHom.comp e.toRingHom = RingHom.id _ := by ext diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index b4890ab73e4bc..72fdec393def4 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -168,11 +168,6 @@ end Monoid section Ring variable [Ring α] {a b : α} {n : ℕ} -/- Porting note (#10618): attribute `simp` removed based on linter report -simp can prove this: - by simp only [even_neg, even_two] --/ --- @[simp] lemma even_neg_two : Even (-2 : α) := by simp only [even_neg, even_two] lemma Odd.neg (hp : Odd a) : Odd (-a) := by @@ -183,11 +178,6 @@ lemma Odd.neg (hp : Odd a) : Odd (-a) := by @[simp] lemma odd_neg : Odd (-a) ↔ Odd a := ⟨fun h ↦ neg_neg a ▸ h.neg, Odd.neg⟩ -/- Porting note (#10618): attribute `simp` removed based on linter report -simp can prove this: - by simp only [odd_neg, odd_one] --/ --- @[simp] lemma odd_neg_one : Odd (-1 : α) := by simp lemma Odd.sub_even (ha : Odd a) (hb : Even b) : Odd (a - b) := by diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index e16be89dd2957..35a508fb82fb3 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -358,8 +358,6 @@ theorem coe_one : ((1 : s) : R) = 1 := theorem coe_pow (x : s) (n : ℕ) : ↑(x ^ n) = (x : R) ^ n := SubmonoidClass.coe_pow x n --- TODO: can be generalized to `AddSubmonoidClass` --- @[simp] -- Porting note (#10618): simp can prove this theorem coe_eq_zero_iff {x : s} : (x : R) = 0 ↔ x = 0 := ⟨fun h => Subtype.ext (Trans.trans h s.coe_zero.symm), fun h => h.symm ▸ s.coe_zero⟩ @@ -1100,11 +1098,9 @@ def inclusion {S T : Subring R} (h : S ≤ T) : S →+* T := theorem range_subtype (s : Subring R) : s.subtype.range = s := SetLike.coe_injective <| (coe_rangeS _).trans Subtype.range_coe --- @[simp] -- Porting note (#10618): simp can prove this theorem range_fst : (fst R S).rangeS = ⊤ := (fst R S).rangeS_top_of_surjective <| Prod.fst_surjective --- @[simp] -- Porting note (#10618): simp can prove this theorem range_snd : (snd R S).rangeS = ⊤ := (snd R S).rangeS_top_of_surjective <| Prod.snd_surjective diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index a67f00afb054d..589f27314eac8 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -152,7 +152,6 @@ instance : SubsemiringClass (Subsemiring R) R where theorem mem_toSubmonoid {s : Subsemiring R} {x : R} : x ∈ s.toSubmonoid ↔ x ∈ s := Iff.rfl --- `@[simp]` -- Porting note (#10618): simp can prove thisrove this theorem mem_carrier {s : Subsemiring R} {x : R} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 935cf00ec3b26..c18ad3b0214ad 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -42,7 +42,6 @@ theorem IsRelPrime.of_squarefree_mul [CommMonoid R] {m n : R} (h : Squarefree (m theorem IsUnit.squarefree [CommMonoid R] {x : R} (h : IsUnit x) : Squarefree x := fun _ hdvd => isUnit_of_mul_isUnit_left (isUnit_of_dvd_unit hdvd h) --- @[simp] -- Porting note (#10618): simp can prove this theorem squarefree_one [CommMonoid R] : Squarefree (1 : R) := isUnit_one.squarefree diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index f4dea0e4e8345..66841a63dca0e 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -95,11 +95,9 @@ theorem unsym_injective : Injective (unsym : αˢʸᵐ → α) := theorem unsym_surjective : Surjective (unsym : αˢʸᵐ → α) := unsym.surjective --- Porting note (#10618): @[simp] can prove this theorem sym_inj {a b : α} : sym a = sym b ↔ a = b := sym_injective.eq_iff --- Porting note (#10618): @[simp] can prove this theorem unsym_inj {a b : αˢʸᵐ} : unsym a = unsym b ↔ a = b := unsym_injective.eq_iff diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index 373542c05d885..5a22f3bbfe005 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -291,7 +291,6 @@ theorem add_eq_left_iff {x y : Tropical R} : x + y = x ↔ x ≤ y := by theorem add_eq_right_iff {x y : Tropical R} : x + y = y ↔ y ≤ x := by rw [trop_add_def, trop_eq_iff_eq_untrop, ← untrop_le_iff, min_eq_right_iff] --- Porting note (#10618): removing `simp`. `simp` can prove it theorem add_self (x : Tropical R) : x + x = x := untrop_injective (min_eq_right le_rfl) @@ -495,7 +494,6 @@ theorem succ_nsmul {R} [LinearOrder R] [OrderTop R] (x : Tropical R) (n : ℕ) : -- Requires `zero_eq_bot` to be true -- lemma add_eq_zero_iff {a b : tropical R} : -- a + b = 1 ↔ a = 1 ∨ b = 1 := sorry --- Porting note (#10618): removing @[simp], `simp` can prove it theorem mul_eq_zero_iff {R : Type*} [LinearOrderedAddCommMonoid R] {a b : Tropical (WithTop R)} : a * b = 0 ↔ a = 0 ∨ b = 0 := by simp [← untrop_inj_iff, WithTop.add_eq_top] From fab5497a2d3ee1fd2526d57093abedf3ffc12f54 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 25 Oct 2024 09:18:09 +0000 Subject: [PATCH 10/29] feat: `Deprecated` is the new `Init` for tech-debts (#18090) There are no more `Init` files, so this PR * removes `Init` counts from the `tech-debt` counters; * uses the same fine-grained approach that `Init` had for `Deprecated`. I also wrote a better script to compute differences that should address the issues that the previous code had. I simply removed the old test, rather than muting it, since I do not imagine anyone adding new `Init` files. --- scripts/technical-debt-metrics.sh | 78 +++++++++++++++++-------------- 1 file changed, 44 insertions(+), 34 deletions(-) diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index ce9a3a5202b16..524c97f135331 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -15,6 +15,28 @@ IFS=$'\n\t' currCommit="${1:-"$(git rev-parse HEAD)"}" refCommit="${2:-"$(git log --pretty=%H --since="$(date -I -d 'last week')" | tail -n -1)"}" +## `computeDiff input` assumes that input consists of lines of the form `value|description` +## where `value` is a number and `description` is the statistic that `value` reports. +## `value` is non-negative, if it refers to the current commit and it is negative otherwise. +## The output is of the form `|current value|difference|description|`. +computeDiff () { + awk -F'|' 'BEGIN{con=1}{ + # order keeps track of maintaining the final order of the counters the same as the input one + # rdict stores the difference of the current value minus the old value + # curr stores the current value, making sure that the number is non-negative and does not start with `-` + if(rdict[$2] == "") {order[con]=$2; con++} + if((0 <= $1+0) && (!($1 ~ "-"))) {curr[$2]=$1} + rdict[$2]+=$1+0 + } END { + # loop over the "sorted" index in `order` + for(ind=1; ind<=con-1; ind++) { + # retrieve the description of the counter + val=order[ind] + # print `|current value|difference|name of counter|` + printf("|%s|%s|%s|\n", curr[val], rdict[val], val)} + }' "${1}" +} + # `tdc` produces a semi-formatted output of the form # ... # |description @@ -34,7 +56,6 @@ titlesAndRegexes=( "maxHeartBeats modifications" "^ *set_option .*maxHeartbeats" ) -printf '|Current number|Change|Type|\n|-:|:-:|:-|\n' for i in ${!titlesAndRegexes[@]}; do # loop on the odd-indexed entries and name each entry and the following if (( (i + 1) % 2 )); then @@ -58,43 +79,32 @@ printf '%s|%s\n' "$(wc -l < scripts/no_lints_prime_decls.txt)" "exceptions for t deprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" printf '%s|%s\n' "$(printf '%s' "${deprecatedFiles}" | wc -l)" "\`Deprecated\` files" -printf '%s|%s\n' "$(printf '%s\n' "${deprecatedFiles}" | grep total | sed 's= total==')" 'total LoC in `Deprecated` files' - -initFiles="$(git ls-files '**/Init/*.lean' | xargs wc -l | sed 's=^ *==')" - -printf '%s|%s\n' "$(printf '%s' "${initFiles}" | wc -l)" "\`Init\` files" -printf '%s|%s\n\n' "$(printf '%s\n' "${initFiles}" | grep total | sed 's= total==')" 'total LoC in `Init` files' - -printf $'```spoiler Changed \'Init\' lines by file\n%s\n```\n' "$( - printf '%s\n' "${initFiles}" | awk 'BEGIN{print("|LoC|Change|File|\n|-:|:-:|-|")} {printf("%s|%s\n", $1, $2)}' - )" +printf '%s|%s\n\n' "$(printf '%s\n' "${deprecatedFiles}" | grep total | sed 's= total==')" 'total LoC in `Deprecated` files' } -# collect the technical debts from the current mathlib -new="$(git checkout -q "${currCommit}" && tdc; git switch -q -)" +# collect the technical debts and the line counts of the deprecated file from the current mathlib +git checkout -q "${currCommit}" +new="$(tdc)" +newDeprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" +git switch -q - -# collect the technical debts from the reference mathlib -- switch to the -old="$(git checkout -q "${refCommit}" && tdc; git switch -q -)" +# collect the technical debts and the line counts of the deprecated file from the reference mathlib +git checkout -q "${refCommit}" +old="$(tdc | sed 's=^[0-9]=-&=')" +oldDeprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *=-=')" +git switch -q - -# place the outputs side-by-side, using `@` as a separator -paste -d@ <(echo "$new") <(echo "${old}") | sed 's=@=|@|=' | - # each line should look like eithe - # [new number]|description@[old number]|descr - # or something that does not start with [number]| - # we split the lines into "words", using `|` as a separator - awk -F'|' ' - # if the first "word" is a number, then we write the 1st entry and compare it with the 4th - ($1+0 == $1) { printf("|%s|%s|%s|\n", $1, $1-$4, $2) } - # otherwise, the line is a "formatting" line, so we simply print it unchanged until we find `@` - !($1+0 == $1) { - for(i=1; i<=NF; i++) { - if ($i == "@") { break } - else { printf("%s|", $i) } - } - print "@" - }' | - # the sequence `@|` ending a line is an artifact of our process and we remove it - sed 's=|@$==' +printf '|Current number|Change|Type|\n|-:|:-:|:-|\n' +printf '%s\n%s\n' "${old}" "${new}" | computeDiff - +deprSummary="$(printf '%s\n%s\n' "${oldDeprecatedFiles}" "${newDeprecatedFiles}" | tr ' ' '|' | computeDiff -)" + +printf $'```spoiler Changed \'Deprecated\' lines by file\n%s\n```\n' "$( + printf '%s\n' "${deprSummary}" | awk -F'|' 'BEGIN{print("|LoC|Change|File|\n|-:|:-:|-|")} + ($4 == "total") {total=$0} + (!($4 == "total")) { + printf("%s\n", $0) + } END {printf("%s\n", total)}' + )" baseURL='https://github.com/leanprover-community/mathlib4/commit' printf '\nCurrent commit [%s](%s)\n' "${currCommit:0:10}" "${baseURL}/${currCommit}" From 8e419e3d18cb1b1aaf91069a76545bd3269ee69b Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 25 Oct 2024 09:55:06 +0000 Subject: [PATCH 11/29] feat (Data/Finset/Basic) : Equivalence between Pi type over Finset.cons and a product (#17004) This PR defines an equivalence between Pi over an augmented finset and a product with Pi over the original finset. --- Mathlib/Data/Finset/Basic.lean | 70 ++++++++++++++++++- .../GroupTheory/SpecificGroups/KleinFour.lean | 4 +- Mathlib/Order/CountableDenseLinearOrder.lean | 6 +- 3 files changed, 73 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 7f0708bdd080c..e3a1af524b590 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -751,6 +751,13 @@ theorem mem_cons_self (a : α) (s : Finset α) {h} : a ∈ cons a s h := theorem cons_val (h : a ∉ s) : (cons a s h).1 = a ::ₘ s.1 := rfl +theorem eq_of_mem_cons_of_not_mem (has : a ∉ s) (h : b ∈ cons a s has) (hb : b ∉ s) : b = a := + (mem_cons.1 h).resolve_right hb + +theorem mem_of_mem_cons_of_ne {s : Finset α} {a : α} {has} {i : α} + (hi : i ∈ cons a s has) (hia : i ≠ a) : i ∈ s := + (mem_cons.1 hi).resolve_left hia + theorem forall_mem_cons (h : a ∉ s) (p : α → Prop) : (∀ x, x ∈ cons a s h → p x) ↔ p a ∧ ∀ x, x ∈ s → p x := by simp only [mem_cons, or_imp, forall_and, forall_eq] @@ -808,6 +815,36 @@ theorem cons_swap (hb : b ∉ s) (ha : a ∉ s.cons b hb) : ha (mem_cons.mpr (.inl ((mem_cons.mp h).elim symm (fun h ↦ False.elim (hb h))))) := eq_of_veq <| Multiset.cons_swap a b s.val +/-- Split the added element of cons off a Pi type. -/ +@[simps!] +def consPiProd (f : α → Type*) (has : a ∉ s) (x : Π i ∈ cons a s has, f i) : f a × Π i ∈ s, f i := + (x a (mem_cons_self a s), fun i hi => x i (mem_cons_of_mem hi)) + +/-- Combine a product with a pi type to pi of cons. -/ +def prodPiCons [DecidableEq α] (f : α → Type*) {a : α} (has : a ∉ s) (x : f a × Π i ∈ s, f i) : + (Π i ∈ cons a s has, f i) := + fun i hi => + if h : i = a then cast (congrArg f h.symm) x.1 else x.2 i (mem_of_mem_cons_of_ne hi h) + +/-- The equivalence between pi types on cons and the product. -/ +def consPiProdEquiv [DecidableEq α] {s : Finset α} (f : α → Type*) {a : α} (has : a ∉ s) : + (Π i ∈ cons a s has, f i) ≃ f a × Π i ∈ s, f i where + toFun := consPiProd f has + invFun := prodPiCons f has + left_inv _ := by + ext i _ + dsimp only [prodPiCons, consPiProd] + by_cases h : i = a + · rw [dif_pos h] + subst h + simp_all only [cast_eq] + · rw [dif_neg h] + right_inv _ := by + ext _ hi + · simp [prodPiCons] + · simp only [consPiProd_snd] + exact dif_neg (ne_of_mem_of_not_mem hi has) + end Cons /-! ### disjoint -/ @@ -955,7 +992,7 @@ theorem mem_insert_of_mem (h : a ∈ s) : a ∈ insert b s := theorem mem_of_mem_insert_of_ne (h : b ∈ insert a s) : b ≠ a → b ∈ s := (mem_insert.1 h).resolve_left -theorem eq_of_not_mem_of_mem_insert (ha : b ∈ insert a s) (hb : b ∉ s) : b = a := +theorem eq_of_mem_insert_of_not_mem (ha : b ∈ insert a s) (hb : b ∉ s) : b = a := (mem_insert.1 ha).resolve_right hb /-- A version of `LawfulSingleton.insert_emptyc_eq` that works with `dsimp`. -/ @@ -1039,7 +1076,7 @@ theorem insert_subset_insert (a : α) {s t : Finset α} (h : s ⊆ t) : insert a simp_rw [← coe_subset]; simp [-coe_subset, ha] theorem insert_inj (ha : a ∉ s) : insert a s = insert b s ↔ a = b := - ⟨fun h => eq_of_not_mem_of_mem_insert (h ▸ mem_insert_self _ _) ha, congr_arg (insert · s)⟩ + ⟨fun h => eq_of_mem_insert_of_not_mem (h ▸ mem_insert_self _ _) ha, congr_arg (insert · s)⟩ theorem insert_inj_on (s : Finset α) : Set.InjOn (fun a => insert a s) sᶜ := fun _ h _ _ => (insert_inj h).1 @@ -1137,6 +1174,35 @@ theorem disjoint_insert_left : Disjoint (insert a s) t ↔ a ∉ t ∧ Disjoint theorem disjoint_insert_right : Disjoint s (insert a t) ↔ a ∉ s ∧ Disjoint s t := disjoint_comm.trans <| by rw [disjoint_insert_left, _root_.disjoint_comm] +/-- Split the added element of insert off a Pi type. -/ +@[simps!] +def insertPiProd (f : α → Type*) (x : Π i ∈ insert a s, f i) : f a × Π i ∈ s, f i := + (x a (mem_insert_self a s), fun i hi => x i (mem_insert_of_mem hi)) + +/-- Combine a product with a pi type to pi of insert. -/ +def prodPiInsert (f : α → Type*) {a : α} (x : f a × Π i ∈ s, f i) : (Π i ∈ insert a s, f i) := + fun i hi => + if h : i = a then cast (congrArg f h.symm) x.1 else x.2 i (mem_of_mem_insert_of_ne hi h) + +/-- The equivalence between pi types on insert and the product. -/ +def insertPiProdEquiv [DecidableEq α] {s : Finset α} (f : α → Type*) {a : α} (has : a ∉ s) : + (Π i ∈ insert a s, f i) ≃ f a × Π i ∈ s, f i where + toFun := insertPiProd f + invFun := prodPiInsert f + left_inv _ := by + ext i _ + dsimp only [prodPiInsert, insertPiProd] + by_cases h : i = a + · rw [dif_pos h] + subst h + simp_all only [cast_eq] + · rw [dif_neg h] + right_inv _ := by + ext _ hi + · simp [prodPiInsert] + · simp only [insertPiProd_snd] + exact dif_neg (ne_of_mem_of_not_mem hi has) + end Insert /-! ### Lattice structure -/ diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index b902f89c46909..e5a725fd32218 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -120,7 +120,7 @@ lemma eq_mul_of_ne_all {x y z : G} (hx : x ≠ 1) (hy : y ≠ 1) (hxy : x ≠ y) (hz : z ≠ 1) (hzx : z ≠ x) (hzy : z ≠ y) : z = x * y := by classical let _ := Fintype.ofFinite G - apply eq_of_not_mem_of_mem_insert <| (eq_finset_univ hx hy hxy).symm ▸ mem_univ _ + apply eq_of_mem_insert_of_not_mem <| (eq_finset_univ hx hy hxy).symm ▸ mem_univ _ simpa only [mem_singleton, mem_insert, not_or] using ⟨hzx, hzy, hz⟩ variable {G₁ G₂ : Type*} [Group G₁] [Group G₂] [IsKleinFour G₁] @@ -146,7 +146,7 @@ def mulEquiv' (e : G₁ ≃ G₂) (he : e 1 = 1) (h : Monoid.exponent G₂ = 2) rw [← Ne, ← e.injective.ne_iff] at hx hy hxy rw [he] at hx hy symm - apply eq_of_not_mem_of_mem_insert <| univ₂.symm ▸ mem_univ _ + apply eq_of_mem_insert_of_not_mem <| univ₂.symm ▸ mem_univ _ simpa using mul_not_mem_of_exponent_two h hx hy hxy /-- Any two `IsKleinFour` groups are isomorphic via any equivalence which sends the identity of one diff --git a/Mathlib/Order/CountableDenseLinearOrder.lean b/Mathlib/Order/CountableDenseLinearOrder.lean index 1ab3a1faebbc2..52b1a34e4be76 100644 --- a/Mathlib/Order/CountableDenseLinearOrder.lean +++ b/Mathlib/Order/CountableDenseLinearOrder.lean @@ -82,16 +82,16 @@ lemma exists_orderEmbedding_insert [DenselyOrdered β] [NoMinOrder β] [NoMaxOrd then if hyS : y ∈ S then simpa only [hxS, hyS, ↓reduceDIte, OrderEmbedding.lt_iff_lt, Subtype.mk_lt_mk] else - obtain rfl := Finset.eq_of_not_mem_of_mem_insert hy hyS + obtain rfl := Finset.eq_of_mem_insert_of_not_mem hy hyS simp only [hxS, hyS, ↓reduceDIte] exact hb _ (Finset.mem_image_of_mem _ (Finset.mem_filter.2 ⟨Finset.mem_attach _ _, hxy⟩)) else - obtain rfl := Finset.eq_of_not_mem_of_mem_insert hx hxS + obtain rfl := Finset.eq_of_mem_insert_of_not_mem hx hxS if hyS : y ∈ S then simp only [hxS, hyS, ↓reduceDIte] exact hb' _ (Finset.mem_image_of_mem _ (Finset.mem_filter.2 ⟨Finset.mem_attach _ _, hxy⟩)) - else simp only [Finset.eq_of_not_mem_of_mem_insert hy hyS, lt_self_iff_false] at hxy + else simp only [Finset.eq_of_mem_insert_of_not_mem hy hyS, lt_self_iff_false] at hxy · ext x simp only [Finset.coe_sort_coe, OrderEmbedding.coe_ofStrictMono, Finset.insert_val, Function.comp_apply, Finset.coe_mem, ↓reduceDIte, Subtype.coe_eta] From 495d87ece1fc97b60cedc4b1c9c4e5e9fbb3d7cd Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 25 Oct 2024 09:55:07 +0000 Subject: [PATCH 12/29] feat: add `cfc_commute_cfc` (#18118) For `cfcHom`, these results are exactly `Commute.map`, but since `cfc` is not actually a homomorphism, it's useful to have this version. --- .../CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean | 5 +++++ .../CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean | 5 +++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index e837b34df4812..73105d7530e9b 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -267,6 +267,11 @@ lemma cfcₙ_cases (P : A → Prop) (a : A) (f : R → R) (h₀ : P 0) · rwa [cfcₙ_apply_of_not_map_zero _ h] · rwa [cfcₙ_apply_of_not_predicate _ h] +lemma cfcₙ_commute_cfcₙ (f g : R → R) (a : A) : Commute (cfcₙ f a) (cfcₙ g a) := by + refine cfcₙ_cases (fun x ↦ Commute x (cfcₙ g a)) a f (by simp) fun hf hf0 ha ↦ ?_ + refine cfcₙ_cases (fun x ↦ Commute _ x) a g (by simp) fun hg hg0 _ ↦ ?_ + exact Commute.all _ _ |>.map _ + variable (R) in include ha in lemma cfcₙ_id : cfcₙ (id : R → R) a = a := diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index c6070f8408685..b63d226c61d6d 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -351,6 +351,11 @@ lemma cfc_cases (P : A → Prop) (a : A) (f : R → R) (h₀ : P 0) · rwa [cfc_apply_of_not_predicate _ h] · rwa [cfc_apply_of_not_continuousOn _ h] +lemma cfc_commute_cfc (f g : R → R) (a : A) : Commute (cfc f a) (cfc g a) := by + refine cfc_cases (fun x ↦ Commute x (cfc g a)) a f (by simp) fun hf ha ↦ ?_ + refine cfc_cases (fun x ↦ Commute _ x) a g (by simp) fun hg _ ↦ ?_ + exact Commute.all _ _ |>.map _ + variable (R) in lemma cfc_id (ha : p a := by cfc_tac) : cfc (id : R → R) a = a := cfc_apply (id : R → R) a ▸ cfcHom_id (p := p) ha From 6ca0131ae189339f939abad453cadc81b2811c42 Mon Sep 17 00:00:00 2001 From: Arend Mellendijk Date: Fri, 25 Oct 2024 10:51:48 +0000 Subject: [PATCH 13/29] chore(Tactic/rename_bvar): silence unused tactic linter and add tests (#18139) `rename_bvar` exists only to rename bound variables inside either the goal state or a local hypothesis. This by definition does not change the current goal at all, so the unused tactic linter should be silenced. Add some tests to document this (and the current behaviour). --- Mathlib/Tactic/Linter/UnusedTactic.lean | 1 + Mathlib/Tactic/RenameBVar.lean | 8 +++--- test/renameBvar.lean | 34 +++++++++++++++++++++++++ 3 files changed, 39 insertions(+), 4 deletions(-) create mode 100644 test/renameBvar.lean diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 67eb1b34f7d9f..37f23d5e87ded 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -93,6 +93,7 @@ initialize allowedRef : IO.Ref (Std.HashSet SyntaxNodeKind) ← |>.insert `change? |>.insert `«tactic#adaptation_note_» |>.insert `tacticSleep_heartbeats_ + |>.insert `Mathlib.Tactic.«tacticRename_bvar_→__» /-- `#allow_unused_tactic` takes an input a space-separated list of identifiers. These identifiers are then allowed by the unused tactic linter: diff --git a/Mathlib/Tactic/RenameBVar.lean b/Mathlib/Tactic/RenameBVar.lean index fe1706f2eacfc..64b015482b45f 100644 --- a/Mathlib/Tactic/RenameBVar.lean +++ b/Mathlib/Tactic/RenameBVar.lean @@ -29,13 +29,13 @@ def renameBVarTarget (mvarId : MVarId) (old new : Name) : MetaM Unit := modifyTarget mvarId fun e ↦ e.renameBVar old new /-- -* `rename_bvar old new` renames all bound variables named `old` to `new` in the target. -* `rename_bvar old new at h` does the same in hypothesis `h`. +* `rename_bvar old → new` renames all bound variables named `old` to `new` in the target. +* `rename_bvar old → new at h` does the same in hypothesis `h`. ```lean example (P : ℕ → ℕ → Prop) (h : ∀ n, ∃ m, P n m) : ∀ l, ∃ m, P l m := by - rename_bvar n q at h -- h is now ∀ (q : ℕ), ∃ (m : ℕ), P q m, - rename_bvar m n -- target is now ∀ (l : ℕ), ∃ (n : ℕ), P k n, + rename_bvar n → q at h -- h is now ∀ (q : ℕ), ∃ (m : ℕ), P q m, + rename_bvar m → n -- target is now ∀ (l : ℕ), ∃ (n : ℕ), P k n, exact h -- Lean does not care about those bound variable names ``` Note: name clashes are resolved automatically. diff --git a/test/renameBvar.lean b/test/renameBvar.lean new file mode 100644 index 0000000000000..07fe22319cebf --- /dev/null +++ b/test/renameBvar.lean @@ -0,0 +1,34 @@ +import Mathlib.Tactic.RenameBVar + +/- This test is somewhat flaky since it depends on the pretty printer. -/ +/-- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (n : ℕ), ∃ m, P n m +⊢ ∀ (l : ℕ), ∃ m, P l m +--- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (q : ℕ), ∃ m, P q m +⊢ ∀ (l : ℕ), ∃ m, P l m +--- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (q : ℕ), ∃ m, P q m +⊢ ∀ (l : ℕ), ∃ n, P l n +--- +info: ℕ : Sort u_1 +P : ℕ → ℕ → Prop +h : ∀ (q : ℕ), ∃ m, P q m +⊢ ∀ (m : ℕ), ∃ n, P m n +-/ +#guard_msgs in +example (P : ℕ → ℕ → Prop) (h : ∀ n, ∃ m, P n m) : ∀ l, ∃ m, P l m := by + trace_state + rename_bvar n → q at h + trace_state + rename_bvar m → n + trace_state + rename_bvar l → m + trace_state + exact h From 086c16c9e1afb5cf1671a508c821c4717fe9050e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 25 Oct 2024 11:44:08 +0000 Subject: [PATCH 14/29] chore: split Topology.GDelta (#18215) --- Mathlib.lean | 3 +- .../Constructions/BorelSpace/Basic.lean | 2 +- Mathlib/Topology/Baire/BaireMeasurable.lean | 2 +- Mathlib/Topology/Baire/Lemmas.lean | 2 +- .../{GDelta.lean => GDelta/Basic.lean} | 37 +++---------- Mathlib/Topology/GDelta/UniformSpace.lean | 55 +++++++++++++++++++ Mathlib/Topology/Instances/Irrational.lean | 1 - Mathlib/Topology/Separation/GDelta.lean | 2 +- Mathlib/Topology/UrysohnsLemma.lean | 2 +- 9 files changed, 69 insertions(+), 37 deletions(-) rename Mathlib/Topology/{GDelta.lean => GDelta/Basic.lean} (87%) create mode 100644 Mathlib/Topology/GDelta/UniformSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index 46bd976d0d3e8..25c3ef177489e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4779,7 +4779,8 @@ import Mathlib.Topology.FiberBundle.IsHomeomorphicTrivialBundle import Mathlib.Topology.FiberBundle.Trivialization import Mathlib.Topology.FiberPartition import Mathlib.Topology.Filter -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.Basic +import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.Germ import Mathlib.Topology.Gluing import Mathlib.Topology.Hom.ContinuousEval diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index c5518f43dd87a..f0fe4fcd0af4a 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.MeasureTheory.Group.Arithmetic -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.Instances.EReal import Mathlib.Topology.Instances.Rat diff --git a/Mathlib/Topology/Baire/BaireMeasurable.lean b/Mathlib/Topology/Baire/BaireMeasurable.lean index 02a160951b95f..1c08bafa7d149 100644 --- a/Mathlib/Topology/Baire/BaireMeasurable.lean +++ b/Mathlib/Topology/Baire/BaireMeasurable.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Felix Weilacher. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Felix Weilacher -/ -import Mathlib.Topology.GDelta +import Mathlib.Topology.GDelta.UniformSpace import Mathlib.Topology.LocallyClosed import Mathlib.MeasureTheory.Constructions.EventuallyMeasurable import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic diff --git a/Mathlib/Topology/Baire/Lemmas.lean b/Mathlib/Topology/Baire/Lemmas.lean index 7091aebfeb216..926a6354ec0f2 100644 --- a/Mathlib/Topology/Baire/Lemmas.lean +++ b/Mathlib/Topology/Baire/Lemmas.lean @@ -3,7 +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.GDelta +import Mathlib.Topology.GDelta.Basic /-! # Baire spaces diff --git a/Mathlib/Topology/GDelta.lean b/Mathlib/Topology/GDelta/Basic.lean similarity index 87% rename from Mathlib/Topology/GDelta.lean rename to Mathlib/Topology/GDelta/Basic.lean index 94a3d1bec4af6..60e2975ce2246 100644 --- a/Mathlib/Topology/GDelta.lean +++ b/Mathlib/Topology/GDelta/Basic.lean @@ -3,8 +3,9 @@ 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, Yury Kudryashov -/ -import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Algebra.Group.Defs import Mathlib.Order.Filter.CountableInter +import Mathlib.Topology.Basic /-! # `Gδ` sets @@ -24,8 +25,7 @@ In this file we define `Gδ` sets and prove their basic properties. ## Main results -We prove that finite or countable intersections of Gδ sets are Gδ sets. We also prove that the -continuity set of a function from a topological space to an (e)metric space is a Gδ set. +We prove that finite or countable intersections of Gδ sets are Gδ sets. - `isClosed_isNowhereDense_iff_compl`: a closed set is nowhere dense iff its complement is open and dense @@ -33,16 +33,19 @@ its complement is open and dense union of nowhere dense sets - subsets of meagre sets are meagre; countable unions of meagre sets are meagre +See `Mathlib.Topology.GDelta.UniformSpace` for the proof that +continuity set of a function from a topological space to a uniform space is a Gδ set. + ## Tags Gδ set, residual set, nowhere dense set, meagre set -/ +assert_not_exists UniformSpace noncomputable section open Topology TopologicalSpace Filter Encodable Set -open scoped Uniformity variable {X Y ι : Type*} {ι' : Sort*} @@ -154,34 +157,8 @@ alias isGδ_biUnion := IsGδ.biUnion theorem IsGδ.iUnion [Finite ι'] {f : ι' → Set X} (h : ∀ i, IsGδ (f i)) : IsGδ (⋃ i, f i) := .sUnion (finite_range _) <| forall_mem_range.2 h -theorem IsClosed.isGδ {X : Type*} [UniformSpace X] [IsCountablyGenerated (𝓤 X)] {s : Set X} - (hs : IsClosed s) : IsGδ s := by - rcases (@uniformity_hasBasis_open X _).exists_antitone_subbasis with ⟨U, hUo, hU, -⟩ - rw [← hs.closure_eq, ← hU.biInter_biUnion_ball] - refine .biInter (to_countable _) fun n _ => IsOpen.isGδ ?_ - exact isOpen_biUnion fun x _ => UniformSpace.isOpen_ball _ (hUo _).2 - end IsGδ -section ContinuousAt - -variable [TopologicalSpace X] - -/-- The set of points where a function is continuous is a Gδ set. -/ -theorem IsGδ.setOf_continuousAt [UniformSpace Y] [IsCountablyGenerated (𝓤 Y)] (f : X → Y) : - IsGδ { x | ContinuousAt f x } := by - obtain ⟨U, _, hU⟩ := (@uniformity_hasBasis_open_symmetric Y _).exists_antitone_subbasis - simp only [Uniform.continuousAt_iff_prod, nhds_prod_eq] - simp only [(nhds_basis_opens _).prod_self.tendsto_iff hU.toHasBasis, forall_prop_of_true, - setOf_forall, id] - refine .iInter fun k ↦ IsOpen.isGδ <| isOpen_iff_mem_nhds.2 fun x ↦ ?_ - rintro ⟨s, ⟨hsx, hso⟩, hsU⟩ - filter_upwards [IsOpen.mem_nhds hso hsx] with _ hy using ⟨s, ⟨hy, hso⟩, hsU⟩ - -@[deprecated (since := "2024-02-15")] alias isGδ_setOf_continuousAt := IsGδ.setOf_continuousAt - -end ContinuousAt - section residual variable [TopologicalSpace X] diff --git a/Mathlib/Topology/GDelta/UniformSpace.lean b/Mathlib/Topology/GDelta/UniformSpace.lean new file mode 100644 index 0000000000000..aa957353a7f05 --- /dev/null +++ b/Mathlib/Topology/GDelta/UniformSpace.lean @@ -0,0 +1,55 @@ +/- +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, Yury Kudryashov +-/ +import Mathlib.Topology.GDelta.Basic +import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Order.Filter.CountableInter + +/-! +# `Gδ` sets and uniform spaces + +## Main results +We prove that the continuity set of a function from a topological space to a uniform space is a +Gδ set. + +-/ + + +noncomputable section + +open Topology TopologicalSpace Filter Encodable Set +open scoped Uniformity + +variable {X Y ι : Type*} {ι' : Sort*} + +section IsGδ + +theorem IsClosed.isGδ {X : Type*} [UniformSpace X] [IsCountablyGenerated (𝓤 X)] {s : Set X} + (hs : IsClosed s) : IsGδ s := by + rcases (@uniformity_hasBasis_open X _).exists_antitone_subbasis with ⟨U, hUo, hU, -⟩ + rw [← hs.closure_eq, ← hU.biInter_biUnion_ball] + refine .biInter (to_countable _) fun n _ => IsOpen.isGδ ?_ + exact isOpen_biUnion fun x _ => UniformSpace.isOpen_ball _ (hUo _).2 + +end IsGδ + +section ContinuousAt + +variable [TopologicalSpace X] + +/-- The set of points where a function is continuous is a Gδ set. -/ +theorem IsGδ.setOf_continuousAt [UniformSpace Y] [IsCountablyGenerated (𝓤 Y)] (f : X → Y) : + IsGδ { x | ContinuousAt f x } := by + obtain ⟨U, _, hU⟩ := (@uniformity_hasBasis_open_symmetric Y _).exists_antitone_subbasis + simp only [Uniform.continuousAt_iff_prod, nhds_prod_eq] + simp only [(nhds_basis_opens _).prod_self.tendsto_iff hU.toHasBasis, forall_prop_of_true, + setOf_forall, id] + refine .iInter fun k ↦ IsOpen.isGδ <| isOpen_iff_mem_nhds.2 fun x ↦ ?_ + rintro ⟨s, ⟨hsx, hso⟩, hsU⟩ + filter_upwards [IsOpen.mem_nhds hso hsx] with _ hy using ⟨s, ⟨hy, hso⟩, hsU⟩ + +@[deprecated (since := "2024-02-15")] alias isGδ_setOf_continuousAt := IsGδ.setOf_continuousAt + +end ContinuousAt diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index ad7ac6309e9fa..f395be578fdff 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import Mathlib.Data.Real.Irrational import Mathlib.Data.Rat.Encodable -import Mathlib.Topology.GDelta import Mathlib.Topology.Separation.GDelta /-! diff --git a/Mathlib/Topology/Separation/GDelta.lean b/Mathlib/Topology/Separation/GDelta.lean index 3b58de08a18ee..e7cf8efc00758 100644 --- a/Mathlib/Topology/Separation/GDelta.lean +++ b/Mathlib/Topology/Separation/GDelta.lean @@ -7,8 +7,8 @@ import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.SigmaCompact import Mathlib.Topology.Connected.TotallyDisconnected import Mathlib.Topology.Inseparable -import Mathlib.Topology.GDelta import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.GDelta.Basic /-! # Separation properties of topological spaces. diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 705eb5ca08358..f645df48176ff 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -6,9 +6,9 @@ Authors: Yury Kudryashov import Mathlib.Analysis.Normed.Affine.AddTorsor import Mathlib.LinearAlgebra.AffineSpace.Ordered import Mathlib.Topology.ContinuousMap.Basic -import Mathlib.Topology.GDelta import Mathlib.Analysis.NormedSpace.FunctionSeries import Mathlib.Analysis.SpecificLimits.Basic +import Mathlib.Topology.GDelta.Basic /-! # Urysohn's lemma From 21128da4eb35f0ab7bb8165a856f11d38de987b3 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Fri, 25 Oct 2024 12:15:34 +0000 Subject: [PATCH 15/29] chore: split StructuredArrow.lean to reduce imports (#18227) Found using the new `lake exe unused`. Co-authored-by: Markus Himmel --- Mathlib.lean | 3 +- .../Adjunction/AdjointFunctorTheorems.lean | 1 + Mathlib/CategoryTheory/Adjunction/Comma.lean | 2 +- .../CategoryTheory/Bicategory/Extension.lean | 2 +- Mathlib/CategoryTheory/Comma/Over.lean | 2 +- .../Basic.lean} | 16 ------ .../Comma/StructuredArrow/Small.lean | 50 +++++++++++++++++++ Mathlib/CategoryTheory/Elements.lean | 2 +- .../Functor/KanExtension/Basic.lean | 2 +- Mathlib/CategoryTheory/Limits/Comma.lean | 2 +- Mathlib/CategoryTheory/Limits/Final.lean | 2 +- .../CategoryTheory/Limits/FinallySmall.lean | 1 + .../Localization/StructuredArrow.lean | 2 +- Mathlib/CategoryTheory/Sites/Sieves.lean | 1 + 14 files changed, 63 insertions(+), 25 deletions(-) rename Mathlib/CategoryTheory/Comma/{StructuredArrow.lean => StructuredArrow/Basic.lean} (97%) create mode 100644 Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean diff --git a/Mathlib.lean b/Mathlib.lean index 25c3ef177489e..a58a3a3939b0d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1541,7 +1541,8 @@ import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.Comma.Over import Mathlib.CategoryTheory.Comma.Presheaf.Basic import Mathlib.CategoryTheory.Comma.Presheaf.Colimit -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic +import Mathlib.CategoryTheory.Comma.StructuredArrow.Small import Mathlib.CategoryTheory.ComposableArrows import Mathlib.CategoryTheory.ConcreteCategory.Basic import Mathlib.CategoryTheory.ConcreteCategory.Bundled diff --git a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean index 3d5e083b0ab9b..6d42354bbaaa3 100644 --- a/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean +++ b/Mathlib/CategoryTheory/Adjunction/AdjointFunctorTheorems.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ +import Mathlib.CategoryTheory.Comma.StructuredArrow.Small import Mathlib.CategoryTheory.Generator import Mathlib.CategoryTheory.Limits.ConeCategory import Mathlib.CategoryTheory.Limits.Constructions.WeaklyInitial diff --git a/Mathlib/CategoryTheory/Adjunction/Comma.lean b/Mathlib/CategoryTheory/Adjunction/Comma.lean index 3b59a2dd44bdc..6f4474b55e705 100644 --- a/Mathlib/CategoryTheory/Adjunction/Comma.lean +++ b/Mathlib/CategoryTheory/Adjunction/Comma.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ import Mathlib.CategoryTheory.Adjunction.Basic -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.PUnit /-! diff --git a/Mathlib/CategoryTheory/Bicategory/Extension.lean b/Mathlib/CategoryTheory/Bicategory/Extension.lean index 898bc3c43dd3e..5f371e5ffbfee 100644 --- a/Mathlib/CategoryTheory/Bicategory/Extension.lean +++ b/Mathlib/CategoryTheory/Bicategory/Extension.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ import Mathlib.CategoryTheory.Bicategory.Basic -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic /-! # Extensions and lifts in bicategories diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 3b9a32d2c8eee..2d0ed2bd61b9c 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Bhavik Mehta -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.PUnit import Mathlib.CategoryTheory.Functor.ReflectsIso import Mathlib.CategoryTheory.Functor.EpiMono diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean similarity index 97% rename from Mathlib/CategoryTheory/Comma/StructuredArrow.lean rename to Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index d86b41669eacf..c0d1729054be6 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -6,8 +6,6 @@ Authors: Adam Topaz, Kim Morrison import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.PUnit import Mathlib.CategoryTheory.Limits.Shapes.Terminal -import Mathlib.CategoryTheory.EssentiallySmall -import Mathlib.Logic.Small.Set /-! # The category of "structured arrows" @@ -331,13 +329,6 @@ noncomputable instance 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 - rw [this] - infer_instance - exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ - /-- A structured arrow is called universal if it is initial. -/ abbrev IsUniversal (f : StructuredArrow S T) := IsInitial f @@ -674,13 +665,6 @@ noncomputable instance 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 - rw [this] - infer_instance - exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ - /-- A costructured arrow is called universal if it is terminal. -/ abbrev IsUniversal (f : CostructuredArrow S T) := IsTerminal f diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean new file mode 100644 index 0000000000000..a31121a52e88a --- /dev/null +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Small.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic +import Mathlib.CategoryTheory.EssentiallySmall +import Mathlib.Logic.Small.Set + +/-! +# Small sets in the category of structured arrows + +Here we prove a technical result about small sets in the category of structured arrows that will +be used in the proof of the Special Adjoint Functor Theorem. +-/ + +namespace CategoryTheory + +-- morphism levels before object levels. See note [CategoryTheory universes]. +universe v₁ v₂ u₁ u₂ + +variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] + +namespace StructuredArrow + +variable {S : D} {T : C ⥤ D} + +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 + rw [this] + infer_instance + exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ + +end StructuredArrow + +namespace CostructuredArrow + +variable {S : C ⥤ D} {T : D} + +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 + rw [this] + infer_instance + exact Set.ext fun X => ⟨fun h => ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by aesop_cat⟩ + +end CostructuredArrow + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index cf261e1d8d7fc..2c682ecb881f9 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.Groupoid import Mathlib.CategoryTheory.PUnit diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 6b966a9ee45ed..7c08315b07111 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.Limits.Shapes.Equivalence import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal diff --git a/Mathlib/CategoryTheory/Limits/Comma.lean b/Mathlib/CategoryTheory/Limits/Comma.lean index a54e3c6a62630..5e5e2804c305b 100644 --- a/Mathlib/CategoryTheory/Limits/Comma.lean +++ b/Mathlib/CategoryTheory/Limits/Comma.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ import Mathlib.CategoryTheory.Comma.Arrow -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.Limits.Constructions.EpiMono import Mathlib.CategoryTheory.Limits.Creates import Mathlib.CategoryTheory.Limits.Unit diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 12c39f6153dc8..8dea5b13733db 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.IsConnected import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Terminal import Mathlib.CategoryTheory.Limits.Shapes.Types diff --git a/Mathlib/CategoryTheory/Limits/FinallySmall.lean b/Mathlib/CategoryTheory/Limits/FinallySmall.lean index 90703593a078a..3613b80d171c2 100644 --- a/Mathlib/CategoryTheory/Limits/FinallySmall.lean +++ b/Mathlib/CategoryTheory/Limits/FinallySmall.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ +import Mathlib.Logic.Small.Set import Mathlib.CategoryTheory.Filtered.Final /-! diff --git a/Mathlib/CategoryTheory/Localization/StructuredArrow.lean b/Mathlib/CategoryTheory/Localization/StructuredArrow.lean index 91c9d508c56f6..96c0af25daccb 100644 --- a/Mathlib/CategoryTheory/Localization/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Localization/StructuredArrow.lean @@ -5,7 +5,7 @@ Authors: Joël Riou -/ import Mathlib.CategoryTheory.Localization.HomEquiv import Mathlib.CategoryTheory.Localization.Opposite -import Mathlib.CategoryTheory.Comma.StructuredArrow +import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic /-! # Induction principles for structured and costructured arrows diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index d974d994459f0..66591c5a8f117 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta, Edward Ayers -/ +import Mathlib.Data.Set.Lattice import Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback /-! From d09d074b8b2b5d8f6a89f259cd8751d5a2a45c9e Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Fri, 25 Oct 2024 12:49:04 +0000 Subject: [PATCH 16/29] chore(NumberField/Embeddings): change `Nr(Real|Complex)Places` to `nr(Real|Complex)Places` (#18140) As noted by @acmepjz, `NrRealPlaces` and `NrComplexPlaces` should be called `nrRealPlaces` and `nrComplexPlaces`. This is fixed in this PR. Co-authored-by: Xavier Roblot <46200072+xroblot@users.noreply.github.com> --- .../NumberTheory/Cyclotomic/Embeddings.lean | 4 +-- .../NumberField/CanonicalEmbedding/Basic.lean | 8 +++--- .../CanonicalEmbedding/ConvexBody.lean | 26 +++++++++---------- .../NumberTheory/NumberField/ClassNumber.lean | 4 +-- .../NumberField/Discriminant/Basic.lean | 22 ++++++++-------- .../NumberTheory/NumberField/Embeddings.lean | 22 +++++++++------- 6 files changed, 45 insertions(+), 41 deletions(-) diff --git a/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean b/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean index 642d6933fad59..a6b2d1fc27387 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Embeddings.lean @@ -30,7 +30,7 @@ of `K`. -/ theorem nrRealPlaces_eq_zero [IsCyclotomicExtension {n} ℚ K] (hn : 2 < n) : haveI := IsCyclotomicExtension.numberField {n} ℚ K - NrRealPlaces K = 0 := by + nrRealPlaces K = 0 := by have := IsCyclotomicExtension.numberField {n} ℚ K apply (IsCyclotomicExtension.zeta_spec n ℚ K).nrRealPlaces_eq_zero_of_two_lt hn @@ -40,7 +40,7 @@ variable (n) of `K`. Note that this uses `1 / 2 = 0` in the cases `n = 1, 2`. -/ theorem nrComplexPlaces_eq_totient_div_two [h : IsCyclotomicExtension {n} ℚ K] : haveI := IsCyclotomicExtension.numberField {n} ℚ K - NrComplexPlaces K = φ n / 2 := by + nrComplexPlaces K = φ n / 2 := by have := IsCyclotomicExtension.numberField {n} ℚ K by_cases hn : 2 < n · obtain ⟨k, hk : φ n = k + k⟩ := totient_even hn diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index f119b3f3072eb..0ff7cd0991a72 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -208,7 +208,7 @@ instance [NumberField K] : Nontrivial (mixedSpace K) := by protected theorem finrank [NumberField K] : finrank ℝ (mixedSpace K) = finrank ℚ K := by classical rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, sum_const, - card_univ, ← NrRealPlaces, ← NrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, + card_univ, ← nrRealPlaces, ← nrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, Nat.add_sub_of_le (Fintype.card_subtype_le _)] @@ -517,7 +517,7 @@ def matrixToStdBasis : Matrix (index K) (index K) ℂ := (blockDiagonal (fun _ => (2 : ℂ)⁻¹ • !![1, 1; - I, I])) theorem det_matrixToStdBasis : - (matrixToStdBasis K).det = (2⁻¹ * I) ^ NrComplexPlaces K := + (matrixToStdBasis K).det = (2⁻¹ * I) ^ nrComplexPlaces K := calc _ = ∏ _k : { w : InfinitePlace K // IsComplex w }, det ((2 : ℂ)⁻¹ • !![1, 1; -I, I]) := by rw [matrixToStdBasis, det_fromBlocks_zero₂₁, det_diagonal, prod_const_one, one_mul, @@ -590,8 +590,8 @@ def latticeBasis : -- and it's a basis since it has the right cardinality refine basisOfLinearIndependentOfCardEqFinrank this ?_ rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_prod, finrank_pi, - finrank_pi_fintype, Complex.finrank_real_complex, sum_const, card_univ, ← NrRealPlaces, - ← NrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, + finrank_pi_fintype, Complex.finrank_real_complex, sum_const, card_univ, ← nrRealPlaces, + ← nrComplexPlaces, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, Nat.add_sub_of_le (Fintype.card_subtype_le _)] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 36a06aabe738b..561216bee288f 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -88,7 +88,7 @@ instance : NoAtoms (volume : Measure (mixedSpace K)) := by /-- The fudge factor that appears in the formula for the volume of `convexBodyLT`. -/ noncomputable abbrev convexBodyLTFactor : ℝ≥0 := - (2 : ℝ≥0) ^ NrRealPlaces K * NNReal.pi ^ NrComplexPlaces K + (2 : ℝ≥0) ^ nrRealPlaces K * NNReal.pi ^ nrComplexPlaces K theorem convexBodyLTFactor_ne_zero : convexBodyLTFactor K ≠ 0 := mul_ne_zero (pow_ne_zero _ two_ne_zero) (pow_ne_zero _ pi_ne_zero) @@ -104,10 +104,10 @@ theorem convexBodyLT_volume : _ = (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (2 * (f x.val))) * ∏ x : {w // InfinitePlace.IsComplex w}, ENNReal.ofReal (f x.val) ^ 2 * NNReal.pi := by simp_rw [volume_eq_prod, prod_prod, volume_pi, pi_pi, Real.volume_ball, Complex.volume_ball] - _ = ((2 : ℝ≥0) ^ NrRealPlaces K + _ = ((2 : ℝ≥0) ^ nrRealPlaces K * (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * ((∏ x : {w // IsComplex w}, ENNReal.ofReal (f x.val) ^ 2) * - NNReal.pi ^ NrComplexPlaces K) := by + NNReal.pi ^ nrComplexPlaces K) := by simp_rw [ofReal_mul (by norm_num : 0 ≤ (2 : ℝ)), Finset.prod_mul_distrib, Finset.prod_const, Finset.card_univ, ofReal_ofNat, ofReal_coe_nnreal, coe_ofNat] _ = (convexBodyLTFactor K) * ((∏ x : {w // InfinitePlace.IsReal w}, .ofReal (f x.val)) * @@ -205,7 +205,7 @@ variable [NumberField K] /-- The fudge factor that appears in the formula for the volume of `convexBodyLT'`. -/ noncomputable abbrev convexBodyLT'Factor : ℝ≥0 := - (2 : ℝ≥0) ^ (NrRealPlaces K + 2) * NNReal.pi ^ (NrComplexPlaces K - 1) + (2 : ℝ≥0) ^ (nrRealPlaces K + 2) * NNReal.pi ^ (nrComplexPlaces K - 1) theorem convexBodyLT'Factor_ne_zero : convexBodyLT'Factor K ≠ 0 := mul_ne_zero (pow_ne_zero _ two_ne_zero) (pow_ne_zero _ pi_ne_zero) @@ -238,10 +238,10 @@ theorem convexBodyLT'_volume : · refine Finset.prod_congr rfl (fun w' hw' ↦ ?_) rw [if_neg (Finset.ne_of_mem_erase hw'), Complex.volume_ball] · simpa only [ite_true] using vol_box (f w₀) - _ = ((2 : ℝ≥0) ^ NrRealPlaces K * + _ = ((2 : ℝ≥0) ^ nrRealPlaces K * (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * ((∏ x ∈ Finset.univ.erase w₀, ENNReal.ofReal (f x.val) ^ 2) * - ↑pi ^ (NrComplexPlaces K - 1) * (4 * (f w₀) ^ 2)) := by + ↑pi ^ (nrComplexPlaces K - 1) * (4 * (f w₀) ^ 2)) := by simp_rw [ofReal_mul (by norm_num : 0 ≤ (2 : ℝ)), Finset.prod_mul_distrib, Finset.prod_const, Finset.card_erase_of_mem (Finset.mem_univ _), Finset.card_univ, ofReal_ofNat, ofReal_coe_nnreal, coe_ofNat] @@ -385,7 +385,7 @@ theorem convexBodySum_compact : IsCompact (convexBodySum K B) := by /-- The fudge factor that appears in the formula for the volume of `convexBodyLt`. -/ noncomputable abbrev convexBodySumFactor : ℝ≥0 := - (2 : ℝ≥0) ^ NrRealPlaces K * (NNReal.pi / 2) ^ NrComplexPlaces K / (finrank ℚ K).factorial + (2 : ℝ≥0) ^ nrRealPlaces K * (NNReal.pi / 2) ^ nrComplexPlaces K / (finrank ℚ K).factorial theorem convexBodySumFactor_ne_zero : convexBodySumFactor K ≠ 0 := by refine div_ne_zero ?_ <| Nat.cast_ne_zero.mpr (Nat.factorial_ne_zero _) @@ -418,7 +418,7 @@ theorem convexBodySum_volume : simp_rw [mixedEmbedding.finrank, div_one, Gamma_nat_eq_factorial, ofReal_div_of_pos (Nat.cast_pos.mpr (Nat.factorial_pos _)), Real.rpow_one, ofReal_natCast] suffices ∫ x : mixedSpace K, exp (-convexBodySumFun x) = - (2 : ℝ) ^ NrRealPlaces K * (π / 2) ^ NrComplexPlaces K by + (2 : ℝ) ^ nrRealPlaces K * (π / 2) ^ nrComplexPlaces K by rw [this, convexBodySumFactor, ofReal_mul (by positivity), ofReal_pow zero_le_two, ofReal_pow (by positivity), ofReal_div_of_pos zero_lt_two, ofReal_ofNat, ← NNReal.coe_real_pi, ofReal_coe_nnreal, coe_div (Nat.cast_ne_zero.mpr @@ -429,17 +429,17 @@ theorem convexBodySum_volume : (∫ x : {w : InfinitePlace K // IsComplex w} → ℂ, ∏ w, exp (- 2 * ‖x w‖)) := by simp_rw [convexBodySumFun_apply', neg_add, ← neg_mul, Finset.mul_sum, ← Finset.sum_neg_distrib, exp_add, exp_sum, ← integral_prod_mul, volume_eq_prod] - _ = (∫ x : ℝ, exp (-|x|)) ^ NrRealPlaces K * - (∫ x : ℂ, Real.exp (-2 * ‖x‖)) ^ NrComplexPlaces K := by + _ = (∫ x : ℝ, exp (-|x|)) ^ nrRealPlaces K * + (∫ x : ℂ, Real.exp (-2 * ‖x‖)) ^ nrComplexPlaces K := by rw [integral_fintype_prod_eq_pow _ (fun x => exp (- ‖x‖)), integral_fintype_prod_eq_pow _ (fun x => exp (- 2 * ‖x‖))] simp_rw [norm_eq_abs] - _ = (2 * Gamma (1 / 1 + 1)) ^ NrRealPlaces K * - (π * (2 : ℝ) ^ (-(2 : ℝ) / 1) * Gamma (2 / 1 + 1)) ^ NrComplexPlaces K := by + _ = (2 * Gamma (1 / 1 + 1)) ^ nrRealPlaces K * + (π * (2 : ℝ) ^ (-(2 : ℝ) / 1) * Gamma (2 / 1 + 1)) ^ nrComplexPlaces K := by rw [integral_comp_abs (f := fun x => exp (- x)), ← integral_exp_neg_rpow zero_lt_one, ← Complex.integral_exp_neg_mul_rpow le_rfl zero_lt_two] simp_rw [Real.rpow_one] - _ = (2 : ℝ) ^ NrRealPlaces K * (π / 2) ^ NrComplexPlaces K := by + _ = (2 : ℝ) ^ nrRealPlaces K * (π / 2) ^ nrComplexPlaces K := by simp_rw [div_one, one_add_one_eq_two, Gamma_add_one two_ne_zero, Gamma_two, mul_one, mul_assoc, ← Real.rpow_add_one two_ne_zero, show (-2 : ℝ) + 1 = -1 by norm_num, Real.rpow_neg_one] diff --git a/Mathlib/NumberTheory/NumberField/ClassNumber.lean b/Mathlib/NumberTheory/NumberField/ClassNumber.lean index 769d1f2df3687..0c87862166478 100644 --- a/Mathlib/NumberTheory/NumberField/ClassNumber.lean +++ b/Mathlib/NumberTheory/NumberField/ClassNumber.lean @@ -46,7 +46,7 @@ open scoped nonZeroDivisors Real theorem exists_ideal_in_class_of_norm_le (C : ClassGroup (𝓞 K)) : ∃ I : (Ideal (𝓞 K))⁰, ClassGroup.mk0 I = C ∧ - Ideal.absNorm (I : Ideal (𝓞 K)) ≤ (4 / π) ^ NrComplexPlaces K * + Ideal.absNorm (I : Ideal (𝓞 K)) ≤ (4 / π) ^ nrComplexPlaces K * ((finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K|) := by obtain ⟨J, hJ⟩ := ClassGroup.mk0_surjective C⁻¹ obtain ⟨_, ⟨a, ha, rfl⟩, h_nz, h_nm⟩ := @@ -70,7 +70,7 @@ theorem exists_ideal_in_class_of_norm_le (C : ClassGroup (𝓞 K)) : exact Nat.cast_pos.mpr <| Nat.pos_of_ne_zero <| Ideal.absNorm_ne_zero_of_nonZeroDivisors J theorem _root_.RingOfIntegers.isPrincipalIdealRing_of_abs_discr_lt - (h : |discr K| < (2 * (π / 4) ^ NrComplexPlaces K * + (h : |discr K| < (2 * (π / 4) ^ nrComplexPlaces K * ((finrank ℚ K) ^ (finrank ℚ K) / (finrank ℚ K).factorial)) ^ 2) : IsPrincipalIdealRing (𝓞 K) := by have : 0 < finrank ℚ K := finrank_pos -- Lean needs to know that for positivity to succeed diff --git a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean index 26397e665bada..69a827b17b1c0 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean @@ -40,7 +40,7 @@ open MeasureTheory MeasureTheory.Measure ZSpan NumberField.mixedEmbedding theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis : volume (fundamentalDomain (latticeBasis K)) = - (2 : ℝ≥0∞)⁻¹ ^ NrComplexPlaces K * sqrt ‖discr K‖₊ := by + (2 : ℝ≥0∞)⁻¹ ^ nrComplexPlaces K * sqrt ‖discr K‖₊ := by let f : Module.Free.ChooseBasisIndex ℤ (𝓞 K) ≃ (K →+* ℂ) := (canonicalEmbedding.latticeBasis K).indexEquiv (Pi.basisFun ℂ _) let e : (index K) ≃ Module.Free.ChooseBasisIndex ℤ (𝓞 K) := (indexEquiv K).trans f.symm @@ -74,7 +74,7 @@ theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : ∃ a ∈ (I : FractionalIdeal (𝓞 K)⁰ K), a ≠ 0 ∧ - |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ NrComplexPlaces K * + |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K| := by -- The smallest possible value for `exists_ne_zero_mem_ideal_of_norm_le` let B := (minkowskiBound K I * (convexBodySumFactor K)⁻¹).toReal ^ (1 / (finrank ℚ K : ℝ)) @@ -92,8 +92,8 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal congr 1 rw [eq_comm] calc - _ = FractionalIdeal.absNorm I.1 * (2 : ℝ)⁻¹ ^ NrComplexPlaces K * sqrt ‖discr K‖₊ * - (2 : ℝ) ^ finrank ℚ K * ((2 : ℝ) ^ NrRealPlaces K * (π / 2) ^ NrComplexPlaces K / + _ = FractionalIdeal.absNorm I.1 * (2 : ℝ)⁻¹ ^ nrComplexPlaces K * sqrt ‖discr K‖₊ * + (2 : ℝ) ^ finrank ℚ K * ((2 : ℝ) ^ nrRealPlaces K * (π / 2) ^ nrComplexPlaces K / (Nat.factorial (finrank ℚ K)))⁻¹ := by simp_rw [minkowskiBound, convexBodySumFactor, volume_fundamentalDomain_fractionalIdealLatticeBasis, @@ -102,19 +102,19 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal rw [ENNReal.toReal_ofReal (Rat.cast_nonneg.mpr (FractionalIdeal.absNorm_nonneg I.1))] simp_rw [NNReal.coe_inv, NNReal.coe_div, NNReal.coe_mul, NNReal.coe_pow, NNReal.coe_div, coe_real_pi, NNReal.coe_ofNat, NNReal.coe_natCast] - _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (finrank ℚ K - NrComplexPlaces K - NrRealPlaces K + - NrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * Nat.factorial (finrank ℚ K) * - π⁻¹ ^ (NrComplexPlaces K) := by + _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (finrank ℚ K - nrComplexPlaces K - nrRealPlaces K + + nrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * Nat.factorial (finrank ℚ K) * + π⁻¹ ^ (nrComplexPlaces K) := by simp_rw [inv_div, div_eq_mul_inv, mul_inv, ← zpow_neg_one, ← zpow_natCast, mul_zpow, ← zpow_mul, neg_one_mul, mul_neg_one, neg_neg, Real.coe_sqrt, coe_nnnorm, sub_eq_add_neg, zpow_add₀ (two_ne_zero : (2 : ℝ) ≠ 0)] ring - _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (2 * NrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * - Nat.factorial (finrank ℚ K) * π⁻¹ ^ (NrComplexPlaces K) := by + _ = FractionalIdeal.absNorm I.1 * (2 : ℝ) ^ (2 * nrComplexPlaces K : ℤ) * Real.sqrt ‖discr K‖ * + Nat.factorial (finrank ℚ K) * π⁻¹ ^ (nrComplexPlaces K) := by congr rw [← card_add_two_mul_card_eq_rank, Nat.cast_add, Nat.cast_mul, Nat.cast_ofNat] ring - _ = FractionalIdeal.absNorm I.1 * (4 / π) ^ NrComplexPlaces K * (finrank ℚ K).factorial * + _ = FractionalIdeal.absNorm I.1 * (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial * Real.sqrt |discr K| := by rw [Int.norm_eq_abs, zpow_mul, show (2 : ℝ) ^ (2 : ℤ) = 4 by norm_cast, div_pow, inv_eq_one_div, div_pow, one_pow, zpow_natCast] @@ -122,7 +122,7 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal theorem exists_ne_zero_mem_ringOfIntegers_of_norm_le_mul_sqrt_discr : ∃ (a : 𝓞 K), a ≠ 0 ∧ - |Algebra.norm ℚ (a : K)| ≤ (4 / π) ^ NrComplexPlaces K * + |Algebra.norm ℚ (a : K)| ≤ (4 / π) ^ nrComplexPlaces K * (finrank ℚ K).factorial / (finrank ℚ K) ^ (finrank ℚ K) * Real.sqrt |discr K| := by obtain ⟨_, h_mem, h_nz, h_nm⟩ := exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr K ↑1 obtain ⟨a, rfl⟩ := (FractionalIdeal.mem_one_iff _).mp h_mem diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index d5e8af77ec71d..f041f1b899502 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -551,22 +551,26 @@ section NumberField variable [NumberField K] /-- The number of infinite real places of the number field `K`. -/ -noncomputable abbrev NrRealPlaces := card { w : InfinitePlace K // IsReal w } +noncomputable abbrev nrRealPlaces := card { w : InfinitePlace K // IsReal w } + +@[deprecated (since := "2024-10-24")] alias NrRealPlaces := nrRealPlaces /-- The number of infinite complex places of the number field `K`. -/ -noncomputable abbrev NrComplexPlaces := card { w : InfinitePlace K // IsComplex w } +noncomputable abbrev nrComplexPlaces := card { w : InfinitePlace K // IsComplex w } + +@[deprecated (since := "2024-10-24")] alias NrComplexPlaces := nrComplexPlaces theorem card_real_embeddings : - card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = NrRealPlaces K := Fintype.card_congr mkReal + card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = nrRealPlaces K := Fintype.card_congr mkReal theorem card_eq_nrRealPlaces_add_nrComplexPlaces : - Fintype.card (InfinitePlace K) = NrRealPlaces K + NrComplexPlaces K := by + Fintype.card (InfinitePlace K) = nrRealPlaces K + nrComplexPlaces K := by convert Fintype.card_subtype_or_disjoint (IsReal (K := K)) (IsComplex (K := K)) (disjoint_isReal_isComplex K) using 1 exact (Fintype.card_of_subtype _ (fun w ↦ ⟨fun _ ↦ isReal_or_isComplex w, fun _ ↦ by simp⟩)).symm theorem card_complex_embeddings : - card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * NrComplexPlaces K := by + card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * nrComplexPlaces K := by suffices ∀ w : { w : InfinitePlace K // IsComplex w }, #{φ : {φ //¬ ComplexEmbedding.IsReal φ} | mkComplex φ = w} = 2 by rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] @@ -583,7 +587,7 @@ theorem card_complex_embeddings : · simp_rw [mult, not_isReal_iff_isComplex.mpr hw, ite_false] theorem card_add_two_mul_card_eq_rank : - NrRealPlaces K + 2 * NrComplexPlaces K = finrank ℚ K := by + nrRealPlaces K + 2 * nrComplexPlaces K = finrank ℚ K := by rw [← card_real_embeddings, ← card_complex_embeddings, Fintype.card_subtype_compl, ← Embeddings.card K ℂ, Nat.add_sub_of_le] exact Fintype.card_subtype_le _ @@ -591,10 +595,10 @@ theorem card_add_two_mul_card_eq_rank : variable {K} theorem nrComplexPlaces_eq_zero_of_finrank_eq_one (h : finrank ℚ K = 1) : - NrComplexPlaces K = 0 := by linarith [card_add_two_mul_card_eq_rank K] + nrComplexPlaces K = 0 := by linarith [card_add_two_mul_card_eq_rank K] theorem nrRealPlaces_eq_one_of_finrank_eq_one (h : finrank ℚ K = 1) : - NrRealPlaces K = 1 := by + nrRealPlaces K = 1 := by have := card_add_two_mul_card_eq_rank K rwa [nrComplexPlaces_eq_zero_of_finrank_eq_one h, h, mul_zero, add_zero] at this @@ -1042,7 +1046,7 @@ namespace IsPrimitiveRoot variable {K : Type*} [Field K] [NumberField K] {ζ : K} {k : ℕ} theorem nrRealPlaces_eq_zero_of_two_lt (hk : 2 < k) (hζ : IsPrimitiveRoot ζ k) : - NumberField.InfinitePlace.NrRealPlaces K = 0 := by + NumberField.InfinitePlace.nrRealPlaces K = 0 := by refine (@Fintype.card_eq_zero_iff _ (_)).2 ⟨fun ⟨w, hwreal⟩ ↦ ?_⟩ rw [NumberField.InfinitePlace.isReal_iff] at hwreal let f := w.embedding From b5a5319a32d9785c6c18f40b803ff6af0ff31077 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 12:58:33 +0000 Subject: [PATCH 17/29] chore(Algebra/Associated): split off Prime and Irreducible (#18224) This PR moves definitions about prime and irreducible elements to their own file. I always found it somewhat weird that associated elements and prime elements had to live together. After this PR: * `Prime`, `Irreducible` and `irreducible_iff_prime` will live in `Algebra/Prime/Defs.lean` * All results on prime and irreducible elements that do not further need associated elements will live in `Algebra/Prime/Lemmas.lean` I considered also splitting up `Associated/Basic.lean` into `Defs`, `Lemmas` and `Prime`, but that might be for a follow-up PR. This is in preparation for further cleaning up `Data/Nat/Prime/Defs.lean`. --- Mathlib.lean | 2 + Mathlib/Algebra/Associated/Basic.lean | 397 +------------------------- Mathlib/Algebra/Prime/Defs.lean | 191 +++++++++++++ Mathlib/Algebra/Prime/Lemmas.lean | 276 ++++++++++++++++++ Mathlib/Data/Nat/Prime/Basic.lean | 5 +- Mathlib/Data/Nat/Prime/Defs.lean | 2 +- Mathlib/RingTheory/Prime.lean | 3 +- 7 files changed, 478 insertions(+), 398 deletions(-) create mode 100644 Mathlib/Algebra/Prime/Defs.lean create mode 100644 Mathlib/Algebra/Prime/Lemmas.lean diff --git a/Mathlib.lean b/Mathlib.lean index a58a3a3939b0d..d885b3af7da93 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -771,6 +771,8 @@ import Mathlib.Algebra.Polynomial.Splits import Mathlib.Algebra.Polynomial.SumIteratedDerivative import Mathlib.Algebra.Polynomial.Taylor import Mathlib.Algebra.Polynomial.UnitTrinomial +import Mathlib.Algebra.Prime.Defs +import Mathlib.Algebra.Prime.Lemmas import Mathlib.Algebra.QuadraticDiscriminant import Mathlib.Algebra.Quandle import Mathlib.Algebra.Quaternion diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index ae0f10085a509..3697b4cbe06ce 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -10,19 +10,12 @@ import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Ring.Units +import Mathlib.Algebra.Prime.Lemmas /-! -# Associated, prime, and irreducible elements. +# Associated elements. -In this file we define the predicate `Prime p` -saying that an element of a commutative monoid with zero is prime. -Namely, `Prime p` means that `p` isn't zero, it isn't a unit, -and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`; - -In decomposition monoids (e.g., `ℕ`, `ℤ`), this predicate is equivalent to `Irreducible`, -however this is not true in general. - -We also define an equivalence relation `Associated` +In this file we define an equivalence relation `Associated` saying that two elements of a monoid differ by a multiplication by a unit. Then we show that the quotient type `Associates` is a monoid and prove basic properties of this quotient. @@ -33,357 +26,6 @@ assert_not_exists Multiset variable {M N : Type*} -section Prime - -variable [CommMonoidWithZero M] - -/-- An element `p` of a commutative monoid with zero (e.g., a ring) is called *prime*, -if it's not zero, not a unit, and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`. -/ -def Prime (p : M) : Prop := - p ≠ 0 ∧ ¬IsUnit p ∧ ∀ a b, p ∣ a * b → p ∣ a ∨ p ∣ b - -namespace Prime - -variable {p : M} (hp : Prime p) -include hp - -theorem ne_zero : p ≠ 0 := - hp.1 - -theorem not_unit : ¬IsUnit p := - hp.2.1 - -theorem not_dvd_one : ¬p ∣ 1 := - mt (isUnit_of_dvd_one ·) hp.not_unit - -theorem ne_one : p ≠ 1 := fun h => hp.2.1 (h.symm ▸ isUnit_one) - -theorem dvd_or_dvd {a b : M} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := - hp.2.2 a b h - -theorem dvd_mul {a b : M} : p ∣ a * b ↔ p ∣ a ∨ p ∣ b := - ⟨hp.dvd_or_dvd, (Or.elim · (dvd_mul_of_dvd_left · _) (dvd_mul_of_dvd_right · _))⟩ - -theorem isPrimal : IsPrimal p := fun _a _b dvd ↦ (hp.dvd_or_dvd dvd).elim - (fun h ↦ ⟨p, 1, h, one_dvd _, (mul_one p).symm⟩) fun h ↦ ⟨1, p, one_dvd _, h, (one_mul p).symm⟩ - -theorem not_dvd_mul {a b : M} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a * b := - hp.dvd_mul.not.mpr <| not_or.mpr ⟨ha, hb⟩ - -theorem dvd_of_dvd_pow {a : M} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by - induction n with - | zero => - rw [pow_zero] at h - have := isUnit_of_dvd_one h - have := not_unit hp - contradiction - | succ n ih => - rw [pow_succ'] at h - rcases dvd_or_dvd hp h with dvd_a | dvd_pow - · assumption - · exact ih dvd_pow - -theorem dvd_pow_iff_dvd {a : M} {n : ℕ} (hn : n ≠ 0) : p ∣ a ^ n ↔ p ∣ a := - ⟨hp.dvd_of_dvd_pow, (dvd_pow · hn)⟩ - -end Prime - -@[simp] -theorem not_prime_zero : ¬Prime (0 : M) := fun h => h.ne_zero rfl - -@[simp] -theorem not_prime_one : ¬Prime (1 : M) := fun h => h.not_unit isUnit_one - -section Map - -variable [CommMonoidWithZero N] {F : Type*} {G : Type*} [FunLike F M N] -variable [MonoidWithZeroHomClass F M N] [FunLike G N M] [MulHomClass G N M] -variable (f : F) (g : G) {p : M} - -theorem comap_prime (hinv : ∀ a, g (f a : N) = a) (hp : Prime (f p)) : Prime p := - ⟨fun h => hp.1 <| by simp [h], fun h => hp.2.1 <| h.map f, fun a b h => by - refine - (hp.2.2 (f a) (f b) <| by - convert map_dvd f h - simp).imp - ?_ ?_ <;> - · intro h - convert ← map_dvd g h <;> apply hinv⟩ - -theorem MulEquiv.prime_iff (e : M ≃* N) : Prime p ↔ Prime (e p) := - ⟨fun h => (comap_prime e.symm e fun a => by simp) <| (e.symm_apply_apply p).substr h, - comap_prime e e.symm fun a => by simp⟩ - -end Map - -end Prime - -theorem Prime.left_dvd_or_dvd_right_of_dvd_mul [CancelCommMonoidWithZero M] {p : M} (hp : Prime p) - {a b : M} : a ∣ p * b → p ∣ a ∨ a ∣ b := by - rintro ⟨c, hc⟩ - rcases hp.2.2 a c (hc ▸ dvd_mul_right _ _) with (h | ⟨x, rfl⟩) - · exact Or.inl h - · rw [mul_left_comm, mul_right_inj' hp.ne_zero] at hc - exact Or.inr (hc.symm ▸ dvd_mul_right _ _) - -theorem Prime.pow_dvd_of_dvd_mul_left [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) - (n : ℕ) (h : ¬p ∣ a) (h' : p ^ n ∣ a * b) : p ^ n ∣ b := by - induction n with - | zero => - rw [pow_zero] - exact one_dvd b - | succ n ih => - obtain ⟨c, rfl⟩ := ih (dvd_trans (pow_dvd_pow p n.le_succ) h') - rw [pow_succ] - apply mul_dvd_mul_left _ ((hp.dvd_or_dvd _).resolve_left h) - rwa [← mul_dvd_mul_iff_left (pow_ne_zero n hp.ne_zero), ← pow_succ, mul_left_comm] - -theorem Prime.pow_dvd_of_dvd_mul_right [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) - (n : ℕ) (h : ¬p ∣ b) (h' : p ^ n ∣ a * b) : p ^ n ∣ a := by - rw [mul_comm] at h' - exact hp.pow_dvd_of_dvd_mul_left n h h' - -theorem Prime.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd [CancelCommMonoidWithZero M] {p a b : M} - {n : ℕ} (hp : Prime p) (hpow : p ^ n.succ ∣ a ^ n.succ * b ^ n) (hb : ¬p ^ 2 ∣ b) : p ∣ a := by - -- Suppose `p ∣ b`, write `b = p * x` and `hy : a ^ n.succ * b ^ n = p ^ n.succ * y`. - rcases hp.dvd_or_dvd ((dvd_pow_self p (Nat.succ_ne_zero n)).trans hpow) with H | hbdiv - · exact hp.dvd_of_dvd_pow H - obtain ⟨x, rfl⟩ := hp.dvd_of_dvd_pow hbdiv - obtain ⟨y, hy⟩ := hpow - -- Then we can divide out a common factor of `p ^ n` from the equation `hy`. - have : a ^ n.succ * x ^ n = p * y := by - refine mul_left_cancel₀ (pow_ne_zero n hp.ne_zero) ?_ - rw [← mul_assoc _ p, ← pow_succ, ← hy, mul_pow, ← mul_assoc (a ^ n.succ), mul_comm _ (p ^ n), - mul_assoc] - -- So `p ∣ a` (and we're done) or `p ∣ x`, which can't be the case since it implies `p^2 ∣ b`. - refine hp.dvd_of_dvd_pow ((hp.dvd_or_dvd ⟨_, this⟩).resolve_right fun hdvdx => hb ?_) - obtain ⟨z, rfl⟩ := hp.dvd_of_dvd_pow hdvdx - rw [pow_two, ← mul_assoc] - exact dvd_mul_right _ _ - -theorem prime_pow_succ_dvd_mul {M : Type*} [CancelCommMonoidWithZero M] {p x y : M} (h : Prime p) - {i : ℕ} (hxy : p ^ (i + 1) ∣ x * y) : p ^ (i + 1) ∣ x ∨ p ∣ y := by - rw [or_iff_not_imp_right] - intro hy - induction i generalizing x with - | zero => rw [pow_one] at hxy ⊢; exact (h.dvd_or_dvd hxy).resolve_right hy - | succ i ih => - rw [pow_succ'] at hxy ⊢ - obtain ⟨x', rfl⟩ := (h.dvd_or_dvd (dvd_of_mul_right_dvd hxy)).resolve_right hy - rw [mul_assoc] at hxy - exact mul_dvd_mul_left p (ih ((mul_dvd_mul_iff_left h.ne_zero).mp hxy)) - -/-- `Irreducible p` states that `p` is non-unit and only factors into units. - -We explicitly avoid stating that `p` is non-zero, this would require a semiring. Assuming only a -monoid allows us to reuse irreducible for associated elements. --/ -structure Irreducible [Monoid M] (p : M) : Prop where - /-- `p` is not a unit -/ - not_unit : ¬IsUnit p - /-- if `p` factors then one factor is a unit -/ - isUnit_or_isUnit' : ∀ a b, p = a * b → IsUnit a ∨ IsUnit b - -namespace Irreducible - -theorem not_dvd_one [CommMonoid M] {p : M} (hp : Irreducible p) : ¬p ∣ 1 := - mt (isUnit_of_dvd_one ·) hp.not_unit - -theorem isUnit_or_isUnit [Monoid M] {p : M} (hp : Irreducible p) {a b : M} (h : p = a * b) : - IsUnit a ∨ IsUnit b := - hp.isUnit_or_isUnit' a b h - -end Irreducible - -theorem irreducible_iff [Monoid M] {p : M} : - Irreducible p ↔ ¬IsUnit p ∧ ∀ a b, p = a * b → IsUnit a ∨ IsUnit b := - ⟨fun h => ⟨h.1, h.2⟩, fun h => ⟨h.1, h.2⟩⟩ - -@[simp] -theorem not_irreducible_one [Monoid M] : ¬Irreducible (1 : M) := by simp [irreducible_iff] - -theorem Irreducible.ne_one [Monoid M] : ∀ {p : M}, Irreducible p → p ≠ 1 - | _, hp, rfl => not_irreducible_one hp - -@[simp] -theorem not_irreducible_zero [MonoidWithZero M] : ¬Irreducible (0 : M) - | ⟨hn0, h⟩ => - have : IsUnit (0 : M) ∨ IsUnit (0 : M) := h 0 0 (mul_zero 0).symm - this.elim hn0 hn0 - -theorem Irreducible.ne_zero [MonoidWithZero M] : ∀ {p : M}, Irreducible p → p ≠ 0 - | _, hp, rfl => not_irreducible_zero hp - -theorem of_irreducible_mul {M} [Monoid M] {x y : M} : Irreducible (x * y) → IsUnit x ∨ IsUnit y - | ⟨_, h⟩ => h _ _ rfl - -theorem not_irreducible_pow {M} [Monoid M] {x : M} {n : ℕ} (hn : n ≠ 1) : - ¬ Irreducible (x ^ n) := by - cases n with - | zero => simp - | succ n => - intro ⟨h₁, h₂⟩ - have := h₂ _ _ (pow_succ _ _) - rw [isUnit_pow_iff (Nat.succ_ne_succ.mp hn), or_self] at this - exact h₁ (this.pow _) - -theorem irreducible_or_factor {M} [Monoid M] (x : M) (h : ¬IsUnit x) : - Irreducible x ∨ ∃ a b, ¬IsUnit a ∧ ¬IsUnit b ∧ a * b = x := by - haveI := Classical.dec - refine or_iff_not_imp_right.2 fun H => ?_ - simp? [h, irreducible_iff] at H ⊢ says - simp only [exists_and_left, not_exists, not_and, irreducible_iff, h, not_false_eq_true, - true_and] at H ⊢ - refine fun a b h => by_contradiction fun o => ?_ - simp? [not_or] at o says simp only [not_or] at o - exact H _ o.1 _ o.2 h.symm - -/-- If `p` and `q` are irreducible, then `p ∣ q` implies `q ∣ p`. -/ -theorem Irreducible.dvd_symm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : - p ∣ q → q ∣ p := by - rintro ⟨q', rfl⟩ - rw [IsUnit.mul_right_dvd (Or.resolve_left (of_irreducible_mul hq) hp.not_unit)] - -theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : - p ∣ q ↔ q ∣ p := - ⟨hp.dvd_symm hq, hq.dvd_symm hp⟩ - -theorem Irreducible.of_map {F : Type*} [Monoid M] [Monoid N] [FunLike F M N] [MonoidHomClass F M N] - {f : F} [IsLocalHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := - ⟨fun hu ↦ hfx.not_unit <| hu.map f, - by rintro p q rfl - exact (hfx.isUnit_or_isUnit <| map_mul f p q).imp (.of_map f _) (.of_map f _)⟩ - -section - -variable [Monoid M] - -theorem irreducible_units_mul (a : Mˣ) (b : M) : Irreducible (↑a * b) ↔ Irreducible b := by - simp only [irreducible_iff, Units.isUnit_units_mul, and_congr_right_iff] - refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ - · rw [← a.isUnit_units_mul] - apply h - rw [mul_assoc, ← HAB] - · rw [← a⁻¹.isUnit_units_mul] - apply h - rw [mul_assoc, ← HAB, Units.inv_mul_cancel_left] - -theorem irreducible_isUnit_mul {a b : M} (h : IsUnit a) : Irreducible (a * b) ↔ Irreducible b := - let ⟨a, ha⟩ := h - ha ▸ irreducible_units_mul a b - -theorem irreducible_mul_units (a : Mˣ) (b : M) : Irreducible (b * ↑a) ↔ Irreducible b := by - simp only [irreducible_iff, Units.isUnit_mul_units, and_congr_right_iff] - refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ - · rw [← Units.isUnit_mul_units B a] - apply h - rw [← mul_assoc, ← HAB] - · rw [← Units.isUnit_mul_units B a⁻¹] - apply h - rw [← mul_assoc, ← HAB, Units.mul_inv_cancel_right] - -theorem irreducible_mul_isUnit {a b : M} (h : IsUnit a) : Irreducible (b * a) ↔ Irreducible b := - let ⟨a, ha⟩ := h - ha ▸ irreducible_mul_units a b - -theorem irreducible_mul_iff {a b : M} : - Irreducible (a * b) ↔ Irreducible a ∧ IsUnit b ∨ Irreducible b ∧ IsUnit a := by - constructor - · refine fun h => Or.imp (fun h' => ⟨?_, h'⟩) (fun h' => ⟨?_, h'⟩) (h.isUnit_or_isUnit rfl).symm - · rwa [irreducible_mul_isUnit h'] at h - · rwa [irreducible_isUnit_mul h'] at h - · rintro (⟨ha, hb⟩ | ⟨hb, ha⟩) - · rwa [irreducible_mul_isUnit hb] - · rwa [irreducible_isUnit_mul ha] - -variable [Monoid N] {F : Type*} [EquivLike F M N] [MulEquivClass F M N] (f : F) - -open MulEquiv - -/-- -Irreducibility is preserved by multiplicative equivalences. -Note that surjective + local hom is not enough. Consider the additive monoids `M = ℕ ⊕ ℕ`, `N = ℕ`, -with a surjective local (additive) hom `f : M →+ N` sending `(m, n)` to `2m + n`. -It is local because the only add unit in `N` is `0`, with preimage `{(0, 0)}` also an add unit. -Then `x = (1, 0)` is irreducible in `M`, but `f x = 2 = 1 + 1` is not irreducible in `N`. --/ -theorem Irreducible.map {x : M} (h : Irreducible x) : Irreducible (f x) := - ⟨fun g ↦ h.not_unit g.of_map, fun a b g ↦ - let f := MulEquivClass.toMulEquiv f - (h.isUnit_or_isUnit (symm_apply_apply f x ▸ map_mul f.symm a b ▸ congrArg f.symm g)).imp - (·.of_map) (·.of_map)⟩ - -theorem MulEquiv.irreducible_iff (f : F) {a : M} : - Irreducible (f a) ↔ Irreducible a := - ⟨Irreducible.of_map, Irreducible.map f⟩ - -end - -section CommMonoid - -variable [CommMonoid M] {a : M} - -theorem Irreducible.not_square (ha : Irreducible a) : ¬IsSquare a := by - rw [isSquare_iff_exists_sq] - rintro ⟨b, rfl⟩ - exact not_irreducible_pow (by decide) ha - -theorem IsSquare.not_irreducible (ha : IsSquare a) : ¬Irreducible a := fun h => h.not_square ha - -end CommMonoid - -section CommMonoidWithZero - -variable [CommMonoidWithZero M] - -theorem Irreducible.prime_of_isPrimal {a : M} - (irr : Irreducible a) (primal : IsPrimal a) : Prime a := - ⟨irr.ne_zero, irr.not_unit, fun a b dvd ↦ by - obtain ⟨d₁, d₂, h₁, h₂, rfl⟩ := primal dvd - exact (of_irreducible_mul irr).symm.imp (·.mul_right_dvd.mpr h₁) (·.mul_left_dvd.mpr h₂)⟩ - -theorem Irreducible.prime [DecompositionMonoid M] {a : M} (irr : Irreducible a) : Prime a := - irr.prime_of_isPrimal (DecompositionMonoid.primal a) - -end CommMonoidWithZero - -section CancelCommMonoidWithZero - -variable [CancelCommMonoidWithZero M] {a p : M} - -protected theorem Prime.irreducible (hp : Prime p) : Irreducible p := - ⟨hp.not_unit, fun a b ↦ by - rintro rfl - exact (hp.dvd_or_dvd dvd_rfl).symm.imp - (isUnit_of_dvd_one <| (mul_dvd_mul_iff_right <| right_ne_zero_of_mul hp.ne_zero).mp <| - dvd_mul_of_dvd_right · _) - (isUnit_of_dvd_one <| (mul_dvd_mul_iff_left <| left_ne_zero_of_mul hp.ne_zero).mp <| - dvd_mul_of_dvd_left · _)⟩ - -theorem irreducible_iff_prime [DecompositionMonoid M] {a : M} : Irreducible a ↔ Prime a := - ⟨Irreducible.prime, Prime.irreducible⟩ - -theorem succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul (hp : Prime p) {a b : M} {k l : ℕ} : - p ^ k ∣ a → p ^ l ∣ b → p ^ (k + l + 1) ∣ a * b → p ^ (k + 1) ∣ a ∨ p ^ (l + 1) ∣ b := - fun ⟨x, hx⟩ ⟨y, hy⟩ ⟨z, hz⟩ => - have h : p ^ (k + l) * (x * y) = p ^ (k + l) * (p * z) := by - simpa [mul_comm, pow_add, hx, hy, mul_assoc, mul_left_comm] using hz - have hp0 : p ^ (k + l) ≠ 0 := pow_ne_zero _ hp.ne_zero - have hpd : p ∣ x * y := ⟨z, by rwa [mul_right_inj' hp0] at h⟩ - (hp.dvd_or_dvd hpd).elim - (fun ⟨d, hd⟩ => Or.inl ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩) - fun ⟨d, hd⟩ => Or.inr ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩ - -theorem Prime.not_square (hp : Prime p) : ¬IsSquare p := - hp.irreducible.not_square - -theorem IsSquare.not_prime (ha : IsSquare a) : ¬Prime a := fun h => h.not_square ha - -theorem not_prime_pow {n : ℕ} (hn : n ≠ 1) : ¬Prime (a ^ n) := fun hp => - not_irreducible_pow hn hp.irreducible - -end CancelCommMonoidWithZero - /-- Two elements of a `Monoid` are `Associated` if one of them is another one multiplied by a unit on the right. -/ def Associated [Monoid M] (x y : M) : Prop := @@ -1107,19 +749,6 @@ end Associates section CommMonoidWithZero -theorem DvdNotUnit.isUnit_of_irreducible_right [CommMonoidWithZero M] {p q : M} - (h : DvdNotUnit p q) (hq : Irreducible q) : IsUnit p := by - obtain ⟨_, x, hx, hx'⟩ := h - exact Or.resolve_right ((irreducible_iff.1 hq).right p x hx') hx - -theorem not_irreducible_of_not_unit_dvdNotUnit [CommMonoidWithZero M] {p q : M} (hp : ¬IsUnit p) - (h : DvdNotUnit p q) : ¬Irreducible q := - mt h.isUnit_of_irreducible_right hp - -theorem DvdNotUnit.not_unit [CommMonoidWithZero M] {p q : M} (hp : DvdNotUnit p q) : ¬IsUnit q := by - obtain ⟨-, x, hx, rfl⟩ := hp - exact fun hc => hx (isUnit_iff_dvd_one.mpr (dvd_of_mul_left_dvd (isUnit_iff_dvd_one.mp hc))) - theorem dvdNotUnit_of_dvdNotUnit_associated [CommMonoidWithZero M] [Nontrivial M] {p q r : M} (h : DvdNotUnit p q) (h' : Associated q r) : DvdNotUnit p r := by obtain ⟨u, rfl⟩ := Associated.symm h' @@ -1144,26 +773,6 @@ theorem DvdNotUnit.not_associated [CancelCommMonoidWithZero M] {p q : M} (h : Dv rcases (mul_right_inj' hp).mp hx' with rfl exact hx a.isUnit -theorem DvdNotUnit.ne [CancelCommMonoidWithZero M] {p q : M} (h : DvdNotUnit p q) : p ≠ q := by - by_contra hcontra - obtain ⟨hp, x, hx', hx''⟩ := h - conv_lhs at hx'' => rw [← hcontra, ← mul_one p] - rw [(mul_left_cancel₀ hp hx'').symm] at hx' - exact hx' isUnit_one - -theorem pow_injective_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) - (hq' : q ≠ 0) : Function.Injective fun n : ℕ => q ^ n := by - refine injective_of_lt_imp_ne fun n m h => DvdNotUnit.ne ⟨pow_ne_zero n hq', q ^ (m - n), ?_, ?_⟩ - · exact not_isUnit_of_not_isUnit_dvd hq (dvd_pow (dvd_refl _) (Nat.sub_pos_of_lt h).ne') - · exact (pow_mul_pow_sub q h.le).symm - -@[deprecated (since := "2024-09-22")] -alias pow_injective_of_not_unit := pow_injective_of_not_isUnit - -theorem pow_inj_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) - (hq' : q ≠ 0) {m n : ℕ} : q ^ m = q ^ n ↔ m = n := - (pow_injective_of_not_isUnit hq hq').eq_iff - theorem dvd_prime_pow [CancelCommMonoidWithZero M] {p q : M} (hp : Prime p) (n : ℕ) : q ∣ p ^ n ↔ ∃ i ≤ n, Associated q (p ^ i) := by induction n generalizing q with diff --git a/Mathlib/Algebra/Prime/Defs.lean b/Mathlib/Algebra/Prime/Defs.lean new file mode 100644 index 0000000000000..90165b72d0050 --- /dev/null +++ b/Mathlib/Algebra/Prime/Defs.lean @@ -0,0 +1,191 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker +-/ +import Mathlib.Algebra.GroupWithZero.Divisibility + +/-! +# Prime and irreducible elements. + +In this file we define the predicate `Prime p` +saying that an element of a commutative monoid with zero is prime. +Namely, `Prime p` means that `p` isn't zero, it isn't a unit, +and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`; + +In decomposition monoids (e.g., `ℕ`, `ℤ`), this predicate is equivalent to `Irreducible` +(see `irreducible_iff_prime`), however this is not true in general. + +## Main definitions + * `Prime`: a prime element of a commutative monoid with zero + * `Irreducible`: an irreducible element of a commutative monoid with zero + +## Main results + * `irreducible_iff_prime`: the two definitions are equivalent in a decomposition monoid. +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists Multiset + +variable {M N : Type*} + +section Prime + +variable [CommMonoidWithZero M] + +/-- An element `p` of a commutative monoid with zero (e.g., a ring) is called *prime*, +if it's not zero, not a unit, and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`. -/ +def Prime (p : M) : Prop := + p ≠ 0 ∧ ¬IsUnit p ∧ ∀ a b, p ∣ a * b → p ∣ a ∨ p ∣ b + +namespace Prime + +variable {p : M} (hp : Prime p) +include hp + +theorem ne_zero : p ≠ 0 := + hp.1 + +theorem not_unit : ¬IsUnit p := + hp.2.1 + +theorem not_dvd_one : ¬p ∣ 1 := + mt (isUnit_of_dvd_one ·) hp.not_unit + +theorem ne_one : p ≠ 1 := fun h => hp.2.1 (h.symm ▸ isUnit_one) + +theorem dvd_or_dvd {a b : M} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := + hp.2.2 a b h + +theorem dvd_mul {a b : M} : p ∣ a * b ↔ p ∣ a ∨ p ∣ b := + ⟨hp.dvd_or_dvd, (Or.elim · (dvd_mul_of_dvd_left · _) (dvd_mul_of_dvd_right · _))⟩ + +theorem isPrimal : IsPrimal p := fun _a _b dvd ↦ (hp.dvd_or_dvd dvd).elim + (fun h ↦ ⟨p, 1, h, one_dvd _, (mul_one p).symm⟩) fun h ↦ ⟨1, p, one_dvd _, h, (one_mul p).symm⟩ + +theorem not_dvd_mul {a b : M} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a * b := + hp.dvd_mul.not.mpr <| not_or.mpr ⟨ha, hb⟩ + +theorem dvd_of_dvd_pow {a : M} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by + induction n with + | zero => + rw [pow_zero] at h + have := isUnit_of_dvd_one h + have := not_unit hp + contradiction + | succ n ih => + rw [pow_succ'] at h + rcases dvd_or_dvd hp h with dvd_a | dvd_pow + · assumption + · exact ih dvd_pow + +theorem dvd_pow_iff_dvd {a : M} {n : ℕ} (hn : n ≠ 0) : p ∣ a ^ n ↔ p ∣ a := + ⟨hp.dvd_of_dvd_pow, (dvd_pow · hn)⟩ + +end Prime + +@[simp] +theorem not_prime_zero : ¬Prime (0 : M) := fun h => h.ne_zero rfl + +@[simp] +theorem not_prime_one : ¬Prime (1 : M) := fun h => h.not_unit isUnit_one + +end Prime + +/-- `Irreducible p` states that `p` is non-unit and only factors into units. + +We explicitly avoid stating that `p` is non-zero, this would require a semiring. Assuming only a +monoid allows us to reuse irreducible for associated elements. +-/ +structure Irreducible [Monoid M] (p : M) : Prop where + /-- `p` is not a unit -/ + not_unit : ¬IsUnit p + /-- if `p` factors then one factor is a unit -/ + isUnit_or_isUnit' : ∀ a b, p = a * b → IsUnit a ∨ IsUnit b + +namespace Irreducible + +theorem not_dvd_one [CommMonoid M] {p : M} (hp : Irreducible p) : ¬p ∣ 1 := + mt (isUnit_of_dvd_one ·) hp.not_unit + +theorem isUnit_or_isUnit [Monoid M] {p : M} (hp : Irreducible p) {a b : M} (h : p = a * b) : + IsUnit a ∨ IsUnit b := + hp.isUnit_or_isUnit' a b h + +end Irreducible + +theorem irreducible_iff [Monoid M] {p : M} : + Irreducible p ↔ ¬IsUnit p ∧ ∀ a b, p = a * b → IsUnit a ∨ IsUnit b := + ⟨fun h => ⟨h.1, h.2⟩, fun h => ⟨h.1, h.2⟩⟩ + +@[simp] +theorem not_irreducible_one [Monoid M] : ¬Irreducible (1 : M) := by simp [irreducible_iff] + +theorem Irreducible.ne_one [Monoid M] : ∀ {p : M}, Irreducible p → p ≠ 1 + | _, hp, rfl => not_irreducible_one hp + +@[simp] +theorem not_irreducible_zero [MonoidWithZero M] : ¬Irreducible (0 : M) + | ⟨hn0, h⟩ => + have : IsUnit (0 : M) ∨ IsUnit (0 : M) := h 0 0 (mul_zero 0).symm + this.elim hn0 hn0 + +theorem Irreducible.ne_zero [MonoidWithZero M] : ∀ {p : M}, Irreducible p → p ≠ 0 + | _, hp, rfl => not_irreducible_zero hp + +theorem of_irreducible_mul {M} [Monoid M] {x y : M} : Irreducible (x * y) → IsUnit x ∨ IsUnit y + | ⟨_, h⟩ => h _ _ rfl + +theorem irreducible_or_factor {M} [Monoid M] (x : M) (h : ¬IsUnit x) : + Irreducible x ∨ ∃ a b, ¬IsUnit a ∧ ¬IsUnit b ∧ a * b = x := by + haveI := Classical.dec + refine or_iff_not_imp_right.2 fun H => ?_ + simp? [h, irreducible_iff] at H ⊢ says + simp only [exists_and_left, not_exists, not_and, irreducible_iff, h, not_false_eq_true, + true_and] at H ⊢ + refine fun a b h => by_contradiction fun o => ?_ + simp? [not_or] at o says simp only [not_or] at o + exact H _ o.1 _ o.2 h.symm + +/-- If `p` and `q` are irreducible, then `p ∣ q` implies `q ∣ p`. -/ +theorem Irreducible.dvd_symm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : + p ∣ q → q ∣ p := by + rintro ⟨q', rfl⟩ + rw [IsUnit.mul_right_dvd (Or.resolve_left (of_irreducible_mul hq) hp.not_unit)] + +theorem Irreducible.dvd_comm [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) : + p ∣ q ↔ q ∣ p := + ⟨hp.dvd_symm hq, hq.dvd_symm hp⟩ + +section CommMonoidWithZero + +variable [CommMonoidWithZero M] + +theorem Irreducible.prime_of_isPrimal {a : M} + (irr : Irreducible a) (primal : IsPrimal a) : Prime a := + ⟨irr.ne_zero, irr.not_unit, fun a b dvd ↦ by + obtain ⟨d₁, d₂, h₁, h₂, rfl⟩ := primal dvd + exact (of_irreducible_mul irr).symm.imp (·.mul_right_dvd.mpr h₁) (·.mul_left_dvd.mpr h₂)⟩ + +theorem Irreducible.prime [DecompositionMonoid M] {a : M} (irr : Irreducible a) : Prime a := + irr.prime_of_isPrimal (DecompositionMonoid.primal a) + +end CommMonoidWithZero + +section CancelCommMonoidWithZero + +variable [CancelCommMonoidWithZero M] {a p : M} + +protected theorem Prime.irreducible (hp : Prime p) : Irreducible p := + ⟨hp.not_unit, fun a b ↦ by + rintro rfl + exact (hp.dvd_or_dvd dvd_rfl).symm.imp + (isUnit_of_dvd_one <| (mul_dvd_mul_iff_right <| right_ne_zero_of_mul hp.ne_zero).mp <| + dvd_mul_of_dvd_right · _) + (isUnit_of_dvd_one <| (mul_dvd_mul_iff_left <| left_ne_zero_of_mul hp.ne_zero).mp <| + dvd_mul_of_dvd_left · _)⟩ + +theorem irreducible_iff_prime [DecompositionMonoid M] {a : M} : Irreducible a ↔ Prime a := + ⟨Irreducible.prime, Prime.irreducible⟩ + +end CancelCommMonoidWithZero diff --git a/Mathlib/Algebra/Prime/Lemmas.lean b/Mathlib/Algebra/Prime/Lemmas.lean new file mode 100644 index 0000000000000..96ee461f3831f --- /dev/null +++ b/Mathlib/Algebra/Prime/Lemmas.lean @@ -0,0 +1,276 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jens Wagemaker +-/ +import Mathlib.Algebra.Group.Commute.Units +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Group.Units.Equiv +import Mathlib.Algebra.GroupWithZero.Hom +import Mathlib.Algebra.Prime.Defs +import Mathlib.Order.Lattice + +/-! +# Associated, prime, and irreducible elements. + +In this file we define the predicate `Prime p` +saying that an element of a commutative monoid with zero is prime. +Namely, `Prime p` means that `p` isn't zero, it isn't a unit, +and `p ∣ a * b → p ∣ a ∨ p ∣ b` for all `a`, `b`; + +In decomposition monoids (e.g., `ℕ`, `ℤ`), this predicate is equivalent to `Irreducible`, +however this is not true in general. + +We also define an equivalence relation `Associated` +saying that two elements of a monoid differ by a multiplication by a unit. +Then we show that the quotient type `Associates` is a monoid +and prove basic properties of this quotient. +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists Multiset + +variable {M N : Type*} + +section Prime + +variable [CommMonoidWithZero M] + +section Map + +variable [CommMonoidWithZero N] {F : Type*} {G : Type*} [FunLike F M N] +variable [MonoidWithZeroHomClass F M N] [FunLike G N M] [MulHomClass G N M] +variable (f : F) (g : G) {p : M} + +theorem comap_prime (hinv : ∀ a, g (f a : N) = a) (hp : Prime (f p)) : Prime p := + ⟨fun h => hp.1 <| by simp [h], fun h => hp.2.1 <| h.map f, fun a b h => by + refine + (hp.2.2 (f a) (f b) <| by + convert map_dvd f h + simp).imp + ?_ ?_ <;> + · intro h + convert ← map_dvd g h <;> apply hinv⟩ + +theorem MulEquiv.prime_iff (e : M ≃* N) : Prime p ↔ Prime (e p) := + ⟨fun h => (comap_prime e.symm e fun a => by simp) <| (e.symm_apply_apply p).substr h, + comap_prime e e.symm fun a => by simp⟩ + +end Map + +end Prime + +theorem Prime.left_dvd_or_dvd_right_of_dvd_mul [CancelCommMonoidWithZero M] {p : M} (hp : Prime p) + {a b : M} : a ∣ p * b → p ∣ a ∨ a ∣ b := by + rintro ⟨c, hc⟩ + rcases hp.2.2 a c (hc ▸ dvd_mul_right _ _) with (h | ⟨x, rfl⟩) + · exact Or.inl h + · rw [mul_left_comm, mul_right_inj' hp.ne_zero] at hc + exact Or.inr (hc.symm ▸ dvd_mul_right _ _) + +theorem Prime.pow_dvd_of_dvd_mul_left [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) + (n : ℕ) (h : ¬p ∣ a) (h' : p ^ n ∣ a * b) : p ^ n ∣ b := by + induction n with + | zero => + rw [pow_zero] + exact one_dvd b + | succ n ih => + obtain ⟨c, rfl⟩ := ih (dvd_trans (pow_dvd_pow p n.le_succ) h') + rw [pow_succ] + apply mul_dvd_mul_left _ ((hp.dvd_or_dvd _).resolve_left h) + rwa [← mul_dvd_mul_iff_left (pow_ne_zero n hp.ne_zero), ← pow_succ, mul_left_comm] + +theorem Prime.pow_dvd_of_dvd_mul_right [CancelCommMonoidWithZero M] {p a b : M} (hp : Prime p) + (n : ℕ) (h : ¬p ∣ b) (h' : p ^ n ∣ a * b) : p ^ n ∣ a := by + rw [mul_comm] at h' + exact hp.pow_dvd_of_dvd_mul_left n h h' + +theorem Prime.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd [CancelCommMonoidWithZero M] {p a b : M} + {n : ℕ} (hp : Prime p) (hpow : p ^ n.succ ∣ a ^ n.succ * b ^ n) (hb : ¬p ^ 2 ∣ b) : p ∣ a := by + -- Suppose `p ∣ b`, write `b = p * x` and `hy : a ^ n.succ * b ^ n = p ^ n.succ * y`. + rcases hp.dvd_or_dvd ((dvd_pow_self p (Nat.succ_ne_zero n)).trans hpow) with H | hbdiv + · exact hp.dvd_of_dvd_pow H + obtain ⟨x, rfl⟩ := hp.dvd_of_dvd_pow hbdiv + obtain ⟨y, hy⟩ := hpow + -- Then we can divide out a common factor of `p ^ n` from the equation `hy`. + have : a ^ n.succ * x ^ n = p * y := by + refine mul_left_cancel₀ (pow_ne_zero n hp.ne_zero) ?_ + rw [← mul_assoc _ p, ← pow_succ, ← hy, mul_pow, ← mul_assoc (a ^ n.succ), mul_comm _ (p ^ n), + mul_assoc] + -- So `p ∣ a` (and we're done) or `p ∣ x`, which can't be the case since it implies `p^2 ∣ b`. + refine hp.dvd_of_dvd_pow ((hp.dvd_or_dvd ⟨_, this⟩).resolve_right fun hdvdx => hb ?_) + obtain ⟨z, rfl⟩ := hp.dvd_of_dvd_pow hdvdx + rw [pow_two, ← mul_assoc] + exact dvd_mul_right _ _ + +theorem prime_pow_succ_dvd_mul {M : Type*} [CancelCommMonoidWithZero M] {p x y : M} (h : Prime p) + {i : ℕ} (hxy : p ^ (i + 1) ∣ x * y) : p ^ (i + 1) ∣ x ∨ p ∣ y := by + rw [or_iff_not_imp_right] + intro hy + induction i generalizing x with + | zero => rw [pow_one] at hxy ⊢; exact (h.dvd_or_dvd hxy).resolve_right hy + | succ i ih => + rw [pow_succ'] at hxy ⊢ + obtain ⟨x', rfl⟩ := (h.dvd_or_dvd (dvd_of_mul_right_dvd hxy)).resolve_right hy + rw [mul_assoc] at hxy + exact mul_dvd_mul_left p (ih ((mul_dvd_mul_iff_left h.ne_zero).mp hxy)) + +theorem not_irreducible_pow {M} [Monoid M] {x : M} {n : ℕ} (hn : n ≠ 1) : + ¬ Irreducible (x ^ n) := by + cases n with + | zero => simp + | succ n => + intro ⟨h₁, h₂⟩ + have := h₂ _ _ (pow_succ _ _) + rw [isUnit_pow_iff (Nat.succ_ne_succ.mp hn), or_self] at this + exact h₁ (this.pow _) + +theorem Irreducible.of_map {F : Type*} [Monoid M] [Monoid N] [FunLike F M N] [MonoidHomClass F M N] + {f : F} [IsLocalHom f] {x} (hfx : Irreducible (f x)) : Irreducible x := + ⟨fun hu ↦ hfx.not_unit <| hu.map f, + by rintro p q rfl + exact (hfx.isUnit_or_isUnit <| map_mul f p q).imp (.of_map f _) (.of_map f _)⟩ + +section + +variable [Monoid M] + +theorem irreducible_units_mul (a : Mˣ) (b : M) : Irreducible (↑a * b) ↔ Irreducible b := by + simp only [irreducible_iff, Units.isUnit_units_mul, and_congr_right_iff] + refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ + · rw [← a.isUnit_units_mul] + apply h + rw [mul_assoc, ← HAB] + · rw [← a⁻¹.isUnit_units_mul] + apply h + rw [mul_assoc, ← HAB, Units.inv_mul_cancel_left] + +theorem irreducible_isUnit_mul {a b : M} (h : IsUnit a) : Irreducible (a * b) ↔ Irreducible b := + let ⟨a, ha⟩ := h + ha ▸ irreducible_units_mul a b + +theorem irreducible_mul_units (a : Mˣ) (b : M) : Irreducible (b * ↑a) ↔ Irreducible b := by + simp only [irreducible_iff, Units.isUnit_mul_units, and_congr_right_iff] + refine fun _ => ⟨fun h A B HAB => ?_, fun h A B HAB => ?_⟩ + · rw [← Units.isUnit_mul_units B a] + apply h + rw [← mul_assoc, ← HAB] + · rw [← Units.isUnit_mul_units B a⁻¹] + apply h + rw [← mul_assoc, ← HAB, Units.mul_inv_cancel_right] + +theorem irreducible_mul_isUnit {a b : M} (h : IsUnit a) : Irreducible (b * a) ↔ Irreducible b := + let ⟨a, ha⟩ := h + ha ▸ irreducible_mul_units a b + +theorem irreducible_mul_iff {a b : M} : + Irreducible (a * b) ↔ Irreducible a ∧ IsUnit b ∨ Irreducible b ∧ IsUnit a := by + constructor + · refine fun h => Or.imp (fun h' => ⟨?_, h'⟩) (fun h' => ⟨?_, h'⟩) (h.isUnit_or_isUnit rfl).symm + · rwa [irreducible_mul_isUnit h'] at h + · rwa [irreducible_isUnit_mul h'] at h + · rintro (⟨ha, hb⟩ | ⟨hb, ha⟩) + · rwa [irreducible_mul_isUnit hb] + · rwa [irreducible_isUnit_mul ha] + +variable [Monoid N] {F : Type*} [EquivLike F M N] [MulEquivClass F M N] (f : F) + +open MulEquiv + +/-- +Irreducibility is preserved by multiplicative equivalences. +Note that surjective + local hom is not enough. Consider the additive monoids `M = ℕ ⊕ ℕ`, `N = ℕ`, +with a surjective local (additive) hom `f : M →+ N` sending `(m, n)` to `2m + n`. +It is local because the only add unit in `N` is `0`, with preimage `{(0, 0)}` also an add unit. +Then `x = (1, 0)` is irreducible in `M`, but `f x = 2 = 1 + 1` is not irreducible in `N`. +-/ +theorem Irreducible.map {x : M} (h : Irreducible x) : Irreducible (f x) := + ⟨fun g ↦ h.not_unit g.of_map, fun a b g ↦ + let f := MulEquivClass.toMulEquiv f + (h.isUnit_or_isUnit (symm_apply_apply f x ▸ map_mul f.symm a b ▸ congrArg f.symm g)).imp + (·.of_map) (·.of_map)⟩ + +theorem MulEquiv.irreducible_iff (f : F) {a : M} : + Irreducible (f a) ↔ Irreducible a := + ⟨Irreducible.of_map, Irreducible.map f⟩ + +end + +section CommMonoid + +variable [CommMonoid M] {a : M} + +theorem Irreducible.not_square (ha : Irreducible a) : ¬IsSquare a := by + rw [isSquare_iff_exists_sq] + rintro ⟨b, rfl⟩ + exact not_irreducible_pow (by decide) ha + +theorem IsSquare.not_irreducible (ha : IsSquare a) : ¬Irreducible a := fun h => h.not_square ha + +end CommMonoid + +section CancelCommMonoidWithZero + +variable [CancelCommMonoidWithZero M] {a p : M} + +theorem succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul (hp : Prime p) {a b : M} {k l : ℕ} : + p ^ k ∣ a → p ^ l ∣ b → p ^ (k + l + 1) ∣ a * b → p ^ (k + 1) ∣ a ∨ p ^ (l + 1) ∣ b := + fun ⟨x, hx⟩ ⟨y, hy⟩ ⟨z, hz⟩ => + have h : p ^ (k + l) * (x * y) = p ^ (k + l) * (p * z) := by + simpa [mul_comm, pow_add, hx, hy, mul_assoc, mul_left_comm] using hz + have hp0 : p ^ (k + l) ≠ 0 := pow_ne_zero _ hp.ne_zero + have hpd : p ∣ x * y := ⟨z, by rwa [mul_right_inj' hp0] at h⟩ + (hp.dvd_or_dvd hpd).elim + (fun ⟨d, hd⟩ => Or.inl ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩) + fun ⟨d, hd⟩ => Or.inr ⟨d, by simp [*, pow_succ, mul_comm, mul_left_comm, mul_assoc]⟩ + +theorem Prime.not_square (hp : Prime p) : ¬IsSquare p := + hp.irreducible.not_square + +theorem IsSquare.not_prime (ha : IsSquare a) : ¬Prime a := fun h => h.not_square ha + +theorem not_prime_pow {n : ℕ} (hn : n ≠ 1) : ¬Prime (a ^ n) := fun hp => + not_irreducible_pow hn hp.irreducible + +end CancelCommMonoidWithZero + +section CommMonoidWithZero + +theorem DvdNotUnit.isUnit_of_irreducible_right [CommMonoidWithZero M] {p q : M} + (h : DvdNotUnit p q) (hq : Irreducible q) : IsUnit p := by + obtain ⟨_, x, hx, hx'⟩ := h + exact Or.resolve_right ((irreducible_iff.1 hq).right p x hx') hx + +theorem not_irreducible_of_not_unit_dvdNotUnit [CommMonoidWithZero M] {p q : M} (hp : ¬IsUnit p) + (h : DvdNotUnit p q) : ¬Irreducible q := + mt h.isUnit_of_irreducible_right hp + +theorem DvdNotUnit.not_unit [CommMonoidWithZero M] {p q : M} (hp : DvdNotUnit p q) : ¬IsUnit q := by + obtain ⟨-, x, hx, rfl⟩ := hp + exact fun hc => hx (isUnit_iff_dvd_one.mpr (dvd_of_mul_left_dvd (isUnit_iff_dvd_one.mp hc))) + +end CommMonoidWithZero + +section CancelCommMonoidWithZero + +theorem DvdNotUnit.ne [CancelCommMonoidWithZero M] {p q : M} (h : DvdNotUnit p q) : p ≠ q := by + by_contra hcontra + obtain ⟨hp, x, hx', hx''⟩ := h + conv_lhs at hx'' => rw [← hcontra, ← mul_one p] + rw [(mul_left_cancel₀ hp hx'').symm] at hx' + exact hx' isUnit_one + +theorem pow_injective_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) + (hq' : q ≠ 0) : Function.Injective fun n : ℕ => q ^ n := by + refine injective_of_lt_imp_ne fun n m h => DvdNotUnit.ne ⟨pow_ne_zero n hq', q ^ (m - n), ?_, ?_⟩ + · exact not_isUnit_of_not_isUnit_dvd hq (dvd_pow (dvd_refl _) (Nat.sub_pos_of_lt h).ne') + · exact (pow_mul_pow_sub q h.le).symm + +@[deprecated (since := "2024-09-22")] +alias pow_injective_of_not_unit := pow_injective_of_not_isUnit + +theorem pow_inj_of_not_isUnit [CancelCommMonoidWithZero M] {q : M} (hq : ¬IsUnit q) + (hq' : q ≠ 0) {m n : ℕ} : q ^ m = q ^ n ↔ m = n := + (pow_injective_of_not_isUnit hq hq').eq_iff + +end CancelCommMonoidWithZero diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index d1a2155b074ae..5a4b6a4fb3db7 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -3,11 +3,12 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ +import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.Ring.Int -import Mathlib.Order.Bounds.Basic import Mathlib.Data.Nat.Factorial.Basic import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Order.Bounds.Basic /-! ## Notable Theorems diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 638f139f80dd6..1741083417a66 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd -import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.Prime.Lemmas import Mathlib.Algebra.Ring.Parity /-! diff --git a/Mathlib/RingTheory/Prime.lean b/Mathlib/RingTheory/Prime.lean index 6fb8a1e0598db..331460fc59ec6 100644 --- a/Mathlib/RingTheory/Prime.lean +++ b/Mathlib/RingTheory/Prime.lean @@ -3,10 +3,11 @@ Copyright (c) 2020 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Algebra.Order.Group.Unbundled.Abs +import Mathlib.Algebra.Prime.Defs +import Mathlib.Algebra.Ring.Units /-! # Prime elements in rings From a1cf0e3bef13bfbc210c074e07fd7bc9b78852d6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 25 Oct 2024 13:43:26 +0000 Subject: [PATCH 18/29] feat: Matrices over finite free modules are finite modules (#18098) The titular result is trivial, but to state the lemmas about the basis this uses, we need to generalize many results about `stdBasisMatrix`. --- Mathlib/Data/Matrix/Basis.lean | 53 ++++++++++++------- .../Dimension/Constructions.lean | 40 +++++++++----- .../FreeModule/Finite/Basic.lean | 12 +++-- Mathlib/LinearAlgebra/Matrix/Ideal.lean | 2 +- Mathlib/LinearAlgebra/Matrix/ToLin.lean | 2 +- Mathlib/LinearAlgebra/StdBasis.lean | 20 +++++++ Mathlib/RingTheory/MatrixAlgebra.lean | 6 +-- 7 files changed, 94 insertions(+), 41 deletions(-) diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 4b4af6da0eab3..23333213c522e 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -22,13 +22,15 @@ namespace Matrix open Matrix variable [DecidableEq l] [DecidableEq m] [DecidableEq n] -variable [Semiring α] + +section Zero +variable [Zero α] /-- `stdBasisMatrix i j a` is the matrix with `a` in the `i`-th row, `j`-th column, and zeroes elsewhere. -/ -def stdBasisMatrix (i : m) (j : n) (a : α) : Matrix m n α := fun i' j' => - if i = i' ∧ j = j' then a else 0 +def stdBasisMatrix (i : m) (j : n) (a : α) : Matrix m n α := + of <| fun i' j' => if i = i' ∧ j = j' then a else 0 theorem stdBasisMatrix_eq_of_single_single (i : m) (j : n) (a : α) : stdBasisMatrix i j a = Matrix.of (Pi.single i (Pi.single j a)) := by @@ -49,12 +51,16 @@ theorem stdBasisMatrix_zero (i : m) (j : n) : stdBasisMatrix i j (0 : α) = 0 := ext simp -theorem stdBasisMatrix_add (i : m) (j : n) (a b : α) : +end Zero + +theorem stdBasisMatrix_add [AddZeroClass α] (i : m) (j : n) (a b : α) : stdBasisMatrix i j (a + b) = stdBasisMatrix i j a + stdBasisMatrix i j b := by - unfold stdBasisMatrix; ext + ext + simp only [stdBasisMatrix, of_apply] split_ifs with h <;> simp [h] -theorem mulVec_stdBasisMatrix [Fintype m] (i : n) (j : m) (c : α) (x : m → α) : +theorem mulVec_stdBasisMatrix [NonUnitalNonAssocSemiring α] [Fintype m] + (i : n) (j : m) (c : α) (x : m → α) : mulVec (stdBasisMatrix i j c) x = Function.update (0 : n → α) i (c * x j) := by ext i' simp [stdBasisMatrix, mulVec, dotProduct] @@ -62,7 +68,7 @@ theorem mulVec_stdBasisMatrix [Fintype m] (i : n) (j : m) (c : α) (x : m → α · simp simp [h, h.symm] -theorem matrix_eq_sum_stdBasisMatrix [Fintype m] [Fintype n] (x : Matrix m n α) : +theorem matrix_eq_sum_stdBasisMatrix [AddCommMonoid α] [Fintype m] [Fintype n] (x : Matrix m n α) : x = ∑ i : m, ∑ j : n, stdBasisMatrix i j (x i j) := by ext i j; symm iterate 2 rw [Finset.sum_apply] @@ -75,7 +81,7 @@ theorem matrix_eq_sum_stdBasisMatrix [Fintype m] [Fintype n] (x : Matrix m n α) @[deprecated (since := "2024-08-11")] alias matrix_eq_sum_std_basis := matrix_eq_sum_stdBasisMatrix -theorem stdBasisMatrix_eq_single_vecMulVec_single (i : m) (j : n) : +theorem stdBasisMatrix_eq_single_vecMulVec_single [MulZeroOneClass α] (i : m) (j : n) : stdBasisMatrix i j (1 : α) = vecMulVec (Pi.single i 1) (Pi.single j 1) := by ext i' j' -- Porting note: lean3 didn't apply `mul_ite`. @@ -84,7 +90,7 @@ theorem stdBasisMatrix_eq_single_vecMulVec_single (i : m) (j : n) : -- TODO: tie this up with the `Basis` machinery of linear algebra -- this is not completely trivial because we are indexing by two types, instead of one @[deprecated stdBasisMatrix_eq_single_vecMulVec_single (since := "2024-08-11")] -theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : +theorem std_basis_eq_basis_mul_basis [MulZeroOneClass α] (i : m) (j : n) : stdBasisMatrix i j (1 : α) = vecMulVec (fun i' => ite (i = i') 1 0) fun j' => ite (j = j') 1 0 := by rw [stdBasisMatrix_eq_single_vecMulVec_single] @@ -92,7 +98,8 @@ theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : -- todo: the old proof used fintypes, I don't know `Finsupp` but this feels generalizable @[elab_as_elim] -protected theorem induction_on' [Finite m] [Finite n] {P : Matrix m n α → Prop} (M : Matrix m n α) +protected theorem induction_on' + [AddCommMonoid α] [Finite m] [Finite n] {P : Matrix m n α → Prop} (M : Matrix m n α) (h_zero : P 0) (h_add : ∀ p q, P p → P q → P (p + q)) (h_std_basis : ∀ (i : m) (j : n) (x : α), P (stdBasisMatrix i j x)) : P M := by cases nonempty_fintype m; cases nonempty_fintype n @@ -102,7 +109,8 @@ protected theorem induction_on' [Finite m] [Finite n] {P : Matrix m n α → Pro apply h_std_basis @[elab_as_elim] -protected theorem induction_on [Finite m] [Finite n] [Nonempty m] [Nonempty n] +protected theorem induction_on + [AddCommMonoid α] [Finite m] [Finite n] [Nonempty m] [Nonempty n] {P : Matrix m n α → Prop} (M : Matrix m n α) (h_add : ∀ p q, P p → P q → P (p + q)) (h_std_basis : ∀ i j x, P (stdBasisMatrix i j x)) : P M := Matrix.induction_on' M @@ -116,7 +124,7 @@ namespace StdBasisMatrix section -variable (i : m) (j : n) (c : α) (i' : m) (j' : n) +variable [Zero α] (i : m) (j : n) (c : α) (i' : m) (j' : n) @[simp] theorem apply_same : stdBasisMatrix i j c i j = c := @@ -124,7 +132,7 @@ theorem apply_same : stdBasisMatrix i j c i j = c := @[simp] theorem apply_of_ne (h : ¬(i = i' ∧ j = j')) : stdBasisMatrix i j c i' j' = 0 := by - simp only [stdBasisMatrix, and_imp, ite_eq_right_iff] + simp only [stdBasisMatrix, and_imp, ite_eq_right_iff, of_apply] tauto @[simp] @@ -138,8 +146,7 @@ theorem apply_of_col_ne (i i' : m) {j j' : n} (hj : j ≠ j') (a : α) : end section - -variable (i j : n) (c : α) (i' j' : n) +variable [Zero α] (i j : n) (c : α) @[simp] theorem diag_zero (h : j ≠ i) : diag (stdBasisMatrix i j c) = 0 := @@ -150,7 +157,10 @@ theorem diag_same : diag (stdBasisMatrix i i c) = Pi.single i c := by ext j by_cases hij : i = j <;> (try rw [hij]) <;> simp [hij] -variable [Fintype n] +end + +section trace +variable [Fintype n] [AddCommMonoid α] (i j : n) (c : α) @[simp] theorem trace_zero (h : j ≠ i) : trace (stdBasisMatrix i j c) = 0 := by @@ -162,6 +172,11 @@ theorem trace_eq : trace (stdBasisMatrix i i c) = c := by -- Porting note: added `-diag_apply` simp [trace, -diag_apply] +end trace + +section mul +variable [Fintype n] [NonUnitalNonAssocSemiring α] (i j : n) (c : α) + @[simp] theorem mul_left_apply_same (b : n) (M : Matrix n n α) : (stdBasisMatrix i j c * M) i b = c * M j b := by simp [mul_apply, stdBasisMatrix] @@ -189,7 +204,7 @@ theorem mul_same (k : n) (d : α) : theorem mul_of_ne {k l : n} (h : j ≠ k) (d : α) : stdBasisMatrix i j c * stdBasisMatrix k l d = 0 := by ext a b - simp only [mul_apply, boole_mul, stdBasisMatrix] + simp only [mul_apply, boole_mul, stdBasisMatrix, of_apply] by_cases h₁ : i = a -- porting note (#10745): was `simp [h₁, h, h.symm]` · simp only [h₁, true_and, mul_ite, ite_mul, zero_mul, mul_zero, ← ite_and, zero_apply] @@ -200,13 +215,13 @@ theorem mul_of_ne {k l : n} (h : j ≠ k) (d : α) : · simp only [h₁, false_and, ite_false, mul_ite, zero_mul, mul_zero, ite_self, Finset.sum_const_zero, zero_apply] -end +end mul end StdBasisMatrix section Commute -variable [Fintype n] +variable [Fintype n] [Semiring α] theorem row_eq_zero_of_commute_stdBasisMatrix {i j k : n} {M : Matrix n n α} (hM : Commute (stdBasisMatrix i j 1) M) (hkj : k ≠ j) : M j k = 0 := by diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index d03d8646c1451..fb4695c2deb46 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -183,25 +183,41 @@ theorem rank_directSum {ι : Type v} (M : ι → Type w) [∀ i : ι, AddCommGro let b : Basis _ R (⨁ i, M i) := DFinsupp.basis fun i => B i simp [← b.mk_eq_rank'', fun i => (B i).mk_eq_rank''] -/-- If `m` and `n` are `Fintype`, the rank of `m × n` matrices is `(#m).lift * (#n).lift`. -/ +/-- If `m` and `n` are finite, the rank of `m × n` matrices over a module `M` is +`(#m).lift * (#n).lift * rank R M`. -/ @[simp] +theorem rank_matrix_module (m : Type w) (n : Type w') [Finite m] [Finite n] : + Module.rank R (Matrix m n M) = + lift.{max v w'} #m * lift.{max v w} #n * lift.{max w w'} (Module.rank R M) := by + cases nonempty_fintype m + cases nonempty_fintype n + obtain ⟨I, b⟩ := Module.Free.exists_basis (R := R) (M := M) + rw [← (b.matrix m n).mk_eq_rank''] + simp only [mk_prod, lift_mul, lift_lift, ← mul_assoc, b.mk_eq_rank''] + + +/-- If `m` and `n` are finite and lie in the same universe, the rank of `m × n` matrices over a +module `M` is `(#m * #n).lift * rank R M`. -/ +@[simp high] +theorem rank_matrix_module' (m n : Type w) [Finite m] [Finite n] : + Module.rank R (Matrix m n M) = + lift.{max v} (#m * #n) * lift.{w} (Module.rank R M) := by + rw [rank_matrix_module, lift_mul, lift_umax.{w, v}] + +/-- If `m` and `n` are finite, the rank of `m × n` matrices is `(#m).lift * (#n).lift`. -/ theorem rank_matrix (m : Type v) (n : Type w) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = Cardinal.lift.{max v w u, v} #m * Cardinal.lift.{max v w u, w} #n := by - cases nonempty_fintype m - cases nonempty_fintype n - have h := (Matrix.stdBasis R m n).mk_eq_rank - rw [← lift_lift.{max v w u, max v w}, lift_inj] at h - simpa using h.symm + rw [rank_matrix_module, rank_self, lift_one, mul_one, ← lift_lift.{v, max u w}, lift_id, + ← lift_lift.{w, max u v}, lift_id] -/-- If `m` and `n` are `Fintype` that lie in the same universe, the rank of `m × n` matrices is +/-- If `m` and `n` are finite and lie in the same universe, the rank of `m × n` matrices is `(#n * #m).lift`. -/ -@[simp high] theorem rank_matrix' (m n : Type v) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = Cardinal.lift.{u} (#m * #n) := by rw [rank_matrix, lift_mul, lift_umax.{v, u}] -/-- If `m` and `n` are `Fintype` that lie in the same universe as `R`, the rank of `m × n` matrices +/-- If `m` and `n` are finite and lie in the same universe as `R`, the rank of `m × n` matrices is `# m * # n`. -/ theorem rank_matrix'' (m n : Type u) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = #m * #n := by simp @@ -228,10 +244,10 @@ theorem finrank_directSum {ι : Type v} [Fintype ι] (M : ι → Type w) [∀ i simp only [finrank, fun i => rank_eq_card_chooseBasisIndex R (M i), rank_directSum, ← mk_sigma, mk_toNat_eq_card, card_sigma] -/-- If `m` and `n` are `Fintype`, the finrank of `m × n` matrices is - `(Fintype.card m) * (Fintype.card n)`. -/ +/-- If `m` and `n` are `Fintype`, the finrank of `m × n` matrices over a module `M` is + `(Fintype.card m) * (Fintype.card n) * finrank R M`. -/ theorem finrank_matrix (m n : Type*) [Fintype m] [Fintype n] : - finrank R (Matrix m n R) = card m * card n := by simp [finrank] + finrank R (Matrix m n M) = card m * card n * finrank R M := by simp [finrank] end Module diff --git a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean index 7cc74f0c2a909..d9da43ef8b1b4 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Finite/Basic.lean @@ -39,9 +39,13 @@ theorem Module.Finite.of_basis {R M ι : Type*} [Semiring R] [AddCommMonoid M] [ refine ⟨⟨Finset.univ.image b, ?_⟩⟩ simp only [Set.image_univ, Finset.coe_univ, Finset.coe_image, Basis.span_eq] -instance Module.Finite.matrix {R : Type u} [Semiring R] - {ι₁ ι₂ : Type*} [_root_.Finite ι₁] [_root_.Finite ι₂] : - Module.Finite R (Matrix ι₁ ι₂ R) := by +instance Module.Finite.matrix {R ι₁ ι₂ M : Type*} + [Semiring R] [AddCommMonoid M] [Module R M] [Module.Free R M] [Module.Finite R M] + [_root_.Finite ι₁] [_root_.Finite ι₂] : + Module.Finite R (Matrix ι₁ ι₂ M) := by cases nonempty_fintype ι₁ cases nonempty_fintype ι₂ - exact Module.Finite.of_basis (Pi.basis fun _ => Pi.basisFun R _) + exact Module.Finite.of_basis <| (Free.chooseBasis _ _).matrix _ _ + +example {ι₁ ι₂ R : Type*} [Semiring R] [Finite ι₁] [Finite ι₂] : + Module.Finite R (Matrix ι₁ ι₂ R) := inferInstance diff --git a/Mathlib/LinearAlgebra/Matrix/Ideal.lean b/Mathlib/LinearAlgebra/Matrix/Ideal.lean index 2d0f2ed989b04..b6372378bf74b 100644 --- a/Mathlib/LinearAlgebra/Matrix/Ideal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Ideal.lean @@ -159,7 +159,7 @@ def equivMatricesOver (i j : n) : TwoSidedIdeal R ≃ TwoSidedIdeal (Matrix n n by_cases hab : a = k ∧ b = l · rcases hab with ⟨ha, hb⟩ subst ha hb - simp only [stdBasisMatrix, and_self, ↓reduceIte, StdBasisMatrix.mul_right_apply_same, + simp only [StdBasisMatrix.apply_same, StdBasisMatrix.mul_right_apply_same, StdBasisMatrix.mul_left_apply_same, one_mul, mul_one] rw [hy2 a b] · conv_lhs => diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index de60812076e72..693066a10400f 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -960,7 +960,7 @@ lemma linearMap_apply_apply (ij : ι₂ × ι₁) (k : ι₁) : (b₁.linearMap b₂ ij) (b₁ k) = if ij.2 = k then b₂ ij.1 else 0 := by have := Classical.decEq ι₂ rw [linearMap_apply, Matrix.stdBasis_eq_stdBasisMatrix, Matrix.toLin_self] - dsimp only [Matrix.stdBasisMatrix] + dsimp only [Matrix.stdBasisMatrix, of_apply] simp_rw [ite_smul, one_smul, zero_smul, ite_and, Finset.sum_ite_eq, Finset.mem_univ, if_true] /-- The standard basis of the endomorphism algebra of a module diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index 3159915328c21..8896d7312a535 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -287,6 +287,26 @@ lemma piEquiv_apply_apply (ι R M : Type*) [Fintype ι] [CommSemiring R] end Module +namespace Basis +variable {ι R M : Type*} (m n : Type*) +variable [Fintype m] [Fintype n] [Semiring R] [AddCommMonoid M] [Module R M] + +/-- The standard basis of `Matrix m n M` given a basis on `M`. -/ +protected noncomputable def matrix (b : Basis ι R M) : + Basis (m × n × ι) R (Matrix m n M) := + Basis.reindex (Pi.basis fun _ : m => Pi.basis fun _ : n => b) + ((Equiv.sigmaEquivProd _ _).trans <| .prodCongr (.refl _) (Equiv.sigmaEquivProd _ _)) + |>.map (Matrix.ofLinearEquiv R) + +variable {n m} + +@[simp] +theorem matrix_apply (b : Basis ι R M) (i : m) (j : n) (k : ι) [DecidableEq m] [DecidableEq n] : + b.matrix m n (i, j, k) = Matrix.stdBasisMatrix i j (b k) := by + simp [Basis.matrix, Matrix.stdBasisMatrix_eq_of_single_single] + +end Basis + namespace Matrix variable (R : Type*) (m n : Type*) [Fintype m] [Finite n] [Semiring R] diff --git a/Mathlib/RingTheory/MatrixAlgebra.lean b/Mathlib/RingTheory/MatrixAlgebra.lean index 4a71967b89d79..0750d23424344 100644 --- a/Mathlib/RingTheory/MatrixAlgebra.lean +++ b/Mathlib/RingTheory/MatrixAlgebra.lean @@ -99,13 +99,11 @@ theorem invFun_algebraMap (M : Matrix n n R) : invFun R A n (M.map (algebraMap R convert Finset.sum_product (β := Matrix n n R) ..; simp theorem right_inv (M : Matrix n n A) : (toFunAlgHom R A n) (invFun R A n M) = M := by - simp only [invFun, map_sum, stdBasisMatrix, apply_ite ↑(algebraMap R A), smul_eq_mul, - mul_boole, toFunAlgHom_apply, RingHom.map_zero, RingHom.map_one, Matrix.map_apply, - Pi.smul_def] + simp only [invFun, map_sum, toFunAlgHom_apply] convert Finset.sum_product (β := Matrix n n A) .. conv_lhs => rw [matrix_eq_sum_stdBasisMatrix M] refine Finset.sum_congr rfl fun i _ => Finset.sum_congr rfl fun j _ => Matrix.ext fun a b => ?_ - simp only [stdBasisMatrix, smul_apply, Matrix.map_apply] + dsimp [stdBasisMatrix] split_ifs <;> aesop theorem left_inv (M : A ⊗[R] Matrix n n R) : invFun R A n (toFunAlgHom R A n M) = M := by From 27e4d0a1b1a1869163126f8e0fdeddd4daf5dbed Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Fri, 25 Oct 2024 14:17:18 +0000 Subject: [PATCH 19/29] feat(Topology/Algebra/Valued/NormedValued): add lemmas for `Valued.toNormedField` (#16758) Add simp lemmas for `Valued.toNormedField`. Prove `isNonarchimedean` for this norm. Add definition `Valued.toNontriviallyNormedField`. --- .../Topology/Algebra/Valued/NormedValued.lean | 50 +++++++++++++++++++ 1 file changed, 50 insertions(+) diff --git a/Mathlib/Topology/Algebra/Valued/NormedValued.lean b/Mathlib/Topology/Algebra/Valued/NormedValued.lean index 249d5250b9868..19c82cd591ba9 100644 --- a/Mathlib/Topology/Algebra/Valued/NormedValued.lean +++ b/Mathlib/Topology/Algebra/Valued/NormedValued.lean @@ -160,4 +160,54 @@ lemma coe_valuation_eq_rankOne_hom_comp_valuation : ⇑NormedField.valuation = h end NormedField +variable {L} {Γ₀} + +namespace toNormedField + +variable {x x' : L} + +@[simp] +theorem norm_le_iff : ‖x‖ ≤ ‖x'‖ ↔ val.v x ≤ val.v x' := + (Valuation.RankOne.strictMono val.v).le_iff_le + +@[simp] +theorem norm_lt_iff : ‖x‖ < ‖x'‖ ↔ val.v x < val.v x' := + (Valuation.RankOne.strictMono val.v).lt_iff_lt + +@[simp] +theorem norm_le_one_iff : ‖x‖ ≤ 1 ↔ val.v x ≤ 1 := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).le_iff_le (b := 1) + +@[simp] +theorem norm_lt_one_iff : ‖x‖ < 1 ↔ val.v x < 1 := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).lt_iff_lt (b := 1) + +@[simp] +theorem one_le_norm_iff : 1 ≤ ‖x‖ ↔ 1 ≤ val.v x := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).le_iff_le (a := 1) + +@[simp] +theorem one_lt_norm_iff : 1 < ‖x‖ ↔ 1 < val.v x := by + simpa only [_root_.map_one] using (Valuation.RankOne.strictMono val.v).lt_iff_lt (a := 1) + +end toNormedField + +/-- +The nontrivially normed field structure determined by a rank one valuation. +-/ +def toNontriviallyNormedField: NontriviallyNormedField L := { + val.toNormedField with + non_trivial := by + obtain ⟨x, hx⟩ := Valuation.RankOne.nontrivial val.v + rcases Valuation.val_le_one_or_val_inv_le_one val.v x with h | h + · use x⁻¹ + simp only [toNormedField.one_lt_norm_iff, map_inv₀, one_lt_inv₀ (zero_lt_iff.mpr hx.1), + lt_of_le_of_ne h hx.2] + · use x + simp only [map_inv₀, inv_le_one₀ <| zero_lt_iff.mpr hx.1] at h + simp only [toNormedField.one_lt_norm_iff, lt_of_le_of_ne h hx.2.symm] +} + +scoped[Valued] attribute [instance] Valued.toNontriviallyNormedField + end Valued From a4411e690f209860046c5e29ee95aac2addf7218 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 25 Oct 2024 14:41:59 +0000 Subject: [PATCH 20/29] chore: remove unused variables (#18226) #17715 --- Mathlib/LinearAlgebra/Dimension/Localization.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Circulant.lean | 2 +- Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean | 2 +- Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean | 2 +- Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean | 2 +- .../Function/ConditionalExpectation/AEMeasurable.lean | 4 +--- .../MeasureTheory/Measure/WithDensityVectorMeasure.lean | 4 ++-- Mathlib/ModelTheory/Graph.lean | 4 ++-- Mathlib/ModelTheory/PartialEquiv.lean | 4 ++-- Mathlib/NumberTheory/Cyclotomic/Gal.lean | 2 +- Mathlib/NumberTheory/Dioph.lean | 4 ++-- Mathlib/Order/Ideal.lean | 8 ++++---- Mathlib/RingTheory/DedekindDomain/Dvr.lean | 2 +- Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean | 2 +- Mathlib/RingTheory/PowerSeries/Basic.lean | 2 +- Mathlib/RingTheory/PowerSeries/Inverse.lean | 2 +- 16 files changed, 23 insertions(+), 25 deletions(-) diff --git a/Mathlib/LinearAlgebra/Dimension/Localization.lean b/Mathlib/LinearAlgebra/Dimension/Localization.lean index b6445fadfe529..594d62bba3e1d 100644 --- a/Mathlib/LinearAlgebra/Dimension/Localization.lean +++ b/Mathlib/LinearAlgebra/Dimension/Localization.lean @@ -120,7 +120,7 @@ end CommRing section Ring -variable {R} [Ring R] [IsDomain R] (S : Submonoid R) +variable {R} [Ring R] [IsDomain R] /-- A domain that is not (left) Ore is of infinite rank. See [cohn_1995] Proposition 1.3.6 -/ diff --git a/Mathlib/LinearAlgebra/Matrix/Circulant.lean b/Mathlib/LinearAlgebra/Matrix/Circulant.lean index c81affc351505..2e87e67a1fa42 100644 --- a/Mathlib/LinearAlgebra/Matrix/Circulant.lean +++ b/Mathlib/LinearAlgebra/Matrix/Circulant.lean @@ -32,7 +32,7 @@ circulant, matrix -/ -variable {α β m n R : Type*} +variable {α β n R : Type*} namespace Matrix diff --git a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean index b339fa3a131cc..ae27ada2167b0 100644 --- a/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean @@ -25,7 +25,7 @@ open Set Function Module noncomputable section -variable {ι R M N : Type*} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] +variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] namespace LinearMap diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index 83a31d388715d..133145051b3f2 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -88,7 +88,7 @@ lemma rnDeriv_pos' [HaveLebesgueDecomposition ν μ] [SigmaFinite μ] (hμν : section rnDeriv_withDensity_leftRight -variable {μ ν : Measure α} {f : α → ℝ≥0∞} +variable {f : α → ℝ≥0∞} /-- Auxiliary lemma for `rnDeriv_withDensity_left`. -/ lemma rnDeriv_withDensity_withDensity_rnDeriv_left (μ ν : Measure α) [SigmaFinite μ] [SigmaFinite ν] diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index a74a832aa9956..7652993309eaf 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -46,7 +46,7 @@ open scoped Classical MeasureTheory NNReal ENNReal open Set -variable {α β : Type*} {m : MeasurableSpace α} {μ ν : MeasureTheory.Measure α} +variable {α : Type*} {m : MeasurableSpace α} {μ : MeasureTheory.Measure α} namespace MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean index bf52c95cfaeaa..95319082eb996 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean @@ -175,12 +175,10 @@ theorem AEStronglyMeasurable'.aeStronglyMeasurable'_of_measurableSpace_le_on {α hf_ind.stronglyMeasurable_of_measurableSpace_le_on hs_m hs fun x hxs => Set.indicator_of_not_mem hxs _ -variable {α F F' 𝕜 : Type*} {p : ℝ≥0∞} [RCLike 𝕜] +variable {α F 𝕜 : Type*} {p : ℝ≥0∞} [RCLike 𝕜] -- 𝕜 for ℝ or ℂ -- F for a Lp submodule [NormedAddCommGroup F] [NormedSpace 𝕜 F] - -- F' for integrals on a Lp submodule - [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] [NormedSpace ℝ F'] section LpMeas diff --git a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean b/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean index 674e34c9ca6f3..095d95dd608a0 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean @@ -26,13 +26,13 @@ noncomputable section open scoped MeasureTheory NNReal ENNReal -variable {α β : Type*} {m : MeasurableSpace α} +variable {α : Type*} {m : MeasurableSpace α} namespace MeasureTheory open TopologicalSpace -variable {μ ν : Measure α} +variable {μ : Measure α} variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] open Classical in diff --git a/Mathlib/ModelTheory/Graph.lean b/Mathlib/ModelTheory/Graph.lean index c4ce433290e9c..ed3a509c17c55 100644 --- a/Mathlib/ModelTheory/Graph.lean +++ b/Mathlib/ModelTheory/Graph.lean @@ -21,7 +21,7 @@ This file defines first-order languages, structures, and theories in graph theor of the theory of simple graphs. -/ -universe u v +universe u namespace FirstOrder @@ -31,7 +31,7 @@ open FirstOrder open Structure -variable {α : Type u} {V : Type v} {n : ℕ} +variable {V : Type u} {n : ℕ} /-! ### Simple Graphs -/ diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean index 3db2f33308397..49e73c8f359f2 100644 --- a/Mathlib/ModelTheory/PartialEquiv.lean +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -40,8 +40,8 @@ namespace FirstOrder namespace Language -variable (L : Language.{u, v}) (M : Type w) (N : Type w') {P : Type*} -variable [L.Structure M] [L.Structure N] [L.Structure P] +variable (L : Language.{u, v}) (M : Type w) (N : Type w') +variable [L.Structure M] [L.Structure N] open FirstOrder Structure Substructure diff --git a/Mathlib/NumberTheory/Cyclotomic/Gal.lean b/Mathlib/NumberTheory/Cyclotomic/Gal.lean index fb381459d36cf..c5f27c52696f2 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Gal.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Gal.lean @@ -152,7 +152,7 @@ end IsCyclotomicExtension section Gal -variable [Field L] (hμ : IsPrimitiveRoot μ n) [Algebra K L] [IsCyclotomicExtension {n} K L] +variable [Field L] [Algebra K L] [IsCyclotomicExtension {n} K L] (h : Irreducible (cyclotomic n K)) {K} /-- `IsCyclotomicExtension.autEquivPow` repackaged in terms of `Gal`. diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index 721b81b5271ba..6c7c378d55eb4 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -67,7 +67,7 @@ Note that this duplicates `MvPolynomial`. section Polynomials -variable {α β γ : Type*} +variable {α β : Type*} /-- A predicate asserting that a function is a multivariate integer polynomial. (We are being a bit lazy here by allowing many representations for multiplication, @@ -432,7 +432,7 @@ end section -variable {α β : Type} {n : ℕ} +variable {α : Type} {n : ℕ} open Vector3 diff --git a/Mathlib/Order/Ideal.lean b/Mathlib/Order/Ideal.lean index 4dae232039074..7c8fe0e4ef08f 100644 --- a/Mathlib/Order/Ideal.lean +++ b/Mathlib/Order/Ideal.lean @@ -92,7 +92,7 @@ variable [LE P] section -variable {I J s t : Ideal P} {x y : P} +variable {I s t : Ideal P} {x : P} theorem toLowerSet_injective : Injective (toLowerSet : Ideal P → LowerSet P) := fun s t _ ↦ by cases s @@ -247,7 +247,7 @@ variable [Preorder P] section -variable {I J : Ideal P} {x y : P} +variable {I : Ideal P} {x y : P} /-- The smallest ideal containing a given element. -/ @[simps] @@ -316,7 +316,7 @@ end SemilatticeSup section SemilatticeSupDirected -variable [SemilatticeSup P] [IsDirected P (· ≥ ·)] {x : P} {I J K s t : Ideal P} +variable [SemilatticeSup P] [IsDirected P (· ≥ ·)] {x : P} {I J s t : Ideal P} /-- The infimum of two ideals of a co-directed order is their intersection. -/ instance : Inf (Ideal P) := @@ -385,7 +385,7 @@ end SemilatticeSupDirected section SemilatticeSupOrderBot -variable [SemilatticeSup P] [OrderBot P] {x : P} {I J K : Ideal P} +variable [SemilatticeSup P] [OrderBot P] {x : P} instance : InfSet (Ideal P) := ⟨fun S ↦ diff --git a/Mathlib/RingTheory/DedekindDomain/Dvr.lean b/Mathlib/RingTheory/DedekindDomain/Dvr.lean index 2edc66bfc7bfa..c7086a988650b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Dvr.lean +++ b/Mathlib/RingTheory/DedekindDomain/Dvr.lean @@ -43,7 +43,7 @@ dedekind domain, dedekind ring -/ -variable (R A K : Type*) [CommRing R] [CommRing A] [IsDomain A] [Field K] +variable (A : Type*) [CommRing A] [IsDomain A] open scoped nonZeroDivisors Polynomial diff --git a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean index 5957f090f0b89..763ee85bc69d5 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/TFAE.lean @@ -25,7 +25,7 @@ Also see `tfae_of_isNoetherianRing_of_localRing_of_isDomain` for a version witho -/ -variable (R : Type*) [CommRing R] (K : Type*) [Field K] [Algebra R K] [IsFractionRing R K] +variable (R : Type*) [CommRing R] open scoped Multiplicative diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 65b589c7373de..8965ee6ddf242 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -758,7 +758,7 @@ namespace Polynomial open Finsupp Polynomial -variable {σ : Type*} {R : Type*} [CommSemiring R] (φ ψ : R[X]) +variable {R : Type*} [CommSemiring R] (φ ψ : R[X]) -- Porting note: added so we can add the `@[coe]` attribute /-- The natural inclusion from polynomials into formal power series. -/ diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index 18c6913c525cd..660923cc4da32 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -274,7 +274,7 @@ theorem map.isLocalHom : IsLocalHom (map f) := @[deprecated (since := "2024-10-10")] alias map.isLocalRingHom := map.isLocalHom -variable [LocalRing R] [LocalRing S] +variable [LocalRing R] instance : LocalRing R⟦X⟧ := { inferInstanceAs <| LocalRing <| MvPowerSeries Unit R with } From aea6f296a35d3b88812a1b424543e1da53822850 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Fri, 25 Oct 2024 14:58:57 +0000 Subject: [PATCH 21/29] chore: split `Limits/Shapes/Terminal.lean` to reduce imports (#18229) The parts that require `HasLimit` stay in `Terminal.lean`, and the parts that only require `IsLimit` go to `IsTerminal.lean`. Found using the new `lake exe unused`. Co-authored-by: Markus Himmel --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Adjunction/Comma.lean | 1 + .../CategoryTheory/Bicategory/Kan/HasKan.lean | 1 + Mathlib/CategoryTheory/Comma/Over.lean | 4 +- .../Comma/StructuredArrow/Basic.lean | 2 +- Mathlib/CategoryTheory/Elements.lean | 3 +- .../CategoryTheory/Endofunctor/Algebra.lean | 2 +- .../Limits/Shapes/IsTerminal.lean | 438 ++++++++++++++++++ .../Limits/Shapes/Terminal.lean | 384 +-------------- Mathlib/CategoryTheory/Monad/Limits.lean | 2 +- Mathlib/CategoryTheory/WithTerminal.lean | 2 +- 11 files changed, 448 insertions(+), 392 deletions(-) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean diff --git a/Mathlib.lean b/Mathlib.lean index d885b3af7da93..05b32bb4e3391 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1733,6 +1733,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts import Mathlib.CategoryTheory.Limits.Shapes.FunctorCategory import Mathlib.CategoryTheory.Limits.Shapes.FunctorToTypes import Mathlib.CategoryTheory.Limits.Shapes.Images +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal import Mathlib.CategoryTheory.Limits.Shapes.KernelPair import Mathlib.CategoryTheory.Limits.Shapes.Kernels import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer diff --git a/Mathlib/CategoryTheory/Adjunction/Comma.lean b/Mathlib/CategoryTheory/Adjunction/Comma.lean index 6f4474b55e705..4147a489b5b31 100644 --- a/Mathlib/CategoryTheory/Adjunction/Comma.lean +++ b/Mathlib/CategoryTheory/Adjunction/Comma.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ +import Mathlib.CategoryTheory.Limits.Shapes.Terminal import Mathlib.CategoryTheory.Adjunction.Basic import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic import Mathlib.CategoryTheory.PUnit diff --git a/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean b/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean index e81755e4324af..40043a409ecfe 100644 --- a/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean +++ b/Mathlib/CategoryTheory/Bicategory/Kan/HasKan.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Yuma Mizuno. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ +import Mathlib.CategoryTheory.Limits.Shapes.Terminal import Mathlib.CategoryTheory.Bicategory.Kan.IsKan /-! diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 2d0ed2bd61b9c..e37bf7ec99866 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -4,9 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Bhavik Mehta -/ import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic -import Mathlib.CategoryTheory.PUnit -import Mathlib.CategoryTheory.Functor.ReflectsIso -import Mathlib.CategoryTheory.Functor.EpiMono +import Mathlib.CategoryTheory.Category.Cat /-! # Over and under categories diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean index c0d1729054be6..7e4231ecc0c06 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow/Basic.lean @@ -5,7 +5,7 @@ Authors: Adam Topaz, Kim Morrison -/ import Mathlib.CategoryTheory.Comma.Basic import Mathlib.CategoryTheory.PUnit -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal /-! # The category of "structured arrows" diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 2c682ecb881f9..7d85416b30755 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.CategoryTheory.Comma.StructuredArrow.Basic -import Mathlib.CategoryTheory.Groupoid -import Mathlib.CategoryTheory.PUnit +import Mathlib.CategoryTheory.Category.Cat /-! # The category of elements diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index 59d61ad43c4c1..1b53aa148eb3a 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joseph Hua. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Bhavik Mehta, Johan Commelin, Reid Barton, Robert Y. Lewis, Joseph Hua -/ -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal /-! diff --git a/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean new file mode 100644 index 0000000000000..7ceb4a6aeb551 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/IsTerminal.lean @@ -0,0 +1,438 @@ +/- +Copyright (c) 2019 Kim Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison, Bhavik Mehta +-/ +import Mathlib.CategoryTheory.PEmpty +import Mathlib.CategoryTheory.Limits.IsLimit +import Mathlib.CategoryTheory.EpiMono +import Mathlib.CategoryTheory.Category.Preorder + +/-! +# Initial and terminal objects in a category. + +In this file we define the predicates `IsTerminal` and `IsInitial` as well as the class +`InitialMonoClass`. + +The classes `HasTerminal` and `HasInitial` and the associated notations for terminal and inital +objects are defined in `Terminal.lean`. + +## References +* [Stacks: Initial and final objects](https://stacks.math.columbia.edu/tag/002B) +-/ + +assert_not_exists CategoryTheory.Limits.HasLimit + +noncomputable section + +universe w w' v v₁ v₂ u u₁ u₂ + +open CategoryTheory + +namespace CategoryTheory.Limits + +variable {C : Type u₁} [Category.{v₁} C] + +/-- Construct a cone for the empty diagram given an object. -/ +@[simps] +def asEmptyCone (X : C) : Cone (Functor.empty.{0} C) := + { pt := X + π := + { app := by aesop_cat } } + +/-- Construct a cocone for the empty diagram given an object. -/ +@[simps] +def asEmptyCocone (X : C) : Cocone (Functor.empty.{0} C) := + { pt := X + ι := + { app := by aesop_cat } } + +/-- `X` is terminal if the cone it induces on the empty diagram is limiting. -/ +abbrev IsTerminal (X : C) := + IsLimit (asEmptyCone X) + +/-- `X` is initial if the cocone it induces on the empty diagram is colimiting. -/ +abbrev IsInitial (X : C) := + IsColimit (asEmptyCocone X) + +/-- An object `Y` is terminal iff for every `X` there is a unique morphism `X ⟶ Y`. -/ +def isTerminalEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (Y : C) : + IsLimit (⟨Y, by aesop_cat, by aesop_cat⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where + toFun t X := + { default := t.lift ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ + uniq := fun f => + t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + invFun u := + { lift := fun s => (u s.pt).default + uniq := fun s _ _ => (u s.pt).2 _ } + left_inv := by dsimp [Function.LeftInverse]; intro x; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.RightInverse,Function.LeftInverse] + intro u; funext X; simp only + +/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` + (as an instance). -/ +def IsTerminal.ofUnique (Y : C) [h : ∀ X : C, Unique (X ⟶ Y)] : IsTerminal Y where + lift s := (h s.pt).default + fac := fun _ ⟨j⟩ => j.elim + +/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` + (as explicit arguments). -/ +def IsTerminal.ofUniqueHom {Y : C} (h : ∀ X : C, X ⟶ Y) (uniq : ∀ (X : C) (m : X ⟶ Y), m = h X) : + IsTerminal Y := + have : ∀ X : C, Unique (X ⟶ Y) := fun X ↦ ⟨⟨h X⟩, uniq X⟩ + IsTerminal.ofUnique Y + +/-- If `α` is a preorder with top, then `⊤` is a terminal object. -/ +def isTerminalTop {α : Type*} [Preorder α] [OrderTop α] : IsTerminal (⊤ : α) := + IsTerminal.ofUnique _ + +/-- Transport a term of type `IsTerminal` across an isomorphism. -/ +def IsTerminal.ofIso {Y Z : C} (hY : IsTerminal Y) (i : Y ≅ Z) : IsTerminal Z := + IsLimit.ofIsoLimit hY + { hom := { hom := i.hom } + inv := { hom := i.inv } } + +/-- If `X` and `Y` are isomorphic, then `X` is terminal iff `Y` is. -/ +def IsTerminal.equivOfIso {X Y : C} (e : X ≅ Y) : + IsTerminal X ≃ IsTerminal Y where + toFun h := IsTerminal.ofIso h e + invFun h := IsTerminal.ofIso h e.symm + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- An object `X` is initial iff for every `Y` there is a unique morphism `X ⟶ Y`. -/ +def isInitialEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (X : C) : + IsColimit (⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where + toFun t X := + { default := t.desc ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ + uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } + invFun u := + { desc := fun s => (u s.pt).default + uniq := fun s _ _ => (u s.pt).2 _ } + left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.RightInverse,Function.LeftInverse] + intro; funext; simp only + +/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` + (as an instance). -/ +def IsInitial.ofUnique (X : C) [h : ∀ Y : C, Unique (X ⟶ Y)] : IsInitial X where + desc s := (h s.pt).default + fac := fun _ ⟨j⟩ => j.elim + +/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` + (as explicit arguments). -/ +def IsInitial.ofUniqueHom {X : C} (h : ∀ Y : C, X ⟶ Y) (uniq : ∀ (Y : C) (m : X ⟶ Y), m = h Y) : + IsInitial X := + have : ∀ Y : C, Unique (X ⟶ Y) := fun Y ↦ ⟨⟨h Y⟩, uniq Y⟩ + IsInitial.ofUnique X + +/-- If `α` is a preorder with bot, then `⊥` is an initial object. -/ +def isInitialBot {α : Type*} [Preorder α] [OrderBot α] : IsInitial (⊥ : α) := + IsInitial.ofUnique _ + +/-- Transport a term of type `is_initial` across an isomorphism. -/ +def IsInitial.ofIso {X Y : C} (hX : IsInitial X) (i : X ≅ Y) : IsInitial Y := + IsColimit.ofIsoColimit hX + { hom := { hom := i.hom } + inv := { hom := i.inv } } + +/-- If `X` and `Y` are isomorphic, then `X` is initial iff `Y` is. -/ +def IsInitial.equivOfIso {X Y : C} (e : X ≅ Y) : + IsInitial X ≃ IsInitial Y where + toFun h := IsInitial.ofIso h e + invFun h := IsInitial.ofIso h e.symm + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- Give the morphism to a terminal object from any other. -/ +def IsTerminal.from {X : C} (t : IsTerminal X) (Y : C) : Y ⟶ X := + t.lift (asEmptyCone Y) + +/-- Any two morphisms to a terminal object are equal. -/ +theorem IsTerminal.hom_ext {X Y : C} (t : IsTerminal X) (f g : Y ⟶ X) : f = g := + IsLimit.hom_ext t (by aesop_cat) + +@[simp] +theorem IsTerminal.comp_from {Z : C} (t : IsTerminal Z) {X Y : C} (f : X ⟶ Y) : + f ≫ t.from Y = t.from X := + t.hom_ext _ _ + +@[simp] +theorem IsTerminal.from_self {X : C} (t : IsTerminal X) : t.from X = 𝟙 X := + t.hom_ext _ _ + +/-- Give the morphism from an initial object to any other. -/ +def IsInitial.to {X : C} (t : IsInitial X) (Y : C) : X ⟶ Y := + t.desc (asEmptyCocone Y) + +/-- Any two morphisms from an initial object are equal. -/ +theorem IsInitial.hom_ext {X Y : C} (t : IsInitial X) (f g : X ⟶ Y) : f = g := + IsColimit.hom_ext t (by aesop_cat) + +@[simp] +theorem IsInitial.to_comp {X : C} (t : IsInitial X) {Y Z : C} (f : Y ⟶ Z) : t.to Y ≫ f = t.to Z := + t.hom_ext _ _ + +@[simp] +theorem IsInitial.to_self {X : C} (t : IsInitial X) : t.to X = 𝟙 X := + t.hom_ext _ _ + +/-- Any morphism from a terminal object is split mono. -/ +theorem IsTerminal.isSplitMono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : IsSplitMono f := + IsSplitMono.mk' ⟨t.from _, t.hom_ext _ _⟩ + +/-- Any morphism to an initial object is split epi. -/ +theorem IsInitial.isSplitEpi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : IsSplitEpi f := + IsSplitEpi.mk' ⟨t.to _, t.hom_ext _ _⟩ + +/-- Any morphism from a terminal object is mono. -/ +theorem IsTerminal.mono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : Mono f := by + haveI := t.isSplitMono_from f; infer_instance + +/-- Any morphism to an initial object is epi. -/ +theorem IsInitial.epi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : Epi f := by + haveI := t.isSplitEpi_to f; infer_instance + +/-- If `T` and `T'` are terminal, they are isomorphic. -/ +@[simps] +def IsTerminal.uniqueUpToIso {T T' : C} (hT : IsTerminal T) (hT' : IsTerminal T') : T ≅ T' where + hom := hT'.from _ + inv := hT.from _ + +/-- If `I` and `I'` are initial, they are isomorphic. -/ +@[simps] +def IsInitial.uniqueUpToIso {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') : I ≅ I' where + hom := hI.to _ + inv := hI'.to _ + +variable (C) + +section Univ + +variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty ⥤ C} + +/-- Being terminal is independent of the empty diagram, its universe, and the cone over it, + as long as the cone points are isomorphic. -/ +def isLimitChangeEmptyCone {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt ≅ c₂.pt) : + IsLimit c₂ where + lift c := hl.lift ⟨c.pt, by aesop_cat, by aesop_cat⟩ ≫ hi.hom + uniq c f _ := by + dsimp + rw [← hl.uniq _ (f ≫ hi.inv) _] + · simp only [Category.assoc, Iso.inv_hom_id, Category.comp_id] + · aesop_cat + +/-- Replacing an empty cone in `IsLimit` by another with the same cone point + is an equivalence. -/ +def isLimitEmptyConeEquiv (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt ≅ c₂.pt) : + IsLimit c₁ ≃ IsLimit c₂ where + toFun hl := isLimitChangeEmptyCone C hl c₂ h + invFun hl := isLimitChangeEmptyCone C hl c₁ h.symm + left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.LeftInverse,Function.RightInverse]; intro + simp only [eq_iff_true_of_subsingleton] + +/-- Being initial is independent of the empty diagram, its universe, and the cocone over it, + as long as the cocone points are isomorphic. -/ +def isColimitChangeEmptyCocone {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) + (hi : c₁.pt ≅ c₂.pt) : IsColimit c₂ where + desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by aesop_cat⟩ + uniq c f _ := by + dsimp + rw [← hl.uniq _ (hi.hom ≫ f) _] + · simp only [Iso.inv_hom_id_assoc] + · aesop_cat + +/-- Replacing an empty cocone in `IsColimit` by another with the same cocone point + is an equivalence. -/ +def isColimitEmptyCoconeEquiv (c₁ : Cocone F₁) (c₂ : Cocone F₂) (h : c₁.pt ≅ c₂.pt) : + IsColimit c₁ ≃ IsColimit c₂ where + toFun hl := isColimitChangeEmptyCocone C hl c₂ h + invFun hl := isColimitChangeEmptyCocone C hl c₁ h.symm + left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] + right_inv := by + dsimp [Function.LeftInverse,Function.RightInverse]; intro + simp only [eq_iff_true_of_subsingleton] + +end Univ + +section + +variable {C} + +/-- An initial object is terminal in the opposite category. -/ +def terminalOpOfInitial {X : C} (t : IsInitial X) : IsTerminal (Opposite.op X) where + lift s := (t.to s.pt.unop).op + uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) + +/-- An initial object in the opposite category is terminal in the original category. -/ +def terminalUnopOfInitial {X : Cᵒᵖ} (t : IsInitial X) : IsTerminal X.unop where + lift s := (t.to (Opposite.op s.pt)).unop + uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) + +/-- A terminal object is initial in the opposite category. -/ +def initialOpOfTerminal {X : C} (t : IsTerminal X) : IsInitial (Opposite.op X) where + desc s := (t.from s.pt.unop).op + uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) + +/-- A terminal object in the opposite category is initial in the original category. -/ +def initialUnopOfTerminal {X : Cᵒᵖ} (t : IsTerminal X) : IsInitial X.unop where + desc s := (t.from (Opposite.op s.pt)).unop + uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) + +/-- A category is an `InitialMonoClass` if the canonical morphism of an initial object is a +monomorphism. In practice, this is most useful when given an arbitrary morphism out of the chosen +initial object, see `initial.mono_from`. +Given a terminal object, this is equivalent to the assumption that the unique morphism from initial +to terminal is a monomorphism, which is the second of Freyd's axioms for an AT category. + +TODO: This is a condition satisfied by categories with zero objects and morphisms. +-/ +class InitialMonoClass (C : Type u₁) [Category.{v₁} C] : Prop where + /-- The map from the (any as stated) initial object to any other object is a + monomorphism -/ + isInitial_mono_from : ∀ {I} (X : C) (hI : IsInitial I), Mono (hI.to X) + +theorem IsInitial.mono_from [InitialMonoClass C] {I} {X : C} (hI : IsInitial I) (f : I ⟶ X) : + Mono f := by + rw [hI.hom_ext f (hI.to X)] + apply InitialMonoClass.isInitial_mono_from + +/-- To show a category is an `InitialMonoClass` it suffices to give an initial object such that +every morphism out of it is a monomorphism. -/ +theorem InitialMonoClass.of_isInitial {I : C} (hI : IsInitial I) (h : ∀ X, Mono (hI.to X)) : + InitialMonoClass C where + isInitial_mono_from {I'} X hI' := by + rw [hI'.hom_ext (hI'.to X) ((hI'.uniqueUpToIso hI).hom ≫ hI.to X)] + apply mono_comp + +/-- To show a category is an `InitialMonoClass` it suffices to show the unique morphism from an +initial object to a terminal object is a monomorphism. -/ +theorem InitialMonoClass.of_isTerminal {I T : C} (hI : IsInitial I) (hT : IsTerminal T) + (_ : Mono (hI.to T)) : InitialMonoClass C := + InitialMonoClass.of_isInitial hI fun X => mono_of_mono_fac (hI.hom_ext (_ ≫ hT.from X) (hI.to T)) + +section Comparison + +variable {D : Type u₂} [Category.{v₂} D] (G : C ⥤ D) + +end Comparison + +variable {J : Type u} [Category.{v} J] + +/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cone for `J`. +In `limitOfDiagramInitial` we show it is a limit cone. -/ +@[simps] +def coneOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : Cone F where + pt := F.obj X + π := + { app := fun j => F.map (tX.to j) + naturality := fun j j' k => by + dsimp + rw [← F.map_comp, Category.id_comp, tX.hom_ext (tX.to j ≫ k) (tX.to j')] } + +/-- From a functor `F : J ⥤ C`, given an initial object of `J`, show the cone +`coneOfDiagramInitial` is a limit. -/ +def limitOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : + IsLimit (coneOfDiagramInitial tX F) where + lift s := s.π.app X + uniq s m w := by + conv_lhs => dsimp + simp_rw [← w X, coneOfDiagramInitial_π_app, tX.hom_ext (tX.to X) (𝟙 _)] + simp + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cone for `J`, +provided that the morphisms in the diagram are isomorphisms. +In `limitOfDiagramTerminal` we show it is a limit cone. -/ +@[simps] +def coneOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cone F where + pt := F.obj X + π := + { app := fun _ => inv (F.map (hX.from _)) + naturality := by + intro i j f + dsimp + simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.id_comp, ← F.map_comp, + hX.hom_ext (hX.from i) (f ≫ hX.from j)] } + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J` and that the morphisms in the +diagram are isomorphisms, show the cone `coneOfDiagramTerminal` is a limit. -/ +def limitOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsLimit (coneOfDiagramTerminal hX F) where + lift S := S.π.app _ + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cocone for `J`. +In `colimitOfDiagramTerminal` we show it is a colimit cocone. -/ +@[simps] +def coconeOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : Cocone F where + pt := F.obj X + ι := + { app := fun j => F.map (tX.from j) + naturality := fun j j' k => by + dsimp + rw [← F.map_comp, Category.comp_id, tX.hom_ext (k ≫ tX.from j') (tX.from j)] } + +/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, show the cocone +`coconeOfDiagramTerminal` is a colimit. -/ +def colimitOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : + IsColimit (coconeOfDiagramTerminal tX F) where + desc s := s.ι.app X + uniq s m w := by + conv_rhs => dsimp -- Porting note: why do I need this much firepower? + rw [← w X, coconeOfDiagramTerminal_ι_app, tX.hom_ext (tX.from X) (𝟙 _)] + simp + +lemma IsColimit.isIso_ι_app_of_isTerminal {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) + (X : J) (hX : IsTerminal X) : + IsIso (c.ι.app X) := by + change IsIso (coconePointUniqueUpToIso (colimitOfDiagramTerminal hX F) hc).hom + infer_instance + +/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cocone for `J`, +provided that the morphisms in the diagram are isomorphisms. +In `colimitOfDiagramInitial` we show it is a colimit cocone. -/ +@[simps] +def coconeOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cocone F where + pt := F.obj X + ι := + { app := fun _ => inv (F.map (hX.to _)) + naturality := by + intro i j f + dsimp + simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.comp_id, ← F.map_comp, + hX.hom_ext (hX.to i ≫ f) (hX.to j)] } + +/-- From a functor `F : J ⥤ C`, given an initial object of `J` and that the morphisms in the +diagram are isomorphisms, show the cone `coconeOfDiagramInitial` is a colimit. -/ +def colimitOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) + [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsColimit (coconeOfDiagramInitial hX F) where + desc S := S.ι.app _ + +lemma IsLimit.isIso_π_app_of_isInitial {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) + (X : J) (hX : IsInitial X) : + IsIso (c.π.app X) := by + change IsIso (conePointUniqueUpToIso hc (limitOfDiagramInitial hX F)).hom + infer_instance + +/-- Any morphism between terminal objects is an isomorphism. -/ +lemma isIso_of_isTerminal {X Y : C} (hX : IsTerminal X) (hY : IsTerminal Y) (f : X ⟶ Y) : + IsIso f := by + refine ⟨⟨IsTerminal.from hX Y, ?_⟩⟩ + simp only [IsTerminal.comp_from, IsTerminal.from_self, true_and] + apply IsTerminal.hom_ext hY + +/-- Any morphism between initial objects is an isomorphism. -/ +lemma isIso_of_isInitial {X Y : C} (hX : IsInitial X) (hY : IsInitial Y) (f : X ⟶ Y) : + IsIso f := by + refine ⟨⟨IsInitial.to hY X, ?_⟩⟩ + simp only [IsInitial.to_comp, IsInitial.to_self, and_true] + apply IsInitial.hom_ext hX + +end + +end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean index 03741e73c82af..d24c05e73cd2c 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Terminal.lean @@ -3,10 +3,8 @@ Copyright (c) 2019 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Bhavik Mehta -/ -import Mathlib.CategoryTheory.PEmpty +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal import Mathlib.CategoryTheory.Limits.HasLimits -import Mathlib.CategoryTheory.EpiMono -import Mathlib.CategoryTheory.Category.Preorder /-! # Initial and terminal objects in a category. @@ -26,180 +24,6 @@ namespace CategoryTheory.Limits variable {C : Type u₁} [Category.{v₁} C] -/-- Construct a cone for the empty diagram given an object. -/ -@[simps] -def asEmptyCone (X : C) : Cone (Functor.empty.{0} C) := - { pt := X - π := - { app := by aesop_cat } } - -/-- Construct a cocone for the empty diagram given an object. -/ -@[simps] -def asEmptyCocone (X : C) : Cocone (Functor.empty.{0} C) := - { pt := X - ι := - { app := by aesop_cat } } - -/-- `X` is terminal if the cone it induces on the empty diagram is limiting. -/ -abbrev IsTerminal (X : C) := - IsLimit (asEmptyCone X) - -/-- `X` is initial if the cocone it induces on the empty diagram is colimiting. -/ -abbrev IsInitial (X : C) := - IsColimit (asEmptyCocone X) - -/-- An object `Y` is terminal iff for every `X` there is a unique morphism `X ⟶ Y`. -/ -def isTerminalEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (Y : C) : - IsLimit (⟨Y, by aesop_cat, by aesop_cat⟩ : Cone F) ≃ ∀ X : C, Unique (X ⟶ Y) where - toFun t X := - { default := t.lift ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ - uniq := fun f => - t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } - invFun u := - { lift := fun s => (u s.pt).default - uniq := fun s _ _ => (u s.pt).2 _ } - left_inv := by dsimp [Function.LeftInverse]; intro x; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.RightInverse,Function.LeftInverse] - intro u; funext X; simp only - -/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` - (as an instance). -/ -def IsTerminal.ofUnique (Y : C) [h : ∀ X : C, Unique (X ⟶ Y)] : IsTerminal Y where - lift s := (h s.pt).default - fac := fun _ ⟨j⟩ => j.elim - -/-- An object `Y` is terminal if for every `X` there is a unique morphism `X ⟶ Y` - (as explicit arguments). -/ -def IsTerminal.ofUniqueHom {Y : C} (h : ∀ X : C, X ⟶ Y) (uniq : ∀ (X : C) (m : X ⟶ Y), m = h X) : - IsTerminal Y := - have : ∀ X : C, Unique (X ⟶ Y) := fun X ↦ ⟨⟨h X⟩, uniq X⟩ - IsTerminal.ofUnique Y - -/-- If `α` is a preorder with top, then `⊤` is a terminal object. -/ -def isTerminalTop {α : Type*} [Preorder α] [OrderTop α] : IsTerminal (⊤ : α) := - IsTerminal.ofUnique _ - -/-- Transport a term of type `IsTerminal` across an isomorphism. -/ -def IsTerminal.ofIso {Y Z : C} (hY : IsTerminal Y) (i : Y ≅ Z) : IsTerminal Z := - IsLimit.ofIsoLimit hY - { hom := { hom := i.hom } - inv := { hom := i.inv } } - -/-- If `X` and `Y` are isomorphic, then `X` is terminal iff `Y` is. -/ -def IsTerminal.equivOfIso {X Y : C} (e : X ≅ Y) : - IsTerminal X ≃ IsTerminal Y where - toFun h := IsTerminal.ofIso h e - invFun h := IsTerminal.ofIso h e.symm - left_inv _ := Subsingleton.elim _ _ - right_inv _ := Subsingleton.elim _ _ - -/-- An object `X` is initial iff for every `Y` there is a unique morphism `X ⟶ Y`. -/ -def isInitialEquivUnique (F : Discrete.{0} PEmpty.{1} ⥤ C) (X : C) : - IsColimit (⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ : Cocone F) ≃ ∀ Y : C, Unique (X ⟶ Y) where - toFun t X := - { default := t.desc ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ - uniq := fun f => t.uniq ⟨X, ⟨by aesop_cat, by aesop_cat⟩⟩ f (by aesop_cat) } - invFun u := - { desc := fun s => (u s.pt).default - uniq := fun s _ _ => (u s.pt).2 _ } - left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.RightInverse,Function.LeftInverse] - intro; funext; simp only - -/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` - (as an instance). -/ -def IsInitial.ofUnique (X : C) [h : ∀ Y : C, Unique (X ⟶ Y)] : IsInitial X where - desc s := (h s.pt).default - fac := fun _ ⟨j⟩ => j.elim - -/-- An object `X` is initial if for every `Y` there is a unique morphism `X ⟶ Y` - (as explicit arguments). -/ -def IsInitial.ofUniqueHom {X : C} (h : ∀ Y : C, X ⟶ Y) (uniq : ∀ (Y : C) (m : X ⟶ Y), m = h Y) : - IsInitial X := - have : ∀ Y : C, Unique (X ⟶ Y) := fun Y ↦ ⟨⟨h Y⟩, uniq Y⟩ - IsInitial.ofUnique X - -/-- If `α` is a preorder with bot, then `⊥` is an initial object. -/ -def isInitialBot {α : Type*} [Preorder α] [OrderBot α] : IsInitial (⊥ : α) := - IsInitial.ofUnique _ - -/-- Transport a term of type `is_initial` across an isomorphism. -/ -def IsInitial.ofIso {X Y : C} (hX : IsInitial X) (i : X ≅ Y) : IsInitial Y := - IsColimit.ofIsoColimit hX - { hom := { hom := i.hom } - inv := { hom := i.inv } } - -/-- If `X` and `Y` are isomorphic, then `X` is initial iff `Y` is. -/ -def IsInitial.equivOfIso {X Y : C} (e : X ≅ Y) : - IsInitial X ≃ IsInitial Y where - toFun h := IsInitial.ofIso h e - invFun h := IsInitial.ofIso h e.symm - left_inv _ := Subsingleton.elim _ _ - right_inv _ := Subsingleton.elim _ _ - -/-- Give the morphism to a terminal object from any other. -/ -def IsTerminal.from {X : C} (t : IsTerminal X) (Y : C) : Y ⟶ X := - t.lift (asEmptyCone Y) - -/-- Any two morphisms to a terminal object are equal. -/ -theorem IsTerminal.hom_ext {X Y : C} (t : IsTerminal X) (f g : Y ⟶ X) : f = g := - IsLimit.hom_ext t (by aesop_cat) - -@[simp] -theorem IsTerminal.comp_from {Z : C} (t : IsTerminal Z) {X Y : C} (f : X ⟶ Y) : - f ≫ t.from Y = t.from X := - t.hom_ext _ _ - -@[simp] -theorem IsTerminal.from_self {X : C} (t : IsTerminal X) : t.from X = 𝟙 X := - t.hom_ext _ _ - -/-- Give the morphism from an initial object to any other. -/ -def IsInitial.to {X : C} (t : IsInitial X) (Y : C) : X ⟶ Y := - t.desc (asEmptyCocone Y) - -/-- Any two morphisms from an initial object are equal. -/ -theorem IsInitial.hom_ext {X Y : C} (t : IsInitial X) (f g : X ⟶ Y) : f = g := - IsColimit.hom_ext t (by aesop_cat) - -@[simp] -theorem IsInitial.to_comp {X : C} (t : IsInitial X) {Y Z : C} (f : Y ⟶ Z) : t.to Y ≫ f = t.to Z := - t.hom_ext _ _ - -@[simp] -theorem IsInitial.to_self {X : C} (t : IsInitial X) : t.to X = 𝟙 X := - t.hom_ext _ _ - -/-- Any morphism from a terminal object is split mono. -/ -theorem IsTerminal.isSplitMono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : IsSplitMono f := - IsSplitMono.mk' ⟨t.from _, t.hom_ext _ _⟩ - -/-- Any morphism to an initial object is split epi. -/ -theorem IsInitial.isSplitEpi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : IsSplitEpi f := - IsSplitEpi.mk' ⟨t.to _, t.hom_ext _ _⟩ - -/-- Any morphism from a terminal object is mono. -/ -theorem IsTerminal.mono_from {X Y : C} (t : IsTerminal X) (f : X ⟶ Y) : Mono f := by - haveI := t.isSplitMono_from f; infer_instance - -/-- Any morphism to an initial object is epi. -/ -theorem IsInitial.epi_to {X Y : C} (t : IsInitial X) (f : Y ⟶ X) : Epi f := by - haveI := t.isSplitEpi_to f; infer_instance - -/-- If `T` and `T'` are terminal, they are isomorphic. -/ -@[simps] -def IsTerminal.uniqueUpToIso {T T' : C} (hT : IsTerminal T) (hT' : IsTerminal T') : T ≅ T' where - hom := hT'.from _ - inv := hT.from _ - -/-- If `I` and `I'` are initial, they are isomorphic. -/ -@[simps] -def IsInitial.uniqueUpToIso {I I' : C} (hI : IsInitial I) (hI' : IsInitial I') : I ≅ I' where - hom := hI.to _ - inv := hI'.to _ - variable (C) /-- A category has a terminal object if it has a limit over the empty diagram. @@ -218,28 +42,6 @@ section Univ variable (X : C) {F₁ : Discrete.{w} PEmpty ⥤ C} {F₂ : Discrete.{w'} PEmpty ⥤ C} -/-- Being terminal is independent of the empty diagram, its universe, and the cone over it, - as long as the cone points are isomorphic. -/ -def isLimitChangeEmptyCone {c₁ : Cone F₁} (hl : IsLimit c₁) (c₂ : Cone F₂) (hi : c₁.pt ≅ c₂.pt) : - IsLimit c₂ where - lift c := hl.lift ⟨c.pt, by aesop_cat, by aesop_cat⟩ ≫ hi.hom - uniq c f _ := by - dsimp - rw [← hl.uniq _ (f ≫ hi.inv) _] - · simp only [Category.assoc, Iso.inv_hom_id, Category.comp_id] - · aesop_cat - -/-- Replacing an empty cone in `IsLimit` by another with the same cone point - is an equivalence. -/ -def isLimitEmptyConeEquiv (c₁ : Cone F₁) (c₂ : Cone F₂) (h : c₁.pt ≅ c₂.pt) : - IsLimit c₁ ≃ IsLimit c₂ where - toFun hl := isLimitChangeEmptyCone C hl c₂ h - invFun hl := isLimitChangeEmptyCone C hl c₁ h.symm - left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.LeftInverse,Function.RightInverse]; intro - simp only [eq_iff_true_of_subsingleton] - theorem hasTerminalChangeDiagram (h : HasLimit F₁) : HasLimit F₂ := ⟨⟨⟨⟨limit F₁, by aesop_cat, by aesop_cat⟩, isLimitChangeEmptyCone C (limit.isLimit F₁) _ (eqToIso rfl)⟩⟩⟩ @@ -248,28 +50,6 @@ theorem hasTerminalChangeUniverse [h : HasLimitsOfShape (Discrete.{w} PEmpty) C] HasLimitsOfShape (Discrete.{w'} PEmpty) C where has_limit _ := hasTerminalChangeDiagram C (h.1 (Functor.empty C)) -/-- Being initial is independent of the empty diagram, its universe, and the cocone over it, - as long as the cocone points are isomorphic. -/ -def isColimitChangeEmptyCocone {c₁ : Cocone F₁} (hl : IsColimit c₁) (c₂ : Cocone F₂) - (hi : c₁.pt ≅ c₂.pt) : IsColimit c₂ where - desc c := hi.inv ≫ hl.desc ⟨c.pt, by aesop_cat, by aesop_cat⟩ - uniq c f _ := by - dsimp - rw [← hl.uniq _ (hi.hom ≫ f) _] - · simp only [Iso.inv_hom_id_assoc] - · aesop_cat - -/-- Replacing an empty cocone in `IsColimit` by another with the same cocone point - is an equivalence. -/ -def isColimitEmptyCoconeEquiv (c₁ : Cocone F₁) (c₂ : Cocone F₂) (h : c₁.pt ≅ c₂.pt) : - IsColimit c₁ ≃ IsColimit c₂ where - toFun hl := isColimitChangeEmptyCocone C hl c₂ h - invFun hl := isColimitChangeEmptyCocone C hl c₁ h.symm - left_inv := by dsimp [Function.LeftInverse]; intro; simp only [eq_iff_true_of_subsingleton] - right_inv := by - dsimp [Function.LeftInverse,Function.RightInverse]; intro - simp only [eq_iff_true_of_subsingleton] - theorem hasInitialChangeDiagram (h : HasColimit F₁) : HasColimit F₂ := ⟨⟨⟨⟨colimit F₁, by aesop_cat, by aesop_cat⟩, isColimitChangeEmptyCocone C (colimit.isColimit F₁) _ (eqToIso rfl)⟩⟩⟩ @@ -379,26 +159,6 @@ instance terminal.isSplitMono_from {Y : C} [HasTerminal C] (f : ⊤_ C ⟶ Y) : instance initial.isSplitEpi_to {Y : C} [HasInitial C] (f : Y ⟶ ⊥_ C) : IsSplitEpi f := IsInitial.isSplitEpi_to initialIsInitial _ -/-- An initial object is terminal in the opposite category. -/ -def terminalOpOfInitial {X : C} (t : IsInitial X) : IsTerminal (Opposite.op X) where - lift s := (t.to s.pt.unop).op - uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) - -/-- An initial object in the opposite category is terminal in the original category. -/ -def terminalUnopOfInitial {X : Cᵒᵖ} (t : IsInitial X) : IsTerminal X.unop where - lift s := (t.to (Opposite.op s.pt)).unop - uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) - -/-- A terminal object is initial in the opposite category. -/ -def initialOpOfTerminal {X : C} (t : IsTerminal X) : IsInitial (Opposite.op X) where - desc s := (t.from s.pt.unop).op - uniq _ _ _ := Quiver.Hom.unop_inj (t.hom_ext _ _) - -/-- A terminal object in the opposite category is initial in the original category. -/ -def initialUnopOfTerminal {X : Cᵒᵖ} (t : IsTerminal X) : IsInitial X.unop where - desc s := (t.from (Opposite.op s.pt)).unop - uniq _ _ _ := Quiver.Hom.op_inj (t.hom_ext _ _) - instance hasInitial_op_of_hasTerminal [HasTerminal C] : HasInitial Cᵒᵖ := (initialOpOfTerminal terminalIsTerminal).hasInitial @@ -459,48 +219,16 @@ theorem ι_colimitConstInitial_hom {J : Type*} [Category J] {C : Type*} [Categor colimit.ι ((CategoryTheory.Functor.const J).obj (⊥_ C)) j ≫ colimitConstInitial.hom = initial.to _ := by aesop_cat -/-- A category is an `InitialMonoClass` if the canonical morphism of an initial object is a -monomorphism. In practice, this is most useful when given an arbitrary morphism out of the chosen -initial object, see `initial.mono_from`. -Given a terminal object, this is equivalent to the assumption that the unique morphism from initial -to terminal is a monomorphism, which is the second of Freyd's axioms for an AT category. - -TODO: This is a condition satisfied by categories with zero objects and morphisms. --/ -class InitialMonoClass (C : Type u₁) [Category.{v₁} C] : Prop where - /-- The map from the (any as stated) initial object to any other object is a - monomorphism -/ - isInitial_mono_from : ∀ {I} (X : C) (hI : IsInitial I), Mono (hI.to X) - -theorem IsInitial.mono_from [InitialMonoClass C] {I} {X : C} (hI : IsInitial I) (f : I ⟶ X) : - Mono f := by - rw [hI.hom_ext f (hI.to X)] - apply InitialMonoClass.isInitial_mono_from - instance (priority := 100) initial.mono_from [HasInitial C] [InitialMonoClass C] (X : C) (f : ⊥_ C ⟶ X) : Mono f := initialIsInitial.mono_from f -/-- To show a category is an `InitialMonoClass` it suffices to give an initial object such that -every morphism out of it is a monomorphism. -/ -theorem InitialMonoClass.of_isInitial {I : C} (hI : IsInitial I) (h : ∀ X, Mono (hI.to X)) : - InitialMonoClass C where - isInitial_mono_from {I'} X hI' := by - rw [hI'.hom_ext (hI'.to X) ((hI'.uniqueUpToIso hI).hom ≫ hI.to X)] - apply mono_comp - /-- To show a category is an `InitialMonoClass` it suffices to show every morphism out of the initial object is a monomorphism. -/ theorem InitialMonoClass.of_initial [HasInitial C] (h : ∀ X : C, Mono (initial.to X)) : InitialMonoClass C := InitialMonoClass.of_isInitial initialIsInitial h -/-- To show a category is an `InitialMonoClass` it suffices to show the unique morphism from an -initial object to a terminal object is a monomorphism. -/ -theorem InitialMonoClass.of_isTerminal {I T : C} (hI : IsInitial I) (hT : IsTerminal T) - (_ : Mono (hI.to T)) : InitialMonoClass C := - InitialMonoClass.of_isInitial hI fun X => mono_of_mono_fac (hI.hom_ext (_ ≫ hT.from X) (hI.to T)) - /-- To show a category is an `InitialMonoClass` it suffices to show the unique morphism from the initial object to a terminal object is a monomorphism. -/ theorem InitialMonoClass.of_terminal [HasInitial C] [HasTerminal C] (h : Mono (initial.to (⊤_ C))) : @@ -531,27 +259,6 @@ end Comparison variable {J : Type u} [Category.{v} J] -/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cone for `J`. -In `limitOfDiagramInitial` we show it is a limit cone. -/ -@[simps] -def coneOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : Cone F where - pt := F.obj X - π := - { app := fun j => F.map (tX.to j) - naturality := fun j j' k => by - dsimp - rw [← F.map_comp, Category.id_comp, tX.hom_ext (tX.to j ≫ k) (tX.to j')] } - -/-- From a functor `F : J ⥤ C`, given an initial object of `J`, show the cone -`coneOfDiagramInitial` is a limit. -/ -def limitOfDiagramInitial {X : J} (tX : IsInitial X) (F : J ⥤ C) : - IsLimit (coneOfDiagramInitial tX F) where - lift s := s.π.app X - uniq s m w := by - conv_lhs => dsimp - simp_rw [← w X, coneOfDiagramInitial_π_app, tX.hom_ext (tX.to X) (𝟙 _)] - simp - instance hasLimit_of_domain_hasInitial [HasInitial J] {F : J ⥤ C} : HasLimit F := HasLimit.mk { cone := _, isLimit := limitOfDiagramInitial (initialIsInitial) F } @@ -562,27 +269,6 @@ to the limit of `F`. -/ abbrev limitOfInitial (F : J ⥤ C) [HasInitial J] : limit F ≅ F.obj (⊥_ J) := IsLimit.conePointUniqueUpToIso (limit.isLimit _) (limitOfDiagramInitial initialIsInitial F) -/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cone for `J`, -provided that the morphisms in the diagram are isomorphisms. -In `limitOfDiagramTerminal` we show it is a limit cone. -/ -@[simps] -def coneOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cone F where - pt := F.obj X - π := - { app := fun _ => inv (F.map (hX.from _)) - naturality := by - intro i j f - dsimp - simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.id_comp, ← F.map_comp, - hX.hom_ext (hX.from i) (f ≫ hX.from j)] } - -/-- From a functor `F : J ⥤ C`, given a terminal object of `J` and that the morphisms in the -diagram are isomorphisms, show the cone `coneOfDiagramTerminal` is a limit. -/ -def limitOfDiagramTerminal {X : J} (hX : IsTerminal X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsLimit (coneOfDiagramTerminal hX F) where - lift S := S.π.app _ - instance hasLimit_of_domain_hasTerminal [HasTerminal J] {F : J ⥤ C} [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : HasLimit F := HasLimit.mk { cone := _, isLimit := limitOfDiagramTerminal (terminalIsTerminal) F } @@ -594,36 +280,9 @@ abbrev limitOfTerminal (F : J ⥤ C) [HasTerminal J] [∀ (i j : J) (f : i ⟶ j limit F ≅ F.obj (⊤_ J) := IsLimit.conePointUniqueUpToIso (limit.isLimit _) (limitOfDiagramTerminal terminalIsTerminal F) -/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, construct a cocone for `J`. -In `colimitOfDiagramTerminal` we show it is a colimit cocone. -/ -@[simps] -def coconeOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : Cocone F where - pt := F.obj X - ι := - { app := fun j => F.map (tX.from j) - naturality := fun j j' k => by - dsimp - rw [← F.map_comp, Category.comp_id, tX.hom_ext (k ≫ tX.from j') (tX.from j)] } - -/-- From a functor `F : J ⥤ C`, given a terminal object of `J`, show the cocone -`coconeOfDiagramTerminal` is a colimit. -/ -def colimitOfDiagramTerminal {X : J} (tX : IsTerminal X) (F : J ⥤ C) : - IsColimit (coconeOfDiagramTerminal tX F) where - desc s := s.ι.app X - uniq s m w := by - conv_rhs => dsimp -- Porting note: why do I need this much firepower? - rw [← w X, coconeOfDiagramTerminal_ι_app, tX.hom_ext (tX.from X) (𝟙 _)] - simp - instance hasColimit_of_domain_hasTerminal [HasTerminal J] {F : J ⥤ C} : HasColimit F := HasColimit.mk { cocone := _, isColimit := colimitOfDiagramTerminal (terminalIsTerminal) F } -lemma IsColimit.isIso_ι_app_of_isTerminal {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) - (X : J) (hX : IsTerminal X) : - IsIso (c.ι.app X) := by - change IsIso (coconePointUniqueUpToIso (colimitOfDiagramTerminal hX F) hc).hom - infer_instance - -- This is reducible to allow usage of lemmas about `cocone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has a terminal object then the image of it is isomorphic to the colimit of `F`. -/ @@ -631,37 +290,10 @@ abbrev colimitOfTerminal (F : J ⥤ C) [HasTerminal J] : colimit F ≅ F.obj ( IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) (colimitOfDiagramTerminal terminalIsTerminal F) -/-- From a functor `F : J ⥤ C`, given an initial object of `J`, construct a cocone for `J`, -provided that the morphisms in the diagram are isomorphisms. -In `colimitOfDiagramInitial` we show it is a colimit cocone. -/ -@[simps] -def coconeOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : Cocone F where - pt := F.obj X - ι := - { app := fun _ => inv (F.map (hX.to _)) - naturality := by - intro i j f - dsimp - simp only [IsIso.eq_inv_comp, IsIso.comp_inv_eq, Category.comp_id, ← F.map_comp, - hX.hom_ext (hX.to i ≫ f) (hX.to j)] } - -/-- From a functor `F : J ⥤ C`, given an initial object of `J` and that the morphisms in the -diagram are isomorphisms, show the cone `coconeOfDiagramInitial` is a colimit. -/ -def colimitOfDiagramInitial {X : J} (hX : IsInitial X) (F : J ⥤ C) - [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : IsColimit (coconeOfDiagramInitial hX F) where - desc S := S.ι.app _ - instance hasColimit_of_domain_hasInitial [HasInitial J] {F : J ⥤ C} [∀ (i j : J) (f : i ⟶ j), IsIso (F.map f)] : HasColimit F := HasColimit.mk { cocone := _, isColimit := colimitOfDiagramInitial (initialIsInitial) F } -lemma IsLimit.isIso_π_app_of_isInitial {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) - (X : J) (hX : IsInitial X) : - IsIso (c.π.app X) := by - change IsIso (conePointUniqueUpToIso hc (limitOfDiagramInitial hX F)).hom - infer_instance - -- This is reducible to allow usage of lemmas about `cocone_point_unique_up_to_iso`. /-- For a functor `F : J ⥤ C`, if `J` has an initial object and all the morphisms in the diagram are isomorphisms, then the image of the initial object is isomorphic to the colimit of `F`. -/ @@ -710,20 +342,6 @@ instance isIso_ι_initial [HasInitial J] (F : J ⥤ C) [∀ (i j : J) (f : i ⟶ IsIso (colimit.ι F (⊥_ J)) := isIso_ι_of_isInitial initialIsInitial F -/-- Any morphism between terminal objects is an isomorphism. -/ -lemma isIso_of_isTerminal {X Y : C} (hX : IsTerminal X) (hY : IsTerminal Y) (f : X ⟶ Y) : - IsIso f := by - refine ⟨⟨IsTerminal.from hX Y, ?_⟩⟩ - simp only [IsTerminal.comp_from, IsTerminal.from_self, true_and] - apply IsTerminal.hom_ext hY - -/-- Any morphism between initial objects is an isomorphism. -/ -lemma isIso_of_isInitial {X Y : C} (hX : IsInitial X) (hY : IsInitial Y) (f : X ⟶ Y) : - IsIso f := by - refine ⟨⟨IsInitial.to hY X, ?_⟩⟩ - simp only [IsInitial.to_comp, IsInitial.to_self, and_true] - apply IsInitial.hom_ext hX - end end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Monad/Limits.lean b/Mathlib/CategoryTheory/Monad/Limits.lean index 1e1128d105453..11d531a830a9b 100644 --- a/Mathlib/CategoryTheory/Monad/Limits.lean +++ b/Mathlib/CategoryTheory/Monad/Limits.lean @@ -5,7 +5,7 @@ Authors: Kim Morrison, Bhavik Mehta, Jack McKoen -/ import Mathlib.CategoryTheory.Monad.Adjunction import Mathlib.CategoryTheory.Adjunction.Limits -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal /-! # Limits and colimits in the category of (co)algebras diff --git a/Mathlib/CategoryTheory/WithTerminal.lean b/Mathlib/CategoryTheory/WithTerminal.lean index 3ac83acd3454e..0b00c0e2f5d8f 100644 --- a/Mathlib/CategoryTheory/WithTerminal.lean +++ b/Mathlib/CategoryTheory/WithTerminal.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith, Adam Topaz -/ -import Mathlib.CategoryTheory.Limits.Shapes.Terminal +import Mathlib.CategoryTheory.Limits.Shapes.IsTerminal import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor /-! From 533ce927a420bd01fa3c66d250c9beb523fcd2a3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 25 Oct 2024 15:10:12 +0000 Subject: [PATCH 22/29] feat(CategoryTheory): constructor for pseudofunctors from a strict locally discrete category (#18201) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this file, we define a constructor `pseudofunctorOfIsLocallyDiscrete` for the type `Pseudofunctor B C` when `C` is any bicategory, and `B` is a strict locally discrete category. Indeed, in this situation, we do not need to care about the field `map₂` of pseudofunctors because all the `2`-morphisms in `B` are identities. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Bicategory/Functor/LocallyDiscrete.lean | 157 ++++++++++++++++++ .../Bicategory/Grothendieck.lean | 1 + .../Bicategory/LocallyDiscrete.lean | 48 ++---- Mathlib/CategoryTheory/DiscreteCategory.lean | 20 +++ 5 files changed, 195 insertions(+), 32 deletions(-) create mode 100644 Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean diff --git a/Mathlib.lean b/Mathlib.lean index 05b32bb4e3391..b0ddfe0730762 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1490,6 +1490,7 @@ import Mathlib.CategoryTheory.Bicategory.End import Mathlib.CategoryTheory.Bicategory.Extension import Mathlib.CategoryTheory.Bicategory.Free import Mathlib.CategoryTheory.Bicategory.Functor.Lax +import Mathlib.CategoryTheory.Bicategory.Functor.LocallyDiscrete import Mathlib.CategoryTheory.Bicategory.Functor.Oplax import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor diff --git a/Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean b/Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean new file mode 100644 index 0000000000000..592c7e48b5551 --- /dev/null +++ b/Mathlib/CategoryTheory/Bicategory/Functor/LocallyDiscrete.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor +import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete + +/-! +# Pseudofunctors from locally discrete bicategories + +This file provides various ways of constructing pseudofunctors from locally discrete +bicategories. + +Firstly, we define the constructors `pseudofunctorOfIsLocallyDiscrete` and +`oplaxFunctorOfIsLocallyDiscrete` for defining pseudofunctors and oplax functors +from a locally discrete bicategories. In this situation, we do not need to care about +the field `map₂`, because all the `2`-morphisms in `B` are identities. + +We also define a specialized constructor `LocallyDiscrete.mkPseudofunctor` when +the source bicategory is of the form `B := LocallyDiscrete B₀` for a category `B₀`. + +We also prove that a functor `F : I ⥤ B` with `B` a strict bicategory can be promoted +to a pseudofunctor (or oplax functor) (`Functor.toPseudofunctor`) with domain +`LocallyDiscrete I`. + +-/ + +namespace CategoryTheory + +open Bicategory + +/-- Constructor for pseudofunctors from a locally discrete bicategory. In that +case, we do not need to provide the `map₂` field of pseudofunctors. -/ +@[simps obj map mapId mapComp] +def pseudofunctorOfIsLocallyDiscrete + {B C : Type*} [Bicategory B] [IsLocallyDiscrete B] [Bicategory C] + (obj : B → C) + (map : ∀ {b b' : B}, (b ⟶ b') → (obj b ⟶ obj b')) + (mapId : ∀ (b : B), map (𝟙 b) ≅ 𝟙 _) + (mapComp : ∀ {b₀ b₁ b₂ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂), map (f ≫ g) ≅ map f ≫ map g) + (map₂_associator : ∀ {b₀ b₁ b₂ b₃ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂) (h : b₂ ⟶ b₃), + (mapComp (f ≫ g) h).hom ≫ + (mapComp f g).hom ▷ map h ≫ (α_ (map f) (map g) (map h)).hom ≫ + map f ◁ (mapComp g h).inv ≫ (mapComp f (g ≫ h)).inv = eqToHom (by simp) := by aesop_cat) + (map₂_left_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + (mapComp (𝟙 b₀) f).hom ≫ (mapId b₀).hom ▷ map f ≫ (λ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) + (map₂_right_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + (mapComp f (𝟙 b₁)).hom ≫ map f ◁ (mapId b₁).hom ≫ (ρ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) : + Pseudofunctor B C where + obj := obj + map := map + map₂ φ := eqToHom (by + obtain rfl := obj_ext_of_isDiscrete φ + dsimp) + mapId := mapId + mapComp := mapComp + map₂_whisker_left _ _ _ η := by + obtain rfl := obj_ext_of_isDiscrete η + simp + map₂_whisker_right η _ := by + obtain rfl := obj_ext_of_isDiscrete η + simp + +/-- Constructor for oplax functors from a locally discrete bicategory. In that +case, we do not need to provide the `map₂` field of oplax functors. -/ +@[simps obj map mapId mapComp] +def oplaxFunctorOfIsLocallyDiscrete + {B C : Type*} [Bicategory B] [IsLocallyDiscrete B] [Bicategory C] + (obj : B → C) + (map : ∀ {b b' : B}, (b ⟶ b') → (obj b ⟶ obj b')) + (mapId : ∀ (b : B), map (𝟙 b) ⟶ 𝟙 _) + (mapComp : ∀ {b₀ b₁ b₂ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂), map (f ≫ g) ⟶ map f ≫ map g) + (map₂_associator : ∀ {b₀ b₁ b₂ b₃ : B} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂) (h : b₂ ⟶ b₃), + eqToHom (by simp) ≫ mapComp f (g ≫ h) ≫ map f ◁ mapComp g h = + mapComp (f ≫ g) h ≫ mapComp f g ▷ map h ≫ (α_ (map f) (map g) (map h)).hom := by + aesop_cat) + (map₂_left_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + mapComp (𝟙 b₀) f ≫ mapId b₀ ▷ map f ≫ (λ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) + (map₂_right_unitor : ∀ {b₀ b₁ : B} (f : b₀ ⟶ b₁), + mapComp f (𝟙 b₁) ≫ map f ◁ mapId b₁ ≫ (ρ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) : + OplaxFunctor B C where + obj := obj + map := map + map₂ φ := eqToHom (by + obtain rfl := obj_ext_of_isDiscrete φ + dsimp) + mapId := mapId + mapComp := mapComp + mapComp_naturality_left η := by + obtain rfl := obj_ext_of_isDiscrete η + simp + mapComp_naturality_right _ _ _ η := by + obtain rfl := obj_ext_of_isDiscrete η + simp + +section + +variable {I B : Type*} [Category I] [Bicategory B] [Strict B] (F : I ⥤ B) + +attribute [local simp] + Strict.leftUnitor_eqToIso Strict.rightUnitor_eqToIso Strict.associator_eqToIso + +/-- +If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can +be promoted to a pseudofunctor from `LocallyDiscrete I` to `B`. +-/ +@[simps! obj map mapId mapComp] +def Functor.toPseudoFunctor : Pseudofunctor (LocallyDiscrete I) B := + pseudofunctorOfIsLocallyDiscrete + (fun ⟨X⟩ ↦ F.obj X) (fun ⟨f⟩ ↦ F.map f) + (fun ⟨X⟩ ↦ eqToIso (by simp)) (fun f g ↦ eqToIso (by simp)) + +/-- +If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can +be promoted to an oplax functor from `LocallyDiscrete I` to `B`. +-/ +@[simps! obj map mapId mapComp] +def Functor.toOplaxFunctor : OplaxFunctor (LocallyDiscrete I) B := + oplaxFunctorOfIsLocallyDiscrete + (fun ⟨X⟩ ↦ F.obj X) (fun ⟨f⟩ ↦ F.map f) + (fun ⟨X⟩ ↦ eqToHom (by simp)) (fun f g ↦ eqToHom (by simp)) + +end + +namespace LocallyDiscrete + +/-- Constructor for pseudofunctors from a locally discrete bicategory. In that +case, we do not need to provide the `map₂` field of pseudofunctors. -/ +@[simps! obj map mapId mapComp] +def mkPseudofunctor {B₀ C : Type*} [Category B₀] [Bicategory C] + (obj : B₀ → C) + (map : ∀ {b b' : B₀}, (b ⟶ b') → (obj b ⟶ obj b')) + (mapId : ∀ (b : B₀), map (𝟙 b) ≅ 𝟙 _) + (mapComp : ∀ {b₀ b₁ b₂ : B₀} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂), map (f ≫ g) ≅ map f ≫ map g) + (map₂_associator : ∀ {b₀ b₁ b₂ b₃ : B₀} (f : b₀ ⟶ b₁) (g : b₁ ⟶ b₂) (h : b₂ ⟶ b₃), + (mapComp (f ≫ g) h).hom ≫ + (mapComp f g).hom ▷ map h ≫ (α_ (map f) (map g) (map h)).hom ≫ + map f ◁ (mapComp g h).inv ≫ (mapComp f (g ≫ h)).inv = eqToHom (by simp) := by aesop_cat) + (map₂_left_unitor : ∀ {b₀ b₁ : B₀} (f : b₀ ⟶ b₁), + (mapComp (𝟙 b₀) f).hom ≫ (mapId b₀).hom ▷ map f ≫ (λ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) + (map₂_right_unitor : ∀ {b₀ b₁ : B₀} (f : b₀ ⟶ b₁), + (mapComp f (𝟙 b₁)).hom ≫ map f ◁ (mapId b₁).hom ≫ (ρ_ (map f)).hom = eqToHom (by simp) := by + aesop_cat) : + Pseudofunctor (LocallyDiscrete B₀) C := + pseudofunctorOfIsLocallyDiscrete (fun b ↦ obj b.as) (fun f ↦ map f.as) + (fun _ ↦ mapId _) (fun _ _ ↦ mapComp _ _) (fun _ _ _ ↦ map₂_associator _ _ _) + (fun _ ↦ map₂_left_unitor _) (fun _ ↦ map₂_right_unitor _) + +end LocallyDiscrete + +end CategoryTheory diff --git a/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean index d7eb112f1f6ba..efb94ebc6c24a 100644 --- a/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Bicategory/Grothendieck.lean @@ -5,6 +5,7 @@ Authors: Calle Sönne -/ import Mathlib.CategoryTheory.Bicategory.LocallyDiscrete +import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor /-! # The Grothendieck construction diff --git a/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean b/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean index 2e93d740e778d..7db3d423d41ad 100644 --- a/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean +++ b/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno, Calle Sönne -/ import Mathlib.CategoryTheory.DiscreteCategory -import Mathlib.CategoryTheory.Bicategory.Functor.Pseudofunctor +import Mathlib.CategoryTheory.Bicategory.Functor.Prelax import Mathlib.CategoryTheory.Bicategory.Strict /-! @@ -21,8 +21,6 @@ namespace CategoryTheory open Bicategory Discrete -open Bicategory - universe w₂ w₁ v₂ v₁ v u₂ u₁ u section @@ -106,35 +104,6 @@ instance locallyDiscreteBicategory.strict : Strict (LocallyDiscrete C) where comp_id _ := Discrete.ext (Category.comp_id _) assoc _ _ _ := Discrete.ext (Category.assoc _ _ _) -attribute [local simp] - Strict.leftUnitor_eqToIso Strict.rightUnitor_eqToIso Strict.associator_eqToIso - -variable {I : Type u₁} [Category.{v₁} I] {B : Type u₂} [Bicategory.{w₂, v₂} B] [Strict B] - -/-- -If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can -be promoted to a pseudofunctor from `LocallyDiscrete I` to `B`. --/ -@[simps] -def Functor.toPseudoFunctor (F : I ⥤ B) : Pseudofunctor (LocallyDiscrete I) B where - obj i := F.obj i.as - map f := F.map f.as - map₂ η := eqToHom (congr_arg _ (LocallyDiscrete.eq_of_hom η)) - mapId i := eqToIso (F.map_id i.as) - mapComp f g := eqToIso (F.map_comp f.as g.as) - -/-- -If `B` is a strict bicategory and `I` is a (1-)category, any functor (of 1-categories) `I ⥤ B` can -be promoted to an oplax functor from `LocallyDiscrete I` to `B`. --/ -@[simps] -def Functor.toOplaxFunctor (F : I ⥤ B) : OplaxFunctor (LocallyDiscrete I) B where - obj i := F.obj i.as - map f := F.map f.as - map₂ η := eqToHom (congr_arg _ (LocallyDiscrete.eq_of_hom η)) - mapId i := eqToHom (F.map_id i.as) - mapComp f g := eqToHom (F.map_comp f.as g.as) - end section @@ -148,6 +117,21 @@ lemma PrelaxFunctor.map₂_eqToHom (F : PrelaxFunctor B C) {a b : B} {f g : a end +namespace Bicategory + +/-- A bicategory is locally discrete if the categories of 1-morphisms are discrete. -/ +abbrev IsLocallyDiscrete (B : Type*) [Bicategory B] := ∀ (b c : B), IsDiscrete (b ⟶ c) + +instance (C : Type*) [Category C] : IsLocallyDiscrete (LocallyDiscrete C) := + fun _ _ ↦ Discrete.isDiscrete _ + +instance (B : Type*) [Bicategory B] [IsLocallyDiscrete B] : Strict B where + id_comp f := obj_ext_of_isDiscrete (leftUnitor f).hom + comp_id f := obj_ext_of_isDiscrete (rightUnitor f).hom + assoc f g h := obj_ext_of_isDiscrete (associator f g h).hom + +end Bicategory + end CategoryTheory section diff --git a/Mathlib/CategoryTheory/DiscreteCategory.lean b/Mathlib/CategoryTheory/DiscreteCategory.lean index a6f0f73828116..1d28f0410aa9e 100644 --- a/Mathlib/CategoryTheory/DiscreteCategory.lean +++ b/Mathlib/CategoryTheory/DiscreteCategory.lean @@ -304,4 +304,24 @@ def piEquivalenceFunctorDiscrete (J : Type u₂) (C : Type u₁) [Category.{v₁ obtain rfl : f = 𝟙 _ := rfl simp))) (by aesop_cat) +/-- A category is discrete when there is at most one morphism between two objects, +in which case they are equal. -/ +class IsDiscrete (C : Type*) [Category C] : Prop where + subsingleton (X Y : C) : Subsingleton (X ⟶ Y) := by infer_instance + eq_of_hom {X Y : C} (f : X ⟶ Y) : X = Y + +attribute [instance] IsDiscrete.subsingleton + +lemma obj_ext_of_isDiscrete {C : Type*} [Category C] [IsDiscrete C] + {X Y : C} (f : X ⟶ Y) : X = Y := IsDiscrete.eq_of_hom f + +instance Discrete.isDiscrete (C : Type*) : IsDiscrete (Discrete C) where + eq_of_hom := by rintro ⟨_⟩ ⟨_⟩ ⟨⟨rfl⟩⟩; rfl + +instance (C : Type*) [Category C] [IsDiscrete C] : IsDiscrete Cᵒᵖ where + eq_of_hom := by + rintro ⟨_⟩ ⟨_⟩ ⟨f⟩ + obtain rfl := obj_ext_of_isDiscrete f + rfl + end CategoryTheory From 7f0670510ccf43e0e595357e31b7b8212cc02c45 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Fri, 25 Oct 2024 15:28:41 +0000 Subject: [PATCH 23/29] feat: add pow_sub_one_dvd_differentIdeal (#18228) From flt-regular. --- .../RingTheory/DedekindDomain/Different.lean | 69 +++++++++++++++++++ Mathlib/RingTheory/Trace/Basic.lean | 19 +++++ 2 files changed, 88 insertions(+) diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index 915f686f016f1..a9e9ad8924c5b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -7,6 +7,8 @@ import Mathlib.RingTheory.DedekindDomain.Ideal import Mathlib.RingTheory.Discriminant import Mathlib.RingTheory.DedekindDomain.IntegralClosure import Mathlib.NumberTheory.KummerDedekind +import Mathlib.RingTheory.IntegralClosure.IntegralRestrict +import Mathlib.RingTheory.Trace.Quotient /-! # The different ideal @@ -568,3 +570,70 @@ lemma aeval_derivative_mem_differentIdeal [NoZeroSMulDivisors A B] refine SetLike.le_def.mp ?_ (Ideal.mem_span_singleton_self _) rw [← conductor_mul_differentIdeal A K L x hx] exact Ideal.mul_le_left + +section + +include K L in +lemma pow_sub_one_dvd_differentIdeal_aux [IsFractionRing B L] [IsDedekindDomain A] + [NoZeroSMulDivisors A B] [Module.Finite A B] + {p : Ideal A} [p.IsMaximal] (P : Ideal B) {e : ℕ} (he : e ≠ 0) (hp : p ≠ ⊥) + (hP : P ^ e ∣ p.map (algebraMap A B)) : P ^ (e - 1) ∣ differentIdeal A B := by + obtain ⟨a, ha⟩ := (pow_dvd_pow _ (Nat.sub_le e 1)).trans hP + have hp' := (Ideal.map_eq_bot_iff_of_injective + (NoZeroSMulDivisors.algebraMap_injective A B)).not.mpr hp + have habot : a ≠ ⊥ := fun ha' ↦ hp' (by simpa [ha'] using ha) + have hPbot : P ≠ ⊥ := by + rintro rfl; apply hp' + rwa [← Ideal.zero_eq_bot, zero_pow he, zero_dvd_iff, Ideal.zero_eq_bot] at hP + have : p.map (algebraMap A B) ∣ a ^ e := by + obtain ⟨b, hb⟩ := hP + apply_fun (· ^ e : Ideal B → _) at ha + apply_fun (· ^ (e - 1) : Ideal B → _) at hb + simp only [mul_pow, ← pow_mul, mul_comm e] at ha hb + conv_lhs at ha => rw [← Nat.sub_add_cancel (Nat.one_le_iff_ne_zero.mpr he)] + rw [pow_add, hb, mul_assoc, mul_right_inj' (pow_ne_zero _ hPbot), pow_one, mul_comm] at ha + exact ⟨_, ha.symm⟩ + suffices ∀ x ∈ a, intTrace A B x ∈ p by + have hP : ((P ^ (e - 1) : _)⁻¹ : FractionalIdeal B⁰ L) = a / p.map (algebraMap A B) := by + apply inv_involutive.injective + simp only [inv_inv, ha, FractionalIdeal.coeIdeal_mul, inv_div, ne_eq, + FractionalIdeal.coeIdeal_eq_zero, mul_div_assoc] + rw [div_self (by simpa), mul_one] + rw [Ideal.dvd_iff_le, differentialIdeal_le_iff (K := K) (L := L) (pow_ne_zero _ hPbot), hP, + Submodule.map_le_iff_le_comap] + intro x hx + rw [Submodule.restrictScalars_mem, FractionalIdeal.mem_coe, + FractionalIdeal.mem_div_iff_of_nonzero (by simpa using hp')] at hx + rw [Submodule.mem_comap, LinearMap.coe_restrictScalars, ← FractionalIdeal.coe_one, + ← div_self (G₀ := FractionalIdeal A⁰ K) (a := p) (by simpa using hp), + FractionalIdeal.mem_coe, FractionalIdeal.mem_div_iff_of_nonzero (by simpa using hp)] + simp only [FractionalIdeal.mem_coeIdeal, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂] at hx + intro y hy' + obtain ⟨y, hy, rfl : algebraMap A K _ = _⟩ := (FractionalIdeal.mem_coeIdeal _).mp hy' + obtain ⟨z, hz, hz'⟩ := hx _ (Ideal.mem_map_of_mem _ hy) + have : trace K L (algebraMap B L z) ∈ (p : FractionalIdeal A⁰ K) := by + rw [← algebraMap_intTrace (A := A)] + exact ⟨intTrace A B z, this z hz, rfl⟩ + rwa [mul_comm, ← smul_eq_mul, ← LinearMap.map_smul, smul_def, mul_comm, + ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply A B L, ← hz'] + intros x hx + rw [← Ideal.Quotient.eq_zero_iff_mem, ← trace_quotient_eq_of_isDedekindDomain, + ← isNilpotent_iff_eq_zero] + refine trace_isNilpotent_of_isNilpotent ⟨e, ?_⟩ + rw [← map_pow, Ideal.Quotient.eq_zero_iff_mem] + exact (Ideal.dvd_iff_le.mp this) <| Ideal.pow_mem_pow hx _ + +lemma pow_sub_one_dvd_differentIdeal [IsDedekindDomain A] [NoZeroSMulDivisors A B] + [Module.Finite A B] [Algebra.IsSeparable (FractionRing A) (FractionRing B)] + {p : Ideal A} [p.IsMaximal] (P : Ideal B) (e : ℕ) (hp : p ≠ ⊥) + (hP : P ^ e ∣ p.map (algebraMap A B)) : P ^ (e - 1) ∣ differentIdeal A B := by + have : IsLocalization (algebraMapSubmonoid B A⁰) (FractionRing B) := + IsIntegralClosure.isLocalization _ (FractionRing A) _ _ + have : FiniteDimensional (FractionRing A) (FractionRing B) := + Module.Finite_of_isLocalization A B _ _ A⁰ + by_cases he : e = 0 + · rw [he, pow_zero]; exact one_dvd _ + exact pow_sub_one_dvd_differentIdeal_aux A (FractionRing A) (FractionRing B) _ he hp hP + +end diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index e06b0141ff923..5d15d4fac0ac3 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -483,3 +483,22 @@ lemma traceForm_dualBasis_powerBasis_eq [FiniteDimensional K L] [Algebra.IsSepar ring end DetNeZero + +section isNilpotent + +namespace Algebra + +/-- The trace of a nilpotent element is nilpotent. -/ +lemma trace_isNilpotent_of_isNilpotent {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] {x : S} + (hx : IsNilpotent x) : IsNilpotent (trace R S x) := by + by_cases hS : ∃ s : Finset S, Nonempty (Basis s R S) + · obtain ⟨s, ⟨b⟩⟩ := hS + have := Module.Finite.of_basis b + have := (Module.free_def R S).mpr ⟨s, ⟨b⟩⟩ + apply LinearMap.isNilpotent_trace_of_isNilpotent (hx.map (lmul R S)) + · rw [trace_eq_zero_of_not_exists_basis _ hS, LinearMap.zero_apply] + exact IsNilpotent.zero + +end Algebra + +end isNilpotent From 159c3b56230633e79ccee2399eab15ebe2284a0e Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Fri, 25 Oct 2024 17:41:50 +0000 Subject: [PATCH 24/29] feat (RingTheory/HahnSeries/Summable) : Pointwise SMul for summable families (#16701) This PR adds a definition and some API for pointwise SMul between summable families. This generalizes the action of a single Hahn series on a summable family. - [ ] depends on: #16649 [sundry API] - [ ] depends on: #16871 [single summable family] - [ ] depends on: #16872 [Equiv of families] --- Mathlib/RingTheory/HahnSeries/Summable.lean | 215 ++++++++++++++------ 1 file changed, 148 insertions(+), 67 deletions(-) diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index c3b6bf05ac681..c9ea82db6087b 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import Mathlib.Algebra.BigOperators.Finprod import Mathlib.RingTheory.HahnSeries.Multiplication /-! @@ -29,7 +28,6 @@ commutative domain. * Remove unnecessary domain hypotheses. * More general summable families, e.g., define the evaluation homomorphism from a power series ring taking `X` to a positive order element. - * Generalize `SMul` to Hahn modules. ## References - [J. van der Hoeven, *Operators on Generalized Power Series*][van_der_hoeven] @@ -42,7 +40,7 @@ open Pointwise noncomputable section -variable {Γ R α β : Type*} +variable {Γ Γ' R V α β : Type*} namespace HahnSeries @@ -232,7 +230,7 @@ end AddCommMonoid section AddCommGroup -variable [PartialOrder Γ] [AddCommGroup R] {α : Type*} {s t : SummableFamily Γ R α} {a : α} +variable [PartialOrder Γ] [AddCommGroup R] {s t : SummableFamily Γ R α} {a : α} instance : Neg (SummableFamily Γ R α) := ⟨fun s => @@ -267,73 +265,155 @@ theorem sub_apply : (s - t) a = s a - t a := end AddCommGroup -section Semiring - -variable [OrderedCancelAddCommMonoid Γ] [Semiring R] {α : Type*} - -instance : SMul (HahnSeries Γ R) (SummableFamily Γ R α) where - smul x s := - { toFun := fun a => x * s a - isPWO_iUnion_support' := by - apply (x.isPWO_support.add s.isPWO_iUnion_support).mono - refine Set.Subset.trans (Set.iUnion_mono fun a => support_mul_subset_add_support) ?_ - intro g - simp only [Set.mem_iUnion, exists_imp] - exact fun a ha => (Set.add_subset_add (Set.Subset.refl _) (Set.subset_iUnion _ a)) ha - finite_co_support' := fun g => by - apply ((addAntidiagonal x.isPWO_support s.isPWO_iUnion_support g).finite_toSet.biUnion' - fun ij _ => ?_).subset fun a ha => ?_ - · exact fun ij _ => Function.support fun a => (s a).coeff ij.2 - · apply s.finite_co_support - · obtain ⟨i, hi, j, hj, rfl⟩ := support_mul_subset_add_support ha - simp only [exists_prop, Set.mem_iUnion, mem_addAntidiagonal, mul_coeff, mem_support, - isPWO_support, Prod.exists] - exact ⟨i, j, mem_coe.2 (mem_addAntidiagonal.2 ⟨hi, Set.mem_iUnion.2 ⟨a, hj⟩, rfl⟩), hj⟩ } +section SMul + +variable [PartialOrder Γ] [PartialOrder Γ'] [AddCommMonoid R] [AddCommMonoid V] [SMulWithZero R V] + +theorem smul_support_subset_prod (s : SummableFamily Γ R α) + (t : SummableFamily Γ' V β) (gh : Γ × Γ') : + (Function.support fun (i : α × β) ↦ (s i.1).coeff gh.1 • (t i.2).coeff gh.2) ⊆ + ((s.finite_co_support' gh.1).prod (t.finite_co_support' gh.2)).toFinset := by + intro _ hab + simp_all only [Function.mem_support, ne_eq, Set.Finite.coe_toFinset, Set.mem_prod, + Set.mem_setOf_eq] + exact ⟨left_ne_zero_of_smul hab, right_ne_zero_of_smul hab⟩ + +theorem smul_support_finite (s : SummableFamily Γ R α) + (t : SummableFamily Γ' V β) (gh : Γ × Γ') : + (Function.support fun (i : α × β) ↦ (s i.1).coeff gh.1 • (t i.2).coeff gh.2).Finite := + Set.Finite.subset (Set.toFinite ((s.finite_co_support' gh.1).prod + (t.finite_co_support' gh.2)).toFinset) (smul_support_subset_prod s t gh) + +variable [VAdd Γ Γ'] [IsOrderedCancelVAdd Γ Γ'] + +open HahnModule + +theorem isPWO_iUnion_support_prod_smul {s : α → HahnSeries Γ R} {t : β → HahnSeries Γ' V} + (hs : (⋃ a, (s a).support).IsPWO) (ht : (⋃ b, (t b).support).IsPWO) : + (⋃ (a : α × β), ((fun a ↦ (of R).symm + ((s a.1) • (of R) (t a.2))) a).support).IsPWO := by + apply (hs.vadd ht).mono + have hsupp : ∀ ab : α × β, support ((fun ab ↦ (of R).symm (s ab.1 • (of R) (t ab.2))) ab) ⊆ + (s ab.1).support +ᵥ (t ab.2).support := by + intro ab + refine Set.Subset.trans (fun x hx => ?_) (support_vaddAntidiagonal_subset_vadd + (hs := (s ab.1).isPWO_support) (ht := (t ab.2).isPWO_support)) + contrapose! hx + simp only [Set.mem_setOf_eq, not_nonempty_iff_eq_empty] at hx + rw [mem_support, not_not, HahnModule.smul_coeff, hx, sum_empty] + refine Set.Subset.trans (Set.iUnion_mono fun a => (hsupp a)) ?_ + simp_all only [Set.iUnion_subset_iff, Prod.forall] + exact fun a b => Set.vadd_subset_vadd (Set.subset_iUnion_of_subset a fun x y ↦ y) + (Set.subset_iUnion_of_subset b fun x y ↦ y) + +theorem finite_co_support_prod_smul (s : SummableFamily Γ R α) + (t : SummableFamily Γ' V β) (g : Γ') : + Finite {(ab : α × β) | + ((fun (ab : α × β) ↦ (of R).symm (s ab.1 • (of R) (t ab.2))) ab).coeff g ≠ 0} := by + apply ((VAddAntidiagonal s.isPWO_iUnion_support t.isPWO_iUnion_support g).finite_toSet.biUnion' + (fun gh _ => smul_support_finite s t gh)).subset _ + exact fun ab hab => by + simp only [smul_coeff, ne_eq, Set.mem_setOf_eq] at hab + obtain ⟨ij, hij⟩ := Finset.exists_ne_zero_of_sum_ne_zero hab + simp only [mem_coe, mem_vaddAntidiagonal, Set.mem_iUnion, mem_support, ne_eq, + Function.mem_support, exists_prop, Prod.exists] + exact ⟨ij.1, ij.2, ⟨⟨ab.1, left_ne_zero_of_smul hij.2⟩, ⟨ab.2, right_ne_zero_of_smul hij.2⟩, + ((mem_vaddAntidiagonal _ _ _).mp hij.1).2.2⟩, hij.2⟩ + +/-- An elementwise scalar multiplication of one summable family on another. -/ +@[simps] +def FamilySMul (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : + (SummableFamily Γ' V (α × β)) where + toFun ab := (of R).symm (s (ab.1) • ((of R) (t (ab.2)))) + isPWO_iUnion_support' := + isPWO_iUnion_support_prod_smul s.isPWO_iUnion_support t.isPWO_iUnion_support + finite_co_support' g := finite_co_support_prod_smul s t g + +theorem sum_vAddAntidiagonal_eq (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') + (a : α × β) : + ∑ x ∈ VAddAntidiagonal (s a.1).isPWO_support' (t a.2).isPWO_support' g, (s a.1).coeff x.1 • + (t a.2).coeff x.2 = ∑ x ∈ VAddAntidiagonal s.isPWO_iUnion_support' t.isPWO_iUnion_support' g, + (s a.1).coeff x.1 • (t a.2).coeff x.2 := by + refine sum_subset (fun gh hgh => ?_) fun gh hgh h => ?_ + · simp_all only [mem_vaddAntidiagonal, Function.mem_support, Set.mem_iUnion, mem_support] + exact ⟨Exists.intro a.1 hgh.1, Exists.intro a.2 hgh.2.1, trivial⟩ + · by_cases hs : (s a.1).coeff gh.1 = 0 + · exact smul_eq_zero_of_left hs ((t a.2).coeff gh.2) + · simp_all + +theorem family_smul_coeff {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] + (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) (g : Γ') : + (FamilySMul s t).hsum.coeff g = ∑ gh ∈ VAddAntidiagonal s.isPWO_iUnion_support + t.isPWO_iUnion_support g, (s.hsum.coeff gh.1) • (t.hsum.coeff gh.2) := by + rw [hsum_coeff] + simp only [hsum_coeff_eq_sum, FamilySMul_toFun, HahnModule.smul_coeff, Equiv.symm_apply_apply] + simp_rw [sum_vAddAntidiagonal_eq, Finset.smul_sum, Finset.sum_smul] + rw [← sum_finsum_comm _ _ <| fun gh _ => smul_support_finite s t gh] + refine sum_congr rfl fun gh _ => ?_ + rw [finsum_eq_sum _ (smul_support_finite s t gh), ← sum_product_right'] + refine sum_subset (fun ab hab => ?_) (fun ab _ hab => by simp_all) + have hsupp := smul_support_subset_prod s t gh + simp_all only [mem_vaddAntidiagonal, Set.mem_iUnion, mem_support, ne_eq, Set.Finite.mem_toFinset, + Function.mem_support, Set.Finite.coe_toFinset, support_subset_iff, Set.mem_prod, + Set.mem_setOf_eq, Prod.forall, coeff_support, mem_product] + exact hsupp ab.1 ab.2 hab + +theorem hsum_family_smul {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] + (s : SummableFamily Γ R α) (t : SummableFamily Γ' V β) : + (FamilySMul s t).hsum = (of R).symm (s.hsum • (of R) (t.hsum)) := by + ext g + rw [family_smul_coeff s t g, HahnModule.smul_coeff, Equiv.symm_apply_apply] + refine Eq.symm (sum_of_injOn (fun a ↦ a) (fun _ _ _ _ h ↦ h) (fun _ hgh => ?_) + (fun gh _ hgh => ?_) fun _ _ => by simp) + · simp_all only [mem_coe, mem_vaddAntidiagonal, mem_support, ne_eq, Set.mem_iUnion, and_true] + constructor + · rw [hsum_coeff_eq_sum] at hgh + have h' := Finset.exists_ne_zero_of_sum_ne_zero hgh.1 + simpa using h' + · by_contra hi + simp_all + · simp only [Set.image_id', mem_coe, mem_vaddAntidiagonal, mem_support, ne_eq, not_and] at hgh + by_cases h : s.hsum.coeff gh.1 = 0 + · exact smul_eq_zero_of_left h (t.hsum.coeff gh.2) + · simp_all + +instance [AddCommMonoid R] [SMulWithZero R V] : SMul (HahnSeries Γ R) (SummableFamily Γ' V β) where + smul x t := Equiv (Equiv.punitProd β) <| FamilySMul (single x) t + +theorem smul_eq {x : HahnSeries Γ R} {t : SummableFamily Γ' V β} : + x • t = Equiv (Equiv.punitProd β) (FamilySMul (single x) t) := + rfl @[simp] -theorem smul_apply {x : HahnSeries Γ R} {s : SummableFamily Γ R α} {a : α} : (x • s) a = x * s a := +theorem smul_apply {x : HahnSeries Γ R} {s : SummableFamily Γ' V α} {a : α} : + (x • s) a = (of R).symm (x • of R (s a)) := rfl -instance : Module (HahnSeries Γ R) (SummableFamily Γ R α) where +@[simp] +theorem hsum_smul_module {R} {V} [Semiring R] [AddCommMonoid V] [Module R V] {x : HahnSeries Γ R} + {s : SummableFamily Γ' V α} : + (x • s).hsum = (of R).symm (x • of R s.hsum) := by + rw [smul_eq, hsum_equiv, hsum_family_smul, hsum_single] + +end SMul + +section Semiring + +variable [OrderedCancelAddCommMonoid Γ] [PartialOrder Γ'] [AddAction Γ Γ'] + [IsOrderedCancelVAdd Γ Γ'] [Semiring R] + +instance [AddCommMonoid V] [Module R V] : Module (HahnSeries Γ R) (SummableFamily Γ' V α) where smul := (· • ·) - smul_zero _ := ext fun _ => mul_zero _ - zero_smul _ := ext fun _ => zero_mul _ - one_smul _ := ext fun _ => one_mul _ - add_smul _ _ _ := ext fun _ => add_mul _ _ _ - smul_add _ _ _ := ext fun _ => mul_add _ _ _ - mul_smul _ _ _ := ext fun _ => mul_assoc _ _ _ + smul_zero _ := ext fun _ => by simp + zero_smul _ := ext fun _ => by simp + one_smul _ := ext fun _ => by rw [smul_apply, HahnModule.one_smul', Equiv.symm_apply_apply] + add_smul _ _ _ := ext fun _ => by simp [add_smul] + smul_add _ _ _ := ext fun _ => by simp + mul_smul _ _ _ := ext fun _ => by simp [HahnModule.instModule.mul_smul] -@[simp] -theorem hsum_smul {x : HahnSeries Γ R} {s : SummableFamily Γ R α} : (x • s).hsum = x * s.hsum := by - ext g - simp only [mul_coeff, hsum_coeff, smul_apply] - refine - (Eq.trans (finsum_congr fun a => ?_) - (finsum_sum_comm (addAntidiagonal x.isPWO_support s.isPWO_iUnion_support g) - (fun i ij => x.coeff (Prod.fst ij) * (s i).coeff ij.snd) ?_)).trans - ?_ - · refine sum_subset (addAntidiagonal_mono_right - (Set.subset_iUnion (fun j => support (toFun s j)) a)) ?_ - rintro ⟨i, j⟩ hU ha - rw [mem_addAntidiagonal] at * - rw [Classical.not_not.1 fun con => ha ⟨hU.1, con, hU.2.2⟩, mul_zero] - · rintro ⟨i, j⟩ _ - refine (s.finite_co_support j).subset ?_ - simp_rw [Function.support_subset_iff', Function.mem_support, Classical.not_not] - intro a ha - rw [ha, mul_zero] - · refine (sum_congr rfl ?_).trans (sum_subset (addAntidiagonal_mono_right ?_) ?_).symm - · rintro ⟨i, j⟩ _ - rw [mul_finsum] - apply s.finite_co_support - · intro x hx - simp only [Set.mem_iUnion, Ne, mem_support] - contrapose! hx - simp [hx] - · rintro ⟨i, j⟩ hU ha - rw [mem_addAntidiagonal] at * - rw [← hsum_coeff, Classical.not_not.1 fun con => ha ⟨hU.1, con, hU.2.2⟩, - mul_zero] +theorem hsum_smul {x : HahnSeries Γ R} {s : SummableFamily Γ R α} : + (x • s).hsum = x * s.hsum := by + rw [hsum_smul_module, of_symm_smul_of_eq_mul] /-- The summation of a `summable_family` as a `LinearMap`. -/ @[simps] @@ -441,6 +521,7 @@ section powers variable [LinearOrderedCancelAddCommMonoid Γ] [CommRing R] [IsDomain R] /-- The powers of an element of positive valuation form a summable family. -/ +@[simps] def powers (x : HahnSeries Γ R) (hx : 0 < x.orderTop) : SummableFamily Γ R ℕ where toFun n := x ^ n isPWO_iUnion_support' := isPWO_iUnion_support_powers hx @@ -486,8 +567,8 @@ theorem embDomain_succ_smul_powers : rw [Set.mem_range, not_exists] exact Nat.succ_ne_zero · refine Eq.trans (embDomain_image _ ⟨Nat.succ, Nat.succ_injective⟩) ?_ - simp only [pow_succ', coe_powers, coe_sub, smul_apply, coe_ofFinsupp, Pi.sub_apply] - rw [Finsupp.single_eq_of_ne n.succ_ne_zero.symm, sub_zero] + rw [smul_apply, powers_toFun, coe_sub, coe_powers, Pi.sub_apply, coe_ofFinsupp, pow_succ', + Finsupp.single_eq_of_ne (Nat.zero_ne_add_one n), sub_zero, of_symm_smul_of_eq_mul] theorem one_sub_self_mul_hsum_powers : (1 - x) * (powers x hx).hsum = 1 := by rw [← hsum_smul, sub_smul 1 x (powers x hx), one_smul, hsum_sub, ← From 6ac1e12b7a0ab36e6738771fb6dfaa8bac131916 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 25 Oct 2024 18:25:24 +0000 Subject: [PATCH 25/29] feat: add `isCompact_{quasi}spectrum` for instances of the continuous functional calculus (#18117) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We already have instances for `CompactSpace`, but it's handy to have these when you need to work with the sets themselves (e.g., to show that `spectrum R a ∪ spectrum R b` is compact). --- .../CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean | 5 +++++ .../CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean | 5 +++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 73105d7530e9b..ed76c4a88dc63 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -106,6 +106,11 @@ variable [TopologicalSemiring R] [ContinuousStar R] [NonUnitalRing A] [StarRing variable [TopologicalSpace A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] variable [instCFCₙ : NonUnitalContinuousFunctionalCalculus R p] +include instCFCₙ in +lemma NonUnitalContinuousFunctionalCalculus.isCompact_quasispectrum (a : A) : + IsCompact (σₙ R a) := + isCompact_iff_compactSpace.mpr inferInstance + lemma NonUnitalStarAlgHom.ext_continuousMap [UniqueNonUnitalContinuousFunctionalCalculus R A] (a : A) (φ ψ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) (hφ : Continuous φ) (hψ : Continuous ψ) (h : φ ⟨.restrict (σₙ R a) <| .id R, rfl⟩ = ψ ⟨.restrict (σₙ R a) <| .id R, rfl⟩) : diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index b63d226c61d6d..b193f3c40b439 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -200,6 +200,11 @@ variable {R A : Type*} {p : A → Prop} [CommSemiring R] [StarRing R] [MetricSpa variable [TopologicalSemiring R] [ContinuousStar R] [TopologicalSpace A] [Ring A] [StarRing A] variable [Algebra R A] [instCFC : ContinuousFunctionalCalculus R p] +include instCFC in +lemma ContinuousFunctionalCalculus.isCompact_spectrum (a : A) : + IsCompact (spectrum R a) := + isCompact_iff_compactSpace.mpr inferInstance + lemma StarAlgHom.ext_continuousMap [UniqueContinuousFunctionalCalculus R A] (a : A) (φ ψ : C(spectrum R a, R) →⋆ₐ[R] A) (hφ : Continuous φ) (hψ : Continuous ψ) (h : φ (.restrict (spectrum R a) <| .id R) = ψ (.restrict (spectrum R a) <| .id R)) : From 970d241d0868926adcfcc688a4a2b846fe732ea3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 25 Oct 2024 18:52:56 +0000 Subject: [PATCH 26/29] refactor: delete IsWellOrderLimitElement.lean (#15904) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I understand this is quite a bold move, and I hope not to come off as confrontational. I argue that `Order/IsWellOrderLimitElement.lean` duplicates existing API, and that all the theorems proven within it already exist in some other form. First, note that the typeclass assumptions `LinearOrder α` + `IsWellOrder α (· < ·)` **very nearly** imply `ConditionallyCompleteLinearOrderBot α` through [`IsWellOrder.conditionallyCompleteLinearOrderBot`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/ConditionallyCompleteLattice/Basic.html#IsWellOrder.conditionallyCompleteLinearOrderBot). The only missing assumption is `OrderBot α`, which actually follows from `Nonempty α` through [`WellFoundedLT.toOrderBot`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/WellFounded.html#WellFoundedLT.toOrderBot). All theorems in this file assume the existence of some `a : α`, so that's covered. Most of the only other file importing this one assumes `OrderBot α` directly. `ConditionallyCompleteLinearOrder α` in turn implies `SuccOrder α` through `ConditionallyCompleteLinearOrder.toSuccOrder` (#15863), so assuming it yields no loss of generality. In fact, doing things this way means we can set nicer def-eqs for `@succ α`, such as `succ o = o + 1` on ordinals. As for `IsWellOrderLimitElement`, we already have [`Order.IsSuccLimit`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit) with only trivial typeclass assumptions. To summarize, here are the correspondences between definitions and theorems in this file and other definitions and theorems in Mathlib, applicable under equivalent typeclasss assumptions: - `wellOrderSucc` → [`Order.succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.succ) - `self_le_wellOrderSucc` → [`Order.le_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.le_succ) - `wellOrderSucc_le` → [`Order.succ_le_of_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.succ_le_of_lt) - `self_lt_wellOrderSucc` → [`Order.lt_succ_of_not_isMax`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.lt_succ_of_not_isMax) - `le_of_lt_wellOrderSucc` → [`Order.le_of_lt_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.le_of_lt_succ). - `IsWellOrderLimitElement`→ [`Order.IsSuccLimit`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit) - `IsWellOrderLimitElement.neq_bot` → (already implied by `IsSuccLimit`) - `IsWellOrderLimitElement.bot_lt` → [`Ne.bot_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/BoundedOrder.html#Ne.bot_lt) - `IsWellOrderLimitElement.wellOrderSucc_lt` → [`Order.IsSuccLimit.succ_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit.succ_lt) - `eq_bot_or_eq_succ_or_isWellOrderLimitElement` → [`Order.isSuccLimitRecOn`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.isSuccLimitRecOn) - `Nat.wellOrderSucc_eq` → [`Nat.succ_eq_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Nat/SuccPred.html#Nat.succ_eq_succ) - `Nat.not_isWellOrderLimitElement` → [`Order.IsSuccLimit.isMin_of_noMax`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit.isMin_of_noMax) Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 - .../CategoryTheory/SmallObject/Iteration.lean | 62 +++++++----- Mathlib/Order/IsWellOrderLimitElement.lean | 98 ------------------- 3 files changed, 35 insertions(+), 126 deletions(-) delete mode 100644 Mathlib/Order/IsWellOrderLimitElement.lean diff --git a/Mathlib.lean b/Mathlib.lean index b0ddfe0730762..0307049fa8db8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3801,7 +3801,6 @@ import Mathlib.Order.Interval.Set.SurjOn import Mathlib.Order.Interval.Set.UnorderedInterval import Mathlib.Order.Interval.Set.WithBotTop import Mathlib.Order.Irreducible -import Mathlib.Order.IsWellOrderLimitElement import Mathlib.Order.Iterate import Mathlib.Order.JordanHolder import Mathlib.Order.KonigLemma diff --git a/Mathlib/CategoryTheory/SmallObject/Iteration.lean b/Mathlib/CategoryTheory/SmallObject/Iteration.lean index a9470418b418d..f4d4af8823a9f 100644 --- a/Mathlib/CategoryTheory/SmallObject/Iteration.lean +++ b/Mathlib/CategoryTheory/SmallObject/Iteration.lean @@ -5,7 +5,8 @@ Authors: Joël Riou -/ import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.Limits.IsLimit -import Mathlib.Order.IsWellOrderLimitElement +import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.SuccPred.Limit /-! # Transfinite iterations of a functor @@ -37,7 +38,7 @@ namespace CategoryTheory open Category Limits variable {C : Type*} [Category C] {Φ : C ⥤ C} (ε : 𝟭 C ⟶ Φ) - {J : Type u} [LinearOrder J] + {J : Type u} [Preorder J] namespace Functor @@ -45,11 +46,11 @@ namespace Iteration variable {j : J} (F : { i // i ≤ j } ⥤ C) -/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨wellOrderSucc i, _⟩` when `F : { i // i ≤ j } ⥤ C` +/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨Order.succ i, _⟩` when `F : { i // i ≤ j } ⥤ C` and `i : J` is such that `i < j`. -/ -noncomputable abbrev mapSucc' [IsWellOrder J (· < ·)] (i : J) (hi : i < j) : - F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ := - F.map (homOfLE (by simpa only [Subtype.mk_le_mk] using self_le_wellOrderSucc i)) +noncomputable abbrev mapSucc' [SuccOrder J] (i : J) (hi : i < j) : + F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ := + F.map <| homOfLE <| Subtype.mk_le_mk.2 <| Order.le_succ i variable {i : J} (hi : i ≤ j) @@ -78,29 +79,26 @@ def coconeOfLE : Cocone (restrictionLT F hi) where end Iteration -variable [IsWellOrder J (· < ·)] [OrderBot J] - /-- The category of `j`th iterations of a functor `Φ` equipped with a natural transformation `ε : 𝟭 C ⟶ Φ`. An object consists of the data of all iterations of `Φ` for `i : J` such that `i ≤ j` (this is the field `F`). Such objects are equipped with data and properties which characterizes the iterations up to a unique isomorphism for the three types of elements: `⊥`, successors, limit elements. -/ -structure Iteration (j : J) where +structure Iteration [Preorder J] [OrderBot J] [SuccOrder J] (j : J) where /-- The data of all `i`th iterations for `i : J` such that `i ≤ j`. -/ F : { i // i ≤ j } ⥤ C ⥤ C /-- The zeroth iteration is the identity functor. -/ isoZero : F.obj ⟨⊥, bot_le⟩ ≅ 𝟭 C /-- The iteration on a successor element is obtained by composition of the previous iteration with `Φ`. -/ - isoSucc (i : J) (hi : i < j) : - F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ ≅ F.obj ⟨i, hi.le⟩ ⋙ Φ + isoSucc (i : J) (hi : i < j) : F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ ≅ F.obj ⟨i, hi.le⟩ ⋙ Φ /-- The natural map from an iteration to its successor is induced by `ε`. -/ mapSucc'_eq (i : J) (hi : i < j) : Iteration.mapSucc' F i hi = whiskerLeft _ ε ≫ (isoSucc i hi).inv /-- If `i` is a limit element, the `i`th iteration is the colimit of `k`th iterations for `k < i`. -/ - isColimit (i : J) [IsWellOrderLimitElement i] (hi : i ≤ j) : - IsColimit (Iteration.coconeOfLE F hi) + isColimit (i : J) (hi : Order.IsSuccLimit i) (hij : i ≤ j) : + IsColimit (Iteration.coconeOfLE F hij) namespace Iteration @@ -109,12 +107,12 @@ variable {j : J} section -variable (iter : Φ.Iteration ε j) +variable [OrderBot J] [SuccOrder J] (iter : Φ.Iteration ε j) /-- For `iter : Φ.Iteration.ε j`, this is the map -`iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨wellOrderSucc i, _⟩` if `i : J` is such that `i < j`. -/ +`iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨Order.succ i, _⟩` if `i : J` is such that `i < j`. -/ noncomputable abbrev mapSucc (i : J) (hi : i < j) : - iter.F.obj ⟨i, hi.le⟩ ⟶ iter.F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ := + iter.F.obj ⟨i, hi.le⟩ ⟶ iter.F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ := mapSucc' iter.F i hi lemma mapSucc_eq (i : J) (hi : i < j) : @@ -123,7 +121,7 @@ lemma mapSucc_eq (i : J) (hi : i < j) : end -variable (iter₁ iter₂ : Φ.Iteration ε j) +variable [OrderBot J] [SuccOrder J] (iter₁ iter₂ : Φ.Iteration ε j) /-- A morphism between two objects `iter₁` and `iter₂` in the category `Φ.Iteration ε j` of `j`th iterations of a functor `Φ` @@ -136,7 +134,7 @@ structure Hom where natTrans_app_zero : natTrans.app ⟨⊥, bot_le⟩ = iter₁.isoZero.hom ≫ iter₂.isoZero.inv := by aesop_cat natTrans_app_succ (i : J) (hi : i < j) : - natTrans.app ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ = (iter₁.isoSucc i hi).hom ≫ + natTrans.app ⟨Order.succ i, Order.succ_le_of_lt hi⟩ = (iter₁.isoSucc i hi).hom ≫ whiskerRight (natTrans.app ⟨i, hi.le⟩) _ ≫ (iter₂.isoSucc i hi).inv := by aesop_cat namespace Hom @@ -172,17 +170,27 @@ instance : Category (Iteration ε j) where id := id comp := comp -instance : Subsingleton (iter₁ ⟶ iter₂) where - allEq f g := ext' (by - let P := fun (i : J) => ∀ (hi : i ≤ j), f.natTrans.app ⟨i, hi⟩ = g.natTrans.app ⟨i, hi⟩ - suffices ∀ (i : J), P i by +instance {J} {j : J} [PartialOrder J] [OrderBot J] [WellFoundedLT J] [SuccOrder J] + {iter₁ iter₂ : Iteration ε j} : + Subsingleton (iter₁ ⟶ iter₂) where + allEq f g := by + apply ext' + suffices ∀ i hi, f.natTrans.app ⟨i, hi⟩ = g.natTrans.app ⟨i, hi⟩ by ext ⟨i, hi⟩ : 2 apply this - refine fun _ => WellFoundedLT.induction _ (fun i hi hi' => ?_) - obtain rfl|⟨i, rfl, hi''⟩|_ := eq_bot_or_eq_succ_or_isWellOrderLimitElement i - · simp only [natTrans_app_zero] - · simp only [Hom.natTrans_app_succ _ i (lt_of_lt_of_le hi'' hi'), hi i hi''] - · exact (iter₁.isColimit i hi').hom_ext (fun ⟨k, hk⟩ => by simp [hi k hk])) + intro i + induction i using SuccOrder.limitRecOn with + | hm j H => + have := H.eq_bot + subst this + simp [natTrans_app_zero] + | hs j H IH => + intro hj + simp [Hom.natTrans_app_succ, IH, (Order.lt_succ_of_not_isMax H).trans_le hj] + | hl j H IH => + apply fun hj ↦ (iter₁.isColimit j H hj).hom_ext ?_ + rintro ⟨k, hk⟩ + simp [IH k hk] end Hom diff --git a/Mathlib/Order/IsWellOrderLimitElement.lean b/Mathlib/Order/IsWellOrderLimitElement.lean deleted file mode 100644 index 2d6bde7b9282b..0000000000000 --- a/Mathlib/Order/IsWellOrderLimitElement.lean +++ /dev/null @@ -1,98 +0,0 @@ -/- -Copyright (c) 2024 Joël Riou. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joël Riou --/ -import Mathlib.Algebra.Order.Group.Nat -import Mathlib.Order.WellFounded - -/-! -# Limit elements in well-ordered types - -This file introduces two main definitions: -- `wellOrderSucc a`: the successor of an element `a` in a well-ordered type -- the typeclass `IsWellOrderLimitElement a` which asserts that an element `a` (in a -well-ordered type) is neither a successor nor the smallest element, i.e. `a` is a limit element - -Then, the lemma `eq_bot_or_eq_succ_or_isWellOrderLimitElement` shows that an element -in a well-ordered type is either `⊥`, a successor, or a limit element. - --/ - -variable {α : Type*} [LinearOrder α] - -section -variable [IsWellOrder α (· < ·)] - -/-- Given an element `a : α` in a well ordered set, this is the successor of `a`, -i.e. the smallest element strictly greater than `a` if it exists (or `a` itself otherwise). -/ -noncomputable def wellOrderSucc (a : α) : α := - (IsWellFounded.wf (r := (· < ·))).succ a - -lemma self_le_wellOrderSucc (a : α) : a ≤ wellOrderSucc a := by - by_cases h : ∃ b, a < b - · exact (IsWellFounded.wf.lt_succ h).le - · dsimp [wellOrderSucc, WellFounded.succ] - rw [dif_neg h] - -lemma wellOrderSucc_le {a b : α} (ha : a < b) : wellOrderSucc a ≤ b := by - dsimp [wellOrderSucc, WellFounded.succ] - rw [dif_pos ⟨_, ha⟩] - exact WellFounded.min_le _ ha - -lemma self_lt_wellOrderSucc {a b : α} (h : a < b) : a < wellOrderSucc a := - IsWellFounded.wf.lt_succ ⟨b, h⟩ - -lemma le_of_lt_wellOrderSucc {a b : α} (h : a < wellOrderSucc b) : a ≤ b := by - by_contra! - simpa using lt_of_lt_of_le h (wellOrderSucc_le this) - -/-- This property of an element `a : α` in a well-ordered type holds iff `a` is a -limit element, i.e. it is not a successor and it is not the smallest element of `α`. -/ -class IsWellOrderLimitElement (a : α) : Prop where - not_bot : ∃ (b : α), b < a - not_succ (b : α) (hb : b < a) : ∃ (c : α), b < c ∧ c < a - -end - -variable (a : α) [ha : IsWellOrderLimitElement a] - -lemma IsWellOrderLimitElement.neq_bot [OrderBot α] : a ≠ ⊥ := by - rintro rfl - obtain ⟨b, hb⟩ := ha.not_bot - simp at hb - -lemma IsWellOrderLimitElement.bot_lt [OrderBot α] : ⊥ < a := by - obtain h|h := eq_or_lt_of_le (@bot_le _ _ _ a) - · exact (IsWellOrderLimitElement.neq_bot a h.symm).elim - · exact h - -variable {a} -variable [IsWellOrder α (· < ·)] - -lemma IsWellOrderLimitElement.wellOrderSucc_lt {b : α} (hb : b < a) : - wellOrderSucc b < a := by - obtain ⟨c, hc₁, hc₂⟩ := ha.not_succ b hb - exact lt_of_le_of_lt (wellOrderSucc_le hc₁) hc₂ - -lemma eq_bot_or_eq_succ_or_isWellOrderLimitElement [OrderBot α] (a : α) : - a = ⊥ ∨ (∃ b, a = wellOrderSucc b ∧ b < a) ∨ IsWellOrderLimitElement a := by - refine or_iff_not_imp_left.2 <| fun h₁ ↦ or_iff_not_imp_left.2 <| fun h₂ ↦ ?_ - refine (IsWellOrderLimitElement.mk ⟨⊥, Ne.bot_lt h₁⟩ fun b hb ↦ ?_) - obtain rfl | h₃ := eq_or_lt_of_le (wellOrderSucc_le hb) - · exact (h₂ ⟨b, rfl, hb⟩).elim - · exact ⟨wellOrderSucc b, self_lt_wellOrderSucc hb, h₃⟩ - -lemma IsWellOrderLimitElement.neq_succ (a : α) (ha : a < wellOrderSucc a) - [IsWellOrderLimitElement (wellOrderSucc a)] : False := by - simpa using IsWellOrderLimitElement.wellOrderSucc_lt ha - -@[simp] -lemma Nat.wellOrderSucc_eq (a : ℕ) : wellOrderSucc a = succ a := - le_antisymm (wellOrderSucc_le (Nat.lt_succ_self a)) - (Nat.succ_le.1 (self_lt_wellOrderSucc (Nat.lt_succ_self a))) - -lemma Nat.not_isWellOrderLimitElement (a : ℕ) [IsWellOrderLimitElement a] : False := by - obtain _|a := a - · simpa using IsWellOrderLimitElement.neq_bot (0 : ℕ) - · simpa using IsWellOrderLimitElement.wellOrderSucc_lt (Nat.lt_succ_self a) From 2a1cffe71f425945911deaebfd02a990443fbb24 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Fri, 25 Oct 2024 19:23:21 +0000 Subject: [PATCH 27/29] feature(LinearAlgebra/QuadraticForm/Basis): toBilinHom (#18212) For each basis `bm`, `QuadraticMap.toBilinHom` is the linear map that sends a quadratic map on a module `M` over `R` to its associated symmetric bilinear map. For free modules, `toBilinHom` is the `toBilin` equivalent of `QuadraticMap.associatedHom` Co-authored-by: Eric Wieser Co-authored-by: Christopher Hoskin --- .../LinearAlgebra/QuadraticForm/Basis.lean | 30 +++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean index c2c99925d7bb1..40f483a2ba81e 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basis.lean @@ -13,6 +13,8 @@ does not require `Invertible (2 : R)`. Unlike that definition, this only works i a basis. -/ +open LinearMap (BilinMap) + namespace QuadraticMap variable {ι R M N} [LinearOrder ι] @@ -60,4 +62,32 @@ theorem _root_.LinearMap.BilinMap.toQuadraticMap_surjective [Module.Free R M] : letI : LinearOrder ι := IsWellOrder.linearOrder WellOrderingRel exact ⟨_, toQuadraticMap_toBilin _ b⟩ +@[simp] +lemma add_toBilin (bm : Basis ι R M) (Q₁ Q₂ : QuadraticMap R M N) : + (Q₁ + Q₂).toBilin bm = Q₁.toBilin bm + Q₂.toBilin bm := by + refine bm.ext fun i => bm.ext fun j => ?_ + obtain h | rfl | h := lt_trichotomy i j + · simp [h.ne, h, toBilin_apply, polar_add] + · simp [toBilin_apply] + · simp [h.ne', h.not_lt, toBilin_apply, polar_add] + +variable (S) [CommSemiring S] [Algebra S R] +variable [Module S N] [IsScalarTower S R N] + +@[simp] +lemma smul_toBilin (bm : Basis ι R M) (s : S) (Q : QuadraticMap R M N) : + (s • Q).toBilin bm = s • Q.toBilin bm := by + refine bm.ext fun i => bm.ext fun j => ?_ + obtain h | rfl | h := lt_trichotomy i j + · simp [h.ne, h, toBilin_apply, polar_smul] + · simp [toBilin_apply] + · simp [h.ne', h.not_lt, toBilin_apply] + +/-- `QuadraticMap.toBilin` as an S-linear map -/ +@[simps] +noncomputable def toBilinHom (bm : Basis ι R M) : QuadraticMap R M N →ₗ[S] BilinMap R M N where + toFun Q := Q.toBilin bm + map_add' := add_toBilin bm + map_smul' := smul_toBilin S bm + end QuadraticMap From eba9c377332bae1a588c85fc07a1a4012ea29299 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 25 Oct 2024 20:03:54 +0000 Subject: [PATCH 28/29] chore(Data/Nat/Prime): split `Defs`; rearrange `Basic` (#18233) This PR reduces the theory included with `Data/Nat/Prime/Defs.lean`, in particular allowing us to define `CharP` without needing theory on ordered monoids. Most of the contents got moved to `Basic.lean`, which would increase the downstream imports for quite a few file. So in turn I split up `Basic.lean` based on the imports and logical structure: * Results related to the factorial go to `Prime/Factorial.lean` * Results related to integers go to `Prime/Int.lean` (which should probably be merged with `Data/Int/NatPrime.lean) * The proof that there are infinitely many primes goes to `Prime/Infinite.lean` * Some results on powers of prime that needed a higher amount of ordered monoid theory go to `Prime/Pow.lean` This is not a universal import win, since some files that imported `Defs.lean` now need to import `Basic.lean` instead, but I think that's worth it for keeping `Defs` files succinct. --- Archive/Imo/Imo2019Q4.lean | 2 +- Mathlib.lean | 4 + Mathlib/Algebra/Order/Floor/Prime.lean | 2 +- Mathlib/Data/Int/NatPrime.lean | 1 + Mathlib/Data/Nat/Choose/Dvd.lean | 2 +- Mathlib/Data/Nat/Factorization/PrimePow.lean | 1 + Mathlib/Data/Nat/Factors.lean | 2 +- Mathlib/Data/Nat/Prime/Basic.lean | 105 +++++------------- Mathlib/Data/Nat/Prime/Defs.lean | 28 +---- Mathlib/Data/Nat/Prime/Factorial.lean | 28 +++++ Mathlib/Data/Nat/Prime/Infinite.lean | 49 ++++++++ Mathlib/Data/Nat/Prime/Int.lean | 47 ++++++++ Mathlib/Data/Nat/Prime/Pow.lean | 31 ++++++ Mathlib/Data/Nat/PrimeFin.lean | 2 +- Mathlib/Dynamics/PeriodicPts.lean | 1 + .../LegendreSymbol/GaussEisensteinLemmas.lean | 3 +- Mathlib/NumberTheory/Multiplicity.lean | 2 +- .../NumberTheory/Padics/PadicVal/Basic.lean | 1 + Mathlib/NumberTheory/Primorial.lean | 2 +- Mathlib/RingTheory/Int/Basic.lean | 4 +- Mathlib/Tactic/NormNum/Prime.lean | 2 +- test/observe.lean | 2 +- 22 files changed, 206 insertions(+), 115 deletions(-) create mode 100644 Mathlib/Data/Nat/Prime/Factorial.lean create mode 100644 Mathlib/Data/Nat/Prime/Infinite.lean create mode 100644 Mathlib/Data/Nat/Prime/Int.lean create mode 100644 Mathlib/Data/Nat/Prime/Pow.lean diff --git a/Archive/Imo/Imo2019Q4.lean b/Archive/Imo/Imo2019Q4.lean index ddf0040f14bd9..fe10780d6586f 100644 --- a/Archive/Imo/Imo2019Q4.lean +++ b/Archive/Imo/Imo2019Q4.lean @@ -5,7 +5,7 @@ Authors: Floris van Doorn -/ import Mathlib.Data.Nat.Factorial.BigOperators import Mathlib.Data.Nat.Multiplicity -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Int import Mathlib.Tactic.IntervalCases import Mathlib.Tactic.GCongr diff --git a/Mathlib.lean b/Mathlib.lean index 0307049fa8db8..61df8bf2227e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2544,6 +2544,10 @@ import Mathlib.Data.Nat.PartENat import Mathlib.Data.Nat.Periodic import Mathlib.Data.Nat.Prime.Basic import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Factorial +import Mathlib.Data.Nat.Prime.Infinite +import Mathlib.Data.Nat.Prime.Int +import Mathlib.Data.Nat.Prime.Pow import Mathlib.Data.Nat.PrimeFin import Mathlib.Data.Nat.Set import Mathlib.Data.Nat.Size diff --git a/Mathlib/Algebra/Order/Floor/Prime.lean b/Mathlib/Algebra/Order/Floor/Prime.lean index f3545648ae511..506029de21811 100644 --- a/Mathlib/Algebra/Order/Floor/Prime.lean +++ b/Mathlib/Algebra/Order/Floor/Prime.lean @@ -3,7 +3,7 @@ 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.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Infinite import Mathlib.Topology.Algebra.Order.Floor /-! diff --git a/Mathlib/Data/Int/NatPrime.lean b/Mathlib/Data/Int/NatPrime.lean index f1cbf001f3416..ee1e51f543988 100644 --- a/Mathlib/Data/Int/NatPrime.lean +++ b/Mathlib/Data/Int/NatPrime.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Bryan Gin-ge Chen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Lacker, Bryan Gin-ge Chen -/ +import Mathlib.Algebra.Group.Int import Mathlib.Data.Nat.Prime.Basic /-! diff --git a/Mathlib/Data/Nat/Choose/Dvd.lean b/Mathlib/Data/Nat/Choose/Dvd.lean index e5cca9d05b77b..0272d236dd30a 100644 --- a/Mathlib/Data/Nat/Choose/Dvd.lean +++ b/Mathlib/Data/Nat/Choose/Dvd.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Patrick Stevens -/ import Mathlib.Data.Nat.Choose.Basic -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Factorial /-! # Divisibility properties of binomial coefficients diff --git a/Mathlib/Data/Nat/Factorization/PrimePow.lean b/Mathlib/Data/Nat/Factorization/PrimePow.lean index 37333d9e42ecf..f36f7a25b7efd 100644 --- a/Mathlib/Data/Nat/Factorization/PrimePow.lean +++ b/Mathlib/Data/Nat/Factorization/PrimePow.lean @@ -5,6 +5,7 @@ Authors: Bhavik Mehta -/ import Mathlib.Algebra.IsPrimePow import Mathlib.Data.Nat.Factorization.Basic +import Mathlib.Data.Nat.Prime.Pow /-! # Prime powers and factorizations diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index c3d6f9ad26ecb..ff8e8166a2c28 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -5,7 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.BigOperators.Ring.List import Mathlib.Data.Nat.GCD.Basic -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Basic import Mathlib.Data.List.Prime import Mathlib.Data.List.Sort import Mathlib.Data.List.Perm.Subperm diff --git a/Mathlib/Data/Nat/Prime/Basic.lean b/Mathlib/Data/Nat/Prime/Basic.lean index 5a4b6a4fb3db7..a0c7d3f9fc836 100644 --- a/Mathlib/Data/Nat/Prime/Basic.lean +++ b/Mathlib/Data/Nat/Prime/Basic.lean @@ -4,22 +4,17 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.Associated.Basic -import Mathlib.Algebra.Order.Monoid.Unbundled.Pow -import Mathlib.Algebra.Ring.Int -import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Algebra.Ring.Parity import Mathlib.Data.Nat.Prime.Defs -import Mathlib.Order.Bounds.Basic /-! -## Notable Theorems +# Prime numbers -- `Nat.exists_infinite_primes`: Euclid's theorem that there exist infinitely many prime numbers. - This also appears as `Nat.not_bddAbove_setOf_prime` and `Nat.infinite_setOf_prime` (the latter - in `Data.Nat.PrimeFin`). +This file develops the theory of prime numbers: natural numbers `p ≥ 2` whose only divisors are +`p` and `1`. -/ - open Bool Subtype open Nat @@ -27,6 +22,29 @@ open Nat namespace Nat variable {n : ℕ} +theorem prime_mul_iff {a b : ℕ} : Nat.Prime (a * b) ↔ a.Prime ∧ b = 1 ∨ b.Prime ∧ a = 1 := by + simp only [irreducible_mul_iff, ← irreducible_iff_nat_prime, Nat.isUnit_iff] + +theorem not_prime_mul {a b : ℕ} (a1 : a ≠ 1) (b1 : b ≠ 1) : ¬Prime (a * b) := by + simp [prime_mul_iff, _root_.not_or, *] + +theorem not_prime_mul' {a b n : ℕ} (h : a * b = n) (h₁ : a ≠ 1) (h₂ : b ≠ 1) : ¬Prime n := + h ▸ not_prime_mul h₁ h₂ + +theorem Prime.dvd_iff_eq {p a : ℕ} (hp : p.Prime) (a1 : a ≠ 1) : a ∣ p ↔ p = a := by + refine ⟨?_, by rintro rfl; rfl⟩ + rintro ⟨j, rfl⟩ + rcases prime_mul_iff.mp hp with (⟨_, rfl⟩ | ⟨_, rfl⟩) + · exact mul_one _ + · exact (a1 rfl).elim + +theorem Prime.eq_two_or_odd {p : ℕ} (hp : Prime p) : p = 2 ∨ p % 2 = 1 := + p.mod_two_eq_zero_or_one.imp_left fun h => + ((hp.eq_one_or_self_of_dvd 2 (dvd_of_mod_eq_zero h)).resolve_left (by decide)).symm + +theorem Prime.eq_two_or_odd' {p : ℕ} (hp : Prime p) : p = 2 ∨ Odd p := + Or.imp_right (fun h => ⟨p / 2, (div_add_mod p 2).symm.trans (congr_arg _ h)⟩) hp.eq_two_or_odd + section theorem Prime.five_le_of_ne_two_of_ne_three {p : ℕ} (hp : p.Prime) (h_two : p ≠ 2) @@ -78,26 +96,6 @@ theorem dvd_of_forall_prime_mul_dvd {a b : ℕ} obtain ⟨p, hp⟩ := exists_prime_and_dvd ha exact _root_.trans (dvd_mul_left a p) (hdvd p hp.1 hp.2) -/-- Euclid's theorem on the **infinitude of primes**. -Here given in the form: for every `n`, there exists a prime number `p ≥ n`. -/ -theorem exists_infinite_primes (n : ℕ) : ∃ p, n ≤ p ∧ Prime p := - let p := minFac (n ! + 1) - have f1 : n ! + 1 ≠ 1 := ne_of_gt <| succ_lt_succ <| factorial_pos _ - have pp : Prime p := minFac_prime f1 - have np : n ≤ p := - le_of_not_ge fun h => - have h₁ : p ∣ n ! := dvd_factorial (minFac_pos _) h - have h₂ : p ∣ 1 := (Nat.dvd_add_iff_right h₁).2 (minFac_dvd _) - pp.not_dvd_one h₂ - ⟨p, np, pp⟩ - -/-- A version of `Nat.exists_infinite_primes` using the `BddAbove` predicate. -/ -theorem not_bddAbove_setOf_prime : ¬BddAbove { p | Prime p } := by - rw [not_bddAbove_iff] - intro n - obtain ⟨p, hi, hp⟩ := exists_infinite_primes n.succ - exact ⟨p, hp, hi⟩ - theorem Prime.even_iff {p : ℕ} (hp : Prime p) : Even p ↔ p = 2 := by rw [even_iff_two_dvd, prime_dvd_prime_iff_eq prime_two hp, eq_comm] @@ -159,18 +157,6 @@ theorem Prime.pow_eq_iff {p a k : ℕ} (hp : p.Prime) : a ^ k = p ↔ a = p ∧ rw [← h] at hp rw [← h, hp.eq_one_of_pow, eq_self_iff_true, _root_.and_true, pow_one] -theorem pow_minFac {n k : ℕ} (hk : k ≠ 0) : (n ^ k).minFac = n.minFac := by - rcases eq_or_ne n 1 with (rfl | hn) - · simp - have hnk : n ^ k ≠ 1 := fun hk' => hn ((pow_eq_one_iff hk).1 hk') - apply (minFac_le_of_dvd (minFac_prime hn).two_le ((minFac_dvd n).pow hk)).antisymm - apply - minFac_le_of_dvd (minFac_prime hnk).two_le - ((minFac_prime hnk).dvd_of_dvd_pow (minFac_dvd _)) - -theorem Prime.pow_minFac {p k : ℕ} (hp : p.Prime) (hk : k ≠ 0) : (p ^ k).minFac = p := by - rw [Nat.pow_minFac hk, hp.minFac_eq] - theorem Prime.mul_eq_prime_sq_iff {x y p : ℕ} (hp : p.Prime) (hx : x ≠ 1) (hy : y ≠ 1) : x * y = p ^ 2 ↔ x = p ∧ y = p := by refine ⟨fun h => ?_, fun ⟨h₁, h₂⟩ => h₁.symm ▸ h₂.symm ▸ (sq _).symm⟩ @@ -197,14 +183,6 @@ theorem Prime.mul_eq_prime_sq_iff {x y p : ℕ} (hp : p.Prime) (hx : x ≠ 1) (h rw [sq, Nat.mul_right_eq_self_iff (Nat.mul_pos hp.pos hp.pos : 0 < a * a)] at h exact h -theorem Prime.dvd_factorial : ∀ {n p : ℕ} (_ : Prime p), p ∣ n ! ↔ p ≤ n - | 0, _, hp => iff_of_false hp.not_dvd_one (not_le_of_lt hp.pos) - | n + 1, p, hp => by - rw [factorial_succ, hp.dvd_mul, Prime.dvd_factorial hp] - exact - ⟨fun h => h.elim (le_of_dvd (succ_pos _)) le_succ_of_le, fun h => - (_root_.lt_or_eq_of_le h).elim (Or.inr ∘ le_of_lt_succ) fun h => Or.inl <| by rw [h]⟩ - theorem Prime.coprime_pow_of_not_dvd {p m a : ℕ} (pp : Prime p) (h : ¬p ∣ a) : Coprime a (p ^ m) := (pp.coprime_iff_not_dvd.2 h).symm.pow_right _ @@ -269,33 +247,4 @@ theorem succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul {p : ℕ} (p_prime : Prime p) { exact hpd5.elim (fun h : p ∣ m / p ^ k => Or.inl <| mul_dvd_of_dvd_div hpm h) fun h : p ∣ n / p ^ l => Or.inr <| mul_dvd_of_dvd_div hpn h -theorem prime_iff_prime_int {p : ℕ} : p.Prime ↔ _root_.Prime (p : ℤ) := - ⟨fun hp => - ⟨Int.natCast_ne_zero_iff_pos.2 hp.pos, mt Int.isUnit_iff_natAbs_eq.1 hp.ne_one, fun a b h => by - rw [← Int.dvd_natAbs, Int.natCast_dvd_natCast, Int.natAbs_mul, hp.dvd_mul] at h - rwa [← Int.dvd_natAbs, Int.natCast_dvd_natCast, ← Int.dvd_natAbs, Int.natCast_dvd_natCast]⟩, - fun hp => - Nat.prime_iff.2 - ⟨Int.natCast_ne_zero.1 hp.1, - (mt Nat.isUnit_iff.1) fun h => by simp [h, not_prime_one] at hp, fun a b => by - simpa only [Int.natCast_dvd_natCast, (Int.ofNat_mul _ _).symm] using hp.2.2 a b⟩⟩ - -/-- Two prime powers with positive exponents are equal only when the primes and the -exponents are equal. -/ -lemma Prime.pow_inj {p q m n : ℕ} (hp : p.Prime) (hq : q.Prime) - (h : p ^ (m + 1) = q ^ (n + 1)) : p = q ∧ m = n := by - have H := dvd_antisymm (Prime.dvd_of_dvd_pow hp <| h ▸ dvd_pow_self p (succ_ne_zero m)) - (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)⟩ - end Nat - -namespace Int - -theorem prime_two : Prime (2 : ℤ) := - Nat.prime_iff_prime_int.mp Nat.prime_two - -theorem prime_three : Prime (3 : ℤ) := - Nat.prime_iff_prime_int.mp Nat.prime_three - -end Int diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 1741083417a66..e6a3e5bd1a467 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd -import Mathlib.Algebra.Prime.Lemmas -import Mathlib.Algebra.Ring.Parity +import Mathlib.Algebra.Prime.Defs +import Mathlib.Algebra.Ring.Nat +import Mathlib.Order.Lattice /-! # Prime numbers @@ -161,29 +162,6 @@ theorem prime_dvd_prime_iff_eq {p q : ℕ} (pp : p.Prime) (qp : q.Prime) : p ∣ theorem Prime.not_dvd_one {p : ℕ} (pp : Prime p) : ¬p ∣ 1 := Irreducible.not_dvd_one pp -theorem prime_mul_iff {a b : ℕ} : Nat.Prime (a * b) ↔ a.Prime ∧ b = 1 ∨ b.Prime ∧ a = 1 := by - simp only [irreducible_mul_iff, ← irreducible_iff_nat_prime, Nat.isUnit_iff] - -theorem not_prime_mul {a b : ℕ} (a1 : a ≠ 1) (b1 : b ≠ 1) : ¬Prime (a * b) := by - simp [prime_mul_iff, _root_.not_or, *] - -theorem not_prime_mul' {a b n : ℕ} (h : a * b = n) (h₁ : a ≠ 1) (h₂ : b ≠ 1) : ¬Prime n := - h ▸ not_prime_mul h₁ h₂ - -theorem Prime.dvd_iff_eq {p a : ℕ} (hp : p.Prime) (a1 : a ≠ 1) : a ∣ p ↔ p = a := by - refine ⟨?_, by rintro rfl; rfl⟩ - rintro ⟨j, rfl⟩ - rcases prime_mul_iff.mp hp with (⟨_, rfl⟩ | ⟨_, rfl⟩) - · exact mul_one _ - · exact (a1 rfl).elim - -theorem Prime.eq_two_or_odd {p : ℕ} (hp : Prime p) : p = 2 ∨ p % 2 = 1 := - p.mod_two_eq_zero_or_one.imp_left fun h => - ((hp.eq_one_or_self_of_dvd 2 (dvd_of_mod_eq_zero h)).resolve_left (by decide)).symm - -theorem Prime.eq_two_or_odd' {p : ℕ} (hp : Prime p) : p = 2 ∨ Odd p := - Or.imp_right (fun h => ⟨p / 2, (div_add_mod p 2).symm.trans (congr_arg _ h)⟩) hp.eq_two_or_odd - section MinFac theorem minFac_lemma (n k : ℕ) (h : ¬n < k * k) : sqrt n - k < sqrt n + 2 - k := diff --git a/Mathlib/Data/Nat/Prime/Factorial.lean b/Mathlib/Data/Nat/Prime/Factorial.lean new file mode 100644 index 0000000000000..4c2f70a84a4c9 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Factorial.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Data.Nat.Prime.Defs + +/-! +# Prime natural numbers and the factorial operator + +-/ + +open Bool Subtype + +open Nat + +namespace Nat + +theorem Prime.dvd_factorial : ∀ {n p : ℕ} (_ : Prime p), p ∣ n ! ↔ p ≤ n + | 0, _, hp => iff_of_false hp.not_dvd_one (not_le_of_lt hp.pos) + | n + 1, p, hp => by + rw [factorial_succ, hp.dvd_mul, Prime.dvd_factorial hp] + exact + ⟨fun h => h.elim (le_of_dvd (succ_pos _)) le_succ_of_le, fun h => + (_root_.lt_or_eq_of_le h).elim (Or.inr ∘ le_of_lt_succ) fun h => Or.inl <| by rw [h]⟩ + +end Nat diff --git a/Mathlib/Data/Nat/Prime/Infinite.lean b/Mathlib/Data/Nat/Prime/Infinite.lean new file mode 100644 index 0000000000000..f67b58631e769 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Infinite.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Data.Nat.Factorial.Basic +import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Order.Bounds.Basic + +/-! +## Notable Theorems + +- `Nat.exists_infinite_primes`: Euclid's theorem that there exist infinitely many prime numbers. + This also appears as `Nat.not_bddAbove_setOf_prime` and `Nat.infinite_setOf_prime` (the latter + in `Data.Nat.PrimeFin`). + +-/ + +open Bool Subtype + +open Nat + +namespace Nat + +section Infinite + +/-- Euclid's theorem on the **infinitude of primes**. +Here given in the form: for every `n`, there exists a prime number `p ≥ n`. -/ +theorem exists_infinite_primes (n : ℕ) : ∃ p, n ≤ p ∧ Prime p := + let p := minFac (n ! + 1) + have f1 : n ! + 1 ≠ 1 := ne_of_gt <| succ_lt_succ <| factorial_pos _ + have pp : Prime p := minFac_prime f1 + have np : n ≤ p := + le_of_not_ge fun h => + have h₁ : p ∣ n ! := dvd_factorial (minFac_pos _) h + have h₂ : p ∣ 1 := (Nat.dvd_add_iff_right h₁).2 (minFac_dvd _) + pp.not_dvd_one h₂ + ⟨p, np, pp⟩ + +/-- A version of `Nat.exists_infinite_primes` using the `BddAbove` predicate. -/ +theorem not_bddAbove_setOf_prime : ¬BddAbove { p | Prime p } := by + rw [not_bddAbove_iff] + intro n + obtain ⟨p, hi, hp⟩ := exists_infinite_primes n.succ + exact ⟨p, hp, hi⟩ + +end Infinite + +end Nat diff --git a/Mathlib/Data/Nat/Prime/Int.lean b/Mathlib/Data/Nat/Prime/Int.lean new file mode 100644 index 0000000000000..c5f0347d3b6c8 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Int.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Ring.Int +import Mathlib.Data.Nat.Prime.Basic + +/-! +# Prime numbers in the naturals and the integers + +TODO: This file can probably be merged with `Data/Nat/Int/NatPrime.lean`. +-/ + + +namespace Nat + +theorem prime_iff_prime_int {p : ℕ} : p.Prime ↔ _root_.Prime (p : ℤ) := + ⟨fun hp => + ⟨Int.natCast_ne_zero_iff_pos.2 hp.pos, mt Int.isUnit_iff_natAbs_eq.1 hp.ne_one, fun a b h => by + rw [← Int.dvd_natAbs, Int.natCast_dvd_natCast, Int.natAbs_mul, hp.dvd_mul] at h + rwa [← Int.dvd_natAbs, Int.natCast_dvd_natCast, ← Int.dvd_natAbs, Int.natCast_dvd_natCast]⟩, + fun hp => + Nat.prime_iff.2 + ⟨Int.natCast_ne_zero.1 hp.1, + (mt Nat.isUnit_iff.1) fun h => by simp [h, not_prime_one] at hp, fun a b => by + simpa only [Int.natCast_dvd_natCast, (Int.ofNat_mul _ _).symm] using hp.2.2 a b⟩⟩ + +/-- Two prime powers with positive exponents are equal only when the primes and the +exponents are equal. -/ +lemma Prime.pow_inj {p q m n : ℕ} (hp : p.Prime) (hq : q.Prime) + (h : p ^ (m + 1) = q ^ (n + 1)) : p = q ∧ m = n := by + have H := dvd_antisymm (Prime.dvd_of_dvd_pow hp <| h ▸ dvd_pow_self p (succ_ne_zero m)) + (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)⟩ + +end Nat + +namespace Int + +theorem prime_two : Prime (2 : ℤ) := + Nat.prime_iff_prime_int.mp Nat.prime_two + +theorem prime_three : Prime (3 : ℤ) := + Nat.prime_iff_prime_int.mp Nat.prime_three + +end Int diff --git a/Mathlib/Data/Nat/Prime/Pow.lean b/Mathlib/Data/Nat/Prime/Pow.lean new file mode 100644 index 0000000000000..a103a434d4ff6 --- /dev/null +++ b/Mathlib/Data/Nat/Prime/Pow.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Data.Nat.Prime.Basic + +/-! +# Prime numbers + +This file develops the theory of prime numbers: natural numbers `p ≥ 2` whose only divisors are +`p` and `1`. + +-/ + +namespace Nat + +theorem pow_minFac {n k : ℕ} (hk : k ≠ 0) : (n ^ k).minFac = n.minFac := by + rcases eq_or_ne n 1 with (rfl | hn) + · simp + have hnk : n ^ k ≠ 1 := fun hk' => hn ((pow_eq_one_iff hk).1 hk') + apply (minFac_le_of_dvd (minFac_prime hn).two_le ((minFac_dvd n).pow hk)).antisymm + apply + minFac_le_of_dvd (minFac_prime hnk).two_le + ((minFac_prime hnk).dvd_of_dvd_pow (minFac_dvd _)) + +theorem Prime.pow_minFac {p k : ℕ} (hp : p.Prime) (hk : k ≠ 0) : (p ^ k).minFac = p := by + rw [Nat.pow_minFac hk, hp.minFac_eq] + +end Nat diff --git a/Mathlib/Data/Nat/PrimeFin.lean b/Mathlib/Data/Nat/PrimeFin.lean index 8f1fe7411bc75..ecb04ce9d0c0a 100644 --- a/Mathlib/Data/Nat/PrimeFin.lean +++ b/Mathlib/Data/Nat/PrimeFin.lean @@ -6,7 +6,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro import Mathlib.Data.Countable.Defs import Mathlib.Data.Nat.Factors import Mathlib.Data.Set.Finite -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Infinite /-! # Prime numbers diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 948033b25c274..6924fe05ed12c 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int import Mathlib.Data.List.Cycle import Mathlib.Data.Nat.GCD.Basic import Mathlib.Data.Nat.Prime.Basic diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index 6927e41b5eeef..1faaafe97442e 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.NumberTheory.LegendreSymbol.Basic import Mathlib.Analysis.Normed.Field.Lemmas +import Mathlib.Data.Nat.Prime.Factorial +import Mathlib.NumberTheory.LegendreSymbol.Basic /-! # Lemmas of Gauss and Eisenstein diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 119163114c7bf..cfa4de5aac58e 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -6,9 +6,9 @@ Authors: Tian Chen, Mantas Bakšys import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Ring.Int +import Mathlib.Data.Nat.Prime.Int import Mathlib.NumberTheory.Padics.PadicVal.Defs import Mathlib.RingTheory.Ideal.Quotient -import Mathlib.Data.Nat.Prime.Basic /-! # Multiplicity in Number Theory diff --git a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean index 6ddb55a0ed27f..d2df11023e6f1 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.NumberTheory.Divisors import Mathlib.NumberTheory.Padics.PadicVal.Defs import Mathlib.Data.Nat.MaxPowDiv import Mathlib.Data.Nat.Multiplicity +import Mathlib.Data.Nat.Prime.Int /-! # `p`-adic Valuation diff --git a/Mathlib/NumberTheory/Primorial.lean b/Mathlib/NumberTheory/Primorial.lean index 7ef909c7b0799..3867a3aa0f893 100644 --- a/Mathlib/NumberTheory/Primorial.lean +++ b/Mathlib/NumberTheory/Primorial.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Data.Nat.Choose.Sum import Mathlib.Data.Nat.Choose.Dvd -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Basic /-! # Primorial diff --git a/Mathlib/RingTheory/Int/Basic.lean b/Mathlib/RingTheory/Int/Basic.lean index 2f194c3abb9aa..be4be78dc6d35 100644 --- a/Mathlib/RingTheory/Int/Basic.lean +++ b/Mathlib/RingTheory/Int/Basic.lean @@ -5,9 +5,9 @@ Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson -/ import Mathlib.Algebra.EuclideanDomain.Basic import Mathlib.Algebra.EuclideanDomain.Int -import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Algebra.GCDMonoid.Nat -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Int +import Mathlib.RingTheory.PrincipalIdealDomain /-! # Divisibility over ℕ and ℤ diff --git a/Mathlib/Tactic/NormNum/Prime.lean b/Mathlib/Tactic/NormNum/Prime.lean index f1e5dc961f2b4..be12f0e31793f 100644 --- a/Mathlib/Tactic/NormNum/Prime.lean +++ b/Mathlib/Tactic/NormNum/Prime.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Tactic.NormNum.Basic -import Mathlib.Data.Nat.Prime.Defs +import Mathlib.Data.Nat.Prime.Basic /-! # `norm_num` extensions on natural numbers diff --git a/test/observe.lean b/test/observe.lean index 9a2b83728aaca..b595d01117739 100644 --- a/test/observe.lean +++ b/test/observe.lean @@ -1,4 +1,4 @@ -import Mathlib.Data.Nat.Prime.Basic +import Mathlib.Data.Nat.Prime.Factorial import Mathlib.Data.Nat.Factorial.Basic open Nat From 1a5abab17a4a8835baf79e49d70edc7f0972f299 Mon Sep 17 00:00:00 2001 From: YnirPaz Date: Fri, 25 Oct 2024 20:32:11 +0000 Subject: [PATCH 29/29] feat(SetTheory/Ordinal/Topology): define Ordinal.IsClosed and Ordinal.IsAcc (#16710) Defined Ordinal.IsClosed and Ordinal.IsAcc and proved basic lemmas about them. This is closedness in the club-sense. --- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 3 + Mathlib/SetTheory/Ordinal/Topology.lean | 136 ++++++++++++++++++++++ Mathlib/Topology/Basic.lean | 11 ++ 3 files changed, 150 insertions(+) diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 6e9b1070b8f6c..865bd245f7301 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -262,6 +262,9 @@ theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ if h : ∃ a, o = succ a then Or.inr (Or.inl h) else Or.inr <| Or.inr ⟨o0, fun _a => (succ_lt_of_not_succ h).2⟩ +theorem isLimit_of_not_succ_of_ne_zero {o : Ordinal} (h : ¬∃ a, o = succ a) (h' : o ≠ 0) : + IsLimit o := ((zero_or_succ_or_limit o).resolve_left h').resolve_left h + -- TODO: this is an iff with `IsSuccPrelimit` theorem IsLimit.sSup_Iio {o : Ordinal} (h : IsLimit o) : sSup (Iio o) = o := by apply (csSup_le' (fun a ha ↦ le_of_lt ha)).antisymm diff --git a/Mathlib/SetTheory/Ordinal/Topology.lean b/Mathlib/SetTheory/Ordinal/Topology.lean index 6993d517742e4..14986ed23a289 100644 --- a/Mathlib/SetTheory/Ordinal/Topology.lean +++ b/Mathlib/SetTheory/Ordinal/Topology.lean @@ -246,4 +246,140 @@ theorem enumOrd_isNormal_iff_isClosed (hs : ¬ BddAbove s) : rw [hb] exact le_bsup.{u, u} _ _ (ha.2 _ hba) +open Set Filter Set.Notation + +/-- An ordinal is an accumulation point of a set of ordinals if it is positive and there +are elements in the set arbitrarily close to the ordinal from below. -/ +def IsAcc (o : Ordinal) (S : Set Ordinal) : Prop := + AccPt o (𝓟 S) + +/-- A set of ordinals is closed below an ordinal if it contains all of +its accumulation points below the ordinal. -/ +def IsClosedBelow (S : Set Ordinal) (o : Ordinal) : Prop := + IsClosed (Iio o ↓∩ S) + +theorem isAcc_iff (o : Ordinal) (S : Set Ordinal) : o.IsAcc S ↔ + o ≠ 0 ∧ ∀ p < o, (S ∩ Ioo p o).Nonempty := by + dsimp [IsAcc] + constructor + · rw [accPt_iff_nhds] + intro h + constructor + · rintro rfl + obtain ⟨x, hx⟩ := h (Iio 1) (Iio_mem_nhds zero_lt_one) + exact hx.2 <| lt_one_iff_zero.mp hx.1.1 + · intro p plt + obtain ⟨x, hx⟩ := h (Ioo p (o + 1)) <| Ioo_mem_nhds plt (lt_succ o) + use x + refine ⟨hx.1.2, ⟨hx.1.1.1, lt_of_le_of_ne ?_ hx.2⟩⟩ + have := hx.1.1.2 + rwa [← succ_eq_add_one, lt_succ_iff] at this + · rw [accPt_iff_nhds] + intro h u umem + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds umem ⟨0, Ordinal.pos_iff_ne_zero.mpr h.1⟩ + obtain ⟨x, hx⟩ := h.2 l hl.1 + use x + exact ⟨⟨hl.2 ⟨hx.2.1, hx.2.2.le⟩, hx.1⟩, hx.2.2.ne⟩ + +theorem IsAcc.forall_lt {o : Ordinal} {S : Set Ordinal} (h : o.IsAcc S) : + ∀ p < o, (S ∩ Ioo p o).Nonempty := ((isAcc_iff _ _).mp h).2 + +theorem IsAcc.pos {o : Ordinal} {S : Set Ordinal} (h : o.IsAcc S) : + 0 < o := Ordinal.pos_iff_ne_zero.mpr ((isAcc_iff _ _).mp h).1 + +theorem IsAcc.isLimit {o : Ordinal} {S : Set Ordinal} (h : o.IsAcc S) : IsLimit o := by + rw [isAcc_iff] at h + refine isLimit_of_not_succ_of_ne_zero (fun ⟨x, hx⟩ ↦ ?_) h.1 + rcases h.2 x (lt_of_lt_of_le (lt_succ x) hx.symm.le) with ⟨p, hp⟩ + exact (hx.symm ▸ (succ_le_iff.mpr hp.2.1)).not_lt hp.2.2 + +theorem IsAcc.mono {o : Ordinal} {S T : Set Ordinal} (h : S ⊆ T) (ho : o.IsAcc S) : + o.IsAcc T := by + rw [isAcc_iff] at * + exact ⟨ho.1, fun p plto ↦ (ho.2 p plto).casesOn fun s hs ↦ ⟨s, h hs.1, hs.2⟩⟩ + +theorem IsAcc.inter_Ioo_nonempty {o : Ordinal} {S : Set Ordinal} (hS : o.IsAcc S) + {p : Ordinal} (hp : p < o) : (S ∩ Ioo p o).Nonempty := hS.forall_lt p hp + +-- todo: prove this for a general linear `SuccOrder`. +theorem accPt_subtype {p o : Ordinal} (S : Set Ordinal) (hpo : p < o) : + AccPt p (𝓟 S) ↔ AccPt ⟨p, hpo⟩ (𝓟 (Iio o ↓∩ S)) := by + constructor + · intro h + have plim : p.IsLimit := IsAcc.isLimit h + rw [accPt_iff_nhds] at * + intro u hu + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds hu ⟨⟨0, plim.pos.trans hpo⟩, plim.pos⟩ + obtain ⟨x, hx⟩ := h (Ioo l (p + 1)) (Ioo_mem_nhds hl.1 (lt_add_one _)) + use ⟨x, lt_of_le_of_lt (lt_succ_iff.mp hx.1.1.2) hpo⟩ + refine ⟨?_, Subtype.coe_ne_coe.mp hx.2⟩ + exact ⟨hl.2 ⟨hx.1.1.1, by exact_mod_cast lt_succ_iff.mp hx.1.1.2⟩, hx.1.2⟩ + · intro h + rw [accPt_iff_nhds] at * + intro u hu + by_cases ho : p + 1 < o + · have ppos : p ≠ 0 := by + rintro rfl + rw [zero_add] at ho + specialize h (Iio ⟨1, ho⟩) (Iio_mem_nhds (Subtype.mk_lt_mk.mpr zero_lt_one)) + obtain ⟨_, h⟩ := h + exact h.2 <| Subtype.mk_eq_mk.mpr (lt_one_iff_zero.mp h.1.1) + have plim : p.IsLimit := by + contrapose! h + obtain ⟨q, hq⟩ := ((zero_or_succ_or_limit p).resolve_left ppos).resolve_right h + use (Ioo ⟨q, ((hq ▸ lt_succ q).trans hpo)⟩ ⟨p + 1, ho⟩) + constructor + · exact Ioo_mem_nhds (by simp only [hq, Subtype.mk_lt_mk, lt_succ]) (lt_succ p) + · intro _ mem + have aux1 := Subtype.mk_lt_mk.mp mem.1.1 + have aux2 := Subtype.mk_lt_mk.mp mem.1.2 + rw [Subtype.mk_eq_mk] + rw [hq] at aux2 ⊢ + exact ((succ_le_iff.mpr aux1).antisymm (le_of_lt_succ aux2)).symm + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds hu ⟨0, plim.pos⟩ + obtain ⟨x, hx⟩ := h (Ioo ⟨l, hl.1.trans hpo⟩ ⟨p + 1, ho⟩) (Ioo_mem_nhds hl.1 (lt_add_one p)) + use x + exact ⟨⟨hl.2 ⟨hx.1.1.1, lt_succ_iff.mp hx.1.1.2⟩, hx.1.2⟩, fun h ↦ hx.2 (SetCoe.ext h)⟩ + have hp : o = p + 1 := (le_succ_iff_eq_or_le.mp (le_of_not_lt ho)).resolve_right + (not_le_of_lt hpo) + have ppos : p ≠ 0 := by + rintro rfl + obtain ⟨x, hx⟩ := h Set.univ univ_mem + have : ↑x < o := x.2 + simp_rw [hp, zero_add, lt_one_iff_zero] at this + exact hx.2 (SetCoe.ext this) + obtain ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds hu ⟨0, Ordinal.pos_iff_ne_zero.mpr ppos⟩ + obtain ⟨x, hx⟩ := h (Ioi ⟨l, hl.1.trans hpo⟩) (Ioi_mem_nhds hl.1) + use x + refine ⟨⟨hl.2 ⟨hx.1.1, ?_⟩, hx.1.2⟩, fun h ↦ hx.2 (SetCoe.ext h)⟩ + rw [← lt_add_one_iff, ← hp] + exact x.2 + +theorem isClosedBelow_iff {S : Set Ordinal} {o : Ordinal} : IsClosedBelow S o ↔ + ∀ p < o, IsAcc p S → p ∈ S := by + dsimp [IsClosedBelow] + constructor + · intro h p plto hp + have : AccPt ⟨p, plto⟩ (𝓟 (Iio o ↓∩ S)) := (accPt_subtype _ _).mp hp + rw [isClosed_iff_clusterPt] at h + exact h ⟨p, plto⟩ this.clusterPt + · intro h + rw [isClosed_iff_clusterPt] + intro r hr + match clusterPt_principal.mp hr with + | .inl h => exact h + | .inr h' => exact h r.1 r.2 <| (accPt_subtype _ _).mpr h' + +alias ⟨IsClosedBelow.forall_lt, _⟩ := isClosedBelow_iff + +theorem IsClosedBelow.sInter {o : Ordinal} {S : Set (Set Ordinal)} + (h : ∀ C ∈ S, IsClosedBelow C o) : IsClosedBelow (⋂₀ S) o := by + rw [isClosedBelow_iff] + intro p plto pAcc C CmemS + exact (h C CmemS).forall_lt p plto (pAcc.mono (sInter_subset_of_mem CmemS)) + +theorem IsClosedBelow.iInter {ι : Type u} {f : ι → Set Ordinal} {o : Ordinal} + (h : ∀ i, IsClosedBelow (f i) o) : IsClosedBelow (⋂ i, f i) o := + IsClosedBelow.sInter fun _ ⟨i, hi⟩ ↦ hi ▸ (h i) + end Ordinal diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index b9f811270fa26..9961f73158dbd 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1031,6 +1031,17 @@ theorem AccPt.mono {F G : Filter X} (h : AccPt x F) (hFG : F ≤ G) : AccPt x G theorem AccPt.clusterPt (x : X) (F : Filter X) (h : AccPt x F) : ClusterPt x F := ((acc_iff_cluster x F).mp h).mono inf_le_right +theorem clusterPt_principal {x : X} {C : Set X} : + ClusterPt x (𝓟 C) ↔ x ∈ C ∨ AccPt x (𝓟 C) := by + constructor + · intro h + by_contra! hc + rw [acc_principal_iff_cluster] at hc + simp_all only [not_false_eq_true, diff_singleton_eq_self, not_true_eq_false, hc.1] + · rintro (h | h) + · exact clusterPt_principal_iff.mpr fun _ mem ↦ ⟨x, ⟨mem_of_mem_nhds mem, h⟩⟩ + · exact h.clusterPt + /-! ### Interior, closure and frontier in terms of neighborhoods -/