From 9ea71b285a927b15158c6305a645796fdc49c433 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 12:46:03 +1100 Subject: [PATCH 01/90] chore: adaptations for nightly-2024-11-18 From 47f0d27ac082d7d45d4efcca5442d6adcbcc0e2b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 02:51:22 +0000 Subject: [PATCH 02/90] chore(SetTheory/Cardinal/Aleph): golf `aleph` theorems (#18204) --- Mathlib/SetTheory/Cardinal/Aleph.lean | 47 ++++++++++++++------------- 1 file changed, 24 insertions(+), 23 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index f95620122c4e1..3970c6fc6cb57 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -207,6 +207,9 @@ theorem omega_le_omega {o₁ o₂ : Ordinal} : ω_ o₁ ≤ ω_ o₂ ↔ o₁ theorem omega_max (o₁ o₂ : Ordinal) : ω_ (max o₁ o₂) = max (ω_ o₁) (ω_ o₂) := omega.monotone.map_max +theorem preOmega_le_omega (o : Ordinal) : preOmega o ≤ ω_ o := + preOmega_le_preOmega.2 (Ordinal.le_add_left _ _) + theorem isInitial_omega (o : Ordinal) : IsInitial (omega o) := isInitial_preOmega _ @@ -298,13 +301,21 @@ theorem preAleph_succ (o : Ordinal) : preAleph (succ o) = succ (preAleph o) := preAleph.map_succ o @[simp] -theorem preAleph_nat : ∀ n : ℕ, preAleph n = n - | 0 => preAleph_zero - | n + 1 => show preAleph (succ n) = n.succ by rw [preAleph_succ, preAleph_nat n, nat_succ] +theorem preAleph_nat (n : ℕ) : preAleph n = n := by + rw [← card_preOmega, preOmega_natCast, card_nat] + +@[simp] +theorem preAleph_omega0 : preAleph ω = ℵ₀ := by + rw [← card_preOmega, preOmega_omega0, card_omega0] +@[simp] theorem preAleph_pos {o : Ordinal} : 0 < preAleph o ↔ 0 < o := by rw [← preAleph_zero, preAleph_lt_preAleph] +@[simp] +theorem aleph0_le_preAleph {o : Ordinal} : ℵ₀ ≤ preAleph o ↔ ω ≤ o := by + rw [← preAleph_omega0, preAleph_le_preAleph] + @[simp] theorem lift_preAleph (o : Ordinal.{u}) : lift.{v} (preAleph o) = preAleph (Ordinal.lift.{v} o) := (preAleph.toInitialSeg.trans liftInitialSeg).eq @@ -328,16 +339,6 @@ theorem preAleph_limit {o : Ordinal} (ho : o.IsLimit) : preAleph o = ⨆ a : Iio rw [preAleph_le_of_isLimit ho] exact fun a ha => le_ciSup (bddAbove_of_small _) (⟨a, ha⟩ : Iio o) -@[simp] -theorem preAleph_omega0 : preAleph ω = ℵ₀ := - eq_of_forall_ge_iff fun c => by - simp only [preAleph_le_of_isLimit isLimit_omega0, lt_omega0, exists_imp, aleph0_le] - exact forall_swap.trans (forall_congr' fun n => by simp only [forall_eq, preAleph_nat]) - -@[simp] -theorem aleph0_le_preAleph {o : Ordinal} : ℵ₀ ≤ preAleph o ↔ ω ≤ o := by - rw [← preAleph_omega0, preAleph_le_preAleph] - /-- The `aleph` function gives the infinite cardinals listed by their ordinal index. `aleph 0 = ℵ₀`, `aleph 1 = succ ℵ₀` is the first uncountable cardinal, and so on. @@ -381,6 +382,9 @@ theorem aleph_max (o₁ o₂ : Ordinal) : ℵ_ (max o₁ o₂) = max (ℵ_ o₁) theorem max_aleph_eq (o₁ o₂ : Ordinal) : max (ℵ_ o₁) (ℵ_ o₂) = ℵ_ (max o₁ o₂) := (aleph_max o₁ o₂).symm +theorem preAleph_le_aleph (o : Ordinal) : preAleph o ≤ ℵ_ o := + preAleph_le_preAleph.2 (Ordinal.le_add_left _ _) + @[simp] theorem aleph_succ (o : Ordinal) : ℵ_ (succ o) = succ (ℵ_ o) := by rw [aleph_eq_preAleph, add_succ, preAleph_succ, aleph_eq_preAleph] @@ -399,16 +403,13 @@ theorem _root_.Ordinal.lift_omega (o : Ordinal.{u}) : simp [omega_eq_preOmega] theorem aleph_limit {o : Ordinal} (ho : o.IsLimit) : ℵ_ o = ⨆ a : Iio o, ℵ_ a := by - apply le_antisymm _ (ciSup_le' _) - · rw [aleph_eq_preAleph, preAleph_limit (ho.add _)] - refine ciSup_mono' (bddAbove_of_small _) ?_ - rintro ⟨i, hi⟩ - cases' lt_or_le i ω with h h - · rcases lt_omega0.1 h with ⟨n, rfl⟩ - use ⟨0, ho.pos⟩ - simpa using (nat_lt_aleph0 n).le - · exact ⟨⟨_, (sub_lt_of_le h).2 hi⟩, preAleph_le_preAleph.2 (le_add_sub _ _)⟩ - · exact fun i => aleph_le_aleph.2 i.2.le + rw [aleph_eq_preAleph, preAleph_limit (isLimit_add ω ho)] + apply le_antisymm <;> + apply ciSup_mono' (bddAbove_of_small _) <;> + intro i + · refine ⟨⟨_, sub_lt_of_lt_add i.2 ho.pos⟩, ?_⟩ + simpa [aleph_eq_preAleph] using le_add_sub _ _ + · exact ⟨⟨_, add_lt_add_left i.2 ω⟩, le_rfl⟩ theorem aleph0_le_aleph (o : Ordinal) : ℵ₀ ≤ ℵ_ o := by rw [aleph_eq_preAleph, aleph0_le_preAleph] From b2244f6fc5832d2e74f3990726e448489cb773bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 04:07:31 +0000 Subject: [PATCH 03/90] =?UTF-8?q?chore(SetTheory/Ordinal/Basic):=20`Ordina?= =?UTF-8?q?l.NeZero.One`=20=E2=86=92=20`Ordinal.instNeZeroOne`=20(#19064)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/SetTheory/Nimber/Basic.lean | 14 ++++++------ Mathlib/SetTheory/Ordinal/Basic.lean | 14 ++++++------ Mathlib/SetTheory/Ordinal/NaturalOps.lean | 26 ++++++++++++----------- 3 files changed, 28 insertions(+), 26 deletions(-) diff --git a/Mathlib/SetTheory/Nimber/Basic.lean b/Mathlib/SetTheory/Nimber/Basic.lean index 28ff884f9c685..9f33020c6c6b5 100644 --- a/Mathlib/SetTheory/Nimber/Basic.lean +++ b/Mathlib/SetTheory/Nimber/Basic.lean @@ -44,16 +44,16 @@ noncomputable section /-! ### Basic casts between `Ordinal` and `Nimber` -/ -/-- A type synonym for ordinals with natural addition and multiplication. -/ +/-- A type synonym for ordinals with nimber addition and multiplication. -/ def Nimber : Type _ := Ordinal deriving Zero, Inhabited, One, Nontrivial, WellFoundedRelation -instance Nimber.linearOrder : LinearOrder Nimber := {Ordinal.linearOrder with} -instance Nimber.succOrder : SuccOrder Nimber := {Ordinal.instSuccOrder with} -instance Nimber.orderBot : OrderBot Nimber := {Ordinal.orderBot with} -instance Nimber.noMaxOrder : NoMaxOrder Nimber := {Ordinal.noMaxOrder with} -instance Nimber.zeroLEOneClass : ZeroLEOneClass Nimber := {Ordinal.zeroLEOneClass with} -instance Nimber.instNeZeroOne : NeZero (1 : Nimber) := Ordinal.NeZero.one +instance Nimber.instLinearOrder : LinearOrder Nimber := Ordinal.instLinearOrder +instance Nimber.instSuccOrder : SuccOrder Nimber := Ordinal.instSuccOrder +instance Nimber.instOrderBot : OrderBot Nimber := Ordinal.instOrderBot +instance Nimber.instNoMaxOrder : NoMaxOrder Nimber := Ordinal.instNoMaxOrder +instance Nimber.instZeroLEOneClass : ZeroLEOneClass Nimber := Ordinal.instZeroLEOneClass +instance Nimber.instNeZeroOne : NeZero (1 : Nimber) := Ordinal.instNeZeroOne /-- The identity function between `Ordinal` and `Nimber`. -/ @[match_pattern] diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 65e5733c47823..5a31bfe919ee9 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -287,7 +287,7 @@ instance partialOrder : PartialOrder Ordinal where Quotient.inductionOn₂ a b fun _ _ ⟨h₁⟩ ⟨h₂⟩ => Quot.sound ⟨InitialSeg.antisymm h₁ h₂⟩ -instance linearOrder : LinearOrder Ordinal := +instance : LinearOrder Ordinal := {inferInstanceAs (PartialOrder Ordinal) with le_total := fun a b => Quotient.inductionOn₂ a b fun ⟨_, r, _⟩ ⟨_, s, _⟩ => (InitialSeg.total r s).recOn (fun f => Or.inl ⟨f⟩) fun f => Or.inr ⟨f⟩ @@ -309,7 +309,7 @@ theorem _root_.PrincipalSeg.ordinal_type_lt {α β} {r : α → α → Prop} {s protected theorem zero_le (o : Ordinal) : 0 ≤ o := inductionOn o fun _ r _ => (InitialSeg.ofIsEmpty _ r).ordinal_type_le -instance orderBot : OrderBot Ordinal where +instance : OrderBot Ordinal where bot := 0 bot_le := Ordinal.zero_le @@ -330,10 +330,10 @@ protected theorem not_lt_zero (o : Ordinal) : ¬o < 0 := theorem eq_zero_or_pos : ∀ a : Ordinal, a = 0 ∨ 0 < a := eq_bot_or_bot_lt -instance zeroLEOneClass : ZeroLEOneClass Ordinal := +instance : ZeroLEOneClass Ordinal := ⟨Ordinal.zero_le _⟩ -instance NeZero.one : NeZero (1 : Ordinal) := +instance instNeZeroOne : NeZero (1 : Ordinal) := ⟨Ordinal.one_ne_zero⟩ theorem type_le_iff {α β} {r : α → α → Prop} {s : β → β → Prop} [IsWellOrder α r] @@ -875,13 +875,13 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b := by · apply (RelEmbedding.ofMonotone (Sum.recOn · f fun _ ↦ f.top) ?_).ordinal_type_le simpa [f.map_rel_iff] using f.lt_top -instance noMaxOrder : NoMaxOrder Ordinal := +instance : NoMaxOrder Ordinal := ⟨fun _ => ⟨_, succ_le_iff'.1 le_rfl⟩⟩ -instance instSuccOrder : SuccOrder Ordinal.{u} := +instance : SuccOrder Ordinal.{u} := SuccOrder.ofSuccLeIff (fun o => o + 1) succ_le_iff' -instance instSuccAddOrder : SuccAddOrder Ordinal := ⟨fun _ => rfl⟩ +instance : SuccAddOrder Ordinal := ⟨fun _ => rfl⟩ @[simp] theorem add_one_eq_succ (o : Ordinal) : o + 1 = succ o := diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index c7ac9850362c8..7463af175343e 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -53,10 +53,12 @@ def NatOrdinal : Type _ := -- Porting note: used to derive LinearOrder & SuccOrder but need to manually define Ordinal deriving Zero, Inhabited, One, WellFoundedRelation -instance NatOrdinal.linearOrder : LinearOrder NatOrdinal := {Ordinal.linearOrder with} -instance NatOrdinal.instSuccOrder : SuccOrder NatOrdinal := {Ordinal.instSuccOrder with} -instance NatOrdinal.orderBot : OrderBot NatOrdinal := {Ordinal.orderBot with} -instance NatOrdinal.noMaxOrder : NoMaxOrder NatOrdinal := {Ordinal.noMaxOrder with} +instance NatOrdinal.instLinearOrder : LinearOrder NatOrdinal := Ordinal.instLinearOrder +instance NatOrdinal.instSuccOrder : SuccOrder NatOrdinal := Ordinal.instSuccOrder +instance NatOrdinal.instOrderBot : OrderBot NatOrdinal := Ordinal.instOrderBot +instance NatOrdinal.instNoMaxOrder : NoMaxOrder NatOrdinal := Ordinal.instNoMaxOrder +instance NatOrdinal.instZeroLEOneClass : ZeroLEOneClass NatOrdinal := Ordinal.instZeroLEOneClass +instance NatOrdinal.instNeZeroOne : NeZero (1 : NatOrdinal) := Ordinal.instNeZeroOne /-- The identity function between `Ordinal` and `NatOrdinal`. -/ @[match_pattern] @@ -316,19 +318,19 @@ open Ordinal NaturalOps instance : Add NatOrdinal := ⟨nadd⟩ instance : SuccAddOrder NatOrdinal := ⟨fun x => (nadd_one x).symm⟩ -instance addLeftStrictMono : AddLeftStrictMono NatOrdinal.{u} := +instance : AddLeftStrictMono NatOrdinal.{u} := ⟨fun a _ _ h => nadd_lt_nadd_left h a⟩ -instance addLeftMono : AddLeftMono NatOrdinal.{u} := +instance : AddLeftMono NatOrdinal.{u} := ⟨fun a _ _ h => nadd_le_nadd_left h a⟩ -instance addLeftReflectLE : AddLeftReflectLE NatOrdinal.{u} := +instance : AddLeftReflectLE NatOrdinal.{u} := ⟨fun a b c h => by by_contra! h' exact h.not_lt (add_lt_add_left h' a)⟩ -instance orderedCancelAddCommMonoid : OrderedCancelAddCommMonoid NatOrdinal := - { NatOrdinal.linearOrder with +instance : OrderedCancelAddCommMonoid NatOrdinal := + { NatOrdinal.instLinearOrder with add := (· + ·) add_assoc := nadd_assoc add_le_add_left := fun _ _ => add_le_add_left @@ -339,7 +341,7 @@ instance orderedCancelAddCommMonoid : OrderedCancelAddCommMonoid NatOrdinal := add_comm := nadd_comm nsmul := nsmulRec } -instance addMonoidWithOne : AddMonoidWithOne NatOrdinal := +instance : AddMonoidWithOne NatOrdinal := AddMonoidWithOne.unary @[deprecated Order.succ_eq_add_one (since := "2024-09-04")] @@ -674,8 +676,8 @@ instance : Mul NatOrdinal := -- Porting note: had to add universe annotations to ensure that the -- two sources lived in the same universe. instance : OrderedCommSemiring NatOrdinal.{u} := - { NatOrdinal.orderedCancelAddCommMonoid.{u}, - NatOrdinal.linearOrder.{u} with + { NatOrdinal.instOrderedCancelAddCommMonoid.{u}, + NatOrdinal.instLinearOrder.{u} with mul := (· * ·) left_distrib := nmul_nadd right_distrib := nadd_nmul From 955e8f97a6372ceeeb97f4acc87f71ae1fea7d85 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 19 Nov 2024 05:23:28 +0000 Subject: [PATCH 04/90] chore: fix defeq abuse between `List.get_mem` and `List.getElem_mem` (#19224) In all these places the goal state is about `getElem` not `List.get`. This will reduce the fallout of leanprover/lean4#6095. --- Mathlib/Data/List/EditDistance/Bounds.lean | 2 +- Mathlib/Data/List/MinMax.lean | 2 +- Mathlib/Data/List/Nodup.lean | 2 +- Mathlib/Data/List/NodupEquivFin.lean | 2 +- Mathlib/Data/List/Permutation.lean | 4 ++-- 5 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/List/EditDistance/Bounds.lean b/Mathlib/Data/List/EditDistance/Bounds.lean index 7a20e9e681c9a..df199af6ccfa1 100644 --- a/Mathlib/Data/List/EditDistance/Bounds.lean +++ b/Mathlib/Data/List/EditDistance/Bounds.lean @@ -79,7 +79,7 @@ theorem le_suffixLevenshtein_append_minimum (xs : List α) (ys₁ ys₂) : theorem suffixLevenshtein_minimum_le_levenshtein_append (xs ys₁ ys₂) : (suffixLevenshtein C xs ys₂).1.minimum ≤ levenshtein C xs (ys₁ ++ ys₂) := by cases ys₁ with - | nil => exact List.minimum_le_of_mem' (List.get_mem _ _ _) + | nil => exact List.minimum_le_of_mem' (List.getElem_mem _) | cons y ys₁ => exact (le_suffixLevenshtein_append_minimum _ _ _).trans (suffixLevenshtein_minimum_le_levenshtein_cons _ _ _) diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 6b18d25cf4107..0237438fa2eda 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -427,7 +427,7 @@ theorem minimum_of_length_pos_le_of_mem (h : a ∈ l) (w : 0 < l.length) : theorem getElem_le_maximum_of_length_pos {i : ℕ} (w : i < l.length) (h := (Nat.zero_lt_of_lt w)) : l[i] ≤ l.maximum_of_length_pos h := by apply le_maximum_of_length_pos_of_mem - exact get_mem l i w + exact getElem_mem _ theorem minimum_of_length_pos_le_getElem {i : ℕ} (w : i < l.length) (h := (Nat.zero_lt_of_lt w)) : l.minimum_of_length_pos h ≤ l[i] := diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index f6b34a9b86b77..1368f8f1890a3 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -122,7 +122,7 @@ theorem not_nodup_of_get_eq_of_ne (xs : List α) (n m : Fin xs.length) theorem indexOf_getElem [DecidableEq α] {l : List α} (H : Nodup l) (i : Nat) (h : i < l.length) : indexOf l[i] l = i := - suffices (⟨indexOf l[i] l, indexOf_lt_length.2 (get_mem _ _ _)⟩ : Fin l.length) = ⟨i, h⟩ + suffices (⟨indexOf l[i] l, indexOf_lt_length.2 (getElem_mem _)⟩ : Fin l.length) = ⟨i, h⟩ from Fin.val_eq_of_eq this nodup_iff_injective_get.1 H (by simp) diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 0dcf11280caae..65a48a57e4574 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -127,7 +127,7 @@ theorem sublist_of_orderEmbedding_get?_eq {l l' : List α} (f : ℕ ↪o ℕ) rw [← List.take_append_drop (f 0 + 1) l', ← List.singleton_append] apply List.Sublist.append _ (IH _ this) rw [List.singleton_sublist, ← h, l'.getElem_take' _ (Nat.lt_succ_self _)] - apply List.get_mem + exact List.getElem_mem _ /-- A `l : List α` is `Sublist l l'` for `l' : List α` iff there is `f`, an order-preserving embedding of `ℕ` into `ℕ` such that diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index d90a8bbe313e9..2ce7b4ce525f3 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -501,13 +501,13 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : rcases lt_trichotomy n m with (ht | ht | ht) · suffices x ∈ bs by exact h x (hb.subset this) rfl rw [← hx', getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hm)] - exact get_mem _ _ _ + exact getElem_mem _ · simp only [ht] at hm' hn' rw [← hm'] at hn' exact H (insertIdx_injective _ _ hn') · suffices x ∈ as by exact h x (ha.subset this) rfl rw [← hx, getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hn)] - exact get_mem _ _ _ + exact getElem_mem _ lemma permutations_take_two (x y : α) (s : List α) : (x :: y :: s).permutations.take 2 = [x :: y :: s, y :: x :: s] := by From 7f8d4bd153dac2a76d5b0396860693a25af10051 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Tue, 19 Nov 2024 05:48:11 +0000 Subject: [PATCH 05/90] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6095 --- lake-manifest.json | 4 ++-- lakefile.lean | 2 +- lean-toolchain | 2 +- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index f151ad4bbb7d2..9316621819494 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "23b4216d1e3a2064058bb415c11e5f6df0417c99", + "rev": "a8815f0d111d6cea6c9301dfdebf32b92e7b8641", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6095", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..2bb0f3940e9cd 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,7 +7,7 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "nightly-testing" +require "leanprover-community" / "batteries" @ git "lean-pr-testing-6095" require "leanprover-community" / "Qq" @ git "nightly-testing" require "leanprover-community" / "aesop" @ git "nightly-testing" require "leanprover-community" / "proofwidgets" @ git "v0.0.46" diff --git a/lean-toolchain b/lean-toolchain index 585547139ef08..33df9a16954bb 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-18 +leanprover/lean4-pr-releases:pr-release-6095 From 9274ddbe2cedc8b212a9e0d0dc92094a830c7abc Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Tue, 19 Nov 2024 07:30:29 +0000 Subject: [PATCH 06/90] feat(Tactic/StacksAttribute): improve comment shown in docstring for stack tags (#19235) Previously it was ```text Stacks Tag 09HK Part 1, finSepDegree variant ``` Now it should be ```text Stacks Tag 09HK (Part 1, finSepDegree variant) ``` which is more compact. --- Mathlib/Tactic/StacksAttribute.lean | 3 ++- MathlibTest/StacksAttribute.lean | 9 +++++++++ 2 files changed, 11 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/StacksAttribute.lean b/Mathlib/Tactic/StacksAttribute.lean index c8f7d829151e7..c7f69363dca6e 100644 --- a/Mathlib/Tactic/StacksAttribute.lean +++ b/Mathlib/Tactic/StacksAttribute.lean @@ -153,7 +153,8 @@ initialize Lean.registerBuiltinAttribute { | _ => throwUnsupportedSyntax let tagStr ← tag.getStacksTag let comment := (comment.map (·.getString)).getD "" - let newDoc := [oldDoc, s!"[{SorK} Tag {tagStr}]({url}/{tagStr})", comment] + let commentInDoc := if comment = "" then "" else s!" ({comment})" + let newDoc := [oldDoc, s!"[{SorK} Tag {tagStr}]({url}/{tagStr}){commentInDoc}"] addDocString decl <| "\n\n".intercalate (newDoc.filter (· != "")) addTagEntry decl database tagStr <| comment } diff --git a/MathlibTest/StacksAttribute.lean b/MathlibTest/StacksAttribute.lean index 92fe5efb40fcd..3b70bdb597610 100644 --- a/MathlibTest/StacksAttribute.lean +++ b/MathlibTest/StacksAttribute.lean @@ -11,6 +11,15 @@ theorem tagged : True := .intro end X +/-- +info: some ([Stacks Tag A04Q](https://stacks.math.columbia.edu/tag/A04Q) (A comment) + +[Kerodon Tag B15R](https://kerodon.net/tag/B15R) (Also a comment)) +-/ +#guard_msgs in +run_cmd + Lean.logInfo m!"{← Lean.findDocString? (← Lean.getEnv) `X.tagged}" + #guard_msgs in @[stacks 0BR2, kerodon 0X12] example : True := .intro From 3a1c2a7c444843bcfdea46b41c78b932217f5612 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Tue, 19 Nov 2024 08:16:25 +0000 Subject: [PATCH 07/90] feat(LinearAlgebra/Matrix): Permanent (#18936) Co-authored-by: Moritz Firsching --- Mathlib.lean | 1 + Mathlib/LinearAlgebra/Matrix/Permanent.lean | 86 +++++++++++++++++++++ 2 files changed, 87 insertions(+) create mode 100644 Mathlib/LinearAlgebra/Matrix/Permanent.lean diff --git a/Mathlib.lean b/Mathlib.lean index e9bea276ba9f6..7566572fa39e6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3327,6 +3327,7 @@ import Mathlib.LinearAlgebra.Matrix.MvPolynomial import Mathlib.LinearAlgebra.Matrix.Nondegenerate import Mathlib.LinearAlgebra.Matrix.NonsingularInverse import Mathlib.LinearAlgebra.Matrix.Orthogonal +import Mathlib.LinearAlgebra.Matrix.Permanent import Mathlib.LinearAlgebra.Matrix.Permutation import Mathlib.LinearAlgebra.Matrix.Polynomial import Mathlib.LinearAlgebra.Matrix.PosDef diff --git a/Mathlib/LinearAlgebra/Matrix/Permanent.lean b/Mathlib/LinearAlgebra/Matrix/Permanent.lean new file mode 100644 index 0000000000000..fa1620053c90a --- /dev/null +++ b/Mathlib/LinearAlgebra/Matrix/Permanent.lean @@ -0,0 +1,86 @@ +/- +Copyright (c) 2024 Moritz Firsching. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Firsching +-/ +import Mathlib.Algebra.BigOperators.GroupWithZero.Finset +import Mathlib.Data.Fintype.Perm +import Mathlib.Data.Matrix.Diagonal +/-! +# Permanent of a matrix + +This file defines the permanent of a matrix, `Matrix.permanent`, and some of its properties. + +## Main definitions + +* `Matrix.permanent`: the permanent of a square matrix, as a sum over permutations + +-/ + +open Equiv Fintype Finset + +namespace Matrix + +variable {n : Type*} [DecidableEq n] [Fintype n] +variable {R : Type*} [CommSemiring R] + +/-- The permanent of a square matrix defined as a sum over all permutations. This is analogous to +the determinant but without alternating signs. -/ +def permanent (M : Matrix n n R) : R := ∑ σ : Perm n, ∏ i, M (σ i) i + +@[simp] +theorem permanent_diagonal {d : n → R} : permanent (diagonal d) = ∏ i, d i := by + refine (sum_eq_single 1 (fun σ _ hσ ↦ ?_) (fun h ↦ (h <| mem_univ _).elim)).trans ?_ + · match not_forall.mp (mt Equiv.ext hσ) with + | ⟨x, hx⟩ => exact Finset.prod_eq_zero (mem_univ x) (if_neg hx) + · simp only [Perm.one_apply, diagonal_apply_eq] + +@[simp] +theorem permanent_zero [Nonempty n] : permanent (0 : Matrix n n R) = 0 := by simp [permanent] + +@[simp] +theorem permanent_one : permanent (1 : Matrix n n R) = 1 := by + rw [← diagonal_one]; simp [-diagonal_one] + +theorem permanent_isEmpty [IsEmpty n] {A : Matrix n n R} : permanent A = 1 := by simp [permanent] + +theorem permanent_eq_one_of_card_eq_zero {A : Matrix n n R} (h : card n = 0) : permanent A = 1 := + haveI : IsEmpty n := card_eq_zero_iff.mp h + permanent_isEmpty + +/-- If `n` has only one element, the permanent of an `n` by `n` matrix is just that element. +Although `Unique` implies `DecidableEq` and `Fintype`, the instances might +not be syntactically equal. Thus, we need to fill in the args explicitly. -/ +@[simp] +theorem permanent_unique {n : Type*} [Unique n] [DecidableEq n] [Fintype n] (A : Matrix n n R) : + permanent A = A default default := by simp [permanent, univ_unique] + +theorem permanent_eq_elem_of_subsingleton [Subsingleton n] (A : Matrix n n R) (k : n) : + permanent A = A k k := by + have := uniqueOfSubsingleton k + convert permanent_unique A + +theorem permanent_eq_elem_of_card_eq_one {A : Matrix n n R} (h : card n = 1) (k : n) : + permanent A = A k k := + haveI : Subsingleton n := card_le_one_iff_subsingleton.mp h.le + permanent_eq_elem_of_subsingleton _ _ + +/-- Transposing a matrix preserves the permanent. -/ +@[simp] +theorem permanent_transpose (M : Matrix n n R) : Mᵀ.permanent = M.permanent := by + refine sum_bijective _ inv_involutive.bijective _ _ ?_ + intro σ + apply Fintype.prod_equiv σ + simp + +/-- Permuting the columns does not change the permanent. -/ +theorem permanent_permute_cols (σ : Perm n) (M : Matrix n n R) : + (M.submatrix σ id).permanent = M.permanent := + (Group.mulLeft_bijective σ).sum_comp fun τ ↦ ∏ i : n, M (τ i) i + +/-- Permuting the rows does not change the permanent. -/ +theorem permanent_permute_rows (σ : Perm n) (M : Matrix n n R) : + (M.submatrix id σ).permanent = M.permanent := by + rw [← permanent_transpose, transpose_submatrix, permanent_permute_cols, permanent_transpose] + +end Matrix From 4363a4db6b7fb119dc9b2e9f229b2cb36a047b08 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 19:23:17 +1100 Subject: [PATCH 08/90] bump batteries --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 9316621819494..cbfa72fcc2d46 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "a8815f0d111d6cea6c9301dfdebf32b92e7b8641", + "rev": "6336d636051c3fbe945113fae80553982f09a87e", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6095", From 809872da938e18d9d335b6c61e41f9c3cd4fac25 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 19 Nov 2024 08:25:54 +0000 Subject: [PATCH 09/90] refactor(CategoryTheory): make PreservesLimit and ReflectsLimit props (#19206) All changes are straightforward. Some universe parameters had to be made explicit, but without changing the level of generality of the definitions/statements. --- .../Algebra/Category/AlgebraCat/Limits.lean | 24 +- Mathlib/Algebra/Category/Grp/AB5.lean | 4 +- .../Category/Grp/FilteredColimits.lean | 18 +- Mathlib/Algebra/Category/Grp/Limits.lean | 60 +- .../Algebra/Category/ModuleCat/Abelian.lean | 16 +- .../Category/ModuleCat/ChangeOfRings.lean | 11 +- .../Algebra/Category/ModuleCat/Colimits.lean | 4 +- .../Category/ModuleCat/FilteredColimits.lean | 12 +- .../Algebra/Category/ModuleCat/Limits.lean | 34 +- .../Category/ModuleCat/Presheaf/Colimits.lean | 16 +- .../Category/ModuleCat/Presheaf/Limits.lean | 16 +- .../ModuleCat/Presheaf/Sheafification.lean | 4 +- .../Category/ModuleCat/Sheaf/Limits.lean | 4 +- .../Category/MonCat/FilteredColimits.lean | 12 +- Mathlib/Algebra/Category/MonCat/Limits.lean | 30 +- .../Category/Ring/FilteredColimits.lean | 33 +- Mathlib/Algebra/Category/Ring/Limits.lean | 88 +-- .../Homology/DerivedCategory/Ext/Basic.lean | 2 +- .../Homology/HomologicalComplexBiprod.lean | 2 +- .../Homology/HomologicalComplexLimits.lean | 28 +- .../ShortComplex/ConcreteCategory.lean | 4 +- .../Homology/ShortComplex/ExactFunctor.lean | 46 +- .../Algebra/Homology/ShortComplex/Limits.lean | 14 +- .../ShortComplex/PreservesHomology.lean | 60 +- Mathlib/AlgebraicGeometry/AffineScheme.lean | 2 +- .../GammaSpecAdjunction.lean | 4 +- Mathlib/AlgebraicGeometry/Limits.lean | 24 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 18 +- .../FundamentalGroupoid/Product.lean | 4 +- Mathlib/CategoryTheory/Abelian/Exact.lean | 14 +- .../Abelian/GrothendieckAxioms.lean | 18 +- Mathlib/CategoryTheory/Abelian/Injective.lean | 8 +- .../CategoryTheory/Abelian/Projective.lean | 8 +- Mathlib/CategoryTheory/Abelian/Transfer.lean | 8 +- Mathlib/CategoryTheory/Adhesive.lean | 8 +- Mathlib/CategoryTheory/Adjunction/Limits.lean | 97 ++- .../CategoryTheory/Category/Cat/Limit.lean | 2 +- .../CategoryTheory/ChosenFiniteProducts.lean | 24 +- .../ChosenFiniteProducts/FunctorCategory.lean | 4 +- Mathlib/CategoryTheory/Closed/Cartesian.lean | 2 +- Mathlib/CategoryTheory/Closed/Ideal.lean | 28 +- Mathlib/CategoryTheory/Closed/Monoidal.lean | 2 +- Mathlib/CategoryTheory/Extensive.lean | 20 +- Mathlib/CategoryTheory/Functor/Flat.lean | 78 +- Mathlib/CategoryTheory/Galois/Action.lean | 6 +- Mathlib/CategoryTheory/Galois/Basic.lean | 14 +- Mathlib/CategoryTheory/Galois/Examples.lean | 2 +- .../CategoryTheory/Galois/GaloisObjects.lean | 2 +- .../Galois/Prorepresentability.lean | 2 +- Mathlib/CategoryTheory/GlueData.lean | 2 +- .../CategoryTheory/GradedObject/Monoidal.lean | 1 + Mathlib/CategoryTheory/Limits/Comma.lean | 9 +- .../CategoryTheory/Limits/ConeCategory.lean | 10 +- Mathlib/CategoryTheory/Limits/Connected.lean | 9 +- .../Limits/Constructions/BinaryProducts.lean | 8 +- .../Limits/Constructions/Equalizers.lean | 9 +- .../FiniteProductsOfBinaryProducts.lean | 36 +- .../LimitsOfProductsAndEqualizers.lean | 47 +- Mathlib/CategoryTheory/Limits/Creates.lean | 79 +- .../CategoryTheory/Limits/ExactFunctor.lean | 36 +- .../FilteredColimitCommutesFiniteLimit.lean | 16 +- Mathlib/CategoryTheory/Limits/FintypeCat.lean | 12 +- .../Limits/FunctorCategory/Basic.lean | 83 ++- Mathlib/CategoryTheory/Limits/Over.lean | 8 +- .../Limits/Preserves/Basic.lean | 679 +++++++++++++----- .../Limits/Preserves/Filtered.lean | 78 +- .../Limits/Preserves/Finite.lean | 158 ++-- .../Limits/Preserves/FunctorCategory.lean | 63 +- .../Limits/Preserves/Limits.lean | 22 +- .../Limits/Preserves/Opposites.lean | 374 +++++----- .../Limits/Preserves/Presheaf.lean | 6 +- .../Preserves/Shapes/BinaryProducts.lean | 16 +- .../Limits/Preserves/Shapes/Biproducts.lean | 51 +- .../Limits/Preserves/Shapes/Equalizers.lean | 28 +- .../Limits/Preserves/Shapes/Kernels.lean | 28 +- .../Limits/Preserves/Shapes/Products.lean | 26 +- .../Limits/Preserves/Shapes/Pullbacks.lean | 46 +- .../Limits/Preserves/Shapes/Terminal.lean | 55 +- .../Limits/Preserves/Shapes/Zero.lean | 34 +- .../Limits/Preserves/Ulift.lean | 7 +- .../Limits/Preserves/Yoneda.lean | 2 +- Mathlib/CategoryTheory/Limits/Presheaf.lean | 19 +- .../Limits/Shapes/NormalMono/Basic.lean | 28 +- Mathlib/CategoryTheory/Limits/VanKampen.lean | 12 +- Mathlib/CategoryTheory/Limits/Yoneda.lean | 46 +- .../Localization/FiniteProducts.lean | 7 +- .../CategoryTheory/Monad/Comonadicity.lean | 10 +- Mathlib/CategoryTheory/Monad/Limits.lean | 54 +- Mathlib/CategoryTheory/Monad/Monadicity.lean | 10 +- .../Monoidal/Internal/Limits.lean | 4 +- .../CategoryTheory/Monoidal/Preadditive.lean | 8 +- .../Preadditive/AdditiveFunctor.lean | 10 +- .../Preadditive/Biproducts.lean | 113 +-- .../CategoryTheory/Preadditive/Injective.lean | 4 +- .../CategoryTheory/Preadditive/LeftExact.lean | 73 +- .../Preadditive/Projective.lean | 4 +- .../Preadditive/Yoneda/Limits.lean | 13 +- Mathlib/CategoryTheory/Sites/Abelian.lean | 2 +- .../Sites/Coherent/ExtensiveSheaves.lean | 19 +- .../Sites/Coherent/LocallySurjective.lean | 2 +- .../Sites/Coherent/SheafComparison.lean | 26 +- Mathlib/CategoryTheory/Sites/Equivalence.lean | 2 +- Mathlib/CategoryTheory/Sites/LeftExact.lean | 80 ++- Mathlib/CategoryTheory/Sites/Preserves.lean | 23 +- Mathlib/CategoryTheory/Sites/Pullback.lean | 6 +- Mathlib/CategoryTheory/Sites/Sheaf.lean | 6 +- .../CategoryTheory/Sites/Sheafification.lean | 2 +- .../CategoryTheory/Triangulated/Functor.lean | 4 +- .../Triangulated/HomologicalFunctor.lean | 6 +- .../Condensed/Discrete/Characterization.lean | 8 +- Mathlib/Condensed/Discrete/Colimit.lean | 4 +- Mathlib/Condensed/Explicit.lean | 18 +- Mathlib/Condensed/Light/Explicit.lean | 6 +- Mathlib/Condensed/TopComparison.lean | 6 +- .../LocallyRingedSpace/HasColimits.lean | 8 +- .../Geometry/RingedSpace/OpenImmersion.lean | 58 +- .../PresheafedSpace/HasColimits.lean | 15 +- .../Geometry/RingedSpace/SheafedSpace.lean | 4 +- .../RepresentationTheory/Action/Basic.lean | 10 +- .../RepresentationTheory/Action/Limits.lean | 52 +- Mathlib/RepresentationTheory/Rep.lean | 4 +- .../Algebra/Category/ProfiniteGrp/Basic.lean | 2 +- .../Category/CompHausLike/Limits.lean | 10 +- .../Category/LightProfinite/Basic.lean | 4 +- .../Topology/Category/Profinite/Basic.lean | 4 +- .../Category/Stonean/Adjunctions.lean | 2 +- .../Category/TopCat/Limits/Basic.lean | 19 +- Mathlib/Topology/Category/TopCat/Yoneda.lean | 4 +- .../SheafCondition/PairwiseIntersections.lean | 4 +- 129 files changed, 2102 insertions(+), 1712 deletions(-) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean index 44fe90192172a..94f312fee5673 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean @@ -155,41 +155,41 @@ instance hasLimits : HasLimits (AlgebraCat.{w} R) := /-- The forgetful functor from R-algebras to rings preserves all limits. -/ -instance forget₂RingPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget₂Ring_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget₂ (AlgebraCat.{w} R) RingCat.{w}) where preservesLimitsOfShape := { preservesLimit := fun {K} ↦ - preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (RingCat.limitConeIsLimit.{v, w} (_ ⋙ forget₂ (AlgebraCat.{w} R) RingCat.{w})) } -instance forget₂RingPreservesLimits : PreservesLimits (forget₂ (AlgebraCat R) RingCat.{w}) := - AlgebraCat.forget₂RingPreservesLimitsOfSize.{w, w} +instance forget₂Ring_preservesLimits : PreservesLimits (forget₂ (AlgebraCat R) RingCat.{w}) := + AlgebraCat.forget₂Ring_preservesLimitsOfSize.{w, w} /-- The forgetful functor from R-algebras to R-modules preserves all limits. -/ -instance forget₂ModulePreservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} +instance forget₂Module_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget₂ (AlgebraCat.{w} R) (ModuleCat.{w} R)) where preservesLimitsOfShape := { preservesLimit := fun {K} ↦ - preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (ModuleCat.HasLimits.limitConeIsLimit (K ⋙ forget₂ (AlgebraCat.{w} R) (ModuleCat.{w} R))) } -instance forget₂ModulePreservesLimits : +instance forget₂Module_preservesLimits : PreservesLimits (forget₂ (AlgebraCat R) (ModuleCat.{w} R)) := - AlgebraCat.forget₂ModulePreservesLimitsOfSize.{w, w} + AlgebraCat.forget₂Module_preservesLimitsOfSize.{w, w} /-- The forgetful functor from R-algebras to types preserves all limits. -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget (AlgebraCat.{w} R)) where preservesLimitsOfShape := { preservesLimit := fun {K} ↦ - preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (Types.Small.limitConeIsLimit.{v} (K ⋙ forget _)) } -instance forgetPreservesLimits : PreservesLimits (forget (AlgebraCat.{w} R)) := - AlgebraCat.forgetPreservesLimitsOfSize.{w, w} +instance forget_preservesLimits : PreservesLimits (forget (AlgebraCat.{w} R)) := + AlgebraCat.forget_preservesLimitsOfSize.{w, w} end AlgebraCat diff --git a/Mathlib/Algebra/Category/Grp/AB5.lean b/Mathlib/Algebra/Category/Grp/AB5.lean index 2edc3d04999e0..a5230f8bf90ff 100644 --- a/Mathlib/Algebra/Category/Grp/AB5.lean +++ b/Mathlib/Algebra/Category/Grp/AB5.lean @@ -27,7 +27,7 @@ variable {J : Type u} [SmallCategory J] [IsFiltered J] noncomputable instance : (colim (J := J) (C := AddCommGrp.{u})).PreservesHomology := - Functor.preservesHomologyOfMapExact _ (fun S hS ↦ by + Functor.preservesHomology_of_map_exact _ (fun S hS ↦ by replace hS := fun j => hS.map ((evaluation _ _).obj j) simp only [ShortComplex.ab_exact_iff_ker_le_range] at hS ⊢ intro x (hx : _ = _) @@ -45,7 +45,7 @@ noncomputable instance : noncomputable instance : PreservesFiniteLimits <| colim (J := J) (C := AddCommGrp.{u}) := by - apply Functor.preservesFiniteLimitsOfPreservesHomology + apply Functor.preservesFiniteLimits_of_preservesHomology instance : HasFilteredColimits (AddCommGrp.{u}) where HasColimitsOfShape := inferInstance diff --git a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean index b79cc51e0f2ed..a31573223addd 100644 --- a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean @@ -128,18 +128,18 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where ((forget Grp).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x -@[to_additive forget₂AddMonPreservesFilteredColimits] -noncomputable instance forget₂MonPreservesFilteredColimits : +@[to_additive forget₂AddMon_preservesFilteredColimits] +noncomputable instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits.{u} (forget₂ Grp.{u} MonCat.{u}) where preserves_filtered_colimits x hx1 _ := letI : Category.{u, u} x := hx1 - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit.{u, u} _)⟩ @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +noncomputable instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget Grp.{u}) := - Limits.compPreservesFilteredColimits (forget₂ Grp MonCat) (forget MonCat.{u}) + Limits.comp_preservesFilteredColimits (forget₂ Grp MonCat) (forget MonCat.{u}) end @@ -198,19 +198,19 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where ((forget CommGrp).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x @[to_additive] -noncomputable instance forget₂GroupPreservesFilteredColimits : +noncomputable instance forget₂Group_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommGrp Grp.{u}) where preserves_filtered_colimits J hJ1 _ := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (Grp.FilteredColimits.colimitCoconeIsColimit.{u, u} (F ⋙ forget₂ CommGrp Grp.{u})) } @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +noncomputable instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget CommGrp.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommGrp Grp) (forget Grp.{u}) + Limits.comp_preservesFilteredColimits (forget₂ CommGrp Grp) (forget Grp.{u}) end diff --git a/Mathlib/Algebra/Category/Grp/Limits.lean b/Mathlib/Algebra/Category/Grp/Limits.lean index 87ad275a4593f..f9226f51a1742 100644 --- a/Mathlib/Algebra/Category/Grp/Limits.lean +++ b/Mathlib/Algebra/Category/Grp/Limits.lean @@ -161,21 +161,21 @@ This means the underlying monoid of a limit can be computed as a limit in the ca This means the underlying additive monoid of a limit can be computed as a limit in the category of additive monoids.", to_additive_relevant_arg 2] -noncomputable instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Mon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ Grp.{u} MonCat.{u}) where preservesLimitsOfShape {J _} := { } @[to_additive] -noncomputable instance forget₂MonPreservesLimits : +instance forget₂Mon_preservesLimits : PreservesLimits (forget₂ Grp.{u} MonCat.{u}) := - Grp.forget₂MonPreservesLimitsOfSize.{u, u} + Grp.forget₂Mon_preservesLimitsOfSize.{u, u} /-- If `J` is `u`-small, the forgetful functor from `Grp.{u}` preserves limits of shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddGrp.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget Grp.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from groups to types preserves all limits. @@ -186,12 +186,12 @@ This means the underlying type of a limit can be computed as a limit in the cate This means the underlying type of a limit can be computed as a limit in the category of types.", to_additive_relevant_arg 2] -noncomputable instance forgetPreservesLimitsOfSize : +instance forget_preservesLimitsOfSize : PreservesLimitsOfSize.{w, v} (forget Grp.{u}) := inferInstance @[to_additive] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget Grp.{u}) := - Grp.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget Grp.{u}) := + Grp.forget_preservesLimitsOfSize.{u, u} end Grp @@ -307,18 +307,18 @@ instance hasLimits : HasLimits CommGrp.{u} := CommGrp.hasLimitsOfSize.{u, u} @[to_additive] -noncomputable instance forget₂GroupPreservesLimit : +instance forget₂Group_preservesLimit : PreservesLimit F (forget₂ CommGrp.{u} Grp.{u}) where - preserves {c} hc := by + preserves {c} hc := ⟨by have : HasLimit (F ⋙ forget₂ CommGrp Grp) := by rw [Grp.hasLimit_iff_small_sections] change Small.{u} (F ⋙ forget CommGrp).sections rw [← CommGrp.hasLimit_iff_small_sections] exact ⟨_, hc⟩ - exact isLimitOfPreserves _ hc + exact isLimitOfPreserves _ hc⟩ @[to_additive] -noncomputable instance forget₂GroupPreservesLimitsOfShape : +instance forget₂Group_preservesLimitsOfShape : PreservesLimitsOfShape J (forget₂ CommGrp.{u} Grp.{u}) where /-- The forgetful functor from commutative groups to groups preserves all limits. @@ -330,19 +330,19 @@ of groups.) (That is, the underlying group could have been computed instead as limits in the category of additive groups.)", to_additive_relevant_arg 2] -noncomputable instance forget₂GroupPreservesLimitsOfSize : +instance forget₂Group_preservesLimitsOfSize : PreservesLimitsOfSize.{w, v} (forget₂ CommGrp.{u} Grp.{u}) where @[to_additive] -noncomputable instance forget₂GroupPreservesLimits : +instance forget₂Group_preservesLimits : PreservesLimits (forget₂ CommGrp Grp.{u}) := - CommGrp.forget₂GroupPreservesLimitsOfSize.{u, u} + CommGrp.forget₂Group_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ -@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsAux +@[to_additive AddCommGrp.forget₂AddCommMon_preservesLimitsAux "An auxiliary declaration to speed up typechecking."] -noncomputable def forget₂CommMonPreservesLimitsAux +noncomputable def forget₂CommMon_preservesLimitsAux [Small.{u} (F ⋙ forget CommGrp).sections] : IsLimit ((forget₂ CommGrp.{u} CommMonCat.{u}).mapCone (limitCone.{v, u} F)) := letI : Small.{u} (Functor.sections ((F ⋙ forget₂ _ CommMonCat) ⋙ forget CommMonCat)) := @@ -351,23 +351,23 @@ noncomputable def forget₂CommMonPreservesLimitsAux /-- If `J` is `u`-small, the forgetful functor from `CommGrp.{u}` to `CommMonCat.{u}` preserves limits of shape `J`. -/ -@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsOfShape +@[to_additive AddCommGrp.forget₂AddCommMon_preservesLimitsOfShape "If `J` is `u`-small, the forgetful functor from `AddCommGrp.{u}` to `AddCommMonCat.{u}` preserves limits of shape `J`."] -noncomputable instance forget₂CommMonPreservesLimitsOfShape [Small.{u} J] : +instance forget₂CommMon_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget₂ CommGrp.{u} CommMonCat.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) - (forget₂CommMonPreservesLimitsAux.{v, u} F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) + (forget₂CommMon_preservesLimitsAux.{v, u} F) /-- The forgetful functor from commutative groups to commutative monoids preserves all limits. (That is, the underlying commutative monoids could have been computed instead as limits in the category of commutative monoids.) -/ -@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsOfSize +@[to_additive AddCommGrp.forget₂AddCommMon_preservesLimitsOfSize "The forgetful functor from additive commutative groups to additive commutative monoids preserves all limits. (That is, the underlying additive commutative monoids could have been computed instead as limits in the category of additive commutative monoids.)"] -noncomputable instance forget₂CommMonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂CommMon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommGrp CommMonCat.{u}) where preservesLimitsOfShape := { } @@ -375,9 +375,9 @@ noncomputable instance forget₂CommMonPreservesLimitsOfSize [UnivLE.{v, u}] : shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddCommGrp.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget CommGrp.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from commutative groups to types preserves all limits. (That is, the @@ -387,16 +387,16 @@ underlying types could have been computed instead as limits in the category of t "The forgetful functor from additive commutative groups to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.)"] -noncomputable instance forgetPreservesLimitsOfSize : +instance forget_preservesLimitsOfSize : PreservesLimitsOfSize.{w, v} (forget CommGrp.{u}) := inferInstance -noncomputable instance _root_.AddCommGrp.forgetPreservesLimits : +noncomputable instance _root_.AddCommGrp.forget_preservesLimits : PreservesLimits (forget AddCommGrp.{u}) := - AddCommGrp.forgetPreservesLimitsOfSize.{u, u} + AddCommGrp.forget_preservesLimitsOfSize.{u, u} @[to_additive existing] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget CommGrp.{u}) := - CommGrp.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance forget_preservesLimits : PreservesLimits (forget CommGrp.{u}) := + CommGrp.forget_preservesLimitsOfSize.{u, u} -- Verify we can form limits indexed over smaller categories. example (f : ℕ → AddCommGrp) : HasProduct f := by infer_instance diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index 401fb715425c1..fcb99e727a6c4 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -82,19 +82,19 @@ instance : HasLimitsOfSize.{v,v} (ModuleCatMax.{v, w} R) := /- We need to put this in this weird spot because we need to know that the category of modules is balanced. -/ -instance forgetReflectsLimitsOfSize : +instance forget_reflectsLimitsOfSize : ReflectsLimitsOfSize.{v, v} (forget (ModuleCatMax.{v, w} R)) := - reflectsLimitsOfReflectsIsomorphisms + reflectsLimits_of_reflectsIsomorphisms -instance forget₂ReflectsLimitsOfSize : +instance forget₂_reflectsLimitsOfSize : ReflectsLimitsOfSize.{v, v} (forget₂ (ModuleCatMax.{v, w} R) AddCommGrp.{max v w}) := - reflectsLimitsOfReflectsIsomorphisms + reflectsLimits_of_reflectsIsomorphisms -instance forgetReflectsLimits : ReflectsLimits (forget (ModuleCat.{v} R)) := - ModuleCat.forgetReflectsLimitsOfSize.{v, v} +instance forget_reflectsLimits : ReflectsLimits (forget (ModuleCat.{v} R)) := + ModuleCat.forget_reflectsLimitsOfSize.{v, v} -instance forget₂ReflectsLimits : ReflectsLimits (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) := - ModuleCat.forget₂ReflectsLimitsOfSize.{v, v} +instance forget₂_reflectsLimits : ReflectsLimits (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) := + ModuleCat.forget₂_reflectsLimitsOfSize.{v, v} end ReflectsLimits diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index 254e152bf938a..4bbcee39a959e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -827,26 +827,25 @@ instance {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* (restrictScalars.{max u₂ w, u₁, u₂} f).IsRightAdjoint := (extendRestrictScalarsAdj f).isRightAdjoint -noncomputable instance preservesLimitRestrictScalars +noncomputable instance preservesLimit_restrictScalars {R : Type*} {S : Type*} [Ring R] [Ring S] (f : R →+* S) {J : Type*} [Category J] (F : J ⥤ ModuleCat.{v} S) [Small.{v} (F ⋙ forget _).sections] : PreservesLimit F (restrictScalars f) := - ⟨fun {c} hc => by + ⟨fun {c} hc => ⟨by have hc' := isLimitOfPreserves (forget₂ _ AddCommGrp) hc - exact isLimitOfReflects (forget₂ _ AddCommGrp) hc'⟩ + exact isLimitOfReflects (forget₂ _ AddCommGrp) hc'⟩⟩ -instance preservesColimitRestrictScalars {R S : Type*} [Ring R] [Ring S] +instance preservesColimit_restrictScalars {R S : Type*} [Ring R] [Ring S] (f : R →+* S) {J : Type*} [Category J] (F : J ⥤ ModuleCat.{v} S) [HasColimit (F ⋙ forget₂ _ AddCommGrp)] : PreservesColimit F (ModuleCat.restrictScalars.{v} f) := by have : HasColimit ((F ⋙ restrictScalars f) ⋙ forget₂ (ModuleCat R) AddCommGrp) := inferInstanceAs (HasColimit (F ⋙ forget₂ _ AddCommGrp)) - apply preservesColimitOfPreservesColimitCocone (HasColimit.isColimitColimitCocone F) + apply preservesColimit_of_preserves_colimit_cocone (HasColimit.isColimitColimitCocone F) apply isColimitOfReflects (forget₂ _ AddCommGrp) apply isColimitOfPreserves (forget₂ (ModuleCat.{v} S) AddCommGrp.{v}) exact HasColimit.isColimitColimitCocone F - end ModuleCat end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean index 726d06d08e305..358c2eb1ea0fc 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean @@ -97,11 +97,11 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) where instance : HasColimit F := ⟨_, isColimitColimitCocone F⟩ noncomputable instance : PreservesColimit F (forget₂ _ AddCommGrp) := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (colimit.isColimit _) noncomputable instance reflectsColimit : ReflectsColimit F (forget₂ (ModuleCat.{w'} R) AddCommGrp) := - reflectsColimitOfReflectsIsomorphisms _ _ + reflectsColimit_of_reflectsIsomorphisms _ _ end HasColimit diff --git a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean index 4ab0442cf4f77..553ce84f4002e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean @@ -170,22 +170,22 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone F) where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget (ModuleCat R))).uniq ((forget (ModuleCat R)).mapCocone t) _ fun j => funext fun x => LinearMap.congr_fun (h j) x -instance forget₂AddCommGroupPreservesFilteredColimits : +instance forget₂AddCommGroup_preservesFilteredColimits : PreservesFilteredColimits (forget₂ (ModuleCat.{u} R) AddCommGrp.{u}) where preserves_filtered_colimits J _ _ := { -- Porting note: without the curly braces for `F` -- here we get a confusing error message about universes. preservesColimit := fun {F : J ⥤ ModuleCat.{u} R} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) (AddCommGrp.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ (ModuleCat.{u} R) AddCommGrp.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget (ModuleCat.{u} R)) := - Limits.compPreservesFilteredColimits (forget₂ (ModuleCat R) AddCommGrp) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget (ModuleCat.{u} R)) := + Limits.comp_preservesFilteredColimits (forget₂ (ModuleCat R) AddCommGrp) (forget AddCommGrp) -instance forgetReflectsFilteredColimits : ReflectsFilteredColimits (forget (ModuleCat.{u} R)) where - reflects_filtered_colimits _ := { reflectsColimit := reflectsColimitOfReflectsIsomorphisms _ _ } +instance forget_reflectsFilteredColimits : ReflectsFilteredColimits (forget (ModuleCat.{u} R)) where + reflects_filtered_colimits _ := { reflectsColimit := reflectsColimit_of_reflectsIsomorphisms _ _ } end diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index dab88fdebc6cf..862547dd00ba8 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -149,7 +149,7 @@ instance (priority := high) hasLimits' : HasLimits (ModuleCat.{u} R) := /-- An auxiliary declaration to speed up typechecking. -/ -def forget₂AddCommGroupPreservesLimitsAux : +def forget₂AddCommGroup_preservesLimitsAux : IsLimit ((forget₂ (ModuleCat R) AddCommGrp).mapCone (limitCone F)) := letI : Small.{w} (Functor.sections ((F ⋙ forget₂ _ AddCommGrp) ⋙ forget _)) := inferInstanceAs <| Small.{w} (Functor.sections (F ⋙ forget (ModuleCat R))) @@ -157,48 +157,48 @@ def forget₂AddCommGroupPreservesLimitsAux : (F ⋙ forget₂ (ModuleCat.{w} R) _ : J ⥤ AddCommGrp.{w}) /-- The forgetful functor from R-modules to abelian groups preserves all limits. -/ -instance forget₂AddCommGroupPreservesLimit : +instance forget₂AddCommGroup_preservesLimit : PreservesLimit F (forget₂ (ModuleCat R) AddCommGrp) := - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) - (forget₂AddCommGroupPreservesLimitsAux F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) + (forget₂AddCommGroup_preservesLimitsAux F) /-- The forgetful functor from R-modules to abelian groups preserves all limits. -/ -instance forget₂AddCommGroupPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget₂AddCommGroup_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget₂ (ModuleCat.{w} R) AddCommGrp.{w}) where preservesLimitsOfShape := { preservesLimit := inferInstance } -instance forget₂AddCommGroupPreservesLimits : +instance forget₂AddCommGroup_preservesLimits : PreservesLimits (forget₂ (ModuleCat R) AddCommGrp.{w}) := - ModuleCat.forget₂AddCommGroupPreservesLimitsOfSize.{w, w} + ModuleCat.forget₂AddCommGroup_preservesLimitsOfSize.{w, w} /-- The forgetful functor from R-modules to types preserves all limits. -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, w}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} (forget (ModuleCat.{w} R)) where preservesLimitsOfShape := - { preservesLimit := fun {K} ↦ preservesLimitOfPreservesLimitCone (limitConeIsLimit K) + { preservesLimit := fun {K} ↦ preservesLimit_of_preserves_limit_cone (limitConeIsLimit K) (Types.Small.limitConeIsLimit.{v} (_ ⋙ forget _)) } -instance forgetPreservesLimits : PreservesLimits (forget (ModuleCat.{w} R)) := - ModuleCat.forgetPreservesLimitsOfSize.{w, w} +instance forget_preservesLimits : PreservesLimits (forget (ModuleCat.{w} R)) := + ModuleCat.forget_preservesLimitsOfSize.{w, w} end -instance forget₂AddCommGroupReflectsLimit : +instance forget₂AddCommGroup_reflectsLimit : ReflectsLimit F (forget₂ (ModuleCat.{w} R) AddCommGrp) where - reflects {c} hc := by + reflects {c} hc := ⟨by have : HasLimit (F ⋙ forget₂ (ModuleCat R) AddCommGrp) := ⟨_, hc⟩ have : Small.{w} (Functor.sections (F ⋙ forget (ModuleCat R))) := by simpa only [AddCommGrp.hasLimit_iff_small_sections] using this - have := reflectsLimitOfReflectsIsomorphisms F (forget₂ (ModuleCat R) AddCommGrp) - exact isLimitOfReflects _ hc + have := reflectsLimit_of_reflectsIsomorphisms F (forget₂ (ModuleCat R) AddCommGrp) + exact isLimitOfReflects _ hc⟩ -instance forget₂AddCommGroupReflectsLimitOfShape : +instance forget₂AddCommGroup_reflectsLimitOfShape : ReflectsLimitsOfShape J (forget₂ (ModuleCat.{w} R) AddCommGrp) where -instance forget₂AddCommGroupReflectsLimitOfSize : +instance forget₂AddCommGroup_reflectsLimitOfSize : ReflectsLimitsOfSize.{t, v} (forget₂ (ModuleCat.{w} R) AddCommGrp) where section DirectLimit diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean index f58ed17c074ed..fa22a86db7365 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean @@ -106,16 +106,16 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) := instance hasColimit : HasColimit F := ⟨_, isColimitColimitCocone F⟩ -noncomputable instance evaluationPreservesColimit (X : Cᵒᵖ) : +instance evaluation_preservesColimit (X : Cᵒᵖ) : PreservesColimit F (evaluation R X) := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (colimit.isColimit _) variable [∀ X, PreservesColimit F (evaluation R X ⋙ forget₂ (ModuleCat (R.obj X)) AddCommGrp)] -noncomputable instance toPresheafPreservesColimit : +instance toPresheaf_preservesColimit : PreservesColimit F (toPresheaf R) := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (Limits.evaluationJointlyReflectsColimits _ (fun X => isColimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGrp) (isColimitColimitCocone F))) @@ -130,10 +130,10 @@ variable [HasColimitsOfShape J AddCommGrp.{v}] instance hasColimitsOfShape : HasColimitsOfShape J (PresheafOfModules.{v} R) where -noncomputable instance evaluationPreservesColimitsOfShape (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesColimitsOfShape (X : Cᵒᵖ) : PreservesColimitsOfShape J (evaluation R X : PresheafOfModules.{v} R ⥤ _) where -noncomputable instance toPresheafPreservesColimitsOfShape : +noncomputable instance toPresheaf_preservesColimitsOfShape : PreservesColimitsOfShape J (toPresheaf.{v} R) where end HasColimitsOfShape @@ -143,10 +143,10 @@ namespace Finite instance hasFiniteColimits : HasFiniteColimits (PresheafOfModules.{v} R) := ⟨fun _ => inferInstance⟩ -noncomputable instance evaluationPreservesFiniteColimits (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesFiniteColimits (X : Cᵒᵖ) : PreservesFiniteColimits (evaluation.{v} R X) where -noncomputable instance toPresheafPreservesFiniteColimits : +noncomputable instance toPresheaf_preservesFiniteColimits : PreservesFiniteColimits (toPresheaf R) where end Finite diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean index 127f29b50a89e..12b4f9fc866eb 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean @@ -118,13 +118,13 @@ noncomputable def isLimitLimitCone : IsLimit (limitCone F) := instance hasLimit : HasLimit F := ⟨_, isLimitLimitCone F⟩ -noncomputable instance evaluationPreservesLimit (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesLimit (X : Cᵒᵖ) : PreservesLimit F (evaluation R X) := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (limit.isLimit _) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (limit.isLimit _) -noncomputable instance toPresheafPreservesLimit : +noncomputable instance toPresheaf_preservesLimit : PreservesLimit F (toPresheaf R) := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (Limits.evaluationJointlyReflectsLimits _ (fun X => isLimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGrp) (isLimitLimitCone F))) @@ -139,10 +139,10 @@ variable [Small.{v} J] instance hasLimitsOfShape : HasLimitsOfShape J (PresheafOfModules.{v} R) where -noncomputable instance evaluationPreservesLimitsOfShape (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesLimitsOfShape (X : Cᵒᵖ) : PreservesLimitsOfShape J (evaluation R X : PresheafOfModules.{v} R ⥤ _) where -noncomputable instance toPresheafPreservesLimitsOfShape : +noncomputable instance toPresheaf_preservesLimitsOfShape : PreservesLimitsOfShape J (toPresheaf.{v} R) where end Small @@ -152,10 +152,10 @@ section Finite instance hasFiniteLimits : HasFiniteLimits (PresheafOfModules.{v} R) := ⟨fun _ => inferInstance⟩ -noncomputable instance evaluationPreservesFiniteLimits (X : Cᵒᵖ) : +noncomputable instance evaluation_preservesFiniteLimits (X : Cᵒᵖ) : PreservesFiniteLimits (evaluation.{v} R X) where -noncomputable instance toPresheafPreservesFiniteLimits : +noncomputable instance toPresheaf_preservesFiniteLimits : PreservesFiniteLimits (toPresheaf.{v} R) where end Finite diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean index 55d42463d8a56..949298d7e1b84 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean @@ -144,7 +144,7 @@ variable [HasSheafify J AddCommGrp.{v}] noncomputable instance : PreservesFiniteLimits (sheafification.{v} α ⋙ SheafOfModules.toSheaf.{v} R) := - compPreservesFiniteLimits (toPresheaf.{v} R₀) (presheafToSheaf J AddCommGrp) + comp_preservesFiniteLimits (toPresheaf.{v} R₀) (presheafToSheaf J AddCommGrp) instance : (SheafOfModules.toSheaf.{v} R ⋙ sheafToPresheaf _ _).ReflectsIsomorphisms := inferInstanceAs (SheafOfModules.forget.{v} R ⋙ toPresheaf _).ReflectsIsomorphisms @@ -156,7 +156,7 @@ noncomputable instance : ReflectsFiniteLimits (SheafOfModules.toSheaf.{v} R) whe reflects _ _ _ := inferInstance noncomputable instance : PreservesFiniteLimits (sheafification.{v} α) := - preservesFiniteLimitsOfReflectsOfPreserves + preservesFiniteLimits_of_reflects_of_preserves (sheafification.{v} α) (SheafOfModules.toSheaf.{v} R) end diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean index d82c0027f62ba..91404085a45b9 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean @@ -111,9 +111,9 @@ noncomputable instance forgetPreservesLimitsOfSize : noncomputable instance : PreservesFiniteLimits (SheafOfModules.toSheaf.{v} R ⋙ sheafToPresheaf _ _) := - compPreservesFiniteLimits (SheafOfModules.forget.{v} R) (PresheafOfModules.toPresheaf R.val) + comp_preservesFiniteLimits (SheafOfModules.forget.{v} R) (PresheafOfModules.toPresheaf R.val) noncomputable instance : PreservesFiniteLimits (SheafOfModules.toSheaf.{v} R) := - preservesFiniteLimitsOfReflectsOfPreserves _ (sheafToPresheaf _ _) + preservesFiniteLimits_of_reflects_of_preserves _ (sheafToPresheaf _ _) end SheafOfModules diff --git a/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean b/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean index db7946f579819..d12f9961dbc01 100644 --- a/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/MonCat/FilteredColimits.lean @@ -296,10 +296,10 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where fun j => funext fun x => DFunLike.congr_fun (i := MonCat.instFunLike _ _) (h j) x) y @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget MonCat.{u}) := ⟨fun J hJ1 _ => letI hJ1' : Category J := hJ1 - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget MonCat.{u}))⟩⟩ end @@ -369,16 +369,16 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where DFunLike.congr_fun (i := CommMonCat.instFunLike _ _) (h j) x @[to_additive forget₂AddMonPreservesFilteredColimits] -noncomputable instance forget₂MonPreservesFilteredColimits : +noncomputable instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommMonCat MonCat.{u}) := ⟨fun J hJ1 _ => letI hJ3 : Category J := hJ1 - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ CommMonCat MonCat.{u}))⟩⟩ @[to_additive] -noncomputable instance forgetPreservesFilteredColimits : +noncomputable instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget CommMonCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommMonCat MonCat) (forget MonCat) + Limits.comp_preservesFilteredColimits (forget₂ CommMonCat MonCat) (forget MonCat) end diff --git a/Mathlib/Algebra/Category/MonCat/Limits.lean b/Mathlib/Algebra/Category/MonCat/Limits.lean index 933d00567f3f1..8145dc0c3be8e 100644 --- a/Mathlib/Algebra/Category/MonCat/Limits.lean +++ b/Mathlib/Algebra/Category/MonCat/Limits.lean @@ -137,9 +137,9 @@ instance hasLimits : HasLimits MonCat.{u} := /-- If `J` is `u`-small, the forgetful functor from `MonCat.{u}` preserves limits of shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddMonCat.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +noncomputable instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget MonCat.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from monoids to types preserves all limits. @@ -149,13 +149,13 @@ This means the underlying type of a limit can be computed as a limit in the cate "The forgetful functor from additive monoids to types preserves all limits.\n\n This means the underlying type of a limit can be computed as a limit in the category of types.", to_additive_relevant_arg 2] -noncomputable instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +noncomputable instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget MonCat.{u}) where preservesLimitsOfShape := { } @[to_additive] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget MonCat.{u}) := - MonCat.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance forget_preservesLimits : PreservesLimits (forget MonCat.{u}) := + MonCat.forget_preservesLimitsOfSize.{u, u} end MonCat @@ -256,22 +256,22 @@ This means the underlying type of a limit can be computed as a limit in the cate This means the underlying type of a limit can be computed as a limit in the category of additive\n monoids.", to_additive_relevant_arg 2] -noncomputable instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Mon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommMonCat.{u} MonCat.{u}) where preservesLimitsOfShape {J} 𝒥 := { } @[to_additive] -noncomputable instance forget₂MonPreservesLimits : +instance forget₂Mon_preservesLimits : PreservesLimits (forget₂ CommMonCat.{u} MonCat.{u}) := - CommMonCat.forget₂MonPreservesLimitsOfSize.{u, u} + CommMonCat.forget₂Mon_preservesLimitsOfSize.{u, u} /-- If `J` is `u`-small, the forgetful functor from `CommMonCat.{u}` preserves limits of shape `J`. -/ @[to_additive "If `J` is `u`-small, the forgetful functor from `AddCommMonCat.{u}`\n preserves limits of shape `J`."] -noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : +instance forget_preservesLimitsOfShape [Small.{u} J] : PreservesLimitsOfShape J (forget CommMonCat.{u}) where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) /-- The forgetful functor from commutative monoids to types preserves all limits. @@ -280,16 +280,16 @@ This means the underlying type of a limit can be computed as a limit in the cate @[to_additive "The forgetful functor from additive commutative monoids to types preserves all\n limits.\n\n This means the underlying type of a limit can be computed as a limit in the category of types."] -noncomputable instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{v, v} (forget CommMonCat.{u}) where preservesLimitsOfShape {_} _ := { } -noncomputable instance _root_.AddCommMonCat.forgetPreservesLimits : +instance _root_.AddCommMonCat.forget_preservesLimits : PreservesLimits (forget AddCommMonCat.{u}) := - AddCommMonCat.forgetPreservesLimitsOfSize.{u, u} + AddCommMonCat.forget_preservesLimitsOfSize.{u, u} @[to_additive existing] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget CommMonCat.{u}) := - CommMonCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget CommMonCat.{u}) := + CommMonCat.forget_preservesLimitsOfSize.{u, u} end CommMonCat diff --git a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean index 9b383c982cbde..197ff8afa8f5c 100644 --- a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean @@ -137,16 +137,16 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget SemiRingCat)).uniq ((forget SemiRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂MonPreservesFilteredColimits : +instance forget₂Mon_preservesFilteredColimits : PreservesFilteredColimits (forget₂ SemiRingCat MonCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ SemiRingCat MonCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget SemiRingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ SemiRingCat MonCat) (forget MonCat.{u}) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget SemiRingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ SemiRingCat MonCat) (forget MonCat.{u}) end @@ -197,17 +197,18 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget CommSemiRingCat)).uniq ((forget CommSemiRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂SemiRingPreservesFilteredColimits : +instance forget₂SemiRing_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (SemiRingCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ CommSemiRingCat SemiRingCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget CommSemiRingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat) +instance forget_preservesFilteredColimits : + PreservesFilteredColimits (forget CommSemiRingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ CommSemiRingCat SemiRingCat) (forget SemiRingCat.{u}) end @@ -259,17 +260,17 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget RingCat)).uniq ((forget RingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂SemiRingPreservesFilteredColimits : +instance forget₂SemiRing_preservesFilteredColimits : PreservesFilteredColimits (forget₂ RingCat SemiRingCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (SemiRingCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ RingCat SemiRingCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget RingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ RingCat SemiRingCat) (forget SemiRingCat.{u}) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget RingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ RingCat SemiRingCat) (forget SemiRingCat.{u}) end @@ -319,16 +320,16 @@ def colimitCoconeIsColimit : IsColimit <| colimitCocone.{v, u} F where (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget CommRingCat)).uniq ((forget CommRingCat).mapCocone t) _ fun j => funext fun x => RingHom.congr_fun (h j) x -instance forget₂RingPreservesFilteredColimits : +instance forget₂Ring_preservesFilteredColimits : PreservesFilteredColimits (forget₂ CommRingCat RingCat.{u}) where preserves_filtered_colimits {J hJ1 _} := letI : Category J := hJ1 { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit.{u, u} F) (RingCat.FilteredColimits.colimitCoconeIsColimit (F ⋙ forget₂ CommRingCat RingCat.{u})) } -instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget CommRingCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommRingCat RingCat) (forget RingCat.{u}) +instance forget_preservesFilteredColimits : PreservesFilteredColimits (forget CommRingCat.{u}) := + Limits.comp_preservesFilteredColimits (forget₂ CommRingCat RingCat) (forget RingCat.{u}) end diff --git a/Mathlib/Algebra/Category/Ring/Limits.lean b/Mathlib/Algebra/Category/Ring/Limits.lean index 13fea077bb9d5..7a46166affcf1 100644 --- a/Mathlib/Algebra/Category/Ring/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Limits.lean @@ -139,16 +139,16 @@ def forget₂AddCommMonPreservesLimitsAux : /-- The forgetful functor from semirings to additive commutative monoids preserves all limits. -/ -instance forget₂AddCommMonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂AddCommMon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ SemiRingCat AddCommMonCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (forget₂AddCommMonPreservesLimitsAux F) } -instance forget₂AddCommMonPreservesLimits : +instance forget₂AddCommMon_preservesLimits : PreservesLimits (forget₂ SemiRingCat AddCommMonCat.{u}) := - SemiRingCat.forget₂AddCommMonPreservesLimitsOfSize.{u, u} + SemiRingCat.forget₂AddCommMon_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ @@ -160,27 +160,27 @@ def forget₂MonPreservesLimitsAux : /-- The forgetful functor from semirings to monoids preserves all limits. -/ -instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Mon_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ SemiRingCat MonCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (forget₂MonPreservesLimitsAux.{v, u} F) } -instance forget₂MonPreservesLimits : PreservesLimits (forget₂ SemiRingCat MonCat.{u}) := - SemiRingCat.forget₂MonPreservesLimitsOfSize.{u, u} +instance forget₂Mon_preservesLimits : PreservesLimits (forget₂ SemiRingCat MonCat.{u}) := + SemiRingCat.forget₂Mon_preservesLimitsOfSize.{u, u} /-- The forgetful functor from semirings to types preserves all limits. -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget SemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Types.Small.limitConeIsLimit.{v, u} (F ⋙ forget _)) } -instance forgetPreservesLimits : PreservesLimits (forget SemiRingCat.{u}) := - SemiRingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget SemiRingCat.{u}) := + SemiRingCat.forget_preservesLimitsOfSize.{u, u} end SemiRingCat @@ -265,29 +265,29 @@ instance hasLimits : HasLimits CommSemiRingCat.{u} := /-- The forgetful functor from rings to semirings preserves all limits. -/ -instance forget₂SemiRingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂SemiRing_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommSemiRingCat SemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (SemiRingCat.HasLimits.limitConeIsLimit (F ⋙ forget₂ _ SemiRingCat)) } -instance forget₂SemiRingPreservesLimits : +instance forget₂SemiRing_preservesLimits : PreservesLimits (forget₂ CommSemiRingCat SemiRingCat.{u}) := - CommSemiRingCat.forget₂SemiRingPreservesLimitsOfSize.{u, u} + CommSemiRingCat.forget₂SemiRing_preservesLimitsOfSize.{u, u} /-- The forgetful functor from rings to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.) -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget CommSemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (Types.Small.limitConeIsLimit.{v, u} _) } -instance forgetPreservesLimits : PreservesLimits (forget CommSemiRingCat.{u}) := - CommSemiRingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget CommSemiRingCat.{u}) := + CommSemiRingCat.forget_preservesLimitsOfSize.{u, u} end CommSemiRingCat @@ -374,15 +374,15 @@ instance hasLimits : HasLimits RingCat.{u} := /-- The forgetful functor from rings to semirings preserves all limits. -/ -instance forget₂SemiRingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂SemiRing_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ RingCat SemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (SemiRingCat.HasLimits.limitConeIsLimit.{v, u} _) } -instance forget₂SemiRingPreservesLimits : PreservesLimits (forget₂ RingCat SemiRingCat.{u}) := - RingCat.forget₂SemiRingPreservesLimitsOfSize.{u, u} +instance forget₂SemiRing_preservesLimits : PreservesLimits (forget₂ RingCat SemiRingCat.{u}) := + RingCat.forget₂SemiRing_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ @@ -396,29 +396,29 @@ def forget₂AddCommGroupPreservesLimitsAux : /-- The forgetful functor from rings to additive commutative groups preserves all limits. -/ -instance forget₂AddCommGroupPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂AddCommGroup_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{v, v} (forget₂ RingCat.{u} AddCommGrp.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (forget₂AddCommGroupPreservesLimitsAux F) } -instance forget₂AddCommGroupPreservesLimits : +instance forget₂AddCommGroup_preservesLimits : PreservesLimits (forget₂ RingCat AddCommGrp.{u}) := - RingCat.forget₂AddCommGroupPreservesLimitsOfSize.{u, u} + RingCat.forget₂AddCommGroup_preservesLimitsOfSize.{u, u} /-- The forgetful functor from rings to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.) -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{v, v} (forget RingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (Types.Small.limitConeIsLimit.{v, u} _) } -instance forgetPreservesLimits : PreservesLimits (forget RingCat.{u}) := - RingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget RingCat.{u}) := + RingCat.forget_preservesLimitsOfSize.{u, u} end RingCat @@ -514,15 +514,15 @@ instance hasLimits : HasLimits CommRingCat.{u} := /-- The forgetful functor from commutative rings to rings preserves all limits. (That is, the underlying rings could have been computed instead as limits in the category of rings.) -/ -instance forget₂RingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂Ring_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommRingCat RingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone.{w, v} (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone.{w, v} (limitConeIsLimit.{v, u} F) (RingCat.limitConeIsLimit.{v, u} _) } -instance forget₂RingPreservesLimits : PreservesLimits (forget₂ CommRingCat RingCat.{u}) := - CommRingCat.forget₂RingPreservesLimitsOfSize.{u, u} +instance forget₂Ring_preservesLimits : PreservesLimits (forget₂ CommRingCat RingCat.{u}) := + CommRingCat.forget₂Ring_preservesLimitsOfSize.{u, u} /-- An auxiliary declaration to speed up typechecking. -/ @@ -536,28 +536,28 @@ def forget₂CommSemiRingPreservesLimitsAux : (That is, the underlying commutative semirings could have been computed instead as limits in the category of commutative semirings.) -/ -instance forget₂CommSemiRingPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget₂CommSemiRing_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget₂ CommRingCat CommSemiRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v, u} F) (forget₂CommSemiRingPreservesLimitsAux.{v, u} F) } -instance forget₂CommSemiRingPreservesLimits : +instance forget₂CommSemiRing_preservesLimits : PreservesLimits (forget₂ CommRingCat CommSemiRingCat.{u}) := - CommRingCat.forget₂CommSemiRingPreservesLimitsOfSize.{u, u} + CommRingCat.forget₂CommSemiRing_preservesLimitsOfSize.{u, u} /-- The forgetful functor from commutative rings to types preserves all limits. (That is, the underlying types could have been computed instead as limits in the category of types.) -/ -instance forgetPreservesLimitsOfSize [UnivLE.{v, u}] : +instance forget_preservesLimitsOfSize [UnivLE.{v, u}] : PreservesLimitsOfSize.{w, v} (forget CommRingCat.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone.{w, v} (limitConeIsLimit.{v, u} F) + preservesLimit_of_preserves_limit_cone.{w, v} (limitConeIsLimit.{v, u} F) (Types.Small.limitConeIsLimit.{v, u} _) } -instance forgetPreservesLimits : PreservesLimits (forget CommRingCat.{u}) := - CommRingCat.forgetPreservesLimitsOfSize.{u, u} +instance forget_preservesLimits : PreservesLimits (forget CommRingCat.{u}) := + CommRingCat.forget_preservesLimitsOfSize.{u, u} end CommRingCat diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean index 774df032bb48c..b5130066b0f40 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext/Basic.lean @@ -251,7 +251,7 @@ lemma mk₀_zero : mk₀ (0 : X ⟶ Y) = 0 := by section -attribute [local instance] preservesBinaryBiproductsOfPreservesBiproducts in +attribute [local instance] preservesBinaryBiproducts_of_preservesBiproducts in lemma biprod_ext {X₁ X₂ : C} {α β : Ext (X₁ ⊞ X₂) Y n} (h₁ : (mk₀ biprod.inl).comp α (zero_add n) = (mk₀ biprod.inl).comp β (zero_add n)) (h₂ : (mk₀ biprod.inr).comp α (zero_add n) = (mk₀ biprod.inr).comp β (zero_add n)) : diff --git a/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean b/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean index cb8b60d15b0e0..48e36902658d6 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplexBiprod.lean @@ -36,7 +36,7 @@ instance (i : ι) : HasColimit ((pair K L) ⋙ (eval C c i)) := by instance : HasBinaryBiproduct K L := HasBinaryBiproduct.of_hasBinaryProduct _ _ instance (i : ι) : PreservesBinaryBiproduct K L (eval C c i) := - preservesBinaryBiproductOfPreservesBinaryProduct _ + preservesBinaryBiproduct_of_preservesBinaryProduct _ /-- The canonical isomorphism `(K ⊞ L).X i ≅ (K.X i) ⊞ (L.X i)`. -/ noncomputable def biprodXIso (i : ι) : (K ⊞ L).X i ≅ (K.X i) ⊞ (L.X i) := diff --git a/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean b/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean index 0e2a446e27264..078140361610a 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplexLimits.lean @@ -79,7 +79,7 @@ noncomputable def isLimitConeOfHasLimitEval : IsLimit (coneOfHasLimitEval F) := instance : HasLimit F := ⟨⟨⟨_, isLimitConeOfHasLimitEval F⟩⟩⟩ noncomputable instance (n : ι) : PreservesLimit F (eval C c n) := - preservesLimitOfPreservesLimitCone (isLimitConeOfHasLimitEval F) (limit.isLimit _) + preservesLimit_of_preserves_limit_cone (isLimitConeOfHasLimitEval F) (limit.isLimit _) end @@ -155,7 +155,7 @@ noncomputable def isColimitCoconeOfHasColimitEval : IsColimit (coconeOfHasColimi instance : HasColimit F := ⟨⟨⟨_, isColimitCoconeOfHasColimitEval F⟩⟩⟩ noncomputable instance (n : ι) : PreservesColimit F (eval C c n) := - preservesColimitOfPreservesColimitCocone (isColimitCoconeOfHasColimitEval F) + preservesColimit_of_preserves_colimit_cocone (isColimitCoconeOfHasColimitEval F) (colimit.isColimit _) end @@ -178,39 +178,39 @@ instance [HasFiniteColimits C] {K L : HomologicalComplex C c} (φ : K ⟶ L) [Ep /-- A functor `D ⥤ HomologicalComplex C c` preserves limits of shape `J` if for any `i`, `G ⋙ eval C c i` does. -/ -def preservesLimitsOfShapeOfEval {D : Type*} [Category D] +lemma preservesLimitsOfShape_of_eval {D : Type*} [Category D] (G : D ⥤ HomologicalComplex C c) (_ : ∀ (i : ι), PreservesLimitsOfShape J (G ⋙ eval C c i)) : PreservesLimitsOfShape J G := - ⟨fun {_} => ⟨fun hs ↦ isLimitOfEval _ _ - (fun i => isLimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩ + ⟨fun {_} => ⟨fun hs ↦ ⟨isLimitOfEval _ _ + (fun i => isLimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩⟩ /-- A functor `D ⥤ HomologicalComplex C c` preserves colimits of shape `J` if for any `i`, `G ⋙ eval C c i` does. -/ -def preservesColimitsOfShapeOfEval {D : Type*} [Category D] +lemma preservesColimitsOfShape_of_eval {D : Type*} [Category D] (G : D ⥤ HomologicalComplex C c) (_ : ∀ (i : ι), PreservesColimitsOfShape J (G ⋙ eval C c i)) : PreservesColimitsOfShape J G := - ⟨fun {_} => ⟨fun hs ↦ isColimitOfEval _ _ - (fun i => isColimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩ + ⟨fun {_} => ⟨fun hs ↦ ⟨isColimitOfEval _ _ + (fun i => isColimitOfPreserves (G ⋙ eval C c i) hs)⟩⟩⟩ section variable [HasZeroObject C] [DecidableEq ι] (i : ι) noncomputable instance : PreservesLimitsOfShape J (single C c i) := - preservesLimitsOfShapeOfEval _ (fun j => by + preservesLimitsOfShape_of_eval _ (fun j => by by_cases h : j = i · subst h - exact preservesLimitsOfShapeOfNatIso (singleCompEvalIsoSelf C c j).symm - · exact Functor.preservesLimitsOfShapeOfIsZero _ (isZero_single_comp_eval C c _ _ h) _) + exact preservesLimitsOfShape_of_natIso (singleCompEvalIsoSelf C c j).symm + · exact Functor.preservesLimitsOfShape_of_isZero _ (isZero_single_comp_eval C c _ _ h) _) noncomputable instance : PreservesColimitsOfShape J (single C c i) := - preservesColimitsOfShapeOfEval _ (fun j => by + preservesColimitsOfShape_of_eval _ (fun j => by by_cases h : j = i · subst h - exact preservesColimitsOfShapeOfNatIso (singleCompEvalIsoSelf C c j).symm - · exact Functor.preservesColimitsOfShapeOfIsZero _ (isZero_single_comp_eval C c _ _ h) _) + exact preservesColimitsOfShape_of_natIso (singleCompEvalIsoSelf C c j).symm + · exact Functor.preservesColimitsOfShape_of_isZero _ (isZero_single_comp_eval C c _ _ h) _) noncomputable instance : PreservesFiniteLimits (single C c i) := ⟨by intros; infer_instance⟩ diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean index 51fd7a82b8752..dfd6dc033079a 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -139,10 +139,10 @@ variable (D : SnakeInput C) lemma δ_apply (x₃ : D.L₀.X₃) (x₂ : D.L₁.X₂) (x₁ : D.L₂.X₁) (h₂ : D.L₁.g x₂ = D.v₀₁.τ₃ x₃) (h₁ : D.L₂.f x₁ = D.v₁₂.τ₂ x₂) : D.δ x₃ = D.v₂₃.τ₁ x₁ := by - have := (forget₂ C Ab).preservesFiniteLimitsOfPreservesHomology + have := (forget₂ C Ab).preservesFiniteLimits_of_preservesHomology have : PreservesFiniteLimits (forget C) := by have : forget₂ C Ab ⋙ forget Ab = forget C := HasForget₂.forget_comp - simpa only [← this] using compPreservesFiniteLimits _ _ + simpa only [← this] using comp_preservesFiniteLimits _ _ have eq := congr_fun ((forget C).congr_map D.snd_δ) (Limits.Concrete.pullbackMk D.L₁.g D.v₀₁.τ₃ x₂ x₃ h₂) have eq₁ := Concrete.pullbackMk_fst D.L₁.g D.v₀₁.τ₃ x₂ x₃ h₂ diff --git a/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean b/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean index d61079ee62930..06141d6538e48 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean @@ -19,9 +19,9 @@ also preserves finite limits and finite colimits. Let `F : C ⥤ D` be an additive functor: -- `Functor.preservesFiniteLimitsOfPreservesHomology`: if `F` preserves homology, then `F` preserves - finite limits. -- `Functor.preservesFiniteColimitsOfPreservesHomology`: if `F` preserves homology, then `F` +- `Functor.preservesFiniteLimits_of_preservesHomology`: if `F` preserves homology, + then `F` preserves finite limits. +- `Functor.preservesFiniteColimits_of_preservesHomology`: if `F` preserves homology, then `F` preserves finite colimits. If we further assume that `C` and `D` are abelian categories, then we have: @@ -63,24 +63,24 @@ variable {C D : Type*} [Category C] [Category D] [Preadditive C] [Preadditive D] (F : C ⥤ D) [F.Additive] [F.PreservesHomology] [HasZeroObject C] /-- An additive functor which preserves homology preserves finite limits. -/ -noncomputable def preservesFiniteLimitsOfPreservesHomology +lemma preservesFiniteLimits_of_preservesHomology [HasFiniteProducts C] [HasKernels C] : PreservesFiniteLimits F := by have := fun {X Y : C} (f : X ⟶ Y) ↦ PreservesHomology.preservesKernel F f have : HasBinaryBiproducts C := HasBinaryBiproducts.of_hasBinaryProducts have : HasEqualizers C := Preadditive.hasEqualizers_of_hasKernels have : HasZeroObject D := ⟨F.obj 0, by rw [IsZero.iff_id_eq_zero, ← F.map_id, id_zero, F.map_zero]⟩ - exact preservesFiniteLimitsOfPreservesKernels F + exact preservesFiniteLimits_of_preservesKernels F /-- An additive which preserves homology preserves finite colimits. -/ -noncomputable def preservesFiniteColimitsOfPreservesHomology +lemma preservesFiniteColimits_of_preservesHomology [HasFiniteCoproducts C] [HasCokernels C] : PreservesFiniteColimits F := by have := fun {X Y : C} (f : X ⟶ Y) ↦ PreservesHomology.preservesCokernel F f have : HasBinaryBiproducts C := HasBinaryBiproducts.of_hasBinaryCoproducts have : HasCoequalizers C := Preadditive.hasCoequalizers_of_hasCokernels have : HasZeroObject D := ⟨F.obj 0, by rw [IsZero.iff_id_eq_zero, ← F.map_id, id_zero, F.map_zero]⟩ - exact preservesFiniteColimitsOfPreservesCokernels F + exact preservesFiniteColimits_of_preservesCokernels F end @@ -112,8 +112,8 @@ lemma preservesFiniteLimits_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).Exact ∧ Mono (F.map S.f), ∀ (S : ShortComplex C), S.Exact ∧ Mono S.f → (S.map F).Exact ∧ Mono (F.map S.f), - ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty <| PreservesLimit (parallelPair f 0) F, - Nonempty <| PreservesFiniteLimits F + ∀ ⦃X Y : C⦄ (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F, + PreservesFiniteLimits F ] := by tfae_have 1 → 2 | hF, S, ⟨hS, hf⟩ => by @@ -130,7 +130,7 @@ lemma preservesFiniteLimits_tfae : List.TFAE tfae_have 2 → 3 | hF, X, Y, f => by - refine ⟨preservesLimitOfPreservesLimitCone (kernelIsKernel f) ?_⟩ + refine preservesLimit_of_preserves_limit_cone (kernelIsKernel f) ?_ apply (KernelFork.isLimitMapConeEquiv _ F).2 let S := ShortComplex.mk _ _ (kernel.condition f) let hS := hF S ⟨exact_kernel f, inferInstance⟩ @@ -139,8 +139,7 @@ lemma preservesFiniteLimits_tfae : List.TFAE tfae_have 3 → 4 | hF => by - have := fun X Y (f : X ⟶ Y) ↦ (hF f).some - exact ⟨preservesFiniteLimitsOfPreservesKernels F⟩ + exact preservesFiniteLimits_of_preservesKernels F tfae_have 4 → 1 | ⟨_⟩, S, hS => @@ -171,8 +170,8 @@ lemma preservesFiniteColimits_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).Exact ∧ Epi (F.map S.g), ∀ (S : ShortComplex C), S.Exact ∧ Epi S.g → (S.map F).Exact ∧ Epi (F.map S.g), - ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty <| PreservesColimit (parallelPair f 0) F, - Nonempty <| PreservesFiniteColimits F + ∀ ⦃X Y : C⦄ (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F, + PreservesFiniteColimits F ] := by tfae_have 1 → 2 | hF, S, ⟨hS, hf⟩ => by @@ -189,7 +188,7 @@ lemma preservesFiniteColimits_tfae : List.TFAE tfae_have 2 → 3 | hF, X, Y, f => by - refine ⟨preservesColimitOfPreservesColimitCocone (cokernelIsCokernel f) ?_⟩ + refine preservesColimit_of_preserves_colimit_cocone (cokernelIsCokernel f) ?_ apply (CokernelCofork.isColimitMapCoconeEquiv _ F).2 let S := ShortComplex.mk _ _ (cokernel.condition f) let hS := hF S ⟨exact_cokernel f, inferInstance⟩ @@ -198,8 +197,7 @@ lemma preservesFiniteColimits_tfae : List.TFAE tfae_have 3 → 4 | hF => by - have := fun X Y (f : X ⟶ Y) ↦ (hF f).some - exact ⟨preservesFiniteColimitsOfPreservesCokernels F⟩ + exact preservesFiniteColimits_of_preservesCokernels F tfae_have 4 → 1 | ⟨_⟩, S, hS => (S.map F).exact_and_epi_g_iff_g_is_cokernel |>.2 @@ -219,18 +217,18 @@ lemma exact_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).ShortExact, ∀ (S : ShortComplex C), S.Exact → (S.map F).Exact, - Nonempty (PreservesHomology F), - Nonempty (PreservesFiniteLimits F) ∧ Nonempty (PreservesFiniteColimits F) + PreservesHomology F, + PreservesFiniteLimits F ∧ PreservesFiniteColimits F ] := by tfae_have 1 → 3 | hF => by refine ⟨fun {X Y} f ↦ ?_, fun {X Y} f ↦ ?_⟩ · have h := (preservesFiniteLimits_tfae F |>.out 0 2 |>.1 fun S hS ↦ And.intro (hF S hS).exact (hF S hS).mono_f) - exact h f |>.some + exact h f · have h := (preservesFiniteColimits_tfae F |>.out 0 2 |>.1 fun S hS ↦ And.intro (hF S hS).exact (hF S hS).epi_g) - exact h f |>.some + exact h f tfae_have 2 → 1 | hF, S, hS => by @@ -241,11 +239,11 @@ lemma exact_tfae : List.TFAE exact ⟨hF S hS.exact⟩ tfae_have 3 → 4 - | ⟨h⟩ => ⟨⟨preservesFiniteLimitsOfPreservesHomology F⟩, - ⟨preservesFiniteColimitsOfPreservesHomology F⟩⟩ + | h => ⟨preservesFiniteLimits_of_preservesHomology F, + preservesFiniteColimits_of_preservesHomology F⟩ tfae_have 4 → 2 - | ⟨⟨h1⟩, ⟨h2⟩⟩, _, h => h.map F + | ⟨h1, h2⟩, _, h => h.map F tfae_finish diff --git a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean index ba94784012560..1f94e0fc71c5e 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Limits.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Limits.lean @@ -14,8 +14,6 @@ import Mathlib.CategoryTheory.Limits.Preserves.Finite In this file, it is shown if a category `C` with zero morphisms has limits of a certain shape `J`, then it is also the case of the category `ShortComplex C`. -TODO (@rioujoel): Do the same for colimits. - -/ namespace CategoryTheory @@ -95,13 +93,13 @@ noncomputable def isLimitLimitCone : IsLimit (limitCone F) := instance hasLimit_of_hasLimitπ : HasLimit F := ⟨⟨⟨_, isLimitLimitCone _⟩⟩⟩ noncomputable instance : PreservesLimit F π₁ := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (isLimitπ₁MapConeLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (isLimitπ₁MapConeLimitCone F) noncomputable instance : PreservesLimit F π₂ := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (isLimitπ₂MapConeLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (isLimitπ₂MapConeLimitCone F) noncomputable instance : PreservesLimit F π₃ := - preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (isLimitπ₃MapConeLimitCone F) + preservesLimit_of_preserves_limit_cone (isLimitLimitCone F) (isLimitπ₃MapConeLimitCone F) end @@ -231,15 +229,15 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) := instance hasColimit_of_hasColimitπ : HasColimit F := ⟨⟨⟨_, isColimitColimitCocone _⟩⟩⟩ noncomputable instance : PreservesColimit F π₁ := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (isColimitπ₁MapCoconeColimitCocone F) noncomputable instance : PreservesColimit F π₂ := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (isColimitπ₂MapCoconeColimitCocone F) noncomputable instance : PreservesColimit F π₃ := - preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) + preservesColimit_of_preserves_colimit_cocone (isColimitColimitCocone F) (isColimitπ₃MapCoconeColimitCocone F) end diff --git a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean index 5f870fcf8b0c0..eacf12a85c953 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/PreservesHomology.lean @@ -34,7 +34,7 @@ namespace Functor variable (F : C ⥤ D) /-- A functor preserves homology when it preserves both kernels and cokernels. -/ -class PreservesHomology (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesHomology (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where /-- the functor preserves kernels -/ preservesKernels ⦃X Y : C⦄ (f : X ⟶ Y) : PreservesLimit (parallelPair f 0) F := by infer_instance @@ -45,12 +45,12 @@ class PreservesHomology (F : C ⥤ D) [PreservesZeroMorphisms F] where variable [PreservesZeroMorphisms F] /-- A functor which preserves homology preserves kernels. -/ -def PreservesHomology.preservesKernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : +lemma PreservesHomology.preservesKernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : PreservesLimit (parallelPair f 0) F := PreservesHomology.preservesKernels _ /-- A functor which preserves homology preserves cokernels. -/ -def PreservesHomology.preservesCokernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : +lemma PreservesHomology.preservesCokernel [F.PreservesHomology] {X Y : C} (f : X ⟶ Y) : PreservesColimit (parallelPair f 0) F := PreservesHomology.preservesCokernels _ @@ -70,7 +70,7 @@ variable (h : S.LeftHomologyData) (F : C ⥤ D) /-- A left homology data `h` of a short complex `S` is preserved by a functor `F` is `F` preserves the kernel of `S.g : S.X₂ ⟶ S.X₃` and the cokernel of `h.f' : S.X₁ ⟶ h.K`. -/ -class IsPreservedBy [F.PreservesZeroMorphisms] where +class IsPreservedBy [F.PreservesZeroMorphisms] : Prop where /-- the functor preserves the kernel of `S.g : S.X₂ ⟶ S.X₃`. -/ g : PreservesLimit (parallelPair S.g 0) F /-- the functor preserves the cokernel of `h.f' : S.X₁ ⟶ h.K`. -/ @@ -78,21 +78,22 @@ class IsPreservedBy [F.PreservesZeroMorphisms] where variable [F.PreservesZeroMorphisms] -noncomputable instance isPreservedByOfPreservesHomology [F.PreservesHomology] : +noncomputable instance isPreservedBy_of_preservesHomology [F.PreservesHomology] : h.IsPreservedBy F where g := Functor.PreservesHomology.preservesKernel _ _ f' := Functor.PreservesHomology.preservesCokernel _ _ variable [h.IsPreservedBy F] +include h in /-- When a left homology data is preserved by a functor `F`, this functor preserves the kernel of `S.g : S.X₂ ⟶ S.X₃`. -/ -def IsPreservedBy.hg : PreservesLimit (parallelPair S.g 0) F := +lemma IsPreservedBy.hg : PreservesLimit (parallelPair S.g 0) F := @IsPreservedBy.g _ _ _ _ _ _ _ h F _ _ /-- When a left homology data `h` is preserved by a functor `F`, this functor preserves the cokernel of `h.f' : S.X₁ ⟶ h.K`. -/ -def IsPreservedBy.hf' : PreservesColimit (parallelPair h.f' 0) F := IsPreservedBy.f' +lemma IsPreservedBy.hf' : PreservesColimit (parallelPair h.f' 0) F := IsPreservedBy.f' /-- When a left homology data `h` of a short complex `S` is preserved by a functor `F`, this is the induced left homology data `h.map F` for the short complex `S.map F`. -/ @@ -149,7 +150,7 @@ variable (h : S.RightHomologyData) (F : C ⥤ D) /-- A right homology data `h` of a short complex `S` is preserved by a functor `F` is `F` preserves the cokernel of `S.f : S.X₁ ⟶ S.X₂` and the kernel of `h.g' : h.Q ⟶ S.X₃`. -/ -class IsPreservedBy [F.PreservesZeroMorphisms] where +class IsPreservedBy [F.PreservesZeroMorphisms] : Prop where /-- the functor preserves the cokernel of `S.f : S.X₁ ⟶ S.X₂`. -/ f : PreservesColimit (parallelPair S.f 0) F /-- the functor preserves the kernel of `h.g' : h.Q ⟶ S.X₃`. -/ @@ -157,21 +158,22 @@ class IsPreservedBy [F.PreservesZeroMorphisms] where variable [F.PreservesZeroMorphisms] -noncomputable instance isPreservedByOfPreservesHomology [F.PreservesHomology] : +noncomputable instance isPreservedBy_of_preservesHomology [F.PreservesHomology] : h.IsPreservedBy F where f := Functor.PreservesHomology.preservesCokernel F _ g' := Functor.PreservesHomology.preservesKernel F _ variable [h.IsPreservedBy F] +include h in /-- When a right homology data is preserved by a functor `F`, this functor preserves the cokernel of `S.f : S.X₁ ⟶ S.X₂`. -/ -def IsPreservedBy.hf : PreservesColimit (parallelPair S.f 0) F := +lemma IsPreservedBy.hf : PreservesColimit (parallelPair S.f 0) F := @IsPreservedBy.f _ _ _ _ _ _ _ h F _ _ /-- When a right homology data `h` is preserved by a functor `F`, this functor preserves the kernel of `h.g' : h.Q ⟶ S.X₃`. -/ -def IsPreservedBy.hg' : PreservesLimit (parallelPair h.g' 0) F := +lemma IsPreservedBy.hg' : PreservesLimit (parallelPair h.g' 0) F := @IsPreservedBy.g' _ _ _ _ _ _ _ h F _ _ /-- When a right homology data `h` of a short complex `S` is preserved by a functor `F`, @@ -256,27 +258,27 @@ variable (F : C ⥤ D) [PreservesZeroMorphisms F] (S : ShortComplex C) {S₁ S /-- A functor preserves the left homology of a short complex `S` if it preserves all the left homology data of `S`. -/ -class PreservesLeftHomologyOf where +class PreservesLeftHomologyOf : Prop where /-- the functor preserves all the left homology data of the short complex -/ isPreservedBy : ∀ (h : S.LeftHomologyData), h.IsPreservedBy F /-- A functor preserves the right homology of a short complex `S` if it preserves all the right homology data of `S`. -/ -class PreservesRightHomologyOf where +class PreservesRightHomologyOf : Prop where /-- the functor preserves all the right homology data of the short complex -/ isPreservedBy : ∀ (h : S.RightHomologyData), h.IsPreservedBy F -noncomputable instance PreservesHomology.preservesLeftHomologyOf [F.PreservesHomology] : +instance PreservesHomology.preservesLeftHomologyOf [F.PreservesHomology] : F.PreservesLeftHomologyOf S := ⟨inferInstance⟩ -noncomputable instance PreservesHomology.preservesRightHomologyOf [F.PreservesHomology] : +instance PreservesHomology.preservesRightHomologyOf [F.PreservesHomology] : F.PreservesRightHomologyOf S := ⟨inferInstance⟩ variable {S} /-- If a functor preserves a certain left homology data of a short complex `S`, then it preserves the left homology of `S`. -/ -def PreservesLeftHomologyOf.mk' (h : S.LeftHomologyData) [h.IsPreservedBy F] : +lemma PreservesLeftHomologyOf.mk' (h : S.LeftHomologyData) [h.IsPreservedBy F] : F.PreservesLeftHomologyOf S where isPreservedBy h' := { g := ShortComplex.LeftHomologyData.IsPreservedBy.hg h F @@ -285,11 +287,11 @@ def PreservesLeftHomologyOf.mk' (h : S.LeftHomologyData) [h.IsPreservedBy F] : let e : parallelPair h.f' 0 ≅ parallelPair h'.f' 0 := parallelPair.ext (Iso.refl _) (ShortComplex.cyclesMapIso' (Iso.refl S) h h') (by simp) (by simp) - exact preservesColimitOfIsoDiagram F e } + exact preservesColimit_of_iso_diagram F e } /-- If a functor preserves a certain right homology data of a short complex `S`, then it preserves the right homology of `S`. -/ -def PreservesRightHomologyOf.mk' (h : S.RightHomologyData) [h.IsPreservedBy F] : +lemma PreservesRightHomologyOf.mk' (h : S.RightHomologyData) [h.IsPreservedBy F] : F.PreservesRightHomologyOf S where isPreservedBy h' := { f := ShortComplex.RightHomologyData.IsPreservedBy.hf h F @@ -298,7 +300,7 @@ def PreservesRightHomologyOf.mk' (h : S.RightHomologyData) [h.IsPreservedBy F] : let e : parallelPair h.g' 0 ≅ parallelPair h'.g' 0 := parallelPair.ext (ShortComplex.opcyclesMapIso' (Iso.refl S) h h') (Iso.refl _) (by simp) (by simp) - exact preservesLimitOfIsoDiagram F e } + exact preservesLimit_of_iso_diagram F e } end Functor @@ -307,11 +309,11 @@ namespace ShortComplex variable {S : ShortComplex C} (h₁ : S.LeftHomologyData) (h₂ : S.RightHomologyData) (F : C ⥤ D) [F.PreservesZeroMorphisms] -instance LeftHomologyData.isPreservedByOfPreserves [F.PreservesLeftHomologyOf S] : +instance LeftHomologyData.isPreservedBy_of_preserves [F.PreservesLeftHomologyOf S] : h₁.IsPreservedBy F := Functor.PreservesLeftHomologyOf.isPreservedBy _ -instance RightHomologyData.isPreservedByOfPreserves [F.PreservesRightHomologyOf S] : +instance RightHomologyData.isPreservedBy_of_preserves [F.PreservesRightHomologyOf S] : h₂.IsPreservedBy F := Functor.PreservesRightHomologyOf.isPreservedBy _ @@ -805,25 +807,25 @@ variable (F : C ⥤ D) [F.PreservesZeroMorphisms] (S : ShortComplex C) /-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved by a functor `F`, then `F` preserves the left homology of `S`. -/ -noncomputable def preservesLeftHomologyOfZerof (hf : S.f = 0) +lemma preservesLeftHomology_of_zero_f (hf : S.f = 0) [PreservesLimit (parallelPair S.g 0) F] : F.PreservesLeftHomologyOf S := ⟨fun h => { g := by infer_instance - f' := Limits.preservesCokernelZero' _ _ + f' := Limits.preservesCokernel_zero' _ _ (by rw [← cancel_mono h.i, h.f'_i, zero_comp, hf]) }⟩ /-- If a short complex `S` is such that `S.g = 0` and that the cokernel of `S.f` is preserved by a functor `F`, then `F` preserves the right homology of `S`. -/ -noncomputable def preservesRightHomologyOfZerog (hg : S.g = 0) +lemma preservesRightHomology_of_zero_g (hg : S.g = 0) [PreservesColimit (parallelPair S.f 0) F] : F.PreservesRightHomologyOf S := ⟨fun h => { f := by infer_instance - g' := Limits.preservesKernelZero' _ _ + g' := Limits.preservesKernel_zero' _ _ (by rw [← cancel_epi h.p, h.p_g', comp_zero, hg]) }⟩ /-- If a short complex `S` is such that `S.g = 0` and that the cokernel of `S.f` is preserved by a functor `F`, then `F` preserves the left homology of `S`. -/ -noncomputable def preservesLeftHomologyOfZerog (hg : S.g = 0) +lemma preservesLeftHomology_of_zero_g (hg : S.g = 0) [PreservesColimit (parallelPair S.f 0) F] : F.PreservesLeftHomologyOf S := ⟨fun h => { g := by @@ -833,11 +835,11 @@ noncomputable def preservesLeftHomologyOfZerog (hg : S.g = 0) have := h.isIso_i hg let e : parallelPair h.f' 0 ≅ parallelPair S.f 0 := parallelPair.ext (Iso.refl _) (asIso h.i) (by aesop_cat) (by aesop_cat) - exact Limits.preservesColimitOfIsoDiagram F e.symm}⟩ + exact Limits.preservesColimit_of_iso_diagram F e.symm}⟩ /-- If a short complex `S` is such that `S.f = 0` and that the kernel of `S.g` is preserved by a functor `F`, then `F` preserves the right homology of `S`. -/ -noncomputable def preservesRightHomologyOfZerof (hf : S.f = 0) +lemma preservesRightHomology_of_zero_f (hf : S.f = 0) [PreservesLimit (parallelPair S.g 0) F] : F.PreservesRightHomologyOf S := ⟨fun h => { f := by @@ -847,7 +849,7 @@ noncomputable def preservesRightHomologyOfZerof (hf : S.f = 0) have := h.isIso_p hf let e : parallelPair S.g 0 ≅ parallelPair h.g' 0 := parallelPair.ext (asIso h.p) (Iso.refl _) (by aesop_cat) (by aesop_cat) - exact Limits.preservesLimitOfIsoDiagram F e }⟩ + exact Limits.preservesLimit_of_iso_diagram F e }⟩ end Functor diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 956a2db40bc22..a83addf8247e5 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -194,7 +194,7 @@ noncomputable instance Γ_preservesLimits : PreservesLimits Γ.{u}.rightOp := in noncomputable instance forgetToScheme_preservesLimits : PreservesLimits forgetToScheme := by apply (config := { allowSynthFailures := true }) - @preservesLimitsOfNatIso _ _ _ _ _ _ + @preservesLimits_of_natIso _ _ _ _ _ _ (isoWhiskerRight equivCommRingCat.unitIso forgetToScheme).symm change PreservesLimits (equivCommRingCat.functor ⋙ Scheme.Spec) infer_instance diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 3b5e437273378..6c8faa057ea86 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -483,10 +483,10 @@ alias ΓSpec.adjunction_unit_map_basicOpen := Scheme.toSpecΓ_preimage_basicOpen /-- Spec preserves limits. -/ instance : Limits.PreservesLimits Spec.toLocallyRingedSpace := - ΓSpec.locallyRingedSpaceAdjunction.rightAdjointPreservesLimits + ΓSpec.locallyRingedSpaceAdjunction.rightAdjoint_preservesLimits instance Spec.preservesLimits : Limits.PreservesLimits Scheme.Spec := - ΓSpec.adjunction.rightAdjointPreservesLimits + ΓSpec.adjunction.rightAdjoint_preservesLimits /-- The functor `Spec.toLocallyRingedSpace : CommRingCatᵒᵖ ⥤ LocallyRingedSpace` is fully faithful.-/ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index d94a986f233f9..6c634b3aadaac 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -197,8 +197,7 @@ instance : CreatesColimitsOfShape (Discrete ι) Scheme.forgetToLocallyRingedSpac intro K exact createsColimitOfIsoDiagram _ (Discrete.natIsoFunctor (F := K)).symm -noncomputable -instance : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop := +instance : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop.{u} := inferInstanceAs (PreservesColimitsOfShape (Discrete ι) (Scheme.forgetToLocallyRingedSpace ⋙ LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget CommRingCat)) @@ -208,12 +207,15 @@ instance : HasCoproducts.{u} Scheme.{u} := instance : HasCoproducts.{0} Scheme.{u} := has_smallest_coproducts_of_hasCoproducts noncomputable -instance {ι : Type} : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop := - preservesColimitsOfShapeOfEquiv (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ +instance {ι : Type} : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToTop.{u} := + preservesColimitsOfShape_of_equiv + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ noncomputable -instance {ι : Type} : PreservesColimitsOfShape (Discrete ι) Scheme.forgetToLocallyRingedSpace := - preservesColimitsOfShapeOfEquiv (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ +instance {ι : Type} : + PreservesColimitsOfShape (Discrete ι) Scheme.forgetToLocallyRingedSpace.{u} := + preservesColimitsOfShape_of_equiv + (Discrete.equivalence Equiv.ulift : Discrete (ULift.{u} ι) ≌ _) _ /-- (Implementation Detail) Coproduct of schemes is isomorphic to the disjoint union. -/ noncomputable @@ -461,16 +463,16 @@ instance (R S : CommRingCatᵒᵖ) : IsIso (coprodComparison Scheme.Spec R S) := noncomputable instance : PreservesColimitsOfShape (Discrete WalkingPair) Scheme.Spec := ⟨fun {_} ↦ - have (X Y : CommRingCatᵒᵖ) := PreservesColimitPair.ofIsoCoprodComparison Scheme.Spec X Y - preservesColimitOfIsoDiagram _ (diagramIsoPair _).symm⟩ + have (X Y : CommRingCatᵒᵖ) := PreservesColimitPair.of_iso_coprod_comparison Scheme.Spec X Y + preservesColimit_of_iso_diagram _ (diagramIsoPair _).symm⟩ noncomputable instance : PreservesColimitsOfShape (Discrete PEmpty.{1}) Scheme.Spec := by have : IsEmpty (Scheme.Spec.obj (⊥_ CommRingCatᵒᵖ)) := @Function.isEmpty _ _ spec_punit_isEmpty (Scheme.Spec.mapIso (initialIsoIsInitial (initialOpOfTerminal CommRingCat.punitIsTerminal))).hom.base - have := preservesInitialOfIso Scheme.Spec (asIso (initial.to _)) - exact preservesColimitsOfShapePemptyOfPreservesInitial _ + have := preservesInitial_of_iso Scheme.Spec (asIso (initial.to _)) + exact preservesColimitsOfShape_pempty_of_preservesInitial _ noncomputable instance {J} [Fintype J] : PreservesColimitsOfShape (Discrete J) Scheme.Spec := @@ -479,7 +481,7 @@ instance {J} [Fintype J] : PreservesColimitsOfShape (Discrete J) Scheme.Spec := noncomputable instance {J : Type*} [Finite J] : PreservesColimitsOfShape (Discrete J) Scheme.Spec := letI := (nonempty_fintype J).some - preservesColimitsOfShapeOfEquiv (Discrete.equivalence (Fintype.equivFin _).symm) _ + preservesColimitsOfShape_of_equiv (Discrete.equivalence (Fintype.equivFin _).symm) _ /-- The canonical map `∐ Spec Rᵢ ⟶ Spec (Π Rᵢ)`. This is an isomorphism when the product is finite. -/ diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index de7842550570e..834a151635fb5 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -454,11 +454,11 @@ instance forgetCreatesPullbackOfRight : CreatesLimit (cospan g f) forget := (PresheafedSpace.IsOpenImmersion.toScheme Y (pullback.fst g.toLRSHom f.toLRSHom).1) (eqToIso (by simp) ≪≫ HasLimit.isoOfNatIso (diagramIsoCospan _).symm) -instance forgetPreservesOfLeft : PreservesLimit (cospan f g) forget := - CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit _ _ +instance forget_preservesOfLeft : PreservesLimit (cospan f g) forget := + CategoryTheory.preservesLimit_of_createsLimit_and_hasLimit _ _ -instance forgetPreservesOfRight : PreservesLimit (cospan g f) forget := - preservesPullbackSymmetry _ _ _ +instance forget_preservesOfRight : PreservesLimit (cospan g f) forget := + preservesPullback_symmetry _ _ _ instance hasPullback_of_left : HasPullback f g := hasLimit_of_created (cospan f g) forget @@ -487,17 +487,17 @@ instance pullback_to_base [IsOpenImmersion g] : -- provided but still class inference fail to find this exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ _ -instance forgetToTopPreservesOfLeft : PreservesLimit (cospan f g) Scheme.forgetToTop := by +instance forgetToTop_preserves_of_left : PreservesLimit (cospan f g) Scheme.forgetToTop := by delta Scheme.forgetToTop - refine @Limits.compPreservesLimit _ _ _ _ _ _ (K := cospan f g) _ _ (F := forget) + refine @Limits.comp_preservesLimit _ _ _ _ _ _ (K := cospan f g) _ _ (F := forget) (G := LocallyRingedSpace.forgetToTop) ?_ ?_ · infer_instance - refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ + refine @preservesLimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ dsimp [LocallyRingedSpace.forgetToTop] infer_instance -instance forgetToTopPreservesOfRight : PreservesLimit (cospan g f) Scheme.forgetToTop := - preservesPullbackSymmetry _ _ _ +instance forgetToTop_preserves_of_right : PreservesLimit (cospan g f) Scheme.forgetToTop := + preservesPullback_symmetry _ _ _ theorem range_pullback_snd_of_left : Set.range (pullback.snd f g).base = (g ⁻¹ᵁ f.opensRange).1 := by diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean index e6ebf8e4a8192..2877b1ad5f01d 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/Product.lean @@ -107,9 +107,9 @@ instance : IsIso (piTopToPiCone X) := Limits.Cones.cone_iso_of_hom_iso (piTopToPiCone X) /-- The fundamental groupoid functor preserves products -/ -def preservesProduct : Limits.PreservesLimit (Discrete.functor X) π := by +lemma preservesProduct : Limits.PreservesLimit (Discrete.functor X) π := by -- Porting note: check universe parameters here - apply Limits.preservesLimitOfPreservesLimitCone (TopCat.piFanIsLimit.{u,u} X) + apply Limits.preservesLimit_of_preserves_limit_cone (TopCat.piFanIsLimit.{u,u} X) apply (Limits.IsLimit.ofConeEquiv (coneDiscreteComp X)).toFun simp only [coneDiscreteComp_obj_mapCone] apply Limits.IsLimit.ofIsoLimit _ (asIso (piTopToPiCone X)).symm diff --git a/Mathlib/CategoryTheory/Abelian/Exact.lean b/Mathlib/CategoryTheory/Abelian/Exact.lean index 04967ba6b7f80..badc3ad7a0d10 100644 --- a/Mathlib/CategoryTheory/Abelian/Exact.lean +++ b/Mathlib/CategoryTheory/Abelian/Exact.lean @@ -244,10 +244,10 @@ theorem preservesEpimorphisms_of_map_exact : L.PreservesEpimorphisms where exact ShortComplex.isoMk (Iso.refl _) (Iso.refl _) (Iso.refl _) /-- A functor which preserves the exactness of short complexes preserves homology. -/ -def preservesHomologyOfMapExact : L.PreservesHomology where +lemma preservesHomology_of_map_exact : L.PreservesHomology where preservesCokernels X Y f := by have := preservesEpimorphisms_of_map_exact _ hL - apply preservesColimitOfPreservesColimitCocone (cokernelIsCokernel f) + apply preservesColimit_of_preserves_colimit_cocone (cokernelIsCokernel f) apply (CokernelCofork.isColimitMapCoconeEquiv _ L).2 have : Epi ((ShortComplex.mk _ _ (cokernel.condition f)).map L).g := by dsimp @@ -256,7 +256,7 @@ def preservesHomologyOfMapExact : L.PreservesHomology where (ShortComplex.exact_of_g_is_cokernel _ (cokernelIsCokernel f))).gIsCokernel preservesKernels X Y f := by have := preservesMonomorphisms_of_map_exact _ hL - apply preservesLimitOfPreservesLimitCone (kernelIsKernel f) + apply preservesLimit_of_preserves_limit_cone (kernelIsKernel f) apply (KernelFork.isLimitMapConeEquiv _ L).2 have : Mono ((ShortComplex.mk _ _ (kernel.condition f)).map L).f := by dsimp @@ -274,10 +274,10 @@ end section /-- A functor preserving zero morphisms, monos, and cokernels preserves homology. -/ -def preservesHomologyOfPreservesMonosAndCokernels [PreservesZeroMorphisms L] +lemma preservesHomology_of_preservesMonos_and_cokernels [PreservesZeroMorphisms L] [PreservesMonomorphisms L] [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) L] : PreservesHomology L := by - apply preservesHomologyOfMapExact + apply preservesHomology_of_map_exact intro S hS let φ : (ShortComplex.mk _ _ (Abelian.comp_coimage_π_eq_zero S.zero)).map L ⟶ S.map L := { τ₁ := 𝟙 _ @@ -291,10 +291,10 @@ def preservesHomologyOfPreservesMonosAndCokernels [PreservesZeroMorphisms L] exact CokernelCofork.mapIsColimit _ ((S.exact_iff_exact_coimage_π).1 hS).gIsCokernel L /-- A functor preserving zero morphisms, epis, and kernels preserves homology. -/ -def preservesHomologyOfPreservesEpisAndKernels [PreservesZeroMorphisms L] +lemma preservesHomology_of_preservesEpis_and_kernels [PreservesZeroMorphisms L] [PreservesEpimorphisms L] [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) L] : PreservesHomology L := by - apply preservesHomologyOfMapExact + apply preservesHomology_of_map_exact intro S hS let φ : S.map L ⟶ (ShortComplex.mk _ _ (Abelian.image_ι_comp_eq_zero S.zero)).map L := { τ₁ := L.map (Abelian.factorThruImage S.f) diff --git a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean index 659a9efa8a888..e11ac111c917c 100644 --- a/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean +++ b/Mathlib/CategoryTheory/Abelian/GrothendieckAxioms.lean @@ -107,24 +107,24 @@ open CoproductsFromFiniteFiltered variable {α : Type w} variable [HasZeroMorphisms C] [HasFiniteBiproducts C] [HasFiniteLimits C] -instance preservesFiniteLimitsLiftToFinset : PreservesFiniteLimits (liftToFinset C α) := - preservesFiniteLimitsOfEvaluation _ fun I => +instance preservesFiniteLimits_liftToFinset : PreservesFiniteLimits (liftToFinset C α) := + preservesFiniteLimits_of_evaluation _ fun I => letI : PreservesFiniteLimits (colim (J := Discrete I) (C := C)) := - preservesFiniteLimitsOfNatIso HasBiproductsOfShape.colimIsoLim.symm + preservesFiniteLimits_of_natIso HasBiproductsOfShape.colimIsoLim.symm letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj (Discrete.functor fun x ↦ ↑x)) := - ⟨fun J _ _ => whiskeringLeftPreservesLimitsOfShape J _⟩ + ⟨fun J _ _ => whiskeringLeft_preservesLimitsOfShape J _⟩ letI : PreservesFiniteLimits ((whiskeringLeft (Discrete I) (Discrete α) C).obj (Discrete.functor (·.val)) ⋙ colim) := - compPreservesFiniteLimits _ _ - preservesFiniteLimitsOfNatIso (liftToFinsetEvaluationIso I).symm + comp_preservesFiniteLimits _ _ + preservesFiniteLimits_of_natIso (liftToFinsetEvaluationIso I).symm /-- A category with finite biproducts and finite limits is AB4 if it is AB5. -/ -def AB4.ofAB5 [HasFilteredColimits C] [AB5 C] : AB4 C where +lemma AB4.of_AB5 [HasFilteredColimits C] [AB5 C] : AB4 C where preservesFiniteLimits J := letI : PreservesFiniteLimits (liftToFinset C J ⋙ colim) := - compPreservesFiniteLimits _ _ - preservesFiniteLimitsOfNatIso (liftToFinsetColimIso) + comp_preservesFiniteLimits _ _ + preservesFiniteLimits_of_natIso (liftToFinsetColimIso) end diff --git a/Mathlib/CategoryTheory/Abelian/Injective.lean b/Mathlib/CategoryTheory/Abelian/Injective.lean index da13360d678e4..b03c8db8b549c 100644 --- a/Mathlib/CategoryTheory/Abelian/Injective.lean +++ b/Mathlib/CategoryTheory/Abelian/Injective.lean @@ -28,15 +28,15 @@ namespace CategoryTheory variable {C : Type u} [Category.{v} C] [Abelian C] /-- The preadditive Yoneda functor on `J` preserves homology if `J` is injective. -/ -instance preservesHomologyPreadditiveYonedaObjOfInjective (J : C) [hJ : Injective J] : +instance preservesHomology_preadditiveYonedaObj_of_injective (J : C) [hJ : Injective J] : (preadditiveYonedaObj J).PreservesHomology := by letI := (injective_iff_preservesEpimorphisms_preadditive_yoneda_obj' J).mp hJ - apply Functor.preservesHomologyOfPreservesEpisAndKernels + apply Functor.preservesHomology_of_preservesEpis_and_kernels /-- The preadditive Yoneda functor on `J` preserves colimits if `J` is injective. -/ -instance preservesFiniteColimitsPreadditiveYonedaObjOfInjective (J : C) [hP : Injective J] : +instance preservesFiniteColimits_preadditiveYonedaObj_of_injective (J : C) [hP : Injective J] : PreservesFiniteColimits (preadditiveYonedaObj J) := by - apply Functor.preservesFiniteColimitsOfPreservesHomology + apply Functor.preservesFiniteColimits_of_preservesHomology /-- An object is injective if its preadditive Yoneda functor preserves finite colimits. -/ theorem injective_of_preservesFiniteColimits_preadditiveYonedaObj (J : C) diff --git a/Mathlib/CategoryTheory/Abelian/Projective.lean b/Mathlib/CategoryTheory/Abelian/Projective.lean index 1eb7a8843324e..af417648361c8 100644 --- a/Mathlib/CategoryTheory/Abelian/Projective.lean +++ b/Mathlib/CategoryTheory/Abelian/Projective.lean @@ -26,18 +26,18 @@ open Limits Projective Opposite variable {C : Type u} [Category.{v} C] [Abelian C] /-- The preadditive Co-Yoneda functor on `P` preserves homology if `P` is projective. -/ -noncomputable instance preservesHomologyPreadditiveCoyonedaObjOfProjective +noncomputable instance preservesHomology_preadditiveCoyonedaObj_of_projective (P : C) [hP : Projective P] : (preadditiveCoyonedaObj (op P)).PreservesHomology := by haveI := (projective_iff_preservesEpimorphisms_preadditiveCoyoneda_obj' P).mp hP haveI := @Functor.preservesEpimorphisms_of_preserves_of_reflects _ _ _ _ _ _ _ _ this _ - apply Functor.preservesHomologyOfPreservesEpisAndKernels + apply Functor.preservesHomology_of_preservesEpis_and_kernels /-- The preadditive Co-Yoneda functor on `P` preserves finite colimits if `P` is projective. -/ -noncomputable instance preservesFiniteColimitsPreadditiveCoyonedaObjOfProjective +noncomputable instance preservesFiniteColimits_preadditiveCoyonedaObj_of_projective (P : C) [hP : Projective P] : PreservesFiniteColimits (preadditiveCoyonedaObj (op P)) := by - apply Functor.preservesFiniteColimitsOfPreservesHomology + apply Functor.preservesFiniteColimits_of_preservesHomology /-- An object is projective if its preadditive Co-Yoneda functor preserves finite colimits. -/ theorem projective_of_preservesFiniteColimits_preadditiveCoyonedaObj (P : C) diff --git a/Mathlib/CategoryTheory/Abelian/Transfer.lean b/Mathlib/CategoryTheory/Abelian/Transfer.lean index 5a3c93413c248..edda974b2f59c 100644 --- a/Mathlib/CategoryTheory/Abelian/Transfer.lean +++ b/Mathlib/CategoryTheory/Abelian/Transfer.lean @@ -59,7 +59,7 @@ theorem hasKernels [PreservesFiniteLimits G] (i : F ⋙ G ≅ 𝟭 C) : HasKerne /-- No point making this an instance, as it requires `i` and `adj`. -/ theorem hasCokernels (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) : HasCokernels C := { has_colimit := fun f => by - have : PreservesColimits G := adj.leftAdjointPreservesColimits + have : PreservesColimits G := adj.leftAdjoint_preservesColimits have := NatIso.naturality_1 i f simp? at this says simp only [Functor.id_obj, Functor.comp_obj, Functor.comp_map, Functor.id_map] at this @@ -74,7 +74,7 @@ def cokernelIso (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) G.obj (cokernel (F.map f)) ≅ cokernel f := by -- We have to write an explicit `PreservesColimits` type here, -- as `leftAdjointPreservesColimits` has universe variables. - have : PreservesColimits G := adj.leftAdjointPreservesColimits + have : PreservesColimits G := adj.leftAdjoint_preservesColimits calc G.obj (cokernel (F.map f)) ≅ cokernel (G.map (F.map f)) := (asIso (cokernelComparison _ G)).symm @@ -87,7 +87,7 @@ variable [Limits.HasKernels C] [PreservesFiniteLimits G] /-- Auxiliary construction for `coimageIsoImage` -/ def coimageIsoImageAux (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : kernel (G.map (cokernel.π (F.map f))) ≅ kernel (cokernel.π f) := by - have : PreservesColimits G := adj.leftAdjointPreservesColimits + have : PreservesColimits G := adj.leftAdjoint_preservesColimits calc kernel (G.map (cokernel.π (F.map f))) ≅ kernel (cokernel.π (G.map (F.map f)) ≫ cokernelComparison (F.map f) G) := @@ -113,7 +113,7 @@ We still need to check that this agrees with the canonical morphism. -/ def coimageIsoImage (i : F ⋙ G ≅ 𝟭 C) (adj : G ⊣ F) {X Y : C} (f : X ⟶ Y) : Abelian.coimage f ≅ Abelian.image f := by - have : PreservesLimits F := adj.rightAdjointPreservesLimits + have : PreservesLimits F := adj.rightAdjoint_preservesLimits calc Abelian.coimage f ≅ cokernel (kernel.ι f) := Iso.refl _ _ ≅ G.obj (cokernel (F.map (kernel.ι f))) := (cokernelIso _ _ i adj _).symm diff --git a/Mathlib/CategoryTheory/Adhesive.lean b/Mathlib/CategoryTheory/Adhesive.lean index 51235c1937b60..1337051cd7b2a 100644 --- a/Mathlib/CategoryTheory/Adhesive.lean +++ b/Mathlib/CategoryTheory/Adhesive.lean @@ -293,9 +293,9 @@ theorem adhesive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) [F.ReflectsIsomorphisms] : Adhesive C := by haveI : ReflectsLimitsOfShape WalkingCospan F := - reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape_of_reflectsIsomorphisms haveI : ReflectsColimitsOfShape WalkingSpan F := - reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape_of_reflectsIsomorphisms exact adhesive_of_preserves_and_reflects F theorem adhesive_of_reflective [HasPullbacks D] [Adhesive C] [HasPullbacks C] [HasPushouts C] @@ -303,8 +303,8 @@ theorem adhesive_of_reflective [HasPullbacks D] [Adhesive C] [HasPullbacks C] [H {Gl : C ⥤ D} {Gr : D ⥤ C} (adj : Gl ⊣ Gr) [Gr.Full] [Gr.Faithful] [PreservesLimitsOfShape WalkingCospan Gl] : Adhesive D := by - have := adj.leftAdjointPreservesColimits - have := adj.rightAdjointPreservesLimits + have := adj.leftAdjoint_preservesColimits + have := adj.rightAdjoint_preservesLimits apply Adhesive.mk (hasPushout_of_mono_left := H₂) intro W X Y Z f g h i _ H have := Adhesive.van_kampen (IsPushout.of_hasPushout (Gr.map f) (Gr.map g)) diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index f57ec0e073cee..dd86591976534 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -9,14 +9,14 @@ import Mathlib.CategoryTheory.Limits.Creates /-! # Adjunctions and limits -A left adjoint preserves colimits (`CategoryTheory.Adjunction.leftAdjointPreservesColimits`), -and a right adjoint preserves limits (`CategoryTheory.Adjunction.rightAdjointPreservesLimits`). +A left adjoint preserves colimits (`CategoryTheory.Adjunction.leftAdjoint_preservesColimits`), +and a right adjoint preserves limits (`CategoryTheory.Adjunction.rightAdjoint_preservesLimits`). Equivalences create and reflect (co)limits. -(`CategoryTheory.Adjunction.isEquivalenceCreatesLimits`, -`CategoryTheory.Adjunction.isEquivalenceCreatesColimits`, -`CategoryTheory.Adjunction.isEquivalenceReflectsLimits`, -`CategoryTheory.Adjunction.isEquivalenceReflectsColimits`,) +(`CategoryTheory.Functor.createsLimitsOfIsEquivalence`, +`CategoryTheory.Functor.createsColimitsOfIsEquivalence`, +`CategoryTheory.Functor.reflectsLimits_of_isEquivalence`, +`CategoryTheory.Functor.reflectsColimits_of_isEquivalence`.) In `CategoryTheory.Adjunction.coconesIso` we show that when `F ⊣ G`, @@ -78,40 +78,53 @@ def functorialityAdjunction : Cocones.functoriality K F ⊣ functorialityRightAd unit := functorialityUnit adj K counit := functorialityCounit adj K +include adj in /-- A left adjoint preserves colimits. See . -/ -def leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F where +lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where preservesColimitsOfShape := { preservesColimit := { preserves := fun hc => - IsColimit.isoUniqueCoconeMorphism.inv fun _ => + ⟨IsColimit.isoUniqueCoconeMorphism.inv fun _ => @Equiv.unique _ _ (IsColimit.isoUniqueCoconeMorphism.hom hc _) - ((adj.functorialityAdjunction _).homEquiv _ _) } } + ((adj.functorialityAdjunction _).homEquiv _ _)⟩ } } + +include adj in +@[deprecated (since := "2024-11-19")] +lemma leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F := + adj.leftAdjoint_preservesColimits noncomputable -instance colimPreservesColimits [HasColimitsOfShape J C] : +instance colim_preservesColimits [HasColimitsOfShape J C] : PreservesColimits (colim (J := J) (C := C)) := - colimConstAdj.leftAdjointPreservesColimits + colimConstAdj.leftAdjoint_preservesColimits -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalencePreservesColimits +noncomputable instance (priority := 100) isEquivalence_preservesColimits (E : C ⥤ D) [E.IsEquivalence] : PreservesColimitsOfSize.{v, u} E := - leftAdjointPreservesColimits E.adjunction + leftAdjoint_preservesColimits E.adjunction -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceReflectsColimits +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.reflectsColimits_of_isEquivalence (E : D ⥤ C) [E.IsEquivalence] : ReflectsColimitsOfSize.{v, u} E where reflectsColimitsOfShape := { reflectsColimit := { reflects := fun t => - (isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm } } + ⟨(isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm⟩ } } + +@[deprecated (since := "2024-11-18")] +lemma isEquivalenceReflectsColimits (E : D ⥤ C) [E.IsEquivalence] : + ReflectsColimitsOfSize.{v, u} E := + Functor.reflectsColimits_of_isEquivalence E -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceCreatesColimits (H : D ⥤ C) +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.createsColimitsOfIsEquivalence (H : D ⥤ C) [H.IsEquivalence] : CreatesColimitsOfSize.{v, u} H where CreatesColimitsOfShape := @@ -120,16 +133,19 @@ noncomputable instance (priority := 100) isEquivalenceCreatesColimits (H : D ⥤ { liftedCocone := mapCoconeInv H c validLift := mapCoconeMapCoconeInv H c } } } +@[deprecated (since := "2024-11-18")] alias isEquivalenceCreatesColimits := + Functor.createsColimitsOfIsEquivalence + -- verify the preserve_colimits instance works as expected: noncomputable example (E : C ⥤ D) [E.IsEquivalence] (c : Cocone K) (h : IsColimit c) : IsColimit (E.mapCocone c) := - PreservesColimit.preserves h + isColimitOfPreserves E h theorem hasColimit_comp_equivalence (E : C ⥤ D) [E.IsEquivalence] [HasColimit K] : HasColimit (K ⋙ E) := HasColimit.mk { cocone := E.mapCocone (colimit.cocone K) - isColimit := PreservesColimit.preserves (colimit.isColimit K) } + isColimit := isColimitOfPreserves _ (colimit.isColimit K) } theorem hasColimit_of_comp_equivalence (E : C ⥤ D) [E.IsEquivalence] [HasColimit (K ⋙ E)] : HasColimit K := @@ -186,40 +202,52 @@ def functorialityAdjunction' : functorialityLeftAdjoint adj K ⊣ Cones.functori unit := functorialityUnit' adj K counit := functorialityCounit' adj K +include adj in /-- A right adjoint preserves limits. See . -/ -def rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G where +lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where preservesLimitsOfShape := { preservesLimit := { preserves := fun hc => - IsLimit.isoUniqueConeMorphism.inv fun _ => + ⟨IsLimit.isoUniqueConeMorphism.inv fun _ => @Equiv.unique _ _ (IsLimit.isoUniqueConeMorphism.hom hc _) - ((adj.functorialityAdjunction' _).homEquiv _ _).symm } } + ((adj.functorialityAdjunction' _).homEquiv _ _).symm⟩ } } -noncomputable -instance limPreservesLimits [HasLimitsOfShape J C] : +include adj in +@[deprecated (since := "2024-11-19")] +lemma rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G := + adj.rightAdjoint_preservesLimits + +instance lim_preservesLimits [HasLimitsOfShape J C] : PreservesLimits (lim (J := J) (C := C)) := - constLimAdj.rightAdjointPreservesLimits + constLimAdj.rightAdjoint_preservesLimits -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalencePreservesLimits +instance (priority := 100) isEquivalencePreservesLimits (E : D ⥤ C) [E.IsEquivalence] : PreservesLimitsOfSize.{v, u} E := - rightAdjointPreservesLimits E.asEquivalence.symm.toAdjunction + rightAdjoint_preservesLimits E.asEquivalence.symm.toAdjunction -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceReflectsLimits +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.reflectsLimits_of_isEquivalence (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E where reflectsLimitsOfShape := { reflectsLimit := { reflects := fun t => - (isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm } } + ⟨(isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm⟩ } } + +@[deprecated (since := "2024-11-18")] +lemma isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : + ReflectsLimitsOfSize.{v, u} E := + Functor.reflectsLimits_of_isEquivalence E -- see Note [lower instance priority] -noncomputable instance (priority := 100) isEquivalenceCreatesLimits (H : D ⥤ C) [H.IsEquivalence] : +noncomputable instance (priority := 100) + _root_.CategoryTheory.Functor.createsLimitsOfIsEquivalence (H : D ⥤ C) [H.IsEquivalence] : CreatesLimitsOfSize.{v, u} H where CreatesLimitsOfShape := { CreatesLimit := @@ -227,15 +255,18 @@ noncomputable instance (priority := 100) isEquivalenceCreatesLimits (H : D ⥤ C { liftedCone := mapConeInv H c validLift := mapConeMapConeInv H c } } } +@[deprecated (since := "2024-11-18")] alias isEquivalenceCreatesLimits := + Functor.createsLimitsOfIsEquivalence + -- verify the preserve_limits instance works as expected: noncomputable example (E : D ⥤ C) [E.IsEquivalence] (c : Cone K) (h : IsLimit c) : IsLimit (E.mapCone c) := - PreservesLimit.preserves h + isLimitOfPreserves E h theorem hasLimit_comp_equivalence (E : D ⥤ C) [E.IsEquivalence] [HasLimit K] : HasLimit (K ⋙ E) := HasLimit.mk { cone := E.mapCone (limit.cone K) - isLimit := PreservesLimit.preserves (limit.isLimit K) } + isLimit := isLimitOfPreserves _ (limit.isLimit K) } theorem hasLimit_of_comp_equivalence (E : D ⥤ C) [E.IsEquivalence] [HasLimit (K ⋙ E)] : HasLimit K := @@ -330,12 +361,12 @@ variable {J C D : Type*} [Category J] [Category C] [Category D] (F : C ⥤ D) noncomputable instance [IsLeftAdjoint F] : PreservesColimitsOfShape J F := - (Adjunction.ofIsLeftAdjoint F).leftAdjointPreservesColimits.preservesColimitsOfShape + (Adjunction.ofIsLeftAdjoint F).leftAdjoint_preservesColimits.preservesColimitsOfShape noncomputable instance [IsLeftAdjoint F] : PreservesColimitsOfSize.{v, u} F where noncomputable instance [IsRightAdjoint F] : PreservesLimitsOfShape J F := - (Adjunction.ofIsRightAdjoint F).rightAdjointPreservesLimits.preservesLimitsOfShape + (Adjunction.ofIsRightAdjoint F).rightAdjoint_preservesLimits.preservesLimitsOfShape noncomputable instance [IsRightAdjoint F] : PreservesLimitsOfSize.{v, u} F where diff --git a/Mathlib/CategoryTheory/Category/Cat/Limit.lean b/Mathlib/CategoryTheory/Category/Cat/Limit.lean index 92300512f30e6..e16a6957e1e3f 100644 --- a/Mathlib/CategoryTheory/Category/Cat/Limit.lean +++ b/Mathlib/CategoryTheory/Category/Cat/Limit.lean @@ -146,7 +146,7 @@ instance : HasLimits Cat.{v, v} where instance : PreservesLimits Cat.objects.{v, v} where preservesLimitsOfShape := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (HasLimits.limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (HasLimits.limitConeIsLimit F) (Limits.IsLimit.ofIsoLimit (limit.isLimit (F ⋙ Cat.objects)) (Cones.ext (by rfl) (by aesop_cat))) } diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean index 80e1a62db2144..0172d99e2112c 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts.lean @@ -246,15 +246,16 @@ lemma map_toUnit_comp_terminalCompariso (A : C) : open Limits /-- If `terminalComparison F` is an Iso, then `F` preserves terminal objects. -/ -noncomputable def preservesLimitEmptyOfIsIsoTerminalComparison [IsIso (terminalComparison F)] : +lemma preservesLimit_empty_of_isIso_terminalComparison [IsIso (terminalComparison F)] : PreservesLimit (Functor.empty.{0} C) F := by - apply preservesLimitOfPreservesLimitCone terminal.isLimit + apply preservesLimit_of_preserves_limit_cone terminal.isLimit apply isLimitChangeEmptyCone D terminal.isLimit exact asIso (terminalComparison F)|>.symm /-- If `F` preserves terminal objects, then `terminalComparison F` is an isomorphism. -/ -def preservesTerminalIso [h : PreservesLimit (Functor.empty.{0} C) F] : F.obj (𝟙_ C) ≅ 𝟙_ D := - (isLimitChangeEmptyCone D (h.preserves terminal.isLimit) (asEmptyCone (F.obj (𝟙_ C))) +noncomputable def preservesTerminalIso [h : PreservesLimit (Functor.empty.{0} C) F] : + F.obj (𝟙_ C) ≅ 𝟙_ D := + (isLimitChangeEmptyCone D (isLimitOfPreserves _ terminal.isLimit) (asEmptyCone (F.obj (𝟙_ C))) (Iso.refl _)).conePointUniqueUpToIso terminal.isLimit @[simp] @@ -415,14 +416,14 @@ variable [PreservesLimit (pair A B) F] /-- If `F` preserves the limit of the pair `(A, B)`, then the binary fan given by `(F.map fst A B, F.map (snd A B))` is a limit cone. -/ -def isLimitChosenFiniteProductsOfPreservesLimits : +noncomputable def isLimitChosenFiniteProductsOfPreservesLimits : IsLimit <| BinaryFan.mk (F.map (fst A B)) (F.map (snd A B)) := mapIsLimitOfPreservesOfIsLimit F (fst _ _) (snd _ _) <| (product A B).isLimit.ofIsoLimit <| isoBinaryFanMk (product A B).cone /-- If `F` preserves the limit of the pair `(A, B)`, then `prodComparison F A B` is an isomorphism. -/ -def prodComparisonIso : F.obj (A ⊗ B) ≅ F.obj A ⊗ F.obj B := +noncomputable def prodComparisonIso : F.obj (A ⊗ B) ≅ F.obj A ⊗ F.obj B := IsLimit.conePointUniqueUpToIso (isLimitChosenFiniteProductsOfPreservesLimits F A B) (product _ _).isLimit @@ -456,9 +457,10 @@ end PreservesLimitPairs section ProdComparisonIso /-- If `prodComparison F A B` is an isomorphism, then `F` preserves the limit of `pair A B`. -/ -noncomputable def preservesLimitPairOfIsIsoProdComparison (A B : C) [IsIso (prodComparison F A B)] : +lemma preservesLimit_pair_of_isIso_prodComparison (A B : C) + [IsIso (prodComparison F A B)] : PreservesLimit (pair A B) F := by - apply preservesLimitOfPreservesLimitCone (product A B).isLimit + apply preservesLimit_of_preserves_limit_cone (product A B).isLimit refine IsLimit.equivOfNatIsoOfIso (pairComp A B F) _ ((product (F.obj A) (F.obj B)).cone.extend (prodComparison F A B)) (BinaryFan.ext (by exact Iso.refl _) ?_ ?_) |>.invFun @@ -470,12 +472,12 @@ noncomputable def preservesLimitPairOfIsIsoProdComparison (A B : C) [IsIso (prod /-- If `prodComparison F A B` is an isomorphism for all `A B` then `F` preserves limits of shape `Discrete (WalkingPair)`. -/ -noncomputable def preservesLimitsOfShapeDiscreteWalkingPairOfIsoProdComparison +lemma preservesLimitsOfShape_discrete_walkingPair_of_isIso_prodComparison [∀ A B, IsIso (prodComparison F A B)] : PreservesLimitsOfShape (Discrete WalkingPair) F := by constructor intro K - refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoPair K).symm ?_ - apply preservesLimitPairOfIsIsoProdComparison + refine @preservesLimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoPair K).symm ?_ + apply preservesLimit_pair_of_isIso_prodComparison end ProdComparisonIso diff --git a/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean b/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean index 3a4eaa80fc91f..68e9344407ae8 100644 --- a/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/ChosenFiniteProducts/FunctorCategory.lean @@ -154,11 +154,11 @@ lemma associator_inv_app (F₁ F₂ F₃ : J ⥤ C) (j : J) : noncomputable instance {K : Type*} [Category K] [HasColimitsOfShape K C] [∀ X : C, PreservesColimitsOfShape K (tensorLeft X)] {F : J ⥤ C} : PreservesColimitsOfShape K (tensorLeft F) := by - apply preservesColimitsOfShapeOfEvaluation + apply preservesColimitsOfShape_of_evaluation intro k haveI : tensorLeft F ⋙ (evaluation J C).obj k ≅ (evaluation J C).obj k ⋙ tensorLeft (F.obj k) := NatIso.ofComponents (fun _ ↦ Iso.refl _) - exact preservesColimitsOfShapeOfNatIso this.symm + exact preservesColimitsOfShape_of_natIso this.symm end Monoidal diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index 90b541c12ab47..6b72829dd536d 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -148,7 +148,7 @@ theorem coev_ev : (coev A).app (A ⟹ B) ≫ (exp A).map ((ev A).app B) = 𝟙 ( end exp instance : PreservesColimits (tensorLeft A) := - (ihom.adjunction A).leftAdjointPreservesColimits + (ihom.adjunction A).leftAdjoint_preservesColimits variable {A} diff --git a/Mathlib/CategoryTheory/Closed/Ideal.lean b/Mathlib/CategoryTheory/Closed/Ideal.lean index 729c2ce13fd84..7421ed490bb18 100644 --- a/Mathlib/CategoryTheory/Closed/Ideal.lean +++ b/Mathlib/CategoryTheory/Closed/Ideal.lean @@ -120,7 +120,7 @@ def reflectiveChosenFiniteProducts [ChosenFiniteProducts C] [Reflective i] : ((reflector i).map (fst (i.obj X) (i.obj Y)) ≫ (reflectorAdjunction i).counit.app _) ((reflector i).map (snd (i.obj X) (i.obj Y)) ≫ (reflectorAdjunction i).counit.app _) isLimit := by - apply (by infer_instance : ReflectsLimit _ i).reflects + apply isLimitOfReflects i apply IsLimit.equivOfNatIsoOfIso (pairComp X Y _) _ _ _|>.invFun (product (i.obj X) (i.obj Y)).isLimit fapply BinaryFan.ext @@ -141,7 +141,7 @@ def reflectiveChosenFiniteProducts [ChosenFiniteProducts C] [Reflective i] : terminal := { cone := Limits.asEmptyCone <| (reflector i).obj (𝟙_ C) isLimit := by - apply (by infer_instance : ReflectsLimit _ i).reflects + apply isLimitOfReflects i apply isLimitChangeEmptyCone _ ChosenFiniteProducts.terminal.isLimit letI : IsIso ((reflectorAdjunction i).unit.app (𝟙_ C)) := by apply Functor.essImage.unit_isIso @@ -195,15 +195,15 @@ def cartesianClosedOfReflective : CartesianClosed D where · symm refine NatIso.ofComponents (fun X => ?_) (fun f => ?_) · haveI := - Adjunction.rightAdjointPreservesLimits.{0, 0} (reflectorAdjunction i) + Adjunction.rightAdjoint_preservesLimits.{0, 0} (reflectorAdjunction i) apply asIso (prodComparison i B X) · dsimp [asIso] rw [prodComparison_natural_whiskerLeft] · apply (exponentialIdealReflective i _).symm } -- It's annoying that I need to do this. -attribute [-instance] CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit - CategoryTheory.preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape +attribute [-instance] CategoryTheory.preservesLimit_of_createsLimit_and_hasLimit + CategoryTheory.preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape /-- We construct a bijection between morphisms `L(A ⊗ B) ⟶ X` and morphisms `LA ⊗ LB ⟶ X`. This bijection has two key properties: @@ -231,8 +231,8 @@ noncomputable def bijection (A B : C) (X : D) : _ ≃ (i.obj ((reflector i).obj A) ⊗ i.obj ((reflector i).obj B) ⟶ i.obj X) := ((exp.adjunction _).homEquiv _ _).symm _ ≃ (i.obj ((reflector i).obj A ⊗ (reflector i).obj B) ⟶ i.obj X) := - haveI : Limits.PreservesLimits i := (reflectorAdjunction i).rightAdjointPreservesLimits - haveI := Limits.preservesSmallestLimitsOfPreservesLimits i + haveI : Limits.PreservesLimits i := (reflectorAdjunction i).rightAdjoint_preservesLimits + haveI := Limits.preservesSmallestLimits_of_preservesLimits i Iso.homCongr (prodComparisonIso _ _ _).symm (Iso.refl (i.obj X)) _ ≃ ((reflector i).obj A ⊗ (reflector i).obj B ⟶ X) := i.fullyFaithfulOfReflective.homEquiv.symm @@ -292,21 +292,21 @@ open Limits If a reflective subcategory is an exponential ideal, then the reflector preserves binary products. This is the converse of `exponentialIdeal_of_preserves_binary_products`. -/ -noncomputable def preservesBinaryProductsOfExponentialIdeal : +lemma preservesBinaryProducts_of_exponentialIdeal : PreservesLimitsOfShape (Discrete WalkingPair) (reflector i) where preservesLimit {K} := - letI := preservesLimitPairOfIsIsoProdComparison + letI := preservesLimit_pair_of_isIso_prodComparison (reflector i) (K.obj ⟨WalkingPair.left⟩) (K.obj ⟨WalkingPair.right⟩) - Limits.preservesLimitOfIsoDiagram _ (diagramIsoPair K).symm + Limits.preservesLimit_of_iso_diagram _ (diagramIsoPair K).symm /-- If a reflective subcategory is an exponential ideal, then the reflector preserves finite products. -/ -noncomputable def preservesFiniteProductsOfExponentialIdeal (J : Type) [Fintype J] : +lemma preservesFiniteProducts_of_exponentialIdeal (J : Type) [Fintype J] : PreservesLimitsOfShape (Discrete J) (reflector i) := by - letI := preservesBinaryProductsOfExponentialIdeal i - letI : PreservesLimitsOfShape _ (reflector i) := leftAdjointPreservesTerminalOfReflective.{0} i - apply preservesFiniteProductsOfPreservesBinaryAndTerminal (reflector i) J + letI := preservesBinaryProducts_of_exponentialIdeal i + letI : PreservesLimitsOfShape _ (reflector i) := leftAdjoint_preservesTerminal_of_reflective.{0} i + apply preservesFiniteProducts_of_preserves_binary_and_terminal (reflector i) J end diff --git a/Mathlib/CategoryTheory/Closed/Monoidal.lean b/Mathlib/CategoryTheory/Closed/Monoidal.lean index fec6606dada4e..4e97aea044bd9 100644 --- a/Mathlib/CategoryTheory/Closed/Monoidal.lean +++ b/Mathlib/CategoryTheory/Closed/Monoidal.lean @@ -115,7 +115,7 @@ end ihom open CategoryTheory.Limits instance : PreservesColimits (tensorLeft A) := - (ihom.adjunction A).leftAdjointPreservesColimits + (ihom.adjunction A).leftAdjoint_preservesColimits variable {A} diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 4b4479017c601..87488121b4c56 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -146,18 +146,18 @@ variable [PreservesPullbacksOfInclusions F] {X Y Z : C} (f : Z ⟶ X ⨿ Y) noncomputable instance preservesPullbackInl' : PreservesLimit (cospan f coprod.inl) F := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ noncomputable instance preservesPullbackInr' : PreservesLimit (cospan f coprod.inr) F := by - apply preservesLimitOfIsoDiagram (K₁ := cospan (f ≫ (coprod.braiding X Y).hom) coprod.inl) + apply preservesLimit_of_iso_diagram (K₁ := cospan (f ≫ (coprod.braiding X Y).hom) coprod.inl) apply cospanExt (Iso.refl _) (Iso.refl _) (coprod.braiding X Y).symm <;> simp noncomputable instance preservesPullbackInr : PreservesLimit (cospan coprod.inr f) F := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ end PreservesPullbacksOfInclusions @@ -354,7 +354,7 @@ theorem finitaryExtensive_of_reflective [∀ X Y (f : X ⟶ Gl.obj Y), PreservesLimit (cospan (Gr.map f) (adj.unit.app Y)) Gl] [PreservesPullbacksOfInclusions Gl] : FinitaryExtensive D := by - have : PreservesColimitsOfSize Gl := adj.leftAdjointPreservesColimits + have : PreservesColimitsOfSize Gl := adj.leftAdjoint_preservesColimits constructor intros X Y c hc apply (IsVanKampenColimit.precompose_isIso_iff @@ -376,20 +376,18 @@ instance finitaryExtensive_functor [HasPullbacks C] [FinitaryExtensive C] : FinitaryExtensive (D ⥤ C) := haveI : HasFiniteCoproducts (D ⥤ C) := ⟨fun _ => Limits.functorCategoryHasColimitsOfShape⟩ ⟨fun c hc => isVanKampenColimit_of_evaluation _ c fun _ => - FinitaryExtensive.vanKampen _ <| PreservesColimit.preserves hc⟩ + FinitaryExtensive.vanKampen _ <| isColimitOfPreserves _ hc⟩ -noncomputable instance {C} [Category C] {D} [Category D] (F : C ⥤ D) {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [IsIso f] : PreservesLimit (cospan f g) F := have := hasPullback_of_left_iso f g - preservesLimitOfPreservesLimitCone (IsPullback.of_hasPullback f g).isLimit + preservesLimit_of_preserves_limit_cone (IsPullback.of_hasPullback f g).isLimit ((isLimitMapConePullbackConeEquiv _ pullback.condition).symm (IsPullback.of_vert_isIso ⟨by simp only [← F.map_comp, pullback.condition]⟩).isLimit) -noncomputable instance {C} [Category C] {D} [Category D] (F : C ⥤ D) {X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) [IsIso g] : PreservesLimit (cospan f g) F := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ theorem finitaryExtensive_of_preserves_and_reflects (F : C ⥤ D) [FinitaryExtensive D] [HasFiniteCoproducts C] [HasPullbacksOfInclusions C] @@ -409,9 +407,9 @@ theorem finitaryExtensive_of_preserves_and_reflects_isomorphism (F : C ⥤ D) [F [HasFiniteCoproducts C] [HasPullbacks C] [PreservesLimitsOfShape WalkingCospan F] [PreservesColimitsOfShape (Discrete WalkingPair) F] [F.ReflectsIsomorphisms] : FinitaryExtensive C := by - haveI : ReflectsLimitsOfShape WalkingCospan F := reflectsLimitsOfShapeOfReflectsIsomorphisms + haveI : ReflectsLimitsOfShape WalkingCospan F := reflectsLimitsOfShape_of_reflectsIsomorphisms haveI : ReflectsColimitsOfShape (Discrete WalkingPair) F := - reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape_of_reflectsIsomorphisms exact finitaryExtensive_of_preserves_and_reflects F end Functor diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index 43e090be3c97b..2d0444d695f18 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -28,14 +28,14 @@ This definition is equivalent to left exact functors (functors that preserves fi * `flat_of_preservesFiniteLimits`: If `F : C ⥤ D` preserves finite limits and `C` has all finite limits, then `F` is flat. -* `preservesFiniteLimitsOfFlat`: If `F : C ⥤ D` is flat, then it preserves all finite limits. -* `preservesFiniteLimitsIffFlat`: If `C` has all finite limits, +* `preservesFiniteLimits_of_flat`: If `F : C ⥤ D` is flat, then it preserves all finite limits. +* `preservesFiniteLimits_iff_flat`: If `C` has all finite limits, then `F` is flat iff `F` is left_exact. -* `lanPreservesFiniteLimitsOfFlat`: If `F : C ⥤ D` is a flat functor between small categories, +* `lan_preservesFiniteLimits_of_flat`: If `F : C ⥤ D` is a flat functor between small categories, then the functor `Lan F.op` between presheaves of sets preserves all finite limits. * `flat_iff_lan_flat`: If `C`, `D` are small and `C` has all finite limits, then `F` is flat iff `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` is flat. -* `preservesFiniteLimitsIffLanPreservesFiniteLimits`: If `C`, `D` are small and `C` has all +* `preservesFiniteLimits_iff_lanPreservesFiniteLimits`: If `C`, `D` are small and `C` has all finite limits, then `F` preserves finite limits iff `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` does. @@ -157,7 +157,7 @@ theorem flat_of_preservesFiniteLimits [HasFiniteLimits C] (F : C ⥤ D) [Preserv theorem coflat_of_preservesFiniteColimits [HasFiniteColimits C] (F : C ⥤ D) [PreservesFiniteColimits F] : RepresentablyCoflat F := - let _ := preservesFiniteLimitsOp F + let _ := preservesFiniteLimits_op F (representablyFlat_op_iff _).1 (flat_of_preservesFiniteLimits _) namespace PreservesFiniteLimitsOfFlat @@ -239,12 +239,13 @@ theorem uniq {K : J ⥤ C} {c : Cone K} (hc : IsLimit c) (s : Cone (K ⋙ F)) end PreservesFiniteLimitsOfFlat /-- Representably flat functors preserve finite limits. -/ -noncomputable def preservesFiniteLimitsOfFlat (F : C ⥤ D) [RepresentablyFlat F] : +lemma preservesFiniteLimits_of_flat (F : C ⥤ D) [RepresentablyFlat F] : PreservesFiniteLimits F := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize intro J _ _; constructor intro K; constructor intro c hc + constructor exact { lift := PreservesFiniteLimitsOfFlat.lift F hc fac := PreservesFiniteLimitsOfFlat.fac F hc @@ -254,25 +255,23 @@ noncomputable def preservesFiniteLimitsOfFlat (F : C ⥤ D) [RepresentablyFlat F · exact PreservesFiniteLimitsOfFlat.fac F hc s } /-- Representably coflat functors preserve finite colimits. -/ -noncomputable def preservesFiniteColimitsOfCoflat (F : C ⥤ D) [RepresentablyCoflat F] : +lemma preservesFiniteColimits_of_coflat (F : C ⥤ D) [RepresentablyCoflat F] : PreservesFiniteColimits F := - letI _ := preservesFiniteLimitsOfFlat F.op - preservesFiniteColimitsOfOp _ + letI _ := preservesFiniteLimits_of_flat F.op + preservesFiniteColimits_of_op _ /-- If `C` is finitely complete, then `F : C ⥤ D` is representably flat iff it preserves finite limits. -/ -noncomputable def preservesFiniteLimitsIffFlat [HasFiniteLimits C] (F : C ⥤ D) : - RepresentablyFlat F ≃ PreservesFiniteLimits F := - equivOfSubsingletonOfSubsingleton - (fun _ => preservesFiniteLimitsOfFlat F) (fun _ => flat_of_preservesFiniteLimits F) +lemma preservesFiniteLimits_iff_flat [HasFiniteLimits C] (F : C ⥤ D) : + RepresentablyFlat F ↔ PreservesFiniteLimits F := + ⟨fun _ ↦ preservesFiniteLimits_of_flat F, fun _ ↦ flat_of_preservesFiniteLimits F⟩ /-- If `C` is finitely cocomplete, then `F : C ⥤ D` is representably coflat iff it preserves finite colmits. -/ -noncomputable def preservesFiniteColimitsIffCoflat [HasFiniteColimits C] (F : C ⥤ D) : - RepresentablyCoflat F ≃ PreservesFiniteColimits F := - equivOfSubsingletonOfSubsingleton - (fun _ => preservesFiniteColimitsOfCoflat F) (fun _ => coflat_of_preservesFiniteColimits F) +lemma preservesFiniteColimits_iff_coflat [HasFiniteColimits C] (F : C ⥤ D) : + RepresentablyCoflat F ↔ PreservesFiniteColimits F := + ⟨fun _ => preservesFiniteColimits_of_coflat F, fun _ => coflat_of_preservesFiniteColimits F⟩ end HasLimit @@ -311,15 +310,15 @@ variable [PreservesLimits (forget E)] /-- If `F : C ⥤ D` is a representably flat functor between small categories, then the functor `Lan F.op` that takes presheaves over `C` to presheaves over `D` preserves finite limits. -/ -noncomputable instance lanPreservesFiniteLimitsOfFlat (F : C ⥤ D) [RepresentablyFlat F] : +noncomputable instance lan_preservesFiniteLimits_of_flat (F : C ⥤ D) [RepresentablyFlat F] : PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ E) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{u₁} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{u₁} intro J _ _ - apply preservesLimitsOfShapeOfEvaluation (F.op.lan : (Cᵒᵖ ⥤ E) ⥤ Dᵒᵖ ⥤ E) J + apply preservesLimitsOfShape_of_evaluation (F.op.lan : (Cᵒᵖ ⥤ E) ⥤ Dᵒᵖ ⥤ E) J intro K haveI : IsFiltered (CostructuredArrow F.op K) := IsFiltered.of_equivalence (structuredArrowOpEquivalence F (unop K)) - exact preservesLimitsOfShapeOfNatIso (lanEvaluationIsoColim _ _ _).symm + exact preservesLimitsOfShape_of_natIso (lanEvaluationIsoColim _ _ _).symm instance lan_flat_of_flat (F : C ⥤ D) [RepresentablyFlat F] : RepresentablyFlat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ E) := @@ -327,7 +326,7 @@ instance lan_flat_of_flat (F : C ⥤ D) [RepresentablyFlat F] : variable [HasFiniteLimits C] -noncomputable instance lanPreservesFiniteLimitsOfPreservesFiniteLimits (F : C ⥤ D) +instance lan_preservesFiniteLimits_of_preservesFiniteLimits (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ E) := by haveI := flat_of_preservesFiniteLimits F infer_instance @@ -335,37 +334,20 @@ noncomputable instance lanPreservesFiniteLimitsOfPreservesFiniteLimits (F : C theorem flat_iff_lan_flat (F : C ⥤ D) : RepresentablyFlat F ↔ RepresentablyFlat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) := ⟨fun _ => inferInstance, fun H => by - haveI := preservesFiniteLimitsOfFlat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) + haveI := preservesFiniteLimits_of_flat (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) haveI : PreservesFiniteLimits F := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{u₁} - intros; apply preservesLimitOfLanPreservesLimit + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{u₁} + intros; apply preservesLimit_of_lan_preservesLimit apply flat_of_preservesFiniteLimits⟩ /-- If `C` is finitely complete, then `F : C ⥤ D` preserves finite limits iff `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` preserves finite limits. -/ -noncomputable def preservesFiniteLimitsIffLanPreservesFiniteLimits (F : C ⥤ D) : - PreservesFiniteLimits F ≃ PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) where - toFun _ := inferInstance - invFun _ := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{u₁} - intros; apply preservesLimitOfLanPreservesLimit - left_inv x := by - -- Porting note: `cases x` and an `unfold` not necessary in lean 4. - -- Remark : in mathlib3 we had `unfold preservesFiniteLimitsOfFlat` - -- but there was no `preservesFiniteLimitsOfFlat` in the goal! Experimentation - -- indicates that it was doing the same as `dsimp only` - dsimp only [preservesFiniteLimitsOfPreservesFiniteLimitsOfSize]; congr - -- Porting note: next line wasn't necessary in lean 3 - subsingleton - right_inv x := by - -- cases x; -- Porting note: not necessary in lean 4 - dsimp only [lanPreservesFiniteLimitsOfPreservesFiniteLimits, - lanPreservesFiniteLimitsOfFlat, - preservesFiniteLimitsOfPreservesFiniteLimitsOfSize] - congr - -- Porting note: next line wasn't necessary in lean 3 - subsingleton +lemma preservesFiniteLimits_iff_lan_preservesFiniteLimits (F : C ⥤ D) : + PreservesFiniteLimits F ↔ PreservesFiniteLimits (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u₁) := + ⟨fun _ ↦ inferInstance, + fun _ ↦ preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{u₁} _ + (fun _ _ _ ↦ preservesLimit_of_lan_preservesLimit _ _)⟩ end SmallCategory diff --git a/Mathlib/CategoryTheory/Galois/Action.lean b/Mathlib/CategoryTheory/Galois/Action.lean index fbf5fbd2f2b86..e45c24948c104 100644 --- a/Mathlib/CategoryTheory/Galois/Action.lean +++ b/Mathlib/CategoryTheory/Galois/Action.lean @@ -69,16 +69,16 @@ instance : Functor.ReflectsIsomorphisms (functorToAction F) where isIso_of_reflects_iso f F noncomputable instance : PreservesFiniteCoproducts (functorToAction F) := - ⟨fun J _ ↦ Action.preservesColimitsOfShapeOfPreserves (functorToAction F) + ⟨fun J _ ↦ Action.preservesColimitsOfShape_of_preserves (functorToAction F) (inferInstanceAs <| PreservesColimitsOfShape (Discrete J) F)⟩ noncomputable instance : PreservesFiniteProducts (functorToAction F) := - ⟨fun J _ ↦ Action.preservesLimitsOfShapeOfPreserves (functorToAction F) + ⟨fun J _ ↦ Action.preservesLimitsOfShape_of_preserves (functorToAction F) (inferInstanceAs <| PreservesLimitsOfShape (Discrete J) F)⟩ noncomputable instance (G : Type*) [Group G] [Finite G] : PreservesColimitsOfShape (SingleObj G) (functorToAction F) := - Action.preservesColimitsOfShapeOfPreserves _ <| + Action.preservesColimitsOfShape_of_preserves _ <| inferInstanceAs <| PreservesColimitsOfShape (SingleObj G) F instance : PreservesIsConnected (functorToAction F) := diff --git a/Mathlib/CategoryTheory/Galois/Basic.lean b/Mathlib/CategoryTheory/Galois/Basic.lean index 131a2e864a5bb..d6afd4303f42b 100644 --- a/Mathlib/CategoryTheory/Galois/Basic.lean +++ b/Mathlib/CategoryTheory/Galois/Basic.lean @@ -131,19 +131,19 @@ attribute [instance] preservesTerminalObjects preservesPullbacks preservesEpis preservesFiniteCoproducts reflectsIsos preservesQuotientsByFiniteGroups noncomputable instance : ReflectsLimitsOfShape (Discrete PEmpty.{1}) F := - reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape_of_reflectsIsomorphisms noncomputable instance : ReflectsColimitsOfShape (Discrete PEmpty.{1}) F := - reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape_of_reflectsIsomorphisms noncomputable instance : PreservesFiniteLimits F := - preservesFiniteLimitsOfPreservesTerminalAndPullbacks F + preservesFiniteLimits_of_preservesTerminal_and_pullbacks F /-- Fiber functors preserve quotients by finite groups in arbitrary universes. -/ -noncomputable instance {G : Type*} [Group G] [Finite G] : +instance {G : Type*} [Group G] [Finite G] : PreservesColimitsOfShape (SingleObj G) F := by choose G' hg hf he using Finite.exists_type_univ_nonempty_mulEquiv G - exact Limits.preservesColimitsOfShapeOfEquiv he.some.toSingleObjEquiv.symm F + exact Limits.preservesColimitsOfShape_of_equiv he.some.toSingleObjEquiv.symm F /-- Fiber functors reflect monomorphisms. -/ instance : ReflectsMonomorphisms F := ReflectsMonomorphisms.mk <| by @@ -170,9 +170,9 @@ section /-- If `F` is a fiber functor and `E` is an equivalence between categories of finite types, then `F ⋙ E` is again a fiber functor. -/ -noncomputable def compRight (E : FintypeCat.{w} ⥤ FintypeCat.{t}) [E.IsEquivalence] : +lemma comp_right (E : FintypeCat.{w} ⥤ FintypeCat.{t}) [E.IsEquivalence] : FiberFunctor (F ⋙ E) where - preservesQuotientsByFiniteGroups _ := compPreservesColimitsOfShape F E + preservesQuotientsByFiniteGroups _ := comp_preservesColimitsOfShape F E end diff --git a/Mathlib/CategoryTheory/Galois/Examples.lean b/Mathlib/CategoryTheory/Galois/Examples.lean index f52783a0ae08a..046ec34c9ba91 100644 --- a/Mathlib/CategoryTheory/Galois/Examples.lean +++ b/Mathlib/CategoryTheory/Galois/Examples.lean @@ -75,7 +75,7 @@ instance [Finite G] : HasColimitsOfShape (SingleObj G) FintypeCat.{w} := by noncomputable instance : PreservesFiniteLimits (forget (Action FintypeCat (MonCat.of G))) := by show PreservesFiniteLimits (Action.forget FintypeCat _ ⋙ FintypeCat.incl) - apply compPreservesFiniteLimits + apply comp_preservesFiniteLimits /-- The category of finite `G`-sets is a `PreGaloisCategory`. -/ instance : PreGaloisCategory (Action FintypeCat (MonCat.of G)) where diff --git a/Mathlib/CategoryTheory/Galois/GaloisObjects.lean b/Mathlib/CategoryTheory/Galois/GaloisObjects.lean index 9c838d2ad07d4..89a2b0fb5a9ac 100644 --- a/Mathlib/CategoryTheory/Galois/GaloisObjects.lean +++ b/Mathlib/CategoryTheory/Galois/GaloisObjects.lean @@ -37,7 +37,7 @@ open Limits Functor noncomputable instance {G : Type v} [Group G] [Finite G] : PreservesColimitsOfShape (SingleObj G) FintypeCat.incl.{w} := by choose G' hg hf e using Finite.exists_type_univ_nonempty_mulEquiv G - exact Limits.preservesColimitsOfShapeOfEquiv (Classical.choice e).toSingleObjEquiv.symm _ + exact Limits.preservesColimitsOfShape_of_equiv (Classical.choice e).toSingleObjEquiv.symm _ /-- A connected object `X` of `C` is Galois if the quotient `X / Aut X` is terminal. -/ class IsGalois {C : Type u₁} [Category.{u₂, u₁} C] [GaloisCategory C] (X : C) diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index fba4b027d12de..554c363a78c89 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -449,7 +449,7 @@ instance FiberFunctor.isPretransitive_of_isConnected (X : C) [IsConnected X] : MulAction.IsPretransitive (Aut F) (F.obj X) where exists_smul_eq x y := by let F' : C ⥤ FintypeCat.{u₂} := F ⋙ FintypeCat.uSwitch.{w, u₂} - letI : FiberFunctor F' := FiberFunctor.compRight _ + letI : FiberFunctor F' := FiberFunctor.comp_right _ let e (Y : C) : F'.obj Y ≃ F.obj Y := (F.obj Y).uSwitchEquiv set x' : F'.obj X := (e X).symm x with hx' set y' : F'.obj X := (e X).symm y with hy' diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 69eb61fe1f99e..936283e96bec5 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -299,7 +299,7 @@ end variable [HasMulticoequalizer D.diagram] [PreservesColimit D.diagram.multispan F] theorem hasColimit_multispan_comp : HasColimit (D.diagram.multispan ⋙ F) := - ⟨⟨⟨_, PreservesColimit.preserves (colimit.isColimit _)⟩⟩⟩ + ⟨⟨⟨_, isColimitOfPreserves _ (colimit.isColimit _)⟩⟩⟩ attribute [local instance] hasColimit_multispan_comp diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index b005a034c6b49..34da9a269a20b 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -546,6 +546,7 @@ lemma triangle : tensorHom (rightUnitor X₁).hom (𝟙 X₃) := by convert mapBifunctor_triangle (curriedAssociatorNatIso C) (𝟙_ C) (rightUnitorNatIso C) (leftUnitorNatIso C) (triangleIndexData I) X₁ X₃ (by simp) + all_goals assumption end Triangle diff --git a/Mathlib/CategoryTheory/Limits/Comma.lean b/Mathlib/CategoryTheory/Limits/Comma.lean index f67006d4500bf..93476394301e1 100644 --- a/Mathlib/CategoryTheory/Limits/Comma.lean +++ b/Mathlib/CategoryTheory/Limits/Comma.lean @@ -48,7 +48,7 @@ limit cone for `F ⋙ snd L R : J ⥤ R` we can build a cone for `F` which will cone. -/ @[simps] -def coneOfPreserves [PreservesLimit (F ⋙ snd L R) R] (c₁ : Cone (F ⋙ fst L R)) +noncomputable def coneOfPreserves [PreservesLimit (F ⋙ snd L R) R] (c₁ : Cone (F ⋙ fst L R)) {c₂ : Cone (F ⋙ snd L R)} (t₂ : IsLimit c₂) : Cone F where pt := { left := c₁.pt @@ -66,7 +66,7 @@ def coneOfPreserves [PreservesLimit (F ⋙ snd L R) R] (c₁ : Cone (F ⋙ fst L /-- Provided that `R` preserves the appropriate limit, then the cone in `coneOfPreserves` is a limit. -/ -def coneOfPreservesIsLimit [PreservesLimit (F ⋙ snd L R) R] {c₁ : Cone (F ⋙ fst L R)} +noncomputable def coneOfPreservesIsLimit [PreservesLimit (F ⋙ snd L R) R] {c₁ : Cone (F ⋙ fst L R)} (t₁ : IsLimit c₁) {c₂ : Cone (F ⋙ snd L R)} (t₂ : IsLimit c₂) : IsLimit (coneOfPreserves F c₁ t₂) where lift s := @@ -95,7 +95,7 @@ a cocone for `F ⋙ snd L R : J ⥤ R` we can build a cocone for `F` which will colimit cocone. -/ @[simps] -def coconeOfPreserves [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ fst L R)} +noncomputable def coconeOfPreserves [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ fst L R)} (t₁ : IsColimit c₁) (c₂ : Cocone (F ⋙ snd L R)) : Cocone F where pt := { left := c₁.pt @@ -113,7 +113,8 @@ def coconeOfPreserves [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ /-- Provided that `L` preserves the appropriate colimit, then the cocone in `coconeOfPreserves` is a colimit. -/ -def coconeOfPreservesIsColimit [PreservesColimit (F ⋙ fst L R) L] {c₁ : Cocone (F ⋙ fst L R)} +noncomputable def coconeOfPreservesIsColimit [PreservesColimit (F ⋙ fst L R) L] + {c₁ : Cocone (F ⋙ fst L R)} (t₁ : IsColimit c₁) {c₂ : Cocone (F ⋙ snd L R)} (t₂ : IsColimit c₂) : IsColimit (coconeOfPreserves F t₁ c₂) where desc s := diff --git a/Mathlib/CategoryTheory/Limits/ConeCategory.lean b/Mathlib/CategoryTheory/Limits/ConeCategory.lean index f7427f3b61376..f15da439d1804 100644 --- a/Mathlib/CategoryTheory/Limits/ConeCategory.lean +++ b/Mathlib/CategoryTheory/Limits/ConeCategory.lean @@ -181,12 +181,12 @@ theorem IsTerminal.from_eq_liftConeMorphism {F : J ⥤ C} {c : Cone F} (hc : IsT (IsLimit.liftConeMorphism_eq_isTerminal_from (c.isLimitEquivIsTerminal.symm hc) s).symm /-- If `G : Cone F ⥤ Cone F'` preserves terminal objects, it preserves limit cones. -/ -def IsLimit.ofPreservesConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') +noncomputable def IsLimit.ofPreservesConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') [PreservesLimit (Functor.empty.{0} _) G] {c : Cone F} (hc : IsLimit c) : IsLimit (G.obj c) := (Cone.isLimitEquivIsTerminal _).symm <| (Cone.isLimitEquivIsTerminal _ hc).isTerminalObj _ _ /-- If `G : Cone F ⥤ Cone F'` reflects terminal objects, it reflects limit cones. -/ -def IsLimit.ofReflectsConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') +noncomputable def IsLimit.ofReflectsConeTerminal {F : J ⥤ C} {F' : K ⥤ D} (G : Cone F ⥤ Cone F') [ReflectsLimit (Functor.empty.{0} _) G] {c : Cone F} (hc : IsLimit (G.obj c)) : IsLimit c := (Cone.isLimitEquivIsTerminal _).symm <| (Cone.isLimitEquivIsTerminal _ hc).isTerminalOfObj _ _ @@ -341,13 +341,15 @@ theorem IsInitial.to_eq_descCoconeMorphism {F : J ⥤ C} {c : Cocone F} (hc : Is (IsColimit.descCoconeMorphism_eq_isInitial_to (c.isColimitEquivIsInitial.symm hc) s).symm /-- If `G : Cocone F ⥤ Cocone F'` preserves initial objects, it preserves colimit cocones. -/ -def IsColimit.ofPreservesCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} (G : Cocone F ⥤ Cocone F') +noncomputable def IsColimit.ofPreservesCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} + (G : Cocone F ⥤ Cocone F') [PreservesColimit (Functor.empty.{0} _) G] {c : Cocone F} (hc : IsColimit c) : IsColimit (G.obj c) := (Cocone.isColimitEquivIsInitial _).symm <| (Cocone.isColimitEquivIsInitial _ hc).isInitialObj _ _ /-- If `G : Cocone F ⥤ Cocone F'` reflects initial objects, it reflects colimit cocones. -/ -def IsColimit.ofReflectsCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} (G : Cocone F ⥤ Cocone F') +noncomputable def IsColimit.ofReflectsCoconeInitial {F : J ⥤ C} {F' : K ⥤ D} + (G : Cocone F ⥤ Cocone F') [ReflectsColimit (Functor.empty.{0} _) G] {c : Cocone F} (hc : IsColimit (G.obj c)) : IsColimit c := (Cocone.isColimitEquivIsInitial _).symm <| diff --git a/Mathlib/CategoryTheory/Limits/Connected.lean b/Mathlib/CategoryTheory/Limits/Connected.lean index 712b12fbf1148..0d9777ebc99b8 100644 --- a/Mathlib/CategoryTheory/Limits/Connected.lean +++ b/Mathlib/CategoryTheory/Limits/Connected.lean @@ -87,12 +87,11 @@ Note that this functor does not preserve the two most obvious disconnected limit `(X × -)` does not preserve products or terminal object, eg `(X ⨯ A) ⨯ (X ⨯ B)` is not isomorphic to `X ⨯ (A ⨯ B)` and `X ⨯ 1` is not isomorphic to `1`. -/ -noncomputable def prodPreservesConnectedLimits [IsConnected J] (X : C) : +lemma prod_preservesConnectedLimits [IsConnected J] (X : C) : PreservesLimitsOfShape J (prod.functor.obj X) where preservesLimit {K} := - { - preserves := fun {c} l => - { lift := fun s => + { preserves := fun {c} l => ⟨{ + lift := fun s => prod.lift (s.π.app (Classical.arbitrary _) ≫ Limits.prod.fst) (l.lift (forgetCone s)) fac := fun s j => by apply Limits.prod.hom_ext @@ -106,6 +105,6 @@ noncomputable def prodPreservesConnectedLimits [IsConnected J] (X : C) : · rw [limit.lift_π] apply l.uniq (forgetCone s) intro j - simp [← L j] } } + simp [← L j] }⟩ } end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean index 3278a8a966db7..4ad7a58d9ac61 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean @@ -93,11 +93,11 @@ theorem hasBinaryProducts_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullb variable {C} /-- A functor that preserves terminal objects and pullbacks preserves binary products. -/ -noncomputable def preservesBinaryProductsOfPreservesTerminalAndPullbacks [HasTerminal C] +lemma preservesBinaryProducts_of_preservesTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] [PreservesLimitsOfShape (Discrete.{0} PEmpty) F] [PreservesLimitsOfShape WalkingCospan F] : PreservesLimitsOfShape (Discrete WalkingPair) F := ⟨fun {K} => - preservesLimitOfPreservesLimitCone (limitConeOfTerminalAndPullbacks K).2 + preservesLimit_of_preserves_limit_cone (limitConeOfTerminalAndPullbacks K).2 (by apply isBinaryProductOfIsTerminalIsPullback _ _ (isLimitOfHasTerminalOfPreservesLimit F) @@ -199,11 +199,11 @@ theorem hasBinaryCoproducts_of_hasInitial_and_pushouts [HasInitial C] [HasPushou variable {C} /-- A functor that preserves initial objects and pushouts preserves binary coproducts. -/ -noncomputable def preservesBinaryCoproductsOfPreservesInitialAndPushouts [HasInitial C] +lemma preservesBinaryCoproducts_of_preservesInitial_and_pushouts [HasInitial C] [HasPushouts C] [PreservesColimitsOfShape (Discrete.{0} PEmpty) F] [PreservesColimitsOfShape WalkingSpan F] : PreservesColimitsOfShape (Discrete WalkingPair) F := ⟨fun {K} => - preservesColimitOfPreservesColimitCocone (colimitCoconeOfInitialAndPushouts K).2 (by + preservesColimit_of_preserves_colimit_cocone (colimitCoconeOfInitialAndPushouts K).2 (by apply isBinaryCoproductOfIsInitialIsPushout _ _ (isColimitOfHasInitialOfPreservesColimit F) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean index aa87a67935f4c..73d519ebd5219 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Equalizers.lean @@ -86,11 +86,12 @@ theorem hasEqualizers_of_hasPullbacks_and_binary_products [HasBinaryProducts C] attribute [local instance] hasPullback_of_preservesPullback /-- A functor that preserves pullbacks and binary products also presrves equalizers. -/ -def preservesEqualizersOfPreservesPullbacksAndBinaryProducts [HasBinaryProducts C] [HasPullbacks C] +lemma preservesEqualizers_of_preservesPullbacks_and_binaryProducts + [HasBinaryProducts C] [HasPullbacks C] [PreservesLimitsOfShape (Discrete WalkingPair) G] [PreservesLimitsOfShape WalkingCospan G] : PreservesLimitsOfShape WalkingParallelPair G := ⟨fun {K} => - preservesLimitOfPreservesLimitCone (equalizerConeIsLimit K) <| + preservesLimit_of_preserves_limit_cone (equalizerConeIsLimit K) <| { lift := fun c => by refine pullback.lift ?_ ?_ ?_ ≫ (PreservesPullback.iso _ _ _ ).inv · exact c.π.app WalkingParallelPair.zero @@ -180,11 +181,11 @@ theorem hasCoequalizers_of_hasPushouts_and_binary_coproducts [HasBinaryCoproduct attribute [local instance] hasPushout_of_preservesPushout /-- A functor that preserves pushouts and binary coproducts also presrves coequalizers. -/ -def preservesCoequalizersOfPreservesPushoutsAndBinaryCoproducts [HasBinaryCoproducts C] +lemma preservesCoequalizers_of_preservesPushouts_and_binaryCoproducts [HasBinaryCoproducts C] [HasPushouts C] [PreservesColimitsOfShape (Discrete WalkingPair) G] [PreservesColimitsOfShape WalkingSpan G] : PreservesColimitsOfShape WalkingParallelPair G := ⟨fun {K} => - preservesColimitOfPreservesColimitCocone (coequalizerCoconeIsColimit K) <| + preservesColimit_of_preserves_colimit_cocone (coequalizerCoconeIsColimit K) <| { desc := fun c => by refine (PreservesPushout.iso _ _ _).inv ≫ pushout.desc ?_ ?_ ?_ · exact c.ι.app WalkingParallelPair.one diff --git a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean index ffb743849fa51..629aa364f506a 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean @@ -117,17 +117,17 @@ variable [HasFiniteProducts.{v} C] /-- If `F` preserves the terminal object and binary products, then it preserves products indexed by `Fin n` for any `n`. -/ -noncomputable def preservesFinOfPreservesBinaryAndTerminal : +lemma preservesFinOfPreservesBinaryAndTerminal : ∀ (n : ℕ) (f : Fin n → C), PreservesLimit (Discrete.functor f) F | 0 => fun f => by letI : PreservesLimitsOfShape (Discrete (Fin 0)) F := - preservesLimitsOfShapeOfEquiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ + preservesLimitsOfShape_of_equiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ infer_instance | n + 1 => by haveI := preservesFinOfPreservesBinaryAndTerminal n intro f apply - preservesLimitOfPreservesLimitCone + preservesLimit_of_preserves_limit_cone (extendFanIsLimit f (limit.isLimit _) (limit.isLimit _)) _ apply (isLimitMapConeFanMkEquiv _ _ _).symm _ let this := @@ -148,20 +148,20 @@ noncomputable def preservesFinOfPreservesBinaryAndTerminal : /-- If `F` preserves the terminal object and binary products, then it preserves limits of shape `Discrete (Fin n)`. -/ -def preservesShapeFinOfPreservesBinaryAndTerminal (n : ℕ) : +lemma preservesShape_fin_of_preserves_binary_and_terminal (n : ℕ) : PreservesLimitsOfShape (Discrete (Fin n)) F where preservesLimit {K} := by let that : (Discrete.functor fun n => K.obj ⟨n⟩) ≅ K := Discrete.natIso fun ⟨i⟩ => Iso.refl _ haveI := preservesFinOfPreservesBinaryAndTerminal F n fun n => K.obj ⟨n⟩ - apply preservesLimitOfIsoDiagram F that + apply preservesLimit_of_iso_diagram F that /-- If `F` preserves the terminal object and binary products then it preserves finite products. -/ -def preservesFiniteProductsOfPreservesBinaryAndTerminal (J : Type*) [Fintype J] : +lemma preservesFiniteProducts_of_preserves_binary_and_terminal (J : Type*) [Fintype J] : PreservesLimitsOfShape (Discrete J) F := by classical let e := Fintype.equivFin J - haveI := preservesShapeFinOfPreservesBinaryAndTerminal F (Fintype.card J) - apply preservesLimitsOfShapeOfEquiv (Discrete.equivalence e).symm + haveI := preservesShape_fin_of_preserves_binary_and_terminal F (Fintype.card J) + apply preservesLimitsOfShape_of_equiv (Discrete.equivalence e).symm end Preserves @@ -249,17 +249,17 @@ variable [HasFiniteCoproducts.{v} C] /-- If `F` preserves the initial object and binary coproducts, then it preserves products indexed by `Fin n` for any `n`. -/ -noncomputable def preservesFinOfPreservesBinaryAndInitial : +lemma preserves_fin_of_preserves_binary_and_initial : ∀ (n : ℕ) (f : Fin n → C), PreservesColimit (Discrete.functor f) F | 0 => fun f => by letI : PreservesColimitsOfShape (Discrete (Fin 0)) F := - preservesColimitsOfShapeOfEquiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ + preservesColimitsOfShape_of_equiv.{0, 0} (Discrete.equivalence finZeroEquiv'.symm) _ infer_instance | n + 1 => by - haveI := preservesFinOfPreservesBinaryAndInitial n + haveI := preserves_fin_of_preserves_binary_and_initial n intro f apply - preservesColimitOfPreservesColimitCocone + preservesColimit_of_preserves_colimit_cocone (extendCofanIsColimit f (colimit.isColimit _) (colimit.isColimit _)) _ apply (isColimitMapCoconeCofanMkEquiv _ _ _).symm _ let this := @@ -279,20 +279,20 @@ noncomputable def preservesFinOfPreservesBinaryAndInitial : /-- If `F` preserves the initial object and binary coproducts, then it preserves colimits of shape `Discrete (Fin n)`. -/ -def preservesShapeFinOfPreservesBinaryAndInitial (n : ℕ) : +lemma preservesShape_fin_of_preserves_binary_and_initial (n : ℕ) : PreservesColimitsOfShape (Discrete (Fin n)) F where preservesColimit {K} := by let that : (Discrete.functor fun n => K.obj ⟨n⟩) ≅ K := Discrete.natIso fun ⟨i⟩ => Iso.refl _ - haveI := preservesFinOfPreservesBinaryAndInitial F n fun n => K.obj ⟨n⟩ - apply preservesColimitOfIsoDiagram F that + haveI := preserves_fin_of_preserves_binary_and_initial F n fun n => K.obj ⟨n⟩ + apply preservesColimit_of_iso_diagram F that /-- If `F` preserves the initial object and binary coproducts then it preserves finite products. -/ -def preservesFiniteCoproductsOfPreservesBinaryAndInitial (J : Type*) [Fintype J] : +lemma preservesFiniteCoproductsOfPreservesBinaryAndInitial (J : Type*) [Fintype J] : PreservesColimitsOfShape (Discrete J) F := by classical let e := Fintype.equivFin J - haveI := preservesShapeFinOfPreservesBinaryAndInitial F (Fintype.card J) - apply preservesColimitsOfShapeOfEquiv (Discrete.equivalence e).symm + haveI := preservesShape_fin_of_preserves_binary_and_initial F (Fintype.card J) + apply preservesColimitsOfShape_of_equiv (Discrete.equivalence e).symm end Preserves diff --git a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean index bd52c4c1ea3b7..353032099523c 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/LimitsOfProductsAndEqualizers.lean @@ -156,7 +156,8 @@ variable (G : C ⥤ D) [PreservesLimitsOfShape WalkingParallelPair G] [PreservesLimitsOfShape (Discrete.{w} (Σp : J × J, p.1 ⟶ p.2)) G] /-- If a functor preserves equalizers and the appropriate products, it preserves limits. -/ -noncomputable def preservesLimitOfPreservesEqualizersAndProduct : PreservesLimitsOfShape J G where +lemma preservesLimit_of_preservesEqualizers_and_product : + PreservesLimitsOfShape J G where preservesLimit {K} := by let P := ∏ᶜ K.obj let Q := ∏ᶜ fun f : Σp : J × J, p.fst ⟶ p.snd => K.obj f.1.2 @@ -164,8 +165,7 @@ noncomputable def preservesLimitOfPreservesEqualizersAndProduct : PreservesLimit let t : P ⟶ Q := Pi.lift fun f => limit.π (Discrete.functor K.obj) ⟨f.1.2⟩ let I := equalizer s t let i : I ⟶ P := equalizer.ι s t - apply - preservesLimitOfPreservesLimitCone + apply preservesLimit_of_preserves_limit_cone (buildIsLimit s t (by simp [s]) (by simp [t]) (limit.isLimit _) (limit.isLimit _) (limit.isLimit _)) apply IsLimit.ofIsoLimit (buildIsLimit _ _ _ _ _ _ _) _ @@ -202,24 +202,24 @@ with type Fintype J does not have forward dependencies, type class resolution ca use this kind of local instance because it will not be able to infer a value for this parameter." Factored out this as new class in `CategoryTheory.Limits.Preserves.Finite` -/ /-- If G preserves equalizers and finite products, it preserves finite limits. -/ -noncomputable def preservesFiniteLimitsOfPreservesEqualizersAndFiniteProducts [HasEqualizers C] +lemma preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts [HasEqualizers C] [HasFiniteProducts C] (G : C ⥤ D) [PreservesLimitsOfShape WalkingParallelPair G] [PreservesFiniteProducts G] : PreservesFiniteLimits G where preservesFiniteLimits := by intro J sJ fJ haveI : Fintype J := inferInstance haveI : Fintype ((p : J × J) × (p.fst ⟶ p.snd)) := inferInstance - apply @preservesLimitOfPreservesEqualizersAndProduct _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ + apply @preservesLimit_of_preservesEqualizers_and_product _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ · apply hasLimitsOfShape_discrete _ _ · apply hasLimitsOfShape_discrete _ · apply PreservesFiniteProducts.preserves _ · apply PreservesFiniteProducts.preserves _ /-- If G preserves equalizers and products, it preserves all limits. -/ -noncomputable def preservesLimitsOfPreservesEqualizersAndProducts [HasEqualizers C] +lemma preservesLimits_of_preservesEqualizers_and_products [HasEqualizers C] [HasProducts.{w} C] (G : C ⥤ D) [PreservesLimitsOfShape WalkingParallelPair G] [∀ J, PreservesLimitsOfShape (Discrete.{w} J) G] : PreservesLimitsOfSize.{w, w} G where - preservesLimitsOfShape := preservesLimitOfPreservesEqualizersAndProduct G + preservesLimitsOfShape := preservesLimit_of_preservesEqualizers_and_product G theorem hasFiniteLimits_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] : HasFiniteLimits C := @@ -230,18 +230,18 @@ theorem hasFiniteLimits_of_hasTerminal_and_pullbacks [HasTerminal C] [HasPullbac (hasBinaryProducts_of_hasTerminal_and_pullbacks C) inferInstance) /-- If G preserves terminal objects and pullbacks, it preserves all finite limits. -/ -noncomputable def preservesFiniteLimitsOfPreservesTerminalAndPullbacks [HasTerminal C] +lemma preservesFiniteLimits_of_preservesTerminal_and_pullbacks [HasTerminal C] [HasPullbacks C] (G : C ⥤ D) [PreservesLimitsOfShape (Discrete.{0} PEmpty) G] [PreservesLimitsOfShape WalkingCospan G] : PreservesFiniteLimits G := by haveI : HasFiniteLimits C := hasFiniteLimits_of_hasTerminal_and_pullbacks haveI : PreservesLimitsOfShape (Discrete WalkingPair) G := - preservesBinaryProductsOfPreservesTerminalAndPullbacks G + preservesBinaryProducts_of_preservesTerminal_and_pullbacks G haveI : PreservesLimitsOfShape WalkingParallelPair G := - preservesEqualizersOfPreservesPullbacksAndBinaryProducts G + preservesEqualizers_of_preservesPullbacks_and_binaryProducts G apply - @preservesFiniteLimitsOfPreservesEqualizersAndFiniteProducts _ _ _ _ _ _ G _ ?_ + @preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts _ _ _ _ _ _ G _ ?_ apply PreservesFiniteProducts.mk - apply preservesFiniteProductsOfPreservesBinaryAndTerminal G + apply preservesFiniteProducts_of_preserves_binary_and_terminal G /-! We now dualize the above constructions, resorting to copy-paste. @@ -372,7 +372,7 @@ variable (G : C ⥤ D) [PreservesColimitsOfShape WalkingParallelPair G] [PreservesColimitsOfShape (Discrete.{w} (Σp : J × J, p.1 ⟶ p.2)) G] /-- If a functor preserves coequalizers and the appropriate coproducts, it preserves colimits. -/ -noncomputable def preservesColimitOfPreservesCoequalizersAndCoproduct : +lemma preservesColimit_of_preservesCoequalizers_and_coproduct : PreservesColimitsOfShape J G where preservesColimit {K} := by let P := ∐ K.obj @@ -381,8 +381,7 @@ noncomputable def preservesColimitOfPreservesCoequalizersAndCoproduct : let t : Q ⟶ P := Sigma.desc fun f => colimit.ι (Discrete.functor K.obj) ⟨f.1.1⟩ let I := coequalizer s t let i : P ⟶ I := coequalizer.π s t - apply - preservesColimitOfPreservesColimitCocone + apply preservesColimit_of_preserves_colimit_cocone (buildIsColimit s t (by simp [s]) (by simp [t]) (colimit.isColimit _) (colimit.isColimit _) (colimit.isColimit _)) apply IsColimit.ofIsoColimit (buildIsColimit _ _ _ _ _ _ _) _ @@ -422,7 +421,7 @@ with type Fintype J does not have forward dependencies, type class resolution ca this kind of local instance because it will not be able to infer a value for this parameter." Factored out this as new class in `CategoryTheory.Limits.Preserves.Finite` -/ /-- If G preserves coequalizers and finite coproducts, it preserves finite colimits. -/ -noncomputable def preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoproducts +lemma preservesFiniteColimits_of_preservesCoequalizers_and_finiteCoproducts [HasCoequalizers C] [HasFiniteCoproducts C] (G : C ⥤ D) [PreservesColimitsOfShape WalkingParallelPair G] [PreservesFiniteCoproducts G] : PreservesFiniteColimits G where @@ -430,17 +429,17 @@ noncomputable def preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoprodu intro J sJ fJ haveI : Fintype J := inferInstance haveI : Fintype ((p : J × J) × (p.fst ⟶ p.snd)) := inferInstance - apply @preservesColimitOfPreservesCoequalizersAndCoproduct _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ + apply @preservesColimit_of_preservesCoequalizers_and_coproduct _ _ _ sJ _ _ ?_ ?_ _ G _ ?_ ?_ · apply hasColimitsOfShape_discrete _ _ · apply hasColimitsOfShape_discrete _ · apply PreservesFiniteCoproducts.preserves _ · apply PreservesFiniteCoproducts.preserves _ /-- If G preserves coequalizers and coproducts, it preserves all colimits. -/ -noncomputable def preservesColimitsOfPreservesCoequalizersAndCoproducts [HasCoequalizers C] +lemma preservesColimits_of_preservesCoequalizers_and_coproducts [HasCoequalizers C] [HasCoproducts.{w} C] (G : C ⥤ D) [PreservesColimitsOfShape WalkingParallelPair G] - [∀ J, PreservesColimitsOfShape (Discrete.{w} J) G] : PreservesColimitsOfSize.{w} G where - preservesColimitsOfShape := preservesColimitOfPreservesCoequalizersAndCoproduct G + [∀ J, PreservesColimitsOfShape (Discrete.{w} J) G] : PreservesColimitsOfSize.{w, w} G where + preservesColimitsOfShape := preservesColimit_of_preservesCoequalizers_and_coproduct G theorem hasFiniteColimits_of_hasInitial_and_pushouts [HasInitial C] [HasPushouts C] : HasFiniteColimits C := @@ -451,16 +450,16 @@ theorem hasFiniteColimits_of_hasInitial_and_pushouts [HasInitial C] [HasPushouts (hasBinaryCoproducts_of_hasInitial_and_pushouts C) inferInstance) /-- If G preserves initial objects and pushouts, it preserves all finite colimits. -/ -noncomputable def preservesFiniteColimitsOfPreservesInitialAndPushouts [HasInitial C] +lemma preservesFiniteColimits_of_preservesInitial_and_pushouts [HasInitial C] [HasPushouts C] (G : C ⥤ D) [PreservesColimitsOfShape (Discrete.{0} PEmpty) G] [PreservesColimitsOfShape WalkingSpan G] : PreservesFiniteColimits G := by haveI : HasFiniteColimits C := hasFiniteColimits_of_hasInitial_and_pushouts haveI : PreservesColimitsOfShape (Discrete WalkingPair) G := - preservesBinaryCoproductsOfPreservesInitialAndPushouts G + preservesBinaryCoproducts_of_preservesInitial_and_pushouts G haveI : PreservesColimitsOfShape (WalkingParallelPair) G := - (preservesCoequalizersOfPreservesPushoutsAndBinaryCoproducts G) + (preservesCoequalizers_of_preservesPushouts_and_binaryCoproducts G) refine - @preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoproducts _ _ _ _ _ _ G _ ?_ + @preservesFiniteColimits_of_preservesCoequalizers_and_finiteCoproducts _ _ _ _ _ _ G _ ?_ apply PreservesFiniteCoproducts.mk apply preservesFiniteCoproductsOfPreservesBinaryAndInitial G diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 3b22aa0a3b6f5..148eafbe8efa1 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -141,7 +141,7 @@ lemma liftedLimitMapsToOriginal_inv_map_π /-- The lifted cone is a limit. -/ def liftedLimitIsLimit {K : J ⥤ C} {F : C ⥤ D} [CreatesLimit K F] {c : Cone (K ⋙ F)} (t : IsLimit c) : IsLimit (liftLimit t) := - ReflectsLimit.reflects (IsLimit.ofIsoLimit t (liftedLimitMapsToOriginal t).symm) + isLimitOfReflects _ (IsLimit.ofIsoLimit t (liftedLimitMapsToOriginal t).symm) /-- If `F` creates the limit of `K` and `K ⋙ F` has a limit, then `K` has a limit. -/ theorem hasLimit_of_created (K : J ⥤ C) (F : C ⥤ D) [HasLimit (K ⋙ F)] [CreatesLimit K F] : @@ -176,7 +176,7 @@ def liftedColimitMapsToOriginal {K : J ⥤ C} {F : C ⥤ D} [CreatesColimit K F] /-- The lifted cocone is a colimit. -/ def liftedColimitIsColimit {K : J ⥤ C} {F : C ⥤ D} [CreatesColimit K F] {c : Cocone (K ⋙ F)} (t : IsColimit c) : IsColimit (liftColimit t) := - ReflectsColimit.reflects (IsColimit.ofIsoColimit t (liftedColimitMapsToOriginal t).symm) + isColimitOfReflects _ (IsColimit.ofIsoColimit t (liftedColimitMapsToOriginal t).symm) /-- If `F` creates the limit of `K` and `K ⋙ F` has a limit, then `K` has a limit. -/ theorem hasColimit_of_created (K : J ⥤ C) (F : C ⥤ D) [HasColimit (K ⋙ F)] [CreatesColimit K F] : @@ -239,7 +239,7 @@ def createsLimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphism (h : ∀ c t, LiftsToLimit K F c t) : CreatesLimit K F where lifts c t := (h c t).toLiftableCone toReflectsLimit := - { reflects := fun {d} hd => by + { reflects := fun {d} hd => ⟨by let d' : Cone K := (h (F.mapCone d) hd).toLiftableCone.liftedCone let i : F.mapCone d' ≅ F.mapCone d := (h (F.mapCone d) hd).toLiftableCone.validLift @@ -251,7 +251,7 @@ def createsLimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphism rw [this] infer_instance haveI : IsIso f := isIso_of_reflects_iso f (Cones.functoriality K F) - exact IsLimit.ofIsoLimit hd' (asIso f).symm } + exact IsLimit.ofIsoLimit hd' (asIso f).symm⟩ } /-- If `F` reflects isomorphisms and we can lift a single limit cone to a limit cone, then `F` creates limits. Note that unlike `createsLimitOfReflectsIso`, to apply this result it is @@ -322,27 +322,44 @@ def createsLimitOfFullyFaithfulOfIso {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Fai /-- A fully faithful functor that preserves a limit that exists also creates the limit. -/ def createsLimitOfFullyFaithfulOfPreserves {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] [HasLimit K] [PreservesLimit K F] : CreatesLimit K F := - createsLimitOfFullyFaithfulOfLift' (PreservesLimit.preserves <| limit.isLimit K) _ (Iso.refl _) + createsLimitOfFullyFaithfulOfLift' (isLimitOfPreserves _ (limit.isLimit K)) _ (Iso.refl _) -- see Note [lower instance priority] /-- `F` preserves the limit of `K` if it creates the limit and `K ⋙ F` has the limit. -/ -instance (priority := 100) preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) preservesLimit_of_createsLimit_and_hasLimit (K : J ⥤ C) (F : C ⥤ D) [CreatesLimit K F] [HasLimit (K ⋙ F)] : PreservesLimit K F where - preserves t := IsLimit.ofIsoLimit (limit.isLimit _) + preserves t := ⟨IsLimit.ofIsoLimit (limit.isLimit _) ((liftedLimitMapsToOriginal (limit.isLimit _)).symm ≪≫ - (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t)) + (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t))⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) + [CreatesLimit K F] [HasLimit (K ⋙ F)] : PreservesLimit K F := + preservesLimit_of_createsLimit_and_hasLimit _ _ -- see Note [lower instance priority] /-- `F` preserves the limit of shape `J` if it creates these limits and `D` has them. -/ -instance (priority := 100) preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape +instance (priority := 100) preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F where +@[deprecated (since := "2024-11-19")] +lemma preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape + (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : + PreservesLimitsOfShape J F := + preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape _ + -- see Note [lower instance priority] /-- `F` preserves limits if it creates limits and `D` has limits. -/ -instance (priority := 100) preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) +instance (priority := 100) preservesLimits_of_createsLimits_and_hasLimits (F : C ⥤ D) [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F where +@[deprecated (since := "2024-11-19")] +lemma preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) + [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : + PreservesLimitsOfSize.{w, w'} F := + preservesLimits_of_createsLimits_and_hasLimits _ + /-- If `F` reflects isomorphisms and we can lift any colimit cocone to a colimit cocone, then `F` creates colimits. In particular here we don't need to assume that F reflects colimits. @@ -351,8 +368,7 @@ def createsColimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphi (h : ∀ c t, LiftsToColimit K F c t) : CreatesColimit K F where lifts c t := (h c t).toLiftableCocone toReflectsColimit := - { - reflects := fun {d} hd => by + { reflects := fun {d} hd => ⟨by let d' : Cocone K := (h (F.mapCocone d) hd).toLiftableCocone.liftedCocone let i : F.mapCocone d' ≅ F.mapCocone d := (h (F.mapCocone d) hd).toLiftableCocone.validLift @@ -364,7 +380,7 @@ def createsColimitOfReflectsIso {K : J ⥤ C} {F : C ⥤ D} [F.ReflectsIsomorphi rw [this] infer_instance haveI := isIso_of_reflects_iso f (Cocones.functoriality K F) - exact IsColimit.ofIsoColimit hd' (asIso f) } + exact IsColimit.ofIsoColimit hd' (asIso f)⟩ } /-- If `F` reflects isomorphisms and we can lift a single colimit cocone to a colimit cocone, then `F` creates limits. Note that unlike `createsColimitOfReflectsIso`, to apply this result it is @@ -407,7 +423,7 @@ def createsColimitOfFullyFaithfulOfLift {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F. /-- A fully faithful functor that preserves a colimit that exists also creates the colimit. -/ def createsColimitOfFullyFaithfulOfPreserves {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.Faithful] [HasColimit K] [PreservesColimit K F] : CreatesColimit K F := - createsColimitOfFullyFaithfulOfLift' (PreservesColimit.preserves <| colimit.isColimit K) _ + createsColimitOfFullyFaithfulOfLift' (isColimitOfPreserves _ (colimit.isColimit K)) _ (Iso.refl _) -- Notice however that even if the isomorphism is `Iso.refl _`, @@ -442,30 +458,47 @@ def createsColimitOfFullyFaithfulOfIso {K : J ⥤ C} {F : C ⥤ D} [F.Full] [F.F -- see Note [lower instance priority] /-- `F` preserves the colimit of `K` if it creates the colimit and `K ⋙ F` has the colimit. -/ -instance (priority := 100) preservesColimitOfCreatesColimitAndHasColimit (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) preservesColimit_of_createsColimit_and_hasColimit (K : J ⥤ C) (F : C ⥤ D) [CreatesColimit K F] [HasColimit (K ⋙ F)] : PreservesColimit K F where preserves t := - IsColimit.ofIsoColimit (colimit.isColimit _) + ⟨IsColimit.ofIsoColimit (colimit.isColimit _) ((liftedColimitMapsToOriginal (colimit.isColimit _)).symm ≪≫ (Cocones.functoriality K F).mapIso - ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t)) + ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t))⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesColimitOfCreatesColimitAndHasColimit (K : J ⥤ C) (F : C ⥤ D) + [CreatesColimit K F] [HasColimit (K ⋙ F)] : PreservesColimit K F := + preservesColimit_of_createsColimit_and_hasColimit _ _ -- see Note [lower instance priority] /-- `F` preserves the colimit of shape `J` if it creates these colimits and `D` has them. -/ -instance (priority := 100) preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape +instance (priority := 100) preservesColimitOfShape_of_createsColimitsOfShape_and_hasColimitsOfShape (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F where +@[deprecated (since := "2024-11-19")] +lemma preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape + (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : + PreservesColimitsOfShape J F := + preservesColimitOfShape_of_createsColimitsOfShape_and_hasColimitsOfShape _ + -- see Note [lower instance priority] /-- `F` preserves limits if it creates limits and `D` has limits. -/ -instance (priority := 100) preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) +instance (priority := 100) preservesColimits_of_createsColimits_and_hasColimits (F : C ⥤ D) [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F where +@[deprecated (since := "2024-11-19")] +lemma preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) + [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : + PreservesColimitsOfSize.{w, w'} F := + preservesColimits_of_createsColimits_and_hasColimits _ + /-- Transfer creation of limits along a natural isomorphism in the diagram. -/ def createsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [CreatesLimit K₁ F] : CreatesLimit K₂ F := - { reflectsLimitOfIsoDiagram F h with + { reflectsLimit_of_iso_diagram F h with lifts := fun c t => let t' := (IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) c).symm t { liftedCone := (Cones.postcompose h.hom).obj (liftLimit t') @@ -485,7 +518,7 @@ def createsLimitOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimit K F] : Crea refine (IsLimit.mapConeEquiv h ?_).uniqueUpToIso t apply IsLimit.ofIsoLimit _ (liftedLimitMapsToOriginal _).symm apply (IsLimit.postcomposeInvEquiv _ _).symm t } - toReflectsLimit := reflectsLimitOfNatIso _ h + toReflectsLimit := reflectsLimit_of_natIso _ h /-- If `F` creates limits of shape `J` and `F ≅ G`, then `G` creates limits of shape `J`. -/ def createsLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimitsOfShape J F] : @@ -499,7 +532,7 @@ def createsLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesLimitsOfSize.{w, /-- Transfer creation of colimits along a natural isomorphism in the diagram. -/ def createsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [CreatesColimit K₁ F] : CreatesColimit K₂ F := - { reflectsColimitOfIsoDiagram F h with + { reflectsColimit_of_iso_diagram F h with lifts := fun c t => let t' := (IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) c).symm t { liftedCocone := (Cocones.precompose h.inv).obj (liftColimit t') @@ -520,7 +553,7 @@ def createsColimitOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesColimit K F] : refine (IsColimit.mapCoconeEquiv h ?_).uniqueUpToIso t apply IsColimit.ofIsoColimit _ (liftedColimitMapsToOriginal _).symm apply (IsColimit.precomposeHomEquiv _ _).symm t } - toReflectsColimit := reflectsColimitOfNatIso _ h + toReflectsColimit := reflectsColimit_of_natIso _ h /-- If `F` creates colimits of shape `J` and `F ≅ G`, then `G` creates colimits of shape `J`. -/ def createsColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [CreatesColimitsOfShape J F] : diff --git a/Mathlib/CategoryTheory/Limits/ExactFunctor.lean b/Mathlib/CategoryTheory/Limits/ExactFunctor.lean index cb283bb686903..811252d88fc96 100644 --- a/Mathlib/CategoryTheory/Limits/ExactFunctor.lean +++ b/Mathlib/CategoryTheory/Limits/ExactFunctor.lean @@ -31,7 +31,7 @@ variable (C) (D) -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled left-exact functors. -/ def LeftExactFunctor := - FullSubcategory fun F : C ⥤ D => Nonempty (PreservesFiniteLimits F) + FullSubcategory fun F : C ⥤ D => PreservesFiniteLimits F instance : Category (LeftExactFunctor C D) := FullSubcategory.category _ @@ -52,7 +52,7 @@ instance : (LeftExactFunctor.forget C D).Faithful := -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] /-- Bundled right-exact functors. -/ def RightExactFunctor := - FullSubcategory fun F : C ⥤ D => Nonempty (PreservesFiniteColimits F) + FullSubcategory fun F : C ⥤ D => PreservesFiniteColimits F instance : Category (RightExactFunctor C D) := FullSubcategory.category _ @@ -74,7 +74,7 @@ instance : (RightExactFunctor.forget C D).Faithful := /-- Bundled exact functors. -/ def ExactFunctor := FullSubcategory fun F : C ⥤ D => - Nonempty (PreservesFiniteLimits F) ∧ Nonempty (PreservesFiniteColimits F) + PreservesFiniteLimits F ∧ PreservesFiniteColimits F instance : Category (ExactFunctor C D) := FullSubcategory.category _ @@ -162,15 +162,15 @@ theorem ExactFunctor.forget_map {F G : C ⥤ₑ D} (α : F ⟶ G) : (ExactFuncto /-- Turn a left exact functor into an object of the category `LeftExactFunctor C D`. -/ def LeftExactFunctor.of (F : C ⥤ D) [PreservesFiniteLimits F] : C ⥤ₗ D := - ⟨F, ⟨inferInstance⟩⟩ + ⟨F, inferInstance⟩ /-- Turn a right exact functor into an object of the category `RightExactFunctor C D`. -/ def RightExactFunctor.of (F : C ⥤ D) [PreservesFiniteColimits F] : C ⥤ᵣ D := - ⟨F, ⟨inferInstance⟩⟩ + ⟨F, inferInstance⟩ /-- Turn an exact functor into an object of the category `ExactFunctor C D`. -/ def ExactFunctor.of (F : C ⥤ D) [PreservesFiniteLimits F] [PreservesFiniteColimits F] : C ⥤ₑ D := - ⟨F, ⟨⟨inferInstance⟩, ⟨inferInstance⟩⟩⟩ + ⟨F, ⟨inferInstance, inferInstance⟩⟩ @[simp] theorem LeftExactFunctor.of_fst (F : C ⥤ D) [PreservesFiniteLimits F] : @@ -200,16 +200,16 @@ theorem ExactFunctor.forget_obj_of (F : C ⥤ D) [PreservesFiniteLimits F] rfl noncomputable instance (F : C ⥤ₗ D) : PreservesFiniteLimits F.obj := - F.property.some + F.property noncomputable instance (F : C ⥤ᵣ D) : PreservesFiniteColimits F.obj := - F.property.some + F.property noncomputable instance (F : C ⥤ₑ D) : PreservesFiniteLimits F.obj := - F.property.1.some + F.property.1 noncomputable instance (F : C ⥤ₑ D) : PreservesFiniteColimits F.obj := - F.property.2.some + F.property.2 variable {E : Type u₃} [Category.{v₃} E] @@ -221,7 +221,7 @@ variable (C D E) @[simps!] def LeftExactFunctor.whiskeringLeft : (C ⥤ₗ D) ⥤ (D ⥤ₗ E) ⥤ (C ⥤ₗ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringLeft C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteLimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteLimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringLeft C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringLeft C D E).map η).naturality f } @@ -236,7 +236,7 @@ def LeftExactFunctor.whiskeringLeft : (C ⥤ₗ D) ⥤ (D ⥤ₗ E) ⥤ (C ⥤ @[simps!] def LeftExactFunctor.whiskeringRight : (D ⥤ₗ E) ⥤ (C ⥤ₗ D) ⥤ (C ⥤ₗ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringRight C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteLimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteLimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringRight C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringRight C D E).map η).naturality f } @@ -251,7 +251,7 @@ def LeftExactFunctor.whiskeringRight : (D ⥤ₗ E) ⥤ (C ⥤ₗ D) ⥤ (C ⥤ @[simps!] def RightExactFunctor.whiskeringLeft : (C ⥤ᵣ D) ⥤ (D ⥤ᵣ E) ⥤ (C ⥤ᵣ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringLeft C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteColimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringLeft C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringLeft C D E).map η).naturality f } @@ -266,7 +266,7 @@ def RightExactFunctor.whiskeringLeft : (C ⥤ᵣ D) ⥤ (D ⥤ᵣ E) ⥤ (C ⥤ @[simps!] def RightExactFunctor.whiskeringRight : (D ⥤ᵣ E) ⥤ (C ⥤ᵣ D) ⥤ (C ⥤ᵣ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringRight C D E).obj F.obj) - (fun G => ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩) + (fun G => by dsimp; exact comp_preservesFiniteColimits _ _) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringRight C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringRight C D E).map η).naturality f } @@ -281,8 +281,8 @@ def RightExactFunctor.whiskeringRight : (D ⥤ᵣ E) ⥤ (C ⥤ᵣ D) ⥤ (C ⥤ @[simps!] def ExactFunctor.whiskeringLeft : (C ⥤ₑ D) ⥤ (D ⥤ₑ E) ⥤ (C ⥤ₑ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringLeft C D E).obj F.obj) - (fun G => ⟨⟨by dsimp; exact compPreservesFiniteLimits _ _⟩, - ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩⟩) + (fun G => ⟨by dsimp; exact comp_preservesFiniteLimits _ _, + by dsimp; exact comp_preservesFiniteColimits _ _⟩) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringLeft C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringLeft C D E).map η).naturality f } @@ -297,8 +297,8 @@ def ExactFunctor.whiskeringLeft : (C ⥤ₑ D) ⥤ (D ⥤ₑ E) ⥤ (C ⥤ₑ E) @[simps!] def ExactFunctor.whiskeringRight : (D ⥤ₑ E) ⥤ (C ⥤ₑ D) ⥤ (C ⥤ₑ E) where obj F := FullSubcategory.lift _ (forget _ _ ⋙ (CategoryTheory.whiskeringRight C D E).obj F.obj) - (fun G => ⟨⟨by dsimp; exact compPreservesFiniteLimits _ _⟩, - ⟨by dsimp; exact compPreservesFiniteColimits _ _⟩⟩) + (fun G => ⟨by dsimp; exact comp_preservesFiniteLimits _ _, + by dsimp; exact comp_preservesFiniteColimits _ _⟩) map {F G} η := { app := fun H => ((CategoryTheory.whiskeringRight C D E).map η).app H.obj naturality := fun _ _ f => ((CategoryTheory.whiskeringRight C D E).map η).naturality f } diff --git a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean index 1702595d0a0b5..494fe94ffdb41 100644 --- a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean +++ b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesFiniteLimit.lean @@ -338,11 +338,11 @@ instance colimitLimitToLimitColimitCone_iso (F : J ⥤ K ⥤ Type v) : infer_instance apply Cones.cone_iso_of_hom_iso -noncomputable instance filteredColimPreservesFiniteLimitsOfTypes : +noncomputable instance filtered_colim_preservesFiniteLimits_of_types : PreservesFiniteLimits (colim : (K ⥤ Type v) ⥤ _) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{v₂} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{v₂} intro J _ _ - refine ⟨fun {F} => ⟨fun {c} hc => IsLimit.ofIsoLimit (limit.isLimit _) ?_⟩⟩ + refine ⟨fun {F} => ⟨fun {c} hc => ⟨IsLimit.ofIsoLimit (limit.isLimit _) ?_⟩⟩⟩ symm trans colim.mapCone (limit.cone F) · exact Functor.mapIso _ (hc.uniqueUpToIso (limit.isLimit F)) @@ -356,20 +356,20 @@ variable [HasLimitsOfShape J C] [HasColimitsOfShape K C] variable [ReflectsLimitsOfShape J (forget C)] [PreservesColimitsOfShape K (forget C)] variable [PreservesLimitsOfShape J (forget C)] -noncomputable instance filteredColimPreservesFiniteLimits : +noncomputable instance filtered_colim_preservesFiniteLimits : PreservesLimitsOfShape J (colim : (K ⥤ C) ⥤ _) := haveI : PreservesLimitsOfShape J ((colim : (K ⥤ C) ⥤ _) ⋙ forget C) := - preservesLimitsOfShapeOfNatIso (preservesColimitNatIso _).symm - preservesLimitsOfShapeOfReflectsOfPreserves _ (forget C) + preservesLimitsOfShape_of_natIso (preservesColimitNatIso _).symm + preservesLimitsOfShape_of_reflects_of_preserves _ (forget C) end -attribute [local instance] reflectsLimitsOfShapeOfReflectsIsomorphisms +attribute [local instance] reflectsLimitsOfShape_of_reflectsIsomorphisms noncomputable instance [PreservesFiniteLimits (forget C)] [PreservesColimitsOfShape K (forget C)] [HasFiniteLimits C] [HasColimitsOfShape K C] [(forget C).ReflectsIsomorphisms] : PreservesFiniteLimits (colim : (K ⥤ C) ⥤ _) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{v} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{v} intro J _ _ infer_instance diff --git a/Mathlib/CategoryTheory/Limits/FintypeCat.lean b/Mathlib/CategoryTheory/Limits/FintypeCat.lean index d856af586c75e..a326c743aabf4 100644 --- a/Mathlib/CategoryTheory/Limits/FintypeCat.lean +++ b/Mathlib/CategoryTheory/Limits/FintypeCat.lean @@ -51,14 +51,14 @@ instance {J : Type} [SmallCategory J] [FinCategory J] : HasLimitsOfShape J Finty instance hasFiniteLimits : HasFiniteLimits FintypeCat.{u} where out _ := inferInstance -noncomputable instance inclusionPreservesFiniteLimits : +noncomputable instance inclusion_preservesFiniteLimits : PreservesFiniteLimits FintypeCat.incl.{u} where preservesFiniteLimits _ := - preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape FintypeCat.incl + preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape FintypeCat.incl /- Help typeclass inference to infer preservation of finite limits for the forgtful functor. -/ noncomputable instance : PreservesFiniteLimits (forget FintypeCat) := - FintypeCat.inclusionPreservesFiniteLimits + FintypeCat.inclusion_preservesFiniteLimits /-- The categorical product of a finite family in `FintypeCat` is equivalent to the product as types. -/ @@ -112,14 +112,14 @@ instance {J : Type} [SmallCategory J] [FinCategory J] : HasColimitsOfShape J Fin instance hasFiniteColimits : HasFiniteColimits FintypeCat.{u} where out _ := inferInstance -noncomputable instance inclusionPreservesFiniteColimits : +noncomputable instance inclusion_preservesFiniteColimits : PreservesFiniteColimits FintypeCat.incl.{u} where preservesFiniteColimits _ := - preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape FintypeCat.incl + preservesColimitOfShape_of_createsColimitsOfShape_and_hasColimitsOfShape FintypeCat.incl /- Help typeclass inference to infer preservation of finite colimits for the forgtful functor. -/ noncomputable instance : PreservesFiniteColimits (forget FintypeCat) := - FintypeCat.inclusionPreservesFiniteColimits + FintypeCat.inclusion_preservesFiniteColimits lemma jointly_surjective {J : Type*} [Category J] [FinCategory J] (F : J ⥤ FintypeCat.{u}) (t : Cocone F) (h : IsColimit t) (x : t.pt) : diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 86b8cb71e4d54..357916b1fe01e 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -12,12 +12,12 @@ import Mathlib.CategoryTheory.Limits.Preserves.Limits We show that if `D` has limits, then the functor category `C ⥤ D` also has limits (`CategoryTheory.Limits.functorCategoryHasLimits`), and the evaluation functors preserve limits -(`CategoryTheory.Limits.evaluationPreservesLimits`) +(`CategoryTheory.Limits.evaluation_preservesLimits`) (and similarly for colimits). We also show that `F : D ⥤ K ⥤ C` preserves (co)limits if it does so for each `k : K` -(`CategoryTheory.Limits.preservesLimitsOfEvaluation` and -`CategoryTheory.Limits.preservesColimitsOfEvaluation`). +(`CategoryTheory.Limits.preservesLimits_of_evaluation` and +`CategoryTheory.Limits.preservesColimits_of_evaluation`). -/ @@ -178,14 +178,14 @@ instance hasLimitCompEvalution (F : J ⥤ K ⥤ C) (k : K) [HasLimit (F.flip.obj HasLimit (F ⋙ (evaluation _ _).obj k) := hasLimitOfIso (F := F.flip.obj k) (Iso.refl _) -instance evaluationPreservesLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] (k : K) : +instance evaluation_preservesLimit (F : J ⥤ K ⥤ C) [∀ k, HasLimit (F.flip.obj k)] (k : K) : PreservesLimit F ((evaluation K C).obj k) := -- Porting note: added a let because X was not inferred let X : (k : K) → LimitCone (F.flip.obj k) := fun k => getLimitCone (F.flip.obj k) - preservesLimitOfPreservesLimitCone (combinedIsLimit _ X) <| + preservesLimit_of_preserves_limit_cone (combinedIsLimit _ X) <| IsLimit.ofIsoLimit (limit.isLimit _) (evaluateCombinedCones F X k).symm -instance evaluationPreservesLimitsOfShape [HasLimitsOfShape J C] (k : K) : +instance evaluation_preservesLimitsOfShape [HasLimitsOfShape J C] (k : K) : PreservesLimitsOfShape J ((evaluation K C).obj k) where preservesLimit := inferInstance @@ -263,14 +263,14 @@ instance hasColimitCompEvaluation (F : J ⥤ K ⥤ C) (k : K) [HasColimit (F.fli HasColimit (F ⋙ (evaluation _ _).obj k) := hasColimitOfIso (F := F.flip.obj k) (Iso.refl _) -instance evaluationPreservesColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] (k : K) : +instance evaluation_preservesColimit (F : J ⥤ K ⥤ C) [∀ k, HasColimit (F.flip.obj k)] (k : K) : PreservesColimit F ((evaluation K C).obj k) := -- Porting note: added a let because X was not inferred let X : (k : K) → ColimitCocone (F.flip.obj k) := fun k => getColimitCocone (F.flip.obj k) - preservesColimitOfPreservesColimitCocone (combinedIsColimit _ X) <| + preservesColimit_of_preserves_colimit_cocone (combinedIsColimit _ X) <| IsColimit.ofIsoColimit (colimit.isColimit _) (evaluateCombinedCocones F X k).symm -instance evaluationPreservesColimitsOfShape [HasColimitsOfShape J C] (k : K) : +instance evaluation_preservesColimitsOfShape [HasColimitsOfShape J C] (k : K) : PreservesColimitsOfShape J ((evaluation K C).obj k) where preservesColimit := inferInstance @@ -353,65 +353,88 @@ instance evaluationPreservesLimits [HasLimits C] (k : K) : preservesLimitsOfShape {_} _𝒥 := inferInstance /-- `F : D ⥤ K ⥤ C` preserves the limit of some `G : J ⥤ D` if it does for each `k : K`. -/ -def preservesLimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) +lemma preservesLimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) (H : ∀ k : K, PreservesLimit G (F ⋙ (evaluation K C).obj k : D ⥤ C)) : PreservesLimit G F := - ⟨fun {c} hc => by + ⟨fun {c} hc => ⟨by apply evaluationJointlyReflectsLimits intro X haveI := H X change IsLimit ((F ⋙ (evaluation K C).obj X).mapCone c) - exact PreservesLimit.preserves hc⟩ + exact isLimitOfPreserves _ hc⟩⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) + (H : ∀ k : K, PreservesLimit G (F ⋙ (evaluation K C).obj k : D ⥤ C)) : + PreservesLimit G F := + preservesLimit_of_evaluation _ _ H /-- `F : D ⥤ K ⥤ C` preserves limits of shape `J` if it does for each `k : K`. -/ -def preservesLimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] +lemma preservesLimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] (_ : ∀ k : K, PreservesLimitsOfShape J (F ⋙ (evaluation K C).obj k)) : PreservesLimitsOfShape J F := - ⟨fun {G} => preservesLimitOfEvaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ + ⟨fun {G} => preservesLimit_of_evaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] + (H : ∀ k : K, PreservesLimitsOfShape J (F ⋙ (evaluation K C).obj k)) : + PreservesLimitsOfShape J F := + preservesLimitsOfShape_of_evaluation _ _ H /-- `F : D ⥤ K ⥤ C` preserves all limits if it does for each `k : K`. -/ -def preservesLimitsOfEvaluation (F : D ⥤ K ⥤ C) +lemma preservesLimits_of_evaluation (F : D ⥤ K ⥤ C) (_ : ∀ k : K, PreservesLimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : PreservesLimitsOfSize.{w', w} F := ⟨fun {L} _ => - preservesLimitsOfShapeOfEvaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ + preservesLimitsOfShape_of_evaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesLimitsOfEvaluation (F : D ⥤ K ⥤ C) + (H : ∀ k : K, PreservesLimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : + PreservesLimitsOfSize.{w', w} F := + preservesLimits_of_evaluation _ H /-- The constant functor `C ⥤ (D ⥤ C)` preserves limits. -/ -instance preservesLimitsConst : PreservesLimitsOfSize.{w', w} (const D : C ⥤ _) := - preservesLimitsOfEvaluation _ fun _ => - preservesLimitsOfNatIso <| Iso.symm <| constCompEvaluationObj _ _ +instance preservesLimits_const : PreservesLimitsOfSize.{w', w} (const D : C ⥤ _) := + preservesLimits_of_evaluation _ fun _ => + preservesLimits_of_natIso <| Iso.symm <| constCompEvaluationObj _ _ -instance evaluationPreservesColimits [HasColimits C] (k : K) : +instance evaluation_preservesColimits [HasColimits C] (k : K) : PreservesColimits ((evaluation K C).obj k) where preservesColimitsOfShape := inferInstance /-- `F : D ⥤ K ⥤ C` preserves the colimit of some `G : J ⥤ D` if it does for each `k : K`. -/ -def preservesColimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) +lemma preservesColimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) (H : ∀ k, PreservesColimit G (F ⋙ (evaluation K C).obj k)) : PreservesColimit G F := - ⟨fun {c} hc => by + ⟨fun {c} hc => ⟨by apply evaluationJointlyReflectsColimits intro X haveI := H X change IsColimit ((F ⋙ (evaluation K C).obj X).mapCocone c) - exact PreservesColimit.preserves hc⟩ + exact isColimitOfPreserves _ hc⟩⟩ + +@[deprecated (since := "2024-11-19")] +lemma preservesColimitOfEvaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) + (H : ∀ k, PreservesColimit G (F ⋙ (evaluation K C).obj k)) : PreservesColimit G F := + preservesColimit_of_evaluation _ _ H /-- `F : D ⥤ K ⥤ C` preserves all colimits of shape `J` if it does for each `k : K`. -/ -def preservesColimitsOfShapeOfEvaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] +lemma preservesColimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Category J] (_ : ∀ k : K, PreservesColimitsOfShape J (F ⋙ (evaluation K C).obj k)) : PreservesColimitsOfShape J F := - ⟨fun {G} => preservesColimitOfEvaluation F G fun _ => PreservesColimitsOfShape.preservesColimit⟩ + ⟨fun {G} => preservesColimit_of_evaluation F G fun _ => PreservesColimitsOfShape.preservesColimit⟩ /-- `F : D ⥤ K ⥤ C` preserves all colimits if it does for each `k : K`. -/ -def preservesColimitsOfEvaluation (F : D ⥤ K ⥤ C) +lemma preservesColimits_of_evaluation (F : D ⥤ K ⥤ C) (_ : ∀ k : K, PreservesColimitsOfSize.{w', w} (F ⋙ (evaluation K C).obj k)) : PreservesColimitsOfSize.{w', w} F := ⟨fun {L} _ => - preservesColimitsOfShapeOfEvaluation F L fun _ => + preservesColimitsOfShape_of_evaluation F L fun _ => PreservesColimitsOfSize.preservesColimitsOfShape⟩ /-- The constant functor `C ⥤ (D ⥤ C)` preserves colimits. -/ -instance preservesColimitsConst : PreservesColimitsOfSize.{w', w} (const D : C ⥤ _) := - preservesColimitsOfEvaluation _ fun _ => - preservesColimitsOfNatIso <| Iso.symm <| constCompEvaluationObj _ _ +instance preservesColimits_const : PreservesColimitsOfSize.{w', w} (const D : C ⥤ _) := + preservesColimits_of_evaluation _ fun _ => + preservesColimits_of_natIso <| Iso.symm <| constCompEvaluationObj _ _ open CategoryTheory.prod diff --git a/Mathlib/CategoryTheory/Limits/Over.lean b/Mathlib/CategoryTheory/Limits/Over.lean index 34e23b9622e0b..e6a0fb7ce116d 100644 --- a/Mathlib/CategoryTheory/Limits/Over.lean +++ b/Mathlib/CategoryTheory/Limits/Over.lean @@ -63,9 +63,9 @@ instance createsColimitsOfSizeMapCompForget {Y : C} (f : X ⟶ Y) : CreatesColimitsOfSize.{w, w'} (map f ⋙ forget Y) := show CreatesColimitsOfSize.{w, w'} (forget X) from inferInstance -instance preservesColimitsOfSizeMap [HasColimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : +instance preservesColimitsOfSize_map [HasColimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : PreservesColimitsOfSize.{w, w'} (map f) := - preservesColimitsOfReflectsOfPreserves (map f) (forget Y) + preservesColimits_of_reflects_of_preserves (map f) (forget Y) /-- If `c` is a colimit cocone, then so is the cocone `c.toOver` with cocone point `𝟙 c.pt`. -/ def isColimitToOver {F : J ⥤ C} {c : Cocone F} (hc : IsColimit c) : IsColimit c.toOver := @@ -110,9 +110,9 @@ instance createLimitsOfSizeMapCompForget {Y : C} (f : X ⟶ Y) : CreatesLimitsOfSize.{w, w'} (map f ⋙ forget X) := show CreatesLimitsOfSize.{w, w'} (forget Y) from inferInstance -instance preservesLimitsOfSizeMap [HasLimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : +instance preservesLimitsOfSize_map [HasLimitsOfSize.{w, w'} C] {Y : C} (f : X ⟶ Y) : PreservesLimitsOfSize.{w, w'} (map f) := - preservesLimitsOfReflectsOfPreserves (map f) (forget X) + preservesLimits_of_reflects_of_preserves (map f) (forget X) /-- If `c` is a limit cone, then so is the cone `c.toUnder` with cone point `𝟙 c.pt`. -/ def isLimitToUnder {F : J ⥤ C} {c : Cone F} (hc : IsLimit c) : IsLimit c.toUnder := diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean index f5abc3a11449f..c77df5a351b5d 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Basic.lean @@ -49,30 +49,30 @@ variable {J : Type w} [Category.{w'} J] {K : J ⥤ C} /-- A functor `F` preserves limits of `K` (written as `PreservesLimit K F`) if `F` maps any limit cone over `K` to a limit cone. -/ -class PreservesLimit (K : J ⥤ C) (F : C ⥤ D) where - preserves : ∀ {c : Cone K}, IsLimit c → IsLimit (F.mapCone c) +class PreservesLimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + preserves {c : Cone K} (hc : IsLimit c) : Nonempty (IsLimit (F.mapCone c)) /-- A functor `F` preserves colimits of `K` (written as `PreservesColimit K F`) if `F` maps any colimit cocone over `K` to a colimit cocone. -/ -class PreservesColimit (K : J ⥤ C) (F : C ⥤ D) where - preserves : ∀ {c : Cocone K}, IsColimit c → IsColimit (F.mapCocone c) +class PreservesColimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + preserves {c : Cocone K} (hc : IsColimit c) : Nonempty (IsColimit (F.mapCocone c)) /-- We say that `F` preserves limits of shape `J` if `F` preserves limits for every diagram `K : J ⥤ C`, i.e., `F` maps limit cones over `K` to limit cones. -/ -class PreservesLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class PreservesLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where preservesLimit : ∀ {K : J ⥤ C}, PreservesLimit K F := by infer_instance /-- We say that `F` preserves colimits of shape `J` if `F` preserves colimits for every diagram `K : J ⥤ C`, i.e., `F` maps colimit cocones over `K` to colimit cocones. -/ -class PreservesColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class PreservesColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where preservesColimit : ∀ {K : J ⥤ C}, PreservesColimit K F := by infer_instance -- This should be used with explicit universe variables. /-- `PreservesLimitsOfSize.{v u} F` means that `F` sends all limit cones over any diagram `J ⥤ C` to limit cones, where `J : Type u` with `[Category.{v} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesLimitsOfSize (F : C ⥤ D) where +class PreservesLimitsOfSize (F : C ⥤ D) : Prop where preservesLimitsOfShape : ∀ {J : Type w} [Category.{w'} J], PreservesLimitsOfShape J F := by infer_instance @@ -85,7 +85,7 @@ abbrev PreservesLimits (F : C ⥤ D) := /-- `PreservesColimitsOfSize.{v u} F` means that `F` sends all colimit cocones over any diagram `J ⥤ C` to colimit cocones, where `J : Type u` with `[Category.{v} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesColimitsOfSize (F : C ⥤ D) where +class PreservesColimitsOfSize (F : C ⥤ D) : Prop where preservesColimitsOfShape : ∀ {J : Type w} [Category.{w'} J], PreservesColimitsOfShape J F := by infer_instance @@ -106,7 +106,7 @@ guide typeclass resolution. -/ def isLimitOfPreserves (F : C ⥤ D) {c : Cone K} (t : IsLimit c) [PreservesLimit K F] : IsLimit (F.mapCone c) := - PreservesLimit.preserves t + (PreservesLimit.preserves t).some /-- A convenience function for `PreservesColimit`, which takes the functor as an explicit argument to @@ -114,33 +114,33 @@ guide typeclass resolution. -/ def isColimitOfPreserves (F : C ⥤ D) {c : Cocone K} (t : IsColimit c) [PreservesColimit K F] : IsColimit (F.mapCocone c) := - PreservesColimit.preserves t + (PreservesColimit.preserves t).some instance preservesLimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (PreservesLimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance preservesColimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (PreservesColimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance preservesLimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (PreservesLimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance preservesColimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (PreservesColimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -instance preserves_limits_subsingleton (F : C ⥤ D) : +instance preservesLimitsOfSize_subsingleton (F : C ⥤ D) : Subsingleton (PreservesLimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -instance preserves_colimits_subsingleton (F : C ⥤ D) : +instance preservesColimitsOfSize_subsingleton (F : C ⥤ D) : Subsingleton (PreservesColimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -instance idPreservesLimits : PreservesLimitsOfSize.{w', w} (𝟭 C) where +instance id_preservesLimitsOfSize : PreservesLimitsOfSize.{w', w} (𝟭 C) where preservesLimitsOfShape {J} 𝒥 := { preservesLimit := fun {K} => @@ -150,7 +150,11 @@ instance idPreservesLimits : PreservesLimitsOfSize.{w', w} (𝟭 C) where cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } -instance idPreservesColimits : PreservesColimitsOfSize.{w', w} (𝟭 C) where +@[deprecated "use id_preservesLimitsOfSize" (since := "2024-11-19")] +lemma idPreservesLimits : PreservesLimitsOfSize.{w', w} (𝟭 C) := + id_preservesLimitsOfSize + +instance id_preservesColimitsOfSize : PreservesColimitsOfSize.{w', w} (𝟭 C) where preservesColimitsOfShape {J} 𝒥 := { preservesColimit := fun {K} => @@ -160,12 +164,15 @@ instance idPreservesColimits : PreservesColimitsOfSize.{w', w} (𝟭 C) where cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } +@[deprecated "use id_preservesColimitsOfSize" (since := "2024-11-19")] +lemma idPreservesColimits : PreservesColimitsOfSize.{w', w} (𝟭 C) := + id_preservesColimitsOfSize + instance [HasLimit K] {F : C ⥤ D} [PreservesLimit K F] : HasLimit (K ⋙ F) where - exists_limit := ⟨⟨F.mapCone (limit.cone K), PreservesLimit.preserves (limit.isLimit K)⟩⟩ + exists_limit := ⟨_, isLimitOfPreserves F (limit.isLimit K)⟩ instance [HasColimit K] {F : C ⥤ D} [PreservesColimit K F] : HasColimit (K ⋙ F) where - exists_colimit := - ⟨⟨F.mapCocone (colimit.cocone K), PreservesColimit.preserves (colimit.isColimit K)⟩⟩ + exists_colimit := ⟨_, isColimitOfPreserves F (colimit.isColimit K)⟩ section @@ -175,172 +182,296 @@ variable (F : C ⥤ D) (G : D ⥤ E) -- Porting note: made this global by removing local attribute [elab_without_expected_type] PreservesLimit.preserves PreservesColimit.preserves -instance compPreservesLimit [PreservesLimit K F] [PreservesLimit (K ⋙ F) G] : - PreservesLimit K (F ⋙ G) := - ⟨fun h => PreservesLimit.preserves (PreservesLimit.preserves h)⟩ +instance comp_preservesLimit [PreservesLimit K F] [PreservesLimit (K ⋙ F) G] : + PreservesLimit K (F ⋙ G) where + preserves hc := ⟨isLimitOfPreserves G (isLimitOfPreserves F hc)⟩ -instance compPreservesLimitsOfShape [PreservesLimitsOfShape J F] [PreservesLimitsOfShape J G] : +instance comp_preservesLimitsOfShape [PreservesLimitsOfShape J F] [PreservesLimitsOfShape J G] : PreservesLimitsOfShape J (F ⋙ G) where -instance compPreservesLimits [PreservesLimitsOfSize.{w', w} F] [PreservesLimitsOfSize.{w', w} G] : +instance comp_preservesLimits [PreservesLimitsOfSize.{w', w} F] [PreservesLimitsOfSize.{w', w} G] : PreservesLimitsOfSize.{w', w} (F ⋙ G) where -instance compPreservesColimit [PreservesColimit K F] [PreservesColimit (K ⋙ F) G] : - PreservesColimit K (F ⋙ G) := - ⟨fun h => PreservesColimit.preserves (PreservesColimit.preserves h)⟩ +instance comp_preservesColimit [PreservesColimit K F] [PreservesColimit (K ⋙ F) G] : + PreservesColimit K (F ⋙ G) where + preserves hc := ⟨isColimitOfPreserves G (isColimitOfPreserves F hc)⟩ -instance compPreservesColimitsOfShape [PreservesColimitsOfShape J F] +instance comp_preservesColimitsOfShape [PreservesColimitsOfShape J F] [PreservesColimitsOfShape J G] : PreservesColimitsOfShape J (F ⋙ G) where -instance compPreservesColimits [PreservesColimitsOfSize.{w', w} F] +instance comp_preservesColimits [PreservesColimitsOfSize.{w', w} F] [PreservesColimitsOfSize.{w', w} G] : PreservesColimitsOfSize.{w', w} (F ⋙ G) where +@[deprecated "use comp_preservesLimit" (since := "2024-11-19")] +lemma compPreservesLimit [PreservesLimit K F] [PreservesLimit (K ⋙ F) G] : + PreservesLimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_preservesLimitsOfShape" (since := "2024-11-19")] +lemma compPreservesLimitsOfShape [PreservesLimitsOfShape J F] [PreservesLimitsOfShape J G] : + PreservesLimitsOfShape J (F ⋙ G) := + comp_preservesLimitsOfShape _ _ + +@[deprecated "use comp_preservesLimits" (since := "2024-11-19")] +lemma compPreservesLimits [PreservesLimitsOfSize.{w', w} F] [PreservesLimitsOfSize.{w', w} G] : + PreservesLimitsOfSize.{w', w} (F ⋙ G) := + comp_preservesLimits _ _ + +@[deprecated "use comp_preservesColimit" (since := "2024-11-19")] +lemma compPreservesColimit [PreservesColimit K F] [PreservesColimit (K ⋙ F) G] : + PreservesColimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_preservesColimitsOfShape" (since := "2024-11-19")] +lemma compPreservesColimitsOfShape [PreservesColimitsOfShape J F] [PreservesColimitsOfShape J G] : + PreservesColimitsOfShape J (F ⋙ G) := + comp_preservesColimitsOfShape _ _ + +@[deprecated "use comp_preservesColimits" (since := "2024-11-19")] +lemma compPreservesColimits [PreservesColimitsOfSize.{w', w} F] + [PreservesColimitsOfSize.{w', w} G] : + PreservesColimitsOfSize.{w', w} (F ⋙ G) := + comp_preservesColimits _ _ + end /-- If F preserves one limit cone for the diagram K, then it preserves any limit cone for K. -/ -def preservesLimitOfPreservesLimitCone {F : C ⥤ D} {t : Cone K} (h : IsLimit t) +lemma preservesLimit_of_preserves_limit_cone {F : C ⥤ D} {t : Cone K} (h : IsLimit t) + (hF : IsLimit (F.mapCone t)) : PreservesLimit K F where + preserves h' := ⟨IsLimit.ofIsoLimit hF (Functor.mapIso _ (IsLimit.uniqueUpToIso h h'))⟩ + +@[deprecated "use preservesLimit_of_preserves_limit_cone" (since := "2024-11-19")] +lemma preservesLimitOfPreservesLimitCone {F : C ⥤ D} {t : Cone K} (h : IsLimit t) (hF : IsLimit (F.mapCone t)) : PreservesLimit K F := - ⟨fun h' => IsLimit.ofIsoLimit hF (Functor.mapIso _ (IsLimit.uniqueUpToIso h h'))⟩ +preservesLimit_of_preserves_limit_cone h hF /-- Transfer preservation of limits along a natural isomorphism in the diagram. -/ -def preservesLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [PreservesLimit K₁ F] : - PreservesLimit K₂ F where - preserves {c} t := by +lemma preservesLimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesLimit K₁ F] : PreservesLimit K₂ F where + preserves {c} t := ⟨by apply IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) _ _ have := (IsLimit.postcomposeInvEquiv h c).symm t apply IsLimit.ofIsoLimit (isLimitOfPreserves F this) - exact Cones.ext (Iso.refl _) + exact Cones.ext (Iso.refl _)⟩ + +@[deprecated "use preservesLimit_of_iso_diagram" (since := "2024-11-19")] +lemma preservesLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesLimit K₁ F] : PreservesLimit K₂ F := + preservesLimit_of_iso_diagram F h /-- Transfer preservation of a limit along a natural isomorphism in the functor. -/ -def preservesLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesLimit K F] : +lemma preservesLimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesLimit K F] : PreservesLimit K G where - preserves t := IsLimit.mapConeEquiv h (PreservesLimit.preserves t) + preserves t := ⟨IsLimit.mapConeEquiv h (isLimitOfPreserves F t)⟩ + +@[deprecated "use preservesLimit_of_natIso" (since := "2024-11-19")] +lemma preservesLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesLimit K F] : + PreservesLimit K G := + preservesLimit_of_natIso K h /-- Transfer preservation of limits of shape along a natural isomorphism in the functor. -/ -def preservesLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfShape J F] : +lemma preservesLimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J G where - preservesLimit {K} := preservesLimitOfNatIso K h + preservesLimit {K} := preservesLimit_of_natIso K h + +@[deprecated "use preservesLimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfShape J F] : + PreservesLimitsOfShape J G := + preservesLimitsOfShape_of_natIso h /-- Transfer preservation of limits along a natural isomorphism in the functor. -/ -def preservesLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} G where - preservesLimitsOfShape {_J} _𝒥₁ := preservesLimitsOfShapeOfNatIso h + preservesLimitsOfShape := preservesLimitsOfShape_of_natIso h + +@[deprecated "use preservesLimits_of_natIso" (since := "2024-11-19")] +lemma preservesLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesLimitsOfSize.{w, w'} F] : + PreservesLimitsOfSize.{w, w'} G := + preservesLimits_of_natIso h /-- Transfer preservation of limits along an equivalence in the shape. -/ -def preservesLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma preservesLimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J' F where preservesLimit {K} := - { preserves := fun {c} t => by + { preserves := fun {c} t => ⟨by let equ := e.invFunIdAssoc (K ⋙ F) have := (isLimitOfPreserves F (t.whiskerEquivalence e)).whiskerEquivalence e.symm apply ((IsLimit.postcomposeHomEquiv equ _).symm this).ofIsoLimit refine Cones.ext (Iso.refl _) fun j => ?_ dsimp - simp [equ, ← Functor.map_comp] } + simp [equ, ← Functor.map_comp]⟩ } + +@[deprecated "use preservesLimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J' F := + preservesLimitsOfShape_of_equiv e F /-- A functor preserving larger limits also preserves smaller limits. -/ -def preservesLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma preservesLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesLimitsOfSize.{w', w₂'} F] : PreservesLimitsOfSize.{w, w₂} F where - preservesLimitsOfShape {J} := preservesLimitsOfShapeOfEquiv + preservesLimitsOfShape {J} := preservesLimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F +@[deprecated "use preservesLimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma preservesLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [PreservesLimitsOfSize.{w', w₂'} F] : PreservesLimitsOfSize.{w, w₂} F := + preservesLimitsOfSize_of_univLE.{w', w₂'} F + -- See library note [dsimp, simp]. -/-- `PreservesLimitsOfSizeShrink.{w w'} F` tries to obtain `PreservesLimitsOfSize.{w w'} F` +/-- `PreservesLimitsOfSize_shrink.{w w'} F` tries to obtain `PreservesLimitsOfSize.{w w'} F` from some other `PreservesLimitsOfSize F`. -/ -def preservesLimitsOfSizeShrink (F : C ⥤ D) [PreservesLimitsOfSize.{max w w₂, max w' w₂'} F] : - PreservesLimitsOfSize.{w, w'} F := preservesLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma preservesLimitsOfSize_shrink (F : C ⥤ D) [PreservesLimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesLimitsOfSize.{w, w'} F := preservesLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use preservesLimitsOfSize_shrink" (since := "2024-11-19")] +lemma PreservesLimitsOfSizeShrink (F : C ⥤ D) [PreservesLimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesLimitsOfSize.{w, w'} F := + preservesLimitsOfSize_shrink F /-- Preserving limits at any universe level implies preserving limits in universe `0`. -/ -def preservesSmallestLimitsOfPreservesLimits (F : C ⥤ D) [PreservesLimitsOfSize.{v₃, u₃} F] : +lemma preservesSmallestLimits_of_preservesLimits (F : C ⥤ D) [PreservesLimitsOfSize.{v₃, u₃} F] : PreservesLimitsOfSize.{0, 0} F := - preservesLimitsOfSizeShrink F + preservesLimitsOfSize_shrink F + +@[deprecated "use preservesSmallestLimits_of_preservesLimits" (since := "2024-11-19")] +lemma preservesSmallestLimitsOfPreservesLimits (F : C ⥤ D) [PreservesLimitsOfSize.{v₃, u₃} F] : + PreservesLimitsOfSize.{0, 0} F := + preservesSmallestLimits_of_preservesLimits F /-- If F preserves one colimit cocone for the diagram K, then it preserves any colimit cocone for K. -/ -def preservesColimitOfPreservesColimitCocone {F : C ⥤ D} {t : Cocone K} (h : IsColimit t) +lemma preservesColimit_of_preserves_colimit_cocone {F : C ⥤ D} {t : Cocone K} (h : IsColimit t) (hF : IsColimit (F.mapCocone t)) : PreservesColimit K F := - ⟨fun h' => IsColimit.ofIsoColimit hF (Functor.mapIso _ (IsColimit.uniqueUpToIso h h'))⟩ + ⟨fun h' => ⟨IsColimit.ofIsoColimit hF (Functor.mapIso _ (IsColimit.uniqueUpToIso h h'))⟩⟩ + +@[deprecated "use preservesColimit_of_preserves_colimit_cocone" (since := "2024-11-19")] +lemma preservesColimitOfPreservesColimitCocone {F : C ⥤ D} {t : Cocone K} (h : IsColimit t) + (hF : IsColimit (F.mapCocone t)) : PreservesColimit K F := +preservesColimit_of_preserves_colimit_cocone h hF /-- Transfer preservation of colimits along a natural isomorphism in the shape. -/ -def preservesColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [PreservesColimit K₁ F] : +lemma preservesColimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesColimit K₁ F] : PreservesColimit K₂ F where - preserves {c} t := by + preserves {c} t := ⟨by apply IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) _ _ have := (IsColimit.precomposeHomEquiv h c).symm t apply IsColimit.ofIsoColimit (isColimitOfPreserves F this) - exact Cocones.ext (Iso.refl _) + exact Cocones.ext (Iso.refl _)⟩ + +@[deprecated "use preservesColimit_of_iso_diagram" (since := "2024-11-19")] +lemma preservesColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [PreservesColimit K₁ F] : + PreservesColimit K₂ F := + preservesColimit_of_iso_diagram F h /-- Transfer preservation of a colimit along a natural isomorphism in the functor. -/ -def preservesColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesColimit K F] : +lemma preservesColimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesColimit K F] : PreservesColimit K G where - preserves t := IsColimit.mapCoconeEquiv h (PreservesColimit.preserves t) + preserves t := ⟨IsColimit.mapCoconeEquiv h (isColimitOfPreserves F t)⟩ + +@[deprecated preservesColimit_of_natIso (since := "2024-11-19")] +lemma preservesColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [PreservesColimit K F] : + PreservesColimit K G := + preservesColimit_of_natIso K h /-- Transfer preservation of colimits of shape along a natural isomorphism in the functor. -/ -def preservesColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : +lemma preservesColimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J G where - preservesColimit {K} := preservesColimitOfNatIso K h + preservesColimit {K} := preservesColimit_of_natIso K h + +@[deprecated "use preservesColimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma preservesColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfShape J F] : + PreservesColimitsOfShape J G := + preservesColimitsOfShape_of_natIso h /-- Transfer preservation of colimits along a natural isomorphism in the functor. -/ -def preservesColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} G where - preservesColimitsOfShape {_J} _𝒥₁ := preservesColimitsOfShapeOfNatIso h + preservesColimitsOfShape {_J} _𝒥₁ := preservesColimitsOfShape_of_natIso h + +@[deprecated "use preservesColimits_of_natIso" (since := "2024-11-19")] +lemma preservesColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesColimitsOfSize.{w, w'} F] : + PreservesColimitsOfSize.{w, w'} G := + preservesColimits_of_natIso h /-- Transfer preservation of colimits along an equivalence in the shape. -/ -def preservesColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma preservesColimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J' F where preservesColimit {K} := - { - preserves := fun {c} t => by + { preserves := fun {c} t => ⟨by let equ := e.invFunIdAssoc (K ⋙ F) have := (isColimitOfPreserves F (t.whiskerEquivalence e)).whiskerEquivalence e.symm apply ((IsColimit.precomposeInvEquiv equ _).symm this).ofIsoColimit refine Cocones.ext (Iso.refl _) fun j => ?_ dsimp - simp [equ, ← Functor.map_comp] } + simp [equ, ← Functor.map_comp]⟩ } + +@[deprecated "use preservesColimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma preservesColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J' F := + preservesColimitsOfShape_of_equiv e F /-- A functor preserving larger colimits also preserves smaller colimits. -/ -def preservesColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma preservesColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesColimitsOfSize.{w', w₂'} F] : PreservesColimitsOfSize.{w, w₂} F where - preservesColimitsOfShape {J} := preservesColimitsOfShapeOfEquiv + preservesColimitsOfShape {J} := preservesColimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F +@[deprecated "use preservesColimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma preservesColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [PreservesColimitsOfSize.{w', w₂'} F] : PreservesColimitsOfSize.{w, w₂} F := + preservesColimitsOfSize_of_univLE.{w', w₂'} F + -- See library note [dsimp, simp]. /-- -`PreservesColimitsOfSizeShrink.{w w'} F` tries to obtain `PreservesColimitsOfSize.{w w'} F` +`PreservesColimitsOfSize_shrink.{w w'} F` tries to obtain `PreservesColimitsOfSize.{w w'} F` from some other `PreservesColimitsOfSize F`. -/ -def preservesColimitsOfSizeShrink (F : C ⥤ D) [PreservesColimitsOfSize.{max w w₂, max w' w₂'} F] : - PreservesColimitsOfSize.{w, w'} F := preservesColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma preservesColimitsOfSize_shrink (F : C ⥤ D) + [PreservesColimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesColimitsOfSize.{w, w'} F := preservesColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use preservesColimitsOfSize_shrink" (since := "2024-11-19")] +lemma PreservesColimitsOfSizeShrink (F : C ⥤ D) + [PreservesColimitsOfSize.{max w w₂, max w' w₂'} F] : + PreservesColimitsOfSize.{w, w'} F := + preservesColimitsOfSize_shrink F /-- Preserving colimits at any universe implies preserving colimits at universe `0`. -/ -def preservesSmallestColimitsOfPreservesColimits (F : C ⥤ D) [PreservesColimitsOfSize.{v₃, u₃} F] : +lemma preservesSmallestColimits_of_preservesColimits (F : C ⥤ D) + [PreservesColimitsOfSize.{v₃, u₃} F] : PreservesColimitsOfSize.{0, 0} F := - preservesColimitsOfSizeShrink F + preservesColimitsOfSize_shrink F + +@[deprecated "use preservesSmallestColimits_of_preservesColimits" (since := "2024-11-19")] +lemma preservesSmallestColimitsOfPreservesColimits (F : C ⥤ D) + [PreservesColimitsOfSize.{v₃, u₃} F] : + PreservesColimitsOfSize.{0, 0} F := + preservesSmallestColimits_of_preservesColimits F /-- A functor `F : C ⥤ D` reflects limits for `K : J ⥤ C` if whenever the image of a cone over `K` under `F` is a limit cone in `D`, the cone was already a limit cone in `C`. Note that we do not assume a priori that `D` actually has any limits. -/ -class ReflectsLimit (K : J ⥤ C) (F : C ⥤ D) where - reflects : ∀ {c : Cone K}, IsLimit (F.mapCone c) → IsLimit c +class ReflectsLimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + reflects {c : Cone K} (hc : IsLimit (F.mapCone c)) : Nonempty (IsLimit c) /-- A functor `F : C ⥤ D` reflects colimits for `K : J ⥤ C` if whenever the image of a cocone over `K` under `F` is a colimit cocone in `D`, the cocone was already a colimit cocone in `C`. Note that we do not assume a priori that `D` actually has any colimits. -/ -class ReflectsColimit (K : J ⥤ C) (F : C ⥤ D) where - reflects : ∀ {c : Cocone K}, IsColimit (F.mapCocone c) → IsColimit c +class ReflectsColimit (K : J ⥤ C) (F : C ⥤ D) : Prop where + reflects {c : Cocone K} (hc : IsColimit (F.mapCocone c)) : Nonempty (IsColimit c) /-- A functor `F : C ⥤ D` reflects limits of shape `J` if whenever the image of a cone over some `K : J ⥤ C` under `F` is a limit cone in `D`, the cone was already a limit cone in `C`. Note that we do not assume a priori that `D` actually has any limits. -/ -class ReflectsLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class ReflectsLimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where reflectsLimit : ∀ {K : J ⥤ C}, ReflectsLimit K F := by infer_instance /-- A functor `F : C ⥤ D` reflects colimits of shape `J` if @@ -348,7 +479,7 @@ whenever the image of a cocone over some `K : J ⥤ C` under `F` is a colimit co the cocone was already a colimit cocone in `C`. Note that we do not assume a priori that `D` actually has any colimits. -/ -class ReflectsColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) where +class ReflectsColimitsOfShape (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Prop where reflectsColimit : ∀ {K : J ⥤ C}, ReflectsColimit K F := by infer_instance -- This should be used with explicit universe variables. @@ -358,7 +489,7 @@ the cone was already a limit cone in `C`. Note that we do not assume a priori that `D` actually has any limits. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsLimitsOfSize (F : C ⥤ D) where +class ReflectsLimitsOfSize (F : C ⥤ D) : Prop where reflectsLimitsOfShape : ∀ {J : Type w} [Category.{w'} J], ReflectsLimitsOfShape J F := by infer_instance @@ -377,7 +508,7 @@ the cocone was already a colimit cocone in `C`. Note that we do not assume a priori that `D` actually has any colimits. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsColimitsOfSize (F : C ⥤ D) where +class ReflectsColimitsOfSize (F : C ⥤ D) : Prop where reflectsColimitsOfShape : ∀ {J : Type w} [Category.{w'} J], ReflectsColimitsOfShape J F := by infer_instance @@ -393,7 +524,8 @@ abbrev ReflectsColimits (F : C ⥤ D) := guide typeclass resolution. -/ def isLimitOfReflects (F : C ⥤ D) {c : Cone K} (t : IsLimit (F.mapCone c)) - [ReflectsLimit K F] : IsLimit c := ReflectsLimit.reflects t + [ReflectsLimit K F] : IsLimit c := + (ReflectsLimit.reflects t).some /-- A convenience function for `ReflectsColimit`, which takes the functor as an explicit argument to @@ -401,176 +533,259 @@ guide typeclass resolution. -/ def isColimitOfReflects (F : C ⥤ D) {c : Cocone K} (t : IsColimit (F.mapCocone c)) [ReflectsColimit K F] : IsColimit c := - ReflectsColimit.reflects t + (ReflectsColimit.reflects t).some instance reflectsLimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (ReflectsLimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflectsColimit_subsingleton (K : J ⥤ C) (F : C ⥤ D) : Subsingleton (ReflectsColimit K F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflectsLimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (ReflectsLimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflectsColimitsOfShape_subsingleton (J : Type w) [Category.{w'} J] (F : C ⥤ D) : Subsingleton (ReflectsColimitsOfShape J F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflects_limits_subsingleton (F : C ⥤ D) : Subsingleton (ReflectsLimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! instance reflects_colimits_subsingleton (F : C ⥤ D) : Subsingleton (ReflectsColimitsOfSize.{w', w} F) := by - constructor; rintro ⟨a⟩ ⟨b⟩; congr!; subsingleton + constructor; rintro ⟨a⟩ ⟨b⟩; congr! -- see Note [lower instance priority] -instance (priority := 100) reflectsLimitOfReflectsLimitsOfShape (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) reflectsLimit_of_reflectsLimitsOfShape (K : J ⥤ C) (F : C ⥤ D) [ReflectsLimitsOfShape J F] : ReflectsLimit K F := ReflectsLimitsOfShape.reflectsLimit -- see Note [lower instance priority] -instance (priority := 100) reflectsColimitOfReflectsColimitsOfShape (K : J ⥤ C) (F : C ⥤ D) +instance (priority := 100) reflectsColimit_of_reflectsColimitsOfShape (K : J ⥤ C) (F : C ⥤ D) [ReflectsColimitsOfShape J F] : ReflectsColimit K F := ReflectsColimitsOfShape.reflectsColimit -- see Note [lower instance priority] -instance (priority := 100) reflectsLimitsOfShapeOfReflectsLimits (J : Type w) [Category.{w'} J] +instance (priority := 100) reflectsLimitsOfShape_of_reflectsLimits (J : Type w) [Category.{w'} J] (F : C ⥤ D) [ReflectsLimitsOfSize.{w', w} F] : ReflectsLimitsOfShape J F := ReflectsLimitsOfSize.reflectsLimitsOfShape -- see Note [lower instance priority] -instance (priority := 100) reflectsColimitsOfShapeOfReflectsColimits (J : Type w) [Category.{w'} J] +instance (priority := 100) reflectsColimitsOfShape_of_reflectsColimits + (J : Type w) [Category.{w'} J] (F : C ⥤ D) [ReflectsColimitsOfSize.{w', w} F] : ReflectsColimitsOfShape J F := ReflectsColimitsOfSize.reflectsColimitsOfShape -instance idReflectsLimits : ReflectsLimitsOfSize.{w, w'} (𝟭 C) where +instance id_reflectsLimits : ReflectsLimitsOfSize.{w, w'} (𝟭 C) where reflectsLimitsOfShape {J} 𝒥 := - { - reflectsLimit := fun {K} => + { reflectsLimit := fun {K} => ⟨fun {c} h => ⟨fun s => h.lift ⟨s.pt, fun j => s.π.app j, fun _ _ f => s.π.naturality f⟩, by cases K; rcases c with ⟨_, _, _⟩; intro s j; cases s; exact h.fac _ j, by cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } -instance idReflectsColimits : ReflectsColimitsOfSize.{w, w'} (𝟭 C) where +@[deprecated "use id_reflectsLimits" (since := "2024-11-19")] +lemma idReflectsLimits : ReflectsLimitsOfSize.{w, w'} (𝟭 C) := id_reflectsLimits + +instance id_reflectsColimits : ReflectsColimitsOfSize.{w, w'} (𝟭 C) where reflectsColimitsOfShape {J} 𝒥 := - { - reflectsColimit := fun {K} => + { reflectsColimit := fun {K} => ⟨fun {c} h => ⟨fun s => h.desc ⟨s.pt, fun j => s.ι.app j, fun _ _ f => s.ι.naturality f⟩, by cases K; rcases c with ⟨_, _, _⟩; intro s j; cases s; exact h.fac _ j, by cases K; rcases c with ⟨_, _, _⟩; intro s m w; rcases s with ⟨_, _, _⟩; exact h.uniq _ m w⟩⟩ } +@[deprecated "use id_reflectsColimits" (since := "2024-11-19")] +lemma idReflectsColimits : ReflectsColimitsOfSize.{w, w'} (𝟭 C) := id_reflectsColimits + section variable {E : Type u₃} [ℰ : Category.{v₃} E] variable (F : C ⥤ D) (G : D ⥤ E) -instance compReflectsLimit [ReflectsLimit K F] [ReflectsLimit (K ⋙ F) G] : +instance comp_reflectsLimit [ReflectsLimit K F] [ReflectsLimit (K ⋙ F) G] : ReflectsLimit K (F ⋙ G) := - ⟨fun h => ReflectsLimit.reflects (ReflectsLimit.reflects h)⟩ + ⟨fun h => ReflectsLimit.reflects (isLimitOfReflects G h)⟩ -instance compReflectsLimitsOfShape [ReflectsLimitsOfShape J F] [ReflectsLimitsOfShape J G] : +instance comp_reflectsLimitsOfShape [ReflectsLimitsOfShape J F] [ReflectsLimitsOfShape J G] : ReflectsLimitsOfShape J (F ⋙ G) where -instance compReflectsLimits [ReflectsLimitsOfSize.{w', w} F] [ReflectsLimitsOfSize.{w', w} G] : +instance comp_reflectsLimits [ReflectsLimitsOfSize.{w', w} F] [ReflectsLimitsOfSize.{w', w} G] : ReflectsLimitsOfSize.{w', w} (F ⋙ G) where -instance compReflectsColimit [ReflectsColimit K F] [ReflectsColimit (K ⋙ F) G] : +instance comp_reflectsColimit [ReflectsColimit K F] [ReflectsColimit (K ⋙ F) G] : ReflectsColimit K (F ⋙ G) := - ⟨fun h => ReflectsColimit.reflects (ReflectsColimit.reflects h)⟩ + ⟨fun h => ReflectsColimit.reflects (isColimitOfReflects G h)⟩ -instance compReflectsColimitsOfShape [ReflectsColimitsOfShape J F] [ReflectsColimitsOfShape J G] : +instance comp_reflectsColimitsOfShape [ReflectsColimitsOfShape J F] [ReflectsColimitsOfShape J G] : ReflectsColimitsOfShape J (F ⋙ G) where -instance compReflectsColimits [ReflectsColimitsOfSize.{w', w} F] +instance comp_reflectsColimits [ReflectsColimitsOfSize.{w', w} F] [ReflectsColimitsOfSize.{w', w} G] : ReflectsColimitsOfSize.{w', w} (F ⋙ G) where +@[deprecated "use comp_reflectsLimit" (since := "2024-11-19")] +lemma compReflectsLimit [ReflectsLimit K F] [ReflectsLimit (K ⋙ F) G] : + ReflectsLimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsLimitsOfShape " (since := "2024-11-19")] +lemma compReflectsLimitsOfShape [ReflectsLimitsOfShape J F] [ReflectsLimitsOfShape J G] : + ReflectsLimitsOfShape J (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsLimits" (since := "2024-11-19")] +lemma compReflectsLimits [ReflectsLimitsOfSize.{w', w} F] [ReflectsLimitsOfSize.{w', w} G] : + ReflectsLimitsOfSize.{w', w} (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsColimit" (since := "2024-11-19")] +lemma compReflectsColimit [ReflectsColimit K F] [ReflectsColimit (K ⋙ F) G] : + ReflectsColimit K (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsColimitsOfShape " (since := "2024-11-19")] +lemma compReflectsColimitsOfShape [ReflectsColimitsOfShape J F] [ReflectsColimitsOfShape J G] : + ReflectsColimitsOfShape J (F ⋙ G) := inferInstance + +@[deprecated "use comp_reflectsColimits" (since := "2024-11-19")] +lemma compReflectsColimits [ReflectsColimitsOfSize.{w', w} F] [ReflectsColimitsOfSize.{w', w} G] : + ReflectsColimitsOfSize.{w', w} (F ⋙ G) := inferInstance + /-- If `F ⋙ G` preserves limits for `K`, and `G` reflects limits for `K ⋙ F`, then `F` preserves limits for `K`. -/ -def preservesLimitOfReflectsOfPreserves [PreservesLimit K (F ⋙ G)] [ReflectsLimit (K ⋙ F) G] : +lemma preservesLimit_of_reflects_of_preserves [PreservesLimit K (F ⋙ G)] [ReflectsLimit (K ⋙ F) G] : PreservesLimit K F := - ⟨fun h => by + ⟨fun h => ⟨by apply isLimitOfReflects G - apply isLimitOfPreserves (F ⋙ G) h⟩ + apply isLimitOfPreserves (F ⋙ G) h⟩⟩ + +@[deprecated "use preservesLimit_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesLimitOfReflectsOfPreserves [PreservesLimit K (F ⋙ G)] [ReflectsLimit (K ⋙ F) G] : + PreservesLimit K F := + preservesLimit_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves limits of shape `J` and `G` reflects limits of shape `J`, then `F` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfReflectsOfPreserves [PreservesLimitsOfShape J (F ⋙ G)] +lemma preservesLimitsOfShape_of_reflects_of_preserves [PreservesLimitsOfShape J (F ⋙ G)] [ReflectsLimitsOfShape J G] : PreservesLimitsOfShape J F where - preservesLimit := preservesLimitOfReflectsOfPreserves F G + preservesLimit := preservesLimit_of_reflects_of_preserves F G + +@[deprecated "use preservesLimitsOfShape_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesLimitsOfShapeOfReflectsOfPreserves [PreservesLimitsOfShape J (F ⋙ G)] + [ReflectsLimitsOfShape J G] : PreservesLimitsOfShape J F := + preservesLimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves limits and `G` reflects limits, then `F` preserves limits. -/ -def preservesLimitsOfReflectsOfPreserves [PreservesLimitsOfSize.{w', w} (F ⋙ G)] +lemma preservesLimits_of_reflects_of_preserves [PreservesLimitsOfSize.{w', w} (F ⋙ G)] [ReflectsLimitsOfSize.{w', w} G] : PreservesLimitsOfSize.{w', w} F where - preservesLimitsOfShape := preservesLimitsOfShapeOfReflectsOfPreserves F G + preservesLimitsOfShape := preservesLimitsOfShape_of_reflects_of_preserves F G + +@[deprecated "use preservesLimits_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesLimitsOfReflectsOfPreserves [PreservesLimitsOfSize.{w', w} (F ⋙ G)] + [ReflectsLimitsOfSize.{w', w} G] : PreservesLimitsOfSize.{w', w} F := + preservesLimits_of_reflects_of_preserves.{w', w} F G /-- Transfer reflection of limits along a natural isomorphism in the diagram. -/ -def reflectsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsLimit K₁ F] : +lemma reflectsLimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsLimit K₁ F] : ReflectsLimit K₂ F where - reflects {c} t := by + reflects {c} t := ⟨by apply IsLimit.postcomposeInvEquiv h c (isLimitOfReflects F _) apply ((IsLimit.postcomposeInvEquiv (isoWhiskerRight h F : _) _).symm t).ofIsoLimit _ - exact Cones.ext (Iso.refl _) + exact Cones.ext (Iso.refl _)⟩ + +@[deprecated "use reflectsLimit_of_iso_diagram" (since := "2024-11-19")] +lemma reflectsLimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsLimit K₁ F] : + ReflectsLimit K₂ F := + reflectsLimit_of_iso_diagram F h /-- Transfer reflection of a limit along a natural isomorphism in the functor. -/ -def reflectsLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimit K F] : +lemma reflectsLimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimit K F] : ReflectsLimit K G where reflects t := ReflectsLimit.reflects (IsLimit.mapConeEquiv h.symm t) +@[deprecated "use reflectsLimit_of_natIso" (since := "2024-11-19")] +lemma reflectsLimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimit K F] : + ReflectsLimit K G := + reflectsLimit_of_natIso K h + /-- Transfer reflection of limits of shape along a natural isomorphism in the functor. -/ -def reflectsLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfShape J F] : +lemma reflectsLimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfShape J F] : ReflectsLimitsOfShape J G where - reflectsLimit {K} := reflectsLimitOfNatIso K h + reflectsLimit {K} := reflectsLimit_of_natIso K h + +@[deprecated "use reflectsLimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma reflectsLimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfShape J F] : + ReflectsLimitsOfShape J G := + reflectsLimitsOfShape_of_natIso h /-- Transfer reflection of limits along a natural isomorphism in the functor. -/ -def reflectsLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfSize.{w', w} F] : +lemma reflectsLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfSize.{w', w} F] : ReflectsLimitsOfSize.{w', w} G where - reflectsLimitsOfShape := reflectsLimitsOfShapeOfNatIso h + reflectsLimitsOfShape := reflectsLimitsOfShape_of_natIso h + +@[deprecated "use reflectsLimits_of_natIso" (since := "2024-11-19")] +lemma reflectsLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsLimitsOfSize.{w', w} F] : + ReflectsLimitsOfSize.{w', w} G := + reflectsLimits_of_natIso h /-- Transfer reflection of limits along an equivalence in the shape. -/ -def reflectsLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma reflectsLimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [ReflectsLimitsOfShape J F] : ReflectsLimitsOfShape J' F where reflectsLimit {K} := - { - reflects := fun {c} t => by + { reflects := fun {c} t => ⟨by apply IsLimit.ofWhiskerEquivalence e apply isLimitOfReflects F apply IsLimit.ofIsoLimit _ (Functor.mapConeWhisker _).symm - exact IsLimit.whiskerEquivalence t _ } + exact IsLimit.whiskerEquivalence t _⟩ } + +@[deprecated "use reflectsLimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma reflectsLimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [ReflectsLimitsOfShape J F] : ReflectsLimitsOfShape J' F := + reflectsLimitsOfShape_of_equiv e F /-- A functor reflecting larger limits also reflects smaller limits. -/ -def reflectsLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma reflectsLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsLimitsOfSize.{w', w₂'} F] : ReflectsLimitsOfSize.{w, w₂} F where - reflectsLimitsOfShape {J} := reflectsLimitsOfShapeOfEquiv + reflectsLimitsOfShape {J} := reflectsLimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F -/-- `reflectsLimitsOfSizeShrink.{w w'} F` tries to obtain `reflectsLimitsOfSize.{w w'} F` +@[deprecated "use reflectsLimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma reflectsLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [ReflectsLimitsOfSize.{w', w₂'} F] : ReflectsLimitsOfSize.{w, w₂} F := + reflectsLimitsOfSize_of_univLE.{w'} F + +/-- `reflectsLimitsOfSize_shrink.{w w'} F` tries to obtain `reflectsLimitsOfSize.{w w'} F` from some other `reflectsLimitsOfSize F`. -/ -def reflectsLimitsOfSizeShrink (F : C ⥤ D) [ReflectsLimitsOfSize.{max w w₂, max w' w₂'} F] : - ReflectsLimitsOfSize.{w, w'} F := reflectsLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma reflectsLimitsOfSize_shrink (F : C ⥤ D) [ReflectsLimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsLimitsOfSize.{w, w'} F := reflectsLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use reflectsLimitsOfSize_shrink" (since := "2024-11-19")] +lemma reflectsLimitsOfSizeShrink (F : C ⥤ D) [ReflectsLimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsLimitsOfSize.{w, w'} F := + reflectsLimitsOfSize_shrink F /-- Reflecting limits at any universe implies reflecting limits at universe `0`. -/ -def reflectsSmallestLimitsOfReflectsLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{v₃, u₃} F] : +lemma reflectsSmallestLimits_of_reflectsLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{v₃, u₃} F] : + ReflectsLimitsOfSize.{0, 0} F := + reflectsLimitsOfSize_shrink F + +@[deprecated "use reflectsSmallestLimits_of_reflectsLimits" (since := "2024-11-19")] +lemma reflectsSmallestLimitsOfReflectsLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{v₃, u₃} F] : ReflectsLimitsOfSize.{0, 0} F := - reflectsLimitsOfSizeShrink F + reflectsSmallestLimits_of_reflectsLimits F /-- If the limit of `F` exists and `G` preserves it, then if `G` reflects isomorphisms then it reflects the limit of `F`. -/ -- Porting note: previous behavior of apply pushed instance holes into hypotheses, this errors -def reflectsLimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] +lemma reflectsLimit_of_reflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] [HasLimit F] [PreservesLimit F G] : ReflectsLimit F G where reflects {c} t := by - suffices IsIso (IsLimit.lift (limit.isLimit F) c) from by - apply IsLimit.ofPointIso (limit.isLimit F) + suffices IsIso (IsLimit.lift (limit.isLimit F) c) from ⟨by + apply IsLimit.ofPointIso (limit.isLimit F)⟩ change IsIso ((Cones.forget _).map ((limit.isLimit F).liftConeMorphism c)) suffices IsIso (IsLimit.liftConeMorphism (limit.isLimit F) c) from by apply (Cones.forget F).map_isIso _ @@ -579,100 +794,174 @@ def reflectsLimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsI apply isIso_of_reflects_iso _ (Cones.functoriality F G) exact t.hom_isIso (isLimitOfPreserves G (limit.isLimit F)) _ +@[deprecated "use reflectsLimit_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsLimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] + [HasLimit F] [PreservesLimit F G] : ReflectsLimit F G := + reflectsLimit_of_reflectsIsomorphisms F G + /-- If `C` has limits of shape `J` and `G` preserves them, then if `G` reflects isomorphisms then it reflects limits of shape `J`. -/ -def reflectsLimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsLimitsOfShape_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasLimitsOfShape J C] [PreservesLimitsOfShape J G] : ReflectsLimitsOfShape J G where - reflectsLimit {F} := reflectsLimitOfReflectsIsomorphisms F G + reflectsLimit {F} := reflectsLimit_of_reflectsIsomorphisms F G + +@[deprecated "use reflectsLimitsOfShape_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsLimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasLimitsOfShape J C] [PreservesLimitsOfShape J G] : ReflectsLimitsOfShape J G := + reflectsLimitsOfShape_of_reflectsIsomorphisms /-- If `C` has limits and `G` preserves limits, then if `G` reflects isomorphisms then it reflects limits. -/ -def reflectsLimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsLimits_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasLimitsOfSize.{w', w} C] [PreservesLimitsOfSize.{w', w} G] : ReflectsLimitsOfSize.{w', w} G where - reflectsLimitsOfShape := reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape := reflectsLimitsOfShape_of_reflectsIsomorphisms + +@[deprecated "use reflectsLimits_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsLimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasLimitsOfSize.{w', w} C] [PreservesLimitsOfSize.{w', w} G] : + ReflectsLimitsOfSize.{w', w} G := + reflectsLimits_of_reflectsIsomorphisms /-- If `F ⋙ G` preserves colimits for `K`, and `G` reflects colimits for `K ⋙ F`, then `F` preserves colimits for `K`. -/ -def preservesColimitOfReflectsOfPreserves [PreservesColimit K (F ⋙ G)] [ReflectsColimit (K ⋙ F) G] : +lemma preservesColimit_of_reflects_of_preserves + [PreservesColimit K (F ⋙ G)] [ReflectsColimit (K ⋙ F) G] : PreservesColimit K F := - ⟨fun {c} h => by + ⟨fun {c} h => ⟨by apply isColimitOfReflects G - apply isColimitOfPreserves (F ⋙ G) h⟩ + apply isColimitOfPreserves (F ⋙ G) h⟩⟩ + +@[deprecated "use preservesColimit_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesColimitOfReflectsOfPreserves + [PreservesColimit K (F ⋙ G)] [ReflectsColimit (K ⋙ F) G] : + PreservesColimit K F := + preservesColimit_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves colimits of shape `J` and `G` reflects colimits of shape `J`, then `F` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfReflectsOfPreserves [PreservesColimitsOfShape J (F ⋙ G)] +lemma preservesColimitsOfShape_of_reflects_of_preserves [PreservesColimitsOfShape J (F ⋙ G)] [ReflectsColimitsOfShape J G] : PreservesColimitsOfShape J F where - preservesColimit := preservesColimitOfReflectsOfPreserves F G + preservesColimit := preservesColimit_of_reflects_of_preserves F G + +@[deprecated "use preservesColimitsOfShape_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesColimitsOfShapeOfReflectsOfPreserves [PreservesColimitsOfShape J (F ⋙ G)] + [ReflectsColimitsOfShape J G] : PreservesColimitsOfShape J F := + preservesColimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves colimits and `G` reflects colimits, then `F` preserves colimits. -/ -def preservesColimitsOfReflectsOfPreserves [PreservesColimitsOfSize.{w', w} (F ⋙ G)] +lemma preservesColimits_of_reflects_of_preserves [PreservesColimitsOfSize.{w', w} (F ⋙ G)] [ReflectsColimitsOfSize.{w', w} G] : PreservesColimitsOfSize.{w', w} F where - preservesColimitsOfShape := preservesColimitsOfShapeOfReflectsOfPreserves F G + preservesColimitsOfShape := preservesColimitsOfShape_of_reflects_of_preserves F G + +@[deprecated "use preservesColimits_of_reflects_of_preserves" (since := "2024-11-19")] +lemma preservesColimitsOfReflectsOfPreserves [PreservesColimitsOfSize.{w', w} (F ⋙ G)] + [ReflectsColimitsOfSize.{w', w} G] : PreservesColimitsOfSize.{w', w} F := + preservesColimits_of_reflects_of_preserves F G /-- Transfer reflection of colimits along a natural isomorphism in the diagram. -/ -def reflectsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) [ReflectsColimit K₁ F] : +lemma reflectsColimit_of_iso_diagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [ReflectsColimit K₁ F] : ReflectsColimit K₂ F where - reflects {c} t := by + reflects {c} t := ⟨by apply IsColimit.precomposeHomEquiv h c (isColimitOfReflects F _) apply ((IsColimit.precomposeHomEquiv (isoWhiskerRight h F : _) _).symm t).ofIsoColimit _ - exact Cocones.ext (Iso.refl _) + exact Cocones.ext (Iso.refl _)⟩ + +@[deprecated "use reflectsColimit_of_iso_diagram" (since := "2024-11-19")] +lemma reflectsColimitOfIsoDiagram {K₁ K₂ : J ⥤ C} (F : C ⥤ D) (h : K₁ ≅ K₂) + [ReflectsColimit K₁ F] : + ReflectsColimit K₂ F := + reflectsColimit_of_iso_diagram F h /-- Transfer reflection of a colimit along a natural isomorphism in the functor. -/ -def reflectsColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimit K F] : +lemma reflectsColimit_of_natIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimit K F] : ReflectsColimit K G where reflects t := ReflectsColimit.reflects (IsColimit.mapCoconeEquiv h.symm t) +@[deprecated "use reflectsColimit_of_natIso" (since := "2024-11-19")] +lemma reflectsColimitOfNatIso (K : J ⥤ C) {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimit K F] : + ReflectsColimit K G := + reflectsColimit_of_natIso K h + /-- Transfer reflection of colimits of shape along a natural isomorphism in the functor. -/ -def reflectsColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfShape J F] : +lemma reflectsColimitsOfShape_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfShape J F] : ReflectsColimitsOfShape J G where - reflectsColimit {K} := reflectsColimitOfNatIso K h + reflectsColimit {K} := reflectsColimit_of_natIso K h + +@[deprecated "use reflectsColimitsOfShape_of_natIso" (since := "2024-11-19")] +lemma reflectsColimitsOfShapeOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfShape J F] : + ReflectsColimitsOfShape J G := + reflectsColimitsOfShape_of_natIso h /-- Transfer reflection of colimits along a natural isomorphism in the functor. -/ -def reflectsColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfSize.{w, w'} F] : +lemma reflectsColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfSize.{w, w'} F] : ReflectsColimitsOfSize.{w, w'} G where - reflectsColimitsOfShape := reflectsColimitsOfShapeOfNatIso h + reflectsColimitsOfShape := reflectsColimitsOfShape_of_natIso h + +@[deprecated "use reflectsColimits_of_natIso" (since := "2024-11-19")] +lemma reflectsColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [ReflectsColimitsOfSize.{w, w'} F] : + ReflectsColimitsOfSize.{w, w'} G := + reflectsColimits_of_natIso h /-- Transfer reflection of colimits along an equivalence in the shape. -/ -def reflectsColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) +lemma reflectsColimitsOfShape_of_equiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) [ReflectsColimitsOfShape J F] : ReflectsColimitsOfShape J' F where reflectsColimit := - { - reflects := fun {c} t => by + { reflects := fun {c} t => ⟨by apply IsColimit.ofWhiskerEquivalence e apply isColimitOfReflects F apply IsColimit.ofIsoColimit _ (Functor.mapCoconeWhisker _).symm - exact IsColimit.whiskerEquivalence t _ } + exact IsColimit.whiskerEquivalence t _⟩ } + +@[deprecated "use reflectsColimitsOfShape_of_equiv" (since := "2024-11-19")] +lemma reflectsColimitsOfShapeOfEquiv {J' : Type w₂} [Category.{w₂'} J'] (e : J ≌ J') (F : C ⥤ D) + [ReflectsColimitsOfShape J F] : ReflectsColimitsOfShape J' F := + reflectsColimitsOfShape_of_equiv e F /-- A functor reflecting larger colimits also reflects smaller colimits. -/ -def reflectsColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] +lemma reflectsColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsColimitsOfSize.{w', w₂'} F] : ReflectsColimitsOfSize.{w, w₂} F where - reflectsColimitsOfShape {J} := reflectsColimitsOfShapeOfEquiv + reflectsColimitsOfShape {J} := reflectsColimitsOfShape_of_equiv ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm F -/-- `reflectsColimitsOfSizeShrink.{w w'} F` tries to obtain `reflectsColimitsOfSize.{w w'} F` +@[deprecated "use reflectsColimitsOfSize_of_univLE" (since := "2024-11-19")] +lemma reflectsColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] + [ReflectsColimitsOfSize.{w', w₂'} F] : ReflectsColimitsOfSize.{w, w₂} F := + reflectsColimitsOfSize_of_univLE.{w'} F + +/-- `reflectsColimitsOfSize_shrink.{w w'} F` tries to obtain `reflectsColimitsOfSize.{w w'} F` from some other `reflectsColimitsOfSize F`. -/ -def reflectsColimitsOfSizeShrink (F : C ⥤ D) [ReflectsColimitsOfSize.{max w w₂, max w' w₂'} F] : - ReflectsColimitsOfSize.{w, w'} F := reflectsColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F +lemma reflectsColimitsOfSize_shrink (F : C ⥤ D) [ReflectsColimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsColimitsOfSize.{w, w'} F := reflectsColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F + +@[deprecated "use reflectsColimitsOfSize_shrink" (since := "2024-11-19")] +lemma reflectsColimitsOfSizeShrink (F : C ⥤ D) [ReflectsColimitsOfSize.{max w w₂, max w' w₂'} F] : + ReflectsColimitsOfSize.{w, w'} F := + reflectsColimitsOfSize_shrink F /-- Reflecting colimits at any universe implies reflecting colimits at universe `0`. -/ -def reflectsSmallestColimitsOfReflectsColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{v₃, u₃} F] : +lemma reflectsSmallestColimits_of_reflectsColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{v₃, u₃} F] : ReflectsColimitsOfSize.{0, 0} F := - reflectsColimitsOfSizeShrink F + reflectsColimitsOfSize_shrink F + +@[deprecated "use reflectsSmallestColimits_of_reflectsColimits" (since := "2024-11-19")] +lemma reflectsSmallestColimitsOfReflectsColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{v₃, u₃} F] : + ReflectsColimitsOfSize.{0, 0} F := + reflectsSmallestColimits_of_reflectsColimits F /-- If the colimit of `F` exists and `G` preserves it, then if `G` reflects isomorphisms then it reflects the colimit of `F`. -/ -- Porting note: previous behavior of apply pushed instance holes into hypotheses, this errors -def reflectsColimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] +lemma reflectsColimit_of_reflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] [HasColimit F] [PreservesColimit F G] : ReflectsColimit F G where reflects {c} t := by - suffices IsIso (IsColimit.desc (colimit.isColimit F) c) from by - apply IsColimit.ofPointIso (colimit.isColimit F) + suffices IsIso (IsColimit.desc (colimit.isColimit F) c) from ⟨by + apply IsColimit.ofPointIso (colimit.isColimit F)⟩ change IsIso ((Cocones.forget _).map ((colimit.isColimit F).descCoconeMorphism c)) suffices IsIso (IsColimit.descCoconeMorphism (colimit.isColimit F) c) from by apply (Cocones.forget F).map_isIso _ @@ -681,50 +970,74 @@ def reflectsColimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.Reflect apply isIso_of_reflects_iso _ (Cocones.functoriality F G) exact (isColimitOfPreserves G (colimit.isColimit F)).hom_isIso t _ +@[deprecated "use reflectsColimit_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsColimitOfReflectsIsomorphisms (F : J ⥤ C) (G : C ⥤ D) [G.ReflectsIsomorphisms] + [HasColimit F] [PreservesColimit F G] : ReflectsColimit F G := + reflectsColimit_of_reflectsIsomorphisms F G + /-- If `C` has colimits of shape `J` and `G` preserves them, then if `G` reflects isomorphisms then it reflects colimits of shape `J`. -/ -def reflectsColimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsColimitsOfShape_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasColimitsOfShape J C] [PreservesColimitsOfShape J G] : ReflectsColimitsOfShape J G where - reflectsColimit {F} := reflectsColimitOfReflectsIsomorphisms F G + reflectsColimit {F} := reflectsColimit_of_reflectsIsomorphisms F G + +@[deprecated "use reflectsColimitsOfShape_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsColimitsOfShapeOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasColimitsOfShape J C] [PreservesColimitsOfShape J G] : ReflectsColimitsOfShape J G := + reflectsColimitsOfShape_of_reflectsIsomorphisms /-- If `C` has colimits and `G` preserves colimits, then if `G` reflects isomorphisms then it reflects colimits. -/ -def reflectsColimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] +lemma reflectsColimits_of_reflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] [HasColimitsOfSize.{w', w} C] [PreservesColimitsOfSize.{w', w} G] : ReflectsColimitsOfSize.{w', w} G where - reflectsColimitsOfShape := reflectsColimitsOfShapeOfReflectsIsomorphisms + reflectsColimitsOfShape := reflectsColimitsOfShape_of_reflectsIsomorphisms + +@[deprecated "use reflectsColimits_of_reflectsIsomorphisms" (since := "2024-11-19")] +lemma reflectsColimitsOfReflectsIsomorphisms {G : C ⥤ D} [G.ReflectsIsomorphisms] + [HasColimitsOfSize.{w', w} C] [PreservesColimitsOfSize.{w', w} G] : + ReflectsColimitsOfSize.{w', w} G := + reflectsColimits_of_reflectsIsomorphisms end variable (F : C ⥤ D) /-- A fully faithful functor reflects limits. -/ -instance fullyFaithfulReflectsLimits [F.Full] [F.Faithful] : ReflectsLimitsOfSize.{w, w'} F where +instance fullyFaithful_reflectsLimits [F.Full] [F.Faithful] : ReflectsLimitsOfSize.{w, w'} F where reflectsLimitsOfShape {J} 𝒥₁ := { reflectsLimit := fun {K} => { reflects := fun {c} t => - (IsLimit.mkConeMorphism fun _ => + ⟨(IsLimit.mkConeMorphism fun _ => (Cones.functoriality K F).preimage (t.liftConeMorphism _)) <| by apply fun s m => (Cones.functoriality K F).map_injective _ intro s m rw [Functor.map_preimage] - apply t.uniq_cone_morphism } } + apply t.uniq_cone_morphism⟩ } } + +@[deprecated "use fullyFaithful_reflectsLimits" (since := "2024-11-19")] +lemma fullyFaithfulReflectsLimits [F.Full] [F.Faithful] : ReflectsLimitsOfSize.{w, w'} F := + inferInstance /-- A fully faithful functor reflects colimits. -/ -instance fullyFaithfulReflectsColimits [F.Full] [F.Faithful] : +instance fullyFaithful_reflectsColimits [F.Full] [F.Faithful] : ReflectsColimitsOfSize.{w, w'} F where reflectsColimitsOfShape {J} 𝒥₁ := { reflectsColimit := fun {K} => { reflects := fun {c} t => - (IsColimit.mkCoconeMorphism fun _ => + ⟨(IsColimit.mkCoconeMorphism fun _ => (Cocones.functoriality K F).preimage (t.descCoconeMorphism _)) <| by apply fun s m => (Cocones.functoriality K F).map_injective _ intro s m rw [Functor.map_preimage] - apply t.uniq_cocone_morphism }} + apply t.uniq_cocone_morphism⟩ }} + +@[deprecated "use fullyFaithful_reflectsColimits" (since := "2024-11-19")] +lemma fullyFaithfulReflectsColimits [F.Full] [F.Faithful] : ReflectsColimitsOfSize.{w, w'} F := + inferInstance end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean b/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean index d1ad0cedcd3f6..f4b69c1ac9a85 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Filtered.lean @@ -36,7 +36,7 @@ section Preserves /-- `PreservesFilteredColimitsOfSize.{w', w} F` means that `F` sends all colimit cocones over any filtered diagram `J ⥤ C` to colimit cocones, where `J : Type w` with `[Category.{w'} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesFilteredColimitsOfSize (F : C ⥤ D) where +class PreservesFilteredColimitsOfSize (F : C ⥤ D) : Prop where preserves_filtered_colimits : ∀ (J : Type w) [Category.{w'} J] [IsFiltered J], PreservesColimitsOfShape J F @@ -45,7 +45,7 @@ A functor is said to preserve filtered colimits, if it preserves all colimits of `J` is a filtered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev PreservesFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev PreservesFilteredColimits (F : C ⥤ D) : Prop := PreservesFilteredColimitsOfSize.{v, v} F attribute [instance 100] PreservesFilteredColimitsOfSize.preserves_filtered_colimits @@ -54,35 +54,35 @@ instance (priority := 100) PreservesColimits.preservesFilteredColimits (F : C [PreservesColimitsOfSize.{w, w'} F] : PreservesFilteredColimitsOfSize.{w, w'} F where preserves_filtered_colimits _ := inferInstance -instance compPreservesFilteredColimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesFilteredColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFilteredColimitsOfSize.{w, w'} F] [PreservesFilteredColimitsOfSize.{w, w'} G] : PreservesFilteredColimitsOfSize.{w, w'} (F ⋙ G) where preserves_filtered_colimits _ := inferInstance /-- A functor preserving larger filtered colimits also preserves smaller filtered colimits. -/ -noncomputable def preservesFilteredColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma preservesFilteredColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesFilteredColimitsOfSize.{w', w₂'} F] : PreservesFilteredColimitsOfSize.{w, w₂} F where preserves_filtered_colimits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsFiltered.of_equivalence e.symm - exact preservesColimitsOfShapeOfEquiv e F + exact preservesColimitsOfShape_of_equiv e F /-- -`PreservesFilteredColimitsOfSizeShrink.{w, w'} F` tries to obtain +`PreservesFilteredColimitsOfSize_shrink.{w, w'} F` tries to obtain `PreservesFilteredColimitsOfSize.{w, w'} F` from some other `PreservesFilteredColimitsOfSize F`. -/ -noncomputable def preservesFilteredColimitsOfSizeShrink (F : C ⥤ D) +lemma preservesFilteredColimitsOfSize_shrink (F : C ⥤ D) [PreservesFilteredColimitsOfSize.{max w w₂, max w' w₂'} F] : PreservesFilteredColimitsOfSize.{w, w'} F := - preservesFilteredColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + preservesFilteredColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Preserving filtered colimits at any universe implies preserving filtered colimits at universe `0`. -/ -noncomputable def preservesSmallestFilteredColimitsOfPreservesFilteredColimits (F : C ⥤ D) +lemma preservesSmallestFilteredColimits_of_preservesFilteredColimits (F : C ⥤ D) [PreservesFilteredColimitsOfSize.{w', w} F] : PreservesFilteredColimitsOfSize.{0, 0} F := - preservesFilteredColimitsOfSizeShrink F + preservesFilteredColimitsOfSize_shrink F end Preserves @@ -92,7 +92,7 @@ section Reflects /-- `ReflectsFilteredColimitsOfSize.{w', w} F` means that whenever the image of a filtered cocone under `F` is a colimit cocone, the original cocone was already a colimit. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsFilteredColimitsOfSize (F : C ⥤ D) where +class ReflectsFilteredColimitsOfSize (F : C ⥤ D) : Prop where reflects_filtered_colimits : ∀ (J : Type w) [Category.{w'} J] [IsFiltered J], ReflectsColimitsOfShape J F @@ -101,7 +101,7 @@ A functor is said to reflect filtered colimits, if it reflects all colimits of s `J` is a filtered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev ReflectsFilteredColimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev ReflectsFilteredColimits (F : C ⥤ D) : Prop := ReflectsFilteredColimitsOfSize.{v, v} F attribute [instance 100] ReflectsFilteredColimitsOfSize.reflects_filtered_colimits @@ -110,35 +110,35 @@ instance (priority := 100) ReflectsColimits.reflectsFilteredColimits (F : C ⥤ [ReflectsColimitsOfSize.{w, w'} F] : ReflectsFilteredColimitsOfSize.{w, w'} F where reflects_filtered_colimits _ := inferInstance -instance compReflectsFilteredColimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsFilteredColimits (F : C ⥤ D) (G : D ⥤ E) [ReflectsFilteredColimitsOfSize.{w, w'} F] [ReflectsFilteredColimitsOfSize.{w, w'} G] : ReflectsFilteredColimitsOfSize.{w, w'} (F ⋙ G) where reflects_filtered_colimits _ := inferInstance /-- A functor reflecting larger filtered colimits also reflects smaller filtered colimits. -/ -noncomputable def reflectsFilteredColimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma reflectsFilteredColimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsFilteredColimitsOfSize.{w', w₂'} F] : ReflectsFilteredColimitsOfSize.{w, w₂} F where reflects_filtered_colimits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsFiltered.of_equivalence e.symm - exact reflectsColimitsOfShapeOfEquiv e F + exact reflectsColimitsOfShape_of_equiv e F /-- -`ReflectsFilteredColimitsOfSizeShrink.{w, w'} F` tries to obtain +`ReflectsFilteredColimitsOfSize_shrink.{w, w'} F` tries to obtain `ReflectsFilteredColimitsOfSize.{w, w'} F` from some other `ReflectsFilteredColimitsOfSize F`. -/ -noncomputable def reflectsFilteredColimitsOfSizeShrink (F : C ⥤ D) +lemma reflectsFilteredColimitsOfSize_shrink (F : C ⥤ D) [ReflectsFilteredColimitsOfSize.{max w w₂, max w' w₂'} F] : ReflectsFilteredColimitsOfSize.{w, w'} F := - reflectsFilteredColimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + reflectsFilteredColimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Reflecting filtered colimits at any universe implies reflecting filtered colimits at universe `0`. -/ -noncomputable def reflectsSmallestFilteredColimitsOfReflectsFilteredColimits (F : C ⥤ D) +lemma reflectsSmallestFilteredColimits_of_reflectsFilteredColimits (F : C ⥤ D) [ReflectsFilteredColimitsOfSize.{w', w} F] : ReflectsFilteredColimitsOfSize.{0, 0} F := - reflectsFilteredColimitsOfSizeShrink F + reflectsFilteredColimitsOfSize_shrink F end Reflects @@ -152,7 +152,7 @@ section Preserves /-- `PreservesCofilteredLimitsOfSize.{w', w} F` means that `F` sends all limit cones over any cofiltered diagram `J ⥤ C` to limit cones, where `J : Type w` with `[Category.{w'} J]`. -/ @[nolint checkUnivs, pp_with_univ] -class PreservesCofilteredLimitsOfSize (F : C ⥤ D) where +class PreservesCofilteredLimitsOfSize (F : C ⥤ D) : Prop where preserves_cofiltered_limits : ∀ (J : Type w) [Category.{w'} J] [IsCofiltered J], PreservesLimitsOfShape J F @@ -161,7 +161,7 @@ A functor is said to preserve cofiltered limits, if it preserves all limits of s `J` is a cofiltered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev PreservesCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev PreservesCofilteredLimits (F : C ⥤ D) : Prop := PreservesCofilteredLimitsOfSize.{v, v} F attribute [instance 100] PreservesCofilteredLimitsOfSize.preserves_cofiltered_limits @@ -170,35 +170,35 @@ instance (priority := 100) PreservesLimits.preservesCofilteredLimits (F : C ⥤ [PreservesLimitsOfSize.{w, w'} F] : PreservesCofilteredLimitsOfSize.{w, w'} F where preserves_cofiltered_limits _ := inferInstance -instance compPreservesCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesCofilteredLimitsOfSize.{w, w'} F] [PreservesCofilteredLimitsOfSize.{w, w'} G] : PreservesCofilteredLimitsOfSize.{w, w'} (F ⋙ G) where preserves_cofiltered_limits _ := inferInstance /-- A functor preserving larger cofiltered limits also preserves smaller cofiltered limits. -/ -noncomputable def preservesCofilteredLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma preservesCofilteredLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [PreservesCofilteredLimitsOfSize.{w', w₂'} F] : PreservesCofilteredLimitsOfSize.{w, w₂} F where preserves_cofiltered_limits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsCofiltered.of_equivalence e.symm - exact preservesLimitsOfShapeOfEquiv e F + exact preservesLimitsOfShape_of_equiv e F /-- `PreservesCofilteredLimitsOfSizeShrink.{w, w'} F` tries to obtain `PreservesCofilteredLimitsOfSize.{w, w'} F` from some other `PreservesCofilteredLimitsOfSize F`. -/ -noncomputable def preservesCofilteredLimitsOfSizeShrink (F : C ⥤ D) +lemma preservesCofilteredLimitsOfSize_shrink (F : C ⥤ D) [PreservesCofilteredLimitsOfSize.{max w w₂, max w' w₂'} F] : PreservesCofilteredLimitsOfSize.{w, w'} F := - preservesCofilteredLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + preservesCofilteredLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Preserving cofiltered limits at any universe implies preserving cofiltered limits at universe `0`. -/ -noncomputable def preservesSmallestCofilteredLimitsOfPreservesCofilteredLimits (F : C ⥤ D) +lemma preservesSmallestCofilteredLimits_of_preservesCofilteredLimits (F : C ⥤ D) [PreservesCofilteredLimitsOfSize.{w', w} F] : PreservesCofilteredLimitsOfSize.{0, 0} F := - preservesCofilteredLimitsOfSizeShrink F + preservesCofilteredLimitsOfSize_shrink F end Preserves @@ -208,7 +208,7 @@ section Reflects /-- `ReflectsCofilteredLimitsOfSize.{w', w} F` means that whenever the image of a cofiltered cone under `F` is a limit cone, the original cone was already a limit. -/ @[nolint checkUnivs, pp_with_univ] -class ReflectsCofilteredLimitsOfSize (F : C ⥤ D) where +class ReflectsCofilteredLimitsOfSize (F : C ⥤ D) : Prop where reflects_cofiltered_limits : ∀ (J : Type w) [Category.{w'} J] [IsCofiltered J], ReflectsLimitsOfShape J F @@ -217,7 +217,7 @@ A functor is said to reflect cofiltered limits, if it reflects all limits of sha `J` is a cofiltered category which is small relative to the universe in which morphisms of the source live. -/ -abbrev ReflectsCofilteredLimits (F : C ⥤ D) : Type max u₁ u₂ (v + 1) := +abbrev ReflectsCofilteredLimits (F : C ⥤ D) : Prop := ReflectsCofilteredLimitsOfSize.{v, v} F attribute [instance 100] ReflectsCofilteredLimitsOfSize.reflects_cofiltered_limits @@ -226,35 +226,35 @@ instance (priority := 100) ReflectsLimits.reflectsCofilteredLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{w, w'} F] : ReflectsCofilteredLimitsOfSize.{w, w'} F where reflects_cofiltered_limits _ := inferInstance -instance compReflectsCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsCofilteredLimits (F : C ⥤ D) (G : D ⥤ E) [ReflectsCofilteredLimitsOfSize.{w, w'} F] [ReflectsCofilteredLimitsOfSize.{w, w'} G] : ReflectsCofilteredLimitsOfSize.{w, w'} (F ⋙ G) where reflects_cofiltered_limits _ := inferInstance /-- A functor reflecting larger cofiltered limits also reflects smaller cofiltered limits. -/ -noncomputable def reflectsCofilteredLimitsOfSizeOfUnivLE (F : C ⥤ D) [UnivLE.{w, w'}] +lemma reflectsCofilteredLimitsOfSize_of_univLE (F : C ⥤ D) [UnivLE.{w, w'}] [UnivLE.{w₂, w₂'}] [ReflectsCofilteredLimitsOfSize.{w', w₂'} F] : ReflectsCofilteredLimitsOfSize.{w, w₂} F where reflects_cofiltered_limits J _ _ := by let e := ((ShrinkHoms.equivalence J).trans <| Shrink.equivalence _).symm haveI := IsCofiltered.of_equivalence e.symm - exact reflectsLimitsOfShapeOfEquiv e F + exact reflectsLimitsOfShape_of_equiv e F /-- -`ReflectsCofilteredLimitsOfSizeShrink.{w, w'} F` tries to obtain +`ReflectsCofilteredLimitsOfSize_shrink.{w, w'} F` tries to obtain `ReflectsCofilteredLimitsOfSize.{w, w'} F` from some other `ReflectsCofilteredLimitsOfSize F`. -/ -noncomputable def reflectsCofilteredLimitsOfSizeShrink (F : C ⥤ D) +lemma reflectsCofilteredLimitsOfSize_shrink (F : C ⥤ D) [ReflectsCofilteredLimitsOfSize.{max w w₂, max w' w₂'} F] : ReflectsCofilteredLimitsOfSize.{w, w'} F := - reflectsCofilteredLimitsOfSizeOfUnivLE.{max w w₂, max w' w₂'} F + reflectsCofilteredLimitsOfSize_of_univLE.{max w w₂, max w' w₂'} F /-- Reflecting cofiltered limits at any universe implies reflecting cofiltered limits at universe `0`. -/ -noncomputable def reflectsSmallestCofilteredLimitsOfReflectsCofilteredLimits (F : C ⥤ D) +lemma reflectsSmallestCofilteredLimits_of_reflectsCofilteredLimits (F : C ⥤ D) [ReflectsCofilteredLimitsOfSize.{w', w} F] : ReflectsCofilteredLimitsOfSize.{0, 0} F := - reflectsCofilteredLimitsOfSizeShrink F + reflectsCofilteredLimitsOfSize_shrink F end Reflects diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean index e852cc4007abb..89f2b68d3e856 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Finite.lean @@ -37,127 +37,127 @@ variable {J : Type w} [SmallCategory J] {K : J ⥤ C} /-- A functor is said to preserve finite limits, if it preserves all limits of shape `J`, where `J : Type` is a finite category. -/ -class PreservesFiniteLimits (F : C ⥤ D) where +class PreservesFiniteLimits (F : C ⥤ D) : Prop where preservesFiniteLimits : ∀ (J : Type) [SmallCategory J] [FinCategory J], PreservesLimitsOfShape J F := by infer_instance attribute [instance] PreservesFiniteLimits.preservesFiniteLimits instance (F : C ⥤ D) : Subsingleton (PreservesFiniteLimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /-- Preserving finite limits also implies preserving limits over finite shapes in higher universes, though through a noncomputable instance. -/ -noncomputable instance (priority := 100) preservesLimitsOfShapeOfPreservesFiniteLimits (F : C ⥤ D) +instance (priority := 100) preservesLimitsOfShapeOfPreservesFiniteLimits (F : C ⥤ D) [PreservesFiniteLimits F] (J : Type w) [SmallCategory J] [FinCategory J] : PreservesLimitsOfShape J F := by - apply preservesLimitsOfShapeOfEquiv (FinCategory.equivAsType J) + apply preservesLimitsOfShape_of_equiv (FinCategory.equivAsType J) -- This is a dangerous instance as it has unbound universe variables. /-- If we preserve limits of some arbitrary size, then we preserve all finite limits. -/ -noncomputable def PreservesLimitsOfSize.preservesFiniteLimits (F : C ⥤ D) +lemma PreservesLimitsOfSize.preservesFiniteLimits (F : C ⥤ D) [PreservesLimitsOfSize.{w, w₂} F] : PreservesFiniteLimits F where preservesFiniteLimits J (sJ : SmallCategory J) fJ := by - haveI := preservesSmallestLimitsOfPreservesLimits F - exact preservesLimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := preservesSmallestLimits_of_preservesLimits F + exact preservesLimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for limits indexed in Type 0. -noncomputable instance (priority := 120) PreservesLimitsOfSize0.preservesFiniteLimits +instance (priority := 120) PreservesLimitsOfSize0.preservesFiniteLimits (F : C ⥤ D) [PreservesLimitsOfSize.{0, 0} F] : PreservesFiniteLimits F := PreservesLimitsOfSize.preservesFiniteLimits F -- An alternative specialization of the dangerous instance for small limits. -noncomputable instance (priority := 120) PreservesLimits.preservesFiniteLimits (F : C ⥤ D) +instance (priority := 120) PreservesLimits.preservesFiniteLimits (F : C ⥤ D) [PreservesLimits F] : PreservesFiniteLimits F := PreservesLimitsOfSize.preservesFiniteLimits F /-- We can always derive `PreservesFiniteLimits C` by showing that we are preserving limits at an arbitrary universe. -/ -def preservesFiniteLimitsOfPreservesFiniteLimitsOfSize (F : C ⥤ D) +lemma preservesFiniteLimits_of_preservesFiniteLimitsOfSize (F : C ⥤ D) (h : ∀ (J : Type w) {𝒥 : SmallCategory J} (_ : @FinCategory J 𝒥), PreservesLimitsOfShape J F) : PreservesFiniteLimits F where preservesFiniteLimits J (_ : SmallCategory J) _ := by letI : Category (ULiftHom (ULift J)) := ULiftHom.category haveI := h (ULiftHom (ULift J)) CategoryTheory.finCategoryUlift - exact preservesLimitsOfShapeOfEquiv (ULiftHomULiftCategory.equiv J).symm F + exact preservesLimitsOfShape_of_equiv (ULiftHomULiftCategory.equiv J).symm F /-- The composition of two left exact functors is left exact. -/ -def compPreservesFiniteLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits F] +lemma comp_preservesFiniteLimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits F] [PreservesFiniteLimits G] : PreservesFiniteLimits (F ⋙ G) := ⟨fun _ _ _ => inferInstance⟩ /-- Transfer preservation of finite limits along a natural isomorphism in the functor. -/ -def preservesFiniteLimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteLimits F] : +lemma preservesFiniteLimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteLimits F] : PreservesFiniteLimits G where - preservesFiniteLimits _ _ _ := preservesLimitsOfShapeOfNatIso h + preservesFiniteLimits _ _ _ := preservesLimitsOfShape_of_natIso h /- Porting note: adding this class because quantified classes don't behave well https://github.com/leanprover-community/mathlib4/pull/2764 -/ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` for `Fintype J` -/ -class PreservesFiniteProducts (F : C ⥤ D) where +class PreservesFiniteProducts (F : C ⥤ D) : Prop where preserves : ∀ (J : Type) [Fintype J], PreservesLimitsOfShape (Discrete J) F attribute [instance] PreservesFiniteProducts.preserves instance (F : C ⥤ D) : Subsingleton (PreservesFiniteProducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -noncomputable instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] +instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] [PreservesFiniteProducts F] : PreservesLimitsOfShape (Discrete J) F := by apply Nonempty.some obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J - exact ⟨preservesLimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F⟩ + exact ⟨preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F⟩ -instance compPreservesFiniteProducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesFiniteProducts (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteProducts F] [PreservesFiniteProducts G] : PreservesFiniteProducts (F ⋙ G) where preserves _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteProducts F where +instance (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteProducts F where preserves _ _ := inferInstance /-- A functor is said to reflect finite limits, if it reflects all limits of shape `J`, where `J : Type` is a finite category. -/ -class ReflectsFiniteLimits (F : C ⥤ D) where +class ReflectsFiniteLimits (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [SmallCategory J] [FinCategory J], ReflectsLimitsOfShape J F := by infer_instance attribute [instance] ReflectsFiniteLimits.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteLimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /- Similarly to preserving finite products, quantified classes don't behave well. -/ /-- A functor `F` preserves finite products if it reflects limits of shape `Discrete J` for finite `J` -/ -class ReflectsFiniteProducts (F : C ⥤ D) where +class ReflectsFiniteProducts (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [Fintype J], ReflectsLimitsOfShape (Discrete J) F attribute [instance] ReflectsFiniteProducts.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteProducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -- This is a dangerous instance as it has unbound universe variables. /-- If we reflect limits of some arbitrary size, then we reflect all finite limits. -/ -noncomputable def ReflectsLimitsOfSize.reflectsFiniteLimits +lemma ReflectsLimitsOfSize.reflectsFiniteLimits (F : C ⥤ D) [ReflectsLimitsOfSize.{w, w₂} F] : ReflectsFiniteLimits F where reflects J (sJ : SmallCategory J) fJ := by - haveI := reflectsSmallestLimitsOfReflectsLimits F - exact reflectsLimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := reflectsSmallestLimits_of_reflectsLimits F + exact reflectsLimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. -noncomputable instance (priority := 120) (F : C ⥤ D) [ReflectsLimitsOfSize.{0, 0} F] : +instance (priority := 120) (F : C ⥤ D) [ReflectsLimitsOfSize.{0, 0} F] : ReflectsFiniteLimits F := ReflectsLimitsOfSize.reflectsFiniteLimits F -- An alternative specialization of the dangerous instance for small colimits. -noncomputable instance (priority := 120) (F : C ⥤ D) +instance (priority := 120) (F : C ⥤ D) [ReflectsLimits F] : ReflectsFiniteLimits F := ReflectsLimitsOfSize.reflectsFiniteLimits F @@ -165,40 +165,40 @@ noncomputable instance (priority := 120) (F : C ⥤ D) If `F ⋙ G` preserves finite limits and `G` reflects finite limits, then `F` preserves finite limits. -/ -def preservesFiniteLimitsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteLimits_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteLimits (F ⋙ G)] [ReflectsFiniteLimits G] : PreservesFiniteLimits F where - preservesFiniteLimits _ _ _ := preservesLimitsOfShapeOfReflectsOfPreserves F G + preservesFiniteLimits _ _ _ := preservesLimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves finite products and `G` reflects finite products, then `F` preserves finite products. -/ -def preservesFiniteProductsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteProducts_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteProducts (F ⋙ G)] [ReflectsFiniteProducts G] : PreservesFiniteProducts F where - preserves _ _ := preservesLimitsOfShapeOfReflectsOfPreserves F G + preserves _ _ := preservesLimitsOfShape_of_reflects_of_preserves F G -noncomputable instance reflectsFiniteLimitsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteLimits_of_reflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteLimits C] [PreservesFiniteLimits F] : ReflectsFiniteLimits F where - reflects _ _ _ := reflectsLimitsOfShapeOfReflectsIsomorphisms + reflects _ _ _ := reflectsLimitsOfShape_of_reflectsIsomorphisms -noncomputable instance reflectsFiniteProductsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteProducts_of_reflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteProducts C] [PreservesFiniteProducts F] : ReflectsFiniteProducts F where - reflects _ _ := reflectsLimitsOfShapeOfReflectsIsomorphisms + reflects _ _ := reflectsLimitsOfShape_of_reflectsIsomorphisms -instance compReflectsFiniteProducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsFiniteProducts (F : C ⥤ D) (G : D ⥤ E) [ReflectsFiniteProducts F] [ReflectsFiniteProducts G] : ReflectsFiniteProducts (F ⋙ G) where reflects _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [ReflectsFiniteLimits F] : ReflectsFiniteProducts F where +instance (F : C ⥤ D) [ReflectsFiniteLimits F] : ReflectsFiniteProducts F where reflects _ _ := inferInstance /-- A functor is said to preserve finite colimits, if it preserves all colimits of shape `J`, where `J : Type` is a finite category. -/ -class PreservesFiniteColimits (F : C ⥤ D) where +class PreservesFiniteColimits (F : C ⥤ D) : Prop where preservesFiniteColimits : ∀ (J : Type) [SmallCategory J] [FinCategory J], PreservesColimitsOfShape J F := by infer_instance @@ -206,111 +206,111 @@ class PreservesFiniteColimits (F : C ⥤ D) where attribute [instance] PreservesFiniteColimits.preservesFiniteColimits instance (F : C ⥤ D) : Subsingleton (PreservesFiniteColimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /-- Preserving finite colimits also implies preserving colimits over finite shapes in higher -universes, though through a noncomputable instance. +universes. -/ -noncomputable instance (priority := 100) preservesColimitsOfShapeOfPreservesFiniteColimits +instance (priority := 100) preservesColimitsOfShapeOfPreservesFiniteColimits (F : C ⥤ D) [PreservesFiniteColimits F] (J : Type w) [SmallCategory J] [FinCategory J] : PreservesColimitsOfShape J F := by - apply preservesColimitsOfShapeOfEquiv (FinCategory.equivAsType J) + apply preservesColimitsOfShape_of_equiv (FinCategory.equivAsType J) -- This is a dangerous instance as it has unbound universe variables. /-- If we preserve colimits of some arbitrary size, then we preserve all finite colimits. -/ -noncomputable def PreservesColimitsOfSize.preservesFiniteColimits (F : C ⥤ D) +lemma PreservesColimitsOfSize.preservesFiniteColimits (F : C ⥤ D) [PreservesColimitsOfSize.{w, w₂} F] : PreservesFiniteColimits F where preservesFiniteColimits J (sJ : SmallCategory J) fJ := by - haveI := preservesSmallestColimitsOfPreservesColimits F - exact preservesColimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := preservesSmallestColimits_of_preservesColimits F + exact preservesColimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. -noncomputable instance (priority := 120) PreservesColimitsOfSize0.preservesFiniteColimits +instance (priority := 120) PreservesColimitsOfSize0.preservesFiniteColimits (F : C ⥤ D) [PreservesColimitsOfSize.{0, 0} F] : PreservesFiniteColimits F := PreservesColimitsOfSize.preservesFiniteColimits F -- An alternative specialization of the dangerous instance for small colimits. -noncomputable instance (priority := 120) PreservesColimits.preservesFiniteColimits (F : C ⥤ D) +instance (priority := 120) PreservesColimits.preservesFiniteColimits (F : C ⥤ D) [PreservesColimits F] : PreservesFiniteColimits F := PreservesColimitsOfSize.preservesFiniteColimits F /-- We can always derive `PreservesFiniteColimits C` by showing that we are preserving colimits at an arbitrary universe. -/ -def preservesFiniteColimitsOfPreservesFiniteColimitsOfSize (F : C ⥤ D) +lemma preservesFiniteColimits_of_preservesFiniteColimitsOfSize (F : C ⥤ D) (h : ∀ (J : Type w) {𝒥 : SmallCategory J} (_ : @FinCategory J 𝒥), PreservesColimitsOfShape J F) : PreservesFiniteColimits F where preservesFiniteColimits J (_ : SmallCategory J) _ := by letI : Category (ULiftHom (ULift J)) := ULiftHom.category haveI := h (ULiftHom (ULift J)) CategoryTheory.finCategoryUlift - exact preservesColimitsOfShapeOfEquiv (ULiftHomULiftCategory.equiv J).symm F + exact preservesColimitsOfShape_of_equiv (ULiftHomULiftCategory.equiv J).symm F /-- The composition of two right exact functors is right exact. -/ -def compPreservesFiniteColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColimits F] +lemma comp_preservesFiniteColimits (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColimits F] [PreservesFiniteColimits G] : PreservesFiniteColimits (F ⋙ G) := ⟨fun _ _ _ => inferInstance⟩ /-- Transfer preservation of finite colimits along a natural isomorphism in the functor. -/ -def preservesFiniteColimitsOfNatIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteColimits F] : +lemma preservesFiniteColimits_of_natIso {F G : C ⥤ D} (h : F ≅ G) [PreservesFiniteColimits F] : PreservesFiniteColimits G where - preservesFiniteColimits _ _ _ := preservesColimitsOfShapeOfNatIso h + preservesFiniteColimits _ _ _ := preservesColimitsOfShape_of_natIso h /- Porting note: adding this class because quantified classes don't behave well https://github.com/leanprover-community/mathlib4/pull/2764 -/ /-- A functor `F` preserves finite products if it preserves all from `Discrete J` for `Fintype J` -/ -class PreservesFiniteCoproducts (F : C ⥤ D) where +class PreservesFiniteCoproducts (F : C ⥤ D) : Prop where /-- preservation of colimits indexed by `Discrete J` when `[Fintype J]` -/ preserves : ∀ (J : Type) [Fintype J], PreservesColimitsOfShape (Discrete J) F attribute [instance] PreservesFiniteCoproducts.preserves instance (F : C ⥤ D) : Subsingleton (PreservesFiniteCoproducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -noncomputable instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] +instance (priority := 100) (F : C ⥤ D) (J : Type u) [Finite J] [PreservesFiniteCoproducts F] : PreservesColimitsOfShape (Discrete J) F := by apply Nonempty.some obtain ⟨n, ⟨e⟩⟩ := Finite.exists_equiv_fin J - exact ⟨preservesColimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F⟩ + exact ⟨preservesColimitsOfShape_of_equiv (Discrete.equivalence e.symm) F⟩ -instance compPreservesFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_preservesFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteCoproducts F] [PreservesFiniteCoproducts G] : PreservesFiniteCoproducts (F ⋙ G) where preserves _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteCoproducts F where +instance (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteCoproducts F where preserves _ _ := inferInstance /-- A functor is said to reflect finite colimits, if it reflects all colimits of shape `J`, where `J : Type` is a finite category. -/ -class ReflectsFiniteColimits (F : C ⥤ D) where +class ReflectsFiniteColimits (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [SmallCategory J] [FinCategory J], ReflectsColimitsOfShape J F := by infer_instance attribute [instance] ReflectsFiniteColimits.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteColimits F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ -- This is a dangerous instance as it has unbound universe variables. /-- If we reflect colimits of some arbitrary size, then we reflect all finite colimits. -/ -noncomputable def ReflectsColimitsOfSize.reflectsFiniteColimits +lemma ReflectsColimitsOfSize.reflectsFiniteColimits (F : C ⥤ D) [ReflectsColimitsOfSize.{w, w₂} F] : ReflectsFiniteColimits F where reflects J (sJ : SmallCategory J) fJ := by - haveI := reflectsSmallestColimitsOfReflectsColimits F - exact reflectsColimitsOfShapeOfEquiv (FinCategory.equivAsType J) F + haveI := reflectsSmallestColimits_of_reflectsColimits F + exact reflectsColimitsOfShape_of_equiv (FinCategory.equivAsType J) F -- Added as a specialization of the dangerous instance above, for colimits indexed in Type 0. -noncomputable instance (priority := 120) (F : C ⥤ D) [ReflectsColimitsOfSize.{0, 0} F] : +instance (priority := 120) (F : C ⥤ D) [ReflectsColimitsOfSize.{0, 0} F] : ReflectsFiniteColimits F := ReflectsColimitsOfSize.reflectsFiniteColimits F -- An alternative specialization of the dangerous instance for small colimits. -noncomputable instance (priority := 120) (F : C ⥤ D) +instance (priority := 120) (F : C ⥤ D) [ReflectsColimits F] : ReflectsFiniteColimits F := ReflectsColimitsOfSize.reflectsFiniteColimits F @@ -319,47 +319,47 @@ noncomputable instance (priority := 120) (F : C ⥤ D) A functor `F` preserves finite coproducts if it reflects colimits of shape `Discrete J` for finite `J` -/ -class ReflectsFiniteCoproducts (F : C ⥤ D) where +class ReflectsFiniteCoproducts (F : C ⥤ D) : Prop where reflects : ∀ (J : Type) [Fintype J], ReflectsColimitsOfShape (Discrete J) F attribute [instance] ReflectsFiniteCoproducts.reflects instance (F : C ⥤ D) : Subsingleton (ReflectsFiniteCoproducts F) := - ⟨fun ⟨a⟩ ⟨b⟩ => by congr; subsingleton⟩ + ⟨fun ⟨a⟩ ⟨b⟩ => by congr⟩ /-- If `F ⋙ G` preserves finite colimits and `G` reflects finite colimits, then `F` preserves finite colimits. -/ -def preservesFiniteColimitsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteColimits_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteColimits (F ⋙ G)] [ReflectsFiniteColimits G] : PreservesFiniteColimits F where - preservesFiniteColimits _ _ _ := preservesColimitsOfShapeOfReflectsOfPreserves F G + preservesFiniteColimits _ _ _ := preservesColimitsOfShape_of_reflects_of_preserves F G /-- If `F ⋙ G` preserves finite coproducts and `G` reflects finite coproducts, then `F` preserves finite coproducts. -/ -noncomputable def preservesFiniteCoproductsOfReflectsOfPreserves (F : C ⥤ D) (G : D ⥤ E) +lemma preservesFiniteCoproducts_of_reflects_of_preserves (F : C ⥤ D) (G : D ⥤ E) [PreservesFiniteCoproducts (F ⋙ G)] [ReflectsFiniteCoproducts G] : PreservesFiniteCoproducts F where - preserves _ _ := preservesColimitsOfShapeOfReflectsOfPreserves F G + preserves _ _ := preservesColimitsOfShape_of_reflects_of_preserves F G -noncomputable instance reflectsFiniteColimitsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteColimitsOfReflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteColimits C] [PreservesFiniteColimits F] : ReflectsFiniteColimits F where - reflects _ _ _ := reflectsColimitsOfShapeOfReflectsIsomorphisms + reflects _ _ _ := reflectsColimitsOfShape_of_reflectsIsomorphisms -noncomputable instance reflectsFiniteCoproductsOfReflectsIsomorphisms (F : C ⥤ D) +instance reflectsFiniteCoproductsOfReflectsIsomorphisms (F : C ⥤ D) [F.ReflectsIsomorphisms] [HasFiniteCoproducts C] [PreservesFiniteCoproducts F] : ReflectsFiniteCoproducts F where - reflects _ _ := reflectsColimitsOfShapeOfReflectsIsomorphisms + reflects _ _ := reflectsColimitsOfShape_of_reflectsIsomorphisms -instance compReflectsFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) +instance comp_reflectsFiniteCoproducts (F : C ⥤ D) (G : D ⥤ E) [ReflectsFiniteCoproducts F] [ReflectsFiniteCoproducts G] : ReflectsFiniteCoproducts (F ⋙ G) where reflects _ _ := inferInstance -noncomputable instance (F : C ⥤ D) [ReflectsFiniteColimits F] : ReflectsFiniteCoproducts F where +instance (F : C ⥤ D) [ReflectsFiniteColimits F] : ReflectsFiniteCoproducts F where reflects _ _ := inferInstance end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean index d3030d269cc41..baf47d8eff349 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/FunctorCategory.lean @@ -54,14 +54,13 @@ work to convert to this version: namely, the natural isomorphism `(evaluation C D).obj k ⋙ prod.functor.obj (F.obj k) ≅ prod.functor.obj F ⋙ (evaluation C D).obj k` -/ -def FunctorCategory.prodPreservesColimits [HasBinaryProducts D] [HasColimits D] +lemma FunctorCategory.prod_preservesColimits [HasBinaryProducts D] [HasColimits D] [∀ X : D, PreservesColimits (prod.functor.obj X)] (F : C ⥤ D) : PreservesColimits (prod.functor.obj F) where preservesColimitsOfShape {J : Type u} [Category.{u, u} J] := { - preservesColimit := fun {K : J ⥤ C ⥤ D} => - ( { - preserves := fun {c : Cocone K} (t : IsColimit c) => by + preservesColimit := fun {K : J ⥤ C ⥤ D} => ({ + preserves := fun {c : Cocone K} (t : IsColimit c) => ⟨by apply evaluationJointlyReflectsColimits _ fun {k} => ?_ change IsColimit ((prod.functor.obj F ⋙ (evaluation _ _).obj k).mapCocone c) let this := @@ -71,7 +70,7 @@ def FunctorCategory.prodPreservesColimits [HasBinaryProducts D] [HasColimits D] · intro G apply asIso (prodComparison ((evaluation C D).obj k) F G) · intro G G' - apply prodComparison_natural ((evaluation C D).obj k) (𝟙 F) } ) } + apply prodComparison_natural ((evaluation C D).obj k) (𝟙 F)⟩ } ) } end @@ -81,43 +80,43 @@ variable {C : Type u₁} [Category.{v₁} C] variable {D : Type u₂} [Category.{v₂} D] variable {E : Type u₃} [Category.{v₃} E] -instance whiskeringLeftPreservesLimitsOfShape (J : Type u) [Category.{v} J] +instance whiskeringLeft_preservesLimitsOfShape (J : Type u) [Category.{v} J] [HasLimitsOfShape J D] (F : C ⥤ E) : PreservesLimitsOfShape J ((whiskeringLeft C E D).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsLimits intro Y change IsLimit (((evaluation E D).obj (F.obj Y)).mapCone c) - exact PreservesLimit.preserves hc⟩⟩ + exact isLimitOfPreserves _ hc⟩⟩⟩ -instance whiskeringLeftPreservesColimitsOfShape (J : Type u) [Category.{v} J] +instance whiskeringLeft_preservesColimitsOfShape (J : Type u) [Category.{v} J] [HasColimitsOfShape J D] (F : C ⥤ E) : PreservesColimitsOfShape J ((whiskeringLeft C E D).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsColimits intro Y change IsColimit (((evaluation E D).obj (F.obj Y)).mapCocone c) - exact PreservesColimit.preserves hc⟩⟩ + exact isColimitOfPreserves _ hc⟩⟩⟩ -instance whiskeringLeftPreservesLimits [HasLimitsOfSize.{w} D] (F : C ⥤ E) : +instance whiskeringLeft_preservesLimits [HasLimitsOfSize.{w, w'} D] (F : C ⥤ E) : PreservesLimitsOfSize.{w, w'} ((whiskeringLeft C E D).obj F) := - ⟨fun {J} _ => whiskeringLeftPreservesLimitsOfShape J F⟩ + ⟨fun {J} _ => whiskeringLeft_preservesLimitsOfShape J F⟩ -instance whiskeringLeftPreservesColimit [HasColimitsOfSize.{w} D] (F : C ⥤ E) : +instance whiskeringLeft_preservesColimit [HasColimitsOfSize.{w, w'} D] (F : C ⥤ E) : PreservesColimitsOfSize.{w, w'} ((whiskeringLeft C E D).obj F) := - ⟨fun {J} _ => whiskeringLeftPreservesColimitsOfShape J F⟩ + ⟨fun {J} _ => whiskeringLeft_preservesColimitsOfShape J F⟩ -instance whiskeringRightPreservesLimitsOfShape {C : Type*} [Category C] {D : Type*} +instance whiskeringRight_preservesLimitsOfShape {C : Type*} [Category C] {D : Type*} [Category D] {E : Type*} [Category E] {J : Type*} [Category J] [HasLimitsOfShape J D] (F : D ⥤ E) [PreservesLimitsOfShape J F] : PreservesLimitsOfShape J ((whiskeringRight C D E).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsLimits _ (fun k => ?_) change IsLimit (((evaluation _ _).obj k ⋙ F).mapCone c) - exact PreservesLimit.preserves hc⟩⟩ + exact isLimitOfPreserves _ hc⟩⟩⟩ /-- Whiskering right and then taking a limit is the same as taking the limit and applying the functor. -/ @@ -143,15 +142,15 @@ theorem limitCompWhiskeringRightIsoLimitComp_hom_whiskerRight_π {C : Type*} [Ca limit.π (G ⋙ (whiskeringRight _ _ _).obj F) j := by simp [← Iso.eq_inv_comp] -instance whiskeringRightPreservesColimitsOfShape {C : Type*} [Category C] {D : Type*} +instance whiskeringRight_preservesColimitsOfShape {C : Type*} [Category C] {D : Type*} [Category D] {E : Type*} [Category E] {J : Type*} [Category J] [HasColimitsOfShape J D] (F : D ⥤ E) [PreservesColimitsOfShape J F] : PreservesColimitsOfShape J ((whiskeringRight C D E).obj F) := ⟨fun {K} => - ⟨fun c {hc} => by + ⟨fun c {hc} => ⟨by apply evaluationJointlyReflectsColimits _ (fun k => ?_) change IsColimit (((evaluation _ _).obj k ⋙ F).mapCocone c) - exact PreservesColimit.preserves hc⟩⟩ + exact isColimitOfPreserves _ hc⟩⟩⟩ /-- Whiskering right and then taking a colimit is the same as taking the colimit and applying the functor. -/ @@ -191,23 +190,23 @@ instance whiskeringRightPreservesColimits {C : Type*} [Category C] {D : Type*} [ -- Porting note: fixed spelling mistake in def /-- If `Lan F.op : (Cᵒᵖ ⥤ Type*) ⥤ (Dᵒᵖ ⥤ Type*)` preserves limits of shape `J`, so will `F`. -/ -noncomputable def preservesLimitOfLanPreservesLimit {C D : Type u} [SmallCategory C] +lemma preservesLimit_of_lan_preservesLimit {C D : Type u} [SmallCategory C] [SmallCategory D] (F : C ⥤ D) (J : Type u) [SmallCategory J] [PreservesLimitsOfShape J (F.op.lan : _ ⥤ Dᵒᵖ ⥤ Type u)] : PreservesLimitsOfShape J F := by - apply @preservesLimitsOfShapeOfReflectsOfPreserves _ _ _ _ _ _ _ _ F yoneda ?_ - exact preservesLimitsOfShapeOfNatIso (Presheaf.compYonedaIsoYonedaCompLan F).symm + apply @preservesLimitsOfShape_of_reflects_of_preserves _ _ _ _ _ _ _ _ F yoneda ?_ + exact preservesLimitsOfShape_of_natIso (Presheaf.compYonedaIsoYonedaCompLan F).symm /-- `F : C ⥤ D ⥤ E` preserves finite limits if it does for each `d : D`. -/ -def preservesFiniteLimitsOfEvaluation {D : Type*} [Category D] {E : Type*} [Category E] +lemma preservesFiniteLimits_of_evaluation {D : Type*} [Category D] {E : Type*} [Category E] (F : C ⥤ D ⥤ E) (h : ∀ d : D, PreservesFiniteLimits (F ⋙ (evaluation D E).obj d)) : PreservesFiniteLimits F := - ⟨fun J _ _ => preservesLimitsOfShapeOfEvaluation F J fun k => (h k).preservesFiniteLimits _⟩ + ⟨fun J _ _ => preservesLimitsOfShape_of_evaluation F J fun k => (h k).preservesFiniteLimits _⟩ /-- `F : C ⥤ D ⥤ E` preserves finite limits if it does for each `d : D`. -/ -def preservesFiniteColimitsOfEvaluation {D : Type*} [Category D] {E : Type*} [Category E] +lemma preservesFiniteColimits_of_evaluation {D : Type*} [Category D] {E : Type*} [Category E] (F : C ⥤ D ⥤ E) (h : ∀ d : D, PreservesFiniteColimits (F ⋙ (evaluation D E).obj d)) : PreservesFiniteColimits F := - ⟨fun J _ _ => preservesColimitsOfShapeOfEvaluation F J fun k => (h k).preservesFiniteColimits _⟩ + ⟨fun J _ _ => preservesColimitsOfShape_of_evaluation F J fun k => (h k).preservesFiniteColimits _⟩ end @@ -224,14 +223,14 @@ variable [HasLimitsOfShape J C] [HasColimitsOfShape K C] variable [PreservesLimitsOfShape J (colim : (K ⥤ C) ⥤ _)] noncomputable instance : PreservesLimitsOfShape J (colim : (K ⥤ D ⥤ C) ⥤ _) := - preservesLimitsOfShapeOfEvaluation _ _ (fun d => + preservesLimitsOfShape_of_evaluation _ _ (fun d => let i : (colim : (K ⥤ D ⥤ C) ⥤ _) ⋙ (evaluation D C).obj d ≅ colimit ((whiskeringRight K (D ⥤ C) C).obj ((evaluation D C).obj d)).flip := NatIso.ofComponents (fun X => (colimitObjIsoColimitCompEvaluation _ _) ≪≫ (by exact HasColimit.isoOfNatIso (Iso.refl _)) ≪≫ (colimitObjIsoColimitCompEvaluation _ _).symm) (fun {F G} η => colimit_obj_ext (fun j => by simp [← NatTrans.comp_app_assoc])) - preservesLimitsOfShapeOfNatIso (i ≪≫ colimitFlipIsoCompColim _).symm) + preservesLimitsOfShape_of_natIso (i ≪≫ colimitFlipIsoCompColim _).symm) end @@ -241,14 +240,14 @@ variable [HasColimitsOfShape J C] [HasLimitsOfShape K C] variable [PreservesColimitsOfShape J (lim : (K ⥤ C) ⥤ _)] noncomputable instance : PreservesColimitsOfShape J (lim : (K ⥤ D ⥤ C) ⥤ _) := - preservesColimitsOfShapeOfEvaluation _ _ (fun d => + preservesColimitsOfShape_of_evaluation _ _ (fun d => let i : (lim : (K ⥤ D ⥤ C) ⥤ _) ⋙ (evaluation D C).obj d ≅ limit ((whiskeringRight K (D ⥤ C) C).obj ((evaluation D C).obj d)).flip := NatIso.ofComponents (fun X => (limitObjIsoLimitCompEvaluation _ _) ≪≫ (by exact HasLimit.isoOfNatIso (Iso.refl _)) ≪≫ (limitObjIsoLimitCompEvaluation _ _).symm) (fun {F G} η => limit_obj_ext (fun j => by simp [← NatTrans.comp_app])) - preservesColimitsOfShapeOfNatIso (i ≪≫ limitFlipIsoCompLim _).symm) + preservesColimitsOfShape_of_natIso (i ≪≫ limitFlipIsoCompLim _).symm) end diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean b/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean index ac77b2c0bcd51..f0517bf3e7d4b 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Limits.lean @@ -38,8 +38,8 @@ variable [PreservesLimit F G] @[simp] theorem preserves_lift_mapCone (c₁ c₂ : Cone F) (t : IsLimit c₁) : - (PreservesLimit.preserves t).lift (G.mapCone c₂) = G.map (t.lift c₂) := - ((PreservesLimit.preserves t).uniq (G.mapCone c₂) _ (by simp [← G.map_comp])).symm + (isLimitOfPreserves G t).lift (G.mapCone c₂) = G.map (t.lift c₂) := + ((isLimitOfPreserves G t).uniq (G.mapCone c₂) _ (by simp [← G.map_comp])).symm variable [HasLimit F] @@ -47,7 +47,7 @@ variable [HasLimit F] to the limit of the functor `F ⋙ G`. -/ def preservesLimitIso : G.obj (limit F) ≅ limit (F ⋙ G) := - (PreservesLimit.preserves (limit.isLimit _)).conePointUniqueUpToIso (limit.isLimit _) + (isLimitOfPreserves G (limit.isLimit _)).conePointUniqueUpToIso (limit.isLimit _) @[reassoc (attr := simp)] theorem preservesLimitIso_hom_π (j) : @@ -98,8 +98,8 @@ variable [HasLimit F] [HasLimit (F ⋙ G)] /-- If the comparison morphism `G.obj (limit F) ⟶ limit (F ⋙ G)` is an isomorphism, then `G` preserves limits of `F`. -/ -def preservesLimitOfIsIsoPost [IsIso (limit.post F G)] : PreservesLimit F G := - preservesLimitOfPreservesLimitCone (limit.isLimit F) (by +lemma preservesLimit_of_isIso_post [IsIso (limit.post F G)] : PreservesLimit F G := + preservesLimit_of_preserves_limit_cone (limit.isLimit F) (by convert IsLimit.ofPointIso (limit.isLimit (F ⋙ G)) assumption) @@ -111,8 +111,8 @@ variable [PreservesColimit F G] @[simp] theorem preserves_desc_mapCocone (c₁ c₂ : Cocone F) (t : IsColimit c₁) : - (PreservesColimit.preserves t).desc (G.mapCocone _) = G.map (t.desc c₂) := - ((PreservesColimit.preserves t).uniq (G.mapCocone _) _ (by simp [← G.map_comp])).symm + (isColimitOfPreserves G t).desc (G.mapCocone _) = G.map (t.desc c₂) := + ((isColimitOfPreserves G t).uniq (G.mapCocone _) _ (by simp [← G.map_comp])).symm variable [HasColimit F] @@ -121,7 +121,7 @@ variable [HasColimit F] to the colimit of the functor `F ⋙ G`. -/ def preservesColimitIso : G.obj (colimit F) ≅ colimit (F ⋙ G) := - (PreservesColimit.preserves (colimit.isColimit _)).coconePointUniqueUpToIso (colimit.isColimit _) + (isColimitOfPreserves G (colimit.isColimit _)).coconePointUniqueUpToIso (colimit.isColimit _) @[reassoc (attr := simp)] theorem ι_preservesColimitIso_inv (j : J) : @@ -134,7 +134,7 @@ alias ι_preservesColimitsIso_inv := ι_preservesColimitIso_inv @[reassoc (attr := simp)] theorem ι_preservesColimitIso_hom (j : J) : G.map (colimit.ι F j) ≫ (preservesColimitIso G F).hom = colimit.ι (F ⋙ G) j := - (PreservesColimit.preserves (colimit.isColimit _)).comp_coconePointUniqueUpToIso_hom _ j + (isColimitOfPreserves G (colimit.isColimit _)).comp_coconePointUniqueUpToIso_hom _ j @[deprecated (since := "2024-10-27")] alias ι_preservesColimitsIso_hom := ι_preservesColimitIso_hom @@ -177,8 +177,8 @@ variable [HasColimit F] [HasColimit (F ⋙ G)] /-- If the comparison morphism `colimit (F ⋙ G) ⟶ G.obj (colimit F)` is an isomorphism, then `G` preserves colimits of `F`. -/ -def preservesColimitOfIsIsoPost [IsIso (colimit.post F G)] : PreservesColimit F G := - preservesColimitOfPreservesColimitCocone (colimit.isColimit F) (by +lemma preservesColimit_of_isIso_post [IsIso (colimit.post F G)] : PreservesColimit F G := + preservesColimit_of_preserves_colimit_cocone (colimit.isColimit F) (by convert IsColimit.ofPointIso (colimit.isColimit (F ⋙ G)) assumption) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean index 99f59eb1dea7a..3c6706cf053d9 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean @@ -28,112 +28,113 @@ variable {J : Type w} [Category.{w'} J] /-- If `F : C ⥤ D` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesColimit K.leftOp F] : +lemma preservesLimit_op (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesColimit K.leftOp F] : PreservesLimit K F.op where preserves {_} hc := - isLimitConeRightOpOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitConeRightOpOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ D` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitOfOp (K : J ⥤ C) (F : C ⥤ D) [PreservesColimit K.op F.op] : +lemma preservesLimit_of_op (K : J ⥤ C) (F : C ⥤ D) [PreservesColimit K.op F.op] : PreservesLimit K F where - preserves {_} hc := isLimitOfOp (isColimitOfPreserves F.op (IsLimit.op hc)) + preserves {_} hc := ⟨isLimitOfOp (isColimitOfPreserves F.op (IsLimit.op hc))⟩ /-- If `F : C ⥤ Dᵒᵖ` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitLeftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F] : +lemma preservesLimit_leftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F] : PreservesLimit K F.leftOp where preserves {_} hc := - isLimitConeUnopOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitConeUnopOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitOfLeftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.op F.leftOp] : +lemma preservesLimit_of_leftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.op F.leftOp] : PreservesLimit K F where - preserves {_} hc := isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.leftOp (IsLimit.op hc)) + preserves {_} hc := + ⟨isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.leftOp (IsLimit.op hc))⟩ /-- If `F : Cᵒᵖ ⥤ D` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitRightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.op F] : +lemma preservesLimit_rightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.op F] : PreservesLimit K F.rightOp where preserves {_} hc := - isLimitConeRightOpOfCocone _ (isColimitOfPreserves F hc.op) + ⟨isLimitConeRightOpOfCocone _ (isColimitOfPreserves F hc.op)⟩ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : Cᵒᵖ ⥤ D` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitOfRightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.leftOp F.rightOp] : +lemma preservesLimit_of_rightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.leftOp F.rightOp] : PreservesLimit K F where preserves {_} hc := - isLimitOfOp (isColimitOfPreserves F.rightOp (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitOfOp (isColimitOfPreserves F.rightOp (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.unop : C ⥤ D` preserves limits of `K : J ⥤ C`. -/ -def preservesLimitUnop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.op F] : +lemma preservesLimit_unop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.op F] : PreservesLimit K F.unop where preserves {_} hc := - isLimitConeUnopOfCocone _ (isColimitOfPreserves F hc.op) + ⟨isLimitConeUnopOfCocone _ (isColimitOfPreserves F hc.op)⟩ /-- If `F.unop : C ⥤ D` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesLimitOfUnop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F.unop] : +lemma preservesLimit_of_unop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F.unop] : PreservesLimit K F where preserves {_} hc := - isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.unop (isColimitCoconeLeftOpOfCone _ hc)) + ⟨isLimitOfCoconeLeftOpOfCone _ (isColimitOfPreserves F.unop (isColimitCoconeLeftOpOfCone _ hc))⟩ /-- If `F : C ⥤ D` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesLimit K.leftOp F] : +lemma preservesColimit_op (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesLimit K.leftOp F] : PreservesColimit K F.op where preserves {_} hc := - isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc))⟩ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ D` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitOfOp (K : J ⥤ C) (F : C ⥤ D) [PreservesLimit K.op F.op] : +lemma preservesColimit_of_op (K : J ⥤ C) (F : C ⥤ D) [PreservesLimit K.op F.op] : PreservesColimit K F where - preserves {_} hc := isColimitOfOp (isLimitOfPreserves F.op (IsColimit.op hc)) + preserves {_} hc := ⟨isColimitOfOp (isLimitOfPreserves F.op (IsColimit.op hc))⟩ /-- If `F : C ⥤ Dᵒᵖ` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitLeftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F] : +lemma preservesColimit_leftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F] : PreservesColimit K F.leftOp where preserves {_} hc := - isColimitCoconeUnopOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitCoconeUnopOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc))⟩ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitOfLeftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.op F.leftOp] : +lemma preservesColimit_of_leftOp (K : J ⥤ C) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.op F.leftOp] : PreservesColimit K F where preserves {_} hc := - isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.leftOp (IsColimit.op hc)) + ⟨isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.leftOp (IsColimit.op hc))⟩ /-- If `F : Cᵒᵖ ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitRightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.op F] : +lemma preservesColimit_rightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.op F] : PreservesColimit K F.rightOp where preserves {_} hc := - isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F hc.op) + ⟨isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F hc.op)⟩ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F : Cᵒᵖ ⥤ D` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitOfRightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.leftOp F.rightOp] : +lemma preservesColimit_of_rightOp (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.leftOp F.rightOp] : PreservesColimit K F where preserves {_} hc := - isColimitOfOp (isLimitOfPreserves F.rightOp (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitOfOp (isLimitOfPreserves F.rightOp (isLimitConeLeftOpOfCocone _ hc))⟩ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.unop : C ⥤ D` preserves colimits of `K : J ⥤ C`. -/ -def preservesColimitUnop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.op F] : +lemma preservesColimit_unop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.op F] : PreservesColimit K F.unop where preserves {_} hc := - isColimitCoconeUnopOfCone _ (isLimitOfPreserves F hc.op) + ⟨isColimitCoconeUnopOfCone _ (isLimitOfPreserves F hc.op)⟩ /-- If `F.unop : C ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ C`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ -def preservesColimitOfUnop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F.unop] : +lemma preservesColimit_of_unop (K : J ⥤ Cᵒᵖ) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F.unop] : PreservesColimit K F where preserves {_} hc := - isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.unop (isLimitConeLeftOpOfCocone _ hc)) + ⟨isColimitOfConeLeftOpOfCocone _ (isLimitOfPreserves F.unop (isLimitConeLeftOpOfCocone _ hc))⟩ section @@ -141,385 +142,388 @@ variable (J) /-- If `F : C ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOp (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.op where preservesLimit {K} := preservesLimitOp K F +lemma preservesLimitsOfShape_op (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.op where preservesLimit {K} := preservesLimit_op K F /-- If `F : C ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.leftOp where preservesLimit {K} := preservesLimitLeftOp K F +lemma preservesLimitsOfShape_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.leftOp where preservesLimit {K} := preservesLimit_leftOp K F /-- If `F : Cᵒᵖ ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.rightOp where preservesLimit {K} := preservesLimitRightOp K F +lemma preservesLimitsOfShape_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.rightOp where preservesLimit {K} := preservesLimit_rightOp K F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F.unop : C ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : - PreservesLimitsOfShape J F.unop where preservesLimit {K} := preservesLimitUnop K F +lemma preservesLimitsOfShape_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : + PreservesLimitsOfShape J F.unop where preservesLimit {K} := preservesLimit_unop K F /-- If `F : C ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOp (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.op where preservesColimit {K} := preservesColimitOp K F +lemma preservesColimitsOfShape_op (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.op where preservesColimit {K} := preservesColimit_op K F /-- If `F : C ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.leftOp where preservesColimit {K} := preservesColimitLeftOp K F +lemma preservesColimitsOfShape_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.leftOp where preservesColimit {K} := preservesColimit_leftOp K F /-- If `F : Cᵒᵖ ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.rightOp where preservesColimit {K} := preservesColimitRightOp K F +lemma preservesColimitsOfShape_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.rightOp where preservesColimit {K} := preservesColimit_rightOp K F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F.unop : C ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : - PreservesColimitsOfShape J F.unop where preservesColimit {K} := preservesColimitUnop K F +lemma preservesColimitsOfShape_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : + PreservesColimitsOfShape J F.unop where preservesColimit {K} := preservesColimit_unop K F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F : C ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfOp (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.op] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfOp K F +lemma preservesLimitsOfShape_of_op (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.op] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_op K F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.leftOp] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfLeftOp K F +lemma preservesLimitsOfShape_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.leftOp] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_leftOp K F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ D` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.rightOp] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfRightOp K F +lemma preservesLimitsOfShape_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F.rightOp] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_rightOp K F /-- If `F.unop : C ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ -def preservesLimitsOfShapeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.unop] : - PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimitOfUnop K F +lemma preservesLimitsOfShape_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F.unop] : + PreservesLimitsOfShape J F where preservesLimit {K} := preservesLimit_of_unop K F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F : C ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfOp (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.op] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfOp K F +lemma preservesColimitsOfShape_of_op (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.op] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_op K F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F : C ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.leftOp] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfLeftOp K F +lemma preservesColimitsOfShape_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.leftOp] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_leftOp K F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ D` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.rightOp] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfRightOp K F +lemma preservesColimitsOfShape_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F.rightOp] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_rightOp K F /-- If `F.unop : C ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ -def preservesColimitsOfShapeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.unop] : - PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimitOfUnop K F +lemma preservesColimitsOfShape_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F.unop] : + PreservesColimitsOfShape J F where preservesColimit {K} := preservesColimit_of_unop K F end /-- If `F : C ⥤ D` preserves colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeOp (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_op (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.op where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.leftOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeLeftOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.rightOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeRightOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F.unop : C ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : +lemma preservesLimitsOfSize_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F] : PreservesLimitsOfSize.{w, w'} F.unop where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeUnop _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_unop _ _ /-- If `F : C ⥤ D` preserves limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeOp (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_op (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.op where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.leftOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeLeftOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.rightOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeRightOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F.unop : C ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : +lemma preservesColimitsOfSize_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F] : PreservesColimitsOfSize.{w, w'} F.unop where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeUnop _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F : C ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeOfOp (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F.op] : +lemma preservesLimitsOfSize_of_op (F : C ⥤ D) [PreservesColimitsOfSize.{w, w'} F.op] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits, then `F : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.leftOp] : +lemma preservesLimitsOfSize_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.leftOp] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfLeftOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits, then `F : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsOfSizeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F.rightOp] : +lemma preservesLimitsOfSize_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfSize.{w, w'} F.rightOp] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfRightOp _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves colimits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfSizeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.unop] : +lemma preservesLimitsOfSize_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSize.{w, w'} F.unop] : PreservesLimitsOfSize.{w, w'} F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfUnop _ _ + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F : C ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeOfOp (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F.op] : +lemma preservesColimitsOfSize_of_op (F : C ⥤ D) [PreservesLimitsOfSize.{w, w'} F.op] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits, then `F : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.leftOp] : +lemma preservesColimitsOfSize_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.leftOp] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfLeftOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits, then `F : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsOfSizeOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F.rightOp] : +lemma preservesColimitsOfSize_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfSize.{w, w'} F.rightOp] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfRightOp _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves limits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfSizeOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.unop] : +lemma preservesColimitsOfSize_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimitsOfSize.{w, w'} F.unop] : PreservesColimitsOfSize.{w, w'} F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfUnop _ _ + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_unop _ _ /-- If `F : C ⥤ D` preserves colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOp (F : C ⥤ D) [PreservesColimits F] : PreservesLimits F.op where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOp _ _ +lemma preservesLimits_op (F : C ⥤ D) [PreservesColimits F] : PreservesLimits F.op where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.leftOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeLeftOp _ _ +lemma preservesLimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.leftOp where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F] : PreservesLimits F.rightOp where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeRightOp _ _ +lemma preservesLimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F] : PreservesLimits F.rightOp where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F.unop : C ⥤ D` preserves limits. -/ -def preservesLimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.unop where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeUnop _ _ +lemma preservesLimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.unop where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_unop _ _ /-- If `F : C ⥤ D` preserves limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def perservesColimitsOp (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOp _ _ +lemma perservesColimits_op (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_op _ _ /-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.leftOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeLeftOp _ _ +lemma preservesColimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.leftOp where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_leftOp _ _ /-- If `F : Cᵒᵖ ⥤ D` preserves limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F] : PreservesColimits F.rightOp where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeRightOp _ _ +lemma preservesColimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F] : + PreservesColimits F.rightOp where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_rightOp _ _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F.unop : C ⥤ D` preserves colimits. -/ -def preservesColimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.unop where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeUnop _ _ +lemma preservesColimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.unop where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits, then `F : C ⥤ D` preserves limits. -/ -def preservesLimitsOfOp (F : C ⥤ D) [PreservesColimits F.op] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfOp _ _ +lemma preservesLimits_of_op (F : C ⥤ D) [PreservesColimits F.op] : PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits, then `F : C ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F.leftOp] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfLeftOp _ _ +lemma preservesLimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F.leftOp] : PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits, then `F : Cᵒᵖ ⥤ D` preserves limits. -/ -def preservesLimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F.rightOp] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfRightOp _ _ +lemma preservesLimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F.rightOp] : + PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves colimits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits. -/ -def preservesLimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F.unop] : PreservesLimits F where - preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOfUnop _ _ +lemma preservesLimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F.unop] : PreservesLimits F where + preservesLimitsOfShape {_} _ := preservesLimitsOfShape_of_unop _ _ /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits, then `F : C ⥤ D` preserves colimits. -/ -def preservesColimitsOfOp (F : C ⥤ D) [PreservesLimits F.op] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfOp _ _ +lemma preservesColimits_of_op (F : C ⥤ D) [PreservesLimits F.op] : PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_op _ _ /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves limits, then `F : C ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F.leftOp] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfLeftOp _ _ +lemma preservesColimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F.leftOp] : + PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_leftOp _ _ /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves limits, then `F : Cᵒᵖ ⥤ D` preserves colimits. -/ -def preservesColimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F.rightOp] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfRightOp _ _ +lemma preservesColimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F.rightOp] : + PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_rightOp _ _ /-- If `F.unop : C ⥤ D` preserves limits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits. -/ -def preservesColimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F.unop] : PreservesColimits F where - preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOfUnop _ _ +lemma preservesColimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F.unop] : PreservesColimits F where + preservesColimitsOfShape {_} _ := preservesColimitsOfShape_of_unop _ _ /-- If `F : C ⥤ D` preserves finite colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsOp (F : C ⥤ D) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_op (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteLimits F.op where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_op J F /-- If `F : C ⥤ Dᵒᵖ` preserves finite colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : PreservesFiniteLimits F.leftOp where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeLeftOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_leftOp J F /-- If `F : Cᵒᵖ ⥤ D` preserves finite colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteLimits F.rightOp where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeRightOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_rightOp J F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits, then `F.unop : C ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : +lemma preservesFiniteLimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : PreservesFiniteLimits F.unop where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeUnop J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_unop J F /-- If `F : C ⥤ D` preserves finite limits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsOp (F : C ⥤ D) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_op (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteColimits F.op where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_op J F /-- If `F : C ⥤ Dᵒᵖ` preserves finite limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : PreservesFiniteColimits F.leftOp where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeLeftOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_leftOp J F /-- If `F : Cᵒᵖ ⥤ D` preserves finite limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteColimits F.rightOp where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeRightOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_rightOp J F /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits, then `F.unop : C ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : +lemma preservesFiniteColimits_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : PreservesFiniteColimits F.unop where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeUnop J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_unop J F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits, then `F : C ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsOfOp (F : C ⥤ D) [PreservesFiniteColimits F.op] : +lemma preservesFiniteLimits_of_op (F : C ⥤ D) [PreservesFiniteColimits F.op] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_op J F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves finite colimits, then `F : C ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F.leftOp] : +lemma preservesFiniteLimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F.leftOp] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfLeftOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_leftOp J F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves finite colimits, then `F : Cᵒᵖ ⥤ D` preserves finite limits. -/ -def preservesFiniteLimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F.rightOp] : +lemma preservesFiniteLimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F.rightOp] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfRightOp J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_rightOp J F /-- If `F.unop : C ⥤ D` preserves finite colimits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits. -/ -def preservesFiniteLimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F.unop] : +lemma preservesFiniteLimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteColimits F.unop] : PreservesFiniteLimits F where - preservesFiniteLimits J _ _ := preservesLimitsOfShapeOfUnop J F + preservesFiniteLimits J _ _ := preservesLimitsOfShape_of_unop J F /-- If `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits, then `F : C ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsOfOp (F : C ⥤ D) [PreservesFiniteLimits F.op] : +lemma preservesFiniteColimits_of_op (F : C ⥤ D) [PreservesFiniteLimits F.op] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_op J F /-- If `F.leftOp : Cᵒᵖ ⥤ D` preserves finite limits, then `F : C ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsOfLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F.leftOp] : +lemma preservesFiniteColimits_of_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F.leftOp] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfLeftOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_leftOp J F /-- If `F.rightOp : C ⥤ Dᵒᵖ` preserves finite limits, then `F : Cᵒᵖ ⥤ D` preserves finite colimits. -/ -def preservesFiniteColimitsOfRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F.rightOp] : +lemma preservesFiniteColimits_of_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F.rightOp] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfRightOp J F + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_rightOp J F /-- If `F.unop : C ⥤ D` preserves finite limits, then `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite colimits. -/ -def preservesFiniteColimitsOfUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F.unop] : +lemma preservesFiniteColimits_of_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits F.unop] : PreservesFiniteColimits F where - preservesFiniteColimits J _ _ := preservesColimitsOfShapeOfUnop J F - + preservesFiniteColimits J _ _ := preservesColimitsOfShape_of_unop J F /-- If `F : C ⥤ D` preserves finite coproducts, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite products. -/ -def preservesFiniteProductsOp (F : C ⥤ D) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_op (F : C ⥤ D) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.op where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_op + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : C ⥤ Dᵒᵖ` preserves finite coproducts, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite products. -/ -def preservesFiniteProductsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.leftOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeLeftOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_leftOp + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ D` preserves finite coproducts, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite products. -/ -def preservesFiniteProductsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.rightOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeRightOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_rightOp + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite coproducts, then `F.unop : C ⥤ D` preserves finite products. -/ -def preservesFiniteProductsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : +lemma preservesFiniteProducts_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : PreservesFiniteProducts F.unop where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeUnop - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesLimitsOfShape_unop + exact preservesColimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : C ⥤ D` preserves finite products, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite coproducts. -/ -def preservesFiniteCoproductsOp (F : C ⥤ D) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_op (F : C ⥤ D) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.op where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeOp - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_op + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : C ⥤ Dᵒᵖ` preserves finite products, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite coproducts. -/ -def preservesFiniteCoproductsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_leftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.leftOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeLeftOp - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_leftOp + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ D` preserves finite products, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite coproducts. -/ -def preservesFiniteCoproductsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_rightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.rightOp where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeRightOp - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_rightOp + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ /-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite products, then `F.unop : C ⥤ D` preserves finite coproducts. -/ -def preservesFiniteCoproductsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : +lemma preservesFiniteCoproducts_unop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : PreservesFiniteCoproducts F.unop where preserves J _ := by - apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeUnop - exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + apply (config := { allowSynthFailures := true }) preservesColimitsOfShape_unop + exact preservesLimitsOfShape_of_equiv (Discrete.opposite J).symm _ end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean index f2a3c01bd9fc0..7750b104f48ec 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Presheaf.lean @@ -172,9 +172,9 @@ attribute [local instance] PreservesFiniteLimitsOfIsFilteredCostructuredArrowYon One direction of Proposition 3.3.13 of [Kashiwara2006]. -/ -noncomputable def preservesFiniteLimitsOfIsFilteredCostructuredArrowYoneda +lemma preservesFiniteLimits_of_isFiltered_costructuredArrow_yoneda [IsFiltered (CostructuredArrow yoneda A)] : PreservesFiniteLimits A where - preservesFiniteLimits _ _ _ := ⟨fun {_} => preservesLimitOfIsIsoPost _ _⟩ + preservesFiniteLimits _ _ _ := ⟨fun {_} => preservesLimit_of_isIso_post _ _⟩ /-- If `C` is a small finitely cocomplete category and `A : Cᵒᵖ ⥤ Type u` is a presheaf, then `CostructuredArrow yoneda A` is filtered if and only if `A` preserves finite limits. @@ -183,7 +183,7 @@ Proposition 3.3.13 of [Kashiwara2006]. -/ theorem isFiltered_costructuredArrow_yoneda_iff_nonempty_preservesFiniteLimits : IsFiltered (CostructuredArrow yoneda A) ↔ Nonempty (PreservesFiniteLimits A) := - ⟨fun _ => ⟨preservesFiniteLimitsOfIsFilteredCostructuredArrowYoneda A⟩, + ⟨fun _ => ⟨preservesFiniteLimits_of_isFiltered_costructuredArrow_yoneda A⟩, fun ⟨_⟩ => isFiltered_costructuredArrow_yoneda_of_preservesFiniteLimits A⟩ end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean index b72dd4a247aae..a9c9bfb413d2c 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/BinaryProducts.lean @@ -47,12 +47,12 @@ def isLimitMapConeBinaryFanEquiv : /-- The property of preserving products expressed in terms of binary fans. -/ def mapIsLimitOfPreservesOfIsLimit [PreservesLimit (pair X Y) G] (l : IsLimit (BinaryFan.mk f g)) : IsLimit (BinaryFan.mk (G.map f) (G.map g)) := - isLimitMapConeBinaryFanEquiv G f g (PreservesLimit.preserves l) + isLimitMapConeBinaryFanEquiv G f g (isLimitOfPreserves G l) /-- The property of reflecting products expressed in terms of binary fans. -/ def isLimitOfReflectsOfMapIsLimit [ReflectsLimit (pair X Y) G] (l : IsLimit (BinaryFan.mk (G.map f) (G.map g))) : IsLimit (BinaryFan.mk f g) := - ReflectsLimit.reflects ((isLimitMapConeBinaryFanEquiv G f g).symm l) + isLimitOfReflects G ((isLimitMapConeBinaryFanEquiv G f g).symm l) variable (X Y) variable [HasBinaryProduct X Y] @@ -69,9 +69,9 @@ variable [HasBinaryProduct (G.obj X) (G.obj Y)] /-- If the product comparison map for `G` at `(X,Y)` is an isomorphism, then `G` preserves the pair of `(X,Y)`. -/ -def PreservesLimitPair.ofIsoProdComparison [i : IsIso (prodComparison G X Y)] : +lemma PreservesLimitPair.of_iso_prod_comparison [i : IsIso (prodComparison G X Y)] : PreservesLimit (pair X Y) G := by - apply preservesLimitOfPreservesLimitCone (prodIsProd X Y) + apply preservesLimit_of_preserves_limit_cone (prodIsProd X Y) apply (isLimitMapConeBinaryFanEquiv _ _ _).symm _ refine @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (pair (G.obj X) (G.obj Y))) ?_ apply i @@ -125,12 +125,12 @@ def isColimitMapCoconeBinaryCofanEquiv : /-- The property of preserving coproducts expressed in terms of binary cofans. -/ def mapIsColimitOfPreservesOfIsColimit [PreservesColimit (pair X Y) G] (l : IsColimit (BinaryCofan.mk f g)) : IsColimit (BinaryCofan.mk (G.map f) (G.map g)) := - isColimitMapCoconeBinaryCofanEquiv G f g (PreservesColimit.preserves l) + isColimitMapCoconeBinaryCofanEquiv G f g (isColimitOfPreserves G l) /-- The property of reflecting coproducts expressed in terms of binary cofans. -/ def isColimitOfReflectsOfMapIsColimit [ReflectsColimit (pair X Y) G] (l : IsColimit (BinaryCofan.mk (G.map f) (G.map g))) : IsColimit (BinaryCofan.mk f g) := - ReflectsColimit.reflects ((isColimitMapCoconeBinaryCofanEquiv G f g).symm l) + isColimitOfReflects G ((isColimitMapCoconeBinaryCofanEquiv G f g).symm l) variable (X Y) variable [HasBinaryCoproduct X Y] @@ -148,9 +148,9 @@ variable [HasBinaryCoproduct (G.obj X) (G.obj Y)] /-- If the coproduct comparison map for `G` at `(X,Y)` is an isomorphism, then `G` preserves the pair of `(X,Y)`. -/ -def PreservesColimitPair.ofIsoCoprodComparison [i : IsIso (coprodComparison G X Y)] : +lemma PreservesColimitPair.of_iso_coprod_comparison [i : IsIso (coprodComparison G X Y)] : PreservesColimit (pair X Y) G := by - apply preservesColimitOfPreservesColimitCocone (coprodIsCoprod X Y) + apply preservesColimit_of_preserves_colimit_cocone (coprodIsCoprod X Y) apply (isColimitMapCoconeBinaryCofanEquiv _ _ _).symm _ refine @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (pair (G.obj X) (G.obj Y))) ?_ apply i diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean index c3d4d473c6c16..550f8c8bea18d 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean @@ -85,8 +85,8 @@ variable {J : Type w₁} {K : Type w₂} /-- A functor `F` preserves biproducts of `f` if `F` maps every bilimit bicone over `f` to a bilimit bicone over `F.obj ∘ f`. -/ -class PreservesBiproduct (f : J → C) (F : C ⥤ D) [PreservesZeroMorphisms F] where - preserves : ∀ {b : Bicone f}, b.IsBilimit → (F.mapBicone b).IsBilimit +class PreservesBiproduct (f : J → C) (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where + preserves : ∀ {b : Bicone f}, b.IsBilimit → Nonempty (F.mapBicone b).IsBilimit attribute [inherit_doc PreservesBiproduct] PreservesBiproduct.preserves @@ -94,13 +94,13 @@ attribute [inherit_doc PreservesBiproduct] PreservesBiproduct.preserves bilimit bicone over `F.obj ∘ f`. -/ def isBilimitOfPreserves {f : J → C} (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproduct f F] {b : Bicone f} (hb : b.IsBilimit) : (F.mapBicone b).IsBilimit := - PreservesBiproduct.preserves hb + (PreservesBiproduct.preserves hb).some variable (J) /-- A functor `F` preserves biproducts of shape `J` if it preserves biproducts of `f` for every `f : J → C`. -/ -class PreservesBiproductsOfShape (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesBiproductsOfShape (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {f : J → C}, PreservesBiproduct f F attribute [inherit_doc PreservesBiproductsOfShape] PreservesBiproductsOfShape.preserves @@ -111,7 +111,7 @@ end Bicone /-- A functor `F` preserves finite biproducts if it preserves biproducts of shape `J` whenever `J` is a fintype. -/ -class PreservesFiniteBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesFiniteBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {J : Type} [Fintype J], PreservesBiproductsOfShape J F attribute [inherit_doc PreservesFiniteBiproducts] PreservesFiniteBiproducts.preserves @@ -121,7 +121,7 @@ attribute [instance 100] PreservesFiniteBiproducts.preserves /-- A functor `F` preserves biproducts if it preserves biproducts of any shape `J` of size `w`. The usual notion of preservation of biproducts is recovered by choosing `w` to be the universe of the morphisms of `C`. -/ -class PreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {J : Type w₁}, PreservesBiproductsOfShape J F attribute [inherit_doc PreservesBiproducts] PreservesBiproducts.preserves @@ -130,22 +130,22 @@ attribute [instance 100] PreservesBiproducts.preserves /-- Preserving biproducts at a bigger universe level implies preserving biproducts at a smaller universe level. -/ -def preservesBiproductsShrink (F : C ⥤ D) [PreservesZeroMorphisms F] +lemma preservesBiproducts_shrink (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproducts.{max w₁ w₂} F] : PreservesBiproducts.{w₁} F := ⟨fun {_} => ⟨fun {_} => ⟨fun {b} ib => - ((F.mapBicone b).whiskerIsBilimitIff _).toFun - (isBilimitOfPreserves F ((b.whiskerIsBilimitIff Equiv.ulift.{w₂}).invFun ib))⟩⟩⟩ + ⟨((F.mapBicone b).whiskerIsBilimitIff _).toFun + (isBilimitOfPreserves F ((b.whiskerIsBilimitIff Equiv.ulift.{w₂}).invFun ib))⟩⟩⟩⟩ instance (priority := 100) preservesFiniteBiproductsOfPreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproducts.{w₁} F] : PreservesFiniteBiproducts F where - preserves {J} _ := by letI := preservesBiproductsShrink.{0} F; infer_instance + preserves {J} _ := by letI := preservesBiproducts_shrink.{0} F; infer_instance /-- A functor `F` preserves binary biproducts of `X` and `Y` if `F` maps every bilimit bicone over `X` and `Y` to a bilimit bicone over `F.obj X` and `F.obj Y`. -/ -class PreservesBinaryBiproduct (X Y : C) (F : C ⥤ D) [PreservesZeroMorphisms F] where - preserves : ∀ {b : BinaryBicone X Y}, b.IsBilimit → (F.mapBinaryBicone b).IsBilimit +class PreservesBinaryBiproduct (X Y : C) (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where + preserves : ∀ {b : BinaryBicone X Y}, b.IsBilimit → Nonempty ((F.mapBinaryBicone b).IsBilimit) attribute [inherit_doc PreservesBinaryBiproduct] PreservesBinaryBiproduct.preserves @@ -154,20 +154,21 @@ attribute [inherit_doc PreservesBinaryBiproduct] PreservesBinaryBiproduct.preser def isBinaryBilimitOfPreserves {X Y : C} (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBinaryBiproduct X Y F] {b : BinaryBicone X Y} (hb : b.IsBilimit) : (F.mapBinaryBicone b).IsBilimit := - PreservesBinaryBiproduct.preserves hb + (PreservesBinaryBiproduct.preserves hb).some /-- A functor `F` preserves binary biproducts if it preserves the binary biproduct of `X` and `Y` for all `X` and `Y`. -/ -class PreservesBinaryBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] where +class PreservesBinaryBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] : Prop where preserves : ∀ {X Y : C}, PreservesBinaryBiproduct X Y F := by infer_instance attribute [inherit_doc PreservesBinaryBiproducts] PreservesBinaryBiproducts.preserves /-- A functor that preserves biproducts of a pair preserves binary biproducts. -/ -def preservesBinaryBiproductOfPreservesBiproduct (F : C ⥤ D) [PreservesZeroMorphisms F] (X Y : C) - [PreservesBiproduct (pairFunction X Y) F] : PreservesBinaryBiproduct X Y F where - preserves {b} hb := - { isLimit := +lemma preservesBinaryBiproduct_of_preservesBiproduct (F : C ⥤ D) + [PreservesZeroMorphisms F] (X Y : C) [PreservesBiproduct (pairFunction X Y) F] : + PreservesBinaryBiproduct X Y F where + preserves {b} hb := ⟨{ + isLimit := IsLimit.ofIsoLimit ((IsLimit.postcomposeHomEquiv (diagramIsoPair _) _).symm (isBilimitOfPreserves F (b.toBiconeIsBilimit.symm hb)).isLimit) <| @@ -178,12 +179,12 @@ def preservesBinaryBiproductOfPreservesBiproduct (F : C ⥤ D) [PreservesZeroMor ((IsColimit.precomposeInvEquiv (diagramIsoPair _) _).symm (isBilimitOfPreserves F (b.toBiconeIsBilimit.symm hb)).isColimit) <| Cocones.ext (Iso.refl _) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp } + rcases j with ⟨⟨⟩⟩ <;> simp }⟩ /-- A functor that preserves biproducts of a pair preserves binary biproducts. -/ -def preservesBinaryBiproductsOfPreservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] +lemma preservesBinaryBiproducts_of_preservesBiproducts (F : C ⥤ D) [PreservesZeroMorphisms F] [PreservesBiproductsOfShape WalkingPair F] : PreservesBinaryBiproducts F where - preserves {X} Y := preservesBinaryBiproductOfPreservesBiproduct F X Y + preserves {X} Y := preservesBinaryBiproduct_of_preservesBiproduct F X Y attribute [instance 100] PreservesBinaryBiproducts.preserves @@ -257,13 +258,13 @@ variable [PreservesZeroMorphisms F] [PreservesBiproduct f F] instance hasBiproduct_of_preserves : HasBiproduct (F.obj ∘ f) := HasBiproduct.mk { bicone := F.mapBicone (biproduct.bicone f) - isBilimit := PreservesBiproduct.preserves (biproduct.isBilimit _) } + isBilimit := isBilimitOfPreserves _ (biproduct.isBilimit _) } /-- If `F` preserves a biproduct, we get a definitionally nice isomorphism `F.obj (⨁ f) ≅ ⨁ (F.obj ∘ f)`. -/ @[simp] def mapBiproduct : F.obj (⨁ f) ≅ ⨁ F.obj ∘ f := - biproduct.uniqueUpToIso _ (PreservesBiproduct.preserves (biproduct.isBilimit _)) + biproduct.uniqueUpToIso _ (isBilimitOfPreserves _ (biproduct.isBilimit _)) theorem mapBiproduct_hom : haveI : HasBiproduct fun j => F.obj (f j) := hasBiproduct_of_preserves F f @@ -341,13 +342,13 @@ variable [PreservesZeroMorphisms F] [PreservesBinaryBiproduct X Y F] instance hasBinaryBiproduct_of_preserves : HasBinaryBiproduct (F.obj X) (F.obj Y) := HasBinaryBiproduct.mk { bicone := F.mapBinaryBicone (BinaryBiproduct.bicone X Y) - isBilimit := PreservesBinaryBiproduct.preserves (BinaryBiproduct.isBilimit _ _) } + isBilimit := isBinaryBilimitOfPreserves F (BinaryBiproduct.isBilimit _ _) } /-- If `F` preserves a binary biproduct, we get a definitionally nice isomorphism `F.obj (X ⊞ Y) ≅ F.obj X ⊞ F.obj Y`. -/ @[simp] def mapBiprod : F.obj (X ⊞ Y) ≅ F.obj X ⊞ F.obj Y := - biprod.uniqueUpToIso _ _ (PreservesBinaryBiproduct.preserves (BinaryBiproduct.isBilimit _ _)) + biprod.uniqueUpToIso _ _ (isBinaryBilimitOfPreserves F (BinaryBiproduct.isBilimit _ _)) theorem mapBiprod_hom : (mapBiprod F X Y).hom = biprod.lift (F.map biprod.fst) (F.map biprod.snd) := rfl diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean index 0eece82e3b677..3e3b781aeb0d8 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Equalizers.lean @@ -46,13 +46,13 @@ def isLimitMapConeForkEquiv : /-- The property of preserving equalizers expressed in terms of forks. -/ def isLimitForkMapOfIsLimit [PreservesLimit (parallelPair f g) G] (l : IsLimit (Fork.ofι h w)) : IsLimit (Fork.ofι (G.map h) (by simp only [← G.map_comp, w]) : Fork (G.map f) (G.map g)) := - isLimitMapConeForkEquiv G w (PreservesLimit.preserves l) + isLimitMapConeForkEquiv G w (isLimitOfPreserves G l) /-- The property of reflecting equalizers expressed in terms of forks. -/ def isLimitOfIsLimitForkMap [ReflectsLimit (parallelPair f g) G] (l : IsLimit (Fork.ofι (G.map h) (by simp only [← G.map_comp, w]) : Fork (G.map f) (G.map g))) : IsLimit (Fork.ofι h w) := - ReflectsLimit.reflects ((isLimitMapConeForkEquiv G w).symm l) + isLimitOfReflects G ((isLimitMapConeForkEquiv G w).symm l) variable (f g) variable [HasEqualizer f g] @@ -72,12 +72,11 @@ variable [HasEqualizer (G.map f) (G.map g)] /-- If the equalizer comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the equalizer of `(f,g)`. -/ -def PreservesEqualizer.ofIsoComparison [i : IsIso (equalizerComparison f g G)] : +lemma PreservesEqualizer.of_iso_comparison [i : IsIso (equalizerComparison f g G)] : PreservesLimit (parallelPair f g) G := by - apply preservesLimitOfPreservesLimitCone (equalizerIsEqualizer f g) + apply preservesLimit_of_preserves_limit_cone (equalizerIsEqualizer f g) apply (isLimitMapConeForkEquiv _ _).symm _ - refine @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (parallelPair (G.map f) (G.map g))) ?_ - apply i + exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (parallelPair (G.map f) (G.map g))) i variable [PreservesLimit (parallelPair f g) G] @@ -128,7 +127,7 @@ def isColimitCoforkMapOfIsColimit [PreservesColimit (parallelPair f g) G] (l : IsColimit (Cofork.ofπ h w)) : IsColimit (Cofork.ofπ (G.map h) (by simp only [← G.map_comp, w]) : Cofork (G.map f) (G.map g)) := - isColimitMapCoconeCoforkEquiv G w (PreservesColimit.preserves l) + isColimitMapCoconeCoforkEquiv G w (isColimitOfPreserves G l) /-- The property of reflecting coequalizers expressed in terms of coforks. -/ def isColimitOfIsColimitCoforkMap [ReflectsColimit (parallelPair f g) G] @@ -136,7 +135,7 @@ def isColimitOfIsColimitCoforkMap [ReflectsColimit (parallelPair f g) G] IsColimit (Cofork.ofπ (G.map h) (by simp only [← G.map_comp, w]) : Cofork (G.map f) (G.map g))) : IsColimit (Cofork.ofπ h w) := - ReflectsColimit.reflects ((isColimitMapCoconeCoforkEquiv G w).symm l) + isColimitOfReflects G ((isColimitMapCoconeCoforkEquiv G w).symm l) variable (f g) variable [HasCoequalizer f g] @@ -155,13 +154,12 @@ variable [HasCoequalizer (G.map f) (G.map g)] /-- If the coequalizer comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the coequalizer of `(f,g)`. -/ -def ofIsoComparison [i : IsIso (coequalizerComparison f g G)] : +lemma of_iso_comparison [i : IsIso (coequalizerComparison f g G)] : PreservesColimit (parallelPair f g) G := by - apply preservesColimitOfPreservesColimitCocone (coequalizerIsCoequalizer f g) + apply preservesColimit_of_preserves_colimit_cocone (coequalizerIsCoequalizer f g) apply (isColimitMapCoconeCoforkEquiv _ _).symm _ - refine - @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (parallelPair (G.map f) (G.map g))) ?_ - apply i + exact + @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (parallelPair (G.map f) (G.map g))) i variable [PreservesColimit (parallelPair f g) G] @@ -227,7 +225,7 @@ theorem map_π_preserves_coequalizer_inv_colimMap_desc {X' Y' : D} (f' g' : X' instance (priority := 1) preservesSplitCoequalizers (f g : X ⟶ Y) [HasSplitCoequalizer f g] : PreservesColimit (parallelPair f g) G := by apply - preservesColimitOfPreservesColimitCocone + preservesColimit_of_preserves_colimit_cocone (HasSplitCoequalizer.isSplitCoequalizer f g).isCoequalizer apply (isColimitMapCoconeCoforkEquiv G _).symm @@ -236,7 +234,7 @@ instance (priority := 1) preservesSplitCoequalizers (f g : X ⟶ Y) [HasSplitCoe instance (priority := 1) preservesSplitEqualizers (f g : X ⟶ Y) [HasSplitEqualizer f g] : PreservesLimit (parallelPair f g) G := by apply - preservesLimitOfPreservesLimitCone + preservesLimit_of_preserves_limit_cone (HasSplitEqualizer.isSplitEqualizer f g).isEqualizer apply (isLimitMapConeForkEquiv G _).symm diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean index b6333a08c0d86..034c5e96b668a 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Kernels.lean @@ -91,7 +91,7 @@ def isLimitForkMapOfIsLimit' [PreservesLimit (parallelPair f 0) G] IsLimit (KernelFork.ofι (G.map h) (by simp only [← G.map_comp, w, Functor.map_zero]) : Fork (G.map f) 0) := - isLimitMapConeForkEquiv' G w (PreservesLimit.preserves l) + isLimitMapConeForkEquiv' G w (isLimitOfPreserves G l) variable (f) variable [HasKernel f] @@ -114,9 +114,9 @@ variable [HasKernel (G.map f)] /-- If the kernel comparison map for `G` at `f` is an isomorphism, then `G` preserves the kernel of `f`. -/ -def PreservesKernel.ofIsoComparison [i : IsIso (kernelComparison f G)] : +lemma PreservesKernel.of_iso_comparison [i : IsIso (kernelComparison f G)] : PreservesLimit (parallelPair f 0) G := by - apply preservesLimitOfPreservesLimitCone (kernelIsKernel f) + apply preservesLimit_of_preserves_limit_cone (kernelIsKernel f) apply (isLimitMapConeForkEquiv' G (kernel.condition f)).symm _ exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (kernelIsKernel (G.map f)) i @@ -212,7 +212,7 @@ def isColimitCoforkMapOfIsColimit' [PreservesColimit (parallelPair f 0) G] IsColimit (CokernelCofork.ofπ (G.map h) (by simp only [← G.map_comp, w, Functor.map_zero]) : Cofork (G.map f) 0) := - isColimitMapCoconeCoforkEquiv' G w (PreservesColimit.preserves l) + isColimitMapCoconeCoforkEquiv' G w (isColimitOfPreserves G l) variable (f) variable [HasCokernel f] @@ -236,9 +236,9 @@ variable [HasCokernel (G.map f)] /-- If the cokernel comparison map for `G` at `f` is an isomorphism, then `G` preserves the cokernel of `f`. -/ -def PreservesCokernel.ofIsoComparison [i : IsIso (cokernelComparison f G)] : +lemma PreservesCokernel.of_iso_comparison [i : IsIso (cokernelComparison f G)] : PreservesColimit (parallelPair f 0) G := by - apply preservesColimitOfPreservesColimitCocone (cokernelIsCokernel f) + apply preservesColimit_of_preserves_colimit_cocone (cokernelIsCokernel f) apply (isColimitMapCoconeCoforkEquiv' G (cokernel.condition f)).symm _ exact @IsColimit.ofPointIso _ _ _ _ _ _ _ (cokernelIsCokernel (G.map f)) i @@ -275,32 +275,32 @@ end Cokernels variable (X Y : C) (G : C ⥤ D) [Functor.PreservesZeroMorphisms G] -noncomputable instance preservesKernelZero : +instance preservesKernel_zero : PreservesLimit (parallelPair (0 : X ⟶ Y) 0) G where - preserves {c} hc := by + preserves {c} hc := ⟨by have := KernelFork.IsLimit.isIso_ι c hc rfl refine (KernelFork.isLimitMapConeEquiv c G).symm ?_ refine IsLimit.ofIsoLimit (KernelFork.IsLimit.ofId _ (G.map_zero _ _)) ?_ - exact (Fork.ext (G.mapIso (asIso (Fork.ι c))).symm (by simp)) + exact (Fork.ext (G.mapIso (asIso (Fork.ι c))).symm (by simp))⟩ -noncomputable instance preservesCokernelZero : +noncomputable instance preservesCokernel_zero : PreservesColimit (parallelPair (0 : X ⟶ Y) 0) G where - preserves {c} hc := by + preserves {c} hc := ⟨by have := CokernelCofork.IsColimit.isIso_π c hc rfl refine (CokernelCofork.isColimitMapCoconeEquiv c G).symm ?_ refine IsColimit.ofIsoColimit (CokernelCofork.IsColimit.ofId _ (G.map_zero _ _)) ?_ - exact (Cofork.ext (G.mapIso (asIso (Cofork.π c))) (by simp)) + exact (Cofork.ext (G.mapIso (asIso (Cofork.π c))) (by simp))⟩ variable {X Y} /-- The kernel of a zero map is preserved by any functor which preserves zero morphisms. -/ -noncomputable def preservesKernelZero' (f : X ⟶ Y) (hf : f = 0) : +lemma preservesKernel_zero' (f : X ⟶ Y) (hf : f = 0) : PreservesLimit (parallelPair f 0) G := by rw [hf] infer_instance /-- The cokernel of a zero map is preserved by any functor which preserves zero morphisms. -/ -noncomputable def preservesCokernelZero' (f : X ⟶ Y) (hf : f = 0) : +lemma preservesCokernel_zero' (f : X ⟶ Y) (hf : f = 0) : PreservesColimit (parallelPair f 0) G := by rw [hf] infer_instance diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean index 9d1105f497bd6..6f15d54ba6cac 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean @@ -46,13 +46,13 @@ def isLimitMapConeFanMkEquiv {P : C} (g : ∀ j, P ⟶ f j) : def isLimitFanMkObjOfIsLimit [PreservesLimit (Discrete.functor f) G] {P : C} (g : ∀ j, P ⟶ f j) (t : IsLimit (Fan.mk _ g)) : IsLimit (Fan.mk (G.obj P) fun j => G.map (g j) : Fan fun j => G.obj (f j)) := - isLimitMapConeFanMkEquiv _ _ _ (PreservesLimit.preserves t) + isLimitMapConeFanMkEquiv _ _ _ (isLimitOfPreserves G t) /-- The property of reflecting products expressed in terms of fans. -/ def isLimitOfIsLimitFanMkObj [ReflectsLimit (Discrete.functor f) G] {P : C} (g : ∀ j, P ⟶ f j) (t : IsLimit (Fan.mk _ fun j => G.map (g j) : Fan fun j => G.obj (f j))) : IsLimit (Fan.mk P g) := - ReflectsLimit.reflects ((isLimitMapConeFanMkEquiv _ _ _).symm t) + isLimitOfReflects G ((isLimitMapConeFanMkEquiv _ _ _).symm t) section @@ -69,13 +69,12 @@ def isLimitOfHasProductOfPreservesLimit [PreservesLimit (Discrete.functor f) G] variable [HasProduct fun j : J => G.obj (f j)] /-- If `pi_comparison G f` is an isomorphism, then `G` preserves the limit of `f`. -/ -def PreservesProduct.ofIsoComparison [i : IsIso (piComparison G f)] : +lemma PreservesProduct.of_iso_comparison [i : IsIso (piComparison G f)] : PreservesLimit (Discrete.functor f) G := by - apply preservesLimitOfPreservesLimitCone (productIsProduct f) + apply preservesLimit_of_preserves_limit_cone (productIsProduct f) apply (isLimitMapConeFanMkEquiv _ _ _).symm _ - refine @IsLimit.ofPointIso _ _ _ _ _ _ _ - (limit.isLimit (Discrete.functor fun j : J => G.obj (f j))) ?_ - apply i + exact @IsLimit.ofPointIso _ _ _ _ _ _ _ + (limit.isLimit (Discrete.functor fun j : J => G.obj (f j))) i variable [PreservesLimit (Discrete.functor f) G] @@ -110,14 +109,14 @@ def isColimitMapCoconeCofanMkEquiv {P : C} (g : ∀ j, f j ⟶ P) : def isColimitCofanMkObjOfIsColimit [PreservesColimit (Discrete.functor f) G] {P : C} (g : ∀ j, f j ⟶ P) (t : IsColimit (Cofan.mk _ g)) : IsColimit (Cofan.mk (G.obj P) fun j => G.map (g j) : Cofan fun j => G.obj (f j)) := - isColimitMapCoconeCofanMkEquiv _ _ _ (PreservesColimit.preserves t) + isColimitMapCoconeCofanMkEquiv _ _ _ (isColimitOfPreserves G t) /-- The property of reflecting coproducts expressed in terms of cofans. -/ def isColimitOfIsColimitCofanMkObj [ReflectsColimit (Discrete.functor f) G] {P : C} (g : ∀ j, f j ⟶ P) (t : IsColimit (Cofan.mk _ fun j => G.map (g j) : Cofan fun j => G.obj (f j))) : IsColimit (Cofan.mk P g) := - ReflectsColimit.reflects ((isColimitMapCoconeCofanMkEquiv _ _ _).symm t) + isColimitOfReflects G ((isColimitMapCoconeCofanMkEquiv _ _ _).symm t) section @@ -133,13 +132,12 @@ def isColimitOfHasCoproductOfPreservesColimit [PreservesColimit (Discrete.functo variable [HasCoproduct fun j : J => G.obj (f j)] /-- If `sigma_comparison G f` is an isomorphism, then `G` preserves the colimit of `f`. -/ -def PreservesCoproduct.ofIsoComparison [i : IsIso (sigmaComparison G f)] : +lemma PreservesCoproduct.of_iso_comparison [i : IsIso (sigmaComparison G f)] : PreservesColimit (Discrete.functor f) G := by - apply preservesColimitOfPreservesColimitCocone (coproductIsCoproduct f) + apply preservesColimit_of_preserves_colimit_cocone (coproductIsCoproduct f) apply (isColimitMapCoconeCofanMkEquiv _ _ _).symm _ - refine @IsColimit.ofPointIso _ _ _ _ _ _ _ - (colimit.isColimit (Discrete.functor fun j : J => G.obj (f j))) ?_ - apply i + exact @IsColimit.ofPointIso _ _ _ _ _ _ _ + (colimit.isColimit (Discrete.functor fun j : J => G.obj (f j))) i variable [PreservesColimit (Discrete.functor f) G] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean index d79fa92a4f909..0eb7d4f49fdba 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Pullbacks.lean @@ -80,13 +80,13 @@ def isLimitPullbackConeMapOfIsLimit [PreservesLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk h k comm)) : have : G.map h ≫ G.map f = G.map k ≫ G.map g := by rw [← G.map_comp, ← G.map_comp,comm] IsLimit (PullbackCone.mk (G.map h) (G.map k) this) := - (PullbackCone.isLimitMapConeEquiv _ G).1 (PreservesLimit.preserves l) + (PullbackCone.isLimitMapConeEquiv _ G).1 (isLimitOfPreserves G l) /-- The property of reflecting pullbacks expressed in terms of binary fans. -/ def isLimitOfIsLimitPullbackConeMap [ReflectsLimit (cospan f g) G] (l : IsLimit (PullbackCone.mk (G.map h) (G.map k) (show G.map h ≫ G.map f = G.map k ≫ G.map g from by simp only [← G.map_comp,comm]))) : IsLimit (PullbackCone.mk h k comm) := - ReflectsLimit.reflects + isLimitOfReflects G ((PullbackCone.isLimitMapConeEquiv (PullbackCone.mk _ _ comm) G).2 l) variable (f g) [PreservesLimit (cospan f g) G] @@ -100,21 +100,21 @@ def isLimitOfHasPullbackOfPreservesLimit [HasPullback f g] : isLimitPullbackConeMapOfIsLimit G _ (pullbackIsPullback f g) /-- If `F` preserves the pullback of `f, g`, it also preserves the pullback of `g, f`. -/ -def preservesPullbackSymmetry : PreservesLimit (cospan g f) G where - preserves {c} hc := by +lemma preservesPullback_symmetry : PreservesLimit (cospan g f) G where + preserves {c} hc := ⟨by apply (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₂} _) _).toFun apply IsLimit.ofIsoLimit _ (PullbackCone.isoMk _).symm apply PullbackCone.isLimitOfFlip apply (isLimitMapConePullbackConeEquiv _ _).toFun - · refine @PreservesLimit.preserves _ _ _ _ _ _ _ _ ?_ _ ?_ + · refine @isLimitOfPreserves _ _ _ _ _ _ _ _ _ ?_ ?_ + · apply PullbackCone.isLimitOfFlip + apply IsLimit.ofIsoLimit _ (PullbackCone.isoMk _) + exact (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₁} _) _).invFun hc · dsimp infer_instance - apply PullbackCone.isLimitOfFlip - apply IsLimit.ofIsoLimit _ (PullbackCone.isoMk _) - exact (IsLimit.postcomposeHomEquiv (diagramIsoCospan.{v₁} _) _).invFun hc · exact (c.π.naturality WalkingCospan.Hom.inr).symm.trans - (c.π.naturality WalkingCospan.Hom.inl : _) + (c.π.naturality WalkingCospan.Hom.inl : _)⟩ theorem hasPullback_of_preservesPullback [HasPullback f g] : HasPullback (G.map f) (G.map g) := ⟨⟨⟨_, isLimitPullbackConeMapOfIsLimit G _ (pullbackIsPullback _ _)⟩⟩⟩ @@ -207,14 +207,14 @@ def isColimitPushoutCoconeMapOfIsColimit [PreservesColimit (span f g) G] (l : IsColimit (PushoutCocone.mk h k comm)) : IsColimit (PushoutCocone.mk (G.map h) (G.map k) (show G.map f ≫ G.map h = G.map g ≫ G.map k from by simp only [← G.map_comp,comm] )) := - isColimitMapCoconePushoutCoconeEquiv G comm (PreservesColimit.preserves l) + isColimitMapCoconePushoutCoconeEquiv G comm (isColimitOfPreserves G l) /-- The property of reflecting pushouts expressed in terms of binary cofans. -/ def isColimitOfIsColimitPushoutCoconeMap [ReflectsColimit (span f g) G] (l : IsColimit (PushoutCocone.mk (G.map h) (G.map k) (show G.map f ≫ G.map h = G.map g ≫ G.map k from by simp only [← G.map_comp,comm]))) : IsColimit (PushoutCocone.mk h k comm) := - ReflectsColimit.reflects ((isColimitMapCoconePushoutCoconeEquiv G comm).symm l) + isColimitOfReflects G ((isColimitMapCoconePushoutCoconeEquiv G comm).symm l) variable (f g) [PreservesColimit (span f g) G] @@ -227,16 +227,16 @@ def isColimitOfHasPushoutOfPreservesColimit [i : HasPushout f g] : isColimitPushoutCoconeMapOfIsColimit G _ (pushoutIsPushout f g) /-- If `F` preserves the pushout of `f, g`, it also preserves the pushout of `g, f`. -/ -def preservesPushoutSymmetry : PreservesColimit (span g f) G where - preserves {c} hc := by +lemma preservesPushout_symmetry : PreservesColimit (span g f) G where + preserves {c} hc := ⟨by apply (IsColimit.precomposeHomEquiv (diagramIsoSpan.{v₂} _).symm _).toFun apply IsColimit.ofIsoColimit _ (PushoutCocone.isoMk _).symm apply PushoutCocone.isColimitOfFlip apply (isColimitMapCoconePushoutCoconeEquiv _ _).toFun - · refine @PreservesColimit.preserves _ _ _ _ _ _ _ _ ?_ _ ?_ -- Porting note: more TC coddling - · dsimp - infer_instance + · refine @isColimitOfPreserves _ _ _ _ _ _ _ _ _ ?_ ?_ -- Porting note: more TC coddling · exact PushoutCocone.flipIsColimit hc + · dsimp + infer_instance⟩ theorem hasPushout_of_preservesPushout [HasPushout f g] : HasPushout (G.map f) (G.map g) := ⟨⟨⟨_, isColimitPushoutCoconeMapOfIsColimit G _ (pushoutIsPushout _ _)⟩⟩⟩ @@ -290,12 +290,11 @@ variable [HasPullback f g] [HasPullback (G.map f) (G.map g)] /-- If the pullback comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the pullback of `(f,g)`. -/ -def PreservesPullback.ofIsoComparison [i : IsIso (pullbackComparison G f g)] : +lemma PreservesPullback.of_iso_comparison [i : IsIso (pullbackComparison G f g)] : PreservesLimit (cospan f g) G := by - apply preservesLimitOfPreservesLimitCone (pullbackIsPullback f g) + apply preservesLimit_of_preserves_limit_cone (pullbackIsPullback f g) apply (isLimitMapConePullbackConeEquiv _ _).symm _ - refine @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (cospan (G.map f) (G.map g))) ?_ - apply i + exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (cospan (G.map f) (G.map g))) i variable [PreservesLimit (cospan f g) G] @@ -312,13 +311,12 @@ variable [HasPushout f g] [HasPushout (G.map f) (G.map g)] /-- If the pushout comparison map for `G` at `(f,g)` is an isomorphism, then `G` preserves the pushout of `(f,g)`. -/ -def PreservesPushout.ofIsoComparison [i : IsIso (pushoutComparison G f g)] : +lemma PreservesPushout.of_iso_comparison [i : IsIso (pushoutComparison G f g)] : PreservesColimit (span f g) G := by - apply preservesColimitOfPreservesColimitCocone (pushoutIsPushout f g) + apply preservesColimit_of_preserves_colimit_cocone (pushoutIsPushout f g) apply (isColimitMapCoconePushoutCoconeEquiv _ _).symm _ -- Porting note: apply no longer creates goals for instances - refine @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (span (G.map f) (G.map g))) ?_ - apply i + exact @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (span (G.map f) (G.map g))) i variable [PreservesColimit (span f g) G] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean index 74e1a1309fbcc..8d97f5f95983f 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Terminal.lean @@ -42,12 +42,12 @@ def isLimitMapConeEmptyConeEquiv : /-- The property of preserving terminal objects expressed in terms of `IsTerminal`. -/ def IsTerminal.isTerminalObj [PreservesLimit (Functor.empty.{0} C) G] (l : IsTerminal X) : IsTerminal (G.obj X) := - isLimitMapConeEmptyConeEquiv G X (PreservesLimit.preserves l) + isLimitMapConeEmptyConeEquiv G X (isLimitOfPreserves G l) /-- The property of reflecting terminal objects expressed in terms of `IsTerminal`. -/ def IsTerminal.isTerminalOfObj [ReflectsLimit (Functor.empty.{0} C) G] (l : IsTerminal (G.obj X)) : IsTerminal X := - ReflectsLimit.reflects ((isLimitMapConeEmptyConeEquiv G X).symm l) + isLimitOfReflects G ((isLimitMapConeEmptyConeEquiv G X).symm l) /-- A functor that preserves and reflects terminal objects induces an equivalence on `IsTerminal`. -/ @@ -60,10 +60,9 @@ def IsTerminal.isTerminalIffObj [PreservesLimit (Functor.empty.{0} C) G] right_inv := by aesop_cat /-- Preserving the terminal object implies preserving all limits of the empty diagram. -/ -def preservesLimitsOfShapePemptyOfPreservesTerminal [PreservesLimit (Functor.empty.{0} C) G] : - PreservesLimitsOfShape (Discrete PEmpty) G where - preservesLimit := - preservesLimitOfIsoDiagram G (Functor.emptyExt (Functor.empty.{0} C) _) +lemma preservesLimitsOfShape_pempty_of_preservesTerminal [PreservesLimit (Functor.empty.{0} C) G] : + PreservesLimitsOfShape (Discrete PEmpty.{1}) G where + preservesLimit := preservesLimit_of_iso_diagram G (Functor.emptyExt (Functor.empty.{0} C) _) variable [HasTerminal C] @@ -90,21 +89,21 @@ variable [HasTerminal D] /-- If the terminal comparison map for `G` is an isomorphism, then `G` preserves terminal objects. -/ -def PreservesTerminal.ofIsoComparison [i : IsIso (terminalComparison G)] : - PreservesLimit (Functor.empty C) G := by - apply preservesLimitOfPreservesLimitCone terminalIsTerminal +lemma PreservesTerminal.of_iso_comparison [i : IsIso (terminalComparison G)] : + PreservesLimit (Functor.empty.{0} C) G := by + apply preservesLimit_of_preserves_limit_cone terminalIsTerminal apply (isLimitMapConeEmptyConeEquiv _ _).symm _ exact @IsLimit.ofPointIso _ _ _ _ _ _ _ (limit.isLimit (Functor.empty.{0} D)) i /-- If there is any isomorphism `G.obj ⊤ ⟶ ⊤`, then `G` preserves terminal objects. -/ -def preservesTerminalOfIsIso (f : G.obj (⊤_ C) ⟶ ⊤_ D) [i : IsIso f] : - PreservesLimit (Functor.empty C) G := by +lemma preservesTerminal_of_isIso (f : G.obj (⊤_ C) ⟶ ⊤_ D) [i : IsIso f] : + PreservesLimit (Functor.empty.{0} C) G := by rw [Subsingleton.elim f (terminalComparison G)] at i - exact PreservesTerminal.ofIsoComparison G + exact PreservesTerminal.of_iso_comparison G /-- If there is any isomorphism `G.obj ⊤ ≅ ⊤`, then `G` preserves terminal objects. -/ -def preservesTerminalOfIso (f : G.obj (⊤_ C) ≅ ⊤_ D) : PreservesLimit (Functor.empty C) G := - preservesTerminalOfIsIso G f.hom +lemma preservesTerminal_of_iso (f : G.obj (⊤_ C) ≅ ⊤_ D) : PreservesLimit (Functor.empty.{0} C) G := + preservesTerminal_of_isIso G f.hom variable [PreservesLimit (Functor.empty.{0} C) G] @@ -134,12 +133,12 @@ def isColimitMapCoconeEmptyCoconeEquiv : /-- The property of preserving initial objects expressed in terms of `IsInitial`. -/ def IsInitial.isInitialObj [PreservesColimit (Functor.empty.{0} C) G] (l : IsInitial X) : IsInitial (G.obj X) := - isColimitMapCoconeEmptyCoconeEquiv G X (PreservesColimit.preserves l) + isColimitMapCoconeEmptyCoconeEquiv G X (isColimitOfPreserves G l) /-- The property of reflecting initial objects expressed in terms of `IsInitial`. -/ def IsInitial.isInitialOfObj [ReflectsColimit (Functor.empty.{0} C) G] (l : IsInitial (G.obj X)) : IsInitial X := - ReflectsColimit.reflects ((isColimitMapCoconeEmptyCoconeEquiv G X).symm l) + isColimitOfReflects G ((isColimitMapCoconeEmptyCoconeEquiv G X).symm l) /-- A functor that preserves and reflects initial objects induces an equivalence on `IsInitial`. -/ def IsInitial.isInitialIffObj [PreservesColimit (Functor.empty.{0} C) G] @@ -151,10 +150,11 @@ def IsInitial.isInitialIffObj [PreservesColimit (Functor.empty.{0} C) G] right_inv := by aesop_cat /-- Preserving the initial object implies preserving all colimits of the empty diagram. -/ -def preservesColimitsOfShapePemptyOfPreservesInitial [PreservesColimit (Functor.empty.{0} C) G] : - PreservesColimitsOfShape (Discrete PEmpty) G where +lemma preservesColimitsOfShape_pempty_of_preservesInitial + [PreservesColimit (Functor.empty.{0} C) G] : + PreservesColimitsOfShape (Discrete PEmpty.{1}) G where preservesColimit := - preservesColimitOfIsoDiagram G (Functor.emptyExt (Functor.empty.{0} C) _) + preservesColimit_of_iso_diagram G (Functor.emptyExt (Functor.empty.{0} C) _) variable [HasInitial C] @@ -181,21 +181,22 @@ variable [HasInitial D] /-- If the initial comparison map for `G` is an isomorphism, then `G` preserves initial objects. -/ -def PreservesInitial.ofIsoComparison [i : IsIso (initialComparison G)] : - PreservesColimit (Functor.empty C) G := by - apply preservesColimitOfPreservesColimitCocone initialIsInitial +lemma PreservesInitial.of_iso_comparison [i : IsIso (initialComparison G)] : + PreservesColimit (Functor.empty.{0} C) G := by + apply preservesColimit_of_preserves_colimit_cocone initialIsInitial apply (isColimitMapCoconeEmptyCoconeEquiv _ _).symm _ exact @IsColimit.ofPointIso _ _ _ _ _ _ _ (colimit.isColimit (Functor.empty.{0} D)) i /-- If there is any isomorphism `⊥ ⟶ G.obj ⊥`, then `G` preserves initial objects. -/ -def preservesInitialOfIsIso (f : ⊥_ D ⟶ G.obj (⊥_ C)) [i : IsIso f] : - PreservesColimit (Functor.empty C) G := by +lemma preservesInitial_of_isIso (f : ⊥_ D ⟶ G.obj (⊥_ C)) [i : IsIso f] : + PreservesColimit (Functor.empty.{0} C) G := by rw [Subsingleton.elim f (initialComparison G)] at i - exact PreservesInitial.ofIsoComparison G + exact PreservesInitial.of_iso_comparison G /-- If there is any isomorphism `⊥ ≅ G.obj ⊥`, then `G` preserves initial objects. -/ -def preservesInitialOfIso (f : ⊥_ D ≅ G.obj (⊥_ C)) : PreservesColimit (Functor.empty C) G := - preservesInitialOfIsIso G f.hom +lemma preservesInitial_of_iso (f : ⊥_ D ≅ G.obj (⊥_ C)) : + PreservesColimit (Functor.empty.{0} C) G := + preservesInitial_of_isIso G f.hom variable [PreservesColimit (Functor.empty.{0} C) G] diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean index 1b3b19cbe4153..9378d2cf0aea9 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean @@ -157,15 +157,15 @@ instance (priority := 100) preservesZeroMorphisms_of_preserves_terminal_object variable (F) /-- Preserving zero morphisms implies preserving terminal objects. -/ -def preservesTerminalObjectOfPreservesZeroMorphisms [PreservesZeroMorphisms F] : - PreservesLimit (Functor.empty C) F := - preservesTerminalOfIso F <| +lemma preservesTerminalObject_of_preservesZeroMorphisms [PreservesZeroMorphisms F] : + PreservesLimit (Functor.empty.{0} C) F := + preservesTerminal_of_iso F <| F.mapIso HasZeroObject.zeroIsoTerminal.symm ≪≫ mapZeroObject F ≪≫ HasZeroObject.zeroIsoTerminal /-- Preserving zero morphisms implies preserving terminal objects. -/ -def preservesInitialObjectOfPreservesZeroMorphisms [PreservesZeroMorphisms F] : - PreservesColimit (Functor.empty C) F := - preservesInitialOfIso F <| +lemma preservesInitialObject_of_preservesZeroMorphisms [PreservesZeroMorphisms F] : + PreservesColimit (Functor.empty.{0} C) F := + preservesInitial_of_iso F <| HasZeroObject.zeroIsoInitial.symm ≪≫ (mapZeroObject F).symm ≪≫ (F.mapIso HasZeroObject.zeroIsoInitial.symm).symm @@ -176,25 +176,27 @@ section variable [HasZeroObject D] [HasZeroMorphisms D] (G : C ⥤ D) (hG : IsZero G) (J : Type*) [Category J] +include hG + /-- A zero functor preserves limits. -/ -def preservesLimitsOfShapeOfIsZero : PreservesLimitsOfShape J G where - preservesLimit {K} := ⟨fun _ => by +lemma preservesLimitsOfShape_of_isZero : PreservesLimitsOfShape J G where + preservesLimit {K} := ⟨fun _ => ⟨by rw [Functor.isZero_iff] at hG - exact IsLimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩ + exact IsLimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩⟩ /-- A zero functor preserves colimits. -/ -def preservesColimitsOfShapeOfIsZero : PreservesColimitsOfShape J G where - preservesColimit {K} := ⟨fun _ => by +lemma preservesColimitsOfShape_of_isZero : PreservesColimitsOfShape J G where + preservesColimit {K} := ⟨fun _ => ⟨by rw [Functor.isZero_iff] at hG - exact IsColimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩ + exact IsColimit.ofIsZero _ ((K ⋙ G).isZero (fun X ↦ hG _)) (hG _)⟩⟩ /-- A zero functor preserves limits. -/ -def preservesLimitsOfSizeOfIsZero : PreservesLimitsOfSize.{v, u} G where - preservesLimitsOfShape := G.preservesLimitsOfShapeOfIsZero hG _ +lemma preservesLimitsOfSize_of_isZero : PreservesLimitsOfSize.{v, u} G where + preservesLimitsOfShape := G.preservesLimitsOfShape_of_isZero hG _ /-- A zero functor preserves colimits. -/ -def preservesColimitsOfSizeOfIsZero : PreservesColimitsOfSize.{v, u} G where - preservesColimitsOfShape := G.preservesColimitsOfShapeOfIsZero hG _ +lemma preservesColimitsOfSize_of_isZero : PreservesColimitsOfSize.{v, u} G where + preservesColimitsOfShape := G.preservesColimitsOfShape_of_isZero hG _ end diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index e11c09ddc138d..eb1da8048a51a 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -42,7 +42,6 @@ instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesLimitsOfShape {J} := { preservesLimit := fun {K} => { preserves := fun {c} hc => by - apply Nonempty.some rw [Types.isLimit_iff ((uliftFunctor.{v, u}).mapCone c)] intro s hs obtain ⟨x, hx₁, hx₂⟩ := (Types.isLimit_iff c).mp ⟨hc⟩ _ ((sectionsEquiv K).symm ⟨s, hs⟩).2 @@ -138,11 +137,11 @@ The functor `uliftFunctor : Type u ⥤ Type (max u v)` preserves colimits of arb noncomputable instance : PreservesColimitsOfSize.{w', w} uliftFunctor.{v, u} where preservesColimitsOfShape {J _} := { preservesColimit := fun {F} ↦ - { preserves := fun {c} hc ↦ - { desc := fun lc x ↦ descFun hc lc x.down + { preserves := fun {c} hc ↦ ⟨{ + desc := fun lc x ↦ descFun hc lc x.down fac := fun lc j ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc _).mp rfl j) uniq := fun lc f hf ↦ by ext ⟨⟩; apply congr_fun ((descFun_spec hc lc (f ∘ ULift.up)).mpr - fun j ↦ funext fun y ↦ congr_fun (hf j) ⟨y⟩) } } } + fun j ↦ funext fun y ↦ congr_fun (hf j) ⟨y⟩) }⟩ } } /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` creates `u`-small colimits. diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean index dabe16ad73949..7861c1e682600 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Yoneda.lean @@ -70,7 +70,7 @@ theorem yonedaYonedaColimit_app_inv {X : C} : ((yonedaYonedaColimit F).app (op X noncomputable instance {X : C} : PreservesColimit F (coyoneda.obj (op (yoneda.obj X))) := by suffices IsIso (colimit.post F (coyoneda.obj (op (yoneda.obj X)))) from - preservesColimitOfIsIsoPost _ _ + preservesColimit_of_isIso_post _ _ suffices colimit.post F (coyoneda.obj (op (yoneda.obj X))) = (colimitObjIsoColimitCompEvaluation _ _).inv ≫ ((yonedaYonedaColimit F).app (op X)).inv from this ▸ inferInstance diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index ca7b9175b9037..9eca881c472cc 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -155,10 +155,11 @@ noncomputable def yonedaAdjunction : L ⊣ restrictedYoneda A := apply yonedaEquiv.injective simp [yonedaEquiv] } +include α in /-- Any left Kan extension along the Yoneda embedding preserves colimits. -/ -noncomputable def preservesColimitsOfSizeOfIsLeftKanExtension : +lemma preservesColimitsOfSize_of_isLeftKanExtension : PreservesColimitsOfSize.{v₃, u₃} L := - (yonedaAdjunction L α).leftAdjointPreservesColimits + (yonedaAdjunction L α).leftAdjoint_preservesColimits lemma isIso_of_isLeftKanExtension : IsIso α := (Functor.isPointwiseLeftKanExtensionOfIsLeftKanExtension _ α).isIso_hom @@ -166,9 +167,9 @@ lemma isIso_of_isLeftKanExtension : IsIso α := variable (A) /-- See Property 2 of https://ncatlab.org/nlab/show/Yoneda+extension#properties. -/ -noncomputable instance preservesColimitsOfSizeLeftKanExtension : +noncomputable instance preservesColimitsOfSize_leftKanExtension : PreservesColimitsOfSize.{v₃, u₃} (yoneda.leftKanExtension A) := - (yonedaAdjunction _ (yoneda.leftKanExtensionUnit A)).leftAdjointPreservesColimits + (yonedaAdjunction _ (yoneda.leftKanExtensionUnit A)).leftAdjoint_preservesColimits instance : IsIso (yoneda.leftKanExtensionUnit A) := isIso_of_isLeftKanExtension _ (yoneda.leftKanExtensionUnit A) @@ -257,12 +258,12 @@ instance [L.IsLeftKanExtension α] : IsIso α := lemma isLeftKanExtension_along_yoneda_iff : L.IsLeftKanExtension α ↔ - (IsIso α ∧ Nonempty (PreservesColimitsOfSize.{v₁, max u₁ v₁} L)) := by + (IsIso α ∧ PreservesColimitsOfSize.{v₁, max u₁ v₁} L) := by constructor · intro - exact ⟨inferInstance, ⟨preservesColimitsOfNatIso - (Functor.leftKanExtensionUnique _ (yoneda.leftKanExtensionUnit A) _ α)⟩⟩ - · rintro ⟨_, ⟨_⟩⟩ + exact ⟨inferInstance, preservesColimits_of_natIso + (Functor.leftKanExtensionUnique _ (yoneda.leftKanExtensionUnit A) _ α)⟩ + · rintro ⟨_, _⟩ apply Functor.LeftExtension.IsPointwiseLeftKanExtension.isLeftKanExtension (E := Functor.LeftExtension.mk _ α) intro P @@ -571,7 +572,7 @@ theorem final_toCostructuredArrow_comp_pre {c : Cocone (F ⋙ yoneda)} (hc : IsC (CostructuredArrow.toOverCompOverEquivPresheafCostructuredArrow c.pt).isoCompInverse apply IsTerminal.ofIso Over.mkIdTerminal - let isc : IsColimit ((Over.forget _).mapCocone _) := PreservesColimit.preserves + let isc : IsColimit ((Over.forget _).mapCocone _) := isColimitOfPreserves _ (colimit.isColimit ((c.toCostructuredArrow ⋙ CostructuredArrow.pre F yoneda c.pt) ⋙ CostructuredArrow.toOver yoneda c.pt)) exact Over.isoMk (hc.coconePointUniqueUpToIso isc) (hc.hom_ext fun i => by simp) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean index e16617bd21fec..8fba6cab9585a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/NormalMono/Basic.lean @@ -59,16 +59,11 @@ def equivalenceReflectsNormalMono {D : Type u₂} [Category.{v₁} D] [HasZeroMo have reassoc' {W : D} (h : hf.Z ⟶ W) : F.map f ≫ hf.g ≫ h = 0 ≫ h := by rw [← Category.assoc, eq_whisker hf.w] simp [reassoc'] - isLimit := - @ReflectsLimit.reflects C _ D _ _ _ _ F _ _ <| - IsLimit.ofConeEquiv (Cones.postcomposeEquivalence (@compNatIso C _ _ _ _ _ D _ _ F _)) <| - IsLimit.ofIsoLimit - (IsLimit.ofIsoLimit - (IsKernel.ofCompIso _ _ (F.objObjPreimageIso hf.Z) (by - simp only [Functor.map_preimage, Category.assoc, Iso.inv_hom_id, Category.comp_id]) - hf.isLimit) - (ofιCongr (Category.comp_id _).symm)) - <| by apply Iso.symm; apply isoOfι -- Porting note: very fiddly unification here + isLimit := isLimitOfReflects F <| + IsLimit.ofConeEquiv (Cones.postcomposeEquivalence (compNatIso F)) <| + (IsLimit.ofIsoLimit (IsKernel.ofCompIso _ _ (F.objObjPreimageIso hf.Z) (by + simp only [Functor.map_preimage, Category.assoc, Iso.inv_hom_id, Category.comp_id]) + hf.isLimit)) (Fork.ext (Iso.refl _) (by simp [compNatIso, Fork.ι])) end @@ -160,14 +155,11 @@ def equivalenceReflectsNormalEpi {D : Type u₂} [Category.{v₁} D] [HasZeroMor W := F.objPreimage hf.W g := F.preimage ((F.objObjPreimageIso hf.W).hom ≫ hf.g) w := F.map_injective <| by simp [hf.w] - isColimit := - ReflectsColimit.reflects <| - IsColimit.ofCoconeEquiv (Cocones.precomposeEquivalence (compNatIso F).symm) <| - IsColimit.ofIsoColimit - (IsColimit.ofIsoColimit - (IsCokernel.ofIsoComp _ _ (F.objObjPreimageIso hf.W).symm (by simp) hf.isColimit) - (ofπCongr (Category.id_comp _).symm)) - <| by apply Iso.symm; apply isoOfπ + isColimit := isColimitOfReflects F <| + IsColimit.ofCoconeEquiv (Cocones.precomposeEquivalence (compNatIso F).symm) <| + (IsColimit.ofIsoColimit + (IsCokernel.ofIsoComp _ _ (F.objObjPreimageIso hf.W).symm (by simp) hf.isColimit) + (Cofork.ext (Iso.refl _) (by simp [compNatIso, Cofork.π]))) end diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index 429dea9a1687b..75e50cc0928fc 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -188,7 +188,7 @@ theorem IsUniversalColimit.of_mapCocone (G : C ⥤ D) {F : J ⥤ C} {c : Cocone [PreservesLimitsOfShape WalkingCospan G] [ReflectsColimitsOfShape J G] (hc : IsUniversalColimit (G.mapCocone c)) : IsUniversalColimit c := fun F' c' α f h hα H ↦ - ⟨ReflectsColimit.reflects (hc (G.mapCocone c') (whiskerRight α G) (G.map f) + ⟨isColimitOfReflects _ (hc (G.mapCocone c') (whiskerRight α G) (G.map f) (by ext j; simpa using G.congr_map (NatTrans.congr_app h j)) (hα.whiskerRight G) (fun j ↦ (H j).map G)).some⟩ @@ -282,7 +282,7 @@ theorem isVanKampenColimit_of_evaluation [HasPullbacks D] [HasColimitsOfShape J · rintro ⟨hc'⟩ j refine ⟨⟨(NatTrans.congr_app e j).symm⟩, ⟨evaluationJointlyReflectsLimits _ ?_⟩⟩ refine fun x => (isLimitMapConePullbackConeEquiv _ _).symm ?_ - exact ((this x).mp ⟨PreservesColimit.preserves hc'⟩ _).isLimit + exact ((this x).mp ⟨isColimitOfPreserves _ hc'⟩ _).isLimit · exact fun H => ⟨evaluationJointlyReflectsColimits _ fun x => ((this x).mpr fun j => (H j).map ((evaluation C D).obj x)).some⟩ @@ -297,8 +297,8 @@ theorem IsUniversalColimit.map_reflective [∀ X (f : X ⟶ Gl.obj c.pt), HasPullback (Gr.map f) (adj.unit.app c.pt)] [∀ X (f : X ⟶ Gl.obj c.pt), PreservesLimit (cospan (Gr.map f) (adj.unit.app c.pt)) Gl] : IsUniversalColimit (Gl.mapCocone c) := by - have := adj.rightAdjointPreservesLimits - have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjointPreservesColimits + have := adj.rightAdjoint_preservesLimits + have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjoint_preservesColimits intros F' c' α f h hα hc' have : HasPullback (Gl.map (Gr.map f)) (Gl.map (adj.unit.app c.pt)) := ⟨⟨_, isLimitPullbackConeMapOfIsLimit _ pullback.condition @@ -401,8 +401,8 @@ theorem IsVanKampenColimit.map_reflective [HasColimitsOfShape J C] [∀ X (f : X ⟶ Gl.obj c.pt), PreservesLimit (cospan (Gr.map f) (adj.unit.app c.pt)) Gl] [∀ X i (f : X ⟶ c.pt), PreservesLimit (cospan f (c.ι.app i)) Gl] : IsVanKampenColimit (Gl.mapCocone c) := by - have := adj.rightAdjointPreservesLimits - have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjointPreservesColimits + have := adj.rightAdjoint_preservesLimits + have : PreservesColimitsOfSize.{u', v'} Gl := adj.leftAdjoint_preservesColimits intro F' c' α f h hα refine ⟨?_, H.isUniversal.map_reflective adj c' α f h hα⟩ intro ⟨hc'⟩ j diff --git a/Mathlib/CategoryTheory/Limits/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Yoneda.lean index ed9f8ff3d339e..fdb2d7ebc8f71 100644 --- a/Mathlib/CategoryTheory/Limits/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Yoneda.lean @@ -80,19 +80,19 @@ def Limits.coneOfSectionCompYoneda (F : J ⥤ Cᵒᵖ) (X : C) pt := Opposite.op X π := compYonedaSectionsEquiv F X s -noncomputable instance yonedaPreservesLimit (F : J ⥤ Cᵒᵖ) (X : C) : +instance yoneda_preservesLimit (F : J ⥤ Cᵒᵖ) (X : C) : PreservesLimit F (yoneda.obj X) where - preserves {c} hc := Nonempty.some (by + preserves {c} hc := by rw [Types.isLimit_iff] intro s hs exact ⟨(hc.lift (Limits.coneOfSectionCompYoneda F X ⟨s, hs⟩)).unop, fun j => Quiver.Hom.op_inj (hc.fac (Limits.coneOfSectionCompYoneda F X ⟨s, hs⟩) j), fun m hm => Quiver.Hom.op_inj (hc.uniq (Limits.coneOfSectionCompYoneda F X ⟨s, hs⟩) _ - (fun j => Quiver.Hom.unop_inj (hm j)))⟩) + (fun j => Quiver.Hom.unop_inj (hm j)))⟩ variable (J) in -noncomputable instance yonedaPreservesLimitsOfShape (X : C) : +noncomputable instance yoneda_preservesLimitsOfShape (X : C) : PreservesLimitsOfShape J (yoneda.obj X) where /-- The yoneda embeddings jointly reflect limits. -/ @@ -125,13 +125,13 @@ def Limits.coneOfSectionCompCoyoneda (F : J ⥤ C) (X : Cᵒᵖ) pt := X.unop π := compCoyonedaSectionsEquiv F X.unop s -noncomputable instance coyonedaPreservesLimit (F : J ⥤ C) (X : Cᵒᵖ) : +instance coyoneda_preservesLimit (F : J ⥤ C) (X : Cᵒᵖ) : PreservesLimit F (coyoneda.obj X) where - preserves {c} hc := Nonempty.some (by + preserves {c} hc := by rw [Types.isLimit_iff] intro s hs exact ⟨hc.lift (Limits.coneOfSectionCompCoyoneda F X ⟨s, hs⟩), hc.fac _, - hc.uniq (Limits.coneOfSectionCompCoyoneda F X ⟨s, hs⟩)⟩) + hc.uniq (Limits.coneOfSectionCompCoyoneda F X ⟨s, hs⟩)⟩ variable (J) in noncomputable instance coyonedaPreservesLimitsOfShape (X : Cᵒᵖ) : @@ -162,31 +162,31 @@ noncomputable def Limits.Cone.isLimitCoyonedaEquiv {F : J ⥤ C} (c : Cone F) : end /-- The yoneda embedding `yoneda.obj X : Cᵒᵖ ⥤ Type v` for `X : C` preserves limits. -/ -noncomputable instance yonedaPreservesLimits (X : C) : +instance yoneda_preservesLimits (X : C) : PreservesLimitsOfSize.{t, w} (yoneda.obj X) where /-- The coyoneda embedding `coyoneda.obj X : C ⥤ Type v` for `X : Cᵒᵖ` preserves limits. -/ -noncomputable instance coyonedaPreservesLimits (X : Cᵒᵖ) : +instance coyoneda_preservesLimits (X : Cᵒᵖ) : PreservesLimitsOfSize.{t, w} (coyoneda.obj X) where -noncomputable instance yonedaFunctorPreservesLimits : +instance yonedaFunctor_preservesLimits : PreservesLimitsOfSize.{t, w} (@yoneda C _) := by - apply preservesLimitsOfEvaluation + apply preservesLimits_of_evaluation intro K change PreservesLimitsOfSize (coyoneda.obj K) infer_instance -noncomputable instance coyonedaFunctorPreservesLimits : +noncomputable instance coyonedaFunctor_preservesLimits : PreservesLimitsOfSize.{t, w} (@coyoneda C _) := by - apply preservesLimitsOfEvaluation + apply preservesLimits_of_evaluation intro K change PreservesLimitsOfSize (yoneda.obj K) infer_instance -noncomputable instance yonedaFunctorReflectsLimits : +noncomputable instance yonedaFunctor_reflectsLimits : ReflectsLimitsOfSize.{t, w} (@yoneda C _) := inferInstance -noncomputable instance coyonedaFunctorReflectsLimits : +noncomputable instance coyonedaFunctor_reflectsLimits : ReflectsLimitsOfSize.{t, w} (@coyoneda C _) := inferInstance namespace Functor @@ -195,15 +195,15 @@ section Representable variable (F : Cᵒᵖ ⥤ Type v) [F.IsRepresentable] {J : Type*} [Category J] -noncomputable instance representablePreservesLimit (G : J ⥤ Cᵒᵖ) : +instance representable_preservesLimit (G : J ⥤ Cᵒᵖ) : PreservesLimit G F := - preservesLimitOfNatIso _ F.reprW + preservesLimit_of_natIso _ F.reprW variable (J) in -noncomputable instance representablePreservesLimitsOfShape : +instance representable_preservesLimitsOfShape : PreservesLimitsOfShape J F where -noncomputable instance representablePreservesLimits : +instance representable_preservesLimits : PreservesLimitsOfSize.{t, w} F where end Representable @@ -212,15 +212,15 @@ section Corepresentable variable (F : C ⥤ Type v) [F.IsCorepresentable] {J : Type*} [Category J] -noncomputable instance corepresentablePreservesLimit (G : J ⥤ C) : +instance corepresentable_preservesLimit (G : J ⥤ C) : PreservesLimit G F := - preservesLimitOfNatIso _ F.coreprW + preservesLimit_of_natIso _ F.coreprW variable (J) in -noncomputable instance corepresentablePreservesLimitsOfShape : +instance corepresentable_preservesLimitsOfShape : PreservesLimitsOfShape J F where -noncomputable instance corepresentablePreservesLimits : +instance corepresentable_preservesLimits : PreservesLimitsOfSize.{t, w} F where end Corepresentable diff --git a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean index 7fe7c9bfe0e8a..d3be845da56a6 100644 --- a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean +++ b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean @@ -99,10 +99,10 @@ lemma hasProductsOfShape (J : Type) [Finite J] [HasProductsOfShape J C] /-- When `C` has finite products indexed by `J`, `W : MorphismProperty C` contains identities and is stable by products indexed by `J`, then any localization functor for `W` preserves finite products indexed by `J`. -/ -noncomputable def preservesProductsOfShape (J : Type) [Finite J] +lemma preservesProductsOfShape (J : Type) [Finite J] [HasProductsOfShape J C] (hW : W.IsStableUnderProductsOfShape J) : PreservesLimitsOfShape (Discrete J) L where - preservesLimit {F} := preservesLimitOfPreservesLimitCone (limit.isLimit F) + preservesLimit {F} := preservesLimit_of_preserves_limit_cone (limit.isLimit F) (HasProductsOfShapeAux.isLimitMapCone L hW F) variable [HasFiniteProducts C] [W.IsStableUnderFiniteProducts] @@ -112,10 +112,11 @@ lemma hasFiniteProducts : HasFiniteProducts D := ⟨fun _ => hasProductsOfShape L W _ (W.isStableUnderProductsOfShape_of_isStableUnderFiniteProducts _)⟩ +include W in /-- When `C` has finite products and `W : MorphismProperty C` contains identities and is stable by finite products, then any localization functor for `W` preserves finite products. -/ -noncomputable def preservesFiniteProducts : +lemma preservesFiniteProducts : PreservesFiniteProducts L where preserves J _ := preservesProductsOfShape L W J (W.isStableUnderProductsOfShape_of_isStableUnderFiniteProducts _) diff --git a/Mathlib/CategoryTheory/Monad/Comonadicity.lean b/Mathlib/CategoryTheory/Monad/Comonadicity.lean index 2649f2cfb69b3..d72e8b6e22780 100644 --- a/Mathlib/CategoryTheory/Monad/Comonadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Comonadicity.lean @@ -223,10 +223,8 @@ Beck's comonadicity theorem, the converse is given in `comonadicOfCreatesFSplitE def createsFSplitEqualizersOfComonadic [ComonadicLeftAdjoint F] ⦃A B⦄ (f g : A ⟶ B) [F.IsCosplitPair f g] : CreatesLimit (parallelPair f g) F := by apply (config := {allowSynthFailures := true}) comonadicCreatesLimitOfPreservesLimit - · apply @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ - dsimp - infer_instance - · apply @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ + all_goals + apply @preservesLimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ dsimp infer_instance @@ -335,7 +333,7 @@ def comonadicOfHasPreservesFSplitEqualizersOfReflectsIsomorphisms [F.ReflectsIso ComonadicLeftAdjoint F := by have : ReflectsLimitOfIsCosplitPair F := ⟨fun f g _ => by have := HasEqualizerOfIsCosplitPair.out F f g - apply reflectsLimitOfReflectsIsomorphisms⟩ + apply reflectsLimit_of_reflectsIsomorphisms⟩ apply comonadicOfHasPreservesReflectsFSplitEqualizers adj end BeckComonadicity @@ -387,7 +385,7 @@ def comonadicOfHasPreservesCoreflexiveEqualizersOfReflectsIsomorphisms : · rw [← G.map_id] simp apply @unitEqualizerOfCoreflectsEqualizer _ _ _ _ _ _ _ _ ?_ - apply reflectsLimitOfReflectsIsomorphisms + apply reflectsLimit_of_reflectsIsomorphisms exact (comparisonAdjunction adj).toEquivalence.symm.isEquivalence_inverse end CoreflexiveComonadicity diff --git a/Mathlib/CategoryTheory/Monad/Limits.lean b/Mathlib/CategoryTheory/Monad/Limits.lean index 11d531a830a9b..8e99e8eace74d 100644 --- a/Mathlib/CategoryTheory/Monad/Limits.lean +++ b/Mathlib/CategoryTheory/Monad/Limits.lean @@ -154,7 +154,7 @@ Define the map `λ : TL ⟶ L`, which will serve as the structure of the coalgeb we will show is the colimiting object. We use the cocone constructed by `c` and the fact that `T` preserves colimits to produce this morphism. -/ -abbrev lambda : ((T : C ⥤ C).mapCocone c).pt ⟶ c.pt := +noncomputable abbrev lambda : ((T : C ⥤ C).mapCocone c).pt ⟶ c.pt := (isColimitOfPreserves _ t).desc (newCocone c) /-- (Impl) The key property defining the map `λ : TL ⟶ L`. -/ @@ -169,7 +169,7 @@ show it satisfies the two algebra laws, which follow from the algebra laws for t our `commuting` lemma. -/ @[simps] -def coconePoint : Algebra T where +noncomputable def coconePoint : Algebra T where A := c.pt a := lambda c t unit := by @@ -189,7 +189,7 @@ def coconePoint : Algebra T where /-- (Impl) Construct the lifted cocone in `Algebra T` which will be colimiting. -/ @[simps] -def liftedCocone : Cocone D where +noncomputable def liftedCocone : Cocone D where pt := coconePoint c t ι := { app := fun j => @@ -203,7 +203,7 @@ def liftedCocone : Cocone D where /-- (Impl) Prove that the lifted cocone is colimiting. -/ @[simps] -def liftedCoconeIsColimit : IsColimit (liftedCocone c t) where +noncomputable def liftedCoconeIsColimit : IsColimit (liftedCocone c t) where desc s := { f := t.desc ((forget T).mapCocone s) h := @@ -293,13 +293,13 @@ noncomputable def monadicCreatesColimitOfPreservesColimit (R : D ⥤ C) (K : J (Adjunction.toMonad (monadicAdjunction R))) (Adjunction.toMonad (monadicAdjunction R)).toFunctor := by dsimp - exact preservesColimitOfIsoDiagram _ i.symm + exact preservesColimit_of_iso_diagram _ i.symm letI : PreservesColimit (((K ⋙ A) ⋙ Monad.forget (Adjunction.toMonad (monadicAdjunction R))) ⋙ (Adjunction.toMonad (monadicAdjunction R)).toFunctor) (Adjunction.toMonad (monadicAdjunction R)).toFunctor := by dsimp - exact preservesColimitOfIsoDiagram _ (isoWhiskerRight i (monadicLeftAdjoint R ⋙ R)).symm + exact preservesColimit_of_iso_diagram _ (isoWhiskerRight i (monadicLeftAdjoint R ⋙ R)).symm letI : CreatesColimit (K ⋙ A) B := CategoryTheory.Monad.forgetCreatesColimit _ letI : CreatesColimit K (A ⋙ B) := CategoryTheory.compCreatesColimit _ _ let e := Monad.comparisonForget (monadicAdjunction R) @@ -309,9 +309,9 @@ noncomputable def monadicCreatesColimitOfPreservesColimit (R : D ⥤ C) (K : J noncomputable def monadicCreatesColimitsOfShapeOfPreservesColimitsOfShape (R : D ⥤ C) [MonadicRightAdjoint R] [PreservesColimitsOfShape J R] : CreatesColimitsOfShape J R := letI : PreservesColimitsOfShape J (monadicLeftAdjoint R) := by - apply (Adjunction.leftAdjointPreservesColimits (monadicAdjunction R)).1 + apply (Adjunction.leftAdjoint_preservesColimits (monadicAdjunction R)).1 letI : PreservesColimitsOfShape J (monadicLeftAdjoint R ⋙ R) := by - apply CategoryTheory.Limits.compPreservesColimitsOfShape _ _ + apply CategoryTheory.Limits.comp_preservesColimitsOfShape _ _ ⟨monadicCreatesColimitOfPreservesColimit _ _⟩ /-- A monadic functor creates colimits if it preserves colimits. -/ @@ -343,7 +343,7 @@ theorem hasColimitsOfShape_of_reflective (R : D ⥤ C) [Reflective R] [HasColimi has_colimit := fun F => by let c := (monadicLeftAdjoint R).mapCocone (colimit.cocone (F ⋙ R)) letI : PreservesColimitsOfShape J _ := - (monadicAdjunction R).leftAdjointPreservesColimits.1 + (monadicAdjunction R).leftAdjoint_preservesColimits.1 let t : IsColimit c := isColimitOfPreserves (monadicLeftAdjoint R) (colimit.isColimit _) apply HasColimit.mk ⟨_, (IsColimit.precomposeInvEquiv _ _).symm t⟩ apply @@ -357,7 +357,7 @@ theorem hasColimits_of_reflective (R : D ⥤ C) [Reflective R] [HasColimitsOfSiz /-- The reflector always preserves terminal objects. Note this in general doesn't apply to any other limit. -/ -noncomputable def leftAdjointPreservesTerminalOfReflective (R : D ⥤ C) [Reflective R] : +lemma leftAdjoint_preservesTerminal_of_reflective (R : D ⥤ C) [Reflective R] : PreservesLimitsOfShape (Discrete.{v} PEmpty) (monadicLeftAdjoint R) where preservesLimit {K} := by let F := Functor.empty.{v} D @@ -366,13 +366,14 @@ noncomputable def leftAdjointPreservesTerminalOfReflective (R : D ⥤ C) [Reflec intro c h haveI : HasLimit (F ⋙ R) := ⟨⟨⟨c, h⟩⟩⟩ haveI : HasLimit F := hasLimit_of_reflective F R + constructor apply isLimitChangeEmptyCone D (limit.isLimit F) apply (asIso ((monadicAdjunction R).counit.app _)).symm.trans apply (monadicLeftAdjoint R).mapIso letI := monadicCreatesLimits.{v, v} R - let A := CategoryTheory.preservesLimitOfCreatesLimitAndHasLimit F R - apply (A.preserves (limit.isLimit F)).conePointUniqueUpToIso h - apply preservesLimitOfIsoDiagram _ (Functor.emptyExt (F ⋙ R) _) + let A := CategoryTheory.preservesLimit_of_createsLimit_and_hasLimit F R + apply (isLimitOfPreserves _ (limit.isLimit F)).conePointUniqueUpToIso h + apply preservesLimit_of_iso_diagram _ (Functor.emptyExt (F ⋙ R) _) end @@ -484,7 +485,7 @@ Define the map `λ : L ⟶ TL`, which will serve as the structure of the algebra we will show is the limiting object. We use the cone constructed by `c` and the fact that `T` preserves limits to produce this morphism. -/ -abbrev lambda : c.pt ⟶ ((T : C ⥤ C).mapCone c).pt := +noncomputable abbrev lambda : c.pt ⟶ ((T : C ⥤ C).mapCone c).pt := (isLimitOfPreserves _ t).lift (newCone c) /-- (Impl) The key property defining the map `λ : L ⟶ TL`. -/ @@ -500,7 +501,7 @@ show it satisfies the two coalgebra laws, which follow from the coalgebra laws f and our `commuting` lemma. -/ @[simps] -def conePoint : Coalgebra T where +noncomputable def conePoint : Coalgebra T where A := c.pt a := lambda c t counit := t.hom_ext fun j ↦ by @@ -516,7 +517,7 @@ def conePoint : Coalgebra T where /-- (Impl) Construct the lifted cone in `Coalgebra T` which will be limiting. -/ @[simps] -def liftedCone : Cone D where +noncomputable def liftedCone : Cone D where pt := conePoint c t π := { app := fun j => @@ -530,7 +531,7 @@ def liftedCone : Cone D where /-- (Impl) Prove that the lifted cone is limiting. -/ @[simps] -def liftedConeIsLimit : IsLimit (liftedCone c t) where +noncomputable def liftedConeIsLimit : IsLimit (liftedCone c t) where lift s := { f := t.lift ((forget T).mapCone s) h := @@ -616,13 +617,13 @@ noncomputable def comonadicCreatesLimitOfPreservesLimit (R : D ⥤ C) (K : J ⥤ (Adjunction.toComonad (comonadicAdjunction R))) (Adjunction.toComonad (comonadicAdjunction R)).toFunctor := by dsimp - exact preservesLimitOfIsoDiagram _ i.symm + exact preservesLimit_of_iso_diagram _ i.symm letI : PreservesLimit (((K ⋙ A) ⋙ Comonad.forget (Adjunction.toComonad (comonadicAdjunction R))) ⋙ (Adjunction.toComonad (comonadicAdjunction R)).toFunctor) (Adjunction.toComonad (comonadicAdjunction R)).toFunctor := by dsimp - exact preservesLimitOfIsoDiagram _ (isoWhiskerRight i (comonadicRightAdjoint R ⋙ R)).symm + exact preservesLimit_of_iso_diagram _ (isoWhiskerRight i (comonadicRightAdjoint R ⋙ R)).symm letI : CreatesLimit (K ⋙ A) B := CategoryTheory.Comonad.forgetCreatesLimit _ letI : CreatesLimit K (A ⋙ B) := CategoryTheory.compCreatesLimit _ _ let e := Comonad.comparisonForget (comonadicAdjunction R) @@ -632,9 +633,9 @@ noncomputable def comonadicCreatesLimitOfPreservesLimit (R : D ⥤ C) (K : J ⥤ noncomputable def comonadicCreatesLimitsOfShapeOfPreservesLimitsOfShape (R : D ⥤ C) [ComonadicLeftAdjoint R] [PreservesLimitsOfShape J R] : CreatesLimitsOfShape J R := letI : PreservesLimitsOfShape J (comonadicRightAdjoint R) := by - apply (Adjunction.rightAdjointPreservesLimits (comonadicAdjunction R)).1 + apply (Adjunction.rightAdjoint_preservesLimits (comonadicAdjunction R)).1 letI : PreservesLimitsOfShape J (comonadicRightAdjoint R ⋙ R) := by - apply CategoryTheory.Limits.compPreservesLimitsOfShape _ _ + apply CategoryTheory.Limits.comp_preservesLimitsOfShape _ _ ⟨comonadicCreatesLimitOfPreservesLimit _ _⟩ /-- A comonadic functor creates limits if it preserves limits. -/ @@ -666,7 +667,7 @@ theorem hasLimitsOfShape_of_coreflective (R : D ⥤ C) [Coreflective R] [HasLimi has_limit := fun F => by let c := (comonadicRightAdjoint R).mapCone (limit.cone (F ⋙ R)) letI : PreservesLimitsOfShape J _ := - (comonadicAdjunction R).rightAdjointPreservesLimits.1 + (comonadicAdjunction R).rightAdjoint_preservesLimits.1 let t : IsLimit c := isLimitOfPreserves (comonadicRightAdjoint R) (limit.isLimit _) apply HasLimit.mk ⟨_, (IsLimit.postcomposeHomEquiv _ _).symm t⟩ apply @@ -680,7 +681,7 @@ theorem hasLimits_of_coreflective (R : D ⥤ C) [Coreflective R] [HasLimitsOfSiz /-- The coreflector always preserves initial objects. Note this in general doesn't apply to any other colimit. -/ -noncomputable def rightAdjointPreservesInitialOfCoreflective (R : D ⥤ C) [Coreflective R] : +lemma rightAdjoint_preservesInitial_of_coreflective (R : D ⥤ C) [Coreflective R] : PreservesColimitsOfShape (Discrete.{v} PEmpty) (comonadicRightAdjoint R) where preservesColimit {K} := by let F := Functor.empty.{v} D @@ -689,13 +690,14 @@ noncomputable def rightAdjointPreservesInitialOfCoreflective (R : D ⥤ C) [Core intro c h haveI : HasColimit (F ⋙ R) := ⟨⟨⟨c, h⟩⟩⟩ haveI : HasColimit F := hasColimit_of_coreflective F R + constructor apply isColimitChangeEmptyCocone D (colimit.isColimit F) apply (asIso ((comonadicAdjunction R).unit.app _)).trans apply (comonadicRightAdjoint R).mapIso letI := comonadicCreatesColimits.{v, v} R - let A := CategoryTheory.preservesColimitOfCreatesColimitAndHasColimit F R - apply (A.preserves (colimit.isColimit F)).coconePointUniqueUpToIso h - apply preservesColimitOfIsoDiagram _ (Functor.emptyExt (F ⋙ R) _) + let A := CategoryTheory.preservesColimit_of_createsColimit_and_hasColimit F R + apply (isColimitOfPreserves _ (colimit.isColimit F)).coconePointUniqueUpToIso h + apply preservesColimit_of_iso_diagram _ (Functor.emptyExt (F ⋙ R) _) end diff --git a/Mathlib/CategoryTheory/Monad/Monadicity.lean b/Mathlib/CategoryTheory/Monad/Monadicity.lean index dcb6b8d573feb..bd9b396629363 100644 --- a/Mathlib/CategoryTheory/Monad/Monadicity.lean +++ b/Mathlib/CategoryTheory/Monad/Monadicity.lean @@ -244,10 +244,8 @@ def createsGSplitCoequalizersOfMonadic [MonadicRightAdjoint G] ⦃A B⦄ (f g : [G.IsSplitPair f g] : CreatesColimit (parallelPair f g) G := by apply (config := {allowSynthFailures := true}) monadicCreatesColimitOfPreservesColimit -- Porting note: oddly (config := {allowSynthFailures := true}) had no effect here and below - · apply @preservesColimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ - dsimp - infer_instance - · apply @preservesColimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ + all_goals + apply @preservesColimit_of_iso_diagram _ _ _ _ _ _ _ _ _ (diagramIsoParallelPair.{v₁} _).symm ?_ dsimp infer_instance @@ -363,7 +361,7 @@ def monadicOfHasPreservesGSplitCoequalizersOfReflectsIsomorphisms [G.ReflectsIso MonadicRightAdjoint G := by have : ReflectsColimitOfIsSplitPair G := ⟨fun f g _ => by have := HasCoequalizerOfIsSplitPair.out G f g - apply reflectsColimitOfReflectsIsomorphisms⟩ + apply reflectsColimit_of_reflectsIsomorphisms⟩ apply monadicOfHasPreservesReflectsGSplitCoequalizers adj end BeckMonadicity @@ -409,7 +407,7 @@ def monadicOfHasPreservesReflexiveCoequalizersOfReflectsIsomorphisms : MonadicRi infer_instance -- Porting note: passing instances through apply @counitCoequalizerOfReflectsCoequalizer _ _ _ _ _ _ _ _ ?_ - apply reflectsColimitOfReflectsIsomorphisms + apply reflectsColimit_of_reflectsIsomorphisms exact (comparisonAdjunction adj).toEquivalence.isEquivalence_inverse end ReflexiveMonadicity diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean index 00f12f9aee66e..84cdce3c91ae3 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Limits.lean @@ -79,9 +79,9 @@ instance hasLimitsOfShape [HasLimitsOfShape J C] : HasLimitsOfShape J (Mon_ C) w { cone := limitCone F isLimit := limitConeIsLimit F } -instance forgetPreservesLimitsOfShape : PreservesLimitsOfShape J (Mon_.forget C) where +instance forget_freservesLimitsOfShape : PreservesLimitsOfShape J (Mon_.forget C) where preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (IsLimit.ofIsoLimit (limit.isLimit (F ⋙ Mon_.forget C)) (forgetMapConeLimitConeIso F).symm) end Mon_ diff --git a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean index 5be7683ef2df1..d7d052e314259 100644 --- a/Mathlib/CategoryTheory/Monoidal/Preadditive.lean +++ b/Mathlib/CategoryTheory/Monoidal/Preadditive.lean @@ -116,20 +116,20 @@ theorem sum_tensor {P Q R S : C} {J : Type*} (s : Finset J) (f : P ⟶ Q) (g : J instance (X : C) : PreservesFiniteBiproducts (tensorLeft X) where preserves {J} := { preserves := fun {f} => - { preserves := fun {b} i => isBilimitOfTotal _ (by + { preserves := fun {b} i => ⟨isBilimitOfTotal _ (by dsimp simp_rw [← id_tensorHom] simp only [← tensor_comp, Category.comp_id, ← tensor_sum, ← tensor_id, - IsBilimit.total i]) } } + IsBilimit.total i])⟩ } } instance (X : C) : PreservesFiniteBiproducts (tensorRight X) where preserves {J} := { preserves := fun {f} => - { preserves := fun {b} i => isBilimitOfTotal _ (by + { preserves := fun {b} i => ⟨isBilimitOfTotal _ (by dsimp simp_rw [← tensorHom_id] simp only [← tensor_comp, Category.comp_id, ← sum_tensor, ← tensor_id, - IsBilimit.total i]) } } + IsBilimit.total i])⟩ } } variable [HasFiniteBiproducts C] diff --git a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean index a5d134fdcfce6..4a0fd32e2e836 100644 --- a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean @@ -157,9 +157,9 @@ instance (priority := 100) preservesFiniteBiproductsOfAdditive [Additive F] : preserves := { preserves := { preserves := fun hb => - isBilimitOfTotal _ (by + ⟨isBilimitOfTotal _ (by simp_rw [F.mapBicone_π, F.mapBicone_ι, ← F.map_comp] - erw [← F.map_sum, ← F.map_id, IsBilimit.total hb])} } + erw [← F.map_sum, ← F.map_id, IsBilimit.total hb])⟩ } } theorem additive_of_preservesBinaryBiproducts [HasBinaryBiproducts C] [PreservesZeroMorphisms F] [PreservesBinaryBiproducts F] : Additive F where @@ -172,7 +172,7 @@ lemma additive_of_preserves_binary_products [HasBinaryProducts C] [PreservesLimitsOfShape (Discrete WalkingPair) F] [F.PreservesZeroMorphisms] : F.Additive := by have : HasBinaryBiproducts C := HasBinaryBiproducts.of_hasBinaryProducts - have := preservesBinaryBiproductsOfPreservesBinaryProducts F + have := preservesBinaryBiproducts_of_preservesBinaryProducts F exact Functor.additive_of_preservesBinaryBiproducts F end @@ -254,9 +254,9 @@ variable [Preadditive D] [HasZeroObject C] [HasZeroObject D] [HasBinaryBiproduct section -attribute [local instance] preservesBinaryBiproductsOfPreservesBinaryProducts +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryProducts -attribute [local instance] preservesBinaryBiproductsOfPreservesBinaryCoproducts +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryCoproducts /-- Turn a left exact functor into an additive functor. -/ def AdditiveFunctor.ofLeftExact : (C ⥤ₗ D) ⥤ C ⥤+ D := diff --git a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean index 5b06f3702ea3b..5cbeb420311dc 100644 --- a/Mathlib/CategoryTheory/Preadditive/Biproducts.lean +++ b/Mathlib/CategoryTheory/Preadditive/Biproducts.lean @@ -825,40 +825,40 @@ variable {J : Type} [Fintype J] /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite products. -/ -def preservesProductOfPreservesBiproduct {f : J → C} [PreservesBiproduct f F] : +lemma preservesProduct_of_preservesBiproduct {f : J → C} [PreservesBiproduct f F] : PreservesLimit (Discrete.functor f) F where preserves hc := - IsLimit.ofIsoLimit + ⟨IsLimit.ofIsoLimit ((IsLimit.postcomposeInvEquiv (Discrete.compNatIsoDiscrete _ _) _).symm (isBilimitOfPreserves F (biconeIsBilimitOfLimitConeOfIsLimit hc)).isLimit) <| - Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ section -attribute [local instance] preservesProductOfPreservesBiproduct +attribute [local instance] preservesProduct_of_preservesBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite products. -/ -def preservesProductsOfShapeOfPreservesBiproductsOfShape [PreservesBiproductsOfShape J F] : +lemma preservesProductsOfShape_of_preservesBiproductsOfShape [PreservesBiproductsOfShape J F] : PreservesLimitsOfShape (Discrete J) F where - preservesLimit {_} := preservesLimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + preservesLimit {_} := preservesLimit_of_iso_diagram _ Discrete.natIsoFunctor.symm end /-- A functor between preadditive categories that preserves (zero morphisms and) finite products preserves finite biproducts. -/ -def preservesBiproductOfPreservesProduct {f : J → C} [PreservesLimit (Discrete.functor f) F] : +lemma preservesBiproduct_of_preservesProduct {f : J → C} [PreservesLimit (Discrete.functor f) F] : PreservesBiproduct f F where preserves {b} hb := - isBilimitOfIsLimit _ <| + ⟨isBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ((IsLimit.postcomposeHomEquiv (Discrete.compNatIsoDiscrete _ _) (F.mapCone b.toCone)).symm (isLimitOfPreserves F hb.isLimit)) <| - Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ /-- If the (product-like) biproduct comparison for `F` and `f` is a monomorphism, then `F` preserves the biproduct of `f`. For the converse, see `mapBiproduct`. -/ -def preservesBiproductOfMonoBiproductComparison {f : J → C} [HasBiproduct f] +lemma preservesBiproduct_of_mono_biproductComparison {f : J → C} [HasBiproduct f] [HasBiproduct (F.obj ∘ f)] [Mono (biproductComparison F f)] : PreservesBiproduct f F := by haveI : HasProduct fun b => F.obj (f b) := by change HasProduct (F.obj ∘ f) @@ -872,101 +872,104 @@ def preservesBiproductOfMonoBiproductComparison {f : J → C} [HasBiproduct f] haveI : IsIso (piComparison F f) := by rw [that] infer_instance - haveI := PreservesProduct.ofIsoComparison F f - apply preservesBiproductOfPreservesProduct + haveI := PreservesProduct.of_iso_comparison F f + apply preservesBiproduct_of_preservesProduct /-- If the (coproduct-like) biproduct comparison for `F` and `f` is an epimorphism, then `F` preserves the biproduct of `F` and `f`. For the converse, see `mapBiproduct`. -/ -def preservesBiproductOfEpiBiproductComparison' {f : J → C} [HasBiproduct f] +lemma preservesBiproduct_of_epi_biproductComparison' {f : J → C} [HasBiproduct f] [HasBiproduct (F.obj ∘ f)] [Epi (biproductComparison' F f)] : PreservesBiproduct f F := by haveI : Epi (splitEpiBiproductComparison F f).section_ := by simpa haveI : IsIso (biproductComparison F f) := IsIso.of_epi_section' (splitEpiBiproductComparison F f) - apply preservesBiproductOfMonoBiproductComparison + apply preservesBiproduct_of_mono_biproductComparison /-- A functor between preadditive categories that preserves (zero morphisms and) finite products preserves finite biproducts. -/ -def preservesBiproductsOfShapeOfPreservesProductsOfShape [PreservesLimitsOfShape (Discrete J) F] : +lemma preservesBiproductsOfShape_of_preservesProductsOfShape + [PreservesLimitsOfShape (Discrete J) F] : PreservesBiproductsOfShape J F where - preserves {_} := preservesBiproductOfPreservesProduct F + preserves {_} := preservesBiproduct_of_preservesProduct F /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite coproducts. -/ -def preservesCoproductOfPreservesBiproduct {f : J → C} [PreservesBiproduct f F] : +lemma preservesCoproduct_of_preservesBiproduct {f : J → C} [PreservesBiproduct f F] : PreservesColimit (Discrete.functor f) F where preserves {c} hc := - IsColimit.ofIsoColimit + ⟨IsColimit.ofIsoColimit ((IsColimit.precomposeHomEquiv (Discrete.compNatIsoDiscrete _ _) _).symm (isBilimitOfPreserves F (biconeIsBilimitOfColimitCoconeOfIsColimit hc)).isColimit) <| - Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ section -attribute [local instance] preservesCoproductOfPreservesBiproduct +attribute [local instance] preservesCoproduct_of_preservesBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts preserves finite coproducts. -/ -def preservesCoproductsOfShapeOfPreservesBiproductsOfShape [PreservesBiproductsOfShape J F] : +lemma preservesCoproductsOfShape_of_preservesBiproductsOfShape [PreservesBiproductsOfShape J F] : PreservesColimitsOfShape (Discrete J) F where - preservesColimit {_} := preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm + preservesColimit {_} := preservesColimit_of_iso_diagram _ Discrete.natIsoFunctor.symm end /-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts preserves finite biproducts. -/ -def preservesBiproductOfPreservesCoproduct {f : J → C} [PreservesColimit (Discrete.functor f) F] : +lemma preservesBiproduct_of_preservesCoproduct {f : J → C} + [PreservesColimit (Discrete.functor f) F] : PreservesBiproduct f F where preserves {b} hb := - isBilimitOfIsColimit _ <| + ⟨isBilimitOfIsColimit _ <| IsColimit.ofIsoColimit ((IsColimit.precomposeInvEquiv (Discrete.compNatIsoDiscrete _ _) (F.mapCocone b.toCocone)).symm (isColimitOfPreserves F hb.isColimit)) <| - Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp) + Cocones.ext (Iso.refl _) (by rintro ⟨⟩; simp)⟩ /-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts preserves finite biproducts. -/ -def preservesBiproductsOfShapeOfPreservesCoproductsOfShape +lemma preservesBiproductsOfShape_of_preservesCoproductsOfShape [PreservesColimitsOfShape (Discrete J) F] : PreservesBiproductsOfShape J F where - preserves {_} := preservesBiproductOfPreservesCoproduct F + preserves {_} := preservesBiproduct_of_preservesCoproduct F end Fintype /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary products. -/ -def preservesBinaryProductOfPreservesBinaryBiproduct {X Y : C} [PreservesBinaryBiproduct X Y F] : +lemma preservesBinaryProduct_of_preservesBinaryBiproduct {X Y : C} + [PreservesBinaryBiproduct X Y F] : PreservesLimit (pair X Y) F where - preserves {c} hc := IsLimit.ofIsoLimit + preserves {c} hc := ⟨IsLimit.ofIsoLimit ((IsLimit.postcomposeInvEquiv (diagramIsoPair _) _).symm (isBinaryBilimitOfPreserves F (binaryBiconeIsBilimitOfLimitConeOfIsLimit hc)).isLimit) <| Cones.ext (by dsimp; rfl) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ section -attribute [local instance] preservesBinaryProductOfPreservesBinaryBiproduct +attribute [local instance] preservesBinaryProduct_of_preservesBinaryBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary products. -/ -def preservesBinaryProductsOfPreservesBinaryBiproducts [PreservesBinaryBiproducts F] : +lemma preservesBinaryProducts_of_preservesBinaryBiproducts [PreservesBinaryBiproducts F] : PreservesLimitsOfShape (Discrete WalkingPair) F where - preservesLimit {_} := preservesLimitOfIsoDiagram _ (diagramIsoPair _).symm + preservesLimit {_} := preservesLimit_of_iso_diagram _ (diagramIsoPair _).symm end /-- A functor between preadditive categories that preserves (zero morphisms and) binary products preserves binary biproducts. -/ -def preservesBinaryBiproductOfPreservesBinaryProduct {X Y : C} [PreservesLimit (pair X Y) F] : +lemma preservesBinaryBiproduct_of_preservesBinaryProduct {X Y : C} [PreservesLimit (pair X Y) F] : PreservesBinaryBiproduct X Y F where - preserves {b} hb := isBinaryBilimitOfIsLimit _ <| IsLimit.ofIsoLimit + preserves {b} hb := ⟨isBinaryBilimitOfIsLimit _ <| IsLimit.ofIsoLimit ((IsLimit.postcomposeHomEquiv (diagramIsoPair _) (F.mapCone b.toCone)).symm (isLimitOfPreserves F hb.isLimit)) <| Cones.ext (by dsimp; rfl) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ /-- If the (product-like) biproduct comparison for `F`, `X` and `Y` is a monomorphism, then `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ -def preservesBinaryBiproductOfMonoBiprodComparison {X Y : C} [HasBinaryBiproduct X Y] +lemma preservesBinaryBiproduct_of_mono_biprodComparison {X Y : C} [HasBinaryBiproduct X Y] [HasBinaryBiproduct (F.obj X) (F.obj Y)] [Mono (biprodComparison F X Y)] : PreservesBinaryBiproduct X Y F := by have that : @@ -977,66 +980,68 @@ def preservesBinaryBiproductOfMonoBiprodComparison {X Y : C} [HasBinaryBiproduct haveI : IsIso (prodComparison F X Y) := by rw [that] infer_instance - haveI := PreservesLimitPair.ofIsoProdComparison F X Y - apply preservesBinaryBiproductOfPreservesBinaryProduct + haveI := PreservesLimitPair.of_iso_prod_comparison F X Y + apply preservesBinaryBiproduct_of_preservesBinaryProduct /-- If the (coproduct-like) biproduct comparison for `F`, `X` and `Y` is an epimorphism, then `F` preserves the biproduct of `X` and `Y`. For the converse, see `mapBiprod`. -/ -def preservesBinaryBiproductOfEpiBiprodComparison' {X Y : C} [HasBinaryBiproduct X Y] +lemma preservesBinaryBiproduct_of_epi_biprodComparison' {X Y : C} [HasBinaryBiproduct X Y] [HasBinaryBiproduct (F.obj X) (F.obj Y)] [Epi (biprodComparison' F X Y)] : PreservesBinaryBiproduct X Y F := by haveI : Epi (splitEpiBiprodComparison F X Y).section_ := by simpa haveI : IsIso (biprodComparison F X Y) := IsIso.of_epi_section' (splitEpiBiprodComparison F X Y) - apply preservesBinaryBiproductOfMonoBiprodComparison + apply preservesBinaryBiproduct_of_mono_biprodComparison /-- A functor between preadditive categories that preserves (zero morphisms and) binary products preserves binary biproducts. -/ -def preservesBinaryBiproductsOfPreservesBinaryProducts +lemma preservesBinaryBiproducts_of_preservesBinaryProducts [PreservesLimitsOfShape (Discrete WalkingPair) F] : PreservesBinaryBiproducts F where - preserves {_} {_} := preservesBinaryBiproductOfPreservesBinaryProduct F + preserves {_} {_} := preservesBinaryBiproduct_of_preservesBinaryProduct F /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary coproducts. -/ -def preservesBinaryCoproductOfPreservesBinaryBiproduct {X Y : C} [PreservesBinaryBiproduct X Y F] : +lemma preservesBinaryCoproduct_of_preservesBinaryBiproduct {X Y : C} + [PreservesBinaryBiproduct X Y F] : PreservesColimit (pair X Y) F where preserves {c} hc := - IsColimit.ofIsoColimit + ⟨IsColimit.ofIsoColimit ((IsColimit.precomposeHomEquiv (diagramIsoPair _) _).symm (isBinaryBilimitOfPreserves F (binaryBiconeIsBilimitOfColimitCoconeOfIsColimit hc)).isColimit) <| Cocones.ext (by dsimp; rfl) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ section -attribute [local instance] preservesBinaryCoproductOfPreservesBinaryBiproduct +attribute [local instance] preservesBinaryCoproduct_of_preservesBinaryBiproduct /-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts preserves binary coproducts. -/ -def preservesBinaryCoproductsOfPreservesBinaryBiproducts [PreservesBinaryBiproducts F] : +lemma preservesBinaryCoproducts_of_preservesBinaryBiproducts [PreservesBinaryBiproducts F] : PreservesColimitsOfShape (Discrete WalkingPair) F where - preservesColimit {_} := preservesColimitOfIsoDiagram _ (diagramIsoPair _).symm + preservesColimit {_} := preservesColimit_of_iso_diagram _ (diagramIsoPair _).symm end /-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts preserves binary biproducts. -/ -def preservesBinaryBiproductOfPreservesBinaryCoproduct {X Y : C} [PreservesColimit (pair X Y) F] : +lemma preservesBinaryBiproduct_of_preservesBinaryCoproduct {X Y : C} + [PreservesColimit (pair X Y) F] : PreservesBinaryBiproduct X Y F where preserves {b} hb := - isBinaryBilimitOfIsColimit _ <| + ⟨isBinaryBilimitOfIsColimit _ <| IsColimit.ofIsoColimit ((IsColimit.precomposeInvEquiv (diagramIsoPair _) (F.mapCocone b.toCocone)).symm (isColimitOfPreserves F hb.isColimit)) <| Cocones.ext (Iso.refl _) fun j => by - rcases j with ⟨⟨⟩⟩ <;> simp + rcases j with ⟨⟨⟩⟩ <;> simp⟩ /-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts preserves binary biproducts. -/ -def preservesBinaryBiproductsOfPreservesBinaryCoproducts +lemma preservesBinaryBiproducts_of_preservesBinaryCoproducts [PreservesColimitsOfShape (Discrete WalkingPair) F] : PreservesBinaryBiproducts F where - preserves {_} {_} := preservesBinaryBiproductOfPreservesBinaryCoproduct F + preserves {_} {_} := preservesBinaryBiproduct_of_preservesBinaryCoproduct F end Limits diff --git a/Mathlib/CategoryTheory/Preadditive/Injective.lean b/Mathlib/CategoryTheory/Preadditive/Injective.lean index 4ee0d211f19f3..d9ce92c4eef09 100644 --- a/Mathlib/CategoryTheory/Preadditive/Injective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Injective.lean @@ -270,7 +270,7 @@ theorem injective_of_map_injective (adj : F ⊣ G) [G.Full] [G.Faithful] (I : D) (hI : Injective (G.obj I)) : Injective I := ⟨fun {X} {Y} f g => by intro - haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjointPreservesLimits + haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjoint_preservesLimits rcases hI.factors (G.map f) (G.map g) with ⟨w,h⟩ use inv (adj.counit.app _) ≫ F.map w ≫ adj.counit.app _ exact G.map_injective (by simpa)⟩ @@ -283,7 +283,7 @@ def mapInjectivePresentation (adj : F ⊣ G) [F.PreservesMonomorphisms] (X : D) injective := adj.map_injective _ I.injective f := G.map I.f mono := by - haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjointPreservesLimits; infer_instance + haveI : PreservesLimitsOfSize.{0, 0} G := adj.rightAdjoint_preservesLimits; infer_instance /-- Given an adjunction `F ⊣ G` such that `F` preserves monomorphisms and is faithful, then any injective presentation of `F(X)` can be pulled back to an injective presentation of `X`. diff --git a/Mathlib/CategoryTheory/Preadditive/LeftExact.lean b/Mathlib/CategoryTheory/Preadditive/LeftExact.lean index 1df5546b1d7c7..750d8578b60ca 100644 --- a/Mathlib/CategoryTheory/Preadditive/LeftExact.lean +++ b/Mathlib/CategoryTheory/Preadditive/LeftExact.lean @@ -55,36 +55,38 @@ def isLimitMapConeBinaryFanOfPreservesKernels {X Y Z : C} (π₁ : Z ⟶ X) (π (isLimitMapConeForkEquiv' F bc.inl_snd (isLimitOfPreserves F hf))).isLimit /-- A kernel preserving functor between preadditive categories preserves any pair being a limit. -/ -def preservesBinaryProductOfPreservesKernels +lemma preservesBinaryProduct_of_preservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] {X Y : C} : PreservesLimit (pair X Y) F where preserves {c} hc := - IsLimit.ofIsoLimit + ⟨IsLimit.ofIsoLimit (isLimitMapConeBinaryFanOfPreservesKernels F _ _ (IsLimit.ofIsoLimit hc (isoBinaryFanMk c))) - ((Cones.functoriality _ F).mapIso (isoBinaryFanMk c).symm) + ((Cones.functoriality _ F).mapIso (isoBinaryFanMk c).symm)⟩ -attribute [local instance] preservesBinaryProductOfPreservesKernels +attribute [local instance] preservesBinaryProduct_of_preservesKernels /-- A kernel preserving functor between preadditive categories preserves binary products. -/ -def preservesBinaryProductsOfPreservesKernels +lemma preservesBinaryProducts_of_preservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] : PreservesLimitsOfShape (Discrete WalkingPair) F where - preservesLimit := preservesLimitOfIsoDiagram F (diagramIsoPair _).symm + preservesLimit := preservesLimit_of_iso_diagram F (diagramIsoPair _).symm -attribute [local instance] preservesBinaryProductsOfPreservesKernels +attribute [local instance] preservesBinaryProducts_of_preservesKernels variable [HasBinaryBiproducts C] /-- A functor between preadditive categories preserves the equalizer of two morphisms if it preserves all kernels. -/ -def preservesEqualizerOfPreservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] +lemma preservesEqualizer_of_preservesKernels + [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] {X Y : C} (f g : X ⟶ Y) : PreservesLimit (parallelPair f g) F := by - letI := preservesBinaryBiproductsOfPreservesBinaryProducts F + letI := preservesBinaryBiproducts_of_preservesBinaryProducts F haveI := additive_of_preservesBinaryBiproducts F constructor; intro c i let c' := isLimitKernelForkOfFork (i.ofIsoLimit (Fork.isoForkOfι c)) dsimp only [kernelForkOfFork_ofι] at c' let iFc := isLimitForkMapOfIsLimit' F _ c' + constructor apply IsLimit.ofIsoLimit _ ((Cones.functoriality _ F).mapIso (Fork.isoForkOfι c).symm) apply (isLimitMapConeForkEquiv F (Fork.condition c)).invFun let p : parallelPair (F.map (f - g)) 0 ≅ parallelPair (F.map f - F.map g) 0 := @@ -96,24 +98,24 @@ def preservesEqualizerOfPreservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimi /-- A functor between preadditive categories preserves all equalizers if it preserves all kernels. -/ -def preservesEqualizersOfPreservesKernels +lemma preservesEqualizers_of_preservesKernels [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] : PreservesLimitsOfShape WalkingParallelPair F where preservesLimit {K} := by - letI := preservesEqualizerOfPreservesKernels F (K.map WalkingParallelPairHom.left) + letI := preservesEqualizer_of_preservesKernels F (K.map WalkingParallelPairHom.left) (K.map WalkingParallelPairHom.right) - apply preservesLimitOfIsoDiagram F (diagramIsoParallelPair K).symm + apply preservesLimit_of_iso_diagram F (diagramIsoParallelPair K).symm /-- A functor between preadditive categories which preserves kernels preserves all finite limits. -/ -def preservesFiniteLimitsOfPreservesKernels [HasFiniteProducts C] [HasEqualizers C] +lemma preservesFiniteLimits_of_preservesKernels [HasFiniteProducts C] [HasEqualizers C] [HasZeroObject C] [HasZeroObject D] [∀ {X Y} (f : X ⟶ Y), PreservesLimit (parallelPair f 0) F] : PreservesFiniteLimits F := by - letI := preservesEqualizersOfPreservesKernels F - letI := preservesTerminalObjectOfPreservesZeroMorphisms F - letI := preservesLimitsOfShapePemptyOfPreservesTerminal F - letI : PreservesFiniteProducts F := ⟨preservesFiniteProductsOfPreservesBinaryAndTerminal F⟩ - exact preservesFiniteLimitsOfPreservesEqualizersAndFiniteProducts F + letI := preservesEqualizers_of_preservesKernels F + letI := preservesTerminalObject_of_preservesZeroMorphisms F + letI := preservesLimitsOfShape_pempty_of_preservesTerminal F + letI : PreservesFiniteProducts F := ⟨preservesFiniteProducts_of_preserves_binary_and_terminal F⟩ + exact preservesFiniteLimits_of_preservesEqualizers_and_finiteProducts F end FiniteLimits @@ -134,39 +136,40 @@ def isColimitMapCoconeBinaryCofanOfPreservesCokernels {X Y Z : C} (ι₁ : X ⟶ /-- A cokernel preserving functor between preadditive categories preserves any pair being a colimit. -/ -def preservesCoproductOfPreservesCokernels +lemma preservesCoproduct_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] {X Y : C} : PreservesColimit (pair X Y) F where preserves {c} hc := - IsColimit.ofIsoColimit + ⟨IsColimit.ofIsoColimit (isColimitMapCoconeBinaryCofanOfPreservesCokernels F _ _ (IsColimit.ofIsoColimit hc (isoBinaryCofanMk c))) - ((Cocones.functoriality _ F).mapIso (isoBinaryCofanMk c).symm) + ((Cocones.functoriality _ F).mapIso (isoBinaryCofanMk c).symm)⟩ -attribute [local instance] preservesCoproductOfPreservesCokernels +attribute [local instance] preservesCoproduct_of_preservesCokernels /-- A cokernel preserving functor between preadditive categories preserves binary coproducts. -/ -def preservesBinaryCoproductsOfPreservesCokernels +lemma preservesBinaryCoproducts_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] : PreservesColimitsOfShape (Discrete WalkingPair) F where - preservesColimit := preservesColimitOfIsoDiagram F (diagramIsoPair _).symm + preservesColimit := preservesColimit_of_iso_diagram F (diagramIsoPair _).symm -attribute [local instance] preservesBinaryCoproductsOfPreservesCokernels +attribute [local instance] preservesBinaryCoproducts_of_preservesCokernels variable [HasBinaryBiproducts C] /-- A functor between preadditive categories preserves the coequalizer of two morphisms if it preserves all cokernels. -/ -def preservesCoequalizerOfPreservesCokernels +lemma preservesCoequalizer_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] {X Y : C} (f g : X ⟶ Y) : PreservesColimit (parallelPair f g) F := by - letI := preservesBinaryBiproductsOfPreservesBinaryCoproducts F + letI := preservesBinaryBiproducts_of_preservesBinaryCoproducts F haveI := additive_of_preservesBinaryBiproducts F constructor intro c i let c' := isColimitCokernelCoforkOfCofork (i.ofIsoColimit (Cofork.isoCoforkOfπ c)) dsimp only [cokernelCoforkOfCofork_ofπ] at c' let iFc := isColimitCoforkMapOfIsColimit' F _ c' + constructor apply IsColimit.ofIsoColimit _ ((Cocones.functoriality _ F).mapIso (Cofork.isoCoforkOfπ c).symm) apply (isColimitMapCoconeCoforkEquiv F (Cofork.condition c)).invFun @@ -179,24 +182,24 @@ def preservesCoequalizerOfPreservesCokernels /-- A functor between preadditive categories preserves all coequalizers if it preserves all kernels. -/ -def preservesCoequalizersOfPreservesCokernels +lemma preservesCoequalizers_of_preservesCokernels [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] : PreservesColimitsOfShape WalkingParallelPair F where preservesColimit {K} := by - letI := preservesCoequalizerOfPreservesCokernels F (K.map Limits.WalkingParallelPairHom.left) + letI := preservesCoequalizer_of_preservesCokernels F (K.map Limits.WalkingParallelPairHom.left) (K.map Limits.WalkingParallelPairHom.right) - apply preservesColimitOfIsoDiagram F (diagramIsoParallelPair K).symm + apply preservesColimit_of_iso_diagram F (diagramIsoParallelPair K).symm /-- A functor between preadditive categories which preserves kernels preserves all finite limits. -/ -def preservesFiniteColimitsOfPreservesCokernels [HasFiniteCoproducts C] [HasCoequalizers C] +lemma preservesFiniteColimits_of_preservesCokernels [HasFiniteCoproducts C] [HasCoequalizers C] [HasZeroObject C] [HasZeroObject D] [∀ {X Y} (f : X ⟶ Y), PreservesColimit (parallelPair f 0) F] : PreservesFiniteColimits F := by - letI := preservesCoequalizersOfPreservesCokernels F - letI := preservesInitialObjectOfPreservesZeroMorphisms F - letI := preservesColimitsOfShapePemptyOfPreservesInitial F + letI := preservesCoequalizers_of_preservesCokernels F + letI := preservesInitialObject_of_preservesZeroMorphisms F + letI := preservesColimitsOfShape_pempty_of_preservesInitial F letI : PreservesFiniteCoproducts F := ⟨preservesFiniteCoproductsOfPreservesBinaryAndInitial F⟩ - exact preservesFiniteColimitsOfPreservesCoequalizersAndFiniteCoproducts F + exact preservesFiniteColimits_of_preservesCoequalizers_and_finiteCoproducts F end FiniteColimits diff --git a/Mathlib/CategoryTheory/Preadditive/Projective.lean b/Mathlib/CategoryTheory/Preadditive/Projective.lean index 813b5d697a416..b576bee9c0e8a 100644 --- a/Mathlib/CategoryTheory/Preadditive/Projective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Projective.lean @@ -199,7 +199,7 @@ theorem map_projective (adj : F ⊣ G) [G.PreservesEpimorphisms] (P : C) (hP : P theorem projective_of_map_projective (adj : F ⊣ G) [F.Full] [F.Faithful] (P : C) (hP : Projective (F.obj P)) : Projective P where factors f g _ := by - haveI := Adjunction.leftAdjointPreservesColimits.{0, 0} adj + haveI := Adjunction.leftAdjoint_preservesColimits.{0, 0} adj rcases (@hP).1 (F.map f) (F.map g) with ⟨f', hf'⟩ use adj.unit.app _ ≫ G.map f' ≫ (inv <| adj.unit.app _) exact F.map_injective (by simpa) @@ -211,7 +211,7 @@ def mapProjectivePresentation (adj : F ⊣ G) [G.PreservesEpimorphisms] (X : C) p := F.obj Y.p projective := adj.map_projective _ Y.projective f := F.map Y.f - epi := have := Adjunction.leftAdjointPreservesColimits.{0, 0} adj; inferInstance + epi := have := Adjunction.leftAdjoint_preservesColimits.{0, 0} adj; inferInstance end Adjunction diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean index 8f0743dd4b1fd..b4c0266ceba89 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Limits.lean @@ -30,21 +30,22 @@ namespace CategoryTheory variable {C : Type u} [Category.{v} C] [Preadditive C] -instance preservesLimitsPreadditiveYonedaObj (X : C) : PreservesLimits (preadditiveYonedaObj X) := +instance preservesLimits_preadditiveYonedaObj (X : C) : PreservesLimits (preadditiveYonedaObj X) := have : PreservesLimits (preadditiveYonedaObj X ⋙ forget _) := (inferInstance : PreservesLimits (yoneda.obj X)) - preservesLimitsOfReflectsOfPreserves _ (forget _) + preservesLimits_of_reflects_of_preserves _ (forget _) -instance preservesLimitsPreadditiveCoyonedaObj (X : Cᵒᵖ) : +instance preservesLimits_preadditiveCoyonedaObj (X : Cᵒᵖ) : PreservesLimits (preadditiveCoyonedaObj X) := have : PreservesLimits (preadditiveCoyonedaObj X ⋙ forget _) := (inferInstance : PreservesLimits (coyoneda.obj X)) - preservesLimitsOfReflectsOfPreserves _ (forget _) + preservesLimits_of_reflects_of_preserves _ (forget _) -instance PreservesLimitsPreadditiveYoneda.obj (X : C) : PreservesLimits (preadditiveYoneda.obj X) := +instance preservesLimits_preadditiveYoneda_obj (X : C) : + PreservesLimits (preadditiveYoneda.obj X) := show PreservesLimits (preadditiveYonedaObj X ⋙ forget₂ _ _) from inferInstance -instance PreservesLimitsPreadditiveCoyoneda.obj (X : Cᵒᵖ) : +instance preservesLimits_preadditiveCoyoneda_obj (X : Cᵒᵖ) : PreservesLimits (preadditiveCoyoneda.obj X) := show PreservesLimits (preadditiveCoyonedaObj X ⋙ forget₂ _ _) from inferInstance diff --git a/Mathlib/CategoryTheory/Sites/Abelian.lean b/Mathlib/CategoryTheory/Sites/Abelian.lean index a4d7e152b10a4..9c89454d03a37 100644 --- a/Mathlib/CategoryTheory/Sites/Abelian.lean +++ b/Mathlib/CategoryTheory/Sites/Abelian.lean @@ -37,7 +37,7 @@ instance sheafIsAbelian : Abelian (Sheaf J D) := let adj := sheafificationAdjunction J D abelianOfAdjunction _ _ (asIso adj.counit) adj -attribute [local instance] preservesBinaryBiproductsOfPreservesBinaryProducts +attribute [local instance] preservesBinaryBiproducts_of_preservesBinaryProducts instance presheafToSheaf_additive : (presheafToSheaf J D).Additive := (presheafToSheaf J D).additive_of_preservesBinaryBiproducts diff --git a/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean b/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean index c4d16b62e04d7..e21fd44b0c627 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/ExtensiveSheaves.lean @@ -90,10 +90,10 @@ theorem Presieve.isSheaf_iff_preservesFiniteProducts (F : Cᵒᵖ ⥤ Type w) : inferInstanceAs <| ∀ (i : α), Mono (Sigma.ι Z i) let i : K ≅ Discrete.functor (fun i ↦ op (Z i)) := Discrete.natIsoFunctor let _ : PreservesLimit (Discrete.functor (fun i ↦ op (Z i))) F := - Presieve.preservesProductOfIsSheafFor F ?_ initialIsInitial _ (coproductIsCoproduct Z) + Presieve.preservesProduct_of_isSheafFor F ?_ initialIsInitial _ (coproductIsCoproduct Z) (FinitaryExtensive.isPullback_initial_to_sigma_ι Z) (hF (Presieve.ofArrows Z (fun i ↦ Sigma.ι Z i)) ?_) - · exact preservesLimitOfIsoDiagram F i.symm + · exact preservesLimit_of_iso_diagram F i.symm · apply hF refine ⟨Empty, inferInstance, Empty.elim, IsEmpty.elim inferInstance, rfl, ⟨default,?_, ?_⟩⟩ · ext b @@ -114,24 +114,23 @@ theorem Presieve.isSheaf_iff_preservesFiniteProducts (F : Cᵒᵖ ⥤ Type w) : A presheaf on a category which is `FinitaryExtensive` is a sheaf iff it preserves finite products. -/ theorem Presheaf.isSheaf_iff_preservesFiniteProducts (F : Cᵒᵖ ⥤ D) : - IsSheaf (extensiveTopology C) F ↔ Nonempty (PreservesFiniteProducts F) := by + IsSheaf (extensiveTopology C) F ↔ PreservesFiniteProducts F := by constructor · intro h rw [IsSheaf] at h - refine ⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩⟩ + refine ⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩ + constructor apply coyonedaJointlyReflectsLimits intro ⟨E⟩ specialize h E rw [Presieve.isSheaf_iff_preservesFiniteProducts] at h have : PreservesLimit K (F.comp (coyoneda.obj ⟨E⟩)) := (h.some.preserves J).preservesLimit - change IsLimit ((F.comp (coyoneda.obj ⟨E⟩)).mapCone c) - apply this.preserves - exact hc - · intro ⟨_⟩ E + exact isLimitOfPreserves (F.comp (coyoneda.obj ⟨E⟩)) hc + · intro _ E rw [Presieve.isSheaf_iff_preservesFiniteProducts] exact ⟨inferInstance⟩ -noncomputable instance (F : Sheaf (extensiveTopology C) D) : PreservesFiniteProducts F.val := - ((Presheaf.isSheaf_iff_preservesFiniteProducts F.val).mp F.cond).some +instance (F : Sheaf (extensiveTopology C) D) : PreservesFiniteProducts F.val := + (Presheaf.isSheaf_iff_preservesFiniteProducts F.val).mp F.cond end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean index 3a79842c88fc9..bd98aeb91cd38 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean @@ -118,7 +118,7 @@ lemma regularTopology.isLocallySurjective_sheafOfTypes [Preregular C] [FinitaryP let i' : ((a : α) → (F.obj ⟨Z a⟩)) ≅ (F.obj ⟨∐ Z⟩) := (Types.productIso _).symm ≪≫ (PreservesProduct.iso F _).symm ≪≫ F.mapIso (opCoproductIsoProduct _).symm refine ⟨∐ Z, Sigma.desc π, inferInstance, i'.hom x, ?_⟩ - have := preservesLimitsOfShapeOfEquiv (Discrete.opposite α).symm G + have := preservesLimitsOfShape_of_equiv (Discrete.opposite α).symm G apply Concrete.isLimit_ext _ (isLimitOfPreserves G (coproductIsCoproduct Z).op) intro ⟨⟨a⟩⟩ simp only [Functor.comp_obj, Functor.op_obj, Discrete.functor_obj, Functor.mapCone_pt, diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index a20deb1ceff5d..e78d9e4dc5cb6 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -239,7 +239,7 @@ theorem isSheaf_coherent_iff_regular_and_extensive [Preregular C] [FinitaryPreEx theorem isSheaf_iff_preservesFiniteProducts_and_equalizerCondition [Preregular C] [FinitaryExtensive C] [h : ∀ {Y X : C} (f : Y ⟶ X) [EffectiveEpi f], HasPullback f f] : - IsSheaf (coherentTopology C) F ↔ Nonempty (PreservesFiniteProducts F) ∧ + IsSheaf (coherentTopology C) F ↔ PreservesFiniteProducts F ∧ EqualizerCondition F := by rw [isSheaf_coherent_iff_regular_and_extensive] exact and_congr (isSheaf_iff_preservesFiniteProducts _) @@ -247,12 +247,12 @@ theorem isSheaf_iff_preservesFiniteProducts_and_equalizerCondition noncomputable instance [Preregular C] [FinitaryExtensive C] (F : Sheaf (coherentTopology C) A) : PreservesFiniteProducts F.val := - ((Presheaf.isSheaf_iff_preservesFiniteProducts F.val).1 - ((Presheaf.isSheaf_coherent_iff_regular_and_extensive F.val).mp F.cond).1).some + (Presheaf.isSheaf_iff_preservesFiniteProducts F.val).1 + ((Presheaf.isSheaf_coherent_iff_regular_and_extensive F.val).mp F.cond).1 theorem isSheaf_iff_preservesFiniteProducts_of_projective [Preregular C] [FinitaryExtensive C] [∀ (X : C), Projective X] : - IsSheaf (coherentTopology C) F ↔ Nonempty (PreservesFiniteProducts F) := by + IsSheaf (coherentTopology C) F ↔ PreservesFiniteProducts F := by rw [isSheaf_coherent_iff_regular_and_extensive, and_iff_left (isSheaf_of_projective F), isSheaf_iff_preservesFiniteProducts] @@ -284,8 +284,8 @@ lemma isSheaf_coherent_of_hasPullbacks_comp [Preregular C] [FinitaryExtensive C] [h : ∀ {Y X : C} (f : Y ⟶ X) [EffectiveEpi f], HasPullback f f] [PreservesFiniteLimits s] (hF : IsSheaf (coherentTopology C) F) : IsSheaf (coherentTopology C) (F ⋙ s) := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition (h := h)] at hF ⊢ - have := hF.1.some - refine ⟨⟨inferInstance⟩, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ + have := hF.1 + refine ⟨inferInstance, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ exact isLimitForkMapOfIsLimit s _ (hF.2 π c hc).some lemma isSheaf_coherent_of_hasPullbacks_of_comp [Preregular C] [FinitaryExtensive C] @@ -293,24 +293,24 @@ lemma isSheaf_coherent_of_hasPullbacks_of_comp [Preregular C] [FinitaryExtensive [ReflectsFiniteLimits s] (hF : IsSheaf (coherentTopology C) (F ⋙ s)) : IsSheaf (coherentTopology C) F := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition (h := h)] at hF ⊢ - refine ⟨⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩⟩, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ - · exact isLimitOfReflects s ((hF.1.some.1 J).1.1 hc) - · exact isLimitOfIsLimitForkMap s _ (hF.2 π c hc).some + obtain ⟨_, hF₂⟩ := hF + refine ⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩, fun _ _ π _ c hc ↦ ⟨?_⟩⟩ + · exact ⟨isLimitOfReflects s (isLimitOfPreserves (F ⋙ s) hc)⟩ + · exact isLimitOfIsLimitForkMap s _ (hF₂ π c hc).some lemma isSheaf_coherent_of_projective_comp [Preregular C] [FinitaryExtensive C] [∀ (X : C), Projective X] [PreservesFiniteProducts s] (hF : IsSheaf (coherentTopology C) F) : IsSheaf (coherentTopology C) (F ⋙ s) := by rw [isSheaf_iff_preservesFiniteProducts_of_projective] at hF ⊢ - have := hF.some - exact ⟨inferInstance⟩ + infer_instance lemma isSheaf_coherent_of_projective_of_comp [Preregular C] [FinitaryExtensive C] [∀ (X : C), Projective X] [ReflectsFiniteProducts s] (hF : IsSheaf (coherentTopology C) (F ⋙ s)) : IsSheaf (coherentTopology C) F := by rw [isSheaf_iff_preservesFiniteProducts_of_projective] at hF ⊢ - refine ⟨⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ ?_⟩⟩⟩⟩ - exact isLimitOfReflects s ((hF.some.1 J).1.1 hc) + exact ⟨fun J _ ↦ ⟨fun {K} ↦ ⟨fun {c} hc ↦ + ⟨isLimitOfReflects s (isLimitOfPreserves (F ⋙ s) hc)⟩⟩⟩⟩ instance [Preregular C] [FinitaryExtensive C] [h : ∀ {Y X : C} (f : Y ⟶ X) [EffectiveEpi f], HasPullback f f] diff --git a/Mathlib/CategoryTheory/Sites/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Equivalence.lean index de7ec32fc5e16..6953348d2088a 100644 --- a/Mathlib/CategoryTheory/Sites/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Equivalence.lean @@ -146,7 +146,7 @@ def transportSheafificationAdjunction : transportAndSheafify J K e A ⊣ sheafTo (transportIsoSheafToPresheaf _ _ _ _) noncomputable instance : PreservesFiniteLimits <| transportAndSheafify J K e A where - preservesFiniteLimits _ := compPreservesLimitsOfShape _ _ + preservesFiniteLimits _ := comp_preservesLimitsOfShape _ _ include K e in /-- Transport `HasSheafify` along an equivalence of sites. -/ diff --git a/Mathlib/CategoryTheory/Sites/LeftExact.lean b/Mathlib/CategoryTheory/Sites/LeftExact.lean index d12346e4674e9..4e6d5411bd86f 100644 --- a/Mathlib/CategoryTheory/Sites/LeftExact.lean +++ b/Mathlib/CategoryTheory/Sites/LeftExact.lean @@ -44,55 +44,61 @@ def coneCompEvaluationOfConeCompDiagramFunctorCompEvaluation {X : C} {K : Type m dsimp [diagramNatTrans] simp only [Multiequalizer.lift_ι, Category.assoc] } +/-- Auxiliary definition for `liftToDiagramLimitObj`. -/ +def liftToDiagramLimitObjAux {X : C} {K : Type max v u} [SmallCategory K] [HasLimitsOfShape K D] + {W : (J.Cover X)ᵒᵖ} (F : K ⥤ Cᵒᵖ ⥤ D) + (E : Cone (F ⋙ J.diagramFunctor D X ⋙ (evaluation (J.Cover X)ᵒᵖ D).obj W)) + (i : (unop W).Arrow) : + E.pt ⟶ (limit F).obj (op i.Y) := + (isLimitOfPreserves ((evaluation Cᵒᵖ D).obj (op i.Y)) (limit.isLimit F)).lift + (coneCompEvaluationOfConeCompDiagramFunctorCompEvaluation.{w, v, u} i E) + +@[reassoc (attr := simp)] +lemma liftToDiagramLimitObjAux_fac {X : C} {K : Type max v u} [SmallCategory K] + [HasLimitsOfShape K D] {W : (J.Cover X)ᵒᵖ} (F : K ⥤ Cᵒᵖ ⥤ D) + (E : Cone (F ⋙ J.diagramFunctor D X ⋙ (evaluation (J.Cover X)ᵒᵖ D).obj W)) + (i : (unop W).Arrow) (k : K) : + liftToDiagramLimitObjAux F E i ≫ (limit.π F k).app (op i.Y) = E.π.app k ≫ + Multiequalizer.ι ((unop W).index (F.obj k)) i := + IsLimit.fac _ _ _ + /-- An auxiliary definition to be used in the proof of the fact that `J.diagramFunctor D X` preserves limits. -/ abbrev liftToDiagramLimitObj {X : C} {K : Type max v u} [SmallCategory K] [HasLimitsOfShape K D] {W : (J.Cover X)ᵒᵖ} (F : K ⥤ Cᵒᵖ ⥤ D) (E : Cone (F ⋙ J.diagramFunctor D X ⋙ (evaluation (J.Cover X)ᵒᵖ D).obj W)) : E.pt ⟶ (J.diagram (limit F) X).obj W := - Multiequalizer.lift ((unop W).index (limit F)) E.pt - (fun i => (isLimitOfPreserves ((evaluation Cᵒᵖ D).obj (op i.Y)) (limit.isLimit F)).lift - (coneCompEvaluationOfConeCompDiagramFunctorCompEvaluation.{w, v, u} i E)) + Multiequalizer.lift ((unop W).index (limit F)) E.pt (liftToDiagramLimitObjAux F E) (by intro i - change (_ ≫ _) ≫ _ = (_ ≫ _) ≫ _ - dsimp [evaluateCombinedCones] - erw [Category.comp_id, Category.comp_id, Category.assoc, Category.assoc, ← - (limit.lift F _).naturality, ← (limit.lift F _).naturality, ← Category.assoc, ← - Category.assoc] - congr 1 - refine limit.hom_ext (fun j => ?_) - erw [Category.assoc, Category.assoc, limit.lift_π, limit.lift_π, limit.lift_π_assoc, - limit.lift_π_assoc, Category.assoc, Category.assoc, Multiequalizer.condition] + dsimp + ext k + dsimp + simp only [Category.assoc, NatTrans.naturality, liftToDiagramLimitObjAux_fac_assoc] + erw [Multiequalizer.condition] rfl) instance preservesLimit_diagramFunctor (X : C) (K : Type max v u) [SmallCategory K] [HasLimitsOfShape K D] (F : K ⥤ Cᵒᵖ ⥤ D) : PreservesLimit F (J.diagramFunctor D X) := - preservesLimitOfEvaluation _ _ fun W => - preservesLimitOfPreservesLimitCone (limit.isLimit _) + preservesLimit_of_evaluation _ _ fun W => + preservesLimit_of_preserves_limit_cone (limit.isLimit _) { lift := fun E => liftToDiagramLimitObj.{w, v, u} F E fac := by intro E k dsimp [diagramNatTrans] refine Multiequalizer.hom_ext _ _ _ (fun a => ?_) - simp only [Multiequalizer.lift_ι, Multiequalizer.lift_ι_assoc, Category.assoc] - change (_ ≫ _) ≫ _ = _ - dsimp [evaluateCombinedCones] - erw [Category.comp_id, Category.assoc, ← NatTrans.comp_app, limit.lift_π, limit.lift_π] - rfl + simp only [Multiequalizer.lift_ι, Multiequalizer.lift_ι_assoc, Category.assoc, + liftToDiagramLimitObjAux_fac] uniq := by intro E m hm refine Multiequalizer.hom_ext _ _ _ (fun a => limit_obj_ext (fun j => ?_)) - delta liftToDiagramLimitObj - erw [Multiequalizer.lift_ι, Category.assoc] - change _ = (_ ≫ _) ≫ _ - dsimp [evaluateCombinedCones] - erw [Category.comp_id, Category.assoc, ← NatTrans.comp_app, limit.lift_π, limit.lift_π] + dsimp [liftToDiagramLimitObj] + rw [Multiequalizer.lift_ι, Category.assoc, liftToDiagramLimitObjAux_fac, ← hm, + Category.assoc] dsimp - rw [← hm] - dsimp [diagramNatTrans] - simp } + rw [limit.lift_π] + dsimp } instance preservesLimitsOfShape_diagramFunctor (X : C) (K : Type max v u) [SmallCategory K] [HasLimitsOfShape K D] : @@ -169,8 +175,8 @@ instance preservesLimitsOfShape_plusFunctor (K : Type max v u) [SmallCategory K] [FinCategory K] [HasLimitsOfShape K D] [PreservesLimitsOfShape K (forget D)] [ReflectsLimitsOfShape K (forget D)] : PreservesLimitsOfShape K (J.plusFunctor D) := by - constructor; intro F; apply preservesLimitOfEvaluation; intro X - apply preservesLimitOfPreservesLimitCone (limit.isLimit F) + constructor; intro F; apply preservesLimit_of_evaluation; intro X + apply preservesLimit_of_preserves_limit_cone (limit.isLimit F) refine ⟨fun S => liftToPlusObjLimitObj.{w, v, u} F X.unop S, ?_, ?_⟩ · intro S k apply liftToPlusObjLimitObj_fac @@ -196,21 +202,21 @@ instance preservesLimitsOfShape_plusFunctor instance preserveFiniteLimits_plusFunctor [HasFiniteLimits D] [PreservesFiniteLimits (forget D)] [(forget D).ReflectsIsomorphisms] : PreservesFiniteLimits (J.plusFunctor D) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{max v u} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{max v u} intro K _ _ - have : ReflectsLimitsOfShape K (forget D) := reflectsLimitsOfShapeOfReflectsIsomorphisms + have : ReflectsLimitsOfShape K (forget D) := reflectsLimitsOfShape_of_reflectsIsomorphisms apply preservesLimitsOfShape_plusFunctor.{w, v, u} instance preservesLimitsOfShape_sheafification (K : Type max v u) [SmallCategory K] [FinCategory K] [HasLimitsOfShape K D] [PreservesLimitsOfShape K (forget D)] [ReflectsLimitsOfShape K (forget D)] : PreservesLimitsOfShape K (J.sheafification D) := - Limits.compPreservesLimitsOfShape _ _ + Limits.comp_preservesLimitsOfShape _ _ instance preservesFiniteLimits_sheafification [HasFiniteLimits D] [PreservesFiniteLimits (forget D)] [(forget D).ReflectsIsomorphisms] : PreservesFiniteLimits (J.sheafification D) := - Limits.compPreservesFiniteLimits _ _ + Limits.comp_preservesFiniteLimits _ _ end CategoryTheory.GrothendieckTopology @@ -240,16 +246,16 @@ instance preservesLimitsOfShape_presheafToSheaf : · intro j j' show Fintype (ULift _) infer_instance - refine @preservesLimitsOfShapeOfEquiv _ _ _ _ _ _ _ _ e.symm _ (show _ from ?_) - constructor; intro F; constructor; intro S hS + refine @preservesLimitsOfShape_of_equiv _ _ _ _ _ _ _ _ e.symm _ (show _ from ?_) + constructor; intro F; constructor; intro S hS; constructor apply isLimitOfReflects (sheafToPresheaf J D) have : ReflectsLimitsOfShape (AsSmall.{max v u} (FinCategory.AsType K)) (forget D) := - reflectsLimitsOfShapeOfReflectsIsomorphisms + reflectsLimitsOfShape_of_reflectsIsomorphisms apply isLimitOfPreserves (J.sheafification D) hS instance preservesfiniteLimits_presheafToSheaf [HasFiniteLimits D] : PreservesFiniteLimits (plusPlusSheaf J D) := by - apply preservesFiniteLimitsOfPreservesFiniteLimitsOfSize.{max v u} + apply preservesFiniteLimits_of_preservesFiniteLimitsOfSize.{max v u} intros infer_instance diff --git a/Mathlib/CategoryTheory/Sites/Preserves.lean b/Mathlib/CategoryTheory/Sites/Preserves.lean index 233115606bae5..ffa55fc78899b 100644 --- a/Mathlib/CategoryTheory/Sites/Preserves.lean +++ b/Mathlib/CategoryTheory/Sites/Preserves.lean @@ -53,14 +53,15 @@ def isTerminal_of_isSheafFor_empty_presieve : IsTerminal (F.obj (op I)) := by choose t h using hF (by tauto) (by tauto) exact ⟨⟨fun _ ↦ t⟩, fun a ↦ by ext; exact h.2 _ (by tauto)⟩ +include hF in /-- If `F` is a presheaf which satisfies the sheaf condition with respect to the empty presieve on the initial object, then `F` preserves terminal objects. -/ -noncomputable -def preservesTerminalOfIsSheafForEmpty (hI : IsInitial I) : PreservesLimit (Functor.empty Cᵒᵖ) F := +lemma preservesTerminal_of_isSheaf_for_empty (hI : IsInitial I) : + PreservesLimit (Functor.empty.{0} Cᵒᵖ) F := have := hI.hasInitial - (preservesTerminalOfIso F + (preservesTerminal_of_iso F ((F.mapIso (terminalIsoIsTerminal (terminalOpOfInitial initialIsInitial)) ≪≫ (F.mapIso (initialIsoIsInitial hI).symm.op) ≪≫ (terminalIsoIsTerminal (isTerminal_of_isSheafFor_empty_presieve I F hF)).symm))) @@ -114,9 +115,8 @@ theorem isSheafFor_of_preservesProduct [PreservesLimit (Discrete.functor (fun x variable [HasInitial C] [∀ i, Mono (c.inj i)] (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) -include hd -include hF hI in +include hd hF hI in /-- The two parallel maps in the equalizer diagram for the sheaf condition corresponding to the inclusion maps in a disjoint coproduct are equal. @@ -129,34 +129,33 @@ theorem firstMap_eq_secondMap : Equalizer.Presieve.Arrows.secondMap] by_cases hi : i = j · rw [hi, Mono.right_cancellation _ _ pullback.condition] - · have := preservesTerminalOfIsSheafForEmpty F hF hI + · have := preservesTerminal_of_isSheaf_for_empty F hF hI apply_fun (F.mapIso ((hd hi).isoPullback).op ≪≫ F.mapIso (terminalIsoIsTerminal (terminalOpOfInitial initialIsInitial)).symm ≪≫ (PreservesTerminal.iso F)).hom using injective_of_mono _ ext ⟨i⟩ exact i.elim +include hc hd hF hI in /-- If `F` is a presheaf which `IsSheafFor` a presieve of arrows and the empty presieve, then it preserves the product corresponding to the presieve of arrows. -/ -noncomputable -def preservesProductOfIsSheafFor - (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) +lemma preservesProduct_of_isSheafFor (hF' : (ofArrows X c.inj).IsSheafFor F) : PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F := by have : HasCoproduct X := ⟨⟨c, hc⟩⟩ - refine @PreservesProduct.ofIsoComparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ + refine @PreservesProduct.of_iso_comparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ rw [piComparison_fac (hc := hc)] refine IsIso.comp_isIso' inferInstance ?_ rw [isIso_iff_bijective, Function.bijective_iff_existsUnique] rw [Equalizer.Presieve.Arrows.sheaf_condition, Limits.Types.type_equalizer_iff_unique] at hF' exact fun b ↦ hF' b (congr_fun (firstMap_eq_secondMap F hF hI c hd) b) -include hF hI hc in +include hc hd hF hI in theorem isSheafFor_iff_preservesProduct : (ofArrows X c.inj).IsSheafFor F ↔ Nonempty (PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F) := by - refine ⟨fun hF' ↦ ⟨preservesProductOfIsSheafFor _ hF hI c hc hd hF'⟩, fun hF' ↦ ?_⟩ + refine ⟨fun hF' ↦ ⟨preservesProduct_of_isSheafFor _ hF hI c hc hd hF'⟩, fun hF' ↦ ?_⟩ let _ := hF'.some exact isSheafFor_of_preservesProduct F c hc diff --git a/Mathlib/CategoryTheory/Sites/Pullback.lean b/Mathlib/CategoryTheory/Sites/Pullback.lean index 335206084917e..ccf9e1aa3527c 100644 --- a/Mathlib/CategoryTheory/Sites/Pullback.lean +++ b/Mathlib/CategoryTheory/Sites/Pullback.lean @@ -41,7 +41,7 @@ variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) variable [ConcreteCategory.{v₁} A] [PreservesLimits (forget A)] [HasColimits A] [HasLimits A] variable [PreservesFilteredColimits (forget A)] [(forget A).ReflectsIsomorphisms] -attribute [local instance] reflectsLimitsOfReflectsIsomorphisms +attribute [local instance] reflectsLimits_of_reflectsIsomorphisms instance {X : C} : IsCofiltered (J.Cover X) := inferInstance @@ -54,8 +54,8 @@ def Functor.sheafPullback : Sheaf J A ⥤ Sheaf K A := instance [RepresentablyFlat G] : PreservesFiniteLimits (G.sheafPullback A J K) := by have : PreservesFiniteLimits (G.op.lan ⋙ presheafToSheaf K A) := - compPreservesFiniteLimits _ _ - apply compPreservesFiniteLimits + comp_preservesFiniteLimits _ _ + apply comp_preservesFiniteLimits /-- The pullback functor is left adjoint to the pushforward functor. -/ def Functor.sheafAdjunctionContinuous [Functor.IsContinuous.{v₁} G J K] : diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index e691068952e92..9ef621b55f223 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -670,7 +670,7 @@ theorem isSheaf_iff_isSheaf' : IsSheaf J P' ↔ IsSheaf' J P' := by rw [Equalizer.Presieve.sheaf_condition] refine ⟨?_⟩ refine isSheafForIsSheafFor' _ _ _ _ ?_ - letI := preservesSmallestLimitsOfPreservesLimits (coyoneda.obj (op U)) + letI := preservesSmallestLimits_of_preservesLimits (coyoneda.obj (op U)) apply isLimitOfPreserves apply Classical.choice (h _ S.arrows _) simpa @@ -692,7 +692,7 @@ theorem isSheaf_comp_of_isSheaf (s : A ⥤ B) [PreservesLimitsOfSize.{v₁, max theorem isSheaf_iff_isSheaf_comp (s : A ⥤ B) [HasLimitsOfSize.{v₁, max v₁ u₁} A] [PreservesLimitsOfSize.{v₁, max v₁ u₁} s] [s.ReflectsIsomorphisms] : IsSheaf J P ↔ IsSheaf J (P ⋙ s) := by - letI : ReflectsLimitsOfSize s := reflectsLimitsOfReflectsIsomorphisms + letI : ReflectsLimitsOfSize s := reflectsLimits_of_reflectsIsomorphisms exact ⟨isSheaf_comp_of_isSheaf J P s, isSheaf_of_isSheaf_comp J P s⟩ /-- @@ -707,7 +707,7 @@ hold. theorem isSheaf_iff_isSheaf_forget (s : A' ⥤ Type max v₁ u₁) [HasLimits A'] [PreservesLimits s] [s.ReflectsIsomorphisms] : IsSheaf J P' ↔ IsSheaf J (P' ⋙ s) := by have : HasLimitsOfSize.{v₁, max v₁ u₁} A' := hasLimitsOfSizeShrink.{_, _, u₁, 0} A' - have : PreservesLimitsOfSize.{v₁, max v₁ u₁} s := preservesLimitsOfSizeShrink.{_, 0, _, u₁} s + have : PreservesLimitsOfSize.{v₁, max v₁ u₁} s := preservesLimitsOfSize_shrink.{_, 0, _, u₁} s apply isSheaf_iff_isSheaf_comp end Concrete diff --git a/Mathlib/CategoryTheory/Sites/Sheafification.lean b/Mathlib/CategoryTheory/Sites/Sheafification.lean index 409605999eeb7..d8da476e5182c 100644 --- a/Mathlib/CategoryTheory/Sites/Sheafification.lean +++ b/Mathlib/CategoryTheory/Sites/Sheafification.lean @@ -55,7 +55,7 @@ theorem HasSheafify.mk' {F : (Cᵒᵖ ⥤ A) ⥤ Sheaf J A} (adj : F ⊣ sheafTo isRightAdjoint := ⟨F, ⟨adj⟩⟩ isLeftExact := ⟨by have : (sheafToPresheaf J A).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ - exact ⟨fun _ _ _ ↦ preservesLimitsOfShapeOfNatIso + exact ⟨fun _ _ _ ↦ preservesLimitsOfShape_of_natIso (adj.leftAdjointUniq (Adjunction.ofIsRightAdjoint (sheafToPresheaf J A)))⟩⟩ /-- The sheafification functor, left adjoint to the inclusion. -/ diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index 9837ef05680cf..884a5b5ce77dc 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -173,8 +173,8 @@ instance (priority := 100) [F.IsTriangulated] : PreservesZeroMorphisms F where noncomputable instance [F.IsTriangulated] : PreservesLimitsOfShape (Discrete WalkingPair) F := by suffices ∀ (X₁ X₃ : C), IsIso (prodComparison F X₁ X₃) by - have := fun (X₁ X₃ : C) ↦ PreservesLimitPair.ofIsoProdComparison F X₁ X₃ - exact ⟨fun {K} ↦ preservesLimitOfIsoDiagram F (diagramIsoPair K).symm⟩ + have := fun (X₁ X₃ : C) ↦ PreservesLimitPair.of_iso_prod_comparison F X₁ X₃ + exact ⟨fun {K} ↦ preservesLimit_of_iso_diagram F (diagramIsoPair K).symm⟩ intro X₁ X₃ let φ : F.mapTriangle.obj (binaryProductTriangle X₁ X₃) ⟶ binaryProductTriangle (F.obj X₁) (F.obj X₃) := diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index fa43e61adcc47..5c1911f19648b 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -122,7 +122,7 @@ lemma mem_homologicalKernel_iff [F.IsHomological] [F.ShiftSequence ℤ] (X : C) noncomputable instance (priority := 100) [F.IsHomological] : PreservesLimitsOfShape (Discrete WalkingPair) F := by suffices ∀ (X₁ X₂ : C), PreservesLimit (pair X₁ X₂) F from - ⟨fun {X} => preservesLimitOfIsoDiagram F (diagramIsoPair X).symm⟩ + ⟨fun {X} => preservesLimit_of_iso_diagram F (diagramIsoPair X).symm⟩ intro X₁ X₂ have : HasBinaryBiproduct (F.obj X₁) (F.obj X₂) := HasBinaryBiproducts.has_binary_biproduct _ _ have : Mono (F.biprodComparison X₁ X₂) := by @@ -137,8 +137,8 @@ noncomputable instance (priority := 100) [F.IsHomological] : simp only [assoc, biprodComparison_fst, zero_comp, ← F.map_comp, biprod.inl_fst, F.map_id, comp_id] at hf rw [hf, zero_comp] - have : PreservesBinaryBiproduct X₁ X₂ F := preservesBinaryBiproductOfMonoBiprodComparison _ - apply Limits.preservesBinaryProductOfPreservesBinaryBiproduct + have : PreservesBinaryBiproduct X₁ X₂ F := preservesBinaryBiproduct_of_mono_biprodComparison _ + apply Limits.preservesBinaryProduct_of_preservesBinaryBiproduct instance (priority := 100) [F.IsHomological] : F.Additive := F.additive_of_preserves_binary_products diff --git a/Mathlib/Condensed/Discrete/Characterization.lean b/Mathlib/Condensed/Discrete/Characterization.lean index 04c5cbbfc2c17..842f3c46d0442 100644 --- a/Mathlib/Condensed/Discrete/Characterization.lean +++ b/Mathlib/Condensed/Discrete/Characterization.lean @@ -147,13 +147,13 @@ theorem isDiscrete_tfae (M : CondensedMod.{u} R) : rw [isDiscrete_iff_isDiscrete_forget, ((CondensedSet.isDiscrete_tfae _).out 0 6:)] intro S letI : PreservesFilteredColimitsOfSize.{u, u} (forget (ModuleCat R)) := - preservesFilteredColimitsOfSizeShrink.{u, u+1, u, u+1} _ + preservesFilteredColimitsOfSize_shrink.{u, u+1, u, u+1} _ exact ⟨isColimitOfPreserves (forget (ModuleCat R)) (h S).some⟩ tfae_have 1 → 7 := by intro h S rw [isDiscrete_iff_isDiscrete_forget, ((CondensedSet.isDiscrete_tfae _).out 0 6:)] at h letI : ReflectsFilteredColimitsOfSize.{u, u} (forget (ModuleCat R)) := - reflectsFilteredColimitsOfSizeShrink.{u, u+1, u, u+1} _ + reflectsFilteredColimitsOfSize_shrink.{u, u+1, u, u+1} _ exact ⟨isColimitOfReflects (forget (ModuleCat R)) (h S).some⟩ tfae_finish @@ -251,13 +251,13 @@ theorem isDiscrete_tfae (M : LightCondMod.{u} R) : rw [isDiscrete_iff_isDiscrete_forget, ((LightCondSet.isDiscrete_tfae _).out 0 5:)] intro S letI : PreservesFilteredColimitsOfSize.{0, 0} (forget (ModuleCat R)) := - preservesFilteredColimitsOfSizeShrink.{0, u, 0, u} _ + preservesFilteredColimitsOfSize_shrink.{0, u, 0, u} _ exact ⟨isColimitOfPreserves (forget (ModuleCat R)) (h S).some⟩ tfae_have 1 → 6 := by intro h S rw [isDiscrete_iff_isDiscrete_forget, ((LightCondSet.isDiscrete_tfae _).out 0 5:)] at h letI : ReflectsFilteredColimitsOfSize.{0, 0} (forget (ModuleCat R)) := - reflectsFilteredColimitsOfSizeShrink.{0, u, 0, u} _ + reflectsFilteredColimitsOfSize_shrink.{0, u, 0, u} _ exact ⟨isColimitOfReflects (forget (ModuleCat R)) (h S).some⟩ tfae_finish diff --git a/Mathlib/Condensed/Discrete/Colimit.lean b/Mathlib/Condensed/Discrete/Colimit.lean index b5f5decb0e322..b991d411c5c9a 100644 --- a/Mathlib/Condensed/Discrete/Colimit.lean +++ b/Mathlib/Condensed/Discrete/Colimit.lean @@ -202,7 +202,7 @@ noncomputable instance (X : Profinite) [Finite X] : let X' := (Countable.toSmall.{0} X).equiv_small.choose let e : X ≃ X' := (Countable.toSmall X).equiv_small.choose_spec.some have : Finite X' := .of_equiv X e - preservesLimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F + preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F /-- Auxiliary definition for `isoFinYoneda`. -/ def isoFinYonedaComponents (X : Profinite.{u}) [Finite X] : @@ -473,7 +473,7 @@ noncomputable instance (X : FintypeCat.{u}) : PreservesLimitsOfShape (Discrete X let X' := (Countable.toSmall.{0} X).equiv_small.choose let e : X ≃ X' := (Countable.toSmall X).equiv_small.choose_spec.some have : Fintype X' := Fintype.ofEquiv X e - preservesLimitsOfShapeOfEquiv (Discrete.equivalence e.symm) F + preservesLimitsOfShape_of_equiv (Discrete.equivalence e.symm) F /-- Auxiliary definition for `isoFinYoneda`. -/ def isoFinYonedaComponents (X : LightProfinite.{u}) [Finite X] : diff --git a/Mathlib/Condensed/Explicit.lean b/Mathlib/Condensed/Explicit.lean index 541c0747e15fc..2efdfa8599be7 100644 --- a/Mathlib/Condensed/Explicit.lean +++ b/Mathlib/Condensed/Explicit.lean @@ -46,7 +46,7 @@ noncomputable def ofSheafStonean val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_of_projective F] - exact ⟨⟨fun _ _ ↦ inferInstance⟩⟩ } + exact ⟨fun _ _ ↦ inferInstance⟩ } /-- The condensed object associated to a presheaf on `Stonean` whose postcomposition with the @@ -62,7 +62,7 @@ noncomputable def ofSheafForgetStonean cond := by apply isSheaf_coherent_of_projective_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_of_projective] - exact ⟨⟨fun _ _ ↦ inferInstance⟩⟩ } + exact ⟨fun _ _ ↦ inferInstance⟩ } /-- The condensed object associated to a presheaf on `Profinite` which preserves finite products and @@ -76,7 +76,7 @@ noncomputable def ofSheafProfinite val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition F] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ } + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ } /-- The condensed object associated to a presheaf on `Profinite` whose postcomposition with the @@ -93,7 +93,7 @@ noncomputable def ofSheafForgetProfinite cond := by apply isSheaf_coherent_of_hasPullbacks_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ } + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ } /-- The condensed object associated to a presheaf on `CompHaus` which preserves finite products and @@ -105,7 +105,7 @@ noncomputable def ofSheafCompHaus val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition F] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- The condensed object associated to a presheaf on `CompHaus` whose postcomposition with the @@ -119,7 +119,7 @@ noncomputable def ofSheafForgetCompHaus cond := by apply isSheaf_coherent_of_hasPullbacks_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- A condensed object satisfies the equalizer condition. -/ theorem equalizerCondition (X : Condensed A) : EqualizerCondition X.val := @@ -128,13 +128,13 @@ theorem equalizerCondition (X : Condensed A) : EqualizerCondition X.val := /-- A condensed object preserves finite products. -/ noncomputable instance (X : Condensed A) : PreservesFiniteProducts X.val := isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp - X.cond |>.1.some + X.cond |>.1 /-- A condensed object regarded as a sheaf on `Profinite` preserves finite products. -/ noncomputable instance (X : Sheaf (coherentTopology Profinite.{u}) A) : PreservesFiniteProducts X.val := isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp - X.cond |>.1.some + X.cond |>.1 /-- A condensed object regarded as a sheaf on `Profinite` satisfies the equalizer condition. -/ theorem equalizerCondition_profinite (X : Sheaf (coherentTopology Profinite.{u}) A) : @@ -144,7 +144,7 @@ theorem equalizerCondition_profinite (X : Sheaf (coherentTopology Profinite.{u}) /-- A condensed object regarded as a sheaf on `Stonean` preserves finite products. -/ noncomputable instance (X : Sheaf (coherentTopology Stonean.{u}) A) : PreservesFiniteProducts X.val := - isSheaf_iff_preservesFiniteProducts_of_projective X.val |>.mp X.cond |>.some + isSheaf_iff_preservesFiniteProducts_of_projective X.val |>.mp X.cond end Condensed diff --git a/Mathlib/Condensed/Light/Explicit.lean b/Mathlib/Condensed/Light/Explicit.lean index 837d7f33f7fab..1c5823f86695f 100644 --- a/Mathlib/Condensed/Light/Explicit.lean +++ b/Mathlib/Condensed/Light/Explicit.lean @@ -42,7 +42,7 @@ noncomputable def ofSheafLightProfinite (F : LightProfinite.{u}ᵒᵖ ⥤ A) [Pr val := F cond := by rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition F] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- The light condensed object associated to a presheaf on `LightProfinite` whose postcomposition with @@ -57,7 +57,7 @@ noncomputable def ofSheafForgetLightProfinite cond := by apply isSheaf_coherent_of_hasPullbacks_of_comp F (CategoryTheory.forget A) rw [isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - exact ⟨⟨⟨fun _ _ ↦ inferInstance⟩⟩, hF⟩ + exact ⟨⟨fun _ _ ↦ inferInstance⟩, hF⟩ /-- A light condensed object satisfies the equalizer condition. -/ theorem equalizerCondition (X : LightCondensed A) : EqualizerCondition X.val := @@ -65,7 +65,7 @@ theorem equalizerCondition (X : LightCondensed A) : EqualizerCondition X.val := /-- A light condensed object preserves finite products. -/ noncomputable instance (X : LightCondensed A) : PreservesFiniteProducts X.val := - isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp X.cond |>.1.some + isSheaf_iff_preservesFiniteProducts_and_equalizerCondition X.val |>.mp X.cond |>.1 end LightCondensed diff --git a/Mathlib/Condensed/TopComparison.lean b/Mathlib/Condensed/TopComparison.lean index 92c2ed342f48f..aea4064452d23 100644 --- a/Mathlib/Condensed/TopComparison.lean +++ b/Mathlib/Condensed/TopComparison.lean @@ -89,8 +89,8 @@ the extensive topology. -/ noncomputable instance [PreservesFiniteCoproducts G] : PreservesFiniteProducts (yonedaPresheaf G X) := - have := preservesFiniteProductsOp G - ⟨fun _ ↦ compPreservesLimitsOfShape G.op (yonedaPresheaf' X)⟩ + have := preservesFiniteProducts_op G + ⟨fun _ ↦ comp_preservesLimitsOfShape G.op (yonedaPresheaf' X)⟩ section @@ -109,7 +109,7 @@ def TopCat.toSheafCompHausLike : cond := by have := CompHausLike.preregular hs rw [Presheaf.isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] - refine ⟨⟨inferInstance⟩, ?_⟩ + refine ⟨inferInstance, ?_⟩ apply (config := { allowSynthFailures := true }) equalizerCondition_yonedaPresheaf (CompHausLike.compHausLikeToTop.{u} P) X intro Z B π he diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index ac53d680e623b..e8f22b8f1f79b 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -119,7 +119,7 @@ instance : HasCoproducts.{u} LocallyRingedSpace.{u} := fun _ => noncomputable instance (J : Type _) : PreservesColimitsOfShape (Discrete.{u} J) forgetToSheafedSpace.{u} := ⟨fun {G} => - preservesColimitOfPreservesColimitCocone (coproductCofanIsColimit G) + preservesColimit_of_preserves_colimit_cocone (coproductCofanIsColimit G) ((colimit.isColimit (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) _).ofIsoColimit (Cocones.ext (Iso.refl _) fun _ => Category.comp_id _))⟩ @@ -320,8 +320,8 @@ noncomputable instance preservesCoequalizer : -- of colimit is provided later suffices PreservesColimit (parallelPair (F.map WalkingParallelPairHom.left) (F.map WalkingParallelPairHom.right)) forgetToSheafedSpace from - preservesColimitOfIsoDiagram _ (diagramIsoParallelPair F).symm - apply preservesColimitOfPreservesColimitCocone (coequalizerCoforkIsColimit _ _) + preservesColimit_of_iso_diagram _ (diagramIsoParallelPair F).symm + apply preservesColimit_of_preserves_colimit_cocone (coequalizerCoforkIsColimit _ _) apply (isColimitMapCoconeCoforkEquiv _ _).symm _ dsimp only [forgetToSheafedSpace] exact coequalizerIsCoequalizer _ _⟩ @@ -333,7 +333,7 @@ instance : HasColimits LocallyRingedSpace := noncomputable instance preservesColimits_forgetToSheafedSpace : PreservesColimits LocallyRingedSpace.forgetToSheafedSpace.{u} := - preservesColimitsOfPreservesCoequalizersAndCoproducts _ + preservesColimits_of_preservesCoequalizers_and_coproducts _ end LocallyRingedSpace diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index b937022700288..d2c8e5bec1786 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -456,8 +456,8 @@ instance pullbackToBaseIsOpenImmersion [IsOpenImmersion g] : rw [← limit.w (cospan f g) WalkingCospan.Hom.inl, cospan_map_inl] infer_instance -instance forgetPreservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := - preservesLimitOfPreservesLimitCone (pullbackConeOfLeftIsLimit f g) +instance forget_preservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := + preservesLimit_of_preserves_limit_cone (pullbackConeOfLeftIsLimit f g) (by apply (IsLimit.postcomposeHomEquiv (diagramIsoCospan _) _).toFun refine (IsLimit.equivIsoLimit ?_).toFun (limit.isLimit (cospan f.base g.base)) @@ -471,8 +471,8 @@ instance forgetPreservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := · exact Category.comp_id _ · exact Category.comp_id _) -instance forgetPreservesLimitsOfRight : PreservesLimit (cospan g f) (forget C) := - preservesPullbackSymmetry (forget C) f g +instance forget_preservesLimitsOfRight : PreservesLimit (cospan g f) (forget C) := + preservesPullback_symmetry (forget C) f g theorem pullback_snd_isIso_of_range_subset (H : Set.range g.base ⊆ Set.range f.base) : IsIso (pullback.snd f g) := by @@ -665,18 +665,20 @@ instance forgetCreatesPullbackOfRight : CreatesLimit (cospan g f) forget := (eqToIso (show pullback _ _ = pullback _ _ by congr) ≪≫ HasLimit.isoOfNatIso (diagramIsoCospan _).symm) -instance sheafedSpaceForgetPreservesOfLeft : PreservesLimit (cospan f g) (SheafedSpace.forget C) := - @Limits.compPreservesLimit _ _ _ _ _ _ (cospan f g) _ _ forget (PresheafedSpace.forget C) +instance sheafedSpace_forgetPreserves_of_left : + PreservesLimit (cospan f g) (SheafedSpace.forget C) := + @Limits.comp_preservesLimit _ _ _ _ _ _ (cospan f g) _ _ forget (PresheafedSpace.forget C) inferInstance <| by have : PreservesLimit (cospan ((cospan f g ⋙ forget).map Hom.inl) ((cospan f g ⋙ forget).map Hom.inr)) (PresheafedSpace.forget C) := by dsimp infer_instance - apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm + apply preservesLimit_of_iso_diagram _ (diagramIsoCospan _).symm -instance sheafedSpaceForgetPreservesOfRight : PreservesLimit (cospan g f) (SheafedSpace.forget C) := - preservesPullbackSymmetry _ _ _ +instance sheafedSpace_forgetPreserves_of_right : + PreservesLimit (cospan g f) (SheafedSpace.forget C) := + preservesPullback_symmetry _ _ _ instance sheafedSpace_hasPullback_of_left : HasPullback f g := hasLimit_of_created (cospan f g) forget @@ -1021,17 +1023,17 @@ instance pullback_to_base_isOpenImmersion [LocallyRingedSpace.IsOpenImmersion g] rw [← limit.w (cospan f g) WalkingCospan.Hom.inl, cospan_map_inl] infer_instance -instance forgetPreservesPullbackOfLeft : +instance forget_preservesPullbackOfLeft : PreservesLimit (cospan f g) LocallyRingedSpace.forgetToSheafedSpace := - preservesLimitOfPreservesLimitCone (pullbackConeOfLeftIsLimit f g) <| by + preservesLimit_of_preserves_limit_cone (pullbackConeOfLeftIsLimit f g) <| by apply (isLimitMapConePullbackConeEquiv _ _).symm.toFun apply isLimitOfIsLimitPullbackConeMap SheafedSpace.forgetToPresheafedSpace exact PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftIsLimit f.1 g.1 -instance forgetToPresheafedSpacePreservesPullbackOfLeft : +instance forgetToPresheafedSpace_preservesPullback_of_left : PreservesLimit (cospan f g) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - preservesLimitOfPreservesLimitCone (pullbackConeOfLeftIsLimit f g) <| by + preservesLimit_of_preserves_limit_cone (pullbackConeOfLeftIsLimit f g) <| by apply (isLimitMapConePullbackConeEquiv _ _).symm.toFun exact PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftIsLimit f.1 g.1 @@ -1040,7 +1042,7 @@ instance forgetToPresheafedSpacePreservesOpenImmersion : ((LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace).map f) := H -instance forgetToTopPreservesPullbackOfLeft : +instance forgetToTop_preservesPullback_of_left : PreservesLimit (cospan f g) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forget _) := by change PreservesLimit _ <| @@ -1056,35 +1058,35 @@ instance forgetToTopPreservesPullbackOfLeft : dsimp; infer_instance have : PreservesLimit (cospan f g ⋙ forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) (PresheafedSpace.forget CommRingCat) := by - apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm - apply Limits.compPreservesLimit + apply preservesLimit_of_iso_diagram _ (diagramIsoCospan _).symm + apply Limits.comp_preservesLimit -instance forgetReflectsPullbackOfLeft : +instance forget_reflectsPullback_of_left : ReflectsLimit (cospan f g) LocallyRingedSpace.forgetToSheafedSpace := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ -instance forgetPreservesPullbackOfRight : +instance forget_preservesPullback_of_right : PreservesLimit (cospan g f) LocallyRingedSpace.forgetToSheafedSpace := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ -instance forgetToPresheafedSpacePreservesPullbackOfRight : +instance forgetToPresheafedSpace_preservesPullback_of_right : PreservesLimit (cospan g f) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - preservesPullbackSymmetry _ _ _ + preservesPullback_symmetry _ _ _ -instance forgetReflectsPullbackOfRight : +instance forget_reflectsPullback_of_right : ReflectsLimit (cospan g f) LocallyRingedSpace.forgetToSheafedSpace := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ -instance forgetToPresheafedSpaceReflectsPullbackOfLeft : +instance forgetToPresheafedSpace_reflectsPullback_of_left : ReflectsLimit (cospan f g) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ -instance forgetToPresheafedSpaceReflectsPullbackOfRight : +instance forgetToPresheafedSpace_reflectsPullback_of_right : ReflectsLimit (cospan g f) (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) := - reflectsLimitOfReflectsIsomorphisms _ _ + reflectsLimit_of_reflectsIsomorphisms _ _ theorem pullback_snd_isIso_of_range_subset (H' : Set.range g.base ⊆ Set.range f.base) : IsIso (pullback.snd f g) := by diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index e178d24170737..4d7c3d0c180bc 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -263,7 +263,7 @@ instance : HasColimitsOfShape J (PresheafedSpace.{_, _, v} C) where has_colimit F := ⟨colimitCocone F, colimitCoconeIsColimit F⟩ instance : PreservesColimitsOfShape J (PresheafedSpace.forget.{u, v, v} C) := - ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) <| by + ⟨fun {F} => preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) <| by apply IsColimit.ofIsoColimit (colimit.isColimit _) fapply Cocones.ext · rfl @@ -278,15 +278,12 @@ instance instHasColimits [HasLimits C] : HasColimits (PresheafedSpace.{_, _, v} /-- The underlying topological space of a colimit of presheaved spaces is the colimit of the underlying topological spaces. -/ -instance forgetPreservesColimits [HasLimits C] : PreservesColimits (PresheafedSpace.forget C) where +instance forget_preservesColimits [HasLimits C] : + PreservesColimits (PresheafedSpace.forget.{_, _, v} C) where preservesColimitsOfShape {J 𝒥} := - { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) - (by apply IsColimit.ofIsoColimit (colimit.isColimit _) - fapply Cocones.ext - · rfl - · intro j - simp) } + { preservesColimit := fun {F} => preservesColimit_of_preserves_colimit_cocone + (colimitCoconeIsColimit F) + (IsColimit.ofIsoColimit (colimit.isColimit _) (Cocones.ext (Iso.refl _))) } /-- The components of the colimit of a diagram of `PresheafedSpace C` is obtained via taking componentwise limits. diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index eace03ab08bbd..1941ecd8cc471 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -206,8 +206,8 @@ noncomputable instance [HasLimits C] : instance [HasLimits C] : HasColimits.{v} (SheafedSpace C) := hasColimits_of_hasColimits_createsColimits forgetToPresheafedSpace -noncomputable instance [HasLimits C] : PreservesColimits (forget C) := - Limits.compPreservesColimits forgetToPresheafedSpace (PresheafedSpace.forget C) +noncomputable instance [HasLimits C] : PreservesColimits (forget.{_, _, v} C) := + Limits.comp_preservesColimits forgetToPresheafedSpace (PresheafedSpace.forget C) section ConcreteCategory diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index 2a8960d7da208..7c61dd710f8eb 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -265,13 +265,13 @@ def functorCategoryEquivalenceCompEvaluation : (functorCategoryEquivalence V G).functor ⋙ (evaluation _ _).obj PUnit.unit ≅ forget V G := Iso.refl _ -noncomputable instance instPreservesLimitsForget [HasLimits V] : - Limits.PreservesLimits (forget V G) := - Limits.preservesLimitsOfNatIso (Action.functorCategoryEquivalenceCompEvaluation V G) +noncomputable instance preservesLimits_forget [HasLimits V] : + PreservesLimits (forget V G) := + Limits.preservesLimits_of_natIso (Action.functorCategoryEquivalenceCompEvaluation V G) -noncomputable instance instPreservesColimitsForget [HasColimits V] : +noncomputable instance preservesColimits_forget [HasColimits V] : PreservesColimits (forget V G) := - preservesColimitsOfNatIso (Action.functorCategoryEquivalenceCompEvaluation V G) + preservesColimits_of_natIso (Action.functorCategoryEquivalenceCompEvaluation V G) -- TODO construct categorical images? end Forget diff --git a/Mathlib/RepresentationTheory/Action/Limits.lean b/Mathlib/RepresentationTheory/Action/Limits.lean index 4e1d7d252a087..104f4244fde0e 100644 --- a/Mathlib/RepresentationTheory/Action/Limits.lean +++ b/Mathlib/RepresentationTheory/Action/Limits.lean @@ -69,80 +69,80 @@ variable {C : Type t₁} [Category.{t₂} C] /-- `F : C ⥤ SingleObj G ⥤ V` preserves the limit of some `K : J ⥤ C` if it does evaluated at `SingleObj.star G`. -/ -private def SingleObj.preservesLimit (F : C ⥤ SingleObj G ⥤ V) +private lemma SingleObj.preservesLimit (F : C ⥤ SingleObj G ⥤ V) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesLimit K (F ⋙ (evaluation (SingleObj G) V).obj (SingleObj.star G))) : PreservesLimit K F := by - apply preservesLimitOfEvaluation + apply preservesLimit_of_evaluation intro _ exact h /-- `F : C ⥤ Action V G` preserves the limit of some `K : J ⥤ C` if if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesLimitOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesLimit_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesLimit K (F ⋙ Action.forget V G)) : PreservesLimit K F := by let F' : C ⥤ SingleObj G ⥤ V := F ⋙ (Action.functorCategoryEquivalence V G).functor have : PreservesLimit K F' := SingleObj.preservesLimit _ _ h - apply preservesLimitOfReflectsOfPreserves F (Action.functorCategoryEquivalence V G).functor + apply preservesLimit_of_reflects_of_preserves F (Action.functorCategoryEquivalence V G).functor /-- `F : C ⥤ Action V G` preserves limits of some shape `J` if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesLimitsOfShapeOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesLimitsOfShape_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (h : PreservesLimitsOfShape J (F ⋙ Action.forget V G)) : PreservesLimitsOfShape J F := by constructor intro K - apply Action.preservesLimitOfPreserves + apply Action.preservesLimit_of_preserves exact PreservesLimitsOfShape.preservesLimit /-- `F : C ⥤ Action V G` preserves limits of some size if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesLimitsOfSizeOfPreserves (F : C ⥤ Action V G) +lemma preservesLimitsOfSize_of_preserves (F : C ⥤ Action V G) (h : PreservesLimitsOfSize.{w₂, w₁} (F ⋙ Action.forget V G)) : PreservesLimitsOfSize.{w₂, w₁} F := by constructor intro J _ - apply Action.preservesLimitsOfShapeOfPreserves + apply Action.preservesLimitsOfShape_of_preserves exact PreservesLimitsOfSize.preservesLimitsOfShape /-- `F : C ⥤ SingleObj G ⥤ V` preserves the colimit of some `K : J ⥤ C` if it does evaluated at `SingleObj.star G`. -/ -private def SingleObj.preservesColimit (F : C ⥤ SingleObj G ⥤ V) +private lemma SingleObj.preservesColimit (F : C ⥤ SingleObj G ⥤ V) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesColimit K (F ⋙ (evaluation (SingleObj G) V).obj (SingleObj.star G))) : PreservesColimit K F := by - apply preservesColimitOfEvaluation + apply preservesColimit_of_evaluation intro _ exact h /-- `F : C ⥤ Action V G` preserves the colimit of some `K : J ⥤ C` if if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesColimitOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesColimit_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (K : J ⥤ C) (h : PreservesColimit K (F ⋙ Action.forget V G)) : PreservesColimit K F := by let F' : C ⥤ SingleObj G ⥤ V := F ⋙ (Action.functorCategoryEquivalence V G).functor have : PreservesColimit K F' := SingleObj.preservesColimit _ _ h - apply preservesColimitOfReflectsOfPreserves F (Action.functorCategoryEquivalence V G).functor + apply preservesColimit_of_reflects_of_preserves F (Action.functorCategoryEquivalence V G).functor /-- `F : C ⥤ Action V G` preserves colimits of some shape `J` if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesColimitsOfShapeOfPreserves (F : C ⥤ Action V G) {J : Type w₁} +lemma preservesColimitsOfShape_of_preserves (F : C ⥤ Action V G) {J : Type w₁} [Category.{w₂} J] (h : PreservesColimitsOfShape J (F ⋙ Action.forget V G)) : PreservesColimitsOfShape J F := by constructor intro K - apply Action.preservesColimitOfPreserves + apply Action.preservesColimit_of_preserves exact PreservesColimitsOfShape.preservesColimit /-- `F : C ⥤ Action V G` preserves colimits of some size if it does after postcomposing with the forgetful functor `Action V G ⥤ V`. -/ -noncomputable def preservesColimitsOfSizeOfPreserves (F : C ⥤ Action V G) +lemma preservesColimitsOfSize_of_preserves (F : C ⥤ Action V G) (h : PreservesColimitsOfSize.{w₂, w₁} (F ⋙ Action.forget V G)) : PreservesColimitsOfSize.{w₂, w₁} F := by constructor intro J _ - apply Action.preservesColimitsOfShapeOfPreserves + apply Action.preservesColimitsOfShape_of_preserves exact PreservesColimitsOfSize.preservesColimitsOfShape end Preservation @@ -168,7 +168,7 @@ noncomputable instance [HasFiniteLimits V] : PreservesFiniteLimits (Action.forge constructor intro _ _ _ infer_instance - apply compPreservesFiniteLimits + apply comp_preservesFiniteLimits noncomputable instance [HasFiniteColimits V] : PreservesFiniteColimits (Action.forget V G) := by show PreservesFiniteColimits ((Action.functorCategoryEquivalence V G).functor ⋙ @@ -177,24 +177,24 @@ noncomputable instance [HasFiniteColimits V] : PreservesFiniteColimits (Action.f constructor intro _ _ _ infer_instance - apply compPreservesFiniteColimits + apply comp_preservesFiniteColimits -noncomputable instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : +instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : ReflectsLimit F (Action.forget V G) where - reflects h := by + reflects h := ⟨by apply isLimitOfReflects ((Action.functorCategoryEquivalence V G).functor) - exact evaluationJointlyReflectsLimits _ (fun _ => h) + exact evaluationJointlyReflectsLimits _ (fun _ => h)⟩ -noncomputable instance {J : Type w₁} [Category.{w₂} J] : +instance {J : Type w₁} [Category.{w₂} J] : ReflectsLimitsOfShape J (Action.forget V G) where -noncomputable instance : ReflectsLimits (Action.forget V G) where +instance : ReflectsLimits (Action.forget V G) where -noncomputable instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : +instance {J : Type w₁} [Category.{w₂} J] (F : J ⥤ Action V G) : ReflectsColimit F (Action.forget V G) where - reflects h := by + reflects h := ⟨by apply isColimitOfReflects ((Action.functorCategoryEquivalence V G).functor) - exact evaluationJointlyReflectsColimits _ (fun _ => h) + exact evaluationJointlyReflectsColimits _ (fun _ => h)⟩ noncomputable instance {J : Type w₁} [Category.{w₂} J] : ReflectsColimitsOfShape J (Action.forget V G) where diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 2af958302e683..d4ddfa12b4c67 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -122,10 +122,10 @@ instance {V : Type u} [AddCommGroup V] [Module k V] (ρ : Representation k G V) -- Porting note: the two following instances were found automatically in mathlib3 noncomputable instance : PreservesLimits (forget₂ (Rep k G) (ModuleCat.{u} k)) := - Action.instPreservesLimitsForget.{u} _ _ + Action.preservesLimits_forget.{u} _ _ noncomputable instance : PreservesColimits (forget₂ (Rep k G) (ModuleCat.{u} k)) := - Action.instPreservesColimitsForget.{u} _ _ + Action.preservesColimits_forget.{u} _ _ /- Porting note: linter complains `simp` unfolds some types in the LHS, so have removed `@[simp]`. -/ diff --git a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean index 803016fe59879..2e6184e78b0ae 100644 --- a/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean +++ b/Mathlib/Topology/Algebra/Category/ProfiniteGrp/Basic.lean @@ -253,7 +253,7 @@ end instance : Limits.PreservesLimits profiniteGrpToProfinite.{u} where preservesLimitsOfShape := { - preservesLimit := fun {F} ↦ CategoryTheory.Limits.preservesLimitOfPreservesLimitCone + preservesLimit := fun {F} ↦ CategoryTheory.Limits.preservesLimit_of_preserves_limit_cone (limitConeIsLimit F) (Profinite.limitConeIsLimit (F ⋙ profiniteGrpToProfinite)) } end ProfiniteGrp diff --git a/Mathlib/Topology/Category/CompHausLike/Limits.lean b/Mathlib/Topology/Category/CompHausLike/Limits.lean index 8f8cb1165e429..7739d95e40a78 100644 --- a/Mathlib/Topology/Category/CompHausLike/Limits.lean +++ b/Mathlib/Topology/Category/CompHausLike/Limits.lean @@ -175,8 +175,8 @@ instance (P) [HasExplicitFiniteCoproducts.{0} P] : PreservesFiniteCoproducts (compHausLikeToTop P) := by refine ⟨fun J hJ ↦ ⟨fun {F} ↦ ?_⟩⟩ suffices PreservesColimit (Discrete.functor (F.obj ∘ Discrete.mk)) (compHausLikeToTop P) from - preservesColimitOfIsoDiagram _ Discrete.natIsoFunctor.symm - apply preservesColimitOfPreservesColimitCocone (CompHausLike.finiteCoproduct.isColimit _) + preservesColimit_of_iso_diagram _ Discrete.natIsoFunctor.symm + apply preservesColimit_of_preserves_colimit_cocone (CompHausLike.finiteCoproduct.isColimit _) exact TopCat.sigmaCofanIsColimit _ /-- The functor to another `CompHausLike` preserves finite coproducts if they exist. -/ @@ -185,7 +185,7 @@ noncomputable instance {P' : TopCat.{u} → Prop} PreservesFiniteCoproducts (toCompHausLike h) := by have : PreservesFiniteCoproducts (toCompHausLike h ⋙ compHausLikeToTop P') := inferInstanceAs (PreservesFiniteCoproducts (compHausLikeToTop _)) - exact preservesFiniteCoproductsOfReflectsOfPreserves (toCompHausLike h) (compHausLikeToTop P') + exact preservesFiniteCoproducts_of_reflects_of_preserves (toCompHausLike h) (compHausLikeToTop P') end FiniteCoproducts @@ -290,7 +290,7 @@ noncomputable instance : CreatesLimit (cospan f g) (compHausLikeToTop P) := by /-- The functor to `TopCat` preserves pullbacks. -/ noncomputable instance : PreservesLimit (cospan f g) (compHausLikeToTop P) := - preservesLimitOfCreatesLimitAndHasLimit _ _ + preservesLimit_of_createsLimit_and_hasLimit _ _ /-- The functor to another `CompHausLike` preserves pullbacks. -/ noncomputable instance {P' : TopCat → Prop} @@ -298,7 +298,7 @@ noncomputable instance {P' : TopCat → Prop} PreservesLimit (cospan f g) (toCompHausLike h) := by have : PreservesLimit (cospan f g) (toCompHausLike h ⋙ compHausLikeToTop P') := inferInstanceAs (PreservesLimit _ (compHausLikeToTop _)) - exact preservesLimitOfReflectsOfPreserves (toCompHausLike h) (compHausLikeToTop P') + exact preservesLimit_of_reflects_of_preserves (toCompHausLike h) (compHausLikeToTop P') variable (P) in /-- diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 3e60e9a5ca8ad..362a23b4df054 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -175,8 +175,8 @@ noncomputable instance createsCountableLimits {J : Type v} [SmallCategory J] [Co instance : HasCountableLimits LightProfinite where out _ := { has_limit := fun F ↦ ⟨limitCone F, limitConeIsLimit F⟩ } -noncomputable instance : PreservesLimitsOfShape ℕᵒᵖ (forget LightProfinite.{u}) := - have : PreservesLimitsOfSize.{0, 0} (forget Profinite.{u}) := preservesLimitsOfSizeShrink _ +instance : PreservesLimitsOfShape ℕᵒᵖ (forget LightProfinite.{u}) := + have : PreservesLimitsOfSize.{0, 0} (forget Profinite.{u}) := preservesLimitsOfSize_shrink _ inferInstanceAs (PreservesLimitsOfShape ℕᵒᵖ (lightToProfinite ⋙ forget Profinite)) variable {X Y : LightProfinite.{u}} (f : X ⟶ Y) diff --git a/Mathlib/Topology/Category/Profinite/Basic.lean b/Mathlib/Topology/Category/Profinite/Basic.lean index 503c4020cd505..e15cf996568bd 100644 --- a/Mathlib/Topology/Category/Profinite/Basic.lean +++ b/Mathlib/Topology/Category/Profinite/Basic.lean @@ -216,8 +216,8 @@ instance hasLimits : Limits.HasLimits Profinite := instance hasColimits : Limits.HasColimits Profinite := hasColimits_of_reflective profiniteToCompHaus -noncomputable instance forgetPreservesLimits : Limits.PreservesLimits (forget Profinite) := by - apply Limits.compPreservesLimits Profinite.toTopCat (forget TopCat) +instance forget_preservesLimits : Limits.PreservesLimits (forget Profinite) := by + apply Limits.comp_preservesLimits Profinite.toTopCat (forget TopCat) theorem epi_iff_surjective {X Y : Profinite.{u}} (f : X ⟶ Y) : Epi f ↔ Function.Surjective f := by constructor diff --git a/Mathlib/Topology/Category/Stonean/Adjunctions.lean b/Mathlib/Topology/Category/Stonean/Adjunctions.lean index dd99b65695aed..da6e4d47f184f 100644 --- a/Mathlib/Topology/Category/Stonean/Adjunctions.lean +++ b/Mathlib/Topology/Category/Stonean/Adjunctions.lean @@ -53,6 +53,6 @@ noncomputable def stoneCechAdjunction : typeToStonean ⊣ (forget Stonean) := /-- The forgetful functor from Stonean spaces, being a right adjoint, preserves limits. -/ noncomputable instance forget.preservesLimits : Limits.PreservesLimits (forget Stonean) := - rightAdjointPreservesLimits stoneCechAdjunction + rightAdjoint_preservesLimits stoneCechAdjunction end Stonean diff --git a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean index e7833bab71d89..4ed86b7b0ecfd 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Basic.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Basic.lean @@ -113,14 +113,15 @@ instance topCat_hasLimitsOfSize : HasLimitsOfSize.{v, v} TopCat.{max v u} where instance topCat_hasLimits : HasLimits TopCat.{u} := TopCat.topCat_hasLimitsOfSize.{u, u} -instance forgetPreservesLimitsOfSize : PreservesLimitsOfSize forget where +instance forget_preservesLimitsOfSize : + PreservesLimitsOfSize.{v, v} (forget : TopCat.{max v u} ⥤ _) where preservesLimitsOfShape {_} := { preservesLimit := fun {F} => - preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v,u} F) + preservesLimit_of_preserves_limit_cone (limitConeIsLimit.{v,u} F) (Types.limitConeIsLimit.{v,u} (F ⋙ forget)) } -instance forgetPreservesLimits : PreservesLimits forget := - TopCat.forgetPreservesLimitsOfSize.{u,u} +instance forget_preservesLimits : PreservesLimits (forget : TopCat.{u} ⥤ _) := + TopCat.forget_preservesLimitsOfSize.{u, u} /-- A choice of colimit cocone for a functor `F : J ⥤ TopCat`. Generally you should just use `colimit.cocone F`, unless you need the actual definition @@ -171,15 +172,15 @@ instance topCat_hasColimitsOfSize : HasColimitsOfSize.{v,v} TopCat.{max v u} whe instance topCat_hasColimits : HasColimits TopCat.{u} := TopCat.topCat_hasColimitsOfSize.{u, u} -instance forgetPreservesColimitsOfSize : - PreservesColimitsOfSize.{v, v} forget where +instance forget_preservesColimitsOfSize : + PreservesColimitsOfSize.{v, v} (forget : TopCat.{max u v} ⥤ _) where preservesColimitsOfShape := { preservesColimit := fun {F} => - preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) + preservesColimit_of_preserves_colimit_cocone (colimitCoconeIsColimit F) (Types.TypeMax.colimitCoconeIsColimit (F ⋙ forget)) } -instance forgetPreservesColimits : PreservesColimits (forget : TopCat.{u} ⥤ Type u) := - TopCat.forgetPreservesColimitsOfSize.{u, u} +instance forget_preservesColimits : PreservesColimits (forget : TopCat.{u} ⥤ Type u) := + TopCat.forget_preservesColimitsOfSize.{u, u} /-- The terminal object of `Top` is `PUnit`. -/ def isTerminalPUnit : IsTerminal (TopCat.of PUnit.{u + 1}) := diff --git a/Mathlib/Topology/Category/TopCat/Yoneda.lean b/Mathlib/Topology/Category/TopCat/Yoneda.lean index d898269d9ac23..f724695c46e33 100644 --- a/Mathlib/Topology/Category/TopCat/Yoneda.lean +++ b/Mathlib/Topology/Category/TopCat/Yoneda.lean @@ -62,9 +62,9 @@ noncomputable instance : PreservesFiniteProducts (yonedaPresheaf'.{w, w'} Y) whe preserves _ _ := { preservesLimit := fun {K} => have : ∀ {α : Type} (X : α → TopCat), PreservesLimit (Discrete.functor (fun x ↦ op (X x))) - (yonedaPresheaf'.{w, w'} Y) := fun X => @PreservesProduct.ofIsoComparison _ _ _ _ + (yonedaPresheaf'.{w, w'} Y) := fun X => @PreservesProduct.of_iso_comparison _ _ _ _ (yonedaPresheaf' Y) _ (fun x ↦ op (X x)) _ _ (by rw [piComparison_fac]; infer_instance) let i : K ≅ Discrete.functor (fun i ↦ op (unop (K.obj ⟨i⟩))) := Discrete.natIsoFunctor - preservesLimitOfIsoDiagram _ i.symm } + preservesLimit_of_iso_diagram _ i.symm } end ContinuousMap diff --git a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean index 18d81939f34e9..196cd67b22551 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean @@ -273,10 +273,10 @@ theorem isSheaf_iff_isSheafPreservesLimitPairwiseIntersections : rw [isSheaf_iff_isSheafPairwiseIntersections] constructor · intro h ι U - exact ⟨preservesLimitOfPreservesLimitCone (Pairwise.coconeIsColimit U).op (h U).some⟩ + exact ⟨preservesLimit_of_preserves_limit_cone (Pairwise.coconeIsColimit U).op (h U).some⟩ · intro h ι U haveI := (h U).some - exact ⟨PreservesLimit.preserves (Pairwise.coconeIsColimit U).op⟩ + exact ⟨isLimitOfPreserves _ (Pairwise.coconeIsColimit U).op⟩ end TopCat.Presheaf From 60a189e5bc5c59e67faa3f3445229e76298ce4dc Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Tue, 19 Nov 2024 08:32:01 +0000 Subject: [PATCH 10/90] chore: bump to nightly-2024-11-19 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 585547139ef08..80611d83b2915 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-18 +leanprover/lean4:nightly-2024-11-19 From 1aefbadf18578666efa249758f32e80fe369e560 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Tue, 19 Nov 2024 08:43:58 +0000 Subject: [PATCH 11/90] feat(Combinatorics/SimpleGraph/Matching): add `IsMatching.exists_of_disjoint_sets_of_card_eq` (#18905) This lemma supports at least two use cases: Matchings in cliques and matchings with a (sub)set of universal vertices. In preparation for Tutte's theorem. Co-authored-by: Pim Otte --- .../Combinatorics/SimpleGraph/Matching.lean | 28 +++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 14f2cbfc960a6..d52240d416332 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -134,6 +134,34 @@ lemma IsMatching.coeSubgraph {G' : Subgraph G} {M : Subgraph G'.coe} (hM : M.IsM · obtain ⟨_, hw', hvw⟩ := (coeSubgraph_adj _ _ _).mp hy rw [← hw.2 ⟨y, hw'⟩ hvw] +lemma IsMatching.exists_of_disjoint_sets_of_equiv {s t : Set V} (h : Disjoint s t) + (f : s ≃ t) (hadj : ∀ v : s, G.Adj v (f v)) : + ∃ M : Subgraph G, M.verts = s ∪ t ∧ M.IsMatching := by + use { + verts := s ∪ t + Adj := fun v w ↦ (∃ h : v ∈ s, f ⟨v, h⟩ = w) ∨ (∃ h : w ∈ s, f ⟨w, h⟩ = v) + adj_sub := by + intro v w h + obtain (⟨hv, rfl⟩ | ⟨hw, rfl⟩) := h + · exact hadj ⟨v, _⟩ + · exact (hadj ⟨w, _⟩).symm + edge_vert := by aesop } + + simp only [Subgraph.IsMatching, Set.mem_union, true_and] + intro v hv + cases' hv with hl hr + · use f ⟨v, hl⟩ + simp only [hl, exists_const, true_or, exists_true_left, true_and] + rintro y (rfl | ⟨hys, rfl⟩) + · rfl + · exact (h.ne_of_mem hl (f ⟨y, hys⟩).coe_prop rfl).elim + · use f.symm ⟨v, hr⟩ + simp only [Subtype.coe_eta, Equiv.apply_symm_apply, Subtype.coe_prop, exists_const, or_true, + true_and] + rintro y (⟨hy, rfl⟩ | ⟨hy, rfl⟩) + · exact (h.ne_of_mem hy hr rfl).elim + · simp + protected lemma IsMatching.map {G' : SimpleGraph W} {M : Subgraph G} (f : G →g G') (hf : Injective f) (hM : M.IsMatching) : (M.map f).IsMatching := by rintro _ ⟨v, hv, rfl⟩ From 8ad393d08a4074b0dff2cac9609c3bd7d78041f2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 08:43:59 +0000 Subject: [PATCH 12/90] feat(SetTheory/ZFC/Basic): results on pairs (#19144) We add a few lemmas on unordered pairs and use them to golf down the injectivity of the Kuratowski pair. --- Mathlib/SetTheory/ZFC/Basic.lean | 49 +++++++++++++++++++------------- 1 file changed, 29 insertions(+), 20 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 8d884f87ac28e..c7a0c35cdfcef 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -892,6 +892,24 @@ theorem singleton_nonempty (u : ZFSet) : ZFSet.Nonempty {u} := theorem mem_pair {x y z : ZFSet.{u}} : x ∈ ({y, z} : ZFSet) ↔ x = y ∨ x = z := by simp +@[simp] +theorem pair_eq_singleton (x : ZFSet) : {x, x} = ({x} : ZFSet) := by + ext + simp + +@[simp] +theorem pair_eq_singleton_iff {x y z : ZFSet} : ({x, y} : ZFSet) = {z} ↔ x = z ∧ y = z := by + refine ⟨fun h ↦ ?_, ?_⟩ + · rw [← mem_singleton, ← mem_singleton] + simp [← h] + · rintro ⟨rfl, rfl⟩ + exact pair_eq_singleton y + +@[simp] +theorem singleton_eq_pair_iff {x y z : ZFSet} : ({x} : ZFSet) = {y, z} ↔ x = y ∧ x = z := by + rw [eq_comm, pair_eq_singleton_iff] + simp_rw [eq_comm] + /-- `omega` is the first infinite von Neumann ordinal -/ def omega : ZFSet := mk PSet.omega @@ -1194,7 +1212,7 @@ theorem toSet_pair (x y : ZFSet.{u}) : (pair x y).toSet = {{x}, {x, y}} := by si /-- A subset of pairs `{(a, b) ∈ x × y | p a b}` -/ def pairSep (p : ZFSet.{u} → ZFSet.{u} → Prop) (x y : ZFSet.{u}) : ZFSet.{u} := - ZFSet.sep (fun z => ∃ a ∈ x, ∃ b ∈ y, z = pair a b ∧ p a b) (powerset (powerset (x ∪ y))) + (powerset (powerset (x ∪ y))).sep fun z => ∃ a ∈ x, ∃ b ∈ y, z = pair a b ∧ p a b @[simp] theorem mem_pairSep {p} {x y z : ZFSet.{u}} : @@ -1207,26 +1225,17 @@ theorem mem_pairSep {p} {x y z : ZFSet.{u}} : exact Or.inl ax · rintro (rfl | rfl) <;> [left; right] <;> assumption -theorem pair_injective : Function.Injective2 pair := fun x x' y y' H => by - have ae := ZFSet.ext_iff.1 H - simp only [pair, mem_pair] at ae - obtain rfl : x = x' := by - cases' (ae {x}).1 (by simp) with h h - · exact singleton_injective h - · have m : x' ∈ ({x} : ZFSet) := by simp [h] - rw [mem_singleton.mp m] - have he : x = y → y = y' := by +theorem pair_injective : Function.Injective2 pair := by + intro x x' y y' H + simp_rw [ZFSet.ext_iff, pair, mem_pair] at H + obtain rfl : x = x' := And.left <| by simpa [or_and_left] using (H {x}).1 (Or.inl rfl) + have he : y = x → y = y' := by rintro rfl - cases' (ae {x, y'}).2 (by simp only [eq_self_iff_true, or_true]) with xy'x xy'xx - · rw [eq_comm, ← mem_singleton, ← xy'x, mem_pair] - exact Or.inr rfl - · simpa [eq_comm] using (ZFSet.ext_iff.1 xy'xx y').1 (by simp) - obtain xyx | xyy' := (ae {x, y}).1 (by simp) - · obtain rfl := mem_singleton.mp ((ZFSet.ext_iff.1 xyx y).1 <| by simp) - simp [he rfl] - · obtain rfl | yy' := mem_pair.mp ((ZFSet.ext_iff.1 xyy' y).1 <| by simp) - · simp [he rfl] - · simp [yy'] + simpa [eq_comm] using H {y, y'} + have hx := H {x, y} + simp_rw [pair_eq_singleton_iff, true_and, or_true, true_iff] at hx + refine ⟨rfl, hx.elim he fun hy ↦ Or.elim ?_ he id⟩ + simpa using ZFSet.ext_iff.1 hy y @[simp] theorem pair_inj {x y x' y' : ZFSet} : pair x y = pair x' y' ↔ x = x' ∧ y = y' := From a5783c915dea554ddb4c2234e50b25a709fb1777 Mon Sep 17 00:00:00 2001 From: Jiang Jiedong Date: Tue, 19 Nov 2024 08:56:06 +0000 Subject: [PATCH 13/90] feat(FieldTheory/Minpoly): minpoly K x splits implies minpoly K (algebraMap K L a - x) splits (#17369) Add `minpoly K x` splits implies `minpoly K (- x)` splits and `minpoly K (algebraMap K L a - x)` splits. This is after #17093 . Co-authored-by: Jiang Jiedong <107380768+jjdishere@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/AlgebraMap.lean | 45 +++++++++++-- Mathlib/Algebra/Polynomial/Splits.lean | 76 +++++++++++++++++----- Mathlib/RingTheory/Adjoin/Field.lean | 12 ++++ 3 files changed, 112 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 6318eb32f3a95..3431db816584a 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -306,22 +306,46 @@ theorem algEquivOfCompEqX_eq_iff (p q p' q' : R[X]) theorem algEquivOfCompEqX_symm (p q : R[X]) (hpq : p.comp q = X) (hqp : q.comp p = X) : (algEquivOfCompEqX p q hpq hqp).symm = algEquivOfCompEqX q p hqp hpq := rfl +/-- The automorphism of the polynomial algebra given by `p(X) ↦ p(a * X + b)`, + with inverse `p(X) ↦ p(a⁻¹ * (X - b))`. -/ +@[simps!] +def algEquivCMulXAddC {R : Type*} [CommRing R] (a b : R) [Invertible a] : R[X] ≃ₐ[R] R[X] := + algEquivOfCompEqX (C a * X + C b) (C ⅟ a * (X - C b)) + (by simp [← C_mul, ← mul_assoc]) (by simp [← C_mul, ← mul_assoc]) + +theorem algEquivCMulXAddC_symm_eq {R : Type*} [CommRing R] (a b : R) [Invertible a] : + (algEquivCMulXAddC a b).symm = algEquivCMulXAddC (⅟ a) (- ⅟ a * b) := by + ext p : 1 + simp only [algEquivCMulXAddC_symm_apply, neg_mul, algEquivCMulXAddC_apply, map_neg, map_mul] + congr + simp [mul_add, sub_eq_add_neg] + /-- The automorphism of the polynomial algebra given by `p(X) ↦ p(X+t)`, with inverse `p(X) ↦ p(X-t)`. -/ @[simps!] -def algEquivAevalXAddC {R} [CommRing R] (t : R) : R[X] ≃ₐ[R] R[X] := +def algEquivAevalXAddC {R : Type*} [CommRing R] (t : R) : R[X] ≃ₐ[R] R[X] := algEquivOfCompEqX (X + C t) (X - C t) (by simp) (by simp) @[simp] -theorem algEquivAevalXAddC_eq_iff {R} [CommRing R] (t t' : R) : +theorem algEquivAevalXAddC_eq_iff {R : Type*} [CommRing R] (t t' : R) : algEquivAevalXAddC t = algEquivAevalXAddC t' ↔ t = t' := by simp [algEquivAevalXAddC] @[simp] -theorem algEquivAevalXAddC_symm {R} [CommRing R] (t : R) : +theorem algEquivAevalXAddC_symm {R : Type*} [CommRing R] (t : R) : (algEquivAevalXAddC t).symm = algEquivAevalXAddC (-t) := by simp [algEquivAevalXAddC, sub_eq_add_neg] +/-- The involutive automorphism of the polynomial algebra given by `p(X) ↦ p(-X)`. -/ +@[simps!] +def algEquivAevalNegX {R : Type*} [CommRing R] : R[X] ≃ₐ[R] R[X] := + algEquivOfCompEqX (-X) (-X) (by simp) (by simp) + +theorem comp_neg_X_comp_neg_X {R : Type*} [CommRing R] (p : R[X]) : + (p.comp (-X)).comp (-X) = p := by + rw [comp_assoc] + simp only [neg_comp, X_comp, neg_neg, comp_X] + theorem aeval_algHom (f : A →ₐ[R] B) (x : A) : aeval (f x) = f.comp (aeval x) := algHom_ext <| by simp only [aeval_X, AlgHom.comp_apply] @@ -570,16 +594,25 @@ lemma comp_X_add_C_eq_zero_iff : p.comp (X + C t) = 0 ↔ p = 0 := lemma comp_X_add_C_ne_zero_iff : p.comp (X + C t) ≠ 0 ↔ p ≠ 0 := comp_X_add_C_eq_zero_iff.not +lemma dvd_comp_C_mul_X_add_C_iff (p q : R[X]) (a b : R) [Invertible a] : + p ∣ q.comp (C a * X + C b) ↔ p.comp (C ⅟ a * (X - C b)) ∣ q := by + convert map_dvd_iff <| algEquivCMulXAddC a b using 2 + simp [← comp_eq_aeval, comp_assoc, ← mul_assoc, ← C_mul] + lemma dvd_comp_X_sub_C_iff (p q : R[X]) (a : R) : p ∣ q.comp (X - C a) ↔ p.comp (X + C a) ∣ q := by - convert (map_dvd_iff <| algEquivAevalXAddC a).symm using 2 - rw [C_eq_algebraMap, algEquivAevalXAddC_apply, ← comp_eq_aeval] - simp [comp_assoc] + let _ := invertibleOne (α := R) + simpa using dvd_comp_C_mul_X_add_C_iff p q 1 (-a) lemma dvd_comp_X_add_C_iff (p q : R[X]) (a : R) : p ∣ q.comp (X + C a) ↔ p.comp (X - C a) ∣ q := by simpa using dvd_comp_X_sub_C_iff p q (-a) +lemma dvd_comp_neg_X_iff (p q : R[X]) : p ∣ q.comp (-X) ↔ p.comp (-X) ∣ q := by + let _ := invertibleOne (α := R) + let _ := invertibleNeg (α := R) 1 + simpa using dvd_comp_C_mul_X_add_C_iff p q (-1) 0 + variable [IsDomain R] lemma units_coeff_zero_smul (c : R[X]ˣ) (p : R[X]) : (c : R[X]).coeff 0 • p = c * p := by diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 98d5a93538fdf..918c15651b4a4 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -139,26 +139,72 @@ theorem splits_X_pow (n : ℕ) : (X ^ n).Splits i := theorem splits_id_iff_splits {f : K[X]} : (f.map i).Splits (RingHom.id L) ↔ f.Splits i := by rw [splits_map_iff, RingHom.id_comp] -theorem Splits.comp_X_sub_C {i : L →+* F} (a : L) {f : L[X]} - (h : f.Splits i) : (f.comp (X - C a)).Splits i := by +variable {i} + +theorem Splits.comp_of_map_degree_le_one {f : K[X]} {p : K[X]} (hd : (p.map i).degree ≤ 1) + (h : f.Splits i) : (f.comp p).Splits i := by + by_cases hzero : map i (f.comp p) = 0 + · exact Or.inl hzero cases h with | inl h0 => - left - simp only [map_eq_zero] at h0 ⊢ - exact h0.symm ▸ zero_comp + exact Or.inl <| map_comp i _ _ ▸ h0.symm ▸ zero_comp | inr h => right intro g irr dvd - rw [map_comp, Polynomial.map_sub, map_X, map_C, dvd_comp_X_sub_C_iff] at dvd - have := h (irr.map (algEquivAevalXAddC _)) dvd - rw [degree_eq_natDegree irr.ne_zero] - rwa [algEquivAevalXAddC_apply, ← comp_eq_aeval, - degree_eq_natDegree (fun h => WithBot.bot_ne_one (h ▸ this)), - natDegree_comp, natDegree_X_add_C, mul_one] at this - -theorem Splits.comp_X_add_C {i : L →+* F} (a : L) {f : L[X]} - (h : f.Splits i) : (f.comp (X + C a)).Splits i := by - simpa only [map_neg, sub_neg_eq_add] using h.comp_X_sub_C (-a) + rw [map_comp] at dvd hzero + cases lt_or_eq_of_le hd with + | inl hd => + rw [eq_C_of_degree_le_zero (Nat.WithBot.lt_one_iff_le_zero.mp hd), comp_C] at dvd hzero + refine False.elim (irr.1 (isUnit_of_dvd_unit dvd ?_)) + simpa using hzero + | inr hd => + let _ := invertibleOfNonzero (leadingCoeff_ne_zero.mpr + (ne_zero_of_degree_gt (n := ⊥) (by rw [hd]; decide))) + rw [eq_X_add_C_of_degree_eq_one hd, dvd_comp_C_mul_X_add_C_iff _ _] at dvd + have := h (irr.map (algEquivCMulXAddC _ _).symm) dvd + rw [degree_eq_natDegree irr.ne_zero] + rwa [algEquivCMulXAddC_symm_apply, ← comp_eq_aeval, + degree_eq_natDegree (fun h => WithBot.bot_ne_one (h ▸ this)), + natDegree_comp, natDegree_C_mul (invertibleInvOf.ne_zero), + natDegree_X_sub_C, mul_one] at this + +theorem splits_iff_comp_splits_of_degree_eq_one {f : K[X]} {p : K[X]} (hd : (p.map i).degree = 1) : + f.Splits i ↔ (f.comp p).Splits i := by + rw [← splits_id_iff_splits, ← splits_id_iff_splits (f := f.comp p), map_comp] + refine ⟨fun h => Splits.comp_of_map_degree_le_one + (le_of_eq (map_id (R := L) ▸ hd)) h, fun h => ?_⟩ + let _ := invertibleOfNonzero (leadingCoeff_ne_zero.mpr + (ne_zero_of_degree_gt (n := ⊥) (by rw [hd]; decide))) + have : (map i f) = ((map i f).comp (map i p)).comp ((C ⅟ (map i p).leadingCoeff * + (X - C ((map i p).coeff 0)))) := by + rw [comp_assoc] + nth_rw 1 [eq_X_add_C_of_degree_eq_one hd] + simp only [coeff_map, invOf_eq_inv, mul_sub, ← C_mul, add_comp, mul_comp, C_comp, X_comp, + ← mul_assoc] + simp + refine this ▸ Splits.comp_of_map_degree_le_one ?_ h + simp [degree_C (inv_ne_zero (Invertible.ne_zero (a := (map i p).leadingCoeff)))] + +/-- +This is a weaker variant of `Splits.comp_of_map_degree_le_one`, +but its conditions are easier to check. +-/ +theorem Splits.comp_of_degree_le_one {f : K[X]} {p : K[X]} (hd : p.degree ≤ 1) + (h : f.Splits i) : (f.comp p).Splits i := + Splits.comp_of_map_degree_le_one ((degree_map_le i _).trans hd) h + +theorem Splits.comp_X_sub_C (a : K) {f : K[X]} + (h : f.Splits i) : (f.comp (X - C a)).Splits i := + Splits.comp_of_degree_le_one (degree_X_sub_C_le _) h + +theorem Splits.comp_X_add_C (a : K) {f : K[X]} + (h : f.Splits i) : (f.comp (X + C a)).Splits i := + Splits.comp_of_degree_le_one (by simpa using degree_X_sub_C_le (-a)) h + +theorem Splits.comp_neg_X {f : K[X]} (h : f.Splits i) : (f.comp (-X)).Splits i := + Splits.comp_of_degree_le_one (by simpa using degree_X_sub_C_le (0 : K)) h + +variable (i) theorem exists_root_of_splits' {f : K[X]} (hs : Splits i f) (hf0 : degree (f.map i) ≠ 0) : ∃ x, eval₂ i x f = 0 := diff --git a/Mathlib/RingTheory/Adjoin/Field.lean b/Mathlib/RingTheory/Adjoin/Field.lean index 25bf90e6eb79c..d791b712416ee 100644 --- a/Mathlib/RingTheory/Adjoin/Field.lean +++ b/Mathlib/RingTheory/Adjoin/Field.lean @@ -97,6 +97,13 @@ theorem IsIntegral.mem_range_algebraMap_of_minpoly_splits [Algebra K L] [IsScala x ∈ (algebraMap K L).range := int.mem_range_algHom_of_minpoly_splits h (IsScalarTower.toAlgHom R K L) +theorem minpoly_neg_splits [Algebra K L] {x : L} (g : (minpoly K x).Splits (algebraMap K L)) : + (minpoly K (-x)).Splits (algebraMap K L) := by + rw [minpoly.neg] + apply splits_mul _ _ g.comp_neg_X + simpa only [map_pow, map_neg, map_one] using + splits_C (algebraMap K L) ((-1) ^ (minpoly K x).natDegree) + theorem minpoly_add_algebraMap_splits [Algebra K L] {x : L} (r : K) (g : (minpoly K x).Splits (algebraMap K L)) : (minpoly K (x + algebraMap K L r)).Splits (algebraMap K L) := by @@ -112,6 +119,11 @@ theorem minpoly_algebraMap_add_splits [Algebra K L] {x : L} (r : K) (minpoly K (algebraMap K L r + x)).Splits (algebraMap K L) := by simpa only [add_comm] using minpoly_add_algebraMap_splits r g +theorem minpoly_algebraMap_sub_splits [Algebra K L] {x : L} (r : K) + (g : (minpoly K x).Splits (algebraMap K L)) : + (minpoly K (algebraMap K L r - x)).Splits (algebraMap K L) := by + simpa only [neg_sub] using minpoly_neg_splits (minpoly_sub_algebraMap_splits r g) + end variable [Algebra K M] [IsScalarTower R K M] {x : M} From 14cea7821c4cffb02af7ef2501761fb51bec6d9b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 20:09:30 +1100 Subject: [PATCH 14/90] fixes for leanprover/lean4#6905 --- Mathlib/Analysis/Calculus/FDeriv/Mul.lean | 2 +- Mathlib/Combinatorics/Enumerative/Composition.lean | 4 ++-- Mathlib/Data/Finset/Sort.lean | 4 ++-- Mathlib/Data/List/Cycle.lean | 10 +++++----- Mathlib/Data/List/NodupEquivFin.lean | 2 +- Mathlib/Data/List/Permutation.lean | 4 ++-- Mathlib/Data/Vector/Mem.lean | 2 +- .../ClassNumber/AdmissibleAbsoluteValue.lean | 2 +- 8 files changed, 15 insertions(+), 15 deletions(-) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean index e8073446374b3..6b5f7dc3f86f0 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean @@ -690,7 +690,7 @@ theorem HasFDerivWithinAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) s x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasFDerivAt_list_prod_finRange'.comp_hasFDerivWithinAt x - (hasFDerivWithinAt_pi.mpr fun i ↦ h l[i] (l.get_mem i i.isLt))) ?_ + (hasFDerivWithinAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem i.isLt))) ?_ ext m simp [← List.map_map] diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 986e40eed5b1b..20b7c32701611 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -150,7 +150,7 @@ theorem sum_blocksFun : ∑ i, c.blocksFun i = n := by conv_rhs => rw [← c.blocks_sum, ← ofFn_blocksFun, sum_ofFn] theorem blocksFun_mem_blocks (i : Fin c.length) : c.blocksFun i ∈ c.blocks := - get_mem _ _ _ + getElem_mem _ @[simp] theorem one_le_blocks {i : ℕ} (h : i ∈ c.blocks) : 1 ≤ i := @@ -158,7 +158,7 @@ theorem one_le_blocks {i : ℕ} (h : i ∈ c.blocks) : 1 ≤ i := @[simp] theorem one_le_blocks' {i : ℕ} (h : i < c.length) : 1 ≤ c.blocks[i] := - c.one_le_blocks (get_mem (blocks c) i h) + c.one_le_blocks (getElem_mem h) @[simp] theorem blocks_pos' (i : ℕ) (h : i < c.length) : 0 < c.blocks[i] := diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index 7e7f3fc62d674..486acd9c00f1b 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -113,7 +113,7 @@ theorem sorted_zero_eq_min'_aux (s : Finset α) (h : 0 < (s.sort (· ≤ ·)).le obtain ⟨i, hi⟩ : ∃ i, l.get i = s.min' H := List.mem_iff_get.1 this rw [← hi] exact (s.sort_sorted (· ≤ ·)).rel_get_of_le (Nat.zero_le i) - · have : l.get ⟨0, h⟩ ∈ s := (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l 0 h) + · have : l.get ⟨0, h⟩ ∈ s := (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.getElem_mem h) exact s.min'_le _ this theorem sorted_zero_eq_min' {s : Finset α} {h : 0 < (s.sort (· ≤ ·)).length} : @@ -130,7 +130,7 @@ theorem sorted_last_eq_max'_aux (s : Finset α) let l := s.sort (· ≤ ·) apply le_antisymm · have : l.get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ ∈ s := - (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.get_mem l _ h) + (Finset.mem_sort (α := α) (· ≤ ·)).1 (List.getElem_mem h) exact s.le_max' _ this · have : s.max' H ∈ l := (Finset.mem_sort (α := α) (· ≤ ·)).mpr (s.max'_mem H) obtain ⟨i, hi⟩ : ∃ i, l.get i = s.max' H := List.mem_iff_get.1 this diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 167f0c1116671..876e1a2130482 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -217,7 +217,7 @@ theorem prev_ne_cons_cons (y z : α) (h : x ∈ y :: z :: l) (hy : x ≠ y) (hz · rw [prev, dif_neg hy, if_neg hz] theorem next_mem (h : x ∈ l) : l.next x h ∈ l := - nextOr_mem (get_mem _ _ _) + nextOr_mem (getElem_mem _) theorem prev_mem (h : x ∈ l) : l.prev x h ∈ l := by cases' l with hd tl @@ -233,7 +233,7 @@ theorem prev_mem (h : x ∈ l) : l.prev x h ∈ l := by · exact mem_cons_of_mem _ (hl _ _) theorem next_get (l : List α) (h : Nodup l) (i : Fin l.length) : - next l (l.get i) (get_mem _ _ _) = + next l (l.get i) (getElem_mem _) = l.get ⟨(i + 1) % l.length, Nat.mod_lt _ (i.1.zero_le.trans_lt i.2)⟩ := match l, h, i with | [], _, i => by simpa using i.2 @@ -254,7 +254,7 @@ theorem next_get (l : List α) (h : Nodup l) (i : Fin l.length) : · subst hi' rw [next_getLast_cons] · simp [hi', get] - · rw [get_cons_succ]; exact get_mem _ _ _ + · rw [get_cons_succ]; exact getElem_mem _ · exact hx' · simp [getLast_eq_getElem] · exact hn.of_cons @@ -271,12 +271,12 @@ theorem next_get (l : List α) (h : Nodup l) (i : Fin l.length) : intro h have := nodup_iff_injective_get.1 hn h simp at this; simp [this] at hi' - · rw [get_cons_succ]; exact get_mem _ _ _ + · rw [get_cons_succ]; exact getElem_mem _ -- Unused variable linter incorrectly reports that `h` is unused here. set_option linter.unusedVariables false in theorem prev_get (l : List α) (h : Nodup l) (i : Fin l.length) : - prev l (l.get i) (get_mem _ _ _) = + prev l (l.get i) (getElem_mem _) = l.get ⟨(i + (l.length - 1)) % l.length, Nat.mod_lt _ i.pos⟩ := match l with | [] => by simpa using i.2 diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 65a48a57e4574..fbcf74a9eec72 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -49,7 +49,7 @@ variable [DecidableEq α] the set of elements of `l`. -/ @[simps] def getEquiv (l : List α) (H : Nodup l) : Fin (length l) ≃ { x // x ∈ l } where - toFun i := ⟨get l i, get_mem l i i.2⟩ + toFun i := ⟨get l i, getElem_mem i.2⟩ invFun x := ⟨indexOf (↑x) l, indexOf_lt_length.2 x.2⟩ left_inv i := by simp only [List.get_indexOf, eq_self_iff_true, Fin.eta, Subtype.coe_mk, H] right_inv x := by simp diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 09092d9f452e8..74a32e39c3263 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -500,13 +500,13 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : simp [hm', ← hn', hn] rcases lt_trichotomy n m with (ht | ht | ht) · suffices x ∈ bs by exact h x (hb.subset this) rfl - rw [← hx', getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hm)] + rw [← hx', getElem_insertIdx_of_lt ht] exact getElem_mem _ · simp only [ht] at hm' hn' rw [← hm'] at hn' exact H (insertIdx_injective _ _ hn') · suffices x ∈ as by exact h x (ha.subset this) rfl - rw [← hx, getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hn)] + rw [← hx, getElem_insertIdx_of_lt ht] exact getElem_mem _ lemma permutations_take_two (x y : α) (s : List α) : diff --git a/Mathlib/Data/Vector/Mem.lean b/Mathlib/Data/Vector/Mem.lean index 3eb0fe6678d82..bd52e003d0ed4 100644 --- a/Mathlib/Data/Vector/Mem.lean +++ b/Mathlib/Data/Vector/Mem.lean @@ -24,7 +24,7 @@ variable {α β : Type*} {n : ℕ} (a a' : α) @[simp] theorem get_mem (i : Fin n) (v : Vector α n) : v.get i ∈ v.toList := by rw [get_eq_get] - exact List.get_mem _ _ _ + exact List.getElem_mem _ theorem mem_iff_get (v : Vector α n) : a ∈ v.toList ↔ ∃ i, v.get i = a := by simp only [List.mem_iff_get, Fin.exists_iff, Vector.get_eq_get] diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean index bdc2a13623ee7..22fc2f2d46e5e 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean @@ -102,7 +102,7 @@ theorem exists_approx_aux (n : ℕ) (h : abv.IsAdmissible) : /-- This proof was nicer prior to https://github.com/leanprover/lean4/pull/4400. Please feel welcome to improve it, by avoiding use of `List.get` in favour of `GetElem`. -/ have : ∀ i h, t ((Finset.univ.filter fun x ↦ t x = s).toList.get ⟨i, h⟩) = s := fun i h ↦ - (Finset.mem_filter.mp (Finset.mem_toList.mp (List.get_mem _ i h))).2 + (Finset.mem_filter.mp (Finset.mem_toList.mp (List.getElem_mem h))).2 simp only [Nat.succ_eq_add_one, Finset.length_toList, List.get_eq_getElem] at this simp only [Nat.succ_eq_add_one, List.get_eq_getElem, Fin.coe_castLE] rw [this _ (Nat.lt_of_le_of_lt (Nat.le_of_lt_succ i₁.2) hs), From 08c691878eacbbe7c022e11194894f840318d608 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 19 Nov 2024 09:20:30 +0000 Subject: [PATCH 15/90] feat(KrullDimension): krullDim_eq_iSup_height_add_coheight_of_nonempty (#19155) Part of #15524. This adds a characterization of the Krull Dimenion based on the sum of height and coheight of elements. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 44 ++++++++++++++++++++++++++----- 1 file changed, 38 insertions(+), 6 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 691b858cf4c86..836303c212832 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -544,6 +544,15 @@ lemma krullDim_eq_of_orderIso (f : α ≃o β) : krullDim α = krullDim β := le_antisymm (iSup_le fun i ↦ le_sSup ⟨i.reverse, rfl⟩) <| iSup_le fun i ↦ le_sSup ⟨i.reverse, rfl⟩ +lemma height_le_krullDim (a : α) : height a ≤ krullDim α := by + have : Nonempty α := ⟨a⟩ + rw [krullDim_eq_iSup_length] + simp only [WithBot.coe_le_coe, iSup_le_iff] + exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + +lemma coheight_le_krullDim (a : α) : coheight a ≤ krullDim α := by + simpa using height_le_krullDim (α := αᵒᵈ) a + /-- The Krull dimension is the supremum of the elements' heights. @@ -560,10 +569,9 @@ lemma krullDim_eq_iSup_height_of_nonempty [Nonempty α] : krullDim α = ↑(⨆ suffices p.length ≤ ⨆ (a : α), height a by exact (WithBot.unbot'_le_iff fun _ => this).mp this apply le_iSup_of_le p.last (length_le_height_last (p := p)) - · rw [krullDim_eq_iSup_length] - simp only [WithBot.coe_le_coe, iSup_le_iff] - intro x - exact height_le fun p _ ↦ le_iSup_of_le p le_rfl + · rw [WithBot.coe_iSup (by bddDefault)] + apply iSup_le + apply height_le_krullDim /-- The Krull dimension is the supremum of the elements' coheights. @@ -576,8 +584,32 @@ version, with the coercion under the supremum. -/ lemma krullDim_eq_iSup_coheight_of_nonempty [Nonempty α] : krullDim α = ↑(⨆ (a : α), coheight a) := by - rw [← krullDim_orderDual] - exact krullDim_eq_iSup_height_of_nonempty (α := αᵒᵈ) + simpa using krullDim_eq_iSup_height_of_nonempty (α := αᵒᵈ) + +/-- +The Krull dimension is the supremum of the elements' height plus coheight. +-/ +lemma krullDim_eq_iSup_height_add_coheight_of_nonempty [Nonempty α] : + krullDim α = ↑(⨆ (a : α), height a + coheight a) := by + apply le_antisymm + · rw [krullDim_eq_iSup_height_of_nonempty, WithBot.coe_le_coe] + apply ciSup_mono (by bddDefault) (by simp) + · wlog hnottop : krullDim α < ⊤ + · simp_all + rw [krullDim_eq_iSup_length, WithBot.coe_le_coe] + apply iSup_le + intro a + have : height a < ⊤ := WithBot.coe_lt_coe.mp (lt_of_le_of_lt (height_le_krullDim a) hnottop) + have : coheight a < ⊤ := WithBot.coe_lt_coe.mp (lt_of_le_of_lt (coheight_le_krullDim a) hnottop) + cases hh : height a with + | top => simp_all + | coe n => + cases hch : coheight a with + | top => simp_all + | coe m => + obtain ⟨p₁, hlast, hlen₁⟩ := exists_series_of_height_eq_coe a hh + obtain ⟨p₂, hhead, hlen₂⟩ := exists_series_of_coheight_eq_coe a hch + apply le_iSup_of_le ((p₁.smash p₂) (by simp [*])) (by simp [*]) /-- The Krull dimension is the supremum of the elements' heights. From 01fc25a265b2d7c926cefe12201483ab2c41dd82 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 19 Nov 2024 09:31:54 +0000 Subject: [PATCH 16/90] chore(Algebra/Bilinear): deprecate duplicated injectivity lemmas (#19231) These are less general versions of results stated elsewhere. This partially deprecates @Vierkantor's https://github.com/leanprover-community/mathlib3/pull/6588. Also tidy the one caller of these lemmas. --- Mathlib/Algebra/Algebra/Bilinear.lean | 23 ++++++++++------------- Mathlib/RingTheory/Ideal/Basis.lean | 4 ++-- 2 files changed, 12 insertions(+), 15 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Bilinear.lean b/Mathlib/Algebra/Algebra/Bilinear.lean index 0ffbf39951dba..02c54d00368ef 100644 --- a/Mathlib/Algebra/Algebra/Bilinear.lean +++ b/Mathlib/Algebra/Algebra/Bilinear.lean @@ -225,22 +225,19 @@ section Ring variable {R A : Type*} [CommSemiring R] [Ring A] [Algebra R A] +@[deprecated mul_right_injective₀ (since := "2024-11-18")] theorem mulLeft_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : - Function.Injective (mulLeft R x) := by - letI : Nontrivial A := ⟨⟨x, 0, hx⟩⟩ - letI := NoZeroDivisors.to_isDomain A - exact mul_right_injective₀ hx + Function.Injective (mulLeft R x) := + mul_right_injective₀ hx +@[deprecated mul_left_injective₀ (since := "2024-11-18")] theorem mulRight_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : - Function.Injective (mulRight R x) := by - letI : Nontrivial A := ⟨⟨x, 0, hx⟩⟩ - letI := NoZeroDivisors.to_isDomain A - exact mul_left_injective₀ hx - -theorem mul_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : Function.Injective (mul R A x) := by - letI : Nontrivial A := ⟨⟨x, 0, hx⟩⟩ - letI := NoZeroDivisors.to_isDomain A - exact mul_right_injective₀ hx + Function.Injective (mulRight R x) := + mul_left_injective₀ hx + +@[deprecated mul_right_injective₀ (since := "2024-11-18")] +theorem mul_injective [NoZeroDivisors A] {x : A} (hx : x ≠ 0) : Function.Injective (mul R A x) := + mul_right_injective₀ hx end Ring diff --git a/Mathlib/RingTheory/Ideal/Basis.lean b/Mathlib/RingTheory/Ideal/Basis.lean index 4453148c5170a..ac26138204089 100644 --- a/Mathlib/RingTheory/Ideal/Basis.lean +++ b/Mathlib/RingTheory/Ideal/Basis.lean @@ -21,7 +21,7 @@ variable {ι R S : Type*} [CommSemiring R] [CommRing S] [IsDomain S] [Algebra R noncomputable def basisSpanSingleton (b : Basis ι R S) {x : S} (hx : x ≠ 0) : Basis ι R (span ({x} : Set S)) := b.map <| - LinearEquiv.ofInjective (Algebra.lmul R S x) (LinearMap.mul_injective hx) ≪≫ₗ + LinearEquiv.ofInjective (LinearMap.mulLeft R x) (mul_right_injective₀ hx) ≪≫ₗ LinearEquiv.ofEq _ _ (by ext @@ -33,7 +33,7 @@ theorem basisSpanSingleton_apply (b : Basis ι R S) {x : S} (hx : x ≠ 0) (i : (basisSpanSingleton b hx i : S) = x * b i := by simp only [basisSpanSingleton, Basis.map_apply, LinearEquiv.trans_apply, Submodule.restrictScalarsEquiv_apply, LinearEquiv.ofInjective_apply, LinearEquiv.coe_ofEq_apply, - LinearEquiv.restrictScalars_apply, Algebra.coe_lmul_eq_mul, LinearMap.mul_apply'] + LinearEquiv.restrictScalars_apply, LinearMap.mulLeft_apply, LinearMap.mul_apply'] @[simp] theorem constr_basisSpanSingleton {N : Type*} [Semiring N] [Module N S] [SMulCommClass R N S] From 487b33d7af3ccc73184c70deb15c976898d9e5c8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 10:10:59 +0000 Subject: [PATCH 17/90] chore: fix some left/right injectivity names (#19239) Follow-up Mathlib fixes from [zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/.60add_right_inj.60/near/482835891). --- Mathlib/Algebra/Group/Invertible/Defs.lean | 12 ++++++------ Mathlib/Algebra/GroupWithZero/Basic.lean | 4 ++-- Mathlib/Analysis/SpecialFunctions/Bernstein.lean | 4 ++-- 3 files changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Group/Invertible/Defs.lean b/Mathlib/Algebra/Group/Invertible/Defs.lean index 58a3021aca712..44a1421f131e7 100644 --- a/Mathlib/Algebra/Group/Invertible/Defs.lean +++ b/Mathlib/Algebra/Group/Invertible/Defs.lean @@ -230,23 +230,23 @@ section variable [Monoid α] {a b c : α} [Invertible c] variable (c) in -theorem mul_right_inj_of_invertible : a * c = b * c ↔ a = b := +theorem mul_left_inj_of_invertible : a * c = b * c ↔ a = b := ⟨fun h => by simpa using congr_arg (· * ⅟c) h, congr_arg (· * _)⟩ variable (c) in -theorem mul_left_inj_of_invertible : c * a = c * b ↔ a = b := +theorem mul_right_inj_of_invertible : c * a = c * b ↔ a = b := ⟨fun h => by simpa using congr_arg (⅟c * ·) h, congr_arg (_ * ·)⟩ theorem invOf_mul_eq_iff_eq_mul_left : ⅟c * a = b ↔ a = c * b := by - rw [← mul_left_inj_of_invertible (c := c), mul_invOf_cancel_left] + rw [← mul_right_inj_of_invertible (c := c), mul_invOf_cancel_left] theorem mul_left_eq_iff_eq_invOf_mul : c * a = b ↔ a = ⅟c * b := by - rw [← mul_left_inj_of_invertible (c := ⅟c), invOf_mul_cancel_left] + rw [← mul_right_inj_of_invertible (c := ⅟c), invOf_mul_cancel_left] theorem mul_invOf_eq_iff_eq_mul_right : a * ⅟c = b ↔ a = b * c := by - rw [← mul_right_inj_of_invertible (c := c), invOf_mul_cancel_right] + rw [← mul_left_inj_of_invertible (c := c), invOf_mul_cancel_right] theorem mul_right_eq_iff_eq_mul_invOf : a * c = b ↔ a = b * ⅟c := by - rw [← mul_right_inj_of_invertible (c := ⅟c), mul_invOf_cancel_right] + rw [← mul_left_inj_of_invertible (c := ⅟c), mul_invOf_cancel_right] end diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index c09c2c9f7742f..f242f9558df45 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -244,11 +244,11 @@ section GroupWithZero variable [GroupWithZero G₀] {a b x : G₀} -theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : +theorem GroupWithZero.mul_right_injective (h : x ≠ 0) : Function.Injective fun y => x * y := fun y y' w => by simpa only [← mul_assoc, inv_mul_cancel₀ h, one_mul] using congr_arg (fun y => x⁻¹ * y) w -theorem GroupWithZero.mul_right_injective (h : x ≠ 0) : +theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : Function.Injective fun y => y * x := fun y y' w => by simpa only [mul_assoc, mul_inv_cancel₀ h, mul_one] using congr_arg (fun y => y * x⁻¹) w diff --git a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean index a38081bf15582..4f6e5c73f04c5 100644 --- a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean +++ b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean @@ -109,8 +109,8 @@ theorem probability (n : ℕ) (x : I) : (∑ k : Fin (n + 1), bernstein n k x) = theorem variance {n : ℕ} (h : 0 < (n : ℝ)) (x : I) : (∑ k : Fin (n + 1), (x - k/ₙ : ℝ) ^ 2 * bernstein n k x) = (x : ℝ) * (1 - x) / n := by have h' : (n : ℝ) ≠ 0 := ne_of_gt h - apply_fun fun x : ℝ => x * n using GroupWithZero.mul_right_injective h' - apply_fun fun x : ℝ => x * n using GroupWithZero.mul_right_injective h' + apply_fun fun x : ℝ => x * n using GroupWithZero.mul_left_injective h' + apply_fun fun x : ℝ => x * n using GroupWithZero.mul_left_injective h' dsimp conv_lhs => simp only [Finset.sum_mul, z] conv_rhs => rw [div_mul_cancel₀ _ h'] From caa5dea7bb98b8587900aa059c01011f969262a1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 21:26:53 +1100 Subject: [PATCH 18/90] bump deps --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index effa2f997dc28..14175c9c38864 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "d99ff6205f5d87394c824868fb493fcdcc4debb5", + "rev": "9762792852464bf9591f6cccfe44f14b3ef54b25", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From 829c1d62a67ea6ccd730ad4616669c76f0999299 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 21:27:53 +1100 Subject: [PATCH 19/90] restore notation --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 12 ++-- .../Morphisms/RingHomProperties.lean | 8 +-- .../Morphisms/Separated.lean | 2 +- Mathlib/AlgebraicGeometry/Noetherian.lean | 2 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 55 ++++++++---------- Mathlib/AlgebraicGeometry/Properties.lean | 14 ++--- Mathlib/AlgebraicGeometry/RationalMap.lean | 2 +- Mathlib/AlgebraicGeometry/Restrict.lean | 58 +++++++++---------- Mathlib/AlgebraicGeometry/SpreadingOut.lean | 4 +- 9 files changed, 76 insertions(+), 81 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 482b93209f28e..a83addf8247e5 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -406,7 +406,7 @@ protected theorem isCompact : include hU in theorem image_of_isOpenImmersion (f : X ⟶ Y) [H : IsOpenImmersion f] : - IsAffineOpen (f ~~ᵁ U) := by + IsAffineOpen (f ''ᵁ U) := by have : IsAffine _ := hU convert isAffineOpen_opensRange (U.ι ≫ f) ext1 @@ -419,7 +419,7 @@ theorem preimage_of_isIso {U : Y.Opens} (hU : IsAffineOpen U) (f : X ⟶ Y) [IsI theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion (f : AlgebraicGeometry.Scheme.Hom X Y) [H : IsOpenImmersion f] {U : X.Opens} : - IsAffineOpen (f ~~ᵁ U) ↔ IsAffineOpen U := by + IsAffineOpen (f ''ᵁ U) ↔ IsAffineOpen U := by refine ⟨fun hU => @isAffine_of_isIso _ _ (IsOpenImmersion.isoOfRangeEq (X.ofRestrict U.isOpenEmbedding ≫ f) (Y.ofRestrict _) ?_).hom ?_ hU, fun hU => hU.image_of_isOpenImmersion f⟩ @@ -434,9 +434,9 @@ the affine open sets containing in the image. -/ @[simps] def _root_.AlgebraicGeometry.IsOpenImmersion.affineOpensEquiv (f : X ⟶ Y) [H : IsOpenImmersion f] : X.affineOpens ≃ { U : Y.affineOpens // U ≤ f.opensRange } where - toFun U := ⟨⟨f ~~ᵁ U, U.2.image_of_isOpenImmersion f⟩, Set.image_subset_range _ _⟩ + toFun U := ⟨⟨f ''ᵁ U, U.2.image_of_isOpenImmersion f⟩, Set.image_subset_range _ _⟩ invFun U := ⟨f ⁻¹ᵁ U, f.isAffineOpen_iff_of_isOpenImmersion.mp (by - rw [show f ~~ᵁ f ⁻¹ᵁ U = U from Opens.ext (Set.image_preimage_eq_of_subset U.2)]; exact U.1.2)⟩ + rw [show f ''ᵁ f ⁻¹ᵁ U = U from Opens.ext (Set.image_preimage_eq_of_subset U.2)]; exact U.1.2)⟩ left_inv _ := Subtype.ext (Opens.ext (Set.preimage_image_eq _ H.base_open.injective)) right_inv U := Subtype.ext (Subtype.ext (Opens.ext (Set.image_preimage_eq_of_subset U.2))) @@ -489,7 +489,7 @@ theorem fromSpec_preimage_basicOpen : rw [fromSpec_preimage_basicOpen', ← basicOpen_eq_of_affine] theorem fromSpec_image_basicOpen : - hU.fromSpec ~~ᵁ (PrimeSpectrum.basicOpen f) = X.basicOpen f := by + hU.fromSpec ''ᵁ (PrimeSpectrum.basicOpen f) = X.basicOpen f := by rw [← hU.fromSpec_preimage_basicOpen] ext1 change hU.fromSpec.base '' (hU.fromSpec.base ⁻¹' (X.basicOpen f : Set X)) = _ @@ -533,7 +533,7 @@ theorem exists_basicOpen_le {V : X.Opens} (x : V) (h : ↑x ∈ U) : (isBasis_basicOpen U).exists_subset_of_mem_open (x.2 : (⟨x, h⟩ : U) ∈ _) ((Opens.map U.inclusion').obj V).isOpen have : - U.ι ~~ᵁ (U.toScheme.basicOpen r) = + U.ι ''ᵁ (U.toScheme.basicOpen r) = X.basicOpen (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen U.ι r).trans ?_ rw [Scheme.basicOpen_res_eq] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index af281dce0055a..6a02465a5cb1a 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -106,10 +106,10 @@ theorem sourceAffineLocally_respectsIso (h₁ : RingHom.RespectsIso P) : (sourceAffineLocally P).toProperty.RespectsIso := by apply AffineTargetMorphismProperty.respectsIso_mk · introv H U - have : IsIso (e.hom.appLE (e.hom ~~ᵁ U) U.1 (e.hom.preimage_image_eq _).ge) := + have : IsIso (e.hom.appLE (e.hom ''ᵁ U) U.1 (e.hom.preimage_image_eq _).ge) := inferInstanceAs (IsIso (e.hom.app _ ≫ X.presheaf.map (eqToHom (e.hom.preimage_image_eq _).symm).op)) - rw [← Scheme.appLE_comp_appLE _ _ ⊤ (e.hom ~~ᵁ U) U.1 le_top (e.hom.preimage_image_eq _).ge, + rw [← Scheme.appLE_comp_appLE _ _ ⊤ (e.hom ''ᵁ U) U.1 le_top (e.hom.preimage_image_eq _).ge, h₁.cancel_right_isIso] exact H ⟨_, U.prop.image_of_isOpenImmersion e.hom⟩ · introv H U @@ -277,10 +277,10 @@ theorem comp_of_isOpenImmersion [IsOpenImmersion f] (H : P g) : P (f ≫ g) := by rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H ⊢ intro U V e - have : IsIso (f.appLE (f ~~ᵁ V) V.1 (f.preimage_image_eq _).ge) := + have : IsIso (f.appLE (f ''ᵁ V) V.1 (f.preimage_image_eq _).ge) := inferInstanceAs (IsIso (f.app _ ≫ X.presheaf.map (eqToHom (f.preimage_image_eq _).symm).op)) - rw [← Scheme.appLE_comp_appLE _ _ _ (f ~~ᵁ V) V.1 + rw [← Scheme.appLE_comp_appLE _ _ _ (f ''ᵁ V) V.1 (Set.image_subset_iff.mpr e) (f.preimage_image_eq _).ge, (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso] exact H _ ⟨_, V.2.image_of_isOpenImmersion _⟩ _ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index 1ff17eb0552fd..a97b12f9fcc9c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -166,7 +166,7 @@ lemma isClosedImmersion_diagonal_restrict_diagonalCoverDiagonalRange IsClosedImmersion (pullback.diagonal f ∣_ diagonalCoverDiagonalRange f 𝒰 𝒱) := by let U : (Σ i, (𝒱 i).J) → (diagonalCoverDiagonalRange f 𝒰 𝒱).toScheme.Opens := fun i ↦ (diagonalCoverDiagonalRange f 𝒰 𝒱).ι ⁻¹ᵁ ((diagonalCover f 𝒰 𝒱).map ⟨i.1, i.2, i.2⟩).opensRange - have hU (i) : (diagonalCoverDiagonalRange f 𝒰 𝒱).ι ~~ᵁ U i = + have hU (i) : (diagonalCoverDiagonalRange f 𝒰 𝒱).ι ''ᵁ U i = ((diagonalCover f 𝒰 𝒱).map ⟨i.1, i.2, i.2⟩).opensRange := by rw [TopologicalSpace.Opens.functor_obj_map_obj, inf_eq_right, Hom.image_top_eq_opensRange, Opens.opensRange_ι] diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index 70d5c5713fb29..a88d237043337 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -146,7 +146,7 @@ theorem isLocallyNoetherian_iff_of_affine_openCover (𝒰 : Scheme.OpenCover.{v, lemma isLocallyNoetherian_of_isOpenImmersion {Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] [IsLocallyNoetherian Y] : IsLocallyNoetherian X := by refine ⟨fun U => ?_⟩ - let V : Y.affineOpens := ⟨f ~~ᵁ U, IsAffineOpen.image_of_isOpenImmersion U.prop _⟩ + let V : Y.affineOpens := ⟨f ''ᵁ U, IsAffineOpen.image_of_isOpenImmersion U.prop _⟩ suffices Γ(X, U) ≅ Γ(Y, V) by convert isNoetherianRing_of_ringEquiv (R := Γ(Y, V)) _ · apply CategoryTheory.Iso.commRingCatIsoToRingEquiv diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index 3c8f3cfec6a45..834a151635fb5 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -87,15 +87,10 @@ def opensRange : Y.Opens := abbrev opensFunctor : X.Opens ⥤ Y.Opens := LocallyRingedSpace.IsOpenImmersion.opensFunctor f.toLRSHom -#adaptation_note -/-- -https://github.com/leanprover/lean4/pull/6012 disallowed the `''ᵁ` notation, -presumably unintentionally. This is a *temporary* workaround. --/ -/-- `f ~~ᵁ U` is notation for the image (as an open set) of `U` under an open immersion `f`. -/ -scoped[AlgebraicGeometry] notation3:90 f:91 " ~~ᵁ " U:90 => (Scheme.Hom.opensFunctor f).obj U +/-- `f ''ᵁ U` is notation for the image (as an open set) of `U` under an open immersion `f`. -/ +scoped[AlgebraicGeometry] notation3:90 f:91 " ''ᵁ " U:90 => (Scheme.Hom.opensFunctor f).obj U -lemma image_le_image_of_le {U V : X.Opens} (e : U ≤ V) : f ~~ᵁ U ≤ f ~~ᵁ V := by +lemma image_le_image_of_le {U V : X.Opens} (e : U ≤ V) : f ''ᵁ U ≤ f ''ᵁ V := by rintro a ⟨u, hu, rfl⟩ exact Set.mem_image_of_mem (⇑f.base) (e hu) @@ -105,12 +100,12 @@ lemma opensFunctor_map_homOfLE {U V : X.Opens} (e : U ≤ V) : rfl @[simp] -lemma image_top_eq_opensRange : f ~~ᵁ ⊤ = f.opensRange := by +lemma image_top_eq_opensRange : f ''ᵁ ⊤ = f.opensRange := by apply Opens.ext simp lemma opensRange_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) - [IsOpenImmersion f] [IsOpenImmersion g] : (f ≫ g).opensRange = g ~~ᵁ f.opensRange := + [IsOpenImmersion f] [IsOpenImmersion g] : (f ≫ g).opensRange = g ''ᵁ f.opensRange := TopologicalSpace.Opens.ext (Set.range_comp g.base f.base) lemma opensRange_of_isIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : @@ -122,36 +117,36 @@ lemma opensRange_comp_of_isIso {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) rw [opensRange_comp, opensRange_of_isIso, image_top_eq_opensRange] @[simp] -lemma preimage_image_eq (U : X.Opens) : f ⁻¹ᵁ f ~~ᵁ U = U := by +lemma preimage_image_eq (U : X.Opens) : f ⁻¹ᵁ f ''ᵁ U = U := by apply Opens.ext simp [Set.preimage_image_eq _ f.isOpenEmbedding.injective] lemma image_le_image_iff (f : X ⟶ Y) [IsOpenImmersion f] (U U' : X.Opens) : - f ~~ᵁ U ≤ f ~~ᵁ U' ↔ U ≤ U' := by + f ''ᵁ U ≤ f ''ᵁ U' ↔ U ≤ U' := by refine ⟨fun h ↦ ?_, image_le_image_of_le f⟩ rw [← preimage_image_eq f U, ← preimage_image_eq f U'] apply preimage_le_preimage_of_le f h -lemma image_preimage_eq_opensRange_inter (U : Y.Opens) : f ~~ᵁ f ⁻¹ᵁ U = f.opensRange ⊓ U := by +lemma image_preimage_eq_opensRange_inter (U : Y.Opens) : f ''ᵁ f ⁻¹ᵁ U = f.opensRange ⊓ U := by apply Opens.ext simp [Set.image_preimage_eq_range_inter] -lemma image_injective : Function.Injective (f ~~ᵁ ·) := by +lemma image_injective : Function.Injective (f ''ᵁ ·) := by intro U V hUV simpa using congrArg (f ⁻¹ᵁ ·) hUV lemma image_iSup {ι : Sort*} (s : ι → X.Opens) : - (f ~~ᵁ ⨆ (i : ι), s i) = ⨆ (i : ι), f ~~ᵁ s i := by + (f ''ᵁ ⨆ (i : ι), s i) = ⨆ (i : ι), f ''ᵁ s i := by ext : 1 simp [Set.image_iUnion] lemma image_iSup₂ {ι : Sort*} {κ : ι → Sort*} (s : (i : ι) → κ i → X.Opens) : - (f ~~ᵁ ⨆ (i : ι), ⨆ (j : κ i), s i j) = ⨆ (i : ι), ⨆ (j : κ i), f ~~ᵁ s i j := by + (f ''ᵁ ⨆ (i : ι), ⨆ (j : κ i), s i j) = ⨆ (i : ι), ⨆ (j : κ i), f ''ᵁ s i j := by ext : 1 simp [Set.image_iUnion₂] /-- The isomorphism `Γ(Y, f(U)) ≅ Γ(X, U)` induced by an open immersion `f : X ⟶ Y`. -/ -def appIso (U) : Γ(Y, f ~~ᵁ U) ≅ Γ(X, U) := +def appIso (U) : Γ(Y, f ''ᵁ U) ≅ Γ(X, U) := (asIso <| LocallyRingedSpace.IsOpenImmersion.invApp f.toLRSHom U).symm @[reassoc (attr := simp)] @@ -161,12 +156,12 @@ theorem appIso_inv_naturality {U V : X.Opens} (i : op U ⟶ op V) : PresheafedSpace.IsOpenImmersion.inv_naturality _ _ theorem appIso_hom (U) : - (f.appIso U).hom = f.app (f ~~ᵁ U) ≫ X.presheaf.map + (f.appIso U).hom = f.app (f ''ᵁ U) ≫ X.presheaf.map (eqToHom (preimage_image_eq f U).symm).op := (PresheafedSpace.IsOpenImmersion.inv_invApp f.toPshHom U).trans (by rw [eqToHom_op]) theorem appIso_hom' (U) : - (f.appIso U).hom = f.appLE (f ~~ᵁ U) U (preimage_image_eq f U).ge := + (f.appIso U).hom = f.appLE (f ''ᵁ U) U (preimage_image_eq f U).ge := f.appIso_hom U @[reassoc (attr := simp)] @@ -184,7 +179,7 @@ theorem app_invApp' (U) (hU : U ≤ f.opensRange) : @[reassoc (attr := simp), elementwise (attr := simp)] theorem appIso_inv_app (U) : - (f.appIso U).inv ≫ f.app (f ~~ᵁ U) = X.presheaf.map (eqToHom (preimage_image_eq f U)).op := + (f.appIso U).inv ≫ f.app (f ''ᵁ U) = X.presheaf.map (eqToHom (preimage_image_eq f U)).op := (PresheafedSpace.IsOpenImmersion.invApp_app _ _).trans (by rw [eqToHom_op]) @[reassoc (attr := simp), elementwise] @@ -200,8 +195,8 @@ lemma appLE_appIso_inv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U : @[reassoc (attr := simp)] lemma appIso_inv_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] {U V : X.Opens} - (e : V ≤ f ⁻¹ᵁ f ~~ᵁ U) : - (f.appIso U).inv ≫ f.appLE (f ~~ᵁ U) V e = + (e : V ≤ f ⁻¹ᵁ f ''ᵁ U) : + (f.appIso U).inv ≫ f.appLE (f ''ᵁ U) V e = X.presheaf.map (homOfLE (by rwa [preimage_image_eq] at e)).op := by simp only [appLE, appIso_inv_app_assoc, eqToHom_op] rw [← Functor.map_comp] @@ -213,7 +208,7 @@ end Scheme.Hom @[simps] def IsOpenImmersion.opensEquiv {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] : X.Opens ≃ { U : Y.Opens // U ≤ f.opensRange } where - toFun U := ⟨f ~~ᵁ U, Set.image_subset_range _ _⟩ + toFun U := ⟨f ''ᵁ U, Set.image_subset_range _ _⟩ invFun U := f ⁻¹ᵁ U left_inv _ := Opens.ext (Set.preimage_image_eq _ f.isOpenEmbedding.injective) right_inv U := Subtype.ext (Opens.ext (Set.image_preimage_eq_of_subset U.2)) @@ -340,7 +335,7 @@ instance IsOpenImmersion.ofRestrict : IsOpenImmersion (X.ofRestrict h) := @[simp] lemma Scheme.ofRestrict_appLE (V W e) : (X.ofRestrict h).appLE V W e = X.presheaf.map - (homOfLE (show X.ofRestrict h ~~ᵁ _ ≤ _ by exact Set.image_subset_iff.mpr e)).op := by + (homOfLE (show X.ofRestrict h ''ᵁ _ ≤ _ by exact Set.image_subset_iff.mpr e)).op := by dsimp [Hom.appLE] exact (X.presheaf.map_comp _ _).symm @@ -353,8 +348,8 @@ lemma Scheme.ofRestrict_appIso (U) : @[simp] lemma Scheme.restrict_presheaf_map (V W) (i : V ⟶ W) : - (X.restrict h).presheaf.map i = X.presheaf.map (homOfLE (show X.ofRestrict h ~~ᵁ W.unop ≤ - X.ofRestrict h ~~ᵁ V.unop from Set.image_subset _ i.unop.le)).op := rfl + (X.restrict h).presheaf.map i = X.presheaf.map (homOfLE (show X.ofRestrict h ''ᵁ W.unop ≤ + X.ofRestrict h ''ᵁ V.unop from Set.image_subset _ i.unop.le)).op := rfl end Restrict @@ -589,7 +584,7 @@ lemma isoOfRangeEq_inv_fac {X Y Z : Scheme.{u}} (f : X ⟶ Z) (g : Y ⟶ Z) theorem app_eq_invApp_app_of_comp_eq_aux {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : U ⟶ X) (fg : Y ⟶ X) (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : U.Opens) : - f ⁻¹ᵁ V = fg ⁻¹ᵁ (g ~~ᵁ V) := by + f ⁻¹ᵁ V = fg ⁻¹ᵁ (g ''ᵁ V) := by subst H rw [Scheme.comp_base, Opens.map_comp_obj] congr 1 @@ -599,7 +594,7 @@ theorem app_eq_invApp_app_of_comp_eq_aux {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : /-- The `fg` argument is to avoid nasty stuff about dependent types. -/ theorem app_eq_appIso_inv_app_of_comp_eq {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : U ⟶ X) (fg : Y ⟶ X) (H : fg = f ≫ g) [h : IsOpenImmersion g] (V : U.Opens) : - f.app V = (g.appIso V).inv ≫ fg.app (g ~~ᵁ V) ≫ Y.presheaf.map + f.app V = (g.appIso V).inv ≫ fg.app (g ''ᵁ V) ≫ Y.presheaf.map (eqToHom <| IsOpenImmersion.app_eq_invApp_app_of_comp_eq_aux f g fg H V).op := by subst H rw [Scheme.comp_app, Category.assoc, Scheme.Hom.appIso_inv_app_assoc, f.naturality_assoc, @@ -608,7 +603,7 @@ theorem app_eq_appIso_inv_app_of_comp_eq {X Y U : Scheme.{u}} (f : Y ⟶ U) (g : theorem lift_app {X Y U : Scheme.{u}} (f : U ⟶ Y) (g : X ⟶ Y) [IsOpenImmersion f] (H) (V : U.Opens) : - (IsOpenImmersion.lift f g H).app V = (f.appIso V).inv ≫ g.app (f ~~ᵁ V) ≫ + (IsOpenImmersion.lift f g H).app V = (f.appIso V).inv ≫ g.app (f ''ᵁ V) ≫ X.presheaf.map (eqToHom <| IsOpenImmersion.app_eq_invApp_app_of_comp_eq_aux _ _ _ (IsOpenImmersion.lift_fac f g H).symm V).op := IsOpenImmersion.app_eq_appIso_inv_app_of_comp_eq _ _ _ (lift_fac _ _ _).symm _ @@ -681,7 +676,7 @@ namespace Scheme theorem image_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) [H : IsOpenImmersion f] {U : X.Opens} (r : Γ(X, U)) : - f ~~ᵁ X.basicOpen r = Y.basicOpen ((f.appIso U).inv r) := by + f ''ᵁ X.basicOpen r = Y.basicOpen ((f.appIso U).inv r) := by have e := Scheme.preimage_basicOpen f ((f.appIso U).inv r) rw [Scheme.Hom.appIso_inv_app_apply, Scheme.basicOpen_res, inf_eq_right.mpr _] at e · rw [← e, f.image_preimage_eq_opensRange_inter, inf_eq_right] diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index b56b6ebd2f65d..800d80f953765 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -75,11 +75,11 @@ theorem isReduced_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenImm [IsReduced Y] : IsReduced X := by constructor intro U - have : U = f ⁻¹ᵁ f ~~ᵁ U := by + have : U = f ⁻¹ᵁ f ''ᵁ U := by ext1; exact (Set.preimage_image_eq _ H.base_open.injective).symm rw [this] - exact isReduced_of_injective (inv <| f.app (f ~~ᵁ U)) - (asIso <| f.app (f ~~ᵁ U) : Γ(Y, f ~~ᵁ U) ≅ _).symm.commRingCatIsoToRingEquiv.injective + exact isReduced_of_injective (inv <| f.app (f ''ᵁ U)) + (asIso <| f.app (f ''ᵁ U) : Γ(Y, f ''ᵁ U) ≅ _).symm.commRingCatIsoToRingEquiv.injective instance {R : CommRingCat.{u}} [H : _root_.IsReduced R] : IsReduced (Spec R) := by apply (config := { allowSynthFailures := true }) isReduced_of_isReduced_stalk @@ -252,13 +252,13 @@ theorem isIntegral_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : IsOpenIm [IsIntegral Y] [Nonempty X] : IsIntegral X := by constructor; · infer_instance intro U hU - have : U = f ⁻¹ᵁ f ~~ᵁ U := by ext1; exact (Set.preimage_image_eq _ H.base_open.injective).symm + have : U = f ⁻¹ᵁ f ''ᵁ U := by ext1; exact (Set.preimage_image_eq _ H.base_open.injective).symm rw [this] - have : IsDomain Γ(Y, f ~~ᵁ U) := by + have : IsDomain Γ(Y, f ''ᵁ U) := by apply (config := { allowSynthFailures := true }) IsIntegral.component_integral exact ⟨⟨_, _, hU.some.prop, rfl⟩⟩ - exact (asIso <| f.app (f ~~ᵁ U) : - Γ(Y, f ~~ᵁ U) ≅ _).symm.commRingCatIsoToRingEquiv.toMulEquiv.isDomain _ + exact (asIso <| f.app (f ''ᵁ U) : + Γ(Y, f ''ᵁ U) ≅ _).symm.commRingCatIsoToRingEquiv.toMulEquiv.isDomain _ instance {R : CommRingCat} [IsDomain R] : IrreducibleSpace (Spec R) := by convert PrimeSpectrum.irreducibleSpace (R := R) diff --git a/Mathlib/AlgebraicGeometry/RationalMap.lean b/Mathlib/AlgebraicGeometry/RationalMap.lean index 4e0f0132a7587..aca42607f4a77 100644 --- a/Mathlib/AlgebraicGeometry/RationalMap.lean +++ b/Mathlib/AlgebraicGeometry/RationalMap.lean @@ -217,7 +217,7 @@ lemma equiv_of_fromSpecStalkOfMem_eq [IrreducibleSpace X] have := spread_out_unique_of_isGermInjective' (X := (f.domain ⊓ g.domain).toScheme) (X.homOfLE inf_le_left ≫ f.hom) (X.homOfLE inf_le_right ≫ g.hom) (x := ⟨x, hxf, hxg⟩) ?_ · obtain ⟨U, hxU, e⟩ := this - refine ⟨(f.domain ⊓ g.domain).ι ~~ᵁ U, ((f.domain ⊓ g.domain).ι ~~ᵁ U).2.dense + refine ⟨(f.domain ⊓ g.domain).ι ''ᵁ U, ((f.domain ⊓ g.domain).ι ''ᵁ U).2.dense ⟨_, ⟨_, hxU, rfl⟩⟩, ((Set.image_subset_range _ _).trans_eq (Subtype.range_val)).trans inf_le_left, ((Set.image_subset_range _ _).trans_eq (Subtype.range_val)).trans inf_le_right, ?_⟩ diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index cfbae16d1cf72..e1e6db017d998 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -57,7 +57,7 @@ instance : IsOpenImmersion U.ι := inferInstanceAs (IsOpenImmersion (X.ofRestric lemma toScheme_carrier : (U : Type u) = (U : Set X) := rfl -lemma toScheme_presheaf_obj (V) : Γ(U, V) = Γ(X, U.ι ~~ᵁ V) := rfl +lemma toScheme_presheaf_obj (V) : Γ(U, V) = Γ(X, U.ι ''ᵁ V) := rfl @[simp] lemma toScheme_presheaf_map {V W} (i : V ⟶ W) : @@ -65,13 +65,13 @@ lemma toScheme_presheaf_map {V W} (i : V ⟶ W) : @[simp] lemma ι_app (V) : U.ι.app V = X.presheaf.map - (homOfLE (x := U.ι ~~ᵁ U.ι ⁻¹ᵁ V) (Set.image_preimage_subset _ _)).op := + (homOfLE (x := U.ι ''ᵁ U.ι ⁻¹ᵁ V) (Set.image_preimage_subset _ _)).op := rfl @[simp] lemma ι_appLE (V W e) : U.ι.appLE V W e = - X.presheaf.map (homOfLE (x := U.ι ~~ᵁ W) (Set.image_subset_iff.mpr ‹_›)).op := by + X.presheaf.map (homOfLE (x := U.ι ''ᵁ W) (Set.image_subset_iff.mpr ‹_›)).op := by simp only [Hom.appLE, ι_app, Functor.op_obj, Opens.carrier_eq_coe, toScheme_presheaf_map, Quiver.Hom.unop_op, Hom.opensFunctor_map_homOfLE, Opens.coe_inclusion', ← Functor.map_comp] rfl @@ -88,10 +88,10 @@ lemma opensRange_ι : U.ι.opensRange = U := lemma range_ι : Set.range U.ι.base = U := Subtype.range_val -lemma ι_image_top : U.ι ~~ᵁ ⊤ = U := +lemma ι_image_top : U.ι ''ᵁ ⊤ = U := U.isOpenEmbedding_obj_top -lemma ι_image_le (W : U.toScheme.Opens) : U.ι ~~ᵁ W ≤ U := by +lemma ι_image_le (W : U.toScheme.Opens) : U.ι ''ᵁ W ≤ U := by simp_rw [← U.ι_image_top] exact U.ι.image_le_image_of_le le_top @@ -105,9 +105,9 @@ instance ι_appLE_isIso : show IsIso (X.presheaf.map (eqToIso U.ι_image_top).hom.op) infer_instance -lemma ι_app_self : U.ι.app U = X.presheaf.map (eqToHom (X := U.ι ~~ᵁ _) (by simp)).op := rfl +lemma ι_app_self : U.ι.app U = X.presheaf.map (eqToHom (X := U.ι ''ᵁ _) (by simp)).op := rfl -lemma eq_presheaf_map_eqToHom {V W : Opens U} (e : U.ι ~~ᵁ V = U.ι ~~ᵁ W) : +lemma eq_presheaf_map_eqToHom {V W : Opens U} (e : U.ι ''ᵁ V = U.ι ''ᵁ W) : X.presheaf.map (eqToHom e).op = U.toScheme.presheaf.map (eqToHom <| U.isOpenEmbedding.functor_obj_injective e).op := rfl @@ -131,12 +131,12 @@ def stalkIso {X : Scheme.{u}} (U : X.Opens) (x : U) : lemma germ_stalkIso_hom {X : Scheme.{u}} (U : X.Opens) {V : U.toScheme.Opens} (x : U) (hx : x ∈ V) : U.toScheme.presheaf.germ V x hx ≫ (U.stalkIso x).hom = - X.presheaf.germ (U.ι ~~ᵁ V) x.1 ⟨x, hx, rfl⟩ := + X.presheaf.germ (U.ι ''ᵁ V) x.1 ⟨x, hx, rfl⟩ := PresheafedSpace.restrictStalkIso_hom_eq_germ _ U.isOpenEmbedding _ _ _ @[reassoc] lemma germ_stalkIso_inv {X : Scheme.{u}} (U : X.Opens) (V : U.toScheme.Opens) (x : U) - (hx : x ∈ V) : X.presheaf.germ (U.ι ~~ᵁ V) x ⟨x, hx, rfl⟩ ≫ + (hx : x ∈ V) : X.presheaf.germ (U.ι ''ᵁ V) x ⟨x, hx, rfl⟩ ≫ (U.stalkIso x).inv = U.toScheme.presheaf.germ V x hx := PresheafedSpace.restrictStalkIso_inv_eq_germ X.toPresheafedSpace U.isOpenEmbedding V x hx @@ -171,7 +171,7 @@ instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : (U.ι.app ⊤).toAlgebra lemma Scheme.map_basicOpen (r : Γ(U, ⊤)) : - U.ι ~~ᵁ U.toScheme.basicOpen r = X.basicOpen + U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen (X.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op r) := by refine (Scheme.image_basicOpen (X.ofRestrict U.isOpenEmbedding) r).trans ?_ rw [← Scheme.basicOpen_res_eq _ _ (eqToHom U.isOpenEmbedding_obj_top).op] @@ -183,11 +183,11 @@ lemma Scheme.map_basicOpen (r : Γ(U, ⊤)) : @[deprecated (since := "2024-10-23")] alias Scheme.map_basicOpen' := Scheme.map_basicOpen lemma Scheme.Opens.ι_image_basicOpen (r : Γ(U, ⊤)) : - U.ι ~~ᵁ U.toScheme.basicOpen r = X.basicOpen r := by + U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen r := by rw [Scheme.map_basicOpen, Scheme.basicOpen_res_eq] lemma Scheme.map_basicOpen_map (r : Γ(X, U)) : - U.ι ~~ᵁ (U.toScheme.basicOpen <| U.topIso.inv r) = X.basicOpen r := by + U.ι ''ᵁ (U.toScheme.basicOpen <| U.topIso.inv r) = X.basicOpen r := by simp only [Scheme.Opens.toScheme_presheaf_obj] rw [Scheme.map_basicOpen, Scheme.basicOpen_res_eq, Scheme.Opens.topIso_inv, Scheme.basicOpen_res_eq X] @@ -223,7 +223,7 @@ theorem Scheme.homOfLE_apply {U V : X.Opens} (e : U ≤ V) (x : U) : rfl theorem Scheme.ι_image_homOfLE_le_ι_image {U V : X.Opens} (e : U ≤ V) (W : Opens V) : - U.ι ~~ᵁ (X.homOfLE e ⁻¹ᵁ W) ≤ V.ι ~~ᵁ W := by + U.ι ''ᵁ (X.homOfLE e ⁻¹ᵁ W) ≤ V.ι ''ᵁ W := by simp only [← SetLike.coe_subset_coe, IsOpenMap.functor_obj_coe, Set.image_subset_iff, Scheme.homOfLE_base, Opens.map_coe, Opens.inclusion'_apply] rintro _ h @@ -233,8 +233,8 @@ theorem Scheme.ι_image_homOfLE_le_ι_image {U V : X.Opens} (e : U ≤ V) (W : O theorem Scheme.homOfLE_app {U V : X.Opens} (e : U ≤ V) (W : Opens V) : (X.homOfLE e).app W = X.presheaf.map (homOfLE <| X.ι_image_homOfLE_le_ι_image e W).op := by - have e₁ := Scheme.congr_app (X.homOfLE_ι e) (V.ι ~~ᵁ W) - have : V.ι ⁻¹ᵁ V.ι ~~ᵁ W = W := W.map_functor_eq (U := V) + have e₁ := Scheme.congr_app (X.homOfLE_ι e) (V.ι ''ᵁ W) + have : V.ι ⁻¹ᵁ V.ι ''ᵁ W = W := W.map_functor_eq (U := V) have e₂ := (X.homOfLE e).naturality (eqToIso this).hom.op have e₃ := e₂.symm.trans e₁ dsimp at e₃ ⊢ @@ -302,23 +302,23 @@ def Scheme.restrictRestrictComm (X : Scheme.{u}) (U V : X.Opens) : simp [Set.image_preimage_eq_inter_range, Set.inter_comm (U : Set X), Set.range_comp] /-- If `f : X ⟶ Y` is an open immersion, then for any `U : X.Opens`, -we have the isomorphism `U ≅ f ~~ᵁ U`. -/ +we have the isomorphism `U ≅ f ''ᵁ U`. -/ noncomputable def Scheme.Hom.isoImage {X Y : Scheme.{u}} (f : X.Hom Y) [IsOpenImmersion f] (U : X.Opens) : - U.toScheme ≅ f ~~ᵁ U := + U.toScheme ≅ f ''ᵁ U := IsOpenImmersion.isoOfRangeEq (Opens.ι _ ≫ f) (Opens.ι _) (by simp [Set.range_comp]) @[reassoc (attr := simp)] lemma Scheme.Hom.isoImage_hom_ι {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : X.Opens) : - (f.isoImage U).hom ≫ (f ~~ᵁ U).ι = U.ι ≫ f := + (f.isoImage U).hom ≫ (f ''ᵁ U).ι = U.ι ≫ f := IsOpenImmersion.isoOfRangeEq_hom_fac _ _ _ @[reassoc (attr := simp)] lemma Scheme.Hom.isoImage_inv_ι {X Y : Scheme.{u}} (f : X ⟶ Y) [IsOpenImmersion f] (U : X.Opens) : - (f.isoImage U).inv ≫ U.ι ≫ f = (f ~~ᵁ U).ι := + (f.isoImage U).inv ≫ U.ι ≫ f = (f ''ᵁ U).ι := IsOpenImmersion.isoOfRangeEq_inv_fac _ _ _ @[deprecated (since := "2024-10-20")] @@ -475,7 +475,7 @@ theorem morphismRestrict_base {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : funext fun x => Subtype.ext (morphismRestrict_base_coe f U x) theorem image_morphismRestrict_preimage {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Opens U) : - (f ⁻¹ᵁ U).ι ~~ᵁ ((f ∣_ U) ⁻¹ᵁ V) = f ⁻¹ᵁ (U.ι ~~ᵁ V) := by + (f ⁻¹ᵁ U).ι ''ᵁ ((f ∣_ U) ⁻¹ᵁ V) = f ⁻¹ᵁ (U.ι ''ᵁ V) := by ext1 ext x constructor @@ -498,14 +498,14 @@ lemma eqToHom_eq_homOfLE {C} [Preorder C] {X Y : C} (e : X = Y) : eqToHom e = ho open Scheme in theorem morphismRestrict_app {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : U.toScheme.Opens) : - (f ∣_ U).app V = f.app (U.ι ~~ᵁ V) ≫ + (f ∣_ U).app V = f.app (U.ι ''ᵁ V) ≫ X.presheaf.map (eqToHom (image_morphismRestrict_preimage f U V)).op := by - have := Scheme.congr_app (morphismRestrict_ι f U) (U.ι ~~ᵁ V) + have := Scheme.congr_app (morphismRestrict_ι f U) (U.ι ''ᵁ V) simp only [Scheme.preimage_comp, Opens.toScheme_presheaf_obj, Hom.app_eq_appLE, comp_appLE, Opens.ι_appLE, eqToHom_op, Opens.toScheme_presheaf_map, eqToHom_unop] at this - have e : U.ι ⁻¹ᵁ (U.ι ~~ᵁ V) = V := + have e : U.ι ⁻¹ᵁ (U.ι ''ᵁ V) = V := Opens.ext (Set.preimage_image_eq _ Subtype.coe_injective) - have e' : (f ∣_ U) ⁻¹ᵁ V = (f ∣_ U) ⁻¹ᵁ U.ι ⁻¹ᵁ U.ι ~~ᵁ V := by rw [e] + have e' : (f ∣_ U) ⁻¹ᵁ V = (f ∣_ U) ⁻¹ᵁ U.ι ⁻¹ᵁ U.ι ''ᵁ V := by rw [e] simp only [Opens.toScheme_presheaf_obj, Hom.app_eq_appLE, eqToHom_op, Hom.appLE_map] rw [← (f ∣_ U).appLE_map' _ e', ← (f ∣_ U).map_appLE' _ e] simp only [Opens.toScheme_presheaf_obj, eqToHom_eq_homOfLE, Opens.toScheme_presheaf_map, @@ -519,7 +519,7 @@ theorem morphismRestrict_app' {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V @[simp] theorem morphismRestrict_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V W e) : - (f ∣_ U).appLE V W e = f.appLE (U.ι ~~ᵁ V) ((f ⁻¹ᵁ U).ι ~~ᵁ W) + (f ∣_ U).appLE V W e = f.appLE (U.ι ''ᵁ V) ((f ⁻¹ᵁ U).ι ''ᵁ W) ((Set.image_subset _ e).trans (image_morphismRestrict_preimage f U V).le) := by rw [Scheme.Hom.appLE, morphismRestrict_app', Scheme.Opens.toScheme_presheaf_map, Scheme.Hom.appLE_map] @@ -557,7 +557,7 @@ def morphismRestrictEq {X Y : Scheme.{u}} (f : X ⟶ Y) {U V : Y.Opens} (e : U = /-- Restricting a morphism twice is isomorphic to one restriction. -/ def morphismRestrictRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : U.toScheme.Opens) : - Arrow.mk (f ∣_ U ∣_ V) ≅ Arrow.mk (f ∣_ U.ι ~~ᵁ V) := by + Arrow.mk (f ∣_ U ∣_ V) ≅ Arrow.mk (f ∣_ U.ι ''ᵁ V) := by refine Arrow.isoMk' _ _ ((Scheme.Opens.ι _).isoImage _ ≪≫ Scheme.isoOfEq _ ?_) ((Scheme.Opens.ι _).isoImage _) ?_ · ext x @@ -653,18 +653,18 @@ lemma resLE_congr (e₁ : U = U') (e₂ : V = V') (P : MorphismProperty Scheme.{ lemma resLE_preimage (f : X ⟶ Y) {U : Y.Opens} {V : X.Opens} (e : V ≤ f ⁻¹ᵁ U) (O : U.toScheme.Opens) : - f.resLE U V e ⁻¹ᵁ O = V.ι ⁻¹ᵁ (f ⁻¹ᵁ U.ι ~~ᵁ O) := by + f.resLE U V e ⁻¹ᵁ O = V.ι ⁻¹ᵁ (f ⁻¹ᵁ U.ι ''ᵁ O) := by rw [← preimage_comp, ← resLE_comp_ι f e, preimage_comp, preimage_image_eq] lemma le_preimage_resLE_iff {U : Y.Opens} {V : X.Opens} (e : V ≤ f ⁻¹ᵁ U) (O : U.toScheme.Opens) (W : V.toScheme.Opens) : - W ≤ (f.resLE U V e) ⁻¹ᵁ O ↔ V.ι ~~ᵁ W ≤ f ⁻¹ᵁ U.ι ~~ᵁ O := by + W ≤ (f.resLE U V e) ⁻¹ᵁ O ↔ V.ι ''ᵁ W ≤ f ⁻¹ᵁ U.ι ''ᵁ O := by simp [resLE_preimage, ← image_le_image_iff V.ι, image_preimage_eq_opensRange_inter, V.ι_image_le] lemma resLE_appLE {U : Y.Opens} {V : X.Opens} (e : V ≤ f ⁻¹ᵁ U) (O : U.toScheme.Opens) (W : V.toScheme.Opens) (e' : W ≤ resLE f U V e ⁻¹ᵁ O) : (f.resLE U V e).appLE O W e' = - f.appLE (U.ι ~~ᵁ O) (V.ι ~~ᵁ W) ((le_preimage_resLE_iff f e O W).mp e') := by + f.appLE (U.ι ''ᵁ O) (V.ι ''ᵁ W) ((le_preimage_resLE_iff f e O W).mp e') := by simp only [appLE, resLE, comp_coeBase, Opens.map_comp_obj, comp_app, morphismRestrict_app', homOfLE_leOfHom, homOfLE_app, Category.assoc, Opens.toScheme_presheaf_map, Quiver.Hom.unop_op, opensFunctor_map_homOfLE] diff --git a/Mathlib/AlgebraicGeometry/SpreadingOut.lean b/Mathlib/AlgebraicGeometry/SpreadingOut.lean index 21b5458d77f0e..b5d7ef83b97b8 100644 --- a/Mathlib/AlgebraicGeometry/SpreadingOut.lean +++ b/Mathlib/AlgebraicGeometry/SpreadingOut.lean @@ -87,7 +87,7 @@ lemma Scheme.exists_le_and_germ_injective (X : Scheme.{u}) (x : X) [X.IsGermInje instance (x : X) [X.IsGermInjectiveAt x] [IsOpenImmersion f] : Y.IsGermInjectiveAt (f.base x) := by obtain ⟨U, hxU, hU, H⟩ := X.exists_germ_injective x - refine ⟨⟨f ~~ᵁ U, ⟨x, hxU, rfl⟩, hU.image_of_isOpenImmersion f, ?_⟩⟩ + refine ⟨⟨f ''ᵁ U, ⟨x, hxU, rfl⟩, hU.image_of_isOpenImmersion f, ?_⟩⟩ refine ((MorphismProperty.injective CommRingCat).cancel_right_of_respectsIso _ (f.stalkMap x)).mp ?_ refine ((MorphismProperty.injective CommRingCat).cancel_left_of_respectsIso @@ -101,7 +101,7 @@ lemma isGermInjectiveAt_iff_of_isOpenImmersion {x : X} [IsOpenImmersion f]: obtain ⟨U, hxU, hU, hU', H⟩ := Y.exists_le_and_germ_injective (f.base x) (V := f.opensRange) ⟨x, rfl⟩ obtain ⟨V, hV⟩ := (IsOpenImmersion.affineOpensEquiv f).surjective ⟨⟨U, hU⟩, hU'⟩ - obtain rfl : f ~~ᵁ V = U := Subtype.eq_iff.mp (Subtype.eq_iff.mp hV) + obtain rfl : f ''ᵁ V = U := Subtype.eq_iff.mp (Subtype.eq_iff.mp hV) obtain ⟨y, hy, e : f.base y = f.base x⟩ := hxU obtain rfl := f.isOpenEmbedding.injective e refine ⟨V, hy, V.2, ?_⟩ From 60930c119a77b2c3ae18d4653cb816889fce70a4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 11:04:08 +0000 Subject: [PATCH 20/90] chore: shortcut instances for `Finite Bool`, `Finite Prop` (#19086) These can be defined much earlier. Also add `WellFoundedLT` shortcut instances. [Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/WellFoundedLT.20Prop.20is.20not.20found.20when.20importing.20too.20much) --- Mathlib/Data/Finite/Defs.lean | 3 +++ Mathlib/Data/Fintype/Card.lean | 7 +++++++ 2 files changed, 10 insertions(+) diff --git a/Mathlib/Data/Finite/Defs.lean b/Mathlib/Data/Finite/Defs.lean index 784e71e12e4f6..dce5c020d7baf 100644 --- a/Mathlib/Data/Finite/Defs.lean +++ b/Mathlib/Data/Finite/Defs.lean @@ -162,6 +162,9 @@ protected theorem Infinite.false [Finite α] (_ : Infinite α) : False := alias ⟨Finite.of_not_infinite, Finite.not_infinite⟩ := not_infinite_iff_finite +instance Bool.instFinite : Finite Bool := .intro finTwoEquiv.symm +instance Prop.instFinite : Finite Prop := .of_equiv _ Equiv.propEquivBool.symm + section Set /-! diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index dc1e60e25f55d..59e57330eb779 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -841,6 +841,13 @@ instance (priority := 100) to_wellFoundedGT [Preorder α] : WellFoundedGT α := end Finite +-- Shortcut instances to make sure those are found even in the presence of other instances +-- See https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/WellFoundedLT.20Prop.20is.20not.20found.20when.20importing.20too.20much +instance Bool.instWellFoundedLT : WellFoundedLT Bool := inferInstance +instance Bool.instWellFoundedGT : WellFoundedGT Bool := inferInstance +instance Prop.instWellFoundedLT : WellFoundedLT Prop := inferInstance +instance Prop.instWellFoundedGT : WellFoundedGT Prop := inferInstance + -- @[nolint fintype_finite] -- Porting note: do we need this? protected theorem Fintype.false [Infinite α] (_h : Fintype α) : False := not_finite α From 2fd09beb18e05ada8bff95d82fdc2e0e6a55ba0a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Tom=C3=A1=C5=A1=20Sk=C5=99ivan?= Date: Tue, 19 Nov 2024 11:04:10 +0000 Subject: [PATCH 21/90] feat(fun_prop): failure cache for `fun_prop` and command to print out `fun_prop` theorems (#19112) Add new cache to `fun_prop` for failed goals. `fun_prop` used to take exponentially long to fail without this cache. Also add command `#print_fun_prop_theorems HAdd.hAdd` prints out all `fun_prop` theorems associated with addition. --- Mathlib/Tactic/FunProp/Core.lean | 50 +++++++++++------ Mathlib/Tactic/FunProp/Elab.lean | 53 ++++++++++++++++++ Mathlib/Tactic/FunProp/Theorems.lean | 2 +- Mathlib/Tactic/FunProp/Types.lean | 4 +- MathlibTest/fun_prop_dev.lean | 83 +++++++++++++++++++++++++--- 5 files changed, 166 insertions(+), 26 deletions(-) diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 7cfc38fde4b73..9f84b5b2c1371 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -507,20 +507,18 @@ def tryTheorems (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some r ← tryTheorem? e thm.thmOrigin funProp then return r | .some (.some (f,g)) => - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposing to later use {←ppOrigin' thm.thmOrigin}" - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposition: {← ppExpr f} ∘ {← ppExpr g}" + trace[Meta.Tactic.fun_prop] + s!"decomposing to later use {←ppOrigin' thm.thmOrigin} as: + ({← ppExpr f}) ∘ ({← ppExpr g})" if let .some r ← applyCompRule funPropDecl e f g funProp then return r | _ => continue else - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposing in args {thm.mainArgs} to later use {←ppOrigin' thm.thmOrigin}" let .some (f,g) ← fData.decompositionOverArgs thm.mainArgs | continue - trace[Debug.Meta.Tactic.fun_prop] - s!"decomposition: {← ppExpr f} ∘ {← ppExpr g}" + trace[Meta.Tactic.fun_prop] + s!"decomposing to later use {←ppOrigin' thm.thmOrigin} as: + ({← ppExpr f}) ∘ ({← ppExpr g})" if let .some r ← applyCompRule funPropDecl e f g funProp then return r -- todo: decompose if uncurried and arguments do not match exactly @@ -594,9 +592,17 @@ def constAppCase (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) return r if let .some (f,g) ← fData.nontrivialDecomposition then + trace[Meta.Tactic.fun_prop] + s!"failed applying `{funPropDecl.funPropName}` theorems for `{funName}` + trying again after decomposing function as: `({← ppExpr f}) ∘ ({← ppExpr g})`" + if let .some r ← applyCompRule funPropDecl e f g funProp then return r else + trace[Meta.Tactic.fun_prop] + s!"failed applying `{funPropDecl.funPropName}` theorems for `{funName}` + now trying to prove `{funPropDecl.funPropName}` from another function property" + if let .some r ← applyTransitionRules e funProp then return r @@ -609,15 +615,27 @@ def cacheResult (e : Expr) (r : Result) : FunPropM Result := do -- return proof? modify (fun s => { s with cache := s.cache.insert e { expr := q(True), proof? := r.proof} }) return r +/-- Cache for failed goals such that `fun_prop` can fail fast next time. -/ +def cacheFailure (e : Expr) : FunPropM Unit := do -- return proof? + modify (fun s => { s with failureCache := s.failureCache.insert e }) + + mutual /-- Main `funProp` function. Returns proof of `e`. -/ partial def funProp (e : Expr) : FunPropM (Option Result) := do let e ← instantiateMVars e - -- check cache + + withTraceNode `Meta.Tactic.fun_prop + (fun r => do pure s!"[{ExceptToEmoji.toEmoji r}] {← ppExpr e}") do + + -- check cache for succesfull goals if let .some { expr := _, proof? := .some proof } := (← get).cache.find? e then - trace[Debug.Meta.Tactic.fun_prop] "cached result for {e}" + trace[Meta.Tactic.fun_prop] "reusing previously found proof for {e}" return .some { proof := proof } + else if (← get).failureCache.contains e then + trace[Meta.Tactic.fun_prop] "skipping proof search, proving {e} was tried already and failed" + return .none else -- take care of forall and let binders and run main match e with @@ -633,9 +651,12 @@ mutual cacheResult e {proof := ← mkLambdaFVars xs r.proof } | .mdata _ e' => funProp e' | _ => - let .some r ← main e - | return none - cacheResult e r + if let .some r ← main e then + cacheResult e r + else + cacheFailure e + return none + /-- Main `funProp` function. Returns proof of `e`. -/ private partial def main (e : Expr) : FunPropM (Option Result) := do @@ -645,9 +666,6 @@ mutual increaseSteps - withTraceNode `Meta.Tactic.fun_prop - (fun r => do pure s!"[{ExceptToEmoji.toEmoji r}] {← ppExpr e}") do - -- if function starts with let bindings move them the top of `e` and try again if f.isLet then return ← letTelescope f fun xs b => do diff --git a/Mathlib/Tactic/FunProp/Elab.lean b/Mathlib/Tactic/FunProp/Elab.lean index 94afec84e9177..6761d53f837c6 100644 --- a/Mathlib/Tactic/FunProp/Elab.lean +++ b/Mathlib/Tactic/FunProp/Elab.lean @@ -83,6 +83,59 @@ def funPropTac : Tactic | _ => throwUnsupportedSyntax + + +/-- Command that printins all function properties attached to a function. + +For example +``` +#print_fun_prop_theorems HAdd.hAdd +``` +might print out +``` +Continuous + continuous_add, args: [4,5], priority: 1000 + continuous_add_left, args: [5], priority: 1000 + continuous_add_right, args [4], priority: 1000 + ... +Diferentiable + Differentiable.add, args: [4,5], priority: 1000 + Differentiable.add_const, args: [4], priority: 1000 + Differentiable.const_add, args: [5], priority: 1000 + ... +``` + +You can also see only theorems about a concrete function property +``` +#print_fun_prop_theorems HAdd.hAdd Continuous +``` +-/ +elab "#print_fun_prop_theorems " funIdent:ident funProp:(ident)? : command => do + + let funName ← ensureNonAmbiguous funIdent (← resolveGlobalConst funIdent) + let funProp? ← funProp.mapM (fun stx => do + ensureNonAmbiguous stx (← resolveGlobalConst stx)) + + let theorems := (functionTheoremsExt.getState (← getEnv)).theorems.findD funName {} + + let logTheorems (funProp : Name) (thms : Array FunctionTheorem) : Command.CommandElabM Unit := do + let mut msg : MessageData := "" + msg := msg ++ m!"{← Meta.ppOrigin (.decl funProp)}" + for thm in thms do + msg := msg ++ m!"\n {← Meta.ppOrigin (.decl thm.thmOrigin.name)}, \ + args: {thm.mainArgs}, form: {thm.form}" + pure () + logInfo msg + + + match funProp? with + | none => + for (funProp,thms) in theorems do + logTheorems funProp thms + | some funProp => + logTheorems funProp (theorems.findD funProp #[]) + + end Meta.FunProp end Mathlib diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 1ab9b3850fec5..b0c184aedc988 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -145,7 +145,7 @@ inductive TheoremForm where /-- TheoremForm to string -/ instance : ToString TheoremForm := - ⟨fun x => match x with | .uncurried => "uncurried" | .comp => "compositional"⟩ + ⟨fun x => match x with | .uncurried => "simple" | .comp => "compositional"⟩ /-- theorem about specific function (either declared constant or free variable) -/ structure FunctionTheorem where diff --git a/Mathlib/Tactic/FunProp/Types.lean b/Mathlib/Tactic/FunProp/Types.lean index 6f12b788e1928..a8af67ac13b91 100644 --- a/Mathlib/Tactic/FunProp/Types.lean +++ b/Mathlib/Tactic/FunProp/Types.lean @@ -97,8 +97,10 @@ structure Context where /-- `fun_prop` state -/ structure State where /-- Simp's cache is used as the `fun_prop` tactic is designed to be used inside of simp and - utilize its cache -/ + utilize its cache. It holds successful goals. -/ cache : Simp.Cache := {} + /-- Cache storing failed goals such that they are not tried again. -/ + failureCache : ExprSet := {} /-- Count the number of steps and stop when maxSteps is reached. -/ numSteps := 0 /-- Log progress and failures messages that should be displayed to the user at the end. -/ diff --git a/MathlibTest/fun_prop_dev.lean b/MathlibTest/fun_prop_dev.lean index f78abae43617a..866e3b6aebefa 100644 --- a/MathlibTest/fun_prop_dev.lean +++ b/MathlibTest/fun_prop_dev.lean @@ -6,6 +6,7 @@ Authors: Tomáš Skřivan import Mathlib.Tactic.FunProp import Mathlib.Logic.Function.Basic import Mathlib.Data.FunLike.Basic +import Mathlib.Tactic.SuccessIfFailWithMsg import Aesop /-! # Tests for the `fun_prop` tactic @@ -536,14 +537,15 @@ example [Add α] (a : α) : -- Test that local theorem is being used /-- info: [Meta.Tactic.fun_prop] [✅️] Con fun x => f x y - [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] - [Meta.Tactic.fun_prop] removing argument to later use this : Con f - [Meta.Tactic.fun_prop] [✅️] applying: Con_comp - [Meta.Tactic.fun_prop] [✅️] Con fun f => f y - [Meta.Tactic.fun_prop] [✅️] applying: Con_apply - [Meta.Tactic.fun_prop] [✅️] Con fun x => f x - [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] - [Meta.Tactic.fun_prop] [✅️] applying: this : Con f + [Meta.Tactic.fun_prop] [✅️] Con fun x => f x y + [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] + [Meta.Tactic.fun_prop] removing argument to later use this : Con f + [Meta.Tactic.fun_prop] [✅️] applying: Con_comp + [Meta.Tactic.fun_prop] [✅️] Con fun f => f y + [Meta.Tactic.fun_prop] [✅️] applying: Con_apply + [Meta.Tactic.fun_prop] [✅️] Con fun x => f x + [Meta.Tactic.fun_prop] candidate local theorems for f #[this : Con f] + [Meta.Tactic.fun_prop] [✅️] applying: this : Con f -/ #guard_msgs in example [Add α] (y : α): @@ -553,3 +555,68 @@ example [Add α] (y : α): have : Con f := by fun_prop set_option trace.Meta.Tactic.fun_prop true in fun_prop + + + +--- pefromance tests - mainly testing fast failure --- +------------------------------------------------------ + + +section PerformanceTests +-- set_option trace.Meta.Tactic.fun_prop true +-- set_option profiler true + +variable {R : Type*} [Add R] [∀ n, OfNat R n] +example (f : R → R) (hf : Con f) : + Con (fun x ↦ f (x + 3)) := by fun_prop -- succeeds in 5ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + 2) := by fun_prop -- succeeds in 6ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + (2 + f (x + 1))) := by fun_prop -- succeeds in 8ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + (2 + f (x + 1)) + x) := by fun_prop -- succeeds in 10ms +example (f : R → R) (hf : Con f) : + Con (fun x ↦ (f (x + 3)) + 2 + f (x + 1) + x + 1) := by fun_prop -- succeeds in 11ms + +-- This used to fail in exponentially increasing time, up to 6s for the last example +-- We set maxHearthbeats to 1000 such that the last three examples should fail if the exponential +-- blow happen again. +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ f (x + 3)) := by + fail_if_success fun_prop + apply silentSorry + +example (f : R → R) : + Con (fun x ↦ (f (x + 3)) + 2) := by + fail_if_success fun_prop + apply silentSorry + +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ ((f (x + 3)) + 2) + f (x + 1)) := by + fail_if_success fun_prop + apply silentSorry + +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ ((f (x + 3)) + 2) + f (x + 1) + x) := by + fail_if_success fun_prop + apply silentSorry + +set_option maxHeartbeats 1000 in +example (f : R → R) : + Con (fun x ↦ ((f (x + 3)) + 2) + f (x + 1) + x + 1) := by + fail_if_success fun_prop + apply silentSorry + +end PerformanceTests + + +/-- +info: Con + add_Con, args: [4, 5], form: simple + add_Con', args: [4, 5], form: compositional +-/ +#guard_msgs in +#print_fun_prop_theorems HAdd.hAdd Con From 1212c0adb03d7db70862a182885d0cf7ddd9dfed Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 19 Nov 2024 11:04:11 +0000 Subject: [PATCH 22/90] chore: some renames `iff_open` -> `iff_isOpen` (#19209) Our naming convention prefers `isOpen`. See the individual commit messages for details. --- .../Constructions/BorelSpace/Basic.lean | 2 +- .../Topology/Algebra/UniformGroup/Defs.lean | 7 ++++++- .../Topology/Connected/LocallyConnected.lean | 18 ++++++++++++------ Mathlib/Topology/Defs/Filter.lean | 4 ++-- Mathlib/Topology/Inseparable.lean | 15 ++++++++++----- Mathlib/Topology/Separation/Basic.lean | 4 ++-- Mathlib/Topology/Sets/Opens.lean | 2 +- Mathlib/Topology/UniformSpace/Basic.lean | 7 +++++-- 8 files changed, 39 insertions(+), 20 deletions(-) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index de06be836259b..23adc61b8b596 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -330,7 +330,7 @@ instance (priority := 100) OpensMeasurableSpace.separatesPoints [T0Space α] : rw [separatesPoints_iff] intro x y hxy apply Inseparable.eq - rw [inseparable_iff_forall_open] + rw [inseparable_iff_forall_isOpen] exact fun s hs => hxy _ hs.measurableSet theorem borel_eq_top_of_countable {α : Type*} [TopologicalSpace α] [T0Space α] [Countable α] : diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean index 1ebc4e5aa7b0d..b593b9d1eb3ca 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -324,7 +324,7 @@ theorem MonoidHom.uniformContinuous_of_continuousAt_one [UniformSpace β] [Group its kernel is open. -/ @[to_additive "A homomorphism from a uniform additive group to a discrete uniform additive group is continuous if and only if its kernel is open."] -theorem UniformGroup.uniformContinuous_iff_open_ker {hom : Type*} [UniformSpace β] +theorem UniformGroup.uniformContinuous_iff_isOpen_ker {hom : Type*} [UniformSpace β] [DiscreteTopology β] [Group β] [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} : UniformContinuous f ↔ IsOpen ((f : α →* β).ker : Set α) := by @@ -334,6 +334,11 @@ theorem UniformGroup.uniformContinuous_iff_open_ker {hom : Type*} [UniformSpace rw [ContinuousAt, nhds_discrete β, map_one, tendsto_pure] exact hf.mem_nhds (map_one f) +@[deprecated (since := "2024-11-18")] alias UniformGroup.uniformContinuous_iff_open_ker := + UniformGroup.uniformContinuous_iff_isOpen_ker +@[deprecated (since := "2024-11-18")] alias UniformAddGroup.uniformContinuous_iff_open_ker := + UniformAddGroup.uniformContinuous_iff_isOpen_ker + @[to_additive] theorem uniformContinuous_monoidHom_of_continuous {hom : Type*} [UniformSpace β] [Group β] [UniformGroup β] [FunLike hom α β] [MonoidHomClass hom α β] {f : hom} (h : Continuous f) : diff --git a/Mathlib/Topology/Connected/LocallyConnected.lean b/Mathlib/Topology/Connected/LocallyConnected.lean index ff3a0ebf66618..2e690ee0fbb2d 100644 --- a/Mathlib/Topology/Connected/LocallyConnected.lean +++ b/Mathlib/Topology/Connected/LocallyConnected.lean @@ -31,15 +31,18 @@ class LocallyConnectedSpace (α : Type*) [TopologicalSpace α] : Prop where /-- Open connected neighborhoods form a basis of the neighborhoods filter. -/ open_connected_basis : ∀ x, (𝓝 x).HasBasis (fun s : Set α => IsOpen s ∧ x ∈ s ∧ IsConnected s) id -theorem locallyConnectedSpace_iff_open_connected_basis : +theorem locallyConnectedSpace_iff_hasBasis_isOpen_isConnected : LocallyConnectedSpace α ↔ ∀ x, (𝓝 x).HasBasis (fun s : Set α => IsOpen s ∧ x ∈ s ∧ IsConnected s) id := ⟨@LocallyConnectedSpace.open_connected_basis _ _, LocallyConnectedSpace.mk⟩ -theorem locallyConnectedSpace_iff_open_connected_subsets : +@[deprecated (since := "2024-11-18")] alias locallyConnectedSpace_iff_open_connected_basis := + locallyConnectedSpace_iff_hasBasis_isOpen_isConnected + +theorem locallyConnectedSpace_iff_subsets_isOpen_isConnected : LocallyConnectedSpace α ↔ ∀ x, ∀ U ∈ 𝓝 x, ∃ V : Set α, V ⊆ U ∧ IsOpen V ∧ x ∈ V ∧ IsConnected V := by - simp_rw [locallyConnectedSpace_iff_open_connected_basis] + simp_rw [locallyConnectedSpace_iff_hasBasis_isOpen_isConnected] refine forall_congr' fun _ => ?_ constructor · intro h U hU @@ -49,10 +52,13 @@ theorem locallyConnectedSpace_iff_open_connected_subsets : let ⟨V, hVU, hV⟩ := h U hU ⟨V, hV, hVU⟩, fun ⟨V, ⟨hV, hxV, _⟩, hVU⟩ => mem_nhds_iff.mpr ⟨V, hVU, hV, hxV⟩⟩⟩ +@[deprecated (since := "2024-11-18")] alias locallyConnectedSpace_iff_open_connected_subsets := + locallyConnectedSpace_iff_subsets_isOpen_isConnected + /-- A space with discrete topology is a locally connected space. -/ instance (priority := 100) DiscreteTopology.toLocallyConnectedSpace (α) [TopologicalSpace α] [DiscreteTopology α] : LocallyConnectedSpace α := - locallyConnectedSpace_iff_open_connected_subsets.2 fun x _U hU => + locallyConnectedSpace_iff_subsets_isOpen_isConnected.2 fun x _U hU => ⟨{x}, singleton_subset_iff.2 <| mem_of_mem_nhds hU, isOpen_discrete _, rfl, isConnected_singleton⟩ @@ -85,7 +91,7 @@ theorem locallyConnectedSpace_iff_connectedComponentIn_open : · intro h exact fun F hF x _ => hF.connectedComponentIn · intro h - rw [locallyConnectedSpace_iff_open_connected_subsets] + rw [locallyConnectedSpace_iff_subsets_isOpen_isConnected] refine fun x U hU => ⟨connectedComponentIn (interior U) x, (connectedComponentIn_subset _ _).trans interior_subset, h _ isOpen_interior x ?_, @@ -95,7 +101,7 @@ theorem locallyConnectedSpace_iff_connectedComponentIn_open : theorem locallyConnectedSpace_iff_connected_subsets : LocallyConnectedSpace α ↔ ∀ (x : α), ∀ U ∈ 𝓝 x, ∃ V ∈ 𝓝 x, IsPreconnected V ∧ V ⊆ U := by constructor - · rw [locallyConnectedSpace_iff_open_connected_subsets] + · rw [locallyConnectedSpace_iff_subsets_isOpen_isConnected] intro h x U hxU rcases h x U hxU with ⟨V, hVU, hV₁, hxV, hV₂⟩ exact ⟨V, hV₁.mem_nhds hxV, hV₂.isPreconnected, hVU⟩ diff --git a/Mathlib/Topology/Defs/Filter.lean b/Mathlib/Topology/Defs/Filter.lean index 462aaeabab057..340ed360bb1da 100644 --- a/Mathlib/Topology/Defs/Filter.lean +++ b/Mathlib/Topology/Defs/Filter.lean @@ -202,8 +202,8 @@ infixl:300 " ⤳ " => Specializes equivalent properties hold: - `𝓝 x = 𝓝 y`; we use this property as the definition; -- for any open set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_open`; -- for any closed set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_closed`; +- for any open set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_forall_isOpen`; +- for any closed set `s`, `x ∈ s ↔ y ∈ s`, see `inseparable_iff_forall_isClosed`; - `x ∈ closure {y}` and `y ∈ closure {x}`, see `inseparable_iff_mem_closure`; - `closure {x} = closure {y}`, see `inseparable_iff_closure_eq`. -/ diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 2c7156a1d00f3..7b0bf5a54bf81 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -433,17 +433,22 @@ theorem Inseparable.specializes' (h : x ~ᵢ y) : y ⤳ x := h.ge theorem Specializes.antisymm (h₁ : x ⤳ y) (h₂ : y ⤳ x) : x ~ᵢ y := le_antisymm h₁ h₂ -theorem inseparable_iff_forall_open : (x ~ᵢ y) ↔ ∀ s : Set X, IsOpen s → (x ∈ s ↔ y ∈ s) := by +theorem inseparable_iff_forall_isOpen : (x ~ᵢ y) ↔ ∀ s : Set X, IsOpen s → (x ∈ s ↔ y ∈ s) := by simp only [inseparable_iff_specializes_and, specializes_iff_forall_open, ← forall_and, ← iff_def, Iff.comm] +@[deprecated (since := "2024-11-18")] alias +inseparable_iff_forall_open := inseparable_iff_forall_isOpen + theorem not_inseparable_iff_exists_open : ¬(x ~ᵢ y) ↔ ∃ s : Set X, IsOpen s ∧ Xor' (x ∈ s) (y ∈ s) := by - simp [inseparable_iff_forall_open, ← xor_iff_not_iff] + simp [inseparable_iff_forall_isOpen, ← xor_iff_not_iff] -theorem inseparable_iff_forall_closed : (x ~ᵢ y) ↔ ∀ s : Set X, IsClosed s → (x ∈ s ↔ y ∈ s) := by +theorem inseparable_iff_forall_isClosed : (x ~ᵢ y) ↔ ∀ s : Set X, IsClosed s → (x ∈ s ↔ y ∈ s) := by simp only [inseparable_iff_specializes_and, specializes_iff_forall_closed, ← forall_and, ← iff_def] +@[deprecated (since := "2024-11-18")] alias +inseparable_iff_forall_closed := inseparable_iff_forall_isClosed theorem inseparable_iff_mem_closure : (x ~ᵢ y) ↔ x ∈ closure ({y} : Set X) ∧ y ∈ closure ({x} : Set X) := @@ -497,10 +502,10 @@ nonrec theorem trans (h₁ : x ~ᵢ y) (h₂ : y ~ᵢ z) : x ~ᵢ z := h₁.tran theorem nhds_eq (h : x ~ᵢ y) : 𝓝 x = 𝓝 y := h theorem mem_open_iff (h : x ~ᵢ y) (hs : IsOpen s) : x ∈ s ↔ y ∈ s := - inseparable_iff_forall_open.1 h s hs + inseparable_iff_forall_isOpen.1 h s hs theorem mem_closed_iff (h : x ~ᵢ y) (hs : IsClosed s) : x ∈ s ↔ y ∈ s := - inseparable_iff_forall_closed.1 h s hs + inseparable_iff_forall_isClosed.1 h s hs theorem map_of_continuousAt (h : x ~ᵢ y) (hx : ContinuousAt f x) (hy : ContinuousAt f y) : f x ~ᵢ f y := diff --git a/Mathlib/Topology/Separation/Basic.lean b/Mathlib/Topology/Separation/Basic.lean index 7698efb12ea4d..b56a3a27aa9a9 100644 --- a/Mathlib/Topology/Separation/Basic.lean +++ b/Mathlib/Topology/Separation/Basic.lean @@ -316,7 +316,7 @@ theorem inseparable_eq_eq [T0Space X] : Inseparable = @Eq X := theorem TopologicalSpace.IsTopologicalBasis.inseparable_iff {b : Set (Set X)} (hb : IsTopologicalBasis b) {x y : X} : Inseparable x y ↔ ∀ s ∈ b, (x ∈ s ↔ y ∈ s) := - ⟨fun h _ hs ↦ inseparable_iff_forall_open.1 h _ (hb.isOpen hs), + ⟨fun h _ hs ↦ inseparable_iff_forall_isOpen.1 h _ (hb.isOpen hs), fun h ↦ hb.nhds_hasBasis.eq_of_same_basis <| by convert hb.nhds_hasBasis using 2 exact and_congr_right (h _)⟩ @@ -328,7 +328,7 @@ theorem TopologicalSpace.IsTopologicalBasis.eq_iff [T0Space X] {b : Set (Set X)} theorem t0Space_iff_exists_isOpen_xor'_mem (X : Type u) [TopologicalSpace X] : T0Space X ↔ Pairwise fun x y => ∃ U : Set X, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := by simp only [t0Space_iff_not_inseparable, xor_iff_not_iff, not_forall, exists_prop, - inseparable_iff_forall_open, Pairwise] + inseparable_iff_forall_isOpen, Pairwise] theorem exists_isOpen_xor'_mem [T0Space X] {x y : X} (h : x ≠ y) : ∃ U : Set X, IsOpen U ∧ Xor' (x ∈ U) (y ∈ U) := diff --git a/Mathlib/Topology/Sets/Opens.lean b/Mathlib/Topology/Sets/Opens.lean index 574a884fa980a..452c6d12ef50e 100644 --- a/Mathlib/Topology/Sets/Opens.lean +++ b/Mathlib/Topology/Sets/Opens.lean @@ -378,7 +378,7 @@ theorem comap_injective [T0Space β] : Injective (comap : C(α, β) → FrameHom fun f g h => ContinuousMap.ext fun a => Inseparable.eq <| - inseparable_iff_forall_open.2 fun s hs => + inseparable_iff_forall_isOpen.2 fun s hs => have : comap f ⟨s, hs⟩ = comap g ⟨s, hs⟩ := DFunLike.congr_fun h ⟨_, hs⟩ show a ∈ f ⁻¹' s ↔ a ∈ g ⁻¹' s from Set.ext_iff.1 (coe_inj.2 this) a diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 9bfe9b67a53a3..47f236596aa6b 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -651,7 +651,7 @@ theorem nhdsWithin_eq_comap_uniformity {x : α} (S : Set α) : 𝓝[S] x = (𝓤 α ⊓ 𝓟 (univ ×ˢ S)).comap (Prod.mk x) := nhdsWithin_eq_comap_uniformity_of_mem (mem_univ _) S -/-- See also `isOpen_iff_open_ball_subset`. -/ +/-- See also `isOpen_iff_isOpen_ball_subset`. -/ theorem isOpen_iff_ball_subset {s : Set α} : IsOpen s ↔ ∀ x ∈ s, ∃ V ∈ 𝓤 α, ball x V ⊆ s := by simp_rw [isOpen_iff_mem_nhds, nhds_eq_comap_uniformity, mem_comap, ball] @@ -863,7 +863,7 @@ theorem mem_uniformity_isClosed {s : Set (α × α)} (h : s ∈ 𝓤 α) : ∃ t let ⟨t, ⟨ht_mem, htc⟩, hts⟩ := uniformity_hasBasis_closed.mem_iff.1 h ⟨t, ht_mem, htc, hts⟩ -theorem isOpen_iff_open_ball_subset {s : Set α} : +theorem isOpen_iff_isOpen_ball_subset {s : Set α} : IsOpen s ↔ ∀ x ∈ s, ∃ V ∈ 𝓤 α, IsOpen V ∧ ball x V ⊆ s := by rw [isOpen_iff_ball_subset] constructor <;> intro h x hx @@ -874,6 +874,9 @@ theorem isOpen_iff_open_ball_subset {s : Set α} : · obtain ⟨V, hV, -, hV'⟩ := h x hx exact ⟨V, hV, hV'⟩ +@[deprecated (since := "2024-11-18")] alias +isOpen_iff_open_ball_subset := isOpen_iff_isOpen_ball_subset + /-- The uniform neighborhoods of all points of a dense set cover the whole space. -/ theorem Dense.biUnion_uniformity_ball {s : Set α} {U : Set (α × α)} (hs : Dense s) (hU : U ∈ 𝓤 α) : ⋃ x ∈ s, ball x U = univ := by From 442a60cbe707e04d02c26db93db0f22e7e19391c Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 19 Nov 2024 11:04:13 +0000 Subject: [PATCH 23/90] chore: update Mathlib dependencies 2024-11-19 (#19242) This PR updates the Mathlib dependencies. --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 2464526042495..cfc65dc4caaa9 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "01f4969b6e861db6a99261ea5eadd5a9bb63011b", + "rev": "485efbc439ee0ebdeae8afb0acd24a5e82e2f771", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From f7b68aa4cf35ead3a763da191bba0ff07de3f96c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 19 Nov 2024 11:44:13 +0000 Subject: [PATCH 24/90] feat(AlgebraicGeometry): ring hom is integral if the induced map between affine lines is closed (#18804) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib/Algebra/Polynomial/Reverse.lean | 8 ++ .../PrimeSpectrum/Basic.lean | 75 +++++++++++++++++-- 2 files changed, 76 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 505b6a0ea5d5a..07d080133f5ee 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -143,6 +143,14 @@ theorem reflect_monomial (N n : ℕ) : reflect N ((X : R[X]) ^ n) = X ^ revAt N @[simp] lemma reflect_one_X : reflect 1 (X : R[X]) = 1 := by simpa using reflect_monomial 1 1 (R := R) +lemma reflect_map {S : Type*} [Semiring S] (f : R →+* S) (p : R[X]) (n : ℕ) : + (p.map f).reflect n = (p.reflect n).map f := by + ext; simp + +@[simp] +lemma reflect_one (n : ℕ) : (1 : R[X]).reflect n = Polynomial.X ^ n := by + rw [← C.map_one, reflect_C, map_one, one_mul] + theorem reflect_mul_induction (cf cg : ℕ) : ∀ N O : ℕ, ∀ f g : R[X], diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index df037ab686364..b1fe379b4c1cc 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -321,13 +321,6 @@ theorem comap_singleton_isClosed_of_surjective (f : R →+* S) (hf : Function.Su haveI : x.asIdeal.IsMaximal := (isClosed_singleton_iff_isMaximal x).1 hx (isClosed_singleton_iff_isMaximal _).2 (Ideal.comap_isMaximal_of_surjective f hf) -theorem comap_singleton_isClosed_of_isIntegral (f : R →+* S) (hf : f.IsIntegral) - (x : PrimeSpectrum S) (hx : IsClosed ({x} : Set (PrimeSpectrum S))) : - IsClosed ({comap f x} : Set (PrimeSpectrum R)) := - have := (isClosed_singleton_iff_isMaximal x).1 hx - (isClosed_singleton_iff_isMaximal _).2 - (Ideal.isMaximal_comap_of_isIntegral_of_isMaximal' f hf x.asIdeal) - theorem image_comap_zeroLocus_eq_zeroLocus_comap (hf : Surjective f) (I : Ideal S) : comap f '' zeroLocus I = zeroLocus (I.comap f) := image_specComap_zeroLocus_eq_zeroLocus_comap _ f hf I @@ -719,6 +712,74 @@ lemma exists_idempotent_basicOpen_eq_of_is_clopen {s : Set (PrimeSpectrum R)} exact PrimeSpectrum.zeroLocus_anti_mono (Set.singleton_subset_iff.mpr <| Ideal.pow_le_self hnz hx) +section IsIntegral + +open Polynomial + +variable {R S : Type*} [CommRing R] [CommRing S] (f : R →+* S) + +theorem isClosedMap_comap_of_isIntegral (hf : f.IsIntegral) : + IsClosedMap (comap f) := by + refine fun s hs ↦ isClosed_image_of_stableUnderSpecialization _ _ hs ?_ + rintro _ y e ⟨x, hx, rfl⟩ + algebraize [f] + obtain ⟨q, hq₁, hq₂, hq₃⟩ := Ideal.exists_ideal_over_prime_of_isIntegral y.asIdeal x.asIdeal + ((le_iff_specializes _ _).mpr e) + refine ⟨⟨q, hq₂⟩, ((le_iff_specializes _ ⟨q, hq₂⟩).mp hq₁).mem_closed hs hx, + PrimeSpectrum.ext hq₃⟩ + +theorem isClosed_comap_singleton_of_isIntegral (hf : f.IsIntegral) + (x : PrimeSpectrum S) (hx : IsClosed ({x} : Set (PrimeSpectrum S))) : + IsClosed ({comap f x} : Set (PrimeSpectrum R)) := by + simpa using isClosedMap_comap_of_isIntegral f hf _ hx + +lemma closure_image_comap_zeroLocus (I : Ideal S) : + closure (comap f '' zeroLocus I) = zeroLocus (I.comap f) := by + apply subset_antisymm + · rw [(isClosed_zeroLocus _).closure_subset_iff, Set.image_subset_iff, preimage_comap_zeroLocus] + exact zeroLocus_anti_mono (Set.image_preimage_subset _ _) + · rintro x (hx : I.comap f ≤ x.asIdeal) + obtain ⟨q, hq₁, hq₂⟩ := Ideal.exists_minimalPrimes_le hx + obtain ⟨p', hp', hp'', rfl⟩ := Ideal.exists_comap_eq_of_mem_minimalPrimes f _ hq₁ + let p'' : PrimeSpectrum S := ⟨p', hp'⟩ + apply isClosed_closure.stableUnderSpecialization ((le_iff_specializes + (comap f ⟨p', hp'⟩) x).mp hq₂) (subset_closure (by exact ⟨_, hp'', rfl⟩)) + +lemma isIntegral_of_isClosedMap_comap_mapRingHom (h : IsClosedMap (comap (mapRingHom f))) : + f.IsIntegral := by + algebraize [f] + suffices Algebra.IsIntegral R S by rwa [Algebra.isIntegral_def] at this + nontriviality R + nontriviality S + constructor + intro r + let p : S[X] := C r * X - 1 + have : (1 : R[X]) ∈ Ideal.span {X} ⊔ (Ideal.span {p}).comap (mapRingHom f) := by + have H := h _ (isClosed_zeroLocus {p}) + rw [← zeroLocus_span, ← closure_eq_iff_isClosed, closure_image_comap_zeroLocus] at H + rw [← Ideal.eq_top_iff_one, sup_comm, ← zeroLocus_empty_iff_eq_top, zeroLocus_sup, H] + suffices ∀ (a : PrimeSpectrum S[X]), p ∈ a.asIdeal → X ∉ a.asIdeal by + simpa [Set.eq_empty_iff_forall_not_mem] + intro q hpq hXq + have : 1 ∈ q.asIdeal := by simpa [p] using (sub_mem (q.asIdeal.mul_mem_left (C r) hXq) hpq) + exact q.2.ne_top (q.asIdeal.eq_top_iff_one.mpr this) + obtain ⟨a, b, hb, e⟩ := Ideal.mem_span_singleton_sup.mp this + obtain ⟨c, hc : b.map (algebraMap R S) = _⟩ := Ideal.mem_span_singleton.mp hb + refine ⟨b.reverse * X ^ (1 + c.natDegree), ?_, ?_⟩ + · refine Monic.mul ?_ (by simp) + have h : b.coeff 0 = 1 := by simpa using congr(($e).coeff 0) + have : b.natTrailingDegree = 0 := by simp [h] + rw [Monic.def, reverse_leadingCoeff, trailingCoeff, this, h] + · have : p.natDegree ≤ 1 := by simpa using natDegree_linear_le (a := r) (b := -1) + rw [eval₂_eq_eval_map, reverse, Polynomial.map_mul, ← reflect_map, Polynomial.map_pow, + map_X, ← revAt_zero (1 + _), ← reflect_monomial, + ← reflect_mul _ _ (natDegree_map_le _ _) (by simp), pow_zero, mul_one, hc, + ← add_assoc, reflect_mul _ _ (this.trans (by simp)) le_rfl, + eval_mul, reflect_sub, reflect_mul _ _ (by simp) (by simp)] + simp [← pow_succ'] + +end IsIntegral + section LocalizationAtMinimal variable {I : Ideal R} [hI : I.IsPrime] From d6d71136020054d771903093d0600839425127b8 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 19 Nov 2024 13:00:52 +0000 Subject: [PATCH 25/90] fix: replace backticks with single quotes in maintainer merge action (#19228) This should prevent mangling the `spoiler` message on Zulip. --- scripts/maintainer_merge_message.sh | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/scripts/maintainer_merge_message.sh b/scripts/maintainer_merge_message.sh index db5a50caa1676..ec4e14808bf86 100755 --- a/scripts/maintainer_merge_message.sh +++ b/scripts/maintainer_merge_message.sh @@ -36,5 +36,10 @@ esac >&2 echo "EVENT_NAME: '${EVENT_NAME}'" >&2 printf 'COMMENT\n%s\nEND COMMENT\n' "${PR_COMMENT}" +# replace backticks in the title with single quotes +unbacktickedTitle="${PR_TITLE//\`/\'}" + +>&2 echo "neat title: '${unbacktickedTitle}'" + printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D}" "${SOURCE}" "${PR}" "${URL}" -printf '```spoiler %s\n%s\n```\n' "${PR_TITLE}" "${PR_COMMENT}" +printf '```spoiler %s\n%s\n```\n' "${unbacktickedTitle}" "${PR_COMMENT}" From 62b241e911f53062101c73e0d8ec9a31076c0916 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 19 Nov 2024 15:12:10 +0000 Subject: [PATCH 26/90] feat(AlgebraicGeometry): define integral morphisms (#19121) also shows that finite = integral + locally of finite type and closed immersion = finte + mono Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../AlgebraicGeometry/Morphisms/Finite.lean | 60 +++++++++++++---- .../AlgebraicGeometry/Morphisms/Integral.lean | 64 +++++++++++++++++++ Mathlib/RingTheory/RingHom/Integral.lean | 32 +++++++++- 4 files changed, 143 insertions(+), 14 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/Integral.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7566572fa39e6..49dd52b773d65 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -938,6 +938,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.Finite import Mathlib.AlgebraicGeometry.Morphisms.FinitePresentation import Mathlib.AlgebraicGeometry.Morphisms.FiniteType import Mathlib.AlgebraicGeometry.Morphisms.Immersion +import Mathlib.AlgebraicGeometry.Morphisms.Integral import Mathlib.AlgebraicGeometry.Morphisms.IsIso import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean index 3a1aca10b2cbd..0f054e90a059e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean @@ -3,9 +3,8 @@ Copyright (c) 2024 Christian Merten. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christian Merten -/ -import Mathlib.AlgebraicGeometry.Morphisms.AffineAnd -import Mathlib.AlgebraicGeometry.Morphisms.FiniteType -import Mathlib.RingTheory.RingHom.Finite +import Mathlib.AlgebraicGeometry.Morphisms.Integral +import Mathlib.Algebra.Category.Ring.Epi /-! @@ -65,15 +64,52 @@ instance (priority := 900) [IsIso f] : IsFinite f := of_isIso @IsFinite f instance {Z : Scheme.{u}} (g : Y ⟶ Z) [IsFinite f] [IsFinite g] : IsFinite (f ≫ g) := IsStableUnderComposition.comp_mem f g ‹IsFinite f› ‹IsFinite g› -instance (priority := 900) [hf : IsFinite f] : LocallyOfFiniteType f := by - have : targetAffineLocally (affineAnd @RingHom.FiniteType) f := by - rw [HasAffineProperty.eq_targetAffineLocally (P := @IsFinite)] at hf - apply targetAffineLocally_affineAnd_le RingHom.Finite.finiteType - exact hf - rw [targetAffineLocally_affineAnd_eq_affineLocally - (HasRingHomProperty.isLocal_ringHomProperty @LocallyOfFiniteType)] at this - rw [HasRingHomProperty.eq_affineLocally (P := @LocallyOfFiniteType)] - exact this.2 +lemma iff_isIntegralHom_and_locallyOfFiniteType : + IsFinite f ↔ IsIntegralHom f ∧ LocallyOfFiniteType f := by + wlog hY : IsAffine Y + · rw [IsLocalAtTarget.iff_of_openCover (P := @IsFinite) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := @IsIntegralHom) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := @LocallyOfFiniteType) Y.affineCover] + simp_rw [this, forall_and] + rw [HasAffineProperty.iff_of_isAffine (P := @IsFinite), + HasAffineProperty.iff_of_isAffine (P := @IsIntegralHom), + RingHom.finite_iff_isIntegral_and_finiteType, ← and_assoc] + refine and_congr_right fun ⟨_, _⟩ ↦ + (HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)).symm + +lemma eq_inf : + @IsFinite = (@IsIntegralHom ⊓ @LocallyOfFiniteType : MorphismProperty Scheme) := by + ext; exact IsFinite.iff_isIntegralHom_and_locallyOfFiniteType _ + +instance (priority := 900) [IsFinite f] : IsIntegralHom f := + ((IsFinite.iff_isIntegralHom_and_locallyOfFiniteType f).mp ‹_›).1 + +instance (priority := 900) [hf : IsFinite f] : LocallyOfFiniteType f := + ((IsFinite.iff_isIntegralHom_and_locallyOfFiniteType f).mp ‹_›).2 + +lemma _root_.AlgebraicGeometry.IsClosedImmersion.iff_isFinite_and_mono : + IsClosedImmersion f ↔ IsFinite f ∧ Mono f := by + wlog hY : IsAffine Y + · show _ ↔ _ ∧ monomorphisms _ f + rw [IsLocalAtTarget.iff_of_openCover (P := @IsFinite) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := @IsClosedImmersion) Y.affineCover, + IsLocalAtTarget.iff_of_openCover (P := monomorphisms _) Y.affineCover] + simp_rw [this, forall_and, monomorphisms] + rw [HasAffineProperty.iff_of_isAffine (P := @IsClosedImmersion), + HasAffineProperty.iff_of_isAffine (P := @IsFinite), + RingHom.surjective_iff_epi_and_finite, @and_comm (Epi _), ← and_assoc] + refine and_congr_right fun ⟨_, _⟩ ↦ + Iff.trans ?_ (arrow_mk_iso_iff (monomorphisms _) (arrowIsoSpecΓOfIsAffine f).symm) + trans Mono (f.app ⊤).op + · exact ⟨fun h ↦ inferInstance, fun h ↦ show Epi (f.app ⊤).op.unop by infer_instance⟩ + exact (Functor.mono_map_iff_mono Scheme.Spec _).symm + +lemma _root_.AlgebraicGeometry.IsClosedImmersion.eq_isFinite_inf_mono : + @IsClosedImmersion = (@IsFinite ⊓ monomorphisms Scheme : MorphismProperty _) := by + ext; exact IsClosedImmersion.iff_isFinite_and_mono _ + +instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsClosedImmersion f] : IsFinite f := + ((IsClosedImmersion.iff_isFinite_and_mono f).mp ‹_›).1 end IsFinite diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean new file mode 100644 index 0000000000000..09ee678d31470 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/Integral.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.AffineAnd +import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion +import Mathlib.RingTheory.RingHom.Integral + +/-! + +# Integral morphisms of schemes + +A morphism of schemes `f : X ⟶ Y` is integral if the preimage +of an arbitrary affine open subset of `Y` is affine and the induced ring map is integral. + +It is equivalent to ask only that `Y` is covered by affine opens whose preimage is affine +and the induced ring map is integral. + +## TODO + +- Show integral = universally closed + affine + +-/ + +universe v u + +open CategoryTheory TopologicalSpace Opposite MorphismProperty + +namespace AlgebraicGeometry + +/-- A morphism of schemes `X ⟶ Y` is finite if +the preimage of any affine open subset of `Y` is affine and the induced ring +hom is finite. -/ +@[mk_iff] +class IsIntegralHom {X Y : Scheme} (f : X ⟶ Y) extends IsAffineHom f : Prop where + integral_app (U : Y.Opens) (hU : IsAffineOpen U) : (f.app U).IsIntegral + +namespace IsIntegralHom + +instance hasAffineProperty : HasAffineProperty @IsIntegralHom + fun X _ f _ ↦ IsAffine X ∧ RingHom.IsIntegral (f.app ⊤) := by + show HasAffineProperty @IsIntegralHom (affineAnd RingHom.IsIntegral) + rw [HasAffineProperty.affineAnd_iff _ RingHom.isIntegral_respectsIso + RingHom.isIntegral_isStableUnderBaseChange.localizationPreserves + RingHom.isIntegral_ofLocalizationSpan] + simp [isIntegralHom_iff] + +instance : IsStableUnderComposition @IsIntegralHom := + HasAffineProperty.affineAnd_isStableUnderComposition (Q := RingHom.IsIntegral) hasAffineProperty + RingHom.isIntegral_stableUnderComposition + +instance : IsStableUnderBaseChange @IsIntegralHom := + HasAffineProperty.affineAnd_isStableUnderBaseChange (Q := RingHom.IsIntegral) hasAffineProperty + RingHom.isIntegral_respectsIso RingHom.isIntegral_isStableUnderBaseChange + +instance : ContainsIdentities @IsIntegralHom := + ⟨fun X ↦ ⟨fun _ _ ↦ by simpa using RingHom.isIntegral_of_surjective _ (Equiv.refl _).surjective⟩⟩ + +instance : IsMultiplicative @IsIntegralHom where + +end IsIntegralHom + +end AlgebraicGeometry diff --git a/Mathlib/RingTheory/RingHom/Integral.lean b/Mathlib/RingTheory/RingHom/Integral.lean index 62662762e06d7..2662ea9b1ccba 100644 --- a/Mathlib/RingTheory/RingHom/Integral.lean +++ b/Mathlib/RingTheory/RingHom/Integral.lean @@ -3,8 +3,8 @@ 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.RingTheory.RingHomProperties -import Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic +import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.RingTheory.Localization.Integral /-! @@ -36,4 +36,32 @@ theorem isIntegral_isStableUnderBaseChange : IsStableUnderBaseChange fun f => f. · intro x y; exact IsIntegral.tmul x (h y) · intro x y hx hy; exact IsIntegral.add hx hy +open Polynomial in +/-- `S` is an integral `R`-algebra if there exists a set `{ r }` that + spans `R` such that each `Sᵣ` is an integral `Rᵣ`-algebra. -/ +theorem isIntegral_ofLocalizationSpan : + OfLocalizationSpan (RingHom.IsIntegral ·) := by + introv R hs H r + letI := f.toAlgebra + show r ∈ (integralClosure R S).toSubmodule + apply Submodule.mem_of_span_eq_top_of_smul_pow_mem _ s hs + rintro ⟨t, ht⟩ + letI := (Localization.awayMap f t).toAlgebra + haveI : IsScalarTower R (Localization.Away t) (Localization.Away (f t)) := .of_algebraMap_eq' + (IsLocalization.lift_comp _).symm + have : _root_.IsIntegral (Localization.Away t) (algebraMap S (Localization.Away (f t)) r) := + H ⟨t, ht⟩ (algebraMap _ _ r) + obtain ⟨⟨_, n, rfl⟩, p, hp, hp'⟩ := this.exists_multiple_integral_of_isLocalization (.powers t) + rw [IsScalarTower.algebraMap_eq R S, Submonoid.smul_def, Algebra.smul_def, + IsScalarTower.algebraMap_apply R S, ← map_mul, ← hom_eval₂, + IsLocalization.map_eq_zero_iff (.powers (f t))] at hp' + obtain ⟨⟨x, m, (rfl : algebraMap R S t ^ m = x)⟩, e⟩ := hp' + by_cases hp' : 1 ≤ p.natDegree; swap + · obtain rfl : p = 1 := eq_one_of_monic_natDegree_zero hp (by nlinarith) + exact ⟨m, by simp [Algebra.smul_def, show algebraMap R S t ^ m = 0 by simpa using e]⟩ + refine ⟨m + n, p.scaleRoots (t ^ m), (monic_scaleRoots_iff _).mpr hp, ?_⟩ + have := p.scaleRoots_eval₂_mul (algebraMap R S) (t ^ n • r) (t ^ m) + simp only [pow_add, ← Algebra.smul_def, mul_smul, ← map_pow] at e this ⊢ + rw [this, ← tsub_add_cancel_of_le hp', pow_succ, mul_smul, e, smul_zero] + end RingHom From e948d859635b67a2d40b53d10bcf906226e2a460 Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 19 Nov 2024 15:27:44 +0000 Subject: [PATCH 27/90] chore: further renames (#19243) Similar to #19209: bring more lemma names in line with the naming convention. --- .../PrimeSpectrum/Basic.lean | 2 +- Mathlib/AlgebraicGeometry/Properties.lean | 2 +- Mathlib/Dynamics/Minimal.lean | 4 +++- Mathlib/MeasureTheory/Group/Action.lean | 2 +- Mathlib/MeasureTheory/Group/Measure.lean | 4 ++-- Mathlib/MeasureTheory/Measure/Regular.lean | 6 ++++-- Mathlib/Topology/ContinuousMap/Algebra.lean | 10 +++++----- Mathlib/Topology/Irreducible.lean | 9 +++++++-- Mathlib/Topology/Maps/Basic.lean | 10 ++++++---- Mathlib/Topology/NoetherianSpace.lean | 6 +++--- .../UniformSpace/CompactConvergence.lean | 19 +++++++++++-------- 11 files changed, 44 insertions(+), 30 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index b1fe379b4c1cc..8121592d7b482 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -156,7 +156,7 @@ theorem isIrreducible_zeroLocus_iff_of_radical (I : Ideal R) (hI : I.IsRadical) apply and_congr · rw [Set.nonempty_iff_ne_empty, Ne, zeroLocus_empty_iff_eq_top] · trans ∀ x y : Ideal R, Z(I) ⊆ Z(x) ∪ Z(y) → Z(I) ⊆ Z(x) ∨ Z(I) ⊆ Z(y) - · simp_rw [isPreirreducible_iff_closed_union_closed, isClosed_iff_zeroLocus_ideal] + · simp_rw [isPreirreducible_iff_isClosed_union_isClosed, isClosed_iff_zeroLocus_ideal] constructor · rintro h x y exact h _ _ ⟨x, rfl⟩ ⟨y, rfl⟩ diff --git a/Mathlib/AlgebraicGeometry/Properties.lean b/Mathlib/AlgebraicGeometry/Properties.lean index 800d80f953765..79091bf28dd7a 100644 --- a/Mathlib/AlgebraicGeometry/Properties.lean +++ b/Mathlib/AlgebraicGeometry/Properties.lean @@ -203,7 +203,7 @@ instance irreducibleSpace_of_isIntegral [IsIntegral X] : IrreducibleSpace X := b replace H : ¬IsPreirreducible (⊤ : Set X) := fun h => H { toPreirreducibleSpace := ⟨h⟩ toNonempty := inferInstance } - simp_rw [isPreirreducible_iff_closed_union_closed, not_forall, not_or] at H + simp_rw [isPreirreducible_iff_isClosed_union_isClosed, not_forall, not_or] at H rcases H with ⟨S, T, hS, hT, h₁, h₂, h₃⟩ erw [not_forall] at h₂ h₃ simp_rw [not_forall] at h₂ h₃ diff --git a/Mathlib/Dynamics/Minimal.lean b/Mathlib/Dynamics/Minimal.lean index ce019b475eaf4..c29a7f2fce92e 100644 --- a/Mathlib/Dynamics/Minimal.lean +++ b/Mathlib/Dynamics/Minimal.lean @@ -93,7 +93,7 @@ theorem eq_empty_or_univ_of_smul_invariant_closed [IsMinimal M α] {s : Set α} hs.closure_eq ▸ (dense_of_nonempty_smul_invariant M hne hsmul).closure_eq @[to_additive] -theorem isMinimal_iff_closed_smul_invariant [ContinuousConstSMul M α] : +theorem isMinimal_iff_isClosed_smul_invariant [ContinuousConstSMul M α] : IsMinimal M α ↔ ∀ s : Set α, IsClosed s → (∀ c : M, c • s ⊆ s) → s = ∅ ∨ s = univ := by constructor · intro _ _ @@ -101,3 +101,5 @@ theorem isMinimal_iff_closed_smul_invariant [ContinuousConstSMul M α] : refine fun H ↦ ⟨fun _ ↦ dense_iff_closure_eq.2 <| (H _ ?_ ?_).resolve_left ?_⟩ exacts [isClosed_closure, fun _ ↦ smul_closure_orbit_subset _ _, (orbit_nonempty _).closure.ne_empty] +@[deprecated (since := "2024-11-19")] alias +isMinimal_iff_closed_smul_invariant := isMinimal_iff_isClosed_smul_invariant diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 04f432bce67b9..0cb212b5273b1 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -307,7 +307,7 @@ variable [Measure.Regular μ] @[to_additive] theorem measure_isOpen_pos_of_smulInvariant_of_ne_zero (hμ : μ ≠ 0) (hU : IsOpen U) (hne : U.Nonempty) : 0 < μ U := - let ⟨_K, hK, hμK⟩ := Regular.exists_compact_not_null.mpr hμ + let ⟨_K, hK, hμK⟩ := Regular.exists_isCompact_not_null.mpr hμ measure_isOpen_pos_of_smulInvariant_of_compact_ne_zero G hK hμK hU hne @[to_additive] diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 05604e5fa4945..8662ff3b7f53c 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -502,7 +502,7 @@ theorem isOpenPosMeasure_of_mulLeftInvariant_of_compact (K : Set G) (hK : IsComp @[to_additive "A nonzero left-invariant regular measure gives positive mass to any open set."] instance (priority := 80) isOpenPosMeasure_of_mulLeftInvariant_of_regular [Regular μ] [NeZero μ] : IsOpenPosMeasure μ := - let ⟨K, hK, h2K⟩ := Regular.exists_compact_not_null.mpr (NeZero.ne μ) + let ⟨K, hK, h2K⟩ := Regular.exists_isCompact_not_null.mpr (NeZero.ne μ) isOpenPosMeasure_of_mulLeftInvariant_of_compact K hK h2K /-- A nonzero left-invariant inner regular measure gives positive mass to any open set. -/ @@ -510,7 +510,7 @@ instance (priority := 80) isOpenPosMeasure_of_mulLeftInvariant_of_regular [Regul instance (priority := 80) isOpenPosMeasure_of_mulLeftInvariant_of_innerRegular [InnerRegular μ] [NeZero μ] : IsOpenPosMeasure μ := - let ⟨K, hK, h2K⟩ := InnerRegular.exists_compact_not_null.mpr (NeZero.ne μ) + let ⟨K, hK, h2K⟩ := InnerRegular.exists_isCompact_not_null.mpr (NeZero.ne μ) isOpenPosMeasure_of_mulLeftInvariant_of_compact K hK h2K @[to_additive] diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index 00bbeafec255e..ed27aa6936522 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -663,9 +663,10 @@ lemma innerRegularWRT_isClosed_isOpen [R1Space α] [OpensMeasurableSpace α] [h exact ⟨closure K, K_comp.closure_subset_of_isOpen hU KU, isClosed_closure, hK.trans_le (measure_mono subset_closure)⟩ -theorem exists_compact_not_null [InnerRegular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by +theorem exists_isCompact_not_null [InnerRegular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by simp_rw [Ne, ← measure_univ_eq_zero, MeasurableSet.univ.measure_eq_iSup_isCompact, ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and] +@[deprecated (since := "2024-11-19")] alias exists_compact_not_null := exists_isCompact_not_null /-- If `μ` is inner regular, then any measurable set can be approximated by a compact subset. See also `MeasurableSet.exists_isCompact_lt_add_of_ne_top`. -/ @@ -986,9 +987,10 @@ theorem _root_.IsOpen.measure_eq_iSup_isCompact ⦃U : Set α⦄ (hU : IsOpen U) [Regular μ] : μ U = ⨆ (K : Set α) (_ : K ⊆ U) (_ : IsCompact K), μ K := Regular.innerRegular.measure_eq_iSup hU -theorem exists_compact_not_null [Regular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by +theorem exists_isCompact_not_null [Regular μ] : (∃ K, IsCompact K ∧ μ K ≠ 0) ↔ μ ≠ 0 := by simp_rw [Ne, ← measure_univ_eq_zero, isOpen_univ.measure_eq_iSup_isCompact, ENNReal.iSup_eq_zero, not_forall, exists_prop, subset_univ, true_and] +@[deprecated (since := "2024-11-19")] alias exists_compact_not_null := exists_isCompact_not_null /-- If `μ` is a regular measure, then any measurable set of finite measure can be approximated by a compact subset. See also `MeasurableSet.exists_isCompact_lt_add` and diff --git a/Mathlib/Topology/ContinuousMap/Algebra.lean b/Mathlib/Topology/ContinuousMap/Algebra.lean index 2cefc72b970be..42a37d593e9a0 100644 --- a/Mathlib/Topology/ContinuousMap/Algebra.lean +++ b/Mathlib/Topology/ContinuousMap/Algebra.lean @@ -351,20 +351,20 @@ instance [CommGroup β] [TopologicalGroup β] : TopologicalGroup C(α, β) where have : UniformGroup β := comm_topologicalGroup_is_uniform rw [continuous_iff_continuousAt] rintro ⟨f, g⟩ - rw [ContinuousAt, tendsto_iff_forall_compact_tendstoUniformlyOn, nhds_prod_eq] + rw [ContinuousAt, tendsto_iff_forall_isCompact_tendstoUniformlyOn, nhds_prod_eq] exact fun K hK => uniformContinuous_mul.comp_tendstoUniformlyOn - ((tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK).prod - (tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK)) + ((tendsto_iff_forall_isCompact_tendstoUniformlyOn.mp Filter.tendsto_id K hK).prod + (tendsto_iff_forall_isCompact_tendstoUniformlyOn.mp Filter.tendsto_id K hK)) continuous_inv := by letI : UniformSpace β := TopologicalGroup.toUniformSpace β have : UniformGroup β := comm_topologicalGroup_is_uniform rw [continuous_iff_continuousAt] intro f - rw [ContinuousAt, tendsto_iff_forall_compact_tendstoUniformlyOn] + rw [ContinuousAt, tendsto_iff_forall_isCompact_tendstoUniformlyOn] exact fun K hK => uniformContinuous_inv.comp_tendstoUniformlyOn - (tendsto_iff_forall_compact_tendstoUniformlyOn.mp Filter.tendsto_id K hK) + (tendsto_iff_forall_isCompact_tendstoUniformlyOn.mp Filter.tendsto_id K hK) /-- If an infinite product of functions in `C(α, β)` converges to `g` (for the compact-open topology), then the pointwise product converges to `g x` for all `x ∈ α`. -/ diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index c5c1339a53ef9..5d383423319f0 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -240,7 +240,7 @@ theorem isIrreducible_iff_sInter : /-- A set is preirreducible if and only if for every cover by two closed sets, it is contained in one of the two covering sets. -/ -theorem isPreirreducible_iff_closed_union_closed : +theorem isPreirreducible_iff_isClosed_union_isClosed : IsPreirreducible s ↔ ∀ z₁ z₂ : Set X, IsClosed z₁ → IsClosed z₂ → s ⊆ z₁ ∪ z₂ → s ⊆ z₁ ∨ s ⊆ z₂ := by refine compl_surjective.forall.trans <| forall_congr' fun z₁ => compl_surjective.forall.trans <| @@ -248,10 +248,12 @@ theorem isPreirreducible_iff_closed_union_closed : simp only [isOpen_compl_iff, ← compl_union, inter_compl_nonempty_iff] refine forall₂_congr fun _ _ => ?_ rw [← and_imp, ← not_or, not_imp_not] +@[deprecated (since := "2024-11-19")] alias +isPreirreducible_iff_closed_union_closed := isPreirreducible_iff_isClosed_union_isClosed /-- A set is irreducible if and only if for every cover by a finite collection of closed sets, it is contained in one of the members of the collection. -/ -theorem isIrreducible_iff_sUnion_closed : +theorem isIrreducible_iff_sUnion_isClosed : IsIrreducible s ↔ ∀ t : Finset (Set X), (∀ z ∈ t, IsClosed z) → (s ⊆ ⋃₀ ↑t) → ∃ z ∈ t, s ⊆ z := by simp only [isIrreducible_iff_sInter] @@ -263,6 +265,9 @@ theorem isIrreducible_iff_sUnion_closed : simp only [not_exists, not_and, ← compl_iInter₂, ← sInter_eq_biInter, subset_compl_iff_disjoint_right, not_disjoint_iff_nonempty_inter] +@[deprecated (since := "2024-11-19")] alias +isIrreducible_iff_sUnion_closed := isIrreducible_iff_sUnion_isClosed + /-- A nonempty open subset of a preirreducible subspace is dense in the subspace. -/ theorem subset_closure_inter_of_isPreirreducible_of_isOpen {S U : Set X} (hS : IsPreirreducible S) (hU : IsOpen U) (h : (S ∩ U).Nonempty) : S ⊆ closure (S ∩ U) := by diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index b71fbe5d98fd9..077466f86f877 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -286,13 +286,15 @@ lemma isQuotientMap_iff : IsQuotientMap f ↔ Surjective f ∧ ∀ s, IsOpen s @[deprecated (since := "2024-10-22")] alias quotientMap_iff := isQuotientMap_iff -theorem isQuotientMap_iff_closed : +theorem isQuotientMap_iff_isClosed : IsQuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsClosed s ↔ IsClosed (f ⁻¹' s) := isQuotientMap_iff.trans <| Iff.rfl.and <| compl_surjective.forall.trans <| by simp only [isOpen_compl_iff, preimage_compl] @[deprecated (since := "2024-10-22")] -alias quotientMap_iff_closed := isQuotientMap_iff_closed +alias quotientMap_iff_closed := isQuotientMap_iff_isClosed +@[deprecated (since := "2024-11-19")] +alias isQuotientMap_iff_closed := isQuotientMap_iff_isClosed namespace IsQuotientMap @@ -326,7 +328,7 @@ protected lemma isOpen_preimage (hf : IsQuotientMap f) {s : Set Y} : IsOpen (f protected theorem isClosed_preimage (hf : IsQuotientMap f) {s : Set Y} : IsClosed (f ⁻¹' s) ↔ IsClosed s := - ((isQuotientMap_iff_closed.1 hf).2 s).symm + ((isQuotientMap_iff_isClosed.1 hf).2 s).symm end IsQuotientMap @@ -489,7 +491,7 @@ theorem isClosed_range (hf : IsClosedMap f) : IsClosed (range f) := theorem isQuotientMap (hcl : IsClosedMap f) (hcont : Continuous f) (hsurj : Surjective f) : IsQuotientMap f := - isQuotientMap_iff_closed.2 ⟨hsurj, fun s => + isQuotientMap_iff_isClosed.2 ⟨hsurj, fun s => ⟨fun hs => hs.preimage hcont, fun hs => hsurj.image_preimage s ▸ hcl _ hs⟩⟩ @[deprecated (since := "2024-10-22")] diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 4008eacfe3e3f..64689fd04f3f0 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -165,7 +165,7 @@ theorem NoetherianSpace.exists_finite_set_closeds_irreducible [NoetherianSpace · by_cases h₁ : IsPreirreducible (s : Set α) · replace h₁ : IsIrreducible (s : Set α) := ⟨Closeds.coe_nonempty.2 h₀, h₁⟩ use {s}; simp [h₁] - · simp only [isPreirreducible_iff_closed_union_closed, not_forall, not_or] at h₁ + · simp only [isPreirreducible_iff_isClosed_union_isClosed, not_forall, not_or] at h₁ obtain ⟨z₁, z₂, hz₁, hz₂, h, hz₁', hz₂'⟩ := h₁ lift z₁ to Closeds α using hz₁ lift z₂ to Closeds α using hz₂ @@ -198,7 +198,7 @@ theorem NoetherianSpace.finite_irreducibleComponents [NoetherianSpace α] : NoetherianSpace.exists_finite_set_isClosed_irreducible isClosed_univ (α := α) refine hSf.subset fun s hs => ?_ lift S to Finset (Set α) using hSf - rcases isIrreducible_iff_sUnion_closed.1 hs.1 S hSc (hSU ▸ Set.subset_univ _) with ⟨t, htS, ht⟩ + rcases isIrreducible_iff_sUnion_isClosed.1 hs.1 S hSc (hSU ▸ Set.subset_univ _) with ⟨t, htS, ht⟩ rwa [ht.antisymm (hs.2 (hSi _ htS) ht)] /-- [Stacks: Lemma 0052 (3)](https://stacks.math.columbia.edu/tag/0052) -/ @@ -213,7 +213,7 @@ theorem NoetherianSpace.exists_open_ne_empty_le_irreducibleComponent [Noetherian let U := Z \ ⋃ (x : ι), x have hU0 : U ≠ ∅ := fun r ↦ by - obtain ⟨Z', hZ'⟩ := isIrreducible_iff_sUnion_closed.mp H.1 hι.toFinset + obtain ⟨Z', hZ'⟩ := isIrreducible_iff_sUnion_isClosed.mp H.1 hι.toFinset (fun z hz ↦ by simp only [Set.Finite.mem_toFinset, Set.mem_diff, Set.mem_singleton_iff] at hz exact isClosed_of_mem_irreducibleComponents _ hz.1) diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index 0a9a0c7520df8..e39e65d2fbda4 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -48,12 +48,12 @@ and also prove its basic properties. a similar statement that uses a basis of entourages of `β` instead of all entourages. It is useful, e.g., if `β` is a metric space. -* `ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn`: +* `ContinuousMap.tendsto_iff_forall_isCompact_tendstoUniformlyOn`: a sequence of functions `Fₙ` in `C(α, β)` converges in the compact-open topology to some `f` iff `Fₙ` converges to `f` uniformly on each compact subset `K` of `α`. * Topology induced by the uniformity described above agrees with the compact-open topology. - This is essentially the same as `ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn`. + This is essentially the same as `ContinuousMap.tendsto_iff_forall_isCompact_tendstoUniformlyOn`. This fact is not available as a separate theorem. Instead, we override the projection of `ContinuousMap.compactConvergenceUniformity` @@ -95,7 +95,7 @@ namespace ContinuousMap /-- Compact-open topology on `C(α, β)` agrees with the topology of uniform convergence on compacts: a family of continuous functions `F i` tends to `f` in the compact-open topology if and only if the `F i` tends to `f` uniformly on all compact sets. -/ -theorem tendsto_iff_forall_compact_tendstoUniformlyOn +theorem tendsto_iff_forall_isCompact_tendstoUniformlyOn {ι : Type u₃} {p : Filter ι} {F : ι → C(α, β)} {f} : Tendsto F p (𝓝 f) ↔ ∀ K, IsCompact K → TendstoUniformlyOn (fun i a => F i a) f p K := by rw [tendsto_nhds_compactOpen] @@ -138,6 +138,9 @@ theorem tendsto_iff_forall_compact_tendstoUniformlyOn -- maps `K` to `U` as well filter_upwards [h K hK V hV] with g hg x hx using hVf _ (mem_image_of_mem f hx) (hg x hx) +@[deprecated (since := "2024-11-19")] alias +tendsto_iff_forall_compact_tendstoUniformlyOn := tendsto_iff_forall_isCompact_tendstoUniformlyOn + /-- Interpret a bundled continuous map as an element of `α →ᵤ[{K | IsCompact K}] β`. We use this map to induce the `UniformSpace` structure on `C(α, β)`. -/ @@ -157,14 +160,14 @@ open UniformSpace in The uniformity comes from `α →ᵤ[{K | IsCompact K}] β` (i.e., `UniformOnFun α β {K | IsCompact K}`) which defines topology of uniform convergence on compact sets. -We use `ContinuousMap.tendsto_iff_forall_compact_tendstoUniformlyOn` +We use `ContinuousMap.tendsto_iff_forall_isCompact_tendstoUniformlyOn` to show that the induced topology agrees with the compact-open topology and replace the topology with `compactOpen` to avoid non-defeq diamonds, see Note [forgetful inheritance]. -/ instance compactConvergenceUniformSpace : UniformSpace C(α, β) := .replaceTopology (.comap toUniformOnFunIsCompact inferInstance) <| by refine TopologicalSpace.ext_nhds fun f ↦ eq_of_forall_le_iff fun l ↦ ?_ - simp_rw [← tendsto_id', tendsto_iff_forall_compact_tendstoUniformlyOn, + simp_rw [← tendsto_id', tendsto_iff_forall_isCompact_tendstoUniformlyOn, nhds_induced, tendsto_comap_iff, UniformOnFun.tendsto_iff_tendstoUniformlyOn] rfl @@ -232,7 +235,7 @@ variable {ι : Type u₃} {p : Filter ι} {F : ι → C(α, β)} {f} /-- Locally uniform convergence implies convergence in the compact-open topology. -/ theorem tendsto_of_tendstoLocallyUniformly (h : TendstoLocallyUniformly (fun i a => F i a) f p) : Tendsto F p (𝓝 f) := by - rw [tendsto_iff_forall_compact_tendstoUniformlyOn] + rw [tendsto_iff_forall_isCompact_tendstoUniformlyOn] intro K hK rw [← tendstoLocallyUniformlyOn_iff_tendstoUniformlyOn_of_compact hK] exact h.tendstoLocallyUniformlyOn @@ -245,7 +248,7 @@ see `ContinuousMap.tendsto_of_tendstoLocallyUniformly`. -/ theorem tendsto_iff_tendstoLocallyUniformly [WeaklyLocallyCompactSpace α] : Tendsto F p (𝓝 f) ↔ TendstoLocallyUniformly (fun i a => F i a) f p := by refine ⟨fun h V hV x ↦ ?_, tendsto_of_tendstoLocallyUniformly⟩ - rw [tendsto_iff_forall_compact_tendstoUniformlyOn] at h + rw [tendsto_iff_forall_isCompact_tendstoUniformlyOn] at h obtain ⟨n, hn₁, hn₂⟩ := exists_compact_mem_nhds x exact ⟨n, hn₂, h n hn₁ V hV⟩ @@ -313,7 +316,7 @@ theorem hasBasis_compactConvergenceUniformity_of_compact : continuous functions on a compact space. -/ theorem tendsto_iff_tendstoUniformly : Tendsto F p (𝓝 f) ↔ TendstoUniformly (fun i a => F i a) f p := by - rw [tendsto_iff_forall_compact_tendstoUniformlyOn, ← tendstoUniformlyOn_univ] + rw [tendsto_iff_forall_isCompact_tendstoUniformlyOn, ← tendstoUniformlyOn_univ] exact ⟨fun h => h univ isCompact_univ, fun h K _hK => h.mono (subset_univ K)⟩ end CompactDomain From fa5742c52990769eba4704185c51c630902c1fa4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 15:36:28 +0000 Subject: [PATCH 28/90] =?UTF-8?q?feat:=20`a=20^=20m=20=E2=89=A4=20b=20^=20?= =?UTF-8?q?n`=20as=20a=20`gcongr`=20lemma=20(#19220)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups --- Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean | 4 ++++ MathlibTest/GCongr/inequalities.lean | 2 +- 2 files changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index 274bb53619f6f..0a60e09d11d81 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -163,6 +163,10 @@ theorem Monotone.pow_const {f : β → M} (hf : Monotone f) : ∀ n : ℕ, Monot @[to_additive nsmul_right_mono] theorem pow_left_mono (n : ℕ) : Monotone fun a : M => a ^ n := monotone_id.pow_const _ +@[to_additive (attr := gcongr)] +lemma pow_le_pow {a b : M} (hab : a ≤ b) (ht : 1 ≤ b) {m n : ℕ} (hmn : m ≤ n) : a ^ m ≤ b ^ n := + (pow_le_pow_left' hab _).trans (pow_le_pow_right' ht hmn) + end CovariantLESwap end Preorder diff --git a/MathlibTest/GCongr/inequalities.lean b/MathlibTest/GCongr/inequalities.lean index 2bce82a11a237..3056ee3b5e6aa 100644 --- a/MathlibTest/GCongr/inequalities.lean +++ b/MathlibTest/GCongr/inequalities.lean @@ -100,7 +100,7 @@ example {a b : ℚ} (h₁ : a < b) (c : ℝ) : (a + c : ℝ) < b + c := by gcong example {k m n : ℤ} (H : m ^ 2 ≤ n ^ 2) : k + m ^ 2 ≤ k + n ^ 2 := by gcongr -- test of behaviour when no lemmas are applicable -example (n k : ℕ) (H : n ^ k + 1 ≤ k ^ n + 1) : n ^ k ≤ k ^ n := by +example (n k : ℕ) (H : n % k + 1 ≤ k % n + 1) : n % k ≤ k % n := by success_if_fail_with_msg "gcongr did not make progress" (gcongr) From 265b1fea02687131c71bdb29a01c8c1da4f1f6f1 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Tue, 19 Nov 2024 16:30:32 +0000 Subject: [PATCH 29/90] feat(Topology/Connected/PathConnected): some instances for locally path-connected spaces (#17064) Shows that locally path-connected spaces are locally connected, and that quotients and disjoint unions of locally path-connected spaces are locally path-connected. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- Mathlib/Topology/Connected/PathConnected.lean | 112 +++++++++++++++++- 1 file changed, 107 insertions(+), 5 deletions(-) diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 07f52c0cebe24..82ff89fe8d4c4 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -1160,13 +1160,18 @@ alias locPathConnected_of_bases := LocPathConnectedSpace.of_bases variable [LocPathConnectedSpace X] +protected theorem IsOpen.pathComponentIn (x : X) (hF : IsOpen F) : + IsOpen (pathComponentIn x F) := by + rw [isOpen_iff_mem_nhds] + intro y hy + let ⟨s, hs⟩ := (path_connected_basis y).mem_iff.mp (hF.mem_nhds (pathComponentIn_subset hy)) + exact mem_of_superset hs.1.1 <| pathComponentIn_congr hy ▸ + hs.1.2.subset_pathComponentIn (mem_of_mem_nhds hs.1.1) hs.2 + /-- In a locally path connected space, each path component is an open set. -/ protected theorem IsOpen.pathComponent (x : X) : IsOpen (pathComponent x) := by - rw [isOpen_iff_mem_nhds] - intro y hxy - rcases (path_connected_basis y).ex_mem with ⟨V, hVy, hVc⟩ - filter_upwards [hVy] with z hz - exact hxy.out.trans (hVc.joinedIn _ (mem_of_mem_nhds hVy) _ hz).joined + rw [← pathComponentIn_univ] + exact isOpen_univ.pathComponentIn _ /-- In a locally path connected space, each path component is a closed set. -/ protected theorem IsClosed.pathComponent (x : X) : IsClosed (pathComponent x) := by @@ -1192,6 +1197,13 @@ theorem pathConnected_subset_basis {U : Set X} (h : IsOpen U) (hx : x ∈ U) : (𝓝 x).HasBasis (fun s : Set X => s ∈ 𝓝 x ∧ IsPathConnected s ∧ s ⊆ U) id := (path_connected_basis x).hasBasis_self_subset (IsOpen.mem_nhds h hx) +theorem isOpen_isPathConnected_basis (x : X) : + (𝓝 x).HasBasis (fun s : Set X ↦ IsOpen s ∧ x ∈ s ∧ IsPathConnected s) id := by + refine ⟨fun s ↦ ⟨fun hs ↦ ?_, fun ⟨u, hu⟩ ↦ mem_nhds_iff.mpr ⟨u, hu.2, hu.1.1, hu.1.2.1⟩⟩⟩ + have ⟨u, hus, hu, hxu⟩ := mem_nhds_iff.mp hs + exact ⟨pathComponentIn x u, ⟨hu.pathComponentIn _, ⟨mem_pathComponentIn_self hxu, + isPathConnected_pathComponentIn hxu⟩⟩, pathComponentIn_subset.trans hus⟩ + theorem Topology.IsOpenEmbedding.locPathConnectedSpace {e : Y → X} (he : IsOpenEmbedding e) : LocPathConnectedSpace Y := have (y : Y) : @@ -1215,4 +1227,94 @@ theorem IsOpen.isConnected_iff_isPathConnected {U : Set X} (U_op : IsOpen U) : haveI := U_op.locPathConnectedSpace exact pathConnectedSpace_iff_connectedSpace.symm +/-- Locally path-connected spaces are locally connected. -/ +instance : LocallyConnectedSpace X := by + refine ⟨forall_imp (fun x h ↦ ⟨fun s ↦ ?_⟩) isOpen_isPathConnected_basis⟩ + refine ⟨fun hs ↦ ?_, fun ⟨u, ⟨hu, hxu, _⟩, hus⟩ ↦ mem_nhds_iff.mpr ⟨u, hus, hu, hxu⟩⟩ + let ⟨u, ⟨hu, hxu, hu'⟩, hus⟩ := (h.mem_iff' s).mp hs + exact ⟨u, ⟨hu, hxu, hu'.isConnected⟩, hus⟩ + +/-- A space is locally path-connected iff all path components of open subsets are open. -/ +lemma locPathConnectedSpace_iff_isOpen_pathComponentIn {X : Type*} [TopologicalSpace X] : + LocPathConnectedSpace X ↔ ∀ (x : X) (u : Set X), IsOpen u → IsOpen (pathComponentIn x u) := + ⟨fun _ _ _ hu ↦ hu.pathComponentIn _, fun h ↦ ⟨fun x ↦ ⟨fun s ↦ by + refine ⟨fun hs ↦ ?_, fun ⟨_, ht⟩ ↦ Filter.mem_of_superset ht.1.1 ht.2⟩ + let ⟨u, hu⟩ := mem_nhds_iff.mp hs + exact ⟨pathComponentIn x u, ⟨(h x u hu.2.1).mem_nhds (mem_pathComponentIn_self hu.2.2), + isPathConnected_pathComponentIn hu.2.2⟩, pathComponentIn_subset.trans hu.1⟩⟩⟩⟩ + +/-- A space is locally path-connected iff all path components of open subsets are neighbourhoods. -/ +lemma locPathConnectedSpace_iff_pathComponentIn_mem_nhds {X : Type*} [TopologicalSpace X] : + LocPathConnectedSpace X ↔ + ∀ x : X, ∀ u : Set X, IsOpen u → x ∈ u → pathComponentIn x u ∈ nhds x := by + rw [locPathConnectedSpace_iff_isOpen_pathComponentIn] + simp_rw [forall_comm (β := Set X), ← imp_forall_iff] + refine forall_congr' fun u ↦ imp_congr_right fun _ ↦ ?_ + exact ⟨fun h x hxu ↦ (h x).mem_nhds (mem_pathComponentIn_self hxu), + fun h x ↦ isOpen_iff_mem_nhds.mpr fun y hy ↦ + pathComponentIn_congr hy ▸ h y <| pathComponentIn_subset hy⟩ + +/-- Any topology coinduced by a locally path-connected topology is locally path-connected. -/ +lemma LocPathConnectedSpace.coinduced {Y : Type*} (f : X → Y) : + @LocPathConnectedSpace Y (.coinduced f ‹_›) := by + let _ := TopologicalSpace.coinduced f ‹_›; have hf : Continuous f := continuous_coinduced_rng + refine locPathConnectedSpace_iff_isOpen_pathComponentIn.mpr fun y u hu ↦ + isOpen_coinduced.mpr <| isOpen_iff_mem_nhds.mpr fun x hx ↦ ?_ + have hx' := preimage_mono pathComponentIn_subset hx + refine mem_nhds_iff.mpr ⟨pathComponentIn x (f ⁻¹' u), ?_, + (hu.preimage hf).pathComponentIn _, mem_pathComponentIn_self hx'⟩ + rw [← image_subset_iff, ← pathComponentIn_congr hx] + exact ((isPathConnected_pathComponentIn hx').image hf).subset_pathComponentIn + ⟨x, mem_pathComponentIn_self hx', rfl⟩ <| + (image_mono pathComponentIn_subset).trans <| u.image_preimage_subset f + +/-- Quotients of locally path-connected spaces are locally path-connected. -/ +lemma Topology.IsQuotientMap.locPathConnectedSpace {f : X → Y} (h : IsQuotientMap f) : + LocPathConnectedSpace Y := + h.2 ▸ LocPathConnectedSpace.coinduced f + +/-- Quotients of locally path-connected spaces are locally path-connected. -/ +instance Quot.locPathConnectedSpace {r : X → X → Prop} : LocPathConnectedSpace (Quot r) := + isQuotientMap_quot_mk.locPathConnectedSpace + +/-- Quotients of locally path-connected spaces are locally path-connected. -/ +instance Quotient.locPathConnectedSpace {s : Setoid X} : LocPathConnectedSpace (Quotient s) := + isQuotientMap_quotient_mk'.locPathConnectedSpace + + +/-- Disjoint unions of locally path-connected spaces are locally path-connected. -/ +instance Sum.locPathConnectedSpace.{u} {X Y : Type u} [TopologicalSpace X] [TopologicalSpace Y] + [LocPathConnectedSpace X] [LocPathConnectedSpace Y] : + LocPathConnectedSpace (X ⊕ Y) := by + rw [locPathConnectedSpace_iff_pathComponentIn_mem_nhds]; intro x u hu hxu; rw [mem_nhds_iff] + obtain x | y := x + · refine ⟨Sum.inl '' (pathComponentIn x (Sum.inl ⁻¹' u)), ?_, ?_, ?_⟩ + · apply IsPathConnected.subset_pathComponentIn + · exact (isPathConnected_pathComponentIn (by exact hxu)).image continuous_inl + · exact ⟨x, mem_pathComponentIn_self hxu, rfl⟩ + · exact (image_mono pathComponentIn_subset).trans (u.image_preimage_subset _) + · exact isOpenMap_inl _ <| (hu.preimage continuous_inl).pathComponentIn _ + · exact ⟨x, mem_pathComponentIn_self hxu, rfl⟩ + · refine ⟨Sum.inr '' (pathComponentIn y (Sum.inr ⁻¹' u)), ?_, ?_, ?_⟩ + · apply IsPathConnected.subset_pathComponentIn + · exact (isPathConnected_pathComponentIn (by exact hxu)).image continuous_inr + · exact ⟨y, mem_pathComponentIn_self hxu, rfl⟩ + · exact (image_mono pathComponentIn_subset).trans (u.image_preimage_subset _) + · exact isOpenMap_inr _ <| (hu.preimage continuous_inr).pathComponentIn _ + · exact ⟨y, mem_pathComponentIn_self hxu, rfl⟩ + + +/-- Disjoint unions of locally path-connected spaces are locally path-connected. -/ +instance Sigma.locPathConnectedSpace {X : ι → Type*} + [(i : ι) → TopologicalSpace (X i)] [(i : ι) → LocPathConnectedSpace (X i)] : + LocPathConnectedSpace ((i : ι) × X i) := by + rw [locPathConnectedSpace_iff_pathComponentIn_mem_nhds]; intro x u hu hxu; rw [mem_nhds_iff] + refine ⟨(Sigma.mk x.1) '' (pathComponentIn x.2 ((Sigma.mk x.1) ⁻¹' u)), ?_, ?_, ?_⟩ + · apply IsPathConnected.subset_pathComponentIn + · exact (isPathConnected_pathComponentIn (by exact hxu)).image continuous_sigmaMk + · exact ⟨x.2, mem_pathComponentIn_self hxu, rfl⟩ + · exact (image_mono pathComponentIn_subset).trans (u.image_preimage_subset _) + · exact isOpenMap_sigmaMk _ <| (hu.preimage continuous_sigmaMk).pathComponentIn _ + · exact ⟨x.2, mem_pathComponentIn_self hxu, rfl⟩ + end LocPathConnectedSpace From d6c2c9157d71b59d98033b31423a0db08f11c4b4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 16:30:33 +0000 Subject: [PATCH 30/90] feat: inclusion-exclusion principle (#17957) It was a long-standing TODO to have this in mathlib. --- Mathlib.lean | 1 + .../Algebra/BigOperators/Group/Finset.lean | 6 + .../Enumerative/InclusionExclusion.lean | 129 ++++++++++++++++++ docs/100.yaml | 7 +- 4 files changed, 142 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean diff --git a/Mathlib.lean b/Mathlib.lean index 49dd52b773d65..3bde3993e0818 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2138,6 +2138,7 @@ import Mathlib.Combinatorics.Enumerative.Composition import Mathlib.Combinatorics.Enumerative.DoubleCounting import Mathlib.Combinatorics.Enumerative.DyckWord import Mathlib.Combinatorics.Enumerative.IncidenceAlgebra +import Mathlib.Combinatorics.Enumerative.InclusionExclusion import Mathlib.Combinatorics.Enumerative.Partition import Mathlib.Combinatorics.HalesJewett import Mathlib.Combinatorics.Hall.Basic diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 29817d8370d97..a8d12379bbcce 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -420,6 +420,12 @@ theorem prod_filter_mul_prod_filter_not have := Classical.decEq α rw [← prod_union (disjoint_filter_filter_neg s s p), filter_union_filter_neg_eq] +@[to_additive] +lemma prod_filter_not_mul_prod_filter (s : Finset α) (p : α → Prop) [DecidablePred p] + [∀ x, Decidable (¬p x)] (f : α → β) : + (∏ x ∈ s.filter fun x ↦ ¬p x, f x) * ∏ x ∈ s.filter p, f x = ∏ x ∈ s, f x := by + rw [mul_comm, prod_filter_mul_prod_filter_not] + section ToList @[to_additive (attr := simp)] diff --git a/Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean b/Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean new file mode 100644 index 0000000000000..314e4eea70693 --- /dev/null +++ b/Mathlib/Combinatorics/Enumerative/InclusionExclusion.lean @@ -0,0 +1,129 @@ +/- +Copyright (c) 2024 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import Mathlib.Algebra.BigOperators.Pi +import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Algebra.Module.BigOperators + +/-! +# Inclusion-exclusion principle + +This file proves several variants of the inclusion-exclusion principle. + +The inclusion-exclusion principle says that the sum/integral of a function over a finite union of +sets can be calculated as the alternating sum over `n > 0` of the sum/integral of the function over +the intersection of `n` of the sets. + +By taking complements, it also says that the sum/integral of a function over a finite intersection +of complements of sets can be calculated as the alternating sum over `n ≥ 0` of the sum/integral of +the function over the intersection of `n` of the sets. + +By taking the function to be constant `1`, we instead get a result about the cardinality/measure of +the sets. + +## Main declarations + +Per the above explanation, this file contains the following variants of inclusion-exclusion: +* `Finset.inclusion_exclusion_sum_biUnion`: Sum of a function over a finite union of sets +* `Finset.inclusion_exclusion_card_biUnion`: Cardinality of a finite union of sets +* `Finset.inclusion_exclusion_sum_inf_compl`: Sum of a function over a finite intersection of + complements of sets +* `Finset.inclusion_exclusion_card_inf_compl`: Cardinality of a finite intersection of + complements of sets + +## TODO + +* Use (a slight variant of) `Finset.prod_indicator_biUnion_sub_indicator` to prove the integral + version of inclusion-exclusion. +* Prove that truncating the series alternatively gives an upper/lower bound to the true value. +-/ + +namespace Finset +variable {ι α G : Type*} [DecidableEq α] [AddCommGroup G] {s : Finset ι} + +lemma prod_indicator_biUnion_sub_indicator (hs : s.Nonempty) (S : ι → Finset α) (a : α) : + ∏ i ∈ s, (Set.indicator (s.biUnion S) 1 a - Set.indicator (S i) 1 a) = (0 : ℤ) := by + by_cases ha : a ∈ s.biUnion S + · obtain ⟨i, hi, ha⟩ := mem_biUnion.1 ha + exact prod_eq_zero hi <| by simp [*, -coe_biUnion] + · obtain ⟨i, hi⟩ := hs + have ha : a ∉ S i := fun h ↦ ha <| subset_biUnion_of_mem _ hi h + exact prod_eq_zero hi <| by simp [*, -coe_biUnion] + +/-- **Inclusion-exclusion principle** for the sum of a function over a union. + +The sum of a function `f` over the union of the `S i` over `i ∈ s` is the alternating sum of the +sums of `f` over the intersections of the `S i`. -/ +theorem inclusion_exclusion_sum_biUnion (s : Finset ι) (S : ι → Finset α) (f : α → G) : + ∑ a ∈ s.biUnion S, f a = ∑ t : s.powerset.filter (·.Nonempty), + (-1) ^ (#t.1 + 1) • ∑ a ∈ t.1.inf' (mem_filter.1 t.2).2 S, f a := by + classical + rw [← sub_eq_zero] + calc + ∑ a ∈ s.biUnion S, f a - ∑ t : s.powerset.filter (·.Nonempty), + (-1) ^ (#t.1 + 1) • ∑ a ∈ t.1.inf' (mem_filter.1 t.2).2 S, f a + = ∑ t : s.powerset.filter (·.Nonempty), + (-1) ^ #t.1 • ∑ a ∈ t.1.inf' (mem_filter.1 t.2).2 S, f a + + ∑ t ∈ s.powerset.filter (¬ ·.Nonempty), (-1) ^ #t • ∑ a ∈ s.biUnion S, f a := by + simp [sub_eq_neg_add, ← sum_neg_distrib, filter_eq', pow_succ] + _ = ∑ t ∈ s.powerset, (-1) ^ #t • + if ht : t.Nonempty then ∑ a ∈ t.inf' ht S, f a else ∑ a ∈ s.biUnion S, f a := by + rw [← sum_attach (filter ..)]; simp [sum_dite, filter_eq', sum_attach] + _ = ∑ a ∈ s.biUnion S, (∏ i ∈ s, (1 - Set.indicator (S i) 1 a : ℤ)) • f a := by + simp only [Int.reduceNeg, mem_coe, prod_sub, sum_comm (s := s.biUnion S), sum_smul, mul_assoc] + congr! with t + split_ifs with ht + · obtain ⟨i, hi⟩ := ht + simp only [prod_const_one, mul_one, prod_indicator_apply] + simp only [smul_sum, Set.indicator, Set.mem_iInter, mem_coe, Pi.one_apply, mul_ite, mul_one, + mul_zero, ite_smul, zero_smul, sum_ite, not_forall, sum_const_zero, add_zero] + congr + aesop + · obtain rfl := not_nonempty_iff_eq_empty.1 ht + simp + _ = ∑ a ∈ s.biUnion S, (∏ i ∈ s, + (Set.indicator (s.biUnion S) 1 a - Set.indicator (S i) 1 a) : ℤ) • f a := by + congr! with t; rw [Set.indicator_of_mem ‹_›, Pi.one_apply] + _ = 0 := by + obtain rfl | hs := s.eq_empty_or_nonempty <;> + simp [-coe_biUnion, prod_indicator_biUnion_sub_indicator, *] + +/-- **Inclusion-exclusion principle** for the cardinality of a union. + +The cardinality of the union of the `S i` over `i ∈ s` is the alternating sum of the cardinalities +of the intersections of the `S i`. -/ +theorem inclusion_exclusion_card_biUnion (s : Finset ι) (S : ι → Finset α) : + #(s.biUnion S) = ∑ t : s.powerset.filter (·.Nonempty), + (-1 : ℤ) ^ (#t.1 + 1) * #(t.1.inf' (mem_filter.1 t.2).2 S) := by + simpa using inclusion_exclusion_sum_biUnion (G := ℤ) s S (f := 1) + +variable [Fintype α] + +/-- **Inclusion-exclusion principle** for the sum of a function over an intersection of complements. + +The sum of a function `f` over the intersection of the complements of the `S i` over `i ∈ s` is the +alternating sum of the sums of `f` over the intersections of the `S i`. -/ +theorem inclusion_exclusion_sum_inf_compl (s : Finset ι) (S : ι → Finset α) (f : α → G) : + ∑ a ∈ s.inf fun i ↦ (S i)ᶜ, f a = ∑ t ∈ s.powerset, (-1) ^ #t • ∑ a ∈ t.inf S, f a := by + classical + calc + ∑ a ∈ s.inf fun i ↦ (S i)ᶜ, f a + = ∑ a, f a - ∑ a ∈ s.biUnion S, f a := by + rw [← Finset.compl_sup, sup_eq_biUnion, eq_sub_iff_add_eq, sum_compl_add_sum] + _ = ∑ t ∈ s.powerset.filter (¬ ·.Nonempty), (-1) ^ #t • ∑ a ∈ t.inf S, f a + + ∑ t ∈ s.powerset.filter (·.Nonempty), (-1) ^ #t • ∑ a ∈ t.inf S, f a := by + simp [← sum_attach (filter ..), inclusion_exclusion_sum_biUnion, inf'_eq_inf, filter_eq', + sub_eq_add_neg, pow_succ] + _ = ∑ t ∈ s.powerset, (-1) ^ #t • ∑ a ∈ t.inf S, f a := sum_filter_not_add_sum_filter .. + +/-- **Inclusion-exclusion principle** for the cardinality of an intersection of complements. + +The cardinality of the intersection of the complements of the `S i` over `i ∈ s` is the +alternating sum of the cardinalities of the intersections of the `S i`. -/ +theorem inclusion_exclusion_card_inf_compl (s : Finset ι) (S : ι → Finset α) : + #(s.inf fun i ↦ (S i)ᶜ) = ∑ t ∈ s.powerset, (-1 : ℤ) ^ #t * #(t.inf S) := by + simpa using inclusion_exclusion_sum_inf_compl (G := ℤ) s S (f := 1) + +end Finset diff --git a/docs/100.yaml b/docs/100.yaml index 670ae3991e330..31f5837e872ca 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -395,9 +395,14 @@ author : Manuel Candales 96: title : Principle of Inclusion/Exclusion + decls : + - Finset.inclusion_exclusion_sum_biUnion + - Finset.inclusion_exclusion_card_biUnion + - Finset.inclusion_exclusion_sum_inf_compl + - Finset.inclusion_exclusion_card_inf_compl links : github : https://github.com/NeilStrickland/lean_lib/blob/f88d162da2f990b87c4d34f5f46bbca2bbc5948e/src/combinatorics/matching.lean#L304 - author : Neil Strickland + author : Neil Strickland (outside mathlib), Yaël Dillies (in mathlib) 97: title : Cramer’s Rule decl : Matrix.mulVec_cramer From 53523844a7ce22746461dc24cb05736dfe43bb91 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 19 Nov 2024 16:30:35 +0000 Subject: [PATCH 31/90] chore: split Topology.UniformSpace.Basic (#19194) and reorganize material in UniformSpace.Compact (moving some material from Basic into it, and splitting parts of the original content out into later files) --- Mathlib.lean | 2 + .../Topology/Algebra/UniformGroup/Basic.lean | 1 + .../Topology/Algebra/UniformGroup/Defs.lean | 2 +- Mathlib/Topology/EMetricSpace/Defs.lean | 1 + Mathlib/Topology/GDelta/UniformSpace.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 1 + Mathlib/Topology/UniformSpace/Basic.lean | 99 +----- Mathlib/Topology/UniformSpace/Cauchy.lean | 2 +- Mathlib/Topology/UniformSpace/Compact.lean | 323 ++++++------------ .../UniformSpace/CompactConvergence.lean | 1 + .../Topology/UniformSpace/HeineCantor.lean | 146 ++++++++ .../Topology/UniformSpace/OfCompactT2.lean | 107 ++++++ Mathlib/Topology/UniformSpace/OfFun.lean | 2 +- Mathlib/Topology/UniformSpace/Separation.lean | 2 +- .../UniformSpace/UniformConvergence.lean | 1 - 15 files changed, 367 insertions(+), 325 deletions(-) create mode 100644 Mathlib/Topology/UniformSpace/HeineCantor.lean create mode 100644 Mathlib/Topology/UniformSpace/OfCompactT2.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3bde3993e0818..866e1168d05e1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5245,7 +5245,9 @@ import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Completion import Mathlib.Topology.UniformSpace.Equicontinuity import Mathlib.Topology.UniformSpace.Equiv +import Mathlib.Topology.UniformSpace.HeineCantor import Mathlib.Topology.UniformSpace.Matrix +import Mathlib.Topology.UniformSpace.OfCompactT2 import Mathlib.Topology.UniformSpace.OfFun import Mathlib.Topology.UniformSpace.Pi import Mathlib.Topology.UniformSpace.Separation diff --git a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean index 8fd7dd61a29d2..4baeea0bda780 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Basic.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.UniformEmbedding import Mathlib.Topology.UniformSpace.CompleteSeparated import Mathlib.Topology.UniformSpace.Compact +import Mathlib.Topology.UniformSpace.HeineCantor import Mathlib.Topology.Algebra.UniformGroup.Defs import Mathlib.Topology.Algebra.Group.Quotient import Mathlib.Topology.DiscreteSubset diff --git a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean index b593b9d1eb3ca..3455d9f14d056 100644 --- a/Mathlib/Topology/Algebra/UniformGroup/Defs.lean +++ b/Mathlib/Topology/Algebra/UniformGroup/Defs.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.Algebra.Group.Basic +import Mathlib.Topology.UniformSpace.Basic /-! # Uniform structure on topological groups diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 1c005ea711864..445f6803e42d1 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas -/ import Mathlib.Data.ENNReal.Inv import Mathlib.Topology.UniformSpace.OfFun +import Mathlib.Topology.Bases /-! # Extended metric spaces diff --git a/Mathlib/Topology/GDelta/UniformSpace.lean b/Mathlib/Topology/GDelta/UniformSpace.lean index 46bdc36f2fdab..3aab2f9dd00c0 100644 --- a/Mathlib/Topology/GDelta/UniformSpace.lean +++ b/Mathlib/Topology/GDelta/UniformSpace.lean @@ -4,8 +4,8 @@ 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 +import Mathlib.Topology.UniformSpace.Basic /-! # `Gδ` sets and uniform spaces diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 19689b2f5cdf6..82a80a1ff7f93 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -6,6 +6,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas import Mathlib.Data.ENNReal.Real import Mathlib.Tactic.Bound.Attribute import Mathlib.Topology.EMetricSpace.Defs +import Mathlib.Topology.UniformSpace.Compact /-! ## Pseudo-metric spaces diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 47f236596aa6b..5534b8ebfe86d 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -5,9 +5,10 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Order.Filter.SmallSets import Mathlib.Tactic.Monotonicity.Basic -import Mathlib.Topology.Compactness.Compact +import Mathlib.Order.Filter.Tendsto import Mathlib.Topology.NhdsSet import Mathlib.Algebra.Group.Defs +import Mathlib.Topology.ContinuousOn /-! # Uniform spaces @@ -1608,102 +1609,6 @@ end Sum end Constructions -/-! -### Compact sets in uniform spaces --/ - -section Compact - -open UniformSpace -variable [UniformSpace α] {K : Set α} - -/-- Let `c : ι → Set α` be an open cover of a compact set `s`. Then there exists an entourage -`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `c i`. -/ -theorem lebesgue_number_lemma {ι : Sort*} {U : ι → Set α} (hK : IsCompact K) - (hopen : ∀ i, IsOpen (U i)) (hcover : K ⊆ ⋃ i, U i) : - ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ i, ball x V ⊆ U i := by - have : ∀ x ∈ K, ∃ i, ∃ V ∈ 𝓤 α, ball x (V ○ V) ⊆ U i := fun x hx ↦ by - obtain ⟨i, hi⟩ := mem_iUnion.1 (hcover hx) - rw [← (hopen i).mem_nhds_iff, nhds_eq_comap_uniformity, ← lift'_comp_uniformity] at hi - exact ⟨i, (((basis_sets _).lift' <| monotone_id.compRel monotone_id).comap _).mem_iff.1 hi⟩ - choose ind W hW hWU using this - rcases hK.elim_nhds_subcover' (fun x hx ↦ ball x (W x hx)) (fun x hx ↦ ball_mem_nhds _ (hW x hx)) - with ⟨t, ht⟩ - refine ⟨⋂ x ∈ t, W x x.2, (biInter_finset_mem _).2 fun x _ ↦ hW x x.2, fun x hx ↦ ?_⟩ - rcases mem_iUnion₂.1 (ht hx) with ⟨y, hyt, hxy⟩ - exact ⟨ind y y.2, fun z hz ↦ hWU _ _ ⟨x, hxy, mem_iInter₂.1 hz _ hyt⟩⟩ - -/-- Let `U : ι → Set α` be an open cover of a compact set `K`. -Then there exists an entourage `V` -such that for each `x ∈ K` its `V`-neighborhood is included in some `U i`. - -Moreover, one can choose an entourage from a given basis. -/ -protected theorem Filter.HasBasis.lebesgue_number_lemma {ι' ι : Sort*} {p : ι' → Prop} - {V : ι' → Set (α × α)} {U : ι → Set α} (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) - (hopen : ∀ j, IsOpen (U j)) (hcover : K ⊆ ⋃ j, U j) : - ∃ i, p i ∧ ∀ x ∈ K, ∃ j, ball x (V i) ⊆ U j := by - refine (hbasis.exists_iff ?_).1 (lebesgue_number_lemma hK hopen hcover) - exact fun s t hst ht x hx ↦ (ht x hx).imp fun i hi ↦ Subset.trans (ball_mono hst _) hi - -/-- Let `c : Set (Set α)` be an open cover of a compact set `s`. Then there exists an entourage -`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `t ∈ c`. -/ -theorem lebesgue_number_lemma_sUnion {S : Set (Set α)} - (hK : IsCompact K) (hopen : ∀ s ∈ S, IsOpen s) (hcover : K ⊆ ⋃₀ S) : - ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ s ∈ S, ball x V ⊆ s := by - rw [sUnion_eq_iUnion] at hcover - simpa using lebesgue_number_lemma hK (by simpa) hcover - -/-- If `K` is a compact set in a uniform space and `{V i | p i}` is a basis of entourages, -then `{⋃ x ∈ K, UniformSpace.ball x (V i) | p i}` is a basis of `𝓝ˢ K`. - -Here "`{s i | p i}` is a basis of a filter `l`" means `Filter.HasBasis l p s`. -/ -theorem IsCompact.nhdsSet_basis_uniformity {p : ι → Prop} {V : ι → Set (α × α)} - (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) : - (𝓝ˢ K).HasBasis p fun i => ⋃ x ∈ K, ball x (V i) where - mem_iff' U := by - constructor - · intro H - have HKU : K ⊆ ⋃ _ : Unit, interior U := by - simpa only [iUnion_const, subset_interior_iff_mem_nhdsSet] using H - obtain ⟨i, hpi, hi⟩ : ∃ i, p i ∧ ⋃ x ∈ K, ball x (V i) ⊆ interior U := by - simpa using hbasis.lebesgue_number_lemma hK (fun _ ↦ isOpen_interior) HKU - exact ⟨i, hpi, hi.trans interior_subset⟩ - · rintro ⟨i, hpi, hi⟩ - refine mem_of_superset (bUnion_mem_nhdsSet fun x _ ↦ ?_) hi - exact ball_mem_nhds _ <| hbasis.mem_of_mem hpi - --- TODO: move to a separate file, golf using the regularity of a uniform space. -theorem Disjoint.exists_uniform_thickening {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) - (h : Disjoint A B) : ∃ V ∈ 𝓤 α, Disjoint (⋃ x ∈ A, ball x V) (⋃ x ∈ B, ball x V) := by - have : Bᶜ ∈ 𝓝ˢ A := hB.isOpen_compl.mem_nhdsSet.mpr h.le_compl_right - rw [(hA.nhdsSet_basis_uniformity (Filter.basis_sets _)).mem_iff] at this - rcases this with ⟨U, hU, hUAB⟩ - rcases comp_symm_mem_uniformity_sets hU with ⟨V, hV, hVsymm, hVU⟩ - refine ⟨V, hV, Set.disjoint_left.mpr fun x => ?_⟩ - simp only [mem_iUnion₂] - rintro ⟨a, ha, hxa⟩ ⟨b, hb, hxb⟩ - rw [mem_ball_symmetry hVsymm] at hxa hxb - exact hUAB (mem_iUnion₂_of_mem ha <| hVU <| mem_comp_of_mem_ball hVsymm hxa hxb) hb - -theorem Disjoint.exists_uniform_thickening_of_basis {p : ι → Prop} {s : ι → Set (α × α)} - (hU : (𝓤 α).HasBasis p s) {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) - (h : Disjoint A B) : ∃ i, p i ∧ Disjoint (⋃ x ∈ A, ball x (s i)) (⋃ x ∈ B, ball x (s i)) := by - rcases h.exists_uniform_thickening hA hB with ⟨V, hV, hVAB⟩ - rcases hU.mem_iff.1 hV with ⟨i, hi, hiV⟩ - exact ⟨i, hi, hVAB.mono (iUnion₂_mono fun a _ => ball_mono hiV a) - (iUnion₂_mono fun b _ => ball_mono hiV b)⟩ - -/-- A useful consequence of the Lebesgue number lemma: given any compact set `K` contained in an -open set `U`, we can find an (open) entourage `V` such that the ball of size `V` about any point of -`K` is contained in `U`. -/ -theorem lebesgue_number_of_compact_open {K U : Set α} (hK : IsCompact K) - (hU : IsOpen U) (hKU : K ⊆ U) : ∃ V ∈ 𝓤 α, IsOpen V ∧ ∀ x ∈ K, UniformSpace.ball x V ⊆ U := - let ⟨V, ⟨hV, hVo⟩, hVU⟩ := - (hK.nhdsSet_basis_uniformity uniformity_hasBasis_open).mem_iff.1 (hU.mem_nhdsSet.2 hKU) - ⟨V, hV, hVo, iUnion₂_subset_iff.1 hVU⟩ - -end Compact - /-! ### Expressing continuity properties in uniform spaces diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 9962bbe2a00d6..b8ae4225ddef4 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -5,8 +5,8 @@ Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Topology.Algebra.Constructions import Mathlib.Topology.Bases -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Algebra.Order.Group.Nat +import Mathlib.Topology.UniformSpace.Basic /-! # Theory of Cauchy filters in uniform spaces. Complete uniform spaces. Totally bounded subsets. diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index f8308127388e0..ed4c24be25ac2 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -1,44 +1,115 @@ /- -Copyright (c) 2020 Patrick Massot. All rights reserved. +Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Patrick Massot, Yury Kudryashov +Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ -import Mathlib.Topology.UniformSpace.UniformConvergence -import Mathlib.Topology.UniformSpace.Equicontinuity -import Mathlib.Topology.Support +import Mathlib.Topology.UniformSpace.Basic +import Mathlib.Topology.Compactness.Compact /-! -# Compact separated uniform spaces - -## Main statements +# Compact sets in uniform spaces * `compactSpace_uniformity`: On a compact uniform space, the topology determines the uniform structure, entourages are exactly the neighborhoods of the diagonal. -* `uniformSpace_of_compact_t2`: every compact T2 topological structure is induced by a uniform - structure. This uniform structure is described in the previous item. - -* **Heine-Cantor** theorem: continuous functions on compact uniform spaces with values in uniform - spaces are automatically uniformly continuous. There are several variations, the main one is - `CompactSpace.uniformContinuous_of_continuous`. - -## Implementation notes - -The construction `uniformSpace_of_compact_t2` is not declared as an instance, as it would badly -loop. - -## Tags - -uniform space, uniform continuity, compact space -/ -open Uniformity Topology Filter UniformSpace Set +universe u v ua ub uc ud + +variable {α : Type ua} {β : Type ub} {γ : Type uc} {δ : Type ud} {ι : Sort*} + +section Compact + +open Uniformity Set Filter UniformSpace +open scoped Topology + +variable [UniformSpace α] {K : Set α} + +/-- Let `c : ι → Set α` be an open cover of a compact set `s`. Then there exists an entourage +`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `c i`. -/ +theorem lebesgue_number_lemma {ι : Sort*} {U : ι → Set α} (hK : IsCompact K) + (hopen : ∀ i, IsOpen (U i)) (hcover : K ⊆ ⋃ i, U i) : + ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ i, ball x V ⊆ U i := by + have : ∀ x ∈ K, ∃ i, ∃ V ∈ 𝓤 α, ball x (V ○ V) ⊆ U i := fun x hx ↦ by + obtain ⟨i, hi⟩ := mem_iUnion.1 (hcover hx) + rw [← (hopen i).mem_nhds_iff, nhds_eq_comap_uniformity, ← lift'_comp_uniformity] at hi + exact ⟨i, (((basis_sets _).lift' <| monotone_id.compRel monotone_id).comap _).mem_iff.1 hi⟩ + choose ind W hW hWU using this + rcases hK.elim_nhds_subcover' (fun x hx ↦ ball x (W x hx)) (fun x hx ↦ ball_mem_nhds _ (hW x hx)) + with ⟨t, ht⟩ + refine ⟨⋂ x ∈ t, W x x.2, (biInter_finset_mem _).2 fun x _ ↦ hW x x.2, fun x hx ↦ ?_⟩ + rcases mem_iUnion₂.1 (ht hx) with ⟨y, hyt, hxy⟩ + exact ⟨ind y y.2, fun z hz ↦ hWU _ _ ⟨x, hxy, mem_iInter₂.1 hz _ hyt⟩⟩ + +/-- Let `U : ι → Set α` be an open cover of a compact set `K`. +Then there exists an entourage `V` +such that for each `x ∈ K` its `V`-neighborhood is included in some `U i`. + +Moreover, one can choose an entourage from a given basis. -/ +protected theorem Filter.HasBasis.lebesgue_number_lemma {ι' ι : Sort*} {p : ι' → Prop} + {V : ι' → Set (α × α)} {U : ι → Set α} (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) + (hopen : ∀ j, IsOpen (U j)) (hcover : K ⊆ ⋃ j, U j) : + ∃ i, p i ∧ ∀ x ∈ K, ∃ j, ball x (V i) ⊆ U j := by + refine (hbasis.exists_iff ?_).1 (lebesgue_number_lemma hK hopen hcover) + exact fun s t hst ht x hx ↦ (ht x hx).imp fun i hi ↦ Subset.trans (ball_mono hst _) hi + +/-- Let `c : Set (Set α)` be an open cover of a compact set `s`. Then there exists an entourage +`n` such that for each `x ∈ s` its `n`-neighborhood is contained in some `t ∈ c`. -/ +theorem lebesgue_number_lemma_sUnion {S : Set (Set α)} + (hK : IsCompact K) (hopen : ∀ s ∈ S, IsOpen s) (hcover : K ⊆ ⋃₀ S) : + ∃ V ∈ 𝓤 α, ∀ x ∈ K, ∃ s ∈ S, ball x V ⊆ s := by + rw [sUnion_eq_iUnion] at hcover + simpa using lebesgue_number_lemma hK (by simpa) hcover + +/-- If `K` is a compact set in a uniform space and `{V i | p i}` is a basis of entourages, +then `{⋃ x ∈ K, UniformSpace.ball x (V i) | p i}` is a basis of `𝓝ˢ K`. + +Here "`{s i | p i}` is a basis of a filter `l`" means `Filter.HasBasis l p s`. -/ +theorem IsCompact.nhdsSet_basis_uniformity {p : ι → Prop} {V : ι → Set (α × α)} + (hbasis : (𝓤 α).HasBasis p V) (hK : IsCompact K) : + (𝓝ˢ K).HasBasis p fun i => ⋃ x ∈ K, ball x (V i) where + mem_iff' U := by + constructor + · intro H + have HKU : K ⊆ ⋃ _ : Unit, interior U := by + simpa only [iUnion_const, subset_interior_iff_mem_nhdsSet] using H + obtain ⟨i, hpi, hi⟩ : ∃ i, p i ∧ ⋃ x ∈ K, ball x (V i) ⊆ interior U := by + simpa using hbasis.lebesgue_number_lemma hK (fun _ ↦ isOpen_interior) HKU + exact ⟨i, hpi, hi.trans interior_subset⟩ + · rintro ⟨i, hpi, hi⟩ + refine mem_of_superset (bUnion_mem_nhdsSet fun x _ ↦ ?_) hi + exact ball_mem_nhds _ <| hbasis.mem_of_mem hpi + +-- TODO: move to a separate file, golf using the regularity of a uniform space. +theorem Disjoint.exists_uniform_thickening {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) + (h : Disjoint A B) : ∃ V ∈ 𝓤 α, Disjoint (⋃ x ∈ A, ball x V) (⋃ x ∈ B, ball x V) := by + have : Bᶜ ∈ 𝓝ˢ A := hB.isOpen_compl.mem_nhdsSet.mpr h.le_compl_right + rw [(hA.nhdsSet_basis_uniformity (Filter.basis_sets _)).mem_iff] at this + rcases this with ⟨U, hU, hUAB⟩ + rcases comp_symm_mem_uniformity_sets hU with ⟨V, hV, hVsymm, hVU⟩ + refine ⟨V, hV, Set.disjoint_left.mpr fun x => ?_⟩ + simp only [mem_iUnion₂] + rintro ⟨a, ha, hxa⟩ ⟨b, hb, hxb⟩ + rw [mem_ball_symmetry hVsymm] at hxa hxb + exact hUAB (mem_iUnion₂_of_mem ha <| hVU <| mem_comp_of_mem_ball hVsymm hxa hxb) hb + +theorem Disjoint.exists_uniform_thickening_of_basis {p : ι → Prop} {s : ι → Set (α × α)} + (hU : (𝓤 α).HasBasis p s) {A B : Set α} (hA : IsCompact A) (hB : IsClosed B) + (h : Disjoint A B) : ∃ i, p i ∧ Disjoint (⋃ x ∈ A, ball x (s i)) (⋃ x ∈ B, ball x (s i)) := by + rcases h.exists_uniform_thickening hA hB with ⟨V, hV, hVAB⟩ + rcases hU.mem_iff.1 hV with ⟨i, hi, hiV⟩ + exact ⟨i, hi, hVAB.mono (iUnion₂_mono fun a _ => ball_mono hiV a) + (iUnion₂_mono fun b _ => ball_mono hiV b)⟩ + +/-- A useful consequence of the Lebesgue number lemma: given any compact set `K` contained in an +open set `U`, we can find an (open) entourage `V` such that the ball of size `V` about any point of +`K` is contained in `U`. -/ +theorem lebesgue_number_of_compact_open {K U : Set α} (hK : IsCompact K) + (hU : IsOpen U) (hKU : K ⊆ U) : ∃ V ∈ 𝓤 α, IsOpen V ∧ ∀ x ∈ K, UniformSpace.ball x V ⊆ U := + let ⟨V, ⟨hV, hVo⟩, hVU⟩ := + (hK.nhdsSet_basis_uniformity uniformity_hasBasis_open).mem_iff.1 (hU.mem_nhdsSet.2 hKU) + ⟨V, hV, hVo, iUnion₂_subset_iff.1 hVU⟩ -variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] - -/-! -### Uniformity on compact spaces --/ /-- On a compact uniform space, the topology determines the uniform structure, entourages are exactly the neighborhoods of the diagonal. -/ @@ -66,196 +137,4 @@ theorem unique_uniformity_of_compact [t : TopologicalSpace γ] [CompactSpace γ] have : @CompactSpace γ u'.toTopologicalSpace := by rwa [h'] rw [@compactSpace_uniformity _ u, compactSpace_uniformity, h, h'] -/-- The unique uniform structure inducing a given compact topological structure. -/ -def uniformSpaceOfCompactT2 [TopologicalSpace γ] [CompactSpace γ] [T2Space γ] : UniformSpace γ where - uniformity := 𝓝ˢ (diagonal γ) - symm := continuous_swap.tendsto_nhdsSet fun _ => Eq.symm - comp := by - /- This is the difficult part of the proof. We need to prove that, for each neighborhood `W` - of the diagonal `Δ`, there exists a smaller neighborhood `V` such that `V ○ V ⊆ W`. - -/ - set 𝓝Δ := 𝓝ˢ (diagonal γ) - -- The filter of neighborhoods of Δ - set F := 𝓝Δ.lift' fun s : Set (γ × γ) => s ○ s - -- Compositions of neighborhoods of Δ - -- If this weren't true, then there would be V ∈ 𝓝Δ such that F ⊓ 𝓟 Vᶜ ≠ ⊥ - rw [le_iff_forall_inf_principal_compl] - intro V V_in - by_contra H - haveI : NeBot (F ⊓ 𝓟 Vᶜ) := ⟨H⟩ - -- Hence compactness would give us a cluster point (x, y) for F ⊓ 𝓟 Vᶜ - obtain ⟨⟨x, y⟩, hxy⟩ : ∃ p : γ × γ, ClusterPt p (F ⊓ 𝓟 Vᶜ) := exists_clusterPt_of_compactSpace _ - -- In particular (x, y) is a cluster point of 𝓟 Vᶜ, hence is not in the interior of V, - -- and a fortiori not in Δ, so x ≠ y - have clV : ClusterPt (x, y) (𝓟 <| Vᶜ) := hxy.of_inf_right - have : (x, y) ∉ interior V := by - have : (x, y) ∈ closure Vᶜ := by rwa [mem_closure_iff_clusterPt] - rwa [closure_compl] at this - have diag_subset : diagonal γ ⊆ interior V := subset_interior_iff_mem_nhdsSet.2 V_in - have x_ne_y : x ≠ y := mt (@diag_subset (x, y)) this - -- Since γ is compact and Hausdorff, it is T₄, hence T₃. - -- So there are closed neighborhoods V₁ and V₂ of x and y contained in - -- disjoint open neighborhoods U₁ and U₂. - obtain - ⟨U₁, _, V₁, V₁_in, U₂, _, V₂, V₂_in, V₁_cl, V₂_cl, U₁_op, U₂_op, VU₁, VU₂, hU₁₂⟩ := - disjoint_nested_nhds x_ne_y - -- We set U₃ := (V₁ ∪ V₂)ᶜ so that W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ is an open - -- neighborhood of Δ. - let U₃ := (V₁ ∪ V₂)ᶜ - have U₃_op : IsOpen U₃ := (V₁_cl.union V₂_cl).isOpen_compl - let W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ - have W_in : W ∈ 𝓝Δ := by - rw [mem_nhdsSet_iff_forall] - rintro ⟨z, z'⟩ (rfl : z = z') - refine IsOpen.mem_nhds ?_ ?_ - · apply_rules [IsOpen.union, IsOpen.prod] - · simp only [W, mem_union, mem_prod, and_self_iff] - exact (_root_.em _).imp_left fun h => union_subset_union VU₁ VU₂ h - -- So W ○ W ∈ F by definition of F - have : W ○ W ∈ F := @mem_lift' _ _ _ (fun s => s ○ s) _ W_in - -- Porting note: was `by simpa only using mem_lift' W_in` - -- And V₁ ×ˢ V₂ ∈ 𝓝 (x, y) - have hV₁₂ : V₁ ×ˢ V₂ ∈ 𝓝 (x, y) := prod_mem_nhds V₁_in V₂_in - -- But (x, y) is also a cluster point of F so (V₁ ×ˢ V₂) ∩ (W ○ W) ≠ ∅ - -- However the construction of W implies (V₁ ×ˢ V₂) ∩ (W ○ W) = ∅. - -- Indeed assume for contradiction there is some (u, v) in the intersection. - obtain ⟨⟨u, v⟩, ⟨u_in, v_in⟩, w, huw, hwv⟩ := clusterPt_iff.mp hxy.of_inf_left hV₁₂ this - -- So u ∈ V₁, v ∈ V₂, and there exists some w such that (u, w) ∈ W and (w ,v) ∈ W. - -- Because u is in V₁ which is disjoint from U₂ and U₃, (u, w) ∈ W forces (u, w) ∈ U₁ ×ˢ U₁. - have uw_in : (u, w) ∈ U₁ ×ˢ U₁ := - (huw.resolve_right fun h => h.1 <| Or.inl u_in).resolve_right fun h => - hU₁₂.le_bot ⟨VU₁ u_in, h.1⟩ - -- Similarly, because v ∈ V₂, (w ,v) ∈ W forces (w, v) ∈ U₂ ×ˢ U₂. - have wv_in : (w, v) ∈ U₂ ×ˢ U₂ := - (hwv.resolve_right fun h => h.2 <| Or.inr v_in).resolve_left fun h => - hU₁₂.le_bot ⟨h.2, VU₂ v_in⟩ - -- Hence w ∈ U₁ ∩ U₂ which is empty. - -- So we have a contradiction - exact hU₁₂.le_bot ⟨uw_in.2, wv_in.1⟩ - nhds_eq_comap_uniformity x := by - simp_rw [nhdsSet_diagonal, comap_iSup, nhds_prod_eq, comap_prod, Function.comp_def, comap_id'] - rw [iSup_split_single _ x, comap_const_of_mem fun V => mem_of_mem_nhds] - suffices ∀ y ≠ x, comap (fun _ : γ ↦ x) (𝓝 y) ⊓ 𝓝 y ≤ 𝓝 x by simpa - intro y hxy - simp [comap_const_of_not_mem (compl_singleton_mem_nhds hxy) (not_not_intro rfl)] - -/-! -### Heine-Cantor theorem --/ - - -/-- Heine-Cantor: a continuous function on a compact uniform space is uniformly -continuous. -/ -theorem CompactSpace.uniformContinuous_of_continuous [CompactSpace α] {f : α → β} - (h : Continuous f) : UniformContinuous f := -calc map (Prod.map f f) (𝓤 α) - = map (Prod.map f f) (𝓝ˢ (diagonal α)) := by rw [nhdsSet_diagonal_eq_uniformity] - _ ≤ 𝓝ˢ (diagonal β) := (h.prodMap h).tendsto_nhdsSet mapsTo_prod_map_diagonal - _ ≤ 𝓤 β := nhdsSet_diagonal_le_uniformity - -/-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly -continuous. -/ -theorem IsCompact.uniformContinuousOn_of_continuous {s : Set α} {f : α → β} (hs : IsCompact s) - (hf : ContinuousOn f s) : UniformContinuousOn f s := by - rw [uniformContinuousOn_iff_restrict] - rw [isCompact_iff_compactSpace] at hs - rw [continuousOn_iff_continuous_restrict] at hf - exact CompactSpace.uniformContinuous_of_continuous hf - -/-- If `s` is compact and `f` is continuous at all points of `s`, then `f` is -"uniformly continuous at the set `s`", i.e. `f x` is close to `f y` whenever `x ∈ s` and `y` is -close to `x` (even if `y` is not itself in `s`, so this is a stronger assertion than -`UniformContinuousOn s`). -/ -theorem IsCompact.uniformContinuousAt_of_continuousAt {r : Set (β × β)} {s : Set α} - (hs : IsCompact s) (f : α → β) (hf : ∀ a ∈ s, ContinuousAt f a) (hr : r ∈ 𝓤 β) : - { x : α × α | x.1 ∈ s → (f x.1, f x.2) ∈ r } ∈ 𝓤 α := by - obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr - choose U hU T hT hb using fun a ha => - exists_mem_nhds_ball_subset_of_mem_nhds ((hf a ha).preimage_mem_nhds <| mem_nhds_left _ ht) - obtain ⟨fs, hsU⟩ := hs.elim_nhds_subcover' U hU - apply mem_of_superset ((biInter_finset_mem fs).2 fun a _ => hT a a.2) - rintro ⟨a₁, a₂⟩ h h₁ - obtain ⟨a, ha, haU⟩ := Set.mem_iUnion₂.1 (hsU h₁) - apply htr - refine ⟨f a, htsymm.mk_mem_comm.1 (hb _ _ _ haU ?_), hb _ _ _ haU ?_⟩ - exacts [mem_ball_self _ (hT a a.2), mem_iInter₂.1 h a ha] - -theorem Continuous.uniformContinuous_of_tendsto_cocompact {f : α → β} {x : β} - (h_cont : Continuous f) (hx : Tendsto f (cocompact α) (𝓝 x)) : UniformContinuous f := - uniformContinuous_def.2 fun r hr => by - obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr - obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (hx <| mem_nhds_left _ ht) - apply - mem_of_superset - (symmetrize_mem_uniformity <| - (hs.uniformContinuousAt_of_continuousAt f fun _ _ => h_cont.continuousAt) <| - symmetrize_mem_uniformity hr) - rintro ⟨b₁, b₂⟩ h - by_cases h₁ : b₁ ∈ s; · exact (h.1 h₁).1 - by_cases h₂ : b₂ ∈ s; · exact (h.2 h₂).2 - apply htr - exact ⟨x, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩ - -@[to_additive] -theorem HasCompactMulSupport.uniformContinuous_of_continuous {f : α → β} [One β] - (h1 : HasCompactMulSupport f) (h2 : Continuous f) : UniformContinuous f := - h2.uniformContinuous_of_tendsto_cocompact h1.is_one_at_infty - -/-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, -`β` is compact and `f` is continuous on `U × (univ : Set β)` for some neighborhood `U` of `x`. -/ -theorem ContinuousOn.tendstoUniformly [LocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] - {f : α → β → γ} {x : α} {U : Set α} (hxU : U ∈ 𝓝 x) (h : ContinuousOn (↿f) (U ×ˢ univ)) : - TendstoUniformly f (f x) (𝓝 x) := by - rcases LocallyCompactSpace.local_compact_nhds _ _ hxU with ⟨K, hxK, hKU, hK⟩ - have : UniformContinuousOn (↿f) (K ×ˢ univ) := - IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) - (h.mono <| prod_mono hKU Subset.rfl) - exact this.tendstoUniformly hxK - -/-- A continuous family of functions `α → β → γ` tends uniformly to its value at `x` -if `α` is weakly locally compact and `β` is compact. -/ -theorem Continuous.tendstoUniformly [WeaklyLocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] - (f : α → β → γ) (h : Continuous ↿f) (x : α) : TendstoUniformly f (f x) (𝓝 x) := - let ⟨K, hK, hxK⟩ := exists_compact_mem_nhds x - have : UniformContinuousOn (↿f) (K ×ˢ univ) := - IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) h.continuousOn - this.tendstoUniformly hxK - -/-- In a product space `α × β`, assume that a function `f` is continuous on `s × k` where `k` is -compact. Then, along the fiber above any `q ∈ s`, `f` is transversely uniformly continuous, i.e., -if `p ∈ s` is close enough to `q`, then `f p x` is uniformly close to `f q x` for all `x ∈ k`. -/ -lemma IsCompact.mem_uniformity_of_prod - {α β E : Type*} [TopologicalSpace α] [TopologicalSpace β] [UniformSpace E] - {f : α → β → E} {s : Set α} {k : Set β} {q : α} {u : Set (E × E)} - (hk : IsCompact k) (hf : ContinuousOn f.uncurry (s ×ˢ k)) (hq : q ∈ s) (hu : u ∈ 𝓤 E) : - ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ k, (f p x, f q x) ∈ u := by - apply hk.induction_on (p := fun t ↦ ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ t, (f p x, f q x) ∈ u) - · exact ⟨univ, univ_mem, by simp⟩ - · intro t' t ht't ⟨v, v_mem, hv⟩ - exact ⟨v, v_mem, fun p hp x hx ↦ hv p hp x (ht't hx)⟩ - · intro t t' ⟨v, v_mem, hv⟩ ⟨v', v'_mem, hv'⟩ - refine ⟨v ∩ v', inter_mem v_mem v'_mem, fun p hp x hx ↦ ?_⟩ - rcases hx with h'x|h'x - · exact hv p hp.1 x h'x - · exact hv' p hp.2 x h'x - · rcases comp_symm_of_uniformity hu with ⟨u', u'_mem, u'_symm, hu'⟩ - intro x hx - obtain ⟨v, hv, w, hw, hvw⟩ : - ∃ v ∈ 𝓝[s] q, ∃ w ∈ 𝓝[k] x, v ×ˢ w ⊆ f.uncurry ⁻¹' {z | (f q x, z) ∈ u'} := - mem_nhdsWithin_prod_iff.1 (hf (q, x) ⟨hq, hx⟩ (mem_nhds_left (f q x) u'_mem)) - refine ⟨w, hw, v, hv, fun p hp y hy ↦ ?_⟩ - have A : (f q x, f p y) ∈ u' := hvw (⟨hp, hy⟩ : (p, y) ∈ v ×ˢ w) - have B : (f q x, f q y) ∈ u' := hvw (⟨mem_of_mem_nhdsWithin hq hv, hy⟩ : (q, y) ∈ v ×ˢ w) - exact hu' (prod_mk_mem_compRel (u'_symm A) B) - -section UniformConvergence - -/-- An equicontinuous family of functions defined on a compact uniform space is automatically -uniformly equicontinuous. -/ -theorem CompactSpace.uniformEquicontinuous_of_equicontinuous {ι : Type*} {F : ι → β → α} - [CompactSpace β] (h : Equicontinuous F) : UniformEquicontinuous F := by - rw [equicontinuous_iff_continuous] at h - rw [uniformEquicontinuous_iff_uniformContinuous] - exact CompactSpace.uniformContinuous_of_continuous h - -end UniformConvergence +end Compact diff --git a/Mathlib/Topology/UniformSpace/CompactConvergence.lean b/Mathlib/Topology/UniformSpace/CompactConvergence.lean index e39e65d2fbda4..d5e4df759ebef 100644 --- a/Mathlib/Topology/UniformSpace/CompactConvergence.lean +++ b/Mathlib/Topology/UniformSpace/CompactConvergence.lean @@ -7,6 +7,7 @@ import Mathlib.Topology.CompactOpen import Mathlib.Topology.LocallyFinite import Mathlib.Topology.Maps.Proper.Basic import Mathlib.Topology.UniformSpace.UniformConvergenceTopology +import Mathlib.Topology.UniformSpace.Compact /-! # Compact convergence (uniform convergence on compact sets) diff --git a/Mathlib/Topology/UniformSpace/HeineCantor.lean b/Mathlib/Topology/UniformSpace/HeineCantor.lean new file mode 100644 index 0000000000000..d97593eba0566 --- /dev/null +++ b/Mathlib/Topology/UniformSpace/HeineCantor.lean @@ -0,0 +1,146 @@ +/- +Copyright (c) 2020 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Yury Kudryashov +-/ +import Mathlib.Topology.UniformSpace.Equicontinuity +import Mathlib.Topology.UniformSpace.Compact +import Mathlib.Topology.Support + +/-! +# Compact separated uniform spaces + +## Main statement + +* **Heine-Cantor** theorem: continuous functions on compact uniform spaces with values in uniform + spaces are automatically uniformly continuous. There are several variations, the main one is + `CompactSpace.uniformContinuous_of_continuous`. + +## Tags + +uniform space, uniform continuity, compact space +-/ + +open Uniformity Topology Filter UniformSpace Set + +variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] + +/-! +### Heine-Cantor theorem +-/ + +/-- Heine-Cantor: a continuous function on a compact uniform space is uniformly +continuous. -/ +theorem CompactSpace.uniformContinuous_of_continuous [CompactSpace α] {f : α → β} + (h : Continuous f) : UniformContinuous f := +calc map (Prod.map f f) (𝓤 α) + = map (Prod.map f f) (𝓝ˢ (diagonal α)) := by rw [nhdsSet_diagonal_eq_uniformity] + _ ≤ 𝓝ˢ (diagonal β) := (h.prodMap h).tendsto_nhdsSet mapsTo_prod_map_diagonal + _ ≤ 𝓤 β := nhdsSet_diagonal_le_uniformity + +/-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly +continuous. -/ +theorem IsCompact.uniformContinuousOn_of_continuous {s : Set α} {f : α → β} (hs : IsCompact s) + (hf : ContinuousOn f s) : UniformContinuousOn f s := by + rw [uniformContinuousOn_iff_restrict] + rw [isCompact_iff_compactSpace] at hs + rw [continuousOn_iff_continuous_restrict] at hf + exact CompactSpace.uniformContinuous_of_continuous hf + +/-- If `s` is compact and `f` is continuous at all points of `s`, then `f` is +"uniformly continuous at the set `s`", i.e. `f x` is close to `f y` whenever `x ∈ s` and `y` is +close to `x` (even if `y` is not itself in `s`, so this is a stronger assertion than +`UniformContinuousOn s`). -/ +theorem IsCompact.uniformContinuousAt_of_continuousAt {r : Set (β × β)} {s : Set α} + (hs : IsCompact s) (f : α → β) (hf : ∀ a ∈ s, ContinuousAt f a) (hr : r ∈ 𝓤 β) : + { x : α × α | x.1 ∈ s → (f x.1, f x.2) ∈ r } ∈ 𝓤 α := by + obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr + choose U hU T hT hb using fun a ha => + exists_mem_nhds_ball_subset_of_mem_nhds ((hf a ha).preimage_mem_nhds <| mem_nhds_left _ ht) + obtain ⟨fs, hsU⟩ := hs.elim_nhds_subcover' U hU + apply mem_of_superset ((biInter_finset_mem fs).2 fun a _ => hT a a.2) + rintro ⟨a₁, a₂⟩ h h₁ + obtain ⟨a, ha, haU⟩ := Set.mem_iUnion₂.1 (hsU h₁) + apply htr + refine ⟨f a, htsymm.mk_mem_comm.1 (hb _ _ _ haU ?_), hb _ _ _ haU ?_⟩ + exacts [mem_ball_self _ (hT a a.2), mem_iInter₂.1 h a ha] + +theorem Continuous.uniformContinuous_of_tendsto_cocompact {f : α → β} {x : β} + (h_cont : Continuous f) (hx : Tendsto f (cocompact α) (𝓝 x)) : UniformContinuous f := + uniformContinuous_def.2 fun r hr => by + obtain ⟨t, ht, htsymm, htr⟩ := comp_symm_mem_uniformity_sets hr + obtain ⟨s, hs, hst⟩ := mem_cocompact.1 (hx <| mem_nhds_left _ ht) + apply + mem_of_superset + (symmetrize_mem_uniformity <| + (hs.uniformContinuousAt_of_continuousAt f fun _ _ => h_cont.continuousAt) <| + symmetrize_mem_uniformity hr) + rintro ⟨b₁, b₂⟩ h + by_cases h₁ : b₁ ∈ s; · exact (h.1 h₁).1 + by_cases h₂ : b₂ ∈ s; · exact (h.2 h₂).2 + apply htr + exact ⟨x, htsymm.mk_mem_comm.1 (hst h₁), hst h₂⟩ + +@[to_additive] +theorem HasCompactMulSupport.uniformContinuous_of_continuous {f : α → β} [One β] + (h1 : HasCompactMulSupport f) (h2 : Continuous f) : UniformContinuous f := + h2.uniformContinuous_of_tendsto_cocompact h1.is_one_at_infty + +/-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, +`β` is compact and `f` is continuous on `U × (univ : Set β)` for some neighborhood `U` of `x`. -/ +theorem ContinuousOn.tendstoUniformly [LocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] + {f : α → β → γ} {x : α} {U : Set α} (hxU : U ∈ 𝓝 x) (h : ContinuousOn (↿f) (U ×ˢ univ)) : + TendstoUniformly f (f x) (𝓝 x) := by + rcases LocallyCompactSpace.local_compact_nhds _ _ hxU with ⟨K, hxK, hKU, hK⟩ + have : UniformContinuousOn (↿f) (K ×ˢ univ) := + IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) + (h.mono <| prod_mono hKU Subset.rfl) + exact this.tendstoUniformly hxK + +/-- A continuous family of functions `α → β → γ` tends uniformly to its value at `x` +if `α` is weakly locally compact and `β` is compact. -/ +theorem Continuous.tendstoUniformly [WeaklyLocallyCompactSpace α] [CompactSpace β] [UniformSpace γ] + (f : α → β → γ) (h : Continuous ↿f) (x : α) : TendstoUniformly f (f x) (𝓝 x) := + let ⟨K, hK, hxK⟩ := exists_compact_mem_nhds x + have : UniformContinuousOn (↿f) (K ×ˢ univ) := + IsCompact.uniformContinuousOn_of_continuous (hK.prod isCompact_univ) h.continuousOn + this.tendstoUniformly hxK + +/-- In a product space `α × β`, assume that a function `f` is continuous on `s × k` where `k` is +compact. Then, along the fiber above any `q ∈ s`, `f` is transversely uniformly continuous, i.e., +if `p ∈ s` is close enough to `q`, then `f p x` is uniformly close to `f q x` for all `x ∈ k`. -/ +lemma IsCompact.mem_uniformity_of_prod + {α β E : Type*} [TopologicalSpace α] [TopologicalSpace β] [UniformSpace E] + {f : α → β → E} {s : Set α} {k : Set β} {q : α} {u : Set (E × E)} + (hk : IsCompact k) (hf : ContinuousOn f.uncurry (s ×ˢ k)) (hq : q ∈ s) (hu : u ∈ 𝓤 E) : + ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ k, (f p x, f q x) ∈ u := by + apply hk.induction_on (p := fun t ↦ ∃ v ∈ 𝓝[s] q, ∀ p ∈ v, ∀ x ∈ t, (f p x, f q x) ∈ u) + · exact ⟨univ, univ_mem, by simp⟩ + · intro t' t ht't ⟨v, v_mem, hv⟩ + exact ⟨v, v_mem, fun p hp x hx ↦ hv p hp x (ht't hx)⟩ + · intro t t' ⟨v, v_mem, hv⟩ ⟨v', v'_mem, hv'⟩ + refine ⟨v ∩ v', inter_mem v_mem v'_mem, fun p hp x hx ↦ ?_⟩ + rcases hx with h'x|h'x + · exact hv p hp.1 x h'x + · exact hv' p hp.2 x h'x + · rcases comp_symm_of_uniformity hu with ⟨u', u'_mem, u'_symm, hu'⟩ + intro x hx + obtain ⟨v, hv, w, hw, hvw⟩ : + ∃ v ∈ 𝓝[s] q, ∃ w ∈ 𝓝[k] x, v ×ˢ w ⊆ f.uncurry ⁻¹' {z | (f q x, z) ∈ u'} := + mem_nhdsWithin_prod_iff.1 (hf (q, x) ⟨hq, hx⟩ (mem_nhds_left (f q x) u'_mem)) + refine ⟨w, hw, v, hv, fun p hp y hy ↦ ?_⟩ + have A : (f q x, f p y) ∈ u' := hvw (⟨hp, hy⟩ : (p, y) ∈ v ×ˢ w) + have B : (f q x, f q y) ∈ u' := hvw (⟨mem_of_mem_nhdsWithin hq hv, hy⟩ : (q, y) ∈ v ×ˢ w) + exact hu' (prod_mk_mem_compRel (u'_symm A) B) + +section UniformConvergence + +/-- An equicontinuous family of functions defined on a compact uniform space is automatically +uniformly equicontinuous. -/ +theorem CompactSpace.uniformEquicontinuous_of_equicontinuous {ι : Type*} {F : ι → β → α} + [CompactSpace β] (h : Equicontinuous F) : UniformEquicontinuous F := by + rw [equicontinuous_iff_continuous] at h + rw [uniformEquicontinuous_iff_uniformContinuous] + exact CompactSpace.uniformContinuous_of_continuous h + +end UniformConvergence diff --git a/Mathlib/Topology/UniformSpace/OfCompactT2.lean b/Mathlib/Topology/UniformSpace/OfCompactT2.lean new file mode 100644 index 0000000000000..5eaf71a852f4f --- /dev/null +++ b/Mathlib/Topology/UniformSpace/OfCompactT2.lean @@ -0,0 +1,107 @@ +/- +Copyright (c) 2020 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Yury Kudryashov +-/ +import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.UniformSpace.Basic + +/-! +# Compact separated uniform spaces + +## Main statement + +* `uniformSpace_of_compact_t2`: every compact T2 topological structure is induced by a uniform + structure. This uniform structure is described by `compactSpace_uniformity`. + +## Implementation notes + +The construction `uniformSpace_of_compact_t2` is not declared as an instance, as it would badly +loop. + +## Tags + +uniform space, uniform continuity, compact space +-/ + +open Uniformity Topology Filter UniformSpace Set + +variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] + +/-! +### Uniformity on compact spaces +-/ + + +/-- The unique uniform structure inducing a given compact topological structure. -/ +def uniformSpaceOfCompactT2 [TopologicalSpace γ] [CompactSpace γ] [T2Space γ] : UniformSpace γ where + uniformity := 𝓝ˢ (diagonal γ) + symm := continuous_swap.tendsto_nhdsSet fun _ => Eq.symm + comp := by + /- This is the difficult part of the proof. We need to prove that, for each neighborhood `W` + of the diagonal `Δ`, there exists a smaller neighborhood `V` such that `V ○ V ⊆ W`. + -/ + set 𝓝Δ := 𝓝ˢ (diagonal γ) + -- The filter of neighborhoods of Δ + set F := 𝓝Δ.lift' fun s : Set (γ × γ) => s ○ s + -- Compositions of neighborhoods of Δ + -- If this weren't true, then there would be V ∈ 𝓝Δ such that F ⊓ 𝓟 Vᶜ ≠ ⊥ + rw [le_iff_forall_inf_principal_compl] + intro V V_in + by_contra H + haveI : NeBot (F ⊓ 𝓟 Vᶜ) := ⟨H⟩ + -- Hence compactness would give us a cluster point (x, y) for F ⊓ 𝓟 Vᶜ + obtain ⟨⟨x, y⟩, hxy⟩ : ∃ p : γ × γ, ClusterPt p (F ⊓ 𝓟 Vᶜ) := exists_clusterPt_of_compactSpace _ + -- In particular (x, y) is a cluster point of 𝓟 Vᶜ, hence is not in the interior of V, + -- and a fortiori not in Δ, so x ≠ y + have clV : ClusterPt (x, y) (𝓟 <| Vᶜ) := hxy.of_inf_right + have : (x, y) ∉ interior V := by + have : (x, y) ∈ closure Vᶜ := by rwa [mem_closure_iff_clusterPt] + rwa [closure_compl] at this + have diag_subset : diagonal γ ⊆ interior V := subset_interior_iff_mem_nhdsSet.2 V_in + have x_ne_y : x ≠ y := mt (@diag_subset (x, y)) this + -- Since γ is compact and Hausdorff, it is T₄, hence T₃. + -- So there are closed neighborhoods V₁ and V₂ of x and y contained in + -- disjoint open neighborhoods U₁ and U₂. + obtain + ⟨U₁, _, V₁, V₁_in, U₂, _, V₂, V₂_in, V₁_cl, V₂_cl, U₁_op, U₂_op, VU₁, VU₂, hU₁₂⟩ := + disjoint_nested_nhds x_ne_y + -- We set U₃ := (V₁ ∪ V₂)ᶜ so that W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ is an open + -- neighborhood of Δ. + let U₃ := (V₁ ∪ V₂)ᶜ + have U₃_op : IsOpen U₃ := (V₁_cl.union V₂_cl).isOpen_compl + let W := U₁ ×ˢ U₁ ∪ U₂ ×ˢ U₂ ∪ U₃ ×ˢ U₃ + have W_in : W ∈ 𝓝Δ := by + rw [mem_nhdsSet_iff_forall] + rintro ⟨z, z'⟩ (rfl : z = z') + refine IsOpen.mem_nhds ?_ ?_ + · apply_rules [IsOpen.union, IsOpen.prod] + · simp only [W, mem_union, mem_prod, and_self_iff] + exact (_root_.em _).imp_left fun h => union_subset_union VU₁ VU₂ h + -- So W ○ W ∈ F by definition of F + have : W ○ W ∈ F := @mem_lift' _ _ _ (fun s => s ○ s) _ W_in + -- Porting note: was `by simpa only using mem_lift' W_in` + -- And V₁ ×ˢ V₂ ∈ 𝓝 (x, y) + have hV₁₂ : V₁ ×ˢ V₂ ∈ 𝓝 (x, y) := prod_mem_nhds V₁_in V₂_in + -- But (x, y) is also a cluster point of F so (V₁ ×ˢ V₂) ∩ (W ○ W) ≠ ∅ + -- However the construction of W implies (V₁ ×ˢ V₂) ∩ (W ○ W) = ∅. + -- Indeed assume for contradiction there is some (u, v) in the intersection. + obtain ⟨⟨u, v⟩, ⟨u_in, v_in⟩, w, huw, hwv⟩ := clusterPt_iff.mp hxy.of_inf_left hV₁₂ this + -- So u ∈ V₁, v ∈ V₂, and there exists some w such that (u, w) ∈ W and (w ,v) ∈ W. + -- Because u is in V₁ which is disjoint from U₂ and U₃, (u, w) ∈ W forces (u, w) ∈ U₁ ×ˢ U₁. + have uw_in : (u, w) ∈ U₁ ×ˢ U₁ := + (huw.resolve_right fun h => h.1 <| Or.inl u_in).resolve_right fun h => + hU₁₂.le_bot ⟨VU₁ u_in, h.1⟩ + -- Similarly, because v ∈ V₂, (w ,v) ∈ W forces (w, v) ∈ U₂ ×ˢ U₂. + have wv_in : (w, v) ∈ U₂ ×ˢ U₂ := + (hwv.resolve_right fun h => h.2 <| Or.inr v_in).resolve_left fun h => + hU₁₂.le_bot ⟨h.2, VU₂ v_in⟩ + -- Hence w ∈ U₁ ∩ U₂ which is empty. + -- So we have a contradiction + exact hU₁₂.le_bot ⟨uw_in.2, wv_in.1⟩ + nhds_eq_comap_uniformity x := by + simp_rw [nhdsSet_diagonal, comap_iSup, nhds_prod_eq, comap_prod, Function.comp_def, comap_id'] + rw [iSup_split_single _ x, comap_const_of_mem fun V => mem_of_mem_nhds] + suffices ∀ y ≠ x, comap (fun _ : γ ↦ x) (𝓝 y) ⊓ 𝓝 y ≤ 𝓝 x by simpa + intro y hxy + simp [comap_const_of_not_mem (compl_singleton_mem_nhds hxy) (not_not_intro rfl)] diff --git a/Mathlib/Topology/UniformSpace/OfFun.lean b/Mathlib/Topology/UniformSpace/OfFun.lean index 6a1ff605f190e..1fa610f447327 100644 --- a/Mathlib/Topology/UniformSpace/OfFun.lean +++ b/Mathlib/Topology/UniformSpace/OfFun.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.UniformSpace.Basic import Mathlib.Algebra.Order.Monoid.Defs +import Mathlib.Topology.UniformSpace.Basic /-! # Construct a `UniformSpace` from a `dist`-like function diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index 672378e6ebef9..331e192d249dc 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov -/ import Mathlib.Tactic.ApplyFun -import Mathlib.Topology.UniformSpace.Basic import Mathlib.Topology.Separation.Basic +import Mathlib.Topology.UniformSpace.Basic /-! # Hausdorff properties of uniform spaces. Separation quotient. diff --git a/Mathlib/Topology/UniformSpace/UniformConvergence.lean b/Mathlib/Topology/UniformSpace/UniformConvergence.lean index aca68d49538ba..48ea1c9d19316 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.UniformSpace.Basic import Mathlib.Topology.UniformSpace.Cauchy /-! From ee430423941cdcdf55a1b5e84f032525cfcef151 Mon Sep 17 00:00:00 2001 From: Peiran Wu <15968905+wupr@users.noreply.github.com> Date: Tue, 19 Nov 2024 17:06:38 +0000 Subject: [PATCH 32/90] feat: define instances of `Finite`, `Fintype`, `DecidableEq` for `MulEquiv` (#17057) --- Mathlib/Algebra/Group/Equiv/Basic.lean | 5 +++++ Mathlib/Data/Finite/Prod.lean | 8 ++++++++ Mathlib/Data/Fintype/Basic.lean | 5 +++++ Mathlib/Data/Fintype/Perm.lean | 14 ++++++++++++-- 4 files changed, 30 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index fc891d4f3cf98..81aff3106d93e 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -111,6 +111,11 @@ infixl:25 " ≃* " => MulEquiv /-- Notation for an `AddEquiv`. -/ infixl:25 " ≃+ " => AddEquiv +@[to_additive] +lemma MulEquiv.toEquiv_injective {α β : Type*} [Mul α] [Mul β] : + Function.Injective (toEquiv : (α ≃* β) → (α ≃ β)) + | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl + /-- `MulEquivClass F A B` states that `F` is a type of multiplication-preserving morphisms. You should extend this class when you extend `MulEquiv`. -/ -- TODO: make this a synonym for MulHomClass? diff --git a/Mathlib/Data/Finite/Prod.lean b/Mathlib/Data/Finite/Prod.lean index 42d9d71fa3edb..c0dbebd034aaa 100644 --- a/Mathlib/Data/Finite/Prod.lean +++ b/Mathlib/Data/Finite/Prod.lean @@ -65,6 +65,14 @@ instance Equiv.finite_right {α β : Sort*} [Finite β] : Finite (α ≃ β) := instance Equiv.finite_left {α β : Sort*} [Finite α] : Finite (α ≃ β) := Finite.of_equiv _ ⟨Equiv.symm, Equiv.symm, Equiv.symm_symm, Equiv.symm_symm⟩ +@[to_additive] +instance MulEquiv.finite_left {α β : Type*} [Mul α] [Mul β] [Finite α] : Finite (α ≃* β) := + Finite.of_injective toEquiv toEquiv_injective + +@[to_additive] +instance MulEquiv.finite_right {α β : Type*} [Mul α] [Mul β] [Finite β] : Finite (α ≃* β) := + Finite.of_injective toEquiv toEquiv_injective + open Set Function variable {γ : Type*} diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 90e35e73ab619..3bda18e430ecb 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -400,6 +400,11 @@ instance decidableEqEquivFintype [DecidableEq β] [Fintype α] : DecidableEq (α instance decidableEqEmbeddingFintype [DecidableEq β] [Fintype α] : DecidableEq (α ↪ β) := fun a b => decidable_of_iff ((a : α → β) = b) Function.Embedding.coe_injective.eq_iff +@[to_additive] +instance decidableEqMulEquivFintype {α β : Type*} [DecidableEq β] [Fintype α] [Mul α] [Mul β] : + DecidableEq (α ≃* β) := + fun a b => decidable_of_iff ((a : α → β) = b) (Injective.eq_iff DFunLike.coe_injective) + end BundledHoms instance decidableInjectiveFintype [DecidableEq α] [DecidableEq β] [Fintype α] : diff --git a/Mathlib/Data/Fintype/Perm.lean b/Mathlib/Data/Fintype/Perm.lean index 4a354e6071e5f..0c2dc9cbbd6df 100644 --- a/Mathlib/Data/Fintype/Perm.lean +++ b/Mathlib/Data/Fintype/Perm.lean @@ -140,7 +140,7 @@ theorem card_perms_of_finset : ∀ s : Finset α, #(permsOfFinset s) = (#s)! := def fintypePerm [Fintype α] : Fintype (Perm α) := ⟨permsOfFinset (@Finset.univ α _), by simp [mem_perms_of_finset_iff]⟩ -instance equivFintype [Fintype α] [Fintype β] : Fintype (α ≃ β) := +instance Equiv.instFintype [Fintype α] [Fintype β] : Fintype (α ≃ β) := if h : Fintype.card β = Fintype.card α then Trunc.recOnSubsingleton (Fintype.truncEquivFin α) fun eα => Trunc.recOnSubsingleton (Fintype.truncEquivFin β) fun eβ => @@ -148,8 +148,18 @@ instance equivFintype [Fintype α] [Fintype β] : Fintype (α ≃ β) := (equivCongr (Equiv.refl α) (eα.trans (Eq.recOn h eβ.symm)) : α ≃ α ≃ (α ≃ β)) else ⟨∅, fun x => False.elim (h (Fintype.card_eq.2 ⟨x.symm⟩))⟩ +@[deprecated (since := "2024-11-19")] alias equivFintype := Equiv.instFintype + +@[to_additive] +instance MulEquiv.instFintype + {α β : Type*} [Mul α] [Mul β] [DecidableEq α] [DecidableEq β] [Fintype α] [Fintype β] : + Fintype (α ≃* β) where + elems := Equiv.instFintype.elems.filterMap + (fun e => if h : ∀ a b : α, e (a * b) = e a * e b then (⟨e, h⟩ : α ≃* β) else none) (by aesop) + complete me := (Finset.mem_filterMap ..).mpr ⟨me.toEquiv, Finset.mem_univ _, by {simp; rfl}⟩ + theorem Fintype.card_perm [Fintype α] : Fintype.card (Perm α) = (Fintype.card α)! := - Subsingleton.elim (@fintypePerm α _ _) (@equivFintype α α _ _ _ _) ▸ card_perms_of_finset _ + Subsingleton.elim (@fintypePerm α _ _) (@Equiv.instFintype α α _ _ _ _) ▸ card_perms_of_finset _ theorem Fintype.card_equiv [Fintype α] [Fintype β] (e : α ≃ β) : Fintype.card (α ≃ β) = (Fintype.card α)! := From 22229e91bbdd842ced13a70475d39215d9429b16 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 19 Nov 2024 17:06:41 +0000 Subject: [PATCH 33/90] feat: `norm_num` extension for `Rat.num` and `Rat.den` (#19099) Closes #18826 Co-authored-by: Eric Wieser --- Mathlib/Tactic/NormNum/GCD.lean | 44 ++++++++++++++++++++++++++++++++- MathlibTest/norm_num_ext.lean | 10 ++++++++ 2 files changed, 53 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/NormNum/GCD.lean b/Mathlib/Tactic/NormNum/GCD.lean index 1e3b522bdd15a..7594701886b26 100644 --- a/Mathlib/Tactic/NormNum/GCD.lean +++ b/Mathlib/Tactic/NormNum/GCD.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Kyle Miller +Authors: Mario Carneiro, Kyle Miller, Eric Wieser -/ import Mathlib.Data.Int.GCD import Mathlib.Tactic.NormNum @@ -210,6 +210,48 @@ def evalIntLCM : NormNumExt where eval {u α} e := do let ⟨ed, pf⟩ := proveIntLCM ex ey return .isNat _ ed q(isInt_lcm $p $q $pf) +theorem isInt_ratNum : ∀ {q : ℚ} {n : ℤ} {n' : ℕ} {d : ℕ}, + IsRat q n d → n.natAbs = n' → n'.gcd d = 1 → IsInt q.num n + | _, n, _, d, ⟨hi, rfl⟩, rfl, h => by + constructor + have : 0 < d := Nat.pos_iff_ne_zero.mpr <| by simpa using hi.ne_zero + simp_rw [Rat.mul_num, Rat.intCast_den, invOf_eq_inv, + Rat.inv_natCast_den_of_pos this, Rat.inv_natCast_num_of_pos this, + Rat.intCast_num, one_mul, mul_one, h, Nat.cast_one, Int.ediv_one, Int.cast_id] + +theorem isNat_ratDen : ∀ {q : ℚ} {n : ℤ} {n' : ℕ} {d : ℕ}, + IsRat q n d → n.natAbs = n' → n'.gcd d = 1 → IsNat q.den d + | _, n, _, d, ⟨hi, rfl⟩, rfl, h => by + constructor + have : 0 < d := Nat.pos_iff_ne_zero.mpr <| by simpa using hi.ne_zero + simp_rw [Rat.mul_den, Rat.intCast_den, invOf_eq_inv, + Rat.inv_natCast_den_of_pos this, Rat.inv_natCast_num_of_pos this, + Rat.intCast_num, one_mul, mul_one, Nat.cast_id, h, Nat.div_one] + +/-- Evaluates the `Rat.num` function. -/ +@[nolint unusedHavesSuffices, norm_num Rat.num _] +def evalRatNum : NormNumExt where eval {u α} e := do + let .proj _ _ (q : Q(ℚ)) ← Meta.whnfR e | failure + have : u =QL 0 := ⟨⟩; have : $α =Q ℤ := ⟨⟩; have : $e =Q Rat.num $q := ⟨⟩ + let ⟨q', n, d, eq⟩ ← deriveRat q (_inst := q(inferInstance)) + let ⟨n', hn⟩ := rawIntLitNatAbs n + -- deriveRat ensures these are coprime, so the gcd will be 1 + let ⟨gcd, pf⟩ := proveNatGCD q($n') q($d) + have : $gcd =Q nat_lit 1 := ⟨⟩ + return .isInt _ n q'.num q(isInt_ratNum $eq $hn $pf) + +/-- Evaluates the `Rat.den` function. -/ +@[nolint unusedHavesSuffices, norm_num Rat.den _] +def evalRatDen : NormNumExt where eval {u α} e := do + let .proj _ _ (q : Q(ℚ)) ← Meta.whnfR e | failure + have : u =QL 0 := ⟨⟩; have : $α =Q ℕ := ⟨⟩; have : $e =Q Rat.den $q := ⟨⟩ + let ⟨q', n, d, eq⟩ ← deriveRat q (_inst := q(inferInstance)) + let ⟨n', hn⟩ := rawIntLitNatAbs n + -- deriveRat ensures these are coprime, so the gcd will be 1 + let ⟨gcd, pf⟩ := proveNatGCD q($n') q($d) + have : $gcd =Q nat_lit 1 := ⟨⟩ + return .isNat _ d q(isNat_ratDen $eq $hn $pf) + end NormNum end Tactic diff --git a/MathlibTest/norm_num_ext.lean b/MathlibTest/norm_num_ext.lean index c864ac4fd8427..de2675f268b14 100644 --- a/MathlibTest/norm_num_ext.lean +++ b/MathlibTest/norm_num_ext.lean @@ -449,3 +449,13 @@ example : (553105253 : ℤ) ∣ 553105253 * 776531401 := by norm_num1 example : ¬ (553105253 : ℤ) ∣ 553105253 * 776531401 + 1 := by norm_num1 end mod + +section num_den + +example : (6 / 15 : ℚ).num = 2 := by norm_num1 +example : (6 / 15 : ℚ).den = 5 := by norm_num1 + +example : (-6 / 15 : ℚ).num = -2 := by norm_num1 +example : (-6 / 15 : ℚ).den = 5 := by norm_num1 + +end num_den From 7242600ff92eb1a664653dc663c7572a944e37ad Mon Sep 17 00:00:00 2001 From: PieterCuijpers Date: Tue, 19 Nov 2024 17:43:39 +0000 Subject: [PATCH 34/90] feat: quantales (#17289) First definition of Quantales - a non-commutative generalization of Locales/Frames. There are still some points of discussion, which would require changes elsewhere in Mathlib, but this definition should be workeable as a starting point (I think). It's my first PR, but I've tried to stay close to what I've seen in CompleteLattices. --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Quantale.lean | 179 ++++++++++++++++++++++++ Mathlib/Tactic/ToAdditive/Frontend.lean | 1 + docs/references.bib | 25 ++++ scripts/noshake.json | 6 +- 5 files changed, 209 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Algebra/Order/Quantale.lean diff --git a/Mathlib.lean b/Mathlib.lean index 866e1168d05e1..7684aafa754e8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -710,6 +710,7 @@ import Mathlib.Algebra.Order.Nonneg.Ring import Mathlib.Algebra.Order.Pi import Mathlib.Algebra.Order.Positive.Field import Mathlib.Algebra.Order.Positive.Ring +import Mathlib.Algebra.Order.Quantale import Mathlib.Algebra.Order.Rearrangement import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Order.Ring.Basic diff --git a/Mathlib/Algebra/Order/Quantale.lean b/Mathlib/Algebra/Order/Quantale.lean new file mode 100644 index 0000000000000..96bfa1624aa35 --- /dev/null +++ b/Mathlib/Algebra/Order/Quantale.lean @@ -0,0 +1,179 @@ +/- +Copyright (c) 2024 Pieter Cuijpers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Pieter Cuijpers +-/ +import Mathlib.Algebra.Group.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.Basic +import Mathlib.Order.CompleteLattice + +/-! +# Theory of quantales + +Quantales are the non-commutative generalization of locales/frames and as such are linked +to point-free topology and order theory. Applications are found throughout logic, +quantum mechanics, and computer science (see e.g. [Vickers1989] and [Mulvey1986]). + +The most general definition of quantale occurring in literature, is that a quantale is a semigroup +distributing over a complete sup-semilattice. In our definition below, we use the fact that +every complete sup-semilattice is in fact a complete lattice, and make constructs defined for +those immediately available. Another view could be to define a quantale as a complete idempotent +semiring, i.e. a complete semiring in which + and sup coincide. However, we will often encounter +additive quantales, i.e. quantales in which the semigroup operator is thought of as addition, in +which case the link with semirings will lead to confusion notationally. + +In this file, we follow the basic definition set out on the wikipedia page on quantales, +using a mixin typeclass to make the special cases of unital, commutative, idempotent, +integral, and involutive quantales easier to add on later. + +## Main definitions + +* `IsQuantale` and `IsAddQuantale` : Typeclass mixin for a (additive) semigroup distributing + over a complete lattice, i.e satisfying `x * (sSup s) = ⨆ y ∈ s, x * y` and + `(sSup s) * y = ⨆ x ∈ s, x * y`; + +* `leftMulResiduation`, `rightMulResiduation`, `leftAddResiduation`, `rightAddResiduation` : + Defining the left- and right- residuations of the semigroup (see notation below). + +* Finally, we provide basic distributivity laws for sSup into iSup and sup, monotonicity of + the semigroup operator, and basic equivalences for left- and right-residuation. + +## Notation + +* `x ⇨ₗ y` : `sSup {z | z * x ≤ y}`, the `leftMulResiduation` of `y` over `x`; + +* `x ⇨ᵣ y` : `sSup {z | x * z ≤ y}`, the `rightMulResiduation` of `y` over `x`; + +## References + + + + + +-/ + +open Function + +/-- An additive quantale is an additive semigroup distributing over a complete lattice. -/ +class IsAddQuantale (α : Type*) [AddSemigroup α] [CompleteLattice α] where + /-- Addition is distributive over join in a quantale -/ + protected add_sSup_distrib (x : α) (s : Set α) : x + sSup s = ⨆ y ∈ s, x + y + /-- Addition is distributive over join in a quantale -/ + protected sSup_add_distrib (s : Set α) (y : α) : sSup s + y = ⨆ x ∈ s, x + y + +/-- A quantale is a semigroup distributing over a complete lattice. -/ +@[to_additive] +class IsQuantale (α : Type*) [Semigroup α] [CompleteLattice α] where + /-- Multiplication is distributive over join in a quantale -/ + protected mul_sSup_distrib (x : α) (s : Set α) : x * sSup s = ⨆ y ∈ s, x * y + /-- Multiplication is distributive over join in a quantale -/ + protected sSup_mul_distrib (s : Set α) (y : α) : sSup s * y = ⨆ x ∈ s, x * y + +section + +variable {α : Type*} {ι : Type*} {x y z : α} {s : Set α} {f : ι → α} +variable [Semigroup α] [CompleteLattice α] [IsQuantale α] + +@[to_additive] +theorem mul_sSup_distrib : x * sSup s = ⨆ y ∈ s, x * y := IsQuantale.mul_sSup_distrib _ _ + +@[to_additive] +theorem sSup_mul_distrib : sSup s * x = ⨆ y ∈ s, y * x := IsQuantale.sSup_mul_distrib _ _ + +end + +namespace IsAddQuantale + +variable {α : Type*} {ι : Type*} {x y z : α} {s : Set α} {f : ι → α} +variable [AddSemigroup α] [CompleteLattice α] [IsAddQuantale α] + +/-- Left- and right- residuation operators on an additive quantale are similar +to the Heyting operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ₗ z ↔ x + y ≤ z` or alternatively `x ⇨ₗ y = sSup { z | z + x ≤ y }`. -/ +def leftAddResiduation (x y : α) := sSup {z | z + x ≤ y} + +/-- Left- and right- residuation operators on an additive quantale are similar +to the Heyting operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ᵣ z ↔ y + x ≤ z` or alternatively `x ⇨ₗ y = sSup { z | x + z ≤ y }`." -/ +def rightAddResiduation (x y : α) := sSup {z | x + z ≤ y} + +@[inherit_doc] +scoped infixr:60 " ⇨ₗ " => leftAddResiduation + +@[inherit_doc] +scoped infixr:60 " ⇨ᵣ " => rightAddResiduation + +end IsAddQuantale + +namespace IsQuantale + +variable {α : Type*} {ι : Type*} {x y z : α} {s : Set α} {f : ι → α} +variable [Semigroup α] [CompleteLattice α] [IsQuantale α] + +/-- Left- and right-residuation operators on an additive quantale are similar to the Heyting +operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ₗ z ↔ x * y ≤ z` or alternatively `x ⇨ₗ y = sSup { z | z * x ≤ y }`. +-/ +@[to_additive existing] +def leftMulResiduation (x y : α) := sSup {z | z * x ≤ y} + +/-- Left- and right- residuation operators on an additive quantale are similar to the Heyting +operator on complete lattices, but for a non-commutative logic. +I.e. `x ≤ y ⇨ᵣ z ↔ y * x ≤ z` or alternatively `x ⇨ₗ y = sSup { z | x * z ≤ y }`. +-/ +@[to_additive existing] +def rightMulResiduation (x y : α) := sSup {z | x * z ≤ y} + +@[inherit_doc, to_additive existing] +scoped infixr:60 " ⇨ₗ " => leftMulResiduation + +@[inherit_doc, to_additive existing] +scoped infixr:60 " ⇨ᵣ " => rightMulResiduation + +@[to_additive] +theorem mul_iSup_distrib : x * ⨆ i, f i = ⨆ i, x * f i := by + rw [iSup, mul_sSup_distrib, iSup_range] + +@[to_additive] +theorem iSup_mul_distrib : (⨆ i, f i) * x = ⨆ i, f i * x := by + rw [iSup, sSup_mul_distrib, iSup_range] + +@[to_additive] +theorem mul_sup_distrib : x * (y ⊔ z) = (x * y) ⊔ (x * z) := by + rw [← iSup_pair, ← sSup_pair, mul_sSup_distrib] + +@[to_additive] +theorem sup_mul_distrib : (x ⊔ y) * z = (x * z) ⊔ (y * z) := by + rw [← (@iSup_pair _ _ _ (fun _? => _? * z) _ _), ← sSup_pair, sSup_mul_distrib] + +@[to_additive] +instance : MulLeftMono α where + elim := by + intro _ _ _; simp only; intro + rw [← left_eq_sup, ← mul_sup_distrib, sup_of_le_left] + trivial + +@[to_additive] +instance : MulRightMono α where + elim := by + intro _ _ _; simp only; intro + rw [← left_eq_sup, ← sup_mul_distrib, sup_of_le_left] + trivial + +@[to_additive] +theorem leftMulResiduation_le_iff_mul_le : x ≤ y ⇨ₗ z ↔ x * y ≤ z where + mp h1 := by + apply le_trans (mul_le_mul_right' h1 _) + simp_all only [leftMulResiduation, sSup_mul_distrib, Set.mem_setOf_eq, + iSup_le_iff, implies_true] + mpr h1 := le_sSup h1 + +@[to_additive] +theorem rightMulResiduation_le_iff_mul_le : x ≤ y ⇨ᵣ z ↔ y * x ≤ z where + mp h1 := by + apply le_trans (mul_le_mul_left' h1 _) + simp_all only [rightMulResiduation, mul_sSup_distrib, Set.mem_setOf_eq, + iSup_le_iff, implies_true] + mpr h1 := le_sSup h1 + +end IsQuantale diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index c7a97764f023a..25e3af975e0bf 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -964,6 +964,7 @@ def nameDict : String → List String | "powers" => ["multiples"] | "multipliable"=> ["summable"] | "gpfree" => ["apfree"] + | "quantale" => ["add", "Quantale"] | x => [x] /-- diff --git a/docs/references.bib b/docs/references.bib index 5e88d18080065..8d65ce04a4a4e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2792,6 +2792,22 @@ @Article{ MR577178 url = {https://doi.org/10.1007/BF00417500} } +@Article{ Mulvey1986, + author = {Mulvey, Christopher J.}, + title = {\&}, + fjournal = {Supplemento ai Rendiconti del Circolo Matem{\`a}tico di + Palermo. Serie II}, + journal = {Suppl. Rend. Circ. Mat. Palermo (2)}, + issn = {1592-9531}, + volume = {12}, + pages = {99--104}, + year = {1986}, + language = {English}, + keywords = {46L60,46L05,81Q10}, + zbmath = {4030272}, + zbl = {0633.46065} +} + @Unpublished{ Naor-2015, author = {Assaf Naor}, title = {Metric Embeddings and Lipschitz Extensions}, @@ -3422,6 +3438,15 @@ @Book{ verdier1996 mrnumber = {1453167} } +@Book{ Vickers1989, + author = {Vickers, Steven}, + title = {Topology via Logic}, + isbn = {0-521-57651-2}, + year = {1989}, + publisher = {University of Cambridge}, + language = {English} +} + @Misc{ vistoli2004, author = {Vistoli, Angelo}, title = {Notes on {Grothendieck} topologies, fibered categories and diff --git a/scripts/noshake.json b/scripts/noshake.json index 23b0f103e1de8..a1d6020a42bd4 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -323,8 +323,7 @@ ["Mathlib.Algebra.MvPolynomial.CommRing", "Mathlib.Algebra.Polynomial.Basic"], "Mathlib.RingTheory.IntegralClosure.IsIntegral.Defs": ["Mathlib.Tactic.Algebraize"], - "Mathlib.RingTheory.Finiteness.Defs": - ["Mathlib.Tactic.Algebraize"], + "Mathlib.RingTheory.Finiteness.Defs": ["Mathlib.Tactic.Algebraize"], "Mathlib.RingTheory.Binomial": ["Mathlib.Algebra.Order.Floor"], "Mathlib.RingTheory.Adjoin.Basic": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], "Mathlib.RepresentationTheory.FdRep": @@ -365,7 +364,6 @@ "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], - "Mathlib.Data.ENat.Lattice": ["Mathlib.Algebra.Group.Action.Defs"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], "Mathlib.Data.Set.Image": ["Batteries.Tactic.Congr", "Mathlib.Data.Set.SymmDiff"], @@ -387,6 +385,7 @@ "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Insert": ["Mathlib.Data.Finset.Attr"], + "Mathlib.Data.ENat.Lattice": ["Mathlib.Algebra.Group.Action.Defs"], "Mathlib.Data.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Bool.Basic": ["Batteries.Tactic.Init"], "Mathlib.Control.Traversable.Instances": ["Mathlib.Control.Applicative"], @@ -426,6 +425,7 @@ ["Mathlib.AlgebraicTopology.AlternatingFaceMapComplex"], "Mathlib.Algebra.Star.Module": ["Mathlib.Algebra.Module.LinearMap.Star"], "Mathlib.Algebra.Ring.CentroidHom": ["Mathlib.Algebra.Algebra.Defs"], + "Mathlib.Algebra.Order.Quantale": ["Mathlib.Tactic.Variable"], "Mathlib.Algebra.Order.CauSeq.Basic": ["Mathlib.Data.Setoid.Basic"], "Mathlib.Algebra.MonoidAlgebra.Basic": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], From fdc34d21aae9609eaae66749df3579f64aaaa2c0 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 19 Nov 2024 17:43:40 +0000 Subject: [PATCH 35/90] feat: `IsSelfAdjoint.smul_iff` (#19216) Also generalizes some typeclass assumptions from `DivisionSemiring` to `GroupWithZero` and renames a few declarations. --- Mathlib/Algebra/Star/Basic.lean | 18 ++++--- Mathlib/Algebra/Star/Pointwise.lean | 3 +- Mathlib/Algebra/Star/SelfAdjoint.lean | 51 ++++++++++++++----- .../Analysis/InnerProductSpace/Symmetric.lean | 2 +- Mathlib/Analysis/RCLike/Basic.lean | 2 +- Mathlib/Data/Complex/Basic.lean | 2 +- 6 files changed, 53 insertions(+), 25 deletions(-) diff --git a/Mathlib/Algebra/Star/Basic.lean b/Mathlib/Algebra/Star/Basic.lean index 7e3efcae9d6b3..04f2f59e38df2 100644 --- a/Mathlib/Algebra/Star/Basic.lean +++ b/Mathlib/Algebra/Star/Basic.lean @@ -3,7 +3,6 @@ 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.Algebra.Field.Defs import Mathlib.Algebra.Group.Invertible.Defs import Mathlib.Algebra.GroupWithZero.Units.Lemmas import Mathlib.Algebra.Regular.Basic @@ -38,7 +37,6 @@ assert_not_exists Rat.instField universe u v w open MulOpposite -open scoped NNRat /-- Notation typeclass (with no default notation!) for an algebraic structure with a star operation. -/ @@ -365,18 +363,22 @@ open scoped ComplexConjugate end CommSemiring @[simp] -theorem star_inv' [DivisionSemiring R] [StarRing R] (x : R) : star x⁻¹ = (star x)⁻¹ := - op_injective <| (map_inv₀ (starRingEquiv : R ≃+* Rᵐᵒᵖ) x).trans (op_inv (star x)).symm +theorem star_inv₀ [GroupWithZero R] [StarMul R] (x : R) : star x⁻¹ = (star x)⁻¹ := + op_injective <| (map_inv₀ (starMulEquiv : R ≃* Rᵐᵒᵖ) x).trans (op_inv (star x)).symm + +@[deprecated (since := "2024-11-18")] alias star_inv' := star_inv₀ @[simp] -theorem star_zpow₀ [DivisionSemiring R] [StarRing R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := - op_injective <| (map_zpow₀ (starRingEquiv : R ≃+* Rᵐᵒᵖ) x z).trans (op_zpow (star x) z).symm +theorem star_zpow₀ [GroupWithZero R] [StarMul R] (x : R) (z : ℤ) : star (x ^ z) = star x ^ z := + op_injective <| (map_zpow₀ (starMulEquiv : R ≃* Rᵐᵒᵖ) x z).trans (op_zpow (star x) z).symm /-- When multiplication is commutative, `star` preserves division. -/ @[simp] -theorem star_div' [Semifield R] [StarRing R] (x y : R) : star (x / y) = star x / star y := by +theorem star_div₀ [CommGroupWithZero R] [StarMul R] (x y : R) : star (x / y) = star x / star y := by apply op_injective - rw [division_def, op_div, mul_comm, star_mul, star_inv', op_mul, op_inv] + rw [division_def, op_div, mul_comm, star_mul, star_inv₀, op_mul, op_inv] + +@[deprecated (since := "2024-11-18")] alias star_div' := star_div₀ /-- Any commutative semiring admits the trivial `*`-structure. diff --git a/Mathlib/Algebra/Star/Pointwise.lean b/Mathlib/Algebra/Star/Pointwise.lean index 75524536ac45c..ee506f2c6f037 100644 --- a/Mathlib/Algebra/Star/Pointwise.lean +++ b/Mathlib/Algebra/Star/Pointwise.lean @@ -6,6 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.Star.Basic import Mathlib.Data.Set.Finite.Basic +import Mathlib.Algebra.Field.Defs /-! # Pointwise star operation on sets @@ -118,7 +119,7 @@ protected theorem star_inv [Group α] [StarMul α] (s : Set α) : s⁻¹⋆ = s protected theorem star_inv' [DivisionSemiring α] [StarRing α] (s : Set α) : s⁻¹⋆ = s⋆⁻¹ := by ext - simp only [mem_star, mem_inv, star_inv'] + simp only [mem_star, mem_inv, star_inv₀] end Set diff --git a/Mathlib/Algebra/Star/SelfAdjoint.lean b/Mathlib/Algebra/Star/SelfAdjoint.lean index ff351d4508158..db9fa87db3a50 100644 --- a/Mathlib/Algebra/Star/SelfAdjoint.lean +++ b/Mathlib/Algebra/Star/SelfAdjoint.lean @@ -235,22 +235,38 @@ protected theorem intCast (z : ℤ) : IsSelfAdjoint (z : R) := end Ring -section DivisionSemiring +section Group -variable [DivisionSemiring R] [StarRing R] +variable [Group R] [StarMul R] @[aesop safe apply] theorem inv {x : R} (hx : IsSelfAdjoint x) : IsSelfAdjoint x⁻¹ := by - simp only [isSelfAdjoint_iff, star_inv', hx.star_eq] + simp only [isSelfAdjoint_iff, star_inv, hx.star_eq] @[aesop safe apply] theorem zpow {x : R} (hx : IsSelfAdjoint x) (n : ℤ) : IsSelfAdjoint (x ^ n) := by + simp only [isSelfAdjoint_iff, star_zpow, hx.star_eq] + +end Group + +section GroupWithZero + +variable [GroupWithZero R] [StarMul R] + +@[aesop safe apply] +theorem inv₀ {x : R} (hx : IsSelfAdjoint x) : IsSelfAdjoint x⁻¹ := by + simp only [isSelfAdjoint_iff, star_inv₀, hx.star_eq] + +@[aesop safe apply] +theorem zpow₀ {x : R} (hx : IsSelfAdjoint x) (n : ℤ) : IsSelfAdjoint (x ^ n) := by simp only [isSelfAdjoint_iff, star_zpow₀, hx.star_eq] -@[simp] -protected lemma nnratCast (q : ℚ≥0) : IsSelfAdjoint (q : R) := star_nnratCast _ +end GroupWithZero -end DivisionSemiring +@[simp] +protected lemma nnratCast [DivisionSemiring R] [StarRing R] (q : ℚ≥0) : + IsSelfAdjoint (q : R) := + star_nnratCast _ section DivisionRing @@ -267,17 +283,26 @@ section Semifield variable [Semifield R] [StarRing R] theorem div {x y : R} (hx : IsSelfAdjoint x) (hy : IsSelfAdjoint y) : IsSelfAdjoint (x / y) := by - simp only [isSelfAdjoint_iff, star_div', hx.star_eq, hy.star_eq] + simp only [isSelfAdjoint_iff, star_div₀, hx.star_eq, hy.star_eq] end Semifield section SMul -variable [Star R] [AddMonoid A] [StarAddMonoid A] [SMul R A] [StarModule R A] - @[aesop safe apply] -theorem smul {r : R} (hr : IsSelfAdjoint r) {x : A} (hx : IsSelfAdjoint x) : - IsSelfAdjoint (r • x) := by simp only [isSelfAdjoint_iff, star_smul, hr.star_eq, hx.star_eq] +theorem smul [Star R] [AddMonoid A] [StarAddMonoid A] [SMul R A] [StarModule R A] + {r : R} (hr : IsSelfAdjoint r) {x : A} (hx : IsSelfAdjoint x) : + IsSelfAdjoint (r • x) := by + simp only [isSelfAdjoint_iff, star_smul, hr.star_eq, hx.star_eq] + +theorem smul_iff [Monoid R] [StarMul R] [AddMonoid A] [StarAddMonoid A] + [MulAction R A] [StarModule R A] {r : R} (hr : IsSelfAdjoint r) (hu : IsUnit r) {x : A} : + IsSelfAdjoint (r • x) ↔ IsSelfAdjoint x := by + refine ⟨fun hrx ↦ ?_, .smul hr⟩ + lift r to Rˣ using hu + rw [← inv_smul_smul r x] + replace hr : IsSelfAdjoint r := Units.ext hr.star_eq + exact hr.inv.smul hrx end SMul @@ -384,7 +409,7 @@ section Field variable [Field R] [StarRing R] instance : Inv (selfAdjoint R) where - inv x := ⟨x.val⁻¹, x.prop.inv⟩ + inv x := ⟨x.val⁻¹, x.prop.inv₀⟩ @[simp, norm_cast] theorem val_inv (x : selfAdjoint R) : ↑x⁻¹ = (x : R)⁻¹ := @@ -398,7 +423,7 @@ theorem val_div (x y : selfAdjoint R) : ↑(x / y) = (x / y : R) := rfl instance : Pow (selfAdjoint R) ℤ where - pow x z := ⟨(x : R) ^ z, x.prop.zpow z⟩ + pow x z := ⟨(x : R) ^ z, x.prop.zpow₀ z⟩ @[simp, norm_cast] theorem val_zpow (x : selfAdjoint R) (z : ℤ) : ↑(x ^ z) = (x : R) ^ z := diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index 2fcd08ccaa66d..fe6587c770dc5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -149,7 +149,7 @@ theorem isSymmetric_iff_inner_map_self_real (T : V →ₗ[ℂ] V) : · intro h x y rw [← inner_conj_symm x (T y)] rw [inner_map_polarization T x y] - simp only [starRingEnd_apply, star_div', star_sub, star_add, star_mul] + simp only [starRingEnd_apply, star_div₀, star_sub, star_add, star_mul] simp only [← starRingEnd_apply] rw [h (x + y), h (x - y), h (x + Complex.I • y), h (x - Complex.I • y)] simp only [Complex.conj_I] diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 5390b9d9b0b33..02a0ac5840750 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -469,7 +469,7 @@ theorem div_im (z w : K) : im (z / w) = im z * re w / normSq w - re z * im w / n @[rclike_simps] -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): was `simp` theorem conj_inv (x : K) : conj x⁻¹ = (conj x)⁻¹ := - star_inv' _ + star_inv₀ _ lemma conj_div (x y : K) : conj (x / y) = conj x / conj y := map_div' conj conj_inv _ _ diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 1762273d53338..e1d6851ed9e2f 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -741,7 +741,7 @@ lemma ofReal_nnqsmul (q : ℚ≥0) (r : ℝ) : ofReal (q • r) = q • r := by lemma ofReal_qsmul (q : ℚ) (r : ℝ) : ofReal (q • r) = q • r := by simp [Rat.smul_def] theorem conj_inv (x : ℂ) : conj x⁻¹ = (conj x)⁻¹ := - star_inv' _ + star_inv₀ _ @[simp, norm_cast] theorem ofReal_div (r s : ℝ) : ((r / s : ℝ) : ℂ) = r / s := map_div₀ ofRealHom r s From 6b9f5f1e22295f304afa59e7d44b3155806eb73b Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 19 Nov 2024 17:43:41 +0000 Subject: [PATCH 36/90] fix: assert_not_exists `CStarRing` that exists but not here (#19226) [Reported on Zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/CI.3A.20noisy.20.22test.20mathlib.22/near/483123791) --- Mathlib/Topology/ContinuousMap/Bounded/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean index 40d279aec7e3a..81b40ea72345a 100644 --- a/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Bounded/Basic.lean @@ -18,7 +18,7 @@ the uniform distance. -/ -assert_not_exists CStarAlgebra.star +assert_not_exists CStarRing noncomputable section From 22fb3a454a37f3bb4cb7320df78c994e575be47d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 17:43:43 +0000 Subject: [PATCH 37/90] chore: make `Finset.subset_union_left` simp (#19249) ... and around. The corresponding `Set` lemmas already are simp From GrowthInGroups --- Mathlib/Data/Finset/Lattice/Basic.lean | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Finset/Lattice/Basic.lean b/Mathlib/Data/Finset/Lattice/Basic.lean index aeb27e8cc2029..14fc97ad84371 100644 --- a/Mathlib/Data/Finset/Lattice/Basic.lean +++ b/Mathlib/Data/Finset/Lattice/Basic.lean @@ -123,9 +123,8 @@ theorem coe_union (s₁ s₂ : Finset α) : ↑(s₁ ∪ s₂) = (s₁ ∪ s₂ theorem union_subset (hs : s ⊆ u) : t ⊆ u → s ∪ t ⊆ u := sup_le <| le_iff_subset.2 hs -theorem subset_union_left {s₁ s₂ : Finset α} : s₁ ⊆ s₁ ∪ s₂ := fun _x => mem_union_left _ - -theorem subset_union_right {s₁ s₂ : Finset α} : s₂ ⊆ s₁ ∪ s₂ := fun _x => mem_union_right _ +@[simp] lemma subset_union_left : s₁ ⊆ s₁ ∪ s₂ := fun _ ↦ mem_union_left _ +@[simp] lemma subset_union_right : s₂ ⊆ s₁ ∪ s₂ := fun _ ↦ mem_union_right _ @[gcongr] theorem union_subset_union (hsu : s ⊆ u) (htv : t ⊆ v) : s ∪ t ⊆ u ∪ v := @@ -212,9 +211,8 @@ theorem mem_of_mem_inter_right {a : α} {s₁ s₂ : Finset α} (h : a ∈ s₁ theorem mem_inter_of_mem {a : α} {s₁ s₂ : Finset α} : a ∈ s₁ → a ∈ s₂ → a ∈ s₁ ∩ s₂ := and_imp.1 mem_inter.2 -theorem inter_subset_left {s₁ s₂ : Finset α} : s₁ ∩ s₂ ⊆ s₁ := fun _a => mem_of_mem_inter_left - -theorem inter_subset_right {s₁ s₂ : Finset α} : s₁ ∩ s₂ ⊆ s₂ := fun _a => mem_of_mem_inter_right +@[simp] lemma inter_subset_left : s₁ ∩ s₂ ⊆ s₁ := fun _ ↦ mem_of_mem_inter_left +@[simp] lemma inter_subset_right : s₁ ∩ s₂ ⊆ s₂ := fun _ ↦ mem_of_mem_inter_right theorem subset_inter {s₁ s₂ u : Finset α} : s₁ ⊆ s₂ → s₁ ⊆ u → s₁ ⊆ s₂ ∩ u := by simp +contextual [subset_iff, mem_inter] From 29f545b8110e796db8de688fa9e3950ce95cf678 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 19 Nov 2024 19:36:06 +0000 Subject: [PATCH 38/90] chore: generalise `div_le_div` to groups with zero (#18917) ... and lemmas around --- Archive/Imo/Imo1998Q2.lean | 2 +- Archive/Imo/Imo2001Q2.lean | 2 +- Mathlib/Algebra/GeomSum.lean | 4 +- Mathlib/Algebra/Order/Chebyshev.lean | 2 +- Mathlib/Algebra/Order/Field/Basic.lean | 69 +++++++------------ .../Order/GroupWithZero/Unbundled.lean | 64 +++++++++++++++-- Mathlib/Algebra/Order/Monovary.lean | 9 ++- .../Algebra/Polynomial/DenomsClearable.lean | 3 +- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/Analytic/OfScalars.lean | 2 +- .../Analysis/Calculus/BumpFunction/Basic.lean | 2 +- .../Calculus/BumpFunction/Normed.lean | 2 +- Mathlib/Analysis/Calculus/MeanValue.lean | 2 +- Mathlib/Analysis/Complex/AbsMax.lean | 2 +- Mathlib/Analysis/Complex/Liouville.lean | 2 +- .../Analysis/Complex/LocallyUniformLimit.lean | 4 +- Mathlib/Analysis/Complex/Periodic.lean | 2 +- Mathlib/Analysis/Complex/Schwarz.lean | 2 +- .../Complex/UpperHalfPlane/Metric.lean | 2 +- Mathlib/Analysis/Convex/Mul.lean | 2 +- Mathlib/Analysis/Convex/Slope.lean | 16 ++--- .../Convex/SpecificFunctions/Basic.lean | 8 +-- Mathlib/Analysis/Convex/Uniform.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 2 +- .../Normed/Ring/SeminormFromBounded.lean | 10 +-- .../Normed/Ring/SeminormFromConst.lean | 6 +- Mathlib/Analysis/NormedSpace/RCLike.lean | 2 +- Mathlib/Analysis/PSeries.lean | 2 +- .../Analysis/SpecialFunctions/Log/Base.lean | 6 +- .../Analysis/SpecialFunctions/Stirling.lean | 4 +- .../SpecialFunctions/Trigonometric/Basic.lean | 6 +- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 4 +- .../Additive/PluenneckeRuzsa.lean | 3 +- .../SetFamily/AhlswedeZhang.lean | 1 + Mathlib/Combinatorics/SetFamily/LYM.lean | 2 +- Mathlib/Data/Finset/Density.lean | 2 +- Mathlib/Data/NNReal/Defs.lean | 4 +- Mathlib/Data/Rat/Cast/Order.lean | 3 +- Mathlib/Data/Real/Pi/Bounds.lean | 4 +- .../RotationNumber/TranslationNumber.lean | 4 +- .../Geometry/Euclidean/Inversion/Basic.lean | 2 +- Mathlib/GroupTheory/OrderOfElement.lean | 1 + .../MeasureTheory/Integral/PeakFunction.lean | 10 +-- Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 4 +- .../DiophantineApproximation.lean | 2 +- Mathlib/NumberTheory/FLT/Basic.lean | 1 + Mathlib/NumberTheory/Harmonic/Defs.lean | 1 + Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean | 2 +- .../NumberTheory/LSeries/HurwitzZetaEven.lean | 4 +- Mathlib/NumberTheory/Modular.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 2 +- Mathlib/NumberTheory/Pell.lean | 3 +- .../Transcendental/Liouville/Basic.lean | 2 +- .../Liouville/LiouvilleNumber.lean | 2 +- .../Transcendental/Liouville/Measure.lean | 4 +- Mathlib/Order/Interval/Set/IsoIoo.lean | 2 +- Mathlib/RingTheory/Multiplicity.lean | 3 +- .../RingTheory/Polynomial/Hermite/Basic.lean | 1 - Mathlib/RingTheory/PowerSeries/WellKnown.lean | 1 - Mathlib/SetTheory/Ordinal/Notation.lean | 1 + Mathlib/SetTheory/Surreal/Dyadic.lean | 2 +- Mathlib/Tactic/LinearCombination.lean | 1 + Mathlib/Tactic/LinearCombination/Lemmas.lean | 2 +- Mathlib/Tactic/NormNum/GCD.lean | 1 + Mathlib/Topology/Connected/PathConnected.lean | 2 +- Mathlib/Topology/MetricSpace/Contracting.lean | 2 +- Mathlib/Topology/TietzeExtension.lean | 2 +- MathlibTest/apply_rules.lean | 2 +- scripts/no_lints_prime_decls.txt | 2 +- scripts/noshake.json | 3 + 70 files changed, 194 insertions(+), 147 deletions(-) diff --git a/Archive/Imo/Imo1998Q2.lean b/Archive/Imo/Imo1998Q2.lean index 5e501114206ba..94378d663c520 100644 --- a/Archive/Imo/Imo1998Q2.lean +++ b/Archive/Imo/Imo1998Q2.lean @@ -202,7 +202,7 @@ end theorem clear_denominators {a b k : ℕ} (ha : 0 < a) (hb : 0 < b) : (b - 1 : ℚ) / (2 * b) ≤ k / a ↔ ((b : ℕ) - 1) * a ≤ k * (2 * b) := by - rw [div_le_div_iff] + rw [div_le_div_iff₀] -- Porting note: proof used to finish with `<;> norm_cast <;> simp [ha, hb]` · convert Nat.cast_le (α := ℚ) · aesop diff --git a/Archive/Imo/Imo2001Q2.lean b/Archive/Imo/Imo2001Q2.lean index a7e630353b751..98dc9266722ff 100644 --- a/Archive/Imo/Imo2001Q2.lean +++ b/Archive/Imo/Imo2001Q2.lean @@ -34,7 +34,7 @@ namespace Imo2001Q2 theorem bound (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a ^ 4 / (a ^ 4 + b ^ 4 + c ^ 4) ≤ a ^ 3 / sqrt ((a ^ 3) ^ 2 + ↑8 * b ^ 3 * c ^ 3) := by - rw [div_le_div_iff (by positivity) (by positivity)] + rw [div_le_div_iff₀ (by positivity) (by positivity)] calc a ^ 4 * sqrt ((a ^ 3) ^ 2 + (8:ℝ) * b ^ 3 * c ^ 3) = a ^ 3 * (a * sqrt ((a ^ 3) ^ 2 + (8:ℝ) * b ^ 3 * c ^ 3)) := by ring _ ≤ a ^ 3 * (a ^ 4 + b ^ 4 + c ^ 4) := ?_ diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 53093778e0883..7e6e9ce83070c 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -7,8 +7,6 @@ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Group.NatPowAssoc import Mathlib.Algebra.Order.BigOperators.Ring.Finset -import Mathlib.Algebra.Order.Field.Basic -import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Ring.Opposite import Mathlib.Tactic.Abel import Mathlib.Algebra.Ring.Regular @@ -382,7 +380,7 @@ theorem geom_sum_Ico_le_of_lt_one [LinearOrderedField α] {x : α} (hx : 0 ≤ x {m n : ℕ} : ∑ i ∈ Ico m n, x ^ i ≤ x ^ m / (1 - x) := by rcases le_or_lt m n with (hmn | hmn) · rw [geom_sum_Ico' h'x.ne hmn] - apply div_le_div (pow_nonneg hx _) _ (sub_pos.2 h'x) le_rfl + apply div_le_div₀ (pow_nonneg hx _) _ (sub_pos.2 h'x) le_rfl simpa using pow_nonneg hx _ · rw [Ico_eq_empty, sum_empty] · apply div_nonneg (pow_nonneg hx _) diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index 25c25f0275d3c..b0be97b492f03 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -165,6 +165,6 @@ theorem sum_div_card_sq_le_sum_sq_div_card : ((∑ i ∈ s, f i) / #s) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) / #s := by obtain rfl | hs := s.eq_empty_or_nonempty · simp - rw [div_pow, div_le_div_iff (by positivity) (by positivity), sq (#s : α), mul_left_comm, + rw [div_pow, div_le_div_iff₀ (by positivity) (by positivity), sq (#s : α), mul_left_comm, ← mul_assoc] exact mul_le_mul_of_nonneg_right sq_sum_le_card_mul_sum_sq (by positivity) diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 6d24b4b52ebbe..5832b57160d05 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -166,27 +166,6 @@ theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := one_le_inv_iff₀ ### Relating two divisions. -/ - -@[mono, gcongr, bound] -lemma div_le_div_of_nonneg_right (hab : a ≤ b) (hc : 0 ≤ c) : a / c ≤ b / c := by - rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] - exact mul_le_mul_of_nonneg_right hab (one_div_nonneg.2 hc) - -@[gcongr, bound] -lemma div_lt_div_of_pos_right (h : a < b) (hc : 0 < c) : a / c < b / c := by - rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] - exact mul_lt_mul_of_pos_right h (one_div_pos.2 hc) - --- Not a `mono` lemma b/c `div_le_div` is strictly more general -@[gcongr] -lemma div_le_div_of_nonneg_left (ha : 0 ≤ a) (hc : 0 < c) (h : c ≤ b) : a / b ≤ a / c := by - rw [div_eq_mul_inv, div_eq_mul_inv] - exact mul_le_mul_of_nonneg_left ((inv_le_inv₀ (hc.trans_le h) hc).mpr h) ha - -@[gcongr, bound] -lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / c := by - simpa only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ (hc.trans h) hc] - @[deprecated (since := "2024-02-16")] alias div_le_div_of_le_of_nonneg := div_le_div_of_nonneg_right @[deprecated (since := "2024-02-16")] alias div_lt_div_of_lt := div_lt_div_of_pos_right @[deprecated (since := "2024-02-16")] alias div_le_div_of_le_left := div_le_div_of_nonneg_left @@ -196,36 +175,39 @@ lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / lemma div_le_div_of_le (hc : 0 ≤ c) (hab : a ≤ b) : a / c ≤ b / c := div_le_div_of_nonneg_right hab hc -theorem div_le_div_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b := - ⟨le_imp_le_of_lt_imp_lt fun hab ↦ div_lt_div_of_pos_right hab hc, - fun hab ↦ div_le_div_of_nonneg_right hab hc.le⟩ +@[deprecated div_le_div_iff_of_pos_right (since := "2024-11-12")] +theorem div_le_div_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b := div_le_div_iff_of_pos_right hc -theorem div_lt_div_right (hc : 0 < c) : a / c < b / c ↔ a < b := - lt_iff_lt_of_le_iff_le <| div_le_div_right hc +@[deprecated div_lt_div_iff_of_pos_right (since := "2024-11-12")] +theorem div_lt_div_right (hc : 0 < c) : a / c < b / c ↔ a < b := div_lt_div_iff_of_pos_right hc -theorem div_lt_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b < a / c ↔ c < b := by - simp only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ hb hc] +@[deprecated div_lt_div_iff_of_pos_left (since := "2024-11-13")] +theorem div_lt_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b < a / c ↔ c < b := + div_lt_div_iff_of_pos_left ha hb hc +@[deprecated div_le_div_iff_of_pos_left (since := "2024-11-12")] theorem div_le_div_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b ≤ a / c ↔ c ≤ b := - le_iff_le_iff_lt_iff_lt.2 (div_lt_div_left ha hc hb) + div_le_div_iff_of_pos_left ha hb hc -theorem div_lt_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b < c / d ↔ a * d < c * b := by - rw [lt_div_iff₀ d0, div_mul_eq_mul_div, div_lt_iff₀ b0] +@[deprecated div_lt_div_iff₀ (since := "2024-11-12")] +theorem div_lt_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b < c / d ↔ a * d < c * b := + div_lt_div_iff₀ b0 d0 -theorem div_le_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := by - rw [le_div_iff₀ d0, div_mul_eq_mul_div, div_le_iff₀ b0] +@[deprecated div_le_div_iff₀ (since := "2024-11-12")] +theorem div_le_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := + div_le_div_iff₀ b0 d0 -@[mono, gcongr, bound] -theorem div_le_div (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hbd : d ≤ b) : a / b ≤ c / d := by - rw [div_le_div_iff (hd.trans_le hbd) hd] - exact mul_le_mul hac hbd hd.le hc +@[deprecated div_le_div₀ (since := "2024-11-12")] +theorem div_le_div (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hbd : d ≤ b) : a / b ≤ c / d := + div_le_div₀ hc hac hd hbd -@[gcongr] +@[deprecated div_lt_div₀ (since := "2024-11-12")] theorem div_lt_div (hac : a < c) (hbd : d ≤ b) (c0 : 0 ≤ c) (d0 : 0 < d) : a / b < c / d := - (div_lt_div_iff (d0.trans_le hbd) d0).2 (mul_lt_mul hac hbd d0 c0) + div_lt_div₀ hac hbd c0 d0 +@[deprecated div_lt_div₀' (since := "2024-11-12")] theorem div_lt_div' (hac : a ≤ c) (hbd : d < b) (c0 : 0 < c) (d0 : 0 < d) : a / b < c / d := - (div_lt_div_iff (d0.trans hbd) d0).2 (mul_lt_mul' hac hbd d0.le c0) + div_lt_div₀' hac hbd c0 d0 /-! ### Relating one division and involving `1` @@ -288,12 +270,12 @@ theorem lt_of_one_div_lt_one_div (ha : 0 < a) (h : 1 / a < 1 / b) : b < a := /-- For the single implications with fewer assumptions, see `one_div_le_one_div_of_le` and `le_of_one_div_le_one_div` -/ theorem one_div_le_one_div (ha : 0 < a) (hb : 0 < b) : 1 / a ≤ 1 / b ↔ b ≤ a := - div_le_div_left zero_lt_one ha hb + div_le_div_iff_of_pos_left zero_lt_one ha hb /-- For the single implications with fewer assumptions, see `one_div_lt_one_div_of_lt` and `lt_of_one_div_lt_one_div` -/ theorem one_div_lt_one_div (ha : 0 < a) (hb : 0 < b) : 1 / a < 1 / b ↔ b < a := - div_lt_div_left zero_lt_one ha hb + div_lt_div_iff_of_pos_left zero_lt_one ha hb theorem one_lt_one_div (h1 : 0 < a) (h2 : a < 1) : 1 < 1 / a := by rwa [lt_one_div (@zero_lt_one α _ _ _ _ _) h1, one_div_one] @@ -705,7 +687,8 @@ theorem div_two_sub_self (a : α) : a / 2 - a = -(a / 2) := by theorem add_sub_div_two_lt (h : a < b) : a + (b - a) / 2 < b := by rwa [← div_sub_div_same, sub_eq_add_neg, add_comm (b / 2), ← add_assoc, ← sub_eq_add_neg, ← - lt_sub_iff_add_lt, sub_self_div_two, sub_self_div_two, div_lt_div_right (zero_lt_two' α)] + lt_sub_iff_add_lt, sub_self_div_two, sub_self_div_two, + div_lt_div_iff_of_pos_right (zero_lt_two' α)] /-- An inequality involving `2`. -/ theorem sub_one_div_inv_le_two (a2 : 2 ≤ a) : (1 - 1 / a)⁻¹ ≤ 2 := by diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index f11dc3c1af754..431015c2b62e0 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1472,6 +1472,11 @@ lemma mul_inv_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : a * b⁻¹ ≤ 1 := lemma div_le_one_of_le₀ (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := div_le_of_le_mul₀ hb zero_le_one <| by rwa [one_mul] +@[mono, gcongr, bound] +lemma div_le_div_of_nonneg_right (hab : a ≤ b) (hc : 0 ≤ c) : a / c ≤ b / c := by + rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] + exact mul_le_mul_of_nonneg_right hab (one_div_nonneg.2 hc) + @[deprecated (since := "2024-08-21")] alias le_div_iff := le_div_iff₀ @[deprecated (since := "2024-08-21")] alias div_le_iff := div_le_iff₀ @@ -1498,6 +1503,12 @@ lemma le_inv_comm₀ (ha : 0 < a) (hb : 0 < b) : a ≤ b⁻¹ ↔ b ≤ a⁻¹ : lemma le_inv_of_le_inv₀ (ha : 0 < a) (h : a ≤ b⁻¹) : b ≤ a⁻¹ := (le_inv_comm₀ ha <| inv_pos.1 <| ha.trans_le h).1 h +-- Not a `mono` lemma b/c `div_le_div₀` is strictly more general +@[gcongr] +lemma div_le_div_of_nonneg_left (ha : 0 ≤ a) (hc : 0 < c) (h : c ≤ b) : a / b ≤ a / c := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_le_mul_of_nonneg_left ((inv_le_inv₀ (hc.trans_le h) hc).mpr h) ha + end MulPosMono section PosMulStrictMono @@ -1599,6 +1610,11 @@ lemma div_lt_iff₀ (hc : 0 < c) : b / c < a ↔ b < a * c := by lemma inv_lt_iff_one_lt_mul₀ (ha : 0 < a) : a⁻¹ < b ↔ 1 < b * a := by rw [← mul_inv_lt_iff₀ ha, one_mul] +@[gcongr, bound] +lemma div_lt_div_of_pos_right (h : a < b) (hc : 0 < c) : a / c < b / c := by + rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] + exact mul_lt_mul_of_pos_right h (one_div_pos.2 hc) + variable [PosMulStrictMono G₀] /-- See `inv_strictAnti₀` for the implication from right-to-left with one fewer assumption. -/ @@ -1623,11 +1639,16 @@ lemma lt_inv_comm₀ (ha : 0 < a) (hb : 0 < b) : a < b⁻¹ ↔ b < a⁻¹ := by lemma lt_inv_of_lt_inv₀ (ha : 0 < a) (h : a < b⁻¹) : b < a⁻¹ := (lt_inv_comm₀ ha <| inv_pos.1 <| ha.trans h).1 h +@[gcongr, bound] +lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / c := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_lt_mul_of_pos_left ((inv_lt_inv₀ (hc.trans h) hc).2 h) ha + end MulPosStrictMono end PartialOrder section LinearOrder -variable [LinearOrder G₀] [ZeroLEOneClass G₀] {a b : G₀} +variable [LinearOrder G₀] [ZeroLEOneClass G₀] {a b c d : G₀} section PosMulMono variable [PosMulMono G₀] @@ -1677,6 +1698,38 @@ lemma zpow_eq_one_iff_right₀ (ha₀ : 0 ≤ a) (ha₁ : a ≠ 1) {n : ℤ} : a · exact zero_zpow_eq_one₀ simpa using zpow_right_inj₀ ha₀ ha₁ (n := 0) +variable [MulPosStrictMono G₀] + +lemma div_le_div_iff_of_pos_right (hc : 0 < c) : a / c ≤ b / c ↔ a ≤ b where + mp := le_imp_le_of_lt_imp_lt fun hab ↦ div_lt_div_of_pos_right hab hc + mpr hab := div_le_div_of_nonneg_right hab hc.le + +lemma div_lt_div_iff_of_pos_right (hc : 0 < c) : a / c < b / c ↔ a < b := + lt_iff_lt_of_le_iff_le <| div_le_div_iff_of_pos_right hc + +lemma div_lt_div_iff_of_pos_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : + a / b < a / c ↔ c < b := by simp only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv₀ hb hc] + +lemma div_le_div_iff_of_pos_left (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : a / b ≤ a / c ↔ c ≤ b := + le_iff_le_iff_lt_iff_lt.2 (div_lt_div_iff_of_pos_left ha hc hb) + +@[mono, gcongr, bound] +lemma div_le_div₀ (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hdb : d ≤ b) : a / b ≤ c / d := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_le_mul hac ((inv_le_inv₀ (hd.trans_le hdb) hd).2 hdb) + (inv_nonneg.2 <| hd.le.trans hdb) hc + +@[gcongr] +lemma div_lt_div₀ (hac : a < c) (hdb : d ≤ b) (hc : 0 ≤ c) (hd : 0 < d) : a / b < c / d := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_lt_mul hac ((inv_le_inv₀ (hd.trans_le hdb) hd).2 hdb) (inv_pos.2 <| hd.trans_le hdb) hc + +/-- See -/ +lemma div_lt_div₀' (hac : a ≤ c) (hdb : d < b) (hc : 0 < c) (hd : 0 < d) : a / b < c / d := by + rw [div_eq_mul_inv, div_eq_mul_inv] + exact mul_lt_mul' hac ((inv_lt_inv₀ (hd.trans hdb) hd).2 hdb) + (inv_nonneg.2 <| hd.le.trans hdb.le) hc + end GroupWithZero.LinearOrder section CommSemigroupHasZero @@ -1722,8 +1775,7 @@ lemma mul_inv_le_iff₀' (hc : 0 < c) : b * c⁻¹ ≤ a ↔ b ≤ c * a := by have := posMulMono_iff_mulPosMono.1 ‹_› rw [mul_inv_le_iff₀ hc, mul_comm] -lemma div_le_div₀ (hb : 0 < b) (hd : 0 < d) : - a / b ≤ c / d ↔ a * d ≤ c * b := by +lemma div_le_div_iff₀ (hb : 0 < b) (hd : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := by have := posMulMono_iff_mulPosMono.1 ‹_› rw [div_le_iff₀ hb, ← mul_div_right_comm, le_div_iff₀ hd] @@ -1751,7 +1803,7 @@ lemma div_le_comm₀ (hb : 0 < b) (hc : 0 < c) : a / b ≤ c ↔ a / c ≤ b := end PosMulMono section PosMulStrictMono -variable [PosMulStrictMono G₀] {a b c : G₀} +variable [PosMulStrictMono G₀] {a b c d : G₀} /-- See `lt_inv_mul_iff₀` for a version with multiplication on the other side. -/ lemma lt_inv_mul_iff₀' (hc : 0 < c) : a < c⁻¹ * b ↔ a * c < b := by @@ -1771,6 +1823,10 @@ lemma mul_inv_lt_iff₀' (hc : 0 < c) : b * c⁻¹ < a ↔ b < c * a := by have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› rw [mul_inv_lt_iff₀ hc, mul_comm] +lemma div_lt_div_iff₀ (hb : 0 < b) (hd : 0 < d) : a / b < c / d ↔ a * d < c * b := by + have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› + rw [div_lt_iff₀ hb, ← mul_div_right_comm, lt_div_iff₀ hd] + /-- See `lt_div_iff₀` for a version with multiplication on the other side. -/ lemma lt_div_iff₀' (hc : 0 < c) : a < b / c ↔ c * a < b := by have := posMulStrictMono_iff_mulPosStrictMono.1 ‹_› diff --git a/Mathlib/Algebra/Order/Monovary.lean b/Mathlib/Algebra/Order/Monovary.lean index ebf26eae2fbc3..9ad0b00b7fa4d 100644 --- a/Mathlib/Algebra/Order/Monovary.lean +++ b/Mathlib/Algebra/Order/Monovary.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Module.Synonym @@ -291,19 +290,19 @@ alias ⟨Antivary.of_inv₀, Antivary.inv₀⟩ := antivary_inv₀ lemma MonovaryOn.div_left₀ (hf₁ : ∀ i ∈ s, 0 ≤ f₁ i) (hf₂ : ∀ i ∈ s, 0 < f₂ i) (h₁ : MonovaryOn f₁ g s) (h₂ : AntivaryOn f₂ g s) : MonovaryOn (f₁ / f₂) g s := - fun _i hi _j hj hij ↦ div_le_div (hf₁ _ hj) (h₁ hi hj hij) (hf₂ _ hj) <| h₂ hi hj hij + fun _i hi _j hj hij ↦ div_le_div₀ (hf₁ _ hj) (h₁ hi hj hij) (hf₂ _ hj) <| h₂ hi hj hij lemma AntivaryOn.div_left₀ (hf₁ : ∀ i ∈ s, 0 ≤ f₁ i) (hf₂ : ∀ i ∈ s, 0 < f₂ i) (h₁ : AntivaryOn f₁ g s) (h₂ : MonovaryOn f₂ g s) : AntivaryOn (f₁ / f₂) g s := - fun _i hi _j hj hij ↦ div_le_div (hf₁ _ hi) (h₁ hi hj hij) (hf₂ _ hi) <| h₂ hi hj hij + fun _i hi _j hj hij ↦ div_le_div₀ (hf₁ _ hi) (h₁ hi hj hij) (hf₂ _ hi) <| h₂ hi hj hij lemma Monovary.div_left₀ (hf₁ : 0 ≤ f₁) (hf₂ : StrongLT 0 f₂) (h₁ : Monovary f₁ g) (h₂ : Antivary f₂ g) : Monovary (f₁ / f₂) g := - fun _i _j hij ↦ div_le_div (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij + fun _i _j hij ↦ div_le_div₀ (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij lemma Antivary.div_left₀ (hf₁ : 0 ≤ f₁) (hf₂ : StrongLT 0 f₂) (h₁ : Antivary f₁ g) (h₂ : Monovary f₂ g) : Antivary (f₁ / f₂) g := - fun _i _j hij ↦ div_le_div (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij + fun _i _j hij ↦ div_le_div₀ (hf₁ _) (h₁ hij) (hf₂ _) <| h₂ hij lemma MonovaryOn.div_right₀ (hg₁ : ∀ i ∈ s, 0 ≤ g₁ i) (hg₂ : ∀ i ∈ s, 0 < g₂ i) (h₁ : MonovaryOn f g₁ s) (h₂ : AntivaryOn f g₂ s) : MonovaryOn f (g₁ / g₂) s := diff --git a/Mathlib/Algebra/Polynomial/DenomsClearable.lean b/Mathlib/Algebra/Polynomial/DenomsClearable.lean index 3cbcf3684ebd3..1f3c6c6b23d0c 100644 --- a/Mathlib/Algebra/Polynomial/DenomsClearable.lean +++ b/Mathlib/Algebra/Polynomial/DenomsClearable.lean @@ -3,8 +3,9 @@ Copyright (c) 2020 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ -import Mathlib.Algebra.Polynomial.EraseLead import Mathlib.Algebra.Algebra.Basic +import Mathlib.Algebra.Order.Ring.Abs +import Mathlib.Algebra.Polynomial.EraseLead /-! # Denominators of evaluation of polynomials at ratios diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index f2a06440fe3ef..3ebd74ae9c5f3 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -214,7 +214,7 @@ theorem lt_radius_of_isBigO (h₀ : r ≠ 0) {a : ℝ} (ha : a ∈ Ioo (-1 : ℝ rw [← pos_iff_ne_zero, ← NNReal.coe_pos] at h₀ lift a to ℝ≥0 using ha.1.le have : (r : ℝ) < r / a := by - simpa only [div_one] using (div_lt_div_left h₀ zero_lt_one ha.1).2 ha.2 + simpa only [div_one] using (div_lt_div_iff_of_pos_left h₀ zero_lt_one ha.1).2 ha.2 norm_cast at this rw [← ENNReal.coe_lt_coe] at this refine this.trans_le (p.le_radius_of_bound C fun n => ?_) diff --git a/Mathlib/Analysis/Analytic/OfScalars.lean b/Mathlib/Analysis/Analytic/OfScalars.lean index 432b903849ad7..688d6468fb486 100644 --- a/Mathlib/Analysis/Analytic/OfScalars.lean +++ b/Mathlib/Analysis/Analytic/OfScalars.lean @@ -250,7 +250,7 @@ theorem ofScalars_radius_eq_zero_of_tendsto [NormOneClass E] cases hc' <;> aesop · filter_upwards [hc.eventually_ge_atTop (2*r⁻¹), eventually_ne_atTop 0] with n hc hn simp only [ofScalars_norm, norm_mul, norm_norm, norm_pow, NNReal.norm_eq] - rw [mul_comm ‖c n‖, ← mul_assoc, ← div_le_div_iff, mul_div_assoc] + rw [mul_comm ‖c n‖, ← mul_assoc, ← div_le_div_iff₀, mul_div_assoc] · convert hc rw [pow_succ, div_mul_cancel_left₀, NNReal.coe_inv] aesop diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean index bba1643f58b06..7c20ab571383f 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean @@ -149,7 +149,7 @@ theorem support_eq : Function.support f = Metric.ball c f.rOut := by simp only [toFun, support_comp_eq_preimage, ContDiffBumpBase.support _ _ f.one_lt_rOut_div_rIn] ext x simp only [mem_ball_iff_norm, sub_zero, norm_smul, mem_preimage, Real.norm_eq_abs, abs_inv, - abs_of_pos f.rIn_pos, ← div_eq_inv_mul, div_lt_div_right f.rIn_pos] + abs_of_pos f.rIn_pos, ← div_eq_inv_mul, div_lt_div_iff_of_pos_right f.rIn_pos] theorem tsupport_eq : tsupport f = closedBall c f.rOut := by simp_rw [tsupport, f.support_eq, closure_ball _ f.rOut_pos.ne'] diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index fd3e53a98f231..540db4e96bee8 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -142,7 +142,7 @@ theorem normed_le_div_measure_closedBall_rOut [IsAddHaarMeasure μ] (K : ℝ) (h · exact f.integral_pos.le · exact f.le_one apply this.trans - rw [div_le_div_iff f.integral_pos, one_mul, ← div_le_iff₀' (pow_pos K_pos _)] + rw [div_le_div_iff₀ f.integral_pos, one_mul, ← div_le_iff₀' (pow_pos K_pos _)] · exact f.measure_closedBall_div_le_integral μ K h · exact ENNReal.toReal_pos (measure_closedBall_pos _ _ f.rOut_pos).ne' measure_closedBall_lt_top.ne diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index c91086f4d4c70..691da7a3e6562 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -117,7 +117,7 @@ theorem image_le_of_liminf_slope_right_lt_deriv_boundary' {f f' : ℝ → ℝ} { (hf'.and_eventually (HB.and (Ioc_mem_nhdsWithin_Ioi ⟨le_rfl, hy⟩))).exists refine ⟨z, ?_, hz⟩ have := (hfz.trans hzB).le - rwa [slope_def_field, slope_def_field, div_le_div_right (sub_pos.2 hz.1), hxB, + rwa [slope_def_field, slope_def_field, div_le_div_iff_of_pos_right (sub_pos.2 hz.1), hxB, sub_le_sub_iff_right] at this /-- General fencing theorem for continuous functions with an estimate on the derivative. diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 7d73410532dd9..9f5038937c5e6 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -132,7 +132,7 @@ theorem norm_max_aux₁ [CompleteSpace F] {f : ℂ → F} {z w : ℂ} exact hz (hsub hζ) show ‖(w - z)⁻¹ • f w‖ < ‖f z‖ / r rw [norm_smul, norm_inv, norm_eq_abs, ← div_eq_inv_mul] - exact (div_lt_div_right hr).2 hw_lt + exact (div_lt_div_iff_of_pos_right hr).2 hw_lt /-! Now we drop the assumption `CompleteSpace F` by embedding `F` into its completion. diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index 1075a8ade0521..2d327b794e416 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -53,7 +53,7 @@ theorem norm_deriv_le_aux [CompleteSpace F] {c : ℂ} {R C : ℝ} {f : ℂ → F have : ∀ z ∈ sphere c R, ‖(z - c) ^ (-2 : ℤ) • f z‖ ≤ C / (R * R) := fun z (hz : abs (z - c) = R) => by simpa [-mul_inv_rev, norm_smul, hz, zpow_two, ← div_eq_inv_mul] using - (div_le_div_right (mul_pos hR hR)).2 (hC z hz) + (div_le_div_iff_of_pos_right (mul_pos hR hR)).2 (hC z hz) calc ‖deriv f c‖ = ‖(2 * π * I : ℂ)⁻¹ • ∮ z in C(c, R), (z - c) ^ (-2 : ℤ) • f z‖ := congr_arg norm (deriv_eq_smul_circleIntegral hR hf) diff --git a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean index 14b603b86f48e..cb464f7dc6892 100644 --- a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean +++ b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean @@ -52,7 +52,7 @@ theorem norm_cderiv_le (hr : 0 < r) (hf : ∀ w ∈ sphere z r, ‖f w‖ ≤ M) intro w hw simp only [mem_sphere_iff_norm, norm_eq_abs] at hw simp only [norm_smul, inv_mul_eq_div, hw, norm_eq_abs, map_inv₀, Complex.abs_pow] - exact div_le_div hM (hf w hw) (sq_pos_of_pos hr) le_rfl + exact div_le_div₀ hM (hf w hw) (sq_pos_of_pos hr) le_rfl have h2 := circleIntegral.norm_integral_le_of_norm_le_const hr.le h1 simp only [cderiv, norm_smul] refine (mul_le_mul le_rfl h2 (norm_nonneg _) (norm_nonneg _)).trans (le_of_eq ?_) @@ -77,7 +77,7 @@ theorem norm_cderiv_lt (hr : 0 < r) (hfM : ∀ w ∈ sphere z r, ‖f w‖ < M) have e2 : ContinuousOn (fun w => ‖f w‖) (sphere z r) := continuous_norm.comp_continuousOn hf obtain ⟨x, hx, hx'⟩ := (isCompact_sphere z r).exists_isMaxOn e1 e2 exact ⟨‖f x‖, hfM x hx, hx'⟩ - exact (norm_cderiv_le hr hL2).trans_lt ((div_lt_div_right hr).mpr hL1) + exact (norm_cderiv_le hr hL2).trans_lt ((div_lt_div_iff_of_pos_right hr).mpr hL1) theorem norm_cderiv_sub_lt (hr : 0 < r) (hfg : ∀ w ∈ sphere z r, ‖f w - g w‖ < M) (hf : ContinuousOn f (sphere z r)) (hg : ContinuousOn g (sphere z r)) : diff --git a/Mathlib/Analysis/Complex/Periodic.lean b/Mathlib/Analysis/Complex/Periodic.lean index 2237eb83d967e..2186d2e121a79 100644 --- a/Mathlib/Analysis/Complex/Periodic.lean +++ b/Mathlib/Analysis/Complex/Periodic.lean @@ -67,7 +67,7 @@ theorem qParam_left_inv_mod_period (hh : h ≠ 0) (z : ℂ) : theorem abs_qParam_lt_iff (hh : 0 < h) (A : ℝ) (z : ℂ) : abs (qParam h z) < Real.exp (-2 * π * A / h) ↔ A < im z := by - rw [abs_qParam, Real.exp_lt_exp, div_lt_div_right hh, mul_lt_mul_left_of_neg] + rw [abs_qParam, Real.exp_lt_exp, div_lt_div_iff_of_pos_right hh, mul_lt_mul_left_of_neg] simpa using Real.pi_pos theorem qParam_tendsto (hh : 0 < h) : Tendsto (qParam h) I∞ (𝓝[≠] 0) := by diff --git a/Mathlib/Analysis/Complex/Schwarz.lean b/Mathlib/Analysis/Complex/Schwarz.lean index 9b43623115be1..91338bfc42214 100644 --- a/Mathlib/Analysis/Complex/Schwarz.lean +++ b/Mathlib/Analysis/Complex/Schwarz.lean @@ -80,7 +80,7 @@ theorem schwarz_aux {f : ℂ → ℂ} (hd : DifferentiableOn ℂ f (ball c R₁) intro z hz have hz' : z ≠ c := ne_of_mem_sphere hz hr₀.ne' rw [dslope_of_ne _ hz', slope_def_module, norm_smul, norm_inv, mem_sphere_iff_norm.1 hz, ← - div_eq_inv_mul, div_le_div_right hr₀, ← dist_eq_norm] + div_eq_inv_mul, div_le_div_iff_of_pos_right hr₀, ← dist_eq_norm] exact le_of_lt (h_maps (mem_ball.2 (by rw [mem_sphere.1 hz]; exact hr.2))) · rw [closure_ball c hr₀.ne', mem_closedBall] exact hr.1.le diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean index 79fd0dfe4e765..88a6d6642a9b9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Metric.lean @@ -79,7 +79,7 @@ protected theorem dist_comm (z w : ℍ) : dist z w = dist w z := by theorem dist_le_iff_le_sinh : dist z w ≤ r ↔ dist (z : ℂ) w / (2 * √(z.im * w.im)) ≤ sinh (r / 2) := by - rw [← div_le_div_right (zero_lt_two' ℝ), ← sinh_le_sinh, sinh_half_dist] + rw [← div_le_div_iff_of_pos_right (zero_lt_two' ℝ), ← sinh_le_sinh, sinh_half_dist] theorem dist_eq_iff_eq_sinh : dist z w = r ↔ dist (z : ℂ) w / (2 * √(z.im * w.im)) = sinh (r / 2) := by diff --git a/Mathlib/Analysis/Convex/Mul.lean b/Mathlib/Analysis/Convex/Mul.lean index 634ffe1a33097..f8a3db0f444b4 100644 --- a/Mathlib/Analysis/Convex/Mul.lean +++ b/Mathlib/Analysis/Convex/Mul.lean @@ -176,7 +176,7 @@ lemma convexOn_zpow : ∀ n : ℤ, ConvexOn 𝕜 (Ioi 0) fun x : 𝕜 ↦ x ^ n refine (convexOn_iff_forall_pos.2 ⟨convex_Ioi _, ?_⟩).pow (fun x (hx : 0 < x) ↦ by positivity) _ rintro x (hx : 0 < x) y (hy : 0 < y) a b ha hb hab field_simp - rw [div_le_div_iff, ← sub_nonneg] + rw [div_le_div_iff₀, ← sub_nonneg] · calc 0 ≤ a * b * (x - y) ^ 2 := by positivity _ = _ := by obtain rfl := eq_sub_of_add_eq hab; ring diff --git a/Mathlib/Analysis/Convex/Slope.lean b/Mathlib/Analysis/Convex/Slope.lean index 38072902d3e91..1ecc6866a18aa 100644 --- a/Mathlib/Analysis/Convex/Slope.lean +++ b/Mathlib/Analysis/Convex/Slope.lean @@ -38,7 +38,7 @@ theorem ConvexOn.slope_mono_adjacent (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx rw [hy] at key replace key := mul_le_mul_of_nonneg_left key hxz.le field_simp [a, b, mul_comm (z - x) _] at key ⊢ - rw [div_le_div_right] + rw [div_le_div_iff_of_pos_right] · linarith · nlinarith @@ -71,7 +71,7 @@ theorem StrictConvexOn.slope_strict_mono_adjacent (hf : StrictConvexOn 𝕜 s f) rw [hy] at key replace key := mul_lt_mul_of_pos_left key hxz field_simp [mul_comm (z - x) _] at key ⊢ - rw [div_lt_div_right] + rw [div_lt_div_iff_of_pos_right] · linarith · nlinarith @@ -100,7 +100,7 @@ theorem convexOn_of_slope_mono_adjacent (hs : Convex 𝕜 s) rw [← one_mul z, ← hab, add_mul] exact add_lt_add_right ((mul_lt_mul_left ha).2 hxz) _ have : (f y - f x) * (z - y) ≤ (f z - f y) * (y - x) := - (div_le_div_iff (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) + (div_le_div_iff₀ (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) have hxz : 0 < z - x := sub_pos.2 (hxy.trans hyz) have ha : (z - y) / (z - x) = a := by rw [eq_comm, ← sub_eq_iff_eq_add'] at hab @@ -145,7 +145,7 @@ theorem strictConvexOn_of_slope_strict_mono_adjacent (hs : Convex 𝕜 s) rw [← one_mul z, ← hab, add_mul] exact add_lt_add_right ((mul_lt_mul_left ha).2 hxz) _ have : (f y - f x) * (z - y) < (f z - f y) * (y - x) := - (div_lt_div_iff (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) + (div_lt_div_iff₀ (sub_pos.2 hxy) (sub_pos.2 hyz)).1 (hf hx hz hxy hyz) have hxz : 0 < z - x := sub_pos.2 (hxy.trans hyz) have ha : (z - y) / (z - x) = a := by rw [eq_comm, ← sub_eq_iff_eq_add'] at hab @@ -238,14 +238,14 @@ theorem ConvexOn.secant_mono_aux2 (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx : (hxy : x < y) (hyz : y < z) : (f y - f x) / (y - x) ≤ (f z - f x) / (z - x) := by have hxy' : 0 < y - x := by linarith have hxz' : 0 < z - x := by linarith - rw [div_le_div_iff hxy' hxz'] + rw [div_le_div_iff₀ hxy' hxz'] linarith only [hf.secant_mono_aux1 hx hz hxy hyz] theorem ConvexOn.secant_mono_aux3 (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : (f z - f x) / (z - x) ≤ (f z - f y) / (z - y) := by have hyz' : 0 < z - y := by linarith have hxz' : 0 < z - x := by linarith - rw [div_le_div_iff hxz' hyz'] + rw [div_le_div_iff₀ hxz' hyz'] linarith only [hf.secant_mono_aux1 hx hz hxy hyz] /-- If `f : 𝕜 → 𝕜` is convex, then for any point `a` the slope of the secant line of `f` through `a` @@ -284,14 +284,14 @@ theorem StrictConvexOn.secant_strict_mono_aux2 (hf : StrictConvexOn 𝕜 s f) {x (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : (f y - f x) / (y - x) < (f z - f x) / (z - x) := by have hxy' : 0 < y - x := by linarith have hxz' : 0 < z - x := by linarith - rw [div_lt_div_iff hxy' hxz'] + rw [div_lt_div_iff₀ hxy' hxz'] linarith only [hf.secant_strict_mono_aux1 hx hz hxy hyz] theorem StrictConvexOn.secant_strict_mono_aux3 (hf : StrictConvexOn 𝕜 s f) {x y z : 𝕜} (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) (hyz : y < z) : (f z - f x) / (z - x) < (f z - f y) / (z - y) := by have hyz' : 0 < z - y := by linarith have hxz' : 0 < z - x := by linarith - rw [div_lt_div_iff hxz' hyz'] + rw [div_lt_div_iff₀ hxz' hyz'] linarith only [hf.secant_strict_mono_aux1 hx hz hxy hyz] /-- If `f : 𝕜 → 𝕜` is strictly convex, then for any point `a` the slope of the secant line of `f` diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean index b293f4c6a7eb3..e58b1bbfd2213 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean @@ -110,7 +110,7 @@ theorem one_add_mul_self_lt_rpow_one_add {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ · rw [add_sub_cancel_left, log_one, sub_zero] · rw [add_sub_cancel_left, div_div, log_one, sub_zero] · apply add_lt_add_left (mul_lt_of_one_lt_left hs' hp) - · rw [← div_lt_iff₀ hp', ← div_lt_div_right hs'] + · rw [← div_lt_iff₀ hp', ← div_lt_div_iff_of_pos_right hs'] convert strictConcaveOn_log_Ioi.secant_strict_mono (zero_lt_one' ℝ) hs1 hs2 hs3 hs4 _ using 1 · rw [add_sub_cancel_left, div_div, log_one, sub_zero] · rw [add_sub_cancel_left, log_one, sub_zero] @@ -149,7 +149,7 @@ theorem rpow_one_add_lt_one_add_mul_self {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ · rw [add_sub_cancel_left, div_div, log_one, sub_zero] · rw [add_sub_cancel_left, log_one, sub_zero] · apply add_lt_add_left (lt_mul_of_lt_one_left hs' hp2) - · rw [← lt_div_iff₀ hp1, ← div_lt_div_right hs'] + · rw [← lt_div_iff₀ hp1, ← div_lt_div_iff_of_pos_right hs'] convert strictConcaveOn_log_Ioi.secant_strict_mono (zero_lt_one' ℝ) hs2 hs1 hs4 hs3 _ using 1 · rw [add_sub_cancel_left, log_one, sub_zero] · rw [add_sub_cancel_left, div_div, log_one, sub_zero] @@ -175,7 +175,7 @@ theorem strictConvexOn_rpow {p : ℝ} (hp : 1 < p) : StrictConvexOn ℝ (Ici 0) have hy' : 0 < y ^ p := rpow_pos_of_pos hy _ trans p * y ^ (p - 1) · have q : 0 < y - x := by rwa [sub_pos] - rw [div_lt_iff₀ q, ← div_lt_div_right hy', _root_.sub_div, div_self hy'.ne', + rw [div_lt_iff₀ q, ← div_lt_div_iff_of_pos_right hy', _root_.sub_div, div_self hy'.ne', ← div_rpow hx hy.le, sub_lt_comm, ← add_sub_cancel_right (x / y) 1, add_comm, add_sub_assoc, ← div_mul_eq_mul_div, mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, mul_assoc, ← div_eq_inv_mul, sub_eq_add_neg, ← mul_neg, ← neg_div, neg_sub, _root_.sub_div, @@ -186,7 +186,7 @@ theorem strictConvexOn_rpow {p : ℝ} (hp : 1 < p) : StrictConvexOn ℝ (Ici 0) · rw [sub_ne_zero] exact ((div_lt_one hy).mpr hxy).ne · have q : 0 < z - y := by rwa [sub_pos] - rw [lt_div_iff₀ q, ← div_lt_div_right hy', _root_.sub_div, div_self hy'.ne', + rw [lt_div_iff₀ q, ← div_lt_div_iff_of_pos_right hy', _root_.sub_div, div_self hy'.ne', ← div_rpow hz hy.le, lt_sub_iff_add_lt', ← add_sub_cancel_right (z / y) 1, add_comm _ 1, add_sub_assoc, ← div_mul_eq_mul_div, mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, mul_assoc, ← div_eq_inv_mul, _root_.sub_div, div_self hy.ne'] diff --git a/Mathlib/Analysis/Convex/Uniform.lean b/Mathlib/Analysis/Convex/Uniform.lean index 152127c0a4736..cbea3a0fbb080 100644 --- a/Mathlib/Analysis/Convex/Uniform.lean +++ b/Mathlib/Analysis/Convex/Uniform.lean @@ -118,7 +118,7 @@ theorem exists_forall_closed_ball_dist_add_le_two_mul_sub (hε : 0 < ε) (r : rw [← div_le_one hr, div_eq_inv_mul, ← norm_smul_of_nonneg (inv_nonneg.2 hr.le)] at hx hy have := h hx hy simp_rw [← smul_add, ← smul_sub, norm_smul_of_nonneg (inv_nonneg.2 hr.le), ← div_eq_inv_mul, - div_le_div_right hr, div_le_iff₀ hr, sub_mul] at this + div_le_div_iff_of_pos_right hr, div_le_iff₀ hr, sub_mul] at this exact this hxy end SeminormedAddCommGroup diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 5dd3a9d8f7ddd..4f081227992a4 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -483,7 +483,7 @@ theorem exp_mem_exp [RCLike 𝕜] [NormedRing A] [NormedAlgebra 𝕜 A] [Complet refine .of_norm_bounded_eventually _ (Real.summable_pow_div_factorial ‖a - ↑ₐ z‖) ?_ filter_upwards [Filter.eventually_cofinite_ne 0] with n hn rw [norm_smul, mul_comm, norm_inv, RCLike.norm_natCast, ← div_eq_mul_inv] - exact div_le_div (pow_nonneg (norm_nonneg _) n) (norm_pow_le' (a - ↑ₐ z) (zero_lt_iff.mpr hn)) + exact div_le_div₀ (pow_nonneg (norm_nonneg _) n) (norm_pow_le' (a - ↑ₐ z) (zero_lt_iff.mpr hn)) (mod_cast Nat.factorial_pos n) (mod_cast Nat.factorial_le (lt_add_one n).le) have h₀ : (∑' n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐ z) ^ (n + 1)) = (a - ↑ₐ z) * b := by simpa only [mul_smul_comm, pow_succ'] using hb.tsum_mul_left (a - ↑ₐ z) diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean index a98f45db8cc03..4aeac560e1e53 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromBounded.lean @@ -189,13 +189,13 @@ theorem seminormFromBounded_mul (f_nonneg : 0 ≤ f) simp_rw [mul_comm, hxyz, zero_div] exact div_nonneg (mul_nonneg (seminormFromBounded_nonneg f_nonneg f_mul y) (f_nonneg _)) (f_nonneg _) - · rw [div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), mul_comm (f (x * z))] + · rw [div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), mul_comm (f (x * z))] by_cases hxz : f (x * z) = 0 · rw [mul_comm x y, mul_assoc, mul_comm y, map_mul_zero_of_map_zero f_nonneg f_mul hxz y] exact mul_nonneg (seminormFromBounded_nonneg f_nonneg f_mul y) (f_nonneg _) · rw [← div_le_iff₀ (lt_of_le_of_ne' (f_nonneg _) hxz)] apply le_ciSup_of_le (seminormFromBounded_bddAbove_range f_nonneg f_mul y) (x * z) - rw [div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hxz), mul_comm x y, mul_assoc] + rw [div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hxz), mul_comm x y, mul_assoc] /-- If `f : R → ℝ` is a nonzero, nonnegative, multiplicatively bounded function, then `seminormFromBounded' f 1 = 1`. -/ @@ -242,7 +242,7 @@ theorem seminormFromBounded_add (f_nonneg : 0 ≤ f) (le_ciSup_of_le (seminormFromBounded_bddAbove_range f_nonneg f_mul y) z (le_refl _))) by_cases hz : f z = 0 · simp only [hz, div_zero, zero_add, le_refl, or_self_iff] - · rw [div_add_div_same, div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul] + · rw [div_add_div_same, div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul] exact f_add _ _ /-- `seminormFromBounded'` is a ring seminorm on `R`. -/ @@ -268,8 +268,8 @@ theorem seminormFromBounded_isNonarchimedean (f_nonneg : 0 ≤ f) · exact Or.inr <| le_ciSup_of_le (seminormFromBounded_bddAbove_range f_nonneg f_mul y) z hfy by_cases hz : f z = 0 · simp only [hz, div_zero, le_refl, or_self_iff] - · rw [div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), - div_le_div_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul, ← le_max_iff] + · rw [div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), + div_le_div_iff_of_pos_right (lt_of_le_of_ne' (f_nonneg _) hz), add_mul, ← le_max_iff] exact hna _ _ /-- If `f : R → ℝ` is a nonnegative, multiplicatively bounded function and `x : R` is diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean index c45e2c357f115..91ff0ab51079c 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean @@ -89,10 +89,10 @@ theorem seminormFromConst_seq_antitone (x : R) : Antitone (seminormFromConst_seq nth_rw 1 [← Nat.add_sub_of_le hmn] rw [pow_add, ← mul_assoc] have hc_pos : 0 < f c := lt_of_le_of_ne (apply_nonneg f _) hc.symm - apply le_trans ((div_le_div_right (pow_pos hc_pos _)).mpr (map_mul_le_mul f _ _)) + apply le_trans ((div_le_div_iff_of_pos_right (pow_pos hc_pos _)).mpr (map_mul_le_mul f _ _)) by_cases heq : m = n · have hnm : n - m = 0 := by rw [heq, Nat.sub_self n] - rw [hnm, heq, div_le_div_right (pow_pos hc_pos _), pow_zero] + rw [hnm, heq, div_le_div_iff_of_pos_right (pow_pos hc_pos _), pow_zero] conv_rhs => rw [← mul_one (f (x * c ^ n))] exact mul_le_mul_of_nonneg_left hf1 (apply_nonneg f _) · have h1 : 1 ≤ n - m := by @@ -150,7 +150,7 @@ def seminormFromConst : RingSeminorm R where (seminormFromConst_isLimit hf1 hc hpm y)) (fun n ↦ ?_) simp only [seminormFromConst_seq] rw [div_mul_div_comm, ← pow_add, two_mul, - div_le_div_right (pow_pos (lt_of_le_of_ne (apply_nonneg f _) hc.symm) _), pow_add, + div_le_div_iff_of_pos_right (pow_pos (lt_of_le_of_ne (apply_nonneg f _) hc.symm) _), pow_add, ← mul_assoc, mul_comm (x * y), ← mul_assoc, mul_assoc, mul_comm (c ^ n)] exact map_mul_le_mul f (x * c ^ n) (y * c ^ n) diff --git a/Mathlib/Analysis/NormedSpace/RCLike.lean b/Mathlib/Analysis/NormedSpace/RCLike.lean index 6b6b64a156b06..f578908f625cc 100644 --- a/Mathlib/Analysis/NormedSpace/RCLike.lean +++ b/Mathlib/Analysis/NormedSpace/RCLike.lean @@ -65,7 +65,7 @@ theorem LinearMap.bound_of_sphere_bound {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : simp only [z_zero, RCLike.ofReal_eq_zero, norm_eq_zero, Ne, not_false_iff] rw [eq, norm_mul, norm_div, RCLike.norm_coe_norm, RCLike.norm_of_nonneg r_pos.le, div_mul_eq_mul_div, div_mul_eq_mul_div, mul_comm] - apply div_le_div _ _ r_pos rfl.ge + apply div_le_div₀ _ _ r_pos rfl.ge · exact mul_nonneg ((norm_nonneg _).trans norm_f_z₁) (norm_nonneg z) apply mul_le_mul norm_f_z₁ rfl.le (norm_nonneg z) ((norm_nonneg _).trans norm_f_z₁) diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 0dc1b5b99f4d7..b077709ddf4da 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -388,7 +388,7 @@ theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : have A : 0 < (n : α) := by simpa using hk.bot_lt.trans_le hn have B : 0 < (n : α) + 1 := by linarith field_simp - rw [div_le_div_iff _ A, ← sub_nonneg] + rw [div_le_div_iff₀ _ A, ← sub_nonneg] · ring_nf rw [add_comm] exact B.le diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index 1adccc79ae6df..5e688563f5ae7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -177,7 +177,7 @@ private theorem b_ne_one' : b ≠ 1 := by linarith @[simp] theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by - rw [logb, logb, div_le_div_right (log_pos hb), log_le_log_iff h h₁] + rw [logb, logb, div_le_div_iff_of_pos_right (log_pos hb), log_le_log_iff h h₁] @[gcongr] theorem logb_le_logb_of_le (h : 0 < x) (hxy : x ≤ y) : logb b x ≤ logb b y := @@ -185,12 +185,12 @@ theorem logb_le_logb_of_le (h : 0 < x) (hxy : x ≤ y) : logb b x ≤ logb b y : @[gcongr] theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y := by - rw [logb, logb, div_lt_div_right (log_pos hb)] + rw [logb, logb, div_lt_div_iff_of_pos_right (log_pos hb)] exact log_lt_log hx hxy @[simp] theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y := by - rw [logb, logb, div_lt_div_right (log_pos hb)] + rw [logb, logb, div_lt_div_iff_of_pos_right (log_pos hb)] exact log_lt_log_iff hx hy theorem logb_le_iff_le_rpow (hx : 0 < x) : logb b x ≤ y ↔ x ≤ b ^ y := by diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 35d9ac4c47d5b..c42129b43cc42 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -122,9 +122,9 @@ theorem log_stirlingSeq_sub_log_stirlingSeq_succ (n : ℕ) : convert H using 1 <;> field_simp [h₃.ne'] refine (log_stirlingSeq_diff_le_geo_sum n).trans ?_ push_cast - rw [div_le_div_iff h₂ h₁] + rw [div_le_div_iff₀ h₂ h₁] field_simp [h₃.ne'] - rw [div_le_div_right h₃] + rw [div_le_div_iff_of_pos_right h₃] ring_nf norm_cast omega diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 4c002f27c027e..68908a896942d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -137,11 +137,11 @@ theorem pi_div_two_le_two : π / 2 ≤ 2 := by exact (Classical.choose_spec exists_cos_eq_zero).1.2 theorem two_le_pi : (2 : ℝ) ≤ π := - (div_le_div_right (show (0 : ℝ) < 2 by norm_num)).1 + (div_le_div_iff_of_pos_right (show (0 : ℝ) < 2 by norm_num)).1 (by rw [div_self (two_ne_zero' ℝ)]; exact one_le_pi_div_two) theorem pi_le_four : π ≤ 4 := - (div_le_div_right (show (0 : ℝ) < 2 by norm_num)).1 + (div_le_div_iff_of_pos_right (show (0 : ℝ) < 2 by norm_num)).1 (calc π / 2 ≤ 2 := pi_div_two_le_two _ = 4 / 2 := by norm_num) @@ -887,7 +887,7 @@ theorem tan_nonpos_of_nonpos_of_neg_pi_div_two_le {x : ℝ} (hx0 : x ≤ 0) (hpx theorem strictMonoOn_tan : StrictMonoOn tan (Ioo (-(π / 2)) (π / 2)) := by rintro x hx y hy hlt rw [tan_eq_sin_div_cos, tan_eq_sin_div_cos, - div_lt_div_iff (cos_pos_of_mem_Ioo hx) (cos_pos_of_mem_Ioo hy), mul_comm, ← sub_pos, ← sin_sub] + div_lt_div_iff₀ (cos_pos_of_mem_Ioo hx) (cos_pos_of_mem_Ioo hy), mul_comm, ← sub_pos, ← sin_sub] exact sin_pos_of_pos_of_lt_pi (sub_pos.2 hlt) <| by linarith [hx.1, hy.2] theorem tan_lt_tan_of_lt_of_lt_pi_div_two {x y : ℝ} (hx₁ : -(π / 2) < x) (hy₂ : y < π / 2) diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 819d71a7534da..d03aabb125742 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -220,7 +220,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc have cpos : 0 < c := zero_lt_one.trans hc have A : (0 : ℝ) < c⁻¹ ^ 2 := sq_pos_of_pos (inv_pos.2 cpos) have B : c ^ 2 * ((1 : ℝ) - c⁻¹ ^ 2)⁻¹ ≤ c ^ 3 * (c - 1)⁻¹ := by - rw [← div_eq_mul_inv, ← div_eq_mul_inv, div_le_div_iff _ (sub_pos.2 hc)] + rw [← div_eq_mul_inv, ← div_eq_mul_inv, div_le_div_iff₀ _ (sub_pos.2 hc)] swap · exact sub_pos.2 (pow_lt_one₀ (inv_nonneg.2 cpos.le) (inv_lt_one_of_one_lt₀ hc) two_ne_zero) have : c ^ 3 = c ^ 2 * c := by ring @@ -287,7 +287,7 @@ theorem sum_div_nat_floor_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : exact fun k hk ↦ hk.trans_le <| Nat.floor_le (by positivity) _ ≤ ∑ i ∈ range N with j < c ^ i, (1 - c⁻¹)⁻¹ ^ 2 * ((1 : ℝ) / (c ^ i) ^ 2) := by refine sum_le_sum fun i _hi => ?_ - rw [mul_div_assoc', mul_one, div_le_div_iff]; rotate_left + rw [mul_div_assoc', mul_one, div_le_div_iff₀]; rotate_left · apply sq_pos_of_pos refine zero_lt_one.trans_le ?_ simp only [Nat.le_floor, one_le_pow₀, hc.le, Nat.one_le_cast, Nat.cast_one] diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index 286aa6a748287..a65c2df01ec20 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, George Shakan -/ import Mathlib.Algebra.Group.Pointwise.Finset.Basic -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Rat import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.Enumerative.DoubleCounting @@ -164,7 +163,7 @@ private theorem mul_aux (hA : A.Nonempty) (hAB : A ⊆ B) have hA₀ : (0 : ℚ≥0) < #A := cast_pos.2 hA.card_pos have hA₀' : (0 : ℚ≥0) < #A' := cast_pos.2 hA'.card_pos exact mod_cast - (div_le_div_iff hA₀ hA₀').1 + (div_le_div_iff₀ hA₀ hA₀').1 (h _ <| mem_erase_of_ne_of_mem hA'.ne_empty <| mem_powerset.2 <| hAA'.trans hAB) /-- **Ruzsa's triangle inequality**. Multiplication version. -/ diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index d2173e136b147..be6ce35efc01e 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies, Vladimir Ivanov import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring import Mathlib.Algebra.Order.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Field.Basic import Mathlib.Data.Finset.Sups import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.Ring diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 4f771b1d44162..67a603abd4165 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -93,7 +93,7 @@ theorem card_div_choose_le_card_shadow_div_choose (hr : r ≠ 0) · rw [choose_eq_zero_of_lt hr', cast_zero, div_zero] exact div_nonneg (cast_nonneg _) (cast_nonneg _) replace h𝒜 := card_mul_le_card_shadow_mul h𝒜 - rw [div_le_div_iff] <;> norm_cast + rw [div_le_div_iff₀] <;> norm_cast · cases' r with r · exact (hr rfl).elim rw [tsub_add_eq_add_tsub hr', add_tsub_add_eq_tsub_right] at h𝒜 diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index b8fcbb8dcf9b9..b8d5585f00ed8 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -3,11 +3,11 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Fintype.Card import Mathlib.Data.NNRat.Order import Mathlib.Data.Rat.Cast.CharZero +import Mathlib.Tactic.Positivity /-! # Density of a finite set diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index ae12ffd52c40f..210cfbe55782c 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -791,13 +791,15 @@ theorem lt_div_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a < b / r ↔ r * a < b : theorem mul_lt_of_lt_div {a b r : ℝ≥0} (h : a < b / r) : a * r < b := (lt_div_iff₀ <| pos_iff_ne_zero.2 fun hr => False.elim <| by simp [hr] at h).1 h +@[deprecated div_le_div_of_nonneg_left (since := "2024-11-12")] theorem div_le_div_left_of_le {a b c : ℝ≥0} (c0 : c ≠ 0) (cb : c ≤ b) : a / b ≤ a / c := div_le_div_of_nonneg_left (zero_le _) c0.bot_lt cb +@[deprecated div_le_div_iff_of_pos_left (since := "2024-11-12")] nonrec theorem div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : 0 < c) : a / b ≤ a / c ↔ c ≤ b := - div_le_div_left a0 b0 c0 + div_le_div_iff_of_pos_left a0 b0 c0 theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) : x ≤ y := le_of_forall_ge_of_dense fun a ha => by diff --git a/Mathlib/Data/Rat/Cast/Order.lean b/Mathlib/Data/Rat/Cast/Order.lean index 637675c471ea0..618e76fd32375 100644 --- a/Mathlib/Data/Rat/Cast/Order.lean +++ b/Mathlib/Data/Rat/Cast/Order.lean @@ -6,7 +6,6 @@ Authors: Johannes Hölzl, Mario Carneiro import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Rat.Cast.CharZero import Mathlib.Tactic.Positivity.Core -import Mathlib.Algebra.Order.Field.Basic /-! # Casts of rational numbers into linear ordered fields. @@ -147,7 +146,7 @@ namespace NNRat variable {K} [LinearOrderedSemifield K] {p q : ℚ≥0} theorem cast_strictMono : StrictMono ((↑) : ℚ≥0 → K) := fun p q h => by - rwa [NNRat.cast_def, NNRat.cast_def, div_lt_div_iff, ← Nat.cast_mul, ← Nat.cast_mul, + rwa [NNRat.cast_def, NNRat.cast_def, div_lt_div_iff₀, ← Nat.cast_mul, ← Nat.cast_mul, Nat.cast_lt (α := K), ← NNRat.lt_def] · simp · simp diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index d04c703754a53..14e62e0f82c7c 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -68,7 +68,7 @@ theorem sqrtTwoAddSeries_step_up (c d : ℕ) {a b n : ℕ} {z : ℝ} (hz : sqrtT have hb' : 0 < (b : ℝ) := Nat.cast_pos.2 hb have hd' : 0 < (d : ℝ) := Nat.cast_pos.2 hd rw [sqrt_le_left (div_nonneg c.cast_nonneg d.cast_nonneg), div_pow, - add_div_eq_mul_add_div _ _ (ne_of_gt hb'), div_le_div_iff hb' (pow_pos hd' _)] + add_div_eq_mul_add_div _ _ (ne_of_gt hb'), div_le_div_iff₀ hb' (pow_pos hd' _)] exact mod_cast h /-- From a lower bound on `sqrtTwoAddSeries 0 n = 2 cos (π / 2 ^ (n+1))` of the form @@ -91,7 +91,7 @@ theorem sqrtTwoAddSeries_step_down (a b : ℕ) {c d n : ℕ} {z : ℝ} apply le_sqrt_of_sq_le have hb' : 0 < (b : ℝ) := Nat.cast_pos.2 hb have hd' : 0 < (d : ℝ) := Nat.cast_pos.2 hd - rw [div_pow, add_div_eq_mul_add_div _ _ (ne_of_gt hd'), div_le_div_iff (pow_pos hb' _) hd'] + rw [div_pow, add_div_eq_mul_add_div _ _ (ne_of_gt hd'), div_le_div_iff₀ (pow_pos hb' _) hd'] exact mod_cast h section Tactic diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 64702155e2477..0b03b2f6c460d 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -600,7 +600,7 @@ theorem tendsto_translationNumber_of_dist_bounded_aux (x : ℕ → ℝ) (C : ℝ · exact fun n => C / 2 ^ n · intro n have : 0 < (2 ^ n : ℝ) := pow_pos zero_lt_two _ - convert (div_le_div_right this).2 (H (2 ^ n)) using 1 + convert (div_le_div_iff_of_pos_right this).2 (H (2 ^ n)) using 1 rw [transnumAuxSeq, Real.dist_eq, ← sub_div, abs_div, abs_of_pos this, Real.dist_eq] · exact mul_zero C ▸ tendsto_const_nhds.mul <| tendsto_inv_atTop_zero.comp <| tendsto_pow_atTop_atTop_of_one_lt one_lt_two @@ -673,7 +673,7 @@ theorem tendsto_translation_number₀' : dsimp have : (0 : ℝ) < n + 1 := n.cast_add_one_pos rw [Real.dist_eq, div_sub' _ _ _ (ne_of_gt this), abs_div, ← Real.dist_eq, abs_of_pos this, - Nat.cast_add_one, div_le_div_right this, ← Nat.cast_add_one] + Nat.cast_add_one, div_le_div_iff_of_pos_right this, ← Nat.cast_add_one] apply dist_pow_map_zero_mul_translationNumber_le theorem tendsto_translation_number₀ : Tendsto (fun n : ℕ => (f ^ n) 0 / n) atTop (𝓝 <| τ f) := diff --git a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean index 4b75f996be818..ef7c492a1d056 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean @@ -194,7 +194,7 @@ theorem mul_dist_le_mul_dist_add_mul_dist (a b c d : P) : rw [dist_inversion_inversion hb hd, dist_inversion_inversion hb hc, dist_inversion_inversion hc hd, one_pow] at H rw [← dist_pos] at hb hc hd - rw [← div_le_div_right (mul_pos hb (mul_pos hc hd))] + rw [← div_le_div_iff_of_pos_right (mul_pos hb (mul_pos hc hd))] convert H using 1 <;> (field_simp [hb.ne', hc.ne', hd.ne', dist_comm a]; ring) end EuclideanGeometry diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 8142e71f6f653..f01189b606a68 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Subgroup.Finite import Mathlib.Algebra.Order.Group.Action +import Mathlib.Algebra.Order.Ring.Abs import Mathlib.GroupTheory.Index import Mathlib.Order.Interval.Set.Infinite diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 7d17886a1c7d2..9c33e5b413492 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -321,11 +321,11 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_measure_n _ ≤ ∫ y in s, c y ^ n ∂μ := setIntegral_mono_set (I n) (J n) (Eventually.of_forall inter_subset_right) simp_rw [φ, ← div_eq_inv_mul, div_pow, div_div] - apply div_le_div (pow_nonneg t_pos n) _ _ B - · exact pow_le_pow_left₀ (hnc _ hx.1) (ht x hx) _ - · apply mul_pos (pow_pos (t_pos.trans_lt tt') _) (ENNReal.toReal_pos (hμ v v_open x₀_v).ne' _) - have : μ (v ∩ s) ≤ μ s := measure_mono inter_subset_right - exact ne_of_lt (lt_of_le_of_lt this hs.measure_lt_top) + have := ENNReal.toReal_pos (hμ v v_open x₀_v).ne' + ((measure_mono inter_subset_right).trans_lt hs.measure_lt_top).ne + gcongr + · exact hnc _ hx.1 + · exact ht x hx have N : Tendsto (fun n => (μ (v ∩ s)).toReal⁻¹ * (t / t') ^ n) atTop (𝓝 ((μ (v ∩ s)).toReal⁻¹ * 0)) := by diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index c09a144500014..546c35be2dda7 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -280,7 +280,7 @@ theorem prehaar_pos (K₀ : PositiveCompacts G) {U : Set G} (hU : (interior U).N theorem prehaar_mono {K₀ : PositiveCompacts G} {U : Set G} (hU : (interior U).Nonempty) {K₁ K₂ : Compacts G} (h : (K₁ : Set G) ⊆ K₂.1) : prehaar (K₀ : Set G) U K₁ ≤ prehaar (K₀ : Set G) U K₂ := by - simp only [prehaar]; rw [div_le_div_right] + simp only [prehaar]; rw [div_le_div_iff_of_pos_right] · exact mod_cast index_mono K₂.2 h hU · exact mod_cast index_pos K₀ hU @@ -293,7 +293,7 @@ theorem prehaar_self {K₀ : PositiveCompacts G} {U : Set G} (hU : (interior U). theorem prehaar_sup_le {K₀ : PositiveCompacts G} {U : Set G} (K₁ K₂ : Compacts G) (hU : (interior U).Nonempty) : prehaar (K₀ : Set G) U (K₁ ⊔ K₂) ≤ prehaar (K₀ : Set G) U K₁ + prehaar (K₀ : Set G) U K₂ := by - simp only [prehaar]; rw [div_add_div_same, div_le_div_right] + simp only [prehaar]; rw [div_add_div_same, div_le_div_iff_of_pos_right] · exact mod_cast index_union_le K₁ K₂ hU · exact mod_cast index_pos K₀ hU diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 604980a89620e..cdccd001950cb 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -408,7 +408,7 @@ private theorem aux₁ : 0 < fract ξ := by have H : (2 * v - 1 : ℝ) < 1 := by refine (mul_lt_iff_lt_one_right hv₀).1 ((inv_lt_inv₀ hv₀ (mul_pos hv₁ hv₂)).1 (h.trans_le' ?_)) have h' : (⌊ξ⌋ : ℝ) - u / v = (⌊ξ⌋ * v - u) / v := by field_simp - rw [h', abs_div, abs_of_pos hv₀, ← one_div, div_le_div_right hv₀] + rw [h', abs_div, abs_of_pos hv₀, ← one_div, div_le_div_iff_of_pos_right hv₀] norm_cast rw [← zero_add (1 : ℤ), add_one_le_iff, abs_pos, sub_ne_zero] rintro rfl diff --git a/Mathlib/NumberTheory/FLT/Basic.lean b/Mathlib/NumberTheory/FLT/Basic.lean index 3630dbe1368d9..c6ebe0078285b 100644 --- a/Mathlib/NumberTheory/FLT/Basic.lean +++ b/Mathlib/NumberTheory/FLT/Basic.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Yaël Dillies, Jineon Baek import Mathlib.Algebra.EuclideanDomain.Int import Mathlib.Algebra.GCDMonoid.Finset import Mathlib.Algebra.GCDMonoid.Nat +import Mathlib.Algebra.Order.Ring.Abs import Mathlib.RingTheory.PrincipalIdealDomain /-! diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index effef54ba3920..51c8f7036e358 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Koundinya Vajjha, Thomas Browning -/ import Mathlib.Algebra.BigOperators.Intervals +import Mathlib.Algebra.Order.Field.Basic import Mathlib.Tactic.Positivity.Finset /-! diff --git a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean index 93d06720ea1b3..ae321a0cce2eb 100644 --- a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean +++ b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean @@ -274,7 +274,7 @@ lemma continuousOn_term_tsum : ContinuousOn term_tsum (Ici 1) := by refine setIntegral_mono_on ?_ ?_ measurableSet_Ioc (fun x hx ↦ ?_) · exact (term_welldef n.succ_pos (zero_lt_one.trans_le hs)).1 · exact (term_welldef n.succ_pos zero_lt_one).1 - · rw [div_le_div_left] -- leave side-goals to end and kill them all together + · rw [div_le_div_iff_of_pos_left] -- leave side-goals to end and kill them all together · apply rpow_le_rpow_of_exponent_le · exact (lt_of_le_of_lt (by simp) hx.1).le · linarith [mem_Ici.mp hs] diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index 5407eebea6353..81dee4c2c5ae9 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -519,7 +519,7 @@ lemma hasSum_int_completedCosZeta (a : ℝ) {s : ℂ} (hs : 1 < re s) : rw [mellin_div_const, completedCosZeta] congr 1 refine ((hurwitzEvenFEPair a).symm.hasMellin (?_ : 1 / 2 < (s / 2).re)).2.symm - rwa [div_ofNat_re, div_lt_div_right two_pos]] + rwa [div_ofNat_re, div_lt_div_iff_of_pos_right two_pos]] refine (hasSum_mellin_pi_mul_sq (zero_lt_one.trans hs) hF ?_).congr_fun fun n ↦ ?_ · apply (((summable_one_div_int_add_rpow 0 s.re).mpr hs).div_const 2).of_norm_bounded intro i @@ -560,7 +560,7 @@ lemma hasSum_int_completedHurwitzZetaEven (a : ℝ) {s : ℂ} (hs : 1 < re s) : ↑(if (a : UnitAddCircle) = 0 then 1 else 0 : ℝ)) / 2) (s / 2) by simp_rw [mellin_div_const, apply_ite ofReal, ofReal_one, ofReal_zero] refine congr_arg (· / 2) ((hurwitzEvenFEPair a).hasMellin (?_ : 1 / 2 < (s / 2).re)).2.symm - rwa [div_ofNat_re, div_lt_div_right two_pos]] + rwa [div_ofNat_re, div_lt_div_iff_of_pos_right two_pos]] refine (hasSum_mellin_pi_mul_sq (zero_lt_one.trans hs) hF ?_).congr_fun fun n ↦ ?_ · simp_rw [← mul_one_div ‖_‖] apply Summable.mul_left diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 6426b4dac0aad..8f86113a1bed0 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -279,7 +279,7 @@ theorem exists_max_im : ∃ g : SL(2, ℤ), ∀ g' : SL(2, ℤ), (g' • z).im obtain ⟨g, -, hg⟩ := bottom_row_surj hp_coprime refine ⟨g, fun g' => ?_⟩ rw [ModularGroup.im_smul_eq_div_normSq, ModularGroup.im_smul_eq_div_normSq, - div_le_div_left] + div_le_div_iff_of_pos_left] · simpa [← hg] using hp (g' 1) (bottom_row_coprime g') · exact z.im_pos · exact normSq_denom_pos g' z diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index e45dc9dad9d42..74e302eb5dbe2 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -267,7 +267,7 @@ private theorem newton_seq_succ_dist (n : ℕ) : newton_seq_norm_eq hnorm _ _ = ‖F.eval (newton_seq n)‖ / ‖F.derivative.eval a‖ := by rw [newton_seq_deriv_norm] _ ≤ ‖F.derivative.eval a‖ ^ 2 * T ^ 2 ^ n / ‖F.derivative.eval a‖ := - ((div_le_div_right (deriv_norm_pos hnorm)).2 (newton_seq_norm_le hnorm _)) + ((div_le_div_iff_of_pos_right (deriv_norm_pos hnorm)).2 (newton_seq_norm_le hnorm _)) _ = ‖F.derivative.eval a‖ * T ^ 2 ^ n := div_sq_cancel _ _ private theorem newton_seq_dist_aux (n : ℕ) : diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index e83f9f210649a..c3559d516dd99 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -337,7 +337,8 @@ theorem exists_of_not_isSquare (h₀ : 0 < d) (hd : ¬IsSquare d) : refine Infinite.mono (fun q h => ?_) (infinite_rat_abs_sub_lt_one_div_den_sq_of_irrational hξ) have h0 : 0 < (q.2 : ℝ) ^ 2 := pow_pos (Nat.cast_pos.mpr q.pos) 2 have h1 : (q.num : ℝ) / (q.den : ℝ) = q := mod_cast q.num_div_den - rw [mem_setOf, abs_sub_comm, ← @Int.cast_lt ℝ, ← div_lt_div_right (abs_pos_of_pos h0)] + rw [mem_setOf, abs_sub_comm, ← @Int.cast_lt ℝ, + ← div_lt_div_iff_of_pos_right (abs_pos_of_pos h0)] push_cast rw [← abs_div, abs_sq, sub_div, mul_div_cancel_right₀ _ h0.ne', ← div_pow, h1, ← sq_sqrt (Int.cast_pos.mpr h₀).le, sq_sub_sq, abs_mul, ← mul_one_div] diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean index 04403810047a7..b3587eda5cecf 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Basic.lean @@ -47,7 +47,7 @@ protected theorem irrational {x : ℝ} (h : Liouville x) : Irrational x := by have bq0 : (0 : ℝ) < b * q := mul_pos (Nat.cast_pos.mpr bN0.bot_lt) qR0 -- At a1, clear denominators... replace a1 : |a * q - b * p| * q ^ (b + 1) < b * q := by - rw [div_sub_div _ _ b0 qR0.ne', abs_div, div_lt_div_iff (abs_pos.mpr bq0.ne') (pow_pos qR0 _), + rw [div_sub_div _ _ b0 qR0.ne', abs_div, div_lt_div_iff₀ (abs_pos.mpr bq0.ne') (pow_pos qR0 _), abs_of_pos bq0, one_mul] at a1 exact mod_cast a1 -- At a0, clear denominators... diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean index 72b1c28fd82b7..2b9b52c60e496 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/LiouvilleNumber.lean @@ -136,7 +136,7 @@ theorem aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : -- [NB: in this block, I do not follow the brace convention for subgoals -- I wait until -- I solve all extraneous goals at once with `exact pow_pos (zero_lt_two.trans_le hm) _`.] -- Clear denominators and massage* - apply (div_le_div_iff _ _).mpr + apply (div_le_div_iff₀ _ _).mpr focus conv_rhs => rw [one_mul, mul_add, pow_add, mul_one, pow_mul, mul_comm, ← pow_mul] -- the second factors coincide, so we prove the inequality of the first factors* diff --git a/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean b/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean index ca97dfa42b9e9..8421fa4e65ef5 100644 --- a/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean +++ b/Mathlib/NumberTheory/Transcendental/Liouville/Measure.lean @@ -58,8 +58,8 @@ theorem setOf_liouvilleWith_subset_aux : _ ≤ (b : ℝ) ^ (2 + 1 / (n + 1 : ℕ) : ℝ) := rpow_le_rpow_of_exponent_le hb (one_le_two.trans ?_) simpa using n.cast_add_one_pos.le - rw [sub_div' _ _ _ hb0.ne', abs_div, abs_of_pos hb0, div_lt_div_right hb0, abs_sub_lt_iff, - sub_lt_iff_lt_add, sub_lt_iff_lt_add, ← sub_lt_iff_lt_add'] at hlt + rw [sub_div' _ _ _ hb0.ne', abs_div, abs_of_pos hb0, div_lt_div_iff_of_pos_right hb0, + abs_sub_lt_iff, sub_lt_iff_lt_add, sub_lt_iff_lt_add, ← sub_lt_iff_lt_add'] at hlt rw [Finset.mem_Icc, ← Int.lt_add_one_iff, ← Int.lt_add_one_iff, ← neg_lt_iff_pos_add, add_comm, ← @Int.cast_lt ℝ, ← @Int.cast_lt ℝ] push_cast diff --git a/Mathlib/Order/Interval/Set/IsoIoo.lean b/Mathlib/Order/Interval/Set/IsoIoo.lean index f44e1e4822571..649d99a424435 100644 --- a/Mathlib/Order/Interval/Set/IsoIoo.lean +++ b/Mathlib/Order/Interval/Set/IsoIoo.lean @@ -35,7 +35,7 @@ def orderIsoIooNegOneOne (k : Type*) [LinearOrderedField k] : k ≃o Ioo (-1 : k · intro x simp only [abs_neg, neg_div] · rintro x (hx : 0 ≤ x) y (hy : 0 ≤ y) hxy - simp [abs_of_nonneg, mul_add, mul_comm x y, div_lt_div_iff, hx.trans_lt (lt_one_add _), + simp [abs_of_nonneg, mul_add, mul_comm x y, div_lt_div_iff₀, hx.trans_lt (lt_one_add _), hy.trans_lt (lt_one_add _), *] · refine fun x ↦ Subtype.ext ?_ have : 0 < 1 - |(x : k)| := sub_pos.2 (abs_lt.2 x.2) diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index 85ab17ad0be89..c095227abd535 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -5,8 +5,9 @@ Authors: Robert Y. Lewis, Chris Hughes, Daniel Weber -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.BigOperators.Group.Finset -import Mathlib.Tactic.Linarith +import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Data.ENat.Basic +import Mathlib.Tactic.Linarith /-! # Multiplicity of a divisor diff --git a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean index fa81775947129..13f33a3208377 100644 --- a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Luke Mantle. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Luke Mantle -/ -import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Polynomial.Derivative import Mathlib.Data.Nat.Factorial.DoubleFactorial diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 292659bb8862f..ef67977a0a5d2 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Rat import Mathlib.Algebra.BigOperators.NatAntidiagonal -import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Data.Nat.Choose.Sum import Mathlib.RingTheory.PowerSeries.Basic diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index f5a5db19acb40..02806888024eb 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Data.Ordering.Lemmas import Mathlib.Data.PNat.Basic import Mathlib.SetTheory.Ordinal.Principal diff --git a/Mathlib/SetTheory/Surreal/Dyadic.lean b/Mathlib/SetTheory/Surreal/Dyadic.lean index 66e64a35705b5..b1dc9275095fb 100644 --- a/Mathlib/SetTheory/Surreal/Dyadic.lean +++ b/Mathlib/SetTheory/Surreal/Dyadic.lean @@ -5,12 +5,12 @@ Authors: Apurva Nakade -/ import Mathlib.Algebra.Algebra.Defs import Mathlib.Algebra.Order.Group.Basic +import Mathlib.Algebra.Ring.Regular import Mathlib.GroupTheory.MonoidLocalization.Away import Mathlib.RingTheory.Localization.Defs import Mathlib.SetTheory.Game.Birthday import Mathlib.SetTheory.Surreal.Multiplication import Mathlib.Tactic.Linarith -import Mathlib.Algebra.Ring.Regular /-! # Dyadic numbers diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index f1c65e78a805b..5c891e021eae9 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Abby J. Goldberg, Mario Carneiro, Heather Macbeth -/ import Mathlib.Tactic.LinearCombination.Lemmas +import Mathlib.Tactic.Positivity.Core import Mathlib.Tactic.Ring import Mathlib.Tactic.Ring.Compare diff --git a/Mathlib/Tactic/LinearCombination/Lemmas.lean b/Mathlib/Tactic/LinearCombination/Lemmas.lean index 2d93dbab83217..e22bec542ae8b 100644 --- a/Mathlib/Tactic/LinearCombination/Lemmas.lean +++ b/Mathlib/Tactic/LinearCombination/Lemmas.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Abby J. Goldberg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Abby J. Goldberg, Mario Carneiro, Heather Macbeth -/ +import Mathlib.Algebra.Order.Field.Defs import Mathlib.Data.Ineq -import Mathlib.Algebra.Order.Field.Basic /-! # Lemmas for the linear_combination tactic diff --git a/Mathlib/Tactic/NormNum/GCD.lean b/Mathlib/Tactic/NormNum/GCD.lean index 7594701886b26..cf43344d0f6e3 100644 --- a/Mathlib/Tactic/NormNum/GCD.lean +++ b/Mathlib/Tactic/NormNum/GCD.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kyle Miller, Eric Wieser -/ +import Mathlib.Algebra.Ring.Divisibility.Basic import Mathlib.Data.Int.GCD import Mathlib.Tactic.NormNum diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 82ff89fe8d4c4..c5d94601f96f7 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -330,7 +330,7 @@ theorem trans_range {a b c : X} (γ₁ : Path a b) (γ₂ : Path b c) : rwa [coe_mk_mk, Function.comp_apply, if_neg h] at hxt · rintro x (⟨⟨t, ht0, ht1⟩, hxt⟩ | ⟨⟨t, ht0, ht1⟩, hxt⟩) · use ⟨t / 2, ⟨by linarith, by linarith⟩⟩ - have : t / 2 ≤ 1 / 2 := (div_le_div_right (zero_lt_two : (0 : ℝ) < 2)).mpr ht1 + have : t / 2 ≤ 1 / 2 := (div_le_div_iff_of_pos_right (zero_lt_two : (0 : ℝ) < 2)).mpr ht1 rw [coe_mk_mk, Function.comp_apply, if_pos this, Subtype.coe_mk] ring_nf rwa [γ₁.extend_extends] diff --git a/Mathlib/Topology/MetricSpace/Contracting.lean b/Mathlib/Topology/MetricSpace/Contracting.lean index e70e951d65294..a4c5c2b74235d 100644 --- a/Mathlib/Topology/MetricSpace/Contracting.lean +++ b/Mathlib/Topology/MetricSpace/Contracting.lean @@ -266,7 +266,7 @@ theorem dist_fixedPoint_fixedPoint_of_dist_le' (g : α → α) {x y} (hx : IsFix dist x y = dist y x := dist_comm x y _ ≤ dist y (f y) / (1 - K) := hf.dist_le_of_fixedPoint y hx _ = dist (f y) (g y) / (1 - K) := by rw [hy.eq, dist_comm] - _ ≤ C / (1 - K) := (div_le_div_right hf.one_sub_K_pos).2 (hfg y) + _ ≤ C / (1 - K) := (div_le_div_iff_of_pos_right hf.one_sub_K_pos).2 (hfg y) variable [Nonempty α] [CompleteSpace α] variable (f) diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 37c54d6ae46ae..84b683fbeb2fe 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -181,7 +181,7 @@ theorem tietze_extension_step (f : X →ᵇ ℝ) (e : C(X, Y)) (he : IsClosedEmb are disjoint, hence by Urysohn's lemma there exists a function `g` that is equal to `-‖f‖ / 3` on the former set and is equal to `‖f‖ / 3` on the latter set. This function `g` satisfies the assertions of the lemma. -/ - have hf3 : -‖f‖ / 3 < ‖f‖ / 3 := (div_lt_div_right h3).2 (Left.neg_lt_self hf) + have hf3 : -‖f‖ / 3 < ‖f‖ / 3 := (div_lt_div_iff_of_pos_right h3).2 (Left.neg_lt_self hf) have hc₁ : IsClosed (e '' (f ⁻¹' Iic (-‖f‖ / 3))) := he.isClosedMap _ (isClosed_Iic.preimage f.continuous) have hc₂ : IsClosed (e '' (f ⁻¹' Ici (‖f‖ / 3))) := diff --git a/MathlibTest/apply_rules.lean b/MathlibTest/apply_rules.lean index d621b7bb85ae8..ab93b4c736a69 100644 --- a/MathlibTest/apply_rules.lean +++ b/MathlibTest/apply_rules.lean @@ -49,4 +49,4 @@ example [LinearOrderedField α] {a b : α} (hb : 0 ≤ b) (hab : a ≤ b) : a / fail_if_success apply_rules (config := { transparency := .reducible }) [mul_le_mul] guard_target = a / 2 ≤ b / 2 - exact div_le_div hb hab zero_lt_two le_rfl + exact div_le_div₀ hb hab zero_lt_two le_rfl diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index f00efff46d47e..7a7e8fb75bfe5 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -1088,7 +1088,6 @@ div_eq_iff_eq_mul' div_eq_of_eq_mul' div_eq_of_eq_mul'' div_le_div'' -div_le_div_iff' div_le_div_left' div_le_div_right' div_left_inj' @@ -1097,6 +1096,7 @@ div_le_iff_le_mul' div_le_iff_of_neg' div_le_one' div_lt_div' +div_lt_div₀' div_lt_div'' div_lt_div_iff' div_lt_div_left' diff --git a/scripts/noshake.json b/scripts/noshake.json index a1d6020a42bd4..736e94310eaf9 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -340,6 +340,7 @@ "Mathlib.Order.BoundedOrder": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical", "Mathlib.Tactic.GCongr.Core"], + "Mathlib.NumberTheory.Harmonic.Defs": ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.NumberTheory.Cyclotomic.Basic": ["Mathlib.Init.Core"], "Mathlib.NumberTheory.ArithmeticFunction": ["Mathlib.Tactic.ArithMult"], "Mathlib.MeasureTheory.MeasurableSpace.Defs": ["Mathlib.Tactic.FunProp.Attr"], @@ -392,6 +393,8 @@ "Mathlib.Control.Monad.Cont": ["Batteries.Lean.Except", "Batteries.Tactic.Congr"], "Mathlib.Control.Basic": ["Mathlib.Logic.Function.Defs"], + "Mathlib.Combinatorics.SimpleGraph.Density": + ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.Combinatorics.SetFamily.AhlswedeZhang": ["Mathlib.Algebra.Order.Field.Basic"], "Mathlib.CategoryTheory.Sites.IsSheafFor": From 4e69bf9a2bbce8bca6666996a53d842121ddf304 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 19 Nov 2024 22:02:12 +0000 Subject: [PATCH 39/90] feat(SetTheory/ZFC/Ordinal): Initial development of ordinals in ZFSet (#15793) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is ported from the Mathlib 3 branch [`von_neumann_v2`](https://github.com/leanprover-community/mathlib3/blob/5b3c1d3d66838f426b44cae25b615f26438c8acb/src/set_theory/zfc/ordinal.lean). This is meant as a small, initial PR to iron out any kinks (with the definition, notation, auto-ported lemmas, etc.) before fully developing the file. Of particular note is the quite bare-bones definition of an ordinal. Surprisingly, our weakened transitivity assumption is equivalent to the much stronger claim that ordinals are well-ordered under `∈`; see #17001 for a proof. --- Mathlib/SetTheory/ZFC/Ordinal.lean | 68 ++++++++++++++++++++++++------ 1 file changed, 54 insertions(+), 14 deletions(-) diff --git a/Mathlib/SetTheory/ZFC/Ordinal.lean b/Mathlib/SetTheory/ZFC/Ordinal.lean index 56dbcfb274d7a..bd816b21e5dac 100644 --- a/Mathlib/SetTheory/ZFC/Ordinal.lean +++ b/Mathlib/SetTheory/ZFC/Ordinal.lean @@ -3,26 +3,27 @@ Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ +import Mathlib.Order.RelIso.Set import Mathlib.SetTheory.ZFC.Basic /-! # Von Neumann ordinals This file works towards the development of von Neumann ordinals, i.e. transitive sets, well-ordered -under `∈`. We currently only have an initial development of transitive sets. +under `∈`. -Further development can be found on the branch `von_neumann_v2`. +We currently only have an initial development of transitive sets and ordinals. Further development +can be found on the Mathlib 3 branch `von_neumann_v2`. ## Definitions - `ZFSet.IsTransitive` means that every element of a set is a subset. +- `ZFSet.IsOrdinal` means that the set is transitive and well-ordered under `∈`. ## TODO -- Define von Neumann ordinals. -- Define the basic arithmetic operations on ordinals from a purely set-theoretic perspective. -- Prove the equivalences between these definitions and those provided in - `SetTheory/Ordinal/Arithmetic.lean`. +- Define von Neumann ordinals and the von Neumann hierarchy. +- Build correspondences between these notions and those of the standard `Ordinal` type. -/ @@ -38,10 +39,12 @@ def IsTransitive (x : ZFSet) : Prop := ∀ y ∈ x, y ⊆ x @[simp] -theorem empty_isTransitive : IsTransitive ∅ := fun y hy => (not_mem_empty y hy).elim +theorem isTransitive_empty : IsTransitive ∅ := fun y hy => (not_mem_empty y hy).elim -theorem IsTransitive.subset_of_mem (h : x.IsTransitive) : y ∈ x → y ⊆ x := - h y +@[deprecated isTransitive_empty (since := "2024-09-21")] +alias empty_isTransitive := isTransitive_empty + +theorem IsTransitive.subset_of_mem (h : x.IsTransitive) : y ∈ x → y ⊆ x := h y theorem isTransitive_iff_mem_trans : z.IsTransitive ↔ ∀ {x y : ZFSet}, x ∈ y → y ∈ z → x ∈ z := ⟨fun h _ _ hx hy => h.subset_of_mem hy hx, fun H _ hx _ hy => H hy hx⟩ @@ -66,7 +69,7 @@ theorem IsTransitive.sUnion' (H : ∀ y ∈ x, IsTransitive y) : protected theorem IsTransitive.union (hx : x.IsTransitive) (hy : y.IsTransitive) : (x ∪ y).IsTransitive := by rw [← sUnion_pair] - apply IsTransitive.sUnion' fun z => _ + apply IsTransitive.sUnion' intro rw [mem_pair] rintro (rfl | rfl) @@ -77,10 +80,12 @@ protected theorem IsTransitive.powerset (h : x.IsTransitive) : (powerset x).IsTr rw [mem_powerset] at hy ⊢ exact h.subset_of_mem (hy hz) -theorem isTransitive_iff_sUnion_subset : x.IsTransitive ↔ (⋃₀ x : ZFSet) ⊆ x := - ⟨fun h y hy => by - rcases mem_sUnion.1 hy with ⟨z, hz, hz'⟩ - exact h.mem_trans hz' hz, fun H _ hy _ hz => H <| mem_sUnion_of_mem hz hy⟩ +theorem isTransitive_iff_sUnion_subset : x.IsTransitive ↔ (⋃₀ x : ZFSet) ⊆ x := by + constructor <;> + intro h y hy + · obtain ⟨z, hz, hz'⟩ := mem_sUnion.1 hy + exact h.mem_trans hz' hz + · exact fun z hz ↦ h <| mem_sUnion_of_mem hz hy alias ⟨IsTransitive.sUnion_subset, _⟩ := isTransitive_iff_sUnion_subset @@ -89,4 +94,39 @@ theorem isTransitive_iff_subset_powerset : x.IsTransitive ↔ x ⊆ powerset x : alias ⟨IsTransitive.subset_powerset, _⟩ := isTransitive_iff_subset_powerset +/-- A set `x` is a von Neumann ordinal when it's a transitive set, that's transitive under `∈`. We +will prove that this further implies that `x` is well-ordered under `∈`. + +The transitivity condition `a ∈ b → b ∈ c → a ∈ c` can be written without assuming `a ∈ x` and +`b ∈ x`. The lemma `isOrdinal_iff_isTrans` shows this condition is equivalent to the usual one. -/ +structure IsOrdinal (x : ZFSet) : Prop where + /-- An ordinal is a transitive set. -/ + isTransitive : x.IsTransitive + /-- The membership operation within an ordinal is transitive. -/ + mem_trans' {y z w : ZFSet} : y ∈ z → z ∈ w → w ∈ x → y ∈ w + +namespace IsOrdinal + +theorem subset_of_mem (h : x.IsOrdinal) : y ∈ x → y ⊆ x := + h.isTransitive.subset_of_mem + +theorem mem_trans (h : z.IsOrdinal) : x ∈ y → y ∈ z → x ∈ z := + h.isTransitive.mem_trans + +protected theorem isTrans (h : x.IsOrdinal) : IsTrans x.toSet (Subrel (· ∈ ·) _) := + ⟨fun _ _ c hab hbc => h.mem_trans' hab hbc c.2⟩ + +end IsOrdinal + +/-- The simplified form of transitivity used within `IsOrdinal` yields an equivalent definition to +the standard one. -/ +theorem isOrdinal_iff_isTrans : + x.IsOrdinal ↔ x.IsTransitive ∧ IsTrans x.toSet (Subrel (· ∈ ·) _) := by + use fun h => ⟨h.isTransitive, h.isTrans⟩ + rintro ⟨h₁, ⟨h₂⟩⟩ + use h₁ + intro y z w hyz hzw hwx + have hzx := h₁.mem_trans hzw hwx + exact h₂ ⟨y, h₁.mem_trans hyz hzx⟩ ⟨z, hzx⟩ ⟨w, hwx⟩ hyz hzw + end ZFSet From 3e72bf57f3bcbb9ab65c5cda0c6c0db55fc1f7dc Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 20 Nov 2024 09:10:52 +1100 Subject: [PATCH 40/90] fixes --- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 1 + Mathlib/Data/List/Permutation.lean | 4 ++-- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 03c9d7579e5f1..5b90caf6e8964 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,6 +537,7 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} +set_option linter.docPrime false in @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 09092d9f452e8..74a32e39c3263 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -500,13 +500,13 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : simp [hm', ← hn', hn] rcases lt_trichotomy n m with (ht | ht | ht) · suffices x ∈ bs by exact h x (hb.subset this) rfl - rw [← hx', getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hm)] + rw [← hx', getElem_insertIdx_of_lt ht] exact getElem_mem _ · simp only [ht] at hm' hn' rw [← hm'] at hn' exact H (insertIdx_injective _ _ hn') · suffices x ∈ as by exact h x (ha.subset this) rfl - rw [← hx, getElem_insertIdx_of_lt _ _ _ _ ht (ht.trans_le hn)] + rw [← hx, getElem_insertIdx_of_lt ht] exact getElem_mem _ lemma permutations_take_two (x y : α) (s : List α) : From c3a2d7bbf53eb210617d830940228487ff55d3d6 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 19 Nov 2024 22:30:50 +0000 Subject: [PATCH 41/90] feat: `#parse` -- a command to parse text and log outputs (#16305) This is a spin-off of #16230. It introduces a function written by Matt that allows to use `#guard_msgs` to check parsing errors. The new command is called `#guard_exceptions` and mimics very closely `#guard_msgs`. As a way of showing how to use it, the PR also adds tests for Stacks Project Tags. Co-authored-by: Eric Wieser Co-authored-by: Matthew Ballard --- Mathlib.lean | 1 + Mathlib/Util/ParseCommand.lean | 57 ++++++++++++++++++++++++++++++++ MathlibTest/StacksAttribute.lean | 16 +++++++++ 3 files changed, 74 insertions(+) create mode 100644 Mathlib/Util/ParseCommand.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7684aafa754e8..3ca25acb72351 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5277,6 +5277,7 @@ import Mathlib.Util.IncludeStr import Mathlib.Util.LongNames import Mathlib.Util.MemoFix import Mathlib.Util.Notation3 +import Mathlib.Util.ParseCommand import Mathlib.Util.Qq import Mathlib.Util.SleepHeartbeats import Mathlib.Util.Superscript diff --git a/Mathlib/Util/ParseCommand.lean b/Mathlib/Util/ParseCommand.lean new file mode 100644 index 0000000000000..d3c71c46e9aac --- /dev/null +++ b/Mathlib/Util/ParseCommand.lean @@ -0,0 +1,57 @@ +/- +Copyright (c) 2024 Matthew Robert Ballard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Matthew Robert Ballard, Damiano Testa +-/ + +import Lean.Elab.Command +import Mathlib.Init + +/-! +# `#parse` -- a command to parse text and log outputs +-/ + +namespace Mathlib.GuardExceptions + +open Lean Parser Elab Command +/-- `captureException env s input` uses the given `Environment` `env` to parse the `String` `input` +using the `ParserFn` `s`. + +This is a variation of `Lean.Parser.runParserCategory`. +-/ +def captureException (env : Environment) (s : ParserFn) (input : String) : Except String Syntax := + let ictx := mkInputContext input "" + let s := s.run ictx { env, options := {} } (getTokenTable env) (mkParserState input) + if !s.allErrors.isEmpty then + .error (s.toErrorMsg ictx) + else if ictx.input.atEnd s.pos then + .ok s.stxStack.back + else + .error ((s.mkError "end of input").toErrorMsg ictx) + +/-- `#parse parserFnId => str` allows to capture parsing exceptions. +`parserFnId` is the identifier of a `ParserFn` and `str` is the string that +`parserFnId` should parse. + +If the parse is successful, then the output is logged; +if the parse is successful, then the output is captured in an exception. + +In either case, `#guard_msgs` can then be used to capture the resulting parsing errors. + +For instance, `#parse` can be used as follows +```lean +/-- error: :1:3: Stacks tags must be exactly 4 characters -/ +#guard_msgs in #parse Mathlib.Stacks.stacksTagFn => "A05" +``` +-/ +syntax (name := parseCmd) "#parse " ident " => " str : command + +@[inherit_doc parseCmd] +elab_rules : command + | `(command| #parse $parserFnId => $str) => do + elabCommand <| ← `(command| + run_cmd do + let exc ← Lean.ofExcept <| captureException (← getEnv) $parserFnId $str + logInfo $str) + +end Mathlib.GuardExceptions diff --git a/MathlibTest/StacksAttribute.lean b/MathlibTest/StacksAttribute.lean index 3b70bdb597610..6fbee7f267bc2 100644 --- a/MathlibTest/StacksAttribute.lean +++ b/MathlibTest/StacksAttribute.lean @@ -1,4 +1,5 @@ import Mathlib.Tactic.StacksAttribute +import Mathlib.Util.ParseCommand /-- info: No tags found. -/ #guard_msgs in @@ -30,6 +31,21 @@ example : True := .intro @[stacks 0X14 "I can also have a comment"] example : True := .intro +@[stacks 0BR2, stacks 0X14 "I can also have a comment"] +example : True := .intro + +@[stacks 0X14 "I can also have a comment"] +example : True := .intro + +/-- error: :1:3: Stacks tags must be exactly 4 characters -/ +#guard_msgs in #parse Mathlib.StacksTag.stacksTagFn => "A05" + +/-- error: :1:4: Stacks tags must consist only of digits and uppercase letters. -/ +#guard_msgs in #parse Mathlib.StacksTag.stacksTagFn => "A05b" + +/-- info: 0BD5 -/ +#guard_msgs in #parse Mathlib.StacksTag.stacksTagFn => "0BD5" + /-- info: [Stacks Tag A04Q](https://stacks.math.columbia.edu/tag/A04Q) corresponds to declaration 'X.tagged'. (A comment) From 25b91b6001b15abdba13e10b8dcf838cfe9d292f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 20 Nov 2024 00:01:29 +0000 Subject: [PATCH 42/90] feat: commutativity of `NonUnital{Star}Algebra.adjoin` (#18612) --- .../Algebra/Algebra/NonUnitalSubalgebra.lean | 51 +++++++++++++ Mathlib/Algebra/Star/NonUnitalSubalgebra.lean | 72 +++++++++++++++++++ 2 files changed, 123 insertions(+) diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index af1d1bc38d9af..bea24677990a7 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -1131,6 +1131,57 @@ end Centralizer end NonUnitalSubalgebra +namespace NonUnitalAlgebra + +open NonUnitalSubalgebra + +variable {R A : Type*} [CommSemiring R] [NonUnitalSemiring A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + +variable (R) in +lemma adjoin_le_centralizer_centralizer (s : Set A) : + adjoin R s ≤ centralizer R (centralizer R s) := + adjoin_le Set.subset_centralizer_centralizer + +lemma commute_of_mem_adjoin_of_forall_mem_commute {a b : A} {s : Set A} + (hb : b ∈ adjoin R s) (h : ∀ b ∈ s, Commute a b) : + Commute a b := by + have : a ∈ centralizer R s := by simpa only [Commute.symm_iff (a := a)] using h + exact adjoin_le_centralizer_centralizer R s hb a this + +lemma commute_of_mem_adjoin_singleton_of_commute {a b c : A} + (hc : c ∈ adjoin R {b}) (h : Commute a b) : + Commute a c := + commute_of_mem_adjoin_of_forall_mem_commute hc <| by simpa + +lemma commute_of_mem_adjoin_self {a b : A} (hb : b ∈ adjoin R {a}) : + Commute a b := + commute_of_mem_adjoin_singleton_of_commute hb rfl + +variable (R) in + +/-- If all elements of `s : Set A` commute pairwise, then `adjoin R s` is a non-unital commutative +semiring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommSemiringOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : + NonUnitalCommSemiring (adjoin R s) := + { (adjoin R s).toNonUnitalSemiring with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ + have := adjoin_le_centralizer_centralizer R s + Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ (this h₁) _ (this h₂) } + +/-- If all elements of `s : Set A` commute pairwise, then `adjoin R s` is a non-unital commutative +ring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommRingOfComm (R : Type*) {A : Type*} [CommRing R] [NonUnitalRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] {s : Set A} + (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) : NonUnitalCommRing (adjoin R s) := + { (adjoin R s).toNonUnitalRing, adjoinNonUnitalCommSemiringOfComm R hcomm with } + +end NonUnitalAlgebra + section Nat variable {R : Type*} [NonUnitalNonAssocSemiring R] diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index 86f059f7cf0fe..0f0476488ad54 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -6,6 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Algebra.NonUnitalSubalgebra import Mathlib.Algebra.Star.StarAlgHom import Mathlib.Algebra.Star.Center +import Mathlib.Algebra.Star.SelfAdjoint /-! # Non-unital Star Subalgebras @@ -1130,6 +1131,77 @@ theorem centralizer_le (s t : Set A) (h : s ⊆ t) : centralizer R t ≤ central theorem centralizer_univ : centralizer R Set.univ = center R A := SetLike.ext' <| by rw [coe_centralizer, Set.univ_union, coe_center, Set.centralizer_univ] +theorem centralizer_toNonUnitalSubalgebra (s : Set A) : + (centralizer R s).toNonUnitalSubalgebra = NonUnitalSubalgebra.centralizer R (s ∪ star s):= + rfl + +theorem coe_centralizer_centralizer (s : Set A) : + (centralizer R (centralizer R s : Set A)) = (s ∪ star s).centralizer.centralizer := by + rw [coe_centralizer, StarMemClass.star_coe_eq, Set.union_self, coe_centralizer] + end Centralizer end NonUnitalStarSubalgebra + +namespace NonUnitalStarAlgebra + +open NonUnitalStarSubalgebra + +variable [CommSemiring R] [StarRing R] +variable [NonUnitalSemiring A] [StarRing A] [Module R A] +variable [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] + +variable (R) in +lemma adjoin_le_centralizer_centralizer (s : Set A) : + adjoin R s ≤ centralizer R (centralizer R s) := by + rw [← toNonUnitalSubalgebra_le_iff, centralizer_toNonUnitalSubalgebra, + adjoin_toNonUnitalSubalgebra] + convert NonUnitalAlgebra.adjoin_le_centralizer_centralizer R (s ∪ star s) + rw [StarMemClass.star_coe_eq] + simp + +lemma commute_of_mem_adjoin_of_forall_mem_commute {a b : A} {s : Set A} + (hb : b ∈ adjoin R s) (h : ∀ b ∈ s, Commute a b) (h_star : ∀ b ∈ s, Commute a (star b)) : + Commute a b := + NonUnitalAlgebra.commute_of_mem_adjoin_of_forall_mem_commute hb fun b hb ↦ + hb.elim (h b) (by simpa using h_star (star b)) + +lemma commute_of_mem_adjoin_singleton_of_commute {a b c : A} + (hc : c ∈ adjoin R {b}) (h : Commute a b) (h_star : Commute a (star b)) : + Commute a c := + commute_of_mem_adjoin_of_forall_mem_commute hc (by simpa) (by simpa) + +lemma commute_of_mem_adjoin_self {a b : A} [IsStarNormal a] (hb : b ∈ adjoin R {a}) : + Commute a b := + commute_of_mem_adjoin_singleton_of_commute hb rfl (isStarNormal_iff a |>.mp inferInstance).symm + +variable (R) in +/-- If all elements of `s : Set A` commute pairwise and with elements of `star s`, then `adjoin R s` +is a non-unital commutative semiring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommSemiringOfComm {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) + (hcomm_star : ∀ a ∈ s, ∀ b ∈ s, a * star b = star b * a) : + NonUnitalCommSemiring (adjoin R s) := + { (adjoin R s).toNonUnitalSemiring with + mul_comm := fun ⟨_, h₁⟩ ⟨_, h₂⟩ ↦ by + have hcomm : ∀ a ∈ s ∪ star s, ∀ b ∈ s ∪ star s, a * b = b * a := fun a ha b hb ↦ + Set.union_star_self_comm (fun _ ha _ hb ↦ hcomm _ hb _ ha) + (fun _ ha _ hb ↦ hcomm_star _ hb _ ha) b hb a ha + have := adjoin_le_centralizer_centralizer R s + apply this at h₁ + apply this at h₂ + rw [← SetLike.mem_coe, coe_centralizer_centralizer] at h₁ h₂ + exact Subtype.ext <| Set.centralizer_centralizer_comm_of_comm hcomm _ h₁ _ h₂ } + +/-- If all elements of `s : Set A` commute pairwise and with elements of `star s`, then `adjoin R s` +is a non-unital commutative ring. + +See note [reducible non-instances]. -/ +abbrev adjoinNonUnitalCommRingOfComm (R : Type*) {A : Type*} [CommRing R] [StarRing R] + [NonUnitalRing A] [StarRing A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + [StarModule R A] {s : Set A} (hcomm : ∀ a ∈ s, ∀ b ∈ s, a * b = b * a) + (hcomm_star : ∀ a ∈ s, ∀ b ∈ s, a * star b = star b * a) : NonUnitalCommRing (adjoin R s) := + { (adjoin R s).toNonUnitalRing, adjoinNonUnitalCommSemiringOfComm R hcomm hcomm_star with } + +end NonUnitalStarAlgebra From 3787191e6f5ffa5b1b02e297313336d0adeec8da Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Wed, 20 Nov 2024 06:37:48 +0000 Subject: [PATCH 43/90] Merge master into nightly-testing --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 14175c9c38864..0df8444539094 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "86d0d0584f5cd165353e2f8a30c455cd0e168ac2", + "rev": "d7caecce0d0f003fd5e9cce9a61f1dd6ba83142b", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From 0e836d6e1a3c5ed008688622e261e19fbef05e0e Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 06:53:43 +0000 Subject: [PATCH 44/90] feat: `maintainer merge?`/`maintainer delegate?` switch for spoiler (#19230) --- .github/workflows/maintainer_merge.yml | 3 ++- scripts/maintainer_merge_message.sh | 17 ++++++++++++----- 2 files changed, 14 insertions(+), 6 deletions(-) diff --git a/.github/workflows/maintainer_merge.yml b/.github/workflows/maintainer_merge.yml index 87ab7583093b7..8be16fcfc4ed6 100644 --- a/.github/workflows/maintainer_merge.yml +++ b/.github/workflows/maintainer_merge.yml @@ -42,7 +42,8 @@ jobs: printf 'Comment:"%s"\n' "${COMMENT}" m_or_d="$(printf '%s' "${COMMENT}" | - sed -n 's=^maintainer *\(merge\|delegate\) *$=\1=p' | head -1)" + # captures `maintainer merge/delegate` as well as an optional `?`, ignoring subsequent spaces + sed -n 's=^maintainer *\(merge\|delegate\)\(?\?\) *$=\1\2=p' | head -1)" printf $'"maintainer delegate" or "maintainer merge" found? \'%s\'\n' "${m_or_d}" diff --git a/scripts/maintainer_merge_message.sh b/scripts/maintainer_merge_message.sh index ec4e14808bf86..0a8290705dd42 100755 --- a/scripts/maintainer_merge_message.sh +++ b/scripts/maintainer_merge_message.sh @@ -36,10 +36,17 @@ esac >&2 echo "EVENT_NAME: '${EVENT_NAME}'" >&2 printf 'COMMENT\n%s\nEND COMMENT\n' "${PR_COMMENT}" -# replace backticks in the title with single quotes -unbacktickedTitle="${PR_TITLE//\`/\'}" +printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D/$'?'/}" "${SOURCE}" "${PR}" "${URL}" +# if `maintainer merge/delegate` is followed by `?`, then print a `spoiler` with the full comment +if [ ${M_or_D: -1} == $'?' ] +then + # replace backticks in the title with single quotes + unbacktickedTitle="${PR_TITLE//\`/\'}" ->&2 echo "neat title: '${unbacktickedTitle}'" + >&2 echo "neat title: '${unbacktickedTitle}'" -printf '%s requested a maintainer **%s** from %s on PR [#%s](%s):\n' "${AUTHOR}" "${M_or_D}" "${SOURCE}" "${PR}" "${URL}" -printf '```spoiler %s\n%s\n```\n' "${unbacktickedTitle}" "${PR_COMMENT}" + printf '```spoiler %s\n%s\n```\n' "${unbacktickedTitle}" "${PR_COMMENT}" +# otherwise, just print the title of the PR +else + printf '> %s\n' "${PR_TITLE}" +fi From 995c25cb9962967328d2ac33a0f46ebfe7a3d897 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 20 Nov 2024 08:15:52 +0000 Subject: [PATCH 45/90] feat(Data/Nat/Nth): `nth_mem_anti` (#19169) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add another lemma for working with `Nat.nth`: ```lean lemma nth_mem_anti {a b : ℕ} (hab : a ≤ b) (h : p (nth p b)) : p (nth p a) := by ``` Feel free to golf if you see a simpler way of proving this. --- Mathlib/Data/Nat/Nth.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index fe2658ba4d808..4cde210602f99 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -252,6 +252,18 @@ theorem le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a · rcases subset_range_nth ha with ⟨n, rfl⟩ rwa [nth_lt_nth hf, Nat.lt_succ_iff, ← nth_le_nth hf] at h +lemma nth_mem_anti {a b : ℕ} (hab : a ≤ b) (h : p (nth p b)) : p (nth p a) := by + by_cases h' : ∀ hf : (setOf p).Finite, a < #hf.toFinset + · exact nth_mem a h' + · simp only [not_forall, not_lt] at h' + have h'b : ∃ hf : (setOf p).Finite, #hf.toFinset ≤ b := by + rcases h' with ⟨hf, ha⟩ + exact ⟨hf, ha.trans hab⟩ + have ha0 : nth p a = 0 := by simp [nth_eq_zero, h'] + have hb0 : nth p b = 0 := by simp [nth_eq_zero, h'b] + rw [ha0] + rwa [hb0] at h + section Count variable (p) [DecidablePred p] From 23f454fb99cf9de12d57d1640ee2b4bc86b25413 Mon Sep 17 00:00:00 2001 From: D-Thomine <100795491+D-Thomine@users.noreply.github.com> Date: Wed, 20 Nov 2024 08:25:05 +0000 Subject: [PATCH 46/90] refactor(EReal): add add/sub lemmas and refactor limsup_add on EReal (#18879) This PR : - adds a few lemmas about addition/subtraction in `Topology.Instances.EReal`. These are translations in `EReal` of similar lemmas about multiplication/division in `ENNReal` (see `Data.ENNReal.Inv`). - renames [EReal.add_le_of_forall_add_le](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.add_le_of_forall_add_le), [EReal.le_add_of_forall_le_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.le_add_of_forall_le_add) to bring them in line with similar lemmas in `Real`/`ENNReal` (see e.g. [add_le_of_forall_lt](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Order/Group/DenselyOrdered.html#add_le_of_forall_lt)) - simplifies significantly the proofs of `add_le_of_forall_lt` and `le_add_of_forall_lt`. - renames [EReal.add_liminf_le_liminf_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.add_liminf_le_liminf_add), [EReal.limsup_add_le_add_limsup](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.limsup_add_le_add_limsup), [EReal.limsup_add_liminf_le_limsup_add](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.limsup_add_liminf_le_limsup_add), [EReal.liminf_add_le_limsup_add_liminf](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.liminf_add_le_limsup_add_liminf) to bring them in line with similar lemmas in `Real`/`ENNReal` (see e.g. [ENNReal.limsup_mul_le](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/Filter/ENNReal.html#ENNReal.limsup_mul_le)) - removes [EReal.le_of_forall_lt_iff_le](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.le_of_forall_lt_iff_le), [EReal.ge_of_forall_gt_iff_ge](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Real/EReal.html#EReal.ge_of_forall_gt_iff_ge), [EReal.limsup_le_iff](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Topology/Instances/EReal.html#EReal.limsup_le_iff) which were very specialized lemmas used only in the older version of the proofs above. --- Mathlib/Data/Real/EReal.lean | 176 +++++++++++------- .../TopologicalEntropy/CoverEntropy.lean | 2 +- .../Integral/VitaliCaratheodory.lean | 2 +- Mathlib/Topology/Instances/EReal.lean | 57 +++--- 4 files changed, 138 insertions(+), 99 deletions(-) diff --git a/Mathlib/Data/Real/EReal.lean b/Mathlib/Data/Real/EReal.lean index 4727e2cbbcc78..723fbe073737f 100644 --- a/Mathlib/Data/Real/EReal.lean +++ b/Mathlib/Data/Real/EReal.lean @@ -923,10 +923,16 @@ def negOrderIso : EReal ≃o ERealᵒᵈ := invFun := fun x => -OrderDual.ofDual x map_rel_iff' := neg_le_neg_iff } -theorem neg_lt_iff_neg_lt {a b : EReal} : -a < b ↔ -b < a := by +theorem neg_lt_comm {a b : EReal} : -a < b ↔ -b < a := by rw [← neg_lt_neg_iff, neg_neg] + +@[deprecated (since := "2024-11-19")] alias neg_lt_iff_neg_lt := neg_lt_comm + +theorem neg_lt_of_neg_lt {a b : EReal} (h : -a < b) : -b < a := neg_lt_comm.1 h + +theorem lt_neg_comm {a b : EReal} : a < -b ↔ b < -a := by rw [← neg_lt_neg_iff, neg_neg] -theorem neg_lt_of_neg_lt {a b : EReal} (h : -a < b) : -b < a := neg_lt_iff_neg_lt.1 h +theorem lt_neg_of_lt_neg {a b : EReal} (h : a < -b) : b < -a := lt_neg_comm.1 h lemma neg_add {x y : EReal} (h1 : x ≠ ⊥ ∨ y ≠ ⊤) (h2 : x ≠ ⊤ ∨ y ≠ ⊥) : - (x + y) = - x - y := by @@ -937,64 +943,6 @@ lemma neg_sub {x y : EReal} (h1 : x ≠ ⊥ ∨ y ≠ ⊥) (h2 : x ≠ ⊤ ∨ y - (x - y) = - x + y := by rw [sub_eq_add_neg, neg_add _ _, sub_eq_add_neg, neg_neg] <;> simp_all -/-! ### Addition and order -/ - -lemma le_of_forall_lt_iff_le {x y : EReal} : (∀ z : ℝ, x < z → y ≤ z) ↔ y ≤ x := by - refine ⟨fun h ↦ WithBot.le_of_forall_lt_iff_le.1 ?_, fun h _ x_z ↦ h.trans x_z.le⟩ - rw [WithTop.forall] - aesop - -lemma ge_of_forall_gt_iff_ge {x y : EReal} : (∀ z : ℝ, z < y → z ≤ x) ↔ y ≤ x := by - refine ⟨fun h ↦ WithBot.ge_of_forall_gt_iff_ge.1 ?_, fun h _ x_z ↦ x_z.le.trans h⟩ - rw [WithTop.forall] - aesop - -/-- This lemma is superseded by `add_le_of_forall_add_le`. -/ -private lemma top_add_le_of_forall_add_le {a b : EReal} (h : ∀ c < ⊤, ∀ d < a, c + d ≤ b) : - ⊤ + a ≤ b := by - induction a with - | h_bot => exact add_bot ⊤ ▸ bot_le - | h_real a => - refine top_add_coe a ▸ le_of_forall_lt_iff_le.1 fun c b_c ↦ ?_ - specialize h (c - a + 1) (coe_lt_top (c - a + 1)) (a - 1) - rw [← coe_one, ← coe_sub, ← coe_sub, ← coe_add, ← coe_add, add_add_sub_cancel, sub_add_cancel, - EReal.coe_lt_coe_iff] at h - exact (not_le_of_lt b_c (h (sub_one_lt a))).rec - | h_top => - refine top_add_top ▸ le_of_forall_lt_iff_le.1 fun c b_c ↦ ?_ - specialize h c (coe_lt_top c) 0 zero_lt_top - rw [add_zero] at h - exact (not_le_of_lt b_c h).rec - -lemma add_le_of_forall_add_le {a b c : EReal} (h : ∀ d < a, ∀ e < b, d + e ≤ c) : a + b ≤ c := by - induction a with - | h_bot => exact bot_add b ▸ bot_le - | h_real a => induction b with - | h_bot => exact add_bot (a : EReal) ▸ bot_le - | h_real b => - refine (@ge_of_forall_gt_iff_ge c (a+b)).1 fun d d_ab ↦ ?_ - rw [← coe_add, EReal.coe_lt_coe_iff] at d_ab - rcases exists_between d_ab with ⟨e, e_d, e_ab⟩ - have key₁ : (a + d - e : ℝ) < (a : EReal) := by apply EReal.coe_lt_coe_iff.2; linarith - have key₂ : (e - a : ℝ) < (b : EReal) := by apply EReal.coe_lt_coe_iff.2; linarith - apply le_of_eq_of_le _ (h (a + d - e) key₁ (e - a) key₂) - rw [← coe_add, ← coe_sub, ← coe_sub, ← coe_add, sub_add_sub_cancel, add_sub_cancel_left] - | h_top => - rw [add_comm (a : EReal) ⊤] - exact top_add_le_of_forall_add_le fun d d_top e e_a ↦ (add_comm d e ▸ h e e_a d d_top) - | h_top => exact top_add_le_of_forall_add_le h - -lemma le_add_of_forall_le_add {a b c : EReal} (h₁ : a ≠ ⊥ ∨ b ≠ ⊤) (h₂ : a ≠ ⊤ ∨ b ≠ ⊥) - (h : ∀ d > a, ∀ e > b, c ≤ d + e) : - c ≤ a + b := by - rw [← neg_le_neg_iff, neg_add h₁ h₂] - refine add_le_of_forall_add_le fun d d_a e e_b ↦ ?_ - have h₃ : d ≠ ⊥ ∨ e ≠ ⊤ := Or.inr (ne_top_of_lt e_b) - have h₄ : d ≠ ⊤ ∨ e ≠ ⊥ := Or.inl (ne_top_of_lt d_a) - rw [← neg_neg d, neg_lt_iff_neg_lt, neg_neg a] at d_a - rw [← neg_neg e, neg_lt_iff_neg_lt, neg_neg b] at e_b - exact le_neg_of_le_neg <| neg_add h₃ h₄ ▸ h (- d) d_a (- e) e_b - /-! ### Subtraction @@ -1023,6 +971,12 @@ theorem top_sub_coe (x : ℝ) : (⊤ : EReal) - x = ⊤ := theorem coe_sub_bot (x : ℝ) : (x : EReal) - ⊥ = ⊤ := rfl +lemma sub_bot {a : EReal} (h : a ≠ ⊥) : a - ⊥ = ⊤ := by + induction a + · simp only [ne_eq, not_true_eq_false] at h + · rw [coe_sub_bot] + · rw [top_sub_bot] + theorem sub_le_sub {x y z t : EReal} (h : x ≤ y) (h' : t ≤ z) : x - z ≤ y - t := add_le_add h (neg_le_neg_iff.2 h') @@ -1047,11 +1001,105 @@ theorem toReal_sub {x y : EReal} (hx : x ≠ ⊤) (h'x : x ≠ ⊥) (hy : y ≠ lift y to ℝ using ⟨hy, h'y⟩ rfl -lemma add_sub_cancel_right {a : EReal} {b : Real} : a + b - b = a := by +lemma sub_add_cancel_left {a : EReal} {b : Real} : a - b + b = a := by induction a - · rw [bot_add b, bot_sub b] + · rw [bot_sub b, bot_add b] · norm_cast; linarith - · rw [top_add_of_ne_bot (coe_ne_bot b), top_sub_coe] + · rw [top_sub_coe b, top_add_coe b] + +lemma add_sub_cancel_right {a : EReal} {b : Real} : a + b - b = a := by + rw [sub_eq_add_neg, add_assoc, add_comm (b : EReal), ← add_assoc, ← sub_eq_add_neg] + exact sub_add_cancel_left + +lemma le_sub_iff_add_le {a b c : EReal} (hb : b ≠ ⊥ ∨ c ≠ ⊥) (ht : b ≠ ⊤ ∨ c ≠ ⊤) : + a ≤ c - b ↔ a + b ≤ c := by + induction b with + | h_bot => + simp only [ne_eq, not_true_eq_false, false_or] at hb + simp only [sub_bot hb, le_top, add_bot, bot_le] + | h_real b => + rw [← (addLECancellable_coe b).add_le_add_iff_right, sub_add_cancel_left] + | h_top => + simp only [ne_eq, not_true_eq_false, false_or, sub_top, le_bot_iff] at ht ⊢ + refine ⟨fun h ↦ h ▸ (bot_add ⊤).symm ▸ bot_le, fun h ↦ ?_⟩ + by_contra ha + exact (h.trans_lt (Ne.lt_top ht)).ne (add_top_iff_ne_bot.2 ha) + +lemma sub_le_iff_le_add {a b c : EReal} (h₁ : b ≠ ⊥ ∨ c ≠ ⊤) (h₂ : b ≠ ⊤ ∨ c ≠ ⊥) : + a - b ≤ c ↔ a ≤ c + b := by + suffices a + (-b) ≤ c ↔ a ≤ c - (-b) by simpa [sub_eq_add_neg] + refine (le_sub_iff_add_le ?_ ?_).symm <;> simpa + +protected theorem lt_sub_iff_add_lt {a b c : EReal} (h₁ : b ≠ ⊥ ∨ c ≠ ⊤) (h₂ : b ≠ ⊤ ∨ c ≠ ⊥) : + c < a - b ↔ c + b < a := + lt_iff_lt_of_le_iff_le (sub_le_iff_le_add h₁ h₂) + +theorem sub_le_of_le_add {a b c : EReal} (h : a ≤ b + c) : a - c ≤ b := by + induction c with + | h_bot => rw [add_bot, le_bot_iff] at h; simp only [h, bot_sub, bot_le] + | h_real c => exact (sub_le_iff_le_add (.inl (coe_ne_bot c)) (.inl (coe_ne_top c))).2 h + | h_top => simp only [sub_top, bot_le] + +/-- See also `EReal.sub_le_of_le_add`.-/ +theorem sub_le_of_le_add' {a b c : EReal} (h : a ≤ b + c) : a - b ≤ c := + sub_le_of_le_add (add_comm b c ▸ h) + +lemma add_le_of_le_sub {a b c : EReal} (h : a ≤ b - c) : a + c ≤ b := by + rw [← neg_neg c] + exact sub_le_of_le_add h + +lemma sub_lt_iff {a b c : EReal} (h₁ : b ≠ ⊥ ∨ c ≠ ⊥) (h₂ : b ≠ ⊤ ∨ c ≠ ⊤) : + c - b < a ↔ c < a + b := + lt_iff_lt_of_le_iff_le (le_sub_iff_add_le h₁ h₂) + +lemma add_lt_of_lt_sub {a b c : EReal} (h : a < b - c) : a + c < b := by + contrapose! h + exact sub_le_of_le_add h + +lemma sub_lt_of_lt_add {a b c : EReal} (h : a < b + c) : a - c < b := + add_lt_of_lt_sub <| by rwa [sub_eq_add_neg, neg_neg] + +/-- See also `EReal.sub_lt_of_lt_add`.-/ +lemma sub_lt_of_lt_add' {a b c : EReal} (h : a < b + c) : a - b < c := + sub_lt_of_lt_add <| by rwa [add_comm] + +/-! ### Addition and order -/ + +lemma le_of_forall_lt_iff_le {x y : EReal} : (∀ z : ℝ, x < z → y ≤ z) ↔ y ≤ x := by + refine ⟨fun h ↦ WithBot.le_of_forall_lt_iff_le.1 ?_, fun h _ x_z ↦ h.trans x_z.le⟩ + rw [WithTop.forall] + aesop + +lemma ge_of_forall_gt_iff_ge {x y : EReal} : (∀ z : ℝ, z < y → z ≤ x) ↔ y ≤ x := by + refine ⟨fun h ↦ WithBot.ge_of_forall_gt_iff_ge.1 ?_, fun h _ x_z ↦ x_z.le.trans h⟩ + rw [WithTop.forall] + aesop + +private lemma exists_lt_add_left {a b c : EReal} (hc : c < a + b) : ∃ a' < a, c < a' + b := by + obtain ⟨a', hc', ha'⟩ := exists_between (sub_lt_of_lt_add hc) + refine ⟨a', ha', (sub_lt_iff (.inl ?_) (.inr hc.ne_top)).1 hc'⟩ + contrapose! hc + exact hc ▸ (add_bot a).symm ▸ bot_le + +private lemma exists_lt_add_right {a b c : EReal} (hc : c < a + b) : ∃ b' < b, c < a + b' := by + simp_rw [add_comm a] at hc ⊢; exact exists_lt_add_left hc + +lemma add_le_of_forall_lt {a b c : EReal} (h : ∀ a' < a, ∀ b' < b, a' + b' ≤ c) : a + b ≤ c := by + refine le_of_forall_ge_of_dense fun d hd ↦ ?_ + obtain ⟨a', ha', hd⟩ := exists_lt_add_left hd + obtain ⟨b', hb', hd⟩ := exists_lt_add_right hd + exact hd.le.trans (h _ ha' _ hb') + +lemma le_add_of_forall_gt {a b c : EReal} (h₁ : a ≠ ⊥ ∨ b ≠ ⊤) (h₂ : a ≠ ⊤ ∨ b ≠ ⊥) + (h : ∀ a' > a, ∀ b' > b, c ≤ a' + b') : c ≤ a + b := by + rw [← neg_le_neg_iff, neg_add h₁ h₂] + exact add_le_of_forall_lt fun a' ha' b' hb' ↦ le_neg_of_le_neg + <| (h (-a') (lt_neg_of_lt_neg ha') (-b') (lt_neg_of_lt_neg hb')).trans_eq + (neg_add (.inr hb'.ne_top) (.inl ha'.ne_top)).symm + +@[deprecated (since := "2024-11-19")] alias top_add_le_of_forall_add_le := add_le_of_forall_lt +@[deprecated (since := "2024-11-19")] alias add_le_of_forall_add_le := add_le_of_forall_lt +@[deprecated (since := "2024-11-19")] alias le_add_of_forall_le_add := le_add_of_forall_gt lemma _root_.ENNReal.toEReal_sub {x y : ℝ≥0∞} (hy_top : y ≠ ∞) (h_le : y ≤ x) : (x - y).toEReal = x.toEReal - y.toEReal := by @@ -1576,7 +1624,7 @@ lemma div_right_distrib_of_nonneg {a b c : EReal} (h : 0 ≤ a) (h' : 0 ≤ b) : (a + b) / c = (a / c) + (b / c) := EReal.right_distrib_of_nonneg h h' -/-! #### Division and Order s -/ +/-! #### Division and Order -/ lemma monotone_div_right_of_nonneg {b : EReal} (h : 0 ≤ b) : Monotone fun a ↦ a / b := fun _ _ h' ↦ mul_le_mul_of_nonneg_right h' (inv_nonneg_of_nonneg h) diff --git a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean index b5c4368d3eecd..f12dd6c32afc6 100644 --- a/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean +++ b/Mathlib/Dynamics/TopologicalEntropy/CoverEntropy.lean @@ -483,7 +483,7 @@ lemma coverEntropyEntourage_le_log_coverMincard_div {T : X → X} {F : Set X} (F eventually_atTop.2 ⟨1, fun m m_pos ↦ log_coverMincard_le_add F_inv U_symm n_pos m_pos⟩ apply ((limsup_le_limsup) key).trans suffices h : atTop.limsup v = 0 by - have := @limsup_add_le_add_limsup ℕ atTop u v + have := @limsup_add_le ℕ atTop u v rw [h, add_zero] at this specialize this (Or.inr EReal.zero_ne_top) (Or.inr EReal.zero_ne_bot) exact this.trans_eq (limsup_const (log (coverMincard T F U n) / n)) diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 2aa0a2e287af9..2365cd675affa 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -527,7 +527,7 @@ theorem exists_upperSemicontinuous_lt_integral_gt [SigmaFinite μ] (f : α → rcases exists_lt_lowerSemicontinuous_integral_lt (fun x => -f x) hf.neg εpos with ⟨g, g_lt_f, gcont, g_integrable, g_lt_top, gint⟩ refine ⟨fun x => -g x, ?_, ?_, ?_, ?_, ?_⟩ - · exact fun x => EReal.neg_lt_iff_neg_lt.1 (by simpa only [EReal.coe_neg] using g_lt_f x) + · exact fun x => EReal.neg_lt_comm.1 (by simpa only [EReal.coe_neg] using g_lt_f x) · exact continuous_neg.comp_lowerSemicontinuous_antitone gcont fun x y hxy => EReal.neg_le_neg_iff.2 hxy diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index cedc98e15dfa0..746c9fd365e5b 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -214,65 +214,56 @@ lemma liminf_neg : liminf (- v) f = - limsup v f := lemma limsup_neg : limsup (- v) f = - liminf v f := EReal.negOrderIso.liminf_apply.symm -lemma add_liminf_le_liminf_add : (liminf u f) + (liminf v f) ≤ liminf (u + v) f := by - refine add_le_of_forall_add_le fun a a_u b b_v ↦ (le_liminf_iff).2 fun c c_ab ↦ ?_ +lemma le_liminf_add : (liminf u f) + (liminf v f) ≤ liminf (u + v) f := by + refine add_le_of_forall_lt fun a a_u b b_v ↦ (le_liminf_iff).2 fun c c_ab ↦ ?_ filter_upwards [eventually_lt_of_lt_liminf a_u, eventually_lt_of_lt_liminf b_v] with x a_x b_x - exact lt_trans c_ab (add_lt_add a_x b_x) + exact c_ab.trans (add_lt_add a_x b_x) -lemma limsup_add_le_add_limsup (h : limsup u f ≠ ⊥ ∨ limsup v f ≠ ⊤) - (h' : limsup u f ≠ ⊤ ∨ limsup v f ≠ ⊥) : +lemma limsup_add_le (h : limsup u f ≠ ⊥ ∨ limsup v f ≠ ⊤) (h' : limsup u f ≠ ⊤ ∨ limsup v f ≠ ⊥) : limsup (u + v) f ≤ (limsup u f) + (limsup v f) := by - refine le_add_of_forall_le_add h h' fun a a_u b b_v ↦ (limsup_le_iff).2 fun c c_ab ↦ ?_ + refine le_add_of_forall_gt h h' fun a a_u b b_v ↦ (limsup_le_iff).2 fun c c_ab ↦ ?_ filter_upwards [eventually_lt_of_limsup_lt a_u, eventually_lt_of_limsup_lt b_v] with x a_x b_x exact (add_lt_add a_x b_x).trans c_ab -lemma limsup_add_liminf_le_limsup_add : (limsup u f) + (liminf v f) ≤ limsup (u + v) f := - add_le_of_forall_add_le fun _ a_u _ b_v ↦ (le_limsup_iff).2 fun _ c_ab ↦ - Frequently.mono (Frequently.and_eventually ((frequently_lt_of_lt_limsup) a_u) - ((eventually_lt_of_lt_liminf) b_v)) fun _ ab_x ↦ c_ab.trans (add_lt_add ab_x.1 ab_x.2) +lemma le_limsup_add : (limsup u f) + (liminf v f) ≤ limsup (u + v) f := + add_le_of_forall_lt fun _ a_u _ b_v ↦ (le_limsup_iff).2 fun _ c_ab ↦ + (((frequently_lt_of_lt_limsup) a_u).and_eventually ((eventually_lt_of_lt_liminf) b_v)).mono + fun _ ab_x ↦ c_ab.trans (add_lt_add ab_x.1 ab_x.2) -lemma liminf_add_le_limsup_add_liminf (h : limsup u f ≠ ⊥ ∨ liminf v f ≠ ⊤) - (h' : limsup u f ≠ ⊤ ∨ liminf v f ≠ ⊥) : +lemma liminf_add_le (h : limsup u f ≠ ⊥ ∨ liminf v f ≠ ⊤) (h' : limsup u f ≠ ⊤ ∨ liminf v f ≠ ⊥) : liminf (u + v) f ≤ (limsup u f) + (liminf v f) := - le_add_of_forall_le_add h h' fun _ a_u _ b_v ↦ (liminf_le_iff).2 fun _ c_ab ↦ - Frequently.mono (Frequently.and_eventually ((frequently_lt_of_liminf_lt) b_v) - ((eventually_lt_of_limsup_lt) a_u)) fun _ ab_x ↦ (add_lt_add ab_x.2 ab_x.1).trans c_ab + le_add_of_forall_gt h h' fun _ a_u _ b_v ↦ (liminf_le_iff).2 fun _ c_ab ↦ + (((frequently_lt_of_liminf_lt) b_v).and_eventually ((eventually_lt_of_limsup_lt) a_u)).mono + fun _ ab_x ↦ (add_lt_add ab_x.2 ab_x.1).trans c_ab + +@[deprecated (since := "2024-11-11")] alias add_liminf_le_liminf_add := le_liminf_add +@[deprecated (since := "2024-11-11")] alias limsup_add_le_add_limsup := limsup_add_le +@[deprecated (since := "2024-11-11")] alias limsup_add_liminf_le_limsup_add := le_limsup_add +@[deprecated (since := "2024-11-11")] alias liminf_add_le_limsup_add_liminf := liminf_add_le variable {a b : EReal} lemma limsup_add_bot_of_ne_top (h : limsup u f = ⊥) (h' : limsup v f ≠ ⊤) : limsup (u + v) f = ⊥ := by - apply le_bot_iff.1 (le_trans (limsup_add_le_add_limsup (Or.inr h') _) _) - · rw [h]; exact Or.inl bot_ne_top + apply le_bot_iff.1 ((limsup_add_le (.inr h') _).trans _) + · rw [h]; exact .inl bot_ne_top · rw [h, bot_add] lemma limsup_add_le_of_le (ha : limsup u f < a) (hb : limsup v f ≤ b) : limsup (u + v) f ≤ a + b := by - rcases eq_top_or_lt_top b with (rfl | h) + rcases eq_top_or_lt_top b with rfl | h · rw [add_top_of_ne_bot ha.ne_bot]; exact le_top - · exact le_trans (limsup_add_le_add_limsup (Or.inr (lt_of_le_of_lt hb h).ne) (Or.inl ha.ne_top)) - (add_le_add ha.le hb) + · exact (limsup_add_le (.inr (hb.trans_lt h).ne) (.inl ha.ne_top)).trans (add_le_add ha.le hb) lemma liminf_add_gt_of_gt (ha : a < liminf u f) (hb : b < liminf v f) : a + b < liminf (u + v) f := - lt_of_lt_of_le (add_lt_add ha hb) add_liminf_le_liminf_add + (add_lt_add ha hb).trans_le le_liminf_add lemma liminf_add_top_of_ne_bot (h : liminf u f = ⊤) (h' : liminf v f ≠ ⊥) : liminf (u + v) f = ⊤ := by - apply top_le_iff.1 (le_trans _ (add_liminf_le_liminf_add)) + apply top_le_iff.1 (le_trans _ le_liminf_add) rw [h, top_add_of_ne_bot h'] -lemma limsup_le_iff {b : EReal} : limsup u f ≤ b ↔ ∀ c : ℝ, b < c → ∀ᶠ a : α in f, u a ≤ c := by - rw [← le_of_forall_lt_iff_le] - refine ⟨?_, ?_⟩ <;> intro h c b_c - · rcases exists_between_coe_real b_c with ⟨d, b_d, d_c⟩ - apply mem_of_superset (eventually_lt_of_limsup_lt (lt_of_le_of_lt (h d b_d) d_c)) - rw [Set.setOf_subset_setOf] - exact fun _ h' ↦ h'.le - · rcases eq_or_neBot f with rfl | _ - · simp only [limsup_bot, bot_le] - · exact (limsup_le_of_le) (h c b_c) - end LimInfSup /-! ### Continuity of addition -/ From 55c020e1b67b5ae4dfe3efcc1f71a0346ef31a9e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 20 Nov 2024 08:25:07 +0000 Subject: [PATCH 47/90] feat(Order/WithBot): Equiv.withBotSubtypeNe (#19251) Needed in #19210 From the Carleson project --- Mathlib/Order/WithBot.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index d49c2f822a4fb..5be2bf90e1157 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -158,6 +158,13 @@ theorem eq_unbot_iff {a : α} {b : WithBot α} (h : b ≠ ⊥) : · simpa using h rfl · simp +/-- The equivalence between the non-bottom elements of `WithBot α` and `α`. -/ +@[simps] def _root_.Equiv.withBotSubtypeNe : {y : WithBot α // y ≠ ⊥} ≃ α where + toFun := fun ⟨x,h⟩ => WithBot.unbot x h + invFun x := ⟨x, WithBot.coe_ne_bot⟩ + left_inv _ := by simp + right_inv _ := by simp + section LE variable [LE α] @@ -729,6 +736,13 @@ theorem eq_untop_iff {a : α} {b : WithTop α} (h : b ≠ ⊤) : a = b.untop h ↔ a = b := WithBot.eq_unbot_iff (α := αᵒᵈ) h +/-- The equivalence between the non-top elements of `WithTop α` and `α`. -/ +@[simps] def _root_.Equiv.withTopSubtypeNe : {y : WithTop α // y ≠ ⊤} ≃ α where + toFun := fun ⟨x,h⟩ => WithTop.untop x h + invFun x := ⟨x, WithTop.coe_ne_top⟩ + left_inv _ := by simp + right_inv _:= by simp + section LE variable [LE α] From 650a3f818f1cd3ac8de8856c884065a6d1f6749e Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Wed, 20 Nov 2024 09:05:39 +0000 Subject: [PATCH 48/90] chore: no `open Foo` after `namespace Foo` (#19223) ...and not opening the same name space twice. Co-authored-by: Moritz Firsching --- Mathlib/Algebra/CubicDiscriminant.lean | 2 -- Mathlib/Algebra/MvPolynomial/Cardinal.lean | 2 -- Mathlib/Algebra/Polynomial/Identities.lean | 2 -- Mathlib/Algebra/Polynomial/Mirror.lean | 2 -- Mathlib/Algebra/Polynomial/Monomial.lean | 2 -- Mathlib/Algebra/Polynomial/Reverse.lean | 4 ++-- Mathlib/Algebra/Polynomial/Splits.lean | 2 -- Mathlib/Algebra/Polynomial/Taylor.lean | 2 -- Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean | 2 -- Mathlib/Analysis/RCLike/Lemmas.lean | 2 -- .../Analysis/SpecialFunctions/Trigonometric/Basic.lean | 2 -- Mathlib/CategoryTheory/Closed/Cartesian.lean | 2 +- Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean | 2 +- Mathlib/CategoryTheory/Sites/Grothendieck.lean | 2 +- Mathlib/CategoryTheory/Sites/PreservesSheafification.lean | 2 +- Mathlib/CategoryTheory/Sites/Pretopology.lean | 2 +- Mathlib/Combinatorics/Enumerative/Catalan.lean | 2 -- Mathlib/Data/Matrix/Basis.lean | 2 -- Mathlib/Data/Matrix/Hadamard.lean | 2 -- Mathlib/Data/Matrix/Kronecker.lean | 2 -- Mathlib/Data/Nat/Choose/Dvd.lean | 2 -- Mathlib/Data/Nat/MaxPowDiv.lean | 2 -- Mathlib/Data/Nat/Periodic.lean | 2 +- Mathlib/Data/Rat/Lemmas.lean | 2 -- Mathlib/Data/Real/Irrational.lean | 2 -- Mathlib/FieldTheory/IntermediateField/Basic.lean | 2 -- Mathlib/LinearAlgebra/CrossProduct.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Symmetric.lean | 2 -- Mathlib/LinearAlgebra/Matrix/Transvection.lean | 2 -- Mathlib/LinearAlgebra/QuadraticForm/Prod.lean | 2 -- .../MeasureTheory/Function/StronglyMeasurable/Basic.lean | 2 -- .../NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean | 2 -- Mathlib/NumberTheory/Primorial.lean | 2 -- Mathlib/RingTheory/Adjoin/PowerBasis.lean | 2 -- Mathlib/RingTheory/EisensteinCriterion.lean | 2 -- Mathlib/RingTheory/IntegralDomain.lean | 2 -- Mathlib/RingTheory/Jacobson.lean | 2 -- Mathlib/RingTheory/MatrixAlgebra.lean | 8 +------- Mathlib/RingTheory/Nullstellensatz.lean | 2 -- Mathlib/RingTheory/Polynomial/Content.lean | 2 -- Mathlib/RingTheory/Polynomial/Dickson.lean | 2 -- Mathlib/RingTheory/Polynomial/Pochhammer.lean | 2 -- Mathlib/RingTheory/Polynomial/ScaleRoots.lean | 2 -- Mathlib/RingTheory/ReesAlgebra.lean | 2 -- Mathlib/SetTheory/Cardinal/Divisibility.lean | 2 -- Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean | 2 -- Mathlib/Topology/Algebra/Polynomial.lean | 2 -- MathlibTest/matrix.lean | 2 -- MathlibTest/norm_cast.lean | 2 -- 51 files changed, 9 insertions(+), 101 deletions(-) diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 42e56fed93eff..4104d2add1285 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -40,8 +40,6 @@ structure Cubic (R : Type*) where namespace Cubic -open Cubic Polynomial - open Polynomial variable {R S F K : Type*} diff --git a/Mathlib/Algebra/MvPolynomial/Cardinal.lean b/Mathlib/Algebra/MvPolynomial/Cardinal.lean index 880c15e0cedd3..3ee977de57b85 100644 --- a/Mathlib/Algebra/MvPolynomial/Cardinal.lean +++ b/Mathlib/Algebra/MvPolynomial/Cardinal.lean @@ -20,8 +20,6 @@ universe u v open Cardinal -open Cardinal - namespace MvPolynomial section TwoUniverses diff --git a/Mathlib/Algebra/Polynomial/Identities.lean b/Mathlib/Algebra/Polynomial/Identities.lean index d33dfaf9c2923..3aa9df0fe8d2a 100644 --- a/Mathlib/Algebra/Polynomial/Identities.lean +++ b/Mathlib/Algebra/Polynomial/Identities.lean @@ -18,8 +18,6 @@ noncomputable section namespace Polynomial -open Polynomial - universe u v w x y z variable {R : Type u} {S : Type v} {T : Type w} {ι : Type x} {k : Type y} {A : Type z} {a b : R} diff --git a/Mathlib/Algebra/Polynomial/Mirror.lean b/Mathlib/Algebra/Polynomial/Mirror.lean index 9f3b754cd0a0a..804488801b7dc 100644 --- a/Mathlib/Algebra/Polynomial/Mirror.lean +++ b/Mathlib/Algebra/Polynomial/Mirror.lean @@ -27,8 +27,6 @@ divisible by `X`. namespace Polynomial -open Polynomial - section Semiring variable {R : Type*} [Semiring R] (p q : R[X]) diff --git a/Mathlib/Algebra/Polynomial/Monomial.lean b/Mathlib/Algebra/Polynomial/Monomial.lean index 535103330d899..f994bf2057858 100644 --- a/Mathlib/Algebra/Polynomial/Monomial.lean +++ b/Mathlib/Algebra/Polynomial/Monomial.lean @@ -16,8 +16,6 @@ noncomputable section namespace Polynomial -open Polynomial - universe u variable {R : Type u} {a b : R} {m n : ℕ} diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 07d080133f5ee..a5ba40706e958 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -19,9 +19,9 @@ coefficients of `f` and `g` do not multiply to zero. namespace Polynomial -open Polynomial Finsupp Finset +open Finsupp Finset -open Polynomial +open scoped Polynomial section Semiring diff --git a/Mathlib/Algebra/Polynomial/Splits.lean b/Mathlib/Algebra/Polynomial/Splits.lean index 918c15651b4a4..211e8294de371 100644 --- a/Mathlib/Algebra/Polynomial/Splits.lean +++ b/Mathlib/Algebra/Polynomial/Splits.lean @@ -33,8 +33,6 @@ variable {R : Type*} {F : Type u} {K : Type v} {L : Type w} namespace Polynomial -open Polynomial - section Splits section CommRing diff --git a/Mathlib/Algebra/Polynomial/Taylor.lean b/Mathlib/Algebra/Polynomial/Taylor.lean index 805deafc3c1c7..f3f2c2b9b9fd7 100644 --- a/Mathlib/Algebra/Polynomial/Taylor.lean +++ b/Mathlib/Algebra/Polynomial/Taylor.lean @@ -26,8 +26,6 @@ noncomputable section namespace Polynomial -open Polynomial - variable {R : Type*} [Semiring R] (r : R) (f : R[X]) /-- The Taylor expansion of a polynomial `f` at `r`. -/ diff --git a/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean b/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean index deb7cba16f499..c112b860d9378 100644 --- a/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean +++ b/Mathlib/Analysis/Normed/Operator/BanachSteinhaus.lean @@ -37,8 +37,6 @@ theorem banach_steinhaus {ι : Type*} [CompleteSpace E] {g : ι → E →SL[σ open ENNReal -open ENNReal - /-- This version of Banach-Steinhaus is stated in terms of suprema of `↑‖·‖₊ : ℝ≥0∞` for convenience. -/ theorem banach_steinhaus_iSup_nnnorm {ι : Type*} [CompleteSpace E] {g : ι → E →SL[σ₁₂] F} diff --git a/Mathlib/Analysis/RCLike/Lemmas.lean b/Mathlib/Analysis/RCLike/Lemmas.lean index 0603a4fda4c16..9ce1f0f40f0a5 100644 --- a/Mathlib/Analysis/RCLike/Lemmas.lean +++ b/Mathlib/Analysis/RCLike/Lemmas.lean @@ -13,8 +13,6 @@ variable {K E : Type*} [RCLike K] namespace Polynomial -open Polynomial - theorem ofReal_eval (p : ℝ[X]) (x : ℝ) : (↑(p.eval x) : K) = aeval (↑x) p := (@aeval_algebraMap_apply_eq_algebraMap_eval ℝ K _ _ _ x p).symm diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 68908a896942d..9dd56c0fcd556 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -201,8 +201,6 @@ end NNReal namespace Real -open Real - @[simp] theorem sin_pi : sin π = 0 := by rw [← mul_div_cancel_left₀ π (two_ne_zero' ℝ), two_mul, add_div, sin_add, cos_pi_div_two]; simp diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index 6b72829dd536d..0607788f02110 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -44,7 +44,7 @@ noncomputable section namespace CategoryTheory -open CategoryTheory Category Limits MonoidalCategory +open Category Limits MonoidalCategory /-- An object `X` is *exponentiable* if `(X × -)` is a left adjoint. We define this as being `Closed` in the cartesian monoidal structure. diff --git a/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean b/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean index 5319308872660..1f67ad02341c3 100644 --- a/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean +++ b/Mathlib/CategoryTheory/FiberedCategory/BasedCategory.lean @@ -29,7 +29,7 @@ universe v₅ u₅ v₄ u₄ v₃ u₃ v₂ u₂ v₁ u₁ namespace CategoryTheory -open CategoryTheory Functor Category NatTrans IsHomLift +open Functor Category NatTrans IsHomLift variable {𝒮 : Type u₁} [Category.{v₁} 𝒮] diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 28c48a20268de..5570569b02e2b 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -53,7 +53,7 @@ universe v₁ u₁ v u namespace CategoryTheory -open CategoryTheory Category +open Category variable (C : Type u) [Category.{v} C] diff --git a/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean b/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean index 52f8413c10939..2c7a0002cec28 100644 --- a/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean +++ b/Mathlib/CategoryTheory/Sites/PreservesSheafification.lean @@ -43,7 +43,7 @@ universe v u namespace CategoryTheory -open CategoryTheory Category Limits +open Category Limits variable {C : Type u} [Category.{v} C] (J : GrothendieckTopology C) {A B : Type*} [Category A] [Category B] (F : A ⥤ B) diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index f2f7dfc06495c..d888a2611acf2 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -35,7 +35,7 @@ noncomputable section namespace CategoryTheory -open CategoryTheory Category Limits Presieve +open Category Limits Presieve variable {C : Type u} [Category.{v} C] [HasPullbacks C] variable (C) diff --git a/Mathlib/Combinatorics/Enumerative/Catalan.lean b/Mathlib/Combinatorics/Enumerative/Catalan.lean index c2ae512bb889f..21a578fab89bf 100644 --- a/Mathlib/Combinatorics/Enumerative/Catalan.lean +++ b/Mathlib/Combinatorics/Enumerative/Catalan.lean @@ -140,8 +140,6 @@ theorem catalan_three : catalan 3 = 5 := by namespace Tree -open Tree - /-- Given two finsets, find all trees that can be formed with left child in `a` and right child in `b` -/ abbrev pairwiseNode (a b : Finset (Tree Unit)) : Finset (Tree Unit) := diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index f9d9157fe8645..7c49644844bcd 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -19,8 +19,6 @@ variable {R α : Type*} namespace Matrix -open Matrix - variable [DecidableEq l] [DecidableEq m] [DecidableEq n] section Zero diff --git a/Mathlib/Data/Matrix/Hadamard.lean b/Mathlib/Data/Matrix/Hadamard.lean index d94a90b94b736..cdaa2360b95c5 100644 --- a/Mathlib/Data/Matrix/Hadamard.lean +++ b/Mathlib/Data/Matrix/Hadamard.lean @@ -34,8 +34,6 @@ variable {α m n R : Type*} namespace Matrix -open Matrix - /-- `Matrix.hadamard` defines the Hadamard product, which is the pointwise product of two matrices of the same size. -/ def hadamard [Mul α] (A : Matrix m n α) (B : Matrix m n α) : Matrix m n α := diff --git a/Mathlib/Data/Matrix/Kronecker.lean b/Mathlib/Data/Matrix/Kronecker.lean index bbbe63c6841d9..1d108492e89cf 100644 --- a/Mathlib/Data/Matrix/Kronecker.lean +++ b/Mathlib/Data/Matrix/Kronecker.lean @@ -43,8 +43,6 @@ These require `open Kronecker`: namespace Matrix - -open Matrix open scoped RightActions variable {R α α' β β' γ γ' : Type*} diff --git a/Mathlib/Data/Nat/Choose/Dvd.lean b/Mathlib/Data/Nat/Choose/Dvd.lean index 0272d236dd30a..37ff1fb69906d 100644 --- a/Mathlib/Data/Nat/Choose/Dvd.lean +++ b/Mathlib/Data/Nat/Choose/Dvd.lean @@ -13,8 +13,6 @@ import Mathlib.Data.Nat.Prime.Factorial namespace Nat -open Nat - namespace Prime variable {p a b k : ℕ} diff --git a/Mathlib/Data/Nat/MaxPowDiv.lean b/Mathlib/Data/Nat/MaxPowDiv.lean index 34905e1420912..c9ea1fcff18cf 100644 --- a/Mathlib/Data/Nat/MaxPowDiv.lean +++ b/Mathlib/Data/Nat/MaxPowDiv.lean @@ -21,8 +21,6 @@ The implementation of `maxPowDiv` improves on the speed of `padicValNat`. namespace Nat -open Nat - /-- Tail recursive function which returns the largest `k : ℕ` such that `p ^ k ∣ n` for any `p : ℕ`. `padicValNat_eq_maxPowDiv` allows the code generator to use this definition for `padicValNat` diff --git a/Mathlib/Data/Nat/Periodic.lean b/Mathlib/Data/Nat/Periodic.lean index 1e2e335bec976..0165a52cf2c76 100644 --- a/Mathlib/Data/Nat/Periodic.lean +++ b/Mathlib/Data/Nat/Periodic.lean @@ -18,7 +18,7 @@ periodic predicates which helps determine their cardinality when filtering inter namespace Nat -open Nat Function +open Function theorem periodic_gcd (a : ℕ) : Periodic (gcd a) a := by simp only [forall_const, gcd_add_self_right, eq_self_iff_true, Periodic] diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index f12050bbcfeaf..dd45de30bdee3 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -16,8 +16,6 @@ import Mathlib.Data.PNat.Defs namespace Rat -open Rat - theorem num_dvd (a) {b : ℤ} (b0 : b ≠ 0) : (a /. b).num ∣ a := by cases' e : a /. b with n d h c rw [Rat.mk'_eq_divInt, divInt_eq_iff b0 (mod_cast h)] at e diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index 62ea491e08723..59bde31ad0787 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -452,8 +452,6 @@ section Polynomial open Polynomial -open Polynomial - variable (x : ℝ) (p : ℤ[X]) theorem one_lt_natDegree_of_irrational_root (hx : Irrational x) (p_nonzero : p ≠ 0) diff --git a/Mathlib/FieldTheory/IntermediateField/Basic.lean b/Mathlib/FieldTheory/IntermediateField/Basic.lean index 994c57c28bf6d..33d4aa8b3bf97 100644 --- a/Mathlib/FieldTheory/IntermediateField/Basic.lean +++ b/Mathlib/FieldTheory/IntermediateField/Basic.lean @@ -37,8 +37,6 @@ intermediate field, field extension -/ -open Polynomial - open Polynomial variable (K L L' : Type*) [Field K] [Field L] [Field L'] [Algebra K L] [Algebra K L'] diff --git a/Mathlib/LinearAlgebra/CrossProduct.lean b/Mathlib/LinearAlgebra/CrossProduct.lean index 8fb6567b4ec21..3093c7a58ad50 100644 --- a/Mathlib/LinearAlgebra/CrossProduct.lean +++ b/Mathlib/LinearAlgebra/CrossProduct.lean @@ -36,8 +36,6 @@ crossproduct -/ -open Matrix - open Matrix variable {R : Type*} [CommRing R] diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean index 7f0eb6a40db51..e2a73551493a8 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Minpoly.lean @@ -28,8 +28,6 @@ open Finset namespace Matrix -open Matrix - variable (M : Matrix n n R) @[simp] diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index dafa3eb44607d..1001a688e6203 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -42,8 +42,6 @@ open Equiv Equiv.Perm Finset Function namespace Matrix -open Matrix - variable {m n : Type*} [DecidableEq n] [Fintype n] [DecidableEq m] [Fintype m] variable {R : Type v} [CommRing R] diff --git a/Mathlib/LinearAlgebra/Matrix/Symmetric.lean b/Mathlib/LinearAlgebra/Matrix/Symmetric.lean index 83d9a2f3dccba..498e22fb3bacc 100644 --- a/Mathlib/LinearAlgebra/Matrix/Symmetric.lean +++ b/Mathlib/LinearAlgebra/Matrix/Symmetric.lean @@ -25,8 +25,6 @@ variable {α β n m R : Type*} namespace Matrix -open Matrix - /-- A matrix `A : Matrix n n α` is "symmetric" if `Aᵀ = A`. -/ def IsSymm (A : Matrix n n α) : Prop := Aᵀ = A diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 1580376e82cc2..1c777ee3fb514 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -63,8 +63,6 @@ universe u₁ u₂ namespace Matrix -open Matrix - variable (n p : Type*) (R : Type u₂) {𝕜 : Type*} [Field 𝕜] variable [DecidableEq n] [DecidableEq p] variable [CommRing R] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index 328b2319e0765..c9ef8e47cf38d 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -37,8 +37,6 @@ variable {ι : Type*} {R : Type*} {M₁ M₂ N₁ N₂ P : Type*} {Mᵢ Nᵢ : namespace QuadraticMap -open QuadraticMap - section Prod section Semiring diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 7ab91eabbf17b..b2da38021c0f9 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -512,8 +512,6 @@ variable [MeasurableSpace α] [TopologicalSpace β] open Filter -open Filter - @[aesop safe 20 (rule_sets := [Measurable])] protected theorem sup [Max β] [ContinuousSup β] (hf : StronglyMeasurable f) (hg : StronglyMeasurable g) : StronglyMeasurable (f ⊔ g) := diff --git a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean index fabb26655f44b..6d1c1c2f19d9b 100644 --- a/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean +++ b/Mathlib/NumberTheory/ClassNumber/AdmissibleCardPowDegree.lean @@ -23,8 +23,6 @@ use to show the class number of the ring of integers of a function field is fini namespace Polynomial -open Polynomial - open AbsoluteValue Real variable {Fq : Type*} [Fintype Fq] diff --git a/Mathlib/NumberTheory/Primorial.lean b/Mathlib/NumberTheory/Primorial.lean index 3867a3aa0f893..d2a7922592c84 100644 --- a/Mathlib/NumberTheory/Primorial.lean +++ b/Mathlib/NumberTheory/Primorial.lean @@ -27,8 +27,6 @@ open Finset open Nat -open Nat - /-- The primorial `n#` of `n` is the product of the primes less than or equal to `n`. -/ def primorial (n : ℕ) : ℕ := ∏ p ∈ range (n + 1) with p.Prime, p diff --git a/Mathlib/RingTheory/Adjoin/PowerBasis.lean b/Mathlib/RingTheory/Adjoin/PowerBasis.lean index 115e9eae03f82..2a3bbcac0b345 100644 --- a/Mathlib/RingTheory/Adjoin/PowerBasis.lean +++ b/Mathlib/RingTheory/Adjoin/PowerBasis.lean @@ -79,8 +79,6 @@ namespace PowerBasis open Polynomial -open Polynomial - variable {R : Type*} [CommRing R] [Algebra R S] [Algebra R K] [IsScalarTower R K S] variable {A : Type*} [CommRing A] [Algebra R A] [Algebra S A] variable [IsScalarTower R S A] {B : PowerBasis S A} diff --git a/Mathlib/RingTheory/EisensteinCriterion.lean b/Mathlib/RingTheory/EisensteinCriterion.lean index 216ad0c0c0a4d..0e246b9233e41 100644 --- a/Mathlib/RingTheory/EisensteinCriterion.lean +++ b/Mathlib/RingTheory/EisensteinCriterion.lean @@ -22,8 +22,6 @@ variable {R : Type*} [CommRing R] namespace Polynomial -open Polynomial - namespace EisensteinCriterionAux -- Section for auxiliary lemmas used in the proof of `irreducible_of_eisenstein_criterion` diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index b292e8b1c1ffc..a44906bcb806f 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -149,8 +149,6 @@ section EuclideanDivision namespace Polynomial -open Polynomial - variable (K : Type) [Field K] [Algebra R[X] K] [IsFractionRing R[X] K] theorem div_eq_quo_add_rem_div (f : R[X]) {g : R[X]} (hg : g.Monic) : diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index 14e36f678fd84..56997e889c66f 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -247,8 +247,6 @@ end Localization namespace Polynomial -open Polynomial - section CommRing -- Porting note: move to better place diff --git a/Mathlib/RingTheory/MatrixAlgebra.lean b/Mathlib/RingTheory/MatrixAlgebra.lean index 0750d23424344..8fcf5dc5f622c 100644 --- a/Mathlib/RingTheory/MatrixAlgebra.lean +++ b/Mathlib/RingTheory/MatrixAlgebra.lean @@ -14,13 +14,7 @@ suppress_compilation universe u v w -open TensorProduct - -open TensorProduct - -open Algebra.TensorProduct - -open Matrix +open TensorProduct Algebra.TensorProduct Matrix variable {R : Type u} [CommSemiring R] variable {A : Type v} [Semiring A] [Algebra R A] diff --git a/Mathlib/RingTheory/Nullstellensatz.lean b/Mathlib/RingTheory/Nullstellensatz.lean index f28499af60736..71d381cfeb340 100644 --- a/Mathlib/RingTheory/Nullstellensatz.lean +++ b/Mathlib/RingTheory/Nullstellensatz.lean @@ -28,8 +28,6 @@ noncomputable section namespace MvPolynomial -open MvPolynomial - variable {k : Type*} [Field k] variable {σ : Type*} diff --git a/Mathlib/RingTheory/Polynomial/Content.lean b/Mathlib/RingTheory/Polynomial/Content.lean index c92ff552d810e..394d392597e7b 100644 --- a/Mathlib/RingTheory/Polynomial/Content.lean +++ b/Mathlib/RingTheory/Polynomial/Content.lean @@ -30,8 +30,6 @@ Let `p : R[X]`. namespace Polynomial -open Polynomial - section Primitive variable {R : Type*} [CommSemiring R] diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index d5d51a3dee789..1872e6ac8a9e2 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -52,8 +52,6 @@ noncomputable section namespace Polynomial -open Polynomial - variable {R S : Type*} [CommRing R] [CommRing S] (k : ℕ) (a : R) /-- `dickson` is the `n`-th (generalised) Dickson polynomial of the `k`-th kind associated to the diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index 63708f272d1ed..d6872afb27cbc 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -38,8 +38,6 @@ universe u v open Polynomial -open Polynomial - section Semiring variable (S : Type u) [Semiring S] diff --git a/Mathlib/RingTheory/Polynomial/ScaleRoots.lean b/Mathlib/RingTheory/Polynomial/ScaleRoots.lean index dce2e27c49300..38ee5b2e41344 100644 --- a/Mathlib/RingTheory/Polynomial/ScaleRoots.lean +++ b/Mathlib/RingTheory/Polynomial/ScaleRoots.lean @@ -20,8 +20,6 @@ variable {R S A K : Type*} namespace Polynomial -open Polynomial - section Semiring variable [Semiring R] [Semiring S] diff --git a/Mathlib/RingTheory/ReesAlgebra.lean b/Mathlib/RingTheory/ReesAlgebra.lean index 127232d147d6d..2069a0781786a 100644 --- a/Mathlib/RingTheory/ReesAlgebra.lean +++ b/Mathlib/RingTheory/ReesAlgebra.lean @@ -30,8 +30,6 @@ variable {R M : Type u} [CommRing R] [AddCommGroup M] [Module R M] (I : Ideal R) open Polynomial -open Polynomial - /-- The Rees algebra of an ideal `I`, defined as the subalgebra of `R[X]` whose `i`-th coefficient falls in `I ^ i`. -/ def reesAlgebra : Subalgebra R R[X] where diff --git a/Mathlib/SetTheory/Cardinal/Divisibility.lean b/Mathlib/SetTheory/Cardinal/Divisibility.lean index 6a176f3246b65..a82232c834d80 100644 --- a/Mathlib/SetTheory/Cardinal/Divisibility.lean +++ b/Mathlib/SetTheory/Cardinal/Divisibility.lean @@ -31,8 +31,6 @@ Note furthermore that no infinite cardinal is irreducible namespace Cardinal -open Cardinal - universe u variable {a b : Cardinal.{u}} {n m : ℕ} diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean index 4f3a24fdd08bc..d339eba913c78 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Basic.lean @@ -105,8 +105,6 @@ end NonarchimedeanGroup namespace NonarchimedeanRing -open NonarchimedeanRing - open NonarchimedeanAddGroup variable {R S : Type*} diff --git a/Mathlib/Topology/Algebra/Polynomial.lean b/Mathlib/Topology/Algebra/Polynomial.lean index 663b8bd81dc20..cda8bf62a44a1 100644 --- a/Mathlib/Topology/Algebra/Polynomial.lean +++ b/Mathlib/Topology/Algebra/Polynomial.lean @@ -37,8 +37,6 @@ open IsAbsoluteValue Filter namespace Polynomial -open Polynomial - section TopologicalSemiring variable {R S : Type*} [Semiring R] [TopologicalSpace R] [TopologicalSemiring R] (p : R[X]) diff --git a/MathlibTest/matrix.lean b/MathlibTest/matrix.lean index d8571679b3a9c..f6e84240f8d07 100644 --- a/MathlibTest/matrix.lean +++ b/MathlibTest/matrix.lean @@ -13,8 +13,6 @@ variable {α β : Type} [Semiring α] [Ring β] namespace Matrix -open Matrix - /-! Test that the dimensions are inferred correctly, even for empty matrices -/ section dimensions diff --git a/MathlibTest/norm_cast.lean b/MathlibTest/norm_cast.lean index cb792593c1ef5..2cd0529030d98 100644 --- a/MathlibTest/norm_cast.lean +++ b/MathlibTest/norm_cast.lean @@ -126,8 +126,6 @@ example (a b : ℕ) (h2 : ((a + b + 0 : ℕ) : ℤ) = 10) : -- example {x : ℚ} : ((x + 42 : ℚ) : ℝ) = x + 42 := by push_cast namespace ENNReal - -open ENNReal lemma half_lt_self_bis {a : ℝ≥0∞} (hz : a ≠ 0) (ht : a ≠ ⊤) : a / 2 < a := by lift a to NNReal using ht have h : (2 : ℝ≥0∞) = ((2 : NNReal) : ℝ≥0∞) := rfl From ef950f1e9fabe40ab2cc6447154355a295c85a74 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 20 Nov 2024 09:05:41 +0000 Subject: [PATCH 49/90] chore: delete unused private lemma in Complex/exponential (#19263) --- Mathlib/Data/Complex/Exponential.lean | 5 ----- 1 file changed, 5 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 87be1b9c712d2..1054c7a046c4b 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -456,9 +456,6 @@ theorem cos_zero : cos 0 = 1 := by simp [cos] @[simp] theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] -private theorem cos_add_aux {a b c d : ℂ} : - (a + b) * (c + d) - (b - a) * (d - c) * -1 = 2 * (a * c + b * d) := by ring - theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] @@ -1498,5 +1495,3 @@ theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = rw [abs_exp, abs_exp, Real.exp_eq_exp] end Complex - -set_option linter.style.longFile 1700 From 2fb27be908b6ad6a43a55a68946651eb967332f1 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 20 Nov 2024 09:38:40 +0000 Subject: [PATCH 50/90] chore(Algebra/Module): further split `Defs.lean` (#18995) The goal is that we can define `Module` without depending on `Units`. The following moves out of `Defs.lean` are included in this PR: * Module structure over Nat and Int go to `Module/NatInt.lean` * Composition with `RingHom`s goes to `Module/RingHom.lean` * Results about units go to `Module/Basic.lean` --- Mathlib.lean | 2 + Mathlib/Algebra/AddConstMap/Basic.lean | 2 +- Mathlib/Algebra/CharZero/Quotient.lean | 2 +- .../Algebra/Group/Submonoid/Pointwise.lean | 1 + Mathlib/Algebra/Homology/TotalComplex.lean | 1 + Mathlib/Algebra/Module/Basic.lean | 8 +- Mathlib/Algebra/Module/Defs.lean | 164 +----------------- Mathlib/Algebra/Module/End.lean | 1 + Mathlib/Algebra/Module/Hom.lean | 1 + Mathlib/Algebra/Module/LinearMap/Defs.lean | 2 + Mathlib/Algebra/Module/NatInt.lean | 146 ++++++++++++++++ Mathlib/Algebra/Module/RingHom.lean | 92 ++++++++++ Mathlib/Algebra/Order/Nonneg/Module.lean | 2 +- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 2 +- Mathlib/CategoryTheory/Preadditive/Basic.lean | 1 + .../SimpleGraph/Regularity/Energy.lean | 2 +- Mathlib/Data/Finsupp/Pointwise.lean | 1 + Mathlib/Data/Int/AbsoluteValue.lean | 2 +- Mathlib/Data/Matrix/Diagonal.lean | 2 + Mathlib/Data/Set/Pointwise/SMul.lean | 1 + Mathlib/GroupTheory/ArchimedeanDensely.lean | 5 +- Mathlib/GroupTheory/Divisible.lean | 1 + Mathlib/GroupTheory/FreeAbelianGroup.lean | 1 + Mathlib/GroupTheory/OrderOfElement.lean | 1 + Mathlib/GroupTheory/QuotientGroup/Basic.lean | 1 + Mathlib/Order/Filter/Germ/Basic.lean | 3 +- 26 files changed, 281 insertions(+), 166 deletions(-) create mode 100644 Mathlib/Algebra/Module/NatInt.lean create mode 100644 Mathlib/Algebra/Module/RingHom.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3ca25acb72351..797b5848ce70d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -520,6 +520,7 @@ import Mathlib.Algebra.Module.LocalizedModule.Int import Mathlib.Algebra.Module.LocalizedModule.IsLocalization import Mathlib.Algebra.Module.LocalizedModule.Submodule import Mathlib.Algebra.Module.MinimalAxioms +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Module.Opposite import Mathlib.Algebra.Module.PID import Mathlib.Algebra.Module.Pi @@ -533,6 +534,7 @@ import Mathlib.Algebra.Module.Presentation.Tautological import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Projective import Mathlib.Algebra.Module.Rat +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Module.SnakeLemma import Mathlib.Algebra.Module.Submodule.Basic import Mathlib.Algebra.Module.Submodule.Bilinear diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index 0bf23cde46517..7dc43d9c08e9a 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.GroupPower.IterateHom -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.Group.Instances import Mathlib.Logic.Function.Iterate diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index a0e14ca270865..d3872fea29c7e 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Order.Group.Unbundled.Int -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt import Mathlib.GroupTheory.QuotientGroup.Defs import Mathlib.Algebra.Group.Subgroup.ZPowers.Basic diff --git a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean index 2c5e53bea3b80..9c5e50a98eb6b 100644 --- a/Mathlib/Algebra/Group/Submonoid/Pointwise.lean +++ b/Mathlib/Algebra/Group/Submonoid/Pointwise.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Hom.End import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Order.WellFoundedSet diff --git a/Mathlib/Algebra/Homology/TotalComplex.lean b/Mathlib/Algebra/Homology/TotalComplex.lean index 645d747de064f..15137d2628a9b 100644 --- a/Mathlib/Algebra/Homology/TotalComplex.lean +++ b/Mathlib/Algebra/Homology/TotalComplex.lean @@ -6,6 +6,7 @@ Authors: Joël Riou import Mathlib.CategoryTheory.Linear.Basic import Mathlib.Algebra.Homology.ComplexShapeSigns import Mathlib.Algebra.Homology.HomologicalBicomplex +import Mathlib.Algebra.Module.Basic /-! # The total complex of a bicomplex diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 0eeee9934e3d3..b653c02aa9f6c 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -6,7 +6,8 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.Group.Indicator -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.NoZeroSMulDivisors.Defs /-! @@ -23,6 +24,11 @@ universe u v variable {α R M M₂ : Type*} +@[simp] +theorem Units.neg_smul [Ring R] [AddCommGroup M] [Module R M] (u : Rˣ) (x : M) : + -u • x = -(u • x) := by + rw [Units.smul_def, Units.val_neg, _root_.neg_smul, Units.smul_def] + @[simp] theorem invOf_two_smul_add_invOf_two_smul (R) [Semiring R] [AddCommMonoid M] [Module R M] [Invertible (2 : R)] (x : M) : diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index c614c036f0977..ebc9ab4254580 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -3,10 +3,11 @@ 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.Group.Action.End +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.Ring.Defs import Mathlib.Algebra.SMulWithZero -import Mathlib.Data.Int.Cast.Lemmas /-! # Modules over a ring @@ -39,7 +40,9 @@ assert_not_exists Field assert_not_exists Invertible assert_not_exists Multiset assert_not_exists Pi.single_smul₀ +assert_not_exists RingHom assert_not_exists Set.indicator +assert_not_exists Units open Function Set @@ -72,14 +75,6 @@ instance (priority := 100) Module.toMulActionWithZero smul_zero := smul_zero zero_smul := Module.zero_smul } -instance AddCommGroup.toNatModule : Module ℕ M where - one_smul := one_nsmul - mul_smul m n a := mul_nsmul' a m n - smul_add n a b := nsmul_add a b n - smul_zero := nsmul_zero - zero_smul := zero_nsmul - add_smul r s x := add_nsmul x r s - theorem add_smul : (r + s) • x = r • x + s • x := Module.add_smul r s x @@ -111,31 +106,7 @@ protected abbrev Function.Surjective.module [AddCommMonoid M₂] [SMul R M₂] ( rcases hf x with ⟨x, rfl⟩ rw [← f.map_zero, ← smul, zero_smul] } -/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →+* S`. - -See also `Function.Surjective.mulActionLeft` and `Function.Surjective.distribMulActionLeft`. --/ -abbrev Function.Surjective.moduleLeft {R S M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] - [Semiring S] [SMul S M] (f : R →+* S) (hf : Function.Surjective f) - (hsmul : ∀ (c) (x : M), f c • x = c • x) : Module S M := - { hf.distribMulActionLeft f.toMonoidHom hsmul with - zero_smul := fun x => by rw [← f.map_zero, hsmul, zero_smul] - add_smul := hf.forall₂.mpr fun a b x => by simp only [← f.map_add, hsmul, add_smul] } - -variable {R} (M) - -/-- Compose a `Module` with a `RingHom`, with action `f s • m`. - -See note [reducible non-instances]. -/ -abbrev Module.compHom [Semiring S] (f : S →+* R) : Module S M := - { MulActionWithZero.compHom M f.toMonoidWithZeroHom, DistribMulAction.compHom M (f : S →* R) with - -- Porting note: the `show f (r + s) • x = f r • x + f s • x` wasn't needed in mathlib3. - -- Somehow, now that `SMul` is heterogeneous, it can't unfold earlier fields of a definition for - -- use in later fields. See - -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication - add_smul := fun r s x => show f (r + s) • x = f r • x + f s • x by simp [add_smul] } - -variable {M} +variable {R} theorem Module.eq_zero_of_zero_eq_one (zero_eq_one : (0 : R) = 1) : x = 0 := by rw [← one_smul R x, ← zero_eq_one, zero_smul] @@ -146,20 +117,9 @@ theorem smul_add_one_sub_smul {R : Type*} [Ring R] [Module R M] {r : R} {m : M} end AddCommMonoid - section AddCommGroup -variable (R M) [Semiring R] [AddCommGroup M] - -instance AddCommGroup.toIntModule : Module ℤ M where - one_smul := one_zsmul - mul_smul m n a := mul_zsmul a m n - smul_add n a b := zsmul_add a b n - smul_zero := zsmul_zero - zero_smul := zero_zsmul - add_smul r s x := add_zsmul x r s - -variable {R M} +variable [Semiring R] [AddCommGroup M] theorem Convex.combo_eq_smul_sub_add [Module R M] {x y : M} {a b : R} (h : a + b = 1) : a • x + b • y = b • (y - x) + x := @@ -187,10 +147,6 @@ theorem neg_smul : -r • x = -(r • x) := theorem neg_smul_neg : -r • -x = r • x := by rw [neg_smul, smul_neg, neg_neg] -@[simp] -theorem Units.neg_smul (u : Rˣ) (x : M) : -u • x = -(u • x) := by - rw [Units.smul_def, Units.val_neg, _root_.neg_smul, Units.smul_def] - variable (R) theorem neg_one_smul (x : M) : (-1 : R) • x = -x := by simp @@ -202,26 +158,6 @@ theorem sub_smul (r s : R) (y : M) : (r - s) • y = r • y - s • y := by end Module -variable (R) - -/-- An `AddCommMonoid` that is a `Module` over a `Ring` carries a natural `AddCommGroup` -structure. -See note [reducible non-instances]. -/ -abbrev Module.addCommMonoidToAddCommGroup - [Ring R] [AddCommMonoid M] [Module R M] : AddCommGroup M := - { (inferInstance : AddCommMonoid M) with - neg := fun a => (-1 : R) • a - neg_add_cancel := fun a => - show (-1 : R) • a + a = 0 by - nth_rw 2 [← one_smul R a] - rw [← add_smul, neg_add_cancel, zero_smul] - zsmul := fun z a => (z : R) • a - zsmul_zero' := fun a => by simpa only [Int.cast_zero] using zero_smul R a - zsmul_succ' := fun z a => by simp [add_comm, add_smul] - zsmul_neg' := fun z a => by simp [← smul_assoc, neg_one_smul] } - -variable {R} - /-- A module over a `Subsingleton` semiring is a `Subsingleton`. We cannot register this as an instance because Lean has no way to guess `R`. -/ protected theorem Module.subsingleton (R M : Type*) [Semiring R] [Subsingleton R] [AddCommMonoid M] @@ -242,87 +178,3 @@ instance (priority := 910) Semiring.toModule [Semiring R] : Module R R where instance [NonUnitalNonAssocSemiring R] : DistribSMul R R where smul_add := left_distrib - -/-- A ring homomorphism `f : R →+* M` defines a module structure by `r • x = f r * x`. -/ -def RingHom.toModule [Semiring R] [Semiring S] (f : R →+* S) : Module R S := - Module.compHom S f - -/-- If the module action of `R` on `S` is compatible with multiplication on `S`, then -`fun x ↦ x • 1` is a ring homomorphism from `R` to `S`. - -This is the `RingHom` version of `MonoidHom.smulOneHom`. - -When `R` is commutative, usually `algebraMap` should be preferred. -/ -@[simps!] def RingHom.smulOneHom - [Semiring R] [NonAssocSemiring S] [Module R S] [IsScalarTower R S S] : R →+* S where - __ := MonoidHom.smulOneHom - map_zero' := zero_smul R 1 - map_add' := (add_smul · · 1) - -/-- A homomorphism between semirings R and S can be equivalently specified by a R-module -structure on S such that S/S/R is a scalar tower. -/ -def ringHomEquivModuleIsScalarTower [Semiring R] [Semiring S] : - (R →+* S) ≃ {_inst : Module R S // IsScalarTower R S S} where - toFun f := ⟨Module.compHom S f, SMul.comp.isScalarTower _⟩ - invFun := fun ⟨_, _⟩ ↦ RingHom.smulOneHom - left_inv f := RingHom.ext fun r ↦ mul_one (f r) - right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| Module.ext <| funext₂ <| smul_one_smul S - -section AddCommMonoid - -variable [Semiring R] [AddCommMonoid M] [Module R M] - -section - -variable (R) - -/-- `nsmul` is equal to any other module structure via a cast. -/ -lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by - induction n with - | zero => rw [Nat.cast_zero, zero_smul, zero_smul] - | succ n ih => rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] - -/-- `nsmul` is equal to any other module structure via a cast. -/ --- See note [no_index around OfNat.ofNat] -lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : - (no_index (OfNat.ofNat n) : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. - -/-- `nsmul` is equal to any other module structure via a cast. -/ -@[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] -lemma nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := (Nat.cast_smul_eq_nsmul ..).symm - -end - -/-- Convert back any exotic `ℕ`-smul to the canonical instance. This should not be needed since in -mathlib all `AddCommMonoid`s should normally have exactly one `ℕ`-module structure by design. --/ -theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : @SMul.smul ℕ M h.toSMul n x = n • x := - Nat.cast_smul_eq_nsmul .. - -/-- All `ℕ`-module structures are equal. Not an instance since in mathlib all `AddCommMonoid` -should normally have exactly one `ℕ`-module structure by design. -/ -def AddCommMonoid.uniqueNatModule : Unique (Module ℕ M) where - default := by infer_instance - uniq P := (Module.ext' P _) fun n => by convert nat_smul_eq_nsmul P n - -instance AddCommMonoid.nat_isScalarTower : IsScalarTower ℕ R M where - smul_assoc n x y := by - induction n with - | zero => simp only [zero_smul] - | succ n ih => simp only [add_smul, one_smul, ih] - -end AddCommMonoid - -theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] - [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] - [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by - simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] - -theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by - rw [nsmul_eq_mul, mul_one] - -theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by - rw [zsmul_eq_mul, mul_one] - -@[deprecated (since := "2024-05-03")] alias Nat.smul_one_eq_coe := Nat.smul_one_eq_cast -@[deprecated (since := "2024-05-03")] alias Int.smul_one_eq_coe := Int.smul_one_eq_cast diff --git a/Mathlib/Algebra/Module/End.lean b/Mathlib/Algebra/Module/End.lean index b1269c071ad3d..405cb931c87f9 100644 --- a/Mathlib/Algebra/Module/End.lean +++ b/Mathlib/Algebra/Module/End.lean @@ -5,6 +5,7 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Group.Hom.End import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt /-! # Module structure and endomorphisms diff --git a/Mathlib/Algebra/Module/Hom.lean b/Mathlib/Algebra/Module/Hom.lean index 66bdfade35658..8149ac54e6b58 100644 --- a/Mathlib/Algebra/Module/Hom.lean +++ b/Mathlib/Algebra/Module/Hom.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Hom.Instances +import Mathlib.Algebra.GroupWithZero.Action.End import Mathlib.Algebra.Module.End import Mathlib.Algebra.Ring.Opposite import Mathlib.GroupTheory.GroupAction.DomAct.Basic diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 893c0c84ca370..c6436036fb7c1 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -5,6 +5,8 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Frédéric Dupuis, Heather Macbeth -/ import Mathlib.Algebra.Group.Hom.Instances +import Mathlib.Algebra.Module.NatInt +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Ring.CompTypeclasses import Mathlib.GroupTheory.GroupAction.Hom diff --git a/Mathlib/Algebra/Module/NatInt.lean b/Mathlib/Algebra/Module/NatInt.lean new file mode 100644 index 0000000000000..d1aab097a652d --- /dev/null +++ b/Mathlib/Algebra/Module/NatInt.lean @@ -0,0 +1,146 @@ +/- +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.Module.Defs +import Mathlib.Data.Int.Cast.Lemmas + +/-! +# Modules over `ℕ` and `ℤ` + +This file concerns modules where the scalars are the natural numbers or the integers. + +## Main definitions + +* `AddCommGroup.toNatModule`: any `AddCommMonoid` is (uniquely) a module over the naturals. + TODO: this name is not right! +* `AddCommGroup.toIntModule`: any `AddCommGroup` is a module over the integers. + +## Main results + + * `AddCommMonoid.uniqueNatModule`: there is an unique `AddCommMonoid ℕ M` structure for any `M` + +## Tags + +semimodule, module, vector space +-/ + +assert_not_exists Field +assert_not_exists Invertible +assert_not_exists Multiset +assert_not_exists Pi.single_smul₀ +assert_not_exists Set.indicator + +open Function Set + +universe u v + +variable {R S M M₂ : Type*} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) + +instance AddCommGroup.toNatModule : Module ℕ M where + one_smul := one_nsmul + mul_smul m n a := mul_nsmul' a m n + smul_add n a b := nsmul_add a b n + smul_zero := nsmul_zero + zero_smul := zero_nsmul + add_smul r s x := add_nsmul x r s + +end AddCommMonoid + +section AddCommGroup + +variable (R M) [Semiring R] [AddCommGroup M] + +instance AddCommGroup.toIntModule : Module ℤ M where + one_smul := one_zsmul + mul_smul m n a := mul_zsmul a m n + smul_add n a b := zsmul_add a b n + smul_zero := zsmul_zero + zero_smul := zero_zsmul + add_smul r s x := add_zsmul x r s + +end AddCommGroup + +variable (R) + +/-- An `AddCommMonoid` that is a `Module` over a `Ring` carries a natural `AddCommGroup` +structure. +See note [reducible non-instances]. -/ +abbrev Module.addCommMonoidToAddCommGroup + [Ring R] [AddCommMonoid M] [Module R M] : AddCommGroup M := + { (inferInstance : AddCommMonoid M) with + neg := fun a => (-1 : R) • a + neg_add_cancel := fun a => + show (-1 : R) • a + a = 0 by + nth_rw 2 [← one_smul R a] + rw [← add_smul, neg_add_cancel, zero_smul] + zsmul := fun z a => (z : R) • a + zsmul_zero' := fun a => by simpa only [Int.cast_zero] using zero_smul R a + zsmul_succ' := fun z a => by simp [add_comm, add_smul] + zsmul_neg' := fun z a => by simp [← smul_assoc, neg_one_smul] } + +variable {R} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] + +section + +variable (R) + +/-- `nsmul` is equal to any other module structure via a cast. -/ +lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by + induction n with + | zero => rw [Nat.cast_zero, zero_smul, zero_smul] + | succ n ih => rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] + +/-- `nsmul` is equal to any other module structure via a cast. -/ +-- See note [no_index around OfNat.ofNat] +lemma ofNat_smul_eq_nsmul (n : ℕ) [n.AtLeastTwo] (b : M) : + (no_index (OfNat.ofNat n) : R) • b = OfNat.ofNat n • b := Nat.cast_smul_eq_nsmul .. + +/-- `nsmul` is equal to any other module structure via a cast. -/ +@[deprecated Nat.cast_smul_eq_nsmul (since := "2024-07-23")] +lemma nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := (Nat.cast_smul_eq_nsmul ..).symm + +end + +/-- Convert back any exotic `ℕ`-smul to the canonical instance. This should not be needed since in +mathlib all `AddCommMonoid`s should normally have exactly one `ℕ`-module structure by design. +-/ +theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : @SMul.smul ℕ M h.toSMul n x = n • x := + Nat.cast_smul_eq_nsmul .. + +/-- All `ℕ`-module structures are equal. Not an instance since in mathlib all `AddCommMonoid` +should normally have exactly one `ℕ`-module structure by design. -/ +def AddCommMonoid.uniqueNatModule : Unique (Module ℕ M) where + default := by infer_instance + uniq P := (Module.ext' P _) fun n => by convert nat_smul_eq_nsmul P n + +instance AddCommMonoid.nat_isScalarTower : IsScalarTower ℕ R M where + smul_assoc n x y := by + induction n with + | zero => simp only [zero_smul] + | succ n ih => simp only [add_smul, one_smul, ih] + +end AddCommMonoid + +theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] + [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] + [Module S M₂] (x : ℕ) (a : M) : f ((x : R) • a) = (x : S) • f a := by + simp only [Nat.cast_smul_eq_nsmul, AddMonoidHom.map_nsmul, map_nsmul] + +theorem Nat.smul_one_eq_cast {R : Type*} [NonAssocSemiring R] (m : ℕ) : m • (1 : R) = ↑m := by + rw [nsmul_eq_mul, mul_one] + +theorem Int.smul_one_eq_cast {R : Type*} [NonAssocRing R] (m : ℤ) : m • (1 : R) = ↑m := by + rw [zsmul_eq_mul, mul_one] + +@[deprecated (since := "2024-05-03")] alias Nat.smul_one_eq_coe := Nat.smul_one_eq_cast +@[deprecated (since := "2024-05-03")] alias Int.smul_one_eq_coe := Int.smul_one_eq_cast diff --git a/Mathlib/Algebra/Module/RingHom.lean b/Mathlib/Algebra/Module/RingHom.lean new file mode 100644 index 0000000000000..ae178799ace69 --- /dev/null +++ b/Mathlib/Algebra/Module/RingHom.lean @@ -0,0 +1,92 @@ +/- +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.Module.Defs +import Mathlib.Algebra.Ring.Hom.Defs + +/-! +# Composing modules with a ring hom + +## Main definitions + + * `Module.compHom`: compose a `Module` with a `RingHom`, with action `f s • m`. + * `RingHom.toModule`: a `RingHom` defines a module structure by `r • x = f r * x`. + +## Tags + +semimodule, module, vector space +-/ + +assert_not_exists Field +assert_not_exists Invertible +assert_not_exists Multiset +assert_not_exists Pi.single_smul₀ +assert_not_exists Set.indicator + +open Function Set + +universe u v + +variable {R S M M₂ : Type*} + +section AddCommMonoid + +variable [Semiring R] [AddCommMonoid M] [Module R M] (r s : R) (x : M) + +variable (R) + +/-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →+* S`. + +See also `Function.Surjective.mulActionLeft` and `Function.Surjective.distribMulActionLeft`. +-/ +abbrev Function.Surjective.moduleLeft {R S M : Type*} [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring S] [SMul S M] (f : R →+* S) (hf : Function.Surjective f) + (hsmul : ∀ (c) (x : M), f c • x = c • x) : Module S M := + { hf.distribMulActionLeft f.toMonoidHom hsmul with + zero_smul := fun x => by rw [← f.map_zero, hsmul, zero_smul] + add_smul := hf.forall₂.mpr fun a b x => by simp only [← f.map_add, hsmul, add_smul] } + +variable {R} (M) + +/-- Compose a `Module` with a `RingHom`, with action `f s • m`. + +See note [reducible non-instances]. -/ +abbrev Module.compHom [Semiring S] (f : S →+* R) : Module S M := + { MulActionWithZero.compHom M f.toMonoidWithZeroHom, DistribMulAction.compHom M (f : S →* R) with + -- Porting note: the `show f (r + s) • x = f r • x + f s • x` wasn't needed in mathlib3. + -- Somehow, now that `SMul` is heterogeneous, it can't unfold earlier fields of a definition for + -- use in later fields. See + -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Heterogeneous.20scalar.20multiplication + add_smul := fun r s x => show f (r + s) • x = f r • x + f s • x by simp [add_smul] } + +variable {M} + +end AddCommMonoid + +/-- A ring homomorphism `f : R →+* M` defines a module structure by `r • x = f r * x`. -/ +def RingHom.toModule [Semiring R] [Semiring S] (f : R →+* S) : Module R S := + Module.compHom S f + +/-- If the module action of `R` on `S` is compatible with multiplication on `S`, then +`fun x ↦ x • 1` is a ring homomorphism from `R` to `S`. + +This is the `RingHom` version of `MonoidHom.smulOneHom`. + +When `R` is commutative, usually `algebraMap` should be preferred. -/ +@[simps!] def RingHom.smulOneHom + [Semiring R] [NonAssocSemiring S] [Module R S] [IsScalarTower R S S] : R →+* S where + __ := MonoidHom.smulOneHom + map_zero' := zero_smul R 1 + map_add' := (add_smul · · 1) + +/-- A homomorphism between semirings R and S can be equivalently specified by a R-module +structure on S such that S/S/R is a scalar tower. -/ +def ringHomEquivModuleIsScalarTower [Semiring R] [Semiring S] : + (R →+* S) ≃ {_inst : Module R S // IsScalarTower R S S} where + toFun f := ⟨Module.compHom S f, SMul.comp.isScalarTower _⟩ + invFun := fun ⟨_, _⟩ ↦ RingHom.smulOneHom + left_inv f := RingHom.ext fun r ↦ mul_one (f r) + right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| Module.ext <| funext₂ <| smul_one_smul S diff --git a/Mathlib/Algebra/Order/Nonneg/Module.lean b/Mathlib/Algebra/Order/Nonneg/Module.lean index 1a388a67825b3..fb4548137a724 100644 --- a/Mathlib/Algebra/Order/Nonneg/Module.lean +++ b/Mathlib/Algebra/Order/Nonneg/Module.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Apurva Nakade. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Apurva Nakade -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Algebra.Order.Nonneg.Ring diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index d955416892c36..bcd9bb5829c17 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Submonoid.Membership -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.RingHom import Mathlib.Algebra.Ring.Action.Subobjects import Mathlib.Algebra.Ring.Equiv import Mathlib.Algebra.Ring.Prod diff --git a/Mathlib/CategoryTheory/Preadditive/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Basic.lean index 0ddcb0081738b..7fade9b2835eb 100644 --- a/Mathlib/CategoryTheory/Preadditive/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Basic.lean @@ -5,6 +5,7 @@ Authors: Markus Himmel, Jakob von Raumer -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.End import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Limits.Shapes.Kernels diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index 452746aebd488..ec9fd519eeeeb 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Combinatorics.SimpleGraph.Density import Mathlib.Data.Rat.BigOperators diff --git a/Mathlib/Data/Finsupp/Pointwise.lean b/Mathlib/Data/Finsupp/Pointwise.lean index 059bf76ad5eed..04b6b2a36e9d5 100644 --- a/Mathlib/Data/Finsupp/Pointwise.lean +++ b/Mathlib/Data/Finsupp/Pointwise.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.InjSurj import Mathlib.Algebra.Ring.Pi import Mathlib.Data.Finsupp.Defs diff --git a/Mathlib/Data/Int/AbsoluteValue.lean b/Mathlib/Data/Int/AbsoluteValue.lean index f061b555444e8..f37f06dffef6a 100644 --- a/Mathlib/Data/Int/AbsoluteValue.lean +++ b/Mathlib/Data/Int/AbsoluteValue.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Algebra.GroupWithZero.Action.Units -import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Module.Basic import Mathlib.Algebra.Order.AbsoluteValue import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Int.Cast.Lemmas diff --git a/Mathlib/Data/Matrix/Diagonal.lean b/Mathlib/Data/Matrix/Diagonal.lean index d9d2d31a41ec1..6e226513da340 100644 --- a/Mathlib/Data/Matrix/Diagonal.lean +++ b/Mathlib/Data/Matrix/Diagonal.lean @@ -3,7 +3,9 @@ Copyright (c) 2018 Ellen Arlt. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ellen Arlt, Blair Shi, Sean Leather, Mario Carneiro, Johan Commelin, Lu-Ming Zhang -/ +import Mathlib.Data.Int.Cast.Lemmas import Mathlib.Data.Matrix.Defs +import Mathlib.Data.Nat.Cast.Basic /-! # Diagonal matrices diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 68200c4096290..1f4a5da3be5cd 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin, Floris van Doorn import Mathlib.Algebra.Group.Pi.Basic import Mathlib.Algebra.Group.Pointwise.Set.Basic import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Ring.Opposite import Mathlib.Algebra.NoZeroSMulDivisors.Defs diff --git a/Mathlib/GroupTheory/ArchimedeanDensely.lean b/Mathlib/GroupTheory/ArchimedeanDensely.lean index 0434db1da4ce6..6e65e049cb576 100644 --- a/Mathlib/GroupTheory/ArchimedeanDensely.lean +++ b/Mathlib/GroupTheory/ArchimedeanDensely.lean @@ -3,12 +3,13 @@ Copyright (c) 2024 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ -import Mathlib.Data.Int.Interval -import Mathlib.GroupTheory.Archimedean import Mathlib.Algebra.Group.Equiv.TypeTags import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.Group.TypeTags import Mathlib.Algebra.Order.Hom.Monoid +import Mathlib.Data.Int.Interval +import Mathlib.GroupTheory.Archimedean /-! # Archimedean groups are either discrete or densely ordered diff --git a/Mathlib/GroupTheory/Divisible.lean b/Mathlib/GroupTheory/Divisible.lean index 99ead5b297aa7..5543e8f8d13df 100644 --- a/Mathlib/GroupTheory/Divisible.lean +++ b/Mathlib/GroupTheory/Divisible.lean @@ -5,6 +5,7 @@ Authors: Jujian Zhang -/ import Mathlib.Algebra.Group.ULift import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Algebra.Module.NatInt import Mathlib.GroupTheory.QuotientGroup.Defs import Mathlib.Tactic.NormNum.Eq diff --git a/Mathlib/GroupTheory/FreeAbelianGroup.lean b/Mathlib/GroupTheory/FreeAbelianGroup.lean index 163b552783eb7..ad652f38aa3fc 100644 --- a/Mathlib/GroupTheory/FreeAbelianGroup.lean +++ b/Mathlib/GroupTheory/FreeAbelianGroup.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.Module.NatInt import Mathlib.GroupTheory.Abelianization import Mathlib.GroupTheory.FreeGroup.Basic diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index f01189b606a68..b380c1f71c410 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer -/ import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.Algebra.Module.NatInt import Mathlib.Algebra.Order.Group.Action import Mathlib.Algebra.Order.Ring.Abs import Mathlib.GroupTheory.Index diff --git a/Mathlib/GroupTheory/QuotientGroup/Basic.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean index 4495ebe68b078..3c1822db6828d 100644 --- a/Mathlib/GroupTheory/QuotientGroup/Basic.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Patrick Massot -- This file is to a certain extent based on `quotient_module.lean` by Johannes Hölzl. import Mathlib.Algebra.Group.Subgroup.Pointwise +import Mathlib.Data.Int.Cast.Lemmas import Mathlib.GroupTheory.Congruence.Hom import Mathlib.GroupTheory.Coset.Basic import Mathlib.GroupTheory.QuotientGroup.Defs diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index d855beedbbf22..21eda99605828 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -3,8 +3,9 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ -import Mathlib.Order.Filter.Tendsto import Mathlib.Algebra.Module.Pi +import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Order.Filter.Tendsto /-! # Germ of a function at a filter From 6095b78b2e3595eb24b62993b94f120a83f7fa91 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 20 Nov 2024 10:13:28 +0000 Subject: [PATCH 51/90] feat(ModelTheory): Ax-Grothendieck (#6468) Co-authored-by: ChrisHughes24 Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> --- Mathlib/FieldTheory/AxGrothendieck.lean | 267 ++++++++++++++++++++---- 1 file changed, 225 insertions(+), 42 deletions(-) diff --git a/Mathlib/FieldTheory/AxGrothendieck.lean b/Mathlib/FieldTheory/AxGrothendieck.lean index 9ddfa55c77de3..b3ce93355bef9 100644 --- a/Mathlib/FieldTheory/AxGrothendieck.lean +++ b/Mathlib/FieldTheory/AxGrothendieck.lean @@ -3,61 +3,244 @@ Copyright (c) 2023 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import Mathlib.Algebra.MvPolynomial.Basic -import Mathlib.Data.Fintype.Card + import Mathlib.RingTheory.Algebraic +import Mathlib.Data.Fintype.Card +import Mathlib.ModelTheory.Algebra.Field.IsAlgClosed +import Mathlib.ModelTheory.Algebra.Ring.Definability +import Mathlib.RingTheory.Polynomial.Basic /-! -# Ax-Grothendieck for algebraic extensions of `ZMod p` +# Ax-Grothendieck -This file proves that if `R` is an algebraic extension of a finite field, -then any injective polynomial map `R^n → R^n` is also surjective. +This file proves that if `K` is an algebraically closed field, +then any injective polynomial map `K^n → K^n` is also surjective. -This proof is required for the true Ax-Grothendieck theorem, which proves the same result -for any algebraically closed field of characteristic zero. +## Main results -## TODO +* `ax_grothendieck_zeroLocus`: If `K` is algebraically closed, `ι` is a finite type, and +`S : Set (ι → K)` is the `zeroLocus` of some ideal of `MvPolynomial ι K`, then any injective +polynomial map `S → S` is also surjective on `S`. +* `ax_grothendieck_univ`: Any injective polynomial map `K^n → K^n` is also surjective if `K` is an +algberaically closed field. +* `ax_grothendieck_of_definable`: Any injective polynomial map `S → S` is also surjective on `S` if +`K` is an algebraically closed field and `S` is a definable subset of `K^n`. +* `ax_grothendieck_of_locally_finite`: any injective polynomial map `R^n → R^n` is also surjective +whenever `R` is an algebraic extension of a finite field. + +## References + +The first order theory of algebraically closed fields, along with the Lefschetz Principle and +the Ax-Grothendieck Theorem were first formalized in Lean 3 by Joseph Hua +[here](https://github.com/Jlh18/ModelTheoryInLean8) with the master's thesis +[here](https://github.com/Jlh18/ModelTheory8Report) -The proof of the theorem for characteristic zero is not in mathlib, but it is at -https://github.com/Jlh18/ModelTheoryInLean8 -/ noncomputable section -open MvPolynomial Finset Function +open MvPolynomial Finset /-- Any injective polynomial map over an algebraic extension of a finite field is surjective. -/ theorem ax_grothendieck_of_locally_finite {ι K R : Type*} [Field K] [Finite K] [CommRing R] - [Finite ι] [Algebra K R] [Algebra.IsAlgebraic K R] (ps : ι → MvPolynomial ι R) - (hinj : Injective fun v i => MvPolynomial.eval v (ps i)) : - Surjective fun v i => MvPolynomial.eval v (ps i) := by + [Finite ι] [Algebra K R] [alg : Algebra.IsAlgebraic K R] (ps : ι → MvPolynomial ι R) + (S : Set (ι → R)) + (hm : S.MapsTo (fun v i => eval v (ps i)) S) + (hinj : S.InjOn (fun v i => eval v (ps i))) : + S.SurjOn (fun v i => eval v (ps i)) S := by + have is_int : ∀ x : R, IsIntegral K x := fun x => isAlgebraic_iff_isIntegral.1 + (alg.isAlgebraic x) + classical + intro v hvS + cases nonempty_fintype ι + /- `s` is the set of all coefficients of the polynomial, as well as all of + the coordinates of `v`, the point I am trying to find the preimage of. -/ + let s : Finset R := + (Finset.biUnion (univ : Finset ι) fun i => (ps i).support.image fun x => coeff x (ps i)) ∪ + (univ : Finset ι).image v + have hv : ∀ i, v i ∈ Algebra.adjoin K (s : Set R) := fun j => + Algebra.subset_adjoin (mem_union_right _ (mem_image.2 ⟨j, mem_univ _, rfl⟩)) + have hs₁ : ∀ (i : ι) (k : ι →₀ ℕ), + k ∈ (ps i).support → coeff k (ps i) ∈ Algebra.adjoin K (s : Set R) := + fun i k hk => Algebra.subset_adjoin + (mem_union_left _ (mem_biUnion.2 ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)) + have := isNoetherian_adjoin_finset s fun x _ => is_int x + have := Module.IsNoetherian.finite K (Algebra.adjoin K (s : Set R)) + have : Finite (Algebra.adjoin K (s : Set R)) := Module.finite_of_finite K + -- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` + let S' : Set (ι → Algebra.adjoin K (s : Set R)) := + (fun v => Subtype.val ∘ v) ⁻¹' S + let res : S' → S' := fun x => ⟨fun i => + ⟨eval (fun j : ι => (x.1 j : R)) (ps i), eval_mem (hs₁ _) fun i => (x.1 i).2⟩, + hm x.2⟩ + have hres_surj : Function.Surjective res := by + rw [← Finite.injective_iff_surjective] + intro x y hxy + ext i + simp only [Subtype.ext_iff, funext_iff] at hxy + exact congr_fun (hinj x.2 y.2 (funext hxy)) i + rcases hres_surj ⟨fun i => ⟨v i, hv i⟩, hvS⟩ with ⟨⟨w, hwS'⟩, hw⟩ + refine ⟨fun i => w i, hwS', ?_⟩ + simpa [Subtype.ext_iff, funext_iff] using hw + +end + +namespace FirstOrder + +open MvPolynomial FreeCommRing Language Field Ring BoundedFormula + +variable {ι α : Type*} [Finite α] {K : Type*} [Field K] [CompatibleRing K] + +/-- The collection of first order formulas corresponding to the Ax-Grothendieck theorem. -/ +noncomputable def genericPolyMapSurjOnOfInjOn [Fintype ι] + (φ : ring.Formula (α ⊕ ι)) + (mons : ι → Finset (ι →₀ ℕ)) : Language.ring.Sentence := + let l1 : ι → Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + fun i => + (termOfFreeCommRing (genericPolyMap mons i)).relabel + (Sum.inl ∘ Sum.map id (fun i => (0, i))) + =' (termOfFreeCommRing (genericPolyMap mons i)).relabel + (Sum.inl ∘ Sum.map id (fun i => (1, i))) + -- p(x) = p(y) as a formula + let f1 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + iInf Finset.univ l1 + let l2 : ι → Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + fun i => .var (Sum.inl (Sum.inr (0, i))) =' .var (Sum.inl (Sum.inr (1, i))) + -- x = y as a formula + let f2 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + iInf Finset.univ l2 + let injOn : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := + Formula.iAlls (γ := Fin 2 × ι) id + (φ.relabel (Sum.map Sum.inl (fun i => (0, i))) ⟹ + φ.relabel (Sum.map Sum.inl (fun i => (1, i))) ⟹ + (f1.imp f2).relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x))) + let l3 : ι → Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + fun i => (termOfFreeCommRing (genericPolyMap mons i)).relabel + (Sum.inl ∘ Sum.map id (fun i => (0, i))) =' + .var (Sum.inl (Sum.inr (1, i))) + let f3 : Language.ring.Formula ((Σ i : ι, mons i) ⊕ (Fin 2 × ι)) := + iInf Finset.univ l3 + let surjOn : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := + Formula.iAlls (γ := ι) id + (Formula.imp (φ.relabel (Sum.map Sum.inl id)) <| + Formula.iExs (γ := ι) + (fun (i : (α ⊕ (Σ i : ι, mons i)) ⊕ (Fin 2 × ι)) => + show ((α ⊕ (Σ i : ι, mons i)) ⊕ ι) ⊕ ι + from Sum.elim (Sum.inl ∘ Sum.inl) + (fun i => if i.1 = 0 then Sum.inr i.2 else (Sum.inl (Sum.inr i.2))) i) + ((φ.relabel (Sum.map Sum.inl (fun i => (0, i)))) ⊓ + (f3.relabel (fun x => (Equiv.sumAssoc _ _ _).symm (Sum.inr x))))) + let mapsTo : Language.ring.Formula (α ⊕ Σ i : ι, mons i) := + Formula.iAlls (γ := ι) id + (Formula.imp (φ.relabel (Sum.map Sum.inl id)) + (φ.subst <| Sum.elim + (fun a => .var (Sum.inl (Sum.inl a))) + (fun i => (termOfFreeCommRing (genericPolyMap mons i)).relabel + (fun i => (Equiv.sumAssoc _ _ _).symm (Sum.inr i))))) + Formula.iAlls (γ := α ⊕ Σ i : ι, mons i) Sum.inr (mapsTo ⟹ injOn ⟹ surjOn) + +theorem realize_genericPolyMapSurjOnOfInjOn + [Fintype ι] (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : + (K ⊨ genericPolyMapSurjOnOfInjOn φ mons) ↔ + ∀ (v : α → K) (p : { p : ι → MvPolynomial ι K // (∀ i, (p i).support ⊆ mons i) }), + let f : (ι → K) → (ι → K) := fun v i => eval v (p.1 i) + let S : Set (ι → K) := fun x => φ.Realize (Sum.elim v x) + S.MapsTo f S → S.InjOn f → S.SurjOn f S := by classical - intro v - cases nonempty_fintype ι - /- `s` is the set of all coefficients of the polynomial, as well as all of - the coordinates of `v`, the point I am trying to find the preimage of. -/ - let s : Finset R := - (Finset.biUnion (univ : Finset ι) fun i => (ps i).support.image fun x => coeff x (ps i)) ∪ - (univ : Finset ι).image v - have hv : ∀ i, v i ∈ Algebra.adjoin K (s : Set R) := fun j => - Algebra.subset_adjoin (mem_union_right _ (mem_image.2 ⟨j, mem_univ _, rfl⟩)) - have hs₁ : ∀ (i : ι) (k : ι →₀ ℕ), - k ∈ (ps i).support → coeff k (ps i) ∈ Algebra.adjoin K (s : Set R) := - fun i k hk => Algebra.subset_adjoin - (mem_union_left _ (mem_biUnion.2 ⟨i, mem_univ _, mem_image_of_mem _ hk⟩)) - letI := isNoetherian_adjoin_finset s fun x _ => Algebra.IsIntegral.isIntegral (R := K) x - letI := Module.IsNoetherian.finite K (Algebra.adjoin K (s : Set R)) - letI : Finite (Algebra.adjoin K (s : Set R)) := Module.finite_of_finite K - -- The restriction of the polynomial map, `ps`, to the subalgebra generated by `s` - let res : (ι → Algebra.adjoin K (s : Set R)) → ι → Algebra.adjoin K (s : Set R) := fun x i => - ⟨eval (fun j : ι => (x j : R)) (ps i), eval_mem (hs₁ _) fun i => (x i).2⟩ - have hres_inj : Injective res := by - intro x y hxy - ext i - simp only [Subtype.ext_iff, funext_iff] at hxy - exact congr_fun (hinj (funext hxy)) i - have hres_surj : Surjective res := Finite.injective_iff_surjective.1 hres_inj - cases' hres_surj fun i => ⟨v i, hv i⟩ with w hw - use fun i => w i - simpa only [Subtype.ext_iff, funext_iff] using hw + have injOnAlt : ∀ {S : Set (ι → K)} (f : (ι → K) → (ι → K)), + S.InjOn f ↔ ∀ x y, x ∈ S → y ∈ S → f x = f y → x = y := by + simp [Set.InjOn]; tauto + simp only [Sentence.Realize, Formula.Realize, genericPolyMapSurjOnOfInjOn, Formula.relabel, + Function.comp_def, Sum.map, id_eq, Equiv.sumAssoc, Equiv.coe_fn_symm_mk, Sum.elim_inr, + realize_iAlls, realize_imp, realize_relabel, Fin.natAdd_zero, realize_subst, realize_iInf, + Finset.mem_univ, realize_bdEqual, Term.realize_relabel, true_imp_iff, + Equiv.forall_congr_left (Equiv.curry (Fin 2) ι K), Equiv.curry_symm_apply, Function.uncurry, + Fin.forall_fin_succ_pi, Fin.forall_fin_zero_pi, realize_iExs, realize_inf, Sum.forall_sum, + Set.MapsTo, Set.mem_def, injOnAlt, funext_iff, Set.SurjOn, Set.image, setOf, + Set.subset_def, Equiv.forall_congr_left (mvPolynomialSupportLEEquiv mons)] + simp +singlePass only [← Sum.elim_comp_inl_inr] + simp [Set.mem_def, Function.comp_def] + +theorem ACF_models_genericPolyMapSurjOnOfInjOn_of_prime [Fintype ι] + {p : ℕ} (hp : p.Prime) (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : + Theory.ACF p ⊨ᵇ genericPolyMapSurjOnOfInjOn φ mons := by + classical + have : Fact p.Prime := ⟨hp⟩ + letI := compatibleRingOfRing (AlgebraicClosure (ZMod p)) + have : CharP (AlgebraicClosure (ZMod p)) p := + charP_of_injective_algebraMap + (RingHom.injective (algebraMap (ZMod p) (AlgebraicClosure (ZMod p)))) p + rw [← (ACF_isComplete (Or.inl hp)).realize_sentence_iff _ + (AlgebraicClosure (ZMod p)), realize_genericPolyMapSurjOnOfInjOn] + rintro v ⟨f, _⟩ + exact ax_grothendieck_of_locally_finite (K := ZMod p) (ι := ι) f _ + +theorem ACF_models_genericPolyMapSurjOnOfInjOn_of_prime_or_zero + [Fintype ι] {p : ℕ} (hp : p.Prime ∨ p = 0) + (φ : ring.Formula (α ⊕ ι)) (mons : ι → Finset (ι →₀ ℕ)) : + Theory.ACF p ⊨ᵇ genericPolyMapSurjOnOfInjOn φ mons := by + rcases hp with hp | rfl + · exact ACF_models_genericPolyMapSurjOnOfInjOn_of_prime hp φ mons + · rw [ACF_zero_realize_iff_infinite_ACF_prime_realize] + convert Set.infinite_univ (α := Nat.Primes) + rw [Set.eq_univ_iff_forall] + intro ⟨p, hp⟩ + exact ACF_models_genericPolyMapSurjOnOfInjOn_of_prime hp φ mons + +end FirstOrder + +open FirstOrder Language Field Ring MvPolynomial + +variable {K ι : Type*} [Field K] [IsAlgClosed K] [Finite ι] + +/-- A slight generalization of the **Ax-Grothendieck** theorem + +If `K` is an algebraically closed field, `ι` is a finite type, and `S` is a definable subset of +`ι → K`, then any injective polynomial map `S → S` is also surjective on `S`. -/ +theorem ax_grothendieck_of_definable [CompatibleRing K] {c : Set K} + (S : Set (ι → K)) (hS : c.Definable Language.ring S) + (ps : ι → MvPolynomial ι K) : + S.MapsTo (fun v i => eval v (ps i)) S → + S.InjOn (fun v i => eval v (ps i)) → + S.SurjOn (fun v i => eval v (ps i)) S := by + letI := Fintype.ofFinite ι + let p : ℕ := ringChar K + have : CharP K p := ⟨ringChar.spec K⟩ + rw [Set.definable_iff_finitely_definable] at hS + rcases hS with ⟨c, _, hS⟩ + rw [Set.definable_iff_exists_formula_sum] at hS + rcases hS with ⟨φ, hφ⟩ + rw [hφ] + have := ACF_models_genericPolyMapSurjOnOfInjOn_of_prime_or_zero + (CharP.char_is_prime_or_zero K p) φ (fun i => (ps i).support) + rw [← (ACF_isComplete (CharP.char_is_prime_or_zero K p)).realize_sentence_iff _ K, + realize_genericPolyMapSurjOnOfInjOn] at this + exact this Subtype.val ⟨ps, fun i => Set.Subset.refl _⟩ + +/-- The **Ax-Grothendieck** theorem + +If `K` is an algebraically closed field, and `S : Set (ι → K)` is the `zeroLocus` of an ideal +of the multivariable polynomial ring, then any injective polynomial map `S → S` is also +surjective on `S`. -/ +theorem ax_grothendieck_zeroLocus + (I : Ideal (MvPolynomial ι K)) + (p : ι → MvPolynomial ι K) : + let S := zeroLocus I + S.MapsTo (fun v i => eval v (p i)) S → + S.InjOn (fun v i => eval v (p i)) → + S.SurjOn (fun v i => eval v (p i)) S := by + letI := compatibleRingOfRing K + intro S + obtain ⟨s, rfl⟩ : I.FG := IsNoetherian.noetherian I + exact ax_grothendieck_of_definable S (mvPolynomial_zeroLocus_definable s) p + +/-- A special case of the **Ax-Grothendieck** theorem + +Any injective polynomial map `K^n → K^n` is also surjective if `K` is an +algberaically closed field. -/ +theorem ax_grothendieck_univ (p : ι → MvPolynomial ι K) : + (fun v i => eval v (p i)).Injective → + (fun v i => eval v (p i)).Surjective := by + simpa [Set.injective_iff_injOn_univ, Set.surjective_iff_surjOn_univ] using + ax_grothendieck_zeroLocus 0 p From ad04df3876a303393c2a7a9e067bcecefce8a678 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 20 Nov 2024 10:44:46 +0000 Subject: [PATCH 52/90] feat: bounds for modular forms of non-positive weights (#18192) We use the Maximum modulus principle to show that modular forms of non-positive weights are bounded by its value of some element of the upper half plane imaginary part at least 1/2. Co-authored-by: Chris Birkbeck Co-authored-by: David Loeffler --- Mathlib.lean | 1 + .../Complex/UpperHalfPlane/Basic.lean | 158 +++++++++--------- Mathlib/NumberTheory/Modular.lean | 34 +++- .../ModularForms/EisensteinSeries/Defs.lean | 4 +- .../NumberTheory/ModularForms/LevelOne.lean | 38 +++++ .../ModularForms/SlashActions.lean | 6 +- 6 files changed, 153 insertions(+), 88 deletions(-) create mode 100644 Mathlib/NumberTheory/ModularForms/LevelOne.lean diff --git a/Mathlib.lean b/Mathlib.lean index 797b5848ce70d..f0d237abd75a7 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3788,6 +3788,7 @@ import Mathlib.NumberTheory.ModularForms.JacobiTheta.Bounds import Mathlib.NumberTheory.ModularForms.JacobiTheta.Manifold import Mathlib.NumberTheory.ModularForms.JacobiTheta.OneVariable import Mathlib.NumberTheory.ModularForms.JacobiTheta.TwoVariable +import Mathlib.NumberTheory.ModularForms.LevelOne import Mathlib.NumberTheory.ModularForms.QExpansion import Mathlib.NumberTheory.ModularForms.SlashActions import Mathlib.NumberTheory.ModularForms.SlashInvariantForms diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 52bf76205cdd1..0ac5af563ff8a 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -268,50 +268,9 @@ instance : MulAction GL(2, ℝ)⁺ ℍ where simp [num, denom] mul_smul := mul_smul' -section ModularScalarTowers - instance SLAction {R : Type*} [CommRing R] [Algebra R ℝ] : MulAction SL(2, R) ℍ := MulAction.compHom ℍ <| SpecialLinearGroup.toGLPos.comp <| map (algebraMap R ℝ) -namespace ModularGroup - -variable (Γ : Subgroup (SpecialLinearGroup (Fin 2) ℤ)) - -/-- Canonical embedding of `SL(2, ℤ)` into `GL(2, ℝ)⁺`. -/ -@[coe] -def coe' : SL(2, ℤ) → GL(2, ℝ)⁺ := fun g => ((g : SL(2, ℝ)) : GL(2, ℝ)⁺) - -instance : Coe SL(2, ℤ) GL(2, ℝ)⁺ := - ⟨coe'⟩ - -@[simp] -theorem coe'_apply_complex {g : SL(2, ℤ)} {i j : Fin 2} : - (Units.val <| Subtype.val <| coe' g) i j = (Subtype.val g i j : ℂ) := - rfl - -@[simp] -theorem det_coe' {g : SL(2, ℤ)} : det (Units.val <| Subtype.val <| coe' g) = 1 := by - simp only [SpecialLinearGroup.coe_GLPos_coe_GL_coe_matrix, SpecialLinearGroup.det_coe, coe'] - -lemma coe_one : UpperHalfPlane.ModularGroup.coe' 1 = 1 := by - simp only [coe', _root_.map_one] - -instance SLOnGLPos : SMul SL(2, ℤ) GL(2, ℝ)⁺ := - ⟨fun s g => s * g⟩ - -theorem SLOnGLPos_smul_apply (s : SL(2, ℤ)) (g : GL(2, ℝ)⁺) (z : ℍ) : - (s • g) • z = ((s : GL(2, ℝ)⁺) * g) • z := - rfl - -instance SL_to_GL_tower : IsScalarTower SL(2, ℤ) GL(2, ℝ)⁺ ℍ where - smul_assoc s g z := by - simp only [SLOnGLPos_smul_apply] - apply mul_smul' - -end ModularGroup - -end ModularScalarTowers - -- Porting note: in the statement, we used to have coercions `↑· : ℝ` -- rather than `algebraMap R ℝ ·`. theorem specialLinearGroup_apply {R : Type*} [CommRing R] [Algebra R ℝ] (g : SL(2, R)) (z : ℍ) : @@ -322,23 +281,23 @@ theorem specialLinearGroup_apply {R : Type*} [CommRing R] [Algebra R ℝ] (g : S (g • z).property := rfl +variable (g : GL(2, ℝ)⁺) (z : ℍ) + @[simp] -theorem coe_smul (g : GL(2, ℝ)⁺) (z : ℍ) : ↑(g • z) = num g z / denom g z := +theorem coe_smul : ↑(g • z) = num g z / denom g z := rfl @[simp] -theorem re_smul (g : GL(2, ℝ)⁺) (z : ℍ) : (g • z).re = (num g z / denom g z).re := +theorem re_smul : (g • z).re = (num g z / denom g z).re := rfl -theorem im_smul (g : GL(2, ℝ)⁺) (z : ℍ) : (g • z).im = (num g z / denom g z).im := +theorem im_smul : (g • z).im = (num g z / denom g z).im := rfl -theorem im_smul_eq_div_normSq (g : GL(2, ℝ)⁺) (z : ℍ) : - (g • z).im = det ↑ₘg * z.im / Complex.normSq (denom g z) := +theorem im_smul_eq_div_normSq : (g • z).im = det ↑ₘg * z.im / Complex.normSq (denom g z) := smulAux'_im g z -theorem c_mul_im_sq_le_normSq_denom (z : ℍ) (g : SL(2, ℝ)) : - (g 1 0 * z.im) ^ 2 ≤ Complex.normSq (denom g z) := by +theorem c_mul_im_sq_le_normSq_denom : (g 1 0 * z.im) ^ 2 ≤ Complex.normSq (denom g z) := by set c := g 1 0 set d := g 1 1 calc @@ -346,40 +305,15 @@ theorem c_mul_im_sq_le_normSq_denom (z : ℍ) (g : SL(2, ℝ)) : _ = Complex.normSq (denom g z) := by dsimp [c, d, denom, Complex.normSq]; ring @[simp] -theorem neg_smul (g : GL(2, ℝ)⁺) (z : ℍ) : -g • z = g • z := by +theorem neg_smul : -g • z = g • z := by ext1 change _ / _ = _ / _ field_simp [denom_ne_zero] simp only [num, denom, Complex.ofReal_neg, neg_mul, GLPos.coe_neg_GL, Units.val_neg, neg_apply] ring_nf -section SLModularAction - -namespace ModularGroup - -variable (g : SL(2, ℤ)) (z : ℍ) (Γ : Subgroup SL(2, ℤ)) - -@[simp] -theorem sl_moeb (A : SL(2, ℤ)) (z : ℍ) : A • z = (A : GL(2, ℝ)⁺) • z := - rfl - -@[simp high] -theorem SL_neg_smul (g : SL(2, ℤ)) (z : ℍ) : -g • z = g • z := by - simp only [coe_GLPos_neg, sl_moeb, coe_int_neg, neg_smul, coe'] - -nonrec theorem im_smul_eq_div_normSq : - (g • z).im = z.im / Complex.normSq (denom g z) := by - convert im_smul_eq_div_normSq g z - simp only [GeneralLinearGroup.val_det_apply, coe_GLPos_coe_GL_coe_matrix, - Int.coe_castRingHom, (g : SL(2, ℝ)).prop, one_mul, coe'] - -theorem denom_apply (g : SL(2, ℤ)) (z : ℍ) : - denom g z = (↑g : Matrix (Fin 2) (Fin 2) ℤ) 1 0 * z + (↑g : Matrix (Fin 2) (Fin 2) ℤ) 1 1 := by - simp [denom, coe'] - -end ModularGroup - -end SLModularAction +lemma denom_one : denom 1 z = 1 := by + simp [denom] section PosRealAction @@ -475,3 +409,75 @@ theorem exists_SL2_smul_eq_of_apply_zero_one_ne_zero (g : SL(2, ℝ)) (hc : g 1 linear_combination (-(z * (c : ℂ) ^ 2) - c * d) * h end UpperHalfPlane + +namespace ModularGroup -- results specific to `SL(2, ℤ)` + +section ModularScalarTowers + +/-- Canonical embedding of `SL(2, ℤ)` into `GL(2, ℝ)⁺`. -/ +@[coe] +def coe (g : SL(2, ℤ)) : GL(2, ℝ)⁺ := ((g : SL(2, ℝ)) : GL(2, ℝ)⁺) + +@[deprecated (since := "2024-11-19")] noncomputable alias coe' := coe + +instance : Coe SL(2, ℤ) GL(2, ℝ)⁺ := + ⟨coe⟩ + +@[simp] +theorem coe_apply_complex {g : SL(2, ℤ)} {i j : Fin 2} : + (Units.val <| Subtype.val <| coe g) i j = (Subtype.val g i j : ℂ) := + rfl + +@[deprecated (since := "2024-11-19")] alias coe'_apply_complex := coe_apply_complex + +@[simp] +theorem det_coe {g : SL(2, ℤ)} : det (Units.val <| Subtype.val <| coe g) = 1 := by + simp only [SpecialLinearGroup.coe_GLPos_coe_GL_coe_matrix, SpecialLinearGroup.det_coe, coe] + +@[deprecated (since := "2024-11-19")] alias det_coe' := det_coe + +lemma coe_one : coe 1 = 1 := by + simp only [coe, _root_.map_one] + +instance SLOnGLPos : SMul SL(2, ℤ) GL(2, ℝ)⁺ := + ⟨fun s g => s * g⟩ + +theorem SLOnGLPos_smul_apply (s : SL(2, ℤ)) (g : GL(2, ℝ)⁺) (z : ℍ) : + (s • g) • z = ((s : GL(2, ℝ)⁺) * g) • z := + rfl + +instance SL_to_GL_tower : IsScalarTower SL(2, ℤ) GL(2, ℝ)⁺ ℍ where + smul_assoc s g z := by + simp only [SLOnGLPos_smul_apply] + apply mul_smul' + +end ModularScalarTowers + +section SLModularAction + +variable (g : SL(2, ℤ)) (z : ℍ) + +@[simp] +theorem sl_moeb (A : SL(2, ℤ)) (z : ℍ) : A • z = (A : GL(2, ℝ)⁺) • z := + rfl + +@[simp high] +theorem SL_neg_smul (g : SL(2, ℤ)) (z : ℍ) : -g • z = g • z := by + simp only [coe_GLPos_neg, sl_moeb, coe_int_neg, neg_smul, coe] + +theorem im_smul_eq_div_normSq : (g • z).im = z.im / Complex.normSq (denom g z) := by + simpa only [coe, coe_GLPos_coe_GL_coe_matrix, (g : SL(2, ℝ)).prop, one_mul] using + z.im_smul_eq_div_normSq g + +theorem denom_apply (g : SL(2, ℤ)) (z : ℍ) : + denom g z = g 1 0 * z + g 1 1 := by + simp [denom, coe] + +@[simp] +lemma denom_S (z : ℍ) : denom S z = z := by + simp only [S, denom_apply, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, + cons_val_one, head_fin_const, Int.cast_one, one_mul, head_cons, Int.cast_zero, add_zero] + +end SLModularAction + +end ModularGroup diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 8f86113a1bed0..f4c849d0063d4 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -239,7 +239,7 @@ theorem smul_eq_lcRow0_add {p : Fin 2 → ℤ} (hp : IsCoprime (p 0) (p 1)) (hg (p 1 : ℂ) * z - p 0 = (p 1 * z - p 0) * ↑(Matrix.det (↑g : Matrix (Fin 2) (Fin 2) ℤ)))] rw [← hg, det_fin_two] simp only [Int.coe_castRingHom, coe_matrix_coe, Int.cast_mul, ofReal_intCast, map_apply, denom, - Int.cast_sub, coe_GLPos_coe_GL_coe_matrix, coe'_apply_complex] + Int.cast_sub, coe_GLPos_coe_GL_coe_matrix, coe_apply_complex] ring theorem tendsto_abs_re_smul {p : Fin 2 → ℤ} (hp : IsCoprime (p 0) (p 1)) : @@ -384,6 +384,11 @@ theorem three_lt_four_mul_im_sq_of_mem_fdo (h : z ∈ 𝒟ᵒ) : 3 < 4 * z.im ^ have := h.2 cases abs_cases z.re <;> nlinarith +/-- non-strict variant of `ModularGroup.three_le_four_mul_im_sq_of_mem_fdo` -/ +theorem three_le_four_mul_im_sq_of_mem_fd {τ : ℍ} (h : τ ∈ 𝒟) : 3 ≤ 4 * τ.im ^ 2 := by + have : 1 ≤ τ.re * τ.re + τ.im * τ.im := by simpa [Complex.normSq_apply] using h.1 + cases abs_cases τ.re <;> nlinarith [h.2] + /-- If `z ∈ 𝒟ᵒ`, and `n : ℤ`, then `|z + n| > 1`. -/ theorem one_lt_normSq_T_zpow_smul (hz : z ∈ 𝒟ᵒ) (n : ℤ) : 1 < normSq (T ^ n • z : ℍ) := by have hz₁ : 1 < z.re * z.re + z.im * z.im := hz.1 @@ -455,20 +460,18 @@ theorem abs_c_le_one (hz : z ∈ 𝒟ᵒ) (hg : g • z ∈ 𝒟ᵒ) : |g 1 0| intro hc replace hc : 0 < c ^ 4 := by change 0 < c ^ (2 * 2); rw [pow_mul]; apply sq_pos_of_pos (sq_pos_of_ne_zero hc) - have h₁ := - mul_lt_mul_of_pos_right + have h₁ := mul_lt_mul_of_pos_right (mul_lt_mul'' (three_lt_four_mul_im_sq_of_mem_fdo hg) (three_lt_four_mul_im_sq_of_mem_fdo hz) - (by linarith) (by linarith)) - hc + (by norm_num) (by norm_num)) hc have h₂ : (c * z.im) ^ 4 / normSq (denom (↑g) z) ^ 2 ≤ 1 := div_le_one_of_le₀ - (pow_four_le_pow_two_of_pow_two_le (UpperHalfPlane.c_mul_im_sq_le_normSq_denom z g)) + (pow_four_le_pow_two_of_pow_two_le (z.c_mul_im_sq_le_normSq_denom g)) (sq_nonneg _) let nsq := normSq (denom g z) calc 9 * c ^ 4 < c ^ 4 * z.im ^ 2 * (g • z).im ^ 2 * 16 := by linarith _ = c ^ 4 * z.im ^ 4 / nsq ^ 2 * 16 := by - rw [ModularGroup.im_smul_eq_div_normSq, div_pow] + rw [im_smul_eq_div_normSq, div_pow] ring _ ≤ 16 := by rw [← mul_pow]; linarith @@ -504,4 +507,21 @@ end UniqueRepresentative end FundamentalDomain +lemma exists_one_half_le_im_smul (τ : ℍ) : ∃ γ : SL(2, ℤ), 1 / 2 ≤ im (γ • τ) := by + obtain ⟨γ, hγ⟩ := exists_smul_mem_fd τ + use γ + nlinarith [three_le_four_mul_im_sq_of_mem_fd hγ, im_pos (γ • τ)] + +/-- For every `τ : ℍ` there is some `γ ∈ SL(2, ℤ)` that sends it to an element whose +imaginary part is at least `1/2` and such that `denom γ τ` has norm at most 1. -/ +lemma exists_one_half_le_im_smul_and_norm_denom_le (τ : ℍ) : + ∃ γ : SL(2, ℤ), 1 / 2 ≤ im (γ • τ) ∧ ‖denom γ τ‖ ≤ 1 := by + rcases le_total (1 / 2) τ.im with h | h + · exact ⟨1, (one_smul SL(2, ℤ) τ).symm ▸ h, by simp only [coe_one, denom_one, norm_one, le_refl]⟩ + · refine (exists_one_half_le_im_smul τ).imp (fun γ hγ ↦ ⟨hγ, ?_⟩) + have h1 : τ.im ≤ (γ • τ).im := h.trans hγ + rw [im_smul_eq_div_normSq, le_div_iff₀ (normSq_denom_pos (↑γ) τ), normSq_eq_norm_sq] at h1 + simpa only [norm_eq_abs, sq_le_one_iff_abs_le_one, Complex.abs_abs] using + (mul_le_iff_le_one_right τ.2).mp h1 + end ModularGroup diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean index 533e632c67595..890e52757badd 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean @@ -109,10 +109,10 @@ def _root_.eisensteinSeries (k : ℤ) (z : ℍ) : ℂ := ∑' x : gammaSet N a, lemma eisensteinSeries_slash_apply (k : ℤ) (γ : SL(2, ℤ)) : eisensteinSeries a k ∣[k] γ = eisensteinSeries (a ᵥ* γ) k := by ext1 z - simp_rw [SL_slash, slash_def, slash, ModularGroup.det_coe', ofReal_one, one_zpow, mul_one, + simp_rw [SL_slash, slash_def, slash, ModularGroup.det_coe, ofReal_one, one_zpow, mul_one, zpow_neg, mul_inv_eq_iff_eq_mul₀ (zpow_ne_zero _ <| z.denom_ne_zero _), mul_comm, eisensteinSeries, ← ModularGroup.sl_moeb, eisSummand_SL2_apply, tsum_mul_left] - erw [(gammaSetEquiv a γ).tsum_eq (eisSummand k · z)] + exact congr_arg (_ * ·) <| (gammaSetEquiv a γ).tsum_eq (eisSummand k · z) /-- The SlashInvariantForm defined by an Eisenstein series of weight `k : ℤ`, level `Γ(N)`, and congruence condition given by `a : Fin 2 → ZMod N`. -/ diff --git a/Mathlib/NumberTheory/ModularForms/LevelOne.lean b/Mathlib/NumberTheory/ModularForms/LevelOne.lean new file mode 100644 index 0000000000000..20ba662c6cc3b --- /dev/null +++ b/Mathlib/NumberTheory/ModularForms/LevelOne.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2024 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ +import Mathlib.NumberTheory.Modular +import Mathlib.NumberTheory.ModularForms.SlashInvariantForms +/-! +# Level one modular forms + +This file contains results specific to modular forms of level one, ie. modular forms for `SL(2, ℤ)`. + +TODO: Add finite-dimensionality of these spaces of modular forms. + +-/ + +open UpperHalfPlane ModularGroup SlashInvariantForm ModularForm Complex MatrixGroups + +lemma SlashInvariantForm.exists_one_half_le_im_and_norm_le {k : ℤ} (hk : k ≤ 0) {F : Type*} + [FunLike F ℍ ℂ] [SlashInvariantFormClass F ⊤ k] (f : F) (τ : ℍ) : + ∃ ξ : ℍ, 1/2 ≤ ξ.im ∧ ‖f τ‖ ≤ ‖f ξ‖ := + let ⟨γ, hγ, hdenom⟩ := exists_one_half_le_im_smul_and_norm_denom_le τ + ⟨γ • τ, hγ, by simpa only [slash_action_eqn'' _ (show γ ∈ ⊤ by tauto), norm_mul, norm_zpow] + using le_mul_of_one_le_left (norm_nonneg _) <| + one_le_zpow_of_nonpos₀ (norm_pos_iff.2 (denom_ne_zero _ _)) hdenom hk⟩ + +/-- If a constant function is modular of weight `k`, then either `k = 0`, or the constant is `0`. -/ +lemma SlashInvariantForm.wt_eq_zero_of_eq_const + {F : Type*} [FunLike F ℍ ℂ] (k : ℤ) [SlashInvariantFormClass F ⊤ k] + {f : F} {c : ℂ} (hf : ∀ τ, f τ = c) : k = 0 ∨ c = 0 := by + have hI := slash_action_eqn'' f (by tauto : ModularGroup.S ∈ ⊤) I + have h2I2 := slash_action_eqn'' f (by tauto : ModularGroup.S ∈ ⊤) ⟨2 * Complex.I, by simp⟩ + simp only [sl_moeb, hf, denom_S, coe_mk_subtype] at hI h2I2 + nth_rw 1 [h2I2] at hI + simp only [mul_zpow, coe_I, mul_eq_mul_right_iff, mul_left_eq_self₀] at hI + refine hI.imp_left (Or.casesOn · (fun H ↦ ?_) (False.elim ∘ zpow_ne_zero k I_ne_zero)) + rwa [← Complex.ofReal_ofNat, ← ofReal_zpow, ← ofReal_one, ofReal_inj, + zpow_eq_one_iff_right₀ (by norm_num) (by norm_num)] at H diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index 4b39fc4a7a847..bc52891b5dffc 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -139,7 +139,7 @@ theorem SL_slash (γ : SL(2, ℤ)) : f ∣[k] γ = f ∣[k] (γ : GL(2, ℝ)⁺) theorem is_invariant_const (A : SL(2, ℤ)) (x : ℂ) : Function.const ℍ x ∣[(0 : ℤ)] A = Function.const ℍ x := by funext - simp only [SL_slash, slash_def, slash, Function.const_apply, det_coe', ofReal_one, zero_sub, + simp only [SL_slash, slash_def, slash, Function.const_apply, det_coe, ofReal_one, zero_sub, zpow_neg, zpow_one, inv_one, mul_one, neg_zero, zpow_zero] /-- The constant function 1 is invariant under any element of `SL(2, ℤ)`. -/ @@ -155,7 +155,7 @@ theorem slash_action_eq'_iff (k : ℤ) (f : ℍ → ℂ) (γ : SL(2, ℤ)) (z : simp only [SL_slash, slash_def, ModularForm.slash] convert inv_mul_eq_iff_eq_mul₀ (G₀ := ℂ) _ using 2 · rw [mul_comm] - simp only [denom, zpow_neg, det_coe', ofReal_one, one_zpow, mul_one, + simp only [denom, zpow_neg, det_coe, ofReal_one, one_zpow, mul_one, sl_moeb] rfl · convert zpow_ne_zero k (denom_ne_zero γ z) @@ -184,7 +184,7 @@ theorem mul_slash_SL2 (k1 k2 : ℤ) (A : SL(2, ℤ)) (f g : ℍ → ℂ) : (f * g) ∣[k1 + k2] (A : GL(2, ℝ)⁺) = ((↑ₘA).det : ℝ) • f ∣[k1] A * g ∣[k2] A := by apply mul_slash - _ = (1 : ℝ) • f ∣[k1] A * g ∣[k2] A := by rw [det_coe'] + _ = (1 : ℝ) • f ∣[k1] A * g ∣[k2] A := by rw [det_coe] _ = f ∣[k1] A * g ∣[k2] A := by rw [one_smul] end From e69c3062daeb950c9dee862d68daa01c08283884 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 10:54:05 +0000 Subject: [PATCH 53/90] chore: deprecate `LinearOrderedCommGroupWithZero` lemmas (#19197) Deprecate almost all lemmas about `LinearOrderedCommGroupWithZero` and remove three months old deprecations. --- Mathlib/Algebra/Order/Archimedean/Basic.lean | 1 + .../Order/GroupWithZero/Canonical.lean | 99 ++++++------------- .../Order/GroupWithZero/Unbundled.lean | 1 - .../Order/GroupWithZero/Unbundled/Lemmas.lean | 32 +++++- Mathlib/Algebra/Order/Hom/Monoid.lean | 1 + Mathlib/RingTheory/LaurentSeries.lean | 3 +- Mathlib/RingTheory/Perfection.lean | 2 +- Mathlib/RingTheory/Valuation/Integral.lean | 2 +- .../Valuation/ValuationSubring.lean | 3 +- .../Algebra/Valued/ValuationTopology.lean | 3 +- .../Topology/Algebra/WithZeroTopology.lean | 2 +- 11 files changed, 70 insertions(+), 79 deletions(-) diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index 2f526eb1aa929..7756989822b78 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.Ring.Pow import Mathlib.Data.Int.LeastGreatest import Mathlib.Data.Rat.Floor diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index e0e51858c4b0b..95813028bae02 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -7,9 +7,8 @@ import Mathlib.Algebra.GroupWithZero.InjSurj import Mathlib.Algebra.GroupWithZero.Units.Equiv import Mathlib.Algebra.GroupWithZero.WithZero import Mathlib.Algebra.Order.AddGroupWithTop -import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.GroupWithZero.Synonym -import Mathlib.Algebra.Order.GroupWithZero.Unbundled +import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas import Mathlib.Algebra.Order.Monoid.Basic import Mathlib.Algebra.Order.Monoid.OrderDual import Mathlib.Algebra.Order.Monoid.TypeTags @@ -130,96 +129,63 @@ instance (priority := 100) LinearOrderedCommGroupWithZero.toMulPosStrictMono : MulPosStrictMono α where elim a b c hbc := by by_contra! h; exact hbc.not_le <| (mul_le_mul_right a.2).1 h -/-- Alias of `one_le_mul'` for unification. -/ -@[deprecated one_le_mul (since := "2024-08-21")] -theorem one_le_mul₀ (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := - one_le_mul ha hb - -@[deprecated mul_le_mul_right (since := "2024-08-21")] -theorem le_of_le_mul_right (h : c ≠ 0) (hab : a * c ≤ b * c) : a ≤ b := - (mul_le_mul_right (zero_lt_iff.2 h)).1 hab - -@[deprecated le_mul_inv_iff₀ (since := "2024-08-21")] -theorem le_mul_inv_of_mul_le (h : c ≠ 0) (hab : a * c ≤ b) : a ≤ b * c⁻¹ := - (le_mul_inv_iff₀ (zero_lt_iff.2 h)).2 hab - -theorem mul_inv_le_of_le_mul (hab : a ≤ b * c) : a * c⁻¹ ≤ b := by - by_cases h : c = 0 - · simp [h] - · exact (mul_le_mul_right (zero_lt_iff.2 h)).1 (by simpa [h] using hab) +@[deprecated mul_inv_le_of_le_mul₀ (since := "2024-11-18")] +theorem mul_inv_le_of_le_mul (hab : a ≤ b * c) : a * c⁻¹ ≤ b := + mul_inv_le_of_le_mul₀ zero_le' zero_le' hab @[simp] theorem Units.zero_lt (u : αˣ) : (0 : α) < u := - zero_lt_iff.2 <| u.ne_zero + zero_lt_iff.2 u.ne_zero +@[deprecated mul_lt_mul_of_le_of_lt_of_nonneg_of_pos (since := "2024-11-18")] theorem mul_lt_mul_of_lt_of_le₀ (hab : a ≤ b) (hb : b ≠ 0) (hcd : c < d) : a * c < b * d := - have hd : d ≠ 0 := ne_zero_of_lt hcd - if ha : a = 0 then by - rw [ha, zero_mul, zero_lt_iff] - exact mul_ne_zero hb hd - else - if hc : c = 0 then by - rw [hc, mul_zero, zero_lt_iff] - exact mul_ne_zero hb hd - else - show Units.mk0 a ha * Units.mk0 c hc < Units.mk0 b hb * Units.mk0 d hd from - mul_lt_mul_of_le_of_lt hab hcd + mul_lt_mul_of_le_of_lt_of_nonneg_of_pos hab hcd zero_le' (zero_lt_iff.2 hb) +@[deprecated mul_lt_mul'' (since := "2024-11-18")] theorem mul_lt_mul₀ (hab : a < b) (hcd : c < d) : a * c < b * d := - mul_lt_mul_of_lt_of_le₀ hab.le (ne_zero_of_lt hab) hcd + mul_lt_mul'' hab hcd zero_le' zero_le' theorem mul_inv_lt_of_lt_mul₀ (h : a < b * c) : a * c⁻¹ < b := by contrapose! h - simpa only [inv_inv] using mul_inv_le_of_le_mul h + simpa only [inv_inv] using mul_inv_le_of_le_mul₀ zero_le' zero_le' h theorem inv_mul_lt_of_lt_mul₀ (h : a < b * c) : b⁻¹ * a < c := by rw [mul_comm] at * exact mul_inv_lt_of_lt_mul₀ h -@[deprecated mul_lt_mul_of_pos_right (since := "2024-08-21")] -theorem mul_lt_right₀ (c : α) (h : a < b) (hc : c ≠ 0) : a * c < b * c := - mul_lt_mul_of_pos_right h (zero_lt_iff.2 hc) - theorem lt_of_mul_lt_mul_of_le₀ (h : a * b < c * d) (hc : 0 < c) (hh : c ≤ a) : b < d := by have ha : a ≠ 0 := ne_of_gt (lt_of_lt_of_le hc hh) rw [← inv_le_inv₀ (zero_lt_iff.2 ha) hc] at hh - have := mul_lt_mul_of_lt_of_le₀ hh (inv_ne_zero (ne_of_gt hc)) h - simpa [inv_mul_cancel_left₀ ha, inv_mul_cancel_left₀ (ne_of_gt hc)] using this + simpa [inv_mul_cancel_left₀ ha, inv_mul_cancel_left₀ hc.ne'] + using mul_lt_mul_of_le_of_lt_of_nonneg_of_pos hh h zero_le' (inv_pos.2 hc) -@[deprecated mul_le_mul_right (since := "2024-08-21")] -theorem mul_le_mul_right₀ (hc : c ≠ 0) : a * c ≤ b * c ↔ a ≤ b := - mul_le_mul_right (zero_lt_iff.2 hc) +@[deprecated div_le_div_iff_of_pos_right (since := "2024-11-18")] +theorem div_le_div_right₀ (hc : c ≠ 0) : a / c ≤ b / c ↔ a ≤ b := + div_le_div_iff_of_pos_right (zero_lt_iff.2 hc) -@[deprecated mul_le_mul_left (since := "2024-08-21")] -theorem mul_le_mul_left₀ (ha : a ≠ 0) : a * b ≤ a * c ↔ b ≤ c := - mul_le_mul_left (zero_lt_iff.2 ha) +@[deprecated div_le_div_iff_of_pos_left (since := "2024-11-18")] +theorem div_le_div_left₀ (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) : a / b ≤ a / c ↔ c ≤ b := + div_le_div_iff_of_pos_left (zero_lt_iff.2 ha) (zero_lt_iff.2 hb) (zero_lt_iff.2 hc) -theorem div_le_div_right₀ (hc : c ≠ 0) : a / c ≤ b / c ↔ a ≤ b := by - rw [div_eq_mul_inv, div_eq_mul_inv, mul_le_mul_right (zero_lt_iff.2 (inv_ne_zero hc))] - -theorem div_le_div_left₀ (ha : a ≠ 0) (hb : b ≠ 0) (hc : c ≠ 0) : a / b ≤ a / c ↔ c ≤ b := by - simp only [div_eq_mul_inv, mul_le_mul_left (zero_lt_iff.2 ha), - inv_le_inv₀ (zero_lt_iff.2 hb) (zero_lt_iff.2 hc)] - -/-- `Equiv.mulLeft₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. - -Note that `OrderIso.mulLeft₀` refers to the `LinearOrderedField` version. -/ -@[simps! (config := { simpRhs := true }) apply toEquiv] -def OrderIso.mulLeft₀' {a : α} (ha : a ≠ 0) : α ≃o α := - { Equiv.mulLeft₀ a ha with map_rel_iff' := mul_le_mul_left (zero_lt_iff.2 ha) } +/-- `Equiv.mulLeft₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. -/ +@[simps! (config := { simpRhs := true }) apply toEquiv, +deprecated OrderIso.mulLeft₀ (since := "2024-11-18")] +def OrderIso.mulLeft₀' {a : α} (ha : a ≠ 0) : α ≃o α := .mulLeft₀ a (zero_lt_iff.2 ha) +set_option linter.deprecated false in +@[deprecated OrderIso.mulLeft₀_symm (since := "2024-11-18")] theorem OrderIso.mulLeft₀'_symm {a : α} (ha : a ≠ 0) : (OrderIso.mulLeft₀' ha).symm = OrderIso.mulLeft₀' (inv_ne_zero ha) := by ext rfl -/-- `Equiv.mulRight₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. - -Note that `OrderIso.mulRight₀` refers to the `LinearOrderedField` version. -/ -@[simps! (config := { simpRhs := true }) apply toEquiv] -def OrderIso.mulRight₀' {a : α} (ha : a ≠ 0) : α ≃o α := - { Equiv.mulRight₀ a ha with map_rel_iff' := mul_le_mul_right (zero_lt_iff.2 ha) } +/-- `Equiv.mulRight₀` as an `OrderIso` on a `LinearOrderedCommGroupWithZero.`. -/ +@[simps! (config := { simpRhs := true }) apply toEquiv, +deprecated OrderIso.mulRight₀ (since := "2024-11-18")] +def OrderIso.mulRight₀' {a : α} (ha : a ≠ 0) : α ≃o α := .mulRight₀ a (zero_lt_iff.2 ha) +set_option linter.deprecated false in +@[deprecated OrderIso.mulRight₀_symm (since := "2024-11-18")] theorem OrderIso.mulRight₀'_symm {a : α} (ha : a ≠ 0) : (OrderIso.mulRight₀' ha).symm = OrderIso.mulRight₀' (inv_ne_zero ha) := by ext @@ -231,9 +197,8 @@ instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := neg_top := inv_zero (G₀ := α) add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : Additive.toMul a ≠ 0) } -lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := by - rw [← one_mul (a ^ n), pow_succ'] - exact mul_lt_mul_of_pos_right ha (pow_pos (zero_lt_one.trans ha) _) +@[deprecated pow_lt_pow_right₀ (since := "2024-11-18")] +lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := pow_lt_pow_right₀ ha n.lt_succ_self end LinearOrderedCommGroupWithZero diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index 431015c2b62e0..bbd7e7951179d 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -1724,7 +1724,6 @@ lemma div_lt_div₀ (hac : a < c) (hdb : d ≤ b) (hc : 0 ≤ c) (hd : 0 < d) : rw [div_eq_mul_inv, div_eq_mul_inv] exact mul_lt_mul hac ((inv_le_inv₀ (hd.trans_le hdb) hd).2 hdb) (inv_pos.2 <| hd.trans_le hdb) hc -/-- See -/ lemma div_lt_div₀' (hac : a ≤ c) (hdb : d < b) (hc : 0 < c) (hd : 0 < d) : a / b < c / d := by rw [div_eq_mul_inv, div_eq_mul_inv] exact mul_lt_mul' hac ((inv_lt_inv₀ (hd.trans hdb) hd).2 hdb) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean index 15331ca9c673a..daa1991832ae7 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled/Lemmas.lean @@ -12,26 +12,48 @@ import Mathlib.Order.Hom.Basic # Multiplication by a positive element as an order isomorphism -/ -variable {G₀} [GroupWithZero G₀] [Preorder G₀] {a : G₀} +namespace OrderIso +variable {G₀} [GroupWithZero G₀] [PartialOrder G₀] + +section left +variable [PosMulMono G₀] [PosMulReflectLE G₀] /-- `Equiv.mulLeft₀` as an order isomorphism. -/ @[simps! (config := { simpRhs := true })] -def OrderIso.mulLeft₀ [PosMulMono G₀] [PosMulReflectLE G₀] (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where +def mulLeft₀ (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where toEquiv := .mulLeft₀ a ha.ne' map_rel_iff' := mul_le_mul_left ha +variable [ZeroLEOneClass G₀] + +lemma mulLeft₀_symm (a : G₀) (ha : 0 < a) : (mulLeft₀ a ha).symm = mulLeft₀ a⁻¹ (inv_pos.2 ha) := by + ext; rfl + +end left + +section right +variable [MulPosMono G₀] [MulPosReflectLE G₀] + /-- `Equiv.mulRight₀` as an order isomorphism. -/ @[simps! (config := { simpRhs := true })] -def OrderIso.mulRight₀ [MulPosMono G₀] [MulPosReflectLE G₀] (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where +def mulRight₀ (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where toEquiv := .mulRight₀ a ha.ne' map_rel_iff' := mul_le_mul_right ha +variable [ZeroLEOneClass G₀] [PosMulReflectLT G₀] + +lemma mulRight₀_symm (a : G₀) (ha : 0 < a) : + (mulRight₀ a ha).symm = mulRight₀ a⁻¹ (inv_pos.2 ha) := by ext; rfl + +end right + /-- `Equiv.divRight₀` as an order isomorphism. -/ @[simps! (config := { simpRhs := true })] -def OrderIso.divRight₀ {G₀} [GroupWithZero G₀] [PartialOrder G₀] [ZeroLEOneClass G₀] - [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] +def divRight₀ [ZeroLEOneClass G₀] [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] (a : G₀) (ha : 0 < a) : G₀ ≃o G₀ where toEquiv := .divRight₀ a ha.ne' map_rel_iff' {b c} := by simp only [Equiv.divRight₀_apply, div_eq_mul_inv] exact mul_le_mul_right (a := a⁻¹) (inv_pos.mpr ha) + +end OrderIso diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index 5fa19423317a2..b40dacaccb13e 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Order.Group.Instances import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.Monoid.Units import Mathlib.Order.Hom.Basic /-! diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index 2d9c6f7c2f02e..d543b3fa95654 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -911,7 +911,8 @@ theorem exists_Polynomial_intValuation_lt (F : K⟦X⟧) (η : ℤₘ₀ˣ) : rw [← valuation_of_algebraMap (K := K⸨X⸩) (PowerSeries.idealX K) (F - F.trunc (d + 1))] apply lt_of_le_of_lt this rw [← mul_one (η : ℤₘ₀), mul_assoc, one_mul] - apply mul_lt_mul_of_lt_of_le₀ (le_refl _) η.ne_zero + gcongr + · exact zero_lt_iff.2 η.ne_zero rw [← WithZero.coe_one, coe_lt_coe, ofAdd_neg, Right.inv_lt_one_iff, ← ofAdd_zero, Multiplicative.ofAdd_lt] exact Int.zero_lt_one diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 013131970238e..7db1bdda6d237 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -424,7 +424,7 @@ theorem mul_ne_zero_of_pow_p_ne_zero {x y : ModP K v O hv p} (hx : x ^ p ≠ 0) rw [← v_p_lt_val hv] at hx hy ⊢ rw [RingHom.map_pow, v.map_pow, ← rpow_lt_rpow_iff h1p, ← rpow_natCast, ← rpow_mul, mul_one_div_cancel (Nat.cast_ne_zero.2 hp.1.ne_zero : (p : ℝ) ≠ 0), rpow_one] at hx hy - rw [RingHom.map_mul, v.map_mul]; refine lt_of_le_of_lt ?_ (mul_lt_mul₀ hx hy) + rw [RingHom.map_mul, v.map_mul]; refine lt_of_le_of_lt ?_ (mul_lt_mul'' hx hy zero_le' zero_le') by_cases hvp : v p = 0 · rw [hvp]; exact zero_le _ replace hvp := zero_lt_iff.2 hvp diff --git a/Mathlib/RingTheory/Valuation/Integral.lean b/Mathlib/RingTheory/Valuation/Integral.lean index c8e1a11e3d970..515de9eca405b 100644 --- a/Mathlib/RingTheory/Valuation/Integral.lean +++ b/Mathlib/RingTheory/Valuation/Integral.lean @@ -37,7 +37,7 @@ theorem mem_of_integral {x : R} (hx : IsIntegral O x) : x ∈ v.integer := rw [eval₂_mul, eval₂_pow, eval₂_C, eval₂_X, v.map_mul, v.map_pow, ← one_mul (v x ^ p.natDegree)] cases' (hv.2 <| p.coeff i).lt_or_eq with hvpi hvpi - · exact mul_lt_mul₀ hvpi (pow_lt_pow_right₀ hvx <| Finset.mem_range.1 hi) + · exact mul_lt_mul'' hvpi (pow_lt_pow_right₀ hvx <| Finset.mem_range.1 hi) zero_le' zero_le' · rw [hvpi, one_mul, one_mul]; exact pow_lt_pow_right₀ hvx (Finset.mem_range.1 hi) protected theorem integralClosure : integralClosure O R = ⊥ := diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index a5ea654a7ca32..867c268d02ac5 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -480,7 +480,8 @@ section nonunits def nonunits : Subsemigroup K where carrier := {x | A.valuation x < 1} -- Porting note: added `Set.mem_setOf.mp` - mul_mem' ha hb := (mul_lt_mul₀ (Set.mem_setOf.mp ha) (Set.mem_setOf.mp hb)).trans_eq <| mul_one _ + mul_mem' ha hb := (mul_lt_mul'' (Set.mem_setOf.mp ha) (Set.mem_setOf.mp hb) + zero_le' zero_le').trans_eq <| mul_one _ theorem mem_nonunits_iff {x : K} : x ∈ A.nonunits ↔ A.valuation x < 1 := Iff.rfl diff --git a/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean b/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean index 8cdff99b6d7fe..621c4dd2159b2 100644 --- a/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean +++ b/Mathlib/Topology/Algebra/Valued/ValuationTopology.lean @@ -42,9 +42,10 @@ theorem subgroups_basis : RingSubgroupsBasis fun γ : Γ₀ˣ => (v.ltAddSubgrou cases' exists_square_le γ with γ₀ h use γ₀ rintro - ⟨r, r_in, s, s_in, rfl⟩ + simp only [ltAddSubgroup, AddSubgroup.coe_set_mk, mem_setOf_eq] at r_in s_in calc (v (r * s) : Γ₀) = v r * v s := Valuation.map_mul _ _ _ - _ < γ₀ * γ₀ := mul_lt_mul₀ r_in s_in + _ < γ₀ * γ₀ := by gcongr <;> exact zero_le' _ ≤ γ := mod_cast h leftMul := by rintro x γ diff --git a/Mathlib/Topology/Algebra/WithZeroTopology.lean b/Mathlib/Topology/Algebra/WithZeroTopology.lean index c115bccac2248..cd47b1c9d0f8f 100644 --- a/Mathlib/Topology/Algebra/WithZeroTopology.lean +++ b/Mathlib/Topology/Algebra/WithZeroTopology.lean @@ -166,7 +166,7 @@ scoped instance (priority := 100) : ContinuousMul Γ₀ where refine ((hasBasis_nhds_zero.prod_nhds hasBasis_nhds_zero).tendsto_iff hasBasis_nhds_zero).2 fun γ hγ => ⟨(γ, 1), ⟨hγ, one_ne_zero⟩, ?_⟩ rintro ⟨x, y⟩ ⟨hx : x < γ, hy : y < 1⟩ - exact (mul_lt_mul₀ hx hy).trans_eq (mul_one γ) + exact (mul_lt_mul'' hx hy zero_le' zero_le').trans_eq (mul_one γ) · rw [zero_mul, nhds_prod_eq, nhds_of_ne_zero hy, prod_pure, tendsto_map'_iff] refine (hasBasis_nhds_zero.tendsto_iff hasBasis_nhds_zero).2 fun γ hγ => ?_ refine ⟨γ / y, div_ne_zero hγ hy, fun x hx => ?_⟩ From 7d4e3bb74bb8562fde3e5bf3f0157f0758e91832 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 20 Nov 2024 11:42:56 +0000 Subject: [PATCH 54/90] chore(discover-lean-pr-testing): properly fetch tags (#19204) Instead of guessing the dates of the latest two tags in the `lean4-nightly` repo, we list the available remote tags and take the last two. We then fetch them with a shallow depth, to avoid cloning the entire repo. After that we parse the commit log between these two tags. --- .../workflows/discover-lean-pr-testing.yml | 33 +++++++++---------- 1 file changed, 15 insertions(+), 18 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 9cc2434aea615..3c414d9e99689 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -24,19 +24,7 @@ jobs: git config --global user.name "github-actions" git config --global user.email "github-actions@github.com" - - name: Read date from lean-toolchain - id: fetch-date - run: | - TODAY=$(grep -oP 'nightly-\K\d{4}-\d{2}-\d{2}' lean-toolchain) - echo "TODAY=$TODAY" >> "$GITHUB_ENV" - - - name: Set YESTERDAY - id: set-yesterday - run: | - YESTERDAY=$(date -d "$TODAY -1 day" +%Y-%m-%d) - echo "YESTERDAY=$YESTERDAY" >> "$GITHUB_ENV" - - - name: Clone lean4-nightly repository and get PRs + - name: Clone lean4 repository and get PRs id: get-prs run: | NIGHTLY_URL="https://github.com/leanprover/lean4-nightly.git" @@ -44,14 +32,22 @@ jobs: # Create a temporary directory for cloning cd "$(mktemp -d)" || exit 1 - # Clone the repository with a depth of 30 (adjust as needed) - git clone --depth 30 "$NIGHTLY_URL" + # Clone the repository with a depth of 1 + git clone --depth 1 "$NIGHTLY_URL" # Navigate to the cloned repository cd lean4-nightly || exit 1 - YESTERDAY="${{ steps.set-yesterday.outputs.yesterday }}" - TODAY="${{ steps.fetch-date.outputs.today }}" - PRS=$(git log --oneline "nightly-$YESTERDAY..nightly-$TODAY" | sed 's/.*(#\([0-9]\+\))$/\1/') + + # Shallow fetch of the last two tags + LAST_TWO_TAGS=$(git ls-remote --tags | grep nightly | tail -2 | cut -f2) + OLD=$(echo "$LAST_TWO_TAGS" | head -1) + NEW=$(echo "$LAST_TWO_TAGS" | tail -1) + git fetch --depth=1 origin tag "$OLD" --no-tags + # Fetch the latest 100 commits prior to the $NEW tag. + # This should cover the terrain between $OLD and $NEW while still being a speedy download. + git fetch --depth=100 origin tag "$NEW" --no-tags + + PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/') # Output the PRs echo "$PRS" @@ -64,6 +60,7 @@ jobs: PRS="${{ steps.get-prs.outputs.prs }}" echo "$PRS" | tr ' ' '\n' > prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) + echo "$MATCHING_BRANCHES" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" From 7600042bca03f6a9066a2607026c72cfc6c034a5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 11:42:58 +0000 Subject: [PATCH 55/90] feat: `IsScalarTower` and `SMulCommClass` instances for pointwise actions of submodules (#19214) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Moves: - `Submodule.smul_def` → `Submodule.setSemiring_smul_def` From GrowthInGroups Co-authored-by: Eric Wieser --- Mathlib/Algebra/Algebra/Operations.lean | 28 ++++++++++++++++--- .../Algebra/Module/Submodule/Pointwise.lean | 2 ++ 2 files changed, 26 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 8decae0eb7f39..de1a4621d5777 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -285,12 +285,14 @@ theorem mul_eq_map₂ : M * N = map₂ (LinearMap.mul R A) M N := le_antisymm (mul_le.mpr fun _m hm _n ↦ apply_mem_map₂ _ hm) (map₂_le.mpr fun _m hm _n ↦ mul_mem_mul hm) -variable (R) +variable (R M N) theorem span_mul_span : span R S * span R T = span R (S * T) := by rw [mul_eq_map₂]; apply map₂_span_span -variable {R} (M N P Q) +lemma mul_def : M * N = span R (M * N : Set A) := by simp [← span_mul_span] + +variable {R} (P Q) protected theorem mul_one : M * 1 = M := by conv_lhs => rw [one_eq_span, ← span_eq M] @@ -353,6 +355,24 @@ theorem comap_op_mul (M N : Submodule R Aᵐᵒᵖ) : comap (↑(opLinearEquiv R : A ≃ₗ[R] Aᵐᵒᵖ) : A →ₗ[R] Aᵐᵒᵖ) M := by simp_rw [comap_equiv_eq_map_symm, map_unop_mul] +section +variable {α : Type*} [Monoid α] [DistribMulAction α A] [SMulCommClass α R A] + +instance [IsScalarTower α A A] : IsScalarTower α (Submodule R A) (Submodule R A) where + smul_assoc a S T := by + rw [← S.span_eq, ← T.span_eq, smul_span, smul_eq_mul, smul_eq_mul, span_mul_span, span_mul_span, + smul_span, smul_mul_assoc] + +instance [SMulCommClass α A A] : SMulCommClass α (Submodule R A) (Submodule R A) where + smul_comm a S T := by + rw [← S.span_eq, ← T.span_eq, smul_span, smul_eq_mul, smul_eq_mul, span_mul_span, span_mul_span, + smul_span, mul_smul_comm] + +instance [SMulCommClass A α A] : SMulCommClass (Submodule R A) α (Submodule R A) := + have := SMulCommClass.symm A α A; .symm .. + +end + section open Pointwise @@ -635,7 +655,7 @@ instance moduleSet : Module (SetSemiring A) (Submodule R A) where variable {R A} -theorem smul_def (s : SetSemiring A) (P : Submodule R A) : +theorem setSemiring_smul_def (s : SetSemiring A) (P : Submodule R A) : s • P = span R (SetSemiring.down (α := A) s) * P := rfl @@ -647,7 +667,7 @@ theorem smul_le_smul {s t : SetSemiring A} {M N : Submodule R A} theorem singleton_smul (a : A) (M : Submodule R A) : Set.up ({a} : Set A) • M = M.map (LinearMap.mulLeft R a) := by conv_lhs => rw [← span_eq M] - rw [smul_def, SetSemiring.down_up, span_mul_span, singleton_mul] + rw [setSemiring_smul_def, SetSemiring.down_up, span_mul_span, singleton_mul] exact (map (LinearMap.mulLeft R a) M).span_eq section Quotient diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index 3e0acab091ae2..ec25efdc44f71 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -228,6 +228,8 @@ theorem smul_sup' (a : α) (S T : Submodule R M) : a • (S ⊔ T) = a • S ⊔ theorem smul_span (a : α) (s : Set M) : a • span R s = span R (a • s) := map_span _ _ +lemma smul_def (a : α) (S : Submodule R M) : a • S = span R (a • S : Set M) := by simp [← smul_span] + theorem span_smul (a : α) (s : Set M) : span R (a • s) = a • span R s := Eq.symm (span_image _).symm From 333a8fa711db4e6b0add704bc4b4bd55656a3f44 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 20 Nov 2024 11:42:59 +0000 Subject: [PATCH 56/90] feat: `Nat.cast_nonpos` (#19247) For simp confluence From GrowthInGroups --- Mathlib/Data/Int/Defs.lean | 3 ++- Mathlib/Data/Nat/Cast/Order/Basic.lean | 2 ++ 2 files changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 6bf06c9ff023a..9fd2c581ce8e9 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -114,7 +114,8 @@ lemma natCast_ne_zero_iff_pos {n : ℕ} : (n : ℤ) ≠ 0 ↔ 0 < n := by omega lemma natCast_succ_pos (n : ℕ) : 0 < (n.succ : ℤ) := natCast_pos.2 n.succ_pos -@[simp] lemma natCast_nonpos_iff {n : ℕ} : (n : ℤ) ≤ 0 ↔ n = 0 := by omega +-- We want to use this lemma earlier than the lemmas simp can prove it with +@[simp, nolint simpNF] lemma natCast_nonpos_iff {n : ℕ} : (n : ℤ) ≤ 0 ↔ n = 0 := by omega lemma natCast_nonneg (n : ℕ) : 0 ≤ (n : ℤ) := ofNat_le.2 (Nat.zero_le _) diff --git a/Mathlib/Data/Nat/Cast/Order/Basic.lean b/Mathlib/Data/Nat/Cast/Order/Basic.lean index 21129456cd441..eec03f2914acb 100644 --- a/Mathlib/Data/Nat/Cast/Order/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Order/Basic.lean @@ -98,6 +98,8 @@ theorem cast_lt_one : (n : α) < 1 ↔ n = 0 := by @[simp, norm_cast] theorem cast_le_one : (n : α) ≤ 1 ↔ n ≤ 1 := by rw [← cast_one, cast_le] +@[simp] lemma cast_nonpos : (n : α) ≤ 0 ↔ n = 0 := by norm_cast; omega + section variable [m.AtLeastTwo] From b6fb044ac0aeee41e114dea81dfa74126bde3aba Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 20 Nov 2024 11:43:00 +0000 Subject: [PATCH 57/90] chore(RingTheory): split `UniqueFactorizationDomain.lean` (#19256) This PR splits the big file `UniqueFactorizationDomain.lean` into many smaller files: * `UniqueFactorizationDomain.Defs` contains the definition of `WfDvdMonoid`, `UniqueFactorizationMonoid` and `factors` * `UniqueFactorizationDomain.Basic` contains basic results on the previous definition such as the equivalence between `UniqueFactorizationMonoid` and the existence of prime factors * `UniqueFactorizationDomain.NormalizedFactors` defines `UniqueFactorizationMonoid.normalizedFactors` and contains basic results. (Or should we merge it into `Defs`/`Basic`?) * `UniqueFactorizationDomain.FactorSet` defines `Associates.FactorSet` and `Associates.factors` * `UniqueFactorizationDomain.Finite` proves that elements of a UFM with finitely many units have finitely many divisors. (This could be moved to `Basic.lean` but I find it niche enough to deserve its own file.) * `UniqueFactorizationDomain.Finsupp` defines `UniqueFactorizationMonoid.factorization` as a `Finsupp` version of `factors` * `UniqueFactorizationDomain.GCDMonoid` defines a GCD on a UFM * `UniqueFactorizationDomain.Ideal` shows UFDs satisfy the ascending chain condition on ideals * `UniqueFactorizationDomain.Multiplicative` contains results on multiplicative properties and functions. * `UniqueFactorizationDomain.Multiplicity` relates `factors.count` and `multiplicity` * `UniqueFactorizationDomain.Nat` contains the `UniqueFactorizationMonoid Nat` instance --- Mathlib.lean | 12 +- Mathlib/Algebra/Squarefree/Basic.lean | 3 +- .../Additive/ErdosGinzburgZiv.lean | 1 - Mathlib/Data/Nat/Squarefree.lean | 1 + Mathlib/NumberTheory/ArithmeticFunction.lean | 4 +- Mathlib/NumberTheory/FLT/Four.lean | 1 + Mathlib/RingTheory/ChainOfDivisors.lean | 3 +- .../DedekindDomain/Factorization.lean | 1 + .../DiscreteValuationRing/Basic.lean | 1 + Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean | 1 + .../RingTheory/Localization/Away/Basic.lean | 2 +- Mathlib/RingTheory/Localization/NumDen.lean | 2 +- .../Noetherian/UniqueFactorizationDomain.lean | 2 +- Mathlib/RingTheory/Polynomial/Radical.lean | 1 + .../Polynomial/UniqueFactorization.lean | 3 + Mathlib/RingTheory/PowerSeries/Inverse.lean | 1 + Mathlib/RingTheory/PrincipalIdealDomain.lean | 1 + Mathlib/RingTheory/Radical.lean | 3 +- Mathlib/RingTheory/RootsOfUnity/Minpoly.lean | 5 +- .../RingTheory/UniqueFactorizationDomain.lean | 2017 ----------------- .../UniqueFactorizationDomain/Basic.lean | 473 ++++ .../UniqueFactorizationDomain/Defs.lean | 205 ++ .../UniqueFactorizationDomain/FactorSet.lean | 647 ++++++ .../UniqueFactorizationDomain/Finite.lean | 53 + .../UniqueFactorizationDomain/Finsupp.lean | 62 + .../UniqueFactorizationDomain/GCDMonoid.lean | 73 + .../UniqueFactorizationDomain/Ideal.lean | 58 + .../Multiplicative.lean | 163 ++ .../Multiplicity.lean | 126 + .../UniqueFactorizationDomain/Nat.lean | 62 + .../NormalizedFactors.lean | 329 +++ MathlibTest/Variable.lean | 8 +- 32 files changed, 2292 insertions(+), 2032 deletions(-) delete mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean create mode 100644 Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean diff --git a/Mathlib.lean b/Mathlib.lean index f0d237abd75a7..0426236be0a65 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4467,7 +4467,17 @@ import Mathlib.RingTheory.TwoSidedIdeal.Instances import Mathlib.RingTheory.TwoSidedIdeal.Kernel import Mathlib.RingTheory.TwoSidedIdeal.Lattice import Mathlib.RingTheory.TwoSidedIdeal.Operations -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Defs +import Mathlib.RingTheory.UniqueFactorizationDomain.FactorSet +import Mathlib.RingTheory.UniqueFactorizationDomain.Finite +import Mathlib.RingTheory.UniqueFactorizationDomain.Finsupp +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid +import Mathlib.RingTheory.UniqueFactorizationDomain.Ideal +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicative +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity +import Mathlib.RingTheory.UniqueFactorizationDomain.Nat +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors import Mathlib.RingTheory.Unramified.Basic import Mathlib.RingTheory.Unramified.Field import Mathlib.RingTheory.Unramified.Finite diff --git a/Mathlib/Algebra/Squarefree/Basic.lean b/Mathlib/Algebra/Squarefree/Basic.lean index 284bc77d8edec..9d68d11fc173a 100644 --- a/Mathlib/Algebra/Squarefree/Basic.lean +++ b/Mathlib/Algebra/Squarefree/Basic.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.RingTheory.Nilpotent.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity /-! # Squarefree elements of monoids diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index 0dc3527d3bad6..5d0d0f094055a 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -6,7 +6,6 @@ Authors: Yaël Dillies import Mathlib.Algebra.BigOperators.Ring import Mathlib.Data.Multiset.Fintype import Mathlib.FieldTheory.ChevalleyWarning -import Mathlib.RingTheory.UniqueFactorizationDomain /-! # The Erdős–Ginzburg–Ziv theorem diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 40e3df63b05fc..53bb8bb29fec5 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -5,6 +5,7 @@ Authors: Aaron Anderson -/ import Mathlib.Algebra.Squarefree.Basic import Mathlib.Data.Nat.Factorization.PrimePow +import Mathlib.RingTheory.UniqueFactorizationDomain.Nat /-! # Lemmas about squarefreeness of natural numbers diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index d6a37afa99adb..5c13f9d53ebe2 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -299,7 +299,7 @@ theorem one_smul' (b : ArithmeticFunction M) : (1 : ArithmeticFunction R) • b intro y ymem ynmem have y1ne : y.fst ≠ 1 := by intro con - simp only [Con, mem_divisorsAntidiagonal, one_mul, Ne] at ymem + simp only [mem_divisorsAntidiagonal, one_mul, Ne] at ymem simp only [mem_singleton, Prod.ext_iff] at ynmem -- Porting note: `tauto` worked from here. cases y @@ -331,7 +331,7 @@ instance instMonoid : Monoid (ArithmeticFunction R) := have y2ne : y.snd ≠ 1 := by intro con cases y; subst con -- Porting note: added - simp only [Con, mem_divisorsAntidiagonal, mul_one, Ne] at ymem + simp only [mem_divisorsAntidiagonal, mul_one, Ne] at ymem simp only [mem_singleton, Prod.ext_iff] at ynmem tauto simp [y2ne] diff --git a/Mathlib/NumberTheory/FLT/Four.lean b/Mathlib/NumberTheory/FLT/Four.lean index e6da8e8d4814f..b6802f986775d 100644 --- a/Mathlib/NumberTheory/FLT/Four.lean +++ b/Mathlib/NumberTheory/FLT/Four.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Paul van Wamelen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul van Wamelen -/ +import Mathlib.Data.Nat.Factors import Mathlib.NumberTheory.FLT.Basic import Mathlib.NumberTheory.PythagoreanTriples import Mathlib.RingTheory.Coprime.Lemmas diff --git a/Mathlib/RingTheory/ChainOfDivisors.lean b/Mathlib/RingTheory/ChainOfDivisors.lean index f22b75532ab93..a1921d9752ce6 100644 --- a/Mathlib/RingTheory/ChainOfDivisors.lean +++ b/Mathlib/RingTheory/ChainOfDivisors.lean @@ -3,10 +3,11 @@ Copyright (c) 2021 Paul Lezeau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Paul Lezeau -/ +import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.Algebra.IsPrimePow import Mathlib.Algebra.Squarefree.Basic -import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.Data.ZMod.Defs +import Mathlib.Order.Atoms import Mathlib.Order.Hom.Bounded /-! diff --git a/Mathlib/RingTheory/DedekindDomain/Factorization.lean b/Mathlib/RingTheory/DedekindDomain/Factorization.lean index 9a837256ab9a2..61803b29f13c4 100644 --- a/Mathlib/RingTheory/DedekindDomain/Factorization.lean +++ b/Mathlib/RingTheory/DedekindDomain/Factorization.lean @@ -5,6 +5,7 @@ Authors: María Inés de Frutos-Fernández -/ import Mathlib.Order.Filter.Cofinite import Mathlib.RingTheory.DedekindDomain.Ideal +import Mathlib.RingTheory.UniqueFactorizationDomain.Finite /-! # Factorization of ideals and fractional ideals of Dedekind domains diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index cbcfd6561f1b7..33a536b988ff1 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -5,6 +5,7 @@ Authors: Kevin Buzzard -/ import Mathlib.RingTheory.AdicCompletion.Basic import Mathlib.RingTheory.LocalRing.MaximalIdeal.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic import Mathlib.RingTheory.Valuation.PrimeMultiplicity import Mathlib.RingTheory.Valuation.ValuationRing diff --git a/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean index 69f9fc150ca4f..bf904f8bef341 100644 --- a/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean +++ b/Mathlib/RingTheory/Ideal/Norm/AbsNorm.lean @@ -11,6 +11,7 @@ import Mathlib.LinearAlgebra.FreeModule.IdealQuotient import Mathlib.RingTheory.DedekindDomain.Dvr import Mathlib.RingTheory.DedekindDomain.Ideal import Mathlib.RingTheory.Norm.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicative /-! diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index 3f7cd7a8e49b6..d587619c72796 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -6,7 +6,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan import Mathlib.GroupTheory.MonoidLocalization.Away import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Localization.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity /-! # Localizations away from an element diff --git a/Mathlib/RingTheory/Localization/NumDen.lean b/Mathlib/RingTheory/Localization/NumDen.lean index 9d72b32932ecc..290ca0e000441 100644 --- a/Mathlib/RingTheory/Localization/NumDen.lean +++ b/Mathlib/RingTheory/Localization/NumDen.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baan -/ import Mathlib.RingTheory.Localization.FractionRing import Mathlib.RingTheory.Localization.Integer -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid /-! # Numerator and denominator in a localization diff --git a/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean index a56bfca5ca1bf..b475161288512 100644 --- a/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/Noetherian/UniqueFactorizationDomain.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.RingTheory.Noetherian.Defs -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Ideal /-! # Noetherian domains have unique factorization diff --git a/Mathlib/RingTheory/Polynomial/Radical.lean b/Mathlib/RingTheory/Polynomial/Radical.lean index 71a22fad74761..4a8f0a57af1e6 100644 --- a/Mathlib/RingTheory/Polynomial/Radical.lean +++ b/Mathlib/RingTheory/Polynomial/Radical.lean @@ -6,6 +6,7 @@ Authors: Jineon Baek, Seewoo Lee import Mathlib.Algebra.Polynomial.FieldDivision import Mathlib.RingTheory.Polynomial.Wronskian import Mathlib.RingTheory.Radical +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicative /-! # Radical of a polynomial diff --git a/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean b/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean index bf2b2308eda3b..6e24a23d71aee 100644 --- a/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean +++ b/Mathlib/RingTheory/Polynomial/UniqueFactorization.lean @@ -5,6 +5,9 @@ Authors: Kenny Lau -/ import Mathlib.RingTheory.Polynomial.Basic import Mathlib.RingTheory.Polynomial.Content +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Finite +import Mathlib.RingTheory.UniqueFactorizationDomain.GCDMonoid /-! # Unique factorization for univariate and multivariate polynomials diff --git a/Mathlib/RingTheory/PowerSeries/Inverse.lean b/Mathlib/RingTheory/PowerSeries/Inverse.lean index c8bef18f0aeb4..f69014616a2de 100644 --- a/Mathlib/RingTheory/PowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/PowerSeries/Inverse.lean @@ -10,6 +10,7 @@ import Mathlib.RingTheory.MvPowerSeries.Inverse import Mathlib.RingTheory.PowerSeries.Basic import Mathlib.RingTheory.PowerSeries.Order import Mathlib.RingTheory.LocalRing.ResidueField.Defs +import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity /-! # Formal power series - Inverses diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 53c6974d87d3a..7cde766527865 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Morenikeji Neri -/ import Mathlib.Algebra.EuclideanDomain.Field +import Mathlib.Algebra.GCDMonoid.Basic import Mathlib.RingTheory.Ideal.Nonunits import Mathlib.RingTheory.Noetherian.UniqueFactorizationDomain diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index dbf09f0b12caa..a4558ba7c2ec9 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jineon Baek, Seewoo Lee -/ import Mathlib.Algebra.EuclideanDomain.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.Coprime.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors /-! # Radical of an element of a unique factorization normalization monoid diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index 992c920e1efc9..34a8597d9b708 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -3,10 +3,11 @@ Copyright (c) 2020 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca, Johan Commelin -/ -import Mathlib.RingTheory.RootsOfUnity.PrimitiveRoots -import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed import Mathlib.Algebra.GCDMonoid.IntegrallyClosed import Mathlib.FieldTheory.Finite.Basic +import Mathlib.FieldTheory.Minpoly.IsIntegrallyClosed +import Mathlib.RingTheory.RootsOfUnity.PrimitiveRoots +import Mathlib.RingTheory.UniqueFactorizationDomain.Nat /-! # Minimal polynomial of roots of unity diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean deleted file mode 100644 index a5f452067a120..0000000000000 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ /dev/null @@ -1,2017 +0,0 @@ -/- -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, Aaron Anderson --/ -import Mathlib.Algebra.BigOperators.Associated -import Mathlib.Algebra.GCDMonoid.Basic -import Mathlib.Data.ENat.Lattice -import Mathlib.Data.Finsupp.Multiset -import Mathlib.Data.Nat.Factors -import Mathlib.RingTheory.Multiplicity -import Mathlib.RingTheory.Ideal.Operations - -/-! - -# Unique factorization - -## Main Definitions -* `WfDvdMonoid` holds for `Monoid`s for which a strict divisibility relation is - well-founded. -* `UniqueFactorizationMonoid` holds for `WfDvdMonoid`s where - `Irreducible` is equivalent to `Prime` - -## Main results -* `Ideal.setOf_isPrincipal_wellFoundedOn_gt`, `WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt` - in a domain, well-foundedness of the strict version of ∣ is equivalent to the ascending - chain condition on principal ideals. - -## TODO -* set up the complete lattice structure on `FactorSet`. - --/ - - -variable {α : Type*} - -local infixl:50 " ~ᵤ " => Associated - -/-- Well-foundedness of the strict version of ∣, which is equivalent to the descending chain -condition on divisibility and to the ascending chain condition on -principal ideals in an integral domain. - -/ -abbrev WfDvdMonoid (α : Type*) [CommMonoidWithZero α] : Prop := - IsWellFounded α DvdNotUnit - -theorem wellFounded_dvdNotUnit {α : Type*} [CommMonoidWithZero α] [h : WfDvdMonoid α] : - WellFounded (DvdNotUnit (α := α)) := - h.wf - -namespace WfDvdMonoid - -variable [CommMonoidWithZero α] - -open Associates Nat - -theorem of_wfDvdMonoid_associates (_ : WfDvdMonoid (Associates α)) : WfDvdMonoid α := - ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).2 wellFounded_dvdNotUnit⟩ - -variable [WfDvdMonoid α] - -instance wfDvdMonoid_associates : WfDvdMonoid (Associates α) := - ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).1 wellFounded_dvdNotUnit⟩ - -theorem wellFoundedLT_associates : WellFoundedLT (Associates α) := - ⟨Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit⟩ - -@[deprecated wellFoundedLT_associates (since := "2024-09-02")] -theorem wellFounded_associates : WellFounded ((· < ·) : Associates α → Associates α → Prop) := - Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit - --- Porting note: elab_as_elim can only be global and cannot be changed on an imported decl --- attribute [local elab_as_elim] WellFounded.fix - -theorem exists_irreducible_factor {a : α} (ha : ¬IsUnit a) (ha0 : a ≠ 0) : - ∃ i, Irreducible i ∧ i ∣ a := - let ⟨b, hs, hr⟩ := wellFounded_dvdNotUnit.has_min { b | b ∣ a ∧ ¬IsUnit b } ⟨a, dvd_rfl, ha⟩ - ⟨b, - ⟨hs.2, fun c d he => - let h := dvd_trans ⟨d, he⟩ hs.1 - or_iff_not_imp_left.2 fun hc => - of_not_not fun hd => hr c ⟨h, hc⟩ ⟨ne_zero_of_dvd_ne_zero ha0 h, d, hd, he⟩⟩, - hs.1⟩ - -@[elab_as_elim] -theorem induction_on_irreducible {P : α → Prop} (a : α) (h0 : P 0) (hu : ∀ u : α, IsUnit u → P u) - (hi : ∀ a i : α, a ≠ 0 → Irreducible i → P a → P (i * a)) : P a := - haveI := Classical.dec - wellFounded_dvdNotUnit.fix - (fun a ih => - if ha0 : a = 0 then ha0.substr h0 - else - if hau : IsUnit a then hu a hau - else - let ⟨i, hii, b, hb⟩ := exists_irreducible_factor hau ha0 - let hb0 : b ≠ 0 := ne_zero_of_dvd_ne_zero ha0 ⟨i, mul_comm i b ▸ hb⟩ - hb.symm ▸ hi b i hb0 hii <| ih b ⟨hb0, i, hii.1, mul_comm i b ▸ hb⟩) - a - -theorem exists_factors (a : α) : - a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ Associated f.prod a := - induction_on_irreducible a (fun h => (h rfl).elim) - (fun _ hu _ => ⟨0, fun _ h => False.elim (Multiset.not_mem_zero _ h), hu.unit, one_mul _⟩) - fun a i ha0 hi ih _ => - let ⟨s, hs⟩ := ih ha0 - ⟨i ::ₘ s, fun b H => (Multiset.mem_cons.1 H).elim (fun h => h.symm ▸ hi) (hs.1 b), by - rw [s.prod_cons i] - exact hs.2.mul_left i⟩ - -theorem not_unit_iff_exists_factors_eq (a : α) (hn0 : a ≠ 0) : - ¬IsUnit a ↔ ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod = a ∧ f ≠ ∅ := - ⟨fun hnu => by - obtain ⟨f, hi, u, rfl⟩ := exists_factors a hn0 - obtain ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero fun h : f = 0 => hnu <| by simp [h] - classical - refine ⟨(f.erase b).cons (b * u), fun a ha => ?_, ?_, Multiset.cons_ne_zero⟩ - · obtain rfl | ha := Multiset.mem_cons.1 ha - exacts [Associated.irreducible ⟨u, rfl⟩ (hi b h), hi a (Multiset.mem_of_mem_erase ha)] - · rw [Multiset.prod_cons, mul_comm b, mul_assoc, Multiset.prod_erase h, mul_comm], - fun ⟨_, hi, he, hne⟩ => - let ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero hne - not_isUnit_of_not_isUnit_dvd (hi b h).not_unit <| he ▸ Multiset.dvd_prod h⟩ - -theorem isRelPrime_of_no_irreducible_factors {x y : α} (nonzero : ¬(x = 0 ∧ y = 0)) - (H : ∀ z : α, Irreducible z → z ∣ x → ¬z ∣ y) : IsRelPrime x y := - isRelPrime_of_no_nonunits_factors nonzero fun _z znu znz zx zy ↦ - have ⟨i, h1, h2⟩ := exists_irreducible_factor znu znz - H i h1 (h2.trans zx) (h2.trans zy) - -end WfDvdMonoid - -theorem WfDvdMonoid.of_wellFoundedLT_associates [CancelCommMonoidWithZero α] - (h : WellFoundedLT (Associates α)) : WfDvdMonoid α := - WfDvdMonoid.of_wfDvdMonoid_associates - ⟨by - convert h.wf - ext - exact Associates.dvdNotUnit_iff_lt⟩ - -@[deprecated WfDvdMonoid.of_wellFoundedLT_associates (since := "2024-09-02")] -theorem WfDvdMonoid.of_wellFounded_associates [CancelCommMonoidWithZero α] - (h : WellFounded ((· < ·) : Associates α → Associates α → Prop)) : WfDvdMonoid α := - WfDvdMonoid.of_wfDvdMonoid_associates - ⟨by - convert h - ext - exact Associates.dvdNotUnit_iff_lt⟩ - -theorem WfDvdMonoid.iff_wellFounded_associates [CancelCommMonoidWithZero α] : - WfDvdMonoid α ↔ WellFoundedLT (Associates α) := - ⟨by apply WfDvdMonoid.wellFoundedLT_associates, WfDvdMonoid.of_wellFoundedLT_associates⟩ - -theorem WfDvdMonoid.max_power_factor' [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} - (h : a₀ ≠ 0) (hx : ¬IsUnit x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := by - obtain ⟨a, ⟨n, rfl⟩, hm⟩ := wellFounded_dvdNotUnit.has_min - {a | ∃ n, x ^ n * a = a₀} ⟨a₀, 0, by rw [pow_zero, one_mul]⟩ - refine ⟨n, a, ?_, rfl⟩; rintro ⟨d, rfl⟩ - exact hm d ⟨n + 1, by rw [pow_succ, mul_assoc]⟩ - ⟨(right_ne_zero_of_mul <| right_ne_zero_of_mul h), x, hx, mul_comm _ _⟩ - -theorem WfDvdMonoid.max_power_factor [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} - (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := - max_power_factor' h hx.not_unit - -theorem multiplicity.finite_of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] - {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : multiplicity.Finite a b := by - obtain ⟨n, c, ndvd, rfl⟩ := WfDvdMonoid.max_power_factor' hb ha - exact ⟨n, by rwa [pow_succ, mul_dvd_mul_iff_left (left_ne_zero_of_mul hb)]⟩ - -section Prio - --- set_option default_priority 100 - --- see Note [default priority] -/-- unique factorization monoids. - -These are defined as `CancelCommMonoidWithZero`s with well-founded strict divisibility -relations, but this is equivalent to more familiar definitions: - -Each element (except zero) is uniquely represented as a multiset of irreducible factors. -Uniqueness is only up to associated elements. - -Each element (except zero) is non-uniquely represented as a multiset -of prime factors. - -To define a UFD using the definition in terms of multisets -of irreducible factors, use the definition `of_exists_unique_irreducible_factors` - -To define a UFD using the definition in terms of multisets -of prime factors, use the definition `of_exists_prime_factors` - --/ -class UniqueFactorizationMonoid (α : Type*) [CancelCommMonoidWithZero α] extends - IsWellFounded α DvdNotUnit : Prop where - protected irreducible_iff_prime : ∀ {a : α}, Irreducible a ↔ Prime a - -instance (priority := 100) ufm_of_decomposition_of_wfDvdMonoid - [CancelCommMonoidWithZero α] [WfDvdMonoid α] [DecompositionMonoid α] : - UniqueFactorizationMonoid α := - { ‹WfDvdMonoid α› with irreducible_iff_prime := irreducible_iff_prime } - -@[deprecated ufm_of_decomposition_of_wfDvdMonoid (since := "2024-02-12")] -theorem ufm_of_gcd_of_wfDvdMonoid [CancelCommMonoidWithZero α] [WfDvdMonoid α] - [DecompositionMonoid α] : UniqueFactorizationMonoid α := - ufm_of_decomposition_of_wfDvdMonoid - -instance Associates.ufm [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : - UniqueFactorizationMonoid (Associates α) := - { (WfDvdMonoid.wfDvdMonoid_associates : WfDvdMonoid (Associates α)) with - irreducible_iff_prime := by - rw [← Associates.irreducible_iff_prime_iff] - apply UniqueFactorizationMonoid.irreducible_iff_prime } - -end Prio - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] - -theorem exists_prime_factors (a : α) : - a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := by - simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] - apply WfDvdMonoid.exists_factors a - -instance : DecompositionMonoid α where - primal a := by - obtain rfl | ha := eq_or_ne a 0; · exact isPrimal_zero - obtain ⟨f, hf, u, rfl⟩ := exists_prime_factors a ha - exact ((Submonoid.isPrimal α).multiset_prod_mem f (hf · ·|>.isPrimal)).mul u.isUnit.isPrimal - -lemma exists_prime_iff : - (∃ (p : α), Prime p) ↔ ∃ (x : α), x ≠ 0 ∧ ¬ IsUnit x := by - refine ⟨fun ⟨p, hp⟩ ↦ ⟨p, hp.ne_zero, hp.not_unit⟩, fun ⟨x, hx₀, hxu⟩ ↦ ?_⟩ - obtain ⟨f, hf, -⟩ := WfDvdMonoid.exists_irreducible_factor hxu hx₀ - exact ⟨f, UniqueFactorizationMonoid.irreducible_iff_prime.mp hf⟩ - -@[elab_as_elim] -theorem induction_on_prime {P : α → Prop} (a : α) (h₁ : P 0) (h₂ : ∀ x : α, IsUnit x → P x) - (h₃ : ∀ a p : α, a ≠ 0 → Prime p → P a → P (p * a)) : P a := by - simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] at h₃ - exact WfDvdMonoid.induction_on_irreducible a h₁ h₂ h₃ - -end UniqueFactorizationMonoid - -theorem prime_factors_unique [CancelCommMonoidWithZero α] : - ∀ {f g : Multiset α}, - (∀ x ∈ f, Prime x) → (∀ x ∈ g, Prime x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g := by - classical - intro f - induction' f using Multiset.induction_on with p f ih - · intros g _ hg h - exact Multiset.rel_zero_left.2 <| - Multiset.eq_zero_of_forall_not_mem fun x hx => - have : IsUnit g.prod := by simpa [associated_one_iff_isUnit] using h.symm - (hg x hx).not_unit <| - isUnit_iff_dvd_one.2 <| (Multiset.dvd_prod hx).trans (isUnit_iff_dvd_one.1 this) - · intros g hf hg hfg - let ⟨b, hbg, hb⟩ := - (exists_associated_mem_of_dvd_prod (hf p (by simp)) fun q hq => hg _ hq) <| - hfg.dvd_iff_dvd_right.1 (show p ∣ (p ::ₘ f).prod by simp) - haveI := Classical.decEq α - rw [← Multiset.cons_erase hbg] - exact - Multiset.Rel.cons hb - (ih (fun q hq => hf _ (by simp [hq])) - (fun {q} (hq : q ∈ g.erase b) => hg q (Multiset.mem_of_mem_erase hq)) - (Associated.of_mul_left - (by rwa [← Multiset.prod_cons, ← Multiset.prod_cons, Multiset.cons_erase hbg]) hb - (hf p (by simp)).ne_zero)) - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] - -theorem factors_unique {f g : Multiset α} (hf : ∀ x ∈ f, Irreducible x) - (hg : ∀ x ∈ g, Irreducible x) (h : f.prod ~ᵤ g.prod) : Multiset.Rel Associated f g := - prime_factors_unique (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hf x hx)) - (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hg x hx)) h - -end UniqueFactorizationMonoid - -/-- If an irreducible has a prime factorization, - then it is an associate of one of its prime factors. -/ -theorem prime_factors_irreducible [CancelCommMonoidWithZero α] {a : α} {f : Multiset α} - (ha : Irreducible a) (pfa : (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) : ∃ p, a ~ᵤ p ∧ f = {p} := by - haveI := Classical.decEq α - refine @Multiset.induction_on _ - (fun g => (g.prod ~ᵤ a) → (∀ b ∈ g, Prime b) → ∃ p, a ~ᵤ p ∧ g = {p}) f ?_ ?_ pfa.2 pfa.1 - · intro h; exact (ha.not_unit (associated_one_iff_isUnit.1 (Associated.symm h))).elim - · rintro p s _ ⟨u, hu⟩ hs - use p - have hs0 : s = 0 := by - by_contra hs0 - obtain ⟨q, hq⟩ := Multiset.exists_mem_of_ne_zero hs0 - apply (hs q (by simp [hq])).2.1 - refine (ha.isUnit_or_isUnit (?_ : _ = p * ↑u * (s.erase q).prod * _)).resolve_left ?_ - · rw [mul_right_comm _ _ q, mul_assoc, ← Multiset.prod_cons, Multiset.cons_erase hq, ← hu, - mul_comm, mul_comm p _, mul_assoc] - simp - apply mt isUnit_of_mul_isUnit_left (mt isUnit_of_mul_isUnit_left _) - apply (hs p (Multiset.mem_cons_self _ _)).2.1 - simp only [mul_one, Multiset.prod_cons, Multiset.prod_zero, hs0] at * - exact ⟨Associated.symm ⟨u, hu⟩, rfl⟩ - -section ExistsPrimeFactors - -variable [CancelCommMonoidWithZero α] -variable (pf : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) -include pf - -theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := - ⟨by - classical - refine RelHomClass.wellFounded - (RelHom.mk ?_ ?_ : (DvdNotUnit : α → α → Prop) →r ((· < ·) : ℕ∞ → ℕ∞ → Prop)) wellFounded_lt - · intro a - by_cases h : a = 0 - · exact ⊤ - exact ↑(Multiset.card (Classical.choose (pf a h))) - rintro a b ⟨ane0, ⟨c, hc, b_eq⟩⟩ - rw [dif_neg ane0] - by_cases h : b = 0 - · simp [h, lt_top_iff_ne_top] - · rw [dif_neg h, Nat.cast_lt] - have cne0 : c ≠ 0 := by - refine mt (fun con => ?_) h - rw [b_eq, con, mul_zero] - calc - Multiset.card (Classical.choose (pf a ane0)) < - _ + Multiset.card (Classical.choose (pf c cne0)) := - lt_add_of_pos_right _ - (Multiset.card_pos.mpr fun con => hc (associated_one_iff_isUnit.mp ?_)) - _ = Multiset.card (Classical.choose (pf a ane0) + Classical.choose (pf c cne0)) := - (Multiset.card_add _ _).symm - _ = Multiset.card (Classical.choose (pf b h)) := - Multiset.card_eq_card_of_rel - (prime_factors_unique ?_ (Classical.choose_spec (pf _ h)).1 ?_) - - · convert (Classical.choose_spec (pf c cne0)).2.symm - rw [con, Multiset.prod_zero] - · intro x hadd - rw [Multiset.mem_add] at hadd - cases' hadd with h h <;> apply (Classical.choose_spec (pf _ _)).1 _ h <;> assumption - · rw [Multiset.prod_add] - trans a * c - · apply Associated.mul_mul <;> apply (Classical.choose_spec (pf _ _)).2 <;> assumption - · rw [← b_eq] - apply (Classical.choose_spec (pf _ _)).2.symm; assumption⟩ - -theorem irreducible_iff_prime_of_exists_prime_factors {p : α} : Irreducible p ↔ Prime p := by - by_cases hp0 : p = 0 - · simp [hp0] - refine ⟨fun h => ?_, Prime.irreducible⟩ - obtain ⟨f, hf⟩ := pf p hp0 - obtain ⟨q, hq, rfl⟩ := prime_factors_irreducible h hf - rw [hq.prime_iff] - exact hf.1 q (Multiset.mem_singleton_self _) - -theorem UniqueFactorizationMonoid.of_exists_prime_factors : UniqueFactorizationMonoid α := - { WfDvdMonoid.of_exists_prime_factors pf with - irreducible_iff_prime := irreducible_iff_prime_of_exists_prime_factors pf } - -end ExistsPrimeFactors - -theorem UniqueFactorizationMonoid.iff_exists_prime_factors [CancelCommMonoidWithZero α] : - UniqueFactorizationMonoid α ↔ - ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := - ⟨fun h => @UniqueFactorizationMonoid.exists_prime_factors _ _ h, - UniqueFactorizationMonoid.of_exists_prime_factors⟩ - -section - -variable {β : Type*} [CancelCommMonoidWithZero α] [CancelCommMonoidWithZero β] - -theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactorizationMonoid α) : - UniqueFactorizationMonoid β := by - rw [UniqueFactorizationMonoid.iff_exists_prime_factors] at hα ⊢ - intro a ha - obtain ⟨w, hp, u, h⟩ := - hα (e.symm a) fun h => - ha <| by - convert← map_zero e - simp [← h] - exact - ⟨w.map e, fun b hb => - let ⟨c, hc, he⟩ := Multiset.mem_map.1 hb - he ▸ e.prime_iff.1 (hp c hc), - Units.map e.toMonoidHom u, - by - rw [Multiset.prod_hom, toMonoidHom_eq_coe, Units.coe_map, MonoidHom.coe_coe, ← map_mul e, h, - apply_symm_apply]⟩ - -theorem MulEquiv.uniqueFactorizationMonoid_iff (e : α ≃* β) : - UniqueFactorizationMonoid α ↔ UniqueFactorizationMonoid β := - ⟨e.uniqueFactorizationMonoid, e.symm.uniqueFactorizationMonoid⟩ - -end - -theorem irreducible_iff_prime_of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] - (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) - (uif : - ∀ f g : Multiset α, - (∀ x ∈ f, Irreducible x) → - (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) - (p : α) : Irreducible p ↔ Prime p := - letI := Classical.decEq α - ⟨ fun hpi => - ⟨hpi.ne_zero, hpi.1, fun a b ⟨x, hx⟩ => - if hab0 : a * b = 0 then - (eq_zero_or_eq_zero_of_mul_eq_zero hab0).elim (fun ha0 => by simp [ha0]) fun hb0 => by - simp [hb0] - else by - have hx0 : x ≠ 0 := fun hx0 => by simp_all - have ha0 : a ≠ 0 := left_ne_zero_of_mul hab0 - have hb0 : b ≠ 0 := right_ne_zero_of_mul hab0 - cases' eif x hx0 with fx hfx - cases' eif a ha0 with fa hfa - cases' eif b hb0 with fb hfb - have h : Multiset.Rel Associated (p ::ₘ fx) (fa + fb) := by - apply uif - · exact fun i hi => (Multiset.mem_cons.1 hi).elim (fun hip => hip.symm ▸ hpi) (hfx.1 _) - · exact fun i hi => (Multiset.mem_add.1 hi).elim (hfa.1 _) (hfb.1 _) - calc - Multiset.prod (p ::ₘ fx) ~ᵤ a * b := by - rw [hx, Multiset.prod_cons]; exact hfx.2.mul_left _ - _ ~ᵤ fa.prod * fb.prod := hfa.2.symm.mul_mul hfb.2.symm - _ = _ := by rw [Multiset.prod_add] - - exact - let ⟨q, hqf, hq⟩ := Multiset.exists_mem_of_rel_of_mem h (Multiset.mem_cons_self p _) - (Multiset.mem_add.1 hqf).elim - (fun hqa => - Or.inl <| hq.dvd_iff_dvd_left.2 <| hfa.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqa)) - fun hqb => - Or.inr <| hq.dvd_iff_dvd_left.2 <| hfb.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqb)⟩, - Prime.irreducible⟩ - -theorem UniqueFactorizationMonoid.of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] - (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) - (uif : - ∀ f g : Multiset α, - (∀ x ∈ f, Irreducible x) → - (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) : - UniqueFactorizationMonoid α := - UniqueFactorizationMonoid.of_exists_prime_factors - (by - convert eif using 7 - simp_rw [irreducible_iff_prime_of_exists_unique_irreducible_factors eif uif]) - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] -variable [UniqueFactorizationMonoid α] - -open Classical in -/-- Noncomputably determines the multiset of prime factors. -/ -noncomputable def factors (a : α) : Multiset α := - if h : a = 0 then 0 else Classical.choose (UniqueFactorizationMonoid.exists_prime_factors a h) - -theorem factors_prod {a : α} (ane0 : a ≠ 0) : Associated (factors a).prod a := by - rw [factors, dif_neg ane0] - exact (Classical.choose_spec (exists_prime_factors a ane0)).2 - -@[simp] -theorem factors_zero : factors (0 : α) = 0 := by simp [factors] - -theorem ne_zero_of_mem_factors {p a : α} (h : p ∈ factors a) : a ≠ 0 := by - rintro rfl - simp at h - -theorem dvd_of_mem_factors {p a : α} (h : p ∈ factors a) : p ∣ a := - dvd_trans (Multiset.dvd_prod h) (Associated.dvd (factors_prod (ne_zero_of_mem_factors h))) - -theorem prime_of_factor {a : α} (x : α) (hx : x ∈ factors a) : Prime x := by - have ane0 := ne_zero_of_mem_factors hx - rw [factors, dif_neg ane0] at hx - exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 x hx - -theorem irreducible_of_factor {a : α} : ∀ x : α, x ∈ factors a → Irreducible x := fun x h => - (prime_of_factor x h).irreducible - -@[simp] -theorem factors_one : factors (1 : α) = 0 := by - nontriviality α using factors - rw [← Multiset.rel_zero_right] - refine factors_unique irreducible_of_factor (fun x hx => (Multiset.not_mem_zero x hx).elim) ?_ - rw [Multiset.prod_zero] - exact factors_prod one_ne_zero - -theorem exists_mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - p ∣ a → ∃ q ∈ factors a, p ~ᵤ q := fun ⟨b, hb⟩ => - have hb0 : b ≠ 0 := fun hb0 => by simp_all - have : Multiset.Rel Associated (p ::ₘ factors b) (factors a) := - factors_unique - (fun _ hx => (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_factor _)) - irreducible_of_factor - (Associated.symm <| - calc - Multiset.prod (factors a) ~ᵤ a := factors_prod ha0 - _ = p * b := hb - _ ~ᵤ Multiset.prod (p ::ₘ factors b) := by - rw [Multiset.prod_cons]; exact (factors_prod hb0).symm.mul_left _ - ) - Multiset.exists_mem_of_rel_of_mem this (by simp) - -theorem exists_mem_factors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : ∃ p, p ∈ factors x := by - obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx - obtain ⟨p, hp, _⟩ := exists_mem_factors_of_dvd hx hp' hp'x - exact ⟨p, hp⟩ - -open Classical in -theorem factors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - Multiset.Rel Associated (factors (x * y)) (factors x + factors y) := by - refine - factors_unique irreducible_of_factor - (fun a ha => - (Multiset.mem_add.mp ha).by_cases (irreducible_of_factor _) (irreducible_of_factor _)) - ((factors_prod (mul_ne_zero hx hy)).trans ?_) - rw [Multiset.prod_add] - exact (Associated.mul_mul (factors_prod hx) (factors_prod hy)).symm - -theorem factors_pow {x : α} (n : ℕ) : - Multiset.Rel Associated (factors (x ^ n)) (n • factors x) := by - match n with - | 0 => rw [zero_smul, pow_zero, factors_one, Multiset.rel_zero_right] - | n+1 => - by_cases h0 : x = 0 - · simp [h0, zero_pow n.succ_ne_zero, smul_zero] - · rw [pow_succ', succ_nsmul'] - refine Multiset.Rel.trans _ (factors_mul h0 (pow_ne_zero n h0)) ?_ - refine Multiset.Rel.add ?_ <| factors_pow n - exact Multiset.rel_refl_of_refl_on fun y _ => Associated.refl _ - -@[simp] -theorem factors_pos (x : α) (hx : x ≠ 0) : 0 < factors x ↔ ¬IsUnit x := by - constructor - · intro h hx - obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' - exact (prime_of_factor _ hp).not_unit (isUnit_of_dvd_unit (dvd_of_mem_factors hp) hx) - · intro h - obtain ⟨p, hp⟩ := exists_mem_factors hx h - exact - bot_lt_iff_ne_bot.mpr - (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) - -open Multiset in -theorem factors_pow_count_prod [DecidableEq α] {x : α} (hx : x ≠ 0) : - (∏ p ∈ (factors x).toFinset, p ^ (factors x).count p) ~ᵤ x := - calc - _ = prod (∑ a ∈ toFinset (factors x), count a (factors x) • {a}) := by - simp only [prod_sum, prod_nsmul, prod_singleton] - _ = prod (factors x) := by rw [toFinset_sum_count_nsmul_eq (factors x)] - _ ~ᵤ x := factors_prod hx - -theorem factors_rel_of_associated {a b : α} (h : Associated a b) : - Multiset.Rel Associated (factors a) (factors b) := by - rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) - · simp - · refine factors_unique irreducible_of_factor irreducible_of_factor ?_ - exact ((factors_prod ha).trans h).trans (factors_prod hb).symm - -end UniqueFactorizationMonoid - -namespace UniqueFactorizationMonoid - -variable [CancelCommMonoidWithZero α] [NormalizationMonoid α] -variable [UniqueFactorizationMonoid α] - -/-- Noncomputably determines the multiset of prime factors. -/ -noncomputable def normalizedFactors (a : α) : Multiset α := - Multiset.map normalize <| factors a - -/-- An arbitrary choice of factors of `x : M` is exactly the (unique) normalized set of factors, -if `M` has a trivial group of units. -/ -@[simp] -theorem factors_eq_normalizedFactors {M : Type*} [CancelCommMonoidWithZero M] - [UniqueFactorizationMonoid M] [Subsingleton Mˣ] (x : M) : factors x = normalizedFactors x := by - unfold normalizedFactors - convert (Multiset.map_id (factors x)).symm - ext p - exact normalize_eq p - -theorem normalizedFactors_prod {a : α} (ane0 : a ≠ 0) : - Associated (normalizedFactors a).prod a := by - rw [normalizedFactors, factors, dif_neg ane0] - refine Associated.trans ?_ (Classical.choose_spec (exists_prime_factors a ane0)).2 - rw [← Associates.mk_eq_mk_iff_associated, ← Associates.prod_mk, ← Associates.prod_mk, - Multiset.map_map] - congr 2 - ext - rw [Function.comp_apply, Associates.mk_normalize] - -theorem prime_of_normalized_factor {a : α} : ∀ x : α, x ∈ normalizedFactors a → Prime x := by - rw [normalizedFactors, factors] - split_ifs with ane0; · simp - intro x hx; rcases Multiset.mem_map.1 hx with ⟨y, ⟨hy, rfl⟩⟩ - rw [(normalize_associated _).prime_iff] - exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 y hy - -theorem irreducible_of_normalized_factor {a : α} : - ∀ x : α, x ∈ normalizedFactors a → Irreducible x := fun x h => - (prime_of_normalized_factor x h).irreducible - -theorem normalize_normalized_factor {a : α} : - ∀ x : α, x ∈ normalizedFactors a → normalize x = x := by - rw [normalizedFactors, factors] - split_ifs with h; · simp - intro x hx - obtain ⟨y, _, rfl⟩ := Multiset.mem_map.1 hx - apply normalize_idem - -theorem normalizedFactors_irreducible {a : α} (ha : Irreducible a) : - normalizedFactors a = {normalize a} := by - obtain ⟨p, a_assoc, hp⟩ := - prime_factors_irreducible ha ⟨prime_of_normalized_factor, normalizedFactors_prod ha.ne_zero⟩ - have p_mem : p ∈ normalizedFactors a := by - rw [hp] - exact Multiset.mem_singleton_self _ - convert hp - rwa [← normalize_normalized_factor p p_mem, normalize_eq_normalize_iff, dvd_dvd_iff_associated] - -theorem normalizedFactors_eq_of_dvd (a : α) : - ∀ᵉ (p ∈ normalizedFactors a) (q ∈ normalizedFactors a), p ∣ q → p = q := by - intro p hp q hq hdvd - convert normalize_eq_normalize hdvd - ((prime_of_normalized_factor _ hp).irreducible.dvd_symm - (prime_of_normalized_factor _ hq).irreducible hdvd) <;> - apply (normalize_normalized_factor _ ‹_›).symm - -theorem exists_mem_normalizedFactors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - p ∣ a → ∃ q ∈ normalizedFactors a, p ~ᵤ q := fun ⟨b, hb⟩ => - have hb0 : b ≠ 0 := fun hb0 => by simp_all - have : Multiset.Rel Associated (p ::ₘ normalizedFactors b) (normalizedFactors a) := - factors_unique - (fun _ hx => - (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_normalized_factor _)) - irreducible_of_normalized_factor - (Associated.symm <| - calc - Multiset.prod (normalizedFactors a) ~ᵤ a := normalizedFactors_prod ha0 - _ = p * b := hb - _ ~ᵤ Multiset.prod (p ::ₘ normalizedFactors b) := by - rw [Multiset.prod_cons] - exact (normalizedFactors_prod hb0).symm.mul_left _ - ) - Multiset.exists_mem_of_rel_of_mem this (by simp) - -theorem exists_mem_normalizedFactors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : - ∃ p, p ∈ normalizedFactors x := by - obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx - obtain ⟨p, hp, _⟩ := exists_mem_normalizedFactors_of_dvd hx hp' hp'x - exact ⟨p, hp⟩ - -@[simp] -theorem normalizedFactors_zero : normalizedFactors (0 : α) = 0 := by - simp [normalizedFactors, factors] - -@[simp] -theorem normalizedFactors_one : normalizedFactors (1 : α) = 0 := by - cases' subsingleton_or_nontrivial α with h h - · dsimp [normalizedFactors, factors] - simp [Subsingleton.elim (1 : α) 0] - · rw [← Multiset.rel_zero_right] - apply factors_unique irreducible_of_normalized_factor - · intro x hx - exfalso - apply Multiset.not_mem_zero x hx - · apply normalizedFactors_prod one_ne_zero - -@[simp] -theorem normalizedFactors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - normalizedFactors (x * y) = normalizedFactors x + normalizedFactors y := by - have h : (normalize : α → α) = Associates.out ∘ Associates.mk := by - ext - rw [Function.comp_apply, Associates.out_mk] - rw [← Multiset.map_id' (normalizedFactors (x * y)), ← Multiset.map_id' (normalizedFactors x), ← - Multiset.map_id' (normalizedFactors y), ← Multiset.map_congr rfl normalize_normalized_factor, ← - Multiset.map_congr rfl normalize_normalized_factor, ← - Multiset.map_congr rfl normalize_normalized_factor, ← Multiset.map_add, h, ← - Multiset.map_map Associates.out, eq_comm, ← Multiset.map_map Associates.out] - refine congr rfl ?_ - apply Multiset.map_mk_eq_map_mk_of_rel - apply factors_unique - · intro x hx - rcases Multiset.mem_add.1 hx with (hx | hx) <;> exact irreducible_of_normalized_factor x hx - · exact irreducible_of_normalized_factor - · rw [Multiset.prod_add] - exact - ((normalizedFactors_prod hx).mul_mul (normalizedFactors_prod hy)).trans - (normalizedFactors_prod (mul_ne_zero hx hy)).symm - -@[simp] -theorem normalizedFactors_pow {x : α} (n : ℕ) : - normalizedFactors (x ^ n) = n • normalizedFactors x := by - induction' n with n ih - · simp - by_cases h0 : x = 0 - · simp [h0, zero_pow n.succ_ne_zero, smul_zero] - rw [pow_succ', succ_nsmul', normalizedFactors_mul h0 (pow_ne_zero _ h0), ih] - -theorem _root_.Irreducible.normalizedFactors_pow {p : α} (hp : Irreducible p) (k : ℕ) : - normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by - rw [UniqueFactorizationMonoid.normalizedFactors_pow, normalizedFactors_irreducible hp, - Multiset.nsmul_singleton] - -theorem normalizedFactors_prod_eq (s : Multiset α) (hs : ∀ a ∈ s, Irreducible a) : - normalizedFactors s.prod = s.map normalize := by - induction' s using Multiset.induction with a s ih - · rw [Multiset.prod_zero, normalizedFactors_one, Multiset.map_zero] - · have ia := hs a (Multiset.mem_cons_self a _) - have ib := fun b h => hs b (Multiset.mem_cons_of_mem h) - obtain rfl | ⟨b, hb⟩ := s.empty_or_exists_mem - · rw [Multiset.cons_zero, Multiset.prod_singleton, Multiset.map_singleton, - normalizedFactors_irreducible ia] - haveI := nontrivial_of_ne b 0 (ib b hb).ne_zero - rw [Multiset.prod_cons, Multiset.map_cons, - normalizedFactors_mul ia.ne_zero (Multiset.prod_ne_zero fun h => (ib 0 h).ne_zero rfl), - normalizedFactors_irreducible ia, ih ib, Multiset.singleton_add] - -theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - x ∣ y ↔ normalizedFactors x ≤ normalizedFactors y := by - constructor - · rintro ⟨c, rfl⟩ - simp [hx, right_ne_zero_of_mul hy] - · rw [← (normalizedFactors_prod hx).dvd_iff_dvd_left, ← - (normalizedFactors_prod hy).dvd_iff_dvd_right] - apply Multiset.prod_dvd_prod_of_le - -theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : - normalizedFactors a = normalizedFactors b := by - unfold normalizedFactors - have h' : ⇑(normalize (α := α)) = Associates.out ∘ Associates.mk := funext Associates.out_mk - rw [h', ← Multiset.map_map, ← Multiset.map_map, - Associates.rel_associated_iff_map_eq_map.mp (factors_rel_of_associated h)] - -theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := - ⟨Associated.normalizedFactors_eq, fun h => - (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ - -theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : - normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by - rw [normalizedFactors_pow, normalizedFactors_irreducible hp, Multiset.nsmul_singleton] - -theorem zero_not_mem_normalizedFactors (x : α) : (0 : α) ∉ normalizedFactors x := fun h => - Prime.ne_zero (prime_of_normalized_factor _ h) rfl - -theorem dvd_of_mem_normalizedFactors {a p : α} (H : p ∈ normalizedFactors a) : p ∣ a := by - by_cases hcases : a = 0 - · rw [hcases] - exact dvd_zero p - · exact dvd_trans (Multiset.dvd_prod H) (Associated.dvd (normalizedFactors_prod hcases)) - -theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) : - p ∈ normalizedFactors x ↔ Prime p ∧ p ∣ x := by - constructor - · intro h - exact ⟨prime_of_normalized_factor p h, dvd_of_mem_normalizedFactors h⟩ - · rintro ⟨hprime, hdvd⟩ - obtain ⟨q, hqmem, hqeq⟩ := exists_mem_normalizedFactors_of_dvd hx hprime.irreducible hdvd - rw [associated_iff_eq] at hqeq - exact hqeq ▸ hqmem - -theorem exists_associated_prime_pow_of_unique_normalized_factor {p r : α} - (h : ∀ {m}, m ∈ normalizedFactors r → m = p) (hr : r ≠ 0) : ∃ i : ℕ, Associated (p ^ i) r := by - use Multiset.card.toFun (normalizedFactors r) - have := UniqueFactorizationMonoid.normalizedFactors_prod hr - rwa [Multiset.eq_replicate_of_mem fun b => h, Multiset.prod_replicate] at this - -theorem normalizedFactors_prod_of_prime [Subsingleton αˣ] {m : Multiset α} - (h : ∀ p ∈ m, Prime p) : normalizedFactors m.prod = m := by - cases subsingleton_or_nontrivial α - · obtain rfl : m = 0 := by - refine Multiset.eq_zero_of_forall_not_mem fun x hx ↦ ?_ - simpa [Subsingleton.elim x 0] using h x hx - simp - · simpa only [← Multiset.rel_eq, ← associated_eq_eq] using - prime_factors_unique prime_of_normalized_factor h - (normalizedFactors_prod (m.prod_ne_zero_of_prime h)) - -theorem mem_normalizedFactors_eq_of_associated {a b c : α} (ha : a ∈ normalizedFactors c) - (hb : b ∈ normalizedFactors c) (h : Associated a b) : a = b := by - rw [← normalize_normalized_factor a ha, ← normalize_normalized_factor b hb, - normalize_eq_normalize_iff] - exact Associated.dvd_dvd h - -@[simp] -theorem normalizedFactors_pos (x : α) (hx : x ≠ 0) : 0 < normalizedFactors x ↔ ¬IsUnit x := by - constructor - · intro h hx - obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' - exact - (prime_of_normalized_factor _ hp).not_unit - (isUnit_of_dvd_unit (dvd_of_mem_normalizedFactors hp) hx) - · intro h - obtain ⟨p, hp⟩ := exists_mem_normalizedFactors hx h - exact - bot_lt_iff_ne_bot.mpr - (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) - -theorem dvdNotUnit_iff_normalizedFactors_lt_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : - DvdNotUnit x y ↔ normalizedFactors x < normalizedFactors y := by - constructor - · rintro ⟨_, c, hc, rfl⟩ - simp only [hx, right_ne_zero_of_mul hy, normalizedFactors_mul, Ne, not_false_iff, - lt_add_iff_pos_right, normalizedFactors_pos, hc] - · intro h - exact - dvdNotUnit_of_dvd_of_not_dvd - ((dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mpr h.le) - (mt (dvd_iff_normalizedFactors_le_normalizedFactors hy hx).mp h.not_le) - -theorem normalizedFactors_multiset_prod (s : Multiset α) (hs : 0 ∉ s) : - normalizedFactors (s.prod) = (s.map normalizedFactors).sum := by - cases subsingleton_or_nontrivial α - · obtain rfl : s = 0 := by - apply Multiset.eq_zero_of_forall_not_mem - intro _ - convert hs - simp - induction s using Multiset.induction with - | empty => simp - | cons _ _ IH => - rw [Multiset.prod_cons, Multiset.map_cons, Multiset.sum_cons, normalizedFactors_mul, IH] - · exact fun h ↦ hs (Multiset.mem_cons_of_mem h) - · exact fun h ↦ hs (h ▸ Multiset.mem_cons_self _ _) - · apply Multiset.prod_ne_zero - exact fun h ↦ hs (Multiset.mem_cons_of_mem h) - -end UniqueFactorizationMonoid - -namespace UniqueFactorizationMonoid - -open scoped Classical - -open Multiset Associates - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] - -/-- Noncomputably defines a `normalizationMonoid` structure on a `UniqueFactorizationMonoid`. -/ -protected noncomputable def normalizationMonoid : NormalizationMonoid α := - normalizationMonoidOfMonoidHomRightInverse - { toFun := fun a : Associates α => - if a = 0 then 0 - else - ((normalizedFactors a).map - (Classical.choose mk_surjective.hasRightInverse : Associates α → α)).prod - map_one' := by nontriviality α; simp - map_mul' := fun x y => by - by_cases hx : x = 0 - · simp [hx] - by_cases hy : y = 0 - · simp [hy] - simp [hx, hy] } - (by - intro x - dsimp - by_cases hx : x = 0 - · simp [hx] - have h : Associates.mkMonoidHom ∘ Classical.choose mk_surjective.hasRightInverse = - (id : Associates α → Associates α) := by - ext x - rw [Function.comp_apply, mkMonoidHom_apply, - Classical.choose_spec mk_surjective.hasRightInverse x] - rfl - rw [if_neg hx, ← mkMonoidHom_apply, MonoidHom.map_multiset_prod, map_map, h, map_id, ← - associated_iff_eq] - apply normalizedFactors_prod hx) - -end UniqueFactorizationMonoid - -namespace UniqueFactorizationMonoid - -variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] - -theorem isRelPrime_iff_no_prime_factors {a b : R} (ha : a ≠ 0) : - IsRelPrime a b ↔ ∀ ⦃d⦄, d ∣ a → d ∣ b → ¬Prime d := - ⟨fun h _ ha hb ↦ (·.not_unit <| h ha hb), fun h ↦ WfDvdMonoid.isRelPrime_of_no_irreducible_factors - (ha ·.1) fun _ irr ha hb ↦ h ha hb (UniqueFactorizationMonoid.irreducible_iff_prime.mp irr)⟩ - -/-- Euclid's lemma: if `a ∣ b * c` and `a` and `c` have no common prime factors, `a ∣ b`. -Compare `IsCoprime.dvd_of_dvd_mul_left`. -/ -theorem dvd_of_dvd_mul_left_of_no_prime_factors {a b c : R} (ha : a ≠ 0) - (h : ∀ ⦃d⦄, d ∣ a → d ∣ c → ¬Prime d) : a ∣ b * c → a ∣ b := - ((isRelPrime_iff_no_prime_factors ha).mpr h).dvd_of_dvd_mul_right - -/-- Euclid's lemma: if `a ∣ b * c` and `a` and `b` have no common prime factors, `a ∣ c`. -Compare `IsCoprime.dvd_of_dvd_mul_right`. -/ -theorem dvd_of_dvd_mul_right_of_no_prime_factors {a b c : R} (ha : a ≠ 0) - (no_factors : ∀ {d}, d ∣ a → d ∣ b → ¬Prime d) : a ∣ b * c → a ∣ c := by - simpa [mul_comm b c] using dvd_of_dvd_mul_left_of_no_prime_factors ha @no_factors - -/-- If `a ≠ 0, b` are elements of a unique factorization domain, then dividing -out their common factor `c'` gives `a'` and `b'` with no factors in common. -/ -theorem exists_reduced_factors : - ∀ a ≠ (0 : R), ∀ b, - ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := by - intro a - refine induction_on_prime a ?_ ?_ ?_ - · intros - contradiction - · intro a a_unit _ b - use a, b, 1 - constructor - · intro p p_dvd_a _ - exact isUnit_of_dvd_unit p_dvd_a a_unit - · simp - · intro a p a_ne_zero p_prime ih_a pa_ne_zero b - by_cases h : p ∣ b - · rcases h with ⟨b, rfl⟩ - obtain ⟨a', b', c', no_factor, ha', hb'⟩ := ih_a a_ne_zero b - refine ⟨a', b', p * c', @no_factor, ?_, ?_⟩ - · rw [mul_assoc, ha'] - · rw [mul_assoc, hb'] - · obtain ⟨a', b', c', coprime, rfl, rfl⟩ := ih_a a_ne_zero b - refine ⟨p * a', b', c', ?_, mul_left_comm _ _ _, rfl⟩ - intro q q_dvd_pa' q_dvd_b' - cases' p_prime.left_dvd_or_dvd_right_of_dvd_mul q_dvd_pa' with p_dvd_q q_dvd_a' - · have : p ∣ c' * b' := dvd_mul_of_dvd_right (p_dvd_q.trans q_dvd_b') _ - contradiction - exact coprime q_dvd_a' q_dvd_b' - -theorem exists_reduced_factors' (a b : R) (hb : b ≠ 0) : - ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := - let ⟨b', a', c', no_factor, hb, ha⟩ := exists_reduced_factors b hb a - ⟨a', b', c', fun _ hpb hpa => no_factor hpa hpb, ha, hb⟩ - -@[deprecated (since := "2024-09-21")] alias pow_right_injective := pow_injective_of_not_isUnit -@[deprecated (since := "2024-09-21")] alias pow_eq_pow_iff := pow_inj_of_not_isUnit - -section multiplicity - -variable [NormalizationMonoid R] - -open Multiset - -section - -theorem le_emultiplicity_iff_replicate_le_normalizedFactors {a b : R} {n : ℕ} (ha : Irreducible a) - (hb : b ≠ 0) : - ↑n ≤ emultiplicity a b ↔ replicate n (normalize a) ≤ normalizedFactors b := by - rw [← pow_dvd_iff_le_emultiplicity] - revert b - induction' n with n ih; · simp - intro b hb - constructor - · rintro ⟨c, rfl⟩ - rw [Ne, pow_succ', mul_assoc, mul_eq_zero, not_or] at hb - rw [pow_succ', mul_assoc, normalizedFactors_mul hb.1 hb.2, replicate_succ, - normalizedFactors_irreducible ha, singleton_add, cons_le_cons_iff, ← ih hb.2] - apply Dvd.intro _ rfl - · rw [Multiset.le_iff_exists_add] - rintro ⟨u, hu⟩ - rw [← (normalizedFactors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate] - exact (Associated.pow_pow <| associated_normalize a).dvd.trans (Dvd.intro u.prod rfl) - -/-- The multiplicity of an irreducible factor of a nonzero element is exactly the number of times -the normalized factor occurs in the `normalizedFactors`. - -See also `count_normalizedFactors_eq` which expands the definition of `multiplicity` -to produce a specification for `count (normalizedFactors _) _`.. --/ -theorem emultiplicity_eq_count_normalizedFactors [DecidableEq R] {a b : R} (ha : Irreducible a) - (hb : b ≠ 0) : emultiplicity a b = (normalizedFactors b).count (normalize a) := by - apply le_antisymm - · apply Order.le_of_lt_add_one - rw [← Nat.cast_one, ← Nat.cast_add, lt_iff_not_ge, ge_iff_le, - le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] - simp - rw [le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] - -end - -/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by -the number of times it divides `x`. - -See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. --/ -theorem count_normalizedFactors_eq [DecidableEq R] {p x : R} (hp : Irreducible p) - (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : - (normalizedFactors x).count p = n := by classical - by_cases hx0 : x = 0 - · simp [hx0] at hlt - apply Nat.cast_injective (R := ℕ∞) - convert (emultiplicity_eq_count_normalizedFactors hp hx0).symm - · exact hnorm.symm - exact (emultiplicity_eq_coe.mpr ⟨hle, hlt⟩).symm - -/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by -the number of times it divides `x`. This is a slightly more general version of -`UniqueFactorizationMonoid.count_normalizedFactors_eq` that allows `p = 0`. - -See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. --/ -theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Irreducible p) - (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : - (normalizedFactors x).count p = n := by - rcases hp with (rfl | hp) - · cases n - · exact count_eq_zero.2 (zero_not_mem_normalizedFactors _) - · rw [zero_pow (Nat.succ_ne_zero _)] at hle hlt - exact absurd hle hlt - · exact count_normalizedFactors_eq hp hnorm hle hlt - -end multiplicity - -/-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ -@[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] -theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : - ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx - -section Multiplicative - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] -variable {β : Type*} [CancelCommMonoidWithZero β] - -theorem prime_pow_coprime_prod_of_coprime_insert [DecidableEq α] {s : Finset α} (i : α → ℕ) (p : α) - (hps : p ∉ s) (is_prime : ∀ q ∈ insert p s, Prime q) - (is_coprime : ∀ᵉ (q ∈ insert p s) (q' ∈ insert p s), q ∣ q' → q = q') : - IsRelPrime (p ^ i p) (∏ p' ∈ s, p' ^ i p') := by - have hp := is_prime _ (Finset.mem_insert_self _ _) - refine (isRelPrime_iff_no_prime_factors <| pow_ne_zero _ hp.ne_zero).mpr ?_ - intro d hdp hdprod hd - apply hps - replace hdp := hd.dvd_of_dvd_pow hdp - obtain ⟨q, q_mem', hdq⟩ := hd.exists_mem_multiset_dvd hdprod - obtain ⟨q, q_mem, rfl⟩ := Multiset.mem_map.mp q_mem' - replace hdq := hd.dvd_of_dvd_pow hdq - have : p ∣ q := dvd_trans (hd.irreducible.dvd_symm hp.irreducible hdp) hdq - convert q_mem using 0 - rw [Finset.mem_val, - is_coprime _ (Finset.mem_insert_self p s) _ (Finset.mem_insert_of_mem q_mem) this] - -/-- If `P` holds for units and powers of primes, -and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, -then `P` holds on a product of powers of distinct primes. -/ -@[elab_as_elim] -theorem induction_on_prime_power {P : α → Prop} (s : Finset α) (i : α → ℕ) - (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) - (h1 : ∀ {x}, IsUnit x → P x) (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) - (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : - P (∏ p ∈ s, p ^ i p) := by - letI := Classical.decEq α - induction' s using Finset.induction_on with p f' hpf' ih - · simpa using h1 isUnit_one - rw [Finset.prod_insert hpf'] - exact - hcp (prime_pow_coprime_prod_of_coprime_insert i p hpf' is_prime is_coprime) - (hpr (i p) (is_prime _ (Finset.mem_insert_self _ _))) - (ih (fun q hq => is_prime _ (Finset.mem_insert_of_mem hq)) fun q hq q' hq' => - is_coprime _ (Finset.mem_insert_of_mem hq) _ (Finset.mem_insert_of_mem hq')) - -/-- If `P` holds for `0`, units and powers of primes, -and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, -then `P` holds on all `a : α`. -/ -@[elab_as_elim] -theorem induction_on_coprime {P : α → Prop} (a : α) (h0 : P 0) (h1 : ∀ {x}, IsUnit x → P x) - (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) - (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : P a := by - letI := Classical.decEq α - have P_of_associated : ∀ {x y}, Associated x y → P x → P y := by - rintro x y ⟨u, rfl⟩ hx - exact hcp (fun p _ hpx => isUnit_of_dvd_unit hpx u.isUnit) hx (h1 u.isUnit) - by_cases ha0 : a = 0 - · rwa [ha0] - haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ - letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid - refine P_of_associated (normalizedFactors_prod ha0) ?_ - rw [← (normalizedFactors a).map_id, Finset.prod_multiset_map_count] - refine induction_on_prime_power _ _ ?_ ?_ @h1 @hpr @hcp <;> simp only [Multiset.mem_toFinset] - · apply prime_of_normalized_factor - · apply normalizedFactors_eq_of_dvd - -/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` -is multiplicative on coprime elements, then `f` is multiplicative on all products of primes. -/ -theorem multiplicative_prime_power {f : α → β} (s : Finset α) (i j : α → ℕ) - (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) - (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) - (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) - (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : - f (∏ p ∈ s, p ^ (i p + j p)) = f (∏ p ∈ s, p ^ i p) * f (∏ p ∈ s, p ^ j p) := by - letI := Classical.decEq α - induction' s using Finset.induction_on with p s hps ih - · simpa using h1 isUnit_one - have hpr_p := is_prime _ (Finset.mem_insert_self _ _) - have hpr_s : ∀ p ∈ s, Prime p := fun p hp => is_prime _ (Finset.mem_insert_of_mem hp) - have hcp_p := fun i => prime_pow_coprime_prod_of_coprime_insert i p hps is_prime is_coprime - have hcp_s : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q := fun p hp q hq => - is_coprime p (Finset.mem_insert_of_mem hp) q (Finset.mem_insert_of_mem hq) - rw [Finset.prod_insert hps, Finset.prod_insert hps, Finset.prod_insert hps, hcp (hcp_p _), - hpr _ hpr_p, hcp (hcp_p _), hpr _ hpr_p, hcp (hcp_p (fun p => i p + j p)), hpr _ hpr_p, - ih hpr_s hcp_s, pow_add, mul_assoc, mul_left_comm (f p ^ j p), mul_assoc] - -/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` -is multiplicative on coprime elements, then `f` is multiplicative everywhere. -/ -theorem multiplicative_of_coprime (f : α → β) (a b : α) (h0 : f 0 = 0) - (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) - (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) - (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : - f (a * b) = f a * f b := by - letI := Classical.decEq α - by_cases ha0 : a = 0 - · rw [ha0, zero_mul, h0, zero_mul] - by_cases hb0 : b = 0 - · rw [hb0, mul_zero, h0, mul_zero] - by_cases hf1 : f 1 = 0 - · calc - f (a * b) = f (a * b * 1) := by rw [mul_one] - _ = 0 := by simp only [h1 isUnit_one, hf1, mul_zero] - _ = f a * f (b * 1) := by simp only [h1 isUnit_one, hf1, mul_zero] - _ = f a * f b := by rw [mul_one] - haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ - letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid - suffices - f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, - p ^ ((normalizedFactors a).count p + (normalizedFactors b).count p)) = - f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, - p ^ (normalizedFactors a).count p) * - f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, - p ^ (normalizedFactors b).count p) by - obtain ⟨ua, a_eq⟩ := normalizedFactors_prod ha0 - obtain ⟨ub, b_eq⟩ := normalizedFactors_prod hb0 - rw [← a_eq, ← b_eq, mul_right_comm (Multiset.prod (normalizedFactors a)) ua - (Multiset.prod (normalizedFactors b) * ub), h1 ua.isUnit, h1 ub.isUnit, h1 ua.isUnit, ← - mul_assoc, h1 ub.isUnit, mul_right_comm _ (f ua), ← mul_assoc] - congr - rw [← (normalizedFactors a).map_id, ← (normalizedFactors b).map_id, - Finset.prod_multiset_map_count, Finset.prod_multiset_map_count, - Finset.prod_subset (Finset.subset_union_left (s₂ := (normalizedFactors b).toFinset)), - Finset.prod_subset (Finset.subset_union_right (s₂ := (normalizedFactors b).toFinset)), ← - Finset.prod_mul_distrib] - · simp_rw [id, ← pow_add, this] - all_goals simp only [Multiset.mem_toFinset] - · intro p _ hpb - simp [hpb] - · intro p _ hpa - simp [hpa] - refine multiplicative_prime_power _ _ _ ?_ ?_ @h1 @hpr @hcp - all_goals simp only [Multiset.mem_toFinset, Finset.mem_union] - · rintro p (hpa | hpb) <;> apply prime_of_normalized_factor <;> assumption - · rintro p (hp | hp) q (hq | hq) hdvd <;> - rw [← normalize_normalized_factor _ hp, ← normalize_normalized_factor _ hq] <;> - exact - normalize_eq_normalize hdvd - ((prime_of_normalized_factor _ hp).irreducible.dvd_symm - (prime_of_normalized_factor _ hq).irreducible hdvd) - -end Multiplicative - -end UniqueFactorizationMonoid - -namespace Associates - -open UniqueFactorizationMonoid Associated Multiset - -variable [CancelCommMonoidWithZero α] - -/-- `FactorSet α` representation elements of unique factorization domain as multisets. -`Multiset α` produced by `normalizedFactors` are only unique up to associated elements, while the -multisets in `FactorSet α` are unique by equality and restricted to irreducible elements. This -gives us a representation of each element as a unique multisets (or the added ⊤ for 0), which has a -complete lattice structure. Infimum is the greatest common divisor and supremum is the least common -multiple. --/ -abbrev FactorSet.{u} (α : Type u) [CancelCommMonoidWithZero α] : Type u := - WithTop (Multiset { a : Associates α // Irreducible a }) - -attribute [local instance] Associated.setoid - -theorem FactorSet.coe_add {a b : Multiset { a : Associates α // Irreducible a }} : - (↑(a + b) : FactorSet α) = a + b := by norm_cast - -theorem FactorSet.sup_add_inf_eq_add [DecidableEq (Associates α)] : - ∀ a b : FactorSet α, a ⊔ b + a ⊓ b = a + b - | ⊤, b => show ⊤ ⊔ b + ⊤ ⊓ b = ⊤ + b by simp - | a, ⊤ => show a ⊔ ⊤ + a ⊓ ⊤ = a + ⊤ by simp - | WithTop.some a, WithTop.some b => - show (a : FactorSet α) ⊔ b + (a : FactorSet α) ⊓ b = a + b by - rw [← WithTop.coe_sup, ← WithTop.coe_inf, ← WithTop.coe_add, ← WithTop.coe_add, - WithTop.coe_eq_coe] - exact Multiset.union_add_inter _ _ - -/-- Evaluates the product of a `FactorSet` to be the product of the corresponding multiset, - or `0` if there is none. -/ -def FactorSet.prod : FactorSet α → Associates α - | ⊤ => 0 - | WithTop.some s => (s.map (↑)).prod - -@[simp] -theorem prod_top : (⊤ : FactorSet α).prod = 0 := - rfl - -@[simp] -theorem prod_coe {s : Multiset { a : Associates α // Irreducible a }} : - FactorSet.prod (s : FactorSet α) = (s.map (↑)).prod := - rfl - -@[simp] -theorem prod_add : ∀ a b : FactorSet α, (a + b).prod = a.prod * b.prod - | ⊤, b => show (⊤ + b).prod = (⊤ : FactorSet α).prod * b.prod by simp - | a, ⊤ => show (a + ⊤).prod = a.prod * (⊤ : FactorSet α).prod by simp - | WithTop.some a, WithTop.some b => by - rw [← FactorSet.coe_add, prod_coe, prod_coe, prod_coe, Multiset.map_add, Multiset.prod_add] - -@[gcongr] -theorem prod_mono : ∀ {a b : FactorSet α}, a ≤ b → a.prod ≤ b.prod - | ⊤, b, h => by - have : b = ⊤ := top_unique h - rw [this, prod_top] - | a, ⊤, _ => show a.prod ≤ (⊤ : FactorSet α).prod by simp - | WithTop.some _, WithTop.some _, h => - prod_le_prod <| Multiset.map_le_map <| WithTop.coe_le_coe.1 <| h - -theorem FactorSet.prod_eq_zero_iff [Nontrivial α] (p : FactorSet α) : p.prod = 0 ↔ p = ⊤ := by - unfold FactorSet at p - induction p -- TODO: `induction_eliminator` doesn't work with `abbrev` - · simp only [eq_self_iff_true, Associates.prod_top] - · rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, - iff_false, not_exists] - exact fun a => not_and_of_not_right _ a.prop.ne_zero - -section count - -variable [DecidableEq (Associates α)] - -/-- `bcount p s` is the multiplicity of `p` in the FactorSet `s` (with bundled `p`)-/ -def bcount (p : { a : Associates α // Irreducible a }) : - FactorSet α → ℕ - | ⊤ => 0 - | WithTop.some s => s.count p - -variable [∀ p : Associates α, Decidable (Irreducible p)] {p : Associates α} - -/-- `count p s` is the multiplicity of the irreducible `p` in the FactorSet `s`. - -If `p` is not irreducible, `count p s` is defined to be `0`. -/ -def count (p : Associates α) : FactorSet α → ℕ := - if hp : Irreducible p then bcount ⟨p, hp⟩ else 0 - -@[simp] -theorem count_some (hp : Irreducible p) (s : Multiset _) : - count p (WithTop.some s) = s.count ⟨p, hp⟩ := by - simp only [count, dif_pos hp, bcount] - -@[simp] -theorem count_zero (hp : Irreducible p) : count p (0 : FactorSet α) = 0 := by - simp only [count, dif_pos hp, bcount, Multiset.count_zero] - -theorem count_reducible (hp : ¬Irreducible p) : count p = 0 := dif_neg hp - -end count - -section Mem - -/-- membership in a FactorSet (bundled version) -/ -def BfactorSetMem : { a : Associates α // Irreducible a } → FactorSet α → Prop - | _, ⊤ => True - | p, some l => p ∈ l - -/-- `FactorSetMem p s` is the predicate that the irreducible `p` is a member of -`s : FactorSet α`. - -If `p` is not irreducible, `p` is not a member of any `FactorSet`. -/ -def FactorSetMem (s : FactorSet α) (p : Associates α) : Prop := - letI : Decidable (Irreducible p) := Classical.dec _ - if hp : Irreducible p then BfactorSetMem ⟨p, hp⟩ s else False - -instance : Membership (Associates α) (FactorSet α) := - ⟨FactorSetMem⟩ - -@[simp] -theorem factorSetMem_eq_mem (p : Associates α) (s : FactorSet α) : FactorSetMem s p = (p ∈ s) := - rfl - -theorem mem_factorSet_top {p : Associates α} {hp : Irreducible p} : p ∈ (⊤ : FactorSet α) := by - dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; exact trivial - -theorem mem_factorSet_some {p : Associates α} {hp : Irreducible p} - {l : Multiset { a : Associates α // Irreducible a }} : - p ∈ (l : FactorSet α) ↔ Subtype.mk p hp ∈ l := by - dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; rfl - -theorem reducible_not_mem_factorSet {p : Associates α} (hp : ¬Irreducible p) (s : FactorSet α) : - ¬p ∈ s := fun h ↦ by - rwa [← factorSetMem_eq_mem, FactorSetMem, dif_neg hp] at h - -theorem irreducible_of_mem_factorSet {p : Associates α} {s : FactorSet α} (h : p ∈ s) : - Irreducible p := - by_contra fun hp ↦ reducible_not_mem_factorSet hp s h - -end Mem - -variable [UniqueFactorizationMonoid α] - -theorem unique' {p q : Multiset (Associates α)} : - (∀ a ∈ p, Irreducible a) → (∀ a ∈ q, Irreducible a) → p.prod = q.prod → p = q := by - apply Multiset.induction_on_multiset_quot p - apply Multiset.induction_on_multiset_quot q - intro s t hs ht eq - refine Multiset.map_mk_eq_map_mk_of_rel (UniqueFactorizationMonoid.factors_unique ?_ ?_ ?_) - · exact fun a ha => irreducible_mk.1 <| hs _ <| Multiset.mem_map_of_mem _ ha - · exact fun a ha => irreducible_mk.1 <| ht _ <| Multiset.mem_map_of_mem _ ha - have eq' : (Quot.mk Setoid.r : α → Associates α) = Associates.mk := funext quot_mk_eq_mk - rwa [eq', prod_mk, prod_mk, mk_eq_mk_iff_associated] at eq - -theorem FactorSet.unique [Nontrivial α] {p q : FactorSet α} (h : p.prod = q.prod) : p = q := by - -- TODO: `induction_eliminator` doesn't work with `abbrev` - unfold FactorSet at p q - induction p <;> induction q - · rfl - · rw [eq_comm, ← FactorSet.prod_eq_zero_iff, ← h, Associates.prod_top] - · rw [← FactorSet.prod_eq_zero_iff, h, Associates.prod_top] - · congr 1 - rw [← Multiset.map_eq_map Subtype.coe_injective] - apply unique' _ _ h <;> - · intro a ha - obtain ⟨⟨a', irred⟩, -, rfl⟩ := Multiset.mem_map.mp ha - rwa [Subtype.coe_mk] - -theorem prod_le_prod_iff_le [Nontrivial α] {p q : Multiset (Associates α)} - (hp : ∀ a ∈ p, Irreducible a) (hq : ∀ a ∈ q, Irreducible a) : p.prod ≤ q.prod ↔ p ≤ q := by - refine ⟨?_, prod_le_prod⟩ - rintro ⟨c, eqc⟩ - refine Multiset.le_iff_exists_add.2 ⟨factors c, unique' hq (fun x hx ↦ ?_) ?_⟩ - · obtain h | h := Multiset.mem_add.1 hx - · exact hp x h - · exact irreducible_of_factor _ h - · rw [eqc, Multiset.prod_add] - congr - refine associated_iff_eq.mp (factors_prod fun hc => ?_).symm - refine not_irreducible_zero (hq _ ?_) - rw [← prod_eq_zero_iff, eqc, hc, mul_zero] - -/-- This returns the multiset of irreducible factors as a `FactorSet`, - a multiset of irreducible associates `WithTop`. -/ -noncomputable def factors' (a : α) : Multiset { a : Associates α // Irreducible a } := - (factors a).pmap (fun a ha => ⟨Associates.mk a, irreducible_mk.2 ha⟩) irreducible_of_factor - -@[simp] -theorem map_subtype_coe_factors' {a : α} : - (factors' a).map (↑) = (factors a).map Associates.mk := by - simp [factors', Multiset.map_pmap, Multiset.pmap_eq_map] - -theorem factors'_cong {a b : α} (h : a ~ᵤ b) : factors' a = factors' b := by - obtain rfl | hb := eq_or_ne b 0 - · rw [associated_zero_iff_eq_zero] at h - rw [h] - have ha : a ≠ 0 := by - contrapose! hb with ha - rw [← associated_zero_iff_eq_zero, ← ha] - exact h.symm - rw [← Multiset.map_eq_map Subtype.coe_injective, map_subtype_coe_factors', - map_subtype_coe_factors', ← rel_associated_iff_map_eq_map] - exact - factors_unique irreducible_of_factor irreducible_of_factor - ((factors_prod ha).trans <| h.trans <| (factors_prod hb).symm) - -/-- This returns the multiset of irreducible factors of an associate as a `FactorSet`, - a multiset of irreducible associates `WithTop`. -/ -noncomputable def factors (a : Associates α) : FactorSet α := by - classical refine if h : a = 0 then ⊤ else Quotient.hrecOn a (fun x _ => factors' x) ?_ h - intro a b hab - apply Function.hfunext - · have : a ~ᵤ 0 ↔ b ~ᵤ 0 := Iff.intro (fun ha0 => hab.symm.trans ha0) fun hb0 => hab.trans hb0 - simp only [associated_zero_iff_eq_zero] at this - simp only [quotient_mk_eq_mk, this, mk_eq_zero] - exact fun ha hb _ => heq_of_eq <| congr_arg some <| factors'_cong hab - -@[simp] -theorem factors_zero : (0 : Associates α).factors = ⊤ := - dif_pos rfl - -@[deprecated (since := "2024-03-16")] alias factors_0 := factors_zero - -@[simp] -theorem factors_mk (a : α) (h : a ≠ 0) : (Associates.mk a).factors = factors' a := by - classical - apply dif_neg - apply mt mk_eq_zero.1 h - -@[simp] -theorem factors_prod (a : Associates α) : a.factors.prod = a := by - rcases Associates.mk_surjective a with ⟨a, rfl⟩ - rcases eq_or_ne a 0 with rfl | ha - · simp - · simp [ha, prod_mk, mk_eq_mk_iff_associated, UniqueFactorizationMonoid.factors_prod, - -Quotient.eq] - -@[simp] -theorem prod_factors [Nontrivial α] (s : FactorSet α) : s.prod.factors = s := - FactorSet.unique <| factors_prod _ - -@[nontriviality] -theorem factors_subsingleton [Subsingleton α] {a : Associates α} : a.factors = ⊤ := by - have : Subsingleton (Associates α) := inferInstance - convert factors_zero - -theorem factors_eq_top_iff_zero {a : Associates α} : a.factors = ⊤ ↔ a = 0 := by - nontriviality α - exact ⟨fun h ↦ by rwa [← factors_prod a, FactorSet.prod_eq_zero_iff], fun h ↦ h ▸ factors_zero⟩ - -@[deprecated (since := "2024-04-16")] alias factors_eq_none_iff_zero := factors_eq_top_iff_zero - -theorem factors_eq_some_iff_ne_zero {a : Associates α} : - (∃ s : Multiset { p : Associates α // Irreducible p }, a.factors = s) ↔ a ≠ 0 := by - simp_rw [@eq_comm _ a.factors, ← WithTop.ne_top_iff_exists] - exact factors_eq_top_iff_zero.not - -theorem eq_of_factors_eq_factors {a b : Associates α} (h : a.factors = b.factors) : a = b := by - have : a.factors.prod = b.factors.prod := by rw [h] - rwa [factors_prod, factors_prod] at this - -theorem eq_of_prod_eq_prod [Nontrivial α] {a b : FactorSet α} (h : a.prod = b.prod) : a = b := by - have : a.prod.factors = b.prod.factors := by rw [h] - rwa [prod_factors, prod_factors] at this - -@[simp] -theorem factors_mul (a b : Associates α) : (a * b).factors = a.factors + b.factors := by - nontriviality α - refine eq_of_prod_eq_prod <| eq_of_factors_eq_factors ?_ - rw [prod_add, factors_prod, factors_prod, factors_prod] - -@[gcongr] -theorem factors_mono : ∀ {a b : Associates α}, a ≤ b → a.factors ≤ b.factors - | s, t, ⟨d, eq⟩ => by rw [eq, factors_mul]; exact le_add_of_nonneg_right bot_le - -@[simp] -theorem factors_le {a b : Associates α} : a.factors ≤ b.factors ↔ a ≤ b := by - refine ⟨fun h ↦ ?_, factors_mono⟩ - have : a.factors.prod ≤ b.factors.prod := prod_mono h - rwa [factors_prod, factors_prod] at this - -section count - -variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] - -theorem eq_factors_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) - (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : - a.factors = b.factors := by - obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha - obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb - rw [h_sa, h_sb] at h ⊢ - rw [WithTop.coe_eq_coe] - have h_count : ∀ (p : Associates α) (hp : Irreducible p), - sa.count ⟨p, hp⟩ = sb.count ⟨p, hp⟩ := by - intro p hp - rw [← count_some, ← count_some, h p hp] - apply Multiset.toFinsupp.injective - ext ⟨p, hp⟩ - rw [Multiset.toFinsupp_apply, Multiset.toFinsupp_apply, h_count p hp] - -theorem eq_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) - (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : a = b := - eq_of_factors_eq_factors (eq_factors_of_eq_counts ha hb h) - -theorem count_le_count_of_factors_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) - (h : a.factors ≤ b.factors) : p.count a.factors ≤ p.count b.factors := by - by_cases ha : a = 0 - · simp_all - obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha - obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb - rw [h_sa, h_sb] at h ⊢ - rw [count_some hp, count_some hp]; rw [WithTop.coe_le_coe] at h - exact Multiset.count_le_of_le _ h - -theorem count_le_count_of_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) (h : a ≤ b) : - p.count a.factors ≤ p.count b.factors := - count_le_count_of_factors_le hb hp <| factors_mono h - -end count - -theorem prod_le [Nontrivial α] {a b : FactorSet α} : a.prod ≤ b.prod ↔ a ≤ b := by - refine ⟨fun h ↦ ?_, prod_mono⟩ - have : a.prod.factors ≤ b.prod.factors := factors_mono h - rwa [prod_factors, prod_factors] at this - -open Classical in -noncomputable instance : Max (Associates α) := - ⟨fun a b => (a.factors ⊔ b.factors).prod⟩ - -open Classical in -noncomputable instance : Min (Associates α) := - ⟨fun a b => (a.factors ⊓ b.factors).prod⟩ - -open Classical in -noncomputable instance : Lattice (Associates α) := - { Associates.instPartialOrder with - sup := (· ⊔ ·) - inf := (· ⊓ ·) - sup_le := fun _ _ c hac hbc => - factors_prod c ▸ prod_mono (sup_le (factors_mono hac) (factors_mono hbc)) - le_sup_left := fun a _ => le_trans (le_of_eq (factors_prod a).symm) <| prod_mono <| le_sup_left - le_sup_right := fun _ b => - le_trans (le_of_eq (factors_prod b).symm) <| prod_mono <| le_sup_right - le_inf := fun a _ _ hac hbc => - factors_prod a ▸ prod_mono (le_inf (factors_mono hac) (factors_mono hbc)) - inf_le_left := fun a _ => le_trans (prod_mono inf_le_left) (le_of_eq (factors_prod a)) - inf_le_right := fun _ b => le_trans (prod_mono inf_le_right) (le_of_eq (factors_prod b)) } - -open Classical in -theorem sup_mul_inf (a b : Associates α) : (a ⊔ b) * (a ⊓ b) = a * b := - show (a.factors ⊔ b.factors).prod * (a.factors ⊓ b.factors).prod = a * b by - nontriviality α - refine eq_of_factors_eq_factors ?_ - rw [← prod_add, prod_factors, factors_mul, FactorSet.sup_add_inf_eq_add] - -theorem dvd_of_mem_factors {a p : Associates α} (hm : p ∈ factors a) : - p ∣ a := by - rcases eq_or_ne a 0 with rfl | ha0 - · exact dvd_zero p - obtain ⟨a0, nza, ha'⟩ := exists_non_zero_rep ha0 - rw [← Associates.factors_prod a] - rw [← ha', factors_mk a0 nza] at hm ⊢ - rw [prod_coe] - apply Multiset.dvd_prod; apply Multiset.mem_map.mpr - exact ⟨⟨p, irreducible_of_mem_factorSet hm⟩, mem_factorSet_some.mp hm, rfl⟩ - -theorem dvd_of_mem_factors' {a : α} {p : Associates α} {hp : Irreducible p} {hz : a ≠ 0} - (h_mem : Subtype.mk p hp ∈ factors' a) : p ∣ Associates.mk a := by - haveI := Classical.decEq (Associates α) - apply dvd_of_mem_factors - rw [factors_mk _ hz] - apply mem_factorSet_some.2 h_mem - -theorem mem_factors'_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : - Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a := by - obtain ⟨q, hq, hpq⟩ := exists_mem_factors_of_dvd ha0 hp hd - apply Multiset.mem_pmap.mpr; use q; use hq - exact Subtype.eq (Eq.symm (mk_eq_mk_iff_associated.mpr hpq)) - -theorem mem_factors'_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a ↔ p ∣ a := by - constructor - · rw [← mk_dvd_mk] - apply dvd_of_mem_factors' - apply ha0 - · apply mem_factors'_of_dvd ha0 hp - -theorem mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : - Associates.mk p ∈ factors (Associates.mk a) := by - rw [factors_mk _ ha0] - exact mem_factorSet_some.mpr (mem_factors'_of_dvd ha0 hp hd) - -theorem mem_factors_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - Associates.mk p ∈ factors (Associates.mk a) ↔ p ∣ a := by - constructor - · rw [← mk_dvd_mk] - apply dvd_of_mem_factors - · apply mem_factors_of_dvd ha0 hp - -open Classical in -theorem exists_prime_dvd_of_not_inf_one {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) - (h : Associates.mk a ⊓ Associates.mk b ≠ 1) : ∃ p : α, Prime p ∧ p ∣ a ∧ p ∣ b := by - have hz : factors (Associates.mk a) ⊓ factors (Associates.mk b) ≠ 0 := by - contrapose! h with hf - change (factors (Associates.mk a) ⊓ factors (Associates.mk b)).prod = 1 - rw [hf] - exact Multiset.prod_zero - rw [factors_mk a ha, factors_mk b hb, ← WithTop.coe_inf] at hz - obtain ⟨⟨p0, p0_irr⟩, p0_mem⟩ := Multiset.exists_mem_of_ne_zero ((mt WithTop.coe_eq_coe.mpr) hz) - rw [Multiset.inf_eq_inter] at p0_mem - obtain ⟨p, rfl⟩ : ∃ p, Associates.mk p = p0 := Quot.exists_rep p0 - refine ⟨p, ?_, ?_, ?_⟩ - · rw [← UniqueFactorizationMonoid.irreducible_iff_prime, ← irreducible_mk] - exact p0_irr - · apply dvd_of_mk_le_mk - apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).left - apply ha - · apply dvd_of_mk_le_mk - apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).right - apply hb - -theorem coprime_iff_inf_one {a b : α} (ha0 : a ≠ 0) (hb0 : b ≠ 0) : - Associates.mk a ⊓ Associates.mk b = 1 ↔ ∀ {d : α}, d ∣ a → d ∣ b → ¬Prime d := by - constructor - · intro hg p ha hb hp - refine (Associates.prime_mk.mpr hp).not_unit (isUnit_of_dvd_one ?_) - rw [← hg] - exact le_inf (mk_le_mk_of_dvd ha) (mk_le_mk_of_dvd hb) - · contrapose - intro hg hc - obtain ⟨p, hp, hpa, hpb⟩ := exists_prime_dvd_of_not_inf_one ha0 hb0 hg - exact hc hpa hpb hp - -theorem factors_self [Nontrivial α] {p : Associates α} (hp : Irreducible p) : - p.factors = WithTop.some {⟨p, hp⟩} := - eq_of_prod_eq_prod - (by rw [factors_prod, FactorSet.prod.eq_def]; dsimp; rw [prod_singleton]) - -theorem factors_prime_pow [Nontrivial α] {p : Associates α} (hp : Irreducible p) (k : ℕ) : - factors (p ^ k) = WithTop.some (Multiset.replicate k ⟨p, hp⟩) := - eq_of_prod_eq_prod - (by - rw [Associates.factors_prod, FactorSet.prod.eq_def] - dsimp; rw [Multiset.map_replicate, Multiset.prod_replicate, Subtype.coe_mk]) - -theorem prime_pow_le_iff_le_bcount [DecidableEq (Associates α)] {m p : Associates α} - (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : p ^ k ≤ m ↔ k ≤ bcount ⟨p, h₂⟩ m.factors := by - rcases Associates.exists_non_zero_rep h₁ with ⟨m, hm, rfl⟩ - have := nontrivial_of_ne _ _ hm - rw [bcount.eq_def, factors_mk, Multiset.le_count_iff_replicate_le, ← factors_le, - factors_prime_pow, factors_mk, WithTop.coe_le_coe] <;> assumption - -@[simp] -theorem factors_one [Nontrivial α] : factors (1 : Associates α) = 0 := by - apply eq_of_prod_eq_prod - rw [Associates.factors_prod] - exact Multiset.prod_zero - -@[simp] -theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : - (a ^ k).factors = k • a.factors := by - induction' k with n h - · rw [zero_nsmul, pow_zero] - exact factors_one - · rw [pow_succ, succ_nsmul, factors_mul, h] - -section count - -variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] - -theorem prime_pow_dvd_iff_le {m p : Associates α} (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : - p ^ k ≤ m ↔ k ≤ count p m.factors := by - rw [count, dif_pos h₂, prime_pow_le_iff_le_bcount h₁] - -theorem le_of_count_ne_zero {m p : Associates α} (h0 : m ≠ 0) (hp : Irreducible p) : - count p m.factors ≠ 0 → p ≤ m := by - nontriviality α - rw [← pos_iff_ne_zero] - intro h - rw [← pow_one p] - apply (prime_pow_dvd_iff_le h0 hp).2 - simpa only - -theorem count_ne_zero_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : - (Associates.mk p).count (Associates.mk a).factors ≠ 0 ↔ p ∣ a := by - nontriviality α - rw [← Associates.mk_le_mk_iff_dvd] - refine - ⟨fun h => - Associates.le_of_count_ne_zero (Associates.mk_ne_zero.mpr ha0) - (Associates.irreducible_mk.mpr hp) h, - fun h => ?_⟩ - rw [← pow_one (Associates.mk p), - Associates.prime_pow_dvd_iff_le (Associates.mk_ne_zero.mpr ha0) - (Associates.irreducible_mk.mpr hp)] at h - exact (zero_lt_one.trans_le h).ne' - -theorem count_self [Nontrivial α] {p : Associates α} - (hp : Irreducible p) : p.count p.factors = 1 := by - simp [factors_self hp, Associates.count_some hp] - -theorem count_eq_zero_of_ne {p q : Associates α} (hp : Irreducible p) - (hq : Irreducible q) (h : p ≠ q) : p.count q.factors = 0 := - not_ne_iff.mp fun h' ↦ h <| associated_iff_eq.mp <| hp.associated_of_dvd hq <| - le_of_count_ne_zero hq.ne_zero hp h' - -theorem count_mul {a : Associates α} (ha : a ≠ 0) {b : Associates α} - (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) : - count p (factors (a * b)) = count p a.factors + count p b.factors := by - obtain ⟨a0, nza, rfl⟩ := exists_non_zero_rep ha - obtain ⟨b0, nzb, rfl⟩ := exists_non_zero_rep hb - rw [factors_mul, factors_mk a0 nza, factors_mk b0 nzb, ← FactorSet.coe_add, count_some hp, - Multiset.count_add, count_some hp, count_some hp] - -theorem count_of_coprime {a : Associates α} (ha : a ≠ 0) - {b : Associates α} (hb : b ≠ 0) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {p : Associates α} - (hp : Irreducible p) : count p a.factors = 0 ∨ count p b.factors = 0 := by - rw [or_iff_not_imp_left, ← Ne] - intro hca - contrapose! hab with hcb - exact ⟨p, le_of_count_ne_zero ha hp hca, le_of_count_ne_zero hb hp hcb, - UniqueFactorizationMonoid.irreducible_iff_prime.mp hp⟩ - -theorem count_mul_of_coprime {a : Associates α} {b : Associates α} - (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : - count p a.factors = 0 ∨ count p a.factors = count p (a * b).factors := by - by_cases ha : a = 0 - · simp [ha] - cases' count_of_coprime ha hb hab hp with hz hb0; · tauto - apply Or.intro_right - rw [count_mul ha hb hp, hb0, add_zero] - -theorem count_mul_of_coprime' {a b : Associates α} {p : Associates α} - (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : - count p (a * b).factors = count p a.factors ∨ count p (a * b).factors = count p b.factors := by - by_cases ha : a = 0 - · simp [ha] - by_cases hb : b = 0 - · simp [hb] - rw [count_mul ha hb hp] - cases' count_of_coprime ha hb hab hp with ha0 hb0 - · apply Or.intro_right - rw [ha0, zero_add] - · apply Or.intro_left - rw [hb0, add_zero] - -theorem dvd_count_of_dvd_count_mul {a b : Associates α} (hb : b ≠ 0) - {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} - (habk : k ∣ count p (a * b).factors) : k ∣ count p a.factors := by - by_cases ha : a = 0 - · simpa [*] using habk - cases' count_of_coprime ha hb hab hp with hz h - · rw [hz] - exact dvd_zero k - · rw [count_mul ha hb hp, h] at habk - exact habk - -theorem count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) - {p : Associates α} (hp : Irreducible p) (k : ℕ) : - count p (a ^ k).factors = k * count p a.factors := by - induction' k with n h - · rw [pow_zero, factors_one, zero_mul, count_zero hp] - · rw [pow_succ', count_mul ha (pow_ne_zero _ ha) hp, h] - ring - -theorem dvd_count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) - {p : Associates α} (hp : Irreducible p) (k : ℕ) : k ∣ count p (a ^ k).factors := by - rw [count_pow ha hp] - apply dvd_mul_right - -theorem is_pow_of_dvd_count {a : Associates α} - (ha : a ≠ 0) {k : ℕ} (hk : ∀ p : Associates α, Irreducible p → k ∣ count p a.factors) : - ∃ b : Associates α, a = b ^ k := by - nontriviality α - obtain ⟨a0, hz, rfl⟩ := exists_non_zero_rep ha - rw [factors_mk a0 hz] at hk - have hk' : ∀ p, p ∈ factors' a0 → k ∣ (factors' a0).count p := by - rintro p - - have pp : p = ⟨p.val, p.2⟩ := by simp only [Subtype.coe_eta] - rw [pp, ← count_some p.2] - exact hk p.val p.2 - obtain ⟨u, hu⟩ := Multiset.exists_smul_of_dvd_count _ hk' - use FactorSet.prod (u : FactorSet α) - apply eq_of_factors_eq_factors - rw [pow_factors, prod_factors, factors_mk a0 hz, hu] - exact WithBot.coe_nsmul u k - -/-- The only divisors of prime powers are prime powers. See `eq_pow_find_of_dvd_irreducible_pow` -for an explicit expression as a p-power (without using `count`). -/ -theorem eq_pow_count_factors_of_dvd_pow {p a : Associates α} - (hp : Irreducible p) {n : ℕ} (h : a ∣ p ^ n) : a = p ^ p.count a.factors := by - nontriviality α - have hph := pow_ne_zero n hp.ne_zero - have ha := ne_zero_of_dvd_ne_zero hph h - apply eq_of_eq_counts ha (pow_ne_zero _ hp.ne_zero) - have eq_zero_of_ne : ∀ q : Associates α, Irreducible q → q ≠ p → _ = 0 := fun q hq h' => - Nat.eq_zero_of_le_zero <| by - convert count_le_count_of_le hph hq h - symm - rw [count_pow hp.ne_zero hq, count_eq_zero_of_ne hq hp h', mul_zero] - intro q hq - rw [count_pow hp.ne_zero hq] - by_cases h : q = p - · rw [h, count_self hp, mul_one] - · rw [count_eq_zero_of_ne hq hp h, mul_zero, eq_zero_of_ne q hq h] - -theorem count_factors_eq_find_of_dvd_pow {a p : Associates α} - (hp : Irreducible p) [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : - @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ = p.count a.factors := by - apply le_antisymm - · refine Nat.find_le ⟨1, ?_⟩ - rw [mul_one] - symm - exact eq_pow_count_factors_of_dvd_pow hp h - · have hph := pow_ne_zero (@Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩) hp.ne_zero - cases' subsingleton_or_nontrivial α with hα hα - · simp [eq_iff_true_of_subsingleton] at hph - convert count_le_count_of_le hph hp (@Nat.find_spec (fun n => a ∣ p ^ n) _ ⟨n, h⟩) - rw [count_pow hp.ne_zero hp, count_self hp, mul_one] - -end count - -theorem eq_pow_of_mul_eq_pow {a b c : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) - (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} (h : a * b = c ^ k) : - ∃ d : Associates α, a = d ^ k := by - classical - nontriviality α - by_cases hk0 : k = 0 - · use 1 - rw [hk0, pow_zero] at h ⊢ - apply (mul_eq_one.1 h).1 - · refine is_pow_of_dvd_count ha fun p hp ↦ ?_ - apply dvd_count_of_dvd_count_mul hb hp hab - rw [h] - apply dvd_count_pow _ hp - rintro rfl - rw [zero_pow hk0] at h - cases mul_eq_zero.mp h <;> contradiction - -/-- The only divisors of prime powers are prime powers. -/ -theorem eq_pow_find_of_dvd_irreducible_pow {a p : Associates α} (hp : Irreducible p) - [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : - a = p ^ @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ := by - classical rw [count_factors_eq_find_of_dvd_pow hp, ← eq_pow_count_factors_of_dvd_pow hp h] - exact h - -end Associates - -section - -open Associates UniqueFactorizationMonoid - -/-- `toGCDMonoid` constructs a GCD monoid out of a unique factorization domain. -/ -noncomputable def UniqueFactorizationMonoid.toGCDMonoid (α : Type*) [CancelCommMonoidWithZero α] - [UniqueFactorizationMonoid α] : GCDMonoid α where - gcd a b := Quot.out (Associates.mk a ⊓ Associates.mk b : Associates α) - lcm a b := Quot.out (Associates.mk a ⊔ Associates.mk b : Associates α) - gcd_dvd_left a b := by - rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] - exact inf_le_left - gcd_dvd_right a b := by - rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] - exact inf_le_right - dvd_gcd {a b c} hac hab := by - rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le, le_inf_iff, - mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] - exact ⟨hac, hab⟩ - lcm_zero_left a := by simp - lcm_zero_right a := by simp - gcd_mul_lcm a b := by - rw [← mk_eq_mk_iff_associated, ← Associates.mk_mul_mk, ← associated_iff_eq, Associates.quot_out, - Associates.quot_out, mul_comm, sup_mul_inf, Associates.mk_mul_mk] - -/-- `toNormalizedGCDMonoid` constructs a GCD monoid out of a normalization on a - unique factorization domain. -/ -noncomputable def UniqueFactorizationMonoid.toNormalizedGCDMonoid (α : Type*) - [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [NormalizationMonoid α] : - NormalizedGCDMonoid α := - { ‹NormalizationMonoid α› with - gcd := fun a b => (Associates.mk a ⊓ Associates.mk b).out - lcm := fun a b => (Associates.mk a ⊔ Associates.mk b).out - gcd_dvd_left := fun a b => (out_dvd_iff a (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_left - gcd_dvd_right := fun a b => - (out_dvd_iff b (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_right - dvd_gcd := fun {a} {b} {c} hac hab => - show a ∣ (Associates.mk c ⊓ Associates.mk b).out by - rw [dvd_out_iff, le_inf_iff, mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] - exact ⟨hac, hab⟩ - lcm_zero_left := fun a => show (⊤ ⊔ Associates.mk a).out = 0 by simp - lcm_zero_right := fun a => show (Associates.mk a ⊔ ⊤).out = 0 by simp - gcd_mul_lcm := fun a b => by - rw [← out_mul, mul_comm, sup_mul_inf, mk_mul_mk, out_mk] - exact normalize_associated (a * b) - normalize_gcd := fun a b => by apply normalize_out _ - normalize_lcm := fun a b => by apply normalize_out _ } - -instance (α) [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : - Nonempty (NormalizedGCDMonoid α) := by - letI := UniqueFactorizationMonoid.normalizationMonoid (α := α) - classical exact ⟨UniqueFactorizationMonoid.toNormalizedGCDMonoid α⟩ - -end - -namespace UniqueFactorizationMonoid - -/-- If `y` is a nonzero element of a unique factorization monoid with finitely -many units (e.g. `ℤ`, `Ideal (ring_of_integers K)`), it has finitely many divisors. -/ -noncomputable def fintypeSubtypeDvd {M : Type*} [CancelCommMonoidWithZero M] - [UniqueFactorizationMonoid M] [Fintype Mˣ] (y : M) (hy : y ≠ 0) : Fintype { x // x ∣ y } := by - haveI : Nontrivial M := ⟨⟨y, 0, hy⟩⟩ - haveI : NormalizationMonoid M := UniqueFactorizationMonoid.normalizationMonoid - haveI := Classical.decEq M - haveI := Classical.decEq (Associates M) - -- We'll show `fun (u : Mˣ) (f ⊆ factors y) ↦ u * Π f` is injective - -- and has image exactly the divisors of `y`. - refine - Fintype.ofFinset - (((normalizedFactors y).powerset.toFinset ×ˢ (Finset.univ : Finset Mˣ)).image fun s => - (s.snd : M) * s.fst.prod) - fun x => ?_ - simp only [exists_prop, Finset.mem_image, Finset.mem_product, Finset.mem_univ, and_true, - Multiset.mem_toFinset, Multiset.mem_powerset, exists_eq_right, Multiset.mem_map] - constructor - · rintro ⟨s, hs, rfl⟩ - show (s.snd : M) * s.fst.prod ∣ y - rw [(unit_associated_one.mul_right s.fst.prod).dvd_iff_dvd_left, one_mul, - ← (normalizedFactors_prod hy).dvd_iff_dvd_right] - exact Multiset.prod_dvd_prod_of_le hs - · rintro (h : x ∣ y) - have hx : x ≠ 0 := by - refine mt (fun hx => ?_) hy - rwa [hx, zero_dvd_iff] at h - obtain ⟨u, hu⟩ := normalizedFactors_prod hx - refine ⟨⟨normalizedFactors x, u⟩, ?_, (mul_comm _ _).trans hu⟩ - exact (dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mp h - -end UniqueFactorizationMonoid - -section Finsupp - -variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] -variable [NormalizationMonoid α] [DecidableEq α] - -open UniqueFactorizationMonoid - -/-- This returns the multiset of irreducible factors as a `Finsupp`. -/ -noncomputable def factorization (n : α) : α →₀ ℕ := - Multiset.toFinsupp (normalizedFactors n) - -theorem factorization_eq_count {n p : α} : - factorization n p = Multiset.count p (normalizedFactors n) := by simp [factorization] - -@[simp] -theorem factorization_zero : factorization (0 : α) = 0 := by simp [factorization] - -@[simp] -theorem factorization_one : factorization (1 : α) = 0 := by simp [factorization] - -/-- The support of `factorization n` is exactly the Finset of normalized factors -/ -@[simp] -theorem support_factorization {n : α} : - (factorization n).support = (normalizedFactors n).toFinset := by - simp [factorization, Multiset.toFinsupp_support] - -/-- For nonzero `a` and `b`, the power of `p` in `a * b` is the sum of the powers in `a` and `b` -/ -@[simp] -theorem factorization_mul {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) : - factorization (a * b) = factorization a + factorization b := by - simp [factorization, normalizedFactors_mul ha hb] - -/-- For any `p`, the power of `p` in `x^n` is `n` times the power in `x` -/ -theorem factorization_pow {x : α} {n : ℕ} : factorization (x ^ n) = n • factorization x := by - ext - simp [factorization] - -theorem associated_of_factorization_eq (a b : α) (ha : a ≠ 0) (hb : b ≠ 0) - (h : factorization a = factorization b) : Associated a b := by - simp_rw [factorization, AddEquiv.apply_eq_iff_eq] at h - rwa [associated_iff_normalizedFactors_eq_normalizedFactors ha hb] - -end Finsupp - -open UniqueFactorizationMonoid in -/-- Every non-zero prime ideal in a unique factorization domain contains a prime element. -/ -theorem Ideal.IsPrime.exists_mem_prime_of_ne_bot {R : Type*} [CommSemiring R] [IsDomain R] - [UniqueFactorizationMonoid R] {I : Ideal R} (hI₂ : I.IsPrime) (hI : I ≠ ⊥) : - ∃ x ∈ I, Prime x := by - classical - obtain ⟨a : R, ha₁ : a ∈ I, ha₂ : a ≠ 0⟩ := Submodule.exists_mem_ne_zero_of_ne_bot hI - replace ha₁ : (factors a).prod ∈ I := by - obtain ⟨u : Rˣ, hu : (factors a).prod * u = a⟩ := factors_prod ha₂ - rwa [← hu, mul_unit_mem_iff_mem _ u.isUnit] at ha₁ - obtain ⟨p : R, hp₁ : p ∈ factors a, hp₂ : p ∈ I⟩ := - (hI₂.multiset_prod_mem_iff_exists_mem <| factors a).1 ha₁ - exact ⟨p, hp₂, prime_of_factor p hp₁⟩ - -namespace Nat - -instance instWfDvdMonoid : WfDvdMonoid ℕ where - wf := by - refine RelHomClass.wellFounded - (⟨fun x : ℕ => if x = 0 then (⊤ : ℕ∞) else x, ?_⟩ : DvdNotUnit →r (· < ·)) wellFounded_lt - intro a b h - cases' a with a - · exfalso - revert h - simp [DvdNotUnit] - cases b - · simpa [succ_ne_zero] using WithTop.coe_lt_top (a + 1) - cases' dvd_and_not_dvd_iff.2 h with h1 h2 - simp only [succ_ne_zero, cast_lt, if_false] - refine lt_of_le_of_ne (Nat.le_of_dvd (Nat.succ_pos _) h1) fun con => h2 ?_ - rw [con] - -instance instUniqueFactorizationMonoid : UniqueFactorizationMonoid ℕ where - irreducible_iff_prime := Nat.irreducible_iff_prime - -open UniqueFactorizationMonoid - -lemma factors_eq : ∀ n : ℕ, normalizedFactors n = n.primeFactorsList - | 0 => by simp - | n + 1 => by - rw [← Multiset.rel_eq, ← associated_eq_eq] - apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor _ - · rw [Multiset.prod_coe, Nat.prod_primeFactorsList n.succ_ne_zero] - exact normalizedFactors_prod n.succ_ne_zero - · intro x hx - rw [Nat.irreducible_iff_prime, ← Nat.prime_iff] - exact Nat.prime_of_mem_primeFactorsList hx - -lemma factors_multiset_prod_of_irreducible {s : Multiset ℕ} (h : ∀ x : ℕ, x ∈ s → Irreducible x) : - normalizedFactors s.prod = s := by - rw [← Multiset.rel_eq, ← associated_eq_eq] - apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor h - (normalizedFactors_prod _) - rw [Ne, Multiset.prod_eq_zero_iff] - exact fun con ↦ not_irreducible_zero (h 0 con) - -end Nat - -section Ideal - -/-- The ascending chain condition on principal ideals holds in a `WfDvdMonoid` domain. -/ -lemma Ideal.setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [WfDvdMonoid α] [IsDomain α] : - {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·) := by - have : {I : Ideal α | I.IsPrincipal} = ((fun a ↦ Ideal.span {a}) '' Set.univ) := by - ext - simp [Submodule.isPrincipal_iff, eq_comm] - rw [this, Set.wellFoundedOn_image, Set.wellFoundedOn_univ] - convert wellFounded_dvdNotUnit (α := α) - ext - exact Ideal.span_singleton_lt_span_singleton - -/-- The ascending chain condition on principal ideals in a domain is sufficient to prove that -the domain is `WfDvdMonoid`. -/ -lemma WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [IsDomain α] - (h : {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·)) : - WfDvdMonoid α := by - have : WellFounded (α := {I : Ideal α // I.IsPrincipal}) (· > ·) := h - constructor - convert InvImage.wf (fun a => ⟨Ideal.span ({a} : Set α), _, rfl⟩) this - ext - exact Ideal.span_singleton_lt_span_singleton.symm - -end Ideal - -set_option linter.style.longFile 2100 diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean new file mode 100644 index 0000000000000..31b9b8426dbed --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Basic.lean @@ -0,0 +1,473 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.Algebra.BigOperators.Associated +import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Algebra.SMulWithZero +import Mathlib.Data.ENat.Basic +import Mathlib.Data.Multiset.OrderedMonoid +import Mathlib.RingTheory.UniqueFactorizationDomain.Defs + +/-! +# Basic results un unique factorization monoids + +## Main results +* `prime_factors_unique`: the prime factors of an element in a cancellative + commutative monoid with zero (e.g. an integral domain) are unique up to associates +* `UniqueFactorizationMonoid.factors_unique`: the irreducible factors of an element + in a unique factorization monoid (e.g. a UFD) are unique up to associates +* `UniqueFactorizationMonoid.iff_exists_prime_factors`: unique factorization exists iff each nonzero + elements factors into a product of primes +* `UniqueFactorizationMonoid.dvd_of_dvd_mul_left_of_no_prime_factors`: Euclid's lemma: + if `a ∣ b * c` and `a` and `c` have no common prime factors, `a ∣ b`. +* `UniqueFactorizationMonoid.dvd_of_dvd_mul_right_of_no_prime_factors`: Euclid's lemma: + if `a ∣ b * c` and `a` and `b` have no common prime factors, `a ∣ c`. +* `UniqueFactorizationMonoid.exists_reduced_factors`: in a UFM, we can divide out a common factor + to get relatively prime elements. +-/ + + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace WfDvdMonoid + +variable [CommMonoidWithZero α] + +open Associates Nat + +theorem of_wfDvdMonoid_associates (_ : WfDvdMonoid (Associates α)) : WfDvdMonoid α := + ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).2 wellFounded_dvdNotUnit⟩ + +variable [WfDvdMonoid α] + +instance wfDvdMonoid_associates : WfDvdMonoid (Associates α) := + ⟨(mk_surjective.wellFounded_iff mk_dvdNotUnit_mk_iff.symm).1 wellFounded_dvdNotUnit⟩ + +theorem wellFoundedLT_associates : WellFoundedLT (Associates α) := + ⟨Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit⟩ + +@[deprecated wellFoundedLT_associates (since := "2024-09-02")] +theorem wellFounded_associates : WellFounded ((· < ·) : Associates α → Associates α → Prop) := + Subrelation.wf dvdNotUnit_of_lt wellFounded_dvdNotUnit + +end WfDvdMonoid + +theorem WfDvdMonoid.of_wellFoundedLT_associates [CancelCommMonoidWithZero α] + (h : WellFoundedLT (Associates α)) : WfDvdMonoid α := + WfDvdMonoid.of_wfDvdMonoid_associates + ⟨by + convert h.wf + ext + exact Associates.dvdNotUnit_iff_lt⟩ + +@[deprecated WfDvdMonoid.of_wellFoundedLT_associates (since := "2024-09-02")] +theorem WfDvdMonoid.of_wellFounded_associates [CancelCommMonoidWithZero α] + (h : WellFounded ((· < ·) : Associates α → Associates α → Prop)) : WfDvdMonoid α := + WfDvdMonoid.of_wfDvdMonoid_associates + ⟨by + convert h + ext + exact Associates.dvdNotUnit_iff_lt⟩ + +theorem WfDvdMonoid.iff_wellFounded_associates [CancelCommMonoidWithZero α] : + WfDvdMonoid α ↔ WellFoundedLT (Associates α) := + ⟨by apply WfDvdMonoid.wellFoundedLT_associates, WfDvdMonoid.of_wellFoundedLT_associates⟩ + +instance Associates.ufm [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : + UniqueFactorizationMonoid (Associates α) := + { (WfDvdMonoid.wfDvdMonoid_associates : WfDvdMonoid (Associates α)) with + irreducible_iff_prime := by + rw [← Associates.irreducible_iff_prime_iff] + apply UniqueFactorizationMonoid.irreducible_iff_prime } + +theorem prime_factors_unique [CancelCommMonoidWithZero α] : + ∀ {f g : Multiset α}, + (∀ x ∈ f, Prime x) → (∀ x ∈ g, Prime x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g := by + classical + intro f + induction' f using Multiset.induction_on with p f ih + · intros g _ hg h + exact Multiset.rel_zero_left.2 <| + Multiset.eq_zero_of_forall_not_mem fun x hx => + have : IsUnit g.prod := by simpa [associated_one_iff_isUnit] using h.symm + (hg x hx).not_unit <| + isUnit_iff_dvd_one.2 <| (Multiset.dvd_prod hx).trans (isUnit_iff_dvd_one.1 this) + · intros g hf hg hfg + let ⟨b, hbg, hb⟩ := + (exists_associated_mem_of_dvd_prod (hf p (by simp)) fun q hq => hg _ hq) <| + hfg.dvd_iff_dvd_right.1 (show p ∣ (p ::ₘ f).prod by simp) + haveI := Classical.decEq α + rw [← Multiset.cons_erase hbg] + exact + Multiset.Rel.cons hb + (ih (fun q hq => hf _ (by simp [hq])) + (fun {q} (hq : q ∈ g.erase b) => hg q (Multiset.mem_of_mem_erase hq)) + (Associated.of_mul_left + (by rwa [← Multiset.prod_cons, ← Multiset.prod_cons, Multiset.cons_erase hbg]) hb + (hf p (by simp)).ne_zero)) + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +theorem factors_unique {f g : Multiset α} (hf : ∀ x ∈ f, Irreducible x) + (hg : ∀ x ∈ g, Irreducible x) (h : f.prod ~ᵤ g.prod) : Multiset.Rel Associated f g := + prime_factors_unique (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hf x hx)) + (fun x hx => UniqueFactorizationMonoid.irreducible_iff_prime.mp (hg x hx)) h + +end UniqueFactorizationMonoid + +/-- If an irreducible has a prime factorization, + then it is an associate of one of its prime factors. -/ +theorem prime_factors_irreducible [CancelCommMonoidWithZero α] {a : α} {f : Multiset α} + (ha : Irreducible a) (pfa : (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) : ∃ p, a ~ᵤ p ∧ f = {p} := by + haveI := Classical.decEq α + refine @Multiset.induction_on _ + (fun g => (g.prod ~ᵤ a) → (∀ b ∈ g, Prime b) → ∃ p, a ~ᵤ p ∧ g = {p}) f ?_ ?_ pfa.2 pfa.1 + · intro h; exact (ha.not_unit (associated_one_iff_isUnit.1 (Associated.symm h))).elim + · rintro p s _ ⟨u, hu⟩ hs + use p + have hs0 : s = 0 := by + by_contra hs0 + obtain ⟨q, hq⟩ := Multiset.exists_mem_of_ne_zero hs0 + apply (hs q (by simp [hq])).2.1 + refine (ha.isUnit_or_isUnit (?_ : _ = p * ↑u * (s.erase q).prod * _)).resolve_left ?_ + · rw [mul_right_comm _ _ q, mul_assoc, ← Multiset.prod_cons, Multiset.cons_erase hq, ← hu, + mul_comm, mul_comm p _, mul_assoc] + simp + apply mt isUnit_of_mul_isUnit_left (mt isUnit_of_mul_isUnit_left _) + apply (hs p (Multiset.mem_cons_self _ _)).2.1 + simp only [mul_one, Multiset.prod_cons, Multiset.prod_zero, hs0] at * + exact ⟨Associated.symm ⟨u, hu⟩, rfl⟩ + +theorem irreducible_iff_prime_of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] + (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) + (uif : + ∀ f g : Multiset α, + (∀ x ∈ f, Irreducible x) → + (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) + (p : α) : Irreducible p ↔ Prime p := + letI := Classical.decEq α + ⟨ fun hpi => + ⟨hpi.ne_zero, hpi.1, fun a b ⟨x, hx⟩ => + if hab0 : a * b = 0 then + (eq_zero_or_eq_zero_of_mul_eq_zero hab0).elim (fun ha0 => by simp [ha0]) fun hb0 => by + simp [hb0] + else by + have hx0 : x ≠ 0 := fun hx0 => by simp_all + have ha0 : a ≠ 0 := left_ne_zero_of_mul hab0 + have hb0 : b ≠ 0 := right_ne_zero_of_mul hab0 + cases' eif x hx0 with fx hfx + cases' eif a ha0 with fa hfa + cases' eif b hb0 with fb hfb + have h : Multiset.Rel Associated (p ::ₘ fx) (fa + fb) := by + apply uif + · exact fun i hi => (Multiset.mem_cons.1 hi).elim (fun hip => hip.symm ▸ hpi) (hfx.1 _) + · exact fun i hi => (Multiset.mem_add.1 hi).elim (hfa.1 _) (hfb.1 _) + calc + Multiset.prod (p ::ₘ fx) ~ᵤ a * b := by + rw [hx, Multiset.prod_cons]; exact hfx.2.mul_left _ + _ ~ᵤ fa.prod * fb.prod := hfa.2.symm.mul_mul hfb.2.symm + _ = _ := by rw [Multiset.prod_add] + + exact + let ⟨q, hqf, hq⟩ := Multiset.exists_mem_of_rel_of_mem h (Multiset.mem_cons_self p _) + (Multiset.mem_add.1 hqf).elim + (fun hqa => + Or.inl <| hq.dvd_iff_dvd_left.2 <| hfa.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqa)) + fun hqb => + Or.inr <| hq.dvd_iff_dvd_left.2 <| hfb.2.dvd_iff_dvd_right.1 (Multiset.dvd_prod hqb)⟩, + Prime.irreducible⟩ + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] +variable [UniqueFactorizationMonoid α] + +@[simp] +theorem factors_one : factors (1 : α) = 0 := by + nontriviality α using factors + rw [← Multiset.rel_zero_right] + refine factors_unique irreducible_of_factor (fun x hx => (Multiset.not_mem_zero x hx).elim) ?_ + rw [Multiset.prod_zero] + exact factors_prod one_ne_zero + +theorem exists_mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + p ∣ a → ∃ q ∈ factors a, p ~ᵤ q := fun ⟨b, hb⟩ => + have hb0 : b ≠ 0 := fun hb0 => by simp_all + have : Multiset.Rel Associated (p ::ₘ factors b) (factors a) := + factors_unique + (fun _ hx => (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_factor _)) + irreducible_of_factor + (Associated.symm <| + calc + Multiset.prod (factors a) ~ᵤ a := factors_prod ha0 + _ = p * b := hb + _ ~ᵤ Multiset.prod (p ::ₘ factors b) := by + rw [Multiset.prod_cons]; exact (factors_prod hb0).symm.mul_left _ + ) + Multiset.exists_mem_of_rel_of_mem this (by simp) + +theorem exists_mem_factors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : ∃ p, p ∈ factors x := by + obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx + obtain ⟨p, hp, _⟩ := exists_mem_factors_of_dvd hx hp' hp'x + exact ⟨p, hp⟩ + +open Classical in +theorem factors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + Multiset.Rel Associated (factors (x * y)) (factors x + factors y) := by + refine + factors_unique irreducible_of_factor + (fun a ha => + (Multiset.mem_add.mp ha).by_cases (irreducible_of_factor _) (irreducible_of_factor _)) + ((factors_prod (mul_ne_zero hx hy)).trans ?_) + rw [Multiset.prod_add] + exact (Associated.mul_mul (factors_prod hx) (factors_prod hy)).symm + +theorem factors_pow {x : α} (n : ℕ) : + Multiset.Rel Associated (factors (x ^ n)) (n • factors x) := by + match n with + | 0 => rw [zero_smul, pow_zero, factors_one, Multiset.rel_zero_right] + | n+1 => + by_cases h0 : x = 0 + · simp [h0, zero_pow n.succ_ne_zero, smul_zero] + · rw [pow_succ', succ_nsmul'] + refine Multiset.Rel.trans _ (factors_mul h0 (pow_ne_zero n h0)) ?_ + refine Multiset.Rel.add ?_ <| factors_pow n + exact Multiset.rel_refl_of_refl_on fun y _ => Associated.refl _ + +@[simp] +theorem factors_pos (x : α) (hx : x ≠ 0) : 0 < factors x ↔ ¬IsUnit x := by + constructor + · intro h hx + obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' + exact (prime_of_factor _ hp).not_unit (isUnit_of_dvd_unit (dvd_of_mem_factors hp) hx) + · intro h + obtain ⟨p, hp⟩ := exists_mem_factors hx h + exact + bot_lt_iff_ne_bot.mpr + (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) + +open Multiset in +theorem factors_pow_count_prod [DecidableEq α] {x : α} (hx : x ≠ 0) : + (∏ p ∈ (factors x).toFinset, p ^ (factors x).count p) ~ᵤ x := + calc + _ = prod (∑ a ∈ toFinset (factors x), count a (factors x) • {a}) := by + simp only [prod_sum, prod_nsmul, prod_singleton] + _ = prod (factors x) := by rw [toFinset_sum_count_nsmul_eq (factors x)] + _ ~ᵤ x := factors_prod hx + +theorem factors_rel_of_associated {a b : α} (h : Associated a b) : + Multiset.Rel Associated (factors a) (factors b) := by + rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) + · simp + · refine factors_unique irreducible_of_factor irreducible_of_factor ?_ + exact ((factors_prod ha).trans h).trans (factors_prod hb).symm + +end UniqueFactorizationMonoid + +namespace Associates + +attribute [local instance] Associated.setoid + +open Multiset UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +theorem unique' {p q : Multiset (Associates α)} : + (∀ a ∈ p, Irreducible a) → (∀ a ∈ q, Irreducible a) → p.prod = q.prod → p = q := by + apply Multiset.induction_on_multiset_quot p + apply Multiset.induction_on_multiset_quot q + intro s t hs ht eq + refine Multiset.map_mk_eq_map_mk_of_rel (UniqueFactorizationMonoid.factors_unique ?_ ?_ ?_) + · exact fun a ha => irreducible_mk.1 <| hs _ <| Multiset.mem_map_of_mem _ ha + · exact fun a ha => irreducible_mk.1 <| ht _ <| Multiset.mem_map_of_mem _ ha + have eq' : (Quot.mk Setoid.r : α → Associates α) = Associates.mk := funext quot_mk_eq_mk + rwa [eq', prod_mk, prod_mk, mk_eq_mk_iff_associated] at eq + +theorem prod_le_prod_iff_le [Nontrivial α] {p q : Multiset (Associates α)} + (hp : ∀ a ∈ p, Irreducible a) (hq : ∀ a ∈ q, Irreducible a) : p.prod ≤ q.prod ↔ p ≤ q := by + refine ⟨?_, prod_le_prod⟩ + rintro ⟨c, eqc⟩ + refine Multiset.le_iff_exists_add.2 ⟨factors c, unique' hq (fun x hx ↦ ?_) ?_⟩ + · obtain h | h := Multiset.mem_add.1 hx + · exact hp x h + · exact irreducible_of_factor _ h + · rw [eqc, Multiset.prod_add] + congr + refine associated_iff_eq.mp (factors_prod fun hc => ?_).symm + refine not_irreducible_zero (hq _ ?_) + rw [← prod_eq_zero_iff, eqc, hc, mul_zero] + +end Associates + +section ExistsPrimeFactors + +variable [CancelCommMonoidWithZero α] +variable (pf : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) +include pf + +theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := + ⟨by + classical + refine RelHomClass.wellFounded + (RelHom.mk ?_ ?_ : (DvdNotUnit : α → α → Prop) →r ((· < ·) : ℕ∞ → ℕ∞ → Prop)) wellFounded_lt + · intro a + by_cases h : a = 0 + · exact ⊤ + exact ↑(Multiset.card (Classical.choose (pf a h))) + rintro a b ⟨ane0, ⟨c, hc, b_eq⟩⟩ + rw [dif_neg ane0] + by_cases h : b = 0 + · simp [h, lt_top_iff_ne_top] + · rw [dif_neg h, Nat.cast_lt] + have cne0 : c ≠ 0 := by + refine mt (fun con => ?_) h + rw [b_eq, con, mul_zero] + calc + Multiset.card (Classical.choose (pf a ane0)) < + _ + Multiset.card (Classical.choose (pf c cne0)) := + lt_add_of_pos_right _ + (Multiset.card_pos.mpr fun con => hc (associated_one_iff_isUnit.mp ?_)) + _ = Multiset.card (Classical.choose (pf a ane0) + Classical.choose (pf c cne0)) := + (Multiset.card_add _ _).symm + _ = Multiset.card (Classical.choose (pf b h)) := + Multiset.card_eq_card_of_rel + (prime_factors_unique ?_ (Classical.choose_spec (pf _ h)).1 ?_) + + · convert (Classical.choose_spec (pf c cne0)).2.symm + rw [con, Multiset.prod_zero] + · intro x hadd + rw [Multiset.mem_add] at hadd + cases' hadd with h h <;> apply (Classical.choose_spec (pf _ _)).1 _ h <;> assumption + · rw [Multiset.prod_add] + trans a * c + · apply Associated.mul_mul <;> apply (Classical.choose_spec (pf _ _)).2 <;> assumption + · rw [← b_eq] + apply (Classical.choose_spec (pf _ _)).2.symm; assumption⟩ + +theorem irreducible_iff_prime_of_exists_prime_factors {p : α} : Irreducible p ↔ Prime p := by + by_cases hp0 : p = 0 + · simp [hp0] + refine ⟨fun h => ?_, Prime.irreducible⟩ + obtain ⟨f, hf⟩ := pf p hp0 + obtain ⟨q, hq, rfl⟩ := prime_factors_irreducible h hf + rw [hq.prime_iff] + exact hf.1 q (Multiset.mem_singleton_self _) + +theorem UniqueFactorizationMonoid.of_exists_prime_factors : UniqueFactorizationMonoid α := + { WfDvdMonoid.of_exists_prime_factors pf with + irreducible_iff_prime := irreducible_iff_prime_of_exists_prime_factors pf } + +end ExistsPrimeFactors + +theorem UniqueFactorizationMonoid.iff_exists_prime_factors [CancelCommMonoidWithZero α] : + UniqueFactorizationMonoid α ↔ + ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := + ⟨fun h => @UniqueFactorizationMonoid.exists_prime_factors _ _ h, + UniqueFactorizationMonoid.of_exists_prime_factors⟩ + +section + +variable {β : Type*} [CancelCommMonoidWithZero α] [CancelCommMonoidWithZero β] + +theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactorizationMonoid α) : + UniqueFactorizationMonoid β := by + rw [UniqueFactorizationMonoid.iff_exists_prime_factors] at hα ⊢ + intro a ha + obtain ⟨w, hp, u, h⟩ := + hα (e.symm a) fun h => + ha <| by + convert← map_zero e + simp [← h] + exact + ⟨w.map e, fun b hb => + let ⟨c, hc, he⟩ := Multiset.mem_map.1 hb + he ▸ e.prime_iff.1 (hp c hc), + Units.map e.toMonoidHom u, + by + rw [Multiset.prod_hom, toMonoidHom_eq_coe, Units.coe_map, MonoidHom.coe_coe, ← map_mul e, h, + apply_symm_apply]⟩ + +theorem MulEquiv.uniqueFactorizationMonoid_iff (e : α ≃* β) : + UniqueFactorizationMonoid α ↔ UniqueFactorizationMonoid β := + ⟨e.uniqueFactorizationMonoid, e.symm.uniqueFactorizationMonoid⟩ + +end + +namespace UniqueFactorizationMonoid + +theorem of_exists_unique_irreducible_factors [CancelCommMonoidWithZero α] + (eif : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod ~ᵤ a) + (uif : + ∀ f g : Multiset α, + (∀ x ∈ f, Irreducible x) → + (∀ x ∈ g, Irreducible x) → f.prod ~ᵤ g.prod → Multiset.Rel Associated f g) : + UniqueFactorizationMonoid α := + UniqueFactorizationMonoid.of_exists_prime_factors + (by + convert eif using 7 + simp_rw [irreducible_iff_prime_of_exists_unique_irreducible_factors eif uif]) + +variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] + +theorem isRelPrime_iff_no_prime_factors {a b : R} (ha : a ≠ 0) : + IsRelPrime a b ↔ ∀ ⦃d⦄, d ∣ a → d ∣ b → ¬Prime d := + ⟨fun h _ ha hb ↦ (·.not_unit <| h ha hb), fun h ↦ WfDvdMonoid.isRelPrime_of_no_irreducible_factors + (ha ·.1) fun _ irr ha hb ↦ h ha hb (UniqueFactorizationMonoid.irreducible_iff_prime.mp irr)⟩ + +/-- Euclid's lemma: if `a ∣ b * c` and `a` and `c` have no common prime factors, `a ∣ b`. +Compare `IsCoprime.dvd_of_dvd_mul_left`. -/ +theorem dvd_of_dvd_mul_left_of_no_prime_factors {a b c : R} (ha : a ≠ 0) + (h : ∀ ⦃d⦄, d ∣ a → d ∣ c → ¬Prime d) : a ∣ b * c → a ∣ b := + ((isRelPrime_iff_no_prime_factors ha).mpr h).dvd_of_dvd_mul_right + +/-- Euclid's lemma: if `a ∣ b * c` and `a` and `b` have no common prime factors, `a ∣ c`. +Compare `IsCoprime.dvd_of_dvd_mul_right`. -/ +theorem dvd_of_dvd_mul_right_of_no_prime_factors {a b c : R} (ha : a ≠ 0) + (no_factors : ∀ {d}, d ∣ a → d ∣ b → ¬Prime d) : a ∣ b * c → a ∣ c := by + simpa [mul_comm b c] using dvd_of_dvd_mul_left_of_no_prime_factors ha @no_factors + +/-- If `a ≠ 0, b` are elements of a unique factorization domain, then dividing +out their common factor `c'` gives `a'` and `b'` with no factors in common. -/ +theorem exists_reduced_factors : + ∀ a ≠ (0 : R), ∀ b, + ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := by + intro a + refine induction_on_prime a ?_ ?_ ?_ + · intros + contradiction + · intro a a_unit _ b + use a, b, 1 + constructor + · intro p p_dvd_a _ + exact isUnit_of_dvd_unit p_dvd_a a_unit + · simp + · intro a p a_ne_zero p_prime ih_a pa_ne_zero b + by_cases h : p ∣ b + · rcases h with ⟨b, rfl⟩ + obtain ⟨a', b', c', no_factor, ha', hb'⟩ := ih_a a_ne_zero b + refine ⟨a', b', p * c', @no_factor, ?_, ?_⟩ + · rw [mul_assoc, ha'] + · rw [mul_assoc, hb'] + · obtain ⟨a', b', c', coprime, rfl, rfl⟩ := ih_a a_ne_zero b + refine ⟨p * a', b', c', ?_, mul_left_comm _ _ _, rfl⟩ + intro q q_dvd_pa' q_dvd_b' + cases' p_prime.left_dvd_or_dvd_right_of_dvd_mul q_dvd_pa' with p_dvd_q q_dvd_a' + · have : p ∣ c' * b' := dvd_mul_of_dvd_right (p_dvd_q.trans q_dvd_b') _ + contradiction + exact coprime q_dvd_a' q_dvd_b' + +theorem exists_reduced_factors' (a b : R) (hb : b ≠ 0) : + ∃ a' b' c', IsRelPrime a' b' ∧ c' * a' = a ∧ c' * b' = b := + let ⟨b', a', c', no_factor, hb, ha⟩ := exists_reduced_factors b hb a + ⟨a', b', c', fun _ hpb hpa => no_factor hpa hpb, ha, hb⟩ + +@[deprecated (since := "2024-09-21")] alias pow_right_injective := pow_injective_of_not_isUnit +@[deprecated (since := "2024-09-21")] alias pow_eq_pow_iff := pow_inj_of_not_isUnit + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean new file mode 100644 index 0000000000000..121cf52e3a622 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Defs.lean @@ -0,0 +1,205 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.Algebra.Associated.Basic +import Mathlib.Algebra.BigOperators.Group.Multiset +import Mathlib.Algebra.Group.Submonoid.Membership +import Mathlib.Order.WellFounded + +/-! +# Unique factorization + +## Main Definitions +* `WfDvdMonoid` holds for `Monoid`s for which a strict divisibility relation is + well-founded. +* `UniqueFactorizationMonoid` holds for `WfDvdMonoid`s where + `Irreducible` is equivalent to `Prime` +-/ + +assert_not_exists Field +assert_not_exists Finsupp +assert_not_exists Ideal + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +/-- Well-foundedness of the strict version of ∣, which is equivalent to the descending chain +condition on divisibility and to the ascending chain condition on +principal ideals in an integral domain. + -/ +abbrev WfDvdMonoid (α : Type*) [CommMonoidWithZero α] : Prop := + IsWellFounded α DvdNotUnit + +theorem wellFounded_dvdNotUnit {α : Type*} [CommMonoidWithZero α] [h : WfDvdMonoid α] : + WellFounded (DvdNotUnit (α := α)) := + h.wf + +namespace WfDvdMonoid + +variable [CommMonoidWithZero α] + +open Associates Nat + +variable [WfDvdMonoid α] + +theorem exists_irreducible_factor {a : α} (ha : ¬IsUnit a) (ha0 : a ≠ 0) : + ∃ i, Irreducible i ∧ i ∣ a := + let ⟨b, hs, hr⟩ := wellFounded_dvdNotUnit.has_min { b | b ∣ a ∧ ¬IsUnit b } ⟨a, dvd_rfl, ha⟩ + ⟨b, + ⟨hs.2, fun c d he => + let h := dvd_trans ⟨d, he⟩ hs.1 + or_iff_not_imp_left.2 fun hc => + of_not_not fun hd => hr c ⟨h, hc⟩ ⟨ne_zero_of_dvd_ne_zero ha0 h, d, hd, he⟩⟩, + hs.1⟩ + +@[elab_as_elim] +theorem induction_on_irreducible {P : α → Prop} (a : α) (h0 : P 0) (hu : ∀ u : α, IsUnit u → P u) + (hi : ∀ a i : α, a ≠ 0 → Irreducible i → P a → P (i * a)) : P a := + haveI := Classical.dec + wellFounded_dvdNotUnit.fix + (fun a ih => + if ha0 : a = 0 then ha0.substr h0 + else + if hau : IsUnit a then hu a hau + else + let ⟨i, hii, b, hb⟩ := exists_irreducible_factor hau ha0 + let hb0 : b ≠ 0 := ne_zero_of_dvd_ne_zero ha0 ⟨i, mul_comm i b ▸ hb⟩ + hb.symm ▸ hi b i hb0 hii <| ih b ⟨hb0, i, hii.1, mul_comm i b ▸ hb⟩) + a + +theorem exists_factors (a : α) : + a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ Associated f.prod a := + induction_on_irreducible a (fun h => (h rfl).elim) + (fun _ hu _ => ⟨0, fun _ h => False.elim (Multiset.not_mem_zero _ h), hu.unit, one_mul _⟩) + fun a i ha0 hi ih _ => + let ⟨s, hs⟩ := ih ha0 + ⟨i ::ₘ s, fun b H => (Multiset.mem_cons.1 H).elim (fun h => h.symm ▸ hi) (hs.1 b), by + rw [s.prod_cons i] + exact hs.2.mul_left i⟩ + +theorem not_unit_iff_exists_factors_eq (a : α) (hn0 : a ≠ 0) : + ¬IsUnit a ↔ ∃ f : Multiset α, (∀ b ∈ f, Irreducible b) ∧ f.prod = a ∧ f ≠ ∅ := + ⟨fun hnu => by + obtain ⟨f, hi, u, rfl⟩ := exists_factors a hn0 + obtain ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero fun h : f = 0 => hnu <| by simp [h] + classical + refine ⟨(f.erase b).cons (b * u), fun a ha => ?_, ?_, Multiset.cons_ne_zero⟩ + · obtain rfl | ha := Multiset.mem_cons.1 ha + exacts [Associated.irreducible ⟨u, rfl⟩ (hi b h), hi a (Multiset.mem_of_mem_erase ha)] + · rw [Multiset.prod_cons, mul_comm b, mul_assoc, Multiset.prod_erase h, mul_comm], + fun ⟨_, hi, he, hne⟩ => + let ⟨b, h⟩ := Multiset.exists_mem_of_ne_zero hne + not_isUnit_of_not_isUnit_dvd (hi b h).not_unit <| he ▸ Multiset.dvd_prod h⟩ + +theorem isRelPrime_of_no_irreducible_factors {x y : α} (nonzero : ¬(x = 0 ∧ y = 0)) + (H : ∀ z : α, Irreducible z → z ∣ x → ¬z ∣ y) : IsRelPrime x y := + isRelPrime_of_no_nonunits_factors nonzero fun _z znu znz zx zy ↦ + have ⟨i, h1, h2⟩ := exists_irreducible_factor znu znz + H i h1 (h2.trans zx) (h2.trans zy) + +end WfDvdMonoid + +section Prio + +-- set_option default_priority 100 + +-- see Note [default priority] +/-- unique factorization monoids. + +These are defined as `CancelCommMonoidWithZero`s with well-founded strict divisibility +relations, but this is equivalent to more familiar definitions: + +Each element (except zero) is uniquely represented as a multiset of irreducible factors. +Uniqueness is only up to associated elements. + +Each element (except zero) is non-uniquely represented as a multiset +of prime factors. + +To define a UFD using the definition in terms of multisets +of irreducible factors, use the definition `of_exists_unique_irreducible_factors` + +To define a UFD using the definition in terms of multisets +of prime factors, use the definition `of_exists_prime_factors` + +-/ +class UniqueFactorizationMonoid (α : Type*) [CancelCommMonoidWithZero α] extends + IsWellFounded α DvdNotUnit : Prop where + protected irreducible_iff_prime : ∀ {a : α}, Irreducible a ↔ Prime a + +instance (priority := 100) ufm_of_decomposition_of_wfDvdMonoid + [CancelCommMonoidWithZero α] [WfDvdMonoid α] [DecompositionMonoid α] : + UniqueFactorizationMonoid α := + { ‹WfDvdMonoid α› with irreducible_iff_prime := irreducible_iff_prime } + +@[deprecated ufm_of_decomposition_of_wfDvdMonoid (since := "2024-02-12")] +theorem ufm_of_gcd_of_wfDvdMonoid [CancelCommMonoidWithZero α] [WfDvdMonoid α] + [DecompositionMonoid α] : UniqueFactorizationMonoid α := + ufm_of_decomposition_of_wfDvdMonoid + +end Prio + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +theorem exists_prime_factors (a : α) : + a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a := by + simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] + apply WfDvdMonoid.exists_factors a + +lemma exists_prime_iff : + (∃ (p : α), Prime p) ↔ ∃ (x : α), x ≠ 0 ∧ ¬ IsUnit x := by + refine ⟨fun ⟨p, hp⟩ ↦ ⟨p, hp.ne_zero, hp.not_unit⟩, fun ⟨x, hx₀, hxu⟩ ↦ ?_⟩ + obtain ⟨f, hf, -⟩ := WfDvdMonoid.exists_irreducible_factor hxu hx₀ + exact ⟨f, UniqueFactorizationMonoid.irreducible_iff_prime.mp hf⟩ + +@[elab_as_elim] +theorem induction_on_prime {P : α → Prop} (a : α) (h₁ : P 0) (h₂ : ∀ x : α, IsUnit x → P x) + (h₃ : ∀ a p : α, a ≠ 0 → Prime p → P a → P (p * a)) : P a := by + simp_rw [← UniqueFactorizationMonoid.irreducible_iff_prime] at h₃ + exact WfDvdMonoid.induction_on_irreducible a h₁ h₂ h₃ + +instance : DecompositionMonoid α where + primal a := by + obtain rfl | ha := eq_or_ne a 0; · exact isPrimal_zero + obtain ⟨f, hf, u, rfl⟩ := exists_prime_factors a ha + exact ((Submonoid.isPrimal α).multiset_prod_mem f (hf · ·|>.isPrimal)).mul u.isUnit.isPrimal + +end UniqueFactorizationMonoid + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] +variable [UniqueFactorizationMonoid α] + +open Classical in +/-- Noncomputably determines the multiset of prime factors. -/ +noncomputable def factors (a : α) : Multiset α := + if h : a = 0 then 0 else Classical.choose (UniqueFactorizationMonoid.exists_prime_factors a h) + +theorem factors_prod {a : α} (ane0 : a ≠ 0) : Associated (factors a).prod a := by + rw [factors, dif_neg ane0] + exact (Classical.choose_spec (exists_prime_factors a ane0)).2 + +@[simp] +theorem factors_zero : factors (0 : α) = 0 := by simp [factors] + +theorem ne_zero_of_mem_factors {p a : α} (h : p ∈ factors a) : a ≠ 0 := by + rintro rfl + simp at h + +theorem dvd_of_mem_factors {p a : α} (h : p ∈ factors a) : p ∣ a := + dvd_trans (Multiset.dvd_prod h) (Associated.dvd (factors_prod (ne_zero_of_mem_factors h))) + +theorem prime_of_factor {a : α} (x : α) (hx : x ∈ factors a) : Prime x := by + have ane0 := ne_zero_of_mem_factors hx + rw [factors, dif_neg ane0] at hx + exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 x hx + +theorem irreducible_of_factor {a : α} : ∀ x : α, x ∈ factors a → Irreducible x := fun x h => + (prime_of_factor x h).irreducible + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean new file mode 100644 index 0000000000000..2dd1d63b54515 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/FactorSet.lean @@ -0,0 +1,647 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.Data.Finsupp.Multiset +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.Tactic.Ring + +/-! +# Set of factors + +## Main definitions +* `Associates.FactorSet`: multiset of factors of an element, unique up to propositional equality. +* `Associates.factors`: determine the `FactorSet` for a given element. + +## TODO +* set up the complete lattice structure on `FactorSet`. + +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace Associates + +open UniqueFactorizationMonoid Associated Multiset + +variable [CancelCommMonoidWithZero α] + +/-- `FactorSet α` representation elements of unique factorization domain as multisets. +`Multiset α` produced by `normalizedFactors` are only unique up to associated elements, while the +multisets in `FactorSet α` are unique by equality and restricted to irreducible elements. This +gives us a representation of each element as a unique multisets (or the added ⊤ for 0), which has a +complete lattice structure. Infimum is the greatest common divisor and supremum is the least common +multiple. +-/ +abbrev FactorSet.{u} (α : Type u) [CancelCommMonoidWithZero α] : Type u := + WithTop (Multiset { a : Associates α // Irreducible a }) + +attribute [local instance] Associated.setoid + +theorem FactorSet.coe_add {a b : Multiset { a : Associates α // Irreducible a }} : + (↑(a + b) : FactorSet α) = a + b := by norm_cast + +theorem FactorSet.sup_add_inf_eq_add [DecidableEq (Associates α)] : + ∀ a b : FactorSet α, a ⊔ b + a ⊓ b = a + b + | ⊤, b => show ⊤ ⊔ b + ⊤ ⊓ b = ⊤ + b by simp + | a, ⊤ => show a ⊔ ⊤ + a ⊓ ⊤ = a + ⊤ by simp + | WithTop.some a, WithTop.some b => + show (a : FactorSet α) ⊔ b + (a : FactorSet α) ⊓ b = a + b by + rw [← WithTop.coe_sup, ← WithTop.coe_inf, ← WithTop.coe_add, ← WithTop.coe_add, + WithTop.coe_eq_coe] + exact Multiset.union_add_inter _ _ + +/-- Evaluates the product of a `FactorSet` to be the product of the corresponding multiset, + or `0` if there is none. -/ +def FactorSet.prod : FactorSet α → Associates α + | ⊤ => 0 + | WithTop.some s => (s.map (↑)).prod + +@[simp] +theorem prod_top : (⊤ : FactorSet α).prod = 0 := + rfl + +@[simp] +theorem prod_coe {s : Multiset { a : Associates α // Irreducible a }} : + FactorSet.prod (s : FactorSet α) = (s.map (↑)).prod := + rfl + +@[simp] +theorem prod_add : ∀ a b : FactorSet α, (a + b).prod = a.prod * b.prod + | ⊤, b => show (⊤ + b).prod = (⊤ : FactorSet α).prod * b.prod by simp + | a, ⊤ => show (a + ⊤).prod = a.prod * (⊤ : FactorSet α).prod by simp + | WithTop.some a, WithTop.some b => by + rw [← FactorSet.coe_add, prod_coe, prod_coe, prod_coe, Multiset.map_add, Multiset.prod_add] + +@[gcongr] +theorem prod_mono : ∀ {a b : FactorSet α}, a ≤ b → a.prod ≤ b.prod + | ⊤, b, h => by + have : b = ⊤ := top_unique h + rw [this, prod_top] + | a, ⊤, _ => show a.prod ≤ (⊤ : FactorSet α).prod by simp + | WithTop.some _, WithTop.some _, h => + prod_le_prod <| Multiset.map_le_map <| WithTop.coe_le_coe.1 <| h + +theorem FactorSet.prod_eq_zero_iff [Nontrivial α] (p : FactorSet α) : p.prod = 0 ↔ p = ⊤ := by + unfold FactorSet at p + induction p -- TODO: `induction_eliminator` doesn't work with `abbrev` + · simp only [eq_self_iff_true, Associates.prod_top] + · rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, + iff_false, not_exists] + exact fun a => not_and_of_not_right _ a.prop.ne_zero + +section count + +variable [DecidableEq (Associates α)] + +/-- `bcount p s` is the multiplicity of `p` in the FactorSet `s` (with bundled `p`)-/ +def bcount (p : { a : Associates α // Irreducible a }) : + FactorSet α → ℕ + | ⊤ => 0 + | WithTop.some s => s.count p + +variable [∀ p : Associates α, Decidable (Irreducible p)] {p : Associates α} + +/-- `count p s` is the multiplicity of the irreducible `p` in the FactorSet `s`. + +If `p` is not irreducible, `count p s` is defined to be `0`. -/ +def count (p : Associates α) : FactorSet α → ℕ := + if hp : Irreducible p then bcount ⟨p, hp⟩ else 0 + +@[simp] +theorem count_some (hp : Irreducible p) (s : Multiset _) : + count p (WithTop.some s) = s.count ⟨p, hp⟩ := by + simp only [count, dif_pos hp, bcount] + +@[simp] +theorem count_zero (hp : Irreducible p) : count p (0 : FactorSet α) = 0 := by + simp only [count, dif_pos hp, bcount, Multiset.count_zero] + +theorem count_reducible (hp : ¬Irreducible p) : count p = 0 := dif_neg hp + +end count + +section Mem + +/-- membership in a FactorSet (bundled version) -/ +def BfactorSetMem : { a : Associates α // Irreducible a } → FactorSet α → Prop + | _, ⊤ => True + | p, some l => p ∈ l + +/-- `FactorSetMem p s` is the predicate that the irreducible `p` is a member of +`s : FactorSet α`. + +If `p` is not irreducible, `p` is not a member of any `FactorSet`. -/ +def FactorSetMem (s : FactorSet α) (p : Associates α) : Prop := + letI : Decidable (Irreducible p) := Classical.dec _ + if hp : Irreducible p then BfactorSetMem ⟨p, hp⟩ s else False + +instance : Membership (Associates α) (FactorSet α) := + ⟨FactorSetMem⟩ + +@[simp] +theorem factorSetMem_eq_mem (p : Associates α) (s : FactorSet α) : FactorSetMem s p = (p ∈ s) := + rfl + +theorem mem_factorSet_top {p : Associates α} {hp : Irreducible p} : p ∈ (⊤ : FactorSet α) := by + dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; exact trivial + +theorem mem_factorSet_some {p : Associates α} {hp : Irreducible p} + {l : Multiset { a : Associates α // Irreducible a }} : + p ∈ (l : FactorSet α) ↔ Subtype.mk p hp ∈ l := by + dsimp only [Membership.mem]; dsimp only [FactorSetMem]; split_ifs; rfl + +theorem reducible_not_mem_factorSet {p : Associates α} (hp : ¬Irreducible p) (s : FactorSet α) : + ¬p ∈ s := fun h ↦ by + rwa [← factorSetMem_eq_mem, FactorSetMem, dif_neg hp] at h + +theorem irreducible_of_mem_factorSet {p : Associates α} {s : FactorSet α} (h : p ∈ s) : + Irreducible p := + by_contra fun hp ↦ reducible_not_mem_factorSet hp s h + +end Mem + +variable [UniqueFactorizationMonoid α] + +theorem FactorSet.unique [Nontrivial α] {p q : FactorSet α} (h : p.prod = q.prod) : p = q := by + -- TODO: `induction_eliminator` doesn't work with `abbrev` + unfold FactorSet at p q + induction p <;> induction q + · rfl + · rw [eq_comm, ← FactorSet.prod_eq_zero_iff, ← h, Associates.prod_top] + · rw [← FactorSet.prod_eq_zero_iff, h, Associates.prod_top] + · congr 1 + rw [← Multiset.map_eq_map Subtype.coe_injective] + apply unique' _ _ h <;> + · intro a ha + obtain ⟨⟨a', irred⟩, -, rfl⟩ := Multiset.mem_map.mp ha + rwa [Subtype.coe_mk] + +/-- This returns the multiset of irreducible factors as a `FactorSet`, + a multiset of irreducible associates `WithTop`. -/ +noncomputable def factors' (a : α) : Multiset { a : Associates α // Irreducible a } := + (factors a).pmap (fun a ha => ⟨Associates.mk a, irreducible_mk.2 ha⟩) irreducible_of_factor + +@[simp] +theorem map_subtype_coe_factors' {a : α} : + (factors' a).map (↑) = (factors a).map Associates.mk := by + simp [factors', Multiset.map_pmap, Multiset.pmap_eq_map] + +theorem factors'_cong {a b : α} (h : a ~ᵤ b) : factors' a = factors' b := by + obtain rfl | hb := eq_or_ne b 0 + · rw [associated_zero_iff_eq_zero] at h + rw [h] + have ha : a ≠ 0 := by + contrapose! hb with ha + rw [← associated_zero_iff_eq_zero, ← ha] + exact h.symm + rw [← Multiset.map_eq_map Subtype.coe_injective, map_subtype_coe_factors', + map_subtype_coe_factors', ← rel_associated_iff_map_eq_map] + exact + factors_unique irreducible_of_factor irreducible_of_factor + ((factors_prod ha).trans <| h.trans <| (factors_prod hb).symm) + +/-- This returns the multiset of irreducible factors of an associate as a `FactorSet`, + a multiset of irreducible associates `WithTop`. -/ +noncomputable def factors (a : Associates α) : FactorSet α := by + classical refine if h : a = 0 then ⊤ else Quotient.hrecOn a (fun x _ => factors' x) ?_ h + intro a b hab + apply Function.hfunext + · have : a ~ᵤ 0 ↔ b ~ᵤ 0 := Iff.intro (fun ha0 => hab.symm.trans ha0) fun hb0 => hab.trans hb0 + simp only [associated_zero_iff_eq_zero] at this + simp only [quotient_mk_eq_mk, this, mk_eq_zero] + exact fun ha hb _ => heq_of_eq <| congr_arg some <| factors'_cong hab + +@[simp] +theorem factors_zero : (0 : Associates α).factors = ⊤ := + dif_pos rfl + +@[deprecated (since := "2024-03-16")] alias factors_0 := factors_zero + +@[simp] +theorem factors_mk (a : α) (h : a ≠ 0) : (Associates.mk a).factors = factors' a := by + classical + apply dif_neg + apply mt mk_eq_zero.1 h + +@[simp] +theorem factors_prod (a : Associates α) : a.factors.prod = a := by + rcases Associates.mk_surjective a with ⟨a, rfl⟩ + rcases eq_or_ne a 0 with rfl | ha + · simp + · simp [ha, prod_mk, mk_eq_mk_iff_associated, UniqueFactorizationMonoid.factors_prod, + -Quotient.eq] + +@[simp] +theorem prod_factors [Nontrivial α] (s : FactorSet α) : s.prod.factors = s := + FactorSet.unique <| factors_prod _ + +@[nontriviality] +theorem factors_subsingleton [Subsingleton α] {a : Associates α} : a.factors = ⊤ := by + have : Subsingleton (Associates α) := inferInstance + convert factors_zero + +theorem factors_eq_top_iff_zero {a : Associates α} : a.factors = ⊤ ↔ a = 0 := by + nontriviality α + exact ⟨fun h ↦ by rwa [← factors_prod a, FactorSet.prod_eq_zero_iff], fun h ↦ h ▸ factors_zero⟩ + +@[deprecated (since := "2024-04-16")] alias factors_eq_none_iff_zero := factors_eq_top_iff_zero + +theorem factors_eq_some_iff_ne_zero {a : Associates α} : + (∃ s : Multiset { p : Associates α // Irreducible p }, a.factors = s) ↔ a ≠ 0 := by + simp_rw [@eq_comm _ a.factors, ← WithTop.ne_top_iff_exists] + exact factors_eq_top_iff_zero.not + +theorem eq_of_factors_eq_factors {a b : Associates α} (h : a.factors = b.factors) : a = b := by + have : a.factors.prod = b.factors.prod := by rw [h] + rwa [factors_prod, factors_prod] at this + +theorem eq_of_prod_eq_prod [Nontrivial α] {a b : FactorSet α} (h : a.prod = b.prod) : a = b := by + have : a.prod.factors = b.prod.factors := by rw [h] + rwa [prod_factors, prod_factors] at this + +@[simp] +theorem factors_mul (a b : Associates α) : (a * b).factors = a.factors + b.factors := by + nontriviality α + refine eq_of_prod_eq_prod <| eq_of_factors_eq_factors ?_ + rw [prod_add, factors_prod, factors_prod, factors_prod] + +@[gcongr] +theorem factors_mono : ∀ {a b : Associates α}, a ≤ b → a.factors ≤ b.factors + | s, t, ⟨d, eq⟩ => by rw [eq, factors_mul]; exact le_add_of_nonneg_right bot_le + +@[simp] +theorem factors_le {a b : Associates α} : a.factors ≤ b.factors ↔ a ≤ b := by + refine ⟨fun h ↦ ?_, factors_mono⟩ + have : a.factors.prod ≤ b.factors.prod := prod_mono h + rwa [factors_prod, factors_prod] at this + +section count + +variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] + +theorem eq_factors_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) + (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : + a.factors = b.factors := by + obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha + obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb + rw [h_sa, h_sb] at h ⊢ + rw [WithTop.coe_eq_coe] + have h_count : ∀ (p : Associates α) (hp : Irreducible p), + sa.count ⟨p, hp⟩ = sb.count ⟨p, hp⟩ := by + intro p hp + rw [← count_some, ← count_some, h p hp] + apply Multiset.toFinsupp.injective + ext ⟨p, hp⟩ + rw [Multiset.toFinsupp_apply, Multiset.toFinsupp_apply, h_count p hp] + +theorem eq_of_eq_counts {a b : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) + (h : ∀ p : Associates α, Irreducible p → p.count a.factors = p.count b.factors) : a = b := + eq_of_factors_eq_factors (eq_factors_of_eq_counts ha hb h) + +theorem count_le_count_of_factors_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) + (h : a.factors ≤ b.factors) : p.count a.factors ≤ p.count b.factors := by + by_cases ha : a = 0 + · simp_all + obtain ⟨sa, h_sa⟩ := factors_eq_some_iff_ne_zero.mpr ha + obtain ⟨sb, h_sb⟩ := factors_eq_some_iff_ne_zero.mpr hb + rw [h_sa, h_sb] at h ⊢ + rw [count_some hp, count_some hp]; rw [WithTop.coe_le_coe] at h + exact Multiset.count_le_of_le _ h + +theorem count_le_count_of_le {a b p : Associates α} (hb : b ≠ 0) (hp : Irreducible p) (h : a ≤ b) : + p.count a.factors ≤ p.count b.factors := + count_le_count_of_factors_le hb hp <| factors_mono h + +end count + +theorem prod_le [Nontrivial α] {a b : FactorSet α} : a.prod ≤ b.prod ↔ a ≤ b := by + refine ⟨fun h ↦ ?_, prod_mono⟩ + have : a.prod.factors ≤ b.prod.factors := factors_mono h + rwa [prod_factors, prod_factors] at this + +open Classical in +noncomputable instance : Max (Associates α) := + ⟨fun a b => (a.factors ⊔ b.factors).prod⟩ + +open Classical in +noncomputable instance : Min (Associates α) := + ⟨fun a b => (a.factors ⊓ b.factors).prod⟩ + +open Classical in +noncomputable instance : Lattice (Associates α) := + { Associates.instPartialOrder with + sup := (· ⊔ ·) + inf := (· ⊓ ·) + sup_le := fun _ _ c hac hbc => + factors_prod c ▸ prod_mono (sup_le (factors_mono hac) (factors_mono hbc)) + le_sup_left := fun a _ => le_trans (le_of_eq (factors_prod a).symm) <| prod_mono <| le_sup_left + le_sup_right := fun _ b => + le_trans (le_of_eq (factors_prod b).symm) <| prod_mono <| le_sup_right + le_inf := fun a _ _ hac hbc => + factors_prod a ▸ prod_mono (le_inf (factors_mono hac) (factors_mono hbc)) + inf_le_left := fun a _ => le_trans (prod_mono inf_le_left) (le_of_eq (factors_prod a)) + inf_le_right := fun _ b => le_trans (prod_mono inf_le_right) (le_of_eq (factors_prod b)) } + +open Classical in +theorem sup_mul_inf (a b : Associates α) : (a ⊔ b) * (a ⊓ b) = a * b := + show (a.factors ⊔ b.factors).prod * (a.factors ⊓ b.factors).prod = a * b by + nontriviality α + refine eq_of_factors_eq_factors ?_ + rw [← prod_add, prod_factors, factors_mul, FactorSet.sup_add_inf_eq_add] + +theorem dvd_of_mem_factors {a p : Associates α} (hm : p ∈ factors a) : + p ∣ a := by + rcases eq_or_ne a 0 with rfl | ha0 + · exact dvd_zero p + obtain ⟨a0, nza, ha'⟩ := exists_non_zero_rep ha0 + rw [← Associates.factors_prod a] + rw [← ha', factors_mk a0 nza] at hm ⊢ + rw [prod_coe] + apply Multiset.dvd_prod; apply Multiset.mem_map.mpr + exact ⟨⟨p, irreducible_of_mem_factorSet hm⟩, mem_factorSet_some.mp hm, rfl⟩ + +theorem dvd_of_mem_factors' {a : α} {p : Associates α} {hp : Irreducible p} {hz : a ≠ 0} + (h_mem : Subtype.mk p hp ∈ factors' a) : p ∣ Associates.mk a := by + haveI := Classical.decEq (Associates α) + apply dvd_of_mem_factors + rw [factors_mk _ hz] + apply mem_factorSet_some.2 h_mem + +theorem mem_factors'_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : + Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a := by + obtain ⟨q, hq, hpq⟩ := exists_mem_factors_of_dvd ha0 hp hd + apply Multiset.mem_pmap.mpr; use q; use hq + exact Subtype.eq (Eq.symm (mk_eq_mk_iff_associated.mpr hpq)) + +theorem mem_factors'_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + Subtype.mk (Associates.mk p) (irreducible_mk.2 hp) ∈ factors' a ↔ p ∣ a := by + constructor + · rw [← mk_dvd_mk] + apply dvd_of_mem_factors' + apply ha0 + · apply mem_factors'_of_dvd ha0 hp + +theorem mem_factors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) (hd : p ∣ a) : + Associates.mk p ∈ factors (Associates.mk a) := by + rw [factors_mk _ ha0] + exact mem_factorSet_some.mpr (mem_factors'_of_dvd ha0 hp hd) + +theorem mem_factors_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + Associates.mk p ∈ factors (Associates.mk a) ↔ p ∣ a := by + constructor + · rw [← mk_dvd_mk] + apply dvd_of_mem_factors + · apply mem_factors_of_dvd ha0 hp + +open Classical in +theorem exists_prime_dvd_of_not_inf_one {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) + (h : Associates.mk a ⊓ Associates.mk b ≠ 1) : ∃ p : α, Prime p ∧ p ∣ a ∧ p ∣ b := by + have hz : factors (Associates.mk a) ⊓ factors (Associates.mk b) ≠ 0 := by + contrapose! h with hf + change (factors (Associates.mk a) ⊓ factors (Associates.mk b)).prod = 1 + rw [hf] + exact Multiset.prod_zero + rw [factors_mk a ha, factors_mk b hb, ← WithTop.coe_inf] at hz + obtain ⟨⟨p0, p0_irr⟩, p0_mem⟩ := Multiset.exists_mem_of_ne_zero ((mt WithTop.coe_eq_coe.mpr) hz) + rw [Multiset.inf_eq_inter] at p0_mem + obtain ⟨p, rfl⟩ : ∃ p, Associates.mk p = p0 := Quot.exists_rep p0 + refine ⟨p, ?_, ?_, ?_⟩ + · rw [← UniqueFactorizationMonoid.irreducible_iff_prime, ← irreducible_mk] + exact p0_irr + · apply dvd_of_mk_le_mk + apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).left + apply ha + · apply dvd_of_mk_le_mk + apply dvd_of_mem_factors' (Multiset.mem_inter.mp p0_mem).right + apply hb + +theorem coprime_iff_inf_one {a b : α} (ha0 : a ≠ 0) (hb0 : b ≠ 0) : + Associates.mk a ⊓ Associates.mk b = 1 ↔ ∀ {d : α}, d ∣ a → d ∣ b → ¬Prime d := by + constructor + · intro hg p ha hb hp + refine (Associates.prime_mk.mpr hp).not_unit (isUnit_of_dvd_one ?_) + rw [← hg] + exact le_inf (mk_le_mk_of_dvd ha) (mk_le_mk_of_dvd hb) + · contrapose + intro hg hc + obtain ⟨p, hp, hpa, hpb⟩ := exists_prime_dvd_of_not_inf_one ha0 hb0 hg + exact hc hpa hpb hp + +theorem factors_self [Nontrivial α] {p : Associates α} (hp : Irreducible p) : + p.factors = WithTop.some {⟨p, hp⟩} := + eq_of_prod_eq_prod + (by rw [factors_prod, FactorSet.prod.eq_def]; dsimp; rw [prod_singleton]) + +theorem factors_prime_pow [Nontrivial α] {p : Associates α} (hp : Irreducible p) (k : ℕ) : + factors (p ^ k) = WithTop.some (Multiset.replicate k ⟨p, hp⟩) := + eq_of_prod_eq_prod + (by + rw [Associates.factors_prod, FactorSet.prod.eq_def] + dsimp; rw [Multiset.map_replicate, Multiset.prod_replicate, Subtype.coe_mk]) + +theorem prime_pow_le_iff_le_bcount [DecidableEq (Associates α)] {m p : Associates α} + (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : p ^ k ≤ m ↔ k ≤ bcount ⟨p, h₂⟩ m.factors := by + rcases Associates.exists_non_zero_rep h₁ with ⟨m, hm, rfl⟩ + have := nontrivial_of_ne _ _ hm + rw [bcount.eq_def, factors_mk, Multiset.le_count_iff_replicate_le, ← factors_le, + factors_prime_pow, factors_mk, WithTop.coe_le_coe] <;> assumption + +@[simp] +theorem factors_one [Nontrivial α] : factors (1 : Associates α) = 0 := by + apply eq_of_prod_eq_prod + rw [Associates.factors_prod] + exact Multiset.prod_zero + +@[simp] +theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : + (a ^ k).factors = k • a.factors := by + induction' k with n h + · rw [zero_nsmul, pow_zero] + exact factors_one + · rw [pow_succ, succ_nsmul, factors_mul, h] + +section count + +variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] + +theorem prime_pow_dvd_iff_le {m p : Associates α} (h₁ : m ≠ 0) (h₂ : Irreducible p) {k : ℕ} : + p ^ k ≤ m ↔ k ≤ count p m.factors := by + rw [count, dif_pos h₂, prime_pow_le_iff_le_bcount h₁] + +theorem le_of_count_ne_zero {m p : Associates α} (h0 : m ≠ 0) (hp : Irreducible p) : + count p m.factors ≠ 0 → p ≤ m := by + nontriviality α + rw [← pos_iff_ne_zero] + intro h + rw [← pow_one p] + apply (prime_pow_dvd_iff_le h0 hp).2 + simpa only + +theorem count_ne_zero_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + (Associates.mk p).count (Associates.mk a).factors ≠ 0 ↔ p ∣ a := by + nontriviality α + rw [← Associates.mk_le_mk_iff_dvd] + refine + ⟨fun h => + Associates.le_of_count_ne_zero (Associates.mk_ne_zero.mpr ha0) + (Associates.irreducible_mk.mpr hp) h, + fun h => ?_⟩ + rw [← pow_one (Associates.mk p), + Associates.prime_pow_dvd_iff_le (Associates.mk_ne_zero.mpr ha0) + (Associates.irreducible_mk.mpr hp)] at h + exact (zero_lt_one.trans_le h).ne' + +theorem count_self [Nontrivial α] {p : Associates α} + (hp : Irreducible p) : p.count p.factors = 1 := by + simp [factors_self hp, Associates.count_some hp] + +theorem count_eq_zero_of_ne {p q : Associates α} (hp : Irreducible p) + (hq : Irreducible q) (h : p ≠ q) : p.count q.factors = 0 := + not_ne_iff.mp fun h' ↦ h <| associated_iff_eq.mp <| hp.associated_of_dvd hq <| + le_of_count_ne_zero hq.ne_zero hp h' + +theorem count_mul {a : Associates α} (ha : a ≠ 0) {b : Associates α} + (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) : + count p (factors (a * b)) = count p a.factors + count p b.factors := by + obtain ⟨a0, nza, rfl⟩ := exists_non_zero_rep ha + obtain ⟨b0, nzb, rfl⟩ := exists_non_zero_rep hb + rw [factors_mul, factors_mk a0 nza, factors_mk b0 nzb, ← FactorSet.coe_add, count_some hp, + Multiset.count_add, count_some hp, count_some hp] + +theorem count_of_coprime {a : Associates α} (ha : a ≠ 0) + {b : Associates α} (hb : b ≠ 0) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {p : Associates α} + (hp : Irreducible p) : count p a.factors = 0 ∨ count p b.factors = 0 := by + rw [or_iff_not_imp_left, ← Ne] + intro hca + contrapose! hab with hcb + exact ⟨p, le_of_count_ne_zero ha hp hca, le_of_count_ne_zero hb hp hcb, + UniqueFactorizationMonoid.irreducible_iff_prime.mp hp⟩ + +theorem count_mul_of_coprime {a : Associates α} {b : Associates α} + (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : + count p a.factors = 0 ∨ count p a.factors = count p (a * b).factors := by + by_cases ha : a = 0 + · simp [ha] + cases' count_of_coprime ha hb hab hp with hz hb0; · tauto + apply Or.intro_right + rw [count_mul ha hb hp, hb0, add_zero] + +theorem count_mul_of_coprime' {a b : Associates α} {p : Associates α} + (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : + count p (a * b).factors = count p a.factors ∨ count p (a * b).factors = count p b.factors := by + by_cases ha : a = 0 + · simp [ha] + by_cases hb : b = 0 + · simp [hb] + rw [count_mul ha hb hp] + cases' count_of_coprime ha hb hab hp with ha0 hb0 + · apply Or.intro_right + rw [ha0, zero_add] + · apply Or.intro_left + rw [hb0, add_zero] + +theorem dvd_count_of_dvd_count_mul {a b : Associates α} (hb : b ≠ 0) + {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} + (habk : k ∣ count p (a * b).factors) : k ∣ count p a.factors := by + by_cases ha : a = 0 + · simpa [*] using habk + cases' count_of_coprime ha hb hab hp with hz h + · rw [hz] + exact dvd_zero k + · rw [count_mul ha hb hp, h] at habk + exact habk + +theorem count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) + {p : Associates α} (hp : Irreducible p) (k : ℕ) : + count p (a ^ k).factors = k * count p a.factors := by + induction' k with n h + · rw [pow_zero, factors_one, zero_mul, count_zero hp] + · rw [pow_succ', count_mul ha (pow_ne_zero _ ha) hp, h] + ring + +theorem dvd_count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) + {p : Associates α} (hp : Irreducible p) (k : ℕ) : k ∣ count p (a ^ k).factors := by + rw [count_pow ha hp] + apply dvd_mul_right + +theorem is_pow_of_dvd_count {a : Associates α} + (ha : a ≠ 0) {k : ℕ} (hk : ∀ p : Associates α, Irreducible p → k ∣ count p a.factors) : + ∃ b : Associates α, a = b ^ k := by + nontriviality α + obtain ⟨a0, hz, rfl⟩ := exists_non_zero_rep ha + rw [factors_mk a0 hz] at hk + have hk' : ∀ p, p ∈ factors' a0 → k ∣ (factors' a0).count p := by + rintro p - + have pp : p = ⟨p.val, p.2⟩ := by simp only [Subtype.coe_eta] + rw [pp, ← count_some p.2] + exact hk p.val p.2 + obtain ⟨u, hu⟩ := Multiset.exists_smul_of_dvd_count _ hk' + use FactorSet.prod (u : FactorSet α) + apply eq_of_factors_eq_factors + rw [pow_factors, prod_factors, factors_mk a0 hz, hu] + exact WithBot.coe_nsmul u k + +/-- The only divisors of prime powers are prime powers. See `eq_pow_find_of_dvd_irreducible_pow` +for an explicit expression as a p-power (without using `count`). -/ +theorem eq_pow_count_factors_of_dvd_pow {p a : Associates α} + (hp : Irreducible p) {n : ℕ} (h : a ∣ p ^ n) : a = p ^ p.count a.factors := by + nontriviality α + have hph := pow_ne_zero n hp.ne_zero + have ha := ne_zero_of_dvd_ne_zero hph h + apply eq_of_eq_counts ha (pow_ne_zero _ hp.ne_zero) + have eq_zero_of_ne : ∀ q : Associates α, Irreducible q → q ≠ p → _ = 0 := fun q hq h' => + Nat.eq_zero_of_le_zero <| by + convert count_le_count_of_le hph hq h + symm + rw [count_pow hp.ne_zero hq, count_eq_zero_of_ne hq hp h', mul_zero] + intro q hq + rw [count_pow hp.ne_zero hq] + by_cases h : q = p + · rw [h, count_self hp, mul_one] + · rw [count_eq_zero_of_ne hq hp h, mul_zero, eq_zero_of_ne q hq h] + +theorem count_factors_eq_find_of_dvd_pow {a p : Associates α} + (hp : Irreducible p) [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : + @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ = p.count a.factors := by + apply le_antisymm + · refine Nat.find_le ⟨1, ?_⟩ + rw [mul_one] + symm + exact eq_pow_count_factors_of_dvd_pow hp h + · have hph := pow_ne_zero (@Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩) hp.ne_zero + cases' subsingleton_or_nontrivial α with hα hα + · simp [eq_iff_true_of_subsingleton] at hph + convert count_le_count_of_le hph hp (@Nat.find_spec (fun n => a ∣ p ^ n) _ ⟨n, h⟩) + rw [count_pow hp.ne_zero hp, count_self hp, mul_one] + +end count + +theorem eq_pow_of_mul_eq_pow {a b c : Associates α} (ha : a ≠ 0) (hb : b ≠ 0) + (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} (h : a * b = c ^ k) : + ∃ d : Associates α, a = d ^ k := by + classical + nontriviality α + by_cases hk0 : k = 0 + · use 1 + rw [hk0, pow_zero] at h ⊢ + apply (mul_eq_one.1 h).1 + · refine is_pow_of_dvd_count ha fun p hp ↦ ?_ + apply dvd_count_of_dvd_count_mul hb hp hab + rw [h] + apply dvd_count_pow _ hp + rintro rfl + rw [zero_pow hk0] at h + cases mul_eq_zero.mp h <;> contradiction + +/-- The only divisors of prime powers are prime powers. -/ +theorem eq_pow_find_of_dvd_irreducible_pow {a p : Associates α} (hp : Irreducible p) + [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : + a = p ^ @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ := by + classical rw [count_factors_eq_find_of_dvd_pow hp, ← eq_pow_count_factors_of_dvd_pow hp h] + exact h + +end Associates diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean new file mode 100644 index 0000000000000..bd5cd5c1e9819 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Finite.lean @@ -0,0 +1,53 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Finiteness of divisors + +## Main results +* `UniqueFactorizationMonoid.fintypeSubtypeDvd`: elements of a UFM with finitely many units have + finitely many divisors. +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace UniqueFactorizationMonoid + +/-- If `y` is a nonzero element of a unique factorization monoid with finitely +many units (e.g. `ℤ`, `Ideal (ring_of_integers K)`), it has finitely many divisors. -/ +noncomputable def fintypeSubtypeDvd {M : Type*} [CancelCommMonoidWithZero M] + [UniqueFactorizationMonoid M] [Fintype Mˣ] (y : M) (hy : y ≠ 0) : Fintype { x // x ∣ y } := by + haveI : Nontrivial M := ⟨⟨y, 0, hy⟩⟩ + haveI : NormalizationMonoid M := UniqueFactorizationMonoid.normalizationMonoid + haveI := Classical.decEq M + haveI := Classical.decEq (Associates M) + -- We'll show `fun (u : Mˣ) (f ⊆ factors y) ↦ u * Π f` is injective + -- and has image exactly the divisors of `y`. + refine + Fintype.ofFinset + (((normalizedFactors y).powerset.toFinset ×ˢ (Finset.univ : Finset Mˣ)).image fun s => + (s.snd : M) * s.fst.prod) + fun x => ?_ + simp only [exists_prop, Finset.mem_image, Finset.mem_product, Finset.mem_univ, and_true, + Multiset.mem_toFinset, Multiset.mem_powerset, exists_eq_right, Multiset.mem_map] + constructor + · rintro ⟨s, hs, rfl⟩ + show (s.snd : M) * s.fst.prod ∣ y + rw [(unit_associated_one.mul_right s.fst.prod).dvd_iff_dvd_left, one_mul, + ← (normalizedFactors_prod hy).dvd_iff_dvd_right] + exact Multiset.prod_dvd_prod_of_le hs + · rintro (h : x ∣ y) + have hx : x ≠ 0 := by + refine mt (fun hx => ?_) hy + rwa [hx, zero_dvd_iff] at h + obtain ⟨u, hu⟩ := normalizedFactors_prod hx + refine ⟨⟨normalizedFactors x, u⟩, ?_, (mul_comm _ _).trans hu⟩ + exact (dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mp h + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean new file mode 100644 index 0000000000000..a901b3c83b1a3 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Finsupp.lean @@ -0,0 +1,62 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.Data.Finsupp.Multiset +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Factors as finsupp + +## Main definitions +* `UniqueFactorizationMonoid.factorization`: the multiset of irreducible factors as a `Finsupp`. +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +section Finsupp + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] +variable [NormalizationMonoid α] [DecidableEq α] + +open UniqueFactorizationMonoid + +/-- This returns the multiset of irreducible factors as a `Finsupp`. -/ +noncomputable def factorization (n : α) : α →₀ ℕ := + Multiset.toFinsupp (normalizedFactors n) + +theorem factorization_eq_count {n p : α} : + factorization n p = Multiset.count p (normalizedFactors n) := by simp [factorization] + +@[simp] +theorem factorization_zero : factorization (0 : α) = 0 := by simp [factorization] + +@[simp] +theorem factorization_one : factorization (1 : α) = 0 := by simp [factorization] + +/-- The support of `factorization n` is exactly the Finset of normalized factors -/ +@[simp] +theorem support_factorization {n : α} : + (factorization n).support = (normalizedFactors n).toFinset := by + simp [factorization, Multiset.toFinsupp_support] + +/-- For nonzero `a` and `b`, the power of `p` in `a * b` is the sum of the powers in `a` and `b` -/ +@[simp] +theorem factorization_mul {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) : + factorization (a * b) = factorization a + factorization b := by + simp [factorization, normalizedFactors_mul ha hb] + +/-- For any `p`, the power of `p` in `x^n` is `n` times the power in `x` -/ +theorem factorization_pow {x : α} {n : ℕ} : factorization (x ^ n) = n • factorization x := by + ext + simp [factorization] + +theorem associated_of_factorization_eq (a b : α) (ha : a ≠ 0) (hb : b ≠ 0) + (h : factorization a = factorization b) : Associated a b := by + simp_rw [factorization, AddEquiv.apply_eq_iff_eq] at h + rwa [associated_iff_normalizedFactors_eq_normalizedFactors ha hb] + +end Finsupp diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean new file mode 100644 index 0000000000000..74940f8f2ed5f --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/GCDMonoid.lean @@ -0,0 +1,73 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain.FactorSet +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Building GCD out of unique factorization + +## Main results +* `UniqueFactorizationMonoid.toGCDMonoid`: choose a GCD monoid structure given unique factorization. +-/ + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +section + +open Associates UniqueFactorizationMonoid + +/-- `toGCDMonoid` constructs a GCD monoid out of a unique factorization domain. -/ +noncomputable def UniqueFactorizationMonoid.toGCDMonoid (α : Type*) [CancelCommMonoidWithZero α] + [UniqueFactorizationMonoid α] : GCDMonoid α where + gcd a b := Quot.out (Associates.mk a ⊓ Associates.mk b : Associates α) + lcm a b := Quot.out (Associates.mk a ⊔ Associates.mk b : Associates α) + gcd_dvd_left a b := by + rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] + exact inf_le_left + gcd_dvd_right a b := by + rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le] + exact inf_le_right + dvd_gcd {a b c} hac hab := by + rw [← mk_dvd_mk, Associates.quot_out, congr_fun₂ dvd_eq_le, le_inf_iff, + mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] + exact ⟨hac, hab⟩ + lcm_zero_left a := by simp + lcm_zero_right a := by simp + gcd_mul_lcm a b := by + rw [← mk_eq_mk_iff_associated, ← Associates.mk_mul_mk, ← associated_iff_eq, Associates.quot_out, + Associates.quot_out, mul_comm, sup_mul_inf, Associates.mk_mul_mk] + +/-- `toNormalizedGCDMonoid` constructs a GCD monoid out of a normalization on a + unique factorization domain. -/ +noncomputable def UniqueFactorizationMonoid.toNormalizedGCDMonoid (α : Type*) + [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] [NormalizationMonoid α] : + NormalizedGCDMonoid α := + { ‹NormalizationMonoid α› with + gcd := fun a b => (Associates.mk a ⊓ Associates.mk b).out + lcm := fun a b => (Associates.mk a ⊔ Associates.mk b).out + gcd_dvd_left := fun a b => (out_dvd_iff a (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_left + gcd_dvd_right := fun a b => + (out_dvd_iff b (Associates.mk a ⊓ Associates.mk b)).2 <| inf_le_right + dvd_gcd := fun {a} {b} {c} hac hab => + show a ∣ (Associates.mk c ⊓ Associates.mk b).out by + rw [dvd_out_iff, le_inf_iff, mk_le_mk_iff_dvd, mk_le_mk_iff_dvd] + exact ⟨hac, hab⟩ + lcm_zero_left := fun a => show (⊤ ⊔ Associates.mk a).out = 0 by simp + lcm_zero_right := fun a => show (Associates.mk a ⊔ ⊤).out = 0 by simp + gcd_mul_lcm := fun a b => by + rw [← out_mul, mul_comm, sup_mul_inf, mk_mul_mk, out_mk] + exact normalize_associated (a * b) + normalize_gcd := fun a b => by apply normalize_out _ + normalize_lcm := fun a b => by apply normalize_out _ } + +instance (α) [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] : + Nonempty (NormalizedGCDMonoid α) := by + letI := UniqueFactorizationMonoid.normalizationMonoid (α := α) + classical exact ⟨UniqueFactorizationMonoid.toNormalizedGCDMonoid α⟩ + +end diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean new file mode 100644 index 0000000000000..76a9fb5c09941 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Ideal.lean @@ -0,0 +1,58 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.RingTheory.Ideal.Operations +import Mathlib.RingTheory.UniqueFactorizationDomain.Defs + +/-! +# Unique factorization and ascending chain condition on ideals + +## Main results +* `Ideal.setOf_isPrincipal_wellFoundedOn_gt`, `WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt` + in a domain, well-foundedness of the strict version of ∣ is equivalent to the ascending + chain condition on principal ideals. +-/ + +variable {α : Type*} + +open UniqueFactorizationMonoid in +/-- Every non-zero prime ideal in a unique factorization domain contains a prime element. -/ +theorem Ideal.IsPrime.exists_mem_prime_of_ne_bot {R : Type*} [CommSemiring R] [IsDomain R] + [UniqueFactorizationMonoid R] {I : Ideal R} (hI₂ : I.IsPrime) (hI : I ≠ ⊥) : + ∃ x ∈ I, Prime x := by + classical + obtain ⟨a : R, ha₁ : a ∈ I, ha₂ : a ≠ 0⟩ := Submodule.exists_mem_ne_zero_of_ne_bot hI + replace ha₁ : (factors a).prod ∈ I := by + obtain ⟨u : Rˣ, hu : (factors a).prod * u = a⟩ := factors_prod ha₂ + rwa [← hu, mul_unit_mem_iff_mem _ u.isUnit] at ha₁ + obtain ⟨p : R, hp₁ : p ∈ factors a, hp₂ : p ∈ I⟩ := + (hI₂.multiset_prod_mem_iff_exists_mem <| factors a).1 ha₁ + exact ⟨p, hp₂, prime_of_factor p hp₁⟩ + +section Ideal + +/-- The ascending chain condition on principal ideals holds in a `WfDvdMonoid` domain. -/ +lemma Ideal.setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [WfDvdMonoid α] [IsDomain α] : + {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·) := by + have : {I : Ideal α | I.IsPrincipal} = ((fun a ↦ Ideal.span {a}) '' Set.univ) := by + ext + simp [Submodule.isPrincipal_iff, eq_comm] + rw [this, Set.wellFoundedOn_image, Set.wellFoundedOn_univ] + convert wellFounded_dvdNotUnit (α := α) + ext + exact Ideal.span_singleton_lt_span_singleton + +/-- The ascending chain condition on principal ideals in a domain is sufficient to prove that +the domain is `WfDvdMonoid`. -/ +lemma WfDvdMonoid.of_setOf_isPrincipal_wellFoundedOn_gt [CommSemiring α] [IsDomain α] + (h : {I : Ideal α | I.IsPrincipal}.WellFoundedOn (· > ·)) : + WfDvdMonoid α := by + have : WellFounded (α := {I : Ideal α // I.IsPrincipal}) (· > ·) := h + constructor + convert InvImage.wf (fun a => ⟨Ideal.span ({a} : Set α), _, rfl⟩) this + ext + exact Ideal.span_singleton_lt_span_singleton.symm + +end Ideal diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean new file mode 100644 index 0000000000000..a8a04d3a9ad0f --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicative.lean @@ -0,0 +1,163 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Multiplicative maps on unique factorization domains + +## Main results +* `UniqueFactorizationMonoid.induction_on_coprime`: if `P` holds for `0`, units and powers of + primes, and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, then `P` holds on all `a : α`. +* `UniqueFactorizationMonoid.multiplicative_of_coprime`: if `f` maps `p ^ i` to `(f p) ^ i` for + primes `p`, and `f` is multiplicative on coprime elements, then `f` is multiplicative everywhere. +-/ + + +variable {α : Type*} + +namespace UniqueFactorizationMonoid + +variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] + +section Multiplicative + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] +variable {β : Type*} [CancelCommMonoidWithZero β] + +theorem prime_pow_coprime_prod_of_coprime_insert [DecidableEq α] {s : Finset α} (i : α → ℕ) (p : α) + (hps : p ∉ s) (is_prime : ∀ q ∈ insert p s, Prime q) + (is_coprime : ∀ᵉ (q ∈ insert p s) (q' ∈ insert p s), q ∣ q' → q = q') : + IsRelPrime (p ^ i p) (∏ p' ∈ s, p' ^ i p') := by + have hp := is_prime _ (Finset.mem_insert_self _ _) + refine (isRelPrime_iff_no_prime_factors <| pow_ne_zero _ hp.ne_zero).mpr ?_ + intro d hdp hdprod hd + apply hps + replace hdp := hd.dvd_of_dvd_pow hdp + obtain ⟨q, q_mem', hdq⟩ := hd.exists_mem_multiset_dvd hdprod + obtain ⟨q, q_mem, rfl⟩ := Multiset.mem_map.mp q_mem' + replace hdq := hd.dvd_of_dvd_pow hdq + have : p ∣ q := dvd_trans (hd.irreducible.dvd_symm hp.irreducible hdp) hdq + convert q_mem using 0 + rw [Finset.mem_val, + is_coprime _ (Finset.mem_insert_self p s) _ (Finset.mem_insert_of_mem q_mem) this] + +/-- If `P` holds for units and powers of primes, +and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, +then `P` holds on a product of powers of distinct primes. -/ +@[elab_as_elim] +theorem induction_on_prime_power {P : α → Prop} (s : Finset α) (i : α → ℕ) + (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) + (h1 : ∀ {x}, IsUnit x → P x) (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) + (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : + P (∏ p ∈ s, p ^ i p) := by + letI := Classical.decEq α + induction' s using Finset.induction_on with p f' hpf' ih + · simpa using h1 isUnit_one + rw [Finset.prod_insert hpf'] + exact + hcp (prime_pow_coprime_prod_of_coprime_insert i p hpf' is_prime is_coprime) + (hpr (i p) (is_prime _ (Finset.mem_insert_self _ _))) + (ih (fun q hq => is_prime _ (Finset.mem_insert_of_mem hq)) fun q hq q' hq' => + is_coprime _ (Finset.mem_insert_of_mem hq) _ (Finset.mem_insert_of_mem hq')) + +/-- If `P` holds for `0`, units and powers of primes, +and `P x ∧ P y` for coprime `x, y` implies `P (x * y)`, +then `P` holds on all `a : α`. -/ +@[elab_as_elim] +theorem induction_on_coprime {P : α → Prop} (a : α) (h0 : P 0) (h1 : ∀ {x}, IsUnit x → P x) + (hpr : ∀ {p} (i : ℕ), Prime p → P (p ^ i)) + (hcp : ∀ {x y}, IsRelPrime x y → P x → P y → P (x * y)) : P a := by + letI := Classical.decEq α + have P_of_associated : ∀ {x y}, Associated x y → P x → P y := by + rintro x y ⟨u, rfl⟩ hx + exact hcp (fun p _ hpx => isUnit_of_dvd_unit hpx u.isUnit) hx (h1 u.isUnit) + by_cases ha0 : a = 0 + · rwa [ha0] + haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ + letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid + refine P_of_associated (normalizedFactors_prod ha0) ?_ + rw [← (normalizedFactors a).map_id, Finset.prod_multiset_map_count] + refine induction_on_prime_power _ _ ?_ ?_ @h1 @hpr @hcp <;> simp only [Multiset.mem_toFinset] + · apply prime_of_normalized_factor + · apply normalizedFactors_eq_of_dvd + +/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` +is multiplicative on coprime elements, then `f` is multiplicative on all products of primes. -/ +theorem multiplicative_prime_power {f : α → β} (s : Finset α) (i j : α → ℕ) + (is_prime : ∀ p ∈ s, Prime p) (is_coprime : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q) + (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) + (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) + (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : + f (∏ p ∈ s, p ^ (i p + j p)) = f (∏ p ∈ s, p ^ i p) * f (∏ p ∈ s, p ^ j p) := by + letI := Classical.decEq α + induction' s using Finset.induction_on with p s hps ih + · simpa using h1 isUnit_one + have hpr_p := is_prime _ (Finset.mem_insert_self _ _) + have hpr_s : ∀ p ∈ s, Prime p := fun p hp => is_prime _ (Finset.mem_insert_of_mem hp) + have hcp_p := fun i => prime_pow_coprime_prod_of_coprime_insert i p hps is_prime is_coprime + have hcp_s : ∀ᵉ (p ∈ s) (q ∈ s), p ∣ q → p = q := fun p hp q hq => + is_coprime p (Finset.mem_insert_of_mem hp) q (Finset.mem_insert_of_mem hq) + rw [Finset.prod_insert hps, Finset.prod_insert hps, Finset.prod_insert hps, hcp (hcp_p _), + hpr _ hpr_p, hcp (hcp_p _), hpr _ hpr_p, hcp (hcp_p (fun p => i p + j p)), hpr _ hpr_p, + ih hpr_s hcp_s, pow_add, mul_assoc, mul_left_comm (f p ^ j p), mul_assoc] + +/-- If `f` maps `p ^ i` to `(f p) ^ i` for primes `p`, and `f` +is multiplicative on coprime elements, then `f` is multiplicative everywhere. -/ +theorem multiplicative_of_coprime (f : α → β) (a b : α) (h0 : f 0 = 0) + (h1 : ∀ {x y}, IsUnit y → f (x * y) = f x * f y) + (hpr : ∀ {p} (i : ℕ), Prime p → f (p ^ i) = f p ^ i) + (hcp : ∀ {x y}, IsRelPrime x y → f (x * y) = f x * f y) : + f (a * b) = f a * f b := by + letI := Classical.decEq α + by_cases ha0 : a = 0 + · rw [ha0, zero_mul, h0, zero_mul] + by_cases hb0 : b = 0 + · rw [hb0, mul_zero, h0, mul_zero] + by_cases hf1 : f 1 = 0 + · calc + f (a * b) = f (a * b * 1) := by rw [mul_one] + _ = 0 := by simp only [h1 isUnit_one, hf1, mul_zero] + _ = f a * f (b * 1) := by simp only [h1 isUnit_one, hf1, mul_zero] + _ = f a * f b := by rw [mul_one] + haveI : Nontrivial α := ⟨⟨_, _, ha0⟩⟩ + letI : NormalizationMonoid α := UniqueFactorizationMonoid.normalizationMonoid + suffices + f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, + p ^ ((normalizedFactors a).count p + (normalizedFactors b).count p)) = + f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, + p ^ (normalizedFactors a).count p) * + f (∏ p ∈ (normalizedFactors a).toFinset ∪ (normalizedFactors b).toFinset, + p ^ (normalizedFactors b).count p) by + obtain ⟨ua, a_eq⟩ := normalizedFactors_prod ha0 + obtain ⟨ub, b_eq⟩ := normalizedFactors_prod hb0 + rw [← a_eq, ← b_eq, mul_right_comm (Multiset.prod (normalizedFactors a)) ua + (Multiset.prod (normalizedFactors b) * ub), h1 ua.isUnit, h1 ub.isUnit, h1 ua.isUnit, ← + mul_assoc, h1 ub.isUnit, mul_right_comm _ (f ua), ← mul_assoc] + congr + rw [← (normalizedFactors a).map_id, ← (normalizedFactors b).map_id, + Finset.prod_multiset_map_count, Finset.prod_multiset_map_count, + Finset.prod_subset (Finset.subset_union_left (s₂ := (normalizedFactors b).toFinset)), + Finset.prod_subset (Finset.subset_union_right (s₂ := (normalizedFactors b).toFinset)), ← + Finset.prod_mul_distrib] + · simp_rw [id, ← pow_add, this] + all_goals simp only [Multiset.mem_toFinset] + · intro p _ hpb + simp [hpb] + · intro p _ hpa + simp [hpa] + refine multiplicative_prime_power _ _ _ ?_ ?_ @h1 @hpr @hcp + all_goals simp only [Multiset.mem_toFinset, Finset.mem_union] + · rintro p (hpa | hpb) <;> apply prime_of_normalized_factor <;> assumption + · rintro p (hp | hp) q (hq | hq) hdvd <;> + rw [← normalize_normalized_factor _ hp, ← normalize_normalized_factor _ hq] <;> + exact + normalize_eq_normalize hdvd + ((prime_of_normalized_factor _ hp).irreducible.dvd_symm + (prime_of_normalized_factor _ hq).irreducible hdvd) + +end Multiplicative + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean new file mode 100644 index 0000000000000..dbc52cb56fad1 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Multiplicity.lean @@ -0,0 +1,126 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.RingTheory.Multiplicity +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Unique factorization and multiplicity + +## Main results + +* `UniqueFactorizationMonoid.emultiplicity_eq_count_normalizedFactors`: The multiplicity of an + irreducible factor of a nonzero element is exactly the number of times the normalized factor + occurs in the `normalizedFactors`. +-/ + + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +theorem WfDvdMonoid.max_power_factor' [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} + (h : a₀ ≠ 0) (hx : ¬IsUnit x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := by + obtain ⟨a, ⟨n, rfl⟩, hm⟩ := wellFounded_dvdNotUnit.has_min + {a | ∃ n, x ^ n * a = a₀} ⟨a₀, 0, by rw [pow_zero, one_mul]⟩ + refine ⟨n, a, ?_, rfl⟩; rintro ⟨d, rfl⟩ + exact hm d ⟨n + 1, by rw [pow_succ, mul_assoc]⟩ + ⟨(right_ne_zero_of_mul <| right_ne_zero_of_mul h), x, hx, mul_comm _ _⟩ + +theorem WfDvdMonoid.max_power_factor [CommMonoidWithZero α] [WfDvdMonoid α] {a₀ x : α} + (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ (n : ℕ) (a : α), ¬x ∣ a ∧ a₀ = x ^ n * a := + max_power_factor' h hx.not_unit + +theorem multiplicity.finite_of_not_isUnit [CancelCommMonoidWithZero α] [WfDvdMonoid α] + {a b : α} (ha : ¬IsUnit a) (hb : b ≠ 0) : multiplicity.Finite a b := by + obtain ⟨n, c, ndvd, rfl⟩ := WfDvdMonoid.max_power_factor' hb ha + exact ⟨n, by rwa [pow_succ, mul_dvd_mul_iff_left (left_ne_zero_of_mul hb)]⟩ + +namespace UniqueFactorizationMonoid + +variable {R : Type*} [CancelCommMonoidWithZero R] [UniqueFactorizationMonoid R] + +section multiplicity + +variable [NormalizationMonoid R] + +open Multiset + +section + +theorem le_emultiplicity_iff_replicate_le_normalizedFactors {a b : R} {n : ℕ} (ha : Irreducible a) + (hb : b ≠ 0) : + ↑n ≤ emultiplicity a b ↔ replicate n (normalize a) ≤ normalizedFactors b := by + rw [← pow_dvd_iff_le_emultiplicity] + revert b + induction' n with n ih; · simp + intro b hb + constructor + · rintro ⟨c, rfl⟩ + rw [Ne, pow_succ', mul_assoc, mul_eq_zero, not_or] at hb + rw [pow_succ', mul_assoc, normalizedFactors_mul hb.1 hb.2, replicate_succ, + normalizedFactors_irreducible ha, singleton_add, cons_le_cons_iff, ← ih hb.2] + apply Dvd.intro _ rfl + · rw [Multiset.le_iff_exists_add] + rintro ⟨u, hu⟩ + rw [← (normalizedFactors_prod hb).dvd_iff_dvd_right, hu, prod_add, prod_replicate] + exact (Associated.pow_pow <| associated_normalize a).dvd.trans (Dvd.intro u.prod rfl) + +/-- The multiplicity of an irreducible factor of a nonzero element is exactly the number of times +the normalized factor occurs in the `normalizedFactors`. + +See also `count_normalizedFactors_eq` which expands the definition of `multiplicity` +to produce a specification for `count (normalizedFactors _) _`.. +-/ +theorem emultiplicity_eq_count_normalizedFactors [DecidableEq R] {a b : R} (ha : Irreducible a) + (hb : b ≠ 0) : emultiplicity a b = (normalizedFactors b).count (normalize a) := by + apply le_antisymm + · apply Order.le_of_lt_add_one + rw [← Nat.cast_one, ← Nat.cast_add, lt_iff_not_ge, ge_iff_le, + le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] + simp + rw [le_emultiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] + +end + +/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by +the number of times it divides `x`. + +See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. +-/ +theorem count_normalizedFactors_eq [DecidableEq R] {p x : R} (hp : Irreducible p) + (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : + (normalizedFactors x).count p = n := by classical + by_cases hx0 : x = 0 + · simp [hx0] at hlt + apply Nat.cast_injective (R := ℕ∞) + convert (emultiplicity_eq_count_normalizedFactors hp hx0).symm + · exact hnorm.symm + exact (emultiplicity_eq_coe.mpr ⟨hle, hlt⟩).symm + +/-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by +the number of times it divides `x`. This is a slightly more general version of +`UniqueFactorizationMonoid.count_normalizedFactors_eq` that allows `p = 0`. + +See also `multiplicity_eq_count_normalizedFactors` if `n` is given by `multiplicity p x`. +-/ +theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Irreducible p) + (hnorm : normalize p = p) {n : ℕ} (hle : p ^ n ∣ x) (hlt : ¬p ^ (n + 1) ∣ x) : + (normalizedFactors x).count p = n := by + rcases hp with (rfl | hp) + · cases n + · exact count_eq_zero.2 (zero_not_mem_normalizedFactors _) + · rw [zero_pow (Nat.succ_ne_zero _)] at hle hlt + exact absurd hle hlt + · exact count_normalizedFactors_eq hp hnorm hle hlt + +end multiplicity + +/-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ +@[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] +theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : + ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx + +end UniqueFactorizationMonoid diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean new file mode 100644 index 0000000000000..680aa8f4249fe --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean @@ -0,0 +1,62 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.Data.ENat.Basic +import Mathlib.Data.Nat.Factors +import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors + +/-! +# Unique factorization of natural numbers + +## Main definitions + +* `Nat.instUniqueFactorizationMonoid`: the natural numbers have unique factorization +-/ + +variable {α : Type*} + +namespace Nat + +instance instWfDvdMonoid : WfDvdMonoid ℕ where + wf := by + refine RelHomClass.wellFounded + (⟨fun x : ℕ => if x = 0 then (⊤ : ℕ∞) else x, ?_⟩ : DvdNotUnit →r (· < ·)) wellFounded_lt + intro a b h + cases' a with a + · exfalso + revert h + simp [DvdNotUnit] + cases b + · simpa [succ_ne_zero] using WithTop.coe_lt_top (a + 1) + cases' dvd_and_not_dvd_iff.2 h with h1 h2 + simp only [succ_ne_zero, cast_lt, if_false] + refine lt_of_le_of_ne (Nat.le_of_dvd (Nat.succ_pos _) h1) fun con => h2 ?_ + rw [con] + +instance instUniqueFactorizationMonoid : UniqueFactorizationMonoid ℕ where + irreducible_iff_prime := Nat.irreducible_iff_prime + +open UniqueFactorizationMonoid + +lemma factors_eq : ∀ n : ℕ, normalizedFactors n = n.primeFactorsList + | 0 => by simp + | n + 1 => by + rw [← Multiset.rel_eq, ← associated_eq_eq] + apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor _ + · rw [Multiset.prod_coe, Nat.prod_primeFactorsList n.succ_ne_zero] + exact normalizedFactors_prod n.succ_ne_zero + · intro x hx + rw [Nat.irreducible_iff_prime, ← Nat.prime_iff] + exact Nat.prime_of_mem_primeFactorsList hx + +lemma factors_multiset_prod_of_irreducible {s : Multiset ℕ} (h : ∀ x : ℕ, x ∈ s → Irreducible x) : + normalizedFactors s.prod = s := by + rw [← Multiset.rel_eq, ← associated_eq_eq] + apply UniqueFactorizationMonoid.factors_unique irreducible_of_normalized_factor h + (normalizedFactors_prod _) + rw [Ne, Multiset.prod_eq_zero_iff] + exact fun con ↦ not_irreducible_zero (h 0 con) + +end Nat diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean new file mode 100644 index 0000000000000..3d30df0c3b813 --- /dev/null +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/NormalizedFactors.lean @@ -0,0 +1,329 @@ +/- +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, Aaron Anderson +-/ +import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic + +/-! +# Unique factorization and normalization + +## Main definitions +* `UniqueFactorizationMonoid.normalizedFactors`: choose a multiset of prime factors that are unique + by normalizing them. +* `UniqueFactorizationMonoid.normalizationMonoid`: choose a way of normalizing the elements of a UFM +-/ + + +variable {α : Type*} + +local infixl:50 " ~ᵤ " => Associated + +namespace UniqueFactorizationMonoid + +variable [CancelCommMonoidWithZero α] [NormalizationMonoid α] +variable [UniqueFactorizationMonoid α] + +/-- Noncomputably determines the multiset of prime factors. -/ +noncomputable def normalizedFactors (a : α) : Multiset α := + Multiset.map normalize <| factors a + +/-- An arbitrary choice of factors of `x : M` is exactly the (unique) normalized set of factors, +if `M` has a trivial group of units. -/ +@[simp] +theorem factors_eq_normalizedFactors {M : Type*} [CancelCommMonoidWithZero M] + [UniqueFactorizationMonoid M] [Subsingleton Mˣ] (x : M) : factors x = normalizedFactors x := by + unfold normalizedFactors + convert (Multiset.map_id (factors x)).symm + ext p + exact normalize_eq p + +theorem normalizedFactors_prod {a : α} (ane0 : a ≠ 0) : + Associated (normalizedFactors a).prod a := by + rw [normalizedFactors, factors, dif_neg ane0] + refine Associated.trans ?_ (Classical.choose_spec (exists_prime_factors a ane0)).2 + rw [← Associates.mk_eq_mk_iff_associated, ← Associates.prod_mk, ← Associates.prod_mk, + Multiset.map_map] + congr 2 + ext + rw [Function.comp_apply, Associates.mk_normalize] + +theorem prime_of_normalized_factor {a : α} : ∀ x : α, x ∈ normalizedFactors a → Prime x := by + rw [normalizedFactors, factors] + split_ifs with ane0; · simp + intro x hx; rcases Multiset.mem_map.1 hx with ⟨y, ⟨hy, rfl⟩⟩ + rw [(normalize_associated _).prime_iff] + exact (Classical.choose_spec (UniqueFactorizationMonoid.exists_prime_factors a ane0)).1 y hy + +theorem irreducible_of_normalized_factor {a : α} : + ∀ x : α, x ∈ normalizedFactors a → Irreducible x := fun x h => + (prime_of_normalized_factor x h).irreducible + +theorem normalize_normalized_factor {a : α} : + ∀ x : α, x ∈ normalizedFactors a → normalize x = x := by + rw [normalizedFactors, factors] + split_ifs with h; · simp + intro x hx + obtain ⟨y, _, rfl⟩ := Multiset.mem_map.1 hx + apply normalize_idem + +theorem normalizedFactors_irreducible {a : α} (ha : Irreducible a) : + normalizedFactors a = {normalize a} := by + obtain ⟨p, a_assoc, hp⟩ := + prime_factors_irreducible ha ⟨prime_of_normalized_factor, normalizedFactors_prod ha.ne_zero⟩ + have p_mem : p ∈ normalizedFactors a := by + rw [hp] + exact Multiset.mem_singleton_self _ + convert hp + rwa [← normalize_normalized_factor p p_mem, normalize_eq_normalize_iff, dvd_dvd_iff_associated] + +theorem normalizedFactors_eq_of_dvd (a : α) : + ∀ᵉ (p ∈ normalizedFactors a) (q ∈ normalizedFactors a), p ∣ q → p = q := by + intro p hp q hq hdvd + convert normalize_eq_normalize hdvd + ((prime_of_normalized_factor _ hp).irreducible.dvd_symm + (prime_of_normalized_factor _ hq).irreducible hdvd) <;> + apply (normalize_normalized_factor _ ‹_›).symm + +theorem exists_mem_normalizedFactors_of_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : + p ∣ a → ∃ q ∈ normalizedFactors a, p ~ᵤ q := fun ⟨b, hb⟩ => + have hb0 : b ≠ 0 := fun hb0 => by simp_all + have : Multiset.Rel Associated (p ::ₘ normalizedFactors b) (normalizedFactors a) := + factors_unique + (fun _ hx => + (Multiset.mem_cons.1 hx).elim (fun h => h.symm ▸ hp) (irreducible_of_normalized_factor _)) + irreducible_of_normalized_factor + (Associated.symm <| + calc + Multiset.prod (normalizedFactors a) ~ᵤ a := normalizedFactors_prod ha0 + _ = p * b := hb + _ ~ᵤ Multiset.prod (p ::ₘ normalizedFactors b) := by + rw [Multiset.prod_cons] + exact (normalizedFactors_prod hb0).symm.mul_left _ + ) + Multiset.exists_mem_of_rel_of_mem this (by simp) + +theorem exists_mem_normalizedFactors {x : α} (hx : x ≠ 0) (h : ¬IsUnit x) : + ∃ p, p ∈ normalizedFactors x := by + obtain ⟨p', hp', hp'x⟩ := WfDvdMonoid.exists_irreducible_factor h hx + obtain ⟨p, hp, _⟩ := exists_mem_normalizedFactors_of_dvd hx hp' hp'x + exact ⟨p, hp⟩ + +@[simp] +theorem normalizedFactors_zero : normalizedFactors (0 : α) = 0 := by + simp [normalizedFactors, factors] + +@[simp] +theorem normalizedFactors_one : normalizedFactors (1 : α) = 0 := by + cases' subsingleton_or_nontrivial α with h h + · dsimp [normalizedFactors, factors] + simp [Subsingleton.elim (1 : α) 0] + · rw [← Multiset.rel_zero_right] + apply factors_unique irreducible_of_normalized_factor + · intro x hx + exfalso + apply Multiset.not_mem_zero x hx + · apply normalizedFactors_prod one_ne_zero + +@[simp] +theorem normalizedFactors_mul {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + normalizedFactors (x * y) = normalizedFactors x + normalizedFactors y := by + have h : (normalize : α → α) = Associates.out ∘ Associates.mk := by + ext + rw [Function.comp_apply, Associates.out_mk] + rw [← Multiset.map_id' (normalizedFactors (x * y)), ← Multiset.map_id' (normalizedFactors x), ← + Multiset.map_id' (normalizedFactors y), ← Multiset.map_congr rfl normalize_normalized_factor, ← + Multiset.map_congr rfl normalize_normalized_factor, ← + Multiset.map_congr rfl normalize_normalized_factor, ← Multiset.map_add, h, ← + Multiset.map_map Associates.out, eq_comm, ← Multiset.map_map Associates.out] + refine congr rfl ?_ + apply Multiset.map_mk_eq_map_mk_of_rel + apply factors_unique + · intro x hx + rcases Multiset.mem_add.1 hx with (hx | hx) <;> exact irreducible_of_normalized_factor x hx + · exact irreducible_of_normalized_factor + · rw [Multiset.prod_add] + exact + ((normalizedFactors_prod hx).mul_mul (normalizedFactors_prod hy)).trans + (normalizedFactors_prod (mul_ne_zero hx hy)).symm + +@[simp] +theorem normalizedFactors_pow {x : α} (n : ℕ) : + normalizedFactors (x ^ n) = n • normalizedFactors x := by + induction' n with n ih + · simp + by_cases h0 : x = 0 + · simp [h0, zero_pow n.succ_ne_zero, smul_zero] + rw [pow_succ', succ_nsmul', normalizedFactors_mul h0 (pow_ne_zero _ h0), ih] + +theorem _root_.Irreducible.normalizedFactors_pow {p : α} (hp : Irreducible p) (k : ℕ) : + normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by + rw [UniqueFactorizationMonoid.normalizedFactors_pow, normalizedFactors_irreducible hp, + Multiset.nsmul_singleton] + +theorem normalizedFactors_prod_eq (s : Multiset α) (hs : ∀ a ∈ s, Irreducible a) : + normalizedFactors s.prod = s.map normalize := by + induction' s using Multiset.induction with a s ih + · rw [Multiset.prod_zero, normalizedFactors_one, Multiset.map_zero] + · have ia := hs a (Multiset.mem_cons_self a _) + have ib := fun b h => hs b (Multiset.mem_cons_of_mem h) + obtain rfl | ⟨b, hb⟩ := s.empty_or_exists_mem + · rw [Multiset.cons_zero, Multiset.prod_singleton, Multiset.map_singleton, + normalizedFactors_irreducible ia] + haveI := nontrivial_of_ne b 0 (ib b hb).ne_zero + rw [Multiset.prod_cons, Multiset.map_cons, + normalizedFactors_mul ia.ne_zero (Multiset.prod_ne_zero fun h => (ib 0 h).ne_zero rfl), + normalizedFactors_irreducible ia, ih ib, Multiset.singleton_add] + +theorem dvd_iff_normalizedFactors_le_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + x ∣ y ↔ normalizedFactors x ≤ normalizedFactors y := by + constructor + · rintro ⟨c, rfl⟩ + simp [hx, right_ne_zero_of_mul hy] + · rw [← (normalizedFactors_prod hx).dvd_iff_dvd_left, ← + (normalizedFactors_prod hy).dvd_iff_dvd_right] + apply Multiset.prod_dvd_prod_of_le + +theorem _root_.Associated.normalizedFactors_eq {a b : α} (h : Associated a b) : + normalizedFactors a = normalizedFactors b := by + unfold normalizedFactors + have h' : ⇑(normalize (α := α)) = Associates.out ∘ Associates.mk := funext Associates.out_mk + rw [h', ← Multiset.map_map, ← Multiset.map_map, + Associates.rel_associated_iff_map_eq_map.mp (factors_rel_of_associated h)] + +theorem associated_iff_normalizedFactors_eq_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + x ~ᵤ y ↔ normalizedFactors x = normalizedFactors y := + ⟨Associated.normalizedFactors_eq, fun h => + (normalizedFactors_prod hx).symm.trans (_root_.trans (by rw [h]) (normalizedFactors_prod hy))⟩ + +theorem normalizedFactors_of_irreducible_pow {p : α} (hp : Irreducible p) (k : ℕ) : + normalizedFactors (p ^ k) = Multiset.replicate k (normalize p) := by + rw [normalizedFactors_pow, normalizedFactors_irreducible hp, Multiset.nsmul_singleton] + +theorem zero_not_mem_normalizedFactors (x : α) : (0 : α) ∉ normalizedFactors x := fun h => + Prime.ne_zero (prime_of_normalized_factor _ h) rfl + +theorem dvd_of_mem_normalizedFactors {a p : α} (H : p ∈ normalizedFactors a) : p ∣ a := by + by_cases hcases : a = 0 + · rw [hcases] + exact dvd_zero p + · exact dvd_trans (Multiset.dvd_prod H) (Associated.dvd (normalizedFactors_prod hcases)) + +theorem mem_normalizedFactors_iff [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) : + p ∈ normalizedFactors x ↔ Prime p ∧ p ∣ x := by + constructor + · intro h + exact ⟨prime_of_normalized_factor p h, dvd_of_mem_normalizedFactors h⟩ + · rintro ⟨hprime, hdvd⟩ + obtain ⟨q, hqmem, hqeq⟩ := exists_mem_normalizedFactors_of_dvd hx hprime.irreducible hdvd + rw [associated_iff_eq] at hqeq + exact hqeq ▸ hqmem + +theorem exists_associated_prime_pow_of_unique_normalized_factor {p r : α} + (h : ∀ {m}, m ∈ normalizedFactors r → m = p) (hr : r ≠ 0) : ∃ i : ℕ, Associated (p ^ i) r := by + use Multiset.card.toFun (normalizedFactors r) + have := UniqueFactorizationMonoid.normalizedFactors_prod hr + rwa [Multiset.eq_replicate_of_mem fun b => h, Multiset.prod_replicate] at this + +theorem normalizedFactors_prod_of_prime [Subsingleton αˣ] {m : Multiset α} + (h : ∀ p ∈ m, Prime p) : normalizedFactors m.prod = m := by + cases subsingleton_or_nontrivial α + · obtain rfl : m = 0 := by + refine Multiset.eq_zero_of_forall_not_mem fun x hx ↦ ?_ + simpa [Subsingleton.elim x 0] using h x hx + simp + · simpa only [← Multiset.rel_eq, ← associated_eq_eq] using + prime_factors_unique prime_of_normalized_factor h + (normalizedFactors_prod (m.prod_ne_zero_of_prime h)) + +theorem mem_normalizedFactors_eq_of_associated {a b c : α} (ha : a ∈ normalizedFactors c) + (hb : b ∈ normalizedFactors c) (h : Associated a b) : a = b := by + rw [← normalize_normalized_factor a ha, ← normalize_normalized_factor b hb, + normalize_eq_normalize_iff] + exact Associated.dvd_dvd h + +@[simp] +theorem normalizedFactors_pos (x : α) (hx : x ≠ 0) : 0 < normalizedFactors x ↔ ¬IsUnit x := by + constructor + · intro h hx + obtain ⟨p, hp⟩ := Multiset.exists_mem_of_ne_zero h.ne' + exact + (prime_of_normalized_factor _ hp).not_unit + (isUnit_of_dvd_unit (dvd_of_mem_normalizedFactors hp) hx) + · intro h + obtain ⟨p, hp⟩ := exists_mem_normalizedFactors hx h + exact + bot_lt_iff_ne_bot.mpr + (mt Multiset.eq_zero_iff_forall_not_mem.mp (not_forall.mpr ⟨p, not_not.mpr hp⟩)) + +theorem dvdNotUnit_iff_normalizedFactors_lt_normalizedFactors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + DvdNotUnit x y ↔ normalizedFactors x < normalizedFactors y := by + constructor + · rintro ⟨_, c, hc, rfl⟩ + simp only [hx, right_ne_zero_of_mul hy, normalizedFactors_mul, Ne, not_false_iff, + lt_add_iff_pos_right, normalizedFactors_pos, hc] + · intro h + exact + dvdNotUnit_of_dvd_of_not_dvd + ((dvd_iff_normalizedFactors_le_normalizedFactors hx hy).mpr h.le) + (mt (dvd_iff_normalizedFactors_le_normalizedFactors hy hx).mp h.not_le) + +theorem normalizedFactors_multiset_prod (s : Multiset α) (hs : 0 ∉ s) : + normalizedFactors (s.prod) = (s.map normalizedFactors).sum := by + cases subsingleton_or_nontrivial α + · obtain rfl : s = 0 := by + apply Multiset.eq_zero_of_forall_not_mem + intro _ + convert hs + simp + induction s using Multiset.induction with + | empty => simp + | cons _ _ IH => + rw [Multiset.prod_cons, Multiset.map_cons, Multiset.sum_cons, normalizedFactors_mul, IH] + · exact fun h ↦ hs (Multiset.mem_cons_of_mem h) + · exact fun h ↦ hs (h ▸ Multiset.mem_cons_self _ _) + · apply Multiset.prod_ne_zero + exact fun h ↦ hs (Multiset.mem_cons_of_mem h) + +end UniqueFactorizationMonoid + +namespace UniqueFactorizationMonoid + +open scoped Classical + +open Multiset Associates + +variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] + +/-- Noncomputably defines a `normalizationMonoid` structure on a `UniqueFactorizationMonoid`. -/ +protected noncomputable def normalizationMonoid : NormalizationMonoid α := + normalizationMonoidOfMonoidHomRightInverse + { toFun := fun a : Associates α => + if a = 0 then 0 + else + ((normalizedFactors a).map + (Classical.choose mk_surjective.hasRightInverse : Associates α → α)).prod + map_one' := by nontriviality α; simp + map_mul' := fun x y => by + by_cases hx : x = 0 + · simp [hx] + by_cases hy : y = 0 + · simp [hy] + simp [hx, hy] } + (by + intro x + dsimp + by_cases hx : x = 0 + · simp [hx] + have h : Associates.mkMonoidHom ∘ Classical.choose mk_surjective.hasRightInverse = + (id : Associates α → Associates α) := by + ext x + rw [Function.comp_apply, mkMonoidHom_apply, + Classical.choose_spec mk_surjective.hasRightInverse x] + rfl + rw [if_neg hx, ← mkMonoidHom_apply, MonoidHom.map_multiset_prod, map_map, h, map_id, ← + associated_iff_eq] + apply normalizedFactors_prod hx) + +end UniqueFactorizationMonoid diff --git a/MathlibTest/Variable.lean b/MathlibTest/Variable.lean index 327d36efd228e..e93a74e81fbec 100644 --- a/MathlibTest/Variable.lean +++ b/MathlibTest/Variable.lean @@ -1,8 +1,10 @@ -import Mathlib.Tactic.Variable -import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.Field.Defs +import Mathlib.Algebra.Module.Defs +import Mathlib.Algebra.Ring.Regular import Mathlib.Algebra.Module.LinearMap.Basic -import Mathlib.RingTheory.UniqueFactorizationDomain +import Mathlib.RingTheory.UniqueFactorizationDomain.Basic +import Mathlib.Tactic.Variable set_option autoImplicit true namespace Tests From 9da26514ae8782fb1fcf86b70fb4acd92c21c7ba Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 20 Nov 2024 11:43:02 +0000 Subject: [PATCH 58/90] chore: bump dependencies (#19279) --- Mathlib/Algebra/Order/Group/Unbundled/Basic.lean | 1 + lake-manifest.json | 2 +- 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 03c9d7579e5f1..5b90caf6e8964 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -537,6 +537,7 @@ section LE variable [LE α] [MulLeftMono α] {a b c d : α} +set_option linter.docPrime false in @[to_additive sub_le_sub_iff] theorem div_le_div_iff' : a / b ≤ c / d ↔ a * d ≤ c * b := by simpa only [div_eq_mul_inv] using mul_inv_le_mul_inv_iff' diff --git a/lake-manifest.json b/lake-manifest.json index cfc65dc4caaa9..d40bdeb3be7db 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "86d0d0584f5cd165353e2f8a30c455cd0e168ac2", + "rev": "d7caecce0d0f003fd5e9cce9a61f1dd6ba83142b", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", From cd4d2185e67b7387e303150dab86521007c56fb6 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 20 Nov 2024 12:32:28 +0000 Subject: [PATCH 59/90] feat(KrullDimension): concrete calculations (#19210) Part of https://github.com/leanprover-community/mathlib4/pull/15524. This adds concrete calculations of the height and Krull Dimension for Nat, Int, WithTop, WithBot and ENat. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 158 +++++++++++++++++++++++++++++- 1 file changed, 156 insertions(+), 2 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 836303c212832..249a8d2cc0301 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang, Fangming Li, Joachim Breitner -/ -import Mathlib.Order.RelSeries -import Mathlib.Order.Minimal +import Mathlib.Algebra.Order.Group.Int import Mathlib.Data.ENat.Lattice +import Mathlib.Order.Minimal +import Mathlib.Order.RelSeries /-! # Krull dimension of a preordered set and height of an element @@ -648,4 +649,157 @@ lemma coheight_bot_eq_krullDim [OrderBot α] : coheight (⊥ : α) = krullDim α end krullDim +/-! +## Concrete calculations +-/ + +section calculations + +variable {α : Type*} [Preorder α] + +/- +These two lemmas could possibly be used to simplify the subsequent calculations, +especially once the `Set.encard` api is richer. + +(Commented out to avoid importing modules purely for `proof_wanted`.) +proof_wanted height_of_linearOrder {α : Type*} [LinearOrder α] (a : α) : + height a = (Set.Iio a).encard + +proof_wanted coheight_of_linearOrder {α : Type*} [LinearOrder α] (a : α) : + coheight a = (Set.Ioi a).encard +-/ + +@[simp] lemma height_nat (n : ℕ) : height n = n := by + induction n using Nat.strongRecOn with | ind n ih => + apply le_antisymm + · apply height_le_coe_iff.mpr + simp +contextual only [ih, Nat.cast_lt, implies_true] + · exact length_le_height_last (p := LTSeries.range n) + +@[simp] lemma coheight_of_noMaxOrder [NoMaxOrder α] (a : α) : coheight a = ⊤ := by + obtain ⟨f, hstrictmono⟩ := Nat.exists_strictMono ↑(Set.Ioi a) + apply coheight_eq_top_iff.mpr + intro m + use {length := m, toFun := fun i => if i = 0 then a else f i, step := ?step } + case h => simp [RelSeries.head] + case step => + intro ⟨i, hi⟩ + by_cases hzero : i = 0 + · subst i + exact (f 1).prop + · suffices f i < f (i + 1) by simp [Fin.ext_iff, hzero, this] + apply hstrictmono + omega + +@[simp] lemma height_of_noMinOrder [NoMinOrder α] (a : α) : height a = ⊤ := + -- Implementation note: Here it's a bit easier to define the coheight variant first + coheight_of_noMaxOrder (α := αᵒᵈ) a + +@[simp] lemma krullDim_of_noMaxOrder [Nonempty α] [NoMaxOrder α] : krullDim α = ⊤ := by + simp [krullDim_eq_iSup_coheight, coheight_of_noMaxOrder] + +@[simp] lemma krullDim_of_noMinOrder [Nonempty α] [NoMinOrder α] : krullDim α = ⊤ := by + simp [krullDim_eq_iSup_height, height_of_noMinOrder] + +lemma coheight_nat (n : ℕ) : coheight n = ⊤ := coheight_of_noMaxOrder .. + +lemma krullDim_nat : krullDim ℕ = ⊤ := krullDim_of_noMaxOrder .. + +lemma height_int (n : ℤ) : height n = ⊤ := height_of_noMinOrder .. + +lemma coheight_int (n : ℤ) : coheight n = ⊤ := coheight_of_noMaxOrder .. + +lemma krullDim_int : krullDim ℤ = ⊤ := krullDim_of_noMaxOrder .. + +@[simp] lemma height_coe_withBot (x : α) : height (x : WithBot α) = height x + 1 := by + apply le_antisymm + · apply height_le + intro p hlast + wlog hlenpos : p.length ≠ 0 + · simp_all + -- essentially p' := (p.drop 1).map unbot + let p' : LTSeries α := { + length := p.length - 1 + toFun := fun ⟨i, hi⟩ => (p ⟨i+1, by omega⟩).unbot (by + apply LT.lt.ne_bot (a := p.head) + apply p.strictMono + exact compare_gt_iff_gt.mp rfl) + step := fun i => by simpa [WithBot.unbot_lt_iff] using p.step ⟨i + 1, by omega⟩ } + have hlast' : p'.last = x := by + simp only [RelSeries.last, Fin.val_last, WithBot.unbot_eq_iff, ← hlast, Fin.last] + congr + omega + suffices p'.length ≤ height p'.last by + simpa [p', hlast'] using this + apply length_le_height_last + · rw [height_add_const] + apply iSup₂_le + intro p hlast + let p' := (p.map _ WithBot.coe_strictMono).cons ⊥ (by simp) + apply le_iSup₂_of_le p' (by simp [p', hlast]) (by simp [p']) + +@[simp] lemma coheight_coe_withTop (x : α) : coheight (x : WithTop α) = coheight x + 1 := + height_coe_withBot (α := αᵒᵈ) x + +@[simp] lemma height_coe_withTop (x : α) : height (x : WithTop α) = height x := by + apply le_antisymm + · apply height_le + intro p hlast + -- essentially p' := p.map untop + let p' : LTSeries α := { + length := p.length + toFun := fun i => (p i).untop (by + apply WithTop.lt_top_iff_ne_top.mp + apply lt_of_le_of_lt + · exact p.monotone (Fin.le_last _) + · rw [RelSeries.last] at hlast + simp [hlast]) + step := fun i => by simpa only [WithTop.untop_lt_iff, WithTop.coe_untop] using p.step i } + have hlast' : p'.last = x := by + simp only [RelSeries.last, Fin.val_last, WithTop.untop_eq_iff, ← hlast] + suffices p'.length ≤ height p'.last by + rw [hlast'] at this + simpa [p'] using this + apply length_le_height_last + · apply height_le + intro p hlast + let p' := p.map _ WithTop.coe_strictMono + apply le_iSup₂_of_le p' (by simp [p', hlast]) (by simp [p']) + +@[simp] lemma coheight_coe_withBot (x : α) : coheight (x : WithBot α) = coheight x := + height_coe_withTop (α := αᵒᵈ) x + +@[simp] lemma krullDim_WithTop [Nonempty α] : krullDim (WithTop α) = krullDim α + 1 := by + rw [← height_top_eq_krullDim, krullDim_eq_iSup_height_of_nonempty, height_eq_iSup_lt_height] + norm_cast + simp_rw [WithTop.lt_top_iff_ne_top] + rw [ENat.iSup_add, iSup_subtype'] + symm + apply Equiv.withTopSubtypeNe.symm.iSup_congr + simp + +@[simp] lemma krullDim_withBot [Nonempty α] : krullDim (WithBot α) = krullDim α + 1 := by + conv_lhs => rw [← krullDim_orderDual] + conv_rhs => rw [← krullDim_orderDual] + exact krullDim_WithTop (α := αᵒᵈ) + +@[simp] +lemma krullDim_enat : krullDim ℕ∞ = ⊤ := by + show (krullDim (WithTop ℕ) = ⊤) + simp only [krullDim_WithTop, krullDim_nat] + rfl + +@[simp] +lemma height_enat (n : ℕ∞) : height n = n := by + cases n with + | top => simp only [← WithBot.coe_eq_coe, height_top_eq_krullDim, krullDim_enat, WithBot.coe_top] + | coe n => exact (height_coe_withTop _).trans (height_nat _) + +@[simp] +lemma coheight_coe_enat (n : ℕ) : coheight (n : ℕ∞) = ⊤ := by + apply (coheight_coe_withTop _).trans + simp only [Nat.cast_id, coheight_nat, top_add] + +end calculations + end Order From b99018cbf7c28b0f2d28a58d3c974f8150366601 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Wed, 20 Nov 2024 12:42:02 +0000 Subject: [PATCH 60/90] chore: bump to nightly-2024-11-20 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 80611d83b2915..118d9e578af27 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-19 +leanprover/lean4:nightly-2024-11-20 From 79ee32ececd77acd6e5c7ef387de3771c9790787 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 20 Nov 2024 12:45:10 +0000 Subject: [PATCH 61/90] feat: cardinality of `Field.Emb` for algebraic extensions of infinite separable degree (#9480) We show that the number of embeddings of an algebraic extension E/F with infinite `sepDegree F E` is equal to `2 ^ sepDegree F E`. Since `separableClosure F E` has the same number of F-embeddings as E, we may assume E/F is separable, so `sepDegree F E = Module.rank F E` ([E:F]). The first part of the proof (up to `two_le_deg`) shows the existence of a transfinite filtration of E/F by subextensions of length equal to [E:F], such that each step is finite and nontrivial (has degree at least 2). The second part (up to `embEquivPi`) shows the embeddings of these subextensions form an inverse system satisfying the conditions of `InverseSystem.globalEquiv`, and therefore the number of embeddings of E/F is equal to the product of the degrees of the steps. We conclude the file with five theorems stating the main result for separable and general algebraic extensions, incorporating the finite case or not. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Separability.20degree.3F/near/411486335) Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + Mathlib/FieldTheory/CardinalEmb.lean | 342 +++++++++++++++++++++++ Mathlib/FieldTheory/SeparableDegree.lean | 25 +- Mathlib/LinearAlgebra/FreeAlgebra.lean | 12 + Mathlib/Order/SuccPred/Basic.lean | 4 +- 5 files changed, 367 insertions(+), 17 deletions(-) create mode 100644 Mathlib/FieldTheory/CardinalEmb.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0426236be0a65..db0e4160d4515 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2916,6 +2916,7 @@ import Mathlib.FieldTheory.AbsoluteGaloisGroup import Mathlib.FieldTheory.Adjoin import Mathlib.FieldTheory.AlgebraicClosure import Mathlib.FieldTheory.AxGrothendieck +import Mathlib.FieldTheory.CardinalEmb import Mathlib.FieldTheory.Cardinality import Mathlib.FieldTheory.ChevalleyWarning import Mathlib.FieldTheory.Differential.Basic diff --git a/Mathlib/FieldTheory/CardinalEmb.lean b/Mathlib/FieldTheory/CardinalEmb.lean new file mode 100644 index 0000000000000..1d4c21d21c883 --- /dev/null +++ b/Mathlib/FieldTheory/CardinalEmb.lean @@ -0,0 +1,342 @@ +/- +Copyright (c) 2024 Junyan Xu. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Junyan Xu +-/ +import Mathlib.FieldTheory.SeparableClosure +import Mathlib.FieldTheory.PurelyInseparable +import Mathlib.LinearAlgebra.FreeAlgebra +import Mathlib.Order.Interval.Set.WithBotTop +import Mathlib.Order.DirectedInverseSystem + +/-! +# Number of embeddings of an algebraic extension of infinite separable degree + +## Main results + +- `Field.Emb.cardinal_eq_two_pow_rank` : if `E/F` is an algebraic separable field extension +of infinite degree, then `#(Field.Emb F E) = 2 ^ Module.rank F E`. +This is in contrast to the case of finite degree, where `#(Field.Emb F E) = Module.rank F E`. + +- `Field.Emb.cardinal_eq_two_pow_sepDegree`: more generally, if `E/F` is an algebraic +extension of infinite separable degree, then `#(Field.Emb F E) = 2 ^ Field.sepDegree F E`. + +## Sketch of the proof + +We use a transfinite recursive construction that is fairly standard in set theory, but the author +has not seen similar arguments elsewhere in mathlib, and some parts proved tricky to formalize. + +The extension `E/F` can be filtered by intermediate fields indexed by a well-order: +simply put a well-order on a basis of `E/F`, and at each step, take the smallest basis element +that is not contained in the intermediate field generated by all previous elements, so that they +generate a strictly larger intermediate field together. This process can extend all the way +to the initial ordinal `ι` of the cardinal `Module.rank F E`, because the dimension of the +subalgebra generated by an infinite set cannot be greater than the cardinality of the set, and +in an algebraic extension, any subalgebra is a field. This is proven as `Algebra.rank_adjoin_le` +and used to show `leastExt` is a total function from `ι` to itself. It is probably +mathematically the most nontrivial part of the whole argument, but turned out easy to +formalize and was done the earliest. + +Once we have the filtration `E⟮ (Module.rank F E).ord.toType + +private local instance : SuccOrder ι := SuccOrder.ofLinearWellFoundedLT ι +local notation i"⁺" => succ i -- Note: conflicts with `PosPart` notation + +/-- A basis of E/F indexed by the initial ordinal. -/ +def wellOrderedBasis : Basis ι F E := + (chooseBasis F E).reindex + (Cardinal.eq.mp <| (mk_ord_toType _).trans <| rank_eq_card_chooseBasisIndex F E).some.symm + +local notation "b" => wellOrderedBasis F E +local notation "Ē" => AlgebraicClosure E + +variable {F E} + +theorem adjoin_basis_eq_top : adjoin F (range b) = ⊤ := + toSubalgebra_injective <| Subalgebra.toSubmodule_injective <| top_unique <| + (Basis.span_eq b).ge.trans <| (Algebra.span_le_adjoin F _).trans <| algebra_adjoin_le_adjoin _ _ + +section Algebraic + +variable [rank_inf : Fact (ℵ₀ ≤ Module.rank F E)] + +lemma noMaxOrder_rank_toType : NoMaxOrder ι := Cardinal.noMaxOrder Fact.out +attribute [local instance] noMaxOrder_rank_toType + +open _root_.Algebra (IsAlgebraic) +variable [IsAlgebraic F E] + +variable (F E) in +/-- `leastExt i` is defined to be the smallest `k : ι` that generates a nontrival extension over +(i.e. does not lie in) the subalgebra (= intermediate field) generated by all previous +`leastExt j`, `j < i`. For cardinality reasons, such `k` always exist if `ι` is infinite. -/ +def leastExt : ι → ι := + wellFounded_lt.fix fun i ih ↦ + let s := range fun j : Iio i ↦ b (ih j j.2) + wellFounded_lt.min {k | b k ∉ adjoin F s} <| by + rw [← compl_setOf, nonempty_compl]; by_contra! + simp_rw [eq_univ_iff_forall, mem_setOf] at this + have := adjoin_le_iff.mpr (range_subset_iff.mpr this) + rw [adjoin_basis_eq_top, ← eq_top_iff] at this + apply_fun Module.rank F at this + refine ne_of_lt ?_ this + conv_rhs => rw [topEquiv.toLinearEquiv.rank_eq] + have := mk_Iio_ord_toType i + obtain eq | lt := rank_inf.out.eq_or_lt + · replace this := mk_lt_aleph0_iff.mp (this.trans_eq eq.symm) + have : FiniteDimensional F (adjoin F s) := + finiteDimensional_adjoin fun x _ ↦ (IsAlgebraic.isAlgebraic x).isIntegral + exact (Module.rank_lt_aleph0 _ _).trans_eq eq + · exact (Subalgebra.equivOfEq _ _ <| adjoin_algebraic_toSubalgebra + fun x _ ↦ IsAlgebraic.isAlgebraic x)|>.toLinearEquiv.rank_eq.trans_lt <| + (Algebra.rank_adjoin_le _).trans_lt (max_lt (mk_range_le.trans_lt this) lt) + +local notation "φ" => leastExt F E + +section +local notation "E⟮<"i"⟯" => adjoin F (b ∘ φ '' Iio i) + +theorem isLeast_leastExt (i : ι) : IsLeast {k | b k ∉ E⟮ Field.Emb (E⟮ filtration i + +variable (F E) in +/-- The functor on `WithTop ι` given by embeddings of `E⟮ rw [← mk_ord_toType (Module.rank F E), ← prod_const'] + exact prod_le_prod _ _ fun i ↦ (Emb.Cardinal.deg_lt_aleph0 _).le + · conv_lhs => rw [← mk_ord_toType (Module.rank F E), ← prod_const'] + exact prod_le_prod _ _ Emb.Cardinal.two_le_deg + +theorem cardinal_eq_of_isSeparable [Algebra.IsSeparable F E] : + #(Field.Emb F E) = (fun c ↦ if ℵ₀ ≤ c then 2 ^ c else c) (Module.rank F E) := by + dsimp only; split_ifs with h + · exact cardinal_eq_two_pow_rank h + rw [not_le, ← IsNoetherian.iff_rank_lt_aleph0] at h + rw [← Module.finrank_eq_rank, ← toNat_eq_iff Module.finrank_pos.ne', + ← Nat.card, ← finSepDegree, finSepDegree_eq_finrank_of_isSeparable] + +theorem cardinal_eq_two_pow_sepDegree [Algebra.IsAlgebraic F E] + (rank_inf : ℵ₀ ≤ sepDegree F E) : #(Field.Emb F E) = 2 ^ sepDegree F E := by + rw [← cardinal_separableClosure, cardinal_eq_two_pow_rank rank_inf] + rfl + +theorem cardinal_eq [Algebra.IsAlgebraic F E] : + #(Field.Emb F E) = (fun c ↦ if ℵ₀ ≤ c then 2 ^ c else c) (sepDegree F E) := by + rw [← cardinal_separableClosure, cardinal_eq_of_isSeparable]; rfl + +end Field.Emb diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 8025c4eff3c76..c6a67370781a2 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -24,26 +24,21 @@ This file contains basics about the separable degree of a field extension. advantage that `Field.Emb F E` lies in the same universe as `E` rather than the maximum over `F` and `E`). Usually denoted by $\operatorname{Emb}_F(E)$ in textbooks. - **Remark:** if `E / F` is not algebraic, then this definition makes no mathematical sense, - and if it is infinite, then its cardinality doesn't behave as expected (namely, not equal to the - field extension degree of `separableClosure F E / F`). For example, if $F = \mathbb{Q}$ and - $E = \mathbb{Q}( \mu_{p^\infty} )$, then $\operatorname{Emb}_F (E)$ is in bijection with - $\operatorname{Gal}(E/F)$, which is isomorphic to - $\mathbb{Z}_p^\times$, which is uncountable, while $[E:F]$ is countable. - - **TODO:** prove or disprove that if `E / F` is algebraic and `Emb F E` is infinite, then - `Field.Emb F E` has cardinality `2 ^ Module.rank F (separableClosure F E)`. - -- `Field.finSepDegree F E`: the (finite) separable degree $[E:F]_s$ of an algebraic extension - `E / F` of fields, defined to be the number of `F`-algebra homomorphisms from `E` to the algebraic +- `Field.finSepDegree F E`: the (finite) separable degree $[E:F]_s$ of an extension `E / F` + of fields, defined to be the number of `F`-algebra homomorphisms from `E` to the algebraic closure of `E`, as a natural number. It is zero if `Field.Emb F E` is not finite. Note that if `E / F` is not algebraic, then this definition makes no mathematical sense. **Remark:** the `Cardinal`-valued, potentially infinite separable degree `Field.sepDegree F E` for a general algebraic extension `E / F` is defined to be the degree of `L / F`, where `L` is - the (relative) separable closure `separableClosure F E` of `F` in `E`, which is not defined in - this file yet. Later we will show that (`Field.finSepDegree_eq`), if `Field.Emb F E` is finite, - then these two definitions coincide. + the separable closure of `F` in `E`, which is not defined in this file yet. Later we + will show that (`Field.finSepDegree_eq`), if `Field.Emb F E` is finite, then these two + definitions coincide. If `E / F` is algebraic with infinite separable degree, we have + `#(Field.Emb F E) = 2 ^ Field.sepDegree F E` instead. + (See `Field.Emb.cardinal_eq_two_pow_sepDegree` in another file.) For example, if + $F = \mathbb{Q}$ and $E = \mathbb{Q}( \mu_{p^\infty} )$, then $\operatorname{Emb}_F (E)$ + is in bijection with $\operatorname{Gal}(E/F)$, which is isomorphic to + $\mathbb{Z}_p^\times$, which is uncountable, whereas $ [E:F] $ is countable. - `Polynomial.natSepDegree`: the separable degree of a polynomial is a natural number, defined to be the number of distinct roots of it over its splitting field. diff --git a/Mathlib/LinearAlgebra/FreeAlgebra.lean b/Mathlib/LinearAlgebra/FreeAlgebra.lean index 02bbf0606d764..727a91d91de6a 100644 --- a/Mathlib/LinearAlgebra/FreeAlgebra.lean +++ b/Mathlib/LinearAlgebra/FreeAlgebra.lean @@ -45,3 +45,15 @@ theorem rank_eq [CommRing R] [Nontrivial R] : Cardinal.lift_umax.{v, u}, FreeMonoid] end FreeAlgebra + +open Cardinal + +theorem Algebra.rank_adjoin_le {R : Type u} {S : Type v} [CommRing R] [Ring S] [Algebra R S] + (s : Set S) : Module.rank R (adjoin R s) ≤ max #s ℵ₀ := by + rw [adjoin_eq_range_freeAlgebra_lift] + cases subsingleton_or_nontrivial R + · rw [rank_subsingleton]; exact one_le_aleph0.trans (le_max_right _ _) + rw [← lift_le.{max u v}] + refine (lift_rank_range_le (FreeAlgebra.lift R ((↑) : s → S)).toLinearMap).trans ?_ + rw [FreeAlgebra.rank_eq, lift_id'.{v,u}, lift_umax.{v,u}, lift_le, max_comm] + exact mk_list_le_max _ diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index 55b6157d4a5db..ac3781d7a4350 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -992,7 +992,7 @@ namespace WithTop section Succ -variable [DecidableEq α] [PartialOrder α] [SuccOrder α] +variable [PartialOrder α] [SuccOrder α] [∀ a : α, Decidable (succ a = a)] instance : SuccOrder (WithTop α) where succ a := @@ -1144,7 +1144,7 @@ end Succ section Pred -variable [DecidableEq α] [PartialOrder α] [PredOrder α] +variable [PartialOrder α] [PredOrder α] [∀ a : α, Decidable (pred a = a)] instance : PredOrder (WithBot α) where pred a := From 9eff8e5e952ea878b6214fe8dbfc2e378534ad85 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 20 Nov 2024 14:06:59 +0100 Subject: [PATCH 62/90] change --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 118d9e578af27..80611d83b2915 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-20 +leanprover/lean4:nightly-2024-11-19 From a8bf48322eec0dbe45a14e6b21f78a82432d10cb Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 20 Nov 2024 14:07:10 +0100 Subject: [PATCH 63/90] revert --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 80611d83b2915..118d9e578af27 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-19 +leanprover/lean4:nightly-2024-11-20 From 159918f85ffa290bd7f3ad8fc1b72a3d41e2d566 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 20 Nov 2024 13:18:31 +0000 Subject: [PATCH 64/90] feat: the monoidal category structure on homological complexes (#15695) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Let `c : ComplexShape I` with `I` an additive monoid. If `c` is equipped with the data and axioms `c.TensorSigns`, then the category `HomologicalComplex C c` can be eqquiped with a monoidal category structure if `C` is a monoidal category such that `C` has certain coproducts and both left/right tensoring commute with these. In particular, we obtain a monoidal category structure on `ChainComplex C ℕ` when `C` is an additive monoidal category. Co-authored-by: Kim Morrison Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Homology/Monoidal.lean | 325 ++++++++++++++++++ .../CategoryTheory/GradedObject/Monoidal.lean | 19 + 3 files changed, 345 insertions(+) create mode 100644 Mathlib/Algebra/Homology/Monoidal.lean diff --git a/Mathlib.lean b/Mathlib.lean index db0e4160d4515..2c342449a9a42 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -419,6 +419,7 @@ import Mathlib.Algebra.Homology.ImageToKernel import Mathlib.Algebra.Homology.Linear import Mathlib.Algebra.Homology.LocalCohomology import Mathlib.Algebra.Homology.Localization +import Mathlib.Algebra.Homology.Monoidal import Mathlib.Algebra.Homology.Opposite import Mathlib.Algebra.Homology.QuasiIso import Mathlib.Algebra.Homology.Refinements diff --git a/Mathlib/Algebra/Homology/Monoidal.lean b/Mathlib/Algebra/Homology/Monoidal.lean new file mode 100644 index 0000000000000..e0e185b0f0bec --- /dev/null +++ b/Mathlib/Algebra/Homology/Monoidal.lean @@ -0,0 +1,325 @@ +/- +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, Kim Morrison +-/ +import Mathlib.Algebra.Homology.BifunctorAssociator +import Mathlib.Algebra.Homology.Single +import Mathlib.CategoryTheory.GradedObject.Monoidal +import Mathlib.CategoryTheory.Monoidal.Transport + +/-! +# The monoidal category structure on homological complexes + +Let `c : ComplexShape I` with `I` an additive monoid. If `c` is equipped +with the data and axioms `c.TensorSigns`, then the category +`HomologicalComplex C c` can be equipped with a monoidal category +structure if `C` is a monoidal category such that `C` has certain +coproducts and both left/right tensoring commute with these. + +In particular, we obtain a monoidal category structure on +`ChainComplex C ℕ` when `C` is an additive monoidal category. + +-/ + +open CategoryTheory Limits MonoidalCategory Category + +namespace HomologicalComplex + +variable {C : Type*} [Category C] [MonoidalCategory C] [Preadditive C] [HasZeroObject C] + [(curriedTensor C).Additive] [∀ (X₁ : C), ((curriedTensor C).obj X₁).Additive] + {I : Type*} [AddMonoid I] {c : ComplexShape I} [c.TensorSigns] + +/-- If `K₁` and `K₂` are two homological complexes, this is the property that +for all `j`, the coproduct of `K₁ i₁ ⊗ K₂ i₂` for `i₁ + i₂ = j` exists. -/ +abbrev HasTensor (K₁ K₂ : HomologicalComplex C c) := HasMapBifunctor K₁ K₂ (curriedTensor C) c + +section + +variable [DecidableEq I] + +/-- The tensor product of two homological complexes. -/ +noncomputable abbrev tensorObj (K₁ K₂ : HomologicalComplex C c) [HasTensor K₁ K₂] : + HomologicalComplex C c := + mapBifunctor K₁ K₂ (curriedTensor C) c + +/-- The inclusion `K₁.X i₁ ⊗ K₂.X i₂ ⟶ (tensorObj K₁ K₂).X j` of a summand in +the tensor product of the homological complexes. -/ +noncomputable abbrev ιTensorObj (K₁ K₂ : HomologicalComplex C c) [HasTensor K₁ K₂] + (i₁ i₂ j : I) (h : i₁ + i₂ = j) : + K₁.X i₁ ⊗ K₂.X i₂ ⟶ (tensorObj K₁ K₂).X j := + ιMapBifunctor K₁ K₂ (curriedTensor C) c i₁ i₂ j h + +/-- The tensor product of two morphisms of homological complexes. -/ +noncomputable abbrev tensorHom {K₁ K₂ L₁ L₂ : HomologicalComplex C c} + (f : K₁ ⟶ L₁) (g : K₂ ⟶ L₂) [HasTensor K₁ K₂] [HasTensor L₁ L₂] : + tensorObj K₁ K₂ ⟶ tensorObj L₁ L₂ := + mapBifunctorMap f g _ _ + +/-- Given three homological complexes `K₁`, `K₂`, and `K₃`, this asserts that for +all `j`, the functor `- ⊗ K₃.X i₃` commutes with the coproduct of +the `K₁.X i₁ ⊗ K₂.X i₂` such that `i₁ + i₂ = j`. -/ +abbrev HasGoodTensor₁₂ (K₁ K₂ K₃ : HomologicalComplex C c) := + HasGoodTrifunctor₁₂Obj (curriedTensor C) (curriedTensor C) K₁ K₂ K₃ c c + +/-- Given three homological complexes `K₁`, `K₂`, and `K₃`, this asserts that for +all `j`, the functor `K₁.X i₁` commutes with the coproduct of +the `K₂.X i₂ ⊗ K₃.X i₃` such that `i₂ + i₃ = j`. -/ +abbrev HasGoodTensor₂₃ (K₁ K₂ K₃ : HomologicalComplex C c) := + HasGoodTrifunctor₂₃Obj (curriedTensor C) (curriedTensor C) K₁ K₂ K₃ c c c + +/-- The associator isomorphism for the tensor product of homological complexes. -/ +noncomputable abbrev associator (K₁ K₂ K₃ : HomologicalComplex C c) + [HasTensor K₁ K₂] [HasTensor K₂ K₃] + [HasTensor (tensorObj K₁ K₂) K₃] [HasTensor K₁ (tensorObj K₂ K₃)] + [HasGoodTensor₁₂ K₁ K₂ K₃] [HasGoodTensor₂₃ K₁ K₂ K₃] : + tensorObj (tensorObj K₁ K₂) K₃ ≅ tensorObj K₁ (tensorObj K₂ K₃) := + mapBifunctorAssociator (curriedAssociatorNatIso C) K₁ K₂ K₃ c c c + +variable (C c) in +/-- The unit of the tensor product of homological complexes. -/ +noncomputable abbrev tensorUnit : HomologicalComplex C c := (single C c 0).obj (𝟙_ C) + +variable (C c) in +/-- As a graded object, the single complex `(single C c 0).obj (𝟙_ C)` identifies +to the unit `(GradedObject.single₀ I).obj (𝟙_ C)` of the tensor product of graded objects. -/ +noncomputable def tensorUnitIso : + (GradedObject.single₀ I).obj (𝟙_ C) ≅ (tensorUnit C c).X := + GradedObject.isoMk _ _ (fun i ↦ + if hi : i = 0 then + (GradedObject.singleObjApplyIsoOfEq (0 : I) (𝟙_ C) i hi).trans + (singleObjXIsoOfEq c 0 (𝟙_ C) i hi).symm + else + { hom := 0 + inv := 0 + hom_inv_id := (GradedObject.isInitialSingleObjApply 0 (𝟙_ C) i hi).hom_ext _ _ + inv_hom_id := (isZero_single_obj_X c 0 (𝟙_ C) i hi).eq_of_src _ _ }) + +end + +instance (K₁ K₂ : HomologicalComplex C c) [GradedObject.HasTensor K₁.X K₂.X] : + HasTensor K₁ K₂ := by + assumption + +instance (K₁ K₂ K₃ : HomologicalComplex C c) + [GradedObject.HasGoodTensor₁₂Tensor K₁.X K₂.X K₃.X] : + HasGoodTensor₁₂ K₁ K₂ K₃ := + inferInstanceAs (GradedObject.HasGoodTensor₁₂Tensor K₁.X K₂.X K₃.X) + +instance (K₁ K₂ K₃ : HomologicalComplex C c) + [GradedObject.HasGoodTensorTensor₂₃ K₁.X K₂.X K₃.X] : + HasGoodTensor₂₃ K₁ K₂ K₃ := + inferInstanceAs (GradedObject.HasGoodTensorTensor₂₃ K₁.X K₂.X K₃.X) + +section + +variable (K : HomologicalComplex C c) [DecidableEq I] + +section + +variable [∀ X₂, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).flip.obj X₂)] + +instance : GradedObject.HasTensor (tensorUnit C c).X K.X := + GradedObject.hasTensor_of_iso (tensorUnitIso C c) (Iso.refl _) + +instance : HasTensor (tensorUnit C c) K := + inferInstanceAs (GradedObject.HasTensor (tensorUnit C c).X K.X) + +@[simp] +lemma unit_tensor_d₁ (i₁ i₂ j : I) : + mapBifunctor.d₁ (tensorUnit C c) K (curriedTensor C) c i₁ i₂ j = 0 := by + by_cases h₁ : c.Rel i₁ (c.next i₁) + · by_cases h₂ : ComplexShape.π c c c (c.next i₁, i₂) = j + · rw [mapBifunctor.d₁_eq _ _ _ _ h₁ _ _ h₂, single_obj_d, Functor.map_zero, + zero_app, zero_comp, smul_zero] + · rw [mapBifunctor.d₁_eq_zero' _ _ _ _ h₁ _ _ h₂] + · rw [mapBifunctor.d₁_eq_zero _ _ _ _ _ _ _ h₁] + +end + +section + +variable [∀ X₁, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).obj X₁)] + +instance : GradedObject.HasTensor K.X (tensorUnit C c).X := + GradedObject.hasTensor_of_iso (Iso.refl _) (tensorUnitIso C c) + +instance : HasTensor K (tensorUnit C c) := + inferInstanceAs (GradedObject.HasTensor K.X (tensorUnit C c).X) + +@[simp] +lemma tensor_unit_d₂ (i₁ i₂ j : I) : + mapBifunctor.d₂ K (tensorUnit C c) (curriedTensor C) c i₁ i₂ j = 0 := by + by_cases h₁ : c.Rel i₂ (c.next i₂) + · by_cases h₂ : ComplexShape.π c c c (i₁, c.next i₂) = j + · rw [mapBifunctor.d₂_eq _ _ _ _ _ h₁ _ h₂, single_obj_d, Functor.map_zero, + zero_comp, smul_zero] + · rw [mapBifunctor.d₂_eq_zero' _ _ _ _ _ h₁ _ h₂] + · rw [mapBifunctor.d₂_eq_zero _ _ _ _ _ _ _ h₁] + +end + +end + +section Unitor + +variable (K : HomologicalComplex C c) [DecidableEq I] + +section LeftUnitor + +variable [∀ X₂, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).flip.obj X₂)] + +/-- Auxiliary definition for `leftUnitor`. -/ +noncomputable def leftUnitor' : + (tensorObj (tensorUnit C c) K).X ≅ K.X := + GradedObject.Monoidal.tensorIso ((tensorUnitIso C c).symm) (Iso.refl _) ≪≫ + GradedObject.Monoidal.leftUnitor K.X + +lemma leftUnitor'_inv (i : I) : + (leftUnitor' K).inv i = (λ_ (K.X i)).inv ≫ ((singleObjXSelf c 0 (𝟙_ C)).inv ▷ (K.X i)) ≫ + ιTensorObj (tensorUnit C c) K 0 i i (zero_add i) := by + dsimp [leftUnitor'] + rw [GradedObject.Monoidal.leftUnitor_inv_apply, assoc, assoc, Iso.cancel_iso_inv_left, + GradedObject.Monoidal.ι_tensorHom] + dsimp + rw [tensorHom_id, ← comp_whiskerRight_assoc] + congr 2 + rw [← cancel_epi (GradedObject.Monoidal.tensorUnit₀ (I := I)).hom, Iso.hom_inv_id_assoc] + dsimp [tensorUnitIso] + rw [dif_pos rfl] + rfl + +@[reassoc] +lemma leftUnitor'_inv_comm (i j : I) : + (leftUnitor' K).inv i ≫ (tensorObj (tensorUnit C c) K).d i j = + K.d i j ≫ (leftUnitor' K).inv j := by + by_cases hij : c.Rel i j + · simp only [leftUnitor'_inv, assoc, mapBifunctor.d_eq, + Preadditive.comp_add, mapBifunctor.ι_D₁, mapBifunctor.ι_D₂, + unit_tensor_d₁, comp_zero, zero_add] + rw [mapBifunctor.d₂_eq _ _ _ _ _ hij _ (by simp)] + dsimp + simp only [ComplexShape.ε_zero, one_smul, ← whisker_exchange_assoc, + id_whiskerLeft, assoc, Iso.inv_hom_id_assoc] + · simp only [shape _ _ _ hij, comp_zero, zero_comp] + +/-- The left unitor for the tensor product of homological complexes. -/ +noncomputable def leftUnitor : + tensorObj (tensorUnit C c) K ≅ K := + Iso.symm (Hom.isoOfComponents (fun i ↦ (GradedObject.eval i).mapIso (leftUnitor' K).symm) + (fun _ _ _ ↦ leftUnitor'_inv_comm _ _ _)) + +end LeftUnitor + +section RightUnitor + +variable [∀ X₁, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).obj X₁)] + +/-- Auxiliary definition for `rightUnitor`. -/ +noncomputable def rightUnitor' : + (tensorObj K (tensorUnit C c)).X ≅ K.X := + GradedObject.Monoidal.tensorIso (Iso.refl _) ((tensorUnitIso C c).symm) ≪≫ + GradedObject.Monoidal.rightUnitor K.X + +lemma rightUnitor'_inv (i : I) : + (rightUnitor' K).inv i = (ρ_ (K.X i)).inv ≫ ((K.X i) ◁ (singleObjXSelf c 0 (𝟙_ C)).inv) ≫ + ιTensorObj K (tensorUnit C c) i 0 i (add_zero i) := by + dsimp [rightUnitor'] + rw [GradedObject.Monoidal.rightUnitor_inv_apply, assoc, assoc, Iso.cancel_iso_inv_left, + GradedObject.Monoidal.ι_tensorHom] + dsimp + rw [id_tensorHom, ← MonoidalCategory.whiskerLeft_comp_assoc] + congr 2 + rw [← cancel_epi (GradedObject.Monoidal.tensorUnit₀ (I := I)).hom, Iso.hom_inv_id_assoc] + dsimp [tensorUnitIso] + rw [dif_pos rfl] + rfl + +lemma rightUnitor'_inv_comm (i j : I) : + (rightUnitor' K).inv i ≫ (tensorObj K (tensorUnit C c)).d i j = + K.d i j ≫ (rightUnitor' K).inv j := by + by_cases hij : c.Rel i j + · simp only [rightUnitor'_inv, assoc, mapBifunctor.d_eq, + Preadditive.comp_add, mapBifunctor.ι_D₁, mapBifunctor.ι_D₂, + tensor_unit_d₂, comp_zero, add_zero] + rw [mapBifunctor.d₁_eq _ _ _ _ hij _ _ (by simp)] + dsimp + simp only [one_smul, whisker_exchange_assoc, + MonoidalCategory.whiskerRight_id, assoc, Iso.inv_hom_id_assoc] + · simp only [shape _ _ _ hij, comp_zero, zero_comp] + +/-- The right unitor for the tensor product of homological complexes. -/ +noncomputable def rightUnitor : + tensorObj K (tensorUnit C c) ≅ K := + Iso.symm (Hom.isoOfComponents (fun i ↦ (GradedObject.eval i).mapIso (rightUnitor' K).symm) + (fun _ _ _ ↦ rightUnitor'_inv_comm _ _ _)) + +end RightUnitor + +end Unitor + +variable (C c) [∀ (X₁ X₂ : GradedObject I C), GradedObject.HasTensor X₁ X₂] + [∀ X₁, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).obj X₁)] + [∀ X₂, PreservesColimit (Functor.empty.{0} C) ((curriedTensor C).flip.obj X₂)] + [∀ (X₁ X₂ X₃ X₄ : GradedObject I C), GradedObject.HasTensor₄ObjExt X₁ X₂ X₃ X₄] + [∀ (X₁ X₂ X₃ : GradedObject I C), GradedObject.HasGoodTensor₁₂Tensor X₁ X₂ X₃] + [∀ (X₁ X₂ X₃ : GradedObject I C), GradedObject.HasGoodTensorTensor₂₃ X₁ X₂ X₃] + [DecidableEq I] + +noncomputable instance monoidalCategoryStruct : + MonoidalCategoryStruct (HomologicalComplex C c) where + tensorObj K₁ K₂ := tensorObj K₁ K₂ + whiskerLeft _ _ _ g := tensorHom (𝟙 _) g + whiskerRight f _ := tensorHom f (𝟙 _) + tensorHom f g := tensorHom f g + tensorUnit := tensorUnit C c + associator K₁ K₂ K₃ := associator K₁ K₂ K₃ + leftUnitor K := leftUnitor K + rightUnitor K := rightUnitor K + +/-- The structure which allows to construct the monoidal category structure +on `HomologicalComplex C c` from the monoidal category structure on +graded objects. -/ +noncomputable def Monoidal.inducingFunctorData : + Monoidal.InducingFunctorData (forget C c) where + μIso _ _ := Iso.refl _ + εIso := tensorUnitIso C c + whiskerLeft_eq K₁ K₂ L₂ g := by + dsimp [forget] + erw [comp_id, id_comp] + rfl + whiskerRight_eq {K₁ L₁} f K₂ := by + dsimp [forget] + erw [comp_id, id_comp] + rfl + tensorHom_eq {K₁ L₁ K₂ L₂} f g := by + dsimp [forget] + erw [comp_id, id_comp] + rfl + associator_eq K₁ K₂ K₃ := by + dsimp [forget] + simp only [tensorHom_id, whiskerRight_tensor, id_whiskerRight, + id_comp, Iso.inv_hom_id, comp_id, assoc] + erw [id_whiskerRight, id_comp, id_comp] + rfl + leftUnitor_eq K := by + dsimp + erw [id_comp] + rfl + rightUnitor_eq K := by + dsimp + rw [assoc] + erw [id_comp] + rfl + +noncomputable instance monoidalCategory : MonoidalCategory (HomologicalComplex C c) := + Monoidal.induced _ (Monoidal.inducingFunctorData C c) + +noncomputable example {D : Type*} [Category D] [Preadditive D] [MonoidalCategory D] + [HasZeroObject D] [HasFiniteCoproducts D] [((curriedTensor D).Additive)] + [∀ (X : D), (((curriedTensor D).obj X).Additive)] + [∀ (X : D), PreservesFiniteCoproducts ((curriedTensor D).obj X)] + [∀ (X : D), PreservesFiniteCoproducts ((curriedTensor D).flip.obj X)] : + MonoidalCategory (ChainComplex D ℕ) := inferInstance + +end HomologicalComplex diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index 34da9a269a20b..b10fab11e47d8 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -38,6 +38,14 @@ the coproduct of the objects `X₁ i ⊗ X₂ j` for `i + j = n` exists. -/ abbrev HasTensor (X₁ X₂ : GradedObject I C) : Prop := HasMap (((mapBifunctor (curriedTensor C) I I).obj X₁).obj X₂) (fun ⟨i, j⟩ => i + j) +lemma hasTensor_of_iso {X₁ X₂ Y₁ Y₂ : GradedObject I C} + (e₁ : X₁ ≅ Y₁) (e₂ : X₂ ≅ Y₂) [HasTensor X₁ X₂] : + HasTensor Y₁ Y₂ := by + let e : ((mapBifunctor (curriedTensor C) I I).obj X₁).obj X₂ ≅ + ((mapBifunctor (curriedTensor C) I I).obj Y₁).obj Y₂ := isoMk _ _ + (fun ⟨i, j⟩ ↦ (eval i).mapIso e₁ ⊗ (eval j).mapIso e₂) + exact hasMap_of_iso e _ + namespace Monoidal /-- The tensor product of two graded objects. -/ @@ -120,6 +128,17 @@ lemma tensor_comp {X₁ X₂ X₃ Y₁ Y₂ Y₃ : GradedObject I C} (f₁ : X apply congr_mapMap simp +/-- The isomorphism `tensorObj X₁ Y₁ ≅ tensorObj X₂ Y₂` induced by isomorphisms of graded +objects `e : X₁ ≅ X₂` and `e' : Y₁ ≅ Y₂`. -/ +@[simps] +noncomputable def tensorIso {X₁ X₂ Y₁ Y₂ : GradedObject I C} (e : X₁ ≅ X₂) (e' : Y₁ ≅ Y₂) + [HasTensor X₁ Y₁] [HasTensor X₂ Y₂] : + tensorObj X₁ Y₁ ≅ tensorObj X₂ Y₂ where + hom := tensorHom e.hom e'.hom + inv := tensorHom e.inv e'.inv + hom_inv_id := by simp only [← tensor_comp, Iso.hom_inv_id, tensor_id] + inv_hom_id := by simp only [← tensor_comp, Iso.inv_hom_id, tensor_id] + lemma tensorHom_def {X₁ X₂ Y₁ Y₂ : GradedObject I C} (f : X₁ ⟶ X₂) (g : Y₁ ⟶ Y₂) [HasTensor X₁ Y₁] [HasTensor X₂ Y₂] [HasTensor X₂ Y₁] : tensorHom f g = whiskerRight f Y₁ ≫ whiskerLeft X₂ g := by From 40b557834ca80e53ba1aed3870919826e1987375 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 20 Nov 2024 13:18:32 +0000 Subject: [PATCH 65/90] feat: expand API around manifold derivatives (#18850) Notably, prove that basic functions are differentiable in the manifold sense, and compute their derivatives. --- Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 23 +++ Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 92 +++++++++ .../Manifold/MFDeriv/SpecificFunctions.lean | 193 +++++++++++++++++- 3 files changed, 303 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 9202595ec5d94..5439b39e461ed 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri -/ import Mathlib.Geometry.Manifold.ContMDiffMap +import Mathlib.Geometry.Manifold.MFDeriv.Basic /-! # Smooth monoid @@ -132,6 +133,28 @@ theorem smooth_mul_left {a : G} : Smooth I I fun b : G => a * b := theorem smooth_mul_right {a : G} : Smooth I I fun b : G => b * a := smooth_id.mul smooth_const +theorem contMDiff_mul_left {a : G} : ContMDiff I I n (a * ·) := smooth_mul_left.contMDiff + +theorem contMDiffAt_mul_left {a b : G} : ContMDiffAt I I n (a * ·) b := + contMDiff_mul_left.contMDiffAt + +theorem mdifferentiable_mul_left {a : G} : MDifferentiable I I (a * ·) := + contMDiff_mul_left.mdifferentiable le_rfl + +theorem mdifferentiableAt_mul_left {a b : G} : MDifferentiableAt I I (a * ·) b := + contMDiffAt_mul_left.mdifferentiableAt le_rfl + +theorem contMDiff_mul_right {a : G} : ContMDiff I I n (· * a) := smooth_mul_right.contMDiff + +theorem contMDiffAt_mul_right {a b : G} : ContMDiffAt I I n (· * a) b := + contMDiff_mul_right.contMDiffAt + +theorem mdifferentiable_mul_right {a : G} : MDifferentiable I I (· * a) := + contMDiff_mul_right.mdifferentiable le_rfl + +theorem mdifferentiableAt_mul_right {a b : G} : MDifferentiableAt I I (· * a) b := + contMDiffAt_mul_right.mdifferentiableAt le_rfl + end variable (I) (g h : G) diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index eecf755b43be2..eb88f2a6c5c87 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -257,4 +257,96 @@ theorem mdifferentiableOn_extChartAt_symm : intro y hy exact (mdifferentiableWithinAt_extChartAt_symm hy).mono (extChartAt_target_subset_range x) +/-- The composition of the derivative of `extChartAt` with the derivative of the inverse of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).target`. -/ +lemma mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm {x : M} + {y : E} (hy : y ∈ (extChartAt I x).target) : + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) ((extChartAt I x).symm y)) ∘L + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) y) = ContinuousLinearMap.id _ _ := by + have U : UniqueMDiffWithinAt 𝓘(𝕜, E) (range ↑I) y := by + apply I.uniqueMDiffOn + exact extChartAt_target_subset_range x hy + have h'y : (extChartAt I x).symm y ∈ (extChartAt I x).source := (extChartAt I x).map_target hy + have h''y : (extChartAt I x).symm y ∈ (chartAt H x).source := by + rwa [← extChartAt_source (I := I)] + rw [← mfderiv_comp_mfderivWithin]; rotate_left + · apply mdifferentiableAt_extChartAt h''y + · exact mdifferentiableWithinAt_extChartAt_symm hy + · exact U + rw [← mfderivWithin_id U] + apply Filter.EventuallyEq.mfderivWithin_eq U + · filter_upwards [extChartAt_target_mem_nhdsWithin_of_mem hy] with z hz + simp only [Function.comp_def, PartialEquiv.right_inv (extChartAt I x) hz, id_eq] + · simp only [Function.comp_def, PartialEquiv.right_inv (extChartAt I x) hy, id_eq] + +/-- The composition of the derivative of `extChartAt` with the derivative of the inverse of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).source`. -/ +lemma mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm' {x : M} + {y : M} (hy : y ∈ (extChartAt I x).source) : + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) y) ∘L + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) (extChartAt I x y)) + = ContinuousLinearMap.id _ _ := by + have : y = (extChartAt I x).symm (extChartAt I x y) := ((extChartAt I x).left_inv hy).symm + convert mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm ((extChartAt I x).map_source hy) + +/-- The composition of the derivative of the inverse of `extChartAt` with the derivative of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).target`. -/ +lemma mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt + {y : E} (hy : y ∈ (extChartAt I x).target) : + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) y) ∘L + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) ((extChartAt I x).symm y)) + = ContinuousLinearMap.id _ _ := by + have h'y : (extChartAt I x).symm y ∈ (extChartAt I x).source := (extChartAt I x).map_target hy + have h''y : (extChartAt I x).symm y ∈ (chartAt H x).source := by + rwa [← extChartAt_source (I := I)] + have U' : UniqueMDiffWithinAt I (extChartAt I x).source ((extChartAt I x).symm y) := + (isOpen_extChartAt_source x).uniqueMDiffWithinAt h'y + have : mfderiv I 𝓘(𝕜, E) (extChartAt I x) ((extChartAt I x).symm y) + = mfderivWithin I 𝓘(𝕜, E) (extChartAt I x) (extChartAt I x).source + ((extChartAt I x).symm y) := by + rw [mfderivWithin_eq_mfderiv U'] + exact mdifferentiableAt_extChartAt h''y + rw [this, ← mfderivWithin_comp_of_eq]; rotate_left + · exact mdifferentiableWithinAt_extChartAt_symm hy + · exact (mdifferentiableAt_extChartAt h''y).mdifferentiableWithinAt + · intro z hz + apply extChartAt_target_subset_range x + exact PartialEquiv.map_source (extChartAt I x) hz + · exact U' + · exact PartialEquiv.right_inv (extChartAt I x) hy + rw [← mfderivWithin_id U'] + apply Filter.EventuallyEq.mfderivWithin_eq U' + · filter_upwards [extChartAt_source_mem_nhdsWithin' h'y] with z hz + simp only [Function.comp_def, PartialEquiv.left_inv (extChartAt I x) hz, id_eq] + · simp only [Function.comp_def, PartialEquiv.right_inv (extChartAt I x) hy, id_eq] + +/-- The composition of the derivative of the inverse of `extChartAt` with the derivative of +`extChartAt` gives the identity. +Version where the basepoint belongs to `(extChartAt I x).source`. -/ +lemma mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt' + {y : M} (hy : y ∈ (extChartAt I x).source) : + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) (extChartAt I x y)) ∘L + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) y) + = ContinuousLinearMap.id _ _ := by + have : y = (extChartAt I x).symm (extChartAt I x y) := ((extChartAt I x).left_inv hy).symm + convert mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt ((extChartAt I x).map_source hy) + +lemma isInvertible_mfderivWithin_extChartAt_symm {y : E} (hy : y ∈ (extChartAt I x).target) : + (mfderivWithin 𝓘(𝕜, E) I (extChartAt I x).symm (range I) y).IsInvertible := + ContinuousLinearMap.IsInvertible.of_inverse + (mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt hy) + (mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm hy) + +lemma isInvertible_mfderiv_extChartAt {y : M} (hy : y ∈ (extChartAt I x).source) : + (mfderiv I 𝓘(𝕜, E) (extChartAt I x) y).IsInvertible := by + have h'y : extChartAt I x y ∈ (extChartAt I x).target := (extChartAt I x).map_source hy + have Z := ContinuousLinearMap.IsInvertible.of_inverse + (mfderiv_extChartAt_comp_mfderivWithin_extChartAt_symm h'y) + (mfderivWithin_extChartAt_symm_comp_mfderiv_extChartAt h'y) + have : (extChartAt I x).symm ((extChartAt I x) y) = y := (extChartAt I x).left_inv hy + rwa [this] at Z + end extChartAt diff --git a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean index 2bbe312b6c383..63e5592d2c537 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/SpecificFunctions.lean @@ -26,14 +26,31 @@ section SpecificFunctions /-! ### Differentiability of specific functions -/ -variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + -- declare a charted space `M` over the pair `(E, H)`. + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} - [TopologicalSpace M] [ChartedSpace H M] {E' : Type*} - [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] + [TopologicalSpace M] [ChartedSpace H M] + -- declare a charted space `M'` over the pair `(E', H')`. + {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] + -- declare a charted space `M''` over the pair `(E'', H'')`. {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + -- declare a charted space `N` over the pair `(F, G)`. + {F : Type*} + [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] + {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] + -- declare a charted space `N'` over the pair `(F', G')`. + {F' : Type*} + [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] + {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] + -- F₁, F₂, F₃, F₄ are normed spaces + {F₁ : Type*} + [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] + [NormedSpace 𝕜 F₂] {F₃ : Type*} [NormedAddCommGroup F₃] [NormedSpace 𝕜 F₃] {F₄ : Type*} + [NormedAddCommGroup F₄] [NormedSpace 𝕜 F₄] namespace ContinuousLinearMap @@ -301,6 +318,128 @@ theorem mfderivWithin_snd {s : Set (M × M')} {x : M × M'} ContinuousLinearMap.snd 𝕜 (TangentSpace I x.1) (TangentSpace I' x.2) := by rw [MDifferentiable.mfderivWithin mdifferentiableAt_snd hxs]; exact mfderiv_snd +theorem MDifferentiableWithinAt.fst {f : N → M × M'} {s : Set N} {x : N} + (hf : MDifferentiableWithinAt J (I.prod I') f s x) : + MDifferentiableWithinAt J I (fun x => (f x).1) s x := + mdifferentiableAt_fst.comp_mdifferentiableWithinAt x hf + +theorem MDifferentiableAt.fst {f : N → M × M'} {x : N} (hf : MDifferentiableAt J (I.prod I') f x) : + MDifferentiableAt J I (fun x => (f x).1) x := + mdifferentiableAt_fst.comp x hf + +theorem MDifferentiable.fst {f : N → M × M'} (hf : MDifferentiable J (I.prod I') f) : + MDifferentiable J I fun x => (f x).1 := + mdifferentiable_fst.comp hf + +theorem MDifferentiableWithinAt.snd {f : N → M × M'} {s : Set N} {x : N} + (hf : MDifferentiableWithinAt J (I.prod I') f s x) : + MDifferentiableWithinAt J I' (fun x => (f x).2) s x := + mdifferentiableAt_snd.comp_mdifferentiableWithinAt x hf + +theorem MDifferentiableAt.snd {f : N → M × M'} {x : N} (hf : MDifferentiableAt J (I.prod I') f x) : + MDifferentiableAt J I' (fun x => (f x).2) x := + mdifferentiableAt_snd.comp x hf + +theorem MDifferentiable.snd {f : N → M × M'} (hf : MDifferentiable J (I.prod I') f) : + MDifferentiable J I' fun x => (f x).2 := + mdifferentiable_snd.comp hf + +theorem mdifferentiableWithinAt_prod_iff (f : M → M' × N') : + MDifferentiableWithinAt I (I'.prod J') f s x ↔ + MDifferentiableWithinAt I I' (Prod.fst ∘ f) s x + ∧ MDifferentiableWithinAt I J' (Prod.snd ∘ f) s x := + ⟨fun h => ⟨h.fst, h.snd⟩, fun h => h.1.prod_mk h.2⟩ + +theorem mdifferentiableWithinAt_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiableWithinAt I 𝓘(𝕜, F₁ × F₂) f s x ↔ + MDifferentiableWithinAt I 𝓘(𝕜, F₁) (Prod.fst ∘ f) s x ∧ + MDifferentiableWithinAt I 𝓘(𝕜, F₂) (Prod.snd ∘ f) s x := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiableWithinAt_prod_iff f + +theorem mdifferentiableAt_prod_iff (f : M → M' × N') : + MDifferentiableAt I (I'.prod J') f x ↔ + MDifferentiableAt I I' (Prod.fst ∘ f) x ∧ MDifferentiableAt I J' (Prod.snd ∘ f) x := by + simp_rw [← mdifferentiableWithinAt_univ]; exact mdifferentiableWithinAt_prod_iff f + +theorem mdifferentiableAt_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiableAt I 𝓘(𝕜, F₁ × F₂) f x ↔ + MDifferentiableAt I 𝓘(𝕜, F₁) (Prod.fst ∘ f) x + ∧ MDifferentiableAt I 𝓘(𝕜, F₂) (Prod.snd ∘ f) x := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiableAt_prod_iff f + +theorem mdifferentiableOn_prod_iff (f : M → M' × N') : + MDifferentiableOn I (I'.prod J') f s ↔ + MDifferentiableOn I I' (Prod.fst ∘ f) s ∧ MDifferentiableOn I J' (Prod.snd ∘ f) s := + ⟨fun h ↦ ⟨fun x hx ↦ ((mdifferentiableWithinAt_prod_iff f).1 (h x hx)).1, + fun x hx ↦ ((mdifferentiableWithinAt_prod_iff f).1 (h x hx)).2⟩ , + fun h x hx ↦ (mdifferentiableWithinAt_prod_iff f).2 ⟨h.1 x hx, h.2 x hx⟩⟩ + +theorem mdifferentiableOn_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiableOn I 𝓘(𝕜, F₁ × F₂) f s ↔ + MDifferentiableOn I 𝓘(𝕜, F₁) (Prod.fst ∘ f) s + ∧ MDifferentiableOn I 𝓘(𝕜, F₂) (Prod.snd ∘ f) s := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiableOn_prod_iff f + +theorem mdifferentiable_prod_iff (f : M → M' × N') : + MDifferentiable I (I'.prod J') f ↔ + MDifferentiable I I' (Prod.fst ∘ f) ∧ MDifferentiable I J' (Prod.snd ∘ f) := + ⟨fun h => ⟨h.fst, h.snd⟩, fun h => by convert h.1.prod_mk h.2⟩ + +theorem mdifferentiable_prod_module_iff (f : M → F₁ × F₂) : + MDifferentiable I 𝓘(𝕜, F₁ × F₂) f ↔ + MDifferentiable I 𝓘(𝕜, F₁) (Prod.fst ∘ f) ∧ MDifferentiable I 𝓘(𝕜, F₂) (Prod.snd ∘ f) := by + rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] + exact mdifferentiable_prod_iff f + + +section prodMap + +variable {f : M → M'} {g : N → N'} {r : Set N} {y : N} + +/-- The product map of two `C^n` functions within a set at a point is `C^n` +within the product set at the product point. -/ +theorem MDifferentiableWithinAt.prod_map' {p : M × N} (hf : MDifferentiableWithinAt I I' f s p.1) + (hg : MDifferentiableWithinAt J J' g r p.2) : + MDifferentiableWithinAt (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) p := + (hf.comp p mdifferentiableWithinAt_fst (prod_subset_preimage_fst _ _)).prod_mk <| + hg.comp p mdifferentiableWithinAt_snd (prod_subset_preimage_snd _ _) + +theorem MDifferentiableWithinAt.prod_map (hf : MDifferentiableWithinAt I I' f s x) + (hg : MDifferentiableWithinAt J J' g r y) : + MDifferentiableWithinAt (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) (x, y) := + MDifferentiableWithinAt.prod_map' hf hg + +theorem MDifferentiableAt.prod_map + (hf : MDifferentiableAt I I' f x) (hg : MDifferentiableAt J J' g y) : + MDifferentiableAt (I.prod J) (I'.prod J') (Prod.map f g) (x, y) := by + rw [← mdifferentiableWithinAt_univ] at * + convert hf.prod_map hg + exact univ_prod_univ.symm + +/-- Variant of `MDifferentiableAt.prod_map` in which the point in the product is given as `p` +instead of a pair `(x, y)`. -/ +theorem MDifferentiableAt.prod_map' {p : M × N} + (hf : MDifferentiableAt I I' f p.1) (hg : MDifferentiableAt J J' g p.2) : + MDifferentiableAt (I.prod J) (I'.prod J') (Prod.map f g) p := by + rcases p with ⟨⟩ + exact hf.prod_map hg + +theorem MDifferentiableOn.prod_map + (hf : MDifferentiableOn I I' f s) (hg : MDifferentiableOn J J' g r) : + MDifferentiableOn (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) := + (hf.comp mdifferentiableOn_fst (prod_subset_preimage_fst _ _)).prod_mk <| + hg.comp mdifferentiableOn_snd (prod_subset_preimage_snd _ _) + +theorem MDifferentiable.prod_map (hf : MDifferentiable I I' f) (hg : MDifferentiable J J' g) : + MDifferentiable (I.prod J) (I'.prod J') (Prod.map f g) := by + intro p + exact (hf p.1).prod_map' (hg p.2) + +end prodMap + @[simp, mfld_simps] theorem tangentMap_prod_snd {p : TangentBundle (I.prod I') (M × M')} : tangentMap (I.prod I') I' Prod.snd p = ⟨p.proj.2, p.2.2⟩ := by @@ -330,22 +469,30 @@ theorem mfderiv_prod_left {x₀ : M} {y₀ : M'} : refine (mdifferentiableAt_id.mfderiv_prod mdifferentiableAt_const).trans ?_ rw [mfderiv_id, mfderiv_const, ContinuousLinearMap.inl] +theorem tangentMap_prod_left {p : TangentBundle I M} {y₀ : M'} : + tangentMap I (I.prod I') (fun x => (x, y₀)) p = ⟨(p.1, y₀), (p.2, 0)⟩ := by + simp only [tangentMap, mfderiv_prod_left, TotalSpace.mk_inj] + rfl + theorem mfderiv_prod_right {x₀ : M} {y₀ : M'} : mfderiv I' (I.prod I') (fun y => (x₀, y)) y₀ = ContinuousLinearMap.inr 𝕜 (TangentSpace I x₀) (TangentSpace I' y₀) := by refine (mdifferentiableAt_const.mfderiv_prod mdifferentiableAt_id).trans ?_ rw [mfderiv_id, mfderiv_const, ContinuousLinearMap.inr] +theorem tangentMap_prod_right {p : TangentBundle I' M'} {x₀ : M} : + tangentMap I' (I.prod I') (fun y => (x₀, y)) p = ⟨(x₀, p.1), (0, p.2)⟩ := by + simp only [tangentMap, mfderiv_prod_right, TotalSpace.mk_inj] + rfl + /-- The total derivative of a function in two variables is the sum of the partial derivatives. Note that to state this (without casts) we need to be able to see through the definition of `TangentSpace`. -/ theorem mfderiv_prod_eq_add {f : M × M' → M''} {p : M × M'} (hf : MDifferentiableAt (I.prod I') I'' f p) : mfderiv (I.prod I') I'' f p = - show E × E' →L[𝕜] E'' from mfderiv (I.prod I') I'' (fun z : M × M' => f (z.1, p.2)) p + mfderiv (I.prod I') I'' (fun z : M × M' => f (p.1, z.2)) p := by - dsimp only erw [mfderiv_comp_of_eq hf (mdifferentiableAt_fst.prod_mk mdifferentiableAt_const) rfl, mfderiv_comp_of_eq hf (mdifferentiableAt_const.prod_mk mdifferentiableAt_snd) rfl, ← ContinuousLinearMap.comp_add, @@ -356,6 +503,42 @@ theorem mfderiv_prod_eq_add {f : M × M' → M''} {p : M × M'} convert ContinuousLinearMap.comp_id <| mfderiv (.prod I I') I'' f (p.1, p.2) exact ContinuousLinearMap.coprod_inl_inr +/-- The total derivative of a function in two variables is the sum of the partial derivatives. + Note that to state this (without casts) we need to be able to see through the definition of + `TangentSpace`. Version in terms of the one-variable derivatives. -/ +theorem mfderiv_prod_eq_add_comp {f : M × M' → M''} {p : M × M'} + (hf : MDifferentiableAt (I.prod I') I'' f p) : + mfderiv (I.prod I') I'' f p = + (mfderiv I I'' (fun z : M => f (z, p.2)) p.1) ∘L (id (ContinuousLinearMap.fst 𝕜 E E') : + (TangentSpace (I.prod I') p) →L[𝕜] (TangentSpace I p.1)) + + (mfderiv I' I'' (fun z : M' => f (p.1, z)) p.2) ∘L (id (ContinuousLinearMap.snd 𝕜 E E') : + (TangentSpace (I.prod I') p) →L[𝕜] (TangentSpace I' p.2)) := by + rw [mfderiv_prod_eq_add hf] + congr + · have : (fun z : M × M' => f (z.1, p.2)) = (fun z : M => f (z, p.2)) ∘ Prod.fst := rfl + rw [this, mfderiv_comp (I' := I)] + · simp only [mfderiv_fst, id_eq] + rfl + · exact hf.comp _ (mdifferentiableAt_id.prod_mk mdifferentiableAt_const) + · exact mdifferentiableAt_fst + · have : (fun z : M × M' => f (p.1, z.2)) = (fun z : M' => f (p.1, z)) ∘ Prod.snd := rfl + rw [this, mfderiv_comp (I' := I')] + · simp only [mfderiv_snd, id_eq] + rfl + · exact hf.comp _ (mdifferentiableAt_const.prod_mk mdifferentiableAt_id) + · exact mdifferentiableAt_snd + +/-- The total derivative of a function in two variables is the sum of the partial derivatives. + Note that to state this (without casts) we need to be able to see through the definition of + `TangentSpace`. Version in terms of the one-variable derivatives. -/ +theorem mfderiv_prod_eq_add_apply {f : M × M' → M''} {p : M × M'} {v : TangentSpace (I.prod I') p} + (hf : MDifferentiableAt (I.prod I') I'' f p) : + mfderiv (I.prod I') I'' f p v = + mfderiv I I'' (fun z : M => f (z, p.2)) p.1 v.1 + + mfderiv I' I'' (fun z : M' => f (p.1, z)) p.2 v.2 := by + rw [mfderiv_prod_eq_add_comp hf] + rfl + end Prod section Arithmetic From fc5f4490ca3100759e2ad673a8e67ea5ad23b3f5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 20 Nov 2024 13:18:33 +0000 Subject: [PATCH 66/90] ci: print the existing update PR, if found (#19218) This reveals that the problem was #18946 being open. --- .github/workflows/update_dependencies.yml | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/.github/workflows/update_dependencies.yml b/.github/workflows/update_dependencies.yml index 48111dc1fac1d..a8ea709c6de9e 100644 --- a/.github/workflows/update_dependencies.yml +++ b/.github/workflows/update_dependencies.yml @@ -40,6 +40,13 @@ jobs: # Only return if PR is still open filterOutClosed: true + - name: Print PR, if found + run: echo "Found PR ${prNumber} at ${prUrl}" + if: steps.PR.outputs.pr_found == 'true' + env: + prNumber: ${{ steps.PR.outputs.number }} + prUrl: ${{ steps.PR.outputs.pr_url }} + - name: Configure Git User if: ${{ !contains(steps.PR.outputs.pr_labels, 'ready-to-merge') }} run: | From 025632cce0c65284113a5c478a7c6303fdb92a49 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 20 Nov 2024 13:28:38 +0000 Subject: [PATCH 67/90] =?UTF-8?q?feat:=20change=20definition=20of=20`HasFT?= =?UTF-8?q?aylorSeries`=20to=20take=20a=20parameter=20in=20`WithTop=20?= =?UTF-8?q?=E2=84=95=E2=88=9E`=20instead=20of=20`=E2=84=95=E2=88=9E`=20(#1?= =?UTF-8?q?8723)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For now this is useless (and even counter-productive, as it means we have to add some casts in some places), but it will make it possible to integrate analytic functions in the smooth hierarchy in #17152. --- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 49 +++---- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 101 +++++++------- .../Calculus/ContDiff/FTaylorSeries.lean | 128 ++++++++++++------ .../Calculus/ContDiff/FaaDiBruno.lean | 14 +- .../Calculus/ContDiff/FiniteDimension.lean | 6 +- .../Analysis/Calculus/ContDiff/RCLike.lean | 3 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 2 +- Mathlib/Analysis/Convolution.lean | 3 +- .../Fourier/FourierTransformDeriv.lean | 22 ++- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 2 +- Mathlib/Data/ENat/Basic.lean | 25 +++- 11 files changed, 216 insertions(+), 139 deletions(-) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 05d77c24f110e..52ac217ea8cf2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -187,7 +187,7 @@ theorem IsBoundedBilinearMap.contDiff (hb : IsBoundedBilinearMap 𝕜 b) : ContD /-- If `f` admits a Taylor series `p` in a set `s`, and `g` is linear, then `g ∘ f` admits a Taylor series whose `k`-th term is given by `g ∘ (p k)`. -/ -theorem HasFTaylorSeriesUpToOn.continuousLinearMap_comp (g : F →L[𝕜] G) +theorem HasFTaylorSeriesUpToOn.continuousLinearMap_comp {n : WithTop ℕ∞} (g : F →L[𝕜] G) (hf : HasFTaylorSeriesUpToOn n f p s) : HasFTaylorSeriesUpToOn n (g ∘ f) (fun x k => g.compContinuousMultilinearMap (p x k)) s where zero_eq x hx := congr_arg g (hf.zero_eq x hx) @@ -221,16 +221,16 @@ theorem ContDiff.continuousLinearMap_comp {f : E → F} (g : F →L[𝕜] G) (hf /-- The iterated derivative within a set of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_left {f : E → F} (g : F →L[𝕜] G) - (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = g.compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := (((hf.ftaylorSeriesWithin hs).continuousLinearMap_comp g).eq_iteratedFDerivWithin_of_uniqueDiffOn - hi hs hx).symm + (mod_cast hi) hs hx).symm /-- The iterated derivative of the composition with a linear map on the left is obtained by applying the linear map to the iterated derivative. -/ theorem ContinuousLinearMap.iteratedFDeriv_comp_left {f : E → F} (g : F →L[𝕜] G) - (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : i ≤ n) : iteratedFDeriv 𝕜 i (g ∘ f) x = g.compContinuousMultilinearMap (iteratedFDeriv 𝕜 i f x) := by simp only [← iteratedFDerivWithin_univ] exact g.iteratedFDerivWithin_comp_left hf.contDiffOn uniqueDiffOn_univ (mem_univ x) hi @@ -262,7 +262,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_left (g : F ≃L[𝕜] G /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative within a set. -/ theorem LinearIsometry.norm_iteratedFDerivWithin_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) - (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {i : ℕ} (hi : i ≤ n) : ‖iteratedFDerivWithin 𝕜 i (g ∘ f) s x‖ = ‖iteratedFDerivWithin 𝕜 i f s x‖ := by have : iteratedFDerivWithin 𝕜 i (g ∘ f) s x = @@ -274,7 +274,7 @@ theorem LinearIsometry.norm_iteratedFDerivWithin_comp_left {f : E → F} (g : F /-- Composition with a linear isometry on the left preserves the norm of the iterated derivative. -/ theorem LinearIsometry.norm_iteratedFDeriv_comp_left {f : E → F} (g : F →ₗᵢ[𝕜] G) - (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiff 𝕜 n f) (x : E) {i : ℕ} (hi : i ≤ n) : ‖iteratedFDeriv 𝕜 i (g ∘ f) x‖ = ‖iteratedFDeriv 𝕜 i f x‖ := by simp only [← iteratedFDerivWithin_univ] exact g.norm_iteratedFDerivWithin_comp_left hf.contDiffOn uniqueDiffOn_univ (mem_univ x) hi @@ -326,8 +326,8 @@ theorem ContinuousLinearEquiv.comp_contDiff_iff (e : F ≃L[𝕜] G) : /-- If `f` admits a Taylor series `p` in a set `s`, and `g` is linear, then `f ∘ g` admits a Taylor series in `g ⁻¹' s`, whose `k`-th term is given by `p k (g v₁, ..., g vₖ)` . -/ -theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap (hf : HasFTaylorSeriesUpToOn n f p s) - (g : G →L[𝕜] E) : +theorem HasFTaylorSeriesUpToOn.compContinuousLinearMap {n : WithTop ℕ∞} + (hf : HasFTaylorSeriesUpToOn n f p s) (g : G →L[𝕜] E) : HasFTaylorSeriesUpToOn n (f ∘ g) (fun x k => (p (g x) k).compContinuousLinearMap fun _ => g) (g ⁻¹' s) := by let A : ∀ m : ℕ, (E[×m]→L[𝕜] F) → G[×m]→L[𝕜] F := fun m h => h.compContinuousLinearMap fun _ => g @@ -372,11 +372,11 @@ theorem ContDiff.comp_continuousLinearMap {f : E → F} {g : G →L[𝕜] E} (hf obtained by composing the iterated derivative with the linear map. -/ theorem ContinuousLinearMap.iteratedFDerivWithin_comp_right {f : E → F} (g : G →L[𝕜] E) (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (h's : UniqueDiffOn 𝕜 (g ⁻¹' s)) {x : G} - (hx : g x ∈ s) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hx : g x ∈ s) {i : ℕ} (hi : i ≤ n) : iteratedFDerivWithin 𝕜 i (f ∘ g) (g ⁻¹' s) x = (iteratedFDerivWithin 𝕜 i f s (g x)).compContinuousLinearMap fun _ => g := (((hf.ftaylorSeriesWithin hs).compContinuousLinearMap g).eq_iteratedFDerivWithin_of_uniqueDiffOn - hi h's hx).symm + (mod_cast hi) h's hx).symm /-- The iterated derivative within a set of the composition with a linear equiv on the right is obtained by composing the iterated derivative with the linear equiv. -/ @@ -406,7 +406,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_right (g : G ≃L[𝕜] /-- The iterated derivative of the composition with a linear map on the right is obtained by composing the iterated derivative with the linear map. -/ theorem ContinuousLinearMap.iteratedFDeriv_comp_right (g : G →L[𝕜] E) {f : E → F} - (hf : ContDiff 𝕜 n f) (x : G) {i : ℕ} (hi : (i : ℕ∞) ≤ n) : + (hf : ContDiff 𝕜 n f) (x : G) {i : ℕ} (hi : i ≤ n) : iteratedFDeriv 𝕜 i (f ∘ g) x = (iteratedFDeriv 𝕜 i f (g x)).compContinuousLinearMap fun _ => g := by simp only [← iteratedFDerivWithin_univ] @@ -463,7 +463,8 @@ theorem ContinuousLinearEquiv.contDiff_comp_iff (e : G ≃L[𝕜] E) : /-- If two functions `f` and `g` admit Taylor series `p` and `q` in a set `s`, then the cartesian product of `f` and `g` admits the cartesian product of `p` and `q` as a Taylor series. -/ -theorem HasFTaylorSeriesUpToOn.prod (hf : HasFTaylorSeriesUpToOn n f p s) {g : E → G} +theorem HasFTaylorSeriesUpToOn.prod {n : WithTop ℕ∞} + (hf : HasFTaylorSeriesUpToOn n f p s) {g : E → G} {q : E → FormalMultilinearSeries 𝕜 E G} (hg : HasFTaylorSeriesUpToOn n g q s) : HasFTaylorSeriesUpToOn n (fun y => (f y, g y)) (fun y k => (p y k).prod (q y k)) s := by set L := fun m => ContinuousMultilinearMap.prodL 𝕜 (fun _ : Fin m => E) F G @@ -1131,7 +1132,7 @@ variable {ι ι' : Type*} [Fintype ι] [Fintype ι'] {F' : ι → Type*} [∀ i, [∀ i, NormedSpace 𝕜 (F' i)] {φ : ∀ i, E → F' i} {p' : ∀ i, E → FormalMultilinearSeries 𝕜 E (F' i)} {Φ : E → ∀ i, F' i} {P' : E → FormalMultilinearSeries 𝕜 E (∀ i, F' i)} -theorem hasFTaylorSeriesUpToOn_pi : +theorem hasFTaylorSeriesUpToOn_pi {n : WithTop ℕ∞} : HasFTaylorSeriesUpToOn n (fun x i => φ i x) (fun x m => ContinuousMultilinearMap.pi fun i => p' i x m) s ↔ ∀ i, HasFTaylorSeriesUpToOn n (φ i) (p' i) s := by @@ -1150,7 +1151,7 @@ theorem hasFTaylorSeriesUpToOn_pi : exact (L m).continuous.comp_continuousOn <| continuousOn_pi.2 fun i => (h i).cont m hm @[simp] -theorem hasFTaylorSeriesUpToOn_pi' : +theorem hasFTaylorSeriesUpToOn_pi' {n : WithTop ℕ∞} : HasFTaylorSeriesUpToOn n Φ P' s ↔ ∀ i, HasFTaylorSeriesUpToOn n (fun x => Φ x i) (fun x m => (@ContinuousLinearMap.proj 𝕜 _ ι F' _ _ _ i).compContinuousMultilinearMap @@ -1204,7 +1205,7 @@ end Pi section Add -theorem HasFTaylorSeriesUpToOn.add {q g} (hf : HasFTaylorSeriesUpToOn n f p s) +theorem HasFTaylorSeriesUpToOn.add {n : WithTop ℕ∞} {q g} (hf : HasFTaylorSeriesUpToOn n f p s) (hg : HasFTaylorSeriesUpToOn n g q s) : HasFTaylorSeriesUpToOn n (f + g) (p + q) s := by convert HasFTaylorSeriesUpToOn.continuousLinearMap_comp (ContinuousLinearMap.fst 𝕜 F F + .snd 𝕜 F F) (hf.prod hg) @@ -1641,12 +1642,13 @@ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : refine ⟨{ y : R | IsUnit y }, ?_, ?_⟩ · simpa [nhdsWithin_univ] using x.nhds · use ftaylorSeriesWithin 𝕜 inverse univ - rw [le_antisymm hm bot_le, hasFTaylorSeriesUpToOn_zero_iff] + have : (m : WithTop ℕ∞) = 0 := by exact_mod_cast le_antisymm hm bot_le + rw [this, hasFTaylorSeriesUpToOn_zero_iff] constructor · rintro _ ⟨x', rfl⟩ exact (inverse_continuousAt x').continuousWithinAt · simp [ftaylorSeriesWithin] - · rw [contDiffAt_succ_iff_hasFDerivAt] + · rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] refine ⟨fun x : R => -mulLeftRight 𝕜 R (inverse x) (inverse x), ?_, ?_⟩ · refine ⟨{ y : R | IsUnit y }, x.nhds, ?_⟩ rintro _ ⟨y, rfl⟩ @@ -1752,7 +1754,7 @@ theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomor · rw [contDiffAt_zero] exact ⟨f.target, IsOpen.mem_nhds f.open_target ha, f.continuousOn_invFun⟩ · obtain ⟨f', ⟨u, hu, hff'⟩, hf'⟩ := contDiffAt_succ_iff_hasFDerivAt.mp hf - rw [contDiffAt_succ_iff_hasFDerivAt] + rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] -- For showing `n.succ` times continuous differentiability (the main inductive step), it -- suffices to produce the derivative and show that it is `n` times continuously differentiable have eq_f₀' : f' (f.symm a) = f₀' := (hff' (f.symm a) (mem_of_mem_nhds hu)).unique hf₀' @@ -1871,7 +1873,7 @@ open ContinuousLinearMap (smulRight) /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (formulated with `derivWithin`) is `C^n`. -/ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) : - ContDiffOn 𝕜 (n + 1 : ℕ) f₂ s₂ ↔ + ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (derivWithin f₂ s₂) s₂ := by rw [contDiffOn_succ_iff_fderivWithin hs, and_congr_right_iff] intro _ @@ -1893,7 +1895,7 @@ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂) /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^n`. -/ theorem contDiffOn_succ_iff_deriv_of_isOpen {n : ℕ} (hs : IsOpen s₂) : - ContDiffOn 𝕜 (n + 1 : ℕ) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by + ContDiffOn 𝕜 (n + 1) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by rw [contDiffOn_succ_iff_derivWithin hs.uniqueDiffOn] exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_isOpen hs) @@ -1944,7 +1946,7 @@ theorem ContDiffOn.continuousOn_deriv_of_isOpen (h : ContDiffOn 𝕜 n f₂ s₂ /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `deriv`) is `C^n`. -/ theorem contDiff_succ_iff_deriv {n : ℕ} : - ContDiff 𝕜 (n + 1 : ℕ) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by + ContDiff 𝕜 (n + 1) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_isOpen, isOpen_univ, differentiableOn_univ] @@ -1992,7 +1994,8 @@ variable [NormedSpace 𝕜' E] [IsScalarTower 𝕜 𝕜' E] variable [NormedSpace 𝕜' F] [IsScalarTower 𝕜 𝕜' F] variable {p' : E → FormalMultilinearSeries 𝕜' E F} -theorem HasFTaylorSeriesUpToOn.restrictScalars (h : HasFTaylorSeriesUpToOn n f p' s) : +theorem HasFTaylorSeriesUpToOn.restrictScalars {n : WithTop ℕ∞} + (h : HasFTaylorSeriesUpToOn n f p' s) : HasFTaylorSeriesUpToOn n f (fun x => (p' x).restrictScalars 𝕜) s where zero_eq x hx := h.zero_eq x hx fderivWithin m hm x hx := by @@ -2017,4 +2020,4 @@ theorem ContDiff.restrict_scalars (h : ContDiff 𝕜' n f) : ContDiff 𝕜 n f : end RestrictScalars -set_option linter.style.longFile 2100 +set_option linter.style.longFile 2200 diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index c7f8b24aa799b..968b9b7d7d1c3 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -124,7 +124,7 @@ For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural ` better, is `C^∞` at `0` within `univ`. -/ def ContDiffWithinAt (n : ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := - ∀ m : ℕ, (m : ℕ∞) ≤ n → ∃ u ∈ 𝓝[insert x s] x, + ∀ m : ℕ, m ≤ n → ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn m f p u variable {𝕜} @@ -132,7 +132,7 @@ variable {𝕜} theorem contDiffWithinAt_nat {n : ℕ} : ContDiffWithinAt 𝕜 n f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn n f p u := - ⟨fun H => H n le_rfl, fun ⟨u, hu, p, hp⟩ _m hm => ⟨u, hu, p, hp.of_le hm⟩⟩ + ⟨fun H => H n le_rfl, fun ⟨u, hu, p, hp⟩ _m hm => ⟨u, hu, p, hp.of_le (mod_cast hm)⟩⟩ theorem ContDiffWithinAt.of_le (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m ≤ n) : ContDiffWithinAt 𝕜 m f s x := fun k hk => h k (le_trans hk hmn) @@ -290,30 +290,32 @@ theorem ContDiffWithinAt.differentiableWithinAt (h : ContDiffWithinAt 𝕜 n f s /-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : - ContDiffWithinAt 𝕜 (n + 1 : ℕ) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, + ContDiffWithinAt 𝕜 (n + 1) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) u x) ∧ ContDiffWithinAt 𝕜 n f' u x := by constructor · intro h rcases h n.succ le_rfl with ⟨u, hu, p, Hp⟩ refine ⟨u, hu, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), fun y hy => - Hp.hasFDerivWithinAt (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)) hy, ?_⟩ + Hp.hasFDerivWithinAt (mod_cast (Nat.le_add_left 1 n)) hy, ?_⟩ intro m hm refine ⟨u, ?_, fun y : E => (p y).shift, ?_⟩ · -- Porting note: without the explicit argument Lean is not sure of the type. convert @self_mem_nhdsWithin _ _ x u have : x ∈ insert x s := by simp exact insert_eq_of_mem (mem_of_mem_nhdsWithin this hu) - · rw [hasFTaylorSeriesUpToOn_succ_iff_right] at Hp - exact Hp.2.2.of_le hm + · rw [show ((n.succ : ℕ) : WithTop ℕ∞) = n + 1 from rfl, + hasFTaylorSeriesUpToOn_succ_iff_right] at Hp + exact Hp.2.2.of_le (mod_cast hm) · rintro ⟨u, hu, f', f'_eq_deriv, Hf'⟩ - rw [contDiffWithinAt_nat] + rw [show (n : ℕ∞) + 1 = (n + 1 : ℕ) from rfl, contDiffWithinAt_nat] rcases Hf' n le_rfl with ⟨v, hv, p', Hp'⟩ refine ⟨v ∩ u, ?_, fun x => (p' x).unshift (f x), ?_⟩ · apply Filter.inter_mem _ hu apply nhdsWithin_le_of_mem hu exact nhdsWithin_mono _ (subset_insert x u) hv - · rw [hasFTaylorSeriesUpToOn_succ_iff_right] + · rw [show ((n.succ : ℕ) : WithTop ℕ∞) = n + 1 from rfl, + hasFTaylorSeriesUpToOn_succ_iff_right] refine ⟨fun y _ => rfl, fun y hy => ?_, ?_⟩ · change HasFDerivWithinAt (fun z => (continuousMultilinearCurryFin0 𝕜 E F).symm (f z)) @@ -340,7 +342,7 @@ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : /-- A version of `contDiffWithinAt_succ_iff_hasFDerivWithinAt` where all derivatives are taken within the same set. -/ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt' {n : ℕ} : - ContDiffWithinAt 𝕜 (n + 1 : ℕ) f s x ↔ + ContDiffWithinAt 𝕜 (n + 1) f s x ↔ ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) s x) ∧ ContDiffWithinAt 𝕜 n f' s x := by refine ⟨fun hf => ?_, ?_⟩ @@ -378,13 +380,13 @@ theorem HasFTaylorSeriesUpToOn.contDiffOn {f' : E → FormalMultilinearSeries intro x hx m hm use s simp only [Set.insert_eq_of_mem hx, self_mem_nhdsWithin, true_and] - exact ⟨f', hf.of_le hm⟩ + exact ⟨f', hf.of_le (mod_cast hm)⟩ theorem ContDiffOn.contDiffWithinAt (h : ContDiffOn 𝕜 n f s) (hx : x ∈ s) : ContDiffWithinAt 𝕜 n f s x := h x hx -theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : (m : ℕ∞) ≤ n) +theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : m ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : ∃ u, IsOpen u ∧ x ∈ u ∧ ContDiffOn 𝕜 m f (insert x s ∩ u) := by rcases h m hm with ⟨t, ht, p, hp⟩ @@ -392,7 +394,7 @@ theorem ContDiffWithinAt.contDiffOn' {m : ℕ} (hm : (m : ℕ∞) ≤ n) rw [inter_comm] at hut exact ⟨u, huo, hxu, (hp.mono hut).contDiffOn⟩ -theorem ContDiffWithinAt.contDiffOn {m : ℕ} (hm : (m : ℕ∞) ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : +theorem ContDiffWithinAt.contDiffOn {m : ℕ} (hm : m ≤ n) (h : ContDiffWithinAt 𝕜 n f s x) : ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ContDiffOn 𝕜 m f u := let ⟨_u, uo, xu, h⟩ := h.contDiffOn' hm ⟨_, inter_mem_nhdsWithin _ (uo.mem_nhds xu), inter_subset_left, h⟩ @@ -466,7 +468,7 @@ theorem contDiffOn_of_locally_contDiffOn /-- A function is `C^(n + 1)` on a domain iff locally, it has a derivative which is `C^n`. -/ theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ ∀ x ∈ s, ∃ u ∈ 𝓝[insert x s] x, ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, HasFDerivWithinAt f (f' x) u x) ∧ ContDiffOn 𝕜 n f' u := by constructor @@ -474,10 +476,11 @@ theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : rcases (h x hx) n.succ le_rfl with ⟨u, hu, p, Hp⟩ refine ⟨u, hu, fun y => (continuousMultilinearCurryFin1 𝕜 E F) (p y 1), fun y hy => - Hp.hasFDerivWithinAt (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)) hy, ?_⟩ - rw [hasFTaylorSeriesUpToOn_succ_iff_right] at Hp + Hp.hasFDerivWithinAt (mod_cast (Nat.le_add_left 1 n)) hy, ?_⟩ + rw [show (n.succ : WithTop ℕ∞) = (n : ℕ) + 1 from rfl, + hasFTaylorSeriesUpToOn_succ_iff_right] at Hp intro z hz m hm - refine ⟨u, ?_, fun x : E => (p x).shift, Hp.2.2.of_le hm⟩ + refine ⟨u, ?_, fun x : E => (p x).shift, Hp.2.2.of_le (mod_cast hm)⟩ -- Porting note: without the explicit arguments `convert` can not determine the type. convert @self_mem_nhdsWithin _ _ z u exact insert_eq_of_mem hz @@ -490,7 +493,7 @@ theorem contDiffOn_succ_iff_hasFDerivWithinAt {n : ℕ} : @[simp] theorem contDiffOn_zero : ContDiffOn 𝕜 0 f s ↔ ContinuousOn f s := by refine ⟨fun H => H.continuousOn, fun H => fun x hx m hm ↦ ?_⟩ - have : (m : ℕ∞) = 0 := le_antisymm hm bot_le + have : (m : WithTop ℕ∞) = 0 := le_antisymm (mod_cast hm) bot_le rw [this] refine ⟨insert x s, self_mem_nhdsWithin, ftaylorSeriesWithin 𝕜 f s, ?_⟩ rw [hasFTaylorSeriesUpToOn_zero_iff] @@ -519,7 +522,8 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, iteratedFDerivWithin_zero_apply] · intro m hm x hx - rcases (h x hx) m.succ (Order.add_one_le_of_lt hm) with ⟨u, hu, p, Hp⟩ + have : m < n := mod_cast hm + rcases (h x hx) m.succ (Order.add_one_le_of_lt this) with ⟨u, hu, p, Hp⟩ rw [insert_eq_of_mem hx] at hu rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [inter_comm] at ho @@ -533,15 +537,15 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs change p y m = iteratedFDerivWithin 𝕜 m f s y rw [← iteratedFDerivWithin_inter_open o_open yo] exact - (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn (WithTop.coe_le_coe.2 (Nat.le_succ m)) + (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn (mod_cast Nat.le_succ m) (hs.inter o_open) ⟨hy, yo⟩ exact - ((Hp.mono ho).fderivWithin m (WithTop.coe_lt_coe.2 (lt_add_one m)) x ⟨hx, xo⟩).congr + ((Hp.mono ho).fderivWithin m (mod_cast lt_add_one m) x ⟨hx, xo⟩).congr (fun y hy => (A y hy).symm) (A x ⟨hx, xo⟩).symm · intro m hm apply continuousOn_of_locally_continuousOn intro x hx - rcases h x hx m hm with ⟨u, hu, p, Hp⟩ + rcases h x hx m (mod_cast hm) with ⟨u, hu, p, Hp⟩ rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [insert_eq_of_mem hx] at ho rw [inter_comm] at ho @@ -554,8 +558,8 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm theorem contDiffOn_of_continuousOn_differentiableOn - (Hcont : ∀ m : ℕ, (m : ℕ∞) ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) - (Hdiff : ∀ m : ℕ, (m : ℕ∞) < n → + (Hcont : ∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) + (Hdiff : ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s) : ContDiffOn 𝕜 n f s := by intro x hx m hm @@ -566,27 +570,27 @@ theorem contDiffOn_of_continuousOn_differentiableOn simp only [ftaylorSeriesWithin, ContinuousMultilinearMap.curry0_apply, iteratedFDerivWithin_zero_apply] · intro k hk y hy - convert (Hdiff k (lt_of_lt_of_le hk hm) y hy).hasFDerivWithinAt + convert (Hdiff k (lt_of_lt_of_le (mod_cast hk) hm) y hy).hasFDerivWithinAt · intro k hk - exact Hcont k (le_trans hk hm) + exact Hcont k (le_trans (mod_cast hk) hm) theorem contDiffOn_of_differentiableOn - (h : ∀ m : ℕ, (m : ℕ∞) ≤ n → DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s) : + (h : ∀ m : ℕ, m ≤ n → DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s) : ContDiffOn 𝕜 n f s := contDiffOn_of_continuousOn_differentiableOn (fun m hm => (h m hm).continuousOn) fun m hm => h m (le_of_lt hm) theorem ContDiffOn.continuousOn_iteratedFDerivWithin {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedFDerivWithin 𝕜 m f s) s := - (h.ftaylorSeriesWithin hs).cont m hmn + (hmn : m ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedFDerivWithin 𝕜 m f s) s := + (h.ftaylorSeriesWithin hs).cont m (mod_cast hmn) theorem ContDiffOn.differentiableOn_iteratedFDerivWithin {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 s) : + (hmn : m < n) (hs : UniqueDiffOn 𝕜 s) : DifferentiableOn 𝕜 (iteratedFDerivWithin 𝕜 m f s) s := fun x hx => - ((h.ftaylorSeriesWithin hs).fderivWithin m hmn x hx).differentiableWithinAt + ((h.ftaylorSeriesWithin hs).fderivWithin m (mod_cast hmn) x hx).differentiableWithinAt theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} - (h : ContDiffWithinAt 𝕜 n f s x) (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : + (h : ContDiffWithinAt 𝕜 n f s x) (hmn : m < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedFDerivWithin 𝕜 m f s) s x := by rcases h.contDiffOn' (Order.add_one_le_of_lt hmn) with ⟨u, uo, xu, hu⟩ set t := insert x s ∩ u @@ -605,14 +609,14 @@ theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} theorem contDiffOn_iff_continuousOn_differentiableOn (hs : UniqueDiffOn 𝕜 s) : ContDiffOn 𝕜 n f s ↔ - (∀ m : ℕ, (m : ℕ∞) ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ - ∀ m : ℕ, (m : ℕ∞) < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := + (∀ m : ℕ, m ≤ n → ContinuousOn (fun x => iteratedFDerivWithin 𝕜 m f s x) s) ∧ + ∀ m : ℕ, m < n → DifferentiableOn 𝕜 (fun x => iteratedFDerivWithin 𝕜 m f s x) s := ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin hm hs, fun _m hm => h.differentiableOn_iteratedFDerivWithin hm hs⟩, fun h => contDiffOn_of_continuousOn_differentiableOn h.1 h.2⟩ theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f s) - (h : ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s) : ContDiffOn 𝕜 (n + 1 : ℕ) f s := by + (h : ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s) : ContDiffOn 𝕜 (n + 1) f s := by intro x hx rw [contDiffWithinAt_succ_iff_hasFDerivWithinAt, insert_eq_of_mem hx] exact @@ -621,7 +625,7 @@ theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (expressed with `fderivWithin`) is `C^n`. -/ theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderivWithin 𝕜 f s y) s := by refine ⟨fun H => ?_, fun h => contDiffOn_succ_of_fderivWithin h.1 h.2⟩ refine ⟨H.differentiableOn (WithTop.coe_le_coe.2 (Nat.le_add_left 1 n)), fun x hx => ?_⟩ @@ -639,7 +643,7 @@ theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : rwa [fderivWithin_inter (o_open.mem_nhds hy.2)] at A theorem contDiffOn_succ_iff_hasFDerivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ ∃ f' : E → E →L[𝕜] F, ContDiffOn 𝕜 n f' s ∧ ∀ x, x ∈ s → HasFDerivWithinAt f (f' x) s x := by rw [contDiffOn_succ_iff_fderivWithin hs] refine ⟨fun h => ⟨fderivWithin 𝕜 f s, h.2, fun x hx => (h.1 x hx).hasFDerivWithinAt⟩, fun h => ?_⟩ @@ -650,7 +654,7 @@ theorem contDiffOn_succ_iff_hasFDerivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s) /-- A function is `C^(n + 1)` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/ theorem contDiffOn_succ_iff_fderiv_of_isOpen {n : ℕ} (hs : IsOpen s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderiv 𝕜 f y) s := by rw [contDiffOn_succ_iff_fderivWithin hs.uniqueDiffOn] exact Iff.rfl.and (contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx) @@ -762,7 +766,7 @@ nonrec lemma ContDiffAt.contDiffOn {m : ℕ} (h : ContDiffAt 𝕜 n f x) (hm : m /-- A function is `C^(n + 1)` at a point iff locally, it has a derivative which is `C^n`. -/ theorem contDiffAt_succ_iff_hasFDerivAt {n : ℕ} : - ContDiffAt 𝕜 (n + 1 : ℕ) f x ↔ + ContDiffAt 𝕜 (n + 1) f x ↔ ∃ f' : E → E →L[𝕜] F, (∃ u ∈ 𝓝 x, ∀ x ∈ u, HasFDerivAt f (f' x) x) ∧ ContDiffAt 𝕜 n f' x := by rw [← contDiffWithinAt_univ, contDiffWithinAt_succ_iff_hasFDerivWithinAt] simp only [nhdsWithin_univ, exists_prop, mem_univ, insert_eq_of_mem] @@ -808,7 +812,8 @@ theorem contDiffOn_univ : ContDiffOn 𝕜 n f univ ↔ ContDiff 𝕜 n f := by rw [← hasFTaylorSeriesUpToOn_univ_iff] exact H.ftaylorSeriesWithin uniqueDiffOn_univ · rintro ⟨p, hp⟩ x _ m hm - exact ⟨univ, Filter.univ_sets _, p, (hp.hasFTaylorSeriesUpToOn univ).of_le hm⟩ + exact ⟨univ, Filter.univ_sets _, p, (hp.hasFTaylorSeriesUpToOn univ).of_le + (mod_cast hm)⟩ theorem contDiff_iff_contDiffAt : ContDiff 𝕜 n f ↔ ∀ x, ContDiffAt 𝕜 n f x := by simp [← contDiffOn_univ, ContDiffOn, ContDiffAt] @@ -839,8 +844,8 @@ theorem contDiffAt_zero : ContDiffAt 𝕜 0 f x ↔ ∃ u ∈ 𝓝 x, Continuous theorem contDiffAt_one_iff : ContDiffAt 𝕜 1 f x ↔ ∃ f' : E → E →L[𝕜] F, ∃ u ∈ 𝓝 x, ContinuousOn f' u ∧ ∀ x ∈ u, HasFDerivAt f (f' x) x := by - simp_rw [show (1 : ℕ∞) = (0 + 1 : ℕ) from (zero_add 1).symm, contDiffAt_succ_iff_hasFDerivAt, - show ((0 : ℕ) : ℕ∞) = 0 from rfl, contDiffAt_zero, + rw [show (1 : ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + simp_rw [show ((0 : ℕ) : ℕ∞) = 0 from rfl, contDiffAt_zero, exists_mem_and_iff antitone_bforall antitone_continuousOn, and_comm] theorem ContDiff.of_le (h : ContDiff 𝕜 n f) (hmn : m ≤ n) : ContDiff 𝕜 m f := @@ -864,7 +869,7 @@ theorem contDiff_iff_forall_nat_le : ContDiff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ /-- A function is `C^(n+1)` iff it has a `C^n` derivative. -/ theorem contDiff_succ_iff_hasFDerivAt {n : ℕ} : - ContDiff 𝕜 (n + 1 : ℕ) f ↔ + ContDiff 𝕜 (n + 1) f ↔ ∃ f' : E → E →L[𝕜] F, ContDiff 𝕜 n f' ∧ ∀ x, HasFDerivAt f (f' x) x := by simp only [← contDiffOn_univ, ← hasFDerivWithinAt_univ, contDiffOn_succ_iff_hasFDerivWithin uniqueDiffOn_univ, Set.mem_univ, forall_true_left] @@ -884,30 +889,30 @@ theorem contDiff_iff_ftaylorSeries : theorem contDiff_iff_continuous_differentiable : ContDiff 𝕜 n f ↔ - (∀ m : ℕ, (m : ℕ∞) ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ - ∀ m : ℕ, (m : ℕ∞) < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by + (∀ m : ℕ, m ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ + ∀ m : ℕ, m < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by simp [contDiffOn_univ.symm, continuous_iff_continuousOn_univ, differentiableOn_univ.symm, iteratedFDerivWithin_univ, contDiffOn_iff_continuousOn_differentiableOn uniqueDiffOn_univ] /-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ -theorem ContDiff.continuous_iteratedFDeriv {m : ℕ} (hm : (m : ℕ∞) ≤ n) (hf : ContDiff 𝕜 n f) : +theorem ContDiff.continuous_iteratedFDeriv {m : ℕ} (hm : m ≤ n) (hf : ContDiff 𝕜 n f) : Continuous fun x => iteratedFDeriv 𝕜 m f x := (contDiff_iff_continuous_differentiable.mp hf).1 m hm /-- If `f` is `C^n` then its `m`-times iterated derivative is differentiable for `m < n`. -/ -theorem ContDiff.differentiable_iteratedFDeriv {m : ℕ} (hm : (m : ℕ∞) < n) (hf : ContDiff 𝕜 n f) : +theorem ContDiff.differentiable_iteratedFDeriv {m : ℕ} (hm : m < n) (hf : ContDiff 𝕜 n f) : Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := (contDiff_iff_continuous_differentiable.mp hf).2 m hm theorem contDiff_of_differentiable_iteratedFDeriv - (h : ∀ m : ℕ, (m : ℕ∞) ≤ n → Differentiable 𝕜 (iteratedFDeriv 𝕜 m f)) : ContDiff 𝕜 n f := + (h : ∀ m : ℕ, m ≤ n → Differentiable 𝕜 (iteratedFDeriv 𝕜 m f)) : ContDiff 𝕜 n f := contDiff_iff_continuous_differentiable.2 ⟨fun m hm => (h m hm).continuous, fun m hm => h m (le_of_lt hm)⟩ /-- A function is `C^(n + 1)` if and only if it is differentiable, and its derivative (formulated in terms of `fderiv`) is `C^n`. -/ theorem contDiff_succ_iff_fderiv {n : ℕ} : - ContDiff 𝕜 (n + 1 : ℕ) f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 n fun y => fderiv 𝕜 f y := by + ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ ContDiff 𝕜 n fun y => fderiv 𝕜 f y := by simp only [← contDiffOn_univ, ← differentiableOn_univ, ← fderivWithin_univ, contDiffOn_succ_iff_fderivWithin uniqueDiffOn_univ] diff --git a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean index 149c86f2c306e..abdac0002a6e2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FTaylorSeries.lean @@ -99,7 +99,7 @@ In this file, we denote `⊤ : ℕ∞` with `∞`. noncomputable section open scoped Classical -open NNReal Topology Filter +open ENat NNReal Topology Filter local notation "∞" => (⊤ : ℕ∞) @@ -115,7 +115,8 @@ universe u uE uF variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - {s t u : Set E} {f f₁ : E → F} {x : E} {m n : ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} + {s t u : Set E} {f f₁ : E → F} {x : E} {m n N : WithTop ℕ∞} + {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Functions with a Taylor series on a domain -/ @@ -125,12 +126,12 @@ derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a pred Notice that `p` does not sum up to `f` on the diagonal (`FormalMultilinearSeries.sum`), even if `f` is analytic and `n = ∞`: an additional `1/m!` factor on the `m`th term is necessary for that. -/ -structure HasFTaylorSeriesUpToOn (n : ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) - (s : Set E) : Prop where +structure HasFTaylorSeriesUpToOn + (n : WithTop ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) (s : Set E) : Prop where zero_eq : ∀ x ∈ s, (p x 0).curry0 = f x protected fderivWithin : ∀ m : ℕ, (m : ℕ∞) < n → ∀ x ∈ s, HasFDerivWithinAt (p · m) (p x m.succ).curryLeft s x - cont : ∀ m : ℕ, (m : ℕ∞) ≤ n → ContinuousOn (p · m) s + cont : ∀ m : ℕ, m ≤ n → ContinuousOn (p · m) s theorem HasFTaylorSeriesUpToOn.zero_eq' (h : HasFTaylorSeriesUpToOn n f p s) {x : E} (hx : x ∈ s) : p x 0 = (continuousMultilinearCurryFin0 𝕜 E F).symm (f x) := by @@ -170,26 +171,31 @@ theorem hasFTaylorSeriesUpToOn_zero_iff : rw [continuousOn_congr this, LinearIsometryEquiv.comp_continuousOn_iff] exact H.1 -theorem hasFTaylorSeriesUpToOn_top_iff : - HasFTaylorSeriesUpToOn ∞ f p s ↔ ∀ n : ℕ, HasFTaylorSeriesUpToOn n f p s := by +theorem hasFTaylorSeriesUpToOn_top_iff_add (hN : ∞ ≤ N) (k : ℕ) : + HasFTaylorSeriesUpToOn N f p s ↔ ∀ n : ℕ, HasFTaylorSeriesUpToOn (n + k : ℕ) f p s := by constructor - · intro H n; exact H.of_le le_top + · intro H n + apply H.of_le (natCast_le_of_coe_top_le_withTop hN _) · intro H constructor · exact (H 0).zero_eq · intro m _ - apply (H m.succ).fderivWithin m (WithTop.coe_lt_coe.2 (lt_add_one m)) + apply (H m.succ).fderivWithin m (by norm_cast; omega) · intro m _ - apply (H m).cont m le_rfl + apply (H m).cont m (by simp) + +theorem hasFTaylorSeriesUpToOn_top_iff (hN : ∞ ≤ N) : + HasFTaylorSeriesUpToOn N f p s ↔ ∀ n : ℕ, HasFTaylorSeriesUpToOn n f p s := by + simpa using hasFTaylorSeriesUpToOn_top_iff_add hN 0 /-- In the case that `n = ∞` we don't need the continuity assumption in `HasFTaylorSeriesUpToOn`. -/ -theorem hasFTaylorSeriesUpToOn_top_iff' : - HasFTaylorSeriesUpToOn ∞ f p s ↔ +theorem hasFTaylorSeriesUpToOn_top_iff' (hN : ∞ ≤ N) : + HasFTaylorSeriesUpToOn N f p s ↔ (∀ x ∈ s, (p x 0).curry0 = f x) ∧ - ∀ m : ℕ, ∀ x ∈ s, HasFDerivWithinAt (fun y => p y m) (p x m.succ).curryLeft s x := + ∀ m : ℕ, ∀ x ∈ s, HasFDerivWithinAt (fun y => p y m) (p x m.succ).curryLeft s x := by -- Everything except for the continuity is trivial: - ⟨fun h => ⟨h.1, fun m => h.2 m (WithTop.coe_lt_top m)⟩, fun h => + refine ⟨fun h => ⟨h.1, fun m => h.2 m (natCast_lt_of_coe_top_le_withTop hN _)⟩, fun h => ⟨h.1, fun m _ => h.2 m, fun m _ x hx => -- The continuity follows from the existence of a derivative: (h.2 m x hx).continuousWithinAt⟩⟩ @@ -241,21 +247,21 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_left {n : ℕ} : (∀ x ∈ s, HasFDerivWithinAt (fun y => p y n) (p x n.succ).curryLeft s x) ∧ ContinuousOn (fun x => p x (n + 1)) s := by constructor - · exact fun h ↦ ⟨h.of_le (WithTop.coe_le_coe.2 (Nat.le_succ n)), - h.fderivWithin _ (WithTop.coe_lt_coe.2 (lt_add_one n)), h.cont (n + 1) le_rfl⟩ + · exact fun h ↦ ⟨h.of_le (mod_cast Nat.le_succ n), + h.fderivWithin _ (mod_cast lt_add_one n), h.cont (n + 1) le_rfl⟩ · intro h constructor · exact h.1.zero_eq · intro m hm by_cases h' : m < n - · exact h.1.fderivWithin m (WithTop.coe_lt_coe.2 h') - · have : m = n := Nat.eq_of_lt_succ_of_not_lt (WithTop.coe_lt_coe.1 hm) h' + · exact h.1.fderivWithin m (mod_cast h') + · have : m = n := Nat.eq_of_lt_succ_of_not_lt (mod_cast hm) h' rw [this] exact h.2.1 · intro m hm by_cases h' : m ≤ n - · apply h.1.cont m (WithTop.coe_le_coe.2 h') - · have : m = n + 1 := le_antisymm (WithTop.coe_le_coe.1 hm) (not_le.1 h') + · apply h.1.cont m (mod_cast h') + · have : m = n + 1 := le_antisymm (mod_cast hm) (not_le.1 h') rw [this] exact h.2.2 @@ -273,8 +279,8 @@ theorem HasFTaylorSeriesUpToOn.shift_of_succ constructor · intro x _ rfl - · intro m (hm : (m : ℕ∞) < n) x (hx : x ∈ s) - have A : (m.succ : ℕ∞) < n.succ := by + · intro m (hm : (m : WithTop ℕ∞) < n) x (hx : x ∈ s) + have A : (m.succ : WithTop ℕ∞) < n.succ := by rw [Nat.cast_lt] at hm ⊢ exact Nat.succ_lt_succ hm change HasFDerivWithinAt (continuousMultilinearCurryRightEquiv' 𝕜 m E F ∘ (p · m.succ)) @@ -284,7 +290,7 @@ theorem HasFTaylorSeriesUpToOn.shift_of_succ ext y v change p x (m + 2) (snoc (cons y (init v)) (v (last _))) = p x (m + 2) (cons y v) rw [← cons_snoc_eq_snoc_cons, snoc_init_self] - · intro m (hm : (m : ℕ∞) ≤ n) + · intro m (hm : (m : WithTop ℕ∞) ≤ n) suffices A : ContinuousOn (p · (m + 1)) s from (continuousMultilinearCurryRightEquiv' 𝕜 m E F).continuous.comp_continuousOn A refine H.cont _ ?_ @@ -292,8 +298,8 @@ theorem HasFTaylorSeriesUpToOn.shift_of_succ exact Nat.succ_le_succ hm /-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` -for `p 1`, which is a derivative of `f`. -/ -theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : +for `p 1`, which is a derivative of `f`. Version for `n : ℕ`. -/ +theorem hasFTaylorSeriesUpToOn_succ_nat_iff_right {n : ℕ} : HasFTaylorSeriesUpToOn (n + 1 : ℕ) f p s ↔ (∀ x ∈ s, (p x 0).curry0 = f x) ∧ (∀ x ∈ s, HasFDerivWithinAt (fun y => p y 0) (p x 1).curryLeft s x) ∧ @@ -306,10 +312,10 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : · rintro ⟨Hzero_eq, Hfderiv_zero, Htaylor⟩ constructor · exact Hzero_eq - · intro m (hm : (m : ℕ∞) < n.succ) x (hx : x ∈ s) + · intro m (hm : (m : WithTop ℕ∞) < n.succ) x (hx : x ∈ s) cases' m with m · exact Hfderiv_zero x hx - · have A : (m : ℕ∞) < n := by + · have A : (m : WithTop ℕ∞) < n := by rw [Nat.cast_lt] at hm ⊢ exact Nat.lt_of_succ_lt_succ hm have : @@ -322,7 +328,7 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : (p x (Nat.succ (Nat.succ m))) (cons y v) = (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))) rw [← cons_snoc_eq_snoc_cons, snoc_init_self] - · intro m (hm : (m : ℕ∞) ≤ n.succ) + · intro m (hm : (m : WithTop ℕ∞) ≤ n.succ) cases' m with m · have : DifferentiableOn 𝕜 (fun x => p x 0) s := fun x hx => (Hfderiv_zero x hx).differentiableWithinAt @@ -332,6 +338,37 @@ theorem hasFTaylorSeriesUpToOn_succ_iff_right {n : ℕ} : rw [Nat.cast_le] at hm ⊢ exact Nat.lt_succ_iff.mp hm +/-- `p` is a Taylor series of `f` up to `⊤` if and only if `p.shift` is a Taylor series up to `⊤` +for `p 1`, which is a derivative of `f`. -/ +theorem hasFTaylorSeriesUpToOn_top_iff_right (hN : ∞ ≤ N) : + HasFTaylorSeriesUpToOn N f p s ↔ + (∀ x ∈ s, (p x 0).curry0 = f x) ∧ + (∀ x ∈ s, HasFDerivWithinAt (fun y => p y 0) (p x 1).curryLeft s x) ∧ + HasFTaylorSeriesUpToOn N (fun x => continuousMultilinearCurryFin1 𝕜 E F (p x 1)) + (fun x => (p x).shift) s := by + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rw [hasFTaylorSeriesUpToOn_top_iff_add hN 1] at h + rw [hasFTaylorSeriesUpToOn_top_iff hN] + exact ⟨(hasFTaylorSeriesUpToOn_succ_nat_iff_right.1 (h 1)).1, + (hasFTaylorSeriesUpToOn_succ_nat_iff_right.1 (h 1)).2.1, + fun n ↦ (hasFTaylorSeriesUpToOn_succ_nat_iff_right.1 (h n)).2.2⟩ + · apply (hasFTaylorSeriesUpToOn_top_iff_add hN 1).2 (fun n ↦ ?_) + rw [hasFTaylorSeriesUpToOn_succ_nat_iff_right] + exact ⟨h.1, h.2.1, (h.2.2).of_le (m := n) (natCast_le_of_coe_top_le_withTop hN n)⟩ + +/-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` +for `p 1`, which is a derivative of `f`. Version for `n : WithTop ℕ∞`. -/ +theorem hasFTaylorSeriesUpToOn_succ_iff_right : + HasFTaylorSeriesUpToOn (n + 1) f p s ↔ + (∀ x ∈ s, (p x 0).curry0 = f x) ∧ + (∀ x ∈ s, HasFDerivWithinAt (fun y => p y 0) (p x 1).curryLeft s x) ∧ + HasFTaylorSeriesUpToOn n (fun x => continuousMultilinearCurryFin1 𝕜 E F (p x 1)) + (fun x => (p x).shift) s := by + match n with + | ⊤ => exact hasFTaylorSeriesUpToOn_top_iff_right (by simp) + | (⊤ : ℕ∞) => exact hasFTaylorSeriesUpToOn_top_iff_right (by simp) + | (n : ℕ) => exact hasFTaylorSeriesUpToOn_succ_nat_iff_right + /-! ### Iterated derivative within a set -/ @@ -537,11 +574,11 @@ theorem iteratedFDerivWithin_inter_open {n : ℕ} (hu : IsOpen u) (hx : x ∈ u) /-- On a set with unique differentiability, any choice of iterated differential has to coincide with the one we have chosen in `iteratedFDerivWithin 𝕜 m f s`. -/ theorem HasFTaylorSeriesUpToOn.eq_iteratedFDerivWithin_of_uniqueDiffOn - (h : HasFTaylorSeriesUpToOn n f p s) {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) + (h : HasFTaylorSeriesUpToOn n f p s) {m : ℕ} (hmn : m ≤ n) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : p x m = iteratedFDerivWithin 𝕜 m f s x := by induction' m with m IH generalizing x · rw [h.zero_eq' hx, iteratedFDerivWithin_zero_eq_comp]; rfl - · have A : (m : ℕ∞) < n := lt_of_lt_of_le (WithTop.coe_lt_coe.2 (lt_add_one m)) hmn + · have A : (m : ℕ∞) < n := lt_of_lt_of_le (mod_cast lt_add_one m) hmn have : HasFDerivWithinAt (fun y : E => iteratedFDerivWithin 𝕜 m f s y) (ContinuousMultilinearMap.curryLeft (p x (Nat.succ m))) s x := @@ -563,11 +600,11 @@ derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a pred Notice that `p` does not sum up to `f` on the diagonal (`FormalMultilinearSeries.sum`), even if `f` is analytic and `n = ∞`: an addition `1/m!` factor on the `m`th term is necessary for that. -/ -structure HasFTaylorSeriesUpTo (n : ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) : - Prop where +structure HasFTaylorSeriesUpTo + (n : WithTop ℕ∞) (f : E → F) (p : E → FormalMultilinearSeries 𝕜 E F) : Prop where zero_eq : ∀ x, (p x 0).curry0 = f x - fderiv : ∀ m : ℕ, (m : ℕ∞) < n → ∀ x, HasFDerivAt (fun y => p y m) (p x m.succ).curryLeft x - cont : ∀ m : ℕ, (m : ℕ∞) ≤ n → Continuous fun x => p x m + fderiv : ∀ m : ℕ, m < n → ∀ x, HasFDerivAt (fun y => p y m) (p x m.succ).curryLeft x + cont : ∀ m : ℕ, m ≤ n → Continuous fun x => p x m theorem HasFTaylorSeriesUpTo.zero_eq' (h : HasFTaylorSeriesUpTo n f p) (x : E) : p x 0 = (continuousMultilinearCurryFin0 𝕜 E F).symm (f x) := by @@ -600,10 +637,13 @@ theorem HasFTaylorSeriesUpTo.hasFTaylorSeriesUpToOn (h : HasFTaylorSeriesUpTo n HasFTaylorSeriesUpToOn n f p s := (hasFTaylorSeriesUpToOn_univ_iff.2 h).mono (subset_univ _) -theorem HasFTaylorSeriesUpTo.ofLe (h : HasFTaylorSeriesUpTo n f p) (hmn : m ≤ n) : +theorem HasFTaylorSeriesUpTo.of_le (h : HasFTaylorSeriesUpTo n f p) (hmn : m ≤ n) : HasFTaylorSeriesUpTo m f p := by rw [← hasFTaylorSeriesUpToOn_univ_iff] at h ⊢; exact h.of_le hmn +@[deprecated (since := "2024-11-07")] +alias HasFTaylorSeriesUpTo.ofLe := HasFTaylorSeriesUpTo.of_le + theorem HasFTaylorSeriesUpTo.continuous (h : HasFTaylorSeriesUpTo n f p) : Continuous f := by rw [← hasFTaylorSeriesUpToOn_univ_iff] at h rw [continuous_iff_continuousOn_univ] @@ -614,17 +654,17 @@ theorem hasFTaylorSeriesUpTo_zero_iff : simp [hasFTaylorSeriesUpToOn_univ_iff.symm, continuous_iff_continuousOn_univ, hasFTaylorSeriesUpToOn_zero_iff] -theorem hasFTaylorSeriesUpTo_top_iff : - HasFTaylorSeriesUpTo ∞ f p ↔ ∀ n : ℕ, HasFTaylorSeriesUpTo n f p := by - simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff] +theorem hasFTaylorSeriesUpTo_top_iff (hN : ∞ ≤ N) : + HasFTaylorSeriesUpTo N f p ↔ ∀ n : ℕ, HasFTaylorSeriesUpTo n f p := by + simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff hN] /-- In the case that `n = ∞` we don't need the continuity assumption in `HasFTaylorSeriesUpTo`. -/ -theorem hasFTaylorSeriesUpTo_top_iff' : - HasFTaylorSeriesUpTo ∞ f p ↔ +theorem hasFTaylorSeriesUpTo_top_iff' (hN : ∞ ≤ N) : + HasFTaylorSeriesUpTo N f p ↔ (∀ x, (p x 0).curry0 = f x) ∧ ∀ (m : ℕ) (x), HasFDerivAt (fun y => p y m) (p x m.succ).curryLeft x := by - simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff', mem_univ, + simp only [← hasFTaylorSeriesUpToOn_univ_iff, hasFTaylorSeriesUpToOn_top_iff' hN, mem_univ, forall_true_left, hasFDerivWithinAt_univ] /-- If a function has a Taylor series at order at least `1`, then the term of order `1` of this @@ -639,15 +679,17 @@ theorem HasFTaylorSeriesUpTo.differentiable (h : HasFTaylorSeriesUpTo n f p) (hn /-- `p` is a Taylor series of `f` up to `n+1` if and only if `p.shift` is a Taylor series up to `n` for `p 1`, which is a derivative of `f`. -/ -theorem hasFTaylorSeriesUpTo_succ_iff_right {n : ℕ} : +theorem hasFTaylorSeriesUpTo_succ_nat_iff_right {n : ℕ} : HasFTaylorSeriesUpTo (n + 1 : ℕ) f p ↔ (∀ x, (p x 0).curry0 = f x) ∧ (∀ x, HasFDerivAt (fun y => p y 0) (p x 1).curryLeft x) ∧ HasFTaylorSeriesUpTo n (fun x => continuousMultilinearCurryFin1 𝕜 E F (p x 1)) fun x => (p x).shift := by - simp only [hasFTaylorSeriesUpToOn_succ_iff_right, ← hasFTaylorSeriesUpToOn_univ_iff, mem_univ, + simp only [hasFTaylorSeriesUpToOn_succ_nat_iff_right, ← hasFTaylorSeriesUpToOn_univ_iff, mem_univ, forall_true_left, hasFDerivWithinAt_univ] +@[deprecated (since := "2024-11-07")] +alias hasFTaylorSeriesUpTo_succ_iff_right := hasFTaylorSeriesUpTo_succ_nat_iff_right /-! ### Iterated derivative -/ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean index 1c2e5249d4f46..582c9ca4fec35 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean @@ -917,7 +917,7 @@ private lemma faaDiBruno_aux2 {m : ℕ} (q : FormalMultilinearSeries 𝕜 F G) /-- *Faa di Bruno* formula: If two functions `g` and `f` have Taylor series up to `n` given by `q` and `p`, then `g ∘ f` also has a Taylor series, given by `q.taylorComp p`. -/ -theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} +theorem HasFTaylorSeriesUpToOn.comp {n : WithTop ℕ∞} {g : F → G} {f : E → F} (hg : HasFTaylorSeriesUpToOn n g q t) (hf : HasFTaylorSeriesUpToOn n f p s) (h : MapsTo f s t) : HasFTaylorSeriesUpToOn n (g ∘ f) (fun x ↦ (q (f x)).taylorComp (p x)) s := by /- One has to check that the `m+1`-th term is the derivative of the `m`-th term. The `m`-th term @@ -940,8 +940,8 @@ theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} change HasFDerivWithinAt (fun y ↦ B (q (f y) c.length) (fun i ↦ p y (c.partSize i))) (∑ i : Option (Fin c.length), ((q (f x)).compAlongOrderedFinpartition (p x) (c.extend i)).curryLeft) s x - have cm : (c.length : ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.length_le c - have cp i : (c.partSize i : ℕ∞) ≤ m := by + have cm : (c.length : WithTop ℕ∞) ≤ m := mod_cast OrderedFinpartition.length_le c + have cp i : (c.partSize i : WithTop ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.partSize_le c i have I i : HasFDerivWithinAt (fun x ↦ p x (c.partSize i)) (p x (c.partSize i).succ).curryLeft s x := @@ -949,7 +949,8 @@ theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} have J : HasFDerivWithinAt (fun x ↦ q x c.length) (q (f x) c.length.succ).curryLeft t (f x) := hg.fderivWithin c.length (cm.trans_lt hm) (f x) (h hx) have K : HasFDerivWithinAt f ((continuousMultilinearCurryFin1 𝕜 E F) (p x 1)) s x := - hf.hasFDerivWithinAt (le_trans le_add_self (Order.add_one_le_of_lt hm)) hx + hf.hasFDerivWithinAt (le_trans (mod_cast Nat.le_add_left 1 m) + (ENat.add_one_natCast_le_withTop_of_lt hm)) hx convert HasFDerivWithinAt.linear_multilinear_comp (J.comp x K h) I B simp only [Nat.succ_eq_add_one, Fintype.sum_option, comp_apply, faaDiBruno_aux1, faaDiBruno_aux2] @@ -975,8 +976,9 @@ theorem HasFTaylorSeriesUpToOn.comp {n : ℕ∞} {g : F → G} {f : E → F} change ContinuousOn ((fun p ↦ B p.1 p.2) ∘ (fun x ↦ (q (f x) c.length, fun i ↦ p x (c.partSize i)))) s apply B.continuous_uncurry_of_multilinear.comp_continuousOn (ContinuousOn.prod ?_ ?_) - · have : (c.length : ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.length_le c + · have : (c.length : WithTop ℕ∞) ≤ m := mod_cast OrderedFinpartition.length_le c exact (hg.cont c.length (this.trans hm)).comp hf.continuousOn h · apply continuousOn_pi.2 (fun i ↦ ?_) - have : (c.partSize i : ℕ∞) ≤ m := by exact_mod_cast OrderedFinpartition.partSize_le c i + have : (c.partSize i : WithTop ℕ∞) ≤ m := by + exact_mod_cast OrderedFinpartition.partSize_le c i exact hf.cont _ (this.trans hm) diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index e5d673c7ffac0..555d81bb107e4 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -54,17 +54,17 @@ often requires an inconvenient need to generalize `F`, which results in universe This lemma avoids these universe issues, but only applies for finite dimensional `E`. -/ theorem contDiff_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} : - ContDiff 𝕜 (n + 1 : ℕ) f ↔ Differentiable 𝕜 f ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by + ContDiff 𝕜 (n + 1) f ↔ Differentiable 𝕜 f ∧ ∀ y, ContDiff 𝕜 n fun x => fderiv 𝕜 f x y := by rw [contDiff_succ_iff_fderiv, contDiff_clm_apply_iff] theorem contDiffOn_succ_of_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} (hf : DifferentiableOn 𝕜 f s) (h : ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s := + ContDiffOn 𝕜 (n + 1) f s := contDiffOn_succ_of_fderivWithin hf <| contDiffOn_clm_apply.mpr h theorem contDiffOn_succ_iff_fderiv_apply [FiniteDimensional 𝕜 D] {n : ℕ} {f : D → E} {s : Set D} (hs : UniqueDiffOn 𝕜 s) : - ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔ + ContDiffOn 𝕜 (n + 1) f s ↔ DifferentiableOn 𝕜 f s ∧ ∀ y, ContDiffOn 𝕜 n (fun x => fderivWithin 𝕜 f s x y) s := by rw [contDiffOn_succ_iff_fderivWithin hs, contDiffOn_clm_apply] diff --git a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean index d9f341edf671a..5f8fdba324bcb 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean @@ -29,7 +29,8 @@ variable {n : ℕ∞} {𝕂 : Type*} [RCLike 𝕂] {E' : Type*} [NormedAddCommGr /-- If a function has a Taylor series at order at least 1, then at points in the interior of the domain of definition, the term of order 1 of this series is a strict derivative of `f`. -/ -theorem HasFTaylorSeriesUpToOn.hasStrictFDerivAt {s : Set E'} {f : E' → F'} {x : E'} +theorem HasFTaylorSeriesUpToOn.hasStrictFDerivAt {n : WithTop ℕ∞} + {s : Set E'} {f : E' → F'} {x : E'} {p : E' → FormalMultilinearSeries 𝕂 E' F'} (hf : HasFTaylorSeriesUpToOn n f p s) (hn : 1 ≤ n) (hs : s ∈ 𝓝 x) : HasStrictFDerivAt f ((continuousMultilinearCurryFin1 𝕂 E' F') (p x 1)) x := hasStrictFDerivAt_of_hasFDerivAt_of_continuousAt (hf.eventually_hasFDerivAt hn hs) <| diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 76c137c4bab78..055c8cb25720e 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -382,7 +382,7 @@ protected theorem AnalyticOn.iteratedFDerivWithin (h : AnalyticOn 𝕜 f s) apply AnalyticOnNhd.comp_analyticOn _ (IH.fderivWithin hu) (mapsTo_univ _ _) apply LinearIsometryEquiv.analyticOnNhd -lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : ℕ∞} +lemma AnalyticOn.hasFTaylorSeriesUpToOn {n : WithTop ℕ∞} (h : AnalyticOn 𝕜 f s) (hu : UniqueDiffOn 𝕜 s) : HasFTaylorSeriesUpToOn n f (ftaylorSeriesWithin 𝕜 f s) s := by refine ⟨fun x _hx ↦ rfl, fun m _hm x hx ↦ ?_, fun m _hm x hx ↦ ?_⟩ diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index a50566576f0f1..e3a80a3cd256e 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1187,7 +1187,8 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} have A : ∀ q₀ : P × G, q₀.1 ∈ s → HasFDerivAt (fun q : P × G => (f ⋆[L, μ] g q.1) q.2) (f' q₀.1 q₀.2) q₀ := hasFDerivAt_convolution_right_with_param L hs hk hgs hf hg.one_of_succ - rw [contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ + rw [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, + contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ constructor · rintro ⟨p, x⟩ ⟨hp, -⟩ exact (A (p, x) hp).differentiableAt.differentiableWithinAt diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 7a53f32811c4c..5b882dad07252 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -440,7 +440,7 @@ lemma integrable_fourierPowSMulRight {n : ℕ} (hf : Integrable (fun v ↦ ‖v filter_upwards with v exact (norm_fourierPowSMulRight_le L f v n).trans (le_of_eq (by ring)) -lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} +lemma hasFTaylorSeriesUpTo_fourierIntegral {N : WithTop ℕ∞} (hf : ∀ (n : ℕ), n ≤ N → Integrable (fun v ↦ ‖v‖^n * ‖f v‖) μ) (h'f : AEStronglyMeasurable f μ) : HasFTaylorSeriesUpTo N (fourierIntegral 𝐞 μ L.toLinearMap₂ f) @@ -455,7 +455,8 @@ lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} have I₁ : Integrable (fun v ↦ fourierPowSMulRight L f v n) μ := integrable_fourierPowSMulRight L (hf n hn.le) h'f have I₂ : Integrable (fun v ↦ ‖v‖ * ‖fourierPowSMulRight L f v n‖) μ := by - apply ((hf (n+1) (Order.add_one_le_of_lt hn)).const_mul ((2 * π * ‖L‖) ^ n)).mono' + apply ((hf (n+1) (ENat.add_one_natCast_le_withTop_of_lt hn)).const_mul + ((2 * π * ‖L‖) ^ n)).mono' (continuous_norm.aestronglyMeasurable.mul (h'f.fourierPowSMulRight L n).norm) filter_upwards with v simp only [Pi.mul_apply, norm_mul, norm_norm] @@ -465,7 +466,7 @@ lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} gcongr; apply norm_fourierPowSMulRight_le _ = (2 * π * ‖L‖) ^ n * (‖v‖ ^ (n + 1) * ‖f v‖) := by rw [pow_succ]; ring have I₃ : Integrable (fun v ↦ fourierPowSMulRight L f v (n + 1)) μ := - integrable_fourierPowSMulRight L (hf (n + 1) (Order.add_one_le_of_lt hn)) h'f + integrable_fourierPowSMulRight L (hf (n + 1) (ENat.add_one_natCast_le_withTop_of_lt hn)) h'f have I₄ : Integrable (fun v ↦ fourierSMulRight L (fun v ↦ fourierPowSMulRight L f v n) v) μ := by apply (I₂.const_mul ((2 * π * ‖L‖))).mono' (h'f.fourierPowSMulRight L n).fourierSMulRight @@ -488,12 +489,22 @@ lemma hasFTaylorSeriesUpTo_fourierIntegral {N : ℕ∞} apply fourierIntegral_continuous Real.continuous_fourierChar (by apply L.continuous₂) exact integrable_fourierPowSMulRight L (hf n hn) h'f +/-- Variant of `hasFTaylorSeriesUpTo_fourierIntegral` in which the smoothness index is restricted +to `ℕ∞` (and so are the inequalities in the assumption `hf`). Avoids normcasting in some +applications. -/ +lemma hasFTaylorSeriesUpTo_fourierIntegral' {N : ℕ∞} + (hf : ∀ (n : ℕ), n ≤ N → Integrable (fun v ↦ ‖v‖^n * ‖f v‖) μ) + (h'f : AEStronglyMeasurable f μ) : + HasFTaylorSeriesUpTo N (fourierIntegral 𝐞 μ L.toLinearMap₂ f) + (fun w n ↦ fourierIntegral 𝐞 μ L.toLinearMap₂ (fun v ↦ fourierPowSMulRight L f v n) w) := + hasFTaylorSeriesUpTo_fourierIntegral _ (fun n hn ↦ hf n (mod_cast hn)) h'f + /-- If `‖v‖^n * ‖f v‖` is integrable for all `n ≤ N`, then the Fourier transform of `f` is `C^N`. -/ theorem contDiff_fourierIntegral {N : ℕ∞} (hf : ∀ (n : ℕ), n ≤ N → Integrable (fun v ↦ ‖v‖^n * ‖f v‖) μ) : ContDiff ℝ N (fourierIntegral 𝐞 μ L.toLinearMap₂ f) := by by_cases h'f : Integrable f μ - · exact (hasFTaylorSeriesUpTo_fourierIntegral L hf h'f.1).contDiff + · exact (hasFTaylorSeriesUpTo_fourierIntegral' L hf h'f.1).contDiff · have : fourierIntegral 𝐞 μ L.toLinearMap₂ f = 0 := by ext w; simp [fourierIntegral, integral, h'f] simpa [this] using contDiff_const @@ -507,7 +518,8 @@ lemma iteratedFDeriv_fourierIntegral {N : ℕ∞} iteratedFDeriv ℝ n (fourierIntegral 𝐞 μ L.toLinearMap₂ f) = fourierIntegral 𝐞 μ L.toLinearMap₂ (fun v ↦ fourierPowSMulRight L f v n) := by ext w : 1 - exact ((hasFTaylorSeriesUpTo_fourierIntegral L hf h'f).eq_iteratedFDeriv hn w).symm + exact ((hasFTaylorSeriesUpTo_fourierIntegral' L hf h'f).eq_iteratedFDeriv + (mod_cast hn) w).symm end SecondCountableTopology diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 6a56975c150b0..d05eed644c71c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -368,7 +368,7 @@ theorem contDiff_rpow_const_of_le {p : ℝ} {n : ℕ} (h : ↑n ≤ p) : · exact contDiff_zero.2 (continuous_id.rpow_const fun x => Or.inr <| by simpa using h) · have h1 : 1 ≤ p := le_trans (by simp) h rw [Nat.cast_succ, ← le_sub_iff_add_le] at h - rw [contDiff_succ_iff_deriv, deriv_rpow_const' h1] + rw [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, contDiff_succ_iff_deriv, deriv_rpow_const' h1] exact ⟨differentiable_rpow_const h1, contDiff_const.mul (ihn h)⟩ theorem contDiffAt_rpow_const_of_le {x p : ℝ} {n : ℕ} (h : ↑n ≤ p) : diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 2c4dde0e12e6a..663a5824cbe31 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -269,6 +269,15 @@ theorem nat_induction {P : ℕ∞ → Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ · exact htop A · exact A _ +lemma add_one_pos : 0 < n + 1 := + succ_def n ▸ Order.bot_lt_succ n + +lemma add_lt_add_iff_right {k : ℕ∞} (h : k ≠ ⊤) : n + k < m + k ↔ n < m := + WithTop.add_lt_add_iff_right h + +lemma add_lt_add_iff_left {k : ℕ∞} (h : k ≠ ⊤) : k + n < k + m ↔ n < m := + WithTop.add_lt_add_iff_left h + protected lemma exists_nat_gt {n : ℕ∞} (hn : n ≠ ⊤) : ∃ m : ℕ, n < m := by lift n to ℕ using hn obtain ⟨m, hm⟩ := exists_gt n @@ -286,6 +295,8 @@ protected lemma le_sub_of_add_le_left (ha : a ≠ ⊤) : a + b ≤ c → b ≤ c protected lemma sub_sub_cancel (h : a ≠ ⊤) (h2 : b ≤ a) : a - (a - b) = b := (addLECancellable_of_ne_top <| ne_top_of_le_ne_top h tsub_le_self).tsub_tsub_cancel_of_le h2 +section withTop_enat + lemma add_one_natCast_le_withTop_of_lt {m : ℕ} {n : WithTop ℕ∞} (h : m < n) : (m + 1 : ℕ) ≤ n := by match n with | ⊤ => exact le_top @@ -294,9 +305,10 @@ lemma add_one_natCast_le_withTop_of_lt {m : ℕ} {n : WithTop ℕ∞} (h : m < n @[simp] lemma coe_top_add_one : ((⊤ : ℕ∞) : WithTop ℕ∞) + 1 = (⊤ : ℕ∞) := rfl -@[simp] lemma add_one_eq_coe_top_iff : n + 1 = (⊤ : ℕ∞) ↔ n = (⊤ : ℕ∞) := by +@[simp] lemma add_one_eq_coe_top_iff {n : WithTop ℕ∞} : n + 1 = (⊤ : ℕ∞) ↔ n = (⊤ : ℕ∞) := by match n with | ⊤ => exact Iff.rfl + | (⊤ : ℕ∞) => simp | (n : ℕ) => norm_cast; simp only [coe_ne_top, iff_false, ne_eq] @[simp] lemma natCast_ne_coe_top (n : ℕ) : (n : WithTop ℕ∞) ≠ (⊤ : ℕ∞) := nofun @@ -308,13 +320,12 @@ lemma one_le_iff_ne_zero_withTop {n : WithTop ℕ∞} : 1 ≤ n ↔ n ≠ 0 := ⟨fun h ↦ (zero_lt_one.trans_le h).ne', fun h ↦ add_one_natCast_le_withTop_of_lt (pos_iff_ne_zero.mpr h)⟩ -lemma add_one_pos : 0 < n + 1 := - succ_def n ▸ Order.bot_lt_succ n +lemma natCast_le_of_coe_top_le_withTop {N : WithTop ℕ∞} (hN : (⊤ : ℕ∞) ≤ N) (n : ℕ) : n ≤ N := + le_trans (mod_cast le_top) hN -lemma add_lt_add_iff_right {k : ℕ∞} (h : k ≠ ⊤) : n + k < m + k ↔ n < m := - WithTop.add_lt_add_iff_right h +lemma natCast_lt_of_coe_top_le_withTop {N : WithTop ℕ∞} (hN : (⊤ : ℕ∞) ≤ N) (n : ℕ) : n < N := + lt_of_lt_of_le (mod_cast lt_add_one n) (natCast_le_of_coe_top_le_withTop hN (n + 1)) -lemma add_lt_add_iff_left {k : ℕ∞} (h : k ≠ ⊤) : k + n < k + m ↔ n < m := - WithTop.add_lt_add_iff_left h +end withTop_enat end ENat From fb23240e8b4409f2f23bcd07aa73c7451f216413 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 13:28:39 +0000 Subject: [PATCH 68/90] feat(CI): add a log of the size of the oleans (#19283) Split off from #16020: it only adds a step that computes the sizes of the oleans. --- .github/build.in.yml | 4 ++++ .github/workflows/bors.yml | 4 ++++ .github/workflows/build.yml | 4 ++++ .github/workflows/build_fork.yml | 4 ++++ 4 files changed, 16 insertions(+) diff --git a/.github/build.in.yml b/.github/build.in.yml index 3afe6166a83f4..cf758b28c47f3 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -156,6 +156,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 598668edd66d3..68c1d03a1f5f2 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -166,6 +166,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 4afc15b2f5939..7edf488cd7923 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -173,6 +173,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 1964290d30bc7..76448210f9cb0 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -170,6 +170,10 @@ jobs: run: | bash -o pipefail -c "env LEAN_ABORT_ON_PANIC=1 lake build --wfail -KCI" + - name: print the sizes of the oleans + run: | + du .lake/build/lib/Mathlib || echo "This code should be unreachable" + - name: upload cache # We only upload the cache if the build started (whether succeeding, failing, or cancelled) # but not if any earlier step failed or was cancelled. From 09242aa0f47b50626ae7552e5695f06fd4068b7a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 20 Nov 2024 14:09:17 +0000 Subject: [PATCH 69/90] feat(KrullDimension): height refactor module docstring (#15524) This concludes a refactoring of and API extension for `height`. Notable changes, spread over 28 prior PRs, include: * Everything put into `namespace Order` * Type of `krullDim` changed from `WithBot (WithTop Nat)` to `WithBot ENat`. * Type of `height` changed from `WithBot (WithTop Nat)` to `ENat`. * Definition of `height` via `LTSeries` rather than krullDim * Plenty of lemmas about `height` * Concrete calculations for `height` and `krullDim` This final change updates the module docstring, and adds two `coheight` lemmas corresponding to existing `height` lemmas that slipped through the cracks of juggling multple PRs. From the Carleson project. --- Mathlib/Order/KrullDimension.lean | 30 +++++++++++++++++++++++++++--- 1 file changed, 27 insertions(+), 3 deletions(-) diff --git a/Mathlib/Order/KrullDimension.lean b/Mathlib/Order/KrullDimension.lean index 249a8d2cc0301..1b435cca65877 100644 --- a/Mathlib/Order/KrullDimension.lean +++ b/Mathlib/Order/KrullDimension.lean @@ -19,16 +19,32 @@ In case that `α` is empty, then its Krull dimension is defined to be negative i length of all series `a₀ < a₁ < ... < aₙ` is unbounded, then its Krull dimension is defined to be positive infinity. -For `a : α`, its height (in `ℕ∞`) is defined to be `sup {n | a₀ < a₁ < ... < aₙ ≤ a}` while its +For `a : α`, its height (in `ℕ∞`) is defined to be `sup {n | a₀ < a₁ < ... < aₙ ≤ a}`, while its coheight is defined to be `sup {n | a ≤ a₀ < a₁ < ... < aₙ}` . ## Main results * The Krull dimension is the same as that of the dual order (`krullDim_orderDual`). -* The Krull dimension is the supremum of the heights of the elements (`krullDim_eq_iSup_height`). +* The Krull dimension is the supremum of the heights of the elements (`krullDim_eq_iSup_height`), + or their coheights (`krullDim_eq_iSup_coheight`), or their sums of height and coheight + (`krullDim_eq_iSup_height_add_coheight_of_nonempty`) -* The height is monotone. +* The height in the dual order equals the coheight, and vice versa. + +* The height is monotone (`height_mono`), and strictly monotone if finite (`height_strictMono`). + +* The coheight is antitone (`coheight_anti`), and strictly antitone if finite + (`coheight_strictAnti`). + +* The height is the supremum of the successor of the height of all smaller elements + (`height_eq_iSup_lt_height`). + +* The elements of height zero are the minimal elements (`height_eq_zero`), and the elements of + height `n` are minimal among those of height `≥ n` (`height_eq_coe_iff_minimal_le_height`). + +* Concrete calculations for the height, coheight and Krull dimension in `ℕ`, `ℤ`, `WithTop`, + `WithBot` and `ℕ∞`. ## Design notes @@ -240,6 +256,9 @@ lemma height_mono : Monotone (α := α) height := lemma coheight_anti : Antitone (α := α) coheight := (height_mono (α := αᵒᵈ)).dual_left +@[gcongr] protected lemma _root_.GCongr.coheight_le_coheight (a b : α) (hba : b ≤ a) : + coheight a ≤ coheight b := coheight_anti hba + private lemma height_add_const (a : α) (n : ℕ∞) : height a + n = ⨆ (p : LTSeries α) (_ : p.last = a), p.length + n := by have hne : Nonempty { p : LTSeries α // p.last = a } := ⟨RelSeries.singleton _ a, rfl⟩ @@ -253,6 +272,11 @@ private lemma height_add_const (a : α) (n : ℕ∞) : have := length_le_height_last (p := p.snoc y (by simp [*])) simpa using this +/- For elements of finite height, `coheight` is strictly antitone. -/ +@[gcongr] lemma coheight_strictAnti {x y : α} (hyx : y < x) (hfin : coheight x < ⊤) : + coheight x < coheight y := + height_strictMono (α := αᵒᵈ) hyx hfin + lemma height_le_height_apply_of_strictMono (f : α → β) (hf : StrictMono f) (x : α) : height x ≤ height (f x) := by simp only [height_eq_iSup_last_eq] From edcdc3e1603a3d61b9e9e2cea650beaa663f26bb Mon Sep 17 00:00:00 2001 From: grhkm21 Date: Wed, 20 Nov 2024 14:09:19 +0000 Subject: [PATCH 70/90] feat(UpperHalfPlane): weaken assumptions of multiple atImInfty lemmas (#16015) In this PR, I weaken some assumptions so that they can be applied more widely (more specifically, to my use case :) ) I also renamed some lemmas because they sounded weird. --- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 2 +- .../FunctionsBoundedAtInfty.lean | 23 +++++++++++-------- Mathlib/NumberTheory/ModularForms/Basic.lean | 5 +--- .../EisensteinSeries/IsBoundedAtImInfty.lean | 2 +- .../Order/Filter/ZeroAndBoundedAtFilter.lean | 9 ++++---- 5 files changed, 21 insertions(+), 20 deletions(-) diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 87c9e471f1eac..c9ef3caef6893 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1149,7 +1149,7 @@ theorem isLittleO_const_iff_isLittleO_one {c : F''} (hc : c ≠ 0) : fun h => h.trans_isBigO <| isBigO_const_const _ hc _⟩ @[simp] -theorem isLittleO_one_iff : f' =o[l] (fun _x => 1 : α → F) ↔ Tendsto f' l (𝓝 0) := by +theorem isLittleO_one_iff {f : α → E'''} : f =o[l] (fun _x => 1 : α → F) ↔ Tendsto f l (𝓝 0) := by simp only [isLittleO_iff, norm_one, mul_one, Metric.nhds_basis_closedBall.tendsto_right_iff, Metric.mem_closedBall, dist_zero_right] diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean index 7ecb4b7001f40..b784bb4935538 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/FunctionsBoundedAtInfty.lean @@ -53,19 +53,22 @@ def zeroAtImInftySubmodule (α : Type*) [NormedField α] : Submodule α (ℍ → def boundedAtImInftySubalgebra (α : Type*) [NormedField α] : Subalgebra α (ℍ → α) := boundedFilterSubalgebra _ atImInfty -nonrec theorem IsBoundedAtImInfty.mul {f g : ℍ → ℂ} (hf : IsBoundedAtImInfty f) - (hg : IsBoundedAtImInfty g) : IsBoundedAtImInfty (f * g) := by - simpa only [Pi.one_apply, mul_one, norm_eq_abs] using hf.mul hg - -theorem bounded_mem (f : ℍ → ℂ) : - IsBoundedAtImInfty f ↔ ∃ M A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ M := by +theorem isBoundedAtImInfty_iff {α : Type*} [Norm α] {f : ℍ → α} : + IsBoundedAtImInfty f ↔ ∃ M A : ℝ, ∀ z : ℍ, A ≤ im z → ‖f z‖ ≤ M := by simp [IsBoundedAtImInfty, BoundedAtFilter, Asymptotics.isBigO_iff, Filter.Eventually, atImInfty_mem] -theorem zero_at_im_infty (f : ℍ → ℂ) : - IsZeroAtImInfty f ↔ ∀ ε : ℝ, 0 < ε → ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ ε := - (atImInfty_basis.tendsto_iff Metric.nhds_basis_closedBall).trans <| by - simp only [true_and, mem_closedBall_zero_iff]; rfl +@[deprecated (since := "2024-08-27")] alias _root_.bounded_mem := isBoundedAtImInfty_iff + +theorem isZeroAtImInfty_iff {α : Type*} [SeminormedAddGroup α] {f : ℍ → α} : + IsZeroAtImInfty f ↔ ∀ ε : ℝ, 0 < ε → ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → ‖f z‖ ≤ ε := + (atImInfty_basis.tendsto_iff Metric.nhds_basis_closedBall).trans <| by simp + +@[deprecated (since := "2024-08-27")] alias _root_.zero_at_im_infty := isZeroAtImInfty_iff + +theorem IsZeroAtImInfty.isBoundedAtImInfty {α : Type*} [SeminormedAddGroup α] {f : ℍ → α} + (hf : IsZeroAtImInfty f) : IsBoundedAtImInfty f := + hf.boundedAtFilter lemma tendsto_comap_im_ofComplex : Tendsto ofComplex (comap Complex.im atTop) atImInfty := by diff --git a/Mathlib/NumberTheory/ModularForms/Basic.lean b/Mathlib/NumberTheory/ModularForms/Basic.lean index d82d0134a4d49..2a7ff22c6f898 100644 --- a/Mathlib/NumberTheory/ModularForms/Basic.lean +++ b/Mathlib/NumberTheory/ModularForms/Basic.lean @@ -227,10 +227,7 @@ def mul {k_1 k_2 : ℤ} {Γ : Subgroup SL(2, ℤ)} (f : ModularForm Γ k_1) (g : toSlashInvariantForm := f.1.mul g.1 holo' := f.holo'.mul g.holo' bdd_at_infty' A := by - -- Porting note: was `by simpa using ...` - -- `mul_slash_SL2` is no longer a `simp` and `simpa only [mul_slash_SL2] using ...` failed - rw [SlashInvariantForm.coe_mul, mul_slash_SL2] - exact (f.bdd_at_infty' A).mul (g.bdd_at_infty' A) + simpa only [coe_mul, mul_slash_SL2] using (f.bdd_at_infty' A).mul (g.bdd_at_infty' A) @[simp] theorem mul_coe {k_1 k_2 : ℤ} {Γ : Subgroup SL(2, ℤ)} (f : ModularForm Γ k_1) diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean index d5fe5f9673fc0..eb74346576451 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/IsBoundedAtImInfty.lean @@ -52,7 +52,7 @@ lemma abs_le_tsum_abs (N : ℕ) (a : Fin 2 → ZMod N) (k : ℤ) (hk : 3 ≤ k) /-- Eisenstein series are bounded at infinity. -/ theorem isBoundedAtImInfty_eisensteinSeries_SIF {N : ℕ+} (a : Fin 2 → ZMod N) {k : ℤ} (hk : 3 ≤ k) (A : SL(2, ℤ)) : IsBoundedAtImInfty ((eisensteinSeries_SIF a k).toFun ∣[k] A) := by - simp_rw [UpperHalfPlane.bounded_mem, eisensteinSeries_SIF] at * + simp_rw [UpperHalfPlane.isBoundedAtImInfty_iff, eisensteinSeries_SIF] at * refine ⟨∑'(x : Fin 2 → ℤ), r ⟨⟨N, 2⟩, Nat.ofNat_pos⟩ ^ (-k) * ‖x‖ ^ (-k), 2, ?_⟩ intro z hz obtain ⟨n, hn⟩ := (ModularGroup_T_zpow_mem_verticalStrip z N.2) diff --git a/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean b/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean index c7c14ca53e212..fd33cad02fcd3 100644 --- a/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean +++ b/Mathlib/Order/Filter/ZeroAndBoundedAtFilter.lean @@ -71,15 +71,16 @@ if `f =O[l] 1`. -/ def BoundedAtFilter [Norm β] (l : Filter α) (f : α → β) : Prop := Asymptotics.IsBigO l f (1 : α → ℝ) -theorem ZeroAtFilter.boundedAtFilter [NormedAddCommGroup β] {l : Filter α} {f : α → β} - (hf : ZeroAtFilter l f) : BoundedAtFilter l f := by - rw [ZeroAtFilter, ← Asymptotics.isLittleO_const_iff (one_ne_zero' ℝ)] at hf - exact hf.isBigO +theorem ZeroAtFilter.boundedAtFilter [SeminormedAddGroup β] {l : Filter α} {f : α → β} + (hf : ZeroAtFilter l f) : BoundedAtFilter l f := + ((Asymptotics.isLittleO_one_iff _).mpr hf).isBigO theorem const_boundedAtFilter [Norm β] (l : Filter α) (c : β) : BoundedAtFilter l (Function.const α c : α → β) := Asymptotics.isBigO_const_const c one_ne_zero l +-- TODO(https://github.com/leanprover-community/mathlib4/issues/19288): Remove all Comm in the next +-- three lemmas. This would require modifying the corresponding general asymptotics lemma. nonrec theorem BoundedAtFilter.add [SeminormedAddCommGroup β] {l : Filter α} {f g : α → β} (hf : BoundedAtFilter l f) (hg : BoundedAtFilter l g) : BoundedAtFilter l (f + g) := by simpa using hf.add hg From f02eb6153c95cc977fdc13456fa9da71348e936b Mon Sep 17 00:00:00 2001 From: FR Date: Wed, 20 Nov 2024 14:09:20 +0000 Subject: [PATCH 71/90] chore: deprecate `Quotient.out'` (#17941) --- Mathlib/Data/Multiset/Basic.lean | 2 +- Mathlib/Data/Quot.lean | 10 ++-- Mathlib/Data/Setoid/Basic.lean | 4 +- Mathlib/Data/Setoid/Partition.lean | 10 ++-- Mathlib/GroupTheory/Complement.lean | 50 ++++++++--------- Mathlib/GroupTheory/Coset/Basic.lean | 17 +++--- Mathlib/GroupTheory/Coset/Defs.lean | 15 +++--- Mathlib/GroupTheory/DoubleCoset.lean | 20 ++++--- Mathlib/GroupTheory/GroupAction/Defs.lean | 8 +-- Mathlib/GroupTheory/GroupAction/Quotient.lean | 53 +++++++++++-------- Mathlib/GroupTheory/Transfer.lean | 14 ++--- .../LinearAlgebra/Projectivization/Basic.lean | 4 +- Mathlib/MeasureTheory/Function/AEEqFun.lean | 2 +- .../MeasureTheory/Measure/Haar/Quotient.lean | 2 +- .../NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/Order/Antisymmetrization.lean | 2 +- Mathlib/Order/RelIso/Basic.lean | 6 ++- .../HomogeneousLocalization.lean | 12 ++--- Mathlib/RingTheory/Ideal/Quotient/Basic.lean | 4 +- Mathlib/RingTheory/Perfection.lean | 4 +- Mathlib/Topology/Algebra/ClosedSubgroup.lean | 4 +- Mathlib/Topology/UniformSpace/Separation.lean | 2 +- 22 files changed, 136 insertions(+), 111 deletions(-) diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index cebed32a1d12a..a594866b4536d 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -412,7 +412,7 @@ section ToList /-- Produces a list of the elements in the multiset using choice. -/ noncomputable def toList (s : Multiset α) := - s.out' + s.out @[simp, norm_cast] theorem coe_toList (s : Multiset α) : (s.toList : Multiset α) = s := diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 253f9ac3dded9..9508faa6a1ad6 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -727,16 +727,12 @@ protected theorem eq' {s₁ : Setoid α} {a b : α} : protected theorem eq'' {a b : α} : @Quotient.mk'' α s₁ a = Quotient.mk'' b ↔ s₁ a b := Quotient.eq -/-- A version of `Quotient.out` taking `{s₁ : Setoid α}` as an implicit argument instead of an -instance argument. -/ -noncomputable def out' (a : Quotient s₁) : α := - Quotient.out a +@[deprecated (since := "2024-10-19")] alias out' := out -@[simp] -theorem out_eq' (q : Quotient s₁) : Quotient.mk'' q.out' = q := +theorem out_eq' (q : Quotient s₁) : Quotient.mk'' q.out = q := q.out_eq -theorem mk_out' (a : α) : s₁ (Quotient.mk'' a : Quotient s₁).out' a := +theorem mk_out' (a : α) : s₁ (Quotient.mk'' a : Quotient s₁).out a := Quotient.exact (Quotient.out_eq _) section diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 5774b37a49f86..2e75e96025a44 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -101,6 +101,8 @@ theorem ker_mk_eq (r : Setoid α) : ker (@Quotient.mk'' _ r) = r := theorem ker_apply_mk_out {f : α → β} (a : α) : f (⟦a⟧ : Quotient (Setoid.ker f)).out = f a := @Quotient.mk_out _ (Setoid.ker f) a +set_option linter.deprecated false in +@[deprecated ker_apply_mk_out (since := "2024-10-19")] theorem ker_apply_mk_out' {f : α → β} (a : α) : f (Quotient.mk _ a : Quotient <| Setoid.ker f).out' = f a := @Quotient.mk_out' _ (Setoid.ker f) a @@ -146,7 +148,7 @@ equivalence relations. -/ @[simps] noncomputable def piQuotientEquiv {ι : Sort*} {α : ι → Sort*} (r : ∀ i, Setoid (α i)) : (∀ i, Quotient (r i)) ≃ Quotient (@piSetoid _ _ r) where - toFun := fun x ↦ Quotient.mk'' fun i ↦ (x i).out' + toFun := fun x ↦ Quotient.mk'' fun i ↦ (x i).out invFun := fun q ↦ Quotient.liftOn' q (fun x i ↦ Quotient.mk'' (x i)) fun x y hxy ↦ by ext i simpa using hxy i diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 92f2964a06ba7..1aa847fee7ff8 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -405,7 +405,7 @@ theorem equivQuotient_index : hs.equivQuotient ∘ hs.index = hs.proj := funext hs.equivQuotient_index_apply /-- A map choosing a representative for each element of the quotient associated to an indexed -partition. This is a computable version of `Quotient.out'` using `IndexedPartition.some`. -/ +partition. This is a computable version of `Quotient.out` using `IndexedPartition.some`. -/ def out : hs.Quotient ↪ α := hs.equivQuotient.symm.toEmbedding.trans ⟨hs.some, Function.LeftInverse.injective hs.index_some⟩ @@ -414,9 +414,11 @@ def out : hs.Quotient ↪ α := theorem out_proj (x : α) : hs.out (hs.proj x) = hs.some (hs.index x) := rfl -/-- The indices of `Quotient.out'` and `IndexedPartition.out` are equal. -/ -theorem index_out' (x : hs.Quotient) : hs.index x.out' = hs.index (hs.out x) := - Quotient.inductionOn' x fun x => (Setoid.ker_apply_mk_out' x).trans (hs.index_some _).symm +/-- The indices of `Quotient.out` and `IndexedPartition.out` are equal. -/ +theorem index_out (x : hs.Quotient) : hs.index x.out = hs.index (hs.out x) := + Quotient.inductionOn' x fun x => (Setoid.ker_apply_mk_out x).trans (hs.index_some _).symm + +@[deprecated (since := "2024-10-19")] alias index_out' := index_out /-- This lemma is analogous to `Quotient.out_eq'`. -/ @[simp] diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 5884f8c7afa96..dc8d224d031ad 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -257,11 +257,11 @@ lemma exists_left_transversal (H : Subgroup G) (g : G) : ∃ S ∈ leftTransversals (H : Set G), g ∈ S := by classical refine - ⟨Set.range (Function.update Quotient.out' _ g), range_mem_leftTransversals fun q => ?_, - Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out'⟩ + ⟨Set.range (Function.update Quotient.out _ g), range_mem_leftTransversals fun q => ?_, + Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out⟩ by_cases hq : q = Quotient.mk'' g - · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out') - · refine (Function.update_noteq ?_ g Quotient.out') ▸ q.out_eq' + · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out) + · refine (Function.update_noteq ?_ g Quotient.out) ▸ q.out_eq' exact hq @[to_additive] @@ -269,11 +269,11 @@ lemma exists_right_transversal (H : Subgroup G) (g : G) : ∃ S ∈ rightTransversals (H : Set G), g ∈ S := by classical refine - ⟨Set.range (Function.update Quotient.out' _ g), range_mem_rightTransversals fun q => ?_, - Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out'⟩ + ⟨Set.range (Function.update Quotient.out _ g), range_mem_rightTransversals fun q => ?_, + Quotient.mk'' g, Function.update_same (Quotient.mk'' g) g Quotient.out⟩ by_cases hq : q = Quotient.mk'' g - · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out') - · exact Eq.trans (congr_arg _ (Function.update_noteq hq g Quotient.out')) q.out_eq' + · exact hq.symm ▸ congr_arg _ (Function.update_same (Quotient.mk'' g) g Quotient.out) + · exact Eq.trans (congr_arg _ (Function.update_noteq hq g Quotient.out)) q.out_eq' /-- Given two subgroups `H' ⊆ H`, there exists a left transversal to `H'` inside `H`. -/ @[to_additive "Given two subgroups `H' ⊆ H`, there exists a transversal to `H'` inside `H`"] @@ -587,11 +587,11 @@ end Action @[to_additive] instance : Inhabited (leftTransversals (H : Set G)) := - ⟨⟨Set.range Quotient.out', range_mem_leftTransversals Quotient.out_eq'⟩⟩ + ⟨⟨Set.range Quotient.out, range_mem_leftTransversals Quotient.out_eq'⟩⟩ @[to_additive] instance : Inhabited (rightTransversals (H : Set G)) := - ⟨⟨Set.range Quotient.out', range_mem_rightTransversals Quotient.out_eq'⟩⟩ + ⟨⟨Set.range Quotient.out, range_mem_rightTransversals Quotient.out_eq'⟩⟩ theorem IsComplement'.isCompl (h : IsComplement' H K) : IsCompl H K := by refine @@ -666,17 +666,17 @@ variable {G : Type u} [Group G] (H : Subgroup G) (g : G) /-- Partition `G ⧸ H` into orbits of the action of `g : G`. -/ noncomputable def quotientEquivSigmaZMod : - G ⧸ H ≃ Σq : orbitRel.Quotient (zpowers g) (G ⧸ H), ZMod (minimalPeriod (g • ·) q.out') := + G ⧸ H ≃ Σq : orbitRel.Quotient (zpowers g) (G ⧸ H), ZMod (minimalPeriod (g • ·) q.out) := (selfEquivSigmaOrbits (zpowers g) (G ⧸ H)).trans - (sigmaCongrRight fun q => orbitZPowersEquiv g q.out') + (sigmaCongrRight fun q => orbitZPowersEquiv g q.out) theorem quotientEquivSigmaZMod_symm_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out')) : - (quotientEquivSigmaZMod H g).symm ⟨q, k⟩ = g ^ (cast k : ℤ) • q.out' := + (k : ZMod (minimalPeriod (g • ·) q.out)) : + (quotientEquivSigmaZMod H g).symm ⟨q, k⟩ = g ^ (cast k : ℤ) • q.out := rfl theorem quotientEquivSigmaZMod_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) (k : ℤ) : - quotientEquivSigmaZMod H g (g ^ k • q.out') = ⟨q, k⟩ := by + quotientEquivSigmaZMod H g (g ^ k • q.out) = ⟨q, k⟩ := by rw [apply_eq_iff_eq_symm_apply, quotientEquivSigmaZMod_symm_apply, ZMod.coe_intCast, zpow_smul_mod_minimalPeriod] @@ -684,16 +684,16 @@ theorem quotientEquivSigmaZMod_apply (q : orbitRel.Quotient (zpowers g) (G ⧸ H in `G ⧸ H`, an element `g ^ k • q₀` is mapped to `g ^ k • g₀` for a fixed choice of representative `g₀` of `q₀`. -/ noncomputable def transferFunction : G ⧸ H → G := fun q => - g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * (quotientEquivSigmaZMod H g q).1.out'.out' + g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * (quotientEquivSigmaZMod H g q).1.out.out theorem transferFunction_apply (q : G ⧸ H) : transferFunction H g q = g ^ (cast (quotientEquivSigmaZMod H g q).2 : ℤ) * - (quotientEquivSigmaZMod H g q).1.out'.out' := + (quotientEquivSigmaZMod H g q).1.out.out := rfl theorem coe_transferFunction (q : G ⧸ H) : ↑(transferFunction H g q) = q := by - rw [transferFunction_apply, ← smul_eq_mul, Quotient.coe_smul_out', + rw [transferFunction_apply, ← smul_eq_mul, Quotient.coe_smul_out, ← quotientEquivSigmaZMod_symm_apply, Sigma.eta, symm_apply_apply] /-- The transfer transversal as a set. Contains elements of the form `g ^ k • g₀` for fixed choices @@ -714,17 +714,17 @@ theorem transferTransversal_apply (q : G ⧸ H) : toEquiv_apply (coe_transferFunction H g) q theorem transferTransversal_apply' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out')) : - ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out')) = - g ^ (cast k : ℤ) * q.out'.out' := by + (k : ZMod (minimalPeriod (g • ·) q.out)) : + ↑(toEquiv (transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + g ^ (cast k : ℤ) * q.out.out := by rw [transferTransversal_apply, transferFunction_apply, ← quotientEquivSigmaZMod_symm_apply, apply_symm_apply] theorem transferTransversal_apply'' (q : orbitRel.Quotient (zpowers g) (G ⧸ H)) - (k : ZMod (minimalPeriod (g • ·) q.out')) : - ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out')) = - if k = 0 then g ^ minimalPeriod (g • ·) q.out' * q.out'.out' - else g ^ (cast k : ℤ) * q.out'.out' := by + (k : ZMod (minimalPeriod (g • ·) q.out)) : + ↑(toEquiv (g • transferTransversal H g).2 (g ^ (cast k : ℤ) • q.out)) = + if k = 0 then g ^ minimalPeriod (g • ·) q.out * q.out.out + else g ^ (cast k : ℤ) * q.out.out := by rw [smul_apply_eq_smul_apply_inv_smul, transferTransversal_apply, transferFunction_apply, ← mul_smul, ← zpow_neg_one, ← zpow_add, quotientEquivSigmaZMod_apply, smul_eq_mul, ← mul_assoc, ← zpow_one_add, Int.cast_add, Int.cast_neg, Int.cast_one, intCast_cast, cast_id', id, ← diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index 5caed356835d7..a93c6f92a2109 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -314,10 +314,13 @@ lemma orbit_mk_eq_smul (x : α) : MulAction.orbitRel.Quotient.orbit (x : α ⧸ simpa [mem_smul_set_iff_inv_smul_mem, ← leftRel_apply, Quotient.eq''] using Setoid.comm' _ @[to_additive] -lemma orbit_eq_out'_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out' • s := by +lemma orbit_eq_out_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x.out • s := by induction x using QuotientGroup.induction_on simp only [orbit_mk_eq_smul, ← eq_class_eq_leftCoset, Quotient.out_eq'] +@[to_additive (attr := deprecated (since := "2024-10-19"))] +alias orbit_eq_out'_smul := orbit_eq_out_smul + end QuotientGroup namespace Subgroup @@ -344,7 +347,7 @@ def rightCosetEquivSubgroup (g : α) : (op g • s : Set α) ≃ s := noncomputable def groupEquivQuotientProdSubgroup : α ≃ (α ⧸ s) × s := calc α ≃ ΣL : α ⧸ s, { x : α // (x : α ⧸ s) = L } := (Equiv.sigmaFiberEquiv QuotientGroup.mk).symm - _ ≃ ΣL : α ⧸ s, (Quotient.out' L • s : Set α) := + _ ≃ ΣL : α ⧸ s, (Quotient.out L • s : Set α) := Equiv.sigmaCongrRight fun L => by rw [← eq_class_eq_leftCoset] show @@ -394,7 +397,7 @@ The constructive version is `quotientEquivProdOfLE'`. -/ @[to_additive (attr := simps!) "If `H ≤ K`, then `G/H ≃ G/K × K/H` nonconstructively. The constructive version is `quotientEquivProdOfLE'`."] noncomputable def quotientEquivProdOfLE (h_le : s ≤ t) : α ⧸ s ≃ (α ⧸ t) × t ⧸ s.subgroupOf t := - quotientEquivProdOfLE' h_le Quotient.out' Quotient.out_eq' + quotientEquivProdOfLE' h_le Quotient.out Quotient.out_eq' /-- If `s ≤ t`, then there is an embedding `s ⧸ H.subgroupOf s ↪ t ⧸ H.subgroupOf t`. -/ @[to_additive "If `s ≤ t`, then there is an embedding @@ -543,11 +546,11 @@ variable [Group α] noncomputable def preimageMkEquivSubgroupProdSet (s : Subgroup α) (t : Set (α ⧸ s)) : QuotientGroup.mk ⁻¹' t ≃ s × t where toFun a := - ⟨⟨((Quotient.out' (QuotientGroup.mk a)) : α)⁻¹ * a, + ⟨⟨((Quotient.out (QuotientGroup.mk a)) : α)⁻¹ * a, leftRel_apply.mp (@Quotient.exact' _ (leftRel s) _ _ <| Quotient.out_eq' _)⟩, ⟨QuotientGroup.mk a, a.2⟩⟩ invFun a := - ⟨Quotient.out' a.2.1 * a.1.1, + ⟨Quotient.out a.2.1 * a.1.1, show QuotientGroup.mk _ ∈ t by rw [mk_mul_of_mem _ a.1.2, out_eq'] exact a.2.2⟩ @@ -559,8 +562,8 @@ open MulAction in @[to_additive "An additive group is made up of a disjoint union of cosets of an additive subgroup."] lemma univ_eq_iUnion_smul (H : Subgroup α) : - (Set.univ (α := α)) = ⋃ x : α ⧸ H, x.out' • (H : Set _) := by - simp_rw [univ_eq_iUnion_orbit H.op, orbit_eq_out'_smul] + (Set.univ (α := α)) = ⋃ x : α ⧸ H, x.out • (H : Set _) := by + simp_rw [univ_eq_iUnion_orbit H.op, orbit_eq_out_smul] rfl end QuotientGroup diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index f70f9cbbba8eb..f43753df9ac06 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -197,17 +197,20 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := @[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq @[to_additive] -theorem out_eq' (a : α ⧸ s) : mk a.out' = a := +theorem out_eq' (a : α ⧸ s) : mk a.out = a := Quotient.out_eq' a variable (s) -/- It can be useful to write `obtain ⟨h, H⟩ := mk_out'_eq_mul ...`, and then `rw [H]` or +/- It can be useful to write `obtain ⟨h, H⟩ := mk_out_eq_mul ...`, and then `rw [H]` or `simp_rw [H]` or `simp only [H]`. In order for `simp_rw` and `simp only` to work, this lemma is - stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out'`. -/ -@[to_additive QuotientAddGroup.mk_out'_eq_mul] -theorem mk_out'_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out' = g * h := - ⟨⟨g⁻¹ * (mk g).out', QuotientGroup.eq.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ + stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out`. -/ +@[to_additive QuotientAddGroup.mk_out_eq_mul] +theorem mk_out_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out = g * h := + ⟨⟨g⁻¹ * (mk g).out, QuotientGroup.eq.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ + +@[to_additive (attr := deprecated (since := "2024-10-19")) QuotientAddGroup.mk_out'_eq_mul] +alias mk_out'_eq_mul := mk_out_eq_mul variable {s} {a b : α} diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index d80688cff9e27..5efccbbe6d6e7 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -101,7 +101,7 @@ theorem rel_bot_eq_right_group_rel (H : Subgroup G) : /-- Create a doset out of an element of `H \ G / K`-/ def quotToDoset (H K : Subgroup G) (q : Quotient (H : Set G) K) : Set G := - doset q.out' H K + doset q.out H K /-- Map from `G` to `H \ G / K`-/ abbrev mk (H K : Subgroup G) (a : G) : Quotient (H : Set G) K := @@ -115,34 +115,38 @@ theorem eq (H K : Subgroup G) (a b : G) : rw [Quotient.eq''] apply rel_iff -theorem out_eq' (H K : Subgroup G) (q : Quotient ↑H ↑K) : mk H K q.out' = q := +theorem out_eq' (H K : Subgroup G) (q : Quotient ↑H ↑K) : mk H K q.out = q := Quotient.out_eq' q -theorem mk_out'_eq_mul (H K : Subgroup G) (g : G) : - ∃ h k : G, h ∈ H ∧ k ∈ K ∧ (mk H K g : Quotient ↑H ↑K).out' = h * g * k := by - have := eq H K (mk H K g : Quotient ↑H ↑K).out' g +theorem mk_out_eq_mul (H K : Subgroup G) (g : G) : + ∃ h k : G, h ∈ H ∧ k ∈ K ∧ (mk H K g : Quotient ↑H ↑K).out = h * g * k := by + have := eq H K (mk H K g : Quotient ↑H ↑K).out g rw [out_eq'] at this obtain ⟨h, h_h, k, hk, T⟩ := this.1 rfl refine ⟨h⁻¹, k⁻¹, H.inv_mem h_h, K.inv_mem hk, eq_mul_inv_of_mul_eq (eq_inv_mul_of_mul_eq ?_)⟩ rw [← mul_assoc, ← T] +@[deprecated (since := "2024-10-19")] alias mk_out'_eq_mul := mk_out_eq_mul + theorem mk_eq_of_doset_eq {H K : Subgroup G} {a b : G} (h : doset a H K = doset b H K) : mk H K a = mk H K b := by rw [eq] exact mem_doset.mp (h.symm ▸ mem_doset_self H K b) -theorem disjoint_out' {H K : Subgroup G} {a b : Quotient H K} : - a ≠ b → Disjoint (doset a.out' H K) (doset b.out' (H : Set G) K) := by +theorem disjoint_out {H K : Subgroup G} {a b : Quotient H K} : + a ≠ b → Disjoint (doset a.out H K) (doset b.out (H : Set G) K) := by contrapose! intro h simpa [out_eq'] using mk_eq_of_doset_eq (eq_of_not_disjoint h) +@[deprecated (since := "2024-10-19")] alias disjoint_out' := disjoint_out + theorem union_quotToDoset (H K : Subgroup G) : ⋃ q, quotToDoset H K q = Set.univ := by ext x simp only [Set.mem_iUnion, quotToDoset, mem_doset, SetLike.mem_coe, exists_prop, Set.mem_univ, iff_true] use mk H K x - obtain ⟨h, k, h3, h4, h5⟩ := mk_out'_eq_mul H K x + obtain ⟨h, k, h3, h4, h5⟩ := mk_out_eq_mul H K x refine ⟨h⁻¹, H.inv_mem h3, k⁻¹, K.inv_mem h4, ?_⟩ simp only [h5, Subgroup.coe_mk, ← mul_assoc, one_mul, inv_mul_cancel, mul_inv_cancel_right] diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index e2621efee5a77..3848050964715 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -431,7 +431,7 @@ nonrec lemma orbitRel.Quotient.orbit_nonempty (x : orbitRel.Quotient G α) : nonrec lemma orbitRel.Quotient.mapsTo_smul_orbit (g : G) (x : orbitRel.Quotient G α) : Set.MapsTo (g • ·) x.orbit x.orbit := by rw [orbitRel.Quotient.orbit_eq_orbit_out x Quotient.out_eq'] - exact mapsTo_smul_orbit g x.out' + exact mapsTo_smul_orbit g x.out @[to_additive] instance (x : orbitRel.Quotient G α) : MulAction G x.orbit where @@ -485,12 +485,12 @@ local notation "Ω" => orbitRel.Quotient G α /-- Decomposition of a type `X` as a disjoint union of its orbits under a group action. This version is expressed in terms of `MulAction.orbitRel.Quotient.orbit` instead of -`MulAction.orbit`, to avoid mentioning `Quotient.out'`. -/ +`MulAction.orbit`, to avoid mentioning `Quotient.out`. -/ @[to_additive "Decomposition of a type `X` as a disjoint union of its orbits under an additive group action. This version is expressed in terms of `AddAction.orbitRel.Quotient.orbit` instead of - `AddAction.orbit`, to avoid mentioning `Quotient.out'`. "] + `AddAction.orbit`, to avoid mentioning `Quotient.out`. "] def selfEquivSigmaOrbits' : α ≃ Σω : Ω, ω.orbit := letI := orbitRel G α calc @@ -503,7 +503,7 @@ def selfEquivSigmaOrbits' : α ≃ Σω : Ω, ω.orbit := @[to_additive "Decomposition of a type `X` as a disjoint union of its orbits under an additive group action."] -def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out' := +def selfEquivSigmaOrbits : α ≃ Σω : Ω, orbit G ω.out := (selfEquivSigmaOrbits' G α).trans <| Equiv.sigmaCongrRight fun _ => Equiv.Set.ofEq <| orbitRel.Quotient.orbit_eq_orbit_out _ Quotient.out_eq' diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index af70bfbcd8453..4c15d5da76105 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -90,19 +90,30 @@ theorem Quotient.smul_coe [QuotientAction β H] (b : β) (a : α) : rfl @[to_additive (attr := simp)] -theorem Quotient.mk_smul_out' [QuotientAction β H] (b : β) (q : α ⧸ H) : - QuotientGroup.mk (b • q.out') = b • q := by rw [← Quotient.smul_mk, QuotientGroup.out_eq'] +theorem Quotient.mk_smul_out [QuotientAction β H] (b : β) (q : α ⧸ H) : + QuotientGroup.mk (b • q.out) = b • q := by rw [← Quotient.smul_mk, QuotientGroup.out_eq'] -- Porting note: removed simp attribute, simp can prove this @[to_additive] -theorem Quotient.coe_smul_out' [QuotientAction β H] (b : β) (q : α ⧸ H) : ↑(b • q.out') = b • q := - Quotient.mk_smul_out' H b q +theorem Quotient.coe_smul_out [QuotientAction β H] (b : β) (q : α ⧸ H) : ↑(b • q.out) = b • q := + Quotient.mk_smul_out H b q -theorem _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ H) : - q.out'⁻¹ * a ^ Function.minimalPeriod (a • ·) q * q.out' ∈ H := by - rw [mul_assoc, ← QuotientGroup.eq, QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out', +theorem _root_.QuotientGroup.out_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ H) : + q.out⁻¹ * a ^ Function.minimalPeriod (a • ·) q * q.out ∈ H := by + rw [mul_assoc, ← QuotientGroup.eq, QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out, eq_comm, pow_smul_eq_iff_minimalPeriod_dvd] +@[to_additive (attr := deprecated (since := "2024-10-19"))] +alias Quotient.mk_smul_out' := Quotient.mk_smul_out + +-- Porting note: removed simp attribute, simp can prove this +@[to_additive (attr := deprecated (since := "2024-10-19"))] +alias Quotient.coe_smul_out' := Quotient.coe_smul_out + +@[deprecated (since := "2024-10-19")] +alias _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem := + QuotientGroup.out_conj_pow_minimalPeriod_mem + end QuotientAction open QuotientGroup @@ -192,14 +203,14 @@ local notation "Ω" => Quotient <| orbitRel α β /-- **Class formula** : given `G` a group acting on `X` and `φ` a function mapping each orbit of `X` under this action (that is, each element of the quotient of `X` by the relation `orbitRel G X`) to an element in this orbit, this gives a (noncomputable) bijection between `X` and the disjoint union -of `G/Stab(φ(ω))` over all orbits `ω`. In most cases you'll want `φ` to be `Quotient.out'`, so we +of `G/Stab(φ(ω))` over all orbits `ω`. In most cases you'll want `φ` to be `Quotient.out`, so we provide `MulAction.selfEquivSigmaOrbitsQuotientStabilizer'` as a special case. -/ @[to_additive "**Class formula** : given `G` an additive group acting on `X` and `φ` a function mapping each orbit of `X` under this action (that is, each element of the quotient of `X` by the relation `orbit_rel G X`) to an element in this orbit, this gives a (noncomputable) bijection between `X` and the disjoint union of `G/Stab(φ(ω))` over all orbits `ω`. In most - cases you'll want `φ` to be `Quotient.out'`, so we provide + cases you'll want `φ` to be `Quotient.out`, so we provide `AddAction.selfEquivSigmaOrbitsQuotientStabilizer'` as a special case. "] noncomputable def selfEquivSigmaOrbitsQuotientStabilizer' {φ : Ω → β} (hφ : LeftInverse Quotient.mk'' φ) : β ≃ Σω : Ω, α ⧸ stabilizer α (φ ω) := @@ -212,11 +223,11 @@ noncomputable def selfEquivSigmaOrbitsQuotientStabilizer' {φ : Ω → β} /-- **Class formula** for a finite group acting on a finite type. See `MulAction.card_eq_sum_card_group_div_card_stabilizer` for a specialized version using -`Quotient.out'`. -/ +`Quotient.out`. -/ @[to_additive "**Class formula** for a finite group acting on a finite type. See `AddAction.card_eq_sum_card_addGroup_div_card_stabilizer` for a specialized version using - `Quotient.out'`."] + `Quotient.out`."] theorem card_eq_sum_card_group_div_card_stabilizer' [Fintype α] [Fintype β] [Fintype Ω] [∀ b : β, Fintype <| stabilizer α b] {φ : Ω → β} (hφ : LeftInverse Quotient.mk'' φ) : Fintype.card β = ∑ ω : Ω, Fintype.card α / Fintype.card (stabilizer α (φ ω)) := by @@ -231,18 +242,18 @@ theorem card_eq_sum_card_group_div_card_stabilizer' [Fintype α] [Fintype β] [F Fintype.card_congr (selfEquivSigmaOrbitsQuotientStabilizer' α β hφ)] /-- **Class formula**. This is a special case of -`MulAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out'`. -/ +`MulAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out`. -/ @[to_additive "**Class formula**. This is a special case of - `AddAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out'`. "] -noncomputable def selfEquivSigmaOrbitsQuotientStabilizer : β ≃ Σω : Ω, α ⧸ stabilizer α ω.out' := + `AddAction.self_equiv_sigma_orbits_quotient_stabilizer'` with `φ = Quotient.out`. "] +noncomputable def selfEquivSigmaOrbitsQuotientStabilizer : β ≃ Σω : Ω, α ⧸ stabilizer α ω.out := selfEquivSigmaOrbitsQuotientStabilizer' α β Quotient.out_eq' /-- **Class formula** for a finite group acting on a finite type. -/ @[to_additive "**Class formula** for a finite group acting on a finite type."] theorem card_eq_sum_card_group_div_card_stabilizer [Fintype α] [Fintype β] [Fintype Ω] [∀ b : β, Fintype <| stabilizer α b] : - Fintype.card β = ∑ ω : Ω, Fintype.card α / Fintype.card (stabilizer α ω.out') := + Fintype.card β = ∑ ω : Ω, Fintype.card α / Fintype.card (stabilizer α ω.out) := card_eq_sum_card_group_div_card_stabilizer' α β Quotient.out_eq' /-- **Burnside's lemma** : a (noncomputable) bijection between the disjoint union of all @@ -259,16 +270,16 @@ noncomputable def sigmaFixedByEquivOrbitsProdGroup : (Σa : α, fixedBy β a) _ ≃ { ba : β × α // ba.2 • ba.1 = ba.1 } := (Equiv.prodComm α β).subtypeEquiv fun _ => Iff.rfl _ ≃ Σb : β, stabilizer α b := Equiv.subtypeProdEquivSigmaSubtype fun (b : β) a => a ∈ stabilizer α b - _ ≃ Σωb : Σω : Ω, orbit α ω.out', stabilizer α (ωb.2 : β) := + _ ≃ Σωb : Σω : Ω, orbit α ω.out, stabilizer α (ωb.2 : β) := (selfEquivSigmaOrbits α β).sigmaCongrLeft' - _ ≃ Σω : Ω, Σb : orbit α ω.out', stabilizer α (b : β) := - Equiv.sigmaAssoc fun (ω : Ω) (b : orbit α ω.out') => stabilizer α (b : β) - _ ≃ Σω : Ω, Σ _ : orbit α ω.out', stabilizer α ω.out' := + _ ≃ Σω : Ω, Σb : orbit α ω.out, stabilizer α (b : β) := + Equiv.sigmaAssoc fun (ω : Ω) (b : orbit α ω.out) => stabilizer α (b : β) + _ ≃ Σω : Ω, Σ _ : orbit α ω.out, stabilizer α ω.out := Equiv.sigmaCongrRight fun _ => Equiv.sigmaCongrRight fun ⟨_, hb⟩ => (stabilizerEquivStabilizerOfOrbitRel hb).toEquiv - _ ≃ Σω : Ω, orbit α ω.out' × stabilizer α ω.out' := + _ ≃ Σω : Ω, orbit α ω.out × stabilizer α ω.out := Equiv.sigmaCongrRight fun _ => Equiv.sigmaEquivProd _ _ - _ ≃ Σ _ : Ω, α := Equiv.sigmaCongrRight fun ω => orbitProdStabilizerEquivGroup α ω.out' + _ ≃ Σ _ : Ω, α := Equiv.sigmaCongrRight fun ω => orbitProdStabilizerEquivGroup α ω.out _ ≃ Ω × α := Equiv.sigmaEquivProd Ω α /-- **Burnside's lemma** : given a finite group `G` acting on a set `X`, the average number of diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index ce1d169cd838e..8b40c9d992fa9 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -103,8 +103,8 @@ theorem transfer_eq_prod_quotient_orbitRel_zpowers_quot [FiniteIndex H] (g : G) transfer ϕ g = ∏ q : Quotient (orbitRel (zpowers g) (G ⧸ H)), ϕ - ⟨q.out'.out'⁻¹ * g ^ Function.minimalPeriod (g • ·) q.out' * q.out'.out', - QuotientGroup.out'_conj_pow_minimalPeriod_mem H g q.out'⟩ := by + ⟨q.out.out⁻¹ * g ^ Function.minimalPeriod (g • ·) q.out * q.out.out, + QuotientGroup.out_conj_pow_minimalPeriod_mem H g q.out⟩ := by classical letI := H.fintypeQuotientOfFiniteIndex calc @@ -115,7 +115,7 @@ theorem transfer_eq_prod_quotient_orbitRel_zpowers_quot [FiniteIndex H] (g : G) refine Fintype.prod_congr _ _ (fun q => ?_) simp only [quotientEquivSigmaZMod_symm_apply, transferTransversal_apply', transferTransversal_apply''] - rw [Fintype.prod_eq_single (0 : ZMod (Function.minimalPeriod (g • ·) q.out')) _] + rw [Fintype.prod_eq_single (0 : ZMod (Function.minimalPeriod (g • ·) q.out)) _] · simp only [if_pos, ZMod.cast_zero, zpow_zero, one_mul, mul_assoc] · intro k hk simp only [if_neg hk, inv_mul_cancel] @@ -133,11 +133,11 @@ theorem transfer_eq_pow_aux (g : G) replace key : ∀ (k : ℕ) (g₀ : G), g₀⁻¹ * g ^ k * g₀ ∈ H → g ^ k ∈ H := fun k g₀ hk => (congr_arg (· ∈ H) (key k g₀ hk)).mp hk replace key : ∀ q : G ⧸ H, g ^ Function.minimalPeriod (g • ·) q ∈ H := fun q => - key (Function.minimalPeriod (g • ·) q) q.out' - (QuotientGroup.out'_conj_pow_minimalPeriod_mem H g q) + key (Function.minimalPeriod (g • ·) q) q.out + (QuotientGroup.out_conj_pow_minimalPeriod_mem H g q) let f : Quotient (orbitRel (zpowers g) (G ⧸ H)) → zpowers g := fun q => - (⟨g, mem_zpowers g⟩ : zpowers g) ^ Function.minimalPeriod (g • ·) q.out' - have hf : ∀ q, f q ∈ H.subgroupOf (zpowers g) := fun q => key q.out' + (⟨g, mem_zpowers g⟩ : zpowers g) ^ Function.minimalPeriod (g • ·) q.out + have hf : ∀ q, f q ∈ H.subgroupOf (zpowers g) := fun q => key q.out replace key := Subgroup.prod_mem (H.subgroupOf (zpowers g)) fun q (_ : q ∈ Finset.univ) => hf q simpa only [f, minimalPeriod_eq_card, Finset.prod_pow_eq_pow_sum, Fintype.card_sigma, diff --git a/Mathlib/LinearAlgebra/Projectivization/Basic.lean b/Mathlib/LinearAlgebra/Projectivization/Basic.lean index 9a12ad4ccb657..38496cfef13f4 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Basic.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Basic.lean @@ -84,10 +84,10 @@ protected lemma lift_mk {α : Type*} (f : { v : V // v ≠ 0 } → α) /-- Choose a representative of `v : Projectivization K V` in `V`. -/ protected noncomputable def rep (v : ℙ K V) : V := - v.out' + v.out theorem rep_nonzero (v : ℙ K V) : v.rep ≠ 0 := - v.out'.2 + v.out.2 @[simp] theorem mk_rep (v : ℙ K V) : mk K v.rep v.rep_nonzero = v := Quotient.out_eq' _ diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index 86720959dec25..38c86bd0a94b3 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -124,7 +124,7 @@ then we choose that one. -/ def cast (f : α →ₘ[μ] β) : α → β := if h : ∃ (b : β), f = mk (const α b) aestronglyMeasurable_const then const α <| Classical.choose h else - AEStronglyMeasurable.mk _ (Quotient.out' f : { f : α → β // AEStronglyMeasurable f μ }).2 + AEStronglyMeasurable.mk _ (Quotient.out f : { f : α → β // AEStronglyMeasurable f μ }).2 /-- A measurable representative of an `AEEqFun` [f] -/ instance instCoeFun : CoeFun (α →ₘ[μ] β) fun _ => α → β := ⟨cast⟩ diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index e552fa14c3286..a782f7e429aeb 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -134,7 +134,7 @@ lemma MeasureTheory.QuotientMeasureEqMeasurePreimage.mulInvariantMeasure_quotien obtain ⟨x₁, h⟩ := @Quotient.exists_rep _ (QuotientGroup.leftRel Γ) x convert measure_preimage_smul μ x₁ A using 1 · rw [← h, Measure.map_apply (measurable_const_mul _) hA] - simp [← MulAction.Quotient.coe_smul_out', ← Quotient.mk''_eq_mk] + simp [← MulAction.Quotient.coe_smul_out, ← Quotient.mk''_eq_mk] exact smulInvariantMeasure_quotient ν variable [Countable Γ] [IsMulRightInvariant ν] [SigmaFinite ν] diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index c735c10548900..31c16252e0c0d 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -465,7 +465,7 @@ def basisUnitLattice : Basis (Fin (rank K)) ℤ (unitLattice K) := units in `basisModTorsion`. -/ def fundSystem : Fin (rank K) → (𝓞 K)ˣ := -- `:)` prevents the `⧸` decaying to a quotient by `leftRel` when we unfold this later - fun i => Quotient.out' (Additive.toMul (basisModTorsion K i):) + fun i => Quotient.out (Additive.toMul (basisModTorsion K i):) theorem fundSystem_mk (i : Fin (rank K)) : Additive.ofMul (QuotientGroup.mk (fundSystem K i)) = (basisModTorsion K i) := by diff --git a/Mathlib/Order/Antisymmetrization.lean b/Mathlib/Order/Antisymmetrization.lean index d098ac0bfffb1..75caf3593c355 100644 --- a/Mathlib/Order/Antisymmetrization.lean +++ b/Mathlib/Order/Antisymmetrization.lean @@ -88,7 +88,7 @@ def toAntisymmetrization : α → Antisymmetrization α r := /-- Get a representative from the antisymmetrization. -/ noncomputable def ofAntisymmetrization : Antisymmetrization α r → α := - Quotient.out' + Quotient.out instance [Inhabited α] : Inhabited (Antisymmetrization α r) := by unfold Antisymmetrization; infer_instance diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index 721ef0208ee6d..cc0083a0b2676 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -363,13 +363,17 @@ noncomputable def Quotient.outRelEmbedding {_ : Setoid α} {r : α → α → Pr refine @fun x y => Quotient.inductionOn₂ x y fun a b => ?_ apply iff_iff_eq.2 (H _ _ _ _ _ _) <;> apply Quotient.mk_out⟩ +set_option linter.deprecated false in /-- `Quotient.out'` as a relation embedding between the lift of a relation and the relation. -/ -@[simps] +@[deprecated Quotient.outRelEmbedding (since := "2024-10-19"), simps] noncomputable def Quotient.out'RelEmbedding {_ : Setoid α} {r : α → α → Prop} (H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ ≈ a₂ → b₁ ≈ b₂ → r a₁ b₁ = r a₂ b₂) : (fun a b => Quotient.liftOn₂' a b r H) ↪r r := { Quotient.outRelEmbedding H with toFun := Quotient.out' } +attribute [deprecated Quotient.outRelEmbedding_apply (since := "2024-10-19")] + Quotient.out'RelEmbedding_apply + @[simp] theorem acc_lift₂_iff {_ : Setoid α} {r : α → α → Prop} {H : ∀ (a₁ b₁ a₂ b₂ : α), a₁ ≈ a₂ → b₁ ≈ b₂ → r a₁ b₁ = r a₂ b₂} {a} : diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 553951dd2ce18..dcab8f6f958a6 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -483,25 +483,25 @@ variable {𝒜} {x} /-- Numerator of an element in `HomogeneousLocalization x`. -/ def num (f : HomogeneousLocalization 𝒜 x) : A := - (Quotient.out' f).num + (Quotient.out f).num /-- Denominator of an element in `HomogeneousLocalization x`. -/ def den (f : HomogeneousLocalization 𝒜 x) : A := - (Quotient.out' f).den + (Quotient.out f).den /-- For an element in `HomogeneousLocalization x`, degree is the natural number `i` such that `𝒜 i` contains both numerator and denominator. -/ def deg (f : HomogeneousLocalization 𝒜 x) : ι := - (Quotient.out' f).deg + (Quotient.out f).deg theorem den_mem (f : HomogeneousLocalization 𝒜 x) : f.den ∈ x := - (Quotient.out' f).den_mem + (Quotient.out f).den_mem theorem num_mem_deg (f : HomogeneousLocalization 𝒜 x) : f.num ∈ 𝒜 f.deg := - (Quotient.out' f).num.2 + (Quotient.out f).num.2 theorem den_mem_deg (f : HomogeneousLocalization 𝒜 x) : f.den ∈ 𝒜 f.deg := - (Quotient.out' f).den.2 + (Quotient.out f).den.2 theorem eq_num_div_den (f : HomogeneousLocalization 𝒜 x) : f.val = Localization.mk f.num ⟨f.den, f.den_mem⟩ := diff --git a/Mathlib/RingTheory/Ideal/Quotient/Basic.lean b/Mathlib/RingTheory/Ideal/Quotient/Basic.lean index 190587bcc84f3..244da7d3aada9 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Basic.lean @@ -191,7 +191,7 @@ noncomputable def piQuotEquiv : ((ι → R) ⧸ I.pi ι) ≃ₗ[R ⧸ I] ι → funext fun i => (Submodule.Quotient.eq' _).2 (QuotientAddGroup.leftRel_apply.mp hab i) map_add' := by rintro ⟨_⟩ ⟨_⟩; rfl map_smul' := by rintro ⟨_⟩ ⟨_⟩; rfl - invFun := fun x ↦ Ideal.Quotient.mk (I.pi ι) fun i ↦ Quotient.out' (x i) + invFun := fun x ↦ Ideal.Quotient.mk (I.pi ι) fun i ↦ Quotient.out (x i) left_inv := by rintro ⟨x⟩ exact Ideal.Quotient.eq.2 fun i => Ideal.Quotient.eq.1 (Quotient.out_eq' _) @@ -216,7 +216,7 @@ end Pi open scoped Pointwise in /-- A ring is made up of a disjoint union of cosets of an ideal. -/ lemma univ_eq_iUnion_image_add {R : Type*} [Ring R] (I : Ideal R) : - (Set.univ (α := R)) = ⋃ x : R ⧸ I, x.out' +ᵥ (I : Set R) := + (Set.univ (α := R)) = ⋃ x : R ⧸ I, x.out +ᵥ (I : Set R) := QuotientAddGroup.univ_eq_iUnion_vadd I.toAddSubgroup lemma _root_.Finite.of_finite_quot_finite_ideal {R : Type*} [Ring R] {I : Ideal R} diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 7db1bdda6d237..7d7e45a3a1929 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -349,13 +349,13 @@ attribute [local instance] Classical.dec /-- For a field `K` with valuation `v : K → ℝ≥0` and ring of integers `O`, a function `O/(p) → ℝ≥0` that sends `0` to `0` and `x + (p)` to `v(x)` as long as `x ∉ (p)`. -/ noncomputable def preVal (x : ModP K v O hv p) : ℝ≥0 := - if x = 0 then 0 else v (algebraMap O K x.out') + if x = 0 then 0 else v (algebraMap O K x.out) variable {K v O hv p} theorem preVal_mk {x : O} (hx : (Ideal.Quotient.mk _ x : ModP K v O hv p) ≠ 0) : preVal K v O hv p (Ideal.Quotient.mk _ x) = v (algebraMap O K x) := by - obtain ⟨r, hr⟩ : ∃ (a : O), a * (p : O) = (Quotient.mk'' x).out' - x := + obtain ⟨r, hr⟩ : ∃ (a : O), a * (p : O) = (Quotient.mk'' x).out - x := Ideal.mem_span_singleton'.1 <| Ideal.Quotient.eq.1 <| Quotient.sound' <| Quotient.mk_out' _ refine (if_neg hx).trans (v.map_eq_of_sub_lt <| lt_of_not_le ?_) erw [← RingHom.map_sub, ← hr, hv.le_iff_dvd] diff --git a/Mathlib/Topology/Algebra/ClosedSubgroup.lean b/Mathlib/Topology/Algebra/ClosedSubgroup.lean index 543b8793ad0fa..f3cf42ab0809f 100644 --- a/Mathlib/Topology/Algebra/ClosedSubgroup.lean +++ b/Mathlib/Topology/Algebra/ClosedSubgroup.lean @@ -102,7 +102,7 @@ lemma isOpen_of_isClosed_of_finiteIndex (H : Subgroup G) [H.FiniteIndex] (h : IsClosed (H : Set G)) : IsOpen (H : Set G) := by apply isClosed_compl_iff.mp convert isClosed_iUnion_of_finite <| fun (x : {x : (G ⧸ H) // x ≠ QuotientGroup.mk 1}) - ↦ IsClosed.smul h (Quotient.out' x.1) + ↦ IsClosed.smul h (Quotient.out x.1) ext x refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · have : QuotientGroup.mk 1 ≠ QuotientGroup.mk (s := H) x := by @@ -110,7 +110,7 @@ lemma isOpen_of_isClosed_of_finiteIndex (H : Subgroup G) [H.FiniteIndex] simpa only [inv_one, one_mul, ne_eq] simp only [ne_eq, Set.mem_iUnion] use ⟨QuotientGroup.mk (s := H) x, this.symm⟩, - (Quotient.out' (QuotientGroup.mk (s := H) x))⁻¹ * x + (Quotient.out (QuotientGroup.mk (s := H) x))⁻¹ * x simp only [SetLike.mem_coe, smul_eq_mul, mul_inv_cancel_left, and_true] exact QuotientGroup.eq.mp <| QuotientGroup.out_eq' (QuotientGroup.mk (s := H) x) · rcases h with ⟨S,⟨y,hS⟩,mem⟩ diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index 331e192d249dc..9b3a324dc67df 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -253,7 +253,7 @@ open Classical in TODO: unify with `SeparationQuotient.lift`. -/ def lift' [T0Space β] (f : α → β) : SeparationQuotient α → β := if hc : UniformContinuous f then lift f fun _ _ h => (h.map hc.continuous).eq - else fun x => f (Nonempty.some ⟨x.out'⟩) + else fun x => f (Nonempty.some ⟨x.out⟩) theorem lift'_mk [T0Space β] {f : α → β} (h : UniformContinuous f) (a : α) : lift' f (mk a) = f a := by rw [lift', dif_pos h, lift_mk] From 185ec9cdb8e7dc5e7ad15854962475f18259c484 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 20 Nov 2024 14:53:32 +0000 Subject: [PATCH 72/90] feat: the Lie bracket of vector fields in vector spaces (#18852) --- Mathlib.lean | 1 + Mathlib/Analysis/Calculus/VectorField.lean | 343 +++++++++++++++++++++ 2 files changed, 344 insertions(+) create mode 100644 Mathlib/Analysis/Calculus/VectorField.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2c342449a9a42..ec30d813df6af 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1172,6 +1172,7 @@ import Mathlib.Analysis.Calculus.SmoothSeries import Mathlib.Analysis.Calculus.TangentCone import Mathlib.Analysis.Calculus.Taylor import Mathlib.Analysis.Calculus.UniformLimitsDeriv +import Mathlib.Analysis.Calculus.VectorField import Mathlib.Analysis.Complex.AbelLimit import Mathlib.Analysis.Complex.AbsMax import Mathlib.Analysis.Complex.Angle diff --git a/Mathlib/Analysis/Calculus/VectorField.lean b/Mathlib/Analysis/Calculus/VectorField.lean new file mode 100644 index 0000000000000..d4e54bc4485f9 --- /dev/null +++ b/Mathlib/Analysis/Calculus/VectorField.lean @@ -0,0 +1,343 @@ +/- +Copyright (c) 2024 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Analysis.Calculus.FDeriv.Symmetric + +/-! +# Vector fields in vector spaces + +We study functions of the form `V : E → E` on a vector space, thinking of these as vector fields. +We define several notions in this context, with the aim to generalize them to vector fields on +manifolds. + +Notably, we define the pullback of a vector field under a map, as +`VectorField.pullback 𝕜 f V x := (fderiv 𝕜 f x).inverse (V (f x))` (together with the same notion +within a set). + +In addition to comprehensive API on this notion, the main result is the following: +* `VectorField.leibniz_identity_lieBracket` is the Leibniz + identity `[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. + +-/ + +open Set +open scoped Topology + +noncomputable section + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] + {V W V₁ W₁ : E → E} {s t : Set E} {x : E} + +/-! +### The Lie bracket of vector fields in a vector space + +We define the Lie bracket of two vector fields, and call it `lieBracket 𝕜 V W x`. We also define +a version localized to sets, `lieBracketWithin 𝕜 V W s x`. We copy the relevant API +of `fderivWithin` and `fderiv` for these notions to get a comprehensive API. +-/ + +namespace VectorField + +variable (𝕜) in +/-- The Lie bracket `[V, W] (x)` of two vector fields at a point, defined as +`DW(x) (V x) - DV(x) (W x)`. -/ +def lieBracket (V W : E → E) (x : E) : E := + fderiv 𝕜 W x (V x) - fderiv 𝕜 V x (W x) + +variable (𝕜) in +/-- The Lie bracket `[V, W] (x)` of two vector fields within a set at a point, defined as +`DW(x) (V x) - DV(x) (W x)` where the derivatives are taken inside `s`. -/ +def lieBracketWithin (V W : E → E) (s : Set E) (x : E) : E := + fderivWithin 𝕜 W s x (V x) - fderivWithin 𝕜 V s x (W x) + +lemma lieBracket_eq : + lieBracket 𝕜 V W = fun x ↦ fderiv 𝕜 W x (V x) - fderiv 𝕜 V x (W x) := rfl + +lemma lieBracketWithin_eq : + lieBracketWithin 𝕜 V W s = + fun x ↦ fderivWithin 𝕜 W s x (V x) - fderivWithin 𝕜 V s x (W x) := rfl + +@[simp] +theorem lieBracketWithin_univ : lieBracketWithin 𝕜 V W univ = lieBracket 𝕜 V W := by + ext1 x + simp [lieBracketWithin, lieBracket] + +lemma lieBracketWithin_eq_zero_of_eq_zero (hV : V x = 0) (hW : W x = 0) : + lieBracketWithin 𝕜 V W s x = 0 := by + simp [lieBracketWithin, hV, hW] + +lemma lieBracket_eq_zero_of_eq_zero (hV : V x = 0) (hW : W x = 0) : + lieBracket 𝕜 V W x = 0 := by + simp [lieBracket, hV, hW] + +lemma lieBracketWithin_smul_left {c : 𝕜} (hV : DifferentiableWithinAt 𝕜 V s x) + (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 (c • V) W s x = + c • lieBracketWithin 𝕜 V W s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add, Pi.smul_apply, map_smul, smul_sub] + rw [fderivWithin_const_smul' hs hV] + rfl + +lemma lieBracket_smul_left {c : 𝕜} (hV : DifferentiableAt 𝕜 V x) : + lieBracket 𝕜 (c • V) W x = c • lieBracket 𝕜 V W x := by + simp only [← differentiableWithinAt_univ, ← lieBracketWithin_univ] at hV ⊢ + exact lieBracketWithin_smul_left hV uniqueDiffWithinAt_univ + +lemma lieBracketWithin_smul_right {c : 𝕜} (hW : DifferentiableWithinAt 𝕜 W s x) + (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 V (c • W) s x = + c • lieBracketWithin 𝕜 V W s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add, Pi.smul_apply, map_smul, smul_sub] + rw [fderivWithin_const_smul' hs hW] + rfl + +lemma lieBracket_smul_right {c : 𝕜} (hW : DifferentiableAt 𝕜 W x) : + lieBracket 𝕜 V (c • W) x = c • lieBracket 𝕜 V W x := by + simp only [← differentiableWithinAt_univ, ← lieBracketWithin_univ] at hW ⊢ + exact lieBracketWithin_smul_right hW uniqueDiffWithinAt_univ + +lemma lieBracketWithin_add_left (hV : DifferentiableWithinAt 𝕜 V s x) + (hV₁ : DifferentiableWithinAt 𝕜 V₁ s x) (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 (V + V₁) W s x = + lieBracketWithin 𝕜 V W s x + lieBracketWithin 𝕜 V₁ W s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add] + rw [fderivWithin_add' hs hV hV₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracket_add_left (hV : DifferentiableAt 𝕜 V x) (hV₁ : DifferentiableAt 𝕜 V₁ x) : + lieBracket 𝕜 (V + V₁) W x = + lieBracket 𝕜 V W x + lieBracket 𝕜 V₁ W x := by + simp only [lieBracket, Pi.add_apply, map_add] + rw [fderiv_add' hV hV₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracketWithin_add_right (hW : DifferentiableWithinAt 𝕜 W s x) + (hW₁ : DifferentiableWithinAt 𝕜 W₁ s x) (hs : UniqueDiffWithinAt 𝕜 s x) : + lieBracketWithin 𝕜 V (W + W₁) s x = + lieBracketWithin 𝕜 V W s x + lieBracketWithin 𝕜 V W₁ s x := by + simp only [lieBracketWithin, Pi.add_apply, map_add] + rw [fderivWithin_add' hs hW hW₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracket_add_right (hW : DifferentiableAt 𝕜 W x) (hW₁ : DifferentiableAt 𝕜 W₁ x) : + lieBracket 𝕜 V (W + W₁) x = + lieBracket 𝕜 V W x + lieBracket 𝕜 V W₁ x := by + simp only [lieBracket, Pi.add_apply, map_add] + rw [fderiv_add' hW hW₁, ContinuousLinearMap.add_apply] + abel + +lemma lieBracketWithin_swap : lieBracketWithin 𝕜 V W s = - lieBracketWithin 𝕜 W V s := by + ext x; simp [lieBracketWithin] + +lemma lieBracket_swap : lieBracket 𝕜 V W x = - lieBracket 𝕜 W V x := by + simp [lieBracket] + +@[simp] lemma lieBracketWithin_self : lieBracketWithin 𝕜 V V s = 0 := by + ext x; simp [lieBracketWithin] + +@[simp] lemma lieBracket_self : lieBracket 𝕜 V V = 0 := by + ext x; simp [lieBracket] + +lemma _root_.ContDiffWithinAt.lieBracketWithin_vectorField + {m n : ℕ∞} (hV : ContDiffWithinAt 𝕜 n V s x) + (hW : ContDiffWithinAt 𝕜 n W s x) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx : x ∈ s) : + ContDiffWithinAt 𝕜 m (lieBracketWithin 𝕜 V W s) s x := by + apply ContDiffWithinAt.sub + · exact ContDiffWithinAt.clm_apply (hW.fderivWithin_right hs hmn hx) + (hV.of_le (le_trans le_self_add hmn)) + · exact ContDiffWithinAt.clm_apply (hV.fderivWithin_right hs hmn hx) + (hW.of_le (le_trans le_self_add hmn)) + +lemma _root_.ContDiffAt.lieBracket_vectorField {m n : ℕ∞} (hV : ContDiffAt 𝕜 n V x) + (hW : ContDiffAt 𝕜 n W x) (hmn : m + 1 ≤ n) : + ContDiffAt 𝕜 m (lieBracket 𝕜 V W) x := by + rw [← contDiffWithinAt_univ] at hV hW ⊢ + simp_rw [← lieBracketWithin_univ] + exact hV.lieBracketWithin_vectorField hW uniqueDiffOn_univ hmn (mem_univ _) + +lemma _root_.ContDiffOn.lieBracketWithin_vectorField {m n : ℕ∞} (hV : ContDiffOn 𝕜 n V s) + (hW : ContDiffOn 𝕜 n W s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : + ContDiffOn 𝕜 m (lieBracketWithin 𝕜 V W s) s := + fun x hx ↦ (hV x hx).lieBracketWithin_vectorField (hW x hx) hs hmn hx + +lemma _root_.ContDiff.lieBracket_vectorField {m n : ℕ∞} (hV : ContDiff 𝕜 n V) + (hW : ContDiff 𝕜 n W) (hmn : m + 1 ≤ n) : + ContDiff 𝕜 m (lieBracket 𝕜 V W) := + contDiff_iff_contDiffAt.2 (fun _ ↦ hV.contDiffAt.lieBracket_vectorField hW.contDiffAt hmn) + +theorem lieBracketWithin_of_mem_nhdsWithin (st : t ∈ 𝓝[s] x) (hs : UniqueDiffWithinAt 𝕜 s x) + (hV : DifferentiableWithinAt 𝕜 V t x) (hW : DifferentiableWithinAt 𝕜 W t x) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := by + simp [lieBracketWithin, fderivWithin_of_mem_nhdsWithin st hs hV, + fderivWithin_of_mem_nhdsWithin st hs hW] + +theorem lieBracketWithin_subset (st : s ⊆ t) (ht : UniqueDiffWithinAt 𝕜 s x) + (hV : DifferentiableWithinAt 𝕜 V t x) (hW : DifferentiableWithinAt 𝕜 W t x) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := + lieBracketWithin_of_mem_nhdsWithin (nhdsWithin_mono _ st self_mem_nhdsWithin) ht hV hW + +theorem lieBracketWithin_inter (ht : t ∈ 𝓝 x) : + lieBracketWithin 𝕜 V W (s ∩ t) x = lieBracketWithin 𝕜 V W s x := by + simp [lieBracketWithin, fderivWithin_inter, ht] + +theorem lieBracketWithin_of_mem_nhds (h : s ∈ 𝓝 x) : + lieBracketWithin 𝕜 V W s x = lieBracket 𝕜 V W x := by + rw [← lieBracketWithin_univ, ← univ_inter s, lieBracketWithin_inter h] + +theorem lieBracketWithin_of_isOpen (hs : IsOpen s) (hx : x ∈ s) : + lieBracketWithin 𝕜 V W s x = lieBracket 𝕜 V W x := + lieBracketWithin_of_mem_nhds (hs.mem_nhds hx) + +theorem lieBracketWithin_eq_lieBracket (hs : UniqueDiffWithinAt 𝕜 s x) + (hV : DifferentiableAt 𝕜 V x) (hW : DifferentiableAt 𝕜 W x) : + lieBracketWithin 𝕜 V W s x = lieBracket 𝕜 V W x := by + simp [lieBracketWithin, lieBracket, fderivWithin_eq_fderiv, hs, hV, hW] + +/-- Variant of `lieBracketWithin_congr_set` where one requires the sets to coincide only in +the complement of a point. -/ +theorem lieBracketWithin_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := by + simp [lieBracketWithin, fderivWithin_congr_set' _ h] + +theorem lieBracketWithin_congr_set (h : s =ᶠ[𝓝 x] t) : + lieBracketWithin 𝕜 V W s x = lieBracketWithin 𝕜 V W t x := + lieBracketWithin_congr_set' x <| h.filter_mono inf_le_left + +/-- Variant of `lieBracketWithin_eventually_congr_set` where one requires the sets to coincide only +in the complement of a point. -/ +theorem lieBracketWithin_eventually_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) : + lieBracketWithin 𝕜 V W s =ᶠ[𝓝 x] lieBracketWithin 𝕜 V W t := + (eventually_nhds_nhdsWithin.2 h).mono fun _ => lieBracketWithin_congr_set' y + +theorem lieBracketWithin_eventually_congr_set (h : s =ᶠ[𝓝 x] t) : + lieBracketWithin 𝕜 V W s =ᶠ[𝓝 x] lieBracketWithin 𝕜 V W t := + lieBracketWithin_eventually_congr_set' x <| h.filter_mono inf_le_left + +theorem _root_.DifferentiableWithinAt.lieBracketWithin_congr_mono + (hV : DifferentiableWithinAt 𝕜 V s x) (hVs : EqOn V₁ V t) (hVx : V₁ x = V x) + (hW : DifferentiableWithinAt 𝕜 W s x) (hWs : EqOn W₁ W t) (hWx : W₁ x = W x) + (hxt : UniqueDiffWithinAt 𝕜 t x) (h₁ : t ⊆ s) : + lieBracketWithin 𝕜 V₁ W₁ t x = lieBracketWithin 𝕜 V W s x := by + simp [lieBracketWithin, hV.fderivWithin_congr_mono, hW.fderivWithin_congr_mono, hVs, hVx, + hWs, hWx, hxt, h₁] + +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq + (hV : V₁ =ᶠ[𝓝[s] x] V) (hxV : V₁ x = V x) (hW : W₁ =ᶠ[𝓝[s] x] W) (hxW : W₁ x = W x) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := by + simp only [lieBracketWithin, hV.fderivWithin_eq hxV, hW.fderivWithin_eq hxW, hxV, hxW] + +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq_of_mem + (hV : V₁ =ᶠ[𝓝[s] x] V) (hW : W₁ =ᶠ[𝓝[s] x] W) (hx : x ∈ s) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + hV.lieBracketWithin_vectorField_eq (mem_of_mem_nhdsWithin hx hV :) + hW (mem_of_mem_nhdsWithin hx hW :) + +/-- If vector fields coincide on a neighborhood of a point within a set, then the Lie brackets +also coincide on a neighborhood of this point within this set. Version where one considers the Lie +bracket within a subset. -/ +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField' + (hV : V₁ =ᶠ[𝓝[s] x] V) (hW : W₁ =ᶠ[𝓝[s] x] W) (ht : t ⊆ s) : + lieBracketWithin 𝕜 V₁ W₁ t =ᶠ[𝓝[s] x] lieBracketWithin 𝕜 V W t := by + filter_upwards [hV.fderivWithin' ht (𝕜 := 𝕜), hW.fderivWithin' ht (𝕜 := 𝕜), hV, hW] + with x hV' hW' hV hW + simp [lieBracketWithin, hV', hW', hV, hW] + +protected theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField + (hV : V₁ =ᶠ[𝓝[s] x] V) (hW : W₁ =ᶠ[𝓝[s] x] W) : + lieBracketWithin 𝕜 V₁ W₁ s =ᶠ[𝓝[s] x] lieBracketWithin 𝕜 V W s := + hV.lieBracketWithin_vectorField' hW Subset.rfl + +protected theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq_of_insert + (hV : V₁ =ᶠ[𝓝[insert x s] x] V) (hW : W₁ =ᶠ[𝓝[insert x s] x] W) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := by + apply mem_of_mem_nhdsWithin (mem_insert x s) (hV.lieBracketWithin_vectorField' hW + (subset_insert x s)) + +theorem _root_.Filter.EventuallyEq.lieBracketWithin_vectorField_eq_nhds + (hV : V₁ =ᶠ[𝓝 x] V) (hW : W₁ =ᶠ[𝓝 x] W) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + (hV.filter_mono nhdsWithin_le_nhds).lieBracketWithin_vectorField_eq hV.self_of_nhds + (hW.filter_mono nhdsWithin_le_nhds) hW.self_of_nhds + +theorem lieBracketWithin_congr + (hV : EqOn V₁ V s) (hVx : V₁ x = V x) (hW : EqOn W₁ W s) (hWx : W₁ x = W x) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + (hV.eventuallyEq.filter_mono inf_le_right).lieBracketWithin_vectorField_eq hVx + (hW.eventuallyEq.filter_mono inf_le_right) hWx + +/-- Version of `lieBracketWithin_congr` in which one assumes that the point belongs to the +given set. -/ +theorem lieBracketWithin_congr' (hV : EqOn V₁ V s) (hW : EqOn W₁ W s) (hx : x ∈ s) : + lieBracketWithin 𝕜 V₁ W₁ s x = lieBracketWithin 𝕜 V W s x := + lieBracketWithin_congr hV (hV hx) hW (hW hx) + +theorem _root_.Filter.EventuallyEq.lieBracket_vectorField_eq + (hV : V₁ =ᶠ[𝓝 x] V) (hW : W₁ =ᶠ[𝓝 x] W) : + lieBracket 𝕜 V₁ W₁ x = lieBracket 𝕜 V W x := by + rw [← lieBracketWithin_univ, ← lieBracketWithin_univ, hV.lieBracketWithin_vectorField_eq_nhds hW] + +protected theorem _root_.Filter.EventuallyEq.lieBracket_vectorField + (hV : V₁ =ᶠ[𝓝 x] V) (hW : W₁ =ᶠ[𝓝 x] W) : lieBracket 𝕜 V₁ W₁ =ᶠ[𝓝 x] lieBracket 𝕜 V W := by + filter_upwards [hV.eventuallyEq_nhds, hW.eventuallyEq_nhds] with y hVy hWy + exact hVy.lieBracket_vectorField_eq hWy + +/-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity +`[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ +lemma leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt + {U V W : E → E} {s : Set E} {x : E} (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) + (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) + (hW : ContDiffWithinAt 𝕜 2 W s x) + (h'U : IsSymmSndFDerivWithinAt 𝕜 U s x) (h'V : IsSymmSndFDerivWithinAt 𝕜 V s x) + (h'W : IsSymmSndFDerivWithinAt 𝕜 W s x) : + lieBracketWithin 𝕜 U (lieBracketWithin 𝕜 V W s) s x = + lieBracketWithin 𝕜 (lieBracketWithin 𝕜 U V s) W s x + + lieBracketWithin 𝕜 V (lieBracketWithin 𝕜 U W s) s x := by + simp only [lieBracketWithin_eq, map_sub] + have aux₁ {U V : E → E} (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) : + DifferentiableWithinAt 𝕜 (fun x ↦ (fderivWithin 𝕜 V s x) (U x)) s x := + have := hV.fderivWithin_right_apply (hU.of_le one_le_two) hs le_rfl hx + this.differentiableWithinAt le_rfl + have aux₂ {U V : E → E} (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) : + fderivWithin 𝕜 (fun y ↦ (fderivWithin 𝕜 U s y) (V y)) s x = + (fderivWithin 𝕜 U s x).comp (fderivWithin 𝕜 V s x) + + (fderivWithin 𝕜 (fderivWithin 𝕜 U s) s x).flip (V x) := by + refine fderivWithin_clm_apply (hs x hx) ?_ (hV.differentiableWithinAt one_le_two) + exact (hU.fderivWithin_right hs le_rfl hx).differentiableWithinAt le_rfl + rw [fderivWithin_sub (hs x hx) (aux₁ hV hW) (aux₁ hW hV)] + rw [fderivWithin_sub (hs x hx) (aux₁ hU hV) (aux₁ hV hU)] + rw [fderivWithin_sub (hs x hx) (aux₁ hU hW) (aux₁ hW hU)] + rw [aux₂ hW hV, aux₂ hV hW, aux₂ hV hU, aux₂ hU hV, aux₂ hW hU, aux₂ hU hW] + simp only [ContinuousLinearMap.coe_sub', Pi.sub_apply, ContinuousLinearMap.add_apply, + ContinuousLinearMap.coe_comp', Function.comp_apply, ContinuousLinearMap.flip_apply, h'V.eq, + h'U.eq, h'W.eq] + abel + +/-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity +`[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ +lemma leibniz_identity_lieBracketWithin [IsRCLikeNormedField 𝕜] {U V W : E → E} {s : Set E} {x : E} + (hs : UniqueDiffOn 𝕜 s) (h'x : x ∈ closure (interior s)) (hx : x ∈ s) + (hU : ContDiffWithinAt 𝕜 2 U s x) (hV : ContDiffWithinAt 𝕜 2 V s x) + (hW : ContDiffWithinAt 𝕜 2 W s x) : + lieBracketWithin 𝕜 U (lieBracketWithin 𝕜 V W s) s x = + lieBracketWithin 𝕜 (lieBracketWithin 𝕜 U V s) W s x + + lieBracketWithin 𝕜 V (lieBracketWithin 𝕜 U W s) s x := by + apply leibniz_identity_lieBracketWithin_of_isSymmSndFDerivWithinAt hs hx hU hV hW + · exact hU.isSymmSndFDerivWithinAt le_rfl hs h'x hx + · exact hV.isSymmSndFDerivWithinAt le_rfl hs h'x hx + · exact hW.isSymmSndFDerivWithinAt le_rfl hs h'x hx + +/-- The Lie bracket of vector fields in vector spaces satisfies the Leibniz identity +`[U, [V, W]] = [[U, V], W] + [V, [U, W]]`. -/ +lemma leibniz_identity_lieBracket [IsRCLikeNormedField 𝕜] {U V W : E → E} {x : E} + (hU : ContDiffAt 𝕜 2 U x) (hV : ContDiffAt 𝕜 2 V x) (hW : ContDiffAt 𝕜 2 W x) : + lieBracket 𝕜 U (lieBracket 𝕜 V W) x = + lieBracket 𝕜 (lieBracket 𝕜 U V) W x + lieBracket 𝕜 V (lieBracket 𝕜 U W) x := by + simp only [← lieBracketWithin_univ, ← contDiffWithinAt_univ] at hU hV hW ⊢ + exact leibniz_identity_lieBracketWithin uniqueDiffOn_univ (by simp) (mem_univ _) hU hV hW + +end VectorField From 92f63d5c1988987aa375830b9b41cbfd269c3625 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 14:53:33 +0000 Subject: [PATCH 73/90] fix(CI): fail if tests are noisy (#19268) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/CI.3A.20noisy.20.22test.20mathlib.22/near/483367292) #19269 adds a test file with `#eval ""` and tests fail on that PR. --- .github/build.in.yml | 2 +- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build_fork.yml | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index cf758b28c47f3..bab96d693f6b2 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -270,7 +270,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 68c1d03a1f5f2..e59f48d2c7267 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -280,7 +280,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 7edf488cd7923..1c778a5720c23 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -287,7 +287,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 76448210f9cb0..e92f87fd50a68 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -284,7 +284,7 @@ jobs: with: linters: gcc run: - lake test + lake --iofail test - name: check for unused imports id: shake From 238cb187bb518855c7150ac44faa919685eb682b Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 15:32:16 +0000 Subject: [PATCH 74/90] feat(CI): merge bors (#19078) Unify the three `bors d`, `bors merge` and `bord r+` actions into a single one. Prefer using `TRIAGE` token for GH API quota. --- .github/workflows/maintainer_bors.yml | 93 +++++++++++++++++++++++++++ 1 file changed, 93 insertions(+) create mode 100644 .github/workflows/maintainer_bors.yml diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml new file mode 100644 index 0000000000000..d81dded626572 --- /dev/null +++ b/.github/workflows/maintainer_bors.yml @@ -0,0 +1,93 @@ +name: Add "ready-to-merge" and "delegated" label + +# triggers the action when +on: + # the PR receives a comment + issue_comment: + types: [created] + # the PR receives a review + pull_request_review: + # whether or not it is accompanied by a comment + types: [submitted] + # the PR receives a review comment + pull_request_review_comment: + types: [created] + +jobs: + add_ready_to_merge_label: + # we set some variables. The ones of the form `${{ X }}${{ Y }}` are typically not + # both set simultaneously: depending on the event that triggers the PR, usually only one is set + env: + AUTHOR: ${{ github.event.comment.user.login }}${{ github.event.review.user.login }} + PR_NUMBER: ${{ github.event.issue.number }}${{ github.event.pull_request.number }} + COMMENT_EVENT: ${{ github.event.comment.body }} + COMMENT_REVIEW: ${{ github.event.review.body }} + PR_TITLE_ISSUE: ${{ github.event.issue.title }} + PR_TITLE_PR: ${{ github.event.pull_request.title }} + PR_URL: ${{ github.event.issue.html_url }}${{ github.event.pull_request.html_url }} + EVENT_NAME: ${{ github.event_name }} + name: Add ready-to-merge or delegated label + runs-on: ubuntu-latest + steps: + - name: Find bors merge/delegate + id: merge_or_delegate + run: | + COMMENT="${COMMENT_EVENT}${COMMENT_REVIEW}" + # we strip `\r` since line endings from GitHub contain this character + COMMENT="${COMMENT//$'\r'/}" + # for debugging, we print some information + printf '%s' "${COMMENT}" | hexdump -cC + printf 'Comment:"%s"\n' "${COMMENT}" + m_or_d="$(printf '%s' "${COMMENT}" | + sed -n 's=^bors *\(merge\|r+\) *$=ready-to-merge=p; s=^bors *d.*=delegated=p' | head -1)" + + printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" + + printf $'mOrD=%s\n' "${m_or_d}" > "${GITHUB_OUTPUT}" + if [ "${AUTHOR}" == 'leanprover-community-mathlib4-bot' ] || + [ "${AUTHOR}" == 'leanprover-community-bot-assistant' ] + then + printf $'bot=true\n' + printf $'bot=true\n' > "${GITHUB_OUTPUT}" + else + printf $'bot=false\n' + printf $'bot=false\n' > "${GITHUB_OUTPUT}" + fi + + - name: Check whether user is a mathlib admin + id: user_permission + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' }} + uses: actions-cool/check-user-permission@v2 + with: + require: 'admin' + # review(_comment) use + # require: 'write' + # token: ${{ secrets.TRIAGE_TOKEN }} + + - name: Add ready-to-merge or delegated label + id: add_label + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + uses: octokit/request-action@v2.x + with: + # check is this ok? was /repos/:repository/issues/:issue_number/labels + route: POST /repos/:repository/issues/${PR_NUMBER}/labels + repository: ${{ github.repository }} + issue_number: ${PR_NUMBER} + labels: '["${{ steps.merge_or_delegate.outputs.mOrD }}"]' + env: + GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} # comment uses ${{ secrets.GITHUB_TOKEN }} + + - if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + id: remove_labels + name: Remove "awaiting-author" + # we use curl rather than octokit/request-action so that the job won't fail + # (and send an annoying email) if the labels don't exist + run: | + curl --request DELETE \ + --url "https://api.github.com/repos/${{ github.repository }}/issues/${PR_NUMBER}/labels/awaiting-author" \ + --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' + # comment uses ${{ secrets.GITHUB_TOKEN }} From 8eed8831892ac64dca2e5d290fdbc2dcfeecfd83 Mon Sep 17 00:00:00 2001 From: sven-manthe Date: Wed, 20 Nov 2024 16:52:58 +0000 Subject: [PATCH 75/90] chore(CategoryTheory/IsConnected): add refl etc annotations (#19017) --- Mathlib/CategoryTheory/IsConnected.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 96bb16ce5927a..f4fa03b9c28a6 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -263,11 +263,11 @@ theorem isConnected_of_isConnected_op [IsConnected Jᵒᵖ] : IsConnected J := def Zag (j₁ j₂ : J) : Prop := Nonempty (j₁ ⟶ j₂) ∨ Nonempty (j₂ ⟶ j₁) -theorem Zag.refl (X : J) : Zag X X := Or.inl ⟨𝟙 _⟩ +@[refl] theorem Zag.refl (X : J) : Zag X X := Or.inl ⟨𝟙 _⟩ theorem zag_symmetric : Symmetric (@Zag J _) := fun _ _ h => h.symm -theorem Zag.symm {j₁ j₂ : J} (h : Zag j₁ j₂) : Zag j₂ j₁ := zag_symmetric h +@[symm] theorem Zag.symm {j₁ j₂ : J} (h : Zag j₁ j₂) : Zag j₂ j₁ := zag_symmetric h theorem Zag.of_hom {j₁ j₂ : J} (f : j₁ ⟶ j₂) : Zag j₁ j₂ := Or.inl ⟨f⟩ @@ -286,11 +286,12 @@ theorem zigzag_equivalence : _root_.Equivalence (@Zigzag J _) := _root_.Equivalence.mk Relation.reflexive_reflTransGen (fun h => zigzag_symmetric h) (fun h g => Relation.transitive_reflTransGen h g) -theorem Zigzag.refl (X : J) : Zigzag X X := zigzag_equivalence.refl _ +@[refl] theorem Zigzag.refl (X : J) : Zigzag X X := zigzag_equivalence.refl _ -theorem Zigzag.symm {j₁ j₂ : J} (h : Zigzag j₁ j₂) : Zigzag j₂ j₁ := zigzag_symmetric h +@[symm] theorem Zigzag.symm {j₁ j₂ : J} (h : Zigzag j₁ j₂) : Zigzag j₂ j₁ := zigzag_symmetric h -theorem Zigzag.trans {j₁ j₂ j₃ : J} (h₁ : Zigzag j₁ j₂) (h₂ : Zigzag j₂ j₃) : Zigzag j₁ j₃ := +@[trans] theorem Zigzag.trans {j₁ j₂ j₃ : J} (h₁ : Zigzag j₁ j₂) (h₂ : Zigzag j₂ j₃) : + Zigzag j₁ j₃ := zigzag_equivalence.trans h₁ h₂ theorem Zigzag.of_zag {j₁ j₂ : J} (h : Zag j₁ j₂) : Zigzag j₁ j₂ := From 46d221cdeabdc58a33e8724bd000be0f730d31ff Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 20 Nov 2024 17:06:21 +0000 Subject: [PATCH 76/90] feat(NumberTheory/LSeries/DirichletContinuation): results on logarithmic derivatives (#19254) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit As a further step in the direction of Dirichlet's Theorem on primes in AP, this PR adds results on (negative) logarithmic derivatives of `L χ` for a non-trivial Dirichlet character `χ` and of `s ↦ (L χ s) * (s - 1)` when `χ` is trivial. See [here on Zulip](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/482543603). Co-authored-by: Michael Stoll <99838730+MichaelStollBayreuth@users.noreply.github.com> --- .../LSeries/DirichletContinuation.lean | 91 ++++++++++++++++++- 1 file changed, 88 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean index e4b323edd1d75..033ad5dde6fc6 100644 --- a/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean +++ b/Mathlib/NumberTheory/LSeries/DirichletContinuation.lean @@ -84,11 +84,13 @@ lemma deriv_LFunction_eq_deriv_LSeries (χ : DirichletCharacter ℂ N) {s : ℂ} The L-function of a Dirichlet character is differentiable, except at `s = 1` if the character is trivial. -/ +@[fun_prop] lemma differentiableAt_LFunction (χ : DirichletCharacter ℂ N) (s : ℂ) (hs : s ≠ 1 ∨ χ ≠ 1) : DifferentiableAt ℂ (LFunction χ) s := ZMod.differentiableAt_LFunction χ s (hs.imp_right χ.sum_eq_zero_of_ne_one) /-- The L-function of a non-trivial Dirichlet character is differentiable everywhere. -/ +@[fun_prop] lemma differentiable_LFunction {χ : DirichletCharacter ℂ N} (hχ : χ ≠ 1) : Differentiable ℂ (LFunction χ) := (differentiableAt_LFunction _ · <| Or.inr hχ) @@ -113,7 +115,7 @@ lemma Even.LFunction_neg_two_mul_nat {χ : DirichletCharacter ℂ N} (hχ : Even ZMod.LFunction_neg_two_mul_nat_sub_one hχ.to_fun n /-! -## Results on changing levels +### Results on changing levels -/ private lemma LFunction_changeLevel_aux {M N : ℕ} [NeZero M] [NeZero N] (hMN : M ∣ N) @@ -156,7 +158,7 @@ lemma LFunction_changeLevel {M N : ℕ} [NeZero M] [NeZero N] (hMN : M ∣ N) · exact LFunction_changeLevel_aux hMN χ h /-! -## The `L`-function of the trivial character mod `N` +### The `L`-function of the trivial character mod `N` -/ /-- The `L`-function of the trivial character mod `N`. -/ @@ -189,7 +191,7 @@ lemma LFunctionTrivChar_residue_one : fun_prop /-! -## Completed L-functions and the functional equation +### Completed L-functions and the functional equation -/ section gammaFactor @@ -307,3 +309,86 @@ theorem completedLFunction_one_sub {χ : DirichletCharacter ℂ N} (hχ : IsPrim end IsPrimitive end DirichletCharacter + +/-! +### The logarithmic derivative of the L-function of a Dirichlet character + +We show that `s ↦ -(L' χ s) / L χ s + 1 / (s - 1)` is continuous outside the zeros of `L χ` +when `χ` is a trivial Dirichlet character and that `-L' χ / L χ` is continuous outside +the zeros of `L χ` when `χ` is nontrivial. +-/ + +namespace DirichletCharacter + +open Complex + +section trivial + +variable (n : ℕ) [NeZero n] + +/-- The function obtained by "multiplying away" the pole of `L χ` for a trivial Dirichlet +character `χ`. Its (negative) logarithmic derivative is used to prove Dirichlet's Theorem +on primes in arithmetic progression. -/ +noncomputable abbrev LFunctionTrivChar₁ : ℂ → ℂ := + Function.update (fun s ↦ (s - 1) * LFunctionTrivChar n s) 1 + (∏ p ∈ n.primeFactors, (1 - (p : ℂ)⁻¹)) + +lemma LFunctionTrivChar₁_apply_one_ne_zero : LFunctionTrivChar₁ n 1 ≠ 0 := by + simp only [Function.update_same] + refine Finset.prod_ne_zero_iff.mpr fun p hp ↦ ?_ + simpa only [ne_eq, sub_ne_zero, one_eq_inv, Nat.cast_eq_one] + using (Nat.prime_of_mem_primeFactors hp).ne_one + +/-- `s ↦ (s - 1) * L χ s` is an entire function when `χ` is a trivial Dirichlet character. -/ +lemma differentiable_LFunctionTrivChar₁ : Differentiable ℂ (LFunctionTrivChar₁ n) := by + rw [← differentiableOn_univ, + ← differentiableOn_compl_singleton_and_continuousAt_iff (c := 1) Filter.univ_mem] + refine ⟨DifferentiableOn.congr (f := fun s ↦ (s - 1) * LFunctionTrivChar n s) + (fun _ hs ↦ DifferentiableAt.differentiableWithinAt <| by fun_prop (disch := simp_all [hs])) + fun _ hs ↦ Function.update_noteq (Set.mem_diff_singleton.mp hs).2 .., + continuousWithinAt_compl_self.mp ?_⟩ + simpa only [continuousWithinAt_compl_self, continuousAt_update_same] + using LFunctionTrivChar_residue_one + +lemma deriv_LFunctionTrivChar₁_apply_of_ne_one {s : ℂ} (hs : s ≠ 1) : + deriv (LFunctionTrivChar₁ n) s = + (s - 1) * deriv (LFunctionTrivChar n) s + LFunctionTrivChar n s := by + have H : deriv (LFunctionTrivChar₁ n) s = + deriv (fun w ↦ (w - 1) * LFunctionTrivChar n w) s := by + refine eventuallyEq_iff_exists_mem.mpr ?_ |>.deriv_eq + exact ⟨_, isOpen_ne.mem_nhds hs, fun _ hw ↦ Function.update_noteq (Set.mem_setOf.mp hw) ..⟩ + rw [H, deriv_mul (by fun_prop) (differentiableAt_LFunction _ s (.inl hs)), deriv_sub_const, + deriv_id'', one_mul, add_comm] + +/-- The negative logarithmtic derivative of `s ↦ (s - 1) * L χ s` for a trivial +Dirichlet character `χ` is continuous away from the zeros of `L χ` (including at `s = 1`). -/ +lemma continuousOn_neg_logDeriv_LFunctionTrivChar₁ : + ContinuousOn (fun s ↦ -deriv (LFunctionTrivChar₁ n) s / LFunctionTrivChar₁ n s) + {s | s = 1 ∨ LFunctionTrivChar n s ≠ 0} := by + simp_rw [neg_div] + have h := differentiable_LFunctionTrivChar₁ n + refine ((h.contDiff.continuous_deriv le_rfl).continuousOn.div + h.continuous.continuousOn fun w hw ↦ ?_).neg + rcases eq_or_ne w 1 with rfl | hw' + · exact LFunctionTrivChar₁_apply_one_ne_zero _ + · rw [LFunctionTrivChar₁, Function.update_noteq hw', mul_ne_zero_iff] + exact ⟨sub_ne_zero_of_ne hw', (Set.mem_setOf.mp hw).resolve_left hw'⟩ + +end trivial + +section nontrivial + +variable {n : ℕ} [NeZero n] {χ : DirichletCharacter ℂ n} + +/-- The negative logarithmic derivative of the L-function of a nontrivial Dirichlet character +is continuous away from the zeros of the L-function. -/ +lemma continuousOn_neg_logDeriv_LFunction_of_nontriv (hχ : χ ≠ 1) : + ContinuousOn (fun s ↦ -deriv (LFunction χ) s / LFunction χ s) {s | LFunction χ s ≠ 0} := by + simp only [neg_div] + have h := differentiable_LFunction hχ + exact ((h.contDiff.continuous_deriv le_rfl).continuousOn.div + h.continuous.continuousOn fun _ hw ↦ hw).neg + +end nontrivial + +end DirichletCharacter From a5d04c838085d6c3554d59278c04126b62314757 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 20 Nov 2024 17:19:12 +0000 Subject: [PATCH 77/90] chore(RingTheory): improve and generalize submodule localization API (#19118) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In Algebra/Module/LocalizedModule/Submodule, we introduce `Submodule.localized₀`, which is localization of a submodule considered as a submodule over the base ring rather than the localization. This allows us to talk about the localization of a submodule without choosing a specific ring localization, just like what `IsLocalizedModule` allows us to do for localization of a module. As applications, `Rₚ` and every hypothesis depending on it are completely removed from the statement of `Submodule.le_of_localization_maximal`, etc. in RingTheory/LocalProperties/Submodule. We also reorder lemmas in this file to make the development more natural and golf the proofs (making use of `Module.eqIdeal` introduced in RingTheory/Ideal/Defs; also add some trivial lemmas. In RingTheory/LocalProperties/Projective, we fix a proof due to removed explicit argument `Rₚ` and remove some unnecessary lines. --- .../Module/LocalizedModule/Submodule.lean | 94 ++++++++---- Mathlib/RingTheory/Ideal/Defs.lean | 9 ++ .../LocalProperties/Projective.lean | 18 +-- .../RingTheory/LocalProperties/Submodule.lean | 139 +++++++++++------- 4 files changed, 163 insertions(+), 97 deletions(-) diff --git a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean index 786335636ddb0..650844736f7ca 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean @@ -28,49 +28,73 @@ Results about localizations of submodules and quotient modules are provided in t open nonZeroDivisors -universe u u' v v' w w' - -variable {R : Type u} (S : Type u') {M : Type v} {N : Type v'} -variable [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] +variable {R S M N : Type*} +variable (S) [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] variable [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] variable (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] variable (M' : Submodule R M) -/-- Let `S` be the localization of `R` at `p` and `N` be the localization of `M` at `p`. -This is the localization of an `R`-submodule of `M` viewed as an `S`-submodule of `N`. -/ -def Submodule.localized' : Submodule S N where +namespace Submodule + +/-- Let `N` be a localization of an `R`-module `M` at `p`. +This is the localization of an `R`-submodule of `M` viewed as an `R`-submodule of `N`. -/ +def localized₀ : Submodule R N where carrier := { x | ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x } - add_mem' := fun {x} {y} ⟨m, hm, s, hx⟩ ⟨n, hn, t, hy⟩ ↦ ⟨t • m + s • n, add_mem (M'.smul_mem t hm) + add_mem' := fun {x y} ⟨m, hm, s, hx⟩ ⟨n, hn, t, hy⟩ ↦ ⟨t • m + s • n, add_mem (M'.smul_mem t hm) (M'.smul_mem s hn), s * t, by rw [← hx, ← hy, IsLocalizedModule.mk'_add_mk']⟩ zero_mem' := ⟨0, zero_mem _, 1, by simp⟩ - smul_mem' := fun r x h ↦ by - have ⟨m, hm, s, hx⟩ := h + smul_mem' r x := by + rintro ⟨m, hm, s, hx⟩ + exact ⟨r • m, smul_mem M' _ hm, s, by rw [IsLocalizedModule.mk'_smul, hx]⟩ + +/-- Let `S` be the localization of `R` at `p` and `N` be a localization of `M` at `p`. +This is the localization of an `R`-submodule of `M` viewed as an `S`-submodule of `N`. -/ +def localized' : Submodule S N where + __ := localized₀ p f M' + smul_mem' := fun r x ⟨m, hm, s, hx⟩ ↦ by have ⟨y, t, hyt⟩ := IsLocalization.mk'_surjective p r exact ⟨y • m, M'.smul_mem y hm, t * s, by simp [← hyt, ← hx, IsLocalizedModule.mk'_smul_mk']⟩ -lemma Submodule.mem_localized' (x : N) : - x ∈ Submodule.localized' S p f M' ↔ ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x := +lemma mem_localized₀ (x : N) : + x ∈ localized₀ p f M' ↔ ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x := + Iff.rfl + +lemma mem_localized' (x : N) : + x ∈ localized' S p f M' ↔ ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x := Iff.rfl +/-- `localized₀` is the same as `localized'` considered as a submodule over the base ring. -/ +lemma restrictScalars_localized' : + (localized' S p f M').restrictScalars R = localized₀ p f M' := + rfl + /-- The localization of an `R`-submodule of `M` at `p` viewed as an `Rₚ`-submodule of `Mₚ`. -/ -abbrev Submodule.localized : Submodule (Localization p) (LocalizedModule p M) := +abbrev localized : Submodule (Localization p) (LocalizedModule p M) := M'.localized' (Localization p) p (LocalizedModule.mkLinearMap p M) @[simp] -lemma Submodule.localized'_bot : (⊥ : Submodule R M).localized' S p f = ⊥ := by +lemma localized₀_bot : (⊥ : Submodule R M).localized₀ p f = ⊥ := by rw [← le_bot_iff] rintro _ ⟨_, rfl, s, rfl⟩ simp only [IsLocalizedModule.mk'_zero, mem_bot] @[simp] -lemma Submodule.localized'_top : (⊤ : Submodule R M).localized' S p f = ⊤ := by +lemma localized'_bot : (⊥ : Submodule R M).localized' S p f = ⊥ := + SetLike.ext' (by apply SetLike.ext'_iff.mp <| Submodule.localized₀_bot p f) + +@[simp] +lemma localized₀_top : (⊤ : Submodule R M).localized₀ p f = ⊤ := by rw [← top_le_iff] rintro x _ obtain ⟨⟨x, s⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f x exact ⟨x, trivial, s, rfl⟩ @[simp] -lemma Submodule.localized'_span (s : Set M) : (span R s).localized' S p f = span S (f '' s) := by +lemma localized'_top : (⊤ : Submodule R M).localized' S p f = ⊤ := + SetLike.ext' (by apply SetLike.ext'_iff.mp <| Submodule.localized₀_top p f) + +@[simp] +lemma localized'_span (s : Set M) : (span R s).localized' S p f = span S (f '' s) := by apply le_antisymm · rintro _ ⟨x, hx, t, rfl⟩ have := IsLocalizedModule.mk'_smul_mk' S f 1 x t 1 @@ -87,29 +111,45 @@ lemma Submodule.localized'_span (s : Set M) : (span R s).localized' S p f = span /-- The localization map of a submodule. -/ @[simps!] -def Submodule.toLocalized' : M' →ₗ[R] M'.localized' S p f := - f.restrict (q := (M'.localized' S p f).restrictScalars R) (fun x hx ↦ ⟨x, hx, 1, by simp⟩) +def toLocalized₀ : M' →ₗ[R] M'.localized₀ p f := f.restrict fun x hx ↦ ⟨x, hx, 1, by simp⟩ /-- The localization map of a submodule. -/ -abbrev Submodule.toLocalized : M' →ₗ[R] M'.localized p := +@[simps!] +def toLocalized' : M' →ₗ[R] M'.localized' S p f := toLocalized₀ p f M' + +/-- The localization map of a submodule. -/ +abbrev toLocalized : M' →ₗ[R] M'.localized p := M'.toLocalized' (Localization p) p (LocalizedModule.mkLinearMap p M) -instance Submodule.isLocalizedModule : IsLocalizedModule p (M'.toLocalized' S p f) where +instance : IsLocalizedModule p (M'.toLocalized₀ p f) where map_units x := by simp_rw [Module.End_isUnit_iff] constructor · exact fun _ _ e ↦ Subtype.ext (IsLocalizedModule.smul_injective f x (congr_arg Subtype.val e)) - · rintro m - use (IsLocalization.mk' S 1 x) • m - rw [Module.algebraMap_end_apply, ← smul_assoc, IsLocalization.smul_mk'_one, - IsLocalization.mk'_self', one_smul] + · rintro ⟨_, m, hm, s, rfl⟩ + refine ⟨⟨IsLocalizedModule.mk' f m (s * x), ⟨_, hm, _, rfl⟩⟩, Subtype.ext ?_⟩ + rw [Module.algebraMap_end_apply, SetLike.val_smul_of_tower, + ← IsLocalizedModule.mk'_smul, ← Submonoid.smul_def, IsLocalizedModule.mk'_cancel_right] surj' := by rintro ⟨y, x, hx, s, rfl⟩ exact ⟨⟨⟨x, hx⟩, s⟩, by ext; simp⟩ exists_of_eq e := by simpa [Subtype.ext_iff] using IsLocalizedModule.exists_of_eq (S := p) (f := f) (congr_arg Subtype.val e) +instance isLocalizedModule : IsLocalizedModule p (M'.toLocalized' S p f) := + inferInstanceAs (IsLocalizedModule p (M'.toLocalized₀ p f)) + +end Submodule + +section Quotient + +variable {R S M N : Type*} +variable (S) [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] +variable [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] +variable (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] +variable (M' : Submodule R M) + /-- The localization map of a quotient module. -/ def Submodule.toLocalizedQuotient' : M ⧸ M' →ₗ[R] N ⧸ M'.localized' S p f := Submodule.mapQ M' ((M'.localized' S p f).restrictScalars R) f (fun x hx ↦ ⟨x, hx, 1, by simp⟩) @@ -147,10 +187,12 @@ instance IsLocalizedModule.toLocalizedQuotient' (M' : Submodule R M) : instance (M' : Submodule R M) : IsLocalizedModule p (M'.toLocalizedQuotient p) := IsLocalizedModule.toLocalizedQuotient' _ _ _ _ +end Quotient + section LinearMap -variable {P : Type w} [AddCommGroup P] [Module R P] -variable {Q : Type w'} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] +variable {P : Type*} [AddCommGroup P] [Module R P] +variable {Q : Type*} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] variable (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] lemma LinearMap.localized'_ker_eq_ker_localizedMap (g : M →ₗ[R] P) : diff --git a/Mathlib/RingTheory/Ideal/Defs.lean b/Mathlib/RingTheory/Ideal/Defs.lean index d2eb82b15fdd4..ba78f8a4b205b 100644 --- a/Mathlib/RingTheory/Ideal/Defs.lean +++ b/Mathlib/RingTheory/Ideal/Defs.lean @@ -68,6 +68,15 @@ theorem unit_mul_mem_iff_mem {x y : α} (hy : IsUnit y) : y * x ∈ I ↔ x ∈ end Ideal +/-- For two elements `m` and `m'` in an `R`-module `M`, the set of elements `r : R` with +equal scalar product with `m` and `m'` is an ideal of `R`. If `M` is a group, this coincides +with the kernel of `LinearMap.toSpanSingleton R M (m - m')`. -/ +def Module.eqIdeal (R) {M} [Semiring R] [AddCommMonoid M] [Module R M] (m m' : M) : Ideal R where + carrier := {r : R | r • m = r • m'} + add_mem' h h' := by simpa [add_smul] using congr($h + $h') + zero_mem' := by simp_rw [Set.mem_setOf, zero_smul] + smul_mem' _ _ h := by simpa [mul_smul] using congr(_ • $h) + end Semiring section CommSemiring diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index d6a190e4ac15f..a2c4615a7942e 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -31,8 +31,8 @@ variable [AddCommGroup N'] [Module R N'] (S : Submonoid R) theorem Module.projective_of_isLocalizedModule {Rₛ Mₛ} [AddCommGroup Mₛ] [Module R Mₛ] [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] (S) (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Projective R M] : - Module.Projective Rₛ Mₛ := - Projective.of_equiv (IsLocalizedModule.isBaseChange S Rₛ f).equiv + Module.Projective Rₛ Mₛ := + Projective.of_equiv (IsLocalizedModule.isBaseChange S Rₛ f).equiv theorem LinearMap.split_surjective_of_localization_maximal (f : M →ₗ[R] N) [Module.FinitePresentation R N] @@ -41,18 +41,8 @@ theorem LinearMap.split_surjective_of_localization_maximal (LocalizedModule.map I.primeCompl f).comp g = LinearMap.id) : ∃ (g : N →ₗ[R] M), f.comp g = LinearMap.id := by show LinearMap.id ∈ LinearMap.range (LinearMap.llcomp R N M N f) - have inst₁ (I : Ideal R) [I.IsMaximal] : - IsLocalizedModule I.primeCompl (LocalizedModule.map (M := N) (N := N) I.primeCompl) := - inferInstance - have inst₂ (I : Ideal R) [I.IsMaximal] : - IsLocalizedModule I.primeCompl (LocalizedModule.map (M := N) (N := M) I.primeCompl) := - inferInstance - apply - @Submodule.mem_of_localization_maximal R (N →ₗ[R] N) _ _ _ - (fun P _ ↦ Localization.AtPrime P) _ _ _ _ _ _ _ _ - (fun P _ ↦ LocalizedModule.map P.primeCompl) - (fun P _ ↦ inst₁ P) - intro I hI + refine Submodule.mem_of_localization_maximal _ (fun P _ ↦ LocalizedModule.map P.primeCompl) _ _ + fun I hI ↦ ?_ rw [LocalizedModule.map_id] have : LinearMap.id ∈ LinearMap.range (LinearMap.llcomp _ (LocalizedModule I.primeCompl N) _ _ (LocalizedModule.map I.primeCompl f)) := H I hI diff --git a/Mathlib/RingTheory/LocalProperties/Submodule.lean b/Mathlib/RingTheory/LocalProperties/Submodule.lean index f4e9274c697bd..2ecaa20be887a 100644 --- a/Mathlib/RingTheory/LocalProperties/Submodule.lean +++ b/Mathlib/RingTheory/LocalProperties/Submodule.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang, David Swinarski -/ import Mathlib.Algebra.Module.LocalizedModule.Submodule -import Mathlib.RingTheory.Ideal.Colon import Mathlib.RingTheory.Localization.AtPrime /-! @@ -15,87 +14,113 @@ In this file, we show that several conditions on submodules can be checked on st open scoped nonZeroDivisors -variable {R A M} [CommRing R] [CommRing A] [AddCommGroup M] [Algebra R A] [Module R M] [Module A M] +variable {R M M₁ : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid M₁] [Module R M₁] + +section maximal variable (Rₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) - [∀ (P : Ideal R) [P.IsMaximal], CommRing (Rₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], CommSemiring (Rₚ P)] [∀ (P : Ideal R) [P.IsMaximal], Algebra R (Rₚ P)] [∀ (P : Ideal R) [P.IsMaximal], IsLocalization.AtPrime (Rₚ P) P] (Mₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) - [∀ (P : Ideal R) [P.IsMaximal], AddCommGroup (Mₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Mₚ P)] [∀ (P : Ideal R) [P.IsMaximal], Module R (Mₚ P)] [∀ (P : Ideal R) [P.IsMaximal], Module (Rₚ P) (Mₚ P)] [∀ (P : Ideal R) [P.IsMaximal], IsScalarTower R (Rₚ P) (Mₚ P)] (f : ∀ (P : Ideal R) [P.IsMaximal], M →ₗ[R] Mₚ P) - [inst : ∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f P)] + (M₁ₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (M₁ₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (M₁ₚ P)] + (f₁ : ∀ (P : Ideal R) [P.IsMaximal], M₁ →ₗ[R] M₁ₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (f₁ P)] + +theorem Submodule.mem_of_localization_maximal (m : M) (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], f P m ∈ N.localized₀ P.primeCompl (f P)) : + m ∈ N := by + let I : Ideal R := N.comap (LinearMap.toSpanSingleton R M m) + suffices I = ⊤ by simpa [I] using I.eq_top_iff_one.mp this + refine Not.imp_symm I.exists_le_maximal fun ⟨P, hP, le⟩ ↦ ?_ + obtain ⟨a, ha, s, e⟩ := h P + rw [← IsLocalizedModule.mk'_one P.primeCompl, IsLocalizedModule.mk'_eq_mk'_iff] at e + obtain ⟨t, ht⟩ := e + simp_rw [smul_smul] at ht + exact (t * s).2 (le <| by apply ht ▸ smul_mem _ _ ha) /-- Let `N₁ N₂ : Submodule R M`. If the localization of `N₁` at each maximal ideal `P` is included in the localization of `N₂` at `P`, then `N₁ ≤ N₂`. -/ theorem Submodule.le_of_localization_maximal {N₁ N₂ : Submodule R M} (h : ∀ (P : Ideal R) [P.IsMaximal], - N₁.localized' (Rₚ P) P.primeCompl (f P) ≤ N₂.localized' (Rₚ P) P.primeCompl (f P)) : - N₁ ≤ N₂ := by - intro x hx - suffices N₂.colon (Submodule.span R {x}) = ⊤ by - simpa using Submodule.mem_colon.mp - (show (1 : R) ∈ N₂.colon (Submodule.span R {x}) from this.symm ▸ Submodule.mem_top) x - (Submodule.mem_span_singleton_self x) - refine Not.imp_symm (N₂.colon (Submodule.span R {x})).exists_le_maximal ?_ - push_neg - intro P hP le - obtain ⟨a, ha, s, e⟩ := h P ⟨x, hx, 1, rfl⟩ - rw [IsLocalizedModule.mk'_eq_mk'_iff] at e - obtain ⟨t, ht⟩ := e - simp at ht - refine (t * s).2 (le (Submodule.mem_colon_singleton.mpr ?_)) - simp only [Submonoid.coe_mul, mul_smul, ← Submonoid.smul_def, ht] - exact N₂.smul_mem _ ha + N₁.localized₀ P.primeCompl (f P) ≤ N₂.localized₀ P.primeCompl (f P)) : + N₁ ≤ N₂ := + fun m hm ↦ mem_of_localization_maximal _ f _ _ fun P hP ↦ h P ⟨m, hm, 1, by simp⟩ /-- Let `N₁ N₂ : Submodule R M`. If the localization of `N₁` at each maximal ideal `P` is equal to the localization of `N₂` at `P`, then `N₁ = N₂`. -/ -theorem Submodule.eq_of_localization_maximal {N₁ N₂ : Submodule R M} +theorem Submodule.eq_of_localization₀_maximal {N₁ N₂ : Submodule R M} (h : ∀ (P : Ideal R) [P.IsMaximal], - N₁.localized' (Rₚ P) P.primeCompl (f P) = N₂.localized' (Rₚ P) P.primeCompl (f P)) : + N₁.localized₀ P.primeCompl (f P) = N₂.localized₀ P.primeCompl (f P)) : N₁ = N₂ := - le_antisymm (Submodule.le_of_localization_maximal Rₚ Mₚ f fun P _ => (h P).le) - (Submodule.le_of_localization_maximal Rₚ Mₚ f fun P _ => (h P).ge) + le_antisymm (Submodule.le_of_localization_maximal Mₚ f fun P _ ↦ (h P).le) + (Submodule.le_of_localization_maximal Mₚ f fun P _ ↦ (h P).ge) /-- A submodule is trivial if its localization at every maximal ideal is trivial. -/ -theorem Submodule.eq_bot_of_localization_maximal (N₁ : Submodule R M) +theorem Submodule.eq_bot_of_localization₀_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized₀ P.primeCompl (f P) = ⊥) : + N = ⊥ := + Submodule.eq_of_localization₀_maximal Mₚ f fun P hP ↦ by simpa using h P + +theorem Submodule.eq_top_of_localization₀_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized₀ P.primeCompl (f P) = ⊤) : + N = ⊤ := + Submodule.eq_of_localization₀_maximal Mₚ f fun P hP ↦ by simpa using h P + +theorem Module.eq_of_localization_maximal (m m' : M) + (h : ∀ (P : Ideal R) [P.IsMaximal], f P m = f P m') : + m = m' := by + by_contra! ne + rw [← one_smul R m, ← one_smul R m'] at ne + have ⟨P, mP, le⟩ := (eqIdeal R m m').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) + have ⟨s, hs⟩ := (IsLocalizedModule.eq_iff_exists P.primeCompl _).mp (h P) + exact s.2 (le hs) + +theorem Module.eq_zero_of_localization_maximal (m : M) + (h : ∀ (P : Ideal R) [P.IsMaximal], f P m = 0) : + m = 0 := + eq_of_localization_maximal _ f _ _ fun P _ ↦ by rw [h, map_zero] + +theorem LinearMap.eq_of_localization_maximal (g g' : M →ₗ[R] M₁) (h : ∀ (P : Ideal R) [P.IsMaximal], - N₁.localized' (Rₚ P) P.primeCompl (f P) = ⊥) : - N₁ = ⊥ := - Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP => by simpa using h P - -theorem Submodule.mem_of_localization_maximal (r : M) (N₁ : Submodule R M) - (h : ∀ (P : Ideal R) [P.IsMaximal], f P r ∈ N₁.localized' (Rₚ P) P.primeCompl (f P)) : - r ∈ N₁ := by - rw [← SetLike.mem_coe, ← Set.singleton_subset_iff, ← Submodule.span_le] - apply Submodule.le_of_localization_maximal Rₚ Mₚ f - intro N₂ hJ - rw [Submodule.localized'_span, Submodule.span_le, Set.image_subset_iff, Set.singleton_subset_iff] - exact h N₂ - -include Rₚ in -theorem Module.eq_zero_of_localization_maximal (r : M) - (h : ∀ (P : Ideal R) [P.IsMaximal], f P r = 0) : - r = 0 := by - rw [← Submodule.mem_bot (R := R)] - apply Submodule.mem_of_localization_maximal Rₚ Mₚ f r ⊥ (by simpa using h) - -include Rₚ in -theorem Module.eq_of_localization_maximal (r s : M) - (h : ∀ (P : Ideal R) [P.IsMaximal], f P r = f P s) : - r = s := by - rw [← sub_eq_zero] - simp_rw [← @sub_eq_zero _ _ (f _ _), ← map_sub] at h - exact Module.eq_zero_of_localization_maximal Rₚ Mₚ f _ h - -include Rₚ f in + IsLocalizedModule.map P.primeCompl (f P) (f₁ P) g = + IsLocalizedModule.map P.primeCompl (f P) (f₁ P) g') : + g = g' := + ext fun x ↦ Module.eq_of_localization_maximal _ f₁ _ _ fun P _ ↦ by + simpa only [IsLocalizedModule.map_apply] using DFunLike.congr_fun (h P) (f P x) + +include f in theorem Module.subsingleton_of_localization_maximal (h : ∀ (P : Ideal R) [P.IsMaximal], Subsingleton (Mₚ P)) : Subsingleton M := by rw [subsingleton_iff_forall_eq 0] intro x - exact Module.eq_of_localization_maximal Rₚ Mₚ f x 0 fun _ _ ↦ Subsingleton.elim _ _ + exact Module.eq_of_localization_maximal Mₚ f x 0 fun _ _ ↦ Subsingleton.elim _ _ + +theorem Submodule.eq_of_localization_maximal {N₁ N₂ : Submodule R M} + (h : ∀ (P : Ideal R) [P.IsMaximal], + N₁.localized' (Rₚ P) P.primeCompl (f P) = N₂.localized' (Rₚ P) P.primeCompl (f P)) : + N₁ = N₂ := + eq_of_localization₀_maximal Mₚ f fun P _ ↦ congr(restrictScalars _ $(h P)) + +theorem Submodule.eq_bot_of_localization_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized' (Rₚ P) P.primeCompl (f P) = ⊥) : + N = ⊥ := + Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP ↦ by simpa using h P + +theorem Submodule.eq_top_of_localization_maximal (N : Submodule R M) + (h : ∀ (P : Ideal R) [P.IsMaximal], N.localized' (Rₚ P) P.primeCompl (f P) = ⊤) : + N = ⊤ := + Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP ↦ by simpa using h P + +end maximal From 8e0d70c35dcdbe42cd6fa06c579a118735380526 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 18:07:47 +0000 Subject: [PATCH 78/90] fix: `>>` and print author (#19300) Use `>>` to set `GITHUB_OUTPUT`, instead of `>`. --- .github/workflows/maintainer_bors.yml | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index d81dded626572..6832b6e007eed 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -42,16 +42,17 @@ jobs: sed -n 's=^bors *\(merge\|r+\) *$=ready-to-merge=p; s=^bors *d.*=delegated=p' | head -1)" printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" + printf $'AUTHOR: \'%s\'\n' "${AUTHOR}" - printf $'mOrD=%s\n' "${m_or_d}" > "${GITHUB_OUTPUT}" + printf $'mOrD=%s\n' "${m_or_d}" >> "${GITHUB_OUTPUT}" if [ "${AUTHOR}" == 'leanprover-community-mathlib4-bot' ] || [ "${AUTHOR}" == 'leanprover-community-bot-assistant' ] then printf $'bot=true\n' - printf $'bot=true\n' > "${GITHUB_OUTPUT}" + printf $'bot=true\n' >> "${GITHUB_OUTPUT}" else printf $'bot=false\n' - printf $'bot=false\n' > "${GITHUB_OUTPUT}" + printf $'bot=false\n' >> "${GITHUB_OUTPUT}" fi - name: Check whether user is a mathlib admin From 9f24be2a7684317446564b4f8656fe70b308c425 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 20 Nov 2024 19:44:13 +0000 Subject: [PATCH 79/90] feat(AlgebraicGeometry): Proj is separated (#19290) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Patience Ablett Co-authored-by: Kevin Buzzard Co-authored-by: Harald Carlens Co-authored-by: Wayne Ng Kwing King Co-authored-by: Michael Schlößer Co-authored-by: Justus Springer Co-authored-by: Jujian Zhang This contribution was created as part of the Durham Computational Algebraic Geometry Workshop Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../ProjectiveSpectrum/Basic.lean | 84 ++++++++++++ .../ProjectiveSpectrum/Proper.lean | 126 ++++++++++++++++++ .../ProjectiveSpectrum/Scheme.lean | 10 ++ .../HomogeneousLocalization.lean | 90 +++++++++++++ .../RingTheory/Localization/Away/Basic.lean | 9 ++ 6 files changed, 320 insertions(+) create mode 100644 Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean diff --git a/Mathlib.lean b/Mathlib.lean index ec30d813df6af..14b5b0c4d7f7a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -966,6 +966,7 @@ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Module import Mathlib.AlgebraicGeometry.PrimeSpectrum.Noetherian import Mathlib.AlgebraicGeometry.PrimeSpectrum.TensorProduct import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Basic +import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Proper import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Scheme import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.StructureSheaf import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Topology diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean index b29a137f71eb2..1ebe7c785f7bb 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Basic.lean @@ -166,6 +166,90 @@ lemma awayι_toSpecZero : awayι 𝒜 f f_deg hm ≫ toSpecZero 𝒜 = ← Spec.map_comp, ← Spec.map_comp, ← Spec.map_comp] rfl +variable {f} +variable {m' : ℕ} {g : A} (g_deg : g ∈ 𝒜 m') (hm' : 0 < m') {x : A} (hx : x = f * g) + +@[reassoc] +lemma awayMap_awayToSection : + CommRingCat.ofHom (awayMap 𝒜 g_deg hx) ≫ awayToSection 𝒜 x = + awayToSection 𝒜 f ≫ (Proj 𝒜).presheaf.map (homOfLE (basicOpen_mono _ _ _ ⟨_, hx⟩)).op := by + ext a + apply Subtype.ext + ext ⟨i, hi⟩ + obtain ⟨⟨n, a, ⟨b, hb'⟩, i, rfl : _ = b⟩, rfl⟩ := mk_surjective a + simp only [CommRingCat.forget_obj, CommRingCat.coe_of, CommRingCat.ofHom, CommRingCat.coe_comp_of, + RingHom.coe_comp, Function.comp_apply, homOfLE_leOfHom] + erw [ProjectiveSpectrum.Proj.awayToSection_apply] + rw [val_awayMap_mk, Localization.mk_eq_mk', IsLocalization.map_mk', + ← Localization.mk_eq_mk'] + refine Localization.mk_eq_mk_iff.mpr ?_ + rw [Localization.r_iff_exists] + use 1 + simp only [OneMemClass.coe_one, RingHom.id_apply, one_mul, hx] + ring + +@[reassoc] +lemma basicOpenToSpec_SpecMap_awayMap : + basicOpenToSpec 𝒜 x ≫ Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) = + (Proj 𝒜).homOfLE (basicOpen_mono _ _ _ ⟨_, hx⟩) ≫ basicOpenToSpec 𝒜 f := by + rw [basicOpenToSpec, Category.assoc, ← Spec.map_comp, awayMap_awayToSection, + Spec.map_comp, Scheme.Opens.toSpecΓ_SpecMap_map_assoc] + rfl + +@[reassoc] +lemma SpecMap_awayMap_awayι : + Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) ≫ awayι 𝒜 f f_deg hm = + awayι 𝒜 x (hx ▸ SetLike.mul_mem_graded f_deg g_deg) (hm.trans_le (m.le_add_right m')) := by + rw [awayι, awayι, Iso.eq_inv_comp, basicOpenIsoSpec_hom, basicOpenToSpec_SpecMap_awayMap_assoc, + ← basicOpenIsoSpec_hom _ _ f_deg hm, Iso.hom_inv_id_assoc, Scheme.homOfLE_ι] + +/-- The isomorphism `D₊(f) ×[Proj 𝒜] D₊(g) ≅ D₊(fg)`. -/ +noncomputable +def pullbackAwayιIso : + Limits.pullback (awayι 𝒜 f f_deg hm) (awayι 𝒜 g g_deg hm') ≅ + Spec (CommRingCat.of (Away 𝒜 x)) := + IsOpenImmersion.isoOfRangeEq (Limits.pullback.fst _ _ ≫ awayι 𝒜 f f_deg hm) + (awayι 𝒜 x (hx ▸ SetLike.mul_mem_graded f_deg g_deg) (hm.trans_le (m.le_add_right m'))) <| by + rw [IsOpenImmersion.range_pullback_to_base_of_left] + show ((awayι 𝒜 f _ _).opensRange ⊓ (awayι 𝒜 g _ _).opensRange).1 = (awayι 𝒜 _ _ _).opensRange.1 + rw [opensRange_awayι, opensRange_awayι, opensRange_awayι, ← basicOpen_mul, hx] + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_hom_awayι : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).hom ≫ + awayι 𝒜 x (hx ▸ SetLike.mul_mem_graded f_deg g_deg) (hm.trans_le (m.le_add_right m')) = + Limits.pullback.fst _ _ ≫ awayι 𝒜 f f_deg hm := + IsOpenImmersion.isoOfRangeEq_hom_fac .. + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_hom_SpecMap_awayMap_left : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).hom ≫ + Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) = Limits.pullback.fst _ _ := by + rw [← cancel_mono (awayι 𝒜 f f_deg hm), ← pullbackAwayιIso_hom_awayι, + Category.assoc, SpecMap_awayMap_awayι] + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_hom_SpecMap_awayMap_right : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).hom ≫ + Spec.map (CommRingCat.ofHom (awayMap 𝒜 f_deg (hx.trans (mul_comm _ _)))) = + Limits.pullback.snd _ _ := by + rw [← cancel_mono (awayι 𝒜 g g_deg hm'), ← Limits.pullback.condition, + ← pullbackAwayιIso_hom_awayι 𝒜 f_deg hm g_deg hm' hx, + Category.assoc, SpecMap_awayMap_awayι] + rfl + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_inv_fst : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).inv ≫ Limits.pullback.fst _ _ = + Spec.map (CommRingCat.ofHom (awayMap 𝒜 g_deg hx)) := by + rw [← pullbackAwayιIso_hom_SpecMap_awayMap_left, Iso.inv_hom_id_assoc] + +@[reassoc (attr := simp)] +lemma pullbackAwayιIso_inv_snd : + (pullbackAwayιIso 𝒜 f_deg hm g_deg hm' hx).inv ≫ Limits.pullback.snd _ _ = + Spec.map (CommRingCat.ofHom (awayMap 𝒜 f_deg (hx.trans (mul_comm _ _)))) := by + rw [← pullbackAwayιIso_hom_SpecMap_awayMap_right, Iso.inv_hom_id_assoc] + open TopologicalSpace.Opens in /-- Given a family of homogeneous elements `f` of positive degree that spans the irrelavent ideal, `Spec (A_f)₀ ⟶ Proj A` forms an affine open cover of `Proj A`. -/ diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean new file mode 100644 index 0000000000000..6235852b32c97 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean @@ -0,0 +1,126 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patience Ablett, Kevin Buzzard, Harald Carlens, Wayne Ng Kwing King, Michael Schlößer, + Justus Springer, Andrew Yang, Jujian Zhang +-/ +import Mathlib.AlgebraicGeometry.Morphisms.Separated +import Mathlib.AlgebraicGeometry.ProjectiveSpectrum.Basic + +/-! +# Properness of `Proj A` + +We show that `Proj 𝒜` is a separated scheme. + +## TODO +- Show that `Proj 𝒜` satisfies the valuative criterion. + +## Notes +This contribution was created as part of the Durham Computational Algebraic Geometry Workshop + +-/ + +namespace AlgebraicGeometry.Proj + +variable {R A : Type*} +variable [CommRing R] [CommRing A] [Algebra R A] +variable (𝒜 : ℕ → Submodule R A) +variable [GradedAlgebra 𝒜] + +open Scheme CategoryTheory Limits pullback HomogeneousLocalization + +lemma lift_awayMapₐ_awayMapₐ_surjective {d e : ℕ} {f : A} (hf : f ∈ 𝒜 d) + {g : A} (hg : g ∈ 𝒜 e) {x : A} (hx : x = f * g) (hd : 0 < d) : + Function.Surjective + (Algebra.TensorProduct.lift (awayMapₐ 𝒜 hg hx) (awayMapₐ 𝒜 hf (hx.trans (mul_comm _ _))) + (fun _ _ ↦ .all _ _)) := by + intro z + obtain ⟨⟨n, ⟨a, ha⟩, ⟨b, hb'⟩, ⟨j, (rfl : _ = b)⟩⟩, rfl⟩ := mk_surjective z + by_cases hfg : (f * g) ^ j = 0 + · use 0 + have := HomogeneousLocalization.subsingleton 𝒜 (x := Submonoid.powers x) ⟨j, hx ▸ hfg⟩ + exact this.elim _ _ + have : n = j * (d + e) := by + apply DirectSum.degree_eq_of_mem_mem 𝒜 hb' + convert SetLike.pow_mem_graded _ _ using 2 + · infer_instance + · exact hx ▸ SetLike.mul_mem_graded hf hg + · exact hx ▸ hfg + let x0 : NumDenSameDeg 𝒜 (.powers f) := + { deg := j * (d * (e + 1)) + num := ⟨a * g ^ (j * (d - 1)), by + convert SetLike.mul_mem_graded ha (SetLike.pow_mem_graded _ hg) using 2 + rw [this] + cases d + · contradiction + · simp; ring⟩ + den := ⟨f ^ (j * (e + 1)), by convert SetLike.pow_mem_graded _ hf using 2; ring⟩ + den_mem := ⟨_,rfl⟩ } + let y0 : NumDenSameDeg 𝒜 (.powers g) := + { deg := j * (d * e) + num := ⟨f ^ (j * e), by convert SetLike.pow_mem_graded _ hf using 2; ring⟩ + den := ⟨g ^ (j * d), by convert SetLike.pow_mem_graded _ hg using 2; ring⟩ + den_mem := ⟨_,rfl⟩ } + use mk x0 ⊗ₜ mk y0 + ext + simp only [Algebra.TensorProduct.lift_tmul, awayMapₐ_apply, val_mul, + val_awayMap_mk, Localization.mk_mul, val_mk, x0, y0] + rw [Localization.mk_eq_mk_iff, Localization.r_iff_exists] + use 1 + simp only [OneMemClass.coe_one, one_mul, Submonoid.mk_mul_mk] + cases d + · contradiction + · simp only [hx, add_tsub_cancel_right] + ring + +open TensorProduct in +instance isSeparated : IsSeparated (toSpecZero 𝒜) := by + refine ⟨IsLocalAtTarget.of_openCover (Pullback.openCoverOfLeftRight + (affineOpenCover 𝒜).openCover (affineOpenCover 𝒜).openCover _ _) ?_⟩ + intro ⟨i, j⟩ + dsimp [Scheme, Cover.pullbackHom] + refine (MorphismProperty.cancel_left_of_respectsIso (P := @IsClosedImmersion) + (f := (pullbackDiagonalMapIdIso ..).inv) _).mp ?_ + let e₁ : pullback ((affineOpenCover 𝒜).map i ≫ toSpecZero 𝒜) + ((affineOpenCover 𝒜).map j ≫ toSpecZero 𝒜) ≅ + Spec (.of <| TensorProduct (𝒜 0) (Away 𝒜 i.2) (Away 𝒜 j.2)) := by + refine pullback.congrHom ?_ ?_ ≪≫ pullbackSpecIso (𝒜 0) (Away 𝒜 i.2) (Away 𝒜 j.2) + · simp [affineOpenCover, openCoverOfISupEqTop, awayι_toSpecZero]; rfl + · simp [affineOpenCover, openCoverOfISupEqTop, awayι_toSpecZero]; rfl + let e₂ : pullback ((affineOpenCover 𝒜).map i) ((affineOpenCover 𝒜).map j) ≅ + Spec (.of <| (Away 𝒜 (i.2 * j.2))) := + pullbackAwayιIso 𝒜 _ _ _ _ rfl + rw [← MorphismProperty.cancel_right_of_respectsIso (P := @IsClosedImmersion) _ e₁.hom, + ← MorphismProperty.cancel_left_of_respectsIso (P := @IsClosedImmersion) e₂.inv] + let F : Away 𝒜 i.2.1 ⊗[𝒜 0] Away 𝒜 j.2.1 →+* Away 𝒜 (i.2.1 * j.2.1) := + (Algebra.TensorProduct.lift (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) + (fun _ _ ↦ .all _ _)).toRingHom + have : Function.Surjective F := lift_awayMapₐ_awayMapₐ_surjective 𝒜 i.2.2 j.2.2 rfl i.1.2 + convert IsClosedImmersion.spec_of_surjective + (CommRingCat.ofHom (R := Away 𝒜 i.2.1 ⊗[𝒜 0] Away 𝒜 j.2.1) F) this using 1 + rw [← cancel_mono (pullbackSpecIso ..).inv] + apply pullback.hom_ext + · simp only [Iso.trans_hom, congrHom_hom, Category.assoc, Iso.hom_inv_id, Category.comp_id, + limit.lift_π, id_eq, eq_mpr_eq_cast, PullbackCone.mk_pt, PullbackCone.mk_π_app, e₂, e₁, + pullbackDiagonalMapIdIso_inv_snd_fst, AlgHom.toRingHom_eq_coe, pullbackSpecIso_inv_fst, + ← Spec.map_comp] + erw [pullbackAwayιIso_inv_fst] + congr 1 + ext x + exact DFunLike.congr_fun (Algebra.TensorProduct.lift_comp_includeLeft + (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) (fun _ _ ↦ .all _ _)).symm x + · simp only [Iso.trans_hom, congrHom_hom, Category.assoc, Iso.hom_inv_id, Category.comp_id, + limit.lift_π, id_eq, eq_mpr_eq_cast, PullbackCone.mk_pt, PullbackCone.mk_π_app, + pullbackDiagonalMapIdIso_inv_snd_snd, AlgHom.toRingHom_eq_coe, pullbackSpecIso_inv_snd, ← + Spec.map_comp, e₂, e₁] + erw [pullbackAwayιIso_inv_snd] + congr 1 + ext x + exact DFunLike.congr_fun (Algebra.TensorProduct.lift_comp_includeRight + (awayMapₐ 𝒜 j.2.2 rfl) (awayMapₐ 𝒜 i.2.2 (mul_comm _ _)) (fun _ _ ↦ .all _ _)).symm x + +@[stacks 01MC] +instance : Scheme.IsSeparated (Proj 𝒜) := + (HasAffineProperty.iff_of_isAffine (P := @IsSeparated)).mp (isSeparated 𝒜) + +end AlgebraicGeometry.Proj diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 462689531697b..781f0783739f0 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -603,6 +603,16 @@ lemma awayToSection_germ (f x hx) : apply (Proj.stalkIso' 𝒜 x).eq_symm_apply.mpr apply Proj.stalkIso'_germ +lemma awayToSection_apply (f : A) (x p) : + (((ProjectiveSpectrum.Proj.awayToSection 𝒜 f).1 x).val p).val = + IsLocalization.map (M := Submonoid.powers f) (T := p.1.1.toIdeal.primeCompl) _ + (RingHom.id _) (Submonoid.powers_le.mpr p.2) x.val := by + obtain ⟨x, rfl⟩ := HomogeneousLocalization.mk_surjective x + show (HomogeneousLocalization.mapId 𝒜 _ _).val = _ + dsimp [HomogeneousLocalization.mapId, HomogeneousLocalization.map] + rw [Localization.mk_eq_mk', Localization.mk_eq_mk', IsLocalization.map_mk'] + rfl + /-- The ring map from `A⁰_ f` to the global sections of the structure sheaf of the projective spectrum of `A` restricted to the basic open set `D(f)`. diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index dcab8f6f958a6..af6dc235b4928 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -6,6 +6,7 @@ Authors: Jujian Zhang, Eric Wieser import Mathlib.Order.Filter.AtTopBot import Mathlib.RingTheory.Localization.AtPrime import Mathlib.RingTheory.GradedAlgebra.Basic +import Mathlib.RingTheory.Localization.Away.Basic /-! # Homogeneous Localization @@ -473,6 +474,11 @@ def fromZeroRingHom : 𝒜 0 →+* HomogeneousLocalization 𝒜 x where map_zero' := rfl map_add' f g := by ext; simp [Localization.add_mk, add_comm f.1 g.1] +instance : Algebra (𝒜 0) (HomogeneousLocalization 𝒜 x) := + (fromZeroRingHom 𝒜 x).toAlgebra + +lemma algebraMap_eq : algebraMap (𝒜 0) (HomogeneousLocalization 𝒜 x) = fromZeroRingHom 𝒜 x := rfl + end HomogeneousLocalization namespace HomogeneousLocalization @@ -627,4 +633,88 @@ lemma map_mk (g : A →+* B) end +section mapAway + +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] +variable {d e : ι} {f : A} (hf : f ∈ 𝒜 d) {g : A} (hg : g ∈ 𝒜 e) +variable {x : A} (hx : x = f * g) + +variable (𝒜) + +/-- Given `f ∣ x`, this is the map `A_{(f)} → A_f → A_x`. We will lift this to a map +`A_{(f)} → A_{(x)}` in `awayMap`. -/ +private def awayMapAux (hx : f ∣ x) : Away 𝒜 f →+* Localization.Away x := + (Localization.awayLift (algebraMap A _) _ + (isUnit_of_dvd_unit (map_dvd _ hx) (IsLocalization.Away.algebraMap_isUnit x))).comp + (algebraMap (Away 𝒜 f) (Localization.Away f)) + +lemma awayMapAux_mk (n a i hi) : + awayMapAux 𝒜 ⟨_, hx⟩ (mk ⟨n, a, ⟨f ^ i, hi⟩, ⟨i, rfl⟩⟩) = + Localization.mk (a * g ^ i) ⟨x ^ i, (Submonoid.mem_powers_iff _ _).mpr ⟨i, rfl⟩⟩ := by + have : algebraMap A (Localization.Away x) f * + (Localization.mk g ⟨f * g, (Submonoid.mem_powers_iff _ _).mpr ⟨1, by simp [hx]⟩⟩) = 1 := by + rw [← Algebra.smul_def, Localization.smul_mk] + exact Localization.mk_self ⟨f*g, _⟩ + simp [awayMapAux] + rw [Localization.awayLift_mk (hv := this), ← Algebra.smul_def, + Localization.mk_pow, Localization.smul_mk] + subst hx + rfl + +include hg in +lemma range_awayMapAux_subset : + Set.range (awayMapAux 𝒜 (f := f) ⟨_, hx⟩) ⊆ Set.range (val (𝒜 := 𝒜)) := by + rintro _ ⟨z, rfl⟩ + obtain ⟨⟨n, ⟨a, ha⟩, ⟨b, hb'⟩, j, rfl : _ = b⟩, rfl⟩ := mk_surjective z + use mk ⟨n+j•e,⟨a*g^j, ?_⟩ ,⟨x^j, ?_⟩, j, rfl⟩ + · simp [awayMapAux_mk 𝒜 (hx := hx)] + · apply SetLike.mul_mem_graded ha + exact SetLike.pow_mem_graded _ hg + · rw [hx, mul_pow] + apply SetLike.mul_mem_graded hb' + exact SetLike.pow_mem_graded _ hg + +/-- Given `x = f * g` with `g` homogeneous of positive degree, +this is the map `A_{(f)} → A_{(x)}` taking `a/f^i` to `ag^i/(fg)^i`. -/ +def awayMap : Away 𝒜 f →+* Away 𝒜 x := by + let e := RingEquiv.ofLeftInverse (f := algebraMap (Away 𝒜 x) (Localization.Away x)) + (h := (val_injective _).hasLeftInverse.choose_spec) + refine RingHom.comp (e.symm.toRingHom.comp (Subring.inclusion ?_)) + (awayMapAux 𝒜 (f := f) ⟨_, hx⟩).rangeRestrict + exact range_awayMapAux_subset 𝒜 hg hx + +lemma val_awayMap_eq_aux (a) : (awayMap 𝒜 hg hx a).val = awayMapAux 𝒜 ⟨_, hx⟩ a := by + let e := RingEquiv.ofLeftInverse (f := algebraMap (Away 𝒜 x) (Localization.Away x)) + (h := (val_injective _).hasLeftInverse.choose_spec) + dsimp [awayMap] + convert_to (e (e.symm ⟨awayMapAux 𝒜 (f := f) ⟨_, hx⟩ a, + range_awayMapAux_subset 𝒜 hg hx ⟨_, rfl⟩⟩)).1 = _ + rw [e.apply_symm_apply] + +lemma val_awayMap (a) : (awayMap 𝒜 hg hx a).val = Localization.awayLift (algebraMap A _) _ + (isUnit_of_dvd_unit (map_dvd _ ⟨_, hx⟩) (IsLocalization.Away.algebraMap_isUnit x)) a.val := by + rw [val_awayMap_eq_aux] + rfl + +lemma awayMap_fromZeroRingHom (a) : + awayMap 𝒜 hg hx (fromZeroRingHom 𝒜 _ a) = fromZeroRingHom 𝒜 _ a := by + ext + simp only [fromZeroRingHom, RingHom.coe_mk, MonoidHom.coe_mk, OneHom.coe_mk, + val_awayMap, val_mk, SetLike.GradeZero.coe_one] + convert IsLocalization.lift_eq _ _ + +lemma val_awayMap_mk (n a i hi) : (awayMap 𝒜 hg hx (mk ⟨n, a, ⟨f ^ i, hi⟩, ⟨i, rfl⟩⟩)).val = + Localization.mk (a * g ^ i) ⟨x ^ i, (Submonoid.mem_powers_iff _ _).mpr ⟨i, rfl⟩⟩ := by + rw [val_awayMap_eq_aux, awayMapAux_mk 𝒜 (hx := hx)] + +/-- Given `x = f * g` with `g` homogeneous of positive degree, +this is the map `A_{(f)} → A_{(x)}` taking `a/f^i` to `ag^i/(fg)^i`. -/ +def awayMapₐ : Away 𝒜 f →ₐ[𝒜 0] Away 𝒜 x where + __ := awayMap 𝒜 hg hx + commutes' _ := awayMap_fromZeroRingHom .. + +@[simp] lemma awayMapₐ_apply (a) : awayMapₐ 𝒜 hg hx a = awayMap 𝒜 hg hx a := rfl + +end mapAway + end HomogeneousLocalization diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index d587619c72796..01a01a331e08b 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -361,6 +361,15 @@ noncomputable abbrev awayLift (f : R →+* P) (r : R) (hr : IsUnit (f r)) : Localization.Away r →+* P := IsLocalization.Away.lift r hr +lemma awayLift_mk {A : Type*} [CommRing A] (f : R →+* A) (r : R) + (a : R) (v : A) (hv : f r * v = 1) (j : ℕ) : + Localization.awayLift f r (isUnit_iff_exists_inv.mpr ⟨v, hv⟩) + (Localization.mk a ⟨r ^ j, j, rfl⟩) = f a * v ^ j := by + rw [Localization.mk_eq_mk'] + erw [IsLocalization.lift_mk'] + rw [Units.mul_inv_eq_iff_eq_mul] + simp [IsUnit.liftRight, mul_assoc, ← mul_pow, (mul_comm _ _).trans hv] + /-- Given a map `f : R →+* S` and an element `r : R`, we may construct a map `Rᵣ →+* Sᵣ`. -/ noncomputable abbrev awayMap (f : R →+* P) (r : R) : Localization.Away r →+* Localization.Away (f r) := From bb8094aa6186788cfc45f171d49e69a785cadb5f Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Wed, 20 Nov 2024 20:06:57 +0000 Subject: [PATCH 80/90] feat(Topology/MetricSpace/Pseudo/Defs): add `dense_iff_iUnion_ball` and `dist_eq_of_dist_zero` (#19294) Upstreamed from the [Carleson](https://github.com/fpvandoorn/carleson) project. Co-authored-by: Floris van Doorn --- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 82a80a1ff7f93..6f5392e44495c 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -1080,6 +1080,10 @@ theorem tendsto_iff_of_dist {f₁ f₂ : ι → α} {p : Filter ι} {a : α} end Real +theorem PseudoMetricSpace.dist_eq_of_dist_zero (x : α) {y z : α} (h : dist y z = 0) : + dist x y = dist x z := + dist_comm y x ▸ dist_comm z x ▸ sub_eq_zero.1 (abs_nonpos_iff.1 (h ▸ abs_dist_sub_le y z x)) + -- Porting note: 3 new lemmas theorem dist_dist_dist_le_left (x y z : α) : dist (dist x z) (dist y z) ≤ dist x y := abs_dist_sub_le .. @@ -1127,6 +1131,10 @@ theorem dense_iff {s : Set α} : Dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s). forall_congr' fun x => by simp only [mem_closure_iff, Set.Nonempty, exists_prop, mem_inter_iff, mem_ball', and_comm] +theorem dense_iff_iUnion_ball (s : Set α) : Dense s ↔ ∀ r > 0, ⋃ c ∈ s, ball c r = univ := by + simp_rw [eq_univ_iff_forall, mem_iUnion, exists_prop, mem_ball, Dense, mem_closure_iff, + forall_comm (α := α)] + theorem denseRange_iff {f : β → α} : DenseRange f ↔ ∀ x, ∀ r > 0, ∃ y, dist x (f y) < r := forall_congr' fun x => by simp only [mem_closure_iff, exists_range_iff] From 3c79dac7ee6ac164a9919f3e21ab38d3e05aef82 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Wed, 20 Nov 2024 20:34:20 +0000 Subject: [PATCH 81/90] feat (LinearAlgebra/RootSystem/Finite): nondegeneracy of canonical bilinear form restricted to root span (#18569) The canonical bilinear form for a root pairing over a linearly ordered commutative ring is nondegenerate when restricted to the span of roots. Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 + .../Order/BigOperators/Ring/Finset.lean | 14 +- Mathlib/LinearAlgebra/Dimension/Basic.lean | 7 + .../LinearAlgebra/Dimension/RankNullity.lean | 12 +- .../Dimension/StrongRankCondition.lean | 8 ++ Mathlib/LinearAlgebra/Dual.lean | 11 ++ Mathlib/LinearAlgebra/RootSystem/Defs.lean | 6 + .../RootSystem/Finite/CanonicalBilinear.lean | 64 ++++++++- .../RootSystem/Finite/Nondegenerate.lean | 128 ++++++++++++++++++ docs/references.bib | 15 ++ 10 files changed, 257 insertions(+), 9 deletions(-) create mode 100644 Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean diff --git a/Mathlib.lean b/Mathlib.lean index 14b5b0c4d7f7a..5922aa04576e9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3391,6 +3391,7 @@ import Mathlib.LinearAlgebra.Reflection import Mathlib.LinearAlgebra.RootSystem.Basic import Mathlib.LinearAlgebra.RootSystem.Defs import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear +import Mathlib.LinearAlgebra.RootSystem.Finite.Nondegenerate import Mathlib.LinearAlgebra.RootSystem.Hom import Mathlib.LinearAlgebra.RootSystem.OfBilinear import Mathlib.LinearAlgebra.RootSystem.RootPairingCat diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index 4bfa3ea69c30d..b51accd2d82ad 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -95,7 +95,6 @@ lemma sum_sq_le_sq_sum_of_nonneg (hf : ∀ i ∈ s, 0 ≤ f i) : · exact hf i hi · exact single_le_sum hf hi - end OrderedSemiring section OrderedCommSemiring @@ -142,6 +141,19 @@ lemma sum_mul_sq_le_sq_mul_sq (s : Finset ι) (f g : ι → R) : sum_le_sum fun i _ ↦ two_mul_le_add_sq (f i * ∑ j ∈ s, g j ^ 2) (g i * ∑ j ∈ s, f j * g j) _ = _ := by simp_rw [sum_add_distrib, mul_pow, ← sum_mul]; ring +theorem sum_mul_self_eq_zero_iff (s : Finset ι) (f : ι → R) : + ∑ i ∈ s, f i * f i = 0 ↔ ∀ i ∈ s, f i = 0 := by + induction s using Finset.cons_induction with + | empty => simp + | cons i s his ih => + simp only [Finset.sum_cons, Finset.mem_cons, forall_eq_or_imp] + refine ⟨fun hc => ?_, fun h => by simpa [h.1] using ih.mpr h.2⟩ + have hi : f i * f i ≤ 0 := by + rw [← hc, le_add_iff_nonneg_right] + exact Finset.sum_nonneg fun i _ ↦ mul_self_nonneg (f i) + have h : f i * f i = 0 := (eq_of_le_of_le (mul_self_nonneg (f i)) hi).symm + exact ⟨zero_eq_mul_self.mp h.symm, ih.mp (by rw [← hc, h, zero_add])⟩ + end LinearOrderedCommSemiring lemma abs_prod [LinearOrderedCommRing R] (s : Finset ι) (f : ι → R) : diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 76b2ee39e5e51..254ddee6fae30 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -353,6 +353,13 @@ theorem rank_subsingleton [Subsingleton R] : Module.rank R M = 1 := by subsingleton · exact hw.trans_eq (Cardinal.mk_singleton _).symm +lemma rank_le_of_smul_regular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] + [IsScalarTower S R M] (L L' : Submodule R M) {s : S} (hr : IsSMulRegular M s) + (h : ∀ x ∈ L, s • x ∈ L') : + Module.rank R L ≤ Module.rank R L' := + ((Algebra.lsmul S R M s).restrict h).rank_le_of_injective <| + fun _ _ h ↦ by simpa using hr (Subtype.ext_iff.mp h) + end end Module diff --git a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean index 79f9779cf6271..cf3c16a79daa0 100644 --- a/Mathlib/LinearAlgebra/Dimension/RankNullity.lean +++ b/Mathlib/LinearAlgebra/Dimension/RankNullity.lean @@ -203,13 +203,23 @@ lemma Submodule.finrank_quotient_add_finrank [Module.Finite R M] (N : Submodule Submodule.finrank_eq_rank] exact HasRankNullity.rank_quotient_add_rank _ - /-- Rank-nullity theorem using `finrank` and subtraction. -/ lemma Submodule.finrank_quotient [Module.Finite R M] {S : Type*} [Ring S] [SMul R S] [Module S M] [IsScalarTower R S M] (N : Submodule S M) : finrank R (M ⧸ N) = finrank R M - finrank R N := by rw [← (N.restrictScalars R).finrank_quotient_add_finrank] exact Nat.eq_sub_of_add_eq rfl +lemma Submodule.disjoint_ker_of_finrank_eq [NoZeroSMulDivisors R M] {N : Type*} [AddCommGroup N] + [Module R N] {L : Submodule R M} [Module.Finite R L] (f : M →ₗ[R] N) + (h : finrank R (L.map f) = finrank R L) : + Disjoint L (LinearMap.ker f) := by + refine disjoint_iff.mpr <| LinearMap.injective_domRestrict_iff.mp <| LinearMap.ker_eq_bot.mp <| + Submodule.rank_eq_zero.mp ?_ + rw [← Submodule.finrank_eq_rank, Nat.cast_eq_zero] + rw [← LinearMap.range_domRestrict] at h + have := (LinearMap.ker (f.domRestrict L)).finrank_quotient_add_finrank + rwa [LinearEquiv.finrank_eq (f.domRestrict L).quotKerEquivRange, h, Nat.add_eq_left] at this + end Finrank section diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index cd51f9aec9d87..0e7836ca78bb4 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -465,4 +465,12 @@ theorem LinearMap.finrank_range_le [Module.Finite R M] (f : M →ₗ[R] M') : finrank R (LinearMap.range f) ≤ finrank R M := finrank_le_finrank_of_rank_le_rank (lift_rank_range_le f) (rank_lt_aleph0 _ _) +theorem LinearMap.finrank_le_of_smul_regular {S : Type*} [CommSemiring S] [Algebra S R] [Module S M] + [IsScalarTower S R M] (L L' : Submodule R M) [Module.Finite R L'] {s : S} + (hr : IsSMulRegular M s) (h : ∀ x ∈ L, s • x ∈ L') : + Module.finrank R L ≤ Module.finrank R L' := by + refine finrank_le_finrank_of_rank_le_rank (lift_le.mpr <| rank_le_of_smul_regular L L' hr h) ?_ + rw [← Module.finrank_eq_rank R L'] + exact nat_lt_aleph0 (finrank R ↥L') + end StrongRankCondition diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index c41abe513c0a7..1ef8aede9c87b 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -715,6 +715,17 @@ instance instFiniteDimensionalOfIsReflexive (K V : Type*) rw [lift_id'] exact lt_trans h₁ h₂ +instance [IsDomain R] : NoZeroSMulDivisors R M := by + refine (noZeroSMulDivisors_iff R M).mpr ?_ + intro r m hrm + rw [or_iff_not_imp_left] + intro hr + suffices Dual.eval R M m = Dual.eval R M 0 from (bijective_dual_eval R M).injective this + ext n + simp only [Dual.eval_apply, map_zero, LinearMap.zero_apply] + suffices r • n m = 0 from eq_zero_of_ne_zero_of_mul_left_eq_zero hr this + rw [← LinearMap.map_smul_of_tower, hrm, LinearMap.map_zero] + end IsReflexive end Module diff --git a/Mathlib/LinearAlgebra/RootSystem/Defs.lean b/Mathlib/LinearAlgebra/RootSystem/Defs.lean index 420de4b340130..d9ab0eb932455 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Defs.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Defs.lean @@ -374,6 +374,12 @@ lemma isReduced_iff : P.IsReduced ↔ ∀ i j : ι, i ≠ j → · exact Or.inl (congrArg P.root h') · exact Or.inr (h i j h' hLin) +/-- The linear span of roots. -/ +abbrev rootSpan := span R (range P.root) + +/-- The linear span of coroots. -/ +abbrev corootSpan := span R (range P.coroot) + /-- The `Weyl group` of a root pairing is the group of automorphisms of the weight space generated by reflections in roots. -/ def weylGroup : Subgroup (M ≃ₗ[R] M) := diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean index d495f1b08bbcc..1f7676b71d4ad 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/CanonicalBilinear.lean @@ -3,8 +3,9 @@ Copyright (c) 2024 Scott Carnahan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Carnahan -/ -import Mathlib.LinearAlgebra.RootSystem.Defs +import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Ring.SumsOfSquares +import Mathlib.LinearAlgebra.RootSystem.Defs /-! # The canonical bilinear form on a finite root pairing @@ -22,27 +23,28 @@ Weyl group. * `Polarization`: A distinguished linear map from the weight space to the coweight space. * `RootForm` : The bilinear form on weight space corresponding to `Polarization`. -## References: - * SGAIII Exp. XXI - * Bourbaki, Lie groups and Lie algebras - ## Main results: * `polarization_self_sum_of_squares` : The inner product of any weight vector is a sum of squares. * `rootForm_reflection_reflection_apply` : `RootForm` is invariant with respect to reflections. * `rootForm_self_smul_coroot`: The inner product of a root with itself times the corresponding coroot is equal to two times Polarization applied to the root. + * `rootForm_self_non_neg`: `RootForm` is positive semidefinite. + +## References: + * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] + * [M. Demazure, *SGA III, Expos\'{e} XXI, Don\'{e}es Radicielles*][demazure1970] ## TODO (possibly in other files) - * Positivity and nondegeneracy * Weyl-invariance * Faithfulness of Weyl group action, and finiteness of Weyl group, for finite root systems. * Relation to Coxeter weight. In particular, positivity constraints for finite root pairings mean we restrict to weights between 0 and 4. -/ -open Function +open Set Function open Module hiding reflection +open Submodule (span) noncomputable section @@ -132,6 +134,54 @@ lemma rootForm_root_self (j : ι) : P.RootForm (P.root j) (P.root j) = ∑ (i : ι), (P.pairing j i) * (P.pairing j i) := by simp [rootForm_apply_apply] +theorem range_polarization_domRestrict_le_span_coroot : + LinearMap.range (P.Polarization.domRestrict P.rootSpan) ≤ P.corootSpan := by + intro y hy + obtain ⟨x, hx⟩ := hy + rw [← hx, LinearMap.domRestrict_apply, Polarization_apply] + refine (mem_span_range_iff_exists_fun R).mpr ?_ + use fun i => (P.toPerfectPairing x) (P.coroot i) + simp + end CommRing +section LinearOrderedCommRing + +variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] + [Module R N] (P : RootPairing ι R M N) + +theorem rootForm_self_non_neg (x : M) : 0 ≤ P.RootForm x x := + IsSumSq.nonneg (P.rootForm_self_sum_of_squares x) + +theorem rootForm_self_zero_iff (x : M) : + P.RootForm x x = 0 ↔ ∀ i, P.coroot' i x = 0 := by + simp only [rootForm_apply_apply, PerfectPairing.toLin_apply, LinearMap.coe_comp, comp_apply, + Polarization_apply, map_sum, map_smul, smul_eq_mul] + convert Finset.sum_mul_self_eq_zero_iff Finset.univ fun i => P.coroot' i x + simp + +lemma rootForm_root_self_pos (j : ι) : + 0 < P.RootForm (P.root j) (P.root j) := by + simp only [LinearMap.coe_mk, AddHom.coe_mk, LinearMap.coe_comp, comp_apply, + rootForm_apply_apply, toLin_toPerfectPairing] + refine Finset.sum_pos' (fun i _ => (sq (P.pairing j i)) ▸ sq_nonneg (P.pairing j i)) ?_ + use j + simp + +lemma prod_rootForm_root_self_pos : + 0 < ∏ i, P.RootForm (P.root i) (P.root i) := + Finset.prod_pos fun i _ => rootForm_root_self_pos P i + +lemma prod_rootForm_smul_coroot_mem_range_domRestrict (i : ι) : + (∏ a : ι, P.RootForm (P.root a) (P.root a)) • P.coroot i ∈ + LinearMap.range (P.Polarization.domRestrict (P.rootSpan)) := by + obtain ⟨c, hc⟩ := Finset.dvd_prod_of_mem (fun a ↦ P.RootForm (P.root a) (P.root a)) + (Finset.mem_univ i) + rw [hc, mul_comm, mul_smul, rootForm_self_smul_coroot] + refine LinearMap.mem_range.mpr ?_ + use ⟨(c • 2 • P.root i), by aesop⟩ + simp + +end LinearOrderedCommRing + end RootPairing diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean new file mode 100644 index 0000000000000..734860e88627e --- /dev/null +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -0,0 +1,128 @@ +/- +Copyright (c) 2024 Scott Carnahan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Carnahan +-/ +import Mathlib.LinearAlgebra.BilinearForm.Basic +import Mathlib.LinearAlgebra.Dimension.Localization +import Mathlib.LinearAlgebra.QuadraticForm.Basic +import Mathlib.LinearAlgebra.RootSystem.Finite.CanonicalBilinear +import Mathlib.LinearAlgebra.RootSystem.RootPositive + +/-! +# Nondegeneracy of the polarization on a finite root pairing + +We show that if the base ring of a finite root pairing is linearly ordered, then the canonical +bilinear form is root-positive and positive-definite on the span of roots. +From these facts, it is easy to show that Coxeter weights in a finite root pairing are bounded +above by 4. Thus, the pairings of roots and coroots in a root pairing are restricted to the +interval `[-4, 4]`. Furthermore, a linearly independent pair of roots cannot have Coxeter weight 4. +For the case of crystallographic root pairings, we are thus reduced to a finite set of possible +options for each pair. +Another application is to the faithfulness of the Weyl group action on roots, and finiteness of the +Weyl group. + +## Main results: + * `RootPairing.rootForm_rootPositive`: `RootForm` is root-positive. + * `RootPairing.polarization_domRestrict_injective`: The polarization restricted to the span of + roots is injective. + * `RootPairing.rootForm_pos_of_nonzero`: `RootForm` is strictly positive on non-zero linear + combinations of roots. This gives us a convenient way to eliminate certain Dynkin diagrams from + the classification, since it suffices to produce a nonzero linear combination of simple roots with + non-positive norm. + +## References: + * [N. Bourbaki, *Lie groups and {L}ie algebras. {C}hapters 4--6*][bourbaki1968] + * [M. Demazure, *SGA III, Expos\'{e} XXI, Don\'{e}es Radicielles*][demazure1970] + +## Todo + * Weyl-invariance of `RootForm` and `CorootForm` + * Faithfulness of Weyl group perm action, and finiteness of Weyl group, over ordered rings. + * Relation to Coxeter weight. In particular, positivity constraints for finite root pairings mean + we restrict to weights between 0 and 4. +-/ + +noncomputable section + +open Set Function +open Module hiding reflection +open Submodule (span) + +namespace RootPairing + +variable {ι R M N : Type*} + +variable [Fintype ι] [LinearOrderedCommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] +[Module R N] (P : RootPairing ι R M N) + +lemma rootForm_rootPositive : IsRootPositive P P.RootForm where + zero_lt_apply_root i := P.rootForm_root_self_pos i + symm := P.rootForm_symmetric + apply_reflection_eq := P.rootForm_reflection_reflection_apply + +instance : Module.Finite R P.rootSpan := Finite.span_of_finite R <| finite_range P.root + +instance : Module.Finite R P.corootSpan := Finite.span_of_finite R <| finite_range P.coroot + +@[simp] +lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : + finrank R (P.rootSpan.map P.Polarization) = finrank R P.corootSpan := by + rw [← LinearMap.range_domRestrict] + apply (Submodule.finrank_mono P.range_polarization_domRestrict_le_span_coroot).antisymm + have : IsReflexive R N := PerfectPairing.reflexive_right P.toPerfectPairing + refine LinearMap.finrank_le_of_smul_regular P.corootSpan + (LinearMap.range (P.Polarization.domRestrict P.rootSpan)) + (smul_right_injective N (Ne.symm (ne_of_lt P.prod_rootForm_root_self_pos))) + fun _ hx => ?_ + obtain ⟨c, hc⟩ := (mem_span_range_iff_exists_fun R).mp hx + rw [← hc, Finset.smul_sum] + simp_rw [smul_smul, mul_comm, ← smul_smul] + exact Submodule.sum_smul_mem (LinearMap.range (P.Polarization.domRestrict P.rootSpan)) c + (fun c _ ↦ prod_rootForm_smul_coroot_mem_range_domRestrict P c) + +/-- An auxiliary lemma en route to `RootPairing.finrank_corootSpan_eq`. -/ +private lemma finrank_corootSpan_le : + finrank R P.corootSpan ≤ finrank R P.rootSpan := by + rw [← finrank_rootSpan_map_polarization_eq_finrank_corootSpan] + exact Submodule.finrank_map_le P.Polarization P.rootSpan + +lemma finrank_corootSpan_eq : + finrank R P.corootSpan = finrank R P.rootSpan := + le_antisymm P.finrank_corootSpan_le P.flip.finrank_corootSpan_le + +lemma disjoint_rootSpan_ker_polarization : + Disjoint P.rootSpan (LinearMap.ker P.Polarization) := by + have : IsReflexive R M := PerfectPairing.reflexive_left P.toPerfectPairing + refine Submodule.disjoint_ker_of_finrank_eq (L := P.rootSpan) P.Polarization ?_ + rw [finrank_rootSpan_map_polarization_eq_finrank_corootSpan, finrank_corootSpan_eq] + +lemma mem_ker_polarization_of_rootForm_self_eq_zero {x : M} (hx : P.RootForm x x = 0) : + x ∈ LinearMap.ker P.Polarization := by + rw [LinearMap.mem_ker, Polarization_apply] + rw [rootForm_self_zero_iff] at hx + exact Fintype.sum_eq_zero _ fun i ↦ by simp [hx i] + +lemma eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero {x : M} + (hx : x ∈ P.rootSpan) (hx' : P.RootForm x x = 0) : + x = 0 := by + rw [← Submodule.mem_bot (R := R), ← P.disjoint_rootSpan_ker_polarization.eq_bot] + exact ⟨hx, P.mem_ker_polarization_of_rootForm_self_eq_zero hx'⟩ + +lemma _root_.RootSystem.rootForm_anisotropic (P : RootSystem ι R M N) : + P.RootForm.toQuadraticMap.Anisotropic := + fun x ↦ P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero <| by + simpa only [rootSpan, P.span_eq_top] using Submodule.mem_top + +lemma rootForm_pos_of_nonzero {x : M} (hx : x ∈ P.rootSpan) (h : x ≠ 0) : + 0 < P.RootForm x x := by + apply (P.rootForm_self_non_neg x).lt_of_ne + contrapose! h + exact eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero P hx h.symm + +lemma rootForm_restrict_nondegenerate : + (P.RootForm.restrict P.rootSpan).Nondegenerate := + LinearMap.IsRefl.nondegenerate_of_separatingLeft (LinearMap.IsSymm.isRefl fun x y => by + simp [rootForm_apply_apply, mul_comm]) fun x h => SetLike.coe_eq_coe.mp + (P.eq_zero_of_mem_rootSpan_of_rootForm_self_eq_zero (Submodule.coe_mem x) (h x)) + +end RootPairing diff --git a/docs/references.bib b/docs/references.bib index 8d65ce04a4a4e..6603b6a5466fe 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1090,6 +1090,21 @@ @InProceedings{ deligne_formulaire mrreviewer = {Jacques Velu} } +@InProceedings{ demazure1970, + author = {Michel Demazure}, + editor = {M. Demazure, A. Grothendieck}, + title = {Expos\'{e} XXI, Don\'{e}es Radicielles}, + booktitle = {S\'{e}minaire de G\'{e}ométrie Alg\'{e}brique du Bois + Marie - 1962-64 - Sch\'{e}mas en groupes - (SGA 3) - vol. + 3, Structure des Sch\'{e}mas en Groupes Reductifs}, + series = {Lecture Notes in Mathematics}, + volume = {153}, + pages = {85--155}, + publisher = {Springer-Verlag}, + year = {1970}, + url = {https://wstein.org/sga/SGA3/Expo21-alpha.pdf} +} + @InProceedings{ demoura2015lean, author = {de Moura, Leonardo and Kong, Soonho and Avigad, Jeremy and van Doorn, Floris and von Raumer, Jakob}, From 961ee9dbff2b14a178fa757f097b5bb27c274e99 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 20 Nov 2024 20:59:58 +0000 Subject: [PATCH 82/90] fix: more explicit PR number in `bors x` actions (#19304) --- .github/workflows/maintainer_bors.yml | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 6832b6e007eed..073477b4fa324 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -43,6 +43,9 @@ jobs: printf $'"bors delegate" or "bors merge" found? \'%s\'\n' "${m_or_d}" printf $'AUTHOR: \'%s\'\n' "${AUTHOR}" + printf $'PR_NUMBER: \'%s\'\n' "${PR_NUMBER}" + printf $'OTHER_NUMBER: \'%s\'\n' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" + printf $'%s' "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" | hexdump -cC printf $'mOrD=%s\n' "${m_or_d}" >> "${GITHUB_OUTPUT}" if [ "${AUTHOR}" == 'leanprover-community-mathlib4-bot' ] || @@ -73,9 +76,9 @@ jobs: uses: octokit/request-action@v2.x with: # check is this ok? was /repos/:repository/issues/:issue_number/labels - route: POST /repos/:repository/issues/${PR_NUMBER}/labels + route: POST /repos/:repository/issues/:issue_number/labels repository: ${{ github.repository }} - issue_number: ${PR_NUMBER} + issue_number: ${{ github.event.issue.number }}${{ github.event.pull_request.number }} labels: '["${{ steps.merge_or_delegate.outputs.mOrD }}"]' env: GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} # comment uses ${{ secrets.GITHUB_TOKEN }} From 33aaac3d047090794ee1627d347864bd8d297987 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:46:49 +1100 Subject: [PATCH 83/90] fixes --- Mathlib/Data/List/TFAE.lean | 2 +- Mathlib/Lean/Meta/KAbstractPositions.lean | 2 +- Mathlib/Tactic/FBinop.lean | 12 ++++++++---- Mathlib/Tactic/Linter/PPRoundtrip.lean | 2 +- Mathlib/Tactic/Variable.lean | 4 ++-- lake-manifest.json | 6 +++--- 6 files changed, 16 insertions(+), 12 deletions(-) diff --git a/Mathlib/Data/List/TFAE.lean b/Mathlib/Data/List/TFAE.lean index 24e43d9612f03..55f77a0f094f9 100644 --- a/Mathlib/Data/List/TFAE.lean +++ b/Mathlib/Data/List/TFAE.lean @@ -62,7 +62,7 @@ theorem tfae_of_cycle {a b} {l : List Prop} (h_chain : List.Chain (· → ·) a theorem TFAE.out {l} (h : TFAE l) (n₁ n₂) {a b} (h₁ : List.get? l n₁ = some a := by rfl) (h₂ : List.get? l n₂ = some b := by rfl) : a ↔ b := - h _ (List.get?_mem h₁) _ (List.get?_mem h₂) + h _ (List.mem_of_get? h₁) _ (List.mem_of_get? h₂) /-- If `P₁ x ↔ ... ↔ Pₙ x` for all `x`, then `(∀ x, P₁ x) ↔ ... ↔ (∀ x, Pₙ x)`. Note: in concrete cases, Lean has trouble finding the list `[P₁, ..., Pₙ]` from the list diff --git a/Mathlib/Lean/Meta/KAbstractPositions.lean b/Mathlib/Lean/Meta/KAbstractPositions.lean index 60d3ded29a824..ffb5899f96568 100644 --- a/Mathlib/Lean/Meta/KAbstractPositions.lean +++ b/Mathlib/Lean/Meta/KAbstractPositions.lean @@ -71,7 +71,7 @@ def viewKAbstractSubExpr (e : Expr) (pos : SubExpr.Pos) : MetaM (Option (Expr × if subExpr.hasLooseBVars then return none let positions ← kabstractPositions subExpr e - let some n := positions.getIdx? pos | unreachable! + let some n := positions.indexOf? pos | unreachable! return some (subExpr, if positions.size == 1 then none else some (n + 1)) /-- Determine whether the result of abstracting `subExpr` from `e` at position `pos` results diff --git a/Mathlib/Tactic/FBinop.lean b/Mathlib/Tactic/FBinop.lean index 3eb83fec8d8d6..96c3a9db63cb6 100644 --- a/Mathlib/Tactic/FBinop.lean +++ b/Mathlib/Tactic/FBinop.lean @@ -185,14 +185,18 @@ private def toExprCore (t : Tree) : TermElabM Expr := do | .term _ trees e => modifyInfoState (fun s => { s with trees := s.trees ++ trees }); return e | .binop ref f lhs rhs => - withRef ref <| withInfoContext' ref (mkInfo := mkTermInfo .anonymous ref) do + withRef ref <| withInfoContext' ref (mkInfo := mkTermInfo .anonymous ref) (do let lhs ← toExprCore lhs let mut rhs ← toExprCore rhs - mkBinOp f lhs rhs + mkBinOp f lhs rhs) + -- FIXME the signature of withInfoContext' has changed, not sure yet what to put here: + (do failure) | .macroExpansion macroName stx stx' nested => - withRef stx <| withInfoContext' stx (mkInfo := mkTermInfo macroName stx) do + withRef stx <| withInfoContext' stx (mkInfo := mkTermInfo macroName stx) (do withMacroExpansion stx stx' do - toExprCore nested + toExprCore nested) + -- FIXME the signature of withInfoContext' has changed, not sure yet what to put here: + (do failure) /-- Try to coerce elements in the tree to `maxS` when needed. -/ private def applyCoe (t : Tree) (maxS : SRec) : TermElabM Tree := do diff --git a/Mathlib/Tactic/Linter/PPRoundtrip.lean b/Mathlib/Tactic/Linter/PPRoundtrip.lean index 9f8e258ecd46f..dced1f5293546 100644 --- a/Mathlib/Tactic/Linter/PPRoundtrip.lean +++ b/Mathlib/Tactic/Linter/PPRoundtrip.lean @@ -60,7 +60,7 @@ def polishSource (s : String) : String × Array Nat := let preWS := split.foldl (init := #[]) fun p q => let txt := q.trimLeft.length (p.push (q.length - txt)).push txt - let preWS := preWS.eraseIdx 0 + let preWS := preWS.eraseIdx! 0 let s := (split.map .trimLeft).filter (· != "") (" ".intercalate (s.filter (!·.isEmpty)), preWS) diff --git a/Mathlib/Tactic/Variable.lean b/Mathlib/Tactic/Variable.lean index c506333c3c075..3f885e361c071 100644 --- a/Mathlib/Tactic/Variable.lean +++ b/Mathlib/Tactic/Variable.lean @@ -159,7 +159,7 @@ partial def completeBinders' (maxSteps : Nat) (gas : Nat) (binders : TSyntaxArray ``bracketedBinder) (toOmit : Array Bool) (i : Nat) : TermElabM (TSyntaxArray ``bracketedBinder × Array Bool) := do - if 0 < gas && i < binders.size then + if h : 0 < gas ∧ i < binders.size then let binder := binders[i]! trace[«variable?»] "\ Have {(← getLCtx).getFVarIds.size} fvars and {(← getLocalInstances).size} local instances. \ @@ -176,7 +176,7 @@ partial def completeBinders' (maxSteps : Nat) (gas : Nat) in binders since they are generated by pretty printing unsatisfied dependencies.\n\n\ Current variable command:{indentD (← `(command| variable $binders'*))}\n\n\ Local context for the unsatisfied dependency:{goalMsg}" - let binders := binders.insertAt! i binder' + let binders := binders.insertIdx i binder' completeBinders' maxSteps (gas - 1) checkRedundant binders toOmit i else let lctx ← getLCtx diff --git a/lake-manifest.json b/lake-manifest.json index 0df8444539094..ab6a0270166c8 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0fc755a2a762bd228db724df926de7d3e2306a34", + "rev": "a75e35f6cb83e773231f793d1cd8112e41804934", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", @@ -45,7 +45,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "24d0b61ede545901e80ee866a3609f9e78080034", + "rev": "2c2069c0e09a0fb8c441f1bef9a4c26d67a4fda5", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9762792852464bf9591f6cccfe44f14b3ef54b25", + "rev": "58b15585c1480b5e9afcded83a36b8ea0a7d4a24", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From b831612cbc094765797ab5bd3d96c5faa8e5cb65 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:48:47 +1100 Subject: [PATCH 84/90] fixes --- Mathlib/Tactic/CC/Addition.lean | 2 +- Mathlib/Tactic/ToAdditive/Frontend.lean | 2 +- lake-manifest.json | 4 ++-- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 2b3349c44376a..8548c1ad7b0c7 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -390,7 +390,7 @@ partial def mkCongrProofCore (lhs rhs : Expr) (heqProofs : Bool) : CCM Expr := d guard (kindsIt[0]! matches .eq) let some p ← getEqProof (lhsArgs[i]'hi.2) (rhsArgs[i]'(ha.symm ▸ hi.2)) | failure lemmaArgs := lemmaArgs.push p - kindsIt := kindsIt.eraseIdx 0 + kindsIt := kindsIt.eraseIdx! 0 let mut r := mkAppN specLemma.proof lemmaArgs if specLemma.heqResult && !heqProofs then r ← mkAppM ``eq_of_heq #[r] diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 25e3af975e0bf..b0278e30a9e9d 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -866,7 +866,7 @@ def additivizeLemmas {m : Type → Type} [Monad m] [MonadError m] [MonadLiftT Co for (nm, lemmas) in names.zip auxLemmas do unless lemmas.size == nLemmas do throwError "{names[0]!} and {nm} do not generate the same number of {desc}." - for (srcLemmas, tgtLemmas) in auxLemmas.zip <| auxLemmas.eraseIdx 0 do + for (srcLemmas, tgtLemmas) in auxLemmas.zip <| auxLemmas.eraseIdx! 0 do for (srcLemma, tgtLemma) in srcLemmas.zip tgtLemmas do insertTranslation srcLemma tgtLemma diff --git a/lake-manifest.json b/lake-manifest.json index ab6a0270166c8..4b7b5e02fb6e4 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -45,7 +45,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "2c2069c0e09a0fb8c441f1bef9a4c26d67a4fda5", + "rev": "0c995c64c9c8e4186e24f85d5d61bc404b378ba6", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "58b15585c1480b5e9afcded83a36b8ea0a7d4a24", + "rev": "9b9f4d0406d00baaae62d1a717b5aa854a2ae51d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From c0d40a2b3e63ec06b546f6fc8e1dc047a49ae057 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:51:38 +1100 Subject: [PATCH 85/90] fix --- Mathlib/Tactic/MoveAdd.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 6acfded0d90f7..12f7051a82e2c 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -195,7 +195,8 @@ def reorderUsing (toReorder : List α) (instructions : List (α × Bool)) : List let uToReorder := (uniquify toReorder).toArray let reorder := uToReorder.qsort fun x y => match uInstructions.find? (Prod.fst · == x), uInstructions.find? (Prod.fst · == y) with - | none, none => (uToReorder.getIdx? x).get! ≤ (uToReorder.getIdx? y).get! + | none, none => + ((uToReorder.indexOf? x).map Fin.val).get! ≤ ((uToReorder.indexOf? y).map Fin.val).get! | _, _ => weight uInstructions x ≤ weight uInstructions y (reorder.map Prod.fst).toList From a881b2f14da2c7a06e5a616a461fa208c8475f74 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:54:33 +1100 Subject: [PATCH 86/90] fixes --- Mathlib/Data/List/Cycle.lean | 2 +- Mathlib/Data/List/NodupEquivFin.lean | 2 +- Mathlib/Data/List/Sym.lean | 2 +- Mathlib/Data/Vector/Mem.lean | 4 +--- 4 files changed, 4 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 167f0c1116671..bf97af552b902 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -217,7 +217,7 @@ theorem prev_ne_cons_cons (y z : α) (h : x ∈ y :: z :: l) (hy : x ≠ y) (hz · rw [prev, dif_neg hy, if_neg hz] theorem next_mem (h : x ∈ l) : l.next x h ∈ l := - nextOr_mem (get_mem _ _ _) + nextOr_mem (getElem_mem _) theorem prev_mem (h : x ∈ l) : l.prev x h ∈ l := by cases' l with hd tl diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 65a48a57e4574..fbcf74a9eec72 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -49,7 +49,7 @@ variable [DecidableEq α] the set of elements of `l`. -/ @[simps] def getEquiv (l : List α) (H : Nodup l) : Fin (length l) ≃ { x // x ∈ l } where - toFun i := ⟨get l i, get_mem l i i.2⟩ + toFun i := ⟨get l i, getElem_mem i.2⟩ invFun x := ⟨indexOf (↑x) l, indexOf_lt_length.2 x.2⟩ left_inv i := by simp only [List.get_indexOf, eq_self_iff_true, Fin.eta, Subtype.coe_mk, H] right_inv x := by simp diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index 1d6228b9258e3..b3981a8ba0333 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -170,7 +170,7 @@ theorem dedup_sym2 [DecidableEq α] (xs : List α) : xs.sym2.dedup = xs.dedup.sy obtain hm | hm := Decidable.em (x ∈ xs) · rw [dedup_cons_of_mem hm, ← ih, dedup_cons_of_mem, List.Subset.dedup_append_right (map_mk_sublist_sym2 _ _ hm).subset] - refine mem_append_of_mem_left _ ?_ + refine mem_append_left _ ?_ rw [mem_map] exact ⟨_, hm, Sym2.eq_swap⟩ · rw [dedup_cons_of_not_mem hm, List.sym2, map_cons, ← ih, dedup_cons_of_not_mem, cons_append, diff --git a/Mathlib/Data/Vector/Mem.lean b/Mathlib/Data/Vector/Mem.lean index 3eb0fe6678d82..9cf9359b65623 100644 --- a/Mathlib/Data/Vector/Mem.lean +++ b/Mathlib/Data/Vector/Mem.lean @@ -22,9 +22,7 @@ namespace Vector variable {α β : Type*} {n : ℕ} (a a' : α) @[simp] -theorem get_mem (i : Fin n) (v : Vector α n) : v.get i ∈ v.toList := by - rw [get_eq_get] - exact List.get_mem _ _ _ +theorem get_mem (i : Fin n) (v : Vector α n) : v.get i ∈ v.toList := List.getElem_mem _ theorem mem_iff_get (v : Vector α n) : a ∈ v.toList ↔ ∃ i, v.get i = a := by simp only [List.mem_iff_get, Fin.exists_iff, Vector.get_eq_get] From 4dc87eb228315b4b0f0d80c33c886fd488658c1a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:55:52 +1100 Subject: [PATCH 87/90] fix batteries dep --- lake-manifest.json | 4 ++-- lakefile.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 9bc0ccddd5573..4b7b5e02fb6e4 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "6336d636051c3fbe945113fae80553982f09a87e", + "rev": "9b9f4d0406d00baaae62d1a717b5aa854a2ae51d", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "lean-pr-testing-6095", + "inputRev": "nightly-testing", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 2bb0f3940e9cd..5c55a15c5c5e2 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,7 +7,7 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "lean-pr-testing-6095" +require "leanprover-community" / "batteries" @ git "nightly-testing" require "leanprover-community" / "Qq" @ git "nightly-testing" require "leanprover-community" / "aesop" @ git "nightly-testing" require "leanprover-community" / "proofwidgets" @ git "v0.0.46" From 7453073758e9f93ad01a831df232e181906158fb Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:56:54 +1100 Subject: [PATCH 88/90] eraseIdx! --- Mathlib/Tactic/Linarith/Frontend.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index 04b1e07d1ac97..552a10b4b0159 100644 --- a/Mathlib/Tactic/Linarith/Frontend.lean +++ b/Mathlib/Tactic/Linarith/Frontend.lean @@ -277,7 +277,7 @@ def runLinarith (cfg : LinarithConfig) (prefType : Option Expr) (g : MVarId) if let some t := prefType then let (i, vs) ← hyp_set.find t proveFalseByLinarith cfg.transparency cfg.oracle cfg.discharger g vs <|> - findLinarithContradiction cfg g ((hyp_set.eraseIdx i).toList.map (·.2)) + findLinarithContradiction cfg g ((hyp_set.eraseIdx! i).toList.map (·.2)) else findLinarithContradiction cfg g (hyp_set.toList.map (·.2)) let mut preprocessors := cfg.preprocessors if cfg.splitNe then From dadfebc218f792393d6c7a5a82799bea167124ed Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 08:57:24 +1100 Subject: [PATCH 89/90] eraseIdxIfInBounds --- Mathlib/Tactic/Linarith/Frontend.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index 552a10b4b0159..a63112c459aff 100644 --- a/Mathlib/Tactic/Linarith/Frontend.lean +++ b/Mathlib/Tactic/Linarith/Frontend.lean @@ -277,7 +277,7 @@ def runLinarith (cfg : LinarithConfig) (prefType : Option Expr) (g : MVarId) if let some t := prefType then let (i, vs) ← hyp_set.find t proveFalseByLinarith cfg.transparency cfg.oracle cfg.discharger g vs <|> - findLinarithContradiction cfg g ((hyp_set.eraseIdx! i).toList.map (·.2)) + findLinarithContradiction cfg g ((hyp_set.eraseIdxIfInBounds i).toList.map (·.2)) else findLinarithContradiction cfg g (hyp_set.toList.map (·.2)) let mut preprocessors := cfg.preprocessors if cfg.splitNe then From 33e7503de887ec81d3b24de62cfce9440ffa55a8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 09:00:46 +1100 Subject: [PATCH 90/90] deprecation --- Mathlib/Testing/Plausible/Functions.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Testing/Plausible/Functions.lean b/Mathlib/Testing/Plausible/Functions.lean index 9a83323198f47..080db229f9115 100644 --- a/Mathlib/Testing/Plausible/Functions.lean +++ b/Mathlib/Testing/Plausible/Functions.lean @@ -202,7 +202,7 @@ theorem List.applyId_zip_eq [DecidableEq α] {xs ys : List α} (h₀ : List.Nodu simp only [getElem?_cons_succ, zip_cons_cons, applyId_cons] at h₂ ⊢ rw [if_neg] · apply xs_ih <;> solve_by_elim [Nat.succ.inj] - · apply h₀; apply List.getElem?_mem h₂ + · apply h₀; apply List.mem_of_getElem? h₂ theorem applyId_mem_iff [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs) (h₁ : xs ~ ys) (x : α) : List.applyId.{u} (xs.zip ys) x ∈ ys ↔ x ∈ xs := by